From 9eb173e4440833e8e1f0a71d73a52986d61fd408 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 2 Jan 2025 09:10:27 +0000 Subject: [PATCH 001/100] fix: ignore `no_index` around `OfNat.ofNat` in `norm_cast` (#6438) This PR ensures `norm_cast` doesn't fail to act in the presence of `no_index` annotations While leanprover/lean4#2867 exists, it is necessary to put `no_index` around `OfNat.ofNat` in simp lemmas. This results in extra `Expr.mdata` nodes, which must be removed before checking for `ofNat` numerals. --- src/Lean/Elab/Tactic/NormCast.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/Lean/Elab/Tactic/NormCast.lean b/src/Lean/Elab/Tactic/NormCast.lean index 57a8d9191d9c..da3e09658a77 100644 --- a/src/Lean/Elab/Tactic/NormCast.lean +++ b/src/Lean/Elab/Tactic/NormCast.lean @@ -63,7 +63,7 @@ def isNumeral? (e : Expr) : Option (Expr × Nat) := if e.isConstOf ``Nat.zero then (mkConst ``Nat, 0) else if let Expr.app (Expr.app (Expr.app (Expr.const ``OfNat.ofNat ..) α ..) - (Expr.lit (Literal.natVal n) ..) ..) .. := e then + (Expr.lit (Literal.natVal n) ..) ..) .. := e.consumeMData then some (α, n) else none From 7d0c0d4d9254f1ad9eaa7428efddb2bc740dcd9a Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Thu, 2 Jan 2025 10:39:18 +0100 Subject: [PATCH 002/100] feat: partial_fixpoint: theory (#6477) This PR adds the necessary domain theory that backs the `partial_fixpoint` feature. Part of #6355. --- src/Init/Internal.lean | 13 + src/Init/Internal/Order.lean | 7 + src/Init/Internal/Order/Basic.lean | 693 +++++++++++++++++++++++++++++ 3 files changed, 713 insertions(+) create mode 100644 src/Init/Internal.lean create mode 100644 src/Init/Internal/Order.lean create mode 100644 src/Init/Internal/Order/Basic.lean diff --git a/src/Init/Internal.lean b/src/Init/Internal.lean new file mode 100644 index 000000000000..2bf2266b69a9 --- /dev/null +++ b/src/Init/Internal.lean @@ -0,0 +1,13 @@ +/- +Copyright (c) 2024 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joachim Breitner +-/ +prelude +import Init.Internal.Order + +/-! +This directory is used for components of the standard library that are either considered +implementation details or not yet ready for public consumption, and that should be available +without explicit import (in contrast to `Std.Internal`) +-/ diff --git a/src/Init/Internal/Order.lean b/src/Init/Internal/Order.lean new file mode 100644 index 000000000000..e30815898646 --- /dev/null +++ b/src/Init/Internal/Order.lean @@ -0,0 +1,7 @@ +/- +Copyright (c) 2024 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joachim Breitner +-/ +prelude +import Init.Internal.Order.Basic diff --git a/src/Init/Internal/Order/Basic.lean b/src/Init/Internal/Order/Basic.lean new file mode 100644 index 000000000000..e94bc3fecc29 --- /dev/null +++ b/src/Init/Internal/Order/Basic.lean @@ -0,0 +1,693 @@ +/- +Copyright (c) 2024 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joachim Breitner +-/ +prelude + +import Init.ByCases +import Init.RCases + +/-! +This module contains some basic definitions and results from domain theory, intended to be used as +the underlying construction of the `partial_fixpoint` feature. It is not meant to be used as a +general purpose library for domain theory, but can be of interest to users who want to extend +the `partial_fixpoint` machinery (e.g. mark more functions as monotone or register more monads). + +This follows the corresponding +[Isabelle development](https://isabelle.in.tum.de/library/HOL/HOL/Partial_Function.html), as also +described in [Alexander Krauss: Recursive Definitions of Monadic Functions](https://www21.in.tum.de/~krauss/papers/mrec.pdf). +-/ + +universe u v w + +namespace Lean.Order + +/-- +A partial order is a reflexive, transitive and antisymmetric relation. + +This is intended to be used in the construction of `partial_fixpoint`, and not meant to be used otherwise. +-/ +class PartialOrder (α : Sort u) where + /-- + A “less-or-equal-to” or “approximates” relation. + + This is intended to be used in the construction of `partial_fixpoint`, and not meant to be used otherwise. + -/ + rel : α → α → Prop + rel_refl : ∀ {x}, rel x x + rel_trans : ∀ {x y z}, rel x y → rel y z → rel x z + rel_antisymm : ∀ {x y}, rel x y → rel y x → x = y + +@[inherit_doc] scoped infix:50 " ⊑ " => PartialOrder.rel + +section PartialOrder + +variable {α : Sort u} [PartialOrder α] + +theorem PartialOrder.rel_of_eq {x y : α} (h : x = y) : x ⊑ y := by cases h; apply rel_refl + +/-- +A chain is a totally ordered set (representing a set as a predicate). + +This is intended to be used in the construction of `partial_fixpoint`, and not meant to be used otherwise. +-/ +def chain (c : α → Prop) : Prop := ∀ x y , c x → c y → x ⊑ y ∨ y ⊑ x + +end PartialOrder + +section CCPO + +/-- +A chain-complete partial order (CCPO) is a partial order where every chain a least upper bound. + +This is intended to be used in the construction of `partial_fixpoint`, and not meant to be used otherwise. +-/ +class CCPO (α : Sort u) extends PartialOrder α where + /-- + The least upper bound of a chain. + + This is intended to be used in the construction of `partial_fixpoint`, and not meant to be used otherwise. + -/ + csup : (α → Prop) → α + csup_spec {c : α → Prop} (hc : chain c) : csup c ⊑ x ↔ (∀ y, c y → y ⊑ x) + +open PartialOrder CCPO + +variable {α : Sort u} [CCPO α] + +theorem csup_le {c : α → Prop} (hchain : chain c) : (∀ y, c y → y ⊑ x) → csup c ⊑ x := + (csup_spec hchain).mpr + +theorem le_csup {c : α → Prop} (hchain : chain c) {y : α} (hy : c y) : y ⊑ csup c := + (csup_spec hchain).mp rel_refl y hy + +/-- +The bottom element is the least upper bound of the empty chain. + +This is intended to be used in the construction of `partial_fixpoint`, and not meant to be used otherwise. +-/ +def bot : α := csup (fun _ => False) + +scoped notation "⊥" => bot + +theorem bot_le (x : α) : ⊥ ⊑ x := by + apply csup_le + · intro x y hx hy; contradiction + · intro x hx; contradiction + +end CCPO + +section monotone + +variable {α : Sort u} [PartialOrder α] +variable {β : Sort v} [PartialOrder β] + +/-- +A function is monotone if if it maps related elements to releated elements. + +This is intended to be used in the construction of `partial_fixpoint`, and not meant to be used otherwise. +-/ +def monotone (f : α → β) : Prop := ∀ x y, x ⊑ y → f x ⊑ f y + +theorem monotone_const (c : β) : monotone (fun (_ : α) => c) := + fun _ _ _ => PartialOrder.rel_refl + +theorem monotone_id : monotone (fun (x : α) => x) := + fun _ _ hxy => hxy + +theorem monotone_compose + {γ : Sort w} [PartialOrder γ] + {f : α → β} {g : β → γ} + (hf : monotone f) (hg : monotone g) : + monotone (fun x => g (f x)) := fun _ _ hxy => hg _ _ (hf _ _ hxy) + +end monotone + +section admissibility + +variable {α : Sort u} [CCPO α] + +open PartialOrder CCPO + +/-- +A predicate is admissable if it can be transferred from the elements of a chain to the chains least +upper bound. Such predicates can be used in fixpoint induction. + +This definition implies `P ⊥`. Sometimes (e.g. in Isabelle) the empty chain is excluded +from this definition, and `P ⊥` is a separate condition of the induction predicate. + +This is intended to be used in the construction of `partial_fixpoint`, and not meant to be used otherwise. +-/ +def admissible (P : α → Prop) := + ∀ (c : α → Prop), chain c → (∀ x, c x → P x) → P (csup c) + +theorem admissible_const_true : admissible (fun (_ : α) => True) := + fun _ _ _ => trivial + +theorem admissible_and (P Q : α → Prop) + (hadm₁ : admissible P) (hadm₂ : admissible Q) : admissible (fun x => P x ∧ Q x) := + fun c hchain h => + ⟨ hadm₁ c hchain fun x hx => (h x hx).1, + hadm₂ c hchain fun x hx => (h x hx).2⟩ + +theorem chain_conj (c P : α → Prop) (hchain : chain c) : chain (fun x => c x ∧ P x) := by + intro x y ⟨hcx, _⟩ ⟨hcy, _⟩ + exact hchain x y hcx hcy + +theorem csup_conj (c P : α → Prop) (hchain : chain c) (h : ∀ x, c x → ∃ y, c y ∧ x ⊑ y ∧ P y) : + csup c = csup (fun x => c x ∧ P x) := by + apply rel_antisymm + · apply csup_le hchain + intro x hcx + obtain ⟨y, hcy, hxy, hPy⟩ := h x hcx + apply rel_trans hxy; clear x hcx hxy + apply le_csup (chain_conj _ _ hchain) ⟨hcy, hPy⟩ + · apply csup_le (chain_conj _ _ hchain) + intro x ⟨hcx, hPx⟩ + apply le_csup hchain hcx + +theorem admissible_or (P Q : α → Prop) + (hadm₁ : admissible P) (hadm₂ : admissible Q) : admissible (fun x => P x ∨ Q x) := by + intro c hchain h + have : (∀ x, c x → ∃ y, c y ∧ x ⊑ y ∧ P y) ∨ (∀ x, c x → ∃ y, c y ∧ x ⊑ y ∧ Q y) := by + open Classical in + apply Decidable.or_iff_not_imp_left.mpr + intro h' + simp only [not_forall, not_imp, not_exists, not_and] at h' + obtain ⟨x, hcx, hx⟩ := h' + intro y hcy + cases hchain x y hcx hcy with + | inl hxy => + refine ⟨y, hcy, rel_refl, ?_⟩ + cases h y hcy with + | inl hPy => exfalso; apply hx y hcy hxy hPy + | inr hQy => assumption + | inr hyx => + refine ⟨x, hcx, hyx , ?_⟩ + cases h x hcx with + | inl hPx => exfalso; apply hx x hcx rel_refl hPx + | inr hQx => assumption + cases this with + | inl hP => + left + rw [csup_conj (h := hP) (hchain := hchain)] + apply hadm₁ _ (chain_conj _ _ hchain) + intro x ⟨hcx, hPx⟩ + exact hPx + | inr hQ => + right + rw [csup_conj (h := hQ) (hchain := hchain)] + apply hadm₂ _ (chain_conj _ _ hchain) + intro x ⟨hcx, hQx⟩ + exact hQx + +def admissible_pi (P : α → β → Prop) + (hadm₁ : ∀ y, admissible (fun x => P x y)) : admissible (fun x => ∀ y, P x y) := + fun c hchain h y => hadm₁ y c hchain fun x hx => h x hx y + +end admissibility + +section fix + +open PartialOrder CCPO + +variable {α : Sort u} [CCPO α] + +variable {c : α → Prop} (hchain : chain c) + +/-- +The transfinite iteration of a function `f` is a set that is `⊥ ` and is closed under application +of `f` and `csup`. + +This is intended to be used in the construction of `partial_fixpoint`, and not meant to be used otherwise. +-/ +inductive iterates (f : α → α) : α → Prop where + | step : iterates f x → iterates f (f x) + | sup {c : α → Prop} (hc : chain c) (hi : ∀ x, c x → iterates f x) : iterates f (csup c) + +theorem chain_iterates {f : α → α} (hf : monotone f) : chain (iterates f) := by + intros x y hx hy + induction hx generalizing y + case step x hx ih => + induction hy + case step y hy _ => + cases ih y hy + · left; apply hf; assumption + · right; apply hf; assumption + case sup c hchain hi ih2 => + show f x ⊑ csup c ∨ csup c ⊑ f x + by_cases h : ∃ z, c z ∧ f x ⊑ z + · left + obtain ⟨z, hz, hfz⟩ := h + apply rel_trans hfz + apply le_csup hchain hz + · right + apply csup_le hchain _ + intro z hz + rw [not_exists] at h + specialize h z + rw [not_and] at h + specialize h hz + cases ih2 z hz + next => contradiction + next => assumption + case sup c hchain hi ih => + show rel (csup c) y ∨ rel y (csup c) + by_cases h : ∃ z, c z ∧ rel y z + · right + obtain ⟨z, hz, hfz⟩ := h + apply rel_trans hfz + apply le_csup hchain hz + · left + apply csup_le hchain _ + intro z hz + rw [not_exists] at h + specialize h z + rw [not_and] at h + specialize h hz + cases ih z hz y hy + next => assumption + next => contradiction + +theorem rel_f_of_iterates {f : α → α} (hf : monotone f) {x : α} (hx : iterates f x) : x ⊑ f x := by + induction hx + case step ih => + apply hf + assumption + case sup c hchain hi ih => + apply csup_le hchain + intro y hy + apply rel_trans (ih y hy) + apply hf + apply le_csup hchain hy + +set_option linter.unusedVariables false in +/-- +The least fixpoint of a monotone function is the least upper bound of its transfinite iteration. + +The `monotone f` assumption is not strictly necessarily for the definition, but without this the +definition is not very meaningful and it simplifies applying theorems like `fix_eq` if every use of +`fix` already has the monotonicty requirement. + +This is intended to be used in the construction of `partial_fixpoint`, and not meant to be used otherwise. +-/ +def fix (f : α → α) (hmono : monotone f) := csup (iterates f) + +/-- +The main fixpoint theorem for fixedpoints of monotone functions in chain-complete partial orders. + +This is intended to be used in the construction of `partial_fixpoint`, and not meant to be used otherwise. +-/ +theorem fix_eq {f : α → α} (hf : monotone f) : fix f hf = f (fix f hf) := by + apply rel_antisymm + · apply rel_f_of_iterates hf + apply iterates.sup (chain_iterates hf) + exact fun _ h => h + · apply le_csup (chain_iterates hf) + apply iterates.step + apply iterates.sup (chain_iterates hf) + intro y hy + exact hy + +/-- +The fixpoint induction theme: An admissible predicate holds for a least fixpoint if it is preserved +by the fixpoint's function. + +This is intended to be used in the construction of `partial_fixpoint`, and not meant to be used otherwise. +-/ +theorem fix_induct {f : α → α} (hf : monotone f) + (motive : α → Prop) (hadm: admissible motive) + (h : ∀ x, motive x → motive (f x)) : motive (fix f hf) := by + apply hadm _ (chain_iterates hf) + intro x hiterates + induction hiterates with + | @step x hiter ih => apply h x ih + | @sup c hchain hiter ih => apply hadm c hchain ih + +end fix + +section fun_order + +open PartialOrder + +variable {α : Sort u} +variable {β : α → Sort v} +variable {γ : Sort w} + +instance instOrderPi [∀ x, PartialOrder (β x)] : PartialOrder (∀ x, β x) where + rel f g := ∀ x, f x ⊑ g x + rel_refl _ := rel_refl + rel_trans hf hg x := rel_trans (hf x) (hg x) + rel_antisymm hf hg := funext (fun x => rel_antisymm (hf x) (hg x)) + +theorem monotone_of_monotone_apply [PartialOrder γ] [∀ x, PartialOrder (β x)] (f : γ → (∀ x, β x)) + (h : ∀ y, monotone (fun x => f x y)) : monotone f := + fun x y hxy z => h z x y hxy + +theorem monotone_apply [PartialOrder γ] [∀ x, PartialOrder (β x)] (a : α) (f : γ → ∀ x, β x) + (h : monotone f) : + monotone (fun x => f x a) := fun _ _ hfg => h _ _ hfg a + +theorem chain_apply [∀ x, PartialOrder (β x)] {c : (∀ x, β x) → Prop} (hc : chain c) (x : α) : + chain (fun y => ∃ f, c f ∧ f x = y) := by + intro _ _ ⟨f, hf, hfeq⟩ ⟨g, hg, hgeq⟩ + subst hfeq; subst hgeq + cases hc f g hf hg + next h => left; apply h x + next h => right; apply h x + +def fun_csup [∀ x, CCPO (β x)] (c : (∀ x, β x) → Prop) (x : α) := + CCPO.csup (fun y => ∃ f, c f ∧ f x = y) + +instance instCCPOPi [∀ x, CCPO (β x)] : CCPO (∀ x, β x) where + csup := fun_csup + csup_spec := by + intro f c hc + constructor + next => + intro hf g hg x + apply rel_trans _ (hf x); clear hf + apply le_csup (chain_apply hc x) + exact ⟨g, hg, rfl⟩ + next => + intro h x + apply csup_le (chain_apply hc x) + intro y ⟨z, hz, hyz⟩ + subst y + apply h z hz + +def admissible_apply [∀ x, CCPO (β x)] (P : ∀ x, β x → Prop) (x : α) + (hadm : admissible (P x)) : admissible (fun (f : ∀ x, β x) => P x (f x)) := by + intro c hchain h + apply hadm _ (chain_apply hchain x) + rintro _ ⟨f, hcf, rfl⟩ + apply h _ hcf + +def admissible_pi_apply [∀ x, CCPO (β x)] (P : ∀ x, β x → Prop) (hadm : ∀ x, admissible (P x)) : + admissible (fun (f : ∀ x, β x) => ∀ x, P x (f x)) := by + apply admissible_pi + intro + apply admissible_apply + apply hadm + +end fun_order + +section monotone_lemmas + +theorem monotone_letFun + {α : Sort u} {β : Sort v} {γ : Sort w} [PartialOrder α] [PartialOrder β] + (v : γ) (k : α → γ → β) + (hmono : ∀ y, monotone (fun x => k x y)) : + monotone fun (x : α) => letFun v (k x) := hmono v + +theorem monotone_ite + {α : Sort u} {β : Sort v} [PartialOrder α] [PartialOrder β] + (c : Prop) [Decidable c] + (k₁ : α → β) (k₂ : α → β) + (hmono₁ : monotone k₁) (hmono₂ : monotone k₂) : + monotone fun x => if c then k₁ x else k₂ x := by + split + · apply hmono₁ + · apply hmono₂ + +theorem monotone_dite + {α : Sort u} {β : Sort v} [PartialOrder α] [PartialOrder β] + (c : Prop) [Decidable c] + (k₁ : α → c → β) (k₂ : α → ¬ c → β) + (hmono₁ : monotone k₁) (hmono₂ : monotone k₂) : + monotone fun x => dite c (k₁ x) (k₂ x) := by + split + · apply monotone_apply _ _ hmono₁ + · apply monotone_apply _ _ hmono₂ + +end monotone_lemmas + +section pprod_order + +open PartialOrder + +variable {α : Sort u} +variable {β : Sort v} +variable {γ : Sort w} + +instance [PartialOrder α] [PartialOrder β] : PartialOrder (α ×' β) where + rel a b := a.1 ⊑ b.1 ∧ a.2 ⊑ b.2 + rel_refl := ⟨rel_refl, rel_refl⟩ + rel_trans ha hb := ⟨rel_trans ha.1 hb.1, rel_trans ha.2 hb.2⟩ + rel_antisymm := fun {a} {b} ha hb => by + cases a; cases b; + dsimp at * + rw [rel_antisymm ha.1 hb.1, rel_antisymm ha.2 hb.2] + +theorem monotone_pprod [PartialOrder α] [PartialOrder β] [PartialOrder γ] + {f : γ → α} {g : γ → β} (hf : monotone f) (hg : monotone g) : + monotone (fun x => PProd.mk (f x) (g x)) := + fun _ _ h12 => ⟨hf _ _ h12, hg _ _ h12⟩ + +theorem monotone_pprod_fst [PartialOrder α] [PartialOrder β] [PartialOrder γ] + {f : γ → α ×' β} (hf : monotone f) : monotone (fun x => (f x).1) := + fun _ _ h12 => (hf _ _ h12).1 + +theorem monotone_pprod_snd [PartialOrder α] [PartialOrder β] [PartialOrder γ] + {f : γ → α ×' β} (hf : monotone f) : monotone (fun x => (f x).2) := + fun _ _ h12 => (hf _ _ h12).2 + +def chain_pprod_fst [CCPO α] [CCPO β] (c : α ×' β → Prop) : α → Prop := fun a => ∃ b, c ⟨a, b⟩ +def chain_pprod_snd [CCPO α] [CCPO β] (c : α ×' β → Prop) : β → Prop := fun b => ∃ a, c ⟨a, b⟩ + +theorem chain.pprod_fst [CCPO α] [CCPO β] (c : α ×' β → Prop) (hchain : chain c) : + chain (chain_pprod_fst c) := by + intro a₁ a₂ ⟨b₁, h₁⟩ ⟨b₂, h₂⟩ + cases hchain ⟨a₁, b₁⟩ ⟨a₂, b₂⟩ h₁ h₂ + case inl h => left; exact h.1 + case inr h => right; exact h.1 + +theorem chain.pprod_snd [CCPO α] [CCPO β] (c : α ×' β → Prop) (hchain : chain c) : + chain (chain_pprod_snd c) := by + intro b₁ b₂ ⟨a₁, h₁⟩ ⟨a₂, h₂⟩ + cases hchain ⟨a₁, b₁⟩ ⟨a₂, b₂⟩ h₁ h₂ + case inl h => left; exact h.2 + case inr h => right; exact h.2 + +instance [CCPO α] [CCPO β] : CCPO (α ×' β) where + csup c := ⟨CCPO.csup (chain_pprod_fst c), CCPO.csup (chain_pprod_snd c)⟩ + csup_spec := by + intro ⟨a, b⟩ c hchain + dsimp + constructor + next => + intro ⟨h₁, h₂⟩ ⟨a', b'⟩ cab + constructor <;> dsimp at * + · apply rel_trans ?_ h₁ + apply le_csup hchain.pprod_fst + exact ⟨b', cab⟩ + · apply rel_trans ?_ h₂ + apply le_csup hchain.pprod_snd + exact ⟨a', cab⟩ + next => + intro h + constructor <;> dsimp + · apply csup_le hchain.pprod_fst + intro a' ⟨b', hcab⟩ + apply (h _ hcab).1 + · apply csup_le hchain.pprod_snd + intro b' ⟨a', hcab⟩ + apply (h _ hcab).2 + +theorem admissible_pprod_fst {α : Sort u} {β : Sort v} [CCPO α] [CCPO β] (P : α → Prop) + (hadm : admissible P) : admissible (fun (x : α ×' β) => P x.1) := by + intro c hchain h + apply hadm _ hchain.pprod_fst + intro x ⟨y, hxy⟩ + apply h ⟨x,y⟩ hxy + +theorem admissible_pprod_snd {α : Sort u} {β : Sort v} [CCPO α] [CCPO β] (P : β → Prop) + (hadm : admissible P) : admissible (fun (x : α ×' β) => P x.2) := by + intro c hchain h + apply hadm _ hchain.pprod_snd + intro y ⟨x, hxy⟩ + apply h ⟨x,y⟩ hxy + +end pprod_order + +section flat_order + +variable {α : Sort u} + +set_option linter.unusedVariables false in +/-- +`FlatOrder b` wraps the type `α` with the flat partial order generated by `∀ x, b ⊑ x`. + +This is intended to be used in the construction of `partial_fixpoint`, and not meant to be used otherwise. +-/ +def FlatOrder {α : Sort u} (b : α) := α + +variable {b : α} + +/-- +The flat partial order generated by `∀ x, b ⊑ x`. + +This is intended to be used in the construction of `partial_fixpoint`, and not meant to be used otherwise. +-/ +inductive FlatOrder.rel : (x y : FlatOrder b) → Prop where + | bot : rel b x + | refl : rel x x + +instance FlatOrder.instOrder : PartialOrder (FlatOrder b) where + rel := rel + rel_refl := .refl + rel_trans {x y z : α} (hxy : rel x y) (hyz : rel y z) := by + cases hxy <;> cases hyz <;> constructor + rel_antisymm {x y : α} (hxy : rel x y) (hyz : rel y x) : x = y := by + cases hxy <;> cases hyz <;> constructor + +open Classical in +private theorem Classical.some_spec₂ {α : Sort _} {p : α → Prop} {h : ∃ a, p a} (q : α → Prop) + (hpq : ∀ a, p a → q a) : q (choose h) := hpq _ <| choose_spec _ + +noncomputable def flat_csup (c : FlatOrder b → Prop) : FlatOrder b := by + by_cases h : ∃ (x : FlatOrder b), c x ∧ x ≠ b + · exact Classical.choose h + · exact b + +noncomputable instance FlatOrder.instCCPO : CCPO (FlatOrder b) where + csup := flat_csup + csup_spec := by + intro x c hc + unfold flat_csup + split + next hex => + apply Classical.some_spec₂ (q := (· ⊑ x ↔ (∀ y, c y → y ⊑ x))) + clear hex + intro z ⟨hz, hnb⟩ + constructor + · intro h y hy + apply PartialOrder.rel_trans _ h; clear h + cases hc y z hy hz + next => assumption + next h => + cases h + · contradiction + · constructor + · intro h + cases h z hz + · contradiction + · constructor + next hnotex => + constructor + · intro h y hy; clear h + suffices y = b by rw [this]; exact rel.bot + rw [not_exists] at hnotex + specialize hnotex y + rw [not_and] at hnotex + specialize hnotex hy + rw [@Classical.not_not] at hnotex + assumption + · intro; exact rel.bot + +theorem admissible_flatOrder (P : FlatOrder b → Prop) (hnot : P b) : admissible P := by + intro c hchain h + by_cases h' : ∃ (x : FlatOrder b), c x ∧ x ≠ b + · simp [CCPO.csup, flat_csup, h'] + apply Classical.some_spec₂ (q := (P ·)) + intro x ⟨hcx, hneb⟩ + apply h x hcx + · simp [CCPO.csup, flat_csup, h', hnot] + +end flat_order + +section mono_bind + +/-- +The class `MonoBind m` indicates that every `m α` has a `PartialOrder`, and that the bind operation +on `m` is monotone in both arguments with regard to that order. + +This is intended to be used in the construction of `partial_fixpoint`, and not meant to be used otherwise. +-/ +class MonoBind (m : Type u → Type v) [Bind m] [∀ α, PartialOrder (m α)] where + bind_mono_left {a₁ a₂ : m α} {f : α → m b} (h : a₁ ⊑ a₂) : a₁ >>= f ⊑ a₂ >>= f + bind_mono_right {a : m α} {f₁ f₂ : α → m b} (h : ∀ x, f₁ x ⊑ f₂ x) : a >>= f₁ ⊑ a >>= f₂ + +theorem monotone_bind + (m : Type u → Type v) [Bind m] [∀ α, PartialOrder (m α)] [MonoBind m] + {α β : Type u} + {γ : Type w} [PartialOrder γ] + (f : γ → m α) (g : γ → α → m β) + (hmono₁ : monotone f) + (hmono₂ : monotone g) : + monotone (fun (x : γ) => f x >>= g x) := by + intro x₁ x₂ hx₁₂ + apply PartialOrder.rel_trans + · apply MonoBind.bind_mono_left (hmono₁ _ _ hx₁₂) + · apply MonoBind.bind_mono_right (fun y => monotone_apply y _ hmono₂ _ _ hx₁₂) + +instance : PartialOrder (Option α) := inferInstanceAs (PartialOrder (FlatOrder none)) +noncomputable instance : CCPO (Option α) := inferInstanceAs (CCPO (FlatOrder none)) +noncomputable instance : MonoBind Option where + bind_mono_left h := by + cases h + · exact FlatOrder.rel.bot + · exact FlatOrder.rel.refl + bind_mono_right h := by + cases ‹Option _› + · exact FlatOrder.rel.refl + · exact h _ + +theorem admissible_eq_some (P : Prop) (y : α) : + admissible (fun (x : Option α) => x = some y → P) := by + apply admissible_flatOrder; simp + +instance [Monad m] [inst : ∀ α, PartialOrder (m α)] : PartialOrder (ExceptT ε m α) := inst _ +instance [Monad m] [∀ α, PartialOrder (m α)] [inst : ∀ α, CCPO (m α)] : CCPO (ExceptT ε m α) := inst _ +instance [Monad m] [∀ α, PartialOrder (m α)] [∀ α, CCPO (m α)] [MonoBind m] : MonoBind (ExceptT ε m) where + bind_mono_left h₁₂ := by + apply MonoBind.bind_mono_left (m := m) + exact h₁₂ + bind_mono_right h₁₂ := by + apply MonoBind.bind_mono_right (m := m) + intro x + cases x + · apply PartialOrder.rel_refl + · apply h₁₂ + +end mono_bind + +namespace Example + +def findF (P : Nat → Bool) (rec : Nat → Option Nat) (x : Nat) : Option Nat := + if P x then + some x + else + rec (x + 1) + +noncomputable def find (P : Nat → Bool) : Nat → Option Nat := fix (findF P) <| by + unfold findF + apply monotone_of_monotone_apply + intro n + split + · apply monotone_const + · apply monotone_apply + apply monotone_id + +theorem find_eq : find P = findF P (find P) := fix_eq .. + +theorem find_spec : ∀ n m, find P n = some m → n ≤ m ∧ P m := by + unfold find + refine fix_induct (motive := fun (f : Nat → Option Nat) => ∀ n m, f n = some m → n ≤ m ∧ P m) _ ?hadm ?hstep + case hadm => + -- apply admissible_pi_apply does not work well, hard to infer everything + exact admissible_pi_apply _ (fun n => admissible_pi _ (fun m => admissible_eq_some _ m)) + case hstep => + intro f ih n m heq + simp only [findF] at heq + split at heq + · simp_all + · obtain ⟨ih1, ih2⟩ := ih _ _ heq + constructor + · exact Nat.le_trans (Nat.le_add_right _ _ ) ih1 + · exact ih2 + +end Example + +end Lean.Order From e9f069146c93bd3ba4b94863ff47e42b6d22b8b8 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Thu, 2 Jan 2025 12:07:05 +0100 Subject: [PATCH 003/100] feat: partial_fixpoint: monotonicity tactic (#6506) This PR adds the `monotonicity` tactic, intended to be used inside the `partial_fixpoint` feature. Part of #6355. --- src/Init.lean | 1 + src/Init/Internal/Order.lean | 1 + src/Init/Internal/Order/Tactic.lean | 20 ++ src/Lean/Elab/Tactic.lean | 1 + src/Lean/Elab/Tactic/Monotonicity.lean | 223 ++++++++++++++++++ .../run/partial_fixpoint_monotonicity.lean | 29 +++ 6 files changed, 275 insertions(+) create mode 100644 src/Init/Internal/Order/Tactic.lean create mode 100644 src/Lean/Elab/Tactic/Monotonicity.lean create mode 100644 tests/lean/run/partial_fixpoint_monotonicity.lean diff --git a/src/Init.lean b/src/Init.lean index 568452373ed2..4a387d46e3a3 100644 --- a/src/Init.lean +++ b/src/Init.lean @@ -37,3 +37,4 @@ import Init.MacroTrace import Init.Grind import Init.While import Init.Syntax +import Init.Internal diff --git a/src/Init/Internal/Order.lean b/src/Init/Internal/Order.lean index e30815898646..e5389b451690 100644 --- a/src/Init/Internal/Order.lean +++ b/src/Init/Internal/Order.lean @@ -5,3 +5,4 @@ Authors: Joachim Breitner -/ prelude import Init.Internal.Order.Basic +import Init.Internal.Order.Tactic diff --git a/src/Init/Internal/Order/Tactic.lean b/src/Init/Internal/Order/Tactic.lean new file mode 100644 index 000000000000..574990e87791 --- /dev/null +++ b/src/Init/Internal/Order/Tactic.lean @@ -0,0 +1,20 @@ +/- +Copyright (c) 2024 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joachim Breitner +-/ + +prelude +import Init.Notation + +namespace Lean.Order +/-- +`monotonicity` performs one compositional step solving `monotone` goals, +using lemma tagged with `@[partial_fixpoint_monotone]`. + +This tactic is mostly used internally by lean in `partial_fixpoint` definitions, but +can be useful on its own for debugging or when proving new `@[partial_fixpoint_monotone]` lemmas. +-/ +scoped syntax (name := monotonicity) "monotonicity" : tactic + +end Lean.Order diff --git a/src/Lean/Elab/Tactic.lean b/src/Lean/Elab/Tactic.lean index 8eec5d47cd17..add9aea409c2 100644 --- a/src/Lean/Elab/Tactic.lean +++ b/src/Lean/Elab/Tactic.lean @@ -45,3 +45,4 @@ import Lean.Elab.Tactic.BVDecide import Lean.Elab.Tactic.BoolToPropSimps import Lean.Elab.Tactic.Classical import Lean.Elab.Tactic.Grind +import Lean.Elab.Tactic.Monotonicity diff --git a/src/Lean/Elab/Tactic/Monotonicity.lean b/src/Lean/Elab/Tactic/Monotonicity.lean new file mode 100644 index 000000000000..50f02f63cbb6 --- /dev/null +++ b/src/Lean/Elab/Tactic/Monotonicity.lean @@ -0,0 +1,223 @@ +/- +Copyright (c) 2024 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joachim Breitner +-/ +prelude +import Lean.Meta.Tactic.Split +import Lean.Elab.RecAppSyntax +import Lean.Elab.Tactic.Basic +import Init.Internal.Order + +namespace Lean.Meta.Monotonicity + +open Lean Meta +open Lean.Order + +partial def headBetaUnderLambda (f : Expr) : Expr := Id.run do + let mut f := f.headBeta + if f.isLambda then + while f.bindingBody!.isHeadBetaTarget do + f := f.updateLambda! f.bindingInfo! f.bindingDomain! f.bindingBody!.headBeta + return f + + +/-- Environment extensions for monotonicity lemmas -/ +builtin_initialize monotoneExt : + SimpleScopedEnvExtension (Name × Array DiscrTree.Key) (DiscrTree Name) ← + registerSimpleScopedEnvExtension { + addEntry := fun dt (n, ks) => dt.insertCore ks n + initial := {} + } + +builtin_initialize registerBuiltinAttribute { + name := `partial_fixpoint_monotone + descr := "monotonicity theorem" + add := fun decl _ kind => MetaM.run' do + let declTy := (← getConstInfo decl).type + let (xs, _, targetTy) ← withReducible <| forallMetaTelescopeReducing declTy + let_expr monotone α inst_α β inst_β f := targetTy | + throwError "@[partial_fixpoint_monotone] attribute only applies to lemmas proving {.ofConstName ``monotone}" + let f := f.headBeta + let f ← if f.isLambda then pure f else etaExpand f + let f := headBetaUnderLambda f + lambdaBoundedTelescope f 1 fun _ e => do + let key ← withReducible <| DiscrTree.mkPath e + monotoneExt.add (decl, key) kind +} + +/-- +Finds tagged monotonicity theorems of the form `monotone (fun x => e)`. +-/ +def findMonoThms (e : Expr) : MetaM (Array Name) := do + (monotoneExt.getState (← getEnv)).getMatch e + +private def defaultFailK (f : Expr) (monoThms : Array Name) : MetaM α := + let extraMsg := if monoThms.isEmpty then m!"" else + m!"Tried to apply {.andList (monoThms.toList.map (m!"'{·}'"))}, but failed." + throwError "Failed to prove monotonicity of:{indentExpr f}\n{extraMsg}" + +private def applyConst (goal : MVarId) (name : Name) : MetaM (List MVarId) := do + mapError (f := (m!"Could not apply {.ofConstName name}:{indentD ·}")) do + goal.applyConst name (cfg := { synthAssignedInstances := false}) + +/-- +Base case for solveMonoStep: Handles goals of the form +``` +monotone (fun f => f.1.2 x y) +``` + +It's tricky to solve them compositionally from the outside in, so here we construct the proof +from the inside out. +-/ +partial def solveMonoCall (α inst_α : Expr) (e : Expr) : MetaM (Option Expr) := do + if e.isApp && !e.appArg!.hasLooseBVars then + let some hmono ← solveMonoCall α inst_α e.appFn! | return none + let hmonoType ← inferType hmono + let_expr monotone _ _ _ inst _ := hmonoType | throwError "solveMonoCall {e}: unexpected type {hmonoType}" + let some inst ← whnfUntil inst ``instOrderPi | throwError "solveMonoCall {e}: unexpected instance {inst}" + let_expr instOrderPi γ δ inst ← inst | throwError "solveMonoCall {e}: whnfUntil failed?{indentExpr inst}" + return ← mkAppOptM ``monotone_apply #[γ, δ, α, inst_α, inst, e.appArg!, none, hmono] + + if e.isProj then + let some hmono ← solveMonoCall α inst_α e.projExpr! | return none + let hmonoType ← inferType hmono + let_expr monotone _ _ _ inst _ := hmonoType | throwError "solveMonoCall {e}: unexpected type {hmonoType}" + let some inst ← whnfUntil inst ``instPartialOrderPProd | throwError "solveMonoCall {e}: unexpected instance {inst}" + let_expr instPartialOrderPProd β γ inst_β inst_γ ← inst | throwError "solveMonoCall {e}: whnfUntil failed?{indentExpr inst}" + let n := if e.projIdx! == 0 then ``monotone_pprod_fst else ``monotone_pprod_snd + return ← mkAppOptM n #[β, γ, α, inst_β, inst_γ, inst_α, none, hmono] + + if e == .bvar 0 then + let hmono ← mkAppOptM ``monotone_id #[α, inst_α] + return some hmono + + return none + + +def solveMonoStep (failK : ∀ {α}, Expr → Array Name → MetaM α := @defaultFailK) (goal : MVarId) : MetaM (List MVarId) := + goal.withContext do + trace[Elab.Tactic.monotonicity] "monotonicity at\n{goal}" + let type ← goal.getType + if type.isForall then + let (_, goal) ← goal.intro1P + return [goal] + + match_expr type with + | monotone α inst_α β inst_β f => + -- Ensure f is not headed by a redex and headed by at least one lambda, and clean some + -- redexes left by some of the lemmas we tend to apply + let f ← instantiateMVars f + let f := f.headBeta + let f ← if f.isLambda then pure f else etaExpand f + let f := headBetaUnderLambda f + let e := f.bindingBody! + + -- No recursive calls left + if !e.hasLooseBVars then + return ← applyConst goal ``monotone_const + + -- NB: `e` is now an open term. + + -- Look through mdata + if e.isMData then + let f' := f.updateLambdaE! f.bindingDomain! e.mdataExpr! + let goal' ← mkFreshExprSyntheticOpaqueMVar (mkApp type.appFn! f') + goal.assign goal' + return [goal'.mvarId!] + + -- Float letE to the environment + if let .letE n t v b _nonDep := e then + if t.hasLooseBVars || v.hasLooseBVars then + failK f #[] + let goal' ← withLetDecl n t v fun x => do + let b' := f.updateLambdaE! f.bindingDomain! (b.instantiate1 x) + let goal' ← mkFreshExprSyntheticOpaqueMVar (mkApp type.appFn! b') + goal.assign (← mkLetFVars #[x] goal') + pure goal' + return [goal'.mvarId!] + + -- Float `letFun` to the environment. + -- `applyConst` tends to reduce the redex + match_expr e with + | letFun γ _ v b => + if γ.hasLooseBVars || v.hasLooseBVars then + failK f #[] + let b' := f.updateLambdaE! f.bindingDomain! b + let p ← mkAppOptM ``monotone_letFun #[α, β, γ, inst_α, inst_β, v, b'] + let new_goals ← mapError (f := (m!"Could not apply {p}:{indentD ·}")) do + goal.apply p + let [new_goal] := new_goals + | throwError "Unexpected number of goals after {.ofConstName ``monotone_letFun}." + let (_, new_goal) ← + if b.isLambda then + new_goal.intro b.bindingName! + else + new_goal.intro1 + return [new_goal] + | _ => pure () + + -- Handle lambdas, preserving the name of the binder + if e.isLambda then + let [new_goal] ← applyConst goal ``monotone_of_monotone_apply + | throwError "Unexpected number of goals after {.ofConstName ``monotone_of_monotone_apply}." + let (_, new_goal) ← new_goal.intro e.bindingName! + return [new_goal] + + -- A recursive call directly here + if e.isBVar then + return ← applyConst goal ``monotone_id + + -- A recursive call + if let some hmono ← solveMonoCall α inst_α e then + trace[Elab.Tactic.monotonicity] "Found recursive call {e}:{indentExpr hmono}" + unless ← goal.checkedAssign hmono do + trace[Elab.Tactic.monotonicity] "Failed to assign {hmono} : {← inferType hmono} to goal" + failK f #[] + return [] + + let monoThms ← withLocalDeclD `f f.bindingDomain! fun f => + -- The discrimination tree does not like open terms + findMonoThms (e.instantiate1 f) + trace[Elab.Tactic.monotonicity] "Found monoThms: {monoThms.map MessageData.ofConstName}" + for monoThm in monoThms do + let new_goals? ← try + let new_goals ← applyConst goal monoThm + trace[Elab.Tactic.monotonicity] "Succeeded with {.ofConstName monoThm}" + pure (some new_goals) + catch e => + trace[Elab.Tactic.monotonicity] "{e.toMessageData}" + pure none + if let some new_goals := new_goals? then + return new_goals + + -- Split match-expressions + if let some info := isMatcherAppCore? (← getEnv) e then + let candidate ← id do + let args := e.getAppArgs + for i in [info.getFirstDiscrPos : info.getFirstDiscrPos + info.numDiscrs] do + if args[i]!.hasLooseBVars then + return false + return true + if candidate then + -- We could be even more deliberate here and use the `lifter` lemmas + -- for the match statements instead of the `split` tactic. + -- For now using `splitMatch` works fine. + return ← Split.splitMatch goal e + + failK f monoThms + | _ => + throwError "Unexpected goal:{goal}" + +partial def solveMono (failK : ∀ {α}, Expr → Array Name → MetaM α := defaultFailK) (goal : MVarId) : MetaM Unit := do + let new_goals ← solveMonoStep failK goal + new_goals.forM (solveMono failK) + +open Elab Tactic in +@[builtin_tactic Lean.Order.monotonicity] +def evalMonotonicity : Tactic := fun _stx => + liftMetaTactic Lean.Meta.Monotonicity.solveMonoStep + +end Lean.Meta.Monotonicity + +builtin_initialize Lean.registerTraceClass `Elab.Tactic.monotonicity diff --git a/tests/lean/run/partial_fixpoint_monotonicity.lean b/tests/lean/run/partial_fixpoint_monotonicity.lean new file mode 100644 index 000000000000..b6dee1799c02 --- /dev/null +++ b/tests/lean/run/partial_fixpoint_monotonicity.lean @@ -0,0 +1,29 @@ +-- Tests the `monotonicity` tactic + +-- These can move to Init after a stage0 update +open Lean.Order in +attribute [partial_fixpoint_monotone] + monotone_ite + monotone_dite + monotone_bind + +/- +Should test that the tactic syntax is scoped, but cannot use #guard_msgs to catch “tactic not known” +errors, it seems: + +/-- +error: unsolved goals +⊢ True +-/ +#guard_msgs in +example : True := by monotonicity + +-/ + +open Lean.Order + +example : monotone (fun (f : Nat → Option Unit) => do {do f 1; f 2; f 3}) := by + repeat monotonicity + +example : monotone (fun (f : Option Unit) => do {do f; f; f}) := by + repeat monotonicity From 092449adb8488a8fcbbebf9fdfcec38deef1eb81 Mon Sep 17 00:00:00 2001 From: Lean stage0 autoupdater <> Date: Thu, 2 Jan 2025 12:29:37 +0000 Subject: [PATCH 004/100] chore: update stage0 --- stage0/stdlib/Init.c | 6 +- stage0/stdlib/Init/Grind/Lemmas.c | 6 +- stage0/stdlib/Init/Grind/Tactics.c | 191 +- stage0/stdlib/Init/Internal.c | 29 + stage0/stdlib/Init/Internal/Order.c | 33 + stage0/stdlib/Init/Internal/Order/Basic.c | 1126 + stage0/stdlib/Init/Internal/Order/Tactic.c | 120 + stage0/stdlib/Lean/Data/PersistentHashSet.c | 96 + stage0/stdlib/Lean/Elab/Calc.c | 6777 ++++-- stage0/stdlib/Lean/Elab/CheckTactic.c | 4107 ++-- stage0/stdlib/Lean/Elab/Deriving.c | 6 +- stage0/stdlib/Lean/Elab/Deriving/ToExpr.c | 16344 ++++++++++++++ stage0/stdlib/Lean/Elab/Structure.c | 2783 ++- stage0/stdlib/Lean/Elab/Tactic.c | 6 +- stage0/stdlib/Lean/Elab/Tactic/Grind.c | 1733 +- stage0/stdlib/Lean/Elab/Tactic/Monotonicity.c | 12007 ++++++++++ stage0/stdlib/Lean/Elab/Tactic/NormCast.c | 185 +- stage0/stdlib/Lean/Meta/Check.c | 2212 +- stage0/stdlib/Lean/Meta/Tactic/Grind.c | 968 +- stage0/stdlib/Lean/Meta/Tactic/Grind/Canon.c | 6402 +++--- stage0/stdlib/Lean/Meta/Tactic/Grind/Core.c | 3663 +-- stage0/stdlib/Lean/Meta/Tactic/Grind/EMatch.c | 14724 ++++++++++++ .../Lean/Meta/Tactic/Grind/EMatchTheorem.c | 18515 ++++++++++++++++ .../Lean/Meta/Tactic/Grind/ForallProp.c | 574 + .../Lean/Meta/Tactic/Grind/Internalize.c | 10018 ++++++--- stage0/stdlib/Lean/Meta/Tactic/Grind/Intro.c | 4860 ++++ stage0/stdlib/Lean/Meta/Tactic/Grind/Inv.c | 1141 +- stage0/stdlib/Lean/Meta/Tactic/Grind/Main.c | 3559 +++ .../Lean/Meta/Tactic/Grind/MarkNestedProofs.c | 315 +- stage0/stdlib/Lean/Meta/Tactic/Grind/PP.c | 429 + .../Lean/Meta/Tactic/Grind/Preprocessor.c | 8623 ------- stage0/stdlib/Lean/Meta/Tactic/Grind/Proj.c | 457 +- stage0/stdlib/Lean/Meta/Tactic/Grind/Proof.c | 1586 +- .../stdlib/Lean/Meta/Tactic/Grind/Propagate.c | 75 +- stage0/stdlib/Lean/Meta/Tactic/Grind/Run.c | 1827 -- stage0/stdlib/Lean/Meta/Tactic/Grind/Simp.c | 1433 +- .../Lean/Meta/Tactic/Grind/TheoremPatterns.c | 8928 -------- stage0/stdlib/Lean/Meta/Tactic/Grind/Types.c | 8820 ++++++-- stage0/stdlib/Lean/ToExpr.c | 2 +- 39 files changed, 105845 insertions(+), 38841 deletions(-) create mode 100644 stage0/stdlib/Init/Internal.c create mode 100644 stage0/stdlib/Init/Internal/Order.c create mode 100644 stage0/stdlib/Init/Internal/Order/Basic.c create mode 100644 stage0/stdlib/Init/Internal/Order/Tactic.c create mode 100644 stage0/stdlib/Lean/Elab/Deriving/ToExpr.c create mode 100644 stage0/stdlib/Lean/Elab/Tactic/Monotonicity.c create mode 100644 stage0/stdlib/Lean/Meta/Tactic/Grind/EMatch.c create mode 100644 stage0/stdlib/Lean/Meta/Tactic/Grind/EMatchTheorem.c create mode 100644 stage0/stdlib/Lean/Meta/Tactic/Grind/ForallProp.c create mode 100644 stage0/stdlib/Lean/Meta/Tactic/Grind/Intro.c create mode 100644 stage0/stdlib/Lean/Meta/Tactic/Grind/Main.c delete mode 100644 stage0/stdlib/Lean/Meta/Tactic/Grind/Preprocessor.c delete mode 100644 stage0/stdlib/Lean/Meta/Tactic/Grind/Run.c delete mode 100644 stage0/stdlib/Lean/Meta/Tactic/Grind/TheoremPatterns.c diff --git a/stage0/stdlib/Init.c b/stage0/stdlib/Init.c index 85020397eb0a..e717eaeb2c53 100644 --- a/stage0/stdlib/Init.c +++ b/stage0/stdlib/Init.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Init -// Imports: Init.Prelude Init.Notation Init.Tactics Init.TacticsExtra Init.ByCases Init.RCases Init.Core Init.Control Init.Data.Basic Init.WF Init.WFTactics Init.Data Init.System Init.Util Init.Dynamic Init.ShareCommon Init.MetaTypes Init.Meta Init.NotationExtra Init.SimpLemmas Init.PropLemmas Init.Hints Init.Conv Init.Guard Init.Simproc Init.SizeOfLemmas Init.BinderPredicates Init.Ext Init.Omega Init.MacroTrace Init.Grind Init.While Init.Syntax +// Imports: Init.Prelude Init.Notation Init.Tactics Init.TacticsExtra Init.ByCases Init.RCases Init.Core Init.Control Init.Data.Basic Init.WF Init.WFTactics Init.Data Init.System Init.Util Init.Dynamic Init.ShareCommon Init.MetaTypes Init.Meta Init.NotationExtra Init.SimpLemmas Init.PropLemmas Init.Hints Init.Conv Init.Guard Init.Simproc Init.SizeOfLemmas Init.BinderPredicates Init.Ext Init.Omega Init.MacroTrace Init.Grind Init.While Init.Syntax Init.Internal #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -46,6 +46,7 @@ lean_object* initialize_Init_MacroTrace(uint8_t builtin, lean_object*); lean_object* initialize_Init_Grind(uint8_t builtin, lean_object*); lean_object* initialize_Init_While(uint8_t builtin, lean_object*); lean_object* initialize_Init_Syntax(uint8_t builtin, lean_object*); +lean_object* initialize_Init_Internal(uint8_t builtin, lean_object*); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Init(uint8_t builtin, lean_object* w) { lean_object * res; @@ -150,6 +151,9 @@ lean_dec_ref(res); res = initialize_Init_Syntax(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +res = initialize_Init_Internal(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/Grind/Lemmas.c b/stage0/stdlib/Init/Grind/Lemmas.c index c70997abd3b5..f0df7c37f12e 100644 --- a/stage0/stdlib/Init/Grind/Lemmas.c +++ b/stage0/stdlib/Init/Grind/Lemmas.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Init.Grind.Lemmas -// Imports: Init.Core Init.SimpLemmas Init.Classical Init.ByCases +// Imports: Init.Core Init.SimpLemmas Init.Classical Init.ByCases Init.Grind.Util #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -17,6 +17,7 @@ lean_object* initialize_Init_Core(uint8_t builtin, lean_object*); lean_object* initialize_Init_SimpLemmas(uint8_t builtin, lean_object*); lean_object* initialize_Init_Classical(uint8_t builtin, lean_object*); lean_object* initialize_Init_ByCases(uint8_t builtin, lean_object*); +lean_object* initialize_Init_Grind_Util(uint8_t builtin, lean_object*); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Init_Grind_Lemmas(uint8_t builtin, lean_object* w) { lean_object * res; @@ -34,6 +35,9 @@ lean_dec_ref(res); res = initialize_Init_ByCases(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +res = initialize_Init_Grind_Util(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/Grind/Tactics.c b/stage0/stdlib/Init/Grind/Tactics.c index e8d386e846d3..4e46c7564577 100644 --- a/stage0/stdlib/Init/Grind/Tactics.c +++ b/stage0/stdlib/Init/Grind/Tactics.c @@ -14,69 +14,163 @@ extern "C" { #endif static lean_object* l_Lean_Parser_Tactic_grind___closed__7; +extern lean_object* l_Lean_Parser_Tactic_optConfig; +static lean_object* l_Lean_Grind_instInhabitedConfig___closed__1; +static lean_object* l_Lean_Parser_Tactic_grind___closed__10; static lean_object* l_Lean_Parser_Tactic_grind___closed__6; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_grind; static lean_object* l_Lean_Grind_instBEqConfig___closed__1; static lean_object* l_Lean_Parser_Tactic_grind___closed__2; LEAN_EXPORT lean_object* l_Lean_Grind_instBEqConfig; -LEAN_EXPORT lean_object* l___private_Init_Grind_Tactics_0__Lean_Grind_beqConfig____x40_Init_Grind_Tactics___hyg_18____boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Tactic_grind___closed__9; static lean_object* l_Lean_Parser_Tactic_grind___closed__1; -LEAN_EXPORT uint8_t l_Lean_Grind_instInhabitedConfig; +lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Grind_instInhabitedConfig; static lean_object* l_Lean_Parser_Tactic_grind___closed__5; static lean_object* l_Lean_Parser_Tactic_grind___closed__3; +uint8_t lean_nat_dec_eq(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_grind___closed__4; +LEAN_EXPORT lean_object* l___private_Init_Grind_Tactics_0__Lean_Grind_beqConfig____x40_Init_Grind_Tactics___hyg_30____boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Tactic_grind___closed__8; lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT uint8_t l___private_Init_Grind_Tactics_0__Lean_Grind_beqConfig____x40_Init_Grind_Tactics___hyg_18_(uint8_t, uint8_t); -static uint8_t _init_l_Lean_Grind_instInhabitedConfig() { +LEAN_EXPORT uint8_t l___private_Init_Grind_Tactics_0__Lean_Grind_beqConfig____x40_Init_Grind_Tactics___hyg_30_(lean_object*, lean_object*); +static lean_object* _init_l_Lean_Grind_instInhabitedConfig___closed__1() { _start: { -uint8_t x_1; +uint8_t x_1; lean_object* x_2; lean_object* x_3; x_1 = 0; +x_2 = lean_unsigned_to_nat(0u); +x_3 = lean_alloc_ctor(0, 4, 1); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_2); +lean_ctor_set(x_3, 2, x_2); +lean_ctor_set(x_3, 3, x_2); +lean_ctor_set_uint8(x_3, sizeof(void*)*4, x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_Grind_instInhabitedConfig() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Grind_instInhabitedConfig___closed__1; return x_1; } } -LEAN_EXPORT uint8_t l___private_Init_Grind_Tactics_0__Lean_Grind_beqConfig____x40_Init_Grind_Tactics___hyg_18_(uint8_t x_1, uint8_t x_2) { +LEAN_EXPORT uint8_t l___private_Init_Grind_Tactics_0__Lean_Grind_beqConfig____x40_Init_Grind_Tactics___hyg_30_(lean_object* x_1, lean_object* x_2) { _start: { -if (x_1 == 0) +uint8_t 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; +x_3 = lean_ctor_get_uint8(x_1, sizeof(void*)*4); +x_4 = lean_ctor_get(x_1, 0); +x_5 = lean_ctor_get(x_1, 1); +x_6 = lean_ctor_get(x_1, 2); +x_7 = lean_ctor_get(x_1, 3); +x_8 = lean_ctor_get_uint8(x_2, sizeof(void*)*4); +x_9 = lean_ctor_get(x_2, 0); +x_10 = lean_ctor_get(x_2, 1); +x_11 = lean_ctor_get(x_2, 2); +x_12 = lean_ctor_get(x_2, 3); +if (x_3 == 0) { -if (x_2 == 0) +if (x_8 == 0) { -uint8_t x_3; -x_3 = 1; -return x_3; +uint8_t x_23; +x_23 = 1; +x_13 = x_23; +goto block_22; } else { -uint8_t x_4; -x_4 = 0; -return x_4; +uint8_t x_24; +x_24 = 0; +x_13 = x_24; +goto block_22; +} +} +else +{ +if (x_8 == 0) +{ +uint8_t x_25; +x_25 = 0; +x_13 = x_25; +goto block_22; +} +else +{ +uint8_t x_26; +x_26 = 1; +x_13 = x_26; +goto block_22; } } +block_22: +{ +if (x_13 == 0) +{ +uint8_t x_14; +x_14 = 0; +return x_14; +} +else +{ +uint8_t x_15; +x_15 = lean_nat_dec_eq(x_4, x_9); +if (x_15 == 0) +{ +uint8_t x_16; +x_16 = 0; +return x_16; +} +else +{ +uint8_t x_17; +x_17 = lean_nat_dec_eq(x_5, x_10); +if (x_17 == 0) +{ +uint8_t x_18; +x_18 = 0; +return x_18; +} +else +{ +uint8_t x_19; +x_19 = lean_nat_dec_eq(x_6, x_11); +if (x_19 == 0) +{ +uint8_t x_20; +x_20 = 0; +return x_20; +} else { -return x_2; +uint8_t x_21; +x_21 = lean_nat_dec_eq(x_7, x_12); +return x_21; } } } -LEAN_EXPORT lean_object* l___private_Init_Grind_Tactics_0__Lean_Grind_beqConfig____x40_Init_Grind_Tactics___hyg_18____boxed(lean_object* x_1, lean_object* x_2) { +} +} +} +} +LEAN_EXPORT lean_object* l___private_Init_Grind_Tactics_0__Lean_Grind_beqConfig____x40_Init_Grind_Tactics___hyg_30____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); +uint8_t x_3; lean_object* x_4; +x_3 = l___private_Init_Grind_Tactics_0__Lean_Grind_beqConfig____x40_Init_Grind_Tactics___hyg_30_(x_1, x_2); lean_dec(x_2); -x_5 = l___private_Init_Grind_Tactics_0__Lean_Grind_beqConfig____x40_Init_Grind_Tactics___hyg_18_(x_3, x_4); -x_6 = lean_box(x_5); -return x_6; +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; } } static lean_object* _init_l_Lean_Grind_instBEqConfig___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Init_Grind_Tactics_0__Lean_Grind_beqConfig____x40_Init_Grind_Tactics___hyg_18____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Init_Grind_Tactics_0__Lean_Grind_beqConfig____x40_Init_Grind_Tactics___hyg_30____boxed), 2, 0); return x_1; } } @@ -135,6 +229,24 @@ return x_5; static lean_object* _init_l_Lean_Parser_Tactic_grind___closed__6() { _start: { +lean_object* x_1; +x_1 = lean_mk_string_unchecked("andthen", 7, 7); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Tactic_grind___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_grind___closed__6; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Tactic_grind___closed__8() { +_start: +{ lean_object* x_1; uint8_t x_2; lean_object* x_3; x_1 = l_Lean_Parser_Tactic_grind___closed__4; x_2 = 0; @@ -144,13 +256,27 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_grind___closed__7() { +static lean_object* _init_l_Lean_Parser_Tactic_grind___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Parser_Tactic_grind___closed__7; +x_2 = l_Lean_Parser_Tactic_grind___closed__8; +x_3 = l_Lean_Parser_Tactic_optConfig; +x_4 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Parser_Tactic_grind___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Tactic_grind___closed__5; -x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Tactic_grind___closed__6; +x_2 = lean_unsigned_to_nat(1022u); +x_3 = l_Lean_Parser_Tactic_grind___closed__9; x_4 = lean_alloc_ctor(3, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -162,7 +288,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_grind() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Tactic_grind___closed__7; +x_1 = l_Lean_Parser_Tactic_grind___closed__10; return x_1; } } @@ -175,7 +301,10 @@ _G_initialized = true; res = initialize_Init_Tactics(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +l_Lean_Grind_instInhabitedConfig___closed__1 = _init_l_Lean_Grind_instInhabitedConfig___closed__1(); +lean_mark_persistent(l_Lean_Grind_instInhabitedConfig___closed__1); l_Lean_Grind_instInhabitedConfig = _init_l_Lean_Grind_instInhabitedConfig(); +lean_mark_persistent(l_Lean_Grind_instInhabitedConfig); l_Lean_Grind_instBEqConfig___closed__1 = _init_l_Lean_Grind_instBEqConfig___closed__1(); lean_mark_persistent(l_Lean_Grind_instBEqConfig___closed__1); l_Lean_Grind_instBEqConfig = _init_l_Lean_Grind_instBEqConfig(); @@ -194,6 +323,12 @@ l_Lean_Parser_Tactic_grind___closed__6 = _init_l_Lean_Parser_Tactic_grind___clos lean_mark_persistent(l_Lean_Parser_Tactic_grind___closed__6); l_Lean_Parser_Tactic_grind___closed__7 = _init_l_Lean_Parser_Tactic_grind___closed__7(); lean_mark_persistent(l_Lean_Parser_Tactic_grind___closed__7); +l_Lean_Parser_Tactic_grind___closed__8 = _init_l_Lean_Parser_Tactic_grind___closed__8(); +lean_mark_persistent(l_Lean_Parser_Tactic_grind___closed__8); +l_Lean_Parser_Tactic_grind___closed__9 = _init_l_Lean_Parser_Tactic_grind___closed__9(); +lean_mark_persistent(l_Lean_Parser_Tactic_grind___closed__9); +l_Lean_Parser_Tactic_grind___closed__10 = _init_l_Lean_Parser_Tactic_grind___closed__10(); +lean_mark_persistent(l_Lean_Parser_Tactic_grind___closed__10); l_Lean_Parser_Tactic_grind = _init_l_Lean_Parser_Tactic_grind(); lean_mark_persistent(l_Lean_Parser_Tactic_grind); return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Init/Internal.c b/stage0/stdlib/Init/Internal.c new file mode 100644 index 000000000000..071ffcd9c2c3 --- /dev/null +++ b/stage0/stdlib/Init/Internal.c @@ -0,0 +1,29 @@ +// Lean compiler output +// Module: Init.Internal +// Imports: Init.Internal.Order +#include +#if defined(__clang__) +#pragma clang diagnostic ignored "-Wunused-parameter" +#pragma clang diagnostic ignored "-Wunused-label" +#elif defined(__GNUC__) && !defined(__CLANG__) +#pragma GCC diagnostic ignored "-Wunused-parameter" +#pragma GCC diagnostic ignored "-Wunused-label" +#pragma GCC diagnostic ignored "-Wunused-but-set-variable" +#endif +#ifdef __cplusplus +extern "C" { +#endif +lean_object* initialize_Init_Internal_Order(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Init_Internal(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_Internal_Order(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/Internal/Order.c b/stage0/stdlib/Init/Internal/Order.c new file mode 100644 index 000000000000..5285d39c0ba1 --- /dev/null +++ b/stage0/stdlib/Init/Internal/Order.c @@ -0,0 +1,33 @@ +// Lean compiler output +// Module: Init.Internal.Order +// Imports: Init.Internal.Order.Basic Init.Internal.Order.Tactic +#include +#if defined(__clang__) +#pragma clang diagnostic ignored "-Wunused-parameter" +#pragma clang diagnostic ignored "-Wunused-label" +#elif defined(__GNUC__) && !defined(__CLANG__) +#pragma GCC diagnostic ignored "-Wunused-parameter" +#pragma GCC diagnostic ignored "-Wunused-label" +#pragma GCC diagnostic ignored "-Wunused-but-set-variable" +#endif +#ifdef __cplusplus +extern "C" { +#endif +lean_object* initialize_Init_Internal_Order_Basic(uint8_t builtin, lean_object*); +lean_object* initialize_Init_Internal_Order_Tactic(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Init_Internal_Order(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_Internal_Order_Basic(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Init_Internal_Order_Tactic(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/Internal/Order/Basic.c b/stage0/stdlib/Init/Internal/Order/Basic.c new file mode 100644 index 000000000000..5b8180dd8563 --- /dev/null +++ b/stage0/stdlib/Init/Internal/Order/Basic.c @@ -0,0 +1,1126 @@ +// Lean compiler output +// Module: Init.Internal.Order.Basic +// Imports: Init.ByCases Init.RCases +#include +#if defined(__clang__) +#pragma clang diagnostic ignored "-Wunused-parameter" +#pragma clang diagnostic ignored "-Wunused-label" +#elif defined(__GNUC__) && !defined(__CLANG__) +#pragma GCC diagnostic ignored "-Wunused-parameter" +#pragma GCC diagnostic ignored "-Wunused-label" +#pragma GCC diagnostic ignored "-Wunused-but-set-variable" +#endif +#ifdef __cplusplus +extern "C" { +#endif +LEAN_EXPORT lean_object* l_Lean_Order_instPartialOrderExceptTOfMonad(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__6; +LEAN_EXPORT lean_object* l_Lean_Order_instCCPOPi___rarg(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Order_instCCPOExceptTOfMonadOfPartialOrder(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Order_term___u2291_____closed__11; +LEAN_EXPORT lean_object* l_Lean_Order_fix(lean_object*); +static lean_object* l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__11; +static lean_object* l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term_u22a5__1___closed__2; +static lean_object* l_Lean_Order_term___u2291_____closed__10; +LEAN_EXPORT lean_object* l_Lean_Order_Example_findF(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Order_fix___rarg(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_replaceRef(lean_object*, lean_object*); +static lean_object* l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__1; +uint8_t l_Lean_Syntax_isOfKind(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Order_FlatOrder_instOrder___boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Order_term_u22a5___closed__4; +static lean_object* l_Lean_Order_term_u22a5___closed__5; +static lean_object* l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__2; +static lean_object* l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term_u22a5__1___closed__1; +LEAN_EXPORT lean_object* l_Lean_Order___aux__Init__Internal__Order__Basic______unexpand__Lean__Order__bot__1(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Order_instPartialOrderPProd(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Order_instPartialOrderExceptTOfMonad___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Order_FlatOrder_instOrder(lean_object*, lean_object*); +static lean_object* l_Lean_Order_term___u2291_____closed__7; +LEAN_EXPORT lean_object* l_Lean_Order___aux__Init__Internal__Order__Basic______unexpand__Lean__Order__PartialOrder__rel__1___boxed(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_SourceInfo_fromRef(lean_object*, uint8_t); +static lean_object* l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__10; +static lean_object* l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__5; +static lean_object* l_Lean_Order_term___u2291_____closed__2; +LEAN_EXPORT lean_object* l_Lean_Order_instOrderPi___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Order_term___u2291_____closed__13; +static lean_object* l_Lean_Order_term___u2291_____closed__6; +lean_object* l_Lean_Syntax_node3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__13; +static lean_object* l_Lean_Order_term___u2291_____closed__4; +LEAN_EXPORT lean_object* l_Lean_Order_fix___rarg___boxed(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_addMacroScope(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term_u22a5__1___closed__3; +LEAN_EXPORT lean_object* l_Lean_Order_instPartialOrderExceptTOfMonad___rarg(lean_object*); +lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); +static lean_object* l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term_u22a5__1___closed__6; +static lean_object* l_Lean_Order_term_u22a5___closed__3; +lean_object* l_Lean_Syntax_node2(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term_u22a5__1___closed__5; +LEAN_EXPORT lean_object* l_Lean_Order_bot(lean_object*); +lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); +static lean_object* l_Lean_Order_term___u2291_____closed__8; +uint8_t l_Lean_Syntax_matchesNull(lean_object*, lean_object*); +static lean_object* l_Lean_Order_term_u22a5___closed__1; +LEAN_EXPORT lean_object* l_Lean_Order_bot___rarg(lean_object*); +static lean_object* l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term_u22a5__1___closed__4; +LEAN_EXPORT lean_object* l_Lean_Order_fun__csup___rarg(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Order_term___u2291_____closed__1; +LEAN_EXPORT lean_object* l_Lean_Order_fun__csup(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Order_term___u2291__; +static lean_object* l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__3; +lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); +static lean_object* l_Lean_Order___aux__Init__Internal__Order__Basic______unexpand__Lean__Order__PartialOrder__rel__1___closed__1; +static lean_object* l_Lean_Order_term_u22a5___closed__2; +lean_object* l_Lean_Syntax_node1(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__9; +LEAN_EXPORT lean_object* l_Lean_Order_instCCPOPProd___rarg(lean_object*, lean_object*); +static lean_object* l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__8; +static lean_object* l_Lean_Order_term___u2291_____closed__5; +LEAN_EXPORT lean_object* l_Lean_Order_term_u22a5; +static lean_object* l_Lean_Order_term___u2291_____closed__3; +static lean_object* l_Lean_Order_term___u2291_____closed__12; +static lean_object* l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__7; +static lean_object* l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__4; +static lean_object* l_Lean_Order___aux__Init__Internal__Order__Basic______unexpand__Lean__Order__PartialOrder__rel__1___closed__2; +static lean_object* l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__14; +LEAN_EXPORT lean_object* l_Lean_Order___aux__Init__Internal__Order__Basic______unexpand__Lean__Order__PartialOrder__rel__1(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_Order_instCCPOPi(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Order___aux__Init__Internal__Order__Basic______unexpand__Lean__Order__bot__1___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term_u22a5__1(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Order_instCCPOExceptTOfMonadOfPartialOrder___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Order_term___u2291_____closed__9; +LEAN_EXPORT lean_object* l_Lean_Order_instPartialOrderPProd___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_nat_add(lean_object*, lean_object*); +static lean_object* l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__12; +LEAN_EXPORT lean_object* l_Lean_Order_instCCPOExceptTOfMonadOfPartialOrder___rarg(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Order_instCCPOPProd(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Order_instCCPOPProd___rarg___lambda__1(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Order_instOrderPi(lean_object*, lean_object*, lean_object*); +lean_object* l_String_toSubstring_x27(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Order_instPartialOrderOption(lean_object*); +static lean_object* _init_l_Lean_Order_term___u2291_____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean", 4, 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Order_term___u2291_____closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Order", 5, 5); +return x_1; +} +} +static lean_object* _init_l_Lean_Order_term___u2291_____closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("term_⊑_", 9, 7); +return x_1; +} +} +static lean_object* _init_l_Lean_Order_term___u2291_____closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Order_term___u2291_____closed__1; +x_2 = l_Lean_Order_term___u2291_____closed__2; +x_3 = l_Lean_Order_term___u2291_____closed__3; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Order_term___u2291_____closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("andthen", 7, 7); +return x_1; +} +} +static lean_object* _init_l_Lean_Order_term___u2291_____closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Order_term___u2291_____closed__5; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Order_term___u2291_____closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(" ⊑ ", 5, 3); +return x_1; +} +} +static lean_object* _init_l_Lean_Order_term___u2291_____closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Order_term___u2291_____closed__7; +x_2 = lean_alloc_ctor(5, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Order_term___u2291_____closed__9() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("term", 4, 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Order_term___u2291_____closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Order_term___u2291_____closed__9; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Order_term___u2291_____closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Order_term___u2291_____closed__10; +x_2 = lean_unsigned_to_nat(51u); +x_3 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Order_term___u2291_____closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Order_term___u2291_____closed__6; +x_2 = l_Lean_Order_term___u2291_____closed__8; +x_3 = l_Lean_Order_term___u2291_____closed__11; +x_4 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Order_term___u2291_____closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Lean_Order_term___u2291_____closed__4; +x_2 = lean_unsigned_to_nat(50u); +x_3 = lean_unsigned_to_nat(51u); +x_4 = l_Lean_Order_term___u2291_____closed__12; +x_5 = lean_alloc_ctor(4, 4, 0); +lean_ctor_set(x_5, 0, x_1); +lean_ctor_set(x_5, 1, x_2); +lean_ctor_set(x_5, 2, x_3); +lean_ctor_set(x_5, 3, x_4); +return x_5; +} +} +static lean_object* _init_l_Lean_Order_term___u2291__() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Order_term___u2291_____closed__13; +return x_1; +} +} +static lean_object* _init_l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Parser", 6, 6); +return x_1; +} +} +static lean_object* _init_l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Term", 4, 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("app", 3, 3); +return x_1; +} +} +static lean_object* _init_l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Lean_Order_term___u2291_____closed__1; +x_2 = l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__1; +x_3 = l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__2; +x_4 = l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__3; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("PartialOrder.rel", 16, 16); +return x_1; +} +} +static lean_object* _init_l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__5; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("PartialOrder", 12, 12); +return x_1; +} +} +static lean_object* _init_l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__8() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("rel", 3, 3); +return x_1; +} +} +static lean_object* _init_l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__7; +x_2 = l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__8; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____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_Order_term___u2291_____closed__1; +x_2 = l_Lean_Order_term___u2291_____closed__2; +x_3 = l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__7; +x_4 = l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__8; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__10; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____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_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__13() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("null", 4, 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__13; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____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_Order_term___u2291_____closed__4; +lean_inc(x_1); +x_5 = l_Lean_Syntax_isOfKind(x_1, x_4); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; +lean_dec(x_2); +lean_dec(x_1); +x_6 = lean_box(1); +x_7 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_7, 0, x_6); +lean_ctor_set(x_7, 1, x_3); +return x_7; +} +else +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; 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_8 = lean_unsigned_to_nat(0u); +x_9 = l_Lean_Syntax_getArg(x_1, x_8); +x_10 = lean_unsigned_to_nat(2u); +x_11 = l_Lean_Syntax_getArg(x_1, x_10); +lean_dec(x_1); +x_12 = lean_ctor_get(x_2, 5); +lean_inc(x_12); +x_13 = 0; +x_14 = l_Lean_SourceInfo_fromRef(x_12, x_13); +lean_dec(x_12); +x_15 = lean_ctor_get(x_2, 2); +lean_inc(x_15); +x_16 = lean_ctor_get(x_2, 1); +lean_inc(x_16); +lean_dec(x_2); +x_17 = l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__9; +x_18 = l_Lean_addMacroScope(x_16, x_17, x_15); +x_19 = l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__6; +x_20 = l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__12; +lean_inc(x_14); +x_21 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_21, 0, x_14); +lean_ctor_set(x_21, 1, x_19); +lean_ctor_set(x_21, 2, x_18); +lean_ctor_set(x_21, 3, x_20); +x_22 = l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__14; +lean_inc(x_14); +x_23 = l_Lean_Syntax_node2(x_14, x_22, x_9, x_11); +x_24 = l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__4; +x_25 = l_Lean_Syntax_node2(x_14, x_24, x_21, 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_3); +return x_26; +} +} +} +static lean_object* _init_l_Lean_Order___aux__Init__Internal__Order__Basic______unexpand__Lean__Order__PartialOrder__rel__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("ident", 5, 5); +return x_1; +} +} +static lean_object* _init_l_Lean_Order___aux__Init__Internal__Order__Basic______unexpand__Lean__Order__PartialOrder__rel__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Order___aux__Init__Internal__Order__Basic______unexpand__Lean__Order__PartialOrder__rel__1___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Order___aux__Init__Internal__Order__Basic______unexpand__Lean__Order__PartialOrder__rel__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_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__4; +lean_inc(x_1); +x_5 = l_Lean_Syntax_isOfKind(x_1, x_4); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; +lean_dec(x_1); +x_6 = lean_box(0); +x_7 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_7, 0, x_6); +lean_ctor_set(x_7, 1, x_3); +return x_7; +} +else +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_8 = lean_unsigned_to_nat(0u); +x_9 = l_Lean_Syntax_getArg(x_1, x_8); +x_10 = l_Lean_Order___aux__Init__Internal__Order__Basic______unexpand__Lean__Order__PartialOrder__rel__1___closed__2; +lean_inc(x_9); +x_11 = l_Lean_Syntax_isOfKind(x_9, x_10); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; +lean_dec(x_9); +lean_dec(x_1); +x_12 = lean_box(0); +x_13 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_3); +return x_13; +} +else +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_14 = lean_unsigned_to_nat(1u); +x_15 = l_Lean_Syntax_getArg(x_1, x_14); +lean_dec(x_1); +x_16 = lean_unsigned_to_nat(2u); +lean_inc(x_15); +x_17 = l_Lean_Syntax_matchesNull(x_15, x_16); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; +lean_dec(x_15); +lean_dec(x_9); +x_18 = lean_box(0); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_3); +return x_19; +} +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_20 = l_Lean_Syntax_getArg(x_15, x_8); +x_21 = l_Lean_Syntax_getArg(x_15, x_14); +lean_dec(x_15); +x_22 = l_Lean_replaceRef(x_9, x_2); +lean_dec(x_9); +x_23 = 0; +x_24 = l_Lean_SourceInfo_fromRef(x_22, x_23); +lean_dec(x_22); +x_25 = l_Lean_Order_term___u2291_____closed__7; +lean_inc(x_24); +x_26 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_26, 0, x_24); +lean_ctor_set(x_26, 1, x_25); +x_27 = l_Lean_Order_term___u2291_____closed__4; +x_28 = l_Lean_Syntax_node3(x_24, x_27, x_20, x_26, x_21); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_3); +return x_29; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Order___aux__Init__Internal__Order__Basic______unexpand__Lean__Order__PartialOrder__rel__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Lean_Order___aux__Init__Internal__Order__Basic______unexpand__Lean__Order__PartialOrder__rel__1(x_1, x_2, x_3); +lean_dec(x_2); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_Order_bot___rarg(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = lean_ctor_get(x_1, 1); +lean_inc(x_2); +lean_dec(x_1); +x_3 = lean_apply_1(x_2, lean_box(0)); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Order_bot(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_Order_bot___rarg), 1, 0); +return x_2; +} +} +static lean_object* _init_l_Lean_Order_term_u22a5___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("term⊥", 7, 5); +return x_1; +} +} +static lean_object* _init_l_Lean_Order_term_u22a5___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Order_term___u2291_____closed__1; +x_2 = l_Lean_Order_term___u2291_____closed__2; +x_3 = l_Lean_Order_term_u22a5___closed__1; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Order_term_u22a5___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("⊥", 3, 1); +return x_1; +} +} +static lean_object* _init_l_Lean_Order_term_u22a5___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Order_term_u22a5___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_Order_term_u22a5___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Order_term_u22a5___closed__2; +x_2 = lean_unsigned_to_nat(1024u); +x_3 = l_Lean_Order_term_u22a5___closed__4; +x_4 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Order_term_u22a5() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Order_term_u22a5___closed__5; +return x_1; +} +} +static lean_object* _init_l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term_u22a5__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("bot", 3, 3); +return x_1; +} +} +static lean_object* _init_l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term_u22a5__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term_u22a5__1___closed__1; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term_u22a5__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_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term_u22a5__1___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term_u22a5__1___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Order_term___u2291_____closed__1; +x_2 = l_Lean_Order_term___u2291_____closed__2; +x_3 = l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term_u22a5__1___closed__1; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term_u22a5__1___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term_u22a5__1___closed__4; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term_u22a5__1___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term_u22a5__1___closed__5; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term_u22a5__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_Order_term_u22a5___closed__2; +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); +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; 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_8 = lean_ctor_get(x_2, 5); +lean_inc(x_8); +x_9 = 0; +x_10 = l_Lean_SourceInfo_fromRef(x_8, x_9); +lean_dec(x_8); +x_11 = lean_ctor_get(x_2, 2); +lean_inc(x_11); +x_12 = lean_ctor_get(x_2, 1); +lean_inc(x_12); +lean_dec(x_2); +x_13 = l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term_u22a5__1___closed__3; +x_14 = l_Lean_addMacroScope(x_12, x_13, x_11); +x_15 = l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term_u22a5__1___closed__2; +x_16 = l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term_u22a5__1___closed__6; +x_17 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_17, 0, x_10); +lean_ctor_set(x_17, 1, x_15); +lean_ctor_set(x_17, 2, x_14); +lean_ctor_set(x_17, 3, x_16); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_17); +lean_ctor_set(x_18, 1, x_3); +return x_18; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Order___aux__Init__Internal__Order__Basic______unexpand__Lean__Order__bot__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_Order___aux__Init__Internal__Order__Basic______unexpand__Lean__Order__PartialOrder__rel__1___closed__2; +lean_inc(x_1); +x_5 = l_Lean_Syntax_isOfKind(x_1, x_4); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; +lean_dec(x_1); +x_6 = lean_box(0); +x_7 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_7, 0, x_6); +lean_ctor_set(x_7, 1, x_3); +return x_7; +} +else +{ +lean_object* x_8; 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_8 = l_Lean_replaceRef(x_1, x_2); +lean_dec(x_1); +x_9 = 0; +x_10 = l_Lean_SourceInfo_fromRef(x_8, x_9); +lean_dec(x_8); +x_11 = l_Lean_Order_term_u22a5___closed__3; +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_Order_term_u22a5___closed__2; +x_14 = l_Lean_Syntax_node1(x_10, x_13, x_12); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_14); +lean_ctor_set(x_15, 1, x_3); +return x_15; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Order___aux__Init__Internal__Order__Basic______unexpand__Lean__Order__bot__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Lean_Order___aux__Init__Internal__Order__Basic______unexpand__Lean__Order__bot__1(x_1, x_2, x_3); +lean_dec(x_2); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_Order_fix___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, 1); +lean_inc(x_4); +lean_dec(x_1); +x_5 = lean_apply_1(x_4, lean_box(0)); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_Order_fix(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_Order_fix___rarg___boxed), 3, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Order_fix___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Lean_Order_fix___rarg(x_1, x_2, x_3); +lean_dec(x_2); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_Order_instOrderPi(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = lean_box(0); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_Order_instOrderPi___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Lean_Order_instOrderPi(x_1, x_2, x_3); +lean_dec(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_Order_fun__csup___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; +x_4 = lean_apply_1(x_1, x_3); +x_5 = lean_ctor_get(x_4, 1); +lean_inc(x_5); +lean_dec(x_4); +x_6 = lean_apply_1(x_5, lean_box(0)); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Lean_Order_fun__csup(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Lean_Order_fun__csup___rarg), 3, 0); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Order_instCCPOPi___rarg(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_2 = lean_box(0); +x_3 = lean_alloc_closure((void*)(l_Lean_Order_fun__csup___rarg), 3, 1); +lean_closure_set(x_3, 0, x_1); +x_4 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_4, 0, x_2); +lean_ctor_set(x_4, 1, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_Order_instCCPOPi(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Lean_Order_instCCPOPi___rarg), 1, 0); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Order_instPartialOrderPProd(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = lean_box(0); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_Order_instPartialOrderPProd___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_Order_instPartialOrderPProd(x_1, x_2, x_3, x_4); +lean_dec(x_4); +lean_dec(x_3); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_Order_instCCPOPProd___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; lean_object* x_8; +x_4 = lean_ctor_get(x_1, 1); +lean_inc(x_4); +lean_dec(x_1); +x_5 = lean_apply_1(x_4, lean_box(0)); +x_6 = lean_ctor_get(x_2, 1); +lean_inc(x_6); +lean_dec(x_2); +x_7 = lean_apply_1(x_6, lean_box(0)); +x_8 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_8, 0, x_5); +lean_ctor_set(x_8, 1, x_7); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_Order_instCCPOPProd___rarg(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = lean_box(0); +x_4 = lean_alloc_closure((void*)(l_Lean_Order_instCCPOPProd___rarg___lambda__1), 3, 2); +lean_closure_set(x_4, 0, x_1); +lean_closure_set(x_4, 1, x_2); +x_5 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_5, 0, x_3); +lean_ctor_set(x_5, 1, x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_Order_instCCPOPProd(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Lean_Order_instCCPOPProd___rarg), 2, 0); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Order_FlatOrder_instOrder(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_box(0); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Order_FlatOrder_instOrder___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_Order_FlatOrder_instOrder(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Order_instPartialOrderOption(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_box(0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Order_instPartialOrderExceptTOfMonad___rarg(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_apply_1(x_1, lean_box(0)); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Order_instPartialOrderExceptTOfMonad(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_Order_instPartialOrderExceptTOfMonad___rarg), 1, 0); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_Order_instPartialOrderExceptTOfMonad___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_Order_instPartialOrderExceptTOfMonad(x_1, x_2, x_3, x_4); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_Order_instCCPOExceptTOfMonadOfPartialOrder___rarg(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_apply_1(x_1, lean_box(0)); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Order_instCCPOExceptTOfMonadOfPartialOrder(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_Order_instCCPOExceptTOfMonadOfPartialOrder___rarg), 1, 0); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Lean_Order_instCCPOExceptTOfMonadOfPartialOrder___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_Order_instCCPOExceptTOfMonadOfPartialOrder(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_5); +lean_dec(x_4); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Lean_Order_Example_findF(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; uint8_t x_5; +lean_inc(x_3); +x_4 = lean_apply_1(x_1, x_3); +x_5 = lean_unbox(x_4); +lean_dec(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_add(x_3, x_6); +lean_dec(x_3); +x_8 = lean_apply_1(x_2, x_7); +return x_8; +} +else +{ +lean_object* x_9; +lean_dec(x_2); +x_9 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_9, 0, x_3); +return x_9; +} +} +} +lean_object* initialize_Init_ByCases(uint8_t builtin, lean_object*); +lean_object* initialize_Init_RCases(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Init_Internal_Order_Basic(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_ByCases(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Init_RCases(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l_Lean_Order_term___u2291_____closed__1 = _init_l_Lean_Order_term___u2291_____closed__1(); +lean_mark_persistent(l_Lean_Order_term___u2291_____closed__1); +l_Lean_Order_term___u2291_____closed__2 = _init_l_Lean_Order_term___u2291_____closed__2(); +lean_mark_persistent(l_Lean_Order_term___u2291_____closed__2); +l_Lean_Order_term___u2291_____closed__3 = _init_l_Lean_Order_term___u2291_____closed__3(); +lean_mark_persistent(l_Lean_Order_term___u2291_____closed__3); +l_Lean_Order_term___u2291_____closed__4 = _init_l_Lean_Order_term___u2291_____closed__4(); +lean_mark_persistent(l_Lean_Order_term___u2291_____closed__4); +l_Lean_Order_term___u2291_____closed__5 = _init_l_Lean_Order_term___u2291_____closed__5(); +lean_mark_persistent(l_Lean_Order_term___u2291_____closed__5); +l_Lean_Order_term___u2291_____closed__6 = _init_l_Lean_Order_term___u2291_____closed__6(); +lean_mark_persistent(l_Lean_Order_term___u2291_____closed__6); +l_Lean_Order_term___u2291_____closed__7 = _init_l_Lean_Order_term___u2291_____closed__7(); +lean_mark_persistent(l_Lean_Order_term___u2291_____closed__7); +l_Lean_Order_term___u2291_____closed__8 = _init_l_Lean_Order_term___u2291_____closed__8(); +lean_mark_persistent(l_Lean_Order_term___u2291_____closed__8); +l_Lean_Order_term___u2291_____closed__9 = _init_l_Lean_Order_term___u2291_____closed__9(); +lean_mark_persistent(l_Lean_Order_term___u2291_____closed__9); +l_Lean_Order_term___u2291_____closed__10 = _init_l_Lean_Order_term___u2291_____closed__10(); +lean_mark_persistent(l_Lean_Order_term___u2291_____closed__10); +l_Lean_Order_term___u2291_____closed__11 = _init_l_Lean_Order_term___u2291_____closed__11(); +lean_mark_persistent(l_Lean_Order_term___u2291_____closed__11); +l_Lean_Order_term___u2291_____closed__12 = _init_l_Lean_Order_term___u2291_____closed__12(); +lean_mark_persistent(l_Lean_Order_term___u2291_____closed__12); +l_Lean_Order_term___u2291_____closed__13 = _init_l_Lean_Order_term___u2291_____closed__13(); +lean_mark_persistent(l_Lean_Order_term___u2291_____closed__13); +l_Lean_Order_term___u2291__ = _init_l_Lean_Order_term___u2291__(); +lean_mark_persistent(l_Lean_Order_term___u2291__); +l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__1 = _init_l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__1(); +lean_mark_persistent(l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__1); +l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__2 = _init_l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__2(); +lean_mark_persistent(l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__2); +l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__3 = _init_l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__3(); +lean_mark_persistent(l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__3); +l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__4 = _init_l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__4(); +lean_mark_persistent(l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__4); +l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__5 = _init_l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__5(); +lean_mark_persistent(l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__5); +l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__6 = _init_l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__6(); +lean_mark_persistent(l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__6); +l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__7 = _init_l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__7(); +lean_mark_persistent(l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__7); +l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__8 = _init_l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__8(); +lean_mark_persistent(l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__8); +l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__9 = _init_l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__9(); +lean_mark_persistent(l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__9); +l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__10 = _init_l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__10(); +lean_mark_persistent(l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__10); +l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__11 = _init_l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__11(); +lean_mark_persistent(l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__11); +l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__12 = _init_l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__12(); +lean_mark_persistent(l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__12); +l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__13 = _init_l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__13(); +lean_mark_persistent(l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__13); +l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__14 = _init_l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__14(); +lean_mark_persistent(l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term___u2291____1___closed__14); +l_Lean_Order___aux__Init__Internal__Order__Basic______unexpand__Lean__Order__PartialOrder__rel__1___closed__1 = _init_l_Lean_Order___aux__Init__Internal__Order__Basic______unexpand__Lean__Order__PartialOrder__rel__1___closed__1(); +lean_mark_persistent(l_Lean_Order___aux__Init__Internal__Order__Basic______unexpand__Lean__Order__PartialOrder__rel__1___closed__1); +l_Lean_Order___aux__Init__Internal__Order__Basic______unexpand__Lean__Order__PartialOrder__rel__1___closed__2 = _init_l_Lean_Order___aux__Init__Internal__Order__Basic______unexpand__Lean__Order__PartialOrder__rel__1___closed__2(); +lean_mark_persistent(l_Lean_Order___aux__Init__Internal__Order__Basic______unexpand__Lean__Order__PartialOrder__rel__1___closed__2); +l_Lean_Order_term_u22a5___closed__1 = _init_l_Lean_Order_term_u22a5___closed__1(); +lean_mark_persistent(l_Lean_Order_term_u22a5___closed__1); +l_Lean_Order_term_u22a5___closed__2 = _init_l_Lean_Order_term_u22a5___closed__2(); +lean_mark_persistent(l_Lean_Order_term_u22a5___closed__2); +l_Lean_Order_term_u22a5___closed__3 = _init_l_Lean_Order_term_u22a5___closed__3(); +lean_mark_persistent(l_Lean_Order_term_u22a5___closed__3); +l_Lean_Order_term_u22a5___closed__4 = _init_l_Lean_Order_term_u22a5___closed__4(); +lean_mark_persistent(l_Lean_Order_term_u22a5___closed__4); +l_Lean_Order_term_u22a5___closed__5 = _init_l_Lean_Order_term_u22a5___closed__5(); +lean_mark_persistent(l_Lean_Order_term_u22a5___closed__5); +l_Lean_Order_term_u22a5 = _init_l_Lean_Order_term_u22a5(); +lean_mark_persistent(l_Lean_Order_term_u22a5); +l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term_u22a5__1___closed__1 = _init_l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term_u22a5__1___closed__1(); +lean_mark_persistent(l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term_u22a5__1___closed__1); +l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term_u22a5__1___closed__2 = _init_l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term_u22a5__1___closed__2(); +lean_mark_persistent(l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term_u22a5__1___closed__2); +l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term_u22a5__1___closed__3 = _init_l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term_u22a5__1___closed__3(); +lean_mark_persistent(l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term_u22a5__1___closed__3); +l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term_u22a5__1___closed__4 = _init_l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term_u22a5__1___closed__4(); +lean_mark_persistent(l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term_u22a5__1___closed__4); +l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term_u22a5__1___closed__5 = _init_l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term_u22a5__1___closed__5(); +lean_mark_persistent(l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term_u22a5__1___closed__5); +l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term_u22a5__1___closed__6 = _init_l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term_u22a5__1___closed__6(); +lean_mark_persistent(l_Lean_Order___aux__Init__Internal__Order__Basic______macroRules__Lean__Order__term_u22a5__1___closed__6); +return lean_io_result_mk_ok(lean_box(0)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Init/Internal/Order/Tactic.c b/stage0/stdlib/Init/Internal/Order/Tactic.c new file mode 100644 index 000000000000..852841c70cd0 --- /dev/null +++ b/stage0/stdlib/Init/Internal/Order/Tactic.c @@ -0,0 +1,120 @@ +// Lean compiler output +// Module: Init.Internal.Order.Tactic +// Imports: Init.Notation +#include +#if defined(__clang__) +#pragma clang diagnostic ignored "-Wunused-parameter" +#pragma clang diagnostic ignored "-Wunused-label" +#elif defined(__GNUC__) && !defined(__CLANG__) +#pragma GCC diagnostic ignored "-Wunused-parameter" +#pragma GCC diagnostic ignored "-Wunused-label" +#pragma GCC diagnostic ignored "-Wunused-but-set-variable" +#endif +#ifdef __cplusplus +extern "C" { +#endif +lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Order_monotonicity___closed__6; +LEAN_EXPORT lean_object* l_Lean_Order_monotonicity; +static lean_object* l_Lean_Order_monotonicity___closed__5; +static lean_object* l_Lean_Order_monotonicity___closed__3; +static lean_object* l_Lean_Order_monotonicity___closed__1; +static lean_object* l_Lean_Order_monotonicity___closed__2; +static lean_object* l_Lean_Order_monotonicity___closed__4; +static lean_object* _init_l_Lean_Order_monotonicity___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean", 4, 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Order_monotonicity___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Order", 5, 5); +return x_1; +} +} +static lean_object* _init_l_Lean_Order_monotonicity___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("monotonicity", 12, 12); +return x_1; +} +} +static lean_object* _init_l_Lean_Order_monotonicity___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Order_monotonicity___closed__1; +x_2 = l_Lean_Order_monotonicity___closed__2; +x_3 = l_Lean_Order_monotonicity___closed__3; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Order_monotonicity___closed__5() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l_Lean_Order_monotonicity___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_Lean_Order_monotonicity___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Order_monotonicity___closed__4; +x_2 = lean_unsigned_to_nat(1024u); +x_3 = l_Lean_Order_monotonicity___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_Order_monotonicity() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Order_monotonicity___closed__6; +return x_1; +} +} +lean_object* initialize_Init_Notation(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Init_Internal_Order_Tactic(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_Notation(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l_Lean_Order_monotonicity___closed__1 = _init_l_Lean_Order_monotonicity___closed__1(); +lean_mark_persistent(l_Lean_Order_monotonicity___closed__1); +l_Lean_Order_monotonicity___closed__2 = _init_l_Lean_Order_monotonicity___closed__2(); +lean_mark_persistent(l_Lean_Order_monotonicity___closed__2); +l_Lean_Order_monotonicity___closed__3 = _init_l_Lean_Order_monotonicity___closed__3(); +lean_mark_persistent(l_Lean_Order_monotonicity___closed__3); +l_Lean_Order_monotonicity___closed__4 = _init_l_Lean_Order_monotonicity___closed__4(); +lean_mark_persistent(l_Lean_Order_monotonicity___closed__4); +l_Lean_Order_monotonicity___closed__5 = _init_l_Lean_Order_monotonicity___closed__5(); +lean_mark_persistent(l_Lean_Order_monotonicity___closed__5); +l_Lean_Order_monotonicity___closed__6 = _init_l_Lean_Order_monotonicity___closed__6(); +lean_mark_persistent(l_Lean_Order_monotonicity___closed__6); +l_Lean_Order_monotonicity = _init_l_Lean_Order_monotonicity(); +lean_mark_persistent(l_Lean_Order_monotonicity); +return lean_io_result_mk_ok(lean_box(0)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Lean/Data/PersistentHashSet.c b/stage0/stdlib/Lean/Data/PersistentHashSet.c index 9e6af094c55c..e0cc107695a4 100644 --- a/stage0/stdlib/Lean/Data/PersistentHashSet.c +++ b/stage0/stdlib/Lean/Data/PersistentHashSet.c @@ -26,10 +26,12 @@ uint8_t lean_usize_dec_eq(size_t, size_t); lean_object* lean_array_fget(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashSet_insert(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at_Lean_PersistentHashSet_fold___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_PersistentHashSet_toList___spec__1___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashSet_contains(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashSet_instInhabited___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux___at_Lean_PersistentHashSet_fold___spec__2(lean_object*, lean_object*, lean_object*); size_t lean_usize_of_nat(lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashSet_toList___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashSet_instInhabited___rarg___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashSet_instEmptyCollection___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_PersistentHashSet_fold___spec__4___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -42,8 +44,10 @@ LEAN_EXPORT lean_object* l_Lean_PersistentHashSet_instEmptyCollection___rarg___b LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at_Lean_PersistentHashSet_fold___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashSet_isEmpty___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at_Lean_PersistentHashSet_foldM___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashSet_toList(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashSet_foldM___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_erase___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_PersistentHashMap_toList___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashSet_fold(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at_Lean_PersistentHashSet_foldM___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashSet_fold___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -55,7 +59,9 @@ uint8_t lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at_Lean_PersistentHashSet_fold___spec__1___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_PersistentHashSet_fold___spec__3___rarg(lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_PersistentHashSet_fold___spec__4___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashSet_toList___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashSet_erase(lean_object*); +lean_object* l_List_reverse___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_PersistentHashSet_fold___spec__4(lean_object*, lean_object*, lean_object*); size_t lean_usize_add(size_t, size_t); lean_object* l_Lean_PersistentHashMap_empty(lean_object*, lean_object*, lean_object*, lean_object*); @@ -65,6 +71,7 @@ uint8_t lean_nat_dec_le(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux___at_Lean_PersistentHashSet_fold___spec__2___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashSet_isEmpty___rarg___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashSet_empty___rarg___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_PersistentHashSet_toList___spec__1(lean_object*); lean_object* lean_nat_add(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentHashSet_isEmpty___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashSet_insert___rarg(lean_object*, lean_object*, lean_object*, lean_object*); @@ -608,6 +615,95 @@ lean_dec(x_1); return x_7; } } +LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_PersistentHashSet_toList___spec__1___rarg(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 = lean_ctor_get(x_5, 0); +lean_inc(x_7); +lean_dec(x_5); +lean_ctor_set(x_1, 1, x_2); +lean_ctor_set(x_1, 0, x_7); +{ +lean_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 = lean_ctor_get(x_9, 0); +lean_inc(x_11); +lean_dec(x_9); +x_12 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_2); +x_1 = x_10; +x_2 = x_12; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_PersistentHashSet_toList___spec__1(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_List_mapTR_loop___at_Lean_PersistentHashSet_toList___spec__1___rarg), 2, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashSet_toList___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; +x_4 = l_Lean_PersistentHashMap_toList___rarg(x_1, x_2, x_3); +x_5 = lean_box(0); +x_6 = l_List_mapTR_loop___at_Lean_PersistentHashSet_toList___spec__1___rarg(x_4, x_5); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashSet_toList(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_PersistentHashSet_toList___rarg___boxed), 3, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashSet_toList___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Lean_PersistentHashSet_toList___rarg(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_4; +} +} lean_object* initialize_Lean_Data_PersistentHashMap(uint8_t builtin, lean_object*); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Lean_Data_PersistentHashSet(uint8_t builtin, lean_object* w) { diff --git a/stage0/stdlib/Lean/Elab/Calc.c b/stage0/stdlib/Lean/Elab/Calc.c index 15068446f903..075301bc5e1c 100644 --- a/stage0/stdlib/Lean/Elab/Calc.c +++ b/stage0/stdlib/Lean/Elab/Calc.c @@ -173,6 +173,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkCalcFirstStepView(lean_object*, lean LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_annotateFirstHoleWithType_go___spec__12(lean_object*, size_t, size_t, 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_annotateFirstHoleWithType_go___lambda__1___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Term_throwCalcFailure(lean_object*); +lean_object* l_Lean_Meta_addPPExplicitToExposeDiff(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_annotateFirstHoleWithType_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 lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_annotateFirstHoleWithType_go___spec__10(lean_object*, size_t, size_t, 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_annotateFirstHoleWithType_go___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -19639,170 +19640,568 @@ lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); -lean_inc(x_1); -x_16 = lean_infer_type(x_1, x_7, x_8, x_9, x_10, x_15); +x_16 = l_Lean_Meta_addPPExplicitToExposeDiff(x_1, x_2, x_7, x_8, x_9, x_10, x_15); if (lean_obj_tag(x_16) == 0) { -lean_object* x_17; lean_object* x_18; lean_object* x_19; +lean_object* x_17; lean_object* x_18; uint8_t x_19; x_17 = lean_ctor_get(x_16, 0); lean_inc(x_17); x_18 = lean_ctor_get(x_16, 1); lean_inc(x_18); lean_dec(x_16); +x_19 = !lean_is_exclusive(x_17); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_17, 0); +x_21 = lean_ctor_get(x_17, 1); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); -lean_inc(x_2); -x_19 = lean_infer_type(x_2, x_7, x_8, x_9, x_10, x_18); -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; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; uint8_t x_47; lean_object* x_48; lean_object* x_49; uint8_t x_50; lean_object* x_51; lean_object* x_52; -x_20 = lean_ctor_get(x_19, 0); lean_inc(x_20); -x_21 = lean_ctor_get(x_19, 1); -lean_inc(x_21); -lean_dec(x_19); -x_22 = l_Lean_Elab_Term_instInhabitedCalcStepView; -x_23 = l_Array_back_x21___rarg(x_22, x_4); -x_24 = lean_ctor_get(x_23, 1); +x_22 = lean_infer_type(x_20, x_7, x_8, x_9, x_10, x_18); +if (lean_obj_tag(x_22) == 0) +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +x_24 = lean_ctor_get(x_22, 1); lean_inc(x_24); -lean_dec(x_23); -x_25 = l_Lean_MessageData_ofExpr(x_1); -x_26 = l___private_Lean_Elab_Calc_0__Lean_Elab_Term_getRelUniv___lambda__1___closed__4; -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_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Term_elabCalcSteps___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_MessageData_ofExpr(x_17); -x_31 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_31, 0, x_29); -lean_ctor_set(x_31, 1, x_30); -x_32 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_32, 0, x_31); -lean_ctor_set(x_32, 1, x_26); -x_33 = l_Lean_indentD(x_32); -x_34 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__3___closed__2; -x_35 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_35, 0, x_34); -lean_ctor_set(x_35, 1, x_33); -x_36 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__3___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_Lean_MessageData_ofExpr(x_2); -x_39 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_39, 0, x_26); -lean_ctor_set(x_39, 1, x_38); -x_40 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_40, 0, x_39); -lean_ctor_set(x_40, 1, x_28); -x_41 = l_Lean_MessageData_ofExpr(x_20); +lean_dec(x_22); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_21); +x_25 = lean_infer_type(x_21, x_7, x_8, x_9, x_10, x_24); +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); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +x_28 = l_Lean_Meta_addPPExplicitToExposeDiff(x_23, x_26, x_7, x_8, x_9, x_10, x_27); +if (lean_obj_tag(x_28) == 0) +{ +lean_object* x_29; lean_object* x_30; uint8_t x_31; +x_29 = lean_ctor_get(x_28, 0); +lean_inc(x_29); +x_30 = lean_ctor_get(x_28, 1); +lean_inc(x_30); +lean_dec(x_28); +x_31 = !lean_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; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; 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_object* x_59; uint8_t x_60; lean_object* x_61; lean_object* x_62; +x_32 = lean_ctor_get(x_29, 0); +x_33 = lean_ctor_get(x_29, 1); +x_34 = l_Lean_Elab_Term_instInhabitedCalcStepView; +x_35 = l_Array_back_x21___rarg(x_34, x_4); +x_36 = lean_ctor_get(x_35, 1); +lean_inc(x_36); +lean_dec(x_35); +x_37 = l_Lean_MessageData_ofExpr(x_20); +x_38 = l___private_Lean_Elab_Calc_0__Lean_Elab_Term_getRelUniv___lambda__1___closed__4; +lean_ctor_set_tag(x_29, 7); +lean_ctor_set(x_29, 1, x_37); +lean_ctor_set(x_29, 0, x_38); +x_39 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Term_elabCalcSteps___spec__1___closed__6; +lean_ctor_set_tag(x_17, 7); +lean_ctor_set(x_17, 1, x_39); +lean_ctor_set(x_17, 0, x_29); +x_40 = l_Lean_MessageData_ofExpr(x_32); +x_41 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_41, 0, x_17); +lean_ctor_set(x_41, 1, x_40); x_42 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_42, 0, x_40); -lean_ctor_set(x_42, 1, x_41); -x_43 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_43, 0, x_42); -lean_ctor_set(x_43, 1, x_26); -x_44 = l_Lean_indentD(x_43); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_38); +x_43 = l_Lean_indentD(x_42); +x_44 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__3___closed__2; x_45 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_45, 0, x_37); -lean_ctor_set(x_45, 1, x_44); -x_46 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_46, 0, x_45); -lean_ctor_set(x_46, 1, x_26); -x_47 = 2; +lean_ctor_set(x_45, 0, x_44); +lean_ctor_set(x_45, 1, x_43); +x_46 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__3___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_Lean_MessageData_ofExpr(x_21); +x_49 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_49, 0, x_38); +lean_ctor_set(x_49, 1, x_48); +x_50 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_50, 0, x_49); +lean_ctor_set(x_50, 1, x_39); +x_51 = l_Lean_MessageData_ofExpr(x_33); +x_52 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_52, 0, x_50); +lean_ctor_set(x_52, 1, x_51); +x_53 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_53, 0, x_52); +lean_ctor_set(x_53, 1, x_38); +x_54 = l_Lean_indentD(x_53); +x_55 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_55, 0, x_47); +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_38); +x_57 = 2; lean_inc(x_9); -x_48 = l_Lean_logAt___at_Lean_Elab_Term_reportUnsolvedGoals___spec__2(x_24, x_46, x_47, x_7, x_8, x_9, x_10, x_21); -lean_dec(x_24); -x_49 = lean_ctor_get(x_48, 1); -lean_inc(x_49); -lean_dec(x_48); -x_50 = 1; -x_51 = lean_box(0); -x_52 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__2(x_3, x_50, x_51, x_7, x_8, x_9, x_10, x_49); -return x_52; +x_58 = l_Lean_logAt___at_Lean_Elab_Term_reportUnsolvedGoals___spec__2(x_36, x_56, x_57, x_7, x_8, x_9, x_10, x_30); +lean_dec(x_36); +x_59 = lean_ctor_get(x_58, 1); +lean_inc(x_59); +lean_dec(x_58); +x_60 = 1; +x_61 = lean_box(0); +x_62 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__2(x_3, x_60, x_61, x_7, x_8, x_9, x_10, x_59); +return x_62; } else { -uint8_t x_53; +lean_object* x_63; 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; uint8_t x_89; lean_object* x_90; lean_object* x_91; uint8_t x_92; lean_object* x_93; lean_object* x_94; +x_63 = lean_ctor_get(x_29, 0); +x_64 = lean_ctor_get(x_29, 1); +lean_inc(x_64); +lean_inc(x_63); +lean_dec(x_29); +x_65 = l_Lean_Elab_Term_instInhabitedCalcStepView; +x_66 = l_Array_back_x21___rarg(x_65, x_4); +x_67 = lean_ctor_get(x_66, 1); +lean_inc(x_67); +lean_dec(x_66); +x_68 = l_Lean_MessageData_ofExpr(x_20); +x_69 = l___private_Lean_Elab_Calc_0__Lean_Elab_Term_getRelUniv___lambda__1___closed__4; +x_70 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_70, 0, x_69); +lean_ctor_set(x_70, 1, x_68); +x_71 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Term_elabCalcSteps___spec__1___closed__6; +lean_ctor_set_tag(x_17, 7); +lean_ctor_set(x_17, 1, x_71); +lean_ctor_set(x_17, 0, x_70); +x_72 = l_Lean_MessageData_ofExpr(x_63); +x_73 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_73, 0, x_17); +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_69); +x_75 = l_Lean_indentD(x_74); +x_76 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__3___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_Term_throwCalcFailure___rarg___lambda__3___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_MessageData_ofExpr(x_21); +x_81 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_81, 0, x_69); +lean_ctor_set(x_81, 1, x_80); +x_82 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_82, 0, x_81); +lean_ctor_set(x_82, 1, x_71); +x_83 = l_Lean_MessageData_ofExpr(x_64); +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 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_85, 0, x_84); +lean_ctor_set(x_85, 1, x_69); +x_86 = l_Lean_indentD(x_85); +x_87 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_87, 0, x_79); +lean_ctor_set(x_87, 1, x_86); +x_88 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_88, 0, x_87); +lean_ctor_set(x_88, 1, x_69); +x_89 = 2; +lean_inc(x_9); +x_90 = l_Lean_logAt___at_Lean_Elab_Term_reportUnsolvedGoals___spec__2(x_67, x_88, x_89, x_7, x_8, x_9, x_10, x_30); +lean_dec(x_67); +x_91 = lean_ctor_get(x_90, 1); +lean_inc(x_91); +lean_dec(x_90); +x_92 = 1; +x_93 = lean_box(0); +x_94 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__2(x_3, x_92, x_93, x_7, x_8, x_9, x_10, x_91); +return x_94; +} +} +else +{ +uint8_t x_95; +lean_free_object(x_17); +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_3); +x_95 = !lean_is_exclusive(x_28); +if (x_95 == 0) +{ +return x_28; +} +else +{ +lean_object* x_96; lean_object* x_97; lean_object* x_98; +x_96 = lean_ctor_get(x_28, 0); +x_97 = lean_ctor_get(x_28, 1); +lean_inc(x_97); +lean_inc(x_96); +lean_dec(x_28); +x_98 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_98, 0, x_96); +lean_ctor_set(x_98, 1, x_97); +return x_98; +} +} +} +else +{ +uint8_t x_99; +lean_dec(x_23); +lean_free_object(x_17); +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_3); +x_99 = !lean_is_exclusive(x_25); +if (x_99 == 0) +{ +return x_25; +} +else +{ +lean_object* x_100; lean_object* x_101; lean_object* x_102; +x_100 = lean_ctor_get(x_25, 0); +x_101 = lean_ctor_get(x_25, 1); +lean_inc(x_101); +lean_inc(x_100); +lean_dec(x_25); +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_103; +lean_free_object(x_17); +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_3); +x_103 = !lean_is_exclusive(x_22); +if (x_103 == 0) +{ +return x_22; +} +else +{ +lean_object* x_104; lean_object* x_105; lean_object* x_106; +x_104 = lean_ctor_get(x_22, 0); +x_105 = lean_ctor_get(x_22, 1); +lean_inc(x_105); +lean_inc(x_104); +lean_dec(x_22); +x_106 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_106, 0, x_104); +lean_ctor_set(x_106, 1, x_105); +return x_106; +} +} +} +else +{ +lean_object* x_107; lean_object* x_108; lean_object* x_109; +x_107 = lean_ctor_get(x_17, 0); +x_108 = lean_ctor_get(x_17, 1); +lean_inc(x_108); +lean_inc(x_107); lean_dec(x_17); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_107); +x_109 = lean_infer_type(x_107, x_7, x_8, x_9, x_10, x_18); +if (lean_obj_tag(x_109) == 0) +{ +lean_object* x_110; lean_object* x_111; lean_object* x_112; +x_110 = lean_ctor_get(x_109, 0); +lean_inc(x_110); +x_111 = lean_ctor_get(x_109, 1); +lean_inc(x_111); +lean_dec(x_109); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_108); +x_112 = lean_infer_type(x_108, x_7, x_8, x_9, x_10, x_111); +if (lean_obj_tag(x_112) == 0) +{ +lean_object* x_113; lean_object* x_114; lean_object* x_115; +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); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +x_115 = l_Lean_Meta_addPPExplicitToExposeDiff(x_110, x_113, x_7, x_8, x_9, x_10, x_114); +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; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; uint8_t x_146; lean_object* x_147; lean_object* x_148; uint8_t x_149; lean_object* x_150; lean_object* x_151; +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 = lean_ctor_get(x_116, 1); +lean_inc(x_119); +if (lean_is_exclusive(x_116)) { + lean_ctor_release(x_116, 0); + lean_ctor_release(x_116, 1); + x_120 = x_116; +} else { + lean_dec_ref(x_116); + x_120 = lean_box(0); +} +x_121 = l_Lean_Elab_Term_instInhabitedCalcStepView; +x_122 = l_Array_back_x21___rarg(x_121, x_4); +x_123 = lean_ctor_get(x_122, 1); +lean_inc(x_123); +lean_dec(x_122); +x_124 = l_Lean_MessageData_ofExpr(x_107); +x_125 = l___private_Lean_Elab_Calc_0__Lean_Elab_Term_getRelUniv___lambda__1___closed__4; +if (lean_is_scalar(x_120)) { + x_126 = lean_alloc_ctor(7, 2, 0); +} else { + x_126 = x_120; + lean_ctor_set_tag(x_126, 7); +} +lean_ctor_set(x_126, 0, x_125); +lean_ctor_set(x_126, 1, x_124); +x_127 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Term_elabCalcSteps___spec__1___closed__6; +x_128 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_128, 0, x_126); +lean_ctor_set(x_128, 1, x_127); +x_129 = l_Lean_MessageData_ofExpr(x_118); +x_130 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_130, 0, x_128); +lean_ctor_set(x_130, 1, x_129); +x_131 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_131, 0, x_130); +lean_ctor_set(x_131, 1, x_125); +x_132 = l_Lean_indentD(x_131); +x_133 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__3___closed__2; +x_134 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_134, 0, x_133); +lean_ctor_set(x_134, 1, x_132); +x_135 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__3___closed__4; +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_MessageData_ofExpr(x_108); +x_138 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_138, 0, x_125); +lean_ctor_set(x_138, 1, x_137); +x_139 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_139, 0, x_138); +lean_ctor_set(x_139, 1, x_127); +x_140 = l_Lean_MessageData_ofExpr(x_119); +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 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_142, 0, x_141); +lean_ctor_set(x_142, 1, x_125); +x_143 = l_Lean_indentD(x_142); +x_144 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_144, 0, x_136); +lean_ctor_set(x_144, 1, x_143); +x_145 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_145, 0, x_144); +lean_ctor_set(x_145, 1, x_125); +x_146 = 2; +lean_inc(x_9); +x_147 = l_Lean_logAt___at_Lean_Elab_Term_reportUnsolvedGoals___spec__2(x_123, x_145, x_146, x_7, x_8, x_9, x_10, x_117); +lean_dec(x_123); +x_148 = lean_ctor_get(x_147, 1); +lean_inc(x_148); +lean_dec(x_147); +x_149 = 1; +x_150 = lean_box(0); +x_151 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__2(x_3, x_149, x_150, x_7, x_8, x_9, x_10, x_148); +return x_151; +} +else +{ +lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; +lean_dec(x_108); +lean_dec(x_107); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_53 = !lean_is_exclusive(x_19); -if (x_53 == 0) +x_152 = lean_ctor_get(x_115, 0); +lean_inc(x_152); +x_153 = lean_ctor_get(x_115, 1); +lean_inc(x_153); +if (lean_is_exclusive(x_115)) { + lean_ctor_release(x_115, 0); + lean_ctor_release(x_115, 1); + x_154 = x_115; +} else { + lean_dec_ref(x_115); + x_154 = lean_box(0); +} +if (lean_is_scalar(x_154)) { + x_155 = lean_alloc_ctor(1, 2, 0); +} else { + x_155 = x_154; +} +lean_ctor_set(x_155, 0, x_152); +lean_ctor_set(x_155, 1, x_153); +return x_155; +} +} +else { -return x_19; +lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; +lean_dec(x_110); +lean_dec(x_108); +lean_dec(x_107); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_3); +x_156 = lean_ctor_get(x_112, 0); +lean_inc(x_156); +x_157 = lean_ctor_get(x_112, 1); +lean_inc(x_157); +if (lean_is_exclusive(x_112)) { + lean_ctor_release(x_112, 0); + lean_ctor_release(x_112, 1); + x_158 = x_112; +} else { + lean_dec_ref(x_112); + x_158 = lean_box(0); +} +if (lean_is_scalar(x_158)) { + x_159 = lean_alloc_ctor(1, 2, 0); +} else { + x_159 = x_158; +} +lean_ctor_set(x_159, 0, x_156); +lean_ctor_set(x_159, 1, x_157); +return x_159; +} } else { -lean_object* x_54; lean_object* x_55; lean_object* x_56; -x_54 = lean_ctor_get(x_19, 0); -x_55 = lean_ctor_get(x_19, 1); -lean_inc(x_55); -lean_inc(x_54); -lean_dec(x_19); -x_56 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_56, 0, x_54); -lean_ctor_set(x_56, 1, x_55); -return x_56; +lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; +lean_dec(x_108); +lean_dec(x_107); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_3); +x_160 = lean_ctor_get(x_109, 0); +lean_inc(x_160); +x_161 = lean_ctor_get(x_109, 1); +lean_inc(x_161); +if (lean_is_exclusive(x_109)) { + lean_ctor_release(x_109, 0); + lean_ctor_release(x_109, 1); + x_162 = x_109; +} else { + lean_dec_ref(x_109); + x_162 = lean_box(0); +} +if (lean_is_scalar(x_162)) { + x_163 = lean_alloc_ctor(1, 2, 0); +} else { + x_163 = x_162; +} +lean_ctor_set(x_163, 0, x_160); +lean_ctor_set(x_163, 1, x_161); +return x_163; } } } else { -uint8_t x_57; +uint8_t x_164; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_57 = !lean_is_exclusive(x_16); -if (x_57 == 0) +x_164 = !lean_is_exclusive(x_16); +if (x_164 == 0) { return x_16; } else { -lean_object* x_58; lean_object* x_59; lean_object* x_60; -x_58 = lean_ctor_get(x_16, 0); -x_59 = lean_ctor_get(x_16, 1); -lean_inc(x_59); -lean_inc(x_58); +lean_object* x_165; lean_object* x_166; lean_object* x_167; +x_165 = lean_ctor_get(x_16, 0); +x_166 = lean_ctor_get(x_16, 1); +lean_inc(x_166); +lean_inc(x_165); lean_dec(x_16); -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_167 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_167, 0, x_165); +lean_ctor_set(x_167, 1, x_166); +return x_167; } } } else { -lean_object* x_61; lean_object* x_62; lean_object* x_63; +lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_dec(x_2); lean_dec(x_1); -x_61 = lean_ctor_get(x_12, 1); -lean_inc(x_61); +x_168 = lean_ctor_get(x_12, 1); +lean_inc(x_168); lean_dec(x_12); -x_62 = lean_box(0); -x_63 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__2(x_3, x_5, x_62, x_7, x_8, x_9, x_10, x_61); -return x_63; +x_169 = lean_box(0); +x_170 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__2(x_3, x_5, x_169, x_7, x_8, x_9, x_10, x_168); +return x_170; } } else { -uint8_t x_64; +uint8_t x_171; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -19810,23 +20209,23 @@ lean_dec(x_7); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_64 = !lean_is_exclusive(x_12); -if (x_64 == 0) +x_171 = !lean_is_exclusive(x_12); +if (x_171 == 0) { return x_12; } else { -lean_object* x_65; lean_object* x_66; lean_object* x_67; -x_65 = lean_ctor_get(x_12, 0); -x_66 = lean_ctor_get(x_12, 1); -lean_inc(x_66); -lean_inc(x_65); +lean_object* x_172; lean_object* x_173; lean_object* x_174; +x_172 = lean_ctor_get(x_12, 0); +x_173 = lean_ctor_get(x_12, 1); +lean_inc(x_173); +lean_inc(x_172); lean_dec(x_12); -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; +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; } } } @@ -20042,133 +20441,546 @@ lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -lean_inc(x_31); -x_60 = lean_infer_type(x_31, x_4, x_5, x_6, x_7, x_59); +x_60 = l_Lean_Meta_addPPExplicitToExposeDiff(x_31, x_47, x_4, x_5, x_6, x_7, x_59); if (lean_obj_tag(x_60) == 0) { -lean_object* x_61; lean_object* x_62; lean_object* x_63; +lean_object* x_61; lean_object* x_62; uint8_t x_63; x_61 = lean_ctor_get(x_60, 0); lean_inc(x_61); x_62 = lean_ctor_get(x_60, 1); lean_inc(x_62); lean_dec(x_60); +x_63 = !lean_is_exclusive(x_61); +if (x_63 == 0) +{ +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_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_64); +x_66 = lean_infer_type(x_64, x_4, x_5, x_6, x_7, x_62); +if (lean_obj_tag(x_66) == 0) +{ +lean_object* x_67; lean_object* x_68; lean_object* x_69; +x_67 = lean_ctor_get(x_66, 0); +lean_inc(x_67); +x_68 = lean_ctor_get(x_66, 1); +lean_inc(x_68); +lean_dec(x_66); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_65); +x_69 = lean_infer_type(x_65, x_4, x_5, x_6, x_7, x_68); +if (lean_obj_tag(x_69) == 0) +{ +lean_object* x_70; lean_object* x_71; lean_object* x_72; +x_70 = lean_ctor_get(x_69, 0); +lean_inc(x_70); +x_71 = lean_ctor_get(x_69, 1); +lean_inc(x_71); +lean_dec(x_69); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +x_72 = l_Lean_Meta_addPPExplicitToExposeDiff(x_67, x_70, x_4, x_5, x_6, x_7, x_71); +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 = !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; uint8_t x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; +x_76 = lean_ctor_get(x_73, 0); +x_77 = lean_ctor_get(x_73, 1); +x_78 = lean_array_get_size(x_1); +x_79 = lean_unsigned_to_nat(0u); +x_80 = lean_nat_dec_lt(x_79, x_78); +lean_dec(x_78); +x_81 = l_Lean_MessageData_ofExpr(x_64); +x_82 = l___private_Lean_Elab_Calc_0__Lean_Elab_Term_getRelUniv___lambda__1___closed__4; +lean_ctor_set_tag(x_73, 7); +lean_ctor_set(x_73, 1, x_81); +lean_ctor_set(x_73, 0, x_82); +x_83 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Term_elabCalcSteps___spec__1___closed__6; +lean_ctor_set_tag(x_61, 7); +lean_ctor_set(x_61, 1, x_83); +lean_ctor_set(x_61, 0, x_73); +x_84 = l_Lean_MessageData_ofExpr(x_76); +lean_ctor_set_tag(x_44, 7); +lean_ctor_set(x_44, 1, x_84); +lean_ctor_set(x_44, 0, x_61); +lean_ctor_set_tag(x_41, 7); +lean_ctor_set(x_41, 1, x_82); +lean_ctor_set(x_41, 0, x_44); +x_85 = l_Lean_indentD(x_41); +x_86 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Term_elabCalcSteps___spec__1___closed__4; +lean_ctor_set_tag(x_33, 7); +lean_ctor_set(x_33, 1, x_85); +lean_ctor_set(x_33, 0, x_86); +x_87 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__3___closed__4; +lean_ctor_set_tag(x_23, 7); +lean_ctor_set(x_23, 1, x_87); +lean_ctor_set(x_23, 0, x_33); +x_88 = l_Lean_MessageData_ofExpr(x_65); +lean_ctor_set_tag(x_22, 7); +lean_ctor_set(x_22, 1, x_88); +lean_ctor_set(x_22, 0, x_82); +lean_ctor_set_tag(x_17, 7); +lean_ctor_set(x_17, 1, x_83); +lean_ctor_set(x_17, 0, x_22); +x_89 = l_Lean_MessageData_ofExpr(x_77); +lean_ctor_set_tag(x_12, 7); +lean_ctor_set(x_12, 1, x_89); +lean_ctor_set(x_12, 0, x_17); +x_90 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_90, 0, x_12); +lean_ctor_set(x_90, 1, x_82); +x_91 = l_Lean_indentD(x_90); +x_92 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_92, 0, x_23); +lean_ctor_set(x_92, 1, x_91); +x_93 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_93, 0, x_92); +lean_ctor_set(x_93, 1, x_82); +if (x_80 == 0) +{ +lean_object* x_94; lean_object* x_95; lean_object* x_96; uint8_t x_97; lean_object* x_98; lean_object* x_99; uint8_t x_100; lean_object* x_101; lean_object* x_102; +x_94 = l_Lean_Elab_Term_instInhabitedCalcStepView; +x_95 = l_outOfBounds___rarg(x_94); +x_96 = lean_ctor_get(x_95, 1); +lean_inc(x_96); +lean_dec(x_95); +x_97 = 2; +lean_inc(x_6); +x_98 = l_Lean_logAt___at_Lean_Elab_Term_reportUnsolvedGoals___spec__2(x_96, x_93, x_97, x_4, x_5, x_6, x_7, x_74); +lean_dec(x_96); +x_99 = lean_ctor_get(x_98, 1); +lean_inc(x_99); +lean_dec(x_98); +x_100 = 1; +x_101 = lean_box(0); +x_102 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__3(x_32, x_48, x_42, x_1, x_100, x_101, x_4, x_5, x_6, x_7, x_99); +return x_102; +} +else +{ +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; +x_103 = lean_array_fget(x_1, x_79); +x_104 = lean_ctor_get(x_103, 1); +lean_inc(x_104); +lean_dec(x_103); +x_105 = 2; +lean_inc(x_6); +x_106 = l_Lean_logAt___at_Lean_Elab_Term_reportUnsolvedGoals___spec__2(x_104, x_93, x_105, x_4, x_5, x_6, x_7, x_74); +lean_dec(x_104); +x_107 = lean_ctor_get(x_106, 1); +lean_inc(x_107); +lean_dec(x_106); +x_108 = 1; +x_109 = lean_box(0); +x_110 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__3(x_32, x_48, x_42, x_1, x_108, x_109, x_4, x_5, x_6, x_7, x_107); +return x_110; +} +} +else +{ +lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; uint8_t x_115; lean_object* x_116; lean_object* x_117; 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_111 = lean_ctor_get(x_73, 0); +x_112 = lean_ctor_get(x_73, 1); +lean_inc(x_112); +lean_inc(x_111); +lean_dec(x_73); +x_113 = lean_array_get_size(x_1); +x_114 = lean_unsigned_to_nat(0u); +x_115 = lean_nat_dec_lt(x_114, x_113); +lean_dec(x_113); +x_116 = l_Lean_MessageData_ofExpr(x_64); +x_117 = l___private_Lean_Elab_Calc_0__Lean_Elab_Term_getRelUniv___lambda__1___closed__4; +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_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Term_elabCalcSteps___spec__1___closed__6; +lean_ctor_set_tag(x_61, 7); +lean_ctor_set(x_61, 1, x_119); +lean_ctor_set(x_61, 0, x_118); +x_120 = l_Lean_MessageData_ofExpr(x_111); +lean_ctor_set_tag(x_44, 7); +lean_ctor_set(x_44, 1, x_120); +lean_ctor_set(x_44, 0, x_61); +lean_ctor_set_tag(x_41, 7); +lean_ctor_set(x_41, 1, x_117); +lean_ctor_set(x_41, 0, x_44); +x_121 = l_Lean_indentD(x_41); +x_122 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Term_elabCalcSteps___spec__1___closed__4; +lean_ctor_set_tag(x_33, 7); +lean_ctor_set(x_33, 1, x_121); +lean_ctor_set(x_33, 0, x_122); +x_123 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__3___closed__4; +lean_ctor_set_tag(x_23, 7); +lean_ctor_set(x_23, 1, x_123); +lean_ctor_set(x_23, 0, x_33); +x_124 = l_Lean_MessageData_ofExpr(x_65); +lean_ctor_set_tag(x_22, 7); +lean_ctor_set(x_22, 1, x_124); +lean_ctor_set(x_22, 0, x_117); +lean_ctor_set_tag(x_17, 7); +lean_ctor_set(x_17, 1, x_119); +lean_ctor_set(x_17, 0, x_22); +x_125 = l_Lean_MessageData_ofExpr(x_112); +lean_ctor_set_tag(x_12, 7); +lean_ctor_set(x_12, 1, x_125); +lean_ctor_set(x_12, 0, x_17); +x_126 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_126, 0, x_12); +lean_ctor_set(x_126, 1, x_117); +x_127 = l_Lean_indentD(x_126); +x_128 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_128, 0, x_23); +lean_ctor_set(x_128, 1, x_127); +x_129 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_129, 0, x_128); +lean_ctor_set(x_129, 1, x_117); +if (x_115 == 0) +{ +lean_object* x_130; lean_object* x_131; lean_object* x_132; uint8_t x_133; lean_object* x_134; lean_object* x_135; uint8_t x_136; lean_object* x_137; lean_object* x_138; +x_130 = l_Lean_Elab_Term_instInhabitedCalcStepView; +x_131 = l_outOfBounds___rarg(x_130); +x_132 = lean_ctor_get(x_131, 1); +lean_inc(x_132); +lean_dec(x_131); +x_133 = 2; +lean_inc(x_6); +x_134 = l_Lean_logAt___at_Lean_Elab_Term_reportUnsolvedGoals___spec__2(x_132, x_129, x_133, x_4, x_5, x_6, x_7, x_74); +lean_dec(x_132); +x_135 = lean_ctor_get(x_134, 1); +lean_inc(x_135); +lean_dec(x_134); +x_136 = 1; +x_137 = lean_box(0); +x_138 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__3(x_32, x_48, x_42, x_1, x_136, x_137, x_4, x_5, x_6, x_7, x_135); +return x_138; +} +else +{ +lean_object* x_139; lean_object* x_140; uint8_t x_141; lean_object* x_142; lean_object* x_143; uint8_t x_144; lean_object* x_145; lean_object* x_146; +x_139 = lean_array_fget(x_1, x_114); +x_140 = lean_ctor_get(x_139, 1); +lean_inc(x_140); +lean_dec(x_139); +x_141 = 2; +lean_inc(x_6); +x_142 = l_Lean_logAt___at_Lean_Elab_Term_reportUnsolvedGoals___spec__2(x_140, x_129, x_141, x_4, x_5, x_6, x_7, x_74); +lean_dec(x_140); +x_143 = lean_ctor_get(x_142, 1); +lean_inc(x_143); +lean_dec(x_142); +x_144 = 1; +x_145 = lean_box(0); +x_146 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__3(x_32, x_48, x_42, x_1, x_144, x_145, x_4, x_5, x_6, x_7, x_143); +return x_146; +} +} +} +else +{ +uint8_t x_147; +lean_free_object(x_61); +lean_dec(x_65); +lean_dec(x_64); +lean_free_object(x_44); +lean_dec(x_48); +lean_free_object(x_41); +lean_dec(x_42); +lean_free_object(x_33); +lean_free_object(x_23); +lean_dec(x_32); +lean_free_object(x_22); +lean_free_object(x_17); +lean_free_object(x_12); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_147 = !lean_is_exclusive(x_72); +if (x_147 == 0) +{ +return x_72; +} +else +{ +lean_object* x_148; lean_object* x_149; lean_object* x_150; +x_148 = lean_ctor_get(x_72, 0); +x_149 = lean_ctor_get(x_72, 1); +lean_inc(x_149); +lean_inc(x_148); +lean_dec(x_72); +x_150 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_150, 0, x_148); +lean_ctor_set(x_150, 1, x_149); +return x_150; +} +} +} +else +{ +uint8_t x_151; +lean_dec(x_67); +lean_free_object(x_61); +lean_dec(x_65); +lean_dec(x_64); +lean_free_object(x_44); +lean_dec(x_48); +lean_free_object(x_41); +lean_dec(x_42); +lean_free_object(x_33); +lean_free_object(x_23); +lean_dec(x_32); +lean_free_object(x_22); +lean_free_object(x_17); +lean_free_object(x_12); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_151 = !lean_is_exclusive(x_69); +if (x_151 == 0) +{ +return x_69; +} +else +{ +lean_object* x_152; lean_object* x_153; lean_object* x_154; +x_152 = lean_ctor_get(x_69, 0); +x_153 = lean_ctor_get(x_69, 1); +lean_inc(x_153); +lean_inc(x_152); +lean_dec(x_69); +x_154 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_154, 0, x_152); +lean_ctor_set(x_154, 1, x_153); +return x_154; +} +} +} +else +{ +uint8_t x_155; +lean_free_object(x_61); +lean_dec(x_65); +lean_dec(x_64); +lean_free_object(x_44); +lean_dec(x_48); +lean_free_object(x_41); +lean_dec(x_42); +lean_free_object(x_33); +lean_free_object(x_23); +lean_dec(x_32); +lean_free_object(x_22); +lean_free_object(x_17); +lean_free_object(x_12); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_155 = !lean_is_exclusive(x_66); +if (x_155 == 0) +{ +return x_66; +} +else +{ +lean_object* x_156; lean_object* x_157; lean_object* x_158; +x_156 = lean_ctor_get(x_66, 0); +x_157 = lean_ctor_get(x_66, 1); +lean_inc(x_157); +lean_inc(x_156); +lean_dec(x_66); +x_158 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_158, 0, x_156); +lean_ctor_set(x_158, 1, x_157); +return x_158; +} +} +} +else +{ +lean_object* x_159; lean_object* x_160; lean_object* x_161; +x_159 = lean_ctor_get(x_61, 0); +x_160 = lean_ctor_get(x_61, 1); +lean_inc(x_160); +lean_inc(x_159); +lean_dec(x_61); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_159); +x_161 = lean_infer_type(x_159, x_4, x_5, x_6, x_7, x_62); +if (lean_obj_tag(x_161) == 0) +{ +lean_object* x_162; lean_object* x_163; lean_object* x_164; +x_162 = lean_ctor_get(x_161, 0); +lean_inc(x_162); +x_163 = lean_ctor_get(x_161, 1); +lean_inc(x_163); +lean_dec(x_161); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_160); +x_164 = lean_infer_type(x_160, x_4, x_5, x_6, x_7, x_163); +if (lean_obj_tag(x_164) == 0) +{ +lean_object* x_165; lean_object* x_166; lean_object* x_167; +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); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -lean_inc(x_47); -x_63 = lean_infer_type(x_47, x_4, x_5, x_6, x_7, x_62); -if (lean_obj_tag(x_63) == 0) +x_167 = l_Lean_Meta_addPPExplicitToExposeDiff(x_162, x_165, x_4, x_5, x_6, x_7, x_166); +if (lean_obj_tag(x_167) == 0) { -lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; uint8_t x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; -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_array_get_size(x_1); -x_67 = lean_unsigned_to_nat(0u); -x_68 = lean_nat_dec_lt(x_67, x_66); -lean_dec(x_66); -x_69 = l_Lean_MessageData_ofExpr(x_31); -x_70 = l___private_Lean_Elab_Calc_0__Lean_Elab_Term_getRelUniv___lambda__1___closed__4; +lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; uint8_t 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_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_168, 0); +lean_inc(x_170); +x_171 = lean_ctor_get(x_168, 1); +lean_inc(x_171); +if (lean_is_exclusive(x_168)) { + lean_ctor_release(x_168, 0); + lean_ctor_release(x_168, 1); + x_172 = x_168; +} else { + lean_dec_ref(x_168); + x_172 = lean_box(0); +} +x_173 = lean_array_get_size(x_1); +x_174 = lean_unsigned_to_nat(0u); +x_175 = lean_nat_dec_lt(x_174, x_173); +lean_dec(x_173); +x_176 = l_Lean_MessageData_ofExpr(x_159); +x_177 = l___private_Lean_Elab_Calc_0__Lean_Elab_Term_getRelUniv___lambda__1___closed__4; +if (lean_is_scalar(x_172)) { + x_178 = lean_alloc_ctor(7, 2, 0); +} else { + x_178 = x_172; + lean_ctor_set_tag(x_178, 7); +} +lean_ctor_set(x_178, 0, x_177); +lean_ctor_set(x_178, 1, x_176); +x_179 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Term_elabCalcSteps___spec__1___closed__6; +x_180 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_180, 0, x_178); +lean_ctor_set(x_180, 1, x_179); +x_181 = l_Lean_MessageData_ofExpr(x_170); lean_ctor_set_tag(x_44, 7); -lean_ctor_set(x_44, 1, x_69); -lean_ctor_set(x_44, 0, x_70); -x_71 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Term_elabCalcSteps___spec__1___closed__6; +lean_ctor_set(x_44, 1, x_181); +lean_ctor_set(x_44, 0, x_180); lean_ctor_set_tag(x_41, 7); -lean_ctor_set(x_41, 1, x_71); +lean_ctor_set(x_41, 1, x_177); lean_ctor_set(x_41, 0, x_44); -x_72 = l_Lean_MessageData_ofExpr(x_61); +x_182 = l_Lean_indentD(x_41); +x_183 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Term_elabCalcSteps___spec__1___closed__4; lean_ctor_set_tag(x_33, 7); -lean_ctor_set(x_33, 1, x_72); -lean_ctor_set(x_33, 0, x_41); +lean_ctor_set(x_33, 1, x_182); +lean_ctor_set(x_33, 0, x_183); +x_184 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__3___closed__4; lean_ctor_set_tag(x_23, 7); -lean_ctor_set(x_23, 1, x_70); +lean_ctor_set(x_23, 1, x_184); lean_ctor_set(x_23, 0, x_33); -x_73 = l_Lean_indentD(x_23); -x_74 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Term_elabCalcSteps___spec__1___closed__4; +x_185 = l_Lean_MessageData_ofExpr(x_160); lean_ctor_set_tag(x_22, 7); -lean_ctor_set(x_22, 1, x_73); -lean_ctor_set(x_22, 0, x_74); -x_75 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__3___closed__4; +lean_ctor_set(x_22, 1, x_185); +lean_ctor_set(x_22, 0, x_177); lean_ctor_set_tag(x_17, 7); -lean_ctor_set(x_17, 1, x_75); +lean_ctor_set(x_17, 1, x_179); lean_ctor_set(x_17, 0, x_22); -x_76 = l_Lean_MessageData_ofExpr(x_47); +x_186 = l_Lean_MessageData_ofExpr(x_171); lean_ctor_set_tag(x_12, 7); -lean_ctor_set(x_12, 1, x_76); -lean_ctor_set(x_12, 0, x_70); -x_77 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_77, 0, x_12); -lean_ctor_set(x_77, 1, x_71); -x_78 = l_Lean_MessageData_ofExpr(x_64); -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 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_80, 0, x_79); -lean_ctor_set(x_80, 1, x_70); -x_81 = l_Lean_indentD(x_80); -x_82 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_82, 0, x_17); -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_70); -if (x_68 == 0) -{ -lean_object* x_84; lean_object* x_85; lean_object* x_86; uint8_t x_87; lean_object* x_88; lean_object* x_89; uint8_t x_90; lean_object* x_91; lean_object* x_92; -x_84 = l_Lean_Elab_Term_instInhabitedCalcStepView; -x_85 = l_outOfBounds___rarg(x_84); -x_86 = lean_ctor_get(x_85, 1); -lean_inc(x_86); -lean_dec(x_85); -x_87 = 2; +lean_ctor_set(x_12, 1, x_186); +lean_ctor_set(x_12, 0, x_17); +x_187 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_187, 0, x_12); +lean_ctor_set(x_187, 1, x_177); +x_188 = l_Lean_indentD(x_187); +x_189 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_189, 0, x_23); +lean_ctor_set(x_189, 1, x_188); +x_190 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_190, 0, x_189); +lean_ctor_set(x_190, 1, x_177); +if (x_175 == 0) +{ +lean_object* x_191; lean_object* x_192; lean_object* x_193; uint8_t x_194; lean_object* x_195; lean_object* x_196; uint8_t x_197; lean_object* x_198; lean_object* x_199; +x_191 = l_Lean_Elab_Term_instInhabitedCalcStepView; +x_192 = l_outOfBounds___rarg(x_191); +x_193 = lean_ctor_get(x_192, 1); +lean_inc(x_193); +lean_dec(x_192); +x_194 = 2; lean_inc(x_6); -x_88 = l_Lean_logAt___at_Lean_Elab_Term_reportUnsolvedGoals___spec__2(x_86, x_83, x_87, x_4, x_5, x_6, x_7, x_65); -lean_dec(x_86); -x_89 = lean_ctor_get(x_88, 1); -lean_inc(x_89); -lean_dec(x_88); -x_90 = 1; -x_91 = lean_box(0); -x_92 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__3(x_32, x_48, x_42, x_1, x_90, x_91, x_4, x_5, x_6, x_7, x_89); -return x_92; +x_195 = l_Lean_logAt___at_Lean_Elab_Term_reportUnsolvedGoals___spec__2(x_193, x_190, x_194, x_4, x_5, x_6, x_7, x_169); +lean_dec(x_193); +x_196 = lean_ctor_get(x_195, 1); +lean_inc(x_196); +lean_dec(x_195); +x_197 = 1; +x_198 = lean_box(0); +x_199 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__3(x_32, x_48, x_42, x_1, x_197, x_198, x_4, x_5, x_6, x_7, x_196); +return x_199; } else { -lean_object* x_93; lean_object* x_94; uint8_t x_95; lean_object* x_96; lean_object* x_97; uint8_t x_98; lean_object* x_99; lean_object* x_100; -x_93 = lean_array_fget(x_1, x_67); -x_94 = lean_ctor_get(x_93, 1); -lean_inc(x_94); -lean_dec(x_93); -x_95 = 2; +lean_object* x_200; lean_object* x_201; uint8_t x_202; lean_object* x_203; lean_object* x_204; uint8_t x_205; lean_object* x_206; lean_object* x_207; +x_200 = lean_array_fget(x_1, x_174); +x_201 = lean_ctor_get(x_200, 1); +lean_inc(x_201); +lean_dec(x_200); +x_202 = 2; lean_inc(x_6); -x_96 = l_Lean_logAt___at_Lean_Elab_Term_reportUnsolvedGoals___spec__2(x_94, x_83, x_95, x_4, x_5, x_6, x_7, x_65); -lean_dec(x_94); -x_97 = lean_ctor_get(x_96, 1); -lean_inc(x_97); -lean_dec(x_96); -x_98 = 1; -x_99 = lean_box(0); -x_100 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__3(x_32, x_48, x_42, x_1, x_98, x_99, x_4, x_5, x_6, x_7, x_97); -return x_100; +x_203 = l_Lean_logAt___at_Lean_Elab_Term_reportUnsolvedGoals___spec__2(x_201, x_190, x_202, x_4, x_5, x_6, x_7, x_169); +lean_dec(x_201); +x_204 = lean_ctor_get(x_203, 1); +lean_inc(x_204); +lean_dec(x_203); +x_205 = 1; +x_206 = lean_box(0); +x_207 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__3(x_32, x_48, x_42, x_1, x_205, x_206, x_4, x_5, x_6, x_7, x_204); +return x_207; } } else { -uint8_t x_101; -lean_dec(x_61); +lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; +lean_dec(x_160); +lean_dec(x_159); lean_free_object(x_44); lean_dec(x_48); -lean_dec(x_47); lean_free_object(x_41); lean_dec(x_42); lean_free_object(x_33); lean_free_object(x_23); lean_dec(x_32); -lean_dec(x_31); lean_free_object(x_22); lean_free_object(x_17); lean_free_object(x_12); @@ -20176,38 +20988,122 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_101 = !lean_is_exclusive(x_63); -if (x_101 == 0) +x_208 = lean_ctor_get(x_167, 0); +lean_inc(x_208); +x_209 = lean_ctor_get(x_167, 1); +lean_inc(x_209); +if (lean_is_exclusive(x_167)) { + lean_ctor_release(x_167, 0); + lean_ctor_release(x_167, 1); + x_210 = x_167; +} else { + lean_dec_ref(x_167); + x_210 = lean_box(0); +} +if (lean_is_scalar(x_210)) { + x_211 = lean_alloc_ctor(1, 2, 0); +} else { + x_211 = x_210; +} +lean_ctor_set(x_211, 0, x_208); +lean_ctor_set(x_211, 1, x_209); +return x_211; +} +} +else { -return x_63; +lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; +lean_dec(x_162); +lean_dec(x_160); +lean_dec(x_159); +lean_free_object(x_44); +lean_dec(x_48); +lean_free_object(x_41); +lean_dec(x_42); +lean_free_object(x_33); +lean_free_object(x_23); +lean_dec(x_32); +lean_free_object(x_22); +lean_free_object(x_17); +lean_free_object(x_12); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_212 = lean_ctor_get(x_164, 0); +lean_inc(x_212); +x_213 = lean_ctor_get(x_164, 1); +lean_inc(x_213); +if (lean_is_exclusive(x_164)) { + lean_ctor_release(x_164, 0); + lean_ctor_release(x_164, 1); + x_214 = x_164; +} else { + lean_dec_ref(x_164); + x_214 = lean_box(0); +} +if (lean_is_scalar(x_214)) { + x_215 = lean_alloc_ctor(1, 2, 0); +} else { + x_215 = x_214; +} +lean_ctor_set(x_215, 0, x_212); +lean_ctor_set(x_215, 1, x_213); +return x_215; +} } else { -lean_object* x_102; lean_object* x_103; lean_object* x_104; -x_102 = lean_ctor_get(x_63, 0); -x_103 = lean_ctor_get(x_63, 1); -lean_inc(x_103); -lean_inc(x_102); -lean_dec(x_63); -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_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; +lean_dec(x_160); +lean_dec(x_159); +lean_free_object(x_44); +lean_dec(x_48); +lean_free_object(x_41); +lean_dec(x_42); +lean_free_object(x_33); +lean_free_object(x_23); +lean_dec(x_32); +lean_free_object(x_22); +lean_free_object(x_17); +lean_free_object(x_12); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_216 = lean_ctor_get(x_161, 0); +lean_inc(x_216); +x_217 = lean_ctor_get(x_161, 1); +lean_inc(x_217); +if (lean_is_exclusive(x_161)) { + lean_ctor_release(x_161, 0); + lean_ctor_release(x_161, 1); + x_218 = x_161; +} else { + lean_dec_ref(x_161); + x_218 = lean_box(0); +} +if (lean_is_scalar(x_218)) { + x_219 = lean_alloc_ctor(1, 2, 0); +} else { + x_219 = x_218; +} +lean_ctor_set(x_219, 0, x_216); +lean_ctor_set(x_219, 1, x_217); +return x_219; } } } else { -uint8_t x_105; +uint8_t x_220; lean_free_object(x_44); lean_dec(x_48); -lean_dec(x_47); lean_free_object(x_41); lean_dec(x_42); lean_free_object(x_33); lean_free_object(x_23); lean_dec(x_32); -lean_dec(x_31); lean_free_object(x_22); lean_free_object(x_17); lean_free_object(x_12); @@ -20215,29 +21111,29 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_105 = !lean_is_exclusive(x_60); -if (x_105 == 0) +x_220 = !lean_is_exclusive(x_60); +if (x_220 == 0) { return x_60; } else { -lean_object* x_106; lean_object* x_107; lean_object* x_108; -x_106 = lean_ctor_get(x_60, 0); -x_107 = lean_ctor_get(x_60, 1); -lean_inc(x_107); -lean_inc(x_106); +lean_object* x_221; lean_object* x_222; lean_object* x_223; +x_221 = lean_ctor_get(x_60, 0); +x_222 = lean_ctor_get(x_60, 1); +lean_inc(x_222); +lean_inc(x_221); lean_dec(x_60); -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_223 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_223, 0, x_221); +lean_ctor_set(x_223, 1, x_222); +return x_223; } } } else { -lean_object* x_109; uint8_t x_110; lean_object* x_111; lean_object* x_112; +lean_object* x_224; uint8_t x_225; lean_object* x_226; lean_object* x_227; lean_free_object(x_44); lean_dec(x_47); lean_free_object(x_41); @@ -20247,18 +21143,18 @@ lean_dec(x_31); lean_free_object(x_22); lean_free_object(x_17); lean_free_object(x_12); -x_109 = lean_ctor_get(x_56, 1); -lean_inc(x_109); +x_224 = lean_ctor_get(x_56, 1); +lean_inc(x_224); lean_dec(x_56); -x_110 = 0; -x_111 = lean_box(0); -x_112 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__3(x_32, x_48, x_42, x_1, x_110, x_111, x_4, x_5, x_6, x_7, x_109); -return x_112; +x_225 = 0; +x_226 = lean_box(0); +x_227 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__3(x_32, x_48, x_42, x_1, x_225, x_226, x_4, x_5, x_6, x_7, x_224); +return x_227; } } else { -uint8_t x_113; +uint8_t x_228; lean_free_object(x_44); lean_dec(x_48); lean_dec(x_47); @@ -20275,30 +21171,30 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_113 = !lean_is_exclusive(x_56); -if (x_113 == 0) +x_228 = !lean_is_exclusive(x_56); +if (x_228 == 0) { return x_56; } else { -lean_object* x_114; lean_object* x_115; lean_object* x_116; -x_114 = lean_ctor_get(x_56, 0); -x_115 = lean_ctor_get(x_56, 1); -lean_inc(x_115); -lean_inc(x_114); +lean_object* x_229; lean_object* x_230; lean_object* x_231; +x_229 = lean_ctor_get(x_56, 0); +x_230 = lean_ctor_get(x_56, 1); +lean_inc(x_230); +lean_inc(x_229); lean_dec(x_56); -x_116 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_116, 0, x_114); -lean_ctor_set(x_116, 1, x_115); -return x_116; +x_231 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_231, 0, x_229); +lean_ctor_set(x_231, 1, x_230); +return x_231; } } } } else { -uint8_t x_117; +uint8_t x_232; lean_free_object(x_44); lean_dec(x_48); lean_dec(x_47); @@ -20318,52 +21214,52 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_117 = !lean_is_exclusive(x_49); -if (x_117 == 0) +x_232 = !lean_is_exclusive(x_49); +if (x_232 == 0) { return x_49; } else { -lean_object* x_118; lean_object* x_119; lean_object* x_120; -x_118 = lean_ctor_get(x_49, 0); -x_119 = lean_ctor_get(x_49, 1); -lean_inc(x_119); -lean_inc(x_118); +lean_object* x_233; lean_object* x_234; lean_object* x_235; +x_233 = lean_ctor_get(x_49, 0); +x_234 = lean_ctor_get(x_49, 1); +lean_inc(x_234); +lean_inc(x_233); lean_dec(x_49); -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; +x_235 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_235, 0, x_233); +lean_ctor_set(x_235, 1, x_234); +return x_235; } } } else { -lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; -x_121 = lean_ctor_get(x_41, 0); -x_122 = lean_ctor_get(x_44, 0); -x_123 = lean_ctor_get(x_44, 1); -lean_inc(x_123); -lean_inc(x_122); +lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; +x_236 = lean_ctor_get(x_41, 0); +x_237 = lean_ctor_get(x_44, 0); +x_238 = lean_ctor_get(x_44, 1); +lean_inc(x_238); +lean_inc(x_237); lean_dec(x_44); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -x_124 = l_Lean_Meta_isExprDefEqGuarded(x_28, x_121, x_4, x_5, x_6, x_7, x_39); -if (lean_obj_tag(x_124) == 0) +x_239 = l_Lean_Meta_isExprDefEqGuarded(x_28, x_236, x_4, x_5, x_6, x_7, x_39); +if (lean_obj_tag(x_239) == 0) { -lean_object* x_125; uint8_t x_126; -x_125 = lean_ctor_get(x_124, 0); -lean_inc(x_125); -x_126 = lean_unbox(x_125); -lean_dec(x_125); -if (x_126 == 0) +lean_object* x_240; uint8_t x_241; +x_240 = lean_ctor_get(x_239, 0); +lean_inc(x_240); +x_241 = lean_unbox(x_240); +lean_dec(x_240); +if (x_241 == 0) { -lean_object* x_127; lean_object* x_128; lean_object* x_129; -lean_dec(x_123); -lean_dec(x_122); +lean_object* x_242; lean_object* x_243; lean_object* x_244; +lean_dec(x_238); +lean_dec(x_237); lean_free_object(x_41); lean_dec(x_42); lean_free_object(x_33); @@ -20373,172 +21269,315 @@ lean_dec(x_31); lean_free_object(x_22); lean_free_object(x_17); lean_free_object(x_12); -x_127 = lean_ctor_get(x_124, 1); -lean_inc(x_127); -lean_dec(x_124); -x_128 = lean_box(0); -x_129 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__1(x_2, x_16, x_3, x_128, x_4, x_5, x_6, x_7, x_127); -return x_129; +x_242 = lean_ctor_get(x_239, 1); +lean_inc(x_242); +lean_dec(x_239); +x_243 = lean_box(0); +x_244 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__1(x_2, x_16, x_3, x_243, x_4, x_5, x_6, x_7, x_242); +return x_244; } else { -lean_object* x_130; lean_object* x_131; +lean_object* x_245; lean_object* x_246; lean_dec(x_16); lean_dec(x_3); lean_dec(x_2); -x_130 = lean_ctor_get(x_124, 1); -lean_inc(x_130); -lean_dec(x_124); +x_245 = lean_ctor_get(x_239, 1); +lean_inc(x_245); +lean_dec(x_239); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -lean_inc(x_122); +lean_inc(x_237); lean_inc(x_31); -x_131 = l_Lean_Meta_isExprDefEqGuarded(x_31, x_122, x_4, x_5, x_6, x_7, x_130); -if (lean_obj_tag(x_131) == 0) +x_246 = l_Lean_Meta_isExprDefEqGuarded(x_31, x_237, x_4, x_5, x_6, x_7, x_245); +if (lean_obj_tag(x_246) == 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_247; uint8_t x_248; +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_134; lean_object* x_135; -x_134 = lean_ctor_get(x_131, 1); -lean_inc(x_134); -lean_dec(x_131); +lean_object* x_249; lean_object* x_250; +x_249 = lean_ctor_get(x_246, 1); +lean_inc(x_249); +lean_dec(x_246); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -lean_inc(x_31); -x_135 = lean_infer_type(x_31, x_4, x_5, x_6, x_7, x_134); -if (lean_obj_tag(x_135) == 0) +x_250 = l_Lean_Meta_addPPExplicitToExposeDiff(x_31, x_237, x_4, x_5, x_6, x_7, x_249); +if (lean_obj_tag(x_250) == 0) { -lean_object* x_136; lean_object* x_137; lean_object* x_138; -x_136 = lean_ctor_get(x_135, 0); -lean_inc(x_136); -x_137 = lean_ctor_get(x_135, 1); -lean_inc(x_137); -lean_dec(x_135); +lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; +x_251 = lean_ctor_get(x_250, 0); +lean_inc(x_251); +x_252 = lean_ctor_get(x_250, 1); +lean_inc(x_252); +lean_dec(x_250); +x_253 = lean_ctor_get(x_251, 0); +lean_inc(x_253); +x_254 = lean_ctor_get(x_251, 1); +lean_inc(x_254); +if (lean_is_exclusive(x_251)) { + lean_ctor_release(x_251, 0); + lean_ctor_release(x_251, 1); + x_255 = x_251; +} else { + lean_dec_ref(x_251); + x_255 = lean_box(0); +} lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -lean_inc(x_122); -x_138 = lean_infer_type(x_122, x_4, x_5, x_6, x_7, x_137); -if (lean_obj_tag(x_138) == 0) +lean_inc(x_253); +x_256 = lean_infer_type(x_253, x_4, x_5, x_6, x_7, x_252); +if (lean_obj_tag(x_256) == 0) { -lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; uint8_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; 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_139 = lean_ctor_get(x_138, 0); -lean_inc(x_139); -x_140 = lean_ctor_get(x_138, 1); -lean_inc(x_140); -lean_dec(x_138); -x_141 = lean_array_get_size(x_1); -x_142 = lean_unsigned_to_nat(0u); -x_143 = lean_nat_dec_lt(x_142, x_141); -lean_dec(x_141); -x_144 = l_Lean_MessageData_ofExpr(x_31); -x_145 = l___private_Lean_Elab_Calc_0__Lean_Elab_Term_getRelUniv___lambda__1___closed__4; -x_146 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_146, 0, x_145); -lean_ctor_set(x_146, 1, x_144); -x_147 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Term_elabCalcSteps___spec__1___closed__6; +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_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_254); +x_259 = lean_infer_type(x_254, x_4, x_5, x_6, x_7, x_258); +if (lean_obj_tag(x_259) == 0) +{ +lean_object* x_260; lean_object* x_261; lean_object* x_262; +x_260 = lean_ctor_get(x_259, 0); +lean_inc(x_260); +x_261 = lean_ctor_get(x_259, 1); +lean_inc(x_261); +lean_dec(x_259); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +x_262 = l_Lean_Meta_addPPExplicitToExposeDiff(x_257, x_260, x_4, x_5, x_6, x_7, x_261); +if (lean_obj_tag(x_262) == 0) +{ +lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; uint8_t x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; +x_263 = lean_ctor_get(x_262, 0); +lean_inc(x_263); +x_264 = lean_ctor_get(x_262, 1); +lean_inc(x_264); +lean_dec(x_262); +x_265 = lean_ctor_get(x_263, 0); +lean_inc(x_265); +x_266 = lean_ctor_get(x_263, 1); +lean_inc(x_266); +if (lean_is_exclusive(x_263)) { + lean_ctor_release(x_263, 0); + lean_ctor_release(x_263, 1); + x_267 = x_263; +} else { + lean_dec_ref(x_263); + x_267 = lean_box(0); +} +x_268 = lean_array_get_size(x_1); +x_269 = lean_unsigned_to_nat(0u); +x_270 = lean_nat_dec_lt(x_269, x_268); +lean_dec(x_268); +x_271 = l_Lean_MessageData_ofExpr(x_253); +x_272 = l___private_Lean_Elab_Calc_0__Lean_Elab_Term_getRelUniv___lambda__1___closed__4; +if (lean_is_scalar(x_267)) { + x_273 = lean_alloc_ctor(7, 2, 0); +} else { + x_273 = x_267; + lean_ctor_set_tag(x_273, 7); +} +lean_ctor_set(x_273, 0, x_272); +lean_ctor_set(x_273, 1, x_271); +x_274 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Term_elabCalcSteps___spec__1___closed__6; +if (lean_is_scalar(x_255)) { + x_275 = lean_alloc_ctor(7, 2, 0); +} else { + x_275 = x_255; + lean_ctor_set_tag(x_275, 7); +} +lean_ctor_set(x_275, 0, x_273); +lean_ctor_set(x_275, 1, x_274); +x_276 = l_Lean_MessageData_ofExpr(x_265); +x_277 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_277, 0, x_275); +lean_ctor_set(x_277, 1, x_276); lean_ctor_set_tag(x_41, 7); -lean_ctor_set(x_41, 1, x_147); -lean_ctor_set(x_41, 0, x_146); -x_148 = l_Lean_MessageData_ofExpr(x_136); +lean_ctor_set(x_41, 1, x_272); +lean_ctor_set(x_41, 0, x_277); +x_278 = l_Lean_indentD(x_41); +x_279 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Term_elabCalcSteps___spec__1___closed__4; lean_ctor_set_tag(x_33, 7); -lean_ctor_set(x_33, 1, x_148); -lean_ctor_set(x_33, 0, x_41); +lean_ctor_set(x_33, 1, x_278); +lean_ctor_set(x_33, 0, x_279); +x_280 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__3___closed__4; lean_ctor_set_tag(x_23, 7); -lean_ctor_set(x_23, 1, x_145); +lean_ctor_set(x_23, 1, x_280); lean_ctor_set(x_23, 0, x_33); -x_149 = l_Lean_indentD(x_23); -x_150 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Term_elabCalcSteps___spec__1___closed__4; +x_281 = l_Lean_MessageData_ofExpr(x_254); lean_ctor_set_tag(x_22, 7); -lean_ctor_set(x_22, 1, x_149); -lean_ctor_set(x_22, 0, x_150); -x_151 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__3___closed__4; +lean_ctor_set(x_22, 1, x_281); +lean_ctor_set(x_22, 0, x_272); lean_ctor_set_tag(x_17, 7); -lean_ctor_set(x_17, 1, x_151); +lean_ctor_set(x_17, 1, x_274); lean_ctor_set(x_17, 0, x_22); -x_152 = l_Lean_MessageData_ofExpr(x_122); +x_282 = l_Lean_MessageData_ofExpr(x_266); lean_ctor_set_tag(x_12, 7); -lean_ctor_set(x_12, 1, x_152); -lean_ctor_set(x_12, 0, x_145); -x_153 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_153, 0, x_12); -lean_ctor_set(x_153, 1, x_147); -x_154 = l_Lean_MessageData_ofExpr(x_139); -x_155 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_155, 0, x_153); -lean_ctor_set(x_155, 1, x_154); -x_156 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_156, 0, x_155); -lean_ctor_set(x_156, 1, x_145); -x_157 = l_Lean_indentD(x_156); -x_158 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_158, 0, x_17); -lean_ctor_set(x_158, 1, x_157); -x_159 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_159, 0, x_158); -lean_ctor_set(x_159, 1, x_145); -if (x_143 == 0) +lean_ctor_set(x_12, 1, x_282); +lean_ctor_set(x_12, 0, x_17); +x_283 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_283, 0, x_12); +lean_ctor_set(x_283, 1, x_272); +x_284 = l_Lean_indentD(x_283); +x_285 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_285, 0, x_23); +lean_ctor_set(x_285, 1, x_284); +x_286 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_286, 0, x_285); +lean_ctor_set(x_286, 1, x_272); +if (x_270 == 0) +{ +lean_object* x_287; lean_object* x_288; lean_object* x_289; uint8_t x_290; lean_object* x_291; lean_object* x_292; uint8_t x_293; lean_object* x_294; lean_object* x_295; +x_287 = l_Lean_Elab_Term_instInhabitedCalcStepView; +x_288 = l_outOfBounds___rarg(x_287); +x_289 = lean_ctor_get(x_288, 1); +lean_inc(x_289); +lean_dec(x_288); +x_290 = 2; +lean_inc(x_6); +x_291 = l_Lean_logAt___at_Lean_Elab_Term_reportUnsolvedGoals___spec__2(x_289, x_286, x_290, x_4, x_5, x_6, x_7, x_264); +lean_dec(x_289); +x_292 = lean_ctor_get(x_291, 1); +lean_inc(x_292); +lean_dec(x_291); +x_293 = 1; +x_294 = lean_box(0); +x_295 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__3(x_32, x_238, x_42, x_1, x_293, x_294, x_4, x_5, x_6, x_7, x_292); +return x_295; +} +else { -lean_object* x_160; lean_object* x_161; lean_object* x_162; uint8_t x_163; lean_object* x_164; lean_object* x_165; uint8_t x_166; lean_object* x_167; lean_object* x_168; -x_160 = l_Lean_Elab_Term_instInhabitedCalcStepView; -x_161 = l_outOfBounds___rarg(x_160); -x_162 = lean_ctor_get(x_161, 1); -lean_inc(x_162); -lean_dec(x_161); -x_163 = 2; +lean_object* x_296; lean_object* x_297; uint8_t x_298; lean_object* x_299; lean_object* x_300; uint8_t x_301; lean_object* x_302; lean_object* x_303; +x_296 = lean_array_fget(x_1, x_269); +x_297 = lean_ctor_get(x_296, 1); +lean_inc(x_297); +lean_dec(x_296); +x_298 = 2; lean_inc(x_6); -x_164 = l_Lean_logAt___at_Lean_Elab_Term_reportUnsolvedGoals___spec__2(x_162, x_159, x_163, x_4, x_5, x_6, x_7, x_140); -lean_dec(x_162); -x_165 = lean_ctor_get(x_164, 1); -lean_inc(x_165); -lean_dec(x_164); -x_166 = 1; -x_167 = lean_box(0); -x_168 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__3(x_32, x_123, x_42, x_1, x_166, x_167, x_4, x_5, x_6, x_7, x_165); -return x_168; +x_299 = l_Lean_logAt___at_Lean_Elab_Term_reportUnsolvedGoals___spec__2(x_297, x_286, x_298, x_4, x_5, x_6, x_7, x_264); +lean_dec(x_297); +x_300 = lean_ctor_get(x_299, 1); +lean_inc(x_300); +lean_dec(x_299); +x_301 = 1; +x_302 = lean_box(0); +x_303 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__3(x_32, x_238, x_42, x_1, x_301, x_302, x_4, x_5, x_6, x_7, x_300); +return x_303; +} +} +else +{ +lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; +lean_dec(x_255); +lean_dec(x_254); +lean_dec(x_253); +lean_dec(x_238); +lean_free_object(x_41); +lean_dec(x_42); +lean_free_object(x_33); +lean_free_object(x_23); +lean_dec(x_32); +lean_free_object(x_22); +lean_free_object(x_17); +lean_free_object(x_12); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_304 = lean_ctor_get(x_262, 0); +lean_inc(x_304); +x_305 = lean_ctor_get(x_262, 1); +lean_inc(x_305); +if (lean_is_exclusive(x_262)) { + lean_ctor_release(x_262, 0); + lean_ctor_release(x_262, 1); + x_306 = x_262; +} else { + lean_dec_ref(x_262); + 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_dec(x_257); +lean_dec(x_255); +lean_dec(x_254); +lean_dec(x_253); +lean_dec(x_238); +lean_free_object(x_41); +lean_dec(x_42); +lean_free_object(x_33); +lean_free_object(x_23); +lean_dec(x_32); +lean_free_object(x_22); +lean_free_object(x_17); +lean_free_object(x_12); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_308 = lean_ctor_get(x_259, 0); +lean_inc(x_308); +x_309 = lean_ctor_get(x_259, 1); +lean_inc(x_309); +if (lean_is_exclusive(x_259)) { + lean_ctor_release(x_259, 0); + lean_ctor_release(x_259, 1); + x_310 = x_259; +} else { + lean_dec_ref(x_259); + x_310 = lean_box(0); +} +if (lean_is_scalar(x_310)) { + x_311 = lean_alloc_ctor(1, 2, 0); +} else { + x_311 = x_310; } -else -{ -lean_object* x_169; lean_object* x_170; uint8_t x_171; lean_object* x_172; lean_object* x_173; uint8_t x_174; lean_object* x_175; lean_object* x_176; -x_169 = lean_array_fget(x_1, x_142); -x_170 = lean_ctor_get(x_169, 1); -lean_inc(x_170); -lean_dec(x_169); -x_171 = 2; -lean_inc(x_6); -x_172 = l_Lean_logAt___at_Lean_Elab_Term_reportUnsolvedGoals___spec__2(x_170, x_159, x_171, x_4, x_5, x_6, x_7, x_140); -lean_dec(x_170); -x_173 = lean_ctor_get(x_172, 1); -lean_inc(x_173); -lean_dec(x_172); -x_174 = 1; -x_175 = lean_box(0); -x_176 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__3(x_32, x_123, x_42, x_1, x_174, x_175, x_4, x_5, x_6, x_7, x_173); -return x_176; +lean_ctor_set(x_311, 0, x_308); +lean_ctor_set(x_311, 1, x_309); +return x_311; } } else { -lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; -lean_dec(x_136); -lean_dec(x_123); -lean_dec(x_122); +lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; +lean_dec(x_255); +lean_dec(x_254); +lean_dec(x_253); +lean_dec(x_238); lean_free_object(x_41); lean_dec(x_42); lean_free_object(x_33); lean_free_object(x_23); lean_dec(x_32); -lean_dec(x_31); lean_free_object(x_22); lean_free_object(x_17); lean_free_object(x_12); @@ -20546,39 +21585,37 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_177 = lean_ctor_get(x_138, 0); -lean_inc(x_177); -x_178 = lean_ctor_get(x_138, 1); -lean_inc(x_178); -if (lean_is_exclusive(x_138)) { - lean_ctor_release(x_138, 0); - lean_ctor_release(x_138, 1); - x_179 = x_138; +x_312 = lean_ctor_get(x_256, 0); +lean_inc(x_312); +x_313 = lean_ctor_get(x_256, 1); +lean_inc(x_313); +if (lean_is_exclusive(x_256)) { + lean_ctor_release(x_256, 0); + lean_ctor_release(x_256, 1); + x_314 = x_256; } else { - lean_dec_ref(x_138); - x_179 = lean_box(0); + lean_dec_ref(x_256); + x_314 = lean_box(0); } -if (lean_is_scalar(x_179)) { - x_180 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_314)) { + x_315 = lean_alloc_ctor(1, 2, 0); } else { - x_180 = x_179; + x_315 = x_314; } -lean_ctor_set(x_180, 0, x_177); -lean_ctor_set(x_180, 1, x_178); -return x_180; +lean_ctor_set(x_315, 0, x_312); +lean_ctor_set(x_315, 1, x_313); +return x_315; } } else { -lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; -lean_dec(x_123); -lean_dec(x_122); +lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; +lean_dec(x_238); lean_free_object(x_41); lean_dec(x_42); lean_free_object(x_33); lean_free_object(x_23); lean_dec(x_32); -lean_dec(x_31); lean_free_object(x_22); lean_free_object(x_17); lean_free_object(x_12); @@ -20586,32 +21623,32 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_181 = lean_ctor_get(x_135, 0); -lean_inc(x_181); -x_182 = lean_ctor_get(x_135, 1); -lean_inc(x_182); -if (lean_is_exclusive(x_135)) { - lean_ctor_release(x_135, 0); - lean_ctor_release(x_135, 1); - x_183 = x_135; +x_316 = lean_ctor_get(x_250, 0); +lean_inc(x_316); +x_317 = lean_ctor_get(x_250, 1); +lean_inc(x_317); +if (lean_is_exclusive(x_250)) { + lean_ctor_release(x_250, 0); + lean_ctor_release(x_250, 1); + x_318 = x_250; } else { - lean_dec_ref(x_135); - x_183 = lean_box(0); + lean_dec_ref(x_250); + x_318 = lean_box(0); } -if (lean_is_scalar(x_183)) { - x_184 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_318)) { + x_319 = lean_alloc_ctor(1, 2, 0); } else { - x_184 = x_183; + x_319 = x_318; } -lean_ctor_set(x_184, 0, x_181); -lean_ctor_set(x_184, 1, x_182); -return x_184; +lean_ctor_set(x_319, 0, x_316); +lean_ctor_set(x_319, 1, x_317); +return x_319; } } else { -lean_object* x_185; uint8_t x_186; lean_object* x_187; lean_object* x_188; -lean_dec(x_122); +lean_object* x_320; uint8_t x_321; lean_object* x_322; lean_object* x_323; +lean_dec(x_237); lean_free_object(x_41); lean_free_object(x_33); lean_free_object(x_23); @@ -20619,20 +21656,20 @@ lean_dec(x_31); lean_free_object(x_22); lean_free_object(x_17); lean_free_object(x_12); -x_185 = lean_ctor_get(x_131, 1); -lean_inc(x_185); -lean_dec(x_131); -x_186 = 0; -x_187 = lean_box(0); -x_188 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__3(x_32, x_123, x_42, x_1, x_186, x_187, x_4, x_5, x_6, x_7, x_185); -return x_188; +x_320 = lean_ctor_get(x_246, 1); +lean_inc(x_320); +lean_dec(x_246); +x_321 = 0; +x_322 = lean_box(0); +x_323 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__3(x_32, x_238, x_42, x_1, x_321, x_322, x_4, x_5, x_6, x_7, x_320); +return x_323; } } else { -lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; -lean_dec(x_123); -lean_dec(x_122); +lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; +lean_dec(x_238); +lean_dec(x_237); lean_free_object(x_41); lean_dec(x_42); lean_free_object(x_33); @@ -20646,34 +21683,34 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_189 = lean_ctor_get(x_131, 0); -lean_inc(x_189); -x_190 = lean_ctor_get(x_131, 1); -lean_inc(x_190); -if (lean_is_exclusive(x_131)) { - lean_ctor_release(x_131, 0); - lean_ctor_release(x_131, 1); - x_191 = x_131; +x_324 = lean_ctor_get(x_246, 0); +lean_inc(x_324); +x_325 = lean_ctor_get(x_246, 1); +lean_inc(x_325); +if (lean_is_exclusive(x_246)) { + lean_ctor_release(x_246, 0); + lean_ctor_release(x_246, 1); + x_326 = x_246; } else { - lean_dec_ref(x_131); - x_191 = lean_box(0); + lean_dec_ref(x_246); + x_326 = lean_box(0); } -if (lean_is_scalar(x_191)) { - x_192 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_326)) { + x_327 = lean_alloc_ctor(1, 2, 0); } else { - x_192 = x_191; + x_327 = x_326; } -lean_ctor_set(x_192, 0, x_189); -lean_ctor_set(x_192, 1, x_190); -return x_192; +lean_ctor_set(x_327, 0, x_324); +lean_ctor_set(x_327, 1, x_325); +return x_327; } } } else { -lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; -lean_dec(x_123); -lean_dec(x_122); +lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; +lean_dec(x_238); +lean_dec(x_237); lean_free_object(x_41); lean_dec(x_42); lean_free_object(x_33); @@ -20690,67 +21727,67 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_193 = lean_ctor_get(x_124, 0); -lean_inc(x_193); -x_194 = lean_ctor_get(x_124, 1); -lean_inc(x_194); -if (lean_is_exclusive(x_124)) { - lean_ctor_release(x_124, 0); - lean_ctor_release(x_124, 1); - x_195 = x_124; +x_328 = lean_ctor_get(x_239, 0); +lean_inc(x_328); +x_329 = lean_ctor_get(x_239, 1); +lean_inc(x_329); +if (lean_is_exclusive(x_239)) { + lean_ctor_release(x_239, 0); + lean_ctor_release(x_239, 1); + x_330 = x_239; } else { - lean_dec_ref(x_124); - x_195 = lean_box(0); + lean_dec_ref(x_239); + x_330 = lean_box(0); } -if (lean_is_scalar(x_195)) { - x_196 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_330)) { + x_331 = lean_alloc_ctor(1, 2, 0); } else { - x_196 = x_195; + x_331 = x_330; } -lean_ctor_set(x_196, 0, x_193); -lean_ctor_set(x_196, 1, x_194); -return x_196; +lean_ctor_set(x_331, 0, x_328); +lean_ctor_set(x_331, 1, x_329); +return x_331; } } } else { -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_197 = lean_ctor_get(x_41, 1); -x_198 = lean_ctor_get(x_41, 0); -lean_inc(x_197); -lean_inc(x_198); +lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; +x_332 = lean_ctor_get(x_41, 1); +x_333 = lean_ctor_get(x_41, 0); +lean_inc(x_332); +lean_inc(x_333); lean_dec(x_41); -x_199 = lean_ctor_get(x_197, 0); -lean_inc(x_199); -x_200 = lean_ctor_get(x_197, 1); -lean_inc(x_200); -if (lean_is_exclusive(x_197)) { - lean_ctor_release(x_197, 0); - lean_ctor_release(x_197, 1); - x_201 = x_197; +x_334 = lean_ctor_get(x_332, 0); +lean_inc(x_334); +x_335 = lean_ctor_get(x_332, 1); +lean_inc(x_335); +if (lean_is_exclusive(x_332)) { + lean_ctor_release(x_332, 0); + lean_ctor_release(x_332, 1); + x_336 = x_332; } else { - lean_dec_ref(x_197); - x_201 = lean_box(0); + lean_dec_ref(x_332); + x_336 = lean_box(0); } lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -x_202 = l_Lean_Meta_isExprDefEqGuarded(x_28, x_198, x_4, x_5, x_6, x_7, x_39); -if (lean_obj_tag(x_202) == 0) +x_337 = l_Lean_Meta_isExprDefEqGuarded(x_28, x_333, x_4, x_5, x_6, x_7, x_39); +if (lean_obj_tag(x_337) == 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_338; uint8_t x_339; +x_338 = lean_ctor_get(x_337, 0); +lean_inc(x_338); +x_339 = lean_unbox(x_338); +lean_dec(x_338); +if (x_339 == 0) { -lean_object* x_205; lean_object* x_206; lean_object* x_207; -lean_dec(x_201); -lean_dec(x_200); -lean_dec(x_199); +lean_object* x_340; lean_object* x_341; lean_object* x_342; +lean_dec(x_336); +lean_dec(x_335); +lean_dec(x_334); lean_dec(x_42); lean_free_object(x_33); lean_free_object(x_23); @@ -20759,177 +21796,237 @@ lean_dec(x_31); lean_free_object(x_22); lean_free_object(x_17); lean_free_object(x_12); -x_205 = lean_ctor_get(x_202, 1); -lean_inc(x_205); -lean_dec(x_202); -x_206 = lean_box(0); -x_207 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__1(x_2, x_16, x_3, x_206, x_4, x_5, x_6, x_7, x_205); -return x_207; +x_340 = lean_ctor_get(x_337, 1); +lean_inc(x_340); +lean_dec(x_337); +x_341 = lean_box(0); +x_342 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__1(x_2, x_16, x_3, x_341, x_4, x_5, x_6, x_7, x_340); +return x_342; } else { -lean_object* x_208; lean_object* x_209; +lean_object* x_343; lean_object* x_344; lean_dec(x_16); lean_dec(x_3); lean_dec(x_2); -x_208 = lean_ctor_get(x_202, 1); -lean_inc(x_208); -lean_dec(x_202); +x_343 = lean_ctor_get(x_337, 1); +lean_inc(x_343); +lean_dec(x_337); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -lean_inc(x_199); +lean_inc(x_334); lean_inc(x_31); -x_209 = l_Lean_Meta_isExprDefEqGuarded(x_31, x_199, x_4, x_5, x_6, x_7, x_208); -if (lean_obj_tag(x_209) == 0) +x_344 = l_Lean_Meta_isExprDefEqGuarded(x_31, x_334, x_4, x_5, x_6, x_7, x_343); +if (lean_obj_tag(x_344) == 0) { -lean_object* x_210; uint8_t x_211; -x_210 = lean_ctor_get(x_209, 0); -lean_inc(x_210); -x_211 = lean_unbox(x_210); -lean_dec(x_210); -if (x_211 == 0) +lean_object* x_345; uint8_t x_346; +x_345 = lean_ctor_get(x_344, 0); +lean_inc(x_345); +x_346 = lean_unbox(x_345); +lean_dec(x_345); +if (x_346 == 0) { -lean_object* x_212; lean_object* x_213; -x_212 = lean_ctor_get(x_209, 1); -lean_inc(x_212); -lean_dec(x_209); +lean_object* x_347; lean_object* x_348; +x_347 = lean_ctor_get(x_344, 1); +lean_inc(x_347); +lean_dec(x_344); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -lean_inc(x_31); -x_213 = lean_infer_type(x_31, x_4, x_5, x_6, x_7, x_212); -if (lean_obj_tag(x_213) == 0) +x_348 = l_Lean_Meta_addPPExplicitToExposeDiff(x_31, x_334, x_4, x_5, x_6, x_7, x_347); +if (lean_obj_tag(x_348) == 0) { -lean_object* x_214; lean_object* x_215; lean_object* x_216; -x_214 = lean_ctor_get(x_213, 0); -lean_inc(x_214); -x_215 = lean_ctor_get(x_213, 1); -lean_inc(x_215); -lean_dec(x_213); +lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; lean_object* x_353; lean_object* x_354; +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 = lean_ctor_get(x_349, 0); +lean_inc(x_351); +x_352 = lean_ctor_get(x_349, 1); +lean_inc(x_352); +if (lean_is_exclusive(x_349)) { + lean_ctor_release(x_349, 0); + lean_ctor_release(x_349, 1); + x_353 = x_349; +} else { + lean_dec_ref(x_349); + x_353 = lean_box(0); +} lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -lean_inc(x_199); -x_216 = lean_infer_type(x_199, x_4, x_5, x_6, x_7, x_215); -if (lean_obj_tag(x_216) == 0) +lean_inc(x_351); +x_354 = lean_infer_type(x_351, x_4, x_5, x_6, x_7, x_350); +if (lean_obj_tag(x_354) == 0) { -lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; uint8_t x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; -x_217 = lean_ctor_get(x_216, 0); -lean_inc(x_217); -x_218 = lean_ctor_get(x_216, 1); -lean_inc(x_218); -lean_dec(x_216); -x_219 = lean_array_get_size(x_1); -x_220 = lean_unsigned_to_nat(0u); -x_221 = lean_nat_dec_lt(x_220, x_219); -lean_dec(x_219); -x_222 = l_Lean_MessageData_ofExpr(x_31); -x_223 = l___private_Lean_Elab_Calc_0__Lean_Elab_Term_getRelUniv___lambda__1___closed__4; -if (lean_is_scalar(x_201)) { - x_224 = lean_alloc_ctor(7, 2, 0); -} else { - x_224 = x_201; - lean_ctor_set_tag(x_224, 7); -} -lean_ctor_set(x_224, 0, x_223); -lean_ctor_set(x_224, 1, x_222); -x_225 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Term_elabCalcSteps___spec__1___closed__6; -x_226 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_226, 0, x_224); -lean_ctor_set(x_226, 1, x_225); -x_227 = l_Lean_MessageData_ofExpr(x_214); +lean_object* x_355; lean_object* x_356; lean_object* x_357; +x_355 = lean_ctor_get(x_354, 0); +lean_inc(x_355); +x_356 = lean_ctor_get(x_354, 1); +lean_inc(x_356); +lean_dec(x_354); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_352); +x_357 = lean_infer_type(x_352, x_4, x_5, x_6, x_7, x_356); +if (lean_obj_tag(x_357) == 0) +{ +lean_object* x_358; lean_object* x_359; lean_object* x_360; +x_358 = lean_ctor_get(x_357, 0); +lean_inc(x_358); +x_359 = lean_ctor_get(x_357, 1); +lean_inc(x_359); +lean_dec(x_357); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +x_360 = l_Lean_Meta_addPPExplicitToExposeDiff(x_355, x_358, x_4, x_5, x_6, x_7, x_359); +if (lean_obj_tag(x_360) == 0) +{ +lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; uint8_t x_368; lean_object* x_369; lean_object* x_370; lean_object* x_371; lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; lean_object* x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; lean_object* x_382; lean_object* x_383; lean_object* x_384; lean_object* x_385; +x_361 = lean_ctor_get(x_360, 0); +lean_inc(x_361); +x_362 = lean_ctor_get(x_360, 1); +lean_inc(x_362); +lean_dec(x_360); +x_363 = lean_ctor_get(x_361, 0); +lean_inc(x_363); +x_364 = lean_ctor_get(x_361, 1); +lean_inc(x_364); +if (lean_is_exclusive(x_361)) { + lean_ctor_release(x_361, 0); + lean_ctor_release(x_361, 1); + x_365 = x_361; +} else { + lean_dec_ref(x_361); + x_365 = lean_box(0); +} +x_366 = lean_array_get_size(x_1); +x_367 = lean_unsigned_to_nat(0u); +x_368 = lean_nat_dec_lt(x_367, x_366); +lean_dec(x_366); +x_369 = l_Lean_MessageData_ofExpr(x_351); +x_370 = l___private_Lean_Elab_Calc_0__Lean_Elab_Term_getRelUniv___lambda__1___closed__4; +if (lean_is_scalar(x_365)) { + x_371 = lean_alloc_ctor(7, 2, 0); +} else { + x_371 = x_365; + lean_ctor_set_tag(x_371, 7); +} +lean_ctor_set(x_371, 0, x_370); +lean_ctor_set(x_371, 1, x_369); +x_372 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Term_elabCalcSteps___spec__1___closed__6; +if (lean_is_scalar(x_353)) { + x_373 = lean_alloc_ctor(7, 2, 0); +} else { + x_373 = x_353; + lean_ctor_set_tag(x_373, 7); +} +lean_ctor_set(x_373, 0, x_371); +lean_ctor_set(x_373, 1, x_372); +x_374 = l_Lean_MessageData_ofExpr(x_363); +if (lean_is_scalar(x_336)) { + x_375 = lean_alloc_ctor(7, 2, 0); +} else { + x_375 = x_336; + lean_ctor_set_tag(x_375, 7); +} +lean_ctor_set(x_375, 0, x_373); +lean_ctor_set(x_375, 1, x_374); +x_376 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_376, 0, x_375); +lean_ctor_set(x_376, 1, x_370); +x_377 = l_Lean_indentD(x_376); +x_378 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Term_elabCalcSteps___spec__1___closed__4; lean_ctor_set_tag(x_33, 7); -lean_ctor_set(x_33, 1, x_227); -lean_ctor_set(x_33, 0, x_226); +lean_ctor_set(x_33, 1, x_377); +lean_ctor_set(x_33, 0, x_378); +x_379 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__3___closed__4; lean_ctor_set_tag(x_23, 7); -lean_ctor_set(x_23, 1, x_223); +lean_ctor_set(x_23, 1, x_379); lean_ctor_set(x_23, 0, x_33); -x_228 = l_Lean_indentD(x_23); -x_229 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Term_elabCalcSteps___spec__1___closed__4; +x_380 = l_Lean_MessageData_ofExpr(x_352); lean_ctor_set_tag(x_22, 7); -lean_ctor_set(x_22, 1, x_228); -lean_ctor_set(x_22, 0, x_229); -x_230 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__3___closed__4; +lean_ctor_set(x_22, 1, x_380); +lean_ctor_set(x_22, 0, x_370); lean_ctor_set_tag(x_17, 7); -lean_ctor_set(x_17, 1, x_230); +lean_ctor_set(x_17, 1, x_372); lean_ctor_set(x_17, 0, x_22); -x_231 = l_Lean_MessageData_ofExpr(x_199); +x_381 = l_Lean_MessageData_ofExpr(x_364); lean_ctor_set_tag(x_12, 7); -lean_ctor_set(x_12, 1, x_231); -lean_ctor_set(x_12, 0, x_223); -x_232 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_232, 0, x_12); -lean_ctor_set(x_232, 1, x_225); -x_233 = l_Lean_MessageData_ofExpr(x_217); -x_234 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_234, 0, x_232); -lean_ctor_set(x_234, 1, x_233); -x_235 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_235, 0, x_234); -lean_ctor_set(x_235, 1, x_223); -x_236 = l_Lean_indentD(x_235); -x_237 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_237, 0, x_17); -lean_ctor_set(x_237, 1, x_236); -x_238 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_238, 0, x_237); -lean_ctor_set(x_238, 1, x_223); -if (x_221 == 0) -{ -lean_object* x_239; lean_object* x_240; lean_object* x_241; uint8_t x_242; lean_object* x_243; lean_object* x_244; uint8_t x_245; lean_object* x_246; lean_object* x_247; -x_239 = l_Lean_Elab_Term_instInhabitedCalcStepView; -x_240 = l_outOfBounds___rarg(x_239); -x_241 = lean_ctor_get(x_240, 1); -lean_inc(x_241); -lean_dec(x_240); -x_242 = 2; +lean_ctor_set(x_12, 1, x_381); +lean_ctor_set(x_12, 0, x_17); +x_382 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_382, 0, x_12); +lean_ctor_set(x_382, 1, x_370); +x_383 = l_Lean_indentD(x_382); +x_384 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_384, 0, x_23); +lean_ctor_set(x_384, 1, x_383); +x_385 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_385, 0, x_384); +lean_ctor_set(x_385, 1, x_370); +if (x_368 == 0) +{ +lean_object* x_386; lean_object* x_387; lean_object* x_388; uint8_t x_389; lean_object* x_390; lean_object* x_391; uint8_t x_392; lean_object* x_393; lean_object* x_394; +x_386 = l_Lean_Elab_Term_instInhabitedCalcStepView; +x_387 = l_outOfBounds___rarg(x_386); +x_388 = lean_ctor_get(x_387, 1); +lean_inc(x_388); +lean_dec(x_387); +x_389 = 2; lean_inc(x_6); -x_243 = l_Lean_logAt___at_Lean_Elab_Term_reportUnsolvedGoals___spec__2(x_241, x_238, x_242, x_4, x_5, x_6, x_7, x_218); -lean_dec(x_241); -x_244 = lean_ctor_get(x_243, 1); -lean_inc(x_244); -lean_dec(x_243); -x_245 = 1; -x_246 = lean_box(0); -x_247 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__3(x_32, x_200, x_42, x_1, x_245, x_246, x_4, x_5, x_6, x_7, x_244); -return x_247; +x_390 = l_Lean_logAt___at_Lean_Elab_Term_reportUnsolvedGoals___spec__2(x_388, x_385, x_389, x_4, x_5, x_6, x_7, x_362); +lean_dec(x_388); +x_391 = lean_ctor_get(x_390, 1); +lean_inc(x_391); +lean_dec(x_390); +x_392 = 1; +x_393 = lean_box(0); +x_394 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__3(x_32, x_335, x_42, x_1, x_392, x_393, x_4, x_5, x_6, x_7, x_391); +return x_394; } else { -lean_object* x_248; lean_object* x_249; uint8_t x_250; lean_object* x_251; lean_object* x_252; uint8_t x_253; lean_object* x_254; lean_object* x_255; -x_248 = lean_array_fget(x_1, x_220); -x_249 = lean_ctor_get(x_248, 1); -lean_inc(x_249); -lean_dec(x_248); -x_250 = 2; +lean_object* x_395; lean_object* x_396; uint8_t x_397; lean_object* x_398; lean_object* x_399; uint8_t x_400; lean_object* x_401; lean_object* x_402; +x_395 = lean_array_fget(x_1, x_367); +x_396 = lean_ctor_get(x_395, 1); +lean_inc(x_396); +lean_dec(x_395); +x_397 = 2; lean_inc(x_6); -x_251 = l_Lean_logAt___at_Lean_Elab_Term_reportUnsolvedGoals___spec__2(x_249, x_238, x_250, x_4, x_5, x_6, x_7, x_218); -lean_dec(x_249); -x_252 = lean_ctor_get(x_251, 1); -lean_inc(x_252); -lean_dec(x_251); -x_253 = 1; -x_254 = lean_box(0); -x_255 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__3(x_32, x_200, x_42, x_1, x_253, x_254, x_4, x_5, x_6, x_7, x_252); -return x_255; +x_398 = l_Lean_logAt___at_Lean_Elab_Term_reportUnsolvedGoals___spec__2(x_396, x_385, x_397, x_4, x_5, x_6, x_7, x_362); +lean_dec(x_396); +x_399 = lean_ctor_get(x_398, 1); +lean_inc(x_399); +lean_dec(x_398); +x_400 = 1; +x_401 = lean_box(0); +x_402 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__3(x_32, x_335, x_42, x_1, x_400, x_401, x_4, x_5, x_6, x_7, x_399); +return x_402; } } else { -lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; -lean_dec(x_214); -lean_dec(x_201); -lean_dec(x_200); -lean_dec(x_199); +lean_object* x_403; lean_object* x_404; lean_object* x_405; lean_object* x_406; +lean_dec(x_353); +lean_dec(x_352); +lean_dec(x_351); +lean_dec(x_336); +lean_dec(x_335); lean_dec(x_42); lean_free_object(x_33); lean_free_object(x_23); lean_dec(x_32); -lean_dec(x_31); lean_free_object(x_22); lean_free_object(x_17); lean_free_object(x_12); @@ -20937,39 +22034,41 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_256 = lean_ctor_get(x_216, 0); -lean_inc(x_256); -x_257 = lean_ctor_get(x_216, 1); -lean_inc(x_257); -if (lean_is_exclusive(x_216)) { - lean_ctor_release(x_216, 0); - lean_ctor_release(x_216, 1); - x_258 = x_216; +x_403 = lean_ctor_get(x_360, 0); +lean_inc(x_403); +x_404 = lean_ctor_get(x_360, 1); +lean_inc(x_404); +if (lean_is_exclusive(x_360)) { + lean_ctor_release(x_360, 0); + lean_ctor_release(x_360, 1); + x_405 = x_360; } else { - lean_dec_ref(x_216); - x_258 = lean_box(0); + lean_dec_ref(x_360); + x_405 = lean_box(0); } -if (lean_is_scalar(x_258)) { - x_259 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_405)) { + x_406 = lean_alloc_ctor(1, 2, 0); } else { - x_259 = x_258; + x_406 = x_405; } -lean_ctor_set(x_259, 0, x_256); -lean_ctor_set(x_259, 1, x_257); -return x_259; +lean_ctor_set(x_406, 0, x_403); +lean_ctor_set(x_406, 1, x_404); +return x_406; } } else { -lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; -lean_dec(x_201); -lean_dec(x_200); -lean_dec(x_199); +lean_object* x_407; lean_object* x_408; lean_object* x_409; lean_object* x_410; +lean_dec(x_355); +lean_dec(x_353); +lean_dec(x_352); +lean_dec(x_351); +lean_dec(x_336); +lean_dec(x_335); lean_dec(x_42); lean_free_object(x_33); lean_free_object(x_23); lean_dec(x_32); -lean_dec(x_31); lean_free_object(x_22); lean_free_object(x_17); lean_free_object(x_12); @@ -20977,54 +22076,133 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_260 = lean_ctor_get(x_213, 0); -lean_inc(x_260); -x_261 = lean_ctor_get(x_213, 1); -lean_inc(x_261); -if (lean_is_exclusive(x_213)) { - lean_ctor_release(x_213, 0); - lean_ctor_release(x_213, 1); - x_262 = x_213; +x_407 = lean_ctor_get(x_357, 0); +lean_inc(x_407); +x_408 = lean_ctor_get(x_357, 1); +lean_inc(x_408); +if (lean_is_exclusive(x_357)) { + lean_ctor_release(x_357, 0); + lean_ctor_release(x_357, 1); + x_409 = x_357; } else { - lean_dec_ref(x_213); - x_262 = lean_box(0); + lean_dec_ref(x_357); + x_409 = lean_box(0); } -if (lean_is_scalar(x_262)) { - x_263 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_409)) { + x_410 = lean_alloc_ctor(1, 2, 0); } else { - x_263 = x_262; + x_410 = x_409; } -lean_ctor_set(x_263, 0, x_260); -lean_ctor_set(x_263, 1, x_261); -return x_263; +lean_ctor_set(x_410, 0, x_407); +lean_ctor_set(x_410, 1, x_408); +return x_410; } } else { -lean_object* x_264; uint8_t x_265; lean_object* x_266; lean_object* x_267; -lean_dec(x_201); -lean_dec(x_199); +lean_object* x_411; lean_object* x_412; lean_object* x_413; lean_object* x_414; +lean_dec(x_353); +lean_dec(x_352); +lean_dec(x_351); +lean_dec(x_336); +lean_dec(x_335); +lean_dec(x_42); +lean_free_object(x_33); +lean_free_object(x_23); +lean_dec(x_32); +lean_free_object(x_22); +lean_free_object(x_17); +lean_free_object(x_12); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_411 = lean_ctor_get(x_354, 0); +lean_inc(x_411); +x_412 = lean_ctor_get(x_354, 1); +lean_inc(x_412); +if (lean_is_exclusive(x_354)) { + lean_ctor_release(x_354, 0); + lean_ctor_release(x_354, 1); + x_413 = x_354; +} else { + lean_dec_ref(x_354); + x_413 = lean_box(0); +} +if (lean_is_scalar(x_413)) { + x_414 = lean_alloc_ctor(1, 2, 0); +} else { + x_414 = x_413; +} +lean_ctor_set(x_414, 0, x_411); +lean_ctor_set(x_414, 1, x_412); +return x_414; +} +} +else +{ +lean_object* x_415; lean_object* x_416; lean_object* x_417; lean_object* x_418; +lean_dec(x_336); +lean_dec(x_335); +lean_dec(x_42); +lean_free_object(x_33); +lean_free_object(x_23); +lean_dec(x_32); +lean_free_object(x_22); +lean_free_object(x_17); +lean_free_object(x_12); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_415 = lean_ctor_get(x_348, 0); +lean_inc(x_415); +x_416 = lean_ctor_get(x_348, 1); +lean_inc(x_416); +if (lean_is_exclusive(x_348)) { + lean_ctor_release(x_348, 0); + lean_ctor_release(x_348, 1); + x_417 = x_348; +} else { + lean_dec_ref(x_348); + x_417 = lean_box(0); +} +if (lean_is_scalar(x_417)) { + x_418 = lean_alloc_ctor(1, 2, 0); +} else { + x_418 = x_417; +} +lean_ctor_set(x_418, 0, x_415); +lean_ctor_set(x_418, 1, x_416); +return x_418; +} +} +else +{ +lean_object* x_419; uint8_t x_420; lean_object* x_421; lean_object* x_422; +lean_dec(x_336); +lean_dec(x_334); lean_free_object(x_33); lean_free_object(x_23); lean_dec(x_31); lean_free_object(x_22); lean_free_object(x_17); lean_free_object(x_12); -x_264 = lean_ctor_get(x_209, 1); -lean_inc(x_264); -lean_dec(x_209); -x_265 = 0; -x_266 = lean_box(0); -x_267 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__3(x_32, x_200, x_42, x_1, x_265, x_266, x_4, x_5, x_6, x_7, x_264); -return x_267; +x_419 = lean_ctor_get(x_344, 1); +lean_inc(x_419); +lean_dec(x_344); +x_420 = 0; +x_421 = lean_box(0); +x_422 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__3(x_32, x_335, x_42, x_1, x_420, x_421, x_4, x_5, x_6, x_7, x_419); +return x_422; } } else { -lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; -lean_dec(x_201); -lean_dec(x_200); -lean_dec(x_199); +lean_object* x_423; lean_object* x_424; lean_object* x_425; lean_object* x_426; +lean_dec(x_336); +lean_dec(x_335); +lean_dec(x_334); lean_dec(x_42); lean_free_object(x_33); lean_free_object(x_23); @@ -21037,35 +22215,35 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_268 = lean_ctor_get(x_209, 0); -lean_inc(x_268); -x_269 = lean_ctor_get(x_209, 1); -lean_inc(x_269); -if (lean_is_exclusive(x_209)) { - lean_ctor_release(x_209, 0); - lean_ctor_release(x_209, 1); - x_270 = x_209; +x_423 = lean_ctor_get(x_344, 0); +lean_inc(x_423); +x_424 = lean_ctor_get(x_344, 1); +lean_inc(x_424); +if (lean_is_exclusive(x_344)) { + lean_ctor_release(x_344, 0); + lean_ctor_release(x_344, 1); + x_425 = x_344; } else { - lean_dec_ref(x_209); - x_270 = lean_box(0); + lean_dec_ref(x_344); + x_425 = lean_box(0); } -if (lean_is_scalar(x_270)) { - x_271 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_425)) { + x_426 = lean_alloc_ctor(1, 2, 0); } else { - x_271 = x_270; + x_426 = x_425; } -lean_ctor_set(x_271, 0, x_268); -lean_ctor_set(x_271, 1, x_269); -return x_271; +lean_ctor_set(x_426, 0, x_423); +lean_ctor_set(x_426, 1, x_424); +return x_426; } } } else { -lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; -lean_dec(x_201); -lean_dec(x_200); -lean_dec(x_199); +lean_object* x_427; lean_object* x_428; lean_object* x_429; lean_object* x_430; +lean_dec(x_336); +lean_dec(x_335); +lean_dec(x_334); lean_dec(x_42); lean_free_object(x_33); lean_free_object(x_23); @@ -21081,271 +22259,331 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_272 = lean_ctor_get(x_202, 0); -lean_inc(x_272); -x_273 = lean_ctor_get(x_202, 1); -lean_inc(x_273); -if (lean_is_exclusive(x_202)) { - lean_ctor_release(x_202, 0); - lean_ctor_release(x_202, 1); - x_274 = x_202; +x_427 = lean_ctor_get(x_337, 0); +lean_inc(x_427); +x_428 = lean_ctor_get(x_337, 1); +lean_inc(x_428); +if (lean_is_exclusive(x_337)) { + lean_ctor_release(x_337, 0); + lean_ctor_release(x_337, 1); + x_429 = x_337; } else { - lean_dec_ref(x_202); - x_274 = lean_box(0); + lean_dec_ref(x_337); + x_429 = lean_box(0); } -if (lean_is_scalar(x_274)) { - x_275 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_429)) { + x_430 = lean_alloc_ctor(1, 2, 0); } else { - x_275 = x_274; + x_430 = x_429; } -lean_ctor_set(x_275, 0, x_272); -lean_ctor_set(x_275, 1, x_273); -return x_275; +lean_ctor_set(x_430, 0, x_427); +lean_ctor_set(x_430, 1, x_428); +return x_430; } } } else { -lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; -x_276 = lean_ctor_get(x_33, 1); -lean_inc(x_276); +lean_object* x_431; lean_object* x_432; lean_object* x_433; lean_object* x_434; lean_object* x_435; lean_object* x_436; lean_object* x_437; lean_object* x_438; lean_object* x_439; lean_object* x_440; +x_431 = lean_ctor_get(x_33, 1); +lean_inc(x_431); lean_dec(x_33); -x_277 = lean_ctor_get(x_34, 0); -lean_inc(x_277); +x_432 = lean_ctor_get(x_34, 0); +lean_inc(x_432); lean_dec(x_34); lean_inc(x_3); lean_inc(x_16); lean_inc(x_2); -x_278 = lean_alloc_closure((void*)(l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__1___boxed), 9, 3); -lean_closure_set(x_278, 0, x_2); -lean_closure_set(x_278, 1, x_16); -lean_closure_set(x_278, 2, x_3); -x_279 = lean_ctor_get(x_277, 1); -lean_inc(x_279); -x_280 = lean_ctor_get(x_277, 0); -lean_inc(x_280); -if (lean_is_exclusive(x_277)) { - lean_ctor_release(x_277, 0); - lean_ctor_release(x_277, 1); - x_281 = x_277; +x_433 = lean_alloc_closure((void*)(l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__1___boxed), 9, 3); +lean_closure_set(x_433, 0, x_2); +lean_closure_set(x_433, 1, x_16); +lean_closure_set(x_433, 2, x_3); +x_434 = lean_ctor_get(x_432, 1); +lean_inc(x_434); +x_435 = lean_ctor_get(x_432, 0); +lean_inc(x_435); +if (lean_is_exclusive(x_432)) { + lean_ctor_release(x_432, 0); + lean_ctor_release(x_432, 1); + x_436 = x_432; } else { - lean_dec_ref(x_277); - x_281 = lean_box(0); + lean_dec_ref(x_432); + x_436 = lean_box(0); } -x_282 = lean_ctor_get(x_279, 0); -lean_inc(x_282); -x_283 = lean_ctor_get(x_279, 1); -lean_inc(x_283); -if (lean_is_exclusive(x_279)) { - lean_ctor_release(x_279, 0); - lean_ctor_release(x_279, 1); - x_284 = x_279; +x_437 = lean_ctor_get(x_434, 0); +lean_inc(x_437); +x_438 = lean_ctor_get(x_434, 1); +lean_inc(x_438); +if (lean_is_exclusive(x_434)) { + lean_ctor_release(x_434, 0); + lean_ctor_release(x_434, 1); + x_439 = x_434; } else { - lean_dec_ref(x_279); - x_284 = lean_box(0); + lean_dec_ref(x_434); + x_439 = lean_box(0); } lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -x_285 = l_Lean_Meta_isExprDefEqGuarded(x_28, x_280, x_4, x_5, x_6, x_7, x_276); -if (lean_obj_tag(x_285) == 0) +x_440 = l_Lean_Meta_isExprDefEqGuarded(x_28, x_435, x_4, x_5, x_6, x_7, x_431); +if (lean_obj_tag(x_440) == 0) { -lean_object* x_286; uint8_t x_287; -x_286 = lean_ctor_get(x_285, 0); -lean_inc(x_286); -x_287 = lean_unbox(x_286); -lean_dec(x_286); -if (x_287 == 0) +lean_object* x_441; uint8_t x_442; +x_441 = lean_ctor_get(x_440, 0); +lean_inc(x_441); +x_442 = lean_unbox(x_441); +lean_dec(x_441); +if (x_442 == 0) { -lean_object* x_288; lean_object* x_289; lean_object* x_290; -lean_dec(x_284); -lean_dec(x_283); -lean_dec(x_282); -lean_dec(x_281); -lean_dec(x_278); +lean_object* x_443; lean_object* x_444; lean_object* x_445; +lean_dec(x_439); +lean_dec(x_438); +lean_dec(x_437); +lean_dec(x_436); +lean_dec(x_433); lean_free_object(x_23); lean_dec(x_32); lean_dec(x_31); lean_free_object(x_22); lean_free_object(x_17); lean_free_object(x_12); -x_288 = lean_ctor_get(x_285, 1); -lean_inc(x_288); -lean_dec(x_285); -x_289 = lean_box(0); -x_290 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__1(x_2, x_16, x_3, x_289, x_4, x_5, x_6, x_7, x_288); -return x_290; +x_443 = lean_ctor_get(x_440, 1); +lean_inc(x_443); +lean_dec(x_440); +x_444 = lean_box(0); +x_445 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__1(x_2, x_16, x_3, x_444, x_4, x_5, x_6, x_7, x_443); +return x_445; } else { -lean_object* x_291; lean_object* x_292; +lean_object* x_446; lean_object* x_447; lean_dec(x_16); lean_dec(x_3); lean_dec(x_2); -x_291 = lean_ctor_get(x_285, 1); -lean_inc(x_291); -lean_dec(x_285); +x_446 = lean_ctor_get(x_440, 1); +lean_inc(x_446); +lean_dec(x_440); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -lean_inc(x_282); +lean_inc(x_437); lean_inc(x_31); -x_292 = l_Lean_Meta_isExprDefEqGuarded(x_31, x_282, x_4, x_5, x_6, x_7, x_291); -if (lean_obj_tag(x_292) == 0) +x_447 = l_Lean_Meta_isExprDefEqGuarded(x_31, x_437, x_4, x_5, x_6, x_7, x_446); +if (lean_obj_tag(x_447) == 0) +{ +lean_object* x_448; uint8_t x_449; +x_448 = lean_ctor_get(x_447, 0); +lean_inc(x_448); +x_449 = lean_unbox(x_448); +lean_dec(x_448); +if (x_449 == 0) +{ +lean_object* x_450; lean_object* x_451; +x_450 = lean_ctor_get(x_447, 1); +lean_inc(x_450); +lean_dec(x_447); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +x_451 = l_Lean_Meta_addPPExplicitToExposeDiff(x_31, x_437, x_4, x_5, x_6, x_7, x_450); +if (lean_obj_tag(x_451) == 0) { -lean_object* x_293; uint8_t x_294; -x_293 = lean_ctor_get(x_292, 0); -lean_inc(x_293); -x_294 = lean_unbox(x_293); -lean_dec(x_293); -if (x_294 == 0) +lean_object* x_452; lean_object* x_453; lean_object* x_454; lean_object* x_455; lean_object* x_456; lean_object* x_457; +x_452 = lean_ctor_get(x_451, 0); +lean_inc(x_452); +x_453 = lean_ctor_get(x_451, 1); +lean_inc(x_453); +lean_dec(x_451); +x_454 = lean_ctor_get(x_452, 0); +lean_inc(x_454); +x_455 = lean_ctor_get(x_452, 1); +lean_inc(x_455); +if (lean_is_exclusive(x_452)) { + lean_ctor_release(x_452, 0); + lean_ctor_release(x_452, 1); + x_456 = x_452; +} else { + lean_dec_ref(x_452); + x_456 = lean_box(0); +} +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_454); +x_457 = lean_infer_type(x_454, x_4, x_5, x_6, x_7, x_453); +if (lean_obj_tag(x_457) == 0) { -lean_object* x_295; lean_object* x_296; -x_295 = lean_ctor_get(x_292, 1); -lean_inc(x_295); -lean_dec(x_292); +lean_object* x_458; lean_object* x_459; lean_object* x_460; +x_458 = lean_ctor_get(x_457, 0); +lean_inc(x_458); +x_459 = lean_ctor_get(x_457, 1); +lean_inc(x_459); +lean_dec(x_457); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -lean_inc(x_31); -x_296 = lean_infer_type(x_31, x_4, x_5, x_6, x_7, x_295); -if (lean_obj_tag(x_296) == 0) +lean_inc(x_455); +x_460 = lean_infer_type(x_455, x_4, x_5, x_6, x_7, x_459); +if (lean_obj_tag(x_460) == 0) { -lean_object* x_297; lean_object* x_298; lean_object* x_299; -x_297 = lean_ctor_get(x_296, 0); -lean_inc(x_297); -x_298 = lean_ctor_get(x_296, 1); -lean_inc(x_298); -lean_dec(x_296); +lean_object* x_461; lean_object* x_462; lean_object* x_463; +x_461 = lean_ctor_get(x_460, 0); +lean_inc(x_461); +x_462 = lean_ctor_get(x_460, 1); +lean_inc(x_462); +lean_dec(x_460); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -lean_inc(x_282); -x_299 = lean_infer_type(x_282, x_4, x_5, x_6, x_7, x_298); -if (lean_obj_tag(x_299) == 0) +x_463 = l_Lean_Meta_addPPExplicitToExposeDiff(x_458, x_461, x_4, x_5, x_6, x_7, x_462); +if (lean_obj_tag(x_463) == 0) { -lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; uint8_t x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; -x_300 = lean_ctor_get(x_299, 0); -lean_inc(x_300); -x_301 = lean_ctor_get(x_299, 1); -lean_inc(x_301); -lean_dec(x_299); -x_302 = lean_array_get_size(x_1); -x_303 = lean_unsigned_to_nat(0u); -x_304 = lean_nat_dec_lt(x_303, x_302); -lean_dec(x_302); -x_305 = l_Lean_MessageData_ofExpr(x_31); -x_306 = l___private_Lean_Elab_Calc_0__Lean_Elab_Term_getRelUniv___lambda__1___closed__4; -if (lean_is_scalar(x_284)) { - x_307 = lean_alloc_ctor(7, 2, 0); -} else { - x_307 = x_284; - lean_ctor_set_tag(x_307, 7); +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; uint8_t x_471; lean_object* x_472; lean_object* x_473; lean_object* x_474; lean_object* x_475; lean_object* x_476; lean_object* x_477; lean_object* x_478; lean_object* x_479; lean_object* x_480; lean_object* x_481; lean_object* x_482; lean_object* x_483; lean_object* x_484; lean_object* x_485; lean_object* x_486; lean_object* x_487; lean_object* x_488; lean_object* x_489; +x_464 = lean_ctor_get(x_463, 0); +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); +lean_inc(x_466); +x_467 = lean_ctor_get(x_464, 1); +lean_inc(x_467); +if (lean_is_exclusive(x_464)) { + lean_ctor_release(x_464, 0); + lean_ctor_release(x_464, 1); + x_468 = x_464; +} else { + lean_dec_ref(x_464); + x_468 = lean_box(0); } -lean_ctor_set(x_307, 0, x_306); -lean_ctor_set(x_307, 1, x_305); -x_308 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Term_elabCalcSteps___spec__1___closed__6; -if (lean_is_scalar(x_281)) { - x_309 = lean_alloc_ctor(7, 2, 0); +x_469 = lean_array_get_size(x_1); +x_470 = lean_unsigned_to_nat(0u); +x_471 = lean_nat_dec_lt(x_470, x_469); +lean_dec(x_469); +x_472 = l_Lean_MessageData_ofExpr(x_454); +x_473 = l___private_Lean_Elab_Calc_0__Lean_Elab_Term_getRelUniv___lambda__1___closed__4; +if (lean_is_scalar(x_468)) { + x_474 = lean_alloc_ctor(7, 2, 0); } else { - x_309 = x_281; - lean_ctor_set_tag(x_309, 7); + x_474 = x_468; + lean_ctor_set_tag(x_474, 7); } -lean_ctor_set(x_309, 0, x_307); -lean_ctor_set(x_309, 1, x_308); -x_310 = l_Lean_MessageData_ofExpr(x_297); -x_311 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_311, 0, x_309); -lean_ctor_set(x_311, 1, x_310); +lean_ctor_set(x_474, 0, x_473); +lean_ctor_set(x_474, 1, x_472); +x_475 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Term_elabCalcSteps___spec__1___closed__6; +if (lean_is_scalar(x_456)) { + x_476 = lean_alloc_ctor(7, 2, 0); +} else { + x_476 = x_456; + lean_ctor_set_tag(x_476, 7); +} +lean_ctor_set(x_476, 0, x_474); +lean_ctor_set(x_476, 1, x_475); +x_477 = l_Lean_MessageData_ofExpr(x_466); +if (lean_is_scalar(x_439)) { + x_478 = lean_alloc_ctor(7, 2, 0); +} else { + x_478 = x_439; + lean_ctor_set_tag(x_478, 7); +} +lean_ctor_set(x_478, 0, x_476); +lean_ctor_set(x_478, 1, x_477); +if (lean_is_scalar(x_436)) { + x_479 = lean_alloc_ctor(7, 2, 0); +} else { + x_479 = x_436; + lean_ctor_set_tag(x_479, 7); +} +lean_ctor_set(x_479, 0, x_478); +lean_ctor_set(x_479, 1, x_473); +x_480 = l_Lean_indentD(x_479); +x_481 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Term_elabCalcSteps___spec__1___closed__4; +x_482 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_482, 0, x_481); +lean_ctor_set(x_482, 1, x_480); +x_483 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__3___closed__4; lean_ctor_set_tag(x_23, 7); -lean_ctor_set(x_23, 1, x_306); -lean_ctor_set(x_23, 0, x_311); -x_312 = l_Lean_indentD(x_23); -x_313 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Term_elabCalcSteps___spec__1___closed__4; +lean_ctor_set(x_23, 1, x_483); +lean_ctor_set(x_23, 0, x_482); +x_484 = l_Lean_MessageData_ofExpr(x_455); lean_ctor_set_tag(x_22, 7); -lean_ctor_set(x_22, 1, x_312); -lean_ctor_set(x_22, 0, x_313); -x_314 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__3___closed__4; +lean_ctor_set(x_22, 1, x_484); +lean_ctor_set(x_22, 0, x_473); lean_ctor_set_tag(x_17, 7); -lean_ctor_set(x_17, 1, x_314); +lean_ctor_set(x_17, 1, x_475); lean_ctor_set(x_17, 0, x_22); -x_315 = l_Lean_MessageData_ofExpr(x_282); +x_485 = l_Lean_MessageData_ofExpr(x_467); lean_ctor_set_tag(x_12, 7); -lean_ctor_set(x_12, 1, x_315); -lean_ctor_set(x_12, 0, x_306); -x_316 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_316, 0, x_12); -lean_ctor_set(x_316, 1, x_308); -x_317 = l_Lean_MessageData_ofExpr(x_300); -x_318 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_318, 0, x_316); -lean_ctor_set(x_318, 1, x_317); -x_319 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_319, 0, x_318); -lean_ctor_set(x_319, 1, x_306); -x_320 = l_Lean_indentD(x_319); -x_321 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_321, 0, x_17); -lean_ctor_set(x_321, 1, x_320); -x_322 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_322, 0, x_321); -lean_ctor_set(x_322, 1, x_306); -if (x_304 == 0) -{ -lean_object* x_323; lean_object* x_324; lean_object* x_325; uint8_t x_326; lean_object* x_327; lean_object* x_328; uint8_t x_329; lean_object* x_330; lean_object* x_331; -x_323 = l_Lean_Elab_Term_instInhabitedCalcStepView; -x_324 = l_outOfBounds___rarg(x_323); -x_325 = lean_ctor_get(x_324, 1); -lean_inc(x_325); -lean_dec(x_324); -x_326 = 2; +lean_ctor_set(x_12, 1, x_485); +lean_ctor_set(x_12, 0, x_17); +x_486 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_486, 0, x_12); +lean_ctor_set(x_486, 1, x_473); +x_487 = l_Lean_indentD(x_486); +x_488 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_488, 0, x_23); +lean_ctor_set(x_488, 1, x_487); +x_489 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_489, 0, x_488); +lean_ctor_set(x_489, 1, x_473); +if (x_471 == 0) +{ +lean_object* x_490; lean_object* x_491; lean_object* x_492; uint8_t x_493; lean_object* x_494; lean_object* x_495; uint8_t x_496; lean_object* x_497; lean_object* x_498; +x_490 = l_Lean_Elab_Term_instInhabitedCalcStepView; +x_491 = l_outOfBounds___rarg(x_490); +x_492 = lean_ctor_get(x_491, 1); +lean_inc(x_492); +lean_dec(x_491); +x_493 = 2; lean_inc(x_6); -x_327 = l_Lean_logAt___at_Lean_Elab_Term_reportUnsolvedGoals___spec__2(x_325, x_322, x_326, x_4, x_5, x_6, x_7, x_301); -lean_dec(x_325); -x_328 = lean_ctor_get(x_327, 1); -lean_inc(x_328); -lean_dec(x_327); -x_329 = 1; -x_330 = lean_box(0); -x_331 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__3(x_32, x_283, x_278, x_1, x_329, x_330, x_4, x_5, x_6, x_7, x_328); -return x_331; +x_494 = l_Lean_logAt___at_Lean_Elab_Term_reportUnsolvedGoals___spec__2(x_492, x_489, x_493, x_4, x_5, x_6, x_7, x_465); +lean_dec(x_492); +x_495 = lean_ctor_get(x_494, 1); +lean_inc(x_495); +lean_dec(x_494); +x_496 = 1; +x_497 = lean_box(0); +x_498 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__3(x_32, x_438, x_433, x_1, x_496, x_497, x_4, x_5, x_6, x_7, x_495); +return x_498; } else { -lean_object* x_332; lean_object* x_333; uint8_t x_334; lean_object* x_335; lean_object* x_336; uint8_t x_337; lean_object* x_338; lean_object* x_339; -x_332 = lean_array_fget(x_1, x_303); -x_333 = lean_ctor_get(x_332, 1); -lean_inc(x_333); -lean_dec(x_332); -x_334 = 2; +lean_object* x_499; lean_object* x_500; uint8_t x_501; lean_object* x_502; lean_object* x_503; uint8_t x_504; lean_object* x_505; lean_object* x_506; +x_499 = lean_array_fget(x_1, x_470); +x_500 = lean_ctor_get(x_499, 1); +lean_inc(x_500); +lean_dec(x_499); +x_501 = 2; lean_inc(x_6); -x_335 = l_Lean_logAt___at_Lean_Elab_Term_reportUnsolvedGoals___spec__2(x_333, x_322, x_334, x_4, x_5, x_6, x_7, x_301); -lean_dec(x_333); -x_336 = lean_ctor_get(x_335, 1); -lean_inc(x_336); -lean_dec(x_335); -x_337 = 1; -x_338 = lean_box(0); -x_339 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__3(x_32, x_283, x_278, x_1, x_337, x_338, x_4, x_5, x_6, x_7, x_336); -return x_339; +x_502 = l_Lean_logAt___at_Lean_Elab_Term_reportUnsolvedGoals___spec__2(x_500, x_489, x_501, x_4, x_5, x_6, x_7, x_465); +lean_dec(x_500); +x_503 = lean_ctor_get(x_502, 1); +lean_inc(x_503); +lean_dec(x_502); +x_504 = 1; +x_505 = lean_box(0); +x_506 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__3(x_32, x_438, x_433, x_1, x_504, x_505, x_4, x_5, x_6, x_7, x_503); +return x_506; } } else { -lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; -lean_dec(x_297); -lean_dec(x_284); -lean_dec(x_283); -lean_dec(x_282); -lean_dec(x_281); -lean_dec(x_278); +lean_object* x_507; lean_object* x_508; lean_object* x_509; lean_object* x_510; +lean_dec(x_456); +lean_dec(x_455); +lean_dec(x_454); +lean_dec(x_439); +lean_dec(x_438); +lean_dec(x_436); +lean_dec(x_433); lean_free_object(x_23); lean_dec(x_32); -lean_dec(x_31); lean_free_object(x_22); lean_free_object(x_17); lean_free_object(x_12); @@ -21353,39 +22591,41 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_340 = lean_ctor_get(x_299, 0); -lean_inc(x_340); -x_341 = lean_ctor_get(x_299, 1); -lean_inc(x_341); -if (lean_is_exclusive(x_299)) { - lean_ctor_release(x_299, 0); - lean_ctor_release(x_299, 1); - x_342 = x_299; +x_507 = lean_ctor_get(x_463, 0); +lean_inc(x_507); +x_508 = lean_ctor_get(x_463, 1); +lean_inc(x_508); +if (lean_is_exclusive(x_463)) { + lean_ctor_release(x_463, 0); + lean_ctor_release(x_463, 1); + x_509 = x_463; } else { - lean_dec_ref(x_299); - x_342 = lean_box(0); + lean_dec_ref(x_463); + x_509 = lean_box(0); } -if (lean_is_scalar(x_342)) { - x_343 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_509)) { + x_510 = lean_alloc_ctor(1, 2, 0); } else { - x_343 = x_342; + x_510 = x_509; } -lean_ctor_set(x_343, 0, x_340); -lean_ctor_set(x_343, 1, x_341); -return x_343; +lean_ctor_set(x_510, 0, x_507); +lean_ctor_set(x_510, 1, x_508); +return x_510; } } else { -lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; -lean_dec(x_284); -lean_dec(x_283); -lean_dec(x_282); -lean_dec(x_281); -lean_dec(x_278); +lean_object* x_511; lean_object* x_512; lean_object* x_513; lean_object* x_514; +lean_dec(x_458); +lean_dec(x_456); +lean_dec(x_455); +lean_dec(x_454); +lean_dec(x_439); +lean_dec(x_438); +lean_dec(x_436); +lean_dec(x_433); lean_free_object(x_23); lean_dec(x_32); -lean_dec(x_31); lean_free_object(x_22); lean_free_object(x_17); lean_free_object(x_12); @@ -21393,59 +22633,78 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_344 = lean_ctor_get(x_296, 0); -lean_inc(x_344); -x_345 = lean_ctor_get(x_296, 1); -lean_inc(x_345); -if (lean_is_exclusive(x_296)) { - lean_ctor_release(x_296, 0); - lean_ctor_release(x_296, 1); - x_346 = x_296; +x_511 = lean_ctor_get(x_460, 0); +lean_inc(x_511); +x_512 = lean_ctor_get(x_460, 1); +lean_inc(x_512); +if (lean_is_exclusive(x_460)) { + lean_ctor_release(x_460, 0); + lean_ctor_release(x_460, 1); + x_513 = x_460; } else { - lean_dec_ref(x_296); - x_346 = lean_box(0); + lean_dec_ref(x_460); + x_513 = lean_box(0); } -if (lean_is_scalar(x_346)) { - x_347 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_513)) { + x_514 = lean_alloc_ctor(1, 2, 0); } else { - x_347 = x_346; + x_514 = x_513; } -lean_ctor_set(x_347, 0, x_344); -lean_ctor_set(x_347, 1, x_345); -return x_347; +lean_ctor_set(x_514, 0, x_511); +lean_ctor_set(x_514, 1, x_512); +return x_514; } } else { -lean_object* x_348; uint8_t x_349; lean_object* x_350; lean_object* x_351; -lean_dec(x_284); -lean_dec(x_282); -lean_dec(x_281); +lean_object* x_515; lean_object* x_516; lean_object* x_517; lean_object* x_518; +lean_dec(x_456); +lean_dec(x_455); +lean_dec(x_454); +lean_dec(x_439); +lean_dec(x_438); +lean_dec(x_436); +lean_dec(x_433); lean_free_object(x_23); -lean_dec(x_31); +lean_dec(x_32); lean_free_object(x_22); lean_free_object(x_17); lean_free_object(x_12); -x_348 = lean_ctor_get(x_292, 1); -lean_inc(x_348); -lean_dec(x_292); -x_349 = 0; -x_350 = lean_box(0); -x_351 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__3(x_32, x_283, x_278, x_1, x_349, x_350, x_4, x_5, x_6, x_7, x_348); -return x_351; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_515 = lean_ctor_get(x_457, 0); +lean_inc(x_515); +x_516 = lean_ctor_get(x_457, 1); +lean_inc(x_516); +if (lean_is_exclusive(x_457)) { + lean_ctor_release(x_457, 0); + lean_ctor_release(x_457, 1); + x_517 = x_457; +} else { + lean_dec_ref(x_457); + x_517 = lean_box(0); +} +if (lean_is_scalar(x_517)) { + x_518 = lean_alloc_ctor(1, 2, 0); +} else { + x_518 = x_517; +} +lean_ctor_set(x_518, 0, x_515); +lean_ctor_set(x_518, 1, x_516); +return x_518; } } else { -lean_object* x_352; lean_object* x_353; lean_object* x_354; lean_object* x_355; -lean_dec(x_284); -lean_dec(x_283); -lean_dec(x_282); -lean_dec(x_281); -lean_dec(x_278); +lean_object* x_519; lean_object* x_520; lean_object* x_521; lean_object* x_522; +lean_dec(x_439); +lean_dec(x_438); +lean_dec(x_436); +lean_dec(x_433); lean_free_object(x_23); lean_dec(x_32); -lean_dec(x_31); lean_free_object(x_22); lean_free_object(x_17); lean_free_object(x_12); @@ -21453,37 +22712,97 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_352 = lean_ctor_get(x_292, 0); -lean_inc(x_352); -x_353 = lean_ctor_get(x_292, 1); -lean_inc(x_353); -if (lean_is_exclusive(x_292)) { - lean_ctor_release(x_292, 0); - lean_ctor_release(x_292, 1); - x_354 = x_292; +x_519 = lean_ctor_get(x_451, 0); +lean_inc(x_519); +x_520 = lean_ctor_get(x_451, 1); +lean_inc(x_520); +if (lean_is_exclusive(x_451)) { + lean_ctor_release(x_451, 0); + lean_ctor_release(x_451, 1); + x_521 = x_451; } else { - lean_dec_ref(x_292); - x_354 = lean_box(0); + lean_dec_ref(x_451); + x_521 = lean_box(0); } -if (lean_is_scalar(x_354)) { - x_355 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_521)) { + x_522 = lean_alloc_ctor(1, 2, 0); } else { - x_355 = x_354; + x_522 = x_521; +} +lean_ctor_set(x_522, 0, x_519); +lean_ctor_set(x_522, 1, x_520); +return x_522; } -lean_ctor_set(x_355, 0, x_352); -lean_ctor_set(x_355, 1, x_353); -return x_355; } +else +{ +lean_object* x_523; uint8_t x_524; lean_object* x_525; lean_object* x_526; +lean_dec(x_439); +lean_dec(x_437); +lean_dec(x_436); +lean_free_object(x_23); +lean_dec(x_31); +lean_free_object(x_22); +lean_free_object(x_17); +lean_free_object(x_12); +x_523 = lean_ctor_get(x_447, 1); +lean_inc(x_523); +lean_dec(x_447); +x_524 = 0; +x_525 = lean_box(0); +x_526 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__3(x_32, x_438, x_433, x_1, x_524, x_525, x_4, x_5, x_6, x_7, x_523); +return x_526; } } else { -lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; -lean_dec(x_284); -lean_dec(x_283); -lean_dec(x_282); -lean_dec(x_281); -lean_dec(x_278); +lean_object* x_527; lean_object* x_528; lean_object* x_529; lean_object* x_530; +lean_dec(x_439); +lean_dec(x_438); +lean_dec(x_437); +lean_dec(x_436); +lean_dec(x_433); +lean_free_object(x_23); +lean_dec(x_32); +lean_dec(x_31); +lean_free_object(x_22); +lean_free_object(x_17); +lean_free_object(x_12); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_527 = lean_ctor_get(x_447, 0); +lean_inc(x_527); +x_528 = lean_ctor_get(x_447, 1); +lean_inc(x_528); +if (lean_is_exclusive(x_447)) { + lean_ctor_release(x_447, 0); + lean_ctor_release(x_447, 1); + x_529 = x_447; +} else { + lean_dec_ref(x_447); + x_529 = lean_box(0); +} +if (lean_is_scalar(x_529)) { + x_530 = lean_alloc_ctor(1, 2, 0); +} else { + x_530 = x_529; +} +lean_ctor_set(x_530, 0, x_527); +lean_ctor_set(x_530, 1, x_528); +return x_530; +} +} +} +else +{ +lean_object* x_531; lean_object* x_532; lean_object* x_533; lean_object* x_534; +lean_dec(x_439); +lean_dec(x_438); +lean_dec(x_437); +lean_dec(x_436); +lean_dec(x_433); lean_free_object(x_23); lean_dec(x_32); lean_dec(x_31); @@ -21497,311 +22816,371 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_356 = lean_ctor_get(x_285, 0); -lean_inc(x_356); -x_357 = lean_ctor_get(x_285, 1); -lean_inc(x_357); -if (lean_is_exclusive(x_285)) { - lean_ctor_release(x_285, 0); - lean_ctor_release(x_285, 1); - x_358 = x_285; +x_531 = lean_ctor_get(x_440, 0); +lean_inc(x_531); +x_532 = lean_ctor_get(x_440, 1); +lean_inc(x_532); +if (lean_is_exclusive(x_440)) { + lean_ctor_release(x_440, 0); + lean_ctor_release(x_440, 1); + x_533 = x_440; } else { - lean_dec_ref(x_285); - x_358 = lean_box(0); + lean_dec_ref(x_440); + x_533 = lean_box(0); } -if (lean_is_scalar(x_358)) { - x_359 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_533)) { + x_534 = lean_alloc_ctor(1, 2, 0); } else { - x_359 = x_358; + x_534 = x_533; } -lean_ctor_set(x_359, 0, x_356); -lean_ctor_set(x_359, 1, x_357); -return x_359; +lean_ctor_set(x_534, 0, x_531); +lean_ctor_set(x_534, 1, x_532); +return x_534; } } } } else { -lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; -x_360 = lean_ctor_get(x_23, 0); -x_361 = lean_ctor_get(x_23, 1); -lean_inc(x_361); -lean_inc(x_360); +lean_object* x_535; lean_object* x_536; lean_object* x_537; lean_object* x_538; +x_535 = lean_ctor_get(x_23, 0); +x_536 = lean_ctor_get(x_23, 1); +lean_inc(x_536); +lean_inc(x_535); lean_dec(x_23); -x_362 = l_Lean_Elab_Term_getCalcRelation_x3f(x_2, x_4, x_5, x_6, x_7, x_25); -x_363 = lean_ctor_get(x_362, 0); -lean_inc(x_363); -if (lean_obj_tag(x_363) == 0) +x_537 = l_Lean_Elab_Term_getCalcRelation_x3f(x_2, x_4, x_5, x_6, x_7, x_25); +x_538 = lean_ctor_get(x_537, 0); +lean_inc(x_538); +if (lean_obj_tag(x_538) == 0) { -lean_object* x_364; lean_object* x_365; lean_object* x_366; -lean_dec(x_361); -lean_dec(x_360); +lean_object* x_539; lean_object* x_540; lean_object* x_541; +lean_dec(x_536); +lean_dec(x_535); lean_free_object(x_22); lean_dec(x_28); lean_free_object(x_17); lean_free_object(x_12); -x_364 = lean_ctor_get(x_362, 1); -lean_inc(x_364); -lean_dec(x_362); -x_365 = lean_box(0); -x_366 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__1(x_2, x_16, x_3, x_365, x_4, x_5, x_6, x_7, x_364); -return x_366; +x_539 = lean_ctor_get(x_537, 1); +lean_inc(x_539); +lean_dec(x_537); +x_540 = lean_box(0); +x_541 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__1(x_2, x_16, x_3, x_540, x_4, x_5, x_6, x_7, x_539); +return x_541; } else { -lean_object* x_367; lean_object* x_368; lean_object* x_369; lean_object* x_370; lean_object* x_371; lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; lean_object* x_376; lean_object* x_377; -x_367 = lean_ctor_get(x_362, 1); -lean_inc(x_367); -if (lean_is_exclusive(x_362)) { - lean_ctor_release(x_362, 0); - lean_ctor_release(x_362, 1); - x_368 = x_362; +lean_object* x_542; lean_object* x_543; lean_object* x_544; lean_object* x_545; lean_object* x_546; lean_object* x_547; lean_object* x_548; lean_object* x_549; lean_object* x_550; lean_object* x_551; lean_object* x_552; +x_542 = lean_ctor_get(x_537, 1); +lean_inc(x_542); +if (lean_is_exclusive(x_537)) { + lean_ctor_release(x_537, 0); + lean_ctor_release(x_537, 1); + x_543 = x_537; } else { - lean_dec_ref(x_362); - x_368 = lean_box(0); + lean_dec_ref(x_537); + x_543 = lean_box(0); } -x_369 = lean_ctor_get(x_363, 0); -lean_inc(x_369); -lean_dec(x_363); +x_544 = lean_ctor_get(x_538, 0); +lean_inc(x_544); +lean_dec(x_538); lean_inc(x_3); lean_inc(x_16); lean_inc(x_2); -x_370 = lean_alloc_closure((void*)(l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__1___boxed), 9, 3); -lean_closure_set(x_370, 0, x_2); -lean_closure_set(x_370, 1, x_16); -lean_closure_set(x_370, 2, x_3); -x_371 = lean_ctor_get(x_369, 1); -lean_inc(x_371); -x_372 = lean_ctor_get(x_369, 0); -lean_inc(x_372); -if (lean_is_exclusive(x_369)) { - lean_ctor_release(x_369, 0); - lean_ctor_release(x_369, 1); - x_373 = x_369; -} else { - lean_dec_ref(x_369); - x_373 = lean_box(0); -} -x_374 = lean_ctor_get(x_371, 0); -lean_inc(x_374); -x_375 = lean_ctor_get(x_371, 1); -lean_inc(x_375); -if (lean_is_exclusive(x_371)) { - lean_ctor_release(x_371, 0); - lean_ctor_release(x_371, 1); - x_376 = x_371; -} else { - lean_dec_ref(x_371); - x_376 = lean_box(0); +x_545 = lean_alloc_closure((void*)(l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__1___boxed), 9, 3); +lean_closure_set(x_545, 0, x_2); +lean_closure_set(x_545, 1, x_16); +lean_closure_set(x_545, 2, x_3); +x_546 = lean_ctor_get(x_544, 1); +lean_inc(x_546); +x_547 = lean_ctor_get(x_544, 0); +lean_inc(x_547); +if (lean_is_exclusive(x_544)) { + lean_ctor_release(x_544, 0); + lean_ctor_release(x_544, 1); + x_548 = x_544; +} else { + lean_dec_ref(x_544); + x_548 = lean_box(0); +} +x_549 = lean_ctor_get(x_546, 0); +lean_inc(x_549); +x_550 = lean_ctor_get(x_546, 1); +lean_inc(x_550); +if (lean_is_exclusive(x_546)) { + lean_ctor_release(x_546, 0); + lean_ctor_release(x_546, 1); + x_551 = x_546; +} else { + lean_dec_ref(x_546); + x_551 = lean_box(0); } lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -x_377 = l_Lean_Meta_isExprDefEqGuarded(x_28, x_372, x_4, x_5, x_6, x_7, x_367); -if (lean_obj_tag(x_377) == 0) +x_552 = l_Lean_Meta_isExprDefEqGuarded(x_28, x_547, x_4, x_5, x_6, x_7, x_542); +if (lean_obj_tag(x_552) == 0) { -lean_object* x_378; uint8_t x_379; -x_378 = lean_ctor_get(x_377, 0); -lean_inc(x_378); -x_379 = lean_unbox(x_378); -lean_dec(x_378); -if (x_379 == 0) +lean_object* x_553; uint8_t x_554; +x_553 = lean_ctor_get(x_552, 0); +lean_inc(x_553); +x_554 = lean_unbox(x_553); +lean_dec(x_553); +if (x_554 == 0) { -lean_object* x_380; lean_object* x_381; lean_object* x_382; -lean_dec(x_376); -lean_dec(x_375); -lean_dec(x_374); -lean_dec(x_373); -lean_dec(x_370); -lean_dec(x_368); -lean_dec(x_361); -lean_dec(x_360); +lean_object* x_555; lean_object* x_556; lean_object* x_557; +lean_dec(x_551); +lean_dec(x_550); +lean_dec(x_549); +lean_dec(x_548); +lean_dec(x_545); +lean_dec(x_543); +lean_dec(x_536); +lean_dec(x_535); lean_free_object(x_22); lean_free_object(x_17); lean_free_object(x_12); -x_380 = lean_ctor_get(x_377, 1); -lean_inc(x_380); -lean_dec(x_377); -x_381 = lean_box(0); -x_382 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__1(x_2, x_16, x_3, x_381, x_4, x_5, x_6, x_7, x_380); -return x_382; +x_555 = lean_ctor_get(x_552, 1); +lean_inc(x_555); +lean_dec(x_552); +x_556 = lean_box(0); +x_557 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__1(x_2, x_16, x_3, x_556, x_4, x_5, x_6, x_7, x_555); +return x_557; } else { -lean_object* x_383; lean_object* x_384; +lean_object* x_558; lean_object* x_559; lean_dec(x_16); lean_dec(x_3); lean_dec(x_2); -x_383 = lean_ctor_get(x_377, 1); -lean_inc(x_383); -lean_dec(x_377); +x_558 = lean_ctor_get(x_552, 1); +lean_inc(x_558); +lean_dec(x_552); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -lean_inc(x_374); -lean_inc(x_360); -x_384 = l_Lean_Meta_isExprDefEqGuarded(x_360, x_374, x_4, x_5, x_6, x_7, x_383); -if (lean_obj_tag(x_384) == 0) +lean_inc(x_549); +lean_inc(x_535); +x_559 = l_Lean_Meta_isExprDefEqGuarded(x_535, x_549, x_4, x_5, x_6, x_7, x_558); +if (lean_obj_tag(x_559) == 0) { -lean_object* x_385; uint8_t x_386; -x_385 = lean_ctor_get(x_384, 0); -lean_inc(x_385); -x_386 = lean_unbox(x_385); -lean_dec(x_385); -if (x_386 == 0) +lean_object* x_560; uint8_t x_561; +x_560 = lean_ctor_get(x_559, 0); +lean_inc(x_560); +x_561 = lean_unbox(x_560); +lean_dec(x_560); +if (x_561 == 0) { -lean_object* x_387; lean_object* x_388; -x_387 = lean_ctor_get(x_384, 1); -lean_inc(x_387); -lean_dec(x_384); +lean_object* x_562; lean_object* x_563; +x_562 = lean_ctor_get(x_559, 1); +lean_inc(x_562); +lean_dec(x_559); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -lean_inc(x_360); -x_388 = lean_infer_type(x_360, x_4, x_5, x_6, x_7, x_387); -if (lean_obj_tag(x_388) == 0) +x_563 = l_Lean_Meta_addPPExplicitToExposeDiff(x_535, x_549, x_4, x_5, x_6, x_7, x_562); +if (lean_obj_tag(x_563) == 0) { -lean_object* x_389; lean_object* x_390; lean_object* x_391; -x_389 = lean_ctor_get(x_388, 0); -lean_inc(x_389); -x_390 = lean_ctor_get(x_388, 1); -lean_inc(x_390); -lean_dec(x_388); +lean_object* x_564; lean_object* x_565; lean_object* x_566; lean_object* x_567; lean_object* x_568; lean_object* x_569; +x_564 = lean_ctor_get(x_563, 0); +lean_inc(x_564); +x_565 = lean_ctor_get(x_563, 1); +lean_inc(x_565); +lean_dec(x_563); +x_566 = lean_ctor_get(x_564, 0); +lean_inc(x_566); +x_567 = lean_ctor_get(x_564, 1); +lean_inc(x_567); +if (lean_is_exclusive(x_564)) { + lean_ctor_release(x_564, 0); + lean_ctor_release(x_564, 1); + x_568 = x_564; +} else { + lean_dec_ref(x_564); + x_568 = lean_box(0); +} lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -lean_inc(x_374); -x_391 = lean_infer_type(x_374, x_4, x_5, x_6, x_7, x_390); -if (lean_obj_tag(x_391) == 0) +lean_inc(x_566); +x_569 = lean_infer_type(x_566, x_4, x_5, x_6, x_7, x_565); +if (lean_obj_tag(x_569) == 0) { -lean_object* x_392; lean_object* x_393; lean_object* x_394; lean_object* x_395; uint8_t x_396; lean_object* x_397; lean_object* x_398; lean_object* x_399; lean_object* x_400; lean_object* x_401; lean_object* x_402; lean_object* x_403; lean_object* x_404; lean_object* x_405; lean_object* x_406; lean_object* x_407; lean_object* x_408; lean_object* x_409; lean_object* x_410; lean_object* x_411; lean_object* x_412; lean_object* x_413; lean_object* x_414; lean_object* x_415; -x_392 = lean_ctor_get(x_391, 0); -lean_inc(x_392); -x_393 = lean_ctor_get(x_391, 1); -lean_inc(x_393); -lean_dec(x_391); -x_394 = lean_array_get_size(x_1); -x_395 = lean_unsigned_to_nat(0u); -x_396 = lean_nat_dec_lt(x_395, x_394); -lean_dec(x_394); -x_397 = l_Lean_MessageData_ofExpr(x_360); -x_398 = l___private_Lean_Elab_Calc_0__Lean_Elab_Term_getRelUniv___lambda__1___closed__4; -if (lean_is_scalar(x_376)) { - x_399 = lean_alloc_ctor(7, 2, 0); +lean_object* x_570; lean_object* x_571; lean_object* x_572; +x_570 = lean_ctor_get(x_569, 0); +lean_inc(x_570); +x_571 = lean_ctor_get(x_569, 1); +lean_inc(x_571); +lean_dec(x_569); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_567); +x_572 = lean_infer_type(x_567, x_4, x_5, x_6, x_7, x_571); +if (lean_obj_tag(x_572) == 0) +{ +lean_object* x_573; lean_object* x_574; lean_object* x_575; +x_573 = lean_ctor_get(x_572, 0); +lean_inc(x_573); +x_574 = lean_ctor_get(x_572, 1); +lean_inc(x_574); +lean_dec(x_572); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +x_575 = l_Lean_Meta_addPPExplicitToExposeDiff(x_570, x_573, x_4, x_5, x_6, x_7, x_574); +if (lean_obj_tag(x_575) == 0) +{ +lean_object* x_576; lean_object* x_577; lean_object* x_578; lean_object* x_579; lean_object* x_580; lean_object* x_581; lean_object* x_582; uint8_t x_583; lean_object* x_584; lean_object* x_585; lean_object* x_586; lean_object* x_587; lean_object* x_588; lean_object* x_589; lean_object* x_590; lean_object* x_591; lean_object* x_592; lean_object* x_593; lean_object* x_594; lean_object* x_595; lean_object* x_596; lean_object* x_597; lean_object* x_598; lean_object* x_599; lean_object* x_600; lean_object* x_601; lean_object* x_602; +x_576 = lean_ctor_get(x_575, 0); +lean_inc(x_576); +x_577 = lean_ctor_get(x_575, 1); +lean_inc(x_577); +lean_dec(x_575); +x_578 = lean_ctor_get(x_576, 0); +lean_inc(x_578); +x_579 = lean_ctor_get(x_576, 1); +lean_inc(x_579); +if (lean_is_exclusive(x_576)) { + lean_ctor_release(x_576, 0); + lean_ctor_release(x_576, 1); + x_580 = x_576; +} else { + lean_dec_ref(x_576); + x_580 = lean_box(0); +} +x_581 = lean_array_get_size(x_1); +x_582 = lean_unsigned_to_nat(0u); +x_583 = lean_nat_dec_lt(x_582, x_581); +lean_dec(x_581); +x_584 = l_Lean_MessageData_ofExpr(x_566); +x_585 = l___private_Lean_Elab_Calc_0__Lean_Elab_Term_getRelUniv___lambda__1___closed__4; +if (lean_is_scalar(x_580)) { + x_586 = lean_alloc_ctor(7, 2, 0); +} else { + x_586 = x_580; + lean_ctor_set_tag(x_586, 7); +} +lean_ctor_set(x_586, 0, x_585); +lean_ctor_set(x_586, 1, x_584); +x_587 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Term_elabCalcSteps___spec__1___closed__6; +if (lean_is_scalar(x_568)) { + x_588 = lean_alloc_ctor(7, 2, 0); +} else { + x_588 = x_568; + lean_ctor_set_tag(x_588, 7); +} +lean_ctor_set(x_588, 0, x_586); +lean_ctor_set(x_588, 1, x_587); +x_589 = l_Lean_MessageData_ofExpr(x_578); +if (lean_is_scalar(x_551)) { + x_590 = lean_alloc_ctor(7, 2, 0); } else { - x_399 = x_376; - lean_ctor_set_tag(x_399, 7); + x_590 = x_551; + lean_ctor_set_tag(x_590, 7); } -lean_ctor_set(x_399, 0, x_398); -lean_ctor_set(x_399, 1, x_397); -x_400 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Term_elabCalcSteps___spec__1___closed__6; -if (lean_is_scalar(x_373)) { - x_401 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_590, 0, x_588); +lean_ctor_set(x_590, 1, x_589); +if (lean_is_scalar(x_548)) { + x_591 = lean_alloc_ctor(7, 2, 0); } else { - x_401 = x_373; - lean_ctor_set_tag(x_401, 7); + x_591 = x_548; + lean_ctor_set_tag(x_591, 7); } -lean_ctor_set(x_401, 0, x_399); -lean_ctor_set(x_401, 1, x_400); -x_402 = l_Lean_MessageData_ofExpr(x_389); -if (lean_is_scalar(x_368)) { - x_403 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_591, 0, x_590); +lean_ctor_set(x_591, 1, x_585); +x_592 = l_Lean_indentD(x_591); +x_593 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Term_elabCalcSteps___spec__1___closed__4; +if (lean_is_scalar(x_543)) { + x_594 = lean_alloc_ctor(7, 2, 0); } else { - x_403 = x_368; - lean_ctor_set_tag(x_403, 7); + x_594 = x_543; + lean_ctor_set_tag(x_594, 7); } -lean_ctor_set(x_403, 0, x_401); -lean_ctor_set(x_403, 1, x_402); -x_404 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_404, 0, x_403); -lean_ctor_set(x_404, 1, x_398); -x_405 = l_Lean_indentD(x_404); -x_406 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Term_elabCalcSteps___spec__1___closed__4; +lean_ctor_set(x_594, 0, x_593); +lean_ctor_set(x_594, 1, x_592); +x_595 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__3___closed__4; +x_596 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_596, 0, x_594); +lean_ctor_set(x_596, 1, x_595); +x_597 = l_Lean_MessageData_ofExpr(x_567); lean_ctor_set_tag(x_22, 7); -lean_ctor_set(x_22, 1, x_405); -lean_ctor_set(x_22, 0, x_406); -x_407 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__3___closed__4; +lean_ctor_set(x_22, 1, x_597); +lean_ctor_set(x_22, 0, x_585); lean_ctor_set_tag(x_17, 7); -lean_ctor_set(x_17, 1, x_407); +lean_ctor_set(x_17, 1, x_587); lean_ctor_set(x_17, 0, x_22); -x_408 = l_Lean_MessageData_ofExpr(x_374); +x_598 = l_Lean_MessageData_ofExpr(x_579); lean_ctor_set_tag(x_12, 7); -lean_ctor_set(x_12, 1, x_408); -lean_ctor_set(x_12, 0, x_398); -x_409 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_409, 0, x_12); -lean_ctor_set(x_409, 1, x_400); -x_410 = l_Lean_MessageData_ofExpr(x_392); -x_411 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_411, 0, x_409); -lean_ctor_set(x_411, 1, x_410); -x_412 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_412, 0, x_411); -lean_ctor_set(x_412, 1, x_398); -x_413 = l_Lean_indentD(x_412); -x_414 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_414, 0, x_17); -lean_ctor_set(x_414, 1, x_413); -x_415 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_415, 0, x_414); -lean_ctor_set(x_415, 1, x_398); -if (x_396 == 0) -{ -lean_object* x_416; lean_object* x_417; lean_object* x_418; uint8_t x_419; lean_object* x_420; lean_object* x_421; uint8_t x_422; lean_object* x_423; lean_object* x_424; -x_416 = l_Lean_Elab_Term_instInhabitedCalcStepView; -x_417 = l_outOfBounds___rarg(x_416); -x_418 = lean_ctor_get(x_417, 1); -lean_inc(x_418); -lean_dec(x_417); -x_419 = 2; +lean_ctor_set(x_12, 1, x_598); +lean_ctor_set(x_12, 0, x_17); +x_599 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_599, 0, x_12); +lean_ctor_set(x_599, 1, x_585); +x_600 = l_Lean_indentD(x_599); +x_601 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_601, 0, x_596); +lean_ctor_set(x_601, 1, x_600); +x_602 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_602, 0, x_601); +lean_ctor_set(x_602, 1, x_585); +if (x_583 == 0) +{ +lean_object* x_603; lean_object* x_604; lean_object* x_605; uint8_t x_606; lean_object* x_607; lean_object* x_608; uint8_t x_609; lean_object* x_610; lean_object* x_611; +x_603 = l_Lean_Elab_Term_instInhabitedCalcStepView; +x_604 = l_outOfBounds___rarg(x_603); +x_605 = lean_ctor_get(x_604, 1); +lean_inc(x_605); +lean_dec(x_604); +x_606 = 2; lean_inc(x_6); -x_420 = l_Lean_logAt___at_Lean_Elab_Term_reportUnsolvedGoals___spec__2(x_418, x_415, x_419, x_4, x_5, x_6, x_7, x_393); -lean_dec(x_418); -x_421 = lean_ctor_get(x_420, 1); -lean_inc(x_421); -lean_dec(x_420); -x_422 = 1; -x_423 = lean_box(0); -x_424 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__3(x_361, x_375, x_370, x_1, x_422, x_423, x_4, x_5, x_6, x_7, x_421); -return x_424; -} -else -{ -lean_object* x_425; lean_object* x_426; uint8_t x_427; lean_object* x_428; lean_object* x_429; uint8_t x_430; lean_object* x_431; lean_object* x_432; -x_425 = lean_array_fget(x_1, x_395); -x_426 = lean_ctor_get(x_425, 1); -lean_inc(x_426); -lean_dec(x_425); -x_427 = 2; +x_607 = l_Lean_logAt___at_Lean_Elab_Term_reportUnsolvedGoals___spec__2(x_605, x_602, x_606, x_4, x_5, x_6, x_7, x_577); +lean_dec(x_605); +x_608 = lean_ctor_get(x_607, 1); +lean_inc(x_608); +lean_dec(x_607); +x_609 = 1; +x_610 = lean_box(0); +x_611 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__3(x_536, x_550, x_545, x_1, x_609, x_610, x_4, x_5, x_6, x_7, x_608); +return x_611; +} +else +{ +lean_object* x_612; lean_object* x_613; uint8_t x_614; lean_object* x_615; lean_object* x_616; uint8_t x_617; lean_object* x_618; lean_object* x_619; +x_612 = lean_array_fget(x_1, x_582); +x_613 = lean_ctor_get(x_612, 1); +lean_inc(x_613); +lean_dec(x_612); +x_614 = 2; lean_inc(x_6); -x_428 = l_Lean_logAt___at_Lean_Elab_Term_reportUnsolvedGoals___spec__2(x_426, x_415, x_427, x_4, x_5, x_6, x_7, x_393); -lean_dec(x_426); -x_429 = lean_ctor_get(x_428, 1); -lean_inc(x_429); -lean_dec(x_428); -x_430 = 1; -x_431 = lean_box(0); -x_432 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__3(x_361, x_375, x_370, x_1, x_430, x_431, x_4, x_5, x_6, x_7, x_429); -return x_432; +x_615 = l_Lean_logAt___at_Lean_Elab_Term_reportUnsolvedGoals___spec__2(x_613, x_602, x_614, x_4, x_5, x_6, x_7, x_577); +lean_dec(x_613); +x_616 = lean_ctor_get(x_615, 1); +lean_inc(x_616); +lean_dec(x_615); +x_617 = 1; +x_618 = lean_box(0); +x_619 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__3(x_536, x_550, x_545, x_1, x_617, x_618, x_4, x_5, x_6, x_7, x_616); +return x_619; } } else { -lean_object* x_433; lean_object* x_434; lean_object* x_435; lean_object* x_436; -lean_dec(x_389); -lean_dec(x_376); -lean_dec(x_375); -lean_dec(x_374); -lean_dec(x_373); -lean_dec(x_370); -lean_dec(x_368); -lean_dec(x_361); -lean_dec(x_360); +lean_object* x_620; lean_object* x_621; lean_object* x_622; lean_object* x_623; +lean_dec(x_568); +lean_dec(x_567); +lean_dec(x_566); +lean_dec(x_551); +lean_dec(x_550); +lean_dec(x_548); +lean_dec(x_545); +lean_dec(x_543); +lean_dec(x_536); lean_free_object(x_22); lean_free_object(x_17); lean_free_object(x_12); @@ -21809,39 +23188,41 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_433 = lean_ctor_get(x_391, 0); -lean_inc(x_433); -x_434 = lean_ctor_get(x_391, 1); -lean_inc(x_434); -if (lean_is_exclusive(x_391)) { - lean_ctor_release(x_391, 0); - lean_ctor_release(x_391, 1); - x_435 = x_391; +x_620 = lean_ctor_get(x_575, 0); +lean_inc(x_620); +x_621 = lean_ctor_get(x_575, 1); +lean_inc(x_621); +if (lean_is_exclusive(x_575)) { + lean_ctor_release(x_575, 0); + lean_ctor_release(x_575, 1); + x_622 = x_575; } else { - lean_dec_ref(x_391); - x_435 = lean_box(0); + lean_dec_ref(x_575); + x_622 = lean_box(0); } -if (lean_is_scalar(x_435)) { - x_436 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_622)) { + x_623 = lean_alloc_ctor(1, 2, 0); } else { - x_436 = x_435; + x_623 = x_622; } -lean_ctor_set(x_436, 0, x_433); -lean_ctor_set(x_436, 1, x_434); -return x_436; +lean_ctor_set(x_623, 0, x_620); +lean_ctor_set(x_623, 1, x_621); +return x_623; } } else { -lean_object* x_437; lean_object* x_438; lean_object* x_439; lean_object* x_440; -lean_dec(x_376); -lean_dec(x_375); -lean_dec(x_374); -lean_dec(x_373); -lean_dec(x_370); -lean_dec(x_368); -lean_dec(x_361); -lean_dec(x_360); +lean_object* x_624; lean_object* x_625; lean_object* x_626; lean_object* x_627; +lean_dec(x_570); +lean_dec(x_568); +lean_dec(x_567); +lean_dec(x_566); +lean_dec(x_551); +lean_dec(x_550); +lean_dec(x_548); +lean_dec(x_545); +lean_dec(x_543); +lean_dec(x_536); lean_free_object(x_22); lean_free_object(x_17); lean_free_object(x_12); @@ -21849,59 +23230,138 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_437 = lean_ctor_get(x_388, 0); -lean_inc(x_437); -x_438 = lean_ctor_get(x_388, 1); -lean_inc(x_438); -if (lean_is_exclusive(x_388)) { - lean_ctor_release(x_388, 0); - lean_ctor_release(x_388, 1); - x_439 = x_388; +x_624 = lean_ctor_get(x_572, 0); +lean_inc(x_624); +x_625 = lean_ctor_get(x_572, 1); +lean_inc(x_625); +if (lean_is_exclusive(x_572)) { + lean_ctor_release(x_572, 0); + lean_ctor_release(x_572, 1); + x_626 = x_572; } else { - lean_dec_ref(x_388); - x_439 = lean_box(0); + lean_dec_ref(x_572); + x_626 = lean_box(0); +} +if (lean_is_scalar(x_626)) { + x_627 = lean_alloc_ctor(1, 2, 0); +} else { + x_627 = x_626; +} +lean_ctor_set(x_627, 0, x_624); +lean_ctor_set(x_627, 1, x_625); +return x_627; +} +} +else +{ +lean_object* x_628; lean_object* x_629; lean_object* x_630; lean_object* x_631; +lean_dec(x_568); +lean_dec(x_567); +lean_dec(x_566); +lean_dec(x_551); +lean_dec(x_550); +lean_dec(x_548); +lean_dec(x_545); +lean_dec(x_543); +lean_dec(x_536); +lean_free_object(x_22); +lean_free_object(x_17); +lean_free_object(x_12); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_628 = lean_ctor_get(x_569, 0); +lean_inc(x_628); +x_629 = lean_ctor_get(x_569, 1); +lean_inc(x_629); +if (lean_is_exclusive(x_569)) { + lean_ctor_release(x_569, 0); + lean_ctor_release(x_569, 1); + x_630 = x_569; +} else { + lean_dec_ref(x_569); + x_630 = lean_box(0); +} +if (lean_is_scalar(x_630)) { + x_631 = lean_alloc_ctor(1, 2, 0); +} else { + x_631 = x_630; +} +lean_ctor_set(x_631, 0, x_628); +lean_ctor_set(x_631, 1, x_629); +return x_631; +} +} +else +{ +lean_object* x_632; lean_object* x_633; lean_object* x_634; lean_object* x_635; +lean_dec(x_551); +lean_dec(x_550); +lean_dec(x_548); +lean_dec(x_545); +lean_dec(x_543); +lean_dec(x_536); +lean_free_object(x_22); +lean_free_object(x_17); +lean_free_object(x_12); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_632 = lean_ctor_get(x_563, 0); +lean_inc(x_632); +x_633 = lean_ctor_get(x_563, 1); +lean_inc(x_633); +if (lean_is_exclusive(x_563)) { + lean_ctor_release(x_563, 0); + lean_ctor_release(x_563, 1); + x_634 = x_563; +} else { + lean_dec_ref(x_563); + x_634 = lean_box(0); } -if (lean_is_scalar(x_439)) { - x_440 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_634)) { + x_635 = lean_alloc_ctor(1, 2, 0); } else { - x_440 = x_439; + x_635 = x_634; } -lean_ctor_set(x_440, 0, x_437); -lean_ctor_set(x_440, 1, x_438); -return x_440; +lean_ctor_set(x_635, 0, x_632); +lean_ctor_set(x_635, 1, x_633); +return x_635; } } else { -lean_object* x_441; uint8_t x_442; lean_object* x_443; lean_object* x_444; -lean_dec(x_376); -lean_dec(x_374); -lean_dec(x_373); -lean_dec(x_368); -lean_dec(x_360); +lean_object* x_636; uint8_t x_637; lean_object* x_638; lean_object* x_639; +lean_dec(x_551); +lean_dec(x_549); +lean_dec(x_548); +lean_dec(x_543); +lean_dec(x_535); lean_free_object(x_22); lean_free_object(x_17); lean_free_object(x_12); -x_441 = lean_ctor_get(x_384, 1); -lean_inc(x_441); -lean_dec(x_384); -x_442 = 0; -x_443 = lean_box(0); -x_444 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__3(x_361, x_375, x_370, x_1, x_442, x_443, x_4, x_5, x_6, x_7, x_441); -return x_444; +x_636 = lean_ctor_get(x_559, 1); +lean_inc(x_636); +lean_dec(x_559); +x_637 = 0; +x_638 = lean_box(0); +x_639 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__3(x_536, x_550, x_545, x_1, x_637, x_638, x_4, x_5, x_6, x_7, x_636); +return x_639; } } else { -lean_object* x_445; lean_object* x_446; lean_object* x_447; lean_object* x_448; -lean_dec(x_376); -lean_dec(x_375); -lean_dec(x_374); -lean_dec(x_373); -lean_dec(x_370); -lean_dec(x_368); -lean_dec(x_361); -lean_dec(x_360); +lean_object* x_640; lean_object* x_641; lean_object* x_642; lean_object* x_643; +lean_dec(x_551); +lean_dec(x_550); +lean_dec(x_549); +lean_dec(x_548); +lean_dec(x_545); +lean_dec(x_543); +lean_dec(x_536); +lean_dec(x_535); lean_free_object(x_22); lean_free_object(x_17); lean_free_object(x_12); @@ -21909,40 +23369,40 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_445 = lean_ctor_get(x_384, 0); -lean_inc(x_445); -x_446 = lean_ctor_get(x_384, 1); -lean_inc(x_446); -if (lean_is_exclusive(x_384)) { - lean_ctor_release(x_384, 0); - lean_ctor_release(x_384, 1); - x_447 = x_384; +x_640 = lean_ctor_get(x_559, 0); +lean_inc(x_640); +x_641 = lean_ctor_get(x_559, 1); +lean_inc(x_641); +if (lean_is_exclusive(x_559)) { + lean_ctor_release(x_559, 0); + lean_ctor_release(x_559, 1); + x_642 = x_559; } else { - lean_dec_ref(x_384); - x_447 = lean_box(0); + lean_dec_ref(x_559); + x_642 = lean_box(0); } -if (lean_is_scalar(x_447)) { - x_448 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_642)) { + x_643 = lean_alloc_ctor(1, 2, 0); } else { - x_448 = x_447; + x_643 = x_642; } -lean_ctor_set(x_448, 0, x_445); -lean_ctor_set(x_448, 1, x_446); -return x_448; +lean_ctor_set(x_643, 0, x_640); +lean_ctor_set(x_643, 1, x_641); +return x_643; } } } else { -lean_object* x_449; lean_object* x_450; lean_object* x_451; lean_object* x_452; -lean_dec(x_376); -lean_dec(x_375); -lean_dec(x_374); -lean_dec(x_373); -lean_dec(x_370); -lean_dec(x_368); -lean_dec(x_361); -lean_dec(x_360); +lean_object* x_644; lean_object* x_645; lean_object* x_646; lean_object* x_647; +lean_dec(x_551); +lean_dec(x_550); +lean_dec(x_549); +lean_dec(x_548); +lean_dec(x_545); +lean_dec(x_543); +lean_dec(x_536); +lean_dec(x_535); lean_free_object(x_22); lean_free_object(x_17); lean_dec(x_16); @@ -21953,468 +23413,609 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_449 = lean_ctor_get(x_377, 0); -lean_inc(x_449); -x_450 = lean_ctor_get(x_377, 1); -lean_inc(x_450); -if (lean_is_exclusive(x_377)) { - lean_ctor_release(x_377, 0); - lean_ctor_release(x_377, 1); - x_451 = x_377; +x_644 = lean_ctor_get(x_552, 0); +lean_inc(x_644); +x_645 = lean_ctor_get(x_552, 1); +lean_inc(x_645); +if (lean_is_exclusive(x_552)) { + lean_ctor_release(x_552, 0); + lean_ctor_release(x_552, 1); + x_646 = x_552; } else { - lean_dec_ref(x_377); - x_451 = lean_box(0); + lean_dec_ref(x_552); + x_646 = lean_box(0); } -if (lean_is_scalar(x_451)) { - x_452 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_646)) { + x_647 = lean_alloc_ctor(1, 2, 0); } else { - x_452 = x_451; + x_647 = x_646; } -lean_ctor_set(x_452, 0, x_449); -lean_ctor_set(x_452, 1, x_450); -return x_452; +lean_ctor_set(x_647, 0, x_644); +lean_ctor_set(x_647, 1, x_645); +return x_647; } } } } 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; -x_453 = lean_ctor_get(x_22, 0); -lean_inc(x_453); +lean_object* x_648; lean_object* x_649; lean_object* x_650; lean_object* x_651; lean_object* x_652; lean_object* x_653; +x_648 = lean_ctor_get(x_22, 0); +lean_inc(x_648); lean_dec(x_22); -x_454 = lean_ctor_get(x_23, 0); -lean_inc(x_454); -x_455 = lean_ctor_get(x_23, 1); -lean_inc(x_455); +x_649 = lean_ctor_get(x_23, 0); +lean_inc(x_649); +x_650 = lean_ctor_get(x_23, 1); +lean_inc(x_650); if (lean_is_exclusive(x_23)) { lean_ctor_release(x_23, 0); lean_ctor_release(x_23, 1); - x_456 = x_23; + x_651 = x_23; } else { lean_dec_ref(x_23); - x_456 = lean_box(0); + x_651 = lean_box(0); } -x_457 = l_Lean_Elab_Term_getCalcRelation_x3f(x_2, x_4, x_5, x_6, x_7, x_25); -x_458 = lean_ctor_get(x_457, 0); -lean_inc(x_458); -if (lean_obj_tag(x_458) == 0) +x_652 = l_Lean_Elab_Term_getCalcRelation_x3f(x_2, x_4, x_5, x_6, x_7, x_25); +x_653 = lean_ctor_get(x_652, 0); +lean_inc(x_653); +if (lean_obj_tag(x_653) == 0) { -lean_object* x_459; lean_object* x_460; lean_object* x_461; -lean_dec(x_456); -lean_dec(x_455); -lean_dec(x_454); -lean_dec(x_453); +lean_object* x_654; lean_object* x_655; lean_object* x_656; +lean_dec(x_651); +lean_dec(x_650); +lean_dec(x_649); +lean_dec(x_648); lean_free_object(x_17); lean_free_object(x_12); -x_459 = lean_ctor_get(x_457, 1); -lean_inc(x_459); -lean_dec(x_457); -x_460 = lean_box(0); -x_461 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__1(x_2, x_16, x_3, x_460, x_4, x_5, x_6, x_7, x_459); -return x_461; +x_654 = lean_ctor_get(x_652, 1); +lean_inc(x_654); +lean_dec(x_652); +x_655 = lean_box(0); +x_656 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__1(x_2, x_16, x_3, x_655, x_4, x_5, x_6, x_7, x_654); +return x_656; } else { -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; -x_462 = lean_ctor_get(x_457, 1); -lean_inc(x_462); -if (lean_is_exclusive(x_457)) { - lean_ctor_release(x_457, 0); - lean_ctor_release(x_457, 1); - x_463 = x_457; +lean_object* x_657; lean_object* x_658; lean_object* x_659; lean_object* x_660; lean_object* x_661; lean_object* x_662; lean_object* x_663; lean_object* x_664; lean_object* x_665; lean_object* x_666; lean_object* x_667; +x_657 = lean_ctor_get(x_652, 1); +lean_inc(x_657); +if (lean_is_exclusive(x_652)) { + lean_ctor_release(x_652, 0); + lean_ctor_release(x_652, 1); + x_658 = x_652; } else { - lean_dec_ref(x_457); - x_463 = lean_box(0); + lean_dec_ref(x_652); + x_658 = lean_box(0); } -x_464 = lean_ctor_get(x_458, 0); -lean_inc(x_464); -lean_dec(x_458); +x_659 = lean_ctor_get(x_653, 0); +lean_inc(x_659); +lean_dec(x_653); lean_inc(x_3); lean_inc(x_16); lean_inc(x_2); -x_465 = lean_alloc_closure((void*)(l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__1___boxed), 9, 3); -lean_closure_set(x_465, 0, x_2); -lean_closure_set(x_465, 1, x_16); -lean_closure_set(x_465, 2, x_3); -x_466 = lean_ctor_get(x_464, 1); -lean_inc(x_466); -x_467 = lean_ctor_get(x_464, 0); -lean_inc(x_467); -if (lean_is_exclusive(x_464)) { - lean_ctor_release(x_464, 0); - lean_ctor_release(x_464, 1); - x_468 = x_464; +x_660 = lean_alloc_closure((void*)(l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__1___boxed), 9, 3); +lean_closure_set(x_660, 0, x_2); +lean_closure_set(x_660, 1, x_16); +lean_closure_set(x_660, 2, x_3); +x_661 = lean_ctor_get(x_659, 1); +lean_inc(x_661); +x_662 = lean_ctor_get(x_659, 0); +lean_inc(x_662); +if (lean_is_exclusive(x_659)) { + lean_ctor_release(x_659, 0); + lean_ctor_release(x_659, 1); + x_663 = x_659; } else { - lean_dec_ref(x_464); - x_468 = lean_box(0); + lean_dec_ref(x_659); + x_663 = lean_box(0); } -x_469 = lean_ctor_get(x_466, 0); -lean_inc(x_469); -x_470 = lean_ctor_get(x_466, 1); -lean_inc(x_470); -if (lean_is_exclusive(x_466)) { - lean_ctor_release(x_466, 0); - lean_ctor_release(x_466, 1); - x_471 = x_466; +x_664 = lean_ctor_get(x_661, 0); +lean_inc(x_664); +x_665 = lean_ctor_get(x_661, 1); +lean_inc(x_665); +if (lean_is_exclusive(x_661)) { + lean_ctor_release(x_661, 0); + lean_ctor_release(x_661, 1); + x_666 = x_661; } else { - lean_dec_ref(x_466); - x_471 = lean_box(0); + lean_dec_ref(x_661); + x_666 = lean_box(0); } lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -x_472 = l_Lean_Meta_isExprDefEqGuarded(x_453, x_467, x_4, x_5, x_6, x_7, x_462); -if (lean_obj_tag(x_472) == 0) +x_667 = l_Lean_Meta_isExprDefEqGuarded(x_648, x_662, x_4, x_5, x_6, x_7, x_657); +if (lean_obj_tag(x_667) == 0) { -lean_object* x_473; uint8_t x_474; -x_473 = lean_ctor_get(x_472, 0); -lean_inc(x_473); -x_474 = lean_unbox(x_473); -lean_dec(x_473); -if (x_474 == 0) +lean_object* x_668; uint8_t x_669; +x_668 = lean_ctor_get(x_667, 0); +lean_inc(x_668); +x_669 = lean_unbox(x_668); +lean_dec(x_668); +if (x_669 == 0) { -lean_object* x_475; lean_object* x_476; lean_object* x_477; -lean_dec(x_471); -lean_dec(x_470); -lean_dec(x_469); -lean_dec(x_468); -lean_dec(x_465); -lean_dec(x_463); -lean_dec(x_456); -lean_dec(x_455); -lean_dec(x_454); +lean_object* x_670; lean_object* x_671; lean_object* x_672; +lean_dec(x_666); +lean_dec(x_665); +lean_dec(x_664); +lean_dec(x_663); +lean_dec(x_660); +lean_dec(x_658); +lean_dec(x_651); +lean_dec(x_650); +lean_dec(x_649); lean_free_object(x_17); lean_free_object(x_12); -x_475 = lean_ctor_get(x_472, 1); -lean_inc(x_475); -lean_dec(x_472); -x_476 = lean_box(0); -x_477 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__1(x_2, x_16, x_3, x_476, x_4, x_5, x_6, x_7, x_475); -return x_477; +x_670 = lean_ctor_get(x_667, 1); +lean_inc(x_670); +lean_dec(x_667); +x_671 = lean_box(0); +x_672 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__1(x_2, x_16, x_3, x_671, x_4, x_5, x_6, x_7, x_670); +return x_672; } else { -lean_object* x_478; lean_object* x_479; +lean_object* x_673; lean_object* x_674; lean_dec(x_16); lean_dec(x_3); lean_dec(x_2); -x_478 = lean_ctor_get(x_472, 1); -lean_inc(x_478); -lean_dec(x_472); +x_673 = lean_ctor_get(x_667, 1); +lean_inc(x_673); +lean_dec(x_667); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -lean_inc(x_469); -lean_inc(x_454); -x_479 = l_Lean_Meta_isExprDefEqGuarded(x_454, x_469, x_4, x_5, x_6, x_7, x_478); -if (lean_obj_tag(x_479) == 0) +lean_inc(x_664); +lean_inc(x_649); +x_674 = l_Lean_Meta_isExprDefEqGuarded(x_649, x_664, x_4, x_5, x_6, x_7, x_673); +if (lean_obj_tag(x_674) == 0) { -lean_object* x_480; uint8_t x_481; -x_480 = lean_ctor_get(x_479, 0); -lean_inc(x_480); -x_481 = lean_unbox(x_480); -lean_dec(x_480); -if (x_481 == 0) +lean_object* x_675; uint8_t x_676; +x_675 = lean_ctor_get(x_674, 0); +lean_inc(x_675); +x_676 = lean_unbox(x_675); +lean_dec(x_675); +if (x_676 == 0) { -lean_object* x_482; lean_object* x_483; -x_482 = lean_ctor_get(x_479, 1); -lean_inc(x_482); -lean_dec(x_479); +lean_object* x_677; lean_object* x_678; +x_677 = lean_ctor_get(x_674, 1); +lean_inc(x_677); +lean_dec(x_674); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -lean_inc(x_454); -x_483 = lean_infer_type(x_454, x_4, x_5, x_6, x_7, x_482); -if (lean_obj_tag(x_483) == 0) +x_678 = l_Lean_Meta_addPPExplicitToExposeDiff(x_649, x_664, x_4, x_5, x_6, x_7, x_677); +if (lean_obj_tag(x_678) == 0) { -lean_object* x_484; lean_object* x_485; lean_object* x_486; -x_484 = lean_ctor_get(x_483, 0); -lean_inc(x_484); -x_485 = lean_ctor_get(x_483, 1); -lean_inc(x_485); -lean_dec(x_483); +lean_object* x_679; lean_object* x_680; lean_object* x_681; lean_object* x_682; lean_object* x_683; lean_object* x_684; +x_679 = lean_ctor_get(x_678, 0); +lean_inc(x_679); +x_680 = lean_ctor_get(x_678, 1); +lean_inc(x_680); +lean_dec(x_678); +x_681 = lean_ctor_get(x_679, 0); +lean_inc(x_681); +x_682 = lean_ctor_get(x_679, 1); +lean_inc(x_682); +if (lean_is_exclusive(x_679)) { + lean_ctor_release(x_679, 0); + lean_ctor_release(x_679, 1); + x_683 = x_679; +} else { + lean_dec_ref(x_679); + x_683 = lean_box(0); +} lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -lean_inc(x_469); -x_486 = lean_infer_type(x_469, x_4, x_5, x_6, x_7, x_485); -if (lean_obj_tag(x_486) == 0) +lean_inc(x_681); +x_684 = lean_infer_type(x_681, x_4, x_5, x_6, x_7, x_680); +if (lean_obj_tag(x_684) == 0) { -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_object* x_500; lean_object* x_501; lean_object* x_502; lean_object* x_503; lean_object* x_504; lean_object* x_505; lean_object* x_506; lean_object* x_507; lean_object* x_508; lean_object* x_509; lean_object* x_510; lean_object* x_511; -x_487 = lean_ctor_get(x_486, 0); -lean_inc(x_487); -x_488 = lean_ctor_get(x_486, 1); -lean_inc(x_488); -lean_dec(x_486); -x_489 = lean_array_get_size(x_1); -x_490 = lean_unsigned_to_nat(0u); -x_491 = lean_nat_dec_lt(x_490, x_489); -lean_dec(x_489); -x_492 = l_Lean_MessageData_ofExpr(x_454); -x_493 = l___private_Lean_Elab_Calc_0__Lean_Elab_Term_getRelUniv___lambda__1___closed__4; -if (lean_is_scalar(x_471)) { - x_494 = lean_alloc_ctor(7, 2, 0); +lean_object* x_685; lean_object* x_686; lean_object* x_687; +x_685 = lean_ctor_get(x_684, 0); +lean_inc(x_685); +x_686 = lean_ctor_get(x_684, 1); +lean_inc(x_686); +lean_dec(x_684); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_682); +x_687 = lean_infer_type(x_682, x_4, x_5, x_6, x_7, x_686); +if (lean_obj_tag(x_687) == 0) +{ +lean_object* x_688; lean_object* x_689; lean_object* x_690; +x_688 = lean_ctor_get(x_687, 0); +lean_inc(x_688); +x_689 = lean_ctor_get(x_687, 1); +lean_inc(x_689); +lean_dec(x_687); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +x_690 = l_Lean_Meta_addPPExplicitToExposeDiff(x_685, x_688, x_4, x_5, x_6, x_7, x_689); +if (lean_obj_tag(x_690) == 0) +{ +lean_object* x_691; lean_object* x_692; lean_object* x_693; lean_object* x_694; lean_object* x_695; lean_object* x_696; lean_object* x_697; uint8_t x_698; lean_object* x_699; lean_object* x_700; lean_object* x_701; lean_object* x_702; lean_object* x_703; lean_object* x_704; lean_object* x_705; lean_object* x_706; lean_object* x_707; lean_object* x_708; lean_object* x_709; lean_object* x_710; lean_object* x_711; lean_object* x_712; lean_object* x_713; lean_object* x_714; lean_object* x_715; lean_object* x_716; lean_object* x_717; lean_object* x_718; +x_691 = lean_ctor_get(x_690, 0); +lean_inc(x_691); +x_692 = lean_ctor_get(x_690, 1); +lean_inc(x_692); +lean_dec(x_690); +x_693 = lean_ctor_get(x_691, 0); +lean_inc(x_693); +x_694 = lean_ctor_get(x_691, 1); +lean_inc(x_694); +if (lean_is_exclusive(x_691)) { + lean_ctor_release(x_691, 0); + lean_ctor_release(x_691, 1); + x_695 = x_691; } else { - x_494 = x_471; - lean_ctor_set_tag(x_494, 7); + lean_dec_ref(x_691); + x_695 = lean_box(0); } -lean_ctor_set(x_494, 0, x_493); -lean_ctor_set(x_494, 1, x_492); -x_495 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Term_elabCalcSteps___spec__1___closed__6; -if (lean_is_scalar(x_468)) { - x_496 = lean_alloc_ctor(7, 2, 0); +x_696 = lean_array_get_size(x_1); +x_697 = lean_unsigned_to_nat(0u); +x_698 = lean_nat_dec_lt(x_697, x_696); +lean_dec(x_696); +x_699 = l_Lean_MessageData_ofExpr(x_681); +x_700 = l___private_Lean_Elab_Calc_0__Lean_Elab_Term_getRelUniv___lambda__1___closed__4; +if (lean_is_scalar(x_695)) { + x_701 = lean_alloc_ctor(7, 2, 0); } else { - x_496 = x_468; - lean_ctor_set_tag(x_496, 7); + x_701 = x_695; + lean_ctor_set_tag(x_701, 7); } -lean_ctor_set(x_496, 0, x_494); -lean_ctor_set(x_496, 1, x_495); -x_497 = l_Lean_MessageData_ofExpr(x_484); -if (lean_is_scalar(x_463)) { - x_498 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_701, 0, x_700); +lean_ctor_set(x_701, 1, x_699); +x_702 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Term_elabCalcSteps___spec__1___closed__6; +if (lean_is_scalar(x_683)) { + x_703 = lean_alloc_ctor(7, 2, 0); } else { - x_498 = x_463; - lean_ctor_set_tag(x_498, 7); + x_703 = x_683; + lean_ctor_set_tag(x_703, 7); } -lean_ctor_set(x_498, 0, x_496); -lean_ctor_set(x_498, 1, x_497); -if (lean_is_scalar(x_456)) { - x_499 = lean_alloc_ctor(7, 2, 0); -} else { - x_499 = x_456; - lean_ctor_set_tag(x_499, 7); -} -lean_ctor_set(x_499, 0, x_498); -lean_ctor_set(x_499, 1, x_493); -x_500 = l_Lean_indentD(x_499); -x_501 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Term_elabCalcSteps___spec__1___closed__4; -x_502 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_502, 0, x_501); -lean_ctor_set(x_502, 1, x_500); -x_503 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__3___closed__4; +lean_ctor_set(x_703, 0, x_701); +lean_ctor_set(x_703, 1, x_702); +x_704 = l_Lean_MessageData_ofExpr(x_693); +if (lean_is_scalar(x_666)) { + x_705 = lean_alloc_ctor(7, 2, 0); +} else { + x_705 = x_666; + lean_ctor_set_tag(x_705, 7); +} +lean_ctor_set(x_705, 0, x_703); +lean_ctor_set(x_705, 1, x_704); +if (lean_is_scalar(x_663)) { + x_706 = lean_alloc_ctor(7, 2, 0); +} else { + x_706 = x_663; + lean_ctor_set_tag(x_706, 7); +} +lean_ctor_set(x_706, 0, x_705); +lean_ctor_set(x_706, 1, x_700); +x_707 = l_Lean_indentD(x_706); +x_708 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Term_elabCalcSteps___spec__1___closed__4; +if (lean_is_scalar(x_658)) { + x_709 = lean_alloc_ctor(7, 2, 0); +} else { + x_709 = x_658; + lean_ctor_set_tag(x_709, 7); +} +lean_ctor_set(x_709, 0, x_708); +lean_ctor_set(x_709, 1, x_707); +x_710 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__3___closed__4; +if (lean_is_scalar(x_651)) { + x_711 = lean_alloc_ctor(7, 2, 0); +} else { + x_711 = x_651; + lean_ctor_set_tag(x_711, 7); +} +lean_ctor_set(x_711, 0, x_709); +lean_ctor_set(x_711, 1, x_710); +x_712 = l_Lean_MessageData_ofExpr(x_682); +x_713 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_713, 0, x_700); +lean_ctor_set(x_713, 1, x_712); lean_ctor_set_tag(x_17, 7); -lean_ctor_set(x_17, 1, x_503); -lean_ctor_set(x_17, 0, x_502); -x_504 = l_Lean_MessageData_ofExpr(x_469); +lean_ctor_set(x_17, 1, x_702); +lean_ctor_set(x_17, 0, x_713); +x_714 = l_Lean_MessageData_ofExpr(x_694); lean_ctor_set_tag(x_12, 7); -lean_ctor_set(x_12, 1, x_504); -lean_ctor_set(x_12, 0, x_493); -x_505 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_505, 0, x_12); -lean_ctor_set(x_505, 1, x_495); -x_506 = l_Lean_MessageData_ofExpr(x_487); -x_507 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_507, 0, x_505); -lean_ctor_set(x_507, 1, x_506); -x_508 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_508, 0, x_507); -lean_ctor_set(x_508, 1, x_493); -x_509 = l_Lean_indentD(x_508); -x_510 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_510, 0, x_17); -lean_ctor_set(x_510, 1, x_509); -x_511 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_511, 0, x_510); -lean_ctor_set(x_511, 1, x_493); -if (x_491 == 0) +lean_ctor_set(x_12, 1, x_714); +lean_ctor_set(x_12, 0, x_17); +x_715 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_715, 0, x_12); +lean_ctor_set(x_715, 1, x_700); +x_716 = l_Lean_indentD(x_715); +x_717 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_717, 0, x_711); +lean_ctor_set(x_717, 1, x_716); +x_718 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_718, 0, x_717); +lean_ctor_set(x_718, 1, x_700); +if (x_698 == 0) { -lean_object* x_512; lean_object* x_513; lean_object* x_514; uint8_t x_515; lean_object* x_516; lean_object* x_517; uint8_t x_518; lean_object* x_519; lean_object* x_520; -x_512 = l_Lean_Elab_Term_instInhabitedCalcStepView; -x_513 = l_outOfBounds___rarg(x_512); -x_514 = lean_ctor_get(x_513, 1); -lean_inc(x_514); -lean_dec(x_513); -x_515 = 2; +lean_object* x_719; lean_object* x_720; lean_object* x_721; uint8_t x_722; lean_object* x_723; lean_object* x_724; uint8_t x_725; lean_object* x_726; lean_object* x_727; +x_719 = l_Lean_Elab_Term_instInhabitedCalcStepView; +x_720 = l_outOfBounds___rarg(x_719); +x_721 = lean_ctor_get(x_720, 1); +lean_inc(x_721); +lean_dec(x_720); +x_722 = 2; lean_inc(x_6); -x_516 = l_Lean_logAt___at_Lean_Elab_Term_reportUnsolvedGoals___spec__2(x_514, x_511, x_515, x_4, x_5, x_6, x_7, x_488); -lean_dec(x_514); -x_517 = lean_ctor_get(x_516, 1); -lean_inc(x_517); -lean_dec(x_516); -x_518 = 1; -x_519 = lean_box(0); -x_520 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__3(x_455, x_470, x_465, x_1, x_518, x_519, x_4, x_5, x_6, x_7, x_517); -return x_520; -} -else -{ -lean_object* x_521; lean_object* x_522; uint8_t x_523; lean_object* x_524; lean_object* x_525; uint8_t x_526; lean_object* x_527; lean_object* x_528; -x_521 = lean_array_fget(x_1, x_490); -x_522 = lean_ctor_get(x_521, 1); -lean_inc(x_522); -lean_dec(x_521); -x_523 = 2; +x_723 = l_Lean_logAt___at_Lean_Elab_Term_reportUnsolvedGoals___spec__2(x_721, x_718, x_722, x_4, x_5, x_6, x_7, x_692); +lean_dec(x_721); +x_724 = lean_ctor_get(x_723, 1); +lean_inc(x_724); +lean_dec(x_723); +x_725 = 1; +x_726 = lean_box(0); +x_727 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__3(x_650, x_665, x_660, x_1, x_725, x_726, x_4, x_5, x_6, x_7, x_724); +return x_727; +} +else +{ +lean_object* x_728; lean_object* x_729; uint8_t x_730; lean_object* x_731; lean_object* x_732; uint8_t x_733; lean_object* x_734; lean_object* x_735; +x_728 = lean_array_fget(x_1, x_697); +x_729 = lean_ctor_get(x_728, 1); +lean_inc(x_729); +lean_dec(x_728); +x_730 = 2; lean_inc(x_6); -x_524 = l_Lean_logAt___at_Lean_Elab_Term_reportUnsolvedGoals___spec__2(x_522, x_511, x_523, x_4, x_5, x_6, x_7, x_488); -lean_dec(x_522); -x_525 = lean_ctor_get(x_524, 1); -lean_inc(x_525); -lean_dec(x_524); -x_526 = 1; -x_527 = lean_box(0); -x_528 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__3(x_455, x_470, x_465, x_1, x_526, x_527, x_4, x_5, x_6, x_7, x_525); -return x_528; +x_731 = l_Lean_logAt___at_Lean_Elab_Term_reportUnsolvedGoals___spec__2(x_729, x_718, x_730, x_4, x_5, x_6, x_7, x_692); +lean_dec(x_729); +x_732 = lean_ctor_get(x_731, 1); +lean_inc(x_732); +lean_dec(x_731); +x_733 = 1; +x_734 = lean_box(0); +x_735 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__3(x_650, x_665, x_660, x_1, x_733, x_734, x_4, x_5, x_6, x_7, x_732); +return x_735; +} +} +else +{ +lean_object* x_736; lean_object* x_737; lean_object* x_738; lean_object* x_739; +lean_dec(x_683); +lean_dec(x_682); +lean_dec(x_681); +lean_dec(x_666); +lean_dec(x_665); +lean_dec(x_663); +lean_dec(x_660); +lean_dec(x_658); +lean_dec(x_651); +lean_dec(x_650); +lean_free_object(x_17); +lean_free_object(x_12); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_736 = lean_ctor_get(x_690, 0); +lean_inc(x_736); +x_737 = lean_ctor_get(x_690, 1); +lean_inc(x_737); +if (lean_is_exclusive(x_690)) { + lean_ctor_release(x_690, 0); + lean_ctor_release(x_690, 1); + x_738 = x_690; +} else { + lean_dec_ref(x_690); + x_738 = lean_box(0); +} +if (lean_is_scalar(x_738)) { + x_739 = lean_alloc_ctor(1, 2, 0); +} else { + x_739 = x_738; +} +lean_ctor_set(x_739, 0, x_736); +lean_ctor_set(x_739, 1, x_737); +return x_739; +} +} +else +{ +lean_object* x_740; lean_object* x_741; lean_object* x_742; lean_object* x_743; +lean_dec(x_685); +lean_dec(x_683); +lean_dec(x_682); +lean_dec(x_681); +lean_dec(x_666); +lean_dec(x_665); +lean_dec(x_663); +lean_dec(x_660); +lean_dec(x_658); +lean_dec(x_651); +lean_dec(x_650); +lean_free_object(x_17); +lean_free_object(x_12); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_740 = lean_ctor_get(x_687, 0); +lean_inc(x_740); +x_741 = lean_ctor_get(x_687, 1); +lean_inc(x_741); +if (lean_is_exclusive(x_687)) { + lean_ctor_release(x_687, 0); + lean_ctor_release(x_687, 1); + x_742 = x_687; +} else { + lean_dec_ref(x_687); + x_742 = lean_box(0); +} +if (lean_is_scalar(x_742)) { + x_743 = lean_alloc_ctor(1, 2, 0); +} else { + x_743 = x_742; +} +lean_ctor_set(x_743, 0, x_740); +lean_ctor_set(x_743, 1, x_741); +return x_743; } } else { -lean_object* x_529; lean_object* x_530; lean_object* x_531; lean_object* x_532; -lean_dec(x_484); -lean_dec(x_471); -lean_dec(x_470); -lean_dec(x_469); -lean_dec(x_468); -lean_dec(x_465); -lean_dec(x_463); -lean_dec(x_456); -lean_dec(x_455); -lean_dec(x_454); +lean_object* x_744; lean_object* x_745; lean_object* x_746; lean_object* x_747; +lean_dec(x_683); +lean_dec(x_682); +lean_dec(x_681); +lean_dec(x_666); +lean_dec(x_665); +lean_dec(x_663); +lean_dec(x_660); +lean_dec(x_658); +lean_dec(x_651); +lean_dec(x_650); lean_free_object(x_17); lean_free_object(x_12); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_529 = lean_ctor_get(x_486, 0); -lean_inc(x_529); -x_530 = lean_ctor_get(x_486, 1); -lean_inc(x_530); -if (lean_is_exclusive(x_486)) { - lean_ctor_release(x_486, 0); - lean_ctor_release(x_486, 1); - x_531 = x_486; +x_744 = lean_ctor_get(x_684, 0); +lean_inc(x_744); +x_745 = lean_ctor_get(x_684, 1); +lean_inc(x_745); +if (lean_is_exclusive(x_684)) { + lean_ctor_release(x_684, 0); + lean_ctor_release(x_684, 1); + x_746 = x_684; } else { - lean_dec_ref(x_486); - x_531 = lean_box(0); + lean_dec_ref(x_684); + x_746 = lean_box(0); } -if (lean_is_scalar(x_531)) { - x_532 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_746)) { + x_747 = lean_alloc_ctor(1, 2, 0); } else { - x_532 = x_531; + x_747 = x_746; } -lean_ctor_set(x_532, 0, x_529); -lean_ctor_set(x_532, 1, x_530); -return x_532; +lean_ctor_set(x_747, 0, x_744); +lean_ctor_set(x_747, 1, x_745); +return x_747; } } else { -lean_object* x_533; lean_object* x_534; lean_object* x_535; lean_object* x_536; -lean_dec(x_471); -lean_dec(x_470); -lean_dec(x_469); -lean_dec(x_468); -lean_dec(x_465); -lean_dec(x_463); -lean_dec(x_456); -lean_dec(x_455); -lean_dec(x_454); +lean_object* x_748; lean_object* x_749; lean_object* x_750; lean_object* x_751; +lean_dec(x_666); +lean_dec(x_665); +lean_dec(x_663); +lean_dec(x_660); +lean_dec(x_658); +lean_dec(x_651); +lean_dec(x_650); lean_free_object(x_17); lean_free_object(x_12); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_533 = lean_ctor_get(x_483, 0); -lean_inc(x_533); -x_534 = lean_ctor_get(x_483, 1); -lean_inc(x_534); -if (lean_is_exclusive(x_483)) { - lean_ctor_release(x_483, 0); - lean_ctor_release(x_483, 1); - x_535 = x_483; -} else { - lean_dec_ref(x_483); - x_535 = lean_box(0); -} -if (lean_is_scalar(x_535)) { - x_536 = lean_alloc_ctor(1, 2, 0); +x_748 = lean_ctor_get(x_678, 0); +lean_inc(x_748); +x_749 = lean_ctor_get(x_678, 1); +lean_inc(x_749); +if (lean_is_exclusive(x_678)) { + lean_ctor_release(x_678, 0); + lean_ctor_release(x_678, 1); + x_750 = x_678; +} else { + lean_dec_ref(x_678); + x_750 = lean_box(0); +} +if (lean_is_scalar(x_750)) { + x_751 = lean_alloc_ctor(1, 2, 0); } else { - x_536 = x_535; + x_751 = x_750; } -lean_ctor_set(x_536, 0, x_533); -lean_ctor_set(x_536, 1, x_534); -return x_536; +lean_ctor_set(x_751, 0, x_748); +lean_ctor_set(x_751, 1, x_749); +return x_751; } } else { -lean_object* x_537; uint8_t x_538; lean_object* x_539; lean_object* x_540; -lean_dec(x_471); -lean_dec(x_469); -lean_dec(x_468); -lean_dec(x_463); -lean_dec(x_456); -lean_dec(x_454); +lean_object* x_752; uint8_t x_753; lean_object* x_754; lean_object* x_755; +lean_dec(x_666); +lean_dec(x_664); +lean_dec(x_663); +lean_dec(x_658); +lean_dec(x_651); +lean_dec(x_649); lean_free_object(x_17); lean_free_object(x_12); -x_537 = lean_ctor_get(x_479, 1); -lean_inc(x_537); -lean_dec(x_479); -x_538 = 0; -x_539 = lean_box(0); -x_540 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__3(x_455, x_470, x_465, x_1, x_538, x_539, x_4, x_5, x_6, x_7, x_537); -return x_540; +x_752 = lean_ctor_get(x_674, 1); +lean_inc(x_752); +lean_dec(x_674); +x_753 = 0; +x_754 = lean_box(0); +x_755 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__3(x_650, x_665, x_660, x_1, x_753, x_754, x_4, x_5, x_6, x_7, x_752); +return x_755; } } else { -lean_object* x_541; lean_object* x_542; lean_object* x_543; lean_object* x_544; -lean_dec(x_471); -lean_dec(x_470); -lean_dec(x_469); -lean_dec(x_468); -lean_dec(x_465); -lean_dec(x_463); -lean_dec(x_456); -lean_dec(x_455); -lean_dec(x_454); +lean_object* x_756; lean_object* x_757; lean_object* x_758; lean_object* x_759; +lean_dec(x_666); +lean_dec(x_665); +lean_dec(x_664); +lean_dec(x_663); +lean_dec(x_660); +lean_dec(x_658); +lean_dec(x_651); +lean_dec(x_650); +lean_dec(x_649); lean_free_object(x_17); lean_free_object(x_12); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_541 = lean_ctor_get(x_479, 0); -lean_inc(x_541); -x_542 = lean_ctor_get(x_479, 1); -lean_inc(x_542); -if (lean_is_exclusive(x_479)) { - lean_ctor_release(x_479, 0); - lean_ctor_release(x_479, 1); - x_543 = x_479; +x_756 = lean_ctor_get(x_674, 0); +lean_inc(x_756); +x_757 = lean_ctor_get(x_674, 1); +lean_inc(x_757); +if (lean_is_exclusive(x_674)) { + lean_ctor_release(x_674, 0); + lean_ctor_release(x_674, 1); + x_758 = x_674; } else { - lean_dec_ref(x_479); - x_543 = lean_box(0); + lean_dec_ref(x_674); + x_758 = lean_box(0); } -if (lean_is_scalar(x_543)) { - x_544 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_758)) { + x_759 = lean_alloc_ctor(1, 2, 0); } else { - x_544 = x_543; + x_759 = x_758; } -lean_ctor_set(x_544, 0, x_541); -lean_ctor_set(x_544, 1, x_542); -return x_544; +lean_ctor_set(x_759, 0, x_756); +lean_ctor_set(x_759, 1, x_757); +return x_759; } } } else { -lean_object* x_545; lean_object* x_546; lean_object* x_547; lean_object* x_548; -lean_dec(x_471); -lean_dec(x_470); -lean_dec(x_469); -lean_dec(x_468); -lean_dec(x_465); -lean_dec(x_463); -lean_dec(x_456); -lean_dec(x_455); -lean_dec(x_454); +lean_object* x_760; lean_object* x_761; lean_object* x_762; lean_object* x_763; +lean_dec(x_666); +lean_dec(x_665); +lean_dec(x_664); +lean_dec(x_663); +lean_dec(x_660); +lean_dec(x_658); +lean_dec(x_651); +lean_dec(x_650); +lean_dec(x_649); lean_free_object(x_17); lean_dec(x_16); lean_free_object(x_12); @@ -22424,484 +24025,625 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_545 = lean_ctor_get(x_472, 0); -lean_inc(x_545); -x_546 = lean_ctor_get(x_472, 1); -lean_inc(x_546); -if (lean_is_exclusive(x_472)) { - lean_ctor_release(x_472, 0); - lean_ctor_release(x_472, 1); - x_547 = x_472; +x_760 = lean_ctor_get(x_667, 0); +lean_inc(x_760); +x_761 = lean_ctor_get(x_667, 1); +lean_inc(x_761); +if (lean_is_exclusive(x_667)) { + lean_ctor_release(x_667, 0); + lean_ctor_release(x_667, 1); + x_762 = x_667; } else { - lean_dec_ref(x_472); - x_547 = lean_box(0); + lean_dec_ref(x_667); + x_762 = lean_box(0); } -if (lean_is_scalar(x_547)) { - x_548 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_762)) { + x_763 = lean_alloc_ctor(1, 2, 0); } else { - x_548 = x_547; + x_763 = x_762; } -lean_ctor_set(x_548, 0, x_545); -lean_ctor_set(x_548, 1, x_546); -return x_548; +lean_ctor_set(x_763, 0, x_760); +lean_ctor_set(x_763, 1, x_761); +return x_763; } } } } else { -lean_object* x_549; lean_object* x_550; lean_object* x_551; lean_object* x_552; lean_object* x_553; lean_object* x_554; lean_object* x_555; lean_object* x_556; -x_549 = lean_ctor_get(x_17, 1); -lean_inc(x_549); +lean_object* x_764; lean_object* x_765; lean_object* x_766; lean_object* x_767; lean_object* x_768; lean_object* x_769; lean_object* x_770; lean_object* x_771; +x_764 = lean_ctor_get(x_17, 1); +lean_inc(x_764); lean_dec(x_17); -x_550 = lean_ctor_get(x_22, 0); -lean_inc(x_550); +x_765 = lean_ctor_get(x_22, 0); +lean_inc(x_765); if (lean_is_exclusive(x_22)) { lean_ctor_release(x_22, 0); lean_ctor_release(x_22, 1); - x_551 = x_22; + x_766 = x_22; } else { lean_dec_ref(x_22); - x_551 = lean_box(0); + x_766 = lean_box(0); } -x_552 = lean_ctor_get(x_23, 0); -lean_inc(x_552); -x_553 = lean_ctor_get(x_23, 1); -lean_inc(x_553); +x_767 = lean_ctor_get(x_23, 0); +lean_inc(x_767); +x_768 = lean_ctor_get(x_23, 1); +lean_inc(x_768); if (lean_is_exclusive(x_23)) { lean_ctor_release(x_23, 0); lean_ctor_release(x_23, 1); - x_554 = x_23; + x_769 = x_23; } else { lean_dec_ref(x_23); - x_554 = lean_box(0); -} -x_555 = l_Lean_Elab_Term_getCalcRelation_x3f(x_2, x_4, x_5, x_6, x_7, x_549); -x_556 = lean_ctor_get(x_555, 0); -lean_inc(x_556); -if (lean_obj_tag(x_556) == 0) -{ -lean_object* x_557; lean_object* x_558; lean_object* x_559; -lean_dec(x_554); -lean_dec(x_553); -lean_dec(x_552); -lean_dec(x_551); -lean_dec(x_550); + x_769 = lean_box(0); +} +x_770 = l_Lean_Elab_Term_getCalcRelation_x3f(x_2, x_4, x_5, x_6, x_7, x_764); +x_771 = lean_ctor_get(x_770, 0); +lean_inc(x_771); +if (lean_obj_tag(x_771) == 0) +{ +lean_object* x_772; lean_object* x_773; lean_object* x_774; +lean_dec(x_769); +lean_dec(x_768); +lean_dec(x_767); +lean_dec(x_766); +lean_dec(x_765); lean_free_object(x_12); -x_557 = lean_ctor_get(x_555, 1); -lean_inc(x_557); -lean_dec(x_555); -x_558 = lean_box(0); -x_559 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__1(x_2, x_16, x_3, x_558, x_4, x_5, x_6, x_7, x_557); -return x_559; +x_772 = lean_ctor_get(x_770, 1); +lean_inc(x_772); +lean_dec(x_770); +x_773 = lean_box(0); +x_774 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__1(x_2, x_16, x_3, x_773, x_4, x_5, x_6, x_7, x_772); +return x_774; } else { -lean_object* x_560; lean_object* x_561; lean_object* x_562; lean_object* x_563; lean_object* x_564; lean_object* x_565; lean_object* x_566; lean_object* x_567; lean_object* x_568; lean_object* x_569; lean_object* x_570; -x_560 = lean_ctor_get(x_555, 1); -lean_inc(x_560); -if (lean_is_exclusive(x_555)) { - lean_ctor_release(x_555, 0); - lean_ctor_release(x_555, 1); - x_561 = x_555; +lean_object* x_775; lean_object* x_776; lean_object* x_777; lean_object* x_778; lean_object* x_779; lean_object* x_780; lean_object* x_781; lean_object* x_782; lean_object* x_783; lean_object* x_784; lean_object* x_785; +x_775 = lean_ctor_get(x_770, 1); +lean_inc(x_775); +if (lean_is_exclusive(x_770)) { + lean_ctor_release(x_770, 0); + lean_ctor_release(x_770, 1); + x_776 = x_770; } else { - lean_dec_ref(x_555); - x_561 = lean_box(0); + lean_dec_ref(x_770); + x_776 = lean_box(0); } -x_562 = lean_ctor_get(x_556, 0); -lean_inc(x_562); -lean_dec(x_556); +x_777 = lean_ctor_get(x_771, 0); +lean_inc(x_777); +lean_dec(x_771); lean_inc(x_3); lean_inc(x_16); lean_inc(x_2); -x_563 = lean_alloc_closure((void*)(l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__1___boxed), 9, 3); -lean_closure_set(x_563, 0, x_2); -lean_closure_set(x_563, 1, x_16); -lean_closure_set(x_563, 2, x_3); -x_564 = lean_ctor_get(x_562, 1); -lean_inc(x_564); -x_565 = lean_ctor_get(x_562, 0); -lean_inc(x_565); -if (lean_is_exclusive(x_562)) { - lean_ctor_release(x_562, 0); - lean_ctor_release(x_562, 1); - x_566 = x_562; +x_778 = lean_alloc_closure((void*)(l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__1___boxed), 9, 3); +lean_closure_set(x_778, 0, x_2); +lean_closure_set(x_778, 1, x_16); +lean_closure_set(x_778, 2, x_3); +x_779 = lean_ctor_get(x_777, 1); +lean_inc(x_779); +x_780 = lean_ctor_get(x_777, 0); +lean_inc(x_780); +if (lean_is_exclusive(x_777)) { + lean_ctor_release(x_777, 0); + lean_ctor_release(x_777, 1); + x_781 = x_777; } else { - lean_dec_ref(x_562); - x_566 = lean_box(0); + lean_dec_ref(x_777); + x_781 = lean_box(0); } -x_567 = lean_ctor_get(x_564, 0); -lean_inc(x_567); -x_568 = lean_ctor_get(x_564, 1); -lean_inc(x_568); -if (lean_is_exclusive(x_564)) { - lean_ctor_release(x_564, 0); - lean_ctor_release(x_564, 1); - x_569 = x_564; +x_782 = lean_ctor_get(x_779, 0); +lean_inc(x_782); +x_783 = lean_ctor_get(x_779, 1); +lean_inc(x_783); +if (lean_is_exclusive(x_779)) { + lean_ctor_release(x_779, 0); + lean_ctor_release(x_779, 1); + x_784 = x_779; } else { - lean_dec_ref(x_564); - x_569 = lean_box(0); + lean_dec_ref(x_779); + x_784 = lean_box(0); } lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -x_570 = l_Lean_Meta_isExprDefEqGuarded(x_550, x_565, x_4, x_5, x_6, x_7, x_560); -if (lean_obj_tag(x_570) == 0) +x_785 = l_Lean_Meta_isExprDefEqGuarded(x_765, x_780, x_4, x_5, x_6, x_7, x_775); +if (lean_obj_tag(x_785) == 0) { -lean_object* x_571; uint8_t x_572; -x_571 = lean_ctor_get(x_570, 0); -lean_inc(x_571); -x_572 = lean_unbox(x_571); -lean_dec(x_571); -if (x_572 == 0) +lean_object* x_786; uint8_t x_787; +x_786 = lean_ctor_get(x_785, 0); +lean_inc(x_786); +x_787 = lean_unbox(x_786); +lean_dec(x_786); +if (x_787 == 0) { -lean_object* x_573; lean_object* x_574; lean_object* x_575; -lean_dec(x_569); -lean_dec(x_568); -lean_dec(x_567); -lean_dec(x_566); -lean_dec(x_563); -lean_dec(x_561); -lean_dec(x_554); -lean_dec(x_553); -lean_dec(x_552); -lean_dec(x_551); +lean_object* x_788; lean_object* x_789; lean_object* x_790; +lean_dec(x_784); +lean_dec(x_783); +lean_dec(x_782); +lean_dec(x_781); +lean_dec(x_778); +lean_dec(x_776); +lean_dec(x_769); +lean_dec(x_768); +lean_dec(x_767); +lean_dec(x_766); lean_free_object(x_12); -x_573 = lean_ctor_get(x_570, 1); -lean_inc(x_573); -lean_dec(x_570); -x_574 = lean_box(0); -x_575 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__1(x_2, x_16, x_3, x_574, x_4, x_5, x_6, x_7, x_573); -return x_575; +x_788 = lean_ctor_get(x_785, 1); +lean_inc(x_788); +lean_dec(x_785); +x_789 = lean_box(0); +x_790 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__1(x_2, x_16, x_3, x_789, x_4, x_5, x_6, x_7, x_788); +return x_790; } else { -lean_object* x_576; lean_object* x_577; +lean_object* x_791; lean_object* x_792; lean_dec(x_16); lean_dec(x_3); lean_dec(x_2); -x_576 = lean_ctor_get(x_570, 1); -lean_inc(x_576); -lean_dec(x_570); +x_791 = lean_ctor_get(x_785, 1); +lean_inc(x_791); +lean_dec(x_785); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -lean_inc(x_567); -lean_inc(x_552); -x_577 = l_Lean_Meta_isExprDefEqGuarded(x_552, x_567, x_4, x_5, x_6, x_7, x_576); -if (lean_obj_tag(x_577) == 0) +lean_inc(x_782); +lean_inc(x_767); +x_792 = l_Lean_Meta_isExprDefEqGuarded(x_767, x_782, x_4, x_5, x_6, x_7, x_791); +if (lean_obj_tag(x_792) == 0) +{ +lean_object* x_793; uint8_t x_794; +x_793 = lean_ctor_get(x_792, 0); +lean_inc(x_793); +x_794 = lean_unbox(x_793); +lean_dec(x_793); +if (x_794 == 0) +{ +lean_object* x_795; lean_object* x_796; +x_795 = lean_ctor_get(x_792, 1); +lean_inc(x_795); +lean_dec(x_792); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +x_796 = l_Lean_Meta_addPPExplicitToExposeDiff(x_767, x_782, x_4, x_5, x_6, x_7, x_795); +if (lean_obj_tag(x_796) == 0) { -lean_object* x_578; uint8_t x_579; -x_578 = lean_ctor_get(x_577, 0); -lean_inc(x_578); -x_579 = lean_unbox(x_578); -lean_dec(x_578); -if (x_579 == 0) -{ -lean_object* x_580; lean_object* x_581; -x_580 = lean_ctor_get(x_577, 1); -lean_inc(x_580); -lean_dec(x_577); +lean_object* x_797; lean_object* x_798; lean_object* x_799; lean_object* x_800; lean_object* x_801; lean_object* x_802; +x_797 = lean_ctor_get(x_796, 0); +lean_inc(x_797); +x_798 = lean_ctor_get(x_796, 1); +lean_inc(x_798); +lean_dec(x_796); +x_799 = lean_ctor_get(x_797, 0); +lean_inc(x_799); +x_800 = lean_ctor_get(x_797, 1); +lean_inc(x_800); +if (lean_is_exclusive(x_797)) { + lean_ctor_release(x_797, 0); + lean_ctor_release(x_797, 1); + x_801 = x_797; +} else { + lean_dec_ref(x_797); + x_801 = lean_box(0); +} lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -lean_inc(x_552); -x_581 = lean_infer_type(x_552, x_4, x_5, x_6, x_7, x_580); -if (lean_obj_tag(x_581) == 0) +lean_inc(x_799); +x_802 = lean_infer_type(x_799, x_4, x_5, x_6, x_7, x_798); +if (lean_obj_tag(x_802) == 0) +{ +lean_object* x_803; lean_object* x_804; lean_object* x_805; +x_803 = lean_ctor_get(x_802, 0); +lean_inc(x_803); +x_804 = lean_ctor_get(x_802, 1); +lean_inc(x_804); +lean_dec(x_802); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_800); +x_805 = lean_infer_type(x_800, x_4, x_5, x_6, x_7, x_804); +if (lean_obj_tag(x_805) == 0) { -lean_object* x_582; lean_object* x_583; lean_object* x_584; -x_582 = lean_ctor_get(x_581, 0); -lean_inc(x_582); -x_583 = lean_ctor_get(x_581, 1); -lean_inc(x_583); -lean_dec(x_581); +lean_object* x_806; lean_object* x_807; lean_object* x_808; +x_806 = lean_ctor_get(x_805, 0); +lean_inc(x_806); +x_807 = lean_ctor_get(x_805, 1); +lean_inc(x_807); +lean_dec(x_805); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -lean_inc(x_567); -x_584 = lean_infer_type(x_567, x_4, x_5, x_6, x_7, x_583); -if (lean_obj_tag(x_584) == 0) +x_808 = l_Lean_Meta_addPPExplicitToExposeDiff(x_803, x_806, x_4, x_5, x_6, x_7, x_807); +if (lean_obj_tag(x_808) == 0) { -lean_object* x_585; lean_object* x_586; lean_object* x_587; lean_object* x_588; uint8_t x_589; lean_object* x_590; lean_object* x_591; lean_object* x_592; lean_object* x_593; lean_object* x_594; lean_object* x_595; lean_object* x_596; lean_object* x_597; lean_object* x_598; lean_object* x_599; lean_object* x_600; lean_object* x_601; lean_object* x_602; lean_object* x_603; lean_object* x_604; lean_object* x_605; lean_object* x_606; lean_object* x_607; lean_object* x_608; lean_object* x_609; lean_object* x_610; -x_585 = lean_ctor_get(x_584, 0); -lean_inc(x_585); -x_586 = lean_ctor_get(x_584, 1); -lean_inc(x_586); -lean_dec(x_584); -x_587 = lean_array_get_size(x_1); -x_588 = lean_unsigned_to_nat(0u); -x_589 = lean_nat_dec_lt(x_588, x_587); -lean_dec(x_587); -x_590 = l_Lean_MessageData_ofExpr(x_552); -x_591 = l___private_Lean_Elab_Calc_0__Lean_Elab_Term_getRelUniv___lambda__1___closed__4; -if (lean_is_scalar(x_569)) { - x_592 = lean_alloc_ctor(7, 2, 0); -} else { - x_592 = x_569; - lean_ctor_set_tag(x_592, 7); -} -lean_ctor_set(x_592, 0, x_591); -lean_ctor_set(x_592, 1, x_590); -x_593 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Term_elabCalcSteps___spec__1___closed__6; -if (lean_is_scalar(x_566)) { - x_594 = lean_alloc_ctor(7, 2, 0); -} else { - x_594 = x_566; - lean_ctor_set_tag(x_594, 7); +lean_object* x_809; lean_object* x_810; lean_object* x_811; lean_object* x_812; lean_object* x_813; lean_object* x_814; lean_object* x_815; uint8_t x_816; lean_object* x_817; lean_object* x_818; lean_object* x_819; lean_object* x_820; lean_object* x_821; lean_object* x_822; lean_object* x_823; lean_object* x_824; lean_object* x_825; lean_object* x_826; lean_object* x_827; lean_object* x_828; lean_object* x_829; lean_object* x_830; lean_object* x_831; lean_object* x_832; lean_object* x_833; lean_object* x_834; lean_object* x_835; lean_object* x_836; lean_object* x_837; +x_809 = lean_ctor_get(x_808, 0); +lean_inc(x_809); +x_810 = lean_ctor_get(x_808, 1); +lean_inc(x_810); +lean_dec(x_808); +x_811 = lean_ctor_get(x_809, 0); +lean_inc(x_811); +x_812 = lean_ctor_get(x_809, 1); +lean_inc(x_812); +if (lean_is_exclusive(x_809)) { + lean_ctor_release(x_809, 0); + lean_ctor_release(x_809, 1); + x_813 = x_809; +} else { + lean_dec_ref(x_809); + x_813 = lean_box(0); +} +x_814 = lean_array_get_size(x_1); +x_815 = lean_unsigned_to_nat(0u); +x_816 = lean_nat_dec_lt(x_815, x_814); +lean_dec(x_814); +x_817 = l_Lean_MessageData_ofExpr(x_799); +x_818 = l___private_Lean_Elab_Calc_0__Lean_Elab_Term_getRelUniv___lambda__1___closed__4; +if (lean_is_scalar(x_813)) { + x_819 = lean_alloc_ctor(7, 2, 0); +} else { + x_819 = x_813; + lean_ctor_set_tag(x_819, 7); +} +lean_ctor_set(x_819, 0, x_818); +lean_ctor_set(x_819, 1, x_817); +x_820 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Term_elabCalcSteps___spec__1___closed__6; +if (lean_is_scalar(x_801)) { + x_821 = lean_alloc_ctor(7, 2, 0); +} else { + x_821 = x_801; + lean_ctor_set_tag(x_821, 7); +} +lean_ctor_set(x_821, 0, x_819); +lean_ctor_set(x_821, 1, x_820); +x_822 = l_Lean_MessageData_ofExpr(x_811); +if (lean_is_scalar(x_784)) { + x_823 = lean_alloc_ctor(7, 2, 0); +} else { + x_823 = x_784; + lean_ctor_set_tag(x_823, 7); +} +lean_ctor_set(x_823, 0, x_821); +lean_ctor_set(x_823, 1, x_822); +if (lean_is_scalar(x_781)) { + x_824 = lean_alloc_ctor(7, 2, 0); +} else { + x_824 = x_781; + lean_ctor_set_tag(x_824, 7); +} +lean_ctor_set(x_824, 0, x_823); +lean_ctor_set(x_824, 1, x_818); +x_825 = l_Lean_indentD(x_824); +x_826 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Term_elabCalcSteps___spec__1___closed__4; +if (lean_is_scalar(x_776)) { + x_827 = lean_alloc_ctor(7, 2, 0); +} else { + x_827 = x_776; + lean_ctor_set_tag(x_827, 7); +} +lean_ctor_set(x_827, 0, x_826); +lean_ctor_set(x_827, 1, x_825); +x_828 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__3___closed__4; +if (lean_is_scalar(x_769)) { + x_829 = lean_alloc_ctor(7, 2, 0); +} else { + x_829 = x_769; + lean_ctor_set_tag(x_829, 7); +} +lean_ctor_set(x_829, 0, x_827); +lean_ctor_set(x_829, 1, x_828); +x_830 = l_Lean_MessageData_ofExpr(x_800); +if (lean_is_scalar(x_766)) { + x_831 = lean_alloc_ctor(7, 2, 0); +} else { + x_831 = x_766; + lean_ctor_set_tag(x_831, 7); +} +lean_ctor_set(x_831, 0, x_818); +lean_ctor_set(x_831, 1, x_830); +x_832 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_832, 0, x_831); +lean_ctor_set(x_832, 1, x_820); +x_833 = l_Lean_MessageData_ofExpr(x_812); +lean_ctor_set_tag(x_12, 7); +lean_ctor_set(x_12, 1, x_833); +lean_ctor_set(x_12, 0, x_832); +x_834 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_834, 0, x_12); +lean_ctor_set(x_834, 1, x_818); +x_835 = l_Lean_indentD(x_834); +x_836 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_836, 0, x_829); +lean_ctor_set(x_836, 1, x_835); +x_837 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_837, 0, x_836); +lean_ctor_set(x_837, 1, x_818); +if (x_816 == 0) +{ +lean_object* x_838; lean_object* x_839; lean_object* x_840; uint8_t x_841; lean_object* x_842; lean_object* x_843; uint8_t x_844; lean_object* x_845; lean_object* x_846; +x_838 = l_Lean_Elab_Term_instInhabitedCalcStepView; +x_839 = l_outOfBounds___rarg(x_838); +x_840 = lean_ctor_get(x_839, 1); +lean_inc(x_840); +lean_dec(x_839); +x_841 = 2; +lean_inc(x_6); +x_842 = l_Lean_logAt___at_Lean_Elab_Term_reportUnsolvedGoals___spec__2(x_840, x_837, x_841, x_4, x_5, x_6, x_7, x_810); +lean_dec(x_840); +x_843 = lean_ctor_get(x_842, 1); +lean_inc(x_843); +lean_dec(x_842); +x_844 = 1; +x_845 = lean_box(0); +x_846 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__3(x_768, x_783, x_778, x_1, x_844, x_845, x_4, x_5, x_6, x_7, x_843); +return x_846; } -lean_ctor_set(x_594, 0, x_592); -lean_ctor_set(x_594, 1, x_593); -x_595 = l_Lean_MessageData_ofExpr(x_582); -if (lean_is_scalar(x_561)) { - x_596 = lean_alloc_ctor(7, 2, 0); -} else { - x_596 = x_561; - lean_ctor_set_tag(x_596, 7); +else +{ +lean_object* x_847; lean_object* x_848; uint8_t x_849; lean_object* x_850; lean_object* x_851; uint8_t x_852; lean_object* x_853; lean_object* x_854; +x_847 = lean_array_fget(x_1, x_815); +x_848 = lean_ctor_get(x_847, 1); +lean_inc(x_848); +lean_dec(x_847); +x_849 = 2; +lean_inc(x_6); +x_850 = l_Lean_logAt___at_Lean_Elab_Term_reportUnsolvedGoals___spec__2(x_848, x_837, x_849, x_4, x_5, x_6, x_7, x_810); +lean_dec(x_848); +x_851 = lean_ctor_get(x_850, 1); +lean_inc(x_851); +lean_dec(x_850); +x_852 = 1; +x_853 = lean_box(0); +x_854 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__3(x_768, x_783, x_778, x_1, x_852, x_853, x_4, x_5, x_6, x_7, x_851); +return x_854; } -lean_ctor_set(x_596, 0, x_594); -lean_ctor_set(x_596, 1, x_595); -if (lean_is_scalar(x_554)) { - x_597 = lean_alloc_ctor(7, 2, 0); +} +else +{ +lean_object* x_855; lean_object* x_856; lean_object* x_857; lean_object* x_858; +lean_dec(x_801); +lean_dec(x_800); +lean_dec(x_799); +lean_dec(x_784); +lean_dec(x_783); +lean_dec(x_781); +lean_dec(x_778); +lean_dec(x_776); +lean_dec(x_769); +lean_dec(x_768); +lean_dec(x_766); +lean_free_object(x_12); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_855 = lean_ctor_get(x_808, 0); +lean_inc(x_855); +x_856 = lean_ctor_get(x_808, 1); +lean_inc(x_856); +if (lean_is_exclusive(x_808)) { + lean_ctor_release(x_808, 0); + lean_ctor_release(x_808, 1); + x_857 = x_808; } else { - x_597 = x_554; - lean_ctor_set_tag(x_597, 7); + lean_dec_ref(x_808); + x_857 = lean_box(0); } -lean_ctor_set(x_597, 0, x_596); -lean_ctor_set(x_597, 1, x_591); -x_598 = l_Lean_indentD(x_597); -x_599 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Term_elabCalcSteps___spec__1___closed__4; -if (lean_is_scalar(x_551)) { - x_600 = lean_alloc_ctor(7, 2, 0); +if (lean_is_scalar(x_857)) { + x_858 = lean_alloc_ctor(1, 2, 0); } else { - x_600 = x_551; - lean_ctor_set_tag(x_600, 7); + x_858 = x_857; +} +lean_ctor_set(x_858, 0, x_855); +lean_ctor_set(x_858, 1, x_856); +return x_858; } -lean_ctor_set(x_600, 0, x_599); -lean_ctor_set(x_600, 1, x_598); -x_601 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__3___closed__4; -x_602 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_602, 0, x_600); -lean_ctor_set(x_602, 1, x_601); -x_603 = l_Lean_MessageData_ofExpr(x_567); -lean_ctor_set_tag(x_12, 7); -lean_ctor_set(x_12, 1, x_603); -lean_ctor_set(x_12, 0, x_591); -x_604 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_604, 0, x_12); -lean_ctor_set(x_604, 1, x_593); -x_605 = l_Lean_MessageData_ofExpr(x_585); -x_606 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_606, 0, x_604); -lean_ctor_set(x_606, 1, x_605); -x_607 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_607, 0, x_606); -lean_ctor_set(x_607, 1, x_591); -x_608 = l_Lean_indentD(x_607); -x_609 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_609, 0, x_602); -lean_ctor_set(x_609, 1, x_608); -x_610 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_610, 0, x_609); -lean_ctor_set(x_610, 1, x_591); -if (x_589 == 0) -{ -lean_object* x_611; lean_object* x_612; lean_object* x_613; uint8_t x_614; lean_object* x_615; lean_object* x_616; uint8_t x_617; lean_object* x_618; lean_object* x_619; -x_611 = l_Lean_Elab_Term_instInhabitedCalcStepView; -x_612 = l_outOfBounds___rarg(x_611); -x_613 = lean_ctor_get(x_612, 1); -lean_inc(x_613); -lean_dec(x_612); -x_614 = 2; -lean_inc(x_6); -x_615 = l_Lean_logAt___at_Lean_Elab_Term_reportUnsolvedGoals___spec__2(x_613, x_610, x_614, x_4, x_5, x_6, x_7, x_586); -lean_dec(x_613); -x_616 = lean_ctor_get(x_615, 1); -lean_inc(x_616); -lean_dec(x_615); -x_617 = 1; -x_618 = lean_box(0); -x_619 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__3(x_553, x_568, x_563, x_1, x_617, x_618, x_4, x_5, x_6, x_7, x_616); -return x_619; } else { -lean_object* x_620; lean_object* x_621; uint8_t x_622; lean_object* x_623; lean_object* x_624; uint8_t x_625; lean_object* x_626; lean_object* x_627; -x_620 = lean_array_fget(x_1, x_588); -x_621 = lean_ctor_get(x_620, 1); -lean_inc(x_621); -lean_dec(x_620); -x_622 = 2; -lean_inc(x_6); -x_623 = l_Lean_logAt___at_Lean_Elab_Term_reportUnsolvedGoals___spec__2(x_621, x_610, x_622, x_4, x_5, x_6, x_7, x_586); -lean_dec(x_621); -x_624 = lean_ctor_get(x_623, 1); -lean_inc(x_624); -lean_dec(x_623); -x_625 = 1; -x_626 = lean_box(0); -x_627 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__3(x_553, x_568, x_563, x_1, x_625, x_626, x_4, x_5, x_6, x_7, x_624); -return x_627; +lean_object* x_859; lean_object* x_860; lean_object* x_861; lean_object* x_862; +lean_dec(x_803); +lean_dec(x_801); +lean_dec(x_800); +lean_dec(x_799); +lean_dec(x_784); +lean_dec(x_783); +lean_dec(x_781); +lean_dec(x_778); +lean_dec(x_776); +lean_dec(x_769); +lean_dec(x_768); +lean_dec(x_766); +lean_free_object(x_12); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_859 = lean_ctor_get(x_805, 0); +lean_inc(x_859); +x_860 = lean_ctor_get(x_805, 1); +lean_inc(x_860); +if (lean_is_exclusive(x_805)) { + lean_ctor_release(x_805, 0); + lean_ctor_release(x_805, 1); + x_861 = x_805; +} else { + lean_dec_ref(x_805); + x_861 = lean_box(0); +} +if (lean_is_scalar(x_861)) { + x_862 = lean_alloc_ctor(1, 2, 0); +} else { + x_862 = x_861; +} +lean_ctor_set(x_862, 0, x_859); +lean_ctor_set(x_862, 1, x_860); +return x_862; } } else { -lean_object* x_628; lean_object* x_629; lean_object* x_630; lean_object* x_631; -lean_dec(x_582); -lean_dec(x_569); -lean_dec(x_568); -lean_dec(x_567); -lean_dec(x_566); -lean_dec(x_563); -lean_dec(x_561); -lean_dec(x_554); -lean_dec(x_553); -lean_dec(x_552); -lean_dec(x_551); +lean_object* x_863; lean_object* x_864; lean_object* x_865; lean_object* x_866; +lean_dec(x_801); +lean_dec(x_800); +lean_dec(x_799); +lean_dec(x_784); +lean_dec(x_783); +lean_dec(x_781); +lean_dec(x_778); +lean_dec(x_776); +lean_dec(x_769); +lean_dec(x_768); +lean_dec(x_766); lean_free_object(x_12); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_628 = lean_ctor_get(x_584, 0); -lean_inc(x_628); -x_629 = lean_ctor_get(x_584, 1); -lean_inc(x_629); -if (lean_is_exclusive(x_584)) { - lean_ctor_release(x_584, 0); - lean_ctor_release(x_584, 1); - x_630 = x_584; +x_863 = lean_ctor_get(x_802, 0); +lean_inc(x_863); +x_864 = lean_ctor_get(x_802, 1); +lean_inc(x_864); +if (lean_is_exclusive(x_802)) { + lean_ctor_release(x_802, 0); + lean_ctor_release(x_802, 1); + x_865 = x_802; } else { - lean_dec_ref(x_584); - x_630 = lean_box(0); + lean_dec_ref(x_802); + x_865 = lean_box(0); } -if (lean_is_scalar(x_630)) { - x_631 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_865)) { + x_866 = lean_alloc_ctor(1, 2, 0); } else { - x_631 = x_630; -} -lean_ctor_set(x_631, 0, x_628); -lean_ctor_set(x_631, 1, x_629); -return x_631; + x_866 = x_865; +} +lean_ctor_set(x_866, 0, x_863); +lean_ctor_set(x_866, 1, x_864); +return x_866; } } else { -lean_object* x_632; lean_object* x_633; lean_object* x_634; lean_object* x_635; -lean_dec(x_569); -lean_dec(x_568); -lean_dec(x_567); -lean_dec(x_566); -lean_dec(x_563); -lean_dec(x_561); -lean_dec(x_554); -lean_dec(x_553); -lean_dec(x_552); -lean_dec(x_551); +lean_object* x_867; lean_object* x_868; lean_object* x_869; lean_object* x_870; +lean_dec(x_784); +lean_dec(x_783); +lean_dec(x_781); +lean_dec(x_778); +lean_dec(x_776); +lean_dec(x_769); +lean_dec(x_768); +lean_dec(x_766); lean_free_object(x_12); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_632 = lean_ctor_get(x_581, 0); -lean_inc(x_632); -x_633 = lean_ctor_get(x_581, 1); -lean_inc(x_633); -if (lean_is_exclusive(x_581)) { - lean_ctor_release(x_581, 0); - lean_ctor_release(x_581, 1); - x_634 = x_581; +x_867 = lean_ctor_get(x_796, 0); +lean_inc(x_867); +x_868 = lean_ctor_get(x_796, 1); +lean_inc(x_868); +if (lean_is_exclusive(x_796)) { + lean_ctor_release(x_796, 0); + lean_ctor_release(x_796, 1); + x_869 = x_796; } else { - lean_dec_ref(x_581); - x_634 = lean_box(0); + lean_dec_ref(x_796); + x_869 = lean_box(0); } -if (lean_is_scalar(x_634)) { - x_635 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_869)) { + x_870 = lean_alloc_ctor(1, 2, 0); } else { - x_635 = x_634; + x_870 = x_869; } -lean_ctor_set(x_635, 0, x_632); -lean_ctor_set(x_635, 1, x_633); -return x_635; +lean_ctor_set(x_870, 0, x_867); +lean_ctor_set(x_870, 1, x_868); +return x_870; } } else { -lean_object* x_636; uint8_t x_637; lean_object* x_638; lean_object* x_639; -lean_dec(x_569); -lean_dec(x_567); -lean_dec(x_566); -lean_dec(x_561); -lean_dec(x_554); -lean_dec(x_552); -lean_dec(x_551); +lean_object* x_871; uint8_t x_872; lean_object* x_873; lean_object* x_874; +lean_dec(x_784); +lean_dec(x_782); +lean_dec(x_781); +lean_dec(x_776); +lean_dec(x_769); +lean_dec(x_767); +lean_dec(x_766); lean_free_object(x_12); -x_636 = lean_ctor_get(x_577, 1); -lean_inc(x_636); -lean_dec(x_577); -x_637 = 0; -x_638 = lean_box(0); -x_639 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__3(x_553, x_568, x_563, x_1, x_637, x_638, x_4, x_5, x_6, x_7, x_636); -return x_639; +x_871 = lean_ctor_get(x_792, 1); +lean_inc(x_871); +lean_dec(x_792); +x_872 = 0; +x_873 = lean_box(0); +x_874 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__3(x_768, x_783, x_778, x_1, x_872, x_873, x_4, x_5, x_6, x_7, x_871); +return x_874; } } else { -lean_object* x_640; lean_object* x_641; lean_object* x_642; lean_object* x_643; -lean_dec(x_569); -lean_dec(x_568); -lean_dec(x_567); -lean_dec(x_566); -lean_dec(x_563); -lean_dec(x_561); -lean_dec(x_554); -lean_dec(x_553); -lean_dec(x_552); -lean_dec(x_551); +lean_object* x_875; lean_object* x_876; lean_object* x_877; lean_object* x_878; +lean_dec(x_784); +lean_dec(x_783); +lean_dec(x_782); +lean_dec(x_781); +lean_dec(x_778); +lean_dec(x_776); +lean_dec(x_769); +lean_dec(x_768); +lean_dec(x_767); +lean_dec(x_766); lean_free_object(x_12); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_640 = lean_ctor_get(x_577, 0); -lean_inc(x_640); -x_641 = lean_ctor_get(x_577, 1); -lean_inc(x_641); -if (lean_is_exclusive(x_577)) { - lean_ctor_release(x_577, 0); - lean_ctor_release(x_577, 1); - x_642 = x_577; +x_875 = lean_ctor_get(x_792, 0); +lean_inc(x_875); +x_876 = lean_ctor_get(x_792, 1); +lean_inc(x_876); +if (lean_is_exclusive(x_792)) { + lean_ctor_release(x_792, 0); + lean_ctor_release(x_792, 1); + x_877 = x_792; } else { - lean_dec_ref(x_577); - x_642 = lean_box(0); + lean_dec_ref(x_792); + x_877 = lean_box(0); } -if (lean_is_scalar(x_642)) { - x_643 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_877)) { + x_878 = lean_alloc_ctor(1, 2, 0); } else { - x_643 = x_642; + x_878 = x_877; } -lean_ctor_set(x_643, 0, x_640); -lean_ctor_set(x_643, 1, x_641); -return x_643; +lean_ctor_set(x_878, 0, x_875); +lean_ctor_set(x_878, 1, x_876); +return x_878; } } } else { -lean_object* x_644; lean_object* x_645; lean_object* x_646; lean_object* x_647; -lean_dec(x_569); -lean_dec(x_568); -lean_dec(x_567); -lean_dec(x_566); -lean_dec(x_563); -lean_dec(x_561); -lean_dec(x_554); -lean_dec(x_553); -lean_dec(x_552); -lean_dec(x_551); +lean_object* x_879; lean_object* x_880; lean_object* x_881; lean_object* x_882; +lean_dec(x_784); +lean_dec(x_783); +lean_dec(x_782); +lean_dec(x_781); +lean_dec(x_778); +lean_dec(x_776); +lean_dec(x_769); +lean_dec(x_768); +lean_dec(x_767); +lean_dec(x_766); lean_dec(x_16); lean_free_object(x_12); lean_dec(x_7); @@ -22910,26 +24652,26 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_644 = lean_ctor_get(x_570, 0); -lean_inc(x_644); -x_645 = lean_ctor_get(x_570, 1); -lean_inc(x_645); -if (lean_is_exclusive(x_570)) { - lean_ctor_release(x_570, 0); - lean_ctor_release(x_570, 1); - x_646 = x_570; +x_879 = lean_ctor_get(x_785, 0); +lean_inc(x_879); +x_880 = lean_ctor_get(x_785, 1); +lean_inc(x_880); +if (lean_is_exclusive(x_785)) { + lean_ctor_release(x_785, 0); + lean_ctor_release(x_785, 1); + x_881 = x_785; } else { - lean_dec_ref(x_570); - x_646 = lean_box(0); + lean_dec_ref(x_785); + x_881 = lean_box(0); } -if (lean_is_scalar(x_646)) { - x_647 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_881)) { + x_882 = lean_alloc_ctor(1, 2, 0); } else { - x_647 = x_646; + x_882 = x_881; } -lean_ctor_set(x_647, 0, x_644); -lean_ctor_set(x_647, 1, x_645); -return x_647; +lean_ctor_set(x_882, 0, x_879); +lean_ctor_set(x_882, 1, x_880); +return x_882; } } } @@ -22937,528 +24679,669 @@ return x_647; } else { -lean_object* x_648; lean_object* x_649; lean_object* x_650; lean_object* x_651; lean_object* x_652; -x_648 = lean_ctor_get(x_12, 0); -x_649 = lean_ctor_get(x_12, 1); -lean_inc(x_649); -lean_inc(x_648); +lean_object* x_883; lean_object* x_884; lean_object* x_885; lean_object* x_886; lean_object* x_887; +x_883 = lean_ctor_get(x_12, 0); +x_884 = lean_ctor_get(x_12, 1); +lean_inc(x_884); +lean_inc(x_883); lean_dec(x_12); -x_650 = l_Lean_Expr_headBeta(x_648); -x_651 = l_Lean_Elab_Term_getCalcRelation_x3f(x_650, x_4, x_5, x_6, x_7, x_649); -x_652 = lean_ctor_get(x_651, 0); -lean_inc(x_652); -if (lean_obj_tag(x_652) == 0) -{ -lean_object* x_653; lean_object* x_654; lean_object* x_655; -lean_dec(x_650); +x_885 = l_Lean_Expr_headBeta(x_883); +x_886 = l_Lean_Elab_Term_getCalcRelation_x3f(x_885, x_4, x_5, x_6, x_7, x_884); +x_887 = lean_ctor_get(x_886, 0); +lean_inc(x_887); +if (lean_obj_tag(x_887) == 0) +{ +lean_object* x_888; lean_object* x_889; lean_object* x_890; +lean_dec(x_885); lean_dec(x_3); lean_dec(x_2); -x_653 = lean_ctor_get(x_651, 1); -lean_inc(x_653); -lean_dec(x_651); -x_654 = l_Lean_Elab_Term_throwCalcFailure___rarg___closed__2; -x_655 = l_panic___at_Lean_Elab_Term_throwCalcFailure___spec__1___rarg(x_654, x_4, x_5, x_6, x_7, x_653); -return x_655; +x_888 = lean_ctor_get(x_886, 1); +lean_inc(x_888); +lean_dec(x_886); +x_889 = l_Lean_Elab_Term_throwCalcFailure___rarg___closed__2; +x_890 = l_panic___at_Lean_Elab_Term_throwCalcFailure___spec__1___rarg(x_889, x_4, x_5, x_6, x_7, x_888); +return x_890; } else { -lean_object* x_656; lean_object* x_657; lean_object* x_658; lean_object* x_659; lean_object* x_660; lean_object* x_661; lean_object* x_662; lean_object* x_663; lean_object* x_664; lean_object* x_665; lean_object* x_666; -x_656 = lean_ctor_get(x_652, 0); -lean_inc(x_656); -lean_dec(x_652); -x_657 = lean_ctor_get(x_656, 1); -lean_inc(x_657); -x_658 = lean_ctor_get(x_651, 1); -lean_inc(x_658); -if (lean_is_exclusive(x_651)) { - lean_ctor_release(x_651, 0); - lean_ctor_release(x_651, 1); - x_659 = x_651; +lean_object* x_891; lean_object* x_892; lean_object* x_893; lean_object* x_894; lean_object* x_895; lean_object* x_896; lean_object* x_897; lean_object* x_898; lean_object* x_899; lean_object* x_900; lean_object* x_901; +x_891 = lean_ctor_get(x_887, 0); +lean_inc(x_891); +lean_dec(x_887); +x_892 = lean_ctor_get(x_891, 1); +lean_inc(x_892); +x_893 = lean_ctor_get(x_886, 1); +lean_inc(x_893); +if (lean_is_exclusive(x_886)) { + lean_ctor_release(x_886, 0); + lean_ctor_release(x_886, 1); + x_894 = x_886; +} else { + lean_dec_ref(x_886); + x_894 = lean_box(0); +} +x_895 = lean_ctor_get(x_891, 0); +lean_inc(x_895); +if (lean_is_exclusive(x_891)) { + lean_ctor_release(x_891, 0); + lean_ctor_release(x_891, 1); + x_896 = x_891; } else { - lean_dec_ref(x_651); - x_659 = lean_box(0); + lean_dec_ref(x_891); + x_896 = lean_box(0); } -x_660 = lean_ctor_get(x_656, 0); -lean_inc(x_660); -if (lean_is_exclusive(x_656)) { - lean_ctor_release(x_656, 0); - lean_ctor_release(x_656, 1); - x_661 = x_656; +x_897 = lean_ctor_get(x_892, 0); +lean_inc(x_897); +x_898 = lean_ctor_get(x_892, 1); +lean_inc(x_898); +if (lean_is_exclusive(x_892)) { + lean_ctor_release(x_892, 0); + lean_ctor_release(x_892, 1); + x_899 = x_892; } else { - lean_dec_ref(x_656); - x_661 = lean_box(0); + lean_dec_ref(x_892); + x_899 = lean_box(0); } -x_662 = lean_ctor_get(x_657, 0); -lean_inc(x_662); -x_663 = lean_ctor_get(x_657, 1); -lean_inc(x_663); -if (lean_is_exclusive(x_657)) { - lean_ctor_release(x_657, 0); - lean_ctor_release(x_657, 1); - x_664 = x_657; -} else { - lean_dec_ref(x_657); - x_664 = lean_box(0); -} -x_665 = l_Lean_Elab_Term_getCalcRelation_x3f(x_2, x_4, x_5, x_6, x_7, x_658); -x_666 = lean_ctor_get(x_665, 0); -lean_inc(x_666); -if (lean_obj_tag(x_666) == 0) -{ -lean_object* x_667; lean_object* x_668; lean_object* x_669; -lean_dec(x_664); -lean_dec(x_663); -lean_dec(x_662); -lean_dec(x_661); -lean_dec(x_660); -lean_dec(x_659); -x_667 = lean_ctor_get(x_665, 1); -lean_inc(x_667); -lean_dec(x_665); -x_668 = lean_box(0); -x_669 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__1(x_2, x_650, x_3, x_668, x_4, x_5, x_6, x_7, x_667); -return x_669; +x_900 = l_Lean_Elab_Term_getCalcRelation_x3f(x_2, x_4, x_5, x_6, x_7, x_893); +x_901 = lean_ctor_get(x_900, 0); +lean_inc(x_901); +if (lean_obj_tag(x_901) == 0) +{ +lean_object* x_902; lean_object* x_903; lean_object* x_904; +lean_dec(x_899); +lean_dec(x_898); +lean_dec(x_897); +lean_dec(x_896); +lean_dec(x_895); +lean_dec(x_894); +x_902 = lean_ctor_get(x_900, 1); +lean_inc(x_902); +lean_dec(x_900); +x_903 = lean_box(0); +x_904 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__1(x_2, x_885, x_3, x_903, x_4, x_5, x_6, x_7, x_902); +return x_904; } else { -lean_object* x_670; lean_object* x_671; lean_object* x_672; lean_object* x_673; lean_object* x_674; lean_object* x_675; lean_object* x_676; lean_object* x_677; lean_object* x_678; lean_object* x_679; lean_object* x_680; -x_670 = lean_ctor_get(x_665, 1); -lean_inc(x_670); -if (lean_is_exclusive(x_665)) { - lean_ctor_release(x_665, 0); - lean_ctor_release(x_665, 1); - x_671 = x_665; +lean_object* x_905; lean_object* x_906; lean_object* x_907; lean_object* x_908; lean_object* x_909; lean_object* x_910; lean_object* x_911; lean_object* x_912; lean_object* x_913; lean_object* x_914; lean_object* x_915; +x_905 = lean_ctor_get(x_900, 1); +lean_inc(x_905); +if (lean_is_exclusive(x_900)) { + lean_ctor_release(x_900, 0); + lean_ctor_release(x_900, 1); + x_906 = x_900; } else { - lean_dec_ref(x_665); - x_671 = lean_box(0); + lean_dec_ref(x_900); + x_906 = lean_box(0); } -x_672 = lean_ctor_get(x_666, 0); -lean_inc(x_672); -lean_dec(x_666); +x_907 = lean_ctor_get(x_901, 0); +lean_inc(x_907); +lean_dec(x_901); lean_inc(x_3); -lean_inc(x_650); +lean_inc(x_885); lean_inc(x_2); -x_673 = lean_alloc_closure((void*)(l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__1___boxed), 9, 3); -lean_closure_set(x_673, 0, x_2); -lean_closure_set(x_673, 1, x_650); -lean_closure_set(x_673, 2, x_3); -x_674 = lean_ctor_get(x_672, 1); -lean_inc(x_674); -x_675 = lean_ctor_get(x_672, 0); -lean_inc(x_675); -if (lean_is_exclusive(x_672)) { - lean_ctor_release(x_672, 0); - lean_ctor_release(x_672, 1); - x_676 = x_672; -} else { - lean_dec_ref(x_672); - x_676 = lean_box(0); -} -x_677 = lean_ctor_get(x_674, 0); -lean_inc(x_677); -x_678 = lean_ctor_get(x_674, 1); -lean_inc(x_678); -if (lean_is_exclusive(x_674)) { - lean_ctor_release(x_674, 0); - lean_ctor_release(x_674, 1); - x_679 = x_674; -} else { - lean_dec_ref(x_674); - x_679 = lean_box(0); +x_908 = lean_alloc_closure((void*)(l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__1___boxed), 9, 3); +lean_closure_set(x_908, 0, x_2); +lean_closure_set(x_908, 1, x_885); +lean_closure_set(x_908, 2, x_3); +x_909 = lean_ctor_get(x_907, 1); +lean_inc(x_909); +x_910 = lean_ctor_get(x_907, 0); +lean_inc(x_910); +if (lean_is_exclusive(x_907)) { + lean_ctor_release(x_907, 0); + lean_ctor_release(x_907, 1); + x_911 = x_907; +} else { + lean_dec_ref(x_907); + x_911 = lean_box(0); +} +x_912 = lean_ctor_get(x_909, 0); +lean_inc(x_912); +x_913 = lean_ctor_get(x_909, 1); +lean_inc(x_913); +if (lean_is_exclusive(x_909)) { + lean_ctor_release(x_909, 0); + lean_ctor_release(x_909, 1); + x_914 = x_909; +} else { + lean_dec_ref(x_909); + x_914 = lean_box(0); } lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -x_680 = l_Lean_Meta_isExprDefEqGuarded(x_660, x_675, x_4, x_5, x_6, x_7, x_670); -if (lean_obj_tag(x_680) == 0) +x_915 = l_Lean_Meta_isExprDefEqGuarded(x_895, x_910, x_4, x_5, x_6, x_7, x_905); +if (lean_obj_tag(x_915) == 0) { -lean_object* x_681; uint8_t x_682; -x_681 = lean_ctor_get(x_680, 0); -lean_inc(x_681); -x_682 = lean_unbox(x_681); -lean_dec(x_681); -if (x_682 == 0) +lean_object* x_916; uint8_t x_917; +x_916 = lean_ctor_get(x_915, 0); +lean_inc(x_916); +x_917 = lean_unbox(x_916); +lean_dec(x_916); +if (x_917 == 0) { -lean_object* x_683; lean_object* x_684; lean_object* x_685; -lean_dec(x_679); -lean_dec(x_678); -lean_dec(x_677); -lean_dec(x_676); -lean_dec(x_673); -lean_dec(x_671); -lean_dec(x_664); -lean_dec(x_663); -lean_dec(x_662); -lean_dec(x_661); -lean_dec(x_659); -x_683 = lean_ctor_get(x_680, 1); -lean_inc(x_683); -lean_dec(x_680); -x_684 = lean_box(0); -x_685 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__1(x_2, x_650, x_3, x_684, x_4, x_5, x_6, x_7, x_683); -return x_685; +lean_object* x_918; lean_object* x_919; lean_object* x_920; +lean_dec(x_914); +lean_dec(x_913); +lean_dec(x_912); +lean_dec(x_911); +lean_dec(x_908); +lean_dec(x_906); +lean_dec(x_899); +lean_dec(x_898); +lean_dec(x_897); +lean_dec(x_896); +lean_dec(x_894); +x_918 = lean_ctor_get(x_915, 1); +lean_inc(x_918); +lean_dec(x_915); +x_919 = lean_box(0); +x_920 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__1(x_2, x_885, x_3, x_919, x_4, x_5, x_6, x_7, x_918); +return x_920; } else { -lean_object* x_686; lean_object* x_687; -lean_dec(x_650); +lean_object* x_921; lean_object* x_922; +lean_dec(x_885); lean_dec(x_3); lean_dec(x_2); -x_686 = lean_ctor_get(x_680, 1); -lean_inc(x_686); -lean_dec(x_680); +x_921 = lean_ctor_get(x_915, 1); +lean_inc(x_921); +lean_dec(x_915); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -lean_inc(x_677); -lean_inc(x_662); -x_687 = l_Lean_Meta_isExprDefEqGuarded(x_662, x_677, x_4, x_5, x_6, x_7, x_686); -if (lean_obj_tag(x_687) == 0) -{ -lean_object* x_688; uint8_t x_689; -x_688 = lean_ctor_get(x_687, 0); -lean_inc(x_688); -x_689 = lean_unbox(x_688); -lean_dec(x_688); -if (x_689 == 0) -{ -lean_object* x_690; lean_object* x_691; -x_690 = lean_ctor_get(x_687, 1); -lean_inc(x_690); -lean_dec(x_687); +lean_inc(x_912); +lean_inc(x_897); +x_922 = l_Lean_Meta_isExprDefEqGuarded(x_897, x_912, x_4, x_5, x_6, x_7, x_921); +if (lean_obj_tag(x_922) == 0) +{ +lean_object* x_923; uint8_t x_924; +x_923 = lean_ctor_get(x_922, 0); +lean_inc(x_923); +x_924 = lean_unbox(x_923); +lean_dec(x_923); +if (x_924 == 0) +{ +lean_object* x_925; lean_object* x_926; +x_925 = lean_ctor_get(x_922, 1); +lean_inc(x_925); +lean_dec(x_922); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -lean_inc(x_662); -x_691 = lean_infer_type(x_662, x_4, x_5, x_6, x_7, x_690); -if (lean_obj_tag(x_691) == 0) +x_926 = l_Lean_Meta_addPPExplicitToExposeDiff(x_897, x_912, x_4, x_5, x_6, x_7, x_925); +if (lean_obj_tag(x_926) == 0) +{ +lean_object* x_927; lean_object* x_928; lean_object* x_929; lean_object* x_930; lean_object* x_931; lean_object* x_932; +x_927 = lean_ctor_get(x_926, 0); +lean_inc(x_927); +x_928 = lean_ctor_get(x_926, 1); +lean_inc(x_928); +lean_dec(x_926); +x_929 = lean_ctor_get(x_927, 0); +lean_inc(x_929); +x_930 = lean_ctor_get(x_927, 1); +lean_inc(x_930); +if (lean_is_exclusive(x_927)) { + lean_ctor_release(x_927, 0); + lean_ctor_release(x_927, 1); + x_931 = x_927; +} else { + lean_dec_ref(x_927); + x_931 = lean_box(0); +} +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_929); +x_932 = lean_infer_type(x_929, x_4, x_5, x_6, x_7, x_928); +if (lean_obj_tag(x_932) == 0) +{ +lean_object* x_933; lean_object* x_934; lean_object* x_935; +x_933 = lean_ctor_get(x_932, 0); +lean_inc(x_933); +x_934 = lean_ctor_get(x_932, 1); +lean_inc(x_934); +lean_dec(x_932); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_930); +x_935 = lean_infer_type(x_930, x_4, x_5, x_6, x_7, x_934); +if (lean_obj_tag(x_935) == 0) { -lean_object* x_692; lean_object* x_693; lean_object* x_694; -x_692 = lean_ctor_get(x_691, 0); -lean_inc(x_692); -x_693 = lean_ctor_get(x_691, 1); -lean_inc(x_693); -lean_dec(x_691); +lean_object* x_936; lean_object* x_937; lean_object* x_938; +x_936 = lean_ctor_get(x_935, 0); +lean_inc(x_936); +x_937 = lean_ctor_get(x_935, 1); +lean_inc(x_937); +lean_dec(x_935); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -lean_inc(x_677); -x_694 = lean_infer_type(x_677, x_4, x_5, x_6, x_7, x_693); -if (lean_obj_tag(x_694) == 0) -{ -lean_object* x_695; lean_object* x_696; lean_object* x_697; lean_object* x_698; uint8_t x_699; lean_object* x_700; lean_object* x_701; lean_object* x_702; lean_object* x_703; lean_object* x_704; lean_object* x_705; lean_object* x_706; lean_object* x_707; lean_object* x_708; lean_object* x_709; lean_object* x_710; lean_object* x_711; lean_object* x_712; lean_object* x_713; lean_object* x_714; lean_object* x_715; lean_object* x_716; lean_object* x_717; lean_object* x_718; lean_object* x_719; lean_object* x_720; lean_object* x_721; -x_695 = lean_ctor_get(x_694, 0); -lean_inc(x_695); -x_696 = lean_ctor_get(x_694, 1); -lean_inc(x_696); -lean_dec(x_694); -x_697 = lean_array_get_size(x_1); -x_698 = lean_unsigned_to_nat(0u); -x_699 = lean_nat_dec_lt(x_698, x_697); -lean_dec(x_697); -x_700 = l_Lean_MessageData_ofExpr(x_662); -x_701 = l___private_Lean_Elab_Calc_0__Lean_Elab_Term_getRelUniv___lambda__1___closed__4; -if (lean_is_scalar(x_679)) { - x_702 = lean_alloc_ctor(7, 2, 0); +x_938 = l_Lean_Meta_addPPExplicitToExposeDiff(x_933, x_936, x_4, x_5, x_6, x_7, x_937); +if (lean_obj_tag(x_938) == 0) +{ +lean_object* x_939; lean_object* x_940; lean_object* x_941; lean_object* x_942; lean_object* x_943; lean_object* x_944; lean_object* x_945; uint8_t x_946; lean_object* x_947; lean_object* x_948; lean_object* x_949; lean_object* x_950; lean_object* x_951; lean_object* x_952; lean_object* x_953; lean_object* x_954; lean_object* x_955; lean_object* x_956; lean_object* x_957; lean_object* x_958; lean_object* x_959; lean_object* x_960; lean_object* x_961; lean_object* x_962; lean_object* x_963; lean_object* x_964; lean_object* x_965; lean_object* x_966; lean_object* x_967; lean_object* x_968; +x_939 = lean_ctor_get(x_938, 0); +lean_inc(x_939); +x_940 = lean_ctor_get(x_938, 1); +lean_inc(x_940); +lean_dec(x_938); +x_941 = lean_ctor_get(x_939, 0); +lean_inc(x_941); +x_942 = lean_ctor_get(x_939, 1); +lean_inc(x_942); +if (lean_is_exclusive(x_939)) { + lean_ctor_release(x_939, 0); + lean_ctor_release(x_939, 1); + x_943 = x_939; } else { - x_702 = x_679; - lean_ctor_set_tag(x_702, 7); + lean_dec_ref(x_939); + x_943 = lean_box(0); } -lean_ctor_set(x_702, 0, x_701); -lean_ctor_set(x_702, 1, x_700); -x_703 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Term_elabCalcSteps___spec__1___closed__6; -if (lean_is_scalar(x_676)) { - x_704 = lean_alloc_ctor(7, 2, 0); +x_944 = lean_array_get_size(x_1); +x_945 = lean_unsigned_to_nat(0u); +x_946 = lean_nat_dec_lt(x_945, x_944); +lean_dec(x_944); +x_947 = l_Lean_MessageData_ofExpr(x_929); +x_948 = l___private_Lean_Elab_Calc_0__Lean_Elab_Term_getRelUniv___lambda__1___closed__4; +if (lean_is_scalar(x_943)) { + x_949 = lean_alloc_ctor(7, 2, 0); } else { - x_704 = x_676; - lean_ctor_set_tag(x_704, 7); + x_949 = x_943; + lean_ctor_set_tag(x_949, 7); } -lean_ctor_set(x_704, 0, x_702); -lean_ctor_set(x_704, 1, x_703); -x_705 = l_Lean_MessageData_ofExpr(x_692); -if (lean_is_scalar(x_671)) { - x_706 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_949, 0, x_948); +lean_ctor_set(x_949, 1, x_947); +x_950 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Term_elabCalcSteps___spec__1___closed__6; +if (lean_is_scalar(x_931)) { + x_951 = lean_alloc_ctor(7, 2, 0); } else { - x_706 = x_671; - lean_ctor_set_tag(x_706, 7); + x_951 = x_931; + lean_ctor_set_tag(x_951, 7); } -lean_ctor_set(x_706, 0, x_704); -lean_ctor_set(x_706, 1, x_705); -if (lean_is_scalar(x_664)) { - x_707 = lean_alloc_ctor(7, 2, 0); -} else { - x_707 = x_664; - lean_ctor_set_tag(x_707, 7); -} -lean_ctor_set(x_707, 0, x_706); -lean_ctor_set(x_707, 1, x_701); -x_708 = l_Lean_indentD(x_707); -x_709 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Term_elabCalcSteps___spec__1___closed__4; -if (lean_is_scalar(x_661)) { - x_710 = lean_alloc_ctor(7, 2, 0); -} else { - x_710 = x_661; - lean_ctor_set_tag(x_710, 7); -} -lean_ctor_set(x_710, 0, x_709); -lean_ctor_set(x_710, 1, x_708); -x_711 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__3___closed__4; -if (lean_is_scalar(x_659)) { - x_712 = lean_alloc_ctor(7, 2, 0); -} else { - x_712 = x_659; - lean_ctor_set_tag(x_712, 7); -} -lean_ctor_set(x_712, 0, x_710); -lean_ctor_set(x_712, 1, x_711); -x_713 = l_Lean_MessageData_ofExpr(x_677); -x_714 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_714, 0, x_701); -lean_ctor_set(x_714, 1, x_713); -x_715 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_715, 0, x_714); -lean_ctor_set(x_715, 1, x_703); -x_716 = l_Lean_MessageData_ofExpr(x_695); -x_717 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_717, 0, x_715); -lean_ctor_set(x_717, 1, x_716); -x_718 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_718, 0, x_717); -lean_ctor_set(x_718, 1, x_701); -x_719 = l_Lean_indentD(x_718); -x_720 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_720, 0, x_712); -lean_ctor_set(x_720, 1, x_719); -x_721 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_721, 0, x_720); -lean_ctor_set(x_721, 1, x_701); -if (x_699 == 0) -{ -lean_object* x_722; lean_object* x_723; lean_object* x_724; uint8_t x_725; lean_object* x_726; lean_object* x_727; uint8_t x_728; lean_object* x_729; lean_object* x_730; -x_722 = l_Lean_Elab_Term_instInhabitedCalcStepView; -x_723 = l_outOfBounds___rarg(x_722); -x_724 = lean_ctor_get(x_723, 1); -lean_inc(x_724); -lean_dec(x_723); -x_725 = 2; +lean_ctor_set(x_951, 0, x_949); +lean_ctor_set(x_951, 1, x_950); +x_952 = l_Lean_MessageData_ofExpr(x_941); +if (lean_is_scalar(x_914)) { + x_953 = lean_alloc_ctor(7, 2, 0); +} else { + x_953 = x_914; + lean_ctor_set_tag(x_953, 7); +} +lean_ctor_set(x_953, 0, x_951); +lean_ctor_set(x_953, 1, x_952); +if (lean_is_scalar(x_911)) { + x_954 = lean_alloc_ctor(7, 2, 0); +} else { + x_954 = x_911; + lean_ctor_set_tag(x_954, 7); +} +lean_ctor_set(x_954, 0, x_953); +lean_ctor_set(x_954, 1, x_948); +x_955 = l_Lean_indentD(x_954); +x_956 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Term_elabCalcSteps___spec__1___closed__4; +if (lean_is_scalar(x_906)) { + x_957 = lean_alloc_ctor(7, 2, 0); +} else { + x_957 = x_906; + lean_ctor_set_tag(x_957, 7); +} +lean_ctor_set(x_957, 0, x_956); +lean_ctor_set(x_957, 1, x_955); +x_958 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__3___closed__4; +if (lean_is_scalar(x_899)) { + x_959 = lean_alloc_ctor(7, 2, 0); +} else { + x_959 = x_899; + lean_ctor_set_tag(x_959, 7); +} +lean_ctor_set(x_959, 0, x_957); +lean_ctor_set(x_959, 1, x_958); +x_960 = l_Lean_MessageData_ofExpr(x_930); +if (lean_is_scalar(x_896)) { + x_961 = lean_alloc_ctor(7, 2, 0); +} else { + x_961 = x_896; + lean_ctor_set_tag(x_961, 7); +} +lean_ctor_set(x_961, 0, x_948); +lean_ctor_set(x_961, 1, x_960); +if (lean_is_scalar(x_894)) { + x_962 = lean_alloc_ctor(7, 2, 0); +} else { + x_962 = x_894; + lean_ctor_set_tag(x_962, 7); +} +lean_ctor_set(x_962, 0, x_961); +lean_ctor_set(x_962, 1, x_950); +x_963 = l_Lean_MessageData_ofExpr(x_942); +x_964 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_964, 0, x_962); +lean_ctor_set(x_964, 1, x_963); +x_965 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_965, 0, x_964); +lean_ctor_set(x_965, 1, x_948); +x_966 = l_Lean_indentD(x_965); +x_967 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_967, 0, x_959); +lean_ctor_set(x_967, 1, x_966); +x_968 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_968, 0, x_967); +lean_ctor_set(x_968, 1, x_948); +if (x_946 == 0) +{ +lean_object* x_969; lean_object* x_970; lean_object* x_971; uint8_t x_972; lean_object* x_973; lean_object* x_974; uint8_t x_975; lean_object* x_976; lean_object* x_977; +x_969 = l_Lean_Elab_Term_instInhabitedCalcStepView; +x_970 = l_outOfBounds___rarg(x_969); +x_971 = lean_ctor_get(x_970, 1); +lean_inc(x_971); +lean_dec(x_970); +x_972 = 2; lean_inc(x_6); -x_726 = l_Lean_logAt___at_Lean_Elab_Term_reportUnsolvedGoals___spec__2(x_724, x_721, x_725, x_4, x_5, x_6, x_7, x_696); -lean_dec(x_724); -x_727 = lean_ctor_get(x_726, 1); -lean_inc(x_727); -lean_dec(x_726); -x_728 = 1; -x_729 = lean_box(0); -x_730 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__3(x_663, x_678, x_673, x_1, x_728, x_729, x_4, x_5, x_6, x_7, x_727); -return x_730; +x_973 = l_Lean_logAt___at_Lean_Elab_Term_reportUnsolvedGoals___spec__2(x_971, x_968, x_972, x_4, x_5, x_6, x_7, x_940); +lean_dec(x_971); +x_974 = lean_ctor_get(x_973, 1); +lean_inc(x_974); +lean_dec(x_973); +x_975 = 1; +x_976 = lean_box(0); +x_977 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__3(x_898, x_913, x_908, x_1, x_975, x_976, x_4, x_5, x_6, x_7, x_974); +return x_977; +} +else +{ +lean_object* x_978; lean_object* x_979; uint8_t x_980; lean_object* x_981; lean_object* x_982; uint8_t x_983; lean_object* x_984; lean_object* x_985; +x_978 = lean_array_fget(x_1, x_945); +x_979 = lean_ctor_get(x_978, 1); +lean_inc(x_979); +lean_dec(x_978); +x_980 = 2; +lean_inc(x_6); +x_981 = l_Lean_logAt___at_Lean_Elab_Term_reportUnsolvedGoals___spec__2(x_979, x_968, x_980, x_4, x_5, x_6, x_7, x_940); +lean_dec(x_979); +x_982 = lean_ctor_get(x_981, 1); +lean_inc(x_982); +lean_dec(x_981); +x_983 = 1; +x_984 = lean_box(0); +x_985 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__3(x_898, x_913, x_908, x_1, x_983, x_984, x_4, x_5, x_6, x_7, x_982); +return x_985; +} } else { -lean_object* x_731; lean_object* x_732; uint8_t x_733; lean_object* x_734; lean_object* x_735; uint8_t x_736; lean_object* x_737; lean_object* x_738; -x_731 = lean_array_fget(x_1, x_698); -x_732 = lean_ctor_get(x_731, 1); -lean_inc(x_732); -lean_dec(x_731); -x_733 = 2; -lean_inc(x_6); -x_734 = l_Lean_logAt___at_Lean_Elab_Term_reportUnsolvedGoals___spec__2(x_732, x_721, x_733, x_4, x_5, x_6, x_7, x_696); -lean_dec(x_732); -x_735 = lean_ctor_get(x_734, 1); -lean_inc(x_735); -lean_dec(x_734); -x_736 = 1; -x_737 = lean_box(0); -x_738 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__3(x_663, x_678, x_673, x_1, x_736, x_737, x_4, x_5, x_6, x_7, x_735); -return x_738; +lean_object* x_986; lean_object* x_987; lean_object* x_988; lean_object* x_989; +lean_dec(x_931); +lean_dec(x_930); +lean_dec(x_929); +lean_dec(x_914); +lean_dec(x_913); +lean_dec(x_911); +lean_dec(x_908); +lean_dec(x_906); +lean_dec(x_899); +lean_dec(x_898); +lean_dec(x_896); +lean_dec(x_894); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_986 = lean_ctor_get(x_938, 0); +lean_inc(x_986); +x_987 = lean_ctor_get(x_938, 1); +lean_inc(x_987); +if (lean_is_exclusive(x_938)) { + lean_ctor_release(x_938, 0); + lean_ctor_release(x_938, 1); + x_988 = x_938; +} else { + lean_dec_ref(x_938); + x_988 = lean_box(0); +} +if (lean_is_scalar(x_988)) { + x_989 = lean_alloc_ctor(1, 2, 0); +} else { + x_989 = x_988; +} +lean_ctor_set(x_989, 0, x_986); +lean_ctor_set(x_989, 1, x_987); +return x_989; } } else { -lean_object* x_739; lean_object* x_740; lean_object* x_741; lean_object* x_742; -lean_dec(x_692); -lean_dec(x_679); -lean_dec(x_678); -lean_dec(x_677); -lean_dec(x_676); -lean_dec(x_673); -lean_dec(x_671); -lean_dec(x_664); -lean_dec(x_663); -lean_dec(x_662); -lean_dec(x_661); -lean_dec(x_659); +lean_object* x_990; lean_object* x_991; lean_object* x_992; lean_object* x_993; +lean_dec(x_933); +lean_dec(x_931); +lean_dec(x_930); +lean_dec(x_929); +lean_dec(x_914); +lean_dec(x_913); +lean_dec(x_911); +lean_dec(x_908); +lean_dec(x_906); +lean_dec(x_899); +lean_dec(x_898); +lean_dec(x_896); +lean_dec(x_894); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_739 = lean_ctor_get(x_694, 0); -lean_inc(x_739); -x_740 = lean_ctor_get(x_694, 1); -lean_inc(x_740); -if (lean_is_exclusive(x_694)) { - lean_ctor_release(x_694, 0); - lean_ctor_release(x_694, 1); - x_741 = x_694; +x_990 = lean_ctor_get(x_935, 0); +lean_inc(x_990); +x_991 = lean_ctor_get(x_935, 1); +lean_inc(x_991); +if (lean_is_exclusive(x_935)) { + lean_ctor_release(x_935, 0); + lean_ctor_release(x_935, 1); + x_992 = x_935; } else { - lean_dec_ref(x_694); - x_741 = lean_box(0); + lean_dec_ref(x_935); + x_992 = lean_box(0); } -if (lean_is_scalar(x_741)) { - x_742 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_992)) { + x_993 = lean_alloc_ctor(1, 2, 0); } else { - x_742 = x_741; + x_993 = x_992; } -lean_ctor_set(x_742, 0, x_739); -lean_ctor_set(x_742, 1, x_740); -return x_742; +lean_ctor_set(x_993, 0, x_990); +lean_ctor_set(x_993, 1, x_991); +return x_993; } } else { -lean_object* x_743; lean_object* x_744; lean_object* x_745; lean_object* x_746; -lean_dec(x_679); -lean_dec(x_678); -lean_dec(x_677); -lean_dec(x_676); -lean_dec(x_673); -lean_dec(x_671); -lean_dec(x_664); -lean_dec(x_663); -lean_dec(x_662); -lean_dec(x_661); -lean_dec(x_659); +lean_object* x_994; lean_object* x_995; lean_object* x_996; lean_object* x_997; +lean_dec(x_931); +lean_dec(x_930); +lean_dec(x_929); +lean_dec(x_914); +lean_dec(x_913); +lean_dec(x_911); +lean_dec(x_908); +lean_dec(x_906); +lean_dec(x_899); +lean_dec(x_898); +lean_dec(x_896); +lean_dec(x_894); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_743 = lean_ctor_get(x_691, 0); -lean_inc(x_743); -x_744 = lean_ctor_get(x_691, 1); -lean_inc(x_744); -if (lean_is_exclusive(x_691)) { - lean_ctor_release(x_691, 0); - lean_ctor_release(x_691, 1); - x_745 = x_691; +x_994 = lean_ctor_get(x_932, 0); +lean_inc(x_994); +x_995 = lean_ctor_get(x_932, 1); +lean_inc(x_995); +if (lean_is_exclusive(x_932)) { + lean_ctor_release(x_932, 0); + lean_ctor_release(x_932, 1); + x_996 = x_932; } else { - lean_dec_ref(x_691); - x_745 = lean_box(0); + lean_dec_ref(x_932); + x_996 = lean_box(0); } -if (lean_is_scalar(x_745)) { - x_746 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_996)) { + x_997 = lean_alloc_ctor(1, 2, 0); } else { - x_746 = x_745; + x_997 = x_996; } -lean_ctor_set(x_746, 0, x_743); -lean_ctor_set(x_746, 1, x_744); -return x_746; +lean_ctor_set(x_997, 0, x_994); +lean_ctor_set(x_997, 1, x_995); +return x_997; } } else { -lean_object* x_747; uint8_t x_748; lean_object* x_749; lean_object* x_750; -lean_dec(x_679); -lean_dec(x_677); -lean_dec(x_676); -lean_dec(x_671); -lean_dec(x_664); -lean_dec(x_662); -lean_dec(x_661); -lean_dec(x_659); -x_747 = lean_ctor_get(x_687, 1); -lean_inc(x_747); -lean_dec(x_687); -x_748 = 0; -x_749 = lean_box(0); -x_750 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__3(x_663, x_678, x_673, x_1, x_748, x_749, x_4, x_5, x_6, x_7, x_747); -return x_750; +lean_object* x_998; lean_object* x_999; lean_object* x_1000; lean_object* x_1001; +lean_dec(x_914); +lean_dec(x_913); +lean_dec(x_911); +lean_dec(x_908); +lean_dec(x_906); +lean_dec(x_899); +lean_dec(x_898); +lean_dec(x_896); +lean_dec(x_894); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_998 = lean_ctor_get(x_926, 0); +lean_inc(x_998); +x_999 = lean_ctor_get(x_926, 1); +lean_inc(x_999); +if (lean_is_exclusive(x_926)) { + lean_ctor_release(x_926, 0); + lean_ctor_release(x_926, 1); + x_1000 = x_926; +} else { + lean_dec_ref(x_926); + x_1000 = lean_box(0); +} +if (lean_is_scalar(x_1000)) { + x_1001 = lean_alloc_ctor(1, 2, 0); +} else { + x_1001 = x_1000; +} +lean_ctor_set(x_1001, 0, x_998); +lean_ctor_set(x_1001, 1, x_999); +return x_1001; } } else { -lean_object* x_751; lean_object* x_752; lean_object* x_753; lean_object* x_754; -lean_dec(x_679); -lean_dec(x_678); -lean_dec(x_677); -lean_dec(x_676); -lean_dec(x_673); -lean_dec(x_671); -lean_dec(x_664); -lean_dec(x_663); -lean_dec(x_662); -lean_dec(x_661); -lean_dec(x_659); +lean_object* x_1002; uint8_t x_1003; lean_object* x_1004; lean_object* x_1005; +lean_dec(x_914); +lean_dec(x_912); +lean_dec(x_911); +lean_dec(x_906); +lean_dec(x_899); +lean_dec(x_897); +lean_dec(x_896); +lean_dec(x_894); +x_1002 = lean_ctor_get(x_922, 1); +lean_inc(x_1002); +lean_dec(x_922); +x_1003 = 0; +x_1004 = lean_box(0); +x_1005 = l_Lean_Elab_Term_throwCalcFailure___rarg___lambda__3(x_898, x_913, x_908, x_1, x_1003, x_1004, x_4, x_5, x_6, x_7, x_1002); +return x_1005; +} +} +else +{ +lean_object* x_1006; lean_object* x_1007; lean_object* x_1008; lean_object* x_1009; +lean_dec(x_914); +lean_dec(x_913); +lean_dec(x_912); +lean_dec(x_911); +lean_dec(x_908); +lean_dec(x_906); +lean_dec(x_899); +lean_dec(x_898); +lean_dec(x_897); +lean_dec(x_896); +lean_dec(x_894); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_751 = lean_ctor_get(x_687, 0); -lean_inc(x_751); -x_752 = lean_ctor_get(x_687, 1); -lean_inc(x_752); -if (lean_is_exclusive(x_687)) { - lean_ctor_release(x_687, 0); - lean_ctor_release(x_687, 1); - x_753 = x_687; +x_1006 = lean_ctor_get(x_922, 0); +lean_inc(x_1006); +x_1007 = lean_ctor_get(x_922, 1); +lean_inc(x_1007); +if (lean_is_exclusive(x_922)) { + lean_ctor_release(x_922, 0); + lean_ctor_release(x_922, 1); + x_1008 = x_922; } else { - lean_dec_ref(x_687); - x_753 = lean_box(0); + lean_dec_ref(x_922); + x_1008 = lean_box(0); } -if (lean_is_scalar(x_753)) { - x_754 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_1008)) { + x_1009 = lean_alloc_ctor(1, 2, 0); } else { - x_754 = x_753; + x_1009 = x_1008; } -lean_ctor_set(x_754, 0, x_751); -lean_ctor_set(x_754, 1, x_752); -return x_754; +lean_ctor_set(x_1009, 0, x_1006); +lean_ctor_set(x_1009, 1, x_1007); +return x_1009; } } } else { -lean_object* x_755; lean_object* x_756; lean_object* x_757; lean_object* x_758; -lean_dec(x_679); -lean_dec(x_678); -lean_dec(x_677); -lean_dec(x_676); -lean_dec(x_673); -lean_dec(x_671); -lean_dec(x_664); -lean_dec(x_663); -lean_dec(x_662); -lean_dec(x_661); -lean_dec(x_659); -lean_dec(x_650); +lean_object* x_1010; lean_object* x_1011; lean_object* x_1012; lean_object* x_1013; +lean_dec(x_914); +lean_dec(x_913); +lean_dec(x_912); +lean_dec(x_911); +lean_dec(x_908); +lean_dec(x_906); +lean_dec(x_899); +lean_dec(x_898); +lean_dec(x_897); +lean_dec(x_896); +lean_dec(x_894); +lean_dec(x_885); 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_755 = lean_ctor_get(x_680, 0); -lean_inc(x_755); -x_756 = lean_ctor_get(x_680, 1); -lean_inc(x_756); -if (lean_is_exclusive(x_680)) { - lean_ctor_release(x_680, 0); - lean_ctor_release(x_680, 1); - x_757 = x_680; +x_1010 = lean_ctor_get(x_915, 0); +lean_inc(x_1010); +x_1011 = lean_ctor_get(x_915, 1); +lean_inc(x_1011); +if (lean_is_exclusive(x_915)) { + lean_ctor_release(x_915, 0); + lean_ctor_release(x_915, 1); + x_1012 = x_915; } else { - lean_dec_ref(x_680); - x_757 = lean_box(0); + lean_dec_ref(x_915); + x_1012 = lean_box(0); } -if (lean_is_scalar(x_757)) { - x_758 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_1012)) { + x_1013 = lean_alloc_ctor(1, 2, 0); } else { - x_758 = x_757; + x_1013 = x_1012; } -lean_ctor_set(x_758, 0, x_755); -lean_ctor_set(x_758, 1, x_756); -return x_758; +lean_ctor_set(x_1013, 0, x_1010); +lean_ctor_set(x_1013, 1, x_1011); +return x_1013; } } } @@ -23466,30 +25349,30 @@ return x_758; } else { -uint8_t x_759; +uint8_t x_1014; 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_759 = !lean_is_exclusive(x_9); -if (x_759 == 0) +x_1014 = !lean_is_exclusive(x_9); +if (x_1014 == 0) { return x_9; } else { -lean_object* x_760; lean_object* x_761; lean_object* x_762; -x_760 = lean_ctor_get(x_9, 0); -x_761 = lean_ctor_get(x_9, 1); -lean_inc(x_761); -lean_inc(x_760); +lean_object* x_1015; lean_object* x_1016; lean_object* x_1017; +x_1015 = lean_ctor_get(x_9, 0); +x_1016 = lean_ctor_get(x_9, 1); +lean_inc(x_1016); +lean_inc(x_1015); lean_dec(x_9); -x_762 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_762, 0, x_760); -lean_ctor_set(x_762, 1, x_761); -return x_762; +x_1017 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1017, 0, x_1015); +lean_ctor_set(x_1017, 1, x_1016); +return x_1017; } } } diff --git a/stage0/stdlib/Lean/Elab/CheckTactic.c b/stage0/stdlib/Lean/Elab/CheckTactic.c index 5367548f2737..807ba7560183 100644 --- a/stage0/stdlib/Lean/Elab/CheckTactic.c +++ b/stage0/stdlib/Lean/Elab/CheckTactic.c @@ -112,6 +112,7 @@ static lean_object* l_Lean_Elab_CheckTactic_expandCheckSimp___closed__7; static lean_object* l___regBuiltin_Lean_Elab_CheckTactic_expandCheckSimp_declRange__1___closed__6; lean_object* l_Lean_Syntax_node4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_CheckTactic_expandCheckSimp_declRange__1___closed__2; +lean_object* l_Lean_Meta_addPPExplicitToExposeDiff(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_CheckTactic_elabCheckTacticFailure_declRange__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_CheckTactic_elabCheckTactic___spec__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__2; @@ -573,6 +574,7 @@ lean_inc(x_56); if (lean_obj_tag(x_56) == 0) { lean_object* x_57; uint8_t x_58; +lean_free_object(x_32); lean_dec(x_3); x_57 = lean_ctor_get(x_43, 1); lean_inc(x_57); @@ -669,1399 +671,1720 @@ x_91 = lean_unbox(x_90); lean_dec(x_90); 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_92; lean_object* x_93; x_92 = lean_ctor_get(x_89, 1); lean_inc(x_92); lean_dec(x_89); -x_93 = l_Lean_indentExpr(x_69); -x_94 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__15; +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +x_93 = l_Lean_Meta_addPPExplicitToExposeDiff(x_69, x_38, x_8, x_9, x_10, x_11, x_92); +if (lean_obj_tag(x_93) == 0) +{ +lean_object* x_94; lean_object* x_95; uint8_t x_96; +x_94 = lean_ctor_get(x_93, 0); +lean_inc(x_94); +x_95 = lean_ctor_get(x_93, 1); +lean_inc(x_95); +lean_dec(x_93); +x_96 = !lean_is_exclusive(x_94); +if (x_96 == 0) +{ +lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; +x_97 = lean_ctor_get(x_94, 0); +x_98 = lean_ctor_get(x_94, 1); +x_99 = l_Lean_indentExpr(x_97); +x_100 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__15; +lean_ctor_set_tag(x_94, 7); +lean_ctor_set(x_94, 1, x_99); +lean_ctor_set(x_94, 0, x_100); +x_101 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__17; lean_ctor_set_tag(x_65, 7); -lean_ctor_set(x_65, 1, x_93); +lean_ctor_set(x_65, 1, x_101); lean_ctor_set(x_65, 0, x_94); -x_95 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__17; +x_102 = l_Lean_indentExpr(x_98); lean_ctor_set_tag(x_46, 7); -lean_ctor_set(x_46, 1, x_95); +lean_ctor_set(x_46, 1, x_102); lean_ctor_set(x_46, 0, x_65); -x_96 = l_Lean_indentExpr(x_38); +x_103 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__8; lean_ctor_set_tag(x_44, 7); -lean_ctor_set(x_44, 1, x_96); -x_97 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__8; -lean_ctor_set_tag(x_32, 7); -lean_ctor_set(x_32, 1, x_97); -lean_ctor_set(x_32, 0, x_44); -x_98 = l_Lean_throwErrorAt___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_throwStuckAtUniverseCnstr___spec__12(x_4, x_32, x_6, x_7, x_8, x_9, x_10, x_11, x_92); +lean_ctor_set(x_44, 1, x_103); +x_104 = l_Lean_throwErrorAt___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_throwStuckAtUniverseCnstr___spec__12(x_4, x_44, x_6, x_7, x_8, x_9, x_10, x_11, x_95); lean_dec(x_11); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -return x_98; +return x_104; +} +else +{ +lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; +x_105 = lean_ctor_get(x_94, 0); +x_106 = lean_ctor_get(x_94, 1); +lean_inc(x_106); +lean_inc(x_105); +lean_dec(x_94); +x_107 = l_Lean_indentExpr(x_105); +x_108 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__15; +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 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__17; +lean_ctor_set_tag(x_65, 7); +lean_ctor_set(x_65, 1, x_110); +lean_ctor_set(x_65, 0, x_109); +x_111 = l_Lean_indentExpr(x_106); +lean_ctor_set_tag(x_46, 7); +lean_ctor_set(x_46, 1, x_111); +lean_ctor_set(x_46, 0, x_65); +x_112 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__8; +lean_ctor_set_tag(x_44, 7); +lean_ctor_set(x_44, 1, x_112); +x_113 = l_Lean_throwErrorAt___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_throwStuckAtUniverseCnstr___spec__12(x_4, x_44, x_6, x_7, x_8, x_9, x_10, x_11, x_95); +lean_dec(x_11); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +return x_113; +} } else { -uint8_t x_99; +uint8_t x_114; +lean_free_object(x_65); +lean_free_object(x_46); +lean_free_object(x_44); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +x_114 = !lean_is_exclusive(x_93); +if (x_114 == 0) +{ +return x_93; +} +else +{ +lean_object* x_115; lean_object* x_116; lean_object* x_117; +x_115 = lean_ctor_get(x_93, 0); +x_116 = lean_ctor_get(x_93, 1); +lean_inc(x_116); +lean_inc(x_115); +lean_dec(x_93); +x_117 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_117, 0, x_115); +lean_ctor_set(x_117, 1, x_116); +return x_117; +} +} +} +else +{ +uint8_t x_118; lean_free_object(x_65); lean_dec(x_69); lean_free_object(x_46); lean_free_object(x_44); lean_dec(x_38); -lean_free_object(x_32); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -x_99 = !lean_is_exclusive(x_89); -if (x_99 == 0) -{ -lean_object* x_100; lean_object* x_101; -x_100 = lean_ctor_get(x_89, 0); -lean_dec(x_100); -x_101 = lean_box(0); -lean_ctor_set(x_89, 0, x_101); +x_118 = !lean_is_exclusive(x_89); +if (x_118 == 0) +{ +lean_object* x_119; lean_object* x_120; +x_119 = lean_ctor_get(x_89, 0); +lean_dec(x_119); +x_120 = lean_box(0); +lean_ctor_set(x_89, 0, x_120); return x_89; } else { -lean_object* x_102; lean_object* x_103; lean_object* x_104; -x_102 = lean_ctor_get(x_89, 1); -lean_inc(x_102); +lean_object* x_121; lean_object* x_122; lean_object* x_123; +x_121 = lean_ctor_get(x_89, 1); +lean_inc(x_121); lean_dec(x_89); -x_103 = lean_box(0); -x_104 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_104, 0, x_103); -lean_ctor_set(x_104, 1, x_102); -return x_104; +x_122 = lean_box(0); +x_123 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_123, 0, x_122); +lean_ctor_set(x_123, 1, x_121); +return x_123; } } } else { -uint8_t x_105; +uint8_t x_124; lean_free_object(x_65); lean_dec(x_69); lean_free_object(x_46); lean_free_object(x_44); lean_dec(x_38); -lean_free_object(x_32); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -x_105 = !lean_is_exclusive(x_89); -if (x_105 == 0) +x_124 = !lean_is_exclusive(x_89); +if (x_124 == 0) { return x_89; } else { -lean_object* x_106; lean_object* x_107; lean_object* x_108; -x_106 = lean_ctor_get(x_89, 0); -x_107 = lean_ctor_get(x_89, 1); -lean_inc(x_107); -lean_inc(x_106); +lean_object* x_125; lean_object* x_126; lean_object* x_127; +x_125 = lean_ctor_get(x_89, 0); +x_126 = lean_ctor_get(x_89, 1); +lean_inc(x_126); +lean_inc(x_125); lean_dec(x_89); -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_127 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_127, 0, x_125); +lean_ctor_set(x_127, 1, x_126); +return x_127; } } } else { -uint8_t x_109; uint8_t x_110; uint8_t x_111; 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; uint8_t x_123; uint8_t x_124; uint8_t x_125; uint8_t x_126; uint8_t x_127; lean_object* x_128; uint64_t x_129; uint64_t x_130; uint64_t x_131; uint64_t x_132; uint64_t x_133; lean_object* x_134; lean_object* x_135; -x_109 = lean_ctor_get_uint8(x_8, sizeof(void*)*7 + 9); -x_110 = lean_ctor_get_uint8(x_8, sizeof(void*)*7 + 10); -x_111 = lean_ctor_get_uint8(x_66, 0); -x_112 = lean_ctor_get_uint8(x_66, 1); -x_113 = lean_ctor_get_uint8(x_66, 2); -x_114 = lean_ctor_get_uint8(x_66, 3); -x_115 = lean_ctor_get_uint8(x_66, 4); -x_116 = lean_ctor_get_uint8(x_66, 5); -x_117 = lean_ctor_get_uint8(x_66, 6); -x_118 = lean_ctor_get_uint8(x_66, 7); -x_119 = lean_ctor_get_uint8(x_66, 8); -x_120 = lean_ctor_get_uint8(x_66, 10); -x_121 = lean_ctor_get_uint8(x_66, 11); -x_122 = lean_ctor_get_uint8(x_66, 12); -x_123 = lean_ctor_get_uint8(x_66, 13); -x_124 = lean_ctor_get_uint8(x_66, 14); -x_125 = lean_ctor_get_uint8(x_66, 15); -x_126 = lean_ctor_get_uint8(x_66, 16); +uint8_t x_128; uint8_t x_129; uint8_t x_130; uint8_t x_131; uint8_t x_132; uint8_t x_133; uint8_t x_134; uint8_t x_135; uint8_t x_136; uint8_t 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; lean_object* x_147; uint64_t x_148; uint64_t x_149; uint64_t x_150; uint64_t x_151; uint64_t x_152; lean_object* x_153; lean_object* x_154; +x_128 = lean_ctor_get_uint8(x_8, sizeof(void*)*7 + 9); +x_129 = lean_ctor_get_uint8(x_8, sizeof(void*)*7 + 10); +x_130 = lean_ctor_get_uint8(x_66, 0); +x_131 = lean_ctor_get_uint8(x_66, 1); +x_132 = lean_ctor_get_uint8(x_66, 2); +x_133 = lean_ctor_get_uint8(x_66, 3); +x_134 = lean_ctor_get_uint8(x_66, 4); +x_135 = lean_ctor_get_uint8(x_66, 5); +x_136 = lean_ctor_get_uint8(x_66, 6); +x_137 = lean_ctor_get_uint8(x_66, 7); +x_138 = lean_ctor_get_uint8(x_66, 8); +x_139 = lean_ctor_get_uint8(x_66, 10); +x_140 = lean_ctor_get_uint8(x_66, 11); +x_141 = lean_ctor_get_uint8(x_66, 12); +x_142 = lean_ctor_get_uint8(x_66, 13); +x_143 = lean_ctor_get_uint8(x_66, 14); +x_144 = lean_ctor_get_uint8(x_66, 15); +x_145 = lean_ctor_get_uint8(x_66, 16); lean_dec(x_66); -x_127 = 2; -x_128 = lean_alloc_ctor(0, 0, 17); -lean_ctor_set_uint8(x_128, 0, x_111); -lean_ctor_set_uint8(x_128, 1, x_112); -lean_ctor_set_uint8(x_128, 2, x_113); -lean_ctor_set_uint8(x_128, 3, x_114); -lean_ctor_set_uint8(x_128, 4, x_115); -lean_ctor_set_uint8(x_128, 5, x_116); -lean_ctor_set_uint8(x_128, 6, x_117); -lean_ctor_set_uint8(x_128, 7, x_118); -lean_ctor_set_uint8(x_128, 8, x_119); -lean_ctor_set_uint8(x_128, 9, x_127); -lean_ctor_set_uint8(x_128, 10, x_120); -lean_ctor_set_uint8(x_128, 11, x_121); -lean_ctor_set_uint8(x_128, 12, x_122); -lean_ctor_set_uint8(x_128, 13, x_123); -lean_ctor_set_uint8(x_128, 14, x_124); -lean_ctor_set_uint8(x_128, 15, x_125); -lean_ctor_set_uint8(x_128, 16, x_126); -x_129 = 2; -x_130 = lean_uint64_shift_right(x_71, x_129); -x_131 = lean_uint64_shift_left(x_130, x_129); -x_132 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__13; -x_133 = lean_uint64_lor(x_131, x_132); -x_134 = lean_alloc_ctor(0, 7, 11); -lean_ctor_set(x_134, 0, x_128); -lean_ctor_set(x_134, 1, x_73); -lean_ctor_set(x_134, 2, x_74); -lean_ctor_set(x_134, 3, x_75); -lean_ctor_set(x_134, 4, x_76); -lean_ctor_set(x_134, 5, x_77); -lean_ctor_set(x_134, 6, x_78); -lean_ctor_set_uint64(x_134, sizeof(void*)*7, x_133); -lean_ctor_set_uint8(x_134, sizeof(void*)*7 + 8, x_72); -lean_ctor_set_uint8(x_134, sizeof(void*)*7 + 9, x_109); -lean_ctor_set_uint8(x_134, sizeof(void*)*7 + 10, x_110); +x_146 = 2; +x_147 = lean_alloc_ctor(0, 0, 17); +lean_ctor_set_uint8(x_147, 0, x_130); +lean_ctor_set_uint8(x_147, 1, x_131); +lean_ctor_set_uint8(x_147, 2, x_132); +lean_ctor_set_uint8(x_147, 3, x_133); +lean_ctor_set_uint8(x_147, 4, x_134); +lean_ctor_set_uint8(x_147, 5, x_135); +lean_ctor_set_uint8(x_147, 6, x_136); +lean_ctor_set_uint8(x_147, 7, x_137); +lean_ctor_set_uint8(x_147, 8, x_138); +lean_ctor_set_uint8(x_147, 9, x_146); +lean_ctor_set_uint8(x_147, 10, x_139); +lean_ctor_set_uint8(x_147, 11, x_140); +lean_ctor_set_uint8(x_147, 12, x_141); +lean_ctor_set_uint8(x_147, 13, x_142); +lean_ctor_set_uint8(x_147, 14, x_143); +lean_ctor_set_uint8(x_147, 15, x_144); +lean_ctor_set_uint8(x_147, 16, x_145); +x_148 = 2; +x_149 = lean_uint64_shift_right(x_71, x_148); +x_150 = lean_uint64_shift_left(x_149, x_148); +x_151 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__13; +x_152 = lean_uint64_lor(x_150, x_151); +x_153 = lean_alloc_ctor(0, 7, 11); +lean_ctor_set(x_153, 0, x_147); +lean_ctor_set(x_153, 1, x_73); +lean_ctor_set(x_153, 2, x_74); +lean_ctor_set(x_153, 3, x_75); +lean_ctor_set(x_153, 4, x_76); +lean_ctor_set(x_153, 5, x_77); +lean_ctor_set(x_153, 6, x_78); +lean_ctor_set_uint64(x_153, sizeof(void*)*7, x_152); +lean_ctor_set_uint8(x_153, sizeof(void*)*7 + 8, x_72); +lean_ctor_set_uint8(x_153, sizeof(void*)*7 + 9, x_128); +lean_ctor_set_uint8(x_153, sizeof(void*)*7 + 10, x_129); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_38); lean_inc(x_69); -x_135 = l_Lean_Meta_isExprDefEq(x_69, x_38, x_134, x_9, x_10, x_11, x_67); -if (lean_obj_tag(x_135) == 0) -{ -lean_object* x_136; uint8_t x_137; -x_136 = lean_ctor_get(x_135, 0); -lean_inc(x_136); -x_137 = lean_unbox(x_136); -lean_dec(x_136); -if (x_137 == 0) -{ -lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; -x_138 = lean_ctor_get(x_135, 1); -lean_inc(x_138); -lean_dec(x_135); -x_139 = l_Lean_indentExpr(x_69); -x_140 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__15; +x_154 = l_Lean_Meta_isExprDefEq(x_69, x_38, x_153, x_9, x_10, x_11, x_67); +if (lean_obj_tag(x_154) == 0) +{ +lean_object* x_155; uint8_t x_156; +x_155 = lean_ctor_get(x_154, 0); +lean_inc(x_155); +x_156 = lean_unbox(x_155); +lean_dec(x_155); +if (x_156 == 0) +{ +lean_object* x_157; lean_object* x_158; +x_157 = lean_ctor_get(x_154, 1); +lean_inc(x_157); +lean_dec(x_154); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +x_158 = l_Lean_Meta_addPPExplicitToExposeDiff(x_69, x_38, x_8, x_9, x_10, x_11, x_157); +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_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; +x_159 = lean_ctor_get(x_158, 0); +lean_inc(x_159); +x_160 = lean_ctor_get(x_158, 1); +lean_inc(x_160); +lean_dec(x_158); +x_161 = lean_ctor_get(x_159, 0); +lean_inc(x_161); +x_162 = lean_ctor_get(x_159, 1); +lean_inc(x_162); +if (lean_is_exclusive(x_159)) { + lean_ctor_release(x_159, 0); + lean_ctor_release(x_159, 1); + x_163 = x_159; +} else { + lean_dec_ref(x_159); + x_163 = lean_box(0); +} +x_164 = l_Lean_indentExpr(x_161); +x_165 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__15; +if (lean_is_scalar(x_163)) { + x_166 = lean_alloc_ctor(7, 2, 0); +} else { + x_166 = x_163; + lean_ctor_set_tag(x_166, 7); +} +lean_ctor_set(x_166, 0, x_165); +lean_ctor_set(x_166, 1, x_164); +x_167 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__17; lean_ctor_set_tag(x_65, 7); -lean_ctor_set(x_65, 1, x_139); -lean_ctor_set(x_65, 0, x_140); -x_141 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__17; +lean_ctor_set(x_65, 1, x_167); +lean_ctor_set(x_65, 0, x_166); +x_168 = l_Lean_indentExpr(x_162); lean_ctor_set_tag(x_46, 7); -lean_ctor_set(x_46, 1, x_141); +lean_ctor_set(x_46, 1, x_168); lean_ctor_set(x_46, 0, x_65); -x_142 = l_Lean_indentExpr(x_38); +x_169 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__8; lean_ctor_set_tag(x_44, 7); -lean_ctor_set(x_44, 1, x_142); -x_143 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__8; -lean_ctor_set_tag(x_32, 7); -lean_ctor_set(x_32, 1, x_143); -lean_ctor_set(x_32, 0, x_44); -x_144 = l_Lean_throwErrorAt___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_throwStuckAtUniverseCnstr___spec__12(x_4, x_32, x_6, x_7, x_8, x_9, x_10, x_11, x_138); +lean_ctor_set(x_44, 1, x_169); +x_170 = l_Lean_throwErrorAt___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_throwStuckAtUniverseCnstr___spec__12(x_4, x_44, x_6, x_7, x_8, x_9, x_10, x_11, x_160); lean_dec(x_11); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -return x_144; +return x_170; +} +else +{ +lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; +lean_free_object(x_65); +lean_free_object(x_46); +lean_free_object(x_44); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +x_171 = lean_ctor_get(x_158, 0); +lean_inc(x_171); +x_172 = lean_ctor_get(x_158, 1); +lean_inc(x_172); +if (lean_is_exclusive(x_158)) { + lean_ctor_release(x_158, 0); + lean_ctor_release(x_158, 1); + x_173 = x_158; +} else { + lean_dec_ref(x_158); + x_173 = lean_box(0); +} +if (lean_is_scalar(x_173)) { + x_174 = lean_alloc_ctor(1, 2, 0); +} else { + x_174 = x_173; +} +lean_ctor_set(x_174, 0, x_171); +lean_ctor_set(x_174, 1, x_172); +return x_174; +} } else { -lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; +lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_free_object(x_65); lean_dec(x_69); lean_free_object(x_46); lean_free_object(x_44); lean_dec(x_38); -lean_free_object(x_32); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -x_145 = lean_ctor_get(x_135, 1); -lean_inc(x_145); -if (lean_is_exclusive(x_135)) { - lean_ctor_release(x_135, 0); - lean_ctor_release(x_135, 1); - x_146 = x_135; +x_175 = lean_ctor_get(x_154, 1); +lean_inc(x_175); +if (lean_is_exclusive(x_154)) { + lean_ctor_release(x_154, 0); + lean_ctor_release(x_154, 1); + x_176 = x_154; } else { - lean_dec_ref(x_135); - x_146 = lean_box(0); + lean_dec_ref(x_154); + x_176 = lean_box(0); } -x_147 = lean_box(0); -if (lean_is_scalar(x_146)) { - x_148 = lean_alloc_ctor(0, 2, 0); +x_177 = lean_box(0); +if (lean_is_scalar(x_176)) { + x_178 = lean_alloc_ctor(0, 2, 0); } else { - x_148 = x_146; + x_178 = x_176; } -lean_ctor_set(x_148, 0, x_147); -lean_ctor_set(x_148, 1, x_145); -return x_148; +lean_ctor_set(x_178, 0, x_177); +lean_ctor_set(x_178, 1, x_175); +return x_178; } } else { -lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; +lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_free_object(x_65); lean_dec(x_69); lean_free_object(x_46); lean_free_object(x_44); lean_dec(x_38); -lean_free_object(x_32); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -x_149 = lean_ctor_get(x_135, 0); -lean_inc(x_149); -x_150 = lean_ctor_get(x_135, 1); -lean_inc(x_150); -if (lean_is_exclusive(x_135)) { - lean_ctor_release(x_135, 0); - lean_ctor_release(x_135, 1); - x_151 = x_135; +x_179 = lean_ctor_get(x_154, 0); +lean_inc(x_179); +x_180 = lean_ctor_get(x_154, 1); +lean_inc(x_180); +if (lean_is_exclusive(x_154)) { + lean_ctor_release(x_154, 0); + lean_ctor_release(x_154, 1); + x_181 = x_154; } else { - lean_dec_ref(x_135); - x_151 = lean_box(0); + lean_dec_ref(x_154); + x_181 = lean_box(0); } -if (lean_is_scalar(x_151)) { - x_152 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_181)) { + x_182 = lean_alloc_ctor(1, 2, 0); } else { - x_152 = x_151; + x_182 = x_181; } -lean_ctor_set(x_152, 0, x_149); -lean_ctor_set(x_152, 1, x_150); -return x_152; +lean_ctor_set(x_182, 0, x_179); +lean_ctor_set(x_182, 1, x_180); +return x_182; } } } else { -lean_object* x_153; uint64_t x_154; uint8_t 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; uint8_t x_162; uint8_t x_163; uint8_t x_164; uint8_t x_165; uint8_t x_166; uint8_t x_167; uint8_t x_168; 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; lean_object* x_180; uint8_t x_181; lean_object* x_182; uint64_t x_183; uint64_t x_184; uint64_t x_185; uint64_t x_186; uint64_t x_187; lean_object* x_188; lean_object* x_189; -x_153 = lean_ctor_get(x_65, 0); -lean_inc(x_153); +lean_object* x_183; uint64_t x_184; uint8_t x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; uint8_t x_192; uint8_t x_193; uint8_t x_194; uint8_t x_195; uint8_t x_196; uint8_t x_197; uint8_t x_198; uint8_t x_199; uint8_t x_200; uint8_t x_201; uint8_t x_202; uint8_t x_203; uint8_t x_204; uint8_t x_205; uint8_t x_206; uint8_t x_207; uint8_t x_208; uint8_t x_209; lean_object* x_210; uint8_t x_211; lean_object* x_212; uint64_t x_213; uint64_t x_214; uint64_t x_215; uint64_t x_216; uint64_t x_217; lean_object* x_218; lean_object* x_219; +x_183 = lean_ctor_get(x_65, 0); +lean_inc(x_183); lean_dec(x_65); -x_154 = lean_ctor_get_uint64(x_8, sizeof(void*)*7); -x_155 = lean_ctor_get_uint8(x_8, sizeof(void*)*7 + 8); -x_156 = lean_ctor_get(x_8, 1); -lean_inc(x_156); -x_157 = lean_ctor_get(x_8, 2); -lean_inc(x_157); -x_158 = lean_ctor_get(x_8, 3); -lean_inc(x_158); -x_159 = lean_ctor_get(x_8, 4); -lean_inc(x_159); -x_160 = lean_ctor_get(x_8, 5); -lean_inc(x_160); -x_161 = lean_ctor_get(x_8, 6); -lean_inc(x_161); -x_162 = lean_ctor_get_uint8(x_8, sizeof(void*)*7 + 9); -x_163 = lean_ctor_get_uint8(x_8, sizeof(void*)*7 + 10); -x_164 = lean_ctor_get_uint8(x_66, 0); -x_165 = lean_ctor_get_uint8(x_66, 1); -x_166 = lean_ctor_get_uint8(x_66, 2); -x_167 = lean_ctor_get_uint8(x_66, 3); -x_168 = lean_ctor_get_uint8(x_66, 4); -x_169 = lean_ctor_get_uint8(x_66, 5); -x_170 = lean_ctor_get_uint8(x_66, 6); -x_171 = lean_ctor_get_uint8(x_66, 7); -x_172 = lean_ctor_get_uint8(x_66, 8); -x_173 = lean_ctor_get_uint8(x_66, 10); -x_174 = lean_ctor_get_uint8(x_66, 11); -x_175 = lean_ctor_get_uint8(x_66, 12); -x_176 = lean_ctor_get_uint8(x_66, 13); -x_177 = lean_ctor_get_uint8(x_66, 14); -x_178 = lean_ctor_get_uint8(x_66, 15); -x_179 = lean_ctor_get_uint8(x_66, 16); +x_184 = lean_ctor_get_uint64(x_8, sizeof(void*)*7); +x_185 = lean_ctor_get_uint8(x_8, sizeof(void*)*7 + 8); +x_186 = lean_ctor_get(x_8, 1); +lean_inc(x_186); +x_187 = lean_ctor_get(x_8, 2); +lean_inc(x_187); +x_188 = lean_ctor_get(x_8, 3); +lean_inc(x_188); +x_189 = lean_ctor_get(x_8, 4); +lean_inc(x_189); +x_190 = lean_ctor_get(x_8, 5); +lean_inc(x_190); +x_191 = lean_ctor_get(x_8, 6); +lean_inc(x_191); +x_192 = lean_ctor_get_uint8(x_8, sizeof(void*)*7 + 9); +x_193 = lean_ctor_get_uint8(x_8, sizeof(void*)*7 + 10); +x_194 = lean_ctor_get_uint8(x_66, 0); +x_195 = lean_ctor_get_uint8(x_66, 1); +x_196 = lean_ctor_get_uint8(x_66, 2); +x_197 = lean_ctor_get_uint8(x_66, 3); +x_198 = lean_ctor_get_uint8(x_66, 4); +x_199 = lean_ctor_get_uint8(x_66, 5); +x_200 = lean_ctor_get_uint8(x_66, 6); +x_201 = lean_ctor_get_uint8(x_66, 7); +x_202 = lean_ctor_get_uint8(x_66, 8); +x_203 = lean_ctor_get_uint8(x_66, 10); +x_204 = lean_ctor_get_uint8(x_66, 11); +x_205 = lean_ctor_get_uint8(x_66, 12); +x_206 = lean_ctor_get_uint8(x_66, 13); +x_207 = lean_ctor_get_uint8(x_66, 14); +x_208 = lean_ctor_get_uint8(x_66, 15); +x_209 = lean_ctor_get_uint8(x_66, 16); if (lean_is_exclusive(x_66)) { - x_180 = x_66; + x_210 = x_66; } else { lean_dec_ref(x_66); - x_180 = lean_box(0); -} -x_181 = 2; -if (lean_is_scalar(x_180)) { - x_182 = lean_alloc_ctor(0, 0, 17); -} else { - x_182 = x_180; -} -lean_ctor_set_uint8(x_182, 0, x_164); -lean_ctor_set_uint8(x_182, 1, x_165); -lean_ctor_set_uint8(x_182, 2, x_166); -lean_ctor_set_uint8(x_182, 3, x_167); -lean_ctor_set_uint8(x_182, 4, x_168); -lean_ctor_set_uint8(x_182, 5, x_169); -lean_ctor_set_uint8(x_182, 6, x_170); -lean_ctor_set_uint8(x_182, 7, x_171); -lean_ctor_set_uint8(x_182, 8, x_172); -lean_ctor_set_uint8(x_182, 9, x_181); -lean_ctor_set_uint8(x_182, 10, x_173); -lean_ctor_set_uint8(x_182, 11, x_174); -lean_ctor_set_uint8(x_182, 12, x_175); -lean_ctor_set_uint8(x_182, 13, x_176); -lean_ctor_set_uint8(x_182, 14, x_177); -lean_ctor_set_uint8(x_182, 15, x_178); -lean_ctor_set_uint8(x_182, 16, x_179); -x_183 = 2; -x_184 = lean_uint64_shift_right(x_154, x_183); -x_185 = lean_uint64_shift_left(x_184, x_183); -x_186 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__13; -x_187 = lean_uint64_lor(x_185, x_186); -x_188 = lean_alloc_ctor(0, 7, 11); -lean_ctor_set(x_188, 0, x_182); -lean_ctor_set(x_188, 1, x_156); -lean_ctor_set(x_188, 2, x_157); -lean_ctor_set(x_188, 3, x_158); -lean_ctor_set(x_188, 4, x_159); -lean_ctor_set(x_188, 5, x_160); -lean_ctor_set(x_188, 6, x_161); -lean_ctor_set_uint64(x_188, sizeof(void*)*7, x_187); -lean_ctor_set_uint8(x_188, sizeof(void*)*7 + 8, x_155); -lean_ctor_set_uint8(x_188, sizeof(void*)*7 + 9, x_162); -lean_ctor_set_uint8(x_188, sizeof(void*)*7 + 10, x_163); + x_210 = lean_box(0); +} +x_211 = 2; +if (lean_is_scalar(x_210)) { + x_212 = lean_alloc_ctor(0, 0, 17); +} else { + x_212 = x_210; +} +lean_ctor_set_uint8(x_212, 0, x_194); +lean_ctor_set_uint8(x_212, 1, x_195); +lean_ctor_set_uint8(x_212, 2, x_196); +lean_ctor_set_uint8(x_212, 3, x_197); +lean_ctor_set_uint8(x_212, 4, x_198); +lean_ctor_set_uint8(x_212, 5, x_199); +lean_ctor_set_uint8(x_212, 6, x_200); +lean_ctor_set_uint8(x_212, 7, x_201); +lean_ctor_set_uint8(x_212, 8, x_202); +lean_ctor_set_uint8(x_212, 9, x_211); +lean_ctor_set_uint8(x_212, 10, x_203); +lean_ctor_set_uint8(x_212, 11, x_204); +lean_ctor_set_uint8(x_212, 12, x_205); +lean_ctor_set_uint8(x_212, 13, x_206); +lean_ctor_set_uint8(x_212, 14, x_207); +lean_ctor_set_uint8(x_212, 15, x_208); +lean_ctor_set_uint8(x_212, 16, x_209); +x_213 = 2; +x_214 = lean_uint64_shift_right(x_184, x_213); +x_215 = lean_uint64_shift_left(x_214, x_213); +x_216 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__13; +x_217 = lean_uint64_lor(x_215, x_216); +x_218 = lean_alloc_ctor(0, 7, 11); +lean_ctor_set(x_218, 0, x_212); +lean_ctor_set(x_218, 1, x_186); +lean_ctor_set(x_218, 2, x_187); +lean_ctor_set(x_218, 3, x_188); +lean_ctor_set(x_218, 4, x_189); +lean_ctor_set(x_218, 5, x_190); +lean_ctor_set(x_218, 6, x_191); +lean_ctor_set_uint64(x_218, sizeof(void*)*7, x_217); +lean_ctor_set_uint8(x_218, sizeof(void*)*7 + 8, x_185); +lean_ctor_set_uint8(x_218, sizeof(void*)*7 + 9, x_192); +lean_ctor_set_uint8(x_218, sizeof(void*)*7 + 10, x_193); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_38); -lean_inc(x_153); -x_189 = l_Lean_Meta_isExprDefEq(x_153, x_38, x_188, x_9, x_10, x_11, x_67); -if (lean_obj_tag(x_189) == 0) +lean_inc(x_183); +x_219 = l_Lean_Meta_isExprDefEq(x_183, x_38, x_218, x_9, x_10, x_11, x_67); +if (lean_obj_tag(x_219) == 0) { -lean_object* x_190; uint8_t x_191; -x_190 = lean_ctor_get(x_189, 0); -lean_inc(x_190); -x_191 = lean_unbox(x_190); -lean_dec(x_190); -if (x_191 == 0) -{ -lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; -x_192 = lean_ctor_get(x_189, 1); -lean_inc(x_192); -lean_dec(x_189); -x_193 = l_Lean_indentExpr(x_153); -x_194 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__15; -x_195 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_195, 0, x_194); -lean_ctor_set(x_195, 1, x_193); -x_196 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__17; +lean_object* x_220; uint8_t x_221; +x_220 = lean_ctor_get(x_219, 0); +lean_inc(x_220); +x_221 = lean_unbox(x_220); +lean_dec(x_220); +if (x_221 == 0) +{ +lean_object* x_222; lean_object* x_223; +x_222 = lean_ctor_get(x_219, 1); +lean_inc(x_222); +lean_dec(x_219); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +x_223 = l_Lean_Meta_addPPExplicitToExposeDiff(x_183, x_38, x_8, x_9, x_10, x_11, x_222); +if (lean_obj_tag(x_223) == 0) +{ +lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; +x_224 = lean_ctor_get(x_223, 0); +lean_inc(x_224); +x_225 = lean_ctor_get(x_223, 1); +lean_inc(x_225); +lean_dec(x_223); +x_226 = lean_ctor_get(x_224, 0); +lean_inc(x_226); +x_227 = lean_ctor_get(x_224, 1); +lean_inc(x_227); +if (lean_is_exclusive(x_224)) { + lean_ctor_release(x_224, 0); + lean_ctor_release(x_224, 1); + x_228 = x_224; +} else { + lean_dec_ref(x_224); + x_228 = lean_box(0); +} +x_229 = l_Lean_indentExpr(x_226); +x_230 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__15; +if (lean_is_scalar(x_228)) { + x_231 = lean_alloc_ctor(7, 2, 0); +} else { + x_231 = x_228; + lean_ctor_set_tag(x_231, 7); +} +lean_ctor_set(x_231, 0, x_230); +lean_ctor_set(x_231, 1, x_229); +x_232 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__17; +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_indentExpr(x_227); lean_ctor_set_tag(x_46, 7); -lean_ctor_set(x_46, 1, x_196); -lean_ctor_set(x_46, 0, x_195); -x_197 = l_Lean_indentExpr(x_38); +lean_ctor_set(x_46, 1, x_234); +lean_ctor_set(x_46, 0, x_233); +x_235 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__8; lean_ctor_set_tag(x_44, 7); -lean_ctor_set(x_44, 1, x_197); -x_198 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__8; -lean_ctor_set_tag(x_32, 7); -lean_ctor_set(x_32, 1, x_198); -lean_ctor_set(x_32, 0, x_44); -x_199 = l_Lean_throwErrorAt___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_throwStuckAtUniverseCnstr___spec__12(x_4, x_32, x_6, x_7, x_8, x_9, x_10, x_11, x_192); +lean_ctor_set(x_44, 1, x_235); +x_236 = l_Lean_throwErrorAt___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_throwStuckAtUniverseCnstr___spec__12(x_4, x_44, x_6, x_7, x_8, x_9, x_10, x_11, x_225); lean_dec(x_11); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -return x_199; +return x_236; } else { -lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; -lean_dec(x_153); +lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; +lean_free_object(x_46); +lean_free_object(x_44); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +x_237 = lean_ctor_get(x_223, 0); +lean_inc(x_237); +x_238 = lean_ctor_get(x_223, 1); +lean_inc(x_238); +if (lean_is_exclusive(x_223)) { + lean_ctor_release(x_223, 0); + lean_ctor_release(x_223, 1); + x_239 = x_223; +} else { + lean_dec_ref(x_223); + x_239 = lean_box(0); +} +if (lean_is_scalar(x_239)) { + x_240 = lean_alloc_ctor(1, 2, 0); +} else { + x_240 = x_239; +} +lean_ctor_set(x_240, 0, x_237); +lean_ctor_set(x_240, 1, x_238); +return x_240; +} +} +else +{ +lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; +lean_dec(x_183); lean_free_object(x_46); lean_free_object(x_44); lean_dec(x_38); -lean_free_object(x_32); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -x_200 = lean_ctor_get(x_189, 1); -lean_inc(x_200); -if (lean_is_exclusive(x_189)) { - lean_ctor_release(x_189, 0); - lean_ctor_release(x_189, 1); - x_201 = x_189; +x_241 = lean_ctor_get(x_219, 1); +lean_inc(x_241); +if (lean_is_exclusive(x_219)) { + lean_ctor_release(x_219, 0); + lean_ctor_release(x_219, 1); + x_242 = x_219; } else { - lean_dec_ref(x_189); - x_201 = lean_box(0); + lean_dec_ref(x_219); + x_242 = lean_box(0); } -x_202 = lean_box(0); -if (lean_is_scalar(x_201)) { - x_203 = lean_alloc_ctor(0, 2, 0); +x_243 = lean_box(0); +if (lean_is_scalar(x_242)) { + x_244 = lean_alloc_ctor(0, 2, 0); } else { - x_203 = x_201; + x_244 = x_242; } -lean_ctor_set(x_203, 0, x_202); -lean_ctor_set(x_203, 1, x_200); -return x_203; +lean_ctor_set(x_244, 0, x_243); +lean_ctor_set(x_244, 1, x_241); +return x_244; } } else { -lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; -lean_dec(x_153); +lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; +lean_dec(x_183); lean_free_object(x_46); lean_free_object(x_44); lean_dec(x_38); -lean_free_object(x_32); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -x_204 = lean_ctor_get(x_189, 0); -lean_inc(x_204); -x_205 = lean_ctor_get(x_189, 1); -lean_inc(x_205); -if (lean_is_exclusive(x_189)) { - lean_ctor_release(x_189, 0); - lean_ctor_release(x_189, 1); - x_206 = x_189; +x_245 = lean_ctor_get(x_219, 0); +lean_inc(x_245); +x_246 = lean_ctor_get(x_219, 1); +lean_inc(x_246); +if (lean_is_exclusive(x_219)) { + lean_ctor_release(x_219, 0); + lean_ctor_release(x_219, 1); + x_247 = x_219; } else { - lean_dec_ref(x_189); - x_206 = lean_box(0); + lean_dec_ref(x_219); + x_247 = lean_box(0); } -if (lean_is_scalar(x_206)) { - x_207 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_247)) { + x_248 = lean_alloc_ctor(1, 2, 0); } else { - x_207 = x_206; + x_248 = x_247; } -lean_ctor_set(x_207, 0, x_204); -lean_ctor_set(x_207, 1, x_205); -return x_207; +lean_ctor_set(x_248, 0, x_245); +lean_ctor_set(x_248, 1, x_246); +return x_248; } } } else { -uint8_t x_208; +uint8_t x_249; lean_free_object(x_46); lean_free_object(x_44); lean_dec(x_38); -lean_free_object(x_32); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -x_208 = !lean_is_exclusive(x_64); -if (x_208 == 0) +x_249 = !lean_is_exclusive(x_64); +if (x_249 == 0) { return x_64; } else { -lean_object* x_209; lean_object* x_210; lean_object* x_211; -x_209 = lean_ctor_get(x_64, 0); -x_210 = lean_ctor_get(x_64, 1); -lean_inc(x_210); -lean_inc(x_209); +lean_object* x_250; lean_object* x_251; lean_object* x_252; +x_250 = lean_ctor_get(x_64, 0); +x_251 = lean_ctor_get(x_64, 1); +lean_inc(x_251); +lean_inc(x_250); lean_dec(x_64); -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; +x_252 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_252, 0, x_250); +lean_ctor_set(x_252, 1, x_251); +return x_252; } } } else { -uint8_t x_212; +uint8_t x_253; lean_free_object(x_46); lean_free_object(x_44); lean_dec(x_38); -lean_free_object(x_32); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -x_212 = !lean_is_exclusive(x_61); -if (x_212 == 0) +x_253 = !lean_is_exclusive(x_61); +if (x_253 == 0) { return x_61; } else { -lean_object* x_213; lean_object* x_214; lean_object* x_215; -x_213 = lean_ctor_get(x_61, 0); -x_214 = lean_ctor_get(x_61, 1); -lean_inc(x_214); -lean_inc(x_213); +lean_object* x_254; lean_object* x_255; lean_object* x_256; +x_254 = lean_ctor_get(x_61, 0); +x_255 = lean_ctor_get(x_61, 1); +lean_inc(x_255); +lean_inc(x_254); lean_dec(x_61); -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; +x_256 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_256, 0, x_254); +lean_ctor_set(x_256, 1, x_255); +return x_256; } } } else { -lean_object* x_216; lean_object* x_217; -x_216 = lean_ctor_get(x_46, 0); -lean_inc(x_216); +lean_object* x_257; lean_object* x_258; +x_257 = lean_ctor_get(x_46, 0); +lean_inc(x_257); lean_dec(x_46); -x_217 = l_Lean_MVarId_getType(x_216, x_8, x_9, x_10, x_11, x_57); -if (lean_obj_tag(x_217) == 0) -{ -lean_object* x_218; lean_object* x_219; lean_object* x_220; -x_218 = lean_ctor_get(x_217, 0); -lean_inc(x_218); -x_219 = lean_ctor_get(x_217, 1); -lean_inc(x_219); -lean_dec(x_217); +x_258 = l_Lean_MVarId_getType(x_257, x_8, x_9, x_10, x_11, x_57); +if (lean_obj_tag(x_258) == 0) +{ +lean_object* x_259; lean_object* x_260; lean_object* x_261; +x_259 = lean_ctor_get(x_258, 0); +lean_inc(x_259); +x_260 = lean_ctor_get(x_258, 1); +lean_inc(x_260); +lean_dec(x_258); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); -x_220 = l_Lean_Meta_CheckTactic_matchCheckGoalType(x_4, x_218, x_8, x_9, x_10, x_11, x_219); -if (lean_obj_tag(x_220) == 0) -{ -lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; uint64_t x_226; uint8_t 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; uint8_t x_235; uint8_t x_236; uint8_t x_237; uint8_t x_238; uint8_t x_239; uint8_t 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; lean_object* x_252; uint8_t x_253; lean_object* x_254; uint64_t x_255; uint64_t x_256; uint64_t x_257; uint64_t x_258; uint64_t x_259; lean_object* x_260; lean_object* x_261; -x_221 = lean_ctor_get(x_220, 0); -lean_inc(x_221); -x_222 = lean_ctor_get(x_8, 0); -lean_inc(x_222); -x_223 = lean_ctor_get(x_220, 1); -lean_inc(x_223); -lean_dec(x_220); -x_224 = lean_ctor_get(x_221, 0); -lean_inc(x_224); -if (lean_is_exclusive(x_221)) { - lean_ctor_release(x_221, 0); - lean_ctor_release(x_221, 1); - x_225 = x_221; -} else { - lean_dec_ref(x_221); - x_225 = lean_box(0); -} -x_226 = lean_ctor_get_uint64(x_8, sizeof(void*)*7); -x_227 = lean_ctor_get_uint8(x_8, sizeof(void*)*7 + 8); -x_228 = lean_ctor_get(x_8, 1); -lean_inc(x_228); -x_229 = lean_ctor_get(x_8, 2); -lean_inc(x_229); -x_230 = lean_ctor_get(x_8, 3); -lean_inc(x_230); -x_231 = lean_ctor_get(x_8, 4); -lean_inc(x_231); -x_232 = lean_ctor_get(x_8, 5); -lean_inc(x_232); -x_233 = lean_ctor_get(x_8, 6); -lean_inc(x_233); -x_234 = lean_ctor_get_uint8(x_8, sizeof(void*)*7 + 9); -x_235 = lean_ctor_get_uint8(x_8, sizeof(void*)*7 + 10); -x_236 = lean_ctor_get_uint8(x_222, 0); -x_237 = lean_ctor_get_uint8(x_222, 1); -x_238 = lean_ctor_get_uint8(x_222, 2); -x_239 = lean_ctor_get_uint8(x_222, 3); -x_240 = lean_ctor_get_uint8(x_222, 4); -x_241 = lean_ctor_get_uint8(x_222, 5); -x_242 = lean_ctor_get_uint8(x_222, 6); -x_243 = lean_ctor_get_uint8(x_222, 7); -x_244 = lean_ctor_get_uint8(x_222, 8); -x_245 = lean_ctor_get_uint8(x_222, 10); -x_246 = lean_ctor_get_uint8(x_222, 11); -x_247 = lean_ctor_get_uint8(x_222, 12); -x_248 = lean_ctor_get_uint8(x_222, 13); -x_249 = lean_ctor_get_uint8(x_222, 14); -x_250 = lean_ctor_get_uint8(x_222, 15); -x_251 = lean_ctor_get_uint8(x_222, 16); -if (lean_is_exclusive(x_222)) { - x_252 = x_222; -} else { - lean_dec_ref(x_222); - x_252 = lean_box(0); -} -x_253 = 2; -if (lean_is_scalar(x_252)) { - x_254 = lean_alloc_ctor(0, 0, 17); -} else { - x_254 = x_252; -} -lean_ctor_set_uint8(x_254, 0, x_236); -lean_ctor_set_uint8(x_254, 1, x_237); -lean_ctor_set_uint8(x_254, 2, x_238); -lean_ctor_set_uint8(x_254, 3, x_239); -lean_ctor_set_uint8(x_254, 4, x_240); -lean_ctor_set_uint8(x_254, 5, x_241); -lean_ctor_set_uint8(x_254, 6, x_242); -lean_ctor_set_uint8(x_254, 7, x_243); -lean_ctor_set_uint8(x_254, 8, x_244); -lean_ctor_set_uint8(x_254, 9, x_253); -lean_ctor_set_uint8(x_254, 10, x_245); -lean_ctor_set_uint8(x_254, 11, x_246); -lean_ctor_set_uint8(x_254, 12, x_247); -lean_ctor_set_uint8(x_254, 13, x_248); -lean_ctor_set_uint8(x_254, 14, x_249); -lean_ctor_set_uint8(x_254, 15, x_250); -lean_ctor_set_uint8(x_254, 16, x_251); -x_255 = 2; -x_256 = lean_uint64_shift_right(x_226, x_255); -x_257 = lean_uint64_shift_left(x_256, x_255); -x_258 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__13; -x_259 = lean_uint64_lor(x_257, x_258); -x_260 = lean_alloc_ctor(0, 7, 11); -lean_ctor_set(x_260, 0, x_254); -lean_ctor_set(x_260, 1, x_228); -lean_ctor_set(x_260, 2, x_229); -lean_ctor_set(x_260, 3, x_230); -lean_ctor_set(x_260, 4, x_231); -lean_ctor_set(x_260, 5, x_232); -lean_ctor_set(x_260, 6, x_233); -lean_ctor_set_uint64(x_260, sizeof(void*)*7, x_259); -lean_ctor_set_uint8(x_260, sizeof(void*)*7 + 8, x_227); -lean_ctor_set_uint8(x_260, sizeof(void*)*7 + 9, x_234); -lean_ctor_set_uint8(x_260, sizeof(void*)*7 + 10, x_235); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_38); -lean_inc(x_224); -x_261 = l_Lean_Meta_isExprDefEq(x_224, x_38, x_260, x_9, x_10, x_11, x_223); +x_261 = l_Lean_Meta_CheckTactic_matchCheckGoalType(x_4, x_259, x_8, x_9, x_10, x_11, x_260); if (lean_obj_tag(x_261) == 0) { -lean_object* x_262; uint8_t x_263; +lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; uint64_t x_267; uint8_t 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; uint8_t x_275; uint8_t x_276; uint8_t x_277; uint8_t x_278; uint8_t x_279; uint8_t x_280; uint8_t x_281; uint8_t x_282; uint8_t x_283; uint8_t x_284; uint8_t x_285; uint8_t x_286; uint8_t x_287; uint8_t x_288; uint8_t x_289; uint8_t x_290; uint8_t x_291; uint8_t x_292; lean_object* x_293; uint8_t x_294; lean_object* x_295; uint64_t x_296; uint64_t x_297; uint64_t x_298; uint64_t x_299; uint64_t x_300; lean_object* x_301; lean_object* x_302; x_262 = lean_ctor_get(x_261, 0); lean_inc(x_262); -x_263 = lean_unbox(x_262); -lean_dec(x_262); -if (x_263 == 0) -{ -lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; +x_263 = lean_ctor_get(x_8, 0); +lean_inc(x_263); x_264 = lean_ctor_get(x_261, 1); lean_inc(x_264); lean_dec(x_261); -x_265 = l_Lean_indentExpr(x_224); -x_266 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__15; -if (lean_is_scalar(x_225)) { - x_267 = lean_alloc_ctor(7, 2, 0); -} else { - x_267 = x_225; - lean_ctor_set_tag(x_267, 7); -} -lean_ctor_set(x_267, 0, x_266); -lean_ctor_set(x_267, 1, x_265); -x_268 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__17; -x_269 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_269, 0, x_267); -lean_ctor_set(x_269, 1, x_268); -x_270 = l_Lean_indentExpr(x_38); +x_265 = lean_ctor_get(x_262, 0); +lean_inc(x_265); +if (lean_is_exclusive(x_262)) { + lean_ctor_release(x_262, 0); + lean_ctor_release(x_262, 1); + x_266 = x_262; +} else { + lean_dec_ref(x_262); + x_266 = lean_box(0); +} +x_267 = lean_ctor_get_uint64(x_8, sizeof(void*)*7); +x_268 = lean_ctor_get_uint8(x_8, sizeof(void*)*7 + 8); +x_269 = lean_ctor_get(x_8, 1); +lean_inc(x_269); +x_270 = lean_ctor_get(x_8, 2); +lean_inc(x_270); +x_271 = lean_ctor_get(x_8, 3); +lean_inc(x_271); +x_272 = lean_ctor_get(x_8, 4); +lean_inc(x_272); +x_273 = lean_ctor_get(x_8, 5); +lean_inc(x_273); +x_274 = lean_ctor_get(x_8, 6); +lean_inc(x_274); +x_275 = lean_ctor_get_uint8(x_8, sizeof(void*)*7 + 9); +x_276 = lean_ctor_get_uint8(x_8, sizeof(void*)*7 + 10); +x_277 = lean_ctor_get_uint8(x_263, 0); +x_278 = lean_ctor_get_uint8(x_263, 1); +x_279 = lean_ctor_get_uint8(x_263, 2); +x_280 = lean_ctor_get_uint8(x_263, 3); +x_281 = lean_ctor_get_uint8(x_263, 4); +x_282 = lean_ctor_get_uint8(x_263, 5); +x_283 = lean_ctor_get_uint8(x_263, 6); +x_284 = lean_ctor_get_uint8(x_263, 7); +x_285 = lean_ctor_get_uint8(x_263, 8); +x_286 = lean_ctor_get_uint8(x_263, 10); +x_287 = lean_ctor_get_uint8(x_263, 11); +x_288 = lean_ctor_get_uint8(x_263, 12); +x_289 = lean_ctor_get_uint8(x_263, 13); +x_290 = lean_ctor_get_uint8(x_263, 14); +x_291 = lean_ctor_get_uint8(x_263, 15); +x_292 = lean_ctor_get_uint8(x_263, 16); +if (lean_is_exclusive(x_263)) { + x_293 = x_263; +} else { + lean_dec_ref(x_263); + x_293 = lean_box(0); +} +x_294 = 2; +if (lean_is_scalar(x_293)) { + x_295 = lean_alloc_ctor(0, 0, 17); +} else { + x_295 = x_293; +} +lean_ctor_set_uint8(x_295, 0, x_277); +lean_ctor_set_uint8(x_295, 1, x_278); +lean_ctor_set_uint8(x_295, 2, x_279); +lean_ctor_set_uint8(x_295, 3, x_280); +lean_ctor_set_uint8(x_295, 4, x_281); +lean_ctor_set_uint8(x_295, 5, x_282); +lean_ctor_set_uint8(x_295, 6, x_283); +lean_ctor_set_uint8(x_295, 7, x_284); +lean_ctor_set_uint8(x_295, 8, x_285); +lean_ctor_set_uint8(x_295, 9, x_294); +lean_ctor_set_uint8(x_295, 10, x_286); +lean_ctor_set_uint8(x_295, 11, x_287); +lean_ctor_set_uint8(x_295, 12, x_288); +lean_ctor_set_uint8(x_295, 13, x_289); +lean_ctor_set_uint8(x_295, 14, x_290); +lean_ctor_set_uint8(x_295, 15, x_291); +lean_ctor_set_uint8(x_295, 16, x_292); +x_296 = 2; +x_297 = lean_uint64_shift_right(x_267, x_296); +x_298 = lean_uint64_shift_left(x_297, x_296); +x_299 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__13; +x_300 = lean_uint64_lor(x_298, x_299); +x_301 = lean_alloc_ctor(0, 7, 11); +lean_ctor_set(x_301, 0, x_295); +lean_ctor_set(x_301, 1, x_269); +lean_ctor_set(x_301, 2, x_270); +lean_ctor_set(x_301, 3, x_271); +lean_ctor_set(x_301, 4, x_272); +lean_ctor_set(x_301, 5, x_273); +lean_ctor_set(x_301, 6, x_274); +lean_ctor_set_uint64(x_301, sizeof(void*)*7, x_300); +lean_ctor_set_uint8(x_301, sizeof(void*)*7 + 8, x_268); +lean_ctor_set_uint8(x_301, sizeof(void*)*7 + 9, x_275); +lean_ctor_set_uint8(x_301, sizeof(void*)*7 + 10, x_276); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_38); +lean_inc(x_265); +x_302 = l_Lean_Meta_isExprDefEq(x_265, x_38, x_301, x_9, x_10, x_11, x_264); +if (lean_obj_tag(x_302) == 0) +{ +lean_object* x_303; uint8_t x_304; +x_303 = lean_ctor_get(x_302, 0); +lean_inc(x_303); +x_304 = lean_unbox(x_303); +lean_dec(x_303); +if (x_304 == 0) +{ +lean_object* x_305; lean_object* x_306; +x_305 = lean_ctor_get(x_302, 1); +lean_inc(x_305); +lean_dec(x_302); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +x_306 = l_Lean_Meta_addPPExplicitToExposeDiff(x_265, x_38, x_8, x_9, x_10, x_11, x_305); +if (lean_obj_tag(x_306) == 0) +{ +lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; +x_307 = lean_ctor_get(x_306, 0); +lean_inc(x_307); +x_308 = lean_ctor_get(x_306, 1); +lean_inc(x_308); +lean_dec(x_306); +x_309 = lean_ctor_get(x_307, 0); +lean_inc(x_309); +x_310 = lean_ctor_get(x_307, 1); +lean_inc(x_310); +if (lean_is_exclusive(x_307)) { + lean_ctor_release(x_307, 0); + lean_ctor_release(x_307, 1); + x_311 = x_307; +} else { + lean_dec_ref(x_307); + x_311 = lean_box(0); +} +x_312 = l_Lean_indentExpr(x_309); +x_313 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__15; +if (lean_is_scalar(x_311)) { + x_314 = lean_alloc_ctor(7, 2, 0); +} else { + x_314 = x_311; + lean_ctor_set_tag(x_314, 7); +} +lean_ctor_set(x_314, 0, x_313); +lean_ctor_set(x_314, 1, x_312); +x_315 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__17; +if (lean_is_scalar(x_266)) { + x_316 = lean_alloc_ctor(7, 2, 0); +} else { + x_316 = x_266; + lean_ctor_set_tag(x_316, 7); +} +lean_ctor_set(x_316, 0, x_314); +lean_ctor_set(x_316, 1, x_315); +x_317 = l_Lean_indentExpr(x_310); +x_318 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_318, 0, x_316); +lean_ctor_set(x_318, 1, x_317); +x_319 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__8; lean_ctor_set_tag(x_44, 7); -lean_ctor_set(x_44, 1, x_270); -lean_ctor_set(x_44, 0, x_269); -x_271 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__8; -lean_ctor_set_tag(x_32, 7); -lean_ctor_set(x_32, 1, x_271); -lean_ctor_set(x_32, 0, x_44); -x_272 = l_Lean_throwErrorAt___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_throwStuckAtUniverseCnstr___spec__12(x_4, x_32, x_6, x_7, x_8, x_9, x_10, x_11, x_264); +lean_ctor_set(x_44, 1, x_319); +lean_ctor_set(x_44, 0, x_318); +x_320 = l_Lean_throwErrorAt___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_throwStuckAtUniverseCnstr___spec__12(x_4, x_44, x_6, x_7, x_8, x_9, x_10, x_11, x_308); lean_dec(x_11); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -return x_272; +return x_320; } else { -lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; -lean_dec(x_225); -lean_dec(x_224); +lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; +lean_dec(x_266); +lean_free_object(x_44); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +x_321 = lean_ctor_get(x_306, 0); +lean_inc(x_321); +x_322 = lean_ctor_get(x_306, 1); +lean_inc(x_322); +if (lean_is_exclusive(x_306)) { + lean_ctor_release(x_306, 0); + lean_ctor_release(x_306, 1); + x_323 = x_306; +} else { + lean_dec_ref(x_306); + x_323 = lean_box(0); +} +if (lean_is_scalar(x_323)) { + x_324 = lean_alloc_ctor(1, 2, 0); +} else { + x_324 = x_323; +} +lean_ctor_set(x_324, 0, x_321); +lean_ctor_set(x_324, 1, x_322); +return x_324; +} +} +else +{ +lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; +lean_dec(x_266); +lean_dec(x_265); lean_free_object(x_44); lean_dec(x_38); -lean_free_object(x_32); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -x_273 = lean_ctor_get(x_261, 1); -lean_inc(x_273); -if (lean_is_exclusive(x_261)) { - lean_ctor_release(x_261, 0); - lean_ctor_release(x_261, 1); - x_274 = x_261; +x_325 = lean_ctor_get(x_302, 1); +lean_inc(x_325); +if (lean_is_exclusive(x_302)) { + lean_ctor_release(x_302, 0); + lean_ctor_release(x_302, 1); + x_326 = x_302; } else { - lean_dec_ref(x_261); - x_274 = lean_box(0); + lean_dec_ref(x_302); + x_326 = lean_box(0); } -x_275 = lean_box(0); -if (lean_is_scalar(x_274)) { - x_276 = lean_alloc_ctor(0, 2, 0); +x_327 = lean_box(0); +if (lean_is_scalar(x_326)) { + x_328 = lean_alloc_ctor(0, 2, 0); } else { - x_276 = x_274; + x_328 = x_326; } -lean_ctor_set(x_276, 0, x_275); -lean_ctor_set(x_276, 1, x_273); -return x_276; +lean_ctor_set(x_328, 0, x_327); +lean_ctor_set(x_328, 1, x_325); +return x_328; } } else { -lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; -lean_dec(x_225); -lean_dec(x_224); +lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; +lean_dec(x_266); +lean_dec(x_265); lean_free_object(x_44); lean_dec(x_38); -lean_free_object(x_32); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -x_277 = lean_ctor_get(x_261, 0); -lean_inc(x_277); -x_278 = lean_ctor_get(x_261, 1); -lean_inc(x_278); -if (lean_is_exclusive(x_261)) { - lean_ctor_release(x_261, 0); - lean_ctor_release(x_261, 1); - x_279 = x_261; +x_329 = lean_ctor_get(x_302, 0); +lean_inc(x_329); +x_330 = lean_ctor_get(x_302, 1); +lean_inc(x_330); +if (lean_is_exclusive(x_302)) { + lean_ctor_release(x_302, 0); + lean_ctor_release(x_302, 1); + x_331 = x_302; } else { - lean_dec_ref(x_261); - x_279 = lean_box(0); + lean_dec_ref(x_302); + x_331 = lean_box(0); } -if (lean_is_scalar(x_279)) { - x_280 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_331)) { + x_332 = lean_alloc_ctor(1, 2, 0); } else { - x_280 = x_279; + x_332 = x_331; } -lean_ctor_set(x_280, 0, x_277); -lean_ctor_set(x_280, 1, x_278); -return x_280; +lean_ctor_set(x_332, 0, x_329); +lean_ctor_set(x_332, 1, x_330); +return x_332; } } else { -lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; +lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_free_object(x_44); lean_dec(x_38); -lean_free_object(x_32); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -x_281 = lean_ctor_get(x_220, 0); -lean_inc(x_281); -x_282 = lean_ctor_get(x_220, 1); -lean_inc(x_282); -if (lean_is_exclusive(x_220)) { - lean_ctor_release(x_220, 0); - lean_ctor_release(x_220, 1); - x_283 = x_220; +x_333 = lean_ctor_get(x_261, 0); +lean_inc(x_333); +x_334 = lean_ctor_get(x_261, 1); +lean_inc(x_334); +if (lean_is_exclusive(x_261)) { + lean_ctor_release(x_261, 0); + lean_ctor_release(x_261, 1); + x_335 = x_261; } else { - lean_dec_ref(x_220); - x_283 = lean_box(0); + lean_dec_ref(x_261); + x_335 = lean_box(0); } -if (lean_is_scalar(x_283)) { - x_284 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_335)) { + x_336 = lean_alloc_ctor(1, 2, 0); } else { - x_284 = x_283; + x_336 = x_335; } -lean_ctor_set(x_284, 0, x_281); -lean_ctor_set(x_284, 1, x_282); -return x_284; +lean_ctor_set(x_336, 0, x_333); +lean_ctor_set(x_336, 1, x_334); +return x_336; } } else { -lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; +lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_free_object(x_44); lean_dec(x_38); -lean_free_object(x_32); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -x_285 = lean_ctor_get(x_217, 0); -lean_inc(x_285); -x_286 = lean_ctor_get(x_217, 1); -lean_inc(x_286); -if (lean_is_exclusive(x_217)) { - lean_ctor_release(x_217, 0); - lean_ctor_release(x_217, 1); - x_287 = x_217; +x_337 = lean_ctor_get(x_258, 0); +lean_inc(x_337); +x_338 = lean_ctor_get(x_258, 1); +lean_inc(x_338); +if (lean_is_exclusive(x_258)) { + lean_ctor_release(x_258, 0); + lean_ctor_release(x_258, 1); + x_339 = x_258; } else { - lean_dec_ref(x_217); - x_287 = lean_box(0); + lean_dec_ref(x_258); + x_339 = lean_box(0); } -if (lean_is_scalar(x_287)) { - x_288 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_339)) { + x_340 = lean_alloc_ctor(1, 2, 0); } else { - x_288 = x_287; + x_340 = x_339; } -lean_ctor_set(x_288, 0, x_285); -lean_ctor_set(x_288, 1, x_286); -return x_288; +lean_ctor_set(x_340, 0, x_337); +lean_ctor_set(x_340, 1, x_338); +return x_340; } } } else { -uint8_t x_289; -x_289 = !lean_is_exclusive(x_46); -if (x_289 == 0) -{ -lean_object* x_290; lean_object* x_291; uint8_t x_292; -x_290 = lean_ctor_get(x_46, 1); -lean_dec(x_290); -x_291 = lean_ctor_get(x_46, 0); -lean_dec(x_291); -x_292 = !lean_is_exclusive(x_56); -if (x_292 == 0) -{ -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_293 = lean_ctor_get(x_56, 1); -lean_dec(x_293); -x_294 = lean_ctor_get(x_56, 0); -lean_dec(x_294); -x_295 = lean_ctor_get(x_43, 1); -lean_inc(x_295); +uint8_t x_341; +x_341 = !lean_is_exclusive(x_46); +if (x_341 == 0) +{ +lean_object* x_342; lean_object* x_343; uint8_t x_344; +x_342 = lean_ctor_get(x_46, 1); +lean_dec(x_342); +x_343 = lean_ctor_get(x_46, 0); +lean_dec(x_343); +x_344 = !lean_is_exclusive(x_56); +if (x_344 == 0) +{ +lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; lean_object* x_353; +x_345 = lean_ctor_get(x_56, 1); +lean_dec(x_345); +x_346 = lean_ctor_get(x_56, 0); +lean_dec(x_346); +x_347 = lean_ctor_get(x_43, 1); +lean_inc(x_347); lean_dec(x_43); -x_296 = l_Lean_MessageData_ofSyntax(x_3); -x_297 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__8; +x_348 = l_Lean_MessageData_ofSyntax(x_3); +x_349 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__8; lean_ctor_set_tag(x_56, 7); -lean_ctor_set(x_56, 1, x_296); -lean_ctor_set(x_56, 0, x_297); -x_298 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__19; +lean_ctor_set(x_56, 1, x_348); +lean_ctor_set(x_56, 0, x_349); +x_350 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__19; lean_ctor_set_tag(x_46, 7); -lean_ctor_set(x_46, 1, x_298); +lean_ctor_set(x_46, 1, x_350); lean_ctor_set(x_46, 0, x_56); -x_299 = l_Lean_indentExpr(x_38); +x_351 = l_Lean_indentExpr(x_38); lean_ctor_set_tag(x_44, 7); -lean_ctor_set(x_44, 1, x_299); -x_300 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__12; +lean_ctor_set(x_44, 1, x_351); +x_352 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__12; lean_ctor_set_tag(x_32, 7); -lean_ctor_set(x_32, 1, x_300); +lean_ctor_set(x_32, 1, x_352); lean_ctor_set(x_32, 0, x_44); -x_301 = l_Lean_throwErrorAt___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_throwStuckAtUniverseCnstr___spec__12(x_4, x_32, x_6, x_7, x_8, x_9, x_10, x_11, x_295); +x_353 = l_Lean_throwErrorAt___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_throwStuckAtUniverseCnstr___spec__12(x_4, x_32, x_6, x_7, x_8, x_9, x_10, x_11, x_347); lean_dec(x_11); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -return x_301; +return x_353; } else { -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_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_dec(x_56); -x_302 = lean_ctor_get(x_43, 1); -lean_inc(x_302); +x_354 = lean_ctor_get(x_43, 1); +lean_inc(x_354); lean_dec(x_43); -x_303 = l_Lean_MessageData_ofSyntax(x_3); -x_304 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__8; -x_305 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_305, 0, x_304); -lean_ctor_set(x_305, 1, x_303); -x_306 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__19; +x_355 = l_Lean_MessageData_ofSyntax(x_3); +x_356 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__8; +x_357 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_357, 0, x_356); +lean_ctor_set(x_357, 1, x_355); +x_358 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__19; lean_ctor_set_tag(x_46, 7); -lean_ctor_set(x_46, 1, x_306); -lean_ctor_set(x_46, 0, x_305); -x_307 = l_Lean_indentExpr(x_38); +lean_ctor_set(x_46, 1, x_358); +lean_ctor_set(x_46, 0, x_357); +x_359 = l_Lean_indentExpr(x_38); lean_ctor_set_tag(x_44, 7); -lean_ctor_set(x_44, 1, x_307); -x_308 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__12; +lean_ctor_set(x_44, 1, x_359); +x_360 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__12; lean_ctor_set_tag(x_32, 7); -lean_ctor_set(x_32, 1, x_308); +lean_ctor_set(x_32, 1, x_360); lean_ctor_set(x_32, 0, x_44); -x_309 = l_Lean_throwErrorAt___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_throwStuckAtUniverseCnstr___spec__12(x_4, x_32, x_6, x_7, x_8, x_9, x_10, x_11, x_302); +x_361 = l_Lean_throwErrorAt___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_throwStuckAtUniverseCnstr___spec__12(x_4, x_32, x_6, x_7, x_8, x_9, x_10, x_11, x_354); lean_dec(x_11); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -return x_309; +return x_361; } } else { -lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; +lean_object* x_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_dec(x_46); if (lean_is_exclusive(x_56)) { lean_ctor_release(x_56, 0); lean_ctor_release(x_56, 1); - x_310 = x_56; + x_362 = x_56; } else { lean_dec_ref(x_56); - x_310 = lean_box(0); + x_362 = lean_box(0); } -x_311 = lean_ctor_get(x_43, 1); -lean_inc(x_311); +x_363 = lean_ctor_get(x_43, 1); +lean_inc(x_363); lean_dec(x_43); -x_312 = l_Lean_MessageData_ofSyntax(x_3); -x_313 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__8; -if (lean_is_scalar(x_310)) { - x_314 = lean_alloc_ctor(7, 2, 0); -} else { - x_314 = x_310; - lean_ctor_set_tag(x_314, 7); -} -lean_ctor_set(x_314, 0, x_313); -lean_ctor_set(x_314, 1, x_312); -x_315 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__19; -x_316 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_316, 0, x_314); -lean_ctor_set(x_316, 1, x_315); -x_317 = l_Lean_indentExpr(x_38); +x_364 = l_Lean_MessageData_ofSyntax(x_3); +x_365 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__8; +if (lean_is_scalar(x_362)) { + x_366 = lean_alloc_ctor(7, 2, 0); +} else { + x_366 = x_362; + lean_ctor_set_tag(x_366, 7); +} +lean_ctor_set(x_366, 0, x_365); +lean_ctor_set(x_366, 1, x_364); +x_367 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__19; +x_368 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_368, 0, x_366); +lean_ctor_set(x_368, 1, x_367); +x_369 = l_Lean_indentExpr(x_38); lean_ctor_set_tag(x_44, 7); -lean_ctor_set(x_44, 1, x_317); -lean_ctor_set(x_44, 0, x_316); -x_318 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__12; +lean_ctor_set(x_44, 1, x_369); +lean_ctor_set(x_44, 0, x_368); +x_370 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__12; lean_ctor_set_tag(x_32, 7); -lean_ctor_set(x_32, 1, x_318); +lean_ctor_set(x_32, 1, x_370); lean_ctor_set(x_32, 0, x_44); -x_319 = l_Lean_throwErrorAt___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_throwStuckAtUniverseCnstr___spec__12(x_4, x_32, x_6, x_7, x_8, x_9, x_10, x_11, x_311); +x_371 = l_Lean_throwErrorAt___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_throwStuckAtUniverseCnstr___spec__12(x_4, x_32, x_6, x_7, x_8, x_9, x_10, x_11, x_363); lean_dec(x_11); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -return x_319; +return x_371; } } } } else { -lean_object* x_320; -x_320 = lean_ctor_get(x_44, 0); -lean_inc(x_320); +lean_object* x_372; +x_372 = lean_ctor_get(x_44, 0); +lean_inc(x_372); lean_dec(x_44); -if (lean_obj_tag(x_320) == 0) +if (lean_obj_tag(x_372) == 0) { -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_321 = lean_ctor_get(x_43, 1); -lean_inc(x_321); +lean_object* x_373; lean_object* x_374; lean_object* x_375; lean_object* x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; +x_373 = lean_ctor_get(x_43, 1); +lean_inc(x_373); lean_dec(x_43); -x_322 = l_Lean_MessageData_ofSyntax(x_3); -x_323 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__8; -x_324 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_324, 0, x_323); -lean_ctor_set(x_324, 1, x_322); -x_325 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__10; +x_374 = l_Lean_MessageData_ofSyntax(x_3); +x_375 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__8; +x_376 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_376, 0, x_375); +lean_ctor_set(x_376, 1, x_374); +x_377 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__10; lean_ctor_set_tag(x_32, 7); -lean_ctor_set(x_32, 1, x_325); -lean_ctor_set(x_32, 0, x_324); -x_326 = l_Lean_indentExpr(x_38); +lean_ctor_set(x_32, 1, x_377); +lean_ctor_set(x_32, 0, x_376); +x_378 = l_Lean_indentExpr(x_38); lean_ctor_set_tag(x_25, 7); -lean_ctor_set(x_25, 1, x_326); +lean_ctor_set(x_25, 1, x_378); lean_ctor_set(x_25, 0, x_32); -x_327 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__12; -x_328 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_328, 0, x_25); -lean_ctor_set(x_328, 1, x_327); -x_329 = l_Lean_throwErrorAt___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_throwStuckAtUniverseCnstr___spec__12(x_4, x_328, x_6, x_7, x_8, x_9, x_10, x_11, x_321); +x_379 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__12; +x_380 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_380, 0, x_25); +lean_ctor_set(x_380, 1, x_379); +x_381 = l_Lean_throwErrorAt___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_throwStuckAtUniverseCnstr___spec__12(x_4, x_380, x_6, x_7, x_8, x_9, x_10, x_11, x_373); lean_dec(x_11); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -return x_329; +return x_381; } else { -lean_object* x_330; +lean_object* x_382; lean_free_object(x_25); -x_330 = lean_ctor_get(x_320, 1); -lean_inc(x_330); -if (lean_obj_tag(x_330) == 0) +x_382 = lean_ctor_get(x_372, 1); +lean_inc(x_382); +if (lean_obj_tag(x_382) == 0) { -lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; +lean_object* x_383; lean_object* x_384; lean_object* x_385; lean_object* x_386; +lean_free_object(x_32); lean_dec(x_3); -x_331 = lean_ctor_get(x_43, 1); -lean_inc(x_331); +x_383 = lean_ctor_get(x_43, 1); +lean_inc(x_383); lean_dec(x_43); -x_332 = lean_ctor_get(x_320, 0); -lean_inc(x_332); -if (lean_is_exclusive(x_320)) { - lean_ctor_release(x_320, 0); - lean_ctor_release(x_320, 1); - x_333 = x_320; -} else { - lean_dec_ref(x_320); - x_333 = lean_box(0); -} -x_334 = l_Lean_MVarId_getType(x_332, x_8, x_9, x_10, x_11, x_331); -if (lean_obj_tag(x_334) == 0) -{ -lean_object* x_335; lean_object* x_336; lean_object* x_337; -x_335 = lean_ctor_get(x_334, 0); -lean_inc(x_335); -x_336 = lean_ctor_get(x_334, 1); -lean_inc(x_336); -lean_dec(x_334); +x_384 = lean_ctor_get(x_372, 0); +lean_inc(x_384); +if (lean_is_exclusive(x_372)) { + lean_ctor_release(x_372, 0); + lean_ctor_release(x_372, 1); + x_385 = x_372; +} else { + lean_dec_ref(x_372); + x_385 = lean_box(0); +} +x_386 = l_Lean_MVarId_getType(x_384, x_8, x_9, x_10, x_11, x_383); +if (lean_obj_tag(x_386) == 0) +{ +lean_object* x_387; lean_object* x_388; lean_object* x_389; +x_387 = lean_ctor_get(x_386, 0); +lean_inc(x_387); +x_388 = lean_ctor_get(x_386, 1); +lean_inc(x_388); +lean_dec(x_386); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); -x_337 = l_Lean_Meta_CheckTactic_matchCheckGoalType(x_4, x_335, x_8, x_9, x_10, x_11, x_336); -if (lean_obj_tag(x_337) == 0) +x_389 = l_Lean_Meta_CheckTactic_matchCheckGoalType(x_4, x_387, x_8, x_9, x_10, x_11, x_388); +if (lean_obj_tag(x_389) == 0) { -lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; uint64_t x_343; uint8_t 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; 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; uint8_t x_360; uint8_t x_361; uint8_t x_362; uint8_t x_363; uint8_t x_364; uint8_t x_365; uint8_t x_366; uint8_t x_367; uint8_t x_368; lean_object* x_369; uint8_t x_370; lean_object* x_371; uint64_t x_372; uint64_t x_373; uint64_t x_374; uint64_t x_375; uint64_t x_376; lean_object* x_377; lean_object* x_378; -x_338 = lean_ctor_get(x_337, 0); -lean_inc(x_338); -x_339 = lean_ctor_get(x_8, 0); -lean_inc(x_339); -x_340 = lean_ctor_get(x_337, 1); -lean_inc(x_340); -lean_dec(x_337); -x_341 = lean_ctor_get(x_338, 0); -lean_inc(x_341); -if (lean_is_exclusive(x_338)) { - lean_ctor_release(x_338, 0); - lean_ctor_release(x_338, 1); - x_342 = x_338; -} else { - lean_dec_ref(x_338); - x_342 = lean_box(0); -} -x_343 = lean_ctor_get_uint64(x_8, sizeof(void*)*7); -x_344 = lean_ctor_get_uint8(x_8, sizeof(void*)*7 + 8); -x_345 = lean_ctor_get(x_8, 1); -lean_inc(x_345); -x_346 = lean_ctor_get(x_8, 2); -lean_inc(x_346); -x_347 = lean_ctor_get(x_8, 3); -lean_inc(x_347); -x_348 = lean_ctor_get(x_8, 4); -lean_inc(x_348); -x_349 = lean_ctor_get(x_8, 5); -lean_inc(x_349); -x_350 = lean_ctor_get(x_8, 6); -lean_inc(x_350); -x_351 = lean_ctor_get_uint8(x_8, sizeof(void*)*7 + 9); -x_352 = lean_ctor_get_uint8(x_8, sizeof(void*)*7 + 10); -x_353 = lean_ctor_get_uint8(x_339, 0); -x_354 = lean_ctor_get_uint8(x_339, 1); -x_355 = lean_ctor_get_uint8(x_339, 2); -x_356 = lean_ctor_get_uint8(x_339, 3); -x_357 = lean_ctor_get_uint8(x_339, 4); -x_358 = lean_ctor_get_uint8(x_339, 5); -x_359 = lean_ctor_get_uint8(x_339, 6); -x_360 = lean_ctor_get_uint8(x_339, 7); -x_361 = lean_ctor_get_uint8(x_339, 8); -x_362 = lean_ctor_get_uint8(x_339, 10); -x_363 = lean_ctor_get_uint8(x_339, 11); -x_364 = lean_ctor_get_uint8(x_339, 12); -x_365 = lean_ctor_get_uint8(x_339, 13); -x_366 = lean_ctor_get_uint8(x_339, 14); -x_367 = lean_ctor_get_uint8(x_339, 15); -x_368 = lean_ctor_get_uint8(x_339, 16); -if (lean_is_exclusive(x_339)) { - x_369 = x_339; -} else { - lean_dec_ref(x_339); - x_369 = lean_box(0); -} -x_370 = 2; -if (lean_is_scalar(x_369)) { - x_371 = lean_alloc_ctor(0, 0, 17); -} else { - x_371 = x_369; -} -lean_ctor_set_uint8(x_371, 0, x_353); -lean_ctor_set_uint8(x_371, 1, x_354); -lean_ctor_set_uint8(x_371, 2, x_355); -lean_ctor_set_uint8(x_371, 3, x_356); -lean_ctor_set_uint8(x_371, 4, x_357); -lean_ctor_set_uint8(x_371, 5, x_358); -lean_ctor_set_uint8(x_371, 6, x_359); -lean_ctor_set_uint8(x_371, 7, x_360); -lean_ctor_set_uint8(x_371, 8, x_361); -lean_ctor_set_uint8(x_371, 9, x_370); -lean_ctor_set_uint8(x_371, 10, x_362); -lean_ctor_set_uint8(x_371, 11, x_363); -lean_ctor_set_uint8(x_371, 12, x_364); -lean_ctor_set_uint8(x_371, 13, x_365); -lean_ctor_set_uint8(x_371, 14, x_366); -lean_ctor_set_uint8(x_371, 15, x_367); -lean_ctor_set_uint8(x_371, 16, x_368); -x_372 = 2; -x_373 = lean_uint64_shift_right(x_343, x_372); -x_374 = lean_uint64_shift_left(x_373, x_372); -x_375 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__13; -x_376 = lean_uint64_lor(x_374, x_375); -x_377 = lean_alloc_ctor(0, 7, 11); -lean_ctor_set(x_377, 0, x_371); -lean_ctor_set(x_377, 1, x_345); -lean_ctor_set(x_377, 2, x_346); -lean_ctor_set(x_377, 3, x_347); -lean_ctor_set(x_377, 4, x_348); -lean_ctor_set(x_377, 5, x_349); -lean_ctor_set(x_377, 6, x_350); -lean_ctor_set_uint64(x_377, sizeof(void*)*7, x_376); -lean_ctor_set_uint8(x_377, sizeof(void*)*7 + 8, x_344); -lean_ctor_set_uint8(x_377, sizeof(void*)*7 + 9, x_351); -lean_ctor_set_uint8(x_377, sizeof(void*)*7 + 10, x_352); +lean_object* x_390; lean_object* x_391; lean_object* x_392; lean_object* x_393; lean_object* x_394; uint64_t x_395; uint8_t 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; uint8_t x_403; uint8_t x_404; uint8_t x_405; uint8_t x_406; uint8_t x_407; uint8_t x_408; uint8_t x_409; uint8_t x_410; uint8_t x_411; uint8_t x_412; uint8_t x_413; uint8_t x_414; uint8_t x_415; uint8_t x_416; uint8_t x_417; uint8_t x_418; uint8_t x_419; uint8_t x_420; lean_object* x_421; uint8_t x_422; lean_object* x_423; uint64_t x_424; uint64_t x_425; uint64_t x_426; uint64_t x_427; uint64_t x_428; lean_object* x_429; lean_object* x_430; +x_390 = lean_ctor_get(x_389, 0); +lean_inc(x_390); +x_391 = lean_ctor_get(x_8, 0); +lean_inc(x_391); +x_392 = lean_ctor_get(x_389, 1); +lean_inc(x_392); +lean_dec(x_389); +x_393 = lean_ctor_get(x_390, 0); +lean_inc(x_393); +if (lean_is_exclusive(x_390)) { + lean_ctor_release(x_390, 0); + lean_ctor_release(x_390, 1); + x_394 = x_390; +} else { + lean_dec_ref(x_390); + x_394 = lean_box(0); +} +x_395 = lean_ctor_get_uint64(x_8, sizeof(void*)*7); +x_396 = lean_ctor_get_uint8(x_8, sizeof(void*)*7 + 8); +x_397 = lean_ctor_get(x_8, 1); +lean_inc(x_397); +x_398 = lean_ctor_get(x_8, 2); +lean_inc(x_398); +x_399 = lean_ctor_get(x_8, 3); +lean_inc(x_399); +x_400 = lean_ctor_get(x_8, 4); +lean_inc(x_400); +x_401 = lean_ctor_get(x_8, 5); +lean_inc(x_401); +x_402 = lean_ctor_get(x_8, 6); +lean_inc(x_402); +x_403 = lean_ctor_get_uint8(x_8, sizeof(void*)*7 + 9); +x_404 = lean_ctor_get_uint8(x_8, sizeof(void*)*7 + 10); +x_405 = lean_ctor_get_uint8(x_391, 0); +x_406 = lean_ctor_get_uint8(x_391, 1); +x_407 = lean_ctor_get_uint8(x_391, 2); +x_408 = lean_ctor_get_uint8(x_391, 3); +x_409 = lean_ctor_get_uint8(x_391, 4); +x_410 = lean_ctor_get_uint8(x_391, 5); +x_411 = lean_ctor_get_uint8(x_391, 6); +x_412 = lean_ctor_get_uint8(x_391, 7); +x_413 = lean_ctor_get_uint8(x_391, 8); +x_414 = lean_ctor_get_uint8(x_391, 10); +x_415 = lean_ctor_get_uint8(x_391, 11); +x_416 = lean_ctor_get_uint8(x_391, 12); +x_417 = lean_ctor_get_uint8(x_391, 13); +x_418 = lean_ctor_get_uint8(x_391, 14); +x_419 = lean_ctor_get_uint8(x_391, 15); +x_420 = lean_ctor_get_uint8(x_391, 16); +if (lean_is_exclusive(x_391)) { + x_421 = x_391; +} else { + lean_dec_ref(x_391); + x_421 = lean_box(0); +} +x_422 = 2; +if (lean_is_scalar(x_421)) { + x_423 = lean_alloc_ctor(0, 0, 17); +} else { + x_423 = x_421; +} +lean_ctor_set_uint8(x_423, 0, x_405); +lean_ctor_set_uint8(x_423, 1, x_406); +lean_ctor_set_uint8(x_423, 2, x_407); +lean_ctor_set_uint8(x_423, 3, x_408); +lean_ctor_set_uint8(x_423, 4, x_409); +lean_ctor_set_uint8(x_423, 5, x_410); +lean_ctor_set_uint8(x_423, 6, x_411); +lean_ctor_set_uint8(x_423, 7, x_412); +lean_ctor_set_uint8(x_423, 8, x_413); +lean_ctor_set_uint8(x_423, 9, x_422); +lean_ctor_set_uint8(x_423, 10, x_414); +lean_ctor_set_uint8(x_423, 11, x_415); +lean_ctor_set_uint8(x_423, 12, x_416); +lean_ctor_set_uint8(x_423, 13, x_417); +lean_ctor_set_uint8(x_423, 14, x_418); +lean_ctor_set_uint8(x_423, 15, x_419); +lean_ctor_set_uint8(x_423, 16, x_420); +x_424 = 2; +x_425 = lean_uint64_shift_right(x_395, x_424); +x_426 = lean_uint64_shift_left(x_425, x_424); +x_427 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__13; +x_428 = lean_uint64_lor(x_426, x_427); +x_429 = lean_alloc_ctor(0, 7, 11); +lean_ctor_set(x_429, 0, x_423); +lean_ctor_set(x_429, 1, x_397); +lean_ctor_set(x_429, 2, x_398); +lean_ctor_set(x_429, 3, x_399); +lean_ctor_set(x_429, 4, x_400); +lean_ctor_set(x_429, 5, x_401); +lean_ctor_set(x_429, 6, x_402); +lean_ctor_set_uint64(x_429, sizeof(void*)*7, x_428); +lean_ctor_set_uint8(x_429, sizeof(void*)*7 + 8, x_396); +lean_ctor_set_uint8(x_429, sizeof(void*)*7 + 9, x_403); +lean_ctor_set_uint8(x_429, sizeof(void*)*7 + 10, x_404); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_38); -lean_inc(x_341); -x_378 = l_Lean_Meta_isExprDefEq(x_341, x_38, x_377, x_9, x_10, x_11, x_340); -if (lean_obj_tag(x_378) == 0) -{ -lean_object* x_379; uint8_t x_380; -x_379 = lean_ctor_get(x_378, 0); -lean_inc(x_379); -x_380 = lean_unbox(x_379); -lean_dec(x_379); -if (x_380 == 0) -{ -lean_object* x_381; lean_object* x_382; lean_object* x_383; lean_object* x_384; lean_object* x_385; lean_object* x_386; lean_object* x_387; lean_object* x_388; lean_object* x_389; lean_object* x_390; -x_381 = lean_ctor_get(x_378, 1); -lean_inc(x_381); -lean_dec(x_378); -x_382 = l_Lean_indentExpr(x_341); -x_383 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__15; -if (lean_is_scalar(x_342)) { - x_384 = lean_alloc_ctor(7, 2, 0); -} else { - x_384 = x_342; - lean_ctor_set_tag(x_384, 7); -} -lean_ctor_set(x_384, 0, x_383); -lean_ctor_set(x_384, 1, x_382); -x_385 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__17; -if (lean_is_scalar(x_333)) { - x_386 = lean_alloc_ctor(7, 2, 0); -} else { - x_386 = x_333; - lean_ctor_set_tag(x_386, 7); -} -lean_ctor_set(x_386, 0, x_384); -lean_ctor_set(x_386, 1, x_385); -x_387 = l_Lean_indentExpr(x_38); -x_388 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_388, 0, x_386); -lean_ctor_set(x_388, 1, x_387); -x_389 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__8; -lean_ctor_set_tag(x_32, 7); -lean_ctor_set(x_32, 1, x_389); -lean_ctor_set(x_32, 0, x_388); -x_390 = l_Lean_throwErrorAt___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_throwStuckAtUniverseCnstr___spec__12(x_4, x_32, x_6, x_7, x_8, x_9, x_10, x_11, x_381); +lean_inc(x_393); +x_430 = l_Lean_Meta_isExprDefEq(x_393, x_38, x_429, x_9, x_10, x_11, x_392); +if (lean_obj_tag(x_430) == 0) +{ +lean_object* x_431; uint8_t x_432; +x_431 = lean_ctor_get(x_430, 0); +lean_inc(x_431); +x_432 = lean_unbox(x_431); +lean_dec(x_431); +if (x_432 == 0) +{ +lean_object* x_433; lean_object* x_434; +x_433 = lean_ctor_get(x_430, 1); +lean_inc(x_433); +lean_dec(x_430); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +x_434 = l_Lean_Meta_addPPExplicitToExposeDiff(x_393, x_38, x_8, x_9, x_10, x_11, x_433); +if (lean_obj_tag(x_434) == 0) +{ +lean_object* x_435; lean_object* x_436; lean_object* x_437; lean_object* x_438; lean_object* x_439; lean_object* x_440; lean_object* x_441; lean_object* x_442; lean_object* x_443; lean_object* x_444; lean_object* x_445; lean_object* x_446; lean_object* x_447; lean_object* x_448; lean_object* x_449; +x_435 = lean_ctor_get(x_434, 0); +lean_inc(x_435); +x_436 = lean_ctor_get(x_434, 1); +lean_inc(x_436); +lean_dec(x_434); +x_437 = lean_ctor_get(x_435, 0); +lean_inc(x_437); +x_438 = lean_ctor_get(x_435, 1); +lean_inc(x_438); +if (lean_is_exclusive(x_435)) { + lean_ctor_release(x_435, 0); + lean_ctor_release(x_435, 1); + x_439 = x_435; +} else { + lean_dec_ref(x_435); + x_439 = lean_box(0); +} +x_440 = l_Lean_indentExpr(x_437); +x_441 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__15; +if (lean_is_scalar(x_439)) { + x_442 = lean_alloc_ctor(7, 2, 0); +} else { + x_442 = x_439; + lean_ctor_set_tag(x_442, 7); +} +lean_ctor_set(x_442, 0, x_441); +lean_ctor_set(x_442, 1, x_440); +x_443 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__17; +if (lean_is_scalar(x_394)) { + x_444 = lean_alloc_ctor(7, 2, 0); +} else { + x_444 = x_394; + lean_ctor_set_tag(x_444, 7); +} +lean_ctor_set(x_444, 0, x_442); +lean_ctor_set(x_444, 1, x_443); +x_445 = l_Lean_indentExpr(x_438); +if (lean_is_scalar(x_385)) { + x_446 = lean_alloc_ctor(7, 2, 0); +} else { + x_446 = x_385; + lean_ctor_set_tag(x_446, 7); +} +lean_ctor_set(x_446, 0, x_444); +lean_ctor_set(x_446, 1, x_445); +x_447 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__8; +x_448 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_448, 0, x_446); +lean_ctor_set(x_448, 1, x_447); +x_449 = l_Lean_throwErrorAt___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_throwStuckAtUniverseCnstr___spec__12(x_4, x_448, x_6, x_7, x_8, x_9, x_10, x_11, x_436); lean_dec(x_11); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -return x_390; +return x_449; } else { -lean_object* x_391; lean_object* x_392; lean_object* x_393; lean_object* x_394; -lean_dec(x_342); -lean_dec(x_341); -lean_dec(x_333); +lean_object* x_450; lean_object* x_451; lean_object* x_452; lean_object* x_453; +lean_dec(x_394); +lean_dec(x_385); +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_450 = lean_ctor_get(x_434, 0); +lean_inc(x_450); +x_451 = lean_ctor_get(x_434, 1); +lean_inc(x_451); +if (lean_is_exclusive(x_434)) { + lean_ctor_release(x_434, 0); + lean_ctor_release(x_434, 1); + x_452 = x_434; +} else { + lean_dec_ref(x_434); + x_452 = lean_box(0); +} +if (lean_is_scalar(x_452)) { + x_453 = lean_alloc_ctor(1, 2, 0); +} else { + x_453 = x_452; +} +lean_ctor_set(x_453, 0, x_450); +lean_ctor_set(x_453, 1, x_451); +return x_453; +} +} +else +{ +lean_object* x_454; lean_object* x_455; lean_object* x_456; lean_object* x_457; +lean_dec(x_394); +lean_dec(x_393); +lean_dec(x_385); lean_dec(x_38); -lean_free_object(x_32); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -x_391 = lean_ctor_get(x_378, 1); -lean_inc(x_391); -if (lean_is_exclusive(x_378)) { - lean_ctor_release(x_378, 0); - lean_ctor_release(x_378, 1); - x_392 = x_378; +x_454 = lean_ctor_get(x_430, 1); +lean_inc(x_454); +if (lean_is_exclusive(x_430)) { + lean_ctor_release(x_430, 0); + lean_ctor_release(x_430, 1); + x_455 = x_430; } else { - lean_dec_ref(x_378); - x_392 = lean_box(0); + lean_dec_ref(x_430); + x_455 = lean_box(0); } -x_393 = lean_box(0); -if (lean_is_scalar(x_392)) { - x_394 = lean_alloc_ctor(0, 2, 0); +x_456 = lean_box(0); +if (lean_is_scalar(x_455)) { + x_457 = lean_alloc_ctor(0, 2, 0); } else { - x_394 = x_392; + x_457 = x_455; } -lean_ctor_set(x_394, 0, x_393); -lean_ctor_set(x_394, 1, x_391); -return x_394; +lean_ctor_set(x_457, 0, x_456); +lean_ctor_set(x_457, 1, x_454); +return x_457; } } else { -lean_object* x_395; lean_object* x_396; lean_object* x_397; lean_object* x_398; -lean_dec(x_342); -lean_dec(x_341); -lean_dec(x_333); +lean_object* x_458; lean_object* x_459; lean_object* x_460; lean_object* x_461; +lean_dec(x_394); +lean_dec(x_393); +lean_dec(x_385); lean_dec(x_38); -lean_free_object(x_32); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -x_395 = lean_ctor_get(x_378, 0); -lean_inc(x_395); -x_396 = lean_ctor_get(x_378, 1); -lean_inc(x_396); -if (lean_is_exclusive(x_378)) { - lean_ctor_release(x_378, 0); - lean_ctor_release(x_378, 1); - x_397 = x_378; +x_458 = lean_ctor_get(x_430, 0); +lean_inc(x_458); +x_459 = lean_ctor_get(x_430, 1); +lean_inc(x_459); +if (lean_is_exclusive(x_430)) { + lean_ctor_release(x_430, 0); + lean_ctor_release(x_430, 1); + x_460 = x_430; } else { - lean_dec_ref(x_378); - x_397 = lean_box(0); + lean_dec_ref(x_430); + x_460 = lean_box(0); } -if (lean_is_scalar(x_397)) { - x_398 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_460)) { + x_461 = lean_alloc_ctor(1, 2, 0); } else { - x_398 = x_397; + x_461 = x_460; } -lean_ctor_set(x_398, 0, x_395); -lean_ctor_set(x_398, 1, x_396); -return x_398; +lean_ctor_set(x_461, 0, x_458); +lean_ctor_set(x_461, 1, x_459); +return x_461; } } else { -lean_object* x_399; lean_object* x_400; lean_object* x_401; lean_object* x_402; -lean_dec(x_333); +lean_object* x_462; lean_object* x_463; lean_object* x_464; lean_object* x_465; +lean_dec(x_385); lean_dec(x_38); -lean_free_object(x_32); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -x_399 = lean_ctor_get(x_337, 0); -lean_inc(x_399); -x_400 = lean_ctor_get(x_337, 1); -lean_inc(x_400); -if (lean_is_exclusive(x_337)) { - lean_ctor_release(x_337, 0); - lean_ctor_release(x_337, 1); - x_401 = x_337; +x_462 = lean_ctor_get(x_389, 0); +lean_inc(x_462); +x_463 = lean_ctor_get(x_389, 1); +lean_inc(x_463); +if (lean_is_exclusive(x_389)) { + lean_ctor_release(x_389, 0); + lean_ctor_release(x_389, 1); + x_464 = x_389; } else { - lean_dec_ref(x_337); - x_401 = lean_box(0); + lean_dec_ref(x_389); + x_464 = lean_box(0); } -if (lean_is_scalar(x_401)) { - x_402 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_464)) { + x_465 = lean_alloc_ctor(1, 2, 0); } else { - x_402 = x_401; + x_465 = x_464; } -lean_ctor_set(x_402, 0, x_399); -lean_ctor_set(x_402, 1, x_400); -return x_402; +lean_ctor_set(x_465, 0, x_462); +lean_ctor_set(x_465, 1, x_463); +return x_465; } } else { -lean_object* x_403; lean_object* x_404; lean_object* x_405; lean_object* x_406; -lean_dec(x_333); +lean_object* x_466; lean_object* x_467; lean_object* x_468; lean_object* x_469; +lean_dec(x_385); lean_dec(x_38); -lean_free_object(x_32); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -x_403 = lean_ctor_get(x_334, 0); -lean_inc(x_403); -x_404 = lean_ctor_get(x_334, 1); -lean_inc(x_404); -if (lean_is_exclusive(x_334)) { - lean_ctor_release(x_334, 0); - lean_ctor_release(x_334, 1); - x_405 = x_334; +x_466 = lean_ctor_get(x_386, 0); +lean_inc(x_466); +x_467 = lean_ctor_get(x_386, 1); +lean_inc(x_467); +if (lean_is_exclusive(x_386)) { + lean_ctor_release(x_386, 0); + lean_ctor_release(x_386, 1); + x_468 = x_386; } else { - lean_dec_ref(x_334); - x_405 = lean_box(0); + lean_dec_ref(x_386); + x_468 = lean_box(0); } -if (lean_is_scalar(x_405)) { - x_406 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_468)) { + x_469 = lean_alloc_ctor(1, 2, 0); } else { - x_406 = x_405; + x_469 = x_468; } -lean_ctor_set(x_406, 0, x_403); -lean_ctor_set(x_406, 1, x_404); -return x_406; +lean_ctor_set(x_469, 0, x_466); +lean_ctor_set(x_469, 1, x_467); +return x_469; } } else { -lean_object* x_407; lean_object* x_408; lean_object* x_409; lean_object* x_410; lean_object* x_411; lean_object* x_412; lean_object* x_413; lean_object* x_414; lean_object* x_415; lean_object* x_416; lean_object* x_417; lean_object* x_418; -if (lean_is_exclusive(x_320)) { - lean_ctor_release(x_320, 0); - lean_ctor_release(x_320, 1); - x_407 = x_320; +lean_object* x_470; lean_object* x_471; lean_object* x_472; lean_object* x_473; lean_object* x_474; lean_object* x_475; lean_object* x_476; lean_object* x_477; lean_object* x_478; lean_object* x_479; lean_object* x_480; lean_object* x_481; +if (lean_is_exclusive(x_372)) { + lean_ctor_release(x_372, 0); + lean_ctor_release(x_372, 1); + x_470 = x_372; } else { - lean_dec_ref(x_320); - x_407 = lean_box(0); + lean_dec_ref(x_372); + x_470 = lean_box(0); } -if (lean_is_exclusive(x_330)) { - lean_ctor_release(x_330, 0); - lean_ctor_release(x_330, 1); - x_408 = x_330; +if (lean_is_exclusive(x_382)) { + lean_ctor_release(x_382, 0); + lean_ctor_release(x_382, 1); + x_471 = x_382; } else { - lean_dec_ref(x_330); - x_408 = lean_box(0); + lean_dec_ref(x_382); + x_471 = lean_box(0); } -x_409 = lean_ctor_get(x_43, 1); -lean_inc(x_409); +x_472 = lean_ctor_get(x_43, 1); +lean_inc(x_472); lean_dec(x_43); -x_410 = l_Lean_MessageData_ofSyntax(x_3); -x_411 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__8; -if (lean_is_scalar(x_408)) { - x_412 = lean_alloc_ctor(7, 2, 0); -} else { - x_412 = x_408; - lean_ctor_set_tag(x_412, 7); -} -lean_ctor_set(x_412, 0, x_411); -lean_ctor_set(x_412, 1, x_410); -x_413 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__19; -if (lean_is_scalar(x_407)) { - x_414 = lean_alloc_ctor(7, 2, 0); -} else { - x_414 = x_407; - lean_ctor_set_tag(x_414, 7); -} -lean_ctor_set(x_414, 0, x_412); -lean_ctor_set(x_414, 1, x_413); -x_415 = l_Lean_indentExpr(x_38); -x_416 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_416, 0, x_414); -lean_ctor_set(x_416, 1, x_415); -x_417 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__12; +x_473 = l_Lean_MessageData_ofSyntax(x_3); +x_474 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__8; +if (lean_is_scalar(x_471)) { + x_475 = lean_alloc_ctor(7, 2, 0); +} else { + x_475 = x_471; + lean_ctor_set_tag(x_475, 7); +} +lean_ctor_set(x_475, 0, x_474); +lean_ctor_set(x_475, 1, x_473); +x_476 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__19; +if (lean_is_scalar(x_470)) { + x_477 = lean_alloc_ctor(7, 2, 0); +} else { + x_477 = x_470; + lean_ctor_set_tag(x_477, 7); +} +lean_ctor_set(x_477, 0, x_475); +lean_ctor_set(x_477, 1, x_476); +x_478 = l_Lean_indentExpr(x_38); +x_479 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_479, 0, x_477); +lean_ctor_set(x_479, 1, x_478); +x_480 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__12; lean_ctor_set_tag(x_32, 7); -lean_ctor_set(x_32, 1, x_417); -lean_ctor_set(x_32, 0, x_416); -x_418 = l_Lean_throwErrorAt___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_throwStuckAtUniverseCnstr___spec__12(x_4, x_32, x_6, x_7, x_8, x_9, x_10, x_11, x_409); +lean_ctor_set(x_32, 1, x_480); +lean_ctor_set(x_32, 0, x_479); +x_481 = l_Lean_throwErrorAt___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_throwStuckAtUniverseCnstr___spec__12(x_4, x_32, x_6, x_7, x_8, x_9, x_10, x_11, x_472); lean_dec(x_11); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -return x_418; +return x_481; } } } } else { -uint8_t x_419; +uint8_t x_482; lean_dec(x_38); lean_free_object(x_32); lean_free_object(x_25); @@ -2072,29 +2395,29 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_3); -x_419 = !lean_is_exclusive(x_43); -if (x_419 == 0) +x_482 = !lean_is_exclusive(x_43); +if (x_482 == 0) { return x_43; } else { -lean_object* x_420; lean_object* x_421; lean_object* x_422; -x_420 = lean_ctor_get(x_43, 0); -x_421 = lean_ctor_get(x_43, 1); -lean_inc(x_421); -lean_inc(x_420); +lean_object* x_483; lean_object* x_484; lean_object* x_485; +x_483 = lean_ctor_get(x_43, 0); +x_484 = lean_ctor_get(x_43, 1); +lean_inc(x_484); +lean_inc(x_483); lean_dec(x_43); -x_422 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_422, 0, x_420); -lean_ctor_set(x_422, 1, x_421); -return x_422; +x_485 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_485, 0, x_483); +lean_ctor_set(x_485, 1, x_484); +return x_485; } } } else { -uint8_t x_423; +uint8_t x_486; lean_free_object(x_32); lean_dec(x_34); lean_free_object(x_25); @@ -2105,510 +2428,574 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_3); -x_423 = !lean_is_exclusive(x_37); -if (x_423 == 0) +x_486 = !lean_is_exclusive(x_37); +if (x_486 == 0) { return x_37; } else { -lean_object* x_424; lean_object* x_425; lean_object* x_426; -x_424 = lean_ctor_get(x_37, 0); -x_425 = lean_ctor_get(x_37, 1); -lean_inc(x_425); -lean_inc(x_424); +lean_object* x_487; lean_object* x_488; lean_object* x_489; +x_487 = lean_ctor_get(x_37, 0); +x_488 = lean_ctor_get(x_37, 1); +lean_inc(x_488); +lean_inc(x_487); lean_dec(x_37); -x_426 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_426, 0, x_424); -lean_ctor_set(x_426, 1, x_425); -return x_426; +x_489 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_489, 0, x_487); +lean_ctor_set(x_489, 1, x_488); +return x_489; } } } else { -lean_object* x_427; lean_object* x_428; lean_object* x_429; lean_object* x_430; -x_427 = lean_ctor_get(x_32, 0); -x_428 = lean_ctor_get(x_32, 1); -lean_inc(x_428); -lean_inc(x_427); +lean_object* x_490; lean_object* x_491; lean_object* x_492; lean_object* x_493; +x_490 = lean_ctor_get(x_32, 0); +x_491 = lean_ctor_get(x_32, 1); +lean_inc(x_491); +lean_inc(x_490); lean_dec(x_32); -x_429 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_429, 0, x_23); +x_492 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_492, 0, x_23); 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_430 = l_Lean_Elab_Term_elabTerm(x_2, x_429, x_14, x_14, x_6, x_7, x_8, x_9, x_10, x_11, x_428); -if (lean_obj_tag(x_430) == 0) -{ -lean_object* x_431; lean_object* x_432; lean_object* x_433; lean_object* x_434; lean_object* x_435; lean_object* x_436; -x_431 = lean_ctor_get(x_430, 0); -lean_inc(x_431); -x_432 = lean_ctor_get(x_430, 1); -lean_inc(x_432); -lean_dec(x_430); -x_433 = l_Lean_Expr_mvarId_x21(x_427); -lean_dec(x_427); -x_434 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__5; -x_435 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__6; +x_493 = l_Lean_Elab_Term_elabTerm(x_2, x_492, x_14, x_14, x_6, x_7, x_8, x_9, x_10, x_11, x_491); +if (lean_obj_tag(x_493) == 0) +{ +lean_object* x_494; lean_object* x_495; lean_object* x_496; lean_object* x_497; lean_object* x_498; lean_object* x_499; +x_494 = lean_ctor_get(x_493, 0); +lean_inc(x_494); +x_495 = lean_ctor_get(x_493, 1); +lean_inc(x_495); +lean_dec(x_493); +x_496 = l_Lean_Expr_mvarId_x21(x_490); +lean_dec(x_490); +x_497 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__5; +x_498 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__6; lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_3); -x_436 = l_Lean_Elab_runTactic(x_433, x_3, x_434, x_435, x_8, x_9, x_10, x_11, x_432); -if (lean_obj_tag(x_436) == 0) +x_499 = l_Lean_Elab_runTactic(x_496, x_3, x_497, x_498, x_8, x_9, x_10, x_11, x_495); +if (lean_obj_tag(x_499) == 0) { -lean_object* x_437; lean_object* x_438; lean_object* x_439; -x_437 = lean_ctor_get(x_436, 0); -lean_inc(x_437); -x_438 = lean_ctor_get(x_437, 0); -lean_inc(x_438); -if (lean_is_exclusive(x_437)) { - lean_ctor_release(x_437, 0); - lean_ctor_release(x_437, 1); - x_439 = x_437; +lean_object* x_500; lean_object* x_501; lean_object* x_502; +x_500 = lean_ctor_get(x_499, 0); +lean_inc(x_500); +x_501 = lean_ctor_get(x_500, 0); +lean_inc(x_501); +if (lean_is_exclusive(x_500)) { + lean_ctor_release(x_500, 0); + lean_ctor_release(x_500, 1); + x_502 = x_500; } else { - lean_dec_ref(x_437); - x_439 = lean_box(0); + lean_dec_ref(x_500); + x_502 = lean_box(0); } -if (lean_obj_tag(x_438) == 0) +if (lean_obj_tag(x_501) == 0) { -lean_object* x_440; lean_object* x_441; lean_object* x_442; lean_object* x_443; lean_object* x_444; lean_object* x_445; lean_object* x_446; lean_object* x_447; lean_object* x_448; lean_object* x_449; -x_440 = lean_ctor_get(x_436, 1); -lean_inc(x_440); -lean_dec(x_436); -x_441 = l_Lean_MessageData_ofSyntax(x_3); -x_442 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__8; -if (lean_is_scalar(x_439)) { - x_443 = lean_alloc_ctor(7, 2, 0); -} else { - x_443 = x_439; - lean_ctor_set_tag(x_443, 7); -} -lean_ctor_set(x_443, 0, x_442); -lean_ctor_set(x_443, 1, x_441); -x_444 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__10; -x_445 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_445, 0, x_443); -lean_ctor_set(x_445, 1, x_444); -x_446 = l_Lean_indentExpr(x_431); +lean_object* x_503; lean_object* x_504; lean_object* x_505; lean_object* x_506; lean_object* x_507; lean_object* x_508; lean_object* x_509; lean_object* x_510; lean_object* x_511; lean_object* x_512; +x_503 = lean_ctor_get(x_499, 1); +lean_inc(x_503); +lean_dec(x_499); +x_504 = l_Lean_MessageData_ofSyntax(x_3); +x_505 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__8; +if (lean_is_scalar(x_502)) { + x_506 = lean_alloc_ctor(7, 2, 0); +} else { + x_506 = x_502; + lean_ctor_set_tag(x_506, 7); +} +lean_ctor_set(x_506, 0, x_505); +lean_ctor_set(x_506, 1, x_504); +x_507 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__10; +x_508 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_508, 0, x_506); +lean_ctor_set(x_508, 1, x_507); +x_509 = l_Lean_indentExpr(x_494); lean_ctor_set_tag(x_25, 7); -lean_ctor_set(x_25, 1, x_446); -lean_ctor_set(x_25, 0, x_445); -x_447 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__12; -x_448 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_448, 0, x_25); -lean_ctor_set(x_448, 1, x_447); -x_449 = l_Lean_throwErrorAt___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_throwStuckAtUniverseCnstr___spec__12(x_4, x_448, x_6, x_7, x_8, x_9, x_10, x_11, x_440); +lean_ctor_set(x_25, 1, x_509); +lean_ctor_set(x_25, 0, x_508); +x_510 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__12; +x_511 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_511, 0, x_25); +lean_ctor_set(x_511, 1, x_510); +x_512 = l_Lean_throwErrorAt___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_throwStuckAtUniverseCnstr___spec__12(x_4, x_511, x_6, x_7, x_8, x_9, x_10, x_11, x_503); lean_dec(x_11); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -return x_449; +return x_512; } else { -lean_object* x_450; +lean_object* x_513; lean_free_object(x_25); -x_450 = lean_ctor_get(x_438, 1); -lean_inc(x_450); -if (lean_obj_tag(x_450) == 0) +x_513 = lean_ctor_get(x_501, 1); +lean_inc(x_513); +if (lean_obj_tag(x_513) == 0) { -lean_object* x_451; lean_object* x_452; lean_object* x_453; lean_object* x_454; +lean_object* x_514; lean_object* x_515; lean_object* x_516; lean_object* x_517; lean_dec(x_3); -x_451 = lean_ctor_get(x_436, 1); -lean_inc(x_451); -lean_dec(x_436); -x_452 = lean_ctor_get(x_438, 0); -lean_inc(x_452); -if (lean_is_exclusive(x_438)) { - lean_ctor_release(x_438, 0); - lean_ctor_release(x_438, 1); - x_453 = x_438; -} else { - lean_dec_ref(x_438); - x_453 = lean_box(0); -} -x_454 = l_Lean_MVarId_getType(x_452, x_8, x_9, x_10, x_11, x_451); -if (lean_obj_tag(x_454) == 0) -{ -lean_object* x_455; lean_object* x_456; lean_object* x_457; -x_455 = lean_ctor_get(x_454, 0); -lean_inc(x_455); -x_456 = lean_ctor_get(x_454, 1); -lean_inc(x_456); -lean_dec(x_454); +x_514 = lean_ctor_get(x_499, 1); +lean_inc(x_514); +lean_dec(x_499); +x_515 = lean_ctor_get(x_501, 0); +lean_inc(x_515); +if (lean_is_exclusive(x_501)) { + lean_ctor_release(x_501, 0); + lean_ctor_release(x_501, 1); + x_516 = x_501; +} else { + lean_dec_ref(x_501); + x_516 = lean_box(0); +} +x_517 = l_Lean_MVarId_getType(x_515, x_8, x_9, x_10, x_11, x_514); +if (lean_obj_tag(x_517) == 0) +{ +lean_object* x_518; lean_object* x_519; lean_object* x_520; +x_518 = lean_ctor_get(x_517, 0); +lean_inc(x_518); +x_519 = lean_ctor_get(x_517, 1); +lean_inc(x_519); +lean_dec(x_517); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); -x_457 = l_Lean_Meta_CheckTactic_matchCheckGoalType(x_4, x_455, x_8, x_9, x_10, x_11, x_456); -if (lean_obj_tag(x_457) == 0) +x_520 = l_Lean_Meta_CheckTactic_matchCheckGoalType(x_4, x_518, x_8, x_9, x_10, x_11, x_519); +if (lean_obj_tag(x_520) == 0) { -lean_object* x_458; lean_object* x_459; lean_object* x_460; lean_object* x_461; lean_object* x_462; uint64_t x_463; uint8_t 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; uint8_t x_471; uint8_t x_472; uint8_t x_473; uint8_t x_474; uint8_t x_475; uint8_t x_476; uint8_t x_477; uint8_t x_478; uint8_t x_479; uint8_t x_480; uint8_t x_481; uint8_t x_482; uint8_t x_483; uint8_t x_484; uint8_t x_485; uint8_t x_486; uint8_t x_487; uint8_t x_488; lean_object* x_489; uint8_t x_490; lean_object* x_491; uint64_t x_492; uint64_t x_493; uint64_t x_494; uint64_t x_495; uint64_t x_496; lean_object* x_497; lean_object* x_498; -x_458 = lean_ctor_get(x_457, 0); -lean_inc(x_458); -x_459 = lean_ctor_get(x_8, 0); -lean_inc(x_459); -x_460 = lean_ctor_get(x_457, 1); -lean_inc(x_460); -lean_dec(x_457); -x_461 = lean_ctor_get(x_458, 0); -lean_inc(x_461); -if (lean_is_exclusive(x_458)) { - lean_ctor_release(x_458, 0); - lean_ctor_release(x_458, 1); - x_462 = x_458; -} else { - lean_dec_ref(x_458); - x_462 = lean_box(0); -} -x_463 = lean_ctor_get_uint64(x_8, sizeof(void*)*7); -x_464 = lean_ctor_get_uint8(x_8, sizeof(void*)*7 + 8); -x_465 = lean_ctor_get(x_8, 1); -lean_inc(x_465); -x_466 = lean_ctor_get(x_8, 2); -lean_inc(x_466); -x_467 = lean_ctor_get(x_8, 3); -lean_inc(x_467); -x_468 = lean_ctor_get(x_8, 4); -lean_inc(x_468); -x_469 = lean_ctor_get(x_8, 5); -lean_inc(x_469); -x_470 = lean_ctor_get(x_8, 6); -lean_inc(x_470); -x_471 = lean_ctor_get_uint8(x_8, sizeof(void*)*7 + 9); -x_472 = lean_ctor_get_uint8(x_8, sizeof(void*)*7 + 10); -x_473 = lean_ctor_get_uint8(x_459, 0); -x_474 = lean_ctor_get_uint8(x_459, 1); -x_475 = lean_ctor_get_uint8(x_459, 2); -x_476 = lean_ctor_get_uint8(x_459, 3); -x_477 = lean_ctor_get_uint8(x_459, 4); -x_478 = lean_ctor_get_uint8(x_459, 5); -x_479 = lean_ctor_get_uint8(x_459, 6); -x_480 = lean_ctor_get_uint8(x_459, 7); -x_481 = lean_ctor_get_uint8(x_459, 8); -x_482 = lean_ctor_get_uint8(x_459, 10); -x_483 = lean_ctor_get_uint8(x_459, 11); -x_484 = lean_ctor_get_uint8(x_459, 12); -x_485 = lean_ctor_get_uint8(x_459, 13); -x_486 = lean_ctor_get_uint8(x_459, 14); -x_487 = lean_ctor_get_uint8(x_459, 15); -x_488 = lean_ctor_get_uint8(x_459, 16); -if (lean_is_exclusive(x_459)) { - x_489 = x_459; -} else { - lean_dec_ref(x_459); - x_489 = lean_box(0); -} -x_490 = 2; -if (lean_is_scalar(x_489)) { - x_491 = lean_alloc_ctor(0, 0, 17); -} else { - x_491 = x_489; -} -lean_ctor_set_uint8(x_491, 0, x_473); -lean_ctor_set_uint8(x_491, 1, x_474); -lean_ctor_set_uint8(x_491, 2, x_475); -lean_ctor_set_uint8(x_491, 3, x_476); -lean_ctor_set_uint8(x_491, 4, x_477); -lean_ctor_set_uint8(x_491, 5, x_478); -lean_ctor_set_uint8(x_491, 6, x_479); -lean_ctor_set_uint8(x_491, 7, x_480); -lean_ctor_set_uint8(x_491, 8, x_481); -lean_ctor_set_uint8(x_491, 9, x_490); -lean_ctor_set_uint8(x_491, 10, x_482); -lean_ctor_set_uint8(x_491, 11, x_483); -lean_ctor_set_uint8(x_491, 12, x_484); -lean_ctor_set_uint8(x_491, 13, x_485); -lean_ctor_set_uint8(x_491, 14, x_486); -lean_ctor_set_uint8(x_491, 15, x_487); -lean_ctor_set_uint8(x_491, 16, x_488); -x_492 = 2; -x_493 = lean_uint64_shift_right(x_463, x_492); -x_494 = lean_uint64_shift_left(x_493, x_492); -x_495 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__13; -x_496 = lean_uint64_lor(x_494, x_495); -x_497 = lean_alloc_ctor(0, 7, 11); -lean_ctor_set(x_497, 0, x_491); -lean_ctor_set(x_497, 1, x_465); -lean_ctor_set(x_497, 2, x_466); -lean_ctor_set(x_497, 3, x_467); -lean_ctor_set(x_497, 4, x_468); -lean_ctor_set(x_497, 5, x_469); -lean_ctor_set(x_497, 6, x_470); -lean_ctor_set_uint64(x_497, sizeof(void*)*7, x_496); -lean_ctor_set_uint8(x_497, sizeof(void*)*7 + 8, x_464); -lean_ctor_set_uint8(x_497, sizeof(void*)*7 + 9, x_471); -lean_ctor_set_uint8(x_497, sizeof(void*)*7 + 10, x_472); +lean_object* x_521; lean_object* x_522; lean_object* x_523; lean_object* x_524; lean_object* x_525; uint64_t x_526; uint8_t x_527; lean_object* x_528; lean_object* x_529; lean_object* x_530; lean_object* x_531; lean_object* x_532; lean_object* x_533; uint8_t x_534; uint8_t x_535; uint8_t x_536; uint8_t x_537; uint8_t x_538; uint8_t x_539; uint8_t x_540; uint8_t x_541; uint8_t x_542; uint8_t x_543; uint8_t x_544; uint8_t x_545; uint8_t x_546; uint8_t x_547; uint8_t x_548; uint8_t x_549; uint8_t x_550; uint8_t x_551; lean_object* x_552; uint8_t x_553; lean_object* x_554; uint64_t x_555; uint64_t x_556; uint64_t x_557; uint64_t x_558; uint64_t x_559; lean_object* x_560; lean_object* x_561; +x_521 = lean_ctor_get(x_520, 0); +lean_inc(x_521); +x_522 = lean_ctor_get(x_8, 0); +lean_inc(x_522); +x_523 = lean_ctor_get(x_520, 1); +lean_inc(x_523); +lean_dec(x_520); +x_524 = lean_ctor_get(x_521, 0); +lean_inc(x_524); +if (lean_is_exclusive(x_521)) { + lean_ctor_release(x_521, 0); + lean_ctor_release(x_521, 1); + x_525 = x_521; +} else { + lean_dec_ref(x_521); + x_525 = lean_box(0); +} +x_526 = lean_ctor_get_uint64(x_8, sizeof(void*)*7); +x_527 = lean_ctor_get_uint8(x_8, sizeof(void*)*7 + 8); +x_528 = lean_ctor_get(x_8, 1); +lean_inc(x_528); +x_529 = lean_ctor_get(x_8, 2); +lean_inc(x_529); +x_530 = lean_ctor_get(x_8, 3); +lean_inc(x_530); +x_531 = lean_ctor_get(x_8, 4); +lean_inc(x_531); +x_532 = lean_ctor_get(x_8, 5); +lean_inc(x_532); +x_533 = lean_ctor_get(x_8, 6); +lean_inc(x_533); +x_534 = lean_ctor_get_uint8(x_8, sizeof(void*)*7 + 9); +x_535 = lean_ctor_get_uint8(x_8, sizeof(void*)*7 + 10); +x_536 = lean_ctor_get_uint8(x_522, 0); +x_537 = lean_ctor_get_uint8(x_522, 1); +x_538 = lean_ctor_get_uint8(x_522, 2); +x_539 = lean_ctor_get_uint8(x_522, 3); +x_540 = lean_ctor_get_uint8(x_522, 4); +x_541 = lean_ctor_get_uint8(x_522, 5); +x_542 = lean_ctor_get_uint8(x_522, 6); +x_543 = lean_ctor_get_uint8(x_522, 7); +x_544 = lean_ctor_get_uint8(x_522, 8); +x_545 = lean_ctor_get_uint8(x_522, 10); +x_546 = lean_ctor_get_uint8(x_522, 11); +x_547 = lean_ctor_get_uint8(x_522, 12); +x_548 = lean_ctor_get_uint8(x_522, 13); +x_549 = lean_ctor_get_uint8(x_522, 14); +x_550 = lean_ctor_get_uint8(x_522, 15); +x_551 = lean_ctor_get_uint8(x_522, 16); +if (lean_is_exclusive(x_522)) { + x_552 = x_522; +} else { + lean_dec_ref(x_522); + x_552 = lean_box(0); +} +x_553 = 2; +if (lean_is_scalar(x_552)) { + x_554 = lean_alloc_ctor(0, 0, 17); +} else { + x_554 = x_552; +} +lean_ctor_set_uint8(x_554, 0, x_536); +lean_ctor_set_uint8(x_554, 1, x_537); +lean_ctor_set_uint8(x_554, 2, x_538); +lean_ctor_set_uint8(x_554, 3, x_539); +lean_ctor_set_uint8(x_554, 4, x_540); +lean_ctor_set_uint8(x_554, 5, x_541); +lean_ctor_set_uint8(x_554, 6, x_542); +lean_ctor_set_uint8(x_554, 7, x_543); +lean_ctor_set_uint8(x_554, 8, x_544); +lean_ctor_set_uint8(x_554, 9, x_553); +lean_ctor_set_uint8(x_554, 10, x_545); +lean_ctor_set_uint8(x_554, 11, x_546); +lean_ctor_set_uint8(x_554, 12, x_547); +lean_ctor_set_uint8(x_554, 13, x_548); +lean_ctor_set_uint8(x_554, 14, x_549); +lean_ctor_set_uint8(x_554, 15, x_550); +lean_ctor_set_uint8(x_554, 16, x_551); +x_555 = 2; +x_556 = lean_uint64_shift_right(x_526, x_555); +x_557 = lean_uint64_shift_left(x_556, x_555); +x_558 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__13; +x_559 = lean_uint64_lor(x_557, x_558); +x_560 = lean_alloc_ctor(0, 7, 11); +lean_ctor_set(x_560, 0, x_554); +lean_ctor_set(x_560, 1, x_528); +lean_ctor_set(x_560, 2, x_529); +lean_ctor_set(x_560, 3, x_530); +lean_ctor_set(x_560, 4, x_531); +lean_ctor_set(x_560, 5, x_532); +lean_ctor_set(x_560, 6, x_533); +lean_ctor_set_uint64(x_560, sizeof(void*)*7, x_559); +lean_ctor_set_uint8(x_560, sizeof(void*)*7 + 8, x_527); +lean_ctor_set_uint8(x_560, sizeof(void*)*7 + 9, x_534); +lean_ctor_set_uint8(x_560, sizeof(void*)*7 + 10, x_535); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); -lean_inc(x_431); -lean_inc(x_461); -x_498 = l_Lean_Meta_isExprDefEq(x_461, x_431, x_497, x_9, x_10, x_11, x_460); -if (lean_obj_tag(x_498) == 0) -{ -lean_object* x_499; uint8_t x_500; -x_499 = lean_ctor_get(x_498, 0); -lean_inc(x_499); -x_500 = lean_unbox(x_499); -lean_dec(x_499); -if (x_500 == 0) +lean_inc(x_494); +lean_inc(x_524); +x_561 = l_Lean_Meta_isExprDefEq(x_524, x_494, x_560, x_9, x_10, x_11, x_523); +if (lean_obj_tag(x_561) == 0) +{ +lean_object* x_562; uint8_t x_563; +x_562 = lean_ctor_get(x_561, 0); +lean_inc(x_562); +x_563 = lean_unbox(x_562); +lean_dec(x_562); +if (x_563 == 0) +{ +lean_object* x_564; lean_object* x_565; +x_564 = lean_ctor_get(x_561, 1); +lean_inc(x_564); +lean_dec(x_561); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +x_565 = l_Lean_Meta_addPPExplicitToExposeDiff(x_524, x_494, x_8, x_9, x_10, x_11, x_564); +if (lean_obj_tag(x_565) == 0) { -lean_object* x_501; lean_object* x_502; lean_object* x_503; lean_object* x_504; lean_object* x_505; lean_object* x_506; lean_object* x_507; lean_object* x_508; lean_object* x_509; lean_object* x_510; lean_object* x_511; -x_501 = lean_ctor_get(x_498, 1); -lean_inc(x_501); -lean_dec(x_498); -x_502 = l_Lean_indentExpr(x_461); -x_503 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__15; -if (lean_is_scalar(x_462)) { - x_504 = lean_alloc_ctor(7, 2, 0); -} else { - x_504 = x_462; - lean_ctor_set_tag(x_504, 7); -} -lean_ctor_set(x_504, 0, x_503); -lean_ctor_set(x_504, 1, x_502); -x_505 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__17; -if (lean_is_scalar(x_453)) { - x_506 = lean_alloc_ctor(7, 2, 0); -} else { - x_506 = x_453; - lean_ctor_set_tag(x_506, 7); -} -lean_ctor_set(x_506, 0, x_504); -lean_ctor_set(x_506, 1, x_505); -x_507 = l_Lean_indentExpr(x_431); -if (lean_is_scalar(x_439)) { - x_508 = lean_alloc_ctor(7, 2, 0); +lean_object* x_566; lean_object* x_567; lean_object* x_568; lean_object* x_569; lean_object* x_570; lean_object* x_571; lean_object* x_572; lean_object* x_573; lean_object* x_574; lean_object* x_575; lean_object* x_576; lean_object* x_577; lean_object* x_578; lean_object* x_579; lean_object* x_580; +x_566 = lean_ctor_get(x_565, 0); +lean_inc(x_566); +x_567 = lean_ctor_get(x_565, 1); +lean_inc(x_567); +lean_dec(x_565); +x_568 = lean_ctor_get(x_566, 0); +lean_inc(x_568); +x_569 = lean_ctor_get(x_566, 1); +lean_inc(x_569); +if (lean_is_exclusive(x_566)) { + lean_ctor_release(x_566, 0); + lean_ctor_release(x_566, 1); + x_570 = x_566; } else { - x_508 = x_439; - lean_ctor_set_tag(x_508, 7); + lean_dec_ref(x_566); + x_570 = lean_box(0); +} +x_571 = l_Lean_indentExpr(x_568); +x_572 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__15; +if (lean_is_scalar(x_570)) { + x_573 = lean_alloc_ctor(7, 2, 0); +} else { + x_573 = x_570; + lean_ctor_set_tag(x_573, 7); +} +lean_ctor_set(x_573, 0, x_572); +lean_ctor_set(x_573, 1, x_571); +x_574 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__17; +if (lean_is_scalar(x_525)) { + x_575 = lean_alloc_ctor(7, 2, 0); +} else { + x_575 = x_525; + lean_ctor_set_tag(x_575, 7); +} +lean_ctor_set(x_575, 0, x_573); +lean_ctor_set(x_575, 1, x_574); +x_576 = l_Lean_indentExpr(x_569); +if (lean_is_scalar(x_516)) { + x_577 = lean_alloc_ctor(7, 2, 0); +} else { + x_577 = x_516; + lean_ctor_set_tag(x_577, 7); +} +lean_ctor_set(x_577, 0, x_575); +lean_ctor_set(x_577, 1, x_576); +x_578 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__8; +if (lean_is_scalar(x_502)) { + x_579 = lean_alloc_ctor(7, 2, 0); +} else { + x_579 = x_502; + lean_ctor_set_tag(x_579, 7); +} +lean_ctor_set(x_579, 0, x_577); +lean_ctor_set(x_579, 1, x_578); +x_580 = l_Lean_throwErrorAt___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_throwStuckAtUniverseCnstr___spec__12(x_4, x_579, x_6, x_7, x_8, x_9, x_10, x_11, x_567); +lean_dec(x_11); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +return x_580; } -lean_ctor_set(x_508, 0, x_506); -lean_ctor_set(x_508, 1, x_507); -x_509 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__8; -x_510 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_510, 0, x_508); -lean_ctor_set(x_510, 1, x_509); -x_511 = l_Lean_throwErrorAt___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_throwStuckAtUniverseCnstr___spec__12(x_4, x_510, x_6, x_7, x_8, x_9, x_10, x_11, x_501); +else +{ +lean_object* x_581; lean_object* x_582; lean_object* x_583; lean_object* x_584; +lean_dec(x_525); +lean_dec(x_516); +lean_dec(x_502); lean_dec(x_11); +lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -return x_511; +lean_dec(x_6); +x_581 = lean_ctor_get(x_565, 0); +lean_inc(x_581); +x_582 = lean_ctor_get(x_565, 1); +lean_inc(x_582); +if (lean_is_exclusive(x_565)) { + lean_ctor_release(x_565, 0); + lean_ctor_release(x_565, 1); + x_583 = x_565; +} else { + lean_dec_ref(x_565); + x_583 = lean_box(0); +} +if (lean_is_scalar(x_583)) { + x_584 = lean_alloc_ctor(1, 2, 0); +} else { + x_584 = x_583; +} +lean_ctor_set(x_584, 0, x_581); +lean_ctor_set(x_584, 1, x_582); +return x_584; +} } else { -lean_object* x_512; lean_object* x_513; lean_object* x_514; lean_object* x_515; -lean_dec(x_462); -lean_dec(x_461); -lean_dec(x_453); -lean_dec(x_439); -lean_dec(x_431); +lean_object* x_585; lean_object* x_586; lean_object* x_587; lean_object* x_588; +lean_dec(x_525); +lean_dec(x_524); +lean_dec(x_516); +lean_dec(x_502); +lean_dec(x_494); 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_512 = lean_ctor_get(x_498, 1); -lean_inc(x_512); -if (lean_is_exclusive(x_498)) { - lean_ctor_release(x_498, 0); - lean_ctor_release(x_498, 1); - x_513 = x_498; +x_585 = lean_ctor_get(x_561, 1); +lean_inc(x_585); +if (lean_is_exclusive(x_561)) { + lean_ctor_release(x_561, 0); + lean_ctor_release(x_561, 1); + x_586 = x_561; } else { - lean_dec_ref(x_498); - x_513 = lean_box(0); + lean_dec_ref(x_561); + x_586 = lean_box(0); } -x_514 = lean_box(0); -if (lean_is_scalar(x_513)) { - x_515 = lean_alloc_ctor(0, 2, 0); +x_587 = lean_box(0); +if (lean_is_scalar(x_586)) { + x_588 = lean_alloc_ctor(0, 2, 0); } else { - x_515 = x_513; + x_588 = x_586; } -lean_ctor_set(x_515, 0, x_514); -lean_ctor_set(x_515, 1, x_512); -return x_515; +lean_ctor_set(x_588, 0, x_587); +lean_ctor_set(x_588, 1, x_585); +return x_588; } } else { -lean_object* x_516; lean_object* x_517; lean_object* x_518; lean_object* x_519; -lean_dec(x_462); -lean_dec(x_461); -lean_dec(x_453); -lean_dec(x_439); -lean_dec(x_431); +lean_object* x_589; lean_object* x_590; lean_object* x_591; lean_object* x_592; +lean_dec(x_525); +lean_dec(x_524); +lean_dec(x_516); +lean_dec(x_502); +lean_dec(x_494); 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_516 = lean_ctor_get(x_498, 0); -lean_inc(x_516); -x_517 = lean_ctor_get(x_498, 1); -lean_inc(x_517); -if (lean_is_exclusive(x_498)) { - lean_ctor_release(x_498, 0); - lean_ctor_release(x_498, 1); - x_518 = x_498; +x_589 = lean_ctor_get(x_561, 0); +lean_inc(x_589); +x_590 = lean_ctor_get(x_561, 1); +lean_inc(x_590); +if (lean_is_exclusive(x_561)) { + lean_ctor_release(x_561, 0); + lean_ctor_release(x_561, 1); + x_591 = x_561; } else { - lean_dec_ref(x_498); - x_518 = lean_box(0); + lean_dec_ref(x_561); + x_591 = lean_box(0); } -if (lean_is_scalar(x_518)) { - x_519 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_591)) { + x_592 = lean_alloc_ctor(1, 2, 0); } else { - x_519 = x_518; + x_592 = x_591; } -lean_ctor_set(x_519, 0, x_516); -lean_ctor_set(x_519, 1, x_517); -return x_519; +lean_ctor_set(x_592, 0, x_589); +lean_ctor_set(x_592, 1, x_590); +return x_592; } } else { -lean_object* x_520; lean_object* x_521; lean_object* x_522; lean_object* x_523; -lean_dec(x_453); -lean_dec(x_439); -lean_dec(x_431); +lean_object* x_593; lean_object* x_594; lean_object* x_595; lean_object* x_596; +lean_dec(x_516); +lean_dec(x_502); +lean_dec(x_494); 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_520 = lean_ctor_get(x_457, 0); -lean_inc(x_520); -x_521 = lean_ctor_get(x_457, 1); -lean_inc(x_521); -if (lean_is_exclusive(x_457)) { - lean_ctor_release(x_457, 0); - lean_ctor_release(x_457, 1); - x_522 = x_457; +x_593 = lean_ctor_get(x_520, 0); +lean_inc(x_593); +x_594 = lean_ctor_get(x_520, 1); +lean_inc(x_594); +if (lean_is_exclusive(x_520)) { + lean_ctor_release(x_520, 0); + lean_ctor_release(x_520, 1); + x_595 = x_520; } else { - lean_dec_ref(x_457); - x_522 = lean_box(0); + lean_dec_ref(x_520); + x_595 = lean_box(0); } -if (lean_is_scalar(x_522)) { - x_523 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_595)) { + x_596 = lean_alloc_ctor(1, 2, 0); } else { - x_523 = x_522; + x_596 = x_595; } -lean_ctor_set(x_523, 0, x_520); -lean_ctor_set(x_523, 1, x_521); -return x_523; +lean_ctor_set(x_596, 0, x_593); +lean_ctor_set(x_596, 1, x_594); +return x_596; } } else { -lean_object* x_524; lean_object* x_525; lean_object* x_526; lean_object* x_527; -lean_dec(x_453); -lean_dec(x_439); -lean_dec(x_431); +lean_object* x_597; lean_object* x_598; lean_object* x_599; lean_object* x_600; +lean_dec(x_516); +lean_dec(x_502); +lean_dec(x_494); 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_524 = lean_ctor_get(x_454, 0); -lean_inc(x_524); -x_525 = lean_ctor_get(x_454, 1); -lean_inc(x_525); -if (lean_is_exclusive(x_454)) { - lean_ctor_release(x_454, 0); - lean_ctor_release(x_454, 1); - x_526 = x_454; +x_597 = lean_ctor_get(x_517, 0); +lean_inc(x_597); +x_598 = lean_ctor_get(x_517, 1); +lean_inc(x_598); +if (lean_is_exclusive(x_517)) { + lean_ctor_release(x_517, 0); + lean_ctor_release(x_517, 1); + x_599 = x_517; } else { - lean_dec_ref(x_454); - x_526 = lean_box(0); + lean_dec_ref(x_517); + x_599 = lean_box(0); } -if (lean_is_scalar(x_526)) { - x_527 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_599)) { + x_600 = lean_alloc_ctor(1, 2, 0); } else { - x_527 = x_526; + x_600 = x_599; } -lean_ctor_set(x_527, 0, x_524); -lean_ctor_set(x_527, 1, x_525); -return x_527; +lean_ctor_set(x_600, 0, x_597); +lean_ctor_set(x_600, 1, x_598); +return x_600; } } else { -lean_object* x_528; lean_object* x_529; lean_object* x_530; lean_object* x_531; lean_object* x_532; lean_object* x_533; lean_object* x_534; lean_object* x_535; lean_object* x_536; lean_object* x_537; lean_object* x_538; lean_object* x_539; lean_object* x_540; -if (lean_is_exclusive(x_438)) { - lean_ctor_release(x_438, 0); - lean_ctor_release(x_438, 1); - x_528 = x_438; +lean_object* x_601; lean_object* x_602; lean_object* x_603; lean_object* x_604; lean_object* x_605; lean_object* x_606; lean_object* x_607; lean_object* x_608; lean_object* x_609; lean_object* x_610; lean_object* x_611; lean_object* x_612; lean_object* x_613; +if (lean_is_exclusive(x_501)) { + lean_ctor_release(x_501, 0); + lean_ctor_release(x_501, 1); + x_601 = x_501; } else { - lean_dec_ref(x_438); - x_528 = lean_box(0); + lean_dec_ref(x_501); + x_601 = lean_box(0); } -if (lean_is_exclusive(x_450)) { - lean_ctor_release(x_450, 0); - lean_ctor_release(x_450, 1); - x_529 = x_450; +if (lean_is_exclusive(x_513)) { + lean_ctor_release(x_513, 0); + lean_ctor_release(x_513, 1); + x_602 = x_513; } else { - lean_dec_ref(x_450); - x_529 = lean_box(0); + lean_dec_ref(x_513); + x_602 = lean_box(0); } -x_530 = lean_ctor_get(x_436, 1); -lean_inc(x_530); -lean_dec(x_436); -x_531 = l_Lean_MessageData_ofSyntax(x_3); -x_532 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__8; -if (lean_is_scalar(x_529)) { - x_533 = lean_alloc_ctor(7, 2, 0); -} else { - x_533 = x_529; - lean_ctor_set_tag(x_533, 7); -} -lean_ctor_set(x_533, 0, x_532); -lean_ctor_set(x_533, 1, x_531); -x_534 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__19; -if (lean_is_scalar(x_528)) { - x_535 = lean_alloc_ctor(7, 2, 0); -} else { - x_535 = x_528; - lean_ctor_set_tag(x_535, 7); -} -lean_ctor_set(x_535, 0, x_533); -lean_ctor_set(x_535, 1, x_534); -x_536 = l_Lean_indentExpr(x_431); -if (lean_is_scalar(x_439)) { - x_537 = lean_alloc_ctor(7, 2, 0); -} else { - x_537 = x_439; - lean_ctor_set_tag(x_537, 7); -} -lean_ctor_set(x_537, 0, x_535); -lean_ctor_set(x_537, 1, x_536); -x_538 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__12; -x_539 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_539, 0, x_537); -lean_ctor_set(x_539, 1, x_538); -x_540 = l_Lean_throwErrorAt___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_throwStuckAtUniverseCnstr___spec__12(x_4, x_539, x_6, x_7, x_8, x_9, x_10, x_11, x_530); +x_603 = lean_ctor_get(x_499, 1); +lean_inc(x_603); +lean_dec(x_499); +x_604 = l_Lean_MessageData_ofSyntax(x_3); +x_605 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__8; +if (lean_is_scalar(x_602)) { + x_606 = lean_alloc_ctor(7, 2, 0); +} else { + x_606 = x_602; + lean_ctor_set_tag(x_606, 7); +} +lean_ctor_set(x_606, 0, x_605); +lean_ctor_set(x_606, 1, x_604); +x_607 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__19; +if (lean_is_scalar(x_601)) { + x_608 = lean_alloc_ctor(7, 2, 0); +} else { + x_608 = x_601; + lean_ctor_set_tag(x_608, 7); +} +lean_ctor_set(x_608, 0, x_606); +lean_ctor_set(x_608, 1, x_607); +x_609 = l_Lean_indentExpr(x_494); +if (lean_is_scalar(x_502)) { + x_610 = lean_alloc_ctor(7, 2, 0); +} else { + x_610 = x_502; + lean_ctor_set_tag(x_610, 7); +} +lean_ctor_set(x_610, 0, x_608); +lean_ctor_set(x_610, 1, x_609); +x_611 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__12; +x_612 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_612, 0, x_610); +lean_ctor_set(x_612, 1, x_611); +x_613 = l_Lean_throwErrorAt___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_throwStuckAtUniverseCnstr___spec__12(x_4, x_612, x_6, x_7, x_8, x_9, x_10, x_11, x_603); lean_dec(x_11); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -return x_540; +return x_613; } } } else { -lean_object* x_541; lean_object* x_542; lean_object* x_543; lean_object* x_544; -lean_dec(x_431); +lean_object* x_614; lean_object* x_615; lean_object* x_616; lean_object* x_617; +lean_dec(x_494); lean_free_object(x_25); lean_dec(x_11); lean_dec(x_10); @@ -2617,32 +3004,32 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_3); -x_541 = lean_ctor_get(x_436, 0); -lean_inc(x_541); -x_542 = lean_ctor_get(x_436, 1); -lean_inc(x_542); -if (lean_is_exclusive(x_436)) { - lean_ctor_release(x_436, 0); - lean_ctor_release(x_436, 1); - x_543 = x_436; +x_614 = lean_ctor_get(x_499, 0); +lean_inc(x_614); +x_615 = lean_ctor_get(x_499, 1); +lean_inc(x_615); +if (lean_is_exclusive(x_499)) { + lean_ctor_release(x_499, 0); + lean_ctor_release(x_499, 1); + x_616 = x_499; } else { - lean_dec_ref(x_436); - x_543 = lean_box(0); + lean_dec_ref(x_499); + x_616 = lean_box(0); } -if (lean_is_scalar(x_543)) { - x_544 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_616)) { + x_617 = lean_alloc_ctor(1, 2, 0); } else { - x_544 = x_543; + x_617 = x_616; } -lean_ctor_set(x_544, 0, x_541); -lean_ctor_set(x_544, 1, x_542); -return x_544; +lean_ctor_set(x_617, 0, x_614); +lean_ctor_set(x_617, 1, x_615); +return x_617; } } else { -lean_object* x_545; lean_object* x_546; lean_object* x_547; lean_object* x_548; -lean_dec(x_427); +lean_object* x_618; lean_object* x_619; lean_object* x_620; lean_object* x_621; +lean_dec(x_490); lean_free_object(x_25); lean_dec(x_11); lean_dec(x_10); @@ -2651,550 +3038,606 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_3); -x_545 = lean_ctor_get(x_430, 0); -lean_inc(x_545); -x_546 = lean_ctor_get(x_430, 1); -lean_inc(x_546); -if (lean_is_exclusive(x_430)) { - lean_ctor_release(x_430, 0); - lean_ctor_release(x_430, 1); - x_547 = x_430; +x_618 = lean_ctor_get(x_493, 0); +lean_inc(x_618); +x_619 = lean_ctor_get(x_493, 1); +lean_inc(x_619); +if (lean_is_exclusive(x_493)) { + lean_ctor_release(x_493, 0); + lean_ctor_release(x_493, 1); + x_620 = x_493; } else { - lean_dec_ref(x_430); - x_547 = lean_box(0); + lean_dec_ref(x_493); + x_620 = lean_box(0); } -if (lean_is_scalar(x_547)) { - x_548 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_620)) { + x_621 = lean_alloc_ctor(1, 2, 0); } else { - x_548 = x_547; + x_621 = x_620; } -lean_ctor_set(x_548, 0, x_545); -lean_ctor_set(x_548, 1, x_546); -return x_548; +lean_ctor_set(x_621, 0, x_618); +lean_ctor_set(x_621, 1, x_619); +return x_621; } } } else { -lean_object* x_549; lean_object* x_550; lean_object* x_551; uint8_t x_552; lean_object* x_553; lean_object* x_554; lean_object* x_555; lean_object* x_556; lean_object* x_557; lean_object* x_558; lean_object* x_559; -x_549 = lean_ctor_get(x_25, 0); -x_550 = lean_ctor_get(x_25, 1); -lean_inc(x_550); -lean_inc(x_549); +lean_object* x_622; lean_object* x_623; lean_object* x_624; uint8_t x_625; lean_object* x_626; lean_object* x_627; lean_object* x_628; lean_object* x_629; lean_object* x_630; lean_object* x_631; lean_object* x_632; +x_622 = lean_ctor_get(x_25, 0); +x_623 = lean_ctor_get(x_25, 1); +lean_inc(x_623); +lean_inc(x_622); lean_dec(x_25); -x_551 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_551, 0, x_549); -x_552 = 0; -x_553 = lean_box(0); +x_624 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_624, 0, x_622); +x_625 = 0; +x_626 = lean_box(0); lean_inc(x_8); -x_554 = l___private_Lean_Meta_Basic_0__Lean_Meta_mkFreshExprMVarImpl(x_551, x_552, x_553, x_8, x_9, x_10, x_11, x_550); -x_555 = lean_ctor_get(x_554, 0); -lean_inc(x_555); -x_556 = lean_ctor_get(x_554, 1); -lean_inc(x_556); -if (lean_is_exclusive(x_554)) { - lean_ctor_release(x_554, 0); - lean_ctor_release(x_554, 1); - x_557 = x_554; -} else { - lean_dec_ref(x_554); - x_557 = lean_box(0); -} -x_558 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_558, 0, x_23); +x_627 = l___private_Lean_Meta_Basic_0__Lean_Meta_mkFreshExprMVarImpl(x_624, x_625, x_626, x_8, x_9, x_10, x_11, x_623); +x_628 = lean_ctor_get(x_627, 0); +lean_inc(x_628); +x_629 = lean_ctor_get(x_627, 1); +lean_inc(x_629); +if (lean_is_exclusive(x_627)) { + lean_ctor_release(x_627, 0); + lean_ctor_release(x_627, 1); + x_630 = x_627; +} else { + lean_dec_ref(x_627); + x_630 = lean_box(0); +} +x_631 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_631, 0, x_23); 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_559 = l_Lean_Elab_Term_elabTerm(x_2, x_558, x_14, x_14, x_6, x_7, x_8, x_9, x_10, x_11, x_556); -if (lean_obj_tag(x_559) == 0) -{ -lean_object* x_560; lean_object* x_561; lean_object* x_562; lean_object* x_563; lean_object* x_564; lean_object* x_565; -x_560 = lean_ctor_get(x_559, 0); -lean_inc(x_560); -x_561 = lean_ctor_get(x_559, 1); -lean_inc(x_561); -lean_dec(x_559); -x_562 = l_Lean_Expr_mvarId_x21(x_555); -lean_dec(x_555); -x_563 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__5; -x_564 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__6; +x_632 = l_Lean_Elab_Term_elabTerm(x_2, x_631, x_14, x_14, x_6, x_7, x_8, x_9, x_10, x_11, x_629); +if (lean_obj_tag(x_632) == 0) +{ +lean_object* x_633; lean_object* x_634; lean_object* x_635; lean_object* x_636; lean_object* x_637; lean_object* x_638; +x_633 = lean_ctor_get(x_632, 0); +lean_inc(x_633); +x_634 = lean_ctor_get(x_632, 1); +lean_inc(x_634); +lean_dec(x_632); +x_635 = l_Lean_Expr_mvarId_x21(x_628); +lean_dec(x_628); +x_636 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__5; +x_637 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__6; lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_3); -x_565 = l_Lean_Elab_runTactic(x_562, x_3, x_563, x_564, x_8, x_9, x_10, x_11, x_561); -if (lean_obj_tag(x_565) == 0) -{ -lean_object* x_566; lean_object* x_567; lean_object* x_568; -x_566 = lean_ctor_get(x_565, 0); -lean_inc(x_566); -x_567 = lean_ctor_get(x_566, 0); -lean_inc(x_567); -if (lean_is_exclusive(x_566)) { - lean_ctor_release(x_566, 0); - lean_ctor_release(x_566, 1); - x_568 = x_566; +x_638 = l_Lean_Elab_runTactic(x_635, x_3, x_636, x_637, x_8, x_9, x_10, x_11, x_634); +if (lean_obj_tag(x_638) == 0) +{ +lean_object* x_639; lean_object* x_640; lean_object* x_641; +x_639 = lean_ctor_get(x_638, 0); +lean_inc(x_639); +x_640 = lean_ctor_get(x_639, 0); +lean_inc(x_640); +if (lean_is_exclusive(x_639)) { + lean_ctor_release(x_639, 0); + lean_ctor_release(x_639, 1); + x_641 = x_639; +} else { + lean_dec_ref(x_639); + x_641 = lean_box(0); +} +if (lean_obj_tag(x_640) == 0) +{ +lean_object* x_642; lean_object* x_643; lean_object* x_644; lean_object* x_645; lean_object* x_646; lean_object* x_647; lean_object* x_648; lean_object* x_649; lean_object* x_650; lean_object* x_651; lean_object* x_652; +x_642 = lean_ctor_get(x_638, 1); +lean_inc(x_642); +lean_dec(x_638); +x_643 = l_Lean_MessageData_ofSyntax(x_3); +x_644 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__8; +if (lean_is_scalar(x_641)) { + x_645 = lean_alloc_ctor(7, 2, 0); } else { - lean_dec_ref(x_566); - x_568 = lean_box(0); + x_645 = x_641; + lean_ctor_set_tag(x_645, 7); } -if (lean_obj_tag(x_567) == 0) -{ -lean_object* x_569; lean_object* x_570; lean_object* x_571; lean_object* x_572; lean_object* x_573; lean_object* x_574; lean_object* x_575; lean_object* x_576; lean_object* x_577; lean_object* x_578; lean_object* x_579; -x_569 = lean_ctor_get(x_565, 1); -lean_inc(x_569); -lean_dec(x_565); -x_570 = l_Lean_MessageData_ofSyntax(x_3); -x_571 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__8; -if (lean_is_scalar(x_568)) { - x_572 = lean_alloc_ctor(7, 2, 0); -} else { - x_572 = x_568; - lean_ctor_set_tag(x_572, 7); -} -lean_ctor_set(x_572, 0, x_571); -lean_ctor_set(x_572, 1, x_570); -x_573 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__10; -if (lean_is_scalar(x_557)) { - x_574 = lean_alloc_ctor(7, 2, 0); -} else { - x_574 = x_557; - lean_ctor_set_tag(x_574, 7); -} -lean_ctor_set(x_574, 0, x_572); -lean_ctor_set(x_574, 1, x_573); -x_575 = l_Lean_indentExpr(x_560); -x_576 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_576, 0, x_574); -lean_ctor_set(x_576, 1, x_575); -x_577 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__12; -x_578 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_578, 0, x_576); -lean_ctor_set(x_578, 1, x_577); -x_579 = l_Lean_throwErrorAt___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_throwStuckAtUniverseCnstr___spec__12(x_4, x_578, x_6, x_7, x_8, x_9, x_10, x_11, x_569); +lean_ctor_set(x_645, 0, x_644); +lean_ctor_set(x_645, 1, x_643); +x_646 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__10; +if (lean_is_scalar(x_630)) { + x_647 = lean_alloc_ctor(7, 2, 0); +} else { + x_647 = x_630; + lean_ctor_set_tag(x_647, 7); +} +lean_ctor_set(x_647, 0, x_645); +lean_ctor_set(x_647, 1, x_646); +x_648 = l_Lean_indentExpr(x_633); +x_649 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_649, 0, x_647); +lean_ctor_set(x_649, 1, x_648); +x_650 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__12; +x_651 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_651, 0, x_649); +lean_ctor_set(x_651, 1, x_650); +x_652 = l_Lean_throwErrorAt___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_throwStuckAtUniverseCnstr___spec__12(x_4, x_651, x_6, x_7, x_8, x_9, x_10, x_11, x_642); lean_dec(x_11); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -return x_579; +return x_652; } else { -lean_object* x_580; -x_580 = lean_ctor_get(x_567, 1); -lean_inc(x_580); -if (lean_obj_tag(x_580) == 0) +lean_object* x_653; +x_653 = lean_ctor_get(x_640, 1); +lean_inc(x_653); +if (lean_obj_tag(x_653) == 0) { -lean_object* x_581; lean_object* x_582; lean_object* x_583; lean_object* x_584; +lean_object* x_654; lean_object* x_655; lean_object* x_656; lean_object* x_657; +lean_dec(x_630); lean_dec(x_3); -x_581 = lean_ctor_get(x_565, 1); -lean_inc(x_581); -lean_dec(x_565); -x_582 = lean_ctor_get(x_567, 0); -lean_inc(x_582); -if (lean_is_exclusive(x_567)) { - lean_ctor_release(x_567, 0); - lean_ctor_release(x_567, 1); - x_583 = x_567; +x_654 = lean_ctor_get(x_638, 1); +lean_inc(x_654); +lean_dec(x_638); +x_655 = lean_ctor_get(x_640, 0); +lean_inc(x_655); +if (lean_is_exclusive(x_640)) { + lean_ctor_release(x_640, 0); + lean_ctor_release(x_640, 1); + x_656 = x_640; } else { - lean_dec_ref(x_567); - x_583 = lean_box(0); + lean_dec_ref(x_640); + x_656 = lean_box(0); } -x_584 = l_Lean_MVarId_getType(x_582, x_8, x_9, x_10, x_11, x_581); -if (lean_obj_tag(x_584) == 0) +x_657 = l_Lean_MVarId_getType(x_655, x_8, x_9, x_10, x_11, x_654); +if (lean_obj_tag(x_657) == 0) { -lean_object* x_585; lean_object* x_586; lean_object* x_587; -x_585 = lean_ctor_get(x_584, 0); -lean_inc(x_585); -x_586 = lean_ctor_get(x_584, 1); -lean_inc(x_586); -lean_dec(x_584); +lean_object* x_658; lean_object* x_659; lean_object* x_660; +x_658 = lean_ctor_get(x_657, 0); +lean_inc(x_658); +x_659 = lean_ctor_get(x_657, 1); +lean_inc(x_659); +lean_dec(x_657); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); -x_587 = l_Lean_Meta_CheckTactic_matchCheckGoalType(x_4, x_585, x_8, x_9, x_10, x_11, x_586); -if (lean_obj_tag(x_587) == 0) -{ -lean_object* x_588; lean_object* x_589; lean_object* x_590; lean_object* x_591; lean_object* x_592; uint64_t x_593; uint8_t x_594; lean_object* x_595; lean_object* x_596; lean_object* x_597; lean_object* x_598; lean_object* x_599; lean_object* x_600; uint8_t x_601; uint8_t x_602; uint8_t x_603; uint8_t x_604; uint8_t x_605; uint8_t x_606; uint8_t x_607; uint8_t x_608; uint8_t x_609; uint8_t x_610; uint8_t x_611; uint8_t x_612; uint8_t x_613; uint8_t x_614; uint8_t x_615; uint8_t x_616; uint8_t x_617; uint8_t x_618; lean_object* x_619; uint8_t x_620; lean_object* x_621; uint64_t x_622; uint64_t x_623; uint64_t x_624; uint64_t x_625; uint64_t x_626; lean_object* x_627; lean_object* x_628; -x_588 = lean_ctor_get(x_587, 0); -lean_inc(x_588); -x_589 = lean_ctor_get(x_8, 0); -lean_inc(x_589); -x_590 = lean_ctor_get(x_587, 1); -lean_inc(x_590); -lean_dec(x_587); -x_591 = lean_ctor_get(x_588, 0); -lean_inc(x_591); -if (lean_is_exclusive(x_588)) { - lean_ctor_release(x_588, 0); - lean_ctor_release(x_588, 1); - x_592 = x_588; -} else { - lean_dec_ref(x_588); - x_592 = lean_box(0); -} -x_593 = lean_ctor_get_uint64(x_8, sizeof(void*)*7); -x_594 = lean_ctor_get_uint8(x_8, sizeof(void*)*7 + 8); -x_595 = lean_ctor_get(x_8, 1); -lean_inc(x_595); -x_596 = lean_ctor_get(x_8, 2); -lean_inc(x_596); -x_597 = lean_ctor_get(x_8, 3); -lean_inc(x_597); -x_598 = lean_ctor_get(x_8, 4); -lean_inc(x_598); -x_599 = lean_ctor_get(x_8, 5); -lean_inc(x_599); -x_600 = lean_ctor_get(x_8, 6); -lean_inc(x_600); -x_601 = lean_ctor_get_uint8(x_8, sizeof(void*)*7 + 9); -x_602 = lean_ctor_get_uint8(x_8, sizeof(void*)*7 + 10); -x_603 = lean_ctor_get_uint8(x_589, 0); -x_604 = lean_ctor_get_uint8(x_589, 1); -x_605 = lean_ctor_get_uint8(x_589, 2); -x_606 = lean_ctor_get_uint8(x_589, 3); -x_607 = lean_ctor_get_uint8(x_589, 4); -x_608 = lean_ctor_get_uint8(x_589, 5); -x_609 = lean_ctor_get_uint8(x_589, 6); -x_610 = lean_ctor_get_uint8(x_589, 7); -x_611 = lean_ctor_get_uint8(x_589, 8); -x_612 = lean_ctor_get_uint8(x_589, 10); -x_613 = lean_ctor_get_uint8(x_589, 11); -x_614 = lean_ctor_get_uint8(x_589, 12); -x_615 = lean_ctor_get_uint8(x_589, 13); -x_616 = lean_ctor_get_uint8(x_589, 14); -x_617 = lean_ctor_get_uint8(x_589, 15); -x_618 = lean_ctor_get_uint8(x_589, 16); -if (lean_is_exclusive(x_589)) { - x_619 = x_589; -} else { - lean_dec_ref(x_589); - x_619 = lean_box(0); -} -x_620 = 2; -if (lean_is_scalar(x_619)) { - x_621 = lean_alloc_ctor(0, 0, 17); -} else { - x_621 = x_619; -} -lean_ctor_set_uint8(x_621, 0, x_603); -lean_ctor_set_uint8(x_621, 1, x_604); -lean_ctor_set_uint8(x_621, 2, x_605); -lean_ctor_set_uint8(x_621, 3, x_606); -lean_ctor_set_uint8(x_621, 4, x_607); -lean_ctor_set_uint8(x_621, 5, x_608); -lean_ctor_set_uint8(x_621, 6, x_609); -lean_ctor_set_uint8(x_621, 7, x_610); -lean_ctor_set_uint8(x_621, 8, x_611); -lean_ctor_set_uint8(x_621, 9, x_620); -lean_ctor_set_uint8(x_621, 10, x_612); -lean_ctor_set_uint8(x_621, 11, x_613); -lean_ctor_set_uint8(x_621, 12, x_614); -lean_ctor_set_uint8(x_621, 13, x_615); -lean_ctor_set_uint8(x_621, 14, x_616); -lean_ctor_set_uint8(x_621, 15, x_617); -lean_ctor_set_uint8(x_621, 16, x_618); -x_622 = 2; -x_623 = lean_uint64_shift_right(x_593, x_622); -x_624 = lean_uint64_shift_left(x_623, x_622); -x_625 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__13; -x_626 = lean_uint64_lor(x_624, x_625); -x_627 = lean_alloc_ctor(0, 7, 11); -lean_ctor_set(x_627, 0, x_621); -lean_ctor_set(x_627, 1, x_595); -lean_ctor_set(x_627, 2, x_596); -lean_ctor_set(x_627, 3, x_597); -lean_ctor_set(x_627, 4, x_598); -lean_ctor_set(x_627, 5, x_599); -lean_ctor_set(x_627, 6, x_600); -lean_ctor_set_uint64(x_627, sizeof(void*)*7, x_626); -lean_ctor_set_uint8(x_627, sizeof(void*)*7 + 8, x_594); -lean_ctor_set_uint8(x_627, sizeof(void*)*7 + 9, x_601); -lean_ctor_set_uint8(x_627, sizeof(void*)*7 + 10, x_602); +x_660 = l_Lean_Meta_CheckTactic_matchCheckGoalType(x_4, x_658, x_8, x_9, x_10, x_11, x_659); +if (lean_obj_tag(x_660) == 0) +{ +lean_object* x_661; lean_object* x_662; lean_object* x_663; lean_object* x_664; lean_object* x_665; uint64_t x_666; uint8_t x_667; lean_object* x_668; lean_object* x_669; lean_object* x_670; lean_object* x_671; lean_object* x_672; lean_object* x_673; uint8_t x_674; uint8_t x_675; uint8_t x_676; uint8_t x_677; uint8_t x_678; uint8_t x_679; uint8_t x_680; uint8_t x_681; uint8_t x_682; uint8_t x_683; uint8_t x_684; uint8_t x_685; uint8_t x_686; uint8_t x_687; uint8_t x_688; uint8_t x_689; uint8_t x_690; uint8_t x_691; lean_object* x_692; uint8_t x_693; lean_object* x_694; uint64_t x_695; uint64_t x_696; uint64_t x_697; uint64_t x_698; uint64_t x_699; lean_object* x_700; lean_object* x_701; +x_661 = lean_ctor_get(x_660, 0); +lean_inc(x_661); +x_662 = lean_ctor_get(x_8, 0); +lean_inc(x_662); +x_663 = lean_ctor_get(x_660, 1); +lean_inc(x_663); +lean_dec(x_660); +x_664 = lean_ctor_get(x_661, 0); +lean_inc(x_664); +if (lean_is_exclusive(x_661)) { + lean_ctor_release(x_661, 0); + lean_ctor_release(x_661, 1); + x_665 = x_661; +} else { + lean_dec_ref(x_661); + x_665 = lean_box(0); +} +x_666 = lean_ctor_get_uint64(x_8, sizeof(void*)*7); +x_667 = lean_ctor_get_uint8(x_8, sizeof(void*)*7 + 8); +x_668 = lean_ctor_get(x_8, 1); +lean_inc(x_668); +x_669 = lean_ctor_get(x_8, 2); +lean_inc(x_669); +x_670 = lean_ctor_get(x_8, 3); +lean_inc(x_670); +x_671 = lean_ctor_get(x_8, 4); +lean_inc(x_671); +x_672 = lean_ctor_get(x_8, 5); +lean_inc(x_672); +x_673 = lean_ctor_get(x_8, 6); +lean_inc(x_673); +x_674 = lean_ctor_get_uint8(x_8, sizeof(void*)*7 + 9); +x_675 = lean_ctor_get_uint8(x_8, sizeof(void*)*7 + 10); +x_676 = lean_ctor_get_uint8(x_662, 0); +x_677 = lean_ctor_get_uint8(x_662, 1); +x_678 = lean_ctor_get_uint8(x_662, 2); +x_679 = lean_ctor_get_uint8(x_662, 3); +x_680 = lean_ctor_get_uint8(x_662, 4); +x_681 = lean_ctor_get_uint8(x_662, 5); +x_682 = lean_ctor_get_uint8(x_662, 6); +x_683 = lean_ctor_get_uint8(x_662, 7); +x_684 = lean_ctor_get_uint8(x_662, 8); +x_685 = lean_ctor_get_uint8(x_662, 10); +x_686 = lean_ctor_get_uint8(x_662, 11); +x_687 = lean_ctor_get_uint8(x_662, 12); +x_688 = lean_ctor_get_uint8(x_662, 13); +x_689 = lean_ctor_get_uint8(x_662, 14); +x_690 = lean_ctor_get_uint8(x_662, 15); +x_691 = lean_ctor_get_uint8(x_662, 16); +if (lean_is_exclusive(x_662)) { + x_692 = x_662; +} else { + lean_dec_ref(x_662); + x_692 = lean_box(0); +} +x_693 = 2; +if (lean_is_scalar(x_692)) { + x_694 = lean_alloc_ctor(0, 0, 17); +} else { + x_694 = x_692; +} +lean_ctor_set_uint8(x_694, 0, x_676); +lean_ctor_set_uint8(x_694, 1, x_677); +lean_ctor_set_uint8(x_694, 2, x_678); +lean_ctor_set_uint8(x_694, 3, x_679); +lean_ctor_set_uint8(x_694, 4, x_680); +lean_ctor_set_uint8(x_694, 5, x_681); +lean_ctor_set_uint8(x_694, 6, x_682); +lean_ctor_set_uint8(x_694, 7, x_683); +lean_ctor_set_uint8(x_694, 8, x_684); +lean_ctor_set_uint8(x_694, 9, x_693); +lean_ctor_set_uint8(x_694, 10, x_685); +lean_ctor_set_uint8(x_694, 11, x_686); +lean_ctor_set_uint8(x_694, 12, x_687); +lean_ctor_set_uint8(x_694, 13, x_688); +lean_ctor_set_uint8(x_694, 14, x_689); +lean_ctor_set_uint8(x_694, 15, x_690); +lean_ctor_set_uint8(x_694, 16, x_691); +x_695 = 2; +x_696 = lean_uint64_shift_right(x_666, x_695); +x_697 = lean_uint64_shift_left(x_696, x_695); +x_698 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__13; +x_699 = lean_uint64_lor(x_697, x_698); +x_700 = lean_alloc_ctor(0, 7, 11); +lean_ctor_set(x_700, 0, x_694); +lean_ctor_set(x_700, 1, x_668); +lean_ctor_set(x_700, 2, x_669); +lean_ctor_set(x_700, 3, x_670); +lean_ctor_set(x_700, 4, x_671); +lean_ctor_set(x_700, 5, x_672); +lean_ctor_set(x_700, 6, x_673); +lean_ctor_set_uint64(x_700, sizeof(void*)*7, x_699); +lean_ctor_set_uint8(x_700, sizeof(void*)*7 + 8, x_667); +lean_ctor_set_uint8(x_700, sizeof(void*)*7 + 9, x_674); +lean_ctor_set_uint8(x_700, sizeof(void*)*7 + 10, x_675); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); -lean_inc(x_560); -lean_inc(x_591); -x_628 = l_Lean_Meta_isExprDefEq(x_591, x_560, x_627, x_9, x_10, x_11, x_590); -if (lean_obj_tag(x_628) == 0) -{ -lean_object* x_629; uint8_t x_630; -x_629 = lean_ctor_get(x_628, 0); -lean_inc(x_629); -x_630 = lean_unbox(x_629); -lean_dec(x_629); -if (x_630 == 0) -{ -lean_object* x_631; lean_object* x_632; lean_object* x_633; lean_object* x_634; lean_object* x_635; lean_object* x_636; lean_object* x_637; lean_object* x_638; lean_object* x_639; lean_object* x_640; lean_object* x_641; -x_631 = lean_ctor_get(x_628, 1); -lean_inc(x_631); -lean_dec(x_628); -x_632 = l_Lean_indentExpr(x_591); -x_633 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__15; -if (lean_is_scalar(x_592)) { - x_634 = lean_alloc_ctor(7, 2, 0); -} else { - x_634 = x_592; - lean_ctor_set_tag(x_634, 7); -} -lean_ctor_set(x_634, 0, x_633); -lean_ctor_set(x_634, 1, x_632); -x_635 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__17; -if (lean_is_scalar(x_583)) { - x_636 = lean_alloc_ctor(7, 2, 0); -} else { - x_636 = x_583; - lean_ctor_set_tag(x_636, 7); -} -lean_ctor_set(x_636, 0, x_634); -lean_ctor_set(x_636, 1, x_635); -x_637 = l_Lean_indentExpr(x_560); -if (lean_is_scalar(x_568)) { - x_638 = lean_alloc_ctor(7, 2, 0); +lean_inc(x_633); +lean_inc(x_664); +x_701 = l_Lean_Meta_isExprDefEq(x_664, x_633, x_700, x_9, x_10, x_11, x_663); +if (lean_obj_tag(x_701) == 0) +{ +lean_object* x_702; uint8_t x_703; +x_702 = lean_ctor_get(x_701, 0); +lean_inc(x_702); +x_703 = lean_unbox(x_702); +lean_dec(x_702); +if (x_703 == 0) +{ +lean_object* x_704; lean_object* x_705; +x_704 = lean_ctor_get(x_701, 1); +lean_inc(x_704); +lean_dec(x_701); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +x_705 = l_Lean_Meta_addPPExplicitToExposeDiff(x_664, x_633, x_8, x_9, x_10, x_11, x_704); +if (lean_obj_tag(x_705) == 0) +{ +lean_object* x_706; lean_object* x_707; lean_object* x_708; lean_object* x_709; lean_object* x_710; lean_object* x_711; lean_object* x_712; lean_object* x_713; lean_object* x_714; lean_object* x_715; lean_object* x_716; lean_object* x_717; lean_object* x_718; lean_object* x_719; lean_object* x_720; +x_706 = lean_ctor_get(x_705, 0); +lean_inc(x_706); +x_707 = lean_ctor_get(x_705, 1); +lean_inc(x_707); +lean_dec(x_705); +x_708 = lean_ctor_get(x_706, 0); +lean_inc(x_708); +x_709 = lean_ctor_get(x_706, 1); +lean_inc(x_709); +if (lean_is_exclusive(x_706)) { + lean_ctor_release(x_706, 0); + lean_ctor_release(x_706, 1); + x_710 = x_706; +} else { + lean_dec_ref(x_706); + x_710 = lean_box(0); +} +x_711 = l_Lean_indentExpr(x_708); +x_712 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__15; +if (lean_is_scalar(x_710)) { + x_713 = lean_alloc_ctor(7, 2, 0); +} else { + x_713 = x_710; + lean_ctor_set_tag(x_713, 7); +} +lean_ctor_set(x_713, 0, x_712); +lean_ctor_set(x_713, 1, x_711); +x_714 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__17; +if (lean_is_scalar(x_665)) { + x_715 = lean_alloc_ctor(7, 2, 0); +} else { + x_715 = x_665; + lean_ctor_set_tag(x_715, 7); +} +lean_ctor_set(x_715, 0, x_713); +lean_ctor_set(x_715, 1, x_714); +x_716 = l_Lean_indentExpr(x_709); +if (lean_is_scalar(x_656)) { + x_717 = lean_alloc_ctor(7, 2, 0); } else { - x_638 = x_568; - lean_ctor_set_tag(x_638, 7); + x_717 = x_656; + lean_ctor_set_tag(x_717, 7); } -lean_ctor_set(x_638, 0, x_636); -lean_ctor_set(x_638, 1, x_637); -x_639 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__8; -if (lean_is_scalar(x_557)) { - x_640 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_717, 0, x_715); +lean_ctor_set(x_717, 1, x_716); +x_718 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__8; +if (lean_is_scalar(x_641)) { + x_719 = lean_alloc_ctor(7, 2, 0); } else { - x_640 = x_557; - lean_ctor_set_tag(x_640, 7); + x_719 = x_641; + lean_ctor_set_tag(x_719, 7); } -lean_ctor_set(x_640, 0, x_638); -lean_ctor_set(x_640, 1, x_639); -x_641 = l_Lean_throwErrorAt___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_throwStuckAtUniverseCnstr___spec__12(x_4, x_640, x_6, x_7, x_8, x_9, x_10, x_11, x_631); +lean_ctor_set(x_719, 0, x_717); +lean_ctor_set(x_719, 1, x_718); +x_720 = l_Lean_throwErrorAt___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_throwStuckAtUniverseCnstr___spec__12(x_4, x_719, x_6, x_7, x_8, x_9, x_10, x_11, x_707); lean_dec(x_11); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -return x_641; +return x_720; } else { -lean_object* x_642; lean_object* x_643; lean_object* x_644; lean_object* x_645; -lean_dec(x_592); -lean_dec(x_591); -lean_dec(x_583); -lean_dec(x_568); -lean_dec(x_560); -lean_dec(x_557); +lean_object* x_721; lean_object* x_722; lean_object* x_723; lean_object* x_724; +lean_dec(x_665); +lean_dec(x_656); +lean_dec(x_641); 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_642 = lean_ctor_get(x_628, 1); -lean_inc(x_642); -if (lean_is_exclusive(x_628)) { - lean_ctor_release(x_628, 0); - lean_ctor_release(x_628, 1); - x_643 = x_628; +x_721 = lean_ctor_get(x_705, 0); +lean_inc(x_721); +x_722 = lean_ctor_get(x_705, 1); +lean_inc(x_722); +if (lean_is_exclusive(x_705)) { + lean_ctor_release(x_705, 0); + lean_ctor_release(x_705, 1); + x_723 = x_705; } else { - lean_dec_ref(x_628); - x_643 = lean_box(0); + lean_dec_ref(x_705); + x_723 = lean_box(0); } -x_644 = lean_box(0); -if (lean_is_scalar(x_643)) { - x_645 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_723)) { + x_724 = lean_alloc_ctor(1, 2, 0); } else { - x_645 = x_643; + x_724 = x_723; } -lean_ctor_set(x_645, 0, x_644); -lean_ctor_set(x_645, 1, x_642); -return x_645; +lean_ctor_set(x_724, 0, x_721); +lean_ctor_set(x_724, 1, x_722); +return x_724; } } else { -lean_object* x_646; lean_object* x_647; lean_object* x_648; lean_object* x_649; -lean_dec(x_592); -lean_dec(x_591); -lean_dec(x_583); -lean_dec(x_568); -lean_dec(x_560); -lean_dec(x_557); +lean_object* x_725; lean_object* x_726; lean_object* x_727; lean_object* x_728; +lean_dec(x_665); +lean_dec(x_664); +lean_dec(x_656); +lean_dec(x_641); +lean_dec(x_633); 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_646 = lean_ctor_get(x_628, 0); -lean_inc(x_646); -x_647 = lean_ctor_get(x_628, 1); -lean_inc(x_647); -if (lean_is_exclusive(x_628)) { - lean_ctor_release(x_628, 0); - lean_ctor_release(x_628, 1); - x_648 = x_628; +x_725 = lean_ctor_get(x_701, 1); +lean_inc(x_725); +if (lean_is_exclusive(x_701)) { + lean_ctor_release(x_701, 0); + lean_ctor_release(x_701, 1); + x_726 = x_701; } else { - lean_dec_ref(x_628); - x_648 = lean_box(0); + lean_dec_ref(x_701); + x_726 = lean_box(0); } -if (lean_is_scalar(x_648)) { - x_649 = lean_alloc_ctor(1, 2, 0); +x_727 = lean_box(0); +if (lean_is_scalar(x_726)) { + x_728 = lean_alloc_ctor(0, 2, 0); } else { - x_649 = x_648; + x_728 = x_726; } -lean_ctor_set(x_649, 0, x_646); -lean_ctor_set(x_649, 1, x_647); -return x_649; +lean_ctor_set(x_728, 0, x_727); +lean_ctor_set(x_728, 1, x_725); +return x_728; } } else { -lean_object* x_650; lean_object* x_651; lean_object* x_652; lean_object* x_653; -lean_dec(x_583); -lean_dec(x_568); -lean_dec(x_560); -lean_dec(x_557); +lean_object* x_729; lean_object* x_730; lean_object* x_731; lean_object* x_732; +lean_dec(x_665); +lean_dec(x_664); +lean_dec(x_656); +lean_dec(x_641); +lean_dec(x_633); 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_650 = lean_ctor_get(x_587, 0); -lean_inc(x_650); -x_651 = lean_ctor_get(x_587, 1); -lean_inc(x_651); -if (lean_is_exclusive(x_587)) { - lean_ctor_release(x_587, 0); - lean_ctor_release(x_587, 1); - x_652 = x_587; +x_729 = lean_ctor_get(x_701, 0); +lean_inc(x_729); +x_730 = lean_ctor_get(x_701, 1); +lean_inc(x_730); +if (lean_is_exclusive(x_701)) { + lean_ctor_release(x_701, 0); + lean_ctor_release(x_701, 1); + x_731 = x_701; } else { - lean_dec_ref(x_587); - x_652 = lean_box(0); + lean_dec_ref(x_701); + x_731 = lean_box(0); } -if (lean_is_scalar(x_652)) { - x_653 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_731)) { + x_732 = lean_alloc_ctor(1, 2, 0); } else { - x_653 = x_652; + x_732 = x_731; } -lean_ctor_set(x_653, 0, x_650); -lean_ctor_set(x_653, 1, x_651); -return x_653; +lean_ctor_set(x_732, 0, x_729); +lean_ctor_set(x_732, 1, x_730); +return x_732; } } else { -lean_object* x_654; lean_object* x_655; lean_object* x_656; lean_object* x_657; -lean_dec(x_583); -lean_dec(x_568); -lean_dec(x_560); -lean_dec(x_557); +lean_object* x_733; lean_object* x_734; lean_object* x_735; lean_object* x_736; +lean_dec(x_656); +lean_dec(x_641); +lean_dec(x_633); 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_654 = lean_ctor_get(x_584, 0); -lean_inc(x_654); -x_655 = lean_ctor_get(x_584, 1); -lean_inc(x_655); -if (lean_is_exclusive(x_584)) { - lean_ctor_release(x_584, 0); - lean_ctor_release(x_584, 1); - x_656 = x_584; +x_733 = lean_ctor_get(x_660, 0); +lean_inc(x_733); +x_734 = lean_ctor_get(x_660, 1); +lean_inc(x_734); +if (lean_is_exclusive(x_660)) { + lean_ctor_release(x_660, 0); + lean_ctor_release(x_660, 1); + x_735 = x_660; } else { - lean_dec_ref(x_584); - x_656 = lean_box(0); + lean_dec_ref(x_660); + x_735 = lean_box(0); } -if (lean_is_scalar(x_656)) { - x_657 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_735)) { + x_736 = lean_alloc_ctor(1, 2, 0); } else { - x_657 = x_656; + x_736 = x_735; } -lean_ctor_set(x_657, 0, x_654); -lean_ctor_set(x_657, 1, x_655); -return x_657; +lean_ctor_set(x_736, 0, x_733); +lean_ctor_set(x_736, 1, x_734); +return x_736; } } else { -lean_object* x_658; lean_object* x_659; lean_object* x_660; lean_object* x_661; lean_object* x_662; lean_object* x_663; lean_object* x_664; lean_object* x_665; lean_object* x_666; lean_object* x_667; lean_object* x_668; lean_object* x_669; lean_object* x_670; -if (lean_is_exclusive(x_567)) { - lean_ctor_release(x_567, 0); - lean_ctor_release(x_567, 1); - x_658 = x_567; +lean_object* x_737; lean_object* x_738; lean_object* x_739; lean_object* x_740; +lean_dec(x_656); +lean_dec(x_641); +lean_dec(x_633); +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_737 = lean_ctor_get(x_657, 0); +lean_inc(x_737); +x_738 = lean_ctor_get(x_657, 1); +lean_inc(x_738); +if (lean_is_exclusive(x_657)) { + lean_ctor_release(x_657, 0); + lean_ctor_release(x_657, 1); + x_739 = x_657; } else { - lean_dec_ref(x_567); - x_658 = lean_box(0); + lean_dec_ref(x_657); + x_739 = lean_box(0); } -if (lean_is_exclusive(x_580)) { - lean_ctor_release(x_580, 0); - lean_ctor_release(x_580, 1); - x_659 = x_580; +if (lean_is_scalar(x_739)) { + x_740 = lean_alloc_ctor(1, 2, 0); } else { - lean_dec_ref(x_580); - x_659 = lean_box(0); + x_740 = x_739; } -x_660 = lean_ctor_get(x_565, 1); -lean_inc(x_660); -lean_dec(x_565); -x_661 = l_Lean_MessageData_ofSyntax(x_3); -x_662 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__8; -if (lean_is_scalar(x_659)) { - x_663 = lean_alloc_ctor(7, 2, 0); -} else { - x_663 = x_659; - lean_ctor_set_tag(x_663, 7); -} -lean_ctor_set(x_663, 0, x_662); -lean_ctor_set(x_663, 1, x_661); -x_664 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__19; -if (lean_is_scalar(x_658)) { - x_665 = lean_alloc_ctor(7, 2, 0); -} else { - x_665 = x_658; - lean_ctor_set_tag(x_665, 7); -} -lean_ctor_set(x_665, 0, x_663); -lean_ctor_set(x_665, 1, x_664); -x_666 = l_Lean_indentExpr(x_560); -if (lean_is_scalar(x_568)) { - x_667 = lean_alloc_ctor(7, 2, 0); -} else { - x_667 = x_568; - lean_ctor_set_tag(x_667, 7); -} -lean_ctor_set(x_667, 0, x_665); -lean_ctor_set(x_667, 1, x_666); -x_668 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__12; -if (lean_is_scalar(x_557)) { - x_669 = lean_alloc_ctor(7, 2, 0); -} else { - x_669 = x_557; - lean_ctor_set_tag(x_669, 7); -} -lean_ctor_set(x_669, 0, x_667); -lean_ctor_set(x_669, 1, x_668); -x_670 = l_Lean_throwErrorAt___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_throwStuckAtUniverseCnstr___spec__12(x_4, x_669, x_6, x_7, x_8, x_9, x_10, x_11, x_660); +lean_ctor_set(x_740, 0, x_737); +lean_ctor_set(x_740, 1, x_738); +return x_740; +} +} +else +{ +lean_object* x_741; lean_object* x_742; lean_object* x_743; lean_object* x_744; lean_object* x_745; lean_object* x_746; lean_object* x_747; lean_object* x_748; lean_object* x_749; lean_object* x_750; lean_object* x_751; lean_object* x_752; lean_object* x_753; +if (lean_is_exclusive(x_640)) { + lean_ctor_release(x_640, 0); + lean_ctor_release(x_640, 1); + x_741 = x_640; +} else { + lean_dec_ref(x_640); + x_741 = lean_box(0); +} +if (lean_is_exclusive(x_653)) { + lean_ctor_release(x_653, 0); + lean_ctor_release(x_653, 1); + x_742 = x_653; +} else { + lean_dec_ref(x_653); + x_742 = lean_box(0); +} +x_743 = lean_ctor_get(x_638, 1); +lean_inc(x_743); +lean_dec(x_638); +x_744 = l_Lean_MessageData_ofSyntax(x_3); +x_745 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__8; +if (lean_is_scalar(x_742)) { + x_746 = lean_alloc_ctor(7, 2, 0); +} else { + x_746 = x_742; + lean_ctor_set_tag(x_746, 7); +} +lean_ctor_set(x_746, 0, x_745); +lean_ctor_set(x_746, 1, x_744); +x_747 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__19; +if (lean_is_scalar(x_741)) { + x_748 = lean_alloc_ctor(7, 2, 0); +} else { + x_748 = x_741; + lean_ctor_set_tag(x_748, 7); +} +lean_ctor_set(x_748, 0, x_746); +lean_ctor_set(x_748, 1, x_747); +x_749 = l_Lean_indentExpr(x_633); +if (lean_is_scalar(x_641)) { + x_750 = lean_alloc_ctor(7, 2, 0); +} else { + x_750 = x_641; + lean_ctor_set_tag(x_750, 7); +} +lean_ctor_set(x_750, 0, x_748); +lean_ctor_set(x_750, 1, x_749); +x_751 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__12; +if (lean_is_scalar(x_630)) { + x_752 = lean_alloc_ctor(7, 2, 0); +} else { + x_752 = x_630; + lean_ctor_set_tag(x_752, 7); +} +lean_ctor_set(x_752, 0, x_750); +lean_ctor_set(x_752, 1, x_751); +x_753 = l_Lean_throwErrorAt___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_throwStuckAtUniverseCnstr___spec__12(x_4, x_752, x_6, x_7, x_8, x_9, x_10, x_11, x_743); lean_dec(x_11); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -return x_670; +return x_753; } } } else { -lean_object* x_671; lean_object* x_672; lean_object* x_673; lean_object* x_674; -lean_dec(x_560); -lean_dec(x_557); +lean_object* x_754; lean_object* x_755; lean_object* x_756; lean_object* x_757; +lean_dec(x_633); +lean_dec(x_630); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -3202,33 +3645,33 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_3); -x_671 = lean_ctor_get(x_565, 0); -lean_inc(x_671); -x_672 = lean_ctor_get(x_565, 1); -lean_inc(x_672); -if (lean_is_exclusive(x_565)) { - lean_ctor_release(x_565, 0); - lean_ctor_release(x_565, 1); - x_673 = x_565; +x_754 = lean_ctor_get(x_638, 0); +lean_inc(x_754); +x_755 = lean_ctor_get(x_638, 1); +lean_inc(x_755); +if (lean_is_exclusive(x_638)) { + lean_ctor_release(x_638, 0); + lean_ctor_release(x_638, 1); + x_756 = x_638; } else { - lean_dec_ref(x_565); - x_673 = lean_box(0); + lean_dec_ref(x_638); + x_756 = lean_box(0); } -if (lean_is_scalar(x_673)) { - x_674 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_756)) { + x_757 = lean_alloc_ctor(1, 2, 0); } else { - x_674 = x_673; + x_757 = x_756; } -lean_ctor_set(x_674, 0, x_671); -lean_ctor_set(x_674, 1, x_672); -return x_674; +lean_ctor_set(x_757, 0, x_754); +lean_ctor_set(x_757, 1, x_755); +return x_757; } } else { -lean_object* x_675; lean_object* x_676; lean_object* x_677; lean_object* x_678; -lean_dec(x_557); -lean_dec(x_555); +lean_object* x_758; lean_object* x_759; lean_object* x_760; lean_object* x_761; +lean_dec(x_630); +lean_dec(x_628); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -3236,32 +3679,32 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_3); -x_675 = lean_ctor_get(x_559, 0); -lean_inc(x_675); -x_676 = lean_ctor_get(x_559, 1); -lean_inc(x_676); -if (lean_is_exclusive(x_559)) { - lean_ctor_release(x_559, 0); - lean_ctor_release(x_559, 1); - x_677 = x_559; +x_758 = lean_ctor_get(x_632, 0); +lean_inc(x_758); +x_759 = lean_ctor_get(x_632, 1); +lean_inc(x_759); +if (lean_is_exclusive(x_632)) { + lean_ctor_release(x_632, 0); + lean_ctor_release(x_632, 1); + x_760 = x_632; } else { - lean_dec_ref(x_559); - x_677 = lean_box(0); + lean_dec_ref(x_632); + x_760 = lean_box(0); } -if (lean_is_scalar(x_677)) { - x_678 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_760)) { + x_761 = lean_alloc_ctor(1, 2, 0); } else { - x_678 = x_677; + x_761 = x_760; } -lean_ctor_set(x_678, 0, x_675); -lean_ctor_set(x_678, 1, x_676); -return x_678; +lean_ctor_set(x_761, 0, x_758); +lean_ctor_set(x_761, 1, x_759); +return x_761; } } } else { -uint8_t x_679; +uint8_t x_762; lean_dec(x_20); lean_dec(x_11); lean_dec(x_10); @@ -3271,29 +3714,29 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_3); lean_dec(x_2); -x_679 = !lean_is_exclusive(x_22); -if (x_679 == 0) +x_762 = !lean_is_exclusive(x_22); +if (x_762 == 0) { return x_22; } else { -lean_object* x_680; lean_object* x_681; lean_object* x_682; -x_680 = lean_ctor_get(x_22, 0); -x_681 = lean_ctor_get(x_22, 1); -lean_inc(x_681); -lean_inc(x_680); +lean_object* x_763; lean_object* x_764; lean_object* x_765; +x_763 = lean_ctor_get(x_22, 0); +x_764 = lean_ctor_get(x_22, 1); +lean_inc(x_764); +lean_inc(x_763); lean_dec(x_22); -x_682 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_682, 0, x_680); -lean_ctor_set(x_682, 1, x_681); -return x_682; +x_765 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_765, 0, x_763); +lean_ctor_set(x_765, 1, x_764); +return x_765; } } } else { -uint8_t x_683; +uint8_t x_766; lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -3302,23 +3745,23 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_3); lean_dec(x_2); -x_683 = !lean_is_exclusive(x_19); -if (x_683 == 0) +x_766 = !lean_is_exclusive(x_19); +if (x_766 == 0) { return x_19; } else { -lean_object* x_684; lean_object* x_685; lean_object* x_686; -x_684 = lean_ctor_get(x_19, 0); -x_685 = lean_ctor_get(x_19, 1); -lean_inc(x_685); -lean_inc(x_684); +lean_object* x_767; lean_object* x_768; lean_object* x_769; +x_767 = lean_ctor_get(x_19, 0); +x_768 = lean_ctor_get(x_19, 1); +lean_inc(x_768); +lean_inc(x_767); lean_dec(x_19); -x_686 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_686, 0, x_684); -lean_ctor_set(x_686, 1, x_685); -return x_686; +x_769 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_769, 0, x_767); +lean_ctor_set(x_769, 1, x_768); +return x_769; } } } diff --git a/stage0/stdlib/Lean/Elab/Deriving.c b/stage0/stdlib/Lean/Elab/Deriving.c index f36564b7da03..98315f27231b 100644 --- a/stage0/stdlib/Lean/Elab/Deriving.c +++ b/stage0/stdlib/Lean/Elab/Deriving.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Lean.Elab.Deriving -// Imports: Lean.Elab.Deriving.Basic Lean.Elab.Deriving.Util Lean.Elab.Deriving.Inhabited Lean.Elab.Deriving.Nonempty Lean.Elab.Deriving.TypeName Lean.Elab.Deriving.BEq Lean.Elab.Deriving.DecEq Lean.Elab.Deriving.Repr Lean.Elab.Deriving.FromToJson Lean.Elab.Deriving.SizeOf Lean.Elab.Deriving.Hashable Lean.Elab.Deriving.Ord +// Imports: Lean.Elab.Deriving.Basic Lean.Elab.Deriving.Util Lean.Elab.Deriving.Inhabited Lean.Elab.Deriving.Nonempty Lean.Elab.Deriving.TypeName Lean.Elab.Deriving.BEq Lean.Elab.Deriving.DecEq Lean.Elab.Deriving.Repr Lean.Elab.Deriving.FromToJson Lean.Elab.Deriving.SizeOf Lean.Elab.Deriving.Hashable Lean.Elab.Deriving.Ord Lean.Elab.Deriving.ToExpr #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -25,6 +25,7 @@ lean_object* initialize_Lean_Elab_Deriving_FromToJson(uint8_t builtin, lean_obje lean_object* initialize_Lean_Elab_Deriving_SizeOf(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Elab_Deriving_Hashable(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Elab_Deriving_Ord(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Elab_Deriving_ToExpr(uint8_t builtin, lean_object*); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Lean_Elab_Deriving(uint8_t builtin, lean_object* w) { lean_object * res; @@ -66,6 +67,9 @@ lean_dec_ref(res); res = initialize_Lean_Elab_Deriving_Ord(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +res = initialize_Lean_Elab_Deriving_ToExpr(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/Elab/Deriving/ToExpr.c b/stage0/stdlib/Lean/Elab/Deriving/ToExpr.c new file mode 100644 index 000000000000..6338a0682504 --- /dev/null +++ b/stage0/stdlib/Lean/Elab/Deriving/ToExpr.c @@ -0,0 +1,16344 @@ +// Lean compiler output +// Module: Lean.Elab.Deriving.ToExpr +// Imports: Lean.Meta.Transform Lean.Elab.Deriving.Basic Lean.Elab.Deriving.Util Lean.ToLevel Lean.ToExpr +#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_Elab_Deriving_ToExpr_updateIndType___closed__2; +static lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__4___closed__2; +static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__4; +static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__40; +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___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*); +static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__9; +lean_object* l_Lean_Syntax_mkNameLit(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_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* l_Lean_Elab_Command_withScope___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__22; +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_Deriving_ToExpr_mkToExprInstanceHandler___spec__2(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__23; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___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_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__13; +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__1___closed__1; +lean_object* lean_mk_empty_array_with_capacity(lean_object*); +static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__7; +LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__3; +static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__3; +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__13; +LEAN_EXPORT lean_object* l_Lean_isInductive___at_Lean_Elab_Deriving_ToExpr_mkToExprInstanceHandler___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__6; +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__1___closed__5; +static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__6; +static lean_object* l_Lean_Elab_Deriving_ToExpr_mkAuxFunctions___closed__1; +LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprInstanceCmds___spec__1(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__8; +static lean_object* l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__12; +static lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__7___lambda__1___closed__3; +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__1___closed__3; +LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___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_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__6; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___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_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__2(lean_object*, size_t, size_t, lean_object*); +static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__8; +static lean_object* l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__6; +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_Deriving_ToExpr_mkToExprInstanceHandler___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_MessageData_ofList(lean_object*); +static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__35; +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_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__5; +static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__22; +static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__5; +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__8; +static lean_object* l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__11; +static lean_object* l_Lean_Elab_Deriving_ToExpr_mkToExprBody___closed__5; +uint8_t lean_usize_dec_eq(size_t, size_t); +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Deriving_ToExpr_updateIndType___spec__1___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_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__10; +static lean_object* l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__14; +static lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__4___closed__1; +lean_object* l_Lean_Syntax_getArgs(lean_object*); +static lean_object* l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__13; +static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__4; +lean_object* l_Lean_Elab_Deriving_mkDiscrs(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Deriving_ToExpr_mkToExprInstanceCmds___closed__2; +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__1___closed__6; +static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__36; +static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__32; +static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__37; +lean_object* l_Lean_Elab_Command_elabCommand(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_array_fget(lean_object*, lean_object*); +static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__26; +static lean_object* l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__8; +static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__2; +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__1___closed__7; +static lean_object* l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__20; +lean_object* lean_environment_find(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_ToExpr_mkInstanceCmds(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__10; +static lean_object* l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__19; +LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___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_Deriving_ToExpr_mkAppNTerm(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__1; +static lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__7___lambda__1___closed__2; +static lean_object* l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__7; +uint8_t l_Lean_Expr_isAppOf(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__34; +static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__39; +lean_object* l_Lean_Syntax_node5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Syntax_isOfKind(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___boxed__const__1; +lean_object* l_Lean_stringToMessageData(lean_object*); +static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__4; +static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__12; +static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__1; +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkAuxFunctions___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__11; +static lean_object* l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__6; +LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Elab_Deriving_mkInductiveApp(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_ToExpr_updateIndType___closed__12; +LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___boxed(lean_object*, 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_ToExpr_mkToExprBody_mkAlts___spec__2___closed__2; +static lean_object* l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__2; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_updateIndType___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__7; +static lean_object* l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__16; +static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__8; +LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Deriving_ToExpr_mkAuxFunctions___closed__2; +LEAN_EXPORT lean_object* l_Lean_isInductive___at_Lean_Elab_Deriving_ToExpr_mkToExprInstanceHandler___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__12; +lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__41; +static lean_object* l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__16; +static lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__7___lambda__1___closed__5; +static lean_object* l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__26; +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__1___closed__4; +size_t lean_usize_of_nat(lean_object*); +static lean_object* l_Lean_getConstInfoCtor___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__2___closed__1; +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___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_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__9; +static lean_object* l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__15; +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Elab_registerDerivingHandler(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__14; +static lean_object* l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__7; +lean_object* l_Lean_Elab_Deriving_mkLet(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__4; +static lean_object* l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__4; +static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__11; +static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__42; +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__9; +LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_ToExpr_mkAppNTerm___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__1___closed__4; +lean_object* l_Array_unzip___rarg(lean_object*); +lean_object* l_Lean_SourceInfo_fromRef(lean_object*, uint8_t); +static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__15; +lean_object* l_Lean_Meta_isType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_MessageData_ofSyntax(lean_object*); +static lean_object* l_Lean_Elab_Deriving_ToExpr_updateIndType___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_Elab_Deriving_ToExpr_updateIndType___boxed(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__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__7___lambda__1___closed__1; +LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_ToExpr_mkToExprInstanceCmds___lambda__1(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_EXPORT lean_object* l_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts(lean_object*, lean_object*, 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_ToExpr_mkToExprBody___closed__1; +static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__1; +LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_ToExpr_mkToExprInstanceCmds___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_mkInductiveApp___spec__1(size_t, size_t, lean_object*); +lean_object* l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__1___closed__1; +static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__5; +static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__6; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__2___boxed(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Elab_getBetterRef(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__3; +LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_ToExpr_mkToExprInstanceCmds(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__9; +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___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*); +static lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__4___closed__3; +lean_object* l_outOfBounds___rarg(lean_object*); +static lean_object* l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__9; +lean_object* l_Lean_quoteNameMk(lean_object*); +lean_object* lean_st_ref_get(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_ToExpr_mkToExprInstanceHandler(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_array_pop(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1(lean_object*, lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__9; +lean_object* l_Lean_Elab_addMacroStack___at_Lean_Elab_Term_instAddErrorMessageContextTermElabM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_ToExpr_mkAuxFunction(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__29; +static lean_object* l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__6; +static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__2; +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__1___closed__3; +static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__5; +static lean_object* l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__3; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_updateIndType___spec__2(lean_object*, size_t, size_t, lean_object*); +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__1___closed__2; +lean_object* lean_array_to_list(lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_getConstInfoCtor___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__2___closed__3; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___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_Name_num___override(lean_object*, lean_object*); +lean_object* l_Lean_Syntax_node3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__10; +lean_object* l_Lean_Elab_Deriving_mkHeader(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___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_Deriving_ToExpr_mkToExprInstanceCmds___closed__5; +static lean_object* l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__13; +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__6___boxed(lean_object**); +static lean_object* l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__17; +static lean_object* l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__10; +static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__34; +lean_object* l_Lean_Elab_Deriving_mkInstImplicitBinders(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +extern lean_object* l_Lean_instInhabitedExpr; +static lean_object* l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__31; +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__1___closed__8; +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___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_Option_set___at_Lean_Elab_Term_withoutMacroStackAtErr___spec__1(lean_object*, lean_object*, uint8_t); +static lean_object* l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__4; +static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__25; +static lean_object* l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__3; +static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__12; +lean_object* l_Lean_addMacroScope(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_ToExpr_mkToExprInstanceHandler___lambda__1(lean_object*); +static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__9; +static lean_object* l_Lean_Elab_Deriving_ToExpr_mkToExprBody___closed__2; +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_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__6___closed__2; +lean_object* l_Lean_Elab_Deriving_mkInductArgNames(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_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___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*); +static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__20; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__2(size_t, size_t, lean_object*); +static lean_object* l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__2; +lean_object* l_Lean_Syntax_node2(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__30; +static lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__7___lambda__1___closed__4; +static lean_object* l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__27; +static lean_object* l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__8; +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Deriving_ToExpr_updateIndType___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__12; +uint8_t l_Lean_Syntax_matchesNull(lean_object*, lean_object*); +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__11; +static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__15; +static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__27; +static lean_object* l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__10; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__3(lean_object*, size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkAuxFunctions___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* l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__2(size_t, size_t, lean_object*); +lean_object* lean_mk_syntax_ident(lean_object*); +static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__7; +static lean_object* l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__29; +static lean_object* l_Lean_Elab_Deriving_ToExpr_mkToExprBody___closed__3; +static lean_object* l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__19; +static lean_object* l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__11; +static lean_object* l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__1; +static lean_object* l_Lean_Elab_Deriving_ToExpr_mkToExprInstanceHandler___closed__1; +static lean_object* l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__15; +static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__16; +lean_object* l_Lean_MessageData_ofConstName(lean_object*, uint8_t); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___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_append___rarg(lean_object*, lean_object*); +static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__28; +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__4___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_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__7; +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__1___closed__2; +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__7___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__30; +static lean_object* l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__13; +static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__2; +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__6; +static lean_object* l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__21; +lean_object* l_Lean_Syntax_node4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__7___boxed(lean_object**); +static lean_object* l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__17; +lean_object* l_Lean_Elab_Command_withFreshMacroScope___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__33; +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__14; +static lean_object* l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__5; +LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049_(lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___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_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__28; +static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__11; +static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__11; +static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__8; +lean_object* l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__2; +static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__19; +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__7___lambda__1___boxed(lean_object**); +lean_object* l_Lean_isTracingEnabledFor___at_Lean_Elab_Term_traceAtCmdPos___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__3; +lean_object* l_Lean_Elab_Deriving_mkImplicitBinders(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Deriving_ToExpr_mkToExprInstanceCmds___closed__3; +LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_ToExpr_mkToExprBody(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Syntax_SepArray_ofElems(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__32; +uint8_t lean_nat_dec_lt(lean_object*, lean_object*); +lean_object* lean_environment_main_module(lean_object*); +static lean_object* l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__4; +static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__13; +static lean_object* l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__4; +static lean_object* l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__11; +lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__18; +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__3(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_Deriving_ToExpr_mkAppNTerm___spec__1___closed__2; +static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__6; +static lean_object* l_Lean_Elab_Deriving_ToExpr_mkToExprBody___closed__4; +static lean_object* l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__9; +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__1; +static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___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_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__38; +static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__24; +static lean_object* l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__11; +static lean_object* l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___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*); +uint8_t l_Array_contains___at_Lean_registerInternalExceptionId___spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___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_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__21; +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__5; +static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__14; +LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_ToExpr_mkAuxFunctions(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_forallTelescopeReducing___at_Lean_Elab_Deriving_mkInductArgNames___spec__2___rarg(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__3; +LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_ToExpr_updateIndType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +extern lean_object* l_Lean_Elab_autoImplicit; +static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__3; +static lean_object* l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__16; +lean_object* l_Array_back_x21___rarg(lean_object*, lean_object*); +lean_object* l_Array_ofSubarray___rarg(lean_object*); +static lean_object* l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__18; +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__12; +static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__33; +lean_object* l_List_reverse___rarg(lean_object*); +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__1___closed__1; +static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__16; +static lean_object* l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__1; +lean_object* l_String_intercalate(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_array_mk(lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__5___boxed(lean_object**); +static lean_object* l_Lean_Elab_Deriving_ToExpr_mkToExprInstanceHandler___lambda__1___closed__1; +static lean_object* l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__5; +static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__13; +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +size_t lean_usize_add(size_t, size_t); +static lean_object* l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__24; +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__1___closed__2; +extern lean_object* l_Lean_instInhabitedName; +static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__7; +lean_object* lean_array_uget(lean_object*, size_t); +size_t lean_array_size(lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___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_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__5; +extern lean_object* l_Lean_instInhabitedInductiveVal; +LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Deriving_ToExpr_mkAuxFunctions___closed__3; +static lean_object* l_Lean_Elab_Deriving_ToExpr_mkToExprInstanceCmds___closed__4; +lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__17; +static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17; +static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__10; +static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__18; +static lean_object* l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__7; +lean_object* lean_string_append(lean_object*, lean_object*); +static lean_object* l_Lean_getConstInfoCtor___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__2___closed__4; +LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_ToExpr_mkToExprInstanceCmds___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_Deriving_ToExpr_updateIndType___closed__8; +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__3; +lean_object* lean_array_get_size(lean_object*); +LEAN_EXPORT lean_object* l_Lean_getConstInfoCtor___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___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_Deriving_ToExpr_mkToExprInstanceCmds___closed__1; +static lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__6___closed__1; +LEAN_EXPORT lean_object* l_Lean_getConstInfoCtor___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___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_ToExpr_mkToTypeExpr___lambda__1___closed__9; +static lean_object* l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__12; +lean_object* lean_infer_type(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t lean_nat_dec_le(lean_object*, lean_object*); +uint8_t lean_usize_dec_lt(size_t, size_t); +static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__12; +static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__1; +lean_object* lean_nat_add(lean_object*, lean_object*); +static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__8; +static lean_object* l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__10; +static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__10; +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__2; +static lean_object* l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__14; +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__23; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__8(size_t, size_t, lean_object*); +lean_object* l_String_toSubstring_x27(lean_object*); +static lean_object* l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__5; +lean_object* lean_array_uset(lean_object*, size_t, lean_object*); +lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__20(size_t, size_t, lean_object*); +static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__13; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__8___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__25; +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__4; +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__14; +static lean_object* l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__14; +static lean_object* l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__10; +static lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__7___lambda__1___closed__6; +lean_object* l_Lean_Elab_Command_liftTermElabM___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__1; +static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__31; +static lean_object* _init_l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean", 4, 4); +return x_1; +} +} +static lean_object* _init_l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Parser", 6, 6); +return x_1; +} +} +static lean_object* _init_l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Term", 4, 4); +return x_1; +} +} +static lean_object* _init_l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("app", 3, 3); +return x_1; +} +} +static lean_object* _init_l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___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_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__1; +x_2 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__2; +x_3 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__3; +x_4 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__4; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Expr.app", 8, 8); +return x_1; +} +} +static lean_object* _init_l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__6; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__8() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Expr", 4, 4); +return x_1; +} +} +static lean_object* _init_l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__8; +x_2 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__4; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__1; +x_2 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__8; +x_3 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__4; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__10; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__10; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__12; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__12; +x_2 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__13; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__15() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__11; +x_2 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__14; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__16() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("null", 4, 4); +return x_1; +} +} +static lean_object* _init_l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__16; +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_Deriving_ToExpr_mkAppNTerm___spec__1(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +uint8_t x_10; +x_10 = lean_usize_dec_eq(x_2, x_3); +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; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; size_t x_30; size_t x_31; +x_11 = lean_array_uget(x_1, x_2); +x_12 = lean_ctor_get(x_7, 5); +lean_inc(x_12); +x_13 = 0; +x_14 = l_Lean_SourceInfo_fromRef(x_12, x_13); +lean_dec(x_12); +x_15 = lean_ctor_get(x_7, 10); +lean_inc(x_15); +x_16 = lean_st_ref_get(x_8, x_9); +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_environment_main_module(x_19); +x_21 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__9; +x_22 = l_Lean_addMacroScope(x_20, x_21, x_15); +x_23 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__7; +x_24 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__15; +lean_inc(x_14); +x_25 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_25, 0, x_14); +lean_ctor_set(x_25, 1, x_23); +lean_ctor_set(x_25, 2, x_22); +lean_ctor_set(x_25, 3, x_24); +x_26 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17; +lean_inc(x_14); +x_27 = l_Lean_Syntax_node2(x_14, x_26, x_4, x_11); +x_28 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__5; +x_29 = l_Lean_Syntax_node2(x_14, x_28, x_25, x_27); +x_30 = 1; +x_31 = lean_usize_add(x_2, x_30); +x_2 = x_31; +x_4 = x_29; +x_9 = x_18; +goto _start; +} +else +{ +lean_object* x_33; +lean_dec(x_7); +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_4); +lean_ctor_set(x_33, 1, x_9); +return x_33; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_ToExpr_mkAppNTerm(lean_object* x_1, 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; +x_8 = lean_array_get_size(x_2); +x_9 = lean_unsigned_to_nat(0u); +x_10 = lean_nat_dec_lt(x_9, x_8); +if (x_10 == 0) +{ +lean_object* x_11; +lean_dec(x_8); +lean_dec(x_5); +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_1); +lean_ctor_set(x_11, 1, x_7); +return x_11; +} +else +{ +uint8_t x_12; +x_12 = lean_nat_dec_le(x_8, x_8); +if (x_12 == 0) +{ +lean_object* x_13; +lean_dec(x_8); +lean_dec(x_5); +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_1); +lean_ctor_set(x_13, 1, x_7); +return x_13; +} +else +{ +size_t x_14; size_t x_15; lean_object* x_16; +x_14 = 0; +x_15 = lean_usize_of_nat(x_8); +lean_dec(x_8); +x_16 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1(x_2, x_14, x_15, x_1, x_3, x_4, x_5, x_6, x_7); +return x_16; +} +} +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___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: +{ +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_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1(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_6); +lean_dec(x_5); +lean_dec(x_1); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_ToExpr_mkAppNTerm___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; +x_8 = l_Lean_Elab_Deriving_ToExpr_mkAppNTerm(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Deriving_ToExpr_updateIndType___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; uint8_t x_13; +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_is_exclusive(x_12); +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_12, 0); +x_15 = lean_ctor_get(x_12, 1); +x_16 = l_Lean_Elab_addMacroStack___at_Lean_Elab_Term_instAddErrorMessageContextTermElabM___spec__1(x_14, x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_15); +lean_dec(x_2); +x_17 = !lean_is_exclusive(x_16); +if (x_17 == 0) +{ +lean_object* x_18; +x_18 = lean_ctor_get(x_16, 0); +lean_ctor_set(x_12, 1, x_18); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set_tag(x_16, 1); +lean_ctor_set(x_16, 0, x_12); +return x_16; +} +else +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_16, 0); +x_20 = lean_ctor_get(x_16, 1); +lean_inc(x_20); +lean_inc(x_19); +lean_dec(x_16); +lean_ctor_set(x_12, 1, x_19); +lean_ctor_set(x_12, 0, x_11); +x_21 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_21, 0, x_12); +lean_ctor_set(x_21, 1, x_20); +return x_21; +} +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_22 = lean_ctor_get(x_12, 0); +x_23 = lean_ctor_get(x_12, 1); +lean_inc(x_23); +lean_inc(x_22); +lean_dec(x_12); +x_24 = l_Lean_Elab_addMacroStack___at_Lean_Elab_Term_instAddErrorMessageContextTermElabM___spec__1(x_22, x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_23); +lean_dec(x_2); +x_25 = lean_ctor_get(x_24, 0); +lean_inc(x_25); +x_26 = lean_ctor_get(x_24, 1); +lean_inc(x_26); +if (lean_is_exclusive(x_24)) { + lean_ctor_release(x_24, 0); + lean_ctor_release(x_24, 1); + x_27 = x_24; +} else { + lean_dec_ref(x_24); + x_27 = lean_box(0); +} +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_11); +lean_ctor_set(x_28, 1, x_25); +if (lean_is_scalar(x_27)) { + x_29 = lean_alloc_ctor(1, 2, 0); +} else { + x_29 = x_27; + lean_ctor_set_tag(x_29, 1); +} +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_26); +return x_29; +} +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_updateIndType___spec__2(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { +_start: +{ +uint8_t x_5; +x_5 = lean_usize_dec_lt(x_3, x_2); +if (x_5 == 0) +{ +return x_4; +} +else +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x_9; size_t x_10; lean_object* x_11; +x_6 = lean_array_uget(x_4, x_3); +x_7 = lean_unsigned_to_nat(0u); +x_8 = lean_array_uset(x_4, x_3, x_7); +x_9 = 1; +x_10 = lean_usize_add(x_3, x_9); +x_11 = lean_array_uset(x_8, x_3, x_6); +x_3 = x_10; +x_4 = x_11; +goto _start; +} +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("(internal error) expecting output of `mkInductiveApp`", 53, 53); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("explicit", 8, 8); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_updateIndType___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_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__1; +x_2 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__2; +x_3 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__3; +x_4 = l_Lean_Elab_Deriving_ToExpr_updateIndType___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_Deriving_ToExpr_updateIndType___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("@", 1, 1); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("explicitUniv", 12, 12); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_updateIndType___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_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__1; +x_2 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__2; +x_3 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__3; +x_4 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__6; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__8() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(".{", 2, 2); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = lean_mk_empty_array_with_capacity(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__10() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("level", 5, 5); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_updateIndType___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_Deriving_ToExpr_updateIndType___closed__10; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__11; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__13() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(",", 1, 1); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__14() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("}", 1, 1); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_ToExpr_updateIndType(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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; size_t x_13; size_t x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_10 = lean_ctor_get(x_1, 0); +lean_inc(x_10); +lean_dec(x_1); +x_11 = lean_ctor_get(x_10, 1); +lean_inc(x_11); +lean_dec(x_10); +x_12 = lean_array_mk(x_11); +x_13 = lean_array_size(x_12); +x_14 = 0; +x_15 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_mkInductiveApp___spec__1(x_13, x_14, x_12); +x_16 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__5; +lean_inc(x_2); +x_17 = l_Lean_Syntax_isOfKind(x_2, x_16); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; +lean_dec(x_15); +lean_dec(x_2); +x_18 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__2; +x_19 = l_Lean_throwError___at_Lean_Elab_Deriving_ToExpr_updateIndType___spec__1(x_18, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_19; +} +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; +x_20 = lean_unsigned_to_nat(0u); +x_21 = l_Lean_Syntax_getArg(x_2, x_20); +x_22 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__4; +lean_inc(x_21); +x_23 = l_Lean_Syntax_isOfKind(x_21, x_22); +if (x_23 == 0) +{ +lean_object* x_24; lean_object* x_25; +lean_dec(x_21); +lean_dec(x_15); +lean_dec(x_2); +x_24 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__2; +x_25 = l_Lean_throwError___at_Lean_Elab_Deriving_ToExpr_updateIndType___spec__1(x_24, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_25; +} +else +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; lean_object* x_32; lean_object* x_33; uint8_t x_34; +lean_dec(x_3); +x_26 = lean_unsigned_to_nat(1u); +x_27 = l_Lean_Syntax_getArg(x_21, x_26); +lean_dec(x_21); +x_28 = l_Lean_Syntax_getArg(x_2, x_26); +lean_dec(x_2); +x_29 = l_Lean_Syntax_getArgs(x_28); +lean_dec(x_28); +x_30 = lean_ctor_get(x_7, 5); +x_31 = 0; +x_32 = l_Lean_SourceInfo_fromRef(x_30, x_31); +x_33 = lean_st_ref_get(x_8, x_9); +x_34 = !lean_is_exclusive(x_33); +if (x_34 == 0) +{ +lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; size_t x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_35 = lean_ctor_get(x_33, 0); +lean_dec(x_35); +x_36 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__5; +lean_inc(x_32); +x_37 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_37, 0, x_32); +lean_ctor_set(x_37, 1, x_36); +x_38 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__8; +lean_inc(x_32); +x_39 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_39, 0, x_32); +lean_ctor_set(x_39, 1, x_38); +x_40 = lean_array_size(x_15); +x_41 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__12; +x_42 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_updateIndType___spec__2(x_41, x_40, x_14, x_15); +x_43 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__13; +x_44 = l_Lean_Syntax_SepArray_ofElems(x_43, x_42); +lean_dec(x_42); +x_45 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__9; +x_46 = l_Array_append___rarg(x_45, x_44); +lean_dec(x_44); +x_47 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17; +lean_inc(x_32); +x_48 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_48, 0, x_32); +lean_ctor_set(x_48, 1, x_47); +lean_ctor_set(x_48, 2, x_46); +x_49 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__14; +lean_inc(x_32); +x_50 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_50, 0, x_32); +lean_ctor_set(x_50, 1, x_49); +x_51 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__7; +lean_inc(x_32); +x_52 = l_Lean_Syntax_node4(x_32, x_51, x_27, x_39, x_48, x_50); +lean_inc(x_32); +x_53 = l_Lean_Syntax_node2(x_32, x_22, x_37, x_52); +x_54 = l_Array_append___rarg(x_45, x_29); +lean_dec(x_29); +lean_inc(x_32); +x_55 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_55, 0, x_32); +lean_ctor_set(x_55, 1, x_47); +lean_ctor_set(x_55, 2, x_54); +x_56 = l_Lean_Syntax_node2(x_32, x_16, x_53, x_55); +lean_ctor_set(x_33, 0, x_56); +return x_33; +} +else +{ +lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; size_t x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; +x_57 = lean_ctor_get(x_33, 1); +lean_inc(x_57); +lean_dec(x_33); +x_58 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__5; +lean_inc(x_32); +x_59 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_59, 0, x_32); +lean_ctor_set(x_59, 1, x_58); +x_60 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__8; +lean_inc(x_32); +x_61 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_61, 0, x_32); +lean_ctor_set(x_61, 1, x_60); +x_62 = lean_array_size(x_15); +x_63 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__12; +x_64 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_updateIndType___spec__2(x_63, x_62, x_14, x_15); +x_65 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__13; +x_66 = l_Lean_Syntax_SepArray_ofElems(x_65, x_64); +lean_dec(x_64); +x_67 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__9; +x_68 = l_Array_append___rarg(x_67, x_66); +lean_dec(x_66); +x_69 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17; +lean_inc(x_32); +x_70 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_70, 0, x_32); +lean_ctor_set(x_70, 1, x_69); +lean_ctor_set(x_70, 2, x_68); +x_71 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__14; +lean_inc(x_32); +x_72 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_72, 0, x_32); +lean_ctor_set(x_72, 1, x_71); +x_73 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__7; +lean_inc(x_32); +x_74 = l_Lean_Syntax_node4(x_32, x_73, x_27, x_61, x_70, x_72); +lean_inc(x_32); +x_75 = l_Lean_Syntax_node2(x_32, x_22, x_59, x_74); +x_76 = l_Array_append___rarg(x_67, x_29); +lean_dec(x_29); +lean_inc(x_32); +x_77 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_77, 0, x_32); +lean_ctor_set(x_77, 1, x_69); +lean_ctor_set(x_77, 2, x_76); +x_78 = l_Lean_Syntax_node2(x_32, x_16, x_75, x_77); +x_79 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_79, 0, x_78); +lean_ctor_set(x_79, 1, x_57); +return x_79; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Deriving_ToExpr_updateIndType___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; +x_9 = l_Lean_throwError___at_Lean_Elab_Deriving_ToExpr_updateIndType___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_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_updateIndType___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +size_t x_5; size_t x_6; lean_object* x_7; +x_5 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_6 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_7 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_updateIndType___spec__2(x_1, x_5, x_6, x_4); +lean_dec(x_1); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_ToExpr_updateIndType___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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_Deriving_ToExpr_updateIndType(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; +} +} +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.toLevel", 12, 12); +return x_1; +} +} +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__1___closed__1; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__1___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("toLevel", 7, 7); +return x_1; +} +} +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__1___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__1; +x_2 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__1___closed__3; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__1___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("ToLevel", 7, 7); +return x_1; +} +} +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__1___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__1; +x_2 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__1___closed__5; +x_3 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__1___closed__3; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__1___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__1___closed__6; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__1___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__1___closed__7; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__1(size_t x_1, size_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_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_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_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; uint8_t x_21; +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_8, 5); +lean_inc(x_16); +x_17 = 0; +x_18 = l_Lean_SourceInfo_fromRef(x_16, x_17); +lean_dec(x_16); +x_19 = lean_ctor_get(x_8, 10); +lean_inc(x_19); +x_20 = lean_st_ref_get(x_9, x_10); +x_21 = !lean_is_exclusive(x_20); +if (x_21 == 0) +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; size_t x_39; size_t x_40; lean_object* x_41; +x_22 = lean_ctor_get(x_20, 0); +x_23 = lean_ctor_get(x_20, 1); +x_24 = lean_ctor_get(x_22, 0); +lean_inc(x_24); +lean_dec(x_22); +x_25 = lean_environment_main_module(x_24); +x_26 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__1___closed__4; +x_27 = l_Lean_addMacroScope(x_25, x_26, x_19); +x_28 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__1___closed__2; +x_29 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__1___closed__8; +lean_inc(x_18); +x_30 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_30, 0, x_18); +lean_ctor_set(x_30, 1, x_28); +lean_ctor_set(x_30, 2, x_27); +lean_ctor_set(x_30, 3, x_29); +x_31 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__8; +lean_inc(x_18); +lean_ctor_set_tag(x_20, 2); +lean_ctor_set(x_20, 1, x_31); +lean_ctor_set(x_20, 0, x_18); +x_32 = lean_mk_syntax_ident(x_13); +x_33 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17; +lean_inc(x_18); +x_34 = l_Lean_Syntax_node1(x_18, x_33, x_32); +x_35 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__14; +lean_inc(x_18); +x_36 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_36, 0, x_18); +lean_ctor_set(x_36, 1, x_35); +x_37 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__7; +x_38 = l_Lean_Syntax_node4(x_18, x_37, x_30, x_20, x_34, x_36); +x_39 = 1; +x_40 = lean_usize_add(x_2, x_39); +x_41 = lean_array_uset(x_15, x_2, x_38); +x_2 = x_40; +x_3 = x_41; +x_10 = x_23; +goto _start; +} +else +{ +lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; size_t x_61; size_t x_62; lean_object* x_63; +x_43 = lean_ctor_get(x_20, 0); +x_44 = lean_ctor_get(x_20, 1); +lean_inc(x_44); +lean_inc(x_43); +lean_dec(x_20); +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_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__1___closed__4; +x_48 = l_Lean_addMacroScope(x_46, x_47, x_19); +x_49 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__1___closed__2; +x_50 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__1___closed__8; +lean_inc(x_18); +x_51 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_51, 0, x_18); +lean_ctor_set(x_51, 1, x_49); +lean_ctor_set(x_51, 2, x_48); +lean_ctor_set(x_51, 3, x_50); +x_52 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__8; +lean_inc(x_18); +x_53 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_53, 0, x_18); +lean_ctor_set(x_53, 1, x_52); +x_54 = lean_mk_syntax_ident(x_13); +x_55 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17; +lean_inc(x_18); +x_56 = l_Lean_Syntax_node1(x_18, x_55, x_54); +x_57 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__14; +lean_inc(x_18); +x_58 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_58, 0, x_18); +lean_ctor_set(x_58, 1, x_57); +x_59 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__7; +x_60 = l_Lean_Syntax_node4(x_18, x_59, x_51, x_53, x_56, x_58); +x_61 = 1; +x_62 = lean_usize_add(x_2, x_61); +x_63 = lean_array_uset(x_15, x_2, x_60); +x_2 = x_62; +x_3 = x_63; +x_10 = x_44; +goto _start; +} +} +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("toExpr", 6, 6); +return x_1; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__1; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("ToExpr", 6, 6); +return x_1; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__1; +x_2 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__4; +x_3 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__1; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__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; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__6; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__8() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("toTypeExpr", 10, 10); +return x_1; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__8; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__8; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__1; +x_2 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__4; +x_3 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__8; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___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_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__12; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +uint8_t x_14; +x_14 = lean_usize_dec_lt(x_5, x_4); +if (x_14 == 0) +{ +lean_object* x_15; +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_6); +lean_ctor_set(x_15, 1, x_13); +return x_15; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_26; +x_16 = lean_array_uget(x_3, x_5); +x_26 = !lean_is_exclusive(x_6); +if (x_26 == 0) +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; uint8_t x_32; +x_27 = lean_ctor_get(x_6, 0); +x_28 = lean_ctor_get(x_6, 1); +x_29 = lean_ctor_get(x_27, 0); +lean_inc(x_29); +x_30 = lean_ctor_get(x_27, 1); +lean_inc(x_30); +x_31 = lean_ctor_get(x_27, 2); +lean_inc(x_31); +x_32 = lean_nat_dec_lt(x_30, x_31); +if (x_32 == 0) +{ +lean_object* x_33; +lean_dec(x_31); +lean_dec(x_30); +lean_dec(x_29); +lean_dec(x_16); +x_33 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_33, 0, x_6); +x_17 = x_33; +x_18 = x_13; +goto block_25; +} +else +{ +uint8_t x_34; +x_34 = !lean_is_exclusive(x_27); +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_27, 2); +lean_dec(x_35); +x_36 = lean_ctor_get(x_27, 1); +lean_dec(x_36); +x_37 = lean_ctor_get(x_27, 0); +lean_dec(x_37); +x_38 = lean_array_fget(x_29, x_30); +x_39 = lean_unsigned_to_nat(1u); +x_40 = lean_nat_add(x_30, x_39); +lean_dec(x_30); +lean_ctor_set(x_27, 1, x_40); +x_41 = lean_mk_syntax_ident(x_16); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +x_42 = l_Lean_Meta_isType(x_38, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_42) == 0) +{ +lean_object* x_43; uint8_t x_44; +x_43 = lean_ctor_get(x_42, 0); +lean_inc(x_43); +x_44 = lean_unbox(x_43); +lean_dec(x_43); +if (x_44 == 0) +{ +lean_object* x_45; lean_object* x_46; uint8_t x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; +x_45 = lean_ctor_get(x_42, 1); +lean_inc(x_45); +lean_dec(x_42); +x_46 = lean_ctor_get(x_11, 5); +lean_inc(x_46); +x_47 = 0; +x_48 = l_Lean_SourceInfo_fromRef(x_46, x_47); +lean_dec(x_46); +x_49 = lean_ctor_get(x_11, 10); +lean_inc(x_49); +x_50 = lean_st_ref_get(x_12, x_45); +x_51 = lean_ctor_get(x_50, 0); +lean_inc(x_51); +x_52 = lean_ctor_get(x_50, 1); +lean_inc(x_52); +lean_dec(x_50); +x_53 = lean_ctor_get(x_51, 0); +lean_inc(x_53); +lean_dec(x_51); +x_54 = lean_environment_main_module(x_53); +x_55 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__3; +x_56 = l_Lean_addMacroScope(x_54, x_55, x_49); +x_57 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__2; +x_58 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__7; +lean_inc(x_48); +x_59 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_59, 0, x_48); +lean_ctor_set(x_59, 1, x_57); +lean_ctor_set(x_59, 2, x_56); +lean_ctor_set(x_59, 3, x_58); +x_60 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17; +lean_inc(x_48); +x_61 = l_Lean_Syntax_node1(x_48, x_60, x_41); +x_62 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__5; +x_63 = l_Lean_Syntax_node2(x_48, x_62, x_59, x_61); +x_64 = lean_array_push(x_28, x_63); +lean_ctor_set(x_6, 1, x_64); +x_65 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_65, 0, x_6); +x_17 = x_65; +x_18 = x_52; +goto block_25; +} +else +{ +lean_object* x_66; lean_object* x_67; uint8_t x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; 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_66 = lean_ctor_get(x_42, 1); +lean_inc(x_66); +lean_dec(x_42); +x_67 = lean_ctor_get(x_11, 5); +lean_inc(x_67); +x_68 = 0; +x_69 = l_Lean_SourceInfo_fromRef(x_67, x_68); +lean_dec(x_67); +x_70 = lean_ctor_get(x_11, 10); +lean_inc(x_70); +x_71 = lean_st_ref_get(x_12, x_66); +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); +lean_dec(x_72); +x_75 = lean_environment_main_module(x_74); +x_76 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__10; +x_77 = l_Lean_addMacroScope(x_75, x_76, x_70); +x_78 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__9; +x_79 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__13; +lean_inc(x_69); +x_80 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_80, 0, x_69); +lean_ctor_set(x_80, 1, x_78); +lean_ctor_set(x_80, 2, x_77); +lean_ctor_set(x_80, 3, x_79); +x_81 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17; +lean_inc(x_69); +x_82 = l_Lean_Syntax_node1(x_69, x_81, x_41); +x_83 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__5; +x_84 = l_Lean_Syntax_node2(x_69, x_83, x_80, x_82); +x_85 = lean_array_push(x_28, x_84); +lean_ctor_set(x_6, 1, x_85); +x_86 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_86, 0, x_6); +x_17 = x_86; +x_18 = x_73; +goto block_25; +} +} +else +{ +uint8_t x_87; +lean_dec(x_41); +lean_dec(x_27); +lean_free_object(x_6); +lean_dec(x_28); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +x_87 = !lean_is_exclusive(x_42); +if (x_87 == 0) +{ +return x_42; +} +else +{ +lean_object* x_88; lean_object* x_89; lean_object* x_90; +x_88 = lean_ctor_get(x_42, 0); +x_89 = lean_ctor_get(x_42, 1); +lean_inc(x_89); +lean_inc(x_88); +lean_dec(x_42); +x_90 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_90, 0, x_88); +lean_ctor_set(x_90, 1, x_89); +return x_90; +} +} +} +else +{ +lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; +lean_dec(x_27); +x_91 = lean_array_fget(x_29, x_30); +x_92 = lean_unsigned_to_nat(1u); +x_93 = lean_nat_add(x_30, x_92); +lean_dec(x_30); +x_94 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_94, 0, x_29); +lean_ctor_set(x_94, 1, x_93); +lean_ctor_set(x_94, 2, x_31); +x_95 = lean_mk_syntax_ident(x_16); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +x_96 = l_Lean_Meta_isType(x_91, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_96) == 0) +{ +lean_object* x_97; uint8_t x_98; +x_97 = lean_ctor_get(x_96, 0); +lean_inc(x_97); +x_98 = lean_unbox(x_97); +lean_dec(x_97); +if (x_98 == 0) +{ +lean_object* x_99; lean_object* x_100; uint8_t x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; +x_99 = lean_ctor_get(x_96, 1); +lean_inc(x_99); +lean_dec(x_96); +x_100 = lean_ctor_get(x_11, 5); +lean_inc(x_100); +x_101 = 0; +x_102 = l_Lean_SourceInfo_fromRef(x_100, x_101); +lean_dec(x_100); +x_103 = lean_ctor_get(x_11, 10); +lean_inc(x_103); +x_104 = lean_st_ref_get(x_12, x_99); +x_105 = lean_ctor_get(x_104, 0); +lean_inc(x_105); +x_106 = lean_ctor_get(x_104, 1); +lean_inc(x_106); +lean_dec(x_104); +x_107 = lean_ctor_get(x_105, 0); +lean_inc(x_107); +lean_dec(x_105); +x_108 = lean_environment_main_module(x_107); +x_109 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__3; +x_110 = l_Lean_addMacroScope(x_108, x_109, x_103); +x_111 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__2; +x_112 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__7; +lean_inc(x_102); +x_113 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_113, 0, x_102); +lean_ctor_set(x_113, 1, x_111); +lean_ctor_set(x_113, 2, x_110); +lean_ctor_set(x_113, 3, x_112); +x_114 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17; +lean_inc(x_102); +x_115 = l_Lean_Syntax_node1(x_102, x_114, x_95); +x_116 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__5; +x_117 = l_Lean_Syntax_node2(x_102, x_116, x_113, x_115); +x_118 = lean_array_push(x_28, x_117); +lean_ctor_set(x_6, 1, x_118); +lean_ctor_set(x_6, 0, x_94); +x_119 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_119, 0, x_6); +x_17 = x_119; +x_18 = x_106; +goto block_25; +} +else +{ +lean_object* x_120; lean_object* x_121; uint8_t x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; +x_120 = lean_ctor_get(x_96, 1); +lean_inc(x_120); +lean_dec(x_96); +x_121 = lean_ctor_get(x_11, 5); +lean_inc(x_121); +x_122 = 0; +x_123 = l_Lean_SourceInfo_fromRef(x_121, x_122); +lean_dec(x_121); +x_124 = lean_ctor_get(x_11, 10); +lean_inc(x_124); +x_125 = lean_st_ref_get(x_12, x_120); +x_126 = lean_ctor_get(x_125, 0); +lean_inc(x_126); +x_127 = lean_ctor_get(x_125, 1); +lean_inc(x_127); +lean_dec(x_125); +x_128 = lean_ctor_get(x_126, 0); +lean_inc(x_128); +lean_dec(x_126); +x_129 = lean_environment_main_module(x_128); +x_130 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__10; +x_131 = l_Lean_addMacroScope(x_129, x_130, x_124); +x_132 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__9; +x_133 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__13; +lean_inc(x_123); +x_134 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_134, 0, x_123); +lean_ctor_set(x_134, 1, x_132); +lean_ctor_set(x_134, 2, x_131); +lean_ctor_set(x_134, 3, x_133); +x_135 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17; +lean_inc(x_123); +x_136 = l_Lean_Syntax_node1(x_123, x_135, x_95); +x_137 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__5; +x_138 = l_Lean_Syntax_node2(x_123, x_137, x_134, x_136); +x_139 = lean_array_push(x_28, x_138); +lean_ctor_set(x_6, 1, x_139); +lean_ctor_set(x_6, 0, x_94); +x_140 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_140, 0, x_6); +x_17 = x_140; +x_18 = x_127; +goto block_25; +} +} +else +{ +lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; +lean_dec(x_95); +lean_dec(x_94); +lean_free_object(x_6); +lean_dec(x_28); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +x_141 = lean_ctor_get(x_96, 0); +lean_inc(x_141); +x_142 = lean_ctor_get(x_96, 1); +lean_inc(x_142); +if (lean_is_exclusive(x_96)) { + lean_ctor_release(x_96, 0); + lean_ctor_release(x_96, 1); + x_143 = x_96; +} else { + lean_dec_ref(x_96); + x_143 = lean_box(0); +} +if (lean_is_scalar(x_143)) { + x_144 = lean_alloc_ctor(1, 2, 0); +} else { + x_144 = x_143; +} +lean_ctor_set(x_144, 0, x_141); +lean_ctor_set(x_144, 1, x_142); +return x_144; +} +} +} +} +else +{ +lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; uint8_t x_150; +x_145 = lean_ctor_get(x_6, 0); +x_146 = lean_ctor_get(x_6, 1); +lean_inc(x_146); +lean_inc(x_145); +lean_dec(x_6); +x_147 = lean_ctor_get(x_145, 0); +lean_inc(x_147); +x_148 = lean_ctor_get(x_145, 1); +lean_inc(x_148); +x_149 = lean_ctor_get(x_145, 2); +lean_inc(x_149); +x_150 = lean_nat_dec_lt(x_148, x_149); +if (x_150 == 0) +{ +lean_object* x_151; lean_object* x_152; +lean_dec(x_149); +lean_dec(x_148); +lean_dec(x_147); +lean_dec(x_16); +x_151 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_151, 0, x_145); +lean_ctor_set(x_151, 1, x_146); +x_152 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_152, 0, x_151); +x_17 = x_152; +x_18 = x_13; +goto block_25; +} +else +{ +lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; +if (lean_is_exclusive(x_145)) { + lean_ctor_release(x_145, 0); + lean_ctor_release(x_145, 1); + lean_ctor_release(x_145, 2); + x_153 = x_145; +} else { + lean_dec_ref(x_145); + x_153 = lean_box(0); +} +x_154 = lean_array_fget(x_147, x_148); +x_155 = lean_unsigned_to_nat(1u); +x_156 = lean_nat_add(x_148, x_155); +lean_dec(x_148); +if (lean_is_scalar(x_153)) { + x_157 = lean_alloc_ctor(0, 3, 0); +} else { + x_157 = x_153; +} +lean_ctor_set(x_157, 0, x_147); +lean_ctor_set(x_157, 1, x_156); +lean_ctor_set(x_157, 2, x_149); +x_158 = lean_mk_syntax_ident(x_16); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +x_159 = l_Lean_Meta_isType(x_154, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_159) == 0) +{ +lean_object* x_160; uint8_t x_161; +x_160 = lean_ctor_get(x_159, 0); +lean_inc(x_160); +x_161 = lean_unbox(x_160); +lean_dec(x_160); +if (x_161 == 0) +{ +lean_object* x_162; 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; lean_object* x_183; +x_162 = lean_ctor_get(x_159, 1); +lean_inc(x_162); +lean_dec(x_159); +x_163 = lean_ctor_get(x_11, 5); +lean_inc(x_163); +x_164 = 0; +x_165 = l_Lean_SourceInfo_fromRef(x_163, x_164); +lean_dec(x_163); +x_166 = lean_ctor_get(x_11, 10); +lean_inc(x_166); +x_167 = lean_st_ref_get(x_12, x_162); +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_168, 0); +lean_inc(x_170); +lean_dec(x_168); +x_171 = lean_environment_main_module(x_170); +x_172 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__3; +x_173 = l_Lean_addMacroScope(x_171, x_172, x_166); +x_174 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__2; +x_175 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__7; +lean_inc(x_165); +x_176 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_176, 0, x_165); +lean_ctor_set(x_176, 1, x_174); +lean_ctor_set(x_176, 2, x_173); +lean_ctor_set(x_176, 3, x_175); +x_177 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17; +lean_inc(x_165); +x_178 = l_Lean_Syntax_node1(x_165, x_177, x_158); +x_179 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__5; +x_180 = l_Lean_Syntax_node2(x_165, x_179, x_176, x_178); +x_181 = lean_array_push(x_146, x_180); +x_182 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_182, 0, x_157); +lean_ctor_set(x_182, 1, x_181); +x_183 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_183, 0, x_182); +x_17 = x_183; +x_18 = x_169; +goto block_25; +} +else +{ +lean_object* x_184; lean_object* x_185; uint8_t x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; +x_184 = lean_ctor_get(x_159, 1); +lean_inc(x_184); +lean_dec(x_159); +x_185 = lean_ctor_get(x_11, 5); +lean_inc(x_185); +x_186 = 0; +x_187 = l_Lean_SourceInfo_fromRef(x_185, x_186); +lean_dec(x_185); +x_188 = lean_ctor_get(x_11, 10); +lean_inc(x_188); +x_189 = lean_st_ref_get(x_12, x_184); +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_190, 0); +lean_inc(x_192); +lean_dec(x_190); +x_193 = lean_environment_main_module(x_192); +x_194 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__10; +x_195 = l_Lean_addMacroScope(x_193, x_194, x_188); +x_196 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__9; +x_197 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__13; +lean_inc(x_187); +x_198 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_198, 0, x_187); +lean_ctor_set(x_198, 1, x_196); +lean_ctor_set(x_198, 2, x_195); +lean_ctor_set(x_198, 3, x_197); +x_199 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17; +lean_inc(x_187); +x_200 = l_Lean_Syntax_node1(x_187, x_199, x_158); +x_201 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__5; +x_202 = l_Lean_Syntax_node2(x_187, x_201, x_198, x_200); +x_203 = lean_array_push(x_146, x_202); +x_204 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_204, 0, x_157); +lean_ctor_set(x_204, 1, x_203); +x_205 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_205, 0, x_204); +x_17 = x_205; +x_18 = x_191; +goto block_25; +} +} +else +{ +lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; +lean_dec(x_158); +lean_dec(x_157); +lean_dec(x_146); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +x_206 = lean_ctor_get(x_159, 0); +lean_inc(x_206); +x_207 = lean_ctor_get(x_159, 1); +lean_inc(x_207); +if (lean_is_exclusive(x_159)) { + lean_ctor_release(x_159, 0); + lean_ctor_release(x_159, 1); + x_208 = x_159; +} else { + lean_dec_ref(x_159); + x_208 = lean_box(0); +} +if (lean_is_scalar(x_208)) { + x_209 = lean_alloc_ctor(1, 2, 0); +} else { + x_209 = x_208; +} +lean_ctor_set(x_209, 0, x_206); +lean_ctor_set(x_209, 1, x_207); +return x_209; +} +} +} +block_25: +{ +if (lean_obj_tag(x_17) == 0) +{ +lean_object* x_19; lean_object* x_20; +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +x_19 = lean_ctor_get(x_17, 0); +lean_inc(x_19); +lean_dec(x_17); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_19); +lean_ctor_set(x_20, 1, x_18); +return x_20; +} +else +{ +lean_object* x_21; size_t x_22; size_t x_23; +x_21 = lean_ctor_get(x_17, 0); +lean_inc(x_21); +lean_dec(x_17); +x_22 = 1; +x_23 = lean_usize_add(x_5, x_22); +x_5 = x_23; +x_6 = x_21; +x_13 = x_18; +goto _start; +} +} +} +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(0); +x_2 = lean_array_mk(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Expr.const", 10, 10); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__2; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("const", 5, 5); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__8; +x_2 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__4; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__1; +x_2 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__8; +x_3 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__4; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__6; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__6; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__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_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__8; +x_2 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___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_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__7; +x_2 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__10; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__12() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("term[_]", 7, 7); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__12; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__14() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("[", 1, 1); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__15() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("]", 1, 1); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__16() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("quotedName", 10, 10); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__17() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__1; +x_2 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__2; +x_3 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__3; +x_4 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__16; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__18() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(".", 1, 1); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__19() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("`", 1, 1); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1(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: +{ +lean_object* x_14; 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; lean_object* x_22; +x_14 = lean_box(0); +x_15 = lean_array_get_size(x_5); +x_16 = lean_unsigned_to_nat(0u); +x_17 = l_Array_toSubarray___rarg(x_5, x_16, x_15); +x_18 = lean_box(0); +x_19 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__1; +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_size(x_1); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +x_22 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2(x_1, x_18, x_1, x_21, x_2, x_20, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_22) == 0) +{ +lean_object* x_23; lean_object* x_24; uint8_t x_25; +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +x_24 = lean_ctor_get(x_22, 1); +lean_inc(x_24); +lean_dec(x_22); +x_25 = !lean_is_exclusive(x_23); +if (x_25 == 0) +{ +lean_object* x_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; uint8_t x_33; +x_26 = lean_ctor_get(x_23, 1); +x_27 = lean_ctor_get(x_23, 0); +lean_dec(x_27); +x_28 = lean_ctor_get(x_11, 5); +lean_inc(x_28); +x_29 = 0; +x_30 = l_Lean_SourceInfo_fromRef(x_28, x_29); +lean_dec(x_28); +x_31 = lean_ctor_get(x_11, 10); +lean_inc(x_31); +x_32 = lean_st_ref_get(x_12, x_24); +x_33 = !lean_is_exclusive(x_32); +if (x_33 == 0) +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; +x_34 = lean_ctor_get(x_32, 0); +x_35 = lean_ctor_get(x_32, 1); +x_36 = lean_ctor_get(x_34, 0); +lean_inc(x_36); +lean_dec(x_34); +x_37 = lean_environment_main_module(x_36); +x_38 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__5; +x_39 = l_Lean_addMacroScope(x_37, x_38, x_31); +x_40 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__3; +x_41 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__11; +lean_inc(x_30); +x_42 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_42, 0, x_30); +lean_ctor_set(x_42, 1, x_40); +lean_ctor_set(x_42, 2, x_39); +lean_ctor_set(x_42, 3, x_41); +x_43 = lean_ctor_get(x_3, 0); +lean_inc(x_43); +lean_dec(x_3); +lean_inc(x_43); +x_44 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_14, x_43); +x_45 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__14; +lean_inc(x_30); +lean_ctor_set_tag(x_32, 2); +lean_ctor_set(x_32, 1, x_45); +lean_ctor_set(x_32, 0, x_30); +x_46 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__13; +x_47 = l_Lean_Syntax_SepArray_ofElems(x_46, x_4); +x_48 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__9; +x_49 = l_Array_append___rarg(x_48, x_47); +lean_dec(x_47); +x_50 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17; +lean_inc(x_30); +x_51 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_51, 0, x_30); +lean_ctor_set(x_51, 1, x_50); +lean_ctor_set(x_51, 2, x_49); +x_52 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__15; +lean_inc(x_30); +lean_ctor_set_tag(x_23, 2); +lean_ctor_set(x_23, 1, x_52); +lean_ctor_set(x_23, 0, x_30); +x_53 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__13; +lean_inc(x_30); +x_54 = l_Lean_Syntax_node3(x_30, x_53, x_32, x_51, x_23); +if (lean_obj_tag(x_44) == 0) +{ +lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; +x_55 = l_Lean_quoteNameMk(x_43); +lean_inc(x_30); +x_56 = l_Lean_Syntax_node2(x_30, x_50, x_55, x_54); +x_57 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__5; +x_58 = l_Lean_Syntax_node2(x_30, x_57, x_42, x_56); +x_59 = l_Lean_Elab_Deriving_ToExpr_mkAppNTerm(x_58, x_26, x_9, x_10, x_11, x_12, x_35); +lean_dec(x_12); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_26); +return x_59; +} +else +{ +lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; +lean_dec(x_43); +x_60 = lean_ctor_get(x_44, 0); +lean_inc(x_60); +lean_dec(x_44); +x_61 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__18; +x_62 = l_String_intercalate(x_61, x_60); +x_63 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__19; +x_64 = lean_string_append(x_63, x_62); +lean_dec(x_62); +x_65 = lean_box(2); +x_66 = l_Lean_Syntax_mkNameLit(x_64, 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_14); +x_68 = lean_array_mk(x_67); +x_69 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__17; +x_70 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_70, 0, x_65); +lean_ctor_set(x_70, 1, x_69); +lean_ctor_set(x_70, 2, x_68); +lean_inc(x_30); +x_71 = l_Lean_Syntax_node2(x_30, x_50, x_70, x_54); +x_72 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__5; +x_73 = l_Lean_Syntax_node2(x_30, x_72, x_42, x_71); +x_74 = l_Lean_Elab_Deriving_ToExpr_mkAppNTerm(x_73, x_26, x_9, x_10, x_11, x_12, x_35); +lean_dec(x_12); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_26); +return x_74; +} +} +else +{ +lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; +x_75 = lean_ctor_get(x_32, 0); +x_76 = lean_ctor_get(x_32, 1); +lean_inc(x_76); +lean_inc(x_75); +lean_dec(x_32); +x_77 = lean_ctor_get(x_75, 0); +lean_inc(x_77); +lean_dec(x_75); +x_78 = lean_environment_main_module(x_77); +x_79 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__5; +x_80 = l_Lean_addMacroScope(x_78, x_79, x_31); +x_81 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__3; +x_82 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__11; +lean_inc(x_30); +x_83 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_83, 0, x_30); +lean_ctor_set(x_83, 1, x_81); +lean_ctor_set(x_83, 2, x_80); +lean_ctor_set(x_83, 3, x_82); +x_84 = lean_ctor_get(x_3, 0); +lean_inc(x_84); +lean_dec(x_3); +lean_inc(x_84); +x_85 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_14, x_84); +x_86 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__14; +lean_inc(x_30); +x_87 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_87, 0, x_30); +lean_ctor_set(x_87, 1, x_86); +x_88 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__13; +x_89 = l_Lean_Syntax_SepArray_ofElems(x_88, x_4); +x_90 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__9; +x_91 = l_Array_append___rarg(x_90, x_89); +lean_dec(x_89); +x_92 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17; +lean_inc(x_30); +x_93 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_93, 0, x_30); +lean_ctor_set(x_93, 1, x_92); +lean_ctor_set(x_93, 2, x_91); +x_94 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__15; +lean_inc(x_30); +lean_ctor_set_tag(x_23, 2); +lean_ctor_set(x_23, 1, x_94); +lean_ctor_set(x_23, 0, x_30); +x_95 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__13; +lean_inc(x_30); +x_96 = l_Lean_Syntax_node3(x_30, x_95, x_87, x_93, x_23); +if (lean_obj_tag(x_85) == 0) +{ +lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; +x_97 = l_Lean_quoteNameMk(x_84); +lean_inc(x_30); +x_98 = l_Lean_Syntax_node2(x_30, x_92, x_97, x_96); +x_99 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__5; +x_100 = l_Lean_Syntax_node2(x_30, x_99, x_83, x_98); +x_101 = l_Lean_Elab_Deriving_ToExpr_mkAppNTerm(x_100, x_26, x_9, x_10, x_11, x_12, x_76); +lean_dec(x_12); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_26); +return x_101; +} +else +{ +lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; +lean_dec(x_84); +x_102 = lean_ctor_get(x_85, 0); +lean_inc(x_102); +lean_dec(x_85); +x_103 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__18; +x_104 = l_String_intercalate(x_103, x_102); +x_105 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__19; +x_106 = lean_string_append(x_105, x_104); +lean_dec(x_104); +x_107 = lean_box(2); +x_108 = l_Lean_Syntax_mkNameLit(x_106, x_107); +x_109 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_109, 0, x_108); +lean_ctor_set(x_109, 1, x_14); +x_110 = lean_array_mk(x_109); +x_111 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__17; +x_112 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_112, 0, x_107); +lean_ctor_set(x_112, 1, x_111); +lean_ctor_set(x_112, 2, x_110); +lean_inc(x_30); +x_113 = l_Lean_Syntax_node2(x_30, x_92, x_112, x_96); +x_114 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__5; +x_115 = l_Lean_Syntax_node2(x_30, x_114, x_83, x_113); +x_116 = l_Lean_Elab_Deriving_ToExpr_mkAppNTerm(x_115, x_26, x_9, x_10, x_11, x_12, x_76); +lean_dec(x_12); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_26); +return x_116; +} +} +} +else +{ +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; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; +x_117 = lean_ctor_get(x_23, 1); +lean_inc(x_117); +lean_dec(x_23); +x_118 = lean_ctor_get(x_11, 5); +lean_inc(x_118); +x_119 = 0; +x_120 = l_Lean_SourceInfo_fromRef(x_118, x_119); +lean_dec(x_118); +x_121 = lean_ctor_get(x_11, 10); +lean_inc(x_121); +x_122 = lean_st_ref_get(x_12, x_24); +x_123 = lean_ctor_get(x_122, 0); +lean_inc(x_123); +x_124 = lean_ctor_get(x_122, 1); +lean_inc(x_124); +if (lean_is_exclusive(x_122)) { + lean_ctor_release(x_122, 0); + lean_ctor_release(x_122, 1); + x_125 = x_122; +} else { + lean_dec_ref(x_122); + x_125 = lean_box(0); +} +x_126 = lean_ctor_get(x_123, 0); +lean_inc(x_126); +lean_dec(x_123); +x_127 = lean_environment_main_module(x_126); +x_128 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__5; +x_129 = l_Lean_addMacroScope(x_127, x_128, x_121); +x_130 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__3; +x_131 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__11; +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_ctor_get(x_3, 0); +lean_inc(x_133); +lean_dec(x_3); +lean_inc(x_133); +x_134 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_14, x_133); +x_135 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__14; +lean_inc(x_120); +if (lean_is_scalar(x_125)) { + x_136 = lean_alloc_ctor(2, 2, 0); +} else { + x_136 = x_125; + lean_ctor_set_tag(x_136, 2); +} +lean_ctor_set(x_136, 0, x_120); +lean_ctor_set(x_136, 1, x_135); +x_137 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__13; +x_138 = l_Lean_Syntax_SepArray_ofElems(x_137, x_4); +x_139 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__9; +x_140 = l_Array_append___rarg(x_139, x_138); +lean_dec(x_138); +x_141 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17; +lean_inc(x_120); +x_142 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_142, 0, x_120); +lean_ctor_set(x_142, 1, x_141); +lean_ctor_set(x_142, 2, x_140); +x_143 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__15; +lean_inc(x_120); +x_144 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_144, 0, x_120); +lean_ctor_set(x_144, 1, x_143); +x_145 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__13; +lean_inc(x_120); +x_146 = l_Lean_Syntax_node3(x_120, x_145, x_136, x_142, x_144); +if (lean_obj_tag(x_134) == 0) +{ +lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; +x_147 = l_Lean_quoteNameMk(x_133); +lean_inc(x_120); +x_148 = l_Lean_Syntax_node2(x_120, x_141, x_147, x_146); +x_149 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__5; +x_150 = l_Lean_Syntax_node2(x_120, x_149, x_132, x_148); +x_151 = l_Lean_Elab_Deriving_ToExpr_mkAppNTerm(x_150, x_117, x_9, x_10, x_11, x_12, x_124); +lean_dec(x_12); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_117); +return x_151; +} +else +{ +lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; +lean_dec(x_133); +x_152 = lean_ctor_get(x_134, 0); +lean_inc(x_152); +lean_dec(x_134); +x_153 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__18; +x_154 = l_String_intercalate(x_153, x_152); +x_155 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__19; +x_156 = lean_string_append(x_155, x_154); +lean_dec(x_154); +x_157 = lean_box(2); +x_158 = l_Lean_Syntax_mkNameLit(x_156, x_157); +x_159 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_159, 0, x_158); +lean_ctor_set(x_159, 1, x_14); +x_160 = lean_array_mk(x_159); +x_161 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__17; +x_162 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_162, 0, x_157); +lean_ctor_set(x_162, 1, x_161); +lean_ctor_set(x_162, 2, x_160); +lean_inc(x_120); +x_163 = l_Lean_Syntax_node2(x_120, x_141, x_162, x_146); +x_164 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__5; +x_165 = l_Lean_Syntax_node2(x_120, x_164, x_132, x_163); +x_166 = l_Lean_Elab_Deriving_ToExpr_mkAppNTerm(x_165, x_117, x_9, x_10, x_11, x_12, x_124); +lean_dec(x_12); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_117); +return x_166; +} +} +} +else +{ +uint8_t x_167; +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_3); +x_167 = !lean_is_exclusive(x_22); +if (x_167 == 0) +{ +return x_22; +} +else +{ +lean_object* x_168; lean_object* x_169; lean_object* x_170; +x_168 = lean_ctor_get(x_22, 0); +x_169 = lean_ctor_get(x_22, 1); +lean_inc(x_169); +lean_inc(x_168); +lean_dec(x_22); +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; +} +} +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___boxed__const__1() { +_start: +{ +size_t x_1; lean_object* x_2; +x_1 = 0; +x_2 = lean_box_usize(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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; 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; uint8_t x_21; lean_object* x_22; +x_10 = lean_ctor_get(x_1, 0); +lean_inc(x_10); +lean_dec(x_1); +x_11 = lean_ctor_get(x_10, 1); +lean_inc(x_11); +x_12 = lean_array_mk(x_11); +x_13 = lean_array_size(x_12); +x_14 = 0; +lean_inc(x_7); +x_15 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__1(x_13, x_14, x_12, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_15, 1); +lean_inc(x_17); +lean_dec(x_15); +x_18 = lean_ctor_get(x_10, 2); +lean_inc(x_18); +x_19 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___boxed__const__1; +x_20 = lean_alloc_closure((void*)(l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___boxed), 13, 4); +lean_closure_set(x_20, 0, x_2); +lean_closure_set(x_20, 1, x_19); +lean_closure_set(x_20, 2, x_10); +lean_closure_set(x_20, 3, x_16); +x_21 = 0; +x_22 = l_Lean_Meta_forallTelescopeReducing___at_Lean_Elab_Deriving_mkInductArgNames___spec__2___rarg(x_18, x_20, x_21, x_3, x_4, x_5, x_6, x_7, x_8, x_17); +return x_22; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___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: +{ +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_Elab_Deriving_ToExpr_mkToTypeExpr___spec__1(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); +lean_dec(x_5); +lean_dec(x_4); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +size_t x_14; size_t x_15; lean_object* x_16; +x_14 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_15 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_16 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2(x_1, x_2, x_3, x_14, x_15, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_16; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___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: +{ +size_t x_14; lean_object* x_15; +x_14 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_15 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1(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_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_1); +return x_15; +} +} +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("proj", 4, 4); +return x_1; +} +} +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__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_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__1; +x_2 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__2; +x_3 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__3; +x_4 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__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_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__1___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__1___closed__3; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__1___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__1___closed__3; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__1(size_t x_1, size_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_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_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_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; uint8_t x_21; +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_8, 5); +lean_inc(x_16); +x_17 = 0; +x_18 = l_Lean_SourceInfo_fromRef(x_16, x_17); +lean_dec(x_16); +x_19 = lean_ctor_get(x_8, 10); +lean_inc(x_19); +x_20 = lean_st_ref_get(x_9, x_10); +x_21 = !lean_is_exclusive(x_20); +if (x_21 == 0) +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; size_t x_34; size_t x_35; lean_object* x_36; +x_22 = lean_ctor_get(x_20, 0); +x_23 = lean_ctor_get(x_20, 1); +x_24 = lean_ctor_get(x_22, 0); +lean_inc(x_24); +lean_dec(x_22); +x_25 = lean_environment_main_module(x_24); +x_26 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__18; +lean_inc(x_18); +lean_ctor_set_tag(x_20, 2); +lean_ctor_set(x_20, 1, x_26); +lean_ctor_set(x_20, 0, x_18); +x_27 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__1___closed__4; +x_28 = l_Lean_addMacroScope(x_25, x_27, x_19); +x_29 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__1___closed__3; +x_30 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__1___closed__8; +lean_inc(x_18); +x_31 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_31, 0, x_18); +lean_ctor_set(x_31, 1, x_29); +lean_ctor_set(x_31, 2, x_28); +lean_ctor_set(x_31, 3, x_30); +x_32 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__1___closed__2; +x_33 = l_Lean_Syntax_node3(x_18, x_32, x_13, x_20, x_31); +x_34 = 1; +x_35 = lean_usize_add(x_2, x_34); +x_36 = lean_array_uset(x_15, x_2, x_33); +x_2 = x_35; +x_3 = x_36; +x_10 = x_23; +goto _start; +} +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; size_t x_51; size_t x_52; lean_object* x_53; +x_38 = lean_ctor_get(x_20, 0); +x_39 = lean_ctor_get(x_20, 1); +lean_inc(x_39); +lean_inc(x_38); +lean_dec(x_20); +x_40 = lean_ctor_get(x_38, 0); +lean_inc(x_40); +lean_dec(x_38); +x_41 = lean_environment_main_module(x_40); +x_42 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__18; +lean_inc(x_18); +x_43 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_43, 0, x_18); +lean_ctor_set(x_43, 1, x_42); +x_44 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__1___closed__4; +x_45 = l_Lean_addMacroScope(x_41, x_44, x_19); +x_46 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__1___closed__3; +x_47 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__1___closed__8; +lean_inc(x_18); +x_48 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_48, 0, x_18); +lean_ctor_set(x_48, 1, x_46); +lean_ctor_set(x_48, 2, x_45); +lean_ctor_set(x_48, 3, x_47); +x_49 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__1___closed__2; +x_50 = l_Lean_Syntax_node3(x_18, x_49, x_13, x_43, x_48); +x_51 = 1; +x_52 = lean_usize_add(x_2, x_51); +x_53 = lean_array_uset(x_15, x_2, x_50); +x_2 = x_52; +x_3 = x_53; +x_10 = x_39; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_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_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_is_exclusive(x_12); +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_12, 0); +x_15 = lean_ctor_get(x_12, 1); +x_16 = l_Lean_Elab_addMacroStack___at_Lean_Elab_Term_instAddErrorMessageContextTermElabM___spec__1(x_14, x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_15); +lean_dec(x_2); +x_17 = !lean_is_exclusive(x_16); +if (x_17 == 0) +{ +lean_object* x_18; +x_18 = lean_ctor_get(x_16, 0); +lean_ctor_set(x_12, 1, x_18); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set_tag(x_16, 1); +lean_ctor_set(x_16, 0, x_12); +return x_16; +} +else +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_16, 0); +x_20 = lean_ctor_get(x_16, 1); +lean_inc(x_20); +lean_inc(x_19); +lean_dec(x_16); +lean_ctor_set(x_12, 1, x_19); +lean_ctor_set(x_12, 0, x_11); +x_21 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_21, 0, x_12); +lean_ctor_set(x_21, 1, x_20); +return x_21; +} +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_22 = lean_ctor_get(x_12, 0); +x_23 = lean_ctor_get(x_12, 1); +lean_inc(x_23); +lean_inc(x_22); +lean_dec(x_12); +x_24 = l_Lean_Elab_addMacroStack___at_Lean_Elab_Term_instAddErrorMessageContextTermElabM___spec__1(x_22, x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_23); +lean_dec(x_2); +x_25 = lean_ctor_get(x_24, 0); +lean_inc(x_25); +x_26 = lean_ctor_get(x_24, 1); +lean_inc(x_26); +if (lean_is_exclusive(x_24)) { + lean_ctor_release(x_24, 0); + lean_ctor_release(x_24, 1); + x_27 = x_24; +} else { + lean_dec_ref(x_24); + x_27 = lean_box(0); +} +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_11); +lean_ctor_set(x_28, 1, x_25); +if (lean_is_scalar(x_27)) { + x_29 = lean_alloc_ctor(1, 2, 0); +} else { + x_29 = x_27; + lean_ctor_set_tag(x_29, 1); +} +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_26); +return x_29; +} +} +} +static lean_object* _init_l_Lean_getConstInfoCtor___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__2___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("'", 1, 1); +return x_1; +} +} +static lean_object* _init_l_Lean_getConstInfoCtor___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__2___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_getConstInfoCtor___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__2___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_getConstInfoCtor___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__2___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("' is not a constructor", 22, 22); +return x_1; +} +} +static lean_object* _init_l_Lean_getConstInfoCtor___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__2___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_getConstInfoCtor___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__2___closed__3; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_getConstInfoCtor___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +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); +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) == 6) +{ +uint8_t x_11; +lean_dec(x_2); +lean_dec(x_1); +x_11 = !lean_is_exclusive(x_9); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; +x_12 = lean_ctor_get(x_9, 0); +lean_dec(x_12); +x_13 = lean_ctor_get(x_10, 0); +lean_inc(x_13); +lean_dec(x_10); +lean_ctor_set(x_9, 0, x_13); +return x_9; +} +else +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_14 = lean_ctor_get(x_9, 1); +lean_inc(x_14); +lean_dec(x_9); +x_15 = lean_ctor_get(x_10, 0); +lean_inc(x_15); +lean_dec(x_10); +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; 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_dec(x_10); +x_17 = lean_ctor_get(x_9, 1); +lean_inc(x_17); +lean_dec(x_9); +x_18 = 0; +x_19 = l_Lean_MessageData_ofConstName(x_1, x_18); +x_20 = l_Lean_getConstInfoCtor___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__2___closed__2; +x_21 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_21, 0, x_20); +lean_ctor_set(x_21, 1, x_19); +x_22 = l_Lean_getConstInfoCtor___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__2___closed__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); +x_24 = l_Lean_throwError___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__3(x_23, x_2, x_3, x_4, x_5, x_6, x_7, x_17); +return x_24; +} +} +else +{ +uint8_t x_25; +lean_dec(x_2); +lean_dec(x_1); +x_25 = !lean_is_exclusive(x_9); +if (x_25 == 0) +{ +return x_9; +} +else +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_26 = lean_ctor_get(x_9, 0); +x_27 = lean_ctor_get(x_9, 1); +lean_inc(x_27); +lean_inc(x_26); +lean_dec(x_9); +x_28 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_28, 0, x_26); +lean_ctor_set(x_28, 1, x_27); +return x_28; +} +} +} +} +static lean_object* _init_l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__4___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("hole", 4, 4); +return x_1; +} +} +static lean_object* _init_l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__4___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__1; +x_2 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__2; +x_3 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__3; +x_4 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__4___closed__1; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__4___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("_", 1, 1); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +lean_object* x_14; uint8_t x_15; +x_14 = lean_ctor_get(x_2, 1); +x_15 = lean_nat_dec_lt(x_4, x_14); +if (x_15 == 0) +{ +lean_object* x_16; +lean_dec(x_4); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_3); +lean_ctor_set(x_16, 1, x_13); +return x_16; +} +else +{ +lean_object* x_17; uint8_t x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_17 = lean_ctor_get(x_11, 5); +x_18 = 0; +x_19 = l_Lean_SourceInfo_fromRef(x_17, x_18); +x_20 = lean_st_ref_get(x_12, x_13); +x_21 = !lean_is_exclusive(x_20); +if (x_21 == 0) +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_22 = lean_ctor_get(x_20, 1); +x_23 = lean_ctor_get(x_20, 0); +lean_dec(x_23); +x_24 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__4___closed__3; +lean_inc(x_19); +lean_ctor_set_tag(x_20, 2); +lean_ctor_set(x_20, 1, x_24); +lean_ctor_set(x_20, 0, x_19); +x_25 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__4___closed__2; +x_26 = l_Lean_Syntax_node1(x_19, x_25, x_20); +x_27 = lean_array_push(x_3, x_26); +x_28 = lean_ctor_get(x_2, 2); +x_29 = lean_nat_add(x_4, x_28); +lean_dec(x_4); +x_3 = x_27; +x_4 = x_29; +x_5 = lean_box(0); +x_6 = lean_box(0); +x_13 = x_22; +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; +x_31 = lean_ctor_get(x_20, 1); +lean_inc(x_31); +lean_dec(x_20); +x_32 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__4___closed__3; +lean_inc(x_19); +x_33 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_33, 0, x_19); +lean_ctor_set(x_33, 1, x_32); +x_34 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__4___closed__2; +x_35 = l_Lean_Syntax_node1(x_19, x_34, x_33); +x_36 = lean_array_push(x_3, x_35); +x_37 = lean_ctor_get(x_2, 2); +x_38 = lean_nat_add(x_4, x_37); +lean_dec(x_4); +x_3 = x_36; +x_4 = x_38; +x_5 = lean_box(0); +x_6 = lean_box(0); +x_13 = x_31; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18) { +_start: +{ +lean_object* x_19; uint8_t x_20; +x_19 = lean_ctor_get(x_7, 1); +x_20 = lean_nat_dec_lt(x_9, x_19); +if (x_20 == 0) +{ +lean_object* x_21; +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_9); +lean_dec(x_3); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_8); +lean_ctor_set(x_21, 1, x_18); +return x_21; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; lean_object* x_34; uint8_t x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_29 = lean_ctor_get(x_8, 0); +lean_inc(x_29); +x_30 = lean_ctor_get(x_8, 1); +lean_inc(x_30); +lean_dec(x_8); +x_31 = lean_ctor_get(x_1, 1); +x_32 = lean_array_get_size(x_31); +x_33 = lean_nat_dec_lt(x_9, x_32); +lean_dec(x_32); +x_34 = lean_ctor_get(x_16, 5); +lean_inc(x_34); +x_35 = 0; +x_36 = l_Lean_SourceInfo_fromRef(x_34, x_35); +lean_dec(x_34); +x_37 = lean_st_ref_get(x_17, x_18); +if (x_33 == 0) +{ +lean_object* x_253; lean_object* x_254; +x_253 = l_Lean_instInhabitedName; +x_254 = l_outOfBounds___rarg(x_253); +x_38 = x_254; +goto block_252; +} +else +{ +lean_object* x_255; +x_255 = lean_array_fget(x_31, x_9); +x_38 = x_255; +goto block_252; +} +block_28: +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_22, 0); +lean_inc(x_24); +lean_dec(x_22); +x_25 = lean_ctor_get(x_7, 2); +x_26 = lean_nat_add(x_9, x_25); +lean_dec(x_9); +x_8 = x_24; +x_9 = x_26; +x_10 = lean_box(0); +x_11 = lean_box(0); +x_18 = x_23; +goto _start; +} +block_252: +{ +uint8_t x_39; +x_39 = !lean_is_exclusive(x_37); +if (x_39 == 0) +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; uint8_t x_48; lean_object* x_49; +x_40 = lean_ctor_get(x_37, 1); +x_41 = lean_ctor_get(x_37, 0); +lean_dec(x_41); +x_42 = lean_mk_syntax_ident(x_38); +x_43 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__4___closed__3; +lean_inc(x_36); +lean_ctor_set_tag(x_37, 2); +lean_ctor_set(x_37, 1, x_43); +lean_ctor_set(x_37, 0, x_36); +x_44 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__4___closed__2; +lean_inc(x_36); +x_45 = l_Lean_Syntax_node1(x_36, x_44, x_37); +x_46 = lean_array_push(x_29, x_45); +x_47 = lean_array_get_size(x_5); +x_48 = lean_nat_dec_lt(x_9, x_47); +lean_dec(x_47); +if (x_48 == 0) +{ +lean_object* x_164; lean_object* x_165; +x_164 = l_Lean_instInhabitedExpr; +x_165 = l_outOfBounds___rarg(x_164); +x_49 = x_165; +goto block_163; +} +else +{ +lean_object* x_166; +x_166 = lean_array_fget(x_5, x_9); +x_49 = x_166; +goto block_163; +} +block_163: +{ +lean_object* x_50; +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_49); +x_50 = lean_infer_type(x_49, x_14, x_15, x_16, x_17, x_40); +if (lean_obj_tag(x_50) == 0) +{ +lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; uint8_t x_55; +x_51 = lean_ctor_get(x_50, 0); +lean_inc(x_51); +x_52 = lean_ctor_get(x_50, 1); +lean_inc(x_52); +lean_dec(x_50); +x_53 = lean_ctor_get(x_2, 0); +x_54 = lean_ctor_get(x_53, 0); +x_55 = l_Lean_Expr_isAppOf(x_51, x_54); +lean_dec(x_51); +if (x_55 == 0) +{ +lean_object* x_56; +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_56 = l_Lean_Meta_isType(x_49, x_14, x_15, x_16, x_17, x_52); +if (lean_obj_tag(x_56) == 0) +{ +lean_object* x_57; uint8_t x_58; +x_57 = lean_ctor_get(x_56, 0); +lean_inc(x_57); +x_58 = lean_unbox(x_57); +lean_dec(x_57); +if (x_58 == 0) +{ +lean_object* x_59; lean_object* x_60; lean_object* x_61; uint8_t x_62; +x_59 = lean_ctor_get(x_56, 1); +lean_inc(x_59); +lean_dec(x_56); +x_60 = lean_ctor_get(x_16, 10); +lean_inc(x_60); +x_61 = lean_st_ref_get(x_17, x_59); +x_62 = !lean_is_exclusive(x_61); +if (x_62 == 0) +{ +lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; 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_63 = lean_ctor_get(x_61, 0); +x_64 = lean_ctor_get(x_61, 1); +x_65 = lean_ctor_get(x_63, 0); +lean_inc(x_65); +lean_dec(x_63); +x_66 = lean_environment_main_module(x_65); +x_67 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__3; +x_68 = l_Lean_addMacroScope(x_66, x_67, x_60); +x_69 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__2; +x_70 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__7; +lean_inc(x_36); +x_71 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_71, 0, x_36); +lean_ctor_set(x_71, 1, x_69); +lean_ctor_set(x_71, 2, x_68); +lean_ctor_set(x_71, 3, x_70); +x_72 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17; +lean_inc(x_36); +x_73 = l_Lean_Syntax_node1(x_36, x_72, x_42); +x_74 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__5; +x_75 = l_Lean_Syntax_node2(x_36, x_74, x_71, x_73); +x_76 = lean_array_push(x_30, x_75); +lean_ctor_set(x_61, 1, x_76); +lean_ctor_set(x_61, 0, x_46); +x_77 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_77, 0, x_61); +x_22 = x_77; +x_23 = x_64; +goto block_28; +} +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; lean_object* x_93; +x_78 = lean_ctor_get(x_61, 0); +x_79 = lean_ctor_get(x_61, 1); +lean_inc(x_79); +lean_inc(x_78); +lean_dec(x_61); +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_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__3; +x_83 = l_Lean_addMacroScope(x_81, x_82, x_60); +x_84 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__2; +x_85 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__7; +lean_inc(x_36); +x_86 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_86, 0, x_36); +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_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17; +lean_inc(x_36); +x_88 = l_Lean_Syntax_node1(x_36, x_87, x_42); +x_89 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__5; +x_90 = l_Lean_Syntax_node2(x_36, x_89, x_86, x_88); +x_91 = lean_array_push(x_30, x_90); +x_92 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_92, 0, x_46); +lean_ctor_set(x_92, 1, x_91); +x_93 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_93, 0, x_92); +x_22 = x_93; +x_23 = x_79; +goto block_28; +} +} +else +{ +lean_object* x_94; lean_object* x_95; lean_object* x_96; uint8_t x_97; +x_94 = lean_ctor_get(x_56, 1); +lean_inc(x_94); +lean_dec(x_56); +x_95 = lean_ctor_get(x_16, 10); +lean_inc(x_95); +x_96 = lean_st_ref_get(x_17, x_94); +x_97 = !lean_is_exclusive(x_96); +if (x_97 == 0) +{ +lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; 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_98 = lean_ctor_get(x_96, 0); +x_99 = lean_ctor_get(x_96, 1); +x_100 = lean_ctor_get(x_98, 0); +lean_inc(x_100); +lean_dec(x_98); +x_101 = lean_environment_main_module(x_100); +x_102 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__10; +x_103 = l_Lean_addMacroScope(x_101, x_102, x_95); +x_104 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__9; +x_105 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__13; +lean_inc(x_36); +x_106 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_106, 0, x_36); +lean_ctor_set(x_106, 1, x_104); +lean_ctor_set(x_106, 2, x_103); +lean_ctor_set(x_106, 3, x_105); +x_107 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17; +lean_inc(x_36); +x_108 = l_Lean_Syntax_node1(x_36, x_107, x_42); +x_109 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__5; +x_110 = l_Lean_Syntax_node2(x_36, x_109, x_106, x_108); +x_111 = lean_array_push(x_30, x_110); +lean_ctor_set(x_96, 1, x_111); +lean_ctor_set(x_96, 0, x_46); +x_112 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_112, 0, x_96); +x_22 = x_112; +x_23 = x_99; +goto block_28; +} +else +{ +lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; +x_113 = lean_ctor_get(x_96, 0); +x_114 = lean_ctor_get(x_96, 1); +lean_inc(x_114); +lean_inc(x_113); +lean_dec(x_96); +x_115 = lean_ctor_get(x_113, 0); +lean_inc(x_115); +lean_dec(x_113); +x_116 = lean_environment_main_module(x_115); +x_117 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__10; +x_118 = l_Lean_addMacroScope(x_116, x_117, x_95); +x_119 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__9; +x_120 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__13; +lean_inc(x_36); +x_121 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_121, 0, x_36); +lean_ctor_set(x_121, 1, x_119); +lean_ctor_set(x_121, 2, x_118); +lean_ctor_set(x_121, 3, x_120); +x_122 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17; +lean_inc(x_36); +x_123 = l_Lean_Syntax_node1(x_36, x_122, x_42); +x_124 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__5; +x_125 = l_Lean_Syntax_node2(x_36, x_124, x_121, x_123); +x_126 = lean_array_push(x_30, x_125); +x_127 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_127, 0, x_46); +lean_ctor_set(x_127, 1, x_126); +x_128 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_128, 0, x_127); +x_22 = x_128; +x_23 = x_114; +goto block_28; +} +} +} +else +{ +uint8_t x_129; +lean_dec(x_46); +lean_dec(x_42); +lean_dec(x_36); +lean_dec(x_30); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_9); +lean_dec(x_3); +x_129 = !lean_is_exclusive(x_56); +if (x_129 == 0) +{ +return x_56; +} +else +{ +lean_object* x_130; lean_object* x_131; lean_object* x_132; +x_130 = lean_ctor_get(x_56, 0); +x_131 = lean_ctor_get(x_56, 1); +lean_inc(x_131); +lean_inc(x_130); +lean_dec(x_56); +x_132 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_132, 0, x_130); +lean_ctor_set(x_132, 1, x_131); +return x_132; +} +} +} +else +{ +lean_object* x_133; uint8_t x_134; +lean_dec(x_49); +x_133 = lean_st_ref_get(x_17, x_52); +x_134 = !lean_is_exclusive(x_133); +if (x_134 == 0) +{ +lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; +x_135 = lean_ctor_get(x_133, 1); +x_136 = lean_ctor_get(x_133, 0); +lean_dec(x_136); +lean_inc(x_3); +x_137 = lean_mk_syntax_ident(x_3); +x_138 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__9; +x_139 = l_Array_append___rarg(x_138, x_4); +x_140 = lean_array_push(x_139, x_42); +x_141 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17; +lean_inc(x_36); +x_142 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_142, 0, x_36); +lean_ctor_set(x_142, 1, x_141); +lean_ctor_set(x_142, 2, x_140); +x_143 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__5; +x_144 = l_Lean_Syntax_node2(x_36, x_143, x_137, x_142); +x_145 = lean_array_push(x_30, x_144); +lean_ctor_set(x_133, 1, x_145); +lean_ctor_set(x_133, 0, x_46); +x_146 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_146, 0, x_133); +x_22 = x_146; +x_23 = x_135; +goto block_28; +} +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_133, 1); +lean_inc(x_147); +lean_dec(x_133); +lean_inc(x_3); +x_148 = lean_mk_syntax_ident(x_3); +x_149 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__9; +x_150 = l_Array_append___rarg(x_149, x_4); +x_151 = lean_array_push(x_150, x_42); +x_152 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17; +lean_inc(x_36); +x_153 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_153, 0, x_36); +lean_ctor_set(x_153, 1, x_152); +lean_ctor_set(x_153, 2, x_151); +x_154 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__5; +x_155 = l_Lean_Syntax_node2(x_36, x_154, x_148, x_153); +x_156 = lean_array_push(x_30, x_155); +x_157 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_157, 0, x_46); +lean_ctor_set(x_157, 1, x_156); +x_158 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_158, 0, x_157); +x_22 = x_158; +x_23 = x_147; +goto block_28; +} +} +} +else +{ +uint8_t x_159; +lean_dec(x_49); +lean_dec(x_46); +lean_dec(x_42); +lean_dec(x_36); +lean_dec(x_30); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_9); +lean_dec(x_3); +x_159 = !lean_is_exclusive(x_50); +if (x_159 == 0) +{ +return x_50; +} +else +{ +lean_object* x_160; lean_object* x_161; lean_object* x_162; +x_160 = lean_ctor_get(x_50, 0); +x_161 = lean_ctor_get(x_50, 1); +lean_inc(x_161); +lean_inc(x_160); +lean_dec(x_50); +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 +{ +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; uint8_t x_175; lean_object* x_176; +x_167 = lean_ctor_get(x_37, 1); +lean_inc(x_167); +lean_dec(x_37); +x_168 = lean_mk_syntax_ident(x_38); +x_169 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__4___closed__3; +lean_inc(x_36); +x_170 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_170, 0, x_36); +lean_ctor_set(x_170, 1, x_169); +x_171 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__4___closed__2; +lean_inc(x_36); +x_172 = l_Lean_Syntax_node1(x_36, x_171, x_170); +x_173 = lean_array_push(x_29, x_172); +x_174 = lean_array_get_size(x_5); +x_175 = lean_nat_dec_lt(x_9, x_174); +lean_dec(x_174); +if (x_175 == 0) +{ +lean_object* x_249; lean_object* x_250; +x_249 = l_Lean_instInhabitedExpr; +x_250 = l_outOfBounds___rarg(x_249); +x_176 = x_250; +goto block_248; +} +else +{ +lean_object* x_251; +x_251 = lean_array_fget(x_5, x_9); +x_176 = x_251; +goto block_248; +} +block_248: +{ +lean_object* x_177; +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_176); +x_177 = lean_infer_type(x_176, x_14, x_15, x_16, x_17, x_167); +if (lean_obj_tag(x_177) == 0) +{ +lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; uint8_t x_182; +x_178 = lean_ctor_get(x_177, 0); +lean_inc(x_178); +x_179 = lean_ctor_get(x_177, 1); +lean_inc(x_179); +lean_dec(x_177); +x_180 = lean_ctor_get(x_2, 0); +x_181 = lean_ctor_get(x_180, 0); +x_182 = l_Lean_Expr_isAppOf(x_178, x_181); +lean_dec(x_178); +if (x_182 == 0) +{ +lean_object* x_183; +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_183 = l_Lean_Meta_isType(x_176, x_14, x_15, x_16, x_17, x_179); +if (lean_obj_tag(x_183) == 0) +{ +lean_object* x_184; uint8_t x_185; +x_184 = lean_ctor_get(x_183, 0); +lean_inc(x_184); +x_185 = lean_unbox(x_184); +lean_dec(x_184); +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; lean_object* x_204; lean_object* x_205; +x_186 = lean_ctor_get(x_183, 1); +lean_inc(x_186); +lean_dec(x_183); +x_187 = lean_ctor_get(x_16, 10); +lean_inc(x_187); +x_188 = lean_st_ref_get(x_17, x_186); +x_189 = lean_ctor_get(x_188, 0); +lean_inc(x_189); +x_190 = lean_ctor_get(x_188, 1); +lean_inc(x_190); +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); +} +x_192 = lean_ctor_get(x_189, 0); +lean_inc(x_192); +lean_dec(x_189); +x_193 = lean_environment_main_module(x_192); +x_194 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__3; +x_195 = l_Lean_addMacroScope(x_193, x_194, x_187); +x_196 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__2; +x_197 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__7; +lean_inc(x_36); +x_198 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_198, 0, x_36); +lean_ctor_set(x_198, 1, x_196); +lean_ctor_set(x_198, 2, x_195); +lean_ctor_set(x_198, 3, x_197); +x_199 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17; +lean_inc(x_36); +x_200 = l_Lean_Syntax_node1(x_36, x_199, x_168); +x_201 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__5; +x_202 = l_Lean_Syntax_node2(x_36, x_201, x_198, x_200); +x_203 = lean_array_push(x_30, x_202); +if (lean_is_scalar(x_191)) { + x_204 = lean_alloc_ctor(0, 2, 0); +} else { + x_204 = x_191; +} +lean_ctor_set(x_204, 0, x_173); +lean_ctor_set(x_204, 1, x_203); +x_205 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_205, 0, x_204); +x_22 = x_205; +x_23 = x_190; +goto block_28; +} +else +{ +lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; +x_206 = lean_ctor_get(x_183, 1); +lean_inc(x_206); +lean_dec(x_183); +x_207 = lean_ctor_get(x_16, 10); +lean_inc(x_207); +x_208 = lean_st_ref_get(x_17, x_206); +x_209 = lean_ctor_get(x_208, 0); +lean_inc(x_209); +x_210 = lean_ctor_get(x_208, 1); +lean_inc(x_210); +if (lean_is_exclusive(x_208)) { + lean_ctor_release(x_208, 0); + lean_ctor_release(x_208, 1); + x_211 = x_208; +} else { + lean_dec_ref(x_208); + x_211 = lean_box(0); +} +x_212 = lean_ctor_get(x_209, 0); +lean_inc(x_212); +lean_dec(x_209); +x_213 = lean_environment_main_module(x_212); +x_214 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__10; +x_215 = l_Lean_addMacroScope(x_213, x_214, x_207); +x_216 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__9; +x_217 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__13; +lean_inc(x_36); +x_218 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_218, 0, x_36); +lean_ctor_set(x_218, 1, x_216); +lean_ctor_set(x_218, 2, x_215); +lean_ctor_set(x_218, 3, x_217); +x_219 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17; +lean_inc(x_36); +x_220 = l_Lean_Syntax_node1(x_36, x_219, x_168); +x_221 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__5; +x_222 = l_Lean_Syntax_node2(x_36, x_221, x_218, x_220); +x_223 = lean_array_push(x_30, x_222); +if (lean_is_scalar(x_211)) { + x_224 = lean_alloc_ctor(0, 2, 0); +} else { + x_224 = x_211; +} +lean_ctor_set(x_224, 0, x_173); +lean_ctor_set(x_224, 1, x_223); +x_225 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_225, 0, x_224); +x_22 = x_225; +x_23 = x_210; +goto block_28; +} +} +else +{ +lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; +lean_dec(x_173); +lean_dec(x_168); +lean_dec(x_36); +lean_dec(x_30); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_9); +lean_dec(x_3); +x_226 = lean_ctor_get(x_183, 0); +lean_inc(x_226); +x_227 = lean_ctor_get(x_183, 1); +lean_inc(x_227); +if (lean_is_exclusive(x_183)) { + lean_ctor_release(x_183, 0); + lean_ctor_release(x_183, 1); + x_228 = x_183; +} else { + lean_dec_ref(x_183); + x_228 = lean_box(0); +} +if (lean_is_scalar(x_228)) { + x_229 = lean_alloc_ctor(1, 2, 0); +} else { + x_229 = x_228; +} +lean_ctor_set(x_229, 0, x_226); +lean_ctor_set(x_229, 1, x_227); +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_dec(x_176); +x_230 = lean_st_ref_get(x_17, x_179); +x_231 = lean_ctor_get(x_230, 1); +lean_inc(x_231); +if (lean_is_exclusive(x_230)) { + lean_ctor_release(x_230, 0); + lean_ctor_release(x_230, 1); + x_232 = x_230; +} else { + lean_dec_ref(x_230); + x_232 = lean_box(0); +} +lean_inc(x_3); +x_233 = lean_mk_syntax_ident(x_3); +x_234 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__9; +x_235 = l_Array_append___rarg(x_234, x_4); +x_236 = lean_array_push(x_235, x_168); +x_237 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17; +lean_inc(x_36); +x_238 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_238, 0, x_36); +lean_ctor_set(x_238, 1, x_237); +lean_ctor_set(x_238, 2, x_236); +x_239 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__5; +x_240 = l_Lean_Syntax_node2(x_36, x_239, x_233, x_238); +x_241 = lean_array_push(x_30, x_240); +if (lean_is_scalar(x_232)) { + x_242 = lean_alloc_ctor(0, 2, 0); +} else { + x_242 = x_232; +} +lean_ctor_set(x_242, 0, x_173); +lean_ctor_set(x_242, 1, x_241); +x_243 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_243, 0, x_242); +x_22 = x_243; +x_23 = x_231; +goto block_28; +} +} +else +{ +lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; +lean_dec(x_176); +lean_dec(x_173); +lean_dec(x_168); +lean_dec(x_36); +lean_dec(x_30); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_9); +lean_dec(x_3); +x_244 = lean_ctor_get(x_177, 0); +lean_inc(x_244); +x_245 = lean_ctor_get(x_177, 1); +lean_inc(x_245); +if (lean_is_exclusive(x_177)) { + lean_ctor_release(x_177, 0); + lean_ctor_release(x_177, 1); + x_246 = x_177; +} else { + lean_dec_ref(x_177); + x_246 = lean_box(0); +} +if (lean_is_scalar(x_246)) { + x_247 = lean_alloc_ctor(1, 2, 0); +} else { + x_247 = x_246; +} +lean_ctor_set(x_247, 0, x_244); +lean_ctor_set(x_247, 1, x_245); +return x_247; +} +} +} +} +} +} +} +static lean_object* _init_l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__6___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("a", 1, 1); +return x_1; +} +} +static lean_object* _init_l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__6___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__6___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18) { +_start: +{ +lean_object* x_19; uint8_t x_20; +x_19 = lean_ctor_get(x_7, 1); +x_20 = lean_nat_dec_lt(x_9, x_19); +if (x_20 == 0) +{ +lean_object* x_21; +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_9); +lean_dec(x_2); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_8); +lean_ctor_set(x_21, 1, x_18); +return x_21; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_161; lean_object* x_162; uint8_t x_163; +x_29 = lean_ctor_get(x_8, 0); +lean_inc(x_29); +x_30 = lean_ctor_get(x_8, 1); +lean_inc(x_30); +lean_dec(x_8); +x_31 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__6___closed__2; +x_32 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_31, x_16, x_17, x_18); +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_mk_syntax_ident(x_33); +lean_inc(x_35); +x_36 = lean_array_push(x_29, x_35); +x_161 = lean_nat_add(x_5, x_9); +x_162 = lean_array_get_size(x_4); +x_163 = lean_nat_dec_lt(x_161, x_162); +lean_dec(x_162); +if (x_163 == 0) +{ +lean_object* x_164; lean_object* x_165; +lean_dec(x_161); +x_164 = l_Lean_instInhabitedExpr; +x_165 = l_outOfBounds___rarg(x_164); +x_37 = x_165; +goto block_160; +} +else +{ +lean_object* x_166; +x_166 = lean_array_fget(x_4, x_161); +lean_dec(x_161); +x_37 = x_166; +goto block_160; +} +block_28: +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_22, 0); +lean_inc(x_24); +lean_dec(x_22); +x_25 = lean_ctor_get(x_7, 2); +x_26 = lean_nat_add(x_9, x_25); +lean_dec(x_9); +x_8 = x_24; +x_9 = x_26; +x_10 = lean_box(0); +x_11 = lean_box(0); +x_18 = x_23; +goto _start; +} +block_160: +{ +lean_object* x_38; +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_37); +x_38 = lean_infer_type(x_37, x_14, x_15, x_16, x_17, x_34); +if (lean_obj_tag(x_38) == 0) +{ +lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; uint8_t x_43; +x_39 = lean_ctor_get(x_38, 0); +lean_inc(x_39); +x_40 = lean_ctor_get(x_38, 1); +lean_inc(x_40); +lean_dec(x_38); +x_41 = lean_ctor_get(x_1, 0); +x_42 = lean_ctor_get(x_41, 0); +x_43 = l_Lean_Expr_isAppOf(x_39, x_42); +lean_dec(x_39); +if (x_43 == 0) +{ +lean_object* x_44; +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_44 = l_Lean_Meta_isType(x_37, x_14, x_15, x_16, x_17, x_40); +if (lean_obj_tag(x_44) == 0) +{ +lean_object* x_45; uint8_t x_46; +x_45 = lean_ctor_get(x_44, 0); +lean_inc(x_45); +x_46 = lean_unbox(x_45); +lean_dec(x_45); +if (x_46 == 0) +{ +lean_object* x_47; lean_object* x_48; uint8_t x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; uint8_t x_53; +x_47 = lean_ctor_get(x_44, 1); +lean_inc(x_47); +lean_dec(x_44); +x_48 = lean_ctor_get(x_16, 5); +lean_inc(x_48); +x_49 = 0; +x_50 = l_Lean_SourceInfo_fromRef(x_48, x_49); +lean_dec(x_48); +x_51 = lean_ctor_get(x_16, 10); +lean_inc(x_51); +x_52 = lean_st_ref_get(x_17, x_47); +x_53 = !lean_is_exclusive(x_52); +if (x_53 == 0) +{ +lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; +x_54 = lean_ctor_get(x_52, 0); +x_55 = lean_ctor_get(x_52, 1); +x_56 = lean_ctor_get(x_54, 0); +lean_inc(x_56); +lean_dec(x_54); +x_57 = lean_environment_main_module(x_56); +x_58 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__3; +x_59 = l_Lean_addMacroScope(x_57, x_58, x_51); +x_60 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__2; +x_61 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__7; +lean_inc(x_50); +x_62 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_62, 0, x_50); +lean_ctor_set(x_62, 1, x_60); +lean_ctor_set(x_62, 2, x_59); +lean_ctor_set(x_62, 3, x_61); +x_63 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17; +lean_inc(x_50); +x_64 = l_Lean_Syntax_node1(x_50, x_63, x_35); +x_65 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__5; +x_66 = l_Lean_Syntax_node2(x_50, x_65, x_62, x_64); +x_67 = lean_array_push(x_30, x_66); +lean_ctor_set(x_52, 1, x_67); +lean_ctor_set(x_52, 0, x_36); +x_68 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_68, 0, x_52); +x_22 = x_68; +x_23 = x_55; +goto block_28; +} +else +{ +lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; +x_69 = lean_ctor_get(x_52, 0); +x_70 = lean_ctor_get(x_52, 1); +lean_inc(x_70); +lean_inc(x_69); +lean_dec(x_52); +x_71 = lean_ctor_get(x_69, 0); +lean_inc(x_71); +lean_dec(x_69); +x_72 = lean_environment_main_module(x_71); +x_73 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__3; +x_74 = l_Lean_addMacroScope(x_72, x_73, x_51); +x_75 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__2; +x_76 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__7; +lean_inc(x_50); +x_77 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_77, 0, x_50); +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_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17; +lean_inc(x_50); +x_79 = l_Lean_Syntax_node1(x_50, x_78, x_35); +x_80 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__5; +x_81 = l_Lean_Syntax_node2(x_50, x_80, x_77, x_79); +x_82 = lean_array_push(x_30, x_81); +x_83 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_83, 0, x_36); +lean_ctor_set(x_83, 1, x_82); +x_84 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_84, 0, x_83); +x_22 = x_84; +x_23 = x_70; +goto block_28; +} +} +else +{ +lean_object* x_85; lean_object* x_86; uint8_t x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; uint8_t x_91; +x_85 = lean_ctor_get(x_44, 1); +lean_inc(x_85); +lean_dec(x_44); +x_86 = lean_ctor_get(x_16, 5); +lean_inc(x_86); +x_87 = 0; +x_88 = l_Lean_SourceInfo_fromRef(x_86, x_87); +lean_dec(x_86); +x_89 = lean_ctor_get(x_16, 10); +lean_inc(x_89); +x_90 = lean_st_ref_get(x_17, x_85); +x_91 = !lean_is_exclusive(x_90); +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_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; +x_92 = lean_ctor_get(x_90, 0); +x_93 = lean_ctor_get(x_90, 1); +x_94 = lean_ctor_get(x_92, 0); +lean_inc(x_94); +lean_dec(x_92); +x_95 = lean_environment_main_module(x_94); +x_96 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__10; +x_97 = l_Lean_addMacroScope(x_95, x_96, x_89); +x_98 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__9; +x_99 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__13; +lean_inc(x_88); +x_100 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_100, 0, x_88); +lean_ctor_set(x_100, 1, x_98); +lean_ctor_set(x_100, 2, x_97); +lean_ctor_set(x_100, 3, x_99); +x_101 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17; +lean_inc(x_88); +x_102 = l_Lean_Syntax_node1(x_88, x_101, x_35); +x_103 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__5; +x_104 = l_Lean_Syntax_node2(x_88, x_103, x_100, x_102); +x_105 = lean_array_push(x_30, x_104); +lean_ctor_set(x_90, 1, x_105); +lean_ctor_set(x_90, 0, x_36); +x_106 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_106, 0, x_90); +x_22 = x_106; +x_23 = x_93; +goto block_28; +} +else +{ +lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; +x_107 = lean_ctor_get(x_90, 0); +x_108 = lean_ctor_get(x_90, 1); +lean_inc(x_108); +lean_inc(x_107); +lean_dec(x_90); +x_109 = lean_ctor_get(x_107, 0); +lean_inc(x_109); +lean_dec(x_107); +x_110 = lean_environment_main_module(x_109); +x_111 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__10; +x_112 = l_Lean_addMacroScope(x_110, x_111, x_89); +x_113 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__9; +x_114 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__13; +lean_inc(x_88); +x_115 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_115, 0, x_88); +lean_ctor_set(x_115, 1, x_113); +lean_ctor_set(x_115, 2, x_112); +lean_ctor_set(x_115, 3, x_114); +x_116 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17; +lean_inc(x_88); +x_117 = l_Lean_Syntax_node1(x_88, x_116, x_35); +x_118 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__5; +x_119 = l_Lean_Syntax_node2(x_88, x_118, x_115, x_117); +x_120 = lean_array_push(x_30, x_119); +x_121 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_121, 0, x_36); +lean_ctor_set(x_121, 1, x_120); +x_122 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_122, 0, x_121); +x_22 = x_122; +x_23 = x_108; +goto block_28; +} +} +} +else +{ +uint8_t x_123; +lean_dec(x_36); +lean_dec(x_35); +lean_dec(x_30); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_9); +lean_dec(x_2); +x_123 = !lean_is_exclusive(x_44); +if (x_123 == 0) +{ +return x_44; +} +else +{ +lean_object* x_124; lean_object* x_125; lean_object* x_126; +x_124 = lean_ctor_get(x_44, 0); +x_125 = lean_ctor_get(x_44, 1); +lean_inc(x_125); +lean_inc(x_124); +lean_dec(x_44); +x_126 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_126, 0, x_124); +lean_ctor_set(x_126, 1, x_125); +return x_126; +} +} +} +else +{ +lean_object* x_127; uint8_t x_128; lean_object* x_129; lean_object* x_130; uint8_t x_131; +lean_dec(x_37); +x_127 = lean_ctor_get(x_16, 5); +lean_inc(x_127); +x_128 = 0; +x_129 = l_Lean_SourceInfo_fromRef(x_127, x_128); +lean_dec(x_127); +x_130 = lean_st_ref_get(x_17, x_40); +x_131 = !lean_is_exclusive(x_130); +if (x_131 == 0) +{ +lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; +x_132 = lean_ctor_get(x_130, 1); +x_133 = lean_ctor_get(x_130, 0); +lean_dec(x_133); +lean_inc(x_2); +x_134 = lean_mk_syntax_ident(x_2); +x_135 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__9; +x_136 = l_Array_append___rarg(x_135, x_3); +x_137 = lean_array_push(x_136, x_35); +x_138 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17; +lean_inc(x_129); +x_139 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_139, 0, x_129); +lean_ctor_set(x_139, 1, x_138); +lean_ctor_set(x_139, 2, x_137); +x_140 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__5; +x_141 = l_Lean_Syntax_node2(x_129, x_140, x_134, x_139); +x_142 = lean_array_push(x_30, x_141); +lean_ctor_set(x_130, 1, x_142); +lean_ctor_set(x_130, 0, x_36); +x_143 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_143, 0, x_130); +x_22 = x_143; +x_23 = x_132; +goto block_28; +} +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; lean_object* x_155; +x_144 = lean_ctor_get(x_130, 1); +lean_inc(x_144); +lean_dec(x_130); +lean_inc(x_2); +x_145 = lean_mk_syntax_ident(x_2); +x_146 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__9; +x_147 = l_Array_append___rarg(x_146, x_3); +x_148 = lean_array_push(x_147, x_35); +x_149 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17; +lean_inc(x_129); +x_150 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_150, 0, x_129); +lean_ctor_set(x_150, 1, x_149); +lean_ctor_set(x_150, 2, x_148); +x_151 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__5; +x_152 = l_Lean_Syntax_node2(x_129, x_151, x_145, x_150); +x_153 = lean_array_push(x_30, x_152); +x_154 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_154, 0, x_36); +lean_ctor_set(x_154, 1, x_153); +x_155 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_155, 0, x_154); +x_22 = x_155; +x_23 = x_144; +goto block_28; +} +} +} +else +{ +uint8_t x_156; +lean_dec(x_37); +lean_dec(x_36); +lean_dec(x_35); +lean_dec(x_30); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_9); +lean_dec(x_2); +x_156 = !lean_is_exclusive(x_38); +if (x_156 == 0) +{ +return x_38; +} +else +{ +lean_object* x_157; lean_object* x_158; lean_object* x_159; +x_157 = lean_ctor_get(x_38, 0); +x_158 = lean_ctor_get(x_38, 1); +lean_inc(x_158); +lean_inc(x_157); +lean_dec(x_38); +x_159 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_159, 0, x_157); +lean_ctor_set(x_159, 1, x_158); +return x_159; +} +} +} +} +} +} +static lean_object* _init_l_List_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__7___lambda__1___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__1; +x_2 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_2, 0, x_1); +lean_ctor_set(x_2, 1, x_1); +return x_2; +} +} +static lean_object* _init_l_List_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__7___lambda__1___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("matchAlt", 8, 8); +return x_1; +} +} +static lean_object* _init_l_List_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__7___lambda__1___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__1; +x_2 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__2; +x_3 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__3; +x_4 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__7___lambda__1___closed__2; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l_List_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__7___lambda__1___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("|", 1, 1); +return x_1; +} +} +static lean_object* _init_l_List_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__7___lambda__1___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(2); +x_2 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__13; +x_3 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_List_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__7___lambda__1___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("=>", 2, 2); +return x_1; +} +} +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___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, size_t x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18) { +_start: +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; +x_19 = lean_box(0); +x_20 = lean_ctor_get(x_1, 2); +x_21 = lean_unsigned_to_nat(0u); +x_22 = lean_unsigned_to_nat(1u); +lean_inc(x_20); +x_23 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_23, 0, x_21); +lean_ctor_set(x_23, 1, x_20); +lean_ctor_set(x_23, 2, x_22); +x_24 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__1; +x_25 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__4(x_23, x_23, x_24, x_21, lean_box(0), lean_box(0), x_12, x_13, x_14, x_15, x_16, x_17, x_18); +lean_dec(x_23); +x_26 = !lean_is_exclusive(x_25); +if (x_26 == 0) +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_27 = lean_ctor_get(x_25, 0); +x_28 = lean_ctor_get(x_25, 1); +x_29 = lean_ctor_get(x_2, 3); +lean_inc(x_29); +x_30 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_30, 0, x_21); +lean_ctor_set(x_30, 1, x_29); +lean_ctor_set(x_30, 2, x_22); +x_31 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__7___lambda__1___closed__1; +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_4); +x_32 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__5(x_3, x_1, x_4, x_5, x_10, x_30, x_30, x_31, x_21, lean_box(0), lean_box(0), x_12, x_13, x_14, x_15, x_16, x_17, x_28); +lean_dec(x_30); +if (lean_obj_tag(x_32) == 0) +{ +lean_object* x_33; lean_object* x_34; uint8_t x_35; +x_33 = lean_ctor_get(x_32, 0); +lean_inc(x_33); +x_34 = lean_ctor_get(x_32, 1); +lean_inc(x_34); +lean_dec(x_32); +x_35 = !lean_is_exclusive(x_33); +if (x_35 == 0) +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_2, 4); +lean_inc(x_36); +x_37 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_37, 0, x_21); +lean_ctor_set(x_37, 1, x_36); +lean_ctor_set(x_37, 2, x_22); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_38 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__6(x_1, x_4, x_5, x_10, x_29, x_37, x_37, x_33, x_21, lean_box(0), lean_box(0), x_12, x_13, x_14, x_15, x_16, x_17, x_34); +lean_dec(x_37); +if (lean_obj_tag(x_38) == 0) +{ +lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; uint8_t x_45; lean_object* x_46; lean_object* x_47; uint8_t x_48; +x_39 = lean_ctor_get(x_38, 0); +lean_inc(x_39); +x_40 = lean_ctor_get(x_38, 1); +lean_inc(x_40); +lean_dec(x_38); +x_41 = lean_ctor_get(x_39, 0); +lean_inc(x_41); +x_42 = lean_ctor_get(x_39, 1); +lean_inc(x_42); +if (lean_is_exclusive(x_39)) { + lean_ctor_release(x_39, 0); + lean_ctor_release(x_39, 1); + x_43 = x_39; +} else { + lean_dec_ref(x_39); + x_43 = lean_box(0); +} +x_44 = lean_ctor_get(x_16, 5); +lean_inc(x_44); +x_45 = 0; +x_46 = l_Lean_SourceInfo_fromRef(x_44, x_45); +lean_dec(x_44); +x_47 = lean_st_ref_get(x_17, x_40); +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; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_117; lean_object* x_118; uint8_t x_119; +x_49 = lean_ctor_get(x_47, 1); +x_50 = lean_ctor_get(x_47, 0); +lean_dec(x_50); +x_51 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__5; +lean_inc(x_46); +lean_ctor_set_tag(x_47, 2); +lean_ctor_set(x_47, 1, x_51); +lean_ctor_set(x_47, 0, x_46); +x_52 = lean_mk_syntax_ident(x_6); +x_53 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__4; +lean_inc(x_46); +x_54 = l_Lean_Syntax_node2(x_46, x_53, x_47, x_52); +x_55 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__9; +x_56 = l_Array_append___rarg(x_55, x_41); +lean_dec(x_41); +x_57 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17; +lean_inc(x_46); +x_58 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_58, 0, x_46); +lean_ctor_set(x_58, 1, x_57); +lean_ctor_set(x_58, 2, x_56); +x_59 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__5; +lean_inc(x_46); +x_60 = l_Lean_Syntax_node2(x_46, x_59, x_54, x_58); +x_61 = lean_array_push(x_27, x_60); +x_117 = lean_ctor_get(x_16, 10); +lean_inc(x_117); +x_118 = lean_st_ref_get(x_17, x_49); +x_119 = !lean_is_exclusive(x_118); +if (x_119 == 0) +{ +lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; +x_120 = lean_ctor_get(x_118, 0); +x_121 = lean_ctor_get(x_118, 1); +x_122 = lean_ctor_get(x_120, 0); +lean_inc(x_122); +lean_dec(x_120); +x_123 = lean_environment_main_module(x_122); +x_124 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__5; +x_125 = l_Lean_addMacroScope(x_123, x_124, x_117); +x_126 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__3; +x_127 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__11; +lean_inc(x_46); +x_128 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_128, 0, x_46); +lean_ctor_set(x_128, 1, x_126); +lean_ctor_set(x_128, 2, x_125); +lean_ctor_set(x_128, 3, x_127); +x_129 = lean_ctor_get(x_8, 0); +lean_inc(x_129); +lean_dec(x_8); +lean_inc(x_129); +x_130 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_19, x_129); +x_131 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__14; +lean_inc(x_46); +lean_ctor_set_tag(x_118, 2); +lean_ctor_set(x_118, 1, x_131); +lean_ctor_set(x_118, 0, x_46); +x_132 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__13; +x_133 = l_Lean_Syntax_SepArray_ofElems(x_132, x_9); +x_134 = l_Array_append___rarg(x_55, x_133); +lean_dec(x_133); +lean_inc(x_46); +x_135 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_135, 0, x_46); +lean_ctor_set(x_135, 1, x_57); +lean_ctor_set(x_135, 2, x_134); +x_136 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__15; +lean_inc(x_46); +lean_ctor_set_tag(x_25, 2); +lean_ctor_set(x_25, 1, x_136); +lean_ctor_set(x_25, 0, x_46); +x_137 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__13; +lean_inc(x_46); +x_138 = l_Lean_Syntax_node3(x_46, x_137, x_118, x_135, x_25); +if (lean_obj_tag(x_130) == 0) +{ +lean_object* x_139; lean_object* x_140; lean_object* x_141; +x_139 = l_Lean_quoteNameMk(x_129); +lean_inc(x_46); +x_140 = l_Lean_Syntax_node2(x_46, x_57, x_139, x_138); +lean_inc(x_46); +x_141 = l_Lean_Syntax_node2(x_46, x_59, x_128, x_140); +x_62 = x_141; +x_63 = x_121; +goto block_116; +} +else +{ +lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; +lean_dec(x_129); +x_142 = lean_ctor_get(x_130, 0); +lean_inc(x_142); +lean_dec(x_130); +x_143 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__18; +x_144 = l_String_intercalate(x_143, x_142); +x_145 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__19; +x_146 = lean_string_append(x_145, x_144); +lean_dec(x_144); +x_147 = lean_box(2); +x_148 = l_Lean_Syntax_mkNameLit(x_146, x_147); +x_149 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_149, 0, x_148); +lean_ctor_set(x_149, 1, x_19); +x_150 = lean_array_mk(x_149); +x_151 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__17; +x_152 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_152, 0, x_147); +lean_ctor_set(x_152, 1, x_151); +lean_ctor_set(x_152, 2, x_150); +lean_inc(x_46); +x_153 = l_Lean_Syntax_node2(x_46, x_57, x_152, x_138); +lean_inc(x_46); +x_154 = l_Lean_Syntax_node2(x_46, x_59, x_128, x_153); +x_62 = x_154; +x_63 = x_121; +goto block_116; +} +} +else +{ +lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; +x_155 = lean_ctor_get(x_118, 0); +x_156 = lean_ctor_get(x_118, 1); +lean_inc(x_156); +lean_inc(x_155); +lean_dec(x_118); +x_157 = lean_ctor_get(x_155, 0); +lean_inc(x_157); +lean_dec(x_155); +x_158 = lean_environment_main_module(x_157); +x_159 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__5; +x_160 = l_Lean_addMacroScope(x_158, x_159, x_117); +x_161 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__3; +x_162 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__11; +lean_inc(x_46); +x_163 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_163, 0, x_46); +lean_ctor_set(x_163, 1, x_161); +lean_ctor_set(x_163, 2, x_160); +lean_ctor_set(x_163, 3, x_162); +x_164 = lean_ctor_get(x_8, 0); +lean_inc(x_164); +lean_dec(x_8); +lean_inc(x_164); +x_165 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_19, x_164); +x_166 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__14; +lean_inc(x_46); +x_167 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_167, 0, x_46); +lean_ctor_set(x_167, 1, x_166); +x_168 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__13; +x_169 = l_Lean_Syntax_SepArray_ofElems(x_168, x_9); +x_170 = l_Array_append___rarg(x_55, x_169); +lean_dec(x_169); +lean_inc(x_46); +x_171 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_171, 0, x_46); +lean_ctor_set(x_171, 1, x_57); +lean_ctor_set(x_171, 2, x_170); +x_172 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__15; +lean_inc(x_46); +lean_ctor_set_tag(x_25, 2); +lean_ctor_set(x_25, 1, x_172); +lean_ctor_set(x_25, 0, x_46); +x_173 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__13; +lean_inc(x_46); +x_174 = l_Lean_Syntax_node3(x_46, x_173, x_167, x_171, x_25); +if (lean_obj_tag(x_165) == 0) +{ +lean_object* x_175; lean_object* x_176; lean_object* x_177; +x_175 = l_Lean_quoteNameMk(x_164); +lean_inc(x_46); +x_176 = l_Lean_Syntax_node2(x_46, x_57, x_175, x_174); +lean_inc(x_46); +x_177 = l_Lean_Syntax_node2(x_46, x_59, x_163, x_176); +x_62 = x_177; +x_63 = x_156; +goto block_116; +} +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; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; +lean_dec(x_164); +x_178 = lean_ctor_get(x_165, 0); +lean_inc(x_178); +lean_dec(x_165); +x_179 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__18; +x_180 = l_String_intercalate(x_179, x_178); +x_181 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__19; +x_182 = lean_string_append(x_181, x_180); +lean_dec(x_180); +x_183 = lean_box(2); +x_184 = l_Lean_Syntax_mkNameLit(x_182, x_183); +x_185 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_185, 0, x_184); +lean_ctor_set(x_185, 1, x_19); +x_186 = lean_array_mk(x_185); +x_187 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__17; +x_188 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_188, 0, x_183); +lean_ctor_set(x_188, 1, x_187); +lean_ctor_set(x_188, 2, x_186); +lean_inc(x_46); +x_189 = l_Lean_Syntax_node2(x_46, x_57, x_188, x_174); +lean_inc(x_46); +x_190 = l_Lean_Syntax_node2(x_46, x_59, x_163, x_189); +x_62 = x_190; +x_63 = x_156; +goto block_116; +} +} +block_116: +{ +lean_object* x_64; uint8_t x_65; +x_64 = l_Lean_Elab_Deriving_ToExpr_mkAppNTerm(x_62, x_42, x_14, x_15, x_16, x_17, x_63); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_42); +x_65 = !lean_is_exclusive(x_64); +if (x_65 == 0) +{ +lean_object* x_66; lean_object* x_67; lean_object* x_68; uint8_t x_69; +x_66 = lean_ctor_get(x_64, 0); +x_67 = lean_ctor_get(x_64, 1); +x_68 = lean_st_ref_get(x_17, x_67); +lean_dec(x_17); +x_69 = !lean_is_exclusive(x_68); +if (x_69 == 0) +{ +lean_object* x_70; lean_object* 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; lean_object* x_81; lean_object* x_82; +x_70 = lean_ctor_get(x_68, 0); +lean_dec(x_70); +x_71 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__7___lambda__1___closed__4; +lean_inc(x_46); +lean_ctor_set_tag(x_64, 2); +lean_ctor_set(x_64, 1, x_71); +lean_ctor_set(x_64, 0, x_46); +x_72 = lean_array_size(x_61); +x_73 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__2(x_72, x_7, x_61); +x_74 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__7___lambda__1___closed__5; +x_75 = l_Lean_mkSepArray(x_73, x_74); +lean_dec(x_73); +x_76 = l_Array_append___rarg(x_55, x_75); +lean_dec(x_75); +lean_inc(x_46); +x_77 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_77, 0, x_46); +lean_ctor_set(x_77, 1, x_57); +lean_ctor_set(x_77, 2, x_76); +lean_inc(x_46); +x_78 = l_Lean_Syntax_node1(x_46, x_57, x_77); +x_79 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__7___lambda__1___closed__6; +lean_inc(x_46); +if (lean_is_scalar(x_43)) { + x_80 = lean_alloc_ctor(2, 2, 0); +} else { + x_80 = x_43; + lean_ctor_set_tag(x_80, 2); +} +lean_ctor_set(x_80, 0, x_46); +lean_ctor_set(x_80, 1, x_79); +x_81 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__7___lambda__1___closed__3; +x_82 = l_Lean_Syntax_node4(x_46, x_81, x_64, x_78, x_80, x_66); +lean_ctor_set(x_68, 0, x_82); +return x_68; +} +else +{ +lean_object* x_83; lean_object* x_84; size_t x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; +x_83 = lean_ctor_get(x_68, 1); +lean_inc(x_83); +lean_dec(x_68); +x_84 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__7___lambda__1___closed__4; +lean_inc(x_46); +lean_ctor_set_tag(x_64, 2); +lean_ctor_set(x_64, 1, x_84); +lean_ctor_set(x_64, 0, x_46); +x_85 = lean_array_size(x_61); +x_86 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__2(x_85, x_7, x_61); +x_87 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__7___lambda__1___closed__5; +x_88 = l_Lean_mkSepArray(x_86, x_87); +lean_dec(x_86); +x_89 = l_Array_append___rarg(x_55, x_88); +lean_dec(x_88); +lean_inc(x_46); +x_90 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_90, 0, x_46); +lean_ctor_set(x_90, 1, x_57); +lean_ctor_set(x_90, 2, x_89); +lean_inc(x_46); +x_91 = l_Lean_Syntax_node1(x_46, x_57, x_90); +x_92 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__7___lambda__1___closed__6; +lean_inc(x_46); +if (lean_is_scalar(x_43)) { + x_93 = lean_alloc_ctor(2, 2, 0); +} else { + x_93 = x_43; + lean_ctor_set_tag(x_93, 2); +} +lean_ctor_set(x_93, 0, x_46); +lean_ctor_set(x_93, 1, x_92); +x_94 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__7___lambda__1___closed__3; +x_95 = l_Lean_Syntax_node4(x_46, x_94, x_64, x_91, x_93, x_66); +x_96 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_96, 0, x_95); +lean_ctor_set(x_96, 1, x_83); +return x_96; +} +} +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; size_t x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; +x_97 = lean_ctor_get(x_64, 0); +x_98 = lean_ctor_get(x_64, 1); +lean_inc(x_98); +lean_inc(x_97); +lean_dec(x_64); +x_99 = lean_st_ref_get(x_17, x_98); +lean_dec(x_17); +x_100 = lean_ctor_get(x_99, 1); +lean_inc(x_100); +if (lean_is_exclusive(x_99)) { + lean_ctor_release(x_99, 0); + lean_ctor_release(x_99, 1); + x_101 = x_99; +} else { + lean_dec_ref(x_99); + x_101 = lean_box(0); +} +x_102 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__7___lambda__1___closed__4; +lean_inc(x_46); +x_103 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_103, 0, x_46); +lean_ctor_set(x_103, 1, x_102); +x_104 = lean_array_size(x_61); +x_105 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__2(x_104, x_7, x_61); +x_106 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__7___lambda__1___closed__5; +x_107 = l_Lean_mkSepArray(x_105, x_106); +lean_dec(x_105); +x_108 = l_Array_append___rarg(x_55, x_107); +lean_dec(x_107); +lean_inc(x_46); +x_109 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_109, 0, x_46); +lean_ctor_set(x_109, 1, x_57); +lean_ctor_set(x_109, 2, x_108); +lean_inc(x_46); +x_110 = l_Lean_Syntax_node1(x_46, x_57, x_109); +x_111 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__7___lambda__1___closed__6; +lean_inc(x_46); +if (lean_is_scalar(x_43)) { + x_112 = lean_alloc_ctor(2, 2, 0); +} else { + x_112 = x_43; + lean_ctor_set_tag(x_112, 2); +} +lean_ctor_set(x_112, 0, x_46); +lean_ctor_set(x_112, 1, x_111); +x_113 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__7___lambda__1___closed__3; +x_114 = l_Lean_Syntax_node4(x_46, x_113, x_103, x_110, x_112, x_97); +if (lean_is_scalar(x_101)) { + x_115 = lean_alloc_ctor(0, 2, 0); +} else { + x_115 = x_101; +} +lean_ctor_set(x_115, 0, x_114); +lean_ctor_set(x_115, 1, x_100); +return x_115; +} +} +} +else +{ +lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_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; +x_191 = lean_ctor_get(x_47, 1); +lean_inc(x_191); +lean_dec(x_47); +x_192 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__5; +lean_inc(x_46); +x_193 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_193, 0, x_46); +lean_ctor_set(x_193, 1, x_192); +x_194 = lean_mk_syntax_ident(x_6); +x_195 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__4; +lean_inc(x_46); +x_196 = l_Lean_Syntax_node2(x_46, x_195, x_193, x_194); +x_197 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__9; +x_198 = l_Array_append___rarg(x_197, x_41); +lean_dec(x_41); +x_199 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17; +lean_inc(x_46); +x_200 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_200, 0, x_46); +lean_ctor_set(x_200, 1, x_199); +lean_ctor_set(x_200, 2, x_198); +x_201 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__5; +lean_inc(x_46); +x_202 = l_Lean_Syntax_node2(x_46, x_201, x_196, x_200); +x_203 = lean_array_push(x_27, x_202); +x_228 = lean_ctor_get(x_16, 10); +lean_inc(x_228); +x_229 = lean_st_ref_get(x_17, x_191); +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; +} else { + lean_dec_ref(x_229); + x_232 = lean_box(0); +} +x_233 = lean_ctor_get(x_230, 0); +lean_inc(x_233); +lean_dec(x_230); +x_234 = lean_environment_main_module(x_233); +x_235 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__5; +x_236 = l_Lean_addMacroScope(x_234, x_235, x_228); +x_237 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__3; +x_238 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__11; +lean_inc(x_46); +x_239 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_239, 0, x_46); +lean_ctor_set(x_239, 1, x_237); +lean_ctor_set(x_239, 2, x_236); +lean_ctor_set(x_239, 3, x_238); +x_240 = lean_ctor_get(x_8, 0); +lean_inc(x_240); +lean_dec(x_8); +lean_inc(x_240); +x_241 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_19, x_240); +x_242 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__14; +lean_inc(x_46); +if (lean_is_scalar(x_232)) { + x_243 = lean_alloc_ctor(2, 2, 0); +} else { + x_243 = x_232; + lean_ctor_set_tag(x_243, 2); +} +lean_ctor_set(x_243, 0, x_46); +lean_ctor_set(x_243, 1, x_242); +x_244 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__13; +x_245 = l_Lean_Syntax_SepArray_ofElems(x_244, x_9); +x_246 = l_Array_append___rarg(x_197, x_245); +lean_dec(x_245); +lean_inc(x_46); +x_247 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_247, 0, x_46); +lean_ctor_set(x_247, 1, x_199); +lean_ctor_set(x_247, 2, x_246); +x_248 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__15; +lean_inc(x_46); +lean_ctor_set_tag(x_25, 2); +lean_ctor_set(x_25, 1, x_248); +lean_ctor_set(x_25, 0, x_46); +x_249 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__13; +lean_inc(x_46); +x_250 = l_Lean_Syntax_node3(x_46, x_249, x_243, x_247, x_25); +if (lean_obj_tag(x_241) == 0) +{ +lean_object* x_251; lean_object* x_252; lean_object* x_253; +x_251 = l_Lean_quoteNameMk(x_240); +lean_inc(x_46); +x_252 = l_Lean_Syntax_node2(x_46, x_199, x_251, x_250); +lean_inc(x_46); +x_253 = l_Lean_Syntax_node2(x_46, x_201, x_239, x_252); +x_204 = x_253; +x_205 = x_231; +goto block_227; +} +else +{ +lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; +lean_dec(x_240); +x_254 = lean_ctor_get(x_241, 0); +lean_inc(x_254); +lean_dec(x_241); +x_255 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__18; +x_256 = l_String_intercalate(x_255, x_254); +x_257 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__19; +x_258 = lean_string_append(x_257, x_256); +lean_dec(x_256); +x_259 = lean_box(2); +x_260 = l_Lean_Syntax_mkNameLit(x_258, x_259); +x_261 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_261, 0, x_260); +lean_ctor_set(x_261, 1, x_19); +x_262 = lean_array_mk(x_261); +x_263 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__17; +x_264 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_264, 0, x_259); +lean_ctor_set(x_264, 1, x_263); +lean_ctor_set(x_264, 2, x_262); +lean_inc(x_46); +x_265 = l_Lean_Syntax_node2(x_46, x_199, x_264, x_250); +lean_inc(x_46); +x_266 = l_Lean_Syntax_node2(x_46, x_201, x_239, x_265); +x_204 = x_266; +x_205 = x_231; +goto block_227; +} +block_227: +{ +lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; size_t 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; +x_206 = l_Lean_Elab_Deriving_ToExpr_mkAppNTerm(x_204, x_42, x_14, x_15, x_16, x_17, x_205); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_42); +x_207 = lean_ctor_get(x_206, 0); +lean_inc(x_207); +x_208 = lean_ctor_get(x_206, 1); +lean_inc(x_208); +if (lean_is_exclusive(x_206)) { + lean_ctor_release(x_206, 0); + lean_ctor_release(x_206, 1); + x_209 = x_206; +} else { + lean_dec_ref(x_206); + x_209 = lean_box(0); +} +x_210 = lean_st_ref_get(x_17, x_208); +lean_dec(x_17); +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_210); + x_212 = lean_box(0); +} +x_213 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__7___lambda__1___closed__4; +lean_inc(x_46); +if (lean_is_scalar(x_209)) { + x_214 = lean_alloc_ctor(2, 2, 0); +} else { + x_214 = x_209; + lean_ctor_set_tag(x_214, 2); +} +lean_ctor_set(x_214, 0, x_46); +lean_ctor_set(x_214, 1, x_213); +x_215 = lean_array_size(x_203); +x_216 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__2(x_215, x_7, x_203); +x_217 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__7___lambda__1___closed__5; +x_218 = l_Lean_mkSepArray(x_216, x_217); +lean_dec(x_216); +x_219 = l_Array_append___rarg(x_197, x_218); +lean_dec(x_218); +lean_inc(x_46); +x_220 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_220, 0, x_46); +lean_ctor_set(x_220, 1, x_199); +lean_ctor_set(x_220, 2, x_219); +lean_inc(x_46); +x_221 = l_Lean_Syntax_node1(x_46, x_199, x_220); +x_222 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__7___lambda__1___closed__6; +lean_inc(x_46); +if (lean_is_scalar(x_43)) { + x_223 = lean_alloc_ctor(2, 2, 0); +} else { + x_223 = x_43; + lean_ctor_set_tag(x_223, 2); +} +lean_ctor_set(x_223, 0, x_46); +lean_ctor_set(x_223, 1, x_222); +x_224 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__7___lambda__1___closed__3; +x_225 = l_Lean_Syntax_node4(x_46, x_224, x_214, x_221, x_223, x_207); +if (lean_is_scalar(x_212)) { + x_226 = lean_alloc_ctor(0, 2, 0); +} else { + x_226 = x_212; +} +lean_ctor_set(x_226, 0, x_225); +lean_ctor_set(x_226, 1, x_211); +return x_226; +} +} +} +else +{ +uint8_t x_267; +lean_free_object(x_25); +lean_dec(x_27); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_8); +lean_dec(x_6); +x_267 = !lean_is_exclusive(x_38); +if (x_267 == 0) +{ +return x_38; +} +else +{ +lean_object* x_268; lean_object* x_269; lean_object* x_270; +x_268 = lean_ctor_get(x_38, 0); +x_269 = lean_ctor_get(x_38, 1); +lean_inc(x_269); +lean_inc(x_268); +lean_dec(x_38); +x_270 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_270, 0, x_268); +lean_ctor_set(x_270, 1, x_269); +return x_270; +} +} +} +else +{ +lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; +x_271 = lean_ctor_get(x_33, 0); +x_272 = lean_ctor_get(x_33, 1); +lean_inc(x_272); +lean_inc(x_271); +lean_dec(x_33); +x_273 = lean_ctor_get(x_2, 4); +lean_inc(x_273); +x_274 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_274, 0, x_21); +lean_ctor_set(x_274, 1, x_273); +lean_ctor_set(x_274, 2, x_22); +x_275 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_275, 0, x_271); +lean_ctor_set(x_275, 1, x_272); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_276 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__6(x_1, x_4, x_5, x_10, x_29, x_274, x_274, x_275, x_21, lean_box(0), lean_box(0), x_12, x_13, x_14, x_15, x_16, x_17, x_34); +lean_dec(x_274); +if (lean_obj_tag(x_276) == 0) +{ +lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; uint8_t x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; +x_277 = lean_ctor_get(x_276, 0); +lean_inc(x_277); +x_278 = lean_ctor_get(x_276, 1); +lean_inc(x_278); +lean_dec(x_276); +x_279 = lean_ctor_get(x_277, 0); +lean_inc(x_279); +x_280 = lean_ctor_get(x_277, 1); +lean_inc(x_280); +if (lean_is_exclusive(x_277)) { + lean_ctor_release(x_277, 0); + lean_ctor_release(x_277, 1); + x_281 = x_277; +} else { + lean_dec_ref(x_277); + x_281 = lean_box(0); +} +x_282 = lean_ctor_get(x_16, 5); +lean_inc(x_282); +x_283 = 0; +x_284 = l_Lean_SourceInfo_fromRef(x_282, x_283); +lean_dec(x_282); +x_285 = lean_st_ref_get(x_17, x_278); +x_286 = lean_ctor_get(x_285, 1); +lean_inc(x_286); +if (lean_is_exclusive(x_285)) { + lean_ctor_release(x_285, 0); + lean_ctor_release(x_285, 1); + x_287 = x_285; +} else { + lean_dec_ref(x_285); + x_287 = lean_box(0); +} +x_288 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__5; +lean_inc(x_284); +if (lean_is_scalar(x_287)) { + x_289 = lean_alloc_ctor(2, 2, 0); +} else { + x_289 = x_287; + lean_ctor_set_tag(x_289, 2); +} +lean_ctor_set(x_289, 0, x_284); +lean_ctor_set(x_289, 1, x_288); +x_290 = lean_mk_syntax_ident(x_6); +x_291 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__4; +lean_inc(x_284); +x_292 = l_Lean_Syntax_node2(x_284, x_291, x_289, x_290); +x_293 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__9; +x_294 = l_Array_append___rarg(x_293, x_279); +lean_dec(x_279); +x_295 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17; +lean_inc(x_284); +x_296 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_296, 0, x_284); +lean_ctor_set(x_296, 1, x_295); +lean_ctor_set(x_296, 2, x_294); +x_297 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__5; +lean_inc(x_284); +x_298 = l_Lean_Syntax_node2(x_284, x_297, x_292, x_296); +x_299 = lean_array_push(x_27, x_298); +x_324 = lean_ctor_get(x_16, 10); +lean_inc(x_324); +x_325 = lean_st_ref_get(x_17, x_286); +x_326 = lean_ctor_get(x_325, 0); +lean_inc(x_326); +x_327 = lean_ctor_get(x_325, 1); +lean_inc(x_327); +if (lean_is_exclusive(x_325)) { + lean_ctor_release(x_325, 0); + lean_ctor_release(x_325, 1); + x_328 = x_325; +} else { + lean_dec_ref(x_325); + x_328 = lean_box(0); +} +x_329 = lean_ctor_get(x_326, 0); +lean_inc(x_329); +lean_dec(x_326); +x_330 = lean_environment_main_module(x_329); +x_331 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__5; +x_332 = l_Lean_addMacroScope(x_330, x_331, x_324); +x_333 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__3; +x_334 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__11; +lean_inc(x_284); +x_335 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_335, 0, x_284); +lean_ctor_set(x_335, 1, x_333); +lean_ctor_set(x_335, 2, x_332); +lean_ctor_set(x_335, 3, x_334); +x_336 = lean_ctor_get(x_8, 0); +lean_inc(x_336); +lean_dec(x_8); +lean_inc(x_336); +x_337 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_19, x_336); +x_338 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__14; +lean_inc(x_284); +if (lean_is_scalar(x_328)) { + x_339 = lean_alloc_ctor(2, 2, 0); +} else { + x_339 = x_328; + lean_ctor_set_tag(x_339, 2); +} +lean_ctor_set(x_339, 0, x_284); +lean_ctor_set(x_339, 1, x_338); +x_340 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__13; +x_341 = l_Lean_Syntax_SepArray_ofElems(x_340, x_9); +x_342 = l_Array_append___rarg(x_293, x_341); +lean_dec(x_341); +lean_inc(x_284); +x_343 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_343, 0, x_284); +lean_ctor_set(x_343, 1, x_295); +lean_ctor_set(x_343, 2, x_342); +x_344 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__15; +lean_inc(x_284); +lean_ctor_set_tag(x_25, 2); +lean_ctor_set(x_25, 1, x_344); +lean_ctor_set(x_25, 0, x_284); +x_345 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__13; +lean_inc(x_284); +x_346 = l_Lean_Syntax_node3(x_284, x_345, x_339, x_343, x_25); +if (lean_obj_tag(x_337) == 0) +{ +lean_object* x_347; lean_object* x_348; lean_object* x_349; +x_347 = l_Lean_quoteNameMk(x_336); +lean_inc(x_284); +x_348 = l_Lean_Syntax_node2(x_284, x_295, x_347, x_346); +lean_inc(x_284); +x_349 = l_Lean_Syntax_node2(x_284, x_297, x_335, x_348); +x_300 = x_349; +x_301 = x_327; +goto block_323; +} +else +{ +lean_object* x_350; lean_object* x_351; lean_object* x_352; lean_object* x_353; lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; +lean_dec(x_336); +x_350 = lean_ctor_get(x_337, 0); +lean_inc(x_350); +lean_dec(x_337); +x_351 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__18; +x_352 = l_String_intercalate(x_351, x_350); +x_353 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__19; +x_354 = lean_string_append(x_353, x_352); +lean_dec(x_352); +x_355 = lean_box(2); +x_356 = l_Lean_Syntax_mkNameLit(x_354, x_355); +x_357 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_357, 0, x_356); +lean_ctor_set(x_357, 1, x_19); +x_358 = lean_array_mk(x_357); +x_359 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__17; +x_360 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_360, 0, x_355); +lean_ctor_set(x_360, 1, x_359); +lean_ctor_set(x_360, 2, x_358); +lean_inc(x_284); +x_361 = l_Lean_Syntax_node2(x_284, x_295, x_360, x_346); +lean_inc(x_284); +x_362 = l_Lean_Syntax_node2(x_284, x_297, x_335, x_361); +x_300 = x_362; +x_301 = x_327; +goto block_323; +} +block_323: +{ +lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; size_t x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; +x_302 = l_Lean_Elab_Deriving_ToExpr_mkAppNTerm(x_300, x_280, x_14, x_15, x_16, x_17, x_301); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_280); +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; +} else { + lean_dec_ref(x_302); + x_305 = lean_box(0); +} +x_306 = lean_st_ref_get(x_17, x_304); +lean_dec(x_17); +x_307 = lean_ctor_get(x_306, 1); +lean_inc(x_307); +if (lean_is_exclusive(x_306)) { + lean_ctor_release(x_306, 0); + lean_ctor_release(x_306, 1); + x_308 = x_306; +} else { + lean_dec_ref(x_306); + x_308 = lean_box(0); +} +x_309 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__7___lambda__1___closed__4; +lean_inc(x_284); +if (lean_is_scalar(x_305)) { + x_310 = lean_alloc_ctor(2, 2, 0); +} else { + x_310 = x_305; + lean_ctor_set_tag(x_310, 2); +} +lean_ctor_set(x_310, 0, x_284); +lean_ctor_set(x_310, 1, x_309); +x_311 = lean_array_size(x_299); +x_312 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__2(x_311, x_7, x_299); +x_313 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__7___lambda__1___closed__5; +x_314 = l_Lean_mkSepArray(x_312, x_313); +lean_dec(x_312); +x_315 = l_Array_append___rarg(x_293, x_314); +lean_dec(x_314); +lean_inc(x_284); +x_316 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_316, 0, x_284); +lean_ctor_set(x_316, 1, x_295); +lean_ctor_set(x_316, 2, x_315); +lean_inc(x_284); +x_317 = l_Lean_Syntax_node1(x_284, x_295, x_316); +x_318 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__7___lambda__1___closed__6; +lean_inc(x_284); +if (lean_is_scalar(x_281)) { + x_319 = lean_alloc_ctor(2, 2, 0); +} else { + x_319 = x_281; + lean_ctor_set_tag(x_319, 2); +} +lean_ctor_set(x_319, 0, x_284); +lean_ctor_set(x_319, 1, x_318); +x_320 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__7___lambda__1___closed__3; +x_321 = l_Lean_Syntax_node4(x_284, x_320, x_310, x_317, x_319, x_303); +if (lean_is_scalar(x_308)) { + x_322 = lean_alloc_ctor(0, 2, 0); +} else { + x_322 = x_308; +} +lean_ctor_set(x_322, 0, x_321); +lean_ctor_set(x_322, 1, x_307); +return x_322; +} +} +else +{ +lean_object* x_363; lean_object* x_364; lean_object* x_365; lean_object* x_366; +lean_free_object(x_25); +lean_dec(x_27); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_8); +lean_dec(x_6); +x_363 = lean_ctor_get(x_276, 0); +lean_inc(x_363); +x_364 = lean_ctor_get(x_276, 1); +lean_inc(x_364); +if (lean_is_exclusive(x_276)) { + lean_ctor_release(x_276, 0); + lean_ctor_release(x_276, 1); + x_365 = x_276; +} else { + lean_dec_ref(x_276); + 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; +} +} +} +else +{ +uint8_t x_367; +lean_free_object(x_25); +lean_dec(x_27); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_4); +x_367 = !lean_is_exclusive(x_32); +if (x_367 == 0) +{ +return x_32; +} +else +{ +lean_object* x_368; lean_object* x_369; lean_object* x_370; +x_368 = lean_ctor_get(x_32, 0); +x_369 = lean_ctor_get(x_32, 1); +lean_inc(x_369); +lean_inc(x_368); +lean_dec(x_32); +x_370 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_370, 0, x_368); +lean_ctor_set(x_370, 1, x_369); +return x_370; +} +} +} +else +{ +lean_object* x_371; lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; lean_object* x_376; +x_371 = lean_ctor_get(x_25, 0); +x_372 = lean_ctor_get(x_25, 1); +lean_inc(x_372); +lean_inc(x_371); +lean_dec(x_25); +x_373 = lean_ctor_get(x_2, 3); +lean_inc(x_373); +x_374 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_374, 0, x_21); +lean_ctor_set(x_374, 1, x_373); +lean_ctor_set(x_374, 2, x_22); +x_375 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__7___lambda__1___closed__1; +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_4); +x_376 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__5(x_3, x_1, x_4, x_5, x_10, x_374, x_374, x_375, x_21, lean_box(0), lean_box(0), x_12, x_13, x_14, x_15, x_16, x_17, x_372); +lean_dec(x_374); +if (lean_obj_tag(x_376) == 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_object* x_385; +x_377 = lean_ctor_get(x_376, 0); +lean_inc(x_377); +x_378 = lean_ctor_get(x_376, 1); +lean_inc(x_378); +lean_dec(x_376); +x_379 = lean_ctor_get(x_377, 0); +lean_inc(x_379); +x_380 = lean_ctor_get(x_377, 1); +lean_inc(x_380); +if (lean_is_exclusive(x_377)) { + lean_ctor_release(x_377, 0); + lean_ctor_release(x_377, 1); + x_381 = x_377; +} else { + lean_dec_ref(x_377); + x_381 = lean_box(0); +} +x_382 = lean_ctor_get(x_2, 4); +lean_inc(x_382); +x_383 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_383, 0, x_21); +lean_ctor_set(x_383, 1, x_382); +lean_ctor_set(x_383, 2, x_22); +if (lean_is_scalar(x_381)) { + x_384 = lean_alloc_ctor(0, 2, 0); +} else { + x_384 = x_381; +} +lean_ctor_set(x_384, 0, x_379); +lean_ctor_set(x_384, 1, x_380); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_385 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__6(x_1, x_4, x_5, x_10, x_373, x_383, x_383, x_384, x_21, lean_box(0), lean_box(0), x_12, x_13, x_14, x_15, x_16, x_17, x_378); +lean_dec(x_383); +if (lean_obj_tag(x_385) == 0) +{ +lean_object* x_386; lean_object* x_387; lean_object* x_388; lean_object* x_389; lean_object* x_390; lean_object* x_391; uint8_t x_392; lean_object* x_393; lean_object* x_394; lean_object* x_395; lean_object* x_396; lean_object* x_397; lean_object* x_398; lean_object* x_399; lean_object* x_400; lean_object* x_401; lean_object* x_402; lean_object* x_403; lean_object* x_404; lean_object* x_405; lean_object* x_406; lean_object* x_407; lean_object* x_408; lean_object* x_409; lean_object* x_410; lean_object* x_433; lean_object* x_434; lean_object* x_435; lean_object* x_436; lean_object* x_437; lean_object* x_438; lean_object* x_439; lean_object* x_440; lean_object* x_441; lean_object* x_442; lean_object* x_443; lean_object* x_444; lean_object* x_445; lean_object* x_446; lean_object* x_447; lean_object* x_448; lean_object* x_449; lean_object* x_450; lean_object* x_451; lean_object* x_452; lean_object* x_453; lean_object* x_454; lean_object* x_455; lean_object* x_456; +x_386 = lean_ctor_get(x_385, 0); +lean_inc(x_386); +x_387 = lean_ctor_get(x_385, 1); +lean_inc(x_387); +lean_dec(x_385); +x_388 = lean_ctor_get(x_386, 0); +lean_inc(x_388); +x_389 = lean_ctor_get(x_386, 1); +lean_inc(x_389); +if (lean_is_exclusive(x_386)) { + lean_ctor_release(x_386, 0); + lean_ctor_release(x_386, 1); + x_390 = x_386; +} else { + lean_dec_ref(x_386); + x_390 = lean_box(0); +} +x_391 = lean_ctor_get(x_16, 5); +lean_inc(x_391); +x_392 = 0; +x_393 = l_Lean_SourceInfo_fromRef(x_391, x_392); +lean_dec(x_391); +x_394 = lean_st_ref_get(x_17, x_387); +x_395 = lean_ctor_get(x_394, 1); +lean_inc(x_395); +if (lean_is_exclusive(x_394)) { + lean_ctor_release(x_394, 0); + lean_ctor_release(x_394, 1); + x_396 = x_394; +} else { + lean_dec_ref(x_394); + x_396 = lean_box(0); +} +x_397 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__5; +lean_inc(x_393); +if (lean_is_scalar(x_396)) { + x_398 = lean_alloc_ctor(2, 2, 0); +} else { + x_398 = x_396; + lean_ctor_set_tag(x_398, 2); +} +lean_ctor_set(x_398, 0, x_393); +lean_ctor_set(x_398, 1, x_397); +x_399 = lean_mk_syntax_ident(x_6); +x_400 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__4; +lean_inc(x_393); +x_401 = l_Lean_Syntax_node2(x_393, x_400, x_398, x_399); +x_402 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__9; +x_403 = l_Array_append___rarg(x_402, x_388); +lean_dec(x_388); +x_404 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17; +lean_inc(x_393); +x_405 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_405, 0, x_393); +lean_ctor_set(x_405, 1, x_404); +lean_ctor_set(x_405, 2, x_403); +x_406 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__5; +lean_inc(x_393); +x_407 = l_Lean_Syntax_node2(x_393, x_406, x_401, x_405); +x_408 = lean_array_push(x_371, x_407); +x_433 = lean_ctor_get(x_16, 10); +lean_inc(x_433); +x_434 = lean_st_ref_get(x_17, x_395); +x_435 = lean_ctor_get(x_434, 0); +lean_inc(x_435); +x_436 = lean_ctor_get(x_434, 1); +lean_inc(x_436); +if (lean_is_exclusive(x_434)) { + lean_ctor_release(x_434, 0); + lean_ctor_release(x_434, 1); + x_437 = x_434; +} else { + lean_dec_ref(x_434); + x_437 = lean_box(0); +} +x_438 = lean_ctor_get(x_435, 0); +lean_inc(x_438); +lean_dec(x_435); +x_439 = lean_environment_main_module(x_438); +x_440 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__5; +x_441 = l_Lean_addMacroScope(x_439, x_440, x_433); +x_442 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__3; +x_443 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__11; +lean_inc(x_393); +x_444 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_444, 0, x_393); +lean_ctor_set(x_444, 1, x_442); +lean_ctor_set(x_444, 2, x_441); +lean_ctor_set(x_444, 3, x_443); +x_445 = lean_ctor_get(x_8, 0); +lean_inc(x_445); +lean_dec(x_8); +lean_inc(x_445); +x_446 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_19, x_445); +x_447 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__14; +lean_inc(x_393); +if (lean_is_scalar(x_437)) { + x_448 = lean_alloc_ctor(2, 2, 0); +} else { + x_448 = x_437; + lean_ctor_set_tag(x_448, 2); +} +lean_ctor_set(x_448, 0, x_393); +lean_ctor_set(x_448, 1, x_447); +x_449 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__13; +x_450 = l_Lean_Syntax_SepArray_ofElems(x_449, x_9); +x_451 = l_Array_append___rarg(x_402, x_450); +lean_dec(x_450); +lean_inc(x_393); +x_452 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_452, 0, x_393); +lean_ctor_set(x_452, 1, x_404); +lean_ctor_set(x_452, 2, x_451); +x_453 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__15; +lean_inc(x_393); +x_454 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_454, 0, x_393); +lean_ctor_set(x_454, 1, x_453); +x_455 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__13; +lean_inc(x_393); +x_456 = l_Lean_Syntax_node3(x_393, x_455, x_448, x_452, x_454); +if (lean_obj_tag(x_446) == 0) +{ +lean_object* x_457; lean_object* x_458; lean_object* x_459; +x_457 = l_Lean_quoteNameMk(x_445); +lean_inc(x_393); +x_458 = l_Lean_Syntax_node2(x_393, x_404, x_457, x_456); +lean_inc(x_393); +x_459 = l_Lean_Syntax_node2(x_393, x_406, x_444, x_458); +x_409 = x_459; +x_410 = x_436; +goto block_432; +} +else +{ +lean_object* x_460; lean_object* x_461; lean_object* x_462; lean_object* x_463; lean_object* x_464; lean_object* x_465; lean_object* x_466; lean_object* x_467; lean_object* x_468; lean_object* x_469; lean_object* x_470; lean_object* x_471; lean_object* x_472; +lean_dec(x_445); +x_460 = lean_ctor_get(x_446, 0); +lean_inc(x_460); +lean_dec(x_446); +x_461 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__18; +x_462 = l_String_intercalate(x_461, x_460); +x_463 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__19; +x_464 = lean_string_append(x_463, x_462); +lean_dec(x_462); +x_465 = lean_box(2); +x_466 = l_Lean_Syntax_mkNameLit(x_464, x_465); +x_467 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_467, 0, x_466); +lean_ctor_set(x_467, 1, x_19); +x_468 = lean_array_mk(x_467); +x_469 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__17; +x_470 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_470, 0, x_465); +lean_ctor_set(x_470, 1, x_469); +lean_ctor_set(x_470, 2, x_468); +lean_inc(x_393); +x_471 = l_Lean_Syntax_node2(x_393, x_404, x_470, x_456); +lean_inc(x_393); +x_472 = l_Lean_Syntax_node2(x_393, x_406, x_444, x_471); +x_409 = x_472; +x_410 = x_436; +goto block_432; +} +block_432: +{ +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; size_t x_420; lean_object* x_421; lean_object* x_422; lean_object* x_423; lean_object* x_424; lean_object* x_425; lean_object* x_426; lean_object* x_427; lean_object* x_428; lean_object* x_429; lean_object* x_430; lean_object* x_431; +x_411 = l_Lean_Elab_Deriving_ToExpr_mkAppNTerm(x_409, x_389, x_14, x_15, x_16, x_17, x_410); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_389); +x_412 = lean_ctor_get(x_411, 0); +lean_inc(x_412); +x_413 = lean_ctor_get(x_411, 1); +lean_inc(x_413); +if (lean_is_exclusive(x_411)) { + lean_ctor_release(x_411, 0); + lean_ctor_release(x_411, 1); + x_414 = x_411; +} else { + lean_dec_ref(x_411); + x_414 = lean_box(0); +} +x_415 = lean_st_ref_get(x_17, x_413); +lean_dec(x_17); +x_416 = lean_ctor_get(x_415, 1); +lean_inc(x_416); +if (lean_is_exclusive(x_415)) { + lean_ctor_release(x_415, 0); + lean_ctor_release(x_415, 1); + x_417 = x_415; +} else { + lean_dec_ref(x_415); + x_417 = lean_box(0); +} +x_418 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__7___lambda__1___closed__4; +lean_inc(x_393); +if (lean_is_scalar(x_414)) { + x_419 = lean_alloc_ctor(2, 2, 0); +} else { + x_419 = x_414; + lean_ctor_set_tag(x_419, 2); +} +lean_ctor_set(x_419, 0, x_393); +lean_ctor_set(x_419, 1, x_418); +x_420 = lean_array_size(x_408); +x_421 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__2(x_420, x_7, x_408); +x_422 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__7___lambda__1___closed__5; +x_423 = l_Lean_mkSepArray(x_421, x_422); +lean_dec(x_421); +x_424 = l_Array_append___rarg(x_402, x_423); +lean_dec(x_423); +lean_inc(x_393); +x_425 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_425, 0, x_393); +lean_ctor_set(x_425, 1, x_404); +lean_ctor_set(x_425, 2, x_424); +lean_inc(x_393); +x_426 = l_Lean_Syntax_node1(x_393, x_404, x_425); +x_427 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__7___lambda__1___closed__6; +lean_inc(x_393); +if (lean_is_scalar(x_390)) { + x_428 = lean_alloc_ctor(2, 2, 0); +} else { + x_428 = x_390; + lean_ctor_set_tag(x_428, 2); +} +lean_ctor_set(x_428, 0, x_393); +lean_ctor_set(x_428, 1, x_427); +x_429 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__7___lambda__1___closed__3; +x_430 = l_Lean_Syntax_node4(x_393, x_429, x_419, x_426, x_428, x_412); +if (lean_is_scalar(x_417)) { + x_431 = lean_alloc_ctor(0, 2, 0); +} else { + x_431 = x_417; +} +lean_ctor_set(x_431, 0, x_430); +lean_ctor_set(x_431, 1, x_416); +return x_431; +} +} +else +{ +lean_object* x_473; lean_object* x_474; lean_object* x_475; lean_object* x_476; +lean_dec(x_371); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_8); +lean_dec(x_6); +x_473 = lean_ctor_get(x_385, 0); +lean_inc(x_473); +x_474 = lean_ctor_get(x_385, 1); +lean_inc(x_474); +if (lean_is_exclusive(x_385)) { + lean_ctor_release(x_385, 0); + lean_ctor_release(x_385, 1); + x_475 = x_385; +} else { + lean_dec_ref(x_385); + x_475 = lean_box(0); +} +if (lean_is_scalar(x_475)) { + x_476 = lean_alloc_ctor(1, 2, 0); +} else { + x_476 = x_475; +} +lean_ctor_set(x_476, 0, x_473); +lean_ctor_set(x_476, 1, x_474); +return x_476; +} +} +else +{ +lean_object* x_477; lean_object* x_478; lean_object* x_479; lean_object* x_480; +lean_dec(x_371); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_4); +x_477 = lean_ctor_get(x_376, 0); +lean_inc(x_477); +x_478 = lean_ctor_get(x_376, 1); +lean_inc(x_478); +if (lean_is_exclusive(x_376)) { + lean_ctor_release(x_376, 0); + lean_ctor_release(x_376, 1); + x_479 = x_376; +} else { + lean_dec_ref(x_376); + x_479 = lean_box(0); +} +if (lean_is_scalar(x_479)) { + x_480 = lean_alloc_ctor(1, 2, 0); +} else { + x_480 = x_479; +} +lean_ctor_set(x_480, 0, x_477); +lean_ctor_set(x_480, 1, x_478); +return x_480; +} +} +} +} +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, 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: +{ +if (lean_obj_tag(x_10) == 0) +{ +lean_object* x_20; +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_11); +lean_ctor_set(x_20, 1, x_19); +return x_20; +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_10, 0); +lean_inc(x_21); +x_22 = lean_ctor_get(x_10, 1); +lean_inc(x_22); +lean_dec(x_10); +lean_inc(x_13); +lean_inc(x_21); +x_23 = l_Lean_getConstInfoCtor___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__2(x_21, x_13, x_14, x_15, x_16, x_17, x_18, x_19); +if (lean_obj_tag(x_23) == 0) +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; uint8_t x_30; lean_object* x_31; +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +x_25 = lean_ctor_get(x_23, 1); +lean_inc(x_25); +lean_dec(x_23); +x_26 = lean_ctor_get(x_24, 0); +lean_inc(x_26); +x_27 = lean_ctor_get(x_26, 2); +lean_inc(x_27); +x_28 = lean_box_usize(x_5); +lean_inc(x_6); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_1); +lean_inc(x_2); +x_29 = lean_alloc_closure((void*)(l_List_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__7___lambda__1___boxed), 18, 9); +lean_closure_set(x_29, 0, x_2); +lean_closure_set(x_29, 1, x_24); +lean_closure_set(x_29, 2, x_1); +lean_closure_set(x_29, 3, x_3); +lean_closure_set(x_29, 4, x_4); +lean_closure_set(x_29, 5, x_21); +lean_closure_set(x_29, 6, x_28); +lean_closure_set(x_29, 7, x_26); +lean_closure_set(x_29, 8, x_6); +x_30 = 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_31 = l_Lean_Meta_forallTelescopeReducing___at_Lean_Elab_Deriving_mkInductArgNames___spec__2___rarg(x_27, x_29, x_30, x_13, x_14, x_15, x_16, x_17, x_18, x_25); +if (lean_obj_tag(x_31) == 0) +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_31, 0); +lean_inc(x_32); +x_33 = lean_ctor_get(x_31, 1); +lean_inc(x_33); +lean_dec(x_31); +x_34 = lean_array_push(x_11, x_32); +x_10 = x_22; +x_11 = x_34; +x_12 = lean_box(0); +x_19 = x_33; +goto _start; +} +else +{ +uint8_t x_36; +lean_dec(x_22); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_36 = !lean_is_exclusive(x_31); +if (x_36 == 0) +{ +return x_31; +} +else +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_37 = lean_ctor_get(x_31, 0); +x_38 = lean_ctor_get(x_31, 1); +lean_inc(x_38); +lean_inc(x_37); +lean_dec(x_31); +x_39 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_39, 0, x_37); +lean_ctor_set(x_39, 1, x_38); +return x_39; +} +} +} +else +{ +uint8_t x_40; +lean_dec(x_22); +lean_dec(x_21); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_40 = !lean_is_exclusive(x_23); +if (x_40 == 0) +{ +return x_23; +} +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_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_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__8(size_t x_1, size_t x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; +x_4 = lean_usize_dec_lt(x_2, x_1); +if (x_4 == 0) +{ +return x_3; +} +else +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; size_t x_8; size_t x_9; lean_object* x_10; +x_5 = lean_array_uget(x_3, x_2); +x_6 = lean_unsigned_to_nat(0u); +x_7 = lean_array_uset(x_3, x_2, x_6); +x_8 = 1; +x_9 = lean_usize_add(x_2, x_8); +x_10 = lean_array_uset(x_7, x_2, x_5); +x_2 = x_9; +x_3 = x_10; +goto _start; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +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; +x_12 = lean_array_size(x_4); +x_13 = 0; +lean_inc(x_9); +lean_inc(x_4); +x_14 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__1(x_12, x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +x_16 = lean_ctor_get(x_14, 1); +lean_inc(x_16); +lean_dec(x_14); +x_17 = lean_ctor_get(x_2, 4); +lean_inc(x_17); +x_18 = lean_box(0); +x_19 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__1; +lean_inc(x_17); +x_20 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__7(x_1, x_2, x_3, x_4, x_13, x_15, x_17, x_18, x_17, x_17, x_19, lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10, x_16); +lean_dec(x_17); +if (lean_obj_tag(x_20) == 0) +{ +uint8_t x_21; +x_21 = !lean_is_exclusive(x_20); +if (x_21 == 0) +{ +lean_object* x_22; size_t x_23; lean_object* x_24; +x_22 = lean_ctor_get(x_20, 0); +x_23 = lean_array_size(x_22); +x_24 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__8(x_23, x_13, x_22); +lean_ctor_set(x_20, 0, x_24); +return x_20; +} +else +{ +lean_object* x_25; lean_object* x_26; size_t 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_dec(x_20); +x_27 = lean_array_size(x_25); +x_28 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__8(x_27, x_13, x_25); +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_30; +x_30 = !lean_is_exclusive(x_20); +if (x_30 == 0) +{ +return x_20; +} +else +{ +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_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; +} +} +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___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: +{ +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_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__1(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); +lean_dec(x_5); +lean_dec(x_4); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; +x_9 = l_Lean_throwError___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___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_getConstInfoCtor___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___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_getConstInfoCtor___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___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_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +lean_object* x_14; +x_14 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_2); +lean_dec(x_1); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__5___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +lean_object* x_18 = _args[17]; +_start: +{ +lean_object* x_19; +x_19 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__5(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +lean_dec(x_13); +lean_dec(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); +return x_19; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__6___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +lean_object* x_18 = _args[17]; +_start: +{ +lean_object* x_19; +x_19 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__6(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +return x_19; +} +} +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__7___lambda__1___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +lean_object* x_18 = _args[17]; +_start: +{ +size_t x_19; lean_object* x_20; +x_19 = lean_unbox_usize(x_7); +lean_dec(x_7); +x_20 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__7___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_19, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_20; +} +} +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__7___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +lean_object* x_18 = _args[17]; +lean_object* x_19 = _args[18]; +_start: +{ +size_t x_20; lean_object* x_21; +x_20 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_21 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__7(x_1, x_2, x_3, x_4, x_20, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +return x_21; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +size_t x_4; size_t x_5; lean_object* x_6; +x_4 = lean_unbox_usize(x_1); +lean_dec(x_1); +x_5 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_6 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__8(x_4, x_5, x_3); +return x_6; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_mkToExprBody___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("match", 5, 5); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_mkToExprBody___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_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__1; +x_2 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__2; +x_3 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__3; +x_4 = l_Lean_Elab_Deriving_ToExpr_mkToExprBody___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_Deriving_ToExpr_mkToExprBody___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("with", 4, 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_mkToExprBody___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("matchAlts", 9, 9); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_mkToExprBody___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_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__1; +x_2 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__2; +x_3 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__3; +x_4 = l_Lean_Elab_Deriving_ToExpr_mkToExprBody___closed__4; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_ToExpr_mkToExprBody(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; uint8_t x_13; +lean_inc(x_2); +lean_inc(x_1); +x_12 = l_Lean_Elab_Deriving_mkDiscrs(x_1, x_2, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_13 = !lean_is_exclusive(x_12); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_14 = lean_ctor_get(x_12, 0); +x_15 = lean_ctor_get(x_12, 1); +lean_inc(x_10); +lean_inc(x_9); +x_16 = l_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_15); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; lean_object* x_21; lean_object* x_22; uint8_t 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); +lean_dec(x_16); +x_19 = lean_ctor_get(x_9, 5); +lean_inc(x_19); +lean_dec(x_9); +x_20 = 0; +x_21 = l_Lean_SourceInfo_fromRef(x_19, x_20); +lean_dec(x_19); +x_22 = lean_st_ref_get(x_10, x_18); +lean_dec(x_10); +x_23 = !lean_is_exclusive(x_22); +if (x_23 == 0) +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; size_t x_29; size_t x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_24 = lean_ctor_get(x_22, 0); +lean_dec(x_24); +x_25 = l_Lean_Elab_Deriving_ToExpr_mkToExprBody___closed__1; +lean_inc(x_21); +lean_ctor_set_tag(x_12, 2); +lean_ctor_set(x_12, 1, x_25); +lean_ctor_set(x_12, 0, x_21); +x_26 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17; +x_27 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__9; +lean_inc(x_21); +x_28 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_28, 0, x_21); +lean_ctor_set(x_28, 1, x_26); +lean_ctor_set(x_28, 2, x_27); +x_29 = lean_array_size(x_14); +x_30 = 0; +x_31 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__20(x_29, x_30, x_14); +x_32 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__7___lambda__1___closed__5; +x_33 = l_Lean_mkSepArray(x_31, x_32); +lean_dec(x_31); +x_34 = l_Array_append___rarg(x_27, x_33); +lean_dec(x_33); +lean_inc(x_21); +x_35 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_35, 0, x_21); +lean_ctor_set(x_35, 1, x_26); +lean_ctor_set(x_35, 2, x_34); +x_36 = l_Lean_Elab_Deriving_ToExpr_mkToExprBody___closed__3; +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_append___rarg(x_27, x_17); +lean_dec(x_17); +lean_inc(x_21); +x_39 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_39, 0, x_21); +lean_ctor_set(x_39, 1, x_26); +lean_ctor_set(x_39, 2, x_38); +x_40 = l_Lean_Elab_Deriving_ToExpr_mkToExprBody___closed__5; +lean_inc(x_21); +x_41 = l_Lean_Syntax_node1(x_21, x_40, x_39); +x_42 = l_Lean_Elab_Deriving_ToExpr_mkToExprBody___closed__2; +lean_inc(x_28); +x_43 = l_Lean_Syntax_node6(x_21, x_42, x_12, x_28, x_28, x_35, x_37, x_41); +lean_ctor_set(x_22, 0, x_43); +return x_22; +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; size_t x_49; size_t x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; +x_44 = lean_ctor_get(x_22, 1); +lean_inc(x_44); +lean_dec(x_22); +x_45 = l_Lean_Elab_Deriving_ToExpr_mkToExprBody___closed__1; +lean_inc(x_21); +lean_ctor_set_tag(x_12, 2); +lean_ctor_set(x_12, 1, x_45); +lean_ctor_set(x_12, 0, x_21); +x_46 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17; +x_47 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__9; +lean_inc(x_21); +x_48 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_48, 0, x_21); +lean_ctor_set(x_48, 1, x_46); +lean_ctor_set(x_48, 2, x_47); +x_49 = lean_array_size(x_14); +x_50 = 0; +x_51 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__20(x_49, x_50, x_14); +x_52 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__7___lambda__1___closed__5; +x_53 = l_Lean_mkSepArray(x_51, x_52); +lean_dec(x_51); +x_54 = l_Array_append___rarg(x_47, x_53); +lean_dec(x_53); +lean_inc(x_21); +x_55 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_55, 0, x_21); +lean_ctor_set(x_55, 1, x_46); +lean_ctor_set(x_55, 2, x_54); +x_56 = l_Lean_Elab_Deriving_ToExpr_mkToExprBody___closed__3; +lean_inc(x_21); +x_57 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_57, 0, x_21); +lean_ctor_set(x_57, 1, x_56); +x_58 = l_Array_append___rarg(x_47, x_17); +lean_dec(x_17); +lean_inc(x_21); +x_59 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_59, 0, x_21); +lean_ctor_set(x_59, 1, x_46); +lean_ctor_set(x_59, 2, x_58); +x_60 = l_Lean_Elab_Deriving_ToExpr_mkToExprBody___closed__5; +lean_inc(x_21); +x_61 = l_Lean_Syntax_node1(x_21, x_60, x_59); +x_62 = l_Lean_Elab_Deriving_ToExpr_mkToExprBody___closed__2; +lean_inc(x_48); +x_63 = l_Lean_Syntax_node6(x_21, x_62, x_12, x_48, x_48, x_55, x_57, x_61); +x_64 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_64, 0, x_63); +lean_ctor_set(x_64, 1, x_44); +return x_64; +} +} +else +{ +uint8_t x_65; +lean_free_object(x_12); +lean_dec(x_14); +lean_dec(x_10); +lean_dec(x_9); +x_65 = !lean_is_exclusive(x_16); +if (x_65 == 0) +{ +return x_16; +} +else +{ +lean_object* x_66; lean_object* x_67; lean_object* x_68; +x_66 = lean_ctor_get(x_16, 0); +x_67 = lean_ctor_get(x_16, 1); +lean_inc(x_67); +lean_inc(x_66); +lean_dec(x_16); +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_69; lean_object* x_70; lean_object* x_71; +x_69 = lean_ctor_get(x_12, 0); +x_70 = lean_ctor_get(x_12, 1); +lean_inc(x_70); +lean_inc(x_69); +lean_dec(x_12); +lean_inc(x_10); +lean_inc(x_9); +x_71 = l_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_70); +if (lean_obj_tag(x_71) == 0) +{ +lean_object* x_72; lean_object* x_73; lean_object* x_74; uint8_t x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; size_t x_85; size_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; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; +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_9, 5); +lean_inc(x_74); +lean_dec(x_9); +x_75 = 0; +x_76 = l_Lean_SourceInfo_fromRef(x_74, x_75); +lean_dec(x_74); +x_77 = lean_st_ref_get(x_10, x_73); +lean_dec(x_10); +x_78 = lean_ctor_get(x_77, 1); +lean_inc(x_78); +if (lean_is_exclusive(x_77)) { + lean_ctor_release(x_77, 0); + lean_ctor_release(x_77, 1); + x_79 = x_77; +} else { + lean_dec_ref(x_77); + x_79 = lean_box(0); +} +x_80 = l_Lean_Elab_Deriving_ToExpr_mkToExprBody___closed__1; +lean_inc(x_76); +x_81 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_81, 0, x_76); +lean_ctor_set(x_81, 1, x_80); +x_82 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17; +x_83 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__9; +lean_inc(x_76); +x_84 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_84, 0, x_76); +lean_ctor_set(x_84, 1, x_82); +lean_ctor_set(x_84, 2, x_83); +x_85 = lean_array_size(x_69); +x_86 = 0; +x_87 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__20(x_85, x_86, x_69); +x_88 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__7___lambda__1___closed__5; +x_89 = l_Lean_mkSepArray(x_87, x_88); +lean_dec(x_87); +x_90 = l_Array_append___rarg(x_83, x_89); +lean_dec(x_89); +lean_inc(x_76); +x_91 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_91, 0, x_76); +lean_ctor_set(x_91, 1, x_82); +lean_ctor_set(x_91, 2, x_90); +x_92 = l_Lean_Elab_Deriving_ToExpr_mkToExprBody___closed__3; +lean_inc(x_76); +x_93 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_93, 0, x_76); +lean_ctor_set(x_93, 1, x_92); +x_94 = l_Array_append___rarg(x_83, x_72); +lean_dec(x_72); +lean_inc(x_76); +x_95 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_95, 0, x_76); +lean_ctor_set(x_95, 1, x_82); +lean_ctor_set(x_95, 2, x_94); +x_96 = l_Lean_Elab_Deriving_ToExpr_mkToExprBody___closed__5; +lean_inc(x_76); +x_97 = l_Lean_Syntax_node1(x_76, x_96, x_95); +x_98 = l_Lean_Elab_Deriving_ToExpr_mkToExprBody___closed__2; +lean_inc(x_84); +x_99 = l_Lean_Syntax_node6(x_76, x_98, x_81, x_84, x_84, x_91, x_93, x_97); +if (lean_is_scalar(x_79)) { + x_100 = lean_alloc_ctor(0, 2, 0); +} else { + x_100 = x_79; +} +lean_ctor_set(x_100, 0, x_99); +lean_ctor_set(x_100, 1, x_78); +return x_100; +} +else +{ +lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; +lean_dec(x_69); +lean_dec(x_10); +lean_dec(x_9); +x_101 = lean_ctor_get(x_71, 0); +lean_inc(x_101); +x_102 = lean_ctor_get(x_71, 1); +lean_inc(x_102); +if (lean_is_exclusive(x_71)) { + lean_ctor_release(x_71, 0); + lean_ctor_release(x_71, 1); + x_103 = x_71; +} else { + lean_dec_ref(x_71); + x_103 = lean_box(0); +} +if (lean_is_scalar(x_103)) { + x_104 = lean_alloc_ctor(1, 2, 0); +} else { + x_104 = x_103; +} +lean_ctor_set(x_104, 0, x_101); +lean_ctor_set(x_104, 1, x_102); +return x_104; +} +} +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("localinst", 9, 9); +return x_1; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___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_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("letDecl", 7, 7); +return x_1; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__1; +x_2 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__2; +x_3 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__3; +x_4 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__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_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("letIdDecl", 9, 9); +return x_1; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__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_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__1; +x_2 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__2; +x_3 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__3; +x_4 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__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_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("typeSpec", 8, 8); +return x_1; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__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_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__1; +x_2 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__2; +x_3 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__3; +x_4 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__7; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__9() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(":", 1, 1); +return x_1; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__4; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__4; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__1; +x_2 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__4; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__12; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__14() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Elab", 4, 4); +return x_1; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__15() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Deriving", 8, 8); +return x_1; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__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_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__1; +x_2 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__14; +x_3 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__15; +x_4 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__4; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__17() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__16; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__18() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__12; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__19() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__18; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__20() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__17; +x_2 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__19; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__21() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__17; +x_2 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__20; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__22() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__17; +x_2 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__21; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__23() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__17; +x_2 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__22; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__24() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__13; +x_2 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__23; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__25() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(":=", 2, 2); +return x_1; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__26() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("structInst", 10, 10); +return x_1; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__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_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__1; +x_2 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__2; +x_3 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__3; +x_4 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__26; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__28() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("{", 1, 1); +return x_1; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__29() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("structInstFields", 16, 16); +return x_1; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__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_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__1; +x_2 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__2; +x_3 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__3; +x_4 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__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_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__31() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("structInstField", 15, 15); +return x_1; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__32() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__1; +x_2 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__2; +x_3 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__3; +x_4 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__31; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__33() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("structInstLVal", 14, 14); +return x_1; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__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_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__1; +x_2 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__2; +x_3 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__3; +x_4 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__33; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__35() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__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; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__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_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__35; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__37() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("structInstFieldDef", 18, 18); +return x_1; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__38() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__1; +x_2 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__2; +x_3 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__3; +x_4 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__37; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__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_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___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_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__40() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__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_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__41() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("optEllipsis", 11, 11); +return x_1; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__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_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__1; +x_2 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__2; +x_3 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__3; +x_4 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__41; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, size_t x_6, size_t x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +_start: +{ +uint8_t x_16; +x_16 = lean_usize_dec_lt(x_7, x_6); +if (x_16 == 0) +{ +lean_object* x_17; +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_1); +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_8); +lean_ctor_set(x_17, 1, x_15); +return x_17; +} +else +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_28; +x_18 = lean_array_uget(x_5, x_7); +x_28 = !lean_is_exclusive(x_8); +if (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_8, 0); +x_30 = lean_ctor_get(x_8, 1); +x_31 = lean_ctor_get(x_29, 0); +lean_inc(x_31); +x_32 = lean_ctor_get(x_29, 1); +lean_inc(x_32); +x_33 = lean_ctor_get(x_29, 2); +lean_inc(x_33); +x_34 = lean_nat_dec_lt(x_32, x_33); +if (x_34 == 0) +{ +lean_object* x_35; +lean_dec(x_33); +lean_dec(x_32); +lean_dec(x_31); +lean_dec(x_18); +x_35 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_35, 0, x_8); +x_19 = x_35; +x_20 = x_15; +goto block_27; +} +else +{ +uint8_t x_36; +x_36 = !lean_is_exclusive(x_29); +if (x_36 == 0) +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_37 = lean_ctor_get(x_29, 2); +lean_dec(x_37); +x_38 = lean_ctor_get(x_29, 1); +lean_dec(x_38); +x_39 = lean_ctor_get(x_29, 0); +lean_dec(x_39); +x_40 = lean_array_fget(x_31, x_32); +x_41 = lean_unsigned_to_nat(1u); +x_42 = lean_nat_add(x_32, x_41); +lean_dec(x_32); +lean_ctor_set(x_29, 1, x_42); +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_18); +x_43 = l_Lean_Elab_Deriving_mkInductArgNames(x_18, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +if (lean_obj_tag(x_43) == 0) +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; uint8_t x_51; +x_44 = lean_ctor_get(x_43, 0); +lean_inc(x_44); +x_45 = lean_ctor_get(x_43, 1); +lean_inc(x_45); +lean_dec(x_43); +x_46 = lean_ctor_get(x_18, 1); +lean_inc(x_46); +x_47 = lean_array_get_size(x_44); +lean_inc(x_46); +x_48 = l_Array_toSubarray___rarg(x_44, x_46, x_47); +x_49 = l_Array_ofSubarray___rarg(x_48); +lean_dec(x_48); +lean_inc(x_49); +x_50 = l_Lean_Elab_Deriving_mkImplicitBinders(x_49, x_9, x_10, x_11, x_12, x_13, x_14, x_45); +x_51 = !lean_is_exclusive(x_50); +if (x_51 == 0) +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; 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_52 = lean_ctor_get(x_50, 0); +x_53 = lean_ctor_get(x_50, 1); +x_54 = lean_unsigned_to_nat(0u); +lean_inc(x_1); +x_55 = l_Array_toSubarray___rarg(x_1, x_54, x_46); +x_56 = l_Array_ofSubarray___rarg(x_55); +lean_dec(x_55); +x_57 = l_Array_append___rarg(x_56, x_49); +lean_dec(x_49); +x_58 = lean_array_get_size(x_57); +x_59 = l_Array_toSubarray___rarg(x_57, x_54, x_58); +x_60 = l_Array_ofSubarray___rarg(x_59); +lean_dec(x_59); +lean_inc(x_18); +x_61 = l_Lean_Elab_Deriving_mkInductiveApp(x_18, x_60, x_9, x_10, x_11, x_12, x_13, x_14, x_53); +x_62 = !lean_is_exclusive(x_61); +if (x_62 == 0) +{ +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_61, 0); +x_64 = lean_ctor_get(x_61, 1); +x_65 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__2; +x_66 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_65, x_13, x_14, x_64); +x_67 = !lean_is_exclusive(x_66); +if (x_67 == 0) +{ +lean_object* x_68; lean_object* x_69; lean_object* x_70; +x_68 = lean_ctor_get(x_66, 0); +x_69 = lean_ctor_get(x_66, 1); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_1); +x_70 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr(x_18, x_1, x_9, x_10, x_11, x_12, x_13, x_14, x_69); +if (lean_obj_tag(x_70) == 0) +{ +lean_object* x_71; lean_object* x_72; lean_object* x_73; uint8_t x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; uint8_t x_78; +x_71 = lean_ctor_get(x_70, 0); +lean_inc(x_71); +x_72 = lean_ctor_get(x_70, 1); +lean_inc(x_72); +lean_dec(x_70); +x_73 = lean_ctor_get(x_13, 5); +lean_inc(x_73); +x_74 = 0; +x_75 = l_Lean_SourceInfo_fromRef(x_73, x_74); +lean_dec(x_73); +x_76 = lean_ctor_get(x_13, 10); +lean_inc(x_76); +x_77 = lean_st_ref_get(x_14, x_72); +x_78 = !lean_is_exclusive(x_77); +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_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; +x_79 = lean_ctor_get(x_77, 0); +x_80 = lean_ctor_get(x_77, 1); +x_81 = lean_ctor_get(x_79, 0); +lean_inc(x_81); +lean_dec(x_79); +x_82 = lean_environment_main_module(x_81); +x_83 = lean_mk_syntax_ident(x_68); +x_84 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__9; +x_85 = l_Array_append___rarg(x_84, x_52); +lean_dec(x_52); +x_86 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17; +lean_inc(x_75); +x_87 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_87, 0, x_75); +lean_ctor_set(x_87, 1, x_86); +lean_ctor_set(x_87, 2, x_85); +x_88 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__9; +lean_inc(x_75); +lean_ctor_set_tag(x_77, 2); +lean_ctor_set(x_77, 1, x_88); +lean_ctor_set(x_77, 0, x_75); +x_89 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__11; +lean_inc(x_76); +lean_inc(x_82); +x_90 = l_Lean_addMacroScope(x_82, x_89, x_76); +x_91 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__10; +x_92 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__24; +lean_inc(x_75); +x_93 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_93, 0, x_75); +lean_ctor_set(x_93, 1, x_91); +lean_ctor_set(x_93, 2, x_90); +lean_ctor_set(x_93, 3, x_92); +lean_inc(x_75); +x_94 = l_Lean_Syntax_node1(x_75, x_86, x_63); +x_95 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__5; +lean_inc(x_75); +x_96 = l_Lean_Syntax_node2(x_75, x_95, x_93, x_94); +x_97 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__8; +lean_inc(x_75); +x_98 = l_Lean_Syntax_node2(x_75, x_97, x_77, x_96); +lean_inc(x_75); +x_99 = l_Lean_Syntax_node1(x_75, x_86, x_98); +x_100 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__25; +lean_inc(x_75); +lean_ctor_set_tag(x_66, 2); +lean_ctor_set(x_66, 1, x_100); +lean_ctor_set(x_66, 0, x_75); +x_101 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__28; +lean_inc(x_75); +lean_ctor_set_tag(x_61, 2); +lean_ctor_set(x_61, 1, x_101); +lean_ctor_set(x_61, 0, x_75); +lean_inc(x_75); +x_102 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_102, 0, x_75); +lean_ctor_set(x_102, 1, x_86); +lean_ctor_set(x_102, 2, x_84); +x_103 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__3; +lean_inc(x_76); +lean_inc(x_82); +x_104 = l_Lean_addMacroScope(x_82, x_103, x_76); +x_105 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__2; +x_106 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__36; +lean_inc(x_75); +x_107 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_107, 0, x_75); +lean_ctor_set(x_107, 1, x_105); +lean_ctor_set(x_107, 2, x_104); +lean_ctor_set(x_107, 3, x_106); +x_108 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__34; +lean_inc(x_102); +lean_inc(x_75); +x_109 = l_Lean_Syntax_node2(x_75, x_108, x_107, x_102); +x_110 = lean_mk_syntax_ident(x_40); +x_111 = l_Array_append___rarg(x_84, x_2); +lean_inc(x_75); +x_112 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_112, 0, x_75); +lean_ctor_set(x_112, 1, x_86); +lean_ctor_set(x_112, 2, x_111); +lean_inc(x_75); +x_113 = l_Lean_Syntax_node2(x_75, x_95, x_110, x_112); +x_114 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__38; +lean_inc(x_66); +lean_inc(x_75); +x_115 = l_Lean_Syntax_node2(x_75, x_114, x_66, x_113); +lean_inc_n(x_102, 2); +lean_inc(x_75); +x_116 = l_Lean_Syntax_node3(x_75, x_86, x_102, x_102, x_115); +x_117 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__32; +lean_inc(x_75); +x_118 = l_Lean_Syntax_node2(x_75, x_117, x_109, x_116); +x_119 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__13; +lean_inc(x_75); +lean_ctor_set_tag(x_50, 2); +lean_ctor_set(x_50, 1, x_119); +lean_ctor_set(x_50, 0, x_75); +x_120 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__10; +x_121 = l_Lean_addMacroScope(x_82, x_120, x_76); +x_122 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__9; +x_123 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__40; +lean_inc(x_75); +x_124 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_124, 0, x_75); +lean_ctor_set(x_124, 1, x_122); +lean_ctor_set(x_124, 2, x_121); +lean_ctor_set(x_124, 3, x_123); +lean_inc(x_102); +lean_inc(x_75); +x_125 = l_Lean_Syntax_node2(x_75, x_108, x_124, x_102); +lean_inc(x_66); +lean_inc(x_75); +x_126 = l_Lean_Syntax_node2(x_75, x_114, x_66, x_71); +lean_inc_n(x_102, 2); +lean_inc(x_75); +x_127 = l_Lean_Syntax_node3(x_75, x_86, x_102, x_102, x_126); +lean_inc(x_75); +x_128 = l_Lean_Syntax_node2(x_75, x_117, x_125, x_127); +lean_inc(x_75); +x_129 = l_Lean_Syntax_node3(x_75, x_86, x_118, x_50, x_128); +x_130 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__30; +lean_inc(x_75); +x_131 = l_Lean_Syntax_node1(x_75, x_130, x_129); +x_132 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__42; +lean_inc(x_102); +lean_inc(x_75); +x_133 = l_Lean_Syntax_node1(x_75, x_132, x_102); +x_134 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__14; +lean_inc(x_75); +x_135 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_135, 0, x_75); +lean_ctor_set(x_135, 1, x_134); +x_136 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__27; +lean_inc(x_102); +lean_inc(x_75); +x_137 = l_Lean_Syntax_node6(x_75, x_136, x_61, x_102, x_131, x_133, x_102, x_135); +x_138 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__6; +lean_inc(x_75); +x_139 = l_Lean_Syntax_node5(x_75, x_138, x_83, x_87, x_99, x_66, x_137); +x_140 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__4; +x_141 = l_Lean_Syntax_node1(x_75, x_140, x_139); +x_142 = lean_array_push(x_30, x_141); +lean_ctor_set(x_8, 1, x_142); +x_143 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_143, 0, x_8); +x_19 = x_143; +x_20 = x_80; +goto block_27; +} +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; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; +x_144 = lean_ctor_get(x_77, 0); +x_145 = lean_ctor_get(x_77, 1); +lean_inc(x_145); +lean_inc(x_144); +lean_dec(x_77); +x_146 = lean_ctor_get(x_144, 0); +lean_inc(x_146); +lean_dec(x_144); +x_147 = lean_environment_main_module(x_146); +x_148 = lean_mk_syntax_ident(x_68); +x_149 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__9; +x_150 = l_Array_append___rarg(x_149, x_52); +lean_dec(x_52); +x_151 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17; +lean_inc(x_75); +x_152 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_152, 0, x_75); +lean_ctor_set(x_152, 1, x_151); +lean_ctor_set(x_152, 2, x_150); +x_153 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__9; +lean_inc(x_75); +x_154 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_154, 0, x_75); +lean_ctor_set(x_154, 1, x_153); +x_155 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__11; +lean_inc(x_76); +lean_inc(x_147); +x_156 = l_Lean_addMacroScope(x_147, x_155, x_76); +x_157 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__10; +x_158 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__24; +lean_inc(x_75); +x_159 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_159, 0, x_75); +lean_ctor_set(x_159, 1, x_157); +lean_ctor_set(x_159, 2, x_156); +lean_ctor_set(x_159, 3, x_158); +lean_inc(x_75); +x_160 = l_Lean_Syntax_node1(x_75, x_151, x_63); +x_161 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__5; +lean_inc(x_75); +x_162 = l_Lean_Syntax_node2(x_75, x_161, x_159, x_160); +x_163 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__8; +lean_inc(x_75); +x_164 = l_Lean_Syntax_node2(x_75, x_163, x_154, x_162); +lean_inc(x_75); +x_165 = l_Lean_Syntax_node1(x_75, x_151, x_164); +x_166 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__25; +lean_inc(x_75); +lean_ctor_set_tag(x_66, 2); +lean_ctor_set(x_66, 1, x_166); +lean_ctor_set(x_66, 0, x_75); +x_167 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__28; +lean_inc(x_75); +lean_ctor_set_tag(x_61, 2); +lean_ctor_set(x_61, 1, x_167); +lean_ctor_set(x_61, 0, x_75); +lean_inc(x_75); +x_168 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_168, 0, x_75); +lean_ctor_set(x_168, 1, x_151); +lean_ctor_set(x_168, 2, x_149); +x_169 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__3; +lean_inc(x_76); +lean_inc(x_147); +x_170 = l_Lean_addMacroScope(x_147, x_169, x_76); +x_171 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__2; +x_172 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__36; +lean_inc(x_75); +x_173 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_173, 0, x_75); +lean_ctor_set(x_173, 1, x_171); +lean_ctor_set(x_173, 2, x_170); +lean_ctor_set(x_173, 3, x_172); +x_174 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__34; +lean_inc(x_168); +lean_inc(x_75); +x_175 = l_Lean_Syntax_node2(x_75, x_174, x_173, x_168); +x_176 = lean_mk_syntax_ident(x_40); +x_177 = l_Array_append___rarg(x_149, x_2); +lean_inc(x_75); +x_178 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_178, 0, x_75); +lean_ctor_set(x_178, 1, x_151); +lean_ctor_set(x_178, 2, x_177); +lean_inc(x_75); +x_179 = l_Lean_Syntax_node2(x_75, x_161, x_176, x_178); +x_180 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__38; +lean_inc(x_66); +lean_inc(x_75); +x_181 = l_Lean_Syntax_node2(x_75, x_180, x_66, x_179); +lean_inc_n(x_168, 2); +lean_inc(x_75); +x_182 = l_Lean_Syntax_node3(x_75, x_151, x_168, x_168, x_181); +x_183 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__32; +lean_inc(x_75); +x_184 = l_Lean_Syntax_node2(x_75, x_183, x_175, x_182); +x_185 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__13; +lean_inc(x_75); +lean_ctor_set_tag(x_50, 2); +lean_ctor_set(x_50, 1, x_185); +lean_ctor_set(x_50, 0, x_75); +x_186 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__10; +x_187 = l_Lean_addMacroScope(x_147, x_186, x_76); +x_188 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__9; +x_189 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__40; +lean_inc(x_75); +x_190 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_190, 0, x_75); +lean_ctor_set(x_190, 1, x_188); +lean_ctor_set(x_190, 2, x_187); +lean_ctor_set(x_190, 3, x_189); +lean_inc(x_168); +lean_inc(x_75); +x_191 = l_Lean_Syntax_node2(x_75, x_174, x_190, x_168); +lean_inc(x_66); +lean_inc(x_75); +x_192 = l_Lean_Syntax_node2(x_75, x_180, x_66, x_71); +lean_inc_n(x_168, 2); +lean_inc(x_75); +x_193 = l_Lean_Syntax_node3(x_75, x_151, x_168, x_168, x_192); +lean_inc(x_75); +x_194 = l_Lean_Syntax_node2(x_75, x_183, x_191, x_193); +lean_inc(x_75); +x_195 = l_Lean_Syntax_node3(x_75, x_151, x_184, x_50, x_194); +x_196 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__30; +lean_inc(x_75); +x_197 = l_Lean_Syntax_node1(x_75, x_196, x_195); +x_198 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__42; +lean_inc(x_168); +lean_inc(x_75); +x_199 = l_Lean_Syntax_node1(x_75, x_198, x_168); +x_200 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__14; +lean_inc(x_75); +x_201 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_201, 0, x_75); +lean_ctor_set(x_201, 1, x_200); +x_202 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__27; +lean_inc(x_168); +lean_inc(x_75); +x_203 = l_Lean_Syntax_node6(x_75, x_202, x_61, x_168, x_197, x_199, x_168, x_201); +x_204 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__6; +lean_inc(x_75); +x_205 = l_Lean_Syntax_node5(x_75, x_204, x_148, x_152, x_165, x_66, x_203); +x_206 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__4; +x_207 = l_Lean_Syntax_node1(x_75, x_206, x_205); +x_208 = lean_array_push(x_30, x_207); +lean_ctor_set(x_8, 1, x_208); +x_209 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_209, 0, x_8); +x_19 = x_209; +x_20 = x_145; +goto block_27; +} +} +else +{ +uint8_t x_210; +lean_free_object(x_66); +lean_dec(x_68); +lean_free_object(x_61); +lean_dec(x_63); +lean_free_object(x_50); +lean_dec(x_52); +lean_dec(x_29); +lean_dec(x_40); +lean_free_object(x_8); +lean_dec(x_30); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_1); +x_210 = !lean_is_exclusive(x_70); +if (x_210 == 0) +{ +return x_70; +} +else +{ +lean_object* x_211; lean_object* x_212; lean_object* x_213; +x_211 = lean_ctor_get(x_70, 0); +x_212 = lean_ctor_get(x_70, 1); +lean_inc(x_212); +lean_inc(x_211); +lean_dec(x_70); +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; +} +} +} +else +{ +lean_object* x_214; lean_object* x_215; lean_object* x_216; +x_214 = lean_ctor_get(x_66, 0); +x_215 = lean_ctor_get(x_66, 1); +lean_inc(x_215); +lean_inc(x_214); +lean_dec(x_66); +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_1); +x_216 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr(x_18, x_1, x_9, x_10, x_11, x_12, x_13, x_14, x_215); +if (lean_obj_tag(x_216) == 0) +{ +lean_object* x_217; lean_object* x_218; lean_object* x_219; uint8_t x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; +x_217 = lean_ctor_get(x_216, 0); +lean_inc(x_217); +x_218 = lean_ctor_get(x_216, 1); +lean_inc(x_218); +lean_dec(x_216); +x_219 = lean_ctor_get(x_13, 5); +lean_inc(x_219); +x_220 = 0; +x_221 = l_Lean_SourceInfo_fromRef(x_219, x_220); +lean_dec(x_219); +x_222 = lean_ctor_get(x_13, 10); +lean_inc(x_222); +x_223 = lean_st_ref_get(x_14, x_218); +x_224 = lean_ctor_get(x_223, 0); +lean_inc(x_224); +x_225 = lean_ctor_get(x_223, 1); +lean_inc(x_225); +if (lean_is_exclusive(x_223)) { + lean_ctor_release(x_223, 0); + lean_ctor_release(x_223, 1); + x_226 = x_223; +} else { + lean_dec_ref(x_223); + x_226 = lean_box(0); +} +x_227 = lean_ctor_get(x_224, 0); +lean_inc(x_227); +lean_dec(x_224); +x_228 = lean_environment_main_module(x_227); +x_229 = lean_mk_syntax_ident(x_214); +x_230 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__9; +x_231 = l_Array_append___rarg(x_230, x_52); +lean_dec(x_52); +x_232 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17; +lean_inc(x_221); +x_233 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_233, 0, x_221); +lean_ctor_set(x_233, 1, x_232); +lean_ctor_set(x_233, 2, x_231); +x_234 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__9; +lean_inc(x_221); +if (lean_is_scalar(x_226)) { + x_235 = lean_alloc_ctor(2, 2, 0); +} else { + x_235 = x_226; + lean_ctor_set_tag(x_235, 2); +} +lean_ctor_set(x_235, 0, x_221); +lean_ctor_set(x_235, 1, x_234); +x_236 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__11; +lean_inc(x_222); +lean_inc(x_228); +x_237 = l_Lean_addMacroScope(x_228, x_236, x_222); +x_238 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__10; +x_239 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__24; +lean_inc(x_221); +x_240 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_240, 0, x_221); +lean_ctor_set(x_240, 1, x_238); +lean_ctor_set(x_240, 2, x_237); +lean_ctor_set(x_240, 3, x_239); +lean_inc(x_221); +x_241 = l_Lean_Syntax_node1(x_221, x_232, x_63); +x_242 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__5; +lean_inc(x_221); +x_243 = l_Lean_Syntax_node2(x_221, x_242, x_240, x_241); +x_244 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__8; +lean_inc(x_221); +x_245 = l_Lean_Syntax_node2(x_221, x_244, x_235, x_243); +lean_inc(x_221); +x_246 = l_Lean_Syntax_node1(x_221, x_232, x_245); +x_247 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__25; +lean_inc(x_221); +x_248 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_248, 0, x_221); +lean_ctor_set(x_248, 1, x_247); +x_249 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__28; +lean_inc(x_221); +lean_ctor_set_tag(x_61, 2); +lean_ctor_set(x_61, 1, x_249); +lean_ctor_set(x_61, 0, x_221); +lean_inc(x_221); +x_250 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_250, 0, x_221); +lean_ctor_set(x_250, 1, x_232); +lean_ctor_set(x_250, 2, x_230); +x_251 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__3; +lean_inc(x_222); +lean_inc(x_228); +x_252 = l_Lean_addMacroScope(x_228, x_251, x_222); +x_253 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__2; +x_254 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__36; +lean_inc(x_221); +x_255 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_255, 0, x_221); +lean_ctor_set(x_255, 1, x_253); +lean_ctor_set(x_255, 2, x_252); +lean_ctor_set(x_255, 3, x_254); +x_256 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__34; +lean_inc(x_250); +lean_inc(x_221); +x_257 = l_Lean_Syntax_node2(x_221, x_256, x_255, x_250); +x_258 = lean_mk_syntax_ident(x_40); +x_259 = l_Array_append___rarg(x_230, x_2); +lean_inc(x_221); +x_260 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_260, 0, x_221); +lean_ctor_set(x_260, 1, x_232); +lean_ctor_set(x_260, 2, x_259); +lean_inc(x_221); +x_261 = l_Lean_Syntax_node2(x_221, x_242, x_258, x_260); +x_262 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__38; +lean_inc(x_248); +lean_inc(x_221); +x_263 = l_Lean_Syntax_node2(x_221, x_262, x_248, x_261); +lean_inc_n(x_250, 2); +lean_inc(x_221); +x_264 = l_Lean_Syntax_node3(x_221, x_232, x_250, x_250, x_263); +x_265 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__32; +lean_inc(x_221); +x_266 = l_Lean_Syntax_node2(x_221, x_265, x_257, x_264); +x_267 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__13; +lean_inc(x_221); +lean_ctor_set_tag(x_50, 2); +lean_ctor_set(x_50, 1, x_267); +lean_ctor_set(x_50, 0, x_221); +x_268 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__10; +x_269 = l_Lean_addMacroScope(x_228, x_268, x_222); +x_270 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__9; +x_271 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__40; +lean_inc(x_221); +x_272 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_272, 0, x_221); +lean_ctor_set(x_272, 1, x_270); +lean_ctor_set(x_272, 2, x_269); +lean_ctor_set(x_272, 3, x_271); +lean_inc(x_250); +lean_inc(x_221); +x_273 = l_Lean_Syntax_node2(x_221, x_256, x_272, x_250); +lean_inc(x_248); +lean_inc(x_221); +x_274 = l_Lean_Syntax_node2(x_221, x_262, x_248, x_217); +lean_inc_n(x_250, 2); +lean_inc(x_221); +x_275 = l_Lean_Syntax_node3(x_221, x_232, x_250, x_250, x_274); +lean_inc(x_221); +x_276 = l_Lean_Syntax_node2(x_221, x_265, x_273, x_275); +lean_inc(x_221); +x_277 = l_Lean_Syntax_node3(x_221, x_232, x_266, x_50, x_276); +x_278 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__30; +lean_inc(x_221); +x_279 = l_Lean_Syntax_node1(x_221, x_278, x_277); +x_280 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__42; +lean_inc(x_250); +lean_inc(x_221); +x_281 = l_Lean_Syntax_node1(x_221, x_280, x_250); +x_282 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__14; +lean_inc(x_221); +x_283 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_283, 0, x_221); +lean_ctor_set(x_283, 1, x_282); +x_284 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__27; +lean_inc(x_250); +lean_inc(x_221); +x_285 = l_Lean_Syntax_node6(x_221, x_284, x_61, x_250, x_279, x_281, x_250, x_283); +x_286 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__6; +lean_inc(x_221); +x_287 = l_Lean_Syntax_node5(x_221, x_286, x_229, x_233, x_246, x_248, x_285); +x_288 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__4; +x_289 = l_Lean_Syntax_node1(x_221, x_288, x_287); +x_290 = lean_array_push(x_30, x_289); +lean_ctor_set(x_8, 1, x_290); +x_291 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_291, 0, x_8); +x_19 = x_291; +x_20 = x_225; +goto block_27; +} +else +{ +lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; +lean_dec(x_214); +lean_free_object(x_61); +lean_dec(x_63); +lean_free_object(x_50); +lean_dec(x_52); +lean_dec(x_29); +lean_dec(x_40); +lean_free_object(x_8); +lean_dec(x_30); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_1); +x_292 = lean_ctor_get(x_216, 0); +lean_inc(x_292); +x_293 = lean_ctor_get(x_216, 1); +lean_inc(x_293); +if (lean_is_exclusive(x_216)) { + lean_ctor_release(x_216, 0); + lean_ctor_release(x_216, 1); + x_294 = x_216; +} else { + lean_dec_ref(x_216); + x_294 = lean_box(0); +} +if (lean_is_scalar(x_294)) { + x_295 = lean_alloc_ctor(1, 2, 0); +} else { + x_295 = x_294; +} +lean_ctor_set(x_295, 0, x_292); +lean_ctor_set(x_295, 1, x_293); +return x_295; +} +} +} +else +{ +lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; +x_296 = lean_ctor_get(x_61, 0); +x_297 = lean_ctor_get(x_61, 1); +lean_inc(x_297); +lean_inc(x_296); +lean_dec(x_61); +x_298 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__2; +x_299 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_298, x_13, x_14, x_297); +x_300 = lean_ctor_get(x_299, 0); +lean_inc(x_300); +x_301 = lean_ctor_get(x_299, 1); +lean_inc(x_301); +if (lean_is_exclusive(x_299)) { + lean_ctor_release(x_299, 0); + lean_ctor_release(x_299, 1); + x_302 = x_299; +} else { + lean_dec_ref(x_299); + x_302 = lean_box(0); +} +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_1); +x_303 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr(x_18, x_1, x_9, x_10, x_11, x_12, x_13, x_14, x_301); +if (lean_obj_tag(x_303) == 0) +{ +lean_object* x_304; lean_object* x_305; lean_object* x_306; uint8_t x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; lean_object* x_353; lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; lean_object* x_369; lean_object* x_370; lean_object* x_371; lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; lean_object* x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; +x_304 = lean_ctor_get(x_303, 0); +lean_inc(x_304); +x_305 = lean_ctor_get(x_303, 1); +lean_inc(x_305); +lean_dec(x_303); +x_306 = lean_ctor_get(x_13, 5); +lean_inc(x_306); +x_307 = 0; +x_308 = l_Lean_SourceInfo_fromRef(x_306, x_307); +lean_dec(x_306); +x_309 = lean_ctor_get(x_13, 10); +lean_inc(x_309); +x_310 = lean_st_ref_get(x_14, x_305); +x_311 = lean_ctor_get(x_310, 0); +lean_inc(x_311); +x_312 = lean_ctor_get(x_310, 1); +lean_inc(x_312); +if (lean_is_exclusive(x_310)) { + lean_ctor_release(x_310, 0); + lean_ctor_release(x_310, 1); + x_313 = x_310; +} else { + lean_dec_ref(x_310); + x_313 = lean_box(0); +} +x_314 = lean_ctor_get(x_311, 0); +lean_inc(x_314); +lean_dec(x_311); +x_315 = lean_environment_main_module(x_314); +x_316 = lean_mk_syntax_ident(x_300); +x_317 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__9; +x_318 = l_Array_append___rarg(x_317, x_52); +lean_dec(x_52); +x_319 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17; +lean_inc(x_308); +x_320 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_320, 0, x_308); +lean_ctor_set(x_320, 1, x_319); +lean_ctor_set(x_320, 2, x_318); +x_321 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__9; +lean_inc(x_308); +if (lean_is_scalar(x_313)) { + x_322 = lean_alloc_ctor(2, 2, 0); +} else { + x_322 = x_313; + lean_ctor_set_tag(x_322, 2); +} +lean_ctor_set(x_322, 0, x_308); +lean_ctor_set(x_322, 1, x_321); +x_323 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__11; +lean_inc(x_309); +lean_inc(x_315); +x_324 = l_Lean_addMacroScope(x_315, x_323, x_309); +x_325 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__10; +x_326 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__24; +lean_inc(x_308); +x_327 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_327, 0, x_308); +lean_ctor_set(x_327, 1, x_325); +lean_ctor_set(x_327, 2, x_324); +lean_ctor_set(x_327, 3, x_326); +lean_inc(x_308); +x_328 = l_Lean_Syntax_node1(x_308, x_319, x_296); +x_329 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__5; +lean_inc(x_308); +x_330 = l_Lean_Syntax_node2(x_308, x_329, x_327, x_328); +x_331 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__8; +lean_inc(x_308); +x_332 = l_Lean_Syntax_node2(x_308, x_331, x_322, x_330); +lean_inc(x_308); +x_333 = l_Lean_Syntax_node1(x_308, x_319, x_332); +x_334 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__25; +lean_inc(x_308); +if (lean_is_scalar(x_302)) { + x_335 = lean_alloc_ctor(2, 2, 0); +} else { + x_335 = x_302; + lean_ctor_set_tag(x_335, 2); +} +lean_ctor_set(x_335, 0, x_308); +lean_ctor_set(x_335, 1, x_334); +x_336 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__28; +lean_inc(x_308); +x_337 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_337, 0, x_308); +lean_ctor_set(x_337, 1, x_336); +lean_inc(x_308); +x_338 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_338, 0, x_308); +lean_ctor_set(x_338, 1, x_319); +lean_ctor_set(x_338, 2, x_317); +x_339 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__3; +lean_inc(x_309); +lean_inc(x_315); +x_340 = l_Lean_addMacroScope(x_315, x_339, x_309); +x_341 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__2; +x_342 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__36; +lean_inc(x_308); +x_343 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_343, 0, x_308); +lean_ctor_set(x_343, 1, x_341); +lean_ctor_set(x_343, 2, x_340); +lean_ctor_set(x_343, 3, x_342); +x_344 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__34; +lean_inc(x_338); +lean_inc(x_308); +x_345 = l_Lean_Syntax_node2(x_308, x_344, x_343, x_338); +x_346 = lean_mk_syntax_ident(x_40); +x_347 = l_Array_append___rarg(x_317, x_2); +lean_inc(x_308); +x_348 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_348, 0, x_308); +lean_ctor_set(x_348, 1, x_319); +lean_ctor_set(x_348, 2, x_347); +lean_inc(x_308); +x_349 = l_Lean_Syntax_node2(x_308, x_329, x_346, x_348); +x_350 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__38; +lean_inc(x_335); +lean_inc(x_308); +x_351 = l_Lean_Syntax_node2(x_308, x_350, x_335, x_349); +lean_inc_n(x_338, 2); +lean_inc(x_308); +x_352 = l_Lean_Syntax_node3(x_308, x_319, x_338, x_338, x_351); +x_353 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__32; +lean_inc(x_308); +x_354 = l_Lean_Syntax_node2(x_308, x_353, x_345, x_352); +x_355 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__13; +lean_inc(x_308); +lean_ctor_set_tag(x_50, 2); +lean_ctor_set(x_50, 1, x_355); +lean_ctor_set(x_50, 0, x_308); +x_356 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__10; +x_357 = l_Lean_addMacroScope(x_315, x_356, x_309); +x_358 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__9; +x_359 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__40; +lean_inc(x_308); +x_360 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_360, 0, x_308); +lean_ctor_set(x_360, 1, x_358); +lean_ctor_set(x_360, 2, x_357); +lean_ctor_set(x_360, 3, x_359); +lean_inc(x_338); +lean_inc(x_308); +x_361 = l_Lean_Syntax_node2(x_308, x_344, x_360, x_338); +lean_inc(x_335); +lean_inc(x_308); +x_362 = l_Lean_Syntax_node2(x_308, x_350, x_335, x_304); +lean_inc_n(x_338, 2); +lean_inc(x_308); +x_363 = l_Lean_Syntax_node3(x_308, x_319, x_338, x_338, x_362); +lean_inc(x_308); +x_364 = l_Lean_Syntax_node2(x_308, x_353, x_361, x_363); +lean_inc(x_308); +x_365 = l_Lean_Syntax_node3(x_308, x_319, x_354, x_50, x_364); +x_366 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__30; +lean_inc(x_308); +x_367 = l_Lean_Syntax_node1(x_308, x_366, x_365); +x_368 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__42; +lean_inc(x_338); +lean_inc(x_308); +x_369 = l_Lean_Syntax_node1(x_308, x_368, x_338); +x_370 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__14; +lean_inc(x_308); +x_371 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_371, 0, x_308); +lean_ctor_set(x_371, 1, x_370); +x_372 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__27; +lean_inc(x_338); +lean_inc(x_308); +x_373 = l_Lean_Syntax_node6(x_308, x_372, x_337, x_338, x_367, x_369, x_338, x_371); +x_374 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__6; +lean_inc(x_308); +x_375 = l_Lean_Syntax_node5(x_308, x_374, x_316, x_320, x_333, x_335, x_373); +x_376 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__4; +x_377 = l_Lean_Syntax_node1(x_308, x_376, x_375); +x_378 = lean_array_push(x_30, x_377); +lean_ctor_set(x_8, 1, x_378); +x_379 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_379, 0, x_8); +x_19 = x_379; +x_20 = x_312; +goto block_27; +} +else +{ +lean_object* x_380; lean_object* x_381; lean_object* x_382; lean_object* x_383; +lean_dec(x_302); +lean_dec(x_300); +lean_dec(x_296); +lean_free_object(x_50); +lean_dec(x_52); +lean_dec(x_29); +lean_dec(x_40); +lean_free_object(x_8); +lean_dec(x_30); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_1); +x_380 = lean_ctor_get(x_303, 0); +lean_inc(x_380); +x_381 = lean_ctor_get(x_303, 1); +lean_inc(x_381); +if (lean_is_exclusive(x_303)) { + lean_ctor_release(x_303, 0); + lean_ctor_release(x_303, 1); + x_382 = x_303; +} else { + lean_dec_ref(x_303); + x_382 = lean_box(0); +} +if (lean_is_scalar(x_382)) { + x_383 = lean_alloc_ctor(1, 2, 0); +} else { + x_383 = x_382; +} +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_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; +x_384 = lean_ctor_get(x_50, 0); +x_385 = lean_ctor_get(x_50, 1); +lean_inc(x_385); +lean_inc(x_384); +lean_dec(x_50); +x_386 = lean_unsigned_to_nat(0u); +lean_inc(x_1); +x_387 = l_Array_toSubarray___rarg(x_1, x_386, x_46); +x_388 = l_Array_ofSubarray___rarg(x_387); +lean_dec(x_387); +x_389 = l_Array_append___rarg(x_388, x_49); +lean_dec(x_49); +x_390 = lean_array_get_size(x_389); +x_391 = l_Array_toSubarray___rarg(x_389, x_386, x_390); +x_392 = l_Array_ofSubarray___rarg(x_391); +lean_dec(x_391); +lean_inc(x_18); +x_393 = l_Lean_Elab_Deriving_mkInductiveApp(x_18, x_392, x_9, x_10, x_11, x_12, x_13, x_14, x_385); +x_394 = lean_ctor_get(x_393, 0); +lean_inc(x_394); +x_395 = lean_ctor_get(x_393, 1); +lean_inc(x_395); +if (lean_is_exclusive(x_393)) { + lean_ctor_release(x_393, 0); + lean_ctor_release(x_393, 1); + x_396 = x_393; +} else { + lean_dec_ref(x_393); + x_396 = lean_box(0); +} +x_397 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__2; +x_398 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_397, x_13, x_14, x_395); +x_399 = lean_ctor_get(x_398, 0); +lean_inc(x_399); +x_400 = lean_ctor_get(x_398, 1); +lean_inc(x_400); +if (lean_is_exclusive(x_398)) { + lean_ctor_release(x_398, 0); + lean_ctor_release(x_398, 1); + x_401 = x_398; +} else { + lean_dec_ref(x_398); + x_401 = lean_box(0); +} +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_1); +x_402 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr(x_18, x_1, x_9, x_10, x_11, x_12, x_13, x_14, x_400); +if (lean_obj_tag(x_402) == 0) +{ +lean_object* x_403; lean_object* x_404; lean_object* x_405; uint8_t x_406; lean_object* x_407; lean_object* x_408; lean_object* x_409; lean_object* x_410; lean_object* x_411; lean_object* x_412; lean_object* x_413; lean_object* x_414; lean_object* x_415; lean_object* x_416; lean_object* x_417; lean_object* x_418; lean_object* x_419; lean_object* x_420; lean_object* x_421; lean_object* x_422; lean_object* x_423; lean_object* x_424; lean_object* x_425; lean_object* x_426; lean_object* x_427; lean_object* x_428; lean_object* x_429; lean_object* x_430; lean_object* x_431; lean_object* x_432; lean_object* x_433; lean_object* x_434; lean_object* x_435; lean_object* x_436; lean_object* x_437; lean_object* x_438; lean_object* x_439; lean_object* x_440; lean_object* x_441; lean_object* x_442; lean_object* x_443; lean_object* x_444; lean_object* x_445; lean_object* x_446; lean_object* x_447; lean_object* x_448; lean_object* x_449; lean_object* x_450; lean_object* x_451; lean_object* x_452; lean_object* x_453; lean_object* x_454; lean_object* x_455; lean_object* x_456; lean_object* x_457; lean_object* x_458; lean_object* x_459; lean_object* x_460; lean_object* x_461; lean_object* x_462; lean_object* x_463; lean_object* x_464; lean_object* x_465; lean_object* x_466; lean_object* x_467; lean_object* x_468; lean_object* x_469; lean_object* x_470; lean_object* x_471; lean_object* x_472; lean_object* x_473; lean_object* x_474; lean_object* x_475; lean_object* x_476; lean_object* x_477; lean_object* x_478; lean_object* x_479; +x_403 = lean_ctor_get(x_402, 0); +lean_inc(x_403); +x_404 = lean_ctor_get(x_402, 1); +lean_inc(x_404); +lean_dec(x_402); +x_405 = lean_ctor_get(x_13, 5); +lean_inc(x_405); +x_406 = 0; +x_407 = l_Lean_SourceInfo_fromRef(x_405, x_406); +lean_dec(x_405); +x_408 = lean_ctor_get(x_13, 10); +lean_inc(x_408); +x_409 = lean_st_ref_get(x_14, x_404); +x_410 = lean_ctor_get(x_409, 0); +lean_inc(x_410); +x_411 = lean_ctor_get(x_409, 1); +lean_inc(x_411); +if (lean_is_exclusive(x_409)) { + lean_ctor_release(x_409, 0); + lean_ctor_release(x_409, 1); + x_412 = x_409; +} else { + lean_dec_ref(x_409); + x_412 = lean_box(0); +} +x_413 = lean_ctor_get(x_410, 0); +lean_inc(x_413); +lean_dec(x_410); +x_414 = lean_environment_main_module(x_413); +x_415 = lean_mk_syntax_ident(x_399); +x_416 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__9; +x_417 = l_Array_append___rarg(x_416, x_384); +lean_dec(x_384); +x_418 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17; +lean_inc(x_407); +x_419 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_419, 0, x_407); +lean_ctor_set(x_419, 1, x_418); +lean_ctor_set(x_419, 2, x_417); +x_420 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__9; +lean_inc(x_407); +if (lean_is_scalar(x_412)) { + x_421 = lean_alloc_ctor(2, 2, 0); +} else { + x_421 = x_412; + lean_ctor_set_tag(x_421, 2); +} +lean_ctor_set(x_421, 0, x_407); +lean_ctor_set(x_421, 1, x_420); +x_422 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__11; +lean_inc(x_408); +lean_inc(x_414); +x_423 = l_Lean_addMacroScope(x_414, x_422, x_408); +x_424 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__10; +x_425 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__24; +lean_inc(x_407); +x_426 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_426, 0, x_407); +lean_ctor_set(x_426, 1, x_424); +lean_ctor_set(x_426, 2, x_423); +lean_ctor_set(x_426, 3, x_425); +lean_inc(x_407); +x_427 = l_Lean_Syntax_node1(x_407, x_418, x_394); +x_428 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__5; +lean_inc(x_407); +x_429 = l_Lean_Syntax_node2(x_407, x_428, x_426, x_427); +x_430 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__8; +lean_inc(x_407); +x_431 = l_Lean_Syntax_node2(x_407, x_430, x_421, x_429); +lean_inc(x_407); +x_432 = l_Lean_Syntax_node1(x_407, x_418, x_431); +x_433 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__25; +lean_inc(x_407); +if (lean_is_scalar(x_401)) { + x_434 = lean_alloc_ctor(2, 2, 0); +} else { + x_434 = x_401; + lean_ctor_set_tag(x_434, 2); +} +lean_ctor_set(x_434, 0, x_407); +lean_ctor_set(x_434, 1, x_433); +x_435 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__28; +lean_inc(x_407); +if (lean_is_scalar(x_396)) { + x_436 = lean_alloc_ctor(2, 2, 0); +} else { + x_436 = x_396; + lean_ctor_set_tag(x_436, 2); +} +lean_ctor_set(x_436, 0, x_407); +lean_ctor_set(x_436, 1, x_435); +lean_inc(x_407); +x_437 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_437, 0, x_407); +lean_ctor_set(x_437, 1, x_418); +lean_ctor_set(x_437, 2, x_416); +x_438 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__3; +lean_inc(x_408); +lean_inc(x_414); +x_439 = l_Lean_addMacroScope(x_414, x_438, x_408); +x_440 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__2; +x_441 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__36; +lean_inc(x_407); +x_442 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_442, 0, x_407); +lean_ctor_set(x_442, 1, x_440); +lean_ctor_set(x_442, 2, x_439); +lean_ctor_set(x_442, 3, x_441); +x_443 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__34; +lean_inc(x_437); +lean_inc(x_407); +x_444 = l_Lean_Syntax_node2(x_407, x_443, x_442, x_437); +x_445 = lean_mk_syntax_ident(x_40); +x_446 = l_Array_append___rarg(x_416, x_2); +lean_inc(x_407); +x_447 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_447, 0, x_407); +lean_ctor_set(x_447, 1, x_418); +lean_ctor_set(x_447, 2, x_446); +lean_inc(x_407); +x_448 = l_Lean_Syntax_node2(x_407, x_428, x_445, x_447); +x_449 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__38; +lean_inc(x_434); +lean_inc(x_407); +x_450 = l_Lean_Syntax_node2(x_407, x_449, x_434, x_448); +lean_inc_n(x_437, 2); +lean_inc(x_407); +x_451 = l_Lean_Syntax_node3(x_407, x_418, x_437, x_437, x_450); +x_452 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__32; +lean_inc(x_407); +x_453 = l_Lean_Syntax_node2(x_407, x_452, x_444, x_451); +x_454 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__13; +lean_inc(x_407); +x_455 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_455, 0, x_407); +lean_ctor_set(x_455, 1, x_454); +x_456 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__10; +x_457 = l_Lean_addMacroScope(x_414, x_456, x_408); +x_458 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__9; +x_459 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__40; +lean_inc(x_407); +x_460 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_460, 0, x_407); +lean_ctor_set(x_460, 1, x_458); +lean_ctor_set(x_460, 2, x_457); +lean_ctor_set(x_460, 3, x_459); +lean_inc(x_437); +lean_inc(x_407); +x_461 = l_Lean_Syntax_node2(x_407, x_443, x_460, x_437); +lean_inc(x_434); +lean_inc(x_407); +x_462 = l_Lean_Syntax_node2(x_407, x_449, x_434, x_403); +lean_inc_n(x_437, 2); +lean_inc(x_407); +x_463 = l_Lean_Syntax_node3(x_407, x_418, x_437, x_437, x_462); +lean_inc(x_407); +x_464 = l_Lean_Syntax_node2(x_407, x_452, x_461, x_463); +lean_inc(x_407); +x_465 = l_Lean_Syntax_node3(x_407, x_418, x_453, x_455, x_464); +x_466 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__30; +lean_inc(x_407); +x_467 = l_Lean_Syntax_node1(x_407, x_466, x_465); +x_468 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__42; +lean_inc(x_437); +lean_inc(x_407); +x_469 = l_Lean_Syntax_node1(x_407, x_468, x_437); +x_470 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__14; +lean_inc(x_407); +x_471 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_471, 0, x_407); +lean_ctor_set(x_471, 1, x_470); +x_472 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__27; +lean_inc(x_437); +lean_inc(x_407); +x_473 = l_Lean_Syntax_node6(x_407, x_472, x_436, x_437, x_467, x_469, x_437, x_471); +x_474 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__6; +lean_inc(x_407); +x_475 = l_Lean_Syntax_node5(x_407, x_474, x_415, x_419, x_432, x_434, x_473); +x_476 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__4; +x_477 = l_Lean_Syntax_node1(x_407, x_476, x_475); +x_478 = lean_array_push(x_30, x_477); +lean_ctor_set(x_8, 1, x_478); +x_479 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_479, 0, x_8); +x_19 = x_479; +x_20 = x_411; +goto block_27; +} +else +{ +lean_object* x_480; lean_object* x_481; lean_object* x_482; lean_object* x_483; +lean_dec(x_401); +lean_dec(x_399); +lean_dec(x_396); +lean_dec(x_394); +lean_dec(x_384); +lean_dec(x_29); +lean_dec(x_40); +lean_free_object(x_8); +lean_dec(x_30); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_1); +x_480 = lean_ctor_get(x_402, 0); +lean_inc(x_480); +x_481 = lean_ctor_get(x_402, 1); +lean_inc(x_481); +if (lean_is_exclusive(x_402)) { + lean_ctor_release(x_402, 0); + lean_ctor_release(x_402, 1); + x_482 = x_402; +} else { + lean_dec_ref(x_402); + x_482 = lean_box(0); +} +if (lean_is_scalar(x_482)) { + x_483 = lean_alloc_ctor(1, 2, 0); +} else { + x_483 = x_482; +} +lean_ctor_set(x_483, 0, x_480); +lean_ctor_set(x_483, 1, x_481); +return x_483; +} +} +} +else +{ +uint8_t x_484; +lean_dec(x_29); +lean_dec(x_40); +lean_free_object(x_8); +lean_dec(x_30); +lean_dec(x_18); +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_1); +x_484 = !lean_is_exclusive(x_43); +if (x_484 == 0) +{ +return x_43; +} +else +{ +lean_object* x_485; lean_object* x_486; lean_object* x_487; +x_485 = lean_ctor_get(x_43, 0); +x_486 = lean_ctor_get(x_43, 1); +lean_inc(x_486); +lean_inc(x_485); +lean_dec(x_43); +x_487 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_487, 0, x_485); +lean_ctor_set(x_487, 1, x_486); +return x_487; +} +} +} +else +{ +lean_object* x_488; lean_object* x_489; lean_object* x_490; lean_object* x_491; lean_object* x_492; +lean_dec(x_29); +x_488 = lean_array_fget(x_31, x_32); +x_489 = lean_unsigned_to_nat(1u); +x_490 = lean_nat_add(x_32, x_489); +lean_dec(x_32); +x_491 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_491, 0, x_31); +lean_ctor_set(x_491, 1, x_490); +lean_ctor_set(x_491, 2, x_33); +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_18); +x_492 = l_Lean_Elab_Deriving_mkInductArgNames(x_18, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +if (lean_obj_tag(x_492) == 0) +{ +lean_object* x_493; lean_object* x_494; lean_object* x_495; lean_object* x_496; lean_object* x_497; lean_object* x_498; lean_object* x_499; lean_object* x_500; lean_object* x_501; lean_object* x_502; lean_object* x_503; lean_object* x_504; lean_object* x_505; lean_object* x_506; lean_object* x_507; lean_object* x_508; lean_object* x_509; lean_object* x_510; lean_object* x_511; lean_object* x_512; lean_object* x_513; lean_object* x_514; lean_object* x_515; lean_object* x_516; lean_object* x_517; lean_object* x_518; lean_object* x_519; +x_493 = lean_ctor_get(x_492, 0); +lean_inc(x_493); +x_494 = lean_ctor_get(x_492, 1); +lean_inc(x_494); +lean_dec(x_492); +x_495 = lean_ctor_get(x_18, 1); +lean_inc(x_495); +x_496 = lean_array_get_size(x_493); +lean_inc(x_495); +x_497 = l_Array_toSubarray___rarg(x_493, x_495, x_496); +x_498 = l_Array_ofSubarray___rarg(x_497); +lean_dec(x_497); +lean_inc(x_498); +x_499 = l_Lean_Elab_Deriving_mkImplicitBinders(x_498, x_9, x_10, x_11, x_12, x_13, x_14, x_494); +x_500 = lean_ctor_get(x_499, 0); +lean_inc(x_500); +x_501 = lean_ctor_get(x_499, 1); +lean_inc(x_501); +if (lean_is_exclusive(x_499)) { + lean_ctor_release(x_499, 0); + lean_ctor_release(x_499, 1); + x_502 = x_499; +} else { + lean_dec_ref(x_499); + x_502 = lean_box(0); +} +x_503 = lean_unsigned_to_nat(0u); +lean_inc(x_1); +x_504 = l_Array_toSubarray___rarg(x_1, x_503, x_495); +x_505 = l_Array_ofSubarray___rarg(x_504); +lean_dec(x_504); +x_506 = l_Array_append___rarg(x_505, x_498); +lean_dec(x_498); +x_507 = lean_array_get_size(x_506); +x_508 = l_Array_toSubarray___rarg(x_506, x_503, x_507); +x_509 = l_Array_ofSubarray___rarg(x_508); +lean_dec(x_508); +lean_inc(x_18); +x_510 = l_Lean_Elab_Deriving_mkInductiveApp(x_18, x_509, x_9, x_10, x_11, x_12, x_13, x_14, x_501); +x_511 = lean_ctor_get(x_510, 0); +lean_inc(x_511); +x_512 = lean_ctor_get(x_510, 1); +lean_inc(x_512); +if (lean_is_exclusive(x_510)) { + lean_ctor_release(x_510, 0); + lean_ctor_release(x_510, 1); + x_513 = x_510; +} else { + lean_dec_ref(x_510); + x_513 = lean_box(0); +} +x_514 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__2; +x_515 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_514, x_13, x_14, x_512); +x_516 = lean_ctor_get(x_515, 0); +lean_inc(x_516); +x_517 = lean_ctor_get(x_515, 1); +lean_inc(x_517); +if (lean_is_exclusive(x_515)) { + lean_ctor_release(x_515, 0); + lean_ctor_release(x_515, 1); + x_518 = x_515; +} else { + lean_dec_ref(x_515); + x_518 = lean_box(0); +} +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_1); +x_519 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr(x_18, x_1, x_9, x_10, x_11, x_12, x_13, x_14, x_517); +if (lean_obj_tag(x_519) == 0) +{ +lean_object* x_520; lean_object* x_521; lean_object* x_522; uint8_t x_523; lean_object* x_524; lean_object* x_525; lean_object* x_526; lean_object* x_527; lean_object* x_528; lean_object* x_529; lean_object* x_530; lean_object* x_531; lean_object* x_532; lean_object* x_533; lean_object* x_534; lean_object* x_535; lean_object* x_536; lean_object* x_537; lean_object* x_538; lean_object* x_539; lean_object* x_540; lean_object* x_541; lean_object* x_542; lean_object* x_543; lean_object* x_544; lean_object* x_545; lean_object* x_546; lean_object* x_547; lean_object* x_548; lean_object* x_549; lean_object* x_550; lean_object* x_551; lean_object* x_552; lean_object* x_553; lean_object* x_554; lean_object* x_555; lean_object* x_556; lean_object* x_557; lean_object* x_558; lean_object* x_559; lean_object* x_560; lean_object* x_561; lean_object* x_562; lean_object* x_563; lean_object* x_564; lean_object* x_565; lean_object* x_566; lean_object* x_567; lean_object* x_568; lean_object* x_569; lean_object* x_570; lean_object* x_571; lean_object* x_572; lean_object* x_573; lean_object* x_574; lean_object* x_575; lean_object* x_576; lean_object* x_577; lean_object* x_578; lean_object* x_579; lean_object* x_580; lean_object* x_581; lean_object* x_582; lean_object* x_583; lean_object* x_584; lean_object* x_585; lean_object* x_586; lean_object* x_587; lean_object* x_588; lean_object* x_589; lean_object* x_590; lean_object* x_591; lean_object* x_592; lean_object* x_593; lean_object* x_594; lean_object* x_595; lean_object* x_596; +x_520 = lean_ctor_get(x_519, 0); +lean_inc(x_520); +x_521 = lean_ctor_get(x_519, 1); +lean_inc(x_521); +lean_dec(x_519); +x_522 = lean_ctor_get(x_13, 5); +lean_inc(x_522); +x_523 = 0; +x_524 = l_Lean_SourceInfo_fromRef(x_522, x_523); +lean_dec(x_522); +x_525 = lean_ctor_get(x_13, 10); +lean_inc(x_525); +x_526 = lean_st_ref_get(x_14, x_521); +x_527 = lean_ctor_get(x_526, 0); +lean_inc(x_527); +x_528 = lean_ctor_get(x_526, 1); +lean_inc(x_528); +if (lean_is_exclusive(x_526)) { + lean_ctor_release(x_526, 0); + lean_ctor_release(x_526, 1); + x_529 = x_526; +} else { + lean_dec_ref(x_526); + x_529 = lean_box(0); +} +x_530 = lean_ctor_get(x_527, 0); +lean_inc(x_530); +lean_dec(x_527); +x_531 = lean_environment_main_module(x_530); +x_532 = lean_mk_syntax_ident(x_516); +x_533 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__9; +x_534 = l_Array_append___rarg(x_533, x_500); +lean_dec(x_500); +x_535 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17; +lean_inc(x_524); +x_536 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_536, 0, x_524); +lean_ctor_set(x_536, 1, x_535); +lean_ctor_set(x_536, 2, x_534); +x_537 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__9; +lean_inc(x_524); +if (lean_is_scalar(x_529)) { + x_538 = lean_alloc_ctor(2, 2, 0); +} else { + x_538 = x_529; + lean_ctor_set_tag(x_538, 2); +} +lean_ctor_set(x_538, 0, x_524); +lean_ctor_set(x_538, 1, x_537); +x_539 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__11; +lean_inc(x_525); +lean_inc(x_531); +x_540 = l_Lean_addMacroScope(x_531, x_539, x_525); +x_541 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__10; +x_542 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__24; +lean_inc(x_524); +x_543 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_543, 0, x_524); +lean_ctor_set(x_543, 1, x_541); +lean_ctor_set(x_543, 2, x_540); +lean_ctor_set(x_543, 3, x_542); +lean_inc(x_524); +x_544 = l_Lean_Syntax_node1(x_524, x_535, x_511); +x_545 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__5; +lean_inc(x_524); +x_546 = l_Lean_Syntax_node2(x_524, x_545, x_543, x_544); +x_547 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__8; +lean_inc(x_524); +x_548 = l_Lean_Syntax_node2(x_524, x_547, x_538, x_546); +lean_inc(x_524); +x_549 = l_Lean_Syntax_node1(x_524, x_535, x_548); +x_550 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__25; +lean_inc(x_524); +if (lean_is_scalar(x_518)) { + x_551 = lean_alloc_ctor(2, 2, 0); +} else { + x_551 = x_518; + lean_ctor_set_tag(x_551, 2); +} +lean_ctor_set(x_551, 0, x_524); +lean_ctor_set(x_551, 1, x_550); +x_552 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__28; +lean_inc(x_524); +if (lean_is_scalar(x_513)) { + x_553 = lean_alloc_ctor(2, 2, 0); +} else { + x_553 = x_513; + lean_ctor_set_tag(x_553, 2); +} +lean_ctor_set(x_553, 0, x_524); +lean_ctor_set(x_553, 1, x_552); +lean_inc(x_524); +x_554 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_554, 0, x_524); +lean_ctor_set(x_554, 1, x_535); +lean_ctor_set(x_554, 2, x_533); +x_555 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__3; +lean_inc(x_525); +lean_inc(x_531); +x_556 = l_Lean_addMacroScope(x_531, x_555, x_525); +x_557 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__2; +x_558 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__36; +lean_inc(x_524); +x_559 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_559, 0, x_524); +lean_ctor_set(x_559, 1, x_557); +lean_ctor_set(x_559, 2, x_556); +lean_ctor_set(x_559, 3, x_558); +x_560 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__34; +lean_inc(x_554); +lean_inc(x_524); +x_561 = l_Lean_Syntax_node2(x_524, x_560, x_559, x_554); +x_562 = lean_mk_syntax_ident(x_488); +x_563 = l_Array_append___rarg(x_533, x_2); +lean_inc(x_524); +x_564 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_564, 0, x_524); +lean_ctor_set(x_564, 1, x_535); +lean_ctor_set(x_564, 2, x_563); +lean_inc(x_524); +x_565 = l_Lean_Syntax_node2(x_524, x_545, x_562, x_564); +x_566 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__38; +lean_inc(x_551); +lean_inc(x_524); +x_567 = l_Lean_Syntax_node2(x_524, x_566, x_551, x_565); +lean_inc_n(x_554, 2); +lean_inc(x_524); +x_568 = l_Lean_Syntax_node3(x_524, x_535, x_554, x_554, x_567); +x_569 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__32; +lean_inc(x_524); +x_570 = l_Lean_Syntax_node2(x_524, x_569, x_561, x_568); +x_571 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__13; +lean_inc(x_524); +if (lean_is_scalar(x_502)) { + x_572 = lean_alloc_ctor(2, 2, 0); +} else { + x_572 = x_502; + lean_ctor_set_tag(x_572, 2); +} +lean_ctor_set(x_572, 0, x_524); +lean_ctor_set(x_572, 1, x_571); +x_573 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__10; +x_574 = l_Lean_addMacroScope(x_531, x_573, x_525); +x_575 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__9; +x_576 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__40; +lean_inc(x_524); +x_577 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_577, 0, x_524); +lean_ctor_set(x_577, 1, x_575); +lean_ctor_set(x_577, 2, x_574); +lean_ctor_set(x_577, 3, x_576); +lean_inc(x_554); +lean_inc(x_524); +x_578 = l_Lean_Syntax_node2(x_524, x_560, x_577, x_554); +lean_inc(x_551); +lean_inc(x_524); +x_579 = l_Lean_Syntax_node2(x_524, x_566, x_551, x_520); +lean_inc_n(x_554, 2); +lean_inc(x_524); +x_580 = l_Lean_Syntax_node3(x_524, x_535, x_554, x_554, x_579); +lean_inc(x_524); +x_581 = l_Lean_Syntax_node2(x_524, x_569, x_578, x_580); +lean_inc(x_524); +x_582 = l_Lean_Syntax_node3(x_524, x_535, x_570, x_572, x_581); +x_583 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__30; +lean_inc(x_524); +x_584 = l_Lean_Syntax_node1(x_524, x_583, x_582); +x_585 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__42; +lean_inc(x_554); +lean_inc(x_524); +x_586 = l_Lean_Syntax_node1(x_524, x_585, x_554); +x_587 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__14; +lean_inc(x_524); +x_588 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_588, 0, x_524); +lean_ctor_set(x_588, 1, x_587); +x_589 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__27; +lean_inc(x_554); +lean_inc(x_524); +x_590 = l_Lean_Syntax_node6(x_524, x_589, x_553, x_554, x_584, x_586, x_554, x_588); +x_591 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__6; +lean_inc(x_524); +x_592 = l_Lean_Syntax_node5(x_524, x_591, x_532, x_536, x_549, x_551, x_590); +x_593 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__4; +x_594 = l_Lean_Syntax_node1(x_524, x_593, x_592); +x_595 = lean_array_push(x_30, x_594); +lean_ctor_set(x_8, 1, x_595); +lean_ctor_set(x_8, 0, x_491); +x_596 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_596, 0, x_8); +x_19 = x_596; +x_20 = x_528; +goto block_27; +} +else +{ +lean_object* x_597; lean_object* x_598; lean_object* x_599; lean_object* x_600; +lean_dec(x_518); +lean_dec(x_516); +lean_dec(x_513); +lean_dec(x_511); +lean_dec(x_502); +lean_dec(x_500); +lean_dec(x_491); +lean_dec(x_488); +lean_free_object(x_8); +lean_dec(x_30); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_1); +x_597 = lean_ctor_get(x_519, 0); +lean_inc(x_597); +x_598 = lean_ctor_get(x_519, 1); +lean_inc(x_598); +if (lean_is_exclusive(x_519)) { + lean_ctor_release(x_519, 0); + lean_ctor_release(x_519, 1); + x_599 = x_519; +} else { + lean_dec_ref(x_519); + x_599 = lean_box(0); +} +if (lean_is_scalar(x_599)) { + x_600 = lean_alloc_ctor(1, 2, 0); +} else { + x_600 = x_599; +} +lean_ctor_set(x_600, 0, x_597); +lean_ctor_set(x_600, 1, x_598); +return x_600; +} +} +else +{ +lean_object* x_601; lean_object* x_602; lean_object* x_603; lean_object* x_604; +lean_dec(x_491); +lean_dec(x_488); +lean_free_object(x_8); +lean_dec(x_30); +lean_dec(x_18); +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_1); +x_601 = lean_ctor_get(x_492, 0); +lean_inc(x_601); +x_602 = lean_ctor_get(x_492, 1); +lean_inc(x_602); +if (lean_is_exclusive(x_492)) { + lean_ctor_release(x_492, 0); + lean_ctor_release(x_492, 1); + x_603 = x_492; +} else { + lean_dec_ref(x_492); + x_603 = lean_box(0); +} +if (lean_is_scalar(x_603)) { + x_604 = lean_alloc_ctor(1, 2, 0); +} else { + x_604 = x_603; +} +lean_ctor_set(x_604, 0, x_601); +lean_ctor_set(x_604, 1, x_602); +return x_604; +} +} +} +} +else +{ +lean_object* x_605; lean_object* x_606; lean_object* x_607; lean_object* x_608; lean_object* x_609; uint8_t x_610; +x_605 = lean_ctor_get(x_8, 0); +x_606 = lean_ctor_get(x_8, 1); +lean_inc(x_606); +lean_inc(x_605); +lean_dec(x_8); +x_607 = lean_ctor_get(x_605, 0); +lean_inc(x_607); +x_608 = lean_ctor_get(x_605, 1); +lean_inc(x_608); +x_609 = lean_ctor_get(x_605, 2); +lean_inc(x_609); +x_610 = lean_nat_dec_lt(x_608, x_609); +if (x_610 == 0) +{ +lean_object* x_611; lean_object* x_612; +lean_dec(x_609); +lean_dec(x_608); +lean_dec(x_607); +lean_dec(x_18); +x_611 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_611, 0, x_605); +lean_ctor_set(x_611, 1, x_606); +x_612 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_612, 0, x_611); +x_19 = x_612; +x_20 = x_15; +goto block_27; +} +else +{ +lean_object* x_613; lean_object* x_614; lean_object* x_615; lean_object* x_616; lean_object* x_617; lean_object* x_618; +if (lean_is_exclusive(x_605)) { + lean_ctor_release(x_605, 0); + lean_ctor_release(x_605, 1); + lean_ctor_release(x_605, 2); + x_613 = x_605; +} else { + lean_dec_ref(x_605); + x_613 = lean_box(0); +} +x_614 = lean_array_fget(x_607, x_608); +x_615 = lean_unsigned_to_nat(1u); +x_616 = lean_nat_add(x_608, x_615); +lean_dec(x_608); +if (lean_is_scalar(x_613)) { + x_617 = lean_alloc_ctor(0, 3, 0); +} else { + x_617 = x_613; +} +lean_ctor_set(x_617, 0, x_607); +lean_ctor_set(x_617, 1, x_616); +lean_ctor_set(x_617, 2, x_609); +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_18); +x_618 = l_Lean_Elab_Deriving_mkInductArgNames(x_18, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +if (lean_obj_tag(x_618) == 0) +{ +lean_object* x_619; lean_object* x_620; lean_object* x_621; lean_object* x_622; lean_object* x_623; lean_object* x_624; lean_object* x_625; lean_object* x_626; lean_object* x_627; lean_object* x_628; lean_object* x_629; lean_object* x_630; lean_object* x_631; lean_object* x_632; lean_object* x_633; lean_object* x_634; lean_object* x_635; lean_object* x_636; lean_object* x_637; lean_object* x_638; lean_object* x_639; lean_object* x_640; lean_object* x_641; lean_object* x_642; lean_object* x_643; lean_object* x_644; lean_object* x_645; +x_619 = lean_ctor_get(x_618, 0); +lean_inc(x_619); +x_620 = lean_ctor_get(x_618, 1); +lean_inc(x_620); +lean_dec(x_618); +x_621 = lean_ctor_get(x_18, 1); +lean_inc(x_621); +x_622 = lean_array_get_size(x_619); +lean_inc(x_621); +x_623 = l_Array_toSubarray___rarg(x_619, x_621, x_622); +x_624 = l_Array_ofSubarray___rarg(x_623); +lean_dec(x_623); +lean_inc(x_624); +x_625 = l_Lean_Elab_Deriving_mkImplicitBinders(x_624, x_9, x_10, x_11, x_12, x_13, x_14, x_620); +x_626 = lean_ctor_get(x_625, 0); +lean_inc(x_626); +x_627 = lean_ctor_get(x_625, 1); +lean_inc(x_627); +if (lean_is_exclusive(x_625)) { + lean_ctor_release(x_625, 0); + lean_ctor_release(x_625, 1); + x_628 = x_625; +} else { + lean_dec_ref(x_625); + x_628 = lean_box(0); +} +x_629 = lean_unsigned_to_nat(0u); +lean_inc(x_1); +x_630 = l_Array_toSubarray___rarg(x_1, x_629, x_621); +x_631 = l_Array_ofSubarray___rarg(x_630); +lean_dec(x_630); +x_632 = l_Array_append___rarg(x_631, x_624); +lean_dec(x_624); +x_633 = lean_array_get_size(x_632); +x_634 = l_Array_toSubarray___rarg(x_632, x_629, x_633); +x_635 = l_Array_ofSubarray___rarg(x_634); +lean_dec(x_634); +lean_inc(x_18); +x_636 = l_Lean_Elab_Deriving_mkInductiveApp(x_18, x_635, x_9, x_10, x_11, x_12, x_13, x_14, x_627); +x_637 = lean_ctor_get(x_636, 0); +lean_inc(x_637); +x_638 = lean_ctor_get(x_636, 1); +lean_inc(x_638); +if (lean_is_exclusive(x_636)) { + lean_ctor_release(x_636, 0); + lean_ctor_release(x_636, 1); + x_639 = x_636; +} else { + lean_dec_ref(x_636); + x_639 = lean_box(0); +} +x_640 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__2; +x_641 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_640, x_13, x_14, x_638); +x_642 = lean_ctor_get(x_641, 0); +lean_inc(x_642); +x_643 = lean_ctor_get(x_641, 1); +lean_inc(x_643); +if (lean_is_exclusive(x_641)) { + lean_ctor_release(x_641, 0); + lean_ctor_release(x_641, 1); + x_644 = x_641; +} else { + lean_dec_ref(x_641); + x_644 = lean_box(0); +} +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_1); +x_645 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr(x_18, x_1, x_9, x_10, x_11, x_12, x_13, x_14, x_643); +if (lean_obj_tag(x_645) == 0) +{ +lean_object* x_646; lean_object* x_647; lean_object* x_648; uint8_t x_649; lean_object* x_650; lean_object* x_651; lean_object* x_652; lean_object* x_653; lean_object* x_654; lean_object* x_655; lean_object* x_656; lean_object* x_657; lean_object* x_658; lean_object* x_659; lean_object* x_660; lean_object* x_661; lean_object* x_662; lean_object* x_663; lean_object* x_664; lean_object* x_665; lean_object* x_666; lean_object* x_667; lean_object* x_668; lean_object* x_669; lean_object* x_670; lean_object* x_671; lean_object* x_672; lean_object* x_673; lean_object* x_674; lean_object* x_675; lean_object* x_676; lean_object* x_677; lean_object* x_678; lean_object* x_679; lean_object* x_680; lean_object* x_681; lean_object* x_682; lean_object* x_683; lean_object* x_684; lean_object* x_685; lean_object* x_686; lean_object* x_687; lean_object* x_688; lean_object* x_689; lean_object* x_690; lean_object* x_691; lean_object* x_692; lean_object* x_693; lean_object* x_694; lean_object* x_695; lean_object* x_696; lean_object* x_697; lean_object* x_698; lean_object* x_699; lean_object* x_700; lean_object* x_701; lean_object* x_702; lean_object* x_703; lean_object* x_704; lean_object* x_705; lean_object* x_706; lean_object* x_707; lean_object* x_708; lean_object* x_709; lean_object* x_710; lean_object* x_711; lean_object* x_712; lean_object* x_713; lean_object* x_714; lean_object* x_715; lean_object* x_716; lean_object* x_717; lean_object* x_718; lean_object* x_719; lean_object* x_720; lean_object* x_721; lean_object* x_722; lean_object* x_723; +x_646 = lean_ctor_get(x_645, 0); +lean_inc(x_646); +x_647 = lean_ctor_get(x_645, 1); +lean_inc(x_647); +lean_dec(x_645); +x_648 = lean_ctor_get(x_13, 5); +lean_inc(x_648); +x_649 = 0; +x_650 = l_Lean_SourceInfo_fromRef(x_648, x_649); +lean_dec(x_648); +x_651 = lean_ctor_get(x_13, 10); +lean_inc(x_651); +x_652 = lean_st_ref_get(x_14, x_647); +x_653 = lean_ctor_get(x_652, 0); +lean_inc(x_653); +x_654 = lean_ctor_get(x_652, 1); +lean_inc(x_654); +if (lean_is_exclusive(x_652)) { + lean_ctor_release(x_652, 0); + lean_ctor_release(x_652, 1); + x_655 = x_652; +} else { + lean_dec_ref(x_652); + x_655 = lean_box(0); +} +x_656 = lean_ctor_get(x_653, 0); +lean_inc(x_656); +lean_dec(x_653); +x_657 = lean_environment_main_module(x_656); +x_658 = lean_mk_syntax_ident(x_642); +x_659 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__9; +x_660 = l_Array_append___rarg(x_659, x_626); +lean_dec(x_626); +x_661 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17; +lean_inc(x_650); +x_662 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_662, 0, x_650); +lean_ctor_set(x_662, 1, x_661); +lean_ctor_set(x_662, 2, x_660); +x_663 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__9; +lean_inc(x_650); +if (lean_is_scalar(x_655)) { + x_664 = lean_alloc_ctor(2, 2, 0); +} else { + x_664 = x_655; + lean_ctor_set_tag(x_664, 2); +} +lean_ctor_set(x_664, 0, x_650); +lean_ctor_set(x_664, 1, x_663); +x_665 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__11; +lean_inc(x_651); +lean_inc(x_657); +x_666 = l_Lean_addMacroScope(x_657, x_665, x_651); +x_667 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__10; +x_668 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__24; +lean_inc(x_650); +x_669 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_669, 0, x_650); +lean_ctor_set(x_669, 1, x_667); +lean_ctor_set(x_669, 2, x_666); +lean_ctor_set(x_669, 3, x_668); +lean_inc(x_650); +x_670 = l_Lean_Syntax_node1(x_650, x_661, x_637); +x_671 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__5; +lean_inc(x_650); +x_672 = l_Lean_Syntax_node2(x_650, x_671, x_669, x_670); +x_673 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__8; +lean_inc(x_650); +x_674 = l_Lean_Syntax_node2(x_650, x_673, x_664, x_672); +lean_inc(x_650); +x_675 = l_Lean_Syntax_node1(x_650, x_661, x_674); +x_676 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__25; +lean_inc(x_650); +if (lean_is_scalar(x_644)) { + x_677 = lean_alloc_ctor(2, 2, 0); +} else { + x_677 = x_644; + lean_ctor_set_tag(x_677, 2); +} +lean_ctor_set(x_677, 0, x_650); +lean_ctor_set(x_677, 1, x_676); +x_678 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__28; +lean_inc(x_650); +if (lean_is_scalar(x_639)) { + x_679 = lean_alloc_ctor(2, 2, 0); +} else { + x_679 = x_639; + lean_ctor_set_tag(x_679, 2); +} +lean_ctor_set(x_679, 0, x_650); +lean_ctor_set(x_679, 1, x_678); +lean_inc(x_650); +x_680 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_680, 0, x_650); +lean_ctor_set(x_680, 1, x_661); +lean_ctor_set(x_680, 2, x_659); +x_681 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__3; +lean_inc(x_651); +lean_inc(x_657); +x_682 = l_Lean_addMacroScope(x_657, x_681, x_651); +x_683 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__2; +x_684 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__36; +lean_inc(x_650); +x_685 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_685, 0, x_650); +lean_ctor_set(x_685, 1, x_683); +lean_ctor_set(x_685, 2, x_682); +lean_ctor_set(x_685, 3, x_684); +x_686 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__34; +lean_inc(x_680); +lean_inc(x_650); +x_687 = l_Lean_Syntax_node2(x_650, x_686, x_685, x_680); +x_688 = lean_mk_syntax_ident(x_614); +x_689 = l_Array_append___rarg(x_659, x_2); +lean_inc(x_650); +x_690 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_690, 0, x_650); +lean_ctor_set(x_690, 1, x_661); +lean_ctor_set(x_690, 2, x_689); +lean_inc(x_650); +x_691 = l_Lean_Syntax_node2(x_650, x_671, x_688, x_690); +x_692 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__38; +lean_inc(x_677); +lean_inc(x_650); +x_693 = l_Lean_Syntax_node2(x_650, x_692, x_677, x_691); +lean_inc_n(x_680, 2); +lean_inc(x_650); +x_694 = l_Lean_Syntax_node3(x_650, x_661, x_680, x_680, x_693); +x_695 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__32; +lean_inc(x_650); +x_696 = l_Lean_Syntax_node2(x_650, x_695, x_687, x_694); +x_697 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__13; +lean_inc(x_650); +if (lean_is_scalar(x_628)) { + x_698 = lean_alloc_ctor(2, 2, 0); +} else { + x_698 = x_628; + lean_ctor_set_tag(x_698, 2); +} +lean_ctor_set(x_698, 0, x_650); +lean_ctor_set(x_698, 1, x_697); +x_699 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__10; +x_700 = l_Lean_addMacroScope(x_657, x_699, x_651); +x_701 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__9; +x_702 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__40; +lean_inc(x_650); +x_703 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_703, 0, x_650); +lean_ctor_set(x_703, 1, x_701); +lean_ctor_set(x_703, 2, x_700); +lean_ctor_set(x_703, 3, x_702); +lean_inc(x_680); +lean_inc(x_650); +x_704 = l_Lean_Syntax_node2(x_650, x_686, x_703, x_680); +lean_inc(x_677); +lean_inc(x_650); +x_705 = l_Lean_Syntax_node2(x_650, x_692, x_677, x_646); +lean_inc_n(x_680, 2); +lean_inc(x_650); +x_706 = l_Lean_Syntax_node3(x_650, x_661, x_680, x_680, x_705); +lean_inc(x_650); +x_707 = l_Lean_Syntax_node2(x_650, x_695, x_704, x_706); +lean_inc(x_650); +x_708 = l_Lean_Syntax_node3(x_650, x_661, x_696, x_698, x_707); +x_709 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__30; +lean_inc(x_650); +x_710 = l_Lean_Syntax_node1(x_650, x_709, x_708); +x_711 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__42; +lean_inc(x_680); +lean_inc(x_650); +x_712 = l_Lean_Syntax_node1(x_650, x_711, x_680); +x_713 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__14; +lean_inc(x_650); +x_714 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_714, 0, x_650); +lean_ctor_set(x_714, 1, x_713); +x_715 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__27; +lean_inc(x_680); +lean_inc(x_650); +x_716 = l_Lean_Syntax_node6(x_650, x_715, x_679, x_680, x_710, x_712, x_680, x_714); +x_717 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__6; +lean_inc(x_650); +x_718 = l_Lean_Syntax_node5(x_650, x_717, x_658, x_662, x_675, x_677, x_716); +x_719 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__4; +x_720 = l_Lean_Syntax_node1(x_650, x_719, x_718); +x_721 = lean_array_push(x_606, x_720); +x_722 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_722, 0, x_617); +lean_ctor_set(x_722, 1, x_721); +x_723 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_723, 0, x_722); +x_19 = x_723; +x_20 = x_654; +goto block_27; +} +else +{ +lean_object* x_724; lean_object* x_725; lean_object* x_726; lean_object* x_727; +lean_dec(x_644); +lean_dec(x_642); +lean_dec(x_639); +lean_dec(x_637); +lean_dec(x_628); +lean_dec(x_626); +lean_dec(x_617); +lean_dec(x_614); +lean_dec(x_606); +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_1); +x_724 = lean_ctor_get(x_645, 0); +lean_inc(x_724); +x_725 = lean_ctor_get(x_645, 1); +lean_inc(x_725); +if (lean_is_exclusive(x_645)) { + lean_ctor_release(x_645, 0); + lean_ctor_release(x_645, 1); + x_726 = x_645; +} else { + lean_dec_ref(x_645); + x_726 = lean_box(0); +} +if (lean_is_scalar(x_726)) { + x_727 = lean_alloc_ctor(1, 2, 0); +} else { + x_727 = x_726; +} +lean_ctor_set(x_727, 0, x_724); +lean_ctor_set(x_727, 1, x_725); +return x_727; +} +} +else +{ +lean_object* x_728; lean_object* x_729; lean_object* x_730; lean_object* x_731; +lean_dec(x_617); +lean_dec(x_614); +lean_dec(x_606); +lean_dec(x_18); +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_1); +x_728 = lean_ctor_get(x_618, 0); +lean_inc(x_728); +x_729 = lean_ctor_get(x_618, 1); +lean_inc(x_729); +if (lean_is_exclusive(x_618)) { + lean_ctor_release(x_618, 0); + lean_ctor_release(x_618, 1); + x_730 = x_618; +} else { + lean_dec_ref(x_618); + x_730 = lean_box(0); +} +if (lean_is_scalar(x_730)) { + x_731 = lean_alloc_ctor(1, 2, 0); +} else { + x_731 = x_730; +} +lean_ctor_set(x_731, 0, x_728); +lean_ctor_set(x_731, 1, x_729); +return x_731; +} +} +} +block_27: +{ +if (lean_obj_tag(x_19) == 0) +{ +lean_object* x_21; lean_object* x_22; +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_1); +x_21 = lean_ctor_get(x_19, 0); +lean_inc(x_21); +lean_dec(x_19); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_20); +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_7, x_24); +x_7 = x_25; +x_8 = x_23; +x_15 = x_20; +goto _start; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; size_t x_19; size_t x_20; lean_object* x_21; +x_11 = lean_ctor_get(x_1, 1); +lean_inc(x_11); +x_12 = lean_array_get_size(x_11); +x_13 = lean_unsigned_to_nat(0u); +x_14 = l_Array_toSubarray___rarg(x_11, x_13, x_12); +x_15 = lean_ctor_get(x_1, 0); +lean_inc(x_15); +lean_dec(x_1); +x_16 = lean_box(0); +x_17 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__1; +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_14); +lean_ctor_set(x_18, 1, x_17); +x_19 = lean_array_size(x_15); +x_20 = 0; +x_21 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1(x_2, x_3, x_15, x_16, x_15, x_19, x_20, x_18, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_15); +if (lean_obj_tag(x_21) == 0) +{ +uint8_t x_22; +x_22 = !lean_is_exclusive(x_21); +if (x_22 == 0) +{ +lean_object* x_23; lean_object* x_24; +x_23 = lean_ctor_get(x_21, 0); +x_24 = lean_ctor_get(x_23, 1); +lean_inc(x_24); +lean_dec(x_23); +lean_ctor_set(x_21, 0, x_24); +return x_21; +} +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_ctor_get(x_25, 1); +lean_inc(x_27); +lean_dec(x_25); +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_27); +lean_ctor_set(x_28, 1, x_26); +return x_28; +} +} +else +{ +uint8_t x_29; +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_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +_start: +{ +size_t x_16; size_t x_17; lean_object* x_18; +x_16 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_17 = lean_unbox_usize(x_7); +lean_dec(x_7); +x_18 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1(x_1, x_2, x_3, x_4, x_5, x_16, x_17, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_18; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_3); +return x_11; +} +} +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("inst", 4, 4); +return x_1; +} +} +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___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_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("explicitBinder", 14, 14); +return x_1; +} +} +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__1; +x_2 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__2; +x_3 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__3; +x_4 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__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_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("(", 1, 1); +return x_1; +} +} +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__1___closed__5; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__1___closed__5; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__1; +x_2 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__1___closed__5; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__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_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__8; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__10; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__10; +x_2 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__11; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__9; +x_2 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__12; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__14() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(")", 1, 1); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1(size_t x_1, size_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_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_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_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +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 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__2; +x_17 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_16, x_8, x_9, x_10); +x_18 = !lean_is_exclusive(x_17); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; +x_19 = lean_ctor_get(x_17, 0); +x_20 = lean_ctor_get(x_17, 1); +x_21 = lean_mk_syntax_ident(x_19); +x_22 = lean_ctor_get(x_8, 5); +lean_inc(x_22); +x_23 = 0; +x_24 = l_Lean_SourceInfo_fromRef(x_22, x_23); +lean_dec(x_22); +x_25 = lean_ctor_get(x_8, 10); +lean_inc(x_25); +x_26 = lean_st_ref_get(x_9, x_20); +x_27 = !lean_is_exclusive(x_26); +if (x_27 == 0) +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; size_t x_57; size_t x_58; lean_object* x_59; +x_28 = lean_ctor_get(x_26, 0); +x_29 = lean_ctor_get(x_26, 1); +x_30 = lean_ctor_get(x_28, 0); +lean_inc(x_30); +lean_dec(x_28); +x_31 = lean_environment_main_module(x_30); +x_32 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__5; +lean_inc(x_24); +lean_ctor_set_tag(x_26, 2); +lean_ctor_set(x_26, 1, x_32); +lean_ctor_set(x_26, 0, x_24); +x_33 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17; +lean_inc(x_21); +lean_inc(x_24); +x_34 = l_Lean_Syntax_node1(x_24, x_33, x_21); +x_35 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__9; +lean_inc(x_24); +lean_ctor_set_tag(x_17, 2); +lean_ctor_set(x_17, 1, x_35); +lean_ctor_set(x_17, 0, x_24); +x_36 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__7; +x_37 = l_Lean_addMacroScope(x_31, x_36, x_25); +x_38 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__6; +x_39 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__13; +lean_inc(x_24); +x_40 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_40, 0, x_24); +lean_ctor_set(x_40, 1, x_38); +lean_ctor_set(x_40, 2, x_37); +lean_ctor_set(x_40, 3, x_39); +x_41 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__8; +lean_inc(x_24); +x_42 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_42, 0, x_24); +lean_ctor_set(x_42, 1, x_41); +x_43 = lean_mk_syntax_ident(x_13); +lean_inc(x_24); +x_44 = l_Lean_Syntax_node1(x_24, x_33, x_43); +x_45 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__14; +lean_inc(x_24); +x_46 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_46, 0, x_24); +lean_ctor_set(x_46, 1, x_45); +x_47 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__7; +lean_inc(x_24); +x_48 = l_Lean_Syntax_node4(x_24, x_47, x_40, x_42, x_44, x_46); +lean_inc(x_24); +x_49 = l_Lean_Syntax_node2(x_24, x_33, x_17, x_48); +x_50 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__9; +lean_inc(x_24); +x_51 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_51, 0, x_24); +lean_ctor_set(x_51, 1, x_33); +lean_ctor_set(x_51, 2, x_50); +x_52 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__14; +lean_inc(x_24); +x_53 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_53, 0, x_24); +lean_ctor_set(x_53, 1, x_52); +x_54 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__4; +x_55 = l_Lean_Syntax_node5(x_24, x_54, x_26, x_34, x_49, x_51, x_53); +x_56 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_56, 0, x_21); +lean_ctor_set(x_56, 1, x_55); +x_57 = 1; +x_58 = lean_usize_add(x_2, x_57); +x_59 = lean_array_uset(x_15, x_2, x_56); +x_2 = x_58; +x_3 = x_59; +x_10 = x_29; +goto _start; +} +else +{ +lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; size_t x_91; size_t x_92; lean_object* x_93; +x_61 = lean_ctor_get(x_26, 0); +x_62 = lean_ctor_get(x_26, 1); +lean_inc(x_62); +lean_inc(x_61); +lean_dec(x_26); +x_63 = lean_ctor_get(x_61, 0); +lean_inc(x_63); +lean_dec(x_61); +x_64 = lean_environment_main_module(x_63); +x_65 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__5; +lean_inc(x_24); +x_66 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_66, 0, x_24); +lean_ctor_set(x_66, 1, x_65); +x_67 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17; +lean_inc(x_21); +lean_inc(x_24); +x_68 = l_Lean_Syntax_node1(x_24, x_67, x_21); +x_69 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__9; +lean_inc(x_24); +lean_ctor_set_tag(x_17, 2); +lean_ctor_set(x_17, 1, x_69); +lean_ctor_set(x_17, 0, x_24); +x_70 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__7; +x_71 = l_Lean_addMacroScope(x_64, x_70, x_25); +x_72 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__6; +x_73 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__13; +lean_inc(x_24); +x_74 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_74, 0, x_24); +lean_ctor_set(x_74, 1, x_72); +lean_ctor_set(x_74, 2, x_71); +lean_ctor_set(x_74, 3, x_73); +x_75 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__8; +lean_inc(x_24); +x_76 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_76, 0, x_24); +lean_ctor_set(x_76, 1, x_75); +x_77 = lean_mk_syntax_ident(x_13); +lean_inc(x_24); +x_78 = l_Lean_Syntax_node1(x_24, x_67, x_77); +x_79 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__14; +lean_inc(x_24); +x_80 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_80, 0, x_24); +lean_ctor_set(x_80, 1, x_79); +x_81 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__7; +lean_inc(x_24); +x_82 = l_Lean_Syntax_node4(x_24, x_81, x_74, x_76, x_78, x_80); +lean_inc(x_24); +x_83 = l_Lean_Syntax_node2(x_24, x_67, x_17, x_82); +x_84 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__9; +lean_inc(x_24); +x_85 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_85, 0, x_24); +lean_ctor_set(x_85, 1, x_67); +lean_ctor_set(x_85, 2, x_84); +x_86 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__14; +lean_inc(x_24); +x_87 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_87, 0, x_24); +lean_ctor_set(x_87, 1, x_86); +x_88 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__4; +x_89 = l_Lean_Syntax_node5(x_24, x_88, x_66, x_68, x_83, x_85, x_87); +x_90 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_90, 0, x_21); +lean_ctor_set(x_90, 1, x_89); +x_91 = 1; +x_92 = lean_usize_add(x_2, x_91); +x_93 = lean_array_uset(x_15, x_2, x_90); +x_2 = x_92; +x_3 = x_93; +x_10 = x_62; +goto _start; +} +} +else +{ +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; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; size_t x_135; size_t x_136; lean_object* x_137; +x_95 = lean_ctor_get(x_17, 0); +x_96 = lean_ctor_get(x_17, 1); +lean_inc(x_96); +lean_inc(x_95); +lean_dec(x_17); +x_97 = lean_mk_syntax_ident(x_95); +x_98 = lean_ctor_get(x_8, 5); +lean_inc(x_98); +x_99 = 0; +x_100 = l_Lean_SourceInfo_fromRef(x_98, x_99); +lean_dec(x_98); +x_101 = lean_ctor_get(x_8, 10); +lean_inc(x_101); +x_102 = lean_st_ref_get(x_9, x_96); +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_102); + x_105 = lean_box(0); +} +x_106 = lean_ctor_get(x_103, 0); +lean_inc(x_106); +lean_dec(x_103); +x_107 = lean_environment_main_module(x_106); +x_108 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__5; +lean_inc(x_100); +if (lean_is_scalar(x_105)) { + x_109 = lean_alloc_ctor(2, 2, 0); +} else { + x_109 = x_105; + lean_ctor_set_tag(x_109, 2); +} +lean_ctor_set(x_109, 0, x_100); +lean_ctor_set(x_109, 1, x_108); +x_110 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17; +lean_inc(x_97); +lean_inc(x_100); +x_111 = l_Lean_Syntax_node1(x_100, x_110, x_97); +x_112 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__9; +lean_inc(x_100); +x_113 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_113, 0, x_100); +lean_ctor_set(x_113, 1, x_112); +x_114 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__7; +x_115 = l_Lean_addMacroScope(x_107, x_114, x_101); +x_116 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__6; +x_117 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__13; +lean_inc(x_100); +x_118 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_118, 0, x_100); +lean_ctor_set(x_118, 1, x_116); +lean_ctor_set(x_118, 2, x_115); +lean_ctor_set(x_118, 3, x_117); +x_119 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__8; +lean_inc(x_100); +x_120 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_120, 0, x_100); +lean_ctor_set(x_120, 1, x_119); +x_121 = lean_mk_syntax_ident(x_13); +lean_inc(x_100); +x_122 = l_Lean_Syntax_node1(x_100, x_110, x_121); +x_123 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__14; +lean_inc(x_100); +x_124 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_124, 0, x_100); +lean_ctor_set(x_124, 1, x_123); +x_125 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__7; +lean_inc(x_100); +x_126 = l_Lean_Syntax_node4(x_100, x_125, x_118, x_120, x_122, x_124); +lean_inc(x_100); +x_127 = l_Lean_Syntax_node2(x_100, x_110, x_113, x_126); +x_128 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__9; +lean_inc(x_100); +x_129 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_129, 0, x_100); +lean_ctor_set(x_129, 1, x_110); +lean_ctor_set(x_129, 2, x_128); +x_130 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__14; +lean_inc(x_100); +x_131 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_131, 0, x_100); +lean_ctor_set(x_131, 1, x_130); +x_132 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__4; +x_133 = l_Lean_Syntax_node5(x_100, x_132, x_109, x_111, x_127, x_129, x_131); +x_134 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_134, 0, x_97); +lean_ctor_set(x_134, 1, x_133); +x_135 = 1; +x_136 = lean_usize_add(x_2, x_135); +x_137 = lean_array_uset(x_15, x_2, x_134); +x_2 = x_136; +x_3 = x_137; +x_10 = x_104; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__2(size_t x_1, size_t x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; +x_4 = lean_usize_dec_lt(x_2, x_1); +if (x_4 == 0) +{ +return x_3; +} +else +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; size_t x_8; size_t x_9; lean_object* x_10; +x_5 = lean_array_uget(x_3, x_2); +x_6 = lean_unsigned_to_nat(0u); +x_7 = lean_array_uset(x_3, x_2, x_6); +x_8 = 1; +x_9 = lean_usize_add(x_2, x_8); +x_10 = lean_array_uset(x_7, x_2, x_5); +x_2 = x_9; +x_3 = x_10; +goto _start; +} +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__3(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { +_start: +{ +uint8_t x_5; +x_5 = lean_usize_dec_lt(x_3, x_2); +if (x_5 == 0) +{ +return x_4; +} +else +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x_9; size_t x_10; lean_object* x_11; +x_6 = lean_array_uget(x_4, x_3); +x_7 = lean_unsigned_to_nat(0u); +x_8 = lean_array_uset(x_4, x_3, x_7); +x_9 = 1; +x_10 = lean_usize_add(x_3, x_9); +x_11 = lean_array_uset(x_8, x_3, x_6); +x_3 = x_10; +x_4 = x_11; +goto _start; +} +} +} +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___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; uint8_t x_13; +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_is_exclusive(x_12); +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_12, 0); +x_15 = lean_ctor_get(x_12, 1); +x_16 = l_Lean_Elab_addMacroStack___at_Lean_Elab_Term_instAddErrorMessageContextTermElabM___spec__1(x_14, x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_15); +lean_dec(x_2); +x_17 = !lean_is_exclusive(x_16); +if (x_17 == 0) +{ +lean_object* x_18; +x_18 = lean_ctor_get(x_16, 0); +lean_ctor_set(x_12, 1, x_18); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set_tag(x_16, 1); +lean_ctor_set(x_16, 0, x_12); +return x_16; +} +else +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_16, 0); +x_20 = lean_ctor_get(x_16, 1); +lean_inc(x_20); +lean_inc(x_19); +lean_dec(x_16); +lean_ctor_set(x_12, 1, x_19); +lean_ctor_set(x_12, 0, x_11); +x_21 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_21, 0, x_12); +lean_ctor_set(x_21, 1, x_20); +return x_21; +} +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_22 = lean_ctor_get(x_12, 0); +x_23 = lean_ctor_get(x_12, 1); +lean_inc(x_23); +lean_inc(x_22); +lean_dec(x_12); +x_24 = l_Lean_Elab_addMacroStack___at_Lean_Elab_Term_instAddErrorMessageContextTermElabM___spec__1(x_22, x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_23); +lean_dec(x_2); +x_25 = lean_ctor_get(x_24, 0); +lean_inc(x_25); +x_26 = lean_ctor_get(x_24, 1); +lean_inc(x_26); +if (lean_is_exclusive(x_24)) { + lean_ctor_release(x_24, 0); + lean_ctor_release(x_24, 1); + x_27 = x_24; +} else { + lean_dec_ref(x_24); + x_27 = lean_box(0); +} +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_11); +lean_ctor_set(x_28, 1, x_25); +if (lean_is_scalar(x_27)) { + x_29 = lean_alloc_ctor(1, 2, 0); +} else { + x_29 = x_27; + lean_ctor_set_tag(x_29, 1); +} +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_26); +return x_29; +} +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("bracketedBinder", 15, 15); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___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_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__1; +x_2 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__2; +x_3 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__3; +x_4 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__1; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__2; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Command", 7, 7); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("declaration", 11, 11); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__1; +x_2 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__2; +x_3 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__4; +x_4 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__5; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("declModifiers", 13, 13); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__1; +x_2 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__2; +x_3 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__4; +x_4 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__7; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__9() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("private", 7, 7); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__1; +x_2 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__2; +x_3 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__4; +x_4 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__9; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__11() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("definition", 10, 10); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__1; +x_2 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__2; +x_3 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__4; +x_4 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__11; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__13() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("def", 3, 3); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__14() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("declId", 6, 6); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__15() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__1; +x_2 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__2; +x_3 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__4; +x_4 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__14; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__16() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("optDeclSig", 10, 10); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__17() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__1; +x_2 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__2; +x_3 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__4; +x_4 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__16; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__18() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__8; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__19() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__8; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__20() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__1; +x_2 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__8; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__21() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__20; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__22() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__20; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__23() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__22; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__24() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__22; +x_2 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__23; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__25() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__21; +x_2 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__24; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__26() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("declValSimple", 13, 13); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__27() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__1; +x_2 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__2; +x_3 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__4; +x_4 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__26; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__28() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Termination", 11, 11); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__29() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("suffix", 6, 6); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__30() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__1; +x_2 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__2; +x_3 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__28; +x_4 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__29; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__31() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("partial", 7, 7); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__32() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__1; +x_2 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__2; +x_3 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__4; +x_4 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__31; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__33() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("(internal error) expecting inst binder", 38, 38); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__34() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__33; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1(lean_object* x_1, lean_object* x_2, size_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +_start: +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; lean_object* x_22; lean_object* x_23; +x_16 = lean_box(0); +x_17 = lean_ctor_get(x_1, 0); +lean_inc(x_17); +lean_dec(x_1); +x_18 = lean_box(0); +x_19 = l_Array_back_x21___rarg(x_18, x_17); +x_20 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__4; +lean_inc(x_19); +x_21 = l_Lean_Syntax_isOfKind(x_19, x_20); +if (x_21 == 0) +{ +lean_object* x_228; lean_object* x_229; uint8_t x_230; +lean_dec(x_19); +lean_dec(x_17); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +x_228 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__34; +x_229 = l_Lean_throwError___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__4(x_228, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +lean_dec(x_13); +x_230 = !lean_is_exclusive(x_229); +if (x_230 == 0) +{ +return x_229; +} +else +{ +lean_object* x_231; lean_object* x_232; lean_object* x_233; +x_231 = lean_ctor_get(x_229, 0); +x_232 = lean_ctor_get(x_229, 1); +lean_inc(x_232); +lean_inc(x_231); +lean_dec(x_229); +x_233 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_233, 0, x_231); +lean_ctor_set(x_233, 1, x_232); +return x_233; +} +} +else +{ +lean_object* x_234; lean_object* x_235; uint8_t x_236; +x_234 = lean_unsigned_to_nat(1u); +x_235 = l_Lean_Syntax_getArg(x_19, x_234); +lean_inc(x_235); +x_236 = l_Lean_Syntax_matchesNull(x_235, x_234); +if (x_236 == 0) +{ +lean_object* x_237; lean_object* x_238; uint8_t x_239; +lean_dec(x_235); +lean_dec(x_19); +lean_dec(x_17); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +x_237 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__34; +x_238 = l_Lean_throwError___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__4(x_237, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +lean_dec(x_13); +x_239 = !lean_is_exclusive(x_238); +if (x_239 == 0) +{ +return x_238; +} +else +{ +lean_object* x_240; lean_object* x_241; lean_object* x_242; +x_240 = lean_ctor_get(x_238, 0); +x_241 = lean_ctor_get(x_238, 1); +lean_inc(x_241); +lean_inc(x_240); +lean_dec(x_238); +x_242 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_242, 0, x_240); +lean_ctor_set(x_242, 1, x_241); +return x_242; +} +} +else +{ +lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; uint8_t x_247; +x_243 = lean_unsigned_to_nat(0u); +x_244 = l_Lean_Syntax_getArg(x_235, x_243); +lean_dec(x_235); +x_245 = lean_unsigned_to_nat(2u); +x_246 = l_Lean_Syntax_getArg(x_19, x_245); +lean_inc(x_246); +x_247 = l_Lean_Syntax_matchesNull(x_246, x_245); +if (x_247 == 0) +{ +lean_object* x_248; lean_object* x_249; uint8_t x_250; +lean_dec(x_246); +lean_dec(x_244); +lean_dec(x_19); +lean_dec(x_17); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +x_248 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__34; +x_249 = l_Lean_throwError___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__4(x_248, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +lean_dec(x_13); +x_250 = !lean_is_exclusive(x_249); +if (x_250 == 0) +{ +return x_249; +} +else +{ +lean_object* x_251; lean_object* x_252; lean_object* x_253; +x_251 = lean_ctor_get(x_249, 0); +x_252 = lean_ctor_get(x_249, 1); +lean_inc(x_252); +lean_inc(x_251); +lean_dec(x_249); +x_253 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_253, 0, x_251); +lean_ctor_set(x_253, 1, x_252); +return x_253; +} +} +else +{ +lean_object* x_254; lean_object* x_255; lean_object* x_256; uint8_t x_257; +x_254 = l_Lean_Syntax_getArg(x_246, x_234); +lean_dec(x_246); +x_255 = lean_unsigned_to_nat(3u); +x_256 = l_Lean_Syntax_getArg(x_19, x_255); +lean_dec(x_19); +x_257 = l_Lean_Syntax_matchesNull(x_256, x_243); +if (x_257 == 0) +{ +lean_object* x_258; lean_object* x_259; uint8_t x_260; +lean_dec(x_254); +lean_dec(x_244); +lean_dec(x_17); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +x_258 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__34; +x_259 = l_Lean_throwError___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__4(x_258, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +lean_dec(x_13); +x_260 = !lean_is_exclusive(x_259); +if (x_260 == 0) +{ +return x_259; +} +else +{ +lean_object* x_261; lean_object* x_262; lean_object* x_263; +x_261 = lean_ctor_get(x_259, 0); +x_262 = lean_ctor_get(x_259, 1); +lean_inc(x_262); +lean_inc(x_261); +lean_dec(x_259); +x_263 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_263, 0, x_261); +lean_ctor_set(x_263, 1, x_262); +return x_263; +} +} +else +{ +lean_object* x_264; +x_264 = l_Lean_Elab_Deriving_ToExpr_updateIndType(x_6, x_254, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +if (lean_obj_tag(x_264) == 0) +{ +lean_object* x_265; lean_object* x_266; lean_object* x_267; uint8_t x_268; lean_object* x_269; lean_object* x_270; uint8_t x_271; +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 = lean_ctor_get(x_13, 5); +lean_inc(x_267); +x_268 = 0; +x_269 = l_Lean_SourceInfo_fromRef(x_267, x_268); +lean_dec(x_267); +x_270 = lean_st_ref_get(x_14, x_266); +x_271 = !lean_is_exclusive(x_270); +if (x_271 == 0) +{ +lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; +x_272 = lean_ctor_get(x_270, 1); +x_273 = lean_ctor_get(x_270, 0); +lean_dec(x_273); +x_274 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__5; +lean_inc(x_269); +lean_ctor_set_tag(x_270, 2); +lean_ctor_set(x_270, 1, x_274); +lean_ctor_set(x_270, 0, x_269); +x_275 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17; +lean_inc(x_269); +x_276 = l_Lean_Syntax_node1(x_269, x_275, x_244); +x_277 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__9; +lean_inc(x_269); +x_278 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_278, 0, x_269); +lean_ctor_set(x_278, 1, x_277); +lean_inc(x_269); +x_279 = l_Lean_Syntax_node2(x_269, x_275, x_278, x_265); +x_280 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__9; +lean_inc(x_269); +x_281 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_281, 0, x_269); +lean_ctor_set(x_281, 1, x_275); +lean_ctor_set(x_281, 2, x_280); +x_282 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__14; +lean_inc(x_269); +x_283 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_283, 0, x_269); +lean_ctor_set(x_283, 1, x_282); +x_284 = l_Lean_Syntax_node5(x_269, x_20, x_270, x_276, x_279, x_281, x_283); +x_22 = x_284; +x_23 = x_272; +goto block_227; +} +else +{ +lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; +x_285 = lean_ctor_get(x_270, 1); +lean_inc(x_285); +lean_dec(x_270); +x_286 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__5; +lean_inc(x_269); +x_287 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_287, 0, x_269); +lean_ctor_set(x_287, 1, x_286); +x_288 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17; +lean_inc(x_269); +x_289 = l_Lean_Syntax_node1(x_269, x_288, x_244); +x_290 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__9; +lean_inc(x_269); +x_291 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_291, 0, x_269); +lean_ctor_set(x_291, 1, x_290); +lean_inc(x_269); +x_292 = l_Lean_Syntax_node2(x_269, x_288, x_291, x_265); +x_293 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__9; +lean_inc(x_269); +x_294 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_294, 0, x_269); +lean_ctor_set(x_294, 1, x_288); +lean_ctor_set(x_294, 2, x_293); +x_295 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__14; +lean_inc(x_269); +x_296 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_296, 0, x_269); +lean_ctor_set(x_296, 1, x_295); +x_297 = l_Lean_Syntax_node5(x_269, x_20, x_287, x_289, x_292, x_294, x_296); +x_22 = x_297; +x_23 = x_285; +goto block_227; +} +} +else +{ +uint8_t x_298; +lean_dec(x_244); +lean_dec(x_17); +lean_dec(x_13); +lean_dec(x_7); +lean_dec(x_5); +x_298 = !lean_is_exclusive(x_264); +if (x_298 == 0) +{ +return x_264; +} +else +{ +lean_object* x_299; lean_object* x_300; lean_object* x_301; +x_299 = lean_ctor_get(x_264, 0); +x_300 = lean_ctor_get(x_264, 1); +lean_inc(x_300); +lean_inc(x_299); +lean_dec(x_264); +x_301 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_301, 0, x_299); +lean_ctor_set(x_301, 1, x_300); +return x_301; +} +} +} +} +} +} +block_227: +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; size_t x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; uint8_t x_32; +x_24 = lean_array_pop(x_17); +x_25 = l_Array_append___rarg(x_24, x_2); +x_26 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_26, 0, x_22); +lean_ctor_set(x_26, 1, x_16); +x_27 = lean_array_mk(x_26); +x_28 = lean_array_size(x_27); +x_29 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__3; +x_30 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__3(x_29, x_28, x_3, x_27); +x_31 = l_Array_append___rarg(x_25, x_30); +lean_dec(x_30); +x_32 = lean_ctor_get_uint8(x_4, sizeof(void*)*2); +if (x_32 == 0) +{ +lean_object* x_33; uint8_t x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; uint8_t x_38; +x_33 = lean_ctor_get(x_13, 5); +lean_inc(x_33); +x_34 = 0; +x_35 = l_Lean_SourceInfo_fromRef(x_33, x_34); +lean_dec(x_33); +x_36 = lean_ctor_get(x_13, 10); +lean_inc(x_36); +lean_dec(x_13); +x_37 = lean_st_ref_get(x_14, x_23); +x_38 = !lean_is_exclusive(x_37); +if (x_38 == 0) +{ +lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; +x_39 = lean_ctor_get(x_37, 0); +x_40 = lean_ctor_get(x_39, 0); +lean_inc(x_40); +lean_dec(x_39); +x_41 = lean_environment_main_module(x_40); +x_42 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17; +x_43 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__9; +lean_inc(x_35); +x_44 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_44, 0, x_35); +lean_ctor_set(x_44, 1, x_42); +lean_ctor_set(x_44, 2, x_43); +x_45 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__9; +lean_inc(x_35); +x_46 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_46, 0, x_35); +lean_ctor_set(x_46, 1, x_45); +x_47 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__10; +lean_inc(x_35); +x_48 = l_Lean_Syntax_node1(x_35, x_47, x_46); +lean_inc(x_35); +x_49 = l_Lean_Syntax_node1(x_35, x_42, x_48); +x_50 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__8; +lean_inc_n(x_44, 5); +lean_inc(x_35); +x_51 = l_Lean_Syntax_node6(x_35, x_50, x_44, x_44, x_49, x_44, x_44, x_44); +x_52 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__13; +lean_inc(x_35); +x_53 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_53, 0, x_35); +lean_ctor_set(x_53, 1, x_52); +x_54 = lean_mk_syntax_ident(x_5); +x_55 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__15; +lean_inc(x_44); +lean_inc(x_35); +x_56 = l_Lean_Syntax_node2(x_35, x_55, x_54, x_44); +x_57 = l_Array_append___rarg(x_43, x_31); +lean_dec(x_31); +lean_inc(x_35); +x_58 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_58, 0, x_35); +lean_ctor_set(x_58, 1, x_42); +lean_ctor_set(x_58, 2, x_57); +x_59 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__9; +lean_inc(x_35); +x_60 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_60, 0, x_35); +lean_ctor_set(x_60, 1, x_59); +x_61 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__19; +x_62 = l_Lean_addMacroScope(x_41, x_61, x_36); +x_63 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__18; +x_64 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__25; +lean_inc(x_35); +x_65 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_65, 0, x_35); +lean_ctor_set(x_65, 1, x_63); +lean_ctor_set(x_65, 2, x_62); +lean_ctor_set(x_65, 3, x_64); +x_66 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__8; +lean_inc(x_35); +x_67 = l_Lean_Syntax_node2(x_35, x_66, x_60, x_65); +lean_inc(x_35); +x_68 = l_Lean_Syntax_node1(x_35, x_42, x_67); +x_69 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__17; +lean_inc(x_35); +x_70 = l_Lean_Syntax_node2(x_35, x_69, x_58, x_68); +x_71 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__25; +lean_inc(x_35); +x_72 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_72, 0, x_35); +lean_ctor_set(x_72, 1, x_71); +x_73 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__30; +lean_inc_n(x_44, 2); +lean_inc(x_35); +x_74 = l_Lean_Syntax_node2(x_35, x_73, x_44, x_44); +x_75 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__27; +lean_inc(x_44); +lean_inc(x_35); +x_76 = l_Lean_Syntax_node4(x_35, x_75, x_72, x_7, x_74, x_44); +x_77 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__12; +lean_inc(x_35); +x_78 = l_Lean_Syntax_node5(x_35, x_77, x_53, x_56, x_70, x_76, x_44); +x_79 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__6; +x_80 = l_Lean_Syntax_node2(x_35, x_79, x_51, x_78); +lean_ctor_set(x_37, 0, x_80); +return x_37; +} +else +{ +lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; +x_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_inc(x_83); +lean_dec(x_81); +x_84 = lean_environment_main_module(x_83); +x_85 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17; +x_86 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__9; +lean_inc(x_35); +x_87 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_87, 0, x_35); +lean_ctor_set(x_87, 1, x_85); +lean_ctor_set(x_87, 2, x_86); +x_88 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__9; +lean_inc(x_35); +x_89 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_89, 0, x_35); +lean_ctor_set(x_89, 1, x_88); +x_90 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__10; +lean_inc(x_35); +x_91 = l_Lean_Syntax_node1(x_35, x_90, x_89); +lean_inc(x_35); +x_92 = l_Lean_Syntax_node1(x_35, x_85, x_91); +x_93 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__8; +lean_inc_n(x_87, 5); +lean_inc(x_35); +x_94 = l_Lean_Syntax_node6(x_35, x_93, x_87, x_87, x_92, x_87, x_87, x_87); +x_95 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__13; +lean_inc(x_35); +x_96 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_96, 0, x_35); +lean_ctor_set(x_96, 1, x_95); +x_97 = lean_mk_syntax_ident(x_5); +x_98 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__15; +lean_inc(x_87); +lean_inc(x_35); +x_99 = l_Lean_Syntax_node2(x_35, x_98, x_97, x_87); +x_100 = l_Array_append___rarg(x_86, x_31); +lean_dec(x_31); +lean_inc(x_35); +x_101 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_101, 0, x_35); +lean_ctor_set(x_101, 1, x_85); +lean_ctor_set(x_101, 2, x_100); +x_102 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__9; +lean_inc(x_35); +x_103 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_103, 0, x_35); +lean_ctor_set(x_103, 1, x_102); +x_104 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__19; +x_105 = l_Lean_addMacroScope(x_84, x_104, x_36); +x_106 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__18; +x_107 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__25; +lean_inc(x_35); +x_108 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_108, 0, x_35); +lean_ctor_set(x_108, 1, x_106); +lean_ctor_set(x_108, 2, x_105); +lean_ctor_set(x_108, 3, x_107); +x_109 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__8; +lean_inc(x_35); +x_110 = l_Lean_Syntax_node2(x_35, x_109, x_103, x_108); +lean_inc(x_35); +x_111 = l_Lean_Syntax_node1(x_35, x_85, x_110); +x_112 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__17; +lean_inc(x_35); +x_113 = l_Lean_Syntax_node2(x_35, x_112, x_101, x_111); +x_114 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__25; +lean_inc(x_35); +x_115 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_115, 0, x_35); +lean_ctor_set(x_115, 1, x_114); +x_116 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__30; +lean_inc_n(x_87, 2); +lean_inc(x_35); +x_117 = l_Lean_Syntax_node2(x_35, x_116, x_87, x_87); +x_118 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__27; +lean_inc(x_87); +lean_inc(x_35); +x_119 = l_Lean_Syntax_node4(x_35, x_118, x_115, x_7, x_117, x_87); +x_120 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__12; +lean_inc(x_35); +x_121 = l_Lean_Syntax_node5(x_35, x_120, x_96, x_99, x_113, x_119, x_87); +x_122 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__6; +x_123 = l_Lean_Syntax_node2(x_35, x_122, x_94, x_121); +x_124 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_124, 0, x_123); +lean_ctor_set(x_124, 1, x_82); +return x_124; +} +} +else +{ +lean_object* x_125; uint8_t x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; uint8_t x_130; +x_125 = lean_ctor_get(x_13, 5); +lean_inc(x_125); +x_126 = 0; +x_127 = l_Lean_SourceInfo_fromRef(x_125, x_126); +lean_dec(x_125); +x_128 = lean_ctor_get(x_13, 10); +lean_inc(x_128); +lean_dec(x_13); +x_129 = lean_st_ref_get(x_14, x_23); +x_130 = !lean_is_exclusive(x_129); +if (x_130 == 0) +{ +lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; +x_131 = lean_ctor_get(x_129, 0); +x_132 = lean_ctor_get(x_131, 0); +lean_inc(x_132); +lean_dec(x_131); +x_133 = lean_environment_main_module(x_132); +x_134 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17; +x_135 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__9; +lean_inc(x_127); +x_136 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_136, 0, x_127); +lean_ctor_set(x_136, 1, x_134); +lean_ctor_set(x_136, 2, x_135); +x_137 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__9; +lean_inc(x_127); +x_138 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_138, 0, x_127); +lean_ctor_set(x_138, 1, x_137); +x_139 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__10; +lean_inc(x_127); +x_140 = l_Lean_Syntax_node1(x_127, x_139, x_138); +lean_inc(x_127); +x_141 = l_Lean_Syntax_node1(x_127, x_134, x_140); +x_142 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__31; +lean_inc(x_127); +x_143 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_143, 0, x_127); +lean_ctor_set(x_143, 1, x_142); +x_144 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__32; +lean_inc(x_127); +x_145 = l_Lean_Syntax_node1(x_127, x_144, x_143); +lean_inc(x_127); +x_146 = l_Lean_Syntax_node1(x_127, x_134, x_145); +x_147 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__8; +lean_inc_n(x_136, 4); +lean_inc(x_127); +x_148 = l_Lean_Syntax_node6(x_127, x_147, x_136, x_136, x_141, x_136, x_136, x_146); +x_149 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__13; +lean_inc(x_127); +x_150 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_150, 0, x_127); +lean_ctor_set(x_150, 1, x_149); +x_151 = lean_mk_syntax_ident(x_5); +x_152 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__15; +lean_inc(x_136); +lean_inc(x_127); +x_153 = l_Lean_Syntax_node2(x_127, x_152, x_151, x_136); +x_154 = l_Array_append___rarg(x_135, x_31); +lean_dec(x_31); +lean_inc(x_127); +x_155 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_155, 0, x_127); +lean_ctor_set(x_155, 1, x_134); +lean_ctor_set(x_155, 2, x_154); +x_156 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__9; +lean_inc(x_127); +x_157 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_157, 0, x_127); +lean_ctor_set(x_157, 1, x_156); +x_158 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__19; +x_159 = l_Lean_addMacroScope(x_133, x_158, x_128); +x_160 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__18; +x_161 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__25; +lean_inc(x_127); +x_162 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_162, 0, x_127); +lean_ctor_set(x_162, 1, x_160); +lean_ctor_set(x_162, 2, x_159); +lean_ctor_set(x_162, 3, x_161); +x_163 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__8; +lean_inc(x_127); +x_164 = l_Lean_Syntax_node2(x_127, x_163, x_157, x_162); +lean_inc(x_127); +x_165 = l_Lean_Syntax_node1(x_127, x_134, x_164); +x_166 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__17; +lean_inc(x_127); +x_167 = l_Lean_Syntax_node2(x_127, x_166, x_155, x_165); +x_168 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__25; +lean_inc(x_127); +x_169 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_169, 0, x_127); +lean_ctor_set(x_169, 1, x_168); +x_170 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__30; +lean_inc_n(x_136, 2); +lean_inc(x_127); +x_171 = l_Lean_Syntax_node2(x_127, x_170, x_136, x_136); +x_172 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__27; +lean_inc(x_136); +lean_inc(x_127); +x_173 = l_Lean_Syntax_node4(x_127, x_172, x_169, x_7, x_171, x_136); +x_174 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__12; +lean_inc(x_127); +x_175 = l_Lean_Syntax_node5(x_127, x_174, x_150, x_153, x_167, x_173, x_136); +x_176 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__6; +x_177 = l_Lean_Syntax_node2(x_127, x_176, x_148, x_175); +lean_ctor_set(x_129, 0, x_177); +return x_129; +} +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; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; +x_178 = lean_ctor_get(x_129, 0); +x_179 = lean_ctor_get(x_129, 1); +lean_inc(x_179); +lean_inc(x_178); +lean_dec(x_129); +x_180 = lean_ctor_get(x_178, 0); +lean_inc(x_180); +lean_dec(x_178); +x_181 = lean_environment_main_module(x_180); +x_182 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17; +x_183 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__9; +lean_inc(x_127); +x_184 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_184, 0, x_127); +lean_ctor_set(x_184, 1, x_182); +lean_ctor_set(x_184, 2, x_183); +x_185 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__9; +lean_inc(x_127); +x_186 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_186, 0, x_127); +lean_ctor_set(x_186, 1, x_185); +x_187 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__10; +lean_inc(x_127); +x_188 = l_Lean_Syntax_node1(x_127, x_187, x_186); +lean_inc(x_127); +x_189 = l_Lean_Syntax_node1(x_127, x_182, x_188); +x_190 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__31; +lean_inc(x_127); +x_191 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_191, 0, x_127); +lean_ctor_set(x_191, 1, x_190); +x_192 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__32; +lean_inc(x_127); +x_193 = l_Lean_Syntax_node1(x_127, x_192, x_191); +lean_inc(x_127); +x_194 = l_Lean_Syntax_node1(x_127, x_182, x_193); +x_195 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__8; +lean_inc_n(x_184, 4); +lean_inc(x_127); +x_196 = l_Lean_Syntax_node6(x_127, x_195, x_184, x_184, x_189, x_184, x_184, x_194); +x_197 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__13; +lean_inc(x_127); +x_198 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_198, 0, x_127); +lean_ctor_set(x_198, 1, x_197); +x_199 = lean_mk_syntax_ident(x_5); +x_200 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__15; +lean_inc(x_184); +lean_inc(x_127); +x_201 = l_Lean_Syntax_node2(x_127, x_200, x_199, x_184); +x_202 = l_Array_append___rarg(x_183, x_31); +lean_dec(x_31); +lean_inc(x_127); +x_203 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_203, 0, x_127); +lean_ctor_set(x_203, 1, x_182); +lean_ctor_set(x_203, 2, x_202); +x_204 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__9; +lean_inc(x_127); +x_205 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_205, 0, x_127); +lean_ctor_set(x_205, 1, x_204); +x_206 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__19; +x_207 = l_Lean_addMacroScope(x_181, x_206, x_128); +x_208 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__18; +x_209 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__25; +lean_inc(x_127); +x_210 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_210, 0, x_127); +lean_ctor_set(x_210, 1, x_208); +lean_ctor_set(x_210, 2, x_207); +lean_ctor_set(x_210, 3, x_209); +x_211 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__8; +lean_inc(x_127); +x_212 = l_Lean_Syntax_node2(x_127, x_211, x_205, x_210); +lean_inc(x_127); +x_213 = l_Lean_Syntax_node1(x_127, x_182, x_212); +x_214 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__17; +lean_inc(x_127); +x_215 = l_Lean_Syntax_node2(x_127, x_214, x_203, x_213); +x_216 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__25; +lean_inc(x_127); +x_217 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_217, 0, x_127); +lean_ctor_set(x_217, 1, x_216); +x_218 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__30; +lean_inc_n(x_184, 2); +lean_inc(x_127); +x_219 = l_Lean_Syntax_node2(x_127, x_218, x_184, x_184); +x_220 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__27; +lean_inc(x_184); +lean_inc(x_127); +x_221 = l_Lean_Syntax_node4(x_127, x_220, x_217, x_7, x_219, x_184); +x_222 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__12; +lean_inc(x_127); +x_223 = l_Lean_Syntax_node5(x_127, x_222, x_198, x_201, x_215, x_221, x_184); +x_224 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__6; +x_225 = l_Lean_Syntax_node2(x_127, x_224, x_196, x_223); +x_226 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_226, 0, x_225); +lean_ctor_set(x_226, 1, x_179); +return x_226; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_ToExpr_mkAuxFunction(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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; lean_object* x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; +x_10 = lean_ctor_get(x_1, 1); +lean_inc(x_10); +x_11 = lean_array_get_size(x_10); +x_12 = lean_nat_dec_lt(x_2, x_11); +lean_dec(x_11); +x_13 = lean_ctor_get(x_1, 0); +lean_inc(x_13); +x_14 = lean_array_get_size(x_13); +x_15 = lean_nat_dec_lt(x_2, x_14); +lean_dec(x_14); +if (x_12 == 0) +{ +lean_object* x_70; lean_object* x_71; +lean_dec(x_10); +x_70 = l_Lean_instInhabitedName; +x_71 = l_outOfBounds___rarg(x_70); +x_16 = x_71; +goto block_69; +} +else +{ +lean_object* x_72; +x_72 = lean_array_fget(x_10, x_2); +lean_dec(x_10); +x_16 = x_72; +goto block_69; +} +block_69: +{ +lean_object* x_17; +if (x_15 == 0) +{ +lean_object* x_66; lean_object* x_67; +lean_dec(x_13); +x_66 = l_Lean_instInhabitedInductiveVal; +x_67 = l_outOfBounds___rarg(x_66); +x_17 = x_67; +goto block_65; +} +else +{ +lean_object* x_68; +x_68 = lean_array_fget(x_13, x_2); +lean_dec(x_13); +x_17 = x_68; +goto block_65; +} +block_65: +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_18 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__12; +x_19 = lean_unsigned_to_nat(1u); +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_17); +x_20 = l_Lean_Elab_Deriving_mkHeader(x_18, x_19, x_17, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +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; size_t x_26; size_t x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; size_t x_34; lean_object* x_35; lean_object* x_36; +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_17, 0); +lean_inc(x_23); +x_24 = lean_ctor_get(x_23, 1); +lean_inc(x_24); +lean_dec(x_23); +x_25 = lean_array_mk(x_24); +x_26 = lean_array_size(x_25); +x_27 = 0; +lean_inc(x_7); +x_28 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1(x_26, x_27, x_25, x_3, x_4, x_5, x_6, x_7, x_8, x_22); +x_29 = lean_ctor_get(x_28, 0); +lean_inc(x_29); +x_30 = lean_ctor_get(x_28, 1); +lean_inc(x_30); +lean_dec(x_28); +x_31 = l_Array_unzip___rarg(x_29); +lean_dec(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 = lean_array_size(x_32); +x_35 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__2(x_34, x_27, x_32); +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_35); +lean_inc(x_16); +lean_inc(x_17); +lean_inc(x_21); +x_36 = l_Lean_Elab_Deriving_ToExpr_mkToExprBody(x_21, x_17, x_16, x_35, x_3, x_4, x_5, x_6, x_7, x_8, x_30); +if (lean_obj_tag(x_36) == 0) +{ +uint8_t x_37; +x_37 = lean_ctor_get_uint8(x_1, sizeof(void*)*2); +if (x_37 == 0) +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; +lean_dec(x_35); +x_38 = lean_ctor_get(x_36, 0); +lean_inc(x_38); +x_39 = lean_ctor_get(x_36, 1); +lean_inc(x_39); +lean_dec(x_36); +x_40 = lean_box(0); +x_41 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1(x_21, x_33, x_27, x_1, x_16, x_17, x_38, x_40, x_3, x_4, x_5, x_6, x_7, x_8, x_39); +lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +lean_dec(x_33); +return x_41; +} +else +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_42 = lean_ctor_get(x_36, 0); +lean_inc(x_42); +x_43 = lean_ctor_get(x_36, 1); +lean_inc(x_43); +lean_dec(x_36); +x_44 = lean_ctor_get(x_21, 1); +lean_inc(x_44); +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_45 = l_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls(x_1, x_44, x_35, x_3, x_4, x_5, x_6, x_7, x_8, x_43); +lean_dec(x_35); +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; +x_46 = lean_ctor_get(x_45, 0); +lean_inc(x_46); +x_47 = lean_ctor_get(x_45, 1); +lean_inc(x_47); +lean_dec(x_45); +x_48 = l_Lean_Elab_Deriving_mkLet(x_46, x_42, x_3, x_4, x_5, x_6, x_7, x_8, x_47); +lean_dec(x_46); +x_49 = lean_ctor_get(x_48, 0); +lean_inc(x_49); +x_50 = lean_ctor_get(x_48, 1); +lean_inc(x_50); +lean_dec(x_48); +x_51 = lean_box(0); +x_52 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1(x_21, x_33, x_27, x_1, x_16, x_17, x_49, x_51, x_3, x_4, x_5, x_6, x_7, x_8, x_50); +lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +lean_dec(x_33); +return x_52; +} +else +{ +uint8_t x_53; +lean_dec(x_42); +lean_dec(x_33); +lean_dec(x_21); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_53 = !lean_is_exclusive(x_45); +if (x_53 == 0) +{ +return x_45; +} +else +{ +lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_54 = lean_ctor_get(x_45, 0); +x_55 = lean_ctor_get(x_45, 1); +lean_inc(x_55); +lean_inc(x_54); +lean_dec(x_45); +x_56 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_56, 0, x_54); +lean_ctor_set(x_56, 1, x_55); +return x_56; +} +} +} +} +else +{ +uint8_t x_57; +lean_dec(x_35); +lean_dec(x_33); +lean_dec(x_21); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_57 = !lean_is_exclusive(x_36); +if (x_57 == 0) +{ +return x_36; +} +else +{ +lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_58 = lean_ctor_get(x_36, 0); +x_59 = lean_ctor_get(x_36, 1); +lean_inc(x_59); +lean_inc(x_58); +lean_dec(x_36); +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_17); +lean_dec(x_16); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_61 = !lean_is_exclusive(x_20); +if (x_61 == 0) +{ +return x_20; +} +else +{ +lean_object* x_62; lean_object* x_63; lean_object* x_64; +x_62 = lean_ctor_get(x_20, 0); +x_63 = lean_ctor_get(x_20, 1); +lean_inc(x_63); +lean_inc(x_62); +lean_dec(x_20); +x_64 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_64, 0, x_62); +lean_ctor_set(x_64, 1, x_63); +return x_64; +} +} +} +} +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___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: +{ +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_Elab_Deriving_ToExpr_mkAuxFunction___spec__1(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); +lean_dec(x_5); +lean_dec(x_4); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +size_t x_4; size_t x_5; lean_object* x_6; +x_4 = lean_unbox_usize(x_1); +lean_dec(x_1); +x_5 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_6 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__2(x_4, x_5, x_3); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +size_t x_5; size_t x_6; lean_object* x_7; +x_5 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_6 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_7 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__3(x_1, x_5, x_6, x_4); +lean_dec(x_1); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___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_Deriving_ToExpr_mkAuxFunction___spec__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +_start: +{ +size_t x_16; lean_object* x_17; +x_16 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_17 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1(x_1, x_2, x_16, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +lean_dec(x_14); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_8); +lean_dec(x_4); +lean_dec(x_2); +return x_17; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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_Deriving_ToExpr_mkAuxFunction(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_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkAuxFunctions___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +lean_object* x_15; uint8_t x_16; +x_15 = lean_ctor_get(x_3, 1); +x_16 = lean_nat_dec_lt(x_5, x_15); +if (x_16 == 0) +{ +lean_object* x_17; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_5); +lean_dec(x_1); +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_4); +lean_ctor_set(x_17, 1, x_14); +return x_17; +} +else +{ +lean_object* x_18; +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_1); +x_18 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction(x_1, x_5, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_18) == 0) +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_18, 1); +lean_inc(x_20); +lean_dec(x_18); +x_21 = lean_array_push(x_4, x_19); +x_22 = lean_ctor_get(x_3, 2); +x_23 = lean_nat_add(x_5, x_22); +lean_dec(x_5); +x_4 = x_21; +x_5 = x_23; +x_6 = lean_box(0); +x_7 = lean_box(0); +x_14 = x_20; +goto _start; +} +else +{ +uint8_t x_25; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_25 = !lean_is_exclusive(x_18); +if (x_25 == 0) +{ +return x_18; +} +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; +} +} +} +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_mkAuxFunctions___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("mutual", 6, 6); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_mkAuxFunctions___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_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__1; +x_2 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__2; +x_3 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__4; +x_4 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunctions___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_Deriving_ToExpr_mkAuxFunctions___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("end", 3, 3); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_ToExpr_mkAuxFunctions(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_9 = lean_ctor_get(x_1, 0); +lean_inc(x_9); +x_10 = lean_array_get_size(x_9); +lean_dec(x_9); +x_11 = lean_unsigned_to_nat(0u); +x_12 = lean_unsigned_to_nat(1u); +x_13 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_13, 0, x_11); +lean_ctor_set(x_13, 1, x_10); +lean_ctor_set(x_13, 2, x_12); +x_14 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__1; +lean_inc(x_7); +lean_inc(x_6); +x_15 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkAuxFunctions___spec__1(x_1, x_13, x_13, x_14, x_11, lean_box(0), lean_box(0), x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_13); +if (lean_obj_tag(x_15) == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; lean_object* x_20; lean_object* x_21; uint8_t 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 = lean_ctor_get(x_6, 5); +lean_inc(x_18); +lean_dec(x_6); +x_19 = 0; +x_20 = l_Lean_SourceInfo_fromRef(x_18, x_19); +lean_dec(x_18); +x_21 = lean_st_ref_get(x_7, x_17); +lean_dec(x_7); +x_22 = !lean_is_exclusive(x_21); +if (x_22 == 0) +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_23 = lean_ctor_get(x_21, 0); +lean_dec(x_23); +x_24 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunctions___closed__1; +lean_inc(x_20); +x_25 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_25, 0, x_20); +lean_ctor_set(x_25, 1, x_24); +x_26 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__9; +x_27 = l_Array_append___rarg(x_26, x_16); +lean_dec(x_16); +x_28 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17; +lean_inc(x_20); +x_29 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_29, 0, x_20); +lean_ctor_set(x_29, 1, x_28); +lean_ctor_set(x_29, 2, x_27); +x_30 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunctions___closed__3; +lean_inc(x_20); +x_31 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_31, 0, x_20); +lean_ctor_set(x_31, 1, x_30); +x_32 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunctions___closed__2; +x_33 = l_Lean_Syntax_node3(x_20, x_32, x_25, x_29, x_31); +lean_ctor_set(x_21, 0, x_33); +return x_21; +} +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; +x_34 = lean_ctor_get(x_21, 1); +lean_inc(x_34); +lean_dec(x_21); +x_35 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunctions___closed__1; +lean_inc(x_20); +x_36 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_36, 0, x_20); +lean_ctor_set(x_36, 1, x_35); +x_37 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__9; +x_38 = l_Array_append___rarg(x_37, x_16); +lean_dec(x_16); +x_39 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17; +lean_inc(x_20); +x_40 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_40, 0, x_20); +lean_ctor_set(x_40, 1, x_39); +lean_ctor_set(x_40, 2, x_38); +x_41 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunctions___closed__3; +lean_inc(x_20); +x_42 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_42, 0, x_20); +lean_ctor_set(x_42, 1, x_41); +x_43 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunctions___closed__2; +x_44 = l_Lean_Syntax_node3(x_20, x_43, x_36, x_40, x_42); +x_45 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_45, 0, x_44); +lean_ctor_set(x_45, 1, x_34); +return x_45; +} +} +else +{ +uint8_t x_46; +lean_dec(x_7); +lean_dec(x_6); +x_46 = !lean_is_exclusive(x_15); +if (x_46 == 0) +{ +return x_15; +} +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_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_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkAuxFunctions___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +lean_object* x_15; +x_15 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkAuxFunctions___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +lean_dec(x_3); +lean_dec(x_2); +return x_15; +} +} +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("instBinder", 10, 10); +return x_1; +} +} +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__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_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__1; +x_2 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__2; +x_3 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__3; +x_4 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__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_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__1(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_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_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_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; +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); +x_17 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__2; +x_18 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_17, x_9, x_10, x_11); +x_19 = !lean_is_exclusive(x_18); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; +x_20 = lean_ctor_get(x_18, 0); +x_21 = lean_ctor_get(x_18, 1); +x_22 = lean_mk_syntax_ident(x_20); +x_23 = lean_ctor_get(x_9, 5); +lean_inc(x_23); +x_24 = 0; +x_25 = l_Lean_SourceInfo_fromRef(x_23, x_24); +lean_dec(x_23); +x_26 = lean_ctor_get(x_9, 10); +lean_inc(x_26); +x_27 = lean_st_ref_get(x_10, x_21); +x_28 = !lean_is_exclusive(x_27); +if (x_28 == 0) +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; size_t x_55; size_t x_56; lean_object* x_57; +x_29 = lean_ctor_get(x_27, 0); +x_30 = lean_ctor_get(x_27, 1); +x_31 = lean_ctor_get(x_29, 0); +lean_inc(x_31); +lean_dec(x_29); +x_32 = lean_environment_main_module(x_31); +x_33 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__14; +lean_inc(x_25); +lean_ctor_set_tag(x_27, 2); +lean_ctor_set(x_27, 1, x_33); +lean_ctor_set(x_27, 0, x_25); +x_34 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__9; +lean_inc(x_25); +lean_ctor_set_tag(x_18, 2); +lean_ctor_set(x_18, 1, x_34); +lean_ctor_set(x_18, 0, x_25); +x_35 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17; +lean_inc(x_22); +lean_inc(x_25); +x_36 = l_Lean_Syntax_node2(x_25, x_35, x_22, x_18); +x_37 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__7; +x_38 = l_Lean_addMacroScope(x_32, x_37, x_26); +x_39 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__6; +x_40 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__13; +lean_inc(x_25); +x_41 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_41, 0, x_25); +lean_ctor_set(x_41, 1, x_39); +lean_ctor_set(x_41, 2, x_38); +lean_ctor_set(x_41, 3, x_40); +x_42 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__8; +lean_inc(x_25); +x_43 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_43, 0, x_25); +lean_ctor_set(x_43, 1, x_42); +x_44 = lean_mk_syntax_ident(x_14); +lean_inc(x_25); +x_45 = l_Lean_Syntax_node1(x_25, x_35, x_44); +x_46 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__14; +lean_inc(x_25); +x_47 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_47, 0, x_25); +lean_ctor_set(x_47, 1, x_46); +x_48 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__7; +lean_inc(x_25); +x_49 = l_Lean_Syntax_node4(x_25, x_48, x_41, x_43, x_45, x_47); +x_50 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__15; +lean_inc(x_25); +x_51 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_51, 0, x_25); +lean_ctor_set(x_51, 1, x_50); +x_52 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__1___closed__2; +x_53 = l_Lean_Syntax_node4(x_25, x_52, x_27, x_36, x_49, x_51); +x_54 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_54, 0, x_22); +lean_ctor_set(x_54, 1, x_53); +x_55 = 1; +x_56 = lean_usize_add(x_3, x_55); +x_57 = lean_array_uset(x_16, x_3, x_54); +x_3 = x_56; +x_4 = x_57; +x_11 = x_30; +goto _start; +} +else +{ +lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; 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; size_t x_86; size_t x_87; lean_object* x_88; +x_59 = lean_ctor_get(x_27, 0); +x_60 = lean_ctor_get(x_27, 1); +lean_inc(x_60); +lean_inc(x_59); +lean_dec(x_27); +x_61 = lean_ctor_get(x_59, 0); +lean_inc(x_61); +lean_dec(x_59); +x_62 = lean_environment_main_module(x_61); +x_63 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__14; +lean_inc(x_25); +x_64 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_64, 0, x_25); +lean_ctor_set(x_64, 1, x_63); +x_65 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__9; +lean_inc(x_25); +lean_ctor_set_tag(x_18, 2); +lean_ctor_set(x_18, 1, x_65); +lean_ctor_set(x_18, 0, x_25); +x_66 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17; +lean_inc(x_22); +lean_inc(x_25); +x_67 = l_Lean_Syntax_node2(x_25, x_66, x_22, x_18); +x_68 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__7; +x_69 = l_Lean_addMacroScope(x_62, x_68, x_26); +x_70 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__6; +x_71 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__13; +lean_inc(x_25); +x_72 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_72, 0, x_25); +lean_ctor_set(x_72, 1, x_70); +lean_ctor_set(x_72, 2, x_69); +lean_ctor_set(x_72, 3, x_71); +x_73 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__8; +lean_inc(x_25); +x_74 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_74, 0, x_25); +lean_ctor_set(x_74, 1, x_73); +x_75 = lean_mk_syntax_ident(x_14); +lean_inc(x_25); +x_76 = l_Lean_Syntax_node1(x_25, x_66, x_75); +x_77 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__14; +lean_inc(x_25); +x_78 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_78, 0, x_25); +lean_ctor_set(x_78, 1, x_77); +x_79 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__7; +lean_inc(x_25); +x_80 = l_Lean_Syntax_node4(x_25, x_79, x_72, x_74, x_76, x_78); +x_81 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__15; +lean_inc(x_25); +x_82 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_82, 0, x_25); +lean_ctor_set(x_82, 1, x_81); +x_83 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__1___closed__2; +x_84 = l_Lean_Syntax_node4(x_25, x_83, x_64, x_67, x_80, x_82); +x_85 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_85, 0, x_22); +lean_ctor_set(x_85, 1, x_84); +x_86 = 1; +x_87 = lean_usize_add(x_3, x_86); +x_88 = lean_array_uset(x_16, x_3, x_85); +x_3 = x_87; +x_4 = x_88; +x_11 = x_60; +goto _start; +} +} +else +{ +lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; uint8_t x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; size_t x_127; size_t x_128; lean_object* x_129; +x_90 = lean_ctor_get(x_18, 0); +x_91 = lean_ctor_get(x_18, 1); +lean_inc(x_91); +lean_inc(x_90); +lean_dec(x_18); +x_92 = lean_mk_syntax_ident(x_90); +x_93 = lean_ctor_get(x_9, 5); +lean_inc(x_93); +x_94 = 0; +x_95 = l_Lean_SourceInfo_fromRef(x_93, x_94); +lean_dec(x_93); +x_96 = lean_ctor_get(x_9, 10); +lean_inc(x_96); +x_97 = lean_st_ref_get(x_10, x_91); +x_98 = lean_ctor_get(x_97, 0); +lean_inc(x_98); +x_99 = lean_ctor_get(x_97, 1); +lean_inc(x_99); +if (lean_is_exclusive(x_97)) { + lean_ctor_release(x_97, 0); + lean_ctor_release(x_97, 1); + x_100 = x_97; +} else { + lean_dec_ref(x_97); + x_100 = lean_box(0); +} +x_101 = lean_ctor_get(x_98, 0); +lean_inc(x_101); +lean_dec(x_98); +x_102 = lean_environment_main_module(x_101); +x_103 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__14; +lean_inc(x_95); +if (lean_is_scalar(x_100)) { + x_104 = lean_alloc_ctor(2, 2, 0); +} else { + x_104 = x_100; + lean_ctor_set_tag(x_104, 2); +} +lean_ctor_set(x_104, 0, x_95); +lean_ctor_set(x_104, 1, x_103); +x_105 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__9; +lean_inc(x_95); +x_106 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_106, 0, x_95); +lean_ctor_set(x_106, 1, x_105); +x_107 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17; +lean_inc(x_92); +lean_inc(x_95); +x_108 = l_Lean_Syntax_node2(x_95, x_107, x_92, x_106); +x_109 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__7; +x_110 = l_Lean_addMacroScope(x_102, x_109, x_96); +x_111 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__6; +x_112 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__13; +lean_inc(x_95); +x_113 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_113, 0, x_95); +lean_ctor_set(x_113, 1, x_111); +lean_ctor_set(x_113, 2, x_110); +lean_ctor_set(x_113, 3, x_112); +x_114 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__8; +lean_inc(x_95); +x_115 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_115, 0, x_95); +lean_ctor_set(x_115, 1, x_114); +x_116 = lean_mk_syntax_ident(x_14); +lean_inc(x_95); +x_117 = l_Lean_Syntax_node1(x_95, x_107, x_116); +x_118 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__14; +lean_inc(x_95); +x_119 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_119, 0, x_95); +lean_ctor_set(x_119, 1, x_118); +x_120 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__7; +lean_inc(x_95); +x_121 = l_Lean_Syntax_node4(x_95, x_120, x_113, x_115, x_117, x_119); +x_122 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__15; +lean_inc(x_95); +x_123 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_123, 0, x_95); +lean_ctor_set(x_123, 1, x_122); +x_124 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__1___closed__2; +x_125 = l_Lean_Syntax_node4(x_95, x_124, x_104, x_108, x_121, x_123); +x_126 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_126, 0, x_92); +lean_ctor_set(x_126, 1, x_125); +x_127 = 1; +x_128 = lean_usize_add(x_3, x_127); +x_129 = lean_array_uset(x_16, x_3, x_126); +x_3 = x_128; +x_4 = x_129; +x_11 = x_99; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__2(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { +_start: +{ +uint8_t x_5; +x_5 = lean_usize_dec_lt(x_3, x_2); +if (x_5 == 0) +{ +return x_4; +} +else +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x_9; size_t x_10; lean_object* x_11; +x_6 = lean_array_uget(x_4, x_3); +x_7 = lean_unsigned_to_nat(0u); +x_8 = lean_array_uset(x_4, x_3, x_7); +x_9 = 1; +x_10 = lean_usize_add(x_3, x_9); +x_11 = lean_array_uset(x_8, x_3, x_6); +x_3 = x_10; +x_4 = x_11; +goto _start; +} +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("instance", 8, 8); +return x_1; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__1; +x_2 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__2; +x_3 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__4; +x_4 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__1; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("attrKind", 8, 8); +return x_1; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__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_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__1; +x_2 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__2; +x_3 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__3; +x_4 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__3; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("declSig", 7, 7); +return x_1; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__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 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__1; +x_2 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__2; +x_3 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__4; +x_4 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__5; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("whereStructInst", 15, 15); +return x_1; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___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_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__1; +x_2 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__2; +x_3 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__4; +x_4 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__7; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__9() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("where", 5, 5); +return x_1; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__10() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("term", 4, 4); +return x_1; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__10; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___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; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, size_t x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +uint8_t x_15; +x_15 = lean_usize_dec_lt(x_6, x_5); +if (x_15 == 0) +{ +lean_object* x_16; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_7); +lean_ctor_set(x_16, 1, x_14); +return x_16; +} +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_27; +x_17 = lean_array_uget(x_4, x_6); +x_27 = !lean_is_exclusive(x_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; uint8_t x_33; +x_28 = lean_ctor_get(x_7, 0); +x_29 = lean_ctor_get(x_7, 1); +x_30 = lean_ctor_get(x_28, 0); +lean_inc(x_30); +x_31 = lean_ctor_get(x_28, 1); +lean_inc(x_31); +x_32 = lean_ctor_get(x_28, 2); +lean_inc(x_32); +x_33 = lean_nat_dec_lt(x_31, x_32); +if (x_33 == 0) +{ +lean_object* x_34; +lean_dec(x_32); +lean_dec(x_31); +lean_dec(x_30); +lean_dec(x_17); +x_34 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_34, 0, x_7); +x_18 = x_34; +x_19 = x_14; +goto block_26; +} +else +{ +uint8_t x_35; +x_35 = !lean_is_exclusive(x_28); +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; uint8_t x_44; +x_36 = lean_ctor_get(x_28, 2); +lean_dec(x_36); +x_37 = lean_ctor_get(x_28, 1); +lean_dec(x_37); +x_38 = lean_ctor_get(x_28, 0); +lean_dec(x_38); +x_39 = lean_array_fget(x_30, x_31); +x_40 = lean_unsigned_to_nat(1u); +x_41 = lean_nat_add(x_31, x_40); +lean_dec(x_31); +lean_ctor_set(x_28, 1, x_41); +x_42 = lean_ctor_get(x_17, 0); +lean_inc(x_42); +x_43 = lean_ctor_get(x_42, 0); +lean_inc(x_43); +x_44 = l_Array_contains___at_Lean_registerInternalExceptionId___spec__1(x_1, x_43); +lean_dec(x_43); +if (x_44 == 0) +{ +lean_object* x_45; +lean_dec(x_42); +lean_dec(x_39); +lean_dec(x_17); +x_45 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_45, 0, x_7); +x_18 = x_45; +x_19 = x_14; +goto block_26; +} +else +{ +lean_object* x_46; +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_17); +x_46 = l_Lean_Elab_Deriving_mkInductArgNames(x_17, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_46) == 0) +{ +lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; +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 = l_Lean_Elab_Deriving_mkImplicitBinders(x_47, x_8, x_9, x_10, x_11, x_12, x_13, x_48); +x_50 = lean_ctor_get(x_49, 0); +lean_inc(x_50); +x_51 = lean_ctor_get(x_49, 1); +lean_inc(x_51); +lean_dec(x_49); +x_52 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__12; +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_47); +lean_inc(x_17); +x_53 = l_Lean_Elab_Deriving_mkInstImplicitBinders(x_52, x_17, x_47, x_8, x_9, x_10, x_11, x_12, x_13, x_51); +if (lean_obj_tag(x_53) == 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; size_t x_60; size_t x_61; lean_object* x_62; uint8_t x_63; +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 = l_Array_append___rarg(x_50, x_54); +lean_dec(x_54); +x_58 = lean_ctor_get(x_42, 1); +lean_inc(x_58); +lean_dec(x_42); +x_59 = lean_array_mk(x_58); +x_60 = lean_array_size(x_59); +x_61 = 0; +lean_inc(x_12); +x_62 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__1(x_56, x_60, x_61, x_59, x_8, x_9, x_10, x_11, x_12, x_13, x_55); +x_63 = !lean_is_exclusive(x_62); +if (x_63 == 0) +{ +lean_object* x_64; lean_object* x_65; lean_object* x_66; uint8_t x_67; +x_64 = lean_ctor_get(x_62, 0); +x_65 = lean_ctor_get(x_62, 1); +x_66 = l_Array_unzip___rarg(x_64); +lean_dec(x_64); +x_67 = !lean_is_exclusive(x_66); +if (x_67 == 0) +{ +lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; uint8_t x_72; +x_68 = lean_ctor_get(x_66, 0); +x_69 = lean_ctor_get(x_66, 1); +x_70 = l_Array_append___rarg(x_57, x_69); +lean_dec(x_69); +lean_inc(x_47); +lean_inc(x_17); +x_71 = l_Lean_Elab_Deriving_mkInductiveApp(x_17, x_47, x_8, x_9, x_10, x_11, x_12, x_13, x_65); +x_72 = !lean_is_exclusive(x_71); +if (x_72 == 0) +{ +lean_object* x_73; lean_object* x_74; lean_object* x_75; +x_73 = lean_ctor_get(x_71, 0); +x_74 = lean_ctor_get(x_71, 1); +lean_inc(x_8); +lean_inc(x_17); +x_75 = l_Lean_Elab_Deriving_ToExpr_updateIndType(x_17, x_73, x_8, x_9, x_10, x_11, x_12, x_13, x_74); +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_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +x_78 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr(x_17, x_47, x_8, x_9, x_10, x_11, x_12, x_13, x_77); +if (lean_obj_tag(x_78) == 0) +{ +lean_object* x_79; lean_object* x_80; lean_object* x_81; uint8_t x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; uint8_t x_86; +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_12, 5); +lean_inc(x_81); +x_82 = 0; +x_83 = l_Lean_SourceInfo_fromRef(x_81, x_82); +lean_dec(x_81); +x_84 = lean_ctor_get(x_12, 10); +lean_inc(x_84); +x_85 = lean_st_ref_get(x_13, x_80); +x_86 = !lean_is_exclusive(x_85); +if (x_86 == 0) +{ +lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; size_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; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; +x_87 = lean_ctor_get(x_85, 0); +x_88 = lean_ctor_get(x_85, 1); +x_89 = lean_ctor_get(x_87, 0); +lean_inc(x_89); +lean_dec(x_87); +x_90 = lean_environment_main_module(x_89); +x_91 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17; +x_92 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__9; +lean_inc(x_83); +x_93 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_93, 0, x_83); +lean_ctor_set(x_93, 1, x_91); +lean_ctor_set(x_93, 2, x_92); +x_94 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__8; +lean_inc_n(x_93, 6); +lean_inc(x_83); +x_95 = l_Lean_Syntax_node6(x_83, x_94, x_93, x_93, x_93, x_93, x_93, x_93); +x_96 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__4; +lean_inc(x_93); +lean_inc(x_83); +x_97 = l_Lean_Syntax_node1(x_83, x_96, x_93); +x_98 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__1; +lean_inc(x_83); +lean_ctor_set_tag(x_85, 2); +lean_ctor_set(x_85, 1, x_98); +lean_ctor_set(x_85, 0, x_83); +x_99 = l_Array_append___rarg(x_92, x_70); +lean_dec(x_70); +lean_inc(x_83); +x_100 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_100, 0, x_83); +lean_ctor_set(x_100, 1, x_91); +lean_ctor_set(x_100, 2, x_99); +x_101 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__9; +lean_inc(x_83); +lean_ctor_set_tag(x_71, 2); +lean_ctor_set(x_71, 1, x_101); +lean_ctor_set(x_71, 0, x_83); +x_102 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__11; +lean_inc(x_84); +lean_inc(x_90); +x_103 = l_Lean_addMacroScope(x_90, x_102, x_84); +x_104 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__10; +x_105 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__24; +lean_inc(x_83); +x_106 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_106, 0, x_83); +lean_ctor_set(x_106, 1, x_104); +lean_ctor_set(x_106, 2, x_103); +lean_ctor_set(x_106, 3, x_105); +lean_inc(x_83); +x_107 = l_Lean_Syntax_node1(x_83, x_91, x_76); +x_108 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__5; +lean_inc(x_83); +x_109 = l_Lean_Syntax_node2(x_83, x_108, x_106, x_107); +x_110 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__8; +lean_inc(x_83); +x_111 = l_Lean_Syntax_node2(x_83, x_110, x_71, x_109); +x_112 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__6; +lean_inc(x_83); +x_113 = l_Lean_Syntax_node2(x_83, x_112, x_100, x_111); +x_114 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__9; +lean_inc(x_83); +lean_ctor_set_tag(x_66, 2); +lean_ctor_set(x_66, 1, x_114); +lean_ctor_set(x_66, 0, x_83); +x_115 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__3; +lean_inc(x_84); +lean_inc(x_90); +x_116 = l_Lean_addMacroScope(x_90, x_115, x_84); +x_117 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__2; +x_118 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__36; +lean_inc(x_83); +x_119 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_119, 0, x_83); +lean_ctor_set(x_119, 1, x_117); +lean_ctor_set(x_119, 2, x_116); +lean_ctor_set(x_119, 3, x_118); +x_120 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__34; +lean_inc(x_93); +lean_inc(x_83); +x_121 = l_Lean_Syntax_node2(x_83, x_120, x_119, x_93); +x_122 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__25; +lean_inc(x_83); +lean_ctor_set_tag(x_62, 2); +lean_ctor_set(x_62, 1, x_122); +lean_ctor_set(x_62, 0, x_83); +x_123 = lean_mk_syntax_ident(x_39); +x_124 = lean_array_size(x_68); +x_125 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__12; +x_126 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__2(x_125, x_124, x_61, x_68); +x_127 = l_Array_append___rarg(x_92, x_126); +lean_dec(x_126); +lean_inc(x_83); +x_128 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_128, 0, x_83); +lean_ctor_set(x_128, 1, x_91); +lean_ctor_set(x_128, 2, x_127); +lean_inc(x_83); +x_129 = l_Lean_Syntax_node2(x_83, x_108, x_123, x_128); +x_130 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__38; +lean_inc(x_62); +lean_inc(x_83); +x_131 = l_Lean_Syntax_node2(x_83, x_130, x_62, x_129); +lean_inc_n(x_93, 2); +lean_inc(x_83); +x_132 = l_Lean_Syntax_node3(x_83, x_91, x_93, x_93, x_131); +x_133 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__32; +lean_inc(x_83); +x_134 = l_Lean_Syntax_node2(x_83, x_133, x_121, x_132); +x_135 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__10; +x_136 = l_Lean_addMacroScope(x_90, x_135, x_84); +x_137 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__9; +x_138 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__40; +lean_inc(x_83); +x_139 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_139, 0, x_83); +lean_ctor_set(x_139, 1, x_137); +lean_ctor_set(x_139, 2, x_136); +lean_ctor_set(x_139, 3, x_138); +lean_inc(x_93); +lean_inc(x_83); +x_140 = l_Lean_Syntax_node2(x_83, x_120, x_139, x_93); +lean_inc(x_83); +x_141 = l_Lean_Syntax_node2(x_83, x_130, x_62, x_79); +lean_inc_n(x_93, 2); +lean_inc(x_83); +x_142 = l_Lean_Syntax_node3(x_83, x_91, x_93, x_93, x_141); +lean_inc(x_83); +x_143 = l_Lean_Syntax_node2(x_83, x_133, x_140, x_142); +lean_inc(x_93); +lean_inc(x_83); +x_144 = l_Lean_Syntax_node3(x_83, x_91, x_134, x_93, x_143); +x_145 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__30; +lean_inc(x_83); +x_146 = l_Lean_Syntax_node1(x_83, x_145, x_144); +x_147 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__8; +lean_inc(x_93); +lean_inc(x_83); +x_148 = l_Lean_Syntax_node3(x_83, x_147, x_66, x_146, x_93); +x_149 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__2; +lean_inc(x_93); +lean_inc(x_83); +x_150 = l_Lean_Syntax_node6(x_83, x_149, x_97, x_85, x_93, x_93, x_113, x_148); +x_151 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__6; +x_152 = l_Lean_Syntax_node2(x_83, x_151, x_95, x_150); +x_153 = lean_array_push(x_29, x_152); +lean_ctor_set(x_7, 1, x_153); +x_154 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_154, 0, x_7); +x_18 = x_154; +x_19 = x_88; +goto block_26; +} +else +{ +lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; size_t x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; +x_155 = lean_ctor_get(x_85, 0); +x_156 = lean_ctor_get(x_85, 1); +lean_inc(x_156); +lean_inc(x_155); +lean_dec(x_85); +x_157 = lean_ctor_get(x_155, 0); +lean_inc(x_157); +lean_dec(x_155); +x_158 = lean_environment_main_module(x_157); +x_159 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17; +x_160 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__9; +lean_inc(x_83); +x_161 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_161, 0, x_83); +lean_ctor_set(x_161, 1, x_159); +lean_ctor_set(x_161, 2, x_160); +x_162 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__8; +lean_inc_n(x_161, 6); +lean_inc(x_83); +x_163 = l_Lean_Syntax_node6(x_83, x_162, x_161, x_161, x_161, x_161, x_161, x_161); +x_164 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__4; +lean_inc(x_161); +lean_inc(x_83); +x_165 = l_Lean_Syntax_node1(x_83, x_164, x_161); +x_166 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__1; +lean_inc(x_83); +x_167 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_167, 0, x_83); +lean_ctor_set(x_167, 1, x_166); +x_168 = l_Array_append___rarg(x_160, x_70); +lean_dec(x_70); +lean_inc(x_83); +x_169 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_169, 0, x_83); +lean_ctor_set(x_169, 1, x_159); +lean_ctor_set(x_169, 2, x_168); +x_170 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__9; +lean_inc(x_83); +lean_ctor_set_tag(x_71, 2); +lean_ctor_set(x_71, 1, x_170); +lean_ctor_set(x_71, 0, x_83); +x_171 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__11; +lean_inc(x_84); +lean_inc(x_158); +x_172 = l_Lean_addMacroScope(x_158, x_171, x_84); +x_173 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__10; +x_174 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__24; +lean_inc(x_83); +x_175 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_175, 0, x_83); +lean_ctor_set(x_175, 1, x_173); +lean_ctor_set(x_175, 2, x_172); +lean_ctor_set(x_175, 3, x_174); +lean_inc(x_83); +x_176 = l_Lean_Syntax_node1(x_83, x_159, x_76); +x_177 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__5; +lean_inc(x_83); +x_178 = l_Lean_Syntax_node2(x_83, x_177, x_175, x_176); +x_179 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__8; +lean_inc(x_83); +x_180 = l_Lean_Syntax_node2(x_83, x_179, x_71, x_178); +x_181 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__6; +lean_inc(x_83); +x_182 = l_Lean_Syntax_node2(x_83, x_181, x_169, x_180); +x_183 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__9; +lean_inc(x_83); +lean_ctor_set_tag(x_66, 2); +lean_ctor_set(x_66, 1, x_183); +lean_ctor_set(x_66, 0, x_83); +x_184 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__3; +lean_inc(x_84); +lean_inc(x_158); +x_185 = l_Lean_addMacroScope(x_158, x_184, x_84); +x_186 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__2; +x_187 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__36; +lean_inc(x_83); +x_188 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_188, 0, x_83); +lean_ctor_set(x_188, 1, x_186); +lean_ctor_set(x_188, 2, x_185); +lean_ctor_set(x_188, 3, x_187); +x_189 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__34; +lean_inc(x_161); +lean_inc(x_83); +x_190 = l_Lean_Syntax_node2(x_83, x_189, x_188, x_161); +x_191 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__25; +lean_inc(x_83); +lean_ctor_set_tag(x_62, 2); +lean_ctor_set(x_62, 1, x_191); +lean_ctor_set(x_62, 0, x_83); +x_192 = lean_mk_syntax_ident(x_39); +x_193 = lean_array_size(x_68); +x_194 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__12; +x_195 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__2(x_194, x_193, x_61, x_68); +x_196 = l_Array_append___rarg(x_160, x_195); +lean_dec(x_195); +lean_inc(x_83); +x_197 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_197, 0, x_83); +lean_ctor_set(x_197, 1, x_159); +lean_ctor_set(x_197, 2, x_196); +lean_inc(x_83); +x_198 = l_Lean_Syntax_node2(x_83, x_177, x_192, x_197); +x_199 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__38; +lean_inc(x_62); +lean_inc(x_83); +x_200 = l_Lean_Syntax_node2(x_83, x_199, x_62, x_198); +lean_inc_n(x_161, 2); +lean_inc(x_83); +x_201 = l_Lean_Syntax_node3(x_83, x_159, x_161, x_161, x_200); +x_202 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__32; +lean_inc(x_83); +x_203 = l_Lean_Syntax_node2(x_83, x_202, x_190, x_201); +x_204 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__10; +x_205 = l_Lean_addMacroScope(x_158, x_204, x_84); +x_206 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__9; +x_207 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__40; +lean_inc(x_83); +x_208 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_208, 0, x_83); +lean_ctor_set(x_208, 1, x_206); +lean_ctor_set(x_208, 2, x_205); +lean_ctor_set(x_208, 3, x_207); +lean_inc(x_161); +lean_inc(x_83); +x_209 = l_Lean_Syntax_node2(x_83, x_189, x_208, x_161); +lean_inc(x_83); +x_210 = l_Lean_Syntax_node2(x_83, x_199, x_62, x_79); +lean_inc_n(x_161, 2); +lean_inc(x_83); +x_211 = l_Lean_Syntax_node3(x_83, x_159, x_161, x_161, x_210); +lean_inc(x_83); +x_212 = l_Lean_Syntax_node2(x_83, x_202, x_209, x_211); +lean_inc(x_161); +lean_inc(x_83); +x_213 = l_Lean_Syntax_node3(x_83, x_159, x_203, x_161, x_212); +x_214 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__30; +lean_inc(x_83); +x_215 = l_Lean_Syntax_node1(x_83, x_214, x_213); +x_216 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__8; +lean_inc(x_161); +lean_inc(x_83); +x_217 = l_Lean_Syntax_node3(x_83, x_216, x_66, x_215, x_161); +x_218 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__2; +lean_inc(x_161); +lean_inc(x_83); +x_219 = l_Lean_Syntax_node6(x_83, x_218, x_165, x_167, x_161, x_161, x_182, x_217); +x_220 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__6; +x_221 = l_Lean_Syntax_node2(x_83, x_220, x_163, x_219); +x_222 = lean_array_push(x_29, x_221); +lean_ctor_set(x_7, 1, x_222); +x_223 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_223, 0, x_7); +x_18 = x_223; +x_19 = x_156; +goto block_26; +} +} +else +{ +uint8_t x_224; +lean_dec(x_76); +lean_free_object(x_71); +lean_dec(x_70); +lean_free_object(x_66); +lean_dec(x_68); +lean_free_object(x_62); +lean_dec(x_28); +lean_dec(x_39); +lean_free_object(x_7); +lean_dec(x_29); +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_224 = !lean_is_exclusive(x_78); +if (x_224 == 0) +{ +return x_78; +} +else +{ +lean_object* x_225; lean_object* x_226; lean_object* x_227; +x_225 = lean_ctor_get(x_78, 0); +x_226 = lean_ctor_get(x_78, 1); +lean_inc(x_226); +lean_inc(x_225); +lean_dec(x_78); +x_227 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_227, 0, x_225); +lean_ctor_set(x_227, 1, x_226); +return x_227; +} +} +} +else +{ +uint8_t x_228; +lean_free_object(x_71); +lean_dec(x_70); +lean_free_object(x_66); +lean_dec(x_68); +lean_free_object(x_62); +lean_dec(x_47); +lean_dec(x_28); +lean_dec(x_39); +lean_free_object(x_7); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +x_228 = !lean_is_exclusive(x_75); +if (x_228 == 0) +{ +return x_75; +} +else +{ +lean_object* x_229; lean_object* x_230; lean_object* x_231; +x_229 = lean_ctor_get(x_75, 0); +x_230 = lean_ctor_get(x_75, 1); +lean_inc(x_230); +lean_inc(x_229); +lean_dec(x_75); +x_231 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_231, 0, x_229); +lean_ctor_set(x_231, 1, x_230); +return x_231; +} +} +} +else +{ +lean_object* x_232; lean_object* x_233; lean_object* x_234; +x_232 = lean_ctor_get(x_71, 0); +x_233 = lean_ctor_get(x_71, 1); +lean_inc(x_233); +lean_inc(x_232); +lean_dec(x_71); +lean_inc(x_8); +lean_inc(x_17); +x_234 = l_Lean_Elab_Deriving_ToExpr_updateIndType(x_17, x_232, x_8, x_9, x_10, x_11, x_12, x_13, x_233); +if (lean_obj_tag(x_234) == 0) +{ +lean_object* x_235; lean_object* x_236; lean_object* x_237; +x_235 = lean_ctor_get(x_234, 0); +lean_inc(x_235); +x_236 = lean_ctor_get(x_234, 1); +lean_inc(x_236); +lean_dec(x_234); +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_237 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr(x_17, x_47, x_8, x_9, x_10, x_11, x_12, x_13, x_236); +if (lean_obj_tag(x_237) == 0) +{ +lean_object* x_238; lean_object* x_239; lean_object* x_240; uint8_t x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; size_t x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; +x_238 = lean_ctor_get(x_237, 0); +lean_inc(x_238); +x_239 = lean_ctor_get(x_237, 1); +lean_inc(x_239); +lean_dec(x_237); +x_240 = lean_ctor_get(x_12, 5); +lean_inc(x_240); +x_241 = 0; +x_242 = l_Lean_SourceInfo_fromRef(x_240, x_241); +lean_dec(x_240); +x_243 = lean_ctor_get(x_12, 10); +lean_inc(x_243); +x_244 = lean_st_ref_get(x_13, x_239); +x_245 = lean_ctor_get(x_244, 0); +lean_inc(x_245); +x_246 = lean_ctor_get(x_244, 1); +lean_inc(x_246); +if (lean_is_exclusive(x_244)) { + lean_ctor_release(x_244, 0); + lean_ctor_release(x_244, 1); + x_247 = x_244; +} else { + lean_dec_ref(x_244); + x_247 = lean_box(0); +} +x_248 = lean_ctor_get(x_245, 0); +lean_inc(x_248); +lean_dec(x_245); +x_249 = lean_environment_main_module(x_248); +x_250 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17; +x_251 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__9; +lean_inc(x_242); +x_252 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_252, 0, x_242); +lean_ctor_set(x_252, 1, x_250); +lean_ctor_set(x_252, 2, x_251); +x_253 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__8; +lean_inc_n(x_252, 6); +lean_inc(x_242); +x_254 = l_Lean_Syntax_node6(x_242, x_253, x_252, x_252, x_252, x_252, x_252, x_252); +x_255 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__4; +lean_inc(x_252); +lean_inc(x_242); +x_256 = l_Lean_Syntax_node1(x_242, x_255, x_252); +x_257 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__1; +lean_inc(x_242); +if (lean_is_scalar(x_247)) { + x_258 = lean_alloc_ctor(2, 2, 0); +} else { + x_258 = x_247; + lean_ctor_set_tag(x_258, 2); +} +lean_ctor_set(x_258, 0, x_242); +lean_ctor_set(x_258, 1, x_257); +x_259 = l_Array_append___rarg(x_251, x_70); +lean_dec(x_70); +lean_inc(x_242); +x_260 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_260, 0, x_242); +lean_ctor_set(x_260, 1, x_250); +lean_ctor_set(x_260, 2, x_259); +x_261 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__9; +lean_inc(x_242); +x_262 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_262, 0, x_242); +lean_ctor_set(x_262, 1, x_261); +x_263 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__11; +lean_inc(x_243); +lean_inc(x_249); +x_264 = l_Lean_addMacroScope(x_249, x_263, x_243); +x_265 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__10; +x_266 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__24; +lean_inc(x_242); +x_267 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_267, 0, x_242); +lean_ctor_set(x_267, 1, x_265); +lean_ctor_set(x_267, 2, x_264); +lean_ctor_set(x_267, 3, x_266); +lean_inc(x_242); +x_268 = l_Lean_Syntax_node1(x_242, x_250, x_235); +x_269 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__5; +lean_inc(x_242); +x_270 = l_Lean_Syntax_node2(x_242, x_269, x_267, x_268); +x_271 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__8; +lean_inc(x_242); +x_272 = l_Lean_Syntax_node2(x_242, x_271, x_262, x_270); +x_273 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__6; +lean_inc(x_242); +x_274 = l_Lean_Syntax_node2(x_242, x_273, x_260, x_272); +x_275 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__9; +lean_inc(x_242); +lean_ctor_set_tag(x_66, 2); +lean_ctor_set(x_66, 1, x_275); +lean_ctor_set(x_66, 0, x_242); +x_276 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__3; +lean_inc(x_243); +lean_inc(x_249); +x_277 = l_Lean_addMacroScope(x_249, x_276, x_243); +x_278 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__2; +x_279 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__36; +lean_inc(x_242); +x_280 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_280, 0, x_242); +lean_ctor_set(x_280, 1, x_278); +lean_ctor_set(x_280, 2, x_277); +lean_ctor_set(x_280, 3, x_279); +x_281 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__34; +lean_inc(x_252); +lean_inc(x_242); +x_282 = l_Lean_Syntax_node2(x_242, x_281, x_280, x_252); +x_283 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__25; +lean_inc(x_242); +lean_ctor_set_tag(x_62, 2); +lean_ctor_set(x_62, 1, x_283); +lean_ctor_set(x_62, 0, x_242); +x_284 = lean_mk_syntax_ident(x_39); +x_285 = lean_array_size(x_68); +x_286 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__12; +x_287 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__2(x_286, x_285, x_61, x_68); +x_288 = l_Array_append___rarg(x_251, x_287); +lean_dec(x_287); +lean_inc(x_242); +x_289 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_289, 0, x_242); +lean_ctor_set(x_289, 1, x_250); +lean_ctor_set(x_289, 2, x_288); +lean_inc(x_242); +x_290 = l_Lean_Syntax_node2(x_242, x_269, x_284, x_289); +x_291 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__38; +lean_inc(x_62); +lean_inc(x_242); +x_292 = l_Lean_Syntax_node2(x_242, x_291, x_62, x_290); +lean_inc_n(x_252, 2); +lean_inc(x_242); +x_293 = l_Lean_Syntax_node3(x_242, x_250, x_252, x_252, x_292); +x_294 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__32; +lean_inc(x_242); +x_295 = l_Lean_Syntax_node2(x_242, x_294, x_282, x_293); +x_296 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__10; +x_297 = l_Lean_addMacroScope(x_249, x_296, x_243); +x_298 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__9; +x_299 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__40; +lean_inc(x_242); +x_300 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_300, 0, x_242); +lean_ctor_set(x_300, 1, x_298); +lean_ctor_set(x_300, 2, x_297); +lean_ctor_set(x_300, 3, x_299); +lean_inc(x_252); +lean_inc(x_242); +x_301 = l_Lean_Syntax_node2(x_242, x_281, x_300, x_252); +lean_inc(x_242); +x_302 = l_Lean_Syntax_node2(x_242, x_291, x_62, x_238); +lean_inc_n(x_252, 2); +lean_inc(x_242); +x_303 = l_Lean_Syntax_node3(x_242, x_250, x_252, x_252, x_302); +lean_inc(x_242); +x_304 = l_Lean_Syntax_node2(x_242, x_294, x_301, x_303); +lean_inc(x_252); +lean_inc(x_242); +x_305 = l_Lean_Syntax_node3(x_242, x_250, x_295, x_252, x_304); +x_306 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__30; +lean_inc(x_242); +x_307 = l_Lean_Syntax_node1(x_242, x_306, x_305); +x_308 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__8; +lean_inc(x_252); +lean_inc(x_242); +x_309 = l_Lean_Syntax_node3(x_242, x_308, x_66, x_307, x_252); +x_310 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__2; +lean_inc(x_252); +lean_inc(x_242); +x_311 = l_Lean_Syntax_node6(x_242, x_310, x_256, x_258, x_252, x_252, x_274, x_309); +x_312 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__6; +x_313 = l_Lean_Syntax_node2(x_242, x_312, x_254, x_311); +x_314 = lean_array_push(x_29, x_313); +lean_ctor_set(x_7, 1, x_314); +x_315 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_315, 0, x_7); +x_18 = x_315; +x_19 = x_246; +goto block_26; +} +else +{ +lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; +lean_dec(x_235); +lean_dec(x_70); +lean_free_object(x_66); +lean_dec(x_68); +lean_free_object(x_62); +lean_dec(x_28); +lean_dec(x_39); +lean_free_object(x_7); +lean_dec(x_29); +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_316 = lean_ctor_get(x_237, 0); +lean_inc(x_316); +x_317 = lean_ctor_get(x_237, 1); +lean_inc(x_317); +if (lean_is_exclusive(x_237)) { + lean_ctor_release(x_237, 0); + lean_ctor_release(x_237, 1); + x_318 = x_237; +} else { + lean_dec_ref(x_237); + x_318 = lean_box(0); +} +if (lean_is_scalar(x_318)) { + x_319 = lean_alloc_ctor(1, 2, 0); +} else { + x_319 = x_318; +} +lean_ctor_set(x_319, 0, x_316); +lean_ctor_set(x_319, 1, x_317); +return x_319; +} +} +else +{ +lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; +lean_dec(x_70); +lean_free_object(x_66); +lean_dec(x_68); +lean_free_object(x_62); +lean_dec(x_47); +lean_dec(x_28); +lean_dec(x_39); +lean_free_object(x_7); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +x_320 = lean_ctor_get(x_234, 0); +lean_inc(x_320); +x_321 = lean_ctor_get(x_234, 1); +lean_inc(x_321); +if (lean_is_exclusive(x_234)) { + lean_ctor_release(x_234, 0); + lean_ctor_release(x_234, 1); + x_322 = x_234; +} else { + lean_dec_ref(x_234); + x_322 = lean_box(0); +} +if (lean_is_scalar(x_322)) { + x_323 = lean_alloc_ctor(1, 2, 0); +} else { + x_323 = x_322; +} +lean_ctor_set(x_323, 0, x_320); +lean_ctor_set(x_323, 1, x_321); +return x_323; +} +} +} +else +{ +lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; +x_324 = lean_ctor_get(x_66, 0); +x_325 = lean_ctor_get(x_66, 1); +lean_inc(x_325); +lean_inc(x_324); +lean_dec(x_66); +x_326 = l_Array_append___rarg(x_57, x_325); +lean_dec(x_325); +lean_inc(x_47); +lean_inc(x_17); +x_327 = l_Lean_Elab_Deriving_mkInductiveApp(x_17, x_47, x_8, x_9, x_10, x_11, x_12, x_13, x_65); +x_328 = lean_ctor_get(x_327, 0); +lean_inc(x_328); +x_329 = lean_ctor_get(x_327, 1); +lean_inc(x_329); +if (lean_is_exclusive(x_327)) { + lean_ctor_release(x_327, 0); + lean_ctor_release(x_327, 1); + x_330 = x_327; +} else { + lean_dec_ref(x_327); + x_330 = lean_box(0); +} +lean_inc(x_8); +lean_inc(x_17); +x_331 = l_Lean_Elab_Deriving_ToExpr_updateIndType(x_17, x_328, x_8, x_9, x_10, x_11, x_12, x_13, x_329); +if (lean_obj_tag(x_331) == 0) +{ +lean_object* x_332; lean_object* x_333; lean_object* x_334; +x_332 = lean_ctor_get(x_331, 0); +lean_inc(x_332); +x_333 = lean_ctor_get(x_331, 1); +lean_inc(x_333); +lean_dec(x_331); +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_334 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr(x_17, x_47, x_8, x_9, x_10, x_11, x_12, x_13, x_333); +if (lean_obj_tag(x_334) == 0) +{ +lean_object* x_335; lean_object* x_336; lean_object* x_337; uint8_t x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; lean_object* x_353; lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; lean_object* x_369; lean_object* x_370; lean_object* x_371; lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; lean_object* x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; lean_object* x_382; size_t x_383; lean_object* x_384; lean_object* x_385; lean_object* x_386; lean_object* x_387; lean_object* x_388; lean_object* x_389; lean_object* x_390; lean_object* x_391; lean_object* x_392; lean_object* x_393; lean_object* x_394; lean_object* x_395; lean_object* x_396; lean_object* x_397; lean_object* x_398; lean_object* x_399; lean_object* x_400; lean_object* x_401; lean_object* x_402; lean_object* x_403; lean_object* x_404; lean_object* x_405; lean_object* x_406; lean_object* x_407; lean_object* x_408; lean_object* x_409; lean_object* x_410; lean_object* x_411; lean_object* x_412; lean_object* x_413; +x_335 = lean_ctor_get(x_334, 0); +lean_inc(x_335); +x_336 = lean_ctor_get(x_334, 1); +lean_inc(x_336); +lean_dec(x_334); +x_337 = lean_ctor_get(x_12, 5); +lean_inc(x_337); +x_338 = 0; +x_339 = l_Lean_SourceInfo_fromRef(x_337, x_338); +lean_dec(x_337); +x_340 = lean_ctor_get(x_12, 10); +lean_inc(x_340); +x_341 = lean_st_ref_get(x_13, x_336); +x_342 = lean_ctor_get(x_341, 0); +lean_inc(x_342); +x_343 = lean_ctor_get(x_341, 1); +lean_inc(x_343); +if (lean_is_exclusive(x_341)) { + lean_ctor_release(x_341, 0); + lean_ctor_release(x_341, 1); + x_344 = x_341; +} else { + lean_dec_ref(x_341); + x_344 = lean_box(0); +} +x_345 = lean_ctor_get(x_342, 0); +lean_inc(x_345); +lean_dec(x_342); +x_346 = lean_environment_main_module(x_345); +x_347 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17; +x_348 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__9; +lean_inc(x_339); +x_349 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_349, 0, x_339); +lean_ctor_set(x_349, 1, x_347); +lean_ctor_set(x_349, 2, x_348); +x_350 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__8; +lean_inc_n(x_349, 6); +lean_inc(x_339); +x_351 = l_Lean_Syntax_node6(x_339, x_350, x_349, x_349, x_349, x_349, x_349, x_349); +x_352 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__4; +lean_inc(x_349); +lean_inc(x_339); +x_353 = l_Lean_Syntax_node1(x_339, x_352, x_349); +x_354 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__1; +lean_inc(x_339); +if (lean_is_scalar(x_344)) { + x_355 = lean_alloc_ctor(2, 2, 0); +} else { + x_355 = x_344; + lean_ctor_set_tag(x_355, 2); +} +lean_ctor_set(x_355, 0, x_339); +lean_ctor_set(x_355, 1, x_354); +x_356 = l_Array_append___rarg(x_348, x_326); +lean_dec(x_326); +lean_inc(x_339); +x_357 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_357, 0, x_339); +lean_ctor_set(x_357, 1, x_347); +lean_ctor_set(x_357, 2, x_356); +x_358 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__9; +lean_inc(x_339); +if (lean_is_scalar(x_330)) { + x_359 = lean_alloc_ctor(2, 2, 0); +} else { + x_359 = x_330; + lean_ctor_set_tag(x_359, 2); +} +lean_ctor_set(x_359, 0, x_339); +lean_ctor_set(x_359, 1, x_358); +x_360 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__11; +lean_inc(x_340); +lean_inc(x_346); +x_361 = l_Lean_addMacroScope(x_346, x_360, x_340); +x_362 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__10; +x_363 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__24; +lean_inc(x_339); +x_364 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_364, 0, x_339); +lean_ctor_set(x_364, 1, x_362); +lean_ctor_set(x_364, 2, x_361); +lean_ctor_set(x_364, 3, x_363); +lean_inc(x_339); +x_365 = l_Lean_Syntax_node1(x_339, x_347, x_332); +x_366 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__5; +lean_inc(x_339); +x_367 = l_Lean_Syntax_node2(x_339, x_366, x_364, x_365); +x_368 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__8; +lean_inc(x_339); +x_369 = l_Lean_Syntax_node2(x_339, x_368, x_359, x_367); +x_370 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__6; +lean_inc(x_339); +x_371 = l_Lean_Syntax_node2(x_339, x_370, x_357, x_369); +x_372 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__9; +lean_inc(x_339); +x_373 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_373, 0, x_339); +lean_ctor_set(x_373, 1, x_372); +x_374 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__3; +lean_inc(x_340); +lean_inc(x_346); +x_375 = l_Lean_addMacroScope(x_346, x_374, x_340); +x_376 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__2; +x_377 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__36; +lean_inc(x_339); +x_378 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_378, 0, x_339); +lean_ctor_set(x_378, 1, x_376); +lean_ctor_set(x_378, 2, x_375); +lean_ctor_set(x_378, 3, x_377); +x_379 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__34; +lean_inc(x_349); +lean_inc(x_339); +x_380 = l_Lean_Syntax_node2(x_339, x_379, x_378, x_349); +x_381 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__25; +lean_inc(x_339); +lean_ctor_set_tag(x_62, 2); +lean_ctor_set(x_62, 1, x_381); +lean_ctor_set(x_62, 0, x_339); +x_382 = lean_mk_syntax_ident(x_39); +x_383 = lean_array_size(x_324); +x_384 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__12; +x_385 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__2(x_384, x_383, x_61, x_324); +x_386 = l_Array_append___rarg(x_348, x_385); +lean_dec(x_385); +lean_inc(x_339); +x_387 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_387, 0, x_339); +lean_ctor_set(x_387, 1, x_347); +lean_ctor_set(x_387, 2, x_386); +lean_inc(x_339); +x_388 = l_Lean_Syntax_node2(x_339, x_366, x_382, x_387); +x_389 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__38; +lean_inc(x_62); +lean_inc(x_339); +x_390 = l_Lean_Syntax_node2(x_339, x_389, x_62, x_388); +lean_inc_n(x_349, 2); +lean_inc(x_339); +x_391 = l_Lean_Syntax_node3(x_339, x_347, x_349, x_349, x_390); +x_392 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__32; +lean_inc(x_339); +x_393 = l_Lean_Syntax_node2(x_339, x_392, x_380, x_391); +x_394 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__10; +x_395 = l_Lean_addMacroScope(x_346, x_394, x_340); +x_396 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__9; +x_397 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__40; +lean_inc(x_339); +x_398 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_398, 0, x_339); +lean_ctor_set(x_398, 1, x_396); +lean_ctor_set(x_398, 2, x_395); +lean_ctor_set(x_398, 3, x_397); +lean_inc(x_349); +lean_inc(x_339); +x_399 = l_Lean_Syntax_node2(x_339, x_379, x_398, x_349); +lean_inc(x_339); +x_400 = l_Lean_Syntax_node2(x_339, x_389, x_62, x_335); +lean_inc_n(x_349, 2); +lean_inc(x_339); +x_401 = l_Lean_Syntax_node3(x_339, x_347, x_349, x_349, x_400); +lean_inc(x_339); +x_402 = l_Lean_Syntax_node2(x_339, x_392, x_399, x_401); +lean_inc(x_349); +lean_inc(x_339); +x_403 = l_Lean_Syntax_node3(x_339, x_347, x_393, x_349, x_402); +x_404 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__30; +lean_inc(x_339); +x_405 = l_Lean_Syntax_node1(x_339, x_404, x_403); +x_406 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__8; +lean_inc(x_349); +lean_inc(x_339); +x_407 = l_Lean_Syntax_node3(x_339, x_406, x_373, x_405, x_349); +x_408 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__2; +lean_inc(x_349); +lean_inc(x_339); +x_409 = l_Lean_Syntax_node6(x_339, x_408, x_353, x_355, x_349, x_349, x_371, x_407); +x_410 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__6; +x_411 = l_Lean_Syntax_node2(x_339, x_410, x_351, x_409); +x_412 = lean_array_push(x_29, x_411); +lean_ctor_set(x_7, 1, x_412); +x_413 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_413, 0, x_7); +x_18 = x_413; +x_19 = x_343; +goto block_26; +} +else +{ +lean_object* x_414; lean_object* x_415; lean_object* x_416; lean_object* x_417; +lean_dec(x_332); +lean_dec(x_330); +lean_dec(x_326); +lean_dec(x_324); +lean_free_object(x_62); +lean_dec(x_28); +lean_dec(x_39); +lean_free_object(x_7); +lean_dec(x_29); +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_414 = lean_ctor_get(x_334, 0); +lean_inc(x_414); +x_415 = lean_ctor_get(x_334, 1); +lean_inc(x_415); +if (lean_is_exclusive(x_334)) { + lean_ctor_release(x_334, 0); + lean_ctor_release(x_334, 1); + x_416 = x_334; +} else { + lean_dec_ref(x_334); + x_416 = lean_box(0); +} +if (lean_is_scalar(x_416)) { + x_417 = lean_alloc_ctor(1, 2, 0); +} else { + x_417 = x_416; +} +lean_ctor_set(x_417, 0, x_414); +lean_ctor_set(x_417, 1, x_415); +return x_417; +} +} +else +{ +lean_object* x_418; lean_object* x_419; lean_object* x_420; lean_object* x_421; +lean_dec(x_330); +lean_dec(x_326); +lean_dec(x_324); +lean_free_object(x_62); +lean_dec(x_47); +lean_dec(x_28); +lean_dec(x_39); +lean_free_object(x_7); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +x_418 = lean_ctor_get(x_331, 0); +lean_inc(x_418); +x_419 = lean_ctor_get(x_331, 1); +lean_inc(x_419); +if (lean_is_exclusive(x_331)) { + lean_ctor_release(x_331, 0); + lean_ctor_release(x_331, 1); + x_420 = x_331; +} else { + lean_dec_ref(x_331); + x_420 = lean_box(0); +} +if (lean_is_scalar(x_420)) { + x_421 = lean_alloc_ctor(1, 2, 0); +} else { + x_421 = x_420; +} +lean_ctor_set(x_421, 0, x_418); +lean_ctor_set(x_421, 1, x_419); +return x_421; +} +} +} +else +{ +lean_object* x_422; lean_object* x_423; lean_object* x_424; lean_object* x_425; lean_object* x_426; lean_object* x_427; lean_object* x_428; lean_object* x_429; lean_object* x_430; lean_object* x_431; lean_object* x_432; lean_object* x_433; +x_422 = lean_ctor_get(x_62, 0); +x_423 = lean_ctor_get(x_62, 1); +lean_inc(x_423); +lean_inc(x_422); +lean_dec(x_62); +x_424 = l_Array_unzip___rarg(x_422); +lean_dec(x_422); +x_425 = lean_ctor_get(x_424, 0); +lean_inc(x_425); +x_426 = lean_ctor_get(x_424, 1); +lean_inc(x_426); +if (lean_is_exclusive(x_424)) { + lean_ctor_release(x_424, 0); + lean_ctor_release(x_424, 1); + x_427 = x_424; +} else { + lean_dec_ref(x_424); + x_427 = lean_box(0); +} +x_428 = l_Array_append___rarg(x_57, x_426); +lean_dec(x_426); +lean_inc(x_47); +lean_inc(x_17); +x_429 = l_Lean_Elab_Deriving_mkInductiveApp(x_17, x_47, x_8, x_9, x_10, x_11, x_12, x_13, x_423); +x_430 = lean_ctor_get(x_429, 0); +lean_inc(x_430); +x_431 = lean_ctor_get(x_429, 1); +lean_inc(x_431); +if (lean_is_exclusive(x_429)) { + lean_ctor_release(x_429, 0); + lean_ctor_release(x_429, 1); + x_432 = x_429; +} else { + lean_dec_ref(x_429); + x_432 = lean_box(0); +} +lean_inc(x_8); +lean_inc(x_17); +x_433 = l_Lean_Elab_Deriving_ToExpr_updateIndType(x_17, x_430, x_8, x_9, x_10, x_11, x_12, x_13, x_431); +if (lean_obj_tag(x_433) == 0) +{ +lean_object* x_434; lean_object* x_435; lean_object* x_436; +x_434 = lean_ctor_get(x_433, 0); +lean_inc(x_434); +x_435 = lean_ctor_get(x_433, 1); +lean_inc(x_435); +lean_dec(x_433); +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_436 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr(x_17, x_47, x_8, x_9, x_10, x_11, x_12, x_13, x_435); +if (lean_obj_tag(x_436) == 0) +{ +lean_object* x_437; lean_object* x_438; lean_object* x_439; uint8_t x_440; lean_object* x_441; lean_object* x_442; lean_object* x_443; lean_object* x_444; lean_object* x_445; lean_object* x_446; lean_object* x_447; lean_object* x_448; lean_object* x_449; lean_object* x_450; lean_object* x_451; lean_object* x_452; lean_object* x_453; lean_object* x_454; lean_object* x_455; lean_object* x_456; lean_object* x_457; lean_object* x_458; lean_object* x_459; lean_object* x_460; lean_object* x_461; lean_object* x_462; lean_object* x_463; lean_object* x_464; lean_object* x_465; lean_object* x_466; lean_object* x_467; lean_object* x_468; lean_object* x_469; lean_object* x_470; lean_object* x_471; lean_object* x_472; lean_object* x_473; lean_object* x_474; lean_object* x_475; lean_object* x_476; lean_object* x_477; lean_object* x_478; lean_object* x_479; lean_object* x_480; lean_object* x_481; lean_object* x_482; lean_object* x_483; lean_object* x_484; lean_object* x_485; size_t x_486; lean_object* x_487; lean_object* x_488; lean_object* x_489; lean_object* x_490; lean_object* x_491; lean_object* x_492; lean_object* x_493; lean_object* x_494; lean_object* x_495; lean_object* x_496; lean_object* x_497; lean_object* x_498; lean_object* x_499; lean_object* x_500; lean_object* x_501; lean_object* x_502; lean_object* x_503; lean_object* x_504; lean_object* x_505; lean_object* x_506; lean_object* x_507; lean_object* x_508; lean_object* x_509; lean_object* x_510; lean_object* x_511; lean_object* x_512; lean_object* x_513; lean_object* x_514; lean_object* x_515; lean_object* x_516; +x_437 = lean_ctor_get(x_436, 0); +lean_inc(x_437); +x_438 = lean_ctor_get(x_436, 1); +lean_inc(x_438); +lean_dec(x_436); +x_439 = lean_ctor_get(x_12, 5); +lean_inc(x_439); +x_440 = 0; +x_441 = l_Lean_SourceInfo_fromRef(x_439, x_440); +lean_dec(x_439); +x_442 = lean_ctor_get(x_12, 10); +lean_inc(x_442); +x_443 = lean_st_ref_get(x_13, x_438); +x_444 = lean_ctor_get(x_443, 0); +lean_inc(x_444); +x_445 = lean_ctor_get(x_443, 1); +lean_inc(x_445); +if (lean_is_exclusive(x_443)) { + lean_ctor_release(x_443, 0); + lean_ctor_release(x_443, 1); + x_446 = x_443; +} else { + lean_dec_ref(x_443); + x_446 = lean_box(0); +} +x_447 = lean_ctor_get(x_444, 0); +lean_inc(x_447); +lean_dec(x_444); +x_448 = lean_environment_main_module(x_447); +x_449 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17; +x_450 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__9; +lean_inc(x_441); +x_451 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_451, 0, x_441); +lean_ctor_set(x_451, 1, x_449); +lean_ctor_set(x_451, 2, x_450); +x_452 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__8; +lean_inc_n(x_451, 6); +lean_inc(x_441); +x_453 = l_Lean_Syntax_node6(x_441, x_452, x_451, x_451, x_451, x_451, x_451, x_451); +x_454 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__4; +lean_inc(x_451); +lean_inc(x_441); +x_455 = l_Lean_Syntax_node1(x_441, x_454, x_451); +x_456 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__1; +lean_inc(x_441); +if (lean_is_scalar(x_446)) { + x_457 = lean_alloc_ctor(2, 2, 0); +} else { + x_457 = x_446; + lean_ctor_set_tag(x_457, 2); +} +lean_ctor_set(x_457, 0, x_441); +lean_ctor_set(x_457, 1, x_456); +x_458 = l_Array_append___rarg(x_450, x_428); +lean_dec(x_428); +lean_inc(x_441); +x_459 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_459, 0, x_441); +lean_ctor_set(x_459, 1, x_449); +lean_ctor_set(x_459, 2, x_458); +x_460 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__9; +lean_inc(x_441); +if (lean_is_scalar(x_432)) { + x_461 = lean_alloc_ctor(2, 2, 0); +} else { + x_461 = x_432; + lean_ctor_set_tag(x_461, 2); +} +lean_ctor_set(x_461, 0, x_441); +lean_ctor_set(x_461, 1, x_460); +x_462 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__11; +lean_inc(x_442); +lean_inc(x_448); +x_463 = l_Lean_addMacroScope(x_448, x_462, x_442); +x_464 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__10; +x_465 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__24; +lean_inc(x_441); +x_466 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_466, 0, x_441); +lean_ctor_set(x_466, 1, x_464); +lean_ctor_set(x_466, 2, x_463); +lean_ctor_set(x_466, 3, x_465); +lean_inc(x_441); +x_467 = l_Lean_Syntax_node1(x_441, x_449, x_434); +x_468 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__5; +lean_inc(x_441); +x_469 = l_Lean_Syntax_node2(x_441, x_468, x_466, x_467); +x_470 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__8; +lean_inc(x_441); +x_471 = l_Lean_Syntax_node2(x_441, x_470, x_461, x_469); +x_472 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__6; +lean_inc(x_441); +x_473 = l_Lean_Syntax_node2(x_441, x_472, x_459, x_471); +x_474 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__9; +lean_inc(x_441); +if (lean_is_scalar(x_427)) { + x_475 = lean_alloc_ctor(2, 2, 0); +} else { + x_475 = x_427; + lean_ctor_set_tag(x_475, 2); +} +lean_ctor_set(x_475, 0, x_441); +lean_ctor_set(x_475, 1, x_474); +x_476 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__3; +lean_inc(x_442); +lean_inc(x_448); +x_477 = l_Lean_addMacroScope(x_448, x_476, x_442); +x_478 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__2; +x_479 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__36; +lean_inc(x_441); +x_480 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_480, 0, x_441); +lean_ctor_set(x_480, 1, x_478); +lean_ctor_set(x_480, 2, x_477); +lean_ctor_set(x_480, 3, x_479); +x_481 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__34; +lean_inc(x_451); +lean_inc(x_441); +x_482 = l_Lean_Syntax_node2(x_441, x_481, x_480, x_451); +x_483 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__25; +lean_inc(x_441); +x_484 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_484, 0, x_441); +lean_ctor_set(x_484, 1, x_483); +x_485 = lean_mk_syntax_ident(x_39); +x_486 = lean_array_size(x_425); +x_487 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__12; +x_488 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__2(x_487, x_486, x_61, x_425); +x_489 = l_Array_append___rarg(x_450, x_488); +lean_dec(x_488); +lean_inc(x_441); +x_490 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_490, 0, x_441); +lean_ctor_set(x_490, 1, x_449); +lean_ctor_set(x_490, 2, x_489); +lean_inc(x_441); +x_491 = l_Lean_Syntax_node2(x_441, x_468, x_485, x_490); +x_492 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__38; +lean_inc(x_484); +lean_inc(x_441); +x_493 = l_Lean_Syntax_node2(x_441, x_492, x_484, x_491); +lean_inc_n(x_451, 2); +lean_inc(x_441); +x_494 = l_Lean_Syntax_node3(x_441, x_449, x_451, x_451, x_493); +x_495 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__32; +lean_inc(x_441); +x_496 = l_Lean_Syntax_node2(x_441, x_495, x_482, x_494); +x_497 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__10; +x_498 = l_Lean_addMacroScope(x_448, x_497, x_442); +x_499 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__9; +x_500 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__40; +lean_inc(x_441); +x_501 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_501, 0, x_441); +lean_ctor_set(x_501, 1, x_499); +lean_ctor_set(x_501, 2, x_498); +lean_ctor_set(x_501, 3, x_500); +lean_inc(x_451); +lean_inc(x_441); +x_502 = l_Lean_Syntax_node2(x_441, x_481, x_501, x_451); +lean_inc(x_441); +x_503 = l_Lean_Syntax_node2(x_441, x_492, x_484, x_437); +lean_inc_n(x_451, 2); +lean_inc(x_441); +x_504 = l_Lean_Syntax_node3(x_441, x_449, x_451, x_451, x_503); +lean_inc(x_441); +x_505 = l_Lean_Syntax_node2(x_441, x_495, x_502, x_504); +lean_inc(x_451); +lean_inc(x_441); +x_506 = l_Lean_Syntax_node3(x_441, x_449, x_496, x_451, x_505); +x_507 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__30; +lean_inc(x_441); +x_508 = l_Lean_Syntax_node1(x_441, x_507, x_506); +x_509 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__8; +lean_inc(x_451); +lean_inc(x_441); +x_510 = l_Lean_Syntax_node3(x_441, x_509, x_475, x_508, x_451); +x_511 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__2; +lean_inc(x_451); +lean_inc(x_441); +x_512 = l_Lean_Syntax_node6(x_441, x_511, x_455, x_457, x_451, x_451, x_473, x_510); +x_513 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__6; +x_514 = l_Lean_Syntax_node2(x_441, x_513, x_453, x_512); +x_515 = lean_array_push(x_29, x_514); +lean_ctor_set(x_7, 1, x_515); +x_516 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_516, 0, x_7); +x_18 = x_516; +x_19 = x_445; +goto block_26; +} +else +{ +lean_object* x_517; lean_object* x_518; lean_object* x_519; lean_object* x_520; +lean_dec(x_434); +lean_dec(x_432); +lean_dec(x_428); +lean_dec(x_427); +lean_dec(x_425); +lean_dec(x_28); +lean_dec(x_39); +lean_free_object(x_7); +lean_dec(x_29); +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_517 = lean_ctor_get(x_436, 0); +lean_inc(x_517); +x_518 = lean_ctor_get(x_436, 1); +lean_inc(x_518); +if (lean_is_exclusive(x_436)) { + lean_ctor_release(x_436, 0); + lean_ctor_release(x_436, 1); + x_519 = x_436; +} else { + lean_dec_ref(x_436); + x_519 = lean_box(0); +} +if (lean_is_scalar(x_519)) { + x_520 = lean_alloc_ctor(1, 2, 0); +} else { + x_520 = x_519; +} +lean_ctor_set(x_520, 0, x_517); +lean_ctor_set(x_520, 1, x_518); +return x_520; +} +} +else +{ +lean_object* x_521; lean_object* x_522; lean_object* x_523; lean_object* x_524; +lean_dec(x_432); +lean_dec(x_428); +lean_dec(x_427); +lean_dec(x_425); +lean_dec(x_47); +lean_dec(x_28); +lean_dec(x_39); +lean_free_object(x_7); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +x_521 = lean_ctor_get(x_433, 0); +lean_inc(x_521); +x_522 = lean_ctor_get(x_433, 1); +lean_inc(x_522); +if (lean_is_exclusive(x_433)) { + lean_ctor_release(x_433, 0); + lean_ctor_release(x_433, 1); + x_523 = x_433; +} else { + lean_dec_ref(x_433); + x_523 = lean_box(0); +} +if (lean_is_scalar(x_523)) { + x_524 = lean_alloc_ctor(1, 2, 0); +} else { + x_524 = x_523; +} +lean_ctor_set(x_524, 0, x_521); +lean_ctor_set(x_524, 1, x_522); +return x_524; +} +} +} +else +{ +uint8_t x_525; +lean_dec(x_50); +lean_dec(x_47); +lean_dec(x_42); +lean_dec(x_28); +lean_dec(x_39); +lean_free_object(x_7); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +x_525 = !lean_is_exclusive(x_53); +if (x_525 == 0) +{ +return x_53; +} +else +{ +lean_object* x_526; lean_object* x_527; lean_object* x_528; +x_526 = lean_ctor_get(x_53, 0); +x_527 = lean_ctor_get(x_53, 1); +lean_inc(x_527); +lean_inc(x_526); +lean_dec(x_53); +x_528 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_528, 0, x_526); +lean_ctor_set(x_528, 1, x_527); +return x_528; +} +} +} +else +{ +uint8_t x_529; +lean_dec(x_42); +lean_dec(x_28); +lean_dec(x_39); +lean_free_object(x_7); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +x_529 = !lean_is_exclusive(x_46); +if (x_529 == 0) +{ +return x_46; +} +else +{ +lean_object* x_530; lean_object* x_531; lean_object* x_532; +x_530 = lean_ctor_get(x_46, 0); +x_531 = lean_ctor_get(x_46, 1); +lean_inc(x_531); +lean_inc(x_530); +lean_dec(x_46); +x_532 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_532, 0, x_530); +lean_ctor_set(x_532, 1, x_531); +return x_532; +} +} +} +} +else +{ +lean_object* x_533; lean_object* x_534; lean_object* x_535; lean_object* x_536; lean_object* x_537; lean_object* x_538; uint8_t x_539; +lean_dec(x_28); +x_533 = lean_array_fget(x_30, x_31); +x_534 = lean_unsigned_to_nat(1u); +x_535 = lean_nat_add(x_31, x_534); +lean_dec(x_31); +x_536 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_536, 0, x_30); +lean_ctor_set(x_536, 1, x_535); +lean_ctor_set(x_536, 2, x_32); +x_537 = lean_ctor_get(x_17, 0); +lean_inc(x_537); +x_538 = lean_ctor_get(x_537, 0); +lean_inc(x_538); +x_539 = l_Array_contains___at_Lean_registerInternalExceptionId___spec__1(x_1, x_538); +lean_dec(x_538); +if (x_539 == 0) +{ +lean_object* x_540; +lean_dec(x_537); +lean_dec(x_533); +lean_dec(x_17); +lean_ctor_set(x_7, 0, x_536); +x_540 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_540, 0, x_7); +x_18 = x_540; +x_19 = x_14; +goto block_26; +} +else +{ +lean_object* x_541; +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_17); +x_541 = l_Lean_Elab_Deriving_mkInductArgNames(x_17, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_541) == 0) +{ +lean_object* x_542; lean_object* x_543; lean_object* x_544; lean_object* x_545; lean_object* x_546; lean_object* x_547; lean_object* x_548; +x_542 = lean_ctor_get(x_541, 0); +lean_inc(x_542); +x_543 = lean_ctor_get(x_541, 1); +lean_inc(x_543); +lean_dec(x_541); +lean_inc(x_542); +x_544 = l_Lean_Elab_Deriving_mkImplicitBinders(x_542, x_8, x_9, x_10, x_11, x_12, x_13, x_543); +x_545 = lean_ctor_get(x_544, 0); +lean_inc(x_545); +x_546 = lean_ctor_get(x_544, 1); +lean_inc(x_546); +lean_dec(x_544); +x_547 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__12; +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_542); +lean_inc(x_17); +x_548 = l_Lean_Elab_Deriving_mkInstImplicitBinders(x_547, x_17, x_542, x_8, x_9, x_10, x_11, x_12, x_13, x_546); +if (lean_obj_tag(x_548) == 0) +{ +lean_object* x_549; lean_object* x_550; lean_object* x_551; lean_object* x_552; lean_object* x_553; lean_object* x_554; size_t x_555; size_t x_556; lean_object* x_557; lean_object* x_558; lean_object* x_559; lean_object* x_560; lean_object* x_561; lean_object* x_562; lean_object* x_563; lean_object* x_564; lean_object* x_565; lean_object* x_566; lean_object* x_567; lean_object* x_568; lean_object* x_569; lean_object* x_570; +x_549 = lean_ctor_get(x_548, 0); +lean_inc(x_549); +x_550 = lean_ctor_get(x_548, 1); +lean_inc(x_550); +lean_dec(x_548); +x_551 = lean_box(0); +x_552 = l_Array_append___rarg(x_545, x_549); +lean_dec(x_549); +x_553 = lean_ctor_get(x_537, 1); +lean_inc(x_553); +lean_dec(x_537); +x_554 = lean_array_mk(x_553); +x_555 = lean_array_size(x_554); +x_556 = 0; +lean_inc(x_12); +x_557 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__1(x_551, x_555, x_556, x_554, x_8, x_9, x_10, x_11, x_12, x_13, x_550); +x_558 = lean_ctor_get(x_557, 0); +lean_inc(x_558); +x_559 = lean_ctor_get(x_557, 1); +lean_inc(x_559); +if (lean_is_exclusive(x_557)) { + lean_ctor_release(x_557, 0); + lean_ctor_release(x_557, 1); + x_560 = x_557; +} else { + lean_dec_ref(x_557); + x_560 = lean_box(0); +} +x_561 = l_Array_unzip___rarg(x_558); +lean_dec(x_558); +x_562 = lean_ctor_get(x_561, 0); +lean_inc(x_562); +x_563 = lean_ctor_get(x_561, 1); +lean_inc(x_563); +if (lean_is_exclusive(x_561)) { + lean_ctor_release(x_561, 0); + lean_ctor_release(x_561, 1); + x_564 = x_561; +} else { + lean_dec_ref(x_561); + x_564 = lean_box(0); +} +x_565 = l_Array_append___rarg(x_552, x_563); +lean_dec(x_563); +lean_inc(x_542); +lean_inc(x_17); +x_566 = l_Lean_Elab_Deriving_mkInductiveApp(x_17, x_542, x_8, x_9, x_10, x_11, x_12, x_13, x_559); +x_567 = lean_ctor_get(x_566, 0); +lean_inc(x_567); +x_568 = lean_ctor_get(x_566, 1); +lean_inc(x_568); +if (lean_is_exclusive(x_566)) { + lean_ctor_release(x_566, 0); + lean_ctor_release(x_566, 1); + x_569 = x_566; +} else { + lean_dec_ref(x_566); + x_569 = lean_box(0); +} +lean_inc(x_8); +lean_inc(x_17); +x_570 = l_Lean_Elab_Deriving_ToExpr_updateIndType(x_17, x_567, x_8, x_9, x_10, x_11, x_12, x_13, x_568); +if (lean_obj_tag(x_570) == 0) +{ +lean_object* x_571; lean_object* x_572; lean_object* x_573; +x_571 = lean_ctor_get(x_570, 0); +lean_inc(x_571); +x_572 = lean_ctor_get(x_570, 1); +lean_inc(x_572); +lean_dec(x_570); +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_573 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr(x_17, x_542, x_8, x_9, x_10, x_11, x_12, x_13, x_572); +if (lean_obj_tag(x_573) == 0) +{ +lean_object* x_574; lean_object* x_575; lean_object* x_576; uint8_t x_577; lean_object* x_578; lean_object* x_579; lean_object* x_580; lean_object* x_581; lean_object* x_582; lean_object* x_583; lean_object* x_584; lean_object* x_585; lean_object* x_586; lean_object* x_587; lean_object* x_588; lean_object* x_589; lean_object* x_590; lean_object* x_591; lean_object* x_592; lean_object* x_593; lean_object* x_594; lean_object* x_595; lean_object* x_596; lean_object* x_597; lean_object* x_598; lean_object* x_599; lean_object* x_600; lean_object* x_601; lean_object* x_602; lean_object* x_603; lean_object* x_604; lean_object* x_605; lean_object* x_606; lean_object* x_607; lean_object* x_608; lean_object* x_609; lean_object* x_610; lean_object* x_611; lean_object* x_612; lean_object* x_613; lean_object* x_614; lean_object* x_615; lean_object* x_616; lean_object* x_617; lean_object* x_618; lean_object* x_619; lean_object* x_620; lean_object* x_621; lean_object* x_622; size_t x_623; lean_object* x_624; lean_object* x_625; lean_object* x_626; lean_object* x_627; lean_object* x_628; lean_object* x_629; lean_object* x_630; lean_object* x_631; lean_object* x_632; lean_object* x_633; lean_object* x_634; lean_object* x_635; lean_object* x_636; lean_object* x_637; lean_object* x_638; lean_object* x_639; lean_object* x_640; lean_object* x_641; lean_object* x_642; lean_object* x_643; lean_object* x_644; lean_object* x_645; lean_object* x_646; lean_object* x_647; lean_object* x_648; lean_object* x_649; lean_object* x_650; lean_object* x_651; lean_object* x_652; lean_object* x_653; +x_574 = lean_ctor_get(x_573, 0); +lean_inc(x_574); +x_575 = lean_ctor_get(x_573, 1); +lean_inc(x_575); +lean_dec(x_573); +x_576 = lean_ctor_get(x_12, 5); +lean_inc(x_576); +x_577 = 0; +x_578 = l_Lean_SourceInfo_fromRef(x_576, x_577); +lean_dec(x_576); +x_579 = lean_ctor_get(x_12, 10); +lean_inc(x_579); +x_580 = lean_st_ref_get(x_13, x_575); +x_581 = lean_ctor_get(x_580, 0); +lean_inc(x_581); +x_582 = lean_ctor_get(x_580, 1); +lean_inc(x_582); +if (lean_is_exclusive(x_580)) { + lean_ctor_release(x_580, 0); + lean_ctor_release(x_580, 1); + x_583 = x_580; +} else { + lean_dec_ref(x_580); + x_583 = lean_box(0); +} +x_584 = lean_ctor_get(x_581, 0); +lean_inc(x_584); +lean_dec(x_581); +x_585 = lean_environment_main_module(x_584); +x_586 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17; +x_587 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__9; +lean_inc(x_578); +x_588 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_588, 0, x_578); +lean_ctor_set(x_588, 1, x_586); +lean_ctor_set(x_588, 2, x_587); +x_589 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__8; +lean_inc_n(x_588, 6); +lean_inc(x_578); +x_590 = l_Lean_Syntax_node6(x_578, x_589, x_588, x_588, x_588, x_588, x_588, x_588); +x_591 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__4; +lean_inc(x_588); +lean_inc(x_578); +x_592 = l_Lean_Syntax_node1(x_578, x_591, x_588); +x_593 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__1; +lean_inc(x_578); +if (lean_is_scalar(x_583)) { + x_594 = lean_alloc_ctor(2, 2, 0); +} else { + x_594 = x_583; + lean_ctor_set_tag(x_594, 2); +} +lean_ctor_set(x_594, 0, x_578); +lean_ctor_set(x_594, 1, x_593); +x_595 = l_Array_append___rarg(x_587, x_565); +lean_dec(x_565); +lean_inc(x_578); +x_596 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_596, 0, x_578); +lean_ctor_set(x_596, 1, x_586); +lean_ctor_set(x_596, 2, x_595); +x_597 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__9; +lean_inc(x_578); +if (lean_is_scalar(x_569)) { + x_598 = lean_alloc_ctor(2, 2, 0); +} else { + x_598 = x_569; + lean_ctor_set_tag(x_598, 2); +} +lean_ctor_set(x_598, 0, x_578); +lean_ctor_set(x_598, 1, x_597); +x_599 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__11; +lean_inc(x_579); +lean_inc(x_585); +x_600 = l_Lean_addMacroScope(x_585, x_599, x_579); +x_601 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__10; +x_602 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__24; +lean_inc(x_578); +x_603 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_603, 0, x_578); +lean_ctor_set(x_603, 1, x_601); +lean_ctor_set(x_603, 2, x_600); +lean_ctor_set(x_603, 3, x_602); +lean_inc(x_578); +x_604 = l_Lean_Syntax_node1(x_578, x_586, x_571); +x_605 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__5; +lean_inc(x_578); +x_606 = l_Lean_Syntax_node2(x_578, x_605, x_603, x_604); +x_607 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__8; +lean_inc(x_578); +x_608 = l_Lean_Syntax_node2(x_578, x_607, x_598, x_606); +x_609 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__6; +lean_inc(x_578); +x_610 = l_Lean_Syntax_node2(x_578, x_609, x_596, x_608); +x_611 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__9; +lean_inc(x_578); +if (lean_is_scalar(x_564)) { + x_612 = lean_alloc_ctor(2, 2, 0); +} else { + x_612 = x_564; + lean_ctor_set_tag(x_612, 2); +} +lean_ctor_set(x_612, 0, x_578); +lean_ctor_set(x_612, 1, x_611); +x_613 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__3; +lean_inc(x_579); +lean_inc(x_585); +x_614 = l_Lean_addMacroScope(x_585, x_613, x_579); +x_615 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__2; +x_616 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__36; +lean_inc(x_578); +x_617 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_617, 0, x_578); +lean_ctor_set(x_617, 1, x_615); +lean_ctor_set(x_617, 2, x_614); +lean_ctor_set(x_617, 3, x_616); +x_618 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__34; +lean_inc(x_588); +lean_inc(x_578); +x_619 = l_Lean_Syntax_node2(x_578, x_618, x_617, x_588); +x_620 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__25; +lean_inc(x_578); +if (lean_is_scalar(x_560)) { + x_621 = lean_alloc_ctor(2, 2, 0); +} else { + x_621 = x_560; + lean_ctor_set_tag(x_621, 2); +} +lean_ctor_set(x_621, 0, x_578); +lean_ctor_set(x_621, 1, x_620); +x_622 = lean_mk_syntax_ident(x_533); +x_623 = lean_array_size(x_562); +x_624 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__12; +x_625 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__2(x_624, x_623, x_556, x_562); +x_626 = l_Array_append___rarg(x_587, x_625); +lean_dec(x_625); +lean_inc(x_578); +x_627 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_627, 0, x_578); +lean_ctor_set(x_627, 1, x_586); +lean_ctor_set(x_627, 2, x_626); +lean_inc(x_578); +x_628 = l_Lean_Syntax_node2(x_578, x_605, x_622, x_627); +x_629 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__38; +lean_inc(x_621); +lean_inc(x_578); +x_630 = l_Lean_Syntax_node2(x_578, x_629, x_621, x_628); +lean_inc_n(x_588, 2); +lean_inc(x_578); +x_631 = l_Lean_Syntax_node3(x_578, x_586, x_588, x_588, x_630); +x_632 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__32; +lean_inc(x_578); +x_633 = l_Lean_Syntax_node2(x_578, x_632, x_619, x_631); +x_634 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__10; +x_635 = l_Lean_addMacroScope(x_585, x_634, x_579); +x_636 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__9; +x_637 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__40; +lean_inc(x_578); +x_638 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_638, 0, x_578); +lean_ctor_set(x_638, 1, x_636); +lean_ctor_set(x_638, 2, x_635); +lean_ctor_set(x_638, 3, x_637); +lean_inc(x_588); +lean_inc(x_578); +x_639 = l_Lean_Syntax_node2(x_578, x_618, x_638, x_588); +lean_inc(x_578); +x_640 = l_Lean_Syntax_node2(x_578, x_629, x_621, x_574); +lean_inc_n(x_588, 2); +lean_inc(x_578); +x_641 = l_Lean_Syntax_node3(x_578, x_586, x_588, x_588, x_640); +lean_inc(x_578); +x_642 = l_Lean_Syntax_node2(x_578, x_632, x_639, x_641); +lean_inc(x_588); +lean_inc(x_578); +x_643 = l_Lean_Syntax_node3(x_578, x_586, x_633, x_588, x_642); +x_644 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__30; +lean_inc(x_578); +x_645 = l_Lean_Syntax_node1(x_578, x_644, x_643); +x_646 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__8; +lean_inc(x_588); +lean_inc(x_578); +x_647 = l_Lean_Syntax_node3(x_578, x_646, x_612, x_645, x_588); +x_648 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__2; +lean_inc(x_588); +lean_inc(x_578); +x_649 = l_Lean_Syntax_node6(x_578, x_648, x_592, x_594, x_588, x_588, x_610, x_647); +x_650 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__6; +x_651 = l_Lean_Syntax_node2(x_578, x_650, x_590, x_649); +x_652 = lean_array_push(x_29, x_651); +lean_ctor_set(x_7, 1, x_652); +lean_ctor_set(x_7, 0, x_536); +x_653 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_653, 0, x_7); +x_18 = x_653; +x_19 = x_582; +goto block_26; +} +else +{ +lean_object* x_654; lean_object* x_655; lean_object* x_656; lean_object* x_657; +lean_dec(x_571); +lean_dec(x_569); +lean_dec(x_565); +lean_dec(x_564); +lean_dec(x_562); +lean_dec(x_560); +lean_dec(x_536); +lean_dec(x_533); +lean_free_object(x_7); +lean_dec(x_29); +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_654 = lean_ctor_get(x_573, 0); +lean_inc(x_654); +x_655 = lean_ctor_get(x_573, 1); +lean_inc(x_655); +if (lean_is_exclusive(x_573)) { + lean_ctor_release(x_573, 0); + lean_ctor_release(x_573, 1); + x_656 = x_573; +} else { + lean_dec_ref(x_573); + x_656 = lean_box(0); +} +if (lean_is_scalar(x_656)) { + x_657 = lean_alloc_ctor(1, 2, 0); +} else { + x_657 = x_656; +} +lean_ctor_set(x_657, 0, x_654); +lean_ctor_set(x_657, 1, x_655); +return x_657; +} +} +else +{ +lean_object* x_658; lean_object* x_659; lean_object* x_660; lean_object* x_661; +lean_dec(x_569); +lean_dec(x_565); +lean_dec(x_564); +lean_dec(x_562); +lean_dec(x_560); +lean_dec(x_542); +lean_dec(x_536); +lean_dec(x_533); +lean_free_object(x_7); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +x_658 = lean_ctor_get(x_570, 0); +lean_inc(x_658); +x_659 = lean_ctor_get(x_570, 1); +lean_inc(x_659); +if (lean_is_exclusive(x_570)) { + lean_ctor_release(x_570, 0); + lean_ctor_release(x_570, 1); + x_660 = x_570; +} else { + lean_dec_ref(x_570); + x_660 = lean_box(0); +} +if (lean_is_scalar(x_660)) { + x_661 = lean_alloc_ctor(1, 2, 0); +} else { + x_661 = x_660; +} +lean_ctor_set(x_661, 0, x_658); +lean_ctor_set(x_661, 1, x_659); +return x_661; +} +} +else +{ +lean_object* x_662; lean_object* x_663; lean_object* x_664; lean_object* x_665; +lean_dec(x_545); +lean_dec(x_542); +lean_dec(x_537); +lean_dec(x_536); +lean_dec(x_533); +lean_free_object(x_7); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +x_662 = lean_ctor_get(x_548, 0); +lean_inc(x_662); +x_663 = lean_ctor_get(x_548, 1); +lean_inc(x_663); +if (lean_is_exclusive(x_548)) { + lean_ctor_release(x_548, 0); + lean_ctor_release(x_548, 1); + x_664 = x_548; +} else { + lean_dec_ref(x_548); + x_664 = lean_box(0); +} +if (lean_is_scalar(x_664)) { + x_665 = lean_alloc_ctor(1, 2, 0); +} else { + x_665 = x_664; +} +lean_ctor_set(x_665, 0, x_662); +lean_ctor_set(x_665, 1, x_663); +return x_665; +} +} +else +{ +lean_object* x_666; lean_object* x_667; lean_object* x_668; lean_object* x_669; +lean_dec(x_537); +lean_dec(x_536); +lean_dec(x_533); +lean_free_object(x_7); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +x_666 = lean_ctor_get(x_541, 0); +lean_inc(x_666); +x_667 = lean_ctor_get(x_541, 1); +lean_inc(x_667); +if (lean_is_exclusive(x_541)) { + lean_ctor_release(x_541, 0); + lean_ctor_release(x_541, 1); + x_668 = x_541; +} else { + lean_dec_ref(x_541); + x_668 = lean_box(0); +} +if (lean_is_scalar(x_668)) { + x_669 = lean_alloc_ctor(1, 2, 0); +} else { + x_669 = x_668; +} +lean_ctor_set(x_669, 0, x_666); +lean_ctor_set(x_669, 1, x_667); +return x_669; +} +} +} +} +} +else +{ +lean_object* x_670; lean_object* x_671; lean_object* x_672; lean_object* x_673; lean_object* x_674; uint8_t x_675; +x_670 = lean_ctor_get(x_7, 0); +x_671 = lean_ctor_get(x_7, 1); +lean_inc(x_671); +lean_inc(x_670); +lean_dec(x_7); +x_672 = lean_ctor_get(x_670, 0); +lean_inc(x_672); +x_673 = lean_ctor_get(x_670, 1); +lean_inc(x_673); +x_674 = lean_ctor_get(x_670, 2); +lean_inc(x_674); +x_675 = lean_nat_dec_lt(x_673, x_674); +if (x_675 == 0) +{ +lean_object* x_676; lean_object* x_677; +lean_dec(x_674); +lean_dec(x_673); +lean_dec(x_672); +lean_dec(x_17); +x_676 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_676, 0, x_670); +lean_ctor_set(x_676, 1, x_671); +x_677 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_677, 0, x_676); +x_18 = x_677; +x_19 = x_14; +goto block_26; +} +else +{ +lean_object* x_678; lean_object* x_679; lean_object* x_680; lean_object* x_681; lean_object* x_682; lean_object* x_683; lean_object* x_684; uint8_t x_685; +if (lean_is_exclusive(x_670)) { + lean_ctor_release(x_670, 0); + lean_ctor_release(x_670, 1); + lean_ctor_release(x_670, 2); + x_678 = x_670; +} else { + lean_dec_ref(x_670); + x_678 = lean_box(0); +} +x_679 = lean_array_fget(x_672, x_673); +x_680 = lean_unsigned_to_nat(1u); +x_681 = lean_nat_add(x_673, x_680); +lean_dec(x_673); +if (lean_is_scalar(x_678)) { + x_682 = lean_alloc_ctor(0, 3, 0); +} else { + x_682 = x_678; +} +lean_ctor_set(x_682, 0, x_672); +lean_ctor_set(x_682, 1, x_681); +lean_ctor_set(x_682, 2, x_674); +x_683 = lean_ctor_get(x_17, 0); +lean_inc(x_683); +x_684 = lean_ctor_get(x_683, 0); +lean_inc(x_684); +x_685 = l_Array_contains___at_Lean_registerInternalExceptionId___spec__1(x_1, x_684); +lean_dec(x_684); +if (x_685 == 0) +{ +lean_object* x_686; lean_object* x_687; +lean_dec(x_683); +lean_dec(x_679); +lean_dec(x_17); +x_686 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_686, 0, x_682); +lean_ctor_set(x_686, 1, x_671); +x_687 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_687, 0, x_686); +x_18 = x_687; +x_19 = x_14; +goto block_26; +} +else +{ +lean_object* x_688; +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_17); +x_688 = l_Lean_Elab_Deriving_mkInductArgNames(x_17, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_688) == 0) +{ +lean_object* x_689; lean_object* x_690; lean_object* x_691; lean_object* x_692; lean_object* x_693; lean_object* x_694; lean_object* x_695; +x_689 = lean_ctor_get(x_688, 0); +lean_inc(x_689); +x_690 = lean_ctor_get(x_688, 1); +lean_inc(x_690); +lean_dec(x_688); +lean_inc(x_689); +x_691 = l_Lean_Elab_Deriving_mkImplicitBinders(x_689, x_8, x_9, x_10, x_11, x_12, x_13, x_690); +x_692 = lean_ctor_get(x_691, 0); +lean_inc(x_692); +x_693 = lean_ctor_get(x_691, 1); +lean_inc(x_693); +lean_dec(x_691); +x_694 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__12; +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_689); +lean_inc(x_17); +x_695 = l_Lean_Elab_Deriving_mkInstImplicitBinders(x_694, x_17, x_689, x_8, x_9, x_10, x_11, x_12, x_13, x_693); +if (lean_obj_tag(x_695) == 0) +{ +lean_object* x_696; lean_object* x_697; lean_object* x_698; lean_object* x_699; lean_object* x_700; lean_object* x_701; size_t x_702; size_t x_703; lean_object* x_704; lean_object* x_705; lean_object* x_706; lean_object* x_707; lean_object* x_708; lean_object* x_709; lean_object* x_710; lean_object* x_711; lean_object* x_712; lean_object* x_713; lean_object* x_714; lean_object* x_715; lean_object* x_716; lean_object* x_717; +x_696 = lean_ctor_get(x_695, 0); +lean_inc(x_696); +x_697 = lean_ctor_get(x_695, 1); +lean_inc(x_697); +lean_dec(x_695); +x_698 = lean_box(0); +x_699 = l_Array_append___rarg(x_692, x_696); +lean_dec(x_696); +x_700 = lean_ctor_get(x_683, 1); +lean_inc(x_700); +lean_dec(x_683); +x_701 = lean_array_mk(x_700); +x_702 = lean_array_size(x_701); +x_703 = 0; +lean_inc(x_12); +x_704 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__1(x_698, x_702, x_703, x_701, x_8, x_9, x_10, x_11, x_12, x_13, x_697); +x_705 = lean_ctor_get(x_704, 0); +lean_inc(x_705); +x_706 = lean_ctor_get(x_704, 1); +lean_inc(x_706); +if (lean_is_exclusive(x_704)) { + lean_ctor_release(x_704, 0); + lean_ctor_release(x_704, 1); + x_707 = x_704; +} else { + lean_dec_ref(x_704); + x_707 = lean_box(0); +} +x_708 = l_Array_unzip___rarg(x_705); +lean_dec(x_705); +x_709 = lean_ctor_get(x_708, 0); +lean_inc(x_709); +x_710 = lean_ctor_get(x_708, 1); +lean_inc(x_710); +if (lean_is_exclusive(x_708)) { + lean_ctor_release(x_708, 0); + lean_ctor_release(x_708, 1); + x_711 = x_708; +} else { + lean_dec_ref(x_708); + x_711 = lean_box(0); +} +x_712 = l_Array_append___rarg(x_699, x_710); +lean_dec(x_710); +lean_inc(x_689); +lean_inc(x_17); +x_713 = l_Lean_Elab_Deriving_mkInductiveApp(x_17, x_689, x_8, x_9, x_10, x_11, x_12, x_13, x_706); +x_714 = lean_ctor_get(x_713, 0); +lean_inc(x_714); +x_715 = lean_ctor_get(x_713, 1); +lean_inc(x_715); +if (lean_is_exclusive(x_713)) { + lean_ctor_release(x_713, 0); + lean_ctor_release(x_713, 1); + x_716 = x_713; +} else { + lean_dec_ref(x_713); + x_716 = lean_box(0); +} +lean_inc(x_8); +lean_inc(x_17); +x_717 = l_Lean_Elab_Deriving_ToExpr_updateIndType(x_17, x_714, x_8, x_9, x_10, x_11, x_12, x_13, x_715); +if (lean_obj_tag(x_717) == 0) +{ +lean_object* x_718; lean_object* x_719; lean_object* x_720; +x_718 = lean_ctor_get(x_717, 0); +lean_inc(x_718); +x_719 = lean_ctor_get(x_717, 1); +lean_inc(x_719); +lean_dec(x_717); +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_720 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr(x_17, x_689, x_8, x_9, x_10, x_11, x_12, x_13, x_719); +if (lean_obj_tag(x_720) == 0) +{ +lean_object* x_721; lean_object* x_722; lean_object* x_723; uint8_t x_724; lean_object* x_725; lean_object* x_726; lean_object* x_727; lean_object* x_728; lean_object* x_729; lean_object* x_730; lean_object* x_731; lean_object* x_732; lean_object* x_733; lean_object* x_734; lean_object* x_735; lean_object* x_736; lean_object* x_737; lean_object* x_738; lean_object* x_739; lean_object* x_740; lean_object* x_741; lean_object* x_742; lean_object* x_743; lean_object* x_744; lean_object* x_745; lean_object* x_746; lean_object* x_747; lean_object* x_748; lean_object* x_749; lean_object* x_750; lean_object* x_751; lean_object* x_752; lean_object* x_753; lean_object* x_754; lean_object* x_755; lean_object* x_756; lean_object* x_757; lean_object* x_758; lean_object* x_759; lean_object* x_760; lean_object* x_761; lean_object* x_762; lean_object* x_763; lean_object* x_764; lean_object* x_765; lean_object* x_766; lean_object* x_767; lean_object* x_768; lean_object* x_769; size_t x_770; lean_object* x_771; lean_object* x_772; lean_object* x_773; lean_object* x_774; lean_object* x_775; lean_object* x_776; lean_object* x_777; lean_object* x_778; lean_object* x_779; lean_object* x_780; lean_object* x_781; lean_object* x_782; lean_object* x_783; lean_object* x_784; lean_object* x_785; lean_object* x_786; lean_object* x_787; lean_object* x_788; lean_object* x_789; lean_object* x_790; lean_object* x_791; lean_object* x_792; lean_object* x_793; lean_object* x_794; lean_object* x_795; lean_object* x_796; lean_object* x_797; lean_object* x_798; lean_object* x_799; lean_object* x_800; lean_object* x_801; +x_721 = lean_ctor_get(x_720, 0); +lean_inc(x_721); +x_722 = lean_ctor_get(x_720, 1); +lean_inc(x_722); +lean_dec(x_720); +x_723 = lean_ctor_get(x_12, 5); +lean_inc(x_723); +x_724 = 0; +x_725 = l_Lean_SourceInfo_fromRef(x_723, x_724); +lean_dec(x_723); +x_726 = lean_ctor_get(x_12, 10); +lean_inc(x_726); +x_727 = lean_st_ref_get(x_13, x_722); +x_728 = lean_ctor_get(x_727, 0); +lean_inc(x_728); +x_729 = lean_ctor_get(x_727, 1); +lean_inc(x_729); +if (lean_is_exclusive(x_727)) { + lean_ctor_release(x_727, 0); + lean_ctor_release(x_727, 1); + x_730 = x_727; +} else { + lean_dec_ref(x_727); + x_730 = lean_box(0); +} +x_731 = lean_ctor_get(x_728, 0); +lean_inc(x_731); +lean_dec(x_728); +x_732 = lean_environment_main_module(x_731); +x_733 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17; +x_734 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__9; +lean_inc(x_725); +x_735 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_735, 0, x_725); +lean_ctor_set(x_735, 1, x_733); +lean_ctor_set(x_735, 2, x_734); +x_736 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__8; +lean_inc_n(x_735, 6); +lean_inc(x_725); +x_737 = l_Lean_Syntax_node6(x_725, x_736, x_735, x_735, x_735, x_735, x_735, x_735); +x_738 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__4; +lean_inc(x_735); +lean_inc(x_725); +x_739 = l_Lean_Syntax_node1(x_725, x_738, x_735); +x_740 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__1; +lean_inc(x_725); +if (lean_is_scalar(x_730)) { + x_741 = lean_alloc_ctor(2, 2, 0); +} else { + x_741 = x_730; + lean_ctor_set_tag(x_741, 2); +} +lean_ctor_set(x_741, 0, x_725); +lean_ctor_set(x_741, 1, x_740); +x_742 = l_Array_append___rarg(x_734, x_712); +lean_dec(x_712); +lean_inc(x_725); +x_743 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_743, 0, x_725); +lean_ctor_set(x_743, 1, x_733); +lean_ctor_set(x_743, 2, x_742); +x_744 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__9; +lean_inc(x_725); +if (lean_is_scalar(x_716)) { + x_745 = lean_alloc_ctor(2, 2, 0); +} else { + x_745 = x_716; + lean_ctor_set_tag(x_745, 2); +} +lean_ctor_set(x_745, 0, x_725); +lean_ctor_set(x_745, 1, x_744); +x_746 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__11; +lean_inc(x_726); +lean_inc(x_732); +x_747 = l_Lean_addMacroScope(x_732, x_746, x_726); +x_748 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__10; +x_749 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__24; +lean_inc(x_725); +x_750 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_750, 0, x_725); +lean_ctor_set(x_750, 1, x_748); +lean_ctor_set(x_750, 2, x_747); +lean_ctor_set(x_750, 3, x_749); +lean_inc(x_725); +x_751 = l_Lean_Syntax_node1(x_725, x_733, x_718); +x_752 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__5; +lean_inc(x_725); +x_753 = l_Lean_Syntax_node2(x_725, x_752, x_750, x_751); +x_754 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__8; +lean_inc(x_725); +x_755 = l_Lean_Syntax_node2(x_725, x_754, x_745, x_753); +x_756 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__6; +lean_inc(x_725); +x_757 = l_Lean_Syntax_node2(x_725, x_756, x_743, x_755); +x_758 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__9; +lean_inc(x_725); +if (lean_is_scalar(x_711)) { + x_759 = lean_alloc_ctor(2, 2, 0); +} else { + x_759 = x_711; + lean_ctor_set_tag(x_759, 2); +} +lean_ctor_set(x_759, 0, x_725); +lean_ctor_set(x_759, 1, x_758); +x_760 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__3; +lean_inc(x_726); +lean_inc(x_732); +x_761 = l_Lean_addMacroScope(x_732, x_760, x_726); +x_762 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__2; +x_763 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__36; +lean_inc(x_725); +x_764 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_764, 0, x_725); +lean_ctor_set(x_764, 1, x_762); +lean_ctor_set(x_764, 2, x_761); +lean_ctor_set(x_764, 3, x_763); +x_765 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__34; +lean_inc(x_735); +lean_inc(x_725); +x_766 = l_Lean_Syntax_node2(x_725, x_765, x_764, x_735); +x_767 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__25; +lean_inc(x_725); +if (lean_is_scalar(x_707)) { + x_768 = lean_alloc_ctor(2, 2, 0); +} else { + x_768 = x_707; + lean_ctor_set_tag(x_768, 2); +} +lean_ctor_set(x_768, 0, x_725); +lean_ctor_set(x_768, 1, x_767); +x_769 = lean_mk_syntax_ident(x_679); +x_770 = lean_array_size(x_709); +x_771 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__12; +x_772 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__2(x_771, x_770, x_703, x_709); +x_773 = l_Array_append___rarg(x_734, x_772); +lean_dec(x_772); +lean_inc(x_725); +x_774 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_774, 0, x_725); +lean_ctor_set(x_774, 1, x_733); +lean_ctor_set(x_774, 2, x_773); +lean_inc(x_725); +x_775 = l_Lean_Syntax_node2(x_725, x_752, x_769, x_774); +x_776 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__38; +lean_inc(x_768); +lean_inc(x_725); +x_777 = l_Lean_Syntax_node2(x_725, x_776, x_768, x_775); +lean_inc_n(x_735, 2); +lean_inc(x_725); +x_778 = l_Lean_Syntax_node3(x_725, x_733, x_735, x_735, x_777); +x_779 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__32; +lean_inc(x_725); +x_780 = l_Lean_Syntax_node2(x_725, x_779, x_766, x_778); +x_781 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__10; +x_782 = l_Lean_addMacroScope(x_732, x_781, x_726); +x_783 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__9; +x_784 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__40; +lean_inc(x_725); +x_785 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_785, 0, x_725); +lean_ctor_set(x_785, 1, x_783); +lean_ctor_set(x_785, 2, x_782); +lean_ctor_set(x_785, 3, x_784); +lean_inc(x_735); +lean_inc(x_725); +x_786 = l_Lean_Syntax_node2(x_725, x_765, x_785, x_735); +lean_inc(x_725); +x_787 = l_Lean_Syntax_node2(x_725, x_776, x_768, x_721); +lean_inc_n(x_735, 2); +lean_inc(x_725); +x_788 = l_Lean_Syntax_node3(x_725, x_733, x_735, x_735, x_787); +lean_inc(x_725); +x_789 = l_Lean_Syntax_node2(x_725, x_779, x_786, x_788); +lean_inc(x_735); +lean_inc(x_725); +x_790 = l_Lean_Syntax_node3(x_725, x_733, x_780, x_735, x_789); +x_791 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__30; +lean_inc(x_725); +x_792 = l_Lean_Syntax_node1(x_725, x_791, x_790); +x_793 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__8; +lean_inc(x_735); +lean_inc(x_725); +x_794 = l_Lean_Syntax_node3(x_725, x_793, x_759, x_792, x_735); +x_795 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__2; +lean_inc(x_735); +lean_inc(x_725); +x_796 = l_Lean_Syntax_node6(x_725, x_795, x_739, x_741, x_735, x_735, x_757, x_794); +x_797 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__6; +x_798 = l_Lean_Syntax_node2(x_725, x_797, x_737, x_796); +x_799 = lean_array_push(x_671, x_798); +x_800 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_800, 0, x_682); +lean_ctor_set(x_800, 1, x_799); +x_801 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_801, 0, x_800); +x_18 = x_801; +x_19 = x_729; +goto block_26; +} +else +{ +lean_object* x_802; lean_object* x_803; lean_object* x_804; lean_object* x_805; +lean_dec(x_718); +lean_dec(x_716); +lean_dec(x_712); +lean_dec(x_711); +lean_dec(x_709); +lean_dec(x_707); +lean_dec(x_682); +lean_dec(x_679); +lean_dec(x_671); +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_802 = lean_ctor_get(x_720, 0); +lean_inc(x_802); +x_803 = lean_ctor_get(x_720, 1); +lean_inc(x_803); +if (lean_is_exclusive(x_720)) { + lean_ctor_release(x_720, 0); + lean_ctor_release(x_720, 1); + x_804 = x_720; +} else { + lean_dec_ref(x_720); + x_804 = lean_box(0); +} +if (lean_is_scalar(x_804)) { + x_805 = lean_alloc_ctor(1, 2, 0); +} else { + x_805 = x_804; +} +lean_ctor_set(x_805, 0, x_802); +lean_ctor_set(x_805, 1, x_803); +return x_805; +} +} +else +{ +lean_object* x_806; lean_object* x_807; lean_object* x_808; lean_object* x_809; +lean_dec(x_716); +lean_dec(x_712); +lean_dec(x_711); +lean_dec(x_709); +lean_dec(x_707); +lean_dec(x_689); +lean_dec(x_682); +lean_dec(x_679); +lean_dec(x_671); +lean_dec(x_17); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +x_806 = lean_ctor_get(x_717, 0); +lean_inc(x_806); +x_807 = lean_ctor_get(x_717, 1); +lean_inc(x_807); +if (lean_is_exclusive(x_717)) { + lean_ctor_release(x_717, 0); + lean_ctor_release(x_717, 1); + x_808 = x_717; +} else { + lean_dec_ref(x_717); + x_808 = lean_box(0); +} +if (lean_is_scalar(x_808)) { + x_809 = lean_alloc_ctor(1, 2, 0); +} else { + x_809 = x_808; +} +lean_ctor_set(x_809, 0, x_806); +lean_ctor_set(x_809, 1, x_807); +return x_809; +} +} +else +{ +lean_object* x_810; lean_object* x_811; lean_object* x_812; lean_object* x_813; +lean_dec(x_692); +lean_dec(x_689); +lean_dec(x_683); +lean_dec(x_682); +lean_dec(x_679); +lean_dec(x_671); +lean_dec(x_17); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +x_810 = lean_ctor_get(x_695, 0); +lean_inc(x_810); +x_811 = lean_ctor_get(x_695, 1); +lean_inc(x_811); +if (lean_is_exclusive(x_695)) { + lean_ctor_release(x_695, 0); + lean_ctor_release(x_695, 1); + x_812 = x_695; +} else { + lean_dec_ref(x_695); + x_812 = lean_box(0); +} +if (lean_is_scalar(x_812)) { + x_813 = lean_alloc_ctor(1, 2, 0); +} else { + x_813 = x_812; +} +lean_ctor_set(x_813, 0, x_810); +lean_ctor_set(x_813, 1, x_811); +return x_813; +} +} +else +{ +lean_object* x_814; lean_object* x_815; lean_object* x_816; lean_object* x_817; +lean_dec(x_683); +lean_dec(x_682); +lean_dec(x_679); +lean_dec(x_671); +lean_dec(x_17); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +x_814 = lean_ctor_get(x_688, 0); +lean_inc(x_814); +x_815 = lean_ctor_get(x_688, 1); +lean_inc(x_815); +if (lean_is_exclusive(x_688)) { + lean_ctor_release(x_688, 0); + lean_ctor_release(x_688, 1); + x_816 = x_688; +} else { + lean_dec_ref(x_688); + x_816 = lean_box(0); +} +if (lean_is_scalar(x_816)) { + x_817 = lean_alloc_ctor(1, 2, 0); +} else { + x_817 = x_816; +} +lean_ctor_set(x_817, 0, x_814); +lean_ctor_set(x_817, 1, x_815); +return x_817; +} +} +} +} +block_26: +{ +if (lean_obj_tag(x_18) == 0) +{ +lean_object* x_20; lean_object* 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); +x_20 = lean_ctor_get(x_18, 0); +lean_inc(x_20); +lean_dec(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; +} +else +{ +lean_object* x_22; size_t x_23; size_t x_24; +x_22 = lean_ctor_get(x_18, 0); +lean_inc(x_22); +lean_dec(x_18); +x_23 = 1; +x_24 = lean_usize_add(x_6, x_23); +x_6 = x_24; +x_7 = x_22; +x_14 = x_19; +goto _start; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_ToExpr_mkInstanceCmds(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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; size_t x_18; size_t x_19; lean_object* x_20; +x_10 = lean_ctor_get(x_1, 1); +lean_inc(x_10); +x_11 = lean_array_get_size(x_10); +x_12 = lean_unsigned_to_nat(0u); +x_13 = l_Array_toSubarray___rarg(x_10, x_12, x_11); +x_14 = lean_ctor_get(x_1, 0); +lean_inc(x_14); +lean_dec(x_1); +x_15 = lean_box(0); +x_16 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__1; +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_13); +lean_ctor_set(x_17, 1, x_16); +x_18 = lean_array_size(x_14); +x_19 = 0; +x_20 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3(x_2, x_14, x_15, x_14, x_18, x_19, x_17, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_14); +if (lean_obj_tag(x_20) == 0) +{ +uint8_t x_21; +x_21 = !lean_is_exclusive(x_20); +if (x_21 == 0) +{ +lean_object* x_22; lean_object* x_23; +x_22 = lean_ctor_get(x_20, 0); +x_23 = lean_ctor_get(x_22, 1); +lean_inc(x_23); +lean_dec(x_22); +lean_ctor_set(x_20, 0, x_23); +return x_20; +} +else +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_24 = lean_ctor_get(x_20, 0); +x_25 = lean_ctor_get(x_20, 1); +lean_inc(x_25); +lean_inc(x_24); +lean_dec(x_20); +x_26 = lean_ctor_get(x_24, 1); +lean_inc(x_26); +lean_dec(x_24); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_25); +return x_27; +} +} +else +{ +uint8_t x_28; +x_28 = !lean_is_exclusive(x_20); +if (x_28 == 0) +{ +return x_20; +} +else +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_29 = lean_ctor_get(x_20, 0); +x_30 = lean_ctor_get(x_20, 1); +lean_inc(x_30); +lean_inc(x_29); +lean_dec(x_20); +x_31 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_31, 0, x_29); +lean_ctor_set(x_31, 1, x_30); +return x_31; +} +} +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___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: +{ +size_t x_12; size_t x_13; lean_object* x_14; +x_12 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_13 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_14 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__1(x_1, x_12, x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_10); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_1); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +size_t x_5; size_t x_6; lean_object* x_7; +x_5 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_6 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_7 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__2(x_1, x_5, x_6, x_4); +lean_dec(x_1); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +size_t x_15; size_t x_16; lean_object* x_17; +x_15 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_16 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_17 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3(x_1, x_2, x_3, x_4, x_15, x_16, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_17; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l_Lean_Elab_Deriving_ToExpr_mkInstanceCmds(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_List_mapTR_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprInstanceCmds___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_ofSyntax(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_ofSyntax(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_EXPORT lean_object* l_Lean_Elab_Deriving_ToExpr_mkToExprInstanceCmds___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_alloc_ctor(0, 2, 0); +lean_ctor_set(x_10, 0, x_1); +lean_ctor_set(x_10, 1, x_9); +return x_10; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_mkToExprInstanceCmds___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__14; +x_2 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__15; +x_3 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__1; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_mkToExprInstanceCmds___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("\n", 1, 1); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_mkToExprInstanceCmds___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Elab_Deriving_ToExpr_mkToExprInstanceCmds___closed__2; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_mkToExprInstanceCmds___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("", 0, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_mkToExprInstanceCmds___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Elab_Deriving_ToExpr_mkToExprInstanceCmds___closed__4; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_ToExpr_mkToExprInstanceCmds(lean_object* x_1, lean_object* x_2, lean_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_58; lean_object* x_59; uint8_t x_60; +x_58 = lean_array_get_size(x_1); +x_59 = lean_unsigned_to_nat(0u); +x_60 = lean_nat_dec_lt(x_59, x_58); +lean_dec(x_58); +if (x_60 == 0) +{ +lean_object* x_61; lean_object* x_62; +x_61 = l_Lean_instInhabitedName; +x_62 = l_outOfBounds___rarg(x_61); +x_9 = x_62; +goto block_57; +} +else +{ +lean_object* x_63; +x_63 = lean_array_fget(x_1, x_59); +x_9 = x_63; +goto block_57; +} +block_57: +{ +lean_object* x_10; lean_object* x_11; +x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__1; +lean_inc(x_2); +x_11 = l_Lean_Elab_Deriving_mkContext(x_10, x_9, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +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); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_12); +x_14 = l_Lean_Elab_Deriving_ToExpr_mkAuxFunctions(x_12, x_2, x_3, x_4, x_5, x_6, x_7, x_13); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +x_16 = lean_ctor_get(x_14, 1); +lean_inc(x_16); +lean_dec(x_14); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +x_17 = l_Lean_Elab_Deriving_ToExpr_mkInstanceCmds(x_12, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_16); +if (lean_obj_tag(x_17) == 0) +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +x_19 = lean_ctor_get(x_17, 1); +lean_inc(x_19); +lean_dec(x_17); +x_20 = lean_box(0); +x_21 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_21, 0, x_15); +lean_ctor_set(x_21, 1, x_20); +x_22 = lean_array_mk(x_21); +x_23 = l_Array_append___rarg(x_22, x_18); +lean_dec(x_18); +x_24 = l_Lean_Elab_Deriving_ToExpr_mkToExprInstanceCmds___closed__1; +x_25 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Term_traceAtCmdPos___spec__1(x_24, x_2, x_3, x_4, x_5, x_6, x_7, x_19); +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) +{ +uint8_t x_28; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_28 = !lean_is_exclusive(x_25); +if (x_28 == 0) +{ +lean_object* x_29; +x_29 = lean_ctor_get(x_25, 0); +lean_dec(x_29); +lean_ctor_set(x_25, 0, x_23); +return x_25; +} +else +{ +lean_object* x_30; lean_object* x_31; +x_30 = lean_ctor_get(x_25, 1); +lean_inc(x_30); +lean_dec(x_25); +x_31 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_31, 0, x_23); +lean_ctor_set(x_31, 1, x_30); +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; uint8_t x_41; +x_32 = lean_ctor_get(x_25, 1); +lean_inc(x_32); +lean_dec(x_25); +lean_inc(x_23); +x_33 = lean_array_to_list(x_23); +x_34 = l_List_mapTR_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprInstanceCmds___spec__1(x_33, x_20); +x_35 = l_Lean_MessageData_ofList(x_34); +x_36 = l_Lean_Elab_Deriving_ToExpr_mkToExprInstanceCmds___closed__3; +x_37 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_37, 0, x_36); +lean_ctor_set(x_37, 1, x_35); +x_38 = l_Lean_Elab_Deriving_ToExpr_mkToExprInstanceCmds___closed__5; +x_39 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_39, 0, x_37); +lean_ctor_set(x_39, 1, x_38); +x_40 = l_Lean_addTrace___at_Lean_Elab_Term_traceAtCmdPos___spec__2(x_24, x_39, x_2, x_3, x_4, x_5, x_6, x_7, x_32); +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_41 = !lean_is_exclusive(x_40); +if (x_41 == 0) +{ +lean_object* x_42; +x_42 = lean_ctor_get(x_40, 0); +lean_dec(x_42); +lean_ctor_set(x_40, 0, x_23); +return x_40; +} +else +{ +lean_object* x_43; lean_object* x_44; +x_43 = lean_ctor_get(x_40, 1); +lean_inc(x_43); +lean_dec(x_40); +x_44 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_44, 0, x_23); +lean_ctor_set(x_44, 1, x_43); +return x_44; +} +} +} +else +{ +uint8_t x_45; +lean_dec(x_15); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_45 = !lean_is_exclusive(x_17); +if (x_45 == 0) +{ +return x_17; +} +else +{ +lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_46 = lean_ctor_get(x_17, 0); +x_47 = lean_ctor_get(x_17, 1); +lean_inc(x_47); +lean_inc(x_46); +lean_dec(x_17); +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_12); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_49 = !lean_is_exclusive(x_14); +if (x_49 == 0) +{ +return x_14; +} +else +{ +lean_object* x_50; lean_object* x_51; lean_object* x_52; +x_50 = lean_ctor_get(x_14, 0); +x_51 = lean_ctor_get(x_14, 1); +lean_inc(x_51); +lean_inc(x_50); +lean_dec(x_14); +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_53; +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_53 = !lean_is_exclusive(x_11); +if (x_53 == 0) +{ +return x_11; +} +else +{ +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_dec(x_11); +x_56 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_56, 0, x_54); +lean_ctor_set(x_56, 1, x_55); +return x_56; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_ToExpr_mkToExprInstanceCmds___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_Deriving_ToExpr_mkToExprInstanceCmds___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_ToExpr_mkToExprInstanceCmds___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; +x_9 = l_Lean_Elab_Deriving_ToExpr_mkToExprInstanceCmds(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_1); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_isInductive___at_Lean_Elab_Deriving_ToExpr_mkToExprInstanceHandler___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; uint8_t x_6; +x_5 = lean_st_ref_get(x_3, x_4); +x_6 = !lean_is_exclusive(x_5); +if (x_6 == 0) +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_7 = lean_ctor_get(x_5, 0); +x_8 = lean_ctor_get(x_7, 0); +lean_inc(x_8); +lean_dec(x_7); +x_9 = lean_environment_find(x_8, x_1); +if (lean_obj_tag(x_9) == 0) +{ +uint8_t x_10; lean_object* x_11; +x_10 = 0; +x_11 = lean_box(x_10); +lean_ctor_set(x_5, 0, x_11); +return x_5; +} +else +{ +lean_object* x_12; +x_12 = lean_ctor_get(x_9, 0); +lean_inc(x_12); +lean_dec(x_9); +if (lean_obj_tag(x_12) == 5) +{ +uint8_t x_13; lean_object* x_14; +lean_dec(x_12); +x_13 = 1; +x_14 = lean_box(x_13); +lean_ctor_set(x_5, 0, x_14); +return x_5; +} +else +{ +uint8_t x_15; lean_object* x_16; +lean_dec(x_12); +x_15 = 0; +x_16 = lean_box(x_15); +lean_ctor_set(x_5, 0, x_16); +return x_5; +} +} +} +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_17 = lean_ctor_get(x_5, 0); +x_18 = lean_ctor_get(x_5, 1); +lean_inc(x_18); +lean_inc(x_17); +lean_dec(x_5); +x_19 = lean_ctor_get(x_17, 0); +lean_inc(x_19); +lean_dec(x_17); +x_20 = lean_environment_find(x_19, x_1); +if (lean_obj_tag(x_20) == 0) +{ +uint8_t x_21; lean_object* x_22; lean_object* x_23; +x_21 = 0; +x_22 = lean_box(x_21); +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_18); +return x_23; +} +else +{ +lean_object* x_24; +x_24 = lean_ctor_get(x_20, 0); +lean_inc(x_24); +lean_dec(x_20); +if (lean_obj_tag(x_24) == 5) +{ +uint8_t x_25; lean_object* x_26; lean_object* x_27; +lean_dec(x_24); +x_25 = 1; +x_26 = lean_box(x_25); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_18); +return x_27; +} +else +{ +uint8_t x_28; lean_object* x_29; lean_object* x_30; +lean_dec(x_24); +x_28 = 0; +x_29 = lean_box(x_28); +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_18); +return x_30; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_Deriving_ToExpr_mkToExprInstanceHandler___spec__2(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +uint8_t x_7; +x_7 = lean_usize_dec_eq(x_2, x_3); +if (x_7 == 0) +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_8 = lean_array_uget(x_1, x_2); +x_9 = l_Lean_isInductive___at_Lean_Elab_Deriving_ToExpr_mkToExprInstanceHandler___spec__1(x_8, x_4, x_5, x_6); +x_10 = lean_ctor_get(x_9, 0); +lean_inc(x_10); +x_11 = lean_unbox(x_10); +lean_dec(x_10); +if (x_11 == 0) +{ +uint8_t x_12; +x_12 = !lean_is_exclusive(x_9); +if (x_12 == 0) +{ +lean_object* x_13; uint8_t x_14; lean_object* x_15; +x_13 = lean_ctor_get(x_9, 0); +lean_dec(x_13); +x_14 = 1; +x_15 = lean_box(x_14); +lean_ctor_set(x_9, 0, x_15); +return x_9; +} +else +{ +lean_object* x_16; uint8_t 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 = 1; +x_18 = lean_box(x_17); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_16); +return x_19; +} +} +else +{ +lean_object* x_20; size_t x_21; size_t x_22; +x_20 = lean_ctor_get(x_9, 1); +lean_inc(x_20); +lean_dec(x_9); +x_21 = 1; +x_22 = lean_usize_add(x_2, x_21); +x_2 = x_22; +x_6 = x_20; +goto _start; +} +} +else +{ +uint8_t x_24; lean_object* x_25; lean_object* x_26; +x_24 = 0; +x_25 = lean_box(x_24); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set(x_26, 1, x_6); +return x_26; +} +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_mkToExprInstanceHandler___lambda__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Elab_autoImplicit; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_ToExpr_mkToExprInstanceHandler___lambda__1(lean_object* x_1) { +_start: +{ +uint8_t x_2; +x_2 = !lean_is_exclusive(x_1); +if (x_2 == 0) +{ +lean_object* x_3; lean_object* x_4; uint8_t x_5; lean_object* x_6; +x_3 = lean_ctor_get(x_1, 1); +x_4 = l_Lean_Elab_Deriving_ToExpr_mkToExprInstanceHandler___lambda__1___closed__1; +x_5 = 1; +x_6 = l_Lean_Option_set___at_Lean_Elab_Term_withoutMacroStackAtErr___spec__1(x_3, x_4, x_5); +lean_ctor_set(x_1, 1, x_6); +return x_1; +} +else +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; lean_object* x_17; uint8_t x_18; lean_object* x_19; lean_object* x_20; +x_7 = lean_ctor_get(x_1, 0); +x_8 = lean_ctor_get(x_1, 1); +x_9 = lean_ctor_get(x_1, 2); +x_10 = lean_ctor_get(x_1, 3); +x_11 = lean_ctor_get(x_1, 4); +x_12 = lean_ctor_get(x_1, 5); +x_13 = lean_ctor_get(x_1, 6); +x_14 = lean_ctor_get(x_1, 7); +x_15 = lean_ctor_get(x_1, 8); +x_16 = lean_ctor_get_uint8(x_1, sizeof(void*)*9); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_dec(x_1); +x_17 = l_Lean_Elab_Deriving_ToExpr_mkToExprInstanceHandler___lambda__1___closed__1; +x_18 = 1; +x_19 = l_Lean_Option_set___at_Lean_Elab_Term_withoutMacroStackAtErr___spec__1(x_8, x_17, x_18); +x_20 = lean_alloc_ctor(0, 9, 1); +lean_ctor_set(x_20, 0, x_7); +lean_ctor_set(x_20, 1, x_19); +lean_ctor_set(x_20, 2, x_9); +lean_ctor_set(x_20, 3, x_10); +lean_ctor_set(x_20, 4, x_11); +lean_ctor_set(x_20, 5, x_12); +lean_ctor_set(x_20, 6, x_13); +lean_ctor_set(x_20, 7, x_14); +lean_ctor_set(x_20, 8, x_15); +lean_ctor_set_uint8(x_20, sizeof(void*)*9, x_16); +return x_20; +} +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_mkToExprInstanceHandler___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Deriving_ToExpr_mkToExprInstanceHandler___lambda__1), 1, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_ToExpr_mkToExprInstanceHandler(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_object* x_7; lean_object* x_44; uint8_t x_45; +x_5 = lean_array_get_size(x_1); +x_44 = lean_unsigned_to_nat(0u); +x_45 = lean_nat_dec_lt(x_44, x_5); +if (x_45 == 0) +{ +uint8_t x_46; +x_46 = 1; +x_6 = x_46; +x_7 = x_4; +goto block_43; +} +else +{ +size_t x_47; size_t x_48; lean_object* x_49; lean_object* x_50; uint8_t x_51; +x_47 = 0; +x_48 = lean_usize_of_nat(x_5); +x_49 = l_Array_anyMUnsafe_any___at_Lean_Elab_Deriving_ToExpr_mkToExprInstanceHandler___spec__2(x_1, x_47, x_48, x_2, x_3, x_4); +x_50 = lean_ctor_get(x_49, 0); +lean_inc(x_50); +x_51 = lean_unbox(x_50); +lean_dec(x_50); +if (x_51 == 0) +{ +lean_object* x_52; uint8_t x_53; +x_52 = lean_ctor_get(x_49, 1); +lean_inc(x_52); +lean_dec(x_49); +x_53 = 1; +x_6 = x_53; +x_7 = x_52; +goto block_43; +} +else +{ +lean_object* x_54; uint8_t x_55; +x_54 = lean_ctor_get(x_49, 1); +lean_inc(x_54); +lean_dec(x_49); +x_55 = 0; +x_6 = x_55; +x_7 = x_54; +goto block_43; +} +} +block_43: +{ +if (x_6 == 0) +{ +uint8_t x_8; lean_object* x_9; lean_object* x_10; +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_8 = 0; +x_9 = lean_box(x_8); +x_10 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_10, 0, x_9); +lean_ctor_set(x_10, 1, x_7); +return x_10; +} +else +{ +lean_object* x_11; uint8_t x_12; +x_11 = lean_unsigned_to_nat(0u); +x_12 = lean_nat_dec_lt(x_11, x_5); +lean_dec(x_5); +if (x_12 == 0) +{ +uint8_t x_13; lean_object* x_14; lean_object* x_15; +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_13 = 0; +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_7); +return x_15; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_alloc_closure((void*)(l_Lean_Elab_Deriving_ToExpr_mkToExprInstanceCmds___boxed), 8, 1); +lean_closure_set(x_16, 0, x_1); +x_17 = lean_alloc_closure((void*)(l_Lean_Elab_Command_liftTermElabM___rarg___boxed), 4, 1); +lean_closure_set(x_17, 0, x_16); +lean_inc(x_3); +lean_inc(x_2); +x_18 = l_Lean_Elab_Command_withFreshMacroScope___rarg(x_17, x_2, x_3, x_7); +if (lean_obj_tag(x_18) == 0) +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; 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); +lean_dec(x_18); +x_21 = lean_box(2); +x_22 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17; +x_23 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_23, 0, x_21); +lean_ctor_set(x_23, 1, x_22); +lean_ctor_set(x_23, 2, x_19); +x_24 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); +lean_closure_set(x_24, 0, x_23); +x_25 = l_Lean_Elab_Deriving_ToExpr_mkToExprInstanceHandler___closed__1; +x_26 = l_Lean_Elab_Command_withScope___rarg(x_25, x_24, x_2, x_3, x_20); +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; uint8_t x_29; lean_object* x_30; +x_28 = lean_ctor_get(x_26, 0); +lean_dec(x_28); +x_29 = 1; +x_30 = lean_box(x_29); +lean_ctor_set(x_26, 0, x_30); +return x_26; +} +else +{ +lean_object* x_31; uint8_t x_32; lean_object* x_33; lean_object* x_34; +x_31 = lean_ctor_get(x_26, 1); +lean_inc(x_31); +lean_dec(x_26); +x_32 = 1; +x_33 = lean_box(x_32); +x_34 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_34, 0, x_33); +lean_ctor_set(x_34, 1, x_31); +return x_34; +} +} +else +{ +uint8_t x_35; +x_35 = !lean_is_exclusive(x_26); +if (x_35 == 0) +{ +return x_26; +} +else +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_26, 0); +x_37 = lean_ctor_get(x_26, 1); +lean_inc(x_37); +lean_inc(x_36); +lean_dec(x_26); +x_38 = lean_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; +lean_dec(x_3); +lean_dec(x_2); +x_39 = !lean_is_exclusive(x_18); +if (x_39 == 0) +{ +return x_18; +} +else +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_40 = lean_ctor_get(x_18, 0); +x_41 = lean_ctor_get(x_18, 1); +lean_inc(x_41); +lean_inc(x_40); +lean_dec(x_18); +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_Lean_isInductive___at_Lean_Elab_Deriving_ToExpr_mkToExprInstanceHandler___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_isInductive___at_Lean_Elab_Deriving_ToExpr_mkToExprInstanceHandler___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_Array_anyMUnsafe_any___at_Lean_Elab_Deriving_ToExpr_mkToExprInstanceHandler___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +size_t x_7; size_t x_8; lean_object* x_9; +x_7 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_8 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_9 = l_Array_anyMUnsafe_any___at_Lean_Elab_Deriving_ToExpr_mkToExprInstanceHandler___spec__2(x_1, x_7, x_8, x_4, x_5, x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +return x_9; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Deriving_ToExpr_mkToExprInstanceHandler), 4, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___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_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__2; +x_2 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__14; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__3; +x_2 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__15; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__4; +x_2 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__4; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("initFn", 6, 6); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__5; +x_2 = l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__6; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__8() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("_@", 2, 2); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__7; +x_2 = l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__8; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__9; +x_2 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___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_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__10; +x_2 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__14; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__11; +x_2 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__15; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__12; +x_2 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__4; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__14() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("_hyg", 4, 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__15() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__13; +x_2 = l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__14; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__16() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__15; +x_2 = lean_unsigned_to_nat(5049u); +x_3 = l_Lean_Name_num___override(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049_(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_2 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__12; +x_3 = l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__1; +x_4 = l_Lean_Elab_registerDerivingHandler(x_2, x_3, x_1); +if (lean_obj_tag(x_4) == 0) +{ +lean_object* x_5; lean_object* x_6; uint8_t x_7; lean_object* x_8; lean_object* x_9; +x_5 = lean_ctor_get(x_4, 1); +lean_inc(x_5); +lean_dec(x_4); +x_6 = l_Lean_Elab_Deriving_ToExpr_mkToExprInstanceCmds___closed__1; +x_7 = 0; +x_8 = l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__16; +x_9 = l_Lean_registerTraceClass(x_6, x_7, x_8, x_5); +return x_9; +} +else +{ +uint8_t x_10; +x_10 = !lean_is_exclusive(x_4); +if (x_10 == 0) +{ +return x_4; +} +else +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_11 = lean_ctor_get(x_4, 0); +x_12 = lean_ctor_get(x_4, 1); +lean_inc(x_12); +lean_inc(x_11); +lean_dec(x_4); +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_object* initialize_Lean_Meta_Transform(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Elab_Deriving_Basic(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Elab_Deriving_Util(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_ToLevel(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_ToExpr(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Lean_Elab_Deriving_ToExpr(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_Transform(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Lean_Elab_Deriving_Basic(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Lean_Elab_Deriving_Util(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Lean_ToLevel(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Lean_ToExpr(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__1 = _init_l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__1(); +lean_mark_persistent(l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__1); +l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__2 = _init_l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__2(); +lean_mark_persistent(l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__2); +l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__3 = _init_l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__3(); +lean_mark_persistent(l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__3); +l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__4 = _init_l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__4(); +lean_mark_persistent(l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__4); +l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__5 = _init_l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__5(); +lean_mark_persistent(l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__5); +l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__6 = _init_l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__6(); +lean_mark_persistent(l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__6); +l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__7 = _init_l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__7(); +lean_mark_persistent(l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__7); +l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__8 = _init_l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__8(); +lean_mark_persistent(l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__8); +l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__9 = _init_l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__9(); +lean_mark_persistent(l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__9); +l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__10 = _init_l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__10(); +lean_mark_persistent(l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__10); +l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__11 = _init_l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__11(); +lean_mark_persistent(l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__11); +l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__12 = _init_l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__12(); +lean_mark_persistent(l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__12); +l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__13 = _init_l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__13(); +lean_mark_persistent(l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__13); +l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__14 = _init_l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__14(); +lean_mark_persistent(l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__14); +l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__15 = _init_l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__15(); +lean_mark_persistent(l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__15); +l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__16 = _init_l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__16(); +lean_mark_persistent(l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__16); +l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17 = _init_l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17(); +lean_mark_persistent(l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17); +l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__1 = _init_l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__1(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__1); +l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__2 = _init_l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__2(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__2); +l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__3 = _init_l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__3(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__3); +l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__4 = _init_l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__4(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__4); +l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__5 = _init_l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__5(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__5); +l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__6 = _init_l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__6(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__6); +l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__7 = _init_l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__7(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__7); +l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__8 = _init_l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__8(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__8); +l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__9 = _init_l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__9(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__9); +l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__10 = _init_l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__10(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__10); +l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__11 = _init_l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__11(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__11); +l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__12 = _init_l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__12(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__12); +l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__13 = _init_l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__13(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__13); +l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__14 = _init_l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__14(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__14); +l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__1___closed__1 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__1___closed__1(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__1___closed__1); +l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__1___closed__2 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__1___closed__2(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__1___closed__2); +l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__1___closed__3 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__1___closed__3(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__1___closed__3); +l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__1___closed__4 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__1___closed__4(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__1___closed__4); +l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__1___closed__5 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__1___closed__5(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__1___closed__5); +l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__1___closed__6 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__1___closed__6(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__1___closed__6); +l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__1___closed__7 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__1___closed__7(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__1___closed__7); +l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__1___closed__8 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__1___closed__8(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__1___closed__8); +l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__1 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__1(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__1); +l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__2 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__2(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__2); +l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__3 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__3(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__3); +l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__4 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__4(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__4); +l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__5 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__5(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__5); +l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__6 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__6(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__6); +l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__7 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__7(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__7); +l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__8 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__8(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__8); +l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__9 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__9(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__9); +l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__10 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__10(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__10); +l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__11 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__11(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__11); +l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__12 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__12(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__12); +l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__13 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__13(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__13); +l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__1 = _init_l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__1(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__1); +l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__2 = _init_l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__2(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__2); +l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__3 = _init_l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__3(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__3); +l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__4 = _init_l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__4(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__4); +l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__5 = _init_l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__5(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__5); +l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__6 = _init_l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__6(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__6); +l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__7 = _init_l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__7(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__7); +l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__8 = _init_l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__8(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__8); +l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__9 = _init_l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__9(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__9); +l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__10 = _init_l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__10(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__10); +l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__11 = _init_l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__11(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__11); +l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__12 = _init_l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__12(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__12); +l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__13 = _init_l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__13(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__13); +l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__14 = _init_l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__14(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__14); +l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__15 = _init_l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__15(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__15); +l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__16 = _init_l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__16(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__16); +l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__17 = _init_l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__17(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__17); +l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__18 = _init_l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__18(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__18); +l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__19 = _init_l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__19(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__19); +l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___boxed__const__1 = _init_l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___boxed__const__1(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___boxed__const__1); +l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__1___closed__1 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__1___closed__1(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__1___closed__1); +l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__1___closed__2 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__1___closed__2(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__1___closed__2); +l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__1___closed__3 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__1___closed__3(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__1___closed__3); +l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__1___closed__4 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__1___closed__4(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__1___closed__4); +l_Lean_getConstInfoCtor___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__2___closed__1 = _init_l_Lean_getConstInfoCtor___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__2___closed__1(); +lean_mark_persistent(l_Lean_getConstInfoCtor___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__2___closed__1); +l_Lean_getConstInfoCtor___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__2___closed__2 = _init_l_Lean_getConstInfoCtor___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__2___closed__2(); +lean_mark_persistent(l_Lean_getConstInfoCtor___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__2___closed__2); +l_Lean_getConstInfoCtor___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__2___closed__3 = _init_l_Lean_getConstInfoCtor___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__2___closed__3(); +lean_mark_persistent(l_Lean_getConstInfoCtor___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__2___closed__3); +l_Lean_getConstInfoCtor___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__2___closed__4 = _init_l_Lean_getConstInfoCtor___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__2___closed__4(); +lean_mark_persistent(l_Lean_getConstInfoCtor___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__2___closed__4); +l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__4___closed__1 = _init_l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__4___closed__1(); +lean_mark_persistent(l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__4___closed__1); +l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__4___closed__2 = _init_l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__4___closed__2(); +lean_mark_persistent(l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__4___closed__2); +l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__4___closed__3 = _init_l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__4___closed__3(); +lean_mark_persistent(l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__4___closed__3); +l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__6___closed__1 = _init_l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__6___closed__1(); +lean_mark_persistent(l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__6___closed__1); +l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__6___closed__2 = _init_l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__6___closed__2(); +lean_mark_persistent(l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__6___closed__2); +l_List_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__7___lambda__1___closed__1 = _init_l_List_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__7___lambda__1___closed__1(); +lean_mark_persistent(l_List_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__7___lambda__1___closed__1); +l_List_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__7___lambda__1___closed__2 = _init_l_List_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__7___lambda__1___closed__2(); +lean_mark_persistent(l_List_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__7___lambda__1___closed__2); +l_List_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__7___lambda__1___closed__3 = _init_l_List_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__7___lambda__1___closed__3(); +lean_mark_persistent(l_List_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__7___lambda__1___closed__3); +l_List_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__7___lambda__1___closed__4 = _init_l_List_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__7___lambda__1___closed__4(); +lean_mark_persistent(l_List_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__7___lambda__1___closed__4); +l_List_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__7___lambda__1___closed__5 = _init_l_List_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__7___lambda__1___closed__5(); +lean_mark_persistent(l_List_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__7___lambda__1___closed__5); +l_List_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__7___lambda__1___closed__6 = _init_l_List_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__7___lambda__1___closed__6(); +lean_mark_persistent(l_List_forIn_x27_loop___at_Lean_Elab_Deriving_ToExpr_mkToExprBody_mkAlts___spec__7___lambda__1___closed__6); +l_Lean_Elab_Deriving_ToExpr_mkToExprBody___closed__1 = _init_l_Lean_Elab_Deriving_ToExpr_mkToExprBody___closed__1(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_mkToExprBody___closed__1); +l_Lean_Elab_Deriving_ToExpr_mkToExprBody___closed__2 = _init_l_Lean_Elab_Deriving_ToExpr_mkToExprBody___closed__2(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_mkToExprBody___closed__2); +l_Lean_Elab_Deriving_ToExpr_mkToExprBody___closed__3 = _init_l_Lean_Elab_Deriving_ToExpr_mkToExprBody___closed__3(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_mkToExprBody___closed__3); +l_Lean_Elab_Deriving_ToExpr_mkToExprBody___closed__4 = _init_l_Lean_Elab_Deriving_ToExpr_mkToExprBody___closed__4(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_mkToExprBody___closed__4); +l_Lean_Elab_Deriving_ToExpr_mkToExprBody___closed__5 = _init_l_Lean_Elab_Deriving_ToExpr_mkToExprBody___closed__5(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_mkToExprBody___closed__5); +l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__1 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__1(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__1); +l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__2 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__2(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__2); +l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__3 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__3(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__3); +l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__4 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__4(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__4); +l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__5 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__5(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__5); +l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__6 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__6(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__6); +l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__7 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__7(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__7); +l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__8 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__8(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__8); +l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__9 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__9(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__9); +l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__10 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__10(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__10); +l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__11 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__11(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__11); +l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__12 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__12(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__12); +l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__13 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__13(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__13); +l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__14 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__14(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__14); +l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__15 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__15(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__15); +l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__16 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__16(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__16); +l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__17 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__17(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__17); +l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__18 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__18(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__18); +l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__19 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__19(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__19); +l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__20 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__20(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__20); +l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__21 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__21(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__21); +l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__22 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__22(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__22); +l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__23 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__23(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__23); +l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__24 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__24(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__24); +l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__25 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__25(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__25); +l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__26 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__26(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__26); +l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__27 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__27(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__27); +l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__28 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__28(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__28); +l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__29 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__29(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__29); +l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__30 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__30(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__30); +l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__31 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__31(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__31); +l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__32 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__32(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__32); +l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__33 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__33(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__33); +l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__34 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__34(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__34); +l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__35 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__35(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__35); +l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__36 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__36(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__36); +l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__37 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__37(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__37); +l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__38 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__38(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__38); +l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__39 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__39(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__39); +l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__40 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__40(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__40); +l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__41 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__41(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__41); +l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__42 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__42(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__42); +l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__1 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__1(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__1); +l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__2 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__2(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__2); +l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__3 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__3(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__3); +l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__4 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__4(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__4); +l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__5 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__5(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__5); +l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__6 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__6(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__6); +l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__7 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__7(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__7); +l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__8 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__8(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__8); +l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__9 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__9(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__9); +l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__10 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__10(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__10); +l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__11 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__11(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__11); +l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__12 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__12(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__12); +l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__13 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__13(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__13); +l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__14 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__14(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__14); +l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__1 = _init_l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__1(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__1); +l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__2 = _init_l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__2(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__2); +l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__3 = _init_l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__3(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__3); +l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__4 = _init_l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__4(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__4); +l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__5 = _init_l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__5(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__5); +l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__6 = _init_l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__6(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__6); +l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__7 = _init_l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__7(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__7); +l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__8 = _init_l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__8(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__8); +l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__9 = _init_l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__9(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__9); +l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__10 = _init_l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__10(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__10); +l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__11 = _init_l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__11(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__11); +l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__12 = _init_l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__12(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__12); +l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__13 = _init_l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__13(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__13); +l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__14 = _init_l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__14(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__14); +l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__15 = _init_l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__15(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__15); +l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__16 = _init_l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__16(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__16); +l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__17 = _init_l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__17(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__17); +l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__18 = _init_l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__18(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__18); +l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__19 = _init_l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__19(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__19); +l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__20 = _init_l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__20(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__20); +l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__21 = _init_l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__21(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__21); +l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__22 = _init_l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__22(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__22); +l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__23 = _init_l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__23(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__23); +l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__24 = _init_l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__24(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__24); +l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__25 = _init_l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__25(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__25); +l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__26 = _init_l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__26(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__26); +l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__27 = _init_l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__27(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__27); +l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__28 = _init_l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__28(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__28); +l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__29 = _init_l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__29(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__29); +l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__30 = _init_l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__30(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__30); +l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__31 = _init_l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__31(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__31); +l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__32 = _init_l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__32(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__32); +l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__33 = _init_l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__33(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__33); +l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__34 = _init_l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__34(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__34); +l_Lean_Elab_Deriving_ToExpr_mkAuxFunctions___closed__1 = _init_l_Lean_Elab_Deriving_ToExpr_mkAuxFunctions___closed__1(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_mkAuxFunctions___closed__1); +l_Lean_Elab_Deriving_ToExpr_mkAuxFunctions___closed__2 = _init_l_Lean_Elab_Deriving_ToExpr_mkAuxFunctions___closed__2(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_mkAuxFunctions___closed__2); +l_Lean_Elab_Deriving_ToExpr_mkAuxFunctions___closed__3 = _init_l_Lean_Elab_Deriving_ToExpr_mkAuxFunctions___closed__3(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_mkAuxFunctions___closed__3); +l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__1___closed__1 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__1___closed__1(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__1___closed__1); +l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__1___closed__2 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__1___closed__2(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__1___closed__2); +l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__1 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__1(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__1); +l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__2 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__2(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__2); +l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__3 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__3(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__3); +l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__4 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__4(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__4); +l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__5 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__5(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__5); +l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__6 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__6(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__6); +l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__7 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__7(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__7); +l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__8 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__8(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__8); +l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__9 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__9(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__9); +l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__10 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__10(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__10); +l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__11 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__11(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__11); +l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__12 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__12(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__12); +l_Lean_Elab_Deriving_ToExpr_mkToExprInstanceCmds___closed__1 = _init_l_Lean_Elab_Deriving_ToExpr_mkToExprInstanceCmds___closed__1(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_mkToExprInstanceCmds___closed__1); +l_Lean_Elab_Deriving_ToExpr_mkToExprInstanceCmds___closed__2 = _init_l_Lean_Elab_Deriving_ToExpr_mkToExprInstanceCmds___closed__2(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_mkToExprInstanceCmds___closed__2); +l_Lean_Elab_Deriving_ToExpr_mkToExprInstanceCmds___closed__3 = _init_l_Lean_Elab_Deriving_ToExpr_mkToExprInstanceCmds___closed__3(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_mkToExprInstanceCmds___closed__3); +l_Lean_Elab_Deriving_ToExpr_mkToExprInstanceCmds___closed__4 = _init_l_Lean_Elab_Deriving_ToExpr_mkToExprInstanceCmds___closed__4(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_mkToExprInstanceCmds___closed__4); +l_Lean_Elab_Deriving_ToExpr_mkToExprInstanceCmds___closed__5 = _init_l_Lean_Elab_Deriving_ToExpr_mkToExprInstanceCmds___closed__5(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_mkToExprInstanceCmds___closed__5); +l_Lean_Elab_Deriving_ToExpr_mkToExprInstanceHandler___lambda__1___closed__1 = _init_l_Lean_Elab_Deriving_ToExpr_mkToExprInstanceHandler___lambda__1___closed__1(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_mkToExprInstanceHandler___lambda__1___closed__1); +l_Lean_Elab_Deriving_ToExpr_mkToExprInstanceHandler___closed__1 = _init_l_Lean_Elab_Deriving_ToExpr_mkToExprInstanceHandler___closed__1(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_mkToExprInstanceHandler___closed__1); +l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__1 = _init_l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__1(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__1); +l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__2 = _init_l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__2(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__2); +l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__3 = _init_l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__3(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__3); +l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__4 = _init_l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__4(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__4); +l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__5 = _init_l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__5(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__5); +l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__6 = _init_l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__6(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__6); +l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__7 = _init_l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__7(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__7); +l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__8 = _init_l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__8(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__8); +l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__9 = _init_l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__9(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__9); +l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__10 = _init_l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__10(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__10); +l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__11 = _init_l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__11(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__11); +l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__12 = _init_l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__12(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__12); +l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__13 = _init_l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__13(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__13); +l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__14 = _init_l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__14(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__14); +l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__15 = _init_l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__15(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__15); +l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__16 = _init_l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__16(); +lean_mark_persistent(l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049____closed__16); +if (builtin) {res = l_Lean_Elab_Deriving_ToExpr_initFn____x40_Lean_Elab_Deriving_ToExpr___hyg_5049_(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/Lean/Elab/Structure.c b/stage0/stdlib/Lean/Elab/Structure.c index afb62c47cdf4..ae938f66b0e8 100644 --- a/stage0/stdlib/Lean/Elab/Structure.c +++ b/stage0/stdlib/Lean/Elab/Structure.c @@ -35,6 +35,7 @@ LEAN_EXPORT uint8_t l___private_Lean_Elab_Structure_0__Lean_Elab_Command_elabFie LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addParentInstances___spec__17___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_structureSyntaxToView___closed__2; static lean_object* l_Lean_Elab_Command_StructView_ctor___closed__4; +static lean_object* l_Lean_addDeclarationRanges___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___spec__8___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_withParents_go___rarg___lambda__5___boxed(lean_object**); 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**); @@ -47,6 +48,7 @@ static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandF 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*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addDefaults___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*); 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*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkRemainingProjections___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_instDecidableEqStructFieldKind___boxed(lean_object*, lean_object*); 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*); @@ -93,6 +95,7 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Structu static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_1023____closed__18; static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_1023____closed__6; 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_Lean_addDeclarationRanges___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___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_registerFailedToInferDefaultValue___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_withParents_go___rarg___closed__3; 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*, lean_object*, lean_object*); @@ -103,6 +106,7 @@ static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStr 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_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_checkResolutionOrder___spec__14___boxed(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_addDefaults___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*); +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_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addParentInstances___spec__26(lean_object*, lean_object*, size_t, size_t); static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_1023____closed__23; @@ -110,6 +114,7 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Stru LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapM___at_Lean_Elab_Command_elabStructureCommand___elambda__1___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_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*); +LEAN_EXPORT lean_object* l_Lean_Elab_getDeclarationRange_x3f___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_checkResolutionOrder___spec__24___closed__9; static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandFields___closed__1; static lean_object* l_Lean_Elab_Command_elabStructureCommand___elambda__1___closed__1; @@ -141,6 +146,7 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Stru 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_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*); static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_checkResultingUniversesForFields___spec__1___closed__4; LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addParentInstances___spec__14(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -196,8 +202,9 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_se LEAN_EXPORT lean_object* l_Lean_throwKernelException___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___spec__3(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__5; extern lean_object* l_Lean_Elab_Command_inductiveElabAttr; -LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___lambda__3(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___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* l_Lean_mkFreshId___at_Lean_Meta_mkFreshExprMVarAt___spec__2___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_addDeclarationRanges___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_copyDefaultValue_x3f_go_x3f___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlFromMAux___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addDefaults___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_mk_array(lean_object*, lean_object*); @@ -209,6 +216,7 @@ static lean_object* l_Lean_Elab_Command_elabStructureCommand___closed__1; lean_object* l_Lean_Syntax_getArgs(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkRemainingProjections___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_checkResolutionOrder___spec__24___closed__8; +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___spec__11(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_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*, lean_object*); static lean_object* l_Lean_Elab_Command_instReprStructFieldInfo___closed__1; static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_checkResolutionOrder___spec__24___closed__2; @@ -270,7 +278,6 @@ uint8_t l_Lean_Linter_getLinterValue(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__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_setStructureParents___at_Lean_Elab_Command_elabStructureCommand___elambda__1___spec__14___closed__2; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addParentInstances___spec__3(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandFields___spec__5___lambda__2___closed__10; LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabStructureCommand___elambda__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* l_Lean_Syntax_node5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -296,6 +303,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_StructFieldInfo_isSubobject___boxed LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addParentInstances___spec__16(size_t, lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_copyNewFieldsFrom_processOveriddenDefaultValues(lean_object*, lean_object*, lean_object*, lean_object*, 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_Structure_0__Lean_Elab_Command_checkResolutionOrder___spec__22(uint8_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*); +uint8_t l_Lean_MapDeclarationExtension_contains___rarg(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_instInhabitedPersistentArrayNode(lean_object*); 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*, lean_object*, lean_object*); @@ -429,6 +437,7 @@ LEAN_EXPORT lean_object* l_Lean_computeStructureResolutionOrder___at___private_L LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_checkResolutionOrder___spec__17(lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addDefaults___spec__9(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_elabFieldTypeValue___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_addDeclarationRanges___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___spec__3___boxed(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__1; static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_1023____closed__1; lean_object* lean_replace_expr(lean_object*, lean_object*); @@ -483,7 +492,6 @@ lean_object* lean_nat_div(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_checkDefaults___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_Elab_Structure_0__Lean_Elab_Command_addProjections(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___lambda__4___closed__3; -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___spec__6(lean_object*, size_t, size_t, lean_object*); static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_addDefaults___closed__5; 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_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_796____closed__6; @@ -506,7 +514,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_co lean_object* l_Lean_LocalContext_setBinderInfo(lean_object*, lean_object*, uint8_t); lean_object* l_Lean_PersistentEnvExtension_getState___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_structureSyntaxToView___spec__2___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___spec__6___boxed(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___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___private_Lean_Elab_Structure_0__Lean_Elab_Command_addParentInstances___spec__20___boxed(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__6; @@ -520,6 +527,7 @@ lean_object* l_Lean_Elab_Term_elabTerm___boxed(lean_object*, lean_object*, lean_ LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_checkDefaults___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_toAttributeKind___boxed(lean_object*, lean_object*, lean_object*); LEAN_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*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_1023____closed__4; 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*); lean_object* l_Lean_Syntax_getKind(lean_object*); @@ -535,7 +543,8 @@ static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_checkDe LEAN_EXPORT lean_object* l_Lean_PersistentArray_foldlM___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_object* l_Lean_MessageData_andList(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_collectUsedFVars___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Elab_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_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_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*); 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*); @@ -606,6 +615,7 @@ static lean_object* l_Lean_Linter_logLint___at___private_Lean_Elab_Structure_0__ LEAN_EXPORT lean_object* l_Lean_computeStructureResolutionOrder_selectParent___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_checkResolutionOrder___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_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; +LEAN_EXPORT lean_object* l_Lean_Elab_addDeclarationRangesFromSyntax___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___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_Data_PersistentArray_0__Lean_PersistentArray_foldlMAux___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addDefaults___spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__5___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabStructureCommand___elambda__1___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -634,6 +644,7 @@ 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_RBNode_insert___at_Lean_NameMap_insert___spec__1___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_computeStructureResolutionOrder___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_checkResolutionOrder___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_addDeclarationRangesFromSyntax___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint64_t l_Lean_Expr_hash(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); @@ -695,7 +706,6 @@ static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Struct LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_withFields_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_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkProjections___boxed(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__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_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_withSaveInfoContext___at_Lean_Elab_Command_elabStructureCommand___elambda__1___spec__2___closed__1; static lean_object* l_Lean_getDocStringText___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__14___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__5___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -743,7 +753,7 @@ LEAN_EXPORT uint8_t l___private_Lean_Elab_Structure_0__Lean_Elab_Command_decEqSt LEAN_EXPORT lean_object* l_Lean_ofExceptKernelException___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___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_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_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_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___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_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___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_Command_checkValidFieldModifier___lambda__3___closed__2; LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_withFields_go___spec__7___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__8___closed__3; @@ -764,6 +774,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_wi uint8_t l_Lean_Environment_contains(lean_object*, lean_object*); uint8_t l_Lean_Level_geq(lean_object*, lean_object*); lean_object* l_Lean_Elab_Command_checkValidInductiveModifier___at_Lean_Elab_Command_elabInductives___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___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*, lean_object*); 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*); lean_object* l_Lean_Option_register___at_Lean_Elab_initFn____x40_Lean_Elab_AutoBound___hyg_6____spec__1(lean_object*, lean_object*, lean_object*, lean_object*); @@ -810,7 +821,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_co 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; 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_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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addDefaults___spec__14(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_instInhabitedStructView___closed__1; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addParentInstances___spec__27(lean_object*, lean_object*, size_t, size_t); @@ -826,6 +837,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_pr static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_copyNewFieldsFrom_copyFields_copy___rarg___closed__4; static lean_object* l_Lean_Elab_elabAttr___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__4___lambda__2___closed__3; lean_object* l_Lean_Elab_Term_addTermInfo_x27(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_FileMap_leanPosToLspPos(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addParentInstances___spec__8(lean_object*, lean_object*, size_t, size_t); 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_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addDefaults___spec__16___closed__2; @@ -839,7 +851,7 @@ LEAN_EXPORT lean_object* l_Lean_computeStructureResolutionOrder_selectParent___a LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkRemainingProjections___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_elabModifiers___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__1___lambda__3___closed__2; static lean_object* l_Lean_Linter_logLint___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandFields___spec__8___closed__3; -LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___lambda__4(lean_object*, 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*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_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(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkRemainingProjections___spec__3___closed__1; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addParentInstances___spec__28___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -857,6 +869,7 @@ static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_withPar static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_1023____closed__28; lean_object* l_Lean_Expr_bindingDomain_x21(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_structureSyntaxToView___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_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_findFieldInfo_x3f___closed__2; static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_withFields_go___rarg___closed__11; LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -884,6 +897,7 @@ static lean_object* l_Lean_setStructureParents___at_Lean_Elab_Command_elabStruct lean_object* l_Lean_CollectMVars_visit(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_checkDefaults___closed__6; static lean_object* l_Lean_Elab_elabAttr___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__4___closed__1; +LEAN_EXPORT lean_object* l_Lean_Elab_addDeclarationRangesFromSyntax___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandFields___closed__4; lean_object* l_Lean_isTracingEnabledFor___at_Lean_Elab_Term_traceAtCmdPos___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); @@ -953,6 +967,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_wi LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addParentInstances___spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_LocalContext_foldlM___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addDefaults___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___private_Lean_Elab_Structure_0__Lean_Elab_Command_copyDefaultValue_x3f_failed___closed__4; +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___spec__10(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addParentInstances___spec__10(lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Elab_withSaveInfoContext___at_Lean_Elab_Command_elabStructureCommand___elambda__1___spec__2___lambda__1___boxed(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(lean_object*); @@ -999,10 +1014,12 @@ static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkRemai uint8_t l_Lean_Expr_isProp(lean_object*); lean_object* lean_nat_mul(lean_object*, lean_object*); lean_object* l_Lean_Elab_Term_registerLevelMVarErrorExprInfo(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_addDeclarationRangesFromSyntax___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_checkResultingUniversesForFields___spec__1___closed__8; static lean_object* l_Lean_Elab_Command_checkValidCtorModifier___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__17___closed__2; static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_copyDefaultValue_x3f_failed___closed__2; lean_object* l_Lean_logWarning___at___private_Lean_Elab_Term_0__Lean_Elab_Term_checkDeprecatedCore___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Syntax_getRange_x3f(lean_object*, uint8_t); static lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandFields___spec__5___lambda__5___closed__3; static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_structureSyntaxToView___spec__2___closed__2; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_withParents_go___spec__2(lean_object*, lean_object*, size_t, size_t); @@ -1016,6 +1033,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_ex LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_withParents_go___rarg___lambda__6___boxed(lean_object**); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Command_elabStructureCommand__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_setReducibilityStatus___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addDefaults___spec__15___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_addDeclarationRanges___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_expandOptDeclSig(lean_object*); extern lean_object* l_Lean_Linter_linter_deprecated; @@ -1031,6 +1049,7 @@ LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Elab_Structu LEAN_EXPORT lean_object* l_Lean_computeStructureResolutionOrder___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addParentInstances___spec__2(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_protectedExt; static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3____closed__20; +extern lean_object* l_Lean_instInhabitedDeclarationRanges; 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; @@ -1157,6 +1176,7 @@ LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Elab_Structu LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_checkResolutionOrder___spec__19(lean_object*, lean_object*, size_t, size_t, lean_object*); static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_checkResolutionOrder___lambda__1___closed__2; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabStructureCommand___elambda__1___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_EXPORT lean_object* l_Lean_Elab_getDeclarationRange_x3f___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_checkValidCtorModifier___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__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*); @@ -1255,6 +1275,7 @@ static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_copyNew 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*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_processSubfields_go___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_getDeclarationRange_x3f___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___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_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_796____closed__4; lean_object* l_panic___at_Lean_Meta_mkPProdFst___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_collectFVars(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1273,6 +1294,7 @@ static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStr static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_registerFailedToInferDefaultValue___closed__3; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addDefaults___spec__5(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapM___at_Lean_Elab_Command_elabStructureCommand___elambda__1___spec__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_getDeclarationRange_x3f___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___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_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_Array_anyMUnsafe_any___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_checkResolutionOrder___spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -25120,7 +25142,703 @@ return x_66; } } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___spec__6(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Elab_getDeclarationRange_x3f___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___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: +{ +uint8_t x_9; lean_object* x_10; +x_9 = 0; +x_10 = l_Lean_Syntax_getRange_x3f(x_1, x_9); +if (lean_obj_tag(x_10) == 0) +{ +lean_object* x_11; lean_object* x_12; +lean_dec(x_6); +x_11 = lean_box(0); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_8); +return x_12; +} +else +{ +uint8_t x_13; +x_13 = !lean_is_exclusive(x_10); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_14 = lean_ctor_get(x_10, 0); +x_15 = lean_ctor_get(x_6, 1); +lean_inc(x_15); +lean_dec(x_6); +x_16 = lean_ctor_get(x_14, 0); +lean_inc(x_16); +lean_inc(x_15); +x_17 = l_Lean_FileMap_toPosition(x_15, x_16); +lean_dec(x_16); +x_18 = lean_ctor_get(x_14, 1); +lean_inc(x_18); +lean_dec(x_14); +lean_inc(x_15); +x_19 = l_Lean_FileMap_toPosition(x_15, x_18); +lean_dec(x_18); +lean_inc(x_17); +x_20 = l_Lean_FileMap_leanPosToLspPos(x_15, x_17); +x_21 = !lean_is_exclusive(x_20); +if (x_21 == 0) +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_22 = lean_ctor_get(x_20, 1); +x_23 = lean_ctor_get(x_20, 0); +lean_dec(x_23); +lean_inc(x_19); +x_24 = l_Lean_FileMap_leanPosToLspPos(x_15, x_19); +lean_dec(x_15); +x_25 = lean_ctor_get(x_24, 1); +lean_inc(x_25); +lean_dec(x_24); +x_26 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_26, 0, x_17); +lean_ctor_set(x_26, 1, x_22); +lean_ctor_set(x_26, 2, x_19); +lean_ctor_set(x_26, 3, x_25); +lean_ctor_set(x_10, 0, x_26); +lean_ctor_set(x_20, 1, x_8); +lean_ctor_set(x_20, 0, x_10); +return x_20; +} +else +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_27 = lean_ctor_get(x_20, 1); +lean_inc(x_27); +lean_dec(x_20); +lean_inc(x_19); +x_28 = l_Lean_FileMap_leanPosToLspPos(x_15, x_19); +lean_dec(x_15); +x_29 = lean_ctor_get(x_28, 1); +lean_inc(x_29); +lean_dec(x_28); +x_30 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_30, 0, x_17); +lean_ctor_set(x_30, 1, x_27); +lean_ctor_set(x_30, 2, x_19); +lean_ctor_set(x_30, 3, x_29); +lean_ctor_set(x_10, 0, x_30); +x_31 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_31, 0, x_10); +lean_ctor_set(x_31, 1, x_8); +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; lean_object* x_45; +x_32 = lean_ctor_get(x_10, 0); +lean_inc(x_32); +lean_dec(x_10); +x_33 = lean_ctor_get(x_6, 1); +lean_inc(x_33); +lean_dec(x_6); +x_34 = lean_ctor_get(x_32, 0); +lean_inc(x_34); +lean_inc(x_33); +x_35 = l_Lean_FileMap_toPosition(x_33, x_34); +lean_dec(x_34); +x_36 = lean_ctor_get(x_32, 1); +lean_inc(x_36); +lean_dec(x_32); +lean_inc(x_33); +x_37 = l_Lean_FileMap_toPosition(x_33, x_36); +lean_dec(x_36); +lean_inc(x_35); +x_38 = l_Lean_FileMap_leanPosToLspPos(x_33, x_35); +x_39 = lean_ctor_get(x_38, 1); +lean_inc(x_39); +if (lean_is_exclusive(x_38)) { + lean_ctor_release(x_38, 0); + lean_ctor_release(x_38, 1); + x_40 = x_38; +} else { + lean_dec_ref(x_38); + x_40 = lean_box(0); +} +lean_inc(x_37); +x_41 = l_Lean_FileMap_leanPosToLspPos(x_33, x_37); +lean_dec(x_33); +x_42 = lean_ctor_get(x_41, 1); +lean_inc(x_42); +lean_dec(x_41); +x_43 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_43, 0, x_35); +lean_ctor_set(x_43, 1, x_39); +lean_ctor_set(x_43, 2, x_37); +lean_ctor_set(x_43, 3, x_42); +x_44 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_44, 0, x_43); +if (lean_is_scalar(x_40)) { + x_45 = lean_alloc_ctor(0, 2, 0); +} else { + x_45 = x_40; +} +lean_ctor_set(x_45, 0, x_44); +lean_ctor_set(x_45, 1, x_8); +return x_45; +} +} +} +} +static lean_object* _init_l_Lean_addDeclarationRanges___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___spec__8___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_declRangeExt; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_addDeclarationRanges___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___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; uint8_t x_11; +x_10 = lean_st_ref_get(x_8, x_9); +x_11 = !lean_is_exclusive(x_10); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_12 = lean_ctor_get(x_10, 0); +x_13 = lean_ctor_get(x_10, 1); +x_14 = lean_ctor_get(x_12, 0); +lean_inc(x_14); +lean_dec(x_12); +x_15 = l_Lean_instInhabitedDeclarationRanges; +x_16 = l_Lean_addDeclarationRanges___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___spec__8___closed__1; +lean_inc(x_1); +x_17 = l_Lean_MapDeclarationExtension_contains___rarg(x_15, x_16, x_14, x_1); +lean_dec(x_14); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; +lean_free_object(x_10); +x_18 = lean_st_ref_take(x_8, x_13); +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_18, 1); +lean_inc(x_20); +lean_dec(x_18); +x_21 = !lean_is_exclusive(x_19); +if (x_21 == 0) +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; 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_22 = lean_ctor_get(x_19, 0); +x_23 = lean_ctor_get(x_19, 4); +lean_dec(x_23); +x_24 = l_Lean_MapDeclarationExtension_insert___rarg(x_16, x_22, x_1, x_2); +x_25 = l_Lean_setEnv___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___spec__5___closed__3; +lean_ctor_set(x_19, 4, x_25); +lean_ctor_set(x_19, 0, x_24); +x_26 = lean_st_ref_set(x_8, x_19, x_20); +x_27 = lean_ctor_get(x_26, 1); +lean_inc(x_27); +lean_dec(x_26); +x_28 = lean_st_ref_take(x_6, x_27); +x_29 = lean_ctor_get(x_28, 0); +lean_inc(x_29); +x_30 = lean_ctor_get(x_28, 1); +lean_inc(x_30); +lean_dec(x_28); +x_31 = !lean_is_exclusive(x_29); +if (x_31 == 0) +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; +x_32 = lean_ctor_get(x_29, 1); +lean_dec(x_32); +x_33 = l_Lean_setEnv___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___spec__5___closed__4; +lean_ctor_set(x_29, 1, x_33); +x_34 = lean_st_ref_set(x_6, 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; +x_41 = lean_ctor_get(x_29, 0); +x_42 = lean_ctor_get(x_29, 2); +x_43 = lean_ctor_get(x_29, 3); +x_44 = lean_ctor_get(x_29, 4); +lean_inc(x_44); +lean_inc(x_43); +lean_inc(x_42); +lean_inc(x_41); +lean_dec(x_29); +x_45 = l_Lean_setEnv___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___spec__5___closed__4; +x_46 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_46, 0, x_41); +lean_ctor_set(x_46, 1, x_45); +lean_ctor_set(x_46, 2, x_42); +lean_ctor_set(x_46, 3, x_43); +lean_ctor_set(x_46, 4, x_44); +x_47 = lean_st_ref_set(x_6, x_46, x_30); +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; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; +x_52 = lean_ctor_get(x_19, 0); +x_53 = lean_ctor_get(x_19, 1); +x_54 = lean_ctor_get(x_19, 2); +x_55 = lean_ctor_get(x_19, 3); +x_56 = lean_ctor_get(x_19, 5); +x_57 = lean_ctor_get(x_19, 6); +x_58 = lean_ctor_get(x_19, 7); +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_19); +x_59 = l_Lean_MapDeclarationExtension_insert___rarg(x_16, x_52, x_1, x_2); +x_60 = l_Lean_setEnv___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___spec__5___closed__3; +x_61 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_61, 0, x_59); +lean_ctor_set(x_61, 1, x_53); +lean_ctor_set(x_61, 2, x_54); +lean_ctor_set(x_61, 3, x_55); +lean_ctor_set(x_61, 4, x_60); +lean_ctor_set(x_61, 5, x_56); +lean_ctor_set(x_61, 6, x_57); +lean_ctor_set(x_61, 7, x_58); +x_62 = lean_st_ref_set(x_8, x_61, x_20); +x_63 = lean_ctor_get(x_62, 1); +lean_inc(x_63); +lean_dec(x_62); +x_64 = lean_st_ref_take(x_6, 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); +x_68 = lean_ctor_get(x_65, 2); +lean_inc(x_68); +x_69 = lean_ctor_get(x_65, 3); +lean_inc(x_69); +x_70 = lean_ctor_get(x_65, 4); +lean_inc(x_70); +if (lean_is_exclusive(x_65)) { + lean_ctor_release(x_65, 0); + lean_ctor_release(x_65, 1); + lean_ctor_release(x_65, 2); + lean_ctor_release(x_65, 3); + lean_ctor_release(x_65, 4); + x_71 = x_65; +} else { + lean_dec_ref(x_65); + x_71 = lean_box(0); +} +x_72 = l_Lean_setEnv___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___spec__5___closed__4; +if (lean_is_scalar(x_71)) { + x_73 = lean_alloc_ctor(0, 5, 0); +} else { + x_73 = x_71; +} +lean_ctor_set(x_73, 0, x_67); +lean_ctor_set(x_73, 1, x_72); +lean_ctor_set(x_73, 2, x_68); +lean_ctor_set(x_73, 3, x_69); +lean_ctor_set(x_73, 4, x_70); +x_74 = lean_st_ref_set(x_6, x_73, x_66); +x_75 = lean_ctor_get(x_74, 1); +lean_inc(x_75); +if (lean_is_exclusive(x_74)) { + lean_ctor_release(x_74, 0); + lean_ctor_release(x_74, 1); + x_76 = x_74; +} else { + lean_dec_ref(x_74); + x_76 = lean_box(0); +} +x_77 = lean_box(0); +if (lean_is_scalar(x_76)) { + x_78 = lean_alloc_ctor(0, 2, 0); +} else { + x_78 = x_76; +} +lean_ctor_set(x_78, 0, x_77); +lean_ctor_set(x_78, 1, x_75); +return x_78; +} +} +else +{ +lean_object* x_79; +lean_dec(x_2); +lean_dec(x_1); +x_79 = lean_box(0); +lean_ctor_set(x_10, 0, x_79); +return x_10; +} +} +else +{ +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_80 = lean_ctor_get(x_10, 0); +x_81 = lean_ctor_get(x_10, 1); +lean_inc(x_81); +lean_inc(x_80); +lean_dec(x_10); +x_82 = lean_ctor_get(x_80, 0); +lean_inc(x_82); +lean_dec(x_80); +x_83 = l_Lean_instInhabitedDeclarationRanges; +x_84 = l_Lean_addDeclarationRanges___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___spec__8___closed__1; +lean_inc(x_1); +x_85 = l_Lean_MapDeclarationExtension_contains___rarg(x_83, x_84, x_82, x_1); +lean_dec(x_82); +if (x_85 == 0) +{ +lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; 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_86 = lean_st_ref_take(x_8, x_81); +x_87 = lean_ctor_get(x_86, 0); +lean_inc(x_87); +x_88 = lean_ctor_get(x_86, 1); +lean_inc(x_88); +lean_dec(x_86); +x_89 = lean_ctor_get(x_87, 0); +lean_inc(x_89); +x_90 = lean_ctor_get(x_87, 1); +lean_inc(x_90); +x_91 = lean_ctor_get(x_87, 2); +lean_inc(x_91); +x_92 = lean_ctor_get(x_87, 3); +lean_inc(x_92); +x_93 = lean_ctor_get(x_87, 5); +lean_inc(x_93); +x_94 = lean_ctor_get(x_87, 6); +lean_inc(x_94); +x_95 = lean_ctor_get(x_87, 7); +lean_inc(x_95); +if (lean_is_exclusive(x_87)) { + lean_ctor_release(x_87, 0); + lean_ctor_release(x_87, 1); + lean_ctor_release(x_87, 2); + lean_ctor_release(x_87, 3); + lean_ctor_release(x_87, 4); + lean_ctor_release(x_87, 5); + lean_ctor_release(x_87, 6); + lean_ctor_release(x_87, 7); + x_96 = x_87; +} else { + lean_dec_ref(x_87); + x_96 = lean_box(0); +} +x_97 = l_Lean_MapDeclarationExtension_insert___rarg(x_84, x_89, x_1, x_2); +x_98 = l_Lean_setEnv___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___spec__5___closed__3; +if (lean_is_scalar(x_96)) { + x_99 = lean_alloc_ctor(0, 8, 0); +} else { + x_99 = x_96; +} +lean_ctor_set(x_99, 0, x_97); +lean_ctor_set(x_99, 1, x_90); +lean_ctor_set(x_99, 2, x_91); +lean_ctor_set(x_99, 3, x_92); +lean_ctor_set(x_99, 4, x_98); +lean_ctor_set(x_99, 5, x_93); +lean_ctor_set(x_99, 6, x_94); +lean_ctor_set(x_99, 7, x_95); +x_100 = lean_st_ref_set(x_8, x_99, x_88); +x_101 = lean_ctor_get(x_100, 1); +lean_inc(x_101); +lean_dec(x_100); +x_102 = lean_st_ref_take(x_6, x_101); +x_103 = lean_ctor_get(x_102, 0); +lean_inc(x_103); +x_104 = lean_ctor_get(x_102, 1); +lean_inc(x_104); +lean_dec(x_102); +x_105 = lean_ctor_get(x_103, 0); +lean_inc(x_105); +x_106 = lean_ctor_get(x_103, 2); +lean_inc(x_106); +x_107 = lean_ctor_get(x_103, 3); +lean_inc(x_107); +x_108 = lean_ctor_get(x_103, 4); +lean_inc(x_108); +if (lean_is_exclusive(x_103)) { + lean_ctor_release(x_103, 0); + lean_ctor_release(x_103, 1); + lean_ctor_release(x_103, 2); + lean_ctor_release(x_103, 3); + lean_ctor_release(x_103, 4); + x_109 = x_103; +} else { + lean_dec_ref(x_103); + x_109 = lean_box(0); +} +x_110 = l_Lean_setEnv___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___spec__5___closed__4; +if (lean_is_scalar(x_109)) { + x_111 = lean_alloc_ctor(0, 5, 0); +} else { + x_111 = x_109; +} +lean_ctor_set(x_111, 0, x_105); +lean_ctor_set(x_111, 1, x_110); +lean_ctor_set(x_111, 2, x_106); +lean_ctor_set(x_111, 3, x_107); +lean_ctor_set(x_111, 4, x_108); +x_112 = lean_st_ref_set(x_6, x_111, x_104); +x_113 = lean_ctor_get(x_112, 1); +lean_inc(x_113); +if (lean_is_exclusive(x_112)) { + lean_ctor_release(x_112, 0); + lean_ctor_release(x_112, 1); + x_114 = x_112; +} else { + lean_dec_ref(x_112); + x_114 = lean_box(0); +} +x_115 = lean_box(0); +if (lean_is_scalar(x_114)) { + x_116 = lean_alloc_ctor(0, 2, 0); +} else { + x_116 = x_114; +} +lean_ctor_set(x_116, 0, x_115); +lean_ctor_set(x_116, 1, x_113); +return x_116; +} +else +{ +lean_object* x_117; lean_object* x_118; +lean_dec(x_2); +lean_dec(x_1); +x_117 = lean_box(0); +x_118 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_118, 0, x_117); +lean_ctor_set(x_118, 1, x_81); +return x_118; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_addDeclarationRangesFromSyntax___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; lean_object* x_12; +lean_inc(x_8); +x_11 = l_Lean_Elab_getDeclarationRange_x3f___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___spec__7(x_2, 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); +if (lean_obj_tag(x_12) == 0) +{ +uint8_t x_13; +lean_dec(x_8); +lean_dec(x_1); +x_13 = !lean_is_exclusive(x_11); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; +x_14 = lean_ctor_get(x_11, 0); +lean_dec(x_14); +x_15 = lean_box(0); +lean_ctor_set(x_11, 0, x_15); +return x_11; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_11, 1); +lean_inc(x_16); +lean_dec(x_11); +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; +x_19 = lean_ctor_get(x_11, 1); +lean_inc(x_19); +lean_dec(x_11); +x_20 = lean_ctor_get(x_12, 0); +lean_inc(x_20); +lean_dec(x_12); +lean_inc(x_8); +x_21 = l_Lean_Elab_getDeclarationRange_x3f___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___spec__7(x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_19); +x_22 = lean_ctor_get(x_21, 0); +lean_inc(x_22); +if (lean_obj_tag(x_22) == 0) +{ +uint8_t x_23; +x_23 = !lean_is_exclusive(x_21); +if (x_23 == 0) +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_21, 1); +x_25 = lean_ctor_get(x_21, 0); +lean_dec(x_25); +lean_inc(x_20); +lean_ctor_set(x_21, 1, x_20); +lean_ctor_set(x_21, 0, x_20); +x_26 = l_Lean_addDeclarationRanges___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___spec__8(x_1, x_21, x_4, x_5, x_6, x_7, x_8, x_9, x_24); +lean_dec(x_8); +return x_26; +} +else +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_21, 1); +lean_inc(x_27); +lean_dec(x_21); +lean_inc(x_20); +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_20); +lean_ctor_set(x_28, 1, x_20); +x_29 = l_Lean_addDeclarationRanges___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___spec__8(x_1, x_28, x_4, x_5, x_6, x_7, x_8, x_9, x_27); +lean_dec(x_8); +return x_29; +} +} +else +{ +uint8_t x_30; +x_30 = !lean_is_exclusive(x_21); +if (x_30 == 0) +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_31 = lean_ctor_get(x_21, 1); +x_32 = lean_ctor_get(x_21, 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_21, 1, x_33); +lean_ctor_set(x_21, 0, x_20); +x_34 = l_Lean_addDeclarationRanges___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___spec__8(x_1, x_21, x_4, x_5, x_6, x_7, x_8, x_9, x_31); +lean_dec(x_8); +return x_34; +} +else +{ +lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_35 = lean_ctor_get(x_21, 1); +lean_inc(x_35); +lean_dec(x_21); +x_36 = lean_ctor_get(x_22, 0); +lean_inc(x_36); +lean_dec(x_22); +x_37 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_37, 0, x_20); +lean_ctor_set(x_37, 1, x_36); +x_38 = l_Lean_addDeclarationRanges___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___spec__8(x_1, x_37, x_4, x_5, x_6, x_7, x_8, x_9, x_35); +lean_dec(x_8); +return x_38; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___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_object* x_13, lean_object* x_14) { +_start: +{ +uint8_t x_15; +x_15 = lean_usize_dec_lt(x_6, x_5); +if (x_15 == 0) +{ +lean_object* x_16; +lean_dec(x_12); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_7); +lean_ctor_set(x_16, 1, x_14); +return x_16; +} +else +{ +lean_object* x_17; uint8_t x_18; +lean_dec(x_7); +x_17 = lean_array_uget(x_4, x_6); +x_18 = l_Lean_Elab_Command_StructFieldInfo_isSubobject(x_17); +if (x_18 == 0) +{ +size_t x_19; size_t x_20; lean_object* x_21; +lean_dec(x_17); +x_19 = 1; +x_20 = lean_usize_add(x_6, x_19); +x_21 = lean_box(0); +x_6 = x_20; +x_7 = x_21; +goto _start; +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; size_t x_28; size_t x_29; lean_object* x_30; +x_23 = lean_ctor_get(x_17, 2); +lean_inc(x_23); +x_24 = lean_ctor_get(x_2, 0); +x_25 = lean_ctor_get(x_17, 0); +lean_inc(x_25); +lean_dec(x_17); +lean_inc(x_12); +x_26 = l_Lean_Elab_addDeclarationRangesFromSyntax___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___spec__6(x_23, x_24, x_25, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +lean_dec(x_25); +x_27 = lean_ctor_get(x_26, 1); +lean_inc(x_27); +lean_dec(x_26); +x_28 = 1; +x_29 = lean_usize_add(x_6, x_28); +x_30 = lean_box(0); +x_6 = x_29; +x_7 = x_30; +x_14 = x_27; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___spec__10(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { _start: { uint8_t x_5; @@ -25153,7 +25871,7 @@ return x_4; } } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, size_t x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___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_object* x_8, 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; @@ -25298,36 +26016,36 @@ x_14 = 0; x_15 = lean_st_ref_get(x_9, x_10); if (x_13 == 0) { -lean_object* x_37; +lean_object* x_46; lean_dec(x_11); -x_37 = l_Lean_Elab_elabAttrs___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__3___closed__1; -x_16 = x_37; -goto block_36; +x_46 = l_Lean_Elab_elabAttrs___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__3___closed__1; +x_16 = x_46; +goto block_45; } else { -uint8_t x_38; -x_38 = lean_nat_dec_le(x_11, x_11); -if (x_38 == 0) +uint8_t x_47; +x_47 = lean_nat_dec_le(x_11, x_11); +if (x_47 == 0) { -lean_object* x_39; +lean_object* x_48; lean_dec(x_11); -x_39 = l_Lean_Elab_elabAttrs___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__3___closed__1; -x_16 = x_39; -goto block_36; +x_48 = l_Lean_Elab_elabAttrs___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__3___closed__1; +x_16 = x_48; +goto block_45; } else { -size_t x_40; lean_object* x_41; lean_object* x_42; -x_40 = lean_usize_of_nat(x_11); +size_t x_49; lean_object* x_50; lean_object* x_51; +x_49 = lean_usize_of_nat(x_11); lean_dec(x_11); -x_41 = l_Lean_Elab_elabAttrs___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__3___closed__1; -x_42 = l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___spec__6(x_1, x_14, x_40, x_41); -x_16 = x_42; -goto block_36; +x_50 = l_Lean_Elab_elabAttrs___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__3___closed__1; +x_51 = l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___spec__10(x_1, x_14, x_49, x_50); +x_16 = x_51; +goto block_45; } } -block_36: +block_45: { lean_object* x_17; lean_object* x_18; size_t x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; lean_object* x_27; lean_object* x_28; x_17 = lean_ctor_get(x_15, 0); @@ -25350,46 +26068,72 @@ x_24 = lean_ctor_get(x_23, 4); lean_inc(x_24); x_25 = lean_array_to_list(x_20); x_26 = lean_ctor_get_uint8(x_23, sizeof(void*)*11); -lean_dec(x_23); x_27 = lean_mk_projections(x_21, x_24, x_25, x_26); lean_inc(x_8); lean_inc(x_4); x_28 = l_Lean_ofExceptKernelException___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___spec__2(x_27, x_4, x_5, x_6, x_7, x_8, x_9, x_18); if (lean_obj_tag(x_28) == 0) { -lean_object* x_29; lean_object* x_30; lean_object* x_31; +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; size_t x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; 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_setEnv___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___spec__5(x_29, x_4, x_5, x_6, x_7, x_8, x_9, x_30); -lean_dec(x_8); +x_32 = lean_ctor_get(x_31, 1); +lean_inc(x_32); +lean_dec(x_31); +x_33 = lean_box(0); +x_34 = lean_array_size(x_1); +x_35 = lean_box(0); +x_36 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___spec__9(x_1, x_23, x_33, x_1, x_34, x_14, x_35, x_4, x_5, x_6, x_7, x_8, x_9, x_32); lean_dec(x_4); -return x_31; +lean_dec(x_23); +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_35); +return x_36; } else { -uint8_t x_32; +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_35); +lean_ctor_set(x_40, 1, x_39); +return x_40; +} +} +else +{ +uint8_t x_41; +lean_dec(x_23); lean_dec(x_8); lean_dec(x_4); -x_32 = !lean_is_exclusive(x_28); -if (x_32 == 0) +x_41 = !lean_is_exclusive(x_28); +if (x_41 == 0) { return x_28; } 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_inc(x_34); -lean_inc(x_33); +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_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; +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; } } } @@ -25488,7 +26232,7 @@ lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -x_59 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___spec__7(x_2, x_15, x_58, x_2, x_17, x_18, x_58, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_59 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___spec__11(x_2, x_15, x_58, x_2, x_17, x_18, 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_60; lean_object* x_61; @@ -25732,7 +26476,71 @@ lean_dec(x_2); return x_9; } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___spec__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Elab_getDeclarationRange_x3f___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___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: +{ +lean_object* x_9; +x_9 = l_Lean_Elab_getDeclarationRange_x3f___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___spec__7(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_7); +lean_dec(x_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_addDeclarationRanges___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___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: +{ +lean_object* x_10; +x_10 = l_Lean_addDeclarationRanges___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___spec__8(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_addDeclarationRangesFromSyntax___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___spec__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_Elab_addDeclarationRangesFromSyntax___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___spec__6(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___spec__9___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +size_t x_15; size_t x_16; lean_object* x_17; +x_15 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_16 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_17 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___spec__9(x_1, x_2, x_3, x_4, x_15, x_16, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +lean_dec(x_13); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_17; +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___spec__10___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { size_t x_5; size_t x_6; lean_object* x_7; @@ -25740,12 +26548,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___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___spec__6(x_1, x_5, x_6, x_4); +x_7 = l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___spec__10(x_1, x_5, x_6, x_4); lean_dec(x_1); return x_7; } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___spec__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___spec__11___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { _start: { size_t x_15; size_t x_16; lean_object* x_17; @@ -25753,7 +26561,7 @@ x_15 = lean_unbox_usize(x_5); lean_dec(x_5); x_16 = lean_unbox_usize(x_6); lean_dec(x_6); -x_17 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___spec__7(x_1, x_2, x_3, x_4, x_15, x_16, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +x_17 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___spec__11(x_1, x_2, x_3, x_4, x_15, x_16, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); lean_dec(x_9); lean_dec(x_8); lean_dec(x_4); @@ -30766,7 +31574,7 @@ static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_s lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_Elab_Command_StructView_ctor___closed__1; x_2 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_setSourceInstImplicit___closed__1; -x_3 = lean_unsigned_to_nat(773u); +x_3 = lean_unsigned_to_nat(776u); x_4 = lean_unsigned_to_nat(9u); x_5 = l_Lean_Elab_Command_StructView_ctor___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -31154,7 +31962,7 @@ static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_Elab_Command_StructView_ctor___closed__1; x_2 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent_copyFields___spec__2___closed__1; -x_3 = lean_unsigned_to_nat(798u); +x_3 = lean_unsigned_to_nat(801u); x_4 = lean_unsigned_to_nat(77u); x_5 = l_Lean_Elab_Command_StructView_ctor___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -31357,254 +32165,885 @@ return x_53; } } } -else -{ -lean_object* x_54; -lean_inc(x_14); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_1); -x_54 = l_Lean_Meta_mkProjection(x_1, x_18, x_11, x_12, x_13, x_14, x_15); -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); -lean_dec(x_54); -x_57 = l_Lean_Expr_app___override(x_10, x_55); -x_58 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_58, 0, x_57); -x_19 = x_58; -x_20 = x_56; -goto block_27; +else +{ +lean_object* x_54; +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_1); +x_54 = l_Lean_Meta_mkProjection(x_1, x_18, x_11, x_12, x_13, x_14, x_15); +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); +lean_dec(x_54); +x_57 = l_Lean_Expr_app___override(x_10, x_55); +x_58 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_58, 0, x_57); +x_19 = x_58; +x_20 = x_56; +goto block_27; +} +else +{ +uint8_t x_59; +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_59 = !lean_is_exclusive(x_54); +if (x_59 == 0) +{ +return x_54; +} +else +{ +lean_object* x_60; lean_object* x_61; lean_object* x_62; +x_60 = lean_ctor_get(x_54, 0); +x_61 = lean_ctor_get(x_54, 1); +lean_inc(x_61); +lean_inc(x_60); +lean_dec(x_54); +x_62 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_62, 0, x_60); +lean_ctor_set(x_62, 1, x_61); +return x_62; +} +} +} +block_27: +{ +if (lean_obj_tag(x_19) == 0) +{ +lean_object* x_21; lean_object* x_22; +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_21 = lean_ctor_get(x_19, 0); +lean_inc(x_21); +lean_dec(x_19); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_20); +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_9, x_24); +x_9 = x_25; +x_10 = x_23; +x_15 = x_20; +goto _start; +} +} +} +} +} +static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent_copyFields___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_Elab_Command_StructView_ctor___closed__1; +x_2 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent_copyFields___spec__2___closed__1; +x_3 = lean_unsigned_to_nat(792u); +x_4 = lean_unsigned_to_nat(68u); +x_5 = l_Lean_Elab_Command_StructView_ctor___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___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent_copyFields(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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_Expr_getAppFn(x_4); +if (lean_obj_tag(x_10) == 4) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; size_t x_27; size_t x_28; lean_object* x_29; +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); +lean_inc(x_2); +x_13 = l_Lean_getStructureCtor(x_2, x_11); +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +lean_dec(x_13); +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +lean_dec(x_14); +x_16 = l_Lean_Expr_const___override(x_15, x_12); +x_17 = lean_unsigned_to_nat(0u); +x_18 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_4, x_17); +x_19 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_getFieldType___lambda__2___closed__3; +lean_inc(x_18); +x_20 = lean_mk_array(x_18, x_19); +x_21 = lean_unsigned_to_nat(1u); +x_22 = lean_nat_sub(x_18, x_21); +lean_dec(x_18); +x_23 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_4, x_20, x_22); +x_24 = l_Lean_mkAppN(x_16, x_23); +lean_dec(x_23); +lean_inc(x_11); +x_25 = l_Lean_getStructureFields(x_2, x_11); +x_26 = lean_box(0); +x_27 = lean_array_size(x_25); +x_28 = 0; +x_29 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent_copyFields___spec__2(x_1, x_2, x_3, x_11, x_25, x_26, x_25, x_27, x_28, x_24, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_25); +if (lean_obj_tag(x_29) == 0) +{ +uint8_t x_30; +x_30 = !lean_is_exclusive(x_29); +if (x_30 == 0) +{ +return x_29; +} +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_29, 0); +x_32 = lean_ctor_get(x_29, 1); +lean_inc(x_32); +lean_inc(x_31); +lean_dec(x_29); +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_31); +lean_ctor_set(x_33, 1, x_32); +return x_33; +} +} +else +{ +uint8_t x_34; +x_34 = !lean_is_exclusive(x_29); +if (x_34 == 0) +{ +return x_29; +} +else +{ +lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_35 = lean_ctor_get(x_29, 0); +x_36 = lean_ctor_get(x_29, 1); +lean_inc(x_36); +lean_inc(x_35); +lean_dec(x_29); +x_37 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_37, 0, x_35); +lean_ctor_set(x_37, 1, x_36); +return x_37; +} +} +} +else +{ +lean_object* x_38; lean_object* x_39; +lean_dec(x_10); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_38 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent_copyFields___closed__1; +x_39 = l_panic___at_Lean_Meta_mkPProdFst___spec__1(x_38, x_5, x_6, x_7, x_8, x_9); +return x_39; +} +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent_copyFields___spec__2___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; +x_12 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent_copyFields___spec__2___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_1); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent_copyFields___spec__2___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent_copyFields___spec__2___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_5); +lean_dec(x_3); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent_copyFields___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +_start: +{ +size_t x_16; size_t x_17; lean_object* x_18; +x_16 = lean_unbox_usize(x_8); +lean_dec(x_8); +x_17 = lean_unbox_usize(x_9); +lean_dec(x_9); +x_18 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent_copyFields___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_16, x_17, x_10, x_11, x_12, x_13, x_14, x_15); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +return x_18; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent_copyFields___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent_copyFields(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_Elab_getDeclarationRange_x3f___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +uint8_t x_7; lean_object* x_8; +x_7 = 0; +x_8 = l_Lean_Syntax_getRange_x3f(x_1, x_7); +if (lean_obj_tag(x_8) == 0) +{ +lean_object* x_9; lean_object* x_10; +lean_dec(x_4); +x_9 = lean_box(0); +x_10 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_10, 0, x_9); +lean_ctor_set(x_10, 1, x_6); +return x_10; +} +else +{ +uint8_t x_11; +x_11 = !lean_is_exclusive(x_8); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; +x_12 = lean_ctor_get(x_8, 0); +x_13 = lean_ctor_get(x_4, 1); +lean_inc(x_13); +lean_dec(x_4); +x_14 = lean_ctor_get(x_12, 0); +lean_inc(x_14); +lean_inc(x_13); +x_15 = l_Lean_FileMap_toPosition(x_13, x_14); +lean_dec(x_14); +x_16 = lean_ctor_get(x_12, 1); +lean_inc(x_16); +lean_dec(x_12); +lean_inc(x_13); +x_17 = l_Lean_FileMap_toPosition(x_13, x_16); +lean_dec(x_16); +lean_inc(x_15); +x_18 = l_Lean_FileMap_leanPosToLspPos(x_13, x_15); +x_19 = !lean_is_exclusive(x_18); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_20 = lean_ctor_get(x_18, 1); +x_21 = lean_ctor_get(x_18, 0); +lean_dec(x_21); +lean_inc(x_17); +x_22 = l_Lean_FileMap_leanPosToLspPos(x_13, x_17); +lean_dec(x_13); +x_23 = lean_ctor_get(x_22, 1); +lean_inc(x_23); +lean_dec(x_22); +x_24 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_24, 0, x_15); +lean_ctor_set(x_24, 1, x_20); +lean_ctor_set(x_24, 2, x_17); +lean_ctor_set(x_24, 3, x_23); +lean_ctor_set(x_8, 0, x_24); +lean_ctor_set(x_18, 1, x_6); +lean_ctor_set(x_18, 0, x_8); +return x_18; +} +else +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_25 = lean_ctor_get(x_18, 1); +lean_inc(x_25); +lean_dec(x_18); +lean_inc(x_17); +x_26 = l_Lean_FileMap_leanPosToLspPos(x_13, x_17); +lean_dec(x_13); +x_27 = lean_ctor_get(x_26, 1); +lean_inc(x_27); +lean_dec(x_26); +x_28 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_28, 0, x_15); +lean_ctor_set(x_28, 1, x_25); +lean_ctor_set(x_28, 2, x_17); +lean_ctor_set(x_28, 3, x_27); +lean_ctor_set(x_8, 0, x_28); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_8); +lean_ctor_set(x_29, 1, x_6); +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; +x_30 = lean_ctor_get(x_8, 0); +lean_inc(x_30); +lean_dec(x_8); +x_31 = lean_ctor_get(x_4, 1); +lean_inc(x_31); +lean_dec(x_4); +x_32 = lean_ctor_get(x_30, 0); +lean_inc(x_32); +lean_inc(x_31); +x_33 = l_Lean_FileMap_toPosition(x_31, x_32); +lean_dec(x_32); +x_34 = lean_ctor_get(x_30, 1); +lean_inc(x_34); +lean_dec(x_30); +lean_inc(x_31); +x_35 = l_Lean_FileMap_toPosition(x_31, x_34); +lean_dec(x_34); +lean_inc(x_33); +x_36 = l_Lean_FileMap_leanPosToLspPos(x_31, x_33); +x_37 = lean_ctor_get(x_36, 1); +lean_inc(x_37); +if (lean_is_exclusive(x_36)) { + lean_ctor_release(x_36, 0); + lean_ctor_release(x_36, 1); + x_38 = x_36; +} else { + lean_dec_ref(x_36); + x_38 = lean_box(0); +} +lean_inc(x_35); +x_39 = l_Lean_FileMap_leanPosToLspPos(x_31, x_35); +lean_dec(x_31); +x_40 = lean_ctor_get(x_39, 1); +lean_inc(x_40); +lean_dec(x_39); +x_41 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_41, 0, x_33); +lean_ctor_set(x_41, 1, x_37); +lean_ctor_set(x_41, 2, x_35); +lean_ctor_set(x_41, 3, x_40); +x_42 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_42, 0, x_41); +if (lean_is_scalar(x_38)) { + x_43 = lean_alloc_ctor(0, 2, 0); +} else { + x_43 = x_38; +} +lean_ctor_set(x_43, 0, x_42); +lean_ctor_set(x_43, 1, x_6); +return x_43; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_addDeclarationRanges___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; uint8_t x_9; +x_8 = lean_st_ref_get(x_6, x_7); +x_9 = !lean_is_exclusive(x_8); +if (x_9 == 0) +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_10 = lean_ctor_get(x_8, 0); +x_11 = lean_ctor_get(x_8, 1); +x_12 = lean_ctor_get(x_10, 0); +lean_inc(x_12); +lean_dec(x_10); +x_13 = l_Lean_instInhabitedDeclarationRanges; +x_14 = l_Lean_addDeclarationRanges___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___spec__8___closed__1; +lean_inc(x_1); +x_15 = l_Lean_MapDeclarationExtension_contains___rarg(x_13, x_14, x_12, x_1); +lean_dec(x_12); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; +lean_free_object(x_8); +x_16 = lean_st_ref_take(x_6, x_11); +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +x_18 = lean_ctor_get(x_16, 1); +lean_inc(x_18); +lean_dec(x_16); +x_19 = !lean_is_exclusive(x_17); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; +x_20 = lean_ctor_get(x_17, 0); +x_21 = lean_ctor_get(x_17, 4); +lean_dec(x_21); +x_22 = l_Lean_MapDeclarationExtension_insert___rarg(x_14, x_20, x_1, x_2); +x_23 = l_Lean_setEnv___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___spec__5___closed__3; +lean_ctor_set(x_17, 4, x_23); +lean_ctor_set(x_17, 0, x_22); +x_24 = lean_st_ref_set(x_6, x_17, x_18); +x_25 = lean_ctor_get(x_24, 1); +lean_inc(x_25); +lean_dec(x_24); +x_26 = lean_st_ref_take(x_4, x_25); +x_27 = lean_ctor_get(x_26, 0); +lean_inc(x_27); +x_28 = lean_ctor_get(x_26, 1); +lean_inc(x_28); +lean_dec(x_26); +x_29 = !lean_is_exclusive(x_27); +if (x_29 == 0) +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; +x_30 = lean_ctor_get(x_27, 1); +lean_dec(x_30); +x_31 = l_Lean_setEnv___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___spec__5___closed__4; +lean_ctor_set(x_27, 1, x_31); +x_32 = lean_st_ref_set(x_4, x_27, x_28); +x_33 = !lean_is_exclusive(x_32); +if (x_33 == 0) +{ +lean_object* x_34; lean_object* x_35; +x_34 = lean_ctor_get(x_32, 0); +lean_dec(x_34); +x_35 = lean_box(0); +lean_ctor_set(x_32, 0, x_35); +return x_32; +} +else +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_32, 1); +lean_inc(x_36); +lean_dec(x_32); +x_37 = lean_box(0); +x_38 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_38, 0, x_37); +lean_ctor_set(x_38, 1, x_36); +return x_38; +} +} +else +{ +lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_39 = lean_ctor_get(x_27, 0); +x_40 = lean_ctor_get(x_27, 2); +x_41 = lean_ctor_get(x_27, 3); +x_42 = lean_ctor_get(x_27, 4); +lean_inc(x_42); +lean_inc(x_41); +lean_inc(x_40); +lean_inc(x_39); +lean_dec(x_27); +x_43 = l_Lean_setEnv___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___spec__5___closed__4; +x_44 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_44, 0, x_39); +lean_ctor_set(x_44, 1, x_43); +lean_ctor_set(x_44, 2, x_40); +lean_ctor_set(x_44, 3, x_41); +lean_ctor_set(x_44, 4, x_42); +x_45 = lean_st_ref_set(x_4, x_44, x_28); +x_46 = lean_ctor_get(x_45, 1); +lean_inc(x_46); +if (lean_is_exclusive(x_45)) { + lean_ctor_release(x_45, 0); + lean_ctor_release(x_45, 1); + x_47 = x_45; +} else { + lean_dec_ref(x_45); + x_47 = lean_box(0); +} +x_48 = lean_box(0); +if (lean_is_scalar(x_47)) { + x_49 = lean_alloc_ctor(0, 2, 0); +} else { + x_49 = x_47; +} +lean_ctor_set(x_49, 0, x_48); +lean_ctor_set(x_49, 1, x_46); +return x_49; +} +} +else +{ +lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; 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_50 = lean_ctor_get(x_17, 0); +x_51 = lean_ctor_get(x_17, 1); +x_52 = lean_ctor_get(x_17, 2); +x_53 = lean_ctor_get(x_17, 3); +x_54 = lean_ctor_get(x_17, 5); +x_55 = lean_ctor_get(x_17, 6); +x_56 = lean_ctor_get(x_17, 7); +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_17); +x_57 = l_Lean_MapDeclarationExtension_insert___rarg(x_14, x_50, x_1, x_2); +x_58 = l_Lean_setEnv___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___spec__5___closed__3; +x_59 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_59, 0, x_57); +lean_ctor_set(x_59, 1, x_51); +lean_ctor_set(x_59, 2, x_52); +lean_ctor_set(x_59, 3, x_53); +lean_ctor_set(x_59, 4, x_58); +lean_ctor_set(x_59, 5, x_54); +lean_ctor_set(x_59, 6, x_55); +lean_ctor_set(x_59, 7, x_56); +x_60 = lean_st_ref_set(x_6, x_59, x_18); +x_61 = lean_ctor_get(x_60, 1); +lean_inc(x_61); +lean_dec(x_60); +x_62 = lean_st_ref_take(x_4, x_61); +x_63 = lean_ctor_get(x_62, 0); +lean_inc(x_63); +x_64 = lean_ctor_get(x_62, 1); +lean_inc(x_64); +lean_dec(x_62); +x_65 = lean_ctor_get(x_63, 0); +lean_inc(x_65); +x_66 = lean_ctor_get(x_63, 2); +lean_inc(x_66); +x_67 = lean_ctor_get(x_63, 3); +lean_inc(x_67); +x_68 = lean_ctor_get(x_63, 4); +lean_inc(x_68); +if (lean_is_exclusive(x_63)) { + lean_ctor_release(x_63, 0); + lean_ctor_release(x_63, 1); + lean_ctor_release(x_63, 2); + lean_ctor_release(x_63, 3); + lean_ctor_release(x_63, 4); + x_69 = x_63; +} else { + lean_dec_ref(x_63); + x_69 = lean_box(0); +} +x_70 = l_Lean_setEnv___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___spec__5___closed__4; +if (lean_is_scalar(x_69)) { + x_71 = lean_alloc_ctor(0, 5, 0); +} else { + x_71 = x_69; +} +lean_ctor_set(x_71, 0, x_65); +lean_ctor_set(x_71, 1, x_70); +lean_ctor_set(x_71, 2, x_66); +lean_ctor_set(x_71, 3, x_67); +lean_ctor_set(x_71, 4, x_68); +x_72 = lean_st_ref_set(x_4, x_71, x_64); +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); +} +x_75 = lean_box(0); +if (lean_is_scalar(x_74)) { + x_76 = lean_alloc_ctor(0, 2, 0); +} else { + x_76 = x_74; +} +lean_ctor_set(x_76, 0, x_75); +lean_ctor_set(x_76, 1, x_73); +return x_76; +} } else { -uint8_t x_59; -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_4); +lean_object* x_77; lean_dec(x_2); lean_dec(x_1); -x_59 = !lean_is_exclusive(x_54); -if (x_59 == 0) -{ -return x_54; +x_77 = lean_box(0); +lean_ctor_set(x_8, 0, x_77); +return x_8; +} } else { -lean_object* x_60; lean_object* x_61; lean_object* x_62; -x_60 = lean_ctor_get(x_54, 0); -x_61 = lean_ctor_get(x_54, 1); -lean_inc(x_61); -lean_inc(x_60); -lean_dec(x_54); -x_62 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_62, 0, x_60); -lean_ctor_set(x_62, 1, x_61); -return x_62; +lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; uint8_t x_83; +x_78 = lean_ctor_get(x_8, 0); +x_79 = lean_ctor_get(x_8, 1); +lean_inc(x_79); +lean_inc(x_78); +lean_dec(x_8); +x_80 = lean_ctor_get(x_78, 0); +lean_inc(x_80); +lean_dec(x_78); +x_81 = l_Lean_instInhabitedDeclarationRanges; +x_82 = l_Lean_addDeclarationRanges___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___spec__8___closed__1; +lean_inc(x_1); +x_83 = l_Lean_MapDeclarationExtension_contains___rarg(x_81, x_82, x_80, x_1); +lean_dec(x_80); +if (x_83 == 0) +{ +lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; +x_84 = lean_st_ref_take(x_6, x_79); +x_85 = lean_ctor_get(x_84, 0); +lean_inc(x_85); +x_86 = lean_ctor_get(x_84, 1); +lean_inc(x_86); +lean_dec(x_84); +x_87 = lean_ctor_get(x_85, 0); +lean_inc(x_87); +x_88 = lean_ctor_get(x_85, 1); +lean_inc(x_88); +x_89 = lean_ctor_get(x_85, 2); +lean_inc(x_89); +x_90 = lean_ctor_get(x_85, 3); +lean_inc(x_90); +x_91 = lean_ctor_get(x_85, 5); +lean_inc(x_91); +x_92 = lean_ctor_get(x_85, 6); +lean_inc(x_92); +x_93 = lean_ctor_get(x_85, 7); +lean_inc(x_93); +if (lean_is_exclusive(x_85)) { + lean_ctor_release(x_85, 0); + lean_ctor_release(x_85, 1); + lean_ctor_release(x_85, 2); + lean_ctor_release(x_85, 3); + lean_ctor_release(x_85, 4); + lean_ctor_release(x_85, 5); + lean_ctor_release(x_85, 6); + lean_ctor_release(x_85, 7); + x_94 = x_85; +} else { + lean_dec_ref(x_85); + x_94 = lean_box(0); } +x_95 = l_Lean_MapDeclarationExtension_insert___rarg(x_82, x_87, x_1, x_2); +x_96 = l_Lean_setEnv___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___spec__5___closed__3; +if (lean_is_scalar(x_94)) { + x_97 = lean_alloc_ctor(0, 8, 0); +} else { + x_97 = x_94; } +lean_ctor_set(x_97, 0, x_95); +lean_ctor_set(x_97, 1, x_88); +lean_ctor_set(x_97, 2, x_89); +lean_ctor_set(x_97, 3, x_90); +lean_ctor_set(x_97, 4, x_96); +lean_ctor_set(x_97, 5, x_91); +lean_ctor_set(x_97, 6, x_92); +lean_ctor_set(x_97, 7, x_93); +x_98 = lean_st_ref_set(x_6, x_97, x_86); +x_99 = lean_ctor_get(x_98, 1); +lean_inc(x_99); +lean_dec(x_98); +x_100 = lean_st_ref_take(x_4, x_99); +x_101 = lean_ctor_get(x_100, 0); +lean_inc(x_101); +x_102 = lean_ctor_get(x_100, 1); +lean_inc(x_102); +lean_dec(x_100); +x_103 = lean_ctor_get(x_101, 0); +lean_inc(x_103); +x_104 = lean_ctor_get(x_101, 2); +lean_inc(x_104); +x_105 = lean_ctor_get(x_101, 3); +lean_inc(x_105); +x_106 = lean_ctor_get(x_101, 4); +lean_inc(x_106); +if (lean_is_exclusive(x_101)) { + lean_ctor_release(x_101, 0); + lean_ctor_release(x_101, 1); + lean_ctor_release(x_101, 2); + lean_ctor_release(x_101, 3); + lean_ctor_release(x_101, 4); + x_107 = x_101; +} else { + lean_dec_ref(x_101); + x_107 = lean_box(0); } -block_27: -{ -if (lean_obj_tag(x_19) == 0) -{ -lean_object* x_21; lean_object* x_22; -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_4); -lean_dec(x_2); -lean_dec(x_1); -x_21 = lean_ctor_get(x_19, 0); -lean_inc(x_21); -lean_dec(x_19); -x_22 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_22, 0, x_21); -lean_ctor_set(x_22, 1, x_20); -return x_22; +x_108 = l_Lean_setEnv___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___spec__5___closed__4; +if (lean_is_scalar(x_107)) { + x_109 = lean_alloc_ctor(0, 5, 0); +} else { + x_109 = x_107; } -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_9, x_24); -x_9 = x_25; -x_10 = x_23; -x_15 = x_20; -goto _start; +lean_ctor_set(x_109, 0, x_103); +lean_ctor_set(x_109, 1, x_108); +lean_ctor_set(x_109, 2, x_104); +lean_ctor_set(x_109, 3, x_105); +lean_ctor_set(x_109, 4, x_106); +x_110 = lean_st_ref_set(x_4, x_109, x_102); +x_111 = lean_ctor_get(x_110, 1); +lean_inc(x_111); +if (lean_is_exclusive(x_110)) { + lean_ctor_release(x_110, 0); + lean_ctor_release(x_110, 1); + x_112 = x_110; +} else { + lean_dec_ref(x_110); + x_112 = lean_box(0); } +x_113 = lean_box(0); +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); +return x_114; } +else +{ +lean_object* x_115; lean_object* x_116; +lean_dec(x_2); +lean_dec(x_1); +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_79); +return x_116; } } -static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent_copyFields___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_Elab_Command_StructView_ctor___closed__1; -x_2 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent_copyFields___spec__2___closed__1; -x_3 = lean_unsigned_to_nat(789u); -x_4 = lean_unsigned_to_nat(68u); -x_5 = l_Lean_Elab_Command_StructView_ctor___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___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent_copyFields(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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_addDeclarationRangesFromSyntax___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___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_10; -x_10 = l_Lean_Expr_getAppFn(x_4); -if (lean_obj_tag(x_10) == 4) -{ -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; size_t x_27; size_t x_28; lean_object* x_29; -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); -lean_inc(x_2); -x_13 = l_Lean_getStructureCtor(x_2, x_11); -x_14 = lean_ctor_get(x_13, 0); -lean_inc(x_14); -lean_dec(x_13); -x_15 = lean_ctor_get(x_14, 0); -lean_inc(x_15); -lean_dec(x_14); -x_16 = l_Lean_Expr_const___override(x_15, x_12); -x_17 = lean_unsigned_to_nat(0u); -x_18 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_4, x_17); -x_19 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_getFieldType___lambda__2___closed__3; -lean_inc(x_18); -x_20 = lean_mk_array(x_18, x_19); -x_21 = lean_unsigned_to_nat(1u); -x_22 = lean_nat_sub(x_18, x_21); -lean_dec(x_18); -x_23 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_4, x_20, x_22); -x_24 = l_Lean_mkAppN(x_16, x_23); -lean_dec(x_23); -lean_inc(x_11); -x_25 = l_Lean_getStructureFields(x_2, x_11); -x_26 = lean_box(0); -x_27 = lean_array_size(x_25); -x_28 = 0; -x_29 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent_copyFields___spec__2(x_1, x_2, x_3, x_11, x_25, x_26, x_25, x_27, x_28, x_24, x_5, x_6, x_7, x_8, x_9); -lean_dec(x_25); -if (lean_obj_tag(x_29) == 0) +lean_object* x_9; lean_object* x_10; +lean_inc(x_6); +x_9 = l_Lean_Elab_getDeclarationRange_x3f___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___spec__2(x_2, x_4, x_5, x_6, x_7, x_8); +x_10 = lean_ctor_get(x_9, 0); +lean_inc(x_10); +if (lean_obj_tag(x_10) == 0) { -uint8_t x_30; -x_30 = !lean_is_exclusive(x_29); -if (x_30 == 0) +uint8_t x_11; +lean_dec(x_6); +lean_dec(x_1); +x_11 = !lean_is_exclusive(x_9); +if (x_11 == 0) { -return x_29; +lean_object* x_12; lean_object* x_13; +x_12 = lean_ctor_get(x_9, 0); +lean_dec(x_12); +x_13 = lean_box(0); +lean_ctor_set(x_9, 0, x_13); +return x_9; } else { -lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_31 = lean_ctor_get(x_29, 0); -x_32 = lean_ctor_get(x_29, 1); -lean_inc(x_32); -lean_inc(x_31); -lean_dec(x_29); -x_33 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_33, 0, x_31); -lean_ctor_set(x_33, 1, x_32); -return x_33; +lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_14 = lean_ctor_get(x_9, 1); +lean_inc(x_14); +lean_dec(x_9); +x_15 = lean_box(0); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_15); +lean_ctor_set(x_16, 1, x_14); +return x_16; } } else { -uint8_t x_34; -x_34 = !lean_is_exclusive(x_29); -if (x_34 == 0) +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_17 = lean_ctor_get(x_9, 1); +lean_inc(x_17); +lean_dec(x_9); +x_18 = lean_ctor_get(x_10, 0); +lean_inc(x_18); +lean_dec(x_10); +lean_inc(x_6); +x_19 = l_Lean_Elab_getDeclarationRange_x3f___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___spec__2(x_3, x_4, x_5, x_6, x_7, x_17); +x_20 = lean_ctor_get(x_19, 0); +lean_inc(x_20); +if (lean_obj_tag(x_20) == 0) { -return x_29; +uint8_t x_21; +x_21 = !lean_is_exclusive(x_19); +if (x_21 == 0) +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_22 = lean_ctor_get(x_19, 1); +x_23 = lean_ctor_get(x_19, 0); +lean_dec(x_23); +lean_inc(x_18); +lean_ctor_set(x_19, 1, x_18); +lean_ctor_set(x_19, 0, x_18); +x_24 = l_Lean_addDeclarationRanges___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___spec__3(x_1, x_19, x_4, x_5, x_6, x_7, x_22); +lean_dec(x_6); +return x_24; } else { -lean_object* x_35; lean_object* x_36; lean_object* x_37; -x_35 = lean_ctor_get(x_29, 0); -x_36 = lean_ctor_get(x_29, 1); -lean_inc(x_36); -lean_inc(x_35); -lean_dec(x_29); -x_37 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_37, 0, x_35); -lean_ctor_set(x_37, 1, x_36); -return x_37; -} +lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_25 = lean_ctor_get(x_19, 1); +lean_inc(x_25); +lean_dec(x_19); +lean_inc(x_18); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_18); +lean_ctor_set(x_26, 1, x_18); +x_27 = l_Lean_addDeclarationRanges___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___spec__3(x_1, x_26, x_4, x_5, x_6, x_7, x_25); +lean_dec(x_6); +return x_27; } } else { -lean_object* x_38; lean_object* x_39; -lean_dec(x_10); -lean_dec(x_4); -lean_dec(x_2); -lean_dec(x_1); -x_38 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent_copyFields___closed__1; -x_39 = l_panic___at_Lean_Meta_mkPProdFst___spec__1(x_38, x_5, x_6, x_7, x_8, x_9); -return x_39; -} -} -} -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent_copyFields___spec__2___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: +uint8_t x_28; +x_28 = !lean_is_exclusive(x_19); +if (x_28 == 0) { -lean_object* x_12; -x_12 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent_copyFields___spec__2___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_29 = lean_ctor_get(x_19, 1); +x_30 = lean_ctor_get(x_19, 0); +lean_dec(x_30); +x_31 = lean_ctor_get(x_20, 0); +lean_inc(x_31); +lean_dec(x_20); +lean_ctor_set(x_19, 1, x_31); +lean_ctor_set(x_19, 0, x_18); +x_32 = l_Lean_addDeclarationRanges___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___spec__3(x_1, x_19, x_4, x_5, x_6, x_7, x_29); lean_dec(x_6); -lean_dec(x_4); -lean_dec(x_1); -return x_12; -} -} -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent_copyFields___spec__2___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_11; -x_11 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent_copyFields___spec__2___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_5); -lean_dec(x_3); -return x_11; -} +return x_32; } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent_copyFields___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { -_start: +else { -size_t x_16; size_t x_17; lean_object* x_18; -x_16 = lean_unbox_usize(x_8); -lean_dec(x_8); -x_17 = lean_unbox_usize(x_9); -lean_dec(x_9); -x_18 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent_copyFields___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_16, x_17, x_10, x_11, x_12, x_13, x_14, x_15); -lean_dec(x_7); +lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_33 = lean_ctor_get(x_19, 1); +lean_inc(x_33); +lean_dec(x_19); +x_34 = lean_ctor_get(x_20, 0); +lean_inc(x_34); +lean_dec(x_20); +x_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_18); +lean_ctor_set(x_35, 1, x_34); +x_36 = l_Lean_addDeclarationRanges___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___spec__3(x_1, x_35, x_4, x_5, x_6, x_7, x_33); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -return x_18; +return x_36; +} } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent_copyFields___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent_copyFields(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 uint8_t l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { @@ -31628,119 +33067,144 @@ return x_7; } } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___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_9; lean_object* x_10; lean_object* x_11; -x_9 = 0; -x_10 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_10, 0, x_1); -lean_ctor_set(x_10, 1, x_2); -lean_ctor_set_uint8(x_10, sizeof(void*)*2, x_9); -x_11 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_11, 0, x_10); -lean_ctor_set(x_11, 1, x_8); -return x_11; +lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_10 = lean_ctor_get(x_1, 0); +x_11 = lean_ctor_get(x_2, 0); +lean_inc(x_3); +x_12 = l_Lean_Elab_addDeclarationRangesFromSyntax___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___spec__1(x_3, x_10, x_11, x_5, x_6, x_7, x_8, x_9); +x_13 = !lean_is_exclusive(x_12); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; uint8_t x_16; lean_object* x_17; +x_14 = lean_ctor_get(x_12, 0); +lean_dec(x_14); +x_15 = lean_ctor_get(x_2, 2); +x_16 = 0; +lean_inc(x_15); +x_17 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_17, 0, x_15); +lean_ctor_set(x_17, 1, x_3); +lean_ctor_set_uint8(x_17, sizeof(void*)*2, x_16); +lean_ctor_set(x_12, 0, x_17); +return x_12; +} +else +{ +lean_object* x_18; lean_object* x_19; uint8_t x_20; lean_object* x_21; lean_object* x_22; +x_18 = lean_ctor_get(x_12, 1); +lean_inc(x_18); +lean_dec(x_12); +x_19 = lean_ctor_get(x_2, 2); +x_20 = 0; +lean_inc(x_19); +x_21 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_21, 0, x_19); +lean_ctor_set(x_21, 1, x_3); +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_18); +return x_22; +} } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _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); -x_11 = l_Lean_Meta_isProp(x_1, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_11) == 0) +x_12 = l_Lean_Meta_isProp(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; 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_object* x_13; lean_object* x_14; uint8_t x_15; +x_13 = lean_ctor_get(x_12, 0); lean_inc(x_13); -lean_dec(x_11); -x_14 = l_Lean_BinderInfo_isInstImplicit(x_4); -if (x_14 == 0) -{ -uint8_t x_15; -x_15 = lean_unbox(x_12); +x_14 = lean_ctor_get(x_12, 1); +lean_inc(x_14); lean_dec(x_12); +x_15 = l_Lean_BinderInfo_isInstImplicit(x_5); if (x_15 == 0) { -uint8_t x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_16 = 0; -lean_inc(x_3); -x_17 = l_Lean_setReducibilityStatus___at___private_Lean_Meta_Constructions_BRecOn_0__Lean_mkBelowFromRec___spec__7(x_3, x_16, x_6, x_7, x_8, x_9, x_13); -x_18 = lean_ctor_get(x_17, 0); -lean_inc(x_18); -x_19 = lean_ctor_get(x_17, 1); +uint8_t x_16; +x_16 = lean_unbox(x_13); +lean_dec(x_13); +if (x_16 == 0) +{ +uint8_t x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_17 = 0; +lean_inc(x_4); +x_18 = l_Lean_setReducibilityStatus___at___private_Lean_Meta_Constructions_BRecOn_0__Lean_mkBelowFromRec___spec__7(x_4, x_17, x_7, x_8, x_9, x_10, x_14); +x_19 = lean_ctor_get(x_18, 0); lean_inc(x_19); -lean_dec(x_17); -x_20 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___lambda__2(x_2, x_3, x_18, x_6, x_7, x_8, x_9, x_19); -lean_dec(x_9); +x_20 = lean_ctor_get(x_18, 1); +lean_inc(x_20); +lean_dec(x_18); +x_21 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___lambda__2(x_2, x_3, x_4, x_19, x_7, x_8, x_9, x_10, x_20); +lean_dec(x_10); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_18); -return x_20; +lean_dec(x_19); +return x_21; } else { -lean_object* x_21; lean_object* x_22; -x_21 = lean_box(0); -x_22 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___lambda__2(x_2, x_3, x_21, x_6, x_7, x_8, x_9, x_13); -lean_dec(x_9); +lean_object* x_22; lean_object* x_23; +x_22 = lean_box(0); +x_23 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___lambda__2(x_2, x_3, x_4, x_22, x_7, x_8, x_9, x_10, x_14); +lean_dec(x_10); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_6); -return x_22; +return x_23; } } else { -lean_object* x_23; lean_object* x_24; -lean_dec(x_12); -x_23 = lean_box(0); -x_24 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___lambda__2(x_2, x_3, x_23, x_6, x_7, x_8, x_9, x_13); -lean_dec(x_9); +lean_object* x_24; lean_object* x_25; +lean_dec(x_13); +x_24 = lean_box(0); +x_25 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___lambda__2(x_2, x_3, x_4, x_24, x_7, x_8, x_9, x_10, x_14); +lean_dec(x_10); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_6); -return x_24; +return x_25; } } else { -uint8_t x_25; +uint8_t 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_3); -lean_dec(x_2); -x_25 = !lean_is_exclusive(x_11); -if (x_25 == 0) +lean_dec(x_4); +x_26 = !lean_is_exclusive(x_12); +if (x_26 == 0) { -return x_11; +return x_12; } else { -lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_26 = lean_ctor_get(x_11, 0); -x_27 = lean_ctor_get(x_11, 1); +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_12, 0); +x_28 = lean_ctor_get(x_12, 1); +lean_inc(x_28); lean_inc(x_27); -lean_inc(x_26); -lean_dec(x_11); -x_28 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set(x_28, 1, x_27); -return x_28; +lean_dec(x_12); +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___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___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, 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, lean_object* x_19, lean_object* x_20) { +LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, uint8_t x_11, 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) { _start: { lean_object* x_21; uint8_t x_22; lean_object* x_23; lean_object* x_24; @@ -31810,10 +33274,10 @@ x_45 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_45, 0, x_44); lean_ctor_set(x_45, 1, x_8); lean_ctor_set(x_45, 2, x_23); -if (x_11 == 0) +if (x_12 == 0) { lean_object* x_46; uint8_t x_47; -x_46 = lean_ctor_get(x_12, 2); +x_46 = lean_ctor_get(x_9, 2); x_47 = lean_ctor_get_uint8(x_46, sizeof(void*)*3 + 3); lean_inc(x_44); lean_ctor_set_tag(x_35, 1); @@ -31843,7 +33307,7 @@ lean_inc(x_53); x_54 = lean_ctor_get(x_52, 1); lean_inc(x_54); lean_dec(x_52); -x_55 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___lambda__3(x_5, x_9, x_44, x_10, x_53, x_16, x_17, x_18, x_19, x_54); +x_55 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___lambda__3(x_5, x_9, x_10, x_44, x_11, x_53, x_16, x_17, x_18, x_19, x_54); lean_dec(x_53); return x_55; } @@ -31855,7 +33319,6 @@ lean_dec(x_19); lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); -lean_dec(x_9); lean_dec(x_5); x_56 = !lean_is_exclusive(x_52); if (x_56 == 0) @@ -31901,7 +33364,7 @@ lean_inc(x_65); x_66 = lean_ctor_get(x_64, 1); lean_inc(x_66); lean_dec(x_64); -x_67 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___lambda__3(x_5, x_9, x_44, x_10, x_65, x_16, x_17, x_18, x_19, x_66); +x_67 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___lambda__3(x_5, x_9, x_10, x_44, x_11, x_65, x_16, x_17, x_18, x_19, x_66); lean_dec(x_65); return x_67; } @@ -31913,7 +33376,6 @@ lean_dec(x_19); lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); -lean_dec(x_9); lean_dec(x_5); x_68 = !lean_is_exclusive(x_64); if (x_68 == 0) @@ -31939,7 +33401,7 @@ return x_71; else { lean_object* x_72; uint8_t x_73; -x_72 = lean_ctor_get(x_12, 2); +x_72 = lean_ctor_get(x_9, 2); x_73 = lean_ctor_get_uint8(x_72, sizeof(void*)*3 + 3); if (x_73 == 0) { @@ -31965,7 +33427,7 @@ lean_inc(x_77); x_78 = lean_ctor_get(x_76, 1); lean_inc(x_78); lean_dec(x_76); -x_79 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___lambda__3(x_5, x_9, x_44, x_10, x_77, x_16, x_17, x_18, x_19, x_78); +x_79 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___lambda__3(x_5, x_9, x_10, x_44, x_11, x_77, x_16, x_17, x_18, x_19, x_78); lean_dec(x_77); return x_79; } @@ -31977,7 +33439,6 @@ lean_dec(x_19); lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); -lean_dec(x_9); lean_dec(x_5); x_80 = !lean_is_exclusive(x_76); if (x_80 == 0) @@ -32024,7 +33485,7 @@ lean_inc(x_87); x_88 = lean_ctor_get(x_86, 1); lean_inc(x_88); lean_dec(x_86); -x_89 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___lambda__3(x_5, x_9, x_44, x_10, x_87, x_16, x_17, x_18, x_19, x_88); +x_89 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___lambda__3(x_5, x_9, x_10, x_44, x_11, x_87, x_16, x_17, x_18, x_19, x_88); lean_dec(x_87); return x_89; } @@ -32036,7 +33497,6 @@ lean_dec(x_19); lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); -lean_dec(x_9); lean_dec(x_5); x_90 = !lean_is_exclusive(x_86); if (x_90 == 0) @@ -32071,7 +33531,6 @@ lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); lean_dec(x_13); -lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_5); @@ -32124,10 +33583,10 @@ x_106 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_106, 0, x_105); lean_ctor_set(x_106, 1, x_8); lean_ctor_set(x_106, 2, x_23); -if (x_11 == 0) +if (x_12 == 0) { lean_object* x_107; uint8_t x_108; lean_object* x_109; -x_107 = lean_ctor_get(x_12, 2); +x_107 = lean_ctor_get(x_9, 2); x_108 = lean_ctor_get_uint8(x_107, sizeof(void*)*3 + 3); lean_inc(x_105); x_109 = lean_alloc_ctor(1, 2, 0); @@ -32157,7 +33616,7 @@ lean_inc(x_115); x_116 = lean_ctor_get(x_114, 1); lean_inc(x_116); lean_dec(x_114); -x_117 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___lambda__3(x_5, x_9, x_105, x_10, x_115, x_16, x_17, x_18, x_19, x_116); +x_117 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___lambda__3(x_5, x_9, x_10, x_105, x_11, x_115, x_16, x_17, x_18, x_19, x_116); lean_dec(x_115); return x_117; } @@ -32169,7 +33628,6 @@ lean_dec(x_19); lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); -lean_dec(x_9); lean_dec(x_5); x_118 = lean_ctor_get(x_114, 0); lean_inc(x_118); @@ -32217,7 +33675,7 @@ lean_inc(x_127); x_128 = lean_ctor_get(x_126, 1); lean_inc(x_128); lean_dec(x_126); -x_129 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___lambda__3(x_5, x_9, x_105, x_10, x_127, x_16, x_17, x_18, x_19, x_128); +x_129 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___lambda__3(x_5, x_9, x_10, x_105, x_11, x_127, x_16, x_17, x_18, x_19, x_128); lean_dec(x_127); return x_129; } @@ -32229,7 +33687,6 @@ lean_dec(x_19); lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); -lean_dec(x_9); lean_dec(x_5); x_130 = lean_ctor_get(x_126, 0); lean_inc(x_130); @@ -32257,7 +33714,7 @@ return x_133; else { lean_object* x_134; uint8_t x_135; -x_134 = lean_ctor_get(x_12, 2); +x_134 = lean_ctor_get(x_9, 2); x_135 = lean_ctor_get_uint8(x_134, sizeof(void*)*3 + 3); if (x_135 == 0) { @@ -32283,7 +33740,7 @@ lean_inc(x_140); x_141 = lean_ctor_get(x_139, 1); lean_inc(x_141); lean_dec(x_139); -x_142 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___lambda__3(x_5, x_9, x_105, x_10, x_140, x_16, x_17, x_18, x_19, x_141); +x_142 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___lambda__3(x_5, x_9, x_10, x_105, x_11, x_140, x_16, x_17, x_18, x_19, x_141); lean_dec(x_140); return x_142; } @@ -32295,7 +33752,6 @@ lean_dec(x_19); lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); -lean_dec(x_9); lean_dec(x_5); x_143 = lean_ctor_get(x_139, 0); lean_inc(x_143); @@ -32344,7 +33800,7 @@ lean_inc(x_151); x_152 = lean_ctor_get(x_150, 1); lean_inc(x_152); lean_dec(x_150); -x_153 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___lambda__3(x_5, x_9, x_105, x_10, x_151, x_16, x_17, x_18, x_19, x_152); +x_153 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___lambda__3(x_5, x_9, x_10, x_105, x_11, x_151, x_16, x_17, x_18, x_19, x_152); lean_dec(x_151); return x_153; } @@ -32356,7 +33812,6 @@ lean_dec(x_19); lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); -lean_dec(x_9); lean_dec(x_5); x_154 = lean_ctor_get(x_150, 0); lean_inc(x_154); @@ -32392,7 +33847,6 @@ lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); lean_dec(x_13); -lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_5); @@ -32429,7 +33883,6 @@ lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); lean_dec(x_13); -lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_5); @@ -32463,7 +33916,6 @@ lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); lean_dec(x_13); -lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_5); @@ -32497,7 +33949,6 @@ lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); lean_dec(x_13); -lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_5); @@ -32577,11 +34028,12 @@ goto block_70; } else { -uint8_t x_71; -lean_inc(x_5); +lean_object* x_71; uint8_t x_72; +x_71 = lean_ctor_get(x_5, 2); +lean_inc(x_71); lean_inc(x_19); -x_71 = lean_is_class(x_19, x_5); -x_30 = x_71; +x_72 = lean_is_class(x_19, x_71); +x_30 = x_72; goto block_70; } block_70: @@ -32730,7 +34182,8 @@ lean_dec(x_38); x_41 = lean_box(0); x_42 = lean_unbox(x_13); lean_dec(x_13); -x_43 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___lambda__4(x_2, x_4, x_19, x_23, x_6, x_26, x_21, x_1, x_5, x_32, x_42, x_20, x_25, x_39, x_41, x_7, x_8, x_9, x_10, x_40); +x_43 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___lambda__4(x_2, x_4, x_19, x_23, x_6, x_26, x_21, x_1, x_20, x_5, x_32, x_42, x_25, x_39, x_41, x_7, x_8, x_9, x_10, x_40); +lean_dec(x_5); lean_dec(x_20); lean_dec(x_26); lean_dec(x_23); @@ -32748,7 +34201,8 @@ x_46 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_setSourceInstImplici x_47 = lean_box(0); x_48 = lean_unbox(x_13); lean_dec(x_13); -x_49 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___lambda__4(x_2, x_4, x_19, x_23, x_6, x_26, x_21, x_1, x_5, x_32, x_48, x_20, x_25, x_46, x_47, x_7, x_8, x_9, x_10, x_45); +x_49 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___lambda__4(x_2, x_4, x_19, x_23, x_6, x_26, x_21, x_1, x_20, x_5, x_32, x_48, x_25, x_46, x_47, x_7, x_8, x_9, x_10, x_45); +lean_dec(x_5); lean_dec(x_20); lean_dec(x_26); lean_dec(x_23); @@ -32797,88 +34251,89 @@ return x_53; } else { -lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; uint8_t x_77; lean_object* x_78; uint8_t 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; -x_72 = lean_ctor_get(x_15, 0); -x_73 = lean_ctor_get(x_15, 1); +lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; uint8_t x_78; lean_object* x_79; uint8_t x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; uint8_t x_84; uint8_t x_85; lean_object* x_86; uint8_t x_87; +x_73 = lean_ctor_get(x_15, 0); +x_74 = lean_ctor_get(x_15, 1); +lean_inc(x_74); lean_inc(x_73); -lean_inc(x_72); lean_dec(x_15); -x_74 = lean_ctor_get(x_72, 0); -lean_inc(x_74); -lean_dec(x_72); -x_75 = lean_ctor_get(x_3, 0); +x_75 = lean_ctor_get(x_73, 0); lean_inc(x_75); -lean_dec(x_3); -x_76 = lean_ctor_get(x_75, 4); -lean_inc(x_76); -x_77 = 1; +lean_dec(x_73); +x_76 = lean_ctor_get(x_3, 0); lean_inc(x_76); -x_78 = l_Lean_getStructureFieldsFlattened(x_74, x_76, x_77); -x_79 = lean_ctor_get_uint8(x_75, sizeof(void*)*11); -x_80 = lean_box(0); +lean_dec(x_3); +x_77 = lean_ctor_get(x_76, 4); +lean_inc(x_77); +x_78 = 1; +lean_inc(x_77); +x_79 = l_Lean_getStructureFieldsFlattened(x_75, x_77, x_78); +x_80 = lean_ctor_get_uint8(x_76, sizeof(void*)*11); +x_81 = lean_box(0); lean_inc(x_4); -x_81 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_81, 0, x_4); -lean_ctor_set(x_81, 1, x_80); -x_82 = lean_array_mk(x_81); -x_83 = 0; -x_84 = 1; +x_82 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_82, 0, x_4); +lean_ctor_set(x_82, 1, x_81); +x_83 = lean_array_mk(x_82); +x_84 = 0; +x_85 = 1; lean_inc(x_6); -x_85 = l_Lean_Meta_mkForallFVars(x_82, x_6, x_83, x_77, x_84, x_7, x_8, x_9, x_10, x_73); -if (x_79 == 0) +x_86 = l_Lean_Meta_mkForallFVars(x_83, x_6, x_84, x_78, x_85, x_7, x_8, x_9, x_10, x_74); +if (x_80 == 0) { -x_86 = x_83; -goto block_126; +x_87 = x_84; +goto block_127; } else { -uint8_t x_127; -lean_inc(x_5); -lean_inc(x_74); -x_127 = lean_is_class(x_74, x_5); -x_86 = x_127; -goto block_126; +lean_object* x_128; uint8_t x_129; +x_128 = lean_ctor_get(x_5, 2); +lean_inc(x_128); +lean_inc(x_75); +x_129 = lean_is_class(x_75, x_128); +x_87 = x_129; +goto block_127; } -block_126: +block_127: { -uint8_t x_87; -if (x_86 == 0) +uint8_t x_88; +if (x_87 == 0) { -x_87 = x_83; -goto block_125; +x_88 = x_84; +goto block_126; } else { -x_87 = x_77; -goto block_125; +x_88 = x_78; +goto block_126; } -block_125: +block_126: { -uint8_t x_88; lean_object* x_89; lean_object* x_90; -if (x_87 == 0) +uint8_t x_89; lean_object* x_90; lean_object* x_91; +if (x_88 == 0) { -if (lean_obj_tag(x_85) == 0) +if (lean_obj_tag(x_86) == 0) { -lean_object* x_111; lean_object* x_112; uint8_t x_113; -x_111 = lean_ctor_get(x_85, 0); -lean_inc(x_111); -x_112 = lean_ctor_get(x_85, 1); +lean_object* x_112; lean_object* x_113; uint8_t x_114; +x_112 = lean_ctor_get(x_86, 0); lean_inc(x_112); -lean_dec(x_85); -x_113 = 0; -x_88 = x_113; -x_89 = x_111; +x_113 = lean_ctor_get(x_86, 1); +lean_inc(x_113); +lean_dec(x_86); +x_114 = 0; +x_89 = x_114; x_90 = x_112; -goto block_110; +x_91 = x_113; +goto block_111; } else { -lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; -lean_dec(x_82); -lean_dec(x_78); +lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; +lean_dec(x_83); +lean_dec(x_79); +lean_dec(x_77); lean_dec(x_76); lean_dec(x_75); -lean_dec(x_74); lean_dec(x_13); lean_dec(x_10); lean_dec(x_9); @@ -32888,52 +34343,52 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_1); -x_114 = lean_ctor_get(x_85, 0); -lean_inc(x_114); -x_115 = lean_ctor_get(x_85, 1); +x_115 = lean_ctor_get(x_86, 0); lean_inc(x_115); -if (lean_is_exclusive(x_85)) { - lean_ctor_release(x_85, 0); - lean_ctor_release(x_85, 1); - x_116 = x_85; +x_116 = lean_ctor_get(x_86, 1); +lean_inc(x_116); +if (lean_is_exclusive(x_86)) { + lean_ctor_release(x_86, 0); + lean_ctor_release(x_86, 1); + x_117 = x_86; } else { - lean_dec_ref(x_85); - x_116 = lean_box(0); + lean_dec_ref(x_86); + x_117 = lean_box(0); } -if (lean_is_scalar(x_116)) { - x_117 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_117)) { + x_118 = lean_alloc_ctor(1, 2, 0); } else { - x_117 = x_116; + x_118 = x_117; } -lean_ctor_set(x_117, 0, x_114); -lean_ctor_set(x_117, 1, x_115); -return x_117; +lean_ctor_set(x_118, 0, x_115); +lean_ctor_set(x_118, 1, x_116); +return x_118; } } else { -if (lean_obj_tag(x_85) == 0) +if (lean_obj_tag(x_86) == 0) { -lean_object* x_118; lean_object* x_119; uint8_t x_120; -x_118 = lean_ctor_get(x_85, 0); -lean_inc(x_118); -x_119 = lean_ctor_get(x_85, 1); +lean_object* x_119; lean_object* x_120; uint8_t x_121; +x_119 = lean_ctor_get(x_86, 0); lean_inc(x_119); -lean_dec(x_85); -x_120 = 3; -x_88 = x_120; -x_89 = x_118; +x_120 = lean_ctor_get(x_86, 1); +lean_inc(x_120); +lean_dec(x_86); +x_121 = 3; +x_89 = x_121; x_90 = x_119; -goto block_110; +x_91 = x_120; +goto block_111; } else { -lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; -lean_dec(x_82); -lean_dec(x_78); +lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; +lean_dec(x_83); +lean_dec(x_79); +lean_dec(x_77); lean_dec(x_76); lean_dec(x_75); -lean_dec(x_74); lean_dec(x_13); lean_dec(x_10); lean_dec(x_9); @@ -32943,85 +34398,87 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_1); -x_121 = lean_ctor_get(x_85, 0); -lean_inc(x_121); -x_122 = lean_ctor_get(x_85, 1); +x_122 = lean_ctor_get(x_86, 0); lean_inc(x_122); -if (lean_is_exclusive(x_85)) { - lean_ctor_release(x_85, 0); - lean_ctor_release(x_85, 1); - x_123 = x_85; +x_123 = lean_ctor_get(x_86, 1); +lean_inc(x_123); +if (lean_is_exclusive(x_86)) { + lean_ctor_release(x_86, 0); + lean_ctor_release(x_86, 1); + x_124 = x_86; } else { - lean_dec_ref(x_85); - x_123 = lean_box(0); + lean_dec_ref(x_86); + 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; } } -block_110: +block_111: { -lean_object* x_91; -x_91 = l_Lean_Meta_mkForallFVars(x_2, x_89, x_83, x_77, x_84, x_7, x_8, x_9, x_10, x_90); -if (lean_obj_tag(x_91) == 0) +lean_object* x_92; +x_92 = l_Lean_Meta_mkForallFVars(x_2, x_90, x_84, x_78, x_85, x_7, x_8, x_9, x_10, x_91); +if (lean_obj_tag(x_92) == 0) { -lean_object* x_92; lean_object* x_93; lean_object* x_94; -x_92 = lean_ctor_get(x_91, 0); -lean_inc(x_92); -x_93 = lean_ctor_get(x_91, 1); +lean_object* x_93; lean_object* x_94; lean_object* x_95; +x_93 = lean_ctor_get(x_92, 0); lean_inc(x_93); -lean_dec(x_91); -x_94 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_92, x_7, x_8, x_9, x_10, x_93); -if (x_87 == 0) +x_94 = lean_ctor_get(x_92, 1); +lean_inc(x_94); +lean_dec(x_92); +x_95 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_93, x_7, x_8, x_9, x_10, x_94); +if (x_88 == 0) { -lean_object* x_95; lean_object* x_96; lean_object* 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_object* x_96; lean_object* x_97; lean_object* x_98; uint8_t x_99; lean_object* x_100; +x_96 = lean_ctor_get(x_95, 0); lean_inc(x_96); -lean_dec(x_94); -x_97 = lean_box(0); -x_98 = lean_unbox(x_13); +x_97 = lean_ctor_get(x_95, 1); +lean_inc(x_97); +lean_dec(x_95); +x_98 = lean_box(0); +x_99 = lean_unbox(x_13); lean_dec(x_13); -x_99 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___lambda__4(x_2, x_4, x_74, x_78, x_6, x_82, x_76, x_1, x_5, x_88, x_98, x_75, x_80, x_95, x_97, x_7, x_8, x_9, x_10, x_96); -lean_dec(x_75); -lean_dec(x_82); -lean_dec(x_78); -return x_99; +x_100 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___lambda__4(x_2, x_4, x_75, x_79, x_6, x_83, x_77, x_1, x_76, x_5, x_89, x_99, x_81, x_96, x_98, x_7, x_8, x_9, x_10, x_97); +lean_dec(x_5); +lean_dec(x_76); +lean_dec(x_83); +lean_dec(x_79); +return x_100; } else { -lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; uint8_t x_104; lean_object* x_105; -x_100 = lean_ctor_get(x_94, 0); -lean_inc(x_100); -x_101 = lean_ctor_get(x_94, 1); +lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; uint8_t x_105; lean_object* x_106; +x_101 = lean_ctor_get(x_95, 0); lean_inc(x_101); -lean_dec(x_94); -x_102 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_setSourceInstImplicit(x_100); -x_103 = lean_box(0); -x_104 = lean_unbox(x_13); +x_102 = lean_ctor_get(x_95, 1); +lean_inc(x_102); +lean_dec(x_95); +x_103 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_setSourceInstImplicit(x_101); +x_104 = lean_box(0); +x_105 = lean_unbox(x_13); lean_dec(x_13); -x_105 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___lambda__4(x_2, x_4, x_74, x_78, x_6, x_82, x_76, x_1, x_5, x_88, x_104, x_75, x_80, x_102, x_103, x_7, x_8, x_9, x_10, x_101); -lean_dec(x_75); -lean_dec(x_82); -lean_dec(x_78); -return x_105; +x_106 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___lambda__4(x_2, x_4, x_75, x_79, x_6, x_83, x_77, x_1, x_76, x_5, x_89, x_105, x_81, x_103, x_104, x_7, x_8, x_9, x_10, x_102); +lean_dec(x_5); +lean_dec(x_76); +lean_dec(x_83); +lean_dec(x_79); +return x_106; } } else { -lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; -lean_dec(x_82); -lean_dec(x_78); +lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; +lean_dec(x_83); +lean_dec(x_79); +lean_dec(x_77); lean_dec(x_76); lean_dec(x_75); -lean_dec(x_74); lean_dec(x_13); lean_dec(x_10); lean_dec(x_9); @@ -33031,26 +34488,26 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_1); -x_106 = lean_ctor_get(x_91, 0); -lean_inc(x_106); -x_107 = lean_ctor_get(x_91, 1); +x_107 = lean_ctor_get(x_92, 0); lean_inc(x_107); -if (lean_is_exclusive(x_91)) { - lean_ctor_release(x_91, 0); - lean_ctor_release(x_91, 1); - x_108 = x_91; +x_108 = lean_ctor_get(x_92, 1); +lean_inc(x_108); +if (lean_is_exclusive(x_92)) { + lean_ctor_release(x_92, 0); + lean_ctor_release(x_92, 1); + x_109 = x_92; } else { - lean_dec_ref(x_91); - x_108 = lean_box(0); + lean_dec_ref(x_92); + x_109 = lean_box(0); } -if (lean_is_scalar(x_108)) { - x_109 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_109)) { + x_110 = lean_alloc_ctor(1, 2, 0); } else { - x_109 = x_108; + x_110 = x_109; } -lean_ctor_set(x_109, 0, x_106); -lean_ctor_set(x_109, 1, x_107); -return x_109; +lean_ctor_set(x_110, 0, x_107); +lean_ctor_set(x_110, 1, x_108); +return x_110; } } } @@ -33059,7 +34516,7 @@ return x_109; } else { -uint8_t x_128; +uint8_t x_130; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -33069,25 +34526,62 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_128 = !lean_is_exclusive(x_12); -if (x_128 == 0) +x_130 = !lean_is_exclusive(x_12); +if (x_130 == 0) { return x_12; } else { -lean_object* x_129; lean_object* x_130; lean_object* x_131; -x_129 = lean_ctor_get(x_12, 0); -x_130 = lean_ctor_get(x_12, 1); -lean_inc(x_130); -lean_inc(x_129); +lean_object* x_131; lean_object* x_132; lean_object* x_133; +x_131 = lean_ctor_get(x_12, 0); +x_132 = lean_ctor_get(x_12, 1); +lean_inc(x_132); +lean_inc(x_131); lean_dec(x_12); -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; +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_EXPORT lean_object* l_Lean_Elab_getDeclarationRange_x3f___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; +x_7 = l_Lean_Elab_getDeclarationRange_x3f___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___spec__2(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_addDeclarationRanges___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; +x_8 = l_Lean_addDeclarationRanges___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___spec__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_addDeclarationRangesFromSyntax___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___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_Elab_addDeclarationRangesFromSyntax___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___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_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_9; } } LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { @@ -33099,28 +34593,31 @@ x_5 = lean_box(x_4); return x_5; } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___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___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { -lean_object* x_9; -x_9 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_7); +lean_object* x_10; +x_10 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_8); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); -return x_9; +lean_dec(x_2); +lean_dec(x_1); +return x_10; } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___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___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___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_11; lean_object* x_12; -x_11 = lean_unbox(x_4); -lean_dec(x_4); -x_12 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___lambda__3(x_1, x_2, x_3, x_11, x_5, x_6, x_7, x_8, x_9, x_10); +uint8_t x_12; lean_object* x_13; +x_12 = lean_unbox(x_5); lean_dec(x_5); -return x_12; +x_13 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___lambda__3(x_1, x_2, x_3, x_4, x_12, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_6); +lean_dec(x_3); +lean_dec(x_2); +return x_13; } } LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___lambda__4___boxed(lean_object** _args) { @@ -33147,13 +34644,14 @@ lean_object* x_20 = _args[19]; _start: { uint8_t x_21; uint8_t x_22; lean_object* x_23; -x_21 = lean_unbox(x_10); -lean_dec(x_10); -x_22 = lean_unbox(x_11); +x_21 = lean_unbox(x_11); lean_dec(x_11); -x_23 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_21, x_22, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20); -lean_dec(x_15); +x_22 = lean_unbox(x_12); lean_dec(x_12); +x_23 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_21, x_22, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20); +lean_dec(x_15); +lean_dec(x_10); +lean_dec(x_9); lean_dec(x_6); lean_dec(x_4); lean_dec(x_1); @@ -33271,7 +34769,7 @@ static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Elab_St lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_Elab_Command_StructView_ctor___closed__1; x_2 = l_Std_Range_forIn_x27_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkRemainingProjections___spec__3___closed__1; -x_3 = lean_unsigned_to_nat(843u); +x_3 = lean_unsigned_to_nat(847u); x_4 = lean_unsigned_to_nat(87u); x_5 = l_Lean_Elab_Command_StructView_ctor___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -33323,7 +34821,7 @@ x_35 = lean_array_fget(x_4, x_12); x_167 = lean_ctor_get_uint8(x_35, sizeof(void*)*4); if (x_167 == 0) { -lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; +lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; x_168 = lean_ctor_get(x_35, 3); lean_inc(x_168); x_169 = l_Lean_instantiateMVars___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__1(x_168, x_15, x_16, x_17, x_18, x_19, x_20, x_21); @@ -33338,31 +34836,30 @@ lean_closure_set(x_172, 0, x_32); x_173 = lean_replace_expr(x_172, x_170); lean_dec(x_170); lean_dec(x_172); -x_174 = lean_ctor_get(x_35, 2); -lean_inc(x_174); lean_inc(x_20); lean_inc(x_19); lean_inc(x_18); lean_inc(x_17); +lean_inc(x_35); lean_inc(x_7); lean_inc(x_3); lean_inc(x_1); -x_175 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent(x_1, x_2, x_3, x_7, x_174, x_173, x_17, x_18, x_19, x_20, x_171); -if (lean_obj_tag(x_175) == 0) +x_174 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent(x_1, x_2, x_3, x_7, x_35, x_173, x_17, x_18, x_19, x_20, x_171); +if (lean_obj_tag(x_174) == 0) { -lean_object* x_176; lean_object* x_177; -x_176 = lean_ctor_get(x_175, 0); +lean_object* x_175; lean_object* x_176; +x_175 = lean_ctor_get(x_174, 0); +lean_inc(x_175); +x_176 = lean_ctor_get(x_174, 1); lean_inc(x_176); -x_177 = lean_ctor_get(x_175, 1); -lean_inc(x_177); -lean_dec(x_175); -x_36 = x_176; -x_37 = x_177; +lean_dec(x_174); +x_36 = x_175; +x_37 = x_176; goto block_166; } else { -uint8_t x_178; +uint8_t x_177; lean_dec(x_35); lean_dec(x_34); lean_dec(x_33); @@ -33379,63 +34876,63 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_3); lean_dec(x_1); -x_178 = !lean_is_exclusive(x_175); -if (x_178 == 0) +x_177 = !lean_is_exclusive(x_174); +if (x_177 == 0) { -return x_175; +return x_174; } else { -lean_object* x_179; lean_object* x_180; lean_object* x_181; -x_179 = lean_ctor_get(x_175, 0); -x_180 = lean_ctor_get(x_175, 1); -lean_inc(x_180); +lean_object* x_178; lean_object* x_179; lean_object* x_180; +x_178 = lean_ctor_get(x_174, 0); +x_179 = lean_ctor_get(x_174, 1); lean_inc(x_179); -lean_dec(x_175); -x_181 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_181, 0, x_179); -lean_ctor_set(x_181, 1, x_180); -return x_181; +lean_inc(x_178); +lean_dec(x_174); +x_180 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_180, 0, x_178); +lean_ctor_set(x_180, 1, x_179); +return x_180; } } } else { -lean_object* x_182; size_t x_183; size_t x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; -x_182 = lean_box(0); -x_183 = lean_array_size(x_5); -x_184 = 0; -x_185 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_findFieldInfo_x3f___closed__1; -x_186 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkRemainingProjections___spec__1(x_5, x_35, x_182, x_185, x_5, x_183, x_184, x_185); -x_187 = lean_ctor_get(x_186, 0); -lean_inc(x_187); -lean_dec(x_186); -if (lean_obj_tag(x_187) == 0) +lean_object* x_181; size_t x_182; size_t x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; +x_181 = lean_box(0); +x_182 = lean_array_size(x_5); +x_183 = 0; +x_184 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_findFieldInfo_x3f___closed__1; +x_185 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkRemainingProjections___spec__1(x_5, x_35, x_181, x_184, x_5, x_182, x_183, x_184); +x_186 = lean_ctor_get(x_185, 0); +lean_inc(x_186); +lean_dec(x_185); +if (lean_obj_tag(x_186) == 0) { -lean_object* x_188; lean_object* x_189; -x_188 = l_Std_Range_forIn_x27_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkRemainingProjections___spec__3___closed__2; +lean_object* x_187; lean_object* x_188; +x_187 = l_Std_Range_forIn_x27_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkRemainingProjections___spec__3___closed__2; lean_inc(x_20); lean_inc(x_19); lean_inc(x_18); lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); -x_189 = l_panic___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkRemainingProjections___spec__2(x_188, x_15, x_16, x_17, x_18, x_19, x_20, x_21); -if (lean_obj_tag(x_189) == 0) +x_188 = l_panic___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkRemainingProjections___spec__2(x_187, x_15, x_16, x_17, x_18, x_19, x_20, x_21); +if (lean_obj_tag(x_188) == 0) { -lean_object* x_190; lean_object* x_191; -x_190 = lean_ctor_get(x_189, 0); +lean_object* x_189; lean_object* x_190; +x_189 = lean_ctor_get(x_188, 0); +lean_inc(x_189); +x_190 = lean_ctor_get(x_188, 1); lean_inc(x_190); -x_191 = lean_ctor_get(x_189, 1); -lean_inc(x_191); -lean_dec(x_189); -x_36 = x_190; -x_37 = x_191; +lean_dec(x_188); +x_36 = x_189; +x_37 = x_190; goto block_166; } else { -uint8_t x_192; +uint8_t x_191; lean_dec(x_35); lean_dec(x_34); lean_dec(x_33); @@ -33452,58 +34949,58 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_3); lean_dec(x_1); -x_192 = !lean_is_exclusive(x_189); -if (x_192 == 0) +x_191 = !lean_is_exclusive(x_188); +if (x_191 == 0) { -return x_189; +return x_188; } else { -lean_object* x_193; lean_object* x_194; lean_object* x_195; -x_193 = lean_ctor_get(x_189, 0); -x_194 = lean_ctor_get(x_189, 1); -lean_inc(x_194); +lean_object* x_192; lean_object* x_193; lean_object* x_194; +x_192 = lean_ctor_get(x_188, 0); +x_193 = lean_ctor_get(x_188, 1); lean_inc(x_193); -lean_dec(x_189); -x_195 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_195, 0, x_193); -lean_ctor_set(x_195, 1, x_194); -return x_195; +lean_inc(x_192); +lean_dec(x_188); +x_194 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_194, 0, x_192); +lean_ctor_set(x_194, 1, x_193); +return x_194; } } } else { -lean_object* x_196; -x_196 = lean_ctor_get(x_187, 0); -lean_inc(x_196); -lean_dec(x_187); -if (lean_obj_tag(x_196) == 0) +lean_object* x_195; +x_195 = lean_ctor_get(x_186, 0); +lean_inc(x_195); +lean_dec(x_186); +if (lean_obj_tag(x_195) == 0) { -lean_object* x_197; lean_object* x_198; -x_197 = l_Std_Range_forIn_x27_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkRemainingProjections___spec__3___closed__2; +lean_object* x_196; lean_object* x_197; +x_196 = l_Std_Range_forIn_x27_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkRemainingProjections___spec__3___closed__2; lean_inc(x_20); lean_inc(x_19); lean_inc(x_18); lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); -x_198 = l_panic___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkRemainingProjections___spec__2(x_197, x_15, x_16, x_17, x_18, x_19, x_20, x_21); -if (lean_obj_tag(x_198) == 0) +x_197 = l_panic___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkRemainingProjections___spec__2(x_196, x_15, x_16, x_17, x_18, x_19, x_20, x_21); +if (lean_obj_tag(x_197) == 0) { -lean_object* x_199; lean_object* x_200; -x_199 = lean_ctor_get(x_198, 0); +lean_object* x_198; lean_object* x_199; +x_198 = lean_ctor_get(x_197, 0); +lean_inc(x_198); +x_199 = lean_ctor_get(x_197, 1); lean_inc(x_199); -x_200 = lean_ctor_get(x_198, 1); -lean_inc(x_200); -lean_dec(x_198); -x_36 = x_199; -x_37 = x_200; +lean_dec(x_197); +x_36 = x_198; +x_37 = x_199; goto block_166; } else { -uint8_t x_201; +uint8_t x_200; lean_dec(x_35); lean_dec(x_34); lean_dec(x_33); @@ -33520,43 +35017,43 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_3); lean_dec(x_1); -x_201 = !lean_is_exclusive(x_198); -if (x_201 == 0) +x_200 = !lean_is_exclusive(x_197); +if (x_200 == 0) { -return x_198; +return x_197; } else { -lean_object* x_202; lean_object* x_203; lean_object* x_204; -x_202 = lean_ctor_get(x_198, 0); -x_203 = lean_ctor_get(x_198, 1); -lean_inc(x_203); +lean_object* x_201; lean_object* x_202; lean_object* x_203; +x_201 = lean_ctor_get(x_197, 0); +x_202 = lean_ctor_get(x_197, 1); lean_inc(x_202); -lean_dec(x_198); -x_204 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_204, 0, x_202); -lean_ctor_set(x_204, 1, x_203); -return x_204; +lean_inc(x_201); +lean_dec(x_197); +x_203 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_203, 0, x_201); +lean_ctor_set(x_203, 1, x_202); +return x_203; } } } else { -lean_object* x_205; lean_object* x_206; lean_object* x_207; uint8_t x_208; lean_object* x_209; -x_205 = lean_ctor_get(x_196, 0); +lean_object* x_204; lean_object* x_205; lean_object* x_206; uint8_t x_207; lean_object* x_208; +x_204 = lean_ctor_get(x_195, 0); +lean_inc(x_204); +lean_dec(x_195); +x_205 = lean_ctor_get(x_35, 2); lean_inc(x_205); -lean_dec(x_196); -x_206 = lean_ctor_get(x_35, 2); +x_206 = lean_ctor_get(x_204, 2); lean_inc(x_206); -x_207 = lean_ctor_get(x_205, 2); -lean_inc(x_207); -lean_dec(x_205); -x_208 = 1; -x_209 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_209, 0, x_206); -lean_ctor_set(x_209, 1, x_207); -lean_ctor_set_uint8(x_209, sizeof(void*)*2, x_208); -x_36 = x_209; +lean_dec(x_204); +x_207 = 1; +x_208 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_208, 0, x_205); +lean_ctor_set(x_208, 1, x_206); +lean_ctor_set_uint8(x_208, sizeof(void*)*2, x_207); +x_36 = x_208; x_37 = x_21; goto block_166; } @@ -50203,6 +51700,8 @@ l_Lean_setEnv___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProject lean_mark_persistent(l_Lean_setEnv___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___spec__5___closed__3); l_Lean_setEnv___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___spec__5___closed__4 = _init_l_Lean_setEnv___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___spec__5___closed__4(); lean_mark_persistent(l_Lean_setEnv___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___spec__5___closed__4); +l_Lean_addDeclarationRanges___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___spec__8___closed__1 = _init_l_Lean_addDeclarationRanges___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___spec__8___closed__1(); +lean_mark_persistent(l_Lean_addDeclarationRanges___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___spec__8___closed__1); l___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___closed__1 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___closed__1(); lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___closed__1); l___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___closed__2 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___closed__2(); diff --git a/stage0/stdlib/Lean/Elab/Tactic.c b/stage0/stdlib/Lean/Elab/Tactic.c index dae76c02ec41..a28af8886d00 100644 --- a/stage0/stdlib/Lean/Elab/Tactic.c +++ b/stage0/stdlib/Lean/Elab/Tactic.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Lean.Elab.Tactic -// Imports: Lean.Elab.Term Lean.Elab.Tactic.Basic Lean.Elab.Tactic.ElabTerm Lean.Elab.Tactic.Induction Lean.Elab.Tactic.Generalize Lean.Elab.Tactic.Injection Lean.Elab.Tactic.Match Lean.Elab.Tactic.Rewrite Lean.Elab.Tactic.Location Lean.Elab.Tactic.SimpTrace Lean.Elab.Tactic.Simp Lean.Elab.Tactic.Simproc Lean.Elab.Tactic.BuiltinTactic Lean.Elab.Tactic.Split Lean.Elab.Tactic.Conv Lean.Elab.Tactic.Delta Lean.Elab.Tactic.Meta Lean.Elab.Tactic.Unfold Lean.Elab.Tactic.Cache Lean.Elab.Tactic.Calc Lean.Elab.Tactic.Congr Lean.Elab.Tactic.Guard Lean.Elab.Tactic.RCases Lean.Elab.Tactic.Repeat Lean.Elab.Tactic.Ext Lean.Elab.Tactic.Change Lean.Elab.Tactic.FalseOrByContra Lean.Elab.Tactic.Omega Lean.Elab.Tactic.Simpa Lean.Elab.Tactic.NormCast Lean.Elab.Tactic.Symm Lean.Elab.Tactic.SolveByElim Lean.Elab.Tactic.LibrarySearch Lean.Elab.Tactic.ShowTerm Lean.Elab.Tactic.Rfl Lean.Elab.Tactic.Rewrites Lean.Elab.Tactic.DiscrTreeKey Lean.Elab.Tactic.BVDecide Lean.Elab.Tactic.BoolToPropSimps Lean.Elab.Tactic.Classical Lean.Elab.Tactic.Grind +// Imports: Lean.Elab.Term Lean.Elab.Tactic.Basic Lean.Elab.Tactic.ElabTerm Lean.Elab.Tactic.Induction Lean.Elab.Tactic.Generalize Lean.Elab.Tactic.Injection Lean.Elab.Tactic.Match Lean.Elab.Tactic.Rewrite Lean.Elab.Tactic.Location Lean.Elab.Tactic.SimpTrace Lean.Elab.Tactic.Simp Lean.Elab.Tactic.Simproc Lean.Elab.Tactic.BuiltinTactic Lean.Elab.Tactic.Split Lean.Elab.Tactic.Conv Lean.Elab.Tactic.Delta Lean.Elab.Tactic.Meta Lean.Elab.Tactic.Unfold Lean.Elab.Tactic.Cache Lean.Elab.Tactic.Calc Lean.Elab.Tactic.Congr Lean.Elab.Tactic.Guard Lean.Elab.Tactic.RCases Lean.Elab.Tactic.Repeat Lean.Elab.Tactic.Ext Lean.Elab.Tactic.Change Lean.Elab.Tactic.FalseOrByContra Lean.Elab.Tactic.Omega Lean.Elab.Tactic.Simpa Lean.Elab.Tactic.NormCast Lean.Elab.Tactic.Symm Lean.Elab.Tactic.SolveByElim Lean.Elab.Tactic.LibrarySearch Lean.Elab.Tactic.ShowTerm Lean.Elab.Tactic.Rfl Lean.Elab.Tactic.Rewrites Lean.Elab.Tactic.DiscrTreeKey Lean.Elab.Tactic.BVDecide Lean.Elab.Tactic.BoolToPropSimps Lean.Elab.Tactic.Classical Lean.Elab.Tactic.Grind Lean.Elab.Tactic.Monotonicity #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -54,6 +54,7 @@ lean_object* initialize_Lean_Elab_Tactic_BVDecide(uint8_t builtin, lean_object*) lean_object* initialize_Lean_Elab_Tactic_BoolToPropSimps(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Elab_Tactic_Classical(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Elab_Tactic_Grind(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Elab_Tactic_Monotonicity(uint8_t builtin, lean_object*); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Lean_Elab_Tactic(uint8_t builtin, lean_object* w) { lean_object * res; @@ -182,6 +183,9 @@ lean_dec_ref(res); res = initialize_Lean_Elab_Tactic_Grind(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +res = initialize_Lean_Elab_Tactic_Monotonicity(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/Elab/Tactic/Grind.c b/stage0/stdlib/Lean/Elab/Tactic/Grind.c index f81cc695fb93..3fc44a87b94b 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/Grind.c +++ b/stage0/stdlib/Lean/Elab/Tactic/Grind.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Lean.Elab.Tactic.Grind -// Imports: Init.Grind.Tactics Lean.Meta.Tactic.Grind Lean.Elab.Command Lean.Elab.Tactic.Basic +// Imports: Init.Grind.Tactics Lean.Meta.Tactic.Grind Lean.Elab.Command Lean.Elab.Tactic.Basic Lean.Elab.Tactic.Config #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -21,21 +21,27 @@ LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Tactic_elabGrindPatt static lean_object* l_Lean_Elab_Tactic_grind___closed__2; static lean_object* l_Lean_throwUnknownConstant___at_Lean_Elab_Tactic_elabGrindPattern___spec__6___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescope___at_Lean_Elab_Tactic_elabGrindPattern___spec__13___rarg(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_elabGrindPattern___closed__7; static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalApplyRfl__1___closed__1; static lean_object* l_Lean_throwUnknownConstant___at_Lean_Elab_Tactic_elabGrindPattern___spec__6___closed__4; +lean_object* l___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig(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_elabGrindConfig___lambda__3___closed__1; LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Tactic_elabGrindPattern___spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_throwError___at_Lean_Elab_Term_mkAuxName___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_ConstantInfo_type(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_elabGrindPattern___spec__1(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_elabGrindConfig___lambda__1___closed__6; +lean_object* l_Lean_Meta_Grind_addEMatchTheorem(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_indentD(lean_object*); +uint8_t l_Lean_Exception_isInterrupt(lean_object*); +static lean_object* l_Lean_Elab_Tactic_elabGrindConfig___lambda__1___closed__2; lean_object* l_Lean_Syntax_formatStxAux(lean_object*, uint8_t, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalApplyRfl__1___closed__4; lean_object* l_Lean_Elab_Term_elabTerm(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_unfoldReducible___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_ensureNonAmbiguous___at_Lean_Elab_Tactic_elabGrindPattern___spec__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); 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_object* l_Lean_Meta_Grind_addTheoremPattern(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Elab_Tactic_tacticElabAttribute; +static lean_object* l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Grind___hyg_5____closed__3; static lean_object* l_Lean_Elab_Tactic_grind___closed__1; extern lean_object* l_Lean_Elab_Term_instMonadTermElabM; lean_object* l_Lean_Syntax_getArgs(lean_object*); @@ -45,6 +51,7 @@ lean_object* l_Lean_instantiateMVars___at_Lean_Elab_Term_MVarErrorInfo_logError_ LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescope___at_Lean_Elab_Tactic_elabGrindPattern___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* l_Lean_Syntax_TSepArray_getElems___rarg(lean_object*); static lean_object* l_Lean_Elab_Tactic_evalApplyRfl___closed__6; +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_elabGrindConfig___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_ensureNonAmbiguous___at_Lean_Elab_Tactic_elabGrindPattern___spec__9___closed__4; static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalApplyRfl__1___closed__3; uint8_t l_Lean_Syntax_isOfKind(lean_object*, lean_object*); @@ -52,35 +59,54 @@ lean_object* l_Lean_stringToMessageData(lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_Elab_Tactic_elabGrindPattern___spec__10(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescope___at_Lean_Elab_Tactic_elabGrindPattern___spec__13___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_resolveGlobalConst___at_Lean_Elab_Tactic_elabGrindPattern___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_Tactic_elabGrindConfig___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_throwErrorAt___at_Lean_Elab_Tactic_elabGrindPattern___spec__11(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_Elab_Tactic_elabGrindConfig___lambda__2___closed__1; +lean_object* l_Lean_Exception_toMessageData(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_elabGrindConfig___lambda__4(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_elabGrindPattern___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Tactic_elabGrindConfig___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_Tactic_evalApplyRfl___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_evalApplyRfl__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalApplyRfl___spec__1___rarg(lean_object*); +uint8_t l_Lean_Expr_hasSyntheticSorry(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescope___at_Lean_Elab_Tactic_elabGrindPattern___spec__13(lean_object*); +lean_object* l_Lean_Parser_Tactic_getConfigItems(lean_object*); lean_object* l_List_filterMapTR_go___at_Lean_preprocessSyntaxAndResolve___spec__1(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_evalApplyRfl___closed__9; +lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_elabGrindPattern___spec__1___rarg(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_elabGrindPattern___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_throwError___at_Lean_Elab_Tactic_elabGrindConfig___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_Tactic_elabGrindPattern___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_expr_abstract(lean_object*, lean_object*); lean_object* l_List_mapTR_loop___at_Lean_filterFieldList___spec__2(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_elabGrindConfig___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_throwUnknownConstant___at_Lean_Elab_Tactic_elabGrindPattern___spec__6___closed__1; lean_object* l_Lean_Core_transform___at_Lean_Meta_Grind_unfoldReducible___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_Tactic_elabGrindConfig___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_elabGrindConfig___lambda__1___closed__5; LEAN_EXPORT lean_object* l_Lean_filterFieldList___at_Lean_Elab_Tactic_elabGrindPattern___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Elab_Command_commandElabAttribute; +static lean_object* l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Grind___hyg_5____closed__4; static lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_elabGrindPattern___spec__1___rarg___closed__2; +static lean_object* l_Lean_Elab_Tactic_evalApplyRfl___closed__8; +static lean_object* l_Lean_Elab_Tactic_elabGrindConfig___lambda__1___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*); static lean_object* l_Lean_Elab_Tactic_evalApplyRfl___closed__5; lean_object* l_Lean_Elab_goalsToMessageData(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_elabGrindPattern___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofFormat(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwUnknownConstant___at_Lean_Elab_Tactic_elabGrindPattern___spec__6___boxed(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*); LEAN_EXPORT lean_object* l_Lean_throwUnknownConstant___at_Lean_Elab_Tactic_elabGrindPattern___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_evalApplyRfl___closed__4; static lean_object* l_Lean_Elab_Tactic_elabGrindPattern___closed__5; static lean_object* l___regBuiltin_Lean_Elab_Tactic_elabGrindPattern__1___closed__2; lean_object* l_Lean_Elab_Tactic_withMainContext___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Command_liftTermElabM___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_st_ref_get(lean_object*, lean_object*); +lean_object* l_Lean_Elab_addMacroStack___at_Lean_Elab_Term_instAddErrorMessageContextTermElabM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalApplyRfl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_List_isEmpty___rarg(lean_object*); static lean_object* l_Lean_ensureNonAmbiguous___at_Lean_Elab_Tactic_elabGrindPattern___spec__9___closed__6; @@ -88,26 +114,40 @@ lean_object* lean_array_to_list(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Tactic_elabGrindPattern___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Term_getDeclName_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_elabGrindPattern___closed__1; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalApplyRfl___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalApplyRfl___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_elabGrindConfig___lambda__3(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_Elab_Tactic_Config_0__Lean_Elab_Tactic_mkConfigItemViews(lean_object*); static lean_object* l_Lean_preprocessSyntaxAndResolve___at_Lean_Elab_Tactic_elabGrindPattern___spec__7___closed__3; +static lean_object* l_Lean_Elab_Tactic_elabGrindConfig___lambda__4___closed__5; +static lean_object* l_Lean_Elab_Tactic_elabGrindConfig___lambda__4___closed__2; lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_elabGrindPattern___closed__6; 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*); LEAN_EXPORT lean_object* l_Lean_filterFieldList___at_Lean_Elab_Tactic_elabGrindPattern___spec__5___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_elabGrindConfig___lambda__4___closed__3; +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_elabGrindConfig___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_resolveGlobalName___at_Lean_Elab_Term_resolveLocalName_loop___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_elabGrindConfig___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_Tactic_elabGrindConfig___lambda__4___closed__1; extern lean_object* l_Std_Format_defWidth; static lean_object* l_Lean_Elab_Tactic_elabGrindPattern___closed__3; static lean_object* l_Lean_preprocessSyntaxAndResolve___at_Lean_Elab_Tactic_elabGrindPattern___spec__7___closed__4; +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_elabGrindConfig(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_ofConstName(lean_object*, uint8_t); +uint8_t l_Lean_Environment_contains(lean_object*, lean_object*); +lean_object* l_Lean_MessageData_ofExpr(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalApplyRfl__1___closed__2; lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_forallTelescopeReducingAuxAux___rarg(uint8_t, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_elabGrindPattern___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_Lean_Meta_Grind_main(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_elabGrindPattern___spec__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_elabGrindConfig___lambda__4___closed__4; +lean_object* l_Lean_Meta_Grind_main(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_elabGrindConfig___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Grind___hyg_5_(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_resolveGlobalConst___at_Lean_Elab_Tactic_elabGrindPattern___spec__3___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_elabGrindPattern__1(lean_object*); static lean_object* l_panic___at_Lean_Elab_Tactic_elabGrindPattern___spec__10___closed__1; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_elabGrindPattern___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_elabGrindPattern___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_Tactic_evalApplyRfl___closed__3; LEAN_EXPORT lean_object* l_Lean_filterFieldList___at_Lean_Elab_Tactic_elabGrindPattern___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_preprocessSyntaxAndResolve___at_Lean_Elab_Tactic_elabGrindPattern___spec__7___closed__1; @@ -115,11 +155,13 @@ static lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_ela static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_elabGrindPattern___spec__12___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*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_elabGrindPattern__1___closed__5; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_grind(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_grind(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_evalApplyRfl___closed__2; static lean_object* l_Lean_Elab_Tactic_elabGrindPattern___closed__4; -static lean_object* l_Lean_Elab_Tactic_grind___closed__3; +lean_object* l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_panic_fn(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Grind___hyg_5____boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Grind___hyg_5____closed__1; LEAN_EXPORT lean_object* l___private_Lean_ResolveName_0__Lean_resolveGlobalConstCore___at_Lean_Elab_Tactic_elabGrindPattern___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_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalApplyRfl___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_throwUnsupportedSyntax___at_Lean_Elab_Tactic_elabGrindPattern___spec__1___boxed(lean_object*, lean_object*); @@ -128,37 +170,902 @@ static lean_object* l_Lean_ensureNonAmbiguous___at_Lean_Elab_Tactic_elabGrindPat static lean_object* l_Lean_ensureNonAmbiguous___at_Lean_Elab_Tactic_elabGrindPattern___spec__9___closed__2; static lean_object* l_Lean_Elab_Tactic_elabGrindPattern___closed__2; LEAN_EXPORT lean_object* l_Lean_resolveGlobalConstNoOverload___at_Lean_Elab_Tactic_elabGrindPattern___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Expr_hasSorry(lean_object*); lean_object* lean_array_mk(lean_object*); size_t lean_usize_add(size_t, size_t); lean_object* l_List_mapTR_loop___at_Lean_ensureNonAmbiguous___spec__2(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_elabGrindPattern__1___closed__1; extern lean_object* l_Lean_instInhabitedName; +static lean_object* l_Lean_Elab_Tactic_elabGrindConfig___lambda__1___closed__4; lean_object* lean_array_uget(lean_object*, size_t); size_t lean_array_size(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_elabGrindPattern__1___closed__6; +lean_object* l_Lean_Elab_Term_addTermInfo(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_elabGrindConfig___lambda__2___closed__2; LEAN_EXPORT lean_object* l___private_Lean_ResolveName_0__Lean_resolveGlobalConstCore___at_Lean_Elab_Tactic_elabGrindPattern___spec__4(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*); -static lean_object* l_Lean_ensureNonAmbiguous___at_Lean_Elab_Tactic_elabGrindPattern___spec__9___closed__7; lean_object* l_Lean_Elab_Tactic_liftMetaFinishingTactic(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_preprocessSyntaxAndResolve___at_Lean_Elab_Tactic_elabGrindPattern___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_string_append(lean_object*, lean_object*); +lean_object* l_Lean_Meta_evalExpr_x27___rarg(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_elabGrindPattern__1___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_elabGrindPattern(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_ensureNonAmbiguous___at_Lean_Elab_Tactic_elabGrindPattern___spec__9___closed__1; extern lean_object* l_Lean_Elab_unsupportedSyntaxExceptionId; uint8_t lean_usize_dec_lt(size_t, size_t); +uint8_t l_Lean_Exception_isRuntime(lean_object*); +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_object* l_Lean_Meta_Grind_unfoldReducible___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalApplyRfl___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_uset(lean_object*, size_t, lean_object*); static lean_object* l_Lean_throwUnknownConstant___at_Lean_Elab_Tactic_elabGrindPattern___spec__6___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Tactic_elabGrindPattern__1___closed__3; +lean_object* l_Lean_MessageData_ofName(lean_object*); static lean_object* l_Lean_Elab_Tactic_evalApplyRfl___closed__7; lean_object* l_List_filterTR_loop___at_Lean_filterFieldList___spec__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_elabGrindPattern___spec__12(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_elabGrindConfig___lambda__1___closed__1; +static lean_object* l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Grind___hyg_5____closed__2; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_elabGrindPattern___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_filterFieldList___at_Lean_Elab_Tactic_elabGrindPattern___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*); +uint8_t l_Array_isEmpty___rarg(lean_object*); lean_object* l_Lean_throwError___at_Lean_Elab_Term_expandDeclId___spec__11(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Grind___hyg_5____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean", 4, 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Grind___hyg_5____closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Grind", 5, 5); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Grind___hyg_5____closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Config", 6, 6); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Grind___hyg_5____closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Grind___hyg_5____closed__1; +x_2 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Grind___hyg_5____closed__2; +x_3 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Grind___hyg_5____closed__3; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Grind___hyg_5_(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; uint8_t x_10; lean_object* x_11; +x_9 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Grind___hyg_5____closed__4; +x_10 = 0; +x_11 = l_Lean_Meta_evalExpr_x27___rarg(x_9, x_1, x_10, x_4, x_5, x_6, x_7, x_8); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Grind___hyg_5____boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; +x_9 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Grind___hyg_5_(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_3); +lean_dec(x_2); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Tactic_elabGrindConfig___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; uint8_t x_13; +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_is_exclusive(x_12); +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_12, 0); +x_15 = lean_ctor_get(x_12, 1); +x_16 = l_Lean_Elab_addMacroStack___at_Lean_Elab_Term_instAddErrorMessageContextTermElabM___spec__1(x_14, x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_15); +lean_dec(x_2); +x_17 = !lean_is_exclusive(x_16); +if (x_17 == 0) +{ +lean_object* x_18; +x_18 = lean_ctor_get(x_16, 0); +lean_ctor_set(x_12, 1, x_18); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set_tag(x_16, 1); +lean_ctor_set(x_16, 0, x_12); +return x_16; +} +else +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_16, 0); +x_20 = lean_ctor_get(x_16, 1); +lean_inc(x_20); +lean_inc(x_19); +lean_dec(x_16); +lean_ctor_set(x_12, 1, x_19); +lean_ctor_set(x_12, 0, x_11); +x_21 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_21, 0, x_12); +lean_ctor_set(x_21, 1, x_20); +return x_21; +} +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_22 = lean_ctor_get(x_12, 0); +x_23 = lean_ctor_get(x_12, 1); +lean_inc(x_23); +lean_inc(x_22); +lean_dec(x_12); +x_24 = l_Lean_Elab_addMacroStack___at_Lean_Elab_Term_instAddErrorMessageContextTermElabM___spec__1(x_22, x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_23); +lean_dec(x_2); +x_25 = lean_ctor_get(x_24, 0); +lean_inc(x_25); +x_26 = lean_ctor_get(x_24, 1); +lean_inc(x_26); +if (lean_is_exclusive(x_24)) { + lean_ctor_release(x_24, 0); + lean_ctor_release(x_24, 1); + x_27 = x_24; +} else { + lean_dec_ref(x_24); + x_27 = lean_box(0); +} +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_11); +lean_ctor_set(x_28, 1, x_25); +if (lean_is_scalar(x_27)) { + x_29 = lean_alloc_ctor(1, 2, 0); +} else { + x_29 = x_27; + lean_ctor_set_tag(x_29, 1); +} +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_26); +return x_29; +} +} +} +static lean_object* _init_l_Lean_Elab_Tactic_elabGrindConfig___lambda__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("error evaluating configuration\n", 31, 31); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Tactic_elabGrindConfig___lambda__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Elab_Tactic_elabGrindConfig___lambda__1___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Elab_Tactic_elabGrindConfig___lambda__1___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("\n\nException: ", 13, 13); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Tactic_elabGrindConfig___lambda__1___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Elab_Tactic_elabGrindConfig___lambda__1___closed__3; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Elab_Tactic_elabGrindConfig___lambda__1___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("", 0, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Tactic_elabGrindConfig___lambda__1___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Elab_Tactic_elabGrindConfig___lambda__1___closed__5; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_elabGrindConfig___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_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_1); +x_10 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Grind___hyg_5_(x_1, 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; +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_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) +{ +lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_16 = lean_ctor_get(x_10, 0); +x_17 = lean_ctor_get(x_10, 1); +x_18 = l_Lean_Exception_isInterrupt(x_16); +if (x_18 == 0) +{ +uint8_t x_19; +x_19 = l_Lean_Exception_isRuntime(x_16); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; +lean_free_object(x_10); +x_20 = l_Lean_MessageData_ofExpr(x_1); +x_21 = l_Lean_indentD(x_20); +x_22 = l_Lean_Elab_Tactic_elabGrindConfig___lambda__1___closed__2; +x_23 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_21); +x_24 = l_Lean_Elab_Tactic_elabGrindConfig___lambda__1___closed__4; +x_25 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_24); +x_26 = l_Lean_Exception_toMessageData(x_16); +x_27 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +x_28 = l_Lean_Elab_Tactic_elabGrindConfig___lambda__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_throwError___at_Lean_Elab_Tactic_elabGrindConfig___spec__1(x_29, x_3, x_4, 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); +x_31 = !lean_is_exclusive(x_30); +if (x_31 == 0) +{ +return x_30; +} +else +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_30, 0); +x_33 = lean_ctor_get(x_30, 1); +lean_inc(x_33); +lean_inc(x_32); +lean_dec(x_30); +x_34 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +return x_34; +} +} +else +{ +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +return x_10; +} +} +else +{ +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +return x_10; +} +} +else +{ +lean_object* x_35; lean_object* x_36; uint8_t x_37; +x_35 = lean_ctor_get(x_10, 0); +x_36 = lean_ctor_get(x_10, 1); +lean_inc(x_36); +lean_inc(x_35); +lean_dec(x_10); +x_37 = l_Lean_Exception_isInterrupt(x_35); +if (x_37 == 0) +{ +uint8_t x_38; +x_38 = l_Lean_Exception_isRuntime(x_35); +if (x_38 == 0) +{ +lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_39 = l_Lean_MessageData_ofExpr(x_1); +x_40 = l_Lean_indentD(x_39); +x_41 = l_Lean_Elab_Tactic_elabGrindConfig___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_elabGrindConfig___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_Exception_toMessageData(x_35); +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_Elab_Tactic_elabGrindConfig___lambda__1___closed__6; +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_Lean_Elab_Tactic_elabGrindConfig___spec__1(x_48, x_3, x_4, x_5, x_6, x_7, x_8, x_36); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +x_50 = lean_ctor_get(x_49, 0); +lean_inc(x_50); +x_51 = lean_ctor_get(x_49, 1); +lean_inc(x_51); +if (lean_is_exclusive(x_49)) { + lean_ctor_release(x_49, 0); + lean_ctor_release(x_49, 1); + x_52 = x_49; +} else { + lean_dec_ref(x_49); + x_52 = lean_box(0); +} +if (lean_is_scalar(x_52)) { + x_53 = lean_alloc_ctor(1, 2, 0); +} else { + x_53 = x_52; +} +lean_ctor_set(x_53, 0, x_50); +lean_ctor_set(x_53, 1, x_51); +return x_53; +} +else +{ +lean_object* x_54; +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +x_54 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_54, 0, x_35); +lean_ctor_set(x_54, 1, x_36); +return x_54; +} +} +else +{ +lean_object* x_55; +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_55 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_55, 0, x_35); +lean_ctor_set(x_55, 1, x_36); +return x_55; +} +} +} +} +} +static lean_object* _init_l_Lean_Elab_Tactic_elabGrindConfig___lambda__2___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("configuration contains 'sorry'", 30, 30); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Tactic_elabGrindConfig___lambda__2___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Elab_Tactic_elabGrindConfig___lambda__2___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_elabGrindConfig___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; +x_10 = l_Lean_Expr_hasSorry(x_1); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; +x_11 = lean_box(0); +x_12 = l_Lean_Elab_Tactic_elabGrindConfig___lambda__1(x_1, x_11, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_12; +} +else +{ +lean_object* x_13; lean_object* x_14; uint8_t x_15; +lean_dec(x_1); +x_13 = l_Lean_Elab_Tactic_elabGrindConfig___lambda__2___closed__2; +x_14 = l_Lean_throwError___at_Lean_Elab_Term_synthesizeInstMVarCore___spec__3(x_13, 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); +x_15 = !lean_is_exclusive(x_14); +if (x_15 == 0) +{ +return x_14; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_14, 0); +x_17 = lean_ctor_get(x_14, 1); +lean_inc(x_17); +lean_inc(x_16); +lean_dec(x_14); +x_18 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_18, 0, x_16); +lean_ctor_set(x_18, 1, x_17); +return x_18; +} +} +} +} +static lean_object* _init_l_Lean_Elab_Tactic_elabGrindConfig___lambda__3___closed__1() { +_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 = lean_unsigned_to_nat(100u); +x_3 = lean_unsigned_to_nat(5u); +x_4 = lean_unsigned_to_nat(1000u); +x_5 = lean_alloc_ctor(0, 4, 1); +lean_ctor_set(x_5, 0, x_2); +lean_ctor_set(x_5, 1, x_3); +lean_ctor_set(x_5, 2, x_3); +lean_ctor_set(x_5, 3, x_4); +lean_ctor_set_uint8(x_5, sizeof(void*)*4, x_1); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_elabGrindConfig___lambda__3(uint8_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; lean_object* x_12; +x_11 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Grind___hyg_5____closed__4; +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___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig(x_1, x_11, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_12) == 0) +{ +uint8_t x_13; +x_13 = !lean_is_exclusive(x_12); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_14 = lean_ctor_get(x_12, 0); +x_15 = lean_ctor_get(x_12, 1); +x_16 = l_Lean_Expr_hasSyntheticSorry(x_14); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; +lean_free_object(x_12); +x_17 = lean_box(0); +x_18 = l_Lean_Elab_Tactic_elabGrindConfig___lambda__2(x_14, x_17, x_4, x_5, x_6, x_7, x_8, x_9, x_15); +lean_dec(x_5); +return x_18; +} +else +{ +lean_object* x_19; +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); +x_19 = l_Lean_Elab_Tactic_elabGrindConfig___lambda__3___closed__1; +lean_ctor_set(x_12, 0, x_19); +return x_12; +} +} +else +{ +lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_20 = lean_ctor_get(x_12, 0); +x_21 = lean_ctor_get(x_12, 1); +lean_inc(x_21); +lean_inc(x_20); +lean_dec(x_12); +x_22 = l_Lean_Expr_hasSyntheticSorry(x_20); +if (x_22 == 0) +{ +lean_object* x_23; lean_object* x_24; +x_23 = lean_box(0); +x_24 = l_Lean_Elab_Tactic_elabGrindConfig___lambda__2(x_20, x_23, x_4, x_5, x_6, x_7, x_8, x_9, x_21); +lean_dec(x_5); +return x_24; +} +else +{ +lean_object* x_25; lean_object* x_26; +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_25 = l_Lean_Elab_Tactic_elabGrindConfig___lambda__3___closed__1; +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set(x_26, 1, x_21); +return x_26; +} +} +} +else +{ +uint8_t x_27; +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_27 = !lean_is_exclusive(x_12); +if (x_27 == 0) +{ +return x_12; +} +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_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; +} +} +} +} +static lean_object* _init_l_Lean_Elab_Tactic_elabGrindConfig___lambda__4___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("error evaluating configuration, environment does not yet contain type ", 70, 70); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Tactic_elabGrindConfig___lambda__4___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Elab_Tactic_elabGrindConfig___lambda__4___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Elab_Tactic_elabGrindConfig___lambda__4___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Grind___hyg_5____closed__4; +x_2 = l_Lean_MessageData_ofName(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Elab_Tactic_elabGrindConfig___lambda__4___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_Tactic_elabGrindConfig___lambda__4___closed__2; +x_2 = l_Lean_Elab_Tactic_elabGrindConfig___lambda__4___closed__3; +x_3 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Tactic_elabGrindConfig___lambda__4___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_Tactic_elabGrindConfig___lambda__4___closed__4; +x_2 = l_Lean_Elab_Tactic_elabGrindConfig___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; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_elabGrindConfig___lambda__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, 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 = lean_st_ref_get(x_9, x_10); +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +x_13 = lean_ctor_get(x_11, 1); +lean_inc(x_13); +lean_dec(x_11); +x_14 = lean_ctor_get(x_12, 0); +lean_inc(x_14); +lean_dec(x_12); +x_15 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Grind___hyg_5____closed__4; +x_16 = l_Lean_Environment_contains(x_14, x_15); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; uint8_t x_19; +lean_dec(x_2); +x_17 = l_Lean_Elab_Tactic_elabGrindConfig___lambda__4___closed__5; +x_18 = l_Lean_throwError___at_Lean_Elab_Term_synthesizeInstMVarCore___spec__3(x_17, 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); +x_19 = !lean_is_exclusive(x_18); +if (x_19 == 0) +{ +return x_18; +} +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_18, 0); +x_21 = lean_ctor_get(x_18, 1); +lean_inc(x_21); +lean_inc(x_20); +lean_dec(x_18); +x_22 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_22, 0, x_20); +lean_ctor_set(x_22, 1, x_21); +return x_22; +} +} +else +{ +lean_object* x_23; lean_object* x_24; +x_23 = lean_box(0); +x_24 = l_Lean_Elab_Tactic_elabGrindConfig___lambda__3(x_1, x_2, x_23, x_4, x_5, x_6, x_7, x_8, x_9, x_13); +return x_24; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_elabGrindConfig(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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; lean_object* x_13; uint8_t x_14; +x_11 = lean_ctor_get_uint8(x_2, sizeof(void*)*1); +lean_inc(x_1); +x_12 = l_Lean_Parser_Tactic_getConfigItems(x_1); +x_13 = l___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_mkConfigItemViews(x_12); +x_14 = l_Array_isEmpty___rarg(x_13); +if (x_14 == 0) +{ +uint8_t x_15; +x_15 = !lean_is_exclusive(x_8); +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_8, 5); +x_17 = l_Lean_replaceRef(x_1, x_16); +lean_dec(x_16); +lean_dec(x_1); +lean_ctor_set(x_8, 5, x_17); +x_18 = lean_box(0); +x_19 = l_Lean_Elab_Tactic_elabGrindConfig___lambda__4(x_11, x_13, x_18, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_19; +} +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t 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; +x_20 = lean_ctor_get(x_8, 0); +x_21 = lean_ctor_get(x_8, 1); +x_22 = lean_ctor_get(x_8, 2); +x_23 = lean_ctor_get(x_8, 3); +x_24 = lean_ctor_get(x_8, 4); +x_25 = lean_ctor_get(x_8, 5); +x_26 = lean_ctor_get(x_8, 6); +x_27 = lean_ctor_get(x_8, 7); +x_28 = lean_ctor_get(x_8, 8); +x_29 = lean_ctor_get(x_8, 9); +x_30 = lean_ctor_get(x_8, 10); +x_31 = lean_ctor_get_uint8(x_8, sizeof(void*)*12); +x_32 = lean_ctor_get(x_8, 11); +x_33 = lean_ctor_get_uint8(x_8, sizeof(void*)*12 + 1); +lean_inc(x_32); +lean_inc(x_30); +lean_inc(x_29); +lean_inc(x_28); +lean_inc(x_27); +lean_inc(x_26); +lean_inc(x_25); +lean_inc(x_24); +lean_inc(x_23); +lean_inc(x_22); +lean_inc(x_21); +lean_inc(x_20); +lean_dec(x_8); +x_34 = l_Lean_replaceRef(x_1, x_25); +lean_dec(x_25); +lean_dec(x_1); +x_35 = lean_alloc_ctor(0, 12, 2); +lean_ctor_set(x_35, 0, x_20); +lean_ctor_set(x_35, 1, x_21); +lean_ctor_set(x_35, 2, x_22); +lean_ctor_set(x_35, 3, x_23); +lean_ctor_set(x_35, 4, x_24); +lean_ctor_set(x_35, 5, x_34); +lean_ctor_set(x_35, 6, x_26); +lean_ctor_set(x_35, 7, x_27); +lean_ctor_set(x_35, 8, x_28); +lean_ctor_set(x_35, 9, x_29); +lean_ctor_set(x_35, 10, x_30); +lean_ctor_set(x_35, 11, x_32); +lean_ctor_set_uint8(x_35, sizeof(void*)*12, x_31); +lean_ctor_set_uint8(x_35, sizeof(void*)*12 + 1, x_33); +x_36 = lean_box(0); +x_37 = l_Lean_Elab_Tactic_elabGrindConfig___lambda__4(x_11, x_13, x_36, x_4, x_5, x_6, x_7, x_35, x_9, x_10); +return x_37; +} +} +else +{ +lean_object* x_38; lean_object* x_39; +lean_dec(x_13); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_38 = l_Lean_Elab_Tactic_elabGrindConfig___lambda__3___closed__1; +x_39 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_10); +return x_39; +} +} +} +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Tactic_elabGrindConfig___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; +x_9 = l_Lean_throwError___at_Lean_Elab_Tactic_elabGrindConfig___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_Tactic_elabGrindConfig___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_Tactic_elabGrindConfig___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_4); +lean_dec(x_2); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_elabGrindConfig___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_Elab_Tactic_elabGrindConfig___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_4); +lean_dec(x_2); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_elabGrindConfig___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +uint8_t x_11; lean_object* x_12; +x_11 = lean_unbox(x_1); +lean_dec(x_1); +x_12 = l_Lean_Elab_Tactic_elabGrindConfig___lambda__3(x_11, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_3); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_elabGrindConfig___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +uint8_t x_11; lean_object* x_12; +x_11 = lean_unbox(x_1); +lean_dec(x_1); +x_12 = l_Lean_Elab_Tactic_elabGrindConfig___lambda__4(x_11, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_3); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_elabGrindConfig___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_Elab_Tactic_elabGrindConfig(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_3); +lean_dec(x_2); +return x_11; +} +} static lean_object* _init_l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_elabGrindPattern___spec__1___rarg___closed__1() { _start: { @@ -698,14 +1605,6 @@ x_1 = lean_mk_string_unchecked("', possible interpretations: ", 29, 29); return x_1; } } -static lean_object* _init_l_Lean_ensureNonAmbiguous___at_Lean_Elab_Tactic_elabGrindPattern___spec__9___closed__7() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("", 0, 0); -return x_1; -} -} LEAN_EXPORT lean_object* l_Lean_ensureNonAmbiguous___at_Lean_Elab_Tactic_elabGrindPattern___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: { @@ -776,7 +1675,7 @@ x_29 = l_List_toString___at_Lean_ensureNoOverload___spec__2(x_28); lean_dec(x_28); x_30 = lean_string_append(x_26, x_29); lean_dec(x_29); -x_31 = l_Lean_ensureNonAmbiguous___at_Lean_Elab_Tactic_elabGrindPattern___spec__9___closed__7; +x_31 = l_Lean_Elab_Tactic_elabGrindConfig___lambda__1___closed__5; x_32 = lean_string_append(x_30, x_31); x_33 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_33, 0, x_32); @@ -863,40 +1762,41 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_unfoldReducible___lambda__3__ return x_1; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_elabGrindPattern___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) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_elabGrindPattern___spec__12(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -uint8_t x_12; -x_12 = lean_usize_dec_lt(x_3, x_2); -if (x_12 == 0) +uint8_t x_13; +x_13 = lean_usize_dec_lt(x_4, x_3); +if (x_13 == 0) { -lean_object* x_13; +lean_object* x_14; +lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -x_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_dec(x_1); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_5); +lean_ctor_set(x_14, 1, x_12); +return x_14; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; lean_object* x_19; -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); -x_17 = lean_box(0); +lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; lean_object* x_19; +x_15 = lean_array_uget(x_5, x_4); +x_16 = lean_unsigned_to_nat(0u); +x_17 = lean_array_uset(x_5, x_4, x_16); x_18 = 1; +lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -lean_inc(x_5); -x_19 = l_Lean_Elab_Term_elabTerm(x_14, x_17, x_18, x_18, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_inc(x_1); +x_19 = l_Lean_Elab_Term_elabTerm(x_15, x_1, x_18, x_18, x_6, x_7, x_8, x_9, x_10, x_11, x_12); if (lean_obj_tag(x_19) == 0) { lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; @@ -905,7 +1805,7 @@ lean_inc(x_20); x_21 = lean_ctor_get(x_19, 1); lean_inc(x_21); lean_dec(x_19); -x_22 = l_Lean_instantiateMVars___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__1(x_20, x_5, x_6, x_7, x_8, x_9, x_10, x_21); +x_22 = l_Lean_instantiateMVars___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__1(x_20, x_6, x_7, x_8, x_9, x_10, x_11, x_21); x_23 = lean_ctor_get(x_22, 0); lean_inc(x_23); x_24 = lean_ctor_get(x_22, 1); @@ -913,11 +1813,11 @@ lean_inc(x_24); lean_dec(x_22); x_25 = l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_elabGrindPattern___spec__12___closed__1; x_26 = l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_elabGrindPattern___spec__12___closed__2; +lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); -lean_inc(x_7); -x_27 = l_Lean_Core_transform___at_Lean_Meta_Grind_unfoldReducible___spec__1(x_23, x_25, x_26, x_7, x_8, x_9, x_10, x_24); +x_27 = l_Lean_Core_transform___at_Lean_Meta_Grind_unfoldReducible___spec__1(x_23, x_25, x_26, x_8, x_9, x_10, x_11, x_24); if (lean_obj_tag(x_27) == 0) { lean_object* x_28; lean_object* x_29; lean_object* x_30; size_t x_31; size_t x_32; lean_object* x_33; @@ -926,26 +1826,27 @@ lean_inc(x_28); x_29 = lean_ctor_get(x_27, 1); lean_inc(x_29); lean_dec(x_27); -x_30 = lean_expr_abstract(x_28, x_1); +x_30 = lean_expr_abstract(x_28, x_2); lean_dec(x_28); x_31 = 1; -x_32 = lean_usize_add(x_3, x_31); -x_33 = lean_array_uset(x_16, x_3, x_30); -x_3 = x_32; -x_4 = x_33; -x_11 = x_29; +x_32 = lean_usize_add(x_4, x_31); +x_33 = lean_array_uset(x_17, x_4, x_30); +x_4 = x_32; +x_5 = x_33; +x_12 = x_29; goto _start; } else { uint8_t x_35; -lean_dec(x_16); +lean_dec(x_17); +lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); +lean_dec(x_1); x_35 = !lean_is_exclusive(x_27); if (x_35 == 0) { @@ -969,13 +1870,14 @@ return x_38; else { uint8_t x_39; -lean_dec(x_16); +lean_dec(x_17); +lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); +lean_dec(x_1); x_39 = !lean_is_exclusive(x_19); if (x_39 == 0) { @@ -1071,56 +1973,56 @@ x_2 = lean_alloc_closure((void*)(l_Lean_Meta_forallTelescope___at_Lean_Elab_Tact return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_elabGrindPattern___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_elabGrindPattern___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -lean_object* x_12; size_t x_13; size_t x_14; lean_object* x_15; -x_12 = l_Lean_Syntax_TSepArray_getElems___rarg(x_1); -x_13 = lean_array_size(x_12); -x_14 = 0; +lean_object* x_13; size_t x_14; size_t x_15; lean_object* x_16; +x_13 = l_Lean_Syntax_TSepArray_getElems___rarg(x_1); +x_14 = lean_array_size(x_13); +x_15 = 0; +lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); -lean_inc(x_7); -x_15 = l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_elabGrindPattern___spec__12(x_3, x_13, x_14, x_12, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_15) == 0) +x_16 = l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_elabGrindPattern___spec__12(x_2, x_4, x_14, x_15, x_13, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +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, 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; +x_17 = lean_ctor_get(x_16, 0); lean_inc(x_17); -lean_dec(x_15); -x_18 = lean_array_get_size(x_3); -x_19 = lean_array_to_list(x_16); -x_20 = l_Lean_Meta_Grind_addTheoremPattern(x_2, x_18, x_19, x_7, x_8, x_9, x_10, x_17); -return x_20; +x_18 = lean_ctor_get(x_16, 1); +lean_inc(x_18); +lean_dec(x_16); +x_19 = lean_array_get_size(x_4); +x_20 = lean_array_to_list(x_17); +x_21 = l_Lean_Meta_Grind_addEMatchTheorem(x_3, x_19, x_20, x_8, x_9, x_10, x_11, x_18); +return x_21; } else { -uint8_t x_21; +uint8_t x_22; +lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_2); -x_21 = !lean_is_exclusive(x_15); -if (x_21 == 0) +lean_dec(x_3); +x_22 = !lean_is_exclusive(x_16); +if (x_22 == 0) { -return x_15; +return x_16; } 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_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_16, 0); +x_24 = lean_ctor_get(x_16, 1); +lean_inc(x_24); lean_inc(x_23); -lean_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_dec(x_16); +x_25 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_24); +return x_25; } } } @@ -1135,6 +2037,7 @@ lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); +lean_inc(x_1); x_10 = l_Lean_resolveGlobalConstNoOverload___at_Lean_Elab_Tactic_elabGrindPattern___spec__2(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9); if (lean_obj_tag(x_10) == 0) { @@ -1146,27 +2049,116 @@ lean_inc(x_12); lean_dec(x_10); lean_inc(x_3); lean_inc(x_11); -x_13 = l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__1(x_11, x_3, x_4, x_5, x_6, x_7, x_8, x_12); +x_13 = l_Lean_mkConstWithLevelParams___at_Lean_Elab_Term_expandDeclId___spec__7(x_11, 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; lean_object* x_17; uint8_t x_18; lean_object* x_19; -x_14 = lean_ctor_get(x_13, 0); -lean_inc(x_14); -x_15 = lean_ctor_get(x_13, 1); -lean_inc(x_15); -lean_dec(x_13); -x_16 = l_Lean_ConstantInfo_type(x_14); -lean_dec(x_14); -x_17 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_elabGrindPattern___lambda__1___boxed), 11, 2); -lean_closure_set(x_17, 0, x_2); -lean_closure_set(x_17, 1, x_11); -x_18 = 0; -x_19 = l_Lean_Meta_forallTelescope___at_Lean_Elab_Tactic_elabGrindPattern___spec__13___rarg(x_16, x_17, x_18, x_3, x_4, x_5, x_6, x_7, x_8, x_15); +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; lean_object* x_19; +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +x_16 = lean_box(0); +x_17 = lean_box(0); +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_19 = l_Lean_Elab_Term_addTermInfo(x_1, x_14, x_16, x_16, x_17, x_18, x_18, x_3, x_4, x_5, x_6, x_7, x_8, x_15); +if (lean_obj_tag(x_19) == 0) +{ +lean_object* x_20; lean_object* x_21; +x_20 = lean_ctor_get(x_19, 1); +lean_inc(x_20); +lean_dec(x_19); +lean_inc(x_3); +lean_inc(x_11); +x_21 = l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__1(x_11, x_3, x_4, x_5, x_6, x_7, x_8, x_20); +if (lean_obj_tag(x_21) == 0) +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_22 = lean_ctor_get(x_21, 0); +lean_inc(x_22); +x_23 = lean_ctor_get(x_21, 1); +lean_inc(x_23); +lean_dec(x_21); +x_24 = l_Lean_ConstantInfo_type(x_22); +lean_dec(x_22); +x_25 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_elabGrindPattern___lambda__1___boxed), 12, 3); +lean_closure_set(x_25, 0, x_2); +lean_closure_set(x_25, 1, x_16); +lean_closure_set(x_25, 2, x_11); +x_26 = l_Lean_Meta_forallTelescope___at_Lean_Elab_Tactic_elabGrindPattern___spec__13___rarg(x_24, x_25, x_18, x_3, x_4, x_5, x_6, x_7, x_8, x_23); +return x_26; +} +else +{ +uint8_t x_27; +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_27 = !lean_is_exclusive(x_21); +if (x_27 == 0) +{ +return x_21; +} +else +{ +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; +} +} +} +else +{ +uint8_t x_31; +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_31 = !lean_is_exclusive(x_19); +if (x_31 == 0) +{ return x_19; } else { -uint8_t x_20; +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_19, 0); +x_33 = lean_ctor_get(x_19, 1); +lean_inc(x_33); +lean_inc(x_32); +lean_dec(x_19); +x_34 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +return x_34; +} +} +} +else +{ +uint8_t x_35; lean_dec(x_11); lean_dec(x_8); lean_dec(x_7); @@ -1175,29 +2167,30 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_20 = !lean_is_exclusive(x_13); -if (x_20 == 0) +lean_dec(x_1); +x_35 = !lean_is_exclusive(x_13); +if (x_35 == 0) { return x_13; } else { -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_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_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; +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_24; +uint8_t x_39; lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -1205,23 +2198,24 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_24 = !lean_is_exclusive(x_10); -if (x_24 == 0) +lean_dec(x_1); +x_39 = !lean_is_exclusive(x_10); +if (x_39 == 0) { return x_10; } else { -lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_25 = lean_ctor_get(x_10, 0); -x_26 = lean_ctor_get(x_10, 1); -lean_inc(x_26); -lean_inc(x_25); +lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_40 = lean_ctor_get(x_10, 0); +x_41 = lean_ctor_get(x_10, 1); +lean_inc(x_41); +lean_inc(x_40); lean_dec(x_10); -x_27 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_27, 0, x_25); -lean_ctor_set(x_27, 1, x_26); -return x_27; +x_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; } } } @@ -1230,19 +2224,11 @@ static lean_object* _init_l_Lean_Elab_Tactic_elabGrindPattern___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean", 4, 4); -return x_1; -} -} -static lean_object* _init_l_Lean_Elab_Tactic_elabGrindPattern___closed__2() { -_start: -{ -lean_object* x_1; x_1 = lean_mk_string_unchecked("Parser", 6, 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_elabGrindPattern___closed__3() { +static lean_object* _init_l_Lean_Elab_Tactic_elabGrindPattern___closed__2() { _start: { lean_object* x_1; @@ -1250,7 +2236,7 @@ x_1 = lean_mk_string_unchecked("Command", 7, 7); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_elabGrindPattern___closed__4() { +static lean_object* _init_l_Lean_Elab_Tactic_elabGrindPattern___closed__3() { _start: { lean_object* x_1; @@ -1258,19 +2244,19 @@ x_1 = lean_mk_string_unchecked("grindPattern", 12, 12); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_elabGrindPattern___closed__5() { +static lean_object* _init_l_Lean_Elab_Tactic_elabGrindPattern___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_Tactic_elabGrindPattern___closed__1; -x_2 = l_Lean_Elab_Tactic_elabGrindPattern___closed__2; -x_3 = l_Lean_Elab_Tactic_elabGrindPattern___closed__3; -x_4 = l_Lean_Elab_Tactic_elabGrindPattern___closed__4; +x_1 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Grind___hyg_5____closed__1; +x_2 = l_Lean_Elab_Tactic_elabGrindPattern___closed__1; +x_3 = l_Lean_Elab_Tactic_elabGrindPattern___closed__2; +x_4 = l_Lean_Elab_Tactic_elabGrindPattern___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_Tactic_elabGrindPattern___closed__6() { +static lean_object* _init_l_Lean_Elab_Tactic_elabGrindPattern___closed__5() { _start: { lean_object* x_1; @@ -1278,12 +2264,12 @@ x_1 = lean_mk_string_unchecked("ident", 5, 5); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_elabGrindPattern___closed__7() { +static lean_object* _init_l_Lean_Elab_Tactic_elabGrindPattern___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic_elabGrindPattern___closed__6; +x_2 = l_Lean_Elab_Tactic_elabGrindPattern___closed__5; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } @@ -1292,7 +2278,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_elabGrindPattern(lean_object* x_1, l _start: { lean_object* x_5; uint8_t x_6; -x_5 = l_Lean_Elab_Tactic_elabGrindPattern___closed__5; +x_5 = l_Lean_Elab_Tactic_elabGrindPattern___closed__4; lean_inc(x_1); x_6 = l_Lean_Syntax_isOfKind(x_1, x_5); if (x_6 == 0) @@ -1307,7 +2293,7 @@ else lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; x_8 = lean_unsigned_to_nat(1u); x_9 = l_Lean_Syntax_getArg(x_1, x_8); -x_10 = l_Lean_Elab_Tactic_elabGrindPattern___closed__7; +x_10 = l_Lean_Elab_Tactic_elabGrindPattern___closed__6; lean_inc(x_9); x_11 = l_Lean_Syntax_isOfKind(x_9, x_10); if (x_11 == 0) @@ -1424,17 +2410,17 @@ lean_dec(x_1); return x_10; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_elabGrindPattern___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) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_elabGrindPattern___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, lean_object* x_12) { _start: { -size_t x_12; size_t x_13; lean_object* x_14; -x_12 = lean_unbox_usize(x_2); -lean_dec(x_2); +size_t x_13; size_t x_14; lean_object* x_15; x_13 = lean_unbox_usize(x_3); lean_dec(x_3); -x_14 = l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_elabGrindPattern___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_1); -return x_14; +x_14 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_15 = l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_elabGrindPattern___spec__12(x_1, x_2, x_13, x_14, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_2); +return x_15; } } LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescope___at_Lean_Elab_Tactic_elabGrindPattern___spec__13___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { @@ -1447,15 +2433,15 @@ x_12 = l_Lean_Meta_forallTelescope___at_Lean_Elab_Tactic_elabGrindPattern___spec return x_12; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_elabGrindPattern___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_elabGrindPattern___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_Lean_Elab_Tactic_elabGrindPattern___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_object* x_13; +x_13 = l_Lean_Elab_Tactic_elabGrindPattern___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); lean_dec(x_4); -lean_dec(x_3); lean_dec(x_1); -return x_12; +return x_13; } } LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_elabGrindPattern___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { @@ -1496,7 +2482,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_elabGrindPattern__1___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l_Lean_Elab_Tactic_elabGrindPattern___closed__1; +x_1 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Grind___hyg_5____closed__1; x_2 = l___regBuiltin_Lean_Elab_Tactic_elabGrindPattern__1___closed__1; x_3 = l___regBuiltin_Lean_Elab_Tactic_elabGrindPattern__1___closed__2; x_4 = l___regBuiltin_Lean_Elab_Tactic_elabGrindPattern__1___closed__3; @@ -1525,7 +2511,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_elabGrindPattern__1(lea { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_2 = l___regBuiltin_Lean_Elab_Tactic_elabGrindPattern__1___closed__5; -x_3 = l_Lean_Elab_Tactic_elabGrindPattern___closed__5; +x_3 = l_Lean_Elab_Tactic_elabGrindPattern___closed__4; x_4 = l___regBuiltin_Lean_Elab_Tactic_elabGrindPattern__1___closed__4; x_5 = l___regBuiltin_Lean_Elab_Tactic_elabGrindPattern__1___closed__6; x_6 = l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(x_2, x_3, x_4, x_5, x_1); @@ -1549,135 +2535,126 @@ x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic_grind___closed__3() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_ensureNonAmbiguous___at_Lean_Elab_Tactic_elabGrindPattern___spec__9___closed__7; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_grind(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_grind(lean_object* x_1, lean_object* x_2, lean_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_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -lean_inc(x_3); -x_8 = l_Lean_Meta_Grind_main(x_1, x_2, x_3, x_4, x_5, x_6, x_7); -if (lean_obj_tag(x_8) == 0) +x_9 = l_Lean_Meta_Grind_main(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_9; -x_9 = !lean_is_exclusive(x_8); -if (x_9 == 0) +uint8_t x_10; +x_10 = !lean_is_exclusive(x_9); +if (x_10 == 0) { -lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_10 = lean_ctor_get(x_8, 0); -x_11 = lean_ctor_get(x_8, 1); -x_12 = l_List_isEmpty___rarg(x_10); -if (x_12 == 0) +lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_11 = lean_ctor_get(x_9, 0); +x_12 = lean_ctor_get(x_9, 1); +x_13 = l_List_isEmpty___rarg(x_11); +if (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_free_object(x_8); -x_13 = l_Lean_Elab_goalsToMessageData(x_10); -x_14 = l_Lean_Elab_Tactic_grind___closed__2; -x_15 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_15, 0, x_14); -lean_ctor_set(x_15, 1, x_13); -x_16 = l_Lean_Elab_Tactic_grind___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); -x_18 = l_Lean_throwError___at_Lean_Meta_setInlineAttribute___spec__1(x_17, x_3, x_4, x_5, x_6, x_11); +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +lean_free_object(x_9); +x_14 = l_Lean_Elab_goalsToMessageData(x_11); +x_15 = l_Lean_Elab_Tactic_grind___closed__2; +x_16 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_16, 0, x_15); +lean_ctor_set(x_16, 1, x_14); +x_17 = l_Lean_Elab_Tactic_elabGrindConfig___lambda__1___closed__6; +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___at_Lean_Meta_setInlineAttribute___spec__1(x_18, x_4, x_5, x_6, x_7, x_12); +lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); -return x_18; +return x_19; } else { -lean_object* x_19; -lean_dec(x_10); +lean_object* x_20; +lean_dec(x_11); +lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); -x_19 = lean_box(0); -lean_ctor_set(x_8, 0, x_19); -return x_8; +x_20 = lean_box(0); +lean_ctor_set(x_9, 0, x_20); +return x_9; } } else { -lean_object* x_20; lean_object* x_21; uint8_t x_22; -x_20 = lean_ctor_get(x_8, 0); -x_21 = lean_ctor_get(x_8, 1); +lean_object* x_21; lean_object* x_22; uint8_t x_23; +x_21 = lean_ctor_get(x_9, 0); +x_22 = lean_ctor_get(x_9, 1); +lean_inc(x_22); lean_inc(x_21); -lean_inc(x_20); -lean_dec(x_8); -x_22 = l_List_isEmpty___rarg(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; -x_23 = l_Lean_Elab_goalsToMessageData(x_20); -x_24 = l_Lean_Elab_Tactic_grind___closed__2; -x_25 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_25, 0, x_24); -lean_ctor_set(x_25, 1, x_23); -x_26 = l_Lean_Elab_Tactic_grind___closed__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 = l_Lean_throwError___at_Lean_Meta_setInlineAttribute___spec__1(x_27, x_3, x_4, x_5, x_6, x_21); +lean_dec(x_9); +x_23 = l_List_isEmpty___rarg(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; +x_24 = l_Lean_Elab_goalsToMessageData(x_21); +x_25 = l_Lean_Elab_Tactic_grind___closed__2; +x_26 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set(x_26, 1, x_24); +x_27 = l_Lean_Elab_Tactic_elabGrindConfig___lambda__1___closed__6; +x_28 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_28, 0, x_26); +lean_ctor_set(x_28, 1, x_27); +x_29 = l_Lean_throwError___at_Lean_Meta_setInlineAttribute___spec__1(x_28, x_4, x_5, x_6, x_7, x_22); +lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); -return x_28; +return x_29; } else { -lean_object* x_29; lean_object* x_30; -lean_dec(x_20); +lean_object* x_30; lean_object* x_31; +lean_dec(x_21); +lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); -x_29 = lean_box(0); -x_30 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_30, 0, x_29); -lean_ctor_set(x_30, 1, x_21); -return x_30; +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_22); +return x_31; } } } else { -uint8_t x_31; +uint8_t x_32; +lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); -x_31 = !lean_is_exclusive(x_8); -if (x_31 == 0) +x_32 = !lean_is_exclusive(x_9); +if (x_32 == 0) { -return x_8; +return x_9; } else { -lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_32 = lean_ctor_get(x_8, 0); -x_33 = lean_ctor_get(x_8, 1); +lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_33 = lean_ctor_get(x_9, 0); +x_34 = lean_ctor_get(x_9, 1); +lean_inc(x_34); lean_inc(x_33); -lean_inc(x_32); -lean_dec(x_8); -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_9); +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; } } } @@ -1701,12 +2678,12 @@ x_9 = lean_alloc_closure((void*)(l_Lean_Elab_throwUnsupportedSyntax___at_Lean_El return x_9; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalApplyRfl___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalApplyRfl___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { -lean_object* x_8; -x_8 = l_Lean_Elab_Tactic_grind(x_2, x_1, x_3, x_4, x_5, x_6, x_7); -return x_8; +lean_object* x_9; +x_9 = l_Lean_Elab_Tactic_grind(x_3, x_1, x_2, x_4, x_5, x_6, x_7, x_8); +return x_9; } } static lean_object* _init_l_Lean_Elab_Tactic_evalApplyRfl___closed__1() { @@ -1721,8 +2698,8 @@ static lean_object* _init_l_Lean_Elab_Tactic_evalApplyRfl___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l_Lean_Elab_Tactic_elabGrindPattern___closed__1; -x_2 = l_Lean_Elab_Tactic_elabGrindPattern___closed__2; +x_1 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Grind___hyg_5____closed__1; +x_2 = l_Lean_Elab_Tactic_elabGrindPattern___closed__1; x_3 = l___regBuiltin_Lean_Elab_Tactic_elabGrindPattern__1___closed__2; x_4 = l_Lean_Elab_Tactic_evalApplyRfl___closed__1; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); @@ -1733,30 +2710,50 @@ static lean_object* _init_l_Lean_Elab_Tactic_evalApplyRfl___closed__3() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("The `grind` tactic is experimental and still under development. Avoid using it in production projects", 101, 101); +x_1 = lean_mk_string_unchecked("optConfig", 9, 9); return x_1; } } static lean_object* _init_l_Lean_Elab_Tactic_evalApplyRfl___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_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Grind___hyg_5____closed__1; +x_2 = l_Lean_Elab_Tactic_elabGrindPattern___closed__1; +x_3 = l___regBuiltin_Lean_Elab_Tactic_elabGrindPattern__1___closed__2; +x_4 = l_Lean_Elab_Tactic_evalApplyRfl___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_Tactic_evalApplyRfl___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("The `grind` tactic is experimental and still under development. Avoid using it in production projects", 101, 101); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Tactic_evalApplyRfl___closed__6() { +_start: +{ lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic_evalApplyRfl___closed__3; +x_1 = l_Lean_Elab_Tactic_evalApplyRfl___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_Elab_Tactic_evalApplyRfl___closed__5() { +static lean_object* _init_l_Lean_Elab_Tactic_evalApplyRfl___closed__7() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic_evalApplyRfl___closed__4; +x_1 = l_Lean_Elab_Tactic_evalApplyRfl___closed__6; x_2 = l_Lean_MessageData_ofFormat(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalApplyRfl___closed__6() { +static lean_object* _init_l_Lean_Elab_Tactic_evalApplyRfl___closed__8() { _start: { lean_object* x_1; @@ -1764,12 +2761,12 @@ x_1 = lean_mk_string_unchecked("_grind", 6, 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalApplyRfl___closed__7() { +static lean_object* _init_l_Lean_Elab_Tactic_evalApplyRfl___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic_evalApplyRfl___closed__6; +x_2 = l_Lean_Elab_Tactic_evalApplyRfl___closed__8; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } @@ -1798,47 +2795,155 @@ return x_13; } else { -lean_object* x_14; uint8_t x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_14 = l_Lean_Elab_Tactic_evalApplyRfl___closed__5; -x_15 = 1; +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_14 = lean_unsigned_to_nat(1u); +x_15 = l_Lean_Syntax_getArg(x_1, x_14); +x_16 = l_Lean_Elab_Tactic_evalApplyRfl___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_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_18 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalApplyRfl___spec__1___rarg(x_10); +return x_18; +} +else +{ +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; +x_19 = l_Lean_Elab_Tactic_evalApplyRfl___closed__7; +x_20 = 1; lean_inc(x_8); -x_16 = l_Lean_logAt___at_Lean_Elab_Tactic_closeUsingOrAdmit___spec__2(x_1, x_14, x_15, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_21 = l_Lean_logAt___at_Lean_Elab_Tactic_closeUsingOrAdmit___spec__2(x_1, x_19, x_20, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); lean_dec(x_1); -x_17 = lean_ctor_get(x_16, 1); -lean_inc(x_17); -lean_dec(x_16); -x_18 = l_Lean_Elab_Term_getDeclName_x3f(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); -if (lean_obj_tag(x_19) == 0) +x_22 = lean_ctor_get(x_21, 1); +lean_inc(x_22); +lean_dec(x_21); +x_23 = l_Lean_Elab_Term_getDeclName_x3f(x_4, x_5, x_6, x_7, x_8, x_9, x_22); +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +x_25 = lean_ctor_get(x_23, 1); +lean_inc(x_25); +lean_dec(x_23); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +x_26 = l_Lean_Elab_Tactic_elabGrindConfig(x_15, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_25); +if (lean_obj_tag(x_24) == 0) { -lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_20 = lean_ctor_get(x_18, 1); -lean_inc(x_20); -lean_dec(x_18); -x_21 = l_Lean_Elab_Tactic_evalApplyRfl___closed__7; -x_22 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_evalApplyRfl___lambda__1), 7, 1); -lean_closure_set(x_22, 0, x_21); -x_23 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_liftMetaFinishingTactic), 10, 1); -lean_closure_set(x_23, 0, x_22); -x_24 = l_Lean_Elab_Tactic_withMainContext___rarg(x_23, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_20); -return x_24; +if (lean_obj_tag(x_26) == 0) +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; +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_evalApplyRfl___closed__9; +x_30 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_evalApplyRfl___lambda__1), 8, 2); +lean_closure_set(x_30, 0, x_27); +lean_closure_set(x_30, 1, x_29); +x_31 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_liftMetaFinishingTactic), 10, 1); +lean_closure_set(x_31, 0, x_30); +x_32 = l_Lean_Elab_Tactic_withMainContext___rarg(x_31, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_28); +return x_32; } else { -lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_25 = lean_ctor_get(x_18, 1); -lean_inc(x_25); -lean_dec(x_18); -x_26 = lean_ctor_get(x_19, 0); -lean_inc(x_26); -lean_dec(x_19); -x_27 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_evalApplyRfl___lambda__1), 7, 1); -lean_closure_set(x_27, 0, x_26); -x_28 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_liftMetaFinishingTactic), 10, 1); -lean_closure_set(x_28, 0, x_27); -x_29 = l_Lean_Elab_Tactic_withMainContext___rarg(x_28, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_25); -return x_29; +uint8_t x_33; +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_33 = !lean_is_exclusive(x_26); +if (x_33 == 0) +{ +return x_26; +} +else +{ +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_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 +{ +if (lean_obj_tag(x_26) == 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_24, 0); +lean_inc(x_37); +lean_dec(x_24); +x_38 = lean_ctor_get(x_26, 0); +lean_inc(x_38); +x_39 = lean_ctor_get(x_26, 1); +lean_inc(x_39); +lean_dec(x_26); +x_40 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_evalApplyRfl___lambda__1), 8, 2); +lean_closure_set(x_40, 0, x_38); +lean_closure_set(x_40, 1, x_37); +x_41 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_liftMetaFinishingTactic), 10, 1); +lean_closure_set(x_41, 0, x_40); +x_42 = l_Lean_Elab_Tactic_withMainContext___rarg(x_41, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_39); +return x_42; +} +else +{ +uint8_t x_43; +lean_dec(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_2); +x_43 = !lean_is_exclusive(x_26); +if (x_43 == 0) +{ +return x_26; +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_26, 0); +x_45 = lean_ctor_get(x_26, 1); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_26); +x_46 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_46, 0, x_44); +lean_ctor_set(x_46, 1, x_45); +return x_46; +} +} +} } } } @@ -1871,7 +2976,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalApplyRfl__1___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l_Lean_Elab_Tactic_elabGrindPattern___closed__1; +x_1 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Grind___hyg_5____closed__1; x_2 = l___regBuiltin_Lean_Elab_Tactic_elabGrindPattern__1___closed__1; x_3 = l___regBuiltin_Lean_Elab_Tactic_elabGrindPattern__1___closed__2; x_4 = l___regBuiltin_Lean_Elab_Tactic_evalApplyRfl__1___closed__1; @@ -1911,6 +3016,7 @@ lean_object* initialize_Init_Grind_Tactics(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_Tactic_Grind(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Elab_Command(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Elab_Tactic_Basic(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Elab_Tactic_Config(uint8_t builtin, lean_object*); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Lean_Elab_Tactic_Grind(uint8_t builtin, lean_object* w) { lean_object * res; @@ -1928,6 +3034,45 @@ lean_dec_ref(res); res = initialize_Lean_Elab_Tactic_Basic(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +res = initialize_Lean_Elab_Tactic_Config(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Grind___hyg_5____closed__1 = _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Grind___hyg_5____closed__1(); +lean_mark_persistent(l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Grind___hyg_5____closed__1); +l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Grind___hyg_5____closed__2 = _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Grind___hyg_5____closed__2(); +lean_mark_persistent(l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Grind___hyg_5____closed__2); +l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Grind___hyg_5____closed__3 = _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Grind___hyg_5____closed__3(); +lean_mark_persistent(l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Grind___hyg_5____closed__3); +l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Grind___hyg_5____closed__4 = _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Grind___hyg_5____closed__4(); +lean_mark_persistent(l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Grind___hyg_5____closed__4); +l_Lean_Elab_Tactic_elabGrindConfig___lambda__1___closed__1 = _init_l_Lean_Elab_Tactic_elabGrindConfig___lambda__1___closed__1(); +lean_mark_persistent(l_Lean_Elab_Tactic_elabGrindConfig___lambda__1___closed__1); +l_Lean_Elab_Tactic_elabGrindConfig___lambda__1___closed__2 = _init_l_Lean_Elab_Tactic_elabGrindConfig___lambda__1___closed__2(); +lean_mark_persistent(l_Lean_Elab_Tactic_elabGrindConfig___lambda__1___closed__2); +l_Lean_Elab_Tactic_elabGrindConfig___lambda__1___closed__3 = _init_l_Lean_Elab_Tactic_elabGrindConfig___lambda__1___closed__3(); +lean_mark_persistent(l_Lean_Elab_Tactic_elabGrindConfig___lambda__1___closed__3); +l_Lean_Elab_Tactic_elabGrindConfig___lambda__1___closed__4 = _init_l_Lean_Elab_Tactic_elabGrindConfig___lambda__1___closed__4(); +lean_mark_persistent(l_Lean_Elab_Tactic_elabGrindConfig___lambda__1___closed__4); +l_Lean_Elab_Tactic_elabGrindConfig___lambda__1___closed__5 = _init_l_Lean_Elab_Tactic_elabGrindConfig___lambda__1___closed__5(); +lean_mark_persistent(l_Lean_Elab_Tactic_elabGrindConfig___lambda__1___closed__5); +l_Lean_Elab_Tactic_elabGrindConfig___lambda__1___closed__6 = _init_l_Lean_Elab_Tactic_elabGrindConfig___lambda__1___closed__6(); +lean_mark_persistent(l_Lean_Elab_Tactic_elabGrindConfig___lambda__1___closed__6); +l_Lean_Elab_Tactic_elabGrindConfig___lambda__2___closed__1 = _init_l_Lean_Elab_Tactic_elabGrindConfig___lambda__2___closed__1(); +lean_mark_persistent(l_Lean_Elab_Tactic_elabGrindConfig___lambda__2___closed__1); +l_Lean_Elab_Tactic_elabGrindConfig___lambda__2___closed__2 = _init_l_Lean_Elab_Tactic_elabGrindConfig___lambda__2___closed__2(); +lean_mark_persistent(l_Lean_Elab_Tactic_elabGrindConfig___lambda__2___closed__2); +l_Lean_Elab_Tactic_elabGrindConfig___lambda__3___closed__1 = _init_l_Lean_Elab_Tactic_elabGrindConfig___lambda__3___closed__1(); +lean_mark_persistent(l_Lean_Elab_Tactic_elabGrindConfig___lambda__3___closed__1); +l_Lean_Elab_Tactic_elabGrindConfig___lambda__4___closed__1 = _init_l_Lean_Elab_Tactic_elabGrindConfig___lambda__4___closed__1(); +lean_mark_persistent(l_Lean_Elab_Tactic_elabGrindConfig___lambda__4___closed__1); +l_Lean_Elab_Tactic_elabGrindConfig___lambda__4___closed__2 = _init_l_Lean_Elab_Tactic_elabGrindConfig___lambda__4___closed__2(); +lean_mark_persistent(l_Lean_Elab_Tactic_elabGrindConfig___lambda__4___closed__2); +l_Lean_Elab_Tactic_elabGrindConfig___lambda__4___closed__3 = _init_l_Lean_Elab_Tactic_elabGrindConfig___lambda__4___closed__3(); +lean_mark_persistent(l_Lean_Elab_Tactic_elabGrindConfig___lambda__4___closed__3); +l_Lean_Elab_Tactic_elabGrindConfig___lambda__4___closed__4 = _init_l_Lean_Elab_Tactic_elabGrindConfig___lambda__4___closed__4(); +lean_mark_persistent(l_Lean_Elab_Tactic_elabGrindConfig___lambda__4___closed__4); +l_Lean_Elab_Tactic_elabGrindConfig___lambda__4___closed__5 = _init_l_Lean_Elab_Tactic_elabGrindConfig___lambda__4___closed__5(); +lean_mark_persistent(l_Lean_Elab_Tactic_elabGrindConfig___lambda__4___closed__5); l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_elabGrindPattern___spec__1___rarg___closed__1 = _init_l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_elabGrindPattern___spec__1___rarg___closed__1(); lean_mark_persistent(l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_elabGrindPattern___spec__1___rarg___closed__1); l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_elabGrindPattern___spec__1___rarg___closed__2 = _init_l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_elabGrindPattern___spec__1___rarg___closed__2(); @@ -1964,8 +3109,6 @@ l_Lean_ensureNonAmbiguous___at_Lean_Elab_Tactic_elabGrindPattern___spec__9___clo lean_mark_persistent(l_Lean_ensureNonAmbiguous___at_Lean_Elab_Tactic_elabGrindPattern___spec__9___closed__5); l_Lean_ensureNonAmbiguous___at_Lean_Elab_Tactic_elabGrindPattern___spec__9___closed__6 = _init_l_Lean_ensureNonAmbiguous___at_Lean_Elab_Tactic_elabGrindPattern___spec__9___closed__6(); lean_mark_persistent(l_Lean_ensureNonAmbiguous___at_Lean_Elab_Tactic_elabGrindPattern___spec__9___closed__6); -l_Lean_ensureNonAmbiguous___at_Lean_Elab_Tactic_elabGrindPattern___spec__9___closed__7 = _init_l_Lean_ensureNonAmbiguous___at_Lean_Elab_Tactic_elabGrindPattern___spec__9___closed__7(); -lean_mark_persistent(l_Lean_ensureNonAmbiguous___at_Lean_Elab_Tactic_elabGrindPattern___spec__9___closed__7); l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_elabGrindPattern___spec__12___closed__1 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_elabGrindPattern___spec__12___closed__1(); lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_elabGrindPattern___spec__12___closed__1); l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_elabGrindPattern___spec__12___closed__2 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_elabGrindPattern___spec__12___closed__2(); @@ -1982,8 +3125,6 @@ l_Lean_Elab_Tactic_elabGrindPattern___closed__5 = _init_l_Lean_Elab_Tactic_elabG lean_mark_persistent(l_Lean_Elab_Tactic_elabGrindPattern___closed__5); l_Lean_Elab_Tactic_elabGrindPattern___closed__6 = _init_l_Lean_Elab_Tactic_elabGrindPattern___closed__6(); lean_mark_persistent(l_Lean_Elab_Tactic_elabGrindPattern___closed__6); -l_Lean_Elab_Tactic_elabGrindPattern___closed__7 = _init_l_Lean_Elab_Tactic_elabGrindPattern___closed__7(); -lean_mark_persistent(l_Lean_Elab_Tactic_elabGrindPattern___closed__7); l___regBuiltin_Lean_Elab_Tactic_elabGrindPattern__1___closed__1 = _init_l___regBuiltin_Lean_Elab_Tactic_elabGrindPattern__1___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Elab_Tactic_elabGrindPattern__1___closed__1); l___regBuiltin_Lean_Elab_Tactic_elabGrindPattern__1___closed__2 = _init_l___regBuiltin_Lean_Elab_Tactic_elabGrindPattern__1___closed__2(); @@ -2003,8 +3144,6 @@ lean_dec_ref(res); lean_mark_persistent(l_Lean_Elab_Tactic_grind___closed__1); l_Lean_Elab_Tactic_grind___closed__2 = _init_l_Lean_Elab_Tactic_grind___closed__2(); lean_mark_persistent(l_Lean_Elab_Tactic_grind___closed__2); -l_Lean_Elab_Tactic_grind___closed__3 = _init_l_Lean_Elab_Tactic_grind___closed__3(); -lean_mark_persistent(l_Lean_Elab_Tactic_grind___closed__3); l_Lean_Elab_Tactic_evalApplyRfl___closed__1 = _init_l_Lean_Elab_Tactic_evalApplyRfl___closed__1(); lean_mark_persistent(l_Lean_Elab_Tactic_evalApplyRfl___closed__1); l_Lean_Elab_Tactic_evalApplyRfl___closed__2 = _init_l_Lean_Elab_Tactic_evalApplyRfl___closed__2(); @@ -2019,6 +3158,10 @@ l_Lean_Elab_Tactic_evalApplyRfl___closed__6 = _init_l_Lean_Elab_Tactic_evalApply lean_mark_persistent(l_Lean_Elab_Tactic_evalApplyRfl___closed__6); l_Lean_Elab_Tactic_evalApplyRfl___closed__7 = _init_l_Lean_Elab_Tactic_evalApplyRfl___closed__7(); lean_mark_persistent(l_Lean_Elab_Tactic_evalApplyRfl___closed__7); +l_Lean_Elab_Tactic_evalApplyRfl___closed__8 = _init_l_Lean_Elab_Tactic_evalApplyRfl___closed__8(); +lean_mark_persistent(l_Lean_Elab_Tactic_evalApplyRfl___closed__8); +l_Lean_Elab_Tactic_evalApplyRfl___closed__9 = _init_l_Lean_Elab_Tactic_evalApplyRfl___closed__9(); +lean_mark_persistent(l_Lean_Elab_Tactic_evalApplyRfl___closed__9); l___regBuiltin_Lean_Elab_Tactic_evalApplyRfl__1___closed__1 = _init_l___regBuiltin_Lean_Elab_Tactic_evalApplyRfl__1___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Elab_Tactic_evalApplyRfl__1___closed__1); l___regBuiltin_Lean_Elab_Tactic_evalApplyRfl__1___closed__2 = _init_l___regBuiltin_Lean_Elab_Tactic_evalApplyRfl__1___closed__2(); diff --git a/stage0/stdlib/Lean/Elab/Tactic/Monotonicity.c b/stage0/stdlib/Lean/Elab/Tactic/Monotonicity.c new file mode 100644 index 000000000000..3c92bd4b0bd7 --- /dev/null +++ b/stage0/stdlib/Lean/Elab/Tactic/Monotonicity.c @@ -0,0 +1,12007 @@ +// Lean compiler output +// Module: Lean.Elab.Tactic.Monotonicity +// Imports: Lean.Meta.Tactic.Split Lean.Elab.RecAppSyntax Lean.Elab.Tactic.Basic Init.Internal.Order +#include +#if defined(__clang__) +#pragma clang diagnostic ignored "-Wunused-parameter" +#pragma clang diagnostic ignored "-Wunused-label" +#elif defined(__GNUC__) && !defined(__CLANG__) +#pragma GCC diagnostic ignored "-Wunused-parameter" +#pragma GCC diagnostic ignored "-Wunused-label" +#pragma GCC diagnostic ignored "-Wunused-but-set-variable" +#endif +#ifdef __cplusplus +extern "C" { +#endif +lean_object* l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_createNodes___rarg(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Expr_bindingName_x21(lean_object*); +static lean_object* l_Lean_Meta_DiscrTree_insertCore___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__1___closed__2; +static lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__23___closed__1; +static lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__15___closed__6; +static size_t l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__3___closed__1; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__3___boxed(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_List_forM___at_Lean_Meta_Monotonicity_solveMono___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__12___closed__1; +static lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__25___closed__2; +static lean_object* l_Lean_Loop_forIn_loop___at_Lean_Meta_Monotonicity_headBetaUnderLambda___spec__1___closed__4; +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__17___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_headBetaUnderLambda(lean_object*); +static lean_object* l_Lean_Meta_Monotonicity_solveMonoCall___lambda__2___closed__4; +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_monotoneExt; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_Monotonicity_solveMonoStep___spec__3___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__6(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__8___boxed(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_Meta_mkConstWithFreshMVarLevels___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_mkAppOptM(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_5433____closed__5; +lean_object* l_Lean_Meta_withLetDecl___at___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_isPerm___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_mk_empty_array_with_capacity(lean_object*); +size_t lean_usize_shift_right(size_t, size_t); +static lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__6___closed__1; +static lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__10; +static lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__6___closed__4; +static lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__7; +static lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__12___closed__9; +lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Monotonicity_solveMonoCall___lambda__3___closed__4; +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMono___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_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__1; +static lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__9___closed__4; +static lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__3; +lean_object* l_Lean_ConstantInfo_type(lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Monotonicity_solveMonoStep___spec__1(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*); +static lean_object* l_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_5433____closed__1; +static lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__18___closed__2; +uint8_t lean_usize_dec_le(size_t, size_t); +lean_object* l_Lean_indentD(lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_BinSearch_0__Array_binInsertAux___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__12(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_projExpr_x21(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__27(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Exception_isInterrupt(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__2; +static lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__6___closed__1; +size_t lean_uint64_to_usize(uint64_t); +uint64_t lean_uint64_lor(uint64_t, uint64_t); +static lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__17; +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__26___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_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__1; +static lean_object* l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_defaultFailK___rarg___closed__5; +static lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__7; +uint8_t l_Lean_Expr_isApp(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Monotonicity_evalMonotonicity___rarg___closed__1; +static lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__1; +extern lean_object* l_Lean_Elab_Tactic_tacticElabAttribute; +static lean_object* l___regBuiltin_Lean_Meta_Monotonicity_evalMonotonicity__1___closed__3; +lean_object* l_Lean_Expr_sort___override(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoCall___lambda__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_MessageData_ofList(lean_object*); +static lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__10; +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static uint64_t l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__14; +LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_applyConst___lambda__1(lean_object*, lean_object*); +lean_object* lean_array_push(lean_object*, lean_object*); +static lean_object* l___regBuiltin_Lean_Meta_Monotonicity_evalMonotonicity__1___closed__4; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166_(lean_object*); +static lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__2; +size_t lean_usize_mul(size_t, size_t); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Monotonicity_solveMonoStep___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_Meta_DiscrTree_instInhabited(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_headBetaUnderLambda___lambda__1(lean_object*, lean_object*); +lean_object* lean_mk_array(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_Monotonicity_solveMonoStep___spec__3(size_t, size_t, lean_object*); +uint8_t lean_usize_dec_eq(size_t, size_t); +static lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__4; +static lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__9; +lean_object* l_Lean_Expr_bvar___override(lean_object*); +lean_object* lean_array_fget(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__20; +static lean_object* l_Lean_Loop_forIn_loop___at_Lean_Meta_Monotonicity_headBetaUnderLambda___spec__1___closed__3; +lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_applyConst___lambda__1___closed__1; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__2(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Elab_Tactic_getMainGoal(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoCall(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__1___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__17(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__27___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Loop_forIn_loop___at_Lean_Meta_Monotonicity_headBetaUnderLambda___spec__1___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoCall___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___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertVal_loop___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__10(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__12; +static lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__23___closed__2; +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Monotonicity_solveMonoCall___lambda__6___closed__1; +static lean_object* l_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_5433____closed__3; +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____spec__1___closed__2; +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Expr_cleanupAnnotations(lean_object*); +LEAN_EXPORT lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_stringToMessageData(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__3(uint64_t, uint64_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_PersistentHashMap_findAux___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__3(lean_object*, size_t, lean_object*); +static lean_object* l_Lean_Meta_Monotonicity_solveMonoCall___lambda__8___closed__2; +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__16(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_5433____closed__7; +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoCall___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_x27_loop___at_Lean_Meta_Monotonicity_solveMonoStep___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___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_defaultFailK___rarg___closed__2; +static lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__10___closed__1; +static lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__6; +LEAN_EXPORT lean_object* l_Lean_Meta_lambdaBoundedTelescope___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____spec__2___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_throwError___at_Lean_Meta_setInlineAttribute___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Exception_toMessageData(lean_object*); +static lean_object* l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_defaultFailK___rarg___closed__6; +lean_object* l_Lean_Expr_appArg_x21(lean_object*); +static lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__25___closed__1; +static lean_object* l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__9___closed__1; +static lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__3; +LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_defaultFailK___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Expr_isBVar(lean_object*); +static lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__7; +static lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__6___closed__2; +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__28(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Expr_isLambda(lean_object*); +static lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__18; +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_object* l_Lean_throwError___at_Lean_Meta_Split_applyMatchSplitter___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +size_t lean_ptr_addr(lean_object*); +lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__28___closed__2; +static lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__9___closed__3; +LEAN_EXPORT lean_object* l_Lean_Meta_lambdaBoundedTelescope___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____spec__2(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_findMonoThms(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__21(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static uint64_t l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__12; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__6(lean_object*, size_t, size_t, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Monotonicity_solveMonoCall___lambda__8___closed__1; +lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_lambdaTelescopeImp___rarg(lean_object*, uint8_t, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_checked_assign(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__5; +static lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__9; +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at_Lean_Meta_Monotonicity_headBetaUnderLambda___spec__1(lean_object*); +LEAN_EXPORT lean_object* l_Array_binInsertM___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__11(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoCall___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__8; +static lean_object* l_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_5433____closed__6; +static lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__11; +static lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__8; +static lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___closed__1; +lean_object* lean_st_ref_take(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Monotonicity_solveMonoCall___lambda__4___closed__4; +static lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__4___closed__1; +uint8_t lean_expr_eqv(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Monotonicity_solveMonoStep___spec__1___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__22; +lean_object* l_Lean_registerSimpleScopedEnvExtension___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_evalMonotonicity___boxed(lean_object*); +uint64_t lean_uint64_shift_right(uint64_t, uint64_t); +static lean_object* l_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_5433____closed__4; +static uint64_t l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__15; +static lean_object* l_Lean_Meta_Monotonicity_solveMonoCall___lambda__4___closed__2; +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoCall___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*); +static lean_object* l_Lean_Meta_Monotonicity_solveMonoCall___lambda__9___closed__1; +static lean_object* l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_applyConst___lambda__1___closed__4; +static lean_object* l_Lean_Meta_Monotonicity_solveMonoCall___lambda__5___closed__2; +lean_object* lean_nat_div(lean_object*, lean_object*); +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*); +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__20___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_registerTraceClass(lean_object*, uint8_t, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_DiscrTree_insertCore___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__1(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__16; +static lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__10; +LEAN_EXPORT lean_object* l_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_5433_(lean_object*); +LEAN_EXPORT lean_object* l_panic___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__13(lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__14; +lean_object* l_Lean_Expr_projIdx_x21(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_defaultFailK___rarg___closed__8; +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__13(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*); +static lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__6___closed__1; +lean_object* l_Lean_MessageData_andList(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_defaultFailK(lean_object*); +lean_object* l_Lean_Meta_Match_MatcherInfo_getFirstDiscrPos(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__22___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_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_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__3; +lean_object* l_Lean_Expr_appArg(lean_object*, lean_object*); +lean_object* l_outOfBounds___rarg(lean_object*); +lean_object* l_Lean_Elab_Tactic_withMainContext___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___regBuiltin_Lean_Meta_Monotonicity_evalMonotonicity__1___closed__5; +LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_applyConst(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_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____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_forIn_x27Unsafe_loop___at_Lean_Meta_Monotonicity_solveMonoStep___spec__1___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Monotonicity_solveMonoCall___lambda__5___closed__4; +lean_object* lean_st_mk_ref(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___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_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_array_to_list(lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Monotonicity_solveMonoStep___spec__1___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__5; +lean_object* l_Lean_Expr_appFnCleanup(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_defaultFailK___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__9; +uint8_t l_Lean_Expr_hasLooseBVars(lean_object*); +uint8_t l_Lean_Expr_isHeadBetaTarget(lean_object*, uint8_t); +lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____spec__1(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Meta_DiscrTree_Key_lt(lean_object*, lean_object*); +static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Monotonicity_solveMonoStep___spec__1___closed__1; +LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Meta_Monotonicity_solveMonoStep___spec__4(lean_object*, lean_object*); +extern lean_object* l_Lean_levelZero; +static lean_object* l_Lean_Meta_Monotonicity_solveMonoCall___lambda__6___closed__2; +static lean_object* l_Lean_Meta_Monotonicity_solveMonoCall___lambda__2___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoCall___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*); +extern lean_object* l_Lean_instInhabitedExpr; +uint8_t l___private_Lean_Meta_DiscrTreeTypes_0__Lean_Meta_DiscrTree_beqKey____x40_Lean_Meta_DiscrTreeTypes___hyg_102_(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__17; +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__15___closed__8; +static lean_object* l___regBuiltin_Lean_Meta_Monotonicity_evalMonotonicity__1___closed__2; +static lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___closed__2; +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__24(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___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_Meta_Monotonicity_headBetaUnderLambda___lambda__1___boxed(lean_object*, lean_object*); +lean_object* l_panic___at_Lean_Expr_appFn_x21___spec__1(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__18(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___boxed(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_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*); +static lean_object* l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_applyConst___lambda__1___closed__2; +static lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__9___closed__7; +LEAN_EXPORT lean_object* l_Lean_Meta_lambdaBoundedTelescope___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____spec__2___rarg(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__6; +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__18___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__21___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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_object* l_Lean_Meta_isMatcherAppCore_x3f(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__1(uint64_t, uint64_t, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Meta_Monotonicity_evalMonotonicity__1(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Expr_appFn_x21(lean_object*); +static lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__19; +static lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__18___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__14(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__6; +static lean_object* l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_defaultFailK___rarg___closed__4; +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_evalMonotonicity___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_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Monotonicity_solveMonoCall___lambda__2___closed__3; +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__25___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__10___closed__2; +lean_object* l_Lean_Meta_DiscrTree_mkPath(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__5; +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_evalMonotonicity(lean_object*); +lean_object* lean_usize_to_nat(size_t); +static lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__4; +lean_object* l_Lean_MessageData_ofConstName(lean_object*, uint8_t); +static lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__15___closed__2; +LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_defaultFailK___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_MessageData_ofExpr(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__9(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__19(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_DiscrTree_insertCore___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__1___closed__1; +static lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__15___closed__7; +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoCall___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_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__4; +static lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___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_EXPORT lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249_(lean_object*); +static lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__10; +static lean_object* l_Lean_Meta_Monotonicity_solveMonoCall___lambda__2___closed__2; +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__20(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_defaultFailK___spec__1(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___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_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____lambda__1(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__21; +static lean_object* l_Lean_Meta_Monotonicity_solveMonoCall___lambda__5___closed__1; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__8(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_defaultFailK___spec__2___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__23; +lean_object* l_Lean_Expr_bindingDomain_x21(lean_object*); +static lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__9___closed__2; +static lean_object* l_Lean_Meta_Monotonicity_solveMonoCall___lambda__4___closed__1; +static lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__15___closed__5; +static lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__8; +static lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__6___closed__3; +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__23(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +extern lean_object* l_Lean_Meta_DiscrTree_instInhabitedKey; +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___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_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___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_Meta_Monotonicity_solveMonoStep___lambda__12___closed__5; +lean_object* l_Lean_Expr_app___override(lean_object*, lean_object*); +lean_object* l_Lean_Meta_mapErrorImp___rarg(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_Meta_Monotonicity_solveMonoCall___lambda__4___closed__3; +static lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__13; +static lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__4; +static lean_object* l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_defaultFailK___rarg___closed__7; +static lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Monotonicity_solveMonoStep___spec__2___closed__1; +uint8_t lean_nat_dec_lt(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_evalMonotonicity___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_Expr_isProj(lean_object*); +lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoCall___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__8; +static uint64_t l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__2; +lean_object* l_id___rarg___boxed(lean_object*); +static lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__9; +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Monotonicity_solveMonoStep___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_indentExpr(lean_object*); +lean_object* l_Lean_PersistentHashMap_mkEmptyEntries(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Monotonicity_solveMonoCall___lambda__3___closed__2; +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_panic___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__13___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___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*); +uint8_t l_Lean_Expr_isConstOf(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__11; +static lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__9___closed__1; +static lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__6; +static lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__12___closed__7; +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_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__9___closed__2; +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__16___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Monotonicity_solveMonoCall___lambda__9___closed__2; +static lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Monotonicity_solveMonoStep___spec__2___closed__2; +static lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__12___closed__3; +lean_object* lean_panic_fn(lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_applyConst___lambda__1___closed__3; +static lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__15___closed__1; +static lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____spec__1___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoCall___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__12___closed__8; +LEAN_EXPORT lean_object* l_Lean_Meta_DiscrTree_insertCore___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__1___boxed(lean_object*, lean_object*, lean_object*); +lean_object* lean_nat_sub(lean_object*, lean_object*); +lean_object* l_Lean_ScopedEnvExtension_getState___rarg(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__15___closed__3; +static lean_object* l_Lean_Loop_forIn_loop___at_Lean_Meta_Monotonicity_headBetaUnderLambda___spec__1___closed__2; +static lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__16; +LEAN_EXPORT lean_object* l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__2___boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__7; +lean_object* l_Lean_MVarId_intro(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___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*); +uint64_t lean_uint64_shift_left(uint64_t, uint64_t); +lean_object* l_Lean_ScopedEnvExtension_addCore___rarg(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoCall___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Monotonicity_solveMonoStep___spec__1___lambda__3(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*); +lean_object* l_Lean_Meta_instantiateMVarsIfMVarApp(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_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Monotonicity_solveMonoStep___spec__1___closed__2; +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_Lean_registerBuiltinAttribute(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__5(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_defaultFailK___rarg___closed__3; +static lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__28___closed__1; +lean_object* l_List_reverse___rarg(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +size_t lean_usize_sub(size_t, size_t); +static lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__2; +lean_object* lean_array_mk(lean_object*); +lean_object* l_List_mapTR_loop___at_Lean_Parser_Tactic_Doc_initFn____x40_Lean_Parser_Tactic_Doc___hyg_2831____spec__8(lean_object*, lean_object*); +static uint64_t l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__13; +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__19___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_add(size_t, size_t); +static lean_object* l_Lean_Meta_Monotonicity_evalMonotonicity___rarg___lambda__1___closed__1; +static lean_object* l_Lean_Meta_DiscrTree_insertCore___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__1___closed__3; +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoCall___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__3; +lean_object* l_Lean_MVarId_applyConst(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_BinSearch_0__Array_binInsertAux___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____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_array_uget(lean_object*, size_t); +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____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*); +size_t lean_array_size(lean_object*); +static lean_object* l_Lean_Meta_Monotonicity_solveMonoCall___lambda__5___closed__3; +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*); +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoCall___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoCall___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_Meta_Monotonicity_solveMonoStep___lambda__14___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +size_t lean_usize_shift_left(size_t, size_t); +lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Expr_bindingBody_x21(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoCall___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_defaultFailK___spec__2(lean_object*); +lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__6___closed__2; +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__2(uint64_t, uint64_t, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__4___closed__2; +static lean_object* l_Lean_Meta_Monotonicity_solveMonoCall___lambda__3___closed__3; +lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Meta_addPPExplicitToExposeDiff_visit___spec__4___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_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__15; +lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__15___closed__4; +static lean_object* l_Lean_Meta_DiscrTree_insertCore___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__1___closed__4; +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_throwError___at___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_defaultFailK___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___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_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Expr_headBeta(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMono(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_Meta_Monotonicity_solveMonoStep___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_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_5433____closed__2; +uint8_t l_Lean_Expr_isMData(lean_object*); +static lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__5___closed__1; +static lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__5; +static lean_object* l_Lean_Meta_Monotonicity_headBetaUnderLambda___closed__1; +lean_object* lean_infer_type(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static size_t l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__3___closed__2; +static lean_object* l___regBuiltin_Lean_Meta_Monotonicity_evalMonotonicity__1___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_Meta_Monotonicity_solveMonoCall___lambda__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__9___closed__6; +lean_object* l_Lean_Meta_mkLambdaFVars(lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_defaultFailK___rarg___closed__1; +uint64_t l_Lean_Meta_TransparencyMode_toUInt64(uint8_t); +lean_object* lean_nat_add(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__12___closed__6; +lean_object* l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(lean_object*); +static lean_object* l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_applyConst___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__22(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_whnfUntil(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__12___closed__2; +uint8_t l_Lean_Exception_isRuntime(lean_object*); +lean_object* l_Lean_Expr_lam___override(lean_object*, lean_object*, lean_object*, uint8_t); +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoCall___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_DiscrTree_getMatch___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_defaultFailK___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Expr_mdataExpr_x21(lean_object*); +lean_object* l_Lean_Expr_bindingInfo_x21(lean_object*); +uint8_t l_Lean_Expr_isForall(lean_object*); +lean_object* l_Lean_Expr_mvarId_x21(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_evalMonotonicity___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_throwError___at_Lean_Attribute_Builtin_ensureNoArgs___spec__2(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Monotonicity_solveMonoStep___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint64_t l___private_Lean_Meta_Basic_0__Lean_Meta_Config_toKey(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_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_object* lean_expr_instantiate1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__7(size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__25(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint64_t l_Lean_Meta_DiscrTree_Key_hash(lean_object*); +LEAN_EXPORT lean_object* l_Array_binInsertM___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Array_insertIdx_loop___rarg(lean_object*, lean_object*, lean_object*); +size_t lean_usize_land(size_t, size_t); +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___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_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_5433____closed__8; +static lean_object* l_Lean_Meta_Monotonicity_solveMonoCall___lambda__3___closed__1; +lean_object* l_Lean_Meta_etaExpand(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__26(lean_object*, 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_Monotonicity_solveMonoCall___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*); +uint8_t l_Array_isEmpty___rarg(lean_object*); +static lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__12___closed__4; +static lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__9___closed__5; +uint8_t l___private_Lean_Expr_0__Lean_beqBinderInfo____x40_Lean_Expr___hyg_406_(uint8_t, uint8_t); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Monotonicity_solveMonoStep___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*); +static lean_object* _init_l_Lean_Loop_forIn_loop___at_Lean_Meta_Monotonicity_headBetaUnderLambda___spec__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.Expr", 9, 9); +return x_1; +} +} +static lean_object* _init_l_Lean_Loop_forIn_loop___at_Lean_Meta_Monotonicity_headBetaUnderLambda___spec__1___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("_private.Lean.Expr.0.Lean.Expr.updateLambda!Impl", 48, 48); +return x_1; +} +} +static lean_object* _init_l_Lean_Loop_forIn_loop___at_Lean_Meta_Monotonicity_headBetaUnderLambda___spec__1___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("lambda expected", 15, 15); +return x_1; +} +} +static lean_object* _init_l_Lean_Loop_forIn_loop___at_Lean_Meta_Monotonicity_headBetaUnderLambda___spec__1___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l_Lean_Loop_forIn_loop___at_Lean_Meta_Monotonicity_headBetaUnderLambda___spec__1___closed__1; +x_2 = l_Lean_Loop_forIn_loop___at_Lean_Meta_Monotonicity_headBetaUnderLambda___spec__1___closed__2; +x_3 = lean_unsigned_to_nat(1804u); +x_4 = lean_unsigned_to_nat(19u); +x_5 = l_Lean_Loop_forIn_loop___at_Lean_Meta_Monotonicity_headBetaUnderLambda___spec__1___closed__3; +x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at_Lean_Meta_Monotonicity_headBetaUnderLambda___spec__1(lean_object* x_1) { +_start: +{ +lean_object* x_2; uint8_t x_3; uint8_t x_4; +x_2 = l_Lean_Expr_bindingBody_x21(x_1); +x_3 = 0; +x_4 = l_Lean_Expr_isHeadBetaTarget(x_2, x_3); +if (x_4 == 0) +{ +lean_dec(x_2); +return x_1; +} +else +{ +if (lean_obj_tag(x_1) == 6) +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; lean_object* x_9; uint8_t x_10; lean_object* x_11; lean_object* x_12; size_t x_13; size_t x_14; uint8_t x_15; +x_5 = lean_ctor_get(x_1, 0); +lean_inc(x_5); +x_6 = lean_ctor_get(x_1, 1); +lean_inc(x_6); +x_7 = lean_ctor_get(x_1, 2); +lean_inc(x_7); +x_8 = lean_ctor_get_uint8(x_1, sizeof(void*)*3 + 8); +x_9 = l_Lean_Expr_bindingInfo_x21(x_1); +x_10 = lean_unbox(x_9); +lean_dec(x_9); +x_11 = l_Lean_Expr_bindingDomain_x21(x_1); +x_12 = l_Lean_Expr_headBeta(x_2); +x_13 = lean_ptr_addr(x_6); +lean_dec(x_6); +x_14 = lean_ptr_addr(x_11); +x_15 = lean_usize_dec_eq(x_13, x_14); +if (x_15 == 0) +{ +lean_object* x_16; +lean_dec(x_7); +lean_dec(x_1); +x_16 = l_Lean_Expr_lam___override(x_5, x_11, x_12, x_10); +x_1 = x_16; +goto _start; +} +else +{ +size_t x_18; size_t x_19; uint8_t x_20; +x_18 = lean_ptr_addr(x_7); +lean_dec(x_7); +x_19 = lean_ptr_addr(x_12); +x_20 = lean_usize_dec_eq(x_18, x_19); +if (x_20 == 0) +{ +lean_object* x_21; +lean_dec(x_1); +x_21 = l_Lean_Expr_lam___override(x_5, x_11, x_12, x_10); +x_1 = x_21; +goto _start; +} +else +{ +uint8_t x_23; +x_23 = l___private_Lean_Expr_0__Lean_beqBinderInfo____x40_Lean_Expr___hyg_406_(x_8, x_10); +if (x_23 == 0) +{ +lean_object* x_24; +lean_dec(x_1); +x_24 = l_Lean_Expr_lam___override(x_5, x_11, x_12, x_10); +x_1 = x_24; +goto _start; +} +else +{ +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_5); +goto _start; +} +} +} +} +else +{ +lean_object* x_27; lean_object* x_28; +lean_dec(x_2); +lean_dec(x_1); +x_27 = l_Lean_Loop_forIn_loop___at_Lean_Meta_Monotonicity_headBetaUnderLambda___spec__1___closed__4; +x_28 = l_panic___at_Lean_Expr_appFn_x21___spec__1(x_27); +x_1 = x_28; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_headBetaUnderLambda___lambda__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_inc(x_1); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_headBetaUnderLambda___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Monotonicity_headBetaUnderLambda___lambda__1___boxed), 2, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_headBetaUnderLambda(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; uint8_t x_4; +x_2 = l_Lean_Expr_headBeta(x_1); +x_3 = l_Lean_Meta_Monotonicity_headBetaUnderLambda___closed__1; +x_4 = l_Lean_Expr_isLambda(x_2); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; +x_5 = lean_box(0); +x_6 = lean_apply_2(x_3, x_2, x_5); +return x_6; +} +else +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_7 = l_Lean_Loop_forIn_loop___at_Lean_Meta_Monotonicity_headBetaUnderLambda___spec__1(x_2); +x_8 = lean_box(0); +x_9 = lean_apply_2(x_3, x_7, x_8); +return x_9; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_headBetaUnderLambda___lambda__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_Meta_Monotonicity_headBetaUnderLambda___lambda__1(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_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 = l___private_Lean_Meta_DiscrTreeTypes_0__Lean_Meta_DiscrTree_beqKey____x40_Lean_Meta_DiscrTreeTypes___hyg_102_(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_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__3___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_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__3___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_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__3___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_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__3(lean_object* x_1, size_t x_2, lean_object* x_3) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +uint8_t x_4; +x_4 = !lean_is_exclusive(x_1); +if (x_4 == 0) +{ +lean_object* x_5; size_t x_6; size_t x_7; size_t x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_5 = lean_ctor_get(x_1, 0); +x_6 = 5; +x_7 = l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__3___closed__2; +x_8 = lean_usize_land(x_2, x_7); +x_9 = lean_usize_to_nat(x_8); +x_10 = lean_box(2); +x_11 = lean_array_get(x_10, x_5, x_9); +lean_dec(x_9); +lean_dec(x_5); +switch (lean_obj_tag(x_11)) { +case 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 = l___private_Lean_Meta_DiscrTreeTypes_0__Lean_Meta_DiscrTree_beqKey____x40_Lean_Meta_DiscrTreeTypes___hyg_102_(x_3, x_12); +lean_dec(x_12); +if (x_14 == 0) +{ +lean_object* x_15; +lean_dec(x_13); +lean_free_object(x_1); +x_15 = lean_box(0); +return x_15; +} +else +{ +lean_ctor_set_tag(x_1, 1); +lean_ctor_set(x_1, 0, x_13); +return x_1; +} +} +case 1: +{ +lean_object* x_16; size_t x_17; +lean_free_object(x_1); +x_16 = lean_ctor_get(x_11, 0); +lean_inc(x_16); +lean_dec(x_11); +x_17 = lean_usize_shift_right(x_2, x_6); +x_1 = x_16; +x_2 = x_17; +goto _start; +} +default: +{ +lean_object* x_19; +lean_free_object(x_1); +x_19 = lean_box(0); +return x_19; +} +} +} +else +{ +lean_object* x_20; size_t x_21; size_t x_22; size_t x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_20 = lean_ctor_get(x_1, 0); +lean_inc(x_20); +lean_dec(x_1); +x_21 = 5; +x_22 = l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__3___closed__2; +x_23 = lean_usize_land(x_2, x_22); +x_24 = lean_usize_to_nat(x_23); +x_25 = lean_box(2); +x_26 = lean_array_get(x_25, x_20, x_24); +lean_dec(x_24); +lean_dec(x_20); +switch (lean_obj_tag(x_26)) { +case 0: +{ +lean_object* x_27; lean_object* x_28; uint8_t 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___private_Lean_Meta_DiscrTreeTypes_0__Lean_Meta_DiscrTree_beqKey____x40_Lean_Meta_DiscrTreeTypes___hyg_102_(x_3, x_27); +lean_dec(x_27); +if (x_29 == 0) +{ +lean_object* x_30; +lean_dec(x_28); +x_30 = lean_box(0); +return x_30; +} +else +{ +lean_object* x_31; +x_31 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_31, 0, x_28); +return x_31; +} +} +case 1: +{ +lean_object* x_32; size_t x_33; +x_32 = lean_ctor_get(x_26, 0); +lean_inc(x_32); +lean_dec(x_26); +x_33 = lean_usize_shift_right(x_2, x_21); +x_1 = x_32; +x_2 = x_33; +goto _start; +} +default: +{ +lean_object* x_35; +x_35 = lean_box(0); +return x_35; +} +} +} +} +else +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_36 = lean_ctor_get(x_1, 0); +lean_inc(x_36); +x_37 = lean_ctor_get(x_1, 1); +lean_inc(x_37); +lean_dec(x_1); +x_38 = lean_unsigned_to_nat(0u); +x_39 = l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__4(x_36, x_37, lean_box(0), x_38, x_3); +lean_dec(x_37); +lean_dec(x_36); +return x_39; +} +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__2(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint64_t x_3; size_t x_4; lean_object* x_5; +x_3 = l_Lean_Meta_DiscrTree_Key_hash(x_2); +x_4 = lean_uint64_to_usize(x_3); +x_5 = l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__3(x_1, x_4, x_2); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__7(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_Meta_DiscrTree_Key_hash(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_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__6(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_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____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; 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 = l___private_Lean_Meta_DiscrTreeTypes_0__Lean_Meta_DiscrTree_beqKey____x40_Lean_Meta_DiscrTreeTypes___hyg_102_(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_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__6___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_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__6(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_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__3___closed__2; +x_11 = lean_usize_land(x_2, x_10); +x_12 = lean_usize_to_nat(x_11); +x_13 = lean_array_get_size(x_7); +x_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 = l___private_Lean_Meta_DiscrTreeTypes_0__Lean_Meta_DiscrTree_beqKey____x40_Lean_Meta_DiscrTreeTypes___hyg_102_(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 = l___private_Lean_Meta_DiscrTreeTypes_0__Lean_Meta_DiscrTree_beqKey____x40_Lean_Meta_DiscrTreeTypes___hyg_102_(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_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__6(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_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__6(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_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__3___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 = l___private_Lean_Meta_DiscrTreeTypes_0__Lean_Meta_DiscrTree_beqKey____x40_Lean_Meta_DiscrTreeTypes___hyg_102_(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_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__6(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_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__8(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_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__6___closed__1; +x_93 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__7(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_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__8(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_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__6___closed__1; +x_107 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__7(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_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint64_t x_4; size_t x_5; size_t x_6; lean_object* x_7; +x_4 = l_Lean_Meta_DiscrTree_Key_hash(x_2); +x_5 = lean_uint64_to_usize(x_4); +x_6 = 1; +x_7 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__6(x_1, x_5, x_6, x_2, x_3); +return x_7; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertVal_loop___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__10(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; uint8_t x_5; +x_4 = lean_array_get_size(x_1); +x_5 = lean_nat_dec_lt(x_3, x_4); +lean_dec(x_4); +if (x_5 == 0) +{ +lean_object* x_6; +lean_dec(x_3); +x_6 = lean_array_push(x_1, x_2); +return x_6; +} +else +{ +lean_object* x_7; uint8_t x_8; +x_7 = lean_array_fget(x_1, x_3); +x_8 = lean_name_eq(x_2, x_7); +lean_dec(x_7); +if (x_8 == 0) +{ +lean_object* x_9; lean_object* x_10; +x_9 = lean_unsigned_to_nat(1u); +x_10 = lean_nat_add(x_3, x_9); +lean_dec(x_3); +x_3 = x_10; +goto _start; +} +else +{ +lean_object* x_12; +x_12 = lean_array_fset(x_1, x_3, x_2); +lean_dec(x_3); +return x_12; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_BinSearch_0__Array_binInsertAux___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__12(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_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_nat_add(x_7, x_8); +x_12 = lean_unsigned_to_nat(2u); +x_13 = lean_nat_div(x_11, x_12); +lean_dec(x_11); +x_14 = lean_array_fget(x_5, x_13); +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +x_16 = lean_ctor_get(x_6, 0); +x_17 = l_Lean_Meta_DiscrTree_Key_lt(x_15, x_16); +if (x_17 == 0) +{ +uint8_t x_18; +lean_dec(x_8); +x_18 = l_Lean_Meta_DiscrTree_Key_lt(x_16, x_15); +lean_dec(x_15); +if (x_18 == 0) +{ +lean_object* x_19; uint8_t x_20; +lean_dec(x_7); +x_19 = lean_array_get_size(x_5); +x_20 = lean_nat_dec_lt(x_13, x_19); +lean_dec(x_19); +if (x_20 == 0) +{ +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_4); +lean_dec(x_2); +return x_5; +} +else +{ +lean_object* x_21; lean_object* x_22; uint8_t x_23; +x_21 = lean_box(0); +x_22 = lean_array_fset(x_5, x_13, x_21); +x_23 = !lean_is_exclusive(x_14); +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; +x_24 = lean_ctor_get(x_14, 1); +x_25 = lean_ctor_get(x_14, 0); +lean_dec(x_25); +x_26 = lean_unsigned_to_nat(1u); +x_27 = lean_nat_add(x_3, x_26); +x_28 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__9(x_1, x_2, x_27, x_24); +lean_dec(x_27); +lean_ctor_set(x_14, 1, x_28); +lean_ctor_set(x_14, 0, x_4); +x_29 = lean_array_fset(x_22, x_13, x_14); +lean_dec(x_13); +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; +x_30 = lean_ctor_get(x_14, 1); +lean_inc(x_30); +lean_dec(x_14); +x_31 = lean_unsigned_to_nat(1u); +x_32 = lean_nat_add(x_3, x_31); +x_33 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__9(x_1, x_2, x_32, x_30); +lean_dec(x_32); +x_34 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_34, 0, x_4); +lean_ctor_set(x_34, 1, x_33); +x_35 = lean_array_fset(x_22, x_13, x_34); +lean_dec(x_13); +return x_35; +} +} +} +else +{ +lean_dec(x_14); +x_8 = x_13; +x_9 = lean_box(0); +x_10 = lean_box(0); +goto _start; +} +} +else +{ +uint8_t x_37; +lean_dec(x_15); +lean_dec(x_14); +x_37 = lean_nat_dec_eq(x_13, x_7); +if (x_37 == 0) +{ +lean_dec(x_7); +x_7 = x_13; +x_9 = lean_box(0); +x_10 = lean_box(0); +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_dec(x_13); +lean_dec(x_8); +x_39 = lean_unsigned_to_nat(1u); +x_40 = lean_nat_add(x_3, x_39); +x_41 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_createNodes___rarg(x_1, x_2, x_40); +lean_dec(x_40); +x_42 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_42, 0, x_4); +lean_ctor_set(x_42, 1, x_41); +x_43 = lean_nat_add(x_7, x_39); +lean_dec(x_7); +x_44 = lean_array_get_size(x_5); +x_45 = lean_array_push(x_5, x_42); +x_46 = l_Array_insertIdx_loop___rarg(x_43, x_45, x_44); +lean_dec(x_43); +return x_46; +} +} +} +} +LEAN_EXPORT lean_object* l_Array_binInsertM___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____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: +{ +lean_object* x_7; lean_object* x_8; uint8_t x_9; +x_7 = lean_array_get_size(x_5); +x_8 = lean_unsigned_to_nat(0u); +x_9 = lean_nat_dec_eq(x_7, x_8); +if (x_9 == 0) +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_10 = lean_array_fget(x_5, x_8); +x_11 = lean_ctor_get(x_6, 0); +x_12 = lean_ctor_get(x_10, 0); +lean_inc(x_12); +x_13 = l_Lean_Meta_DiscrTree_Key_lt(x_11, x_12); +if (x_13 == 0) +{ +uint8_t x_14; +x_14 = l_Lean_Meta_DiscrTree_Key_lt(x_12, x_11); +lean_dec(x_12); +if (x_14 == 0) +{ +uint8_t x_15; +x_15 = lean_nat_dec_lt(x_8, x_7); +lean_dec(x_7); +if (x_15 == 0) +{ +lean_dec(x_10); +lean_dec(x_4); +lean_dec(x_2); +return x_5; +} +else +{ +lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_16 = lean_box(0); +x_17 = lean_array_fset(x_5, x_8, x_16); +x_18 = !lean_is_exclusive(x_10); +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; +x_19 = lean_ctor_get(x_10, 1); +x_20 = lean_ctor_get(x_10, 0); +lean_dec(x_20); +x_21 = lean_unsigned_to_nat(1u); +x_22 = lean_nat_add(x_3, x_21); +x_23 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__9(x_1, x_2, x_22, x_19); +lean_dec(x_22); +lean_ctor_set(x_10, 1, x_23); +lean_ctor_set(x_10, 0, x_4); +x_24 = lean_array_fset(x_17, x_8, x_10); +return x_24; +} +else +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_25 = lean_ctor_get(x_10, 1); +lean_inc(x_25); +lean_dec(x_10); +x_26 = lean_unsigned_to_nat(1u); +x_27 = lean_nat_add(x_3, x_26); +x_28 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__9(x_1, x_2, x_27, x_25); +lean_dec(x_27); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_4); +lean_ctor_set(x_29, 1, x_28); +x_30 = lean_array_fset(x_17, x_8, x_29); +return x_30; +} +} +} +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; +lean_dec(x_10); +x_31 = lean_unsigned_to_nat(1u); +x_32 = lean_nat_sub(x_7, x_31); +x_33 = lean_array_fget(x_5, x_32); +x_34 = lean_ctor_get(x_33, 0); +lean_inc(x_34); +x_35 = l_Lean_Meta_DiscrTree_Key_lt(x_34, x_11); +if (x_35 == 0) +{ +uint8_t x_36; +x_36 = l_Lean_Meta_DiscrTree_Key_lt(x_11, x_34); +lean_dec(x_34); +if (x_36 == 0) +{ +uint8_t x_37; +x_37 = lean_nat_dec_lt(x_32, x_7); +lean_dec(x_7); +if (x_37 == 0) +{ +lean_dec(x_33); +lean_dec(x_32); +lean_dec(x_4); +lean_dec(x_2); +return x_5; +} +else +{ +lean_object* x_38; lean_object* x_39; uint8_t x_40; +x_38 = lean_box(0); +x_39 = lean_array_fset(x_5, x_32, x_38); +x_40 = !lean_is_exclusive(x_33); +if (x_40 == 0) +{ +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_33, 1); +x_42 = lean_ctor_get(x_33, 0); +lean_dec(x_42); +x_43 = lean_nat_add(x_3, x_31); +x_44 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__9(x_1, x_2, x_43, x_41); +lean_dec(x_43); +lean_ctor_set(x_33, 1, x_44); +lean_ctor_set(x_33, 0, x_4); +x_45 = lean_array_fset(x_39, x_32, x_33); +lean_dec(x_32); +return x_45; +} +else +{ +lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_46 = lean_ctor_get(x_33, 1); +lean_inc(x_46); +lean_dec(x_33); +x_47 = lean_nat_add(x_3, x_31); +x_48 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__9(x_1, x_2, x_47, x_46); +lean_dec(x_47); +x_49 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_49, 0, x_4); +lean_ctor_set(x_49, 1, x_48); +x_50 = lean_array_fset(x_39, x_32, x_49); +lean_dec(x_32); +return x_50; +} +} +} +else +{ +lean_object* x_51; +lean_dec(x_33); +lean_dec(x_7); +x_51 = l___private_Init_Data_Array_BinSearch_0__Array_binInsertAux___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__12(x_1, x_2, x_3, x_4, x_5, x_6, x_8, x_32, lean_box(0), lean_box(0)); +return x_51; +} +} +else +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; +lean_dec(x_34); +lean_dec(x_33); +lean_dec(x_32); +lean_dec(x_7); +x_52 = lean_nat_add(x_3, x_31); +x_53 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_createNodes___rarg(x_1, x_2, x_52); +lean_dec(x_52); +x_54 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_54, 0, x_4); +lean_ctor_set(x_54, 1, x_53); +x_55 = lean_array_push(x_5, x_54); +return x_55; +} +} +} +else +{ +lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; +lean_dec(x_12); +lean_dec(x_10); +x_56 = lean_unsigned_to_nat(1u); +x_57 = lean_nat_add(x_3, x_56); +x_58 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_createNodes___rarg(x_1, x_2, x_57); +lean_dec(x_57); +x_59 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_59, 0, x_4); +lean_ctor_set(x_59, 1, x_58); +x_60 = lean_array_push(x_5, x_59); +x_61 = l_Array_insertIdx_loop___rarg(x_8, x_60, x_7); +return x_61; +} +} +else +{ +lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; +lean_dec(x_7); +x_62 = lean_unsigned_to_nat(1u); +x_63 = lean_nat_add(x_3, x_62); +x_64 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_createNodes___rarg(x_1, x_2, x_63); +lean_dec(x_63); +x_65 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_65, 0, x_4); +lean_ctor_set(x_65, 1, x_64); +x_66 = lean_array_push(x_5, x_65); +return x_66; +} +} +} +static lean_object* _init_l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__9___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(0); +x_2 = lean_array_mk(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__9___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__9___closed__1; +x_2 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_2, 0, x_1); +lean_ctor_set(x_2, 1, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__9(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_4); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; +x_6 = lean_ctor_get(x_4, 0); +x_7 = lean_ctor_get(x_4, 1); +x_8 = lean_array_get_size(x_1); +x_9 = lean_nat_dec_lt(x_3, x_8); +lean_dec(x_8); +if (x_9 == 0) +{ +lean_object* x_10; lean_object* x_11; +x_10 = lean_unsigned_to_nat(0u); +x_11 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertVal_loop___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__10(x_6, x_2, x_10); +lean_ctor_set(x_4, 0, x_11); +return x_4; +} +else +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_12 = lean_array_fget(x_1, x_3); +x_13 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__9___closed__2; +lean_inc(x_12); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_12); +lean_ctor_set(x_14, 1, x_13); +x_15 = l_Array_binInsertM___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__11(x_1, x_2, x_3, x_12, x_7, x_14); +lean_dec(x_14); +lean_ctor_set(x_4, 1, x_15); +return x_4; +} +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; +x_16 = lean_ctor_get(x_4, 0); +x_17 = lean_ctor_get(x_4, 1); +lean_inc(x_17); +lean_inc(x_16); +lean_dec(x_4); +x_18 = lean_array_get_size(x_1); +x_19 = lean_nat_dec_lt(x_3, x_18); +lean_dec(x_18); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_unsigned_to_nat(0u); +x_21 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertVal_loop___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__10(x_16, x_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_17); +return x_22; +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_23 = lean_array_fget(x_1, x_3); +x_24 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__9___closed__2; +lean_inc(x_23); +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 = l_Array_binInsertM___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__11(x_1, x_2, x_3, x_23, x_17, x_25); +lean_dec(x_25); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_16); +lean_ctor_set(x_27, 1, x_26); +return x_27; +} +} +} +} +static lean_object* _init_l_panic___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__13___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Meta_DiscrTree_instInhabited(lean_box(0)); +return x_1; +} +} +LEAN_EXPORT lean_object* l_panic___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__13(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_panic___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__13___closed__1; +x_3 = lean_panic_fn(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_DiscrTree_insertCore___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.Meta.DiscrTree", 19, 19); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_DiscrTree_insertCore___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__1___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.Meta.DiscrTree.insertCore", 30, 30); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_DiscrTree_insertCore___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__1___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("invalid key sequence", 20, 20); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_DiscrTree_insertCore___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__1___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l_Lean_Meta_DiscrTree_insertCore___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__1___closed__1; +x_2 = l_Lean_Meta_DiscrTree_insertCore___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__1___closed__2; +x_3 = lean_unsigned_to_nat(481u); +x_4 = lean_unsigned_to_nat(23u); +x_5 = l_Lean_Meta_DiscrTree_insertCore___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__1___closed__3; +x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_DiscrTree_insertCore___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; +x_4 = l_Array_isEmpty___rarg(x_2); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; uint8_t x_7; +x_5 = lean_array_get_size(x_2); +x_6 = lean_unsigned_to_nat(0u); +x_7 = lean_nat_dec_lt(x_6, x_5); +lean_dec(x_5); +if (x_7 == 0) +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_8 = l_Lean_Meta_DiscrTree_instInhabitedKey; +x_9 = l_outOfBounds___rarg(x_8); +lean_inc(x_1); +x_10 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__2(x_1, x_9); +if (lean_obj_tag(x_10) == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_11 = lean_unsigned_to_nat(1u); +x_12 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_createNodes___rarg(x_2, x_3, x_11); +x_13 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__5(x_1, x_9, x_12); +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_10, 0); +lean_inc(x_14); +lean_dec(x_10); +x_15 = lean_unsigned_to_nat(1u); +x_16 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__9(x_2, x_3, x_15, x_14); +x_17 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__5(x_1, x_9, x_16); +return x_17; +} +} +else +{ +lean_object* x_18; lean_object* x_19; +x_18 = lean_array_fget(x_2, x_6); +lean_inc(x_1); +x_19 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__2(x_1, x_18); +if (lean_obj_tag(x_19) == 0) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_unsigned_to_nat(1u); +x_21 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_createNodes___rarg(x_2, x_3, x_20); +x_22 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__5(x_1, x_18, x_21); +return x_22; +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_23 = lean_ctor_get(x_19, 0); +lean_inc(x_23); +lean_dec(x_19); +x_24 = lean_unsigned_to_nat(1u); +x_25 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__9(x_2, x_3, x_24, x_23); +x_26 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__5(x_1, x_18, x_25); +return x_26; +} +} +} +else +{ +lean_object* x_27; lean_object* x_28; +lean_dec(x_3); +lean_dec(x_1); +x_27 = l_Lean_Meta_DiscrTree_insertCore___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__1___closed__4; +x_28 = l_panic___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__13(x_27); +return x_28; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____lambda__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = lean_ctor_get(x_2, 0); +lean_inc(x_3); +x_4 = lean_ctor_get(x_2, 1); +lean_inc(x_4); +lean_dec(x_2); +x_5 = l_Lean_Meta_DiscrTree_insertCore___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__1(x_1, x_4, x_3); +lean_dec(x_4); +return x_5; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean", 4, 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Meta", 4, 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Monotonicity", 12, 12); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("monotoneExt", 11, 11); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____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_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__1; +x_2 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__2; +x_3 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__3; +x_4 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____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_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__6() { +_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_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____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_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__8() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____lambda__1), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__9() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_id___rarg___boxed), 1, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____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_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__5; +x_2 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__8; +x_3 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__7; +x_4 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__9; +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; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166_(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__10; +x_3 = l_Lean_registerSimpleScopedEnvExtension___rarg(x_2, x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; +x_6 = l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__4(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_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__3___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_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__3(x_1, x_4, x_3); +lean_dec(x_3); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__2___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__2(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +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_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__7(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_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__6___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_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__6(x_1, x_6, x_7, x_4, x_5); +return x_8; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_BinSearch_0__Array_binInsertAux___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__12___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l___private_Init_Data_Array_BinSearch_0__Array_binInsertAux___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__12(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_6); +lean_dec(x_3); +lean_dec(x_1); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Array_binInsertM___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____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_Array_binInsertM___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__11(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_6); +lean_dec(x_3); +lean_dec(x_1); +return x_7; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__9___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__9(x_1, x_2, x_3, x_4); +lean_dec(x_3); +lean_dec(x_1); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_DiscrTree_insertCore___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Lean_Meta_DiscrTree_insertCore___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__1(x_1, x_2, x_3); +lean_dec(x_2); +return x_4; +} +} +static lean_object* _init_l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____spec__1___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____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; +} +} +static lean_object* _init_l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____spec__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__7; +x_2 = lean_alloc_ctor(0, 6, 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); +lean_ctor_set(x_2, 4, x_1); +lean_ctor_set(x_2, 5, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____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, 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_6, 6); +lean_inc(x_9); +lean_dec(x_6); +x_10 = lean_st_ref_take(x_7, 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_13 = !lean_is_exclusive(x_11); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; +x_14 = lean_ctor_get(x_11, 0); +x_15 = lean_ctor_get(x_11, 4); +lean_dec(x_15); +x_16 = l_Lean_ScopedEnvExtension_addCore___rarg(x_14, x_1, x_2, x_3, x_9); +x_17 = l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____spec__1___closed__1; +lean_ctor_set(x_11, 4, x_17); +lean_ctor_set(x_11, 0, x_16); +x_18 = lean_st_ref_set(x_7, x_11, x_12); +x_19 = lean_ctor_get(x_18, 1); +lean_inc(x_19); +lean_dec(x_18); +x_20 = lean_st_ref_take(x_5, x_19); +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; uint8_t x_27; +x_24 = lean_ctor_get(x_21, 1); +lean_dec(x_24); +x_25 = l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____spec__1___closed__2; +lean_ctor_set(x_21, 1, x_25); +x_26 = lean_st_ref_set(x_5, x_21, x_22); +x_27 = !lean_is_exclusive(x_26); +if (x_27 == 0) +{ +lean_object* x_28; lean_object* x_29; +x_28 = lean_ctor_get(x_26, 0); +lean_dec(x_28); +x_29 = lean_box(0); +lean_ctor_set(x_26, 0, x_29); +return x_26; +} +else +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_30 = lean_ctor_get(x_26, 1); +lean_inc(x_30); +lean_dec(x_26); +x_31 = lean_box(0); +x_32 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_30); +return x_32; +} +} +else +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_33 = lean_ctor_get(x_21, 0); +x_34 = lean_ctor_get(x_21, 2); +x_35 = lean_ctor_get(x_21, 3); +x_36 = lean_ctor_get(x_21, 4); +lean_inc(x_36); +lean_inc(x_35); +lean_inc(x_34); +lean_inc(x_33); +lean_dec(x_21); +x_37 = l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____spec__1___closed__2; +x_38 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_38, 0, x_33); +lean_ctor_set(x_38, 1, x_37); +lean_ctor_set(x_38, 2, x_34); +lean_ctor_set(x_38, 3, x_35); +lean_ctor_set(x_38, 4, x_36); +x_39 = lean_st_ref_set(x_5, x_38, x_22); +x_40 = lean_ctor_get(x_39, 1); +lean_inc(x_40); +if (lean_is_exclusive(x_39)) { + lean_ctor_release(x_39, 0); + lean_ctor_release(x_39, 1); + x_41 = x_39; +} else { + lean_dec_ref(x_39); + x_41 = lean_box(0); +} +x_42 = lean_box(0); +if (lean_is_scalar(x_41)) { + x_43 = lean_alloc_ctor(0, 2, 0); +} else { + x_43 = x_41; +} +lean_ctor_set(x_43, 0, x_42); +lean_ctor_set(x_43, 1, x_40); +return x_43; +} +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; 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_44 = lean_ctor_get(x_11, 0); +x_45 = lean_ctor_get(x_11, 1); +x_46 = lean_ctor_get(x_11, 2); +x_47 = lean_ctor_get(x_11, 3); +x_48 = lean_ctor_get(x_11, 5); +x_49 = lean_ctor_get(x_11, 6); +x_50 = lean_ctor_get(x_11, 7); +lean_inc(x_50); +lean_inc(x_49); +lean_inc(x_48); +lean_inc(x_47); +lean_inc(x_46); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_11); +x_51 = l_Lean_ScopedEnvExtension_addCore___rarg(x_44, x_1, x_2, x_3, x_9); +x_52 = l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____spec__1___closed__1; +x_53 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_53, 0, x_51); +lean_ctor_set(x_53, 1, x_45); +lean_ctor_set(x_53, 2, x_46); +lean_ctor_set(x_53, 3, x_47); +lean_ctor_set(x_53, 4, x_52); +lean_ctor_set(x_53, 5, x_48); +lean_ctor_set(x_53, 6, x_49); +lean_ctor_set(x_53, 7, x_50); +x_54 = lean_st_ref_set(x_7, x_53, x_12); +x_55 = lean_ctor_get(x_54, 1); +lean_inc(x_55); +lean_dec(x_54); +x_56 = lean_st_ref_take(x_5, x_55); +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_inc(x_59); +x_60 = lean_ctor_get(x_57, 2); +lean_inc(x_60); +x_61 = lean_ctor_get(x_57, 3); +lean_inc(x_61); +x_62 = lean_ctor_get(x_57, 4); +lean_inc(x_62); +if (lean_is_exclusive(x_57)) { + lean_ctor_release(x_57, 0); + lean_ctor_release(x_57, 1); + lean_ctor_release(x_57, 2); + lean_ctor_release(x_57, 3); + lean_ctor_release(x_57, 4); + x_63 = x_57; +} else { + lean_dec_ref(x_57); + x_63 = lean_box(0); +} +x_64 = l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____spec__1___closed__2; +if (lean_is_scalar(x_63)) { + x_65 = lean_alloc_ctor(0, 5, 0); +} else { + x_65 = x_63; +} +lean_ctor_set(x_65, 0, x_59); +lean_ctor_set(x_65, 1, x_64); +lean_ctor_set(x_65, 2, x_60); +lean_ctor_set(x_65, 3, x_61); +lean_ctor_set(x_65, 4, x_62); +x_66 = lean_st_ref_set(x_5, x_65, x_58); +x_67 = lean_ctor_get(x_66, 1); +lean_inc(x_67); +if (lean_is_exclusive(x_66)) { + lean_ctor_release(x_66, 0); + lean_ctor_release(x_66, 1); + x_68 = x_66; +} else { + lean_dec_ref(x_66); + x_68 = lean_box(0); +} +x_69 = lean_box(0); +if (lean_is_scalar(x_68)) { + x_70 = lean_alloc_ctor(0, 2, 0); +} else { + x_70 = x_68; +} +lean_ctor_set(x_70, 0, x_69); +lean_ctor_set(x_70, 1, x_67); +return x_70; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_lambdaBoundedTelescope___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____spec__2___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; uint8_t x_11; lean_object* x_12; +x_10 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_10, 0, x_2); +x_11 = 0; +x_12 = l___private_Lean_Meta_Basic_0__Lean_Meta_lambdaTelescopeImp___rarg(x_1, x_11, x_10, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_10); +if (lean_obj_tag(x_12) == 0) +{ +uint8_t x_13; +x_13 = !lean_is_exclusive(x_12); +if (x_13 == 0) +{ +return x_12; +} +else +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_14 = lean_ctor_get(x_12, 0); +x_15 = lean_ctor_get(x_12, 1); +lean_inc(x_15); +lean_inc(x_14); +lean_dec(x_12); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_14); +lean_ctor_set(x_16, 1, x_15); +return x_16; +} +} +else +{ +uint8_t x_17; +x_17 = !lean_is_exclusive(x_12); +if (x_17 == 0) +{ +return x_12; +} +else +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_18 = lean_ctor_get(x_12, 0); +x_19 = lean_ctor_get(x_12, 1); +lean_inc(x_19); +lean_inc(x_18); +lean_dec(x_12); +x_20 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_20, 0, x_18); +lean_ctor_set(x_20, 1, x_19); +return x_20; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_lambdaBoundedTelescope___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____spec__2(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_Meta_lambdaBoundedTelescope___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____spec__2___rarg___boxed), 9, 0); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Meta_Monotonicity_monotoneExt; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__1(uint64_t x_1, uint64_t x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; uint64_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; lean_object* x_20; uint8_t x_21; +x_12 = lean_ctor_get(x_7, 0); +lean_inc(x_12); +x_13 = lean_ctor_get_uint64(x_7, sizeof(void*)*7); +x_14 = lean_ctor_get_uint8(x_7, sizeof(void*)*7 + 8); +x_15 = lean_ctor_get(x_7, 1); +lean_inc(x_15); +x_16 = lean_ctor_get(x_7, 2); +lean_inc(x_16); +x_17 = lean_ctor_get(x_7, 3); +lean_inc(x_17); +x_18 = lean_ctor_get(x_7, 4); +lean_inc(x_18); +x_19 = lean_ctor_get(x_7, 5); +lean_inc(x_19); +x_20 = lean_ctor_get(x_7, 6); +lean_inc(x_20); +x_21 = !lean_is_exclusive(x_12); +if (x_21 == 0) +{ +uint8_t x_22; uint8_t x_23; uint8_t x_24; uint64_t x_25; uint64_t x_26; uint64_t x_27; lean_object* x_28; uint8_t x_29; lean_object* x_30; +x_22 = lean_ctor_get_uint8(x_7, sizeof(void*)*7 + 9); +x_23 = lean_ctor_get_uint8(x_7, sizeof(void*)*7 + 10); +x_24 = 2; +lean_ctor_set_uint8(x_12, 9, x_24); +x_25 = lean_uint64_shift_right(x_13, x_1); +x_26 = lean_uint64_shift_left(x_25, x_1); +x_27 = lean_uint64_lor(x_26, x_2); +x_28 = lean_alloc_ctor(0, 7, 11); +lean_ctor_set(x_28, 0, x_12); +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_19); +lean_ctor_set(x_28, 6, x_20); +lean_ctor_set_uint64(x_28, sizeof(void*)*7, x_27); +lean_ctor_set_uint8(x_28, sizeof(void*)*7 + 8, x_14); +lean_ctor_set_uint8(x_28, sizeof(void*)*7 + 9, x_22); +lean_ctor_set_uint8(x_28, sizeof(void*)*7 + 10, x_23); +x_29 = 0; +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +x_30 = l_Lean_Meta_DiscrTree_mkPath(x_6, x_29, x_28, x_8, x_9, x_10, x_11); +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; +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_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_3); +lean_ctor_set(x_33, 1, x_31); +x_34 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__1___closed__1; +x_35 = l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____spec__1(x_34, x_33, x_4, x_7, x_8, x_9, x_10, x_32); +lean_dec(x_10); +lean_dec(x_8); +lean_dec(x_7); +return x_35; +} +else +{ +uint8_t x_36; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_3); +x_36 = !lean_is_exclusive(x_30); +if (x_36 == 0) +{ +return x_30; +} +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_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; +} +} +} +else +{ +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; 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; lean_object* x_59; uint64_t x_60; uint64_t x_61; uint64_t x_62; lean_object* x_63; uint8_t x_64; lean_object* x_65; +x_40 = lean_ctor_get_uint8(x_7, sizeof(void*)*7 + 9); +x_41 = lean_ctor_get_uint8(x_7, sizeof(void*)*7 + 10); +x_42 = lean_ctor_get_uint8(x_12, 0); +x_43 = lean_ctor_get_uint8(x_12, 1); +x_44 = lean_ctor_get_uint8(x_12, 2); +x_45 = lean_ctor_get_uint8(x_12, 3); +x_46 = lean_ctor_get_uint8(x_12, 4); +x_47 = lean_ctor_get_uint8(x_12, 5); +x_48 = lean_ctor_get_uint8(x_12, 6); +x_49 = lean_ctor_get_uint8(x_12, 7); +x_50 = lean_ctor_get_uint8(x_12, 8); +x_51 = lean_ctor_get_uint8(x_12, 10); +x_52 = lean_ctor_get_uint8(x_12, 11); +x_53 = lean_ctor_get_uint8(x_12, 12); +x_54 = lean_ctor_get_uint8(x_12, 13); +x_55 = lean_ctor_get_uint8(x_12, 14); +x_56 = lean_ctor_get_uint8(x_12, 15); +x_57 = lean_ctor_get_uint8(x_12, 16); +lean_dec(x_12); +x_58 = 2; +x_59 = lean_alloc_ctor(0, 0, 17); +lean_ctor_set_uint8(x_59, 0, x_42); +lean_ctor_set_uint8(x_59, 1, x_43); +lean_ctor_set_uint8(x_59, 2, x_44); +lean_ctor_set_uint8(x_59, 3, x_45); +lean_ctor_set_uint8(x_59, 4, x_46); +lean_ctor_set_uint8(x_59, 5, x_47); +lean_ctor_set_uint8(x_59, 6, x_48); +lean_ctor_set_uint8(x_59, 7, x_49); +lean_ctor_set_uint8(x_59, 8, x_50); +lean_ctor_set_uint8(x_59, 9, x_58); +lean_ctor_set_uint8(x_59, 10, x_51); +lean_ctor_set_uint8(x_59, 11, x_52); +lean_ctor_set_uint8(x_59, 12, x_53); +lean_ctor_set_uint8(x_59, 13, x_54); +lean_ctor_set_uint8(x_59, 14, x_55); +lean_ctor_set_uint8(x_59, 15, x_56); +lean_ctor_set_uint8(x_59, 16, x_57); +x_60 = lean_uint64_shift_right(x_13, x_1); +x_61 = lean_uint64_shift_left(x_60, x_1); +x_62 = lean_uint64_lor(x_61, x_2); +x_63 = lean_alloc_ctor(0, 7, 11); +lean_ctor_set(x_63, 0, x_59); +lean_ctor_set(x_63, 1, x_15); +lean_ctor_set(x_63, 2, x_16); +lean_ctor_set(x_63, 3, x_17); +lean_ctor_set(x_63, 4, x_18); +lean_ctor_set(x_63, 5, x_19); +lean_ctor_set(x_63, 6, x_20); +lean_ctor_set_uint64(x_63, sizeof(void*)*7, x_62); +lean_ctor_set_uint8(x_63, sizeof(void*)*7 + 8, x_14); +lean_ctor_set_uint8(x_63, sizeof(void*)*7 + 9, x_40); +lean_ctor_set_uint8(x_63, sizeof(void*)*7 + 10, x_41); +x_64 = 0; +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +x_65 = l_Lean_Meta_DiscrTree_mkPath(x_6, x_64, x_63, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_65) == 0) +{ +lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; +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(0, 2, 0); +lean_ctor_set(x_68, 0, x_3); +lean_ctor_set(x_68, 1, x_66); +x_69 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__1___closed__1; +x_70 = l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____spec__1(x_69, x_68, x_4, x_7, x_8, x_9, x_10, x_67); +lean_dec(x_10); +lean_dec(x_8); +lean_dec(x_7); +return x_70; +} +else +{ +lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_3); +x_71 = lean_ctor_get(x_65, 0); +lean_inc(x_71); +x_72 = lean_ctor_get(x_65, 1); +lean_inc(x_72); +if (lean_is_exclusive(x_65)) { + lean_ctor_release(x_65, 0); + lean_ctor_release(x_65, 1); + x_73 = x_65; +} else { + lean_dec_ref(x_65); + 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(x_74, 0, x_71); +lean_ctor_set(x_74, 1, x_72); +return x_74; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__2(uint64_t x_1, uint64_t x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; lean_object* x_18; +x_11 = l_Lean_Meta_Monotonicity_headBetaUnderLambda(x_5); +x_12 = lean_box_uint64(x_1); +x_13 = lean_box_uint64(x_2); +x_14 = lean_box(x_4); +x_15 = lean_alloc_closure((void*)(l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__1___boxed), 11, 4); +lean_closure_set(x_15, 0, x_12); +lean_closure_set(x_15, 1, x_13); +lean_closure_set(x_15, 2, x_3); +lean_closure_set(x_15, 3, x_14); +x_16 = lean_unsigned_to_nat(1u); +x_17 = 0; +x_18 = l_Lean_Meta_lambdaBoundedTelescope___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____spec__2___rarg(x_11, x_16, x_15, x_17, x_6, x_7, x_8, x_9, x_10); +return x_18; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__3(uint64_t x_1, uint64_t x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +lean_object* x_15; uint8_t x_16; +x_15 = l_Lean_Expr_headBeta(x_9); +x_16 = l_Lean_Expr_isLambda(x_15); +if (x_16 == 0) +{ +lean_object* x_17; +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +x_17 = l_Lean_Meta_etaExpand(x_15, x_10, x_11, x_12, x_13, x_14); +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_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_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__2(x_1, x_2, x_3, x_4, x_18, x_10, x_11, x_12, x_13, x_19); +return x_20; +} +else +{ +uint8_t x_21; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_3); +x_21 = !lean_is_exclusive(x_17); +if (x_21 == 0) +{ +return x_17; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_22 = lean_ctor_get(x_17, 0); +x_23 = lean_ctor_get(x_17, 1); +lean_inc(x_23); +lean_inc(x_22); +lean_dec(x_17); +x_24 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_24, 0, x_22); +lean_ctor_set(x_24, 1, x_23); +return x_24; +} +} +} +else +{ +lean_object* x_25; +x_25 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__2(x_1, x_2, x_3, x_4, x_15, x_10, x_11, x_12, x_13, x_14); +return x_25; +} +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("@[partial_fixpoint_monotone] attribute only applies to lemmas proving ", 70, 70); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Order", 5, 5); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("monotone", 8, 8); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__1; +x_2 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__3; +x_3 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__4; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__6() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__5; +x_2 = 0; +x_3 = l_Lean_MessageData_ofConstName(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__2; +x_2 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___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_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__8() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("", 0, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__8; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__7; +x_2 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__9; +x_3 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; lean_object* x_8; +x_7 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__10; +x_8 = l_Lean_throwError___at_Lean_Meta_setInlineAttribute___spec__1(x_7, x_2, x_3, x_4, x_5, x_6); +return x_8; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__1() { +_start: +{ +uint8_t x_1; uint8_t x_2; uint8_t x_3; uint8_t x_4; uint8_t x_5; lean_object* x_6; +x_1 = 0; +x_2 = 1; +x_3 = 1; +x_4 = 0; +x_5 = 2; +x_6 = lean_alloc_ctor(0, 0, 17); +lean_ctor_set_uint8(x_6, 0, x_1); +lean_ctor_set_uint8(x_6, 1, x_1); +lean_ctor_set_uint8(x_6, 2, x_1); +lean_ctor_set_uint8(x_6, 3, x_1); +lean_ctor_set_uint8(x_6, 4, x_1); +lean_ctor_set_uint8(x_6, 5, x_2); +lean_ctor_set_uint8(x_6, 6, x_2); +lean_ctor_set_uint8(x_6, 7, x_1); +lean_ctor_set_uint8(x_6, 8, x_2); +lean_ctor_set_uint8(x_6, 9, x_3); +lean_ctor_set_uint8(x_6, 10, x_4); +lean_ctor_set_uint8(x_6, 11, x_2); +lean_ctor_set_uint8(x_6, 12, x_2); +lean_ctor_set_uint8(x_6, 13, x_2); +lean_ctor_set_uint8(x_6, 14, x_5); +lean_ctor_set_uint8(x_6, 15, x_2); +lean_ctor_set_uint8(x_6, 16, x_2); +return x_6; +} +} +static uint64_t _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__2() { +_start: +{ +lean_object* x_1; uint64_t x_2; +x_1 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__1; +x_2 = l___private_Lean_Meta_Basic_0__Lean_Meta_Config_toKey(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(32u); +x_2 = lean_mk_empty_array_with_capacity(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__3; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__5() { +_start: +{ +size_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = 5; +x_2 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__4; +x_3 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__3; +x_4 = lean_unsigned_to_nat(0u); +x_5 = lean_alloc_ctor(0, 4, sizeof(size_t)*1); +lean_ctor_set(x_5, 0, x_2); +lean_ctor_set(x_5, 1, x_3); +lean_ctor_set(x_5, 2, x_4); +lean_ctor_set(x_5, 3, x_4); +lean_ctor_set_usize(x_5, 4, x_1); +return x_5; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__7; +x_2 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__5; +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; uint64_t x_4; uint8_t x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_1 = lean_box(0); +x_2 = lean_box(0); +x_3 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__1; +x_4 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__2; +x_5 = 0; +x_6 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__6; +x_7 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__9___closed__1; +x_8 = lean_unsigned_to_nat(0u); +x_9 = lean_alloc_ctor(0, 7, 11); +lean_ctor_set(x_9, 0, x_3); +lean_ctor_set(x_9, 1, x_1); +lean_ctor_set(x_9, 2, x_6); +lean_ctor_set(x_9, 3, x_7); +lean_ctor_set(x_9, 4, x_2); +lean_ctor_set(x_9, 5, x_8); +lean_ctor_set(x_9, 6, x_2); +lean_ctor_set_uint64(x_9, sizeof(void*)*7, x_4); +lean_ctor_set_uint8(x_9, sizeof(void*)*7 + 8, x_5); +lean_ctor_set_uint8(x_9, sizeof(void*)*7 + 9, x_5); +lean_ctor_set_uint8(x_9, sizeof(void*)*7 + 10, x_5); +return x_9; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_unsigned_to_nat(0u); +x_2 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__7; +x_3 = lean_alloc_ctor(0, 9, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_1); +lean_ctor_set(x_3, 2, x_1); +lean_ctor_set(x_3, 3, x_2); +lean_ctor_set(x_3, 4, x_2); +lean_ctor_set(x_3, 5, x_2); +lean_ctor_set(x_3, 6, x_2); +lean_ctor_set(x_3, 7, x_2); +lean_ctor_set(x_3, 8, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__7; +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_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___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 = lean_box(0); +x_2 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__8; +x_3 = l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____spec__1___closed__2; +x_4 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__5; +x_5 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__9; +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; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__11() { +_start: +{ +uint8_t x_1; uint8_t x_2; uint8_t x_3; uint8_t x_4; uint8_t x_5; lean_object* x_6; +x_1 = 0; +x_2 = 1; +x_3 = 2; +x_4 = 0; +x_5 = 2; +x_6 = lean_alloc_ctor(0, 0, 17); +lean_ctor_set_uint8(x_6, 0, x_1); +lean_ctor_set_uint8(x_6, 1, x_1); +lean_ctor_set_uint8(x_6, 2, x_1); +lean_ctor_set_uint8(x_6, 3, x_1); +lean_ctor_set_uint8(x_6, 4, x_1); +lean_ctor_set_uint8(x_6, 5, x_2); +lean_ctor_set_uint8(x_6, 6, x_2); +lean_ctor_set_uint8(x_6, 7, x_1); +lean_ctor_set_uint8(x_6, 8, x_2); +lean_ctor_set_uint8(x_6, 9, x_3); +lean_ctor_set_uint8(x_6, 10, x_4); +lean_ctor_set_uint8(x_6, 11, x_2); +lean_ctor_set_uint8(x_6, 12, x_2); +lean_ctor_set_uint8(x_6, 13, x_2); +lean_ctor_set_uint8(x_6, 14, x_5); +lean_ctor_set_uint8(x_6, 15, x_2); +lean_ctor_set_uint8(x_6, 16, x_2); +return x_6; +} +} +static uint64_t _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__12() { +_start: +{ +uint64_t x_1; uint64_t x_2; uint64_t x_3; +x_1 = 2; +x_2 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__2; +x_3 = lean_uint64_shift_right(x_2, x_1); +return x_3; +} +} +static uint64_t _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__13() { +_start: +{ +uint64_t x_1; uint64_t x_2; uint64_t x_3; +x_1 = 2; +x_2 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__12; +x_3 = lean_uint64_shift_left(x_2, x_1); +return x_3; +} +} +static uint64_t _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__14() { +_start: +{ +uint8_t x_1; uint64_t x_2; +x_1 = 2; +x_2 = l_Lean_Meta_TransparencyMode_toUInt64(x_1); +return x_2; +} +} +static uint64_t _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__15() { +_start: +{ +uint64_t x_1; uint64_t x_2; uint64_t x_3; +x_1 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__13; +x_2 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__14; +x_3 = lean_uint64_lor(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__16() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; uint64_t x_4; uint8_t x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_1 = lean_box(0); +x_2 = lean_box(0); +x_3 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__11; +x_4 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__15; +x_5 = 0; +x_6 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__6; +x_7 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__9___closed__1; +x_8 = lean_unsigned_to_nat(0u); +x_9 = lean_alloc_ctor(0, 7, 11); +lean_ctor_set(x_9, 0, x_3); +lean_ctor_set(x_9, 1, x_1); +lean_ctor_set(x_9, 2, x_6); +lean_ctor_set(x_9, 3, x_7); +lean_ctor_set(x_9, 4, x_2); +lean_ctor_set(x_9, 5, x_8); +lean_ctor_set(x_9, 6, x_2); +lean_ctor_set_uint64(x_9, sizeof(void*)*7, x_4); +lean_ctor_set_uint8(x_9, sizeof(void*)*7 + 8, x_5); +lean_ctor_set_uint8(x_9, sizeof(void*)*7 + 9, x_5); +lean_ctor_set_uint8(x_9, sizeof(void*)*7 + 10, x_5); +return x_9; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__17() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___boxed), 6, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____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) { +_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; +x_7 = lean_box(0); +x_8 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__10; +x_9 = lean_st_mk_ref(x_8, x_6); +x_10 = lean_ctor_get(x_9, 0); +lean_inc(x_10); +x_11 = lean_ctor_get(x_9, 1); +lean_inc(x_11); +lean_dec(x_9); +x_12 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__7; +lean_inc(x_1); +x_13 = l_Lean_getConstInfo___at_Lean_Meta_mkConstWithFreshMVarLevels___spec__1(x_1, x_12, x_10, x_4, x_5, x_11); +if (lean_obj_tag(x_13) == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint64_t x_17; uint8_t x_18; uint8_t x_19; lean_object* x_20; lean_object* x_21; +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +x_16 = l_Lean_ConstantInfo_type(x_14); +lean_dec(x_14); +x_17 = 2; +x_18 = 1; +x_19 = 0; +x_20 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__16; +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_10); +x_21 = l___private_Lean_Meta_Basic_0__Lean_Meta_forallMetaTelescopeReducingAux(x_16, x_18, x_7, x_19, x_20, x_10, x_4, x_5, x_15); +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; uint8_t x_28; +x_22 = lean_ctor_get(x_21, 0); +lean_inc(x_22); +x_23 = lean_ctor_get(x_22, 1); +lean_inc(x_23); +lean_dec(x_22); +x_24 = lean_ctor_get(x_21, 1); +lean_inc(x_24); +lean_dec(x_21); +x_25 = lean_ctor_get(x_23, 1); +lean_inc(x_25); +lean_dec(x_23); +x_26 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__17; +x_27 = l_Lean_Expr_cleanupAnnotations(x_25); +x_28 = l_Lean_Expr_isApp(x_27); +if (x_28 == 0) +{ +lean_object* x_29; lean_object* x_30; +lean_dec(x_27); +lean_dec(x_1); +x_29 = lean_box(0); +lean_inc(x_10); +x_30 = lean_apply_6(x_26, x_29, x_12, x_10, x_4, x_5, x_24); +if (lean_obj_tag(x_30) == 0) +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; uint8_t x_34; +x_31 = lean_ctor_get(x_30, 0); +lean_inc(x_31); +x_32 = lean_ctor_get(x_30, 1); +lean_inc(x_32); +lean_dec(x_30); +x_33 = lean_st_ref_get(x_10, x_32); +lean_dec(x_10); +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_31); +return x_33; +} +else +{ +lean_object* x_36; lean_object* x_37; +x_36 = lean_ctor_get(x_33, 1); +lean_inc(x_36); +lean_dec(x_33); +x_37 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_37, 0, x_31); +lean_ctor_set(x_37, 1, x_36); +return x_37; +} +} +else +{ +uint8_t x_38; +lean_dec(x_10); +x_38 = !lean_is_exclusive(x_30); +if (x_38 == 0) +{ +return x_30; +} +else +{ +lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_39 = lean_ctor_get(x_30, 0); +x_40 = lean_ctor_get(x_30, 1); +lean_inc(x_40); +lean_inc(x_39); +lean_dec(x_30); +x_41 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_41, 0, x_39); +lean_ctor_set(x_41, 1, x_40); +return x_41; +} +} +} +else +{ +lean_object* x_42; lean_object* x_43; uint8_t x_44; +x_42 = l_Lean_Expr_appArg(x_27, lean_box(0)); +x_43 = l_Lean_Expr_appFnCleanup(x_27, lean_box(0)); +x_44 = l_Lean_Expr_isApp(x_43); +if (x_44 == 0) +{ +lean_object* x_45; lean_object* x_46; +lean_dec(x_43); +lean_dec(x_42); +lean_dec(x_1); +x_45 = lean_box(0); +lean_inc(x_10); +x_46 = lean_apply_6(x_26, x_45, x_12, x_10, x_4, x_5, x_24); +if (lean_obj_tag(x_46) == 0) +{ +lean_object* x_47; lean_object* x_48; lean_object* x_49; uint8_t x_50; +x_47 = lean_ctor_get(x_46, 0); +lean_inc(x_47); +x_48 = lean_ctor_get(x_46, 1); +lean_inc(x_48); +lean_dec(x_46); +x_49 = lean_st_ref_get(x_10, x_48); +lean_dec(x_10); +x_50 = !lean_is_exclusive(x_49); +if (x_50 == 0) +{ +lean_object* x_51; +x_51 = lean_ctor_get(x_49, 0); +lean_dec(x_51); +lean_ctor_set(x_49, 0, x_47); +return x_49; +} +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_47); +lean_ctor_set(x_53, 1, x_52); +return x_53; +} +} +else +{ +uint8_t x_54; +lean_dec(x_10); +x_54 = !lean_is_exclusive(x_46); +if (x_54 == 0) +{ +return x_46; +} +else +{ +lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_55 = lean_ctor_get(x_46, 0); +x_56 = lean_ctor_get(x_46, 1); +lean_inc(x_56); +lean_inc(x_55); +lean_dec(x_46); +x_57 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_57, 0, x_55); +lean_ctor_set(x_57, 1, x_56); +return x_57; +} +} +} +else +{ +lean_object* x_58; lean_object* x_59; uint8_t x_60; +x_58 = l_Lean_Expr_appArg(x_43, lean_box(0)); +x_59 = l_Lean_Expr_appFnCleanup(x_43, lean_box(0)); +x_60 = l_Lean_Expr_isApp(x_59); +if (x_60 == 0) +{ +lean_object* x_61; lean_object* x_62; +lean_dec(x_59); +lean_dec(x_58); +lean_dec(x_42); +lean_dec(x_1); +x_61 = lean_box(0); +lean_inc(x_10); +x_62 = lean_apply_6(x_26, x_61, x_12, x_10, x_4, x_5, x_24); +if (lean_obj_tag(x_62) == 0) +{ +lean_object* x_63; lean_object* x_64; lean_object* x_65; uint8_t 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); +lean_dec(x_62); +x_65 = lean_st_ref_get(x_10, x_64); +lean_dec(x_10); +x_66 = !lean_is_exclusive(x_65); +if (x_66 == 0) +{ +lean_object* x_67; +x_67 = lean_ctor_get(x_65, 0); +lean_dec(x_67); +lean_ctor_set(x_65, 0, x_63); +return x_65; +} +else +{ +lean_object* x_68; lean_object* x_69; +x_68 = lean_ctor_get(x_65, 1); +lean_inc(x_68); +lean_dec(x_65); +x_69 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_69, 0, x_63); +lean_ctor_set(x_69, 1, x_68); +return x_69; +} +} +else +{ +uint8_t x_70; +lean_dec(x_10); +x_70 = !lean_is_exclusive(x_62); +if (x_70 == 0) +{ +return x_62; +} +else +{ +lean_object* x_71; lean_object* x_72; lean_object* x_73; +x_71 = lean_ctor_get(x_62, 0); +x_72 = lean_ctor_get(x_62, 1); +lean_inc(x_72); +lean_inc(x_71); +lean_dec(x_62); +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_74; lean_object* x_75; uint8_t x_76; +x_74 = l_Lean_Expr_appArg(x_59, lean_box(0)); +x_75 = l_Lean_Expr_appFnCleanup(x_59, lean_box(0)); +x_76 = l_Lean_Expr_isApp(x_75); +if (x_76 == 0) +{ +lean_object* x_77; lean_object* x_78; +lean_dec(x_75); +lean_dec(x_74); +lean_dec(x_58); +lean_dec(x_42); +lean_dec(x_1); +x_77 = lean_box(0); +lean_inc(x_10); +x_78 = lean_apply_6(x_26, x_77, x_12, x_10, x_4, x_5, x_24); +if (lean_obj_tag(x_78) == 0) +{ +lean_object* x_79; lean_object* x_80; lean_object* x_81; uint8_t x_82; +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_st_ref_get(x_10, x_80); +lean_dec(x_10); +x_82 = !lean_is_exclusive(x_81); +if (x_82 == 0) +{ +lean_object* x_83; +x_83 = lean_ctor_get(x_81, 0); +lean_dec(x_83); +lean_ctor_set(x_81, 0, x_79); +return x_81; +} +else +{ +lean_object* x_84; lean_object* x_85; +x_84 = lean_ctor_get(x_81, 1); +lean_inc(x_84); +lean_dec(x_81); +x_85 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_85, 0, x_79); +lean_ctor_set(x_85, 1, x_84); +return x_85; +} +} +else +{ +uint8_t x_86; +lean_dec(x_10); +x_86 = !lean_is_exclusive(x_78); +if (x_86 == 0) +{ +return x_78; +} +else +{ +lean_object* x_87; lean_object* x_88; lean_object* x_89; +x_87 = lean_ctor_get(x_78, 0); +x_88 = lean_ctor_get(x_78, 1); +lean_inc(x_88); +lean_inc(x_87); +lean_dec(x_78); +x_89 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_89, 0, x_87); +lean_ctor_set(x_89, 1, x_88); +return x_89; +} +} +} +else +{ +lean_object* x_90; lean_object* x_91; uint8_t x_92; +x_90 = l_Lean_Expr_appArg(x_75, lean_box(0)); +x_91 = l_Lean_Expr_appFnCleanup(x_75, lean_box(0)); +x_92 = l_Lean_Expr_isApp(x_91); +if (x_92 == 0) +{ +lean_object* x_93; lean_object* x_94; +lean_dec(x_91); +lean_dec(x_90); +lean_dec(x_74); +lean_dec(x_58); +lean_dec(x_42); +lean_dec(x_1); +x_93 = lean_box(0); +lean_inc(x_10); +x_94 = lean_apply_6(x_26, x_93, x_12, x_10, x_4, x_5, x_24); +if (lean_obj_tag(x_94) == 0) +{ +lean_object* x_95; lean_object* x_96; lean_object* x_97; uint8_t x_98; +x_95 = lean_ctor_get(x_94, 0); +lean_inc(x_95); +x_96 = lean_ctor_get(x_94, 1); +lean_inc(x_96); +lean_dec(x_94); +x_97 = lean_st_ref_get(x_10, x_96); +lean_dec(x_10); +x_98 = !lean_is_exclusive(x_97); +if (x_98 == 0) +{ +lean_object* x_99; +x_99 = lean_ctor_get(x_97, 0); +lean_dec(x_99); +lean_ctor_set(x_97, 0, x_95); +return x_97; +} +else +{ +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_alloc_ctor(0, 2, 0); +lean_ctor_set(x_101, 0, x_95); +lean_ctor_set(x_101, 1, x_100); +return x_101; +} +} +else +{ +uint8_t x_102; +lean_dec(x_10); +x_102 = !lean_is_exclusive(x_94); +if (x_102 == 0) +{ +return x_94; +} +else +{ +lean_object* x_103; lean_object* x_104; lean_object* x_105; +x_103 = lean_ctor_get(x_94, 0); +x_104 = lean_ctor_get(x_94, 1); +lean_inc(x_104); +lean_inc(x_103); +lean_dec(x_94); +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 +{ +lean_object* x_106; lean_object* x_107; lean_object* x_108; uint8_t x_109; +x_106 = l_Lean_Expr_appArg(x_91, lean_box(0)); +x_107 = l_Lean_Expr_appFnCleanup(x_91, lean_box(0)); +x_108 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__5; +x_109 = l_Lean_Expr_isConstOf(x_107, x_108); +lean_dec(x_107); +if (x_109 == 0) +{ +lean_object* x_110; lean_object* x_111; +lean_dec(x_106); +lean_dec(x_90); +lean_dec(x_74); +lean_dec(x_58); +lean_dec(x_42); +lean_dec(x_1); +x_110 = lean_box(0); +lean_inc(x_10); +x_111 = lean_apply_6(x_26, x_110, x_12, x_10, x_4, x_5, x_24); +if (lean_obj_tag(x_111) == 0) +{ +lean_object* x_112; lean_object* x_113; lean_object* x_114; uint8_t x_115; +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_st_ref_get(x_10, x_113); +lean_dec(x_10); +x_115 = !lean_is_exclusive(x_114); +if (x_115 == 0) +{ +lean_object* x_116; +x_116 = lean_ctor_get(x_114, 0); +lean_dec(x_116); +lean_ctor_set(x_114, 0, x_112); +return x_114; +} +else +{ +lean_object* x_117; lean_object* x_118; +x_117 = lean_ctor_get(x_114, 1); +lean_inc(x_117); +lean_dec(x_114); +x_118 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_118, 0, x_112); +lean_ctor_set(x_118, 1, x_117); +return x_118; +} +} +else +{ +uint8_t x_119; +lean_dec(x_10); +x_119 = !lean_is_exclusive(x_111); +if (x_119 == 0) +{ +return x_111; +} +else +{ +lean_object* x_120; lean_object* x_121; lean_object* x_122; +x_120 = lean_ctor_get(x_111, 0); +x_121 = lean_ctor_get(x_111, 1); +lean_inc(x_121); +lean_inc(x_120); +lean_dec(x_111); +x_122 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_122, 0, x_120); +lean_ctor_set(x_122, 1, x_121); +return x_122; +} +} +} +else +{ +uint64_t x_123; lean_object* x_124; +x_123 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__14; +lean_inc(x_10); +x_124 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__3(x_17, x_123, x_1, x_3, x_106, x_90, x_74, x_58, x_42, x_12, x_10, x_4, x_5, x_24); +lean_dec(x_58); +lean_dec(x_74); +lean_dec(x_90); +lean_dec(x_106); +if (lean_obj_tag(x_124) == 0) +{ +lean_object* x_125; lean_object* x_126; lean_object* x_127; uint8_t x_128; +x_125 = lean_ctor_get(x_124, 0); +lean_inc(x_125); +x_126 = lean_ctor_get(x_124, 1); +lean_inc(x_126); +lean_dec(x_124); +x_127 = lean_st_ref_get(x_10, x_126); +lean_dec(x_10); +x_128 = !lean_is_exclusive(x_127); +if (x_128 == 0) +{ +lean_object* x_129; +x_129 = lean_ctor_get(x_127, 0); +lean_dec(x_129); +lean_ctor_set(x_127, 0, x_125); +return x_127; +} +else +{ +lean_object* x_130; lean_object* x_131; +x_130 = lean_ctor_get(x_127, 1); +lean_inc(x_130); +lean_dec(x_127); +x_131 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_131, 0, x_125); +lean_ctor_set(x_131, 1, x_130); +return x_131; +} +} +else +{ +uint8_t x_132; +lean_dec(x_10); +x_132 = !lean_is_exclusive(x_124); +if (x_132 == 0) +{ +return x_124; +} +else +{ +lean_object* x_133; lean_object* x_134; lean_object* x_135; +x_133 = lean_ctor_get(x_124, 0); +x_134 = lean_ctor_get(x_124, 1); +lean_inc(x_134); +lean_inc(x_133); +lean_dec(x_124); +x_135 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_135, 0, x_133); +lean_ctor_set(x_135, 1, x_134); +return x_135; +} +} +} +} +} +} +} +} +} +else +{ +uint8_t x_136; +lean_dec(x_10); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_136 = !lean_is_exclusive(x_21); +if (x_136 == 0) +{ +return x_21; +} +else +{ +lean_object* x_137; lean_object* x_138; lean_object* x_139; +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_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); +return x_139; +} +} +} +else +{ +uint8_t x_140; +lean_dec(x_10); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_140 = !lean_is_exclusive(x_13); +if (x_140 == 0) +{ +return x_13; +} +else +{ +lean_object* x_141; lean_object* x_142; lean_object* x_143; +x_141 = lean_ctor_get(x_13, 0); +x_142 = lean_ctor_get(x_13, 1); +lean_inc(x_142); +lean_inc(x_141); +lean_dec(x_13); +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; +} +} +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__6___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("attribute cannot be erased", 26, 26); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__6___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__6___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; lean_object* x_6; +x_5 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__6___closed__2; +x_6 = l_Lean_throwError___at_Lean_Attribute_Builtin_ensureNoArgs___spec__2(x_5, x_2, x_3, x_4); +return x_6; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__1; +x_2 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__2; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__2; +x_2 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__3; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("initFn", 6, 6); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__3; +x_2 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__4; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("_@", 2, 2); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__5; +x_2 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__6; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__7; +x_2 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__9() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Elab", 4, 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__8; +x_2 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__9; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__11() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Tactic", 6, 6); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__10; +x_2 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__11; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__12; +x_2 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__3; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__14() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("_hyg", 4, 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__15() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__13; +x_2 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__14; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__16() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__15; +x_2 = lean_unsigned_to_nat(249u); +x_3 = l_Lean_Name_num___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__17() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("partial_fixpoint_monotone", 25, 25); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__18() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__17; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__19() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("monotonicity theorem", 20, 20); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__20() { +_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_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__16; +x_2 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__18; +x_3 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__19; +x_4 = 0; +x_5 = lean_alloc_ctor(0, 3, 1); +lean_ctor_set(x_5, 0, x_1); +lean_ctor_set(x_5, 1, x_2); +lean_ctor_set(x_5, 2, x_3); +lean_ctor_set_uint8(x_5, sizeof(void*)*3, x_4); +return x_5; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__21() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___boxed), 6, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__22() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__6___boxed), 4, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__23() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__20; +x_2 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__21; +x_3 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__22; +x_4 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249_(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__23; +x_3 = l_Lean_registerBuiltinAttribute(x_2, x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +uint8_t x_9; lean_object* x_10; +x_9 = lean_unbox(x_3); +lean_dec(x_3); +x_10 = l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____spec__1(x_1, x_2, x_9, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_7); +lean_dec(x_5); +lean_dec(x_4); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_lambdaBoundedTelescope___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____spec__2___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +uint8_t x_10; lean_object* x_11; +x_10 = lean_unbox(x_4); +lean_dec(x_4); +x_11 = l_Lean_Meta_lambdaBoundedTelescope___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____spec__2___rarg(x_1, x_2, x_3, x_10, x_5, x_6, x_7, x_8, x_9); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____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: +{ +uint64_t x_12; uint64_t x_13; uint8_t x_14; lean_object* x_15; +x_12 = lean_unbox_uint64(x_1); +lean_dec(x_1); +x_13 = lean_unbox_uint64(x_2); +lean_dec(x_2); +x_14 = lean_unbox(x_4); +lean_dec(x_4); +x_15 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__1(x_12, x_13, x_3, x_14, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_5); +return x_15; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____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: +{ +uint64_t x_11; uint64_t x_12; uint8_t x_13; lean_object* x_14; +x_11 = lean_unbox_uint64(x_1); +lean_dec(x_1); +x_12 = lean_unbox_uint64(x_2); +lean_dec(x_2); +x_13 = lean_unbox(x_4); +lean_dec(x_4); +x_14 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__2(x_11, x_12, x_3, x_13, x_5, x_6, x_7, x_8, x_9, x_10); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____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: +{ +uint64_t x_15; uint64_t x_16; uint8_t x_17; lean_object* x_18; +x_15 = lean_unbox_uint64(x_1); +lean_dec(x_1); +x_16 = lean_unbox_uint64(x_2); +lean_dec(x_2); +x_17 = lean_unbox(x_4); +lean_dec(x_4); +x_18 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__3(x_15, x_16, x_3, x_17, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +return x_18; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; +x_7 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +uint8_t x_7; lean_object* x_8; +x_7 = lean_unbox(x_3); +lean_dec(x_3); +x_8 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5(x_1, x_2, x_7, x_4, x_5, x_6); +lean_dec(x_2); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__6(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_Monotonicity_findMonoThms(lean_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_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 = l_panic___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__13___closed__1; +x_12 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__1___closed__1; +x_13 = l_Lean_ScopedEnvExtension_getState___rarg(x_11, x_12, x_10); +lean_dec(x_10); +x_14 = l_Lean_Meta_DiscrTree_getMatch___rarg(x_13, x_1, x_2, x_3, x_4, x_5, x_9); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_defaultFailK___spec__1___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; lean_object* x_8; uint8_t x_9; +x_7 = lean_ctor_get(x_4, 5); +x_8 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_1, x_2, x_3, x_4, x_5, x_6); +x_9 = !lean_is_exclusive(x_8); +if (x_9 == 0) +{ +lean_object* x_10; lean_object* x_11; +x_10 = lean_ctor_get(x_8, 0); +lean_inc(x_7); +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_7); +lean_ctor_set(x_11, 1, x_10); +lean_ctor_set_tag(x_8, 1); +lean_ctor_set(x_8, 0, x_11); +return x_8; +} +else +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_12 = lean_ctor_get(x_8, 0); +x_13 = lean_ctor_get(x_8, 1); +lean_inc(x_13); +lean_inc(x_12); +lean_dec(x_8); +lean_inc(x_7); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_7); +lean_ctor_set(x_14, 1, x_12); +x_15 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_15, 0, x_14); +lean_ctor_set(x_15, 1, x_13); +return x_15; +} +} +} +LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_defaultFailK___spec__1(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_throwError___at___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_defaultFailK___spec__1___rarg___boxed), 6, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_defaultFailK___spec__2___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; lean_object* x_8; uint8_t x_9; +x_7 = lean_ctor_get(x_4, 5); +x_8 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_1, x_2, x_3, x_4, x_5, x_6); +x_9 = !lean_is_exclusive(x_8); +if (x_9 == 0) +{ +lean_object* x_10; lean_object* x_11; +x_10 = lean_ctor_get(x_8, 0); +lean_inc(x_7); +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_7); +lean_ctor_set(x_11, 1, x_10); +lean_ctor_set_tag(x_8, 1); +lean_ctor_set(x_8, 0, x_11); +return x_8; +} +else +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_12 = lean_ctor_get(x_8, 0); +x_13 = lean_ctor_get(x_8, 1); +lean_inc(x_13); +lean_inc(x_12); +lean_dec(x_8); +lean_inc(x_7); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_7); +lean_ctor_set(x_14, 1, x_12); +x_15 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_15, 0, x_14); +lean_ctor_set(x_15, 1, x_13); +return x_15; +} +} +} +LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_defaultFailK___spec__2(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_throwError___at___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_defaultFailK___spec__2___rarg___boxed), 6, 0); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_defaultFailK___rarg___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Failed to prove monotonicity of:", 32, 32); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_defaultFailK___rarg___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_defaultFailK___rarg___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_defaultFailK___rarg___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("\n", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_defaultFailK___rarg___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_defaultFailK___rarg___closed__3; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_defaultFailK___rarg___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Tried to apply ", 15, 15); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_defaultFailK___rarg___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_defaultFailK___rarg___closed__5; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_defaultFailK___rarg___closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(", but failed.", 13, 13); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_defaultFailK___rarg___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_defaultFailK___rarg___closed__7; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_defaultFailK___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; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_8 = l_Array_isEmpty___rarg(x_2); +x_9 = l_Lean_indentExpr(x_1); +x_10 = l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_defaultFailK___rarg___closed__2; +x_11 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_11, 0, x_10); +lean_ctor_set(x_11, 1, x_9); +x_12 = l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_defaultFailK___rarg___closed__4; +x_13 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_13, 0, x_11); +lean_ctor_set(x_13, 1, x_12); +if (x_8 == 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; +x_14 = lean_array_to_list(x_2); +x_15 = lean_box(0); +x_16 = l_List_mapTR_loop___at_Lean_Parser_Tactic_Doc_initFn____x40_Lean_Parser_Tactic_Doc___hyg_2831____spec__8(x_14, x_15); +x_17 = l_Lean_MessageData_andList(x_16); +x_18 = l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_defaultFailK___rarg___closed__6; +x_19 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_17); +x_20 = l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_defaultFailK___rarg___closed__8; +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_13); +lean_ctor_set(x_22, 1, x_21); +x_23 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__9; +x_24 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_24, 0, x_22); +lean_ctor_set(x_24, 1, x_23); +x_25 = l_Lean_throwError___at___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_defaultFailK___spec__1___rarg(x_24, x_3, x_4, x_5, x_6, x_7); +return x_25; +} +else +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +lean_dec(x_2); +x_26 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__9; +x_27 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_27, 0, x_13); +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_26); +x_29 = l_Lean_throwError___at___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_defaultFailK___spec__2___rarg(x_28, x_3, x_4, x_5, x_6, x_7); +return x_29; +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_defaultFailK(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_defaultFailK___rarg___boxed), 7, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_defaultFailK___spec__1___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; +x_7 = l_Lean_throwError___at___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_defaultFailK___spec__1___rarg(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_defaultFailK___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) { +_start: +{ +lean_object* x_7; +x_7 = l_Lean_throwError___at___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_defaultFailK___spec__2___rarg(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_7; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_defaultFailK___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___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_defaultFailK___rarg(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; +} +} +static lean_object* _init_l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_applyConst___lambda__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Could not apply ", 16, 16); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_applyConst___lambda__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_applyConst___lambda__1___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_applyConst___lambda__1___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(":", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_applyConst___lambda__1___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_applyConst___lambda__1___closed__3; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_applyConst___lambda__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_3 = 0; +x_4 = l_Lean_MessageData_ofConstName(x_1, x_3); +x_5 = l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_applyConst___lambda__1___closed__2; +x_6 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_6, 0, x_5); +lean_ctor_set(x_6, 1, x_4); +x_7 = l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_applyConst___lambda__1___closed__4; +x_8 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_8, 0, x_6); +lean_ctor_set(x_8, 1, x_7); +x_9 = l_Lean_indentD(x_2); +x_10 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_10, 0, x_8); +lean_ctor_set(x_10, 1, x_9); +x_11 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__9; +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; +} +} +static lean_object* _init_l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_applyConst___closed__1() { +_start: +{ +uint8_t x_1; uint8_t x_2; uint8_t x_3; lean_object* x_4; +x_1 = 0; +x_2 = 0; +x_3 = 1; +x_4 = lean_alloc_ctor(0, 0, 4); +lean_ctor_set_uint8(x_4, 0, x_1); +lean_ctor_set_uint8(x_4, 1, x_2); +lean_ctor_set_uint8(x_4, 2, x_2); +lean_ctor_set_uint8(x_4, 3, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_applyConst(lean_object* x_1, 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; +x_8 = l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_applyConst___closed__1; +lean_inc(x_2); +x_9 = lean_alloc_closure((void*)(l_Lean_MVarId_applyConst), 8, 3); +lean_closure_set(x_9, 0, x_1); +lean_closure_set(x_9, 1, x_2); +lean_closure_set(x_9, 2, x_8); +x_10 = lean_alloc_closure((void*)(l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_applyConst___lambda__1), 2, 1); +lean_closure_set(x_10, 0, x_2); +x_11 = l_Lean_Meta_mapErrorImp___rarg(x_9, x_10, x_3, x_4, x_5, x_6, x_7); +if (lean_obj_tag(x_11) == 0) +{ +uint8_t x_12; +x_12 = !lean_is_exclusive(x_11); +if (x_12 == 0) +{ +return x_11; +} +else +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_13 = lean_ctor_get(x_11, 0); +x_14 = lean_ctor_get(x_11, 1); +lean_inc(x_14); +lean_inc(x_13); +lean_dec(x_11); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_13); +lean_ctor_set(x_15, 1, x_14); +return x_15; +} +} +else +{ +uint8_t x_16; +x_16 = !lean_is_exclusive(x_11); +if (x_16 == 0) +{ +return x_11; +} +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_17 = lean_ctor_get(x_11, 0); +x_18 = lean_ctor_get(x_11, 1); +lean_inc(x_18); +lean_inc(x_17); +lean_dec(x_11); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_17); +lean_ctor_set(x_19, 1, x_18); +return x_19; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoCall___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; lean_object* x_8; +x_7 = lean_box(0); +x_8 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_8, 0, x_7); +lean_ctor_set(x_8, 1, x_6); +return x_8; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_solveMonoCall___lambda__2___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Monotonicity_solveMonoCall___lambda__1___boxed), 6, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_solveMonoCall___lambda__2___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = l_Lean_Expr_bvar___override(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_solveMonoCall___lambda__2___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("monotone_id", 11, 11); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_solveMonoCall___lambda__2___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__1; +x_2 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__3; +x_3 = l_Lean_Meta_Monotonicity_solveMonoCall___lambda__2___closed__3; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoCall___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; lean_object* x_11; uint8_t x_12; +x_10 = l_Lean_Meta_Monotonicity_solveMonoCall___lambda__2___closed__1; +x_11 = l_Lean_Meta_Monotonicity_solveMonoCall___lambda__2___closed__2; +x_12 = lean_expr_eqv(x_1, x_11); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; +lean_dec(x_3); +lean_dec(x_2); +x_13 = lean_box(0); +x_14 = lean_apply_6(x_10, x_13, x_5, x_6, x_7, x_8, x_9); +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; +x_15 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_15, 0, x_2); +x_16 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_16, 0, x_3); +x_17 = lean_box(0); +x_18 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_18, 0, x_16); +lean_ctor_set(x_18, 1, x_17); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_15); +lean_ctor_set(x_19, 1, x_18); +x_20 = lean_array_mk(x_19); +x_21 = l_Lean_Meta_Monotonicity_solveMonoCall___lambda__2___closed__4; +x_22 = l_Lean_Meta_mkAppOptM(x_21, x_20, x_5, x_6, x_7, x_8, x_9); +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; lean_object* x_25; +x_24 = lean_ctor_get(x_22, 0); +x_25 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_22, 0, x_25); +return x_22; +} +else +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_26 = lean_ctor_get(x_22, 0); +x_27 = lean_ctor_get(x_22, 1); +lean_inc(x_27); +lean_inc(x_26); +lean_dec(x_22); +x_28 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_28, 0, x_26); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_27); +return x_29; +} +} +else +{ +uint8_t x_30; +x_30 = !lean_is_exclusive(x_22); +if (x_30 == 0) +{ +return x_22; +} +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_22, 0); +x_32 = lean_ctor_get(x_22, 1); +lean_inc(x_32); +lean_inc(x_31); +lean_dec(x_22); +x_33 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_33, 0, x_31); +lean_ctor_set(x_33, 1, x_32); +return x_33; +} +} +} +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_solveMonoCall___lambda__3___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("monotone_pprod_snd", 18, 18); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_solveMonoCall___lambda__3___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__1; +x_2 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__3; +x_3 = l_Lean_Meta_Monotonicity_solveMonoCall___lambda__3___closed__1; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_solveMonoCall___lambda__3___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("monotone_pprod_fst", 18, 18); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_solveMonoCall___lambda__3___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__1; +x_2 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__3; +x_3 = l_Lean_Meta_Monotonicity_solveMonoCall___lambda__3___closed__3; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoCall___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +lean_object* x_14; lean_object* x_15; 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; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; 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_14 = l_Lean_Expr_projIdx_x21(x_1); +x_15 = lean_unsigned_to_nat(0u); +x_16 = lean_nat_dec_eq(x_14, x_15); +lean_dec(x_14); +x_17 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_17, 0, x_5); +x_18 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_18, 0, x_6); +x_19 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_19, 0, x_2); +x_20 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_20, 0, x_7); +x_21 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_21, 0, x_8); +x_22 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_22, 0, x_3); +x_23 = lean_box(0); +x_24 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_24, 0, x_4); +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 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_27, 0, x_23); +lean_ctor_set(x_27, 1, x_26); +x_28 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_28, 0, x_22); +lean_ctor_set(x_28, 1, x_27); +x_29 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_29, 0, x_21); +lean_ctor_set(x_29, 1, x_28); +x_30 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_30, 0, x_20); +lean_ctor_set(x_30, 1, x_29); +x_31 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_31, 0, x_19); +lean_ctor_set(x_31, 1, x_30); +x_32 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_32, 0, x_18); +lean_ctor_set(x_32, 1, x_31); +x_33 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_33, 0, x_17); +lean_ctor_set(x_33, 1, x_32); +x_34 = lean_array_mk(x_33); +if (x_16 == 0) +{ +lean_object* x_35; lean_object* x_36; +x_35 = l_Lean_Meta_Monotonicity_solveMonoCall___lambda__3___closed__2; +x_36 = l_Lean_Meta_mkAppOptM(x_35, x_34, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_36) == 0) +{ +uint8_t x_37; +x_37 = !lean_is_exclusive(x_36); +if (x_37 == 0) +{ +lean_object* x_38; lean_object* x_39; +x_38 = lean_ctor_get(x_36, 0); +x_39 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_36, 0, x_39); +return x_36; +} +else +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_40 = lean_ctor_get(x_36, 0); +x_41 = lean_ctor_get(x_36, 1); +lean_inc(x_41); +lean_inc(x_40); +lean_dec(x_36); +x_42 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_42, 0, 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_41); +return x_43; +} +} +else +{ +uint8_t x_44; +x_44 = !lean_is_exclusive(x_36); +if (x_44 == 0) +{ +return x_36; +} +else +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_45 = lean_ctor_get(x_36, 0); +x_46 = lean_ctor_get(x_36, 1); +lean_inc(x_46); +lean_inc(x_45); +lean_dec(x_36); +x_47 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_47, 0, x_45); +lean_ctor_set(x_47, 1, x_46); +return x_47; +} +} +} +else +{ +lean_object* x_48; lean_object* x_49; +x_48 = l_Lean_Meta_Monotonicity_solveMonoCall___lambda__3___closed__4; +x_49 = l_Lean_Meta_mkAppOptM(x_48, x_34, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_49) == 0) +{ +uint8_t x_50; +x_50 = !lean_is_exclusive(x_49); +if (x_50 == 0) +{ +lean_object* x_51; lean_object* x_52; +x_51 = lean_ctor_get(x_49, 0); +x_52 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_52, 0, x_51); +lean_ctor_set(x_49, 0, x_52); +return x_49; +} +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_alloc_ctor(1, 1, 0); +lean_ctor_set(x_55, 0, 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; +} +} +else +{ +uint8_t x_57; +x_57 = !lean_is_exclusive(x_49); +if (x_57 == 0) +{ +return x_49; +} +else +{ +lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_58 = lean_ctor_get(x_49, 0); +x_59 = lean_ctor_get(x_49, 1); +lean_inc(x_59); +lean_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; +} +} +} +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_solveMonoCall___lambda__4___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("solveMonoCall ", 14, 14); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_solveMonoCall___lambda__4___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Monotonicity_solveMonoCall___lambda__4___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_solveMonoCall___lambda__4___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(": whnfUntil failed\?", 19, 19); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_solveMonoCall___lambda__4___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Monotonicity_solveMonoCall___lambda__4___closed__3; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoCall___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; 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_10 = l_Lean_MessageData_ofExpr(x_1); +x_11 = l_Lean_Meta_Monotonicity_solveMonoCall___lambda__4___closed__2; +x_12 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_10); +x_13 = l_Lean_Meta_Monotonicity_solveMonoCall___lambda__4___closed__4; +x_14 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_14, 0, x_12); +lean_ctor_set(x_14, 1, x_13); +x_15 = l_Lean_indentExpr(x_2); +x_16 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_16, 0, x_14); +lean_ctor_set(x_16, 1, x_15); +x_17 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__9; +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___at___private_Lean_Meta_RecursorInfo_0__Lean_Meta_getMajorPosDepElim___spec__3(x_18, x_5, x_6, x_7, x_8, x_9); +x_20 = !lean_is_exclusive(x_19); +if (x_20 == 0) +{ +return x_19; +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_19, 0); +x_22 = lean_ctor_get(x_19, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_19); +x_23 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_23, 0, x_21); +lean_ctor_set(x_23, 1, x_22); +return x_23; +} +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_solveMonoCall___lambda__5___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("instPartialOrderPProd", 21, 21); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_solveMonoCall___lambda__5___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__1; +x_2 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__3; +x_3 = l_Lean_Meta_Monotonicity_solveMonoCall___lambda__5___closed__1; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_solveMonoCall___lambda__5___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(": unexpected instance ", 22, 22); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_solveMonoCall___lambda__5___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Monotonicity_solveMonoCall___lambda__5___closed__3; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoCall___lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; lean_object* x_13; +x_12 = l_Lean_Meta_Monotonicity_solveMonoCall___lambda__5___closed__2; +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +x_13 = l_Lean_Meta_whnfUntil(x_6, x_12, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_13) == 0) +{ +lean_object* x_14; +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; 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_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +x_16 = l_Lean_MessageData_ofExpr(x_1); +x_17 = l_Lean_Meta_Monotonicity_solveMonoCall___lambda__4___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_Lean_Meta_Monotonicity_solveMonoCall___lambda__5___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_Lean_MessageData_ofExpr(x_6); +x_22 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_22, 0, x_20); +lean_ctor_set(x_22, 1, x_21); +x_23 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__9; +x_24 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_24, 0, x_22); +lean_ctor_set(x_24, 1, x_23); +x_25 = l_Lean_throwError___at___private_Lean_Meta_RecursorInfo_0__Lean_Meta_getMajorPosDepElim___spec__3(x_24, x_7, x_8, x_9, x_10, x_15); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +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_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; +lean_dec(x_6); +x_30 = lean_ctor_get(x_13, 1); +lean_inc(x_30); +lean_dec(x_13); +x_31 = lean_ctor_get(x_14, 0); +lean_inc(x_31); +lean_dec(x_14); +lean_inc(x_31); +x_32 = l_Lean_Meta_instantiateMVarsIfMVarApp(x_31, x_7, x_8, x_9, x_10, x_30); +x_33 = lean_ctor_get(x_32, 0); +lean_inc(x_33); +x_34 = lean_ctor_get(x_32, 1); +lean_inc(x_34); +lean_dec(x_32); +x_35 = l_Lean_Expr_cleanupAnnotations(x_33); +x_36 = l_Lean_Expr_isApp(x_35); +if (x_36 == 0) +{ +lean_object* x_37; lean_object* x_38; +lean_dec(x_35); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_37 = lean_box(0); +x_38 = l_Lean_Meta_Monotonicity_solveMonoCall___lambda__4(x_1, x_31, x_2, x_37, 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); +return x_38; +} +else +{ +lean_object* x_39; lean_object* x_40; uint8_t x_41; +x_39 = l_Lean_Expr_appArg(x_35, lean_box(0)); +x_40 = l_Lean_Expr_appFnCleanup(x_35, lean_box(0)); +x_41 = l_Lean_Expr_isApp(x_40); +if (x_41 == 0) +{ +lean_object* x_42; lean_object* x_43; +lean_dec(x_40); +lean_dec(x_39); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_42 = lean_box(0); +x_43 = l_Lean_Meta_Monotonicity_solveMonoCall___lambda__4(x_1, x_31, x_2, x_42, 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); +return x_43; +} +else +{ +lean_object* x_44; lean_object* x_45; uint8_t x_46; +x_44 = l_Lean_Expr_appArg(x_40, lean_box(0)); +x_45 = l_Lean_Expr_appFnCleanup(x_40, lean_box(0)); +x_46 = l_Lean_Expr_isApp(x_45); +if (x_46 == 0) +{ +lean_object* x_47; lean_object* x_48; +lean_dec(x_45); +lean_dec(x_44); +lean_dec(x_39); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_47 = lean_box(0); +x_48 = l_Lean_Meta_Monotonicity_solveMonoCall___lambda__4(x_1, x_31, x_2, x_47, 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); +return x_48; +} +else +{ +lean_object* x_49; lean_object* x_50; uint8_t x_51; +x_49 = l_Lean_Expr_appArg(x_45, lean_box(0)); +x_50 = l_Lean_Expr_appFnCleanup(x_45, lean_box(0)); +x_51 = l_Lean_Expr_isApp(x_50); +if (x_51 == 0) +{ +lean_object* x_52; lean_object* x_53; +lean_dec(x_50); +lean_dec(x_49); +lean_dec(x_44); +lean_dec(x_39); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_52 = lean_box(0); +x_53 = l_Lean_Meta_Monotonicity_solveMonoCall___lambda__4(x_1, x_31, x_2, x_52, 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); +return x_53; +} +else +{ +lean_object* x_54; lean_object* x_55; uint8_t x_56; +x_54 = l_Lean_Expr_appArg(x_50, lean_box(0)); +x_55 = l_Lean_Expr_appFnCleanup(x_50, lean_box(0)); +x_56 = l_Lean_Expr_isConstOf(x_55, x_12); +lean_dec(x_55); +if (x_56 == 0) +{ +lean_object* x_57; lean_object* x_58; +lean_dec(x_54); +lean_dec(x_49); +lean_dec(x_44); +lean_dec(x_39); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_57 = lean_box(0); +x_58 = l_Lean_Meta_Monotonicity_solveMonoCall___lambda__4(x_1, x_31, x_2, x_57, 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); +return x_58; +} +else +{ +lean_object* x_59; +lean_dec(x_31); +x_59 = l_Lean_Meta_Monotonicity_solveMonoCall___lambda__3(x_1, x_3, x_4, x_5, x_54, x_49, x_44, x_39, x_7, x_8, x_9, x_10, x_34); +lean_dec(x_1); +return x_59; +} +} +} +} +} +} +} +else +{ +uint8_t x_60; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_60 = !lean_is_exclusive(x_13); +if (x_60 == 0) +{ +return x_13; +} +else +{ +lean_object* x_61; lean_object* x_62; lean_object* x_63; +x_61 = lean_ctor_get(x_13, 0); +x_62 = lean_ctor_get(x_13, 1); +lean_inc(x_62); +lean_inc(x_61); +lean_dec(x_13); +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; +} +} +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_solveMonoCall___lambda__6___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(": unexpected type ", 18, 18); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_solveMonoCall___lambda__6___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Monotonicity_solveMonoCall___lambda__6___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoCall___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; 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_10 = l_Lean_MessageData_ofExpr(x_1); +x_11 = l_Lean_Meta_Monotonicity_solveMonoCall___lambda__4___closed__2; +x_12 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_10); +x_13 = l_Lean_Meta_Monotonicity_solveMonoCall___lambda__6___closed__2; +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_MessageData_ofExpr(x_2); +x_16 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_16, 0, x_14); +lean_ctor_set(x_16, 1, x_15); +x_17 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__9; +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___at___private_Lean_Meta_RecursorInfo_0__Lean_Meta_getMajorPosDepElim___spec__3(x_18, x_5, x_6, x_7, x_8, x_9); +x_20 = !lean_is_exclusive(x_19); +if (x_20 == 0) +{ +return x_19; +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_19, 0); +x_22 = lean_ctor_get(x_19, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_19); +x_23 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_23, 0, x_21); +lean_ctor_set(x_23, 1, x_22); +return x_23; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoCall___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) { +_start: +{ +lean_object* x_10; uint8_t x_11; +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_10 = lean_alloc_closure((void*)(l_Lean_Meta_Monotonicity_solveMonoCall___lambda__2___boxed), 9, 3); +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 = l_Lean_Expr_isProj(x_1); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; +lean_dec(x_10); +x_12 = lean_box(0); +x_13 = l_Lean_Meta_Monotonicity_solveMonoCall___lambda__2(x_1, x_2, x_3, x_12, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_1); +return x_13; +} +else +{ +lean_object* x_14; lean_object* x_15; +x_14 = l_Lean_Expr_projExpr_x21(x_1); +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_Meta_Monotonicity_solveMonoCall(x_2, x_3, x_14, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_15) == 0) +{ +lean_object* x_16; +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +if (lean_obj_tag(x_16) == 0) +{ +uint8_t x_17; +lean_dec(x_10); +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_17 = !lean_is_exclusive(x_15); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; +x_18 = lean_ctor_get(x_15, 0); +lean_dec(x_18); +x_19 = lean_box(0); +lean_ctor_set(x_15, 0, x_19); +return x_15; +} +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_15, 1); +lean_inc(x_20); +lean_dec(x_15); +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_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_15, 1); +lean_inc(x_23); +lean_dec(x_15); +x_24 = lean_ctor_get(x_16, 0); +lean_inc(x_24); +lean_dec(x_16); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_24); +x_25 = lean_infer_type(x_24, x_5, x_6, x_7, x_8, x_23); +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); +lean_inc(x_26); +x_28 = l_Lean_Expr_cleanupAnnotations(x_26); +x_29 = l_Lean_Expr_isApp(x_28); +if (x_29 == 0) +{ +lean_object* x_30; lean_object* x_31; +lean_dec(x_28); +lean_dec(x_24); +lean_dec(x_3); +lean_dec(x_2); +x_30 = lean_box(0); +x_31 = l_Lean_Meta_Monotonicity_solveMonoCall___lambda__6(x_1, x_26, x_10, x_30, x_5, x_6, x_7, x_8, x_27); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_10); +return x_31; +} +else +{ +lean_object* x_32; uint8_t x_33; +x_32 = l_Lean_Expr_appFnCleanup(x_28, lean_box(0)); +x_33 = l_Lean_Expr_isApp(x_32); +if (x_33 == 0) +{ +lean_object* x_34; lean_object* x_35; +lean_dec(x_32); +lean_dec(x_24); +lean_dec(x_3); +lean_dec(x_2); +x_34 = lean_box(0); +x_35 = l_Lean_Meta_Monotonicity_solveMonoCall___lambda__6(x_1, x_26, x_10, x_34, x_5, x_6, x_7, x_8, x_27); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_10); +return x_35; +} +else +{ +lean_object* x_36; lean_object* x_37; uint8_t x_38; +x_36 = l_Lean_Expr_appArg(x_32, lean_box(0)); +x_37 = l_Lean_Expr_appFnCleanup(x_32, lean_box(0)); +x_38 = l_Lean_Expr_isApp(x_37); +if (x_38 == 0) +{ +lean_object* x_39; lean_object* x_40; +lean_dec(x_37); +lean_dec(x_36); +lean_dec(x_24); +lean_dec(x_3); +lean_dec(x_2); +x_39 = lean_box(0); +x_40 = l_Lean_Meta_Monotonicity_solveMonoCall___lambda__6(x_1, x_26, x_10, x_39, x_5, x_6, x_7, x_8, x_27); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_10); +return x_40; +} +else +{ +lean_object* x_41; uint8_t x_42; +x_41 = l_Lean_Expr_appFnCleanup(x_37, lean_box(0)); +x_42 = l_Lean_Expr_isApp(x_41); +if (x_42 == 0) +{ +lean_object* x_43; lean_object* x_44; +lean_dec(x_41); +lean_dec(x_36); +lean_dec(x_24); +lean_dec(x_3); +lean_dec(x_2); +x_43 = lean_box(0); +x_44 = l_Lean_Meta_Monotonicity_solveMonoCall___lambda__6(x_1, x_26, x_10, x_43, x_5, x_6, x_7, x_8, x_27); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_10); +return x_44; +} +else +{ +lean_object* x_45; uint8_t x_46; +x_45 = l_Lean_Expr_appFnCleanup(x_41, lean_box(0)); +x_46 = l_Lean_Expr_isApp(x_45); +if (x_46 == 0) +{ +lean_object* x_47; lean_object* x_48; +lean_dec(x_45); +lean_dec(x_36); +lean_dec(x_24); +lean_dec(x_3); +lean_dec(x_2); +x_47 = lean_box(0); +x_48 = l_Lean_Meta_Monotonicity_solveMonoCall___lambda__6(x_1, x_26, x_10, x_47, x_5, x_6, x_7, x_8, x_27); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_10); +return x_48; +} +else +{ +lean_object* x_49; lean_object* x_50; uint8_t x_51; +x_49 = l_Lean_Expr_appFnCleanup(x_45, lean_box(0)); +x_50 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__5; +x_51 = l_Lean_Expr_isConstOf(x_49, x_50); +lean_dec(x_49); +if (x_51 == 0) +{ +lean_object* x_52; lean_object* x_53; +lean_dec(x_36); +lean_dec(x_24); +lean_dec(x_3); +lean_dec(x_2); +x_52 = lean_box(0); +x_53 = l_Lean_Meta_Monotonicity_solveMonoCall___lambda__6(x_1, x_26, x_10, x_52, x_5, x_6, x_7, x_8, x_27); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_10); +return x_53; +} +else +{ +lean_object* x_54; +lean_dec(x_26); +x_54 = l_Lean_Meta_Monotonicity_solveMonoCall___lambda__5(x_1, x_10, x_2, x_3, x_24, x_36, x_5, x_6, x_7, x_8, x_27); +lean_dec(x_10); +return x_54; +} +} +} +} +} +} +} +else +{ +uint8_t x_55; +lean_dec(x_24); +lean_dec(x_10); +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_55 = !lean_is_exclusive(x_25); +if (x_55 == 0) +{ +return x_25; +} +else +{ +lean_object* x_56; lean_object* x_57; lean_object* x_58; +x_56 = lean_ctor_get(x_25, 0); +x_57 = lean_ctor_get(x_25, 1); +lean_inc(x_57); +lean_inc(x_56); +lean_dec(x_25); +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; +lean_dec(x_10); +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_59 = !lean_is_exclusive(x_15); +if (x_59 == 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_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; +} +} +} +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_solveMonoCall___lambda__8___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("monotone_apply", 14, 14); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_solveMonoCall___lambda__8___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__1; +x_2 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__3; +x_3 = l_Lean_Meta_Monotonicity_solveMonoCall___lambda__8___closed__1; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoCall___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_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; 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_13 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_13, 0, x_5); +x_14 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_14, 0, x_6); +x_15 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_15, 0, x_1); +x_16 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_16, 0, x_2); +x_17 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_17, 0, x_7); +x_18 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_18, 0, x_3); +x_19 = lean_box(0); +x_20 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_20, 0, x_4); +x_21 = lean_box(0); +x_22 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_22, 0, x_20); +lean_ctor_set(x_22, 1, x_21); +x_23 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_23, 0, x_19); +lean_ctor_set(x_23, 1, x_22); +x_24 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_24, 0, x_18); +lean_ctor_set(x_24, 1, x_23); +x_25 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_25, 0, x_17); +lean_ctor_set(x_25, 1, x_24); +x_26 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_26, 0, x_16); +lean_ctor_set(x_26, 1, x_25); +x_27 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_27, 0, x_15); +lean_ctor_set(x_27, 1, x_26); +x_28 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_28, 0, x_14); +lean_ctor_set(x_28, 1, x_27); +x_29 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_29, 0, x_13); +lean_ctor_set(x_29, 1, x_28); +x_30 = lean_array_mk(x_29); +x_31 = l_Lean_Meta_Monotonicity_solveMonoCall___lambda__8___closed__2; +x_32 = l_Lean_Meta_mkAppOptM(x_31, x_30, x_8, x_9, x_10, x_11, x_12); +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; lean_object* x_35; +x_34 = lean_ctor_get(x_32, 0); +x_35 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_35, 0, x_34); +lean_ctor_set(x_32, 0, x_35); +return x_32; +} +else +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_36 = lean_ctor_get(x_32, 0); +x_37 = lean_ctor_get(x_32, 1); +lean_inc(x_37); +lean_inc(x_36); +lean_dec(x_32); +x_38 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_38, 0, x_36); +x_39 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_37); +return x_39; +} +} +else +{ +uint8_t x_40; +x_40 = !lean_is_exclusive(x_32); +if (x_40 == 0) +{ +return x_32; +} +else +{ +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_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; +} +} +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_solveMonoCall___lambda__9___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("instOrderPi", 11, 11); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_solveMonoCall___lambda__9___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__1; +x_2 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__3; +x_3 = l_Lean_Meta_Monotonicity_solveMonoCall___lambda__9___closed__1; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoCall___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_13; lean_object* x_14; +x_13 = l_Lean_Meta_Monotonicity_solveMonoCall___lambda__9___closed__2; +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +x_14 = l_Lean_Meta_whnfUntil(x_7, x_13, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +if (lean_obj_tag(x_15) == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; 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_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_16 = lean_ctor_get(x_14, 1); +lean_inc(x_16); +lean_dec(x_14); +x_17 = l_Lean_MessageData_ofExpr(x_1); +x_18 = l_Lean_Meta_Monotonicity_solveMonoCall___lambda__4___closed__2; +x_19 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_17); +x_20 = l_Lean_Meta_Monotonicity_solveMonoCall___lambda__5___closed__4; +x_21 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_21, 0, x_19); +lean_ctor_set(x_21, 1, x_20); +x_22 = l_Lean_MessageData_ofExpr(x_7); +x_23 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_23, 0, x_21); +lean_ctor_set(x_23, 1, x_22); +x_24 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___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); +x_26 = l_Lean_throwError___at___private_Lean_Meta_RecursorInfo_0__Lean_Meta_getMajorPosDepElim___spec__3(x_25, x_8, x_9, x_10, x_11, x_16); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(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; +} +} +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; +lean_dec(x_7); +x_31 = lean_ctor_get(x_14, 1); +lean_inc(x_31); +lean_dec(x_14); +x_32 = lean_ctor_get(x_15, 0); +lean_inc(x_32); +lean_dec(x_15); +lean_inc(x_32); +x_33 = l_Lean_Meta_instantiateMVarsIfMVarApp(x_32, x_8, x_9, x_10, x_11, x_31); +x_34 = lean_ctor_get(x_33, 0); +lean_inc(x_34); +x_35 = lean_ctor_get(x_33, 1); +lean_inc(x_35); +lean_dec(x_33); +x_36 = l_Lean_Expr_cleanupAnnotations(x_34); +x_37 = l_Lean_Expr_isApp(x_36); +if (x_37 == 0) +{ +lean_object* x_38; lean_object* x_39; +lean_dec(x_36); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_38 = lean_box(0); +x_39 = l_Lean_Meta_Monotonicity_solveMonoCall___lambda__4(x_1, x_32, x_2, x_38, x_8, x_9, x_10, x_11, x_35); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +return x_39; +} +else +{ +lean_object* x_40; lean_object* x_41; uint8_t x_42; +x_40 = l_Lean_Expr_appArg(x_36, lean_box(0)); +x_41 = l_Lean_Expr_appFnCleanup(x_36, lean_box(0)); +x_42 = l_Lean_Expr_isApp(x_41); +if (x_42 == 0) +{ +lean_object* x_43; lean_object* x_44; +lean_dec(x_41); +lean_dec(x_40); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_43 = lean_box(0); +x_44 = l_Lean_Meta_Monotonicity_solveMonoCall___lambda__4(x_1, x_32, x_2, x_43, x_8, x_9, x_10, x_11, x_35); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +return x_44; +} +else +{ +lean_object* x_45; lean_object* x_46; uint8_t x_47; +x_45 = l_Lean_Expr_appArg(x_41, lean_box(0)); +x_46 = l_Lean_Expr_appFnCleanup(x_41, lean_box(0)); +x_47 = l_Lean_Expr_isApp(x_46); +if (x_47 == 0) +{ +lean_object* x_48; lean_object* x_49; +lean_dec(x_46); +lean_dec(x_45); +lean_dec(x_40); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_48 = lean_box(0); +x_49 = l_Lean_Meta_Monotonicity_solveMonoCall___lambda__4(x_1, x_32, x_2, x_48, x_8, x_9, x_10, x_11, x_35); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +return x_49; +} +else +{ +lean_object* x_50; lean_object* x_51; uint8_t x_52; +x_50 = l_Lean_Expr_appArg(x_46, lean_box(0)); +x_51 = l_Lean_Expr_appFnCleanup(x_46, lean_box(0)); +x_52 = l_Lean_Expr_isConstOf(x_51, x_13); +lean_dec(x_51); +if (x_52 == 0) +{ +lean_object* x_53; lean_object* x_54; +lean_dec(x_50); +lean_dec(x_45); +lean_dec(x_40); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_53 = lean_box(0); +x_54 = l_Lean_Meta_Monotonicity_solveMonoCall___lambda__4(x_1, x_32, x_2, x_53, x_8, x_9, x_10, x_11, x_35); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +return x_54; +} +else +{ +lean_object* x_55; +lean_dec(x_32); +lean_dec(x_1); +x_55 = l_Lean_Meta_Monotonicity_solveMonoCall___lambda__8(x_3, x_4, x_5, x_6, x_50, x_45, x_40, x_8, x_9, x_10, x_11, x_35); +return x_55; +} +} +} +} +} +} +else +{ +uint8_t x_56; +lean_dec(x_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_56 = !lean_is_exclusive(x_14); +if (x_56 == 0) +{ +return x_14; +} +else +{ +lean_object* x_57; lean_object* x_58; lean_object* x_59; +x_57 = lean_ctor_get(x_14, 0); +x_58 = lean_ctor_get(x_14, 1); +lean_inc(x_58); +lean_inc(x_57); +lean_dec(x_14); +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_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoCall(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; uint8_t x_10; +lean_inc(x_2); +lean_inc(x_1); +lean_inc(x_3); +x_9 = lean_alloc_closure((void*)(l_Lean_Meta_Monotonicity_solveMonoCall___lambda__7___boxed), 9, 3); +lean_closure_set(x_9, 0, x_3); +lean_closure_set(x_9, 1, x_1); +lean_closure_set(x_9, 2, x_2); +x_10 = l_Lean_Expr_isApp(x_3); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; +lean_dec(x_9); +x_11 = lean_box(0); +x_12 = l_Lean_Meta_Monotonicity_solveMonoCall___lambda__7(x_3, x_1, x_2, x_11, x_4, x_5, x_6, x_7, x_8); +return x_12; +} +else +{ +lean_object* x_13; uint8_t x_14; +x_13 = l_Lean_Expr_appArg_x21(x_3); +x_14 = l_Lean_Expr_hasLooseBVars(x_13); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; +x_15 = l_Lean_Expr_appFn_x21(x_3); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_2); +lean_inc(x_1); +x_16 = l_Lean_Meta_Monotonicity_solveMonoCall(x_1, x_2, x_15, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_17; +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +if (lean_obj_tag(x_17) == 0) +{ +uint8_t x_18; +lean_dec(x_13); +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_18 = !lean_is_exclusive(x_16); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; +x_19 = lean_ctor_get(x_16, 0); +lean_dec(x_19); +x_20 = lean_box(0); +lean_ctor_set(x_16, 0, x_20); +return x_16; +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_16, 1); +lean_inc(x_21); +lean_dec(x_16); +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; +x_24 = lean_ctor_get(x_16, 1); +lean_inc(x_24); +lean_dec(x_16); +x_25 = lean_ctor_get(x_17, 0); +lean_inc(x_25); +lean_dec(x_17); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_25); +x_26 = lean_infer_type(x_25, x_4, x_5, x_6, x_7, x_24); +if (lean_obj_tag(x_26) == 0) +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; uint8_t x_30; +x_27 = lean_ctor_get(x_26, 0); +lean_inc(x_27); +x_28 = lean_ctor_get(x_26, 1); +lean_inc(x_28); +lean_dec(x_26); +lean_inc(x_27); +x_29 = l_Lean_Expr_cleanupAnnotations(x_27); +x_30 = l_Lean_Expr_isApp(x_29); +if (x_30 == 0) +{ +lean_object* x_31; lean_object* x_32; +lean_dec(x_29); +lean_dec(x_25); +lean_dec(x_13); +lean_dec(x_2); +lean_dec(x_1); +x_31 = lean_box(0); +x_32 = l_Lean_Meta_Monotonicity_solveMonoCall___lambda__6(x_3, x_27, x_9, x_31, x_4, x_5, x_6, x_7, x_28); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_9); +return x_32; +} +else +{ +lean_object* x_33; uint8_t x_34; +x_33 = l_Lean_Expr_appFnCleanup(x_29, lean_box(0)); +x_34 = l_Lean_Expr_isApp(x_33); +if (x_34 == 0) +{ +lean_object* x_35; lean_object* x_36; +lean_dec(x_33); +lean_dec(x_25); +lean_dec(x_13); +lean_dec(x_2); +lean_dec(x_1); +x_35 = lean_box(0); +x_36 = l_Lean_Meta_Monotonicity_solveMonoCall___lambda__6(x_3, x_27, x_9, x_35, x_4, x_5, x_6, x_7, x_28); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_9); +return x_36; +} +else +{ +lean_object* x_37; lean_object* x_38; uint8_t x_39; +x_37 = l_Lean_Expr_appArg(x_33, lean_box(0)); +x_38 = l_Lean_Expr_appFnCleanup(x_33, lean_box(0)); +x_39 = l_Lean_Expr_isApp(x_38); +if (x_39 == 0) +{ +lean_object* x_40; lean_object* x_41; +lean_dec(x_38); +lean_dec(x_37); +lean_dec(x_25); +lean_dec(x_13); +lean_dec(x_2); +lean_dec(x_1); +x_40 = lean_box(0); +x_41 = l_Lean_Meta_Monotonicity_solveMonoCall___lambda__6(x_3, x_27, x_9, x_40, x_4, x_5, x_6, x_7, x_28); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_9); +return x_41; +} +else +{ +lean_object* x_42; uint8_t x_43; +x_42 = l_Lean_Expr_appFnCleanup(x_38, lean_box(0)); +x_43 = l_Lean_Expr_isApp(x_42); +if (x_43 == 0) +{ +lean_object* x_44; lean_object* x_45; +lean_dec(x_42); +lean_dec(x_37); +lean_dec(x_25); +lean_dec(x_13); +lean_dec(x_2); +lean_dec(x_1); +x_44 = lean_box(0); +x_45 = l_Lean_Meta_Monotonicity_solveMonoCall___lambda__6(x_3, x_27, x_9, x_44, x_4, x_5, x_6, x_7, x_28); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_9); +return x_45; +} +else +{ +lean_object* x_46; uint8_t x_47; +x_46 = l_Lean_Expr_appFnCleanup(x_42, lean_box(0)); +x_47 = l_Lean_Expr_isApp(x_46); +if (x_47 == 0) +{ +lean_object* x_48; lean_object* x_49; +lean_dec(x_46); +lean_dec(x_37); +lean_dec(x_25); +lean_dec(x_13); +lean_dec(x_2); +lean_dec(x_1); +x_48 = lean_box(0); +x_49 = l_Lean_Meta_Monotonicity_solveMonoCall___lambda__6(x_3, x_27, x_9, x_48, x_4, x_5, x_6, x_7, x_28); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_9); +return x_49; +} +else +{ +lean_object* x_50; lean_object* x_51; uint8_t x_52; +x_50 = l_Lean_Expr_appFnCleanup(x_46, lean_box(0)); +x_51 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__5; +x_52 = l_Lean_Expr_isConstOf(x_50, x_51); +lean_dec(x_50); +if (x_52 == 0) +{ +lean_object* x_53; lean_object* x_54; +lean_dec(x_37); +lean_dec(x_25); +lean_dec(x_13); +lean_dec(x_2); +lean_dec(x_1); +x_53 = lean_box(0); +x_54 = l_Lean_Meta_Monotonicity_solveMonoCall___lambda__6(x_3, x_27, x_9, x_53, x_4, x_5, x_6, x_7, x_28); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_9); +return x_54; +} +else +{ +lean_object* x_55; +lean_dec(x_27); +x_55 = l_Lean_Meta_Monotonicity_solveMonoCall___lambda__9(x_3, x_9, x_1, x_2, x_13, x_25, x_37, x_4, x_5, x_6, x_7, x_28); +lean_dec(x_9); +return x_55; +} +} +} +} +} +} +} +else +{ +uint8_t x_56; +lean_dec(x_25); +lean_dec(x_13); +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_56 = !lean_is_exclusive(x_26); +if (x_56 == 0) +{ +return x_26; +} +else +{ +lean_object* x_57; lean_object* x_58; lean_object* x_59; +x_57 = lean_ctor_get(x_26, 0); +x_58 = lean_ctor_get(x_26, 1); +lean_inc(x_58); +lean_inc(x_57); +lean_dec(x_26); +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_60; +lean_dec(x_13); +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_60 = !lean_is_exclusive(x_16); +if (x_60 == 0) +{ +return x_16; +} +else +{ +lean_object* x_61; lean_object* x_62; lean_object* x_63; +x_61 = lean_ctor_get(x_16, 0); +x_62 = lean_ctor_get(x_16, 1); +lean_inc(x_62); +lean_inc(x_61); +lean_dec(x_16); +x_63 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_63, 0, x_61); +lean_ctor_set(x_63, 1, x_62); +return x_63; +} +} +} +else +{ +lean_object* x_64; lean_object* x_65; +lean_dec(x_13); +lean_dec(x_9); +x_64 = lean_box(0); +x_65 = l_Lean_Meta_Monotonicity_solveMonoCall___lambda__7(x_3, x_1, x_2, x_64, x_4, x_5, x_6, x_7, x_8); +return x_65; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoCall___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; +x_7 = l_Lean_Meta_Monotonicity_solveMonoCall___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoCall___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_Meta_Monotonicity_solveMonoCall___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_4); +lean_dec(x_1); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoCall___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: +{ +lean_object* x_14; +x_14 = l_Lean_Meta_Monotonicity_solveMonoCall___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_1); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoCall___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l_Lean_Meta_Monotonicity_solveMonoCall___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoCall___lambda__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; +x_12 = l_Lean_Meta_Monotonicity_solveMonoCall___lambda__5(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_2); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoCall___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) { +_start: +{ +lean_object* x_10; +x_10 = l_Lean_Meta_Monotonicity_solveMonoCall___lambda__6(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoCall___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) { +_start: +{ +lean_object* x_10; +x_10 = l_Lean_Meta_Monotonicity_solveMonoCall___lambda__7(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_Meta_Monotonicity_solveMonoCall___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: +{ +lean_object* x_13; +x_13 = l_Lean_Meta_Monotonicity_solveMonoCall___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); +lean_dec(x_2); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Monotonicity_solveMonoStep___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) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_8; lean_object* x_9; +x_8 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_8, 0, x_1); +x_9 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_9, 0, x_8); +lean_ctor_set(x_9, 1, x_7); +return x_9; +} +else +{ +uint8_t x_10; +lean_dec(x_1); +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; +x_11 = lean_box(0); +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, 1, 0); +lean_ctor_set(x_13, 0, x_12); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_7); +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; +x_15 = lean_ctor_get(x_2, 0); +lean_inc(x_15); +lean_dec(x_2); +x_16 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_16, 0, x_15); +x_17 = lean_box(0); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_16); +lean_ctor_set(x_18, 1, x_17); +x_19 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_19, 0, x_18); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_19); +lean_ctor_set(x_20, 1, x_7); +return x_20; +} +} +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Monotonicity_solveMonoStep___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) { +_start: +{ +lean_object* x_8; +x_8 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_8, 0, x_1); +lean_ctor_set(x_8, 1, x_7); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Monotonicity_solveMonoStep___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) { +_start: +{ +lean_object* x_8; lean_object* x_9; +x_8 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_8, 0, x_1); +x_9 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_9, 0, x_8); +lean_ctor_set(x_9, 1, x_7); +return x_9; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Monotonicity_solveMonoStep___spec__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Succeeded with ", 15, 15); +return x_1; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Monotonicity_solveMonoStep___spec__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Monotonicity_solveMonoStep___spec__1___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Monotonicity_solveMonoStep___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, size_t x_8, size_t x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +_start: +{ +uint8_t x_16; +x_16 = lean_usize_dec_lt(x_9, x_8); +if (x_16 == 0) +{ +lean_object* x_17; +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_2); +lean_dec(x_1); +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_10); +lean_ctor_set(x_17, 1, x_15); +return x_17; +} +else +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_28; lean_object* x_29; lean_object* x_63; +lean_dec(x_10); +x_18 = lean_array_uget(x_7, x_9); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_18); +lean_inc(x_1); +x_63 = l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_applyConst(x_1, x_18, x_11, x_12, x_13, x_14, x_15); +if (lean_obj_tag(x_63) == 0) +{ +uint8_t x_64; +x_64 = !lean_is_exclusive(x_63); +if (x_64 == 0) +{ +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_63, 0); +x_66 = lean_ctor_get(x_63, 1); +lean_inc(x_2); +x_67 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_2, x_11, x_12, x_13, x_14, x_66); +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) +{ +lean_object* x_70; 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_free_object(x_63); +lean_dec(x_18); +x_70 = lean_ctor_get(x_67, 1); +lean_inc(x_70); +lean_dec(x_67); +x_71 = lean_box(0); +x_72 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Monotonicity_solveMonoStep___spec__1___lambda__3(x_65, x_71, x_11, x_12, x_13, x_14, x_70); +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); +lean_inc(x_6); +x_75 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Monotonicity_solveMonoStep___spec__1___lambda__1(x_6, x_73, x_11, x_12, x_13, x_14, 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_19 = x_76; +x_20 = x_77; +goto block_27; +} +else +{ +uint8_t x_78; +x_78 = !lean_is_exclusive(x_67); +if (x_78 == 0) +{ +lean_object* x_79; lean_object* x_80; uint8_t x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; +x_79 = lean_ctor_get(x_67, 1); +x_80 = lean_ctor_get(x_67, 0); +lean_dec(x_80); +x_81 = 0; +x_82 = l_Lean_MessageData_ofConstName(x_18, x_81); +x_83 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Monotonicity_solveMonoStep___spec__1___closed__2; +lean_ctor_set_tag(x_67, 7); +lean_ctor_set(x_67, 1, x_82); +lean_ctor_set(x_67, 0, x_83); +x_84 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__9; +lean_ctor_set_tag(x_63, 7); +lean_ctor_set(x_63, 1, x_84); +lean_ctor_set(x_63, 0, x_67); +lean_inc(x_2); +x_85 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_2, x_63, x_11, x_12, x_13, x_14, x_79); +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_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Monotonicity_solveMonoStep___spec__1___lambda__3(x_65, x_86, x_11, x_12, x_13, x_14, x_87); +lean_dec(x_86); +x_89 = lean_ctor_get(x_88, 0); +lean_inc(x_89); +x_90 = lean_ctor_get(x_88, 1); +lean_inc(x_90); +lean_dec(x_88); +lean_inc(x_6); +x_91 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Monotonicity_solveMonoStep___spec__1___lambda__1(x_6, x_89, x_11, x_12, x_13, x_14, x_90); +x_92 = lean_ctor_get(x_91, 0); +lean_inc(x_92); +x_93 = lean_ctor_get(x_91, 1); +lean_inc(x_93); +lean_dec(x_91); +x_19 = x_92; +x_20 = x_93; +goto block_27; +} +else +{ +lean_object* x_94; uint8_t x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; +x_94 = lean_ctor_get(x_67, 1); +lean_inc(x_94); +lean_dec(x_67); +x_95 = 0; +x_96 = l_Lean_MessageData_ofConstName(x_18, x_95); +x_97 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Monotonicity_solveMonoStep___spec__1___closed__2; +x_98 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_98, 0, x_97); +lean_ctor_set(x_98, 1, x_96); +x_99 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__9; +lean_ctor_set_tag(x_63, 7); +lean_ctor_set(x_63, 1, x_99); +lean_ctor_set(x_63, 0, x_98); +lean_inc(x_2); +x_100 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_2, x_63, x_11, x_12, x_13, x_14, x_94); +x_101 = lean_ctor_get(x_100, 0); +lean_inc(x_101); +x_102 = lean_ctor_get(x_100, 1); +lean_inc(x_102); +lean_dec(x_100); +x_103 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Monotonicity_solveMonoStep___spec__1___lambda__3(x_65, x_101, x_11, x_12, x_13, x_14, x_102); +lean_dec(x_101); +x_104 = lean_ctor_get(x_103, 0); +lean_inc(x_104); +x_105 = lean_ctor_get(x_103, 1); +lean_inc(x_105); +lean_dec(x_103); +lean_inc(x_6); +x_106 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Monotonicity_solveMonoStep___spec__1___lambda__1(x_6, x_104, x_11, x_12, x_13, x_14, x_105); +x_107 = lean_ctor_get(x_106, 0); +lean_inc(x_107); +x_108 = lean_ctor_get(x_106, 1); +lean_inc(x_108); +lean_dec(x_106); +x_19 = x_107; +x_20 = x_108; +goto block_27; +} +} +} +else +{ +lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; uint8_t x_113; +x_109 = lean_ctor_get(x_63, 0); +x_110 = lean_ctor_get(x_63, 1); +lean_inc(x_110); +lean_inc(x_109); +lean_dec(x_63); +lean_inc(x_2); +x_111 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_2, x_11, x_12, x_13, x_14, x_110); +x_112 = lean_ctor_get(x_111, 0); +lean_inc(x_112); +x_113 = lean_unbox(x_112); +lean_dec(x_112); +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_dec(x_18); +x_114 = lean_ctor_get(x_111, 1); +lean_inc(x_114); +lean_dec(x_111); +x_115 = lean_box(0); +x_116 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Monotonicity_solveMonoStep___spec__1___lambda__3(x_109, x_115, x_11, x_12, x_13, x_14, x_114); +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); +lean_inc(x_6); +x_119 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Monotonicity_solveMonoStep___spec__1___lambda__1(x_6, x_117, x_11, x_12, x_13, x_14, x_118); +x_120 = lean_ctor_get(x_119, 0); +lean_inc(x_120); +x_121 = lean_ctor_get(x_119, 1); +lean_inc(x_121); +lean_dec(x_119); +x_19 = x_120; +x_20 = x_121; +goto block_27; +} +else +{ +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_122 = lean_ctor_get(x_111, 1); +lean_inc(x_122); +if (lean_is_exclusive(x_111)) { + lean_ctor_release(x_111, 0); + lean_ctor_release(x_111, 1); + x_123 = x_111; +} else { + lean_dec_ref(x_111); + x_123 = lean_box(0); +} +x_124 = 0; +x_125 = l_Lean_MessageData_ofConstName(x_18, x_124); +x_126 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Monotonicity_solveMonoStep___spec__1___closed__2; +if (lean_is_scalar(x_123)) { + x_127 = lean_alloc_ctor(7, 2, 0); +} else { + x_127 = x_123; + lean_ctor_set_tag(x_127, 7); +} +lean_ctor_set(x_127, 0, x_126); +lean_ctor_set(x_127, 1, x_125); +x_128 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__9; +x_129 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_129, 0, x_127); +lean_ctor_set(x_129, 1, x_128); +lean_inc(x_2); +x_130 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_2, x_129, x_11, x_12, x_13, x_14, x_122); +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 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Monotonicity_solveMonoStep___spec__1___lambda__3(x_109, x_131, x_11, x_12, x_13, x_14, x_132); +lean_dec(x_131); +x_134 = lean_ctor_get(x_133, 0); +lean_inc(x_134); +x_135 = lean_ctor_get(x_133, 1); +lean_inc(x_135); +lean_dec(x_133); +lean_inc(x_6); +x_136 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Monotonicity_solveMonoStep___spec__1___lambda__1(x_6, x_134, x_11, x_12, x_13, x_14, x_135); +x_137 = lean_ctor_get(x_136, 0); +lean_inc(x_137); +x_138 = lean_ctor_get(x_136, 1); +lean_inc(x_138); +lean_dec(x_136); +x_19 = x_137; +x_20 = x_138; +goto block_27; +} +} +} +else +{ +lean_object* x_139; lean_object* x_140; +lean_dec(x_18); +x_139 = lean_ctor_get(x_63, 0); +lean_inc(x_139); +x_140 = lean_ctor_get(x_63, 1); +lean_inc(x_140); +lean_dec(x_63); +x_28 = x_139; +x_29 = x_140; +goto block_62; +} +block_27: +{ +if (lean_obj_tag(x_19) == 0) +{ +lean_object* x_21; lean_object* x_22; +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_2); +lean_dec(x_1); +x_21 = lean_ctor_get(x_19, 0); +lean_inc(x_21); +lean_dec(x_19); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_20); +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_9, x_24); +x_9 = x_25; +x_10 = x_23; +x_15 = x_20; +goto _start; +} +} +block_62: +{ +uint8_t x_30; +x_30 = l_Lean_Exception_isInterrupt(x_28); +if (x_30 == 0) +{ +uint8_t x_31; +x_31 = l_Lean_Exception_isRuntime(x_28); +if (x_31 == 0) +{ +lean_object* x_32; lean_object* x_33; uint8_t x_34; +lean_inc(x_2); +x_32 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_2, x_11, x_12, x_13, x_14, x_29); +x_33 = lean_ctor_get(x_32, 0); +lean_inc(x_33); +x_34 = lean_unbox(x_33); +lean_dec(x_33); +if (x_34 == 0) +{ +lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; +lean_dec(x_28); +x_35 = lean_ctor_get(x_32, 1); +lean_inc(x_35); +lean_dec(x_32); +lean_inc(x_5); +lean_inc(x_6); +x_36 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Monotonicity_solveMonoStep___spec__1___lambda__1(x_6, x_5, x_11, x_12, x_13, x_14, x_35); +x_37 = lean_ctor_get(x_36, 0); +lean_inc(x_37); +x_38 = lean_ctor_get(x_36, 1); +lean_inc(x_38); +lean_dec(x_36); +x_19 = x_37; +x_20 = x_38; +goto block_27; +} +else +{ +uint8_t x_39; +x_39 = !lean_is_exclusive(x_32); +if (x_39 == 0) +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_40 = lean_ctor_get(x_32, 1); +x_41 = lean_ctor_get(x_32, 0); +lean_dec(x_41); +x_42 = l_Lean_Exception_toMessageData(x_28); +x_43 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__9; +lean_ctor_set_tag(x_32, 7); +lean_ctor_set(x_32, 1, x_42); +lean_ctor_set(x_32, 0, x_43); +x_44 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_44, 0, x_32); +lean_ctor_set(x_44, 1, x_43); +lean_inc(x_2); +x_45 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_2, x_44, x_11, x_12, x_13, x_14, x_40); +x_46 = lean_ctor_get(x_45, 1); +lean_inc(x_46); +lean_dec(x_45); +lean_inc(x_5); +lean_inc(x_6); +x_47 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Monotonicity_solveMonoStep___spec__1___lambda__1(x_6, x_5, x_11, x_12, x_13, x_14, x_46); +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_19 = x_48; +x_20 = x_49; +goto block_27; +} +else +{ +lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; +x_50 = lean_ctor_get(x_32, 1); +lean_inc(x_50); +lean_dec(x_32); +x_51 = l_Lean_Exception_toMessageData(x_28); +x_52 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__9; +x_53 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_53, 0, x_52); +lean_ctor_set(x_53, 1, x_51); +x_54 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_54, 0, x_53); +lean_ctor_set(x_54, 1, x_52); +lean_inc(x_2); +x_55 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_2, x_54, x_11, x_12, x_13, x_14, x_50); +x_56 = lean_ctor_get(x_55, 1); +lean_inc(x_56); +lean_dec(x_55); +lean_inc(x_5); +lean_inc(x_6); +x_57 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Monotonicity_solveMonoStep___spec__1___lambda__1(x_6, x_5, x_11, x_12, x_13, x_14, x_56); +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_19 = x_58; +x_20 = x_59; +goto block_27; +} +} +} +else +{ +lean_object* x_60; +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_2); +lean_dec(x_1); +x_60 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_60, 0, x_28); +lean_ctor_set(x_60, 1, x_29); +return x_60; +} +} +else +{ +lean_object* x_61; +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_2); +lean_dec(x_1); +x_61 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_61, 0, x_28); +lean_ctor_set(x_61, 1, x_29); +return x_61; +} +} +} +} +} +static lean_object* _init_l_Std_Range_forIn_x27_loop___at_Lean_Meta_Monotonicity_solveMonoStep___spec__2___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; +} +} +static lean_object* _init_l_Std_Range_forIn_x27_loop___at_Lean_Meta_Monotonicity_solveMonoStep___spec__2___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Monotonicity_solveMonoStep___spec__2___closed__1; +x_2 = lean_box(0); +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Monotonicity_solveMonoStep___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) { +_start: +{ +lean_object* x_14; uint8_t x_15; +x_14 = lean_ctor_get(x_4, 1); +x_15 = lean_nat_dec_lt(x_6, x_14); +if (x_15 == 0) +{ +lean_object* x_16; +lean_dec(x_6); +lean_dec(x_3); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_5); +lean_ctor_set(x_16, 1, x_13); +return x_16; +} +else +{ +lean_object* x_17; uint8_t x_18; +lean_dec(x_5); +x_17 = lean_array_get_size(x_1); +x_18 = lean_nat_dec_lt(x_6, x_17); +lean_dec(x_17); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_19 = l_Lean_instInhabitedExpr; +x_20 = l_outOfBounds___rarg(x_19); +x_21 = l_Lean_Expr_hasLooseBVars(x_20); +lean_dec(x_20); +if (x_21 == 0) +{ +lean_object* x_22; lean_object* x_23; +x_22 = lean_ctor_get(x_4, 2); +x_23 = lean_nat_add(x_6, x_22); +lean_dec(x_6); +lean_inc(x_3); +{ +lean_object* _tmp_4 = x_3; +lean_object* _tmp_5 = x_23; +lean_object* _tmp_6 = lean_box(0); +lean_object* _tmp_7 = lean_box(0); +x_5 = _tmp_4; +x_6 = _tmp_5; +x_7 = _tmp_6; +x_8 = _tmp_7; +} +goto _start; +} +else +{ +lean_object* x_25; lean_object* x_26; +lean_dec(x_6); +lean_dec(x_3); +x_25 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Monotonicity_solveMonoStep___spec__2___closed__2; +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set(x_26, 1, x_13); +return x_26; +} +} +else +{ +lean_object* x_27; uint8_t x_28; +x_27 = lean_array_fget(x_1, x_6); +x_28 = l_Lean_Expr_hasLooseBVars(x_27); +lean_dec(x_27); +if (x_28 == 0) +{ +lean_object* x_29; lean_object* x_30; +x_29 = lean_ctor_get(x_4, 2); +x_30 = lean_nat_add(x_6, x_29); +lean_dec(x_6); +lean_inc(x_3); +{ +lean_object* _tmp_4 = x_3; +lean_object* _tmp_5 = x_30; +lean_object* _tmp_6 = lean_box(0); +lean_object* _tmp_7 = lean_box(0); +x_5 = _tmp_4; +x_6 = _tmp_5; +x_7 = _tmp_6; +x_8 = _tmp_7; +} +goto _start; +} +else +{ +lean_object* x_32; lean_object* x_33; +lean_dec(x_6); +lean_dec(x_3); +x_32 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Monotonicity_solveMonoStep___spec__2___closed__2; +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_13); +return x_33; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_Monotonicity_solveMonoStep___spec__3(size_t x_1, size_t x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; +x_4 = lean_usize_dec_lt(x_2, x_1); +if (x_4 == 0) +{ +return x_3; +} +else +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; lean_object* x_9; size_t x_10; size_t x_11; lean_object* x_12; +x_5 = lean_array_uget(x_3, x_2); +x_6 = lean_unsigned_to_nat(0u); +x_7 = lean_array_uset(x_3, x_2, x_6); +x_8 = 0; +x_9 = l_Lean_MessageData_ofConstName(x_5, x_8); +x_10 = 1; +x_11 = lean_usize_add(x_2, x_10); +x_12 = lean_array_uset(x_7, x_2, x_9); +x_2 = x_11; +x_3 = x_12; +goto _start; +} +} +} +LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Meta_Monotonicity_solveMonoStep___spec__4(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; +x_5 = lean_ctor_get(x_1, 1); +lean_ctor_set(x_1, 1, x_2); +{ +lean_object* _tmp_0 = x_5; +lean_object* _tmp_1 = x_1; +x_1 = _tmp_0; +x_2 = _tmp_1; +} +goto _start; +} +else +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_7 = lean_ctor_get(x_1, 0); +x_8 = lean_ctor_get(x_1, 1); +lean_inc(x_8); +lean_inc(x_7); +lean_dec(x_1); +x_9 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_9, 0, x_7); +lean_ctor_set(x_9, 1, x_2); +x_1 = x_8; +x_2 = x_9; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; lean_object* x_9; +x_8 = lean_expr_instantiate1(x_1, x_2); +x_9 = l_Lean_Meta_Monotonicity_findMonoThms(x_8, x_3, x_4, x_5, x_6, x_7); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___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_8(x_1, lean_box(0), x_2, x_3, x_5, x_6, x_7, x_8, x_9); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +uint8_t x_7; lean_object* x_8; lean_object* x_9; +x_7 = 1; +x_8 = lean_box(x_7); +x_9 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_9, 0, x_8); +lean_ctor_set(x_9, 1, x_6); +return x_9; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__4___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_levelZero; +x_2 = l_Lean_Expr_sort___override(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__4___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Monotonicity_solveMonoStep___lambda__3___boxed), 6, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___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) { +_start: +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_13 = lean_st_ref_get(x_11, x_12); +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +x_16 = lean_ctor_get(x_14, 0); +lean_inc(x_16); +lean_dec(x_14); +x_17 = l_Lean_Meta_isMatcherAppCore_x3f(x_16, x_4); +lean_dec(x_16); +if (lean_obj_tag(x_17) == 0) +{ +lean_object* x_18; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_18 = lean_apply_8(x_1, lean_box(0), x_2, x_3, x_8, x_9, x_10, x_11, x_15); +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; lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_19 = lean_ctor_get(x_17, 0); +lean_inc(x_19); +lean_dec(x_17); +x_20 = lean_unsigned_to_nat(0u); +x_21 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_4, x_20); +x_22 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__4___closed__1; +lean_inc(x_21); +x_23 = lean_mk_array(x_21, x_22); +x_24 = lean_unsigned_to_nat(1u); +x_25 = lean_nat_sub(x_21, x_24); +lean_dec(x_21); +lean_inc(x_4); +x_26 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_4, x_23, x_25); +x_27 = l_Lean_Meta_Match_MatcherInfo_getFirstDiscrPos(x_19); +x_28 = lean_ctor_get(x_19, 1); +lean_inc(x_28); +lean_dec(x_19); +x_29 = lean_nat_add(x_27, x_28); +lean_dec(x_28); +lean_inc(x_27); +x_30 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_30, 0, x_27); +lean_ctor_set(x_30, 1, x_29); +lean_ctor_set(x_30, 2, x_24); +lean_inc(x_5); +x_31 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Monotonicity_solveMonoStep___spec__2(x_26, x_30, x_5, x_30, x_5, x_27, lean_box(0), lean_box(0), x_8, x_9, x_10, x_11, x_15); +lean_dec(x_30); +lean_dec(x_26); +x_32 = lean_ctor_get(x_31, 0); +lean_inc(x_32); +x_33 = lean_ctor_get(x_32, 0); +lean_inc(x_33); +lean_dec(x_32); +if (lean_obj_tag(x_33) == 0) +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_34 = lean_ctor_get(x_31, 1); +lean_inc(x_34); +lean_dec(x_31); +x_35 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__4___closed__2; +x_36 = lean_box(0); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +x_37 = lean_apply_6(x_35, x_36, x_8, x_9, x_10, x_11, x_34); +if (lean_obj_tag(x_37) == 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_40; lean_object* x_41; +lean_dec(x_6); +lean_dec(x_4); +x_40 = lean_ctor_get(x_37, 1); +lean_inc(x_40); +lean_dec(x_37); +x_41 = lean_apply_8(x_1, lean_box(0), x_2, x_3, x_8, x_9, x_10, x_11, x_40); +return x_41; +} +else +{ +lean_object* x_42; lean_object* x_43; +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_42 = lean_ctor_get(x_37, 1); +lean_inc(x_42); +lean_dec(x_37); +x_43 = l_Lean_Meta_Split_splitMatch(x_6, x_4, x_8, x_9, x_10, x_11, x_42); +if (lean_obj_tag(x_43) == 0) +{ +uint8_t x_44; +x_44 = !lean_is_exclusive(x_43); +if (x_44 == 0) +{ +return x_43; +} +else +{ +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(0, 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_48; +x_48 = !lean_is_exclusive(x_43); +if (x_48 == 0) +{ +return x_43; +} +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_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; +} +} +} +} +else +{ +uint8_t x_52; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_52 = !lean_is_exclusive(x_37); +if (x_52 == 0) +{ +return x_37; +} +else +{ +lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_53 = lean_ctor_get(x_37, 0); +x_54 = lean_ctor_get(x_37, 1); +lean_inc(x_54); +lean_inc(x_53); +lean_dec(x_37); +x_55 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_55, 0, x_53); +lean_ctor_set(x_55, 1, x_54); +return x_55; +} +} +} +else +{ +lean_object* x_56; uint8_t x_57; +x_56 = lean_ctor_get(x_33, 0); +lean_inc(x_56); +lean_dec(x_33); +x_57 = lean_unbox(x_56); +lean_dec(x_56); +if (x_57 == 0) +{ +lean_object* x_58; lean_object* x_59; +lean_dec(x_6); +lean_dec(x_4); +x_58 = lean_ctor_get(x_31, 1); +lean_inc(x_58); +lean_dec(x_31); +x_59 = lean_apply_8(x_1, lean_box(0), x_2, x_3, x_8, x_9, x_10, x_11, x_58); +return x_59; +} +else +{ +lean_object* x_60; lean_object* x_61; +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_60 = lean_ctor_get(x_31, 1); +lean_inc(x_60); +lean_dec(x_31); +x_61 = l_Lean_Meta_Split_splitMatch(x_6, x_4, x_8, x_9, x_10, x_11, x_60); +if (lean_obj_tag(x_61) == 0) +{ +uint8_t x_62; +x_62 = !lean_is_exclusive(x_61); +if (x_62 == 0) +{ +return x_61; +} +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(0, 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_66; +x_66 = !lean_is_exclusive(x_61); +if (x_66 == 0) +{ +return x_61; +} +else +{ +lean_object* x_67; lean_object* x_68; lean_object* x_69; +x_67 = lean_ctor_get(x_61, 0); +x_68 = lean_ctor_get(x_61, 1); +lean_inc(x_68); +lean_inc(x_67); +lean_dec(x_61); +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; +} +} +} +} +} +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__5___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = lean_box(0); +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___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) { +_start: +{ +lean_object* x_13; lean_object* x_14; size_t x_15; size_t x_16; lean_object* x_17; lean_object* x_18; +x_13 = lean_box(0); +x_14 = lean_box(0); +x_15 = lean_array_size(x_1); +x_16 = 0; +x_17 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__5___closed__1; +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_2); +x_18 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Monotonicity_solveMonoStep___spec__1(x_2, x_3, x_1, x_13, x_14, x_17, x_1, x_15, x_16, x_17, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_18) == 0) +{ +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_19, 0); +lean_inc(x_20); +lean_dec(x_19); +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_18, 1); +lean_inc(x_21); +lean_dec(x_18); +x_22 = lean_box(0); +x_23 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__4(x_4, x_5, x_1, x_6, x_17, x_2, x_22, x_8, x_9, x_10, x_11, x_21); +return x_23; +} +else +{ +uint8_t x_24; +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_2); +lean_dec(x_1); +x_24 = !lean_is_exclusive(x_18); +if (x_24 == 0) +{ +lean_object* x_25; lean_object* x_26; +x_25 = lean_ctor_get(x_18, 0); +lean_dec(x_25); +x_26 = lean_ctor_get(x_20, 0); +lean_inc(x_26); +lean_dec(x_20); +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_ctor_get(x_18, 1); +lean_inc(x_27); +lean_dec(x_18); +x_28 = lean_ctor_get(x_20, 0); +lean_inc(x_28); +lean_dec(x_20); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_27); +return x_29; +} +} +} +else +{ +uint8_t x_30; +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_2); +lean_dec(x_1); +x_30 = !lean_is_exclusive(x_18); +if (x_30 == 0) +{ +return x_18; +} +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_18, 0); +x_32 = lean_ctor_get(x_18, 1); +lean_inc(x_32); +lean_inc(x_31); +lean_dec(x_18); +x_33 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_33, 0, x_31); +lean_ctor_set(x_33, 1, x_32); +return x_33; +} +} +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__6___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("f", 1, 1); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_solveMonoStep___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_Meta_Monotonicity_solveMonoStep___lambda__6___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__6___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Found monoThms: ", 16, 16); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__6___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__6___closed__3; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___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; lean_object* x_13; lean_object* x_14; uint8_t x_15; uint8_t x_16; lean_object* x_17; +x_12 = l_Lean_Expr_bindingDomain_x21(x_1); +lean_inc(x_2); +x_13 = lean_alloc_closure((void*)(l_Lean_Meta_Monotonicity_solveMonoStep___lambda__1___boxed), 7, 1); +lean_closure_set(x_13, 0, x_2); +x_14 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__6___closed__2; +x_15 = 0; +x_16 = 0; +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +x_17 = l_Lean_Meta_withLocalDecl___at_Lean_Meta_addPPExplicitToExposeDiff_visit___spec__4___rarg(x_14, x_15, x_12, x_13, x_16, 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; lean_object* x_20; lean_object* x_21; uint8_t 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); +lean_inc(x_3); +x_20 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_3, x_7, x_8, x_9, x_10, x_19); +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; +x_23 = lean_ctor_get(x_20, 1); +lean_inc(x_23); +lean_dec(x_20); +x_24 = lean_box(0); +x_25 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__5(x_18, x_4, x_3, x_5, x_1, x_2, x_24, x_7, x_8, x_9, x_10, x_23); +return x_25; +} +else +{ +uint8_t x_26; +x_26 = !lean_is_exclusive(x_20); +if (x_26 == 0) +{ +lean_object* x_27; lean_object* x_28; size_t x_29; size_t x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_27 = lean_ctor_get(x_20, 1); +x_28 = lean_ctor_get(x_20, 0); +lean_dec(x_28); +x_29 = lean_array_size(x_18); +x_30 = 0; +lean_inc(x_18); +x_31 = l_Array_mapMUnsafe_map___at_Lean_Meta_Monotonicity_solveMonoStep___spec__3(x_29, x_30, x_18); +x_32 = lean_array_to_list(x_31); +x_33 = lean_box(0); +x_34 = l_List_mapTR_loop___at_Lean_Meta_Monotonicity_solveMonoStep___spec__4(x_32, x_33); +x_35 = l_Lean_MessageData_ofList(x_34); +x_36 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__6___closed__4; +lean_ctor_set_tag(x_20, 7); +lean_ctor_set(x_20, 1, x_35); +lean_ctor_set(x_20, 0, x_36); +x_37 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__9; +x_38 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_38, 0, x_20); +lean_ctor_set(x_38, 1, x_37); +lean_inc(x_3); +x_39 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_3, x_38, x_7, x_8, x_9, x_10, x_27); +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_Meta_Monotonicity_solveMonoStep___lambda__5(x_18, x_4, x_3, x_5, x_1, x_2, x_40, x_7, x_8, x_9, x_10, x_41); +lean_dec(x_40); +return x_42; +} +else +{ +lean_object* x_43; size_t x_44; size_t x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; +x_43 = lean_ctor_get(x_20, 1); +lean_inc(x_43); +lean_dec(x_20); +x_44 = lean_array_size(x_18); +x_45 = 0; +lean_inc(x_18); +x_46 = l_Array_mapMUnsafe_map___at_Lean_Meta_Monotonicity_solveMonoStep___spec__3(x_44, x_45, x_18); +x_47 = lean_array_to_list(x_46); +x_48 = lean_box(0); +x_49 = l_List_mapTR_loop___at_Lean_Meta_Monotonicity_solveMonoStep___spec__4(x_47, x_48); +x_50 = l_Lean_MessageData_ofList(x_49); +x_51 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__6___closed__4; +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_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__9; +x_54 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_54, 0, x_52); +lean_ctor_set(x_54, 1, x_53); +lean_inc(x_3); +x_55 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_3, x_54, x_7, x_8, x_9, x_10, x_43); +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_Meta_Monotonicity_solveMonoStep___lambda__5(x_18, x_4, x_3, x_5, x_1, x_2, x_56, x_7, x_8, x_9, x_10, x_57); +lean_dec(x_56); +return x_58; +} +} +} +else +{ +uint8_t x_59; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_59 = !lean_is_exclusive(x_17); +if (x_59 == 0) +{ +return x_17; +} +else +{ +lean_object* x_60; lean_object* x_61; lean_object* x_62; +x_60 = lean_ctor_get(x_17, 0); +x_61 = lean_ctor_get(x_17, 1); +lean_inc(x_61); +lean_inc(x_60); +lean_dec(x_17); +x_62 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_62, 0, x_60); +lean_ctor_set(x_62, 1, x_61); +return x_62; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___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) { +_start: +{ +lean_object* x_7; lean_object* x_8; +x_7 = lean_box(0); +x_8 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_8, 0, x_7); +lean_ctor_set(x_8, 1, x_6); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___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) { +_start: +{ +lean_object* x_10; lean_object* x_11; +x_10 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__9___closed__1; +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +x_11 = lean_apply_8(x_1, lean_box(0), x_2, x_10, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_11) == 0) +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; +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_6(x_3, x_12, x_5, x_6, x_7, x_8, x_13); +return x_14; +} +else +{ +uint8_t x_15; +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +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; +} +} +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__9___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Monotonicity_solveMonoStep___lambda__7___boxed), 6, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__9___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Failed to assign ", 17, 17); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__9___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__9___closed__2; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__9___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(" : ", 3, 3); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__9___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__9___closed__4; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__9___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(" to goal", 8, 8); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__9___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__9___closed__6; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__9(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_2); +x_12 = lean_checked_assign(x_1, x_2, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_12) == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_12, 1); +lean_inc(x_14); +lean_dec(x_12); +x_15 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__9___closed__1; +x_16 = lean_unbox(x_13); +lean_dec(x_13); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; uint8_t x_19; +lean_inc(x_3); +x_17 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_3, x_7, x_8, x_9, x_10, x_14); +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +x_19 = lean_unbox(x_18); +lean_dec(x_18); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; +lean_dec(x_3); +lean_dec(x_2); +x_20 = lean_ctor_get(x_17, 1); +lean_inc(x_20); +lean_dec(x_17); +x_21 = lean_box(0); +x_22 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__8(x_4, x_5, x_15, x_21, x_7, x_8, x_9, x_10, x_20); +return x_22; +} +else +{ +uint8_t x_23; +x_23 = !lean_is_exclusive(x_17); +if (x_23 == 0) +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_17, 1); +x_25 = lean_ctor_get(x_17, 0); +lean_dec(x_25); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_2); +x_26 = lean_infer_type(x_2, x_7, x_8, x_9, x_10, x_24); +if (lean_obj_tag(x_26) == 0) +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; 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_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_MessageData_ofExpr(x_2); +x_30 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__9___closed__3; +lean_ctor_set_tag(x_17, 7); +lean_ctor_set(x_17, 1, x_29); +lean_ctor_set(x_17, 0, x_30); +x_31 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__9___closed__5; +x_32 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_32, 0, x_17); +lean_ctor_set(x_32, 1, x_31); +x_33 = l_Lean_MessageData_ofExpr(x_27); +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_Meta_Monotonicity_solveMonoStep___lambda__9___closed__7; +x_36 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_36, 0, x_34); +lean_ctor_set(x_36, 1, x_35); +x_37 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_3, x_36, x_7, x_8, x_9, x_10, x_28); +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_Monotonicity_solveMonoStep___lambda__8(x_4, x_5, x_15, x_38, x_7, x_8, x_9, x_10, x_39); +lean_dec(x_38); +return x_40; +} +else +{ +uint8_t x_41; +lean_free_object(x_17); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_41 = !lean_is_exclusive(x_26); +if (x_41 == 0) +{ +return x_26; +} +else +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_42 = lean_ctor_get(x_26, 0); +x_43 = lean_ctor_get(x_26, 1); +lean_inc(x_43); +lean_inc(x_42); +lean_dec(x_26); +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_45; lean_object* x_46; +x_45 = lean_ctor_get(x_17, 1); +lean_inc(x_45); +lean_dec(x_17); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_2); +x_46 = lean_infer_type(x_2, x_7, x_8, x_9, x_10, x_45); +if (lean_obj_tag(x_46) == 0) +{ +lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; +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_MessageData_ofExpr(x_2); +x_50 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__9___closed__3; +x_51 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_51, 0, x_50); +lean_ctor_set(x_51, 1, x_49); +x_52 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__9___closed__5; +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_Lean_MessageData_ofExpr(x_47); +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_Meta_Monotonicity_solveMonoStep___lambda__9___closed__7; +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_Meta_processPostponed_loop___spec__2(x_3, x_57, x_7, x_8, x_9, x_10, x_48); +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_Monotonicity_solveMonoStep___lambda__8(x_4, x_5, x_15, x_59, x_7, x_8, x_9, x_10, x_60); +lean_dec(x_59); +return x_61; +} +else +{ +lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_62 = lean_ctor_get(x_46, 0); +lean_inc(x_62); +x_63 = lean_ctor_get(x_46, 1); +lean_inc(x_63); +if (lean_is_exclusive(x_46)) { + lean_ctor_release(x_46, 0); + lean_ctor_release(x_46, 1); + x_64 = x_46; +} else { + lean_dec_ref(x_46); + x_64 = lean_box(0); +} +if (lean_is_scalar(x_64)) { + x_65 = lean_alloc_ctor(1, 2, 0); +} else { + x_65 = x_64; +} +lean_ctor_set(x_65, 0, x_62); +lean_ctor_set(x_65, 1, x_63); +return x_65; +} +} +} +} +else +{ +lean_object* x_66; lean_object* x_67; +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_66 = lean_box(0); +x_67 = lean_apply_6(x_15, x_66, x_7, x_8, x_9, x_10, x_14); +return x_67; +} +} +else +{ +uint8_t x_68; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_68 = !lean_is_exclusive(x_12); +if (x_68 == 0) +{ +return x_12; +} +else +{ +lean_object* x_69; lean_object* x_70; lean_object* x_71; +x_69 = lean_ctor_get(x_12, 0); +x_70 = lean_ctor_get(x_12, 1); +lean_inc(x_70); +lean_inc(x_69); +lean_dec(x_12); +x_71 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_71, 0, x_69); +lean_ctor_set(x_71, 1, x_70); +return x_71; +} +} +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__10___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Found recursive call ", 21, 21); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__10___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__10___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___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); +lean_inc(x_3); +x_14 = l_Lean_Meta_Monotonicity_solveMonoCall(x_1, x_2, x_3, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +if (lean_obj_tag(x_15) == 0) +{ +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_Meta_Monotonicity_solveMonoStep___lambda__6(x_4, x_3, x_5, x_6, x_7, x_17, x_9, x_10, x_11, x_12, x_16); +return x_18; +} +else +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; +x_19 = lean_ctor_get(x_14, 1); +lean_inc(x_19); +lean_dec(x_14); +x_20 = lean_ctor_get(x_15, 0); +lean_inc(x_20); +lean_dec(x_15); +lean_inc(x_5); +x_21 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_5, x_9, x_10, x_11, x_12, x_19); +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_24; lean_object* x_25; lean_object* x_26; +lean_dec(x_3); +x_24 = lean_ctor_get(x_21, 1); +lean_inc(x_24); +lean_dec(x_21); +x_25 = lean_box(0); +x_26 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__9(x_6, x_20, x_5, x_7, x_4, x_25, x_9, x_10, x_11, x_12, x_24); +return x_26; +} +else +{ +uint8_t x_27; +x_27 = !lean_is_exclusive(x_21); +if (x_27 == 0) +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_28 = lean_ctor_get(x_21, 1); +x_29 = lean_ctor_get(x_21, 0); +lean_dec(x_29); +x_30 = l_Lean_MessageData_ofExpr(x_3); +x_31 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__10___closed__2; +lean_ctor_set_tag(x_21, 7); +lean_ctor_set(x_21, 1, x_30); +lean_ctor_set(x_21, 0, x_31); +x_32 = l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_applyConst___lambda__1___closed__4; +x_33 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_33, 0, x_21); +lean_ctor_set(x_33, 1, x_32); +lean_inc(x_20); +x_34 = l_Lean_indentExpr(x_20); +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_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__9; +x_37 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_37, 0, x_35); +lean_ctor_set(x_37, 1, x_36); +lean_inc(x_5); +x_38 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_5, x_37, x_9, x_10, x_11, x_12, x_28); +x_39 = lean_ctor_get(x_38, 0); +lean_inc(x_39); +x_40 = lean_ctor_get(x_38, 1); +lean_inc(x_40); +lean_dec(x_38); +x_41 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__9(x_6, x_20, x_5, x_7, x_4, x_39, x_9, x_10, x_11, x_12, x_40); +lean_dec(x_39); +return x_41; +} +else +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_42 = lean_ctor_get(x_21, 1); +lean_inc(x_42); +lean_dec(x_21); +x_43 = l_Lean_MessageData_ofExpr(x_3); +x_44 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__10___closed__2; +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___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_applyConst___lambda__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); +lean_inc(x_20); +x_48 = l_Lean_indentExpr(x_20); +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_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__9; +x_51 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_51, 0, x_49); +lean_ctor_set(x_51, 1, x_50); +lean_inc(x_5); +x_52 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_5, x_51, x_9, x_10, x_11, x_12, 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); +x_55 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__9(x_6, x_20, x_5, x_7, x_4, x_53, x_9, x_10, x_11, x_12, x_54); +lean_dec(x_53); +return x_55; +} +} +} +} +else +{ +uint8_t x_56; +lean_dec(x_12); +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); +x_56 = !lean_is_exclusive(x_14); +if (x_56 == 0) +{ +return x_14; +} +else +{ +lean_object* x_57; lean_object* x_58; lean_object* x_59; +x_57 = lean_ctor_get(x_14, 0); +x_58 = lean_ctor_get(x_14, 1); +lean_inc(x_58); +lean_inc(x_57); +lean_dec(x_14); +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_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__11(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +uint8_t x_14; +x_14 = l_Lean_Expr_isBVar(x_3); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; +x_15 = lean_box(0); +x_16 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__10(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_15, x_9, x_10, x_11, x_12, x_13); +return x_16; +} +else +{ +lean_object* x_17; lean_object* x_18; +lean_dec(x_7); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_17 = l_Lean_Meta_Monotonicity_solveMonoCall___lambda__2___closed__4; +x_18 = l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_applyConst(x_6, x_17, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_18) == 0) +{ +uint8_t x_19; +x_19 = !lean_is_exclusive(x_18); +if (x_19 == 0) +{ +return x_18; +} +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_18, 0); +x_21 = lean_ctor_get(x_18, 1); +lean_inc(x_21); +lean_inc(x_20); +lean_dec(x_18); +x_22 = lean_alloc_ctor(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_23; +x_23 = !lean_is_exclusive(x_18); +if (x_23 == 0) +{ +return x_18; +} +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_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; +} +} +} +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__12___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("monotone_of_monotone_apply", 26, 26); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__12___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__1; +x_2 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__3; +x_3 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__12___closed__1; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__12___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Unexpected number of goals after ", 33, 33); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__12___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__12___closed__3; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__12___closed__5() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__12___closed__2; +x_2 = 0; +x_3 = l_Lean_MessageData_ofConstName(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__12___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__12___closed__4; +x_2 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__12___closed__5; +x_3 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__12___closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(".", 1, 1); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__12___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__12___closed__7; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__12___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__12___closed__6; +x_2 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__12___closed__8; +x_3 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__12(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +uint8_t x_14; +x_14 = l_Lean_Expr_isLambda(x_3); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; +x_15 = lean_box(0); +x_16 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__11(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_15, x_9, x_10, x_11, x_12, x_13); +return x_16; +} +else +{ +lean_object* x_17; lean_object* x_18; +lean_dec(x_7); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_17 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__12___closed__2; +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +x_18 = l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_applyConst(x_6, x_17, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_18) == 0) +{ +lean_object* x_19; +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +if (lean_obj_tag(x_19) == 0) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; +lean_dec(x_3); +x_20 = lean_ctor_get(x_18, 1); +lean_inc(x_20); +lean_dec(x_18); +x_21 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__12___closed__9; +x_22 = l_Lean_throwError___at___private_Lean_Meta_RecursorInfo_0__Lean_Meta_getMajorPosDepElim___spec__3(x_21, x_9, x_10, x_11, x_12, x_20); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +x_23 = !lean_is_exclusive(x_22); +if (x_23 == 0) +{ +return x_22; +} +else +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_22, 0); +x_25 = lean_ctor_get(x_22, 1); +lean_inc(x_25); +lean_inc(x_24); +lean_dec(x_22); +x_26 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_26, 0, x_24); +lean_ctor_set(x_26, 1, x_25); +return x_26; +} +} +else +{ +lean_object* x_27; +x_27 = lean_ctor_get(x_19, 1); +lean_inc(x_27); +if (lean_obj_tag(x_27) == 0) +{ +lean_object* x_28; uint8_t x_29; +x_28 = lean_ctor_get(x_18, 1); +lean_inc(x_28); +lean_dec(x_18); +x_29 = !lean_is_exclusive(x_19); +if (x_29 == 0) +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_30 = lean_ctor_get(x_19, 0); +x_31 = lean_ctor_get(x_19, 1); +lean_dec(x_31); +x_32 = l_Lean_Expr_bindingName_x21(x_3); +lean_dec(x_3); +x_33 = l_Lean_MVarId_intro(x_30, x_32, x_9, x_10, x_11, x_12, 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; lean_object* x_36; lean_object* x_37; +x_35 = lean_ctor_get(x_33, 0); +x_36 = lean_ctor_get(x_35, 1); +lean_inc(x_36); +lean_dec(x_35); +x_37 = lean_box(0); +lean_ctor_set(x_19, 1, x_37); +lean_ctor_set(x_19, 0, x_36); +lean_ctor_set(x_33, 0, x_19); +return x_33; +} +else +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_38 = lean_ctor_get(x_33, 0); +x_39 = lean_ctor_get(x_33, 1); +lean_inc(x_39); +lean_inc(x_38); +lean_dec(x_33); +x_40 = lean_ctor_get(x_38, 1); +lean_inc(x_40); +lean_dec(x_38); +x_41 = lean_box(0); +lean_ctor_set(x_19, 1, x_41); +lean_ctor_set(x_19, 0, x_40); +x_42 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_42, 0, x_19); +lean_ctor_set(x_42, 1, x_39); +return x_42; +} +} +else +{ +uint8_t x_43; +lean_free_object(x_19); +x_43 = !lean_is_exclusive(x_33); +if (x_43 == 0) +{ +return x_33; +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_33, 0); +x_45 = lean_ctor_get(x_33, 1); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_33); +x_46 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_46, 0, x_44); +lean_ctor_set(x_46, 1, x_45); +return x_46; +} +} +} +else +{ +lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_47 = lean_ctor_get(x_19, 0); +lean_inc(x_47); +lean_dec(x_19); +x_48 = l_Lean_Expr_bindingName_x21(x_3); +lean_dec(x_3); +x_49 = l_Lean_MVarId_intro(x_47, x_48, x_9, x_10, x_11, x_12, x_28); +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; +x_50 = lean_ctor_get(x_49, 0); +lean_inc(x_50); +x_51 = lean_ctor_get(x_49, 1); +lean_inc(x_51); +if (lean_is_exclusive(x_49)) { + lean_ctor_release(x_49, 0); + lean_ctor_release(x_49, 1); + x_52 = x_49; +} else { + lean_dec_ref(x_49); + x_52 = lean_box(0); +} +x_53 = lean_ctor_get(x_50, 1); +lean_inc(x_53); +lean_dec(x_50); +x_54 = lean_box(0); +x_55 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_55, 0, x_53); +lean_ctor_set(x_55, 1, x_54); +if (lean_is_scalar(x_52)) { + x_56 = lean_alloc_ctor(0, 2, 0); +} else { + x_56 = x_52; +} +lean_ctor_set(x_56, 0, x_55); +lean_ctor_set(x_56, 1, x_51); +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_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; +} +} +} +else +{ +lean_object* x_61; lean_object* x_62; lean_object* x_63; uint8_t x_64; +lean_dec(x_27); +lean_dec(x_19); +lean_dec(x_3); +x_61 = lean_ctor_get(x_18, 1); +lean_inc(x_61); +lean_dec(x_18); +x_62 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__12___closed__9; +x_63 = l_Lean_throwError___at___private_Lean_Meta_RecursorInfo_0__Lean_Meta_getMajorPosDepElim___spec__3(x_62, x_9, x_10, x_11, x_12, x_61); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +x_64 = !lean_is_exclusive(x_63); +if (x_64 == 0) +{ +return x_63; +} +else +{ +lean_object* x_65; lean_object* x_66; lean_object* x_67; +x_65 = lean_ctor_get(x_63, 0); +x_66 = lean_ctor_get(x_63, 1); +lean_inc(x_66); +lean_inc(x_65); +lean_dec(x_63); +x_67 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_67, 0, x_65); +lean_ctor_set(x_67, 1, x_66); +return x_67; +} +} +} +} +else +{ +uint8_t x_68; +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_3); +x_68 = !lean_is_exclusive(x_18); +if (x_68 == 0) +{ +return x_18; +} +else +{ +lean_object* x_69; lean_object* x_70; lean_object* x_71; +x_69 = lean_ctor_get(x_18, 0); +x_70 = lean_ctor_get(x_18, 1); +lean_inc(x_70); +lean_inc(x_69); +lean_dec(x_18); +x_71 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_71, 0, x_69); +lean_ctor_set(x_71, 1, x_70); +return x_71; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__13(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; +x_3 = l_Lean_MessageData_ofExpr(x_1); +x_4 = l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_applyConst___lambda__1___closed__2; +x_5 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_5, 0, x_4); +lean_ctor_set(x_5, 1, x_3); +x_6 = l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_applyConst___lambda__1___closed__4; +x_7 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_7, 0, x_5); +lean_ctor_set(x_7, 1, x_6); +x_8 = l_Lean_indentD(x_2); +x_9 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_9, 0, x_7); +lean_ctor_set(x_9, 1, x_8); +x_10 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__9; +x_11 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_11, 0, x_9); +lean_ctor_set(x_11, 1, x_10); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__14(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +uint8_t x_8; +x_8 = !lean_is_exclusive(x_2); +if (x_8 == 0) +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_9 = lean_ctor_get(x_2, 1); +x_10 = lean_ctor_get(x_2, 0); +lean_dec(x_10); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 1, x_1); +lean_ctor_set(x_2, 0, x_9); +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_2); +lean_ctor_set(x_11, 1, x_7); +return x_11; +} +else +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_12 = lean_ctor_get(x_2, 1); +lean_inc(x_12); +lean_dec(x_2); +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_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_7); +return x_14; +} +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__15___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("monotone_letFun", 15, 15); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__15___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__1; +x_2 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__3; +x_3 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__15___closed__1; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__15___closed__3() { +_start: +{ +uint8_t x_1; uint8_t x_2; uint8_t x_3; lean_object* x_4; +x_1 = 0; +x_2 = 1; +x_3 = 0; +x_4 = lean_alloc_ctor(0, 0, 4); +lean_ctor_set_uint8(x_4, 0, x_1); +lean_ctor_set_uint8(x_4, 1, x_2); +lean_ctor_set_uint8(x_4, 2, x_3); +lean_ctor_set_uint8(x_4, 3, x_2); +return x_4; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__15___closed__4() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__15___closed__2; +x_2 = 0; +x_3 = l_Lean_MessageData_ofConstName(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__15___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__12___closed__4; +x_2 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__15___closed__4; +x_3 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__15___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__15___closed__5; +x_2 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__12___closed__8; +x_3 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__15___closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.Expr.updateLambdaE!", 24, 24); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__15___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; +x_1 = l_Lean_Loop_forIn_loop___at_Lean_Meta_Monotonicity_headBetaUnderLambda___spec__1___closed__1; +x_2 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__15___closed__7; +x_3 = lean_unsigned_to_nat(1815u); +x_4 = lean_unsigned_to_nat(20u); +x_5 = l_Lean_Loop_forIn_loop___at_Lean_Meta_Monotonicity_headBetaUnderLambda___spec__1___closed__3; +x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__15(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +_start: +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_17 = l_Lean_Expr_bindingDomain_x21(x_1); +x_18 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_18, 0, x_2); +x_19 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_19, 0, x_3); +x_20 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_20, 0, x_4); +x_21 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_21, 0, x_5); +x_22 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_22, 0, x_6); +x_23 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_23, 0, x_7); +x_24 = lean_box(0); +if (lean_obj_tag(x_1) == 6) +{ +lean_object* x_89; lean_object* x_90; lean_object* x_91; uint8_t x_92; lean_object* x_93; size_t x_94; size_t x_95; uint8_t x_96; +x_89 = lean_ctor_get(x_1, 0); +lean_inc(x_89); +x_90 = lean_ctor_get(x_1, 1); +lean_inc(x_90); +x_91 = lean_ctor_get(x_1, 2); +lean_inc(x_91); +x_92 = lean_ctor_get_uint8(x_1, sizeof(void*)*3 + 8); +lean_dec(x_1); +lean_inc(x_91); +lean_inc(x_90); +lean_inc(x_89); +x_93 = l_Lean_Expr_lam___override(x_89, x_90, x_91, x_92); +x_94 = lean_ptr_addr(x_90); +lean_dec(x_90); +x_95 = lean_ptr_addr(x_17); +x_96 = lean_usize_dec_eq(x_94, x_95); +if (x_96 == 0) +{ +lean_object* x_97; +lean_dec(x_93); +lean_dec(x_91); +lean_inc(x_10); +x_97 = l_Lean_Expr_lam___override(x_89, x_17, x_10, x_92); +x_25 = x_97; +goto block_88; +} +else +{ +size_t x_98; size_t x_99; uint8_t x_100; +x_98 = lean_ptr_addr(x_91); +lean_dec(x_91); +x_99 = lean_ptr_addr(x_10); +x_100 = lean_usize_dec_eq(x_98, x_99); +if (x_100 == 0) +{ +lean_object* x_101; +lean_dec(x_93); +lean_inc(x_10); +x_101 = l_Lean_Expr_lam___override(x_89, x_17, x_10, x_92); +x_25 = x_101; +goto block_88; +} +else +{ +uint8_t x_102; +x_102 = l___private_Lean_Expr_0__Lean_beqBinderInfo____x40_Lean_Expr___hyg_406_(x_92, x_92); +if (x_102 == 0) +{ +lean_object* x_103; +lean_dec(x_93); +lean_inc(x_10); +x_103 = l_Lean_Expr_lam___override(x_89, x_17, x_10, x_92); +x_25 = x_103; +goto block_88; +} +else +{ +lean_dec(x_89); +lean_dec(x_17); +x_25 = x_93; +goto block_88; +} +} +} +} +else +{ +lean_object* x_104; lean_object* x_105; +lean_dec(x_17); +lean_dec(x_1); +x_104 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__15___closed__8; +x_105 = l_panic___at_Lean_Expr_appFn_x21___spec__1(x_104); +x_25 = x_105; +goto block_88; +} +block_88: +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; 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 = lean_alloc_ctor(1, 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_24); +x_28 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_28, 0, x_23); +lean_ctor_set(x_28, 1, x_27); +x_29 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_29, 0, x_22); +lean_ctor_set(x_29, 1, x_28); +x_30 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_30, 0, x_21); +lean_ctor_set(x_30, 1, x_29); +x_31 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_31, 0, x_20); +lean_ctor_set(x_31, 1, x_30); +x_32 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_32, 0, x_19); +lean_ctor_set(x_32, 1, x_31); +x_33 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_33, 0, x_18); +lean_ctor_set(x_33, 1, x_32); +x_34 = lean_array_mk(x_33); +x_35 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__15___closed__2; +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +x_36 = l_Lean_Meta_mkAppOptM(x_35, x_34, x_12, x_13, x_14, x_15, x_16); +if (lean_obj_tag(x_36) == 0) +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_37 = lean_ctor_get(x_36, 0); +lean_inc(x_37); +x_38 = lean_ctor_get(x_36, 1); +lean_inc(x_38); +lean_dec(x_36); +x_39 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__15___closed__3; +lean_inc(x_37); +x_40 = lean_alloc_closure((void*)(l_Lean_MVarId_apply), 8, 3); +lean_closure_set(x_40, 0, x_8); +lean_closure_set(x_40, 1, x_37); +lean_closure_set(x_40, 2, x_39); +x_41 = lean_alloc_closure((void*)(l_Lean_Meta_Monotonicity_solveMonoStep___lambda__13), 2, 1); +lean_closure_set(x_41, 0, x_37); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +x_42 = l_Lean_Meta_mapErrorImp___rarg(x_40, x_41, x_12, x_13, x_14, x_15, x_38); +if (lean_obj_tag(x_42) == 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_44; lean_object* x_45; lean_object* x_46; uint8_t x_47; +lean_dec(x_10); +x_44 = lean_ctor_get(x_42, 1); +lean_inc(x_44); +lean_dec(x_42); +x_45 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__15___closed__6; +x_46 = l_Lean_throwError___at_Lean_Meta_setInlineAttribute___spec__1(x_45, x_12, x_13, x_14, x_15, x_44); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +x_47 = !lean_is_exclusive(x_46); +if (x_47 == 0) +{ +return x_46; +} +else +{ +lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_48 = lean_ctor_get(x_46, 0); +x_49 = lean_ctor_get(x_46, 1); +lean_inc(x_49); +lean_inc(x_48); +lean_dec(x_46); +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; +x_51 = lean_ctor_get(x_43, 1); +lean_inc(x_51); +if (lean_obj_tag(x_51) == 0) +{ +lean_object* x_52; lean_object* x_53; uint8_t x_54; +x_52 = lean_ctor_get(x_42, 1); +lean_inc(x_52); +lean_dec(x_42); +x_53 = lean_ctor_get(x_43, 0); +lean_inc(x_53); +lean_dec(x_43); +x_54 = l_Lean_Expr_isLambda(x_10); +if (x_54 == 0) +{ +uint8_t x_55; lean_object* x_56; +lean_dec(x_10); +x_55 = 0; +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +x_56 = l_Lean_Meta_intro1Core(x_53, x_55, x_12, x_13, x_14, x_15, x_52); +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_56, 1); +lean_inc(x_58); +lean_dec(x_56); +x_59 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__14(x_24, x_57, x_12, x_13, x_14, x_15, x_58); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +return x_59; +} +else +{ +uint8_t x_60; +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +x_60 = !lean_is_exclusive(x_56); +if (x_60 == 0) +{ +return x_56; +} +else +{ +lean_object* x_61; lean_object* x_62; lean_object* x_63; +x_61 = lean_ctor_get(x_56, 0); +x_62 = lean_ctor_get(x_56, 1); +lean_inc(x_62); +lean_inc(x_61); +lean_dec(x_56); +x_63 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_63, 0, x_61); +lean_ctor_set(x_63, 1, x_62); +return x_63; +} +} +} +else +{ +lean_object* x_64; lean_object* x_65; +x_64 = l_Lean_Expr_bindingName_x21(x_10); +lean_dec(x_10); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +x_65 = l_Lean_MVarId_intro(x_53, x_64, x_12, x_13, x_14, x_15, x_52); +if (lean_obj_tag(x_65) == 0) +{ +lean_object* x_66; lean_object* x_67; lean_object* x_68; +x_66 = lean_ctor_get(x_65, 0); +lean_inc(x_66); +x_67 = lean_ctor_get(x_65, 1); +lean_inc(x_67); +lean_dec(x_65); +x_68 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__14(x_24, x_66, x_12, x_13, x_14, x_15, x_67); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +return x_68; +} +else +{ +uint8_t x_69; +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +x_69 = !lean_is_exclusive(x_65); +if (x_69 == 0) +{ +return x_65; +} +else +{ +lean_object* x_70; lean_object* x_71; lean_object* x_72; +x_70 = lean_ctor_get(x_65, 0); +x_71 = lean_ctor_get(x_65, 1); +lean_inc(x_71); +lean_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 +{ +lean_object* x_73; lean_object* x_74; lean_object* x_75; uint8_t x_76; +lean_dec(x_51); +lean_dec(x_43); +lean_dec(x_10); +x_73 = lean_ctor_get(x_42, 1); +lean_inc(x_73); +lean_dec(x_42); +x_74 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__15___closed__6; +x_75 = l_Lean_throwError___at_Lean_Meta_setInlineAttribute___spec__1(x_74, x_12, x_13, x_14, x_15, x_73); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +x_76 = !lean_is_exclusive(x_75); +if (x_76 == 0) +{ +return x_75; +} +else +{ +lean_object* x_77; lean_object* x_78; lean_object* x_79; +x_77 = lean_ctor_get(x_75, 0); +x_78 = lean_ctor_get(x_75, 1); +lean_inc(x_78); +lean_inc(x_77); +lean_dec(x_75); +x_79 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_79, 0, x_77); +lean_ctor_set(x_79, 1, x_78); +return x_79; +} +} +} +} +else +{ +uint8_t x_80; +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_10); +x_80 = !lean_is_exclusive(x_42); +if (x_80 == 0) +{ +return x_42; +} +else +{ +lean_object* x_81; lean_object* x_82; lean_object* x_83; +x_81 = lean_ctor_get(x_42, 0); +x_82 = lean_ctor_get(x_42, 1); +lean_inc(x_82); +lean_inc(x_81); +lean_dec(x_42); +x_83 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_83, 0, x_81); +lean_ctor_set(x_83, 1, x_82); +return x_83; +} +} +} +else +{ +uint8_t x_84; +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_10); +lean_dec(x_8); +x_84 = !lean_is_exclusive(x_36); +if (x_84 == 0) +{ +return x_36; +} +else +{ +lean_object* x_85; lean_object* x_86; lean_object* x_87; +x_85 = lean_ctor_get(x_36, 0); +x_86 = lean_ctor_get(x_36, 1); +lean_inc(x_86); +lean_inc(x_85); +lean_dec(x_36); +x_87 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_87, 0, x_85); +lean_ctor_set(x_87, 1, x_86); +return x_87; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__16(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +_start: +{ +uint8_t x_17; +x_17 = l_Lean_Expr_hasLooseBVars(x_9); +if (x_17 == 0) +{ +uint8_t x_18; +x_18 = l_Lean_Expr_hasLooseBVars(x_10); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; +lean_dec(x_8); +x_19 = lean_box(0); +x_20 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__15(x_1, x_2, x_3, x_9, x_4, x_5, x_10, x_6, x_7, x_11, x_19, x_12, x_13, x_14, x_15, x_16); +return x_20; +} +else +{ +lean_object* x_21; lean_object* x_22; +x_21 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__9___closed__1; +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_1); +x_22 = lean_apply_8(x_8, lean_box(0), x_1, x_21, x_12, x_13, x_14, x_15, x_16); +if (lean_obj_tag(x_22) == 0) +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +x_24 = lean_ctor_get(x_22, 1); +lean_inc(x_24); +lean_dec(x_22); +x_25 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__15(x_1, x_2, x_3, x_9, x_4, x_5, x_10, x_6, x_7, x_11, x_23, x_12, x_13, x_14, x_15, x_24); +lean_dec(x_23); +return x_25; +} +else +{ +uint8_t x_26; +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_26 = !lean_is_exclusive(x_22); +if (x_26 == 0) +{ +return x_22; +} +else +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_22, 0); +x_28 = lean_ctor_get(x_22, 1); +lean_inc(x_28); +lean_inc(x_27); +lean_dec(x_22); +x_29 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_29, 0, x_27); +lean_ctor_set(x_29, 1, x_28); +return x_29; +} +} +} +} +else +{ +lean_object* x_30; lean_object* x_31; +x_30 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__9___closed__1; +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_1); +x_31 = lean_apply_8(x_8, lean_box(0), x_1, x_30, x_12, x_13, x_14, x_15, x_16); +if (lean_obj_tag(x_31) == 0) +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_31, 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_Meta_Monotonicity_solveMonoStep___lambda__15(x_1, x_2, x_3, x_9, x_4, x_5, x_10, x_6, x_7, x_11, x_32, x_12, x_13, x_14, x_15, x_33); +lean_dec(x_32); +return x_34; +} +else +{ +uint8_t x_35; +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_35 = !lean_is_exclusive(x_31); +if (x_35 == 0) +{ +return x_31; +} +else +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_31, 0); +x_37 = lean_ctor_get(x_31, 1); +lean_inc(x_37); +lean_inc(x_36); +lean_dec(x_31); +x_38 = 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_Monotonicity_solveMonoStep___lambda__17(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; lean_object* x_9; +x_8 = lean_box(0); +x_9 = lean_apply_6(x_1, x_8, x_3, x_4, x_5, x_6, x_7); +return x_9; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__18___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("letFun", 6, 6); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__18___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_Monotonicity_solveMonoStep___lambda__18___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__18(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +_start: +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_16 = lean_alloc_closure((void*)(l_Lean_Meta_Monotonicity_solveMonoStep___lambda__12___boxed), 13, 7); +lean_closure_set(x_16, 0, x_1); +lean_closure_set(x_16, 1, x_2); +lean_closure_set(x_16, 2, x_3); +lean_closure_set(x_16, 3, x_4); +lean_closure_set(x_16, 4, x_5); +lean_closure_set(x_16, 5, x_6); +lean_closure_set(x_16, 6, x_7); +lean_inc(x_3); +x_17 = l_Lean_Meta_instantiateMVarsIfMVarApp(x_3, x_11, x_12, x_13, x_14, x_15); +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +x_19 = lean_ctor_get(x_17, 1); +lean_inc(x_19); +lean_dec(x_17); +x_20 = l_Lean_Expr_cleanupAnnotations(x_18); +x_21 = l_Lean_Expr_isApp(x_20); +if (x_21 == 0) +{ +lean_object* x_22; lean_object* x_23; +lean_dec(x_20); +lean_dec(x_16); +lean_dec(x_9); +lean_dec(x_8); +x_22 = lean_box(0); +x_23 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__12(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_22, x_11, x_12, x_13, x_14, x_19); +return x_23; +} +else +{ +lean_object* x_24; lean_object* x_25; uint8_t x_26; +x_24 = l_Lean_Expr_appArg(x_20, lean_box(0)); +x_25 = l_Lean_Expr_appFnCleanup(x_20, lean_box(0)); +x_26 = l_Lean_Expr_isApp(x_25); +if (x_26 == 0) +{ +lean_object* x_27; lean_object* x_28; +lean_dec(x_25); +lean_dec(x_24); +lean_dec(x_16); +lean_dec(x_9); +lean_dec(x_8); +x_27 = lean_box(0); +x_28 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__12(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_27, x_11, x_12, x_13, x_14, x_19); +return x_28; +} +else +{ +lean_object* x_29; lean_object* x_30; uint8_t x_31; +x_29 = l_Lean_Expr_appArg(x_25, lean_box(0)); +x_30 = l_Lean_Expr_appFnCleanup(x_25, lean_box(0)); +x_31 = l_Lean_Expr_isApp(x_30); +if (x_31 == 0) +{ +lean_object* x_32; lean_object* x_33; +lean_dec(x_30); +lean_dec(x_29); +lean_dec(x_24); +lean_dec(x_16); +lean_dec(x_9); +lean_dec(x_8); +x_32 = lean_box(0); +x_33 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__12(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_32, x_11, x_12, x_13, x_14, x_19); +return x_33; +} +else +{ +lean_object* x_34; uint8_t x_35; +x_34 = l_Lean_Expr_appFnCleanup(x_30, lean_box(0)); +x_35 = l_Lean_Expr_isApp(x_34); +if (x_35 == 0) +{ +lean_object* x_36; lean_object* x_37; +lean_dec(x_34); +lean_dec(x_29); +lean_dec(x_24); +lean_dec(x_16); +lean_dec(x_9); +lean_dec(x_8); +x_36 = lean_box(0); +x_37 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__12(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_36, x_11, x_12, x_13, x_14, x_19); +return x_37; +} +else +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; uint8_t x_41; +x_38 = l_Lean_Expr_appArg(x_34, lean_box(0)); +x_39 = l_Lean_Expr_appFnCleanup(x_34, lean_box(0)); +x_40 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__18___closed__2; +x_41 = l_Lean_Expr_isConstOf(x_39, x_40); +lean_dec(x_39); +if (x_41 == 0) +{ +lean_object* x_42; lean_object* x_43; +lean_dec(x_38); +lean_dec(x_29); +lean_dec(x_24); +lean_dec(x_16); +lean_dec(x_9); +lean_dec(x_8); +x_42 = lean_box(0); +x_43 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__12(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_42, x_11, x_12, x_13, x_14, x_19); +return x_43; +} +else +{ +lean_object* x_44; +lean_dec(x_5); +lean_dec(x_3); +x_44 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__16(x_4, x_1, x_8, x_2, x_9, x_6, x_16, x_7, x_38, x_29, x_24, x_11, x_12, x_13, x_14, x_19); +lean_dec(x_16); +return x_44; +} +} +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__19(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_11 = l_Lean_Expr_bindingDomain_x21(x_1); +x_12 = lean_expr_instantiate1(x_2, x_5); +x_13 = l_Lean_Expr_appFn_x21(x_3); +if (lean_obj_tag(x_1) == 6) +{ +lean_object* x_58; lean_object* x_59; lean_object* x_60; uint8_t x_61; lean_object* x_62; size_t x_63; size_t x_64; uint8_t x_65; +x_58 = lean_ctor_get(x_1, 0); +lean_inc(x_58); +x_59 = lean_ctor_get(x_1, 1); +lean_inc(x_59); +x_60 = lean_ctor_get(x_1, 2); +lean_inc(x_60); +x_61 = lean_ctor_get_uint8(x_1, sizeof(void*)*3 + 8); +lean_dec(x_1); +lean_inc(x_60); +lean_inc(x_59); +lean_inc(x_58); +x_62 = l_Lean_Expr_lam___override(x_58, x_59, x_60, x_61); +x_63 = lean_ptr_addr(x_59); +lean_dec(x_59); +x_64 = lean_ptr_addr(x_11); +x_65 = lean_usize_dec_eq(x_63, x_64); +if (x_65 == 0) +{ +lean_object* x_66; +lean_dec(x_62); +lean_dec(x_60); +x_66 = l_Lean_Expr_lam___override(x_58, x_11, x_12, x_61); +x_14 = x_66; +goto block_57; +} +else +{ +size_t x_67; size_t x_68; uint8_t x_69; +x_67 = lean_ptr_addr(x_60); +lean_dec(x_60); +x_68 = lean_ptr_addr(x_12); +x_69 = lean_usize_dec_eq(x_67, x_68); +if (x_69 == 0) +{ +lean_object* x_70; +lean_dec(x_62); +x_70 = l_Lean_Expr_lam___override(x_58, x_11, x_12, x_61); +x_14 = x_70; +goto block_57; +} +else +{ +uint8_t x_71; +x_71 = l___private_Lean_Expr_0__Lean_beqBinderInfo____x40_Lean_Expr___hyg_406_(x_61, x_61); +if (x_71 == 0) +{ +lean_object* x_72; +lean_dec(x_62); +x_72 = l_Lean_Expr_lam___override(x_58, x_11, x_12, x_61); +x_14 = x_72; +goto block_57; +} +else +{ +lean_dec(x_58); +lean_dec(x_12); +lean_dec(x_11); +x_14 = x_62; +goto block_57; +} +} +} +} +else +{ +lean_object* x_73; lean_object* x_74; +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_1); +x_73 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__15___closed__8; +x_74 = l_panic___at_Lean_Expr_appFn_x21___spec__1(x_73); +x_14 = x_74; +goto block_57; +} +block_57: +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_15 = l_Lean_Expr_app___override(x_13, x_14); +x_16 = lean_box(0); +lean_inc(x_6); +x_17 = l_Lean_Meta_mkFreshExprSyntheticOpaqueMVar(x_15, x_16, x_6, x_7, x_8, x_9, x_10); +x_18 = !lean_is_exclusive(x_17); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; uint8_t x_24; uint8_t x_25; lean_object* x_26; +x_19 = lean_ctor_get(x_17, 0); +x_20 = lean_ctor_get(x_17, 1); +x_21 = lean_box(0); +lean_ctor_set_tag(x_17, 1); +lean_ctor_set(x_17, 1, x_21); +lean_ctor_set(x_17, 0, x_5); +x_22 = lean_array_mk(x_17); +x_23 = 0; +x_24 = 1; +x_25 = 1; +lean_inc(x_19); +x_26 = l_Lean_Meta_mkLambdaFVars(x_22, x_19, x_23, x_24, x_23, x_25, x_6, x_7, x_8, x_9, x_20); +lean_dec(x_22); +if (lean_obj_tag(x_26) == 0) +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; uint8_t x_30; +x_27 = lean_ctor_get(x_26, 0); +lean_inc(x_27); +x_28 = lean_ctor_get(x_26, 1); +lean_inc(x_28); +lean_dec(x_26); +x_29 = l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_4, x_27, x_6, x_7, x_8, x_9, x_28); +lean_dec(x_6); +x_30 = !lean_is_exclusive(x_29); +if (x_30 == 0) +{ +lean_object* x_31; +x_31 = lean_ctor_get(x_29, 0); +lean_dec(x_31); +lean_ctor_set(x_29, 0, x_19); +return x_29; +} +else +{ +lean_object* x_32; lean_object* x_33; +x_32 = lean_ctor_get(x_29, 1); +lean_inc(x_32); +lean_dec(x_29); +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_19); +lean_ctor_set(x_33, 1, x_32); +return x_33; +} +} +else +{ +uint8_t x_34; +lean_dec(x_19); +lean_dec(x_6); +lean_dec(x_4); +x_34 = !lean_is_exclusive(x_26); +if (x_34 == 0) +{ +return x_26; +} +else +{ +lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_35 = lean_ctor_get(x_26, 0); +x_36 = lean_ctor_get(x_26, 1); +lean_inc(x_36); +lean_inc(x_35); +lean_dec(x_26); +x_37 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_37, 0, x_35); +lean_ctor_set(x_37, 1, x_36); +return x_37; +} +} +} +else +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; uint8_t x_43; uint8_t x_44; uint8_t x_45; lean_object* x_46; +x_38 = lean_ctor_get(x_17, 0); +x_39 = lean_ctor_get(x_17, 1); +lean_inc(x_39); +lean_inc(x_38); +lean_dec(x_17); +x_40 = lean_box(0); +x_41 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_41, 0, x_5); +lean_ctor_set(x_41, 1, x_40); +x_42 = lean_array_mk(x_41); +x_43 = 0; +x_44 = 1; +x_45 = 1; +lean_inc(x_38); +x_46 = l_Lean_Meta_mkLambdaFVars(x_42, x_38, x_43, x_44, x_43, x_45, x_6, x_7, x_8, x_9, x_39); +lean_dec(x_42); +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; +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_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_4, x_47, x_6, x_7, x_8, x_9, x_48); +lean_dec(x_6); +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; +} +lean_ctor_set(x_52, 0, x_38); +lean_ctor_set(x_52, 1, x_50); +return x_52; +} +else +{ +lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; +lean_dec(x_38); +lean_dec(x_6); +lean_dec(x_4); +x_53 = lean_ctor_get(x_46, 0); +lean_inc(x_53); +x_54 = lean_ctor_get(x_46, 1); +lean_inc(x_54); +if (lean_is_exclusive(x_46)) { + lean_ctor_release(x_46, 0); + lean_ctor_release(x_46, 1); + x_55 = x_46; +} else { + lean_dec_ref(x_46); + 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_Meta_Monotonicity_solveMonoStep___lambda__20(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +lean_object* x_14; uint8_t x_15; lean_object* x_16; +x_14 = lean_alloc_closure((void*)(l_Lean_Meta_Monotonicity_solveMonoStep___lambda__19___boxed), 10, 4); +lean_closure_set(x_14, 0, x_1); +lean_closure_set(x_14, 1, x_2); +lean_closure_set(x_14, 2, x_3); +lean_closure_set(x_14, 3, x_4); +x_15 = 0; +x_16 = l_Lean_Meta_withLetDecl___at___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_isPerm___spec__1___rarg(x_5, x_6, x_7, x_14, x_15, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_16) == 0) +{ +uint8_t x_17; +x_17 = !lean_is_exclusive(x_16); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_18 = lean_ctor_get(x_16, 0); +x_19 = l_Lean_Expr_mvarId_x21(x_18); +lean_dec(x_18); +x_20 = lean_box(0); +x_21 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_21, 0, x_19); +lean_ctor_set(x_21, 1, x_20); +lean_ctor_set(x_16, 0, x_21); +return x_16; +} +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_16, 0); +x_23 = lean_ctor_get(x_16, 1); +lean_inc(x_23); +lean_inc(x_22); +lean_dec(x_16); +x_24 = l_Lean_Expr_mvarId_x21(x_22); +lean_dec(x_22); +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 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_23); +return x_27; +} +} +else +{ +uint8_t x_28; +x_28 = !lean_is_exclusive(x_16); +if (x_28 == 0) +{ +return x_16; +} +else +{ +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_dec(x_16); +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_Meta_Monotonicity_solveMonoStep___lambda__21(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +_start: +{ +if (lean_obj_tag(x_3) == 8) +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_5); +lean_dec(x_2); +lean_dec(x_1); +x_17 = lean_ctor_get(x_3, 0); +lean_inc(x_17); +x_18 = lean_ctor_get(x_3, 1); +lean_inc(x_18); +x_19 = lean_ctor_get(x_3, 2); +lean_inc(x_19); +x_20 = lean_ctor_get(x_3, 3); +lean_inc(x_20); +lean_dec(x_3); +x_21 = l_Lean_Expr_hasLooseBVars(x_18); +if (x_21 == 0) +{ +uint8_t x_22; +x_22 = l_Lean_Expr_hasLooseBVars(x_19); +if (x_22 == 0) +{ +lean_object* x_23; lean_object* x_24; +lean_dec(x_7); +x_23 = lean_box(0); +x_24 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__20(x_4, x_20, x_10, x_6, x_17, x_18, x_19, x_23, x_12, x_13, x_14, x_15, x_16); +return x_24; +} +else +{ +lean_object* x_25; lean_object* x_26; +x_25 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__9___closed__1; +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_4); +x_26 = lean_apply_8(x_7, lean_box(0), x_4, x_25, x_12, x_13, x_14, x_15, x_16); +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_Monotonicity_solveMonoStep___lambda__20(x_4, x_20, x_10, x_6, x_17, x_18, x_19, x_27, x_12, x_13, x_14, x_15, x_28); +lean_dec(x_27); +return x_29; +} +else +{ +uint8_t x_30; +lean_dec(x_20); +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_10); +lean_dec(x_6); +lean_dec(x_4); +x_30 = !lean_is_exclusive(x_26); +if (x_30 == 0) +{ +return x_26; +} +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_26, 0); +x_32 = lean_ctor_get(x_26, 1); +lean_inc(x_32); +lean_inc(x_31); +lean_dec(x_26); +x_33 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_33, 0, x_31); +lean_ctor_set(x_33, 1, x_32); +return x_33; +} +} +} +} +else +{ +lean_object* x_34; lean_object* x_35; +x_34 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__9___closed__1; +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_4); +x_35 = lean_apply_8(x_7, lean_box(0), x_4, x_34, x_12, x_13, x_14, x_15, x_16); +if (lean_obj_tag(x_35) == 0) +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_35, 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_Monotonicity_solveMonoStep___lambda__20(x_4, x_20, x_10, x_6, x_17, x_18, x_19, x_36, x_12, x_13, x_14, x_15, x_37); +lean_dec(x_36); +return x_38; +} +else +{ +uint8_t x_39; +lean_dec(x_20); +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_10); +lean_dec(x_6); +lean_dec(x_4); +x_39 = !lean_is_exclusive(x_35); +if (x_39 == 0) +{ +return x_35; +} +else +{ +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_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 +{ +lean_object* x_43; lean_object* x_44; +lean_dec(x_10); +x_43 = lean_box(0); +x_44 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__18(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_43, x_12, x_13, x_14, x_15, x_16); +return x_44; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__22(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +_start: +{ +uint8_t x_17; +x_17 = l_Lean_Expr_isMData(x_3); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; +x_18 = lean_box(0); +x_19 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__21(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_18, x_12, x_13, x_14, x_15, x_16); +return x_19; +} +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_5); +lean_dec(x_2); +lean_dec(x_1); +x_20 = l_Lean_Expr_bindingDomain_x21(x_4); +x_21 = l_Lean_Expr_mdataExpr_x21(x_3); +lean_dec(x_3); +x_22 = l_Lean_Expr_appFn_x21(x_10); +lean_dec(x_10); +if (lean_obj_tag(x_4) == 6) +{ +lean_object* x_49; lean_object* x_50; lean_object* x_51; uint8_t x_52; lean_object* x_53; size_t x_54; size_t x_55; uint8_t x_56; +x_49 = lean_ctor_get(x_4, 0); +lean_inc(x_49); +x_50 = lean_ctor_get(x_4, 1); +lean_inc(x_50); +x_51 = lean_ctor_get(x_4, 2); +lean_inc(x_51); +x_52 = lean_ctor_get_uint8(x_4, sizeof(void*)*3 + 8); +lean_dec(x_4); +lean_inc(x_51); +lean_inc(x_50); +lean_inc(x_49); +x_53 = l_Lean_Expr_lam___override(x_49, x_50, x_51, x_52); +x_54 = lean_ptr_addr(x_50); +lean_dec(x_50); +x_55 = lean_ptr_addr(x_20); +x_56 = lean_usize_dec_eq(x_54, x_55); +if (x_56 == 0) +{ +lean_object* x_57; +lean_dec(x_53); +lean_dec(x_51); +x_57 = l_Lean_Expr_lam___override(x_49, x_20, x_21, x_52); +x_23 = x_57; +goto block_48; +} +else +{ +size_t x_58; size_t x_59; uint8_t x_60; +x_58 = lean_ptr_addr(x_51); +lean_dec(x_51); +x_59 = lean_ptr_addr(x_21); +x_60 = lean_usize_dec_eq(x_58, x_59); +if (x_60 == 0) +{ +lean_object* x_61; +lean_dec(x_53); +x_61 = l_Lean_Expr_lam___override(x_49, x_20, x_21, x_52); +x_23 = x_61; +goto block_48; +} +else +{ +uint8_t x_62; +x_62 = l___private_Lean_Expr_0__Lean_beqBinderInfo____x40_Lean_Expr___hyg_406_(x_52, x_52); +if (x_62 == 0) +{ +lean_object* x_63; +lean_dec(x_53); +x_63 = l_Lean_Expr_lam___override(x_49, x_20, x_21, x_52); +x_23 = x_63; +goto block_48; +} +else +{ +lean_dec(x_49); +lean_dec(x_21); +lean_dec(x_20); +x_23 = x_53; +goto block_48; +} +} +} +} +else +{ +lean_object* x_64; lean_object* x_65; +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_4); +x_64 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__15___closed__8; +x_65 = l_panic___at_Lean_Expr_appFn_x21___spec__1(x_64); +x_23 = x_65; +goto block_48; +} +block_48: +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; +x_24 = l_Lean_Expr_app___override(x_22, x_23); +x_25 = lean_box(0); +lean_inc(x_12); +x_26 = l_Lean_Meta_mkFreshExprSyntheticOpaqueMVar(x_24, x_25, x_12, x_13, x_14, x_15, x_16); +x_27 = !lean_is_exclusive(x_26); +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_26, 0); +x_29 = lean_ctor_get(x_26, 1); +lean_inc(x_28); +x_30 = l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_6, x_28, x_12, x_13, x_14, x_15, x_29); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +x_31 = !lean_is_exclusive(x_30); +if (x_31 == 0) +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_30, 0); +lean_dec(x_32); +x_33 = l_Lean_Expr_mvarId_x21(x_28); +lean_dec(x_28); +x_34 = lean_box(0); +lean_ctor_set_tag(x_26, 1); +lean_ctor_set(x_26, 1, x_34); +lean_ctor_set(x_26, 0, x_33); +lean_ctor_set(x_30, 0, x_26); +return x_30; +} +else +{ +lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_35 = lean_ctor_get(x_30, 1); +lean_inc(x_35); +lean_dec(x_30); +x_36 = l_Lean_Expr_mvarId_x21(x_28); +lean_dec(x_28); +x_37 = lean_box(0); +lean_ctor_set_tag(x_26, 1); +lean_ctor_set(x_26, 1, x_37); +lean_ctor_set(x_26, 0, x_36); +x_38 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_38, 0, x_26); +lean_ctor_set(x_38, 1, x_35); +return x_38; +} +} +else +{ +lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_39 = lean_ctor_get(x_26, 0); +x_40 = lean_ctor_get(x_26, 1); +lean_inc(x_40); +lean_inc(x_39); +lean_dec(x_26); +lean_inc(x_39); +x_41 = l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_6, x_39, x_12, x_13, x_14, x_15, x_40); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +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); +} +x_44 = l_Lean_Expr_mvarId_x21(x_39); +lean_dec(x_39); +x_45 = lean_box(0); +x_46 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_46, 0, x_44); +lean_ctor_set(x_46, 1, x_45); +if (lean_is_scalar(x_43)) { + x_47 = lean_alloc_ctor(0, 2, 0); +} else { + x_47 = x_43; +} +lean_ctor_set(x_47, 0, x_46); +lean_ctor_set(x_47, 1, x_42); +return x_47; +} +} +} +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__23___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("monotone_const", 14, 14); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__23___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__1; +x_2 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__3; +x_3 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__23___closed__1; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__23(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_15 = l_Lean_Meta_Monotonicity_headBetaUnderLambda(x_9); +x_16 = l_Lean_Expr_bindingBody_x21(x_15); +x_17 = l_Lean_Expr_hasLooseBVars(x_16); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; +lean_dec(x_16); +lean_dec(x_15); +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_18 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__23___closed__2; +x_19 = l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_applyConst(x_4, x_18, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_19) == 0) +{ +uint8_t x_20; +x_20 = !lean_is_exclusive(x_19); +if (x_20 == 0) +{ +return x_19; +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_19, 0); +x_22 = lean_ctor_get(x_19, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_19); +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_21); +lean_ctor_set(x_23, 1, x_22); +return x_23; +} +} +else +{ +uint8_t x_24; +x_24 = !lean_is_exclusive(x_19); +if (x_24 == 0) +{ +return x_19; +} +else +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_25 = lean_ctor_get(x_19, 0); +x_26 = lean_ctor_get(x_19, 1); +lean_inc(x_26); +lean_inc(x_25); +lean_dec(x_19); +x_27 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +return x_27; +} +} +} +else +{ +lean_object* x_28; lean_object* x_29; +x_28 = lean_box(0); +x_29 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__22(x_1, x_2, x_16, x_15, x_3, x_4, x_5, x_6, x_7, x_8, x_28, x_10, x_11, x_12, x_13, x_14); +return x_29; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__24(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_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; uint8_t x_19; +x_15 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_9, x_10, x_11, x_12, x_13, x_14); +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_15, 1); +lean_inc(x_17); +lean_dec(x_15); +x_18 = l_Lean_Expr_headBeta(x_16); +x_19 = l_Lean_Expr_isLambda(x_18); +if (x_19 == 0) +{ +lean_object* x_20; +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +x_20 = l_Lean_Meta_etaExpand(x_18, x_10, x_11, x_12, x_13, x_17); +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_Meta_Monotonicity_solveMonoStep___lambda__23(x_5, x_6, x_1, x_2, x_3, x_7, x_8, x_4, x_21, x_10, x_11, x_12, x_13, x_22); +return x_23; +} +else +{ +uint8_t x_24; +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_3); +lean_dec(x_2); +lean_dec(x_1); +x_24 = !lean_is_exclusive(x_20); +if (x_24 == 0) +{ +return x_20; +} +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; +} +} +} +else +{ +lean_object* x_28; +x_28 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__23(x_5, x_6, x_1, x_2, x_3, x_7, x_8, x_4, x_18, x_10, x_11, x_12, x_13, x_17); +return x_28; +} +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__25___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Unexpected goal:", 16, 16); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__25___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__25___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__25(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_8 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_8, 0, x_1); +x_9 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__25___closed__2; +x_10 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_10, 0, x_9); +lean_ctor_set(x_10, 1, x_8); +x_11 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__9; +x_12 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_12, 0, x_10); +lean_ctor_set(x_12, 1, x_11); +x_13 = l_Lean_throwError___at_Lean_Meta_Split_applyMatchSplitter___spec__1(x_12, x_3, x_4, x_5, x_6, x_7); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__26(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; +lean_inc(x_1); +x_11 = l_Lean_Meta_instantiateMVarsIfMVarApp(x_1, x_6, x_7, x_8, x_9, x_10); +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +x_13 = lean_ctor_get(x_11, 1); +lean_inc(x_13); +lean_dec(x_11); +x_14 = l_Lean_Expr_cleanupAnnotations(x_12); +x_15 = l_Lean_Expr_isApp(x_14); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; +lean_dec(x_14); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_16 = lean_box(0); +x_17 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__25(x_3, x_16, 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); +return x_17; +} +else +{ +lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_18 = l_Lean_Expr_appArg(x_14, lean_box(0)); +x_19 = l_Lean_Expr_appFnCleanup(x_14, lean_box(0)); +x_20 = l_Lean_Expr_isApp(x_19); +if (x_20 == 0) +{ +lean_object* x_21; lean_object* x_22; +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_21 = lean_box(0); +x_22 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__25(x_3, x_21, 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); +return x_22; +} +else +{ +lean_object* x_23; lean_object* x_24; uint8_t x_25; +x_23 = l_Lean_Expr_appArg(x_19, lean_box(0)); +x_24 = l_Lean_Expr_appFnCleanup(x_19, lean_box(0)); +x_25 = l_Lean_Expr_isApp(x_24); +if (x_25 == 0) +{ +lean_object* x_26; lean_object* x_27; +lean_dec(x_24); +lean_dec(x_23); +lean_dec(x_18); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_26 = lean_box(0); +x_27 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__25(x_3, x_26, 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); +return x_27; +} +else +{ +lean_object* x_28; lean_object* x_29; uint8_t x_30; +x_28 = l_Lean_Expr_appArg(x_24, lean_box(0)); +x_29 = l_Lean_Expr_appFnCleanup(x_24, lean_box(0)); +x_30 = l_Lean_Expr_isApp(x_29); +if (x_30 == 0) +{ +lean_object* x_31; lean_object* x_32; +lean_dec(x_29); +lean_dec(x_28); +lean_dec(x_23); +lean_dec(x_18); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_31 = lean_box(0); +x_32 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__25(x_3, x_31, 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); +return x_32; +} +else +{ +lean_object* x_33; lean_object* x_34; uint8_t x_35; +x_33 = l_Lean_Expr_appArg(x_29, lean_box(0)); +x_34 = l_Lean_Expr_appFnCleanup(x_29, lean_box(0)); +x_35 = l_Lean_Expr_isApp(x_34); +if (x_35 == 0) +{ +lean_object* x_36; lean_object* x_37; +lean_dec(x_34); +lean_dec(x_33); +lean_dec(x_28); +lean_dec(x_23); +lean_dec(x_18); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_36 = lean_box(0); +x_37 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__25(x_3, x_36, 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); +return x_37; +} +else +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; uint8_t x_41; +x_38 = l_Lean_Expr_appArg(x_34, lean_box(0)); +x_39 = l_Lean_Expr_appFnCleanup(x_34, lean_box(0)); +x_40 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__5; +x_41 = l_Lean_Expr_isConstOf(x_39, x_40); +lean_dec(x_39); +if (x_41 == 0) +{ +lean_object* x_42; lean_object* x_43; +lean_dec(x_38); +lean_dec(x_33); +lean_dec(x_28); +lean_dec(x_23); +lean_dec(x_18); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_42 = lean_box(0); +x_43 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__25(x_3, x_42, 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); +return x_43; +} +else +{ +lean_object* x_44; +x_44 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__24(x_2, x_3, x_4, x_1, x_38, x_33, x_28, x_23, x_18, x_6, x_7, x_8, x_9, x_13); +return x_44; +} +} +} +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__27(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +lean_inc(x_1); +x_10 = l_Lean_MVarId_getType(x_1, 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; +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_isForall(x_11); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; +x_14 = lean_box(0); +x_15 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__26(x_11, x_2, x_1, x_3, x_14, x_5, x_6, x_7, x_8, x_12); +return x_15; +} +else +{ +uint8_t x_16; lean_object* x_17; +lean_dec(x_11); +lean_dec(x_3); +lean_dec(x_2); +x_16 = 1; +x_17 = l_Lean_Meta_intro1Core(x_1, x_16, x_5, x_6, x_7, x_8, x_12); +if (lean_obj_tag(x_17) == 0) +{ +uint8_t x_18; +x_18 = !lean_is_exclusive(x_17); +if (x_18 == 0) +{ +lean_object* x_19; 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_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_19, 1); +x_22 = lean_ctor_get(x_19, 0); +lean_dec(x_22); +x_23 = lean_box(0); +lean_ctor_set_tag(x_19, 1); +lean_ctor_set(x_19, 1, x_23); +lean_ctor_set(x_19, 0, x_21); +return x_17; +} +else +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_19, 1); +lean_inc(x_24); +lean_dec(x_19); +x_25 = lean_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); +lean_ctor_set(x_17, 0, x_26); +return x_17; +} +} +else +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_27 = lean_ctor_get(x_17, 0); +x_28 = lean_ctor_get(x_17, 1); +lean_inc(x_28); +lean_inc(x_27); +lean_dec(x_17); +x_29 = lean_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; +} else { + lean_dec_ref(x_27); + x_30 = lean_box(0); +} +x_31 = lean_box(0); +if (lean_is_scalar(x_30)) { + x_32 = lean_alloc_ctor(1, 2, 0); +} else { + x_32 = x_30; + lean_ctor_set_tag(x_32, 1); +} +lean_ctor_set(x_32, 0, x_29); +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_28); +return x_33; +} +} +else +{ +uint8_t x_34; +x_34 = !lean_is_exclusive(x_17); +if (x_34 == 0) +{ +return x_17; +} +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; +} +} +} +} +else +{ +uint8_t x_38; +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_is_exclusive(x_10); +if (x_38 == 0) +{ +return x_10; +} +else +{ +lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_39 = lean_ctor_get(x_10, 0); +x_40 = lean_ctor_get(x_10, 1); +lean_inc(x_40); +lean_inc(x_39); +lean_dec(x_10); +x_41 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_41, 0, x_39); +lean_ctor_set(x_41, 1, x_40); +return x_41; +} +} +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__28___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("monotonicity at\n", 16, 16); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__28___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__28___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__28(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; lean_object* x_10; uint8_t x_11; +lean_inc(x_1); +x_9 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_1, x_4, x_5, x_6, x_7, x_8); +x_10 = lean_ctor_get(x_9, 0); +lean_inc(x_10); +x_11 = lean_unbox(x_10); +lean_dec(x_10); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_12 = lean_ctor_get(x_9, 1); +lean_inc(x_12); +lean_dec(x_9); +x_13 = lean_box(0); +x_14 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__27(x_2, x_1, x_3, x_13, x_4, x_5, x_6, x_7, x_12); +return x_14; +} +else +{ +uint8_t x_15; +x_15 = !lean_is_exclusive(x_9); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_16 = lean_ctor_get(x_9, 1); +x_17 = lean_ctor_get(x_9, 0); +lean_dec(x_17); +lean_inc(x_2); +x_18 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_18, 0, x_2); +x_19 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__28___closed__2; +lean_ctor_set_tag(x_9, 7); +lean_ctor_set(x_9, 1, x_18); +lean_ctor_set(x_9, 0, x_19); +x_20 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__9; +x_21 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_21, 0, x_9); +lean_ctor_set(x_21, 1, x_20); +lean_inc(x_1); +x_22 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_1, x_21, x_4, x_5, x_6, x_7, x_16); +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +x_24 = lean_ctor_get(x_22, 1); +lean_inc(x_24); +lean_dec(x_22); +x_25 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__27(x_2, x_1, x_3, x_23, x_4, x_5, x_6, x_7, x_24); +lean_dec(x_23); +return x_25; +} +else +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_26 = lean_ctor_get(x_9, 1); +lean_inc(x_26); +lean_dec(x_9); +lean_inc(x_2); +x_27 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_27, 0, x_2); +x_28 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__28___closed__2; +x_29 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_27); +x_30 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__9; +x_31 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_31, 0, x_29); +lean_ctor_set(x_31, 1, x_30); +lean_inc(x_1); +x_32 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_1, x_31, x_4, x_5, x_6, x_7, x_26); +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_Monotonicity_solveMonoStep___lambda__27(x_2, x_1, x_3, x_33, x_4, x_5, x_6, x_7, x_34); +lean_dec(x_33); +return x_35; +} +} +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_solveMonoStep___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("monotonicity", 12, 12); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_solveMonoStep___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__9; +x_2 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__11; +x_3 = l_Lean_Meta_Monotonicity_solveMonoStep___closed__1; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_8 = l_Lean_Meta_Monotonicity_solveMonoStep___closed__2; +lean_inc(x_2); +x_9 = lean_alloc_closure((void*)(l_Lean_Meta_Monotonicity_solveMonoStep___lambda__28), 8, 3); +lean_closure_set(x_9, 0, x_8); +lean_closure_set(x_9, 1, x_2); +lean_closure_set(x_9, 2, x_1); +x_10 = l_Lean_MVarId_withContext___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___spec__2___rarg(x_2, x_9, x_3, x_4, x_5, x_6, x_7); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Monotonicity_solveMonoStep___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) { +_start: +{ +lean_object* x_8; +x_8 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Monotonicity_solveMonoStep___spec__1___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Monotonicity_solveMonoStep___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) { +_start: +{ +lean_object* x_8; +x_8 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Monotonicity_solveMonoStep___spec__1___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Monotonicity_solveMonoStep___spec__1___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; +x_8 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Monotonicity_solveMonoStep___spec__1___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Monotonicity_solveMonoStep___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +_start: +{ +size_t x_16; size_t x_17; lean_object* x_18; +x_16 = lean_unbox_usize(x_8); +lean_dec(x_8); +x_17 = lean_unbox_usize(x_9); +lean_dec(x_9); +x_18 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Monotonicity_solveMonoStep___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_16, x_17, x_10, x_11, x_12, x_13, x_14, x_15); +lean_dec(x_7); +lean_dec(x_4); +lean_dec(x_3); +return x_18; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Monotonicity_solveMonoStep___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +lean_object* x_14; +x_14 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Monotonicity_solveMonoStep___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); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_Monotonicity_solveMonoStep___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +size_t x_4; size_t x_5; lean_object* x_6; +x_4 = lean_unbox_usize(x_1); +lean_dec(x_1); +x_5 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_6 = l_Array_mapMUnsafe_map___at_Lean_Meta_Monotonicity_solveMonoStep___spec__3(x_4, x_5, x_3); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; +x_8 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__1(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_Meta_Monotonicity_solveMonoStep___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_Meta_Monotonicity_solveMonoStep___lambda__2(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_Meta_Monotonicity_solveMonoStep___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; +x_7 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +x_13 = l_Lean_Meta_Monotonicity_solveMonoStep___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); +lean_dec(x_7); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___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) { +_start: +{ +lean_object* x_13; +x_13 = l_Lean_Meta_Monotonicity_solveMonoStep___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); +lean_dec(x_7); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___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) { +_start: +{ +lean_object* x_12; +x_12 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__6(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_6); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___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) { +_start: +{ +lean_object* x_7; +x_7 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__7(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___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) { +_start: +{ +lean_object* x_10; +x_10 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__8(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_Meta_Monotonicity_solveMonoStep___lambda__9___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; +x_12 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__9(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_6); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___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_Meta_Monotonicity_solveMonoStep___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_8); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__11___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +lean_object* x_14; +x_14 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__11(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_8); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__12___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +lean_object* x_14; +x_14 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__12(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_8); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__14___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; +x_8 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__14(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__15___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +_start: +{ +lean_object* x_17; +x_17 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__15(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +lean_dec(x_11); +lean_dec(x_9); +return x_17; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__16___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +_start: +{ +lean_object* x_17; +x_17 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__16(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +lean_dec(x_7); +return x_17; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__17___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; +x_8 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__17(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_2); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__18___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +_start: +{ +lean_object* x_16; +x_16 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__18(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +lean_dec(x_10); +return x_16; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__19___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__19(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_3); +lean_dec(x_2); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__20___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +lean_object* x_14; +x_14 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__20(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_8); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__21___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +_start: +{ +lean_object* x_17; +x_17 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__21(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +lean_dec(x_11); +return x_17; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__22___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +_start: +{ +lean_object* x_17; +x_17 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__22(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +lean_dec(x_11); +return x_17; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__25___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; +x_8 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__25(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__26___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__26(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; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMonoStep___lambda__27___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l_Lean_Meta_Monotonicity_solveMonoStep___lambda__27(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_forM___at_Lean_Meta_Monotonicity_solveMono___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_8; lean_object* x_9; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_8 = lean_box(0); +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; +} +else +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_10 = lean_ctor_get(x_2, 0); +lean_inc(x_10); +x_11 = lean_ctor_get(x_2, 1); +lean_inc(x_11); +lean_dec(x_2); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_1); +x_12 = l_Lean_Meta_Monotonicity_solveMono(x_1, x_10, x_3, x_4, x_5, x_6, x_7); +if (lean_obj_tag(x_12) == 0) +{ +lean_object* x_13; +x_13 = lean_ctor_get(x_12, 1); +lean_inc(x_13); +lean_dec(x_12); +x_2 = x_11; +x_7 = x_13; +goto _start; +} +else +{ +uint8_t x_15; +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_15 = !lean_is_exclusive(x_12); +if (x_15 == 0) +{ +return x_12; +} +else +{ +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_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; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMono___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, lean_box(0), x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_solveMono(lean_object* x_1, 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_Monotonicity_solveMono___lambda__1), 9, 1); +lean_closure_set(x_8, 0, x_1); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_8); +x_9 = l_Lean_Meta_Monotonicity_solveMonoStep(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; 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_inc(x_11); +lean_dec(x_9); +x_12 = l_List_forM___at_Lean_Meta_Monotonicity_solveMono___spec__1(x_8, x_10, x_3, x_4, x_5, x_6, x_11); +return x_12; +} +else +{ +uint8_t x_13; +lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_13 = !lean_is_exclusive(x_9); +if (x_13 == 0) +{ +return x_9; +} +else +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; +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_alloc_ctor(1, 2, 0); +lean_ctor_set(x_16, 0, x_14); +lean_ctor_set(x_16, 1, x_15); +return x_16; +} +} +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_evalMonotonicity___rarg___lambda__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_defaultFailK), 1, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_evalMonotonicity___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; +x_10 = l_Lean_Elab_Tactic_getMainGoal(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_10) == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; 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); +x_13 = l_Lean_Meta_Monotonicity_evalMonotonicity___rarg___lambda__1___closed__1; +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +x_14 = l_Lean_Meta_Monotonicity_solveMonoStep(x_13, x_11, 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; lean_object* x_17; +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +x_16 = lean_ctor_get(x_14, 1); +lean_inc(x_16); +lean_dec(x_14); +x_17 = l_Lean_Elab_Tactic_replaceMainGoal(x_15, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_16); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +if (lean_obj_tag(x_17) == 0) +{ +uint8_t x_18; +x_18 = !lean_is_exclusive(x_17); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; +x_19 = lean_ctor_get(x_17, 0); +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 +{ +uint8_t x_24; +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; +} +} +} +else +{ +uint8_t x_28; +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +x_28 = !lean_is_exclusive(x_14); +if (x_28 == 0) +{ +return x_14; +} +else +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_29 = lean_ctor_get(x_14, 0); +x_30 = lean_ctor_get(x_14, 1); +lean_inc(x_30); +lean_inc(x_29); +lean_dec(x_14); +x_31 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_31, 0, x_29); +lean_ctor_set(x_31, 1, x_30); +return x_31; +} +} +} +else +{ +uint8_t x_32; +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +x_32 = !lean_is_exclusive(x_10); +if (x_32 == 0) +{ +return x_10; +} +else +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_33 = lean_ctor_get(x_10, 0); +x_34 = lean_ctor_get(x_10, 1); +lean_inc(x_34); +lean_inc(x_33); +lean_dec(x_10); +x_35 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_35, 0, x_33); +lean_ctor_set(x_35, 1, x_34); +return x_35; +} +} +} +} +static lean_object* _init_l_Lean_Meta_Monotonicity_evalMonotonicity___rarg___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Monotonicity_evalMonotonicity___rarg___lambda__1___boxed), 9, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_evalMonotonicity___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_Meta_Monotonicity_evalMonotonicity___rarg___closed__1; +x_11 = l_Lean_Elab_Tactic_withMainContext___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_EXPORT lean_object* l_Lean_Meta_Monotonicity_evalMonotonicity(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_Meta_Monotonicity_evalMonotonicity___rarg), 9, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_evalMonotonicity___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_Lean_Meta_Monotonicity_evalMonotonicity___rarg___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Monotonicity_evalMonotonicity___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Lean_Meta_Monotonicity_evalMonotonicity(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l___regBuiltin_Lean_Meta_Monotonicity_evalMonotonicity__1___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__1; +x_2 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__3; +x_3 = l_Lean_Meta_Monotonicity_solveMonoStep___closed__1; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l___regBuiltin_Lean_Meta_Monotonicity_evalMonotonicity__1___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("evalMonotonicity", 16, 16); +return x_1; +} +} +static lean_object* _init_l___regBuiltin_Lean_Meta_Monotonicity_evalMonotonicity__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_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__1; +x_2 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__2; +x_3 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__3; +x_4 = l___regBuiltin_Lean_Meta_Monotonicity_evalMonotonicity__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___regBuiltin_Lean_Meta_Monotonicity_evalMonotonicity__1___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Elab_Tactic_tacticElabAttribute; +return x_1; +} +} +static lean_object* _init_l___regBuiltin_Lean_Meta_Monotonicity_evalMonotonicity__1___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Monotonicity_evalMonotonicity___boxed), 1, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Meta_Monotonicity_evalMonotonicity__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; +x_2 = l___regBuiltin_Lean_Meta_Monotonicity_evalMonotonicity__1___closed__4; +x_3 = l___regBuiltin_Lean_Meta_Monotonicity_evalMonotonicity__1___closed__1; +x_4 = l___regBuiltin_Lean_Meta_Monotonicity_evalMonotonicity__1___closed__3; +x_5 = l___regBuiltin_Lean_Meta_Monotonicity_evalMonotonicity__1___closed__5; +x_6 = l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(x_2, x_3, x_4, x_5, x_1); +return x_6; +} +} +static lean_object* _init_l_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_5433____closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__4; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_5433____closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_5433____closed__1; +x_2 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__6; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_5433____closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_5433____closed__2; +x_2 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_5433____closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_5433____closed__3; +x_2 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__9; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_5433____closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_5433____closed__4; +x_2 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__11; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_5433____closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_5433____closed__5; +x_2 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__3; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_5433____closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_5433____closed__6; +x_2 = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__14; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_5433____closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_5433____closed__7; +x_2 = lean_unsigned_to_nat(5433u); +x_3 = l_Lean_Name_num___override(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_5433_(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_Monotonicity_solveMonoStep___closed__2; +x_3 = 0; +x_4 = l_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_5433____closed__8; +x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); +return x_5; +} +} +lean_object* initialize_Lean_Meta_Tactic_Split(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Elab_RecAppSyntax(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Elab_Tactic_Basic(uint8_t builtin, lean_object*); +lean_object* initialize_Init_Internal_Order(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Lean_Elab_Tactic_Monotonicity(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_Tactic_Split(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Lean_Elab_RecAppSyntax(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Lean_Elab_Tactic_Basic(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Init_Internal_Order(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l_Lean_Loop_forIn_loop___at_Lean_Meta_Monotonicity_headBetaUnderLambda___spec__1___closed__1 = _init_l_Lean_Loop_forIn_loop___at_Lean_Meta_Monotonicity_headBetaUnderLambda___spec__1___closed__1(); +lean_mark_persistent(l_Lean_Loop_forIn_loop___at_Lean_Meta_Monotonicity_headBetaUnderLambda___spec__1___closed__1); +l_Lean_Loop_forIn_loop___at_Lean_Meta_Monotonicity_headBetaUnderLambda___spec__1___closed__2 = _init_l_Lean_Loop_forIn_loop___at_Lean_Meta_Monotonicity_headBetaUnderLambda___spec__1___closed__2(); +lean_mark_persistent(l_Lean_Loop_forIn_loop___at_Lean_Meta_Monotonicity_headBetaUnderLambda___spec__1___closed__2); +l_Lean_Loop_forIn_loop___at_Lean_Meta_Monotonicity_headBetaUnderLambda___spec__1___closed__3 = _init_l_Lean_Loop_forIn_loop___at_Lean_Meta_Monotonicity_headBetaUnderLambda___spec__1___closed__3(); +lean_mark_persistent(l_Lean_Loop_forIn_loop___at_Lean_Meta_Monotonicity_headBetaUnderLambda___spec__1___closed__3); +l_Lean_Loop_forIn_loop___at_Lean_Meta_Monotonicity_headBetaUnderLambda___spec__1___closed__4 = _init_l_Lean_Loop_forIn_loop___at_Lean_Meta_Monotonicity_headBetaUnderLambda___spec__1___closed__4(); +lean_mark_persistent(l_Lean_Loop_forIn_loop___at_Lean_Meta_Monotonicity_headBetaUnderLambda___spec__1___closed__4); +l_Lean_Meta_Monotonicity_headBetaUnderLambda___closed__1 = _init_l_Lean_Meta_Monotonicity_headBetaUnderLambda___closed__1(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_headBetaUnderLambda___closed__1); +l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__3___closed__1 = _init_l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__3___closed__1(); +l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__3___closed__2 = _init_l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__3___closed__2(); +l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__6___closed__1 = _init_l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__6___closed__1(); +lean_mark_persistent(l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__6___closed__1); +l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__9___closed__1 = _init_l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__9___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__9___closed__1); +l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__9___closed__2 = _init_l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__9___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__9___closed__2); +l_panic___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__13___closed__1 = _init_l_panic___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__13___closed__1(); +lean_mark_persistent(l_panic___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__13___closed__1); +l_Lean_Meta_DiscrTree_insertCore___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__1___closed__1 = _init_l_Lean_Meta_DiscrTree_insertCore___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__1___closed__1(); +lean_mark_persistent(l_Lean_Meta_DiscrTree_insertCore___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__1___closed__1); +l_Lean_Meta_DiscrTree_insertCore___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__1___closed__2 = _init_l_Lean_Meta_DiscrTree_insertCore___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__1___closed__2(); +lean_mark_persistent(l_Lean_Meta_DiscrTree_insertCore___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__1___closed__2); +l_Lean_Meta_DiscrTree_insertCore___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__1___closed__3 = _init_l_Lean_Meta_DiscrTree_insertCore___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__1___closed__3(); +lean_mark_persistent(l_Lean_Meta_DiscrTree_insertCore___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__1___closed__3); +l_Lean_Meta_DiscrTree_insertCore___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__1___closed__4 = _init_l_Lean_Meta_DiscrTree_insertCore___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__1___closed__4(); +lean_mark_persistent(l_Lean_Meta_DiscrTree_insertCore___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____spec__1___closed__4); +l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__1 = _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__1(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__1); +l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__2 = _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__2(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__2); +l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__3 = _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__3(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__3); +l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__4 = _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__4(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__4); +l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__5 = _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__5(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__5); +l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__6 = _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__6(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__6); +l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__7 = _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__7(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__7); +l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__8 = _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__8(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__8); +l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__9 = _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__9(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__9); +l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__10 = _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__10(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166____closed__10); +if (builtin) {res = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_166_(lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +l_Lean_Meta_Monotonicity_monotoneExt = lean_io_result_get_value(res); +lean_mark_persistent(l_Lean_Meta_Monotonicity_monotoneExt); +lean_dec_ref(res); +}l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____spec__1___closed__1 = _init_l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____spec__1___closed__1(); +lean_mark_persistent(l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____spec__1___closed__1); +l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____spec__1___closed__2 = _init_l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____spec__1___closed__2(); +lean_mark_persistent(l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____spec__1___closed__2); +l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__1___closed__1 = _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__1___closed__1(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__1___closed__1); +l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__1 = _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__1(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__1); +l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__2 = _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__2(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__2); +l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__3 = _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__3(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__3); +l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__4 = _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__4(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__4); +l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__5 = _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__5(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__5); +l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__6 = _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__6(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__6); +l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__7 = _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__7(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__7); +l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__8 = _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__8(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__8); +l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__9 = _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__9(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__9); +l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__10 = _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__10(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__4___closed__10); +l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__1 = _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__1(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__1); +l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__2 = _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__2(); +l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__3 = _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__3(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__3); +l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__4 = _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__4(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__4); +l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__5 = _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__5(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__5); +l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__6 = _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__6(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__6); +l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__7 = _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__7(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__7); +l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__8 = _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__8(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__8); +l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__9 = _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__9(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__9); +l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__10 = _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__10(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__10); +l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__11 = _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__11(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__11); +l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__12 = _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__12(); +l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__13 = _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__13(); +l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__14 = _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__14(); +l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__15 = _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__15(); +l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__16 = _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__16(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__16); +l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__17 = _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__17(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__5___closed__17); +l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__6___closed__1 = _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__6___closed__1(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__6___closed__1); +l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__6___closed__2 = _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__6___closed__2(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____lambda__6___closed__2); +l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__1 = _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__1(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__1); +l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__2 = _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__2(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__2); +l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__3 = _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__3(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__3); +l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__4 = _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__4(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__4); +l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__5 = _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__5(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__5); +l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__6 = _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__6(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__6); +l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__7 = _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__7(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__7); +l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__8 = _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__8(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__8); +l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__9 = _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__9(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__9); +l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__10 = _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__10(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__10); +l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__11 = _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__11(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__11); +l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__12 = _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__12(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__12); +l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__13 = _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__13(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__13); +l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__14 = _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__14(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__14); +l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__15 = _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__15(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__15); +l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__16 = _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__16(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__16); +l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__17 = _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__17(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__17); +l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__18 = _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__18(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__18); +l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__19 = _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__19(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__19); +l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__20 = _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__20(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__20); +l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__21 = _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__21(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__21); +l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__22 = _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__22(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__22); +l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__23 = _init_l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__23(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249____closed__23); +if (builtin) {res = l_Lean_Meta_Monotonicity_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_249_(lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +}l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_defaultFailK___rarg___closed__1 = _init_l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_defaultFailK___rarg___closed__1(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_defaultFailK___rarg___closed__1); +l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_defaultFailK___rarg___closed__2 = _init_l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_defaultFailK___rarg___closed__2(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_defaultFailK___rarg___closed__2); +l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_defaultFailK___rarg___closed__3 = _init_l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_defaultFailK___rarg___closed__3(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_defaultFailK___rarg___closed__3); +l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_defaultFailK___rarg___closed__4 = _init_l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_defaultFailK___rarg___closed__4(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_defaultFailK___rarg___closed__4); +l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_defaultFailK___rarg___closed__5 = _init_l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_defaultFailK___rarg___closed__5(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_defaultFailK___rarg___closed__5); +l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_defaultFailK___rarg___closed__6 = _init_l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_defaultFailK___rarg___closed__6(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_defaultFailK___rarg___closed__6); +l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_defaultFailK___rarg___closed__7 = _init_l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_defaultFailK___rarg___closed__7(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_defaultFailK___rarg___closed__7); +l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_defaultFailK___rarg___closed__8 = _init_l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_defaultFailK___rarg___closed__8(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_defaultFailK___rarg___closed__8); +l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_applyConst___lambda__1___closed__1 = _init_l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_applyConst___lambda__1___closed__1(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_applyConst___lambda__1___closed__1); +l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_applyConst___lambda__1___closed__2 = _init_l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_applyConst___lambda__1___closed__2(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_applyConst___lambda__1___closed__2); +l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_applyConst___lambda__1___closed__3 = _init_l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_applyConst___lambda__1___closed__3(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_applyConst___lambda__1___closed__3); +l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_applyConst___lambda__1___closed__4 = _init_l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_applyConst___lambda__1___closed__4(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_applyConst___lambda__1___closed__4); +l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_applyConst___closed__1 = _init_l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_applyConst___closed__1(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Monotonicity_0__Lean_Meta_Monotonicity_applyConst___closed__1); +l_Lean_Meta_Monotonicity_solveMonoCall___lambda__2___closed__1 = _init_l_Lean_Meta_Monotonicity_solveMonoCall___lambda__2___closed__1(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_solveMonoCall___lambda__2___closed__1); +l_Lean_Meta_Monotonicity_solveMonoCall___lambda__2___closed__2 = _init_l_Lean_Meta_Monotonicity_solveMonoCall___lambda__2___closed__2(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_solveMonoCall___lambda__2___closed__2); +l_Lean_Meta_Monotonicity_solveMonoCall___lambda__2___closed__3 = _init_l_Lean_Meta_Monotonicity_solveMonoCall___lambda__2___closed__3(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_solveMonoCall___lambda__2___closed__3); +l_Lean_Meta_Monotonicity_solveMonoCall___lambda__2___closed__4 = _init_l_Lean_Meta_Monotonicity_solveMonoCall___lambda__2___closed__4(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_solveMonoCall___lambda__2___closed__4); +l_Lean_Meta_Monotonicity_solveMonoCall___lambda__3___closed__1 = _init_l_Lean_Meta_Monotonicity_solveMonoCall___lambda__3___closed__1(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_solveMonoCall___lambda__3___closed__1); +l_Lean_Meta_Monotonicity_solveMonoCall___lambda__3___closed__2 = _init_l_Lean_Meta_Monotonicity_solveMonoCall___lambda__3___closed__2(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_solveMonoCall___lambda__3___closed__2); +l_Lean_Meta_Monotonicity_solveMonoCall___lambda__3___closed__3 = _init_l_Lean_Meta_Monotonicity_solveMonoCall___lambda__3___closed__3(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_solveMonoCall___lambda__3___closed__3); +l_Lean_Meta_Monotonicity_solveMonoCall___lambda__3___closed__4 = _init_l_Lean_Meta_Monotonicity_solveMonoCall___lambda__3___closed__4(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_solveMonoCall___lambda__3___closed__4); +l_Lean_Meta_Monotonicity_solveMonoCall___lambda__4___closed__1 = _init_l_Lean_Meta_Monotonicity_solveMonoCall___lambda__4___closed__1(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_solveMonoCall___lambda__4___closed__1); +l_Lean_Meta_Monotonicity_solveMonoCall___lambda__4___closed__2 = _init_l_Lean_Meta_Monotonicity_solveMonoCall___lambda__4___closed__2(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_solveMonoCall___lambda__4___closed__2); +l_Lean_Meta_Monotonicity_solveMonoCall___lambda__4___closed__3 = _init_l_Lean_Meta_Monotonicity_solveMonoCall___lambda__4___closed__3(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_solveMonoCall___lambda__4___closed__3); +l_Lean_Meta_Monotonicity_solveMonoCall___lambda__4___closed__4 = _init_l_Lean_Meta_Monotonicity_solveMonoCall___lambda__4___closed__4(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_solveMonoCall___lambda__4___closed__4); +l_Lean_Meta_Monotonicity_solveMonoCall___lambda__5___closed__1 = _init_l_Lean_Meta_Monotonicity_solveMonoCall___lambda__5___closed__1(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_solveMonoCall___lambda__5___closed__1); +l_Lean_Meta_Monotonicity_solveMonoCall___lambda__5___closed__2 = _init_l_Lean_Meta_Monotonicity_solveMonoCall___lambda__5___closed__2(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_solveMonoCall___lambda__5___closed__2); +l_Lean_Meta_Monotonicity_solveMonoCall___lambda__5___closed__3 = _init_l_Lean_Meta_Monotonicity_solveMonoCall___lambda__5___closed__3(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_solveMonoCall___lambda__5___closed__3); +l_Lean_Meta_Monotonicity_solveMonoCall___lambda__5___closed__4 = _init_l_Lean_Meta_Monotonicity_solveMonoCall___lambda__5___closed__4(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_solveMonoCall___lambda__5___closed__4); +l_Lean_Meta_Monotonicity_solveMonoCall___lambda__6___closed__1 = _init_l_Lean_Meta_Monotonicity_solveMonoCall___lambda__6___closed__1(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_solveMonoCall___lambda__6___closed__1); +l_Lean_Meta_Monotonicity_solveMonoCall___lambda__6___closed__2 = _init_l_Lean_Meta_Monotonicity_solveMonoCall___lambda__6___closed__2(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_solveMonoCall___lambda__6___closed__2); +l_Lean_Meta_Monotonicity_solveMonoCall___lambda__8___closed__1 = _init_l_Lean_Meta_Monotonicity_solveMonoCall___lambda__8___closed__1(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_solveMonoCall___lambda__8___closed__1); +l_Lean_Meta_Monotonicity_solveMonoCall___lambda__8___closed__2 = _init_l_Lean_Meta_Monotonicity_solveMonoCall___lambda__8___closed__2(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_solveMonoCall___lambda__8___closed__2); +l_Lean_Meta_Monotonicity_solveMonoCall___lambda__9___closed__1 = _init_l_Lean_Meta_Monotonicity_solveMonoCall___lambda__9___closed__1(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_solveMonoCall___lambda__9___closed__1); +l_Lean_Meta_Monotonicity_solveMonoCall___lambda__9___closed__2 = _init_l_Lean_Meta_Monotonicity_solveMonoCall___lambda__9___closed__2(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_solveMonoCall___lambda__9___closed__2); +l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Monotonicity_solveMonoStep___spec__1___closed__1 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Monotonicity_solveMonoStep___spec__1___closed__1(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Monotonicity_solveMonoStep___spec__1___closed__1); +l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Monotonicity_solveMonoStep___spec__1___closed__2 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Monotonicity_solveMonoStep___spec__1___closed__2(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Monotonicity_solveMonoStep___spec__1___closed__2); +l_Std_Range_forIn_x27_loop___at_Lean_Meta_Monotonicity_solveMonoStep___spec__2___closed__1 = _init_l_Std_Range_forIn_x27_loop___at_Lean_Meta_Monotonicity_solveMonoStep___spec__2___closed__1(); +lean_mark_persistent(l_Std_Range_forIn_x27_loop___at_Lean_Meta_Monotonicity_solveMonoStep___spec__2___closed__1); +l_Std_Range_forIn_x27_loop___at_Lean_Meta_Monotonicity_solveMonoStep___spec__2___closed__2 = _init_l_Std_Range_forIn_x27_loop___at_Lean_Meta_Monotonicity_solveMonoStep___spec__2___closed__2(); +lean_mark_persistent(l_Std_Range_forIn_x27_loop___at_Lean_Meta_Monotonicity_solveMonoStep___spec__2___closed__2); +l_Lean_Meta_Monotonicity_solveMonoStep___lambda__4___closed__1 = _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__4___closed__1(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_solveMonoStep___lambda__4___closed__1); +l_Lean_Meta_Monotonicity_solveMonoStep___lambda__4___closed__2 = _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__4___closed__2(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_solveMonoStep___lambda__4___closed__2); +l_Lean_Meta_Monotonicity_solveMonoStep___lambda__5___closed__1 = _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__5___closed__1(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_solveMonoStep___lambda__5___closed__1); +l_Lean_Meta_Monotonicity_solveMonoStep___lambda__6___closed__1 = _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__6___closed__1(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_solveMonoStep___lambda__6___closed__1); +l_Lean_Meta_Monotonicity_solveMonoStep___lambda__6___closed__2 = _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__6___closed__2(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_solveMonoStep___lambda__6___closed__2); +l_Lean_Meta_Monotonicity_solveMonoStep___lambda__6___closed__3 = _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__6___closed__3(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_solveMonoStep___lambda__6___closed__3); +l_Lean_Meta_Monotonicity_solveMonoStep___lambda__6___closed__4 = _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__6___closed__4(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_solveMonoStep___lambda__6___closed__4); +l_Lean_Meta_Monotonicity_solveMonoStep___lambda__9___closed__1 = _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__9___closed__1(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_solveMonoStep___lambda__9___closed__1); +l_Lean_Meta_Monotonicity_solveMonoStep___lambda__9___closed__2 = _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__9___closed__2(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_solveMonoStep___lambda__9___closed__2); +l_Lean_Meta_Monotonicity_solveMonoStep___lambda__9___closed__3 = _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__9___closed__3(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_solveMonoStep___lambda__9___closed__3); +l_Lean_Meta_Monotonicity_solveMonoStep___lambda__9___closed__4 = _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__9___closed__4(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_solveMonoStep___lambda__9___closed__4); +l_Lean_Meta_Monotonicity_solveMonoStep___lambda__9___closed__5 = _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__9___closed__5(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_solveMonoStep___lambda__9___closed__5); +l_Lean_Meta_Monotonicity_solveMonoStep___lambda__9___closed__6 = _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__9___closed__6(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_solveMonoStep___lambda__9___closed__6); +l_Lean_Meta_Monotonicity_solveMonoStep___lambda__9___closed__7 = _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__9___closed__7(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_solveMonoStep___lambda__9___closed__7); +l_Lean_Meta_Monotonicity_solveMonoStep___lambda__10___closed__1 = _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__10___closed__1(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_solveMonoStep___lambda__10___closed__1); +l_Lean_Meta_Monotonicity_solveMonoStep___lambda__10___closed__2 = _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__10___closed__2(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_solveMonoStep___lambda__10___closed__2); +l_Lean_Meta_Monotonicity_solveMonoStep___lambda__12___closed__1 = _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__12___closed__1(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_solveMonoStep___lambda__12___closed__1); +l_Lean_Meta_Monotonicity_solveMonoStep___lambda__12___closed__2 = _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__12___closed__2(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_solveMonoStep___lambda__12___closed__2); +l_Lean_Meta_Monotonicity_solveMonoStep___lambda__12___closed__3 = _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__12___closed__3(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_solveMonoStep___lambda__12___closed__3); +l_Lean_Meta_Monotonicity_solveMonoStep___lambda__12___closed__4 = _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__12___closed__4(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_solveMonoStep___lambda__12___closed__4); +l_Lean_Meta_Monotonicity_solveMonoStep___lambda__12___closed__5 = _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__12___closed__5(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_solveMonoStep___lambda__12___closed__5); +l_Lean_Meta_Monotonicity_solveMonoStep___lambda__12___closed__6 = _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__12___closed__6(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_solveMonoStep___lambda__12___closed__6); +l_Lean_Meta_Monotonicity_solveMonoStep___lambda__12___closed__7 = _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__12___closed__7(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_solveMonoStep___lambda__12___closed__7); +l_Lean_Meta_Monotonicity_solveMonoStep___lambda__12___closed__8 = _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__12___closed__8(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_solveMonoStep___lambda__12___closed__8); +l_Lean_Meta_Monotonicity_solveMonoStep___lambda__12___closed__9 = _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__12___closed__9(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_solveMonoStep___lambda__12___closed__9); +l_Lean_Meta_Monotonicity_solveMonoStep___lambda__15___closed__1 = _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__15___closed__1(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_solveMonoStep___lambda__15___closed__1); +l_Lean_Meta_Monotonicity_solveMonoStep___lambda__15___closed__2 = _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__15___closed__2(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_solveMonoStep___lambda__15___closed__2); +l_Lean_Meta_Monotonicity_solveMonoStep___lambda__15___closed__3 = _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__15___closed__3(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_solveMonoStep___lambda__15___closed__3); +l_Lean_Meta_Monotonicity_solveMonoStep___lambda__15___closed__4 = _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__15___closed__4(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_solveMonoStep___lambda__15___closed__4); +l_Lean_Meta_Monotonicity_solveMonoStep___lambda__15___closed__5 = _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__15___closed__5(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_solveMonoStep___lambda__15___closed__5); +l_Lean_Meta_Monotonicity_solveMonoStep___lambda__15___closed__6 = _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__15___closed__6(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_solveMonoStep___lambda__15___closed__6); +l_Lean_Meta_Monotonicity_solveMonoStep___lambda__15___closed__7 = _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__15___closed__7(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_solveMonoStep___lambda__15___closed__7); +l_Lean_Meta_Monotonicity_solveMonoStep___lambda__15___closed__8 = _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__15___closed__8(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_solveMonoStep___lambda__15___closed__8); +l_Lean_Meta_Monotonicity_solveMonoStep___lambda__18___closed__1 = _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__18___closed__1(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_solveMonoStep___lambda__18___closed__1); +l_Lean_Meta_Monotonicity_solveMonoStep___lambda__18___closed__2 = _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__18___closed__2(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_solveMonoStep___lambda__18___closed__2); +l_Lean_Meta_Monotonicity_solveMonoStep___lambda__23___closed__1 = _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__23___closed__1(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_solveMonoStep___lambda__23___closed__1); +l_Lean_Meta_Monotonicity_solveMonoStep___lambda__23___closed__2 = _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__23___closed__2(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_solveMonoStep___lambda__23___closed__2); +l_Lean_Meta_Monotonicity_solveMonoStep___lambda__25___closed__1 = _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__25___closed__1(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_solveMonoStep___lambda__25___closed__1); +l_Lean_Meta_Monotonicity_solveMonoStep___lambda__25___closed__2 = _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__25___closed__2(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_solveMonoStep___lambda__25___closed__2); +l_Lean_Meta_Monotonicity_solveMonoStep___lambda__28___closed__1 = _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__28___closed__1(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_solveMonoStep___lambda__28___closed__1); +l_Lean_Meta_Monotonicity_solveMonoStep___lambda__28___closed__2 = _init_l_Lean_Meta_Monotonicity_solveMonoStep___lambda__28___closed__2(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_solveMonoStep___lambda__28___closed__2); +l_Lean_Meta_Monotonicity_solveMonoStep___closed__1 = _init_l_Lean_Meta_Monotonicity_solveMonoStep___closed__1(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_solveMonoStep___closed__1); +l_Lean_Meta_Monotonicity_solveMonoStep___closed__2 = _init_l_Lean_Meta_Monotonicity_solveMonoStep___closed__2(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_solveMonoStep___closed__2); +l_Lean_Meta_Monotonicity_evalMonotonicity___rarg___lambda__1___closed__1 = _init_l_Lean_Meta_Monotonicity_evalMonotonicity___rarg___lambda__1___closed__1(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_evalMonotonicity___rarg___lambda__1___closed__1); +l_Lean_Meta_Monotonicity_evalMonotonicity___rarg___closed__1 = _init_l_Lean_Meta_Monotonicity_evalMonotonicity___rarg___closed__1(); +lean_mark_persistent(l_Lean_Meta_Monotonicity_evalMonotonicity___rarg___closed__1); +l___regBuiltin_Lean_Meta_Monotonicity_evalMonotonicity__1___closed__1 = _init_l___regBuiltin_Lean_Meta_Monotonicity_evalMonotonicity__1___closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Meta_Monotonicity_evalMonotonicity__1___closed__1); +l___regBuiltin_Lean_Meta_Monotonicity_evalMonotonicity__1___closed__2 = _init_l___regBuiltin_Lean_Meta_Monotonicity_evalMonotonicity__1___closed__2(); +lean_mark_persistent(l___regBuiltin_Lean_Meta_Monotonicity_evalMonotonicity__1___closed__2); +l___regBuiltin_Lean_Meta_Monotonicity_evalMonotonicity__1___closed__3 = _init_l___regBuiltin_Lean_Meta_Monotonicity_evalMonotonicity__1___closed__3(); +lean_mark_persistent(l___regBuiltin_Lean_Meta_Monotonicity_evalMonotonicity__1___closed__3); +l___regBuiltin_Lean_Meta_Monotonicity_evalMonotonicity__1___closed__4 = _init_l___regBuiltin_Lean_Meta_Monotonicity_evalMonotonicity__1___closed__4(); +lean_mark_persistent(l___regBuiltin_Lean_Meta_Monotonicity_evalMonotonicity__1___closed__4); +l___regBuiltin_Lean_Meta_Monotonicity_evalMonotonicity__1___closed__5 = _init_l___regBuiltin_Lean_Meta_Monotonicity_evalMonotonicity__1___closed__5(); +lean_mark_persistent(l___regBuiltin_Lean_Meta_Monotonicity_evalMonotonicity__1___closed__5); +if (builtin) {res = l___regBuiltin_Lean_Meta_Monotonicity_evalMonotonicity__1(lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +}l_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_5433____closed__1 = _init_l_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_5433____closed__1(); +lean_mark_persistent(l_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_5433____closed__1); +l_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_5433____closed__2 = _init_l_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_5433____closed__2(); +lean_mark_persistent(l_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_5433____closed__2); +l_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_5433____closed__3 = _init_l_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_5433____closed__3(); +lean_mark_persistent(l_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_5433____closed__3); +l_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_5433____closed__4 = _init_l_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_5433____closed__4(); +lean_mark_persistent(l_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_5433____closed__4); +l_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_5433____closed__5 = _init_l_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_5433____closed__5(); +lean_mark_persistent(l_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_5433____closed__5); +l_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_5433____closed__6 = _init_l_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_5433____closed__6(); +lean_mark_persistent(l_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_5433____closed__6); +l_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_5433____closed__7 = _init_l_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_5433____closed__7(); +lean_mark_persistent(l_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_5433____closed__7); +l_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_5433____closed__8 = _init_l_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_5433____closed__8(); +lean_mark_persistent(l_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_5433____closed__8); +if (builtin) {res = l_initFn____x40_Lean_Elab_Tactic_Monotonicity___hyg_5433_(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/Lean/Elab/Tactic/NormCast.c b/stage0/stdlib/Lean/Elab/Tactic/NormCast.c index 79edd14162a2..f4b6f22866bc 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/NormCast.c +++ b/stage0/stdlib/Lean/Elab/Tactic/NormCast.c @@ -448,6 +448,7 @@ static lean_object* l_Lean_Elab_Tactic_NormCast_proveEqUsing___closed__4; static lean_object* l___regBuiltin_Lean_Elab_Tactic_NormCast_evalConvNormCast_declRange__1___closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_NormCast_upwardAndElim___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Meta_Simp_defaultMaxSteps; +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_NormCast_isNumeral_x3f___boxed(lean_object*); lean_object* l_Lean_Meta_Simp_Result_mkEqTrans(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__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_NormCast_elabAddElim_declRange__1___closed__4; @@ -501,6 +502,7 @@ static lean_object* l_Lean_Elab_Tactic_NormCast_derive___lambda__12___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Tactic_NormCast_evalNormCast0_declRange__1___closed__5; static lean_object* l_Lean_Elab_Tactic_NormCast_initFn____x40_Lean_Elab_Tactic_NormCast___hyg_5____closed__6; uint64_t l___private_Lean_Meta_Basic_0__Lean_Meta_Config_toKey(lean_object*); +lean_object* l_Lean_Expr_consumeMData(lean_object*); static lean_object* l_Lean_Elab_Tactic_NormCast_derive___lambda__8___closed__4; static lean_object* l_Lean_Elab_Tactic_NormCast_proveEqUsing___closed__8; lean_object* l_Lean_MessageData_ofName(lean_object*); @@ -4524,155 +4526,146 @@ x_2 = l_Lean_Elab_Tactic_NormCast_isNumeral_x3f___closed__3; x_3 = l_Lean_Expr_isConstOf(x_1, x_2); if (x_3 == 0) { -if (lean_obj_tag(x_1) == 5) -{ lean_object* x_4; -x_4 = lean_ctor_get(x_1, 0); -lean_inc(x_4); -lean_dec(x_1); +x_4 = l_Lean_Expr_consumeMData(x_1); if (lean_obj_tag(x_4) == 5) { lean_object* x_5; x_5 = lean_ctor_get(x_4, 0); lean_inc(x_5); +lean_dec(x_4); if (lean_obj_tag(x_5) == 5) { lean_object* x_6; x_6 = lean_ctor_get(x_5, 0); lean_inc(x_6); -if (lean_obj_tag(x_6) == 4) +if (lean_obj_tag(x_6) == 5) { lean_object* x_7; x_7 = lean_ctor_get(x_6, 0); lean_inc(x_7); -lean_dec(x_6); -if (lean_obj_tag(x_7) == 1) +if (lean_obj_tag(x_7) == 4) { lean_object* x_8; x_8 = lean_ctor_get(x_7, 0); lean_inc(x_8); +lean_dec(x_7); if (lean_obj_tag(x_8) == 1) { lean_object* x_9; x_9 = lean_ctor_get(x_8, 0); lean_inc(x_9); -if (lean_obj_tag(x_9) == 0) +if (lean_obj_tag(x_9) == 1) { -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; -x_10 = lean_ctor_get(x_4, 1); +lean_object* x_10; +x_10 = lean_ctor_get(x_9, 0); lean_inc(x_10); -lean_dec(x_4); +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; uint8_t x_16; x_11 = lean_ctor_get(x_5, 1); lean_inc(x_11); lean_dec(x_5); -x_12 = lean_ctor_get(x_7, 1); +x_12 = lean_ctor_get(x_6, 1); lean_inc(x_12); -lean_dec(x_7); +lean_dec(x_6); x_13 = lean_ctor_get(x_8, 1); lean_inc(x_13); lean_dec(x_8); -x_14 = l_Lean_Elab_Tactic_NormCast_isNumeral_x3f___closed__4; -x_15 = lean_string_dec_eq(x_13, x_14); -lean_dec(x_13); -if (x_15 == 0) +x_14 = lean_ctor_get(x_9, 1); +lean_inc(x_14); +lean_dec(x_9); +x_15 = l_Lean_Elab_Tactic_NormCast_isNumeral_x3f___closed__4; +x_16 = lean_string_dec_eq(x_14, x_15); +lean_dec(x_14); +if (x_16 == 0) { -lean_object* x_16; +lean_object* x_17; +lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); -lean_dec(x_10); -x_16 = lean_box(0); -return x_16; +x_17 = lean_box(0); +return x_17; } else { -lean_object* x_17; uint8_t x_18; -x_17 = l_Lean_Elab_Tactic_NormCast_isNumeral_x3f___closed__5; -x_18 = lean_string_dec_eq(x_12, x_17); -lean_dec(x_12); -if (x_18 == 0) +lean_object* x_18; uint8_t x_19; +x_18 = l_Lean_Elab_Tactic_NormCast_isNumeral_x3f___closed__5; +x_19 = lean_string_dec_eq(x_13, x_18); +lean_dec(x_13); +if (x_19 == 0) { -lean_object* x_19; +lean_object* x_20; +lean_dec(x_12); lean_dec(x_11); -lean_dec(x_10); -x_19 = lean_box(0); -return x_19; +x_20 = lean_box(0); +return x_20; } else { -if (lean_obj_tag(x_10) == 9) +if (lean_obj_tag(x_11) == 9) { -lean_object* x_20; -x_20 = lean_ctor_get(x_10, 0); -lean_inc(x_20); -lean_dec(x_10); -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_21; +x_21 = lean_ctor_get(x_11, 0); +lean_inc(x_21); +lean_dec(x_11); +if (lean_obj_tag(x_21) == 0) { -lean_object* x_22; lean_object* x_23; -x_22 = lean_ctor_get(x_20, 0); -x_23 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_23, 0, x_11); -lean_ctor_set(x_23, 1, x_22); -lean_ctor_set_tag(x_20, 1); -lean_ctor_set(x_20, 0, x_23); -return x_20; -} -else +uint8_t x_22; +x_22 = !lean_is_exclusive(x_21); +if (x_22 == 0) { -lean_object* x_24; lean_object* x_25; lean_object* x_26; -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_11); -lean_ctor_set(x_25, 1, x_24); -x_26 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_26, 0, x_25); -return x_26; -} +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_12); +lean_ctor_set(x_24, 1, x_23); +lean_ctor_set_tag(x_21, 1); +lean_ctor_set(x_21, 0, x_24); +return x_21; } else { -lean_object* x_27; -lean_dec(x_20); -lean_dec(x_11); -x_27 = lean_box(0); +lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_25 = lean_ctor_get(x_21, 0); +lean_inc(x_25); +lean_dec(x_21); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_12); +lean_ctor_set(x_26, 1, x_25); +x_27 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_27, 0, x_26); return x_27; } } else { lean_object* x_28; -lean_dec(x_11); -lean_dec(x_10); +lean_dec(x_21); +lean_dec(x_12); x_28 = lean_box(0); return x_28; } } -} -} else { lean_object* x_29; -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -lean_dec(x_4); +lean_dec(x_12); +lean_dec(x_11); x_29 = lean_box(0); return x_29; } } +} +} else { lean_object* x_30; +lean_dec(x_10); +lean_dec(x_9); lean_dec(x_8); -lean_dec(x_7); +lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); x_30 = lean_box(0); return x_30; } @@ -4680,9 +4673,10 @@ return x_30; else { lean_object* x_31; -lean_dec(x_7); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); x_31 = lean_box(0); return x_31; } @@ -4690,9 +4684,9 @@ return x_31; else { lean_object* x_32; +lean_dec(x_8); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); x_32 = lean_box(0); return x_32; } @@ -4700,8 +4694,9 @@ return x_32; else { lean_object* x_33; +lean_dec(x_7); +lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); x_33 = lean_box(0); return x_33; } @@ -4709,7 +4704,8 @@ return x_33; else { lean_object* x_34; -lean_dec(x_4); +lean_dec(x_6); +lean_dec(x_5); x_34 = lean_box(0); return x_34; } @@ -4717,7 +4713,7 @@ return x_34; else { lean_object* x_35; -lean_dec(x_1); +lean_dec(x_5); x_35 = lean_box(0); return x_35; } @@ -4725,11 +4721,27 @@ return x_35; else { lean_object* x_36; -lean_dec(x_1); -x_36 = l_Lean_Elab_Tactic_NormCast_isNumeral_x3f___closed__9; +lean_dec(x_4); +x_36 = lean_box(0); return x_36; } } +else +{ +lean_object* x_37; +x_37 = l_Lean_Elab_Tactic_NormCast_isNumeral_x3f___closed__9; +return x_37; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_NormCast_isNumeral_x3f___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Lean_Elab_Tactic_NormCast_isNumeral_x3f(x_1); +lean_dec(x_1); +return x_2; +} } LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Tactic_NormCast_splittingProcedure___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: @@ -7287,7 +7299,6 @@ if (x_73 == 0) { lean_object* x_74; lean_dec(x_70); -lean_inc(x_1); x_74 = l_Lean_Elab_Tactic_NormCast_isNumeral_x3f(x_1); if (lean_obj_tag(x_74) == 0) { @@ -7686,7 +7697,6 @@ if (x_142 == 0) { lean_object* x_143; lean_dec(x_139); -lean_inc(x_3); x_143 = l_Lean_Elab_Tactic_NormCast_isNumeral_x3f(x_3); if (lean_obj_tag(x_143) == 0) { @@ -11643,7 +11653,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_NormCast_numeralToCoe(lean_object* x _start: { lean_object* x_7; -lean_inc(x_1); x_7 = l_Lean_Elab_Tactic_NormCast_isNumeral_x3f(x_1); if (lean_obj_tag(x_7) == 0) { diff --git a/stage0/stdlib/Lean/Meta/Check.c b/stage0/stdlib/Lean/Meta/Check.c index 1c92077d57d7..693f81aa806a 100644 --- a/stage0/stdlib/Lean/Meta/Check.c +++ b/stage0/stdlib/Lean/Meta/Check.c @@ -19,7 +19,6 @@ static lean_object* l_Lean_Meta_addPPExplicitToExposeDiff_visit___lambda__4___cl LEAN_EXPORT lean_object* l_Lean_Meta_lambdaLetTelescope___at___private_Lean_Meta_Check_0__Lean_Meta_checkAux_checkLambdaLet___spec__2___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_check___lambda__1___closed__8; LEAN_EXPORT uint8_t l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Check_0__Lean_Meta_checkAux_check___spec__2(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692____closed__6; lean_object* l_Lean_getConstInfo___at_Lean_Meta_mkConstWithFreshMVarLevels___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__8; lean_object* l_Lean_mkAppN(lean_object*, lean_object*); @@ -69,7 +68,6 @@ lean_object* lean_array_fget(lean_object*, lean_object*); lean_object* l_Lean_Expr_setAppPPExplicit(lean_object*); lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_addPPExplicitToExposeDiff_visit___lambda__3___closed__1; -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692____closed__7; static lean_object* l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__7; lean_object* l_Lean_Expr_fvarId_x21(lean_object*); static lean_object* l_Lean_Meta_check___lambda__1___closed__2; @@ -78,10 +76,10 @@ uint8_t lean_float_decLt(double, double); static lean_object* l_panic___at_Lean_Meta_throwLetTypeMismatchMessage___spec__1___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_addPPExplicitToExposeDiff_visit___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_addPPExplicitToExposeDiff_visit___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_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____closed__14; LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Meta_check___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_check___closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_check___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692____closed__14; lean_object* lean_io_get_num_heartbeats(lean_object*); static double l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___lambda__2___closed__1; lean_object* l_Nat_nextPowerOfTwo_go(lean_object*, lean_object*, lean_object*); @@ -90,13 +88,15 @@ lean_object* l_Lean_stringToMessageData(lean_object*); static lean_object* l_Lean_Meta_throwAppTypeMismatch___rarg___lambda__1___closed__6; static lean_object* l_Lean_Meta_addPPExplicitToExposeDiff_visit___closed__2; static lean_object* l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__1; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____closed__11; LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_lambdaLetTelescope___at___private_Lean_Meta_Check_0__Lean_Meta_checkAux_checkLambdaLet___spec__2___rarg(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___lambda__4___closed__2; lean_object* l_Lean_throwError___at_Lean_Meta_setInlineAttribute___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Exception_toMessageData(lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____closed__2; lean_object* l_Lean_Meta_throwFunctionExpected___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692____closed__4; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____closed__10; LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Meta_Check_0__Lean_Meta_checkAux_check___spec__5(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_addPPExplicitToExposeDiff_visit___lambda__6(lean_object*, 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*); size_t lean_ptr_addr(lean_object*); @@ -105,9 +105,10 @@ LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_throwAppTypeMismatch__ static double l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___lambda__4___closed__5; static lean_object* l_Lean_Meta_addPPExplicitToExposeDiff___closed__3; size_t lean_usize_of_nat(lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692____closed__9; lean_object* lean_expr_abstract(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____closed__6; lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_lambdaTelescopeImp___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_Check___hyg_4731____closed__1; static double l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___lambda__4___closed__4; LEAN_EXPORT lean_object* l___private_Lean_Meta_Check_0__Lean_Meta_ensureType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_local_ctx_find(lean_object*, lean_object*); @@ -118,11 +119,9 @@ static lean_object* l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__9; uint64_t lean_uint64_shift_right(uint64_t, uint64_t); lean_object* lean_nat_div(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_throwLetTypeMismatchMessage(lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692____closed__3; lean_object* l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(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*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692____closed__11; static lean_object* l_Lean_Meta_addPPExplicitToExposeDiff___closed__2; LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at_Lean_Meta_addPPExplicitToExposeDiff_visit___spec__2___boxed(lean_object*, lean_object*); lean_object* l_Lean_Expr_setPPExplicit(lean_object*, uint8_t); @@ -144,23 +143,23 @@ LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_throwLetTypeMismatchMe extern lean_object* l_Lean_checkEmoji; lean_object* lean_st_ref_get(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_throwAppTypeMismatch(lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692____closed__5; lean_object* lean_st_mk_ref(lean_object*, lean_object*); static lean_object* l_Lean_Meta_throwAppTypeMismatch___rarg___lambda__1___closed__4; lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_addPPExplicitToExposeDiff_visit___spec__1___lambda__1___boxed(lean_object**); LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_addPPExplicitToExposeDiff_visit___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____closed__7; LEAN_EXPORT lean_object* l_panic___at_Lean_Meta_throwLetTypeMismatchMessage___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_throwAppTypeMismatch___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_levelZero; lean_object* lean_io_mono_nanos_now(lean_object*); extern lean_object* l_Lean_instInhabitedExpr; -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692____closed__2; lean_object* l_Lean_Meta_whnfD(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint64_t l_Lean_Expr_hash(lean_object*); lean_object* l_Lean_Meta_isLabeledSorry_x3f(lean_object*); lean_object* l_List_lengthTRAux___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_throwLetTypeMismatchMessage___spec__3(lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____closed__13; static lean_object* l_Lean_Meta_addPPExplicitToExposeDiff_visit___lambda__4___closed__5; LEAN_EXPORT lean_object* l_Lean_Meta_lambdaLetTelescope___at___private_Lean_Meta_Check_0__Lean_Meta_checkAux_checkLambdaLet___spec__2(lean_object*); uint8_t lean_name_eq(lean_object*, lean_object*); @@ -177,6 +176,7 @@ lean_object* l_Lean_Meta_saveState___rarg(lean_object*, lean_object*, lean_objec LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescope___at___private_Lean_Meta_Check_0__Lean_Meta_checkAux_checkForall___spec__2(lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_check___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_checkProj___closed__3; +LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731_(lean_object*); static lean_object* l_Lean_Meta_addPPExplicitToExposeDiff_visit___lambda__4___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_Check_0__Lean_Meta_checkAux_checkForall___lambda__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_check___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*); @@ -194,13 +194,14 @@ LEAN_EXPORT lean_object* l_Lean_withoutModifyingState___at_Lean_Meta_addPPExplic lean_object* l_Lean_MessageData_ofExpr(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Check_0__Lean_Meta_checkAux_checkLambdaLet___lambda__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_forallTelescopeReducingAuxAux___rarg(uint8_t, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____closed__3; double l_Float_ofScientific(lean_object*, uint8_t, lean_object*); extern lean_object* l_Lean_crossEmoji; LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at_Lean_Meta_addPPExplicitToExposeDiff_visit___spec__3___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692_(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Check_0__Lean_Meta_checkAux(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_addPPExplicitToExposeDiff(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_addPPExplicitToExposeDiff___lambda__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_Check___hyg_4731____closed__4; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Check_0__Lean_Meta_checkAux_checkLambdaLet___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_addPPExplicitToExposeDiff_visit___closed__1; static lean_object* l_Lean_Meta_addPPExplicitToExposeDiff___closed__4; @@ -210,8 +211,10 @@ lean_object* l_Lean_Expr_app___override(lean_object*, lean_object*); static lean_object* l_Lean_Meta_throwAppTypeMismatch___rarg___lambda__1___closed__1; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Check_0__Lean_Meta_checkAux_checkLambdaLet(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____closed__8; lean_object* l_Lean_Meta_getFVarLocalDecl(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_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____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_Lean_Meta_check___lambda__1___closed__4; static lean_object* l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__5; @@ -226,7 +229,6 @@ static lean_object* l_Lean_Meta_check___lambda__1___closed__7; LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescope___at___private_Lean_Meta_Check_0__Lean_Meta_checkAux_checkForall___spec__2___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Check_0__Lean_Meta_checkAux_checkForall___closed__1; uint8_t l___private_Lean_Data_DeclarationRange_0__Lean_decEqDeclarationLocation____x40_Lean_Data_DeclarationRange___hyg_600_(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692____closed__10; LEAN_EXPORT lean_object* l_Lean_Meta_mkHasTypeButIsExpectedMsg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_checkProj___closed__2; static lean_object* l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___lambda__4___closed__3; @@ -243,7 +245,6 @@ lean_object* l_Lean_Expr_setOption___at_Lean_Expr_setPPExplicit___spec__1(lean_o static lean_object* l_Lean_Meta_check___lambda__1___closed__6; LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_addPPExplicitToExposeDiff_visit___spec__1___lambda__2___boxed(lean_object**); lean_object* lean_nat_sub(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692____closed__8; static lean_object* l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___lambda__3___closed__1; lean_object* l_Lean_Expr_getAppFn(lean_object*); static lean_object* l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___lambda__4___closed__1; @@ -274,12 +275,11 @@ lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at___private_Lean_Meta_Check_0__Lean_Meta_checkAux_check___spec__1___boxed(lean_object*, lean_object*); lean_object* l_Lean_Expr_setPPFunBinderTypes(lean_object*, uint8_t); lean_object* l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692____closed__12; -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692____closed__13; static lean_object* l_Lean_Meta_throwAppTypeMismatch___rarg___lambda__1___closed__5; lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_lambdaLetTelescope___at___private_Lean_Meta_Check_0__Lean_Meta_checkAux_checkLambdaLet___spec__2___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Meta_addPPExplicitToExposeDiff_visit___spec__4___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_Check___hyg_4731____closed__9; lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_checkProj___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_check___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -293,6 +293,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_addPPExplicitToExposeDiff_visit___lambda__3 uint64_t l_Lean_Meta_TransparencyMode_toUInt64(uint8_t); lean_object* lean_nat_add(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Check_0__Lean_Meta_checkAux_checkForall___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____closed__5; LEAN_EXPORT lean_object* l___private_Lean_Meta_Check_0__Lean_Meta_checkAux_checkForall___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_withTraceNode___at_Lean_Meta_check___spec__1___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Exception_isRuntime(lean_object*); @@ -307,7 +308,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Check_0__Lean_Meta_checkAux_check LEAN_EXPORT lean_object* l_Lean_Meta_addPPExplicitToExposeDiff_visit___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_uset(lean_object*, size_t, lean_object*); static lean_object* l___private_Lean_Meta_Check_0__Lean_Meta_checkAux___closed__1; -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692____closed__1; lean_object* lean_expr_instantiate1(lean_object*, lean_object*); static lean_object* l_Lean_Meta_mkHasTypeButIsExpectedMsg___closed__3; LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand_go___at___private_Lean_Meta_Check_0__Lean_Meta_checkAux_check___spec__4(lean_object*, lean_object*, lean_object*); @@ -371,634 +371,304 @@ return x_17; } } } -static lean_object* _init_l_panic___at_Lean_Meta_throwLetTypeMismatchMessage___spec__1___rarg___closed__1() { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Check_0__Lean_Meta_checkConstant(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Meta_instInhabitedMetaM___boxed), 5, 1); -lean_closure_set(x_1, 0, lean_box(0)); -return x_1; -} -} -LEAN_EXPORT lean_object* l_panic___at_Lean_Meta_throwLetTypeMismatchMessage___spec__1___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: +lean_object* x_8; +lean_inc(x_1); +x_8 = l_Lean_getConstInfo___at_Lean_Meta_mkConstWithFreshMVarLevels___spec__1(x_1, x_3, x_4, x_5, x_6, x_7); +if (lean_obj_tag(x_8) == 0) { -lean_object* x_7; lean_object* x_8; lean_object* x_9; -x_7 = l_panic___at_Lean_Meta_throwLetTypeMismatchMessage___spec__1___rarg___closed__1; -x_8 = lean_panic_fn(x_7, x_1); -x_9 = lean_apply_5(x_8, x_2, x_3, x_4, x_5, x_6); -return x_9; -} +uint8_t x_9; +x_9 = !lean_is_exclusive(x_8); +if (x_9 == 0) +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_10 = lean_ctor_get(x_8, 0); +x_11 = lean_ctor_get(x_8, 1); +x_12 = lean_unsigned_to_nat(0u); +x_13 = l_List_lengthTRAux___rarg(x_2, x_12); +x_14 = l_Lean_ConstantInfo_levelParams(x_10); +lean_dec(x_10); +x_15 = l_List_lengthTRAux___rarg(x_14, x_12); +lean_dec(x_14); +x_16 = lean_nat_dec_eq(x_13, x_15); +lean_dec(x_15); +lean_dec(x_13); +if (x_16 == 0) +{ +lean_object* x_17; +lean_free_object(x_8); +x_17 = l_Lean_Meta_throwIncorrectNumberOfLevels___rarg(x_1, x_2, x_3, x_4, x_5, x_6, x_11); +return x_17; } -LEAN_EXPORT lean_object* l_panic___at_Lean_Meta_throwLetTypeMismatchMessage___spec__1(lean_object* x_1) { -_start: +else { -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_panic___at_Lean_Meta_throwLetTypeMismatchMessage___spec__1___rarg), 6, 0); -return x_2; +lean_object* x_18; +lean_dec(x_2); +lean_dec(x_1); +x_18 = lean_box(0); +lean_ctor_set(x_8, 0, x_18); +return x_8; } } -LEAN_EXPORT lean_object* l_panic___at_Lean_Meta_throwLetTypeMismatchMessage___spec__2___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: +else { -lean_object* x_7; lean_object* x_8; lean_object* x_9; -x_7 = l_panic___at_Lean_Meta_throwLetTypeMismatchMessage___spec__1___rarg___closed__1; -x_8 = lean_panic_fn(x_7, x_1); -x_9 = lean_apply_5(x_8, x_2, x_3, x_4, x_5, x_6); -return x_9; -} +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; +x_19 = lean_ctor_get(x_8, 0); +x_20 = lean_ctor_get(x_8, 1); +lean_inc(x_20); +lean_inc(x_19); +lean_dec(x_8); +x_21 = lean_unsigned_to_nat(0u); +x_22 = l_List_lengthTRAux___rarg(x_2, x_21); +x_23 = l_Lean_ConstantInfo_levelParams(x_19); +lean_dec(x_19); +x_24 = l_List_lengthTRAux___rarg(x_23, x_21); +lean_dec(x_23); +x_25 = lean_nat_dec_eq(x_22, x_24); +lean_dec(x_24); +lean_dec(x_22); +if (x_25 == 0) +{ +lean_object* x_26; +x_26 = l_Lean_Meta_throwIncorrectNumberOfLevels___rarg(x_1, x_2, x_3, x_4, x_5, x_6, x_20); +return x_26; } -LEAN_EXPORT lean_object* l_panic___at_Lean_Meta_throwLetTypeMismatchMessage___spec__2(lean_object* x_1) { -_start: +else { -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_panic___at_Lean_Meta_throwLetTypeMismatchMessage___spec__2___rarg), 6, 0); -return x_2; +lean_object* x_27; lean_object* x_28; +lean_dec(x_2); +lean_dec(x_1); +x_27 = lean_box(0); +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_27); +lean_ctor_set(x_28, 1, x_20); +return x_28; } } -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_throwLetTypeMismatchMessage___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) { -_start: +} +else { -lean_object* x_7; lean_object* x_8; uint8_t x_9; -x_7 = lean_ctor_get(x_4, 5); -x_8 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_1, x_2, x_3, x_4, x_5, x_6); -x_9 = !lean_is_exclusive(x_8); -if (x_9 == 0) +uint8_t x_29; +lean_dec(x_2); +lean_dec(x_1); +x_29 = !lean_is_exclusive(x_8); +if (x_29 == 0) { -lean_object* x_10; lean_object* x_11; -x_10 = lean_ctor_get(x_8, 0); -lean_inc(x_7); -x_11 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_11, 0, x_7); -lean_ctor_set(x_11, 1, x_10); -lean_ctor_set_tag(x_8, 1); -lean_ctor_set(x_8, 0, x_11); return x_8; } else { -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_12 = lean_ctor_get(x_8, 0); -x_13 = lean_ctor_get(x_8, 1); -lean_inc(x_13); -lean_inc(x_12); +lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_30 = lean_ctor_get(x_8, 0); +x_31 = lean_ctor_get(x_8, 1); +lean_inc(x_31); +lean_inc(x_30); lean_dec(x_8); -lean_inc(x_7); -x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_7); -lean_ctor_set(x_14, 1, x_12); -x_15 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_15, 0, x_14); -lean_ctor_set(x_15, 1, x_13); -return x_15; -} +x_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_throwError___at_Lean_Meta_throwLetTypeMismatchMessage___spec__3(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_throwError___at_Lean_Meta_throwLetTypeMismatchMessage___spec__3___rarg___boxed), 6, 0); -return x_2; } } -static lean_object* _init_l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__1() { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Check_0__Lean_Meta_checkConstant___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean.Meta.Check", 15, 15); -return x_1; +lean_object* x_8; +x_8 = l___private_Lean_Meta_Check_0__Lean_Meta_checkConstant(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; } } -static lean_object* _init_l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__2() { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Check_0__Lean_Meta_getFunctionDomain(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean.Meta.throwLetTypeMismatchMessage", 37, 37); -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__3() { -_start: +lean_object* x_7; +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_7 = lean_infer_type(x_1, x_2, x_3, x_4, x_5, x_6); +if (lean_obj_tag(x_7) == 0) { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("unreachable code has been reached", 33, 33); -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__4() { -_start: +lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_8 = lean_ctor_get(x_7, 0); +lean_inc(x_8); +x_9 = lean_ctor_get(x_7, 1); +lean_inc(x_9); +lean_dec(x_7); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +x_10 = l_Lean_Meta_whnfD(x_8, x_2, x_3, x_4, x_5, x_9); +if (lean_obj_tag(x_10) == 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_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__1; -x_2 = l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__2; -x_3 = lean_unsigned_to_nat(26u); -x_4 = lean_unsigned_to_nat(9u); -x_5 = l_Lean_Meta_throwLetTypeMismatchMessage___rarg___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_11; +x_11 = lean_ctor_get(x_10, 0); +lean_inc(x_11); +if (lean_obj_tag(x_11) == 7) +{ +uint8_t x_12; +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_12 = !lean_is_exclusive(x_10); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; lean_object* x_17; +x_13 = lean_ctor_get(x_10, 0); +lean_dec(x_13); +x_14 = lean_ctor_get(x_11, 1); +lean_inc(x_14); +x_15 = lean_ctor_get_uint8(x_11, sizeof(void*)*3 + 8); +lean_dec(x_11); +x_16 = lean_box(x_15); +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_14); +lean_ctor_set(x_17, 1, x_16); +lean_ctor_set(x_10, 0, x_17); +return x_10; } -static lean_object* _init_l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__5() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("invalid let declaration, term", 29, 29); -return x_1; +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_10, 1); +lean_inc(x_18); +lean_dec(x_10); +x_19 = lean_ctor_get(x_11, 1); +lean_inc(x_19); +x_20 = lean_ctor_get_uint8(x_11, sizeof(void*)*3 + 8); +lean_dec(x_11); +x_21 = lean_box(x_20); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_19); +lean_ctor_set(x_22, 1, x_21); +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_18); +return x_23; } } -static lean_object* _init_l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__6() { -_start: +else { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__5; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__7() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("\nhas type", 9, 9); -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__8() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__7; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__9() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("\nbut is expected to have type", 29, 29); -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__10() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__9; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__11() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("", 0, 0); -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__12() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__11; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_throwLetTypeMismatchMessage___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; -x_7 = lean_ctor_get(x_2, 2); -lean_inc(x_7); -x_8 = lean_local_ctx_find(x_7, x_1); -if (lean_obj_tag(x_8) == 0) -{ -lean_object* x_9; lean_object* x_10; -x_9 = l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__4; -x_10 = l_panic___at_Lean_Meta_throwLetTypeMismatchMessage___spec__1___rarg(x_9, x_2, x_3, x_4, x_5, x_6); -return x_10; -} -else -{ -lean_object* x_11; -x_11 = lean_ctor_get(x_8, 0); -lean_inc(x_11); -lean_dec(x_8); -if (lean_obj_tag(x_11) == 0) -{ -lean_object* x_12; lean_object* x_13; -lean_dec(x_11); -x_12 = l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__4; -x_13 = l_panic___at_Lean_Meta_throwLetTypeMismatchMessage___spec__2___rarg(x_12, x_2, x_3, x_4, x_5, x_6); -return x_13; -} -else -{ -lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_14 = lean_ctor_get(x_11, 3); -lean_inc(x_14); -x_15 = lean_ctor_get(x_11, 4); -lean_inc(x_15); +lean_object* x_24; lean_object* x_25; lean_dec(x_11); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -lean_inc(x_2); -lean_inc(x_15); -x_16 = lean_infer_type(x_15, x_2, x_3, x_4, x_5, x_6); -if (lean_obj_tag(x_16) == 0) -{ -lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; -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_indentExpr(x_15); -x_20 = l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__6; -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_Meta_throwLetTypeMismatchMessage___rarg___closed__8; -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_indentExpr(x_17); -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_Meta_throwLetTypeMismatchMessage___rarg___closed__10; -x_27 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_27, 0, x_25); -lean_ctor_set(x_27, 1, x_26); -x_28 = l_Lean_indentExpr(x_14); -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_Meta_throwLetTypeMismatchMessage___rarg___closed__12; -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_Meta_throwLetTypeMismatchMessage___spec__3___rarg(x_31, x_2, x_3, x_4, x_5, x_18); +x_24 = lean_ctor_get(x_10, 1); +lean_inc(x_24); +lean_dec(x_10); +x_25 = l_Lean_Meta_throwFunctionExpected___rarg(x_1, x_2, x_3, x_4, x_5, x_24); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -return x_32; +return x_25; +} } else { -uint8_t x_33; -lean_dec(x_15); -lean_dec(x_14); +uint8_t x_26; lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_33 = !lean_is_exclusive(x_16); -if (x_33 == 0) +lean_dec(x_1); +x_26 = !lean_is_exclusive(x_10); +if (x_26 == 0) { -return x_16; +return x_10; } else { -lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_34 = lean_ctor_get(x_16, 0); -x_35 = lean_ctor_get(x_16, 1); -lean_inc(x_35); -lean_inc(x_34); -lean_dec(x_16); -x_36 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_36, 0, x_34); -lean_ctor_set(x_36, 1, x_35); -return x_36; -} -} -} -} -} +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_10, 0); +x_28 = lean_ctor_get(x_10, 1); +lean_inc(x_28); +lean_inc(x_27); +lean_dec(x_10); +x_29 = 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_throwLetTypeMismatchMessage(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_Meta_throwLetTypeMismatchMessage___rarg), 6, 0); -return x_2; } } -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_throwLetTypeMismatchMessage___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: +else { -lean_object* x_7; -x_7 = l_Lean_throwError___at_Lean_Meta_throwLetTypeMismatchMessage___spec__3___rarg(x_1, x_2, x_3, x_4, x_5, x_6); +uint8_t x_30; lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); +lean_dec(x_1); +x_30 = !lean_is_exclusive(x_7); +if (x_30 == 0) +{ return x_7; } +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_7, 0); +x_32 = lean_ctor_get(x_7, 1); +lean_inc(x_32); +lean_inc(x_31); +lean_dec(x_7); +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___private_Lean_Meta_Check_0__Lean_Meta_checkConstant(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +} +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_addPPExplicitToExposeDiff_visit___spec__1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17) { _start: { -lean_object* x_8; -lean_inc(x_1); -x_8 = l_Lean_getConstInfo___at_Lean_Meta_mkConstWithFreshMVarLevels___spec__1(x_1, x_3, x_4, x_5, x_6, x_7); -if (lean_obj_tag(x_8) == 0) +if (lean_obj_tag(x_8) == 7) { -uint8_t x_9; -x_9 = !lean_is_exclusive(x_8); -if (x_9 == 0) +if (lean_obj_tag(x_9) == 7) { -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_8, 0); -x_11 = lean_ctor_get(x_8, 1); -x_12 = lean_unsigned_to_nat(0u); -x_13 = l_List_lengthTRAux___rarg(x_2, x_12); -x_14 = l_Lean_ConstantInfo_levelParams(x_10); -lean_dec(x_10); -x_15 = l_List_lengthTRAux___rarg(x_14, x_12); -lean_dec(x_14); -x_16 = lean_nat_dec_eq(x_13, x_15); -lean_dec(x_15); -lean_dec(x_13); -if (x_16 == 0) +lean_object* x_18; uint8_t x_19; lean_object* x_20; uint8_t x_21; uint8_t x_22; lean_object* x_23; uint8_t x_24; lean_object* x_25; +lean_dec(x_2); +lean_dec(x_1); +x_18 = lean_ctor_get(x_8, 2); +lean_inc(x_18); +x_19 = lean_ctor_get_uint8(x_8, sizeof(void*)*3 + 8); +lean_dec(x_8); +x_20 = lean_ctor_get(x_9, 2); +lean_inc(x_20); +x_21 = lean_ctor_get_uint8(x_9, sizeof(void*)*3 + 8); +lean_dec(x_9); +x_22 = lean_nat_dec_lt(x_3, x_4); +x_23 = lean_array_get_size(x_5); +x_24 = lean_nat_dec_lt(x_3, x_23); +lean_dec(x_23); +if (x_22 == 0) { -lean_object* x_17; -lean_free_object(x_8); -x_17 = l_Lean_Meta_throwIncorrectNumberOfLevels___rarg(x_1, x_2, x_3, x_4, x_5, x_6, x_11); -return x_17; +lean_object* x_167; lean_object* x_168; +x_167 = l_Lean_instInhabitedExpr; +x_168 = l_outOfBounds___rarg(x_167); +x_25 = x_168; +goto block_166; } else { -lean_object* x_18; -lean_dec(x_2); -lean_dec(x_1); -x_18 = lean_box(0); -lean_ctor_set(x_8, 0, x_18); -return x_8; -} +lean_object* x_169; +x_169 = lean_array_fget(x_7, x_3); +x_25 = x_169; +goto block_166; } -else -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; -x_19 = lean_ctor_get(x_8, 0); -x_20 = lean_ctor_get(x_8, 1); -lean_inc(x_20); -lean_inc(x_19); -lean_dec(x_8); -x_21 = lean_unsigned_to_nat(0u); -x_22 = l_List_lengthTRAux___rarg(x_2, x_21); -x_23 = l_Lean_ConstantInfo_levelParams(x_19); -lean_dec(x_19); -x_24 = l_List_lengthTRAux___rarg(x_23, x_21); -lean_dec(x_23); -x_25 = lean_nat_dec_eq(x_22, x_24); -lean_dec(x_24); -lean_dec(x_22); -if (x_25 == 0) -{ -lean_object* x_26; -x_26 = l_Lean_Meta_throwIncorrectNumberOfLevels___rarg(x_1, x_2, x_3, x_4, x_5, x_6, x_20); -return x_26; -} -else -{ -lean_object* x_27; lean_object* x_28; -lean_dec(x_2); -lean_dec(x_1); -x_27 = lean_box(0); -x_28 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_28, 0, x_27); -lean_ctor_set(x_28, 1, x_20); -return x_28; -} -} -} -else -{ -uint8_t x_29; -lean_dec(x_2); -lean_dec(x_1); -x_29 = !lean_is_exclusive(x_8); -if (x_29 == 0) -{ -return x_8; -} -else -{ -lean_object* x_30; lean_object* x_31; lean_object* x_32; -x_30 = lean_ctor_get(x_8, 0); -x_31 = lean_ctor_get(x_8, 1); -lean_inc(x_31); -lean_inc(x_30); -lean_dec(x_8); -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___private_Lean_Meta_Check_0__Lean_Meta_checkConstant___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: -{ -lean_object* x_8; -x_8 = l___private_Lean_Meta_Check_0__Lean_Meta_checkConstant(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_Check_0__Lean_Meta_getFunctionDomain(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -lean_object* x_7; -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -lean_inc(x_2); -lean_inc(x_1); -x_7 = lean_infer_type(x_1, x_2, x_3, x_4, x_5, x_6); -if (lean_obj_tag(x_7) == 0) -{ -lean_object* x_8; lean_object* x_9; lean_object* x_10; -x_8 = lean_ctor_get(x_7, 0); -lean_inc(x_8); -x_9 = lean_ctor_get(x_7, 1); -lean_inc(x_9); -lean_dec(x_7); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -lean_inc(x_2); -x_10 = l_Lean_Meta_whnfD(x_8, x_2, x_3, x_4, x_5, x_9); -if (lean_obj_tag(x_10) == 0) -{ -lean_object* x_11; -x_11 = lean_ctor_get(x_10, 0); -lean_inc(x_11); -if (lean_obj_tag(x_11) == 7) -{ -uint8_t x_12; -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_12 = !lean_is_exclusive(x_10); -if (x_12 == 0) -{ -lean_object* x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; lean_object* x_17; -x_13 = lean_ctor_get(x_10, 0); -lean_dec(x_13); -x_14 = lean_ctor_get(x_11, 1); -lean_inc(x_14); -x_15 = lean_ctor_get_uint8(x_11, sizeof(void*)*3 + 8); -lean_dec(x_11); -x_16 = lean_box(x_15); -x_17 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_17, 0, x_14); -lean_ctor_set(x_17, 1, x_16); -lean_ctor_set(x_10, 0, x_17); -return x_10; -} -else -{ -lean_object* x_18; lean_object* x_19; uint8_t x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_18 = lean_ctor_get(x_10, 1); -lean_inc(x_18); -lean_dec(x_10); -x_19 = lean_ctor_get(x_11, 1); -lean_inc(x_19); -x_20 = lean_ctor_get_uint8(x_11, sizeof(void*)*3 + 8); -lean_dec(x_11); -x_21 = lean_box(x_20); -x_22 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_22, 0, x_19); -lean_ctor_set(x_22, 1, x_21); -x_23 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_23, 0, x_22); -lean_ctor_set(x_23, 1, x_18); -return x_23; -} -} -else -{ -lean_object* x_24; lean_object* x_25; -lean_dec(x_11); -x_24 = lean_ctor_get(x_10, 1); -lean_inc(x_24); -lean_dec(x_10); -x_25 = l_Lean_Meta_throwFunctionExpected___rarg(x_1, x_2, x_3, x_4, x_5, x_24); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_25; -} -} -else -{ -uint8_t x_26; -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_26 = !lean_is_exclusive(x_10); -if (x_26 == 0) -{ -return x_10; -} -else -{ -lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_27 = lean_ctor_get(x_10, 0); -x_28 = lean_ctor_get(x_10, 1); -lean_inc(x_28); -lean_inc(x_27); -lean_dec(x_10); -x_29 = 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_30; -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_30 = !lean_is_exclusive(x_7); -if (x_30 == 0) -{ -return x_7; -} -else -{ -lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_31 = lean_ctor_get(x_7, 0); -x_32 = lean_ctor_get(x_7, 1); -lean_inc(x_32); -lean_inc(x_31); -lean_dec(x_7); -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_Std_Range_forIn_x27_loop___at_Lean_Meta_addPPExplicitToExposeDiff_visit___spec__1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17) { -_start: -{ -if (lean_obj_tag(x_8) == 7) -{ -if (lean_obj_tag(x_9) == 7) -{ -lean_object* x_18; uint8_t x_19; lean_object* x_20; uint8_t x_21; uint8_t x_22; lean_object* x_23; uint8_t x_24; lean_object* x_25; -lean_dec(x_2); -lean_dec(x_1); -x_18 = lean_ctor_get(x_8, 2); -lean_inc(x_18); -x_19 = lean_ctor_get_uint8(x_8, sizeof(void*)*3 + 8); -lean_dec(x_8); -x_20 = lean_ctor_get(x_9, 2); -lean_inc(x_20); -x_21 = lean_ctor_get_uint8(x_9, sizeof(void*)*3 + 8); -lean_dec(x_9); -x_22 = lean_nat_dec_lt(x_3, x_4); -x_23 = lean_array_get_size(x_5); -x_24 = lean_nat_dec_lt(x_3, x_23); -lean_dec(x_23); -if (x_22 == 0) -{ -lean_object* x_167; lean_object* x_168; -x_167 = l_Lean_instInhabitedExpr; -x_168 = l_outOfBounds___rarg(x_167); -x_25 = x_168; -goto block_166; -} -else -{ -lean_object* x_169; -x_169 = lean_array_fget(x_7, x_3); -x_25 = x_169; -goto block_166; -} -block_166: +block_166: { lean_object* x_26; lean_object* x_27; x_26 = lean_expr_instantiate1(x_18, x_25); @@ -6144,538 +5814,956 @@ goto block_18; } else { -lean_object* x_628; lean_object* x_629; -lean_dec(x_589); -lean_dec(x_588); -x_628 = lean_ctor_get(x_590, 0); -lean_inc(x_628); -x_629 = lean_ctor_get(x_590, 1); -lean_inc(x_629); -lean_dec(x_590); -x_19 = x_628; -x_20 = x_629; -goto block_28; +lean_object* x_628; lean_object* x_629; +lean_dec(x_589); +lean_dec(x_588); +x_628 = lean_ctor_get(x_590, 0); +lean_inc(x_628); +x_629 = lean_ctor_get(x_590, 1); +lean_inc(x_629); +lean_dec(x_590); +x_19 = x_628; +x_20 = x_629; +goto block_28; +} +} +else +{ +lean_object* x_630; lean_object* x_631; lean_object* x_632; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_630 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_630, 0, x_1); +lean_ctor_set(x_630, 1, x_2); +x_631 = lean_box(0); +x_632 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_632, 0, x_630); +lean_ctor_set(x_632, 1, x_631); +x_8 = x_632; +x_9 = x_7; +goto block_18; +} +} +} +block_18: +{ +if (lean_obj_tag(x_8) == 0) +{ +uint8_t x_10; +x_10 = !lean_is_exclusive(x_8); +if (x_10 == 0) +{ +lean_object* x_11; +x_11 = lean_ctor_get(x_8, 1); +lean_dec(x_11); +lean_ctor_set(x_8, 1, x_9); +return x_8; +} +else +{ +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_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_9); +return x_13; +} +} +else +{ +uint8_t x_14; +x_14 = !lean_is_exclusive(x_8); +if (x_14 == 0) +{ +lean_object* x_15; +x_15 = lean_ctor_get(x_8, 1); +lean_dec(x_15); +lean_ctor_set_tag(x_8, 0); +lean_ctor_set(x_8, 1, x_9); +return x_8; +} +else +{ +lean_object* x_16; lean_object* x_17; +x_16 = lean_ctor_get(x_8, 0); +lean_inc(x_16); +lean_dec(x_8); +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_16); +lean_ctor_set(x_17, 1, x_9); +return x_17; +} +} +} +block_28: +{ +uint8_t x_21; +x_21 = l_Lean_Exception_isInterrupt(x_19); +if (x_21 == 0) +{ +uint8_t x_22; +x_22 = l_Lean_Exception_isRuntime(x_19); +if (x_22 == 0) +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; +lean_dec(x_19); +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_1); +lean_ctor_set(x_23, 1, x_2); +x_24 = lean_box(0); +x_25 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_24); +x_8 = x_25; +x_9 = x_20; +goto block_18; +} +else +{ +lean_object* x_26; +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; +} +} +else +{ +lean_object* x_27; +lean_dec(x_2); +lean_dec(x_1); +x_27 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_27, 0, x_19); +lean_ctor_set(x_27, 1, x_20); +return x_27; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_addPPExplicitToExposeDiff_visit___spec__1___lambda__1___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +_start: +{ +lean_object* x_18; +x_18 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_addPPExplicitToExposeDiff_visit___spec__1___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +lean_dec(x_12); +lean_dec(x_7); +lean_dec(x_5); +lean_dec(x_4); +return x_18; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_addPPExplicitToExposeDiff_visit___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]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +_start: +{ +lean_object* x_18; +x_18 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_addPPExplicitToExposeDiff_visit___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, x_13, x_14, x_15, x_16, x_17); +lean_dec(x_12); +lean_dec(x_7); +lean_dec(x_5); +lean_dec(x_4); +return x_18; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_addPPExplicitToExposeDiff_visit___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]; +_start: +{ +lean_object* x_18; +x_18 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_addPPExplicitToExposeDiff_visit___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_18; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at_Lean_Meta_addPPExplicitToExposeDiff_visit___spec__2___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at_Lean_Meta_addPPExplicitToExposeDiff_visit___spec__2(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at_Lean_Meta_addPPExplicitToExposeDiff_visit___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_159____at_Lean_Meta_addPPExplicitToExposeDiff_visit___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_Lean_Meta_withLocalDecl___at_Lean_Meta_addPPExplicitToExposeDiff_visit___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: +{ +uint8_t x_11; uint8_t x_12; lean_object* x_13; +x_11 = lean_unbox(x_2); +lean_dec(x_2); +x_12 = lean_unbox(x_5); +lean_dec(x_5); +x_13 = l_Lean_Meta_withLocalDecl___at_Lean_Meta_addPPExplicitToExposeDiff_visit___spec__4___rarg(x_1, x_11, x_3, x_4, x_12, x_6, x_7, x_8, x_9, x_10); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_addPPExplicitToExposeDiff_visit___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_Meta_addPPExplicitToExposeDiff_visit___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_addPPExplicitToExposeDiff_visit___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +x_13 = l_Lean_Meta_addPPExplicitToExposeDiff_visit___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_7); +lean_dec(x_4); +lean_dec(x_3); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_addPPExplicitToExposeDiff_visit___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +_start: +{ +lean_object* x_16; +x_16 = l_Lean_Meta_addPPExplicitToExposeDiff_visit___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +lean_dec(x_10); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_16; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_addPPExplicitToExposeDiff_visit___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +_start: +{ +lean_object* x_16; +x_16 = l_Lean_Meta_addPPExplicitToExposeDiff_visit___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +lean_dec(x_10); +lean_dec(x_4); +lean_dec(x_3); +return x_16; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_addPPExplicitToExposeDiff_visit___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: +{ +uint8_t x_15; uint8_t x_16; lean_object* x_17; +x_15 = lean_unbox(x_5); +lean_dec(x_5); +x_16 = lean_unbox(x_8); +lean_dec(x_8); +x_17 = l_Lean_Meta_addPPExplicitToExposeDiff_visit___lambda__5(x_1, x_2, x_3, x_4, x_15, x_6, x_7, x_16, x_9, x_10, x_11, x_12, x_13, x_14); +lean_dec(x_2); +lean_dec(x_1); +return x_17; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_addPPExplicitToExposeDiff_visit___lambda__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +uint8_t x_15; uint8_t x_16; lean_object* x_17; +x_15 = lean_unbox(x_5); +lean_dec(x_5); +x_16 = lean_unbox(x_8); +lean_dec(x_8); +x_17 = l_Lean_Meta_addPPExplicitToExposeDiff_visit___lambda__6(x_1, x_2, x_3, x_4, x_15, x_6, x_7, x_16, x_9, x_10, x_11, x_12, x_13, x_14); +lean_dec(x_2); +lean_dec(x_1); +return x_17; +} +} +LEAN_EXPORT lean_object* l_Lean_withoutModifyingState___at_Lean_Meta_addPPExplicitToExposeDiff___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; lean_object* x_9; lean_object* x_10; +x_7 = l_Lean_Meta_saveState___rarg(x_3, x_4, x_5, x_6); +x_8 = lean_ctor_get(x_7, 0); +lean_inc(x_8); +x_9 = lean_ctor_get(x_7, 1); +lean_inc(x_9); +lean_dec(x_7); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +x_10 = lean_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 = l_Lean_Meta_SavedState_restore(x_8, x_2, x_3, x_4, x_5, x_12); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_8); +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) +{ +lean_object* x_15; +x_15 = lean_ctor_get(x_13, 0); +lean_dec(x_15); +lean_ctor_set(x_13, 0, x_11); +return x_13; +} +else +{ +lean_object* x_16; lean_object* x_17; +x_16 = lean_ctor_get(x_13, 1); +lean_inc(x_16); +lean_dec(x_13); +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_11); +lean_ctor_set(x_17, 1, x_16); +return x_17; +} +} +else +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_18 = lean_ctor_get(x_10, 0); +lean_inc(x_18); +x_19 = lean_ctor_get(x_10, 1); +lean_inc(x_19); +lean_dec(x_10); +x_20 = l_Lean_Meta_SavedState_restore(x_8, x_2, x_3, x_4, x_5, x_19); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_8); +x_21 = !lean_is_exclusive(x_20); +if (x_21 == 0) +{ +lean_object* x_22; +x_22 = lean_ctor_get(x_20, 0); +lean_dec(x_22); +lean_ctor_set_tag(x_20, 1); +lean_ctor_set(x_20, 0, x_18); +return x_20; +} +else +{ +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(1, 2, 0); +lean_ctor_set(x_24, 0, x_18); +lean_ctor_set(x_24, 1, x_23); +return x_24; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_addPPExplicitToExposeDiff___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_8 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_1, x_3, x_4, x_5, x_6, x_7); +x_9 = lean_ctor_get(x_8, 0); +lean_inc(x_9); +x_10 = lean_ctor_get(x_8, 1); +lean_inc(x_10); +lean_dec(x_8); +x_11 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_2, x_3, x_4, x_5, x_6, x_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_Meta_addPPExplicitToExposeDiff_visit(x_9, x_12, x_3, x_4, x_5, x_6, x_13); +return x_14; +} +} +static lean_object* _init_l_Lean_Meta_addPPExplicitToExposeDiff___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("all", 3, 3); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_addPPExplicitToExposeDiff___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_addPPExplicitToExposeDiff_visit___lambda__3___closed__4; +x_2 = l_Lean_Meta_addPPExplicitToExposeDiff___closed__1; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_addPPExplicitToExposeDiff___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("explicit", 8, 8); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_addPPExplicitToExposeDiff___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_addPPExplicitToExposeDiff_visit___lambda__3___closed__4; +x_2 = l_Lean_Meta_addPPExplicitToExposeDiff___closed__3; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_addPPExplicitToExposeDiff(lean_object* x_1, 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; uint8_t x_11; +x_8 = lean_ctor_get(x_5, 2); +lean_inc(x_8); +x_9 = l_Lean_Meta_addPPExplicitToExposeDiff___closed__2; +x_10 = 0; +x_11 = l_Lean_KVMap_getBool(x_8, x_9, x_10); +if (x_11 == 0) +{ +lean_object* x_12; uint8_t x_13; +x_12 = l_Lean_Meta_addPPExplicitToExposeDiff___closed__4; +x_13 = l_Lean_KVMap_getBool(x_8, x_12, x_10); +lean_dec(x_8); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; +x_14 = lean_alloc_closure((void*)(l_Lean_Meta_addPPExplicitToExposeDiff___lambda__1), 7, 2); +lean_closure_set(x_14, 0, x_1); +lean_closure_set(x_14, 1, x_2); +x_15 = l_Lean_withoutModifyingState___at_Lean_Meta_addPPExplicitToExposeDiff___spec__1(x_14, x_3, x_4, x_5, x_6, x_7); +return x_15; +} +else +{ +lean_object* x_16; lean_object* x_17; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_1); +lean_ctor_set(x_16, 1, x_2); +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_16); +lean_ctor_set(x_17, 1, x_7); +return x_17; } } else { -lean_object* x_630; lean_object* x_631; lean_object* x_632; +lean_object* x_18; lean_object* x_19; +lean_dec(x_8); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_630 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_630, 0, x_1); -lean_ctor_set(x_630, 1, x_2); -x_631 = lean_box(0); -x_632 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_632, 0, x_630); -lean_ctor_set(x_632, 1, x_631); -x_8 = x_632; -x_9 = x_7; -goto block_18; +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_1); +lean_ctor_set(x_18, 1, x_2); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_7); +return x_19; } } } -block_18: -{ -if (lean_obj_tag(x_8) == 0) -{ -uint8_t x_10; -x_10 = !lean_is_exclusive(x_8); -if (x_10 == 0) +static lean_object* _init_l_panic___at_Lean_Meta_throwLetTypeMismatchMessage___spec__1___rarg___closed__1() { +_start: { -lean_object* x_11; -x_11 = lean_ctor_get(x_8, 1); -lean_dec(x_11); -lean_ctor_set(x_8, 1, x_9); -return x_8; +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_instInhabitedMetaM___boxed), 5, 1); +lean_closure_set(x_1, 0, lean_box(0)); +return x_1; } -else +} +LEAN_EXPORT lean_object* l_panic___at_Lean_Meta_throwLetTypeMismatchMessage___spec__1___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: { -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_alloc_ctor(0, 2, 0); -lean_ctor_set(x_13, 0, x_12); -lean_ctor_set(x_13, 1, x_9); -return x_13; +lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_7 = l_panic___at_Lean_Meta_throwLetTypeMismatchMessage___spec__1___rarg___closed__1; +x_8 = lean_panic_fn(x_7, x_1); +x_9 = lean_apply_5(x_8, x_2, x_3, x_4, x_5, x_6); +return x_9; } } -else -{ -uint8_t x_14; -x_14 = !lean_is_exclusive(x_8); -if (x_14 == 0) +LEAN_EXPORT lean_object* l_panic___at_Lean_Meta_throwLetTypeMismatchMessage___spec__1(lean_object* x_1) { +_start: { -lean_object* x_15; -x_15 = lean_ctor_get(x_8, 1); -lean_dec(x_15); -lean_ctor_set_tag(x_8, 0); -lean_ctor_set(x_8, 1, x_9); -return x_8; +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_panic___at_Lean_Meta_throwLetTypeMismatchMessage___spec__1___rarg), 6, 0); +return x_2; } -else -{ -lean_object* x_16; lean_object* x_17; -x_16 = lean_ctor_get(x_8, 0); -lean_inc(x_16); -lean_dec(x_8); -x_17 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_17, 0, x_16); -lean_ctor_set(x_17, 1, x_9); -return x_17; } +LEAN_EXPORT lean_object* l_panic___at_Lean_Meta_throwLetTypeMismatchMessage___spec__2___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_7 = l_panic___at_Lean_Meta_throwLetTypeMismatchMessage___spec__1___rarg___closed__1; +x_8 = lean_panic_fn(x_7, x_1); +x_9 = lean_apply_5(x_8, x_2, x_3, x_4, x_5, x_6); +return x_9; } } -block_28: +LEAN_EXPORT lean_object* l_panic___at_Lean_Meta_throwLetTypeMismatchMessage___spec__2(lean_object* x_1) { +_start: { -uint8_t x_21; -x_21 = l_Lean_Exception_isInterrupt(x_19); -if (x_21 == 0) +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_panic___at_Lean_Meta_throwLetTypeMismatchMessage___spec__2___rarg), 6, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_throwLetTypeMismatchMessage___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) { +_start: { -uint8_t x_22; -x_22 = l_Lean_Exception_isRuntime(x_19); -if (x_22 == 0) +lean_object* x_7; lean_object* x_8; uint8_t x_9; +x_7 = lean_ctor_get(x_4, 5); +x_8 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_1, x_2, x_3, x_4, x_5, x_6); +x_9 = !lean_is_exclusive(x_8); +if (x_9 == 0) { -lean_object* x_23; lean_object* x_24; lean_object* x_25; -lean_dec(x_19); -x_23 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_23, 0, x_1); -lean_ctor_set(x_23, 1, x_2); -x_24 = lean_box(0); -x_25 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_25, 0, x_23); -lean_ctor_set(x_25, 1, x_24); -x_8 = x_25; -x_9 = x_20; -goto block_18; +lean_object* x_10; lean_object* x_11; +x_10 = lean_ctor_get(x_8, 0); +lean_inc(x_7); +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_7); +lean_ctor_set(x_11, 1, x_10); +lean_ctor_set_tag(x_8, 1); +lean_ctor_set(x_8, 0, x_11); +return x_8; } else { -lean_object* x_26; -lean_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; -} +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_12 = lean_ctor_get(x_8, 0); +x_13 = lean_ctor_get(x_8, 1); +lean_inc(x_13); +lean_inc(x_12); +lean_dec(x_8); +lean_inc(x_7); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_7); +lean_ctor_set(x_14, 1, x_12); +x_15 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_15, 0, x_14); +lean_ctor_set(x_15, 1, x_13); +return x_15; } -else -{ -lean_object* x_27; -lean_dec(x_2); -lean_dec(x_1); -x_27 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_27, 0, x_19); -lean_ctor_set(x_27, 1, x_20); -return x_27; } } +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_throwLetTypeMismatchMessage___spec__3(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_throwError___at_Lean_Meta_throwLetTypeMismatchMessage___spec__3___rarg___boxed), 6, 0); +return x_2; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_addPPExplicitToExposeDiff_visit___spec__1___lambda__1___boxed(lean_object** _args) { -lean_object* x_1 = _args[0]; -lean_object* x_2 = _args[1]; -lean_object* x_3 = _args[2]; -lean_object* x_4 = _args[3]; -lean_object* x_5 = _args[4]; -lean_object* x_6 = _args[5]; -lean_object* x_7 = _args[6]; -lean_object* x_8 = _args[7]; -lean_object* x_9 = _args[8]; -lean_object* x_10 = _args[9]; -lean_object* x_11 = _args[10]; -lean_object* x_12 = _args[11]; -lean_object* x_13 = _args[12]; -lean_object* x_14 = _args[13]; -lean_object* x_15 = _args[14]; -lean_object* x_16 = _args[15]; -lean_object* x_17 = _args[16]; +static lean_object* _init_l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__1() { _start: { -lean_object* x_18; -x_18 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_addPPExplicitToExposeDiff_visit___spec__1___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_5); -lean_dec(x_4); -return x_18; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.Meta.Check", 15, 15); +return x_1; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_addPPExplicitToExposeDiff_visit___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]; -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]; +static lean_object* _init_l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__2() { _start: { -lean_object* x_18; -x_18 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_addPPExplicitToExposeDiff_visit___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, x_13, x_14, x_15, x_16, x_17); -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_5); -lean_dec(x_4); -return x_18; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.Meta.throwLetTypeMismatchMessage", 37, 37); +return x_1; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_addPPExplicitToExposeDiff_visit___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]; +static lean_object* _init_l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__3() { _start: { -lean_object* x_18; -x_18 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_addPPExplicitToExposeDiff_visit___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); -lean_dec(x_8); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -return x_18; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("unreachable code has been reached", 33, 33); +return x_1; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at_Lean_Meta_addPPExplicitToExposeDiff_visit___spec__2___boxed(lean_object* x_1, lean_object* x_2) { +static lean_object* _init_l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__4() { _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_159____at_Lean_Meta_addPPExplicitToExposeDiff_visit___spec__2(x_1, x_2); -lean_dec(x_2); -lean_dec(x_1); -x_4 = lean_box(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; lean_object* x_6; +x_1 = l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__1; +x_2 = l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__2; +x_3 = lean_unsigned_to_nat(179u); +x_4 = lean_unsigned_to_nat(9u); +x_5 = l_Lean_Meta_throwLetTypeMismatchMessage___rarg___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___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at_Lean_Meta_addPPExplicitToExposeDiff_visit___spec__3___boxed(lean_object* x_1, lean_object* x_2) { +static lean_object* _init_l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__5() { _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_159____at_Lean_Meta_addPPExplicitToExposeDiff_visit___spec__3(x_1, x_2); -lean_dec(x_2); -lean_dec(x_1); -x_4 = lean_box(x_3); -return x_4; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("invalid let declaration, term", 29, 29); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Meta_addPPExplicitToExposeDiff_visit___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) { +static lean_object* _init_l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__6() { _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_5); -lean_dec(x_5); -x_13 = l_Lean_Meta_withLocalDecl___at_Lean_Meta_addPPExplicitToExposeDiff_visit___spec__4___rarg(x_1, x_11, x_3, x_4, x_12, x_6, x_7, x_8, x_9, x_10); -return x_13; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__5; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Meta_addPPExplicitToExposeDiff_visit___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +static lean_object* _init_l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__7() { _start: { -lean_object* x_12; -x_12 = l_Lean_Meta_addPPExplicitToExposeDiff_visit___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -return x_12; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("\nhas type", 9, 9); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Meta_addPPExplicitToExposeDiff_visit___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +static lean_object* _init_l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__8() { _start: { -lean_object* x_13; -x_13 = l_Lean_Meta_addPPExplicitToExposeDiff_visit___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_7); -lean_dec(x_4); -lean_dec(x_3); -return x_13; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__7; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Meta_addPPExplicitToExposeDiff_visit___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +static lean_object* _init_l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__9() { _start: { -lean_object* x_16; -x_16 = l_Lean_Meta_addPPExplicitToExposeDiff_visit___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); -lean_dec(x_10); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -return x_16; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("\nbut is expected to have type", 29, 29); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Meta_addPPExplicitToExposeDiff_visit___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +static lean_object* _init_l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__10() { _start: { -lean_object* x_16; -x_16 = l_Lean_Meta_addPPExplicitToExposeDiff_visit___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); -lean_dec(x_10); -lean_dec(x_4); -lean_dec(x_3); -return x_16; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__9; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Meta_addPPExplicitToExposeDiff_visit___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) { +static lean_object* _init_l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__11() { _start: { -uint8_t x_15; uint8_t x_16; lean_object* x_17; -x_15 = lean_unbox(x_5); -lean_dec(x_5); -x_16 = lean_unbox(x_8); -lean_dec(x_8); -x_17 = l_Lean_Meta_addPPExplicitToExposeDiff_visit___lambda__5(x_1, x_2, x_3, x_4, x_15, x_6, x_7, x_16, x_9, x_10, x_11, x_12, x_13, x_14); -lean_dec(x_2); -lean_dec(x_1); -return x_17; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("", 0, 0); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Meta_addPPExplicitToExposeDiff_visit___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) { +static lean_object* _init_l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__12() { _start: { -uint8_t x_15; uint8_t x_16; lean_object* x_17; -x_15 = lean_unbox(x_5); -lean_dec(x_5); -x_16 = lean_unbox(x_8); -lean_dec(x_8); -x_17 = l_Lean_Meta_addPPExplicitToExposeDiff_visit___lambda__6(x_1, x_2, x_3, x_4, x_15, x_6, x_7, x_16, x_9, x_10, x_11, x_12, x_13, x_14); -lean_dec(x_2); -lean_dec(x_1); -return x_17; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__11; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } } -LEAN_EXPORT lean_object* l_Lean_withoutModifyingState___at_Lean_Meta_addPPExplicitToExposeDiff___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_Meta_throwLetTypeMismatchMessage___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; -x_7 = l_Lean_Meta_saveState___rarg(x_3, x_4, x_5, x_6); -x_8 = lean_ctor_get(x_7, 0); -lean_inc(x_8); -x_9 = lean_ctor_get(x_7, 1); -lean_inc(x_9); -lean_dec(x_7); +lean_object* x_7; lean_object* x_8; +x_7 = lean_ctor_get(x_2, 2); +lean_inc(x_7); +x_8 = lean_local_ctx_find(x_7, x_1); +if (lean_obj_tag(x_8) == 0) +{ +lean_object* x_9; lean_object* x_10; +x_9 = l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__4; +x_10 = l_panic___at_Lean_Meta_throwLetTypeMismatchMessage___spec__1___rarg(x_9, x_2, x_3, x_4, x_5, x_6); +return x_10; +} +else +{ +lean_object* x_11; +x_11 = lean_ctor_get(x_8, 0); +lean_inc(x_11); +lean_dec(x_8); +if (lean_obj_tag(x_11) == 0) +{ +lean_object* x_12; lean_object* x_13; +lean_dec(x_11); +x_12 = l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__4; +x_13 = l_panic___at_Lean_Meta_throwLetTypeMismatchMessage___spec__2___rarg(x_12, x_2, x_3, x_4, x_5, x_6); +return x_13; +} +else +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_14 = lean_ctor_get(x_11, 3); +lean_inc(x_14); +x_15 = lean_ctor_get(x_11, 4); +lean_inc(x_15); +lean_dec(x_11); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); -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_inc(x_15); +x_16 = lean_infer_type(x_15, x_2, x_3, x_4, x_5, x_6); +if (lean_obj_tag(x_16) == 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 = l_Lean_Meta_SavedState_restore(x_8, x_2, x_3, x_4, x_5, x_12); +lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +x_18 = lean_ctor_get(x_16, 1); +lean_inc(x_18); +lean_dec(x_16); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +x_19 = l_Lean_Meta_addPPExplicitToExposeDiff(x_17, x_14, x_2, x_3, x_4, x_5, x_18); +if (lean_obj_tag(x_19) == 0) +{ +lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_20 = lean_ctor_get(x_19, 0); +lean_inc(x_20); +x_21 = lean_ctor_get(x_19, 1); +lean_inc(x_21); +lean_dec(x_19); +x_22 = !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; lean_object* x_31; 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_23 = lean_ctor_get(x_20, 0); +x_24 = lean_ctor_get(x_20, 1); +x_25 = l_Lean_indentExpr(x_15); +x_26 = l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__6; +lean_ctor_set_tag(x_20, 7); +lean_ctor_set(x_20, 1, x_25); +lean_ctor_set(x_20, 0, x_26); +x_27 = l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__8; +x_28 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_28, 0, x_20); +lean_ctor_set(x_28, 1, x_27); +x_29 = l_Lean_indentExpr(x_23); +x_30 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_30, 0, x_28); +lean_ctor_set(x_30, 1, x_29); +x_31 = l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__10; +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_indentExpr(x_24); +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_Meta_throwLetTypeMismatchMessage___rarg___closed__12; +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_throwError___at_Lean_Meta_throwLetTypeMismatchMessage___spec__3___rarg(x_36, x_2, x_3, x_4, x_5, x_21); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -lean_dec(x_8); -x_14 = !lean_is_exclusive(x_13); -if (x_14 == 0) -{ -lean_object* x_15; -x_15 = lean_ctor_get(x_13, 0); -lean_dec(x_15); -lean_ctor_set(x_13, 0, x_11); -return x_13; +return x_37; } else { -lean_object* x_16; lean_object* x_17; -x_16 = lean_ctor_get(x_13, 1); -lean_inc(x_16); -lean_dec(x_13); -x_17 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_17, 0, x_11); -lean_ctor_set(x_17, 1, x_16); -return x_17; +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_38 = lean_ctor_get(x_20, 0); +x_39 = lean_ctor_get(x_20, 1); +lean_inc(x_39); +lean_inc(x_38); +lean_dec(x_20); +x_40 = l_Lean_indentExpr(x_15); +x_41 = l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__6; +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_Meta_throwLetTypeMismatchMessage___rarg___closed__8; +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_indentExpr(x_38); +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_Meta_throwLetTypeMismatchMessage___rarg___closed__10; +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); +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_Meta_throwLetTypeMismatchMessage___rarg___closed__12; +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_Meta_throwLetTypeMismatchMessage___spec__3___rarg(x_52, x_2, x_3, x_4, x_5, x_21); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_53; } } else { -lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; -x_18 = lean_ctor_get(x_10, 0); -lean_inc(x_18); -x_19 = lean_ctor_get(x_10, 1); -lean_inc(x_19); -lean_dec(x_10); -x_20 = l_Lean_Meta_SavedState_restore(x_8, x_2, x_3, x_4, x_5, x_19); +uint8_t x_54; +lean_dec(x_15); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -lean_dec(x_8); -x_21 = !lean_is_exclusive(x_20); -if (x_21 == 0) +x_54 = !lean_is_exclusive(x_19); +if (x_54 == 0) { -lean_object* x_22; -x_22 = lean_ctor_get(x_20, 0); -lean_dec(x_22); -lean_ctor_set_tag(x_20, 1); -lean_ctor_set(x_20, 0, x_18); -return x_20; +return x_19; } else { -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(1, 2, 0); -lean_ctor_set(x_24, 0, x_18); -lean_ctor_set(x_24, 1, x_23); -return x_24; -} +lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_55 = lean_ctor_get(x_19, 0); +x_56 = lean_ctor_get(x_19, 1); +lean_inc(x_56); +lean_inc(x_55); +lean_dec(x_19); +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_EXPORT lean_object* l_Lean_Meta_addPPExplicitToExposeDiff___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: +else { -lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_8 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_1, x_3, x_4, x_5, x_6, x_7); -x_9 = lean_ctor_get(x_8, 0); -lean_inc(x_9); -x_10 = lean_ctor_get(x_8, 1); -lean_inc(x_10); -lean_dec(x_8); -x_11 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_2, x_3, x_4, x_5, x_6, x_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_Meta_addPPExplicitToExposeDiff_visit(x_9, x_12, x_3, x_4, x_5, x_6, x_13); -return x_14; -} -} -static lean_object* _init_l_Lean_Meta_addPPExplicitToExposeDiff___closed__1() { -_start: +uint8_t x_58; +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_58 = !lean_is_exclusive(x_16); +if (x_58 == 0) { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("all", 3, 3); -return x_1; -} +return x_16; } -static lean_object* _init_l_Lean_Meta_addPPExplicitToExposeDiff___closed__2() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_addPPExplicitToExposeDiff_visit___lambda__3___closed__4; -x_2 = l_Lean_Meta_addPPExplicitToExposeDiff___closed__1; -x_3 = l_Lean_Name_mkStr2(x_1, x_2); -return x_3; +lean_object* x_59; lean_object* x_60; lean_object* x_61; +x_59 = lean_ctor_get(x_16, 0); +x_60 = lean_ctor_get(x_16, 1); +lean_inc(x_60); +lean_inc(x_59); +lean_dec(x_16); +x_61 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_61, 0, x_59); +lean_ctor_set(x_61, 1, x_60); +return x_61; } } -static lean_object* _init_l_Lean_Meta_addPPExplicitToExposeDiff___closed__3() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("explicit", 8, 8); -return x_1; } } -static lean_object* _init_l_Lean_Meta_addPPExplicitToExposeDiff___closed__4() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_addPPExplicitToExposeDiff_visit___lambda__3___closed__4; -x_2 = l_Lean_Meta_addPPExplicitToExposeDiff___closed__3; -x_3 = l_Lean_Name_mkStr2(x_1, x_2); -return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Meta_addPPExplicitToExposeDiff(lean_object* x_1, lean_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_throwLetTypeMismatchMessage(lean_object* x_1) { _start: { -lean_object* x_8; lean_object* x_9; uint8_t x_10; uint8_t x_11; -x_8 = lean_ctor_get(x_5, 2); -lean_inc(x_8); -x_9 = l_Lean_Meta_addPPExplicitToExposeDiff___closed__2; -x_10 = 0; -x_11 = l_Lean_KVMap_getBool(x_8, x_9, x_10); -if (x_11 == 0) -{ -lean_object* x_12; uint8_t x_13; -x_12 = l_Lean_Meta_addPPExplicitToExposeDiff___closed__4; -x_13 = l_Lean_KVMap_getBool(x_8, x_12, x_10); -lean_dec(x_8); -if (x_13 == 0) -{ -lean_object* x_14; lean_object* x_15; -x_14 = lean_alloc_closure((void*)(l_Lean_Meta_addPPExplicitToExposeDiff___lambda__1), 7, 2); -lean_closure_set(x_14, 0, x_1); -lean_closure_set(x_14, 1, x_2); -x_15 = l_Lean_withoutModifyingState___at_Lean_Meta_addPPExplicitToExposeDiff___spec__1(x_14, x_3, x_4, x_5, x_6, x_7); -return x_15; -} -else -{ -lean_object* x_16; lean_object* x_17; -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -x_16 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_16, 0, x_1); -lean_ctor_set(x_16, 1, x_2); -x_17 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_17, 0, x_16); -lean_ctor_set(x_17, 1, x_7); -return x_17; +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_Meta_throwLetTypeMismatchMessage___rarg), 6, 0); +return x_2; } } -else +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_throwLetTypeMismatchMessage___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: { -lean_object* x_18; lean_object* x_19; -lean_dec(x_8); -lean_dec(x_6); +lean_object* x_7; +x_7 = l_Lean_throwError___at_Lean_Meta_throwLetTypeMismatchMessage___spec__3___rarg(x_1, x_2, x_3, x_4, x_5, x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_18 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_18, 0, x_1); -lean_ctor_set(x_18, 1, x_2); -x_19 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_19, 0, x_18); -lean_ctor_set(x_19, 1, x_7); -return x_19; -} +lean_dec(x_2); +return x_7; } } static lean_object* _init_l_Lean_Meta_mkHasTypeButIsExpectedMsg___closed__1() { @@ -16993,7 +17081,7 @@ return x_30; } } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692____closed__1() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____closed__1() { _start: { lean_object* x_1; @@ -17001,27 +17089,27 @@ x_1 = lean_mk_string_unchecked("Lean", 4, 4); return x_1; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692____closed__2() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____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_Check___hyg_4692____closed__1; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____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_Check___hyg_4692____closed__3() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692____closed__2; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____closed__2; x_2 = l_Lean_Meta_check___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_Check___hyg_4692____closed__4() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____closed__4() { _start: { lean_object* x_1; @@ -17029,17 +17117,17 @@ x_1 = lean_mk_string_unchecked("initFn", 6, 6); return x_1; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692____closed__5() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692____closed__3; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692____closed__4; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____closed__3; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____closed__4; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692____closed__6() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____closed__6() { _start: { lean_object* x_1; @@ -17047,37 +17135,37 @@ x_1 = lean_mk_string_unchecked("_@", 2, 2); return x_1; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692____closed__7() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692____closed__5; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692____closed__6; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____closed__5; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____closed__6; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692____closed__8() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692____closed__7; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692____closed__1; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____closed__7; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____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_Check___hyg_4692____closed__9() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692____closed__8; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____closed__8; x_2 = l_Lean_Meta_check___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_Check___hyg_4692____closed__10() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____closed__10() { _start: { lean_object* x_1; @@ -17085,17 +17173,17 @@ x_1 = lean_mk_string_unchecked("Check", 5, 5); return x_1; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692____closed__11() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692____closed__9; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692____closed__10; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____closed__9; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____closed__10; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692____closed__12() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____closed__12() { _start: { lean_object* x_1; @@ -17103,33 +17191,33 @@ x_1 = lean_mk_string_unchecked("_hyg", 4, 4); return x_1; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692____closed__13() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692____closed__11; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692____closed__12; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____closed__11; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____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_Check___hyg_4692____closed__14() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692____closed__13; -x_2 = lean_unsigned_to_nat(4692u); +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____closed__13; +x_2 = lean_unsigned_to_nat(4731u); 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_Check___hyg_4692_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731_(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_check___closed__3; x_3 = 0; -x_4 = l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692____closed__14; +x_4 = l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____closed__14; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); return x_5; } @@ -17147,32 +17235,6 @@ lean_dec_ref(res); res = initialize_Lean_Meta_Sorry(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l_panic___at_Lean_Meta_throwLetTypeMismatchMessage___spec__1___rarg___closed__1 = _init_l_panic___at_Lean_Meta_throwLetTypeMismatchMessage___spec__1___rarg___closed__1(); -lean_mark_persistent(l_panic___at_Lean_Meta_throwLetTypeMismatchMessage___spec__1___rarg___closed__1); -l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__1 = _init_l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__1(); -lean_mark_persistent(l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__1); -l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__2 = _init_l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__2(); -lean_mark_persistent(l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__2); -l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__3 = _init_l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__3(); -lean_mark_persistent(l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__3); -l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__4 = _init_l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__4(); -lean_mark_persistent(l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__4); -l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__5 = _init_l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__5(); -lean_mark_persistent(l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__5); -l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__6 = _init_l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__6(); -lean_mark_persistent(l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__6); -l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__7 = _init_l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__7(); -lean_mark_persistent(l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__7); -l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__8 = _init_l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__8(); -lean_mark_persistent(l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__8); -l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__9 = _init_l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__9(); -lean_mark_persistent(l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__9); -l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__10 = _init_l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__10(); -lean_mark_persistent(l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__10); -l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__11 = _init_l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__11(); -lean_mark_persistent(l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__11); -l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__12 = _init_l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__12(); -lean_mark_persistent(l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__12); l_Std_Range_forIn_x27_loop___at_Lean_Meta_addPPExplicitToExposeDiff_visit___spec__1___lambda__2___closed__1 = _init_l_Std_Range_forIn_x27_loop___at_Lean_Meta_addPPExplicitToExposeDiff_visit___spec__1___lambda__2___closed__1(); l_Lean_Meta_addPPExplicitToExposeDiff_visit___lambda__3___closed__1 = _init_l_Lean_Meta_addPPExplicitToExposeDiff_visit___lambda__3___closed__1(); lean_mark_persistent(l_Lean_Meta_addPPExplicitToExposeDiff_visit___lambda__3___closed__1); @@ -17208,6 +17270,32 @@ l_Lean_Meta_addPPExplicitToExposeDiff___closed__3 = _init_l_Lean_Meta_addPPExpli lean_mark_persistent(l_Lean_Meta_addPPExplicitToExposeDiff___closed__3); l_Lean_Meta_addPPExplicitToExposeDiff___closed__4 = _init_l_Lean_Meta_addPPExplicitToExposeDiff___closed__4(); lean_mark_persistent(l_Lean_Meta_addPPExplicitToExposeDiff___closed__4); +l_panic___at_Lean_Meta_throwLetTypeMismatchMessage___spec__1___rarg___closed__1 = _init_l_panic___at_Lean_Meta_throwLetTypeMismatchMessage___spec__1___rarg___closed__1(); +lean_mark_persistent(l_panic___at_Lean_Meta_throwLetTypeMismatchMessage___spec__1___rarg___closed__1); +l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__1 = _init_l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__1(); +lean_mark_persistent(l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__1); +l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__2 = _init_l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__2(); +lean_mark_persistent(l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__2); +l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__3 = _init_l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__3(); +lean_mark_persistent(l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__3); +l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__4 = _init_l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__4(); +lean_mark_persistent(l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__4); +l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__5 = _init_l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__5(); +lean_mark_persistent(l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__5); +l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__6 = _init_l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__6(); +lean_mark_persistent(l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__6); +l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__7 = _init_l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__7(); +lean_mark_persistent(l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__7); +l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__8 = _init_l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__8(); +lean_mark_persistent(l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__8); +l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__9 = _init_l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__9(); +lean_mark_persistent(l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__9); +l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__10 = _init_l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__10(); +lean_mark_persistent(l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__10); +l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__11 = _init_l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__11(); +lean_mark_persistent(l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__11); +l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__12 = _init_l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__12(); +lean_mark_persistent(l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__12); l_Lean_Meta_mkHasTypeButIsExpectedMsg___closed__1 = _init_l_Lean_Meta_mkHasTypeButIsExpectedMsg___closed__1(); lean_mark_persistent(l_Lean_Meta_mkHasTypeButIsExpectedMsg___closed__1); l_Lean_Meta_mkHasTypeButIsExpectedMsg___closed__2 = _init_l_Lean_Meta_mkHasTypeButIsExpectedMsg___closed__2(); @@ -17281,35 +17369,35 @@ l_Lean_Meta_check___closed__2 = _init_l_Lean_Meta_check___closed__2(); lean_mark_persistent(l_Lean_Meta_check___closed__2); l_Lean_Meta_check___closed__3 = _init_l_Lean_Meta_check___closed__3(); lean_mark_persistent(l_Lean_Meta_check___closed__3); -l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692____closed__1 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692____closed__1(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692____closed__1); -l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692____closed__2 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692____closed__2(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692____closed__2); -l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692____closed__3 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692____closed__3(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692____closed__3); -l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692____closed__4 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692____closed__4(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692____closed__4); -l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692____closed__5 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692____closed__5(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692____closed__5); -l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692____closed__6 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692____closed__6(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692____closed__6); -l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692____closed__7 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692____closed__7(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692____closed__7); -l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692____closed__8 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692____closed__8(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692____closed__8); -l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692____closed__9 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692____closed__9(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692____closed__9); -l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692____closed__10 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692____closed__10(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692____closed__10); -l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692____closed__11 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692____closed__11(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692____closed__11); -l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692____closed__12 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692____closed__12(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692____closed__12); -l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692____closed__13 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692____closed__13(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692____closed__13); -l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692____closed__14 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692____closed__14(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692____closed__14); -if (builtin) {res = l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4692_(lean_io_mk_world()); +l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____closed__1 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____closed__1(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____closed__1); +l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____closed__2 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____closed__2(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____closed__2); +l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____closed__3 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____closed__3(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____closed__3); +l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____closed__4 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____closed__4(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____closed__4); +l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____closed__5 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____closed__5(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____closed__5); +l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____closed__6 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____closed__6(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____closed__6); +l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____closed__7 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____closed__7(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____closed__7); +l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____closed__8 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____closed__8(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____closed__8); +l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____closed__9 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____closed__9(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____closed__9); +l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____closed__10 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____closed__10(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____closed__10); +l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____closed__11 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____closed__11(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____closed__11); +l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____closed__12 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____closed__12(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____closed__12); +l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____closed__13 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____closed__13(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____closed__13); +l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____closed__14 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____closed__14(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731____closed__14); +if (builtin) {res = l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4731_(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/Grind.c b/stage0/stdlib/Lean/Meta/Tactic/Grind.c index f6c2115f76a0..08ec47f8a228 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Grind.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Grind.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Lean.Meta.Tactic.Grind -// Imports: Lean.Meta.Tactic.Grind.Attr Lean.Meta.Tactic.Grind.RevertAll Lean.Meta.Tactic.Grind.Types Lean.Meta.Tactic.Grind.Preprocessor Lean.Meta.Tactic.Grind.Util Lean.Meta.Tactic.Grind.Cases Lean.Meta.Tactic.Grind.Injection Lean.Meta.Tactic.Grind.Core Lean.Meta.Tactic.Grind.Canon Lean.Meta.Tactic.Grind.MarkNestedProofs Lean.Meta.Tactic.Grind.Inv Lean.Meta.Tactic.Grind.Proof Lean.Meta.Tactic.Grind.Propagate Lean.Meta.Tactic.Grind.PP Lean.Meta.Tactic.Grind.Simp Lean.Meta.Tactic.Grind.Ctor Lean.Meta.Tactic.Grind.Parser Lean.Meta.Tactic.Grind.TheoremPatterns +// Imports: Lean.Meta.Tactic.Grind.Attr Lean.Meta.Tactic.Grind.RevertAll Lean.Meta.Tactic.Grind.Types Lean.Meta.Tactic.Grind.Util Lean.Meta.Tactic.Grind.Cases Lean.Meta.Tactic.Grind.Injection Lean.Meta.Tactic.Grind.Core Lean.Meta.Tactic.Grind.Canon Lean.Meta.Tactic.Grind.MarkNestedProofs Lean.Meta.Tactic.Grind.Inv Lean.Meta.Tactic.Grind.Proof Lean.Meta.Tactic.Grind.Propagate Lean.Meta.Tactic.Grind.PP Lean.Meta.Tactic.Grind.Simp Lean.Meta.Tactic.Grind.Ctor Lean.Meta.Tactic.Grind.Parser Lean.Meta.Tactic.Grind.EMatchTheorem Lean.Meta.Tactic.Grind.EMatch Lean.Meta.Tactic.Grind.Main #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -13,75 +13,92 @@ #ifdef __cplusplus extern "C" { #endif -static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_299____closed__1; -static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_188____closed__1; -static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_336____closed__2; -static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_40____closed__2; -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_151_(lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_40____closed__3; -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_262_(lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_225____closed__3; -static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__2; -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_336_(lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__14; -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_188_(lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__15; -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_114_(lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__10; -static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_188____closed__2; -static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_299____closed__3; -static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_151____closed__2; -static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_373____closed__3; -static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_336____closed__3; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_78____closed__2; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__8; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__9; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_300____closed__3; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__1; +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_226_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_412_(lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__14; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__10; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_226____closed__3; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__2; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_337____closed__3; +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_375_(lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__13; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_375____closed__1; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_523____closed__2; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_152____closed__1; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_189____closed__2; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__11; +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_449_(lean_object*); lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_40____closed__1; -static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_225____closed__1; +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_115_(lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_263____closed__3; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_115____closed__3; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_523____closed__1; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__3; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_41____closed__3; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_337____closed__1; lean_object* l_Lean_registerTraceClass(lean_object*, uint8_t, lean_object*, lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__4; -static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_410____closed__3; -static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_336____closed__1; -static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_373____closed__2; -static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__9; -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_373_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3_(lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_262____closed__1; +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_152_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_300_(lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_486____closed__3; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_375____closed__3; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_226____closed__1; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_375____closed__2; lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_262____closed__3; -static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_77____closed__1; -static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__8; -static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_225____closed__2; -static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__5; -static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__17; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_449____closed__3; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_412____closed__3; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__6; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_300____closed__1; +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_523_(lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_412____closed__2; +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_78_(lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_40_(lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_114____closed__2; -static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__13; -static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__1; -static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_188____closed__3; -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_299_(lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_373____closed__1; -static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__11; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__5; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_189____closed__3; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_263____closed__2; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_523____closed__3; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_41____closed__1; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__12; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_560____closed__2; +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_486_(lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_78____closed__3; +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_189_(lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_486____closed__2; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_189____closed__1; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_226____closed__2; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__16; lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_410____closed__1; -static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_114____closed__3; -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_410_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_225_(lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_77____closed__3; -static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__7; -static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_114____closed__1; -static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_151____closed__3; -static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__12; -static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_299____closed__2; -static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_262____closed__2; -static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__6; -static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_410____closed__2; -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_77_(lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_151____closed__1; -static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__18; -static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__3; -static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__16; -static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_77____closed__2; -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__1() { +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_152____closed__3; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_560____closed__1; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_449____closed__2; +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_41_(lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_486____closed__1; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_337____closed__2; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__15; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_263____closed__1; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_300____closed__2; +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_560_(lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_449____closed__1; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__17; +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_337_(lean_object*); +lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_115____closed__1; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__18; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_412____closed__1; +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4_(lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__7; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__4; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_115____closed__2; +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_263_(lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_152____closed__2; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_78____closed__1; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_560____closed__3; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_41____closed__2; +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__1() { _start: { lean_object* x_1; @@ -89,17 +106,17 @@ x_1 = lean_mk_string_unchecked("grind", 5, 5); return x_1; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__2() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____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_Meta_Tactic_Grind___hyg_3____closed__1; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__3() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__3() { _start: { lean_object* x_1; @@ -107,17 +124,17 @@ x_1 = lean_mk_string_unchecked("Lean", 4, 4); return x_1; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__4() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__4() { _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_Meta_Tactic_Grind___hyg_3____closed__3; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__3; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__5() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__5() { _start: { lean_object* x_1; @@ -125,17 +142,17 @@ x_1 = lean_mk_string_unchecked("initFn", 6, 6); return x_1; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__6() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__4; -x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__5; +x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__4; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__5; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__7() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__7() { _start: { lean_object* x_1; @@ -143,27 +160,27 @@ x_1 = lean_mk_string_unchecked("_@", 2, 2); return x_1; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__8() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__6; -x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__7; +x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__6; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__7; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__9() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__8; -x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__3; +x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__8; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__3; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__10() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__10() { _start: { lean_object* x_1; @@ -171,17 +188,17 @@ x_1 = lean_mk_string_unchecked("Meta", 4, 4); return x_1; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__11() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__9; -x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__10; +x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__9; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__10; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__12() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__12() { _start: { lean_object* x_1; @@ -189,17 +206,17 @@ x_1 = lean_mk_string_unchecked("Tactic", 6, 6); return x_1; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__13() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__11; -x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__12; +x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__11; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__12; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__14() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__14() { _start: { lean_object* x_1; @@ -207,17 +224,17 @@ x_1 = lean_mk_string_unchecked("Grind", 5, 5); return x_1; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__15() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__13; -x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__14; +x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__13; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__14; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__16() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__16() { _start: { lean_object* x_1; @@ -225,464 +242,627 @@ x_1 = lean_mk_string_unchecked("_hyg", 4, 4); return x_1; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__17() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__15; -x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__16; +x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__15; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__16; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__18() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__17; -x_2 = lean_unsigned_to_nat(3u); +x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__17; +x_2 = lean_unsigned_to_nat(4u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4_(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_Meta_Tactic_Grind___hyg_3____closed__2; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__2; x_3 = 0; -x_4 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__18; +x_4 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__18; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); return x_5; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_40____closed__1() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_41____closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("eq", 2, 2); +x_1 = lean_mk_string_unchecked("assert", 6, 6); return x_1; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_40____closed__2() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_41____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__1; -x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_40____closed__1; +x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__1; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_41____closed__1; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_40____closed__3() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_41____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__17; -x_2 = lean_unsigned_to_nat(40u); +x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__17; +x_2 = lean_unsigned_to_nat(41u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_40_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_41_(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_Meta_Tactic_Grind___hyg_40____closed__2; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_41____closed__2; x_3 = 0; -x_4 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_40____closed__3; +x_4 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_41____closed__3; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); return x_5; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_77____closed__1() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_78____closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("issues", 6, 6); +x_1 = lean_mk_string_unchecked("eqc", 3, 3); return x_1; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_77____closed__2() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_78____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__1; -x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_77____closed__1; +x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__1; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_78____closed__1; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_77____closed__3() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_78____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__17; -x_2 = lean_unsigned_to_nat(77u); +x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__17; +x_2 = lean_unsigned_to_nat(78u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_77_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_78_(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_Meta_Tactic_Grind___hyg_77____closed__2; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_78____closed__2; x_3 = 0; -x_4 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_77____closed__3; +x_4 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_78____closed__3; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); return x_5; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_114____closed__1() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_115____closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("add", 3, 3); +x_1 = lean_mk_string_unchecked("internalize", 11, 11); return x_1; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_114____closed__2() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_115____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__1; -x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_114____closed__1; +x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__1; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_115____closed__1; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_114____closed__3() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_115____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__17; -x_2 = lean_unsigned_to_nat(114u); +x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__17; +x_2 = lean_unsigned_to_nat(115u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_114_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_115_(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_Meta_Tactic_Grind___hyg_114____closed__2; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_115____closed__2; x_3 = 0; -x_4 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_114____closed__3; +x_4 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_115____closed__3; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); return x_5; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_151____closed__1() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_152____closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("pre", 3, 3); +x_1 = lean_mk_string_unchecked("ematch", 6, 6); return x_1; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_151____closed__2() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_152____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__1; -x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_151____closed__1; +x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__1; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_152____closed__1; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_151____closed__3() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_152____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__17; -x_2 = lean_unsigned_to_nat(151u); +x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__17; +x_2 = lean_unsigned_to_nat(152u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_151_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_152_(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_Meta_Tactic_Grind___hyg_151____closed__2; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_152____closed__2; x_3 = 0; -x_4 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_151____closed__3; +x_4 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_152____closed__3; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); return x_5; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_188____closed__1() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_189____closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("debug", 5, 5); +x_1 = lean_mk_string_unchecked("pattern", 7, 7); return x_1; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_188____closed__2() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_189____closed__2() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__1; -x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_188____closed__1; -x_3 = l_Lean_Name_mkStr2(x_1, x_2); -return x_3; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__1; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_152____closed__1; +x_3 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_189____closed__1; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_188____closed__3() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_189____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__17; -x_2 = lean_unsigned_to_nat(188u); +x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__17; +x_2 = lean_unsigned_to_nat(189u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_188_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_189_(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_Meta_Tactic_Grind___hyg_188____closed__2; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_189____closed__2; x_3 = 0; -x_4 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_188____closed__3; +x_4 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_189____closed__3; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); return x_5; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_225____closed__1() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_226____closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("proofs", 6, 6); +x_1 = lean_mk_string_unchecked("instance", 8, 8); return x_1; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_225____closed__2() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_226____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__1; -x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_188____closed__1; -x_3 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_225____closed__1; +x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__1; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_152____closed__1; +x_3 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_226____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_225____closed__3() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_226____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__17; -x_2 = lean_unsigned_to_nat(225u); +x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__17; +x_2 = lean_unsigned_to_nat(226u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_225_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_226_(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_Meta_Tactic_Grind___hyg_225____closed__2; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_226____closed__2; x_3 = 0; -x_4 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_225____closed__3; +x_4 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_226____closed__3; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); return x_5; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_262____closed__1() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_263____closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("simp", 4, 4); +x_1 = lean_mk_string_unchecked("assignment", 10, 10); return x_1; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_262____closed__2() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_263____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_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__1; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_152____closed__1; +x_3 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_226____closed__1; +x_4 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_263____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_initFn____x40_Lean_Meta_Tactic_Grind___hyg_263____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__1; -x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_262____closed__1; +x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__17; +x_2 = lean_unsigned_to_nat(263u); +x_3 = l_Lean_Name_num___override(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_263_(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_Meta_Tactic_Grind___hyg_263____closed__2; +x_3 = 0; +x_4 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_263____closed__3; +x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); +return x_5; +} +} +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_300____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("issues", 6, 6); +return x_1; +} +} +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_300____closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__1; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_300____closed__1; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_262____closed__3() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_300____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__17; -x_2 = lean_unsigned_to_nat(262u); +x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__17; +x_2 = lean_unsigned_to_nat(300u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_262_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_300_(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_Meta_Tactic_Grind___hyg_262____closed__2; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_300____closed__2; x_3 = 0; -x_4 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_262____closed__3; +x_4 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_300____closed__3; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); return x_5; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_299____closed__1() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_337____closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("congr", 5, 5); +x_1 = lean_mk_string_unchecked("simp", 4, 4); return x_1; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_299____closed__2() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_337____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__1; -x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_299____closed__1; +x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__1; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_337____closed__1; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_299____closed__3() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_337____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__17; -x_2 = lean_unsigned_to_nat(299u); +x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__17; +x_2 = lean_unsigned_to_nat(337u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_299_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_337_(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_Meta_Tactic_Grind___hyg_299____closed__2; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_337____closed__2; x_3 = 0; -x_4 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_299____closed__3; +x_4 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_337____closed__3; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); return x_5; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_336____closed__1() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_375____closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("proof", 5, 5); +x_1 = lean_mk_string_unchecked("debug", 5, 5); return x_1; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_336____closed__2() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_375____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__1; -x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_336____closed__1; +x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__1; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_375____closed__1; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_336____closed__3() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_375____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__17; -x_2 = lean_unsigned_to_nat(336u); +x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__17; +x_2 = lean_unsigned_to_nat(375u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_336_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_375_(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_Meta_Tactic_Grind___hyg_336____closed__2; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_375____closed__2; x_3 = 0; -x_4 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_336____closed__3; +x_4 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_375____closed__3; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); return x_5; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_373____closed__1() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_412____closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("detail", 6, 6); +x_1 = lean_mk_string_unchecked("proofs", 6, 6); return x_1; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_373____closed__2() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_412____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__1; -x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_336____closed__1; -x_3 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_373____closed__1; +x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__1; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_375____closed__1; +x_3 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_412____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_373____closed__3() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_412____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__17; -x_2 = lean_unsigned_to_nat(373u); +x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__17; +x_2 = lean_unsigned_to_nat(412u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_373_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_412_(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_Meta_Tactic_Grind___hyg_373____closed__2; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_412____closed__2; x_3 = 0; -x_4 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_373____closed__3; +x_4 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_412____closed__3; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); return x_5; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_410____closed__1() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_449____closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("pattern", 7, 7); +x_1 = lean_mk_string_unchecked("congr", 5, 5); return x_1; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_410____closed__2() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_449____closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__1; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_375____closed__1; +x_3 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_449____closed__1; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_449____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__1; -x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_410____closed__1; -x_3 = l_Lean_Name_mkStr2(x_1, x_2); +x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__17; +x_2 = lean_unsigned_to_nat(449u); +x_3 = l_Lean_Name_num___override(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_449_(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_Meta_Tactic_Grind___hyg_449____closed__2; +x_3 = 0; +x_4 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_449____closed__3; +x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); +return x_5; +} +} +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_486____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("proof", 5, 5); +return x_1; +} +} +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_486____closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__1; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_375____closed__1; +x_3 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_486____closed__1; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_486____closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__17; +x_2 = lean_unsigned_to_nat(486u); +x_3 = l_Lean_Name_num___override(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_486_(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_Meta_Tactic_Grind___hyg_486____closed__2; +x_3 = 0; +x_4 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_486____closed__3; +x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); +return x_5; +} +} +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_523____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("proj", 4, 4); +return x_1; +} +} +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_523____closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__1; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_375____closed__1; +x_3 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_523____closed__1; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_523____closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__17; +x_2 = lean_unsigned_to_nat(523u); +x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_410____closed__3() { +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_523_(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_Meta_Tactic_Grind___hyg_523____closed__2; +x_3 = 0; +x_4 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_523____closed__3; +x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); +return x_5; +} +} +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_560____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("parent", 6, 6); +return x_1; +} +} +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_560____closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__1; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_375____closed__1; +x_3 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_560____closed__1; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_560____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__17; -x_2 = lean_unsigned_to_nat(410u); +x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__17; +x_2 = lean_unsigned_to_nat(560u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_410_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_560_(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_Meta_Tactic_Grind___hyg_410____closed__2; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_560____closed__2; x_3 = 0; -x_4 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_410____closed__3; +x_4 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_560____closed__3; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); return x_5; } @@ -690,7 +870,6 @@ return x_5; lean_object* initialize_Lean_Meta_Tactic_Grind_Attr(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_Tactic_Grind_RevertAll(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_Tactic_Grind_Types(uint8_t builtin, lean_object*); -lean_object* initialize_Lean_Meta_Tactic_Grind_Preprocessor(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_Tactic_Grind_Util(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_Tactic_Grind_Cases(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_Tactic_Grind_Injection(uint8_t builtin, lean_object*); @@ -704,7 +883,9 @@ lean_object* initialize_Lean_Meta_Tactic_Grind_PP(uint8_t builtin, lean_object*) lean_object* initialize_Lean_Meta_Tactic_Grind_Simp(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_Tactic_Grind_Ctor(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_Tactic_Grind_Parser(uint8_t builtin, lean_object*); -lean_object* initialize_Lean_Meta_Tactic_Grind_TheoremPatterns(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Meta_Tactic_Grind_EMatchTheorem(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Meta_Tactic_Grind_EMatch(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Meta_Tactic_Grind_Main(uint8_t builtin, lean_object*); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Lean_Meta_Tactic_Grind(uint8_t builtin, lean_object* w) { lean_object * res; @@ -719,9 +900,6 @@ lean_dec_ref(res); res = initialize_Lean_Meta_Tactic_Grind_Types(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -res = initialize_Lean_Meta_Tactic_Grind_Preprocessor(builtin, lean_io_mk_world()); -if (lean_io_result_is_error(res)) return res; -lean_dec_ref(res); res = initialize_Lean_Meta_Tactic_Grind_Util(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); @@ -761,145 +939,187 @@ lean_dec_ref(res); res = initialize_Lean_Meta_Tactic_Grind_Parser(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -res = initialize_Lean_Meta_Tactic_Grind_TheoremPatterns(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_Grind___hyg_3____closed__1 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__1(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__1); -l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__2 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__2(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__2); -l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__3 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__3(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__3); -l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__4 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__4(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__4); -l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__5 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__5(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__5); -l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__6 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__6(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__6); -l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__7 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__7(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__7); -l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__8 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__8(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__8); -l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__9 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__9(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__9); -l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__10 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__10(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__10); -l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__11 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__11(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__11); -l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__12 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__12(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__12); -l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__13 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__13(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__13); -l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__14 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__14(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__14); -l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__15 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__15(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__15); -l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__16 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__16(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__16); -l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__17 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__17(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__17); -l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__18 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__18(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3____closed__18); -if (builtin) {res = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_3_(lean_io_mk_world()); -if (lean_io_result_is_error(res)) return res; -lean_dec_ref(res); -}l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_40____closed__1 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_40____closed__1(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_40____closed__1); -l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_40____closed__2 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_40____closed__2(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_40____closed__2); -l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_40____closed__3 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_40____closed__3(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_40____closed__3); -if (builtin) {res = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_40_(lean_io_mk_world()); -if (lean_io_result_is_error(res)) return res; -lean_dec_ref(res); -}l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_77____closed__1 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_77____closed__1(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_77____closed__1); -l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_77____closed__2 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_77____closed__2(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_77____closed__2); -l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_77____closed__3 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_77____closed__3(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_77____closed__3); -if (builtin) {res = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_77_(lean_io_mk_world()); -if (lean_io_result_is_error(res)) return res; -lean_dec_ref(res); -}l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_114____closed__1 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_114____closed__1(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_114____closed__1); -l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_114____closed__2 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_114____closed__2(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_114____closed__2); -l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_114____closed__3 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_114____closed__3(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_114____closed__3); -if (builtin) {res = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_114_(lean_io_mk_world()); -if (lean_io_result_is_error(res)) return res; -lean_dec_ref(res); -}l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_151____closed__1 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_151____closed__1(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_151____closed__1); -l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_151____closed__2 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_151____closed__2(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_151____closed__2); -l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_151____closed__3 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_151____closed__3(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_151____closed__3); -if (builtin) {res = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_151_(lean_io_mk_world()); -if (lean_io_result_is_error(res)) return res; -lean_dec_ref(res); -}l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_188____closed__1 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_188____closed__1(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_188____closed__1); -l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_188____closed__2 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_188____closed__2(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_188____closed__2); -l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_188____closed__3 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_188____closed__3(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_188____closed__3); -if (builtin) {res = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_188_(lean_io_mk_world()); -if (lean_io_result_is_error(res)) return res; -lean_dec_ref(res); -}l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_225____closed__1 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_225____closed__1(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_225____closed__1); -l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_225____closed__2 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_225____closed__2(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_225____closed__2); -l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_225____closed__3 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_225____closed__3(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_225____closed__3); -if (builtin) {res = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_225_(lean_io_mk_world()); -if (lean_io_result_is_error(res)) return res; -lean_dec_ref(res); -}l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_262____closed__1 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_262____closed__1(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_262____closed__1); -l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_262____closed__2 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_262____closed__2(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_262____closed__2); -l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_262____closed__3 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_262____closed__3(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_262____closed__3); -if (builtin) {res = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_262_(lean_io_mk_world()); -if (lean_io_result_is_error(res)) return res; -lean_dec_ref(res); -}l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_299____closed__1 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_299____closed__1(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_299____closed__1); -l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_299____closed__2 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_299____closed__2(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_299____closed__2); -l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_299____closed__3 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_299____closed__3(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_299____closed__3); -if (builtin) {res = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_299_(lean_io_mk_world()); -if (lean_io_result_is_error(res)) return res; -lean_dec_ref(res); -}l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_336____closed__1 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_336____closed__1(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_336____closed__1); -l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_336____closed__2 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_336____closed__2(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_336____closed__2); -l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_336____closed__3 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_336____closed__3(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_336____closed__3); -if (builtin) {res = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_336_(lean_io_mk_world()); -if (lean_io_result_is_error(res)) return res; -lean_dec_ref(res); -}l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_373____closed__1 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_373____closed__1(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_373____closed__1); -l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_373____closed__2 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_373____closed__2(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_373____closed__2); -l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_373____closed__3 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_373____closed__3(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_373____closed__3); -if (builtin) {res = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_373_(lean_io_mk_world()); -if (lean_io_result_is_error(res)) return res; -lean_dec_ref(res); -}l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_410____closed__1 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_410____closed__1(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_410____closed__1); -l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_410____closed__2 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_410____closed__2(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_410____closed__2); -l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_410____closed__3 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_410____closed__3(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_410____closed__3); -if (builtin) {res = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_410_(lean_io_mk_world()); +res = initialize_Lean_Meta_Tactic_Grind_EMatchTheorem(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Lean_Meta_Tactic_Grind_EMatch(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Lean_Meta_Tactic_Grind_Main(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_Grind___hyg_4____closed__1 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__1(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__1); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__2 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__2(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__2); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__3 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__3(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__3); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__4 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__4(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__4); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__5 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__5(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__5); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__6 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__6(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__6); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__7 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__7(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__7); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__8 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__8(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__8); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__9 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__9(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__9); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__10 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__10(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__10); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__11 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__11(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__11); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__12 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__12(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__12); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__13 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__13(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__13); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__14 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__14(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__14); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__15 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__15(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__15); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__16 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__16(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__16); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__17 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__17(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__17); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__18 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__18(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__18); +if (builtin) {res = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4_(lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +}l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_41____closed__1 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_41____closed__1(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_41____closed__1); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_41____closed__2 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_41____closed__2(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_41____closed__2); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_41____closed__3 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_41____closed__3(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_41____closed__3); +if (builtin) {res = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_41_(lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +}l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_78____closed__1 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_78____closed__1(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_78____closed__1); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_78____closed__2 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_78____closed__2(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_78____closed__2); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_78____closed__3 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_78____closed__3(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_78____closed__3); +if (builtin) {res = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_78_(lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +}l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_115____closed__1 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_115____closed__1(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_115____closed__1); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_115____closed__2 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_115____closed__2(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_115____closed__2); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_115____closed__3 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_115____closed__3(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_115____closed__3); +if (builtin) {res = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_115_(lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +}l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_152____closed__1 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_152____closed__1(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_152____closed__1); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_152____closed__2 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_152____closed__2(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_152____closed__2); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_152____closed__3 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_152____closed__3(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_152____closed__3); +if (builtin) {res = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_152_(lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +}l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_189____closed__1 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_189____closed__1(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_189____closed__1); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_189____closed__2 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_189____closed__2(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_189____closed__2); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_189____closed__3 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_189____closed__3(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_189____closed__3); +if (builtin) {res = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_189_(lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +}l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_226____closed__1 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_226____closed__1(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_226____closed__1); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_226____closed__2 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_226____closed__2(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_226____closed__2); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_226____closed__3 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_226____closed__3(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_226____closed__3); +if (builtin) {res = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_226_(lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +}l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_263____closed__1 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_263____closed__1(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_263____closed__1); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_263____closed__2 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_263____closed__2(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_263____closed__2); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_263____closed__3 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_263____closed__3(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_263____closed__3); +if (builtin) {res = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_263_(lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +}l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_300____closed__1 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_300____closed__1(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_300____closed__1); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_300____closed__2 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_300____closed__2(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_300____closed__2); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_300____closed__3 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_300____closed__3(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_300____closed__3); +if (builtin) {res = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_300_(lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +}l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_337____closed__1 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_337____closed__1(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_337____closed__1); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_337____closed__2 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_337____closed__2(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_337____closed__2); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_337____closed__3 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_337____closed__3(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_337____closed__3); +if (builtin) {res = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_337_(lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +}l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_375____closed__1 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_375____closed__1(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_375____closed__1); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_375____closed__2 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_375____closed__2(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_375____closed__2); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_375____closed__3 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_375____closed__3(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_375____closed__3); +if (builtin) {res = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_375_(lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +}l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_412____closed__1 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_412____closed__1(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_412____closed__1); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_412____closed__2 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_412____closed__2(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_412____closed__2); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_412____closed__3 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_412____closed__3(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_412____closed__3); +if (builtin) {res = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_412_(lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +}l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_449____closed__1 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_449____closed__1(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_449____closed__1); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_449____closed__2 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_449____closed__2(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_449____closed__2); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_449____closed__3 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_449____closed__3(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_449____closed__3); +if (builtin) {res = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_449_(lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +}l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_486____closed__1 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_486____closed__1(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_486____closed__1); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_486____closed__2 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_486____closed__2(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_486____closed__2); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_486____closed__3 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_486____closed__3(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_486____closed__3); +if (builtin) {res = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_486_(lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +}l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_523____closed__1 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_523____closed__1(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_523____closed__1); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_523____closed__2 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_523____closed__2(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_523____closed__2); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_523____closed__3 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_523____closed__3(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_523____closed__3); +if (builtin) {res = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_523_(lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +}l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_560____closed__1 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_560____closed__1(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_560____closed__1); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_560____closed__2 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_560____closed__2(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_560____closed__2); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_560____closed__3 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_560____closed__3(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_560____closed__3); +if (builtin) {res = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_560_(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/Grind/Canon.c b/stage0/stdlib/Lean/Meta/Tactic/Grind/Canon.c index fb9018ad0508..29d5704f10e6 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Grind/Canon.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Grind/Canon.c @@ -221,9 +221,10 @@ static lean_object* _init_l_Lean_Meta_Grind_Canon_instInhabitedState___closed__3 { lean_object* x_1; lean_object* x_2; x_1 = l_Lean_Meta_Grind_Canon_instInhabitedState___closed__2; -x_2 = lean_alloc_ctor(0, 2, 0); +x_2 = lean_alloc_ctor(0, 3, 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); return x_2; } } @@ -2505,398 +2506,359 @@ return x_50; else { uint8_t x_51; -x_51 = !lean_is_exclusive(x_21); -if (x_51 == 0) -{ -lean_object* x_52; lean_object* x_53; uint8_t x_54; -x_52 = lean_ctor_get(x_21, 1); -x_53 = lean_ctor_get(x_21, 0); -lean_dec(x_53); -lean_inc(x_1); -x_54 = l_Lean_Expr_fvarsSubset(x_19, x_1); -if (x_54 == 0) -{ -lean_object* x_55; lean_object* x_56; -lean_free_object(x_21); -lean_free_object(x_7); -x_55 = lean_box(0); -lean_inc(x_14); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_1); -lean_inc(x_4); -x_56 = l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1(x_2, x_4, x_1, x_19, x_55, x_10, x_11, x_12, x_13, x_14, x_52); -if (lean_obj_tag(x_56) == 0) -{ -lean_object* x_57; lean_object* x_58; -x_57 = lean_ctor_get(x_56, 0); -lean_inc(x_57); -x_58 = lean_ctor_get(x_57, 0); -lean_inc(x_58); -if (lean_obj_tag(x_58) == 0) -{ -uint8_t x_59; lean_dec(x_20); lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_4); -lean_dec(x_1); -x_59 = !lean_is_exclusive(x_56); -if (x_59 == 0) +x_51 = !lean_is_exclusive(x_21); +if (x_51 == 0) { -lean_object* x_60; uint8_t x_61; -x_60 = lean_ctor_get(x_56, 0); -lean_dec(x_60); -x_61 = !lean_is_exclusive(x_57); -if (x_61 == 0) +lean_object* x_52; uint8_t x_53; +x_52 = lean_ctor_get(x_21, 0); +lean_dec(x_52); +x_53 = !lean_is_exclusive(x_10); +if (x_53 == 0) { -lean_object* x_62; lean_object* x_63; -x_62 = lean_ctor_get(x_57, 0); -lean_dec(x_62); -x_63 = lean_ctor_get(x_58, 0); -lean_inc(x_63); -lean_dec(x_58); -lean_ctor_set(x_57, 0, x_63); -return x_56; +lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; +x_54 = lean_ctor_get(x_10, 1); +lean_inc(x_19); +x_55 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_54, x_1, x_19); +lean_ctor_set(x_10, 1, x_55); +x_56 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_56, 0, x_19); +x_57 = lean_box(0); +lean_ctor_set_tag(x_7, 0); +lean_ctor_set(x_7, 1, x_57); +lean_ctor_set(x_7, 0, 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_10); +lean_ctor_set(x_21, 0, x_58); +return x_21; } else { -lean_object* x_64; lean_object* x_65; lean_object* x_66; -x_64 = lean_ctor_get(x_57, 1); -lean_inc(x_64); -lean_dec(x_57); -x_65 = lean_ctor_get(x_58, 0); -lean_inc(x_65); -lean_dec(x_58); +lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_59 = lean_ctor_get(x_10, 0); +x_60 = lean_ctor_get(x_10, 1); +x_61 = lean_ctor_get(x_10, 2); +lean_inc(x_61); +lean_inc(x_60); +lean_inc(x_59); +lean_dec(x_10); +lean_inc(x_19); +x_62 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_60, x_1, x_19); +x_63 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_63, 0, x_59); +lean_ctor_set(x_63, 1, x_62); +lean_ctor_set(x_63, 2, x_61); +x_64 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_64, 0, x_19); +x_65 = lean_box(0); +lean_ctor_set_tag(x_7, 0); +lean_ctor_set(x_7, 1, x_65); +lean_ctor_set(x_7, 0, 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_64); -lean_ctor_set(x_56, 0, x_66); -return x_56; +lean_ctor_set(x_66, 0, x_7); +lean_ctor_set(x_66, 1, x_63); +lean_ctor_set(x_21, 0, x_66); +return x_21; } } else { -lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; -x_67 = lean_ctor_get(x_56, 1); +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_67 = lean_ctor_get(x_21, 1); lean_inc(x_67); -lean_dec(x_56); -x_68 = lean_ctor_get(x_57, 1); +lean_dec(x_21); +x_68 = lean_ctor_get(x_10, 0); lean_inc(x_68); -if (lean_is_exclusive(x_57)) { - lean_ctor_release(x_57, 0); - lean_ctor_release(x_57, 1); - x_69 = x_57; -} else { - lean_dec_ref(x_57); - x_69 = lean_box(0); -} -x_70 = lean_ctor_get(x_58, 0); +x_69 = lean_ctor_get(x_10, 1); +lean_inc(x_69); +x_70 = lean_ctor_get(x_10, 2); lean_inc(x_70); -lean_dec(x_58); -if (lean_is_scalar(x_69)) { - x_71 = lean_alloc_ctor(0, 2, 0); +if (lean_is_exclusive(x_10)) { + lean_ctor_release(x_10, 0); + lean_ctor_release(x_10, 1); + lean_ctor_release(x_10, 2); + x_71 = x_10; } else { - x_71 = x_69; + lean_dec_ref(x_10); + x_71 = lean_box(0); } -lean_ctor_set(x_71, 0, x_70); -lean_ctor_set(x_71, 1, x_68); -x_72 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_72, 0, x_71); -lean_ctor_set(x_72, 1, x_67); -return x_72; +lean_inc(x_19); +x_72 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_69, x_1, x_19); +if (lean_is_scalar(x_71)) { + x_73 = lean_alloc_ctor(0, 3, 0); +} else { + x_73 = x_71; } +lean_ctor_set(x_73, 0, x_68); +lean_ctor_set(x_73, 1, x_72); +lean_ctor_set(x_73, 2, x_70); +x_74 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_74, 0, x_19); +x_75 = lean_box(0); +lean_ctor_set_tag(x_7, 0); +lean_ctor_set(x_7, 1, x_75); +lean_ctor_set(x_7, 0, x_74); +x_76 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_76, 0, x_7); +lean_ctor_set(x_76, 1, x_73); +x_77 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_77, 0, x_76); +lean_ctor_set(x_77, 1, x_67); +return x_77; } -else -{ -lean_object* x_73; lean_object* x_74; lean_object* x_75; -x_73 = lean_ctor_get(x_56, 1); -lean_inc(x_73); -lean_dec(x_56); -x_74 = lean_ctor_get(x_57, 1); -lean_inc(x_74); -lean_dec(x_57); -x_75 = lean_ctor_get(x_58, 0); -lean_inc(x_75); -lean_dec(x_58); -x_7 = x_20; -x_8 = x_75; -x_9 = lean_box(0); -x_10 = x_74; -x_15 = x_73; -goto _start; } } else { -uint8_t x_77; +uint8_t x_78; +lean_free_object(x_7); lean_dec(x_20); +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_4); lean_dec(x_1); -x_77 = !lean_is_exclusive(x_56); -if (x_77 == 0) +x_78 = !lean_is_exclusive(x_21); +if (x_78 == 0) { -return x_56; +return x_21; } else { -lean_object* x_78; lean_object* x_79; lean_object* x_80; -x_78 = lean_ctor_get(x_56, 0); -x_79 = lean_ctor_get(x_56, 1); +lean_object* x_79; lean_object* x_80; lean_object* x_81; +x_79 = lean_ctor_get(x_21, 0); +x_80 = lean_ctor_get(x_21, 1); +lean_inc(x_80); lean_inc(x_79); -lean_inc(x_78); -lean_dec(x_56); -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_dec(x_21); +x_81 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_81, 0, x_79); +lean_ctor_set(x_81, 1, x_80); +return x_81; } } } else { -uint8_t x_81; -lean_dec(x_20); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_4); -x_81 = !lean_is_exclusive(x_10); -if (x_81 == 0) +lean_object* x_82; lean_object* x_83; lean_object* x_84; +x_82 = lean_ctor_get(x_7, 0); +x_83 = lean_ctor_get(x_7, 1); +lean_inc(x_83); +lean_inc(x_82); +lean_dec(x_7); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_82); +lean_inc(x_1); +x_84 = l_Lean_Meta_isExprDefEq(x_1, x_82, x_11, x_12, x_13, x_14, x_15); +if (lean_obj_tag(x_84) == 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_10, 1); -lean_inc(x_19); -x_83 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_82, x_1, x_19); -lean_ctor_set(x_10, 1, x_83); -x_84 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_84, 0, x_19); -x_85 = lean_box(0); -lean_ctor_set_tag(x_7, 0); -lean_ctor_set(x_7, 1, x_85); -lean_ctor_set(x_7, 0, x_84); -x_86 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_86, 0, x_7); -lean_ctor_set(x_86, 1, x_10); -lean_ctor_set(x_21, 0, x_86); -return x_21; -} -else +lean_object* x_85; uint8_t x_86; +x_85 = lean_ctor_get(x_84, 0); +lean_inc(x_85); +x_86 = lean_unbox(x_85); +lean_dec(x_85); +if (x_86 == 0) { -lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; -x_87 = lean_ctor_get(x_10, 0); -x_88 = lean_ctor_get(x_10, 1); -lean_inc(x_88); +lean_object* x_87; lean_object* x_88; lean_object* x_89; +x_87 = lean_ctor_get(x_84, 1); lean_inc(x_87); -lean_dec(x_10); -lean_inc(x_19); -x_89 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_88, x_1, x_19); -x_90 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_90, 0, x_87); -lean_ctor_set(x_90, 1, x_89); -x_91 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_91, 0, x_19); -x_92 = lean_box(0); -lean_ctor_set_tag(x_7, 0); -lean_ctor_set(x_7, 1, x_92); -lean_ctor_set(x_7, 0, x_91); -x_93 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_93, 0, x_7); -lean_ctor_set(x_93, 1, x_90); -lean_ctor_set(x_21, 0, x_93); -return x_21; -} -} -} -else -{ -lean_object* x_94; uint8_t x_95; -x_94 = lean_ctor_get(x_21, 1); -lean_inc(x_94); -lean_dec(x_21); -lean_inc(x_1); -x_95 = l_Lean_Expr_fvarsSubset(x_19, x_1); -if (x_95 == 0) -{ -lean_object* x_96; lean_object* x_97; -lean_free_object(x_7); -x_96 = lean_box(0); +lean_dec(x_84); +x_88 = lean_box(0); lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); lean_inc(x_1); lean_inc(x_4); -x_97 = l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1(x_2, x_4, x_1, x_19, x_96, x_10, x_11, x_12, x_13, x_14, x_94); -if (lean_obj_tag(x_97) == 0) +x_89 = l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1(x_2, x_4, x_1, x_82, x_88, x_10, x_11, x_12, x_13, x_14, x_87); +if (lean_obj_tag(x_89) == 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_90; lean_object* x_91; +x_90 = lean_ctor_get(x_89, 0); +lean_inc(x_90); +x_91 = lean_ctor_get(x_90, 0); +lean_inc(x_91); +if (lean_obj_tag(x_91) == 0) { -lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; -lean_dec(x_20); +lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; +lean_dec(x_83); lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_4); lean_dec(x_1); -x_100 = lean_ctor_get(x_97, 1); -lean_inc(x_100); -if (lean_is_exclusive(x_97)) { - lean_ctor_release(x_97, 0); - lean_ctor_release(x_97, 1); - x_101 = x_97; +x_92 = lean_ctor_get(x_89, 1); +lean_inc(x_92); +if (lean_is_exclusive(x_89)) { + lean_ctor_release(x_89, 0); + lean_ctor_release(x_89, 1); + x_93 = x_89; } else { - lean_dec_ref(x_97); - x_101 = lean_box(0); + lean_dec_ref(x_89); + x_93 = lean_box(0); } -x_102 = lean_ctor_get(x_98, 1); -lean_inc(x_102); -if (lean_is_exclusive(x_98)) { - lean_ctor_release(x_98, 0); - lean_ctor_release(x_98, 1); - x_103 = x_98; +x_94 = lean_ctor_get(x_90, 1); +lean_inc(x_94); +if (lean_is_exclusive(x_90)) { + lean_ctor_release(x_90, 0); + lean_ctor_release(x_90, 1); + x_95 = x_90; } else { - lean_dec_ref(x_98); - x_103 = lean_box(0); + lean_dec_ref(x_90); + x_95 = lean_box(0); } -x_104 = lean_ctor_get(x_99, 0); -lean_inc(x_104); -lean_dec(x_99); -if (lean_is_scalar(x_103)) { - x_105 = lean_alloc_ctor(0, 2, 0); +x_96 = lean_ctor_get(x_91, 0); +lean_inc(x_96); +lean_dec(x_91); +if (lean_is_scalar(x_95)) { + x_97 = lean_alloc_ctor(0, 2, 0); } else { - x_105 = x_103; + x_97 = x_95; } -lean_ctor_set(x_105, 0, x_104); -lean_ctor_set(x_105, 1, x_102); -if (lean_is_scalar(x_101)) { - x_106 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_97, 0, x_96); +lean_ctor_set(x_97, 1, x_94); +if (lean_is_scalar(x_93)) { + x_98 = lean_alloc_ctor(0, 2, 0); } else { - x_106 = x_101; + x_98 = x_93; } -lean_ctor_set(x_106, 0, x_105); -lean_ctor_set(x_106, 1, x_100); -return x_106; +lean_ctor_set(x_98, 0, x_97); +lean_ctor_set(x_98, 1, x_92); +return x_98; } else { -lean_object* x_107; lean_object* x_108; lean_object* x_109; -x_107 = lean_ctor_get(x_97, 1); -lean_inc(x_107); -lean_dec(x_97); -x_108 = lean_ctor_get(x_98, 1); -lean_inc(x_108); -lean_dec(x_98); -x_109 = lean_ctor_get(x_99, 0); -lean_inc(x_109); -lean_dec(x_99); -x_7 = x_20; -x_8 = x_109; +lean_object* x_99; lean_object* x_100; lean_object* x_101; +x_99 = lean_ctor_get(x_89, 1); +lean_inc(x_99); +lean_dec(x_89); +x_100 = lean_ctor_get(x_90, 1); +lean_inc(x_100); +lean_dec(x_90); +x_101 = lean_ctor_get(x_91, 0); +lean_inc(x_101); +lean_dec(x_91); +x_7 = x_83; +x_8 = x_101; x_9 = lean_box(0); -x_10 = x_108; -x_15 = x_107; +x_10 = x_100; +x_15 = x_99; goto _start; } } else { -lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; -lean_dec(x_20); +lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; +lean_dec(x_83); lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_4); lean_dec(x_1); -x_111 = lean_ctor_get(x_97, 0); -lean_inc(x_111); -x_112 = lean_ctor_get(x_97, 1); -lean_inc(x_112); -if (lean_is_exclusive(x_97)) { - lean_ctor_release(x_97, 0); - lean_ctor_release(x_97, 1); - x_113 = x_97; +x_103 = lean_ctor_get(x_89, 0); +lean_inc(x_103); +x_104 = lean_ctor_get(x_89, 1); +lean_inc(x_104); +if (lean_is_exclusive(x_89)) { + lean_ctor_release(x_89, 0); + lean_ctor_release(x_89, 1); + x_105 = x_89; } else { - lean_dec_ref(x_97); - x_113 = lean_box(0); + lean_dec_ref(x_89); + x_105 = lean_box(0); } -if (lean_is_scalar(x_113)) { - x_114 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_105)) { + x_106 = lean_alloc_ctor(1, 2, 0); } else { - x_114 = x_113; + x_106 = x_105; } -lean_ctor_set(x_114, 0, x_111); -lean_ctor_set(x_114, 1, x_112); -return x_114; +lean_ctor_set(x_106, 0, x_103); +lean_ctor_set(x_106, 1, x_104); +return x_106; } } else { -lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; -lean_dec(x_20); +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_dec(x_83); lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_4); -x_115 = lean_ctor_get(x_10, 0); -lean_inc(x_115); -x_116 = lean_ctor_get(x_10, 1); -lean_inc(x_116); +x_107 = lean_ctor_get(x_84, 1); +lean_inc(x_107); +if (lean_is_exclusive(x_84)) { + lean_ctor_release(x_84, 0); + lean_ctor_release(x_84, 1); + x_108 = x_84; +} else { + lean_dec_ref(x_84); + x_108 = lean_box(0); +} +x_109 = lean_ctor_get(x_10, 0); +lean_inc(x_109); +x_110 = lean_ctor_get(x_10, 1); +lean_inc(x_110); +x_111 = lean_ctor_get(x_10, 2); +lean_inc(x_111); if (lean_is_exclusive(x_10)) { lean_ctor_release(x_10, 0); lean_ctor_release(x_10, 1); - x_117 = x_10; + lean_ctor_release(x_10, 2); + x_112 = x_10; } else { lean_dec_ref(x_10); - x_117 = lean_box(0); + x_112 = lean_box(0); } -lean_inc(x_19); -x_118 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_116, x_1, x_19); -if (lean_is_scalar(x_117)) { - x_119 = lean_alloc_ctor(0, 2, 0); +lean_inc(x_82); +x_113 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_110, x_1, x_82); +if (lean_is_scalar(x_112)) { + x_114 = lean_alloc_ctor(0, 3, 0); } else { - x_119 = x_117; -} -lean_ctor_set(x_119, 0, x_115); -lean_ctor_set(x_119, 1, x_118); -x_120 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_120, 0, x_19); -x_121 = lean_box(0); -lean_ctor_set_tag(x_7, 0); -lean_ctor_set(x_7, 1, x_121); -lean_ctor_set(x_7, 0, x_120); -x_122 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_122, 0, x_7); -lean_ctor_set(x_122, 1, x_119); -x_123 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_123, 0, x_122); -lean_ctor_set(x_123, 1, x_94); -return x_123; + x_114 = x_112; } +lean_ctor_set(x_114, 0, x_109); +lean_ctor_set(x_114, 1, x_113); +lean_ctor_set(x_114, 2, x_111); +x_115 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_115, 0, x_82); +x_116 = lean_box(0); +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_117); +lean_ctor_set(x_118, 1, x_114); +if (lean_is_scalar(x_108)) { + x_119 = lean_alloc_ctor(0, 2, 0); +} else { + x_119 = x_108; } +lean_ctor_set(x_119, 0, x_118); +lean_ctor_set(x_119, 1, x_107); +return x_119; } } else { -uint8_t x_124; -lean_free_object(x_7); -lean_dec(x_20); -lean_dec(x_19); +lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; +lean_dec(x_83); +lean_dec(x_82); lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); @@ -2904,405 +2866,44 @@ lean_dec(x_11); lean_dec(x_10); lean_dec(x_4); lean_dec(x_1); -x_124 = !lean_is_exclusive(x_21); -if (x_124 == 0) -{ -return x_21; +x_120 = lean_ctor_get(x_84, 0); +lean_inc(x_120); +x_121 = lean_ctor_get(x_84, 1); +lean_inc(x_121); +if (lean_is_exclusive(x_84)) { + lean_ctor_release(x_84, 0); + lean_ctor_release(x_84, 1); + x_122 = x_84; +} else { + lean_dec_ref(x_84); + x_122 = lean_box(0); } -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; +if (lean_is_scalar(x_122)) { + x_123 = lean_alloc_ctor(1, 2, 0); +} else { + x_123 = x_122; +} +lean_ctor_set(x_123, 0, x_120); +lean_ctor_set(x_123, 1, x_121); +return x_123; } } } -else +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_Canon_canonElemCore___spec__13(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_128; lean_object* x_129; lean_object* x_130; -x_128 = lean_ctor_get(x_7, 0); -x_129 = lean_ctor_get(x_7, 1); -lean_inc(x_129); -lean_inc(x_128); +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); -lean_inc(x_14); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_128); -lean_inc(x_1); -x_130 = l_Lean_Meta_isExprDefEq(x_1, x_128, x_11, x_12, x_13, x_14, x_15); -if (lean_obj_tag(x_130) == 0) +if (x_8 == 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_133; lean_object* x_134; lean_object* x_135; -x_133 = lean_ctor_get(x_130, 1); -lean_inc(x_133); -lean_dec(x_130); -x_134 = lean_box(0); -lean_inc(x_14); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_1); -lean_inc(x_4); -x_135 = l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1(x_2, x_4, x_1, x_128, x_134, x_10, x_11, x_12, x_13, x_14, x_133); -if (lean_obj_tag(x_135) == 0) -{ -lean_object* x_136; lean_object* x_137; -x_136 = lean_ctor_get(x_135, 0); -lean_inc(x_136); -x_137 = lean_ctor_get(x_136, 0); -lean_inc(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_object* x_142; lean_object* x_143; lean_object* x_144; -lean_dec(x_129); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_4); -lean_dec(x_1); -x_138 = lean_ctor_get(x_135, 1); -lean_inc(x_138); -if (lean_is_exclusive(x_135)) { - lean_ctor_release(x_135, 0); - lean_ctor_release(x_135, 1); - x_139 = x_135; -} else { - lean_dec_ref(x_135); - x_139 = lean_box(0); -} -x_140 = lean_ctor_get(x_136, 1); -lean_inc(x_140); -if (lean_is_exclusive(x_136)) { - lean_ctor_release(x_136, 0); - lean_ctor_release(x_136, 1); - x_141 = x_136; -} else { - lean_dec_ref(x_136); - x_141 = lean_box(0); -} -x_142 = lean_ctor_get(x_137, 0); -lean_inc(x_142); -lean_dec(x_137); -if (lean_is_scalar(x_141)) { - x_143 = lean_alloc_ctor(0, 2, 0); -} else { - x_143 = x_141; -} -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_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; -x_145 = lean_ctor_get(x_135, 1); -lean_inc(x_145); -lean_dec(x_135); -x_146 = lean_ctor_get(x_136, 1); -lean_inc(x_146); -lean_dec(x_136); -x_147 = lean_ctor_get(x_137, 0); -lean_inc(x_147); -lean_dec(x_137); -x_7 = x_129; -x_8 = x_147; -x_9 = lean_box(0); -x_10 = x_146; -x_15 = x_145; -goto _start; -} -} -else -{ -lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; -lean_dec(x_129); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_4); -lean_dec(x_1); -x_149 = lean_ctor_get(x_135, 0); -lean_inc(x_149); -x_150 = lean_ctor_get(x_135, 1); -lean_inc(x_150); -if (lean_is_exclusive(x_135)) { - lean_ctor_release(x_135, 0); - lean_ctor_release(x_135, 1); - x_151 = x_135; -} else { - lean_dec_ref(x_135); - x_151 = lean_box(0); -} -if (lean_is_scalar(x_151)) { - x_152 = lean_alloc_ctor(1, 2, 0); -} else { - x_152 = x_151; -} -lean_ctor_set(x_152, 0, x_149); -lean_ctor_set(x_152, 1, x_150); -return x_152; -} -} -else -{ -lean_object* x_153; lean_object* x_154; uint8_t x_155; -x_153 = lean_ctor_get(x_130, 1); -lean_inc(x_153); -if (lean_is_exclusive(x_130)) { - lean_ctor_release(x_130, 0); - lean_ctor_release(x_130, 1); - x_154 = x_130; -} else { - lean_dec_ref(x_130); - x_154 = lean_box(0); -} -lean_inc(x_1); -x_155 = l_Lean_Expr_fvarsSubset(x_128, x_1); -if (x_155 == 0) -{ -lean_object* x_156; lean_object* x_157; -lean_dec(x_154); -x_156 = lean_box(0); -lean_inc(x_14); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_1); -lean_inc(x_4); -x_157 = l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1(x_2, x_4, x_1, x_128, x_156, x_10, x_11, x_12, x_13, x_14, x_153); -if (lean_obj_tag(x_157) == 0) -{ -lean_object* x_158; lean_object* x_159; -x_158 = lean_ctor_get(x_157, 0); -lean_inc(x_158); -x_159 = lean_ctor_get(x_158, 0); -lean_inc(x_159); -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; lean_object* x_166; -lean_dec(x_129); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_4); -lean_dec(x_1); -x_160 = lean_ctor_get(x_157, 1); -lean_inc(x_160); -if (lean_is_exclusive(x_157)) { - lean_ctor_release(x_157, 0); - lean_ctor_release(x_157, 1); - x_161 = x_157; -} else { - lean_dec_ref(x_157); - x_161 = lean_box(0); -} -x_162 = lean_ctor_get(x_158, 1); -lean_inc(x_162); -if (lean_is_exclusive(x_158)) { - lean_ctor_release(x_158, 0); - lean_ctor_release(x_158, 1); - x_163 = x_158; -} else { - lean_dec_ref(x_158); - x_163 = lean_box(0); -} -x_164 = lean_ctor_get(x_159, 0); -lean_inc(x_164); -lean_dec(x_159); -if (lean_is_scalar(x_163)) { - x_165 = lean_alloc_ctor(0, 2, 0); -} else { - x_165 = x_163; -} -lean_ctor_set(x_165, 0, x_164); -lean_ctor_set(x_165, 1, x_162); -if (lean_is_scalar(x_161)) { - x_166 = lean_alloc_ctor(0, 2, 0); -} else { - x_166 = x_161; -} -lean_ctor_set(x_166, 0, x_165); -lean_ctor_set(x_166, 1, x_160); -return x_166; -} -else -{ -lean_object* x_167; lean_object* x_168; lean_object* x_169; -x_167 = lean_ctor_get(x_157, 1); -lean_inc(x_167); -lean_dec(x_157); -x_168 = lean_ctor_get(x_158, 1); -lean_inc(x_168); -lean_dec(x_158); -x_169 = lean_ctor_get(x_159, 0); -lean_inc(x_169); -lean_dec(x_159); -x_7 = x_129; -x_8 = x_169; -x_9 = lean_box(0); -x_10 = x_168; -x_15 = x_167; -goto _start; -} -} -else -{ -lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; -lean_dec(x_129); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_4); -lean_dec(x_1); -x_171 = lean_ctor_get(x_157, 0); -lean_inc(x_171); -x_172 = lean_ctor_get(x_157, 1); -lean_inc(x_172); -if (lean_is_exclusive(x_157)) { - lean_ctor_release(x_157, 0); - lean_ctor_release(x_157, 1); - x_173 = x_157; -} else { - lean_dec_ref(x_157); - x_173 = lean_box(0); -} -if (lean_is_scalar(x_173)) { - x_174 = lean_alloc_ctor(1, 2, 0); -} else { - x_174 = x_173; -} -lean_ctor_set(x_174, 0, x_171); -lean_ctor_set(x_174, 1, x_172); -return x_174; -} -} -else -{ -lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; -lean_dec(x_129); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_4); -x_175 = lean_ctor_get(x_10, 0); -lean_inc(x_175); -x_176 = lean_ctor_get(x_10, 1); -lean_inc(x_176); -if (lean_is_exclusive(x_10)) { - lean_ctor_release(x_10, 0); - lean_ctor_release(x_10, 1); - x_177 = x_10; -} else { - lean_dec_ref(x_10); - x_177 = lean_box(0); -} -lean_inc(x_128); -x_178 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_176, x_1, x_128); -if (lean_is_scalar(x_177)) { - x_179 = lean_alloc_ctor(0, 2, 0); -} else { - x_179 = x_177; -} -lean_ctor_set(x_179, 0, x_175); -lean_ctor_set(x_179, 1, x_178); -x_180 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_180, 0, x_128); -x_181 = lean_box(0); -x_182 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_182, 0, x_180); -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_179); -if (lean_is_scalar(x_154)) { - x_184 = lean_alloc_ctor(0, 2, 0); -} else { - x_184 = x_154; -} -lean_ctor_set(x_184, 0, x_183); -lean_ctor_set(x_184, 1, x_153); -return x_184; -} -} -} -else -{ -lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; -lean_dec(x_129); -lean_dec(x_128); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_4); -lean_dec(x_1); -x_185 = lean_ctor_get(x_130, 0); -lean_inc(x_185); -x_186 = lean_ctor_get(x_130, 1); -lean_inc(x_186); -if (lean_is_exclusive(x_130)) { - lean_ctor_release(x_130, 0); - lean_ctor_release(x_130, 1); - x_187 = x_130; -} else { - lean_dec_ref(x_130); - x_187 = lean_box(0); -} -if (lean_is_scalar(x_187)) { - x_188 = lean_alloc_ctor(1, 2, 0); -} else { - x_188 = x_187; -} -lean_ctor_set(x_188, 0, x_185); -lean_ctor_set(x_188, 1, x_186); -return x_188; -} -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_Canon_canonElemCore___spec__13(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_dec(x_5); +return x_6; +} +else { lean_object* x_9; lean_object* x_10; size_t x_11; size_t x_12; size_t x_13; size_t x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint64_t x_19; uint64_t x_20; uint64_t x_21; size_t x_22; size_t x_23; lean_object* x_24; x_9 = lean_array_fget(x_2, x_5); @@ -4128,30 +3729,60 @@ return x_5; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_canonElemCore___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_11 = lean_ctor_get(x_5, 0); -lean_inc(x_11); +uint8_t x_11; +x_11 = !lean_is_exclusive(x_5); +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; +x_12 = lean_ctor_get(x_5, 0); +x_13 = lean_ctor_get(x_5, 1); lean_inc(x_1); -x_12 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_12, 0, x_1); -lean_ctor_set(x_12, 1, x_2); -x_13 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__11(x_11, x_3, x_12); -x_14 = lean_ctor_get(x_5, 1); -lean_inc(x_14); -lean_dec(x_5); +x_14 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_14, 0, x_1); +lean_ctor_set(x_14, 1, x_2); +x_15 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__11(x_12, x_3, x_14); lean_inc_n(x_1, 2); -x_15 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_14, x_1, x_1); -x_16 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_16, 0, x_13); -lean_ctor_set(x_16, 1, x_15); +x_16 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_13, x_1, x_1); +lean_ctor_set(x_5, 1, x_16); +lean_ctor_set(x_5, 0, x_15); x_17 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_17, 0, x_1); -lean_ctor_set(x_17, 1, x_16); +lean_ctor_set(x_17, 1, x_5); x_18 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_18, 0, x_17); lean_ctor_set(x_18, 1, x_10); 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; +x_19 = lean_ctor_get(x_5, 0); +x_20 = lean_ctor_get(x_5, 1); +x_21 = lean_ctor_get(x_5, 2); +lean_inc(x_21); +lean_inc(x_20); +lean_inc(x_19); +lean_dec(x_5); +lean_inc(x_1); +x_22 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_22, 0, x_1); +lean_ctor_set(x_22, 1, x_2); +x_23 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__11(x_19, x_3, x_22); +lean_inc_n(x_1, 2); +x_24 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_20, x_1, x_1); +x_25 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_24); +lean_ctor_set(x_25, 2, x_21); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_1); +lean_ctor_set(x_26, 1, x_25); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_10); +return x_27; +} +} } static lean_object* _init_l_Lean_Meta_Grind_Canon_canonElemCore___lambda__2___closed__1() { _start: @@ -7210,565 +6841,165 @@ LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canon switch (lean_obj_tag(x_2)) { case 0: { -lean_object* x_12; lean_object* x_46; uint8_t x_47; +lean_object* x_12; lean_object* x_43; lean_object* x_77; uint8_t x_78; lean_dec(x_4); -x_46 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__4; -x_47 = l_Lean_Expr_isConstOf(x_2, x_46); -if (x_47 == 0) +x_77 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__4; +x_78 = l_Lean_Expr_isConstOf(x_2, x_77); +if (x_78 == 0) { -lean_object* x_48; -x_48 = lean_box(0); -x_12 = x_48; -goto block_45; +lean_object* x_79; +x_79 = lean_box(0); +x_43 = x_79; +goto block_76; } else { -lean_object* x_49; lean_object* x_50; uint8_t x_51; -x_49 = lean_array_get_size(x_3); -x_50 = lean_unsigned_to_nat(2u); -x_51 = lean_nat_dec_eq(x_49, x_50); -if (x_51 == 0) +lean_object* x_80; lean_object* x_81; uint8_t x_82; +x_80 = lean_array_get_size(x_3); +x_81 = lean_unsigned_to_nat(2u); +x_82 = lean_nat_dec_eq(x_80, x_81); +if (x_82 == 0) { -lean_object* x_52; -lean_dec(x_49); -x_52 = lean_box(0); -x_12 = x_52; -goto block_45; +lean_object* x_83; +lean_dec(x_80); +x_83 = lean_box(0); +x_43 = x_83; +goto block_76; } else { -lean_object* x_53; uint8_t x_54; -x_53 = lean_unsigned_to_nat(0u); -x_54 = lean_nat_dec_lt(x_53, x_49); -lean_dec(x_49); -if (x_54 == 0) +lean_object* x_84; uint8_t x_85; +x_84 = lean_unsigned_to_nat(0u); +x_85 = lean_nat_dec_lt(x_84, x_80); +lean_dec(x_80); +if (x_85 == 0) +{ +lean_object* x_86; lean_object* x_87; +x_86 = l_Lean_instInhabitedExpr; +x_87 = l_outOfBounds___rarg(x_86); +x_12 = x_87; +goto block_42; +} +else +{ +lean_object* x_88; +x_88 = lean_array_fget(x_3, x_84); +x_12 = x_88; +goto block_42; +} +} +} +block_42: { -lean_object* x_55; lean_object* x_56; lean_object* x_57; -x_55 = l_Lean_instInhabitedExpr; -x_56 = l_outOfBounds___rarg(x_55); +lean_object* x_13; lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_5); -lean_inc(x_56); -x_57 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_56, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_57) == 0) +lean_inc(x_12); +x_13 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_12, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_13) == 0) { -lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; size_t x_62; size_t 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); -x_60 = lean_ctor_get(x_58, 0); -lean_inc(x_60); -x_61 = lean_ctor_get(x_58, 1); -lean_inc(x_61); -lean_dec(x_58); -x_62 = lean_ptr_addr(x_56); -lean_dec(x_56); -x_63 = lean_ptr_addr(x_60); -x_64 = lean_usize_dec_eq(x_62, x_63); -if (x_64 == 0) +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +x_16 = lean_ctor_get(x_14, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_14, 1); +lean_inc(x_17); +lean_dec(x_14); +x_18 = lean_ctor_get(x_17, 2); +lean_inc(x_18); +lean_inc(x_18); +x_19 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_18, x_16); +if (lean_obj_tag(x_19) == 0) { -lean_object* x_65; -lean_dec(x_60); -lean_dec(x_3); -lean_dec(x_2); -lean_inc(x_1); -x_65 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_61, x_7, x_8, x_9, x_10, x_59); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_65; -} -else -{ -lean_object* x_66; lean_object* x_67; lean_object* x_68; -x_66 = lean_array_set(x_3, x_53, x_60); -x_67 = l_Lean_mkAppN(x_2, x_66); -lean_dec(x_66); -x_68 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_67, x_5, x_61, x_7, x_8, x_9, x_10, x_59); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_68; -} -} -else -{ -uint8_t x_69; -lean_dec(x_56); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_69 = !lean_is_exclusive(x_57); -if (x_69 == 0) -{ -return x_57; -} -else -{ -lean_object* x_70; lean_object* x_71; lean_object* x_72; -x_70 = lean_ctor_get(x_57, 0); -x_71 = lean_ctor_get(x_57, 1); -lean_inc(x_71); -lean_inc(x_70); -lean_dec(x_57); -x_72 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_72, 0, x_70); -lean_ctor_set(x_72, 1, x_71); -return x_72; -} -} -} -else -{ -lean_object* x_73; lean_object* x_74; -x_73 = lean_array_fget(x_3, x_53); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_5); -lean_inc(x_73); -x_74 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_73, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_74) == 0) -{ -lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; size_t x_79; size_t x_80; uint8_t x_81; -x_75 = lean_ctor_get(x_74, 0); -lean_inc(x_75); -x_76 = lean_ctor_get(x_74, 1); -lean_inc(x_76); -lean_dec(x_74); -x_77 = lean_ctor_get(x_75, 0); -lean_inc(x_77); -x_78 = lean_ctor_get(x_75, 1); -lean_inc(x_78); -lean_dec(x_75); -x_79 = lean_ptr_addr(x_73); -lean_dec(x_73); -x_80 = lean_ptr_addr(x_77); -x_81 = lean_usize_dec_eq(x_79, x_80); -if (x_81 == 0) -{ -lean_object* x_82; -lean_dec(x_77); -lean_dec(x_3); -lean_dec(x_2); -lean_inc(x_1); -x_82 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_78, x_7, x_8, x_9, x_10, x_76); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_82; -} -else -{ -lean_object* x_83; lean_object* x_84; lean_object* x_85; -x_83 = lean_array_set(x_3, x_53, x_77); -x_84 = l_Lean_mkAppN(x_2, x_83); -lean_dec(x_83); -x_85 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_84, x_5, x_78, x_7, x_8, x_9, x_10, x_76); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_85; -} -} -else -{ -uint8_t x_86; -lean_dec(x_73); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_86 = !lean_is_exclusive(x_74); -if (x_86 == 0) -{ -return x_74; -} -else -{ -lean_object* x_87; lean_object* x_88; lean_object* x_89; -x_87 = lean_ctor_get(x_74, 0); -x_88 = lean_ctor_get(x_74, 1); -lean_inc(x_88); -lean_inc(x_87); -lean_dec(x_74); -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; -} -} -} -} -} -block_45: -{ -lean_object* x_13; +size_t x_20; size_t x_21; uint8_t x_22; +x_20 = lean_ptr_addr(x_12); lean_dec(x_12); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_2); -x_13 = l_Lean_Meta_getFunInfo(x_2, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_13) == 0) +x_21 = lean_ptr_addr(x_16); +x_22 = lean_usize_dec_eq(x_20, x_21); +if (x_22 == 0) { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; -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_object* x_23; lean_object* x_24; 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_unsigned_to_nat(0u); lean_inc(x_16); -lean_dec(x_14); -x_17 = lean_array_get_size(x_3); -x_18 = lean_unsigned_to_nat(0u); -x_19 = lean_unsigned_to_nat(1u); -x_20 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_20, 0, x_18); -lean_ctor_set(x_20, 1, x_17); -lean_ctor_set(x_20, 2, x_19); -x_21 = 0; -x_22 = lean_box(x_21); -x_23 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_23, 0, x_3); -lean_ctor_set(x_23, 1, x_22); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_5); -lean_inc(x_2); -x_24 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8(x_2, x_16, x_20, x_20, x_23, x_18, lean_box(0), lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10, x_15); -lean_dec(x_20); -lean_dec(x_16); -if (lean_obj_tag(x_24) == 0) -{ -lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; -x_25 = lean_ctor_get(x_24, 0); -lean_inc(x_25); -x_26 = lean_ctor_get(x_25, 0); +x_24 = lean_array_set(x_3, x_23, x_16); +x_25 = l_Lean_mkAppN(x_2, x_24); +lean_dec(x_24); +x_26 = lean_ctor_get(x_17, 0); lean_inc(x_26); -x_27 = lean_ctor_get(x_26, 1); +x_27 = lean_ctor_get(x_17, 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_object* x_31; -lean_dec(x_26); -lean_dec(x_2); -x_29 = lean_ctor_get(x_24, 1); -lean_inc(x_29); -lean_dec(x_24); -x_30 = lean_ctor_get(x_25, 1); -lean_inc(x_30); -lean_dec(x_25); -lean_inc(x_1); -x_31 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_30, x_7, x_8, x_9, x_10, x_29); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_31; -} -else -{ -lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_32 = lean_ctor_get(x_24, 1); -lean_inc(x_32); -lean_dec(x_24); -x_33 = lean_ctor_get(x_25, 1); -lean_inc(x_33); -lean_dec(x_25); -x_34 = lean_ctor_get(x_26, 0); -lean_inc(x_34); -lean_dec(x_26); -x_35 = l_Lean_mkAppN(x_2, x_34); -lean_dec(x_34); -x_36 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_35, x_5, x_33, x_7, x_8, x_9, x_10, x_32); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_36; -} -} -else -{ -uint8_t x_37; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -lean_dec(x_2); -lean_dec(x_1); -x_37 = !lean_is_exclusive(x_24); -if (x_37 == 0) -{ -return x_24; -} -else -{ -lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_38 = lean_ctor_get(x_24, 0); -x_39 = lean_ctor_get(x_24, 1); -lean_inc(x_39); -lean_inc(x_38); -lean_dec(x_24); -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_17); +lean_inc(x_25); +x_28 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_18, x_16, x_25); +x_29 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_29, 0, x_26); +lean_ctor_set(x_29, 1, x_27); +lean_ctor_set(x_29, 2, x_28); +x_30 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_25, x_5, x_29, x_7, x_8, x_9, x_10, x_15); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_6); lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_41 = !lean_is_exclusive(x_13); -if (x_41 == 0) -{ -return x_13; -} -else -{ -lean_object* x_42; lean_object* x_43; lean_object* x_44; -x_42 = lean_ctor_get(x_13, 0); -x_43 = lean_ctor_get(x_13, 1); -lean_inc(x_43); -lean_inc(x_42); -lean_dec(x_13); -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; -} -} -} -} -case 1: -{ -lean_object* x_90; lean_object* x_124; uint8_t x_125; -lean_dec(x_4); -x_124 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__4; -x_125 = l_Lean_Expr_isConstOf(x_2, x_124); -if (x_125 == 0) -{ -lean_object* x_126; -x_126 = lean_box(0); -x_90 = x_126; -goto block_123; -} -else -{ -lean_object* x_127; lean_object* x_128; uint8_t x_129; -x_127 = lean_array_get_size(x_3); -x_128 = lean_unsigned_to_nat(2u); -x_129 = lean_nat_dec_eq(x_127, x_128); -if (x_129 == 0) -{ -lean_object* x_130; -lean_dec(x_127); -x_130 = lean_box(0); -x_90 = x_130; -goto block_123; +return x_30; } else { -lean_object* x_131; uint8_t x_132; -x_131 = lean_unsigned_to_nat(0u); -x_132 = lean_nat_dec_lt(x_131, x_127); -lean_dec(x_127); -if (x_132 == 0) -{ -lean_object* x_133; lean_object* x_134; lean_object* x_135; -x_133 = l_Lean_instInhabitedExpr; -x_134 = l_outOfBounds___rarg(x_133); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_5); -lean_inc(x_134); -x_135 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_134, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_135) == 0) -{ -lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; size_t x_140; size_t x_141; uint8_t x_142; -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 = lean_ctor_get(x_136, 0); -lean_inc(x_138); -x_139 = lean_ctor_get(x_136, 1); -lean_inc(x_139); -lean_dec(x_136); -x_140 = lean_ptr_addr(x_134); -lean_dec(x_134); -x_141 = lean_ptr_addr(x_138); -x_142 = lean_usize_dec_eq(x_140, x_141); -if (x_142 == 0) -{ -lean_object* x_143; -lean_dec(x_138); +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_dec(x_3); lean_dec(x_2); +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); lean_inc(x_1); -x_143 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_139, x_7, x_8, x_9, x_10, x_137); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_143; -} -else -{ -lean_object* x_144; lean_object* x_145; lean_object* x_146; -x_144 = lean_array_set(x_3, x_131, x_138); -x_145 = l_Lean_mkAppN(x_2, x_144); -lean_dec(x_144); -x_146 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_145, x_5, x_139, x_7, x_8, x_9, x_10, x_137); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_146; -} -} -else -{ -uint8_t x_147; -lean_dec(x_134); +x_33 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_18, x_16, x_1); +x_34 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_34, 0, x_31); +lean_ctor_set(x_34, 1, x_32); +lean_ctor_set(x_34, 2, x_33); +lean_inc(x_1); +x_35 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_34, x_7, x_8, x_9, x_10, x_15); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_147 = !lean_is_exclusive(x_135); -if (x_147 == 0) -{ -return x_135; -} -else -{ -lean_object* x_148; lean_object* x_149; lean_object* x_150; -x_148 = lean_ctor_get(x_135, 0); -x_149 = lean_ctor_get(x_135, 1); -lean_inc(x_149); -lean_inc(x_148); -lean_dec(x_135); -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; -} +return x_35; } } else { -lean_object* x_151; lean_object* x_152; -x_151 = lean_array_fget(x_3, x_131); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_5); -lean_inc(x_151); -x_152 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_151, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_152) == 0) -{ -lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; size_t x_157; size_t x_158; uint8_t x_159; -x_153 = lean_ctor_get(x_152, 0); -lean_inc(x_153); -x_154 = lean_ctor_get(x_152, 1); -lean_inc(x_154); -lean_dec(x_152); -x_155 = lean_ctor_get(x_153, 0); -lean_inc(x_155); -x_156 = lean_ctor_get(x_153, 1); -lean_inc(x_156); -lean_dec(x_153); -x_157 = lean_ptr_addr(x_151); -lean_dec(x_151); -x_158 = lean_ptr_addr(x_155); -x_159 = lean_usize_dec_eq(x_157, x_158); -if (x_159 == 0) -{ -lean_object* x_160; -lean_dec(x_155); +lean_object* x_36; lean_object* x_37; +lean_dec(x_18); +lean_dec(x_16); +lean_dec(x_12); lean_dec(x_3); lean_dec(x_2); -lean_inc(x_1); -x_160 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_156, x_7, x_8, x_9, x_10, x_154); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_160; -} -else -{ -lean_object* x_161; lean_object* x_162; lean_object* x_163; -x_161 = lean_array_set(x_3, x_131, x_155); -x_162 = l_Lean_mkAppN(x_2, x_161); -lean_dec(x_161); -x_163 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_162, x_5, x_156, x_7, x_8, x_9, x_10, x_154); +x_36 = lean_ctor_get(x_19, 0); +lean_inc(x_36); +lean_dec(x_19); +x_37 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_36, x_5, x_17, x_7, x_8, x_9, x_10, x_15); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_5); -return x_163; +return x_37; } } else { -uint8_t x_164; -lean_dec(x_151); +uint8_t x_38; +lean_dec(x_12); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -7777,127 +7008,125 @@ lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_164 = !lean_is_exclusive(x_152); -if (x_164 == 0) +x_38 = !lean_is_exclusive(x_13); +if (x_38 == 0) { -return x_152; +return x_13; } else { -lean_object* x_165; lean_object* x_166; lean_object* x_167; -x_165 = lean_ctor_get(x_152, 0); -x_166 = lean_ctor_get(x_152, 1); -lean_inc(x_166); -lean_inc(x_165); -lean_dec(x_152); -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; -} -} +lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_39 = lean_ctor_get(x_13, 0); +x_40 = lean_ctor_get(x_13, 1); +lean_inc(x_40); +lean_inc(x_39); +lean_dec(x_13); +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; } } } -block_123: +block_76: { -lean_object* x_91; -lean_dec(x_90); +lean_object* x_44; +lean_dec(x_43); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_2); -x_91 = l_Lean_Meta_getFunInfo(x_2, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_91) == 0) +x_44 = l_Lean_Meta_getFunInfo(x_2, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_44) == 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; uint8_t x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; -x_92 = lean_ctor_get(x_91, 0); -lean_inc(x_92); -x_93 = lean_ctor_get(x_91, 1); -lean_inc(x_93); -lean_dec(x_91); -x_94 = lean_ctor_get(x_92, 0); -lean_inc(x_94); -lean_dec(x_92); -x_95 = lean_array_get_size(x_3); -x_96 = lean_unsigned_to_nat(0u); -x_97 = lean_unsigned_to_nat(1u); -x_98 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_98, 0, x_96); -lean_ctor_set(x_98, 1, x_95); -lean_ctor_set(x_98, 2, x_97); -x_99 = 0; -x_100 = lean_box(x_99); -x_101 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_101, 0, x_3); -lean_ctor_set(x_101, 1, x_100); +lean_object* x_45; 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; +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 = lean_array_get_size(x_3); +x_49 = lean_unsigned_to_nat(0u); +x_50 = lean_unsigned_to_nat(1u); +x_51 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_51, 0, x_49); +lean_ctor_set(x_51, 1, x_48); +lean_ctor_set(x_51, 2, x_50); +x_52 = 0; +x_53 = lean_box(x_52); +x_54 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_54, 0, x_3); +lean_ctor_set(x_54, 1, x_53); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_5); lean_inc(x_2); -x_102 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8(x_2, x_94, x_98, x_98, x_101, x_96, lean_box(0), lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10, x_93); -lean_dec(x_98); -lean_dec(x_94); -if (lean_obj_tag(x_102) == 0) +x_55 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8(x_2, x_47, x_51, x_51, x_54, x_49, lean_box(0), lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10, x_46); +lean_dec(x_51); +lean_dec(x_47); +if (lean_obj_tag(x_55) == 0) { -lean_object* x_103; lean_object* x_104; lean_object* x_105; uint8_t x_106; -x_103 = lean_ctor_get(x_102, 0); -lean_inc(x_103); -x_104 = lean_ctor_get(x_103, 0); -lean_inc(x_104); -x_105 = lean_ctor_get(x_104, 1); -lean_inc(x_105); -x_106 = lean_unbox(x_105); -lean_dec(x_105); -if (x_106 == 0) +lean_object* x_56; lean_object* x_57; lean_object* x_58; uint8_t x_59; +x_56 = lean_ctor_get(x_55, 0); +lean_inc(x_56); +x_57 = lean_ctor_get(x_56, 0); +lean_inc(x_57); +x_58 = lean_ctor_get(x_57, 1); +lean_inc(x_58); +x_59 = lean_unbox(x_58); +lean_dec(x_58); +if (x_59 == 0) { -lean_object* x_107; lean_object* x_108; lean_object* x_109; -lean_dec(x_104); +lean_object* x_60; lean_object* x_61; lean_object* x_62; +lean_dec(x_57); lean_dec(x_2); -x_107 = lean_ctor_get(x_102, 1); -lean_inc(x_107); -lean_dec(x_102); -x_108 = lean_ctor_get(x_103, 1); -lean_inc(x_108); -lean_dec(x_103); +x_60 = lean_ctor_get(x_55, 1); +lean_inc(x_60); +lean_dec(x_55); +x_61 = lean_ctor_get(x_56, 1); +lean_inc(x_61); +lean_dec(x_56); lean_inc(x_1); -x_109 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_108, x_7, x_8, x_9, x_10, x_107); +x_62 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_61, x_7, x_8, x_9, x_10, x_60); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_5); -return x_109; +return x_62; } -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_102, 1); -lean_inc(x_110); -lean_dec(x_102); -x_111 = lean_ctor_get(x_103, 1); -lean_inc(x_111); -lean_dec(x_103); -x_112 = lean_ctor_get(x_104, 0); -lean_inc(x_112); -lean_dec(x_104); -x_113 = l_Lean_mkAppN(x_2, x_112); -lean_dec(x_112); -x_114 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_113, x_5, x_111, x_7, x_8, x_9, x_10, x_110); +else +{ +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_55, 1); +lean_inc(x_63); +lean_dec(x_55); +x_64 = lean_ctor_get(x_56, 1); +lean_inc(x_64); +lean_dec(x_56); +x_65 = lean_ctor_get(x_57, 0); +lean_inc(x_65); +lean_dec(x_57); +x_66 = l_Lean_mkAppN(x_2, x_65); +lean_dec(x_65); +x_67 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_66, x_5, x_64, x_7, x_8, x_9, x_10, x_63); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_5); -return x_114; +return x_67; } } else { -uint8_t x_115; +uint8_t x_68; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -7905,29 +7134,29 @@ lean_dec(x_7); lean_dec(x_5); lean_dec(x_2); lean_dec(x_1); -x_115 = !lean_is_exclusive(x_102); -if (x_115 == 0) +x_68 = !lean_is_exclusive(x_55); +if (x_68 == 0) { -return x_102; +return x_55; } else { -lean_object* x_116; lean_object* x_117; lean_object* x_118; -x_116 = lean_ctor_get(x_102, 0); -x_117 = lean_ctor_get(x_102, 1); -lean_inc(x_117); -lean_inc(x_116); -lean_dec(x_102); -x_118 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_118, 0, x_116); -lean_ctor_set(x_118, 1, x_117); -return x_118; +lean_object* x_69; lean_object* x_70; lean_object* x_71; +x_69 = lean_ctor_get(x_55, 0); +x_70 = lean_ctor_get(x_55, 1); +lean_inc(x_70); +lean_inc(x_69); +lean_dec(x_55); +x_71 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_71, 0, x_69); +lean_ctor_set(x_71, 1, x_70); +return x_71; } } } else { -uint8_t x_119; +uint8_t x_72; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -7937,213 +7166,188 @@ lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_119 = !lean_is_exclusive(x_91); -if (x_119 == 0) +x_72 = !lean_is_exclusive(x_44); +if (x_72 == 0) { -return x_91; +return x_44; } else { -lean_object* x_120; lean_object* x_121; lean_object* x_122; -x_120 = lean_ctor_get(x_91, 0); -x_121 = lean_ctor_get(x_91, 1); -lean_inc(x_121); -lean_inc(x_120); -lean_dec(x_91); -x_122 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_122, 0, x_120); -lean_ctor_set(x_122, 1, x_121); -return x_122; +lean_object* x_73; lean_object* x_74; lean_object* x_75; +x_73 = lean_ctor_get(x_44, 0); +x_74 = lean_ctor_get(x_44, 1); +lean_inc(x_74); +lean_inc(x_73); +lean_dec(x_44); +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; } } } } -case 2: +case 1: { -lean_object* x_168; lean_object* x_202; uint8_t x_203; +lean_object* x_89; lean_object* x_120; lean_object* x_154; uint8_t x_155; lean_dec(x_4); -x_202 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__4; -x_203 = l_Lean_Expr_isConstOf(x_2, x_202); -if (x_203 == 0) +x_154 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__4; +x_155 = l_Lean_Expr_isConstOf(x_2, x_154); +if (x_155 == 0) { -lean_object* x_204; -x_204 = lean_box(0); -x_168 = x_204; -goto block_201; +lean_object* x_156; +x_156 = lean_box(0); +x_120 = x_156; +goto block_153; } else { -lean_object* x_205; lean_object* x_206; uint8_t x_207; -x_205 = lean_array_get_size(x_3); -x_206 = lean_unsigned_to_nat(2u); -x_207 = lean_nat_dec_eq(x_205, x_206); -if (x_207 == 0) +lean_object* x_157; lean_object* x_158; uint8_t x_159; +x_157 = lean_array_get_size(x_3); +x_158 = lean_unsigned_to_nat(2u); +x_159 = lean_nat_dec_eq(x_157, x_158); +if (x_159 == 0) { -lean_object* x_208; -lean_dec(x_205); -x_208 = lean_box(0); -x_168 = x_208; -goto block_201; +lean_object* x_160; +lean_dec(x_157); +x_160 = lean_box(0); +x_120 = x_160; +goto block_153; } else { -lean_object* x_209; uint8_t x_210; -x_209 = lean_unsigned_to_nat(0u); -x_210 = lean_nat_dec_lt(x_209, x_205); -lean_dec(x_205); -if (x_210 == 0) +lean_object* x_161; uint8_t x_162; +x_161 = lean_unsigned_to_nat(0u); +x_162 = lean_nat_dec_lt(x_161, x_157); +lean_dec(x_157); +if (x_162 == 0) +{ +lean_object* x_163; lean_object* x_164; +x_163 = l_Lean_instInhabitedExpr; +x_164 = l_outOfBounds___rarg(x_163); +x_89 = x_164; +goto block_119; +} +else +{ +lean_object* x_165; +x_165 = lean_array_fget(x_3, x_161); +x_89 = x_165; +goto block_119; +} +} +} +block_119: { -lean_object* x_211; lean_object* x_212; lean_object* x_213; -x_211 = l_Lean_instInhabitedExpr; -x_212 = l_outOfBounds___rarg(x_211); +lean_object* x_90; lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_5); -lean_inc(x_212); -x_213 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_212, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_213) == 0) -{ -lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; size_t x_218; size_t x_219; uint8_t x_220; -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_ctor_get(x_214, 0); -lean_inc(x_216); -x_217 = lean_ctor_get(x_214, 1); -lean_inc(x_217); -lean_dec(x_214); -x_218 = lean_ptr_addr(x_212); -lean_dec(x_212); -x_219 = lean_ptr_addr(x_216); -x_220 = lean_usize_dec_eq(x_218, x_219); -if (x_220 == 0) +lean_inc(x_89); +x_90 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_89, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_90) == 0) { -lean_object* x_221; -lean_dec(x_216); -lean_dec(x_3); -lean_dec(x_2); -lean_inc(x_1); -x_221 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_217, x_7, x_8, x_9, x_10, x_215); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_221; -} -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; +x_91 = lean_ctor_get(x_90, 0); +lean_inc(x_91); +x_92 = lean_ctor_get(x_90, 1); +lean_inc(x_92); +lean_dec(x_90); +x_93 = lean_ctor_get(x_91, 0); +lean_inc(x_93); +x_94 = lean_ctor_get(x_91, 1); +lean_inc(x_94); +lean_dec(x_91); +x_95 = lean_ctor_get(x_94, 2); +lean_inc(x_95); +lean_inc(x_95); +x_96 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_95, x_93); +if (lean_obj_tag(x_96) == 0) { -lean_object* x_222; lean_object* x_223; lean_object* x_224; -x_222 = lean_array_set(x_3, x_209, x_216); -x_223 = l_Lean_mkAppN(x_2, x_222); -lean_dec(x_222); -x_224 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_223, x_5, x_217, x_7, x_8, x_9, x_10, x_215); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_224; -} -} -else +size_t x_97; size_t x_98; uint8_t x_99; +x_97 = lean_ptr_addr(x_89); +lean_dec(x_89); +x_98 = lean_ptr_addr(x_93); +x_99 = lean_usize_dec_eq(x_97, x_98); +if (x_99 == 0) { -uint8_t x_225; -lean_dec(x_212); +lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; +x_100 = lean_unsigned_to_nat(0u); +lean_inc(x_93); +x_101 = lean_array_set(x_3, x_100, x_93); +x_102 = l_Lean_mkAppN(x_2, x_101); +lean_dec(x_101); +x_103 = lean_ctor_get(x_94, 0); +lean_inc(x_103); +x_104 = lean_ctor_get(x_94, 1); +lean_inc(x_104); +lean_dec(x_94); +lean_inc(x_102); +x_105 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_95, x_93, x_102); +x_106 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_106, 0, x_103); +lean_ctor_set(x_106, 1, x_104); +lean_ctor_set(x_106, 2, x_105); +x_107 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_102, x_5, x_106, x_7, x_8, x_9, x_10, x_92); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_225 = !lean_is_exclusive(x_213); -if (x_225 == 0) -{ -return x_213; -} -else -{ -lean_object* x_226; lean_object* x_227; lean_object* x_228; -x_226 = lean_ctor_get(x_213, 0); -x_227 = lean_ctor_get(x_213, 1); -lean_inc(x_227); -lean_inc(x_226); -lean_dec(x_213); -x_228 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_228, 0, x_226); -lean_ctor_set(x_228, 1, x_227); -return x_228; -} -} +return x_107; } else { -lean_object* x_229; lean_object* x_230; -x_229 = lean_array_fget(x_3, x_209); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_5); -lean_inc(x_229); -x_230 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_229, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_230) == 0) -{ -lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; size_t x_235; size_t x_236; uint8_t 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); -lean_dec(x_230); -x_233 = lean_ctor_get(x_231, 0); -lean_inc(x_233); -x_234 = lean_ctor_get(x_231, 1); -lean_inc(x_234); -lean_dec(x_231); -x_235 = lean_ptr_addr(x_229); -lean_dec(x_229); -x_236 = lean_ptr_addr(x_233); -x_237 = lean_usize_dec_eq(x_235, x_236); -if (x_237 == 0) -{ -lean_object* x_238; -lean_dec(x_233); +lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_dec(x_3); lean_dec(x_2); +x_108 = lean_ctor_get(x_94, 0); +lean_inc(x_108); +x_109 = lean_ctor_get(x_94, 1); +lean_inc(x_109); +lean_dec(x_94); lean_inc(x_1); -x_238 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_234, x_7, x_8, x_9, x_10, x_232); +x_110 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_95, x_93, x_1); +x_111 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_111, 0, x_108); +lean_ctor_set(x_111, 1, x_109); +lean_ctor_set(x_111, 2, x_110); +lean_inc(x_1); +x_112 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_111, x_7, x_8, x_9, x_10, x_92); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_5); -return x_238; +return x_112; +} } else { -lean_object* x_239; lean_object* x_240; lean_object* x_241; -x_239 = lean_array_set(x_3, x_209, x_233); -x_240 = l_Lean_mkAppN(x_2, x_239); -lean_dec(x_239); -x_241 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_240, x_5, x_234, x_7, x_8, x_9, x_10, x_232); +lean_object* x_113; lean_object* x_114; +lean_dec(x_95); +lean_dec(x_93); +lean_dec(x_89); +lean_dec(x_3); +lean_dec(x_2); +x_113 = lean_ctor_get(x_96, 0); +lean_inc(x_113); +lean_dec(x_96); +x_114 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_113, x_5, x_94, x_7, x_8, x_9, x_10, x_92); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_5); -return x_241; +return x_114; } } else { -uint8_t x_242; -lean_dec(x_229); +uint8_t x_115; +lean_dec(x_89); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -8152,127 +7356,125 @@ lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_242 = !lean_is_exclusive(x_230); -if (x_242 == 0) +x_115 = !lean_is_exclusive(x_90); +if (x_115 == 0) { -return x_230; +return x_90; } else { -lean_object* x_243; lean_object* x_244; lean_object* x_245; -x_243 = lean_ctor_get(x_230, 0); -x_244 = lean_ctor_get(x_230, 1); -lean_inc(x_244); -lean_inc(x_243); -lean_dec(x_230); -x_245 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_245, 0, x_243); -lean_ctor_set(x_245, 1, x_244); -return x_245; -} -} +lean_object* x_116; lean_object* x_117; lean_object* x_118; +x_116 = lean_ctor_get(x_90, 0); +x_117 = lean_ctor_get(x_90, 1); +lean_inc(x_117); +lean_inc(x_116); +lean_dec(x_90); +x_118 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_118, 0, x_116); +lean_ctor_set(x_118, 1, x_117); +return x_118; } } } -block_201: +block_153: { -lean_object* x_169; -lean_dec(x_168); +lean_object* x_121; +lean_dec(x_120); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_2); -x_169 = l_Lean_Meta_getFunInfo(x_2, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_169) == 0) +x_121 = l_Lean_Meta_getFunInfo(x_2, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_121) == 0) { -lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; uint8_t x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; -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); -lean_dec(x_170); -x_173 = lean_array_get_size(x_3); -x_174 = lean_unsigned_to_nat(0u); -x_175 = lean_unsigned_to_nat(1u); -x_176 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_176, 0, x_174); -lean_ctor_set(x_176, 1, x_173); -lean_ctor_set(x_176, 2, x_175); -x_177 = 0; -x_178 = lean_box(x_177); -x_179 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_179, 0, x_3); -lean_ctor_set(x_179, 1, x_178); +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_object* x_130; lean_object* x_131; lean_object* x_132; +x_122 = lean_ctor_get(x_121, 0); +lean_inc(x_122); +x_123 = lean_ctor_get(x_121, 1); +lean_inc(x_123); +lean_dec(x_121); +x_124 = lean_ctor_get(x_122, 0); +lean_inc(x_124); +lean_dec(x_122); +x_125 = lean_array_get_size(x_3); +x_126 = lean_unsigned_to_nat(0u); +x_127 = lean_unsigned_to_nat(1u); +x_128 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_128, 0, x_126); +lean_ctor_set(x_128, 1, x_125); +lean_ctor_set(x_128, 2, x_127); +x_129 = 0; +x_130 = lean_box(x_129); +x_131 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_131, 0, x_3); +lean_ctor_set(x_131, 1, x_130); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_5); lean_inc(x_2); -x_180 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8(x_2, x_172, x_176, x_176, x_179, x_174, lean_box(0), lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10, x_171); -lean_dec(x_176); -lean_dec(x_172); -if (lean_obj_tag(x_180) == 0) +x_132 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8(x_2, x_124, x_128, x_128, x_131, x_126, lean_box(0), lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10, x_123); +lean_dec(x_128); +lean_dec(x_124); +if (lean_obj_tag(x_132) == 0) { -lean_object* x_181; lean_object* x_182; lean_object* x_183; uint8_t x_184; -x_181 = lean_ctor_get(x_180, 0); -lean_inc(x_181); -x_182 = lean_ctor_get(x_181, 0); -lean_inc(x_182); -x_183 = lean_ctor_get(x_182, 1); -lean_inc(x_183); -x_184 = lean_unbox(x_183); -lean_dec(x_183); -if (x_184 == 0) +lean_object* x_133; lean_object* x_134; lean_object* x_135; uint8_t x_136; +x_133 = lean_ctor_get(x_132, 0); +lean_inc(x_133); +x_134 = lean_ctor_get(x_133, 0); +lean_inc(x_134); +x_135 = lean_ctor_get(x_134, 1); +lean_inc(x_135); +x_136 = lean_unbox(x_135); +lean_dec(x_135); +if (x_136 == 0) { -lean_object* x_185; lean_object* x_186; lean_object* x_187; -lean_dec(x_182); +lean_object* x_137; lean_object* x_138; lean_object* x_139; +lean_dec(x_134); lean_dec(x_2); -x_185 = lean_ctor_get(x_180, 1); -lean_inc(x_185); -lean_dec(x_180); -x_186 = lean_ctor_get(x_181, 1); -lean_inc(x_186); -lean_dec(x_181); +x_137 = lean_ctor_get(x_132, 1); +lean_inc(x_137); +lean_dec(x_132); +x_138 = lean_ctor_get(x_133, 1); +lean_inc(x_138); +lean_dec(x_133); lean_inc(x_1); -x_187 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_186, x_7, x_8, x_9, x_10, x_185); +x_139 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_138, x_7, x_8, x_9, x_10, x_137); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_5); -return x_187; +return x_139; } else { -lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; -x_188 = lean_ctor_get(x_180, 1); -lean_inc(x_188); -lean_dec(x_180); -x_189 = lean_ctor_get(x_181, 1); -lean_inc(x_189); -lean_dec(x_181); -x_190 = lean_ctor_get(x_182, 0); -lean_inc(x_190); -lean_dec(x_182); -x_191 = l_Lean_mkAppN(x_2, x_190); -lean_dec(x_190); -x_192 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_191, x_5, x_189, x_7, x_8, x_9, x_10, x_188); +lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; +x_140 = lean_ctor_get(x_132, 1); +lean_inc(x_140); +lean_dec(x_132); +x_141 = lean_ctor_get(x_133, 1); +lean_inc(x_141); +lean_dec(x_133); +x_142 = lean_ctor_get(x_134, 0); +lean_inc(x_142); +lean_dec(x_134); +x_143 = l_Lean_mkAppN(x_2, x_142); +lean_dec(x_142); +x_144 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_143, x_5, x_141, x_7, x_8, x_9, x_10, x_140); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_5); -return x_192; +return x_144; } } else { -uint8_t x_193; +uint8_t x_145; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -8280,29 +7482,29 @@ lean_dec(x_7); lean_dec(x_5); lean_dec(x_2); lean_dec(x_1); -x_193 = !lean_is_exclusive(x_180); -if (x_193 == 0) +x_145 = !lean_is_exclusive(x_132); +if (x_145 == 0) { -return x_180; +return x_132; } else { -lean_object* x_194; lean_object* x_195; lean_object* x_196; -x_194 = lean_ctor_get(x_180, 0); -x_195 = lean_ctor_get(x_180, 1); -lean_inc(x_195); -lean_inc(x_194); -lean_dec(x_180); -x_196 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_196, 0, x_194); -lean_ctor_set(x_196, 1, x_195); -return x_196; +lean_object* x_146; lean_object* x_147; lean_object* x_148; +x_146 = lean_ctor_get(x_132, 0); +x_147 = lean_ctor_get(x_132, 1); +lean_inc(x_147); +lean_inc(x_146); +lean_dec(x_132); +x_148 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_148, 0, x_146); +lean_ctor_set(x_148, 1, x_147); +return x_148; } } } else { -uint8_t x_197; +uint8_t x_149; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -8312,213 +7514,188 @@ lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_197 = !lean_is_exclusive(x_169); -if (x_197 == 0) +x_149 = !lean_is_exclusive(x_121); +if (x_149 == 0) { -return x_169; +return x_121; } else { -lean_object* x_198; lean_object* x_199; lean_object* x_200; -x_198 = lean_ctor_get(x_169, 0); -x_199 = lean_ctor_get(x_169, 1); -lean_inc(x_199); -lean_inc(x_198); -lean_dec(x_169); -x_200 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_200, 0, x_198); -lean_ctor_set(x_200, 1, x_199); -return x_200; +lean_object* x_150; lean_object* x_151; lean_object* x_152; +x_150 = lean_ctor_get(x_121, 0); +x_151 = lean_ctor_get(x_121, 1); +lean_inc(x_151); +lean_inc(x_150); +lean_dec(x_121); +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; } } } } -case 3: +case 2: { -lean_object* x_246; lean_object* x_280; uint8_t x_281; +lean_object* x_166; lean_object* x_197; lean_object* x_231; uint8_t x_232; lean_dec(x_4); -x_280 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__4; -x_281 = l_Lean_Expr_isConstOf(x_2, x_280); -if (x_281 == 0) +x_231 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__4; +x_232 = l_Lean_Expr_isConstOf(x_2, x_231); +if (x_232 == 0) { -lean_object* x_282; -x_282 = lean_box(0); -x_246 = x_282; -goto block_279; +lean_object* x_233; +x_233 = lean_box(0); +x_197 = x_233; +goto block_230; } else { -lean_object* x_283; lean_object* x_284; uint8_t x_285; -x_283 = lean_array_get_size(x_3); -x_284 = lean_unsigned_to_nat(2u); -x_285 = lean_nat_dec_eq(x_283, x_284); -if (x_285 == 0) +lean_object* x_234; lean_object* x_235; uint8_t x_236; +x_234 = lean_array_get_size(x_3); +x_235 = lean_unsigned_to_nat(2u); +x_236 = lean_nat_dec_eq(x_234, x_235); +if (x_236 == 0) { -lean_object* x_286; -lean_dec(x_283); -x_286 = lean_box(0); -x_246 = x_286; -goto block_279; +lean_object* x_237; +lean_dec(x_234); +x_237 = lean_box(0); +x_197 = x_237; +goto block_230; } else { -lean_object* x_287; uint8_t x_288; -x_287 = lean_unsigned_to_nat(0u); -x_288 = lean_nat_dec_lt(x_287, x_283); -lean_dec(x_283); -if (x_288 == 0) -{ -lean_object* x_289; lean_object* x_290; lean_object* x_291; -x_289 = l_Lean_instInhabitedExpr; -x_290 = l_outOfBounds___rarg(x_289); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_5); -lean_inc(x_290); -x_291 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_290, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_291) == 0) +lean_object* x_238; uint8_t x_239; +x_238 = lean_unsigned_to_nat(0u); +x_239 = lean_nat_dec_lt(x_238, x_234); +lean_dec(x_234); +if (x_239 == 0) { -lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; size_t x_296; size_t x_297; uint8_t x_298; -x_292 = lean_ctor_get(x_291, 0); -lean_inc(x_292); -x_293 = lean_ctor_get(x_291, 1); -lean_inc(x_293); -lean_dec(x_291); -x_294 = lean_ctor_get(x_292, 0); -lean_inc(x_294); -x_295 = lean_ctor_get(x_292, 1); -lean_inc(x_295); -lean_dec(x_292); -x_296 = lean_ptr_addr(x_290); -lean_dec(x_290); -x_297 = lean_ptr_addr(x_294); -x_298 = lean_usize_dec_eq(x_296, x_297); -if (x_298 == 0) -{ -lean_object* x_299; -lean_dec(x_294); -lean_dec(x_3); -lean_dec(x_2); -lean_inc(x_1); -x_299 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_295, x_7, x_8, x_9, x_10, x_293); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_299; +lean_object* x_240; lean_object* x_241; +x_240 = l_Lean_instInhabitedExpr; +x_241 = l_outOfBounds___rarg(x_240); +x_166 = x_241; +goto block_196; } else { -lean_object* x_300; lean_object* x_301; lean_object* x_302; -x_300 = lean_array_set(x_3, x_287, x_294); -x_301 = l_Lean_mkAppN(x_2, x_300); -lean_dec(x_300); -x_302 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_301, x_5, x_295, x_7, x_8, x_9, x_10, x_293); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_302; +lean_object* x_242; +x_242 = lean_array_fget(x_3, x_238); +x_166 = x_242; +goto block_196; } } -else -{ -uint8_t x_303; -lean_dec(x_290); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_303 = !lean_is_exclusive(x_291); -if (x_303 == 0) -{ -return x_291; } -else +block_196: { -lean_object* x_304; lean_object* x_305; lean_object* x_306; -x_304 = lean_ctor_get(x_291, 0); -x_305 = lean_ctor_get(x_291, 1); -lean_inc(x_305); -lean_inc(x_304); -lean_dec(x_291); -x_306 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_306, 0, x_304); -lean_ctor_set(x_306, 1, x_305); -return x_306; -} -} +lean_object* x_167; +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_5); +lean_inc(x_166); +x_167 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_166, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_167) == 0) +{ +lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; +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_168, 0); +lean_inc(x_170); +x_171 = lean_ctor_get(x_168, 1); +lean_inc(x_171); +lean_dec(x_168); +x_172 = lean_ctor_get(x_171, 2); +lean_inc(x_172); +lean_inc(x_172); +x_173 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_172, x_170); +if (lean_obj_tag(x_173) == 0) +{ +size_t x_174; size_t x_175; uint8_t x_176; +x_174 = lean_ptr_addr(x_166); +lean_dec(x_166); +x_175 = lean_ptr_addr(x_170); +x_176 = lean_usize_dec_eq(x_174, x_175); +if (x_176 == 0) +{ +lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; +x_177 = lean_unsigned_to_nat(0u); +lean_inc(x_170); +x_178 = lean_array_set(x_3, x_177, x_170); +x_179 = l_Lean_mkAppN(x_2, x_178); +lean_dec(x_178); +x_180 = lean_ctor_get(x_171, 0); +lean_inc(x_180); +x_181 = lean_ctor_get(x_171, 1); +lean_inc(x_181); +lean_dec(x_171); +lean_inc(x_179); +x_182 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_172, x_170, x_179); +x_183 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_183, 0, x_180); +lean_ctor_set(x_183, 1, x_181); +lean_ctor_set(x_183, 2, x_182); +x_184 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_179, x_5, x_183, x_7, x_8, x_9, x_10, x_169); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_5); +return x_184; } else { -lean_object* x_307; lean_object* x_308; -x_307 = lean_array_fget(x_3, x_287); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_5); -lean_inc(x_307); -x_308 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_307, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_308) == 0) -{ -lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; size_t x_313; size_t x_314; uint8_t x_315; -x_309 = lean_ctor_get(x_308, 0); -lean_inc(x_309); -x_310 = lean_ctor_get(x_308, 1); -lean_inc(x_310); -lean_dec(x_308); -x_311 = lean_ctor_get(x_309, 0); -lean_inc(x_311); -x_312 = lean_ctor_get(x_309, 1); -lean_inc(x_312); -lean_dec(x_309); -x_313 = lean_ptr_addr(x_307); -lean_dec(x_307); -x_314 = lean_ptr_addr(x_311); -x_315 = lean_usize_dec_eq(x_313, x_314); -if (x_315 == 0) -{ -lean_object* x_316; -lean_dec(x_311); +lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_dec(x_3); lean_dec(x_2); +x_185 = lean_ctor_get(x_171, 0); +lean_inc(x_185); +x_186 = lean_ctor_get(x_171, 1); +lean_inc(x_186); +lean_dec(x_171); +lean_inc(x_1); +x_187 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_172, x_170, x_1); +x_188 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_188, 0, x_185); +lean_ctor_set(x_188, 1, x_186); +lean_ctor_set(x_188, 2, x_187); lean_inc(x_1); -x_316 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_312, x_7, x_8, x_9, x_10, x_310); +x_189 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_188, x_7, x_8, x_9, x_10, x_169); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_5); -return x_316; +return x_189; +} } else { -lean_object* x_317; lean_object* x_318; lean_object* x_319; -x_317 = lean_array_set(x_3, x_287, x_311); -x_318 = l_Lean_mkAppN(x_2, x_317); -lean_dec(x_317); -x_319 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_318, x_5, x_312, x_7, x_8, x_9, x_10, x_310); +lean_object* x_190; lean_object* x_191; +lean_dec(x_172); +lean_dec(x_170); +lean_dec(x_166); +lean_dec(x_3); +lean_dec(x_2); +x_190 = lean_ctor_get(x_173, 0); +lean_inc(x_190); +lean_dec(x_173); +x_191 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_190, x_5, x_171, x_7, x_8, x_9, x_10, x_169); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_5); -return x_319; +return x_191; } } else { -uint8_t x_320; -lean_dec(x_307); +uint8_t x_192; +lean_dec(x_166); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -8527,127 +7704,125 @@ lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_320 = !lean_is_exclusive(x_308); -if (x_320 == 0) +x_192 = !lean_is_exclusive(x_167); +if (x_192 == 0) { -return x_308; +return x_167; } else { -lean_object* x_321; lean_object* x_322; lean_object* x_323; -x_321 = lean_ctor_get(x_308, 0); -x_322 = lean_ctor_get(x_308, 1); -lean_inc(x_322); -lean_inc(x_321); -lean_dec(x_308); -x_323 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_323, 0, x_321); -lean_ctor_set(x_323, 1, x_322); -return x_323; -} -} +lean_object* x_193; lean_object* x_194; lean_object* x_195; +x_193 = lean_ctor_get(x_167, 0); +x_194 = lean_ctor_get(x_167, 1); +lean_inc(x_194); +lean_inc(x_193); +lean_dec(x_167); +x_195 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_195, 0, x_193); +lean_ctor_set(x_195, 1, x_194); +return x_195; } } } -block_279: +block_230: { -lean_object* x_247; -lean_dec(x_246); +lean_object* x_198; +lean_dec(x_197); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_2); -x_247 = l_Lean_Meta_getFunInfo(x_2, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_247) == 0) +x_198 = l_Lean_Meta_getFunInfo(x_2, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_198) == 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; uint8_t x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; -x_248 = lean_ctor_get(x_247, 0); -lean_inc(x_248); -x_249 = lean_ctor_get(x_247, 1); -lean_inc(x_249); -lean_dec(x_247); -x_250 = lean_ctor_get(x_248, 0); -lean_inc(x_250); -lean_dec(x_248); -x_251 = lean_array_get_size(x_3); -x_252 = lean_unsigned_to_nat(0u); -x_253 = lean_unsigned_to_nat(1u); -x_254 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_254, 0, x_252); -lean_ctor_set(x_254, 1, x_251); -lean_ctor_set(x_254, 2, x_253); -x_255 = 0; -x_256 = lean_box(x_255); -x_257 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_257, 0, x_3); -lean_ctor_set(x_257, 1, x_256); +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; lean_object* x_207; lean_object* x_208; lean_object* x_209; +x_199 = lean_ctor_get(x_198, 0); +lean_inc(x_199); +x_200 = lean_ctor_get(x_198, 1); +lean_inc(x_200); +lean_dec(x_198); +x_201 = lean_ctor_get(x_199, 0); +lean_inc(x_201); +lean_dec(x_199); +x_202 = lean_array_get_size(x_3); +x_203 = lean_unsigned_to_nat(0u); +x_204 = lean_unsigned_to_nat(1u); +x_205 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_205, 0, x_203); +lean_ctor_set(x_205, 1, x_202); +lean_ctor_set(x_205, 2, x_204); +x_206 = 0; +x_207 = lean_box(x_206); +x_208 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_208, 0, x_3); +lean_ctor_set(x_208, 1, x_207); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_5); lean_inc(x_2); -x_258 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8(x_2, x_250, x_254, x_254, x_257, x_252, lean_box(0), lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10, x_249); -lean_dec(x_254); -lean_dec(x_250); -if (lean_obj_tag(x_258) == 0) -{ -lean_object* x_259; lean_object* x_260; lean_object* x_261; uint8_t x_262; -x_259 = lean_ctor_get(x_258, 0); -lean_inc(x_259); -x_260 = lean_ctor_get(x_259, 0); -lean_inc(x_260); -x_261 = lean_ctor_get(x_260, 1); -lean_inc(x_261); -x_262 = lean_unbox(x_261); -lean_dec(x_261); -if (x_262 == 0) -{ -lean_object* x_263; lean_object* x_264; lean_object* x_265; -lean_dec(x_260); +x_209 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8(x_2, x_201, x_205, x_205, x_208, x_203, lean_box(0), lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10, x_200); +lean_dec(x_205); +lean_dec(x_201); +if (lean_obj_tag(x_209) == 0) +{ +lean_object* x_210; lean_object* x_211; lean_object* x_212; uint8_t x_213; +x_210 = lean_ctor_get(x_209, 0); +lean_inc(x_210); +x_211 = lean_ctor_get(x_210, 0); +lean_inc(x_211); +x_212 = lean_ctor_get(x_211, 1); +lean_inc(x_212); +x_213 = lean_unbox(x_212); +lean_dec(x_212); +if (x_213 == 0) +{ +lean_object* x_214; lean_object* x_215; lean_object* x_216; +lean_dec(x_211); lean_dec(x_2); -x_263 = lean_ctor_get(x_258, 1); -lean_inc(x_263); -lean_dec(x_258); -x_264 = lean_ctor_get(x_259, 1); -lean_inc(x_264); -lean_dec(x_259); +x_214 = lean_ctor_get(x_209, 1); +lean_inc(x_214); +lean_dec(x_209); +x_215 = lean_ctor_get(x_210, 1); +lean_inc(x_215); +lean_dec(x_210); lean_inc(x_1); -x_265 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_264, x_7, x_8, x_9, x_10, x_263); +x_216 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_215, x_7, x_8, x_9, x_10, x_214); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_5); -return x_265; +return x_216; } else { -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_258, 1); -lean_inc(x_266); -lean_dec(x_258); -x_267 = lean_ctor_get(x_259, 1); -lean_inc(x_267); -lean_dec(x_259); -x_268 = lean_ctor_get(x_260, 0); -lean_inc(x_268); -lean_dec(x_260); -x_269 = l_Lean_mkAppN(x_2, x_268); -lean_dec(x_268); -x_270 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_269, x_5, x_267, x_7, x_8, x_9, x_10, x_266); +lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; +x_217 = lean_ctor_get(x_209, 1); +lean_inc(x_217); +lean_dec(x_209); +x_218 = lean_ctor_get(x_210, 1); +lean_inc(x_218); +lean_dec(x_210); +x_219 = lean_ctor_get(x_211, 0); +lean_inc(x_219); +lean_dec(x_211); +x_220 = l_Lean_mkAppN(x_2, x_219); +lean_dec(x_219); +x_221 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_220, x_5, x_218, x_7, x_8, x_9, x_10, x_217); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_5); -return x_270; +return x_221; } } else { -uint8_t x_271; +uint8_t x_222; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -8655,29 +7830,29 @@ lean_dec(x_7); lean_dec(x_5); lean_dec(x_2); lean_dec(x_1); -x_271 = !lean_is_exclusive(x_258); -if (x_271 == 0) +x_222 = !lean_is_exclusive(x_209); +if (x_222 == 0) { -return x_258; +return x_209; } else { -lean_object* x_272; lean_object* x_273; lean_object* x_274; -x_272 = lean_ctor_get(x_258, 0); -x_273 = lean_ctor_get(x_258, 1); -lean_inc(x_273); -lean_inc(x_272); -lean_dec(x_258); -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; +lean_object* x_223; lean_object* x_224; lean_object* x_225; +x_223 = lean_ctor_get(x_209, 0); +x_224 = lean_ctor_get(x_209, 1); +lean_inc(x_224); +lean_inc(x_223); +lean_dec(x_209); +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; } } } else { -uint8_t x_275; +uint8_t x_226; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -8687,123 +7862,188 @@ lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_275 = !lean_is_exclusive(x_247); -if (x_275 == 0) +x_226 = !lean_is_exclusive(x_198); +if (x_226 == 0) { -return x_247; +return x_198; } else { -lean_object* x_276; lean_object* x_277; lean_object* x_278; -x_276 = lean_ctor_get(x_247, 0); -x_277 = lean_ctor_get(x_247, 1); -lean_inc(x_277); -lean_inc(x_276); -lean_dec(x_247); -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; +lean_object* x_227; lean_object* x_228; lean_object* x_229; +x_227 = lean_ctor_get(x_198, 0); +x_228 = lean_ctor_get(x_198, 1); +lean_inc(x_228); +lean_inc(x_227); +lean_dec(x_198); +x_229 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_229, 0, x_227); +lean_ctor_set(x_229, 1, x_228); +return x_229; } } } } -case 4: +case 3: { -lean_object* x_324; lean_object* x_358; uint8_t x_359; +lean_object* x_243; lean_object* x_274; lean_object* x_308; uint8_t x_309; lean_dec(x_4); -x_358 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__4; -x_359 = l_Lean_Expr_isConstOf(x_2, x_358); -if (x_359 == 0) +x_308 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__4; +x_309 = l_Lean_Expr_isConstOf(x_2, x_308); +if (x_309 == 0) +{ +lean_object* x_310; +x_310 = lean_box(0); +x_274 = x_310; +goto block_307; +} +else +{ +lean_object* x_311; lean_object* x_312; uint8_t x_313; +x_311 = lean_array_get_size(x_3); +x_312 = lean_unsigned_to_nat(2u); +x_313 = lean_nat_dec_eq(x_311, x_312); +if (x_313 == 0) { -lean_object* x_360; -x_360 = lean_box(0); -x_324 = x_360; -goto block_357; +lean_object* x_314; +lean_dec(x_311); +x_314 = lean_box(0); +x_274 = x_314; +goto block_307; } else { -lean_object* x_361; lean_object* x_362; uint8_t x_363; -x_361 = lean_array_get_size(x_3); -x_362 = lean_unsigned_to_nat(2u); -x_363 = lean_nat_dec_eq(x_361, x_362); -if (x_363 == 0) +lean_object* x_315; uint8_t x_316; +x_315 = lean_unsigned_to_nat(0u); +x_316 = lean_nat_dec_lt(x_315, x_311); +lean_dec(x_311); +if (x_316 == 0) { -lean_object* x_364; -lean_dec(x_361); -x_364 = lean_box(0); -x_324 = x_364; -goto block_357; +lean_object* x_317; lean_object* x_318; +x_317 = l_Lean_instInhabitedExpr; +x_318 = l_outOfBounds___rarg(x_317); +x_243 = x_318; +goto block_273; } else { -lean_object* x_365; uint8_t x_366; -x_365 = lean_unsigned_to_nat(0u); -x_366 = lean_nat_dec_lt(x_365, x_361); -lean_dec(x_361); -if (x_366 == 0) +lean_object* x_319; +x_319 = lean_array_fget(x_3, x_315); +x_243 = x_319; +goto block_273; +} +} +} +block_273: { -lean_object* x_367; lean_object* x_368; lean_object* x_369; -x_367 = l_Lean_instInhabitedExpr; -x_368 = l_outOfBounds___rarg(x_367); +lean_object* x_244; lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_5); -lean_inc(x_368); -x_369 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_368, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_369) == 0) -{ -lean_object* x_370; lean_object* x_371; lean_object* x_372; lean_object* x_373; size_t x_374; size_t x_375; uint8_t x_376; -x_370 = lean_ctor_get(x_369, 0); -lean_inc(x_370); -x_371 = lean_ctor_get(x_369, 1); -lean_inc(x_371); -lean_dec(x_369); -x_372 = lean_ctor_get(x_370, 0); -lean_inc(x_372); -x_373 = lean_ctor_get(x_370, 1); -lean_inc(x_373); -lean_dec(x_370); -x_374 = lean_ptr_addr(x_368); -lean_dec(x_368); -x_375 = lean_ptr_addr(x_372); -x_376 = lean_usize_dec_eq(x_374, x_375); -if (x_376 == 0) +lean_inc(x_243); +x_244 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_243, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_244) == 0) +{ +lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; +x_245 = lean_ctor_get(x_244, 0); +lean_inc(x_245); +x_246 = lean_ctor_get(x_244, 1); +lean_inc(x_246); +lean_dec(x_244); +x_247 = lean_ctor_get(x_245, 0); +lean_inc(x_247); +x_248 = lean_ctor_get(x_245, 1); +lean_inc(x_248); +lean_dec(x_245); +x_249 = lean_ctor_get(x_248, 2); +lean_inc(x_249); +lean_inc(x_249); +x_250 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_249, x_247); +if (lean_obj_tag(x_250) == 0) +{ +size_t x_251; size_t x_252; uint8_t x_253; +x_251 = lean_ptr_addr(x_243); +lean_dec(x_243); +x_252 = lean_ptr_addr(x_247); +x_253 = lean_usize_dec_eq(x_251, x_252); +if (x_253 == 0) +{ +lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; +x_254 = lean_unsigned_to_nat(0u); +lean_inc(x_247); +x_255 = lean_array_set(x_3, x_254, x_247); +x_256 = l_Lean_mkAppN(x_2, x_255); +lean_dec(x_255); +x_257 = lean_ctor_get(x_248, 0); +lean_inc(x_257); +x_258 = lean_ctor_get(x_248, 1); +lean_inc(x_258); +lean_dec(x_248); +lean_inc(x_256); +x_259 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_249, x_247, x_256); +x_260 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_260, 0, x_257); +lean_ctor_set(x_260, 1, x_258); +lean_ctor_set(x_260, 2, x_259); +x_261 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_256, x_5, x_260, x_7, x_8, x_9, x_10, x_246); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_5); +return x_261; +} +else { -lean_object* x_377; -lean_dec(x_372); +lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_dec(x_3); lean_dec(x_2); +x_262 = lean_ctor_get(x_248, 0); +lean_inc(x_262); +x_263 = lean_ctor_get(x_248, 1); +lean_inc(x_263); +lean_dec(x_248); +lean_inc(x_1); +x_264 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_249, x_247, x_1); +x_265 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_265, 0, x_262); +lean_ctor_set(x_265, 1, x_263); +lean_ctor_set(x_265, 2, x_264); lean_inc(x_1); -x_377 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_373, x_7, x_8, x_9, x_10, x_371); +x_266 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_265, x_7, x_8, x_9, x_10, x_246); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_5); -return x_377; +return x_266; +} } else { -lean_object* x_378; lean_object* x_379; lean_object* x_380; -x_378 = lean_array_set(x_3, x_365, x_372); -x_379 = l_Lean_mkAppN(x_2, x_378); -lean_dec(x_378); -x_380 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_379, x_5, x_373, x_7, x_8, x_9, x_10, x_371); +lean_object* x_267; lean_object* x_268; +lean_dec(x_249); +lean_dec(x_247); +lean_dec(x_243); +lean_dec(x_3); +lean_dec(x_2); +x_267 = lean_ctor_get(x_250, 0); +lean_inc(x_267); +lean_dec(x_250); +x_268 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_267, x_5, x_248, x_7, x_8, x_9, x_10, x_246); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_5); -return x_380; +return x_268; } } else { -uint8_t x_381; -lean_dec(x_368); +uint8_t x_269; +lean_dec(x_243); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -8812,480 +8052,711 @@ lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_381 = !lean_is_exclusive(x_369); -if (x_381 == 0) +x_269 = !lean_is_exclusive(x_244); +if (x_269 == 0) { -return x_369; +return x_244; } else { -lean_object* x_382; lean_object* x_383; lean_object* x_384; -x_382 = lean_ctor_get(x_369, 0); -x_383 = lean_ctor_get(x_369, 1); -lean_inc(x_383); -lean_inc(x_382); -lean_dec(x_369); -x_384 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_384, 0, x_382); -lean_ctor_set(x_384, 1, x_383); -return x_384; +lean_object* x_270; lean_object* x_271; lean_object* x_272; +x_270 = lean_ctor_get(x_244, 0); +x_271 = lean_ctor_get(x_244, 1); +lean_inc(x_271); +lean_inc(x_270); +lean_dec(x_244); +x_272 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_272, 0, x_270); +lean_ctor_set(x_272, 1, x_271); +return x_272; } } } -else +block_307: { -lean_object* x_385; lean_object* x_386; -x_385 = lean_array_fget(x_3, x_365); +lean_object* x_275; +lean_dec(x_274); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); -lean_inc(x_5); -lean_inc(x_385); -x_386 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_385, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_386) == 0) -{ -lean_object* x_387; lean_object* x_388; lean_object* x_389; lean_object* x_390; size_t x_391; size_t x_392; uint8_t x_393; -x_387 = lean_ctor_get(x_386, 0); -lean_inc(x_387); -x_388 = lean_ctor_get(x_386, 1); -lean_inc(x_388); -lean_dec(x_386); -x_389 = lean_ctor_get(x_387, 0); -lean_inc(x_389); -x_390 = lean_ctor_get(x_387, 1); -lean_inc(x_390); -lean_dec(x_387); -x_391 = lean_ptr_addr(x_385); -lean_dec(x_385); -x_392 = lean_ptr_addr(x_389); -x_393 = lean_usize_dec_eq(x_391, x_392); -if (x_393 == 0) +lean_inc(x_2); +x_275 = l_Lean_Meta_getFunInfo(x_2, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_275) == 0) { -lean_object* x_394; -lean_dec(x_389); -lean_dec(x_3); +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; uint8_t x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; +x_276 = lean_ctor_get(x_275, 0); +lean_inc(x_276); +x_277 = lean_ctor_get(x_275, 1); +lean_inc(x_277); +lean_dec(x_275); +x_278 = lean_ctor_get(x_276, 0); +lean_inc(x_278); +lean_dec(x_276); +x_279 = lean_array_get_size(x_3); +x_280 = lean_unsigned_to_nat(0u); +x_281 = lean_unsigned_to_nat(1u); +x_282 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_282, 0, x_280); +lean_ctor_set(x_282, 1, x_279); +lean_ctor_set(x_282, 2, x_281); +x_283 = 0; +x_284 = lean_box(x_283); +x_285 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_285, 0, x_3); +lean_ctor_set(x_285, 1, x_284); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_5); +lean_inc(x_2); +x_286 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8(x_2, x_278, x_282, x_282, x_285, x_280, lean_box(0), lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10, x_277); +lean_dec(x_282); +lean_dec(x_278); +if (lean_obj_tag(x_286) == 0) +{ +lean_object* x_287; lean_object* x_288; lean_object* x_289; uint8_t x_290; +x_287 = lean_ctor_get(x_286, 0); +lean_inc(x_287); +x_288 = lean_ctor_get(x_287, 0); +lean_inc(x_288); +x_289 = lean_ctor_get(x_288, 1); +lean_inc(x_289); +x_290 = lean_unbox(x_289); +lean_dec(x_289); +if (x_290 == 0) +{ +lean_object* x_291; lean_object* x_292; lean_object* x_293; +lean_dec(x_288); lean_dec(x_2); +x_291 = lean_ctor_get(x_286, 1); +lean_inc(x_291); +lean_dec(x_286); +x_292 = lean_ctor_get(x_287, 1); +lean_inc(x_292); +lean_dec(x_287); lean_inc(x_1); -x_394 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_390, x_7, x_8, x_9, x_10, x_388); +x_293 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_292, x_7, x_8, x_9, x_10, x_291); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_5); +return x_293; +} +else +{ +lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; +x_294 = lean_ctor_get(x_286, 1); +lean_inc(x_294); +lean_dec(x_286); +x_295 = lean_ctor_get(x_287, 1); +lean_inc(x_295); +lean_dec(x_287); +x_296 = lean_ctor_get(x_288, 0); +lean_inc(x_296); +lean_dec(x_288); +x_297 = l_Lean_mkAppN(x_2, x_296); +lean_dec(x_296); +x_298 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_297, x_5, x_295, x_7, x_8, x_9, x_10, x_294); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_5); -return x_394; +return x_298; +} } else { -lean_object* x_395; lean_object* x_396; lean_object* x_397; -x_395 = lean_array_set(x_3, x_365, x_389); -x_396 = l_Lean_mkAppN(x_2, x_395); -lean_dec(x_395); -x_397 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_396, x_5, x_390, x_7, x_8, x_9, x_10, x_388); +uint8_t x_299; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_5); -return x_397; +lean_dec(x_2); +lean_dec(x_1); +x_299 = !lean_is_exclusive(x_286); +if (x_299 == 0) +{ +return x_286; +} +else +{ +lean_object* x_300; lean_object* x_301; lean_object* x_302; +x_300 = lean_ctor_get(x_286, 0); +x_301 = lean_ctor_get(x_286, 1); +lean_inc(x_301); +lean_inc(x_300); +lean_dec(x_286); +x_302 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_302, 0, x_300); +lean_ctor_set(x_302, 1, x_301); +return x_302; +} } } else { -uint8_t x_398; -lean_dec(x_385); +uint8_t x_303; 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_398 = !lean_is_exclusive(x_386); -if (x_398 == 0) +x_303 = !lean_is_exclusive(x_275); +if (x_303 == 0) { -return x_386; +return x_275; } else { -lean_object* x_399; lean_object* x_400; lean_object* x_401; -x_399 = lean_ctor_get(x_386, 0); -x_400 = lean_ctor_get(x_386, 1); -lean_inc(x_400); -lean_inc(x_399); -lean_dec(x_386); -x_401 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_401, 0, x_399); -lean_ctor_set(x_401, 1, x_400); -return x_401; +lean_object* x_304; lean_object* x_305; lean_object* x_306; +x_304 = lean_ctor_get(x_275, 0); +x_305 = lean_ctor_get(x_275, 1); +lean_inc(x_305); +lean_inc(x_304); +lean_dec(x_275); +x_306 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_306, 0, x_304); +lean_ctor_set(x_306, 1, x_305); +return x_306; } } } } +case 4: +{ +lean_object* x_320; lean_object* x_351; lean_object* x_385; uint8_t x_386; +lean_dec(x_4); +x_385 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__4; +x_386 = l_Lean_Expr_isConstOf(x_2, x_385); +if (x_386 == 0) +{ +lean_object* x_387; +x_387 = lean_box(0); +x_351 = x_387; +goto block_384; } -block_357: +else { -lean_object* x_325; -lean_dec(x_324); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_2); -x_325 = l_Lean_Meta_getFunInfo(x_2, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_325) == 0) +lean_object* x_388; lean_object* x_389; uint8_t x_390; +x_388 = lean_array_get_size(x_3); +x_389 = lean_unsigned_to_nat(2u); +x_390 = lean_nat_dec_eq(x_388, x_389); +if (x_390 == 0) { -lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; uint8_t x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; -x_326 = lean_ctor_get(x_325, 0); -lean_inc(x_326); -x_327 = lean_ctor_get(x_325, 1); -lean_inc(x_327); -lean_dec(x_325); -x_328 = lean_ctor_get(x_326, 0); -lean_inc(x_328); -lean_dec(x_326); -x_329 = lean_array_get_size(x_3); -x_330 = lean_unsigned_to_nat(0u); -x_331 = lean_unsigned_to_nat(1u); -x_332 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_332, 0, x_330); -lean_ctor_set(x_332, 1, x_329); -lean_ctor_set(x_332, 2, x_331); -x_333 = 0; -x_334 = lean_box(x_333); -x_335 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_335, 0, x_3); -lean_ctor_set(x_335, 1, x_334); +lean_object* x_391; +lean_dec(x_388); +x_391 = lean_box(0); +x_351 = x_391; +goto block_384; +} +else +{ +lean_object* x_392; uint8_t x_393; +x_392 = lean_unsigned_to_nat(0u); +x_393 = lean_nat_dec_lt(x_392, x_388); +lean_dec(x_388); +if (x_393 == 0) +{ +lean_object* x_394; lean_object* x_395; +x_394 = l_Lean_instInhabitedExpr; +x_395 = l_outOfBounds___rarg(x_394); +x_320 = x_395; +goto block_350; +} +else +{ +lean_object* x_396; +x_396 = lean_array_fget(x_3, x_392); +x_320 = x_396; +goto block_350; +} +} +} +block_350: +{ +lean_object* x_321; lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_5); -lean_inc(x_2); -x_336 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8(x_2, x_328, x_332, x_332, x_335, x_330, lean_box(0), lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10, x_327); -lean_dec(x_332); -lean_dec(x_328); -if (lean_obj_tag(x_336) == 0) -{ -lean_object* x_337; lean_object* x_338; lean_object* x_339; uint8_t x_340; -x_337 = lean_ctor_get(x_336, 0); -lean_inc(x_337); -x_338 = lean_ctor_get(x_337, 0); -lean_inc(x_338); -x_339 = lean_ctor_get(x_338, 1); -lean_inc(x_339); -x_340 = lean_unbox(x_339); -lean_dec(x_339); -if (x_340 == 0) +lean_inc(x_320); +x_321 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_320, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_321) == 0) { -lean_object* x_341; lean_object* x_342; lean_object* x_343; -lean_dec(x_338); -lean_dec(x_2); -x_341 = lean_ctor_get(x_336, 1); -lean_inc(x_341); -lean_dec(x_336); -x_342 = lean_ctor_get(x_337, 1); -lean_inc(x_342); -lean_dec(x_337); -lean_inc(x_1); -x_343 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_342, x_7, x_8, x_9, x_10, x_341); +lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; +x_322 = lean_ctor_get(x_321, 0); +lean_inc(x_322); +x_323 = lean_ctor_get(x_321, 1); +lean_inc(x_323); +lean_dec(x_321); +x_324 = lean_ctor_get(x_322, 0); +lean_inc(x_324); +x_325 = lean_ctor_get(x_322, 1); +lean_inc(x_325); +lean_dec(x_322); +x_326 = lean_ctor_get(x_325, 2); +lean_inc(x_326); +lean_inc(x_326); +x_327 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_326, x_324); +if (lean_obj_tag(x_327) == 0) +{ +size_t x_328; size_t x_329; uint8_t x_330; +x_328 = lean_ptr_addr(x_320); +lean_dec(x_320); +x_329 = lean_ptr_addr(x_324); +x_330 = lean_usize_dec_eq(x_328, x_329); +if (x_330 == 0) +{ +lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; +x_331 = lean_unsigned_to_nat(0u); +lean_inc(x_324); +x_332 = lean_array_set(x_3, x_331, x_324); +x_333 = l_Lean_mkAppN(x_2, x_332); +lean_dec(x_332); +x_334 = lean_ctor_get(x_325, 0); +lean_inc(x_334); +x_335 = lean_ctor_get(x_325, 1); +lean_inc(x_335); +lean_dec(x_325); +lean_inc(x_333); +x_336 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_326, x_324, x_333); +x_337 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_337, 0, x_334); +lean_ctor_set(x_337, 1, x_335); +lean_ctor_set(x_337, 2, x_336); +x_338 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_333, x_5, x_337, x_7, x_8, x_9, x_10, x_323); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_5); -return x_343; +return x_338; } else { -lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; -x_344 = lean_ctor_get(x_336, 1); -lean_inc(x_344); -lean_dec(x_336); -x_345 = lean_ctor_get(x_337, 1); -lean_inc(x_345); -lean_dec(x_337); -x_346 = lean_ctor_get(x_338, 0); -lean_inc(x_346); -lean_dec(x_338); -x_347 = l_Lean_mkAppN(x_2, x_346); -lean_dec(x_346); -x_348 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_347, x_5, x_345, x_7, x_8, x_9, x_10, x_344); +lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; +lean_dec(x_3); +lean_dec(x_2); +x_339 = lean_ctor_get(x_325, 0); +lean_inc(x_339); +x_340 = lean_ctor_get(x_325, 1); +lean_inc(x_340); +lean_dec(x_325); +lean_inc(x_1); +x_341 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_326, x_324, x_1); +x_342 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_342, 0, x_339); +lean_ctor_set(x_342, 1, x_340); +lean_ctor_set(x_342, 2, x_341); +lean_inc(x_1); +x_343 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_342, x_7, x_8, x_9, x_10, x_323); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_5); -return x_348; +return x_343; } } else { -uint8_t x_349; +lean_object* x_344; lean_object* x_345; +lean_dec(x_326); +lean_dec(x_324); +lean_dec(x_320); +lean_dec(x_3); +lean_dec(x_2); +x_344 = lean_ctor_get(x_327, 0); +lean_inc(x_344); +lean_dec(x_327); +x_345 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_344, x_5, x_325, x_7, x_8, x_9, x_10, x_323); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_5); -lean_dec(x_2); -lean_dec(x_1); -x_349 = !lean_is_exclusive(x_336); -if (x_349 == 0) -{ -return x_336; -} -else -{ -lean_object* x_350; lean_object* x_351; lean_object* x_352; -x_350 = lean_ctor_get(x_336, 0); -x_351 = lean_ctor_get(x_336, 1); -lean_inc(x_351); -lean_inc(x_350); -lean_dec(x_336); -x_352 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_352, 0, x_350); -lean_ctor_set(x_352, 1, x_351); -return x_352; -} +return x_345; } } else { -uint8_t x_353; +uint8_t x_346; +lean_dec(x_320); 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_353 = !lean_is_exclusive(x_325); -if (x_353 == 0) +x_346 = !lean_is_exclusive(x_321); +if (x_346 == 0) { -return x_325; +return x_321; } else { -lean_object* x_354; lean_object* x_355; lean_object* x_356; -x_354 = lean_ctor_get(x_325, 0); -x_355 = lean_ctor_get(x_325, 1); -lean_inc(x_355); -lean_inc(x_354); -lean_dec(x_325); -x_356 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_356, 0, x_354); -lean_ctor_set(x_356, 1, x_355); -return x_356; -} -} -} -} -case 5: -{ -lean_object* x_402; lean_object* x_403; lean_object* x_404; lean_object* x_405; lean_object* x_406; -x_402 = lean_ctor_get(x_2, 0); -lean_inc(x_402); -x_403 = lean_ctor_get(x_2, 1); -lean_inc(x_403); -lean_dec(x_2); -x_404 = lean_array_set(x_3, x_4, x_403); -x_405 = lean_unsigned_to_nat(1u); -x_406 = lean_nat_sub(x_4, x_405); -lean_dec(x_4); -x_2 = x_402; -x_3 = x_404; -x_4 = x_406; -goto _start; +lean_object* x_347; lean_object* x_348; lean_object* x_349; +x_347 = lean_ctor_get(x_321, 0); +x_348 = lean_ctor_get(x_321, 1); +lean_inc(x_348); +lean_inc(x_347); +lean_dec(x_321); +x_349 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_349, 0, x_347); +lean_ctor_set(x_349, 1, x_348); +return x_349; } -case 6: -{ -lean_object* x_408; lean_object* x_442; uint8_t x_443; -lean_dec(x_4); -x_442 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__4; -x_443 = l_Lean_Expr_isConstOf(x_2, x_442); -if (x_443 == 0) -{ -lean_object* x_444; -x_444 = lean_box(0); -x_408 = x_444; -goto block_441; } -else -{ -lean_object* x_445; lean_object* x_446; uint8_t x_447; -x_445 = lean_array_get_size(x_3); -x_446 = lean_unsigned_to_nat(2u); -x_447 = lean_nat_dec_eq(x_445, x_446); -if (x_447 == 0) -{ -lean_object* x_448; -lean_dec(x_445); -x_448 = lean_box(0); -x_408 = x_448; -goto block_441; } -else +block_384: { -lean_object* x_449; uint8_t x_450; -x_449 = lean_unsigned_to_nat(0u); -x_450 = lean_nat_dec_lt(x_449, x_445); -lean_dec(x_445); -if (x_450 == 0) +lean_object* x_352; +lean_dec(x_351); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_2); +x_352 = l_Lean_Meta_getFunInfo(x_2, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_352) == 0) { -lean_object* x_451; lean_object* x_452; lean_object* x_453; -x_451 = l_Lean_instInhabitedExpr; -x_452 = l_outOfBounds___rarg(x_451); +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; uint8_t 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_355 = lean_ctor_get(x_353, 0); +lean_inc(x_355); +lean_dec(x_353); +x_356 = lean_array_get_size(x_3); +x_357 = lean_unsigned_to_nat(0u); +x_358 = lean_unsigned_to_nat(1u); +x_359 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_359, 0, x_357); +lean_ctor_set(x_359, 1, x_356); +lean_ctor_set(x_359, 2, x_358); +x_360 = 0; +x_361 = lean_box(x_360); +x_362 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_362, 0, x_3); +lean_ctor_set(x_362, 1, x_361); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_5); -lean_inc(x_452); -x_453 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_452, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_453) == 0) -{ -lean_object* x_454; lean_object* x_455; lean_object* x_456; lean_object* x_457; size_t x_458; size_t x_459; uint8_t x_460; -x_454 = lean_ctor_get(x_453, 0); -lean_inc(x_454); -x_455 = lean_ctor_get(x_453, 1); -lean_inc(x_455); -lean_dec(x_453); -x_456 = lean_ctor_get(x_454, 0); -lean_inc(x_456); -x_457 = lean_ctor_get(x_454, 1); -lean_inc(x_457); -lean_dec(x_454); -x_458 = lean_ptr_addr(x_452); -lean_dec(x_452); -x_459 = lean_ptr_addr(x_456); -x_460 = lean_usize_dec_eq(x_458, x_459); -if (x_460 == 0) -{ -lean_object* x_461; -lean_dec(x_456); -lean_dec(x_3); +lean_inc(x_2); +x_363 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8(x_2, x_355, x_359, x_359, x_362, x_357, lean_box(0), lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10, x_354); +lean_dec(x_359); +lean_dec(x_355); +if (lean_obj_tag(x_363) == 0) +{ +lean_object* x_364; lean_object* x_365; lean_object* x_366; uint8_t x_367; +x_364 = lean_ctor_get(x_363, 0); +lean_inc(x_364); +x_365 = lean_ctor_get(x_364, 0); +lean_inc(x_365); +x_366 = lean_ctor_get(x_365, 1); +lean_inc(x_366); +x_367 = lean_unbox(x_366); +lean_dec(x_366); +if (x_367 == 0) +{ +lean_object* x_368; lean_object* x_369; lean_object* x_370; +lean_dec(x_365); lean_dec(x_2); +x_368 = lean_ctor_get(x_363, 1); +lean_inc(x_368); +lean_dec(x_363); +x_369 = lean_ctor_get(x_364, 1); +lean_inc(x_369); +lean_dec(x_364); lean_inc(x_1); -x_461 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_457, x_7, x_8, x_9, x_10, x_455); +x_370 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_369, x_7, x_8, x_9, x_10, x_368); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_5); +return x_370; +} +else +{ +lean_object* x_371; lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; +x_371 = lean_ctor_get(x_363, 1); +lean_inc(x_371); +lean_dec(x_363); +x_372 = lean_ctor_get(x_364, 1); +lean_inc(x_372); +lean_dec(x_364); +x_373 = lean_ctor_get(x_365, 0); +lean_inc(x_373); +lean_dec(x_365); +x_374 = l_Lean_mkAppN(x_2, x_373); +lean_dec(x_373); +x_375 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_374, x_5, x_372, x_7, x_8, x_9, x_10, x_371); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_5); +return x_375; +} +} +else +{ +uint8_t x_376; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_5); -return x_461; +lean_dec(x_2); +lean_dec(x_1); +x_376 = !lean_is_exclusive(x_363); +if (x_376 == 0) +{ +return x_363; } else { -lean_object* x_462; lean_object* x_463; lean_object* x_464; -x_462 = lean_array_set(x_3, x_449, x_456); -x_463 = l_Lean_mkAppN(x_2, x_462); -lean_dec(x_462); -x_464 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_463, x_5, x_457, x_7, x_8, x_9, x_10, x_455); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_464; +lean_object* x_377; lean_object* x_378; lean_object* x_379; +x_377 = lean_ctor_get(x_363, 0); +x_378 = lean_ctor_get(x_363, 1); +lean_inc(x_378); +lean_inc(x_377); +lean_dec(x_363); +x_379 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_379, 0, x_377); +lean_ctor_set(x_379, 1, x_378); +return x_379; +} } } else { -uint8_t x_465; -lean_dec(x_452); +uint8_t x_380; 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_465 = !lean_is_exclusive(x_453); -if (x_465 == 0) +x_380 = !lean_is_exclusive(x_352); +if (x_380 == 0) { -return x_453; +return x_352; } else { -lean_object* x_466; lean_object* x_467; lean_object* x_468; -x_466 = lean_ctor_get(x_453, 0); -x_467 = lean_ctor_get(x_453, 1); -lean_inc(x_467); -lean_inc(x_466); -lean_dec(x_453); -x_468 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_468, 0, x_466); -lean_ctor_set(x_468, 1, x_467); -return x_468; +lean_object* x_381; lean_object* x_382; lean_object* x_383; +x_381 = lean_ctor_get(x_352, 0); +x_382 = lean_ctor_get(x_352, 1); +lean_inc(x_382); +lean_inc(x_381); +lean_dec(x_352); +x_383 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_383, 0, x_381); +lean_ctor_set(x_383, 1, x_382); +return x_383; +} +} +} +} +case 5: +{ +lean_object* x_397; lean_object* x_398; lean_object* x_399; lean_object* x_400; lean_object* x_401; +x_397 = lean_ctor_get(x_2, 0); +lean_inc(x_397); +x_398 = lean_ctor_get(x_2, 1); +lean_inc(x_398); +lean_dec(x_2); +x_399 = lean_array_set(x_3, x_4, x_398); +x_400 = lean_unsigned_to_nat(1u); +x_401 = lean_nat_sub(x_4, x_400); +lean_dec(x_4); +x_2 = x_397; +x_3 = x_399; +x_4 = x_401; +goto _start; +} +case 6: +{ +lean_object* x_403; lean_object* x_434; lean_object* x_468; uint8_t x_469; +lean_dec(x_4); +x_468 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__4; +x_469 = l_Lean_Expr_isConstOf(x_2, x_468); +if (x_469 == 0) +{ +lean_object* x_470; +x_470 = lean_box(0); +x_434 = x_470; +goto block_467; } +else +{ +lean_object* x_471; lean_object* x_472; uint8_t x_473; +x_471 = lean_array_get_size(x_3); +x_472 = lean_unsigned_to_nat(2u); +x_473 = lean_nat_dec_eq(x_471, x_472); +if (x_473 == 0) +{ +lean_object* x_474; +lean_dec(x_471); +x_474 = lean_box(0); +x_434 = x_474; +goto block_467; } +else +{ +lean_object* x_475; uint8_t x_476; +x_475 = lean_unsigned_to_nat(0u); +x_476 = lean_nat_dec_lt(x_475, x_471); +lean_dec(x_471); +if (x_476 == 0) +{ +lean_object* x_477; lean_object* x_478; +x_477 = l_Lean_instInhabitedExpr; +x_478 = l_outOfBounds___rarg(x_477); +x_403 = x_478; +goto block_433; } else { -lean_object* x_469; lean_object* x_470; -x_469 = lean_array_fget(x_3, x_449); +lean_object* x_479; +x_479 = lean_array_fget(x_3, x_475); +x_403 = x_479; +goto block_433; +} +} +} +block_433: +{ +lean_object* x_404; lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_5); -lean_inc(x_469); -x_470 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_469, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_470) == 0) -{ -lean_object* x_471; lean_object* x_472; lean_object* x_473; lean_object* x_474; size_t x_475; size_t x_476; uint8_t x_477; -x_471 = lean_ctor_get(x_470, 0); -lean_inc(x_471); -x_472 = lean_ctor_get(x_470, 1); -lean_inc(x_472); -lean_dec(x_470); -x_473 = lean_ctor_get(x_471, 0); -lean_inc(x_473); -x_474 = lean_ctor_get(x_471, 1); -lean_inc(x_474); -lean_dec(x_471); -x_475 = lean_ptr_addr(x_469); -lean_dec(x_469); -x_476 = lean_ptr_addr(x_473); -x_477 = lean_usize_dec_eq(x_475, x_476); -if (x_477 == 0) -{ -lean_object* x_478; -lean_dec(x_473); +lean_inc(x_403); +x_404 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_403, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_404) == 0) +{ +lean_object* x_405; lean_object* x_406; lean_object* x_407; lean_object* x_408; lean_object* x_409; lean_object* x_410; +x_405 = lean_ctor_get(x_404, 0); +lean_inc(x_405); +x_406 = lean_ctor_get(x_404, 1); +lean_inc(x_406); +lean_dec(x_404); +x_407 = lean_ctor_get(x_405, 0); +lean_inc(x_407); +x_408 = lean_ctor_get(x_405, 1); +lean_inc(x_408); +lean_dec(x_405); +x_409 = lean_ctor_get(x_408, 2); +lean_inc(x_409); +lean_inc(x_409); +x_410 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_409, x_407); +if (lean_obj_tag(x_410) == 0) +{ +size_t x_411; size_t x_412; uint8_t x_413; +x_411 = lean_ptr_addr(x_403); +lean_dec(x_403); +x_412 = lean_ptr_addr(x_407); +x_413 = lean_usize_dec_eq(x_411, x_412); +if (x_413 == 0) +{ +lean_object* x_414; lean_object* x_415; lean_object* x_416; lean_object* x_417; lean_object* x_418; lean_object* x_419; lean_object* x_420; lean_object* x_421; +x_414 = lean_unsigned_to_nat(0u); +lean_inc(x_407); +x_415 = lean_array_set(x_3, x_414, x_407); +x_416 = l_Lean_mkAppN(x_2, x_415); +lean_dec(x_415); +x_417 = lean_ctor_get(x_408, 0); +lean_inc(x_417); +x_418 = lean_ctor_get(x_408, 1); +lean_inc(x_418); +lean_dec(x_408); +lean_inc(x_416); +x_419 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_409, x_407, x_416); +x_420 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_420, 0, x_417); +lean_ctor_set(x_420, 1, x_418); +lean_ctor_set(x_420, 2, x_419); +x_421 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_416, x_5, x_420, x_7, x_8, x_9, x_10, x_406); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_5); +return x_421; +} +else +{ +lean_object* x_422; lean_object* x_423; lean_object* x_424; lean_object* x_425; lean_object* x_426; lean_dec(x_3); lean_dec(x_2); +x_422 = lean_ctor_get(x_408, 0); +lean_inc(x_422); +x_423 = lean_ctor_get(x_408, 1); +lean_inc(x_423); +lean_dec(x_408); lean_inc(x_1); -x_478 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_474, x_7, x_8, x_9, x_10, x_472); +x_424 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_409, x_407, x_1); +x_425 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_425, 0, x_422); +lean_ctor_set(x_425, 1, x_423); +lean_ctor_set(x_425, 2, x_424); +lean_inc(x_1); +x_426 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_425, x_7, x_8, x_9, x_10, x_406); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_5); -return x_478; +return x_426; +} } else { -lean_object* x_479; lean_object* x_480; lean_object* x_481; -x_479 = lean_array_set(x_3, x_449, x_473); -x_480 = l_Lean_mkAppN(x_2, x_479); -lean_dec(x_479); -x_481 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_480, x_5, x_474, x_7, x_8, x_9, x_10, x_472); +lean_object* x_427; lean_object* x_428; +lean_dec(x_409); +lean_dec(x_407); +lean_dec(x_403); +lean_dec(x_3); +lean_dec(x_2); +x_427 = lean_ctor_get(x_410, 0); +lean_inc(x_427); +lean_dec(x_410); +x_428 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_427, x_5, x_408, x_7, x_8, x_9, x_10, x_406); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_5); -return x_481; +return x_428; } } else { -uint8_t x_482; -lean_dec(x_469); +uint8_t x_429; +lean_dec(x_403); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -9294,127 +8765,125 @@ lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_482 = !lean_is_exclusive(x_470); -if (x_482 == 0) +x_429 = !lean_is_exclusive(x_404); +if (x_429 == 0) { -return x_470; +return x_404; } else { -lean_object* x_483; lean_object* x_484; lean_object* x_485; -x_483 = lean_ctor_get(x_470, 0); -x_484 = lean_ctor_get(x_470, 1); -lean_inc(x_484); -lean_inc(x_483); -lean_dec(x_470); -x_485 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_485, 0, x_483); -lean_ctor_set(x_485, 1, x_484); -return x_485; -} -} +lean_object* x_430; lean_object* x_431; lean_object* x_432; +x_430 = lean_ctor_get(x_404, 0); +x_431 = lean_ctor_get(x_404, 1); +lean_inc(x_431); +lean_inc(x_430); +lean_dec(x_404); +x_432 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_432, 0, x_430); +lean_ctor_set(x_432, 1, x_431); +return x_432; } } } -block_441: +block_467: { -lean_object* x_409; -lean_dec(x_408); +lean_object* x_435; +lean_dec(x_434); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_2); -x_409 = l_Lean_Meta_getFunInfo(x_2, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_409) == 0) -{ -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; uint8_t x_417; lean_object* x_418; lean_object* x_419; lean_object* x_420; -x_410 = lean_ctor_get(x_409, 0); -lean_inc(x_410); -x_411 = lean_ctor_get(x_409, 1); -lean_inc(x_411); -lean_dec(x_409); -x_412 = lean_ctor_get(x_410, 0); -lean_inc(x_412); -lean_dec(x_410); -x_413 = lean_array_get_size(x_3); -x_414 = lean_unsigned_to_nat(0u); -x_415 = lean_unsigned_to_nat(1u); -x_416 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_416, 0, x_414); -lean_ctor_set(x_416, 1, x_413); -lean_ctor_set(x_416, 2, x_415); -x_417 = 0; -x_418 = lean_box(x_417); -x_419 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_419, 0, x_3); -lean_ctor_set(x_419, 1, x_418); +x_435 = l_Lean_Meta_getFunInfo(x_2, x_7, x_8, x_9, x_10, x_11); +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; uint8_t x_443; lean_object* x_444; lean_object* x_445; lean_object* x_446; +x_436 = lean_ctor_get(x_435, 0); +lean_inc(x_436); +x_437 = lean_ctor_get(x_435, 1); +lean_inc(x_437); +lean_dec(x_435); +x_438 = lean_ctor_get(x_436, 0); +lean_inc(x_438); +lean_dec(x_436); +x_439 = lean_array_get_size(x_3); +x_440 = lean_unsigned_to_nat(0u); +x_441 = lean_unsigned_to_nat(1u); +x_442 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_442, 0, x_440); +lean_ctor_set(x_442, 1, x_439); +lean_ctor_set(x_442, 2, x_441); +x_443 = 0; +x_444 = lean_box(x_443); +x_445 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_445, 0, x_3); +lean_ctor_set(x_445, 1, x_444); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_5); lean_inc(x_2); -x_420 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8(x_2, x_412, x_416, x_416, x_419, x_414, lean_box(0), lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10, x_411); -lean_dec(x_416); -lean_dec(x_412); -if (lean_obj_tag(x_420) == 0) -{ -lean_object* x_421; lean_object* x_422; lean_object* x_423; uint8_t x_424; -x_421 = lean_ctor_get(x_420, 0); -lean_inc(x_421); -x_422 = lean_ctor_get(x_421, 0); -lean_inc(x_422); -x_423 = lean_ctor_get(x_422, 1); -lean_inc(x_423); -x_424 = lean_unbox(x_423); -lean_dec(x_423); -if (x_424 == 0) +x_446 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8(x_2, x_438, x_442, x_442, x_445, x_440, lean_box(0), lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10, x_437); +lean_dec(x_442); +lean_dec(x_438); +if (lean_obj_tag(x_446) == 0) +{ +lean_object* x_447; lean_object* x_448; lean_object* x_449; uint8_t x_450; +x_447 = lean_ctor_get(x_446, 0); +lean_inc(x_447); +x_448 = lean_ctor_get(x_447, 0); +lean_inc(x_448); +x_449 = lean_ctor_get(x_448, 1); +lean_inc(x_449); +x_450 = lean_unbox(x_449); +lean_dec(x_449); +if (x_450 == 0) { -lean_object* x_425; lean_object* x_426; lean_object* x_427; -lean_dec(x_422); +lean_object* x_451; lean_object* x_452; lean_object* x_453; +lean_dec(x_448); lean_dec(x_2); -x_425 = lean_ctor_get(x_420, 1); -lean_inc(x_425); -lean_dec(x_420); -x_426 = lean_ctor_get(x_421, 1); -lean_inc(x_426); -lean_dec(x_421); +x_451 = lean_ctor_get(x_446, 1); +lean_inc(x_451); +lean_dec(x_446); +x_452 = lean_ctor_get(x_447, 1); +lean_inc(x_452); +lean_dec(x_447); lean_inc(x_1); -x_427 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_426, x_7, x_8, x_9, x_10, x_425); +x_453 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_452, x_7, x_8, x_9, x_10, x_451); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_5); -return x_427; +return x_453; } else { -lean_object* x_428; lean_object* x_429; lean_object* x_430; lean_object* x_431; lean_object* x_432; -x_428 = lean_ctor_get(x_420, 1); -lean_inc(x_428); -lean_dec(x_420); -x_429 = lean_ctor_get(x_421, 1); -lean_inc(x_429); -lean_dec(x_421); -x_430 = lean_ctor_get(x_422, 0); -lean_inc(x_430); -lean_dec(x_422); -x_431 = l_Lean_mkAppN(x_2, x_430); -lean_dec(x_430); -x_432 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_431, x_5, x_429, x_7, x_8, x_9, x_10, x_428); +lean_object* x_454; lean_object* x_455; lean_object* x_456; lean_object* x_457; lean_object* x_458; +x_454 = lean_ctor_get(x_446, 1); +lean_inc(x_454); +lean_dec(x_446); +x_455 = lean_ctor_get(x_447, 1); +lean_inc(x_455); +lean_dec(x_447); +x_456 = lean_ctor_get(x_448, 0); +lean_inc(x_456); +lean_dec(x_448); +x_457 = l_Lean_mkAppN(x_2, x_456); +lean_dec(x_456); +x_458 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_457, x_5, x_455, x_7, x_8, x_9, x_10, x_454); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_5); -return x_432; +return x_458; } } else { -uint8_t x_433; +uint8_t x_459; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -9422,29 +8891,29 @@ lean_dec(x_7); lean_dec(x_5); lean_dec(x_2); lean_dec(x_1); -x_433 = !lean_is_exclusive(x_420); -if (x_433 == 0) +x_459 = !lean_is_exclusive(x_446); +if (x_459 == 0) { -return x_420; +return x_446; } else { -lean_object* x_434; lean_object* x_435; lean_object* x_436; -x_434 = lean_ctor_get(x_420, 0); -x_435 = lean_ctor_get(x_420, 1); -lean_inc(x_435); -lean_inc(x_434); -lean_dec(x_420); -x_436 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_436, 0, x_434); -lean_ctor_set(x_436, 1, x_435); -return x_436; +lean_object* x_460; lean_object* x_461; lean_object* x_462; +x_460 = lean_ctor_get(x_446, 0); +x_461 = lean_ctor_get(x_446, 1); +lean_inc(x_461); +lean_inc(x_460); +lean_dec(x_446); +x_462 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_462, 0, x_460); +lean_ctor_set(x_462, 1, x_461); +return x_462; } } } else { -uint8_t x_437; +uint8_t x_463; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -9454,213 +8923,188 @@ lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_437 = !lean_is_exclusive(x_409); -if (x_437 == 0) +x_463 = !lean_is_exclusive(x_435); +if (x_463 == 0) { -return x_409; +return x_435; } else { -lean_object* x_438; lean_object* x_439; lean_object* x_440; -x_438 = lean_ctor_get(x_409, 0); -x_439 = lean_ctor_get(x_409, 1); -lean_inc(x_439); -lean_inc(x_438); -lean_dec(x_409); -x_440 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_440, 0, x_438); -lean_ctor_set(x_440, 1, x_439); -return x_440; +lean_object* x_464; lean_object* x_465; lean_object* x_466; +x_464 = lean_ctor_get(x_435, 0); +x_465 = lean_ctor_get(x_435, 1); +lean_inc(x_465); +lean_inc(x_464); +lean_dec(x_435); +x_466 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_466, 0, x_464); +lean_ctor_set(x_466, 1, x_465); +return x_466; } } } } case 7: { -lean_object* x_486; lean_object* x_520; uint8_t x_521; +lean_object* x_480; lean_object* x_511; lean_object* x_545; uint8_t x_546; lean_dec(x_4); -x_520 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__4; -x_521 = l_Lean_Expr_isConstOf(x_2, x_520); -if (x_521 == 0) +x_545 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__4; +x_546 = l_Lean_Expr_isConstOf(x_2, x_545); +if (x_546 == 0) { -lean_object* x_522; -x_522 = lean_box(0); -x_486 = x_522; -goto block_519; +lean_object* x_547; +x_547 = lean_box(0); +x_511 = x_547; +goto block_544; } else { -lean_object* x_523; lean_object* x_524; uint8_t x_525; -x_523 = lean_array_get_size(x_3); -x_524 = lean_unsigned_to_nat(2u); -x_525 = lean_nat_dec_eq(x_523, x_524); -if (x_525 == 0) +lean_object* x_548; lean_object* x_549; uint8_t x_550; +x_548 = lean_array_get_size(x_3); +x_549 = lean_unsigned_to_nat(2u); +x_550 = lean_nat_dec_eq(x_548, x_549); +if (x_550 == 0) { -lean_object* x_526; -lean_dec(x_523); -x_526 = lean_box(0); -x_486 = x_526; -goto block_519; +lean_object* x_551; +lean_dec(x_548); +x_551 = lean_box(0); +x_511 = x_551; +goto block_544; } else { -lean_object* x_527; uint8_t x_528; -x_527 = lean_unsigned_to_nat(0u); -x_528 = lean_nat_dec_lt(x_527, x_523); -lean_dec(x_523); -if (x_528 == 0) +lean_object* x_552; uint8_t x_553; +x_552 = lean_unsigned_to_nat(0u); +x_553 = lean_nat_dec_lt(x_552, x_548); +lean_dec(x_548); +if (x_553 == 0) +{ +lean_object* x_554; lean_object* x_555; +x_554 = l_Lean_instInhabitedExpr; +x_555 = l_outOfBounds___rarg(x_554); +x_480 = x_555; +goto block_510; +} +else +{ +lean_object* x_556; +x_556 = lean_array_fget(x_3, x_552); +x_480 = x_556; +goto block_510; +} +} +} +block_510: { -lean_object* x_529; lean_object* x_530; lean_object* x_531; -x_529 = l_Lean_instInhabitedExpr; -x_530 = l_outOfBounds___rarg(x_529); +lean_object* x_481; lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_5); -lean_inc(x_530); -x_531 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_530, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_531) == 0) -{ -lean_object* x_532; lean_object* x_533; lean_object* x_534; lean_object* x_535; size_t x_536; size_t x_537; uint8_t x_538; -x_532 = lean_ctor_get(x_531, 0); -lean_inc(x_532); -x_533 = lean_ctor_get(x_531, 1); -lean_inc(x_533); -lean_dec(x_531); -x_534 = lean_ctor_get(x_532, 0); -lean_inc(x_534); -x_535 = lean_ctor_get(x_532, 1); -lean_inc(x_535); -lean_dec(x_532); -x_536 = lean_ptr_addr(x_530); -lean_dec(x_530); -x_537 = lean_ptr_addr(x_534); -x_538 = lean_usize_dec_eq(x_536, x_537); -if (x_538 == 0) -{ -lean_object* x_539; -lean_dec(x_534); -lean_dec(x_3); -lean_dec(x_2); -lean_inc(x_1); -x_539 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_535, x_7, x_8, x_9, x_10, x_533); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_539; -} -else +lean_inc(x_480); +x_481 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_480, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_481) == 0) +{ +lean_object* x_482; lean_object* x_483; lean_object* x_484; lean_object* x_485; lean_object* x_486; lean_object* x_487; +x_482 = lean_ctor_get(x_481, 0); +lean_inc(x_482); +x_483 = lean_ctor_get(x_481, 1); +lean_inc(x_483); +lean_dec(x_481); +x_484 = lean_ctor_get(x_482, 0); +lean_inc(x_484); +x_485 = lean_ctor_get(x_482, 1); +lean_inc(x_485); +lean_dec(x_482); +x_486 = lean_ctor_get(x_485, 2); +lean_inc(x_486); +lean_inc(x_486); +x_487 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_486, x_484); +if (lean_obj_tag(x_487) == 0) { -lean_object* x_540; lean_object* x_541; lean_object* x_542; -x_540 = lean_array_set(x_3, x_527, x_534); -x_541 = l_Lean_mkAppN(x_2, x_540); -lean_dec(x_540); -x_542 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_541, x_5, x_535, x_7, x_8, x_9, x_10, x_533); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_542; -} -} -else +size_t x_488; size_t x_489; uint8_t x_490; +x_488 = lean_ptr_addr(x_480); +lean_dec(x_480); +x_489 = lean_ptr_addr(x_484); +x_490 = lean_usize_dec_eq(x_488, x_489); +if (x_490 == 0) { -uint8_t x_543; -lean_dec(x_530); +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; +x_491 = lean_unsigned_to_nat(0u); +lean_inc(x_484); +x_492 = lean_array_set(x_3, x_491, x_484); +x_493 = l_Lean_mkAppN(x_2, x_492); +lean_dec(x_492); +x_494 = lean_ctor_get(x_485, 0); +lean_inc(x_494); +x_495 = lean_ctor_get(x_485, 1); +lean_inc(x_495); +lean_dec(x_485); +lean_inc(x_493); +x_496 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_486, x_484, x_493); +x_497 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_497, 0, x_494); +lean_ctor_set(x_497, 1, x_495); +lean_ctor_set(x_497, 2, x_496); +x_498 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_493, x_5, x_497, x_7, x_8, x_9, x_10, x_483); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_543 = !lean_is_exclusive(x_531); -if (x_543 == 0) -{ -return x_531; -} -else -{ -lean_object* x_544; lean_object* x_545; lean_object* x_546; -x_544 = lean_ctor_get(x_531, 0); -x_545 = lean_ctor_get(x_531, 1); -lean_inc(x_545); -lean_inc(x_544); -lean_dec(x_531); -x_546 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_546, 0, x_544); -lean_ctor_set(x_546, 1, x_545); -return x_546; -} -} +return x_498; } else { -lean_object* x_547; lean_object* x_548; -x_547 = lean_array_fget(x_3, x_527); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_5); -lean_inc(x_547); -x_548 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_547, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_548) == 0) -{ -lean_object* x_549; lean_object* x_550; lean_object* x_551; lean_object* x_552; size_t x_553; size_t x_554; uint8_t x_555; -x_549 = lean_ctor_get(x_548, 0); -lean_inc(x_549); -x_550 = lean_ctor_get(x_548, 1); -lean_inc(x_550); -lean_dec(x_548); -x_551 = lean_ctor_get(x_549, 0); -lean_inc(x_551); -x_552 = lean_ctor_get(x_549, 1); -lean_inc(x_552); -lean_dec(x_549); -x_553 = lean_ptr_addr(x_547); -lean_dec(x_547); -x_554 = lean_ptr_addr(x_551); -x_555 = lean_usize_dec_eq(x_553, x_554); -if (x_555 == 0) -{ -lean_object* x_556; -lean_dec(x_551); +lean_object* x_499; lean_object* x_500; lean_object* x_501; lean_object* x_502; lean_object* x_503; lean_dec(x_3); lean_dec(x_2); +x_499 = lean_ctor_get(x_485, 0); +lean_inc(x_499); +x_500 = lean_ctor_get(x_485, 1); +lean_inc(x_500); +lean_dec(x_485); lean_inc(x_1); -x_556 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_552, x_7, x_8, x_9, x_10, x_550); +x_501 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_486, x_484, x_1); +x_502 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_502, 0, x_499); +lean_ctor_set(x_502, 1, x_500); +lean_ctor_set(x_502, 2, x_501); +lean_inc(x_1); +x_503 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_502, x_7, x_8, x_9, x_10, x_483); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_5); -return x_556; +return x_503; +} } else { -lean_object* x_557; lean_object* x_558; lean_object* x_559; -x_557 = lean_array_set(x_3, x_527, x_551); -x_558 = l_Lean_mkAppN(x_2, x_557); -lean_dec(x_557); -x_559 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_558, x_5, x_552, x_7, x_8, x_9, x_10, x_550); +lean_object* x_504; lean_object* x_505; +lean_dec(x_486); +lean_dec(x_484); +lean_dec(x_480); +lean_dec(x_3); +lean_dec(x_2); +x_504 = lean_ctor_get(x_487, 0); +lean_inc(x_504); +lean_dec(x_487); +x_505 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_504, x_5, x_485, x_7, x_8, x_9, x_10, x_483); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_5); -return x_559; +return x_505; } } else { -uint8_t x_560; -lean_dec(x_547); +uint8_t x_506; +lean_dec(x_480); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -9669,127 +9113,125 @@ lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_560 = !lean_is_exclusive(x_548); -if (x_560 == 0) +x_506 = !lean_is_exclusive(x_481); +if (x_506 == 0) { -return x_548; +return x_481; } else { -lean_object* x_561; lean_object* x_562; lean_object* x_563; -x_561 = lean_ctor_get(x_548, 0); -x_562 = lean_ctor_get(x_548, 1); -lean_inc(x_562); -lean_inc(x_561); -lean_dec(x_548); -x_563 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_563, 0, x_561); -lean_ctor_set(x_563, 1, x_562); -return x_563; -} -} +lean_object* x_507; lean_object* x_508; lean_object* x_509; +x_507 = lean_ctor_get(x_481, 0); +x_508 = lean_ctor_get(x_481, 1); +lean_inc(x_508); +lean_inc(x_507); +lean_dec(x_481); +x_509 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_509, 0, x_507); +lean_ctor_set(x_509, 1, x_508); +return x_509; } } } -block_519: +block_544: { -lean_object* x_487; -lean_dec(x_486); +lean_object* x_512; +lean_dec(x_511); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_2); -x_487 = l_Lean_Meta_getFunInfo(x_2, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_487) == 0) +x_512 = l_Lean_Meta_getFunInfo(x_2, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_512) == 0) { -lean_object* x_488; lean_object* x_489; lean_object* x_490; lean_object* x_491; lean_object* x_492; lean_object* x_493; lean_object* x_494; uint8_t x_495; lean_object* x_496; lean_object* x_497; lean_object* x_498; -x_488 = lean_ctor_get(x_487, 0); -lean_inc(x_488); -x_489 = lean_ctor_get(x_487, 1); -lean_inc(x_489); -lean_dec(x_487); -x_490 = lean_ctor_get(x_488, 0); -lean_inc(x_490); -lean_dec(x_488); -x_491 = lean_array_get_size(x_3); -x_492 = lean_unsigned_to_nat(0u); -x_493 = lean_unsigned_to_nat(1u); -x_494 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_494, 0, x_492); -lean_ctor_set(x_494, 1, x_491); -lean_ctor_set(x_494, 2, x_493); -x_495 = 0; -x_496 = lean_box(x_495); -x_497 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_497, 0, x_3); -lean_ctor_set(x_497, 1, x_496); +lean_object* x_513; lean_object* x_514; lean_object* x_515; lean_object* x_516; lean_object* x_517; lean_object* x_518; lean_object* x_519; uint8_t x_520; lean_object* x_521; lean_object* x_522; lean_object* x_523; +x_513 = lean_ctor_get(x_512, 0); +lean_inc(x_513); +x_514 = lean_ctor_get(x_512, 1); +lean_inc(x_514); +lean_dec(x_512); +x_515 = lean_ctor_get(x_513, 0); +lean_inc(x_515); +lean_dec(x_513); +x_516 = lean_array_get_size(x_3); +x_517 = lean_unsigned_to_nat(0u); +x_518 = lean_unsigned_to_nat(1u); +x_519 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_519, 0, x_517); +lean_ctor_set(x_519, 1, x_516); +lean_ctor_set(x_519, 2, x_518); +x_520 = 0; +x_521 = lean_box(x_520); +x_522 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_522, 0, x_3); +lean_ctor_set(x_522, 1, x_521); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_5); lean_inc(x_2); -x_498 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8(x_2, x_490, x_494, x_494, x_497, x_492, lean_box(0), lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10, x_489); -lean_dec(x_494); -lean_dec(x_490); -if (lean_obj_tag(x_498) == 0) -{ -lean_object* x_499; lean_object* x_500; lean_object* x_501; uint8_t x_502; -x_499 = lean_ctor_get(x_498, 0); -lean_inc(x_499); -x_500 = lean_ctor_get(x_499, 0); -lean_inc(x_500); -x_501 = lean_ctor_get(x_500, 1); -lean_inc(x_501); -x_502 = lean_unbox(x_501); -lean_dec(x_501); -if (x_502 == 0) -{ -lean_object* x_503; lean_object* x_504; lean_object* x_505; -lean_dec(x_500); +x_523 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8(x_2, x_515, x_519, x_519, x_522, x_517, lean_box(0), lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10, x_514); +lean_dec(x_519); +lean_dec(x_515); +if (lean_obj_tag(x_523) == 0) +{ +lean_object* x_524; lean_object* x_525; lean_object* x_526; uint8_t x_527; +x_524 = lean_ctor_get(x_523, 0); +lean_inc(x_524); +x_525 = lean_ctor_get(x_524, 0); +lean_inc(x_525); +x_526 = lean_ctor_get(x_525, 1); +lean_inc(x_526); +x_527 = lean_unbox(x_526); +lean_dec(x_526); +if (x_527 == 0) +{ +lean_object* x_528; lean_object* x_529; lean_object* x_530; +lean_dec(x_525); lean_dec(x_2); -x_503 = lean_ctor_get(x_498, 1); -lean_inc(x_503); -lean_dec(x_498); -x_504 = lean_ctor_get(x_499, 1); -lean_inc(x_504); -lean_dec(x_499); +x_528 = lean_ctor_get(x_523, 1); +lean_inc(x_528); +lean_dec(x_523); +x_529 = lean_ctor_get(x_524, 1); +lean_inc(x_529); +lean_dec(x_524); lean_inc(x_1); -x_505 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_504, x_7, x_8, x_9, x_10, x_503); +x_530 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_529, x_7, x_8, x_9, x_10, x_528); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_5); -return x_505; +return x_530; } else { -lean_object* x_506; lean_object* x_507; lean_object* x_508; lean_object* x_509; lean_object* x_510; -x_506 = lean_ctor_get(x_498, 1); -lean_inc(x_506); -lean_dec(x_498); -x_507 = lean_ctor_get(x_499, 1); -lean_inc(x_507); -lean_dec(x_499); -x_508 = lean_ctor_get(x_500, 0); -lean_inc(x_508); -lean_dec(x_500); -x_509 = l_Lean_mkAppN(x_2, x_508); -lean_dec(x_508); -x_510 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_509, x_5, x_507, x_7, x_8, x_9, x_10, x_506); +lean_object* x_531; lean_object* x_532; lean_object* x_533; lean_object* x_534; lean_object* x_535; +x_531 = lean_ctor_get(x_523, 1); +lean_inc(x_531); +lean_dec(x_523); +x_532 = lean_ctor_get(x_524, 1); +lean_inc(x_532); +lean_dec(x_524); +x_533 = lean_ctor_get(x_525, 0); +lean_inc(x_533); +lean_dec(x_525); +x_534 = l_Lean_mkAppN(x_2, x_533); +lean_dec(x_533); +x_535 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_534, x_5, x_532, x_7, x_8, x_9, x_10, x_531); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_5); -return x_510; +return x_535; } } else { -uint8_t x_511; +uint8_t x_536; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -9797,29 +9239,29 @@ lean_dec(x_7); lean_dec(x_5); lean_dec(x_2); lean_dec(x_1); -x_511 = !lean_is_exclusive(x_498); -if (x_511 == 0) +x_536 = !lean_is_exclusive(x_523); +if (x_536 == 0) { -return x_498; +return x_523; } else { -lean_object* x_512; lean_object* x_513; lean_object* x_514; -x_512 = lean_ctor_get(x_498, 0); -x_513 = lean_ctor_get(x_498, 1); -lean_inc(x_513); -lean_inc(x_512); -lean_dec(x_498); -x_514 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_514, 0, x_512); -lean_ctor_set(x_514, 1, x_513); -return x_514; +lean_object* x_537; lean_object* x_538; lean_object* x_539; +x_537 = lean_ctor_get(x_523, 0); +x_538 = lean_ctor_get(x_523, 1); +lean_inc(x_538); +lean_inc(x_537); +lean_dec(x_523); +x_539 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_539, 0, x_537); +lean_ctor_set(x_539, 1, x_538); +return x_539; } } } else { -uint8_t x_515; +uint8_t x_540; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -9829,213 +9271,188 @@ lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_515 = !lean_is_exclusive(x_487); -if (x_515 == 0) +x_540 = !lean_is_exclusive(x_512); +if (x_540 == 0) { -return x_487; +return x_512; } else { -lean_object* x_516; lean_object* x_517; lean_object* x_518; -x_516 = lean_ctor_get(x_487, 0); -x_517 = lean_ctor_get(x_487, 1); -lean_inc(x_517); -lean_inc(x_516); -lean_dec(x_487); -x_518 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_518, 0, x_516); -lean_ctor_set(x_518, 1, x_517); -return x_518; +lean_object* x_541; lean_object* x_542; lean_object* x_543; +x_541 = lean_ctor_get(x_512, 0); +x_542 = lean_ctor_get(x_512, 1); +lean_inc(x_542); +lean_inc(x_541); +lean_dec(x_512); +x_543 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_543, 0, x_541); +lean_ctor_set(x_543, 1, x_542); +return x_543; } } } } case 8: { -lean_object* x_564; lean_object* x_598; uint8_t x_599; +lean_object* x_557; lean_object* x_588; lean_object* x_622; uint8_t x_623; lean_dec(x_4); -x_598 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__4; -x_599 = l_Lean_Expr_isConstOf(x_2, x_598); -if (x_599 == 0) +x_622 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__4; +x_623 = l_Lean_Expr_isConstOf(x_2, x_622); +if (x_623 == 0) { -lean_object* x_600; -x_600 = lean_box(0); -x_564 = x_600; -goto block_597; +lean_object* x_624; +x_624 = lean_box(0); +x_588 = x_624; +goto block_621; } else { -lean_object* x_601; lean_object* x_602; uint8_t x_603; -x_601 = lean_array_get_size(x_3); -x_602 = lean_unsigned_to_nat(2u); -x_603 = lean_nat_dec_eq(x_601, x_602); -if (x_603 == 0) +lean_object* x_625; lean_object* x_626; uint8_t x_627; +x_625 = lean_array_get_size(x_3); +x_626 = lean_unsigned_to_nat(2u); +x_627 = lean_nat_dec_eq(x_625, x_626); +if (x_627 == 0) { -lean_object* x_604; -lean_dec(x_601); -x_604 = lean_box(0); -x_564 = x_604; -goto block_597; +lean_object* x_628; +lean_dec(x_625); +x_628 = lean_box(0); +x_588 = x_628; +goto block_621; } else { -lean_object* x_605; uint8_t x_606; -x_605 = lean_unsigned_to_nat(0u); -x_606 = lean_nat_dec_lt(x_605, x_601); -lean_dec(x_601); -if (x_606 == 0) -{ -lean_object* x_607; lean_object* x_608; lean_object* x_609; -x_607 = l_Lean_instInhabitedExpr; -x_608 = l_outOfBounds___rarg(x_607); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_5); -lean_inc(x_608); -x_609 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_608, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_609) == 0) +lean_object* x_629; uint8_t x_630; +x_629 = lean_unsigned_to_nat(0u); +x_630 = lean_nat_dec_lt(x_629, x_625); +lean_dec(x_625); +if (x_630 == 0) { -lean_object* x_610; lean_object* x_611; lean_object* x_612; lean_object* x_613; size_t x_614; size_t x_615; uint8_t x_616; -x_610 = lean_ctor_get(x_609, 0); -lean_inc(x_610); -x_611 = lean_ctor_get(x_609, 1); -lean_inc(x_611); -lean_dec(x_609); -x_612 = lean_ctor_get(x_610, 0); -lean_inc(x_612); -x_613 = lean_ctor_get(x_610, 1); -lean_inc(x_613); -lean_dec(x_610); -x_614 = lean_ptr_addr(x_608); -lean_dec(x_608); -x_615 = lean_ptr_addr(x_612); -x_616 = lean_usize_dec_eq(x_614, x_615); -if (x_616 == 0) -{ -lean_object* x_617; -lean_dec(x_612); -lean_dec(x_3); -lean_dec(x_2); -lean_inc(x_1); -x_617 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_613, x_7, x_8, x_9, x_10, x_611); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_617; +lean_object* x_631; lean_object* x_632; +x_631 = l_Lean_instInhabitedExpr; +x_632 = l_outOfBounds___rarg(x_631); +x_557 = x_632; +goto block_587; } else { -lean_object* x_618; lean_object* x_619; lean_object* x_620; -x_618 = lean_array_set(x_3, x_605, x_612); -x_619 = l_Lean_mkAppN(x_2, x_618); -lean_dec(x_618); -x_620 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_619, x_5, x_613, x_7, x_8, x_9, x_10, x_611); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_620; +lean_object* x_633; +x_633 = lean_array_fget(x_3, x_629); +x_557 = x_633; +goto block_587; } } -else +} +block_587: +{ +lean_object* x_558; +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_5); +lean_inc(x_557); +x_558 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_557, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_558) == 0) +{ +lean_object* x_559; lean_object* x_560; lean_object* x_561; lean_object* x_562; lean_object* x_563; lean_object* x_564; +x_559 = lean_ctor_get(x_558, 0); +lean_inc(x_559); +x_560 = lean_ctor_get(x_558, 1); +lean_inc(x_560); +lean_dec(x_558); +x_561 = lean_ctor_get(x_559, 0); +lean_inc(x_561); +x_562 = lean_ctor_get(x_559, 1); +lean_inc(x_562); +lean_dec(x_559); +x_563 = lean_ctor_get(x_562, 2); +lean_inc(x_563); +lean_inc(x_563); +x_564 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_563, x_561); +if (lean_obj_tag(x_564) == 0) +{ +size_t x_565; size_t x_566; uint8_t x_567; +x_565 = lean_ptr_addr(x_557); +lean_dec(x_557); +x_566 = lean_ptr_addr(x_561); +x_567 = lean_usize_dec_eq(x_565, x_566); +if (x_567 == 0) { -uint8_t x_621; -lean_dec(x_608); +lean_object* x_568; lean_object* x_569; lean_object* x_570; lean_object* x_571; lean_object* x_572; lean_object* x_573; lean_object* x_574; lean_object* x_575; +x_568 = lean_unsigned_to_nat(0u); +lean_inc(x_561); +x_569 = lean_array_set(x_3, x_568, x_561); +x_570 = l_Lean_mkAppN(x_2, x_569); +lean_dec(x_569); +x_571 = lean_ctor_get(x_562, 0); +lean_inc(x_571); +x_572 = lean_ctor_get(x_562, 1); +lean_inc(x_572); +lean_dec(x_562); +lean_inc(x_570); +x_573 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_563, x_561, x_570); +x_574 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_574, 0, x_571); +lean_ctor_set(x_574, 1, x_572); +lean_ctor_set(x_574, 2, x_573); +x_575 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_570, x_5, x_574, x_7, x_8, x_9, x_10, x_560); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_621 = !lean_is_exclusive(x_609); -if (x_621 == 0) -{ -return x_609; -} -else -{ -lean_object* x_622; lean_object* x_623; lean_object* x_624; -x_622 = lean_ctor_get(x_609, 0); -x_623 = lean_ctor_get(x_609, 1); -lean_inc(x_623); -lean_inc(x_622); -lean_dec(x_609); -x_624 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_624, 0, x_622); -lean_ctor_set(x_624, 1, x_623); -return x_624; -} -} +return x_575; } else { -lean_object* x_625; lean_object* x_626; -x_625 = lean_array_fget(x_3, x_605); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_5); -lean_inc(x_625); -x_626 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_625, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_626) == 0) -{ -lean_object* x_627; lean_object* x_628; lean_object* x_629; lean_object* x_630; size_t x_631; size_t x_632; uint8_t x_633; -x_627 = lean_ctor_get(x_626, 0); -lean_inc(x_627); -x_628 = lean_ctor_get(x_626, 1); -lean_inc(x_628); -lean_dec(x_626); -x_629 = lean_ctor_get(x_627, 0); -lean_inc(x_629); -x_630 = lean_ctor_get(x_627, 1); -lean_inc(x_630); -lean_dec(x_627); -x_631 = lean_ptr_addr(x_625); -lean_dec(x_625); -x_632 = lean_ptr_addr(x_629); -x_633 = lean_usize_dec_eq(x_631, x_632); -if (x_633 == 0) -{ -lean_object* x_634; -lean_dec(x_629); +lean_object* x_576; lean_object* x_577; lean_object* x_578; lean_object* x_579; lean_object* x_580; lean_dec(x_3); lean_dec(x_2); +x_576 = lean_ctor_get(x_562, 0); +lean_inc(x_576); +x_577 = lean_ctor_get(x_562, 1); +lean_inc(x_577); +lean_dec(x_562); lean_inc(x_1); -x_634 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_630, x_7, x_8, x_9, x_10, x_628); +x_578 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_563, x_561, x_1); +x_579 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_579, 0, x_576); +lean_ctor_set(x_579, 1, x_577); +lean_ctor_set(x_579, 2, x_578); +lean_inc(x_1); +x_580 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_579, x_7, x_8, x_9, x_10, x_560); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_5); -return x_634; +return x_580; +} } else { -lean_object* x_635; lean_object* x_636; lean_object* x_637; -x_635 = lean_array_set(x_3, x_605, x_629); -x_636 = l_Lean_mkAppN(x_2, x_635); -lean_dec(x_635); -x_637 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_636, x_5, x_630, x_7, x_8, x_9, x_10, x_628); +lean_object* x_581; lean_object* x_582; +lean_dec(x_563); +lean_dec(x_561); +lean_dec(x_557); +lean_dec(x_3); +lean_dec(x_2); +x_581 = lean_ctor_get(x_564, 0); +lean_inc(x_581); +lean_dec(x_564); +x_582 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_581, x_5, x_562, x_7, x_8, x_9, x_10, x_560); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_5); -return x_637; +return x_582; } } else { -uint8_t x_638; -lean_dec(x_625); +uint8_t x_583; +lean_dec(x_557); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -10044,127 +9461,125 @@ lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_638 = !lean_is_exclusive(x_626); -if (x_638 == 0) +x_583 = !lean_is_exclusive(x_558); +if (x_583 == 0) { -return x_626; +return x_558; } else { -lean_object* x_639; lean_object* x_640; lean_object* x_641; -x_639 = lean_ctor_get(x_626, 0); -x_640 = lean_ctor_get(x_626, 1); -lean_inc(x_640); -lean_inc(x_639); -lean_dec(x_626); -x_641 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_641, 0, x_639); -lean_ctor_set(x_641, 1, x_640); -return x_641; -} -} +lean_object* x_584; lean_object* x_585; lean_object* x_586; +x_584 = lean_ctor_get(x_558, 0); +x_585 = lean_ctor_get(x_558, 1); +lean_inc(x_585); +lean_inc(x_584); +lean_dec(x_558); +x_586 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_586, 0, x_584); +lean_ctor_set(x_586, 1, x_585); +return x_586; } } } -block_597: +block_621: { -lean_object* x_565; -lean_dec(x_564); +lean_object* x_589; +lean_dec(x_588); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_2); -x_565 = l_Lean_Meta_getFunInfo(x_2, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_565) == 0) -{ -lean_object* x_566; lean_object* x_567; lean_object* x_568; lean_object* x_569; lean_object* x_570; lean_object* x_571; lean_object* x_572; uint8_t x_573; lean_object* x_574; lean_object* x_575; lean_object* x_576; -x_566 = lean_ctor_get(x_565, 0); -lean_inc(x_566); -x_567 = lean_ctor_get(x_565, 1); -lean_inc(x_567); -lean_dec(x_565); -x_568 = lean_ctor_get(x_566, 0); -lean_inc(x_568); -lean_dec(x_566); -x_569 = lean_array_get_size(x_3); -x_570 = lean_unsigned_to_nat(0u); -x_571 = lean_unsigned_to_nat(1u); -x_572 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_572, 0, x_570); -lean_ctor_set(x_572, 1, x_569); -lean_ctor_set(x_572, 2, x_571); -x_573 = 0; -x_574 = lean_box(x_573); -x_575 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_575, 0, x_3); -lean_ctor_set(x_575, 1, x_574); +x_589 = l_Lean_Meta_getFunInfo(x_2, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_589) == 0) +{ +lean_object* x_590; lean_object* x_591; lean_object* x_592; lean_object* x_593; lean_object* x_594; lean_object* x_595; lean_object* x_596; uint8_t x_597; lean_object* x_598; lean_object* x_599; lean_object* x_600; +x_590 = lean_ctor_get(x_589, 0); +lean_inc(x_590); +x_591 = lean_ctor_get(x_589, 1); +lean_inc(x_591); +lean_dec(x_589); +x_592 = lean_ctor_get(x_590, 0); +lean_inc(x_592); +lean_dec(x_590); +x_593 = lean_array_get_size(x_3); +x_594 = lean_unsigned_to_nat(0u); +x_595 = lean_unsigned_to_nat(1u); +x_596 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_596, 0, x_594); +lean_ctor_set(x_596, 1, x_593); +lean_ctor_set(x_596, 2, x_595); +x_597 = 0; +x_598 = lean_box(x_597); +x_599 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_599, 0, x_3); +lean_ctor_set(x_599, 1, x_598); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_5); lean_inc(x_2); -x_576 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8(x_2, x_568, x_572, x_572, x_575, x_570, lean_box(0), lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10, x_567); -lean_dec(x_572); -lean_dec(x_568); -if (lean_obj_tag(x_576) == 0) -{ -lean_object* x_577; lean_object* x_578; lean_object* x_579; uint8_t x_580; -x_577 = lean_ctor_get(x_576, 0); -lean_inc(x_577); -x_578 = lean_ctor_get(x_577, 0); -lean_inc(x_578); -x_579 = lean_ctor_get(x_578, 1); -lean_inc(x_579); -x_580 = lean_unbox(x_579); -lean_dec(x_579); -if (x_580 == 0) -{ -lean_object* x_581; lean_object* x_582; lean_object* x_583; -lean_dec(x_578); +x_600 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8(x_2, x_592, x_596, x_596, x_599, x_594, lean_box(0), lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10, x_591); +lean_dec(x_596); +lean_dec(x_592); +if (lean_obj_tag(x_600) == 0) +{ +lean_object* x_601; lean_object* x_602; lean_object* x_603; uint8_t x_604; +x_601 = lean_ctor_get(x_600, 0); +lean_inc(x_601); +x_602 = lean_ctor_get(x_601, 0); +lean_inc(x_602); +x_603 = lean_ctor_get(x_602, 1); +lean_inc(x_603); +x_604 = lean_unbox(x_603); +lean_dec(x_603); +if (x_604 == 0) +{ +lean_object* x_605; lean_object* x_606; lean_object* x_607; +lean_dec(x_602); lean_dec(x_2); -x_581 = lean_ctor_get(x_576, 1); -lean_inc(x_581); -lean_dec(x_576); -x_582 = lean_ctor_get(x_577, 1); -lean_inc(x_582); -lean_dec(x_577); +x_605 = lean_ctor_get(x_600, 1); +lean_inc(x_605); +lean_dec(x_600); +x_606 = lean_ctor_get(x_601, 1); +lean_inc(x_606); +lean_dec(x_601); lean_inc(x_1); -x_583 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_582, x_7, x_8, x_9, x_10, x_581); +x_607 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_606, x_7, x_8, x_9, x_10, x_605); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_5); -return x_583; +return x_607; } else { -lean_object* x_584; lean_object* x_585; lean_object* x_586; lean_object* x_587; lean_object* x_588; -x_584 = lean_ctor_get(x_576, 1); -lean_inc(x_584); -lean_dec(x_576); -x_585 = lean_ctor_get(x_577, 1); -lean_inc(x_585); -lean_dec(x_577); -x_586 = lean_ctor_get(x_578, 0); -lean_inc(x_586); -lean_dec(x_578); -x_587 = l_Lean_mkAppN(x_2, x_586); -lean_dec(x_586); -x_588 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_587, x_5, x_585, x_7, x_8, x_9, x_10, x_584); +lean_object* x_608; lean_object* x_609; lean_object* x_610; lean_object* x_611; lean_object* x_612; +x_608 = lean_ctor_get(x_600, 1); +lean_inc(x_608); +lean_dec(x_600); +x_609 = lean_ctor_get(x_601, 1); +lean_inc(x_609); +lean_dec(x_601); +x_610 = lean_ctor_get(x_602, 0); +lean_inc(x_610); +lean_dec(x_602); +x_611 = l_Lean_mkAppN(x_2, x_610); +lean_dec(x_610); +x_612 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_611, x_5, x_609, x_7, x_8, x_9, x_10, x_608); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_5); -return x_588; +return x_612; } } else { -uint8_t x_589; +uint8_t x_613; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -10172,29 +9587,29 @@ lean_dec(x_7); lean_dec(x_5); lean_dec(x_2); lean_dec(x_1); -x_589 = !lean_is_exclusive(x_576); -if (x_589 == 0) +x_613 = !lean_is_exclusive(x_600); +if (x_613 == 0) { -return x_576; +return x_600; } else { -lean_object* x_590; lean_object* x_591; lean_object* x_592; -x_590 = lean_ctor_get(x_576, 0); -x_591 = lean_ctor_get(x_576, 1); -lean_inc(x_591); -lean_inc(x_590); -lean_dec(x_576); -x_592 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_592, 0, x_590); -lean_ctor_set(x_592, 1, x_591); -return x_592; +lean_object* x_614; lean_object* x_615; lean_object* x_616; +x_614 = lean_ctor_get(x_600, 0); +x_615 = lean_ctor_get(x_600, 1); +lean_inc(x_615); +lean_inc(x_614); +lean_dec(x_600); +x_616 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_616, 0, x_614); +lean_ctor_set(x_616, 1, x_615); +return x_616; } } } else { -uint8_t x_593; +uint8_t x_617; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -10204,213 +9619,188 @@ lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_593 = !lean_is_exclusive(x_565); -if (x_593 == 0) +x_617 = !lean_is_exclusive(x_589); +if (x_617 == 0) { -return x_565; +return x_589; } else { -lean_object* x_594; lean_object* x_595; lean_object* x_596; -x_594 = lean_ctor_get(x_565, 0); -x_595 = lean_ctor_get(x_565, 1); -lean_inc(x_595); -lean_inc(x_594); -lean_dec(x_565); -x_596 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_596, 0, x_594); -lean_ctor_set(x_596, 1, x_595); -return x_596; +lean_object* x_618; lean_object* x_619; lean_object* x_620; +x_618 = lean_ctor_get(x_589, 0); +x_619 = lean_ctor_get(x_589, 1); +lean_inc(x_619); +lean_inc(x_618); +lean_dec(x_589); +x_620 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_620, 0, x_618); +lean_ctor_set(x_620, 1, x_619); +return x_620; } } } } case 9: { -lean_object* x_642; lean_object* x_676; uint8_t x_677; +lean_object* x_634; lean_object* x_665; lean_object* x_699; uint8_t x_700; lean_dec(x_4); -x_676 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__4; -x_677 = l_Lean_Expr_isConstOf(x_2, x_676); -if (x_677 == 0) +x_699 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__4; +x_700 = l_Lean_Expr_isConstOf(x_2, x_699); +if (x_700 == 0) { -lean_object* x_678; -x_678 = lean_box(0); -x_642 = x_678; -goto block_675; +lean_object* x_701; +x_701 = lean_box(0); +x_665 = x_701; +goto block_698; } else { -lean_object* x_679; lean_object* x_680; uint8_t x_681; -x_679 = lean_array_get_size(x_3); -x_680 = lean_unsigned_to_nat(2u); -x_681 = lean_nat_dec_eq(x_679, x_680); -if (x_681 == 0) +lean_object* x_702; lean_object* x_703; uint8_t x_704; +x_702 = lean_array_get_size(x_3); +x_703 = lean_unsigned_to_nat(2u); +x_704 = lean_nat_dec_eq(x_702, x_703); +if (x_704 == 0) { -lean_object* x_682; -lean_dec(x_679); -x_682 = lean_box(0); -x_642 = x_682; -goto block_675; +lean_object* x_705; +lean_dec(x_702); +x_705 = lean_box(0); +x_665 = x_705; +goto block_698; } else { -lean_object* x_683; uint8_t x_684; -x_683 = lean_unsigned_to_nat(0u); -x_684 = lean_nat_dec_lt(x_683, x_679); -lean_dec(x_679); -if (x_684 == 0) -{ -lean_object* x_685; lean_object* x_686; lean_object* x_687; -x_685 = l_Lean_instInhabitedExpr; -x_686 = l_outOfBounds___rarg(x_685); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_5); -lean_inc(x_686); -x_687 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_686, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_687) == 0) -{ -lean_object* x_688; lean_object* x_689; lean_object* x_690; lean_object* x_691; size_t x_692; size_t x_693; uint8_t x_694; -x_688 = lean_ctor_get(x_687, 0); -lean_inc(x_688); -x_689 = lean_ctor_get(x_687, 1); -lean_inc(x_689); -lean_dec(x_687); -x_690 = lean_ctor_get(x_688, 0); -lean_inc(x_690); -x_691 = lean_ctor_get(x_688, 1); -lean_inc(x_691); -lean_dec(x_688); -x_692 = lean_ptr_addr(x_686); -lean_dec(x_686); -x_693 = lean_ptr_addr(x_690); -x_694 = lean_usize_dec_eq(x_692, x_693); -if (x_694 == 0) +lean_object* x_706; uint8_t x_707; +x_706 = lean_unsigned_to_nat(0u); +x_707 = lean_nat_dec_lt(x_706, x_702); +lean_dec(x_702); +if (x_707 == 0) { -lean_object* x_695; -lean_dec(x_690); -lean_dec(x_3); -lean_dec(x_2); -lean_inc(x_1); -x_695 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_691, x_7, x_8, x_9, x_10, x_689); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_695; +lean_object* x_708; lean_object* x_709; +x_708 = l_Lean_instInhabitedExpr; +x_709 = l_outOfBounds___rarg(x_708); +x_634 = x_709; +goto block_664; } else { -lean_object* x_696; lean_object* x_697; lean_object* x_698; -x_696 = lean_array_set(x_3, x_683, x_690); -x_697 = l_Lean_mkAppN(x_2, x_696); -lean_dec(x_696); -x_698 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_697, x_5, x_691, x_7, x_8, x_9, x_10, x_689); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_698; +lean_object* x_710; +x_710 = lean_array_fget(x_3, x_706); +x_634 = x_710; +goto block_664; } } -else +} +block_664: { -uint8_t x_699; -lean_dec(x_686); +lean_object* x_635; +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_5); +lean_inc(x_634); +x_635 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_634, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_635) == 0) +{ +lean_object* x_636; lean_object* x_637; lean_object* x_638; lean_object* x_639; lean_object* x_640; lean_object* x_641; +x_636 = lean_ctor_get(x_635, 0); +lean_inc(x_636); +x_637 = lean_ctor_get(x_635, 1); +lean_inc(x_637); +lean_dec(x_635); +x_638 = lean_ctor_get(x_636, 0); +lean_inc(x_638); +x_639 = lean_ctor_get(x_636, 1); +lean_inc(x_639); +lean_dec(x_636); +x_640 = lean_ctor_get(x_639, 2); +lean_inc(x_640); +lean_inc(x_640); +x_641 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_640, x_638); +if (lean_obj_tag(x_641) == 0) +{ +size_t x_642; size_t x_643; uint8_t x_644; +x_642 = lean_ptr_addr(x_634); +lean_dec(x_634); +x_643 = lean_ptr_addr(x_638); +x_644 = lean_usize_dec_eq(x_642, x_643); +if (x_644 == 0) +{ +lean_object* x_645; lean_object* x_646; lean_object* x_647; lean_object* x_648; lean_object* x_649; lean_object* x_650; lean_object* x_651; lean_object* x_652; +x_645 = lean_unsigned_to_nat(0u); +lean_inc(x_638); +x_646 = lean_array_set(x_3, x_645, x_638); +x_647 = l_Lean_mkAppN(x_2, x_646); +lean_dec(x_646); +x_648 = lean_ctor_get(x_639, 0); +lean_inc(x_648); +x_649 = lean_ctor_get(x_639, 1); +lean_inc(x_649); +lean_dec(x_639); +lean_inc(x_647); +x_650 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_640, x_638, x_647); +x_651 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_651, 0, x_648); +lean_ctor_set(x_651, 1, x_649); +lean_ctor_set(x_651, 2, x_650); +x_652 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_647, x_5, x_651, x_7, x_8, x_9, x_10, x_637); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_699 = !lean_is_exclusive(x_687); -if (x_699 == 0) -{ -return x_687; -} -else -{ -lean_object* x_700; lean_object* x_701; lean_object* x_702; -x_700 = lean_ctor_get(x_687, 0); -x_701 = lean_ctor_get(x_687, 1); -lean_inc(x_701); -lean_inc(x_700); -lean_dec(x_687); -x_702 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_702, 0, x_700); -lean_ctor_set(x_702, 1, x_701); -return x_702; -} -} +return x_652; } else { -lean_object* x_703; lean_object* x_704; -x_703 = lean_array_fget(x_3, x_683); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_5); -lean_inc(x_703); -x_704 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_703, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_704) == 0) -{ -lean_object* x_705; lean_object* x_706; lean_object* x_707; lean_object* x_708; size_t x_709; size_t x_710; uint8_t x_711; -x_705 = lean_ctor_get(x_704, 0); -lean_inc(x_705); -x_706 = lean_ctor_get(x_704, 1); -lean_inc(x_706); -lean_dec(x_704); -x_707 = lean_ctor_get(x_705, 0); -lean_inc(x_707); -x_708 = lean_ctor_get(x_705, 1); -lean_inc(x_708); -lean_dec(x_705); -x_709 = lean_ptr_addr(x_703); -lean_dec(x_703); -x_710 = lean_ptr_addr(x_707); -x_711 = lean_usize_dec_eq(x_709, x_710); -if (x_711 == 0) -{ -lean_object* x_712; -lean_dec(x_707); +lean_object* x_653; lean_object* x_654; lean_object* x_655; lean_object* x_656; lean_object* x_657; lean_dec(x_3); lean_dec(x_2); +x_653 = lean_ctor_get(x_639, 0); +lean_inc(x_653); +x_654 = lean_ctor_get(x_639, 1); +lean_inc(x_654); +lean_dec(x_639); +lean_inc(x_1); +x_655 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_640, x_638, x_1); +x_656 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_656, 0, x_653); +lean_ctor_set(x_656, 1, x_654); +lean_ctor_set(x_656, 2, x_655); lean_inc(x_1); -x_712 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_708, x_7, x_8, x_9, x_10, x_706); +x_657 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_656, x_7, x_8, x_9, x_10, x_637); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_5); -return x_712; +return x_657; +} } else { -lean_object* x_713; lean_object* x_714; lean_object* x_715; -x_713 = lean_array_set(x_3, x_683, x_707); -x_714 = l_Lean_mkAppN(x_2, x_713); -lean_dec(x_713); -x_715 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_714, x_5, x_708, x_7, x_8, x_9, x_10, x_706); +lean_object* x_658; lean_object* x_659; +lean_dec(x_640); +lean_dec(x_638); +lean_dec(x_634); +lean_dec(x_3); +lean_dec(x_2); +x_658 = lean_ctor_get(x_641, 0); +lean_inc(x_658); +lean_dec(x_641); +x_659 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_658, x_5, x_639, x_7, x_8, x_9, x_10, x_637); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_5); -return x_715; +return x_659; } } else { -uint8_t x_716; -lean_dec(x_703); +uint8_t x_660; +lean_dec(x_634); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -10419,127 +9809,125 @@ lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_716 = !lean_is_exclusive(x_704); -if (x_716 == 0) +x_660 = !lean_is_exclusive(x_635); +if (x_660 == 0) { -return x_704; +return x_635; } else { -lean_object* x_717; lean_object* x_718; lean_object* x_719; -x_717 = lean_ctor_get(x_704, 0); -x_718 = lean_ctor_get(x_704, 1); -lean_inc(x_718); -lean_inc(x_717); -lean_dec(x_704); -x_719 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_719, 0, x_717); -lean_ctor_set(x_719, 1, x_718); -return x_719; -} -} +lean_object* x_661; lean_object* x_662; lean_object* x_663; +x_661 = lean_ctor_get(x_635, 0); +x_662 = lean_ctor_get(x_635, 1); +lean_inc(x_662); +lean_inc(x_661); +lean_dec(x_635); +x_663 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_663, 0, x_661); +lean_ctor_set(x_663, 1, x_662); +return x_663; } } } -block_675: +block_698: { -lean_object* x_643; -lean_dec(x_642); +lean_object* x_666; +lean_dec(x_665); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_2); -x_643 = l_Lean_Meta_getFunInfo(x_2, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_643) == 0) -{ -lean_object* x_644; lean_object* x_645; lean_object* x_646; lean_object* x_647; lean_object* x_648; lean_object* x_649; lean_object* x_650; uint8_t x_651; lean_object* x_652; lean_object* x_653; lean_object* x_654; -x_644 = lean_ctor_get(x_643, 0); -lean_inc(x_644); -x_645 = lean_ctor_get(x_643, 1); -lean_inc(x_645); -lean_dec(x_643); -x_646 = lean_ctor_get(x_644, 0); -lean_inc(x_646); -lean_dec(x_644); -x_647 = lean_array_get_size(x_3); -x_648 = lean_unsigned_to_nat(0u); -x_649 = lean_unsigned_to_nat(1u); -x_650 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_650, 0, x_648); -lean_ctor_set(x_650, 1, x_647); -lean_ctor_set(x_650, 2, x_649); -x_651 = 0; -x_652 = lean_box(x_651); -x_653 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_653, 0, x_3); -lean_ctor_set(x_653, 1, x_652); +x_666 = l_Lean_Meta_getFunInfo(x_2, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_666) == 0) +{ +lean_object* x_667; lean_object* x_668; lean_object* x_669; lean_object* x_670; lean_object* x_671; lean_object* x_672; lean_object* x_673; uint8_t x_674; lean_object* x_675; lean_object* x_676; lean_object* x_677; +x_667 = lean_ctor_get(x_666, 0); +lean_inc(x_667); +x_668 = lean_ctor_get(x_666, 1); +lean_inc(x_668); +lean_dec(x_666); +x_669 = lean_ctor_get(x_667, 0); +lean_inc(x_669); +lean_dec(x_667); +x_670 = lean_array_get_size(x_3); +x_671 = lean_unsigned_to_nat(0u); +x_672 = lean_unsigned_to_nat(1u); +x_673 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_673, 0, x_671); +lean_ctor_set(x_673, 1, x_670); +lean_ctor_set(x_673, 2, x_672); +x_674 = 0; +x_675 = lean_box(x_674); +x_676 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_676, 0, x_3); +lean_ctor_set(x_676, 1, x_675); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_5); lean_inc(x_2); -x_654 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8(x_2, x_646, x_650, x_650, x_653, x_648, lean_box(0), lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10, x_645); -lean_dec(x_650); -lean_dec(x_646); -if (lean_obj_tag(x_654) == 0) -{ -lean_object* x_655; lean_object* x_656; lean_object* x_657; uint8_t x_658; -x_655 = lean_ctor_get(x_654, 0); -lean_inc(x_655); -x_656 = lean_ctor_get(x_655, 0); -lean_inc(x_656); -x_657 = lean_ctor_get(x_656, 1); -lean_inc(x_657); -x_658 = lean_unbox(x_657); -lean_dec(x_657); -if (x_658 == 0) -{ -lean_object* x_659; lean_object* x_660; lean_object* x_661; -lean_dec(x_656); +x_677 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8(x_2, x_669, x_673, x_673, x_676, x_671, lean_box(0), lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10, x_668); +lean_dec(x_673); +lean_dec(x_669); +if (lean_obj_tag(x_677) == 0) +{ +lean_object* x_678; lean_object* x_679; lean_object* x_680; uint8_t x_681; +x_678 = lean_ctor_get(x_677, 0); +lean_inc(x_678); +x_679 = lean_ctor_get(x_678, 0); +lean_inc(x_679); +x_680 = lean_ctor_get(x_679, 1); +lean_inc(x_680); +x_681 = lean_unbox(x_680); +lean_dec(x_680); +if (x_681 == 0) +{ +lean_object* x_682; lean_object* x_683; lean_object* x_684; +lean_dec(x_679); lean_dec(x_2); -x_659 = lean_ctor_get(x_654, 1); -lean_inc(x_659); -lean_dec(x_654); -x_660 = lean_ctor_get(x_655, 1); -lean_inc(x_660); -lean_dec(x_655); +x_682 = lean_ctor_get(x_677, 1); +lean_inc(x_682); +lean_dec(x_677); +x_683 = lean_ctor_get(x_678, 1); +lean_inc(x_683); +lean_dec(x_678); lean_inc(x_1); -x_661 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_660, x_7, x_8, x_9, x_10, x_659); +x_684 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_683, x_7, x_8, x_9, x_10, x_682); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_5); -return x_661; +return x_684; } else { -lean_object* x_662; lean_object* x_663; lean_object* x_664; lean_object* x_665; lean_object* x_666; -x_662 = lean_ctor_get(x_654, 1); -lean_inc(x_662); -lean_dec(x_654); -x_663 = lean_ctor_get(x_655, 1); -lean_inc(x_663); -lean_dec(x_655); -x_664 = lean_ctor_get(x_656, 0); -lean_inc(x_664); -lean_dec(x_656); -x_665 = l_Lean_mkAppN(x_2, x_664); -lean_dec(x_664); -x_666 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_665, x_5, x_663, x_7, x_8, x_9, x_10, x_662); +lean_object* x_685; lean_object* x_686; lean_object* x_687; lean_object* x_688; lean_object* x_689; +x_685 = lean_ctor_get(x_677, 1); +lean_inc(x_685); +lean_dec(x_677); +x_686 = lean_ctor_get(x_678, 1); +lean_inc(x_686); +lean_dec(x_678); +x_687 = lean_ctor_get(x_679, 0); +lean_inc(x_687); +lean_dec(x_679); +x_688 = l_Lean_mkAppN(x_2, x_687); +lean_dec(x_687); +x_689 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_688, x_5, x_686, x_7, x_8, x_9, x_10, x_685); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_5); -return x_666; +return x_689; } } else { -uint8_t x_667; +uint8_t x_690; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -10547,29 +9935,29 @@ lean_dec(x_7); lean_dec(x_5); lean_dec(x_2); lean_dec(x_1); -x_667 = !lean_is_exclusive(x_654); -if (x_667 == 0) +x_690 = !lean_is_exclusive(x_677); +if (x_690 == 0) { -return x_654; +return x_677; } else { -lean_object* x_668; lean_object* x_669; lean_object* x_670; -x_668 = lean_ctor_get(x_654, 0); -x_669 = lean_ctor_get(x_654, 1); -lean_inc(x_669); -lean_inc(x_668); -lean_dec(x_654); -x_670 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_670, 0, x_668); -lean_ctor_set(x_670, 1, x_669); -return x_670; +lean_object* x_691; lean_object* x_692; lean_object* x_693; +x_691 = lean_ctor_get(x_677, 0); +x_692 = lean_ctor_get(x_677, 1); +lean_inc(x_692); +lean_inc(x_691); +lean_dec(x_677); +x_693 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_693, 0, x_691); +lean_ctor_set(x_693, 1, x_692); +return x_693; } } } else { -uint8_t x_671; +uint8_t x_694; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -10579,213 +9967,188 @@ lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_671 = !lean_is_exclusive(x_643); -if (x_671 == 0) +x_694 = !lean_is_exclusive(x_666); +if (x_694 == 0) { -return x_643; +return x_666; } else { -lean_object* x_672; lean_object* x_673; lean_object* x_674; -x_672 = lean_ctor_get(x_643, 0); -x_673 = lean_ctor_get(x_643, 1); -lean_inc(x_673); -lean_inc(x_672); -lean_dec(x_643); -x_674 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_674, 0, x_672); -lean_ctor_set(x_674, 1, x_673); -return x_674; +lean_object* x_695; lean_object* x_696; lean_object* x_697; +x_695 = lean_ctor_get(x_666, 0); +x_696 = lean_ctor_get(x_666, 1); +lean_inc(x_696); +lean_inc(x_695); +lean_dec(x_666); +x_697 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_697, 0, x_695); +lean_ctor_set(x_697, 1, x_696); +return x_697; } } } } case 10: { -lean_object* x_720; lean_object* x_754; uint8_t x_755; +lean_object* x_711; lean_object* x_742; lean_object* x_776; uint8_t x_777; lean_dec(x_4); -x_754 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__4; -x_755 = l_Lean_Expr_isConstOf(x_2, x_754); -if (x_755 == 0) +x_776 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__4; +x_777 = l_Lean_Expr_isConstOf(x_2, x_776); +if (x_777 == 0) { -lean_object* x_756; -x_756 = lean_box(0); -x_720 = x_756; -goto block_753; +lean_object* x_778; +x_778 = lean_box(0); +x_742 = x_778; +goto block_775; } else { -lean_object* x_757; lean_object* x_758; uint8_t x_759; -x_757 = lean_array_get_size(x_3); -x_758 = lean_unsigned_to_nat(2u); -x_759 = lean_nat_dec_eq(x_757, x_758); -if (x_759 == 0) +lean_object* x_779; lean_object* x_780; uint8_t x_781; +x_779 = lean_array_get_size(x_3); +x_780 = lean_unsigned_to_nat(2u); +x_781 = lean_nat_dec_eq(x_779, x_780); +if (x_781 == 0) { -lean_object* x_760; -lean_dec(x_757); -x_760 = lean_box(0); -x_720 = x_760; -goto block_753; +lean_object* x_782; +lean_dec(x_779); +x_782 = lean_box(0); +x_742 = x_782; +goto block_775; } else { -lean_object* x_761; uint8_t x_762; -x_761 = lean_unsigned_to_nat(0u); -x_762 = lean_nat_dec_lt(x_761, x_757); -lean_dec(x_757); -if (x_762 == 0) -{ -lean_object* x_763; lean_object* x_764; lean_object* x_765; -x_763 = l_Lean_instInhabitedExpr; -x_764 = l_outOfBounds___rarg(x_763); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_5); -lean_inc(x_764); -x_765 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_764, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_765) == 0) -{ -lean_object* x_766; lean_object* x_767; lean_object* x_768; lean_object* x_769; size_t x_770; size_t x_771; uint8_t x_772; -x_766 = lean_ctor_get(x_765, 0); -lean_inc(x_766); -x_767 = lean_ctor_get(x_765, 1); -lean_inc(x_767); -lean_dec(x_765); -x_768 = lean_ctor_get(x_766, 0); -lean_inc(x_768); -x_769 = lean_ctor_get(x_766, 1); -lean_inc(x_769); -lean_dec(x_766); -x_770 = lean_ptr_addr(x_764); -lean_dec(x_764); -x_771 = lean_ptr_addr(x_768); -x_772 = lean_usize_dec_eq(x_770, x_771); -if (x_772 == 0) +lean_object* x_783; uint8_t x_784; +x_783 = lean_unsigned_to_nat(0u); +x_784 = lean_nat_dec_lt(x_783, x_779); +lean_dec(x_779); +if (x_784 == 0) { -lean_object* x_773; -lean_dec(x_768); -lean_dec(x_3); -lean_dec(x_2); -lean_inc(x_1); -x_773 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_769, x_7, x_8, x_9, x_10, x_767); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_773; +lean_object* x_785; lean_object* x_786; +x_785 = l_Lean_instInhabitedExpr; +x_786 = l_outOfBounds___rarg(x_785); +x_711 = x_786; +goto block_741; } else { -lean_object* x_774; lean_object* x_775; lean_object* x_776; -x_774 = lean_array_set(x_3, x_761, x_768); -x_775 = l_Lean_mkAppN(x_2, x_774); -lean_dec(x_774); -x_776 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_775, x_5, x_769, x_7, x_8, x_9, x_10, x_767); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_776; +lean_object* x_787; +x_787 = lean_array_fget(x_3, x_783); +x_711 = x_787; +goto block_741; } } -else +} +block_741: { -uint8_t x_777; -lean_dec(x_764); +lean_object* x_712; +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_5); +lean_inc(x_711); +x_712 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_711, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_712) == 0) +{ +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_713 = lean_ctor_get(x_712, 0); +lean_inc(x_713); +x_714 = lean_ctor_get(x_712, 1); +lean_inc(x_714); +lean_dec(x_712); +x_715 = lean_ctor_get(x_713, 0); +lean_inc(x_715); +x_716 = lean_ctor_get(x_713, 1); +lean_inc(x_716); +lean_dec(x_713); +x_717 = lean_ctor_get(x_716, 2); +lean_inc(x_717); +lean_inc(x_717); +x_718 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_717, x_715); +if (lean_obj_tag(x_718) == 0) +{ +size_t x_719; size_t x_720; uint8_t x_721; +x_719 = lean_ptr_addr(x_711); +lean_dec(x_711); +x_720 = lean_ptr_addr(x_715); +x_721 = lean_usize_dec_eq(x_719, x_720); +if (x_721 == 0) +{ +lean_object* x_722; lean_object* x_723; lean_object* x_724; lean_object* x_725; lean_object* x_726; lean_object* x_727; lean_object* x_728; lean_object* x_729; +x_722 = lean_unsigned_to_nat(0u); +lean_inc(x_715); +x_723 = lean_array_set(x_3, x_722, x_715); +x_724 = l_Lean_mkAppN(x_2, x_723); +lean_dec(x_723); +x_725 = lean_ctor_get(x_716, 0); +lean_inc(x_725); +x_726 = lean_ctor_get(x_716, 1); +lean_inc(x_726); +lean_dec(x_716); +lean_inc(x_724); +x_727 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_717, x_715, x_724); +x_728 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_728, 0, x_725); +lean_ctor_set(x_728, 1, x_726); +lean_ctor_set(x_728, 2, x_727); +x_729 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_724, x_5, x_728, x_7, x_8, x_9, x_10, x_714); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_777 = !lean_is_exclusive(x_765); -if (x_777 == 0) -{ -return x_765; -} -else -{ -lean_object* x_778; lean_object* x_779; lean_object* x_780; -x_778 = lean_ctor_get(x_765, 0); -x_779 = lean_ctor_get(x_765, 1); -lean_inc(x_779); -lean_inc(x_778); -lean_dec(x_765); -x_780 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_780, 0, x_778); -lean_ctor_set(x_780, 1, x_779); -return x_780; -} -} +return x_729; } else { -lean_object* x_781; lean_object* x_782; -x_781 = lean_array_fget(x_3, x_761); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_5); -lean_inc(x_781); -x_782 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_781, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_782) == 0) -{ -lean_object* x_783; lean_object* x_784; lean_object* x_785; lean_object* x_786; size_t x_787; size_t x_788; uint8_t x_789; -x_783 = lean_ctor_get(x_782, 0); -lean_inc(x_783); -x_784 = lean_ctor_get(x_782, 1); -lean_inc(x_784); -lean_dec(x_782); -x_785 = lean_ctor_get(x_783, 0); -lean_inc(x_785); -x_786 = lean_ctor_get(x_783, 1); -lean_inc(x_786); -lean_dec(x_783); -x_787 = lean_ptr_addr(x_781); -lean_dec(x_781); -x_788 = lean_ptr_addr(x_785); -x_789 = lean_usize_dec_eq(x_787, x_788); -if (x_789 == 0) -{ -lean_object* x_790; -lean_dec(x_785); +lean_object* x_730; lean_object* x_731; lean_object* x_732; lean_object* x_733; lean_object* x_734; lean_dec(x_3); lean_dec(x_2); +x_730 = lean_ctor_get(x_716, 0); +lean_inc(x_730); +x_731 = lean_ctor_get(x_716, 1); +lean_inc(x_731); +lean_dec(x_716); +lean_inc(x_1); +x_732 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_717, x_715, x_1); +x_733 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_733, 0, x_730); +lean_ctor_set(x_733, 1, x_731); +lean_ctor_set(x_733, 2, x_732); lean_inc(x_1); -x_790 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_786, x_7, x_8, x_9, x_10, x_784); +x_734 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_733, x_7, x_8, x_9, x_10, x_714); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_5); -return x_790; +return x_734; +} } else { -lean_object* x_791; lean_object* x_792; lean_object* x_793; -x_791 = lean_array_set(x_3, x_761, x_785); -x_792 = l_Lean_mkAppN(x_2, x_791); -lean_dec(x_791); -x_793 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_792, x_5, x_786, x_7, x_8, x_9, x_10, x_784); +lean_object* x_735; lean_object* x_736; +lean_dec(x_717); +lean_dec(x_715); +lean_dec(x_711); +lean_dec(x_3); +lean_dec(x_2); +x_735 = lean_ctor_get(x_718, 0); +lean_inc(x_735); +lean_dec(x_718); +x_736 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_735, x_5, x_716, x_7, x_8, x_9, x_10, x_714); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_5); -return x_793; +return x_736; } } else { -uint8_t x_794; -lean_dec(x_781); +uint8_t x_737; +lean_dec(x_711); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -10794,127 +10157,125 @@ lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_794 = !lean_is_exclusive(x_782); -if (x_794 == 0) +x_737 = !lean_is_exclusive(x_712); +if (x_737 == 0) { -return x_782; +return x_712; } else { -lean_object* x_795; lean_object* x_796; lean_object* x_797; -x_795 = lean_ctor_get(x_782, 0); -x_796 = lean_ctor_get(x_782, 1); -lean_inc(x_796); -lean_inc(x_795); -lean_dec(x_782); -x_797 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_797, 0, x_795); -lean_ctor_set(x_797, 1, x_796); -return x_797; -} -} +lean_object* x_738; lean_object* x_739; lean_object* x_740; +x_738 = lean_ctor_get(x_712, 0); +x_739 = lean_ctor_get(x_712, 1); +lean_inc(x_739); +lean_inc(x_738); +lean_dec(x_712); +x_740 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_740, 0, x_738); +lean_ctor_set(x_740, 1, x_739); +return x_740; } } } -block_753: +block_775: { -lean_object* x_721; -lean_dec(x_720); +lean_object* x_743; +lean_dec(x_742); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_2); -x_721 = l_Lean_Meta_getFunInfo(x_2, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_721) == 0) -{ -lean_object* x_722; lean_object* x_723; lean_object* x_724; lean_object* x_725; lean_object* x_726; lean_object* x_727; lean_object* x_728; uint8_t x_729; lean_object* x_730; lean_object* x_731; lean_object* x_732; -x_722 = lean_ctor_get(x_721, 0); -lean_inc(x_722); -x_723 = lean_ctor_get(x_721, 1); -lean_inc(x_723); -lean_dec(x_721); -x_724 = lean_ctor_get(x_722, 0); -lean_inc(x_724); -lean_dec(x_722); -x_725 = lean_array_get_size(x_3); -x_726 = lean_unsigned_to_nat(0u); -x_727 = lean_unsigned_to_nat(1u); -x_728 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_728, 0, x_726); -lean_ctor_set(x_728, 1, x_725); -lean_ctor_set(x_728, 2, x_727); -x_729 = 0; -x_730 = lean_box(x_729); -x_731 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_731, 0, x_3); -lean_ctor_set(x_731, 1, x_730); +x_743 = l_Lean_Meta_getFunInfo(x_2, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_743) == 0) +{ +lean_object* x_744; lean_object* x_745; lean_object* x_746; lean_object* x_747; lean_object* x_748; lean_object* x_749; lean_object* x_750; uint8_t x_751; lean_object* x_752; lean_object* x_753; lean_object* x_754; +x_744 = lean_ctor_get(x_743, 0); +lean_inc(x_744); +x_745 = lean_ctor_get(x_743, 1); +lean_inc(x_745); +lean_dec(x_743); +x_746 = lean_ctor_get(x_744, 0); +lean_inc(x_746); +lean_dec(x_744); +x_747 = lean_array_get_size(x_3); +x_748 = lean_unsigned_to_nat(0u); +x_749 = lean_unsigned_to_nat(1u); +x_750 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_750, 0, x_748); +lean_ctor_set(x_750, 1, x_747); +lean_ctor_set(x_750, 2, x_749); +x_751 = 0; +x_752 = lean_box(x_751); +x_753 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_753, 0, x_3); +lean_ctor_set(x_753, 1, x_752); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_5); lean_inc(x_2); -x_732 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8(x_2, x_724, x_728, x_728, x_731, x_726, lean_box(0), lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10, x_723); -lean_dec(x_728); -lean_dec(x_724); -if (lean_obj_tag(x_732) == 0) -{ -lean_object* x_733; lean_object* x_734; lean_object* x_735; uint8_t x_736; -x_733 = lean_ctor_get(x_732, 0); -lean_inc(x_733); -x_734 = lean_ctor_get(x_733, 0); -lean_inc(x_734); -x_735 = lean_ctor_get(x_734, 1); -lean_inc(x_735); -x_736 = lean_unbox(x_735); -lean_dec(x_735); -if (x_736 == 0) +x_754 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8(x_2, x_746, x_750, x_750, x_753, x_748, lean_box(0), lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10, x_745); +lean_dec(x_750); +lean_dec(x_746); +if (lean_obj_tag(x_754) == 0) +{ +lean_object* x_755; lean_object* x_756; lean_object* x_757; uint8_t x_758; +x_755 = lean_ctor_get(x_754, 0); +lean_inc(x_755); +x_756 = lean_ctor_get(x_755, 0); +lean_inc(x_756); +x_757 = lean_ctor_get(x_756, 1); +lean_inc(x_757); +x_758 = lean_unbox(x_757); +lean_dec(x_757); +if (x_758 == 0) { -lean_object* x_737; lean_object* x_738; lean_object* x_739; -lean_dec(x_734); +lean_object* x_759; lean_object* x_760; lean_object* x_761; +lean_dec(x_756); lean_dec(x_2); -x_737 = lean_ctor_get(x_732, 1); -lean_inc(x_737); -lean_dec(x_732); -x_738 = lean_ctor_get(x_733, 1); -lean_inc(x_738); -lean_dec(x_733); +x_759 = lean_ctor_get(x_754, 1); +lean_inc(x_759); +lean_dec(x_754); +x_760 = lean_ctor_get(x_755, 1); +lean_inc(x_760); +lean_dec(x_755); lean_inc(x_1); -x_739 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_738, x_7, x_8, x_9, x_10, x_737); +x_761 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_760, x_7, x_8, x_9, x_10, x_759); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_5); -return x_739; +return x_761; } else { -lean_object* x_740; lean_object* x_741; lean_object* x_742; lean_object* x_743; lean_object* x_744; -x_740 = lean_ctor_get(x_732, 1); -lean_inc(x_740); -lean_dec(x_732); -x_741 = lean_ctor_get(x_733, 1); -lean_inc(x_741); -lean_dec(x_733); -x_742 = lean_ctor_get(x_734, 0); -lean_inc(x_742); -lean_dec(x_734); -x_743 = l_Lean_mkAppN(x_2, x_742); -lean_dec(x_742); -x_744 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_743, x_5, x_741, x_7, x_8, x_9, x_10, x_740); +lean_object* x_762; lean_object* x_763; lean_object* x_764; lean_object* x_765; lean_object* x_766; +x_762 = lean_ctor_get(x_754, 1); +lean_inc(x_762); +lean_dec(x_754); +x_763 = lean_ctor_get(x_755, 1); +lean_inc(x_763); +lean_dec(x_755); +x_764 = lean_ctor_get(x_756, 0); +lean_inc(x_764); +lean_dec(x_756); +x_765 = l_Lean_mkAppN(x_2, x_764); +lean_dec(x_764); +x_766 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_765, x_5, x_763, x_7, x_8, x_9, x_10, x_762); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_5); -return x_744; +return x_766; } } else { -uint8_t x_745; +uint8_t x_767; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -10922,29 +10283,29 @@ lean_dec(x_7); lean_dec(x_5); lean_dec(x_2); lean_dec(x_1); -x_745 = !lean_is_exclusive(x_732); -if (x_745 == 0) +x_767 = !lean_is_exclusive(x_754); +if (x_767 == 0) { -return x_732; +return x_754; } else { -lean_object* x_746; lean_object* x_747; lean_object* x_748; -x_746 = lean_ctor_get(x_732, 0); -x_747 = lean_ctor_get(x_732, 1); -lean_inc(x_747); -lean_inc(x_746); -lean_dec(x_732); -x_748 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_748, 0, x_746); -lean_ctor_set(x_748, 1, x_747); -return x_748; +lean_object* x_768; lean_object* x_769; lean_object* x_770; +x_768 = lean_ctor_get(x_754, 0); +x_769 = lean_ctor_get(x_754, 1); +lean_inc(x_769); +lean_inc(x_768); +lean_dec(x_754); +x_770 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_770, 0, x_768); +lean_ctor_set(x_770, 1, x_769); +return x_770; } } } else { -uint8_t x_749; +uint8_t x_771; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -10954,213 +10315,188 @@ lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_749 = !lean_is_exclusive(x_721); -if (x_749 == 0) +x_771 = !lean_is_exclusive(x_743); +if (x_771 == 0) { -return x_721; +return x_743; } else { -lean_object* x_750; lean_object* x_751; lean_object* x_752; -x_750 = lean_ctor_get(x_721, 0); -x_751 = lean_ctor_get(x_721, 1); -lean_inc(x_751); -lean_inc(x_750); -lean_dec(x_721); -x_752 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_752, 0, x_750); -lean_ctor_set(x_752, 1, x_751); -return x_752; +lean_object* x_772; lean_object* x_773; lean_object* x_774; +x_772 = lean_ctor_get(x_743, 0); +x_773 = lean_ctor_get(x_743, 1); +lean_inc(x_773); +lean_inc(x_772); +lean_dec(x_743); +x_774 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_774, 0, x_772); +lean_ctor_set(x_774, 1, x_773); +return x_774; } } } } default: { -lean_object* x_798; lean_object* x_832; uint8_t x_833; +lean_object* x_788; lean_object* x_819; lean_object* x_853; uint8_t x_854; lean_dec(x_4); -x_832 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__4; -x_833 = l_Lean_Expr_isConstOf(x_2, x_832); -if (x_833 == 0) +x_853 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__4; +x_854 = l_Lean_Expr_isConstOf(x_2, x_853); +if (x_854 == 0) { -lean_object* x_834; -x_834 = lean_box(0); -x_798 = x_834; -goto block_831; +lean_object* x_855; +x_855 = lean_box(0); +x_819 = x_855; +goto block_852; } else { -lean_object* x_835; lean_object* x_836; uint8_t x_837; -x_835 = lean_array_get_size(x_3); -x_836 = lean_unsigned_to_nat(2u); -x_837 = lean_nat_dec_eq(x_835, x_836); -if (x_837 == 0) +lean_object* x_856; lean_object* x_857; uint8_t x_858; +x_856 = lean_array_get_size(x_3); +x_857 = lean_unsigned_to_nat(2u); +x_858 = lean_nat_dec_eq(x_856, x_857); +if (x_858 == 0) { -lean_object* x_838; -lean_dec(x_835); -x_838 = lean_box(0); -x_798 = x_838; -goto block_831; +lean_object* x_859; +lean_dec(x_856); +x_859 = lean_box(0); +x_819 = x_859; +goto block_852; } else { -lean_object* x_839; uint8_t x_840; -x_839 = lean_unsigned_to_nat(0u); -x_840 = lean_nat_dec_lt(x_839, x_835); -lean_dec(x_835); -if (x_840 == 0) +lean_object* x_860; uint8_t x_861; +x_860 = lean_unsigned_to_nat(0u); +x_861 = lean_nat_dec_lt(x_860, x_856); +lean_dec(x_856); +if (x_861 == 0) { -lean_object* x_841; lean_object* x_842; lean_object* x_843; -x_841 = l_Lean_instInhabitedExpr; -x_842 = l_outOfBounds___rarg(x_841); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_5); -lean_inc(x_842); -x_843 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_842, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_843) == 0) -{ -lean_object* x_844; lean_object* x_845; lean_object* x_846; lean_object* x_847; size_t x_848; size_t x_849; uint8_t x_850; -x_844 = lean_ctor_get(x_843, 0); -lean_inc(x_844); -x_845 = lean_ctor_get(x_843, 1); -lean_inc(x_845); -lean_dec(x_843); -x_846 = lean_ctor_get(x_844, 0); -lean_inc(x_846); -x_847 = lean_ctor_get(x_844, 1); -lean_inc(x_847); -lean_dec(x_844); -x_848 = lean_ptr_addr(x_842); -lean_dec(x_842); -x_849 = lean_ptr_addr(x_846); -x_850 = lean_usize_dec_eq(x_848, x_849); -if (x_850 == 0) -{ -lean_object* x_851; -lean_dec(x_846); -lean_dec(x_3); -lean_dec(x_2); -lean_inc(x_1); -x_851 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_847, x_7, x_8, x_9, x_10, x_845); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_851; +lean_object* x_862; lean_object* x_863; +x_862 = l_Lean_instInhabitedExpr; +x_863 = l_outOfBounds___rarg(x_862); +x_788 = x_863; +goto block_818; } else { -lean_object* x_852; lean_object* x_853; lean_object* x_854; -x_852 = lean_array_set(x_3, x_839, x_846); -x_853 = l_Lean_mkAppN(x_2, x_852); -lean_dec(x_852); -x_854 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_853, x_5, x_847, x_7, x_8, x_9, x_10, x_845); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_854; +lean_object* x_864; +x_864 = lean_array_fget(x_3, x_860); +x_788 = x_864; +goto block_818; } } -else +} +block_818: { -uint8_t x_855; -lean_dec(x_842); +lean_object* x_789; +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_5); +lean_inc(x_788); +x_789 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_788, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_789) == 0) +{ +lean_object* x_790; lean_object* x_791; lean_object* x_792; lean_object* x_793; lean_object* x_794; lean_object* x_795; +x_790 = lean_ctor_get(x_789, 0); +lean_inc(x_790); +x_791 = lean_ctor_get(x_789, 1); +lean_inc(x_791); +lean_dec(x_789); +x_792 = lean_ctor_get(x_790, 0); +lean_inc(x_792); +x_793 = lean_ctor_get(x_790, 1); +lean_inc(x_793); +lean_dec(x_790); +x_794 = lean_ctor_get(x_793, 2); +lean_inc(x_794); +lean_inc(x_794); +x_795 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_794, x_792); +if (lean_obj_tag(x_795) == 0) +{ +size_t x_796; size_t x_797; uint8_t x_798; +x_796 = lean_ptr_addr(x_788); +lean_dec(x_788); +x_797 = lean_ptr_addr(x_792); +x_798 = lean_usize_dec_eq(x_796, x_797); +if (x_798 == 0) +{ +lean_object* x_799; lean_object* x_800; lean_object* x_801; lean_object* x_802; lean_object* x_803; lean_object* x_804; lean_object* x_805; lean_object* x_806; +x_799 = lean_unsigned_to_nat(0u); +lean_inc(x_792); +x_800 = lean_array_set(x_3, x_799, x_792); +x_801 = l_Lean_mkAppN(x_2, x_800); +lean_dec(x_800); +x_802 = lean_ctor_get(x_793, 0); +lean_inc(x_802); +x_803 = lean_ctor_get(x_793, 1); +lean_inc(x_803); +lean_dec(x_793); +lean_inc(x_801); +x_804 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_794, x_792, x_801); +x_805 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_805, 0, x_802); +lean_ctor_set(x_805, 1, x_803); +lean_ctor_set(x_805, 2, x_804); +x_806 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_801, x_5, x_805, x_7, x_8, x_9, x_10, x_791); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_855 = !lean_is_exclusive(x_843); -if (x_855 == 0) -{ -return x_843; -} -else -{ -lean_object* x_856; lean_object* x_857; lean_object* x_858; -x_856 = lean_ctor_get(x_843, 0); -x_857 = lean_ctor_get(x_843, 1); -lean_inc(x_857); -lean_inc(x_856); -lean_dec(x_843); -x_858 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_858, 0, x_856); -lean_ctor_set(x_858, 1, x_857); -return x_858; -} -} +return x_806; } else { -lean_object* x_859; lean_object* x_860; -x_859 = lean_array_fget(x_3, x_839); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_5); -lean_inc(x_859); -x_860 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_859, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_860) == 0) -{ -lean_object* x_861; lean_object* x_862; lean_object* x_863; lean_object* x_864; size_t x_865; size_t x_866; uint8_t x_867; -x_861 = lean_ctor_get(x_860, 0); -lean_inc(x_861); -x_862 = lean_ctor_get(x_860, 1); -lean_inc(x_862); -lean_dec(x_860); -x_863 = lean_ctor_get(x_861, 0); -lean_inc(x_863); -x_864 = lean_ctor_get(x_861, 1); -lean_inc(x_864); -lean_dec(x_861); -x_865 = lean_ptr_addr(x_859); -lean_dec(x_859); -x_866 = lean_ptr_addr(x_863); -x_867 = lean_usize_dec_eq(x_865, x_866); -if (x_867 == 0) -{ -lean_object* x_868; -lean_dec(x_863); +lean_object* x_807; lean_object* x_808; lean_object* x_809; lean_object* x_810; lean_object* x_811; lean_dec(x_3); lean_dec(x_2); +x_807 = lean_ctor_get(x_793, 0); +lean_inc(x_807); +x_808 = lean_ctor_get(x_793, 1); +lean_inc(x_808); +lean_dec(x_793); lean_inc(x_1); -x_868 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_864, x_7, x_8, x_9, x_10, x_862); +x_809 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_794, x_792, x_1); +x_810 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_810, 0, x_807); +lean_ctor_set(x_810, 1, x_808); +lean_ctor_set(x_810, 2, x_809); +lean_inc(x_1); +x_811 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_810, x_7, x_8, x_9, x_10, x_791); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_5); -return x_868; +return x_811; +} } else { -lean_object* x_869; lean_object* x_870; lean_object* x_871; -x_869 = lean_array_set(x_3, x_839, x_863); -x_870 = l_Lean_mkAppN(x_2, x_869); -lean_dec(x_869); -x_871 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_870, x_5, x_864, x_7, x_8, x_9, x_10, x_862); +lean_object* x_812; lean_object* x_813; +lean_dec(x_794); +lean_dec(x_792); +lean_dec(x_788); +lean_dec(x_3); +lean_dec(x_2); +x_812 = lean_ctor_get(x_795, 0); +lean_inc(x_812); +lean_dec(x_795); +x_813 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_812, x_5, x_793, x_7, x_8, x_9, x_10, x_791); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_5); -return x_871; +return x_813; } } else { -uint8_t x_872; -lean_dec(x_859); +uint8_t x_814; +lean_dec(x_788); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -11169,127 +10505,125 @@ lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_872 = !lean_is_exclusive(x_860); -if (x_872 == 0) +x_814 = !lean_is_exclusive(x_789); +if (x_814 == 0) { -return x_860; +return x_789; } else { -lean_object* x_873; lean_object* x_874; lean_object* x_875; -x_873 = lean_ctor_get(x_860, 0); -x_874 = lean_ctor_get(x_860, 1); -lean_inc(x_874); -lean_inc(x_873); -lean_dec(x_860); -x_875 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_875, 0, x_873); -lean_ctor_set(x_875, 1, x_874); -return x_875; -} -} +lean_object* x_815; lean_object* x_816; lean_object* x_817; +x_815 = lean_ctor_get(x_789, 0); +x_816 = lean_ctor_get(x_789, 1); +lean_inc(x_816); +lean_inc(x_815); +lean_dec(x_789); +x_817 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_817, 0, x_815); +lean_ctor_set(x_817, 1, x_816); +return x_817; } } } -block_831: +block_852: { -lean_object* x_799; -lean_dec(x_798); +lean_object* x_820; +lean_dec(x_819); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_2); -x_799 = l_Lean_Meta_getFunInfo(x_2, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_799) == 0) -{ -lean_object* x_800; lean_object* x_801; lean_object* x_802; lean_object* x_803; lean_object* x_804; lean_object* x_805; lean_object* x_806; uint8_t x_807; lean_object* x_808; lean_object* x_809; lean_object* x_810; -x_800 = lean_ctor_get(x_799, 0); -lean_inc(x_800); -x_801 = lean_ctor_get(x_799, 1); -lean_inc(x_801); -lean_dec(x_799); -x_802 = lean_ctor_get(x_800, 0); -lean_inc(x_802); -lean_dec(x_800); -x_803 = lean_array_get_size(x_3); -x_804 = lean_unsigned_to_nat(0u); -x_805 = lean_unsigned_to_nat(1u); -x_806 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_806, 0, x_804); -lean_ctor_set(x_806, 1, x_803); -lean_ctor_set(x_806, 2, x_805); -x_807 = 0; -x_808 = lean_box(x_807); -x_809 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_809, 0, x_3); -lean_ctor_set(x_809, 1, x_808); +x_820 = l_Lean_Meta_getFunInfo(x_2, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_820) == 0) +{ +lean_object* x_821; lean_object* x_822; lean_object* x_823; lean_object* x_824; lean_object* x_825; lean_object* x_826; lean_object* x_827; uint8_t x_828; lean_object* x_829; lean_object* x_830; lean_object* x_831; +x_821 = lean_ctor_get(x_820, 0); +lean_inc(x_821); +x_822 = lean_ctor_get(x_820, 1); +lean_inc(x_822); +lean_dec(x_820); +x_823 = lean_ctor_get(x_821, 0); +lean_inc(x_823); +lean_dec(x_821); +x_824 = lean_array_get_size(x_3); +x_825 = lean_unsigned_to_nat(0u); +x_826 = lean_unsigned_to_nat(1u); +x_827 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_827, 0, x_825); +lean_ctor_set(x_827, 1, x_824); +lean_ctor_set(x_827, 2, x_826); +x_828 = 0; +x_829 = lean_box(x_828); +x_830 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_830, 0, x_3); +lean_ctor_set(x_830, 1, x_829); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_5); lean_inc(x_2); -x_810 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8(x_2, x_802, x_806, x_806, x_809, x_804, lean_box(0), lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10, x_801); -lean_dec(x_806); -lean_dec(x_802); -if (lean_obj_tag(x_810) == 0) -{ -lean_object* x_811; lean_object* x_812; lean_object* x_813; uint8_t x_814; -x_811 = lean_ctor_get(x_810, 0); -lean_inc(x_811); -x_812 = lean_ctor_get(x_811, 0); -lean_inc(x_812); -x_813 = lean_ctor_get(x_812, 1); -lean_inc(x_813); -x_814 = lean_unbox(x_813); -lean_dec(x_813); -if (x_814 == 0) -{ -lean_object* x_815; lean_object* x_816; lean_object* x_817; -lean_dec(x_812); +x_831 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8(x_2, x_823, x_827, x_827, x_830, x_825, lean_box(0), lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10, x_822); +lean_dec(x_827); +lean_dec(x_823); +if (lean_obj_tag(x_831) == 0) +{ +lean_object* x_832; lean_object* x_833; lean_object* x_834; uint8_t x_835; +x_832 = lean_ctor_get(x_831, 0); +lean_inc(x_832); +x_833 = lean_ctor_get(x_832, 0); +lean_inc(x_833); +x_834 = lean_ctor_get(x_833, 1); +lean_inc(x_834); +x_835 = lean_unbox(x_834); +lean_dec(x_834); +if (x_835 == 0) +{ +lean_object* x_836; lean_object* x_837; lean_object* x_838; +lean_dec(x_833); lean_dec(x_2); -x_815 = lean_ctor_get(x_810, 1); -lean_inc(x_815); -lean_dec(x_810); -x_816 = lean_ctor_get(x_811, 1); -lean_inc(x_816); -lean_dec(x_811); +x_836 = lean_ctor_get(x_831, 1); +lean_inc(x_836); +lean_dec(x_831); +x_837 = lean_ctor_get(x_832, 1); +lean_inc(x_837); +lean_dec(x_832); lean_inc(x_1); -x_817 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_816, x_7, x_8, x_9, x_10, x_815); +x_838 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_837, x_7, x_8, x_9, x_10, x_836); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_5); -return x_817; +return x_838; } else { -lean_object* x_818; lean_object* x_819; lean_object* x_820; lean_object* x_821; lean_object* x_822; -x_818 = lean_ctor_get(x_810, 1); -lean_inc(x_818); -lean_dec(x_810); -x_819 = lean_ctor_get(x_811, 1); -lean_inc(x_819); -lean_dec(x_811); -x_820 = lean_ctor_get(x_812, 0); -lean_inc(x_820); -lean_dec(x_812); -x_821 = l_Lean_mkAppN(x_2, x_820); -lean_dec(x_820); -x_822 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_821, x_5, x_819, x_7, x_8, x_9, x_10, x_818); +lean_object* x_839; lean_object* x_840; lean_object* x_841; lean_object* x_842; lean_object* x_843; +x_839 = lean_ctor_get(x_831, 1); +lean_inc(x_839); +lean_dec(x_831); +x_840 = lean_ctor_get(x_832, 1); +lean_inc(x_840); +lean_dec(x_832); +x_841 = lean_ctor_get(x_833, 0); +lean_inc(x_841); +lean_dec(x_833); +x_842 = l_Lean_mkAppN(x_2, x_841); +lean_dec(x_841); +x_843 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_842, x_5, x_840, x_7, x_8, x_9, x_10, x_839); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_5); -return x_822; +return x_843; } } else { -uint8_t x_823; +uint8_t x_844; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -11297,29 +10631,29 @@ lean_dec(x_7); lean_dec(x_5); lean_dec(x_2); lean_dec(x_1); -x_823 = !lean_is_exclusive(x_810); -if (x_823 == 0) +x_844 = !lean_is_exclusive(x_831); +if (x_844 == 0) { -return x_810; +return x_831; } else { -lean_object* x_824; lean_object* x_825; lean_object* x_826; -x_824 = lean_ctor_get(x_810, 0); -x_825 = lean_ctor_get(x_810, 1); -lean_inc(x_825); -lean_inc(x_824); -lean_dec(x_810); -x_826 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_826, 0, x_824); -lean_ctor_set(x_826, 1, x_825); -return x_826; +lean_object* x_845; lean_object* x_846; lean_object* x_847; +x_845 = lean_ctor_get(x_831, 0); +x_846 = lean_ctor_get(x_831, 1); +lean_inc(x_846); +lean_inc(x_845); +lean_dec(x_831); +x_847 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_847, 0, x_845); +lean_ctor_set(x_847, 1, x_846); +return x_847; } } } else { -uint8_t x_827; +uint8_t x_848; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -11329,23 +10663,23 @@ lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_827 = !lean_is_exclusive(x_799); -if (x_827 == 0) +x_848 = !lean_is_exclusive(x_820); +if (x_848 == 0) { -return x_799; +return x_820; } else { -lean_object* x_828; lean_object* x_829; lean_object* x_830; -x_828 = lean_ctor_get(x_799, 0); -x_829 = lean_ctor_get(x_799, 1); -lean_inc(x_829); -lean_inc(x_828); -lean_dec(x_799); -x_830 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_830, 0, x_828); -lean_ctor_set(x_830, 1, x_829); -return x_830; +lean_object* x_849; lean_object* x_850; lean_object* x_851; +x_849 = lean_ctor_get(x_820, 0); +x_850 = lean_ctor_get(x_820, 1); +lean_inc(x_850); +lean_inc(x_849); +lean_dec(x_820); +x_851 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_851, 0, x_849); +lean_ctor_set(x_851, 1, x_850); +return x_851; } } } @@ -11443,7 +10777,7 @@ static lean_object* _init_l_Lean_Meta_Grind_Canon_canonImpl_visit___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_Grind_Canon_canonImpl_visit___closed__1; x_2 = l_Lean_Meta_Grind_Canon_canonImpl_visit___closed__2; -x_3 = lean_unsigned_to_nat(113u); +x_3 = lean_unsigned_to_nat(115u); x_4 = lean_unsigned_to_nat(18u); x_5 = l_Lean_Meta_Grind_Canon_canonImpl_visit___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); diff --git a/stage0/stdlib/Lean/Meta/Tactic/Grind/Core.c b/stage0/stdlib/Lean/Meta/Tactic/Grind/Core.c index bf0206f87bc8..80d6a7357d90 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Grind/Core.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Grind/Core.c @@ -15,42 +15,39 @@ extern "C" { #endif lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_add_go___closed__3; +static lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__5; +static lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__3; static lean_object* l_Lean_Meta_Grind_add___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___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_Grind_addHEq(lean_object*, lean_object*, lean_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_Grind_Core_0__Lean_Meta_Grind_addEqStep___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___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_shift_right(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_add___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_Meta_Grind_add_go___closed__12; static lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_closeGoalWithValuesEq___closed__5; lean_object* l_Lean_Meta_mkEqTrue(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__7(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_add_go___lambda__4(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__2___closed__1; -static lean_object* l_Lean_Meta_Grind_add_go___closed__11; -static lean_object* l_Lean_Meta_Grind_add_go___closed__13; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_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*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__8___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep(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___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqCore_processTodo___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_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_invertTrans___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_mkAppB(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__2; -static lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__8; -static lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__6; size_t lean_uint64_to_usize(uint64_t); static lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_closeGoalWithValuesEq___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_add___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_add_go___closed__9; lean_object* l_Lean_Meta_Grind_propagateCtor(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_isApp(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_popNextEq_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Meta_Grind_isCongruent(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Core_checkSystem(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_add___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_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__1; static lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__3; static size_t l_Lean_PersistentHashMap_eraseAux___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__2___closed__2; -static lean_object* l_Lean_Meta_Grind_add_go___lambda__4___closed__1; static lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_closeGoalWithValuesEq___closed__11; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint64_t l_Lean_Meta_Grind_congrHash(lean_object*, lean_object*); @@ -59,24 +56,21 @@ lean_object* lean_array_fget(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___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* l_Lean_Meta_Grind_propagateDown(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_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__2___closed__2; +LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_reinsertParents___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_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_updateMT___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_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_closeGoalWithValuesEq(lean_object*, lean_object*, 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_Grind_add_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*, lean_object*, lean_object*); +static lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_reinsertParents___spec__1___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_add_go___closed__8; lean_object* l_Lean_Meta_Grind_setENode(lean_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_cleanupAnnotations(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_updateRoots_loop___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_stringToMessageData(lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__9___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__6(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_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addEq(lean_object*, lean_object*, 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_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__7; -static lean_object* l_Lean_Meta_Grind_add_go___closed__10; LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_updateMT___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_closeGoalWithTrueEqFalse___closed__4; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_add_goEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__7; lean_object* l_Lean_Level_ofNat(lean_object*); lean_object* l_Lean_Expr_appArg_x21(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_add_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*, lean_object*, lean_object*); @@ -86,17 +80,16 @@ static lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_a LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_updateMT(lean_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_Grind_Core_0__Lean_Meta_Grind_updateMT___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_Grind_getRoot(lean_object*, 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_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_reinsertParents___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___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_updateRoots(lean_object*, 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_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__5; -static lean_object* l_Lean_Meta_Grind_add_go___lambda__4___closed__3; +lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MVarId_isAssigned___at_Lean_Meta_Grind_closeGoal___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___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__6; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___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_mkEqMP(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_getFalseExpr___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_closeGoalWithTrueEqFalse___closed__2; -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___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*, lean_object*); -static lean_object* l_Lean_Meta_Grind_add_go___lambda__4___closed__2; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_isUnaryNode___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_reinsertParents___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_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_invertTrans(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -107,11 +100,12 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Gr static lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_closeGoalWithValuesEq___closed__12; LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_eraseAux___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__2(lean_object*, lean_object*, size_t, lean_object*); -static lean_object* l_Lean_Meta_Grind_add_go___closed__6; static lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_closeGoalWithValuesEq___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqCore_processTodo(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_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_closeGoal(lean_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_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofFormat(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_popNextEq_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_popNextEq_x3f___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -122,6 +116,7 @@ lean_object* l_Lean_Meta_Grind_getTrueExpr___rarg(lean_object*, lean_object*, le lean_object* lean_st_ref_get(lean_object*, lean_object*); lean_object* lean_array_pop(lean_object*); static lean_object* l_Lean_Meta_Grind_add___lambda__3___closed__1; +lean_object* l_Lean_Meta_mkHEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_mkEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_add___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* l_Lean_Expr_appFnCleanup(lean_object*, lean_object*); @@ -134,8 +129,7 @@ lean_object* l_Lean_Meta_Grind_checkInvariants(uint8_t, lean_object*, lean_objec LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_resetNewEqs___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_ppState(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_Grind_add_go___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__9___closed__3; -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__8(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_FVarId_getType(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*); static size_t l_Lean_PersistentHashMap_eraseAux___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__2___closed__1; @@ -154,12 +148,14 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Gr static lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqCore_processTodo___closed__1; static lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_closeGoalWithValuesEq___closed__10; lean_object* l_Array_eraseIdx___rarg(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_updateRoots_loop(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_usize_to_nat(size_t); static lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__5; +static lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_reinsertParents___spec__1___closed__2; lean_object* l_Lean_MessageData_ofExpr(lean_object*); lean_object* l_Lean_Meta_Grind_copyParentsTo(lean_object*, lean_object*, 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_Grind_add_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_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addHypothesis(lean_object*, lean_object*, 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_Grind_add___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___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqCore___boxed(lean_object*, lean_object*, lean_object*, 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_Grind_Core_0__Lean_Meta_Grind_closeGoalWithValuesEq___closed__4; @@ -167,15 +163,16 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Grind_add___lambda__2___boxed(lean_object*, lean_object* lean_grind_mk_eq_proof(lean_object*, 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_Grind_Core_0__Lean_Meta_Grind_closeGoalWithValuesEq___closed__7; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__5(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___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___lambda__1___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__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___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_closeGoalWithTrueEqFalse___closed__1; -static lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__9___closed__5; static lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_closeGoalWithValuesEq___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkApp3(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_mkEqFalseProof(lean_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_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__2___boxed(lean_object**); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_updateRoots_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_object*); +static lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__6; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__2(lean_object*, 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*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_closeGoalWithTrueEqFalse(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -197,6 +194,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Gr static lean_object* l_Lean_Meta_Grind_add_go___closed__2; lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_getENode(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__4; 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_Grind_add_goEq___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_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_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___boxed(lean_object**); @@ -208,15 +206,11 @@ size_t lean_usize_sub(size_t, size_t); lean_object* lean_array_mk(lean_object*); lean_object* l_Lean_Meta_Grind_getParentsAndReset(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_isTrue(lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__9___closed__2; -static lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__5; static lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_popNextEq_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* l_Lean_Meta_mkEqFalse(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_add_go___closed__7; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___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_fvar___override(lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__9___closed__4; static lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__4; LEAN_EXPORT lean_object* l_Array_indexOfAux___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); @@ -226,26 +220,25 @@ static lean_object* l_Lean_Meta_Grind_add_go___closed__1; static lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__2___closed__3; LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___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___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__2; -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addHyp(lean_object*, lean_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_Grind_Core_0__Lean_Meta_Grind_addEqStep___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___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__8___closed__2; lean_object* lean_array_get_size(lean_object*); +static lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__7; static lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_closeGoalWithValuesEq___closed__9; lean_object* lean_nat_add(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__6; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__6___boxed(lean_object**); LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___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_Grind_add_go___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*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_closeGoalWithValuesEq___closed__13; static lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_closeGoalWithTrueEqFalse___closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_add_go___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_add_go___closed__4; static lean_object* l_Lean_Meta_Grind_add___lambda__3___closed__2; lean_object* lean_array_get(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__3___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_updateRoots_loop___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* l_Lean_Meta_Grind_isInconsistent(lean_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_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___boxed(lean_object**); -static lean_object* l_Lean_Meta_Grind_add_go___closed__5; +static lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__2; lean_object* l_Lean_Meta_Grind_internalize(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_land(size_t, size_t); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_invertTrans_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*, lean_object*, lean_object*); @@ -335,7 +328,6 @@ lean_inc(x_17); lean_dec(x_14); x_18 = lean_box(0); x_19 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_invertTrans_go___lambda__1(x_15, x_3, x_4, x_2, x_1, x_18, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_17); -lean_dec(x_1); return x_19; } else @@ -369,7 +361,6 @@ lean_inc(x_28); lean_dec(x_26); x_29 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_invertTrans_go___lambda__1(x_15, x_3, x_4, x_2, x_1, x_27, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_28); lean_dec(x_27); -lean_dec(x_1); return x_29; } else @@ -414,7 +405,6 @@ lean_inc(x_37); lean_dec(x_35); x_38 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_invertTrans_go___lambda__1(x_15, x_3, x_4, x_2, x_1, x_36, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_37); lean_dec(x_36); -lean_dec(x_1); return x_38; } else @@ -470,7 +460,6 @@ lean_inc(x_49); lean_dec(x_47); x_50 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_invertTrans_go___lambda__1(x_15, x_3, x_4, x_2, x_1, x_48, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_49); lean_dec(x_48); -lean_dec(x_1); return x_50; } else @@ -517,7 +506,6 @@ lean_inc(x_58); lean_dec(x_56); x_59 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_invertTrans_go___lambda__1(x_15, x_3, x_4, x_2, x_1, x_57, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_58); lean_dec(x_57); -lean_dec(x_1); return x_59; } else @@ -595,7 +583,6 @@ lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); return x_17; } } @@ -1030,6 +1017,226 @@ x_7 = l_Lean_PersistentHashMap_eraseAux___at___private_Lean_Meta_Tactic_Grind_Co return x_7; } } +static lean_object* _init_l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___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_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; +x_12 = lean_st_ref_take(x_3, x_11); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_12, 1); +lean_inc(x_14); +lean_dec(x_12); +x_15 = lean_ctor_get(x_13, 0); +lean_inc(x_15); +x_16 = lean_ctor_get(x_13, 1); +lean_inc(x_16); +x_17 = lean_ctor_get(x_13, 2); +lean_inc(x_17); +x_18 = lean_ctor_get(x_13, 3); +lean_inc(x_18); +x_19 = lean_ctor_get(x_13, 4); +lean_inc(x_19); +x_20 = lean_ctor_get(x_13, 5); +lean_inc(x_20); +x_21 = lean_ctor_get_uint8(x_13, sizeof(void*)*14); +x_22 = lean_ctor_get(x_13, 6); +lean_inc(x_22); +x_23 = lean_ctor_get(x_13, 7); +lean_inc(x_23); +x_24 = lean_ctor_get(x_13, 8); +lean_inc(x_24); +x_25 = lean_ctor_get(x_13, 9); +lean_inc(x_25); +x_26 = lean_ctor_get(x_13, 10); +lean_inc(x_26); +x_27 = lean_ctor_get(x_13, 11); +lean_inc(x_27); +x_28 = lean_ctor_get(x_13, 12); +lean_inc(x_28); +x_29 = lean_ctor_get(x_13, 13); +lean_inc(x_29); +lean_inc(x_13); +x_30 = l_Lean_PersistentHashMap_erase___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__1(x_13, x_18, x_1); +x_31 = !lean_is_exclusive(x_13); +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; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; uint8_t x_47; +x_32 = lean_ctor_get(x_13, 13); +lean_dec(x_32); +x_33 = lean_ctor_get(x_13, 12); +lean_dec(x_33); +x_34 = lean_ctor_get(x_13, 11); +lean_dec(x_34); +x_35 = lean_ctor_get(x_13, 10); +lean_dec(x_35); +x_36 = lean_ctor_get(x_13, 9); +lean_dec(x_36); +x_37 = lean_ctor_get(x_13, 8); +lean_dec(x_37); +x_38 = lean_ctor_get(x_13, 7); +lean_dec(x_38); +x_39 = lean_ctor_get(x_13, 6); +lean_dec(x_39); +x_40 = lean_ctor_get(x_13, 5); +lean_dec(x_40); +x_41 = lean_ctor_get(x_13, 4); +lean_dec(x_41); +x_42 = lean_ctor_get(x_13, 3); +lean_dec(x_42); +x_43 = lean_ctor_get(x_13, 2); +lean_dec(x_43); +x_44 = lean_ctor_get(x_13, 1); +lean_dec(x_44); +x_45 = lean_ctor_get(x_13, 0); +lean_dec(x_45); +lean_ctor_set(x_13, 3, x_30); +x_46 = lean_st_ref_set(x_3, x_13, x_14); +x_47 = !lean_is_exclusive(x_46); +if (x_47 == 0) +{ +lean_object* x_48; lean_object* x_49; +x_48 = lean_ctor_get(x_46, 0); +lean_dec(x_48); +x_49 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___lambda__1___closed__1; +lean_ctor_set(x_46, 0, x_49); +return x_46; +} +else +{ +lean_object* x_50; lean_object* x_51; lean_object* x_52; +x_50 = lean_ctor_get(x_46, 1); +lean_inc(x_50); +lean_dec(x_46); +x_51 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___lambda__1___closed__1; +x_52 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_52, 0, x_51); +lean_ctor_set(x_52, 1, x_50); +return x_52; +} +} +else +{ +lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; +lean_dec(x_13); +x_53 = lean_alloc_ctor(0, 14, 1); +lean_ctor_set(x_53, 0, x_15); +lean_ctor_set(x_53, 1, x_16); +lean_ctor_set(x_53, 2, x_17); +lean_ctor_set(x_53, 3, x_30); +lean_ctor_set(x_53, 4, x_19); +lean_ctor_set(x_53, 5, x_20); +lean_ctor_set(x_53, 6, x_22); +lean_ctor_set(x_53, 7, x_23); +lean_ctor_set(x_53, 8, x_24); +lean_ctor_set(x_53, 9, x_25); +lean_ctor_set(x_53, 10, x_26); +lean_ctor_set(x_53, 11, x_27); +lean_ctor_set(x_53, 12, x_28); +lean_ctor_set(x_53, 13, x_29); +lean_ctor_set_uint8(x_53, sizeof(void*)*14, x_21); +x_54 = lean_st_ref_set(x_3, x_53, x_14); +x_55 = lean_ctor_get(x_54, 1); +lean_inc(x_55); +if (lean_is_exclusive(x_54)) { + lean_ctor_release(x_54, 0); + lean_ctor_release(x_54, 1); + x_56 = x_54; +} else { + lean_dec_ref(x_54); + x_56 = lean_box(0); +} +x_57 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___lambda__1___closed__1; +if (lean_is_scalar(x_56)) { + x_58 = lean_alloc_ctor(0, 2, 0); +} else { + x_58 = x_56; +} +lean_ctor_set(x_58, 0, x_57); +lean_ctor_set(x_58, 1, x_55); +return x_58; +} +} +} +static lean_object* _init_l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("grind", 5, 5); +return x_1; +} +} +static lean_object* _init_l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("debug", 5, 5); +return x_1; +} +} +static lean_object* _init_l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("parent", 6, 6); +return x_1; +} +} +static lean_object* _init_l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__1; +x_2 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__2; +x_3 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__3; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("remove: ", 8, 8); +return x_1; +} +} +static lean_object* _init_l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__5; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("", 0, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__7; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { @@ -1045,7 +1252,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; lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; uint8_t x_32; +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; x_14 = lean_ctor_get(x_1, 0); lean_inc(x_14); x_15 = lean_ctor_get(x_1, 1); @@ -1054,92 +1261,207 @@ x_16 = lean_ctor_get(x_1, 3); lean_inc(x_16); lean_dec(x_1); x_17 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4(x_14, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -x_18 = lean_ctor_get(x_17, 1); -lean_inc(x_18); -lean_dec(x_17); -x_19 = lean_st_ref_take(x_3, x_18); -x_20 = lean_ctor_get(x_19, 0); -lean_inc(x_20); -x_21 = lean_ctor_get(x_19, 1); -lean_inc(x_21); -lean_dec(x_19); -x_22 = lean_ctor_get(x_20, 0); -lean_inc(x_22); -x_23 = lean_ctor_get(x_20, 1); +x_18 = !lean_is_exclusive(x_17); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_19 = lean_ctor_get(x_17, 1); +x_20 = lean_ctor_get(x_17, 0); +lean_dec(x_20); +x_21 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__4; +x_22 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_21, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_19); +x_23 = lean_ctor_get(x_22, 0); lean_inc(x_23); -x_24 = lean_ctor_get(x_20, 2); -lean_inc(x_24); -x_25 = lean_ctor_get(x_20, 3); +x_24 = lean_unbox(x_23); +lean_dec(x_23); +if (x_24 == 0) +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +lean_free_object(x_17); +x_25 = lean_ctor_get(x_22, 1); lean_inc(x_25); -x_26 = lean_ctor_get(x_20, 4); -lean_inc(x_26); -x_27 = lean_ctor_get(x_20, 5); -lean_inc(x_27); -x_28 = lean_ctor_get_uint8(x_20, sizeof(void*)*8); -x_29 = lean_ctor_get(x_20, 6); +lean_dec(x_22); +x_26 = lean_box(0); +x_27 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___lambda__1(x_15, x_26, x_3, x_4, 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_ctor_get(x_27, 1); lean_inc(x_29); -x_30 = lean_ctor_get(x_20, 7); +lean_dec(x_27); +x_30 = lean_ctor_get(x_28, 0); lean_inc(x_30); -lean_inc(x_20); -x_31 = l_Lean_PersistentHashMap_erase___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__1(x_20, x_25, x_15); -x_32 = !lean_is_exclusive(x_20); +lean_dec(x_28); +x_1 = x_16; +x_2 = x_30; +x_11 = x_29; +goto _start; +} +else +{ +uint8_t x_32; +x_32 = !lean_is_exclusive(x_22); if (x_32 == 0) { -lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; -x_33 = lean_ctor_get(x_20, 7); -lean_dec(x_33); -x_34 = lean_ctor_get(x_20, 6); +lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; 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_22, 1); +x_34 = lean_ctor_get(x_22, 0); lean_dec(x_34); -x_35 = lean_ctor_get(x_20, 5); -lean_dec(x_35); -x_36 = lean_ctor_get(x_20, 4); -lean_dec(x_36); -x_37 = lean_ctor_get(x_20, 3); -lean_dec(x_37); -x_38 = lean_ctor_get(x_20, 2); +lean_inc(x_15); +x_35 = l_Lean_MessageData_ofExpr(x_15); +x_36 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__6; +lean_ctor_set_tag(x_22, 7); +lean_ctor_set(x_22, 1, x_35); +lean_ctor_set(x_22, 0, x_36); +x_37 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; +lean_ctor_set_tag(x_17, 7); +lean_ctor_set(x_17, 1, x_37); +lean_ctor_set(x_17, 0, x_22); +x_38 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_21, x_17, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, 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); -x_39 = lean_ctor_get(x_20, 1); +x_41 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___lambda__1(x_15, x_39, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_40); lean_dec(x_39); -x_40 = lean_ctor_get(x_20, 0); -lean_dec(x_40); -lean_ctor_set(x_20, 3, x_31); -x_41 = lean_st_ref_set(x_3, x_20, x_21); -x_42 = lean_ctor_get(x_41, 1); +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_43 = lean_box(0); +x_44 = lean_ctor_get(x_42, 0); +lean_inc(x_44); +lean_dec(x_42); x_1 = x_16; -x_2 = x_43; -x_11 = x_42; +x_2 = x_44; +x_11 = x_43; goto _start; } else { -lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; -lean_dec(x_20); -x_45 = lean_alloc_ctor(0, 8, 1); -lean_ctor_set(x_45, 0, x_22); -lean_ctor_set(x_45, 1, x_23); -lean_ctor_set(x_45, 2, x_24); -lean_ctor_set(x_45, 3, x_31); -lean_ctor_set(x_45, 4, x_26); -lean_ctor_set(x_45, 5, x_27); -lean_ctor_set(x_45, 6, x_29); -lean_ctor_set(x_45, 7, x_30); -lean_ctor_set_uint8(x_45, sizeof(void*)*8, x_28); -x_46 = lean_st_ref_set(x_3, x_45, x_21); -x_47 = lean_ctor_get(x_46, 1); -lean_inc(x_47); -lean_dec(x_46); -x_48 = lean_box(0); +lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_46 = lean_ctor_get(x_22, 1); +lean_inc(x_46); +lean_dec(x_22); +lean_inc(x_15); +x_47 = l_Lean_MessageData_ofExpr(x_15); +x_48 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__6; +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_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; +lean_ctor_set_tag(x_17, 7); +lean_ctor_set(x_17, 1, x_50); +lean_ctor_set(x_17, 0, x_49); +x_51 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_21, x_17, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_46); +x_52 = lean_ctor_get(x_51, 0); +lean_inc(x_52); +x_53 = lean_ctor_get(x_51, 1); +lean_inc(x_53); +lean_dec(x_51); +x_54 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___lambda__1(x_15, x_52, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_53); +lean_dec(x_52); +x_55 = lean_ctor_get(x_54, 0); +lean_inc(x_55); +x_56 = lean_ctor_get(x_54, 1); +lean_inc(x_56); +lean_dec(x_54); +x_57 = lean_ctor_get(x_55, 0); +lean_inc(x_57); +lean_dec(x_55); +x_1 = x_16; +x_2 = x_57; +x_11 = x_56; +goto _start; +} +} +} +else +{ +lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; uint8_t x_63; +x_59 = lean_ctor_get(x_17, 1); +lean_inc(x_59); +lean_dec(x_17); +x_60 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__4; +x_61 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_60, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_59); +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_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; +x_64 = lean_ctor_get(x_61, 1); +lean_inc(x_64); +lean_dec(x_61); +x_65 = lean_box(0); +x_66 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___lambda__1(x_15, x_65, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_64); +x_67 = lean_ctor_get(x_66, 0); +lean_inc(x_67); +x_68 = lean_ctor_get(x_66, 1); +lean_inc(x_68); +lean_dec(x_66); +x_69 = lean_ctor_get(x_67, 0); +lean_inc(x_69); +lean_dec(x_67); +x_1 = x_16; +x_2 = x_69; +x_11 = x_68; +goto _start; +} +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; 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_71 = lean_ctor_get(x_61, 1); +lean_inc(x_71); +if (lean_is_exclusive(x_61)) { + lean_ctor_release(x_61, 0); + lean_ctor_release(x_61, 1); + x_72 = x_61; +} else { + lean_dec_ref(x_61); + x_72 = lean_box(0); +} +lean_inc(x_15); +x_73 = l_Lean_MessageData_ofExpr(x_15); +x_74 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__6; +if (lean_is_scalar(x_72)) { + x_75 = lean_alloc_ctor(7, 2, 0); +} else { + x_75 = x_72; + lean_ctor_set_tag(x_75, 7); +} +lean_ctor_set(x_75, 0, x_74); +lean_ctor_set(x_75, 1, x_73); +x_76 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; +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_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_60, x_77, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_71); +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 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___lambda__1(x_15, x_79, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_80); +lean_dec(x_79); +x_82 = lean_ctor_get(x_81, 0); +lean_inc(x_82); +x_83 = lean_ctor_get(x_81, 1); +lean_inc(x_83); +lean_dec(x_81); +x_84 = lean_ctor_get(x_82, 0); +lean_inc(x_84); +lean_dec(x_82); x_1 = x_16; -x_2 = x_48; -x_11 = x_47; +x_2 = x_84; +x_11 = x_83; goto _start; } } } } +} LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { @@ -1194,6 +1516,23 @@ x_6 = l_Lean_PersistentHashMap_eraseAux___at___private_Lean_Meta_Tactic_Grind_Co return x_6; } } +LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; +x_12 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_12; +} +} LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { @@ -1227,17 +1566,89 @@ lean_dec(x_1); return x_11; } } -LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_reinsertParents___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_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_reinsertParents___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: { -if (lean_obj_tag(x_1) == 0) +lean_object* x_12; +x_12 = l_Lean_Meta_Grind_addCongrTable(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_12) == 0) { -lean_object* x_12; lean_object* x_13; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); +uint8_t x_13; +x_13 = !lean_is_exclusive(x_12); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; +x_14 = lean_ctor_get(x_12, 0); +lean_dec(x_14); +x_15 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___lambda__1___closed__1; +lean_ctor_set(x_12, 0, x_15); +return x_12; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_12, 1); +lean_inc(x_16); +lean_dec(x_12); +x_17 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___lambda__1___closed__1; +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_17); +lean_ctor_set(x_18, 1, x_16); +return x_18; +} +} +else +{ +uint8_t x_19; +x_19 = !lean_is_exclusive(x_12); +if (x_19 == 0) +{ +return x_12; +} +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_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; +} +} +} +} +static lean_object* _init_l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_reinsertParents___spec__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("reinsert: ", 10, 10); +return x_1; +} +} +static lean_object* _init_l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_reinsertParents___spec__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_reinsertParents___spec__1___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_reinsertParents___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) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_12; lean_object* x_13; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); @@ -1307,11 +1718,24 @@ return x_22; } else { -lean_object* x_23; lean_object* x_24; +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; lean_dec(x_18); x_23 = lean_ctor_get(x_17, 1); lean_inc(x_23); lean_dec(x_17); +x_24 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__4; +x_25 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_24, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_23); +x_26 = lean_ctor_get(x_25, 0); +lean_inc(x_26); +x_27 = lean_unbox(x_26); +lean_dec(x_26); +if (x_27 == 0) +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_28 = lean_ctor_get(x_25, 1); +lean_inc(x_28); +lean_dec(x_25); +x_29 = lean_box(0); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); @@ -1320,22 +1744,15 @@ lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -x_24 = l_Lean_Meta_Grind_addCongrTable(x_15, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_23); -if (lean_obj_tag(x_24) == 0) +x_30 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_reinsertParents___spec__1___lambda__1(x_15, x_29, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_28); +if (lean_obj_tag(x_30) == 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 = lean_box(0); -x_1 = x_16; -x_2 = x_26; -x_11 = x_25; -goto _start; -} -else +lean_object* x_31; +x_31 = lean_ctor_get(x_30, 0); +lean_inc(x_31); +if (lean_obj_tag(x_31) == 0) { -uint8_t x_28; +uint8_t x_32; lean_dec(x_16); lean_dec(x_10); lean_dec(x_9); @@ -1345,32 +1762,45 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_28 = !lean_is_exclusive(x_24); -if (x_28 == 0) +x_32 = !lean_is_exclusive(x_30); +if (x_32 == 0) { -return x_24; +lean_object* x_33; +x_33 = lean_ctor_get(x_30, 0); +lean_dec(x_33); +return x_30; } else { -lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_29 = lean_ctor_get(x_24, 0); -x_30 = lean_ctor_get(x_24, 1); -lean_inc(x_30); -lean_inc(x_29); -lean_dec(x_24); -x_31 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_31, 0, x_29); -lean_ctor_set(x_31, 1, x_30); -return x_31; +lean_object* x_34; lean_object* x_35; +x_34 = lean_ctor_get(x_30, 1); +lean_inc(x_34); +lean_dec(x_30); +x_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_31); +lean_ctor_set(x_35, 1, x_34); +return x_35; } } +else +{ +lean_object* x_36; lean_object* x_37; +x_36 = lean_ctor_get(x_30, 1); +lean_inc(x_36); +lean_dec(x_30); +x_37 = lean_ctor_get(x_31, 0); +lean_inc(x_37); +lean_dec(x_31); +x_1 = x_16; +x_2 = x_37; +x_11 = x_36; +goto _start; } } else { -uint8_t x_32; +uint8_t x_39; lean_dec(x_16); -lean_dec(x_15); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -1379,96 +1809,378 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_32 = !lean_is_exclusive(x_17); -if (x_32 == 0) +x_39 = !lean_is_exclusive(x_30); +if (x_39 == 0) { -return x_17; +return x_30; } else { -lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_33 = lean_ctor_get(x_17, 0); -x_34 = lean_ctor_get(x_17, 1); -lean_inc(x_34); -lean_inc(x_33); -lean_dec(x_17); -x_35 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -return x_35; -} -} +lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_40 = lean_ctor_get(x_30, 0); +x_41 = lean_ctor_get(x_30, 1); +lean_inc(x_41); +lean_inc(x_40); +lean_dec(x_30); +x_42 = lean_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_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_reinsertParents(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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; -x_11 = lean_box(0); -x_12 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_reinsertParents___spec__1(x_1, x_11, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_12) == 0) +uint8_t x_43; +x_43 = !lean_is_exclusive(x_25); +if (x_43 == 0) { -uint8_t x_13; -x_13 = !lean_is_exclusive(x_12); -if (x_13 == 0) +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_44 = lean_ctor_get(x_25, 1); +x_45 = lean_ctor_get(x_25, 0); +lean_dec(x_45); +lean_inc(x_15); +x_46 = l_Lean_MessageData_ofExpr(x_15); +x_47 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_reinsertParents___spec__1___closed__2; +lean_ctor_set_tag(x_25, 7); +lean_ctor_set(x_25, 1, x_46); +lean_ctor_set(x_25, 0, x_47); +x_48 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; +x_49 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_49, 0, x_25); +lean_ctor_set(x_49, 1, x_48); +x_50 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_24, x_49, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_44); +x_51 = lean_ctor_get(x_50, 0); +lean_inc(x_51); +x_52 = lean_ctor_get(x_50, 1); +lean_inc(x_52); +lean_dec(x_50); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_53 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_reinsertParents___spec__1___lambda__1(x_15, x_51, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_52); +lean_dec(x_51); +if (lean_obj_tag(x_53) == 0) { -lean_object* x_14; -x_14 = lean_ctor_get(x_12, 0); -lean_dec(x_14); -lean_ctor_set(x_12, 0, x_11); -return x_12; +lean_object* x_54; +x_54 = lean_ctor_get(x_53, 0); +lean_inc(x_54); +if (lean_obj_tag(x_54) == 0) +{ +uint8_t x_55; +lean_dec(x_16); +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_55 = !lean_is_exclusive(x_53); +if (x_55 == 0) +{ +lean_object* x_56; +x_56 = lean_ctor_get(x_53, 0); +lean_dec(x_56); +return x_53; } 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_11); -lean_ctor_set(x_16, 1, x_15); -return x_16; +lean_object* x_57; lean_object* x_58; +x_57 = lean_ctor_get(x_53, 1); +lean_inc(x_57); +lean_dec(x_53); +x_58 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_58, 0, x_54); +lean_ctor_set(x_58, 1, x_57); +return x_58; } } else { -uint8_t x_17; -x_17 = !lean_is_exclusive(x_12); -if (x_17 == 0) -{ -return x_12; +lean_object* x_59; lean_object* x_60; +x_59 = lean_ctor_get(x_53, 1); +lean_inc(x_59); +lean_dec(x_53); +x_60 = lean_ctor_get(x_54, 0); +lean_inc(x_60); +lean_dec(x_54); +x_1 = x_16; +x_2 = x_60; +x_11 = x_59; +goto _start; +} } else { -lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_18 = lean_ctor_get(x_12, 0); -x_19 = lean_ctor_get(x_12, 1); -lean_inc(x_19); -lean_inc(x_18); -lean_dec(x_12); -x_20 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_20, 0, x_18); -lean_ctor_set(x_20, 1, x_19); -return x_20; -} -} -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_closeGoalWithTrueEqFalse___closed__1() { -_start: +uint8_t x_62; +lean_dec(x_16); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_62 = !lean_is_exclusive(x_53); +if (x_62 == 0) { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("True", 4, 4); -return x_1; -} +return x_53; } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_closeGoalWithTrueEqFalse___closed__2() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("intro", 5, 5); -return x_1; +lean_object* x_63; lean_object* x_64; lean_object* x_65; +x_63 = lean_ctor_get(x_53, 0); +x_64 = lean_ctor_get(x_53, 1); +lean_inc(x_64); +lean_inc(x_63); +lean_dec(x_53); +x_65 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_65, 0, x_63); +lean_ctor_set(x_65, 1, x_64); +return x_65; +} +} +} +else +{ +lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; +x_66 = lean_ctor_get(x_25, 1); +lean_inc(x_66); +lean_dec(x_25); +lean_inc(x_15); +x_67 = l_Lean_MessageData_ofExpr(x_15); +x_68 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_reinsertParents___spec__1___closed__2; +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_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; +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_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_24, x_71, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_66); +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); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_75 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_reinsertParents___spec__1___lambda__1(x_15, x_73, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_74); +lean_dec(x_73); +if (lean_obj_tag(x_75) == 0) +{ +lean_object* x_76; +x_76 = lean_ctor_get(x_75, 0); +lean_inc(x_76); +if (lean_obj_tag(x_76) == 0) +{ +lean_object* x_77; lean_object* x_78; lean_object* x_79; +lean_dec(x_16); +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_77 = lean_ctor_get(x_75, 1); +lean_inc(x_77); +if (lean_is_exclusive(x_75)) { + lean_ctor_release(x_75, 0); + lean_ctor_release(x_75, 1); + x_78 = x_75; +} else { + lean_dec_ref(x_75); + x_78 = lean_box(0); +} +if (lean_is_scalar(x_78)) { + x_79 = lean_alloc_ctor(0, 2, 0); +} else { + x_79 = x_78; +} +lean_ctor_set(x_79, 0, x_76); +lean_ctor_set(x_79, 1, x_77); +return x_79; +} +else +{ +lean_object* x_80; lean_object* x_81; +x_80 = lean_ctor_get(x_75, 1); +lean_inc(x_80); +lean_dec(x_75); +x_81 = lean_ctor_get(x_76, 0); +lean_inc(x_81); +lean_dec(x_76); +x_1 = x_16; +x_2 = x_81; +x_11 = x_80; +goto _start; +} +} +else +{ +lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; +lean_dec(x_16); +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_75, 0); +lean_inc(x_83); +x_84 = lean_ctor_get(x_75, 1); +lean_inc(x_84); +if (lean_is_exclusive(x_75)) { + lean_ctor_release(x_75, 0); + lean_ctor_release(x_75, 1); + x_85 = x_75; +} else { + lean_dec_ref(x_75); + 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; +} +} +} +} +} +else +{ +uint8_t x_87; +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_87 = !lean_is_exclusive(x_17); +if (x_87 == 0) +{ +return x_17; +} +else +{ +lean_object* x_88; lean_object* x_89; lean_object* x_90; +x_88 = lean_ctor_get(x_17, 0); +x_89 = lean_ctor_get(x_17, 1); +lean_inc(x_89); +lean_inc(x_88); +lean_dec(x_17); +x_90 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_90, 0, x_88); +lean_ctor_set(x_90, 1, x_89); +return x_90; +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_reinsertParents(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; lean_object* x_12; +x_11 = lean_box(0); +x_12 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_reinsertParents___spec__1(x_1, x_11, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_12) == 0) +{ +uint8_t x_13; +x_13 = !lean_is_exclusive(x_12); +if (x_13 == 0) +{ +lean_object* x_14; +x_14 = lean_ctor_get(x_12, 0); +lean_dec(x_14); +lean_ctor_set(x_12, 0, x_11); +return x_12; +} +else +{ +lean_object* x_15; lean_object* x_16; +x_15 = lean_ctor_get(x_12, 1); +lean_inc(x_15); +lean_dec(x_12); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_11); +lean_ctor_set(x_16, 1, x_15); +return x_16; +} +} +else +{ +uint8_t x_17; +x_17 = !lean_is_exclusive(x_12); +if (x_17 == 0) +{ +return x_12; +} +else +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_18 = lean_ctor_get(x_12, 0); +x_19 = lean_ctor_get(x_12, 1); +lean_inc(x_19); +lean_inc(x_18); +lean_dec(x_12); +x_20 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_20, 0, x_18); +lean_ctor_set(x_20, 1, x_19); +return x_20; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_reinsertParents___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_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_reinsertParents___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_2); +return x_12; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_closeGoalWithTrueEqFalse___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("True", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_closeGoalWithTrueEqFalse___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("intro", 5, 5); +return x_1; } } static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_closeGoalWithTrueEqFalse___closed__3() { @@ -2130,6 +2842,7 @@ else lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_inc(x_1); lean_ctor_set(x_26, 9, x_1); +lean_inc(x_16); x_42 = l_Lean_Meta_Grind_setENode(x_16, x_26, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_27); x_43 = lean_ctor_get(x_42, 1); lean_inc(x_43); @@ -2244,6 +2957,7 @@ lean_ctor_set_uint8(x_70, sizeof(void*)*10 + 1, x_60); lean_ctor_set_uint8(x_70, sizeof(void*)*10 + 2, x_61); lean_ctor_set_uint8(x_70, sizeof(void*)*10 + 3, x_62); lean_ctor_set_uint8(x_70, sizeof(void*)*10 + 4, x_63); +lean_inc(x_16); x_71 = l_Lean_Meta_Grind_setENode(x_16, x_70, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_27); x_72 = lean_ctor_get(x_71, 1); lean_inc(x_72); @@ -2512,6 +3226,7 @@ lean_dec(x_18); lean_inc(x_2); lean_inc(x_17); lean_ctor_set(x_14, 2, x_2); +lean_inc(x_3); x_19 = l_Lean_Meta_Grind_setENode(x_3, x_14, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_15); x_20 = lean_ctor_get(x_19, 1); lean_inc(x_20); @@ -2638,6 +3353,7 @@ lean_ctor_set_uint8(x_50, sizeof(void*)*10 + 1, x_43); lean_ctor_set_uint8(x_50, sizeof(void*)*10 + 2, x_44); lean_ctor_set_uint8(x_50, sizeof(void*)*10 + 3, x_45); lean_ctor_set_uint8(x_50, sizeof(void*)*10 + 4, x_46); +lean_inc(x_3); x_51 = l_Lean_Meta_Grind_setENode(x_3, x_50, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_15); x_52 = lean_ctor_get(x_51, 1); lean_inc(x_52); @@ -3182,6 +3898,7 @@ lean_ctor_set_uint8(x_61, sizeof(void*)*10 + 1, x_49); lean_ctor_set_uint8(x_61, sizeof(void*)*10 + 2, x_50); lean_ctor_set_uint8(x_61, sizeof(void*)*10 + 3, x_59); lean_ctor_set_uint8(x_61, sizeof(void*)*10 + 4, x_60); +lean_inc(x_7); x_62 = l_Lean_Meta_Grind_setENode(x_7, x_61, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_57); x_63 = lean_ctor_get(x_62, 1); lean_inc(x_63); @@ -3397,6 +4114,7 @@ lean_ctor_set_uint8(x_107, sizeof(void*)*10 + 1, x_94); lean_ctor_set_uint8(x_107, sizeof(void*)*10 + 2, x_95); lean_ctor_set_uint8(x_107, sizeof(void*)*10 + 3, x_105); lean_ctor_set_uint8(x_107, sizeof(void*)*10 + 4, x_106); +lean_inc(x_7); x_108 = l_Lean_Meta_Grind_setENode(x_7, x_107, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_103); x_109 = lean_ctor_get(x_108, 1); lean_inc(x_109); @@ -3514,7 +4232,9 @@ lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); +lean_dec(x_7); lean_dec(x_6); +lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); x_132 = !lean_is_exclusive(x_23); @@ -3548,7 +4268,9 @@ lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); +lean_dec(x_7); lean_dec(x_6); +lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); @@ -3577,7 +4299,7 @@ static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_G _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("", 0, 0); +x_1 = lean_mk_string_unchecked(" new root ", 10, 10); return x_1; } } @@ -3594,7 +4316,7 @@ static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_G _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked(" new root ", 10, 10); +x_1 = lean_mk_string_unchecked(", ", 2, 2); return x_1; } } @@ -3607,23 +4329,6 @@ x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__5() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked(", ", 2, 2); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__6() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__5; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; -} -} LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, 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) { _start: { @@ -3653,6 +4358,7 @@ lean_inc(x_24); lean_ctor_set(x_2, 5, x_28); lean_ctor_set(x_2, 4, x_27); lean_ctor_set_uint8(x_2, sizeof(void*)*10, x_5); +lean_inc(x_1); x_29 = l_Lean_Meta_Grind_setENode(x_1, x_2, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_22); x_30 = !lean_is_exclusive(x_29); if (x_30 == 0) @@ -3715,9 +4421,7 @@ lean_inc(x_48); lean_dec(x_45); x_49 = lean_box(0); x_50 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__2(x_40, x_33, x_9, x_24, x_35, x_34, x_42, x_10, x_37, x_36, x_49, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_48); -lean_dec(x_42); lean_dec(x_35); -lean_dec(x_24); return x_50; } else @@ -3762,11 +4466,11 @@ lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean x_67 = lean_ctor_get(x_65, 0); x_68 = lean_ctor_get(x_65, 1); x_69 = l_Lean_MessageData_ofFormat(x_56); -x_70 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__2; +x_70 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; lean_ctor_set_tag(x_65, 7); lean_ctor_set(x_65, 1, x_69); lean_ctor_set(x_65, 0, x_70); -x_71 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__4; +x_71 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__2; lean_ctor_set_tag(x_58, 7); lean_ctor_set(x_58, 1, x_71); lean_ctor_set(x_58, 0, x_65); @@ -3774,7 +4478,7 @@ x_72 = l_Lean_MessageData_ofFormat(x_60); lean_ctor_set_tag(x_54, 7); lean_ctor_set(x_54, 1, x_72); lean_ctor_set(x_54, 0, x_58); -x_73 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__6; +x_73 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__4; lean_ctor_set_tag(x_45, 7); lean_ctor_set(x_45, 1, x_73); lean_ctor_set(x_45, 0, x_54); @@ -3793,9 +4497,7 @@ lean_inc(x_77); lean_dec(x_75); x_78 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__2(x_40, x_33, x_9, x_24, x_35, x_34, x_42, x_10, x_37, x_36, x_76, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_77); lean_dec(x_76); -lean_dec(x_42); lean_dec(x_35); -lean_dec(x_24); return x_78; } else @@ -3807,11 +4509,11 @@ lean_inc(x_80); lean_inc(x_79); lean_dec(x_65); x_81 = l_Lean_MessageData_ofFormat(x_56); -x_82 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__2; +x_82 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; x_83 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_83, 0, x_82); lean_ctor_set(x_83, 1, x_81); -x_84 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__4; +x_84 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__2; lean_ctor_set_tag(x_58, 7); lean_ctor_set(x_58, 1, x_84); lean_ctor_set(x_58, 0, x_83); @@ -3819,7 +4521,7 @@ x_85 = l_Lean_MessageData_ofFormat(x_60); lean_ctor_set_tag(x_54, 7); lean_ctor_set(x_54, 1, x_85); lean_ctor_set(x_54, 0, x_58); -x_86 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__6; +x_86 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__4; lean_ctor_set_tag(x_45, 7); lean_ctor_set(x_45, 1, x_86); lean_ctor_set(x_45, 0, x_54); @@ -3838,9 +4540,7 @@ lean_inc(x_90); lean_dec(x_88); x_91 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__2(x_40, x_33, x_9, x_24, x_35, x_34, x_42, x_10, x_37, x_36, x_89, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_90); lean_dec(x_89); -lean_dec(x_42); lean_dec(x_35); -lean_dec(x_24); return x_91; } } @@ -3922,7 +4622,7 @@ if (lean_is_exclusive(x_101)) { x_104 = lean_box(0); } x_105 = l_Lean_MessageData_ofFormat(x_56); -x_106 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__2; +x_106 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; if (lean_is_scalar(x_104)) { x_107 = lean_alloc_ctor(7, 2, 0); } else { @@ -3931,7 +4631,7 @@ if (lean_is_scalar(x_104)) { } lean_ctor_set(x_107, 0, x_106); lean_ctor_set(x_107, 1, x_105); -x_108 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__4; +x_108 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__2; x_109 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_109, 0, x_107); lean_ctor_set(x_109, 1, x_108); @@ -3939,7 +4639,7 @@ x_110 = l_Lean_MessageData_ofFormat(x_96); lean_ctor_set_tag(x_54, 7); lean_ctor_set(x_54, 1, x_110); lean_ctor_set(x_54, 0, x_109); -x_111 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__6; +x_111 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__4; lean_ctor_set_tag(x_45, 7); lean_ctor_set(x_45, 1, x_111); lean_ctor_set(x_45, 0, x_54); @@ -3958,9 +4658,7 @@ lean_inc(x_115); lean_dec(x_113); x_116 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__2(x_40, x_33, x_9, x_24, x_35, x_34, x_42, x_10, x_37, x_36, x_114, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_115); lean_dec(x_114); -lean_dec(x_42); lean_dec(x_35); -lean_dec(x_24); return x_116; } else @@ -4056,7 +4754,7 @@ if (lean_is_exclusive(x_130)) { x_133 = lean_box(0); } x_134 = l_Lean_MessageData_ofFormat(x_121); -x_135 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__2; +x_135 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; if (lean_is_scalar(x_133)) { x_136 = lean_alloc_ctor(7, 2, 0); } else { @@ -4065,7 +4763,7 @@ if (lean_is_scalar(x_133)) { } lean_ctor_set(x_136, 0, x_135); lean_ctor_set(x_136, 1, x_134); -x_137 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__4; +x_137 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__2; if (lean_is_scalar(x_126)) { x_138 = lean_alloc_ctor(7, 2, 0); } else { @@ -4078,7 +4776,7 @@ x_139 = l_Lean_MessageData_ofFormat(x_124); x_140 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_140, 0, x_138); lean_ctor_set(x_140, 1, x_139); -x_141 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__6; +x_141 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__4; lean_ctor_set_tag(x_45, 7); lean_ctor_set(x_45, 1, x_141); lean_ctor_set(x_45, 0, x_140); @@ -4097,9 +4795,7 @@ lean_inc(x_145); lean_dec(x_143); x_146 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__2(x_40, x_33, x_9, x_24, x_35, x_34, x_42, x_10, x_37, x_36, x_144, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_145); lean_dec(x_144); -lean_dec(x_42); lean_dec(x_35); -lean_dec(x_24); return x_146; } else @@ -4206,7 +4902,7 @@ if (lean_is_exclusive(x_163)) { x_166 = lean_box(0); } x_167 = l_Lean_MessageData_ofFormat(x_153); -x_168 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__2; +x_168 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; if (lean_is_scalar(x_166)) { x_169 = lean_alloc_ctor(7, 2, 0); } else { @@ -4215,7 +4911,7 @@ if (lean_is_scalar(x_166)) { } lean_ctor_set(x_169, 0, x_168); lean_ctor_set(x_169, 1, x_167); -x_170 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__4; +x_170 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__2; if (lean_is_scalar(x_159)) { x_171 = lean_alloc_ctor(7, 2, 0); } else { @@ -4233,7 +4929,7 @@ if (lean_is_scalar(x_155)) { } lean_ctor_set(x_173, 0, x_171); lean_ctor_set(x_173, 1, x_172); -x_174 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__6; +x_174 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__4; x_175 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_175, 0, x_173); lean_ctor_set(x_175, 1, x_174); @@ -4252,9 +4948,7 @@ lean_inc(x_179); lean_dec(x_177); x_180 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__2(x_40, x_33, x_9, x_24, x_35, x_34, x_42, x_10, x_37, x_36, x_178, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_179); lean_dec(x_178); -lean_dec(x_42); lean_dec(x_35); -lean_dec(x_24); return x_180; } else @@ -4393,9 +5087,7 @@ lean_inc(x_197); lean_dec(x_194); x_198 = lean_box(0); x_199 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__2(x_189, x_33, x_9, x_24, x_35, x_34, x_191, x_10, x_37, x_36, x_198, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_197); -lean_dec(x_191); lean_dec(x_35); -lean_dec(x_24); return x_199; } else @@ -4461,7 +5153,7 @@ if (lean_is_exclusive(x_213)) { x_216 = lean_box(0); } x_217 = l_Lean_MessageData_ofFormat(x_203); -x_218 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__2; +x_218 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; if (lean_is_scalar(x_216)) { x_219 = lean_alloc_ctor(7, 2, 0); } else { @@ -4470,7 +5162,7 @@ if (lean_is_scalar(x_216)) { } lean_ctor_set(x_219, 0, x_218); lean_ctor_set(x_219, 1, x_217); -x_220 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__4; +x_220 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__2; if (lean_is_scalar(x_209)) { x_221 = lean_alloc_ctor(7, 2, 0); } else { @@ -4488,7 +5180,7 @@ if (lean_is_scalar(x_205)) { } lean_ctor_set(x_223, 0, x_221); lean_ctor_set(x_223, 1, x_222); -x_224 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__6; +x_224 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__4; if (lean_is_scalar(x_201)) { x_225 = lean_alloc_ctor(7, 2, 0); } else { @@ -4512,9 +5204,7 @@ lean_inc(x_230); lean_dec(x_228); x_231 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__2(x_189, x_33, x_9, x_24, x_35, x_34, x_191, x_10, x_37, x_36, x_229, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_230); lean_dec(x_229); -lean_dec(x_191); lean_dec(x_35); -lean_dec(x_24); return x_231; } else @@ -4674,9 +5364,7 @@ lean_inc(x_256); lean_dec(x_253); x_257 = lean_box(0); x_258 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__2(x_247, x_241, x_9, x_24, x_243, x_242, x_250, x_10, x_245, x_244, x_257, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_256); -lean_dec(x_250); lean_dec(x_243); -lean_dec(x_24); return x_258; } else @@ -4742,7 +5430,7 @@ if (lean_is_exclusive(x_272)) { x_275 = lean_box(0); } x_276 = l_Lean_MessageData_ofFormat(x_262); -x_277 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__2; +x_277 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; if (lean_is_scalar(x_275)) { x_278 = lean_alloc_ctor(7, 2, 0); } else { @@ -4751,7 +5439,7 @@ if (lean_is_scalar(x_275)) { } lean_ctor_set(x_278, 0, x_277); lean_ctor_set(x_278, 1, x_276); -x_279 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__4; +x_279 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__2; if (lean_is_scalar(x_268)) { x_280 = lean_alloc_ctor(7, 2, 0); } else { @@ -4769,7 +5457,7 @@ if (lean_is_scalar(x_264)) { } lean_ctor_set(x_282, 0, x_280); lean_ctor_set(x_282, 1, x_281); -x_283 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__6; +x_283 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__4; if (lean_is_scalar(x_260)) { x_284 = lean_alloc_ctor(7, 2, 0); } else { @@ -4798,9 +5486,7 @@ lean_inc(x_290); lean_dec(x_288); x_291 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__2(x_247, x_241, x_9, x_24, x_243, x_242, x_250, x_10, x_245, x_244, x_289, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_290); lean_dec(x_289); -lean_dec(x_250); lean_dec(x_243); -lean_dec(x_24); return x_291; } else @@ -4940,6 +5626,7 @@ lean_ctor_set_uint8(x_314, sizeof(void*)*10 + 1, x_305); lean_ctor_set_uint8(x_314, sizeof(void*)*10 + 2, x_306); lean_ctor_set_uint8(x_314, sizeof(void*)*10 + 3, x_307); lean_ctor_set_uint8(x_314, sizeof(void*)*10 + 4, x_308); +lean_inc(x_1); x_315 = l_Lean_Meta_Grind_setENode(x_1, x_314, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_22); x_316 = lean_ctor_get(x_315, 1); lean_inc(x_316); @@ -5011,9 +5698,7 @@ lean_inc(x_333); lean_dec(x_330); x_334 = lean_box(0); x_335 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__2(x_324, x_318, x_9, x_302, x_320, x_319, x_327, x_10, x_322, x_321, x_334, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_333); -lean_dec(x_327); lean_dec(x_320); -lean_dec(x_302); return x_335; } else @@ -5079,7 +5764,7 @@ if (lean_is_exclusive(x_349)) { x_352 = lean_box(0); } x_353 = l_Lean_MessageData_ofFormat(x_339); -x_354 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__2; +x_354 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; if (lean_is_scalar(x_352)) { x_355 = lean_alloc_ctor(7, 2, 0); } else { @@ -5088,7 +5773,7 @@ if (lean_is_scalar(x_352)) { } lean_ctor_set(x_355, 0, x_354); lean_ctor_set(x_355, 1, x_353); -x_356 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__4; +x_356 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__2; if (lean_is_scalar(x_345)) { x_357 = lean_alloc_ctor(7, 2, 0); } else { @@ -5106,7 +5791,7 @@ if (lean_is_scalar(x_341)) { } lean_ctor_set(x_359, 0, x_357); lean_ctor_set(x_359, 1, x_358); -x_360 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__6; +x_360 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__4; if (lean_is_scalar(x_337)) { x_361 = lean_alloc_ctor(7, 2, 0); } else { @@ -5140,9 +5825,7 @@ lean_inc(x_367); lean_dec(x_365); x_368 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__2(x_324, x_318, x_9, x_302, x_320, x_319, x_327, x_10, x_322, x_321, x_366, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_367); lean_dec(x_366); -lean_dec(x_327); lean_dec(x_320); -lean_dec(x_302); return x_368; } else @@ -5282,30 +5965,14 @@ return x_380; static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__1() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("grind", 5, 5); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("debug", 5, 5); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__3() { -_start: -{ lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__1; -x_2 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__2; +x_1 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__1; +x_2 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__2; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__4() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__2() { _start: { lean_object* x_1; @@ -5313,16 +5980,16 @@ x_1 = lean_mk_string_unchecked("adding ", 7, 7); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__5() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__3() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__4; +x_1 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__2; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__6() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__4() { _start: { lean_object* x_1; @@ -5330,11 +5997,11 @@ x_1 = lean_mk_string_unchecked(" ↦ ", 5, 3); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__7() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__5() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__6; +x_1 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__4; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } @@ -5343,7 +6010,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Gr _start: { lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; -x_19 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__3; +x_19 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__1; x_20 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_19, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); x_21 = lean_ctor_get(x_20, 0); lean_inc(x_21); @@ -5384,11 +6051,11 @@ lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean x_35 = lean_ctor_get(x_33, 0); x_36 = lean_ctor_get(x_33, 1); x_37 = l_Lean_MessageData_ofFormat(x_31); -x_38 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__5; +x_38 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__3; lean_ctor_set_tag(x_33, 7); lean_ctor_set(x_33, 1, x_37); lean_ctor_set(x_33, 0, x_38); -x_39 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__7; +x_39 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__5; lean_ctor_set_tag(x_29, 7); lean_ctor_set(x_29, 1, x_39); lean_ctor_set(x_29, 0, x_33); @@ -5396,7 +6063,7 @@ x_40 = l_Lean_MessageData_ofFormat(x_35); lean_ctor_set_tag(x_20, 7); lean_ctor_set(x_20, 1, x_40); lean_ctor_set(x_20, 0, x_29); -x_41 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__2; +x_41 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; x_42 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_42, 0, x_20); lean_ctor_set(x_42, 1, x_41); @@ -5419,11 +6086,11 @@ lean_inc(x_48); lean_inc(x_47); lean_dec(x_33); x_49 = l_Lean_MessageData_ofFormat(x_31); -x_50 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__5; +x_50 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__3; x_51 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_51, 0, x_50); lean_ctor_set(x_51, 1, x_49); -x_52 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__7; +x_52 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__5; lean_ctor_set_tag(x_29, 7); lean_ctor_set(x_29, 1, x_52); lean_ctor_set(x_29, 0, x_51); @@ -5431,7 +6098,7 @@ x_53 = l_Lean_MessageData_ofFormat(x_47); lean_ctor_set_tag(x_20, 7); lean_ctor_set(x_20, 1, x_53); lean_ctor_set(x_20, 0, x_29); -x_54 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__2; +x_54 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; x_55 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_55, 0, x_20); lean_ctor_set(x_55, 1, x_54); @@ -5468,7 +6135,7 @@ if (lean_is_exclusive(x_62)) { x_65 = lean_box(0); } x_66 = l_Lean_MessageData_ofFormat(x_60); -x_67 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__5; +x_67 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__3; if (lean_is_scalar(x_65)) { x_68 = lean_alloc_ctor(7, 2, 0); } else { @@ -5477,7 +6144,7 @@ if (lean_is_scalar(x_65)) { } lean_ctor_set(x_68, 0, x_67); lean_ctor_set(x_68, 1, x_66); -x_69 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__7; +x_69 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__5; x_70 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_70, 0, x_68); lean_ctor_set(x_70, 1, x_69); @@ -5485,7 +6152,7 @@ x_71 = l_Lean_MessageData_ofFormat(x_63); lean_ctor_set_tag(x_20, 7); lean_ctor_set(x_20, 1, x_71); lean_ctor_set(x_20, 0, x_70); -x_72 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__2; +x_72 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; x_73 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_73, 0, x_20); lean_ctor_set(x_73, 1, x_72); @@ -5533,7 +6200,7 @@ if (lean_is_exclusive(x_83)) { x_86 = lean_box(0); } x_87 = l_Lean_MessageData_ofFormat(x_80); -x_88 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__5; +x_88 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__3; if (lean_is_scalar(x_86)) { x_89 = lean_alloc_ctor(7, 2, 0); } else { @@ -5542,7 +6209,7 @@ if (lean_is_scalar(x_86)) { } lean_ctor_set(x_89, 0, x_88); lean_ctor_set(x_89, 1, x_87); -x_90 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__7; +x_90 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__5; if (lean_is_scalar(x_82)) { x_91 = lean_alloc_ctor(7, 2, 0); } else { @@ -5555,7 +6222,7 @@ x_92 = l_Lean_MessageData_ofFormat(x_84); x_93 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_93, 0, x_91); lean_ctor_set(x_93, 1, x_92); -x_94 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__2; +x_94 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; x_95 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_95, 0, x_93); lean_ctor_set(x_95, 1, x_94); @@ -5622,9 +6289,7 @@ x_23 = lean_unbox(x_10); lean_dec(x_10); x_24 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_21, x_22, x_23, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20); lean_dec(x_11); -lean_dec(x_7); lean_dec(x_5); -lean_dec(x_4); return x_24; } } @@ -5729,7 +6394,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Gr _start: { lean_object* x_11; lean_object* x_12; uint8_t x_13; -x_11 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__3; +x_11 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__1; x_12 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_11, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); x_13 = !lean_is_exclusive(x_12); if (x_13 == 0) @@ -5773,7 +6438,7 @@ x_24 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___la lean_ctor_set_tag(x_12, 7); lean_ctor_set(x_12, 1, x_23); lean_ctor_set(x_12, 0, x_24); -x_25 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__2; +x_25 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; x_26 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_26, 0, x_12); lean_ctor_set(x_26, 1, x_25); @@ -5862,7 +6527,7 @@ x_45 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___la 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_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__2; +x_47 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; x_48 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_48, 0, x_46); lean_ctor_set(x_48, 1, x_47); @@ -6887,653 +7552,774 @@ return x_61; } } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__8(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_11; lean_object* x_12; -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; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__9___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__8___boxed), 10, 0); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__9___closed__2() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__8___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked(" and ", 5, 5); +x_1 = lean_mk_string_unchecked("eqc", 3, 3); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__9___closed__3() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__8___closed__2() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__9___closed__2; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__8___closed__1; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__9___closed__4() { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__8(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked(" are already in the same equivalence class", 42, 42); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__9___closed__5() { -_start: +lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_17 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__8___closed__2; +x_18 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_17, x_8, x_9, 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_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__9___closed__4; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; -} +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___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__7(x_1, x_2, x_3, x_4, x_5, x_6, x_22, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_21); +return x_23; } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___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, lean_object* x_13, lean_object* x_14) { -_start: +else { -lean_object* x_15; -lean_inc(x_1); -x_15 = l_Lean_Meta_Grind_getENode(x_1, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_15) == 0) +if (x_4 == 0) { -lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_16 = lean_ctor_get(x_15, 0); -lean_inc(x_16); -x_17 = lean_ctor_get(x_15, 1); -lean_inc(x_17); -lean_dec(x_15); -lean_inc(x_2); -x_18 = l_Lean_Meta_Grind_getENode(x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_17); -if (lean_obj_tag(x_18) == 0) +uint8_t x_24; +x_24 = !lean_is_exclusive(x_18); +if (x_24 == 0) { -lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; -x_19 = lean_ctor_get(x_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_16, 2); -lean_inc(x_21); -x_22 = lean_ctor_get(x_19, 2); -lean_inc(x_22); -x_23 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_21, x_22); -lean_dec(x_22); -lean_dec(x_21); -if (x_23 == 0) +lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_25 = lean_ctor_get(x_18, 1); +x_26 = lean_ctor_get(x_18, 0); +lean_dec(x_26); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_6); +lean_inc(x_5); +x_27 = l_Lean_Meta_mkEq(x_5, x_6, x_12, x_13, x_14, x_15, x_25); +if (lean_obj_tag(x_27) == 0) { -lean_object* x_24; lean_object* x_25; -x_24 = lean_box(0); -x_25 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__7(x_16, x_19, x_3, x_4, x_1, x_2, x_24, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_20); -return 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; lean_object* x_34; lean_object* x_35; lean_object* x_36; +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_MessageData_ofExpr(x_28); +x_31 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; +lean_ctor_set_tag(x_18, 7); +lean_ctor_set(x_18, 1, x_30); +lean_ctor_set(x_18, 0, x_31); +x_32 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_32, 0, x_18); +lean_ctor_set(x_32, 1, x_31); +x_33 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_17, x_32, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_29); +x_34 = lean_ctor_get(x_33, 0); +lean_inc(x_34); +x_35 = lean_ctor_get(x_33, 1); +lean_inc(x_35); +lean_dec(x_33); +x_36 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__7(x_1, x_2, x_3, x_4, x_5, x_6, x_34, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_35); +lean_dec(x_34); +return x_36; } else { -lean_object* x_26; lean_object* x_27; uint8_t x_28; -lean_dec(x_19); -lean_dec(x_16); +uint8_t x_37; +lean_free_object(x_18); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_5); lean_dec(x_3); -x_26 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__3; -x_27 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_26, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_20); -x_28 = !lean_is_exclusive(x_27); -if (x_28 == 0) -{ -lean_object* x_29; lean_object* x_30; lean_object* x_31; uint8_t x_32; -x_29 = lean_ctor_get(x_27, 0); -x_30 = lean_ctor_get(x_27, 1); -x_31 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__9___closed__1; -x_32 = lean_unbox(x_29); -lean_dec(x_29); -if (x_32 == 0) -{ -lean_object* x_33; lean_object* x_34; -lean_free_object(x_27); lean_dec(x_2); lean_dec(x_1); -x_33 = lean_box(0); -x_34 = lean_apply_10(x_31, x_33, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_30); -return x_34; +x_37 = !lean_is_exclusive(x_27); +if (x_37 == 0) +{ +return x_27; } else { -lean_object* x_35; uint8_t x_36; -x_35 = l_Lean_Meta_Grind_ppENodeRef(x_1, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_30); -lean_dec(x_1); -x_36 = !lean_is_exclusive(x_35); -if (x_36 == 0) +lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_38 = lean_ctor_get(x_27, 0); +x_39 = lean_ctor_get(x_27, 1); +lean_inc(x_39); +lean_inc(x_38); +lean_dec(x_27); +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_37; lean_object* x_38; lean_object* x_39; uint8_t x_40; -x_37 = lean_ctor_get(x_35, 0); -x_38 = lean_ctor_get(x_35, 1); -x_39 = l_Lean_Meta_Grind_ppENodeRef(x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_38); -lean_dec(x_2); -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_18, 1); +lean_inc(x_41); +lean_dec(x_18); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_6); +lean_inc(x_5); +x_42 = l_Lean_Meta_mkEq(x_5, x_6, x_12, x_13, x_14, x_15, x_41); +if (lean_obj_tag(x_42) == 0) { -lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; -x_41 = lean_ctor_get(x_39, 0); -x_42 = lean_ctor_get(x_39, 1); -x_43 = l_Lean_MessageData_ofFormat(x_37); -x_44 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__2; -lean_ctor_set_tag(x_39, 7); -lean_ctor_set(x_39, 1, x_43); -lean_ctor_set(x_39, 0, x_44); -x_45 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__9___closed__3; -lean_ctor_set_tag(x_35, 7); -lean_ctor_set(x_35, 1, x_45); -lean_ctor_set(x_35, 0, x_39); -x_46 = l_Lean_MessageData_ofFormat(x_41); -lean_ctor_set_tag(x_27, 7); -lean_ctor_set(x_27, 1, x_46); -lean_ctor_set(x_27, 0, x_35); -x_47 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__9___closed__5; +lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; 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_42, 0); +lean_inc(x_43); +x_44 = lean_ctor_get(x_42, 1); +lean_inc(x_44); +lean_dec(x_42); +x_45 = l_Lean_MessageData_ofExpr(x_43); +x_46 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; +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 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_48, 0, x_27); -lean_ctor_set(x_48, 1, x_47); -x_49 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_26, x_48, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_42); +lean_ctor_set(x_48, 0, x_47); +lean_ctor_set(x_48, 1, x_46); +x_49 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_17, x_48, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_44); x_50 = lean_ctor_get(x_49, 0); lean_inc(x_50); x_51 = lean_ctor_get(x_49, 1); lean_inc(x_51); lean_dec(x_49); -x_52 = lean_apply_10(x_31, x_50, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_51); +x_52 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__7(x_1, x_2, x_3, x_4, x_5, x_6, x_50, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_51); +lean_dec(x_50); return x_52; } else { -lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; -x_53 = lean_ctor_get(x_39, 0); -x_54 = lean_ctor_get(x_39, 1); -lean_inc(x_54); -lean_inc(x_53); -lean_dec(x_39); -x_55 = l_Lean_MessageData_ofFormat(x_37); -x_56 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__2; -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___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__9___closed__3; -lean_ctor_set_tag(x_35, 7); -lean_ctor_set(x_35, 1, x_58); -lean_ctor_set(x_35, 0, x_57); -x_59 = l_Lean_MessageData_ofFormat(x_53); -lean_ctor_set_tag(x_27, 7); -lean_ctor_set(x_27, 1, x_59); -lean_ctor_set(x_27, 0, x_35); -x_60 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__9___closed__5; -x_61 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_61, 0, x_27); -lean_ctor_set(x_61, 1, x_60); -x_62 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_26, x_61, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_54); -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_apply_10(x_31, x_63, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_64); -return x_65; -} -} -else -{ -lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; -x_66 = lean_ctor_get(x_35, 0); -x_67 = lean_ctor_get(x_35, 1); -lean_inc(x_67); -lean_inc(x_66); -lean_dec(x_35); -x_68 = l_Lean_Meta_Grind_ppENodeRef(x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_67); +lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; +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_6); +lean_dec(x_5); +lean_dec(x_3); lean_dec(x_2); -x_69 = lean_ctor_get(x_68, 0); -lean_inc(x_69); -x_70 = lean_ctor_get(x_68, 1); -lean_inc(x_70); -if (lean_is_exclusive(x_68)) { - lean_ctor_release(x_68, 0); - lean_ctor_release(x_68, 1); - x_71 = x_68; +lean_dec(x_1); +x_53 = lean_ctor_get(x_42, 0); +lean_inc(x_53); +x_54 = lean_ctor_get(x_42, 1); +lean_inc(x_54); +if (lean_is_exclusive(x_42)) { + lean_ctor_release(x_42, 0); + lean_ctor_release(x_42, 1); + x_55 = x_42; } else { - lean_dec_ref(x_68); - x_71 = lean_box(0); + lean_dec_ref(x_42); + x_55 = lean_box(0); } -x_72 = l_Lean_MessageData_ofFormat(x_66); -x_73 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__2; -if (lean_is_scalar(x_71)) { - x_74 = lean_alloc_ctor(7, 2, 0); +if (lean_is_scalar(x_55)) { + x_56 = lean_alloc_ctor(1, 2, 0); } else { - x_74 = x_71; - lean_ctor_set_tag(x_74, 7); + x_56 = x_55; } -lean_ctor_set(x_74, 0, x_73); -lean_ctor_set(x_74, 1, x_72); -x_75 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__9___closed__3; -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_MessageData_ofFormat(x_69); -lean_ctor_set_tag(x_27, 7); -lean_ctor_set(x_27, 1, x_77); -lean_ctor_set(x_27, 0, x_76); -x_78 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__9___closed__5; -x_79 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_79, 0, x_27); -lean_ctor_set(x_79, 1, x_78); -x_80 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_26, x_79, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_70); -x_81 = lean_ctor_get(x_80, 0); -lean_inc(x_81); -x_82 = lean_ctor_get(x_80, 1); -lean_inc(x_82); -lean_dec(x_80); -x_83 = lean_apply_10(x_31, x_81, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_82); -return x_83; +lean_ctor_set(x_56, 0, x_53); +lean_ctor_set(x_56, 1, x_54); +return x_56; } } } else { -lean_object* x_84; lean_object* x_85; lean_object* x_86; uint8_t x_87; -x_84 = lean_ctor_get(x_27, 0); -x_85 = lean_ctor_get(x_27, 1); -lean_inc(x_85); -lean_inc(x_84); -lean_dec(x_27); -x_86 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__9___closed__1; -x_87 = lean_unbox(x_84); -lean_dec(x_84); -if (x_87 == 0) +uint8_t x_57; +x_57 = !lean_is_exclusive(x_18); +if (x_57 == 0) { -lean_object* x_88; lean_object* x_89; -lean_dec(x_2); -lean_dec(x_1); -x_88 = lean_box(0); -x_89 = lean_apply_10(x_86, x_88, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_85); -return x_89; -} -else +lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_58 = lean_ctor_get(x_18, 1); +x_59 = lean_ctor_get(x_18, 0); +lean_dec(x_59); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_6); +lean_inc(x_5); +x_60 = l_Lean_Meta_mkHEq(x_5, x_6, x_12, x_13, x_14, x_15, x_58); +if (lean_obj_tag(x_60) == 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; -x_90 = l_Lean_Meta_Grind_ppENodeRef(x_1, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_85); -lean_dec(x_1); -x_91 = lean_ctor_get(x_90, 0); -lean_inc(x_91); -x_92 = lean_ctor_get(x_90, 1); -lean_inc(x_92); -if (lean_is_exclusive(x_90)) { - lean_ctor_release(x_90, 0); - lean_ctor_release(x_90, 1); - x_93 = x_90; -} else { - lean_dec_ref(x_90); - x_93 = lean_box(0); -} -x_94 = l_Lean_Meta_Grind_ppENodeRef(x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_92); -lean_dec(x_2); -x_95 = lean_ctor_get(x_94, 0); -lean_inc(x_95); -x_96 = lean_ctor_get(x_94, 1); -lean_inc(x_96); -if (lean_is_exclusive(x_94)) { - lean_ctor_release(x_94, 0); - lean_ctor_release(x_94, 1); - x_97 = x_94; -} else { - lean_dec_ref(x_94); - x_97 = lean_box(0); -} -x_98 = l_Lean_MessageData_ofFormat(x_91); -x_99 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__2; -if (lean_is_scalar(x_97)) { - x_100 = lean_alloc_ctor(7, 2, 0); -} else { - x_100 = x_97; - lean_ctor_set_tag(x_100, 7); -} -lean_ctor_set(x_100, 0, x_99); -lean_ctor_set(x_100, 1, x_98); -x_101 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__9___closed__3; -if (lean_is_scalar(x_93)) { - x_102 = lean_alloc_ctor(7, 2, 0); -} else { - x_102 = x_93; - lean_ctor_set_tag(x_102, 7); -} -lean_ctor_set(x_102, 0, x_100); -lean_ctor_set(x_102, 1, x_101); -x_103 = l_Lean_MessageData_ofFormat(x_95); -x_104 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_104, 0, x_102); -lean_ctor_set(x_104, 1, x_103); -x_105 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__9___closed__5; -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_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_26, x_106, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_96); -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_apply_10(x_86, x_108, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_109); -return x_110; -} -} -} +lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; +x_61 = lean_ctor_get(x_60, 0); +lean_inc(x_61); +x_62 = lean_ctor_get(x_60, 1); +lean_inc(x_62); +lean_dec(x_60); +x_63 = l_Lean_MessageData_ofExpr(x_61); +x_64 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; +lean_ctor_set_tag(x_18, 7); +lean_ctor_set(x_18, 1, x_63); +lean_ctor_set(x_18, 0, x_64); +x_65 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_65, 0, x_18); +lean_ctor_set(x_65, 1, x_64); +x_66 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_17, x_65, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, 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 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__7(x_1, x_2, x_3, x_4, x_5, x_6, x_67, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_68); +lean_dec(x_67); +return x_69; } else { -uint8_t x_111; -lean_dec(x_16); +uint8_t x_70; +lean_free_object(x_18); +lean_dec(x_15); +lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); +lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_111 = !lean_is_exclusive(x_18); -if (x_111 == 0) +x_70 = !lean_is_exclusive(x_60); +if (x_70 == 0) { -return x_18; +return x_60; } else { -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_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; +lean_object* x_71; lean_object* x_72; lean_object* x_73; +x_71 = lean_ctor_get(x_60, 0); +x_72 = lean_ctor_get(x_60, 1); +lean_inc(x_72); +lean_inc(x_71); +lean_dec(x_60); +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_74; lean_object* x_75; +x_74 = lean_ctor_get(x_18, 1); +lean_inc(x_74); +lean_dec(x_18); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_6); +lean_inc(x_5); +x_75 = l_Lean_Meta_mkHEq(x_5, x_6, x_12, x_13, x_14, x_15, x_74); +if (lean_obj_tag(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; +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_MessageData_ofExpr(x_76); +x_79 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; +x_80 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_80, 0, x_79); +lean_ctor_set(x_80, 1, x_78); +x_81 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_81, 0, x_80); +lean_ctor_set(x_81, 1, x_79); +x_82 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_17, x_81, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_77); +x_83 = lean_ctor_get(x_82, 0); +lean_inc(x_83); +x_84 = lean_ctor_get(x_82, 1); +lean_inc(x_84); +lean_dec(x_82); +x_85 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__7(x_1, x_2, x_3, x_4, x_5, x_6, x_83, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_84); +lean_dec(x_83); +return x_85; } else { -uint8_t x_115; +lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; +lean_dec(x_15); +lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); +lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_115 = !lean_is_exclusive(x_15); -if (x_115 == 0) -{ -return x_15; +x_86 = lean_ctor_get(x_75, 0); +lean_inc(x_86); +x_87 = lean_ctor_get(x_75, 1); +lean_inc(x_87); +if (lean_is_exclusive(x_75)) { + lean_ctor_release(x_75, 0); + lean_ctor_release(x_75, 1); + x_88 = x_75; +} else { + lean_dec_ref(x_75); + x_88 = lean_box(0); +} +if (lean_is_scalar(x_88)) { + x_89 = lean_alloc_ctor(1, 2, 0); +} else { + x_89 = x_88; +} +lean_ctor_set(x_89, 0, x_86); +lean_ctor_set(x_89, 1, x_87); +return x_89; +} } -else -{ -lean_object* x_116; lean_object* x_117; lean_object* x_118; -x_116 = lean_ctor_get(x_15, 0); -x_117 = lean_ctor_get(x_15, 1); -lean_inc(x_117); -lean_inc(x_116); -lean_dec(x_15); -x_118 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_118, 0, x_116); -lean_ctor_set(x_118, 1, x_117); -return x_118; } } } } +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___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) { +_start: +{ +lean_object* x_11; lean_object* x_12; +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; +} +} static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("eq", 2, 2); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__9___boxed), 10, 0); return x_1; } } static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__2() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__1; -x_2 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__1; -x_3 = l_Lean_Name_mkStr2(x_1, x_2); -return x_3; +lean_object* x_1; +x_1 = lean_mk_string_unchecked(" and ", 5, 5); +return x_1; } } static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__3() { _start: { +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__2; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__4() { +_start: +{ lean_object* x_1; -x_1 = lean_mk_string_unchecked(" ", 1, 1); +x_1 = lean_mk_string_unchecked(" are already in the same equivalence class", 42, 42); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__4() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__5() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__3; +x_1 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__4; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__5() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("=", 1, 1); -return x_1; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +lean_object* x_14; +lean_inc(x_1); +x_14 = l_Lean_Meta_Grind_getENode(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +x_16 = lean_ctor_get(x_14, 1); +lean_inc(x_16); +lean_dec(x_14); +lean_inc(x_2); +x_17 = l_Lean_Meta_Grind_getENode(x_2, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, 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; uint8_t 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_ctor_get(x_15, 2); +lean_inc(x_20); +x_21 = lean_ctor_get(x_18, 2); +lean_inc(x_21); +x_22 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_20, x_21); +lean_dec(x_21); +lean_dec(x_20); +if (x_22 == 0) +{ +lean_object* x_23; lean_object* x_24; +x_23 = lean_box(0); +x_24 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__8(x_15, x_18, x_3, x_4, x_1, x_2, x_23, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_19); +return x_24; +} +else +{ +lean_object* x_25; lean_object* x_26; uint8_t x_27; +lean_dec(x_18); +lean_dec(x_15); +lean_dec(x_3); +x_25 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__1; +x_26 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_25, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_19); +x_27 = !lean_is_exclusive(x_26); +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_26, 0); +x_29 = lean_ctor_get(x_26, 1); +x_30 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__1; +x_31 = lean_unbox(x_28); +lean_dec(x_28); +if (x_31 == 0) +{ +lean_object* x_32; lean_object* x_33; +lean_free_object(x_26); +lean_dec(x_2); +lean_dec(x_1); +x_32 = lean_box(0); +x_33 = lean_apply_10(x_30, x_32, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_29); +return x_33; +} +else +{ +lean_object* x_34; uint8_t x_35; +x_34 = l_Lean_Meta_Grind_ppENodeRef(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_29); +lean_dec(x_1); +x_35 = !lean_is_exclusive(x_34); +if (x_35 == 0) +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; uint8_t x_39; +x_36 = lean_ctor_get(x_34, 0); +x_37 = lean_ctor_get(x_34, 1); +x_38 = l_Lean_Meta_Grind_ppENodeRef(x_2, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_37); +lean_dec(x_2); +x_39 = !lean_is_exclusive(x_38); +if (x_39 == 0) +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_40 = lean_ctor_get(x_38, 0); +x_41 = lean_ctor_get(x_38, 1); +x_42 = l_Lean_MessageData_ofFormat(x_36); +x_43 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; +lean_ctor_set_tag(x_38, 7); +lean_ctor_set(x_38, 1, x_42); +lean_ctor_set(x_38, 0, x_43); +x_44 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__3; +lean_ctor_set_tag(x_34, 7); +lean_ctor_set(x_34, 1, x_44); +lean_ctor_set(x_34, 0, x_38); +x_45 = l_Lean_MessageData_ofFormat(x_40); +lean_ctor_set_tag(x_26, 7); +lean_ctor_set(x_26, 1, x_45); +lean_ctor_set(x_26, 0, x_34); +x_46 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__5; +x_47 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_47, 0, x_26); +lean_ctor_set(x_47, 1, x_46); +x_48 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_25, x_47, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_41); +x_49 = lean_ctor_get(x_48, 0); +lean_inc(x_49); +x_50 = lean_ctor_get(x_48, 1); +lean_inc(x_50); +lean_dec(x_48); +x_51 = lean_apply_10(x_30, x_49, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_50); +return x_51; +} +else +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; +x_52 = lean_ctor_get(x_38, 0); +x_53 = lean_ctor_get(x_38, 1); +lean_inc(x_53); +lean_inc(x_52); +lean_dec(x_38); +x_54 = l_Lean_MessageData_ofFormat(x_36); +x_55 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; +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___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__3; +lean_ctor_set_tag(x_34, 7); +lean_ctor_set(x_34, 1, x_57); +lean_ctor_set(x_34, 0, x_56); +x_58 = l_Lean_MessageData_ofFormat(x_52); +lean_ctor_set_tag(x_26, 7); +lean_ctor_set(x_26, 1, x_58); +lean_ctor_set(x_26, 0, x_34); +x_59 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__5; +x_60 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_60, 0, x_26); +lean_ctor_set(x_60, 1, x_59); +x_61 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_25, x_60, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_53); +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_apply_10(x_30, x_62, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_63); +return x_64; +} +} +else +{ +lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; +x_65 = lean_ctor_get(x_34, 0); +x_66 = lean_ctor_get(x_34, 1); +lean_inc(x_66); +lean_inc(x_65); +lean_dec(x_34); +x_67 = l_Lean_Meta_Grind_ppENodeRef(x_2, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_66); +lean_dec(x_2); +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_67); + x_70 = lean_box(0); +} +x_71 = l_Lean_MessageData_ofFormat(x_65); +x_72 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; +if (lean_is_scalar(x_70)) { + x_73 = lean_alloc_ctor(7, 2, 0); +} else { + x_73 = x_70; + lean_ctor_set_tag(x_73, 7); +} +lean_ctor_set(x_73, 0, x_72); +lean_ctor_set(x_73, 1, x_71); +x_74 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__3; +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_MessageData_ofFormat(x_68); +lean_ctor_set_tag(x_26, 7); +lean_ctor_set(x_26, 1, x_76); +lean_ctor_set(x_26, 0, x_75); +x_77 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__5; +x_78 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_78, 0, x_26); +lean_ctor_set(x_78, 1, x_77); +x_79 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_25, x_78, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_69); +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_apply_10(x_30, x_80, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_81); +return x_82; +} } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__6() { -_start: +else { -lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__5; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; -} +lean_object* x_83; lean_object* x_84; lean_object* x_85; uint8_t x_86; +x_83 = lean_ctor_get(x_26, 0); +x_84 = lean_ctor_get(x_26, 1); +lean_inc(x_84); +lean_inc(x_83); +lean_dec(x_26); +x_85 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__1; +x_86 = lean_unbox(x_83); +lean_dec(x_83); +if (x_86 == 0) +{ +lean_object* x_87; lean_object* x_88; +lean_dec(x_2); +lean_dec(x_1); +x_87 = lean_box(0); +x_88 = lean_apply_10(x_85, x_87, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_84); +return x_88; } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__7() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("≡", 3, 1); -return x_1; +lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; +x_89 = l_Lean_Meta_Grind_ppENodeRef(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_84); +lean_dec(x_1); +x_90 = lean_ctor_get(x_89, 0); +lean_inc(x_90); +x_91 = lean_ctor_get(x_89, 1); +lean_inc(x_91); +if (lean_is_exclusive(x_89)) { + lean_ctor_release(x_89, 0); + lean_ctor_release(x_89, 1); + x_92 = x_89; +} else { + lean_dec_ref(x_89); + x_92 = lean_box(0); +} +x_93 = l_Lean_Meta_Grind_ppENodeRef(x_2, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_91); +lean_dec(x_2); +x_94 = lean_ctor_get(x_93, 0); +lean_inc(x_94); +x_95 = lean_ctor_get(x_93, 1); +lean_inc(x_95); +if (lean_is_exclusive(x_93)) { + lean_ctor_release(x_93, 0); + lean_ctor_release(x_93, 1); + x_96 = x_93; +} else { + lean_dec_ref(x_93); + x_96 = lean_box(0); } +x_97 = l_Lean_MessageData_ofFormat(x_90); +x_98 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; +if (lean_is_scalar(x_96)) { + x_99 = lean_alloc_ctor(7, 2, 0); +} else { + x_99 = x_96; + lean_ctor_set_tag(x_99, 7); +} +lean_ctor_set(x_99, 0, x_98); +lean_ctor_set(x_99, 1, x_97); +x_100 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__3; +if (lean_is_scalar(x_92)) { + x_101 = lean_alloc_ctor(7, 2, 0); +} else { + x_101 = x_92; + lean_ctor_set_tag(x_101, 7); +} +lean_ctor_set(x_101, 0, x_99); +lean_ctor_set(x_101, 1, x_100); +x_102 = l_Lean_MessageData_ofFormat(x_94); +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___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__5; +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_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_25, x_105, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_95); +x_107 = lean_ctor_get(x_106, 0); +lean_inc(x_107); +x_108 = lean_ctor_get(x_106, 1); +lean_inc(x_108); +lean_dec(x_106); +x_109 = lean_apply_10(x_85, x_107, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_108); +return x_109; } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__8() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__7; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { -_start: -{ -lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; -x_14 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__2; -x_15 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_14, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -x_16 = lean_ctor_get(x_15, 0); -lean_inc(x_16); -x_17 = lean_unbox(x_16); -lean_dec(x_16); -if (x_17 == 0) -{ -lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_18 = lean_ctor_get(x_15, 1); -lean_inc(x_18); -lean_dec(x_15); -x_19 = lean_box(0); -x_20 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__9(x_1, x_2, x_3, x_4, x_19, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_18); -return x_20; } else { -uint8_t x_21; -x_21 = !lean_is_exclusive(x_15); -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; -x_22 = lean_ctor_get(x_15, 1); -x_23 = lean_ctor_get(x_15, 0); -lean_dec(x_23); -lean_inc(x_1); -x_24 = l_Lean_MessageData_ofExpr(x_1); -x_25 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__2; -lean_ctor_set_tag(x_15, 7); -lean_ctor_set(x_15, 1, x_24); -lean_ctor_set(x_15, 0, x_25); -x_26 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__4; -x_27 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_27, 0, x_15); -lean_ctor_set(x_27, 1, x_26); -lean_inc(x_2); -x_28 = l_Lean_MessageData_ofExpr(x_2); -if (x_4 == 0) +uint8_t x_110; +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); +lean_dec(x_1); +x_110 = !lean_is_exclusive(x_17); +if (x_110 == 0) { -lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; -x_29 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__6; -x_30 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_30, 0, x_27); -lean_ctor_set(x_30, 1, x_29); -x_31 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_31, 0, x_30); -lean_ctor_set(x_31, 1, x_26); -x_32 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_32, 0, x_31); -lean_ctor_set(x_32, 1, x_28); -x_33 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_33, 0, x_32); -lean_ctor_set(x_33, 1, x_25); -x_34 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_14, x_33, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_22); -x_35 = lean_ctor_get(x_34, 0); -lean_inc(x_35); -x_36 = lean_ctor_get(x_34, 1); -lean_inc(x_36); -lean_dec(x_34); -x_37 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__9(x_1, x_2, x_3, x_4, x_35, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_36); -lean_dec(x_35); -return x_37; +return x_17; } else { -lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; -x_38 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__8; -x_39 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_39, 0, x_27); -lean_ctor_set(x_39, 1, x_38); -x_40 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_40, 0, x_39); -lean_ctor_set(x_40, 1, x_26); -x_41 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_41, 0, x_40); -lean_ctor_set(x_41, 1, x_28); -x_42 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_42, 0, x_41); -lean_ctor_set(x_42, 1, x_25); -x_43 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_14, x_42, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_22); -x_44 = lean_ctor_get(x_43, 0); -lean_inc(x_44); -x_45 = lean_ctor_get(x_43, 1); -lean_inc(x_45); -lean_dec(x_43); -x_46 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__9(x_1, x_2, x_3, x_4, x_44, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_45); -lean_dec(x_44); -return x_46; +lean_object* x_111; lean_object* x_112; lean_object* x_113; +x_111 = lean_ctor_get(x_17, 0); +x_112 = lean_ctor_get(x_17, 1); +lean_inc(x_112); +lean_inc(x_111); +lean_dec(x_17); +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_47; 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_47 = lean_ctor_get(x_15, 1); -lean_inc(x_47); -lean_dec(x_15); -lean_inc(x_1); -x_48 = l_Lean_MessageData_ofExpr(x_1); -x_49 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__2; -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___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__4; -x_52 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_52, 0, x_50); -lean_ctor_set(x_52, 1, x_51); -lean_inc(x_2); -x_53 = l_Lean_MessageData_ofExpr(x_2); -if (x_4 == 0) +uint8_t x_114; +lean_dec(x_12); +lean_dec(x_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_114 = !lean_is_exclusive(x_14); +if (x_114 == 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 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__6; -x_55 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_55, 0, x_52); -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_51); -x_57 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_57, 0, x_56); -lean_ctor_set(x_57, 1, x_53); -x_58 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_58, 0, x_57); -lean_ctor_set(x_58, 1, x_49); -x_59 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_14, x_58, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_47); -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___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__9(x_1, x_2, x_3, x_4, x_60, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_61); -lean_dec(x_60); -return x_62; +return x_14; } 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; -x_63 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__8; -x_64 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_64, 0, x_52); -lean_ctor_set(x_64, 1, x_63); -x_65 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_65, 0, x_64); -lean_ctor_set(x_65, 1, x_51); -x_66 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_66, 0, x_65); -lean_ctor_set(x_66, 1, x_53); -x_67 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_67, 0, x_66); -lean_ctor_set(x_67, 1, x_49); -x_68 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_14, x_67, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_47); -x_69 = lean_ctor_get(x_68, 0); -lean_inc(x_69); -x_70 = lean_ctor_get(x_68, 1); -lean_inc(x_70); -lean_dec(x_68); -x_71 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__9(x_1, x_2, x_3, x_4, x_69, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_70); -lean_dec(x_69); -return x_71; -} +lean_object* x_115; lean_object* x_116; lean_object* x_117; +x_115 = lean_ctor_get(x_14, 0); +x_116 = lean_ctor_get(x_14, 1); +lean_inc(x_116); +lean_inc(x_115); +lean_dec(x_14); +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; } } } @@ -7637,11 +8423,22 @@ lean_dec(x_7); return x_18; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___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_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___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, lean_object* x_15, lean_object* x_16) { +_start: +{ +uint8_t x_17; lean_object* x_18; +x_17 = lean_unbox(x_4); +lean_dec(x_4); +x_18 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__8(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_7); +return x_18; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___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) { _start: { lean_object* x_11; -x_11 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__8(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_11 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__9(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -7654,17 +8451,6 @@ lean_dec(x_1); return x_11; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__9___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { -_start: -{ -uint8_t x_15; lean_object* x_16; -x_15 = lean_unbox(x_4); -lean_dec(x_4); -x_16 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__9(x_1, x_2, x_3, x_15, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -lean_dec(x_5); -return x_16; -} -} LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, 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: { @@ -7728,15 +8514,27 @@ return x_22; } else { -lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; 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_23 = lean_ctor_get(x_11, 0); x_24 = lean_ctor_get(x_11, 1); x_25 = lean_ctor_get(x_11, 2); x_26 = lean_ctor_get(x_11, 3); x_27 = lean_ctor_get(x_11, 4); -x_28 = lean_ctor_get_uint8(x_11, sizeof(void*)*8); +x_28 = lean_ctor_get_uint8(x_11, sizeof(void*)*14); x_29 = lean_ctor_get(x_11, 6); x_30 = lean_ctor_get(x_11, 7); +x_31 = lean_ctor_get(x_11, 8); +x_32 = lean_ctor_get(x_11, 9); +x_33 = lean_ctor_get(x_11, 10); +x_34 = lean_ctor_get(x_11, 11); +x_35 = lean_ctor_get(x_11, 12); +x_36 = lean_ctor_get(x_11, 13); +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_27); @@ -7745,37 +8543,43 @@ lean_inc(x_25); lean_inc(x_24); lean_inc(x_23); lean_dec(x_11); -x_31 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_resetNewEqs___closed__1; -x_32 = lean_alloc_ctor(0, 8, 1); -lean_ctor_set(x_32, 0, x_23); -lean_ctor_set(x_32, 1, x_24); -lean_ctor_set(x_32, 2, x_25); -lean_ctor_set(x_32, 3, x_26); -lean_ctor_set(x_32, 4, x_27); -lean_ctor_set(x_32, 5, x_31); -lean_ctor_set(x_32, 6, x_29); -lean_ctor_set(x_32, 7, x_30); -lean_ctor_set_uint8(x_32, sizeof(void*)*8, x_28); -x_33 = lean_st_ref_set(x_1, x_32, x_12); -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_37 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_resetNewEqs___closed__1; +x_38 = lean_alloc_ctor(0, 14, 1); +lean_ctor_set(x_38, 0, x_23); +lean_ctor_set(x_38, 1, x_24); +lean_ctor_set(x_38, 2, x_25); +lean_ctor_set(x_38, 3, x_26); +lean_ctor_set(x_38, 4, x_27); +lean_ctor_set(x_38, 5, x_37); +lean_ctor_set(x_38, 6, x_29); +lean_ctor_set(x_38, 7, x_30); +lean_ctor_set(x_38, 8, x_31); +lean_ctor_set(x_38, 9, x_32); +lean_ctor_set(x_38, 10, x_33); +lean_ctor_set(x_38, 11, x_34); +lean_ctor_set(x_38, 12, x_35); +lean_ctor_set(x_38, 13, x_36); +lean_ctor_set_uint8(x_38, sizeof(void*)*14, x_28); +x_39 = lean_st_ref_set(x_1, x_38, x_12); +x_40 = lean_ctor_get(x_39, 1); +lean_inc(x_40); +if (lean_is_exclusive(x_39)) { + lean_ctor_release(x_39, 0); + lean_ctor_release(x_39, 1); + x_41 = x_39; } else { - lean_dec_ref(x_33); - x_35 = lean_box(0); + lean_dec_ref(x_39); + x_41 = lean_box(0); } -x_36 = lean_box(0); -if (lean_is_scalar(x_35)) { - x_37 = lean_alloc_ctor(0, 2, 0); +x_42 = lean_box(0); +if (lean_is_scalar(x_41)) { + x_43 = lean_alloc_ctor(0, 2, 0); } else { - x_37 = x_35; + x_43 = x_41; } -lean_ctor_set(x_37, 0, x_36); -lean_ctor_set(x_37, 1, x_34); -return x_37; +lean_ctor_set(x_43, 0, x_42); +lean_ctor_set(x_43, 1, x_40); +return x_43; } } } @@ -7867,16 +8671,28 @@ 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; 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; +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; lean_object* x_41; 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_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); x_32 = lean_ctor_get(x_17, 5); -x_33 = lean_ctor_get_uint8(x_17, sizeof(void*)*8); +x_33 = lean_ctor_get_uint8(x_17, sizeof(void*)*14); x_34 = lean_ctor_get(x_17, 6); x_35 = lean_ctor_get(x_17, 7); +x_36 = lean_ctor_get(x_17, 8); +x_37 = lean_ctor_get(x_17, 9); +x_38 = lean_ctor_get(x_17, 10); +x_39 = lean_ctor_get(x_17, 11); +x_40 = lean_ctor_get(x_17, 12); +x_41 = lean_ctor_get(x_17, 13); +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); lean_inc(x_32); @@ -7886,134 +8702,164 @@ lean_inc(x_29); lean_inc(x_28); lean_inc(x_27); lean_dec(x_17); -x_36 = lean_array_pop(x_32); -x_37 = lean_alloc_ctor(0, 8, 1); -lean_ctor_set(x_37, 0, x_27); -lean_ctor_set(x_37, 1, x_28); -lean_ctor_set(x_37, 2, x_29); -lean_ctor_set(x_37, 3, x_30); -lean_ctor_set(x_37, 4, x_31); -lean_ctor_set(x_37, 5, x_36); -lean_ctor_set(x_37, 6, x_34); -lean_ctor_set(x_37, 7, x_35); -lean_ctor_set_uint8(x_37, sizeof(void*)*8, x_33); -x_38 = lean_st_ref_set(x_1, x_37, x_18); -x_39 = lean_ctor_get(x_38, 1); -lean_inc(x_39); -if (lean_is_exclusive(x_38)) { - lean_ctor_release(x_38, 0); - lean_ctor_release(x_38, 1); - x_40 = x_38; +x_42 = lean_array_pop(x_32); +x_43 = lean_alloc_ctor(0, 14, 1); +lean_ctor_set(x_43, 0, x_27); +lean_ctor_set(x_43, 1, x_28); +lean_ctor_set(x_43, 2, x_29); +lean_ctor_set(x_43, 3, x_30); +lean_ctor_set(x_43, 4, x_31); +lean_ctor_set(x_43, 5, x_42); +lean_ctor_set(x_43, 6, x_34); +lean_ctor_set(x_43, 7, x_35); +lean_ctor_set(x_43, 8, x_36); +lean_ctor_set(x_43, 9, x_37); +lean_ctor_set(x_43, 10, x_38); +lean_ctor_set(x_43, 11, x_39); +lean_ctor_set(x_43, 12, x_40); +lean_ctor_set(x_43, 13, x_41); +lean_ctor_set_uint8(x_43, sizeof(void*)*14, x_33); +x_44 = lean_st_ref_set(x_1, x_43, x_18); +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_38); - x_40 = lean_box(0); + lean_dec_ref(x_44); + x_46 = lean_box(0); } -if (lean_is_scalar(x_40)) { - x_41 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_46)) { + x_47 = lean_alloc_ctor(0, 2, 0); } else { - x_41 = x_40; + x_47 = x_46; } -lean_ctor_set(x_41, 0, x_15); -lean_ctor_set(x_41, 1, x_39); -return x_41; +lean_ctor_set(x_47, 0, x_15); +lean_ctor_set(x_47, 1, x_45); +return x_47; } } } else { -lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; -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_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_48 = lean_ctor_get(x_10, 0); +x_49 = lean_ctor_get(x_10, 1); +lean_inc(x_49); +lean_inc(x_48); lean_dec(x_10); -x_44 = lean_ctor_get(x_42, 5); -lean_inc(x_44); -lean_dec(x_42); -x_45 = l_Array_back_x3f___rarg(x_44); -lean_dec(x_44); -if (lean_obj_tag(x_45) == 0) +x_50 = lean_ctor_get(x_48, 5); +lean_inc(x_50); +lean_dec(x_48); +x_51 = l_Array_back_x3f___rarg(x_50); +lean_dec(x_50); +if (lean_obj_tag(x_51) == 0) { -lean_object* x_46; -x_46 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_46, 0, x_45); -lean_ctor_set(x_46, 1, x_43); -return x_46; +lean_object* x_52; +x_52 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_52, 0, x_51); +lean_ctor_set(x_52, 1, x_49); +return x_52; } 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; 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; -x_47 = lean_st_ref_take(x_1, x_43); -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); -x_51 = lean_ctor_get(x_48, 1); -lean_inc(x_51); -x_52 = lean_ctor_get(x_48, 2); -lean_inc(x_52); -x_53 = lean_ctor_get(x_48, 3); -lean_inc(x_53); -x_54 = lean_ctor_get(x_48, 4); +lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; uint8_t x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; +x_53 = lean_st_ref_take(x_1, x_49); +x_54 = lean_ctor_get(x_53, 0); lean_inc(x_54); -x_55 = lean_ctor_get(x_48, 5); +x_55 = lean_ctor_get(x_53, 1); lean_inc(x_55); -x_56 = lean_ctor_get_uint8(x_48, sizeof(void*)*8); -x_57 = lean_ctor_get(x_48, 6); +lean_dec(x_53); +x_56 = lean_ctor_get(x_54, 0); +lean_inc(x_56); +x_57 = lean_ctor_get(x_54, 1); lean_inc(x_57); -x_58 = lean_ctor_get(x_48, 7); +x_58 = lean_ctor_get(x_54, 2); lean_inc(x_58); -if (lean_is_exclusive(x_48)) { - lean_ctor_release(x_48, 0); - lean_ctor_release(x_48, 1); - lean_ctor_release(x_48, 2); - lean_ctor_release(x_48, 3); - lean_ctor_release(x_48, 4); - lean_ctor_release(x_48, 5); - lean_ctor_release(x_48, 6); - lean_ctor_release(x_48, 7); - x_59 = x_48; +x_59 = lean_ctor_get(x_54, 3); +lean_inc(x_59); +x_60 = lean_ctor_get(x_54, 4); +lean_inc(x_60); +x_61 = lean_ctor_get(x_54, 5); +lean_inc(x_61); +x_62 = lean_ctor_get_uint8(x_54, sizeof(void*)*14); +x_63 = lean_ctor_get(x_54, 6); +lean_inc(x_63); +x_64 = lean_ctor_get(x_54, 7); +lean_inc(x_64); +x_65 = lean_ctor_get(x_54, 8); +lean_inc(x_65); +x_66 = lean_ctor_get(x_54, 9); +lean_inc(x_66); +x_67 = lean_ctor_get(x_54, 10); +lean_inc(x_67); +x_68 = lean_ctor_get(x_54, 11); +lean_inc(x_68); +x_69 = lean_ctor_get(x_54, 12); +lean_inc(x_69); +x_70 = lean_ctor_get(x_54, 13); +lean_inc(x_70); +if (lean_is_exclusive(x_54)) { + lean_ctor_release(x_54, 0); + lean_ctor_release(x_54, 1); + lean_ctor_release(x_54, 2); + lean_ctor_release(x_54, 3); + lean_ctor_release(x_54, 4); + lean_ctor_release(x_54, 5); + lean_ctor_release(x_54, 6); + lean_ctor_release(x_54, 7); + lean_ctor_release(x_54, 8); + lean_ctor_release(x_54, 9); + lean_ctor_release(x_54, 10); + lean_ctor_release(x_54, 11); + lean_ctor_release(x_54, 12); + lean_ctor_release(x_54, 13); + x_71 = x_54; } else { - lean_dec_ref(x_48); - x_59 = lean_box(0); + lean_dec_ref(x_54); + x_71 = lean_box(0); } -x_60 = lean_array_pop(x_55); -if (lean_is_scalar(x_59)) { - x_61 = lean_alloc_ctor(0, 8, 1); +x_72 = lean_array_pop(x_61); +if (lean_is_scalar(x_71)) { + x_73 = lean_alloc_ctor(0, 14, 1); } else { - x_61 = x_59; -} -lean_ctor_set(x_61, 0, x_50); -lean_ctor_set(x_61, 1, x_51); -lean_ctor_set(x_61, 2, x_52); -lean_ctor_set(x_61, 3, x_53); -lean_ctor_set(x_61, 4, x_54); -lean_ctor_set(x_61, 5, x_60); -lean_ctor_set(x_61, 6, x_57); -lean_ctor_set(x_61, 7, x_58); -lean_ctor_set_uint8(x_61, sizeof(void*)*8, x_56); -x_62 = lean_st_ref_set(x_1, x_61, x_49); -x_63 = lean_ctor_get(x_62, 1); -lean_inc(x_63); -if (lean_is_exclusive(x_62)) { - lean_ctor_release(x_62, 0); - lean_ctor_release(x_62, 1); - x_64 = x_62; + x_73 = x_71; +} +lean_ctor_set(x_73, 0, x_56); +lean_ctor_set(x_73, 1, x_57); +lean_ctor_set(x_73, 2, x_58); +lean_ctor_set(x_73, 3, x_59); +lean_ctor_set(x_73, 4, x_60); +lean_ctor_set(x_73, 5, x_72); +lean_ctor_set(x_73, 6, x_63); +lean_ctor_set(x_73, 7, x_64); +lean_ctor_set(x_73, 8, x_65); +lean_ctor_set(x_73, 9, x_66); +lean_ctor_set(x_73, 10, x_67); +lean_ctor_set(x_73, 11, x_68); +lean_ctor_set(x_73, 12, x_69); +lean_ctor_set(x_73, 13, x_70); +lean_ctor_set_uint8(x_73, sizeof(void*)*14, x_62); +x_74 = lean_st_ref_set(x_1, x_73, x_55); +x_75 = lean_ctor_get(x_74, 1); +lean_inc(x_75); +if (lean_is_exclusive(x_74)) { + lean_ctor_release(x_74, 0); + lean_ctor_release(x_74, 1); + x_76 = x_74; } else { - lean_dec_ref(x_62); - x_64 = lean_box(0); + lean_dec_ref(x_74); + x_76 = lean_box(0); } -if (lean_is_scalar(x_64)) { - x_65 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_76)) { + x_77 = lean_alloc_ctor(0, 2, 0); } else { - x_65 = x_64; + x_77 = x_76; } -lean_ctor_set(x_65, 0, x_45); -lean_ctor_set(x_65, 1, x_63); -return x_65; +lean_ctor_set(x_77, 0, x_51); +lean_ctor_set(x_77, 1, x_75); +return x_77; } } } @@ -8055,12 +8901,20 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Gr _start: { lean_object* x_11; lean_object* x_12; -x_11 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_popNextEq_x3f(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); +x_11 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__1; +x_12 = l_Lean_Core_checkSystem(x_11, x_8, x_9, x_10); if (lean_obj_tag(x_12) == 0) { -uint8_t x_13; +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___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_popNextEq_x3f(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_13); +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +if (lean_obj_tag(x_15) == 0) +{ +uint8_t x_16; lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -8069,46 +8923,46 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_13 = !lean_is_exclusive(x_11); -if (x_13 == 0) +x_16 = !lean_is_exclusive(x_14); +if (x_16 == 0) { -lean_object* x_14; lean_object* x_15; -x_14 = lean_ctor_get(x_11, 0); -lean_dec(x_14); -x_15 = lean_box(0); -lean_ctor_set(x_11, 0, x_15); -return x_11; +lean_object* x_17; lean_object* x_18; +x_17 = lean_ctor_get(x_14, 0); +lean_dec(x_17); +x_18 = lean_box(0); +lean_ctor_set(x_14, 0, x_18); +return x_14; } else { -lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_16 = lean_ctor_get(x_11, 1); -lean_inc(x_16); -lean_dec(x_11); -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_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_14, 1); +lean_inc(x_19); +lean_dec(x_14); +x_20 = lean_box(0); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_20); +lean_ctor_set(x_21, 1, x_19); +return x_21; } } else { -lean_object* x_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; -x_19 = lean_ctor_get(x_12, 0); -lean_inc(x_19); -lean_dec(x_12); -x_20 = lean_ctor_get(x_11, 1); -lean_inc(x_20); -lean_dec(x_11); -x_21 = lean_ctor_get(x_19, 0); -lean_inc(x_21); -x_22 = lean_ctor_get(x_19, 1); +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; +x_22 = lean_ctor_get(x_15, 0); lean_inc(x_22); -x_23 = lean_ctor_get(x_19, 2); +lean_dec(x_15); +x_23 = lean_ctor_get(x_14, 1); lean_inc(x_23); -x_24 = lean_ctor_get_uint8(x_19, sizeof(void*)*3); -lean_dec(x_19); +lean_dec(x_14); +x_24 = lean_ctor_get(x_22, 0); +lean_inc(x_24); +x_25 = lean_ctor_get(x_22, 1); +lean_inc(x_25); +x_26 = lean_ctor_get(x_22, 2); +lean_inc(x_26); +x_27 = lean_ctor_get_uint8(x_22, sizeof(void*)*3); +lean_dec(x_22); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); @@ -8117,19 +8971,19 @@ lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); -x_25 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep(x_21, x_22, x_23, x_24, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_20); -if (lean_obj_tag(x_25) == 0) +x_28 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep(x_24, x_25, x_26, x_27, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_23); +if (lean_obj_tag(x_28) == 0) { -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___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqCore_processTodo(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_26); -return x_27; +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___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqCore_processTodo(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_29); +return x_30; } else { -uint8_t x_28; +uint8_t x_31; lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -8138,25 +8992,56 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_28 = !lean_is_exclusive(x_25); -if (x_28 == 0) +x_31 = !lean_is_exclusive(x_28); +if (x_31 == 0) { -return x_25; +return x_28; } else { -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_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; +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_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; +} +} } } +else +{ +uint8_t x_35; +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_35 = !lean_is_exclusive(x_12); +if (x_35 == 0) +{ +return x_12; +} +else +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_12, 0); +x_37 = lean_ctor_get(x_12, 1); +lean_inc(x_37); +lean_inc(x_36); +lean_dec(x_12); +x_38 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_38, 0, x_36); +lean_ctor_set(x_38, 1, x_37); +return x_38; +} } } } @@ -8697,451 +9582,193 @@ x_38 = !lean_is_exclusive(x_33); if (x_38 == 0) { return x_33; -} -else -{ -lean_object* x_39; lean_object* x_40; lean_object* x_41; -x_39 = lean_ctor_get(x_33, 0); -x_40 = lean_ctor_get(x_33, 1); -lean_inc(x_40); -lean_inc(x_39); -lean_dec(x_33); -x_41 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_41, 0, x_39); -lean_ctor_set(x_41, 1, x_40); -return x_41; -} -} -} -} -else -{ -uint8_t x_42; -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_4); -lean_dec(x_1); -x_42 = !lean_is_exclusive(x_15); -if (x_42 == 0) -{ -return x_15; -} -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_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; -} -} -} -} -static lean_object* _init_l_Lean_Meta_Grind_add_go___lambda__4___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_closeGoalWithValuesEq___closed__4; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Meta_Grind_add_go___lambda__4___closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("HEq", 3, 3); -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_Grind_add_go___lambda__4___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_Grind_add_go___lambda__4___closed__2; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_add_go___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { -_start: -{ -lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; -lean_inc(x_1); -x_15 = l_Lean_Meta_instantiateMVarsIfMVarApp(x_1, x_10, x_11, x_12, x_13, x_14); -x_16 = lean_ctor_get(x_15, 0); -lean_inc(x_16); -x_17 = lean_ctor_get(x_15, 1); -lean_inc(x_17); -lean_dec(x_15); -x_18 = l_Lean_Expr_cleanupAnnotations(x_16); -x_19 = l_Lean_Expr_isApp(x_18); -if (x_19 == 0) -{ -lean_object* x_20; lean_object* x_21; -lean_dec(x_18); -x_20 = lean_box(0); -x_21 = l_Lean_Meta_Grind_add_go___lambda__3(x_1, x_3, x_4, x_2, x_20, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_17); -return x_21; -} -else -{ -lean_object* x_22; lean_object* x_23; uint8_t x_24; -x_22 = l_Lean_Expr_appArg(x_18, lean_box(0)); -x_23 = l_Lean_Expr_appFnCleanup(x_18, lean_box(0)); -x_24 = l_Lean_Expr_isApp(x_23); -if (x_24 == 0) -{ -lean_object* x_25; lean_object* x_26; -lean_dec(x_23); -lean_dec(x_22); -x_25 = lean_box(0); -x_26 = l_Lean_Meta_Grind_add_go___lambda__3(x_1, x_3, x_4, x_2, x_25, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_17); -return x_26; -} -else -{ -lean_object* x_27; lean_object* x_28; uint8_t x_29; -x_27 = l_Lean_Expr_appArg(x_23, lean_box(0)); -x_28 = l_Lean_Expr_appFnCleanup(x_23, lean_box(0)); -x_29 = l_Lean_Expr_isApp(x_28); -if (x_29 == 0) -{ -lean_object* x_30; lean_object* x_31; -lean_dec(x_28); -lean_dec(x_27); -lean_dec(x_22); -x_30 = lean_box(0); -x_31 = l_Lean_Meta_Grind_add_go___lambda__3(x_1, x_3, x_4, x_2, x_30, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_17); -return x_31; -} -else -{ -lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; -x_32 = l_Lean_Expr_appArg(x_28, lean_box(0)); -x_33 = l_Lean_Expr_appFnCleanup(x_28, lean_box(0)); -x_34 = l_Lean_Meta_Grind_add_go___lambda__4___closed__1; -x_35 = l_Lean_Expr_isConstOf(x_33, x_34); -if (x_35 == 0) -{ -uint8_t x_36; -lean_dec(x_27); -x_36 = l_Lean_Expr_isApp(x_33); -if (x_36 == 0) -{ -lean_object* x_37; lean_object* x_38; -lean_dec(x_33); -lean_dec(x_32); -lean_dec(x_22); -x_37 = lean_box(0); -x_38 = l_Lean_Meta_Grind_add_go___lambda__3(x_1, x_3, x_4, x_2, x_37, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_17); -return x_38; -} -else -{ -lean_object* x_39; lean_object* x_40; uint8_t x_41; -x_39 = l_Lean_Expr_appFnCleanup(x_33, lean_box(0)); -x_40 = l_Lean_Meta_Grind_add_go___lambda__4___closed__3; -x_41 = l_Lean_Expr_isConstOf(x_39, x_40); -lean_dec(x_39); -if (x_41 == 0) -{ -lean_object* x_42; lean_object* x_43; -lean_dec(x_32); -lean_dec(x_22); -x_42 = lean_box(0); -x_43 = l_Lean_Meta_Grind_add_go___lambda__3(x_1, x_3, x_4, x_2, x_42, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_17); -return x_43; -} -else -{ -uint8_t x_44; lean_object* x_45; -x_44 = 1; -x_45 = l_Lean_Meta_Grind_add_goEq(x_2, x_3, x_1, x_32, x_22, x_4, x_44, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_17); -return x_45; -} -} -} -else -{ -uint8_t x_46; lean_object* x_47; -lean_dec(x_33); -lean_dec(x_32); -x_46 = 0; -x_47 = l_Lean_Meta_Grind_add_goEq(x_2, x_3, x_1, x_27, x_22, x_4, x_46, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_17); -return x_47; -} -} -} -} -} -} -static lean_object* _init_l_Lean_Meta_Grind_add_go___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("add", 3, 3); -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_Grind_add_go___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__1; -x_2 = l_Lean_Meta_Grind_add_go___closed__1; -x_3 = l_Lean_Name_mkStr2(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Meta_Grind_add_go___closed__3() { -_start: +} +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("isNeg: ", 7, 7); -return x_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_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_Grind_add_go___closed__4() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_Grind_add_go___closed__3; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; } } -static lean_object* _init_l_Lean_Meta_Grind_add_go___closed__5() { -_start: +else { -lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_closeGoalWithValuesEq___closed__13; -x_2 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; -} +uint8_t x_42; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_1); +x_42 = !lean_is_exclusive(x_15); +if (x_42 == 0) +{ +return x_15; } -static lean_object* _init_l_Lean_Meta_Grind_add_go___closed__6() { -_start: +else { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_Grind_add_go___closed__5; -x_2 = l_Lean_MessageData_ofFormat(x_1); -return x_2; +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_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; } } -static lean_object* _init_l_Lean_Meta_Grind_add_go___closed__7() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_Grind_add_go___closed__4; -x_2 = l_Lean_Meta_Grind_add_go___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_Lean_Meta_Grind_add_go___closed__8() { +static lean_object* _init_l_Lean_Meta_Grind_add_go___closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_Grind_add_go___closed__7; -x_2 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___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); +x_1 = lean_box(0); +x_2 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_closeGoalWithValuesEq___closed__4; +x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_Grind_add_go___closed__9() { +static lean_object* _init_l_Lean_Meta_Grind_add_go___closed__2() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("true", 4, 4); +x_1 = lean_mk_string_unchecked("HEq", 3, 3); return x_1; } } -static lean_object* _init_l_Lean_Meta_Grind_add_go___closed__10() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_Grind_add_go___closed__9; -x_2 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Meta_Grind_add_go___closed__11() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_Grind_add_go___closed__10; -x_2 = l_Lean_MessageData_ofFormat(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Meta_Grind_add_go___closed__12() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_Grind_add_go___closed__4; -x_2 = l_Lean_Meta_Grind_add_go___closed__11; -x_3 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set(x_3, 1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Meta_Grind_add_go___closed__13() { +static lean_object* _init_l_Lean_Meta_Grind_add_go___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_Grind_add_go___closed__12; -x_2 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___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); +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind_add_go___closed__2; +x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } LEAN_EXPORT lean_object* l_Lean_Meta_Grind_add_go(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { -lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; -x_14 = l_Lean_Meta_Grind_add_go___closed__2; -x_15 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_14, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -x_16 = lean_ctor_get(x_15, 0); +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +lean_inc(x_3); +x_14 = l_Lean_Meta_instantiateMVarsIfMVarApp(x_3, x_9, x_10, x_11, x_12, x_13); +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +x_16 = lean_ctor_get(x_14, 1); lean_inc(x_16); -x_17 = lean_unbox(x_16); -lean_dec(x_16); -if (x_17 == 0) +lean_dec(x_14); +x_17 = l_Lean_Expr_cleanupAnnotations(x_15); +x_18 = l_Lean_Expr_isApp(x_17); +if (x_18 == 0) { -lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_18 = lean_ctor_get(x_15, 1); -lean_inc(x_18); -lean_dec(x_15); +lean_object* x_19; lean_object* x_20; +lean_dec(x_17); x_19 = lean_box(0); -x_20 = l_Lean_Meta_Grind_add_go___lambda__4(x_3, x_1, x_2, x_4, x_19, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_18); +x_20 = l_Lean_Meta_Grind_add_go___lambda__3(x_3, x_2, x_4, x_1, x_19, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_16); return x_20; } else { -uint8_t x_21; -x_21 = !lean_is_exclusive(x_15); -if (x_21 == 0) -{ -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_15, 1); -x_23 = lean_ctor_get(x_15, 0); -lean_dec(x_23); -lean_inc(x_3); -x_24 = l_Lean_MessageData_ofExpr(x_3); -if (x_4 == 0) +lean_object* x_21; lean_object* x_22; uint8_t x_23; +x_21 = l_Lean_Expr_appArg(x_17, lean_box(0)); +x_22 = l_Lean_Expr_appFnCleanup(x_17, lean_box(0)); +x_23 = l_Lean_Expr_isApp(x_22); +if (x_23 == 0) { -lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_25 = l_Lean_Meta_Grind_add_go___closed__8; -lean_ctor_set_tag(x_15, 7); -lean_ctor_set(x_15, 1, x_24); -lean_ctor_set(x_15, 0, x_25); -x_26 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__2; -x_27 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_27, 0, x_15); -lean_ctor_set(x_27, 1, x_26); -x_28 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_14, x_27, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_22); -x_29 = lean_ctor_get(x_28, 0); -lean_inc(x_29); -x_30 = lean_ctor_get(x_28, 1); -lean_inc(x_30); -lean_dec(x_28); -x_31 = l_Lean_Meta_Grind_add_go___lambda__4(x_3, x_1, x_2, x_4, x_29, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_30); -lean_dec(x_29); -return x_31; +lean_object* x_24; lean_object* x_25; +lean_dec(x_22); +lean_dec(x_21); +x_24 = lean_box(0); +x_25 = l_Lean_Meta_Grind_add_go___lambda__3(x_3, x_2, x_4, x_1, x_24, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_16); +return x_25; } else { -lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; -x_32 = l_Lean_Meta_Grind_add_go___closed__13; -lean_ctor_set_tag(x_15, 7); -lean_ctor_set(x_15, 1, x_24); -lean_ctor_set(x_15, 0, x_32); -x_33 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__2; -x_34 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_34, 0, x_15); -lean_ctor_set(x_34, 1, x_33); -x_35 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_14, x_34, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_22); -x_36 = lean_ctor_get(x_35, 0); -lean_inc(x_36); -x_37 = lean_ctor_get(x_35, 1); -lean_inc(x_37); -lean_dec(x_35); -x_38 = l_Lean_Meta_Grind_add_go___lambda__4(x_3, x_1, x_2, x_4, x_36, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_37); -lean_dec(x_36); -return x_38; +lean_object* x_26; lean_object* x_27; uint8_t x_28; +x_26 = l_Lean_Expr_appArg(x_22, lean_box(0)); +x_27 = l_Lean_Expr_appFnCleanup(x_22, lean_box(0)); +x_28 = l_Lean_Expr_isApp(x_27); +if (x_28 == 0) +{ +lean_object* x_29; lean_object* x_30; +lean_dec(x_27); +lean_dec(x_26); +lean_dec(x_21); +x_29 = lean_box(0); +x_30 = l_Lean_Meta_Grind_add_go___lambda__3(x_3, x_2, x_4, x_1, x_29, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_16); +return x_30; } +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; uint8_t x_34; +x_31 = l_Lean_Expr_appArg(x_27, lean_box(0)); +x_32 = l_Lean_Expr_appFnCleanup(x_27, lean_box(0)); +x_33 = l_Lean_Meta_Grind_add_go___closed__1; +x_34 = l_Lean_Expr_isConstOf(x_32, x_33); +if (x_34 == 0) +{ +uint8_t x_35; +lean_dec(x_26); +x_35 = l_Lean_Expr_isApp(x_32); +if (x_35 == 0) +{ +lean_object* x_36; lean_object* x_37; +lean_dec(x_32); +lean_dec(x_31); +lean_dec(x_21); +x_36 = lean_box(0); +x_37 = l_Lean_Meta_Grind_add_go___lambda__3(x_3, x_2, x_4, x_1, x_36, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_16); +return x_37; } else { -lean_object* x_39; lean_object* x_40; -x_39 = lean_ctor_get(x_15, 1); -lean_inc(x_39); -lean_dec(x_15); -lean_inc(x_3); -x_40 = l_Lean_MessageData_ofExpr(x_3); -if (x_4 == 0) +lean_object* x_38; lean_object* x_39; uint8_t x_40; +x_38 = l_Lean_Expr_appFnCleanup(x_32, lean_box(0)); +x_39 = l_Lean_Meta_Grind_add_go___closed__3; +x_40 = l_Lean_Expr_isConstOf(x_38, x_39); +lean_dec(x_38); +if (x_40 == 0) { -lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; -x_41 = l_Lean_Meta_Grind_add_go___closed__8; -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___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__2; -x_44 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_44, 0, x_42); -lean_ctor_set(x_44, 1, x_43); -x_45 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_14, x_44, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_39); -x_46 = lean_ctor_get(x_45, 0); -lean_inc(x_46); -x_47 = lean_ctor_get(x_45, 1); -lean_inc(x_47); -lean_dec(x_45); -x_48 = l_Lean_Meta_Grind_add_go___lambda__4(x_3, x_1, x_2, x_4, x_46, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_47); -lean_dec(x_46); -return x_48; +lean_object* x_41; lean_object* x_42; +lean_dec(x_31); +lean_dec(x_21); +x_41 = lean_box(0); +x_42 = l_Lean_Meta_Grind_add_go___lambda__3(x_3, x_2, x_4, x_1, x_41, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_16); +return x_42; } 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 = l_Lean_Meta_Grind_add_go___closed__13; -x_50 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_50, 0, x_49); -lean_ctor_set(x_50, 1, x_40); -x_51 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__2; -x_52 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_52, 0, x_50); -lean_ctor_set(x_52, 1, x_51); -x_53 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_14, x_52, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_39); -x_54 = lean_ctor_get(x_53, 0); -lean_inc(x_54); -x_55 = lean_ctor_get(x_53, 1); -lean_inc(x_55); -lean_dec(x_53); -x_56 = l_Lean_Meta_Grind_add_go___lambda__4(x_3, x_1, x_2, x_4, x_54, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_55); -lean_dec(x_54); -return x_56; +uint8_t x_43; lean_object* x_44; +x_43 = 1; +x_44 = l_Lean_Meta_Grind_add_goEq(x_1, x_2, x_3, x_31, x_21, x_4, x_43, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_16); +return x_44; +} +} +} +else +{ +uint8_t x_45; lean_object* x_46; +lean_dec(x_32); +lean_dec(x_31); +x_45 = 0; +x_46 = l_Lean_Meta_Grind_add_goEq(x_1, x_2, x_3, x_26, x_21, x_4, x_45, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_16); +return x_46; +} } } } @@ -9178,17 +9805,6 @@ lean_dec(x_5); return x_16; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_add_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, lean_object* x_12, lean_object* x_13, lean_object* x_14) { -_start: -{ -uint8_t x_15; lean_object* x_16; -x_15 = lean_unbox(x_4); -lean_dec(x_4); -x_16 = l_Lean_Meta_Grind_add_go___lambda__4(x_1, x_2, x_3, x_15, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -lean_dec(x_5); -return x_16; -} -} LEAN_EXPORT lean_object* l_Lean_Meta_Grind_add_go___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { @@ -9343,24 +9959,25 @@ static lean_object* _init_l_Lean_Meta_Grind_add___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked(" : ", 3, 3); +x_1 = lean_mk_string_unchecked("assert", 6, 6); return x_1; } } static lean_object* _init_l_Lean_Meta_Grind_add___closed__2() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_Grind_add___closed__1; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__1; +x_2 = l_Lean_Meta_Grind_add___closed__1; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; } } LEAN_EXPORT lean_object* l_Lean_Meta_Grind_add(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; -x_13 = l_Lean_Meta_Grind_add_go___closed__2; +x_13 = l_Lean_Meta_Grind_add___closed__2; x_14 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); x_15 = lean_ctor_get(x_14, 0); lean_inc(x_15); @@ -9382,71 +9999,53 @@ uint8_t x_20; x_20 = !lean_is_exclusive(x_14); if (x_20 == 0) { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; +lean_object* x_21; lean_object* x_22; lean_object* x_23; 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_14, 1); x_22 = lean_ctor_get(x_14, 0); lean_dec(x_22); -lean_inc(x_2); -x_23 = l_Lean_MessageData_ofExpr(x_2); -x_24 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__2; +lean_inc(x_1); +x_23 = l_Lean_MessageData_ofExpr(x_1); +x_24 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; lean_ctor_set_tag(x_14, 7); lean_ctor_set(x_14, 1, x_23); lean_ctor_set(x_14, 0, x_24); -x_25 = l_Lean_Meta_Grind_add___closed__2; -x_26 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_26, 0, x_14); -lean_ctor_set(x_26, 1, x_25); -lean_inc(x_1); -x_27 = l_Lean_MessageData_ofExpr(x_1); -x_28 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set(x_28, 1, x_27); -x_29 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_29, 0, x_28); -lean_ctor_set(x_29, 1, x_24); -x_30 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_13, x_29, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_21); -x_31 = lean_ctor_get(x_30, 0); -lean_inc(x_31); -x_32 = lean_ctor_get(x_30, 1); -lean_inc(x_32); -lean_dec(x_30); -x_33 = l_Lean_Meta_Grind_add___lambda__4(x_1, x_2, x_3, x_31, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_32); -lean_dec(x_31); -return x_33; +x_25 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_25, 0, x_14); +lean_ctor_set(x_25, 1, x_24); +x_26 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_13, x_25, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_21); +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_Grind_add___lambda__4(x_1, x_2, x_3, x_27, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_28); +lean_dec(x_27); +return x_29; } 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_14, 1); -lean_inc(x_34); +lean_object* x_30; lean_object* x_31; lean_object* x_32; 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_14, 1); +lean_inc(x_30); lean_dec(x_14); -lean_inc(x_2); -x_35 = l_Lean_MessageData_ofExpr(x_2); -x_36 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__2; -x_37 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_37, 0, x_36); -lean_ctor_set(x_37, 1, x_35); -x_38 = l_Lean_Meta_Grind_add___closed__2; -x_39 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_39, 0, x_37); -lean_ctor_set(x_39, 1, x_38); lean_inc(x_1); -x_40 = l_Lean_MessageData_ofExpr(x_1); -x_41 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_41, 0, x_39); -lean_ctor_set(x_41, 1, x_40); -x_42 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_42, 0, x_41); -lean_ctor_set(x_42, 1, x_36); -x_43 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_13, x_42, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_34); -x_44 = lean_ctor_get(x_43, 0); -lean_inc(x_44); -x_45 = lean_ctor_get(x_43, 1); -lean_inc(x_45); -lean_dec(x_43); -x_46 = l_Lean_Meta_Grind_add___lambda__4(x_1, x_2, x_3, x_44, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_45); -lean_dec(x_44); -return x_46; +x_31 = l_Lean_MessageData_ofExpr(x_1); +x_32 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; +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 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_34, 0, x_33); +lean_ctor_set(x_34, 1, x_32); +x_35 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_13, x_34, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, 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_Grind_add___lambda__4(x_1, x_2, x_3, x_36, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_37); +lean_dec(x_36); +return x_38; } } } @@ -9478,7 +10077,7 @@ lean_dec(x_4); return x_14; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addHyp(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_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_Grind_addHypothesis(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, 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; @@ -9566,6 +10165,28 @@ if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); l_Lean_PersistentHashMap_eraseAux___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__2___closed__1 = _init_l_Lean_PersistentHashMap_eraseAux___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__2___closed__1(); l_Lean_PersistentHashMap_eraseAux___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__2___closed__2 = _init_l_Lean_PersistentHashMap_eraseAux___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__2___closed__2(); +l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___lambda__1___closed__1 = _init_l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___lambda__1___closed__1(); +lean_mark_persistent(l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___lambda__1___closed__1); +l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__1 = _init_l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__1(); +lean_mark_persistent(l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__1); +l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__2 = _init_l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__2(); +lean_mark_persistent(l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__2); +l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__3 = _init_l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__3(); +lean_mark_persistent(l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__3); +l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__4 = _init_l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__4(); +lean_mark_persistent(l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__4); +l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__5 = _init_l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__5(); +lean_mark_persistent(l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__5); +l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__6 = _init_l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__6(); +lean_mark_persistent(l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__6); +l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__7 = _init_l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__7(); +lean_mark_persistent(l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__7); +l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8 = _init_l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8(); +lean_mark_persistent(l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8); +l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_reinsertParents___spec__1___closed__1 = _init_l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_reinsertParents___spec__1___closed__1(); +lean_mark_persistent(l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_reinsertParents___spec__1___closed__1); +l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_reinsertParents___spec__1___closed__2 = _init_l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_reinsertParents___spec__1___closed__2(); +lean_mark_persistent(l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_reinsertParents___spec__1___closed__2); l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_closeGoalWithTrueEqFalse___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_closeGoalWithTrueEqFalse___closed__1(); lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_closeGoalWithTrueEqFalse___closed__1); l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_closeGoalWithTrueEqFalse___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_closeGoalWithTrueEqFalse___closed__2(); @@ -9614,10 +10235,6 @@ l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__3); l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__4(); lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__4); -l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__5 = _init_l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__5(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__5); -l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__6 = _init_l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__6(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__6); l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__1(); lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__1); l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__2(); @@ -9628,10 +10245,6 @@ l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__4); l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__5 = _init_l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__5(); lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__5); -l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__6 = _init_l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__6(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__6); -l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__7 = _init_l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__7(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__7); l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__2___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__2___closed__1(); lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__2___closed__1); l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__2___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__2___closed__2(); @@ -9640,16 +10253,10 @@ l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__2 lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__2___closed__3); l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__3___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__3___closed__1(); lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__3___closed__1); -l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__9___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__9___closed__1(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__9___closed__1); -l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__9___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__9___closed__2(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__9___closed__2); -l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__9___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__9___closed__3(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__9___closed__3); -l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__9___closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__9___closed__4(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__9___closed__4); -l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__9___closed__5 = _init_l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__9___closed__5(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__9___closed__5); +l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__8___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__8___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__8___closed__1); +l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__8___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__8___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__8___closed__2); l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__1(); lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__1); l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__2(); @@ -9660,48 +10267,16 @@ l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__4 lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__4); l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__5 = _init_l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__5(); lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__5); -l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__6 = _init_l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__6(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__6); -l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__7 = _init_l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__7(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__7); -l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__8 = _init_l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__8(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__8); l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_resetNewEqs___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_resetNewEqs___closed__1(); lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_resetNewEqs___closed__1); l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqCore_processTodo___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqCore_processTodo___closed__1(); lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqCore_processTodo___closed__1); -l_Lean_Meta_Grind_add_go___lambda__4___closed__1 = _init_l_Lean_Meta_Grind_add_go___lambda__4___closed__1(); -lean_mark_persistent(l_Lean_Meta_Grind_add_go___lambda__4___closed__1); -l_Lean_Meta_Grind_add_go___lambda__4___closed__2 = _init_l_Lean_Meta_Grind_add_go___lambda__4___closed__2(); -lean_mark_persistent(l_Lean_Meta_Grind_add_go___lambda__4___closed__2); -l_Lean_Meta_Grind_add_go___lambda__4___closed__3 = _init_l_Lean_Meta_Grind_add_go___lambda__4___closed__3(); -lean_mark_persistent(l_Lean_Meta_Grind_add_go___lambda__4___closed__3); l_Lean_Meta_Grind_add_go___closed__1 = _init_l_Lean_Meta_Grind_add_go___closed__1(); lean_mark_persistent(l_Lean_Meta_Grind_add_go___closed__1); l_Lean_Meta_Grind_add_go___closed__2 = _init_l_Lean_Meta_Grind_add_go___closed__2(); lean_mark_persistent(l_Lean_Meta_Grind_add_go___closed__2); l_Lean_Meta_Grind_add_go___closed__3 = _init_l_Lean_Meta_Grind_add_go___closed__3(); lean_mark_persistent(l_Lean_Meta_Grind_add_go___closed__3); -l_Lean_Meta_Grind_add_go___closed__4 = _init_l_Lean_Meta_Grind_add_go___closed__4(); -lean_mark_persistent(l_Lean_Meta_Grind_add_go___closed__4); -l_Lean_Meta_Grind_add_go___closed__5 = _init_l_Lean_Meta_Grind_add_go___closed__5(); -lean_mark_persistent(l_Lean_Meta_Grind_add_go___closed__5); -l_Lean_Meta_Grind_add_go___closed__6 = _init_l_Lean_Meta_Grind_add_go___closed__6(); -lean_mark_persistent(l_Lean_Meta_Grind_add_go___closed__6); -l_Lean_Meta_Grind_add_go___closed__7 = _init_l_Lean_Meta_Grind_add_go___closed__7(); -lean_mark_persistent(l_Lean_Meta_Grind_add_go___closed__7); -l_Lean_Meta_Grind_add_go___closed__8 = _init_l_Lean_Meta_Grind_add_go___closed__8(); -lean_mark_persistent(l_Lean_Meta_Grind_add_go___closed__8); -l_Lean_Meta_Grind_add_go___closed__9 = _init_l_Lean_Meta_Grind_add_go___closed__9(); -lean_mark_persistent(l_Lean_Meta_Grind_add_go___closed__9); -l_Lean_Meta_Grind_add_go___closed__10 = _init_l_Lean_Meta_Grind_add_go___closed__10(); -lean_mark_persistent(l_Lean_Meta_Grind_add_go___closed__10); -l_Lean_Meta_Grind_add_go___closed__11 = _init_l_Lean_Meta_Grind_add_go___closed__11(); -lean_mark_persistent(l_Lean_Meta_Grind_add_go___closed__11); -l_Lean_Meta_Grind_add_go___closed__12 = _init_l_Lean_Meta_Grind_add_go___closed__12(); -lean_mark_persistent(l_Lean_Meta_Grind_add_go___closed__12); -l_Lean_Meta_Grind_add_go___closed__13 = _init_l_Lean_Meta_Grind_add_go___closed__13(); -lean_mark_persistent(l_Lean_Meta_Grind_add_go___closed__13); l_Lean_Meta_Grind_add___lambda__3___closed__1 = _init_l_Lean_Meta_Grind_add___lambda__3___closed__1(); lean_mark_persistent(l_Lean_Meta_Grind_add___lambda__3___closed__1); l_Lean_Meta_Grind_add___lambda__3___closed__2 = _init_l_Lean_Meta_Grind_add___lambda__3___closed__2(); diff --git a/stage0/stdlib/Lean/Meta/Tactic/Grind/EMatch.c b/stage0/stdlib/Lean/Meta/Tactic/Grind/EMatch.c new file mode 100644 index 000000000000..b573249175a5 --- /dev/null +++ b/stage0/stdlib/Lean/Meta/Tactic/Grind/EMatch.c @@ -0,0 +1,14724 @@ +// Lean compiler output +// Module: Lean.Meta.Tactic.Grind.EMatch +// Imports: Lean.Meta.Tactic.Grind.Types Lean.Meta.Tactic.Grind.Intro +#include +#if defined(__clang__) +#pragma clang diagnostic ignored "-Wunused-parameter" +#pragma clang diagnostic ignored "-Wunused-label" +#elif defined(__GNUC__) && !defined(__CLANG__) +#pragma GCC diagnostic ignored "-Wunused-parameter" +#pragma GCC diagnostic ignored "-Wunused-label" +#pragma GCC diagnostic ignored "-Wunused-but-set-variable" +#endif +#ifdef __cplusplus +extern "C" { +#endif +lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___boxed(lean_object*, lean_object*, lean_object*, lean_object*, 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_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___spec__1___closed__2; +LEAN_EXPORT lean_object* l_Lean_mkConstWithLevelParams___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___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___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem_synthesizeInstance(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_EMatch_ematchTheorem(lean_object*, lean_object*, lean_object*, 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_Grind_EMatch_ematchTheorem___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_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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*); +extern lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare; +static lean_object* l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2; +lean_object* lean_mk_empty_array_with_capacity(lean_object*); +lean_object* l_Lean_mkAppN(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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_EXPORT lean_object* l_Lean_Meta_Grind_ematchStar(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_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; +lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(lean_object*, lean_object*); +lean_object* l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___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___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___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_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___boxed(lean_object**); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_EMatch_instInhabitedContext; +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___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_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_ConstantInfo_levelParams(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___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___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__2; +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_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_main___spec__1___boxed(lean_object**); +static lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___spec__1___closed__6; +uint64_t lean_uint64_lor(uint64_t, uint64_t); +static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__3; +uint8_t l_Lean_Expr_isApp(lean_object*); +lean_object* l_Lean_Meta_trySynthInstance(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_MessageData_ofList(lean_object*); +lean_object* l_Lean_Meta_Grind_iterate(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_push___rarg(lean_object*, lean_object*); +lean_object* lean_array_push(lean_object*, lean_object*); +lean_object* l_Array_toSubarray___rarg(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Core_checkSystem(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_ematchStar___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_EMatch_ematchTheorem_tryAll(lean_object*, lean_object*, lean_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_Grind_ematch___closed__1; +uint8_t lean_usize_dec_eq(size_t, size_t); +static lean_object* l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5___closed__4; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_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_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___closed__4; +lean_object* l_instInhabitedReaderT___rarg___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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*); +static lean_object* l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5___closed__1; +lean_object* l_Lean_Expr_bvar___override(lean_object*); +lean_object* lean_array_fget(lean_object*, lean_object*); +lean_object* l_Lean_Meta_Grind_getMaxGeneration___rarg(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___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__8___closed__1; +extern lean_object* l_Lean_Meta_Grind_grind_debug_proofs; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__4; +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_Grind_EMatch_ematchTheorems___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* lean_environment_find(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_EMatch_ematchTheorems___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___boxed(lean_object**); +LEAN_EXPORT lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_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___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__3; +LEAN_EXPORT lean_object* l_ReaderT_read___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___spec__1___closed__3; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematchAndAssert_x3f___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_EMatch_ematchTheorem___lambda__1___closed__2; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__3; +lean_object* l_Lean_stringToMessageData(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_EMatch_M_run_x27___rarg___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_Grind_assertNext_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_List_head_x21___rarg(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___closed__1; +static lean_object* l_Lean_Meta_Grind_ematchAndAssert_x3f___closed__1; +static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___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___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__2; +LEAN_EXPORT lean_object* l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Expr_appArg_x21(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematchAndAssert_x3f___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_Expr_isBVar(lean_object*); +lean_object* l_List_mapTR_loop___at_Lean_mkConstWithLevelParams___spec__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__8___closed__2; +static lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___spec__1___closed__5; +LEAN_EXPORT lean_object* l_Lean_mkConstWithLevelParams___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___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*); +extern lean_object* l_instInhabitedPUnit; +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Expr_hasMVar(lean_object*); +lean_object* l_Lean_Meta_Grind_EMatchTheorem_getProofWithFreshMVarLevels(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentArray_forMAux___at_Lean_Meta_Grind_EMatch_ematchTheorems___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___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__5; +lean_object* l_List_appendTR___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_ReaderT_read___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem(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_Lean_Meta_Grind_ematch___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___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__2; +lean_object* lean_checked_assign(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Array_reverse___rarg(lean_object*); +lean_object* lean_st_ref_take(lean_object*, lean_object*); +static lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___spec__1___closed__4; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_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*); +static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__7; +static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__3; +uint8_t lean_expr_eqv(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_withNewMCtxDepth___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__12___rarg(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_EMatch_instInhabitedState; +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__10(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint64_t lean_uint64_shift_right(uint64_t, uint64_t); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices(lean_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_ofSyntax(lean_object*); +static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_MVarId_getType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_withNewMCtxDepth___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__12___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_x3f___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1(size_t, size_t, lean_object*); +static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__8; +lean_object* l_Lean_Meta_Grind_checkMaxInstancesExceeded(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3___closed__3; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematch___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_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__2; +lean_object* l_Lean_Meta_Grind_groundPattern_x3f(lean_object*); +static lean_object* l_Lean_Meta_Grind_EMatch_instInhabitedCnstr___closed__1; +static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__9; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__3; +lean_object* l_Lean_MessageData_ofFormat(lean_object*); +lean_object* l_Lean_PersistentArray_append___rarg(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_EMatch_instInhabitedCnstr___closed__2; +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_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_main(lean_object*, lean_object*, lean_object*, lean_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_Grind_EMatch_instInhabitedChoice___closed__2; +lean_object* l_outOfBounds___rarg(lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_ematchAndAssert_x3f___closed__2; +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_main___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_ematch___closed__2; +lean_object* lean_st_ref_get(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_EMatch_instInhabitedContext___closed__1; +LEAN_EXPORT lean_object* l_Lean_PersistentArray_forMAux___at_Lean_Meta_Grind_EMatch_ematchTheorems___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*); +uint8_t l_List_isEmpty___rarg(lean_object*); +lean_object* lean_st_mk_ref(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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*, lean_object*); +lean_object* lean_array_to_list(lean_object*); +static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___lambda__1___closed__3; +static lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___spec__1___closed__1; +LEAN_EXPORT lean_object* l_Lean_PersistentArray_forM___at_Lean_Meta_Grind_EMatch_ematchTheorems___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_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___closed__1; +lean_object* l_Lean_Expr_toHeadIndex(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData(lean_object*); +static lean_object* l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5___closed__2; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___closed__2; +lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withNewMCtxDepthImp___rarg(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Expr_constName_x21(lean_object*); +extern lean_object* l_Lean_instInhabitedExpr; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__4___boxed(lean_object**); +lean_object* l_Lean_MessageData_ofConst(lean_object*); +LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_EMatch_instInhabitedCnstr; +uint8_t l_Lean_BinderInfo_isInstImplicit(uint8_t); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__5; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___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_Meta_Grind_EMatch_ematchTheorem___lambda__1___closed__1; +LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___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_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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_Lean_Meta_Grind_EMatch_M_run_x27___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_Grind_ematchAndAssert_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_Grind_ematchAndAssert_x3f___lambda__4___closed__1; +static lean_object* l_Lean_Meta_Grind_EMatch_instInhabitedChoice___closed__1; +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_main___spec__1___lambda__1___boxed(lean_object**); +lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3___closed__2; +uint8_t l_Lean_Meta_Grind_isSameExpr_unsafe__1(lean_object*, lean_object*); +static uint64_t l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___closed__2; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_eqvFunctions___boxed(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_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_main___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematch___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_Meta_Grind_EMatch_instInhabitedContext___closed__2; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___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_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___closed__3; +lean_object* l_Lean_Expr_appFn_x21(lean_object*); +static lean_object* l_Lean_Meta_Grind_ematch___closed__3; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__4; +extern lean_object* l_Lean_Meta_instMonadMetaM; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_EMatch_ematchTheorem___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_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__9___boxed(lean_object**); +static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___spec__1___closed__1; +uint8_t l_Lean_PersistentHashMap_contains___at_Lean_MVarId_isAssigned___spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematchAndAssert_x3f___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_MessageData_ofConstName(lean_object*, uint8_t); +lean_object* l_OptionT_instMonad___rarg(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematchAndAssert_x3f___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___boxed(lean_object**); +lean_object* l_Lean_MessageData_ofExpr(lean_object*); +static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__4; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance(lean_object*, lean_object*, lean_object*, lean_object*, lean_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_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_List_tail_x21___rarg(lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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*); +static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__6; +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_Grind_EMatch_ematchTheorems___spec__4(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +double l_Float_ofScientific(lean_object*, uint8_t, lean_object*); +LEAN_EXPORT lean_object* l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___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_Grind_EMatch_ematchTheorems(lean_object*, lean_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_forallMetaBoundedTelescope(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_Grind_EMatch_ematchTheorems___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_Grind_getNext(lean_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___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5___closed__3; +lean_object* l_Lean_PersistentHashMap_find_x3f___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___boxed(lean_object*, lean_object*, lean_object*, 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_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__4; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_EMatch_M_run_x27(lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__5; +lean_object* l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___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_ReaderT_bind___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__11___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematch___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___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___boxed(lean_object**); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematchAndAssert_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_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___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___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_x3f___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t lean_nat_dec_eq(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_eqvFunctions(lean_object*, lean_object*); +lean_object* l_Lean_Meta_Grind_addTheoremInstance(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___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*); +uint8_t lean_nat_dec_lt(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_unassigned___closed__1; +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__6___boxed(lean_object**); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__7; +static lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__2___closed__2; +lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); +lean_object* l_Lean_indentExpr(lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__8; +static double l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__1; +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__7___boxed(lean_object**); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__1; +lean_object* l_Lean_Meta_Grind_markTheoremInstance(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Expr_isConstOf(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___lambda__2___closed__1; +lean_object* l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_panic_fn(lean_object*, lean_object*); +lean_object* l_Lean_Meta_Grind_isEqv(lean_object*, 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_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__2___closed__3; +static lean_object* l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3___closed__1; +lean_object* lean_nat_sub(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Expr_getAppFn(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___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___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__1; +LEAN_EXPORT lean_object* l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___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*); +uint64_t lean_uint64_shift_left(uint64_t, uint64_t); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_Grind_getENode(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_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___closed__2; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematchAndAssert_x3f___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_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_List_reverse___rarg(lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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*); +static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__5; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_array_mk(lean_object*); +lean_object* l_Lean_instantiateMVarsCore(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_withNewMCtxDepth___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__12(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematch(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_Grind_EMatch_ematchTheorems___spec__3(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__4(lean_object*, lean_object*); +lean_object* l_Lean_isTracingEnabledForCore(lean_object*, lean_object*, lean_object*); +size_t lean_usize_add(size_t, size_t); +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___spec__1___boxed(lean_object**); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch(lean_object*, lean_object*, lean_object*, lean_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_object* l_Lean_Expr_fvar___override(lean_object*); +size_t lean_array_size(lean_object*); +lean_object* l_instInhabitedOfMonad___rarg(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__5___closed__1; +lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_ReaderT_bind___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__11(lean_object*, lean_object*); +lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_unassigned___closed__3; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__6; +static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___lambda__1___closed__1; +static size_t l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__1; +lean_object* lean_string_append(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__5___boxed(lean_object**); +lean_object* lean_array_get_size(lean_object*); +static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___lambda__1___closed__2; +LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_EMatch_instInhabitedContext___closed__3; +lean_object* lean_infer_type(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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*); +uint8_t lean_nat_dec_le(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_EMatch_instInhabitedChoice; +uint8_t lean_usize_dec_lt(size_t, size_t); +LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Meta_Grind_EMatch_ematchTheorem_tryAll___spec__1(lean_object*, lean_object*); +lean_object* l_Lean_Expr_bvarIdx_x21(lean_object*); +lean_object* l_Lean_Meta_mkLambdaFVars(lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint64_t l_Lean_Meta_TransparencyMode_toUInt64(uint8_t); +lean_object* lean_nat_add(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_MVarId_isAssigned___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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*); +uint8_t l_Lean_Expr_isConst(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_unassigned; +uint8_t l_Lean_Expr_isFVar(lean_object*); +static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__1; +lean_object* l_Lean_Expr_mvarId_x21(lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentArray_forM___at_Lean_Meta_Grind_EMatch_ematchTheorems___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_unassigned___closed__2; +lean_object* lean_array_uset(lean_object*, size_t, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_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_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__2___closed__1; +LEAN_EXPORT lean_object* l_Lean_MVarId_isAssigned___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__6; +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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_object*, lean_object*); +lean_object* l_Lean_Meta_isProof(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___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___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_ReaderT_instMonad___rarg(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___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*); +static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__4; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematchAndAssert_x3f___lambda__3(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_Grind_EMatch_instInhabitedCnstr___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = l_Lean_Expr_bvar___override(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Grind_EMatch_instInhabitedCnstr___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_EMatch_instInhabitedCnstr___closed__1; +x_2 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_2, 0, x_1); +lean_ctor_set(x_2, 1, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Grind_EMatch_instInhabitedCnstr() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Meta_Grind_EMatch_instInhabitedCnstr___closed__2; +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_unassigned___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("[grind_unassigned]", 18, 18); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_unassigned___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_unassigned___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_unassigned___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_unassigned___closed__2; +x_3 = l_Lean_Expr_const___override(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_unassigned() { +_start: +{ +lean_object* x_1; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_unassigned___closed__3; +return x_1; +} +} +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("", 0, 0); +return x_1; +} +} +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("_", 1, 1); +return x_1; +} +} +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__3; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1(size_t x_1, size_t x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; +x_4 = lean_usize_dec_lt(x_2, x_1); +if (x_4 == 0) +{ +return x_3; +} +else +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; size_t x_10; size_t x_11; +x_5 = lean_array_uget(x_3, x_2); +x_6 = lean_unsigned_to_nat(0u); +x_7 = lean_array_uset(x_3, x_2, x_6); +x_8 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_unassigned; +x_9 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_5, x_8); +x_10 = 1; +x_11 = lean_usize_add(x_2, x_10); +if (x_9 == 0) +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_12 = l_Lean_MessageData_ofExpr(x_5); +x_13 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; +x_14 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_12); +x_15 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_15, 0, x_14); +lean_ctor_set(x_15, 1, x_13); +x_16 = lean_array_uset(x_7, x_2, x_15); +x_2 = x_11; +x_3 = x_16; +goto _start; +} +else +{ +lean_object* x_18; lean_object* x_19; +lean_dec(x_5); +x_18 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__4; +x_19 = lean_array_uset(x_7, x_2, x_18); +x_2 = x_11; +x_3 = x_19; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData(lean_object* x_1) { +_start: +{ +lean_object* x_2; size_t x_3; size_t x_4; lean_object* x_5; +x_2 = l_Array_reverse___rarg(x_1); +x_3 = lean_array_size(x_2); +x_4 = 0; +x_5 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1(x_3, x_4, x_2); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +size_t x_4; size_t x_5; lean_object* x_6; +x_4 = lean_unbox_usize(x_1); +lean_dec(x_1); +x_5 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1(x_4, x_5, x_3); +return x_6; +} +} +static lean_object* _init_l_Lean_Meta_Grind_EMatch_instInhabitedChoice___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = lean_mk_empty_array_with_capacity(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Grind_EMatch_instInhabitedChoice___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = lean_box(0); +x_2 = lean_unsigned_to_nat(0u); +x_3 = l_Lean_Meta_Grind_EMatch_instInhabitedChoice___closed__1; +x_4 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Meta_Grind_EMatch_instInhabitedChoice() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Meta_Grind_EMatch_instInhabitedChoice___closed__2; +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_EMatch_instInhabitedContext___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(0); +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Grind_EMatch_instInhabitedContext___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 = lean_box(0); +x_2 = l_Lean_Meta_Grind_EMatch_instInhabitedChoice___closed__1; +x_3 = l_Lean_Meta_Grind_EMatch_instInhabitedCnstr___closed__1; +x_4 = lean_unsigned_to_nat(0u); +x_5 = l_Lean_Meta_Grind_EMatch_instInhabitedContext___closed__1; +x_6 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_6, 0, x_2); +lean_ctor_set(x_6, 1, x_3); +lean_ctor_set(x_6, 2, x_4); +lean_ctor_set(x_6, 3, x_1); +lean_ctor_set(x_6, 4, x_1); +lean_ctor_set(x_6, 5, x_5); +return x_6; +} +} +static lean_object* _init_l_Lean_Meta_Grind_EMatch_instInhabitedContext___closed__3() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; +x_1 = 0; +x_2 = l_Lean_Meta_Grind_EMatch_instInhabitedContext___closed__2; +x_3 = lean_alloc_ctor(0, 1, 1); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind_EMatch_instInhabitedContext() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Meta_Grind_EMatch_instInhabitedContext___closed__3; +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_EMatch_instInhabitedState() { +_start: +{ +lean_object* x_1; +x_1 = lean_box(0); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_EMatch_M_run_x27___rarg___closed__1() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; +x_1 = 1; +x_2 = l_Lean_Meta_Grind_EMatch_instInhabitedContext___closed__2; +x_3 = lean_alloc_ctor(0, 1, 1); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_EMatch_M_run_x27___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; +x_11 = lean_box(0); +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_inc(x_14); +lean_dec(x_12); +x_15 = l_Lean_Meta_Grind_EMatch_M_run_x27___rarg___closed__1; +lean_inc(x_13); +x_16 = lean_apply_11(x_1, x_15, x_13, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_14); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +x_18 = lean_ctor_get(x_16, 1); +lean_inc(x_18); +lean_dec(x_16); +x_19 = lean_st_ref_get(x_13, x_18); +lean_dec(x_13); +x_20 = !lean_is_exclusive(x_19); +if (x_20 == 0) +{ +lean_object* x_21; +x_21 = lean_ctor_get(x_19, 0); +lean_dec(x_21); +lean_ctor_set(x_19, 0, x_17); +return x_19; +} +else +{ +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_24; +lean_dec(x_13); +x_24 = !lean_is_exclusive(x_16); +if (x_24 == 0) +{ +return x_16; +} +else +{ +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_dec(x_16); +x_27 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +return x_27; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_EMatch_M_run_x27(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_EMatch_M_run_x27___rarg), 10, 0); +return x_2; +} +} +static lean_object* _init_l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___spec__1___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_instMonadMetaM; +x_2 = l_ReaderT_instMonad___rarg(x_1); +return x_2; +} +} +static lean_object* _init_l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___spec__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___spec__1___closed__1; +x_2 = l_ReaderT_instMonad___rarg(x_1); +return x_2; +} +} +static lean_object* _init_l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___spec__1___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___spec__1___closed__2; +x_2 = l_ReaderT_instMonad___rarg(x_1); +return x_2; +} +} +static lean_object* _init_l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___spec__1___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___spec__1___closed__3; +x_2 = l_ReaderT_instMonad___rarg(x_1); +return x_2; +} +} +static lean_object* _init_l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___spec__1___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___spec__1___closed__4; +x_2 = l_OptionT_instMonad___rarg(x_1); +return x_2; +} +} +static lean_object* _init_l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___spec__1___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___spec__1___closed__5; +x_2 = l_Lean_Meta_Grind_EMatch_instInhabitedChoice; +x_3 = l_instInhabitedOfMonad___rarg(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_11 = l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___spec__1___closed__6; +x_12 = lean_panic_fn(x_11, x_1); +x_13 = lean_apply_9(x_12, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_13; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.Meta.Tactic.Grind.EMatch", 29, 29); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("_private.Lean.Meta.Tactic.Grind.EMatch.0.Lean.Meta.Grind.EMatch.assign\?", 71, 71); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("unreachable code has been reached", 33, 33); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___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___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___closed__2; +x_3 = lean_unsigned_to_nat(80u); +x_4 = lean_unsigned_to_nat(4u); +x_5 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___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___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_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, 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; +x_13 = lean_ctor_get(x_1, 0); +lean_inc(x_13); +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_array_get_size(x_15); +x_17 = lean_nat_dec_lt(x_2, x_16); +lean_dec(x_16); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_3); +lean_dec(x_1); +x_18 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___closed__4; +x_19 = l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___spec__1(x_18, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_19; +} +else +{ +lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_20 = lean_array_fget(x_15, x_2); +x_21 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_unassigned; +x_22 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_20, x_21); +if (x_22 == 0) +{ +lean_object* x_23; +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +x_23 = l_Lean_Meta_Grind_isEqv(x_20, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +if (lean_obj_tag(x_23) == 0) +{ +lean_object* x_24; uint8_t x_25; +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +x_25 = lean_unbox(x_24); +lean_dec(x_24); +if (x_25 == 0) +{ +uint8_t x_26; +lean_dec(x_1); +x_26 = !lean_is_exclusive(x_23); +if (x_26 == 0) +{ +lean_object* x_27; lean_object* x_28; +x_27 = lean_ctor_get(x_23, 0); +lean_dec(x_27); +x_28 = lean_box(0); +lean_ctor_set(x_23, 0, x_28); +return x_23; +} +else +{ +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_box(0); +x_31 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_31, 0, x_30); +lean_ctor_set(x_31, 1, x_29); +return x_31; +} +} +else +{ +uint8_t x_32; +x_32 = !lean_is_exclusive(x_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_alloc_ctor(1, 1, 0); +lean_ctor_set(x_34, 0, x_1); +lean_ctor_set(x_23, 0, x_34); +return x_23; +} +else +{ +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_36 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_36, 0, x_1); +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; +lean_dec(x_1); +x_38 = !lean_is_exclusive(x_23); +if (x_38 == 0) +{ +return x_23; +} +else +{ +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; +} +} +} +else +{ +uint8_t x_42; +lean_dec(x_20); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_42 = !lean_is_exclusive(x_1); +if (x_42 == 0) +{ +lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_43 = lean_ctor_get(x_1, 2); +lean_dec(x_43); +x_44 = lean_ctor_get(x_1, 1); +lean_dec(x_44); +x_45 = lean_ctor_get(x_1, 0); +lean_dec(x_45); +x_46 = lean_array_fset(x_15, x_2, x_3); +lean_ctor_set(x_1, 2, x_46); +x_47 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_47, 0, x_1); +x_48 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_48, 0, x_47); +lean_ctor_set(x_48, 1, x_12); +return x_48; +} +else +{ +lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; +lean_dec(x_1); +x_49 = lean_array_fset(x_15, x_2, x_3); +x_50 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_50, 0, x_13); +lean_ctor_set(x_50, 1, x_14); +lean_ctor_set(x_50, 2, x_49); +x_51 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_51, 0, x_50); +x_52 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_52, 0, x_51); +lean_ctor_set(x_52, 1, x_12); +return x_52; +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_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, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +x_13 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f(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_2); +return x_13; +} +} +LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_eqvFunctions(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = l_Lean_Expr_isFVar(x_1); +if (x_3 == 0) +{ +uint8_t x_4; +x_4 = l_Lean_Expr_isConst(x_1); +if (x_4 == 0) +{ +uint8_t x_5; +x_5 = 0; +return x_5; +} +else +{ +lean_object* x_6; uint8_t x_7; +x_6 = l_Lean_Expr_constName_x21(x_1); +x_7 = l_Lean_Expr_isConstOf(x_2, x_6); +lean_dec(x_6); +return x_7; +} +} +else +{ +uint8_t x_8; +x_8 = lean_expr_eqv(x_1, x_2); +if (x_8 == 0) +{ +uint8_t x_9; +x_9 = l_Lean_Expr_isConst(x_1); +if (x_9 == 0) +{ +uint8_t x_10; +x_10 = 0; +return x_10; +} +else +{ +lean_object* x_11; uint8_t x_12; +x_11 = l_Lean_Expr_constName_x21(x_1); +x_12 = l_Lean_Expr_isConstOf(x_2, x_11); +lean_dec(x_11); +return x_12; +} +} +else +{ +uint8_t x_13; +x_13 = 1; +return x_13; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_eqvFunctions___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_eqvFunctions(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_x3f___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_14 = l_Lean_Expr_appArg_x21(x_1); +x_15 = l_Lean_Expr_appArg_x21(x_2); +x_16 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare; +x_17 = lean_expr_eqv(x_14, x_16); +if (x_17 == 0) +{ +uint8_t x_18; +x_18 = l_Lean_Expr_isBVar(x_14); +if (x_18 == 0) +{ +lean_object* x_19; +x_19 = l_Lean_Meta_Grind_groundPattern_x3f(x_14); +if (lean_obj_tag(x_19) == 0) +{ +lean_object* x_20; uint8_t x_21; +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_14); +lean_ctor_set(x_20, 1, x_15); +x_21 = !lean_is_exclusive(x_3); +if (x_21 == 0) +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_22 = lean_ctor_get(x_3, 0); +x_23 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_23, 0, x_20); +lean_ctor_set(x_23, 1, x_22); +lean_ctor_set(x_3, 0, x_23); +x_24 = l_Lean_Expr_appFn_x21(x_1); +x_25 = l_Lean_Expr_appFn_x21(x_2); +x_26 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_x3f(x_3, x_24, x_25, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_25); +lean_dec(x_24); +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; +x_27 = lean_ctor_get(x_3, 0); +x_28 = lean_ctor_get(x_3, 1); +x_29 = lean_ctor_get(x_3, 2); +lean_inc(x_29); +lean_inc(x_28); +lean_inc(x_27); +lean_dec(x_3); +x_30 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_30, 0, x_20); +lean_ctor_set(x_30, 1, x_27); +x_31 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_31, 0, x_30); +lean_ctor_set(x_31, 1, x_28); +lean_ctor_set(x_31, 2, x_29); +x_32 = l_Lean_Expr_appFn_x21(x_1); +x_33 = l_Lean_Expr_appFn_x21(x_2); +x_34 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_x3f(x_31, x_32, x_33, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_33); +lean_dec(x_32); +return x_34; +} +} +else +{ +lean_object* x_35; lean_object* x_36; +lean_dec(x_14); +x_35 = lean_ctor_get(x_19, 0); +lean_inc(x_35); +lean_dec(x_19); +x_36 = l_Lean_Meta_Grind_isEqv(x_35, x_15, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_36) == 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) +{ +uint8_t x_39; +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +x_39 = !lean_is_exclusive(x_36); +if (x_39 == 0) +{ +lean_object* x_40; lean_object* x_41; +x_40 = lean_ctor_get(x_36, 0); +lean_dec(x_40); +x_41 = lean_box(0); +lean_ctor_set(x_36, 0, x_41); +return x_36; +} +else +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_42 = lean_ctor_get(x_36, 1); +lean_inc(x_42); +lean_dec(x_36); +x_43 = lean_box(0); +x_44 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_44, 0, x_43); +lean_ctor_set(x_44, 1, x_42); +return x_44; +} +} +else +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_45 = lean_ctor_get(x_36, 1); +lean_inc(x_45); +lean_dec(x_36); +x_46 = l_Lean_Expr_appFn_x21(x_1); +x_47 = l_Lean_Expr_appFn_x21(x_2); +x_48 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_x3f(x_3, x_46, x_47, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_45); +lean_dec(x_47); +lean_dec(x_46); +return x_48; +} +} +else +{ +uint8_t x_49; +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +x_49 = !lean_is_exclusive(x_36); +if (x_49 == 0) +{ +return x_36; +} +else +{ +lean_object* x_50; lean_object* x_51; lean_object* x_52; +x_50 = lean_ctor_get(x_36, 0); +x_51 = lean_ctor_get(x_36, 1); +lean_inc(x_51); +lean_inc(x_50); +lean_dec(x_36); +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_object* x_54; +x_53 = l_Lean_Expr_bvarIdx_x21(x_14); +lean_dec(x_14); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +x_54 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f(x_3, x_53, x_15, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_53); +if (lean_obj_tag(x_54) == 0) +{ +lean_object* x_55; +x_55 = lean_ctor_get(x_54, 0); +lean_inc(x_55); +if (lean_obj_tag(x_55) == 0) +{ +uint8_t x_56; +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +x_56 = !lean_is_exclusive(x_54); +if (x_56 == 0) +{ +lean_object* x_57; lean_object* x_58; +x_57 = lean_ctor_get(x_54, 0); +lean_dec(x_57); +x_58 = lean_box(0); +lean_ctor_set(x_54, 0, x_58); +return x_54; +} +else +{ +lean_object* x_59; lean_object* x_60; lean_object* x_61; +x_59 = lean_ctor_get(x_54, 1); +lean_inc(x_59); +lean_dec(x_54); +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_59); +return x_61; +} +} +else +{ +lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_62 = lean_ctor_get(x_54, 1); +lean_inc(x_62); +lean_dec(x_54); +x_63 = lean_ctor_get(x_55, 0); +lean_inc(x_63); +lean_dec(x_55); +x_64 = l_Lean_Expr_appFn_x21(x_1); +x_65 = l_Lean_Expr_appFn_x21(x_2); +x_66 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_x3f(x_63, x_64, x_65, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_62); +lean_dec(x_65); +lean_dec(x_64); +return x_66; +} +} +else +{ +uint8_t x_67; +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +x_67 = !lean_is_exclusive(x_54); +if (x_67 == 0) +{ +return x_54; +} +else +{ +lean_object* x_68; lean_object* x_69; lean_object* x_70; +x_68 = lean_ctor_get(x_54, 0); +x_69 = lean_ctor_get(x_54, 1); +lean_inc(x_69); +lean_inc(x_68); +lean_dec(x_54); +x_70 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_70, 0, x_68); +lean_ctor_set(x_70, 1, x_69); +return x_70; +} +} +} +} +else +{ +lean_object* x_71; lean_object* x_72; lean_object* x_73; +lean_dec(x_15); +lean_dec(x_14); +x_71 = l_Lean_Expr_appFn_x21(x_1); +x_72 = l_Lean_Expr_appFn_x21(x_2); +x_73 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_x3f(x_3, x_71, x_72, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_72); +lean_dec(x_71); +return x_73; +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_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, lean_object* x_11, lean_object* x_12) { +_start: +{ +uint8_t x_13; +x_13 = l_Lean_Expr_isApp(x_2); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_14 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_14, 0, x_1); +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; +} +else +{ +lean_object* x_16; lean_object* x_17; +x_16 = lean_box(0); +x_17 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_x3f___lambda__1(x_2, x_3, x_1, x_16, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_17; +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_x3f___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +lean_object* x_14; +x_14 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_x3f___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +return x_14; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_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, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +x_13 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_x3f(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_3); +lean_dec(x_2); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +lean_object* x_15; +x_15 = l_Lean_Meta_Grind_getNext(x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_15) == 0) +{ +uint8_t x_16; +x_16 = !lean_is_exclusive(x_15); +if (x_16 == 0) +{ +lean_object* x_17; uint8_t x_18; +x_17 = lean_ctor_get(x_15, 0); +x_18 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_17, x_1); +if (x_18 == 0) +{ +lean_object* x_19; +x_19 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_19, 0, x_17); +lean_ctor_set(x_15, 0, x_19); +return x_15; +} +else +{ +lean_object* x_20; +x_20 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_20, 0, x_17); +lean_ctor_set(x_15, 0, x_20); +return x_15; +} +} +else +{ +lean_object* x_21; lean_object* x_22; uint8_t 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 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_21, x_1); +if (x_23 == 0) +{ +lean_object* x_24; lean_object* x_25; +x_24 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_24, 0, x_21); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_22); +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_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); +return x_27; +} +} +} +else +{ +uint8_t x_28; +x_28 = !lean_is_exclusive(x_15); +if (x_28 == 0) +{ +return x_15; +} +else +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_29 = lean_ctor_get(x_15, 0); +x_30 = lean_ctor_get(x_15, 1); +lean_inc(x_30); +lean_inc(x_29); +lean_dec(x_15); +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_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18) { +_start: +{ +lean_object* x_19; lean_object* x_35; +lean_inc(x_7); +x_35 = l_Lean_Meta_Grind_getENode(x_7, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +if (lean_obj_tag(x_35) == 0) +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; uint8_t x_109; +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, 8); +lean_inc(x_38); +x_109 = lean_nat_dec_le(x_38, x_4); +if (x_109 == 0) +{ +lean_object* x_110; lean_object* x_111; +lean_dec(x_38); +lean_dec(x_36); +x_110 = lean_box(0); +x_111 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_3, x_7, x_110, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_37); +x_19 = x_111; +goto block_34; +} +else +{ +uint8_t x_112; +x_112 = lean_ctor_get_uint8(x_36, sizeof(void*)*10 + 4); +if (x_112 == 0) +{ +lean_object* x_113; uint8_t x_114; +x_113 = lean_ctor_get(x_36, 3); +lean_inc(x_113); +lean_dec(x_36); +x_114 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_7, x_113); +lean_dec(x_113); +if (x_114 == 0) +{ +lean_object* x_115; lean_object* x_116; +lean_dec(x_38); +x_115 = lean_box(0); +x_116 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_3, x_7, x_115, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_37); +x_19 = x_116; +goto block_34; +} +else +{ +lean_object* x_117; uint8_t x_118; +x_117 = l_Lean_Expr_getAppFn(x_7); +x_118 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_eqvFunctions(x_5, x_117); +lean_dec(x_117); +if (x_118 == 0) +{ +lean_object* x_119; lean_object* x_120; +lean_dec(x_38); +x_119 = lean_box(0); +x_120 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_3, x_7, x_119, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_37); +x_19 = x_120; +goto block_34; +} +else +{ +lean_object* x_121; lean_object* x_122; uint8_t x_123; +x_121 = lean_unsigned_to_nat(0u); +x_122 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_7, x_121); +x_123 = lean_nat_dec_eq(x_122, x_6); +lean_dec(x_122); +if (x_123 == 0) +{ +lean_object* x_124; lean_object* x_125; +lean_dec(x_38); +x_124 = lean_box(0); +x_125 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_3, x_7, x_124, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_37); +x_19 = x_125; +goto block_34; +} +else +{ +lean_object* x_126; +x_126 = lean_box(0); +x_39 = x_126; +goto block_108; +} +} +} +} +else +{ +lean_object* x_127; uint8_t x_128; +lean_dec(x_36); +x_127 = l_Lean_Expr_getAppFn(x_7); +x_128 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_eqvFunctions(x_5, x_127); +lean_dec(x_127); +if (x_128 == 0) +{ +lean_object* x_129; lean_object* x_130; +lean_dec(x_38); +x_129 = lean_box(0); +x_130 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_3, x_7, x_129, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_37); +x_19 = x_130; +goto block_34; +} +else +{ +lean_object* x_131; lean_object* x_132; uint8_t x_133; +x_131 = lean_unsigned_to_nat(0u); +x_132 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_7, x_131); +x_133 = lean_nat_dec_eq(x_132, x_6); +lean_dec(x_132); +if (x_133 == 0) +{ +lean_object* x_134; lean_object* x_135; +lean_dec(x_38); +x_134 = lean_box(0); +x_135 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_3, x_7, x_134, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_37); +x_19 = x_135; +goto block_34; +} +else +{ +lean_object* x_136; +x_136 = lean_box(0); +x_39 = x_136; +goto block_108; +} +} +} +} +block_108: +{ +lean_object* x_40; +lean_dec(x_39); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_1); +x_40 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_x3f(x_1, x_2, x_7, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_37); +if (lean_obj_tag(x_40) == 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_42; lean_object* x_43; lean_object* x_44; +lean_dec(x_38); +x_42 = lean_ctor_get(x_40, 1); +lean_inc(x_42); +lean_dec(x_40); +x_43 = lean_box(0); +x_44 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_3, x_7, x_43, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_42); +x_19 = x_44; +goto block_34; +} +else +{ +lean_object* x_45; lean_object* x_46; uint8_t x_47; +x_45 = lean_ctor_get(x_41, 0); +lean_inc(x_45); +lean_dec(x_41); +x_46 = lean_ctor_get(x_40, 1); +lean_inc(x_46); +lean_dec(x_40); +x_47 = !lean_is_exclusive(x_45); +if (x_47 == 0) +{ +lean_object* x_48; uint8_t x_49; +x_48 = lean_ctor_get(x_45, 1); +x_49 = lean_nat_dec_le(x_38, x_48); +if (x_49 == 0) +{ +lean_object* x_50; uint8_t x_51; +lean_dec(x_48); +lean_ctor_set(x_45, 1, x_38); +x_50 = lean_st_ref_take(x_9, x_46); +x_51 = !lean_is_exclusive(x_50); +if (x_51 == 0) +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_52 = lean_ctor_get(x_50, 0); +x_53 = lean_ctor_get(x_50, 1); +lean_ctor_set_tag(x_50, 1); +lean_ctor_set(x_50, 1, x_52); +lean_ctor_set(x_50, 0, x_45); +x_54 = lean_st_ref_set(x_9, x_50, x_53); +x_55 = lean_ctor_get(x_54, 1); +lean_inc(x_55); +lean_dec(x_54); +x_56 = lean_box(0); +x_57 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_3, x_7, x_56, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_55); +x_19 = x_57; +goto block_34; +} +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; +x_58 = lean_ctor_get(x_50, 0); +x_59 = lean_ctor_get(x_50, 1); +lean_inc(x_59); +lean_inc(x_58); +lean_dec(x_50); +x_60 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_60, 0, x_45); +lean_ctor_set(x_60, 1, x_58); +x_61 = lean_st_ref_set(x_9, x_60, x_59); +x_62 = lean_ctor_get(x_61, 1); +lean_inc(x_62); +lean_dec(x_61); +x_63 = lean_box(0); +x_64 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_3, x_7, x_63, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_62); +x_19 = x_64; +goto block_34; +} +} +else +{ +lean_object* x_65; uint8_t x_66; +lean_dec(x_38); +x_65 = lean_st_ref_take(x_9, x_46); +x_66 = !lean_is_exclusive(x_65); +if (x_66 == 0) +{ +lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; +x_67 = lean_ctor_get(x_65, 0); +x_68 = lean_ctor_get(x_65, 1); +lean_ctor_set_tag(x_65, 1); +lean_ctor_set(x_65, 1, x_67); +lean_ctor_set(x_65, 0, x_45); +x_69 = lean_st_ref_set(x_9, x_65, x_68); +x_70 = lean_ctor_get(x_69, 1); +lean_inc(x_70); +lean_dec(x_69); +x_71 = lean_box(0); +x_72 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_3, x_7, x_71, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_70); +x_19 = x_72; +goto block_34; +} +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; +x_73 = lean_ctor_get(x_65, 0); +x_74 = lean_ctor_get(x_65, 1); +lean_inc(x_74); +lean_inc(x_73); +lean_dec(x_65); +x_75 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_75, 0, x_45); +lean_ctor_set(x_75, 1, x_73); +x_76 = lean_st_ref_set(x_9, x_75, x_74); +x_77 = lean_ctor_get(x_76, 1); +lean_inc(x_77); +lean_dec(x_76); +x_78 = lean_box(0); +x_79 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_3, x_7, x_78, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_77); +x_19 = x_79; +goto block_34; +} +} +} +else +{ +lean_object* x_80; lean_object* x_81; lean_object* x_82; uint8_t x_83; +x_80 = lean_ctor_get(x_45, 0); +x_81 = lean_ctor_get(x_45, 1); +x_82 = lean_ctor_get(x_45, 2); +lean_inc(x_82); +lean_inc(x_81); +lean_inc(x_80); +lean_dec(x_45); +x_83 = lean_nat_dec_le(x_38, x_81); +if (x_83 == 0) +{ +lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; +lean_dec(x_81); +x_84 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_84, 0, x_80); +lean_ctor_set(x_84, 1, x_38); +lean_ctor_set(x_84, 2, x_82); +x_85 = lean_st_ref_take(x_9, x_46); +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); +} +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_86); +x_90 = lean_st_ref_set(x_9, x_89, x_87); +x_91 = lean_ctor_get(x_90, 1); +lean_inc(x_91); +lean_dec(x_90); +x_92 = lean_box(0); +x_93 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_3, x_7, x_92, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_91); +x_19 = x_93; +goto block_34; +} +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_dec(x_38); +x_94 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_94, 0, x_80); +lean_ctor_set(x_94, 1, x_81); +lean_ctor_set(x_94, 2, x_82); +x_95 = lean_st_ref_take(x_9, x_46); +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_95); + x_98 = lean_box(0); +} +if (lean_is_scalar(x_98)) { + x_99 = lean_alloc_ctor(1, 2, 0); +} else { + x_99 = x_98; + lean_ctor_set_tag(x_99, 1); +} +lean_ctor_set(x_99, 0, x_94); +lean_ctor_set(x_99, 1, x_96); +x_100 = lean_st_ref_set(x_9, x_99, x_97); +x_101 = lean_ctor_get(x_100, 1); +lean_inc(x_101); +lean_dec(x_100); +x_102 = lean_box(0); +x_103 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_3, x_7, x_102, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_101); +x_19 = x_103; +goto block_34; +} +} +} +} +else +{ +uint8_t x_104; +lean_dec(x_38); +lean_dec(x_7); +x_104 = !lean_is_exclusive(x_40); +if (x_104 == 0) +{ +x_19 = x_40; +goto block_34; +} +else +{ +lean_object* x_105; lean_object* x_106; lean_object* x_107; +x_105 = lean_ctor_get(x_40, 0); +x_106 = lean_ctor_get(x_40, 1); +lean_inc(x_106); +lean_inc(x_105); +lean_dec(x_40); +x_107 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_107, 0, x_105); +lean_ctor_set(x_107, 1, x_106); +x_19 = x_107; +goto block_34; +} +} +} +} +else +{ +uint8_t x_137; +lean_dec(x_7); +x_137 = !lean_is_exclusive(x_35); +if (x_137 == 0) +{ +x_19 = x_35; +goto block_34; +} +else +{ +lean_object* x_138; lean_object* x_139; lean_object* x_140; +x_138 = lean_ctor_get(x_35, 0); +x_139 = lean_ctor_get(x_35, 1); +lean_inc(x_139); +lean_inc(x_138); +lean_dec(x_35); +x_140 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_140, 0, x_138); +lean_ctor_set(x_140, 1, x_139); +x_19 = x_140; +goto block_34; +} +} +block_34: +{ +if (lean_obj_tag(x_19) == 0) +{ +lean_object* x_20; +x_20 = lean_ctor_get(x_19, 0); +lean_inc(x_20); +if (lean_obj_tag(x_20) == 0) +{ +uint8_t x_21; +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_1); +x_21 = !lean_is_exclusive(x_19); +if (x_21 == 0) +{ +lean_object* x_22; lean_object* x_23; +x_22 = lean_ctor_get(x_19, 0); +lean_dec(x_22); +x_23 = lean_ctor_get(x_20, 0); +lean_inc(x_23); +lean_dec(x_20); +lean_ctor_set(x_19, 0, x_23); +return x_19; +} +else +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_19, 1); +lean_inc(x_24); +lean_dec(x_19); +x_25 = lean_ctor_get(x_20, 0); +lean_inc(x_25); +lean_dec(x_20); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set(x_26, 1, x_24); +return x_26; +} +} +else +{ +lean_object* x_27; lean_object* x_28; +x_27 = lean_ctor_get(x_19, 1); +lean_inc(x_27); +lean_dec(x_19); +x_28 = lean_ctor_get(x_20, 0); +lean_inc(x_28); +lean_dec(x_20); +x_7 = x_28; +x_18 = x_27; +goto _start; +} +} +else +{ +uint8_t x_30; +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_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; +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_15 = l_Lean_Meta_Grind_getMaxGeneration___rarg(x_8, x_9, x_10, x_11, x_12, x_13, x_14); +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_15, 1); +lean_inc(x_17); +lean_dec(x_15); +x_18 = l_Lean_Expr_getAppFn(x_2); +x_19 = lean_unsigned_to_nat(0u); +x_20 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_2, x_19); +lean_inc(x_3); +x_21 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1(x_1, x_2, x_3, x_16, x_18, x_20, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_17); +lean_dec(x_20); +lean_dec(x_18); +lean_dec(x_16); +lean_dec(x_3); +if (lean_obj_tag(x_21) == 0) +{ +uint8_t x_22; +x_22 = !lean_is_exclusive(x_21); +if (x_22 == 0) +{ +lean_object* x_23; lean_object* x_24; +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_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_28; +x_28 = !lean_is_exclusive(x_21); +if (x_28 == 0) +{ +return x_21; +} +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; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +lean_object* x_15; +x_15 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +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); +return x_15; +} +} +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___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]; +_start: +{ +lean_object* x_19; +x_19 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_19; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +lean_object* x_15; +x_15 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +return x_15; +} +} +static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___spec__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_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19, lean_object* x_20) { +_start: +{ +if (lean_obj_tag(x_7) == 0) +{ +lean_object* x_21; +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_1); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_8); +lean_ctor_set(x_21, 1, x_20); +return x_21; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_30; +lean_dec(x_8); +x_22 = lean_ctor_get(x_7, 0); +lean_inc(x_22); +x_23 = lean_ctor_get(x_7, 1); +lean_inc(x_23); +if (lean_is_exclusive(x_7)) { + lean_ctor_release(x_7, 0); + lean_ctor_release(x_7, 1); + x_24 = x_7; +} else { + lean_dec_ref(x_7); + x_24 = lean_box(0); +} +lean_inc(x_22); +x_30 = l_Lean_Meta_Grind_getENode(x_22, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20); +if (lean_obj_tag(x_30) == 0) +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_83; +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, 8); +lean_inc(x_33); +x_83 = lean_nat_dec_le(x_33, x_4); +if (x_83 == 0) +{ +lean_object* x_84; +lean_dec(x_33); +lean_dec(x_31); +lean_dec(x_24); +lean_dec(x_22); +x_84 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___spec__1___closed__1; +x_25 = x_84; +x_26 = x_32; +goto block_29; +} +else +{ +uint8_t x_85; +x_85 = lean_ctor_get_uint8(x_31, sizeof(void*)*10 + 4); +if (x_85 == 0) +{ +lean_object* x_86; uint8_t x_87; +x_86 = lean_ctor_get(x_31, 3); +lean_inc(x_86); +lean_dec(x_31); +x_87 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_86, x_22); +lean_dec(x_86); +if (x_87 == 0) +{ +lean_object* x_88; +lean_dec(x_33); +lean_dec(x_24); +lean_dec(x_22); +x_88 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___spec__1___closed__1; +x_25 = x_88; +x_26 = x_32; +goto block_29; +} +else +{ +lean_object* x_89; +x_89 = lean_box(0); +x_34 = x_89; +goto block_82; +} +} +else +{ +lean_object* x_90; +lean_dec(x_31); +x_90 = lean_box(0); +x_34 = x_90; +goto block_82; +} +} +block_82: +{ +lean_object* x_35; +lean_dec(x_34); +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_1); +x_35 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_x3f(x_1, x_2, x_22, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_32); +lean_dec(x_22); +if (lean_obj_tag(x_35) == 0) +{ +lean_object* x_36; +x_36 = lean_ctor_get(x_35, 0); +lean_inc(x_36); +if (lean_obj_tag(x_36) == 0) +{ +lean_object* x_37; lean_object* x_38; +lean_dec(x_33); +lean_dec(x_24); +x_37 = lean_ctor_get(x_35, 1); +lean_inc(x_37); +lean_dec(x_35); +x_38 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___spec__1___closed__1; +x_25 = x_38; +x_26 = x_37; +goto block_29; +} +else +{ +lean_object* x_39; lean_object* x_40; uint8_t x_41; +x_39 = lean_ctor_get(x_36, 0); +lean_inc(x_39); +lean_dec(x_36); +x_40 = lean_ctor_get(x_35, 1); +lean_inc(x_40); +lean_dec(x_35); +x_41 = !lean_is_exclusive(x_39); +if (x_41 == 0) +{ +lean_object* x_42; uint8_t x_43; +x_42 = lean_ctor_get(x_39, 1); +x_43 = lean_nat_dec_le(x_33, x_42); +if (x_43 == 0) +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; +lean_dec(x_42); +lean_ctor_set(x_39, 1, x_33); +x_44 = lean_st_ref_take(x_11, x_40); +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); +if (lean_is_scalar(x_24)) { + x_47 = lean_alloc_ctor(1, 2, 0); +} else { + x_47 = x_24; +} +lean_ctor_set(x_47, 0, x_39); +lean_ctor_set(x_47, 1, x_45); +x_48 = lean_st_ref_set(x_11, x_47, x_46); +x_49 = lean_ctor_get(x_48, 1); +lean_inc(x_49); +lean_dec(x_48); +x_50 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___spec__1___closed__1; +x_25 = x_50; +x_26 = x_49; +goto block_29; +} +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_dec(x_33); +x_51 = lean_st_ref_take(x_11, x_40); +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); +if (lean_is_scalar(x_24)) { + x_54 = lean_alloc_ctor(1, 2, 0); +} else { + x_54 = x_24; +} +lean_ctor_set(x_54, 0, x_39); +lean_ctor_set(x_54, 1, x_52); +x_55 = lean_st_ref_set(x_11, x_54, x_53); +x_56 = lean_ctor_get(x_55, 1); +lean_inc(x_56); +lean_dec(x_55); +x_57 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___spec__1___closed__1; +x_25 = x_57; +x_26 = x_56; +goto block_29; +} +} +else +{ +lean_object* x_58; lean_object* x_59; lean_object* x_60; uint8_t x_61; +x_58 = lean_ctor_get(x_39, 0); +x_59 = lean_ctor_get(x_39, 1); +x_60 = lean_ctor_get(x_39, 2); +lean_inc(x_60); +lean_inc(x_59); +lean_inc(x_58); +lean_dec(x_39); +x_61 = lean_nat_dec_le(x_33, x_59); +if (x_61 == 0) +{ +lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; +lean_dec(x_59); +x_62 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_62, 0, x_58); +lean_ctor_set(x_62, 1, x_33); +lean_ctor_set(x_62, 2, x_60); +x_63 = lean_st_ref_take(x_11, x_40); +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); +if (lean_is_scalar(x_24)) { + x_66 = lean_alloc_ctor(1, 2, 0); +} else { + x_66 = x_24; +} +lean_ctor_set(x_66, 0, x_62); +lean_ctor_set(x_66, 1, x_64); +x_67 = lean_st_ref_set(x_11, x_66, x_65); +x_68 = lean_ctor_get(x_67, 1); +lean_inc(x_68); +lean_dec(x_67); +x_69 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___spec__1___closed__1; +x_25 = x_69; +x_26 = x_68; +goto block_29; +} +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_33); +x_70 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_70, 0, x_58); +lean_ctor_set(x_70, 1, x_59); +lean_ctor_set(x_70, 2, x_60); +x_71 = lean_st_ref_take(x_11, x_40); +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); +if (lean_is_scalar(x_24)) { + x_74 = lean_alloc_ctor(1, 2, 0); +} else { + x_74 = x_24; +} +lean_ctor_set(x_74, 0, x_70); +lean_ctor_set(x_74, 1, x_72); +x_75 = lean_st_ref_set(x_11, x_74, x_73); +x_76 = lean_ctor_get(x_75, 1); +lean_inc(x_76); +lean_dec(x_75); +x_77 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___spec__1___closed__1; +x_25 = x_77; +x_26 = x_76; +goto block_29; +} +} +} +} +else +{ +uint8_t x_78; +lean_dec(x_33); +lean_dec(x_24); +lean_dec(x_23); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_1); +x_78 = !lean_is_exclusive(x_35); +if (x_78 == 0) +{ +return x_35; +} +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_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; +} +} +} +} +else +{ +uint8_t x_91; +lean_dec(x_24); +lean_dec(x_23); +lean_dec(x_22); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_1); +x_91 = !lean_is_exclusive(x_30); +if (x_91 == 0) +{ +return x_30; +} +else +{ +lean_object* x_92; lean_object* x_93; lean_object* x_94; +x_92 = lean_ctor_get(x_30, 0); +x_93 = lean_ctor_get(x_30, 1); +lean_inc(x_93); +lean_inc(x_92); +lean_dec(x_30); +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; +} +} +block_29: +{ +lean_object* x_27; +x_27 = lean_ctor_get(x_25, 0); +lean_inc(x_27); +lean_dec(x_25); +x_7 = x_23; +x_8 = x_27; +x_9 = lean_box(0); +x_20 = x_26; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, 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_st_ref_get(x_5, x_13); +x_15 = !lean_is_exclusive(x_14); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_16 = lean_ctor_get(x_14, 0); +x_17 = lean_ctor_get(x_14, 1); +x_18 = lean_ctor_get(x_16, 4); +lean_inc(x_18); +lean_dec(x_16); +lean_inc(x_2); +x_19 = l_Lean_Expr_toHeadIndex(x_2); +x_20 = l_Lean_PersistentHashMap_find_x3f___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__1(x_18, x_19); +lean_dec(x_19); +if (lean_obj_tag(x_20) == 0) +{ +lean_object* 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_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_2); +lean_dec(x_1); +x_21 = lean_box(0); +lean_ctor_set(x_14, 0, x_21); +return x_14; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +lean_free_object(x_14); +x_22 = lean_ctor_get(x_20, 0); +lean_inc(x_22); +lean_dec(x_20); +x_23 = l_Lean_Meta_Grind_getMaxGeneration___rarg(x_7, x_8, x_9, x_10, x_11, x_12, x_17); +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +x_25 = lean_ctor_get(x_23, 1); +lean_inc(x_25); +lean_dec(x_23); +x_26 = lean_box(0); +x_27 = lean_box(0); +lean_inc(x_22); +x_28 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___spec__1(x_1, x_2, x_22, x_24, x_26, x_22, x_22, x_27, lean_box(0), x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_25); +lean_dec(x_24); +lean_dec(x_22); +lean_dec(x_2); +if (lean_obj_tag(x_28) == 0) +{ +uint8_t x_29; +x_29 = !lean_is_exclusive(x_28); +if (x_29 == 0) +{ +lean_object* x_30; +x_30 = lean_ctor_get(x_28, 0); +lean_dec(x_30); +lean_ctor_set(x_28, 0, x_27); +return x_28; +} +else +{ +lean_object* x_31; lean_object* x_32; +x_31 = lean_ctor_get(x_28, 1); +lean_inc(x_31); +lean_dec(x_28); +x_32 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_32, 0, x_27); +lean_ctor_set(x_32, 1, x_31); +return x_32; +} +} +else +{ +uint8_t x_33; +x_33 = !lean_is_exclusive(x_28); +if (x_33 == 0) +{ +return x_28; +} +else +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_28, 0); +x_35 = lean_ctor_get(x_28, 1); +lean_inc(x_35); +lean_inc(x_34); +lean_dec(x_28); +x_36 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_36, 0, x_34); +lean_ctor_set(x_36, 1, x_35); +return x_36; +} +} +} +} +else +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_37 = lean_ctor_get(x_14, 0); +x_38 = lean_ctor_get(x_14, 1); +lean_inc(x_38); +lean_inc(x_37); +lean_dec(x_14); +x_39 = lean_ctor_get(x_37, 4); +lean_inc(x_39); +lean_dec(x_37); +lean_inc(x_2); +x_40 = l_Lean_Expr_toHeadIndex(x_2); +x_41 = l_Lean_PersistentHashMap_find_x3f___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__1(x_39, x_40); +lean_dec(x_40); +if (lean_obj_tag(x_41) == 0) +{ +lean_object* x_42; lean_object* 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_2); +lean_dec(x_1); +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 +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_44 = lean_ctor_get(x_41, 0); +lean_inc(x_44); +lean_dec(x_41); +x_45 = l_Lean_Meta_Grind_getMaxGeneration___rarg(x_7, x_8, x_9, x_10, x_11, x_12, x_38); +x_46 = lean_ctor_get(x_45, 0); +lean_inc(x_46); +x_47 = lean_ctor_get(x_45, 1); +lean_inc(x_47); +lean_dec(x_45); +x_48 = lean_box(0); +x_49 = lean_box(0); +lean_inc(x_44); +x_50 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___spec__1(x_1, x_2, x_44, x_46, x_48, x_44, x_44, x_49, lean_box(0), x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_47); +lean_dec(x_46); +lean_dec(x_44); +lean_dec(x_2); +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, 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); +} +if (lean_is_scalar(x_52)) { + x_53 = lean_alloc_ctor(0, 2, 0); +} else { + x_53 = x_52; +} +lean_ctor_set(x_53, 0, x_49); +lean_ctor_set(x_53, 1, x_51); +return x_53; +} +else +{ +lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_54 = lean_ctor_get(x_50, 0); +lean_inc(x_54); +x_55 = lean_ctor_get(x_50, 1); +lean_inc(x_55); +if (lean_is_exclusive(x_50)) { + lean_ctor_release(x_50, 0); + lean_ctor_release(x_50, 1); + x_56 = x_50; +} else { + lean_dec_ref(x_50); + x_56 = lean_box(0); +} +if (lean_is_scalar(x_56)) { + x_57 = lean_alloc_ctor(1, 2, 0); +} else { + x_57 = x_56; +} +lean_ctor_set(x_57, 0, x_54); +lean_ctor_set(x_57, 1, x_55); +return x_57; +} +} +} +} +} +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___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_object* x_19 = _args[18]; +lean_object* x_20 = _args[19]; +_start: +{ +lean_object* x_21; +x_21 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_21; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, 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___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_4); +lean_dec(x_3); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +uint8_t x_13; +x_13 = l_Lean_Expr_hasMVar(x_1); +if (x_13 == 0) +{ +lean_object* x_14; +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_1); +lean_ctor_set(x_14, 1, x_12); +return x_14; +} +else +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; +x_15 = lean_st_ref_get(x_9, 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 = lean_ctor_get(x_16, 0); +lean_inc(x_18); +lean_dec(x_16); +x_19 = l_Lean_instantiateMVarsCore(x_18, x_1); +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_st_ref_take(x_9, x_17); +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; uint8_t x_28; +x_26 = lean_ctor_get(x_23, 0); +lean_dec(x_26); +lean_ctor_set(x_23, 0, x_21); +x_27 = lean_st_ref_set(x_9, x_23, x_24); +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_20); +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_20); +lean_ctor_set(x_31, 1, x_30); +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; +x_32 = lean_ctor_get(x_23, 1); +x_33 = lean_ctor_get(x_23, 2); +x_34 = lean_ctor_get(x_23, 3); +x_35 = lean_ctor_get(x_23, 4); +lean_inc(x_35); +lean_inc(x_34); +lean_inc(x_33); +lean_inc(x_32); +lean_dec(x_23); +x_36 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_36, 0, x_21); +lean_ctor_set(x_36, 1, x_32); +lean_ctor_set(x_36, 2, x_33); +lean_ctor_set(x_36, 3, x_34); +lean_ctor_set(x_36, 4, x_35); +x_37 = lean_st_ref_set(x_9, x_36, x_24); +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); +} +if (lean_is_scalar(x_39)) { + x_40 = lean_alloc_ctor(0, 2, 0); +} else { + x_40 = x_39; +} +lean_ctor_set(x_40, 0, x_20); +lean_ctor_set(x_40, 1, x_38); +return x_40; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_13 = lean_ctor_get(x_10, 2); +x_14 = l_Lean_isTracingEnabledForCore(x_1, x_13, x_12); +x_15 = !lean_is_exclusive(x_14); +if (x_15 == 0) +{ +return x_14; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_14, 0); +x_17 = lean_ctor_get(x_14, 1); +lean_inc(x_17); +lean_inc(x_16); +lean_dec(x_14); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_16); +lean_ctor_set(x_18, 1, x_17); +return x_18; +} +} +} +LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___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) { +_start: +{ +lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_13 = lean_ctor_get(x_10, 5); +x_14 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_1, x_8, x_9, x_10, x_11, x_12); +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_inc(x_13); +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_13); +lean_ctor_set(x_17, 1, x_16); +lean_ctor_set_tag(x_14, 1); +lean_ctor_set(x_14, 0, x_17); +return x_14; +} +else +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_18 = lean_ctor_get(x_14, 0); +x_19 = lean_ctor_get(x_14, 1); +lean_inc(x_19); +lean_inc(x_18); +lean_dec(x_14); +lean_inc(x_13); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_13); +lean_ctor_set(x_20, 1, x_18); +x_21 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_21, 0, x_20); +lean_ctor_set(x_21, 1, x_19); +return x_21; +} +} +} +static lean_object* _init_l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("unknown constant '", 18, 18); +return x_1; +} +} +static lean_object* _init_l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("'", 1, 1); +return x_1; +} +} +static lean_object* _init_l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5___closed__3; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; uint8_t x_14; +x_13 = lean_st_ref_get(x_11, x_12); +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; +x_15 = lean_ctor_get(x_13, 0); +x_16 = lean_ctor_get(x_13, 1); +x_17 = lean_ctor_get(x_15, 0); +lean_inc(x_17); +lean_dec(x_15); +lean_inc(x_1); +x_18 = lean_environment_find(x_17, x_1); +if (lean_obj_tag(x_18) == 0) +{ +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_free_object(x_13); +x_19 = 0; +x_20 = l_Lean_MessageData_ofConstName(x_1, x_19); +x_21 = l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5___closed__2; +x_22 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_20); +x_23 = l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5___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_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__6(x_24, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_16); +return x_25; +} +else +{ +lean_object* x_26; +lean_dec(x_1); +x_26 = lean_ctor_get(x_18, 0); +lean_inc(x_26); +lean_dec(x_18); +lean_ctor_set(x_13, 0, x_26); +return x_13; +} +} +else +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +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 = lean_ctor_get(x_27, 0); +lean_inc(x_29); +lean_dec(x_27); +lean_inc(x_1); +x_30 = lean_environment_find(x_29, x_1); +if (lean_obj_tag(x_30) == 0) +{ +uint8_t x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_31 = 0; +x_32 = l_Lean_MessageData_ofConstName(x_1, x_31); +x_33 = l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5___closed__2; +x_34 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_34, 0, x_33); +lean_ctor_set(x_34, 1, x_32); +x_35 = l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5___closed__4; +x_36 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_36, 0, x_34); +lean_ctor_set(x_36, 1, x_35); +x_37 = l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__6(x_36, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_28); +return x_37; +} +else +{ +lean_object* x_38; lean_object* x_39; +lean_dec(x_1); +x_38 = lean_ctor_get(x_30, 0); +lean_inc(x_38); +lean_dec(x_30); +x_39 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_28); +return x_39; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_mkConstWithLevelParams___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +lean_inc(x_1); +x_13 = l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___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); +if (lean_obj_tag(x_13) == 0) +{ +uint8_t x_14; +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_15 = lean_ctor_get(x_13, 0); +x_16 = l_Lean_ConstantInfo_levelParams(x_15); +lean_dec(x_15); +x_17 = lean_box(0); +x_18 = l_List_mapTR_loop___at_Lean_mkConstWithLevelParams___spec__1(x_16, x_17); +x_19 = l_Lean_Expr_const___override(x_1, x_18); +lean_ctor_set(x_13, 0, x_19); +return x_13; +} +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_20 = lean_ctor_get(x_13, 0); +x_21 = lean_ctor_get(x_13, 1); +lean_inc(x_21); +lean_inc(x_20); +lean_dec(x_13); +x_22 = l_Lean_ConstantInfo_levelParams(x_20); +lean_dec(x_20); +x_23 = lean_box(0); +x_24 = l_List_mapTR_loop___at_Lean_mkConstWithLevelParams___spec__1(x_22, x_23); +x_25 = l_Lean_Expr_const___override(x_1, x_24); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set(x_26, 1, x_21); +return x_26; +} +} +else +{ +uint8_t x_27; +lean_dec(x_1); +x_27 = !lean_is_exclusive(x_13); +if (x_27 == 0) +{ +return x_13; +} +else +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_28 = lean_ctor_get(x_13, 0); +x_29 = lean_ctor_get(x_13, 1); +lean_inc(x_29); +lean_inc(x_28); +lean_dec(x_13); +x_30 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_30, 0, x_28); +lean_ctor_set(x_30, 1, x_29); +return x_30; +} +} +} +} +static lean_object* _init_l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("[unknown]", 9, 9); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__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_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3___closed__2; +x_2 = l_Lean_MessageData_ofFormat(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +switch (lean_obj_tag(x_1)) { +case 0: +{ +lean_object* x_13; lean_object* x_14; +x_13 = lean_ctor_get(x_1, 0); +lean_inc(x_13); +lean_dec(x_1); +x_14 = l_Lean_mkConstWithLevelParams___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__4(x_13, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_14) == 0) +{ +uint8_t x_15; +x_15 = !lean_is_exclusive(x_14); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; +x_16 = lean_ctor_get(x_14, 0); +x_17 = l_Lean_MessageData_ofConst(x_16); +lean_ctor_set(x_14, 0, x_17); +return x_14; +} +else +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_18 = lean_ctor_get(x_14, 0); +x_19 = lean_ctor_get(x_14, 1); +lean_inc(x_19); +lean_inc(x_18); +lean_dec(x_14); +x_20 = l_Lean_MessageData_ofConst(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; +} +} +else +{ +uint8_t x_22; +x_22 = !lean_is_exclusive(x_14); +if (x_22 == 0) +{ +return x_14; +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_14, 0); +x_24 = lean_ctor_get(x_14, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_14); +x_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 1: +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_26 = lean_ctor_get(x_1, 0); +lean_inc(x_26); +lean_dec(x_1); +x_27 = l_Lean_Expr_fvar___override(x_26); +x_28 = l_Lean_MessageData_ofExpr(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_12); +return x_29; +} +case 2: +{ +uint8_t x_30; +x_30 = !lean_is_exclusive(x_1); +if (x_30 == 0) +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_1, 1); +x_32 = lean_ctor_get(x_1, 0); +lean_dec(x_32); +x_33 = l_Lean_MessageData_ofSyntax(x_31); +lean_ctor_set_tag(x_1, 0); +lean_ctor_set(x_1, 1, x_12); +lean_ctor_set(x_1, 0, x_33); +return x_1; +} +else +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_1, 1); +lean_inc(x_34); +lean_dec(x_1); +x_35 = l_Lean_MessageData_ofSyntax(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_12); +return x_36; +} +} +default: +{ +lean_object* x_37; lean_object* x_38; +x_37 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3___closed__3; +x_38 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_38, 0, x_37); +lean_ctor_set(x_38, 1, x_12); +return x_38; +} +} +} +} +static double _init_l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___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_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(0); +x_2 = lean_array_mk(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_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; uint8_t x_21; +x_14 = lean_ctor_get(x_11, 5); +x_15 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_2, x_9, x_10, x_11, x_12, x_13); +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_15, 1); +lean_inc(x_17); +lean_dec(x_15); +x_18 = lean_st_ref_take(x_12, x_17); +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_19, 3); +lean_inc(x_20); +x_21 = !lean_is_exclusive(x_18); +if (x_21 == 0) +{ +lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_22 = lean_ctor_get(x_18, 1); +x_23 = lean_ctor_get(x_18, 0); +lean_dec(x_23); +x_24 = !lean_is_exclusive(x_19); +if (x_24 == 0) +{ +lean_object* x_25; uint8_t x_26; +x_25 = lean_ctor_get(x_19, 3); +lean_dec(x_25); +x_26 = !lean_is_exclusive(x_20); +if (x_26 == 0) +{ +lean_object* x_27; double 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; uint8_t x_36; +x_27 = lean_ctor_get(x_20, 0); +x_28 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__1; +x_29 = 0; +x_30 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__1; +x_31 = lean_alloc_ctor(0, 2, 17); +lean_ctor_set(x_31, 0, x_1); +lean_ctor_set(x_31, 1, x_30); +lean_ctor_set_float(x_31, sizeof(void*)*2, x_28); +lean_ctor_set_float(x_31, sizeof(void*)*2 + 8, x_28); +lean_ctor_set_uint8(x_31, sizeof(void*)*2 + 16, x_29); +x_32 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2; +x_33 = lean_alloc_ctor(9, 3, 0); +lean_ctor_set(x_33, 0, x_31); +lean_ctor_set(x_33, 1, x_16); +lean_ctor_set(x_33, 2, x_32); +lean_inc(x_14); +lean_ctor_set(x_18, 1, x_33); +lean_ctor_set(x_18, 0, x_14); +x_34 = l_Lean_PersistentArray_push___rarg(x_27, x_18); +lean_ctor_set(x_20, 0, x_34); +x_35 = lean_st_ref_set(x_12, x_19, x_22); +x_36 = !lean_is_exclusive(x_35); +if (x_36 == 0) +{ +lean_object* x_37; lean_object* x_38; +x_37 = lean_ctor_get(x_35, 0); +lean_dec(x_37); +x_38 = lean_box(0); +lean_ctor_set(x_35, 0, x_38); +return x_35; +} +else +{ +lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_39 = lean_ctor_get(x_35, 1); +lean_inc(x_39); +lean_dec(x_35); +x_40 = lean_box(0); +x_41 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_41, 0, x_40); +lean_ctor_set(x_41, 1, x_39); +return x_41; +} +} +else +{ +uint64_t x_42; lean_object* x_43; double 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; +x_42 = lean_ctor_get_uint64(x_20, sizeof(void*)*1); +x_43 = lean_ctor_get(x_20, 0); +lean_inc(x_43); +lean_dec(x_20); +x_44 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__1; +x_45 = 0; +x_46 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__1; +x_47 = lean_alloc_ctor(0, 2, 17); +lean_ctor_set(x_47, 0, x_1); +lean_ctor_set(x_47, 1, x_46); +lean_ctor_set_float(x_47, sizeof(void*)*2, x_44); +lean_ctor_set_float(x_47, sizeof(void*)*2 + 8, x_44); +lean_ctor_set_uint8(x_47, sizeof(void*)*2 + 16, x_45); +x_48 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2; +x_49 = lean_alloc_ctor(9, 3, 0); +lean_ctor_set(x_49, 0, x_47); +lean_ctor_set(x_49, 1, x_16); +lean_ctor_set(x_49, 2, x_48); +lean_inc(x_14); +lean_ctor_set(x_18, 1, x_49); +lean_ctor_set(x_18, 0, x_14); +x_50 = l_Lean_PersistentArray_push___rarg(x_43, x_18); +x_51 = lean_alloc_ctor(0, 1, 8); +lean_ctor_set(x_51, 0, x_50); +lean_ctor_set_uint64(x_51, sizeof(void*)*1, x_42); +lean_ctor_set(x_19, 3, x_51); +x_52 = lean_st_ref_set(x_12, x_19, x_22); +x_53 = lean_ctor_get(x_52, 1); +lean_inc(x_53); +if (lean_is_exclusive(x_52)) { + lean_ctor_release(x_52, 0); + lean_ctor_release(x_52, 1); + x_54 = x_52; +} else { + lean_dec_ref(x_52); + x_54 = lean_box(0); +} +x_55 = lean_box(0); +if (lean_is_scalar(x_54)) { + x_56 = lean_alloc_ctor(0, 2, 0); +} else { + x_56 = x_54; +} +lean_ctor_set(x_56, 0, x_55); +lean_ctor_set(x_56, 1, x_53); +return x_56; +} +} +else +{ +lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; uint64_t x_64; lean_object* x_65; lean_object* x_66; double x_67; uint8_t x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; +x_57 = lean_ctor_get(x_19, 0); +x_58 = lean_ctor_get(x_19, 1); +x_59 = lean_ctor_get(x_19, 2); +x_60 = lean_ctor_get(x_19, 4); +x_61 = lean_ctor_get(x_19, 5); +x_62 = lean_ctor_get(x_19, 6); +x_63 = lean_ctor_get(x_19, 7); +lean_inc(x_63); +lean_inc(x_62); +lean_inc(x_61); +lean_inc(x_60); +lean_inc(x_59); +lean_inc(x_58); +lean_inc(x_57); +lean_dec(x_19); +x_64 = lean_ctor_get_uint64(x_20, sizeof(void*)*1); +x_65 = lean_ctor_get(x_20, 0); +lean_inc(x_65); +if (lean_is_exclusive(x_20)) { + lean_ctor_release(x_20, 0); + x_66 = x_20; +} else { + lean_dec_ref(x_20); + x_66 = lean_box(0); +} +x_67 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__1; +x_68 = 0; +x_69 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__1; +x_70 = lean_alloc_ctor(0, 2, 17); +lean_ctor_set(x_70, 0, x_1); +lean_ctor_set(x_70, 1, x_69); +lean_ctor_set_float(x_70, sizeof(void*)*2, x_67); +lean_ctor_set_float(x_70, sizeof(void*)*2 + 8, x_67); +lean_ctor_set_uint8(x_70, sizeof(void*)*2 + 16, x_68); +x_71 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2; +x_72 = lean_alloc_ctor(9, 3, 0); +lean_ctor_set(x_72, 0, x_70); +lean_ctor_set(x_72, 1, x_16); +lean_ctor_set(x_72, 2, x_71); +lean_inc(x_14); +lean_ctor_set(x_18, 1, x_72); +lean_ctor_set(x_18, 0, x_14); +x_73 = l_Lean_PersistentArray_push___rarg(x_65, x_18); +if (lean_is_scalar(x_66)) { + x_74 = lean_alloc_ctor(0, 1, 8); +} else { + x_74 = x_66; +} +lean_ctor_set(x_74, 0, x_73); +lean_ctor_set_uint64(x_74, sizeof(void*)*1, x_64); +x_75 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_75, 0, x_57); +lean_ctor_set(x_75, 1, x_58); +lean_ctor_set(x_75, 2, x_59); +lean_ctor_set(x_75, 3, x_74); +lean_ctor_set(x_75, 4, x_60); +lean_ctor_set(x_75, 5, x_61); +lean_ctor_set(x_75, 6, x_62); +lean_ctor_set(x_75, 7, x_63); +x_76 = lean_st_ref_set(x_12, x_75, x_22); +x_77 = lean_ctor_get(x_76, 1); +lean_inc(x_77); +if (lean_is_exclusive(x_76)) { + lean_ctor_release(x_76, 0); + lean_ctor_release(x_76, 1); + x_78 = x_76; +} else { + lean_dec_ref(x_76); + x_78 = lean_box(0); +} +x_79 = lean_box(0); +if (lean_is_scalar(x_78)) { + x_80 = lean_alloc_ctor(0, 2, 0); +} else { + x_80 = x_78; +} +lean_ctor_set(x_80, 0, x_79); +lean_ctor_set(x_80, 1, x_77); +return x_80; +} +} +else +{ +lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; uint64_t x_90; lean_object* x_91; lean_object* x_92; double x_93; uint8_t x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; +x_81 = lean_ctor_get(x_18, 1); +lean_inc(x_81); +lean_dec(x_18); +x_82 = lean_ctor_get(x_19, 0); +lean_inc(x_82); +x_83 = lean_ctor_get(x_19, 1); +lean_inc(x_83); +x_84 = lean_ctor_get(x_19, 2); +lean_inc(x_84); +x_85 = lean_ctor_get(x_19, 4); +lean_inc(x_85); +x_86 = lean_ctor_get(x_19, 5); +lean_inc(x_86); +x_87 = lean_ctor_get(x_19, 6); +lean_inc(x_87); +x_88 = lean_ctor_get(x_19, 7); +lean_inc(x_88); +if (lean_is_exclusive(x_19)) { + lean_ctor_release(x_19, 0); + lean_ctor_release(x_19, 1); + lean_ctor_release(x_19, 2); + lean_ctor_release(x_19, 3); + lean_ctor_release(x_19, 4); + lean_ctor_release(x_19, 5); + lean_ctor_release(x_19, 6); + lean_ctor_release(x_19, 7); + x_89 = x_19; +} else { + lean_dec_ref(x_19); + x_89 = lean_box(0); +} +x_90 = lean_ctor_get_uint64(x_20, sizeof(void*)*1); +x_91 = lean_ctor_get(x_20, 0); +lean_inc(x_91); +if (lean_is_exclusive(x_20)) { + lean_ctor_release(x_20, 0); + x_92 = x_20; +} else { + lean_dec_ref(x_20); + x_92 = lean_box(0); +} +x_93 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__1; +x_94 = 0; +x_95 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__1; +x_96 = lean_alloc_ctor(0, 2, 17); +lean_ctor_set(x_96, 0, x_1); +lean_ctor_set(x_96, 1, x_95); +lean_ctor_set_float(x_96, sizeof(void*)*2, x_93); +lean_ctor_set_float(x_96, sizeof(void*)*2 + 8, x_93); +lean_ctor_set_uint8(x_96, sizeof(void*)*2 + 16, x_94); +x_97 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2; +x_98 = lean_alloc_ctor(9, 3, 0); +lean_ctor_set(x_98, 0, x_96); +lean_ctor_set(x_98, 1, x_16); +lean_ctor_set(x_98, 2, x_97); +lean_inc(x_14); +x_99 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_99, 0, x_14); +lean_ctor_set(x_99, 1, x_98); +x_100 = l_Lean_PersistentArray_push___rarg(x_91, x_99); +if (lean_is_scalar(x_92)) { + x_101 = lean_alloc_ctor(0, 1, 8); +} else { + x_101 = x_92; +} +lean_ctor_set(x_101, 0, x_100); +lean_ctor_set_uint64(x_101, sizeof(void*)*1, x_90); +if (lean_is_scalar(x_89)) { + x_102 = lean_alloc_ctor(0, 8, 0); +} else { + x_102 = x_89; +} +lean_ctor_set(x_102, 0, x_82); +lean_ctor_set(x_102, 1, x_83); +lean_ctor_set(x_102, 2, x_84); +lean_ctor_set(x_102, 3, x_101); +lean_ctor_set(x_102, 4, x_85); +lean_ctor_set(x_102, 5, x_86); +lean_ctor_set(x_102, 6, x_87); +lean_ctor_set(x_102, 7, x_88); +x_103 = lean_st_ref_set(x_12, x_102, x_81); +x_104 = lean_ctor_get(x_103, 1); +lean_inc(x_104); +if (lean_is_exclusive(x_103)) { + lean_ctor_release(x_103, 0); + lean_ctor_release(x_103, 1); + x_105 = x_103; +} else { + lean_dec_ref(x_103); + x_105 = lean_box(0); +} +x_106 = lean_box(0); +if (lean_is_scalar(x_105)) { + x_107 = lean_alloc_ctor(0, 2, 0); +} else { + x_107 = x_105; +} +lean_ctor_set(x_107, 0, x_106); +lean_ctor_set(x_107, 1, x_104); +return x_107; +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___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; +x_16 = l_Lean_Meta_Grind_addTheoremInstance(x_1, x_2, x_3, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +return x_16; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("grind", 5, 5); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("ematch", 6, 6); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("instance", 8, 8); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__2; +x_3 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__3; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(": ", 2, 2); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__5; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +_start: +{ +lean_object* x_16; +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_1); +x_16 = lean_infer_type(x_1, x_11, x_12, x_13, x_14, x_15); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +x_18 = lean_ctor_get(x_16, 1); +lean_inc(x_18); +lean_dec(x_16); +x_19 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__4; +x_20 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__2(x_19, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_18); +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_dec(x_3); +x_23 = lean_ctor_get(x_20, 1); +lean_inc(x_23); +lean_dec(x_20); +x_24 = l_Lean_Meta_Grind_addTheoremInstance(x_1, x_17, x_2, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_23); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +return x_24; +} +else +{ +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; +x_26 = lean_ctor_get(x_20, 1); +x_27 = lean_ctor_get(x_20, 0); +lean_dec(x_27); +x_28 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_3, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_26); +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; 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_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_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; +lean_ctor_set_tag(x_20, 7); +lean_ctor_set(x_20, 1, x_29); +lean_ctor_set(x_20, 0, x_31); +x_32 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__6; +x_33 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_33, 0, x_20); +lean_ctor_set(x_33, 1, x_32); +lean_inc(x_17); +x_34 = l_Lean_MessageData_ofExpr(x_17); +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 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set(x_36, 1, x_31); +x_37 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_19, x_36, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_30); +x_38 = lean_ctor_get(x_37, 1); +lean_inc(x_38); +lean_dec(x_37); +x_39 = l_Lean_Meta_Grind_addTheoremInstance(x_1, x_17, x_2, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_38); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +return x_39; +} +else +{ +uint8_t x_40; +lean_free_object(x_20); +lean_dec(x_17); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_2); +lean_dec(x_1); +x_40 = !lean_is_exclusive(x_28); +if (x_40 == 0) +{ +return x_28; +} +else +{ +lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_41 = lean_ctor_get(x_28, 0); +x_42 = lean_ctor_get(x_28, 1); +lean_inc(x_42); +lean_inc(x_41); +lean_dec(x_28); +x_43 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_43, 0, x_41); +lean_ctor_set(x_43, 1, x_42); +return x_43; +} +} +} +else +{ +lean_object* x_44; lean_object* x_45; +x_44 = lean_ctor_get(x_20, 1); +lean_inc(x_44); +lean_dec(x_20); +x_45 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_3, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, 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; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_46 = lean_ctor_get(x_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_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___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_46); +x_50 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__6; +x_51 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_51, 0, x_49); +lean_ctor_set(x_51, 1, x_50); +lean_inc(x_17); +x_52 = l_Lean_MessageData_ofExpr(x_17); +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_48); +x_55 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_19, x_54, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_47); +x_56 = lean_ctor_get(x_55, 1); +lean_inc(x_56); +lean_dec(x_55); +x_57 = l_Lean_Meta_Grind_addTheoremInstance(x_1, x_17, x_2, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_56); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +return x_57; +} +else +{ +lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; +lean_dec(x_17); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_2); +lean_dec(x_1); +x_58 = lean_ctor_get(x_45, 0); +lean_inc(x_58); +x_59 = lean_ctor_get(x_45, 1); +lean_inc(x_59); +if (lean_is_exclusive(x_45)) { + lean_ctor_release(x_45, 0); + lean_ctor_release(x_45, 1); + x_60 = x_45; +} else { + lean_dec_ref(x_45); + x_60 = lean_box(0); +} +if (lean_is_scalar(x_60)) { + x_61 = lean_alloc_ctor(1, 2, 0); +} else { + x_61 = x_60; +} +lean_ctor_set(x_61, 0, x_58); +lean_ctor_set(x_61, 1, x_59); +return x_61; +} +} +} +} +else +{ +uint8_t x_62; +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_62 = !lean_is_exclusive(x_16); +if (x_62 == 0) +{ +return x_16; +} +else +{ +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_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; +} +} +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Meta_Grind_grind_debug_proofs; +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, 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; +x_15 = l_Lean_instantiateMVars___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__1(x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_15, 1); +lean_inc(x_17); +lean_dec(x_15); +x_18 = lean_ctor_get(x_12, 2); +lean_inc(x_18); +x_19 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___closed__1; +x_20 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_18, x_19); +lean_dec(x_18); +if (x_20 == 0) +{ +lean_object* x_21; lean_object* x_22; +x_21 = lean_box(0); +x_22 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2(x_16, x_3, x_1, x_21, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_17); +return x_22; +} +else +{ +lean_object* x_23; +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_16); +x_23 = l_Lean_Meta_check(x_16, x_10, x_11, x_12, x_13, x_17); +if (lean_obj_tag(x_23) == 0) +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +x_25 = lean_ctor_get(x_23, 1); +lean_inc(x_25); +lean_dec(x_23); +x_26 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2(x_16, x_3, x_1, x_24, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_25); +lean_dec(x_24); +return x_26; +} +else +{ +uint8_t x_27; +lean_dec(x_16); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_3); +lean_dec(x_1); +x_27 = !lean_is_exclusive(x_23); +if (x_27 == 0) +{ +return x_23; +} +else +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_28 = lean_ctor_get(x_23, 0); +x_29 = lean_ctor_get(x_23, 1); +lean_inc(x_29); +lean_inc(x_28); +lean_dec(x_23); +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_instantiateMVars___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +x_13 = l_Lean_instantiateMVars___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +x_13 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___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) { +_start: +{ +lean_object* x_13; +x_13 = l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__6(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +x_13 = l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___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); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Lean_mkConstWithLevelParams___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +x_13 = l_Lean_mkConstWithLevelParams___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +x_13 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +lean_object* x_14; +x_14 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_14; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +_start: +{ +lean_object* x_16; +x_16 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_16; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +_start: +{ +lean_object* x_16; +x_16 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_16; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +lean_object* x_15; +x_15 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_15; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem_synthesizeInstance(lean_object* x_1, 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_box(0); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_9 = l_Lean_Meta_trySynthInstance(x_2, x_8, 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) == 1) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_11 = lean_ctor_get(x_9, 1); +lean_inc(x_11); +lean_dec(x_9); +x_12 = lean_ctor_get(x_10, 0); +lean_inc(x_12); +lean_dec(x_10); +x_13 = l_Lean_Meta_isExprDefEq(x_1, x_12, x_3, x_4, x_5, x_6, x_11); +return x_13; +} +else +{ +uint8_t x_14; +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_14 = !lean_is_exclusive(x_9); +if (x_14 == 0) +{ +lean_object* x_15; uint8_t x_16; lean_object* x_17; +x_15 = lean_ctor_get(x_9, 0); +lean_dec(x_15); +x_16 = 0; +x_17 = lean_box(x_16); +lean_ctor_set(x_9, 0, x_17); +return x_9; +} +else +{ +lean_object* x_18; uint8_t x_19; lean_object* x_20; lean_object* x_21; +x_18 = lean_ctor_get(x_9, 1); +lean_inc(x_18); +lean_dec(x_9); +x_19 = 0; +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; +} +} +} +else +{ +uint8_t x_22; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_22 = !lean_is_exclusive(x_9); +if (x_22 == 0) +{ +return x_9; +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_9, 0); +x_24 = lean_ctor_get(x_9, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_9); +x_25 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_24); +return x_25; +} +} +} +} +LEAN_EXPORT lean_object* l_ReaderT_read___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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) { +_start: +{ +lean_object* x_12; +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_1); +lean_ctor_set(x_12, 1, x_11); +return x_12; +} +} +static lean_object* _init_l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__2___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___spec__1___closed__4; +x_2 = l_ReaderT_instMonad___rarg(x_1); +return x_2; +} +} +static lean_object* _init_l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__2___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__2___closed__1; +x_2 = l_instInhabitedPUnit; +x_3 = l_instInhabitedOfMonad___rarg(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__2___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__2___closed__2; +x_2 = lean_alloc_closure((void*)(l_instInhabitedReaderT___rarg___boxed), 2, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_13 = l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__2___closed__3; +x_14 = lean_panic_fn(x_13, x_1); +x_15 = lean_apply_11(x_14, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_15; +} +} +static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___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; +} +} +static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___lambda__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___lambda__1___closed__1; +x_2 = lean_box(0); +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___lambda__1___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___lambda__1___closed__2; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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) { +_start: +{ +lean_object* x_13; lean_object* x_14; +x_13 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___lambda__1___closed__3; +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_12); +return x_14; +} +} +static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("issues", 6, 6); +return x_1; +} +} +static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__1; +x_2 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__1; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___lambda__1___boxed), 12, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("type error constructing proof for ", 34, 34); +return x_1; +} +} +static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__4; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("\nwhen assigning metavariable ", 29, 29); +return x_1; +} +} +static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__6; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__8() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(" with ", 6, 6); +return x_1; +} +} +static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__8; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19, lean_object* x_20, lean_object* x_21, lean_object* x_22, lean_object* x_23) { +_start: +{ +lean_object* x_24; uint8_t x_25; +x_24 = lean_ctor_get(x_8, 1); +x_25 = lean_nat_dec_lt(x_10, x_24); +if (x_25 == 0) +{ +lean_object* x_26; +lean_dec(x_22); +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_10); +lean_dec(x_7); +lean_dec(x_1); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_9); +lean_ctor_set(x_26, 1, x_23); +return x_26; +} +else +{ +lean_object* x_27; lean_object* x_28; lean_object* x_36; lean_object* x_165; lean_object* x_166; lean_object* x_167; uint8_t x_168; +lean_dec(x_9); +x_165 = lean_nat_sub(x_3, x_10); +x_166 = lean_unsigned_to_nat(1u); +x_167 = lean_nat_sub(x_165, x_166); +lean_dec(x_165); +x_168 = lean_nat_dec_lt(x_167, x_4); +if (x_168 == 0) +{ +lean_object* x_169; lean_object* x_170; +lean_dec(x_167); +x_169 = l_Lean_instInhabitedExpr; +x_170 = l_outOfBounds___rarg(x_169); +x_36 = x_170; +goto block_164; +} +else +{ +lean_object* x_171; +x_171 = lean_array_fget(x_2, x_167); +lean_dec(x_167); +x_36 = x_171; +goto block_164; +} +block_35: +{ +if (lean_obj_tag(x_27) == 0) +{ +lean_object* x_29; lean_object* x_30; +lean_dec(x_22); +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_10); +lean_dec(x_7); +lean_dec(x_1); +x_29 = lean_ctor_get(x_27, 0); +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; +} +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_27, 0); +lean_inc(x_31); +lean_dec(x_27); +x_32 = lean_ctor_get(x_8, 2); +x_33 = lean_nat_add(x_10, x_32); +lean_dec(x_10); +x_9 = x_31; +x_10 = x_33; +x_11 = lean_box(0); +x_12 = lean_box(0); +x_23 = x_28; +goto _start; +} +} +block_164: +{ +lean_object* x_37; uint8_t x_38; +x_37 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_unassigned; +x_38 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_36, x_37); +if (x_38 == 0) +{ +lean_object* x_39; uint8_t x_40; lean_object* x_41; lean_object* x_130; lean_object* x_131; +x_39 = lean_array_fget(x_5, x_10); +x_130 = l_Lean_Expr_mvarId_x21(x_39); +lean_inc(x_130); +x_131 = l_Lean_MVarId_getType(x_130, x_19, x_20, x_21, x_22, x_23); +if (lean_obj_tag(x_131) == 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_inc(x_22); +lean_inc(x_21); +lean_inc(x_20); +lean_inc(x_19); +lean_inc(x_36); +x_134 = lean_infer_type(x_36, x_19, x_20, x_21, x_22, x_133); +if (lean_obj_tag(x_134) == 0) +{ +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_22); +lean_inc(x_21); +lean_inc(x_20); +lean_inc(x_19); +x_137 = l_Lean_Meta_isExprDefEq(x_132, x_135, x_19, x_20, x_21, x_22, x_136); +if (lean_obj_tag(x_137) == 0) +{ +lean_object* x_138; uint8_t x_139; +x_138 = lean_ctor_get(x_137, 0); +lean_inc(x_138); +x_139 = lean_unbox(x_138); +if (x_139 == 0) +{ +lean_object* x_140; uint8_t x_141; +lean_dec(x_130); +x_140 = lean_ctor_get(x_137, 1); +lean_inc(x_140); +lean_dec(x_137); +x_141 = lean_unbox(x_138); +lean_dec(x_138); +x_40 = x_141; +x_41 = x_140; +goto block_129; +} +else +{ +lean_object* x_142; lean_object* x_143; +lean_dec(x_138); +x_142 = lean_ctor_get(x_137, 1); +lean_inc(x_142); +lean_dec(x_137); +lean_inc(x_22); +lean_inc(x_21); +lean_inc(x_20); +lean_inc(x_19); +lean_inc(x_36); +x_143 = lean_checked_assign(x_130, x_36, x_19, x_20, x_21, x_22, x_142); +if (lean_obj_tag(x_143) == 0) +{ +lean_object* x_144; lean_object* x_145; uint8_t x_146; +x_144 = lean_ctor_get(x_143, 0); +lean_inc(x_144); +x_145 = lean_ctor_get(x_143, 1); +lean_inc(x_145); +lean_dec(x_143); +x_146 = lean_unbox(x_144); +lean_dec(x_144); +x_40 = x_146; +x_41 = x_145; +goto block_129; +} +else +{ +uint8_t x_147; +lean_dec(x_39); +lean_dec(x_36); +lean_dec(x_22); +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_10); +lean_dec(x_7); +lean_dec(x_1); +x_147 = !lean_is_exclusive(x_143); +if (x_147 == 0) +{ +return x_143; +} +else +{ +lean_object* x_148; lean_object* x_149; lean_object* x_150; +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); +x_150 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_150, 0, x_148); +lean_ctor_set(x_150, 1, x_149); +return x_150; +} +} +} +} +else +{ +uint8_t x_151; +lean_dec(x_130); +lean_dec(x_39); +lean_dec(x_36); +lean_dec(x_22); +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_10); +lean_dec(x_7); +lean_dec(x_1); +x_151 = !lean_is_exclusive(x_137); +if (x_151 == 0) +{ +return x_137; +} +else +{ +lean_object* x_152; lean_object* x_153; lean_object* x_154; +x_152 = lean_ctor_get(x_137, 0); +x_153 = lean_ctor_get(x_137, 1); +lean_inc(x_153); +lean_inc(x_152); +lean_dec(x_137); +x_154 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_154, 0, x_152); +lean_ctor_set(x_154, 1, x_153); +return x_154; +} +} +} +else +{ +uint8_t x_155; +lean_dec(x_132); +lean_dec(x_130); +lean_dec(x_39); +lean_dec(x_36); +lean_dec(x_22); +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_10); +lean_dec(x_7); +lean_dec(x_1); +x_155 = !lean_is_exclusive(x_134); +if (x_155 == 0) +{ +return x_134; +} +else +{ +lean_object* x_156; lean_object* x_157; lean_object* x_158; +x_156 = lean_ctor_get(x_134, 0); +x_157 = lean_ctor_get(x_134, 1); +lean_inc(x_157); +lean_inc(x_156); +lean_dec(x_134); +x_158 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_158, 0, x_156); +lean_ctor_set(x_158, 1, x_157); +return x_158; +} +} +} +else +{ +uint8_t x_159; +lean_dec(x_130); +lean_dec(x_39); +lean_dec(x_36); +lean_dec(x_22); +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_10); +lean_dec(x_7); +lean_dec(x_1); +x_159 = !lean_is_exclusive(x_131); +if (x_159 == 0) +{ +return x_131; +} +else +{ +lean_object* x_160; lean_object* x_161; lean_object* x_162; +x_160 = lean_ctor_get(x_131, 0); +x_161 = lean_ctor_get(x_131, 1); +lean_inc(x_161); +lean_inc(x_160); +lean_dec(x_131); +x_162 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_162, 0, x_160); +lean_ctor_set(x_162, 1, x_161); +return x_162; +} +} +block_129: +{ +if (x_40 == 0) +{ +lean_object* x_42; lean_object* x_43; uint8_t x_44; +x_42 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__2; +x_43 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__2(x_42, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_41); +x_44 = !lean_is_exclusive(x_43); +if (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_43, 0); +x_46 = lean_ctor_get(x_43, 1); +x_47 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__3; +x_48 = lean_unbox(x_45); +lean_dec(x_45); +if (x_48 == 0) +{ +lean_object* x_49; lean_object* x_50; +lean_free_object(x_43); +lean_dec(x_39); +lean_dec(x_36); +x_49 = lean_box(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_inc(x_14); +lean_inc(x_13); +x_50 = lean_apply_12(x_47, x_49, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_46); +if (lean_obj_tag(x_50) == 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_inc(x_52); +lean_dec(x_50); +x_27 = x_51; +x_28 = x_52; +goto block_35; +} +else +{ +uint8_t x_53; +lean_dec(x_22); +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_10); +lean_dec(x_7); +lean_dec(x_1); +x_53 = !lean_is_exclusive(x_50); +if (x_53 == 0) +{ +return x_50; +} +else +{ +lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_54 = lean_ctor_get(x_50, 0); +x_55 = lean_ctor_get(x_50, 1); +lean_inc(x_55); +lean_inc(x_54); +lean_dec(x_50); +x_56 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_56, 0, x_54); +lean_ctor_set(x_56, 1, x_55); +return x_56; +} +} +} +else +{ +lean_object* x_57; lean_object* x_58; +x_57 = lean_ctor_get(x_1, 5); +lean_inc(x_57); +x_58 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_57, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_46); +if (lean_obj_tag(x_58) == 0) +{ +lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; 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_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_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__5; +lean_ctor_set_tag(x_43, 7); +lean_ctor_set(x_43, 1, x_59); +lean_ctor_set(x_43, 0, x_61); +x_62 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__7; +x_63 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_63, 0, x_43); +lean_ctor_set(x_63, 1, x_62); +x_64 = l_Lean_MessageData_ofExpr(x_39); +x_65 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_65, 0, x_63); +lean_ctor_set(x_65, 1, x_64); +x_66 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__9; +x_67 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_67, 0, x_65); +lean_ctor_set(x_67, 1, x_66); +x_68 = l_Lean_indentExpr(x_36); +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 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; +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_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_42, x_71, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_60); +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); +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_inc(x_13); +x_75 = lean_apply_12(x_47, x_73, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_74); +if (lean_obj_tag(x_75) == 0) +{ +lean_object* x_76; lean_object* x_77; +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_27 = x_76; +x_28 = x_77; +goto block_35; +} +else +{ +uint8_t x_78; +lean_dec(x_22); +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_10); +lean_dec(x_7); +lean_dec(x_1); +x_78 = !lean_is_exclusive(x_75); +if (x_78 == 0) +{ +return x_75; +} +else +{ +lean_object* x_79; lean_object* x_80; lean_object* x_81; +x_79 = lean_ctor_get(x_75, 0); +x_80 = lean_ctor_get(x_75, 1); +lean_inc(x_80); +lean_inc(x_79); +lean_dec(x_75); +x_81 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_81, 0, x_79); +lean_ctor_set(x_81, 1, x_80); +return x_81; +} +} +} +else +{ +uint8_t x_82; +lean_free_object(x_43); +lean_dec(x_39); +lean_dec(x_36); +lean_dec(x_22); +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_10); +lean_dec(x_7); +lean_dec(x_1); +x_82 = !lean_is_exclusive(x_58); +if (x_82 == 0) +{ +return x_58; +} +else +{ +lean_object* x_83; lean_object* x_84; lean_object* x_85; +x_83 = lean_ctor_get(x_58, 0); +x_84 = lean_ctor_get(x_58, 1); +lean_inc(x_84); +lean_inc(x_83); +lean_dec(x_58); +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_86; lean_object* x_87; lean_object* x_88; uint8_t x_89; +x_86 = lean_ctor_get(x_43, 0); +x_87 = lean_ctor_get(x_43, 1); +lean_inc(x_87); +lean_inc(x_86); +lean_dec(x_43); +x_88 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__3; +x_89 = lean_unbox(x_86); +lean_dec(x_86); +if (x_89 == 0) +{ +lean_object* x_90; lean_object* x_91; +lean_dec(x_39); +lean_dec(x_36); +x_90 = lean_box(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_inc(x_14); +lean_inc(x_13); +x_91 = lean_apply_12(x_88, x_90, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_87); +if (lean_obj_tag(x_91) == 0) +{ +lean_object* x_92; lean_object* x_93; +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_27 = x_92; +x_28 = x_93; +goto block_35; +} +else +{ +lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; +lean_dec(x_22); +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_10); +lean_dec(x_7); +lean_dec(x_1); +x_94 = lean_ctor_get(x_91, 0); +lean_inc(x_94); +x_95 = lean_ctor_get(x_91, 1); +lean_inc(x_95); +if (lean_is_exclusive(x_91)) { + lean_ctor_release(x_91, 0); + lean_ctor_release(x_91, 1); + x_96 = x_91; +} else { + lean_dec_ref(x_91); + x_96 = lean_box(0); +} +if (lean_is_scalar(x_96)) { + x_97 = lean_alloc_ctor(1, 2, 0); +} else { + x_97 = x_96; +} +lean_ctor_set(x_97, 0, x_94); +lean_ctor_set(x_97, 1, x_95); +return x_97; +} +} +else +{ +lean_object* x_98; lean_object* x_99; +x_98 = lean_ctor_get(x_1, 5); +lean_inc(x_98); +x_99 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_98, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_87); +if (lean_obj_tag(x_99) == 0) +{ +lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; +x_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_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__5; +x_103 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_103, 0, x_102); +lean_ctor_set(x_103, 1, x_100); +x_104 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__7; +x_105 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_105, 0, x_103); +lean_ctor_set(x_105, 1, x_104); +x_106 = l_Lean_MessageData_ofExpr(x_39); +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_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__9; +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_indentExpr(x_36); +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 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; +x_113 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_113, 0, x_111); +lean_ctor_set(x_113, 1, x_112); +x_114 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_42, x_113, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_101); +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); +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_inc(x_13); +x_117 = lean_apply_12(x_88, x_115, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_116); +if (lean_obj_tag(x_117) == 0) +{ +lean_object* x_118; lean_object* x_119; +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_27 = x_118; +x_28 = x_119; +goto block_35; +} +else +{ +lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; +lean_dec(x_22); +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_10); +lean_dec(x_7); +lean_dec(x_1); +x_120 = lean_ctor_get(x_117, 0); +lean_inc(x_120); +x_121 = lean_ctor_get(x_117, 1); +lean_inc(x_121); +if (lean_is_exclusive(x_117)) { + lean_ctor_release(x_117, 0); + lean_ctor_release(x_117, 1); + x_122 = x_117; +} else { + lean_dec_ref(x_117); + x_122 = lean_box(0); +} +if (lean_is_scalar(x_122)) { + x_123 = lean_alloc_ctor(1, 2, 0); +} else { + x_123 = 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_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; +lean_dec(x_39); +lean_dec(x_36); +lean_dec(x_22); +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_10); +lean_dec(x_7); +lean_dec(x_1); +x_124 = lean_ctor_get(x_99, 0); +lean_inc(x_124); +x_125 = lean_ctor_get(x_99, 1); +lean_inc(x_125); +if (lean_is_exclusive(x_99)) { + lean_ctor_release(x_99, 0); + lean_ctor_release(x_99, 1); + x_126 = x_99; +} else { + lean_dec_ref(x_99); + x_126 = lean_box(0); +} +if (lean_is_scalar(x_126)) { + x_127 = lean_alloc_ctor(1, 2, 0); +} else { + x_127 = x_126; +} +lean_ctor_set(x_127, 0, x_124); +lean_ctor_set(x_127, 1, x_125); +return x_127; +} +} +} +} +else +{ +lean_object* x_128; +lean_dec(x_39); +lean_dec(x_36); +lean_inc(x_7); +x_128 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_128, 0, x_7); +x_27 = x_128; +x_28 = x_41; +goto block_35; +} +} +} +else +{ +lean_object* x_163; +lean_dec(x_36); +lean_inc(x_7); +x_163 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_163, 0, x_7); +x_27 = x_163; +x_28 = x_23; +goto block_35; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_MVarId_isAssigned___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; uint8_t x_14; +x_13 = lean_st_ref_get(x_9, x_12); +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; lean_object* x_19; +x_15 = lean_ctor_get(x_13, 0); +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +lean_dec(x_15); +x_17 = lean_ctor_get(x_16, 7); +lean_inc(x_17); +lean_dec(x_16); +x_18 = l_Lean_PersistentHashMap_contains___at_Lean_MVarId_isAssigned___spec__1(x_17, x_1); +x_19 = lean_box(x_18); +lean_ctor_set(x_13, 0, x_19); +return x_13; +} +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; lean_object* x_25; lean_object* x_26; +x_20 = lean_ctor_get(x_13, 0); +x_21 = lean_ctor_get(x_13, 1); +lean_inc(x_21); +lean_inc(x_20); +lean_dec(x_13); +x_22 = lean_ctor_get(x_20, 0); +lean_inc(x_22); +lean_dec(x_20); +x_23 = lean_ctor_get(x_22, 7); +lean_inc(x_23); +lean_dec(x_22); +x_24 = l_Lean_PersistentHashMap_contains___at_Lean_MVarId_isAssigned___spec__1(x_23, x_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_21); +return x_26; +} +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_14 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___lambda__1___closed__1; +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_14); +lean_ctor_set(x_15, 1, x_1); +x_16 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_16, 0, 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; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("failed to synthesize instance when instantiating ", 49, 49); +return x_1; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, size_t x_6, size_t x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, 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; +x_20 = lean_usize_dec_lt(x_7, x_6); +if (x_20 == 0) +{ +lean_object* x_21; +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_3); +lean_dec(x_1); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_8); +lean_ctor_set(x_21, 1, x_19); +return x_21; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_32; +x_22 = lean_array_uget(x_5, x_7); +x_32 = !lean_is_exclusive(x_8); +if (x_32 == 0) +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; uint8_t x_38; +x_33 = lean_ctor_get(x_8, 1); +x_34 = lean_ctor_get(x_8, 0); +lean_dec(x_34); +x_35 = lean_ctor_get(x_33, 0); +lean_inc(x_35); +x_36 = lean_ctor_get(x_33, 1); +lean_inc(x_36); +x_37 = lean_ctor_get(x_33, 2); +lean_inc(x_37); +x_38 = lean_nat_dec_lt(x_36, x_37); +if (x_38 == 0) +{ +lean_object* x_39; +lean_dec(x_37); +lean_dec(x_36); +lean_dec(x_35); +lean_dec(x_22); +lean_inc(x_3); +lean_ctor_set(x_8, 0, x_3); +x_39 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_39, 0, x_8); +x_23 = x_39; +x_24 = x_19; +goto block_31; +} +else +{ +uint8_t x_40; +x_40 = !lean_is_exclusive(x_33); +if (x_40 == 0) +{ +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; uint8_t x_50; +x_41 = lean_ctor_get(x_33, 2); +lean_dec(x_41); +x_42 = lean_ctor_get(x_33, 1); +lean_dec(x_42); +x_43 = lean_ctor_get(x_33, 0); +lean_dec(x_43); +x_44 = lean_array_fget(x_35, x_36); +x_45 = lean_unbox(x_44); +lean_dec(x_44); +x_46 = lean_unsigned_to_nat(1u); +x_47 = lean_nat_add(x_36, x_46); +lean_dec(x_36); +lean_ctor_set(x_33, 1, x_47); +x_48 = l_Lean_Expr_mvarId_x21(x_22); +x_49 = l_Lean_MVarId_isAssigned___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__4(x_48, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19); +lean_dec(x_48); +x_50 = !lean_is_exclusive(x_49); +if (x_50 == 0) +{ +lean_object* x_51; lean_object* x_52; uint8_t x_53; +x_51 = lean_ctor_get(x_49, 0); +x_52 = lean_ctor_get(x_49, 1); +x_53 = l_Lean_BinderInfo_isInstImplicit(x_45); +if (x_53 == 0) +{ +lean_object* x_54; +lean_free_object(x_49); +lean_dec(x_51); +lean_dec(x_22); +lean_inc(x_3); +lean_ctor_set(x_8, 0, x_3); +x_54 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_54, 0, x_8); +x_23 = x_54; +x_24 = x_52; +goto block_31; +} +else +{ +uint8_t x_55; +x_55 = lean_unbox(x_51); +lean_dec(x_51); +if (x_55 == 0) +{ +lean_object* x_56; +lean_inc(x_18); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_22); +x_56 = lean_infer_type(x_22, x_15, x_16, x_17, x_18, x_52); +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_56, 1); +lean_inc(x_58); +lean_dec(x_56); +lean_inc(x_18); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_57); +x_59 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem_synthesizeInstance(x_22, x_57, x_15, x_16, x_17, x_18, 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; lean_object* x_64; lean_object* x_65; uint8_t x_66; +lean_free_object(x_8); +x_62 = lean_ctor_get(x_59, 1); +lean_inc(x_62); +lean_dec(x_59); +x_63 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__2; +x_64 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__2(x_63, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_62); +x_65 = lean_ctor_get(x_64, 0); +lean_inc(x_65); +x_66 = lean_unbox(x_65); +lean_dec(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_dec(x_57); +lean_free_object(x_49); +x_67 = lean_ctor_get(x_64, 1); +lean_inc(x_67); +lean_dec(x_64); +x_68 = lean_box(0); +x_69 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___lambda__1(x_33, x_68, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_67); +x_70 = lean_ctor_get(x_69, 0); +lean_inc(x_70); +x_71 = lean_ctor_get(x_69, 1); +lean_inc(x_71); +lean_dec(x_69); +x_23 = x_70; +x_24 = x_71; +goto block_31; +} +else +{ +uint8_t x_72; +x_72 = !lean_is_exclusive(x_64); +if (x_72 == 0) +{ +lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; +x_73 = lean_ctor_get(x_64, 1); +x_74 = lean_ctor_get(x_64, 0); +lean_dec(x_74); +x_75 = lean_ctor_get(x_1, 5); +lean_inc(x_75); +x_76 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_75, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_73); +if (lean_obj_tag(x_76) == 0) +{ +lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; +x_77 = lean_ctor_get(x_76, 0); +lean_inc(x_77); +x_78 = lean_ctor_get(x_76, 1); +lean_inc(x_78); +lean_dec(x_76); +x_79 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___closed__2; +lean_ctor_set_tag(x_64, 7); +lean_ctor_set(x_64, 1, x_77); +lean_ctor_set(x_64, 0, x_79); +x_80 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; +lean_ctor_set_tag(x_49, 7); +lean_ctor_set(x_49, 1, x_80); +lean_ctor_set(x_49, 0, x_64); +x_81 = l_Lean_indentExpr(x_57); +x_82 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_82, 0, x_49); +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_80); +x_84 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_63, x_83, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_78); +x_85 = lean_ctor_get(x_84, 0); +lean_inc(x_85); +x_86 = lean_ctor_get(x_84, 1); +lean_inc(x_86); +lean_dec(x_84); +x_87 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___lambda__1(x_33, x_85, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_86); +lean_dec(x_85); +x_88 = lean_ctor_get(x_87, 0); +lean_inc(x_88); +x_89 = lean_ctor_get(x_87, 1); +lean_inc(x_89); +lean_dec(x_87); +x_23 = x_88; +x_24 = x_89; +goto block_31; +} +else +{ +uint8_t x_90; +lean_free_object(x_64); +lean_dec(x_57); +lean_free_object(x_49); +lean_dec(x_33); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_3); +lean_dec(x_1); +x_90 = !lean_is_exclusive(x_76); +if (x_90 == 0) +{ +return x_76; +} +else +{ +lean_object* x_91; lean_object* x_92; lean_object* x_93; +x_91 = lean_ctor_get(x_76, 0); +x_92 = lean_ctor_get(x_76, 1); +lean_inc(x_92); +lean_inc(x_91); +lean_dec(x_76); +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_94; lean_object* x_95; lean_object* x_96; +x_94 = lean_ctor_get(x_64, 1); +lean_inc(x_94); +lean_dec(x_64); +x_95 = lean_ctor_get(x_1, 5); +lean_inc(x_95); +x_96 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_95, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_94); +if (lean_obj_tag(x_96) == 0) +{ +lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; 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_97 = lean_ctor_get(x_96, 0); +lean_inc(x_97); +x_98 = lean_ctor_get(x_96, 1); +lean_inc(x_98); +lean_dec(x_96); +x_99 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___closed__2; +x_100 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_100, 0, x_99); +lean_ctor_set(x_100, 1, x_97); +x_101 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; +lean_ctor_set_tag(x_49, 7); +lean_ctor_set(x_49, 1, x_101); +lean_ctor_set(x_49, 0, x_100); +x_102 = l_Lean_indentExpr(x_57); +x_103 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_103, 0, x_49); +lean_ctor_set(x_103, 1, x_102); +x_104 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_104, 0, x_103); +lean_ctor_set(x_104, 1, x_101); +x_105 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_63, x_104, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_98); +x_106 = lean_ctor_get(x_105, 0); +lean_inc(x_106); +x_107 = lean_ctor_get(x_105, 1); +lean_inc(x_107); +lean_dec(x_105); +x_108 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___lambda__1(x_33, x_106, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_107); +lean_dec(x_106); +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_23 = x_109; +x_24 = x_110; +goto block_31; +} +else +{ +lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; +lean_dec(x_57); +lean_free_object(x_49); +lean_dec(x_33); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_3); +lean_dec(x_1); +x_111 = lean_ctor_get(x_96, 0); +lean_inc(x_111); +x_112 = lean_ctor_get(x_96, 1); +lean_inc(x_112); +if (lean_is_exclusive(x_96)) { + lean_ctor_release(x_96, 0); + lean_ctor_release(x_96, 1); + x_113 = x_96; +} else { + lean_dec_ref(x_96); + x_113 = lean_box(0); +} +if (lean_is_scalar(x_113)) { + x_114 = lean_alloc_ctor(1, 2, 0); +} else { + x_114 = x_113; +} +lean_ctor_set(x_114, 0, x_111); +lean_ctor_set(x_114, 1, x_112); +return x_114; +} +} +} +} +else +{ +lean_object* x_115; lean_object* x_116; +lean_dec(x_57); +lean_free_object(x_49); +x_115 = lean_ctor_get(x_59, 1); +lean_inc(x_115); +lean_dec(x_59); +lean_inc(x_3); +lean_ctor_set(x_8, 0, x_3); +x_116 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_116, 0, x_8); +x_23 = x_116; +x_24 = x_115; +goto block_31; +} +} +else +{ +uint8_t x_117; +lean_dec(x_57); +lean_free_object(x_49); +lean_dec(x_33); +lean_free_object(x_8); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_3); +lean_dec(x_1); +x_117 = !lean_is_exclusive(x_59); +if (x_117 == 0) +{ +return x_59; +} +else +{ +lean_object* x_118; lean_object* x_119; lean_object* x_120; +x_118 = lean_ctor_get(x_59, 0); +x_119 = lean_ctor_get(x_59, 1); +lean_inc(x_119); +lean_inc(x_118); +lean_dec(x_59); +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_121; +lean_free_object(x_49); +lean_dec(x_33); +lean_free_object(x_8); +lean_dec(x_22); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_3); +lean_dec(x_1); +x_121 = !lean_is_exclusive(x_56); +if (x_121 == 0) +{ +return x_56; +} +else +{ +lean_object* x_122; lean_object* x_123; lean_object* x_124; +x_122 = lean_ctor_get(x_56, 0); +x_123 = lean_ctor_get(x_56, 1); +lean_inc(x_123); +lean_inc(x_122); +lean_dec(x_56); +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_125; +lean_free_object(x_49); +lean_dec(x_22); +lean_inc(x_3); +lean_ctor_set(x_8, 0, x_3); +x_125 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_125, 0, x_8); +x_23 = x_125; +x_24 = x_52; +goto block_31; +} +} +} +else +{ +lean_object* x_126; lean_object* x_127; uint8_t x_128; +x_126 = lean_ctor_get(x_49, 0); +x_127 = lean_ctor_get(x_49, 1); +lean_inc(x_127); +lean_inc(x_126); +lean_dec(x_49); +x_128 = l_Lean_BinderInfo_isInstImplicit(x_45); +if (x_128 == 0) +{ +lean_object* x_129; +lean_dec(x_126); +lean_dec(x_22); +lean_inc(x_3); +lean_ctor_set(x_8, 0, x_3); +x_129 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_129, 0, x_8); +x_23 = x_129; +x_24 = x_127; +goto block_31; +} +else +{ +uint8_t x_130; +x_130 = lean_unbox(x_126); +lean_dec(x_126); +if (x_130 == 0) +{ +lean_object* x_131; +lean_inc(x_18); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_22); +x_131 = lean_infer_type(x_22, x_15, x_16, x_17, x_18, x_127); +if (lean_obj_tag(x_131) == 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_inc(x_18); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_132); +x_134 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem_synthesizeInstance(x_22, x_132, x_15, x_16, x_17, x_18, x_133); +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) +{ +lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; uint8_t x_141; +lean_free_object(x_8); +x_137 = lean_ctor_get(x_134, 1); +lean_inc(x_137); +lean_dec(x_134); +x_138 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__2; +x_139 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__2(x_138, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_137); +x_140 = lean_ctor_get(x_139, 0); +lean_inc(x_140); +x_141 = lean_unbox(x_140); +lean_dec(x_140); +if (x_141 == 0) +{ +lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; +lean_dec(x_132); +x_142 = lean_ctor_get(x_139, 1); +lean_inc(x_142); +lean_dec(x_139); +x_143 = lean_box(0); +x_144 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___lambda__1(x_33, x_143, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_142); +x_145 = lean_ctor_get(x_144, 0); +lean_inc(x_145); +x_146 = lean_ctor_get(x_144, 1); +lean_inc(x_146); +lean_dec(x_144); +x_23 = x_145; +x_24 = x_146; +goto block_31; +} +else +{ +lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; +x_147 = lean_ctor_get(x_139, 1); +lean_inc(x_147); +if (lean_is_exclusive(x_139)) { + lean_ctor_release(x_139, 0); + lean_ctor_release(x_139, 1); + x_148 = x_139; +} else { + lean_dec_ref(x_139); + x_148 = lean_box(0); +} +x_149 = lean_ctor_get(x_1, 5); +lean_inc(x_149); +x_150 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_149, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_147); +if (lean_obj_tag(x_150) == 0) +{ +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; +x_151 = lean_ctor_get(x_150, 0); +lean_inc(x_151); +x_152 = lean_ctor_get(x_150, 1); +lean_inc(x_152); +lean_dec(x_150); +x_153 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___closed__2; +if (lean_is_scalar(x_148)) { + x_154 = lean_alloc_ctor(7, 2, 0); +} else { + x_154 = x_148; + lean_ctor_set_tag(x_154, 7); +} +lean_ctor_set(x_154, 0, x_153); +lean_ctor_set(x_154, 1, x_151); +x_155 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; +x_156 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_156, 0, x_154); +lean_ctor_set(x_156, 1, x_155); +x_157 = l_Lean_indentExpr(x_132); +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 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_159, 0, x_158); +lean_ctor_set(x_159, 1, x_155); +x_160 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_138, x_159, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_152); +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_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___lambda__1(x_33, x_161, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_162); +lean_dec(x_161); +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_23 = x_164; +x_24 = x_165; +goto block_31; +} +else +{ +lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; +lean_dec(x_148); +lean_dec(x_132); +lean_dec(x_33); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_3); +lean_dec(x_1); +x_166 = lean_ctor_get(x_150, 0); +lean_inc(x_166); +x_167 = lean_ctor_get(x_150, 1); +lean_inc(x_167); +if (lean_is_exclusive(x_150)) { + lean_ctor_release(x_150, 0); + lean_ctor_release(x_150, 1); + x_168 = x_150; +} else { + lean_dec_ref(x_150); + x_168 = lean_box(0); +} +if (lean_is_scalar(x_168)) { + x_169 = lean_alloc_ctor(1, 2, 0); +} else { + x_169 = x_168; +} +lean_ctor_set(x_169, 0, x_166); +lean_ctor_set(x_169, 1, x_167); +return x_169; +} +} +} +else +{ +lean_object* x_170; lean_object* x_171; +lean_dec(x_132); +x_170 = lean_ctor_get(x_134, 1); +lean_inc(x_170); +lean_dec(x_134); +lean_inc(x_3); +lean_ctor_set(x_8, 0, x_3); +x_171 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_171, 0, x_8); +x_23 = x_171; +x_24 = x_170; +goto block_31; +} +} +else +{ +lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; +lean_dec(x_132); +lean_dec(x_33); +lean_free_object(x_8); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_3); +lean_dec(x_1); +x_172 = lean_ctor_get(x_134, 0); +lean_inc(x_172); +x_173 = lean_ctor_get(x_134, 1); +lean_inc(x_173); +if (lean_is_exclusive(x_134)) { + lean_ctor_release(x_134, 0); + lean_ctor_release(x_134, 1); + x_174 = x_134; +} else { + lean_dec_ref(x_134); + x_174 = lean_box(0); +} +if (lean_is_scalar(x_174)) { + x_175 = lean_alloc_ctor(1, 2, 0); +} else { + x_175 = x_174; +} +lean_ctor_set(x_175, 0, x_172); +lean_ctor_set(x_175, 1, x_173); +return x_175; +} +} +else +{ +lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; +lean_dec(x_33); +lean_free_object(x_8); +lean_dec(x_22); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_3); +lean_dec(x_1); +x_176 = lean_ctor_get(x_131, 0); +lean_inc(x_176); +x_177 = lean_ctor_get(x_131, 1); +lean_inc(x_177); +if (lean_is_exclusive(x_131)) { + lean_ctor_release(x_131, 0); + lean_ctor_release(x_131, 1); + x_178 = x_131; +} else { + lean_dec_ref(x_131); + x_178 = lean_box(0); +} +if (lean_is_scalar(x_178)) { + x_179 = lean_alloc_ctor(1, 2, 0); +} else { + x_179 = x_178; +} +lean_ctor_set(x_179, 0, x_176); +lean_ctor_set(x_179, 1, x_177); +return x_179; +} +} +else +{ +lean_object* x_180; +lean_dec(x_22); +lean_inc(x_3); +lean_ctor_set(x_8, 0, x_3); +x_180 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_180, 0, x_8); +x_23 = x_180; +x_24 = x_127; +goto block_31; +} +} +} +} +else +{ +lean_object* x_181; uint8_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; lean_object* x_190; uint8_t x_191; +lean_dec(x_33); +x_181 = lean_array_fget(x_35, x_36); +x_182 = lean_unbox(x_181); +lean_dec(x_181); +x_183 = lean_unsigned_to_nat(1u); +x_184 = lean_nat_add(x_36, x_183); +lean_dec(x_36); +x_185 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_185, 0, x_35); +lean_ctor_set(x_185, 1, x_184); +lean_ctor_set(x_185, 2, x_37); +x_186 = l_Lean_Expr_mvarId_x21(x_22); +x_187 = l_Lean_MVarId_isAssigned___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__4(x_186, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19); +lean_dec(x_186); +x_188 = lean_ctor_get(x_187, 0); +lean_inc(x_188); +x_189 = lean_ctor_get(x_187, 1); +lean_inc(x_189); +if (lean_is_exclusive(x_187)) { + lean_ctor_release(x_187, 0); + lean_ctor_release(x_187, 1); + x_190 = x_187; +} else { + lean_dec_ref(x_187); + x_190 = lean_box(0); +} +x_191 = l_Lean_BinderInfo_isInstImplicit(x_182); +if (x_191 == 0) +{ +lean_object* x_192; +lean_dec(x_190); +lean_dec(x_188); +lean_dec(x_22); +lean_inc(x_3); +lean_ctor_set(x_8, 1, x_185); +lean_ctor_set(x_8, 0, x_3); +x_192 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_192, 0, x_8); +x_23 = x_192; +x_24 = x_189; +goto block_31; +} +else +{ +uint8_t x_193; +x_193 = lean_unbox(x_188); +lean_dec(x_188); +if (x_193 == 0) +{ +lean_object* x_194; +lean_inc(x_18); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_22); +x_194 = lean_infer_type(x_22, x_15, x_16, x_17, x_18, x_189); +if (lean_obj_tag(x_194) == 0) +{ +lean_object* x_195; lean_object* x_196; lean_object* x_197; +x_195 = lean_ctor_get(x_194, 0); +lean_inc(x_195); +x_196 = lean_ctor_get(x_194, 1); +lean_inc(x_196); +lean_dec(x_194); +lean_inc(x_18); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_195); +x_197 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem_synthesizeInstance(x_22, x_195, x_15, x_16, x_17, x_18, x_196); +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_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; uint8_t x_204; +lean_free_object(x_8); +x_200 = lean_ctor_get(x_197, 1); +lean_inc(x_200); +lean_dec(x_197); +x_201 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__2; +x_202 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__2(x_201, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_200); +x_203 = lean_ctor_get(x_202, 0); +lean_inc(x_203); +x_204 = lean_unbox(x_203); +lean_dec(x_203); +if (x_204 == 0) +{ +lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; +lean_dec(x_195); +lean_dec(x_190); +x_205 = lean_ctor_get(x_202, 1); +lean_inc(x_205); +lean_dec(x_202); +x_206 = lean_box(0); +x_207 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___lambda__1(x_185, x_206, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_205); +x_208 = lean_ctor_get(x_207, 0); +lean_inc(x_208); +x_209 = lean_ctor_get(x_207, 1); +lean_inc(x_209); +lean_dec(x_207); +x_23 = x_208; +x_24 = x_209; +goto block_31; +} +else +{ +lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; +x_210 = lean_ctor_get(x_202, 1); +lean_inc(x_210); +if (lean_is_exclusive(x_202)) { + lean_ctor_release(x_202, 0); + lean_ctor_release(x_202, 1); + x_211 = x_202; +} else { + lean_dec_ref(x_202); + x_211 = lean_box(0); +} +x_212 = lean_ctor_get(x_1, 5); +lean_inc(x_212); +x_213 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_212, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_210); +if (lean_obj_tag(x_213) == 0) +{ +lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; +x_214 = lean_ctor_get(x_213, 0); +lean_inc(x_214); +x_215 = lean_ctor_get(x_213, 1); +lean_inc(x_215); +lean_dec(x_213); +x_216 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___closed__2; +if (lean_is_scalar(x_211)) { + x_217 = lean_alloc_ctor(7, 2, 0); +} else { + x_217 = x_211; + lean_ctor_set_tag(x_217, 7); +} +lean_ctor_set(x_217, 0, x_216); +lean_ctor_set(x_217, 1, x_214); +x_218 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; +if (lean_is_scalar(x_190)) { + x_219 = lean_alloc_ctor(7, 2, 0); +} else { + x_219 = x_190; + lean_ctor_set_tag(x_219, 7); +} +lean_ctor_set(x_219, 0, x_217); +lean_ctor_set(x_219, 1, x_218); +x_220 = l_Lean_indentExpr(x_195); +x_221 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_221, 0, x_219); +lean_ctor_set(x_221, 1, x_220); +x_222 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_222, 0, x_221); +lean_ctor_set(x_222, 1, x_218); +x_223 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_201, x_222, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_215); +x_224 = lean_ctor_get(x_223, 0); +lean_inc(x_224); +x_225 = lean_ctor_get(x_223, 1); +lean_inc(x_225); +lean_dec(x_223); +x_226 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___lambda__1(x_185, x_224, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_225); +lean_dec(x_224); +x_227 = lean_ctor_get(x_226, 0); +lean_inc(x_227); +x_228 = lean_ctor_get(x_226, 1); +lean_inc(x_228); +lean_dec(x_226); +x_23 = x_227; +x_24 = x_228; +goto block_31; +} +else +{ +lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; +lean_dec(x_211); +lean_dec(x_195); +lean_dec(x_190); +lean_dec(x_185); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_3); +lean_dec(x_1); +x_229 = lean_ctor_get(x_213, 0); +lean_inc(x_229); +x_230 = lean_ctor_get(x_213, 1); +lean_inc(x_230); +if (lean_is_exclusive(x_213)) { + lean_ctor_release(x_213, 0); + lean_ctor_release(x_213, 1); + x_231 = x_213; +} else { + lean_dec_ref(x_213); + x_231 = lean_box(0); +} +if (lean_is_scalar(x_231)) { + x_232 = lean_alloc_ctor(1, 2, 0); +} else { + x_232 = x_231; +} +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_195); +lean_dec(x_190); +x_233 = lean_ctor_get(x_197, 1); +lean_inc(x_233); +lean_dec(x_197); +lean_inc(x_3); +lean_ctor_set(x_8, 1, x_185); +lean_ctor_set(x_8, 0, x_3); +x_234 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_234, 0, x_8); +x_23 = x_234; +x_24 = x_233; +goto block_31; +} +} +else +{ +lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; +lean_dec(x_195); +lean_dec(x_190); +lean_dec(x_185); +lean_free_object(x_8); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_3); +lean_dec(x_1); +x_235 = lean_ctor_get(x_197, 0); +lean_inc(x_235); +x_236 = lean_ctor_get(x_197, 1); +lean_inc(x_236); +if (lean_is_exclusive(x_197)) { + lean_ctor_release(x_197, 0); + lean_ctor_release(x_197, 1); + x_237 = x_197; +} else { + lean_dec_ref(x_197); + x_237 = lean_box(0); +} +if (lean_is_scalar(x_237)) { + x_238 = lean_alloc_ctor(1, 2, 0); +} else { + x_238 = x_237; +} +lean_ctor_set(x_238, 0, x_235); +lean_ctor_set(x_238, 1, x_236); +return x_238; +} +} +else +{ +lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; +lean_dec(x_190); +lean_dec(x_185); +lean_free_object(x_8); +lean_dec(x_22); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_3); +lean_dec(x_1); +x_239 = lean_ctor_get(x_194, 0); +lean_inc(x_239); +x_240 = lean_ctor_get(x_194, 1); +lean_inc(x_240); +if (lean_is_exclusive(x_194)) { + lean_ctor_release(x_194, 0); + lean_ctor_release(x_194, 1); + x_241 = x_194; +} else { + lean_dec_ref(x_194); + x_241 = lean_box(0); +} +if (lean_is_scalar(x_241)) { + x_242 = lean_alloc_ctor(1, 2, 0); +} else { + x_242 = x_241; +} +lean_ctor_set(x_242, 0, x_239); +lean_ctor_set(x_242, 1, x_240); +return x_242; +} +} +else +{ +lean_object* x_243; +lean_dec(x_190); +lean_dec(x_22); +lean_inc(x_3); +lean_ctor_set(x_8, 1, x_185); +lean_ctor_set(x_8, 0, x_3); +x_243 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_243, 0, x_8); +x_23 = x_243; +x_24 = x_189; +goto block_31; +} +} +} +} +} +else +{ +lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; uint8_t x_248; +x_244 = lean_ctor_get(x_8, 1); +lean_inc(x_244); +lean_dec(x_8); +x_245 = lean_ctor_get(x_244, 0); +lean_inc(x_245); +x_246 = lean_ctor_get(x_244, 1); +lean_inc(x_246); +x_247 = lean_ctor_get(x_244, 2); +lean_inc(x_247); +x_248 = lean_nat_dec_lt(x_246, x_247); +if (x_248 == 0) +{ +lean_object* x_249; lean_object* x_250; +lean_dec(x_247); +lean_dec(x_246); +lean_dec(x_245); +lean_dec(x_22); +lean_inc(x_3); +x_249 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_249, 0, x_3); +lean_ctor_set(x_249, 1, x_244); +x_250 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_250, 0, x_249); +x_23 = x_250; +x_24 = x_19; +goto block_31; +} +else +{ +lean_object* x_251; lean_object* x_252; uint8_t 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; uint8_t x_262; +if (lean_is_exclusive(x_244)) { + lean_ctor_release(x_244, 0); + lean_ctor_release(x_244, 1); + lean_ctor_release(x_244, 2); + x_251 = x_244; +} else { + lean_dec_ref(x_244); + x_251 = lean_box(0); +} +x_252 = lean_array_fget(x_245, x_246); +x_253 = lean_unbox(x_252); +lean_dec(x_252); +x_254 = lean_unsigned_to_nat(1u); +x_255 = lean_nat_add(x_246, x_254); +lean_dec(x_246); +if (lean_is_scalar(x_251)) { + x_256 = lean_alloc_ctor(0, 3, 0); +} else { + x_256 = x_251; +} +lean_ctor_set(x_256, 0, x_245); +lean_ctor_set(x_256, 1, x_255); +lean_ctor_set(x_256, 2, x_247); +x_257 = l_Lean_Expr_mvarId_x21(x_22); +x_258 = l_Lean_MVarId_isAssigned___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__4(x_257, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19); +lean_dec(x_257); +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); +} +x_262 = l_Lean_BinderInfo_isInstImplicit(x_253); +if (x_262 == 0) +{ +lean_object* x_263; lean_object* x_264; +lean_dec(x_261); +lean_dec(x_259); +lean_dec(x_22); +lean_inc(x_3); +x_263 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_263, 0, x_3); +lean_ctor_set(x_263, 1, x_256); +x_264 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_264, 0, x_263); +x_23 = x_264; +x_24 = x_260; +goto block_31; +} +else +{ +uint8_t x_265; +x_265 = lean_unbox(x_259); +lean_dec(x_259); +if (x_265 == 0) +{ +lean_object* x_266; +lean_inc(x_18); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_22); +x_266 = lean_infer_type(x_22, x_15, x_16, x_17, x_18, x_260); +if (lean_obj_tag(x_266) == 0) +{ +lean_object* x_267; lean_object* x_268; lean_object* x_269; +x_267 = lean_ctor_get(x_266, 0); +lean_inc(x_267); +x_268 = lean_ctor_get(x_266, 1); +lean_inc(x_268); +lean_dec(x_266); +lean_inc(x_18); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_267); +x_269 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem_synthesizeInstance(x_22, x_267, x_15, x_16, x_17, x_18, x_268); +if (lean_obj_tag(x_269) == 0) +{ +lean_object* x_270; uint8_t x_271; +x_270 = lean_ctor_get(x_269, 0); +lean_inc(x_270); +x_271 = lean_unbox(x_270); +lean_dec(x_270); +if (x_271 == 0) +{ +lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; uint8_t x_276; +x_272 = lean_ctor_get(x_269, 1); +lean_inc(x_272); +lean_dec(x_269); +x_273 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__2; +x_274 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__2(x_273, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_272); +x_275 = lean_ctor_get(x_274, 0); +lean_inc(x_275); +x_276 = lean_unbox(x_275); +lean_dec(x_275); +if (x_276 == 0) +{ +lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; +lean_dec(x_267); +lean_dec(x_261); +x_277 = lean_ctor_get(x_274, 1); +lean_inc(x_277); +lean_dec(x_274); +x_278 = lean_box(0); +x_279 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___lambda__1(x_256, x_278, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_277); +x_280 = lean_ctor_get(x_279, 0); +lean_inc(x_280); +x_281 = lean_ctor_get(x_279, 1); +lean_inc(x_281); +lean_dec(x_279); +x_23 = x_280; +x_24 = x_281; +goto block_31; +} +else +{ +lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; +x_282 = lean_ctor_get(x_274, 1); +lean_inc(x_282); +if (lean_is_exclusive(x_274)) { + lean_ctor_release(x_274, 0); + lean_ctor_release(x_274, 1); + x_283 = x_274; +} else { + lean_dec_ref(x_274); + x_283 = lean_box(0); +} +x_284 = lean_ctor_get(x_1, 5); +lean_inc(x_284); +x_285 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_284, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_282); +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, 0); +lean_inc(x_286); +x_287 = lean_ctor_get(x_285, 1); +lean_inc(x_287); +lean_dec(x_285); +x_288 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___closed__2; +if (lean_is_scalar(x_283)) { + x_289 = lean_alloc_ctor(7, 2, 0); +} else { + x_289 = x_283; + lean_ctor_set_tag(x_289, 7); +} +lean_ctor_set(x_289, 0, x_288); +lean_ctor_set(x_289, 1, x_286); +x_290 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; +if (lean_is_scalar(x_261)) { + x_291 = lean_alloc_ctor(7, 2, 0); +} else { + x_291 = x_261; + lean_ctor_set_tag(x_291, 7); +} +lean_ctor_set(x_291, 0, x_289); +lean_ctor_set(x_291, 1, x_290); +x_292 = l_Lean_indentExpr(x_267); +x_293 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_293, 0, x_291); +lean_ctor_set(x_293, 1, x_292); +x_294 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_294, 0, x_293); +lean_ctor_set(x_294, 1, x_290); +x_295 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_273, x_294, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_287); +x_296 = lean_ctor_get(x_295, 0); +lean_inc(x_296); +x_297 = lean_ctor_get(x_295, 1); +lean_inc(x_297); +lean_dec(x_295); +x_298 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___lambda__1(x_256, x_296, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_297); +lean_dec(x_296); +x_299 = lean_ctor_get(x_298, 0); +lean_inc(x_299); +x_300 = lean_ctor_get(x_298, 1); +lean_inc(x_300); +lean_dec(x_298); +x_23 = x_299; +x_24 = x_300; +goto block_31; +} +else +{ +lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; +lean_dec(x_283); +lean_dec(x_267); +lean_dec(x_261); +lean_dec(x_256); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_3); +lean_dec(x_1); +x_301 = lean_ctor_get(x_285, 0); +lean_inc(x_301); +x_302 = lean_ctor_get(x_285, 1); +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; +} else { + lean_dec_ref(x_285); + x_303 = lean_box(0); +} +if (lean_is_scalar(x_303)) { + x_304 = lean_alloc_ctor(1, 2, 0); +} else { + x_304 = x_303; +} +lean_ctor_set(x_304, 0, x_301); +lean_ctor_set(x_304, 1, x_302); +return x_304; +} +} +} +else +{ +lean_object* x_305; lean_object* x_306; lean_object* x_307; +lean_dec(x_267); +lean_dec(x_261); +x_305 = lean_ctor_get(x_269, 1); +lean_inc(x_305); +lean_dec(x_269); +lean_inc(x_3); +x_306 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_306, 0, x_3); +lean_ctor_set(x_306, 1, x_256); +x_307 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_307, 0, x_306); +x_23 = x_307; +x_24 = x_305; +goto block_31; +} +} +else +{ +lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; +lean_dec(x_267); +lean_dec(x_261); +lean_dec(x_256); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_3); +lean_dec(x_1); +x_308 = lean_ctor_get(x_269, 0); +lean_inc(x_308); +x_309 = lean_ctor_get(x_269, 1); +lean_inc(x_309); +if (lean_is_exclusive(x_269)) { + lean_ctor_release(x_269, 0); + lean_ctor_release(x_269, 1); + x_310 = x_269; +} else { + lean_dec_ref(x_269); + x_310 = lean_box(0); +} +if (lean_is_scalar(x_310)) { + x_311 = lean_alloc_ctor(1, 2, 0); +} else { + x_311 = x_310; +} +lean_ctor_set(x_311, 0, x_308); +lean_ctor_set(x_311, 1, x_309); +return x_311; +} +} +else +{ +lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; +lean_dec(x_261); +lean_dec(x_256); +lean_dec(x_22); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_3); +lean_dec(x_1); +x_312 = lean_ctor_get(x_266, 0); +lean_inc(x_312); +x_313 = lean_ctor_get(x_266, 1); +lean_inc(x_313); +if (lean_is_exclusive(x_266)) { + lean_ctor_release(x_266, 0); + lean_ctor_release(x_266, 1); + x_314 = x_266; +} else { + lean_dec_ref(x_266); + x_314 = lean_box(0); +} +if (lean_is_scalar(x_314)) { + x_315 = lean_alloc_ctor(1, 2, 0); +} else { + x_315 = x_314; +} +lean_ctor_set(x_315, 0, x_312); +lean_ctor_set(x_315, 1, x_313); +return x_315; +} +} +else +{ +lean_object* x_316; lean_object* x_317; +lean_dec(x_261); +lean_dec(x_22); +lean_inc(x_3); +x_316 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_316, 0, x_3); +lean_ctor_set(x_316, 1, x_256); +x_317 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_317, 0, x_316); +x_23 = x_317; +x_24 = x_260; +goto block_31; +} +} +} +} +block_31: +{ +if (lean_obj_tag(x_23) == 0) +{ +lean_object* x_25; lean_object* x_26; +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_3); +lean_dec(x_1); +x_25 = lean_ctor_get(x_23, 0); +lean_inc(x_25); +lean_dec(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_24); +return x_26; +} +else +{ +lean_object* x_27; size_t x_28; size_t x_29; +x_27 = lean_ctor_get(x_23, 0); +lean_inc(x_27); +lean_dec(x_23); +x_28 = 1; +x_29 = lean_usize_add(x_7, x_28); +x_7 = x_29; +x_8 = x_27; +x_19 = x_24; +goto _start; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, size_t x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18) { +_start: +{ +uint8_t x_19; +x_19 = lean_usize_dec_lt(x_6, x_5); +if (x_19 == 0) +{ +lean_object* x_20; +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_3); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_7); +lean_ctor_set(x_20, 1, x_18); +return x_20; +} +else +{ +lean_object* x_21; lean_object* x_22; +lean_dec(x_7); +x_21 = lean_array_uget(x_4, x_6); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_21); +x_22 = l_Lean_Meta_isProof(x_21, x_14, x_15, x_16, x_17, x_18); +if (lean_obj_tag(x_22) == 0) +{ +lean_object* x_23; uint8_t x_24; +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +x_24 = lean_unbox(x_23); +lean_dec(x_23); +if (x_24 == 0) +{ +uint8_t x_25; +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_3); +x_25 = !lean_is_exclusive(x_22); +if (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_22, 0); +lean_dec(x_26); +x_27 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_27, 0, x_21); +x_28 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_28, 0, x_27); +x_29 = lean_box(0); +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_28); +lean_ctor_set(x_30, 1, x_29); +lean_ctor_set(x_22, 0, x_30); +return x_22; +} +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_31 = lean_ctor_get(x_22, 1); +lean_inc(x_31); +lean_dec(x_22); +x_32 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_32, 0, x_21); +x_33 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_33, 0, x_32); +x_34 = lean_box(0); +x_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_33); +lean_ctor_set(x_35, 1, x_34); +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set(x_36, 1, x_31); +return x_36; +} +} +else +{ +lean_object* x_37; size_t x_38; size_t x_39; +lean_dec(x_21); +x_37 = lean_ctor_get(x_22, 1); +lean_inc(x_37); +lean_dec(x_22); +x_38 = 1; +x_39 = lean_usize_add(x_6, x_38); +lean_inc(x_3); +{ +size_t _tmp_5 = x_39; +lean_object* _tmp_6 = x_3; +lean_object* _tmp_17 = x_37; +x_6 = _tmp_5; +x_7 = _tmp_6; +x_18 = _tmp_17; +} +goto _start; +} +} +else +{ +uint8_t x_41; +lean_dec(x_21); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_3); +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_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, size_t x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18) { +_start: +{ +uint8_t x_19; +x_19 = lean_usize_dec_lt(x_6, x_5); +if (x_19 == 0) +{ +lean_object* x_20; +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_3); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_7); +lean_ctor_set(x_20, 1, x_18); +return x_20; +} +else +{ +lean_object* x_21; lean_object* x_22; +lean_dec(x_7); +x_21 = lean_array_uget(x_4, x_6); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_21); +x_22 = l_Lean_Meta_isProof(x_21, x_14, x_15, x_16, x_17, x_18); +if (lean_obj_tag(x_22) == 0) +{ +lean_object* x_23; uint8_t x_24; +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +x_24 = lean_unbox(x_23); +lean_dec(x_23); +if (x_24 == 0) +{ +uint8_t x_25; +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_3); +x_25 = !lean_is_exclusive(x_22); +if (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_22, 0); +lean_dec(x_26); +x_27 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_27, 0, x_21); +x_28 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_28, 0, x_27); +x_29 = lean_box(0); +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_28); +lean_ctor_set(x_30, 1, x_29); +lean_ctor_set(x_22, 0, x_30); +return x_22; +} +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_31 = lean_ctor_get(x_22, 1); +lean_inc(x_31); +lean_dec(x_22); +x_32 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_32, 0, x_21); +x_33 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_33, 0, x_32); +x_34 = lean_box(0); +x_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_33); +lean_ctor_set(x_35, 1, x_34); +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set(x_36, 1, x_31); +return x_36; +} +} +else +{ +lean_object* x_37; size_t x_38; size_t x_39; +lean_dec(x_21); +x_37 = lean_ctor_get(x_22, 1); +lean_inc(x_37); +lean_dec(x_22); +x_38 = 1; +x_39 = lean_usize_add(x_6, x_38); +lean_inc(x_3); +{ +size_t _tmp_5 = x_39; +lean_object* _tmp_6 = x_3; +lean_object* _tmp_17 = x_37; +x_6 = _tmp_5; +x_7 = _tmp_6; +x_18 = _tmp_17; +} +goto _start; +} +} +else +{ +uint8_t x_41; +lean_dec(x_21); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_3); +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_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__8(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +_start: +{ +uint8_t x_16; +x_16 = lean_usize_dec_eq(x_2, x_3); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_17 = lean_array_uget(x_1, x_2); +x_18 = l_Lean_Expr_mvarId_x21(x_17); +x_19 = l_Lean_MVarId_isAssigned___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__4(x_18, 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_18); +x_20 = lean_ctor_get(x_19, 0); +lean_inc(x_20); +x_21 = lean_unbox(x_20); +lean_dec(x_20); +if (x_21 == 0) +{ +lean_object* x_22; lean_object* x_23; size_t x_24; size_t x_25; +x_22 = lean_ctor_get(x_19, 1); +lean_inc(x_22); +lean_dec(x_19); +x_23 = lean_array_push(x_4, x_17); +x_24 = 1; +x_25 = lean_usize_add(x_2, x_24); +x_2 = x_25; +x_4 = x_23; +x_15 = x_22; +goto _start; +} +else +{ +lean_object* x_27; size_t x_28; size_t x_29; +lean_dec(x_17); +x_27 = lean_ctor_get(x_19, 1); +lean_inc(x_27); +lean_dec(x_19); +x_28 = 1; +x_29 = lean_usize_add(x_2, x_28); +x_2 = x_29; +x_15 = x_27; +goto _start; +} +} +else +{ +lean_object* x_31; +x_31 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_31, 0, x_4); +lean_ctor_set(x_31, 1, x_15); +return x_31; +} +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18) { +_start: +{ +uint8_t x_19; +x_19 = lean_usize_dec_lt(x_6, x_5); +if (x_19 == 0) +{ +lean_object* x_20; +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_3); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_7); +lean_ctor_set(x_20, 1, x_18); +return x_20; +} +else +{ +lean_object* x_21; lean_object* x_22; +lean_dec(x_7); +x_21 = lean_array_uget(x_4, x_6); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_21); +x_22 = l_Lean_Meta_isProof(x_21, x_14, x_15, x_16, x_17, x_18); +if (lean_obj_tag(x_22) == 0) +{ +lean_object* x_23; uint8_t x_24; +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +x_24 = lean_unbox(x_23); +lean_dec(x_23); +if (x_24 == 0) +{ +uint8_t x_25; +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_3); +x_25 = !lean_is_exclusive(x_22); +if (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_22, 0); +lean_dec(x_26); +x_27 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_27, 0, x_21); +x_28 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_28, 0, x_27); +x_29 = lean_box(0); +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_28); +lean_ctor_set(x_30, 1, x_29); +lean_ctor_set(x_22, 0, x_30); +return x_22; +} +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_31 = lean_ctor_get(x_22, 1); +lean_inc(x_31); +lean_dec(x_22); +x_32 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_32, 0, x_21); +x_33 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_33, 0, x_32); +x_34 = lean_box(0); +x_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_33); +lean_ctor_set(x_35, 1, x_34); +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set(x_36, 1, x_31); +return x_36; +} +} +else +{ +lean_object* x_37; size_t x_38; size_t x_39; +lean_dec(x_21); +x_37 = lean_ctor_get(x_22, 1); +lean_inc(x_37); +lean_dec(x_22); +x_38 = 1; +x_39 = lean_usize_add(x_6, x_38); +lean_inc(x_3); +{ +size_t _tmp_5 = x_39; +lean_object* _tmp_6 = x_3; +lean_object* _tmp_17 = x_37; +x_6 = _tmp_5; +x_7 = _tmp_6; +x_18 = _tmp_17; +} +goto _start; +} +} +else +{ +uint8_t x_41; +lean_dec(x_21); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_3); +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_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +uint8_t x_15; +x_15 = lean_usize_dec_eq(x_2, x_3); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_16 = lean_array_uget(x_1, x_2); +x_17 = l_Lean_Expr_mvarId_x21(x_16); +lean_dec(x_16); +x_18 = l_Lean_MVarId_isAssigned___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__4(x_17, 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_17); +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_unbox(x_19); +lean_dec(x_19); +if (x_20 == 0) +{ +uint8_t x_21; +x_21 = !lean_is_exclusive(x_18); +if (x_21 == 0) +{ +lean_object* x_22; uint8_t x_23; lean_object* x_24; +x_22 = lean_ctor_get(x_18, 0); +lean_dec(x_22); +x_23 = 1; +x_24 = lean_box(x_23); +lean_ctor_set(x_18, 0, x_24); +return x_18; +} +else +{ +lean_object* x_25; uint8_t x_26; lean_object* x_27; lean_object* x_28; +x_25 = lean_ctor_get(x_18, 1); +lean_inc(x_25); +lean_dec(x_18); +x_26 = 1; +x_27 = lean_box(x_26); +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_27); +lean_ctor_set(x_28, 1, x_25); +return x_28; +} +} +else +{ +lean_object* x_29; size_t x_30; size_t x_31; +x_29 = lean_ctor_get(x_18, 1); +lean_inc(x_29); +lean_dec(x_18); +x_30 = 1; +x_31 = lean_usize_add(x_2, x_30); +x_2 = x_31; +x_14 = x_29; +goto _start; +} +} +else +{ +uint8_t x_33; lean_object* x_34; lean_object* x_35; +x_33 = 0; +x_34 = lean_box(x_33); +x_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_34); +lean_ctor_set(x_35, 1, x_14); +return x_35; +} +} +} +LEAN_EXPORT lean_object* l_ReaderT_bind___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__11___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +lean_object* x_14; +lean_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); +lean_inc(x_3); +x_14 = lean_apply_11(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +x_16 = lean_ctor_get(x_14, 1); +lean_inc(x_16); +lean_dec(x_14); +x_17 = lean_apply_12(x_2, x_15, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_16); +return x_17; +} +else +{ +uint8_t x_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_2); +x_18 = !lean_is_exclusive(x_14); +if (x_18 == 0) +{ +return x_14; +} +else +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_14, 0); +x_20 = lean_ctor_get(x_14, 1); +lean_inc(x_20); +lean_inc(x_19); +lean_dec(x_14); +x_21 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_21, 0, x_19); +lean_ctor_set(x_21, 1, x_20); +return x_21; +} +} +} +} +LEAN_EXPORT lean_object* l_ReaderT_bind___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__11(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_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__11___rarg), 13, 0); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_withNewMCtxDepth___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__12___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_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +lean_object* x_14; lean_object* x_15; +x_14 = lean_apply_6(x_1, x_3, x_4, x_5, x_6, x_7, x_8); +x_15 = l___private_Lean_Meta_Basic_0__Lean_Meta_withNewMCtxDepthImp___rarg(x_2, x_14, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_15) == 0) +{ +uint8_t x_16; +x_16 = !lean_is_exclusive(x_15); +if (x_16 == 0) +{ +return x_15; +} +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_17 = lean_ctor_get(x_15, 0); +x_18 = lean_ctor_get(x_15, 1); +lean_inc(x_18); +lean_inc(x_17); +lean_dec(x_15); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_17); +lean_ctor_set(x_19, 1, x_18); +return x_19; +} +} +else +{ +uint8_t x_20; +x_20 = !lean_is_exclusive(x_15); +if (x_20 == 0) +{ +return x_15; +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_15, 0); +x_22 = lean_ctor_get(x_15, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_15); +x_23 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_23, 0, x_21); +lean_ctor_set(x_23, 1, x_22); +return x_23; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_withNewMCtxDepth___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__12(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_Meta_withNewMCtxDepth___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__12___rarg___boxed), 13, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +_start: +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; uint8_t x_21; uint8_t x_22; lean_object* x_23; +x_17 = l_Lean_instantiateMVars___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__1(x_1, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +x_18 = lean_ctor_get(x_17, 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 = 1; +x_22 = 0; +x_23 = l_Lean_Meta_mkLambdaFVars(x_2, x_18, x_20, x_21, x_20, x_22, x_12, x_13, x_14, x_15, x_19); +if (lean_obj_tag(x_23) == 0) +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +x_25 = lean_ctor_get(x_23, 1); +lean_inc(x_25); +lean_dec(x_23); +x_26 = lean_ctor_get(x_3, 5); +lean_inc(x_26); +lean_dec(x_3); +x_27 = lean_ctor_get(x_4, 1); +lean_inc(x_27); +lean_dec(x_4); +x_28 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance(x_26, x_24, x_27, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_25); +return x_28; +} +else +{ +uint8_t x_29; +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_4); +lean_dec(x_3); +x_29 = !lean_is_exclusive(x_23); +if (x_29 == 0) +{ +return x_23; +} +else +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_30 = lean_ctor_get(x_23, 0); +x_31 = lean_ctor_get(x_23, 1); +lean_inc(x_31); +lean_inc(x_30); +lean_dec(x_23); +x_32 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_32, 0, x_30); +lean_ctor_set(x_32, 1, x_31); +return x_32; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +lean_object* x_14; +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; +} +} +static size_t _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__1() { +_start: +{ +lean_object* x_1; size_t x_2; +x_1 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2; +x_2 = lean_array_size(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("failed to instantiate ", 22, 22); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__2; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(", failed to instantiate non propositional argument with type", 60, 60); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__4; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, size_t x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19, lean_object* x_20, lean_object* x_21) { +_start: +{ +lean_object* x_22; uint8_t x_23; uint8_t x_24; lean_object* x_25; +x_22 = lean_unsigned_to_nat(0u); +x_23 = lean_nat_dec_lt(x_22, x_1); +if (x_23 == 0) +{ +uint8_t x_292; +x_292 = 1; +x_24 = x_292; +x_25 = x_21; +goto block_291; +} +else +{ +size_t x_293; lean_object* x_294; lean_object* x_295; uint8_t x_296; +x_293 = lean_usize_of_nat(x_1); +x_294 = l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__10(x_3, x_8, x_293, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21); +x_295 = lean_ctor_get(x_294, 0); +lean_inc(x_295); +x_296 = lean_unbox(x_295); +lean_dec(x_295); +if (x_296 == 0) +{ +lean_object* x_297; uint8_t x_298; +x_297 = lean_ctor_get(x_294, 1); +lean_inc(x_297); +lean_dec(x_294); +x_298 = 1; +x_24 = x_298; +x_25 = x_297; +goto block_291; +} +else +{ +lean_object* x_299; uint8_t x_300; +x_299 = lean_ctor_get(x_294, 1); +lean_inc(x_299); +lean_dec(x_294); +x_300 = 0; +x_24 = x_300; +x_25 = x_299; +goto block_291; +} +} +block_291: +{ +if (x_24 == 0) +{ +lean_object* x_26; +x_26 = l_Lean_mkAppN(x_2, x_3); +if (x_23 == 0) +{ +lean_object* x_27; lean_object* x_28; lean_object* x_101; size_t x_102; lean_object* x_103; +x_101 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2; +x_102 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__1; +lean_inc(x_20); +lean_inc(x_19); +lean_inc(x_18); +lean_inc(x_17); +lean_inc(x_7); +x_103 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__6(x_6, x_101, x_7, x_101, x_102, x_8, x_7, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_25); +if (lean_obj_tag(x_103) == 0) +{ +lean_object* x_104; lean_object* x_105; +x_104 = lean_ctor_get(x_103, 0); +lean_inc(x_104); +x_105 = lean_ctor_get(x_104, 0); +lean_inc(x_105); +lean_dec(x_104); +if (lean_obj_tag(x_105) == 0) +{ +lean_object* x_106; +x_106 = lean_ctor_get(x_103, 1); +lean_inc(x_106); +lean_dec(x_103); +x_27 = x_9; +x_28 = x_106; +goto block_100; +} +else +{ +lean_object* x_107; lean_object* x_108; +lean_dec(x_9); +x_107 = lean_ctor_get(x_103, 1); +lean_inc(x_107); +lean_dec(x_103); +x_108 = lean_ctor_get(x_105, 0); +lean_inc(x_108); +lean_dec(x_105); +x_27 = x_108; +x_28 = x_107; +goto block_100; +} +} +else +{ +uint8_t x_109; +lean_dec(x_26); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_9); +lean_dec(x_5); +lean_dec(x_4); +x_109 = !lean_is_exclusive(x_103); +if (x_109 == 0) +{ +return x_103; +} +else +{ +lean_object* x_110; lean_object* x_111; lean_object* x_112; +x_110 = lean_ctor_get(x_103, 0); +x_111 = lean_ctor_get(x_103, 1); +lean_inc(x_111); +lean_inc(x_110); +lean_dec(x_103); +x_112 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_112, 0, x_110); +lean_ctor_set(x_112, 1, x_111); +return x_112; +} +} +block_100: +{ +if (lean_obj_tag(x_27) == 0) +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_29 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2; +x_30 = lean_box(0); +x_31 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__1(x_26, x_29, x_4, x_5, x_30, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_28); +return x_31; +} +else +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; +x_32 = lean_ctor_get(x_27, 0); +lean_inc(x_32); +lean_dec(x_27); +x_33 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__2; +x_34 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__2(x_33, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_28); +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_object* x_40; +lean_dec(x_32); +x_37 = lean_ctor_get(x_34, 1); +lean_inc(x_37); +lean_dec(x_34); +x_38 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2; +x_39 = lean_box(0); +x_40 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__1(x_26, x_38, x_4, x_5, x_39, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_37); +return x_40; +} +else +{ +uint8_t x_41; +x_41 = !lean_is_exclusive(x_34); +if (x_41 == 0) +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_42 = lean_ctor_get(x_34, 1); +x_43 = lean_ctor_get(x_34, 0); +lean_dec(x_43); +x_44 = lean_ctor_get(x_4, 5); +lean_inc(x_44); +x_45 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_44, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_42); +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, 0); +lean_inc(x_46); +x_47 = lean_ctor_get(x_45, 1); +lean_inc(x_47); +lean_dec(x_45); +lean_inc(x_20); +lean_inc(x_19); +lean_inc(x_18); +lean_inc(x_17); +x_48 = lean_infer_type(x_32, x_17, x_18, x_19, x_20, x_47); +if (lean_obj_tag(x_48) == 0) +{ +lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; +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___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__3; +lean_ctor_set_tag(x_34, 7); +lean_ctor_set(x_34, 1, x_46); +lean_ctor_set(x_34, 0, x_51); +x_52 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__5; +x_53 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_53, 0, x_34); +lean_ctor_set(x_53, 1, x_52); +x_54 = l_Lean_indentExpr(x_49); +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_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; +x_57 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_57, 0, x_55); +lean_ctor_set(x_57, 1, x_56); +x_58 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_33, x_57, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_50); +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_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2; +x_62 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__1(x_26, x_61, x_4, x_5, x_59, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_60); +lean_dec(x_59); +return x_62; +} +else +{ +uint8_t x_63; +lean_dec(x_46); +lean_free_object(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_5); +lean_dec(x_4); +x_63 = !lean_is_exclusive(x_48); +if (x_63 == 0) +{ +return x_48; +} +else +{ +lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_64 = lean_ctor_get(x_48, 0); +x_65 = lean_ctor_get(x_48, 1); +lean_inc(x_65); +lean_inc(x_64); +lean_dec(x_48); +x_66 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_66, 0, x_64); +lean_ctor_set(x_66, 1, x_65); +return x_66; +} +} +} +else +{ +uint8_t x_67; +lean_free_object(x_34); +lean_dec(x_32); +lean_dec(x_26); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_5); +lean_dec(x_4); +x_67 = !lean_is_exclusive(x_45); +if (x_67 == 0) +{ +return x_45; +} +else +{ +lean_object* x_68; lean_object* x_69; lean_object* x_70; +x_68 = lean_ctor_get(x_45, 0); +x_69 = lean_ctor_get(x_45, 1); +lean_inc(x_69); +lean_inc(x_68); +lean_dec(x_45); +x_70 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_70, 0, x_68); +lean_ctor_set(x_70, 1, x_69); +return x_70; +} +} +} +else +{ +lean_object* x_71; lean_object* x_72; lean_object* x_73; +x_71 = lean_ctor_get(x_34, 1); +lean_inc(x_71); +lean_dec(x_34); +x_72 = lean_ctor_get(x_4, 5); +lean_inc(x_72); +x_73 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_72, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_71); +if (lean_obj_tag(x_73) == 0) +{ +lean_object* x_74; lean_object* x_75; lean_object* x_76; +x_74 = lean_ctor_get(x_73, 0); +lean_inc(x_74); +x_75 = lean_ctor_get(x_73, 1); +lean_inc(x_75); +lean_dec(x_73); +lean_inc(x_20); +lean_inc(x_19); +lean_inc(x_18); +lean_inc(x_17); +x_76 = lean_infer_type(x_32, x_17, x_18, x_19, x_20, x_75); +if (lean_obj_tag(x_76) == 0) +{ +lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; +x_77 = lean_ctor_get(x_76, 0); +lean_inc(x_77); +x_78 = lean_ctor_get(x_76, 1); +lean_inc(x_78); +lean_dec(x_76); +x_79 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__3; +x_80 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_80, 0, x_79); +lean_ctor_set(x_80, 1, x_74); +x_81 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__5; +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_indentExpr(x_77); +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_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; +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_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_33, x_86, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_78); +x_88 = lean_ctor_get(x_87, 0); +lean_inc(x_88); +x_89 = lean_ctor_get(x_87, 1); +lean_inc(x_89); +lean_dec(x_87); +x_90 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2; +x_91 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__1(x_26, x_90, x_4, x_5, x_88, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_89); +lean_dec(x_88); +return x_91; +} +else +{ +lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; +lean_dec(x_74); +lean_dec(x_26); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_5); +lean_dec(x_4); +x_92 = lean_ctor_get(x_76, 0); +lean_inc(x_92); +x_93 = lean_ctor_get(x_76, 1); +lean_inc(x_93); +if (lean_is_exclusive(x_76)) { + lean_ctor_release(x_76, 0); + lean_ctor_release(x_76, 1); + x_94 = x_76; +} else { + lean_dec_ref(x_76); + x_94 = lean_box(0); +} +if (lean_is_scalar(x_94)) { + x_95 = lean_alloc_ctor(1, 2, 0); +} else { + x_95 = x_94; +} +lean_ctor_set(x_95, 0, x_92); +lean_ctor_set(x_95, 1, x_93); +return x_95; +} +} +else +{ +lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; +lean_dec(x_32); +lean_dec(x_26); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_5); +lean_dec(x_4); +x_96 = lean_ctor_get(x_73, 0); +lean_inc(x_96); +x_97 = lean_ctor_get(x_73, 1); +lean_inc(x_97); +if (lean_is_exclusive(x_73)) { + lean_ctor_release(x_73, 0); + lean_ctor_release(x_73, 1); + x_98 = x_73; +} else { + lean_dec_ref(x_73); + x_98 = lean_box(0); +} +if (lean_is_scalar(x_98)) { + x_99 = lean_alloc_ctor(1, 2, 0); +} else { + x_99 = x_98; +} +lean_ctor_set(x_99, 0, x_96); +lean_ctor_set(x_99, 1, x_97); +return x_99; +} +} +} +} +} +} +else +{ +uint8_t x_113; +x_113 = lean_nat_dec_le(x_1, x_1); +if (x_113 == 0) +{ +lean_object* x_114; lean_object* x_115; lean_object* x_188; size_t x_189; lean_object* x_190; +x_188 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2; +x_189 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__1; +lean_inc(x_20); +lean_inc(x_19); +lean_inc(x_18); +lean_inc(x_17); +lean_inc(x_7); +x_190 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__7(x_6, x_188, x_7, x_188, x_189, x_8, x_7, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_25); +if (lean_obj_tag(x_190) == 0) +{ +lean_object* x_191; lean_object* x_192; +x_191 = lean_ctor_get(x_190, 0); +lean_inc(x_191); +x_192 = lean_ctor_get(x_191, 0); +lean_inc(x_192); +lean_dec(x_191); +if (lean_obj_tag(x_192) == 0) +{ +lean_object* x_193; +x_193 = lean_ctor_get(x_190, 1); +lean_inc(x_193); +lean_dec(x_190); +x_114 = x_9; +x_115 = x_193; +goto block_187; +} +else +{ +lean_object* x_194; lean_object* x_195; +lean_dec(x_9); +x_194 = lean_ctor_get(x_190, 1); +lean_inc(x_194); +lean_dec(x_190); +x_195 = lean_ctor_get(x_192, 0); +lean_inc(x_195); +lean_dec(x_192); +x_114 = x_195; +x_115 = x_194; +goto block_187; +} +} +else +{ +uint8_t x_196; +lean_dec(x_26); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_9); +lean_dec(x_5); +lean_dec(x_4); +x_196 = !lean_is_exclusive(x_190); +if (x_196 == 0) +{ +return x_190; +} +else +{ +lean_object* x_197; lean_object* x_198; lean_object* x_199; +x_197 = lean_ctor_get(x_190, 0); +x_198 = lean_ctor_get(x_190, 1); +lean_inc(x_198); +lean_inc(x_197); +lean_dec(x_190); +x_199 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_199, 0, x_197); +lean_ctor_set(x_199, 1, x_198); +return x_199; +} +} +block_187: +{ +if (lean_obj_tag(x_114) == 0) +{ +lean_object* x_116; lean_object* x_117; lean_object* x_118; +x_116 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2; +x_117 = lean_box(0); +x_118 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__1(x_26, x_116, x_4, x_5, x_117, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_115); +return x_118; +} +else +{ +lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; uint8_t x_123; +x_119 = lean_ctor_get(x_114, 0); +lean_inc(x_119); +lean_dec(x_114); +x_120 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__2; +x_121 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__2(x_120, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_115); +x_122 = lean_ctor_get(x_121, 0); +lean_inc(x_122); +x_123 = lean_unbox(x_122); +lean_dec(x_122); +if (x_123 == 0) +{ +lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; +lean_dec(x_119); +x_124 = lean_ctor_get(x_121, 1); +lean_inc(x_124); +lean_dec(x_121); +x_125 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2; +x_126 = lean_box(0); +x_127 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__1(x_26, x_125, x_4, x_5, x_126, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_124); +return x_127; +} +else +{ +uint8_t x_128; +x_128 = !lean_is_exclusive(x_121); +if (x_128 == 0) +{ +lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; +x_129 = lean_ctor_get(x_121, 1); +x_130 = lean_ctor_get(x_121, 0); +lean_dec(x_130); +x_131 = lean_ctor_get(x_4, 5); +lean_inc(x_131); +x_132 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_131, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_129); +if (lean_obj_tag(x_132) == 0) +{ +lean_object* x_133; lean_object* x_134; lean_object* x_135; +x_133 = lean_ctor_get(x_132, 0); +lean_inc(x_133); +x_134 = lean_ctor_get(x_132, 1); +lean_inc(x_134); +lean_dec(x_132); +lean_inc(x_20); +lean_inc(x_19); +lean_inc(x_18); +lean_inc(x_17); +x_135 = lean_infer_type(x_119, x_17, x_18, x_19, x_20, x_134); +if (lean_obj_tag(x_135) == 0) +{ +lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; +x_136 = lean_ctor_get(x_135, 0); +lean_inc(x_136); +x_137 = lean_ctor_get(x_135, 1); +lean_inc(x_137); +lean_dec(x_135); +x_138 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__3; +lean_ctor_set_tag(x_121, 7); +lean_ctor_set(x_121, 1, x_133); +lean_ctor_set(x_121, 0, x_138); +x_139 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__5; +x_140 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_140, 0, x_121); +lean_ctor_set(x_140, 1, x_139); +x_141 = l_Lean_indentExpr(x_136); +x_142 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_142, 0, x_140); +lean_ctor_set(x_142, 1, x_141); +x_143 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; +x_144 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_144, 0, x_142); +lean_ctor_set(x_144, 1, x_143); +x_145 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_120, x_144, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, 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_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2; +x_149 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__1(x_26, x_148, x_4, x_5, x_146, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_147); +lean_dec(x_146); +return x_149; +} +else +{ +uint8_t x_150; +lean_dec(x_133); +lean_free_object(x_121); +lean_dec(x_26); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_5); +lean_dec(x_4); +x_150 = !lean_is_exclusive(x_135); +if (x_150 == 0) +{ +return x_135; +} +else +{ +lean_object* x_151; lean_object* x_152; lean_object* x_153; +x_151 = lean_ctor_get(x_135, 0); +x_152 = lean_ctor_get(x_135, 1); +lean_inc(x_152); +lean_inc(x_151); +lean_dec(x_135); +x_153 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_153, 0, x_151); +lean_ctor_set(x_153, 1, x_152); +return x_153; +} +} +} +else +{ +uint8_t x_154; +lean_free_object(x_121); +lean_dec(x_119); +lean_dec(x_26); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_5); +lean_dec(x_4); +x_154 = !lean_is_exclusive(x_132); +if (x_154 == 0) +{ +return x_132; +} +else +{ +lean_object* x_155; lean_object* x_156; lean_object* x_157; +x_155 = lean_ctor_get(x_132, 0); +x_156 = lean_ctor_get(x_132, 1); +lean_inc(x_156); +lean_inc(x_155); +lean_dec(x_132); +x_157 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_157, 0, x_155); +lean_ctor_set(x_157, 1, x_156); +return x_157; +} +} +} +else +{ +lean_object* x_158; lean_object* x_159; lean_object* x_160; +x_158 = lean_ctor_get(x_121, 1); +lean_inc(x_158); +lean_dec(x_121); +x_159 = lean_ctor_get(x_4, 5); +lean_inc(x_159); +x_160 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_159, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_158); +if (lean_obj_tag(x_160) == 0) +{ +lean_object* x_161; lean_object* x_162; lean_object* x_163; +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); +lean_inc(x_20); +lean_inc(x_19); +lean_inc(x_18); +lean_inc(x_17); +x_163 = lean_infer_type(x_119, x_17, x_18, x_19, x_20, x_162); +if (lean_obj_tag(x_163) == 0) +{ +lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; +x_164 = lean_ctor_get(x_163, 0); +lean_inc(x_164); +x_165 = lean_ctor_get(x_163, 1); +lean_inc(x_165); +lean_dec(x_163); +x_166 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__3; +x_167 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_167, 0, x_166); +lean_ctor_set(x_167, 1, x_161); +x_168 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__5; +x_169 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_169, 0, x_167); +lean_ctor_set(x_169, 1, x_168); +x_170 = l_Lean_indentExpr(x_164); +x_171 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_171, 0, x_169); +lean_ctor_set(x_171, 1, x_170); +x_172 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; +x_173 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_173, 0, x_171); +lean_ctor_set(x_173, 1, x_172); +x_174 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_120, x_173, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_165); +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_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2; +x_178 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__1(x_26, x_177, x_4, x_5, x_175, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_176); +lean_dec(x_175); +return x_178; +} +else +{ +lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; +lean_dec(x_161); +lean_dec(x_26); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_5); +lean_dec(x_4); +x_179 = lean_ctor_get(x_163, 0); +lean_inc(x_179); +x_180 = lean_ctor_get(x_163, 1); +lean_inc(x_180); +if (lean_is_exclusive(x_163)) { + lean_ctor_release(x_163, 0); + lean_ctor_release(x_163, 1); + x_181 = x_163; +} else { + lean_dec_ref(x_163); + x_181 = lean_box(0); +} +if (lean_is_scalar(x_181)) { + x_182 = lean_alloc_ctor(1, 2, 0); +} else { + x_182 = x_181; +} +lean_ctor_set(x_182, 0, x_179); +lean_ctor_set(x_182, 1, x_180); +return x_182; +} +} +else +{ +lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; +lean_dec(x_119); +lean_dec(x_26); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_5); +lean_dec(x_4); +x_183 = lean_ctor_get(x_160, 0); +lean_inc(x_183); +x_184 = lean_ctor_get(x_160, 1); +lean_inc(x_184); +if (lean_is_exclusive(x_160)) { + lean_ctor_release(x_160, 0); + lean_ctor_release(x_160, 1); + x_185 = x_160; +} else { + lean_dec_ref(x_160); + x_185 = lean_box(0); +} +if (lean_is_scalar(x_185)) { + x_186 = lean_alloc_ctor(1, 2, 0); +} else { + x_186 = x_185; +} +lean_ctor_set(x_186, 0, x_183); +lean_ctor_set(x_186, 1, x_184); +return x_186; +} +} +} +} +} +} +else +{ +size_t x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; size_t x_206; lean_object* x_207; lean_object* x_208; lean_object* x_277; +x_200 = lean_usize_of_nat(x_1); +x_201 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2; +x_202 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__8(x_3, x_8, x_200, x_201, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_25); +x_203 = lean_ctor_get(x_202, 0); +lean_inc(x_203); +x_204 = lean_ctor_get(x_202, 1); +lean_inc(x_204); +if (lean_is_exclusive(x_202)) { + lean_ctor_release(x_202, 0); + lean_ctor_release(x_202, 1); + x_205 = x_202; +} else { + lean_dec_ref(x_202); + x_205 = lean_box(0); +} +x_206 = lean_array_size(x_203); +lean_inc(x_20); +lean_inc(x_19); +lean_inc(x_18); +lean_inc(x_17); +lean_inc(x_7); +x_277 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__9(x_6, x_203, x_7, x_203, x_206, x_8, x_7, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_204); +if (lean_obj_tag(x_277) == 0) +{ +lean_object* x_278; lean_object* x_279; +x_278 = lean_ctor_get(x_277, 0); +lean_inc(x_278); +x_279 = lean_ctor_get(x_278, 0); +lean_inc(x_279); +lean_dec(x_278); +if (lean_obj_tag(x_279) == 0) +{ +lean_object* x_280; +x_280 = lean_ctor_get(x_277, 1); +lean_inc(x_280); +lean_dec(x_277); +x_207 = x_9; +x_208 = x_280; +goto block_276; +} +else +{ +lean_object* x_281; lean_object* x_282; +lean_dec(x_9); +x_281 = lean_ctor_get(x_277, 1); +lean_inc(x_281); +lean_dec(x_277); +x_282 = lean_ctor_get(x_279, 0); +lean_inc(x_282); +lean_dec(x_279); +x_207 = x_282; +x_208 = x_281; +goto block_276; +} +} +else +{ +uint8_t x_283; +lean_dec(x_205); +lean_dec(x_203); +lean_dec(x_26); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_9); +lean_dec(x_5); +lean_dec(x_4); +x_283 = !lean_is_exclusive(x_277); +if (x_283 == 0) +{ +return x_277; +} +else +{ +lean_object* x_284; lean_object* x_285; lean_object* x_286; +x_284 = lean_ctor_get(x_277, 0); +x_285 = lean_ctor_get(x_277, 1); +lean_inc(x_285); +lean_inc(x_284); +lean_dec(x_277); +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; +} +} +block_276: +{ +if (lean_obj_tag(x_207) == 0) +{ +lean_object* x_209; lean_object* x_210; +lean_dec(x_205); +x_209 = lean_box(0); +x_210 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__1(x_26, x_203, x_4, x_5, x_209, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_208); +lean_dec(x_203); +return x_210; +} +else +{ +lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; uint8_t x_215; +x_211 = lean_ctor_get(x_207, 0); +lean_inc(x_211); +lean_dec(x_207); +x_212 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__2; +x_213 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__2(x_212, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_208); +x_214 = lean_ctor_get(x_213, 0); +lean_inc(x_214); +x_215 = lean_unbox(x_214); +lean_dec(x_214); +if (x_215 == 0) +{ +lean_object* x_216; lean_object* x_217; lean_object* x_218; +lean_dec(x_211); +lean_dec(x_205); +x_216 = lean_ctor_get(x_213, 1); +lean_inc(x_216); +lean_dec(x_213); +x_217 = lean_box(0); +x_218 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__1(x_26, x_203, x_4, x_5, x_217, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_216); +lean_dec(x_203); +return x_218; +} +else +{ +uint8_t x_219; +x_219 = !lean_is_exclusive(x_213); +if (x_219 == 0) +{ +lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; +x_220 = lean_ctor_get(x_213, 1); +x_221 = lean_ctor_get(x_213, 0); +lean_dec(x_221); +x_222 = lean_ctor_get(x_4, 5); +lean_inc(x_222); +x_223 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_222, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_220); +if (lean_obj_tag(x_223) == 0) +{ +lean_object* x_224; lean_object* x_225; lean_object* x_226; +x_224 = lean_ctor_get(x_223, 0); +lean_inc(x_224); +x_225 = lean_ctor_get(x_223, 1); +lean_inc(x_225); +lean_dec(x_223); +lean_inc(x_20); +lean_inc(x_19); +lean_inc(x_18); +lean_inc(x_17); +x_226 = lean_infer_type(x_211, x_17, x_18, x_19, x_20, x_225); +if (lean_obj_tag(x_226) == 0) +{ +lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; +x_227 = lean_ctor_get(x_226, 0); +lean_inc(x_227); +x_228 = lean_ctor_get(x_226, 1); +lean_inc(x_228); +lean_dec(x_226); +x_229 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__3; +lean_ctor_set_tag(x_213, 7); +lean_ctor_set(x_213, 1, x_224); +lean_ctor_set(x_213, 0, x_229); +x_230 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__5; +if (lean_is_scalar(x_205)) { + x_231 = lean_alloc_ctor(7, 2, 0); +} else { + x_231 = x_205; + lean_ctor_set_tag(x_231, 7); +} +lean_ctor_set(x_231, 0, x_213); +lean_ctor_set(x_231, 1, x_230); +x_232 = l_Lean_indentExpr(x_227); +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_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; +x_235 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_235, 0, x_233); +lean_ctor_set(x_235, 1, x_234); +x_236 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_212, x_235, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_228); +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 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__1(x_26, x_203, x_4, x_5, x_237, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_238); +lean_dec(x_237); +lean_dec(x_203); +return x_239; +} +else +{ +uint8_t x_240; +lean_dec(x_224); +lean_free_object(x_213); +lean_dec(x_205); +lean_dec(x_203); +lean_dec(x_26); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_5); +lean_dec(x_4); +x_240 = !lean_is_exclusive(x_226); +if (x_240 == 0) +{ +return x_226; +} +else +{ +lean_object* x_241; lean_object* x_242; lean_object* x_243; +x_241 = lean_ctor_get(x_226, 0); +x_242 = lean_ctor_get(x_226, 1); +lean_inc(x_242); +lean_inc(x_241); +lean_dec(x_226); +x_243 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_243, 0, x_241); +lean_ctor_set(x_243, 1, x_242); +return x_243; +} +} +} +else +{ +uint8_t x_244; +lean_free_object(x_213); +lean_dec(x_211); +lean_dec(x_205); +lean_dec(x_203); +lean_dec(x_26); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_5); +lean_dec(x_4); +x_244 = !lean_is_exclusive(x_223); +if (x_244 == 0) +{ +return x_223; +} +else +{ +lean_object* x_245; lean_object* x_246; lean_object* x_247; +x_245 = lean_ctor_get(x_223, 0); +x_246 = lean_ctor_get(x_223, 1); +lean_inc(x_246); +lean_inc(x_245); +lean_dec(x_223); +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; +} +} +} +else +{ +lean_object* x_248; lean_object* x_249; lean_object* x_250; +x_248 = lean_ctor_get(x_213, 1); +lean_inc(x_248); +lean_dec(x_213); +x_249 = lean_ctor_get(x_4, 5); +lean_inc(x_249); +x_250 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_249, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_248); +if (lean_obj_tag(x_250) == 0) +{ +lean_object* x_251; lean_object* x_252; lean_object* x_253; +x_251 = lean_ctor_get(x_250, 0); +lean_inc(x_251); +x_252 = lean_ctor_get(x_250, 1); +lean_inc(x_252); +lean_dec(x_250); +lean_inc(x_20); +lean_inc(x_19); +lean_inc(x_18); +lean_inc(x_17); +x_253 = lean_infer_type(x_211, x_17, x_18, x_19, x_20, x_252); +if (lean_obj_tag(x_253) == 0) +{ +lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; +x_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_253); +x_256 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__3; +x_257 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_257, 0, x_256); +lean_ctor_set(x_257, 1, x_251); +x_258 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__5; +if (lean_is_scalar(x_205)) { + x_259 = lean_alloc_ctor(7, 2, 0); +} else { + x_259 = x_205; + lean_ctor_set_tag(x_259, 7); +} +lean_ctor_set(x_259, 0, x_257); +lean_ctor_set(x_259, 1, x_258); +x_260 = l_Lean_indentExpr(x_254); +x_261 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_261, 0, x_259); +lean_ctor_set(x_261, 1, x_260); +x_262 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; +x_263 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_263, 0, x_261); +lean_ctor_set(x_263, 1, x_262); +x_264 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_212, x_263, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_255); +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___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__1(x_26, x_203, x_4, x_5, x_265, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_266); +lean_dec(x_265); +lean_dec(x_203); +return x_267; +} +else +{ +lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; +lean_dec(x_251); +lean_dec(x_205); +lean_dec(x_203); +lean_dec(x_26); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_5); +lean_dec(x_4); +x_268 = lean_ctor_get(x_253, 0); +lean_inc(x_268); +x_269 = lean_ctor_get(x_253, 1); +lean_inc(x_269); +if (lean_is_exclusive(x_253)) { + lean_ctor_release(x_253, 0); + lean_ctor_release(x_253, 1); + x_270 = x_253; +} else { + lean_dec_ref(x_253); + x_270 = lean_box(0); +} +if (lean_is_scalar(x_270)) { + x_271 = lean_alloc_ctor(1, 2, 0); +} else { + x_271 = x_270; +} +lean_ctor_set(x_271, 0, x_268); +lean_ctor_set(x_271, 1, x_269); +return x_271; +} +} +else +{ +lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; +lean_dec(x_211); +lean_dec(x_205); +lean_dec(x_203); +lean_dec(x_26); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_5); +lean_dec(x_4); +x_272 = lean_ctor_get(x_250, 0); +lean_inc(x_272); +x_273 = lean_ctor_get(x_250, 1); +lean_inc(x_273); +if (lean_is_exclusive(x_250)) { + lean_ctor_release(x_250, 0); + lean_ctor_release(x_250, 1); + x_274 = x_250; +} else { + lean_dec_ref(x_250); + x_274 = lean_box(0); +} +if (lean_is_scalar(x_274)) { + x_275 = lean_alloc_ctor(1, 2, 0); +} else { + x_275 = x_274; +} +lean_ctor_set(x_275, 0, x_272); +lean_ctor_set(x_275, 1, x_273); +return x_275; +} +} +} +} +} +} +} +} +else +{ +lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; +lean_dec(x_9); +lean_dec(x_7); +x_287 = lean_ctor_get(x_4, 5); +lean_inc(x_287); +lean_dec(x_4); +x_288 = l_Lean_mkAppN(x_2, x_3); +x_289 = lean_ctor_get(x_5, 1); +lean_inc(x_289); +lean_dec(x_5); +x_290 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance(x_287, x_288, x_289, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_25); +return x_290; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19, lean_object* x_20) { +_start: +{ +lean_object* x_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_21 = lean_array_get_size(x_1); +x_22 = lean_unsigned_to_nat(0u); +x_23 = l_Array_toSubarray___rarg(x_1, x_22, x_21); +x_24 = lean_box(0); +lean_inc(x_2); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_2); +lean_ctor_set(x_25, 1, x_23); +x_26 = lean_array_size(x_3); +x_27 = 0; +lean_inc(x_19); +lean_inc(x_18); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_2); +lean_inc(x_4); +x_28 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5(x_4, x_3, x_2, x_24, x_3, x_26, x_27, x_25, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20); +if (lean_obj_tag(x_28) == 0) +{ +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); +if (lean_obj_tag(x_30) == 0) +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_28, 1); +lean_inc(x_31); +lean_dec(x_28); +x_32 = lean_box(0); +x_33 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3(x_5, x_6, x_3, x_4, x_7, x_24, x_8, x_27, x_2, x_32, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_31); +return x_33; +} +else +{ +uint8_t x_34; +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_2); +x_34 = !lean_is_exclusive(x_28); +if (x_34 == 0) +{ +lean_object* x_35; lean_object* x_36; +x_35 = lean_ctor_get(x_28, 0); +lean_dec(x_35); +x_36 = lean_ctor_get(x_30, 0); +lean_inc(x_36); +lean_dec(x_30); +lean_ctor_set(x_28, 0, x_36); +return x_28; +} +else +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_37 = lean_ctor_get(x_28, 1); +lean_inc(x_37); +lean_dec(x_28); +x_38 = lean_ctor_get(x_30, 0); +lean_inc(x_38); +lean_dec(x_30); +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_40; +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_2); +x_40 = !lean_is_exclusive(x_28); +if (x_40 == 0) +{ +return x_28; +} +else +{ +lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_41 = lean_ctor_get(x_28, 0); +x_42 = lean_ctor_get(x_28, 1); +lean_inc(x_42); +lean_inc(x_41); +lean_dec(x_28); +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; +} +} +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__5___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = lean_box(0); +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19, lean_object* x_20) { +_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; +x_21 = lean_array_get_size(x_1); +x_22 = lean_unsigned_to_nat(0u); +x_23 = lean_unsigned_to_nat(1u); +lean_inc(x_21); +x_24 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_24, 0, x_22); +lean_ctor_set(x_24, 1, x_21); +lean_ctor_set(x_24, 2, x_23); +x_25 = lean_box(0); +x_26 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__5___closed__1; +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_inc(x_2); +x_27 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3(x_2, x_3, x_4, x_5, x_1, x_24, x_26, x_24, x_26, x_22, lean_box(0), lean_box(0), x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20); +lean_dec(x_24); +if (lean_obj_tag(x_27) == 0) +{ +lean_object* x_28; lean_object* x_29; +x_28 = lean_ctor_get(x_27, 0); +lean_inc(x_28); +x_29 = lean_ctor_get(x_28, 0); +lean_inc(x_29); +lean_dec(x_28); +if (lean_obj_tag(x_29) == 0) +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_30 = lean_ctor_get(x_27, 1); +lean_inc(x_30); +lean_dec(x_27); +x_31 = lean_box(0); +x_32 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__4(x_6, x_25, x_1, x_2, x_21, x_7, x_8, x_26, x_31, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_30); +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_21); +return x_32; +} +else +{ +uint8_t x_33; +lean_dec(x_21); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_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_2); +x_33 = !lean_is_exclusive(x_27); +if (x_33 == 0) +{ +lean_object* x_34; lean_object* x_35; +x_34 = lean_ctor_get(x_27, 0); +lean_dec(x_34); +x_35 = lean_ctor_get(x_29, 0); +lean_inc(x_35); +lean_dec(x_29); +lean_ctor_set(x_27, 0, x_35); +return x_27; +} +else +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_27, 1); +lean_inc(x_36); +lean_dec(x_27); +x_37 = lean_ctor_get(x_29, 0); +lean_inc(x_37); +lean_dec(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_36); +return x_38; +} +} +} +else +{ +uint8_t x_39; +lean_dec(x_21); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_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_2); +x_39 = !lean_is_exclusive(x_27); +if (x_39 == 0) +{ +return x_27; +} +else +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_40 = lean_ctor_get(x_27, 0); +x_41 = lean_ctor_get(x_27, 1); +lean_inc(x_41); +lean_inc(x_40); +lean_dec(x_27); +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_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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) { +_start: +{ +lean_object* x_13; lean_object* x_14; +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; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("assertion violation: ", 21, 21); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("c.assignment.size == numParams\n ", 33, 33); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__2; +x_3 = lean_string_append(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("_private.Lean.Meta.Tactic.Grind.EMatch.0.Lean.Meta.Grind.EMatch.instantiateTheorem", 82, 82); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__4; +x_3 = lean_unsigned_to_nat(172u); +x_4 = lean_unsigned_to_nat(2u); +x_5 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__3; +x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); +return x_6; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___boxed), 12, 0); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("unexpected number of parameters at ", 35, 35); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__7; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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_object* x_15) { +_start: +{ +lean_object* x_16; +lean_inc(x_1); +x_16 = l_Lean_Meta_Grind_EMatchTheorem_getProofWithFreshMVarLevels(x_1, x_11, x_12, x_13, x_14, x_15); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +x_18 = lean_ctor_get(x_16, 1); +lean_inc(x_18); +lean_dec(x_16); +x_19 = lean_ctor_get(x_1, 2); +lean_inc(x_19); +x_20 = lean_array_get_size(x_2); +x_21 = lean_nat_dec_eq(x_20, x_19); +if (x_21 == 0) +{ +lean_object* x_22; lean_object* x_23; +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_17); +lean_dec(x_3); +lean_dec(x_1); +x_22 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__5; +x_23 = l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__2(x_22, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_18); +return x_23; +} +else +{ +lean_object* x_24; +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_17); +x_24 = lean_infer_type(x_17, x_11, x_12, x_13, x_14, x_18); +if (lean_obj_tag(x_24) == 0) +{ +lean_object* x_25; lean_object* x_26; uint8_t x_27; lean_object* x_28; +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 = 0; +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_19); +x_28 = l_Lean_Meta_forallMetaBoundedTelescope(x_25, x_19, x_27, x_11, x_12, x_13, x_14, x_26); +if (lean_obj_tag(x_28) == 0) +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; +x_29 = lean_ctor_get(x_28, 0); +lean_inc(x_29); +x_30 = lean_ctor_get(x_29, 1); +lean_inc(x_30); +x_31 = lean_ctor_get(x_28, 1); +lean_inc(x_31); +lean_dec(x_28); +x_32 = lean_ctor_get(x_29, 0); +lean_inc(x_32); +lean_dec(x_29); +x_33 = !lean_is_exclusive(x_30); +if (x_33 == 0) +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; +x_34 = lean_ctor_get(x_30, 0); +x_35 = lean_ctor_get(x_30, 1); +lean_dec(x_35); +x_36 = lean_array_get_size(x_32); +x_37 = lean_nat_dec_eq(x_36, x_19); +lean_dec(x_36); +if (x_37 == 0) +{ +lean_object* x_38; lean_object* x_39; uint8_t x_40; +lean_dec(x_34); +lean_dec(x_32); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_17); +lean_dec(x_3); +x_38 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__2; +x_39 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__2(x_38, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_31); +x_40 = !lean_is_exclusive(x_39); +if (x_40 == 0) +{ +lean_object* x_41; lean_object* x_42; lean_object* x_43; uint8_t x_44; +x_41 = lean_ctor_get(x_39, 0); +x_42 = lean_ctor_get(x_39, 1); +x_43 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__6; +x_44 = lean_unbox(x_41); +lean_dec(x_41); +if (x_44 == 0) +{ +lean_object* x_45; lean_object* x_46; +lean_free_object(x_39); +lean_free_object(x_30); +lean_dec(x_1); +x_45 = lean_box(0); +x_46 = lean_apply_12(x_43, x_45, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_42); +return x_46; +} +else +{ +lean_object* x_47; lean_object* x_48; +x_47 = lean_ctor_get(x_1, 5); +lean_inc(x_47); +lean_dec(x_1); +x_48 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_47, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_42); +if (lean_obj_tag(x_48) == 0) +{ +lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; +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___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__8; +lean_ctor_set_tag(x_39, 7); +lean_ctor_set(x_39, 1, x_49); +lean_ctor_set(x_39, 0, x_51); +x_52 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; +lean_ctor_set_tag(x_30, 7); +lean_ctor_set(x_30, 1, x_52); +lean_ctor_set(x_30, 0, x_39); +x_53 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_38, x_30, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, 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); +x_56 = lean_apply_12(x_43, x_54, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_55); +return x_56; +} +else +{ +uint8_t x_57; +lean_free_object(x_39); +lean_free_object(x_30); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +x_57 = !lean_is_exclusive(x_48); +if (x_57 == 0) +{ +return x_48; +} +else +{ +lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_58 = lean_ctor_get(x_48, 0); +x_59 = lean_ctor_get(x_48, 1); +lean_inc(x_59); +lean_inc(x_58); +lean_dec(x_48); +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; +x_61 = lean_ctor_get(x_39, 0); +x_62 = lean_ctor_get(x_39, 1); +lean_inc(x_62); +lean_inc(x_61); +lean_dec(x_39); +x_63 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__6; +x_64 = lean_unbox(x_61); +lean_dec(x_61); +if (x_64 == 0) +{ +lean_object* x_65; lean_object* x_66; +lean_free_object(x_30); +lean_dec(x_1); +x_65 = lean_box(0); +x_66 = lean_apply_12(x_63, x_65, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_62); +return x_66; +} +else +{ +lean_object* x_67; lean_object* x_68; +x_67 = lean_ctor_get(x_1, 5); +lean_inc(x_67); +lean_dec(x_1); +x_68 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_67, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_62); +if (lean_obj_tag(x_68) == 0) +{ +lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; +x_69 = lean_ctor_get(x_68, 0); +lean_inc(x_69); +x_70 = lean_ctor_get(x_68, 1); +lean_inc(x_70); +lean_dec(x_68); +x_71 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__8; +x_72 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_72, 0, x_71); +lean_ctor_set(x_72, 1, x_69); +x_73 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; +lean_ctor_set_tag(x_30, 7); +lean_ctor_set(x_30, 1, x_73); +lean_ctor_set(x_30, 0, x_72); +x_74 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_38, x_30, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, 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_apply_12(x_63, x_75, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_76); +return x_77; +} +else +{ +lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; +lean_free_object(x_30); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +x_78 = lean_ctor_get(x_68, 0); +lean_inc(x_78); +x_79 = lean_ctor_get(x_68, 1); +lean_inc(x_79); +if (lean_is_exclusive(x_68)) { + lean_ctor_release(x_68, 0); + lean_ctor_release(x_68, 1); + x_80 = x_68; +} else { + lean_dec_ref(x_68); + x_80 = lean_box(0); +} +if (lean_is_scalar(x_80)) { + x_81 = lean_alloc_ctor(1, 2, 0); +} else { + x_81 = x_80; +} +lean_ctor_set(x_81, 0, x_78); +lean_ctor_set(x_81, 1, x_79); +return x_81; +} +} +} +} +else +{ +lean_object* x_82; lean_object* x_83; +lean_free_object(x_30); +x_82 = lean_box(0); +x_83 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__5(x_32, x_1, x_2, x_19, x_20, x_34, x_17, x_3, x_82, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_31); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_32); +return x_83; +} +} +else +{ +lean_object* x_84; lean_object* x_85; uint8_t x_86; +x_84 = lean_ctor_get(x_30, 0); +lean_inc(x_84); +lean_dec(x_30); +x_85 = lean_array_get_size(x_32); +x_86 = lean_nat_dec_eq(x_85, x_19); +lean_dec(x_85); +if (x_86 == 0) +{ +lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; uint8_t x_93; +lean_dec(x_84); +lean_dec(x_32); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_17); +lean_dec(x_3); +x_87 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__2; +x_88 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__2(x_87, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_31); +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 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__6; +x_93 = lean_unbox(x_89); +lean_dec(x_89); +if (x_93 == 0) +{ +lean_object* x_94; lean_object* x_95; +lean_dec(x_91); +lean_dec(x_1); +x_94 = lean_box(0); +x_95 = lean_apply_12(x_92, x_94, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_90); +return x_95; +} +else +{ +lean_object* x_96; lean_object* x_97; +x_96 = lean_ctor_get(x_1, 5); +lean_inc(x_96); +lean_dec(x_1); +x_97 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_96, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_90); +if (lean_obj_tag(x_97) == 0) +{ +lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; +x_98 = lean_ctor_get(x_97, 0); +lean_inc(x_98); +x_99 = lean_ctor_get(x_97, 1); +lean_inc(x_99); +lean_dec(x_97); +x_100 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__8; +if (lean_is_scalar(x_91)) { + x_101 = lean_alloc_ctor(7, 2, 0); +} else { + x_101 = x_91; + lean_ctor_set_tag(x_101, 7); +} +lean_ctor_set(x_101, 0, x_100); +lean_ctor_set(x_101, 1, x_98); +x_102 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; +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_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_87, x_103, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_99); +x_105 = lean_ctor_get(x_104, 0); +lean_inc(x_105); +x_106 = lean_ctor_get(x_104, 1); +lean_inc(x_106); +lean_dec(x_104); +x_107 = lean_apply_12(x_92, x_105, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_106); +return x_107; +} +else +{ +lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; +lean_dec(x_91); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +x_108 = lean_ctor_get(x_97, 0); +lean_inc(x_108); +x_109 = lean_ctor_get(x_97, 1); +lean_inc(x_109); +if (lean_is_exclusive(x_97)) { + lean_ctor_release(x_97, 0); + lean_ctor_release(x_97, 1); + x_110 = x_97; +} else { + lean_dec_ref(x_97); + x_110 = lean_box(0); +} +if (lean_is_scalar(x_110)) { + x_111 = lean_alloc_ctor(1, 2, 0); +} else { + x_111 = x_110; +} +lean_ctor_set(x_111, 0, x_108); +lean_ctor_set(x_111, 1, x_109); +return x_111; +} +} +} +else +{ +lean_object* x_112; lean_object* x_113; +x_112 = lean_box(0); +x_113 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__5(x_32, x_1, x_2, x_19, x_20, x_84, x_17, x_3, x_112, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_31); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_32); +return x_113; +} +} +} +else +{ +uint8_t x_114; +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_17); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +x_114 = !lean_is_exclusive(x_28); +if (x_114 == 0) +{ +return x_28; +} +else +{ +lean_object* x_115; lean_object* x_116; lean_object* x_117; +x_115 = lean_ctor_get(x_28, 0); +x_116 = lean_ctor_get(x_28, 1); +lean_inc(x_116); +lean_inc(x_115); +lean_dec(x_28); +x_117 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_117, 0, x_115); +lean_ctor_set(x_117, 1, x_116); +return x_117; +} +} +} +else +{ +uint8_t x_118; +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_17); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +x_118 = !lean_is_exclusive(x_24); +if (x_118 == 0) +{ +return x_24; +} +else +{ +lean_object* x_119; lean_object* x_120; lean_object* x_121; +x_119 = lean_ctor_get(x_24, 0); +x_120 = lean_ctor_get(x_24, 1); +lean_inc(x_120); +lean_inc(x_119); +lean_dec(x_24); +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_122; +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +x_122 = !lean_is_exclusive(x_16); +if (x_122 == 0) +{ +return x_16; +} +else +{ +lean_object* x_123; lean_object* x_124; lean_object* x_125; +x_123 = lean_ctor_get(x_16, 0); +x_124 = lean_ctor_get(x_16, 1); +lean_inc(x_124); +lean_inc(x_123); +lean_dec(x_16); +x_125 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_125, 0, x_123); +lean_ctor_set(x_125, 1, x_124); +return x_125; +} +} +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__8___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("assignment", 10, 10); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__8___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_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__2; +x_3 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__3; +x_4 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__8___closed__1; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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) { +_start: +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; +x_16 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__8___closed__2; +x_17 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__2(x_16, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +x_19 = lean_unbox(x_18); +lean_dec(x_18); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_17, 1); +lean_inc(x_20); +lean_dec(x_17); +x_21 = lean_box(0); +x_22 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7(x_1, x_2, x_3, x_21, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_20); +lean_dec(x_2); +return x_22; +} +else +{ +uint8_t x_23; +x_23 = !lean_is_exclusive(x_17); +if (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_17, 1); +x_25 = lean_ctor_get(x_17, 0); +lean_dec(x_25); +x_26 = lean_ctor_get(x_1, 5); +lean_inc(x_26); +x_27 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_26, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_24); +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; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_28 = lean_ctor_get(x_27, 0); +lean_inc(x_28); +x_29 = lean_ctor_get(x_27, 1); +lean_inc(x_29); +lean_dec(x_27); +x_30 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; +lean_ctor_set_tag(x_17, 7); +lean_ctor_set(x_17, 1, x_28); +lean_ctor_set(x_17, 0, x_30); +x_31 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__6; +x_32 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_32, 0, x_17); +lean_ctor_set(x_32, 1, x_31); +lean_inc(x_2); +x_33 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData(x_2); +x_34 = lean_array_to_list(x_33); +x_35 = lean_box(0); +x_36 = l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__4(x_34, x_35); +x_37 = l_Lean_MessageData_ofList(x_36); +x_38 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_38, 0, x_32); +lean_ctor_set(x_38, 1, x_37); +x_39 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_30); +x_40 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_16, x_39, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_29); +x_41 = lean_ctor_get(x_40, 0); +lean_inc(x_41); +x_42 = lean_ctor_get(x_40, 1); +lean_inc(x_42); +lean_dec(x_40); +x_43 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7(x_1, x_2, x_3, x_41, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_42); +lean_dec(x_41); +lean_dec(x_2); +return x_43; +} +else +{ +uint8_t x_44; +lean_free_object(x_17); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_44 = !lean_is_exclusive(x_27); +if (x_44 == 0) +{ +return x_27; +} +else +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_45 = lean_ctor_get(x_27, 0); +x_46 = lean_ctor_get(x_27, 1); +lean_inc(x_46); +lean_inc(x_45); +lean_dec(x_27); +x_47 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_47, 0, x_45); +lean_ctor_set(x_47, 1, x_46); +return x_47; +} +} +} +else +{ +lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_48 = lean_ctor_get(x_17, 1); +lean_inc(x_48); +lean_dec(x_17); +x_49 = lean_ctor_get(x_1, 5); +lean_inc(x_49); +x_50 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_49, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_48); +if (lean_obj_tag(x_50) == 0) +{ +lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; 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_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_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; +x_54 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_54, 0, x_53); +lean_ctor_set(x_54, 1, x_51); +x_55 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__6; +x_56 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_56, 0, x_54); +lean_ctor_set(x_56, 1, x_55); +lean_inc(x_2); +x_57 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData(x_2); +x_58 = lean_array_to_list(x_57); +x_59 = lean_box(0); +x_60 = l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__4(x_58, x_59); +x_61 = l_Lean_MessageData_ofList(x_60); +x_62 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_62, 0, x_56); +lean_ctor_set(x_62, 1, x_61); +x_63 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_63, 0, x_62); +lean_ctor_set(x_63, 1, x_53); +x_64 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_16, x_63, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_52); +x_65 = lean_ctor_get(x_64, 0); +lean_inc(x_65); +x_66 = lean_ctor_get(x_64, 1); +lean_inc(x_66); +lean_dec(x_64); +x_67 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7(x_1, x_2, x_3, x_65, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_66); +lean_dec(x_65); +lean_dec(x_2); +return x_67; +} +else +{ +lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_68 = lean_ctor_get(x_50, 0); +lean_inc(x_68); +x_69 = lean_ctor_get(x_50, 1); +lean_inc(x_69); +if (lean_is_exclusive(x_50)) { + lean_ctor_release(x_50, 0); + lean_ctor_release(x_50, 1); + x_70 = x_50; +} else { + lean_dec_ref(x_50); + 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; +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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; uint8_t x_19; +x_14 = lean_ctor_get(x_2, 0); +lean_inc(x_14); +lean_dec(x_2); +x_15 = lean_ctor_get(x_14, 1); +lean_inc(x_15); +x_16 = lean_ctor_get(x_1, 2); +lean_inc(x_16); +lean_inc(x_16); +x_17 = l_Lean_Meta_Grind_markTheoremInstance(x_15, x_16, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +x_19 = lean_unbox(x_18); +lean_dec(x_18); +if (x_19 == 0) +{ +uint8_t x_20; +lean_dec(x_16); +lean_dec(x_14); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_20 = !lean_is_exclusive(x_17); +if (x_20 == 0) +{ +lean_object* x_21; lean_object* x_22; +x_21 = lean_ctor_get(x_17, 0); +lean_dec(x_21); +x_22 = lean_box(0); +lean_ctor_set(x_17, 0, x_22); +return x_17; +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_17, 1); +lean_inc(x_23); +lean_dec(x_17); +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_26; lean_object* x_27; lean_object* x_28; +x_26 = lean_ctor_get(x_17, 1); +lean_inc(x_26); +lean_dec(x_17); +x_27 = lean_box(0); +x_28 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__8(x_14, x_16, x_1, x_27, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_26); +return x_28; +} +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_ReaderT_read___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__1___boxed), 11, 0); +return x_1; +} +} +static uint64_t _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___closed__2() { +_start: +{ +uint8_t x_1; uint64_t x_2; +x_1 = 1; +x_2 = l_Lean_Meta_TransparencyMode_toUInt64(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_13 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__9), 13, 1); +lean_closure_set(x_13, 0, x_1); +x_14 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___closed__1; +x_15 = lean_alloc_closure((void*)(l_ReaderT_bind___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__11___rarg), 13, 2); +lean_closure_set(x_15, 0, x_14); +lean_closure_set(x_15, 1, x_13); +x_16 = !lean_is_exclusive(x_8); +if (x_16 == 0) +{ +lean_object* x_17; uint8_t x_18; +x_17 = lean_ctor_get(x_8, 0); +x_18 = !lean_is_exclusive(x_17); +if (x_18 == 0) +{ +uint64_t x_19; uint8_t x_20; uint64_t x_21; uint64_t x_22; uint64_t x_23; uint64_t x_24; uint64_t x_25; uint8_t x_26; lean_object* x_27; +x_19 = lean_ctor_get_uint64(x_8, sizeof(void*)*7); +x_20 = 1; +lean_ctor_set_uint8(x_17, 9, x_20); +x_21 = 2; +x_22 = lean_uint64_shift_right(x_19, x_21); +x_23 = lean_uint64_shift_left(x_22, x_21); +x_24 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___closed__2; +x_25 = lean_uint64_lor(x_23, x_24); +lean_ctor_set_uint64(x_8, sizeof(void*)*7, x_25); +x_26 = 0; +x_27 = l_Lean_Meta_withNewMCtxDepth___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__12___rarg(x_15, x_26, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_27) == 0) +{ +uint8_t x_28; +x_28 = !lean_is_exclusive(x_27); +if (x_28 == 0) +{ +return x_27; +} +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_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; +} +} +else +{ +uint8_t x_32; +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 +{ +uint64_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; 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; uint64_t x_55; uint64_t x_56; uint64_t x_57; uint64_t x_58; uint64_t x_59; uint8_t x_60; lean_object* x_61; +x_36 = lean_ctor_get_uint64(x_8, sizeof(void*)*7); +x_37 = lean_ctor_get_uint8(x_17, 0); +x_38 = lean_ctor_get_uint8(x_17, 1); +x_39 = lean_ctor_get_uint8(x_17, 2); +x_40 = lean_ctor_get_uint8(x_17, 3); +x_41 = lean_ctor_get_uint8(x_17, 4); +x_42 = lean_ctor_get_uint8(x_17, 5); +x_43 = lean_ctor_get_uint8(x_17, 6); +x_44 = lean_ctor_get_uint8(x_17, 7); +x_45 = lean_ctor_get_uint8(x_17, 8); +x_46 = lean_ctor_get_uint8(x_17, 10); +x_47 = lean_ctor_get_uint8(x_17, 11); +x_48 = lean_ctor_get_uint8(x_17, 12); +x_49 = lean_ctor_get_uint8(x_17, 13); +x_50 = lean_ctor_get_uint8(x_17, 14); +x_51 = lean_ctor_get_uint8(x_17, 15); +x_52 = lean_ctor_get_uint8(x_17, 16); +lean_dec(x_17); +x_53 = 1; +x_54 = lean_alloc_ctor(0, 0, 17); +lean_ctor_set_uint8(x_54, 0, x_37); +lean_ctor_set_uint8(x_54, 1, x_38); +lean_ctor_set_uint8(x_54, 2, x_39); +lean_ctor_set_uint8(x_54, 3, x_40); +lean_ctor_set_uint8(x_54, 4, x_41); +lean_ctor_set_uint8(x_54, 5, x_42); +lean_ctor_set_uint8(x_54, 6, x_43); +lean_ctor_set_uint8(x_54, 7, x_44); +lean_ctor_set_uint8(x_54, 8, x_45); +lean_ctor_set_uint8(x_54, 9, x_53); +lean_ctor_set_uint8(x_54, 10, x_46); +lean_ctor_set_uint8(x_54, 11, x_47); +lean_ctor_set_uint8(x_54, 12, x_48); +lean_ctor_set_uint8(x_54, 13, x_49); +lean_ctor_set_uint8(x_54, 14, x_50); +lean_ctor_set_uint8(x_54, 15, x_51); +lean_ctor_set_uint8(x_54, 16, x_52); +x_55 = 2; +x_56 = lean_uint64_shift_right(x_36, x_55); +x_57 = lean_uint64_shift_left(x_56, x_55); +x_58 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___closed__2; +x_59 = lean_uint64_lor(x_57, x_58); +lean_ctor_set(x_8, 0, x_54); +lean_ctor_set_uint64(x_8, sizeof(void*)*7, x_59); +x_60 = 0; +x_61 = l_Lean_Meta_withNewMCtxDepth___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__12___rarg(x_15, x_60, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_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_61, 0); +lean_inc(x_66); +x_67 = lean_ctor_get(x_61, 1); +lean_inc(x_67); +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_61); + x_68 = lean_box(0); +} +if (lean_is_scalar(x_68)) { + x_69 = lean_alloc_ctor(1, 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; uint64_t x_71; 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; 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; 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; lean_object* x_97; uint8_t x_98; lean_object* x_99; uint64_t x_100; uint64_t x_101; uint64_t x_102; uint64_t x_103; uint64_t x_104; lean_object* x_105; uint8_t x_106; lean_object* x_107; +x_70 = lean_ctor_get(x_8, 0); +x_71 = lean_ctor_get_uint64(x_8, sizeof(void*)*7); +x_72 = lean_ctor_get_uint8(x_8, sizeof(void*)*7 + 8); +x_73 = lean_ctor_get(x_8, 1); +x_74 = lean_ctor_get(x_8, 2); +x_75 = lean_ctor_get(x_8, 3); +x_76 = lean_ctor_get(x_8, 4); +x_77 = lean_ctor_get(x_8, 5); +x_78 = lean_ctor_get(x_8, 6); +x_79 = lean_ctor_get_uint8(x_8, sizeof(void*)*7 + 9); +x_80 = lean_ctor_get_uint8(x_8, sizeof(void*)*7 + 10); +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_70); +lean_dec(x_8); +x_81 = lean_ctor_get_uint8(x_70, 0); +x_82 = lean_ctor_get_uint8(x_70, 1); +x_83 = lean_ctor_get_uint8(x_70, 2); +x_84 = lean_ctor_get_uint8(x_70, 3); +x_85 = lean_ctor_get_uint8(x_70, 4); +x_86 = lean_ctor_get_uint8(x_70, 5); +x_87 = lean_ctor_get_uint8(x_70, 6); +x_88 = lean_ctor_get_uint8(x_70, 7); +x_89 = lean_ctor_get_uint8(x_70, 8); +x_90 = lean_ctor_get_uint8(x_70, 10); +x_91 = lean_ctor_get_uint8(x_70, 11); +x_92 = lean_ctor_get_uint8(x_70, 12); +x_93 = lean_ctor_get_uint8(x_70, 13); +x_94 = lean_ctor_get_uint8(x_70, 14); +x_95 = lean_ctor_get_uint8(x_70, 15); +x_96 = lean_ctor_get_uint8(x_70, 16); +if (lean_is_exclusive(x_70)) { + x_97 = x_70; +} else { + lean_dec_ref(x_70); + x_97 = lean_box(0); +} +x_98 = 1; +if (lean_is_scalar(x_97)) { + x_99 = lean_alloc_ctor(0, 0, 17); +} else { + x_99 = x_97; +} +lean_ctor_set_uint8(x_99, 0, x_81); +lean_ctor_set_uint8(x_99, 1, x_82); +lean_ctor_set_uint8(x_99, 2, x_83); +lean_ctor_set_uint8(x_99, 3, x_84); +lean_ctor_set_uint8(x_99, 4, x_85); +lean_ctor_set_uint8(x_99, 5, x_86); +lean_ctor_set_uint8(x_99, 6, x_87); +lean_ctor_set_uint8(x_99, 7, x_88); +lean_ctor_set_uint8(x_99, 8, x_89); +lean_ctor_set_uint8(x_99, 9, x_98); +lean_ctor_set_uint8(x_99, 10, x_90); +lean_ctor_set_uint8(x_99, 11, x_91); +lean_ctor_set_uint8(x_99, 12, x_92); +lean_ctor_set_uint8(x_99, 13, x_93); +lean_ctor_set_uint8(x_99, 14, x_94); +lean_ctor_set_uint8(x_99, 15, x_95); +lean_ctor_set_uint8(x_99, 16, x_96); +x_100 = 2; +x_101 = lean_uint64_shift_right(x_71, x_100); +x_102 = lean_uint64_shift_left(x_101, x_100); +x_103 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___closed__2; +x_104 = lean_uint64_lor(x_102, x_103); +x_105 = lean_alloc_ctor(0, 7, 11); +lean_ctor_set(x_105, 0, x_99); +lean_ctor_set(x_105, 1, x_73); +lean_ctor_set(x_105, 2, x_74); +lean_ctor_set(x_105, 3, x_75); +lean_ctor_set(x_105, 4, x_76); +lean_ctor_set(x_105, 5, x_77); +lean_ctor_set(x_105, 6, x_78); +lean_ctor_set_uint64(x_105, sizeof(void*)*7, x_104); +lean_ctor_set_uint8(x_105, sizeof(void*)*7 + 8, x_72); +lean_ctor_set_uint8(x_105, sizeof(void*)*7 + 9, x_79); +lean_ctor_set_uint8(x_105, sizeof(void*)*7 + 10, x_80); +x_106 = 0; +x_107 = l_Lean_Meta_withNewMCtxDepth___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__12___rarg(x_15, x_106, x_2, x_3, x_4, x_5, x_6, x_7, x_105, x_9, x_10, x_11, x_12); +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); +if (lean_is_exclusive(x_107)) { + lean_ctor_release(x_107, 0); + lean_ctor_release(x_107, 1); + x_110 = x_107; +} else { + lean_dec_ref(x_107); + x_110 = lean_box(0); +} +if (lean_is_scalar(x_110)) { + x_111 = lean_alloc_ctor(0, 2, 0); +} else { + x_111 = x_110; +} +lean_ctor_set(x_111, 0, x_108); +lean_ctor_set(x_111, 1, x_109); +return x_111; +} +else +{ +lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; +x_112 = lean_ctor_get(x_107, 0); +lean_inc(x_112); +x_113 = lean_ctor_get(x_107, 1); +lean_inc(x_113); +if (lean_is_exclusive(x_107)) { + lean_ctor_release(x_107, 0); + lean_ctor_release(x_107, 1); + x_114 = x_107; +} else { + lean_dec_ref(x_107); + x_114 = lean_box(0); +} +if (lean_is_scalar(x_114)) { + x_115 = lean_alloc_ctor(1, 2, 0); +} else { + x_115 = x_114; +} +lean_ctor_set(x_115, 0, x_112); +lean_ctor_set(x_115, 1, x_113); +return x_115; +} +} +} +} +LEAN_EXPORT lean_object* l_ReaderT_read___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; +x_12 = l_ReaderT_read___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +x_13 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +lean_object* x_18 = _args[17]; +lean_object* x_19 = _args[18]; +lean_object* x_20 = _args[19]; +lean_object* x_21 = _args[20]; +lean_object* x_22 = _args[21]; +lean_object* x_23 = _args[22]; +_start: +{ +lean_object* x_24; +x_24 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_23); +lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_24; +} +} +LEAN_EXPORT lean_object* l_Lean_MVarId_isAssigned___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +x_13 = l_Lean_MVarId_isAssigned___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +lean_object* x_14; +x_14 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +lean_object* x_18 = _args[17]; +lean_object* x_19 = _args[18]; +_start: +{ +size_t x_20; size_t x_21; lean_object* x_22; +x_20 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_21 = lean_unbox_usize(x_7); +lean_dec(x_7); +x_22 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5(x_1, x_2, x_3, x_4, x_5, x_20, x_21, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +return x_22; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__6___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +lean_object* x_18 = _args[17]; +_start: +{ +size_t x_19; size_t x_20; lean_object* x_21; +x_19 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_20 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_21 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__6(x_1, x_2, x_3, x_4, x_19, x_20, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +return x_21; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__7___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +lean_object* x_18 = _args[17]; +_start: +{ +size_t x_19; size_t x_20; lean_object* x_21; +x_19 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_20 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_21 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__7(x_1, x_2, x_3, x_4, x_19, x_20, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +return x_21; +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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) { +_start: +{ +size_t x_16; size_t x_17; lean_object* x_18; +x_16 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_17 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_18 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__8(x_1, x_16, x_17, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_1); +return x_18; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__9___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +lean_object* x_18 = _args[17]; +_start: +{ +size_t x_19; size_t x_20; lean_object* x_21; +x_19 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_20 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_21 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__9(x_1, x_2, x_3, x_4, x_19, x_20, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +return x_21; +} +} +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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) { +_start: +{ +size_t x_15; size_t x_16; lean_object* x_17; +x_15 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_16 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_17 = l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__10(x_1, x_15, x_16, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +return x_17; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_withNewMCtxDepth___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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, 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_2); +lean_dec(x_2); +x_15 = l_Lean_Meta_withNewMCtxDepth___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__12___rarg(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); +return x_15; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +_start: +{ +lean_object* x_17; +x_17 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_2); +return x_17; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +lean_object* x_14; +x_14 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_14; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +lean_object* x_18 = _args[17]; +lean_object* x_19 = _args[18]; +lean_object* x_20 = _args[19]; +lean_object* x_21 = _args[20]; +_start: +{ +size_t x_22; lean_object* x_23; +x_22 = lean_unbox_usize(x_8); +lean_dec(x_8); +x_23 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_22, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_6); +lean_dec(x_3); +lean_dec(x_1); +return x_23; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__4___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_object* x_19 = _args[18]; +lean_object* x_20 = _args[19]; +_start: +{ +lean_object* x_21; +x_21 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20); +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_5); +lean_dec(x_3); +return x_21; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__5___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +lean_object* x_18 = _args[17]; +lean_object* x_19 = _args[18]; +lean_object* x_20 = _args[19]; +_start: +{ +lean_object* x_21; +x_21 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__5(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20); +lean_dec(x_9); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +return x_21; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +x_13 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_13; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +_start: +{ +lean_object* x_16; +x_16 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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); +lean_dec(x_4); +lean_dec(x_2); +return x_16; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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, lean_object* x_15) { +_start: +{ +lean_object* x_16; +x_16 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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); +lean_dec(x_4); +return x_16; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +x_13 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_13; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___lambda__2___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___lambda__1___boxed), 12, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; 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_13 = lean_st_ref_take(x_3, 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_Meta_Grind_EMatch_instInhabitedChoice; +x_17 = l_List_head_x21___rarg(x_16, x_14); +x_18 = l_List_tail_x21___rarg(x_14); +lean_dec(x_14); +x_19 = lean_st_ref_set(x_3, x_18, x_15); +x_20 = lean_ctor_get(x_19, 1); +lean_inc(x_20); +lean_dec(x_19); +x_21 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___lambda__2___closed__1; +x_22 = lean_ctor_get(x_17, 0); +lean_inc(x_22); +if (lean_obj_tag(x_22) == 0) +{ +lean_object* x_23; +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +x_23 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem(x_17, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_20); +if (lean_obj_tag(x_23) == 0) +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +x_25 = lean_ctor_get(x_23, 1); +lean_inc(x_25); +lean_dec(x_23); +x_26 = lean_apply_12(x_21, x_24, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_25); +return x_26; +} +else +{ +uint8_t x_27; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_27 = !lean_is_exclusive(x_23); +if (x_27 == 0) +{ +return x_23; +} +else +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_28 = lean_ctor_get(x_23, 0); +x_29 = lean_ctor_get(x_23, 1); +lean_inc(x_29); +lean_inc(x_28); +lean_dec(x_23); +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; +x_31 = lean_ctor_get(x_22, 0); +lean_inc(x_31); +if (lean_obj_tag(x_31) == 0) +{ +uint8_t x_32; +x_32 = !lean_is_exclusive(x_17); +if (x_32 == 0) +{ +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_17, 0); +lean_dec(x_33); +x_34 = lean_ctor_get(x_22, 1); +lean_inc(x_34); +lean_dec(x_22); +x_35 = lean_ctor_get(x_31, 0); +lean_inc(x_35); +x_36 = lean_ctor_get(x_31, 1); +lean_inc(x_36); +lean_dec(x_31); +lean_ctor_set(x_17, 0, x_34); +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_37 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch(x_17, x_35, x_36, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_20); +lean_dec(x_35); +if (lean_obj_tag(x_37) == 0) +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_38 = lean_ctor_get(x_37, 0); +lean_inc(x_38); +x_39 = lean_ctor_get(x_37, 1); +lean_inc(x_39); +lean_dec(x_37); +x_40 = lean_apply_12(x_21, x_38, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_39); +return x_40; +} +else +{ +uint8_t x_41; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_41 = !lean_is_exclusive(x_37); +if (x_41 == 0) +{ +return x_37; +} +else +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_42 = lean_ctor_get(x_37, 0); +x_43 = lean_ctor_get(x_37, 1); +lean_inc(x_43); +lean_inc(x_42); +lean_dec(x_37); +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_45; 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_45 = lean_ctor_get(x_17, 1); +x_46 = lean_ctor_get(x_17, 2); +lean_inc(x_46); +lean_inc(x_45); +lean_dec(x_17); +x_47 = lean_ctor_get(x_22, 1); +lean_inc(x_47); +lean_dec(x_22); +x_48 = lean_ctor_get(x_31, 0); +lean_inc(x_48); +x_49 = lean_ctor_get(x_31, 1); +lean_inc(x_49); +lean_dec(x_31); +x_50 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_50, 0, x_47); +lean_ctor_set(x_50, 1, x_45); +lean_ctor_set(x_50, 2, x_46); +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_51 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch(x_50, x_48, x_49, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_20); +lean_dec(x_48); +if (lean_obj_tag(x_51) == 0) +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; +x_52 = lean_ctor_get(x_51, 0); +lean_inc(x_52); +x_53 = lean_ctor_get(x_51, 1); +lean_inc(x_53); +lean_dec(x_51); +x_54 = lean_apply_12(x_21, x_52, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_53); +return x_54; +} +else +{ +lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; +lean_dec(x_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_55 = lean_ctor_get(x_51, 0); +lean_inc(x_55); +x_56 = lean_ctor_get(x_51, 1); +lean_inc(x_56); +if (lean_is_exclusive(x_51)) { + lean_ctor_release(x_51, 0); + lean_ctor_release(x_51, 1); + x_57 = x_51; +} else { + lean_dec_ref(x_51); + x_57 = lean_box(0); +} +if (lean_is_scalar(x_57)) { + x_58 = lean_alloc_ctor(1, 2, 0); +} else { + x_58 = x_57; +} +lean_ctor_set(x_58, 0, x_55); +lean_ctor_set(x_58, 1, x_56); +return x_58; +} +} +} +else +{ +uint8_t x_59; +x_59 = !lean_is_exclusive(x_17); +if (x_59 == 0) +{ +lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; +x_60 = lean_ctor_get(x_17, 0); +lean_dec(x_60); +x_61 = lean_ctor_get(x_22, 1); +lean_inc(x_61); +lean_dec(x_22); +x_62 = lean_ctor_get(x_31, 0); +lean_inc(x_62); +lean_dec(x_31); +lean_ctor_set(x_17, 0, x_61); +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_63 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue(x_17, x_62, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_20); +if (lean_obj_tag(x_63) == 0) +{ +lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_64 = lean_ctor_get(x_63, 0); +lean_inc(x_64); +x_65 = lean_ctor_get(x_63, 1); +lean_inc(x_65); +lean_dec(x_63); +x_66 = lean_apply_12(x_21, x_64, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_65); +return x_66; +} +else +{ +uint8_t x_67; +lean_dec(x_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_67 = !lean_is_exclusive(x_63); +if (x_67 == 0) +{ +return x_63; +} +else +{ +lean_object* x_68; lean_object* x_69; lean_object* x_70; +x_68 = lean_ctor_get(x_63, 0); +x_69 = lean_ctor_get(x_63, 1); +lean_inc(x_69); +lean_inc(x_68); +lean_dec(x_63); +x_70 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_70, 0, x_68); +lean_ctor_set(x_70, 1, x_69); +return x_70; +} +} +} +else +{ +lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; +x_71 = lean_ctor_get(x_17, 1); +x_72 = lean_ctor_get(x_17, 2); +lean_inc(x_72); +lean_inc(x_71); +lean_dec(x_17); +x_73 = lean_ctor_get(x_22, 1); +lean_inc(x_73); +lean_dec(x_22); +x_74 = lean_ctor_get(x_31, 0); +lean_inc(x_74); +lean_dec(x_31); +x_75 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_75, 0, x_73); +lean_ctor_set(x_75, 1, x_71); +lean_ctor_set(x_75, 2, x_72); +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_76 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue(x_75, x_74, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_20); +if (lean_obj_tag(x_76) == 0) +{ +lean_object* x_77; lean_object* x_78; lean_object* x_79; +x_77 = lean_ctor_get(x_76, 0); +lean_inc(x_77); +x_78 = lean_ctor_get(x_76, 1); +lean_inc(x_78); +lean_dec(x_76); +x_79 = lean_apply_12(x_21, x_77, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_78); +return x_79; +} +else +{ +lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; +lean_dec(x_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_80 = lean_ctor_get(x_76, 0); +lean_inc(x_80); +x_81 = lean_ctor_get(x_76, 1); +lean_inc(x_81); +if (lean_is_exclusive(x_76)) { + lean_ctor_release(x_76, 0); + lean_ctor_release(x_76, 1); + x_82 = x_76; +} else { + lean_dec_ref(x_76); + x_82 = lean_box(0); +} +if (lean_is_scalar(x_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; +} +} +} +} +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___lambda__2___boxed), 12, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, 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_st_ref_get(x_2, x_11); +x_13 = !lean_is_exclusive(x_12); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_14 = lean_ctor_get(x_12, 0); +x_15 = lean_ctor_get(x_12, 1); +x_16 = l_List_isEmpty___rarg(x_14); +lean_dec(x_14); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; +lean_free_object(x_12); +x_17 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__2; +x_18 = l_Lean_Core_checkSystem(x_17, x_9, x_10, x_15); +if (lean_obj_tag(x_18) == 0) +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_19 = lean_ctor_get(x_18, 1); +lean_inc(x_19); +lean_dec(x_18); +x_20 = l_Lean_Meta_Grind_checkMaxInstancesExceeded(x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_19); +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; +x_23 = lean_ctor_get(x_20, 1); +lean_inc(x_23); +lean_dec(x_20); +x_24 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___closed__1; +x_25 = lean_box(0); +x_26 = lean_apply_12(x_24, x_25, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_23); +return x_26; +} +else +{ +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_3); +lean_dec(x_2); +lean_dec(x_1); +x_27 = !lean_is_exclusive(x_20); +if (x_27 == 0) +{ +lean_object* x_28; lean_object* x_29; +x_28 = lean_ctor_get(x_20, 0); +lean_dec(x_28); +x_29 = lean_box(0); +lean_ctor_set(x_20, 0, x_29); +return x_20; +} +else +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_30 = lean_ctor_get(x_20, 1); +lean_inc(x_30); +lean_dec(x_20); +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_33; +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_33 = !lean_is_exclusive(x_18); +if (x_33 == 0) +{ +return x_18; +} +else +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_18, 0); +x_35 = lean_ctor_get(x_18, 1); +lean_inc(x_35); +lean_inc(x_34); +lean_dec(x_18); +x_36 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_36, 0, x_34); +lean_ctor_set(x_36, 1, x_35); +return x_36; +} +} +} +else +{ +lean_object* x_37; +lean_dec(x_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_37 = lean_box(0); +lean_ctor_set(x_12, 0, x_37); +return x_12; +} +} +else +{ +lean_object* x_38; lean_object* x_39; uint8_t 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 = l_List_isEmpty___rarg(x_38); +lean_dec(x_38); +if (x_40 == 0) +{ +lean_object* x_41; lean_object* x_42; +x_41 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__2; +x_42 = l_Lean_Core_checkSystem(x_41, x_9, x_10, x_39); +if (lean_obj_tag(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_42, 1); +lean_inc(x_43); +lean_dec(x_42); +x_44 = l_Lean_Meta_Grind_checkMaxInstancesExceeded(x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_43); +x_45 = lean_ctor_get(x_44, 0); +lean_inc(x_45); +x_46 = lean_unbox(x_45); +lean_dec(x_45); +if (x_46 == 0) +{ +lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_47 = lean_ctor_get(x_44, 1); +lean_inc(x_47); +lean_dec(x_44); +x_48 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___closed__1; +x_49 = lean_box(0); +x_50 = lean_apply_12(x_48, x_49, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_47); +return x_50; +} +else +{ +lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; +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_51 = lean_ctor_get(x_44, 1); +lean_inc(x_51); +if (lean_is_exclusive(x_44)) { + lean_ctor_release(x_44, 0); + lean_ctor_release(x_44, 1); + x_52 = x_44; +} else { + lean_dec_ref(x_44); + x_52 = lean_box(0); +} +x_53 = lean_box(0); +if (lean_is_scalar(x_52)) { + x_54 = lean_alloc_ctor(0, 2, 0); +} else { + x_54 = x_52; +} +lean_ctor_set(x_54, 0, x_53); +lean_ctor_set(x_54, 1, x_51); +return x_54; +} +} +else +{ +lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_55 = lean_ctor_get(x_42, 0); +lean_inc(x_55); +x_56 = lean_ctor_get(x_42, 1); +lean_inc(x_56); +if (lean_is_exclusive(x_42)) { + lean_ctor_release(x_42, 0); + lean_ctor_release(x_42, 1); + x_57 = x_42; +} else { + lean_dec_ref(x_42); + x_57 = lean_box(0); +} +if (lean_is_scalar(x_57)) { + x_58 = lean_alloc_ctor(1, 2, 0); +} else { + x_58 = x_57; +} +lean_ctor_set(x_58, 0, x_55); +lean_ctor_set(x_58, 1, x_56); +return x_58; +} +} +else +{ +lean_object* x_59; lean_object* x_60; +lean_dec(x_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_59 = lean_box(0); +x_60 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_60, 0, x_59); +lean_ctor_set(x_60, 1, x_39); +return x_60; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___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_13; +x_13 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___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_1); +return x_13; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +x_13 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_1); +return x_13; +} +} +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_main___spec__1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, uint8_t x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19) { +_start: +{ +lean_object* x_20; +lean_inc(x_1); +x_20 = l_Lean_Meta_Grind_getENode(x_1, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19); +if (lean_obj_tag(x_20) == 0) +{ +uint8_t x_21; +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_89; +x_22 = lean_ctor_get(x_20, 0); +x_23 = lean_ctor_get(x_20, 1); +x_89 = lean_ctor_get_uint8(x_22, sizeof(void*)*10 + 4); +if (x_89 == 0) +{ +lean_object* x_90; uint8_t x_91; +x_90 = lean_ctor_get(x_22, 3); +lean_inc(x_90); +x_91 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_90, x_1); +lean_dec(x_90); +if (x_91 == 0) +{ +lean_object* x_92; +lean_dec(x_22); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_92 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_92, 0, x_5); +lean_ctor_set(x_20, 0, x_92); +return x_20; +} +else +{ +if (x_6 == 0) +{ +lean_object* x_93; +lean_free_object(x_20); +x_93 = lean_box(0); +x_24 = x_93; +goto block_88; +} +else +{ +lean_object* x_94; uint8_t x_95; +x_94 = lean_ctor_get(x_22, 9); +lean_inc(x_94); +x_95 = lean_nat_dec_eq(x_94, x_7); +lean_dec(x_94); +if (x_95 == 0) +{ +lean_object* x_96; +lean_dec(x_22); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_96 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_96, 0, x_5); +lean_ctor_set(x_20, 0, x_96); +return x_20; +} +else +{ +lean_object* x_97; +lean_free_object(x_20); +x_97 = lean_box(0); +x_24 = x_97; +goto block_88; +} +} +} +} +else +{ +if (x_6 == 0) +{ +lean_object* x_98; +lean_free_object(x_20); +x_98 = lean_box(0); +x_24 = x_98; +goto block_88; +} +else +{ +lean_object* x_99; uint8_t x_100; +x_99 = lean_ctor_get(x_22, 9); +lean_inc(x_99); +x_100 = lean_nat_dec_eq(x_99, x_7); +lean_dec(x_99); +if (x_100 == 0) +{ +lean_object* x_101; +lean_dec(x_22); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_101 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_101, 0, x_5); +lean_ctor_set(x_20, 0, x_101); +return x_20; +} +else +{ +lean_object* x_102; +lean_free_object(x_20); +x_102 = lean_box(0); +x_24 = x_102; +goto block_88; +} +} +} +block_88: +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; +lean_dec(x_24); +x_25 = lean_ctor_get(x_22, 8); +lean_inc(x_25); +lean_dec(x_22); +x_26 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_26, 0, x_2); +lean_ctor_set(x_26, 1, x_25); +lean_ctor_set(x_26, 2, x_3); +lean_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); +x_27 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_x3f(x_26, x_4, x_1, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_23); +lean_dec(x_1); +if (lean_obj_tag(x_27) == 0) +{ +lean_object* x_28; +x_28 = lean_ctor_get(x_27, 0); +lean_inc(x_28); +if (lean_obj_tag(x_28) == 0) +{ +uint8_t x_29; +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +x_29 = !lean_is_exclusive(x_27); +if (x_29 == 0) +{ +lean_object* x_30; lean_object* x_31; +x_30 = lean_ctor_get(x_27, 0); +lean_dec(x_30); +x_31 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_31, 0, x_5); +lean_ctor_set(x_27, 0, x_31); +return x_27; +} +else +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_27, 1); +lean_inc(x_32); +lean_dec(x_27); +x_33 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_33, 0, x_5); +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; uint8_t x_36; +x_35 = lean_ctor_get(x_27, 1); +lean_inc(x_35); +lean_dec(x_27); +x_36 = !lean_is_exclusive(x_28); +if (x_36 == 0) +{ +lean_object* x_37; lean_object* x_38; uint8_t x_39; +x_37 = lean_ctor_get(x_28, 0); +x_38 = lean_st_ref_take(x_10, x_35); +x_39 = !lean_is_exclusive(x_38); +if (x_39 == 0) +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_40 = lean_ctor_get(x_38, 1); +x_41 = lean_ctor_get(x_38, 0); +lean_dec(x_41); +x_42 = lean_box(0); +lean_ctor_set_tag(x_38, 1); +lean_ctor_set(x_38, 1, x_42); +lean_ctor_set(x_38, 0, x_37); +x_43 = lean_st_ref_set(x_10, x_38, x_40); +x_44 = lean_ctor_get(x_43, 1); +lean_inc(x_44); +lean_dec(x_43); +x_45 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices(x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_44); +if (lean_obj_tag(x_45) == 0) +{ +uint8_t x_46; +x_46 = !lean_is_exclusive(x_45); +if (x_46 == 0) +{ +lean_object* x_47; +x_47 = lean_ctor_get(x_45, 0); +lean_dec(x_47); +lean_ctor_set(x_28, 0, x_5); +lean_ctor_set(x_45, 0, x_28); +return x_45; +} +else +{ +lean_object* x_48; lean_object* x_49; +x_48 = lean_ctor_get(x_45, 1); +lean_inc(x_48); +lean_dec(x_45); +lean_ctor_set(x_28, 0, x_5); +x_49 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_49, 0, x_28); +lean_ctor_set(x_49, 1, x_48); +return x_49; +} +} +else +{ +uint8_t x_50; +lean_free_object(x_28); +lean_dec(x_5); +x_50 = !lean_is_exclusive(x_45); +if (x_50 == 0) +{ +return x_45; +} +else +{ +lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_51 = lean_ctor_get(x_45, 0); +x_52 = lean_ctor_get(x_45, 1); +lean_inc(x_52); +lean_inc(x_51); +lean_dec(x_45); +x_53 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_53, 0, x_51); +lean_ctor_set(x_53, 1, x_52); +return x_53; +} +} +} +else +{ +lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; +x_54 = lean_ctor_get(x_38, 1); +lean_inc(x_54); +lean_dec(x_38); +x_55 = lean_box(0); +x_56 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_56, 0, x_37); +lean_ctor_set(x_56, 1, x_55); +x_57 = lean_st_ref_set(x_10, x_56, x_54); +x_58 = lean_ctor_get(x_57, 1); +lean_inc(x_58); +lean_dec(x_57); +x_59 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices(x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_58); +if (lean_obj_tag(x_59) == 0) +{ +lean_object* x_60; lean_object* x_61; lean_object* x_62; +x_60 = lean_ctor_get(x_59, 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); +} +lean_ctor_set(x_28, 0, x_5); +if (lean_is_scalar(x_61)) { + x_62 = lean_alloc_ctor(0, 2, 0); +} else { + x_62 = x_61; +} +lean_ctor_set(x_62, 0, x_28); +lean_ctor_set(x_62, 1, x_60); +return x_62; +} +else +{ +lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; +lean_free_object(x_28); +lean_dec(x_5); +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); +} +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; +} +} +} +else +{ +lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; +x_67 = lean_ctor_get(x_28, 0); +lean_inc(x_67); +lean_dec(x_28); +x_68 = lean_st_ref_take(x_10, x_35); +x_69 = lean_ctor_get(x_68, 1); +lean_inc(x_69); +if (lean_is_exclusive(x_68)) { + lean_ctor_release(x_68, 0); + lean_ctor_release(x_68, 1); + x_70 = x_68; +} else { + lean_dec_ref(x_68); + x_70 = lean_box(0); +} +x_71 = lean_box(0); +if (lean_is_scalar(x_70)) { + x_72 = lean_alloc_ctor(1, 2, 0); +} else { + x_72 = x_70; + lean_ctor_set_tag(x_72, 1); +} +lean_ctor_set(x_72, 0, x_67); +lean_ctor_set(x_72, 1, x_71); +x_73 = lean_st_ref_set(x_10, x_72, x_69); +x_74 = lean_ctor_get(x_73, 1); +lean_inc(x_74); +lean_dec(x_73); +x_75 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices(x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_74); +if (lean_obj_tag(x_75) == 0) +{ +lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; +x_76 = lean_ctor_get(x_75, 1); +lean_inc(x_76); +if (lean_is_exclusive(x_75)) { + lean_ctor_release(x_75, 0); + lean_ctor_release(x_75, 1); + x_77 = x_75; +} else { + lean_dec_ref(x_75); + x_77 = lean_box(0); +} +x_78 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_78, 0, x_5); +if (lean_is_scalar(x_77)) { + x_79 = lean_alloc_ctor(0, 2, 0); +} else { + x_79 = x_77; +} +lean_ctor_set(x_79, 0, x_78); +lean_ctor_set(x_79, 1, x_76); +return x_79; +} +else +{ +lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; +lean_dec(x_5); +x_80 = lean_ctor_get(x_75, 0); +lean_inc(x_80); +x_81 = lean_ctor_get(x_75, 1); +lean_inc(x_81); +if (lean_is_exclusive(x_75)) { + lean_ctor_release(x_75, 0); + lean_ctor_release(x_75, 1); + x_82 = x_75; +} else { + lean_dec_ref(x_75); + 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 +{ +uint8_t x_84; +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_5); +x_84 = !lean_is_exclusive(x_27); +if (x_84 == 0) +{ +return x_27; +} +else +{ +lean_object* x_85; lean_object* x_86; lean_object* x_87; +x_85 = lean_ctor_get(x_27, 0); +x_86 = lean_ctor_get(x_27, 1); +lean_inc(x_86); +lean_inc(x_85); +lean_dec(x_27); +x_87 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_87, 0, x_85); +lean_ctor_set(x_87, 1, x_86); +return x_87; +} +} +} +} +else +{ +lean_object* x_103; lean_object* x_104; lean_object* x_105; uint8_t x_138; +x_103 = lean_ctor_get(x_20, 0); +x_104 = lean_ctor_get(x_20, 1); +lean_inc(x_104); +lean_inc(x_103); +lean_dec(x_20); +x_138 = lean_ctor_get_uint8(x_103, sizeof(void*)*10 + 4); +if (x_138 == 0) +{ +lean_object* x_139; uint8_t x_140; +x_139 = lean_ctor_get(x_103, 3); +lean_inc(x_139); +x_140 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_139, x_1); +lean_dec(x_139); +if (x_140 == 0) +{ +lean_object* x_141; lean_object* x_142; +lean_dec(x_103); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_141 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_141, 0, x_5); +x_142 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_142, 0, x_141); +lean_ctor_set(x_142, 1, x_104); +return x_142; +} +else +{ +if (x_6 == 0) +{ +lean_object* x_143; +x_143 = lean_box(0); +x_105 = x_143; +goto block_137; +} +else +{ +lean_object* x_144; uint8_t x_145; +x_144 = lean_ctor_get(x_103, 9); +lean_inc(x_144); +x_145 = lean_nat_dec_eq(x_144, x_7); +lean_dec(x_144); +if (x_145 == 0) +{ +lean_object* x_146; lean_object* x_147; +lean_dec(x_103); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_146 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_146, 0, x_5); +x_147 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_147, 0, x_146); +lean_ctor_set(x_147, 1, x_104); +return x_147; +} +else +{ +lean_object* x_148; +x_148 = lean_box(0); +x_105 = x_148; +goto block_137; +} +} +} +} +else +{ +if (x_6 == 0) +{ +lean_object* x_149; +x_149 = lean_box(0); +x_105 = x_149; +goto block_137; +} +else +{ +lean_object* x_150; uint8_t x_151; +x_150 = lean_ctor_get(x_103, 9); +lean_inc(x_150); +x_151 = lean_nat_dec_eq(x_150, x_7); +lean_dec(x_150); +if (x_151 == 0) +{ +lean_object* x_152; lean_object* x_153; +lean_dec(x_103); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_152 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_152, 0, x_5); +x_153 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_153, 0, x_152); +lean_ctor_set(x_153, 1, x_104); +return x_153; +} +else +{ +lean_object* x_154; +x_154 = lean_box(0); +x_105 = x_154; +goto block_137; +} +} +} +block_137: +{ +lean_object* x_106; lean_object* x_107; lean_object* x_108; +lean_dec(x_105); +x_106 = lean_ctor_get(x_103, 8); +lean_inc(x_106); +lean_dec(x_103); +x_107 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_107, 0, x_2); +lean_ctor_set(x_107, 1, x_106); +lean_ctor_set(x_107, 2, x_3); +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); +x_108 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_x3f(x_107, x_4, x_1, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_104); +lean_dec(x_1); +if (lean_obj_tag(x_108) == 0) +{ +lean_object* x_109; +x_109 = lean_ctor_get(x_108, 0); +lean_inc(x_109); +if (lean_obj_tag(x_109) == 0) +{ +lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +x_110 = lean_ctor_get(x_108, 1); +lean_inc(x_110); +if (lean_is_exclusive(x_108)) { + lean_ctor_release(x_108, 0); + lean_ctor_release(x_108, 1); + x_111 = x_108; +} else { + lean_dec_ref(x_108); + x_111 = lean_box(0); +} +x_112 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_112, 0, x_5); +if (lean_is_scalar(x_111)) { + x_113 = lean_alloc_ctor(0, 2, 0); +} else { + x_113 = x_111; +} +lean_ctor_set(x_113, 0, x_112); +lean_ctor_set(x_113, 1, x_110); +return x_113; +} +else +{ +lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; +x_114 = lean_ctor_get(x_108, 1); +lean_inc(x_114); +lean_dec(x_108); +x_115 = lean_ctor_get(x_109, 0); +lean_inc(x_115); +if (lean_is_exclusive(x_109)) { + lean_ctor_release(x_109, 0); + x_116 = x_109; +} else { + lean_dec_ref(x_109); + x_116 = lean_box(0); +} +x_117 = lean_st_ref_take(x_10, x_114); +x_118 = lean_ctor_get(x_117, 1); +lean_inc(x_118); +if (lean_is_exclusive(x_117)) { + lean_ctor_release(x_117, 0); + lean_ctor_release(x_117, 1); + x_119 = x_117; +} else { + lean_dec_ref(x_117); + x_119 = lean_box(0); +} +x_120 = lean_box(0); +if (lean_is_scalar(x_119)) { + x_121 = lean_alloc_ctor(1, 2, 0); +} else { + x_121 = x_119; + lean_ctor_set_tag(x_121, 1); +} +lean_ctor_set(x_121, 0, x_115); +lean_ctor_set(x_121, 1, x_120); +x_122 = lean_st_ref_set(x_10, x_121, x_118); +x_123 = lean_ctor_get(x_122, 1); +lean_inc(x_123); +lean_dec(x_122); +x_124 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices(x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_123); +if (lean_obj_tag(x_124) == 0) +{ +lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; +x_125 = lean_ctor_get(x_124, 1); +lean_inc(x_125); +if (lean_is_exclusive(x_124)) { + lean_ctor_release(x_124, 0); + lean_ctor_release(x_124, 1); + x_126 = x_124; +} else { + lean_dec_ref(x_124); + x_126 = lean_box(0); +} +if (lean_is_scalar(x_116)) { + x_127 = lean_alloc_ctor(1, 1, 0); +} else { + x_127 = x_116; +} +lean_ctor_set(x_127, 0, x_5); +if (lean_is_scalar(x_126)) { + x_128 = lean_alloc_ctor(0, 2, 0); +} else { + x_128 = x_126; +} +lean_ctor_set(x_128, 0, x_127); +lean_ctor_set(x_128, 1, x_125); +return x_128; +} +else +{ +lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; +lean_dec(x_116); +lean_dec(x_5); +x_129 = lean_ctor_get(x_124, 0); +lean_inc(x_129); +x_130 = lean_ctor_get(x_124, 1); +lean_inc(x_130); +if (lean_is_exclusive(x_124)) { + lean_ctor_release(x_124, 0); + lean_ctor_release(x_124, 1); + x_131 = x_124; +} else { + lean_dec_ref(x_124); + x_131 = lean_box(0); +} +if (lean_is_scalar(x_131)) { + x_132 = lean_alloc_ctor(1, 2, 0); +} else { + x_132 = x_131; +} +lean_ctor_set(x_132, 0, x_129); +lean_ctor_set(x_132, 1, x_130); +return x_132; +} +} +} +else +{ +lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_5); +x_133 = lean_ctor_get(x_108, 0); +lean_inc(x_133); +x_134 = lean_ctor_get(x_108, 1); +lean_inc(x_134); +if (lean_is_exclusive(x_108)) { + lean_ctor_release(x_108, 0); + lean_ctor_release(x_108, 1); + x_135 = x_108; +} else { + lean_dec_ref(x_108); + x_135 = lean_box(0); +} +if (lean_is_scalar(x_135)) { + x_136 = lean_alloc_ctor(1, 2, 0); +} else { + x_136 = x_135; +} +lean_ctor_set(x_136, 0, x_133); +lean_ctor_set(x_136, 1, x_134); +return x_136; +} +} +} +} +else +{ +uint8_t x_155; +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_155 = !lean_is_exclusive(x_20); +if (x_155 == 0) +{ +return x_20; +} +else +{ +lean_object* x_156; lean_object* x_157; lean_object* x_158; +x_156 = lean_ctor_get(x_20, 0); +x_157 = lean_ctor_get(x_20, 1); +lean_inc(x_157); +lean_inc(x_156); +lean_dec(x_20); +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_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_main___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19, lean_object* x_20, lean_object* x_21, lean_object* x_22, lean_object* x_23) { +_start: +{ +if (lean_obj_tag(x_10) == 0) +{ +lean_object* x_24; +lean_dec(x_22); +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_8); +lean_dec(x_4); +lean_dec(x_2); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_11); +lean_ctor_set(x_24, 1, x_23); +return x_24; +} +else +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; +lean_dec(x_11); +x_25 = lean_ctor_get(x_10, 0); +lean_inc(x_25); +x_26 = lean_ctor_get(x_10, 1); +lean_inc(x_26); +lean_dec(x_10); +x_27 = l_Lean_Meta_Grind_checkMaxInstancesExceeded(x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_23); +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) +{ +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); +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_inc(x_13); +lean_inc(x_8); +lean_inc(x_4); +lean_inc(x_2); +x_32 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_main___spec__1___lambda__1(x_25, x_2, x_4, x_1, x_8, x_5, x_6, x_31, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_30); +if (lean_obj_tag(x_32) == 0) +{ +lean_object* x_33; +x_33 = lean_ctor_get(x_32, 0); +lean_inc(x_33); +if (lean_obj_tag(x_33) == 0) +{ +uint8_t x_34; +lean_dec(x_26); +lean_dec(x_22); +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_8); +lean_dec(x_4); +lean_dec(x_2); +x_34 = !lean_is_exclusive(x_32); +if (x_34 == 0) +{ +lean_object* x_35; lean_object* x_36; +x_35 = lean_ctor_get(x_32, 0); +lean_dec(x_35); +x_36 = lean_ctor_get(x_33, 0); +lean_inc(x_36); +lean_dec(x_33); +lean_ctor_set(x_32, 0, x_36); +return x_32; +} +else +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_37 = lean_ctor_get(x_32, 1); +lean_inc(x_37); +lean_dec(x_32); +x_38 = lean_ctor_get(x_33, 0); +lean_inc(x_38); +lean_dec(x_33); +x_39 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_37); +return x_39; +} +} +else +{ +lean_object* x_40; lean_object* x_41; +x_40 = lean_ctor_get(x_32, 1); +lean_inc(x_40); +lean_dec(x_32); +x_41 = lean_ctor_get(x_33, 0); +lean_inc(x_41); +lean_dec(x_33); +x_10 = x_26; +x_11 = x_41; +x_12 = lean_box(0); +x_23 = x_40; +goto _start; +} +} +else +{ +uint8_t x_43; +lean_dec(x_26); +lean_dec(x_22); +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_8); +lean_dec(x_4); +lean_dec(x_2); +x_43 = !lean_is_exclusive(x_32); +if (x_43 == 0) +{ +return x_32; +} +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_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; +} +} +} +else +{ +uint8_t x_47; +lean_dec(x_26); +lean_dec(x_25); +lean_dec(x_22); +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_8); +lean_dec(x_4); +lean_dec(x_2); +x_47 = !lean_is_exclusive(x_27); +if (x_47 == 0) +{ +lean_object* x_48; lean_object* x_49; +x_48 = lean_ctor_get(x_27, 0); +lean_dec(x_48); +x_49 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___lambda__1___closed__2; +lean_ctor_set(x_27, 0, x_49); +return x_27; +} +else +{ +lean_object* x_50; lean_object* x_51; lean_object* x_52; +x_50 = lean_ctor_get(x_27, 1); +lean_inc(x_50); +lean_dec(x_27); +x_51 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___lambda__1___closed__2; +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_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_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_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_st_ref_get(x_5, x_13); +x_15 = !lean_is_exclusive(x_14); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_16 = lean_ctor_get(x_14, 0); +x_17 = lean_ctor_get(x_14, 1); +x_18 = lean_ctor_get(x_16, 4); +lean_inc(x_18); +lean_dec(x_16); +lean_inc(x_1); +x_19 = l_Lean_Expr_toHeadIndex(x_1); +x_20 = l_Lean_PersistentHashMap_find_x3f___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__1(x_18, x_19); +lean_dec(x_19); +if (lean_obj_tag(x_20) == 0) +{ +lean_object* 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_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_21 = lean_box(0); +lean_ctor_set(x_14, 0, x_21); +return x_14; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; +lean_free_object(x_14); +x_22 = lean_ctor_get(x_20, 0); +lean_inc(x_22); +lean_dec(x_20); +x_23 = lean_ctor_get(x_3, 0); +lean_inc(x_23); +x_24 = lean_ctor_get(x_23, 2); +lean_inc(x_24); +lean_dec(x_23); +x_25 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_unassigned; +x_26 = lean_mk_array(x_24, x_25); +x_27 = lean_ctor_get_uint8(x_3, sizeof(void*)*1); +x_28 = lean_st_ref_get(x_5, x_17); +x_29 = lean_ctor_get(x_28, 0); +lean_inc(x_29); +x_30 = lean_ctor_get(x_28, 1); +lean_inc(x_30); +lean_dec(x_28); +x_31 = lean_ctor_get(x_29, 6); +lean_inc(x_31); +lean_dec(x_29); +x_32 = lean_box(0); +x_33 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__5___closed__1; +lean_inc(x_22); +x_34 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_main___spec__1(x_1, x_2, x_22, x_26, x_27, x_31, x_32, x_33, x_22, x_22, x_33, lean_box(0), x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_30); +lean_dec(x_31); +lean_dec(x_22); +lean_dec(x_1); +if (lean_obj_tag(x_34) == 0) +{ +lean_object* x_35; lean_object* x_36; +x_35 = lean_ctor_get(x_34, 0); +lean_inc(x_35); +x_36 = lean_ctor_get(x_35, 0); +lean_inc(x_36); +lean_dec(x_35); +if (lean_obj_tag(x_36) == 0) +{ +uint8_t x_37; +x_37 = !lean_is_exclusive(x_34); +if (x_37 == 0) +{ +lean_object* x_38; lean_object* x_39; +x_38 = lean_ctor_get(x_34, 0); +lean_dec(x_38); +x_39 = lean_box(0); +lean_ctor_set(x_34, 0, x_39); +return x_34; +} +else +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_40 = lean_ctor_get(x_34, 1); +lean_inc(x_40); +lean_dec(x_34); +x_41 = lean_box(0); +x_42 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_40); +return x_42; +} +} +else +{ +uint8_t x_43; +x_43 = !lean_is_exclusive(x_34); +if (x_43 == 0) +{ +lean_object* x_44; lean_object* x_45; +x_44 = lean_ctor_get(x_34, 0); +lean_dec(x_44); +x_45 = lean_ctor_get(x_36, 0); +lean_inc(x_45); +lean_dec(x_36); +lean_ctor_set(x_34, 0, x_45); +return x_34; +} +else +{ +lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_46 = lean_ctor_get(x_34, 1); +lean_inc(x_46); +lean_dec(x_34); +x_47 = lean_ctor_get(x_36, 0); +lean_inc(x_47); +lean_dec(x_36); +x_48 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_48, 0, x_47); +lean_ctor_set(x_48, 1, x_46); +return x_48; +} +} +} +else +{ +uint8_t x_49; +x_49 = !lean_is_exclusive(x_34); +if (x_49 == 0) +{ +return x_34; +} +else +{ +lean_object* x_50; lean_object* x_51; lean_object* x_52; +x_50 = lean_ctor_get(x_34, 0); +x_51 = lean_ctor_get(x_34, 1); +lean_inc(x_51); +lean_inc(x_50); +lean_dec(x_34); +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_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; +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_dec(x_14); +x_55 = lean_ctor_get(x_53, 4); +lean_inc(x_55); +lean_dec(x_53); +lean_inc(x_1); +x_56 = l_Lean_Expr_toHeadIndex(x_1); +x_57 = l_Lean_PersistentHashMap_find_x3f___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__1(x_55, x_56); +lean_dec(x_56); +if (lean_obj_tag(x_57) == 0) +{ +lean_object* x_58; lean_object* x_59; +lean_dec(x_12); +lean_dec(x_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_58 = lean_box(0); +x_59 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_59, 0, x_58); +lean_ctor_set(x_59, 1, x_54); +return x_59; +} +else +{ +lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; uint8_t x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; +x_60 = lean_ctor_get(x_57, 0); +lean_inc(x_60); +lean_dec(x_57); +x_61 = lean_ctor_get(x_3, 0); +lean_inc(x_61); +x_62 = lean_ctor_get(x_61, 2); +lean_inc(x_62); +lean_dec(x_61); +x_63 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_unassigned; +x_64 = lean_mk_array(x_62, x_63); +x_65 = lean_ctor_get_uint8(x_3, sizeof(void*)*1); +x_66 = lean_st_ref_get(x_5, x_54); +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, 6); +lean_inc(x_69); +lean_dec(x_67); +x_70 = lean_box(0); +x_71 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__5___closed__1; +lean_inc(x_60); +x_72 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_main___spec__1(x_1, x_2, x_60, x_64, x_65, x_69, x_70, x_71, x_60, x_60, x_71, lean_box(0), x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_68); +lean_dec(x_69); +lean_dec(x_60); +lean_dec(x_1); +if (lean_obj_tag(x_72) == 0) +{ +lean_object* x_73; lean_object* x_74; +x_73 = lean_ctor_get(x_72, 0); +lean_inc(x_73); +x_74 = lean_ctor_get(x_73, 0); +lean_inc(x_74); +lean_dec(x_73); +if (lean_obj_tag(x_74) == 0) +{ +lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; +x_75 = lean_ctor_get(x_72, 1); +lean_inc(x_75); +if (lean_is_exclusive(x_72)) { + lean_ctor_release(x_72, 0); + lean_ctor_release(x_72, 1); + x_76 = x_72; +} else { + lean_dec_ref(x_72); + x_76 = lean_box(0); +} +x_77 = lean_box(0); +if (lean_is_scalar(x_76)) { + x_78 = lean_alloc_ctor(0, 2, 0); +} else { + x_78 = x_76; +} +lean_ctor_set(x_78, 0, x_77); +lean_ctor_set(x_78, 1, x_75); +return x_78; +} +else +{ +lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; +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_72); + x_80 = lean_box(0); +} +x_81 = lean_ctor_get(x_74, 0); +lean_inc(x_81); +lean_dec(x_74); +if (lean_is_scalar(x_80)) { + x_82 = lean_alloc_ctor(0, 2, 0); +} else { + x_82 = x_80; +} +lean_ctor_set(x_82, 0, x_81); +lean_ctor_set(x_82, 1, x_79); +return x_82; +} +} +else +{ +lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; +x_83 = lean_ctor_get(x_72, 0); +lean_inc(x_83); +x_84 = lean_ctor_get(x_72, 1); +lean_inc(x_84); +if (lean_is_exclusive(x_72)) { + lean_ctor_release(x_72, 0); + lean_ctor_release(x_72, 1); + x_85 = x_72; +} else { + lean_dec_ref(x_72); + 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_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_main___spec__1___lambda__1___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +lean_object* x_18 = _args[17]; +lean_object* x_19 = _args[18]; +_start: +{ +uint8_t x_20; lean_object* x_21; +x_20 = lean_unbox(x_6); +lean_dec(x_6); +x_21 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_main___spec__1___lambda__1(x_1, x_2, x_3, x_4, x_5, x_20, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_4); +return x_21; +} +} +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_main___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_object* x_19 = _args[18]; +lean_object* x_20 = _args[19]; +lean_object* x_21 = _args[20]; +lean_object* x_22 = _args[21]; +lean_object* x_23 = _args[22]; +_start: +{ +uint8_t x_24; lean_object* x_25; +x_24 = lean_unbox(x_5); +lean_dec(x_5); +x_25 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_main___spec__1(x_1, x_2, x_3, x_4, x_24, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_23); +lean_dec(x_9); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_3); +lean_dec(x_1); +return x_25; +} +} +LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Meta_Grind_EMatch_ematchTheorem_tryAll___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 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_7, 0, x_5); +lean_ctor_set(x_1, 1, x_2); +lean_ctor_set(x_1, 0, x_7); +{ +lean_object* _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 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_11, 0, x_9); +x_12 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_2); +x_1 = x_10; +x_2 = x_12; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_EMatch_ematchTheorem_tryAll(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_14; lean_object* 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_4); +lean_dec(x_3); +lean_dec(x_2); +x_14 = lean_box(0); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_14); +lean_ctor_set(x_15, 1, x_13); +return x_15; +} +else +{ +uint8_t x_16; +x_16 = !lean_is_exclusive(x_1); +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; +x_17 = lean_ctor_get(x_1, 0); +x_18 = lean_ctor_get(x_1, 1); +lean_inc(x_2); +x_19 = l_List_reverse___rarg(x_2); +x_20 = lean_box(0); +lean_inc(x_18); +x_21 = l_List_mapTR_loop___at_Lean_Meta_Grind_EMatch_ematchTheorem_tryAll___spec__1(x_18, x_20); +x_22 = l_List_appendTR___rarg(x_19, x_21); +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); +lean_inc(x_3); +lean_inc(x_17); +x_23 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_main(x_17, x_22, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_23) == 0) +{ +lean_object* x_24; lean_object* x_25; +x_24 = lean_ctor_get(x_23, 1); +lean_inc(x_24); +lean_dec(x_23); +x_25 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_25, 0, x_17); +lean_ctor_set(x_1, 1, x_2); +lean_ctor_set(x_1, 0, x_25); +{ +lean_object* _tmp_0 = x_18; +lean_object* _tmp_1 = x_1; +lean_object* _tmp_12 = x_24; +x_1 = _tmp_0; +x_2 = _tmp_1; +x_13 = _tmp_12; +} +goto _start; +} +else +{ +uint8_t x_27; +lean_free_object(x_1); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_27 = !lean_is_exclusive(x_23); +if (x_27 == 0) +{ +return x_23; +} +else +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_28 = lean_ctor_get(x_23, 0); +x_29 = lean_ctor_get(x_23, 1); +lean_inc(x_29); +lean_inc(x_28); +lean_dec(x_23); +x_30 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_30, 0, x_28); +lean_ctor_set(x_30, 1, x_29); +return x_30; +} +} +} +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_31 = lean_ctor_get(x_1, 0); +x_32 = lean_ctor_get(x_1, 1); +lean_inc(x_32); +lean_inc(x_31); +lean_dec(x_1); +lean_inc(x_2); +x_33 = l_List_reverse___rarg(x_2); +x_34 = lean_box(0); +lean_inc(x_32); +x_35 = l_List_mapTR_loop___at_Lean_Meta_Grind_EMatch_ematchTheorem_tryAll___spec__1(x_32, x_34); +x_36 = l_List_appendTR___rarg(x_33, x_35); +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); +lean_inc(x_3); +lean_inc(x_31); +x_37 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_main(x_31, x_36, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_37) == 0) +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_38 = lean_ctor_get(x_37, 1); +lean_inc(x_38); +lean_dec(x_37); +x_39 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_39, 0, x_31); +x_40 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_40, 0, x_39); +lean_ctor_set(x_40, 1, x_2); +x_1 = x_32; +x_2 = x_40; +x_13 = x_38; +goto _start; +} +else +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; +lean_dec(x_32); +lean_dec(x_31); +lean_dec(x_12); +lean_dec(x_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_42 = lean_ctor_get(x_37, 0); +lean_inc(x_42); +x_43 = lean_ctor_get(x_37, 1); +lean_inc(x_43); +if (lean_is_exclusive(x_37)) { + lean_ctor_release(x_37, 0); + lean_ctor_release(x_37, 1); + x_44 = x_37; +} else { + lean_dec_ref(x_37); + x_44 = lean_box(0); +} +if (lean_is_scalar(x_44)) { + x_45 = lean_alloc_ctor(1, 2, 0); +} else { + x_45 = x_44; +} +lean_ctor_set(x_45, 0, x_42); +lean_ctor_set(x_45, 1, x_43); +return x_45; +} +} +} +} +} +static lean_object* _init_l_Lean_Meta_Grind_EMatch_ematchTheorem___lambda__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.Meta.Grind.EMatch.ematchTheorem", 36, 36); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_EMatch_ematchTheorem___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; lean_object* x_6; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___closed__1; +x_2 = l_Lean_Meta_Grind_EMatch_ematchTheorem___lambda__1___closed__1; +x_3 = lean_unsigned_to_nat(242u); +x_4 = lean_unsigned_to_nat(22u); +x_5 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___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_Meta_Grind_EMatch_ematchTheorem___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; uint8_t x_15; +x_14 = lean_ctor_get(x_1, 3); +lean_inc(x_14); +x_15 = !lean_is_exclusive(x_3); +if (x_15 == 0) +{ +uint8_t x_16; lean_object* x_17; +x_16 = lean_ctor_get_uint8(x_3, sizeof(void*)*1); +x_17 = lean_ctor_get(x_3, 0); +lean_dec(x_17); +lean_ctor_set(x_3, 0, x_1); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_18; lean_object* x_19; +x_18 = l_Lean_Meta_Grind_EMatch_ematchTheorem___lambda__1___closed__2; +x_19 = l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__2(x_18, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_19; +} +else +{ +lean_object* x_20; +x_20 = lean_ctor_get(x_14, 1); +lean_inc(x_20); +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_14, 0); +lean_inc(x_21); +lean_dec(x_14); +x_22 = lean_box(0); +x_23 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_main(x_21, x_22, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_23; +} +else +{ +if (x_16 == 0) +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_24 = lean_ctor_get(x_14, 0); +lean_inc(x_24); +lean_dec(x_14); +x_25 = lean_box(0); +x_26 = l_List_mapTR_loop___at_Lean_Meta_Grind_EMatch_ematchTheorem_tryAll___spec__1(x_20, x_25); +x_27 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_main(x_24, x_26, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_27; +} +else +{ +lean_object* x_28; lean_object* x_29; +lean_dec(x_20); +x_28 = lean_box(0); +x_29 = l_Lean_Meta_Grind_EMatch_ematchTheorem_tryAll(x_14, x_28, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_29; +} +} +} +} +else +{ +uint8_t x_30; lean_object* x_31; +x_30 = lean_ctor_get_uint8(x_3, sizeof(void*)*1); +lean_dec(x_3); +x_31 = lean_alloc_ctor(0, 1, 1); +lean_ctor_set(x_31, 0, x_1); +lean_ctor_set_uint8(x_31, sizeof(void*)*1, x_30); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_32; lean_object* x_33; +x_32 = l_Lean_Meta_Grind_EMatch_ematchTheorem___lambda__1___closed__2; +x_33 = l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__2(x_32, x_31, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_33; +} +else +{ +lean_object* x_34; +x_34 = lean_ctor_get(x_14, 1); +lean_inc(x_34); +if (lean_obj_tag(x_34) == 0) +{ +lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_35 = lean_ctor_get(x_14, 0); +lean_inc(x_35); +lean_dec(x_14); +x_36 = lean_box(0); +x_37 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_main(x_35, x_36, x_31, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_37; +} +else +{ +if (x_30 == 0) +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_38 = lean_ctor_get(x_14, 0); +lean_inc(x_38); +lean_dec(x_14); +x_39 = lean_box(0); +x_40 = l_List_mapTR_loop___at_Lean_Meta_Grind_EMatch_ematchTheorem_tryAll___spec__1(x_34, x_39); +x_41 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_main(x_38, x_40, x_31, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_41; +} +else +{ +lean_object* x_42; lean_object* x_43; +lean_dec(x_34); +x_42 = lean_box(0); +x_43 = l_Lean_Meta_Grind_EMatch_ematchTheorem_tryAll(x_14, x_42, x_31, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_43; +} +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_EMatch_ematchTheorem(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, 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; uint8_t x_15; +x_13 = l_Lean_Meta_Grind_checkMaxInstancesExceeded(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_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_13, 1); +lean_inc(x_16); +lean_dec(x_13); +x_17 = lean_box(0); +x_18 = l_Lean_Meta_Grind_EMatch_ematchTheorem___lambda__1(x_1, x_17, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_16); +return x_18; +} +else +{ +uint8_t x_19; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_19 = !lean_is_exclusive(x_13); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; +x_20 = lean_ctor_get(x_13, 0); +lean_dec(x_20); +x_21 = lean_box(0); +lean_ctor_set(x_13, 0, x_21); +return x_13; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_22 = lean_ctor_get(x_13, 1); +lean_inc(x_22); +lean_dec(x_13); +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_EXPORT lean_object* l_Lean_Meta_Grind_EMatch_ematchTheorem___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +lean_object* x_14; +x_14 = l_Lean_Meta_Grind_EMatch_ematchTheorem___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_2); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_Grind_EMatch_ematchTheorems___spec__3(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +_start: +{ +uint8_t x_16; +x_16 = lean_usize_dec_eq(x_2, x_3); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; +lean_dec(x_4); +x_17 = lean_array_uget(x_1, x_2); +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); +lean_inc(x_5); +x_18 = l_Lean_PersistentArray_forMAux___at_Lean_Meta_Grind_EMatch_ematchTheorems___spec__2(x_17, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +lean_dec(x_17); +if (lean_obj_tag(x_18) == 0) +{ +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); +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_15 = x_20; +goto _start; +} +else +{ +uint8_t x_24; +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +x_24 = !lean_is_exclusive(x_18); +if (x_24 == 0) +{ +return x_18; +} +else +{ +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_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_28; +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_4); +lean_ctor_set(x_28, 1, x_15); +return x_28; +} +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_Grind_EMatch_ematchTheorems___spec__4(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +_start: +{ +uint8_t x_16; +x_16 = lean_usize_dec_eq(x_2, x_3); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; +lean_dec(x_4); +x_17 = lean_array_uget(x_1, x_2); +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); +lean_inc(x_5); +x_18 = l_Lean_Meta_Grind_EMatch_ematchTheorem(x_17, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +if (lean_obj_tag(x_18) == 0) +{ +lean_object* x_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); +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_15 = x_20; +goto _start; +} +else +{ +uint8_t x_24; +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +x_24 = !lean_is_exclusive(x_18); +if (x_24 == 0) +{ +return x_18; +} +else +{ +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_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_28; +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_4); +lean_ctor_set(x_28, 1, x_15); +return x_28; +} +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentArray_forMAux___at_Lean_Meta_Grind_EMatch_ematchTheorems___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_13 = lean_ctor_get(x_1, 0); +x_14 = lean_array_get_size(x_13); +x_15 = lean_unsigned_to_nat(0u); +x_16 = lean_nat_dec_lt(x_15, x_14); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; +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_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_12); +return x_18; +} +else +{ +uint8_t x_19; +x_19 = lean_nat_dec_le(x_14, x_14); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_20 = lean_box(0); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_20); +lean_ctor_set(x_21, 1, x_12); +return x_21; +} +else +{ +size_t x_22; size_t x_23; lean_object* x_24; lean_object* x_25; +x_22 = 0; +x_23 = lean_usize_of_nat(x_14); +lean_dec(x_14); +x_24 = lean_box(0); +x_25 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_Grind_EMatch_ematchTheorems___spec__3(x_13, x_22, x_23, x_24, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_25; +} +} +} +else +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; +x_26 = lean_ctor_get(x_1, 0); +x_27 = lean_array_get_size(x_26); +x_28 = lean_unsigned_to_nat(0u); +x_29 = lean_nat_dec_lt(x_28, x_27); +if (x_29 == 0) +{ +lean_object* x_30; lean_object* x_31; +lean_dec(x_27); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +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_12); +return x_31; +} +else +{ +uint8_t x_32; +x_32 = lean_nat_dec_le(x_27, x_27); +if (x_32 == 0) +{ +lean_object* x_33; lean_object* x_34; +lean_dec(x_27); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +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_12); +return x_34; +} +else +{ +size_t x_35; size_t x_36; lean_object* x_37; lean_object* x_38; +x_35 = 0; +x_36 = lean_usize_of_nat(x_27); +lean_dec(x_27); +x_37 = lean_box(0); +x_38 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_Grind_EMatch_ematchTheorems___spec__4(x_26, x_35, x_36, x_37, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_38; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentArray_forM___at_Lean_Meta_Grind_EMatch_ematchTheorems___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: +{ +lean_object* x_13; lean_object* x_14; +x_13 = lean_ctor_get(x_1, 0); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +x_14 = l_Lean_PersistentArray_forMAux___at_Lean_Meta_Grind_EMatch_ematchTheorems___spec__2(x_13, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_14) == 0) +{ +uint8_t x_15; +x_15 = !lean_is_exclusive(x_14); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_16 = lean_ctor_get(x_14, 1); +x_17 = lean_ctor_get(x_14, 0); +lean_dec(x_17); +x_18 = lean_ctor_get(x_1, 1); +x_19 = lean_array_get_size(x_18); +x_20 = lean_unsigned_to_nat(0u); +x_21 = lean_nat_dec_lt(x_20, x_19); +if (x_21 == 0) +{ +lean_object* x_22; +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); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_22 = lean_box(0); +lean_ctor_set(x_14, 0, x_22); +return x_14; +} +else +{ +uint8_t x_23; +x_23 = lean_nat_dec_le(x_19, x_19); +if (x_23 == 0) +{ +lean_object* x_24; +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); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_24 = lean_box(0); +lean_ctor_set(x_14, 0, x_24); +return x_14; +} +else +{ +size_t x_25; size_t x_26; lean_object* x_27; lean_object* x_28; +lean_free_object(x_14); +x_25 = 0; +x_26 = lean_usize_of_nat(x_19); +lean_dec(x_19); +x_27 = lean_box(0); +x_28 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_Grind_EMatch_ematchTheorems___spec__4(x_18, x_25, x_26, x_27, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_16); +return x_28; +} +} +} +else +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; +x_29 = lean_ctor_get(x_14, 1); +lean_inc(x_29); +lean_dec(x_14); +x_30 = lean_ctor_get(x_1, 1); +x_31 = lean_array_get_size(x_30); +x_32 = lean_unsigned_to_nat(0u); +x_33 = lean_nat_dec_lt(x_32, x_31); +if (x_33 == 0) +{ +lean_object* x_34; lean_object* x_35; +lean_dec(x_31); +lean_dec(x_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_34 = lean_box(0); +x_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_34); +lean_ctor_set(x_35, 1, x_29); +return x_35; +} +else +{ +uint8_t x_36; +x_36 = lean_nat_dec_le(x_31, x_31); +if (x_36 == 0) +{ +lean_object* x_37; lean_object* x_38; +lean_dec(x_31); +lean_dec(x_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_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_29); +return x_38; +} +else +{ +size_t x_39; size_t x_40; lean_object* x_41; lean_object* x_42; +x_39 = 0; +x_40 = lean_usize_of_nat(x_31); +lean_dec(x_31); +x_41 = lean_box(0); +x_42 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_Grind_EMatch_ematchTheorems___spec__4(x_30, x_39, x_40, x_41, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_29); +return x_42; +} +} +} +} +else +{ +uint8_t x_43; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_43 = !lean_is_exclusive(x_14); +if (x_43 == 0) +{ +return x_14; +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_14, 0); +x_45 = lean_ctor_get(x_14, 1); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_14); +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_EXPORT lean_object* l_Lean_Meta_Grind_EMatch_ematchTheorems(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +x_13 = l_Lean_PersistentArray_forM___at_Lean_Meta_Grind_EMatch_ematchTheorems___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); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_Grind_EMatch_ematchTheorems___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +_start: +{ +size_t x_16; size_t x_17; lean_object* x_18; +x_16 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_17 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_18 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_Grind_EMatch_ematchTheorems___spec__3(x_1, x_16, x_17, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +lean_dec(x_1); +return x_18; +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_Grind_EMatch_ematchTheorems___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +_start: +{ +size_t x_16; size_t x_17; lean_object* x_18; +x_16 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_17 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_18 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_Grind_EMatch_ematchTheorems___spec__4(x_1, x_16, x_17, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +lean_dec(x_1); +return x_18; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentArray_forMAux___at_Lean_Meta_Grind_EMatch_ematchTheorems___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +x_13 = l_Lean_PersistentArray_forMAux___at_Lean_Meta_Grind_EMatch_ematchTheorems___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_1); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentArray_forM___at_Lean_Meta_Grind_EMatch_ematchTheorems___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +x_13 = l_Lean_PersistentArray_forM___at_Lean_Meta_Grind_EMatch_ematchTheorems___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_1); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_EMatch_ematchTheorems___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +x_13 = l_Lean_Meta_Grind_EMatch_ematchTheorems(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_1); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematch___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +uint8_t x_13; +x_13 = !lean_is_exclusive(x_2); +if (x_13 == 0) +{ +uint8_t x_14; lean_object* x_15; +x_14 = 1; +lean_ctor_set_uint8(x_2, sizeof(void*)*1, x_14); +x_15 = l_Lean_PersistentArray_forM___at_Lean_Meta_Grind_EMatch_ematchTheorems___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); +return x_15; +} +else +{ +lean_object* x_16; uint8_t x_17; lean_object* x_18; lean_object* x_19; +x_16 = lean_ctor_get(x_2, 0); +lean_inc(x_16); +lean_dec(x_2); +x_17 = 1; +x_18 = lean_alloc_ctor(0, 1, 1); +lean_ctor_set(x_18, 0, x_16); +lean_ctor_set_uint8(x_18, sizeof(void*)*1, x_17); +x_19 = l_Lean_PersistentArray_forM___at_Lean_Meta_Grind_EMatch_ematchTheorems___spec__1(x_1, x_18, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_19; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematch___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +uint8_t x_14; +x_14 = !lean_is_exclusive(x_3); +if (x_14 == 0) +{ +uint8_t x_15; lean_object* x_16; +x_15 = 0; +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_15); +x_16 = l_Lean_PersistentArray_forM___at_Lean_Meta_Grind_EMatch_ematchTheorems___spec__1(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_16; +} +else +{ +lean_object* x_17; uint8_t x_18; lean_object* x_19; lean_object* x_20; +x_17 = lean_ctor_get(x_3, 0); +lean_inc(x_17); +lean_dec(x_3); +x_18 = 0; +x_19 = lean_alloc_ctor(0, 1, 1); +lean_ctor_set(x_19, 0, x_17); +lean_ctor_set_uint8(x_19, sizeof(void*)*1, x_18); +x_20 = l_Lean_PersistentArray_forM___at_Lean_Meta_Grind_EMatch_ematchTheorems___spec__1(x_1, x_19, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_20; +} +} +} +static lean_object* _init_l_Lean_Meta_Grind_ematch___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(32u); +x_2 = lean_mk_empty_array_with_capacity(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Grind_ematch___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_ematch___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_Meta_Grind_ematch___closed__3() { +_start: +{ +size_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = 5; +x_2 = l_Lean_Meta_Grind_ematch___closed__2; +x_3 = l_Lean_Meta_Grind_ematch___closed__1; +x_4 = lean_unsigned_to_nat(0u); +x_5 = lean_alloc_ctor(0, 4, sizeof(size_t)*1); +lean_ctor_set(x_5, 0, x_2); +lean_ctor_set(x_5, 1, x_3); +lean_ctor_set(x_5, 2, x_4); +lean_ctor_set(x_5, 3, x_4); +lean_ctor_set_usize(x_5, 4, x_1); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematch(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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_Meta_Grind_checkMaxInstancesExceeded(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_11 = lean_ctor_get(x_10, 0); +lean_inc(x_11); +x_12 = lean_unbox(x_11); +lean_dec(x_11); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_13 = lean_ctor_get(x_10, 1); +lean_inc(x_13); +lean_dec(x_10); +x_14 = lean_st_ref_get(x_1, x_13); +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +x_16 = lean_ctor_get(x_14, 1); +lean_inc(x_16); +lean_dec(x_14); +x_17 = lean_st_ref_get(x_1, x_16); +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +x_19 = lean_ctor_get(x_17, 1); +lean_inc(x_19); +lean_dec(x_17); +x_20 = lean_ctor_get(x_15, 8); +lean_inc(x_20); +lean_dec(x_15); +x_21 = lean_ctor_get(x_18, 9); +lean_inc(x_21); +lean_dec(x_18); +x_22 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_ematch___lambda__1___boxed), 12, 1); +lean_closure_set(x_22, 0, x_20); +x_23 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_ematch___lambda__2___boxed), 13, 1); +lean_closure_set(x_23, 0, x_21); +x_24 = lean_alloc_closure((void*)(l_ReaderT_bind___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__11___rarg), 13, 2); +lean_closure_set(x_24, 0, x_22); +lean_closure_set(x_24, 1, x_23); +lean_inc(x_1); +x_25 = l_Lean_Meta_Grind_EMatch_M_run_x27___rarg(x_24, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_19); +if (lean_obj_tag(x_25) == 0) +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; uint8_t x_30; +x_26 = lean_ctor_get(x_25, 1); +lean_inc(x_26); +lean_dec(x_25); +x_27 = lean_st_ref_take(x_1, x_26); +x_28 = lean_ctor_get(x_27, 0); +lean_inc(x_28); +x_29 = lean_ctor_get(x_27, 1); +lean_inc(x_29); +lean_dec(x_27); +x_30 = !lean_is_exclusive(x_28); +if (x_30 == 0) +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; uint8_t x_39; +x_31 = lean_ctor_get(x_28, 6); +x_32 = lean_ctor_get(x_28, 8); +x_33 = lean_ctor_get(x_28, 9); +x_34 = lean_unsigned_to_nat(1u); +x_35 = lean_nat_add(x_31, x_34); +lean_dec(x_31); +x_36 = l_Lean_PersistentArray_append___rarg(x_32, x_33); +lean_dec(x_33); +x_37 = l_Lean_Meta_Grind_ematch___closed__3; +lean_ctor_set(x_28, 9, x_37); +lean_ctor_set(x_28, 8, x_36); +lean_ctor_set(x_28, 6, x_35); +x_38 = lean_st_ref_set(x_1, x_28, x_29); +lean_dec(x_1); +x_39 = !lean_is_exclusive(x_38); +if (x_39 == 0) +{ +lean_object* x_40; lean_object* x_41; +x_40 = lean_ctor_get(x_38, 0); +lean_dec(x_40); +x_41 = lean_box(0); +lean_ctor_set(x_38, 0, x_41); +return x_38; +} +else +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_42 = lean_ctor_get(x_38, 1); +lean_inc(x_42); +lean_dec(x_38); +x_43 = lean_box(0); +x_44 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_44, 0, x_43); +lean_ctor_set(x_44, 1, x_42); +return x_44; +} +} +else +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; 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; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; +x_45 = lean_ctor_get(x_28, 0); +x_46 = lean_ctor_get(x_28, 1); +x_47 = lean_ctor_get(x_28, 2); +x_48 = lean_ctor_get(x_28, 3); +x_49 = lean_ctor_get(x_28, 4); +x_50 = lean_ctor_get(x_28, 5); +x_51 = lean_ctor_get_uint8(x_28, sizeof(void*)*14); +x_52 = lean_ctor_get(x_28, 6); +x_53 = lean_ctor_get(x_28, 7); +x_54 = lean_ctor_get(x_28, 8); +x_55 = lean_ctor_get(x_28, 9); +x_56 = lean_ctor_get(x_28, 10); +x_57 = lean_ctor_get(x_28, 11); +x_58 = lean_ctor_get(x_28, 12); +x_59 = lean_ctor_get(x_28, 13); +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_50); +lean_inc(x_49); +lean_inc(x_48); +lean_inc(x_47); +lean_inc(x_46); +lean_inc(x_45); +lean_dec(x_28); +x_60 = lean_unsigned_to_nat(1u); +x_61 = lean_nat_add(x_52, x_60); +lean_dec(x_52); +x_62 = l_Lean_PersistentArray_append___rarg(x_54, x_55); +lean_dec(x_55); +x_63 = l_Lean_Meta_Grind_ematch___closed__3; +x_64 = lean_alloc_ctor(0, 14, 1); +lean_ctor_set(x_64, 0, x_45); +lean_ctor_set(x_64, 1, x_46); +lean_ctor_set(x_64, 2, x_47); +lean_ctor_set(x_64, 3, x_48); +lean_ctor_set(x_64, 4, x_49); +lean_ctor_set(x_64, 5, x_50); +lean_ctor_set(x_64, 6, x_61); +lean_ctor_set(x_64, 7, x_53); +lean_ctor_set(x_64, 8, x_62); +lean_ctor_set(x_64, 9, x_63); +lean_ctor_set(x_64, 10, x_56); +lean_ctor_set(x_64, 11, x_57); +lean_ctor_set(x_64, 12, x_58); +lean_ctor_set(x_64, 13, x_59); +lean_ctor_set_uint8(x_64, sizeof(void*)*14, x_51); +x_65 = lean_st_ref_set(x_1, x_64, x_29); +lean_dec(x_1); +x_66 = lean_ctor_get(x_65, 1); +lean_inc(x_66); +if (lean_is_exclusive(x_65)) { + lean_ctor_release(x_65, 0); + lean_ctor_release(x_65, 1); + x_67 = x_65; +} else { + lean_dec_ref(x_65); + x_67 = lean_box(0); +} +x_68 = lean_box(0); +if (lean_is_scalar(x_67)) { + x_69 = lean_alloc_ctor(0, 2, 0); +} else { + x_69 = x_67; +} +lean_ctor_set(x_69, 0, x_68); +lean_ctor_set(x_69, 1, x_66); +return x_69; +} +} +else +{ +uint8_t x_70; +lean_dec(x_1); +x_70 = !lean_is_exclusive(x_25); +if (x_70 == 0) +{ +return x_25; +} +else +{ +lean_object* x_71; lean_object* x_72; lean_object* x_73; +x_71 = lean_ctor_get(x_25, 0); +x_72 = lean_ctor_get(x_25, 1); +lean_inc(x_72); +lean_inc(x_71); +lean_dec(x_25); +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_74; +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_74 = !lean_is_exclusive(x_10); +if (x_74 == 0) +{ +lean_object* x_75; lean_object* x_76; +x_75 = lean_ctor_get(x_10, 0); +lean_dec(x_75); +x_76 = lean_box(0); +lean_ctor_set(x_10, 0, x_76); +return x_10; +} +else +{ +lean_object* x_77; lean_object* x_78; lean_object* x_79; +x_77 = lean_ctor_get(x_10, 1); +lean_inc(x_77); +lean_dec(x_10); +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_EXPORT lean_object* l_Lean_Meta_Grind_ematch___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_13; +x_13 = l_Lean_Meta_Grind_ematch___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_1); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematch___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +lean_object* x_14; +x_14 = l_Lean_Meta_Grind_ematch___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_2); +lean_dec(x_1); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematchAndAssert_x3f___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; uint8_t x_11; +x_10 = lean_st_mk_ref(x_1, x_9); +x_11 = !lean_is_exclusive(x_10); +if (x_11 == 0) +{ +return x_10; +} +else +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_12 = lean_ctor_get(x_10, 0); +x_13 = lean_ctor_get(x_10, 1); +lean_inc(x_13); +lean_inc(x_12); +lean_dec(x_10); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_12); +lean_ctor_set(x_14, 1, x_13); +return x_14; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematchAndAssert_x3f___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +lean_inc(x_1); +x_10 = l_Lean_Meta_Grind_ematch(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_10) == 0) +{ +lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_11 = lean_ctor_get(x_10, 1); +lean_inc(x_11); +lean_dec(x_10); +x_12 = lean_st_ref_get(x_1, x_11); +x_13 = !lean_is_exclusive(x_12); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_14 = lean_ctor_get(x_12, 1); +x_15 = lean_st_ref_get(x_1, x_14); +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, 0); +lean_ctor_set(x_12, 1, x_17); +lean_ctor_set(x_15, 0, x_12); +return x_15; +} +else +{ +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_12, 1, x_18); +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 +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_21 = lean_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_st_ref_get(x_1, x_22); +lean_dec(x_1); +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +x_25 = lean_ctor_get(x_23, 1); +lean_inc(x_25); +if (lean_is_exclusive(x_23)) { + lean_ctor_release(x_23, 0); + lean_ctor_release(x_23, 1); + x_26 = x_23; +} else { + lean_dec_ref(x_23); + x_26 = lean_box(0); +} +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_21); +lean_ctor_set(x_27, 1, x_24); +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; +} +} +else +{ +uint8_t x_29; +lean_dec(x_1); +x_29 = !lean_is_exclusive(x_10); +if (x_29 == 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_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; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematchAndAssert_x3f___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +uint8_t x_10; +x_10 = !lean_is_exclusive(x_1); +if (x_10 == 0) +{ +lean_object* x_11; +x_11 = lean_ctor_get(x_1, 1); +lean_dec(x_11); +lean_ctor_set(x_1, 1, x_9); +return x_1; +} +else +{ +lean_object* x_12; lean_object* x_13; +x_12 = lean_ctor_get(x_1, 0); +lean_inc(x_12); +lean_dec(x_1); +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_9); +return x_13; +} +} +} +static lean_object* _init_l_Lean_Meta_Grind_ematchAndAssert_x3f___lambda__4___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_assertNext_x3f), 9, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematchAndAssert_x3f___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; lean_object* x_12; +x_11 = l_Lean_Meta_Grind_ematchAndAssert_x3f___lambda__4___closed__1; +x_12 = l_Lean_Meta_Grind_iterate(x_1, 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) +{ +uint8_t x_13; +x_13 = !lean_is_exclusive(x_12); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; +x_14 = lean_ctor_get(x_12, 0); +x_15 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_15, 0, x_14); +lean_ctor_set(x_12, 0, x_15); +return x_12; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_16 = lean_ctor_get(x_12, 0); +x_17 = lean_ctor_get(x_12, 1); +lean_inc(x_17); +lean_inc(x_16); +lean_dec(x_12); +x_18 = lean_alloc_ctor(1, 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; +} +} +else +{ +uint8_t x_20; +x_20 = !lean_is_exclusive(x_12); +if (x_20 == 0) +{ +return x_12; +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_12, 0); +x_22 = lean_ctor_get(x_12, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_12); +x_23 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_23, 0, x_21); +lean_ctor_set(x_23, 1, x_22); +return x_23; +} +} +} +} +static lean_object* _init_l_Lean_Meta_Grind_ematchAndAssert_x3f___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_ematchAndAssert_x3f___lambda__2), 9, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_ematchAndAssert_x3f___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_ematchAndAssert_x3f___lambda__3___boxed), 9, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematchAndAssert_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) { +_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; +x_10 = lean_ctor_get(x_1, 11); +lean_inc(x_10); +x_11 = lean_ctor_get(x_1, 0); +lean_inc(x_11); +x_12 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_ematchAndAssert_x3f___lambda__1___boxed), 9, 1); +lean_closure_set(x_12, 0, x_1); +x_13 = l_Lean_Meta_Grind_ematchAndAssert_x3f___closed__1; +x_14 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); +lean_closure_set(x_14, 0, x_12); +lean_closure_set(x_14, 1, x_13); +x_15 = l_Lean_Meta_Grind_ematchAndAssert_x3f___closed__2; +x_16 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); +lean_closure_set(x_16, 0, x_14); +lean_closure_set(x_16, 1, 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); +lean_inc(x_2); +x_17 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(x_11, x_16, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_17) == 0) +{ +uint8_t x_18; +x_18 = !lean_is_exclusive(x_17); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_19 = lean_ctor_get(x_17, 0); +x_20 = lean_ctor_get(x_17, 1); +x_21 = lean_ctor_get(x_19, 11); +lean_inc(x_21); +x_22 = lean_nat_dec_eq(x_21, x_10); +lean_dec(x_10); +lean_dec(x_21); +if (x_22 == 0) +{ +lean_object* x_23; lean_object* x_24; +lean_free_object(x_17); +x_23 = lean_box(0); +x_24 = l_Lean_Meta_Grind_ematchAndAssert_x3f___lambda__4(x_19, x_23, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_20); +return x_24; +} +else +{ +lean_object* x_25; +lean_dec(x_19); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_25 = lean_box(0); +lean_ctor_set(x_17, 0, x_25); +return x_17; +} +} +else +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; +x_26 = lean_ctor_get(x_17, 0); +x_27 = lean_ctor_get(x_17, 1); +lean_inc(x_27); +lean_inc(x_26); +lean_dec(x_17); +x_28 = lean_ctor_get(x_26, 11); +lean_inc(x_28); +x_29 = lean_nat_dec_eq(x_28, x_10); +lean_dec(x_10); +lean_dec(x_28); +if (x_29 == 0) +{ +lean_object* x_30; lean_object* x_31; +x_30 = lean_box(0); +x_31 = l_Lean_Meta_Grind_ematchAndAssert_x3f___lambda__4(x_26, x_30, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_27); +return x_31; +} +else +{ +lean_object* x_32; lean_object* x_33; +lean_dec(x_26); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_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_27); +return x_33; +} +} +} +else +{ +uint8_t x_34; +lean_dec(x_10); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_34 = !lean_is_exclusive(x_17); +if (x_34 == 0) +{ +return x_17; +} +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_EXPORT lean_object* l_Lean_Meta_Grind_ematchAndAssert_x3f___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l_Lean_Meta_Grind_ematchAndAssert_x3f___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematchAndAssert_x3f___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l_Lean_Meta_Grind_ematchAndAssert_x3f___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematchAndAssert_x3f___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_11; +x_11 = l_Lean_Meta_Grind_ematchAndAssert_x3f___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_2); +return x_11; +} +} +static lean_object* _init_l_Lean_Meta_Grind_ematchStar___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_ematchAndAssert_x3f), 9, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematchStar(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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_Meta_Grind_ematchStar___closed__1; +x_11 = l_Lean_Meta_Grind_iterate(x_1, x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_11; +} +} +lean_object* initialize_Lean_Meta_Tactic_Grind_Types(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Meta_Tactic_Grind_Intro(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Lean_Meta_Tactic_Grind_EMatch(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_Tactic_Grind_Types(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Lean_Meta_Tactic_Grind_Intro(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l_Lean_Meta_Grind_EMatch_instInhabitedCnstr___closed__1 = _init_l_Lean_Meta_Grind_EMatch_instInhabitedCnstr___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_EMatch_instInhabitedCnstr___closed__1); +l_Lean_Meta_Grind_EMatch_instInhabitedCnstr___closed__2 = _init_l_Lean_Meta_Grind_EMatch_instInhabitedCnstr___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_EMatch_instInhabitedCnstr___closed__2); +l_Lean_Meta_Grind_EMatch_instInhabitedCnstr = _init_l_Lean_Meta_Grind_EMatch_instInhabitedCnstr(); +lean_mark_persistent(l_Lean_Meta_Grind_EMatch_instInhabitedCnstr); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_unassigned___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_unassigned___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_unassigned___closed__1); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_unassigned___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_unassigned___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_unassigned___closed__2); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_unassigned___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_unassigned___closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_unassigned___closed__3); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_unassigned = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_unassigned(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_unassigned); +l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__1 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__1(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__1); +l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2); +l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__3 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__3(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__3); +l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__4 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__4(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__4); +l_Lean_Meta_Grind_EMatch_instInhabitedChoice___closed__1 = _init_l_Lean_Meta_Grind_EMatch_instInhabitedChoice___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_EMatch_instInhabitedChoice___closed__1); +l_Lean_Meta_Grind_EMatch_instInhabitedChoice___closed__2 = _init_l_Lean_Meta_Grind_EMatch_instInhabitedChoice___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_EMatch_instInhabitedChoice___closed__2); +l_Lean_Meta_Grind_EMatch_instInhabitedChoice = _init_l_Lean_Meta_Grind_EMatch_instInhabitedChoice(); +lean_mark_persistent(l_Lean_Meta_Grind_EMatch_instInhabitedChoice); +l_Lean_Meta_Grind_EMatch_instInhabitedContext___closed__1 = _init_l_Lean_Meta_Grind_EMatch_instInhabitedContext___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_EMatch_instInhabitedContext___closed__1); +l_Lean_Meta_Grind_EMatch_instInhabitedContext___closed__2 = _init_l_Lean_Meta_Grind_EMatch_instInhabitedContext___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_EMatch_instInhabitedContext___closed__2); +l_Lean_Meta_Grind_EMatch_instInhabitedContext___closed__3 = _init_l_Lean_Meta_Grind_EMatch_instInhabitedContext___closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind_EMatch_instInhabitedContext___closed__3); +l_Lean_Meta_Grind_EMatch_instInhabitedContext = _init_l_Lean_Meta_Grind_EMatch_instInhabitedContext(); +lean_mark_persistent(l_Lean_Meta_Grind_EMatch_instInhabitedContext); +l_Lean_Meta_Grind_EMatch_instInhabitedState = _init_l_Lean_Meta_Grind_EMatch_instInhabitedState(); +lean_mark_persistent(l_Lean_Meta_Grind_EMatch_instInhabitedState); +l_Lean_Meta_Grind_EMatch_M_run_x27___rarg___closed__1 = _init_l_Lean_Meta_Grind_EMatch_M_run_x27___rarg___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_EMatch_M_run_x27___rarg___closed__1); +l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___spec__1___closed__1 = _init_l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___spec__1___closed__1(); +lean_mark_persistent(l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___spec__1___closed__1); +l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___spec__1___closed__2 = _init_l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___spec__1___closed__2(); +lean_mark_persistent(l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___spec__1___closed__2); +l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___spec__1___closed__3 = _init_l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___spec__1___closed__3(); +lean_mark_persistent(l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___spec__1___closed__3); +l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___spec__1___closed__4 = _init_l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___spec__1___closed__4(); +lean_mark_persistent(l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___spec__1___closed__4); +l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___spec__1___closed__5 = _init_l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___spec__1___closed__5(); +lean_mark_persistent(l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___spec__1___closed__5); +l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___spec__1___closed__6 = _init_l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___spec__1___closed__6(); +lean_mark_persistent(l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___spec__1___closed__6); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___closed__1); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___closed__2); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___closed__3); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___closed__4(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___closed__4); +l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___spec__1___closed__1 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___spec__1___closed__1(); +lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___spec__1___closed__1); +l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5___closed__1 = _init_l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5___closed__1(); +lean_mark_persistent(l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5___closed__1); +l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5___closed__2 = _init_l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5___closed__2(); +lean_mark_persistent(l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5___closed__2); +l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5___closed__3 = _init_l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5___closed__3(); +lean_mark_persistent(l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5___closed__3); +l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5___closed__4 = _init_l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5___closed__4(); +lean_mark_persistent(l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5___closed__4); +l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3___closed__1 = _init_l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3___closed__1); +l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3___closed__2 = _init_l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3___closed__2); +l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3___closed__3 = _init_l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3___closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3___closed__3); +l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__1 = _init_l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__1(); +l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2 = _init_l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2(); +lean_mark_persistent(l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__1); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__2); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__3); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__4(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__4); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__5 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__5(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__5); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__6 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__6(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__6); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___closed__1); +l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__2___closed__1 = _init_l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__2___closed__1(); +lean_mark_persistent(l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__2___closed__1); +l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__2___closed__2 = _init_l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__2___closed__2(); +lean_mark_persistent(l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__2___closed__2); +l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__2___closed__3 = _init_l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__2___closed__3(); +lean_mark_persistent(l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__2___closed__3); +l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___lambda__1___closed__1 = _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___lambda__1___closed__1(); +lean_mark_persistent(l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___lambda__1___closed__1); +l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___lambda__1___closed__2 = _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___lambda__1___closed__2(); +lean_mark_persistent(l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___lambda__1___closed__2); +l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___lambda__1___closed__3 = _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___lambda__1___closed__3(); +lean_mark_persistent(l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___lambda__1___closed__3); +l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__1 = _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__1(); +lean_mark_persistent(l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__1); +l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__2 = _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__2(); +lean_mark_persistent(l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__2); +l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__3 = _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__3(); +lean_mark_persistent(l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__3); +l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__4 = _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__4(); +lean_mark_persistent(l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__4); +l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__5 = _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__5(); +lean_mark_persistent(l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__5); +l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__6 = _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__6(); +lean_mark_persistent(l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__6); +l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__7 = _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__7(); +lean_mark_persistent(l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__7); +l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__8 = _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__8(); +lean_mark_persistent(l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__8); +l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__9 = _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__9(); +lean_mark_persistent(l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__9); +l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___closed__1 = _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___closed__1(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___closed__1); +l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___closed__2 = _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___closed__2(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___closed__2); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__1(); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__2); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__3); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__4(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__4); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__5 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__5(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__5); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__5___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__5___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__5___closed__1); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__1); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__2); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__3); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__4(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__4); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__5 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__5(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__5); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__6 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__6(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__6); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__7 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__7(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__7); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__8 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__8(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__8); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__8___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__8___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__8___closed__1); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__8___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__8___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__8___closed__2); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___closed__1); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___closed__2(); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___lambda__2___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___lambda__2___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___lambda__2___closed__1); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___closed__1); +l_Lean_Meta_Grind_EMatch_ematchTheorem___lambda__1___closed__1 = _init_l_Lean_Meta_Grind_EMatch_ematchTheorem___lambda__1___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_EMatch_ematchTheorem___lambda__1___closed__1); +l_Lean_Meta_Grind_EMatch_ematchTheorem___lambda__1___closed__2 = _init_l_Lean_Meta_Grind_EMatch_ematchTheorem___lambda__1___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_EMatch_ematchTheorem___lambda__1___closed__2); +l_Lean_Meta_Grind_ematch___closed__1 = _init_l_Lean_Meta_Grind_ematch___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_ematch___closed__1); +l_Lean_Meta_Grind_ematch___closed__2 = _init_l_Lean_Meta_Grind_ematch___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_ematch___closed__2); +l_Lean_Meta_Grind_ematch___closed__3 = _init_l_Lean_Meta_Grind_ematch___closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind_ematch___closed__3); +l_Lean_Meta_Grind_ematchAndAssert_x3f___lambda__4___closed__1 = _init_l_Lean_Meta_Grind_ematchAndAssert_x3f___lambda__4___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_ematchAndAssert_x3f___lambda__4___closed__1); +l_Lean_Meta_Grind_ematchAndAssert_x3f___closed__1 = _init_l_Lean_Meta_Grind_ematchAndAssert_x3f___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_ematchAndAssert_x3f___closed__1); +l_Lean_Meta_Grind_ematchAndAssert_x3f___closed__2 = _init_l_Lean_Meta_Grind_ematchAndAssert_x3f___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_ematchAndAssert_x3f___closed__2); +l_Lean_Meta_Grind_ematchStar___closed__1 = _init_l_Lean_Meta_Grind_ematchStar___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_ematchStar___closed__1); +return lean_io_result_mk_ok(lean_box(0)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Lean/Meta/Tactic/Grind/EMatchTheorem.c b/stage0/stdlib/Lean/Meta/Tactic/Grind/EMatchTheorem.c new file mode 100644 index 000000000000..3186aa6411c1 --- /dev/null +++ b/stage0/stdlib/Lean/Meta/Tactic/Grind/EMatchTheorem.c @@ -0,0 +1,18515 @@ +// Lean compiler output +// Module: Lean.Meta.Tactic.Grind.EMatchTheorem +// Imports: Lean.HeadIndex Lean.PrettyPrinter Lean.Util.FoldConsts Lean.Util.CollectFVars Lean.Meta.Basic Lean.Meta.InferType +#include +#if defined(__clang__) +#pragma clang diagnostic ignored "-Wunused-parameter" +#pragma clang diagnostic ignored "-Wunused-label" +#elif defined(__GNUC__) && !defined(__CLANG__) +#pragma GCC diagnostic ignored "-Wunused-parameter" +#pragma GCC diagnostic ignored "-Wunused-label" +#pragma GCC diagnostic ignored "-Wunused-but-set-variable" +#endif +#ifdef __cplusplus +extern "C" { +#endif +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__2___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addEMatchTheorem___lambda__1___boxed(lean_object*); +lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ematchTheoremsExt; +lean_object* l___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2170_(lean_object*, lean_object*); +lean_object* l_Lean_Name_reprPrec(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_ppPattern___closed__4; +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__15; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare; +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__10; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___boxed(lean_object*, 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_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___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_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__5(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__3; +lean_object* l_Lean_getConstInfo___at_Lean_Meta_mkConstWithFreshMVarLevels___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__6(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); +static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___lambda__1___closed__1; +static lean_object* l_Lean_Meta_Grind_addEMatchTheorem___closed__1; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__6; +lean_object* lean_mk_empty_array_with_capacity(lean_object*); +lean_object* l_Lean_mkAppN(lean_object*, lean_object*); +size_t lean_usize_shift_right(size_t, size_t); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addEMatchTheorem___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_Meta_Grind_mkGroundPattern___closed__3; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage(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*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___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_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687_(lean_object*); +static lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___closed__1; +static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__1; +LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_Meta_Grind_NormalizePattern_main___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__8___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_List_all___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__15; +static lean_object* l_Lean_Meta_Grind_Origin_pp___rarg___closed__2; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_EMatchTheorem_getProofWithFreshMVarLevels___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__1___boxed(lean_object*, lean_object*); +static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___closed__2; +LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___spec__1___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___spec__2(lean_object*, lean_object*, lean_object*, size_t, size_t); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__19; +uint8_t lean_usize_dec_le(size_t, size_t); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addEMatchTheorem(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_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__7(size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__4(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_ConstantInfo_levelParams(lean_object*); +lean_object* l_Lean_indentD(lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__1; +LEAN_EXPORT uint8_t l_Lean_Meta_Grind_isPatternDontCare(lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__6___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint64_t lean_uint64_of_nat(lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__6(lean_object*, size_t, size_t, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instInhabitedOrigin; +size_t lean_uint64_to_usize(uint64_t); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__11; +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___spec__1___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__18___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__19(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Expr_isApp(lean_object*); +LEAN_EXPORT lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__2(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Expr_sort___override(lean_object*); +lean_object* l_Lean_MessageData_ofList(lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3___closed__2; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___closed__2; +lean_object* lean_array_push(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_addEMatchTheorem___closed__5; +lean_object* l_Array_toSubarray___rarg(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_addEMatchTheorem___closed__6; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__5; +size_t lean_usize_mul(size_t, size_t); +lean_object* l_Lean_RBNode_fold___at_Lean_RBMap_size___spec__1___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_RBTree_ofList___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__4(lean_object*); +static lean_object* l_Lean_Meta_Grind_Origin_pp___rarg___closed__3; +lean_object* lean_mk_array(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__5; +uint8_t lean_usize_dec_eq(size_t, size_t); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___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_PersistentHashMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____spec__1___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51_(lean_object*, lean_object*); +static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__1___closed__1; +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand_go___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__3(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_mapTR_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__8(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_ppPattern___closed__8; +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__16; +static lean_object* l_Lean_Meta_Grind_mkGroundPattern___closed__1; +static lean_object* l_Lean_Meta_Grind_instInhabitedOrigin___closed__1; +lean_object* l_Lean_Expr_bvar___override(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___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_array_fget(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___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_Grind_addEMatchTheorem___closed__11; +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__17; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); +static size_t l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__1; +static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__1___closed__2; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Expr_fvarId_x21(lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__10(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); +static size_t l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__2; +LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3(lean_object*, size_t, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__4; +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__9(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); +lean_object* l_Lean_Meta_forallTelescopeReducing___at_Lean_Meta_getParamNames___spec__2___rarg(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__4___boxed(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_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Expr_isAppOf(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Origin_pp(lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__6; +lean_object* l_Nat_nextPowerOfTwo_go(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__7; +lean_object* l_Lean_stringToMessageData(lean_object*); +static lean_object* l_Lean_Meta_Grind_NormalizePattern_main___closed__2; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instReprOrigin; +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___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_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___lambda__1___closed__2; +static lean_object* l_Lean_Meta_Grind_instReprOrigin___closed__1; +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__3___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__4(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_instInhabitedEMatchTheorem___closed__3; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instInhabitedEMatchTheorem; +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_Grind_addEMatchTheorem___closed__8; +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___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_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__2___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_addEMatchTheorem___closed__9; +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_ppPattern___closed__3; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isPatternDontCare___boxed(lean_object*); +uint8_t l_Lean_Expr_isBVar(lean_object*); +lean_object* l_List_mapTR_loop___at_Lean_mkConstWithLevelParams___spec__1(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_addEMatchTheorem___closed__2; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__4; +static lean_object* l_Lean_Meta_Grind_mkGroundPattern___closed__2; +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__2(lean_object*); +uint8_t l_Lean_Expr_hasMVar(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_EMatchTheorem_getProofWithFreshMVarLevels(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Origin_key___boxed(lean_object*); +LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__3(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare___closed__3; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_EMatchTheorems_insert(lean_object*, lean_object*); +lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Origin_pp___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__20(lean_object*, size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__10(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); +static lean_object* l_Lean_Meta_Grind_NormalizePattern_main___closed__3; +static lean_object* l_Lean_Meta_Grind_NormalizePattern_main___closed__4; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3___closed__3; +size_t lean_usize_of_nat(lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__6; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3___closed__1; +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___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_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__14___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__15___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_mapTR_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__6___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__7(lean_object*, lean_object*, size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___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_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__14; +static lean_object* l_Lean_Meta_Grind_ppPattern___closed__7; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__2; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_foundBVar___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldrM___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__2___boxed(lean_object*, lean_object*); +lean_object* lean_st_ref_take(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__2; +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__13; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__18; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__1; +uint8_t lean_expr_eqv(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_registerSimpleScopedEnvExtension___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isAtomicPattern___boxed(lean_object*); +lean_object* l_Lean_Meta_forallBoundedTelescope___at_Lean_Meta_arrowDomainsN___spec__6___rarg(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___closed__3; +uint64_t lean_uint64_shift_right(uint64_t, uint64_t); +lean_object* lean_nat_to_int(lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand_go___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__2(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_MessageData_ofSyntax(lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___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*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__8; +lean_object* lean_nat_div(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_NormalizePattern_main___closed__5; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addEMatchTheorem___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_Grind_addEMatchTheorem___lambda__3___closed__4; +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFunMask(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_groundPattern_x3f(lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_RBNode_findCore___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__2(lean_object*, lean_object*); +static lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__2___closed__2; +LEAN_EXPORT uint8_t l_Lean_Meta_Grind_addEMatchTheorem___lambda__1(lean_object*); +lean_object* l_Lean_MessageData_ofFormat(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFunMask___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__7; +lean_object* l_Lean_FVarId_getDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_outOfBounds___rarg(lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___closed__1; +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getEMatchTheorems(lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__8; +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___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_st_ref_get(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__18; +static lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__2___closed__1; +static lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__3___closed__2; +uint8_t l_List_isEmpty___rarg(lean_object*); +lean_object* lean_st_mk_ref(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isForbidden(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___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_array_to_list(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Expr_hasLooseBVars(lean_object*); +lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); +static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___closed__3; +uint8_t l_List_elem___at_Lean_Meta_Occurrences_contains___spec__1(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_ppPattern___closed__2; +lean_object* l_Lean_Expr_toHeadIndex(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___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___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__3; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_NormalizePattern_main(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_ppPattern___closed__11; +lean_object* l_Lean_mkConstWithLevelParams___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Expr_constName_x21(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___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; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFunMask___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFn_x3f(lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__5; +static lean_object* l_Lean_Meta_Grind_addEMatchTheorem___lambda__2___closed__1; +LEAN_EXPORT uint8_t l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__1(lean_object*, lean_object*); +lean_object* l_Lean_MessageData_ofConst(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getEMatchTheorems___boxed(lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__9; +lean_object* l_Lean_FVarId_getType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_mkAnnotation(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__9; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__13; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare___closed__2; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_foundBVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__1(lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_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_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__4(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*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__4(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); +lean_object* l_Lean_Expr_instantiateLevelParamsArray(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__12___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__13(lean_object*, lean_object*, size_t, size_t, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__4; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___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_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__2; +static lean_object* l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__6; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames; +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__3; +static lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__1___closed__1; +lean_object* l_Lean_CollectFVars_main(lean_object*, lean_object*); +extern lean_object* l_Lean_Meta_instMonadMetaM; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Origin_pp___rarg___lambda__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__10; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___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___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFunMask___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_instInhabitedEMatchTheorem___closed__1; +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_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*, lean_object*, lean_object*); +static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___closed__1; +static lean_object* l_Lean_Meta_Grind_Origin_key___closed__2; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__7; +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___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_usize_to_nat(size_t); +lean_object* l_Lean_Meta_instInhabitedMetaM___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___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*); +static lean_object* l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____spec__1___closed__2; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__14; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___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_Meta_Grind_EMatchTheorems_insert___closed__2; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___closed__1; +lean_object* l_Lean_MessageData_ofExpr(lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____spec__1; +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__3(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); +static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__4; +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__12(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__14___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__15(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__2(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isForbidden___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addEMatchTheorem___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_Grind_mkGroundPattern(lean_object*); +static lean_object* l_Lean_Meta_Grind_ppPattern___closed__9; +LEAN_EXPORT lean_object* l_List_mapTR_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__3(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldrM___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__2(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__5; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3___closed__4; +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__14; +LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isAtomicPattern(lean_object*); +LEAN_EXPORT lean_object* l_panic___at_Lean_Meta_Grind_addEMatchTheorem___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___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_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__7(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__16; +lean_object* l_Lean_PersistentHashMap_instInhabited(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_panic___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__1(lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__10___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__11(lean_object*, lean_object*, size_t, size_t, lean_object*); +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__20; +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___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_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__8; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__11; +uint8_t lean_nat_dec_eq(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__4; +uint8_t lean_nat_dec_lt(lean_object*, lean_object*); +static lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__2___closed__1; +lean_object* l_id___rarg___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__3; +lean_object* l_Lean_PersistentHashMap_mkEmptyEntries(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__6(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__20___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ppPattern(lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___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_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_EMatchTheorems_insert___closed__1; +LEAN_EXPORT uint8_t l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_foundBVar___spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__5(lean_object*, lean_object*, lean_object*); +lean_object* lean_array_set(lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3; +uint64_t l_Lean_Name_hash___override(lean_object*); +lean_object* l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Array_contains___at_Lean_registerInternalExceptionId___spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__3(lean_object*, lean_object*, lean_object*); +uint64_t lean_uint64_xor(uint64_t, uint64_t); +extern lean_object* l_Id_instMonad; +lean_object* l_Repr_addAppParen(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_panic_fn(lean_object*, lean_object*); +lean_object* l_Lean_annotation_x3f(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__6; +lean_object* lean_nat_sub(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_ScopedEnvExtension_getState___rarg(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Expr_getAppFn(lean_object*); +static lean_object* l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__1; +static lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__3___closed__1; +lean_object* lean_nat_mul(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__11(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); +lean_object* l_Lean_Meta_isTypeFormer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_panic___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__1___closed__2; +lean_object* l_Lean_ScopedEnvExtension_addCore___rarg(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*); +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); +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*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___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 l_Lean_LocalDecl_binderInfo(lean_object*); +lean_object* l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__12; +static lean_object* l_Lean_Meta_Grind_NormalizePattern_main___closed__1; +static lean_object* l_Lean_Meta_Grind_EMatchTheorems_insert___closed__3; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__17; +lean_object* l_List_reverse___rarg(lean_object*); +static lean_object* l_Lean_Meta_Grind_Origin_key___closed__1; +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__19; +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__11; +size_t lean_usize_sub(size_t, size_t); +lean_object* lean_array_mk(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_foundBVar___spec__1___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___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_Grind_ppPattern___closed__1; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__10; +static lean_object* l_Lean_Meta_Grind_addEMatchTheorem___closed__10; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__4(lean_object*, lean_object*); +size_t lean_usize_add(size_t, size_t); +LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__2; +lean_object* lean_array_uget(lean_object*, size_t); +lean_object* l_Lean_Expr_fvar___override(lean_object*); +static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___closed__1; +size_t lean_array_size(lean_object*); +static lean_object* l_Lean_Meta_Grind_addEMatchTheorem___closed__12; +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__10___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_panic___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__1___closed__1; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__13; +lean_object* l_instInhabitedOfMonad___rarg(lean_object*, lean_object*); +static lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__6___closed__1; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__3; +lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +size_t lean_usize_shift_left(size_t, size_t); +LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isGroundPattern(lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__18___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__19___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l___private_Lean_HeadIndex_0__Lean_beqHeadIndex____x40_Lean_HeadIndex___hyg_69_(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFn_x3f___boxed(lean_object*); +lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__9; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__15; +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_string_append(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_ppPattern___closed__10; +static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__2; +lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__2(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__17; +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__12; +extern lean_object* l_Lean_Name_instBEq; +lean_object* lean_array_get_size(lean_object*); +static lean_object* l_Lean_Meta_Grind_addEMatchTheorem___closed__3; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__16; +static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__3; +static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1; +static lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Origin_key(lean_object*); +lean_object* lean_infer_type(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isGroundPattern___boxed(lean_object*); +uint8_t lean_nat_dec_le(lean_object*, lean_object*); +uint8_t lean_usize_dec_lt(size_t, size_t); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__12___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__13___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFunMask___spec__1(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__8___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___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*); +static lean_object* l_Lean_Meta_Grind_ppPattern___closed__5; +lean_object* lean_nat_add(lean_object*, lean_object*); +lean_object* l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(lean_object*); +lean_object* l_Lean_MessageData_bracket(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___lambda__1___closed__1; +uint8_t l_Lean_Expr_isConst(lean_object*); +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___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*); +static lean_object* l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__5; +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__17(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint64_t l_Lean_HeadIndex_hash(lean_object*); +uint8_t l_Lean_Expr_isFVar(lean_object*); +static lean_object* l_Lean_Meta_Grind_EMatchTheorems_insert___closed__4; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getEMatchTheorems___rarg___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFunMask___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_ppPattern___closed__6; +lean_object* l_Lean_Meta_mkConstWithFreshMVarLevels(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_Origin_pp___rarg___closed__1; +static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___closed__2; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__8(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_groundPattern_x3f___boxed(lean_object*); +lean_object* lean_array_uset(lean_object*, size_t, lean_object*); +lean_object* l_Lean_MessageData_ofName(lean_object*); +lean_object* lean_array_get(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__12; +static lean_object* l_Lean_Meta_Grind_addEMatchTheorem___closed__4; +lean_object* l___private_Init_Data_Repr_0__Nat_reprFast(lean_object*); +static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare___closed__1; +static lean_object* l_Lean_Meta_Grind_addEMatchTheorem___closed__7; +static lean_object* l_Lean_Meta_Grind_instInhabitedEMatchTheorem___closed__2; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__7; +lean_object* l_Lean_Meta_isProof(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_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* l_ReaderT_instMonad___rarg(lean_object*); +size_t lean_usize_land(size_t, size_t); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__1; +extern lean_object* l_Lean_instHashableName; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getEMatchTheorems___rarg(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___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*); +uint8_t l_Array_isEmpty___rarg(lean_object*); +lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_openAbstractMVarsResult___spec__1(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* _init_l_Lean_Meta_Grind_instInhabitedOrigin___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(0); +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Grind_instInhabitedOrigin() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Meta_Grind_instInhabitedOrigin___closed__1; +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.Meta.Grind.Origin.decl", 27, 27); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__1; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__2; +x_2 = lean_box(1); +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__4() { +_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_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__5() { +_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_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.Meta.Grind.Origin.fvar", 27, 27); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__6; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__7; +x_2 = lean_box(1); +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__9() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.Meta.Grind.Origin.stx", 26, 26); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____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_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__10; +x_2 = lean_box(1); +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__12() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.Meta.Grind.Origin.other", 28, 28); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____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_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__4; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__13; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__15() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__14; +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_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__16() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__5; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__13; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__17() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__16; +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_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51_(lean_object* x_1, lean_object* x_2) { +_start: +{ +switch (lean_obj_tag(x_1)) { +case 0: +{ +lean_object* x_3; lean_object* x_4; uint8_t x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_3 = lean_ctor_get(x_1, 0); +lean_inc(x_3); +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_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____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_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____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); +x_11 = 0; +x_12 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_12, 0, x_10); +lean_ctor_set_uint8(x_12, sizeof(void*)*1, x_11); +x_13 = l_Repr_addAppParen(x_12, x_2); +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_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____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); +x_16 = 0; +x_17 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_17, 0, x_15); +lean_ctor_set_uint8(x_17, sizeof(void*)*1, x_16); +x_18 = l_Repr_addAppParen(x_17, x_2); +return x_18; +} +} +case 1: +{ +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_1, 0); +lean_inc(x_19); +lean_dec(x_1); +x_20 = lean_unsigned_to_nat(1024u); +x_21 = lean_nat_dec_le(x_20, x_2); +x_22 = l_Lean_Name_reprPrec(x_19, x_20); +x_23 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__8; +x_24 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set(x_24, 1, x_22); +if (x_21 == 0) +{ +lean_object* x_25; lean_object* x_26; uint8_t x_27; lean_object* x_28; lean_object* x_29; +x_25 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__4; +x_26 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set(x_26, 1, x_24); +x_27 = 0; +x_28 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_28, 0, x_26); +lean_ctor_set_uint8(x_28, sizeof(void*)*1, x_27); +x_29 = l_Repr_addAppParen(x_28, x_2); +return x_29; +} +else +{ +lean_object* x_30; lean_object* x_31; uint8_t x_32; lean_object* x_33; lean_object* x_34; +x_30 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__5; +x_31 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_31, 0, x_30); +lean_ctor_set(x_31, 1, x_24); +x_32 = 0; +x_33 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_33, 0, x_31); +lean_ctor_set_uint8(x_33, sizeof(void*)*1, x_32); +x_34 = l_Repr_addAppParen(x_33, x_2); +return x_34; +} +} +case 2: +{ +uint8_t x_35; +x_35 = !lean_is_exclusive(x_1); +if (x_35 == 0) +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; uint8_t x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_36 = lean_ctor_get(x_1, 0); +x_37 = lean_ctor_get(x_1, 1); +x_38 = lean_unsigned_to_nat(1024u); +x_39 = lean_nat_dec_le(x_38, x_2); +x_40 = l_Lean_Name_reprPrec(x_36, x_38); +x_41 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__11; +lean_ctor_set_tag(x_1, 5); +lean_ctor_set(x_1, 1, x_40); +lean_ctor_set(x_1, 0, x_41); +x_42 = lean_box(1); +x_43 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_43, 0, x_1); +lean_ctor_set(x_43, 1, x_42); +x_44 = l___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2170_(x_37, x_38); +x_45 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +if (x_39 == 0) +{ +lean_object* x_46; lean_object* x_47; uint8_t x_48; lean_object* x_49; lean_object* x_50; +x_46 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__4; +x_47 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_47, 0, x_46); +lean_ctor_set(x_47, 1, x_45); +x_48 = 0; +x_49 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_49, 0, x_47); +lean_ctor_set_uint8(x_49, sizeof(void*)*1, x_48); +x_50 = l_Repr_addAppParen(x_49, x_2); +return x_50; +} +else +{ +lean_object* x_51; lean_object* x_52; uint8_t x_53; lean_object* x_54; lean_object* x_55; +x_51 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____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_45); +x_53 = 0; +x_54 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_54, 0, x_52); +lean_ctor_set_uint8(x_54, sizeof(void*)*1, x_53); +x_55 = l_Repr_addAppParen(x_54, x_2); +return x_55; +} +} +else +{ +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; +x_56 = lean_ctor_get(x_1, 0); +x_57 = lean_ctor_get(x_1, 1); +lean_inc(x_57); +lean_inc(x_56); +lean_dec(x_1); +x_58 = lean_unsigned_to_nat(1024u); +x_59 = lean_nat_dec_le(x_58, x_2); +x_60 = l_Lean_Name_reprPrec(x_56, x_58); +x_61 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__11; +x_62 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_62, 0, x_61); +lean_ctor_set(x_62, 1, x_60); +x_63 = lean_box(1); +x_64 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_64, 0, x_62); +lean_ctor_set(x_64, 1, x_63); +x_65 = l___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2170_(x_57, x_58); +x_66 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_66, 0, x_64); +lean_ctor_set(x_66, 1, x_65); +if (x_59 == 0) +{ +lean_object* x_67; lean_object* x_68; uint8_t x_69; lean_object* x_70; lean_object* x_71; +x_67 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__4; +x_68 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_68, 0, x_67); +lean_ctor_set(x_68, 1, x_66); +x_69 = 0; +x_70 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_70, 0, x_68); +lean_ctor_set_uint8(x_70, sizeof(void*)*1, x_69); +x_71 = l_Repr_addAppParen(x_70, x_2); +return x_71; +} +else +{ +lean_object* x_72; lean_object* x_73; uint8_t x_74; lean_object* x_75; lean_object* x_76; +x_72 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__5; +x_73 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_73, 0, x_72); +lean_ctor_set(x_73, 1, x_66); +x_74 = 0; +x_75 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_75, 0, x_73); +lean_ctor_set_uint8(x_75, sizeof(void*)*1, x_74); +x_76 = l_Repr_addAppParen(x_75, x_2); +return x_76; +} +} +} +default: +{ +lean_object* x_77; uint8_t x_78; +x_77 = lean_unsigned_to_nat(1024u); +x_78 = lean_nat_dec_le(x_77, x_2); +if (x_78 == 0) +{ +lean_object* x_79; lean_object* x_80; +x_79 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__15; +x_80 = l_Repr_addAppParen(x_79, x_2); +return x_80; +} +else +{ +lean_object* x_81; lean_object* x_82; +x_81 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__17; +x_82 = l_Repr_addAppParen(x_81, x_2); +return x_82; +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51_(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind_instReprOrigin___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_instReprOrigin() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Meta_Grind_instReprOrigin___closed__1; +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_Origin_key___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("other", 5, 5); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_Origin_key___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_Grind_Origin_key___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Origin_key(lean_object* x_1) { +_start: +{ +if (lean_obj_tag(x_1) == 3) +{ +lean_object* x_2; +x_2 = l_Lean_Meta_Grind_Origin_key___closed__2; +return x_2; +} +else +{ +lean_object* x_3; +x_3 = lean_ctor_get(x_1, 0); +lean_inc(x_3); +return x_3; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Origin_key___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Lean_Meta_Grind_Origin_key(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Origin_pp___rarg___lambda__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_3 = lean_ctor_get(x_1, 0); +lean_inc(x_3); +lean_dec(x_1); +x_4 = lean_ctor_get(x_3, 1); +lean_inc(x_4); +lean_dec(x_3); +x_5 = l_Lean_MessageData_ofConst(x_2); +x_6 = lean_apply_2(x_4, lean_box(0), x_5); +return x_6; +} +} +static lean_object* _init_l_Lean_Meta_Grind_Origin_pp___rarg___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("[unknown]", 9, 9); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_Origin_pp___rarg___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_Origin_pp___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_Meta_Grind_Origin_pp___rarg___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_Origin_pp___rarg___closed__2; +x_2 = l_Lean_MessageData_ofFormat(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Origin_pp___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +switch (lean_obj_tag(x_4)) { +case 0: +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_5 = lean_ctor_get(x_4, 0); +lean_inc(x_5); +lean_dec(x_4); +x_6 = lean_ctor_get(x_1, 1); +lean_inc(x_6); +lean_inc(x_1); +x_7 = l_Lean_mkConstWithLevelParams___rarg(x_1, x_2, x_3, x_5); +x_8 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_Origin_pp___rarg___lambda__1), 2, 1); +lean_closure_set(x_8, 0, x_1); +x_9 = lean_apply_4(x_6, lean_box(0), lean_box(0), x_7, x_8); +return x_9; +} +case 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_dec(x_3); +lean_dec(x_2); +x_10 = lean_ctor_get(x_4, 0); +lean_inc(x_10); +lean_dec(x_4); +x_11 = lean_ctor_get(x_1, 0); +lean_inc(x_11); +lean_dec(x_1); +x_12 = lean_ctor_get(x_11, 1); +lean_inc(x_12); +lean_dec(x_11); +x_13 = l_Lean_Expr_fvar___override(x_10); +x_14 = l_Lean_MessageData_ofExpr(x_13); +x_15 = lean_apply_2(x_12, lean_box(0), x_14); +return x_15; +} +case 2: +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +lean_dec(x_3); +lean_dec(x_2); +x_16 = lean_ctor_get(x_4, 1); +lean_inc(x_16); +lean_dec(x_4); +x_17 = lean_ctor_get(x_1, 0); +lean_inc(x_17); +lean_dec(x_1); +x_18 = lean_ctor_get(x_17, 1); +lean_inc(x_18); +lean_dec(x_17); +x_19 = l_Lean_MessageData_ofSyntax(x_16); +x_20 = lean_apply_2(x_18, lean_box(0), x_19); +return x_20; +} +default: +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +lean_dec(x_3); +lean_dec(x_2); +x_21 = lean_ctor_get(x_1, 0); +lean_inc(x_21); +lean_dec(x_1); +x_22 = lean_ctor_get(x_21, 1); +lean_inc(x_22); +lean_dec(x_21); +x_23 = l_Lean_Meta_Grind_Origin_pp___rarg___closed__3; +x_24 = lean_apply_2(x_22, lean_box(0), x_23); +return x_24; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Origin_pp(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_Origin_pp___rarg), 4, 0); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Grind_instInhabitedEMatchTheorem___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = lean_mk_empty_array_with_capacity(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Grind_instInhabitedEMatchTheorem___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = l_Lean_Expr_bvar___override(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Grind_instInhabitedEMatchTheorem___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind_instInhabitedEMatchTheorem___closed__1; +x_3 = l_Lean_Meta_Grind_instInhabitedEMatchTheorem___closed__2; +x_4 = lean_unsigned_to_nat(0u); +x_5 = l_Lean_Meta_Grind_instInhabitedOrigin___closed__1; +x_6 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_6, 0, x_2); +lean_ctor_set(x_6, 1, x_3); +lean_ctor_set(x_6, 2, x_4); +lean_ctor_set(x_6, 3, x_1); +lean_ctor_set(x_6, 4, x_1); +lean_ctor_set(x_6, 5, x_5); +return x_6; +} +} +static lean_object* _init_l_Lean_Meta_Grind_instInhabitedEMatchTheorem() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Meta_Grind_instInhabitedEMatchTheorem___closed__3; +return x_1; +} +} +static lean_object* _init_l_panic___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__1___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Name_instBEq; +x_2 = l_Lean_instHashableName; +x_3 = l_Lean_PersistentHashMap_instInhabited(lean_box(0), lean_box(0), x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_panic___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Id_instMonad; +x_2 = l_panic___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__1___closed__1; +x_3 = l_instInhabitedOfMonad___rarg(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_panic___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__1(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_panic___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__1___closed__2; +x_3 = lean_panic_fn(x_2, x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_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_Grind_EMatchTheorems_insert___spec__3___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_Grind_EMatchTheorems_insert___spec__3___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_Grind_EMatchTheorems_insert___spec__3___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_Grind_EMatchTheorems_insert___spec__3(lean_object* x_1, size_t x_2, lean_object* x_3) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +uint8_t x_4; +x_4 = !lean_is_exclusive(x_1); +if (x_4 == 0) +{ +lean_object* x_5; size_t x_6; size_t x_7; size_t x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_5 = lean_ctor_get(x_1, 0); +x_6 = 5; +x_7 = l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__2; +x_8 = lean_usize_land(x_2, x_7); +x_9 = lean_usize_to_nat(x_8); +x_10 = lean_box(2); +x_11 = lean_array_get(x_10, x_5, x_9); +lean_dec(x_9); +lean_dec(x_5); +switch (lean_obj_tag(x_11)) { +case 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_name_eq(x_3, x_12); +lean_dec(x_12); +if (x_14 == 0) +{ +lean_object* x_15; +lean_dec(x_13); +lean_free_object(x_1); +x_15 = lean_box(0); +return x_15; +} +else +{ +lean_ctor_set_tag(x_1, 1); +lean_ctor_set(x_1, 0, x_13); +return x_1; +} +} +case 1: +{ +lean_object* x_16; size_t x_17; +lean_free_object(x_1); +x_16 = lean_ctor_get(x_11, 0); +lean_inc(x_16); +lean_dec(x_11); +x_17 = lean_usize_shift_right(x_2, x_6); +x_1 = x_16; +x_2 = x_17; +goto _start; +} +default: +{ +lean_object* x_19; +lean_free_object(x_1); +x_19 = lean_box(0); +return x_19; +} +} +} +else +{ +lean_object* x_20; size_t x_21; size_t x_22; size_t x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_20 = lean_ctor_get(x_1, 0); +lean_inc(x_20); +lean_dec(x_1); +x_21 = 5; +x_22 = l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__2; +x_23 = lean_usize_land(x_2, x_22); +x_24 = lean_usize_to_nat(x_23); +x_25 = lean_box(2); +x_26 = lean_array_get(x_25, x_20, x_24); +lean_dec(x_24); +lean_dec(x_20); +switch (lean_obj_tag(x_26)) { +case 0: +{ +lean_object* x_27; lean_object* x_28; uint8_t x_29; +x_27 = lean_ctor_get(x_26, 0); +lean_inc(x_27); +x_28 = lean_ctor_get(x_26, 1); +lean_inc(x_28); +lean_dec(x_26); +x_29 = lean_name_eq(x_3, x_27); +lean_dec(x_27); +if (x_29 == 0) +{ +lean_object* x_30; +lean_dec(x_28); +x_30 = lean_box(0); +return x_30; +} +else +{ +lean_object* x_31; +x_31 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_31, 0, x_28); +return x_31; +} +} +case 1: +{ +lean_object* x_32; size_t x_33; +x_32 = lean_ctor_get(x_26, 0); +lean_inc(x_32); +lean_dec(x_26); +x_33 = lean_usize_shift_right(x_2, x_21); +x_1 = x_32; +x_2 = x_33; +goto _start; +} +default: +{ +lean_object* x_35; +x_35 = lean_box(0); +return x_35; +} +} +} +} +else +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_36 = lean_ctor_get(x_1, 0); +lean_inc(x_36); +x_37 = lean_ctor_get(x_1, 1); +lean_inc(x_37); +lean_dec(x_1); +x_38 = lean_unsigned_to_nat(0u); +x_39 = l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__4(x_36, x_37, lean_box(0), x_38, x_3); +lean_dec(x_37); +lean_dec(x_36); +return x_39; +} +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__2(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint64_t x_3; size_t x_4; lean_object* x_5; +x_3 = l_Lean_Name_hash___override(x_2); +x_4 = lean_uint64_to_usize(x_3); +x_5 = l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3(x_1, x_4, x_2); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__7(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_Grind_EMatchTheorems_insert___spec__6(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_Grind_EMatchTheorems_insert___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; 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_Grind_EMatchTheorems_insert___spec__6___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_Grind_EMatchTheorems_insert___spec__6(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_Grind_EMatchTheorems_insert___spec__3___closed__2; +x_11 = lean_usize_land(x_2, x_10); +x_12 = lean_usize_to_nat(x_11); +x_13 = lean_array_get_size(x_7); +x_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_Grind_EMatchTheorems_insert___spec__6(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_Grind_EMatchTheorems_insert___spec__6(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_Grind_EMatchTheorems_insert___spec__3___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_Grind_EMatchTheorems_insert___spec__6(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_Grind_EMatchTheorems_insert___spec__8(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_Grind_EMatchTheorems_insert___spec__6___closed__1; +x_93 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__7(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_Grind_EMatchTheorems_insert___spec__8(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_Grind_EMatchTheorems_insert___spec__6___closed__1; +x_107 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__7(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_Grind_EMatchTheorems_insert___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint64_t x_4; size_t x_5; size_t x_6; lean_object* x_7; +x_4 = l_Lean_Name_hash___override(x_2); +x_5 = lean_uint64_to_usize(x_4); +x_6 = 1; +x_7 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__6(x_1, x_5, x_6, x_2, x_3); +return x_7; +} +} +static lean_object* _init_l_Lean_Meta_Grind_EMatchTheorems_insert___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.Meta.Tactic.Grind.EMatchTheorem", 36, 36); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_EMatchTheorems_insert___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.Meta.Grind.EMatchTheorems.insert", 37, 37); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_EMatchTheorems_insert___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("unreachable code has been reached", 33, 33); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_EMatchTheorems_insert___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_Meta_Grind_EMatchTheorems_insert___closed__1; +x_2 = l_Lean_Meta_Grind_EMatchTheorems_insert___closed__2; +x_3 = lean_unsigned_to_nat(64u); +x_4 = lean_unsigned_to_nat(6u); +x_5 = l_Lean_Meta_Grind_EMatchTheorems_insert___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_Meta_Grind_EMatchTheorems_insert(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_ctor_get(x_2, 4); +lean_inc(x_3); +if (lean_obj_tag(x_3) == 0) +{ +lean_object* x_4; lean_object* x_5; +lean_dec(x_2); +lean_dec(x_1); +x_4 = l_Lean_Meta_Grind_EMatchTheorems_insert___closed__4; +x_5 = l_panic___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__1(x_4); +return x_5; +} +else +{ +lean_object* x_6; +x_6 = lean_ctor_get(x_3, 0); +lean_inc(x_6); +if (lean_obj_tag(x_6) == 2) +{ +uint8_t x_7; +x_7 = !lean_is_exclusive(x_2); +if (x_7 == 0) +{ +lean_object* x_8; uint8_t x_9; +x_8 = lean_ctor_get(x_2, 4); +lean_dec(x_8); +x_9 = !lean_is_exclusive(x_3); +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_3, 1); +x_11 = lean_ctor_get(x_3, 0); +lean_dec(x_11); +x_12 = lean_ctor_get(x_6, 0); +lean_inc(x_12); +lean_dec(x_6); +lean_ctor_set(x_2, 4, x_10); +lean_inc(x_1); +x_13 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__2(x_1, x_12); +if (lean_obj_tag(x_13) == 0) +{ +lean_object* x_14; lean_object* x_15; +x_14 = lean_box(0); +lean_ctor_set(x_3, 1, x_14); +lean_ctor_set(x_3, 0, x_2); +x_15 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__5(x_1, x_12, x_3); +return x_15; +} +else +{ +lean_object* x_16; lean_object* x_17; +x_16 = lean_ctor_get(x_13, 0); +lean_inc(x_16); +lean_dec(x_13); +lean_ctor_set(x_3, 1, x_16); +lean_ctor_set(x_3, 0, x_2); +x_17 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__5(x_1, x_12, x_3); +return x_17; +} +} +else +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_18 = lean_ctor_get(x_3, 1); +lean_inc(x_18); +lean_dec(x_3); +x_19 = lean_ctor_get(x_6, 0); +lean_inc(x_19); +lean_dec(x_6); +lean_ctor_set(x_2, 4, x_18); +lean_inc(x_1); +x_20 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__2(x_1, x_19); +if (lean_obj_tag(x_20) == 0) +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_box(0); +x_22 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_22, 0, x_2); +lean_ctor_set(x_22, 1, x_21); +x_23 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__5(x_1, x_19, 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, 0); +lean_inc(x_24); +lean_dec(x_20); +x_25 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_25, 0, x_2); +lean_ctor_set(x_25, 1, x_24); +x_26 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__5(x_1, x_19, x_25); +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; +x_27 = lean_ctor_get(x_2, 0); +x_28 = lean_ctor_get(x_2, 1); +x_29 = lean_ctor_get(x_2, 2); +x_30 = lean_ctor_get(x_2, 3); +x_31 = lean_ctor_get(x_2, 5); +lean_inc(x_31); +lean_inc(x_30); +lean_inc(x_29); +lean_inc(x_28); +lean_inc(x_27); +lean_dec(x_2); +x_32 = lean_ctor_get(x_3, 1); +lean_inc(x_32); +if (lean_is_exclusive(x_3)) { + lean_ctor_release(x_3, 0); + lean_ctor_release(x_3, 1); + x_33 = x_3; +} else { + lean_dec_ref(x_3); + x_33 = lean_box(0); +} +x_34 = lean_ctor_get(x_6, 0); +lean_inc(x_34); +lean_dec(x_6); +x_35 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_35, 0, x_27); +lean_ctor_set(x_35, 1, x_28); +lean_ctor_set(x_35, 2, x_29); +lean_ctor_set(x_35, 3, x_30); +lean_ctor_set(x_35, 4, x_32); +lean_ctor_set(x_35, 5, x_31); +lean_inc(x_1); +x_36 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__2(x_1, x_34); +if (lean_obj_tag(x_36) == 0) +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_37 = lean_box(0); +if (lean_is_scalar(x_33)) { + x_38 = lean_alloc_ctor(1, 2, 0); +} else { + x_38 = x_33; +} +lean_ctor_set(x_38, 0, x_35); +lean_ctor_set(x_38, 1, x_37); +x_39 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__5(x_1, x_34, x_38); +return x_39; +} +else +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_40 = lean_ctor_get(x_36, 0); +lean_inc(x_40); +lean_dec(x_36); +if (lean_is_scalar(x_33)) { + x_41 = lean_alloc_ctor(1, 2, 0); +} else { + x_41 = x_33; +} +lean_ctor_set(x_41, 0, x_35); +lean_ctor_set(x_41, 1, x_40); +x_42 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__5(x_1, x_34, x_41); +return x_42; +} +} +} +else +{ +lean_object* x_43; lean_object* x_44; +lean_dec(x_6); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_43 = l_Lean_Meta_Grind_EMatchTheorems_insert___closed__4; +x_44 = l_panic___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__1(x_43); +return x_44; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; +x_6 = l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__4(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_Grind_EMatchTheorems_insert___spec__3___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_Grind_EMatchTheorems_insert___spec__3(x_1, x_4, x_3); +lean_dec(x_3); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__2___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__2(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +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_Grind_EMatchTheorems_insert___spec__7(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_Grind_EMatchTheorems_insert___spec__6___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_Grind_EMatchTheorems_insert___spec__6(x_1, x_6, x_7, x_4, x_5); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_EMatchTheorem_getProofWithFreshMVarLevels(lean_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_ctor_get(x_1, 1); +lean_inc(x_7); +x_8 = l_Lean_Expr_isConst(x_7); +if (x_8 == 0) +{ +lean_object* x_9; uint8_t x_10; +x_9 = lean_ctor_get(x_1, 0); +lean_inc(x_9); +lean_dec(x_1); +x_10 = l_Array_isEmpty___rarg(x_9); +if (x_10 == 0) +{ +size_t x_11; size_t x_12; lean_object* x_13; uint8_t x_14; +x_11 = lean_array_size(x_9); +x_12 = 0; +lean_inc(x_9); +x_13 = l_Array_mapMUnsafe_map___at_Lean_Meta_openAbstractMVarsResult___spec__1(x_11, x_12, x_9, x_2, x_3, x_4, x_5, x_6); +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; +x_15 = lean_ctor_get(x_13, 0); +x_16 = l_Lean_Expr_instantiateLevelParamsArray(x_7, x_9, x_15); +lean_dec(x_7); +lean_ctor_set(x_13, 0, x_16); +return x_13; +} +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_17 = lean_ctor_get(x_13, 0); +x_18 = lean_ctor_get(x_13, 1); +lean_inc(x_18); +lean_inc(x_17); +lean_dec(x_13); +x_19 = l_Lean_Expr_instantiateLevelParamsArray(x_7, x_9, x_17); +lean_dec(x_7); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_19); +lean_ctor_set(x_20, 1, x_18); +return x_20; +} +} +else +{ +lean_object* x_21; +lean_dec(x_9); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_7); +lean_ctor_set(x_21, 1, x_6); +return x_21; +} +} +else +{ +lean_object* x_22; uint8_t x_23; +x_22 = lean_ctor_get(x_1, 0); +lean_inc(x_22); +lean_dec(x_1); +x_23 = l_Array_isEmpty___rarg(x_22); +if (x_23 == 0) +{ +size_t x_24; size_t x_25; lean_object* x_26; uint8_t x_27; +x_24 = lean_array_size(x_22); +x_25 = 0; +lean_inc(x_22); +x_26 = l_Array_mapMUnsafe_map___at_Lean_Meta_openAbstractMVarsResult___spec__1(x_24, x_25, x_22, x_2, x_3, x_4, x_5, x_6); +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); +x_29 = l_Lean_Expr_instantiateLevelParamsArray(x_7, x_22, x_28); +lean_dec(x_7); +lean_ctor_set(x_26, 0, x_29); +return x_26; +} +else +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_30 = lean_ctor_get(x_26, 0); +x_31 = lean_ctor_get(x_26, 1); +lean_inc(x_31); +lean_inc(x_30); +lean_dec(x_26); +x_32 = l_Lean_Expr_instantiateLevelParamsArray(x_7, x_22, x_30); +lean_dec(x_7); +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_31); +return x_33; +} +} +else +{ +lean_object* x_34; lean_object* x_35; +lean_dec(x_22); +x_34 = l_Lean_Expr_constName_x21(x_7); +lean_inc(x_34); +x_35 = l_Lean_getConstInfo___at_Lean_Meta_mkConstWithFreshMVarLevels___spec__1(x_34, x_2, x_3, x_4, x_5, x_6); +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; lean_object* x_39; uint8_t x_40; +x_37 = lean_ctor_get(x_35, 0); +x_38 = lean_ctor_get(x_35, 1); +x_39 = l_Lean_ConstantInfo_levelParams(x_37); +lean_dec(x_37); +x_40 = l_List_isEmpty___rarg(x_39); +lean_dec(x_39); +if (x_40 == 0) +{ +lean_object* x_41; +lean_free_object(x_35); +lean_dec(x_7); +x_41 = l_Lean_Meta_mkConstWithFreshMVarLevels(x_34, x_2, x_3, x_4, x_5, x_38); +return x_41; +} +else +{ +lean_dec(x_34); +lean_ctor_set(x_35, 0, x_7); +return x_35; +} +} +else +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; uint8_t x_45; +x_42 = lean_ctor_get(x_35, 0); +x_43 = lean_ctor_get(x_35, 1); +lean_inc(x_43); +lean_inc(x_42); +lean_dec(x_35); +x_44 = l_Lean_ConstantInfo_levelParams(x_42); +lean_dec(x_42); +x_45 = l_List_isEmpty___rarg(x_44); +lean_dec(x_44); +if (x_45 == 0) +{ +lean_object* x_46; +lean_dec(x_7); +x_46 = l_Lean_Meta_mkConstWithFreshMVarLevels(x_34, x_2, x_3, x_4, x_5, x_43); +return x_46; +} +else +{ +lean_object* x_47; +lean_dec(x_34); +x_47 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_47, 0, x_7); +lean_ctor_set(x_47, 1, x_43); +return x_47; +} +} +} +else +{ +uint8_t x_48; +lean_dec(x_34); +lean_dec(x_7); +x_48 = !lean_is_exclusive(x_35); +if (x_48 == 0) +{ +return x_35; +} +else +{ +lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_49 = lean_ctor_get(x_35, 0); +x_50 = lean_ctor_get(x_35, 1); +lean_inc(x_50); +lean_inc(x_49); +lean_dec(x_35); +x_51 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_51, 0, x_49); +lean_ctor_set(x_51, 1, x_50); +return x_51; +} +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_EMatchTheorem_getProofWithFreshMVarLevels___boxed(lean_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_Grind_EMatchTheorem_getProofWithFreshMVarLevels(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_7; +} +} +static lean_object* _init_l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____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_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____spec__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____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_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____spec__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____spec__1___closed__2; +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("_private", 8, 8); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___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_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___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_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean", 4, 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__2; +x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__3; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Meta", 4, 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__4; +x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__5; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Tactic", 6, 6); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__6; +x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__7; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__9() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Grind", 5, 5); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__8; +x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__9; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__11() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("EMatchTheorem", 13, 13); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__10; +x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__11; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__12; +x_2 = lean_unsigned_to_nat(0u); +x_3 = l_Lean_Name_num___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__13; +x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__3; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__15() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__14; +x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__5; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__16() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__15; +x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__9; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__17() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("ematchTheoremsExt", 17, 17); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__18() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__16; +x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__17; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__19() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_EMatchTheorems_insert), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__20() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_id___rarg___boxed), 1, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687_(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__18; +x_3 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__19; +x_4 = l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____spec__1; +x_5 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__20; +x_6 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_6, 0, x_2); +lean_ctor_set(x_6, 1, x_3); +lean_ctor_set(x_6, 2, x_4); +lean_ctor_set(x_6, 3, x_5); +x_7 = l_Lean_registerSimpleScopedEnvExtension___rarg(x_6, x_1); +return x_7; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Eq", 2, 2); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("HEq", 3, 3); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___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_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__3; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Iff", 3, 3); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___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_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__5; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("And", 3, 3); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___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_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__7; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__9() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Or", 2, 2); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___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_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__9; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__11() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Not", 3, 3); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___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_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__11; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___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_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__12; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__10; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__13; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__15() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__8; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__14; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__16() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__6; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__15; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__17() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__4; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__16; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__18() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__2; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__17; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__19() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__18; +x_2 = lean_array_mk(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames() { +_start: +{ +lean_object* x_1; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__19; +return x_1; +} +} +LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isForbidden(lean_object* x_1) { +_start: +{ +lean_object* x_2; uint8_t x_3; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames; +x_3 = l_Array_contains___at_Lean_registerInternalExceptionId___spec__1(x_2, x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isForbidden___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isForbidden(x_1); +lean_dec(x_1); +x_3 = lean_box(x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("[grind_dontcare]", 16, 16); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare___closed__2; +x_3 = l_Lean_Expr_const___override(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare() { +_start: +{ +lean_object* x_1; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare___closed__3; +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_mkGroundPattern___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("grind", 5, 5); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_mkGroundPattern___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("ground_pat", 10, 10); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_mkGroundPattern___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_mkGroundPattern___closed__1; +x_2 = l_Lean_Meta_Grind_mkGroundPattern___closed__2; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkGroundPattern(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Lean_Meta_Grind_mkGroundPattern___closed__3; +x_3 = l_Lean_mkAnnotation(x_2, x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_groundPattern_x3f(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Lean_Meta_Grind_mkGroundPattern___closed__3; +x_3 = l_Lean_annotation_x3f(x_2, x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_groundPattern_x3f___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Lean_Meta_Grind_groundPattern_x3f(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isGroundPattern(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Lean_Meta_Grind_groundPattern_x3f(x_1); +if (lean_obj_tag(x_2) == 0) +{ +uint8_t x_3; +x_3 = 0; +return x_3; +} +else +{ +uint8_t x_4; +lean_dec(x_2); +x_4 = 1; +return x_4; +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isGroundPattern___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isGroundPattern(x_1); +lean_dec(x_1); +x_3 = lean_box(x_2); +return x_3; +} +} +LEAN_EXPORT uint8_t l_Lean_Meta_Grind_isPatternDontCare(lean_object* x_1) { +_start: +{ +lean_object* x_2; uint8_t x_3; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare; +x_3 = lean_expr_eqv(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isPatternDontCare___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = l_Lean_Meta_Grind_isPatternDontCare(x_1); +lean_dec(x_1); +x_3 = lean_box(x_2); +return x_3; +} +} +LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isAtomicPattern(lean_object* x_1) { +_start: +{ +uint8_t x_2; +x_2 = l_Lean_Expr_isBVar(x_1); +if (x_2 == 0) +{ +lean_object* x_3; uint8_t x_4; +x_3 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare; +x_4 = lean_expr_eqv(x_1, x_3); +if (x_4 == 0) +{ +uint8_t x_5; +x_5 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isGroundPattern(x_1); +return x_5; +} +else +{ +uint8_t x_6; +x_6 = 1; +return x_6; +} +} +else +{ +uint8_t x_7; +x_7 = 1; +return x_7; +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isAtomicPattern___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isAtomicPattern(x_1); +lean_dec(x_1); +x_3 = lean_box(x_2); +return x_3; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(" ", 1, 1); +return x_1; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___closed__1; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___closed__2; +x_2 = l_Lean_MessageData_ofFormat(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_4 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___closed__3; +x_5 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_5, 0, x_2); +lean_ctor_set(x_5, 1, x_4); +x_6 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_6, 0, x_5); +lean_ctor_set(x_6, 1, x_1); +x_7 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_7, 0, x_6); +return x_7; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___boxed), 3, 0); +return x_1; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("(", 1, 1); +return x_1; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(")", 1, 1); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___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) { +_start: +{ +uint8_t x_7; +x_7 = lean_usize_dec_lt(x_5, x_4); +if (x_7 == 0) +{ +return x_6; +} +else +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_8 = lean_array_uget(x_3, x_5); +lean_inc(x_8); +x_9 = l_Lean_Meta_Grind_ppPattern(x_8); +x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1; +x_11 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isAtomicPattern(x_8); +lean_dec(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 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2; +x_13 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3; +x_14 = l_Lean_MessageData_bracket(x_12, x_9, x_13); +x_15 = lean_box(0); +x_16 = lean_apply_3(x_10, x_14, x_6, x_15); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_17; +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +lean_dec(x_16); +return x_17; +} +else +{ +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_5, x_19); +x_5 = x_20; +x_6 = x_18; +goto _start; +} +} +else +{ +lean_object* x_22; lean_object* x_23; +x_22 = lean_box(0); +x_23 = lean_apply_3(x_10, x_9, x_6, x_22); +if (lean_obj_tag(x_23) == 0) +{ +lean_object* x_24; +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +lean_dec(x_23); +return x_24; +} +else +{ +lean_object* x_25; size_t x_26; size_t x_27; +x_25 = lean_ctor_get(x_23, 0); +lean_inc(x_25); +lean_dec(x_23); +x_26 = 1; +x_27 = lean_usize_add(x_5, x_26); +x_5 = x_27; +x_6 = x_25; +goto _start; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { +_start: +{ +uint8_t x_7; +x_7 = lean_usize_dec_lt(x_5, x_4); +if (x_7 == 0) +{ +return x_6; +} +else +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_8 = lean_array_uget(x_3, x_5); +lean_inc(x_8); +x_9 = l_Lean_Meta_Grind_ppPattern(x_8); +x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1; +x_11 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isAtomicPattern(x_8); +lean_dec(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 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2; +x_13 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3; +x_14 = l_Lean_MessageData_bracket(x_12, x_9, x_13); +x_15 = lean_box(0); +x_16 = lean_apply_3(x_10, x_14, x_6, x_15); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_17; +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +lean_dec(x_16); +return x_17; +} +else +{ +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_5, x_19); +x_5 = x_20; +x_6 = x_18; +goto _start; +} +} +else +{ +lean_object* x_22; lean_object* x_23; +x_22 = lean_box(0); +x_23 = lean_apply_3(x_10, x_9, x_6, x_22); +if (lean_obj_tag(x_23) == 0) +{ +lean_object* x_24; +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +lean_dec(x_23); +return x_24; +} +else +{ +lean_object* x_25; size_t x_26; size_t x_27; +x_25 = lean_ctor_get(x_23, 0); +lean_inc(x_25); +lean_dec(x_23); +x_26 = 1; +x_27 = lean_usize_add(x_5, x_26); +x_5 = x_27; +x_6 = x_25; +goto _start; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { +_start: +{ +uint8_t x_7; +x_7 = lean_usize_dec_lt(x_5, x_4); +if (x_7 == 0) +{ +return x_6; +} +else +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_8 = lean_array_uget(x_3, x_5); +lean_inc(x_8); +x_9 = l_Lean_Meta_Grind_ppPattern(x_8); +x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1; +x_11 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isAtomicPattern(x_8); +lean_dec(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 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2; +x_13 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3; +x_14 = l_Lean_MessageData_bracket(x_12, x_9, x_13); +x_15 = lean_box(0); +x_16 = lean_apply_3(x_10, x_14, x_6, x_15); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_17; +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +lean_dec(x_16); +return x_17; +} +else +{ +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_5, x_19); +x_5 = x_20; +x_6 = x_18; +goto _start; +} +} +else +{ +lean_object* x_22; lean_object* x_23; +x_22 = lean_box(0); +x_23 = lean_apply_3(x_10, x_9, x_6, x_22); +if (lean_obj_tag(x_23) == 0) +{ +lean_object* x_24; +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +lean_dec(x_23); +return x_24; +} +else +{ +lean_object* x_25; size_t x_26; size_t x_27; +x_25 = lean_ctor_get(x_23, 0); +lean_inc(x_25); +lean_dec(x_23); +x_26 = 1; +x_27 = lean_usize_add(x_5, x_26); +x_5 = x_27; +x_6 = x_25; +goto _start; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { +_start: +{ +uint8_t x_7; +x_7 = lean_usize_dec_lt(x_5, x_4); +if (x_7 == 0) +{ +return x_6; +} +else +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_8 = lean_array_uget(x_3, x_5); +lean_inc(x_8); +x_9 = l_Lean_Meta_Grind_ppPattern(x_8); +x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1; +x_11 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isAtomicPattern(x_8); +lean_dec(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 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2; +x_13 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3; +x_14 = l_Lean_MessageData_bracket(x_12, x_9, x_13); +x_15 = lean_box(0); +x_16 = lean_apply_3(x_10, x_14, x_6, x_15); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_17; +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +lean_dec(x_16); +return x_17; +} +else +{ +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_5, x_19); +x_5 = x_20; +x_6 = x_18; +goto _start; +} +} +else +{ +lean_object* x_22; lean_object* x_23; +x_22 = lean_box(0); +x_23 = lean_apply_3(x_10, x_9, x_6, x_22); +if (lean_obj_tag(x_23) == 0) +{ +lean_object* x_24; +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +lean_dec(x_23); +return x_24; +} +else +{ +lean_object* x_25; size_t x_26; size_t x_27; +x_25 = lean_ctor_get(x_23, 0); +lean_inc(x_25); +lean_dec(x_23); +x_26 = 1; +x_27 = lean_usize_add(x_5, x_26); +x_5 = x_27; +x_6 = x_25; +goto _start; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { +_start: +{ +uint8_t x_7; +x_7 = lean_usize_dec_lt(x_5, x_4); +if (x_7 == 0) +{ +return x_6; +} +else +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_8 = lean_array_uget(x_3, x_5); +lean_inc(x_8); +x_9 = l_Lean_Meta_Grind_ppPattern(x_8); +x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1; +x_11 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isAtomicPattern(x_8); +lean_dec(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 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2; +x_13 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3; +x_14 = l_Lean_MessageData_bracket(x_12, x_9, x_13); +x_15 = lean_box(0); +x_16 = lean_apply_3(x_10, x_14, x_6, x_15); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_17; +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +lean_dec(x_16); +return x_17; +} +else +{ +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_5, x_19); +x_5 = x_20; +x_6 = x_18; +goto _start; +} +} +else +{ +lean_object* x_22; lean_object* x_23; +x_22 = lean_box(0); +x_23 = lean_apply_3(x_10, x_9, x_6, x_22); +if (lean_obj_tag(x_23) == 0) +{ +lean_object* x_24; +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +lean_dec(x_23); +return x_24; +} +else +{ +lean_object* x_25; size_t x_26; size_t x_27; +x_25 = lean_ctor_get(x_23, 0); +lean_inc(x_25); +lean_dec(x_23); +x_26 = 1; +x_27 = lean_usize_add(x_5, x_26); +x_5 = x_27; +x_6 = x_25; +goto _start; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { +_start: +{ +uint8_t x_7; +x_7 = lean_usize_dec_lt(x_5, x_4); +if (x_7 == 0) +{ +return x_6; +} +else +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_8 = lean_array_uget(x_3, x_5); +lean_inc(x_8); +x_9 = l_Lean_Meta_Grind_ppPattern(x_8); +x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1; +x_11 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isAtomicPattern(x_8); +lean_dec(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 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2; +x_13 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3; +x_14 = l_Lean_MessageData_bracket(x_12, x_9, x_13); +x_15 = lean_box(0); +x_16 = lean_apply_3(x_10, x_14, x_6, x_15); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_17; +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +lean_dec(x_16); +return x_17; +} +else +{ +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_5, x_19); +x_5 = x_20; +x_6 = x_18; +goto _start; +} +} +else +{ +lean_object* x_22; lean_object* x_23; +x_22 = lean_box(0); +x_23 = lean_apply_3(x_10, x_9, x_6, x_22); +if (lean_obj_tag(x_23) == 0) +{ +lean_object* x_24; +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +lean_dec(x_23); +return x_24; +} +else +{ +lean_object* x_25; size_t x_26; size_t x_27; +x_25 = lean_ctor_get(x_23, 0); +lean_inc(x_25); +lean_dec(x_23); +x_26 = 1; +x_27 = lean_usize_add(x_5, x_26); +x_5 = x_27; +x_6 = x_25; +goto _start; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { +_start: +{ +uint8_t x_7; +x_7 = lean_usize_dec_lt(x_5, x_4); +if (x_7 == 0) +{ +return x_6; +} +else +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_8 = lean_array_uget(x_3, x_5); +lean_inc(x_8); +x_9 = l_Lean_Meta_Grind_ppPattern(x_8); +x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1; +x_11 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isAtomicPattern(x_8); +lean_dec(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 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2; +x_13 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3; +x_14 = l_Lean_MessageData_bracket(x_12, x_9, x_13); +x_15 = lean_box(0); +x_16 = lean_apply_3(x_10, x_14, x_6, x_15); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_17; +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +lean_dec(x_16); +return x_17; +} +else +{ +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_5, x_19); +x_5 = x_20; +x_6 = x_18; +goto _start; +} +} +else +{ +lean_object* x_22; lean_object* x_23; +x_22 = lean_box(0); +x_23 = lean_apply_3(x_10, x_9, x_6, x_22); +if (lean_obj_tag(x_23) == 0) +{ +lean_object* x_24; +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +lean_dec(x_23); +return x_24; +} +else +{ +lean_object* x_25; size_t x_26; size_t x_27; +x_25 = lean_ctor_get(x_23, 0); +lean_inc(x_25); +lean_dec(x_23); +x_26 = 1; +x_27 = lean_usize_add(x_5, x_26); +x_5 = x_27; +x_6 = x_25; +goto _start; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__8(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { +_start: +{ +uint8_t x_7; +x_7 = lean_usize_dec_lt(x_5, x_4); +if (x_7 == 0) +{ +return x_6; +} +else +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_8 = lean_array_uget(x_3, x_5); +lean_inc(x_8); +x_9 = l_Lean_Meta_Grind_ppPattern(x_8); +x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1; +x_11 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isAtomicPattern(x_8); +lean_dec(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 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2; +x_13 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3; +x_14 = l_Lean_MessageData_bracket(x_12, x_9, x_13); +x_15 = lean_box(0); +x_16 = lean_apply_3(x_10, x_14, x_6, x_15); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_17; +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +lean_dec(x_16); +return x_17; +} +else +{ +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_5, x_19); +x_5 = x_20; +x_6 = x_18; +goto _start; +} +} +else +{ +lean_object* x_22; lean_object* x_23; +x_22 = lean_box(0); +x_23 = lean_apply_3(x_10, x_9, x_6, x_22); +if (lean_obj_tag(x_23) == 0) +{ +lean_object* x_24; +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +lean_dec(x_23); +return x_24; +} +else +{ +lean_object* x_25; size_t x_26; size_t x_27; +x_25 = lean_ctor_get(x_23, 0); +lean_inc(x_25); +lean_dec(x_23); +x_26 = 1; +x_27 = lean_usize_add(x_5, x_26); +x_5 = x_27; +x_6 = x_25; +goto _start; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__9(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { +_start: +{ +uint8_t x_7; +x_7 = lean_usize_dec_lt(x_5, x_4); +if (x_7 == 0) +{ +return x_6; +} +else +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_8 = lean_array_uget(x_3, x_5); +lean_inc(x_8); +x_9 = l_Lean_Meta_Grind_ppPattern(x_8); +x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1; +x_11 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isAtomicPattern(x_8); +lean_dec(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 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2; +x_13 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3; +x_14 = l_Lean_MessageData_bracket(x_12, x_9, x_13); +x_15 = lean_box(0); +x_16 = lean_apply_3(x_10, x_14, x_6, x_15); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_17; +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +lean_dec(x_16); +return x_17; +} +else +{ +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_5, x_19); +x_5 = x_20; +x_6 = x_18; +goto _start; +} +} +else +{ +lean_object* x_22; lean_object* x_23; +x_22 = lean_box(0); +x_23 = lean_apply_3(x_10, x_9, x_6, x_22); +if (lean_obj_tag(x_23) == 0) +{ +lean_object* x_24; +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +lean_dec(x_23); +return x_24; +} +else +{ +lean_object* x_25; size_t x_26; size_t x_27; +x_25 = lean_ctor_get(x_23, 0); +lean_inc(x_25); +lean_dec(x_23); +x_26 = 1; +x_27 = lean_usize_add(x_5, x_26); +x_5 = x_27; +x_6 = x_25; +goto _start; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__10(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { +_start: +{ +uint8_t x_7; +x_7 = lean_usize_dec_lt(x_5, x_4); +if (x_7 == 0) +{ +return x_6; +} +else +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_8 = lean_array_uget(x_3, x_5); +lean_inc(x_8); +x_9 = l_Lean_Meta_Grind_ppPattern(x_8); +x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1; +x_11 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isAtomicPattern(x_8); +lean_dec(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 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2; +x_13 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3; +x_14 = l_Lean_MessageData_bracket(x_12, x_9, x_13); +x_15 = lean_box(0); +x_16 = lean_apply_3(x_10, x_14, x_6, x_15); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_17; +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +lean_dec(x_16); +return x_17; +} +else +{ +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_5, x_19); +x_5 = x_20; +x_6 = x_18; +goto _start; +} +} +else +{ +lean_object* x_22; lean_object* x_23; +x_22 = lean_box(0); +x_23 = lean_apply_3(x_10, x_9, x_6, x_22); +if (lean_obj_tag(x_23) == 0) +{ +lean_object* x_24; +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +lean_dec(x_23); +return x_24; +} +else +{ +lean_object* x_25; size_t x_26; size_t x_27; +x_25 = lean_ctor_get(x_23, 0); +lean_inc(x_25); +lean_dec(x_23); +x_26 = 1; +x_27 = lean_usize_add(x_5, x_26); +x_5 = x_27; +x_6 = x_25; +goto _start; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__11(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { +_start: +{ +uint8_t x_7; +x_7 = lean_usize_dec_lt(x_5, x_4); +if (x_7 == 0) +{ +return x_6; +} +else +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_8 = lean_array_uget(x_3, x_5); +lean_inc(x_8); +x_9 = l_Lean_Meta_Grind_ppPattern(x_8); +x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1; +x_11 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isAtomicPattern(x_8); +lean_dec(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 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2; +x_13 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3; +x_14 = l_Lean_MessageData_bracket(x_12, x_9, x_13); +x_15 = lean_box(0); +x_16 = lean_apply_3(x_10, x_14, x_6, x_15); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_17; +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +lean_dec(x_16); +return x_17; +} +else +{ +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_5, x_19); +x_5 = x_20; +x_6 = x_18; +goto _start; +} +} +else +{ +lean_object* x_22; lean_object* x_23; +x_22 = lean_box(0); +x_23 = lean_apply_3(x_10, x_9, x_6, x_22); +if (lean_obj_tag(x_23) == 0) +{ +lean_object* x_24; +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +lean_dec(x_23); +return x_24; +} +else +{ +lean_object* x_25; size_t x_26; size_t x_27; +x_25 = lean_ctor_get(x_23, 0); +lean_inc(x_25); +lean_dec(x_23); +x_26 = 1; +x_27 = lean_usize_add(x_5, x_26); +x_5 = x_27; +x_6 = x_25; +goto _start; +} +} +} +} +} +static lean_object* _init_l_Lean_Meta_Grind_ppPattern___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("#", 1, 1); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_ppPattern___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_ppPattern___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Grind_ppPattern___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("", 0, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_ppPattern___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_ppPattern___closed__3; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Grind_ppPattern___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_levelZero; +x_2 = l_Lean_Expr_sort___override(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Grind_ppPattern___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("\?", 1, 1); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_ppPattern___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_ppPattern___closed__6; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Grind_ppPattern___closed__8() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("`[", 2, 2); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_ppPattern___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_ppPattern___closed__8; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Grind_ppPattern___closed__10() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("]", 1, 1); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_ppPattern___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_ppPattern___closed__10; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ppPattern(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Lean_Meta_Grind_groundPattern_x3f(x_1); +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_3; uint8_t x_4; +x_3 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare; +x_4 = lean_expr_eqv(x_1, x_3); +if (x_4 == 0) +{ +switch (lean_obj_tag(x_1)) { +case 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; +x_5 = lean_ctor_get(x_1, 0); +lean_inc(x_5); +lean_dec(x_1); +x_6 = l___private_Init_Data_Repr_0__Nat_reprFast(x_5); +x_7 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_7, 0, x_6); +x_8 = l_Lean_MessageData_ofFormat(x_7); +x_9 = l_Lean_Meta_Grind_ppPattern___closed__2; +x_10 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_10, 0, x_9); +lean_ctor_set(x_10, 1, x_8); +x_11 = l_Lean_Meta_Grind_ppPattern___closed__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; +} +case 1: +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; size_t x_26; size_t x_27; lean_object* x_28; +x_13 = l_Lean_Expr_getAppFn(x_1); +x_14 = l_Lean_MessageData_ofExpr(x_13); +x_15 = l_Lean_Meta_Grind_ppPattern___closed__4; +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 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_17, 0, x_16); +lean_ctor_set(x_17, 1, x_15); +x_18 = lean_unsigned_to_nat(0u); +x_19 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_18); +x_20 = l_Lean_Meta_Grind_ppPattern___closed__5; +lean_inc(x_19); +x_21 = lean_mk_array(x_19, x_20); +x_22 = lean_unsigned_to_nat(1u); +x_23 = lean_nat_sub(x_19, x_22); +lean_dec(x_19); +x_24 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_21, x_23); +x_25 = lean_box(0); +x_26 = lean_array_size(x_24); +x_27 = 0; +x_28 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1(x_24, x_25, x_24, x_26, x_27, x_17); +lean_dec(x_24); +return x_28; +} +case 2: +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; 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_29 = l_Lean_Expr_getAppFn(x_1); +x_30 = l_Lean_MessageData_ofExpr(x_29); +x_31 = l_Lean_Meta_Grind_ppPattern___closed__4; +x_32 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_30); +x_33 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_31); +x_34 = lean_unsigned_to_nat(0u); +x_35 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_34); +x_36 = l_Lean_Meta_Grind_ppPattern___closed__5; +lean_inc(x_35); +x_37 = lean_mk_array(x_35, x_36); +x_38 = lean_unsigned_to_nat(1u); +x_39 = lean_nat_sub(x_35, x_38); +lean_dec(x_35); +x_40 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_37, x_39); +x_41 = lean_box(0); +x_42 = lean_array_size(x_40); +x_43 = 0; +x_44 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__2(x_40, x_41, x_40, x_42, x_43, x_33); +lean_dec(x_40); +return x_44; +} +case 3: +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; size_t x_58; size_t x_59; lean_object* x_60; +x_45 = l_Lean_Expr_getAppFn(x_1); +x_46 = l_Lean_MessageData_ofExpr(x_45); +x_47 = l_Lean_Meta_Grind_ppPattern___closed__4; +x_48 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_48, 0, x_47); +lean_ctor_set(x_48, 1, x_46); +x_49 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_49, 0, x_48); +lean_ctor_set(x_49, 1, x_47); +x_50 = lean_unsigned_to_nat(0u); +x_51 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_50); +x_52 = l_Lean_Meta_Grind_ppPattern___closed__5; +lean_inc(x_51); +x_53 = lean_mk_array(x_51, x_52); +x_54 = lean_unsigned_to_nat(1u); +x_55 = lean_nat_sub(x_51, x_54); +lean_dec(x_51); +x_56 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_53, x_55); +x_57 = lean_box(0); +x_58 = lean_array_size(x_56); +x_59 = 0; +x_60 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__3(x_56, x_57, x_56, x_58, x_59, x_49); +lean_dec(x_56); +return x_60; +} +case 4: +{ +lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; size_t x_74; size_t x_75; lean_object* x_76; +x_61 = l_Lean_Expr_getAppFn(x_1); +x_62 = l_Lean_MessageData_ofExpr(x_61); +x_63 = l_Lean_Meta_Grind_ppPattern___closed__4; +x_64 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_64, 0, x_63); +lean_ctor_set(x_64, 1, x_62); +x_65 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_65, 0, x_64); +lean_ctor_set(x_65, 1, x_63); +x_66 = lean_unsigned_to_nat(0u); +x_67 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_66); +x_68 = l_Lean_Meta_Grind_ppPattern___closed__5; +lean_inc(x_67); +x_69 = lean_mk_array(x_67, x_68); +x_70 = lean_unsigned_to_nat(1u); +x_71 = lean_nat_sub(x_67, x_70); +lean_dec(x_67); +x_72 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_69, x_71); +x_73 = lean_box(0); +x_74 = lean_array_size(x_72); +x_75 = 0; +x_76 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__4(x_72, x_73, x_72, x_74, x_75, x_65); +lean_dec(x_72); +return x_76; +} +case 5: +{ +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; size_t x_90; size_t x_91; lean_object* x_92; +x_77 = l_Lean_Expr_getAppFn(x_1); +x_78 = l_Lean_MessageData_ofExpr(x_77); +x_79 = l_Lean_Meta_Grind_ppPattern___closed__4; +x_80 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_80, 0, x_79); +lean_ctor_set(x_80, 1, x_78); +x_81 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_81, 0, x_80); +lean_ctor_set(x_81, 1, x_79); +x_82 = lean_unsigned_to_nat(0u); +x_83 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_82); +x_84 = l_Lean_Meta_Grind_ppPattern___closed__5; +lean_inc(x_83); +x_85 = lean_mk_array(x_83, x_84); +x_86 = lean_unsigned_to_nat(1u); +x_87 = lean_nat_sub(x_83, x_86); +lean_dec(x_83); +x_88 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_85, x_87); +x_89 = lean_box(0); +x_90 = lean_array_size(x_88); +x_91 = 0; +x_92 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__5(x_88, x_89, x_88, x_90, x_91, x_81); +lean_dec(x_88); +return x_92; +} +case 6: +{ +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; size_t x_106; size_t x_107; lean_object* x_108; +x_93 = l_Lean_Expr_getAppFn(x_1); +x_94 = l_Lean_MessageData_ofExpr(x_93); +x_95 = l_Lean_Meta_Grind_ppPattern___closed__4; +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_alloc_ctor(7, 2, 0); +lean_ctor_set(x_97, 0, x_96); +lean_ctor_set(x_97, 1, x_95); +x_98 = lean_unsigned_to_nat(0u); +x_99 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_98); +x_100 = l_Lean_Meta_Grind_ppPattern___closed__5; +lean_inc(x_99); +x_101 = lean_mk_array(x_99, x_100); +x_102 = lean_unsigned_to_nat(1u); +x_103 = lean_nat_sub(x_99, x_102); +lean_dec(x_99); +x_104 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_101, x_103); +x_105 = lean_box(0); +x_106 = lean_array_size(x_104); +x_107 = 0; +x_108 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__6(x_104, x_105, x_104, x_106, x_107, x_97); +lean_dec(x_104); +return x_108; +} +case 7: +{ +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; size_t x_122; size_t x_123; lean_object* x_124; +x_109 = l_Lean_Expr_getAppFn(x_1); +x_110 = l_Lean_MessageData_ofExpr(x_109); +x_111 = l_Lean_Meta_Grind_ppPattern___closed__4; +x_112 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_112, 0, x_111); +lean_ctor_set(x_112, 1, x_110); +x_113 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_113, 0, x_112); +lean_ctor_set(x_113, 1, x_111); +x_114 = lean_unsigned_to_nat(0u); +x_115 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_114); +x_116 = l_Lean_Meta_Grind_ppPattern___closed__5; +lean_inc(x_115); +x_117 = lean_mk_array(x_115, x_116); +x_118 = lean_unsigned_to_nat(1u); +x_119 = lean_nat_sub(x_115, x_118); +lean_dec(x_115); +x_120 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_117, x_119); +x_121 = lean_box(0); +x_122 = lean_array_size(x_120); +x_123 = 0; +x_124 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__7(x_120, x_121, x_120, x_122, x_123, x_113); +lean_dec(x_120); +return x_124; +} +case 8: +{ +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; size_t x_138; size_t x_139; lean_object* x_140; +x_125 = l_Lean_Expr_getAppFn(x_1); +x_126 = l_Lean_MessageData_ofExpr(x_125); +x_127 = l_Lean_Meta_Grind_ppPattern___closed__4; +x_128 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_128, 0, x_127); +lean_ctor_set(x_128, 1, x_126); +x_129 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_129, 0, x_128); +lean_ctor_set(x_129, 1, x_127); +x_130 = lean_unsigned_to_nat(0u); +x_131 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_130); +x_132 = l_Lean_Meta_Grind_ppPattern___closed__5; +lean_inc(x_131); +x_133 = lean_mk_array(x_131, x_132); +x_134 = lean_unsigned_to_nat(1u); +x_135 = lean_nat_sub(x_131, x_134); +lean_dec(x_131); +x_136 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_133, x_135); +x_137 = lean_box(0); +x_138 = lean_array_size(x_136); +x_139 = 0; +x_140 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__8(x_136, x_137, x_136, x_138, x_139, x_129); +lean_dec(x_136); +return x_140; +} +case 9: +{ +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; size_t x_154; size_t x_155; lean_object* x_156; +x_141 = l_Lean_Expr_getAppFn(x_1); +x_142 = l_Lean_MessageData_ofExpr(x_141); +x_143 = l_Lean_Meta_Grind_ppPattern___closed__4; +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 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_145, 0, x_144); +lean_ctor_set(x_145, 1, x_143); +x_146 = lean_unsigned_to_nat(0u); +x_147 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_146); +x_148 = l_Lean_Meta_Grind_ppPattern___closed__5; +lean_inc(x_147); +x_149 = lean_mk_array(x_147, x_148); +x_150 = lean_unsigned_to_nat(1u); +x_151 = lean_nat_sub(x_147, x_150); +lean_dec(x_147); +x_152 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_149, x_151); +x_153 = lean_box(0); +x_154 = lean_array_size(x_152); +x_155 = 0; +x_156 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__9(x_152, x_153, x_152, x_154, x_155, x_145); +lean_dec(x_152); +return x_156; +} +case 10: +{ +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; size_t x_170; size_t x_171; lean_object* x_172; +x_157 = l_Lean_Expr_getAppFn(x_1); +x_158 = l_Lean_MessageData_ofExpr(x_157); +x_159 = l_Lean_Meta_Grind_ppPattern___closed__4; +x_160 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_160, 0, x_159); +lean_ctor_set(x_160, 1, x_158); +x_161 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_161, 0, x_160); +lean_ctor_set(x_161, 1, x_159); +x_162 = lean_unsigned_to_nat(0u); +x_163 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_162); +x_164 = l_Lean_Meta_Grind_ppPattern___closed__5; +lean_inc(x_163); +x_165 = lean_mk_array(x_163, x_164); +x_166 = lean_unsigned_to_nat(1u); +x_167 = lean_nat_sub(x_163, x_166); +lean_dec(x_163); +x_168 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_165, x_167); +x_169 = lean_box(0); +x_170 = lean_array_size(x_168); +x_171 = 0; +x_172 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__10(x_168, x_169, x_168, x_170, x_171, x_161); +lean_dec(x_168); +return x_172; +} +default: +{ +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; size_t x_186; size_t x_187; lean_object* x_188; +x_173 = l_Lean_Expr_getAppFn(x_1); +x_174 = l_Lean_MessageData_ofExpr(x_173); +x_175 = l_Lean_Meta_Grind_ppPattern___closed__4; +x_176 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_176, 0, x_175); +lean_ctor_set(x_176, 1, x_174); +x_177 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_177, 0, x_176); +lean_ctor_set(x_177, 1, x_175); +x_178 = lean_unsigned_to_nat(0u); +x_179 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_178); +x_180 = l_Lean_Meta_Grind_ppPattern___closed__5; +lean_inc(x_179); +x_181 = lean_mk_array(x_179, x_180); +x_182 = lean_unsigned_to_nat(1u); +x_183 = lean_nat_sub(x_179, x_182); +lean_dec(x_179); +x_184 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_181, x_183); +x_185 = lean_box(0); +x_186 = lean_array_size(x_184); +x_187 = 0; +x_188 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__11(x_184, x_185, x_184, x_186, x_187, x_177); +lean_dec(x_184); +return x_188; +} +} +} +else +{ +lean_object* x_189; +lean_dec(x_1); +x_189 = l_Lean_Meta_Grind_ppPattern___closed__7; +return x_189; +} +} +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_dec(x_1); +x_190 = lean_ctor_get(x_2, 0); +lean_inc(x_190); +lean_dec(x_2); +x_191 = l_Lean_MessageData_ofExpr(x_190); +x_192 = l_Lean_Meta_Grind_ppPattern___closed__9; +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_Meta_Grind_ppPattern___closed__11; +x_195 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_195, 0, x_193); +lean_ctor_set(x_195, 1, x_194); +return x_195; +} +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1(x_1, x_2, x_3); +lean_dec(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___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_7; size_t x_8; lean_object* x_9; +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_9 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1(x_1, x_2, x_3, x_7, x_8, x_6); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +size_t x_7; size_t x_8; lean_object* x_9; +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_9 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__2(x_1, x_2, x_3, x_7, x_8, x_6); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +size_t x_7; size_t x_8; lean_object* x_9; +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_9 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__3(x_1, x_2, x_3, x_7, x_8, x_6); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +size_t x_7; size_t x_8; lean_object* x_9; +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_9 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__4(x_1, x_2, x_3, x_7, x_8, x_6); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +size_t x_7; size_t x_8; lean_object* x_9; +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_9 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__5(x_1, x_2, x_3, x_7, x_8, x_6); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___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; size_t x_8; lean_object* x_9; +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_9 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__6(x_1, x_2, x_3, x_7, x_8, x_6); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +size_t x_7; size_t x_8; lean_object* x_9; +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_9 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__7(x_1, x_2, x_3, x_7, x_8, x_6); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +size_t x_7; size_t x_8; lean_object* x_9; +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_9 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__8(x_1, x_2, x_3, x_7, x_8, x_6); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___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) { +_start: +{ +size_t x_7; size_t x_8; lean_object* x_9; +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_9 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__9(x_1, x_2, x_3, x_7, x_8, x_6); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___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: +{ +size_t x_7; size_t x_8; lean_object* x_9; +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_9 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__10(x_1, x_2, x_3, x_7, x_8, x_6); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___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: +{ +size_t x_7; size_t x_8; lean_object* x_9; +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_9 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__11(x_1, x_2, x_3, x_7, x_8, x_6); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_9; +} +} +LEAN_EXPORT uint8_t l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +uint8_t x_3; +x_3 = 0; +return x_3; +} +else +{ +lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_4 = lean_ctor_get(x_2, 0); +x_5 = lean_ctor_get(x_2, 2); +x_6 = l___private_Lean_HeadIndex_0__Lean_beqHeadIndex____x40_Lean_HeadIndex___hyg_69_(x_4, x_1); +if (x_6 == 0) +{ +x_2 = x_5; +goto _start; +} +else +{ +uint8_t x_8; +x_8 = 1; +return x_8; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__4(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +return x_1; +} +else +{ +uint8_t x_3; +x_3 = !lean_is_exclusive(x_2); +if (x_3 == 0) +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; uint64_t x_7; uint64_t x_8; uint64_t x_9; uint64_t x_10; uint64_t x_11; uint64_t x_12; uint64_t x_13; size_t x_14; size_t x_15; size_t x_16; size_t x_17; size_t x_18; lean_object* x_19; lean_object* x_20; +x_4 = lean_ctor_get(x_2, 0); +x_5 = lean_ctor_get(x_2, 2); +x_6 = lean_array_get_size(x_1); +x_7 = l_Lean_HeadIndex_hash(x_4); +x_8 = 32; +x_9 = lean_uint64_shift_right(x_7, x_8); +x_10 = lean_uint64_xor(x_7, x_9); +x_11 = 16; +x_12 = lean_uint64_shift_right(x_10, x_11); +x_13 = lean_uint64_xor(x_10, x_12); +x_14 = lean_uint64_to_usize(x_13); +x_15 = lean_usize_of_nat(x_6); +lean_dec(x_6); +x_16 = 1; +x_17 = lean_usize_sub(x_15, x_16); +x_18 = lean_usize_land(x_14, x_17); +x_19 = lean_array_uget(x_1, x_18); +lean_ctor_set(x_2, 2, x_19); +x_20 = lean_array_uset(x_1, x_18, x_2); +x_1 = x_20; +x_2 = x_5; +goto _start; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint64_t x_26; uint64_t x_27; uint64_t x_28; uint64_t x_29; uint64_t x_30; uint64_t x_31; uint64_t x_32; size_t x_33; size_t x_34; size_t x_35; size_t x_36; size_t x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_22 = lean_ctor_get(x_2, 0); +x_23 = lean_ctor_get(x_2, 1); +x_24 = lean_ctor_get(x_2, 2); +lean_inc(x_24); +lean_inc(x_23); +lean_inc(x_22); +lean_dec(x_2); +x_25 = lean_array_get_size(x_1); +x_26 = l_Lean_HeadIndex_hash(x_22); +x_27 = 32; +x_28 = lean_uint64_shift_right(x_26, x_27); +x_29 = lean_uint64_xor(x_26, x_28); +x_30 = 16; +x_31 = lean_uint64_shift_right(x_29, x_30); +x_32 = lean_uint64_xor(x_29, x_31); +x_33 = lean_uint64_to_usize(x_32); +x_34 = lean_usize_of_nat(x_25); +lean_dec(x_25); +x_35 = 1; +x_36 = lean_usize_sub(x_34, x_35); +x_37 = lean_usize_land(x_33, x_36); +x_38 = lean_array_uget(x_1, x_37); +x_39 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_39, 0, x_22); +lean_ctor_set(x_39, 1, x_23); +lean_ctor_set(x_39, 2, x_38); +x_40 = lean_array_uset(x_1, x_37, x_39); +x_1 = x_40; +x_2 = x_24; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand_go___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; uint8_t x_5; +x_4 = lean_array_get_size(x_2); +x_5 = lean_nat_dec_lt(x_1, x_4); +lean_dec(x_4); +if (x_5 == 0) +{ +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +else +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_6 = lean_array_fget(x_2, x_1); +x_7 = lean_box(0); +x_8 = lean_array_fset(x_2, x_1, x_7); +x_9 = l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__4(x_3, x_6); +x_10 = lean_unsigned_to_nat(1u); +x_11 = lean_nat_add(x_1, x_10); +lean_dec(x_1); +x_1 = x_11; +x_2 = x_8; +x_3 = x_9; +goto _start; +} +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__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; +x_2 = lean_array_get_size(x_1); +x_3 = lean_unsigned_to_nat(2u); +x_4 = lean_nat_mul(x_2, x_3); +lean_dec(x_2); +x_5 = lean_box(0); +x_6 = lean_mk_array(x_4, x_5); +x_7 = lean_unsigned_to_nat(0u); +x_8 = l_Std_DHashMap_Internal_Raw_u2080_expand_go___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__3(x_7, x_1, x_6); +return x_8; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_8 = lean_st_ref_get(x_2, x_7); +x_9 = lean_ctor_get(x_8, 0); +lean_inc(x_9); +x_10 = lean_ctor_get(x_9, 1); +lean_inc(x_10); +lean_dec(x_9); +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; uint64_t x_16; uint64_t x_17; uint64_t x_18; uint64_t x_19; uint64_t x_20; uint64_t x_21; uint64_t x_22; size_t x_23; size_t x_24; size_t x_25; size_t x_26; size_t x_27; lean_object* x_28; uint8_t x_29; +x_12 = lean_ctor_get(x_8, 1); +x_13 = lean_ctor_get(x_8, 0); +lean_dec(x_13); +x_14 = lean_ctor_get(x_10, 1); +lean_inc(x_14); +lean_dec(x_10); +x_15 = lean_array_get_size(x_14); +x_16 = l_Lean_HeadIndex_hash(x_1); +x_17 = 32; +x_18 = lean_uint64_shift_right(x_16, x_17); +x_19 = lean_uint64_xor(x_16, x_18); +x_20 = 16; +x_21 = lean_uint64_shift_right(x_19, x_20); +x_22 = lean_uint64_xor(x_19, x_21); +x_23 = lean_uint64_to_usize(x_22); +x_24 = lean_usize_of_nat(x_15); +lean_dec(x_15); +x_25 = 1; +x_26 = lean_usize_sub(x_24, x_25); +x_27 = lean_usize_land(x_23, x_26); +x_28 = lean_array_uget(x_14, x_27); +lean_dec(x_14); +x_29 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__1(x_1, x_28); +lean_dec(x_28); +if (x_29 == 0) +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; +lean_free_object(x_8); +x_30 = lean_st_ref_take(x_2, x_12); +x_31 = lean_ctor_get(x_30, 0); +lean_inc(x_31); +x_32 = lean_ctor_get(x_30, 1); +lean_inc(x_32); +lean_dec(x_30); +x_33 = !lean_is_exclusive(x_31); +if (x_33 == 0) +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; +x_34 = lean_ctor_get(x_31, 0); +x_35 = lean_ctor_get(x_31, 1); +lean_inc(x_1); +x_36 = lean_array_push(x_34, x_1); +x_37 = !lean_is_exclusive(x_35); +if (x_37 == 0) +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; size_t x_41; size_t x_42; size_t x_43; lean_object* x_44; uint8_t x_45; +x_38 = lean_ctor_get(x_35, 0); +x_39 = lean_ctor_get(x_35, 1); +x_40 = lean_array_get_size(x_39); +x_41 = lean_usize_of_nat(x_40); +lean_dec(x_40); +x_42 = lean_usize_sub(x_41, x_25); +x_43 = lean_usize_land(x_23, x_42); +x_44 = lean_array_uget(x_39, x_43); +x_45 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__1(x_1, x_44); +if (x_45 == 0) +{ +lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; uint8_t x_56; +x_46 = lean_unsigned_to_nat(1u); +x_47 = lean_nat_add(x_38, x_46); +lean_dec(x_38); +x_48 = lean_box(0); +x_49 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_49, 0, x_1); +lean_ctor_set(x_49, 1, x_48); +lean_ctor_set(x_49, 2, x_44); +x_50 = lean_array_uset(x_39, x_43, x_49); +x_51 = lean_unsigned_to_nat(4u); +x_52 = lean_nat_mul(x_47, x_51); +x_53 = lean_unsigned_to_nat(3u); +x_54 = lean_nat_div(x_52, x_53); +lean_dec(x_52); +x_55 = lean_array_get_size(x_50); +x_56 = lean_nat_dec_le(x_54, x_55); +lean_dec(x_55); +lean_dec(x_54); +if (x_56 == 0) +{ +lean_object* x_57; lean_object* x_58; uint8_t x_59; +x_57 = l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__2(x_50); +lean_ctor_set(x_35, 1, x_57); +lean_ctor_set(x_35, 0, x_47); +lean_ctor_set(x_31, 0, x_36); +x_58 = lean_st_ref_set(x_2, x_31, x_32); +x_59 = !lean_is_exclusive(x_58); +if (x_59 == 0) +{ +lean_object* x_60; +x_60 = lean_ctor_get(x_58, 0); +lean_dec(x_60); +lean_ctor_set(x_58, 0, x_48); +return x_58; +} +else +{ +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(0, 2, 0); +lean_ctor_set(x_62, 0, x_48); +lean_ctor_set(x_62, 1, x_61); +return x_62; +} +} +else +{ +lean_object* x_63; uint8_t x_64; +lean_ctor_set(x_35, 1, x_50); +lean_ctor_set(x_35, 0, x_47); +lean_ctor_set(x_31, 0, x_36); +x_63 = lean_st_ref_set(x_2, x_31, x_32); +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(x_63, 0, x_48); +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(0, 2, 0); +lean_ctor_set(x_67, 0, x_48); +lean_ctor_set(x_67, 1, x_66); +return x_67; +} +} +} +else +{ +lean_object* x_68; uint8_t x_69; +lean_dec(x_44); +lean_dec(x_1); +lean_ctor_set(x_31, 0, x_36); +x_68 = lean_st_ref_set(x_2, x_31, x_32); +x_69 = !lean_is_exclusive(x_68); +if (x_69 == 0) +{ +lean_object* x_70; lean_object* x_71; +x_70 = lean_ctor_get(x_68, 0); +lean_dec(x_70); +x_71 = lean_box(0); +lean_ctor_set(x_68, 0, x_71); +return x_68; +} +else +{ +lean_object* x_72; lean_object* x_73; lean_object* x_74; +x_72 = lean_ctor_get(x_68, 1); +lean_inc(x_72); +lean_dec(x_68); +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_72); +return x_74; +} +} +} +else +{ +lean_object* x_75; lean_object* x_76; lean_object* x_77; size_t x_78; size_t x_79; size_t x_80; lean_object* x_81; uint8_t x_82; +x_75 = lean_ctor_get(x_35, 0); +x_76 = lean_ctor_get(x_35, 1); +lean_inc(x_76); +lean_inc(x_75); +lean_dec(x_35); +x_77 = lean_array_get_size(x_76); +x_78 = lean_usize_of_nat(x_77); +lean_dec(x_77); +x_79 = lean_usize_sub(x_78, x_25); +x_80 = lean_usize_land(x_23, x_79); +x_81 = lean_array_uget(x_76, x_80); +x_82 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__1(x_1, x_81); +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; uint8_t x_93; +x_83 = lean_unsigned_to_nat(1u); +x_84 = lean_nat_add(x_75, x_83); +lean_dec(x_75); +x_85 = lean_box(0); +x_86 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_86, 0, x_1); +lean_ctor_set(x_86, 1, x_85); +lean_ctor_set(x_86, 2, x_81); +x_87 = lean_array_uset(x_76, x_80, x_86); +x_88 = lean_unsigned_to_nat(4u); +x_89 = lean_nat_mul(x_84, x_88); +x_90 = lean_unsigned_to_nat(3u); +x_91 = lean_nat_div(x_89, x_90); +lean_dec(x_89); +x_92 = lean_array_get_size(x_87); +x_93 = lean_nat_dec_le(x_91, x_92); +lean_dec(x_92); +lean_dec(x_91); +if (x_93 == 0) +{ +lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; +x_94 = l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__2(x_87); +x_95 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_95, 0, x_84); +lean_ctor_set(x_95, 1, x_94); +lean_ctor_set(x_31, 1, x_95); +lean_ctor_set(x_31, 0, x_36); +x_96 = lean_st_ref_set(x_2, x_31, x_32); +x_97 = lean_ctor_get(x_96, 1); +lean_inc(x_97); +if (lean_is_exclusive(x_96)) { + lean_ctor_release(x_96, 0); + lean_ctor_release(x_96, 1); + x_98 = x_96; +} else { + lean_dec_ref(x_96); + x_98 = lean_box(0); +} +if (lean_is_scalar(x_98)) { + x_99 = lean_alloc_ctor(0, 2, 0); +} else { + x_99 = x_98; +} +lean_ctor_set(x_99, 0, x_85); +lean_ctor_set(x_99, 1, x_97); +return x_99; +} +else +{ +lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; +x_100 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_100, 0, x_84); +lean_ctor_set(x_100, 1, x_87); +lean_ctor_set(x_31, 1, x_100); +lean_ctor_set(x_31, 0, x_36); +x_101 = lean_st_ref_set(x_2, x_31, x_32); +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; +} else { + lean_dec_ref(x_101); + x_103 = lean_box(0); +} +if (lean_is_scalar(x_103)) { + x_104 = lean_alloc_ctor(0, 2, 0); +} else { + x_104 = x_103; +} +lean_ctor_set(x_104, 0, x_85); +lean_ctor_set(x_104, 1, x_102); +return x_104; +} +} +else +{ +lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; +lean_dec(x_81); +lean_dec(x_1); +x_105 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_105, 0, x_75); +lean_ctor_set(x_105, 1, x_76); +lean_ctor_set(x_31, 1, x_105); +lean_ctor_set(x_31, 0, x_36); +x_106 = lean_st_ref_set(x_2, x_31, x_32); +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); +} +x_109 = lean_box(0); +if (lean_is_scalar(x_108)) { + x_110 = lean_alloc_ctor(0, 2, 0); +} else { + x_110 = x_108; +} +lean_ctor_set(x_110, 0, x_109); +lean_ctor_set(x_110, 1, x_107); +return x_110; +} +} +} +else +{ +lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; size_t x_119; size_t x_120; size_t x_121; lean_object* x_122; uint8_t x_123; +x_111 = lean_ctor_get(x_31, 0); +x_112 = lean_ctor_get(x_31, 1); +x_113 = lean_ctor_get(x_31, 2); +lean_inc(x_113); +lean_inc(x_112); +lean_inc(x_111); +lean_dec(x_31); +lean_inc(x_1); +x_114 = lean_array_push(x_111, x_1); +x_115 = lean_ctor_get(x_112, 0); +lean_inc(x_115); +x_116 = lean_ctor_get(x_112, 1); +lean_inc(x_116); +if (lean_is_exclusive(x_112)) { + lean_ctor_release(x_112, 0); + lean_ctor_release(x_112, 1); + x_117 = x_112; +} else { + lean_dec_ref(x_112); + x_117 = lean_box(0); +} +x_118 = lean_array_get_size(x_116); +x_119 = lean_usize_of_nat(x_118); +lean_dec(x_118); +x_120 = lean_usize_sub(x_119, x_25); +x_121 = lean_usize_land(x_23, x_120); +x_122 = lean_array_uget(x_116, x_121); +x_123 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__1(x_1, x_122); +if (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; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; uint8_t x_134; +x_124 = lean_unsigned_to_nat(1u); +x_125 = lean_nat_add(x_115, x_124); +lean_dec(x_115); +x_126 = lean_box(0); +x_127 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_127, 0, x_1); +lean_ctor_set(x_127, 1, x_126); +lean_ctor_set(x_127, 2, x_122); +x_128 = lean_array_uset(x_116, x_121, x_127); +x_129 = lean_unsigned_to_nat(4u); +x_130 = lean_nat_mul(x_125, x_129); +x_131 = lean_unsigned_to_nat(3u); +x_132 = lean_nat_div(x_130, x_131); +lean_dec(x_130); +x_133 = lean_array_get_size(x_128); +x_134 = lean_nat_dec_le(x_132, x_133); +lean_dec(x_133); +lean_dec(x_132); +if (x_134 == 0) +{ +lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; +x_135 = l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__2(x_128); +if (lean_is_scalar(x_117)) { + x_136 = lean_alloc_ctor(0, 2, 0); +} else { + x_136 = x_117; +} +lean_ctor_set(x_136, 0, x_125); +lean_ctor_set(x_136, 1, x_135); +x_137 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_137, 0, x_114); +lean_ctor_set(x_137, 1, x_136); +lean_ctor_set(x_137, 2, x_113); +x_138 = lean_st_ref_set(x_2, x_137, x_32); +x_139 = lean_ctor_get(x_138, 1); +lean_inc(x_139); +if (lean_is_exclusive(x_138)) { + lean_ctor_release(x_138, 0); + lean_ctor_release(x_138, 1); + x_140 = x_138; +} else { + lean_dec_ref(x_138); + x_140 = lean_box(0); +} +if (lean_is_scalar(x_140)) { + x_141 = lean_alloc_ctor(0, 2, 0); +} else { + x_141 = x_140; +} +lean_ctor_set(x_141, 0, x_126); +lean_ctor_set(x_141, 1, x_139); +return x_141; +} +else +{ +lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; +if (lean_is_scalar(x_117)) { + x_142 = lean_alloc_ctor(0, 2, 0); +} else { + x_142 = x_117; +} +lean_ctor_set(x_142, 0, x_125); +lean_ctor_set(x_142, 1, x_128); +x_143 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_143, 0, x_114); +lean_ctor_set(x_143, 1, x_142); +lean_ctor_set(x_143, 2, x_113); +x_144 = lean_st_ref_set(x_2, x_143, x_32); +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_126); +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_object* x_152; lean_object* x_153; lean_object* x_154; +lean_dec(x_122); +lean_dec(x_1); +if (lean_is_scalar(x_117)) { + x_148 = lean_alloc_ctor(0, 2, 0); +} else { + x_148 = x_117; +} +lean_ctor_set(x_148, 0, x_115); +lean_ctor_set(x_148, 1, x_116); +x_149 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_149, 0, x_114); +lean_ctor_set(x_149, 1, x_148); +lean_ctor_set(x_149, 2, x_113); +x_150 = lean_st_ref_set(x_2, x_149, x_32); +x_151 = lean_ctor_get(x_150, 1); +lean_inc(x_151); +if (lean_is_exclusive(x_150)) { + lean_ctor_release(x_150, 0); + lean_ctor_release(x_150, 1); + x_152 = x_150; +} else { + lean_dec_ref(x_150); + x_152 = lean_box(0); +} +x_153 = lean_box(0); +if (lean_is_scalar(x_152)) { + x_154 = lean_alloc_ctor(0, 2, 0); +} else { + x_154 = x_152; +} +lean_ctor_set(x_154, 0, x_153); +lean_ctor_set(x_154, 1, x_151); +return x_154; +} +} +} +else +{ +lean_object* x_155; +lean_dec(x_1); +x_155 = lean_box(0); +lean_ctor_set(x_8, 0, x_155); +return x_8; +} +} +else +{ +lean_object* x_156; lean_object* x_157; lean_object* x_158; uint64_t x_159; uint64_t x_160; uint64_t x_161; uint64_t x_162; uint64_t x_163; uint64_t x_164; uint64_t x_165; size_t x_166; size_t x_167; size_t x_168; size_t x_169; size_t x_170; lean_object* x_171; uint8_t x_172; +x_156 = lean_ctor_get(x_8, 1); +lean_inc(x_156); +lean_dec(x_8); +x_157 = lean_ctor_get(x_10, 1); +lean_inc(x_157); +lean_dec(x_10); +x_158 = lean_array_get_size(x_157); +x_159 = l_Lean_HeadIndex_hash(x_1); +x_160 = 32; +x_161 = lean_uint64_shift_right(x_159, x_160); +x_162 = lean_uint64_xor(x_159, x_161); +x_163 = 16; +x_164 = lean_uint64_shift_right(x_162, x_163); +x_165 = lean_uint64_xor(x_162, x_164); +x_166 = lean_uint64_to_usize(x_165); +x_167 = lean_usize_of_nat(x_158); +lean_dec(x_158); +x_168 = 1; +x_169 = lean_usize_sub(x_167, x_168); +x_170 = lean_usize_land(x_166, x_169); +x_171 = lean_array_uget(x_157, x_170); +lean_dec(x_157); +x_172 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__1(x_1, x_171); +lean_dec(x_171); +if (x_172 == 0) +{ +lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; size_t x_185; size_t x_186; size_t x_187; lean_object* x_188; uint8_t x_189; +x_173 = lean_st_ref_take(x_2, x_156); +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); +if (lean_is_exclusive(x_174)) { + lean_ctor_release(x_174, 0); + lean_ctor_release(x_174, 1); + lean_ctor_release(x_174, 2); + x_179 = x_174; +} else { + lean_dec_ref(x_174); + x_179 = lean_box(0); +} +lean_inc(x_1); +x_180 = lean_array_push(x_176, x_1); +x_181 = lean_ctor_get(x_177, 0); +lean_inc(x_181); +x_182 = lean_ctor_get(x_177, 1); +lean_inc(x_182); +if (lean_is_exclusive(x_177)) { + lean_ctor_release(x_177, 0); + lean_ctor_release(x_177, 1); + x_183 = x_177; +} else { + lean_dec_ref(x_177); + x_183 = lean_box(0); +} +x_184 = lean_array_get_size(x_182); +x_185 = lean_usize_of_nat(x_184); +lean_dec(x_184); +x_186 = lean_usize_sub(x_185, x_168); +x_187 = lean_usize_land(x_166, x_186); +x_188 = lean_array_uget(x_182, x_187); +x_189 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__1(x_1, x_188); +if (x_189 == 0) +{ +lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; uint8_t x_200; +x_190 = lean_unsigned_to_nat(1u); +x_191 = lean_nat_add(x_181, x_190); +lean_dec(x_181); +x_192 = lean_box(0); +x_193 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_193, 0, x_1); +lean_ctor_set(x_193, 1, x_192); +lean_ctor_set(x_193, 2, x_188); +x_194 = lean_array_uset(x_182, x_187, x_193); +x_195 = lean_unsigned_to_nat(4u); +x_196 = lean_nat_mul(x_191, x_195); +x_197 = lean_unsigned_to_nat(3u); +x_198 = lean_nat_div(x_196, x_197); +lean_dec(x_196); +x_199 = lean_array_get_size(x_194); +x_200 = lean_nat_dec_le(x_198, x_199); +lean_dec(x_199); +lean_dec(x_198); +if (x_200 == 0) +{ +lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; +x_201 = l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__2(x_194); +if (lean_is_scalar(x_183)) { + x_202 = lean_alloc_ctor(0, 2, 0); +} else { + x_202 = x_183; +} +lean_ctor_set(x_202, 0, x_191); +lean_ctor_set(x_202, 1, x_201); +if (lean_is_scalar(x_179)) { + x_203 = lean_alloc_ctor(0, 3, 0); +} else { + x_203 = x_179; +} +lean_ctor_set(x_203, 0, x_180); +lean_ctor_set(x_203, 1, x_202); +lean_ctor_set(x_203, 2, x_178); +x_204 = lean_st_ref_set(x_2, x_203, x_175); +x_205 = lean_ctor_get(x_204, 1); +lean_inc(x_205); +if (lean_is_exclusive(x_204)) { + lean_ctor_release(x_204, 0); + lean_ctor_release(x_204, 1); + x_206 = x_204; +} else { + lean_dec_ref(x_204); + x_206 = lean_box(0); +} +if (lean_is_scalar(x_206)) { + x_207 = lean_alloc_ctor(0, 2, 0); +} else { + x_207 = x_206; +} +lean_ctor_set(x_207, 0, x_192); +lean_ctor_set(x_207, 1, x_205); +return x_207; +} +else +{ +lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; +if (lean_is_scalar(x_183)) { + x_208 = lean_alloc_ctor(0, 2, 0); +} else { + x_208 = x_183; +} +lean_ctor_set(x_208, 0, x_191); +lean_ctor_set(x_208, 1, x_194); +if (lean_is_scalar(x_179)) { + x_209 = lean_alloc_ctor(0, 3, 0); +} else { + x_209 = x_179; +} +lean_ctor_set(x_209, 0, x_180); +lean_ctor_set(x_209, 1, x_208); +lean_ctor_set(x_209, 2, x_178); +x_210 = lean_st_ref_set(x_2, x_209, x_175); +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_210); + x_212 = lean_box(0); +} +if (lean_is_scalar(x_212)) { + x_213 = lean_alloc_ctor(0, 2, 0); +} else { + x_213 = x_212; +} +lean_ctor_set(x_213, 0, x_192); +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_dec(x_188); +lean_dec(x_1); +if (lean_is_scalar(x_183)) { + x_214 = lean_alloc_ctor(0, 2, 0); +} else { + x_214 = x_183; +} +lean_ctor_set(x_214, 0, x_181); +lean_ctor_set(x_214, 1, x_182); +if (lean_is_scalar(x_179)) { + x_215 = lean_alloc_ctor(0, 3, 0); +} else { + x_215 = x_179; +} +lean_ctor_set(x_215, 0, x_180); +lean_ctor_set(x_215, 1, x_214); +lean_ctor_set(x_215, 2, x_178); +x_216 = lean_st_ref_set(x_2, x_215, x_175); +x_217 = lean_ctor_get(x_216, 1); +lean_inc(x_217); +if (lean_is_exclusive(x_216)) { + lean_ctor_release(x_216, 0); + lean_ctor_release(x_216, 1); + x_218 = x_216; +} else { + lean_dec_ref(x_216); + x_218 = lean_box(0); +} +x_219 = lean_box(0); +if (lean_is_scalar(x_218)) { + x_220 = lean_alloc_ctor(0, 2, 0); +} else { + x_220 = x_218; +} +lean_ctor_set(x_220, 0, x_219); +lean_ctor_set(x_220, 1, x_217); +return x_220; +} +} +else +{ +lean_object* x_221; lean_object* x_222; +lean_dec(x_1); +x_221 = lean_box(0); +x_222 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_222, 0, x_221); +lean_ctor_set(x_222, 1, x_156); +return x_222; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__1(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; +x_8 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_8; +} +} +LEAN_EXPORT uint8_t l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_foundBVar___spec__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +uint8_t x_3; +x_3 = 0; +return x_3; +} +else +{ +lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_4 = lean_ctor_get(x_2, 0); +x_5 = lean_ctor_get(x_2, 2); +x_6 = lean_nat_dec_eq(x_4, x_1); +if (x_6 == 0) +{ +x_2 = x_5; +goto _start; +} +else +{ +uint8_t x_8; +x_8 = 1; +return x_8; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_foundBVar(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_8 = lean_st_ref_get(x_2, x_7); +x_9 = lean_ctor_get(x_8, 0); +lean_inc(x_9); +x_10 = lean_ctor_get(x_9, 2); +lean_inc(x_10); +lean_dec(x_9); +x_11 = !lean_is_exclusive(x_8); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; uint64_t x_15; uint64_t x_16; uint64_t x_17; uint64_t x_18; uint64_t x_19; uint64_t x_20; uint64_t x_21; size_t x_22; size_t x_23; size_t x_24; size_t x_25; size_t x_26; lean_object* x_27; uint8_t x_28; lean_object* x_29; +x_12 = lean_ctor_get(x_8, 0); +lean_dec(x_12); +x_13 = lean_ctor_get(x_10, 1); +lean_inc(x_13); +lean_dec(x_10); +x_14 = lean_array_get_size(x_13); +x_15 = lean_uint64_of_nat(x_1); +x_16 = 32; +x_17 = lean_uint64_shift_right(x_15, x_16); +x_18 = lean_uint64_xor(x_15, x_17); +x_19 = 16; +x_20 = lean_uint64_shift_right(x_18, x_19); +x_21 = lean_uint64_xor(x_18, x_20); +x_22 = lean_uint64_to_usize(x_21); +x_23 = lean_usize_of_nat(x_14); +lean_dec(x_14); +x_24 = 1; +x_25 = lean_usize_sub(x_23, x_24); +x_26 = lean_usize_land(x_22, x_25); +x_27 = lean_array_uget(x_13, x_26); +lean_dec(x_13); +x_28 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_foundBVar___spec__1(x_1, x_27); +lean_dec(x_27); +x_29 = lean_box(x_28); +lean_ctor_set(x_8, 0, x_29); +return x_8; +} +else +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; uint64_t x_33; uint64_t x_34; uint64_t x_35; uint64_t x_36; uint64_t x_37; uint64_t x_38; uint64_t x_39; size_t x_40; size_t x_41; size_t x_42; size_t x_43; size_t x_44; lean_object* x_45; uint8_t x_46; lean_object* x_47; lean_object* x_48; +x_30 = lean_ctor_get(x_8, 1); +lean_inc(x_30); +lean_dec(x_8); +x_31 = lean_ctor_get(x_10, 1); +lean_inc(x_31); +lean_dec(x_10); +x_32 = lean_array_get_size(x_31); +x_33 = lean_uint64_of_nat(x_1); +x_34 = 32; +x_35 = lean_uint64_shift_right(x_33, x_34); +x_36 = lean_uint64_xor(x_33, x_35); +x_37 = 16; +x_38 = lean_uint64_shift_right(x_36, x_37); +x_39 = lean_uint64_xor(x_36, x_38); +x_40 = lean_uint64_to_usize(x_39); +x_41 = lean_usize_of_nat(x_32); +lean_dec(x_32); +x_42 = 1; +x_43 = lean_usize_sub(x_41, x_42); +x_44 = lean_usize_land(x_40, x_43); +x_45 = lean_array_uget(x_31, x_44); +lean_dec(x_31); +x_46 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_foundBVar___spec__1(x_1, x_45); +lean_dec(x_45); +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_30); +return x_48; +} +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_foundBVar___spec__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_foundBVar___spec__1(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_foundBVar___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; +x_8 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_foundBVar(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +if (lean_obj_tag(x_3) == 0) +{ +lean_dec(x_1); +return x_2; +} +else +{ +uint8_t x_4; +x_4 = !lean_is_exclusive(x_3); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint64_t x_9; uint64_t x_10; uint64_t x_11; uint64_t x_12; uint64_t x_13; uint64_t x_14; uint64_t x_15; size_t x_16; size_t x_17; size_t x_18; size_t x_19; size_t x_20; lean_object* x_21; lean_object* x_22; +x_5 = lean_ctor_get(x_3, 0); +x_6 = lean_ctor_get(x_3, 2); +x_7 = lean_array_get_size(x_2); +lean_inc(x_1); +lean_inc(x_5); +x_8 = lean_apply_1(x_1, x_5); +x_9 = lean_unbox_uint64(x_8); +lean_dec(x_8); +x_10 = 32; +x_11 = lean_uint64_shift_right(x_9, x_10); +x_12 = lean_uint64_xor(x_9, x_11); +x_13 = 16; +x_14 = lean_uint64_shift_right(x_12, x_13); +x_15 = lean_uint64_xor(x_12, x_14); +x_16 = lean_uint64_to_usize(x_15); +x_17 = lean_usize_of_nat(x_7); +lean_dec(x_7); +x_18 = 1; +x_19 = lean_usize_sub(x_17, x_18); +x_20 = lean_usize_land(x_16, x_19); +x_21 = lean_array_uget(x_2, x_20); +lean_ctor_set(x_3, 2, x_21); +x_22 = lean_array_uset(x_2, x_20, x_3); +x_2 = x_22; +x_3 = x_6; +goto _start; +} +else +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; uint64_t x_29; uint64_t x_30; uint64_t x_31; uint64_t x_32; uint64_t x_33; uint64_t x_34; uint64_t x_35; size_t x_36; size_t x_37; size_t x_38; size_t x_39; size_t x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_24 = lean_ctor_get(x_3, 0); +x_25 = lean_ctor_get(x_3, 1); +x_26 = lean_ctor_get(x_3, 2); +lean_inc(x_26); +lean_inc(x_25); +lean_inc(x_24); +lean_dec(x_3); +x_27 = lean_array_get_size(x_2); +lean_inc(x_1); +lean_inc(x_24); +x_28 = lean_apply_1(x_1, x_24); +x_29 = lean_unbox_uint64(x_28); +lean_dec(x_28); +x_30 = 32; +x_31 = lean_uint64_shift_right(x_29, x_30); +x_32 = lean_uint64_xor(x_29, x_31); +x_33 = 16; +x_34 = lean_uint64_shift_right(x_32, x_33); +x_35 = lean_uint64_xor(x_32, x_34); +x_36 = lean_uint64_to_usize(x_35); +x_37 = lean_usize_of_nat(x_27); +lean_dec(x_27); +x_38 = 1; +x_39 = lean_usize_sub(x_37, x_38); +x_40 = lean_usize_land(x_36, x_39); +x_41 = lean_array_uget(x_2, x_40); +x_42 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_42, 0, x_24); +lean_ctor_set(x_42, 1, x_25); +lean_ctor_set(x_42, 2, x_41); +x_43 = lean_array_uset(x_2, x_40, x_42); +x_2 = x_43; +x_3 = x_26; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__3___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__4(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +return x_1; +} +else +{ +uint8_t x_3; +x_3 = !lean_is_exclusive(x_2); +if (x_3 == 0) +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; uint64_t x_7; uint64_t x_8; uint64_t x_9; uint64_t x_10; uint64_t x_11; uint64_t x_12; uint64_t x_13; size_t x_14; size_t x_15; size_t x_16; size_t x_17; size_t x_18; lean_object* x_19; lean_object* x_20; +x_4 = lean_ctor_get(x_2, 0); +x_5 = lean_ctor_get(x_2, 2); +x_6 = lean_array_get_size(x_1); +x_7 = lean_uint64_of_nat(x_4); +x_8 = 32; +x_9 = lean_uint64_shift_right(x_7, x_8); +x_10 = lean_uint64_xor(x_7, x_9); +x_11 = 16; +x_12 = lean_uint64_shift_right(x_10, x_11); +x_13 = lean_uint64_xor(x_10, x_12); +x_14 = lean_uint64_to_usize(x_13); +x_15 = lean_usize_of_nat(x_6); +lean_dec(x_6); +x_16 = 1; +x_17 = lean_usize_sub(x_15, x_16); +x_18 = lean_usize_land(x_14, x_17); +x_19 = lean_array_uget(x_1, x_18); +lean_ctor_set(x_2, 2, x_19); +x_20 = lean_array_uset(x_1, x_18, x_2); +x_1 = x_20; +x_2 = x_5; +goto _start; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint64_t x_26; uint64_t x_27; uint64_t x_28; uint64_t x_29; uint64_t x_30; uint64_t x_31; uint64_t x_32; size_t x_33; size_t x_34; size_t x_35; size_t x_36; size_t x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_22 = lean_ctor_get(x_2, 0); +x_23 = lean_ctor_get(x_2, 1); +x_24 = lean_ctor_get(x_2, 2); +lean_inc(x_24); +lean_inc(x_23); +lean_inc(x_22); +lean_dec(x_2); +x_25 = lean_array_get_size(x_1); +x_26 = lean_uint64_of_nat(x_22); +x_27 = 32; +x_28 = lean_uint64_shift_right(x_26, x_27); +x_29 = lean_uint64_xor(x_26, x_28); +x_30 = 16; +x_31 = lean_uint64_shift_right(x_29, x_30); +x_32 = lean_uint64_xor(x_29, x_31); +x_33 = lean_uint64_to_usize(x_32); +x_34 = lean_usize_of_nat(x_25); +lean_dec(x_25); +x_35 = 1; +x_36 = lean_usize_sub(x_34, x_35); +x_37 = lean_usize_land(x_33, x_36); +x_38 = lean_array_uget(x_1, x_37); +x_39 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_39, 0, x_22); +lean_ctor_set(x_39, 1, x_23); +lean_ctor_set(x_39, 2, x_38); +x_40 = lean_array_uset(x_1, x_37, x_39); +x_1 = x_40; +x_2 = x_24; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand_go___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; uint8_t x_5; +x_4 = lean_array_get_size(x_2); +x_5 = lean_nat_dec_lt(x_1, x_4); +lean_dec(x_4); +if (x_5 == 0) +{ +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +else +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_6 = lean_array_fget(x_2, x_1); +x_7 = lean_box(0); +x_8 = lean_array_fset(x_2, x_1, x_7); +x_9 = l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__3___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__4(x_3, x_6); +x_10 = lean_unsigned_to_nat(1u); +x_11 = lean_nat_add(x_1, x_10); +lean_dec(x_1); +x_1 = x_11; +x_2 = x_8; +x_3 = x_9; +goto _start; +} +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__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; lean_object* x_7; lean_object* x_8; +x_2 = lean_array_get_size(x_1); +x_3 = lean_unsigned_to_nat(2u); +x_4 = lean_nat_mul(x_2, x_3); +lean_dec(x_2); +x_5 = lean_box(0); +x_6 = lean_mk_array(x_4, x_5); +x_7 = lean_unsigned_to_nat(0u); +x_8 = l_Std_DHashMap_Internal_Raw_u2080_expand_go___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__2(x_7, x_1, x_6); +return x_8; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar(lean_object* x_1, 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; +x_8 = lean_st_ref_take(x_2, x_7); +x_9 = lean_ctor_get(x_8, 0); +lean_inc(x_9); +x_10 = lean_ctor_get(x_9, 2); +lean_inc(x_10); +x_11 = lean_ctor_get(x_8, 1); +lean_inc(x_11); +lean_dec(x_8); +x_12 = !lean_is_exclusive(x_9); +if (x_12 == 0) +{ +lean_object* x_13; uint8_t x_14; +x_13 = lean_ctor_get(x_9, 2); +lean_dec(x_13); +x_14 = !lean_is_exclusive(x_10); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; uint64_t x_18; uint64_t x_19; uint64_t x_20; uint64_t x_21; uint64_t x_22; uint64_t x_23; uint64_t x_24; size_t x_25; size_t x_26; size_t x_27; size_t x_28; size_t x_29; lean_object* x_30; uint8_t x_31; +x_15 = lean_ctor_get(x_10, 0); +x_16 = lean_ctor_get(x_10, 1); +x_17 = lean_array_get_size(x_16); +x_18 = lean_uint64_of_nat(x_1); +x_19 = 32; +x_20 = lean_uint64_shift_right(x_18, x_19); +x_21 = lean_uint64_xor(x_18, x_20); +x_22 = 16; +x_23 = lean_uint64_shift_right(x_21, x_22); +x_24 = lean_uint64_xor(x_21, x_23); +x_25 = lean_uint64_to_usize(x_24); +x_26 = lean_usize_of_nat(x_17); +lean_dec(x_17); +x_27 = 1; +x_28 = lean_usize_sub(x_26, x_27); +x_29 = lean_usize_land(x_25, x_28); +x_30 = lean_array_uget(x_16, x_29); +x_31 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_foundBVar___spec__1(x_1, x_30); +if (x_31 == 0) +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; uint8_t x_42; +x_32 = lean_unsigned_to_nat(1u); +x_33 = lean_nat_add(x_15, x_32); +lean_dec(x_15); +x_34 = lean_box(0); +x_35 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_35, 0, x_1); +lean_ctor_set(x_35, 1, x_34); +lean_ctor_set(x_35, 2, x_30); +x_36 = lean_array_uset(x_16, x_29, x_35); +x_37 = lean_unsigned_to_nat(4u); +x_38 = lean_nat_mul(x_33, x_37); +x_39 = lean_unsigned_to_nat(3u); +x_40 = lean_nat_div(x_38, x_39); +lean_dec(x_38); +x_41 = lean_array_get_size(x_36); +x_42 = lean_nat_dec_le(x_40, x_41); +lean_dec(x_41); +lean_dec(x_40); +if (x_42 == 0) +{ +lean_object* x_43; lean_object* x_44; uint8_t x_45; +x_43 = l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__1(x_36); +lean_ctor_set(x_10, 1, x_43); +lean_ctor_set(x_10, 0, x_33); +x_44 = lean_st_ref_set(x_2, x_9, x_11); +x_45 = !lean_is_exclusive(x_44); +if (x_45 == 0) +{ +lean_object* x_46; +x_46 = lean_ctor_get(x_44, 0); +lean_dec(x_46); +lean_ctor_set(x_44, 0, x_34); +return x_44; +} +else +{ +lean_object* x_47; lean_object* x_48; +x_47 = lean_ctor_get(x_44, 1); +lean_inc(x_47); +lean_dec(x_44); +x_48 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_48, 0, x_34); +lean_ctor_set(x_48, 1, x_47); +return x_48; +} +} +else +{ +lean_object* x_49; uint8_t x_50; +lean_ctor_set(x_10, 1, x_36); +lean_ctor_set(x_10, 0, x_33); +x_49 = lean_st_ref_set(x_2, x_9, x_11); +x_50 = !lean_is_exclusive(x_49); +if (x_50 == 0) +{ +lean_object* x_51; +x_51 = lean_ctor_get(x_49, 0); +lean_dec(x_51); +lean_ctor_set(x_49, 0, x_34); +return x_49; +} +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_34); +lean_ctor_set(x_53, 1, x_52); +return x_53; +} +} +} +else +{ +lean_object* x_54; uint8_t x_55; +lean_dec(x_30); +lean_dec(x_1); +x_54 = lean_st_ref_set(x_2, x_9, x_11); +x_55 = !lean_is_exclusive(x_54); +if (x_55 == 0) +{ +lean_object* x_56; lean_object* x_57; +x_56 = lean_ctor_get(x_54, 0); +lean_dec(x_56); +x_57 = lean_box(0); +lean_ctor_set(x_54, 0, x_57); +return x_54; +} +else +{ +lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_58 = lean_ctor_get(x_54, 1); +lean_inc(x_58); +lean_dec(x_54); +x_59 = lean_box(0); +x_60 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_60, 0, x_59); +lean_ctor_set(x_60, 1, x_58); +return x_60; +} +} +} +else +{ +lean_object* x_61; lean_object* x_62; lean_object* x_63; uint64_t x_64; uint64_t x_65; uint64_t x_66; uint64_t x_67; uint64_t x_68; uint64_t x_69; uint64_t x_70; size_t x_71; size_t x_72; size_t x_73; size_t x_74; size_t x_75; lean_object* x_76; uint8_t x_77; +x_61 = lean_ctor_get(x_10, 0); +x_62 = lean_ctor_get(x_10, 1); +lean_inc(x_62); +lean_inc(x_61); +lean_dec(x_10); +x_63 = lean_array_get_size(x_62); +x_64 = lean_uint64_of_nat(x_1); +x_65 = 32; +x_66 = lean_uint64_shift_right(x_64, x_65); +x_67 = lean_uint64_xor(x_64, x_66); +x_68 = 16; +x_69 = lean_uint64_shift_right(x_67, x_68); +x_70 = lean_uint64_xor(x_67, x_69); +x_71 = lean_uint64_to_usize(x_70); +x_72 = lean_usize_of_nat(x_63); +lean_dec(x_63); +x_73 = 1; +x_74 = lean_usize_sub(x_72, x_73); +x_75 = lean_usize_land(x_71, x_74); +x_76 = lean_array_uget(x_62, x_75); +x_77 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_foundBVar___spec__1(x_1, x_76); +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; uint8_t x_88; +x_78 = lean_unsigned_to_nat(1u); +x_79 = lean_nat_add(x_61, x_78); +lean_dec(x_61); +x_80 = lean_box(0); +x_81 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_81, 0, x_1); +lean_ctor_set(x_81, 1, x_80); +lean_ctor_set(x_81, 2, x_76); +x_82 = lean_array_uset(x_62, x_75, x_81); +x_83 = lean_unsigned_to_nat(4u); +x_84 = lean_nat_mul(x_79, x_83); +x_85 = lean_unsigned_to_nat(3u); +x_86 = lean_nat_div(x_84, x_85); +lean_dec(x_84); +x_87 = lean_array_get_size(x_82); +x_88 = lean_nat_dec_le(x_86, x_87); +lean_dec(x_87); +lean_dec(x_86); +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; +x_89 = l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__1(x_82); +x_90 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_90, 0, x_79); +lean_ctor_set(x_90, 1, x_89); +lean_ctor_set(x_9, 2, x_90); +x_91 = lean_st_ref_set(x_2, x_9, x_11); +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_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; +x_95 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_95, 0, x_79); +lean_ctor_set(x_95, 1, x_82); +lean_ctor_set(x_9, 2, x_95); +x_96 = lean_st_ref_set(x_2, x_9, x_11); +x_97 = lean_ctor_get(x_96, 1); +lean_inc(x_97); +if (lean_is_exclusive(x_96)) { + lean_ctor_release(x_96, 0); + lean_ctor_release(x_96, 1); + x_98 = x_96; +} else { + lean_dec_ref(x_96); + x_98 = lean_box(0); +} +if (lean_is_scalar(x_98)) { + x_99 = lean_alloc_ctor(0, 2, 0); +} else { + x_99 = x_98; +} +lean_ctor_set(x_99, 0, x_80); +lean_ctor_set(x_99, 1, x_97); +return x_99; +} +} +else +{ +lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; +lean_dec(x_76); +lean_dec(x_1); +x_100 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_100, 0, x_61); +lean_ctor_set(x_100, 1, x_62); +lean_ctor_set(x_9, 2, x_100); +x_101 = lean_st_ref_set(x_2, x_9, x_11); +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; +} else { + lean_dec_ref(x_101); + x_103 = lean_box(0); +} +x_104 = lean_box(0); +if (lean_is_scalar(x_103)) { + x_105 = lean_alloc_ctor(0, 2, 0); +} else { + x_105 = x_103; +} +lean_ctor_set(x_105, 0, x_104); +lean_ctor_set(x_105, 1, x_102); +return x_105; +} +} +} +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; uint64_t x_112; uint64_t x_113; uint64_t x_114; uint64_t x_115; uint64_t x_116; uint64_t x_117; uint64_t x_118; size_t x_119; size_t x_120; size_t x_121; size_t x_122; size_t x_123; lean_object* x_124; uint8_t x_125; +x_106 = lean_ctor_get(x_9, 0); +x_107 = lean_ctor_get(x_9, 1); +lean_inc(x_107); +lean_inc(x_106); +lean_dec(x_9); +x_108 = lean_ctor_get(x_10, 0); +lean_inc(x_108); +x_109 = lean_ctor_get(x_10, 1); +lean_inc(x_109); +if (lean_is_exclusive(x_10)) { + lean_ctor_release(x_10, 0); + lean_ctor_release(x_10, 1); + x_110 = x_10; +} else { + lean_dec_ref(x_10); + x_110 = lean_box(0); +} +x_111 = lean_array_get_size(x_109); +x_112 = lean_uint64_of_nat(x_1); +x_113 = 32; +x_114 = lean_uint64_shift_right(x_112, x_113); +x_115 = lean_uint64_xor(x_112, x_114); +x_116 = 16; +x_117 = lean_uint64_shift_right(x_115, x_116); +x_118 = lean_uint64_xor(x_115, x_117); +x_119 = lean_uint64_to_usize(x_118); +x_120 = lean_usize_of_nat(x_111); +lean_dec(x_111); +x_121 = 1; +x_122 = lean_usize_sub(x_120, x_121); +x_123 = lean_usize_land(x_119, x_122); +x_124 = lean_array_uget(x_109, x_123); +x_125 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_foundBVar___spec__1(x_1, x_124); +if (x_125 == 0) +{ +lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; uint8_t x_136; +x_126 = lean_unsigned_to_nat(1u); +x_127 = lean_nat_add(x_108, x_126); +lean_dec(x_108); +x_128 = lean_box(0); +x_129 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_129, 0, x_1); +lean_ctor_set(x_129, 1, x_128); +lean_ctor_set(x_129, 2, x_124); +x_130 = lean_array_uset(x_109, x_123, x_129); +x_131 = lean_unsigned_to_nat(4u); +x_132 = lean_nat_mul(x_127, x_131); +x_133 = lean_unsigned_to_nat(3u); +x_134 = lean_nat_div(x_132, x_133); +lean_dec(x_132); +x_135 = lean_array_get_size(x_130); +x_136 = lean_nat_dec_le(x_134, x_135); +lean_dec(x_135); +lean_dec(x_134); +if (x_136 == 0) +{ +lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; +x_137 = l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__1(x_130); +if (lean_is_scalar(x_110)) { + x_138 = lean_alloc_ctor(0, 2, 0); +} else { + x_138 = x_110; +} +lean_ctor_set(x_138, 0, x_127); +lean_ctor_set(x_138, 1, x_137); +x_139 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_139, 0, x_106); +lean_ctor_set(x_139, 1, x_107); +lean_ctor_set(x_139, 2, x_138); +x_140 = lean_st_ref_set(x_2, x_139, x_11); +x_141 = lean_ctor_get(x_140, 1); +lean_inc(x_141); +if (lean_is_exclusive(x_140)) { + lean_ctor_release(x_140, 0); + lean_ctor_release(x_140, 1); + x_142 = x_140; +} else { + lean_dec_ref(x_140); + x_142 = lean_box(0); +} +if (lean_is_scalar(x_142)) { + x_143 = lean_alloc_ctor(0, 2, 0); +} else { + x_143 = x_142; +} +lean_ctor_set(x_143, 0, x_128); +lean_ctor_set(x_143, 1, x_141); +return x_143; +} +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; +if (lean_is_scalar(x_110)) { + x_144 = lean_alloc_ctor(0, 2, 0); +} else { + x_144 = x_110; +} +lean_ctor_set(x_144, 0, x_127); +lean_ctor_set(x_144, 1, x_130); +x_145 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_145, 0, x_106); +lean_ctor_set(x_145, 1, x_107); +lean_ctor_set(x_145, 2, x_144); +x_146 = lean_st_ref_set(x_2, x_145, x_11); +x_147 = lean_ctor_get(x_146, 1); +lean_inc(x_147); +if (lean_is_exclusive(x_146)) { + lean_ctor_release(x_146, 0); + lean_ctor_release(x_146, 1); + x_148 = x_146; +} else { + lean_dec_ref(x_146); + x_148 = lean_box(0); +} +if (lean_is_scalar(x_148)) { + x_149 = lean_alloc_ctor(0, 2, 0); +} else { + x_149 = x_148; +} +lean_ctor_set(x_149, 0, x_128); +lean_ctor_set(x_149, 1, x_147); +return x_149; +} +} +else +{ +lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; +lean_dec(x_124); +lean_dec(x_1); +if (lean_is_scalar(x_110)) { + x_150 = lean_alloc_ctor(0, 2, 0); +} else { + x_150 = x_110; +} +lean_ctor_set(x_150, 0, x_108); +lean_ctor_set(x_150, 1, x_109); +x_151 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_151, 0, x_106); +lean_ctor_set(x_151, 1, x_107); +lean_ctor_set(x_151, 2, x_150); +x_152 = lean_st_ref_set(x_2, x_151, x_11); +x_153 = lean_ctor_get(x_152, 1); +lean_inc(x_153); +if (lean_is_exclusive(x_152)) { + lean_ctor_release(x_152, 0); + lean_ctor_release(x_152, 1); + x_154 = x_152; +} else { + lean_dec_ref(x_152); + x_154 = lean_box(0); +} +x_155 = lean_box(0); +if (lean_is_scalar(x_154)) { + x_156 = lean_alloc_ctor(0, 2, 0); +} else { + x_156 = x_154; +} +lean_ctor_set(x_156, 0, x_155); +lean_ctor_set(x_156, 1, x_153); +return x_156; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; +x_8 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_8; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFn_x3f(lean_object* x_1) { +_start: +{ +uint8_t x_2; +x_2 = l_Lean_Expr_isApp(x_1); +if (x_2 == 0) +{ +lean_object* x_3; +x_3 = lean_box(0); +return x_3; +} +else +{ +lean_object* x_4; +x_4 = l_Lean_Expr_getAppFn(x_1); +switch (lean_obj_tag(x_4)) { +case 1: +{ +lean_object* x_5; +x_5 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_5, 0, x_4); +return x_5; +} +case 4: +{ +lean_object* x_6; lean_object* x_7; uint8_t x_8; +x_6 = lean_ctor_get(x_4, 0); +lean_inc(x_6); +x_7 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames; +x_8 = l_Array_contains___at_Lean_registerInternalExceptionId___spec__1(x_7, x_6); +lean_dec(x_6); +if (x_8 == 0) +{ +lean_object* x_9; +x_9 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_9, 0, x_4); +return x_9; +} +else +{ +lean_object* x_10; +lean_dec(x_4); +x_10 = lean_box(0); +return x_10; +} +} +default: +{ +lean_object* x_11; +lean_dec(x_4); +x_11 = lean_box(0); +return x_11; +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFn_x3f___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFn_x3f(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFunMask___spec__1(size_t x_1, size_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +uint8_t x_9; +x_9 = lean_usize_dec_lt(x_2, x_1); +if (x_9 == 0) +{ +lean_object* x_10; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_10 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_10, 0, x_3); +lean_ctor_set(x_10, 1, x_8); +return x_10; +} +else +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; lean_object* x_15; uint8_t x_22; lean_object* x_23; lean_object* x_38; +x_11 = lean_array_uget(x_3, x_2); +x_12 = lean_unsigned_to_nat(0u); +x_13 = lean_array_uset(x_3, x_2, x_12); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_11); +x_38 = l_Lean_Meta_isTypeFormer(x_11, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_38) == 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); +if (x_40 == 0) +{ +lean_object* x_41; lean_object* x_42; +lean_dec(x_39); +x_41 = lean_ctor_get(x_38, 1); +lean_inc(x_41); +lean_dec(x_38); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_11); +x_42 = l_Lean_Meta_isProof(x_11, x_4, x_5, x_6, x_7, x_41); +if (lean_obj_tag(x_42) == 0) +{ +lean_object* x_43; lean_object* x_44; uint8_t x_45; +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_unbox(x_43); +lean_dec(x_43); +x_22 = x_45; +x_23 = x_44; +goto block_37; +} +else +{ +uint8_t x_46; +lean_dec(x_13); +lean_dec(x_11); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_46 = !lean_is_exclusive(x_42); +if (x_46 == 0) +{ +return x_42; +} +else +{ +lean_object* x_47; lean_object* x_48; lean_object* x_49; +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_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; uint8_t x_51; +x_50 = lean_ctor_get(x_38, 1); +lean_inc(x_50); +lean_dec(x_38); +x_51 = lean_unbox(x_39); +lean_dec(x_39); +x_22 = x_51; +x_23 = x_50; +goto block_37; +} +} +else +{ +uint8_t x_52; +lean_dec(x_13); +lean_dec(x_11); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_52 = !lean_is_exclusive(x_38); +if (x_52 == 0) +{ +return x_38; +} +else +{ +lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_53 = lean_ctor_get(x_38, 0); +x_54 = lean_ctor_get(x_38, 1); +lean_inc(x_54); +lean_inc(x_53); +lean_dec(x_38); +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; +} +} +block_21: +{ +size_t x_16; size_t x_17; lean_object* x_18; lean_object* x_19; +x_16 = 1; +x_17 = lean_usize_add(x_2, x_16); +x_18 = lean_box(x_14); +x_19 = lean_array_uset(x_13, x_2, x_18); +x_2 = x_17; +x_3 = x_19; +x_8 = x_15; +goto _start; +} +block_37: +{ +if (x_22 == 0) +{ +lean_object* x_24; lean_object* x_25; +x_24 = l_Lean_Expr_fvarId_x21(x_11); +lean_dec(x_11); +lean_inc(x_4); +x_25 = l_Lean_FVarId_getDecl(x_24, x_4, x_5, x_6, x_7, x_23); +if (lean_obj_tag(x_25) == 0) +{ +lean_object* x_26; lean_object* x_27; uint8_t 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_LocalDecl_binderInfo(x_26); +lean_dec(x_26); +x_29 = lean_box(x_28); +if (lean_obj_tag(x_29) == 3) +{ +uint8_t x_30; +x_30 = 1; +x_14 = x_30; +x_15 = x_27; +goto block_21; +} +else +{ +uint8_t x_31; +lean_dec(x_29); +x_31 = 0; +x_14 = x_31; +x_15 = x_27; +goto block_21; +} +} +else +{ +uint8_t x_32; +lean_dec(x_13); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_32 = !lean_is_exclusive(x_25); +if (x_32 == 0) +{ +return x_25; +} +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_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; +} +} +} +else +{ +uint8_t x_36; +lean_dec(x_11); +x_36 = 1; +x_14 = x_36; +x_15 = x_23; +goto block_21; +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFunMask___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +size_t x_8; size_t x_9; lean_object* x_10; +x_8 = lean_array_size(x_1); +x_9 = 0; +x_10 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFunMask___spec__1(x_8, x_9, x_1, x_3, x_4, x_5, x_6, x_7); +return x_10; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFunMask___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFunMask___lambda__1___boxed), 7, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFunMask(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_8 = lean_infer_type(x_1, x_3, x_4, x_5, x_6, x_7); +if (lean_obj_tag(x_8) == 0) +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; +x_9 = lean_ctor_get(x_8, 0); +lean_inc(x_9); +x_10 = lean_ctor_get(x_8, 1); +lean_inc(x_10); +lean_dec(x_8); +x_11 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_11, 0, x_2); +x_12 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFunMask___closed__1; +x_13 = 0; +x_14 = l_Lean_Meta_forallBoundedTelescope___at_Lean_Meta_arrowDomainsN___spec__6___rarg(x_9, x_11, x_12, x_13, x_3, x_4, x_5, x_6, x_10); +return x_14; +} +else +{ +uint8_t x_15; +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_8); +if (x_15 == 0) +{ +return x_8; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_8, 0); +x_17 = lean_ctor_get(x_8, 1); +lean_inc(x_17); +lean_inc(x_16); +lean_dec(x_8); +x_18 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_18, 0, x_16); +lean_ctor_set(x_18, 1, x_17); +return x_18; +} +} +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFunMask___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +size_t x_9; size_t x_10; lean_object* x_11; +x_9 = lean_unbox_usize(x_1); +lean_dec(x_1); +x_10 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_11 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFunMask___spec__1(x_9, x_10, x_3, x_4, x_5, x_6, x_7, x_8); +return x_11; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFunMask___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; +x_8 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFunMask___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_2); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; lean_object* x_9; uint8_t x_10; +x_8 = lean_ctor_get(x_5, 5); +x_9 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_1, x_3, x_4, x_5, x_6, x_7); +x_10 = !lean_is_exclusive(x_9); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; +x_11 = lean_ctor_get(x_9, 0); +lean_inc(x_8); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_8); +lean_ctor_set(x_12, 1, x_11); +lean_ctor_set_tag(x_9, 1); +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); +lean_inc(x_8); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_8); +lean_ctor_set(x_15, 1, x_13); +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_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__2___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_10 = lean_array_set(x_2, x_1, x_3); +x_11 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_11, 0, 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_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +lean_object* x_14; uint8_t x_15; +x_14 = lean_ctor_get(x_3, 1); +x_15 = lean_nat_dec_lt(x_5, x_14); +if (x_15 == 0) +{ +lean_object* 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_5); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_4); +lean_ctor_set(x_16, 1, x_13); +return x_16; +} +else +{ +lean_object* x_17; lean_object* x_35; uint8_t x_36; lean_object* x_37; uint8_t x_38; lean_object* x_39; +x_35 = lean_array_get_size(x_4); +x_36 = lean_nat_dec_lt(x_5, x_35); +lean_dec(x_35); +x_37 = lean_array_get_size(x_1); +x_38 = lean_nat_dec_lt(x_5, x_37); +lean_dec(x_37); +if (x_36 == 0) +{ +lean_object* x_93; lean_object* x_94; +x_93 = l_Lean_instInhabitedExpr; +x_94 = l_outOfBounds___rarg(x_93); +x_39 = x_94; +goto block_92; +} +else +{ +lean_object* x_95; +x_95 = lean_array_fget(x_4, x_5); +x_39 = x_95; +goto block_92; +} +block_34: +{ +if (lean_obj_tag(x_17) == 0) +{ +lean_object* x_18; +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +if (lean_obj_tag(x_18) == 0) +{ +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_5); +x_19 = !lean_is_exclusive(x_17); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; +x_20 = lean_ctor_get(x_17, 0); +lean_dec(x_20); +x_21 = lean_ctor_get(x_18, 0); +lean_inc(x_21); +lean_dec(x_18); +lean_ctor_set(x_17, 0, x_21); +return x_17; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_22 = lean_ctor_get(x_17, 1); +lean_inc(x_22); +lean_dec(x_17); +x_23 = lean_ctor_get(x_18, 0); +lean_inc(x_23); +lean_dec(x_18); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set(x_24, 1, x_22); +return x_24; +} +} +else +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +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_3, 2); +x_28 = lean_nat_add(x_5, x_27); +lean_dec(x_5); +x_4 = x_26; +x_5 = x_28; +x_6 = lean_box(0); +x_7 = lean_box(0); +x_13 = x_25; +goto _start; +} +} +else +{ +uint8_t x_30; +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_5); +x_30 = !lean_is_exclusive(x_17); +if (x_30 == 0) +{ +return x_17; +} +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_17, 0); +x_32 = lean_ctor_get(x_17, 1); +lean_inc(x_32); +lean_inc(x_31); +lean_dec(x_17); +x_33 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_33, 0, x_31); +lean_ctor_set(x_33, 1, x_32); +return x_33; +} +} +} +block_92: +{ +uint8_t x_40; uint8_t x_41; +x_40 = l_Lean_Expr_hasLooseBVars(x_39); +if (x_38 == 0) +{ +uint8_t x_89; +x_89 = 0; +x_41 = x_89; +goto block_88; +} +else +{ +lean_object* x_90; uint8_t x_91; +x_90 = lean_array_fget(x_1, x_5); +x_91 = lean_unbox(x_90); +lean_dec(x_90); +x_41 = x_91; +goto block_88; +} +block_88: +{ +if (x_40 == 0) +{ +uint8_t x_42; +x_42 = l_Lean_Expr_hasMVar(x_39); +if (x_42 == 0) +{ +lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_43 = l_Lean_Meta_Grind_mkGroundPattern(x_39); +x_44 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__2___lambda__1(x_5, x_4, x_43, x_8, x_9, x_10, x_11, x_12, x_13); +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 = lean_ctor_get(x_3, 2); +x_49 = lean_nat_add(x_5, x_48); +lean_dec(x_5); +x_4 = x_47; +x_5 = x_49; +x_6 = lean_box(0); +x_7 = lean_box(0); +x_13 = x_46; +goto _start; +} +else +{ +lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; +lean_dec(x_39); +x_51 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare; +x_52 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__2___lambda__1(x_5, x_4, x_51, x_8, x_9, x_10, x_11, x_12, x_13); +x_53 = lean_ctor_get(x_52, 0); +lean_inc(x_53); +x_54 = lean_ctor_get(x_52, 1); +lean_inc(x_54); +lean_dec(x_52); +x_55 = lean_ctor_get(x_53, 0); +lean_inc(x_55); +lean_dec(x_53); +x_56 = lean_ctor_get(x_3, 2); +x_57 = lean_nat_add(x_5, x_56); +lean_dec(x_5); +x_4 = x_55; +x_5 = x_57; +x_6 = lean_box(0); +x_7 = lean_box(0); +x_13 = x_54; +goto _start; +} +} +else +{ +if (lean_obj_tag(x_39) == 0) +{ +lean_object* x_59; lean_object* x_60; +x_59 = lean_ctor_get(x_39, 0); +lean_inc(x_59); +x_60 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_foundBVar(x_59, x_8, x_9, x_10, x_11, x_12, x_13); +if (x_41 == 0) +{ +lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; +x_61 = lean_ctor_get(x_60, 1); +lean_inc(x_61); +lean_dec(x_60); +x_62 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar(x_59, x_8, x_9, x_10, x_11, x_12, x_61); +x_63 = lean_ctor_get(x_62, 1); +lean_inc(x_63); +lean_dec(x_62); +x_64 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__2___lambda__1(x_5, x_4, x_39, x_8, x_9, x_10, x_11, x_12, x_63); +x_17 = x_64; +goto block_34; +} +else +{ +lean_object* x_65; uint8_t x_66; +x_65 = lean_ctor_get(x_60, 0); +lean_inc(x_65); +x_66 = lean_unbox(x_65); +lean_dec(x_65); +if (x_66 == 0) +{ +lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; +x_67 = lean_ctor_get(x_60, 1); +lean_inc(x_67); +lean_dec(x_60); +x_68 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar(x_59, x_8, x_9, x_10, x_11, x_12, x_67); +x_69 = lean_ctor_get(x_68, 1); +lean_inc(x_69); +lean_dec(x_68); +x_70 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__2___lambda__1(x_5, x_4, x_39, x_8, x_9, x_10, x_11, x_12, x_69); +x_17 = x_70; +goto block_34; +} +else +{ +lean_object* x_71; lean_object* x_72; lean_object* x_73; +lean_dec(x_59); +lean_dec(x_39); +x_71 = lean_ctor_get(x_60, 1); +lean_inc(x_71); +lean_dec(x_60); +x_72 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare; +x_73 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__2___lambda__1(x_5, x_4, x_72, x_8, x_9, x_10, x_11, x_12, x_71); +x_17 = x_73; +goto block_34; +} +} +} +else +{ +if (x_41 == 0) +{ +lean_object* x_74; +x_74 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFn_x3f(x_39); +if (lean_obj_tag(x_74) == 0) +{ +lean_object* x_75; lean_object* x_76; +lean_dec(x_39); +x_75 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare; +x_76 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__2___lambda__1(x_5, x_4, x_75, x_8, x_9, x_10, x_11, x_12, x_13); +x_17 = x_76; +goto block_34; +} +else +{ +uint8_t x_77; lean_object* x_78; +lean_dec(x_74); +x_77 = 0; +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +x_78 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go(x_39, x_77, x_8, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_78) == 0) +{ +lean_object* x_79; lean_object* x_80; lean_object* x_81; +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 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__2___lambda__1(x_5, x_4, x_79, x_8, x_9, x_10, x_11, x_12, x_80); +x_17 = x_81; +goto block_34; +} +else +{ +uint8_t x_82; +lean_dec(x_4); +x_82 = !lean_is_exclusive(x_78); +if (x_82 == 0) +{ +x_17 = x_78; +goto block_34; +} +else +{ +lean_object* x_83; lean_object* x_84; lean_object* x_85; +x_83 = lean_ctor_get(x_78, 0); +x_84 = lean_ctor_get(x_78, 1); +lean_inc(x_84); +lean_inc(x_83); +lean_dec(x_78); +x_85 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_85, 0, x_83); +lean_ctor_set(x_85, 1, x_84); +x_17 = x_85; +goto block_34; +} +} +} +} +else +{ +lean_object* x_86; lean_object* x_87; +lean_dec(x_39); +x_86 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare; +x_87 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__2___lambda__1(x_5, x_4, x_86, x_8, x_9, x_10, x_11, x_12, x_13); +x_17 = x_87; +goto block_34; +} +} +} +} +} +} +} +} +static lean_object* _init_l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__3___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_instMonadMetaM; +x_2 = l_ReaderT_instMonad___rarg(x_1); +return x_2; +} +} +static lean_object* _init_l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__3___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__3___closed__1; +x_2 = l_Lean_instInhabitedExpr; +x_3 = l_instInhabitedOfMonad___rarg(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_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) { +_start: +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_8 = l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__3___closed__2; +x_9 = lean_panic_fn(x_8, x_1); +x_10 = lean_apply_6(x_9, x_2, x_3, x_4, x_5, x_6, x_7); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_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) { +_start: +{ +lean_object* x_8; lean_object* x_9; uint8_t x_10; +x_8 = lean_ctor_get(x_5, 5); +x_9 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_1, x_3, x_4, x_5, x_6, x_7); +x_10 = !lean_is_exclusive(x_9); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; +x_11 = lean_ctor_get(x_9, 0); +lean_inc(x_8); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_8); +lean_ctor_set(x_12, 1, x_11); +lean_ctor_set_tag(x_9, 1); +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); +lean_inc(x_8); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_8); +lean_ctor_set(x_15, 1, x_13); +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; +} +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("invalid pattern, (non-forbidden) application expected", 53, 53); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("assertion violation: ", 21, 21); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("f.isConst || f.isFVar\n ", 24, 24); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__3; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__4; +x_3 = lean_string_append(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("_private.Lean.Meta.Tactic.Grind.EMatchTheorem.0.Lean.Meta.Grind.NormalizePattern.go", 83, 83); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l_Lean_Meta_Grind_EMatchTheorems_insert___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__6; +x_3 = lean_unsigned_to_nat(177u); +x_4 = lean_unsigned_to_nat(2u); +x_5 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__5; +x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); +return x_6; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_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; +x_9 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFn_x3f(x_1); +if (lean_obj_tag(x_9) == 0) +{ +lean_object* x_10; lean_object* x_11; +lean_dec(x_1); +x_10 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__2; +x_11 = l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__1(x_10, 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_11; +} +else +{ +lean_object* x_12; lean_object* x_13; uint8_t x_46; +x_12 = lean_ctor_get(x_9, 0); +lean_inc(x_12); +lean_dec(x_9); +x_46 = l_Lean_Expr_isConst(x_12); +if (x_46 == 0) +{ +uint8_t x_47; +x_47 = l_Lean_Expr_isFVar(x_12); +if (x_47 == 0) +{ +lean_object* x_48; lean_object* x_49; +lean_dec(x_12); +lean_dec(x_1); +x_48 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__7; +x_49 = l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__3(x_48, x_3, x_4, x_5, x_6, x_7, x_8); +return x_49; +} +else +{ +lean_object* x_50; +x_50 = lean_box(0); +x_13 = x_50; +goto block_45; +} +} +else +{ +lean_object* x_51; +x_51 = lean_box(0); +x_13 = x_51; +goto block_45; +} +block_45: +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +lean_dec(x_13); +lean_inc(x_12); +x_14 = l_Lean_Expr_toHeadIndex(x_12); +x_15 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol(x_14, x_3, x_4, x_5, x_6, x_7, x_8); +x_16 = lean_ctor_get(x_15, 1); +lean_inc(x_16); +lean_dec(x_15); +x_17 = lean_unsigned_to_nat(0u); +x_18 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_17); +x_19 = l_Lean_Meta_Grind_ppPattern___closed__5; +lean_inc(x_18); +x_20 = lean_mk_array(x_18, x_19); +x_21 = lean_unsigned_to_nat(1u); +x_22 = lean_nat_sub(x_18, x_21); +lean_dec(x_18); +x_23 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_20, x_22); +x_24 = lean_array_get_size(x_23); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_24); +lean_inc(x_12); +x_25 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFunMask(x_12, x_24, x_4, x_5, x_6, x_7, x_16); +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 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_28, 0, x_17); +lean_ctor_set(x_28, 1, x_24); +lean_ctor_set(x_28, 2, x_21); +x_29 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__2(x_26, x_28, x_28, x_23, x_17, lean_box(0), lean_box(0), x_3, x_4, x_5, x_6, x_7, x_27); +lean_dec(x_28); +lean_dec(x_26); +if (lean_obj_tag(x_29) == 0) +{ +uint8_t x_30; +x_30 = !lean_is_exclusive(x_29); +if (x_30 == 0) +{ +lean_object* x_31; lean_object* x_32; +x_31 = lean_ctor_get(x_29, 0); +x_32 = l_Lean_mkAppN(x_12, x_31); +lean_dec(x_31); +lean_ctor_set(x_29, 0, x_32); +return x_29; +} +else +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_33 = lean_ctor_get(x_29, 0); +x_34 = lean_ctor_get(x_29, 1); +lean_inc(x_34); +lean_inc(x_33); +lean_dec(x_29); +x_35 = l_Lean_mkAppN(x_12, x_33); +lean_dec(x_33); +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 +{ +uint8_t x_37; +lean_dec(x_12); +x_37 = !lean_is_exclusive(x_29); +if (x_37 == 0) +{ +return x_29; +} +else +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_38 = lean_ctor_get(x_29, 0); +x_39 = lean_ctor_get(x_29, 1); +lean_inc(x_39); +lean_inc(x_38); +lean_dec(x_29); +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_24); +lean_dec(x_23); +lean_dec(x_12); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_41 = !lean_is_exclusive(x_25); +if (x_41 == 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_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; +} +} +} +} +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("invalid pattern, it does not have pattern variables", 51, 51); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +if (x_2 == 0) +{ +lean_object* x_9; lean_object* x_10; +x_9 = lean_box(0); +x_10 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1(x_1, x_9, x_3, x_4, x_5, x_6, x_7, x_8); +return x_10; +} +else +{ +uint8_t x_11; +x_11 = l_Lean_Expr_hasLooseBVars(x_1); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; uint8_t x_14; +lean_dec(x_1); +x_12 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___closed__2; +x_13 = l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__4(x_12, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) +{ +return x_13; +} +else +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_ctor_get(x_13, 0); +x_16 = lean_ctor_get(x_13, 1); +lean_inc(x_16); +lean_inc(x_15); +lean_dec(x_13); +x_17 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_17, 0, x_15); +lean_ctor_set(x_17, 1, x_16); +return x_17; +} +} +else +{ +lean_object* x_18; lean_object* x_19; +x_18 = lean_box(0); +x_19 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1(x_1, x_18, x_3, x_4, x_5, x_6, x_7, x_8); +return x_19; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_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, lean_object* x_7) { +_start: +{ +lean_object* x_8; +x_8 = l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__2___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__2___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +lean_object* x_14; +x_14 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_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) { +_start: +{ +lean_object* x_8; +x_8 = l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_8; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_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) { +_start: +{ +lean_object* x_9; +x_9 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_2); +return x_9; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_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) { +_start: +{ +uint8_t x_9; lean_object* x_10; +x_9 = lean_unbox(x_2); +lean_dec(x_2); +x_10 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go(x_1, x_9, x_3, x_4, x_5, x_6, x_7, x_8); +return x_10; +} +} +LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_Meta_Grind_NormalizePattern_main___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_9; lean_object* x_10; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_9 = l_List_reverse___rarg(x_2); +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_1); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; uint8_t x_14; lean_object* x_15; +x_12 = lean_ctor_get(x_1, 0); +x_13 = lean_ctor_get(x_1, 1); +x_14 = 0; +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_15 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go(x_12, 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; +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_15, 1); +lean_inc(x_17); +lean_dec(x_15); +lean_ctor_set(x_1, 1, x_2); +lean_ctor_set(x_1, 0, x_16); +{ +lean_object* _tmp_0 = x_13; +lean_object* _tmp_1 = x_1; +lean_object* _tmp_7 = x_17; +x_1 = _tmp_0; +x_2 = _tmp_1; +x_8 = _tmp_7; +} +goto _start; +} +else +{ +uint8_t x_19; +lean_free_object(x_1); +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_19 = !lean_is_exclusive(x_15); +if (x_19 == 0) +{ +return x_15; +} +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_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; +} +} +} +else +{ +lean_object* x_23; lean_object* x_24; uint8_t x_25; lean_object* x_26; +x_23 = lean_ctor_get(x_1, 0); +x_24 = lean_ctor_get(x_1, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_1); +x_25 = 0; +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_26 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go(x_23, x_25, x_3, x_4, x_5, x_6, x_7, x_8); +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 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_29, 0, x_27); +lean_ctor_set(x_29, 1, x_2); +x_1 = x_24; +x_2 = x_29; +x_8 = x_28; +goto _start; +} +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; +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_2); +x_31 = lean_ctor_get(x_26, 0); +lean_inc(x_31); +x_32 = lean_ctor_get(x_26, 1); +lean_inc(x_32); +if (lean_is_exclusive(x_26)) { + lean_ctor_release(x_26, 0); + lean_ctor_release(x_26, 1); + x_33 = x_26; +} else { + lean_dec_ref(x_26); + x_33 = lean_box(0); +} +if (lean_is_scalar(x_33)) { + x_34 = lean_alloc_ctor(1, 2, 0); +} else { + x_34 = x_33; +} +lean_ctor_set(x_34, 0, x_31); +lean_ctor_set(x_34, 1, x_32); +return x_34; +} +} +} +} +} +static lean_object* _init_l_Lean_Meta_Grind_NormalizePattern_main___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(0); +x_2 = lean_array_mk(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Grind_NormalizePattern_main___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_unsigned_to_nat(10u); +x_2 = lean_unsigned_to_nat(1u); +x_3 = l_Nat_nextPowerOfTwo_go(x_1, x_2, lean_box(0)); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind_NormalizePattern_main___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_Grind_NormalizePattern_main___closed__2; +x_3 = lean_mk_array(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind_NormalizePattern_main___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_unsigned_to_nat(0u); +x_2 = l_Lean_Meta_Grind_NormalizePattern_main___closed__3; +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind_NormalizePattern_main___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_NormalizePattern_main___closed__1; +x_2 = l_Lean_Meta_Grind_NormalizePattern_main___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_2); +lean_ctor_set(x_3, 2, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_NormalizePattern_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) { +_start: +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; +x_7 = lean_box(0); +x_8 = l_Lean_Meta_Grind_NormalizePattern_main___closed__5; +x_9 = lean_st_mk_ref(x_8, 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; +x_11 = lean_ctor_get(x_9, 0); +x_12 = lean_ctor_get(x_9, 1); +lean_inc(x_11); +x_13 = l_List_mapM_loop___at_Lean_Meta_Grind_NormalizePattern_main___spec__1(x_1, x_7, x_11, x_2, x_3, x_4, x_5, 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_st_ref_get(x_11, x_15); +lean_dec(x_11); +x_17 = !lean_is_exclusive(x_16); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_18 = lean_ctor_get(x_16, 0); +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_array_to_list(x_19); +x_21 = lean_ctor_get(x_18, 2); +lean_inc(x_21); +lean_dec(x_18); +lean_ctor_set(x_9, 1, x_21); +lean_ctor_set(x_9, 0, x_20); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_14); +lean_ctor_set(x_22, 1, x_9); +lean_ctor_set(x_16, 0, x_22); +return x_16; +} +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; +x_23 = lean_ctor_get(x_16, 0); +x_24 = lean_ctor_get(x_16, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_16); +x_25 = lean_ctor_get(x_23, 0); +lean_inc(x_25); +x_26 = lean_array_to_list(x_25); +x_27 = lean_ctor_get(x_23, 2); +lean_inc(x_27); +lean_dec(x_23); +lean_ctor_set(x_9, 1, x_27); +lean_ctor_set(x_9, 0, x_26); +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_14); +lean_ctor_set(x_28, 1, x_9); +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 +{ +uint8_t x_30; +lean_free_object(x_9); +lean_dec(x_11); +x_30 = !lean_is_exclusive(x_13); +if (x_30 == 0) +{ +return x_13; +} +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_13, 0); +x_32 = lean_ctor_get(x_13, 1); +lean_inc(x_32); +lean_inc(x_31); +lean_dec(x_13); +x_33 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_33, 0, x_31); +lean_ctor_set(x_33, 1, x_32); +return x_33; +} +} +} +else +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_9, 0); +x_35 = lean_ctor_get(x_9, 1); +lean_inc(x_35); +lean_inc(x_34); +lean_dec(x_9); +lean_inc(x_34); +x_36 = l_List_mapM_loop___at_Lean_Meta_Grind_NormalizePattern_main___spec__1(x_1, x_7, x_34, x_2, x_3, x_4, x_5, x_35); +if (lean_obj_tag(x_36) == 0) +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_37 = lean_ctor_get(x_36, 0); +lean_inc(x_37); +x_38 = lean_ctor_get(x_36, 1); +lean_inc(x_38); +lean_dec(x_36); +x_39 = lean_st_ref_get(x_34, x_38); +lean_dec(x_34); +x_40 = lean_ctor_get(x_39, 0); +lean_inc(x_40); +x_41 = lean_ctor_get(x_39, 1); +lean_inc(x_41); +if (lean_is_exclusive(x_39)) { + lean_ctor_release(x_39, 0); + lean_ctor_release(x_39, 1); + x_42 = x_39; +} else { + lean_dec_ref(x_39); + x_42 = lean_box(0); +} +x_43 = lean_ctor_get(x_40, 0); +lean_inc(x_43); +x_44 = lean_array_to_list(x_43); +x_45 = lean_ctor_get(x_40, 2); +lean_inc(x_45); +lean_dec(x_40); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_44); +lean_ctor_set(x_46, 1, x_45); +x_47 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_47, 0, x_37); +lean_ctor_set(x_47, 1, x_46); +if (lean_is_scalar(x_42)) { + x_48 = lean_alloc_ctor(0, 2, 0); +} else { + x_48 = x_42; +} +lean_ctor_set(x_48, 0, x_47); +lean_ctor_set(x_48, 1, x_41); +return x_48; +} +else +{ +lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; +lean_dec(x_34); +x_49 = lean_ctor_get(x_36, 0); +lean_inc(x_49); +x_50 = lean_ctor_get(x_36, 1); +lean_inc(x_50); +if (lean_is_exclusive(x_36)) { + lean_ctor_release(x_36, 0); + lean_ctor_release(x_36, 1); + x_51 = x_36; +} else { + lean_dec_ref(x_36); + x_51 = lean_box(0); +} +if (lean_is_scalar(x_51)) { + x_52 = lean_alloc_ctor(1, 2, 0); +} else { + x_52 = x_51; +} +lean_ctor_set(x_52, 0, x_49); +lean_ctor_set(x_52, 1, x_50); +return x_52; +} +} +} +} +LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___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) { +_start: +{ +uint8_t x_7; +x_7 = lean_usize_dec_eq(x_5, x_6); +if (x_7 == 0) +{ +lean_object* x_8; lean_object* x_9; +x_8 = lean_array_uget(x_4, x_5); +x_9 = l_Lean_RBNode_findCore___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__2(x_1, x_8); +if (lean_obj_tag(x_9) == 0) +{ +size_t x_10; size_t x_11; +lean_dec(x_8); +x_10 = 1; +x_11 = lean_usize_add(x_5, x_10); +x_5 = x_11; +goto _start; +} +else +{ +lean_object* x_13; +lean_dec(x_9); +x_13 = l_Lean_RBNode_findCore___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__2(x_2, x_8); +lean_dec(x_8); +if (lean_obj_tag(x_13) == 0) +{ +uint8_t x_14; +x_14 = 1; +return x_14; +} +else +{ +size_t x_15; size_t x_16; +lean_dec(x_13); +x_15 = 1; +x_16 = lean_usize_add(x_5, x_15); +x_5 = x_16; +goto _start; +} +} +} +else +{ +uint8_t x_18; +x_18 = 0; +return x_18; +} +} +} +LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___spec__1___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___spec__2(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; +x_7 = lean_array_uget(x_3, x_4); +x_8 = l_Lean_RBNode_findCore___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__2(x_1, x_7); +if (lean_obj_tag(x_8) == 0) +{ +size_t x_9; size_t x_10; +lean_dec(x_7); +x_9 = 1; +x_10 = lean_usize_add(x_4, x_9); +x_4 = x_10; +goto _start; +} +else +{ +lean_object* x_12; +lean_dec(x_8); +x_12 = l_Lean_RBNode_findCore___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__2(x_2, x_7); +lean_dec(x_7); +if (lean_obj_tag(x_12) == 0) +{ +uint8_t x_13; +x_13 = 1; +return x_13; +} +else +{ +size_t x_14; size_t x_15; +lean_dec(x_12); +x_14 = 1; +x_15 = lean_usize_add(x_4, x_14); +x_4 = x_15; +goto _start; +} +} +} +else +{ +uint8_t x_17; +x_17 = 0; +return x_17; +} +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; +x_3 = l_Lean_Meta_Grind_NormalizePattern_main___closed__1; +x_4 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_4, 0, x_2); +lean_ctor_set(x_4, 1, x_1); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; +x_4 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___closed__1; +x_5 = l_Lean_CollectFVars_main(x_3, x_4); +x_6 = lean_ctor_get(x_5, 2); +lean_inc(x_6); +lean_dec(x_5); +x_7 = lean_array_get_size(x_6); +x_8 = lean_unsigned_to_nat(0u); +x_9 = lean_nat_dec_lt(x_8, x_7); +if (x_9 == 0) +{ +uint8_t x_10; +lean_dec(x_7); +lean_dec(x_6); +x_10 = 1; +return x_10; +} +else +{ +size_t x_11; size_t x_12; uint8_t x_13; +x_11 = 0; +x_12 = lean_usize_of_nat(x_7); +lean_dec(x_7); +x_13 = l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___spec__1___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___spec__2(x_1, x_2, x_6, x_11, x_12); +lean_dec(x_6); +if (x_13 == 0) +{ +uint8_t x_14; +x_14 = 1; +return x_14; +} +else +{ +uint8_t x_15; +x_15 = 0; +return x_15; +} +} +} +} +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___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_7; size_t x_8; uint8_t x_9; lean_object* x_10; +x_7 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_8 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_9 = l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___spec__1(x_1, x_2, x_3, x_4, x_7, x_8); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_10 = lean_box(x_9); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___spec__1___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +size_t x_6; size_t x_7; 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___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___spec__1___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___spec__2(x_1, x_2, x_3, x_6, x_7); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_9 = lean_box(x_8); +return x_9; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___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___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars(x_1, x_2, x_3); +lean_dec(x_2); +lean_dec(x_1); +x_5 = lean_box(x_4); +return x_5; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__1___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; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__1___closed__1; +x_2 = lean_box(0); +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___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, size_t x_7, size_t x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +uint8_t x_15; +x_15 = lean_usize_dec_lt(x_8, x_7); +if (x_15 == 0) +{ +lean_object* x_16; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_5); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_9); +lean_ctor_set(x_16, 1, x_14); +return x_16; +} +else +{ +lean_object* x_17; lean_object* x_18; +lean_dec(x_9); +x_17 = lean_array_uget(x_6, x_8); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +x_18 = lean_infer_type(x_17, x_10, x_11, x_12, x_13, x_14); +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_ctor_get(x_18, 1); +x_22 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars(x_1, x_2, x_20); +if (x_22 == 0) +{ +lean_object* x_23; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_5); +x_23 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__1___closed__2; +lean_ctor_set(x_18, 0, x_23); +return x_18; +} +else +{ +size_t x_24; size_t x_25; +lean_free_object(x_18); +x_24 = 1; +x_25 = lean_usize_add(x_8, x_24); +lean_inc(x_5); +{ +size_t _tmp_7 = x_25; +lean_object* _tmp_8 = x_5; +lean_object* _tmp_13 = x_21; +x_8 = _tmp_7; +x_9 = _tmp_8; +x_14 = _tmp_13; +} +goto _start; +} +} +else +{ +lean_object* x_27; lean_object* x_28; uint8_t x_29; +x_27 = lean_ctor_get(x_18, 0); +x_28 = lean_ctor_get(x_18, 1); +lean_inc(x_28); +lean_inc(x_27); +lean_dec(x_18); +x_29 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars(x_1, x_2, x_27); +if (x_29 == 0) +{ +lean_object* x_30; lean_object* x_31; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_5); +x_30 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__1___closed__2; +x_31 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_31, 0, x_30); +lean_ctor_set(x_31, 1, x_28); +return x_31; +} +else +{ +size_t x_32; size_t x_33; +x_32 = 1; +x_33 = lean_usize_add(x_8, x_32); +lean_inc(x_5); +{ +size_t _tmp_7 = x_33; +lean_object* _tmp_8 = x_5; +lean_object* _tmp_13 = x_28; +x_8 = _tmp_7; +x_9 = _tmp_8; +x_14 = _tmp_13; +} +goto _start; +} +} +} +else +{ +uint8_t x_35; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_5); +x_35 = !lean_is_exclusive(x_18); +if (x_35 == 0) +{ +return x_18; +} +else +{ +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_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; +} +} +} +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("semiOutParam", 12, 12); +return x_1; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("outParam", 8, 8); +return x_1; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__3; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___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, size_t x_7, size_t x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +uint8_t x_15; +x_15 = lean_usize_dec_lt(x_8, x_7); +if (x_15 == 0) +{ +lean_object* x_16; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_4); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_9); +lean_ctor_set(x_16, 1, x_14); +return x_16; +} +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_27; +x_17 = lean_array_uget(x_6, x_8); +x_27 = !lean_is_exclusive(x_9); +if (x_27 == 0) +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; +x_28 = lean_ctor_get(x_9, 1); +x_29 = lean_ctor_get(x_9, 0); +lean_dec(x_29); +x_30 = lean_ctor_get(x_28, 0); +lean_inc(x_30); +x_31 = lean_ctor_get(x_28, 1); +lean_inc(x_31); +x_32 = lean_ctor_get(x_28, 2); +lean_inc(x_32); +x_33 = lean_nat_dec_lt(x_31, x_32); +if (x_33 == 0) +{ +lean_object* x_34; +lean_dec(x_32); +lean_dec(x_31); +lean_dec(x_30); +lean_dec(x_17); +lean_inc(x_4); +lean_ctor_set(x_9, 0, x_4); +x_34 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_34, 0, x_9); +x_18 = x_34; +x_19 = x_14; +goto block_26; +} +else +{ +uint8_t x_35; +x_35 = !lean_is_exclusive(x_28); +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; +x_36 = lean_ctor_get(x_28, 2); +lean_dec(x_36); +x_37 = lean_ctor_get(x_28, 1); +lean_dec(x_37); +x_38 = lean_ctor_get(x_28, 0); +lean_dec(x_38); +x_39 = lean_array_fget(x_30, x_31); +x_40 = lean_unsigned_to_nat(1u); +x_41 = lean_nat_add(x_31, x_40); +lean_dec(x_31); +lean_ctor_set(x_28, 1, x_41); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +x_42 = lean_infer_type(x_17, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(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_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_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__2; +x_46 = l_Lean_Expr_isAppOf(x_43, x_45); +if (x_46 == 0) +{ +lean_object* x_47; uint8_t x_48; +x_47 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__4; +x_48 = l_Lean_Expr_isAppOf(x_43, x_47); +lean_dec(x_43); +if (x_48 == 0) +{ +uint8_t x_49; +x_49 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars(x_1, x_2, x_39); +if (x_49 == 0) +{ +lean_object* x_50; lean_object* x_51; +x_50 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__1___closed__1; +lean_ctor_set(x_9, 0, x_50); +x_51 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_51, 0, x_9); +x_18 = x_51; +x_19 = x_44; +goto block_26; +} +else +{ +lean_object* x_52; +lean_inc(x_4); +lean_ctor_set(x_9, 0, x_4); +x_52 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_52, 0, x_9); +x_18 = x_52; +x_19 = x_44; +goto block_26; +} +} +else +{ +lean_object* x_53; +lean_dec(x_39); +lean_inc(x_4); +lean_ctor_set(x_9, 0, x_4); +x_53 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_53, 0, x_9); +x_18 = x_53; +x_19 = x_44; +goto block_26; +} +} +else +{ +lean_object* x_54; +lean_dec(x_43); +lean_dec(x_39); +lean_inc(x_4); +lean_ctor_set(x_9, 0, x_4); +x_54 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_54, 0, x_9); +x_18 = x_54; +x_19 = x_44; +goto block_26; +} +} +else +{ +uint8_t x_55; +lean_dec(x_28); +lean_dec(x_39); +lean_free_object(x_9); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_4); +x_55 = !lean_is_exclusive(x_42); +if (x_55 == 0) +{ +return x_42; +} +else +{ +lean_object* x_56; lean_object* x_57; lean_object* x_58; +x_56 = lean_ctor_get(x_42, 0); +x_57 = lean_ctor_get(x_42, 1); +lean_inc(x_57); +lean_inc(x_56); +lean_dec(x_42); +x_58 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_58, 0, x_56); +lean_ctor_set(x_58, 1, x_57); +return x_58; +} +} +} +else +{ +lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; +lean_dec(x_28); +x_59 = lean_array_fget(x_30, x_31); +x_60 = lean_unsigned_to_nat(1u); +x_61 = lean_nat_add(x_31, x_60); +lean_dec(x_31); +x_62 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_62, 0, x_30); +lean_ctor_set(x_62, 1, x_61); +lean_ctor_set(x_62, 2, x_32); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +x_63 = lean_infer_type(x_17, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_63) == 0) +{ +lean_object* x_64; lean_object* x_65; lean_object* x_66; uint8_t 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 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__2; +x_67 = l_Lean_Expr_isAppOf(x_64, x_66); +if (x_67 == 0) +{ +lean_object* x_68; uint8_t x_69; +x_68 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__4; +x_69 = l_Lean_Expr_isAppOf(x_64, x_68); +lean_dec(x_64); +if (x_69 == 0) +{ +uint8_t x_70; +x_70 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars(x_1, x_2, x_59); +if (x_70 == 0) +{ +lean_object* x_71; lean_object* x_72; +x_71 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__1___closed__1; +lean_ctor_set(x_9, 1, x_62); +lean_ctor_set(x_9, 0, x_71); +x_72 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_72, 0, x_9); +x_18 = x_72; +x_19 = x_65; +goto block_26; +} +else +{ +lean_object* x_73; +lean_inc(x_4); +lean_ctor_set(x_9, 1, x_62); +lean_ctor_set(x_9, 0, x_4); +x_73 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_73, 0, x_9); +x_18 = x_73; +x_19 = x_65; +goto block_26; +} +} +else +{ +lean_object* x_74; +lean_dec(x_59); +lean_inc(x_4); +lean_ctor_set(x_9, 1, x_62); +lean_ctor_set(x_9, 0, x_4); +x_74 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_74, 0, x_9); +x_18 = x_74; +x_19 = x_65; +goto block_26; +} +} +else +{ +lean_object* x_75; +lean_dec(x_64); +lean_dec(x_59); +lean_inc(x_4); +lean_ctor_set(x_9, 1, x_62); +lean_ctor_set(x_9, 0, x_4); +x_75 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_75, 0, x_9); +x_18 = x_75; +x_19 = x_65; +goto block_26; +} +} +else +{ +lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; +lean_dec(x_62); +lean_dec(x_59); +lean_free_object(x_9); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_4); +x_76 = lean_ctor_get(x_63, 0); +lean_inc(x_76); +x_77 = lean_ctor_get(x_63, 1); +lean_inc(x_77); +if (lean_is_exclusive(x_63)) { + lean_ctor_release(x_63, 0); + lean_ctor_release(x_63, 1); + x_78 = x_63; +} else { + lean_dec_ref(x_63); + 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 +{ +lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; uint8_t x_84; +x_80 = lean_ctor_get(x_9, 1); +lean_inc(x_80); +lean_dec(x_9); +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_80, 2); +lean_inc(x_83); +x_84 = lean_nat_dec_lt(x_82, x_83); +if (x_84 == 0) +{ +lean_object* x_85; lean_object* x_86; +lean_dec(x_83); +lean_dec(x_82); +lean_dec(x_81); +lean_dec(x_17); +lean_inc(x_4); +x_85 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_85, 0, x_4); +lean_ctor_set(x_85, 1, x_80); +x_86 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_86, 0, x_85); +x_18 = x_86; +x_19 = x_14; +goto block_26; +} +else +{ +lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; +if (lean_is_exclusive(x_80)) { + lean_ctor_release(x_80, 0); + lean_ctor_release(x_80, 1); + lean_ctor_release(x_80, 2); + x_87 = x_80; +} else { + lean_dec_ref(x_80); + x_87 = lean_box(0); +} +x_88 = lean_array_fget(x_81, x_82); +x_89 = lean_unsigned_to_nat(1u); +x_90 = lean_nat_add(x_82, x_89); +lean_dec(x_82); +if (lean_is_scalar(x_87)) { + x_91 = lean_alloc_ctor(0, 3, 0); +} else { + x_91 = x_87; +} +lean_ctor_set(x_91, 0, x_81); +lean_ctor_set(x_91, 1, x_90); +lean_ctor_set(x_91, 2, x_83); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +x_92 = lean_infer_type(x_17, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_92) == 0) +{ +lean_object* x_93; lean_object* x_94; lean_object* x_95; uint8_t x_96; +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_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__2; +x_96 = l_Lean_Expr_isAppOf(x_93, x_95); +if (x_96 == 0) +{ +lean_object* x_97; uint8_t x_98; +x_97 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__4; +x_98 = l_Lean_Expr_isAppOf(x_93, x_97); +lean_dec(x_93); +if (x_98 == 0) +{ +uint8_t x_99; +x_99 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars(x_1, x_2, x_88); +if (x_99 == 0) +{ +lean_object* x_100; lean_object* x_101; lean_object* x_102; +x_100 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__1___closed__1; +x_101 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_101, 0, x_100); +lean_ctor_set(x_101, 1, x_91); +x_102 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_102, 0, x_101); +x_18 = x_102; +x_19 = x_94; +goto block_26; +} +else +{ +lean_object* x_103; lean_object* x_104; +lean_inc(x_4); +x_103 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_103, 0, x_4); +lean_ctor_set(x_103, 1, x_91); +x_104 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_104, 0, x_103); +x_18 = x_104; +x_19 = x_94; +goto block_26; +} +} +else +{ +lean_object* x_105; lean_object* x_106; +lean_dec(x_88); +lean_inc(x_4); +x_105 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_105, 0, x_4); +lean_ctor_set(x_105, 1, x_91); +x_106 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_106, 0, x_105); +x_18 = x_106; +x_19 = x_94; +goto block_26; +} +} +else +{ +lean_object* x_107; lean_object* x_108; +lean_dec(x_93); +lean_dec(x_88); +lean_inc(x_4); +x_107 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_107, 0, x_4); +lean_ctor_set(x_107, 1, x_91); +x_108 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_108, 0, x_107); +x_18 = x_108; +x_19 = x_94; +goto block_26; +} +} +else +{ +lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; +lean_dec(x_91); +lean_dec(x_88); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_4); +x_109 = lean_ctor_get(x_92, 0); +lean_inc(x_109); +x_110 = lean_ctor_get(x_92, 1); +lean_inc(x_110); +if (lean_is_exclusive(x_92)) { + lean_ctor_release(x_92, 0); + lean_ctor_release(x_92, 1); + x_111 = x_92; +} else { + lean_dec_ref(x_92); + x_111 = lean_box(0); +} +if (lean_is_scalar(x_111)) { + x_112 = lean_alloc_ctor(1, 2, 0); +} else { + x_112 = x_111; +} +lean_ctor_set(x_112, 0, x_109); +lean_ctor_set(x_112, 1, x_110); +return x_112; +} +} +} +block_26: +{ +if (lean_obj_tag(x_18) == 0) +{ +lean_object* x_20; lean_object* x_21; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_4); +x_20 = lean_ctor_get(x_18, 0); +lean_inc(x_20); +lean_dec(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; +} +else +{ +lean_object* x_22; size_t x_23; size_t x_24; +x_22 = lean_ctor_get(x_18, 0); +lean_inc(x_22); +lean_dec(x_18); +x_23 = 1; +x_24 = lean_usize_add(x_8, x_23); +x_8 = x_24; +x_9 = x_22; +x_14 = x_19; +goto _start; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___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) { +_start: +{ +uint8_t x_7; lean_object* x_8; lean_object* x_9; +x_7 = 1; +x_8 = lean_box(x_7); +x_9 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_9, 0, x_8); +lean_ctor_set(x_9, 1, x_6); +return x_9; +} +} +static lean_object* _init_l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__2___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__1___boxed), 6, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, size_t x_6, lean_object* x_7, 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; size_t x_18; lean_object* x_19; +x_14 = lean_array_get_size(x_1); +x_15 = lean_unsigned_to_nat(0u); +x_16 = l_Array_toSubarray___rarg(x_1, x_15, x_14); +lean_inc(x_2); +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_2); +lean_ctor_set(x_17, 1, x_16); +x_18 = lean_array_size(x_7); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +x_19 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2(x_3, x_4, x_5, x_2, x_7, x_7, x_18, x_6, x_17, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_19) == 0) +{ +lean_object* x_20; lean_object* x_21; +x_20 = lean_ctor_get(x_19, 0); +lean_inc(x_20); +x_21 = lean_ctor_get(x_20, 0); +lean_inc(x_21); +lean_dec(x_20); +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_19, 1); +lean_inc(x_22); +lean_dec(x_19); +x_23 = l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__2___closed__1; +x_24 = lean_box(0); +x_25 = lean_apply_6(x_23, x_24, x_9, x_10, x_11, x_12, x_22); +return x_25; +} +else +{ +uint8_t x_26; +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +x_26 = !lean_is_exclusive(x_19); +if (x_26 == 0) +{ +lean_object* x_27; lean_object* x_28; +x_27 = lean_ctor_get(x_19, 0); +lean_dec(x_27); +x_28 = lean_ctor_get(x_21, 0); +lean_inc(x_28); +lean_dec(x_21); +lean_ctor_set(x_19, 0, x_28); +return x_19; +} +else +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_29 = lean_ctor_get(x_19, 1); +lean_inc(x_29); +lean_dec(x_19); +x_30 = lean_ctor_get(x_21, 0); +lean_inc(x_30); +lean_dec(x_21); +x_31 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_31, 0, x_30); +lean_ctor_set(x_31, 1, x_29); +return x_31; +} +} +} +else +{ +uint8_t x_32; +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +x_32 = !lean_is_exclusive(x_19); +if (x_32 == 0) +{ +return x_19; +} +else +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_33 = lean_ctor_get(x_19, 0); +x_34 = lean_ctor_get(x_19, 1); +lean_inc(x_34); +lean_inc(x_33); +lean_dec(x_19); +x_35 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_35, 0, x_33); +lean_ctor_set(x_35, 1, x_34); +return x_35; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, size_t x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +lean_object* x_15; +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +x_15 = lean_infer_type(x_1, x_10, x_11, x_12, x_13, x_14); +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; uint8_t x_23; lean_object* x_24; +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_unsigned_to_nat(0u); +x_19 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_2, x_18); +x_20 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_20, 0, x_19); +x_21 = lean_box_usize(x_8); +x_22 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__2___boxed), 13, 6); +lean_closure_set(x_22, 0, x_3); +lean_closure_set(x_22, 1, x_4); +lean_closure_set(x_22, 2, x_5); +lean_closure_set(x_22, 3, x_6); +lean_closure_set(x_22, 4, x_7); +lean_closure_set(x_22, 5, x_21); +x_23 = 0; +x_24 = l_Lean_Meta_forallBoundedTelescope___at_Lean_Meta_arrowDomainsN___spec__6___rarg(x_16, x_20, x_22, x_23, x_10, x_11, x_12, x_13, x_17); +return x_24; +} +else +{ +uint8_t x_25; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_25 = !lean_is_exclusive(x_15); +if (x_25 == 0) +{ +return x_15; +} +else +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_26 = lean_ctor_get(x_15, 0); +x_27 = lean_ctor_get(x_15, 1); +lean_inc(x_27); +lean_inc(x_26); +lean_dec(x_15); +x_28 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_28, 0, x_26); +lean_ctor_set(x_28, 1, x_27); +return x_28; +} +} +} +} +static lean_object* _init_l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = lean_box(0); +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +if (lean_obj_tag(x_5) == 5) +{ +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_5, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_5, 1); +lean_inc(x_14); +lean_dec(x_5); +x_15 = lean_array_set(x_6, x_7, x_14); +x_16 = lean_unsigned_to_nat(1u); +x_17 = lean_nat_sub(x_7, x_16); +lean_dec(x_7); +x_5 = x_13; +x_6 = x_15; +x_7 = x_17; +goto _start; +} +else +{ +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_dec(x_7); +x_19 = lean_box(0); +x_20 = lean_box(0); +x_21 = lean_array_size(x_3); +x_22 = 0; +x_23 = l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___closed__1; +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +x_24 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__1(x_1, x_2, x_3, x_19, x_23, x_3, x_21, x_22, x_23, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_24) == 0) +{ +lean_object* x_25; lean_object* x_26; +x_25 = lean_ctor_get(x_24, 0); +lean_inc(x_25); +x_26 = lean_ctor_get(x_25, 0); +lean_inc(x_26); +lean_dec(x_25); +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_24, 1); +lean_inc(x_27); +lean_dec(x_24); +x_28 = lean_box(0); +x_29 = l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__3(x_5, x_4, x_6, x_20, x_1, x_2, x_19, x_22, x_28, x_8, x_9, x_10, x_11, x_27); +return x_29; +} +else +{ +uint8_t x_30; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_2); +lean_dec(x_1); +x_30 = !lean_is_exclusive(x_24); +if (x_30 == 0) +{ +lean_object* x_31; lean_object* x_32; +x_31 = lean_ctor_get(x_24, 0); +lean_dec(x_31); +x_32 = lean_ctor_get(x_26, 0); +lean_inc(x_32); +lean_dec(x_26); +lean_ctor_set(x_24, 0, x_32); +return x_24; +} +else +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_33 = lean_ctor_get(x_24, 1); +lean_inc(x_33); +lean_dec(x_24); +x_34 = lean_ctor_get(x_26, 0); +lean_inc(x_34); +lean_dec(x_26); +x_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_34); +lean_ctor_set(x_35, 1, x_33); +return x_35; +} +} +} +else +{ +uint8_t x_36; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_2); +lean_dec(x_1); +x_36 = !lean_is_exclusive(x_24); +if (x_36 == 0) +{ +return x_24; +} +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_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_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_10 = lean_unsigned_to_nat(0u); +x_11 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_4, x_10); +x_12 = l_Lean_Meta_Grind_ppPattern___closed__5; +lean_inc(x_11); +x_13 = lean_mk_array(x_11, x_12); +x_14 = lean_unsigned_to_nat(1u); +x_15 = lean_nat_sub(x_11, x_14); +lean_dec(x_11); +lean_inc(x_4); +x_16 = l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3(x_1, x_2, x_3, x_4, x_4, x_13, x_15, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_4); +return x_16; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; uint8_t x_10; lean_object* x_11; +x_9 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___lambda__1___boxed), 9, 2); +lean_closure_set(x_9, 0, x_1); +lean_closure_set(x_9, 1, x_2); +x_10 = 0; +x_11 = l_Lean_Meta_forallTelescopeReducing___at_Lean_Meta_getParamNames___spec__2___rarg(x_3, x_9, x_10, x_4, x_5, x_6, x_7, x_8); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +size_t x_15; size_t x_16; lean_object* x_17; +x_15 = lean_unbox_usize(x_7); +lean_dec(x_7); +x_16 = lean_unbox_usize(x_8); +lean_dec(x_8); +x_17 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_15, x_16, x_9, x_10, x_11, x_12, x_13, x_14); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_17; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +size_t x_15; size_t x_16; lean_object* x_17; +x_15 = lean_unbox_usize(x_7); +lean_dec(x_7); +x_16 = lean_unbox_usize(x_8); +lean_dec(x_8); +x_17 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_15, x_16, x_9, x_10, x_11, x_12, x_13, x_14); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_17; +} +} +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___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) { +_start: +{ +lean_object* x_7; +x_7 = l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +size_t x_14; lean_object* x_15; +x_14 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_15 = l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__2(x_1, x_2, x_3, x_4, x_5, x_14, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_15; +} +} +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +size_t x_15; lean_object* x_16; +x_15 = lean_unbox_usize(x_8); +lean_dec(x_8); +x_16 = l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_15, x_9, x_10, x_11, x_12, x_13, x_14); +lean_dec(x_9); +lean_dec(x_2); +return x_16; +} +} +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +x_13 = l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_4); +lean_dec(x_3); +return x_13; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___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_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___lambda__1(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_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_instInhabitedMetaM___boxed), 5, 1); +lean_closure_set(x_1, 0, lean_box(0)); +return x_1; +} +} +LEAN_EXPORT lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___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; lean_object* x_9; +x_7 = l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__1___closed__1; +x_8 = lean_panic_fn(x_7, x_1); +x_9 = lean_apply_5(x_8, x_2, x_3, x_4, x_5, x_6); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldrM___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__2(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +lean_inc(x_1); +return x_1; +} +else +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_3 = lean_ctor_get(x_2, 0); +x_4 = lean_ctor_get(x_2, 2); +x_5 = l_Std_DHashMap_Internal_AssocList_foldrM___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__2(x_1, x_4); +lean_inc(x_3); +x_6 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_6, 0, x_3); +lean_ctor_set(x_6, 1, x_5); +return x_6; +} +} +} +LEAN_EXPORT lean_object* l_List_mapTR_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__3(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_Expr_fvarId_x21(x_5); +lean_dec(x_5); +lean_ctor_set(x_1, 1, x_2); +lean_ctor_set(x_1, 0, x_7); +{ +lean_object* _tmp_0 = x_6; +lean_object* _tmp_1 = x_1; +x_1 = _tmp_0; +x_2 = _tmp_1; +} +goto _start; +} +else +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_9 = lean_ctor_get(x_1, 0); +x_10 = lean_ctor_get(x_1, 1); +lean_inc(x_10); +lean_inc(x_9); +lean_dec(x_1); +x_11 = l_Lean_Expr_fvarId_x21(x_9); +lean_dec(x_9); +x_12 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_2); +x_1 = x_10; +x_2 = x_12; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_RBTree_ofList___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__4(lean_object* x_1) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_2; +x_2 = lean_box(0); +return x_2; +} +else +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_3 = lean_ctor_get(x_1, 0); +lean_inc(x_3); +x_4 = lean_ctor_get(x_1, 1); +lean_inc(x_4); +lean_dec(x_1); +x_5 = l_Lean_RBTree_ofList___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__4(x_4); +x_6 = lean_box(0); +x_7 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_5, x_3, x_6); +return x_7; +} +} +} +LEAN_EXPORT lean_object* l_List_mapTR_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__5(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_4) == 0) +{ +lean_object* x_6; +x_6 = l_List_reverse___rarg(x_5); +return x_6; +} +else +{ +uint8_t x_7; +x_7 = !lean_is_exclusive(x_4); +if (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; +x_8 = lean_ctor_get(x_4, 0); +x_9 = lean_ctor_get(x_4, 1); +x_10 = lean_nat_sub(x_1, x_8); +lean_dec(x_8); +x_11 = lean_unsigned_to_nat(1u); +x_12 = lean_nat_sub(x_10, x_11); +lean_dec(x_10); +x_13 = lean_nat_dec_lt(x_12, x_3); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; +lean_dec(x_12); +x_14 = l_Lean_instInhabitedExpr; +x_15 = l_outOfBounds___rarg(x_14); +x_16 = l_Lean_Expr_fvarId_x21(x_15); +lean_dec(x_15); +lean_ctor_set(x_4, 1, x_5); +lean_ctor_set(x_4, 0, x_16); +{ +lean_object* _tmp_3 = x_9; +lean_object* _tmp_4 = x_4; +x_4 = _tmp_3; +x_5 = _tmp_4; +} +goto _start; +} +else +{ +lean_object* x_18; lean_object* x_19; +x_18 = lean_array_fget(x_2, x_12); +lean_dec(x_12); +x_19 = l_Lean_Expr_fvarId_x21(x_18); +lean_dec(x_18); +lean_ctor_set(x_4, 1, x_5); +lean_ctor_set(x_4, 0, x_19); +{ +lean_object* _tmp_3 = x_9; +lean_object* _tmp_4 = x_4; +x_4 = _tmp_3; +x_5 = _tmp_4; +} +goto _start; +} +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; +x_21 = lean_ctor_get(x_4, 0); +x_22 = lean_ctor_get(x_4, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_4); +x_23 = lean_nat_sub(x_1, x_21); +lean_dec(x_21); +x_24 = lean_unsigned_to_nat(1u); +x_25 = lean_nat_sub(x_23, x_24); +lean_dec(x_23); +x_26 = lean_nat_dec_lt(x_25, x_3); +if (x_26 == 0) +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +lean_dec(x_25); +x_27 = l_Lean_instInhabitedExpr; +x_28 = l_outOfBounds___rarg(x_27); +x_29 = l_Lean_Expr_fvarId_x21(x_28); +lean_dec(x_28); +x_30 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_5); +x_4 = x_22; +x_5 = x_30; +goto _start; +} +else +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_array_fget(x_2, x_25); +lean_dec(x_25); +x_33 = l_Lean_Expr_fvarId_x21(x_32); +lean_dec(x_32); +x_34 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_34, 0, x_33); +lean_ctor_set(x_34, 1, x_5); +x_4 = x_22; +x_5 = x_34; +goto _start; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { +_start: +{ +uint8_t x_7; +x_7 = lean_usize_dec_eq(x_4, x_5); +if (x_7 == 0) +{ +lean_object* x_8; size_t x_9; size_t x_10; lean_object* x_11; +x_8 = lean_array_uget(x_3, x_4); +x_9 = 1; +x_10 = lean_usize_add(x_4, x_9); +x_11 = l_Lean_RBNode_findCore___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__2(x_2, x_8); +if (lean_obj_tag(x_11) == 0) +{ +lean_dec(x_8); +x_4 = x_10; +goto _start; +} +else +{ +lean_object* x_13; lean_object* x_14; +lean_dec(x_11); +x_13 = lean_box(0); +x_14 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_6, x_8, x_13); +x_4 = x_10; +x_6 = x_14; +goto _start; +} +} +else +{ +return x_6; +} +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__6___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__7(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; size_t x_8; size_t x_9; lean_object* x_10; +x_7 = lean_array_uget(x_2, x_3); +x_8 = 1; +x_9 = lean_usize_add(x_3, x_8); +x_10 = l_Lean_RBNode_findCore___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__2(x_1, x_7); +if (lean_obj_tag(x_10) == 0) +{ +lean_dec(x_7); +x_3 = x_9; +goto _start; +} +else +{ +lean_object* x_12; lean_object* x_13; +lean_dec(x_10); +x_12 = lean_box(0); +x_13 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_5, x_7, x_12); +x_3 = x_9; +x_5 = x_13; +goto _start; +} +} +else +{ +return x_5; +} +} +} +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___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) { +_start: +{ +if (lean_obj_tag(x_7) == 0) +{ +lean_object* x_15; +lean_dec(x_10); +lean_dec(x_1); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_8); +lean_ctor_set(x_15, 1, x_14); +return x_15; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_7, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_7, 1); +lean_inc(x_17); +lean_dec(x_7); +lean_inc(x_10); +x_18 = l_Lean_FVarId_getType(x_16, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_18) == 0) +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_18, 1); +lean_inc(x_20); +lean_dec(x_18); +x_21 = lean_box(0); +lean_inc(x_1); +x_22 = lean_array_mk(x_1); +x_23 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; +x_24 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set(x_24, 1, x_21); +lean_ctor_set(x_24, 2, x_22); +x_25 = l_Lean_CollectFVars_main(x_19, x_24); +x_26 = lean_ctor_get(x_25, 2); +lean_inc(x_26); +lean_dec(x_25); +x_27 = lean_array_get_size(x_26); +x_28 = lean_unsigned_to_nat(0u); +x_29 = lean_nat_dec_lt(x_28, x_27); +if (x_29 == 0) +{ +lean_dec(x_27); +lean_dec(x_26); +x_7 = x_17; +x_9 = lean_box(0); +x_14 = x_20; +goto _start; +} +else +{ +uint8_t x_31; +x_31 = lean_nat_dec_le(x_27, x_27); +if (x_31 == 0) +{ +lean_dec(x_27); +lean_dec(x_26); +x_7 = x_17; +x_9 = lean_box(0); +x_14 = x_20; +goto _start; +} +else +{ +size_t x_33; size_t x_34; lean_object* x_35; +x_33 = 0; +x_34 = lean_usize_of_nat(x_27); +lean_dec(x_27); +x_35 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__6___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__7(x_3, x_26, x_33, x_34, x_8); +lean_dec(x_26); +x_7 = x_17; +x_8 = x_35; +x_9 = lean_box(0); +x_14 = x_20; +goto _start; +} +} +} +else +{ +uint8_t x_37; +lean_dec(x_17); +lean_dec(x_10); +lean_dec(x_8); +lean_dec(x_1); +x_37 = !lean_is_exclusive(x_18); +if (x_37 == 0) +{ +return x_18; +} +else +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_38 = lean_ctor_get(x_18, 0); +x_39 = lean_ctor_get(x_18, 1); +lean_inc(x_39); +lean_inc(x_38); +lean_dec(x_18); +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_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__8___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__9(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +if (lean_obj_tag(x_6) == 0) +{ +lean_object* x_14; +lean_dec(x_9); +lean_dec(x_1); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_7); +lean_ctor_set(x_14, 1, x_13); +return x_14; +} +else +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_ctor_get(x_6, 0); +lean_inc(x_15); +x_16 = lean_ctor_get(x_6, 1); +lean_inc(x_16); +lean_dec(x_6); +lean_inc(x_9); +x_17 = l_Lean_FVarId_getType(x_15, x_9, x_10, x_11, x_12, 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_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_17, 0); +lean_inc(x_18); +x_19 = lean_ctor_get(x_17, 1); +lean_inc(x_19); +lean_dec(x_17); +x_20 = lean_box(0); +lean_inc(x_1); +x_21 = lean_array_mk(x_1); +x_22 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; +x_23 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_20); +lean_ctor_set(x_23, 2, x_21); +x_24 = l_Lean_CollectFVars_main(x_18, x_23); +x_25 = lean_ctor_get(x_24, 2); +lean_inc(x_25); +lean_dec(x_24); +x_26 = lean_array_get_size(x_25); +x_27 = lean_unsigned_to_nat(0u); +x_28 = lean_nat_dec_lt(x_27, x_26); +if (x_28 == 0) +{ +lean_dec(x_26); +lean_dec(x_25); +x_6 = x_16; +x_8 = lean_box(0); +x_13 = x_19; +goto _start; +} +else +{ +uint8_t x_30; +x_30 = lean_nat_dec_le(x_26, x_26); +if (x_30 == 0) +{ +lean_dec(x_26); +lean_dec(x_25); +x_6 = x_16; +x_8 = lean_box(0); +x_13 = x_19; +goto _start; +} +else +{ +size_t x_32; size_t x_33; lean_object* x_34; +x_32 = 0; +x_33 = lean_usize_of_nat(x_26); +lean_dec(x_26); +x_34 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__6___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__7(x_2, x_25, x_32, x_33, x_7); +lean_dec(x_25); +x_6 = x_16; +x_7 = x_34; +x_8 = lean_box(0); +x_13 = x_19; +goto _start; +} +} +} +else +{ +uint8_t x_36; +lean_dec(x_16); +lean_dec(x_9); +lean_dec(x_7); +lean_dec(x_1); +x_36 = !lean_is_exclusive(x_17); +if (x_36 == 0) +{ +return x_17; +} +else +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_37 = lean_ctor_get(x_17, 0); +x_38 = lean_ctor_get(x_17, 1); +lean_inc(x_38); +lean_inc(x_37); +lean_dec(x_17); +x_39 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_39, 0, x_37); +lean_ctor_set(x_39, 1, x_38); +return x_39; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__10(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { +_start: +{ +uint8_t x_7; +x_7 = lean_usize_dec_eq(x_4, x_5); +if (x_7 == 0) +{ +lean_object* x_8; size_t x_9; size_t x_10; lean_object* x_11; +x_8 = lean_array_uget(x_3, x_4); +x_9 = 1; +x_10 = lean_usize_add(x_4, x_9); +x_11 = l_Lean_RBNode_findCore___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__2(x_2, x_8); +if (lean_obj_tag(x_11) == 0) +{ +lean_dec(x_8); +x_4 = x_10; +goto _start; +} +else +{ +lean_object* x_13; lean_object* x_14; +lean_dec(x_11); +x_13 = lean_box(0); +x_14 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_6, x_8, x_13); +x_4 = x_10; +x_6 = x_14; +goto _start; +} +} +else +{ +return x_6; +} +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__10___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__11(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5) { +_start: +{ +uint8_t x_6; +x_6 = lean_usize_dec_eq(x_3, x_4); +if (x_6 == 0) +{ +lean_object* x_7; size_t x_8; size_t x_9; lean_object* x_10; +x_7 = lean_array_uget(x_2, x_3); +x_8 = 1; +x_9 = lean_usize_add(x_3, x_8); +x_10 = l_Lean_RBNode_findCore___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__2(x_1, x_7); +if (lean_obj_tag(x_10) == 0) +{ +lean_dec(x_7); +x_3 = x_9; +goto _start; +} +else +{ +lean_object* x_12; lean_object* x_13; +lean_dec(x_10); +x_12 = lean_box(0); +x_13 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_5, x_7, x_12); +x_3 = x_9; +x_5 = x_13; +goto _start; +} +} +else +{ +return x_5; +} +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__12(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { +_start: +{ +uint8_t x_7; +x_7 = lean_usize_dec_eq(x_4, x_5); +if (x_7 == 0) +{ +lean_object* x_8; size_t x_9; size_t x_10; lean_object* x_11; +x_8 = lean_array_uget(x_3, x_4); +x_9 = 1; +x_10 = lean_usize_add(x_4, x_9); +x_11 = l_Lean_RBNode_findCore___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__2(x_2, x_8); +if (lean_obj_tag(x_11) == 0) +{ +lean_dec(x_8); +x_4 = x_10; +goto _start; +} +else +{ +lean_object* x_13; lean_object* x_14; +lean_dec(x_11); +x_13 = lean_box(0); +x_14 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_6, x_8, x_13); +x_4 = x_10; +x_6 = x_14; +goto _start; +} +} +else +{ +return x_6; +} +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__12___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__13(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5) { +_start: +{ +uint8_t x_6; +x_6 = lean_usize_dec_eq(x_3, x_4); +if (x_6 == 0) +{ +lean_object* x_7; size_t x_8; size_t x_9; lean_object* x_10; +x_7 = lean_array_uget(x_2, x_3); +x_8 = 1; +x_9 = lean_usize_add(x_3, x_8); +x_10 = l_Lean_RBNode_findCore___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__2(x_1, x_7); +if (lean_obj_tag(x_10) == 0) +{ +lean_dec(x_7); +x_3 = x_9; +goto _start; +} +else +{ +lean_object* x_12; lean_object* x_13; +lean_dec(x_10); +x_12 = lean_box(0); +x_13 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_5, x_7, x_12); +x_3 = x_9; +x_5 = x_13; +goto _start; +} +} +else +{ +return x_5; +} +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__14(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, size_t x_7, size_t x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +uint8_t x_15; +x_15 = lean_usize_dec_lt(x_8, x_7); +if (x_15 == 0) +{ +lean_object* x_16; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_4); +lean_dec(x_2); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_9); +lean_ctor_set(x_16, 1, x_14); +return x_16; +} +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_25; +x_17 = lean_array_uget(x_6, x_8); +x_25 = !lean_is_exclusive(x_9); +if (x_25 == 0) +{ +lean_object* x_26; uint8_t x_27; +x_26 = lean_ctor_get(x_9, 1); +x_27 = !lean_is_exclusive(x_26); +if (x_27 == 0) +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_28 = lean_ctor_get(x_9, 0); +x_29 = lean_ctor_get(x_26, 0); +x_30 = lean_ctor_get(x_26, 1); +x_31 = l_Lean_Expr_fvarId_x21(x_17); +x_32 = l_Lean_RBNode_findCore___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__2(x_30, x_31); +if (lean_obj_tag(x_32) == 0) +{ +lean_object* x_33; +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +x_33 = lean_infer_type(x_17, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_33) == 0) +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_33, 0); +lean_inc(x_34); +x_35 = lean_ctor_get(x_33, 1); +lean_inc(x_35); +lean_dec(x_33); +x_36 = l_Lean_RBNode_findCore___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__2(x_28, x_31); +if (lean_obj_tag(x_36) == 0) +{ +lean_object* x_37; +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_34); +x_37 = l_Lean_Meta_isProp(x_34, x_10, x_11, x_12, x_13, x_35); +if (lean_obj_tag(x_37) == 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_40; lean_object* x_41; +x_40 = lean_ctor_get(x_37, 1); +lean_inc(x_40); +lean_dec(x_37); +lean_inc(x_10); +lean_inc(x_31); +x_41 = l_Lean_FVarId_getDecl(x_31, x_10, x_11, x_12, x_13, x_40); +if (lean_obj_tag(x_41) == 0) +{ +lean_object* x_42; lean_object* x_43; uint8_t 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 = l_Lean_LocalDecl_binderInfo(x_42); +lean_dec(x_42); +x_45 = lean_box(x_44); +if (lean_obj_tag(x_45) == 3) +{ +lean_object* x_46; +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_34); +lean_inc(x_28); +lean_inc(x_4); +x_46 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized(x_4, x_28, x_34, x_10, x_11, x_12, x_13, x_43); +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); +lean_dec(x_47); +if (x_48 == 0) +{ +lean_object* x_49; lean_object* x_50; +lean_dec(x_34); +lean_dec(x_31); +x_49 = lean_ctor_get(x_46, 1); +lean_inc(x_49); +lean_dec(x_46); +x_50 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_50, 0, x_9); +x_18 = x_50; +x_19 = x_49; +goto block_24; +} +else +{ +lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; uint8_t x_62; lean_object* x_63; +lean_dec(x_29); +x_51 = lean_ctor_get(x_46, 1); +lean_inc(x_51); +lean_dec(x_46); +x_52 = lean_box(0); +lean_inc(x_31); +x_53 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_28, x_31, x_52); +x_54 = lean_box(0); +lean_inc(x_2); +x_55 = lean_array_mk(x_2); +x_56 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; +x_57 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_57, 0, x_56); +lean_ctor_set(x_57, 1, x_54); +lean_ctor_set(x_57, 2, x_55); +x_58 = l_Lean_CollectFVars_main(x_34, x_57); +x_59 = lean_ctor_get(x_58, 2); +lean_inc(x_59); +lean_dec(x_58); +x_60 = lean_array_get_size(x_59); +x_61 = lean_unsigned_to_nat(0u); +x_62 = lean_nat_dec_lt(x_61, x_60); +x_63 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_30, x_31, x_52); +if (x_62 == 0) +{ +uint8_t x_64; lean_object* x_65; lean_object* x_66; +lean_dec(x_60); +lean_dec(x_59); +x_64 = 1; +x_65 = lean_box(x_64); +lean_ctor_set(x_26, 1, x_63); +lean_ctor_set(x_26, 0, x_65); +lean_ctor_set(x_9, 0, x_53); +x_66 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_66, 0, x_9); +x_18 = x_66; +x_19 = x_51; +goto block_24; +} +else +{ +uint8_t x_67; +x_67 = lean_nat_dec_le(x_60, x_60); +if (x_67 == 0) +{ +uint8_t x_68; lean_object* x_69; lean_object* x_70; +lean_dec(x_60); +lean_dec(x_59); +x_68 = 1; +x_69 = lean_box(x_68); +lean_ctor_set(x_26, 1, x_63); +lean_ctor_set(x_26, 0, x_69); +lean_ctor_set(x_9, 0, x_53); +x_70 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_70, 0, x_9); +x_18 = x_70; +x_19 = x_51; +goto block_24; +} +else +{ +size_t x_71; size_t x_72; lean_object* x_73; uint8_t x_74; lean_object* x_75; lean_object* x_76; +x_71 = 0; +x_72 = lean_usize_of_nat(x_60); +lean_dec(x_60); +x_73 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__10___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__11(x_4, x_59, x_71, x_72, x_53); +lean_dec(x_59); +x_74 = 1; +x_75 = lean_box(x_74); +lean_ctor_set(x_26, 1, x_63); +lean_ctor_set(x_26, 0, x_75); +lean_ctor_set(x_9, 0, x_73); +x_76 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_76, 0, x_9); +x_18 = x_76; +x_19 = x_51; +goto block_24; +} +} +} +} +else +{ +uint8_t x_77; +lean_dec(x_34); +lean_dec(x_31); +lean_free_object(x_26); +lean_dec(x_30); +lean_dec(x_29); +lean_free_object(x_9); +lean_dec(x_28); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_4); +lean_dec(x_2); +x_77 = !lean_is_exclusive(x_46); +if (x_77 == 0) +{ +return x_46; +} +else +{ +lean_object* x_78; lean_object* x_79; lean_object* x_80; +x_78 = lean_ctor_get(x_46, 0); +x_79 = lean_ctor_get(x_46, 1); +lean_inc(x_79); +lean_inc(x_78); +lean_dec(x_46); +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_81; +lean_dec(x_45); +lean_dec(x_34); +lean_dec(x_31); +x_81 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_81, 0, x_9); +x_18 = x_81; +x_19 = x_43; +goto block_24; +} +} +else +{ +uint8_t x_82; +lean_dec(x_34); +lean_dec(x_31); +lean_free_object(x_26); +lean_dec(x_30); +lean_dec(x_29); +lean_free_object(x_9); +lean_dec(x_28); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_4); +lean_dec(x_2); +x_82 = !lean_is_exclusive(x_41); +if (x_82 == 0) +{ +return x_41; +} +else +{ +lean_object* x_83; lean_object* x_84; lean_object* x_85; +x_83 = lean_ctor_get(x_41, 0); +x_84 = lean_ctor_get(x_41, 1); +lean_inc(x_84); +lean_inc(x_83); +lean_dec(x_41); +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_86; uint8_t x_87; +x_86 = lean_ctor_get(x_37, 1); +lean_inc(x_86); +lean_dec(x_37); +x_87 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars(x_4, x_28, x_34); +if (x_87 == 0) +{ +lean_object* x_88; +lean_dec(x_31); +x_88 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_88, 0, x_9); +x_18 = x_88; +x_19 = x_86; +goto block_24; +} +else +{ +lean_object* x_89; lean_object* x_90; lean_object* x_91; uint8_t x_92; lean_object* x_93; lean_object* x_94; +lean_dec(x_29); +x_89 = lean_box(0); +lean_inc(x_31); +x_90 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_28, x_31, x_89); +x_91 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_30, x_31, x_89); +x_92 = 1; +x_93 = lean_box(x_92); +lean_ctor_set(x_26, 1, x_91); +lean_ctor_set(x_26, 0, x_93); +lean_ctor_set(x_9, 0, x_90); +x_94 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_94, 0, x_9); +x_18 = x_94; +x_19 = x_86; +goto block_24; +} +} +} +else +{ +uint8_t x_95; +lean_dec(x_34); +lean_dec(x_31); +lean_free_object(x_26); +lean_dec(x_30); +lean_dec(x_29); +lean_free_object(x_9); +lean_dec(x_28); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_4); +lean_dec(x_2); +x_95 = !lean_is_exclusive(x_37); +if (x_95 == 0) +{ +return x_37; +} +else +{ +lean_object* x_96; lean_object* x_97; lean_object* x_98; +x_96 = lean_ctor_get(x_37, 0); +x_97 = lean_ctor_get(x_37, 1); +lean_inc(x_97); +lean_inc(x_96); +lean_dec(x_37); +x_98 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_98, 0, x_96); +lean_ctor_set(x_98, 1, x_97); +return x_98; +} +} +} +else +{ +uint8_t x_99; +lean_dec(x_29); +x_99 = !lean_is_exclusive(x_36); +if (x_99 == 0) +{ +lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; uint8_t x_109; lean_object* x_110; lean_object* x_111; +x_100 = lean_ctor_get(x_36, 0); +lean_dec(x_100); +x_101 = lean_box(0); +lean_inc(x_2); +x_102 = lean_array_mk(x_2); +x_103 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; +x_104 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_104, 0, x_103); +lean_ctor_set(x_104, 1, x_101); +lean_ctor_set(x_104, 2, x_102); +x_105 = l_Lean_CollectFVars_main(x_34, x_104); +x_106 = lean_ctor_get(x_105, 2); +lean_inc(x_106); +lean_dec(x_105); +x_107 = lean_array_get_size(x_106); +x_108 = lean_unsigned_to_nat(0u); +x_109 = lean_nat_dec_lt(x_108, x_107); +x_110 = lean_box(0); +x_111 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_30, x_31, x_110); +if (x_109 == 0) +{ +uint8_t x_112; lean_object* x_113; +lean_dec(x_107); +lean_dec(x_106); +x_112 = 1; +x_113 = lean_box(x_112); +lean_ctor_set(x_26, 1, x_111); +lean_ctor_set(x_26, 0, x_113); +lean_ctor_set(x_36, 0, x_9); +x_18 = x_36; +x_19 = x_35; +goto block_24; +} +else +{ +uint8_t x_114; +x_114 = lean_nat_dec_le(x_107, x_107); +if (x_114 == 0) +{ +uint8_t x_115; lean_object* x_116; +lean_dec(x_107); +lean_dec(x_106); +x_115 = 1; +x_116 = lean_box(x_115); +lean_ctor_set(x_26, 1, x_111); +lean_ctor_set(x_26, 0, x_116); +lean_ctor_set(x_36, 0, x_9); +x_18 = x_36; +x_19 = x_35; +goto block_24; +} +else +{ +size_t x_117; size_t x_118; lean_object* x_119; uint8_t x_120; lean_object* x_121; +x_117 = 0; +x_118 = lean_usize_of_nat(x_107); +lean_dec(x_107); +x_119 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__12___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__13(x_4, x_106, x_117, x_118, x_28); +lean_dec(x_106); +x_120 = 1; +x_121 = lean_box(x_120); +lean_ctor_set(x_26, 1, x_111); +lean_ctor_set(x_26, 0, x_121); +lean_ctor_set(x_9, 0, x_119); +lean_ctor_set(x_36, 0, x_9); +x_18 = x_36; +x_19 = x_35; +goto block_24; +} +} +} +else +{ +lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; uint8_t x_130; lean_object* x_131; lean_object* x_132; +lean_dec(x_36); +x_122 = lean_box(0); +lean_inc(x_2); +x_123 = lean_array_mk(x_2); +x_124 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; +x_125 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_125, 0, x_124); +lean_ctor_set(x_125, 1, x_122); +lean_ctor_set(x_125, 2, x_123); +x_126 = l_Lean_CollectFVars_main(x_34, x_125); +x_127 = lean_ctor_get(x_126, 2); +lean_inc(x_127); +lean_dec(x_126); +x_128 = lean_array_get_size(x_127); +x_129 = lean_unsigned_to_nat(0u); +x_130 = lean_nat_dec_lt(x_129, x_128); +x_131 = lean_box(0); +x_132 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_30, x_31, x_131); +if (x_130 == 0) +{ +uint8_t x_133; lean_object* x_134; lean_object* x_135; +lean_dec(x_128); +lean_dec(x_127); +x_133 = 1; +x_134 = lean_box(x_133); +lean_ctor_set(x_26, 1, x_132); +lean_ctor_set(x_26, 0, x_134); +x_135 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_135, 0, x_9); +x_18 = x_135; +x_19 = x_35; +goto block_24; +} +else +{ +uint8_t x_136; +x_136 = lean_nat_dec_le(x_128, x_128); +if (x_136 == 0) +{ +uint8_t x_137; lean_object* x_138; lean_object* x_139; +lean_dec(x_128); +lean_dec(x_127); +x_137 = 1; +x_138 = lean_box(x_137); +lean_ctor_set(x_26, 1, x_132); +lean_ctor_set(x_26, 0, x_138); +x_139 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_139, 0, x_9); +x_18 = x_139; +x_19 = x_35; +goto block_24; +} +else +{ +size_t x_140; size_t x_141; lean_object* x_142; uint8_t x_143; lean_object* x_144; lean_object* x_145; +x_140 = 0; +x_141 = lean_usize_of_nat(x_128); +lean_dec(x_128); +x_142 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__12___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__13(x_4, x_127, x_140, x_141, x_28); +lean_dec(x_127); +x_143 = 1; +x_144 = lean_box(x_143); +lean_ctor_set(x_26, 1, x_132); +lean_ctor_set(x_26, 0, x_144); +lean_ctor_set(x_9, 0, x_142); +x_145 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_145, 0, x_9); +x_18 = x_145; +x_19 = x_35; +goto block_24; +} +} +} +} +} +else +{ +uint8_t x_146; +lean_dec(x_31); +lean_free_object(x_26); +lean_dec(x_30); +lean_dec(x_29); +lean_free_object(x_9); +lean_dec(x_28); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_4); +lean_dec(x_2); +x_146 = !lean_is_exclusive(x_33); +if (x_146 == 0) +{ +return x_33; +} +else +{ +lean_object* x_147; lean_object* x_148; lean_object* x_149; +x_147 = lean_ctor_get(x_33, 0); +x_148 = lean_ctor_get(x_33, 1); +lean_inc(x_148); +lean_inc(x_147); +lean_dec(x_33); +x_149 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_149, 0, x_147); +lean_ctor_set(x_149, 1, x_148); +return x_149; +} +} +} +else +{ +uint8_t x_150; +lean_dec(x_31); +lean_dec(x_17); +x_150 = !lean_is_exclusive(x_32); +if (x_150 == 0) +{ +lean_object* x_151; +x_151 = lean_ctor_get(x_32, 0); +lean_dec(x_151); +lean_ctor_set(x_32, 0, x_9); +x_18 = x_32; +x_19 = x_14; +goto block_24; +} +else +{ +lean_object* x_152; +lean_dec(x_32); +x_152 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_152, 0, x_9); +x_18 = x_152; +x_19 = x_14; +goto block_24; +} +} +} +else +{ +lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; +x_153 = lean_ctor_get(x_9, 0); +x_154 = lean_ctor_get(x_26, 0); +x_155 = lean_ctor_get(x_26, 1); +lean_inc(x_155); +lean_inc(x_154); +lean_dec(x_26); +x_156 = l_Lean_Expr_fvarId_x21(x_17); +x_157 = l_Lean_RBNode_findCore___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__2(x_155, x_156); +if (lean_obj_tag(x_157) == 0) +{ +lean_object* x_158; +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +x_158 = lean_infer_type(x_17, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_158) == 0) +{ +lean_object* x_159; lean_object* x_160; lean_object* x_161; +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_RBNode_findCore___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__2(x_153, x_156); +if (lean_obj_tag(x_161) == 0) +{ +lean_object* x_162; +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_159); +x_162 = l_Lean_Meta_isProp(x_159, x_10, x_11, x_12, x_13, x_160); +if (lean_obj_tag(x_162) == 0) +{ +lean_object* x_163; uint8_t x_164; +x_163 = lean_ctor_get(x_162, 0); +lean_inc(x_163); +x_164 = lean_unbox(x_163); +lean_dec(x_163); +if (x_164 == 0) +{ +lean_object* x_165; lean_object* x_166; +x_165 = lean_ctor_get(x_162, 1); +lean_inc(x_165); +lean_dec(x_162); +lean_inc(x_10); +lean_inc(x_156); +x_166 = l_Lean_FVarId_getDecl(x_156, x_10, x_11, x_12, x_13, x_165); +if (lean_obj_tag(x_166) == 0) +{ +lean_object* x_167; lean_object* x_168; uint8_t x_169; lean_object* x_170; +x_167 = lean_ctor_get(x_166, 0); +lean_inc(x_167); +x_168 = lean_ctor_get(x_166, 1); +lean_inc(x_168); +lean_dec(x_166); +x_169 = l_Lean_LocalDecl_binderInfo(x_167); +lean_dec(x_167); +x_170 = lean_box(x_169); +if (lean_obj_tag(x_170) == 3) +{ +lean_object* x_171; +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_159); +lean_inc(x_153); +lean_inc(x_4); +x_171 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized(x_4, x_153, x_159, x_10, x_11, x_12, x_13, x_168); +if (lean_obj_tag(x_171) == 0) +{ +lean_object* x_172; uint8_t x_173; +x_172 = lean_ctor_get(x_171, 0); +lean_inc(x_172); +x_173 = lean_unbox(x_172); +lean_dec(x_172); +if (x_173 == 0) +{ +lean_object* x_174; lean_object* x_175; lean_object* x_176; +lean_dec(x_159); +lean_dec(x_156); +x_174 = lean_ctor_get(x_171, 1); +lean_inc(x_174); +lean_dec(x_171); +x_175 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_175, 0, x_154); +lean_ctor_set(x_175, 1, x_155); +lean_ctor_set(x_9, 1, x_175); +x_176 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_176, 0, x_9); +x_18 = x_176; +x_19 = x_174; +goto block_24; +} +else +{ +lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; uint8_t x_188; lean_object* x_189; +lean_dec(x_154); +x_177 = lean_ctor_get(x_171, 1); +lean_inc(x_177); +lean_dec(x_171); +x_178 = lean_box(0); +lean_inc(x_156); +x_179 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_153, x_156, x_178); +x_180 = lean_box(0); +lean_inc(x_2); +x_181 = lean_array_mk(x_2); +x_182 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; +x_183 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_183, 0, x_182); +lean_ctor_set(x_183, 1, x_180); +lean_ctor_set(x_183, 2, x_181); +x_184 = l_Lean_CollectFVars_main(x_159, x_183); +x_185 = lean_ctor_get(x_184, 2); +lean_inc(x_185); +lean_dec(x_184); +x_186 = lean_array_get_size(x_185); +x_187 = lean_unsigned_to_nat(0u); +x_188 = lean_nat_dec_lt(x_187, x_186); +x_189 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_155, x_156, x_178); +if (x_188 == 0) +{ +uint8_t x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; +lean_dec(x_186); +lean_dec(x_185); +x_190 = 1; +x_191 = lean_box(x_190); +x_192 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_192, 0, x_191); +lean_ctor_set(x_192, 1, x_189); +lean_ctor_set(x_9, 1, x_192); +lean_ctor_set(x_9, 0, x_179); +x_193 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_193, 0, x_9); +x_18 = x_193; +x_19 = x_177; +goto block_24; +} +else +{ +uint8_t x_194; +x_194 = lean_nat_dec_le(x_186, x_186); +if (x_194 == 0) +{ +uint8_t x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; +lean_dec(x_186); +lean_dec(x_185); +x_195 = 1; +x_196 = lean_box(x_195); +x_197 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_197, 0, x_196); +lean_ctor_set(x_197, 1, x_189); +lean_ctor_set(x_9, 1, x_197); +lean_ctor_set(x_9, 0, x_179); +x_198 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_198, 0, x_9); +x_18 = x_198; +x_19 = x_177; +goto block_24; +} +else +{ +size_t x_199; size_t x_200; lean_object* x_201; uint8_t x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; +x_199 = 0; +x_200 = lean_usize_of_nat(x_186); +lean_dec(x_186); +x_201 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__10___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__11(x_4, x_185, x_199, x_200, x_179); +lean_dec(x_185); +x_202 = 1; +x_203 = lean_box(x_202); +x_204 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_204, 0, x_203); +lean_ctor_set(x_204, 1, x_189); +lean_ctor_set(x_9, 1, x_204); +lean_ctor_set(x_9, 0, x_201); +x_205 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_205, 0, x_9); +x_18 = x_205; +x_19 = x_177; +goto block_24; +} +} +} +} +else +{ +lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; +lean_dec(x_159); +lean_dec(x_156); +lean_dec(x_155); +lean_dec(x_154); +lean_free_object(x_9); +lean_dec(x_153); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_4); +lean_dec(x_2); +x_206 = lean_ctor_get(x_171, 0); +lean_inc(x_206); +x_207 = lean_ctor_get(x_171, 1); +lean_inc(x_207); +if (lean_is_exclusive(x_171)) { + lean_ctor_release(x_171, 0); + lean_ctor_release(x_171, 1); + x_208 = x_171; +} else { + lean_dec_ref(x_171); + x_208 = lean_box(0); +} +if (lean_is_scalar(x_208)) { + x_209 = lean_alloc_ctor(1, 2, 0); +} else { + x_209 = x_208; +} +lean_ctor_set(x_209, 0, x_206); +lean_ctor_set(x_209, 1, x_207); +return x_209; +} +} +else +{ +lean_object* x_210; lean_object* x_211; +lean_dec(x_170); +lean_dec(x_159); +lean_dec(x_156); +x_210 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_210, 0, x_154); +lean_ctor_set(x_210, 1, x_155); +lean_ctor_set(x_9, 1, x_210); +x_211 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_211, 0, x_9); +x_18 = x_211; +x_19 = x_168; +goto block_24; +} +} +else +{ +lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; +lean_dec(x_159); +lean_dec(x_156); +lean_dec(x_155); +lean_dec(x_154); +lean_free_object(x_9); +lean_dec(x_153); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_4); +lean_dec(x_2); +x_212 = lean_ctor_get(x_166, 0); +lean_inc(x_212); +x_213 = lean_ctor_get(x_166, 1); +lean_inc(x_213); +if (lean_is_exclusive(x_166)) { + lean_ctor_release(x_166, 0); + lean_ctor_release(x_166, 1); + x_214 = x_166; +} else { + lean_dec_ref(x_166); + x_214 = lean_box(0); +} +if (lean_is_scalar(x_214)) { + x_215 = lean_alloc_ctor(1, 2, 0); +} else { + x_215 = x_214; +} +lean_ctor_set(x_215, 0, x_212); +lean_ctor_set(x_215, 1, x_213); +return x_215; +} +} +else +{ +lean_object* x_216; uint8_t x_217; +x_216 = lean_ctor_get(x_162, 1); +lean_inc(x_216); +lean_dec(x_162); +x_217 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars(x_4, x_153, x_159); +if (x_217 == 0) +{ +lean_object* x_218; lean_object* x_219; +lean_dec(x_156); +x_218 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_218, 0, x_154); +lean_ctor_set(x_218, 1, x_155); +lean_ctor_set(x_9, 1, x_218); +x_219 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_219, 0, x_9); +x_18 = x_219; +x_19 = x_216; +goto block_24; +} +else +{ +lean_object* x_220; lean_object* x_221; lean_object* x_222; uint8_t x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; +lean_dec(x_154); +x_220 = lean_box(0); +lean_inc(x_156); +x_221 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_153, x_156, x_220); +x_222 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_155, x_156, x_220); +x_223 = 1; +x_224 = lean_box(x_223); +x_225 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_225, 0, x_224); +lean_ctor_set(x_225, 1, x_222); +lean_ctor_set(x_9, 1, x_225); +lean_ctor_set(x_9, 0, x_221); +x_226 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_226, 0, x_9); +x_18 = x_226; +x_19 = x_216; +goto block_24; +} +} +} +else +{ +lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; +lean_dec(x_159); +lean_dec(x_156); +lean_dec(x_155); +lean_dec(x_154); +lean_free_object(x_9); +lean_dec(x_153); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_4); +lean_dec(x_2); +x_227 = lean_ctor_get(x_162, 0); +lean_inc(x_227); +x_228 = lean_ctor_get(x_162, 1); +lean_inc(x_228); +if (lean_is_exclusive(x_162)) { + lean_ctor_release(x_162, 0); + lean_ctor_release(x_162, 1); + x_229 = x_162; +} else { + lean_dec_ref(x_162); + x_229 = lean_box(0); +} +if (lean_is_scalar(x_229)) { + x_230 = lean_alloc_ctor(1, 2, 0); +} else { + x_230 = x_229; +} +lean_ctor_set(x_230, 0, x_227); +lean_ctor_set(x_230, 1, x_228); +return x_230; +} +} +else +{ +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; uint8_t x_240; lean_object* x_241; lean_object* x_242; +lean_dec(x_154); +if (lean_is_exclusive(x_161)) { + lean_ctor_release(x_161, 0); + x_231 = x_161; +} else { + lean_dec_ref(x_161); + x_231 = lean_box(0); +} +x_232 = lean_box(0); +lean_inc(x_2); +x_233 = lean_array_mk(x_2); +x_234 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; +x_235 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_235, 0, x_234); +lean_ctor_set(x_235, 1, x_232); +lean_ctor_set(x_235, 2, x_233); +x_236 = l_Lean_CollectFVars_main(x_159, x_235); +x_237 = lean_ctor_get(x_236, 2); +lean_inc(x_237); +lean_dec(x_236); +x_238 = lean_array_get_size(x_237); +x_239 = lean_unsigned_to_nat(0u); +x_240 = lean_nat_dec_lt(x_239, x_238); +x_241 = lean_box(0); +x_242 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_155, x_156, x_241); +if (x_240 == 0) +{ +uint8_t x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; +lean_dec(x_238); +lean_dec(x_237); +x_243 = 1; +x_244 = lean_box(x_243); +x_245 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_245, 0, x_244); +lean_ctor_set(x_245, 1, x_242); +lean_ctor_set(x_9, 1, x_245); +if (lean_is_scalar(x_231)) { + x_246 = lean_alloc_ctor(1, 1, 0); +} else { + x_246 = x_231; +} +lean_ctor_set(x_246, 0, x_9); +x_18 = x_246; +x_19 = x_160; +goto block_24; +} +else +{ +uint8_t x_247; +x_247 = lean_nat_dec_le(x_238, x_238); +if (x_247 == 0) +{ +uint8_t x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; +lean_dec(x_238); +lean_dec(x_237); +x_248 = 1; +x_249 = lean_box(x_248); +x_250 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_250, 0, x_249); +lean_ctor_set(x_250, 1, x_242); +lean_ctor_set(x_9, 1, x_250); +if (lean_is_scalar(x_231)) { + x_251 = lean_alloc_ctor(1, 1, 0); +} else { + x_251 = x_231; +} +lean_ctor_set(x_251, 0, x_9); +x_18 = x_251; +x_19 = x_160; +goto block_24; +} +else +{ +size_t x_252; size_t x_253; lean_object* x_254; uint8_t x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; +x_252 = 0; +x_253 = lean_usize_of_nat(x_238); +lean_dec(x_238); +x_254 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__12___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__13(x_4, x_237, x_252, x_253, x_153); +lean_dec(x_237); +x_255 = 1; +x_256 = lean_box(x_255); +x_257 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_257, 0, x_256); +lean_ctor_set(x_257, 1, x_242); +lean_ctor_set(x_9, 1, x_257); +lean_ctor_set(x_9, 0, x_254); +if (lean_is_scalar(x_231)) { + x_258 = lean_alloc_ctor(1, 1, 0); +} else { + x_258 = x_231; +} +lean_ctor_set(x_258, 0, x_9); +x_18 = x_258; +x_19 = x_160; +goto block_24; +} +} +} +} +else +{ +lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; +lean_dec(x_156); +lean_dec(x_155); +lean_dec(x_154); +lean_free_object(x_9); +lean_dec(x_153); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_4); +lean_dec(x_2); +x_259 = lean_ctor_get(x_158, 0); +lean_inc(x_259); +x_260 = lean_ctor_get(x_158, 1); +lean_inc(x_260); +if (lean_is_exclusive(x_158)) { + lean_ctor_release(x_158, 0); + lean_ctor_release(x_158, 1); + x_261 = x_158; +} else { + lean_dec_ref(x_158); + x_261 = lean_box(0); +} +if (lean_is_scalar(x_261)) { + x_262 = lean_alloc_ctor(1, 2, 0); +} else { + x_262 = x_261; +} +lean_ctor_set(x_262, 0, x_259); +lean_ctor_set(x_262, 1, x_260); +return x_262; +} +} +else +{ +lean_object* x_263; lean_object* x_264; lean_object* x_265; +lean_dec(x_156); +lean_dec(x_17); +if (lean_is_exclusive(x_157)) { + lean_ctor_release(x_157, 0); + x_263 = x_157; +} else { + lean_dec_ref(x_157); + x_263 = lean_box(0); +} +x_264 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_264, 0, x_154); +lean_ctor_set(x_264, 1, x_155); +lean_ctor_set(x_9, 1, x_264); +if (lean_is_scalar(x_263)) { + x_265 = lean_alloc_ctor(1, 1, 0); +} else { + x_265 = x_263; +} +lean_ctor_set(x_265, 0, x_9); +x_18 = x_265; +x_19 = x_14; +goto block_24; +} +} +} +else +{ +lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; +x_266 = lean_ctor_get(x_9, 1); +x_267 = lean_ctor_get(x_9, 0); +lean_inc(x_266); +lean_inc(x_267); +lean_dec(x_9); +x_268 = lean_ctor_get(x_266, 0); +lean_inc(x_268); +x_269 = lean_ctor_get(x_266, 1); +lean_inc(x_269); +if (lean_is_exclusive(x_266)) { + lean_ctor_release(x_266, 0); + lean_ctor_release(x_266, 1); + x_270 = x_266; +} else { + lean_dec_ref(x_266); + x_270 = lean_box(0); +} +x_271 = l_Lean_Expr_fvarId_x21(x_17); +x_272 = l_Lean_RBNode_findCore___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__2(x_269, x_271); +if (lean_obj_tag(x_272) == 0) +{ +lean_object* x_273; +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +x_273 = lean_infer_type(x_17, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_273) == 0) +{ +lean_object* x_274; lean_object* x_275; lean_object* x_276; +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 = l_Lean_RBNode_findCore___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__2(x_267, x_271); +if (lean_obj_tag(x_276) == 0) +{ +lean_object* x_277; +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_274); +x_277 = l_Lean_Meta_isProp(x_274, x_10, x_11, x_12, x_13, x_275); +if (lean_obj_tag(x_277) == 0) +{ +lean_object* x_278; uint8_t x_279; +x_278 = lean_ctor_get(x_277, 0); +lean_inc(x_278); +x_279 = lean_unbox(x_278); +lean_dec(x_278); +if (x_279 == 0) +{ +lean_object* x_280; lean_object* x_281; +x_280 = lean_ctor_get(x_277, 1); +lean_inc(x_280); +lean_dec(x_277); +lean_inc(x_10); +lean_inc(x_271); +x_281 = l_Lean_FVarId_getDecl(x_271, x_10, x_11, x_12, x_13, x_280); +if (lean_obj_tag(x_281) == 0) +{ +lean_object* x_282; lean_object* x_283; uint8_t x_284; lean_object* x_285; +x_282 = lean_ctor_get(x_281, 0); +lean_inc(x_282); +x_283 = lean_ctor_get(x_281, 1); +lean_inc(x_283); +lean_dec(x_281); +x_284 = l_Lean_LocalDecl_binderInfo(x_282); +lean_dec(x_282); +x_285 = lean_box(x_284); +if (lean_obj_tag(x_285) == 3) +{ +lean_object* x_286; +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_274); +lean_inc(x_267); +lean_inc(x_4); +x_286 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized(x_4, x_267, x_274, x_10, x_11, x_12, x_13, x_283); +if (lean_obj_tag(x_286) == 0) +{ +lean_object* x_287; uint8_t x_288; +x_287 = lean_ctor_get(x_286, 0); +lean_inc(x_287); +x_288 = lean_unbox(x_287); +lean_dec(x_287); +if (x_288 == 0) +{ +lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; +lean_dec(x_274); +lean_dec(x_271); +x_289 = lean_ctor_get(x_286, 1); +lean_inc(x_289); +lean_dec(x_286); +if (lean_is_scalar(x_270)) { + x_290 = lean_alloc_ctor(0, 2, 0); +} else { + x_290 = x_270; +} +lean_ctor_set(x_290, 0, x_268); +lean_ctor_set(x_290, 1, x_269); +x_291 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_291, 0, x_267); +lean_ctor_set(x_291, 1, x_290); +x_292 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_292, 0, x_291); +x_18 = x_292; +x_19 = x_289; +goto block_24; +} +else +{ +lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; uint8_t x_304; lean_object* x_305; +lean_dec(x_268); +x_293 = lean_ctor_get(x_286, 1); +lean_inc(x_293); +lean_dec(x_286); +x_294 = lean_box(0); +lean_inc(x_271); +x_295 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_267, x_271, x_294); +x_296 = lean_box(0); +lean_inc(x_2); +x_297 = lean_array_mk(x_2); +x_298 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; +x_299 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_299, 0, x_298); +lean_ctor_set(x_299, 1, x_296); +lean_ctor_set(x_299, 2, x_297); +x_300 = l_Lean_CollectFVars_main(x_274, x_299); +x_301 = lean_ctor_get(x_300, 2); +lean_inc(x_301); +lean_dec(x_300); +x_302 = lean_array_get_size(x_301); +x_303 = lean_unsigned_to_nat(0u); +x_304 = lean_nat_dec_lt(x_303, x_302); +x_305 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_269, x_271, x_294); +if (x_304 == 0) +{ +uint8_t x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; +lean_dec(x_302); +lean_dec(x_301); +x_306 = 1; +x_307 = lean_box(x_306); +if (lean_is_scalar(x_270)) { + x_308 = lean_alloc_ctor(0, 2, 0); +} else { + x_308 = x_270; +} +lean_ctor_set(x_308, 0, x_307); +lean_ctor_set(x_308, 1, x_305); +x_309 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_309, 0, x_295); +lean_ctor_set(x_309, 1, x_308); +x_310 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_310, 0, x_309); +x_18 = x_310; +x_19 = x_293; +goto block_24; +} +else +{ +uint8_t x_311; +x_311 = lean_nat_dec_le(x_302, x_302); +if (x_311 == 0) +{ +uint8_t x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; +lean_dec(x_302); +lean_dec(x_301); +x_312 = 1; +x_313 = lean_box(x_312); +if (lean_is_scalar(x_270)) { + x_314 = lean_alloc_ctor(0, 2, 0); +} else { + x_314 = x_270; +} +lean_ctor_set(x_314, 0, x_313); +lean_ctor_set(x_314, 1, x_305); +x_315 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_315, 0, x_295); +lean_ctor_set(x_315, 1, x_314); +x_316 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_316, 0, x_315); +x_18 = x_316; +x_19 = x_293; +goto block_24; +} +else +{ +size_t x_317; size_t x_318; lean_object* x_319; uint8_t x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; +x_317 = 0; +x_318 = lean_usize_of_nat(x_302); +lean_dec(x_302); +x_319 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__10___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__11(x_4, x_301, x_317, x_318, x_295); +lean_dec(x_301); +x_320 = 1; +x_321 = lean_box(x_320); +if (lean_is_scalar(x_270)) { + x_322 = lean_alloc_ctor(0, 2, 0); +} else { + x_322 = x_270; +} +lean_ctor_set(x_322, 0, x_321); +lean_ctor_set(x_322, 1, x_305); +x_323 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_323, 0, x_319); +lean_ctor_set(x_323, 1, x_322); +x_324 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_324, 0, x_323); +x_18 = x_324; +x_19 = x_293; +goto block_24; +} +} +} +} +else +{ +lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; +lean_dec(x_274); +lean_dec(x_271); +lean_dec(x_270); +lean_dec(x_269); +lean_dec(x_268); +lean_dec(x_267); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_4); +lean_dec(x_2); +x_325 = lean_ctor_get(x_286, 0); +lean_inc(x_325); +x_326 = lean_ctor_get(x_286, 1); +lean_inc(x_326); +if (lean_is_exclusive(x_286)) { + lean_ctor_release(x_286, 0); + lean_ctor_release(x_286, 1); + x_327 = x_286; +} else { + lean_dec_ref(x_286); + x_327 = lean_box(0); +} +if (lean_is_scalar(x_327)) { + x_328 = lean_alloc_ctor(1, 2, 0); +} else { + x_328 = x_327; +} +lean_ctor_set(x_328, 0, x_325); +lean_ctor_set(x_328, 1, x_326); +return x_328; +} +} +else +{ +lean_object* x_329; lean_object* x_330; lean_object* x_331; +lean_dec(x_285); +lean_dec(x_274); +lean_dec(x_271); +if (lean_is_scalar(x_270)) { + x_329 = lean_alloc_ctor(0, 2, 0); +} else { + x_329 = x_270; +} +lean_ctor_set(x_329, 0, x_268); +lean_ctor_set(x_329, 1, x_269); +x_330 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_330, 0, x_267); +lean_ctor_set(x_330, 1, x_329); +x_331 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_331, 0, x_330); +x_18 = x_331; +x_19 = x_283; +goto block_24; +} +} +else +{ +lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; +lean_dec(x_274); +lean_dec(x_271); +lean_dec(x_270); +lean_dec(x_269); +lean_dec(x_268); +lean_dec(x_267); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_4); +lean_dec(x_2); +x_332 = lean_ctor_get(x_281, 0); +lean_inc(x_332); +x_333 = lean_ctor_get(x_281, 1); +lean_inc(x_333); +if (lean_is_exclusive(x_281)) { + lean_ctor_release(x_281, 0); + lean_ctor_release(x_281, 1); + x_334 = x_281; +} else { + lean_dec_ref(x_281); + x_334 = lean_box(0); +} +if (lean_is_scalar(x_334)) { + x_335 = lean_alloc_ctor(1, 2, 0); +} else { + x_335 = x_334; +} +lean_ctor_set(x_335, 0, x_332); +lean_ctor_set(x_335, 1, x_333); +return x_335; +} +} +else +{ +lean_object* x_336; uint8_t x_337; +x_336 = lean_ctor_get(x_277, 1); +lean_inc(x_336); +lean_dec(x_277); +x_337 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars(x_4, x_267, x_274); +if (x_337 == 0) +{ +lean_object* x_338; lean_object* x_339; lean_object* x_340; +lean_dec(x_271); +if (lean_is_scalar(x_270)) { + x_338 = lean_alloc_ctor(0, 2, 0); +} else { + x_338 = x_270; +} +lean_ctor_set(x_338, 0, x_268); +lean_ctor_set(x_338, 1, x_269); +x_339 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_339, 0, x_267); +lean_ctor_set(x_339, 1, x_338); +x_340 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_340, 0, x_339); +x_18 = x_340; +x_19 = x_336; +goto block_24; +} +else +{ +lean_object* x_341; lean_object* x_342; lean_object* x_343; uint8_t x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; +lean_dec(x_268); +x_341 = lean_box(0); +lean_inc(x_271); +x_342 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_267, x_271, x_341); +x_343 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_269, x_271, x_341); +x_344 = 1; +x_345 = lean_box(x_344); +if (lean_is_scalar(x_270)) { + x_346 = lean_alloc_ctor(0, 2, 0); +} else { + x_346 = x_270; +} +lean_ctor_set(x_346, 0, x_345); +lean_ctor_set(x_346, 1, x_343); +x_347 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_347, 0, x_342); +lean_ctor_set(x_347, 1, x_346); +x_348 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_348, 0, x_347); +x_18 = x_348; +x_19 = x_336; +goto block_24; +} +} +} +else +{ +lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; +lean_dec(x_274); +lean_dec(x_271); +lean_dec(x_270); +lean_dec(x_269); +lean_dec(x_268); +lean_dec(x_267); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_4); +lean_dec(x_2); +x_349 = lean_ctor_get(x_277, 0); +lean_inc(x_349); +x_350 = lean_ctor_get(x_277, 1); +lean_inc(x_350); +if (lean_is_exclusive(x_277)) { + lean_ctor_release(x_277, 0); + lean_ctor_release(x_277, 1); + x_351 = x_277; +} else { + lean_dec_ref(x_277); + x_351 = lean_box(0); +} +if (lean_is_scalar(x_351)) { + x_352 = lean_alloc_ctor(1, 2, 0); +} else { + x_352 = x_351; +} +lean_ctor_set(x_352, 0, x_349); +lean_ctor_set(x_352, 1, x_350); +return x_352; +} +} +else +{ +lean_object* x_353; lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; uint8_t x_362; lean_object* x_363; lean_object* x_364; +lean_dec(x_268); +if (lean_is_exclusive(x_276)) { + lean_ctor_release(x_276, 0); + x_353 = x_276; +} else { + lean_dec_ref(x_276); + x_353 = lean_box(0); +} +x_354 = lean_box(0); +lean_inc(x_2); +x_355 = lean_array_mk(x_2); +x_356 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; +x_357 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_357, 0, x_356); +lean_ctor_set(x_357, 1, x_354); +lean_ctor_set(x_357, 2, x_355); +x_358 = l_Lean_CollectFVars_main(x_274, x_357); +x_359 = lean_ctor_get(x_358, 2); +lean_inc(x_359); +lean_dec(x_358); +x_360 = lean_array_get_size(x_359); +x_361 = lean_unsigned_to_nat(0u); +x_362 = lean_nat_dec_lt(x_361, x_360); +x_363 = lean_box(0); +x_364 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_269, x_271, x_363); +if (x_362 == 0) +{ +uint8_t x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; lean_object* x_369; +lean_dec(x_360); +lean_dec(x_359); +x_365 = 1; +x_366 = lean_box(x_365); +if (lean_is_scalar(x_270)) { + x_367 = lean_alloc_ctor(0, 2, 0); +} else { + x_367 = x_270; +} +lean_ctor_set(x_367, 0, x_366); +lean_ctor_set(x_367, 1, x_364); +x_368 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_368, 0, x_267); +lean_ctor_set(x_368, 1, x_367); +if (lean_is_scalar(x_353)) { + x_369 = lean_alloc_ctor(1, 1, 0); +} else { + x_369 = x_353; +} +lean_ctor_set(x_369, 0, x_368); +x_18 = x_369; +x_19 = x_275; +goto block_24; +} +else +{ +uint8_t x_370; +x_370 = lean_nat_dec_le(x_360, x_360); +if (x_370 == 0) +{ +uint8_t x_371; lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; +lean_dec(x_360); +lean_dec(x_359); +x_371 = 1; +x_372 = lean_box(x_371); +if (lean_is_scalar(x_270)) { + x_373 = lean_alloc_ctor(0, 2, 0); +} else { + x_373 = x_270; +} +lean_ctor_set(x_373, 0, x_372); +lean_ctor_set(x_373, 1, x_364); +x_374 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_374, 0, x_267); +lean_ctor_set(x_374, 1, x_373); +if (lean_is_scalar(x_353)) { + x_375 = lean_alloc_ctor(1, 1, 0); +} else { + x_375 = x_353; +} +lean_ctor_set(x_375, 0, x_374); +x_18 = x_375; +x_19 = x_275; +goto block_24; +} +else +{ +size_t x_376; size_t x_377; lean_object* x_378; uint8_t x_379; lean_object* x_380; lean_object* x_381; lean_object* x_382; lean_object* x_383; +x_376 = 0; +x_377 = lean_usize_of_nat(x_360); +lean_dec(x_360); +x_378 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__12___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__13(x_4, x_359, x_376, x_377, x_267); +lean_dec(x_359); +x_379 = 1; +x_380 = lean_box(x_379); +if (lean_is_scalar(x_270)) { + x_381 = lean_alloc_ctor(0, 2, 0); +} else { + x_381 = x_270; +} +lean_ctor_set(x_381, 0, x_380); +lean_ctor_set(x_381, 1, x_364); +x_382 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_382, 0, x_378); +lean_ctor_set(x_382, 1, x_381); +if (lean_is_scalar(x_353)) { + x_383 = lean_alloc_ctor(1, 1, 0); +} else { + x_383 = x_353; +} +lean_ctor_set(x_383, 0, x_382); +x_18 = x_383; +x_19 = x_275; +goto block_24; +} +} +} +} +else +{ +lean_object* x_384; lean_object* x_385; lean_object* x_386; lean_object* x_387; +lean_dec(x_271); +lean_dec(x_270); +lean_dec(x_269); +lean_dec(x_268); +lean_dec(x_267); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_4); +lean_dec(x_2); +x_384 = lean_ctor_get(x_273, 0); +lean_inc(x_384); +x_385 = lean_ctor_get(x_273, 1); +lean_inc(x_385); +if (lean_is_exclusive(x_273)) { + lean_ctor_release(x_273, 0); + lean_ctor_release(x_273, 1); + x_386 = x_273; +} else { + lean_dec_ref(x_273); + x_386 = lean_box(0); +} +if (lean_is_scalar(x_386)) { + x_387 = lean_alloc_ctor(1, 2, 0); +} else { + x_387 = x_386; +} +lean_ctor_set(x_387, 0, x_384); +lean_ctor_set(x_387, 1, x_385); +return x_387; +} +} +else +{ +lean_object* x_388; lean_object* x_389; lean_object* x_390; lean_object* x_391; +lean_dec(x_271); +lean_dec(x_17); +if (lean_is_exclusive(x_272)) { + lean_ctor_release(x_272, 0); + x_388 = x_272; +} else { + lean_dec_ref(x_272); + x_388 = lean_box(0); +} +if (lean_is_scalar(x_270)) { + x_389 = lean_alloc_ctor(0, 2, 0); +} else { + x_389 = x_270; +} +lean_ctor_set(x_389, 0, x_268); +lean_ctor_set(x_389, 1, x_269); +x_390 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_390, 0, x_267); +lean_ctor_set(x_390, 1, x_389); +if (lean_is_scalar(x_388)) { + x_391 = lean_alloc_ctor(1, 1, 0); +} else { + x_391 = x_388; +} +lean_ctor_set(x_391, 0, x_390); +x_18 = x_391; +x_19 = x_14; +goto block_24; +} +} +block_24: +{ +lean_object* x_20; size_t x_21; size_t x_22; +x_20 = lean_ctor_get(x_18, 0); +lean_inc(x_20); +lean_dec(x_18); +x_21 = 1; +x_22 = lean_usize_add(x_8, x_21); +x_8 = x_22; +x_9 = x_20; +x_14 = x_19; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__14___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__15(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, size_t x_6, size_t x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +uint8_t x_14; +x_14 = lean_usize_dec_lt(x_7, x_6); +if (x_14 == 0) +{ +lean_object* x_15; +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_3); +lean_dec(x_2); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_8); +lean_ctor_set(x_15, 1, x_13); +return x_15; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_24; +x_16 = lean_array_uget(x_5, x_7); +x_24 = !lean_is_exclusive(x_8); +if (x_24 == 0) +{ +lean_object* x_25; uint8_t x_26; +x_25 = lean_ctor_get(x_8, 1); +x_26 = !lean_is_exclusive(x_25); +if (x_26 == 0) +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_27 = lean_ctor_get(x_8, 0); +x_28 = lean_ctor_get(x_25, 0); +x_29 = lean_ctor_get(x_25, 1); +x_30 = l_Lean_Expr_fvarId_x21(x_16); +x_31 = l_Lean_RBNode_findCore___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__2(x_29, x_30); +if (lean_obj_tag(x_31) == 0) +{ +lean_object* x_32; +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +x_32 = lean_infer_type(x_16, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_32) == 0) +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_33 = lean_ctor_get(x_32, 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_RBNode_findCore___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__2(x_27, x_30); +if (lean_obj_tag(x_35) == 0) +{ +lean_object* x_36; +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_33); +x_36 = l_Lean_Meta_isProp(x_33, x_9, x_10, x_11, x_12, x_34); +if (lean_obj_tag(x_36) == 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_39; lean_object* x_40; +x_39 = lean_ctor_get(x_36, 1); +lean_inc(x_39); +lean_dec(x_36); +lean_inc(x_9); +lean_inc(x_30); +x_40 = l_Lean_FVarId_getDecl(x_30, x_9, x_10, x_11, x_12, x_39); +if (lean_obj_tag(x_40) == 0) +{ +lean_object* x_41; lean_object* x_42; uint8_t 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 = l_Lean_LocalDecl_binderInfo(x_41); +lean_dec(x_41); +x_44 = lean_box(x_43); +if (lean_obj_tag(x_44) == 3) +{ +lean_object* x_45; +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_33); +lean_inc(x_27); +lean_inc(x_3); +x_45 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized(x_3, x_27, x_33, x_9, x_10, x_11, x_12, x_42); +if (lean_obj_tag(x_45) == 0) +{ +lean_object* x_46; uint8_t x_47; +x_46 = lean_ctor_get(x_45, 0); +lean_inc(x_46); +x_47 = lean_unbox(x_46); +lean_dec(x_46); +if (x_47 == 0) +{ +lean_object* x_48; lean_object* x_49; +lean_dec(x_33); +lean_dec(x_30); +x_48 = lean_ctor_get(x_45, 1); +lean_inc(x_48); +lean_dec(x_45); +x_49 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_49, 0, x_8); +x_17 = x_49; +x_18 = x_48; +goto block_23; +} +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; uint8_t x_61; lean_object* x_62; +lean_dec(x_28); +x_50 = lean_ctor_get(x_45, 1); +lean_inc(x_50); +lean_dec(x_45); +x_51 = lean_box(0); +lean_inc(x_30); +x_52 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_27, x_30, x_51); +x_53 = lean_box(0); +lean_inc(x_2); +x_54 = lean_array_mk(x_2); +x_55 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; +x_56 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_56, 0, x_55); +lean_ctor_set(x_56, 1, x_53); +lean_ctor_set(x_56, 2, x_54); +x_57 = l_Lean_CollectFVars_main(x_33, x_56); +x_58 = lean_ctor_get(x_57, 2); +lean_inc(x_58); +lean_dec(x_57); +x_59 = lean_array_get_size(x_58); +x_60 = lean_unsigned_to_nat(0u); +x_61 = lean_nat_dec_lt(x_60, x_59); +x_62 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_29, x_30, x_51); +if (x_61 == 0) +{ +uint8_t x_63; lean_object* x_64; lean_object* x_65; +lean_dec(x_59); +lean_dec(x_58); +x_63 = 1; +x_64 = lean_box(x_63); +lean_ctor_set(x_25, 1, x_62); +lean_ctor_set(x_25, 0, x_64); +lean_ctor_set(x_8, 0, x_52); +x_65 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_65, 0, x_8); +x_17 = x_65; +x_18 = x_50; +goto block_23; +} +else +{ +uint8_t x_66; +x_66 = lean_nat_dec_le(x_59, x_59); +if (x_66 == 0) +{ +uint8_t x_67; lean_object* x_68; lean_object* x_69; +lean_dec(x_59); +lean_dec(x_58); +x_67 = 1; +x_68 = lean_box(x_67); +lean_ctor_set(x_25, 1, x_62); +lean_ctor_set(x_25, 0, x_68); +lean_ctor_set(x_8, 0, x_52); +x_69 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_69, 0, x_8); +x_17 = x_69; +x_18 = x_50; +goto block_23; +} +else +{ +size_t x_70; size_t x_71; lean_object* x_72; uint8_t x_73; lean_object* x_74; lean_object* x_75; +x_70 = 0; +x_71 = lean_usize_of_nat(x_59); +lean_dec(x_59); +x_72 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__10___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__11(x_3, x_58, x_70, x_71, x_52); +lean_dec(x_58); +x_73 = 1; +x_74 = lean_box(x_73); +lean_ctor_set(x_25, 1, x_62); +lean_ctor_set(x_25, 0, x_74); +lean_ctor_set(x_8, 0, x_72); +x_75 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_75, 0, x_8); +x_17 = x_75; +x_18 = x_50; +goto block_23; +} +} +} +} +else +{ +uint8_t x_76; +lean_dec(x_33); +lean_dec(x_30); +lean_free_object(x_25); +lean_dec(x_29); +lean_dec(x_28); +lean_free_object(x_8); +lean_dec(x_27); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_3); +lean_dec(x_2); +x_76 = !lean_is_exclusive(x_45); +if (x_76 == 0) +{ +return x_45; +} +else +{ +lean_object* x_77; lean_object* x_78; lean_object* x_79; +x_77 = lean_ctor_get(x_45, 0); +x_78 = lean_ctor_get(x_45, 1); +lean_inc(x_78); +lean_inc(x_77); +lean_dec(x_45); +x_79 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_79, 0, x_77); +lean_ctor_set(x_79, 1, x_78); +return x_79; +} +} +} +else +{ +lean_object* x_80; +lean_dec(x_44); +lean_dec(x_33); +lean_dec(x_30); +x_80 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_80, 0, x_8); +x_17 = x_80; +x_18 = x_42; +goto block_23; +} +} +else +{ +uint8_t x_81; +lean_dec(x_33); +lean_dec(x_30); +lean_free_object(x_25); +lean_dec(x_29); +lean_dec(x_28); +lean_free_object(x_8); +lean_dec(x_27); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_3); +lean_dec(x_2); +x_81 = !lean_is_exclusive(x_40); +if (x_81 == 0) +{ +return x_40; +} +else +{ +lean_object* x_82; lean_object* x_83; lean_object* x_84; +x_82 = lean_ctor_get(x_40, 0); +x_83 = lean_ctor_get(x_40, 1); +lean_inc(x_83); +lean_inc(x_82); +lean_dec(x_40); +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_85; uint8_t x_86; +x_85 = lean_ctor_get(x_36, 1); +lean_inc(x_85); +lean_dec(x_36); +x_86 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars(x_3, x_27, x_33); +if (x_86 == 0) +{ +lean_object* x_87; +lean_dec(x_30); +x_87 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_87, 0, x_8); +x_17 = x_87; +x_18 = x_85; +goto block_23; +} +else +{ +lean_object* x_88; lean_object* x_89; lean_object* x_90; uint8_t x_91; lean_object* x_92; lean_object* x_93; +lean_dec(x_28); +x_88 = lean_box(0); +lean_inc(x_30); +x_89 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_27, x_30, x_88); +x_90 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_29, x_30, x_88); +x_91 = 1; +x_92 = lean_box(x_91); +lean_ctor_set(x_25, 1, x_90); +lean_ctor_set(x_25, 0, x_92); +lean_ctor_set(x_8, 0, x_89); +x_93 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_93, 0, x_8); +x_17 = x_93; +x_18 = x_85; +goto block_23; +} +} +} +else +{ +uint8_t x_94; +lean_dec(x_33); +lean_dec(x_30); +lean_free_object(x_25); +lean_dec(x_29); +lean_dec(x_28); +lean_free_object(x_8); +lean_dec(x_27); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_3); +lean_dec(x_2); +x_94 = !lean_is_exclusive(x_36); +if (x_94 == 0) +{ +return x_36; +} +else +{ +lean_object* x_95; lean_object* x_96; lean_object* x_97; +x_95 = lean_ctor_get(x_36, 0); +x_96 = lean_ctor_get(x_36, 1); +lean_inc(x_96); +lean_inc(x_95); +lean_dec(x_36); +x_97 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_97, 0, x_95); +lean_ctor_set(x_97, 1, x_96); +return x_97; +} +} +} +else +{ +uint8_t x_98; +lean_dec(x_28); +x_98 = !lean_is_exclusive(x_35); +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; uint8_t x_108; lean_object* x_109; lean_object* x_110; +x_99 = lean_ctor_get(x_35, 0); +lean_dec(x_99); +x_100 = lean_box(0); +lean_inc(x_2); +x_101 = lean_array_mk(x_2); +x_102 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; +x_103 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_103, 0, x_102); +lean_ctor_set(x_103, 1, x_100); +lean_ctor_set(x_103, 2, x_101); +x_104 = l_Lean_CollectFVars_main(x_33, x_103); +x_105 = lean_ctor_get(x_104, 2); +lean_inc(x_105); +lean_dec(x_104); +x_106 = lean_array_get_size(x_105); +x_107 = lean_unsigned_to_nat(0u); +x_108 = lean_nat_dec_lt(x_107, x_106); +x_109 = lean_box(0); +x_110 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_29, x_30, x_109); +if (x_108 == 0) +{ +uint8_t x_111; lean_object* x_112; +lean_dec(x_106); +lean_dec(x_105); +x_111 = 1; +x_112 = lean_box(x_111); +lean_ctor_set(x_25, 1, x_110); +lean_ctor_set(x_25, 0, x_112); +lean_ctor_set(x_35, 0, x_8); +x_17 = x_35; +x_18 = x_34; +goto block_23; +} +else +{ +uint8_t x_113; +x_113 = lean_nat_dec_le(x_106, x_106); +if (x_113 == 0) +{ +uint8_t x_114; lean_object* x_115; +lean_dec(x_106); +lean_dec(x_105); +x_114 = 1; +x_115 = lean_box(x_114); +lean_ctor_set(x_25, 1, x_110); +lean_ctor_set(x_25, 0, x_115); +lean_ctor_set(x_35, 0, x_8); +x_17 = x_35; +x_18 = x_34; +goto block_23; +} +else +{ +size_t x_116; size_t x_117; lean_object* x_118; uint8_t x_119; lean_object* x_120; +x_116 = 0; +x_117 = lean_usize_of_nat(x_106); +lean_dec(x_106); +x_118 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__12___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__13(x_3, x_105, x_116, x_117, x_27); +lean_dec(x_105); +x_119 = 1; +x_120 = lean_box(x_119); +lean_ctor_set(x_25, 1, x_110); +lean_ctor_set(x_25, 0, x_120); +lean_ctor_set(x_8, 0, x_118); +lean_ctor_set(x_35, 0, x_8); +x_17 = x_35; +x_18 = x_34; +goto block_23; +} +} +} +else +{ +lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; uint8_t x_129; lean_object* x_130; lean_object* x_131; +lean_dec(x_35); +x_121 = lean_box(0); +lean_inc(x_2); +x_122 = lean_array_mk(x_2); +x_123 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; +x_124 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_124, 0, x_123); +lean_ctor_set(x_124, 1, x_121); +lean_ctor_set(x_124, 2, x_122); +x_125 = l_Lean_CollectFVars_main(x_33, x_124); +x_126 = lean_ctor_get(x_125, 2); +lean_inc(x_126); +lean_dec(x_125); +x_127 = lean_array_get_size(x_126); +x_128 = lean_unsigned_to_nat(0u); +x_129 = lean_nat_dec_lt(x_128, x_127); +x_130 = lean_box(0); +x_131 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_29, x_30, x_130); +if (x_129 == 0) +{ +uint8_t x_132; lean_object* x_133; lean_object* x_134; +lean_dec(x_127); +lean_dec(x_126); +x_132 = 1; +x_133 = lean_box(x_132); +lean_ctor_set(x_25, 1, x_131); +lean_ctor_set(x_25, 0, x_133); +x_134 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_134, 0, x_8); +x_17 = x_134; +x_18 = x_34; +goto block_23; +} +else +{ +uint8_t x_135; +x_135 = lean_nat_dec_le(x_127, x_127); +if (x_135 == 0) +{ +uint8_t x_136; lean_object* x_137; lean_object* x_138; +lean_dec(x_127); +lean_dec(x_126); +x_136 = 1; +x_137 = lean_box(x_136); +lean_ctor_set(x_25, 1, x_131); +lean_ctor_set(x_25, 0, x_137); +x_138 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_138, 0, x_8); +x_17 = x_138; +x_18 = x_34; +goto block_23; +} +else +{ +size_t x_139; size_t x_140; lean_object* x_141; uint8_t x_142; lean_object* x_143; lean_object* x_144; +x_139 = 0; +x_140 = lean_usize_of_nat(x_127); +lean_dec(x_127); +x_141 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__12___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__13(x_3, x_126, x_139, x_140, x_27); +lean_dec(x_126); +x_142 = 1; +x_143 = lean_box(x_142); +lean_ctor_set(x_25, 1, x_131); +lean_ctor_set(x_25, 0, x_143); +lean_ctor_set(x_8, 0, x_141); +x_144 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_144, 0, x_8); +x_17 = x_144; +x_18 = x_34; +goto block_23; +} +} +} +} +} +else +{ +uint8_t x_145; +lean_dec(x_30); +lean_free_object(x_25); +lean_dec(x_29); +lean_dec(x_28); +lean_free_object(x_8); +lean_dec(x_27); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_3); +lean_dec(x_2); +x_145 = !lean_is_exclusive(x_32); +if (x_145 == 0) +{ +return x_32; +} +else +{ +lean_object* x_146; lean_object* x_147; lean_object* x_148; +x_146 = lean_ctor_get(x_32, 0); +x_147 = lean_ctor_get(x_32, 1); +lean_inc(x_147); +lean_inc(x_146); +lean_dec(x_32); +x_148 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_148, 0, x_146); +lean_ctor_set(x_148, 1, x_147); +return x_148; +} +} +} +else +{ +uint8_t x_149; +lean_dec(x_30); +lean_dec(x_16); +x_149 = !lean_is_exclusive(x_31); +if (x_149 == 0) +{ +lean_object* x_150; +x_150 = lean_ctor_get(x_31, 0); +lean_dec(x_150); +lean_ctor_set(x_31, 0, x_8); +x_17 = x_31; +x_18 = x_13; +goto block_23; +} +else +{ +lean_object* x_151; +lean_dec(x_31); +x_151 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_151, 0, x_8); +x_17 = x_151; +x_18 = x_13; +goto block_23; +} +} +} +else +{ +lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; +x_152 = lean_ctor_get(x_8, 0); +x_153 = lean_ctor_get(x_25, 0); +x_154 = lean_ctor_get(x_25, 1); +lean_inc(x_154); +lean_inc(x_153); +lean_dec(x_25); +x_155 = l_Lean_Expr_fvarId_x21(x_16); +x_156 = l_Lean_RBNode_findCore___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__2(x_154, x_155); +if (lean_obj_tag(x_156) == 0) +{ +lean_object* x_157; +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +x_157 = lean_infer_type(x_16, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_157) == 0) +{ +lean_object* x_158; lean_object* x_159; lean_object* x_160; +x_158 = lean_ctor_get(x_157, 0); +lean_inc(x_158); +x_159 = lean_ctor_get(x_157, 1); +lean_inc(x_159); +lean_dec(x_157); +x_160 = l_Lean_RBNode_findCore___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__2(x_152, x_155); +if (lean_obj_tag(x_160) == 0) +{ +lean_object* x_161; +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_158); +x_161 = l_Lean_Meta_isProp(x_158, x_9, x_10, x_11, x_12, x_159); +if (lean_obj_tag(x_161) == 0) +{ +lean_object* x_162; uint8_t x_163; +x_162 = lean_ctor_get(x_161, 0); +lean_inc(x_162); +x_163 = lean_unbox(x_162); +lean_dec(x_162); +if (x_163 == 0) +{ +lean_object* x_164; lean_object* x_165; +x_164 = lean_ctor_get(x_161, 1); +lean_inc(x_164); +lean_dec(x_161); +lean_inc(x_9); +lean_inc(x_155); +x_165 = l_Lean_FVarId_getDecl(x_155, x_9, x_10, x_11, x_12, x_164); +if (lean_obj_tag(x_165) == 0) +{ +lean_object* x_166; lean_object* x_167; uint8_t x_168; lean_object* x_169; +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_Lean_LocalDecl_binderInfo(x_166); +lean_dec(x_166); +x_169 = lean_box(x_168); +if (lean_obj_tag(x_169) == 3) +{ +lean_object* x_170; +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_158); +lean_inc(x_152); +lean_inc(x_3); +x_170 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized(x_3, x_152, x_158, x_9, x_10, x_11, x_12, x_167); +if (lean_obj_tag(x_170) == 0) +{ +lean_object* x_171; uint8_t x_172; +x_171 = lean_ctor_get(x_170, 0); +lean_inc(x_171); +x_172 = lean_unbox(x_171); +lean_dec(x_171); +if (x_172 == 0) +{ +lean_object* x_173; lean_object* x_174; lean_object* x_175; +lean_dec(x_158); +lean_dec(x_155); +x_173 = lean_ctor_get(x_170, 1); +lean_inc(x_173); +lean_dec(x_170); +x_174 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_174, 0, x_153); +lean_ctor_set(x_174, 1, x_154); +lean_ctor_set(x_8, 1, x_174); +x_175 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_175, 0, x_8); +x_17 = x_175; +x_18 = x_173; +goto block_23; +} +else +{ +lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; uint8_t x_187; lean_object* x_188; +lean_dec(x_153); +x_176 = lean_ctor_get(x_170, 1); +lean_inc(x_176); +lean_dec(x_170); +x_177 = lean_box(0); +lean_inc(x_155); +x_178 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_152, x_155, x_177); +x_179 = lean_box(0); +lean_inc(x_2); +x_180 = lean_array_mk(x_2); +x_181 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; +x_182 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_182, 0, x_181); +lean_ctor_set(x_182, 1, x_179); +lean_ctor_set(x_182, 2, x_180); +x_183 = l_Lean_CollectFVars_main(x_158, x_182); +x_184 = lean_ctor_get(x_183, 2); +lean_inc(x_184); +lean_dec(x_183); +x_185 = lean_array_get_size(x_184); +x_186 = lean_unsigned_to_nat(0u); +x_187 = lean_nat_dec_lt(x_186, x_185); +x_188 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_154, x_155, x_177); +if (x_187 == 0) +{ +uint8_t x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; +lean_dec(x_185); +lean_dec(x_184); +x_189 = 1; +x_190 = lean_box(x_189); +x_191 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_191, 0, x_190); +lean_ctor_set(x_191, 1, x_188); +lean_ctor_set(x_8, 1, x_191); +lean_ctor_set(x_8, 0, x_178); +x_192 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_192, 0, x_8); +x_17 = x_192; +x_18 = x_176; +goto block_23; +} +else +{ +uint8_t x_193; +x_193 = lean_nat_dec_le(x_185, x_185); +if (x_193 == 0) +{ +uint8_t x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; +lean_dec(x_185); +lean_dec(x_184); +x_194 = 1; +x_195 = lean_box(x_194); +x_196 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_196, 0, x_195); +lean_ctor_set(x_196, 1, x_188); +lean_ctor_set(x_8, 1, x_196); +lean_ctor_set(x_8, 0, x_178); +x_197 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_197, 0, x_8); +x_17 = x_197; +x_18 = x_176; +goto block_23; +} +else +{ +size_t x_198; size_t x_199; lean_object* x_200; uint8_t x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; +x_198 = 0; +x_199 = lean_usize_of_nat(x_185); +lean_dec(x_185); +x_200 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__10___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__11(x_3, x_184, x_198, x_199, x_178); +lean_dec(x_184); +x_201 = 1; +x_202 = lean_box(x_201); +x_203 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_203, 0, x_202); +lean_ctor_set(x_203, 1, x_188); +lean_ctor_set(x_8, 1, x_203); +lean_ctor_set(x_8, 0, x_200); +x_204 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_204, 0, x_8); +x_17 = x_204; +x_18 = x_176; +goto block_23; +} +} +} +} +else +{ +lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; +lean_dec(x_158); +lean_dec(x_155); +lean_dec(x_154); +lean_dec(x_153); +lean_free_object(x_8); +lean_dec(x_152); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_3); +lean_dec(x_2); +x_205 = lean_ctor_get(x_170, 0); +lean_inc(x_205); +x_206 = lean_ctor_get(x_170, 1); +lean_inc(x_206); +if (lean_is_exclusive(x_170)) { + lean_ctor_release(x_170, 0); + lean_ctor_release(x_170, 1); + x_207 = x_170; +} else { + lean_dec_ref(x_170); + x_207 = lean_box(0); +} +if (lean_is_scalar(x_207)) { + x_208 = lean_alloc_ctor(1, 2, 0); +} else { + x_208 = x_207; +} +lean_ctor_set(x_208, 0, x_205); +lean_ctor_set(x_208, 1, x_206); +return x_208; +} +} +else +{ +lean_object* x_209; lean_object* x_210; +lean_dec(x_169); +lean_dec(x_158); +lean_dec(x_155); +x_209 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_209, 0, x_153); +lean_ctor_set(x_209, 1, x_154); +lean_ctor_set(x_8, 1, x_209); +x_210 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_210, 0, x_8); +x_17 = x_210; +x_18 = x_167; +goto block_23; +} +} +else +{ +lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; +lean_dec(x_158); +lean_dec(x_155); +lean_dec(x_154); +lean_dec(x_153); +lean_free_object(x_8); +lean_dec(x_152); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_3); +lean_dec(x_2); +x_211 = lean_ctor_get(x_165, 0); +lean_inc(x_211); +x_212 = lean_ctor_get(x_165, 1); +lean_inc(x_212); +if (lean_is_exclusive(x_165)) { + lean_ctor_release(x_165, 0); + lean_ctor_release(x_165, 1); + x_213 = x_165; +} else { + lean_dec_ref(x_165); + 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_215; uint8_t x_216; +x_215 = lean_ctor_get(x_161, 1); +lean_inc(x_215); +lean_dec(x_161); +x_216 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars(x_3, x_152, x_158); +if (x_216 == 0) +{ +lean_object* x_217; lean_object* x_218; +lean_dec(x_155); +x_217 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_217, 0, x_153); +lean_ctor_set(x_217, 1, x_154); +lean_ctor_set(x_8, 1, x_217); +x_218 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_218, 0, x_8); +x_17 = x_218; +x_18 = x_215; +goto block_23; +} +else +{ +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; +lean_dec(x_153); +x_219 = lean_box(0); +lean_inc(x_155); +x_220 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_152, x_155, x_219); +x_221 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_154, x_155, x_219); +x_222 = 1; +x_223 = lean_box(x_222); +x_224 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_224, 0, x_223); +lean_ctor_set(x_224, 1, x_221); +lean_ctor_set(x_8, 1, x_224); +lean_ctor_set(x_8, 0, x_220); +x_225 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_225, 0, x_8); +x_17 = x_225; +x_18 = x_215; +goto block_23; +} +} +} +else +{ +lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; +lean_dec(x_158); +lean_dec(x_155); +lean_dec(x_154); +lean_dec(x_153); +lean_free_object(x_8); +lean_dec(x_152); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_3); +lean_dec(x_2); +x_226 = lean_ctor_get(x_161, 0); +lean_inc(x_226); +x_227 = lean_ctor_get(x_161, 1); +lean_inc(x_227); +if (lean_is_exclusive(x_161)) { + lean_ctor_release(x_161, 0); + lean_ctor_release(x_161, 1); + x_228 = x_161; +} else { + lean_dec_ref(x_161); + x_228 = lean_box(0); +} +if (lean_is_scalar(x_228)) { + x_229 = lean_alloc_ctor(1, 2, 0); +} else { + x_229 = x_228; +} +lean_ctor_set(x_229, 0, x_226); +lean_ctor_set(x_229, 1, x_227); +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; uint8_t x_239; lean_object* x_240; lean_object* x_241; +lean_dec(x_153); +if (lean_is_exclusive(x_160)) { + lean_ctor_release(x_160, 0); + x_230 = x_160; +} else { + lean_dec_ref(x_160); + x_230 = lean_box(0); +} +x_231 = lean_box(0); +lean_inc(x_2); +x_232 = lean_array_mk(x_2); +x_233 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; +x_234 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_234, 0, x_233); +lean_ctor_set(x_234, 1, x_231); +lean_ctor_set(x_234, 2, x_232); +x_235 = l_Lean_CollectFVars_main(x_158, x_234); +x_236 = lean_ctor_get(x_235, 2); +lean_inc(x_236); +lean_dec(x_235); +x_237 = lean_array_get_size(x_236); +x_238 = lean_unsigned_to_nat(0u); +x_239 = lean_nat_dec_lt(x_238, x_237); +x_240 = lean_box(0); +x_241 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_154, x_155, x_240); +if (x_239 == 0) +{ +uint8_t x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; +lean_dec(x_237); +lean_dec(x_236); +x_242 = 1; +x_243 = lean_box(x_242); +x_244 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_244, 0, x_243); +lean_ctor_set(x_244, 1, x_241); +lean_ctor_set(x_8, 1, x_244); +if (lean_is_scalar(x_230)) { + x_245 = lean_alloc_ctor(1, 1, 0); +} else { + x_245 = x_230; +} +lean_ctor_set(x_245, 0, x_8); +x_17 = x_245; +x_18 = x_159; +goto block_23; +} +else +{ +uint8_t x_246; +x_246 = lean_nat_dec_le(x_237, x_237); +if (x_246 == 0) +{ +uint8_t x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; +lean_dec(x_237); +lean_dec(x_236); +x_247 = 1; +x_248 = lean_box(x_247); +x_249 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_249, 0, x_248); +lean_ctor_set(x_249, 1, x_241); +lean_ctor_set(x_8, 1, x_249); +if (lean_is_scalar(x_230)) { + x_250 = lean_alloc_ctor(1, 1, 0); +} else { + x_250 = x_230; +} +lean_ctor_set(x_250, 0, x_8); +x_17 = x_250; +x_18 = x_159; +goto block_23; +} +else +{ +size_t x_251; size_t x_252; lean_object* x_253; uint8_t x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; +x_251 = 0; +x_252 = lean_usize_of_nat(x_237); +lean_dec(x_237); +x_253 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__12___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__13(x_3, x_236, x_251, x_252, x_152); +lean_dec(x_236); +x_254 = 1; +x_255 = lean_box(x_254); +x_256 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_256, 0, x_255); +lean_ctor_set(x_256, 1, x_241); +lean_ctor_set(x_8, 1, x_256); +lean_ctor_set(x_8, 0, x_253); +if (lean_is_scalar(x_230)) { + x_257 = lean_alloc_ctor(1, 1, 0); +} else { + x_257 = x_230; +} +lean_ctor_set(x_257, 0, x_8); +x_17 = x_257; +x_18 = x_159; +goto block_23; +} +} +} +} +else +{ +lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; +lean_dec(x_155); +lean_dec(x_154); +lean_dec(x_153); +lean_free_object(x_8); +lean_dec(x_152); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_3); +lean_dec(x_2); +x_258 = lean_ctor_get(x_157, 0); +lean_inc(x_258); +x_259 = lean_ctor_get(x_157, 1); +lean_inc(x_259); +if (lean_is_exclusive(x_157)) { + lean_ctor_release(x_157, 0); + lean_ctor_release(x_157, 1); + x_260 = x_157; +} else { + lean_dec_ref(x_157); + x_260 = lean_box(0); +} +if (lean_is_scalar(x_260)) { + x_261 = lean_alloc_ctor(1, 2, 0); +} else { + x_261 = x_260; +} +lean_ctor_set(x_261, 0, x_258); +lean_ctor_set(x_261, 1, x_259); +return x_261; +} +} +else +{ +lean_object* x_262; lean_object* x_263; lean_object* x_264; +lean_dec(x_155); +lean_dec(x_16); +if (lean_is_exclusive(x_156)) { + lean_ctor_release(x_156, 0); + x_262 = x_156; +} else { + lean_dec_ref(x_156); + x_262 = lean_box(0); +} +x_263 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_263, 0, x_153); +lean_ctor_set(x_263, 1, x_154); +lean_ctor_set(x_8, 1, x_263); +if (lean_is_scalar(x_262)) { + x_264 = lean_alloc_ctor(1, 1, 0); +} else { + x_264 = x_262; +} +lean_ctor_set(x_264, 0, x_8); +x_17 = x_264; +x_18 = x_13; +goto block_23; +} +} +} +else +{ +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_265 = lean_ctor_get(x_8, 1); +x_266 = lean_ctor_get(x_8, 0); +lean_inc(x_265); +lean_inc(x_266); +lean_dec(x_8); +x_267 = lean_ctor_get(x_265, 0); +lean_inc(x_267); +x_268 = lean_ctor_get(x_265, 1); +lean_inc(x_268); +if (lean_is_exclusive(x_265)) { + lean_ctor_release(x_265, 0); + lean_ctor_release(x_265, 1); + x_269 = x_265; +} else { + lean_dec_ref(x_265); + x_269 = lean_box(0); +} +x_270 = l_Lean_Expr_fvarId_x21(x_16); +x_271 = l_Lean_RBNode_findCore___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__2(x_268, x_270); +if (lean_obj_tag(x_271) == 0) +{ +lean_object* x_272; +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +x_272 = lean_infer_type(x_16, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_272) == 0) +{ +lean_object* x_273; lean_object* x_274; lean_object* x_275; +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_275 = l_Lean_RBNode_findCore___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__2(x_266, x_270); +if (lean_obj_tag(x_275) == 0) +{ +lean_object* x_276; +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_273); +x_276 = l_Lean_Meta_isProp(x_273, x_9, x_10, x_11, x_12, x_274); +if (lean_obj_tag(x_276) == 0) +{ +lean_object* x_277; uint8_t x_278; +x_277 = lean_ctor_get(x_276, 0); +lean_inc(x_277); +x_278 = lean_unbox(x_277); +lean_dec(x_277); +if (x_278 == 0) +{ +lean_object* x_279; lean_object* x_280; +x_279 = lean_ctor_get(x_276, 1); +lean_inc(x_279); +lean_dec(x_276); +lean_inc(x_9); +lean_inc(x_270); +x_280 = l_Lean_FVarId_getDecl(x_270, x_9, x_10, x_11, x_12, x_279); +if (lean_obj_tag(x_280) == 0) +{ +lean_object* x_281; lean_object* x_282; uint8_t x_283; lean_object* x_284; +x_281 = lean_ctor_get(x_280, 0); +lean_inc(x_281); +x_282 = lean_ctor_get(x_280, 1); +lean_inc(x_282); +lean_dec(x_280); +x_283 = l_Lean_LocalDecl_binderInfo(x_281); +lean_dec(x_281); +x_284 = lean_box(x_283); +if (lean_obj_tag(x_284) == 3) +{ +lean_object* x_285; +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_273); +lean_inc(x_266); +lean_inc(x_3); +x_285 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized(x_3, x_266, x_273, x_9, x_10, x_11, x_12, x_282); +if (lean_obj_tag(x_285) == 0) +{ +lean_object* x_286; uint8_t x_287; +x_286 = lean_ctor_get(x_285, 0); +lean_inc(x_286); +x_287 = lean_unbox(x_286); +lean_dec(x_286); +if (x_287 == 0) +{ +lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; +lean_dec(x_273); +lean_dec(x_270); +x_288 = lean_ctor_get(x_285, 1); +lean_inc(x_288); +lean_dec(x_285); +if (lean_is_scalar(x_269)) { + x_289 = lean_alloc_ctor(0, 2, 0); +} else { + x_289 = x_269; +} +lean_ctor_set(x_289, 0, x_267); +lean_ctor_set(x_289, 1, x_268); +x_290 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_290, 0, x_266); +lean_ctor_set(x_290, 1, x_289); +x_291 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_291, 0, x_290); +x_17 = x_291; +x_18 = x_288; +goto block_23; +} +else +{ +lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; uint8_t x_303; lean_object* x_304; +lean_dec(x_267); +x_292 = lean_ctor_get(x_285, 1); +lean_inc(x_292); +lean_dec(x_285); +x_293 = lean_box(0); +lean_inc(x_270); +x_294 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_266, x_270, x_293); +x_295 = lean_box(0); +lean_inc(x_2); +x_296 = lean_array_mk(x_2); +x_297 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; +x_298 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_298, 0, x_297); +lean_ctor_set(x_298, 1, x_295); +lean_ctor_set(x_298, 2, x_296); +x_299 = l_Lean_CollectFVars_main(x_273, x_298); +x_300 = lean_ctor_get(x_299, 2); +lean_inc(x_300); +lean_dec(x_299); +x_301 = lean_array_get_size(x_300); +x_302 = lean_unsigned_to_nat(0u); +x_303 = lean_nat_dec_lt(x_302, x_301); +x_304 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_268, x_270, x_293); +if (x_303 == 0) +{ +uint8_t x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; +lean_dec(x_301); +lean_dec(x_300); +x_305 = 1; +x_306 = lean_box(x_305); +if (lean_is_scalar(x_269)) { + x_307 = lean_alloc_ctor(0, 2, 0); +} else { + x_307 = x_269; +} +lean_ctor_set(x_307, 0, x_306); +lean_ctor_set(x_307, 1, x_304); +x_308 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_308, 0, x_294); +lean_ctor_set(x_308, 1, x_307); +x_309 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_309, 0, x_308); +x_17 = x_309; +x_18 = x_292; +goto block_23; +} +else +{ +uint8_t x_310; +x_310 = lean_nat_dec_le(x_301, x_301); +if (x_310 == 0) +{ +uint8_t x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; +lean_dec(x_301); +lean_dec(x_300); +x_311 = 1; +x_312 = lean_box(x_311); +if (lean_is_scalar(x_269)) { + x_313 = lean_alloc_ctor(0, 2, 0); +} else { + x_313 = x_269; +} +lean_ctor_set(x_313, 0, x_312); +lean_ctor_set(x_313, 1, x_304); +x_314 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_314, 0, x_294); +lean_ctor_set(x_314, 1, x_313); +x_315 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_315, 0, x_314); +x_17 = x_315; +x_18 = x_292; +goto block_23; +} +else +{ +size_t x_316; size_t x_317; lean_object* x_318; uint8_t x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; +x_316 = 0; +x_317 = lean_usize_of_nat(x_301); +lean_dec(x_301); +x_318 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__10___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__11(x_3, x_300, x_316, x_317, x_294); +lean_dec(x_300); +x_319 = 1; +x_320 = lean_box(x_319); +if (lean_is_scalar(x_269)) { + x_321 = lean_alloc_ctor(0, 2, 0); +} else { + x_321 = x_269; +} +lean_ctor_set(x_321, 0, x_320); +lean_ctor_set(x_321, 1, x_304); +x_322 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_322, 0, x_318); +lean_ctor_set(x_322, 1, x_321); +x_323 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_323, 0, x_322); +x_17 = x_323; +x_18 = x_292; +goto block_23; +} +} +} +} +else +{ +lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; +lean_dec(x_273); +lean_dec(x_270); +lean_dec(x_269); +lean_dec(x_268); +lean_dec(x_267); +lean_dec(x_266); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_3); +lean_dec(x_2); +x_324 = lean_ctor_get(x_285, 0); +lean_inc(x_324); +x_325 = lean_ctor_get(x_285, 1); +lean_inc(x_325); +if (lean_is_exclusive(x_285)) { + lean_ctor_release(x_285, 0); + lean_ctor_release(x_285, 1); + x_326 = x_285; +} else { + lean_dec_ref(x_285); + x_326 = lean_box(0); +} +if (lean_is_scalar(x_326)) { + x_327 = lean_alloc_ctor(1, 2, 0); +} else { + x_327 = x_326; +} +lean_ctor_set(x_327, 0, x_324); +lean_ctor_set(x_327, 1, x_325); +return x_327; +} +} +else +{ +lean_object* x_328; lean_object* x_329; lean_object* x_330; +lean_dec(x_284); +lean_dec(x_273); +lean_dec(x_270); +if (lean_is_scalar(x_269)) { + x_328 = lean_alloc_ctor(0, 2, 0); +} else { + x_328 = x_269; +} +lean_ctor_set(x_328, 0, x_267); +lean_ctor_set(x_328, 1, x_268); +x_329 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_329, 0, x_266); +lean_ctor_set(x_329, 1, x_328); +x_330 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_330, 0, x_329); +x_17 = x_330; +x_18 = x_282; +goto block_23; +} +} +else +{ +lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; +lean_dec(x_273); +lean_dec(x_270); +lean_dec(x_269); +lean_dec(x_268); +lean_dec(x_267); +lean_dec(x_266); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_3); +lean_dec(x_2); +x_331 = lean_ctor_get(x_280, 0); +lean_inc(x_331); +x_332 = lean_ctor_get(x_280, 1); +lean_inc(x_332); +if (lean_is_exclusive(x_280)) { + lean_ctor_release(x_280, 0); + lean_ctor_release(x_280, 1); + x_333 = x_280; +} else { + lean_dec_ref(x_280); + x_333 = lean_box(0); +} +if (lean_is_scalar(x_333)) { + x_334 = lean_alloc_ctor(1, 2, 0); +} else { + x_334 = 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_335; uint8_t x_336; +x_335 = lean_ctor_get(x_276, 1); +lean_inc(x_335); +lean_dec(x_276); +x_336 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars(x_3, x_266, x_273); +if (x_336 == 0) +{ +lean_object* x_337; lean_object* x_338; lean_object* x_339; +lean_dec(x_270); +if (lean_is_scalar(x_269)) { + x_337 = lean_alloc_ctor(0, 2, 0); +} else { + x_337 = x_269; +} +lean_ctor_set(x_337, 0, x_267); +lean_ctor_set(x_337, 1, x_268); +x_338 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_338, 0, x_266); +lean_ctor_set(x_338, 1, x_337); +x_339 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_339, 0, x_338); +x_17 = x_339; +x_18 = x_335; +goto block_23; +} +else +{ +lean_object* x_340; lean_object* x_341; lean_object* x_342; uint8_t x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; +lean_dec(x_267); +x_340 = lean_box(0); +lean_inc(x_270); +x_341 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_266, x_270, x_340); +x_342 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_268, x_270, x_340); +x_343 = 1; +x_344 = lean_box(x_343); +if (lean_is_scalar(x_269)) { + x_345 = lean_alloc_ctor(0, 2, 0); +} else { + x_345 = x_269; +} +lean_ctor_set(x_345, 0, x_344); +lean_ctor_set(x_345, 1, x_342); +x_346 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_346, 0, x_341); +lean_ctor_set(x_346, 1, x_345); +x_347 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_347, 0, x_346); +x_17 = x_347; +x_18 = x_335; +goto block_23; +} +} +} +else +{ +lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; +lean_dec(x_273); +lean_dec(x_270); +lean_dec(x_269); +lean_dec(x_268); +lean_dec(x_267); +lean_dec(x_266); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_3); +lean_dec(x_2); +x_348 = lean_ctor_get(x_276, 0); +lean_inc(x_348); +x_349 = lean_ctor_get(x_276, 1); +lean_inc(x_349); +if (lean_is_exclusive(x_276)) { + lean_ctor_release(x_276, 0); + lean_ctor_release(x_276, 1); + x_350 = x_276; +} else { + lean_dec_ref(x_276); + x_350 = lean_box(0); +} +if (lean_is_scalar(x_350)) { + x_351 = lean_alloc_ctor(1, 2, 0); +} else { + x_351 = x_350; +} +lean_ctor_set(x_351, 0, x_348); +lean_ctor_set(x_351, 1, x_349); +return x_351; +} +} +else +{ +lean_object* x_352; lean_object* x_353; lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; uint8_t x_361; lean_object* x_362; lean_object* x_363; +lean_dec(x_267); +if (lean_is_exclusive(x_275)) { + lean_ctor_release(x_275, 0); + x_352 = x_275; +} else { + lean_dec_ref(x_275); + x_352 = lean_box(0); +} +x_353 = lean_box(0); +lean_inc(x_2); +x_354 = lean_array_mk(x_2); +x_355 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; +x_356 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_356, 0, x_355); +lean_ctor_set(x_356, 1, x_353); +lean_ctor_set(x_356, 2, x_354); +x_357 = l_Lean_CollectFVars_main(x_273, x_356); +x_358 = lean_ctor_get(x_357, 2); +lean_inc(x_358); +lean_dec(x_357); +x_359 = lean_array_get_size(x_358); +x_360 = lean_unsigned_to_nat(0u); +x_361 = lean_nat_dec_lt(x_360, x_359); +x_362 = lean_box(0); +x_363 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_268, x_270, x_362); +if (x_361 == 0) +{ +uint8_t x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; +lean_dec(x_359); +lean_dec(x_358); +x_364 = 1; +x_365 = lean_box(x_364); +if (lean_is_scalar(x_269)) { + x_366 = lean_alloc_ctor(0, 2, 0); +} else { + x_366 = x_269; +} +lean_ctor_set(x_366, 0, x_365); +lean_ctor_set(x_366, 1, x_363); +x_367 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_367, 0, x_266); +lean_ctor_set(x_367, 1, x_366); +if (lean_is_scalar(x_352)) { + x_368 = lean_alloc_ctor(1, 1, 0); +} else { + x_368 = x_352; +} +lean_ctor_set(x_368, 0, x_367); +x_17 = x_368; +x_18 = x_274; +goto block_23; +} +else +{ +uint8_t x_369; +x_369 = lean_nat_dec_le(x_359, x_359); +if (x_369 == 0) +{ +uint8_t x_370; lean_object* x_371; lean_object* x_372; lean_object* x_373; lean_object* x_374; +lean_dec(x_359); +lean_dec(x_358); +x_370 = 1; +x_371 = lean_box(x_370); +if (lean_is_scalar(x_269)) { + x_372 = lean_alloc_ctor(0, 2, 0); +} else { + x_372 = x_269; +} +lean_ctor_set(x_372, 0, x_371); +lean_ctor_set(x_372, 1, x_363); +x_373 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_373, 0, x_266); +lean_ctor_set(x_373, 1, x_372); +if (lean_is_scalar(x_352)) { + x_374 = lean_alloc_ctor(1, 1, 0); +} else { + x_374 = x_352; +} +lean_ctor_set(x_374, 0, x_373); +x_17 = x_374; +x_18 = x_274; +goto block_23; +} +else +{ +size_t x_375; size_t x_376; lean_object* x_377; uint8_t x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; lean_object* x_382; +x_375 = 0; +x_376 = lean_usize_of_nat(x_359); +lean_dec(x_359); +x_377 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__12___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__13(x_3, x_358, x_375, x_376, x_266); +lean_dec(x_358); +x_378 = 1; +x_379 = lean_box(x_378); +if (lean_is_scalar(x_269)) { + x_380 = lean_alloc_ctor(0, 2, 0); +} else { + x_380 = x_269; +} +lean_ctor_set(x_380, 0, x_379); +lean_ctor_set(x_380, 1, x_363); +x_381 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_381, 0, x_377); +lean_ctor_set(x_381, 1, x_380); +if (lean_is_scalar(x_352)) { + x_382 = lean_alloc_ctor(1, 1, 0); +} else { + x_382 = x_352; +} +lean_ctor_set(x_382, 0, x_381); +x_17 = x_382; +x_18 = x_274; +goto block_23; +} +} +} +} +else +{ +lean_object* x_383; lean_object* x_384; lean_object* x_385; lean_object* x_386; +lean_dec(x_270); +lean_dec(x_269); +lean_dec(x_268); +lean_dec(x_267); +lean_dec(x_266); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_3); +lean_dec(x_2); +x_383 = lean_ctor_get(x_272, 0); +lean_inc(x_383); +x_384 = lean_ctor_get(x_272, 1); +lean_inc(x_384); +if (lean_is_exclusive(x_272)) { + lean_ctor_release(x_272, 0); + lean_ctor_release(x_272, 1); + x_385 = x_272; +} else { + lean_dec_ref(x_272); + x_385 = lean_box(0); +} +if (lean_is_scalar(x_385)) { + x_386 = lean_alloc_ctor(1, 2, 0); +} else { + x_386 = x_385; +} +lean_ctor_set(x_386, 0, x_383); +lean_ctor_set(x_386, 1, x_384); +return x_386; +} +} +else +{ +lean_object* x_387; lean_object* x_388; lean_object* x_389; lean_object* x_390; +lean_dec(x_270); +lean_dec(x_16); +if (lean_is_exclusive(x_271)) { + lean_ctor_release(x_271, 0); + x_387 = x_271; +} else { + lean_dec_ref(x_271); + x_387 = lean_box(0); +} +if (lean_is_scalar(x_269)) { + x_388 = lean_alloc_ctor(0, 2, 0); +} else { + x_388 = x_269; +} +lean_ctor_set(x_388, 0, x_267); +lean_ctor_set(x_388, 1, x_268); +x_389 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_389, 0, x_266); +lean_ctor_set(x_389, 1, x_388); +if (lean_is_scalar(x_387)) { + x_390 = lean_alloc_ctor(1, 1, 0); +} else { + x_390 = x_387; +} +lean_ctor_set(x_390, 0, x_389); +x_17 = x_390; +x_18 = x_13; +goto block_23; +} +} +block_23: +{ +lean_object* x_19; size_t x_20; size_t x_21; +x_19 = lean_ctor_get(x_17, 0); +lean_inc(x_19); +lean_dec(x_17); +x_20 = 1; +x_21 = lean_usize_add(x_7, x_20); +x_7 = x_21; +x_8 = x_19; +x_13 = x_18; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___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: +{ +if (x_1 == 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_box(x_1); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_2); +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_3); +lean_ctor_set(x_13, 1, x_12); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_4); +lean_ctor_set(x_14, 1, x_13); +x_15 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_15, 0, x_14); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_15); +lean_ctor_set(x_16, 1, x_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; +x_17 = lean_box(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_2); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_3); +lean_ctor_set(x_19, 1, x_18); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_4); +lean_ctor_set(x_20, 1, x_19); +x_21 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_21, 0, x_20); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_10); +return x_22; +} +} +} +static lean_object* _init_l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___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_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___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, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; uint8_t x_14; +x_13 = lean_ctor_get(x_7, 1); +lean_inc(x_13); +lean_dec(x_7); +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) +{ +lean_object* x_15; uint8_t x_16; +x_15 = lean_ctor_get(x_13, 1); +x_16 = !lean_is_exclusive(x_15); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; uint8_t x_19; lean_object* x_20; size_t x_21; size_t x_22; lean_object* x_23; +x_17 = lean_ctor_get(x_15, 0); +lean_dec(x_17); +x_18 = lean_box(0); +x_19 = 0; +x_20 = lean_box(x_19); +lean_ctor_set(x_15, 0, x_20); +x_21 = lean_array_size(x_2); +x_22 = 0; +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_5); +lean_inc(x_3); +x_23 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__14___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__15(x_2, x_3, x_5, x_18, x_2, x_21, x_22, x_13, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_23) == 0) +{ +lean_object* x_24; lean_object* x_25; uint8_t x_26; +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_is_exclusive(x_23); +if (x_26 == 0) +{ +lean_object* x_27; lean_object* x_28; uint8_t x_29; +x_27 = lean_ctor_get(x_23, 1); +x_28 = lean_ctor_get(x_23, 0); +lean_dec(x_28); +x_29 = !lean_is_exclusive(x_24); +if (x_29 == 0) +{ +lean_object* x_30; lean_object* x_31; uint8_t x_32; +x_30 = lean_ctor_get(x_24, 0); +x_31 = lean_ctor_get(x_24, 1); +lean_dec(x_31); +x_32 = !lean_is_exclusive(x_25); +if (x_32 == 0) +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; +x_33 = lean_ctor_get(x_25, 0); +x_34 = lean_ctor_get(x_25, 1); +x_35 = lean_unsigned_to_nat(0u); +x_36 = l_Lean_RBNode_fold___at_Lean_RBMap_size___spec__1___rarg(x_35, x_30); +x_37 = lean_nat_dec_eq(x_36, x_1); +lean_dec(x_36); +if (x_37 == 0) +{ +lean_object* x_38; uint8_t x_39; lean_object* x_40; lean_object* x_41; +lean_free_object(x_25); +lean_free_object(x_24); +lean_free_object(x_23); +x_38 = lean_box(0); +x_39 = lean_unbox(x_33); +lean_dec(x_33); +lean_inc(x_6); +x_40 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___lambda__1(x_39, x_34, x_30, x_6, x_38, x_8, x_9, x_10, x_11, x_27); +x_41 = lean_ctor_get(x_40, 0); +lean_inc(x_41); +if (lean_obj_tag(x_41) == 0) +{ +uint8_t x_42; +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_3); +x_42 = !lean_is_exclusive(x_40); +if (x_42 == 0) +{ +lean_object* x_43; lean_object* x_44; +x_43 = lean_ctor_get(x_40, 0); +lean_dec(x_43); +x_44 = lean_ctor_get(x_41, 0); +lean_inc(x_44); +lean_dec(x_41); +lean_ctor_set(x_40, 0, x_44); +return x_40; +} +else +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_45 = lean_ctor_get(x_40, 1); +lean_inc(x_45); +lean_dec(x_40); +x_46 = lean_ctor_get(x_41, 0); +lean_inc(x_46); +lean_dec(x_41); +x_47 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_47, 0, x_46); +lean_ctor_set(x_47, 1, x_45); +return x_47; +} +} +else +{ +lean_object* x_48; lean_object* x_49; +x_48 = lean_ctor_get(x_40, 1); +lean_inc(x_48); +lean_dec(x_40); +x_49 = lean_ctor_get(x_41, 0); +lean_inc(x_49); +lean_dec(x_41); +x_7 = x_49; +x_12 = x_48; +goto _start; +} +} +else +{ +lean_object* x_51; lean_object* x_52; +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_3); +x_51 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___closed__1; +x_52 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_52, 0, x_51); +lean_ctor_set(x_52, 1, x_24); +lean_ctor_set(x_23, 0, x_52); +return x_23; +} +} +else +{ +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_25, 0); +x_54 = lean_ctor_get(x_25, 1); +lean_inc(x_54); +lean_inc(x_53); +lean_dec(x_25); +x_55 = lean_unsigned_to_nat(0u); +x_56 = l_Lean_RBNode_fold___at_Lean_RBMap_size___spec__1___rarg(x_55, x_30); +x_57 = lean_nat_dec_eq(x_56, x_1); +lean_dec(x_56); +if (x_57 == 0) +{ +lean_object* x_58; uint8_t x_59; lean_object* x_60; lean_object* x_61; +lean_free_object(x_24); +lean_free_object(x_23); +x_58 = lean_box(0); +x_59 = lean_unbox(x_53); +lean_dec(x_53); +lean_inc(x_6); +x_60 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___lambda__1(x_59, x_54, x_30, x_6, x_58, x_8, x_9, x_10, x_11, x_27); +x_61 = lean_ctor_get(x_60, 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; +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_3); +x_62 = lean_ctor_get(x_60, 1); +lean_inc(x_62); +if (lean_is_exclusive(x_60)) { + lean_ctor_release(x_60, 0); + lean_ctor_release(x_60, 1); + x_63 = x_60; +} else { + lean_dec_ref(x_60); + x_63 = lean_box(0); +} +x_64 = lean_ctor_get(x_61, 0); +lean_inc(x_64); +lean_dec(x_61); +if (lean_is_scalar(x_63)) { + x_65 = lean_alloc_ctor(0, 2, 0); +} else { + x_65 = x_63; +} +lean_ctor_set(x_65, 0, x_64); +lean_ctor_set(x_65, 1, x_62); +return x_65; +} +else +{ +lean_object* x_66; lean_object* x_67; +x_66 = lean_ctor_get(x_60, 1); +lean_inc(x_66); +lean_dec(x_60); +x_67 = lean_ctor_get(x_61, 0); +lean_inc(x_67); +lean_dec(x_61); +x_7 = x_67; +x_12 = x_66; +goto _start; +} +} +else +{ +lean_object* x_69; lean_object* x_70; lean_object* x_71; +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_3); +x_69 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_69, 0, x_53); +lean_ctor_set(x_69, 1, x_54); +lean_ctor_set(x_24, 1, x_69); +x_70 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___closed__1; +x_71 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_71, 0, x_70); +lean_ctor_set(x_71, 1, x_24); +lean_ctor_set(x_23, 0, x_71); +return x_23; +} +} +} +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; uint8_t x_78; +x_72 = lean_ctor_get(x_24, 0); +lean_inc(x_72); +lean_dec(x_24); +x_73 = lean_ctor_get(x_25, 0); +lean_inc(x_73); +x_74 = lean_ctor_get(x_25, 1); +lean_inc(x_74); +if (lean_is_exclusive(x_25)) { + lean_ctor_release(x_25, 0); + lean_ctor_release(x_25, 1); + x_75 = x_25; +} else { + lean_dec_ref(x_25); + x_75 = lean_box(0); +} +x_76 = lean_unsigned_to_nat(0u); +x_77 = l_Lean_RBNode_fold___at_Lean_RBMap_size___spec__1___rarg(x_76, x_72); +x_78 = lean_nat_dec_eq(x_77, x_1); +lean_dec(x_77); +if (x_78 == 0) +{ +lean_object* x_79; uint8_t x_80; lean_object* x_81; lean_object* x_82; +lean_dec(x_75); +lean_free_object(x_23); +x_79 = lean_box(0); +x_80 = lean_unbox(x_73); +lean_dec(x_73); +lean_inc(x_6); +x_81 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___lambda__1(x_80, x_74, x_72, x_6, x_79, x_8, x_9, x_10, x_11, x_27); +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; lean_object* x_85; lean_object* x_86; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +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; +} else { + lean_dec_ref(x_81); + x_84 = lean_box(0); +} +x_85 = lean_ctor_get(x_82, 0); +lean_inc(x_85); +lean_dec(x_82); +if (lean_is_scalar(x_84)) { + x_86 = lean_alloc_ctor(0, 2, 0); +} else { + x_86 = x_84; +} +lean_ctor_set(x_86, 0, x_85); +lean_ctor_set(x_86, 1, x_83); +return x_86; +} +else +{ +lean_object* x_87; lean_object* x_88; +x_87 = lean_ctor_get(x_81, 1); +lean_inc(x_87); +lean_dec(x_81); +x_88 = lean_ctor_get(x_82, 0); +lean_inc(x_88); +lean_dec(x_82); +x_7 = x_88; +x_12 = x_87; +goto _start; +} +} +else +{ +lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; +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_3); +if (lean_is_scalar(x_75)) { + x_90 = lean_alloc_ctor(0, 2, 0); +} else { + x_90 = x_75; +} +lean_ctor_set(x_90, 0, x_73); +lean_ctor_set(x_90, 1, x_74); +x_91 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_91, 0, x_72); +lean_ctor_set(x_91, 1, x_90); +x_92 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___closed__1; +x_93 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_93, 0, x_92); +lean_ctor_set(x_93, 1, x_91); +lean_ctor_set(x_23, 0, x_93); +return x_23; +} +} +} +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; uint8_t x_102; +x_94 = lean_ctor_get(x_23, 1); +lean_inc(x_94); +lean_dec(x_23); +x_95 = lean_ctor_get(x_24, 0); +lean_inc(x_95); +if (lean_is_exclusive(x_24)) { + lean_ctor_release(x_24, 0); + lean_ctor_release(x_24, 1); + x_96 = x_24; +} else { + lean_dec_ref(x_24); + x_96 = lean_box(0); +} +x_97 = lean_ctor_get(x_25, 0); +lean_inc(x_97); +x_98 = lean_ctor_get(x_25, 1); +lean_inc(x_98); +if (lean_is_exclusive(x_25)) { + lean_ctor_release(x_25, 0); + lean_ctor_release(x_25, 1); + x_99 = x_25; +} else { + lean_dec_ref(x_25); + x_99 = lean_box(0); +} +x_100 = lean_unsigned_to_nat(0u); +x_101 = l_Lean_RBNode_fold___at_Lean_RBMap_size___spec__1___rarg(x_100, x_95); +x_102 = lean_nat_dec_eq(x_101, x_1); +lean_dec(x_101); +if (x_102 == 0) +{ +lean_object* x_103; uint8_t x_104; lean_object* x_105; lean_object* x_106; +lean_dec(x_99); +lean_dec(x_96); +x_103 = lean_box(0); +x_104 = lean_unbox(x_97); +lean_dec(x_97); +lean_inc(x_6); +x_105 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___lambda__1(x_104, x_98, x_95, x_6, x_103, x_8, x_9, x_10, x_11, x_94); +x_106 = lean_ctor_get(x_105, 0); +lean_inc(x_106); +if (lean_obj_tag(x_106) == 0) +{ +lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; +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_3); +x_107 = lean_ctor_get(x_105, 1); +lean_inc(x_107); +if (lean_is_exclusive(x_105)) { + lean_ctor_release(x_105, 0); + lean_ctor_release(x_105, 1); + x_108 = x_105; +} else { + lean_dec_ref(x_105); + x_108 = lean_box(0); +} +x_109 = lean_ctor_get(x_106, 0); +lean_inc(x_109); +lean_dec(x_106); +if (lean_is_scalar(x_108)) { + x_110 = lean_alloc_ctor(0, 2, 0); +} else { + x_110 = x_108; +} +lean_ctor_set(x_110, 0, x_109); +lean_ctor_set(x_110, 1, x_107); +return x_110; +} +else +{ +lean_object* x_111; lean_object* x_112; +x_111 = lean_ctor_get(x_105, 1); +lean_inc(x_111); +lean_dec(x_105); +x_112 = lean_ctor_get(x_106, 0); +lean_inc(x_112); +lean_dec(x_106); +x_7 = x_112; +x_12 = x_111; +goto _start; +} +} +else +{ +lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; +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_3); +if (lean_is_scalar(x_99)) { + x_114 = lean_alloc_ctor(0, 2, 0); +} else { + x_114 = x_99; +} +lean_ctor_set(x_114, 0, x_97); +lean_ctor_set(x_114, 1, x_98); +if (lean_is_scalar(x_96)) { + x_115 = lean_alloc_ctor(0, 2, 0); +} else { + x_115 = x_96; +} +lean_ctor_set(x_115, 0, x_95); +lean_ctor_set(x_115, 1, x_114); +x_116 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___closed__1; +x_117 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_117, 0, x_116); +lean_ctor_set(x_117, 1, x_115); +x_118 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_118, 0, x_117); +lean_ctor_set(x_118, 1, x_94); +return x_118; +} +} +} +else +{ +uint8_t x_119; +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_3); +x_119 = !lean_is_exclusive(x_23); +if (x_119 == 0) +{ +return x_23; +} +else +{ +lean_object* x_120; lean_object* x_121; lean_object* x_122; +x_120 = lean_ctor_get(x_23, 0); +x_121 = lean_ctor_get(x_23, 1); +lean_inc(x_121); +lean_inc(x_120); +lean_dec(x_23); +x_122 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_122, 0, x_120); +lean_ctor_set(x_122, 1, x_121); +return x_122; +} +} +} +else +{ +lean_object* x_123; lean_object* x_124; uint8_t x_125; lean_object* x_126; lean_object* x_127; size_t x_128; size_t x_129; lean_object* x_130; +x_123 = lean_ctor_get(x_15, 1); +lean_inc(x_123); +lean_dec(x_15); +x_124 = lean_box(0); +x_125 = 0; +x_126 = lean_box(x_125); +x_127 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_127, 0, x_126); +lean_ctor_set(x_127, 1, x_123); +lean_ctor_set(x_13, 1, x_127); +x_128 = lean_array_size(x_2); +x_129 = 0; +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_5); +lean_inc(x_3); +x_130 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__14___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__15(x_2, x_3, x_5, x_124, x_2, x_128, x_129, x_13, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_130) == 0) +{ +lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; uint8_t x_142; +x_131 = lean_ctor_get(x_130, 0); +lean_inc(x_131); +x_132 = lean_ctor_get(x_131, 1); +lean_inc(x_132); +x_133 = lean_ctor_get(x_130, 1); +lean_inc(x_133); +if (lean_is_exclusive(x_130)) { + lean_ctor_release(x_130, 0); + lean_ctor_release(x_130, 1); + x_134 = x_130; +} else { + lean_dec_ref(x_130); + x_134 = lean_box(0); +} +x_135 = lean_ctor_get(x_131, 0); +lean_inc(x_135); +if (lean_is_exclusive(x_131)) { + lean_ctor_release(x_131, 0); + lean_ctor_release(x_131, 1); + x_136 = x_131; +} else { + lean_dec_ref(x_131); + x_136 = lean_box(0); +} +x_137 = lean_ctor_get(x_132, 0); +lean_inc(x_137); +x_138 = lean_ctor_get(x_132, 1); +lean_inc(x_138); +if (lean_is_exclusive(x_132)) { + lean_ctor_release(x_132, 0); + lean_ctor_release(x_132, 1); + x_139 = x_132; +} else { + lean_dec_ref(x_132); + x_139 = lean_box(0); +} +x_140 = lean_unsigned_to_nat(0u); +x_141 = l_Lean_RBNode_fold___at_Lean_RBMap_size___spec__1___rarg(x_140, x_135); +x_142 = lean_nat_dec_eq(x_141, x_1); +lean_dec(x_141); +if (x_142 == 0) +{ +lean_object* x_143; uint8_t x_144; lean_object* x_145; lean_object* x_146; +lean_dec(x_139); +lean_dec(x_136); +lean_dec(x_134); +x_143 = lean_box(0); +x_144 = lean_unbox(x_137); +lean_dec(x_137); +lean_inc(x_6); +x_145 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___lambda__1(x_144, x_138, x_135, x_6, x_143, x_8, x_9, x_10, x_11, x_133); +x_146 = lean_ctor_get(x_145, 0); +lean_inc(x_146); +if (lean_obj_tag(x_146) == 0) +{ +lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; +lean_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_3); +x_147 = lean_ctor_get(x_145, 1); +lean_inc(x_147); +if (lean_is_exclusive(x_145)) { + lean_ctor_release(x_145, 0); + lean_ctor_release(x_145, 1); + x_148 = x_145; +} else { + lean_dec_ref(x_145); + x_148 = lean_box(0); +} +x_149 = lean_ctor_get(x_146, 0); +lean_inc(x_149); +lean_dec(x_146); +if (lean_is_scalar(x_148)) { + x_150 = lean_alloc_ctor(0, 2, 0); +} else { + x_150 = x_148; +} +lean_ctor_set(x_150, 0, x_149); +lean_ctor_set(x_150, 1, x_147); +return x_150; +} +else +{ +lean_object* x_151; lean_object* x_152; +x_151 = lean_ctor_get(x_145, 1); +lean_inc(x_151); +lean_dec(x_145); +x_152 = lean_ctor_get(x_146, 0); +lean_inc(x_152); +lean_dec(x_146); +x_7 = x_152; +x_12 = x_151; +goto _start; +} +} +else +{ +lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; +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_3); +if (lean_is_scalar(x_139)) { + x_154 = lean_alloc_ctor(0, 2, 0); +} else { + x_154 = x_139; +} +lean_ctor_set(x_154, 0, x_137); +lean_ctor_set(x_154, 1, x_138); +if (lean_is_scalar(x_136)) { + x_155 = lean_alloc_ctor(0, 2, 0); +} else { + x_155 = x_136; +} +lean_ctor_set(x_155, 0, x_135); +lean_ctor_set(x_155, 1, x_154); +x_156 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___closed__1; +x_157 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_157, 0, x_156); +lean_ctor_set(x_157, 1, x_155); +if (lean_is_scalar(x_134)) { + x_158 = lean_alloc_ctor(0, 2, 0); +} else { + x_158 = x_134; +} +lean_ctor_set(x_158, 0, x_157); +lean_ctor_set(x_158, 1, x_133); +return x_158; +} +} +else +{ +lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; +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_3); +x_159 = lean_ctor_get(x_130, 0); +lean_inc(x_159); +x_160 = lean_ctor_get(x_130, 1); +lean_inc(x_160); +if (lean_is_exclusive(x_130)) { + lean_ctor_release(x_130, 0); + lean_ctor_release(x_130, 1); + x_161 = x_130; +} else { + lean_dec_ref(x_130); + x_161 = lean_box(0); +} +if (lean_is_scalar(x_161)) { + x_162 = lean_alloc_ctor(1, 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; lean_object* x_167; uint8_t x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; size_t x_172; size_t x_173; lean_object* x_174; +x_163 = lean_ctor_get(x_13, 1); +x_164 = lean_ctor_get(x_13, 0); +lean_inc(x_163); +lean_inc(x_164); +lean_dec(x_13); +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; +} else { + lean_dec_ref(x_163); + x_166 = lean_box(0); +} +x_167 = lean_box(0); +x_168 = 0; +x_169 = lean_box(x_168); +if (lean_is_scalar(x_166)) { + x_170 = lean_alloc_ctor(0, 2, 0); +} else { + x_170 = x_166; +} +lean_ctor_set(x_170, 0, x_169); +lean_ctor_set(x_170, 1, x_165); +x_171 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_171, 0, x_164); +lean_ctor_set(x_171, 1, x_170); +x_172 = lean_array_size(x_2); +x_173 = 0; +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_5); +lean_inc(x_3); +x_174 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__14___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__15(x_2, x_3, x_5, x_167, x_2, x_172, x_173, x_171, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_174) == 0) +{ +lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; uint8_t x_186; +x_175 = lean_ctor_get(x_174, 0); +lean_inc(x_175); +x_176 = lean_ctor_get(x_175, 1); +lean_inc(x_176); +x_177 = lean_ctor_get(x_174, 1); +lean_inc(x_177); +if (lean_is_exclusive(x_174)) { + lean_ctor_release(x_174, 0); + lean_ctor_release(x_174, 1); + x_178 = x_174; +} else { + lean_dec_ref(x_174); + x_178 = lean_box(0); +} +x_179 = lean_ctor_get(x_175, 0); +lean_inc(x_179); +if (lean_is_exclusive(x_175)) { + lean_ctor_release(x_175, 0); + lean_ctor_release(x_175, 1); + x_180 = x_175; +} else { + lean_dec_ref(x_175); + x_180 = lean_box(0); +} +x_181 = lean_ctor_get(x_176, 0); +lean_inc(x_181); +x_182 = lean_ctor_get(x_176, 1); +lean_inc(x_182); +if (lean_is_exclusive(x_176)) { + lean_ctor_release(x_176, 0); + lean_ctor_release(x_176, 1); + x_183 = x_176; +} else { + lean_dec_ref(x_176); + x_183 = lean_box(0); +} +x_184 = lean_unsigned_to_nat(0u); +x_185 = l_Lean_RBNode_fold___at_Lean_RBMap_size___spec__1___rarg(x_184, x_179); +x_186 = lean_nat_dec_eq(x_185, x_1); +lean_dec(x_185); +if (x_186 == 0) +{ +lean_object* x_187; uint8_t x_188; lean_object* x_189; lean_object* x_190; +lean_dec(x_183); +lean_dec(x_180); +lean_dec(x_178); +x_187 = lean_box(0); +x_188 = lean_unbox(x_181); +lean_dec(x_181); +lean_inc(x_6); +x_189 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___lambda__1(x_188, x_182, x_179, x_6, x_187, x_8, x_9, x_10, x_11, x_177); +x_190 = lean_ctor_get(x_189, 0); +lean_inc(x_190); +if (lean_obj_tag(x_190) == 0) +{ +lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; +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_3); +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; +} else { + lean_dec_ref(x_189); + x_192 = lean_box(0); +} +x_193 = lean_ctor_get(x_190, 0); +lean_inc(x_193); +lean_dec(x_190); +if (lean_is_scalar(x_192)) { + x_194 = lean_alloc_ctor(0, 2, 0); +} else { + x_194 = x_192; +} +lean_ctor_set(x_194, 0, x_193); +lean_ctor_set(x_194, 1, x_191); +return x_194; +} +else +{ +lean_object* x_195; lean_object* x_196; +x_195 = lean_ctor_get(x_189, 1); +lean_inc(x_195); +lean_dec(x_189); +x_196 = lean_ctor_get(x_190, 0); +lean_inc(x_196); +lean_dec(x_190); +x_7 = x_196; +x_12 = x_195; +goto _start; +} +} +else +{ +lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; +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_3); +if (lean_is_scalar(x_183)) { + x_198 = lean_alloc_ctor(0, 2, 0); +} else { + x_198 = x_183; +} +lean_ctor_set(x_198, 0, x_181); +lean_ctor_set(x_198, 1, x_182); +if (lean_is_scalar(x_180)) { + x_199 = lean_alloc_ctor(0, 2, 0); +} else { + x_199 = x_180; +} +lean_ctor_set(x_199, 0, x_179); +lean_ctor_set(x_199, 1, x_198); +x_200 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___closed__1; +x_201 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_201, 0, x_200); +lean_ctor_set(x_201, 1, x_199); +if (lean_is_scalar(x_178)) { + x_202 = lean_alloc_ctor(0, 2, 0); +} else { + x_202 = x_178; +} +lean_ctor_set(x_202, 0, x_201); +lean_ctor_set(x_202, 1, x_177); +return x_202; +} +} +else +{ +lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; +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_3); +x_203 = lean_ctor_get(x_174, 0); +lean_inc(x_203); +x_204 = lean_ctor_get(x_174, 1); +lean_inc(x_204); +if (lean_is_exclusive(x_174)) { + lean_ctor_release(x_174, 0); + lean_ctor_release(x_174, 1); + x_205 = x_174; +} else { + lean_dec_ref(x_174); + x_205 = lean_box(0); +} +if (lean_is_scalar(x_205)) { + x_206 = lean_alloc_ctor(1, 2, 0); +} else { + x_206 = x_205; +} +lean_ctor_set(x_206, 0, x_203); +lean_ctor_set(x_206, 1, x_204); +return x_206; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__17(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; uint8_t x_13; +x_12 = lean_ctor_get(x_6, 1); +lean_inc(x_12); +lean_dec(x_6); +x_13 = !lean_is_exclusive(x_12); +if (x_13 == 0) +{ +lean_object* x_14; uint8_t x_15; +x_14 = lean_ctor_get(x_12, 1); +x_15 = !lean_is_exclusive(x_14); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; uint8_t x_18; lean_object* x_19; size_t x_20; size_t x_21; lean_object* x_22; +x_16 = lean_ctor_get(x_14, 0); +lean_dec(x_16); +x_17 = lean_box(0); +x_18 = 0; +x_19 = lean_box(x_18); +lean_ctor_set(x_14, 0, x_19); +x_20 = lean_array_size(x_2); +x_21 = 0; +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_4); +lean_inc(x_3); +x_22 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__14___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__15(x_2, x_3, x_4, x_17, x_2, x_20, x_21, x_12, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_22) == 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_is_exclusive(x_22); +if (x_25 == 0) +{ +lean_object* x_26; lean_object* x_27; uint8_t x_28; +x_26 = lean_ctor_get(x_22, 1); +x_27 = lean_ctor_get(x_22, 0); +lean_dec(x_27); +x_28 = !lean_is_exclusive(x_23); +if (x_28 == 0) +{ +lean_object* x_29; lean_object* x_30; uint8_t x_31; +x_29 = lean_ctor_get(x_23, 0); +x_30 = lean_ctor_get(x_23, 1); +lean_dec(x_30); +x_31 = !lean_is_exclusive(x_24); +if (x_31 == 0) +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; +x_32 = lean_ctor_get(x_24, 0); +x_33 = lean_ctor_get(x_24, 1); +x_34 = lean_unsigned_to_nat(0u); +x_35 = l_Lean_RBNode_fold___at_Lean_RBMap_size___spec__1___rarg(x_34, x_29); +x_36 = lean_nat_dec_eq(x_35, x_1); +lean_dec(x_35); +if (x_36 == 0) +{ +lean_object* x_37; uint8_t x_38; lean_object* x_39; lean_object* x_40; +lean_free_object(x_24); +lean_free_object(x_23); +lean_free_object(x_22); +x_37 = lean_box(0); +x_38 = lean_unbox(x_32); +lean_dec(x_32); +lean_inc(x_5); +x_39 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___lambda__1(x_38, x_33, x_29, x_5, x_37, x_7, x_8, x_9, x_10, x_26); +x_40 = lean_ctor_get(x_39, 0); +lean_inc(x_40); +if (lean_obj_tag(x_40) == 0) +{ +uint8_t x_41; +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_41 = !lean_is_exclusive(x_39); +if (x_41 == 0) +{ +lean_object* x_42; lean_object* x_43; +x_42 = lean_ctor_get(x_39, 0); +lean_dec(x_42); +x_43 = lean_ctor_get(x_40, 0); +lean_inc(x_43); +lean_dec(x_40); +lean_ctor_set(x_39, 0, x_43); +return x_39; +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_39, 1); +lean_inc(x_44); +lean_dec(x_39); +x_45 = lean_ctor_get(x_40, 0); +lean_inc(x_45); +lean_dec(x_40); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_44); +return x_46; +} +} +else +{ +lean_object* x_47; lean_object* x_48; +x_47 = lean_ctor_get(x_39, 1); +lean_inc(x_47); +lean_dec(x_39); +x_48 = lean_ctor_get(x_40, 0); +lean_inc(x_48); +lean_dec(x_40); +x_6 = x_48; +x_11 = x_47; +goto _start; +} +} +else +{ +lean_object* x_50; lean_object* x_51; +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_50 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___closed__1; +x_51 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_51, 0, x_50); +lean_ctor_set(x_51, 1, x_23); +lean_ctor_set(x_22, 0, x_51); +return x_22; +} +} +else +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; uint8_t x_56; +x_52 = lean_ctor_get(x_24, 0); +x_53 = lean_ctor_get(x_24, 1); +lean_inc(x_53); +lean_inc(x_52); +lean_dec(x_24); +x_54 = lean_unsigned_to_nat(0u); +x_55 = l_Lean_RBNode_fold___at_Lean_RBMap_size___spec__1___rarg(x_54, x_29); +x_56 = lean_nat_dec_eq(x_55, x_1); +lean_dec(x_55); +if (x_56 == 0) +{ +lean_object* x_57; uint8_t x_58; lean_object* x_59; lean_object* x_60; +lean_free_object(x_23); +lean_free_object(x_22); +x_57 = lean_box(0); +x_58 = lean_unbox(x_52); +lean_dec(x_52); +lean_inc(x_5); +x_59 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___lambda__1(x_58, x_53, x_29, x_5, x_57, x_7, x_8, x_9, x_10, x_26); +x_60 = lean_ctor_get(x_59, 0); +lean_inc(x_60); +if (lean_obj_tag(x_60) == 0) +{ +lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; +lean_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_61 = lean_ctor_get(x_59, 1); +lean_inc(x_61); +if (lean_is_exclusive(x_59)) { + lean_ctor_release(x_59, 0); + lean_ctor_release(x_59, 1); + x_62 = x_59; +} else { + lean_dec_ref(x_59); + x_62 = lean_box(0); +} +x_63 = lean_ctor_get(x_60, 0); +lean_inc(x_63); +lean_dec(x_60); +if (lean_is_scalar(x_62)) { + x_64 = lean_alloc_ctor(0, 2, 0); +} else { + x_64 = x_62; +} +lean_ctor_set(x_64, 0, x_63); +lean_ctor_set(x_64, 1, x_61); +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); +x_66 = lean_ctor_get(x_60, 0); +lean_inc(x_66); +lean_dec(x_60); +x_6 = x_66; +x_11 = x_65; +goto _start; +} +} +else +{ +lean_object* x_68; lean_object* x_69; lean_object* x_70; +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_68 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_68, 0, x_52); +lean_ctor_set(x_68, 1, x_53); +lean_ctor_set(x_23, 1, x_68); +x_69 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___closed__1; +x_70 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_70, 0, x_69); +lean_ctor_set(x_70, 1, x_23); +lean_ctor_set(x_22, 0, x_70); +return x_22; +} +} +} +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; uint8_t x_77; +x_71 = lean_ctor_get(x_23, 0); +lean_inc(x_71); +lean_dec(x_23); +x_72 = lean_ctor_get(x_24, 0); +lean_inc(x_72); +x_73 = lean_ctor_get(x_24, 1); +lean_inc(x_73); +if (lean_is_exclusive(x_24)) { + lean_ctor_release(x_24, 0); + lean_ctor_release(x_24, 1); + x_74 = x_24; +} else { + lean_dec_ref(x_24); + x_74 = lean_box(0); +} +x_75 = lean_unsigned_to_nat(0u); +x_76 = l_Lean_RBNode_fold___at_Lean_RBMap_size___spec__1___rarg(x_75, x_71); +x_77 = lean_nat_dec_eq(x_76, x_1); +lean_dec(x_76); +if (x_77 == 0) +{ +lean_object* x_78; uint8_t x_79; lean_object* x_80; lean_object* x_81; +lean_dec(x_74); +lean_free_object(x_22); +x_78 = lean_box(0); +x_79 = lean_unbox(x_72); +lean_dec(x_72); +lean_inc(x_5); +x_80 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___lambda__1(x_79, x_73, x_71, x_5, x_78, x_7, x_8, x_9, x_10, x_26); +x_81 = lean_ctor_get(x_80, 0); +lean_inc(x_81); +if (lean_obj_tag(x_81) == 0) +{ +lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_82 = lean_ctor_get(x_80, 1); +lean_inc(x_82); +if (lean_is_exclusive(x_80)) { + lean_ctor_release(x_80, 0); + lean_ctor_release(x_80, 1); + x_83 = x_80; +} else { + lean_dec_ref(x_80); + x_83 = lean_box(0); +} +x_84 = lean_ctor_get(x_81, 0); +lean_inc(x_84); +lean_dec(x_81); +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; +x_86 = lean_ctor_get(x_80, 1); +lean_inc(x_86); +lean_dec(x_80); +x_87 = lean_ctor_get(x_81, 0); +lean_inc(x_87); +lean_dec(x_81); +x_6 = x_87; +x_11 = x_86; +goto _start; +} +} +else +{ +lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; +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); +if (lean_is_scalar(x_74)) { + x_89 = lean_alloc_ctor(0, 2, 0); +} else { + x_89 = x_74; +} +lean_ctor_set(x_89, 0, x_72); +lean_ctor_set(x_89, 1, x_73); +x_90 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_90, 0, x_71); +lean_ctor_set(x_90, 1, x_89); +x_91 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___closed__1; +x_92 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_92, 0, x_91); +lean_ctor_set(x_92, 1, x_90); +lean_ctor_set(x_22, 0, x_92); +return x_22; +} +} +} +else +{ +lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; uint8_t x_101; +x_93 = lean_ctor_get(x_22, 1); +lean_inc(x_93); +lean_dec(x_22); +x_94 = lean_ctor_get(x_23, 0); +lean_inc(x_94); +if (lean_is_exclusive(x_23)) { + lean_ctor_release(x_23, 0); + lean_ctor_release(x_23, 1); + x_95 = x_23; +} else { + lean_dec_ref(x_23); + x_95 = lean_box(0); +} +x_96 = lean_ctor_get(x_24, 0); +lean_inc(x_96); +x_97 = lean_ctor_get(x_24, 1); +lean_inc(x_97); +if (lean_is_exclusive(x_24)) { + lean_ctor_release(x_24, 0); + lean_ctor_release(x_24, 1); + x_98 = x_24; +} else { + lean_dec_ref(x_24); + x_98 = lean_box(0); +} +x_99 = lean_unsigned_to_nat(0u); +x_100 = l_Lean_RBNode_fold___at_Lean_RBMap_size___spec__1___rarg(x_99, x_94); +x_101 = lean_nat_dec_eq(x_100, x_1); +lean_dec(x_100); +if (x_101 == 0) +{ +lean_object* x_102; uint8_t x_103; lean_object* x_104; lean_object* x_105; +lean_dec(x_98); +lean_dec(x_95); +x_102 = lean_box(0); +x_103 = lean_unbox(x_96); +lean_dec(x_96); +lean_inc(x_5); +x_104 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___lambda__1(x_103, x_97, x_94, x_5, x_102, x_7, x_8, x_9, x_10, x_93); +x_105 = lean_ctor_get(x_104, 0); +lean_inc(x_105); +if (lean_obj_tag(x_105) == 0) +{ +lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; +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_106 = lean_ctor_get(x_104, 1); +lean_inc(x_106); +if (lean_is_exclusive(x_104)) { + lean_ctor_release(x_104, 0); + lean_ctor_release(x_104, 1); + x_107 = x_104; +} else { + lean_dec_ref(x_104); + x_107 = lean_box(0); +} +x_108 = lean_ctor_get(x_105, 0); +lean_inc(x_108); +lean_dec(x_105); +if (lean_is_scalar(x_107)) { + x_109 = lean_alloc_ctor(0, 2, 0); +} else { + x_109 = x_107; +} +lean_ctor_set(x_109, 0, x_108); +lean_ctor_set(x_109, 1, x_106); +return x_109; +} +else +{ +lean_object* x_110; lean_object* x_111; +x_110 = lean_ctor_get(x_104, 1); +lean_inc(x_110); +lean_dec(x_104); +x_111 = lean_ctor_get(x_105, 0); +lean_inc(x_111); +lean_dec(x_105); +x_6 = x_111; +x_11 = x_110; +goto _start; +} +} +else +{ +lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +if (lean_is_scalar(x_98)) { + x_113 = lean_alloc_ctor(0, 2, 0); +} else { + x_113 = x_98; +} +lean_ctor_set(x_113, 0, x_96); +lean_ctor_set(x_113, 1, x_97); +if (lean_is_scalar(x_95)) { + x_114 = lean_alloc_ctor(0, 2, 0); +} else { + x_114 = x_95; +} +lean_ctor_set(x_114, 0, x_94); +lean_ctor_set(x_114, 1, x_113); +x_115 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___closed__1; +x_116 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_116, 0, x_115); +lean_ctor_set(x_116, 1, x_114); +x_117 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_117, 0, x_116); +lean_ctor_set(x_117, 1, x_93); +return x_117; +} +} +} +else +{ +uint8_t x_118; +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_118 = !lean_is_exclusive(x_22); +if (x_118 == 0) +{ +return x_22; +} +else +{ +lean_object* x_119; lean_object* x_120; lean_object* x_121; +x_119 = lean_ctor_get(x_22, 0); +x_120 = lean_ctor_get(x_22, 1); +lean_inc(x_120); +lean_inc(x_119); +lean_dec(x_22); +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 +{ +lean_object* x_122; lean_object* x_123; uint8_t x_124; lean_object* x_125; lean_object* x_126; size_t x_127; size_t x_128; lean_object* x_129; +x_122 = lean_ctor_get(x_14, 1); +lean_inc(x_122); +lean_dec(x_14); +x_123 = lean_box(0); +x_124 = 0; +x_125 = lean_box(x_124); +x_126 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_126, 0, x_125); +lean_ctor_set(x_126, 1, x_122); +lean_ctor_set(x_12, 1, x_126); +x_127 = lean_array_size(x_2); +x_128 = 0; +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_4); +lean_inc(x_3); +x_129 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__14___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__15(x_2, x_3, x_4, x_123, x_2, x_127, x_128, x_12, x_7, x_8, x_9, x_10, x_11); +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; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; uint8_t x_141; +x_130 = lean_ctor_get(x_129, 0); +lean_inc(x_130); +x_131 = lean_ctor_get(x_130, 1); +lean_inc(x_131); +x_132 = lean_ctor_get(x_129, 1); +lean_inc(x_132); +if (lean_is_exclusive(x_129)) { + lean_ctor_release(x_129, 0); + lean_ctor_release(x_129, 1); + x_133 = x_129; +} else { + lean_dec_ref(x_129); + x_133 = lean_box(0); +} +x_134 = lean_ctor_get(x_130, 0); +lean_inc(x_134); +if (lean_is_exclusive(x_130)) { + lean_ctor_release(x_130, 0); + lean_ctor_release(x_130, 1); + x_135 = x_130; +} else { + lean_dec_ref(x_130); + x_135 = lean_box(0); +} +x_136 = lean_ctor_get(x_131, 0); +lean_inc(x_136); +x_137 = lean_ctor_get(x_131, 1); +lean_inc(x_137); +if (lean_is_exclusive(x_131)) { + lean_ctor_release(x_131, 0); + lean_ctor_release(x_131, 1); + x_138 = x_131; +} else { + lean_dec_ref(x_131); + x_138 = lean_box(0); +} +x_139 = lean_unsigned_to_nat(0u); +x_140 = l_Lean_RBNode_fold___at_Lean_RBMap_size___spec__1___rarg(x_139, x_134); +x_141 = lean_nat_dec_eq(x_140, x_1); +lean_dec(x_140); +if (x_141 == 0) +{ +lean_object* x_142; uint8_t x_143; lean_object* x_144; lean_object* x_145; +lean_dec(x_138); +lean_dec(x_135); +lean_dec(x_133); +x_142 = lean_box(0); +x_143 = lean_unbox(x_136); +lean_dec(x_136); +lean_inc(x_5); +x_144 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___lambda__1(x_143, x_137, x_134, x_5, x_142, x_7, x_8, x_9, x_10, x_132); +x_145 = lean_ctor_get(x_144, 0); +lean_inc(x_145); +if (lean_obj_tag(x_145) == 0) +{ +lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; +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_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; +} else { + lean_dec_ref(x_144); + x_147 = lean_box(0); +} +x_148 = lean_ctor_get(x_145, 0); +lean_inc(x_148); +lean_dec(x_145); +if (lean_is_scalar(x_147)) { + x_149 = lean_alloc_ctor(0, 2, 0); +} else { + x_149 = x_147; +} +lean_ctor_set(x_149, 0, x_148); +lean_ctor_set(x_149, 1, x_146); +return x_149; +} +else +{ +lean_object* x_150; lean_object* x_151; +x_150 = lean_ctor_get(x_144, 1); +lean_inc(x_150); +lean_dec(x_144); +x_151 = lean_ctor_get(x_145, 0); +lean_inc(x_151); +lean_dec(x_145); +x_6 = x_151; +x_11 = x_150; +goto _start; +} +} +else +{ +lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; +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); +if (lean_is_scalar(x_138)) { + x_153 = lean_alloc_ctor(0, 2, 0); +} else { + x_153 = x_138; +} +lean_ctor_set(x_153, 0, x_136); +lean_ctor_set(x_153, 1, x_137); +if (lean_is_scalar(x_135)) { + x_154 = lean_alloc_ctor(0, 2, 0); +} else { + x_154 = x_135; +} +lean_ctor_set(x_154, 0, x_134); +lean_ctor_set(x_154, 1, x_153); +x_155 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___closed__1; +x_156 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_156, 0, x_155); +lean_ctor_set(x_156, 1, x_154); +if (lean_is_scalar(x_133)) { + x_157 = lean_alloc_ctor(0, 2, 0); +} else { + x_157 = x_133; +} +lean_ctor_set(x_157, 0, x_156); +lean_ctor_set(x_157, 1, x_132); +return x_157; +} +} +else +{ +lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; +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_158 = lean_ctor_get(x_129, 0); +lean_inc(x_158); +x_159 = lean_ctor_get(x_129, 1); +lean_inc(x_159); +if (lean_is_exclusive(x_129)) { + lean_ctor_release(x_129, 0); + lean_ctor_release(x_129, 1); + x_160 = x_129; +} else { + lean_dec_ref(x_129); + x_160 = lean_box(0); +} +if (lean_is_scalar(x_160)) { + x_161 = lean_alloc_ctor(1, 2, 0); +} else { + x_161 = x_160; +} +lean_ctor_set(x_161, 0, x_158); +lean_ctor_set(x_161, 1, x_159); +return x_161; +} +} +} +else +{ +lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; uint8_t x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; size_t x_171; size_t x_172; lean_object* x_173; +x_162 = lean_ctor_get(x_12, 1); +x_163 = lean_ctor_get(x_12, 0); +lean_inc(x_162); +lean_inc(x_163); +lean_dec(x_12); +x_164 = lean_ctor_get(x_162, 1); +lean_inc(x_164); +if (lean_is_exclusive(x_162)) { + lean_ctor_release(x_162, 0); + lean_ctor_release(x_162, 1); + x_165 = x_162; +} else { + lean_dec_ref(x_162); + x_165 = lean_box(0); +} +x_166 = lean_box(0); +x_167 = 0; +x_168 = lean_box(x_167); +if (lean_is_scalar(x_165)) { + x_169 = lean_alloc_ctor(0, 2, 0); +} else { + x_169 = x_165; +} +lean_ctor_set(x_169, 0, x_168); +lean_ctor_set(x_169, 1, x_164); +x_170 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_170, 0, x_163); +lean_ctor_set(x_170, 1, x_169); +x_171 = lean_array_size(x_2); +x_172 = 0; +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_4); +lean_inc(x_3); +x_173 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__14___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__15(x_2, x_3, x_4, x_166, x_2, x_171, x_172, x_170, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_173) == 0) +{ +lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; uint8_t x_185; +x_174 = lean_ctor_get(x_173, 0); +lean_inc(x_174); +x_175 = lean_ctor_get(x_174, 1); +lean_inc(x_175); +x_176 = lean_ctor_get(x_173, 1); +lean_inc(x_176); +if (lean_is_exclusive(x_173)) { + lean_ctor_release(x_173, 0); + lean_ctor_release(x_173, 1); + x_177 = x_173; +} else { + lean_dec_ref(x_173); + x_177 = lean_box(0); +} +x_178 = lean_ctor_get(x_174, 0); +lean_inc(x_178); +if (lean_is_exclusive(x_174)) { + lean_ctor_release(x_174, 0); + lean_ctor_release(x_174, 1); + x_179 = x_174; +} else { + lean_dec_ref(x_174); + x_179 = lean_box(0); +} +x_180 = lean_ctor_get(x_175, 0); +lean_inc(x_180); +x_181 = lean_ctor_get(x_175, 1); +lean_inc(x_181); +if (lean_is_exclusive(x_175)) { + lean_ctor_release(x_175, 0); + lean_ctor_release(x_175, 1); + x_182 = x_175; +} else { + lean_dec_ref(x_175); + x_182 = lean_box(0); +} +x_183 = lean_unsigned_to_nat(0u); +x_184 = l_Lean_RBNode_fold___at_Lean_RBMap_size___spec__1___rarg(x_183, x_178); +x_185 = lean_nat_dec_eq(x_184, x_1); +lean_dec(x_184); +if (x_185 == 0) +{ +lean_object* x_186; uint8_t x_187; lean_object* x_188; lean_object* x_189; +lean_dec(x_182); +lean_dec(x_179); +lean_dec(x_177); +x_186 = lean_box(0); +x_187 = lean_unbox(x_180); +lean_dec(x_180); +lean_inc(x_5); +x_188 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___lambda__1(x_187, x_181, x_178, x_5, x_186, x_7, x_8, x_9, x_10, x_176); +x_189 = lean_ctor_get(x_188, 0); +lean_inc(x_189); +if (lean_obj_tag(x_189) == 0) +{ +lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; +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_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); +} +x_192 = lean_ctor_get(x_189, 0); +lean_inc(x_192); +lean_dec(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); +return x_193; +} +else +{ +lean_object* x_194; lean_object* x_195; +x_194 = lean_ctor_get(x_188, 1); +lean_inc(x_194); +lean_dec(x_188); +x_195 = lean_ctor_get(x_189, 0); +lean_inc(x_195); +lean_dec(x_189); +x_6 = x_195; +x_11 = x_194; +goto _start; +} +} +else +{ +lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; +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); +if (lean_is_scalar(x_182)) { + x_197 = lean_alloc_ctor(0, 2, 0); +} else { + x_197 = x_182; +} +lean_ctor_set(x_197, 0, x_180); +lean_ctor_set(x_197, 1, x_181); +if (lean_is_scalar(x_179)) { + x_198 = lean_alloc_ctor(0, 2, 0); +} else { + x_198 = x_179; +} +lean_ctor_set(x_198, 0, x_178); +lean_ctor_set(x_198, 1, x_197); +x_199 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___closed__1; +x_200 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_200, 0, x_199); +lean_ctor_set(x_200, 1, x_198); +if (lean_is_scalar(x_177)) { + x_201 = lean_alloc_ctor(0, 2, 0); +} else { + x_201 = x_177; +} +lean_ctor_set(x_201, 0, x_200); +lean_ctor_set(x_201, 1, x_176); +return x_201; +} +} +else +{ +lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; +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_202 = lean_ctor_get(x_173, 0); +lean_inc(x_202); +x_203 = lean_ctor_get(x_173, 1); +lean_inc(x_203); +if (lean_is_exclusive(x_173)) { + lean_ctor_release(x_173, 0); + lean_ctor_release(x_173, 1); + x_204 = x_173; +} else { + lean_dec_ref(x_173); + x_204 = lean_box(0); +} +if (lean_is_scalar(x_204)) { + x_205 = lean_alloc_ctor(1, 2, 0); +} else { + x_205 = x_204; +} +lean_ctor_set(x_205, 0, x_202); +lean_ctor_set(x_205, 1, x_203); +return x_205; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___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) { +_start: +{ +lean_object* x_15; uint8_t x_16; +x_15 = lean_ctor_get(x_5, 1); +x_16 = lean_nat_dec_lt(x_7, x_15); +if (x_16 == 0) +{ +lean_object* x_17; +lean_dec(x_7); +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_6); +lean_ctor_set(x_17, 1, x_14); +return x_17; +} +else +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_18 = lean_array_fget(x_1, x_7); +x_19 = l_Lean_Expr_fvarId_x21(x_18); +lean_dec(x_18); +x_20 = l_Lean_RBNode_findCore___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__2(x_3, x_19); +lean_dec(x_19); +if (lean_obj_tag(x_20) == 0) +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +lean_inc(x_7); +x_21 = lean_array_push(x_6, x_7); +x_22 = lean_ctor_get(x_5, 2); +x_23 = lean_nat_add(x_7, x_22); +lean_dec(x_7); +x_6 = x_21; +x_7 = x_23; +x_8 = lean_box(0); +x_9 = lean_box(0); +goto _start; +} +else +{ +lean_object* x_25; lean_object* x_26; +lean_dec(x_20); +x_25 = lean_ctor_get(x_5, 2); +x_26 = lean_nat_add(x_7, x_25); +lean_dec(x_7); +x_7 = x_26; +x_8 = lean_box(0); +x_9 = lean_box(0); +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__18___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__19(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_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_ctor_get(x_4, 1); +x_15 = lean_nat_dec_lt(x_6, x_14); +if (x_15 == 0) +{ +lean_object* x_16; +lean_dec(x_6); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_5); +lean_ctor_set(x_16, 1, x_13); +return x_16; +} +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_17 = lean_array_fget(x_1, x_6); +x_18 = l_Lean_Expr_fvarId_x21(x_17); +lean_dec(x_17); +x_19 = l_Lean_RBNode_findCore___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__2(x_2, x_18); +lean_dec(x_18); +if (lean_obj_tag(x_19) == 0) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; +lean_inc(x_6); +x_20 = lean_array_push(x_5, x_6); +x_21 = lean_ctor_get(x_4, 2); +x_22 = lean_nat_add(x_6, x_21); +lean_dec(x_6); +x_5 = x_20; +x_6 = x_22; +x_7 = lean_box(0); +x_8 = lean_box(0); +goto _start; +} +else +{ +lean_object* x_24; lean_object* x_25; +lean_dec(x_19); +x_24 = lean_ctor_get(x_4, 2); +x_25 = lean_nat_add(x_6, x_24); +lean_dec(x_6); +x_6 = x_25; +x_7 = lean_box(0); +x_8 = lean_box(0); +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__20(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_eq(x_2, x_3); +if (x_5 == 0) +{ +size_t x_6; size_t x_7; lean_object* x_8; lean_object* x_9; +x_6 = 1; +x_7 = lean_usize_sub(x_2, x_6); +x_8 = lean_array_uget(x_1, x_7); +x_9 = l_Std_DHashMap_Internal_AssocList_foldrM___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__2(x_4, x_8); +lean_dec(x_8); +lean_dec(x_4); +x_2 = x_7; +x_4 = x_9; +goto _start; +} +else +{ +return x_4; +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_11 = lean_array_mk(x_1); +x_12 = lean_unsigned_to_nat(0u); +x_13 = lean_unsigned_to_nat(1u); +x_14 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_14, 0, x_12); +lean_ctor_set(x_14, 1, x_2); +lean_ctor_set(x_14, 2, x_13); +x_15 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__18___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__19(x_3, x_4, x_14, x_14, x_11, x_12, lean_box(0), lean_box(0), x_6, x_7, x_8, x_9, x_10); +lean_dec(x_14); +x_16 = !lean_is_exclusive(x_15); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_17 = lean_ctor_get(x_15, 0); +x_18 = lean_array_to_list(x_17); +x_19 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_15, 0, x_19); +return x_15; +} +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_20 = lean_ctor_get(x_15, 0); +x_21 = lean_ctor_get(x_15, 1); +lean_inc(x_21); +lean_inc(x_20); +lean_dec(x_15); +x_22 = lean_array_to_list(x_20); +x_23 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_23, 0, x_22); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set(x_24, 1, x_21); +return x_24; +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +lean_object* x_14; 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_14 = lean_box(0); +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_1); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_7); +lean_ctor_set(x_18, 1, 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); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_4); +x_20 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__17(x_2, x_3, x_4, x_5, x_14, x_19, x_9, x_10, x_11, x_12, x_13); +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_21, 1); +lean_inc(x_22); +x_23 = lean_ctor_get(x_21, 0); +lean_inc(x_23); +lean_dec(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; +x_24 = lean_ctor_get(x_20, 1); +lean_inc(x_24); +lean_dec(x_20); +x_25 = lean_ctor_get(x_22, 0); +lean_inc(x_25); +lean_dec(x_22); +x_26 = lean_box(0); +x_27 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__1(x_4, x_6, x_3, x_25, x_26, x_9, x_10, x_11, x_12, x_24); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_25); +return x_27; +} +else +{ +uint8_t x_28; +lean_dec(x_22); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_6); +lean_dec(x_4); +x_28 = !lean_is_exclusive(x_20); +if (x_28 == 0) +{ +lean_object* x_29; lean_object* x_30; +x_29 = lean_ctor_get(x_20, 0); +lean_dec(x_29); +x_30 = lean_ctor_get(x_23, 0); +lean_inc(x_30); +lean_dec(x_23); +lean_ctor_set(x_20, 0, x_30); +return x_20; +} +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_20, 1); +lean_inc(x_31); +lean_dec(x_20); +x_32 = lean_ctor_get(x_23, 0); +lean_inc(x_32); +lean_dec(x_23); +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_31); +return x_33; +} +} +} +else +{ +uint8_t x_34; +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_6); +lean_dec(x_4); +x_34 = !lean_is_exclusive(x_20); +if (x_34 == 0) +{ +return x_20; +} +else +{ +lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_35 = lean_ctor_get(x_20, 0); +x_36 = lean_ctor_get(x_20, 1); +lean_inc(x_36); +lean_inc(x_35); +lean_dec(x_20); +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; +} +} +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("numParams == xs.size\n ", 25, 25); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__3; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3___closed__1; +x_3 = lean_string_append(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("_private.Lean.Meta.Tactic.Grind.EMatchTheorem.0.Lean.Meta.Grind.checkCoverage", 77, 77); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___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; lean_object* x_6; +x_1 = l_Lean_Meta_Grind_EMatchTheorems_insert___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3___closed__3; +x_3 = lean_unsigned_to_nat(263u); +x_4 = lean_unsigned_to_nat(4u); +x_5 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3___closed__2; +x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); +return x_6; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___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; uint8_t x_11; +x_10 = lean_array_get_size(x_3); +x_11 = lean_nat_dec_eq(x_1, x_10); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; +lean_dec(x_10); +lean_dec(x_3); +x_12 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3___closed__4; +x_13 = l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__1(x_12, x_5, x_6, x_7, x_8, x_9); +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_48; uint8_t x_49; +x_14 = lean_box(0); +x_15 = lean_ctor_get(x_2, 1); +x_16 = lean_array_get_size(x_15); +lean_inc(x_3); +x_17 = lean_array_to_list(x_3); +x_18 = l_List_mapTR_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__3(x_17, x_14); +x_19 = l_Lean_RBTree_ofList___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__4(x_18); +x_20 = lean_box(0); +x_48 = lean_unsigned_to_nat(0u); +x_49 = lean_nat_dec_lt(x_48, x_16); +if (x_49 == 0) +{ +lean_dec(x_16); +x_21 = x_14; +goto block_47; +} +else +{ +size_t x_50; size_t x_51; lean_object* x_52; +x_50 = lean_usize_of_nat(x_16); +lean_dec(x_16); +x_51 = 0; +x_52 = l_Array_foldrMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__20(x_15, x_50, x_51, x_14); +x_21 = x_52; +goto block_47; +} +block_47: +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_22 = l_List_mapTR_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__5(x_1, x_3, x_10, x_21, x_14); +lean_inc(x_22); +x_23 = l_Lean_RBTree_ofList___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__4(x_22); +lean_inc(x_5); +lean_inc(x_23); +lean_inc(x_22); +x_24 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__8___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__9(x_14, x_19, x_20, x_22, x_22, x_22, x_23, lean_box(0), x_5, x_6, x_7, x_8, x_9); +lean_dec(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; lean_object* x_28; lean_object* x_29; uint8_t x_30; +x_26 = lean_ctor_get(x_24, 0); +x_27 = lean_ctor_get(x_24, 1); +x_28 = lean_unsigned_to_nat(0u); +x_29 = l_Lean_RBNode_fold___at_Lean_RBMap_size___spec__1___rarg(x_28, x_26); +x_30 = lean_nat_dec_eq(x_29, x_1); +lean_dec(x_29); +if (x_30 == 0) +{ +lean_object* x_31; lean_object* x_32; +lean_free_object(x_24); +x_31 = lean_box(0); +x_32 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__2(x_23, x_1, x_3, x_14, x_19, x_10, x_26, x_31, x_5, x_6, x_7, x_8, x_27); +lean_dec(x_3); +return x_32; +} +else +{ +lean_object* x_33; +lean_dec(x_26); +lean_dec(x_23); +lean_dec(x_19); +lean_dec(x_10); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +x_33 = lean_box(0); +lean_ctor_set(x_24, 0, x_33); +return x_24; +} +} +else +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; uint8_t x_38; +x_34 = lean_ctor_get(x_24, 0); +x_35 = lean_ctor_get(x_24, 1); +lean_inc(x_35); +lean_inc(x_34); +lean_dec(x_24); +x_36 = lean_unsigned_to_nat(0u); +x_37 = l_Lean_RBNode_fold___at_Lean_RBMap_size___spec__1___rarg(x_36, x_34); +x_38 = lean_nat_dec_eq(x_37, x_1); +lean_dec(x_37); +if (x_38 == 0) +{ +lean_object* x_39; lean_object* x_40; +x_39 = lean_box(0); +x_40 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__2(x_23, x_1, x_3, x_14, x_19, x_10, x_34, x_39, x_5, x_6, x_7, x_8, x_35); +lean_dec(x_3); +return x_40; +} +else +{ +lean_object* x_41; lean_object* x_42; +lean_dec(x_34); +lean_dec(x_23); +lean_dec(x_19); +lean_dec(x_10); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +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_35); +return x_42; +} +} +} +else +{ +uint8_t x_43; +lean_dec(x_23); +lean_dec(x_19); +lean_dec(x_10); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +x_43 = !lean_is_exclusive(x_24); +if (x_43 == 0) +{ +return x_24; +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_24, 0); +x_45 = lean_ctor_get(x_24, 1); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_24); +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_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +x_10 = lean_infer_type(x_1, 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; uint8_t x_15; lean_object* x_16; +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_2); +x_13 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_13, 0, x_2); +x_14 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3___boxed), 9, 2); +lean_closure_set(x_14, 0, x_2); +lean_closure_set(x_14, 1, x_3); +x_15 = 0; +x_16 = l_Lean_Meta_forallBoundedTelescope___at_Lean_Meta_arrowDomainsN___spec__6___rarg(x_11, x_13, x_14, x_15, x_5, x_6, x_7, x_8, x_12); +return x_16; +} +else +{ +uint8_t x_17; +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_17 = !lean_is_exclusive(x_10); +if (x_17 == 0) +{ +return x_10; +} +else +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_18 = lean_ctor_get(x_10, 0); +x_19 = lean_ctor_get(x_10, 1); +lean_inc(x_19); +lean_inc(x_18); +lean_dec(x_10); +x_20 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_20, 0, x_18); +lean_ctor_set(x_20, 1, x_19); +return x_20; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage(lean_object* x_1, lean_object* x_2, lean_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_ctor_get(x_3, 0); +lean_inc(x_9); +x_10 = lean_nat_dec_eq(x_9, x_2); +lean_dec(x_9); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; +x_11 = lean_box(0); +x_12 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__4(x_1, x_2, x_3, x_11, x_4, x_5, x_6, x_7, x_8); +return x_12; +} +else +{ +lean_object* x_13; lean_object* x_14; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_13 = lean_box(0); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_8); +return x_14; +} +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldrM___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__2___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_DHashMap_Internal_AssocList_foldrM___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__2(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_List_mapTR_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__5___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_mapTR_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__5(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___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; size_t x_8; lean_object* x_9; +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_9 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__6(x_1, x_2, x_3, x_7, x_8, x_6); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__6___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__7___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___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__6___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__7(x_1, x_2, x_6, x_7, x_5); +lean_dec(x_2); +lean_dec(x_1); +return x_8; +} +} +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___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) { +_start: +{ +lean_object* x_15; +x_15 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___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); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_15; +} +} +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__8___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__9___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +lean_object* x_14; +x_14 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__8___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__9(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___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: +{ +size_t x_7; size_t x_8; lean_object* x_9; +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_9 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__10(x_1, x_2, x_3, x_7, x_8, x_6); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__10___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__11___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +size_t x_6; size_t x_7; lean_object* x_8; +x_6 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__10___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__11(x_1, x_2, x_6, x_7, x_5); +lean_dec(x_2); +lean_dec(x_1); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___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: +{ +size_t x_7; size_t x_8; lean_object* x_9; +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_9 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__12(x_1, x_2, x_3, x_7, x_8, x_6); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__12___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__13___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +size_t x_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___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__12___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__13(x_1, x_2, x_6, x_7, x_5); +lean_dec(x_2); +lean_dec(x_1); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__14___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +size_t x_15; size_t x_16; lean_object* x_17; +x_15 = lean_unbox_usize(x_7); +lean_dec(x_7); +x_16 = lean_unbox_usize(x_8); +lean_dec(x_8); +x_17 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__14(x_1, x_2, x_3, x_4, x_5, x_6, x_15, x_16, x_9, x_10, x_11, x_12, x_13, x_14); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +return x_17; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__14___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__15___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +size_t x_14; size_t x_15; lean_object* x_16; +x_14 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_15 = lean_unbox_usize(x_7); +lean_dec(x_7); +x_16 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__14___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__15(x_1, x_2, x_3, x_4, x_5, x_14, x_15, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +return x_16; +} +} +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___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_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___lambda__1(x_11, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +x_13 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__17___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; +x_12 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__17(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_2); +lean_dec(x_1); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___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) { +_start: +{ +lean_object* x_15; +x_15 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___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); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_15; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__18___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__19___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +lean_object* x_14; +x_14 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__18___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__19(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__20___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_foldrMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__20(x_1, x_5, x_6, x_4); +lean_dec(x_1); +return x_7; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___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_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_11; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +lean_object* x_14; +x_14 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_8); +lean_dec(x_3); +lean_dec(x_2); +return x_14; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___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_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +return x_10; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_4); +return x_10; +} +} +static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___lambda__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(" : ", 3, 3); +return x_1; +} +} +static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___lambda__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___lambda__1___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___lambda__1(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +lean_inc(x_1); +x_11 = lean_infer_type(x_1, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_11) == 0) +{ +uint8_t x_12; +x_12 = !lean_is_exclusive(x_11); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; 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 = l_Lean_MessageData_ofExpr(x_1); +lean_inc(x_2); +x_15 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_15, 0, x_2); +lean_ctor_set(x_15, 1, x_14); +x_16 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___lambda__1___closed__2; +x_17 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_17, 0, x_15); +lean_ctor_set(x_17, 1, x_16); +x_18 = l_Lean_MessageData_ofExpr(x_13); +x_19 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_19, 0, x_17); +lean_ctor_set(x_19, 1, x_18); +x_20 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_20, 0, x_19); +lean_ctor_set(x_20, 1, x_2); +x_21 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_21, 0, x_4); +lean_ctor_set(x_21, 1, x_20); +x_22 = lean_box(x_3); +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_21); +x_24 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set(x_11, 0, x_24); +return x_11; +} +else +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; +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 = l_Lean_MessageData_ofExpr(x_1); +lean_inc(x_2); +x_28 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_28, 0, x_2); +lean_ctor_set(x_28, 1, x_27); +x_29 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___lambda__1___closed__2; +x_30 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_30, 0, x_28); +lean_ctor_set(x_30, 1, x_29); +x_31 = l_Lean_MessageData_ofExpr(x_25); +x_32 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_32, 0, x_30); +lean_ctor_set(x_32, 1, x_31); +x_33 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_2); +x_34 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_34, 0, x_4); +lean_ctor_set(x_34, 1, x_33); +x_35 = lean_box(x_3); +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set(x_36, 1, x_34); +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_26); +return x_38; +} +} +else +{ +uint8_t x_39; +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_39 = !lean_is_exclusive(x_11); +if (x_39 == 0) +{ +return x_11; +} +else +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; +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_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); +return x_42; +} +} +} +} +static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("\n", 1, 1); +return x_1; +} +} +static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___closed__1; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___closed__2; +x_2 = l_Lean_MessageData_ofFormat(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +lean_object* x_15; uint8_t x_16; +x_15 = lean_ctor_get(x_5, 1); +x_16 = lean_nat_dec_lt(x_7, x_15); +if (x_16 == 0) +{ +lean_object* x_17; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_7); +lean_dec(x_3); +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_6); +lean_ctor_set(x_17, 1, x_14); +return x_17; +} +else +{ +uint8_t x_18; +x_18 = !lean_is_exclusive(x_6); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_19 = lean_ctor_get(x_6, 0); +x_20 = lean_ctor_get(x_6, 1); +x_21 = l_List_elem___at_Lean_Meta_Occurrences_contains___spec__1(x_7, x_1); +if (x_21 == 0) +{ +lean_object* x_22; lean_object* x_23; +x_22 = lean_ctor_get(x_5, 2); +x_23 = lean_nat_add(x_7, x_22); +lean_dec(x_7); +x_7 = x_23; +x_8 = lean_box(0); +x_9 = lean_box(0); +goto _start; +} +else +{ +lean_object* x_25; uint8_t x_26; +lean_free_object(x_6); +x_25 = lean_array_fget(x_2, x_7); +x_26 = lean_unbox(x_19); +if (x_26 == 0) +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; uint8_t x_30; lean_object* x_31; +x_27 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___closed__3; +x_28 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_28, 0, x_20); +lean_ctor_set(x_28, 1, x_27); +x_29 = lean_box(0); +x_30 = lean_unbox(x_19); +lean_dec(x_19); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_3); +x_31 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___lambda__1(x_25, x_3, x_30, x_28, x_29, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_31) == 0) +{ +lean_object* x_32; +x_32 = lean_ctor_get(x_31, 0); +lean_inc(x_32); +if (lean_obj_tag(x_32) == 0) +{ +uint8_t x_33; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_7); +lean_dec(x_3); +x_33 = !lean_is_exclusive(x_31); +if (x_33 == 0) +{ +lean_object* x_34; lean_object* x_35; +x_34 = lean_ctor_get(x_31, 0); +lean_dec(x_34); +x_35 = lean_ctor_get(x_32, 0); +lean_inc(x_35); +lean_dec(x_32); +lean_ctor_set(x_31, 0, x_35); +return x_31; +} +else +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_31, 1); +lean_inc(x_36); +lean_dec(x_31); +x_37 = lean_ctor_get(x_32, 0); +lean_inc(x_37); +lean_dec(x_32); +x_38 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_38, 0, x_37); +lean_ctor_set(x_38, 1, x_36); +return x_38; +} +} +else +{ +lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_39 = lean_ctor_get(x_31, 1); +lean_inc(x_39); +lean_dec(x_31); +x_40 = lean_ctor_get(x_32, 0); +lean_inc(x_40); +lean_dec(x_32); +x_41 = lean_ctor_get(x_5, 2); +x_42 = lean_nat_add(x_7, x_41); +lean_dec(x_7); +x_6 = x_40; +x_7 = x_42; +x_8 = lean_box(0); +x_9 = lean_box(0); +x_14 = x_39; +goto _start; +} +} +else +{ +uint8_t x_44; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_7); +lean_dec(x_3); +x_44 = !lean_is_exclusive(x_31); +if (x_44 == 0) +{ +return x_31; +} +else +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_45 = lean_ctor_get(x_31, 0); +x_46 = lean_ctor_get(x_31, 1); +lean_inc(x_46); +lean_inc(x_45); +lean_dec(x_31); +x_47 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_47, 0, x_45); +lean_ctor_set(x_47, 1, x_46); +return x_47; +} +} +} +else +{ +uint8_t x_48; lean_object* x_49; lean_object* x_50; +lean_dec(x_19); +x_48 = 0; +x_49 = lean_box(0); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_3); +x_50 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___lambda__1(x_25, x_3, x_48, x_20, x_49, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_50) == 0) +{ +lean_object* x_51; +x_51 = lean_ctor_get(x_50, 0); +lean_inc(x_51); +if (lean_obj_tag(x_51) == 0) +{ +uint8_t x_52; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_7); +lean_dec(x_3); +x_52 = !lean_is_exclusive(x_50); +if (x_52 == 0) +{ +lean_object* x_53; lean_object* x_54; +x_53 = lean_ctor_get(x_50, 0); +lean_dec(x_53); +x_54 = lean_ctor_get(x_51, 0); +lean_inc(x_54); +lean_dec(x_51); +lean_ctor_set(x_50, 0, x_54); +return x_50; +} +else +{ +lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_55 = lean_ctor_get(x_50, 1); +lean_inc(x_55); +lean_dec(x_50); +x_56 = lean_ctor_get(x_51, 0); +lean_inc(x_56); +lean_dec(x_51); +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_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; +x_58 = lean_ctor_get(x_50, 1); +lean_inc(x_58); +lean_dec(x_50); +x_59 = lean_ctor_get(x_51, 0); +lean_inc(x_59); +lean_dec(x_51); +x_60 = lean_ctor_get(x_5, 2); +x_61 = lean_nat_add(x_7, x_60); +lean_dec(x_7); +x_6 = x_59; +x_7 = x_61; +x_8 = lean_box(0); +x_9 = lean_box(0); +x_14 = x_58; +goto _start; +} +} +else +{ +uint8_t x_63; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_7); +lean_dec(x_3); +x_63 = !lean_is_exclusive(x_50); +if (x_63 == 0) +{ +return x_50; +} +else +{ +lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_64 = lean_ctor_get(x_50, 0); +x_65 = lean_ctor_get(x_50, 1); +lean_inc(x_65); +lean_inc(x_64); +lean_dec(x_50); +x_66 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_66, 0, x_64); +lean_ctor_set(x_66, 1, x_65); +return x_66; +} +} +} +} +} +else +{ +lean_object* x_67; lean_object* x_68; uint8_t x_69; +x_67 = lean_ctor_get(x_6, 0); +x_68 = lean_ctor_get(x_6, 1); +lean_inc(x_68); +lean_inc(x_67); +lean_dec(x_6); +x_69 = l_List_elem___at_Lean_Meta_Occurrences_contains___spec__1(x_7, x_1); +if (x_69 == 0) +{ +lean_object* x_70; lean_object* x_71; lean_object* x_72; +x_70 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_70, 0, x_67); +lean_ctor_set(x_70, 1, x_68); +x_71 = lean_ctor_get(x_5, 2); +x_72 = lean_nat_add(x_7, x_71); +lean_dec(x_7); +x_6 = x_70; +x_7 = x_72; +x_8 = lean_box(0); +x_9 = lean_box(0); +goto _start; +} +else +{ +lean_object* x_74; uint8_t x_75; +x_74 = lean_array_fget(x_2, x_7); +x_75 = lean_unbox(x_67); +if (x_75 == 0) +{ +lean_object* x_76; lean_object* x_77; lean_object* x_78; uint8_t x_79; lean_object* x_80; +x_76 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___closed__3; +x_77 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_77, 0, x_68); +lean_ctor_set(x_77, 1, x_76); +x_78 = lean_box(0); +x_79 = lean_unbox(x_67); +lean_dec(x_67); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_3); +x_80 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___lambda__1(x_74, x_3, x_79, x_77, x_78, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_80) == 0) +{ +lean_object* x_81; +x_81 = lean_ctor_get(x_80, 0); +lean_inc(x_81); +if (lean_obj_tag(x_81) == 0) +{ +lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_7); +lean_dec(x_3); +x_82 = lean_ctor_get(x_80, 1); +lean_inc(x_82); +if (lean_is_exclusive(x_80)) { + lean_ctor_release(x_80, 0); + lean_ctor_release(x_80, 1); + x_83 = x_80; +} else { + lean_dec_ref(x_80); + x_83 = lean_box(0); +} +x_84 = lean_ctor_get(x_81, 0); +lean_inc(x_84); +lean_dec(x_81); +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; +x_86 = lean_ctor_get(x_80, 1); +lean_inc(x_86); +lean_dec(x_80); +x_87 = lean_ctor_get(x_81, 0); +lean_inc(x_87); +lean_dec(x_81); +x_88 = lean_ctor_get(x_5, 2); +x_89 = lean_nat_add(x_7, x_88); +lean_dec(x_7); +x_6 = x_87; +x_7 = x_89; +x_8 = lean_box(0); +x_9 = lean_box(0); +x_14 = x_86; +goto _start; +} +} +else +{ +lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_7); +lean_dec(x_3); +x_91 = lean_ctor_get(x_80, 0); +lean_inc(x_91); +x_92 = lean_ctor_get(x_80, 1); +lean_inc(x_92); +if (lean_is_exclusive(x_80)) { + lean_ctor_release(x_80, 0); + lean_ctor_release(x_80, 1); + x_93 = x_80; +} else { + lean_dec_ref(x_80); + x_93 = lean_box(0); +} +if (lean_is_scalar(x_93)) { + x_94 = lean_alloc_ctor(1, 2, 0); +} else { + x_94 = x_93; +} +lean_ctor_set(x_94, 0, x_91); +lean_ctor_set(x_94, 1, x_92); +return x_94; +} +} +else +{ +uint8_t x_95; lean_object* x_96; lean_object* x_97; +lean_dec(x_67); +x_95 = 0; +x_96 = lean_box(0); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_3); +x_97 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___lambda__1(x_74, x_3, x_95, x_68, x_96, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_97) == 0) +{ +lean_object* x_98; +x_98 = lean_ctor_get(x_97, 0); +lean_inc(x_98); +if (lean_obj_tag(x_98) == 0) +{ +lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_7); +lean_dec(x_3); +x_99 = lean_ctor_get(x_97, 1); +lean_inc(x_99); +if (lean_is_exclusive(x_97)) { + lean_ctor_release(x_97, 0); + lean_ctor_release(x_97, 1); + x_100 = x_97; +} else { + lean_dec_ref(x_97); + x_100 = lean_box(0); +} +x_101 = lean_ctor_get(x_98, 0); +lean_inc(x_101); +lean_dec(x_98); +if (lean_is_scalar(x_100)) { + x_102 = lean_alloc_ctor(0, 2, 0); +} else { + x_102 = x_100; +} +lean_ctor_set(x_102, 0, x_101); +lean_ctor_set(x_102, 1, x_99); +return x_102; +} +else +{ +lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; +x_103 = lean_ctor_get(x_97, 1); +lean_inc(x_103); +lean_dec(x_97); +x_104 = lean_ctor_get(x_98, 0); +lean_inc(x_104); +lean_dec(x_98); +x_105 = lean_ctor_get(x_5, 2); +x_106 = lean_nat_add(x_7, x_105); +lean_dec(x_7); +x_6 = x_104; +x_7 = x_106; +x_8 = lean_box(0); +x_9 = lean_box(0); +x_14 = x_103; +goto _start; +} +} +else +{ +lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_7); +lean_dec(x_3); +x_108 = lean_ctor_get(x_97, 0); +lean_inc(x_108); +x_109 = lean_ctor_get(x_97, 1); +lean_inc(x_109); +if (lean_is_exclusive(x_97)) { + lean_ctor_release(x_97, 0); + lean_ctor_release(x_97, 1); + x_110 = x_97; +} else { + lean_dec_ref(x_97); + x_110 = lean_box(0); +} +if (lean_is_scalar(x_110)) { + x_111 = lean_alloc_ctor(1, 2, 0); +} else { + x_111 = x_110; +} +lean_ctor_set(x_111, 0, x_108); +lean_ctor_set(x_111, 1, x_109); +return x_111; +} +} +} +} +} +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___lambda__1___closed__1() { +_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_Grind_ppPattern___closed__4; +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_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_9 = lean_array_get_size(x_2); +x_10 = lean_unsigned_to_nat(0u); +x_11 = lean_unsigned_to_nat(1u); +x_12 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_12, 0, x_10); +lean_ctor_set(x_12, 1, x_9); +lean_ctor_set(x_12, 2, x_11); +x_13 = l_Lean_Meta_Grind_ppPattern___closed__4; +x_14 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___lambda__1___closed__1; +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +x_15 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1(x_1, x_2, x_13, x_12, x_12, x_14, x_10, lean_box(0), lean_box(0), x_4, x_5, x_6, x_7, x_8); +lean_dec(x_12); +if (lean_obj_tag(x_15) == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_16 = lean_ctor_get(x_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, 1); +lean_inc(x_18); +lean_dec(x_16); +x_19 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_18, x_4, x_5, x_6, x_7, x_17); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_19; +} +else +{ +uint8_t x_20; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_20 = !lean_is_exclusive(x_15); +if (x_20 == 0) +{ +return x_15; +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_15, 0); +x_22 = lean_ctor_get(x_15, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_15); +x_23 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_23, 0, x_21); +lean_ctor_set(x_23, 1, x_22); +return x_23; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +x_9 = lean_infer_type(x_1, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_9) == 0) +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; lean_object* x_15; +x_10 = lean_ctor_get(x_9, 0); +lean_inc(x_10); +x_11 = lean_ctor_get(x_9, 1); +lean_inc(x_11); +lean_dec(x_9); +x_12 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_12, 0, x_2); +x_13 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___lambda__1___boxed), 8, 1); +lean_closure_set(x_13, 0, x_3); +x_14 = 0; +x_15 = l_Lean_Meta_forallBoundedTelescope___at_Lean_Meta_arrowDomainsN___spec__6___rarg(x_10, x_12, x_13, x_14, x_4, x_5, x_6, x_7, x_11); +return x_15; +} +else +{ +uint8_t x_16; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_16 = !lean_is_exclusive(x_9); +if (x_16 == 0) +{ +return x_9; +} +else +{ +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(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_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___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) { +_start: +{ +uint8_t x_11; lean_object* x_12; +x_11 = lean_unbox(x_3); +lean_dec(x_3); +x_12 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___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_5); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +lean_object* x_15; +x_15 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +return x_15; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; +x_9 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_9; +} +} +LEAN_EXPORT lean_object* l_panic___at_Lean_Meta_Grind_addEMatchTheorem___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; lean_object* x_9; +x_7 = l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__1___closed__1; +x_8 = lean_panic_fn(x_7, x_1); +x_9 = lean_apply_5(x_8, x_2, x_3, x_4, x_5, x_6); +return x_9; +} +} +static lean_object* _init_l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__2___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____spec__1___closed__2; +x_2 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_2, 0, x_1); +lean_ctor_set(x_2, 1, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__2___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____spec__1___closed__2; +x_2 = lean_alloc_ctor(0, 6, 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); +lean_ctor_set(x_2, 4, x_1); +lean_ctor_set(x_2, 5, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___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) { +_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_6, 6); +lean_inc(x_9); +lean_dec(x_6); +x_10 = lean_st_ref_take(x_7, 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_13 = !lean_is_exclusive(x_11); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; +x_14 = lean_ctor_get(x_11, 0); +x_15 = lean_ctor_get(x_11, 4); +lean_dec(x_15); +x_16 = l_Lean_ScopedEnvExtension_addCore___rarg(x_14, x_1, x_2, x_3, x_9); +x_17 = l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__2___closed__1; +lean_ctor_set(x_11, 4, x_17); +lean_ctor_set(x_11, 0, x_16); +x_18 = lean_st_ref_set(x_7, x_11, x_12); +x_19 = lean_ctor_get(x_18, 1); +lean_inc(x_19); +lean_dec(x_18); +x_20 = lean_st_ref_take(x_5, x_19); +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; uint8_t x_27; +x_24 = lean_ctor_get(x_21, 1); +lean_dec(x_24); +x_25 = l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__2___closed__2; +lean_ctor_set(x_21, 1, x_25); +x_26 = lean_st_ref_set(x_5, x_21, x_22); +x_27 = !lean_is_exclusive(x_26); +if (x_27 == 0) +{ +lean_object* x_28; lean_object* x_29; +x_28 = lean_ctor_get(x_26, 0); +lean_dec(x_28); +x_29 = lean_box(0); +lean_ctor_set(x_26, 0, x_29); +return x_26; +} +else +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_30 = lean_ctor_get(x_26, 1); +lean_inc(x_30); +lean_dec(x_26); +x_31 = lean_box(0); +x_32 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_30); +return x_32; +} +} +else +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_33 = lean_ctor_get(x_21, 0); +x_34 = lean_ctor_get(x_21, 2); +x_35 = lean_ctor_get(x_21, 3); +x_36 = lean_ctor_get(x_21, 4); +lean_inc(x_36); +lean_inc(x_35); +lean_inc(x_34); +lean_inc(x_33); +lean_dec(x_21); +x_37 = l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__2___closed__2; +x_38 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_38, 0, x_33); +lean_ctor_set(x_38, 1, x_37); +lean_ctor_set(x_38, 2, x_34); +lean_ctor_set(x_38, 3, x_35); +lean_ctor_set(x_38, 4, x_36); +x_39 = lean_st_ref_set(x_5, x_38, x_22); +x_40 = lean_ctor_get(x_39, 1); +lean_inc(x_40); +if (lean_is_exclusive(x_39)) { + lean_ctor_release(x_39, 0); + lean_ctor_release(x_39, 1); + x_41 = x_39; +} else { + lean_dec_ref(x_39); + x_41 = lean_box(0); +} +x_42 = lean_box(0); +if (lean_is_scalar(x_41)) { + x_43 = lean_alloc_ctor(0, 2, 0); +} else { + x_43 = x_41; +} +lean_ctor_set(x_43, 0, x_42); +lean_ctor_set(x_43, 1, x_40); +return x_43; +} +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; 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_44 = lean_ctor_get(x_11, 0); +x_45 = lean_ctor_get(x_11, 1); +x_46 = lean_ctor_get(x_11, 2); +x_47 = lean_ctor_get(x_11, 3); +x_48 = lean_ctor_get(x_11, 5); +x_49 = lean_ctor_get(x_11, 6); +x_50 = lean_ctor_get(x_11, 7); +lean_inc(x_50); +lean_inc(x_49); +lean_inc(x_48); +lean_inc(x_47); +lean_inc(x_46); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_11); +x_51 = l_Lean_ScopedEnvExtension_addCore___rarg(x_44, x_1, x_2, x_3, x_9); +x_52 = l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__2___closed__1; +x_53 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_53, 0, x_51); +lean_ctor_set(x_53, 1, x_45); +lean_ctor_set(x_53, 2, x_46); +lean_ctor_set(x_53, 3, x_47); +lean_ctor_set(x_53, 4, x_52); +lean_ctor_set(x_53, 5, x_48); +lean_ctor_set(x_53, 6, x_49); +lean_ctor_set(x_53, 7, x_50); +x_54 = lean_st_ref_set(x_7, x_53, x_12); +x_55 = lean_ctor_get(x_54, 1); +lean_inc(x_55); +lean_dec(x_54); +x_56 = lean_st_ref_take(x_5, x_55); +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_inc(x_59); +x_60 = lean_ctor_get(x_57, 2); +lean_inc(x_60); +x_61 = lean_ctor_get(x_57, 3); +lean_inc(x_61); +x_62 = lean_ctor_get(x_57, 4); +lean_inc(x_62); +if (lean_is_exclusive(x_57)) { + lean_ctor_release(x_57, 0); + lean_ctor_release(x_57, 1); + lean_ctor_release(x_57, 2); + lean_ctor_release(x_57, 3); + lean_ctor_release(x_57, 4); + x_63 = x_57; +} else { + lean_dec_ref(x_57); + x_63 = lean_box(0); +} +x_64 = l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__2___closed__2; +if (lean_is_scalar(x_63)) { + x_65 = lean_alloc_ctor(0, 5, 0); +} else { + x_65 = x_63; +} +lean_ctor_set(x_65, 0, x_59); +lean_ctor_set(x_65, 1, x_64); +lean_ctor_set(x_65, 2, x_60); +lean_ctor_set(x_65, 3, x_61); +lean_ctor_set(x_65, 4, x_62); +x_66 = lean_st_ref_set(x_5, x_65, x_58); +x_67 = lean_ctor_get(x_66, 1); +lean_inc(x_67); +if (lean_is_exclusive(x_66)) { + lean_ctor_release(x_66, 0); + lean_ctor_release(x_66, 1); + x_68 = x_66; +} else { + lean_dec_ref(x_66); + x_68 = lean_box(0); +} +x_69 = lean_box(0); +if (lean_is_scalar(x_68)) { + x_70 = lean_alloc_ctor(0, 2, 0); +} else { + x_70 = x_68; +} +lean_ctor_set(x_70, 0, x_69); +lean_ctor_set(x_70, 1, x_67); +return x_70; +} +} +} +LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__3(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_Meta_Grind_ppPattern(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_Meta_Grind_ppPattern(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_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__4(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; +x_5 = lean_ctor_get(x_1, 1); +lean_ctor_set(x_1, 1, x_2); +{ +lean_object* _tmp_0 = x_5; +lean_object* _tmp_1 = x_1; +x_1 = _tmp_0; +x_2 = _tmp_1; +} +goto _start; +} +else +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_7 = lean_ctor_get(x_1, 0); +x_8 = lean_ctor_get(x_1, 1); +lean_inc(x_8); +lean_inc(x_7); +lean_dec(x_1); +x_9 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_9, 0, x_7); +lean_ctor_set(x_9, 1, x_2); +x_1 = x_8; +x_2 = x_9; +goto _start; +} +} +} +} +LEAN_EXPORT uint8_t l_Lean_Meta_Grind_addEMatchTheorem___lambda__1(lean_object* x_1) { +_start: +{ +if (lean_obj_tag(x_1) == 2) +{ +uint8_t x_2; +x_2 = 1; +return x_2; +} +else +{ +uint8_t x_3; +x_3 = 0; +return x_3; +} +} +} +static lean_object* _init_l_Lean_Meta_Grind_addEMatchTheorem___lambda__2___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ematchTheoremsExt; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addEMatchTheorem___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; lean_object* x_18; +x_13 = lean_array_mk(x_1); +x_14 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_14, 0, x_2); +x_15 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_15, 0, x_13); +lean_ctor_set(x_15, 1, x_3); +lean_ctor_set(x_15, 2, x_4); +lean_ctor_set(x_15, 3, x_5); +lean_ctor_set(x_15, 4, x_6); +lean_ctor_set(x_15, 5, x_14); +x_16 = l_Lean_Meta_Grind_addEMatchTheorem___lambda__2___closed__1; +x_17 = 0; +x_18 = l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__2(x_16, x_15, x_17, x_8, x_9, x_10, x_11, x_12); +return x_18; +} +} +static lean_object* _init_l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("invalid pattern(s) for `", 24, 24); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("`", 1, 1); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__3; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("\nthe following theorem parameters cannot be instantiated:", 57, 57); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__5; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addEMatchTheorem___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +lean_object* x_14; +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_2); +lean_inc(x_1); +x_14 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage(x_1, x_2, x_3, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +if (lean_obj_tag(x_15) == 0) +{ +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_Meta_Grind_addEMatchTheorem___lambda__2(x_4, x_5, x_1, x_2, x_6, x_7, x_17, x_9, x_10, x_11, x_12, x_16); +lean_dec(x_12); +lean_dec(x_10); +lean_dec(x_9); +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_dec(x_7); +x_19 = lean_ctor_get(x_14, 1); +lean_inc(x_19); +lean_dec(x_14); +x_20 = lean_ctor_get(x_15, 0); +lean_inc(x_20); +lean_dec(x_15); +lean_inc(x_4); +x_21 = l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__3(x_6, x_4); +x_22 = l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__4(x_21, x_4); +x_23 = l_Lean_MessageData_ofList(x_22); +x_24 = l_Lean_Meta_Grind_ppPattern___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); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +x_27 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt(x_1, x_2, x_20, x_9, x_10, x_11, x_12, x_19); +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; lean_object* x_36; 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_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_MessageData_ofName(x_5); +x_31 = l_Lean_Meta_Grind_addEMatchTheorem___lambda__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_Meta_Grind_addEMatchTheorem___lambda__3___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_indentD(x_26); +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_Meta_Grind_addEMatchTheorem___lambda__3___closed__6; +x_38 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_38, 0, x_36); +lean_ctor_set(x_38, 1, x_37); +x_39 = l_Lean_indentD(x_28); +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_alloc_ctor(7, 2, 0); +lean_ctor_set(x_41, 0, x_40); +lean_ctor_set(x_41, 1, x_24); +x_42 = l_Lean_throwError___at___private_Lean_Meta_InferType_0__Lean_Meta_inferProjType___spec__1(x_41, x_9, x_10, x_11, x_12, x_29); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +x_43 = !lean_is_exclusive(x_42); +if (x_43 == 0) +{ +return x_42; +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_42, 0); +x_45 = lean_ctor_get(x_42, 1); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_42); +x_46 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_46, 0, x_44); +lean_ctor_set(x_46, 1, x_45); +return x_46; +} +} +else +{ +uint8_t x_47; +lean_dec(x_26); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_5); +x_47 = !lean_is_exclusive(x_27); +if (x_47 == 0) +{ +return x_27; +} +else +{ +lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_48 = lean_ctor_get(x_27, 0); +x_49 = lean_ctor_get(x_27, 1); +lean_inc(x_49); +lean_inc(x_48); +lean_dec(x_27); +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_51; +lean_dec(x_12); +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_2); +lean_dec(x_1); +x_51 = !lean_is_exclusive(x_14); +if (x_51 == 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_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; +} +} +} +} +static lean_object* _init_l_Lean_Meta_Grind_addEMatchTheorem___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("` is not a theorem, you cannot assign patterns to non-theorems for the `grind` tactic", 85, 85); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_addEMatchTheorem___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_addEMatchTheorem___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Grind_addEMatchTheorem___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_addEMatchTheorem___lambda__1___boxed), 1, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_addEMatchTheorem___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("symbols.all fun s => s matches .const _\n ", 42, 42); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_addEMatchTheorem___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__3; +x_2 = l_Lean_Meta_Grind_addEMatchTheorem___closed__4; +x_3 = lean_string_append(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind_addEMatchTheorem___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.Meta.Grind.addEMatchTheorem", 32, 32); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_addEMatchTheorem___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l_Lean_Meta_Grind_EMatchTheorems_insert___closed__1; +x_2 = l_Lean_Meta_Grind_addEMatchTheorem___closed__6; +x_3 = lean_unsigned_to_nat(340u); +x_4 = lean_unsigned_to_nat(2u); +x_5 = l_Lean_Meta_Grind_addEMatchTheorem___closed__5; +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_Meta_Grind_addEMatchTheorem___closed__8() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("ematch", 6, 6); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_addEMatchTheorem___closed__9() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("pattern", 7, 7); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_addEMatchTheorem___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Meta_Grind_mkGroundPattern___closed__1; +x_2 = l_Lean_Meta_Grind_addEMatchTheorem___closed__8; +x_3 = l_Lean_Meta_Grind_addEMatchTheorem___closed__9; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Meta_Grind_addEMatchTheorem___closed__11() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(": ", 2, 2); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_addEMatchTheorem___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_addEMatchTheorem___closed__11; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addEMatchTheorem(lean_object* x_1, lean_object* x_2, lean_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 = l_Lean_getConstInfo___at_Lean_Meta_mkConstWithFreshMVarLevels___spec__1(x_1, x_4, x_5, x_6, x_7, x_8); +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) == 2) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_11 = lean_ctor_get(x_9, 1); +lean_inc(x_11); +lean_dec(x_9); +x_12 = lean_ctor_get(x_10, 0); +lean_inc(x_12); +lean_dec(x_10); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +lean_dec(x_12); +x_14 = lean_ctor_get(x_13, 1); +lean_inc(x_14); +lean_dec(x_13); +x_15 = lean_box(0); +x_16 = l_List_mapTR_loop___at_Lean_mkConstWithLevelParams___spec__1(x_14, x_15); +lean_inc(x_1); +x_17 = l_Lean_Expr_const___override(x_1, x_16); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +x_18 = l_Lean_Meta_Grind_NormalizePattern_main(x_3, x_4, x_5, x_6, x_7, x_11); +if (lean_obj_tag(x_18) == 0) +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_19, 1); +lean_inc(x_20); +x_21 = lean_ctor_get(x_18, 1); +lean_inc(x_21); +lean_dec(x_18); +x_22 = !lean_is_exclusive(x_19); +if (x_22 == 0) +{ +lean_object* x_23; lean_object* x_24; uint8_t x_25; +x_23 = lean_ctor_get(x_19, 0); +x_24 = lean_ctor_get(x_19, 1); +lean_dec(x_24); +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); +x_27 = lean_ctor_get(x_20, 1); +x_28 = l_Lean_Meta_Grind_addEMatchTheorem___closed__3; +lean_inc(x_26); +x_29 = l_List_all___rarg(x_26, x_28); +if (x_29 == 0) +{ +lean_object* x_30; lean_object* x_31; +lean_free_object(x_20); +lean_dec(x_27); +lean_dec(x_26); +lean_free_object(x_19); +lean_dec(x_23); +lean_dec(x_17); +lean_dec(x_2); +lean_dec(x_1); +x_30 = l_Lean_Meta_Grind_addEMatchTheorem___closed__7; +x_31 = l_panic___at_Lean_Meta_Grind_addEMatchTheorem___spec__1(x_30, x_4, x_5, x_6, x_7, x_21); +return x_31; +} +else +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; +x_32 = l_Lean_Meta_Grind_addEMatchTheorem___closed__10; +x_33 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_32, x_4, x_5, x_6, x_7, x_21); +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_36; lean_object* x_37; lean_object* x_38; +lean_free_object(x_20); +lean_free_object(x_19); +x_36 = lean_ctor_get(x_33, 1); +lean_inc(x_36); +lean_dec(x_33); +x_37 = lean_box(0); +x_38 = l_Lean_Meta_Grind_addEMatchTheorem___lambda__3(x_17, x_2, x_27, x_15, x_1, x_23, x_26, x_37, x_4, x_5, x_6, x_7, x_36); +return x_38; +} +else +{ +uint8_t x_39; +x_39 = !lean_is_exclusive(x_33); +if (x_39 == 0) +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; +x_40 = lean_ctor_get(x_33, 1); +x_41 = lean_ctor_get(x_33, 0); +lean_dec(x_41); +lean_inc(x_1); +x_42 = l_Lean_MessageData_ofName(x_1); +x_43 = l_Lean_Meta_Grind_ppPattern___closed__4; +lean_ctor_set_tag(x_33, 7); +lean_ctor_set(x_33, 1, x_42); +lean_ctor_set(x_33, 0, x_43); +x_44 = l_Lean_Meta_Grind_addEMatchTheorem___closed__12; +lean_ctor_set_tag(x_20, 7); +lean_ctor_set(x_20, 1, x_44); +lean_ctor_set(x_20, 0, x_33); +lean_inc(x_23); +x_45 = l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__3(x_23, x_15); +x_46 = l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__4(x_45, x_15); +x_47 = l_Lean_MessageData_ofList(x_46); +lean_ctor_set_tag(x_19, 7); +lean_ctor_set(x_19, 1, x_47); +lean_ctor_set(x_19, 0, x_20); +x_48 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_48, 0, x_19); +lean_ctor_set(x_48, 1, x_43); +x_49 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_32, x_48, x_4, x_5, x_6, x_7, x_40); +x_50 = lean_ctor_get(x_49, 0); +lean_inc(x_50); +x_51 = lean_ctor_get(x_49, 1); +lean_inc(x_51); +lean_dec(x_49); +x_52 = l_Lean_Meta_Grind_addEMatchTheorem___lambda__3(x_17, x_2, x_27, x_15, x_1, x_23, x_26, x_50, x_4, x_5, x_6, x_7, x_51); +lean_dec(x_50); +return x_52; +} +else +{ +lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; +x_53 = lean_ctor_get(x_33, 1); +lean_inc(x_53); +lean_dec(x_33); +lean_inc(x_1); +x_54 = l_Lean_MessageData_ofName(x_1); +x_55 = l_Lean_Meta_Grind_ppPattern___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 = l_Lean_Meta_Grind_addEMatchTheorem___closed__12; +lean_ctor_set_tag(x_20, 7); +lean_ctor_set(x_20, 1, x_57); +lean_ctor_set(x_20, 0, x_56); +lean_inc(x_23); +x_58 = l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__3(x_23, x_15); +x_59 = l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__4(x_58, x_15); +x_60 = l_Lean_MessageData_ofList(x_59); +lean_ctor_set_tag(x_19, 7); +lean_ctor_set(x_19, 1, x_60); +lean_ctor_set(x_19, 0, x_20); +x_61 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_61, 0, x_19); +lean_ctor_set(x_61, 1, x_55); +x_62 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_32, x_61, x_4, x_5, x_6, x_7, x_53); +x_63 = lean_ctor_get(x_62, 0); +lean_inc(x_63); +x_64 = lean_ctor_get(x_62, 1); +lean_inc(x_64); +lean_dec(x_62); +x_65 = l_Lean_Meta_Grind_addEMatchTheorem___lambda__3(x_17, x_2, x_27, x_15, x_1, x_23, x_26, x_63, x_4, x_5, x_6, x_7, x_64); +lean_dec(x_63); +return x_65; +} +} +} +} +else +{ +lean_object* x_66; lean_object* x_67; lean_object* x_68; uint8_t x_69; +x_66 = lean_ctor_get(x_20, 0); +x_67 = lean_ctor_get(x_20, 1); +lean_inc(x_67); +lean_inc(x_66); +lean_dec(x_20); +x_68 = l_Lean_Meta_Grind_addEMatchTheorem___closed__3; +lean_inc(x_66); +x_69 = l_List_all___rarg(x_66, x_68); +if (x_69 == 0) +{ +lean_object* x_70; lean_object* x_71; +lean_dec(x_67); +lean_dec(x_66); +lean_free_object(x_19); +lean_dec(x_23); +lean_dec(x_17); +lean_dec(x_2); +lean_dec(x_1); +x_70 = l_Lean_Meta_Grind_addEMatchTheorem___closed__7; +x_71 = l_panic___at_Lean_Meta_Grind_addEMatchTheorem___spec__1(x_70, x_4, x_5, x_6, x_7, x_21); +return x_71; +} +else +{ +lean_object* x_72; lean_object* x_73; lean_object* x_74; uint8_t x_75; +x_72 = l_Lean_Meta_Grind_addEMatchTheorem___closed__10; +x_73 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_72, x_4, x_5, x_6, x_7, x_21); +x_74 = lean_ctor_get(x_73, 0); +lean_inc(x_74); +x_75 = lean_unbox(x_74); +lean_dec(x_74); +if (x_75 == 0) +{ +lean_object* x_76; lean_object* x_77; lean_object* x_78; +lean_free_object(x_19); +x_76 = lean_ctor_get(x_73, 1); +lean_inc(x_76); +lean_dec(x_73); +x_77 = lean_box(0); +x_78 = l_Lean_Meta_Grind_addEMatchTheorem___lambda__3(x_17, x_2, x_67, x_15, x_1, x_23, x_66, x_77, x_4, x_5, x_6, x_7, x_76); +return x_78; +} +else +{ +lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; 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_79 = lean_ctor_get(x_73, 1); +lean_inc(x_79); +if (lean_is_exclusive(x_73)) { + lean_ctor_release(x_73, 0); + lean_ctor_release(x_73, 1); + x_80 = x_73; +} else { + lean_dec_ref(x_73); + x_80 = lean_box(0); +} +lean_inc(x_1); +x_81 = l_Lean_MessageData_ofName(x_1); +x_82 = l_Lean_Meta_Grind_ppPattern___closed__4; +if (lean_is_scalar(x_80)) { + x_83 = lean_alloc_ctor(7, 2, 0); +} else { + x_83 = x_80; + lean_ctor_set_tag(x_83, 7); +} +lean_ctor_set(x_83, 0, x_82); +lean_ctor_set(x_83, 1, x_81); +x_84 = l_Lean_Meta_Grind_addEMatchTheorem___closed__12; +x_85 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_85, 0, x_83); +lean_ctor_set(x_85, 1, x_84); +lean_inc(x_23); +x_86 = l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__3(x_23, x_15); +x_87 = l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__4(x_86, x_15); +x_88 = l_Lean_MessageData_ofList(x_87); +lean_ctor_set_tag(x_19, 7); +lean_ctor_set(x_19, 1, x_88); +lean_ctor_set(x_19, 0, x_85); +x_89 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_89, 0, x_19); +lean_ctor_set(x_89, 1, x_82); +x_90 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_72, x_89, x_4, x_5, x_6, x_7, x_79); +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 = l_Lean_Meta_Grind_addEMatchTheorem___lambda__3(x_17, x_2, x_67, x_15, x_1, x_23, x_66, x_91, x_4, x_5, x_6, x_7, x_92); +lean_dec(x_91); +return x_93; +} +} +} +} +else +{ +lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; uint8_t x_99; +x_94 = lean_ctor_get(x_19, 0); +lean_inc(x_94); +lean_dec(x_19); +x_95 = lean_ctor_get(x_20, 0); +lean_inc(x_95); +x_96 = lean_ctor_get(x_20, 1); +lean_inc(x_96); +if (lean_is_exclusive(x_20)) { + lean_ctor_release(x_20, 0); + lean_ctor_release(x_20, 1); + x_97 = x_20; +} else { + lean_dec_ref(x_20); + x_97 = lean_box(0); +} +x_98 = l_Lean_Meta_Grind_addEMatchTheorem___closed__3; +lean_inc(x_95); +x_99 = l_List_all___rarg(x_95, x_98); +if (x_99 == 0) +{ +lean_object* x_100; lean_object* x_101; +lean_dec(x_97); +lean_dec(x_96); +lean_dec(x_95); +lean_dec(x_94); +lean_dec(x_17); +lean_dec(x_2); +lean_dec(x_1); +x_100 = l_Lean_Meta_Grind_addEMatchTheorem___closed__7; +x_101 = l_panic___at_Lean_Meta_Grind_addEMatchTheorem___spec__1(x_100, x_4, x_5, x_6, x_7, x_21); +return x_101; +} +else +{ +lean_object* x_102; lean_object* x_103; lean_object* x_104; uint8_t x_105; +x_102 = l_Lean_Meta_Grind_addEMatchTheorem___closed__10; +x_103 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_102, x_4, x_5, x_6, x_7, x_21); +x_104 = lean_ctor_get(x_103, 0); +lean_inc(x_104); +x_105 = lean_unbox(x_104); +lean_dec(x_104); +if (x_105 == 0) +{ +lean_object* x_106; lean_object* x_107; lean_object* x_108; +lean_dec(x_97); +x_106 = lean_ctor_get(x_103, 1); +lean_inc(x_106); +lean_dec(x_103); +x_107 = lean_box(0); +x_108 = l_Lean_Meta_Grind_addEMatchTheorem___lambda__3(x_17, x_2, x_96, x_15, x_1, x_94, x_95, x_107, x_4, x_5, x_6, x_7, x_106); +return x_108; +} +else +{ +lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; 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_ctor_get(x_103, 1); +lean_inc(x_109); +if (lean_is_exclusive(x_103)) { + lean_ctor_release(x_103, 0); + lean_ctor_release(x_103, 1); + x_110 = x_103; +} else { + lean_dec_ref(x_103); + x_110 = lean_box(0); +} +lean_inc(x_1); +x_111 = l_Lean_MessageData_ofName(x_1); +x_112 = l_Lean_Meta_Grind_ppPattern___closed__4; +if (lean_is_scalar(x_110)) { + x_113 = lean_alloc_ctor(7, 2, 0); +} else { + x_113 = x_110; + lean_ctor_set_tag(x_113, 7); +} +lean_ctor_set(x_113, 0, x_112); +lean_ctor_set(x_113, 1, x_111); +x_114 = l_Lean_Meta_Grind_addEMatchTheorem___closed__12; +if (lean_is_scalar(x_97)) { + x_115 = lean_alloc_ctor(7, 2, 0); +} else { + x_115 = x_97; + lean_ctor_set_tag(x_115, 7); +} +lean_ctor_set(x_115, 0, x_113); +lean_ctor_set(x_115, 1, x_114); +lean_inc(x_94); +x_116 = l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__3(x_94, x_15); +x_117 = l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__4(x_116, x_15); +x_118 = l_Lean_MessageData_ofList(x_117); +x_119 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_119, 0, x_115); +lean_ctor_set(x_119, 1, x_118); +x_120 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_120, 0, x_119); +lean_ctor_set(x_120, 1, x_112); +x_121 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_102, x_120, x_4, x_5, x_6, x_7, x_109); +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 = l_Lean_Meta_Grind_addEMatchTheorem___lambda__3(x_17, x_2, x_96, x_15, x_1, x_94, x_95, x_122, x_4, x_5, x_6, x_7, x_123); +lean_dec(x_122); +return x_124; +} +} +} +} +else +{ +uint8_t x_125; +lean_dec(x_17); +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_125 = !lean_is_exclusive(x_18); +if (x_125 == 0) +{ +return x_18; +} +else +{ +lean_object* x_126; lean_object* x_127; lean_object* x_128; +x_126 = lean_ctor_get(x_18, 0); +x_127 = lean_ctor_get(x_18, 1); +lean_inc(x_127); +lean_inc(x_126); +lean_dec(x_18); +x_128 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_128, 0, x_126); +lean_ctor_set(x_128, 1, x_127); +return x_128; +} +} +} +else +{ +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_dec(x_10); +lean_dec(x_3); +lean_dec(x_2); +x_129 = lean_ctor_get(x_9, 1); +lean_inc(x_129); +lean_dec(x_9); +x_130 = l_Lean_MessageData_ofName(x_1); +x_131 = l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__4; +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_Meta_Grind_addEMatchTheorem___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_throwError___at_Lean_Meta_setInlineAttribute___spec__1(x_134, x_4, x_5, x_6, x_7, x_129); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_135; +} +} +else +{ +uint8_t x_136; +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_136 = !lean_is_exclusive(x_9); +if (x_136 == 0) +{ +return x_9; +} +else +{ +lean_object* x_137; lean_object* x_138; lean_object* x_139; +x_137 = lean_ctor_get(x_9, 0); +x_138 = lean_ctor_get(x_9, 1); +lean_inc(x_138); +lean_inc(x_137); +lean_dec(x_9); +x_139 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_139, 0, x_137); +lean_ctor_set(x_139, 1, x_138); +return x_139; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___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: +{ +uint8_t x_9; lean_object* x_10; +x_9 = lean_unbox(x_3); +lean_dec(x_3); +x_10 = l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___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_5); +lean_dec(x_4); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addEMatchTheorem___lambda__1___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = l_Lean_Meta_Grind_addEMatchTheorem___lambda__1(x_1); +lean_dec(x_1); +x_3 = lean_box(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addEMatchTheorem___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +x_13 = l_Lean_Meta_Grind_addEMatchTheorem___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_11); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addEMatchTheorem___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: +{ +lean_object* x_14; +x_14 = l_Lean_Meta_Grind_addEMatchTheorem___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_8); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getEMatchTheorems___rarg(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; uint8_t x_4; +x_3 = lean_st_ref_get(x_1, x_2); +x_4 = !lean_is_exclusive(x_3); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_5 = lean_ctor_get(x_3, 0); +x_6 = lean_ctor_get(x_5, 0); +lean_inc(x_6); +lean_dec(x_5); +x_7 = l_panic___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__1___closed__1; +x_8 = l_Lean_Meta_Grind_addEMatchTheorem___lambda__2___closed__1; +x_9 = l_Lean_ScopedEnvExtension_getState___rarg(x_7, x_8, x_6); +lean_dec(x_6); +lean_ctor_set(x_3, 0, x_9); +return x_3; +} +else +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_10 = lean_ctor_get(x_3, 0); +x_11 = lean_ctor_get(x_3, 1); +lean_inc(x_11); +lean_inc(x_10); +lean_dec(x_3); +x_12 = lean_ctor_get(x_10, 0); +lean_inc(x_12); +lean_dec(x_10); +x_13 = l_panic___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__1___closed__1; +x_14 = l_Lean_Meta_Grind_addEMatchTheorem___lambda__2___closed__1; +x_15 = l_Lean_ScopedEnvExtension_getState___rarg(x_13, x_14, x_12); +lean_dec(x_12); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_15); +lean_ctor_set(x_16, 1, x_11); +return x_16; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getEMatchTheorems(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_getEMatchTheorems___rarg___boxed), 2, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getEMatchTheorems___rarg___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_Meta_Grind_getEMatchTheorems___rarg(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getEMatchTheorems___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Lean_Meta_Grind_getEMatchTheorems(x_1); +lean_dec(x_1); +return x_2; +} +} +lean_object* initialize_Lean_HeadIndex(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_PrettyPrinter(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Util_FoldConsts(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Util_CollectFVars(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Meta_Basic(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Meta_InferType(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Lean_Meta_Tactic_Grind_EMatchTheorem(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_HeadIndex(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +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_Util_FoldConsts(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Lean_Util_CollectFVars(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_InferType(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l_Lean_Meta_Grind_instInhabitedOrigin___closed__1 = _init_l_Lean_Meta_Grind_instInhabitedOrigin___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_instInhabitedOrigin___closed__1); +l_Lean_Meta_Grind_instInhabitedOrigin = _init_l_Lean_Meta_Grind_instInhabitedOrigin(); +lean_mark_persistent(l_Lean_Meta_Grind_instInhabitedOrigin); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__1); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__2); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__3); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__4(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__4); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__5 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__5(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__5); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__6 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__6(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__6); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__7 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__7(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__7); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__8 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__8(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__8); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__9 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__9(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__9); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__10 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__10(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__10); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__11 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__11(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__11); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__12 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__12(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__12); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__13 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__13(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__13); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__14 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__14(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__14); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__15 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__15(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__15); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__16 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__16(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__16); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__17 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__17(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__17); +l_Lean_Meta_Grind_instReprOrigin___closed__1 = _init_l_Lean_Meta_Grind_instReprOrigin___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_instReprOrigin___closed__1); +l_Lean_Meta_Grind_instReprOrigin = _init_l_Lean_Meta_Grind_instReprOrigin(); +lean_mark_persistent(l_Lean_Meta_Grind_instReprOrigin); +l_Lean_Meta_Grind_Origin_key___closed__1 = _init_l_Lean_Meta_Grind_Origin_key___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_Origin_key___closed__1); +l_Lean_Meta_Grind_Origin_key___closed__2 = _init_l_Lean_Meta_Grind_Origin_key___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_Origin_key___closed__2); +l_Lean_Meta_Grind_Origin_pp___rarg___closed__1 = _init_l_Lean_Meta_Grind_Origin_pp___rarg___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_Origin_pp___rarg___closed__1); +l_Lean_Meta_Grind_Origin_pp___rarg___closed__2 = _init_l_Lean_Meta_Grind_Origin_pp___rarg___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_Origin_pp___rarg___closed__2); +l_Lean_Meta_Grind_Origin_pp___rarg___closed__3 = _init_l_Lean_Meta_Grind_Origin_pp___rarg___closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind_Origin_pp___rarg___closed__3); +l_Lean_Meta_Grind_instInhabitedEMatchTheorem___closed__1 = _init_l_Lean_Meta_Grind_instInhabitedEMatchTheorem___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_instInhabitedEMatchTheorem___closed__1); +l_Lean_Meta_Grind_instInhabitedEMatchTheorem___closed__2 = _init_l_Lean_Meta_Grind_instInhabitedEMatchTheorem___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_instInhabitedEMatchTheorem___closed__2); +l_Lean_Meta_Grind_instInhabitedEMatchTheorem___closed__3 = _init_l_Lean_Meta_Grind_instInhabitedEMatchTheorem___closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind_instInhabitedEMatchTheorem___closed__3); +l_Lean_Meta_Grind_instInhabitedEMatchTheorem = _init_l_Lean_Meta_Grind_instInhabitedEMatchTheorem(); +lean_mark_persistent(l_Lean_Meta_Grind_instInhabitedEMatchTheorem); +l_panic___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__1___closed__1 = _init_l_panic___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__1___closed__1(); +lean_mark_persistent(l_panic___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__1___closed__1); +l_panic___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__1___closed__2 = _init_l_panic___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__1___closed__2(); +lean_mark_persistent(l_panic___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__1___closed__2); +l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__1 = _init_l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__1(); +l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__2 = _init_l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__2(); +l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__6___closed__1 = _init_l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__6___closed__1(); +lean_mark_persistent(l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__6___closed__1); +l_Lean_Meta_Grind_EMatchTheorems_insert___closed__1 = _init_l_Lean_Meta_Grind_EMatchTheorems_insert___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_EMatchTheorems_insert___closed__1); +l_Lean_Meta_Grind_EMatchTheorems_insert___closed__2 = _init_l_Lean_Meta_Grind_EMatchTheorems_insert___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_EMatchTheorems_insert___closed__2); +l_Lean_Meta_Grind_EMatchTheorems_insert___closed__3 = _init_l_Lean_Meta_Grind_EMatchTheorems_insert___closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind_EMatchTheorems_insert___closed__3); +l_Lean_Meta_Grind_EMatchTheorems_insert___closed__4 = _init_l_Lean_Meta_Grind_EMatchTheorems_insert___closed__4(); +lean_mark_persistent(l_Lean_Meta_Grind_EMatchTheorems_insert___closed__4); +l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____spec__1___closed__1 = _init_l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____spec__1___closed__1(); +lean_mark_persistent(l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____spec__1___closed__1); +l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____spec__1___closed__2 = _init_l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____spec__1___closed__2(); +lean_mark_persistent(l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____spec__1___closed__2); +l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____spec__1 = _init_l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____spec__1(); +lean_mark_persistent(l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____spec__1); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__1 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__1); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__2 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__2); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__3 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__3); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__4 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__4(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__4); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__5 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__5(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__5); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__6 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__6(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__6); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__7 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__7(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__7); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__8 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__8(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__8); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__9 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__9(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__9); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__10 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__10(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__10); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__11 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__11(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__11); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__12 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__12(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__12); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__13 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__13(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__13); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__14 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__14(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__14); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__15 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__15(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__15); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__16 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__16(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__16); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__17 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__17(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__17); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__18 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__18(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__18); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__19 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__19(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__19); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__20 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__20(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__20); +if (builtin) {res = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687_(lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ematchTheoremsExt = lean_io_result_get_value(res); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ematchTheoremsExt); +lean_dec_ref(res); +}l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__1); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__2); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__3); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__4(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__4); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__5 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__5(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__5); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__6 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__6(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__6); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__7 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__7(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__7); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__8 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__8(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__8); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__9 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__9(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__9); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__10 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__10(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__10); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__11 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__11(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__11); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__12 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__12(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__12); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__13 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__13(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__13); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__14 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__14(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__14); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__15 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__15(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__15); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__16 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__16(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__16); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__17 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__17(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__17); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__18 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__18(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__18); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__19 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__19(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__19); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare___closed__1); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare___closed__2); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare___closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare___closed__3); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare); +l_Lean_Meta_Grind_mkGroundPattern___closed__1 = _init_l_Lean_Meta_Grind_mkGroundPattern___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_mkGroundPattern___closed__1); +l_Lean_Meta_Grind_mkGroundPattern___closed__2 = _init_l_Lean_Meta_Grind_mkGroundPattern___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_mkGroundPattern___closed__2); +l_Lean_Meta_Grind_mkGroundPattern___closed__3 = _init_l_Lean_Meta_Grind_mkGroundPattern___closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind_mkGroundPattern___closed__3); +l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___closed__1 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___closed__1(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___closed__1); +l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___closed__2 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___closed__2(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___closed__2); +l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___closed__3 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___closed__3(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___closed__3); +l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1); +l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2); +l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3); +l_Lean_Meta_Grind_ppPattern___closed__1 = _init_l_Lean_Meta_Grind_ppPattern___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_ppPattern___closed__1); +l_Lean_Meta_Grind_ppPattern___closed__2 = _init_l_Lean_Meta_Grind_ppPattern___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_ppPattern___closed__2); +l_Lean_Meta_Grind_ppPattern___closed__3 = _init_l_Lean_Meta_Grind_ppPattern___closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind_ppPattern___closed__3); +l_Lean_Meta_Grind_ppPattern___closed__4 = _init_l_Lean_Meta_Grind_ppPattern___closed__4(); +lean_mark_persistent(l_Lean_Meta_Grind_ppPattern___closed__4); +l_Lean_Meta_Grind_ppPattern___closed__5 = _init_l_Lean_Meta_Grind_ppPattern___closed__5(); +lean_mark_persistent(l_Lean_Meta_Grind_ppPattern___closed__5); +l_Lean_Meta_Grind_ppPattern___closed__6 = _init_l_Lean_Meta_Grind_ppPattern___closed__6(); +lean_mark_persistent(l_Lean_Meta_Grind_ppPattern___closed__6); +l_Lean_Meta_Grind_ppPattern___closed__7 = _init_l_Lean_Meta_Grind_ppPattern___closed__7(); +lean_mark_persistent(l_Lean_Meta_Grind_ppPattern___closed__7); +l_Lean_Meta_Grind_ppPattern___closed__8 = _init_l_Lean_Meta_Grind_ppPattern___closed__8(); +lean_mark_persistent(l_Lean_Meta_Grind_ppPattern___closed__8); +l_Lean_Meta_Grind_ppPattern___closed__9 = _init_l_Lean_Meta_Grind_ppPattern___closed__9(); +lean_mark_persistent(l_Lean_Meta_Grind_ppPattern___closed__9); +l_Lean_Meta_Grind_ppPattern___closed__10 = _init_l_Lean_Meta_Grind_ppPattern___closed__10(); +lean_mark_persistent(l_Lean_Meta_Grind_ppPattern___closed__10); +l_Lean_Meta_Grind_ppPattern___closed__11 = _init_l_Lean_Meta_Grind_ppPattern___closed__11(); +lean_mark_persistent(l_Lean_Meta_Grind_ppPattern___closed__11); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFunMask___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFunMask___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFunMask___closed__1); +l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__3___closed__1 = _init_l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__3___closed__1(); +lean_mark_persistent(l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__3___closed__1); +l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__3___closed__2 = _init_l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__3___closed__2(); +lean_mark_persistent(l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__3___closed__2); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__1); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__2); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__3); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__4(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__4); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__5 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__5(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__5); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__6 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__6(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__6); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__7 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__7(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__7); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___closed__1); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___closed__2); +l_Lean_Meta_Grind_NormalizePattern_main___closed__1 = _init_l_Lean_Meta_Grind_NormalizePattern_main___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_NormalizePattern_main___closed__1); +l_Lean_Meta_Grind_NormalizePattern_main___closed__2 = _init_l_Lean_Meta_Grind_NormalizePattern_main___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_NormalizePattern_main___closed__2); +l_Lean_Meta_Grind_NormalizePattern_main___closed__3 = _init_l_Lean_Meta_Grind_NormalizePattern_main___closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind_NormalizePattern_main___closed__3); +l_Lean_Meta_Grind_NormalizePattern_main___closed__4 = _init_l_Lean_Meta_Grind_NormalizePattern_main___closed__4(); +lean_mark_persistent(l_Lean_Meta_Grind_NormalizePattern_main___closed__4); +l_Lean_Meta_Grind_NormalizePattern_main___closed__5 = _init_l_Lean_Meta_Grind_NormalizePattern_main___closed__5(); +lean_mark_persistent(l_Lean_Meta_Grind_NormalizePattern_main___closed__5); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___closed__1); +l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__1___closed__1 = _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__1___closed__1(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__1___closed__1); +l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__1___closed__2 = _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__1___closed__2(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__1___closed__2); +l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__1 = _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__1(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__1); +l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__2 = _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__2(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__2); +l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__3 = _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__3(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__3); +l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__4 = _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__4(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__4); +l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__2___closed__1 = _init_l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__2___closed__1(); +lean_mark_persistent(l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__2___closed__1); +l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___closed__1 = _init_l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___closed__1(); +lean_mark_persistent(l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___closed__1); +l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__1___closed__1 = _init_l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__1___closed__1(); +lean_mark_persistent(l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__1___closed__1); +l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___closed__1 = _init_l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___closed__1(); +lean_mark_persistent(l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___closed__1); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3___closed__1); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3___closed__2); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3___closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3___closed__3); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3___closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3___closed__4(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3___closed__4); +l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___lambda__1___closed__1 = _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___lambda__1___closed__1(); +lean_mark_persistent(l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___lambda__1___closed__1); +l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___lambda__1___closed__2 = _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___lambda__1___closed__2(); +lean_mark_persistent(l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___lambda__1___closed__2); +l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___closed__1 = _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___closed__1(); +lean_mark_persistent(l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___closed__1); +l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___closed__2 = _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___closed__2(); +lean_mark_persistent(l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___closed__2); +l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___closed__3 = _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___closed__3(); +lean_mark_persistent(l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___closed__3); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___lambda__1___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___lambda__1___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___lambda__1___closed__1); +l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__2___closed__1 = _init_l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__2___closed__1(); +lean_mark_persistent(l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__2___closed__1); +l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__2___closed__2 = _init_l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__2___closed__2(); +lean_mark_persistent(l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__2___closed__2); +l_Lean_Meta_Grind_addEMatchTheorem___lambda__2___closed__1 = _init_l_Lean_Meta_Grind_addEMatchTheorem___lambda__2___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_addEMatchTheorem___lambda__2___closed__1); +l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__1 = _init_l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__1); +l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__2 = _init_l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__2); +l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__3 = _init_l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__3); +l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__4 = _init_l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__4(); +lean_mark_persistent(l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__4); +l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__5 = _init_l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__5(); +lean_mark_persistent(l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__5); +l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__6 = _init_l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__6(); +lean_mark_persistent(l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__6); +l_Lean_Meta_Grind_addEMatchTheorem___closed__1 = _init_l_Lean_Meta_Grind_addEMatchTheorem___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_addEMatchTheorem___closed__1); +l_Lean_Meta_Grind_addEMatchTheorem___closed__2 = _init_l_Lean_Meta_Grind_addEMatchTheorem___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_addEMatchTheorem___closed__2); +l_Lean_Meta_Grind_addEMatchTheorem___closed__3 = _init_l_Lean_Meta_Grind_addEMatchTheorem___closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind_addEMatchTheorem___closed__3); +l_Lean_Meta_Grind_addEMatchTheorem___closed__4 = _init_l_Lean_Meta_Grind_addEMatchTheorem___closed__4(); +lean_mark_persistent(l_Lean_Meta_Grind_addEMatchTheorem___closed__4); +l_Lean_Meta_Grind_addEMatchTheorem___closed__5 = _init_l_Lean_Meta_Grind_addEMatchTheorem___closed__5(); +lean_mark_persistent(l_Lean_Meta_Grind_addEMatchTheorem___closed__5); +l_Lean_Meta_Grind_addEMatchTheorem___closed__6 = _init_l_Lean_Meta_Grind_addEMatchTheorem___closed__6(); +lean_mark_persistent(l_Lean_Meta_Grind_addEMatchTheorem___closed__6); +l_Lean_Meta_Grind_addEMatchTheorem___closed__7 = _init_l_Lean_Meta_Grind_addEMatchTheorem___closed__7(); +lean_mark_persistent(l_Lean_Meta_Grind_addEMatchTheorem___closed__7); +l_Lean_Meta_Grind_addEMatchTheorem___closed__8 = _init_l_Lean_Meta_Grind_addEMatchTheorem___closed__8(); +lean_mark_persistent(l_Lean_Meta_Grind_addEMatchTheorem___closed__8); +l_Lean_Meta_Grind_addEMatchTheorem___closed__9 = _init_l_Lean_Meta_Grind_addEMatchTheorem___closed__9(); +lean_mark_persistent(l_Lean_Meta_Grind_addEMatchTheorem___closed__9); +l_Lean_Meta_Grind_addEMatchTheorem___closed__10 = _init_l_Lean_Meta_Grind_addEMatchTheorem___closed__10(); +lean_mark_persistent(l_Lean_Meta_Grind_addEMatchTheorem___closed__10); +l_Lean_Meta_Grind_addEMatchTheorem___closed__11 = _init_l_Lean_Meta_Grind_addEMatchTheorem___closed__11(); +lean_mark_persistent(l_Lean_Meta_Grind_addEMatchTheorem___closed__11); +l_Lean_Meta_Grind_addEMatchTheorem___closed__12 = _init_l_Lean_Meta_Grind_addEMatchTheorem___closed__12(); +lean_mark_persistent(l_Lean_Meta_Grind_addEMatchTheorem___closed__12); +return lean_io_result_mk_ok(lean_box(0)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Lean/Meta/Tactic/Grind/ForallProp.c b/stage0/stdlib/Lean/Meta/Tactic/Grind/ForallProp.c new file mode 100644 index 000000000000..65165f8e3b54 --- /dev/null +++ b/stage0/stdlib/Lean/Meta/Tactic/Grind/ForallProp.c @@ -0,0 +1,574 @@ +// Lean compiler output +// Module: Lean.Meta.Tactic.Grind.ForallProp +// Imports: Init.Grind.Lemmas Lean.Meta.Tactic.Grind.Types Lean.Meta.Tactic.Grind.Internalize Lean.Meta.Tactic.Grind.Simp +#include +#if defined(__clang__) +#pragma clang diagnostic ignored "-Wunused-parameter" +#pragma clang diagnostic ignored "-Wunused-label" +#elif defined(__GNUC__) && !defined(__CLANG__) +#pragma GCC diagnostic ignored "-Wunused-parameter" +#pragma GCC diagnostic ignored "-Wunused-label" +#pragma GCC diagnostic ignored "-Wunused-but-set-variable" +#endif +#ifdef __cplusplus +extern "C" { +#endif +lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); +lean_object* l_Lean_Meta_Simp_Result_getProof(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_mkAppB(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_Grind_pushEqCore(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__6; +static lean_object* l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__2; +static lean_object* l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__8; +lean_object* l_Lean_Meta_Grind_simp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__5; +static lean_object* l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__3; +static lean_object* l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__1; +lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); +lean_object* l_Lean_Meta_Grind_getGeneration(lean_object*, 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_Grind_propagateForallProp___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__4; +lean_object* l_Lean_Meta_Grind_mkEqTrueProof(lean_object*, 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_Grind_propagateForallProp(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_Grind_isEqTrue(lean_object*, 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_Grind_propagateForallProp___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_Expr_lam___override(lean_object*, lean_object*, lean_object*, uint8_t); +lean_object* l_Lean_mkApp5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_expr_instantiate1(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__7; +lean_object* l_Lean_Meta_Grind_internalize(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* _init_l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("of_eq_true", 10, 10); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__2; +x_3 = l_Lean_Expr_const___override(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean", 4, 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Grind", 5, 5); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("forall_propagator", 17, 17); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__4; +x_2 = l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__5; +x_3 = l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__6; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__7; +x_3 = l_Lean_Expr_const___override(x_2, x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateForallProp___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +_start: +{ +lean_object* x_16; +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_1); +x_16 = l_Lean_Meta_Grind_mkEqTrueProof(x_1, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +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_Meta_Grind_propagateForallProp___lambda__1___closed__3; +lean_inc(x_17); +lean_inc(x_1); +x_20 = l_Lean_mkAppB(x_19, x_1, x_17); +x_21 = lean_expr_instantiate1(x_2, x_20); +lean_dec(x_20); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_9); +x_22 = l_Lean_Meta_Grind_simp(x_21, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_18); +if (lean_obj_tag(x_22) == 0) +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +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_1); +x_25 = l_Lean_Expr_lam___override(x_3, x_1, x_2, x_4); +x_26 = lean_ctor_get(x_23, 0); +lean_inc(x_26); +lean_inc(x_5); +x_27 = l_Lean_Meta_Grind_getGeneration(x_5, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_24); +if (lean_obj_tag(x_27) == 0) +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_28 = lean_ctor_get(x_27, 0); +lean_inc(x_28); +x_29 = lean_ctor_get(x_27, 1); +lean_inc(x_29); +lean_dec(x_27); +lean_inc(x_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_26); +x_30 = l_Lean_Meta_Grind_internalize(x_26, x_28, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_29); +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); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +x_32 = l_Lean_Meta_Simp_Result_getProof(x_23, x_11, x_12, x_13, x_14, x_31); +if (lean_obj_tag(x_32) == 0) +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; lean_object* x_38; +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_Grind_propagateForallProp___lambda__1___closed__8; +lean_inc(x_26); +x_36 = l_Lean_mkApp5(x_35, x_1, x_25, x_26, x_17, x_33); +x_37 = 0; +x_38 = l_Lean_Meta_Grind_pushEqCore(x_5, x_26, x_36, x_37, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_34); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +return x_38; +} +else +{ +uint8_t x_39; +lean_dec(x_26); +lean_dec(x_25); +lean_dec(x_17); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_5); +lean_dec(x_1); +x_39 = !lean_is_exclusive(x_32); +if (x_39 == 0) +{ +return x_32; +} +else +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_40 = lean_ctor_get(x_32, 0); +x_41 = lean_ctor_get(x_32, 1); +lean_inc(x_41); +lean_inc(x_40); +lean_dec(x_32); +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_43; +lean_dec(x_26); +lean_dec(x_25); +lean_dec(x_23); +lean_dec(x_17); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_5); +lean_dec(x_1); +x_43 = !lean_is_exclusive(x_30); +if (x_43 == 0) +{ +return x_30; +} +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_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; +} +} +} +else +{ +uint8_t x_47; +lean_dec(x_26); +lean_dec(x_25); +lean_dec(x_23); +lean_dec(x_17); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_5); +lean_dec(x_1); +x_47 = !lean_is_exclusive(x_27); +if (x_47 == 0) +{ +return x_27; +} +else +{ +lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_48 = lean_ctor_get(x_27, 0); +x_49 = lean_ctor_get(x_27, 1); +lean_inc(x_49); +lean_inc(x_48); +lean_dec(x_27); +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_51; +lean_dec(x_17); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_51 = !lean_is_exclusive(x_22); +if (x_51 == 0) +{ +return x_22; +} +else +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; +x_52 = lean_ctor_get(x_22, 0); +x_53 = lean_ctor_get(x_22, 1); +lean_inc(x_53); +lean_inc(x_52); +lean_dec(x_22); +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_55; +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_5); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_55 = !lean_is_exclusive(x_16); +if (x_55 == 0) +{ +return x_16; +} +else +{ +lean_object* x_56; lean_object* x_57; lean_object* x_58; +x_56 = lean_ctor_get(x_16, 0); +x_57 = lean_ctor_get(x_16, 1); +lean_inc(x_57); +lean_inc(x_56); +lean_dec(x_16); +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_EXPORT lean_object* l_Lean_Meta_Grind_propagateForallProp(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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) == 7) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; lean_object* x_15; +x_11 = lean_ctor_get(x_1, 0); +lean_inc(x_11); +x_12 = lean_ctor_get(x_1, 1); +lean_inc(x_12); +x_13 = lean_ctor_get(x_1, 2); +lean_inc(x_13); +x_14 = lean_ctor_get_uint8(x_1, sizeof(void*)*3 + 8); +lean_inc(x_12); +x_15 = l_Lean_Meta_Grind_isEqTrue(x_12, 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; 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) +{ +uint8_t x_18; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_18 = !lean_is_exclusive(x_15); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; +x_19 = lean_ctor_get(x_15, 0); +lean_dec(x_19); +x_20 = lean_box(0); +lean_ctor_set(x_15, 0, x_20); +return x_15; +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_15, 1); +lean_inc(x_21); +lean_dec(x_15); +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; +x_24 = lean_ctor_get(x_15, 1); +lean_inc(x_24); +lean_dec(x_15); +x_25 = lean_box(0); +x_26 = l_Lean_Meta_Grind_propagateForallProp___lambda__1(x_12, x_13, x_11, x_14, x_1, x_25, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_24); +return x_26; +} +} +else +{ +uint8_t x_27; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_27 = !lean_is_exclusive(x_15); +if (x_27 == 0) +{ +return x_15; +} +else +{ +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_dec(x_15); +x_30 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_30, 0, x_28); +lean_ctor_set(x_30, 1, x_29); +return x_30; +} +} +} +else +{ +lean_object* x_31; lean_object* x_32; +lean_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_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_10); +return x_32; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateForallProp___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +_start: +{ +uint8_t x_16; lean_object* x_17; +x_16 = lean_unbox(x_4); +lean_dec(x_4); +x_17 = l_Lean_Meta_Grind_propagateForallProp___lambda__1(x_1, x_2, x_3, x_16, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +lean_dec(x_6); +return x_17; +} +} +lean_object* initialize_Init_Grind_Lemmas(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Meta_Tactic_Grind_Types(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Meta_Tactic_Grind_Internalize(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Meta_Tactic_Grind_Simp(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Lean_Meta_Tactic_Grind_ForallProp(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_Grind_Lemmas(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Lean_Meta_Tactic_Grind_Types(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Lean_Meta_Tactic_Grind_Internalize(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Lean_Meta_Tactic_Grind_Simp(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__1 = _init_l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__1); +l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__2 = _init_l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__2); +l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__3 = _init_l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__3); +l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__4 = _init_l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__4(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__4); +l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__5 = _init_l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__5(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__5); +l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__6 = _init_l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__6(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__6); +l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__7 = _init_l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__7(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__7); +l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__8 = _init_l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__8(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__8); +return lean_io_result_mk_ok(lean_box(0)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Lean/Meta/Tactic/Grind/Internalize.c b/stage0/stdlib/Lean/Meta/Tactic/Grind/Internalize.c index a07cb790a064..c0a9bbeba790 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Grind/Internalize.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Grind/Internalize.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Lean.Meta.Tactic.Grind.Internalize -// Imports: Init.Grind.Util Lean.Meta.LitValues Lean.Meta.Tactic.Grind.Types +// Imports: Init.Grind.Util Lean.Meta.LitValues Lean.Meta.Tactic.Grind.Types Lean.Meta.Tactic.Grind.Util #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -13,30 +13,45 @@ #ifdef __cplusplus extern "C" { #endif +extern lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare; static lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_addCongrTable___spec__5___closed__1; static lean_object* l_Lean_Meta_Grind_addCongrTable___closed__1; static lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__3; +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_internalize___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_PersistentHashMap_find_x3f___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__8; +lean_object* l_Lean_mkAppN(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__4___boxed(lean_object*, lean_object*, lean_object*); size_t lean_usize_shift_right(size_t, size_t); static lean_object* l_Lean_Meta_Grind_addCongrTable___closed__9; +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___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_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__5(lean_object*, size_t, size_t, lean_object*, lean_object*); lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__2___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__2; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__3; +lean_object* l_Lean_Meta_Grind_eraseIrrelevantMData___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_usize_dec_le(size_t, size_t); static lean_object* l_Lean_Meta_Grind_addCongrTable___closed__4; +lean_object* l_Lean_Meta_isProp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__1; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_addCongrTable___spec__6(lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__2; +lean_object* l_Lean_Meta_Grind_unfoldReducible___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_filterTR_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__6(lean_object*, lean_object*, lean_object*); size_t lean_uint64_to_usize(uint64_t); lean_object* l_Lean_Meta_Grind_mkENode(lean_object*, 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_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__4; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findEntry_x3f___at_Lean_Meta_Grind_addCongrTable___spec__1(lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Meta_Grind_isCongruent(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__6(size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_internalize___lambda__2___closed__7; lean_object* l_Lean_Expr_sort___override(lean_object*); +lean_object* l_Lean_MessageData_ofList(lean_object*); lean_object* l_Lean_PersistentArray_push___rarg(lean_object*, lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_contains___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__3(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_internalize___lambda__2___closed__4; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addCongrTable___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*); size_t lean_usize_mul(size_t, size_t); @@ -44,32 +59,54 @@ static lean_object* l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addCongrTable___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_Meta_Grind_internalize___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_mk_array(lean_object*, lean_object*); +static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__2; uint64_t l_Lean_Meta_Grind_congrHash(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addCongrTable___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_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__4; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns(lean_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*); +static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__9; +lean_object* l_Lean_Meta_Grind_normalizeLevels___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_pushEqCore(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_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__1; static lean_object* l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__2; lean_object* l_Lean_Meta_Grind_setENode(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_stringToMessageData(lean_object*); +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___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_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___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_PersistentHashMap_findEntryAtAux___at_Lean_Meta_Grind_addCongrTable___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_internalize___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_internalize___lambda__3___closed__1; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_eraseAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__2(lean_object*, size_t, lean_object*); +static lean_object* l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__7; static lean_object* l_Lean_Meta_Grind_internalize___lambda__2___closed__3; lean_object* l_Lean_Meta_Grind_propagateUp(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_isBVar(lean_object*); extern lean_object* l_instInhabitedPUnit; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addCongrTable___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__3(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__4; static lean_object* l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9___closed__2; +lean_object* l_Lean_Meta_Grind_EMatchTheorems_insert(lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); static size_t l_Lean_PersistentHashMap_findEntryAux___at_Lean_Meta_Grind_addCongrTable___spec__2___closed__2; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_addCongrTable___spec__5(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_internalize___lambda__2___closed__2; +lean_object* l_Lean_PersistentHashMap_isUnaryNode___rarg(lean_object*); +lean_object* l_Lean_Core_transform___at_Lean_Meta_Grind_unfoldReducible___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_take(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_addCongrTable___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t lean_expr_eqv(lean_object*, lean_object*); +static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__3; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addCongrTable___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* l_Lean_Meta_Grind_groundPattern_x3f(lean_object*); +static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__1___closed__1; static lean_object* l_Lean_Meta_Grind_addCongrTable___closed__5; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_outOfBounds___rarg(lean_object*); +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__4(lean_object*, size_t, lean_object*); static lean_object* l_Lean_Meta_Grind_internalize___lambda__2___closed__1; lean_object* lean_st_ref_get(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -88,16 +125,21 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Grind_internalize___lambda__1(lean_object*, static lean_object* l_Lean_Meta_Grind_addCongrTable___closed__7; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_addCongrTable___spec__7(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_PersistentHashMap_insert___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__4(lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Meta_Grind_isSameExpr_unsafe__1(lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_alreadyInternalized(lean_object*, 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_Meta_Grind_addCongrTable___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___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_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__2(lean_object*, size_t, lean_object*); static lean_object* l_Lean_Meta_Grind_addCongrTable___closed__6; +static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__7; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_addCongrTable___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Meta_instMonadMetaM; +lean_object* l_Array_eraseIdx___rarg(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_hasSameType(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___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -105,66 +147,99 @@ lean_object* l_Lean_MessageData_ofExpr(lean_object*); static lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__1; double l_Float_ofScientific(lean_object*, uint8_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___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* l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__2(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___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*, lean_object*); +lean_object* l_Lean_Core_transform___at_Lean_Core_betaReduce___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___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_PersistentHashMap_find_x3f___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__1(lean_object*, lean_object*); +lean_object* l_Lean_Meta_Grind_mkGroundPattern(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findEntryAtAux___at_Lean_Meta_Grind_addCongrTable___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_internalize___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_EXPORT lean_object* l_Lean_Meta_Grind_internalize___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__5; LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___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*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); +lean_object* l_Lean_Meta_Grind_shareCommon(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_containsAtAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_registerParent(lean_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_indentExpr(lean_object*); lean_object* l_Lean_PersistentHashMap_mkEmptyEntries(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___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*); uint8_t l_Lean_Expr_isConstOf(lean_object*, lean_object*); static lean_object* l_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__5; lean_object* lean_array_set(lean_object*, lean_object*, lean_object*); +uint64_t l_Lean_Name_hash___override(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_PersistentHashMap_contains___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__3___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___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_panic_fn(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_internalize___lambda__2___closed__6; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_erase___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__1___boxed(lean_object*, lean_object*); lean_object* lean_nat_sub(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findEntryAux___at_Lean_Meta_Grind_addCongrTable___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_getAppFn(lean_object*); +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___boxed(lean_object**); +lean_object* l_Array_indexOfAux___at_Lean_MetavarContext_setMVarUserName___spec__3(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_mkCollisionNode___rarg(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_getENode(lean_object*, 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_Grind_internalize___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_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___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*); static lean_object* l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__6; lean_object* l_Lean_Meta_Grind_mkENodeCore(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_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_internalize___spec__2___boxed(lean_object**); +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAtAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_internalize___lambda__2___closed__5; +lean_object* l_List_reverse___rarg(lean_object*); static double l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9___closed__1; size_t lean_usize_sub(size_t, size_t); lean_object* lean_array_mk(lean_object*); static lean_object* l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addCongrTable___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_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__6; +lean_object* l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__4(lean_object*, lean_object*); lean_object* l_Lean_isTracingEnabledForCore(lean_object*, lean_object*, lean_object*); size_t lean_usize_add(size_t, size_t); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_eraseAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__2___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9___closed__3; +lean_object* lean_array_uget(lean_object*, size_t); +size_t lean_array_size(lean_object*); lean_object* l_instInhabitedOfMonad___rarg(lean_object*, lean_object*); lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern(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_shift_left(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__7(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__1; uint8_t l___private_Lean_HeadIndex_0__Lean_beqHeadIndex____x40_Lean_HeadIndex___hyg_69_(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_addCongrTable___closed__8; +lean_object* l_Lean_Meta_Grind_canon(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*); static lean_object* l_Lean_Meta_Grind_addCongrTable___closed__2; +static lean_object* l_Lean_Meta_Grind_internalize___lambda__3___closed__2; +LEAN_EXPORT lean_object* l_List_mapM_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_Grind_Origin_key(lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_erase___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__1(lean_object*, lean_object*); +uint8_t lean_usize_dec_lt(size_t, size_t); static lean_object* l_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__3; lean_object* lean_nat_add(lean_object*, lean_object*); extern lean_object* l_Lean_Meta_Grind_congrPlaceholderProof; lean_object* l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(lean_object*); -uint8_t l_Lean_Expr_isConst(lean_object*); uint64_t l_Lean_HeadIndex_hash(lean_object*); static lean_object* l_Lean_Meta_Grind_addCongrTable___closed__3; static lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__2; +lean_object* l_Lean_Meta_Grind_unfoldReducible___lambda__2(lean_object*, lean_object*, 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*); lean_object* lean_array_get(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_isLitValue(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__10; lean_object* l_ReaderT_instMonad___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_internalize(lean_object*, lean_object*, 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_Grind_internalize___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*); size_t lean_usize_land(size_t, size_t); +static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__5; static lean_object* l_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__4; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findEntryAtAux___at_Lean_Meta_Grind_addCongrTable___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: @@ -1485,7 +1560,6 @@ lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); -lean_dec(x_1); return x_19; } else @@ -1536,7 +1610,6 @@ lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); -lean_dec(x_1); return x_35; } } @@ -1583,21 +1656,30 @@ static lean_object* _init_l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__ _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("congr", 5, 5); +x_1 = lean_mk_string_unchecked("debug", 5, 5); return x_1; } } static lean_object* _init_l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__3() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("congr", 5, 5); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__1; x_2 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__2; -x_3 = l_Lean_Name_mkStr2(x_1, x_2); -return x_3; +x_3 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__3; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; } } -static lean_object* _init_l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__4() { +static lean_object* _init_l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__5() { _start: { lean_object* x_1; lean_object* x_2; @@ -1606,7 +1688,7 @@ x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__5() { +static lean_object* _init_l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__6() { _start: { lean_object* x_1; @@ -1614,11 +1696,11 @@ x_1 = lean_mk_string_unchecked(" = ", 3, 3); return x_1; } } -static lean_object* _init_l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__6() { +static lean_object* _init_l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__7() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__5; +x_1 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__6; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } @@ -1627,7 +1709,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addCongrTable___lambda__2(lean_object _start: { lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; -x_13 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__3; +x_13 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__4; x_14 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); x_15 = lean_ctor_get(x_14, 0); lean_inc(x_15); @@ -1655,11 +1737,11 @@ x_22 = lean_ctor_get(x_14, 0); lean_dec(x_22); lean_inc(x_1); x_23 = l_Lean_MessageData_ofExpr(x_1); -x_24 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__4; +x_24 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__5; lean_ctor_set_tag(x_14, 7); lean_ctor_set(x_14, 1, x_23); lean_ctor_set(x_14, 0, x_24); -x_25 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__6; +x_25 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__7; x_26 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_26, 0, x_14); lean_ctor_set(x_26, 1, x_25); @@ -1689,11 +1771,11 @@ lean_inc(x_34); lean_dec(x_14); lean_inc(x_1); x_35 = l_Lean_MessageData_ofExpr(x_1); -x_36 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__4; +x_36 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__5; x_37 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_37, 0, x_36); lean_ctor_set(x_37, 1, x_35); -x_38 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__6; +x_38 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__7; x_39 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_39, 0, x_37); lean_ctor_set(x_39, 1, x_38); @@ -1823,7 +1905,7 @@ lean_inc(x_1); x_16 = l_Lean_PersistentHashMap_findEntry_x3f___at_Lean_Meta_Grind_addCongrTable___spec__1(x_13, x_15, x_1); if (lean_obj_tag(x_16) == 0) { -lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t 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_17; lean_object* x_18; lean_object* x_19; 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; 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_free_object(x_11); lean_dec(x_9); lean_dec(x_8); @@ -1850,259 +1932,289 @@ x_24 = lean_ctor_get(x_18, 4); lean_inc(x_24); x_25 = lean_ctor_get(x_18, 5); lean_inc(x_25); -x_26 = lean_ctor_get_uint8(x_18, sizeof(void*)*8); +x_26 = lean_ctor_get_uint8(x_18, sizeof(void*)*14); x_27 = lean_ctor_get(x_18, 6); lean_inc(x_27); x_28 = lean_ctor_get(x_18, 7); lean_inc(x_28); -x_29 = lean_box(0); +x_29 = lean_ctor_get(x_18, 8); +lean_inc(x_29); +x_30 = lean_ctor_get(x_18, 9); +lean_inc(x_30); +x_31 = lean_ctor_get(x_18, 10); +lean_inc(x_31); +x_32 = lean_ctor_get(x_18, 11); +lean_inc(x_32); +x_33 = lean_ctor_get(x_18, 12); +lean_inc(x_33); +x_34 = lean_ctor_get(x_18, 13); +lean_inc(x_34); +x_35 = lean_box(0); lean_inc(x_18); -x_30 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_addCongrTable___spec__4(x_18, x_23, x_1, x_29); -x_31 = !lean_is_exclusive(x_18); -if (x_31 == 0) +x_36 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_addCongrTable___spec__4(x_18, x_23, x_1, x_35); +x_37 = !lean_is_exclusive(x_18); +if (x_37 == 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; uint8_t x_41; -x_32 = lean_ctor_get(x_18, 7); -lean_dec(x_32); -x_33 = lean_ctor_get(x_18, 6); -lean_dec(x_33); -x_34 = lean_ctor_get(x_18, 5); -lean_dec(x_34); -x_35 = lean_ctor_get(x_18, 4); -lean_dec(x_35); -x_36 = lean_ctor_get(x_18, 3); -lean_dec(x_36); -x_37 = lean_ctor_get(x_18, 2); -lean_dec(x_37); -x_38 = lean_ctor_get(x_18, 1); +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; uint8_t x_53; +x_38 = lean_ctor_get(x_18, 13); lean_dec(x_38); -x_39 = lean_ctor_get(x_18, 0); +x_39 = lean_ctor_get(x_18, 12); lean_dec(x_39); -lean_ctor_set(x_18, 3, x_30); -x_40 = lean_st_ref_set(x_2, x_18, x_19); +x_40 = lean_ctor_get(x_18, 11); +lean_dec(x_40); +x_41 = lean_ctor_get(x_18, 10); +lean_dec(x_41); +x_42 = lean_ctor_get(x_18, 9); +lean_dec(x_42); +x_43 = lean_ctor_get(x_18, 8); +lean_dec(x_43); +x_44 = lean_ctor_get(x_18, 7); +lean_dec(x_44); +x_45 = lean_ctor_get(x_18, 6); +lean_dec(x_45); +x_46 = lean_ctor_get(x_18, 5); +lean_dec(x_46); +x_47 = lean_ctor_get(x_18, 4); +lean_dec(x_47); +x_48 = lean_ctor_get(x_18, 3); +lean_dec(x_48); +x_49 = lean_ctor_get(x_18, 2); +lean_dec(x_49); +x_50 = lean_ctor_get(x_18, 1); +lean_dec(x_50); +x_51 = lean_ctor_get(x_18, 0); +lean_dec(x_51); +lean_ctor_set(x_18, 3, x_36); +x_52 = lean_st_ref_set(x_2, x_18, x_19); lean_dec(x_2); -x_41 = !lean_is_exclusive(x_40); -if (x_41 == 0) +x_53 = !lean_is_exclusive(x_52); +if (x_53 == 0) { -lean_object* x_42; -x_42 = lean_ctor_get(x_40, 0); -lean_dec(x_42); -lean_ctor_set(x_40, 0, x_29); -return x_40; +lean_object* x_54; +x_54 = lean_ctor_get(x_52, 0); +lean_dec(x_54); +lean_ctor_set(x_52, 0, x_35); +return x_52; } else { -lean_object* x_43; lean_object* x_44; -x_43 = lean_ctor_get(x_40, 1); -lean_inc(x_43); -lean_dec(x_40); -x_44 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_44, 0, x_29); -lean_ctor_set(x_44, 1, x_43); -return x_44; +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_35); +lean_ctor_set(x_56, 1, x_55); +return x_56; } } else { -lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_dec(x_18); -x_45 = lean_alloc_ctor(0, 8, 1); -lean_ctor_set(x_45, 0, x_20); -lean_ctor_set(x_45, 1, x_21); -lean_ctor_set(x_45, 2, x_22); -lean_ctor_set(x_45, 3, x_30); -lean_ctor_set(x_45, 4, x_24); -lean_ctor_set(x_45, 5, x_25); -lean_ctor_set(x_45, 6, x_27); -lean_ctor_set(x_45, 7, x_28); -lean_ctor_set_uint8(x_45, sizeof(void*)*8, x_26); -x_46 = lean_st_ref_set(x_2, x_45, x_19); +x_57 = lean_alloc_ctor(0, 14, 1); +lean_ctor_set(x_57, 0, x_20); +lean_ctor_set(x_57, 1, x_21); +lean_ctor_set(x_57, 2, x_22); +lean_ctor_set(x_57, 3, x_36); +lean_ctor_set(x_57, 4, x_24); +lean_ctor_set(x_57, 5, x_25); +lean_ctor_set(x_57, 6, x_27); +lean_ctor_set(x_57, 7, x_28); +lean_ctor_set(x_57, 8, x_29); +lean_ctor_set(x_57, 9, x_30); +lean_ctor_set(x_57, 10, x_31); +lean_ctor_set(x_57, 11, x_32); +lean_ctor_set(x_57, 12, x_33); +lean_ctor_set(x_57, 13, x_34); +lean_ctor_set_uint8(x_57, sizeof(void*)*14, x_26); +x_58 = lean_st_ref_set(x_2, x_57, x_19); lean_dec(x_2); -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; +x_59 = lean_ctor_get(x_58, 1); +lean_inc(x_59); +if (lean_is_exclusive(x_58)) { + lean_ctor_release(x_58, 0); + lean_ctor_release(x_58, 1); + x_60 = x_58; } else { - lean_dec_ref(x_46); - x_48 = lean_box(0); + lean_dec_ref(x_58); + x_60 = lean_box(0); } -if (lean_is_scalar(x_48)) { - x_49 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_60)) { + x_61 = lean_alloc_ctor(0, 2, 0); } else { - x_49 = x_48; + x_61 = x_60; } -lean_ctor_set(x_49, 0, x_29); -lean_ctor_set(x_49, 1, x_47); -return x_49; +lean_ctor_set(x_61, 0, x_35); +lean_ctor_set(x_61, 1, x_59); +return x_61; } } else { -lean_object* x_50; uint8_t x_51; -x_50 = lean_ctor_get(x_16, 0); -lean_inc(x_50); +lean_object* x_62; uint8_t x_63; +x_62 = lean_ctor_get(x_16, 0); +lean_inc(x_62); lean_dec(x_16); -x_51 = !lean_is_exclusive(x_50); -if (x_51 == 0) +x_63 = !lean_is_exclusive(x_62); +if (x_63 == 0) { -lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; uint8_t x_56; -x_52 = lean_ctor_get(x_50, 0); -x_53 = lean_ctor_get(x_50, 1); -lean_dec(x_53); -x_54 = l_Lean_Expr_getAppFn(x_1); -x_55 = l_Lean_Expr_getAppFn(x_52); -x_56 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_54, x_55); -if (x_56 == 0) +lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; uint8_t x_68; +x_64 = lean_ctor_get(x_62, 0); +x_65 = lean_ctor_get(x_62, 1); +lean_dec(x_65); +x_66 = l_Lean_Expr_getAppFn(x_1); +x_67 = l_Lean_Expr_getAppFn(x_64); +x_68 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_66, x_67); +if (x_68 == 0) { -lean_object* x_57; +lean_object* x_69; lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -x_57 = l_Lean_Meta_Grind_hasSameType(x_54, x_55, x_6, x_7, x_8, x_9, x_14); -if (lean_obj_tag(x_57) == 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) +x_69 = l_Lean_Meta_Grind_hasSameType(x_66, x_67, x_6, x_7, x_8, x_9, x_14); +if (lean_obj_tag(x_69) == 0) { -lean_object* x_60; lean_object* x_61; lean_object* x_62; uint8_t x_63; -x_60 = lean_ctor_get(x_57, 1); -lean_inc(x_60); -lean_dec(x_57); -x_61 = l_Lean_Meta_Grind_addCongrTable___closed__2; -x_62 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_61, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_60); -x_63 = !lean_is_exclusive(x_62); -if (x_63 == 0) +lean_object* x_70; uint8_t x_71; +x_70 = lean_ctor_get(x_69, 0); +lean_inc(x_70); +x_71 = lean_unbox(x_70); +lean_dec(x_70); +if (x_71 == 0) { -lean_object* x_64; lean_object* x_65; lean_object* x_66; uint8_t x_67; -x_64 = lean_ctor_get(x_62, 0); -x_65 = lean_ctor_get(x_62, 1); -x_66 = l_Lean_Meta_Grind_addCongrTable___closed__3; -x_67 = lean_unbox(x_64); -lean_dec(x_64); -if (x_67 == 0) +lean_object* x_72; lean_object* x_73; lean_object* x_74; uint8_t x_75; +x_72 = lean_ctor_get(x_69, 1); +lean_inc(x_72); +lean_dec(x_69); +x_73 = l_Lean_Meta_Grind_addCongrTable___closed__2; +x_74 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_73, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_72); +x_75 = !lean_is_exclusive(x_74); +if (x_75 == 0) +{ +lean_object* x_76; lean_object* x_77; lean_object* x_78; uint8_t x_79; +x_76 = lean_ctor_get(x_74, 0); +x_77 = lean_ctor_get(x_74, 1); +x_78 = l_Lean_Meta_Grind_addCongrTable___closed__3; +x_79 = lean_unbox(x_76); +lean_dec(x_76); +if (x_79 == 0) { -lean_object* x_68; lean_object* x_69; +lean_object* x_80; lean_object* x_81; +lean_free_object(x_74); lean_free_object(x_62); -lean_free_object(x_50); -lean_dec(x_52); +lean_dec(x_64); lean_free_object(x_11); lean_dec(x_1); -x_68 = lean_box(0); -x_69 = lean_apply_10(x_66, x_68, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_65); -return x_69; +x_80 = lean_box(0); +x_81 = lean_apply_10(x_78, x_80, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_77); +return x_81; } else { -lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; -x_70 = l_Lean_indentExpr(x_1); -x_71 = l_Lean_Meta_Grind_addCongrTable___closed__5; +lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; +x_82 = l_Lean_indentExpr(x_1); +x_83 = l_Lean_Meta_Grind_addCongrTable___closed__5; +lean_ctor_set_tag(x_74, 7); +lean_ctor_set(x_74, 1, x_82); +lean_ctor_set(x_74, 0, x_83); +x_84 = l_Lean_Meta_Grind_addCongrTable___closed__7; lean_ctor_set_tag(x_62, 7); -lean_ctor_set(x_62, 1, x_70); -lean_ctor_set(x_62, 0, x_71); -x_72 = l_Lean_Meta_Grind_addCongrTable___closed__7; -lean_ctor_set_tag(x_50, 7); -lean_ctor_set(x_50, 1, x_72); -lean_ctor_set(x_50, 0, x_62); -x_73 = l_Lean_indentExpr(x_52); +lean_ctor_set(x_62, 1, x_84); +lean_ctor_set(x_62, 0, x_74); +x_85 = l_Lean_indentExpr(x_64); lean_ctor_set_tag(x_11, 7); -lean_ctor_set(x_11, 1, x_73); -lean_ctor_set(x_11, 0, x_50); -x_74 = l_Lean_Meta_Grind_addCongrTable___closed__9; -x_75 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_75, 0, x_11); -lean_ctor_set(x_75, 1, x_74); -x_76 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_61, x_75, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_65); -x_77 = lean_ctor_get(x_76, 0); -lean_inc(x_77); -x_78 = lean_ctor_get(x_76, 1); -lean_inc(x_78); -lean_dec(x_76); -x_79 = lean_apply_10(x_66, x_77, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_78); -return x_79; +lean_ctor_set(x_11, 1, x_85); +lean_ctor_set(x_11, 0, x_62); +x_86 = l_Lean_Meta_Grind_addCongrTable___closed__9; +x_87 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_87, 0, x_11); +lean_ctor_set(x_87, 1, x_86); +x_88 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_73, x_87, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, 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); +lean_dec(x_88); +x_91 = lean_apply_10(x_78, x_89, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_90); +return x_91; } } else { -lean_object* x_80; lean_object* x_81; lean_object* x_82; uint8_t x_83; -x_80 = lean_ctor_get(x_62, 0); -x_81 = lean_ctor_get(x_62, 1); -lean_inc(x_81); -lean_inc(x_80); -lean_dec(x_62); -x_82 = l_Lean_Meta_Grind_addCongrTable___closed__3; -x_83 = lean_unbox(x_80); -lean_dec(x_80); -if (x_83 == 0) +lean_object* x_92; lean_object* x_93; lean_object* x_94; uint8_t x_95; +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_dec(x_74); +x_94 = l_Lean_Meta_Grind_addCongrTable___closed__3; +x_95 = lean_unbox(x_92); +lean_dec(x_92); +if (x_95 == 0) { -lean_object* x_84; lean_object* x_85; -lean_free_object(x_50); -lean_dec(x_52); +lean_object* x_96; lean_object* x_97; +lean_free_object(x_62); +lean_dec(x_64); lean_free_object(x_11); lean_dec(x_1); -x_84 = lean_box(0); -x_85 = lean_apply_10(x_82, x_84, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_81); -return x_85; +x_96 = lean_box(0); +x_97 = lean_apply_10(x_94, x_96, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_93); +return x_97; } 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 = l_Lean_indentExpr(x_1); -x_87 = l_Lean_Meta_Grind_addCongrTable___closed__5; -x_88 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_88, 0, x_87); -lean_ctor_set(x_88, 1, x_86); -x_89 = l_Lean_Meta_Grind_addCongrTable___closed__7; -lean_ctor_set_tag(x_50, 7); -lean_ctor_set(x_50, 1, x_89); -lean_ctor_set(x_50, 0, x_88); -x_90 = l_Lean_indentExpr(x_52); +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_98 = l_Lean_indentExpr(x_1); +x_99 = l_Lean_Meta_Grind_addCongrTable___closed__5; +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 = l_Lean_Meta_Grind_addCongrTable___closed__7; +lean_ctor_set_tag(x_62, 7); +lean_ctor_set(x_62, 1, x_101); +lean_ctor_set(x_62, 0, x_100); +x_102 = l_Lean_indentExpr(x_64); lean_ctor_set_tag(x_11, 7); -lean_ctor_set(x_11, 1, x_90); -lean_ctor_set(x_11, 0, x_50); -x_91 = l_Lean_Meta_Grind_addCongrTable___closed__9; -x_92 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_92, 0, x_11); -lean_ctor_set(x_92, 1, x_91); -x_93 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_61, x_92, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_81); -x_94 = lean_ctor_get(x_93, 0); -lean_inc(x_94); -x_95 = lean_ctor_get(x_93, 1); -lean_inc(x_95); -lean_dec(x_93); -x_96 = lean_apply_10(x_82, x_94, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_95); -return x_96; +lean_ctor_set(x_11, 1, x_102); +lean_ctor_set(x_11, 0, x_62); +x_103 = l_Lean_Meta_Grind_addCongrTable___closed__9; +x_104 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_104, 0, x_11); +lean_ctor_set(x_104, 1, x_103); +x_105 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_73, x_104, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_93); +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_apply_10(x_94, x_106, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_107); +return x_108; } } } else { -lean_object* x_97; lean_object* x_98; lean_object* x_99; -lean_free_object(x_50); +lean_object* x_109; lean_object* x_110; lean_object* x_111; +lean_free_object(x_62); lean_free_object(x_11); -x_97 = lean_ctor_get(x_57, 1); -lean_inc(x_97); -lean_dec(x_57); -x_98 = lean_box(0); -x_99 = l_Lean_Meta_Grind_addCongrTable___lambda__2(x_1, x_52, x_98, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_97); +x_109 = lean_ctor_get(x_69, 1); +lean_inc(x_109); +lean_dec(x_69); +x_110 = lean_box(0); +x_111 = l_Lean_Meta_Grind_addCongrTable___lambda__2(x_1, x_64, x_110, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_109); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -return x_99; +return x_111; } } else { -uint8_t x_100; -lean_free_object(x_50); -lean_dec(x_52); +uint8_t x_112; +lean_free_object(x_62); +lean_dec(x_64); lean_free_object(x_11); lean_dec(x_9); lean_dec(x_8); @@ -2113,155 +2225,155 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_100 = !lean_is_exclusive(x_57); -if (x_100 == 0) +x_112 = !lean_is_exclusive(x_69); +if (x_112 == 0) { -return x_57; +return x_69; } else { -lean_object* x_101; lean_object* x_102; lean_object* x_103; -x_101 = lean_ctor_get(x_57, 0); -x_102 = lean_ctor_get(x_57, 1); -lean_inc(x_102); -lean_inc(x_101); -lean_dec(x_57); -x_103 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_103, 0, x_101); -lean_ctor_set(x_103, 1, x_102); -return x_103; +lean_object* x_113; lean_object* x_114; lean_object* x_115; +x_113 = lean_ctor_get(x_69, 0); +x_114 = lean_ctor_get(x_69, 1); +lean_inc(x_114); +lean_inc(x_113); +lean_dec(x_69); +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_104; lean_object* x_105; -lean_dec(x_55); -lean_dec(x_54); -lean_free_object(x_50); +lean_object* x_116; lean_object* x_117; +lean_dec(x_67); +lean_dec(x_66); +lean_free_object(x_62); lean_free_object(x_11); -x_104 = lean_box(0); -x_105 = l_Lean_Meta_Grind_addCongrTable___lambda__2(x_1, x_52, x_104, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_14); +x_116 = lean_box(0); +x_117 = l_Lean_Meta_Grind_addCongrTable___lambda__2(x_1, x_64, x_116, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_14); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -return x_105; +return x_117; } } else { -lean_object* x_106; lean_object* x_107; lean_object* x_108; uint8_t x_109; -x_106 = lean_ctor_get(x_50, 0); -lean_inc(x_106); -lean_dec(x_50); -x_107 = l_Lean_Expr_getAppFn(x_1); -x_108 = l_Lean_Expr_getAppFn(x_106); -x_109 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_107, x_108); -if (x_109 == 0) +lean_object* x_118; lean_object* x_119; lean_object* x_120; uint8_t x_121; +x_118 = lean_ctor_get(x_62, 0); +lean_inc(x_118); +lean_dec(x_62); +x_119 = l_Lean_Expr_getAppFn(x_1); +x_120 = l_Lean_Expr_getAppFn(x_118); +x_121 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_119, x_120); +if (x_121 == 0) { -lean_object* x_110; +lean_object* x_122; lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -x_110 = l_Lean_Meta_Grind_hasSameType(x_107, x_108, x_6, x_7, x_8, x_9, x_14); -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_object* x_117; lean_object* x_118; lean_object* x_119; uint8_t x_120; -x_113 = lean_ctor_get(x_110, 1); -lean_inc(x_113); -lean_dec(x_110); -x_114 = l_Lean_Meta_Grind_addCongrTable___closed__2; -x_115 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_114, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, 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); -if (lean_is_exclusive(x_115)) { - lean_ctor_release(x_115, 0); - lean_ctor_release(x_115, 1); - x_118 = x_115; +x_122 = l_Lean_Meta_Grind_hasSameType(x_119, x_120, x_6, x_7, x_8, x_9, x_14); +if (lean_obj_tag(x_122) == 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); +lean_dec(x_123); +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; uint8_t x_132; +x_125 = lean_ctor_get(x_122, 1); +lean_inc(x_125); +lean_dec(x_122); +x_126 = l_Lean_Meta_Grind_addCongrTable___closed__2; +x_127 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_126, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_125); +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; } else { - lean_dec_ref(x_115); - x_118 = lean_box(0); + lean_dec_ref(x_127); + x_130 = lean_box(0); } -x_119 = l_Lean_Meta_Grind_addCongrTable___closed__3; -x_120 = lean_unbox(x_116); -lean_dec(x_116); -if (x_120 == 0) +x_131 = l_Lean_Meta_Grind_addCongrTable___closed__3; +x_132 = lean_unbox(x_128); +lean_dec(x_128); +if (x_132 == 0) { -lean_object* x_121; lean_object* x_122; +lean_object* x_133; lean_object* x_134; +lean_dec(x_130); lean_dec(x_118); -lean_dec(x_106); lean_free_object(x_11); lean_dec(x_1); -x_121 = lean_box(0); -x_122 = lean_apply_10(x_119, x_121, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_117); -return x_122; +x_133 = lean_box(0); +x_134 = lean_apply_10(x_131, x_133, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_129); +return x_134; } else { -lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; -x_123 = l_Lean_indentExpr(x_1); -x_124 = l_Lean_Meta_Grind_addCongrTable___closed__5; -if (lean_is_scalar(x_118)) { - x_125 = lean_alloc_ctor(7, 2, 0); +lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; +x_135 = l_Lean_indentExpr(x_1); +x_136 = l_Lean_Meta_Grind_addCongrTable___closed__5; +if (lean_is_scalar(x_130)) { + x_137 = lean_alloc_ctor(7, 2, 0); } else { - x_125 = x_118; - lean_ctor_set_tag(x_125, 7); -} -lean_ctor_set(x_125, 0, x_124); -lean_ctor_set(x_125, 1, x_123); -x_126 = l_Lean_Meta_Grind_addCongrTable___closed__7; -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_indentExpr(x_106); + x_137 = x_130; + lean_ctor_set_tag(x_137, 7); +} +lean_ctor_set(x_137, 0, x_136); +lean_ctor_set(x_137, 1, x_135); +x_138 = l_Lean_Meta_Grind_addCongrTable___closed__7; +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_indentExpr(x_118); lean_ctor_set_tag(x_11, 7); -lean_ctor_set(x_11, 1, x_128); -lean_ctor_set(x_11, 0, x_127); -x_129 = l_Lean_Meta_Grind_addCongrTable___closed__9; -x_130 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_130, 0, x_11); -lean_ctor_set(x_130, 1, x_129); -x_131 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_114, x_130, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_117); -x_132 = lean_ctor_get(x_131, 0); -lean_inc(x_132); -x_133 = lean_ctor_get(x_131, 1); -lean_inc(x_133); -lean_dec(x_131); -x_134 = lean_apply_10(x_119, x_132, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_133); -return x_134; +lean_ctor_set(x_11, 1, x_140); +lean_ctor_set(x_11, 0, x_139); +x_141 = l_Lean_Meta_Grind_addCongrTable___closed__9; +x_142 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_142, 0, x_11); +lean_ctor_set(x_142, 1, x_141); +x_143 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_126, x_142, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_129); +x_144 = lean_ctor_get(x_143, 0); +lean_inc(x_144); +x_145 = lean_ctor_get(x_143, 1); +lean_inc(x_145); +lean_dec(x_143); +x_146 = lean_apply_10(x_131, x_144, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_145); +return x_146; } } else { -lean_object* x_135; lean_object* x_136; lean_object* x_137; +lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_free_object(x_11); -x_135 = lean_ctor_get(x_110, 1); -lean_inc(x_135); -lean_dec(x_110); -x_136 = lean_box(0); -x_137 = l_Lean_Meta_Grind_addCongrTable___lambda__2(x_1, x_106, x_136, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_135); +x_147 = lean_ctor_get(x_122, 1); +lean_inc(x_147); +lean_dec(x_122); +x_148 = lean_box(0); +x_149 = l_Lean_Meta_Grind_addCongrTable___lambda__2(x_1, x_118, x_148, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_147); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -return x_137; +return x_149; } } else { -lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; -lean_dec(x_106); +lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; +lean_dec(x_118); lean_free_object(x_11); lean_dec(x_9); lean_dec(x_8); @@ -2272,60 +2384,60 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_138 = lean_ctor_get(x_110, 0); -lean_inc(x_138); -x_139 = lean_ctor_get(x_110, 1); -lean_inc(x_139); -if (lean_is_exclusive(x_110)) { - lean_ctor_release(x_110, 0); - lean_ctor_release(x_110, 1); - x_140 = x_110; +x_150 = lean_ctor_get(x_122, 0); +lean_inc(x_150); +x_151 = lean_ctor_get(x_122, 1); +lean_inc(x_151); +if (lean_is_exclusive(x_122)) { + lean_ctor_release(x_122, 0); + lean_ctor_release(x_122, 1); + x_152 = x_122; } else { - lean_dec_ref(x_110); - x_140 = lean_box(0); + lean_dec_ref(x_122); + x_152 = lean_box(0); } -if (lean_is_scalar(x_140)) { - x_141 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_152)) { + x_153 = lean_alloc_ctor(1, 2, 0); } else { - x_141 = x_140; + x_153 = x_152; } -lean_ctor_set(x_141, 0, x_138); -lean_ctor_set(x_141, 1, x_139); -return x_141; +lean_ctor_set(x_153, 0, x_150); +lean_ctor_set(x_153, 1, x_151); +return x_153; } } else { -lean_object* x_142; lean_object* x_143; -lean_dec(x_108); -lean_dec(x_107); +lean_object* x_154; lean_object* x_155; +lean_dec(x_120); +lean_dec(x_119); lean_free_object(x_11); -x_142 = lean_box(0); -x_143 = l_Lean_Meta_Grind_addCongrTable___lambda__2(x_1, x_106, x_142, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_14); +x_154 = lean_box(0); +x_155 = l_Lean_Meta_Grind_addCongrTable___lambda__2(x_1, x_118, x_154, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_14); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -return x_143; +return x_155; } } } } else { -lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; -x_144 = lean_ctor_get(x_11, 0); -x_145 = lean_ctor_get(x_11, 1); -lean_inc(x_145); -lean_inc(x_144); +lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; +x_156 = lean_ctor_get(x_11, 0); +x_157 = lean_ctor_get(x_11, 1); +lean_inc(x_157); +lean_inc(x_156); lean_dec(x_11); -x_146 = lean_ctor_get(x_144, 3); -lean_inc(x_146); +x_158 = lean_ctor_get(x_156, 3); +lean_inc(x_158); lean_inc(x_1); -x_147 = l_Lean_PersistentHashMap_findEntry_x3f___at_Lean_Meta_Grind_addCongrTable___spec__1(x_144, x_146, x_1); -if (lean_obj_tag(x_147) == 0) +x_159 = l_Lean_PersistentHashMap_findEntry_x3f___at_Lean_Meta_Grind_addCongrTable___spec__1(x_156, x_158, x_1); +if (lean_obj_tag(x_159) == 0) { -lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; 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; lean_object* x_167; +lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; uint8_t x_169; lean_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_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -2333,210 +2445,234 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_148 = lean_st_ref_take(x_2, x_145); -x_149 = lean_ctor_get(x_148, 0); -lean_inc(x_149); -x_150 = lean_ctor_get(x_148, 1); -lean_inc(x_150); -lean_dec(x_148); -x_151 = lean_ctor_get(x_149, 0); -lean_inc(x_151); -x_152 = lean_ctor_get(x_149, 1); -lean_inc(x_152); -x_153 = lean_ctor_get(x_149, 2); -lean_inc(x_153); -x_154 = lean_ctor_get(x_149, 3); -lean_inc(x_154); -x_155 = lean_ctor_get(x_149, 4); -lean_inc(x_155); -x_156 = lean_ctor_get(x_149, 5); -lean_inc(x_156); -x_157 = lean_ctor_get_uint8(x_149, sizeof(void*)*8); -x_158 = lean_ctor_get(x_149, 6); -lean_inc(x_158); -x_159 = lean_ctor_get(x_149, 7); -lean_inc(x_159); -x_160 = lean_box(0); -lean_inc(x_149); -x_161 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_addCongrTable___spec__4(x_149, x_154, x_1, x_160); -if (lean_is_exclusive(x_149)) { - lean_ctor_release(x_149, 0); - lean_ctor_release(x_149, 1); - lean_ctor_release(x_149, 2); - lean_ctor_release(x_149, 3); - lean_ctor_release(x_149, 4); - lean_ctor_release(x_149, 5); - lean_ctor_release(x_149, 6); - lean_ctor_release(x_149, 7); - x_162 = x_149; +x_160 = lean_st_ref_take(x_2, x_157); +x_161 = lean_ctor_get(x_160, 0); +lean_inc(x_161); +x_162 = lean_ctor_get(x_160, 1); +lean_inc(x_162); +lean_dec(x_160); +x_163 = lean_ctor_get(x_161, 0); +lean_inc(x_163); +x_164 = lean_ctor_get(x_161, 1); +lean_inc(x_164); +x_165 = lean_ctor_get(x_161, 2); +lean_inc(x_165); +x_166 = lean_ctor_get(x_161, 3); +lean_inc(x_166); +x_167 = lean_ctor_get(x_161, 4); +lean_inc(x_167); +x_168 = lean_ctor_get(x_161, 5); +lean_inc(x_168); +x_169 = lean_ctor_get_uint8(x_161, sizeof(void*)*14); +x_170 = lean_ctor_get(x_161, 6); +lean_inc(x_170); +x_171 = lean_ctor_get(x_161, 7); +lean_inc(x_171); +x_172 = lean_ctor_get(x_161, 8); +lean_inc(x_172); +x_173 = lean_ctor_get(x_161, 9); +lean_inc(x_173); +x_174 = lean_ctor_get(x_161, 10); +lean_inc(x_174); +x_175 = lean_ctor_get(x_161, 11); +lean_inc(x_175); +x_176 = lean_ctor_get(x_161, 12); +lean_inc(x_176); +x_177 = lean_ctor_get(x_161, 13); +lean_inc(x_177); +x_178 = lean_box(0); +lean_inc(x_161); +x_179 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_addCongrTable___spec__4(x_161, x_166, x_1, x_178); +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); + lean_ctor_release(x_161, 5); + lean_ctor_release(x_161, 6); + lean_ctor_release(x_161, 7); + lean_ctor_release(x_161, 8); + lean_ctor_release(x_161, 9); + lean_ctor_release(x_161, 10); + lean_ctor_release(x_161, 11); + lean_ctor_release(x_161, 12); + lean_ctor_release(x_161, 13); + x_180 = x_161; } else { - lean_dec_ref(x_149); - x_162 = lean_box(0); + lean_dec_ref(x_161); + x_180 = lean_box(0); } -if (lean_is_scalar(x_162)) { - x_163 = lean_alloc_ctor(0, 8, 1); +if (lean_is_scalar(x_180)) { + x_181 = lean_alloc_ctor(0, 14, 1); } else { - x_163 = x_162; -} -lean_ctor_set(x_163, 0, x_151); -lean_ctor_set(x_163, 1, x_152); -lean_ctor_set(x_163, 2, x_153); -lean_ctor_set(x_163, 3, x_161); -lean_ctor_set(x_163, 4, x_155); -lean_ctor_set(x_163, 5, x_156); -lean_ctor_set(x_163, 6, x_158); -lean_ctor_set(x_163, 7, x_159); -lean_ctor_set_uint8(x_163, sizeof(void*)*8, x_157); -x_164 = lean_st_ref_set(x_2, x_163, x_150); + x_181 = x_180; +} +lean_ctor_set(x_181, 0, x_163); +lean_ctor_set(x_181, 1, x_164); +lean_ctor_set(x_181, 2, x_165); +lean_ctor_set(x_181, 3, x_179); +lean_ctor_set(x_181, 4, x_167); +lean_ctor_set(x_181, 5, x_168); +lean_ctor_set(x_181, 6, x_170); +lean_ctor_set(x_181, 7, x_171); +lean_ctor_set(x_181, 8, x_172); +lean_ctor_set(x_181, 9, x_173); +lean_ctor_set(x_181, 10, x_174); +lean_ctor_set(x_181, 11, x_175); +lean_ctor_set(x_181, 12, x_176); +lean_ctor_set(x_181, 13, x_177); +lean_ctor_set_uint8(x_181, sizeof(void*)*14, x_169); +x_182 = lean_st_ref_set(x_2, x_181, x_162); lean_dec(x_2); -x_165 = lean_ctor_get(x_164, 1); -lean_inc(x_165); -if (lean_is_exclusive(x_164)) { - lean_ctor_release(x_164, 0); - lean_ctor_release(x_164, 1); - x_166 = x_164; +x_183 = lean_ctor_get(x_182, 1); +lean_inc(x_183); +if (lean_is_exclusive(x_182)) { + lean_ctor_release(x_182, 0); + lean_ctor_release(x_182, 1); + x_184 = x_182; } else { - lean_dec_ref(x_164); - x_166 = lean_box(0); + lean_dec_ref(x_182); + x_184 = lean_box(0); } -if (lean_is_scalar(x_166)) { - x_167 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_184)) { + x_185 = lean_alloc_ctor(0, 2, 0); } else { - x_167 = x_166; + x_185 = x_184; } -lean_ctor_set(x_167, 0, x_160); -lean_ctor_set(x_167, 1, x_165); -return x_167; +lean_ctor_set(x_185, 0, x_178); +lean_ctor_set(x_185, 1, x_183); +return x_185; } else { -lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; uint8_t x_173; -x_168 = lean_ctor_get(x_147, 0); -lean_inc(x_168); -lean_dec(x_147); -x_169 = lean_ctor_get(x_168, 0); -lean_inc(x_169); -if (lean_is_exclusive(x_168)) { - lean_ctor_release(x_168, 0); - lean_ctor_release(x_168, 1); - x_170 = x_168; +lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; uint8_t x_191; +x_186 = lean_ctor_get(x_159, 0); +lean_inc(x_186); +lean_dec(x_159); +x_187 = lean_ctor_get(x_186, 0); +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_168); - x_170 = lean_box(0); + lean_dec_ref(x_186); + x_188 = lean_box(0); } -x_171 = l_Lean_Expr_getAppFn(x_1); -x_172 = l_Lean_Expr_getAppFn(x_169); -x_173 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_171, x_172); -if (x_173 == 0) +x_189 = l_Lean_Expr_getAppFn(x_1); +x_190 = l_Lean_Expr_getAppFn(x_187); +x_191 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_189, x_190); +if (x_191 == 0) { -lean_object* x_174; +lean_object* x_192; lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -x_174 = l_Lean_Meta_Grind_hasSameType(x_171, x_172, x_6, x_7, x_8, x_9, x_145); -if (lean_obj_tag(x_174) == 0) -{ -lean_object* x_175; uint8_t x_176; -x_175 = lean_ctor_get(x_174, 0); -lean_inc(x_175); -x_176 = lean_unbox(x_175); -lean_dec(x_175); -if (x_176 == 0) +x_192 = l_Lean_Meta_Grind_hasSameType(x_189, x_190, x_6, x_7, x_8, x_9, x_157); +if (lean_obj_tag(x_192) == 0) { -lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; uint8_t x_184; -x_177 = lean_ctor_get(x_174, 1); -lean_inc(x_177); -lean_dec(x_174); -x_178 = l_Lean_Meta_Grind_addCongrTable___closed__2; -x_179 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_178, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_177); -x_180 = lean_ctor_get(x_179, 0); -lean_inc(x_180); -x_181 = lean_ctor_get(x_179, 1); -lean_inc(x_181); -if (lean_is_exclusive(x_179)) { - lean_ctor_release(x_179, 0); - lean_ctor_release(x_179, 1); - x_182 = x_179; +lean_object* x_193; uint8_t x_194; +x_193 = lean_ctor_get(x_192, 0); +lean_inc(x_193); +x_194 = lean_unbox(x_193); +lean_dec(x_193); +if (x_194 == 0) +{ +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; uint8_t x_202; +x_195 = lean_ctor_get(x_192, 1); +lean_inc(x_195); +lean_dec(x_192); +x_196 = l_Lean_Meta_Grind_addCongrTable___closed__2; +x_197 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_196, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_195); +x_198 = lean_ctor_get(x_197, 0); +lean_inc(x_198); +x_199 = lean_ctor_get(x_197, 1); +lean_inc(x_199); +if (lean_is_exclusive(x_197)) { + lean_ctor_release(x_197, 0); + lean_ctor_release(x_197, 1); + x_200 = x_197; } else { - lean_dec_ref(x_179); - x_182 = lean_box(0); + lean_dec_ref(x_197); + x_200 = lean_box(0); } -x_183 = l_Lean_Meta_Grind_addCongrTable___closed__3; -x_184 = lean_unbox(x_180); -lean_dec(x_180); -if (x_184 == 0) +x_201 = l_Lean_Meta_Grind_addCongrTable___closed__3; +x_202 = lean_unbox(x_198); +lean_dec(x_198); +if (x_202 == 0) { -lean_object* x_185; lean_object* x_186; -lean_dec(x_182); -lean_dec(x_170); -lean_dec(x_169); +lean_object* x_203; lean_object* x_204; +lean_dec(x_200); +lean_dec(x_188); +lean_dec(x_187); lean_dec(x_1); -x_185 = lean_box(0); -x_186 = lean_apply_10(x_183, x_185, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_181); -return x_186; +x_203 = lean_box(0); +x_204 = lean_apply_10(x_201, x_203, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_199); +return x_204; } else { -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; -x_187 = l_Lean_indentExpr(x_1); -x_188 = l_Lean_Meta_Grind_addCongrTable___closed__5; -if (lean_is_scalar(x_182)) { - x_189 = lean_alloc_ctor(7, 2, 0); +lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; +x_205 = l_Lean_indentExpr(x_1); +x_206 = l_Lean_Meta_Grind_addCongrTable___closed__5; +if (lean_is_scalar(x_200)) { + x_207 = lean_alloc_ctor(7, 2, 0); } else { - x_189 = x_182; - lean_ctor_set_tag(x_189, 7); -} -lean_ctor_set(x_189, 0, x_188); -lean_ctor_set(x_189, 1, x_187); -x_190 = l_Lean_Meta_Grind_addCongrTable___closed__7; -if (lean_is_scalar(x_170)) { - x_191 = lean_alloc_ctor(7, 2, 0); + x_207 = x_200; + lean_ctor_set_tag(x_207, 7); +} +lean_ctor_set(x_207, 0, x_206); +lean_ctor_set(x_207, 1, x_205); +x_208 = l_Lean_Meta_Grind_addCongrTable___closed__7; +if (lean_is_scalar(x_188)) { + x_209 = lean_alloc_ctor(7, 2, 0); } else { - x_191 = x_170; - lean_ctor_set_tag(x_191, 7); -} -lean_ctor_set(x_191, 0, x_189); -lean_ctor_set(x_191, 1, x_190); -x_192 = l_Lean_indentExpr(x_169); -x_193 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_193, 0, x_191); -lean_ctor_set(x_193, 1, x_192); -x_194 = l_Lean_Meta_Grind_addCongrTable___closed__9; -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_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_178, x_195, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_181); -x_197 = lean_ctor_get(x_196, 0); -lean_inc(x_197); -x_198 = lean_ctor_get(x_196, 1); -lean_inc(x_198); -lean_dec(x_196); -x_199 = lean_apply_10(x_183, x_197, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_198); -return x_199; -} -} -else -{ -lean_object* x_200; lean_object* x_201; lean_object* x_202; -lean_dec(x_170); -x_200 = lean_ctor_get(x_174, 1); -lean_inc(x_200); -lean_dec(x_174); -x_201 = lean_box(0); -x_202 = l_Lean_Meta_Grind_addCongrTable___lambda__2(x_1, x_169, x_201, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_200); + x_209 = x_188; + lean_ctor_set_tag(x_209, 7); +} +lean_ctor_set(x_209, 0, x_207); +lean_ctor_set(x_209, 1, x_208); +x_210 = l_Lean_indentExpr(x_187); +x_211 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_211, 0, x_209); +lean_ctor_set(x_211, 1, x_210); +x_212 = l_Lean_Meta_Grind_addCongrTable___closed__9; +x_213 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_213, 0, x_211); +lean_ctor_set(x_213, 1, x_212); +x_214 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_196, x_213, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_199); +x_215 = lean_ctor_get(x_214, 0); +lean_inc(x_215); +x_216 = lean_ctor_get(x_214, 1); +lean_inc(x_216); +lean_dec(x_214); +x_217 = lean_apply_10(x_201, x_215, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_216); +return x_217; +} +} +else +{ +lean_object* x_218; lean_object* x_219; lean_object* x_220; +lean_dec(x_188); +x_218 = lean_ctor_get(x_192, 1); +lean_inc(x_218); +lean_dec(x_192); +x_219 = lean_box(0); +x_220 = l_Lean_Meta_Grind_addCongrTable___lambda__2(x_1, x_187, x_219, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_218); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -return x_202; +return x_220; } } else { -lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; -lean_dec(x_170); -lean_dec(x_169); +lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; +lean_dec(x_188); +lean_dec(x_187); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -2546,41 +2682,41 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_203 = lean_ctor_get(x_174, 0); -lean_inc(x_203); -x_204 = lean_ctor_get(x_174, 1); -lean_inc(x_204); -if (lean_is_exclusive(x_174)) { - lean_ctor_release(x_174, 0); - lean_ctor_release(x_174, 1); - x_205 = x_174; +x_221 = lean_ctor_get(x_192, 0); +lean_inc(x_221); +x_222 = lean_ctor_get(x_192, 1); +lean_inc(x_222); +if (lean_is_exclusive(x_192)) { + lean_ctor_release(x_192, 0); + lean_ctor_release(x_192, 1); + x_223 = x_192; } else { - lean_dec_ref(x_174); - x_205 = lean_box(0); + lean_dec_ref(x_192); + x_223 = lean_box(0); } -if (lean_is_scalar(x_205)) { - x_206 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_223)) { + x_224 = lean_alloc_ctor(1, 2, 0); } else { - x_206 = x_205; + x_224 = x_223; } -lean_ctor_set(x_206, 0, x_203); -lean_ctor_set(x_206, 1, x_204); -return x_206; +lean_ctor_set(x_224, 0, x_221); +lean_ctor_set(x_224, 1, x_222); +return x_224; } } else { -lean_object* x_207; lean_object* x_208; -lean_dec(x_172); -lean_dec(x_171); -lean_dec(x_170); -x_207 = lean_box(0); -x_208 = l_Lean_Meta_Grind_addCongrTable___lambda__2(x_1, x_169, x_207, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_145); +lean_object* x_225; lean_object* x_226; +lean_dec(x_190); +lean_dec(x_189); +lean_dec(x_188); +x_225 = lean_box(0); +x_226 = l_Lean_Meta_Grind_addCongrTable___lambda__2(x_1, x_187, x_225, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_157); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -return x_208; +return x_226; } } } @@ -3492,7 +3628,7 @@ return x_36; } else { -lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; uint8_t x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; +lean_object* x_37; 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; lean_object* x_47; 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_37 = lean_ctor_get(x_12, 1); x_38 = lean_ctor_get(x_14, 0); x_39 = lean_ctor_get(x_14, 1); @@ -3500,9 +3636,21 @@ x_40 = lean_ctor_get(x_14, 2); x_41 = lean_ctor_get(x_14, 3); x_42 = lean_ctor_get(x_14, 4); x_43 = lean_ctor_get(x_14, 5); -x_44 = lean_ctor_get_uint8(x_14, sizeof(void*)*8); +x_44 = lean_ctor_get_uint8(x_14, sizeof(void*)*14); x_45 = lean_ctor_get(x_14, 6); x_46 = lean_ctor_get(x_14, 7); +x_47 = lean_ctor_get(x_14, 8); +x_48 = lean_ctor_get(x_14, 9); +x_49 = lean_ctor_get(x_14, 10); +x_50 = lean_ctor_get(x_14, 11); +x_51 = lean_ctor_get(x_14, 12); +x_52 = lean_ctor_get(x_14, 13); +lean_inc(x_52); +lean_inc(x_51); +lean_inc(x_50); +lean_inc(x_49); +lean_inc(x_48); +lean_inc(x_47); lean_inc(x_46); lean_inc(x_45); lean_inc(x_43); @@ -3513,217 +3661,259 @@ lean_inc(x_39); lean_inc(x_38); lean_dec(x_14); lean_inc(x_42); -x_47 = l_Lean_PersistentHashMap_find_x3f___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__1(x_42, x_11); -if (lean_obj_tag(x_47) == 0) +x_53 = l_Lean_PersistentHashMap_find_x3f___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__1(x_42, x_11); +if (lean_obj_tag(x_53) == 0) { -lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; -x_48 = lean_box(0); +lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; +x_54 = lean_box(0); lean_ctor_set_tag(x_12, 1); -lean_ctor_set(x_12, 1, x_48); +lean_ctor_set(x_12, 1, x_54); lean_ctor_set(x_12, 0, x_1); -x_49 = l_Lean_PersistentHashMap_insert___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__4(x_42, x_11, x_12); -x_50 = lean_alloc_ctor(0, 8, 1); -lean_ctor_set(x_50, 0, x_38); -lean_ctor_set(x_50, 1, x_39); -lean_ctor_set(x_50, 2, x_40); -lean_ctor_set(x_50, 3, x_41); -lean_ctor_set(x_50, 4, x_49); -lean_ctor_set(x_50, 5, x_43); -lean_ctor_set(x_50, 6, x_45); -lean_ctor_set(x_50, 7, x_46); -lean_ctor_set_uint8(x_50, sizeof(void*)*8, x_44); -x_51 = lean_st_ref_set(x_2, x_50, x_37); -x_52 = lean_ctor_get(x_51, 1); -lean_inc(x_52); -if (lean_is_exclusive(x_51)) { - lean_ctor_release(x_51, 0); - lean_ctor_release(x_51, 1); - x_53 = x_51; +x_55 = l_Lean_PersistentHashMap_insert___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__4(x_42, x_11, x_12); +x_56 = lean_alloc_ctor(0, 14, 1); +lean_ctor_set(x_56, 0, x_38); +lean_ctor_set(x_56, 1, x_39); +lean_ctor_set(x_56, 2, x_40); +lean_ctor_set(x_56, 3, x_41); +lean_ctor_set(x_56, 4, x_55); +lean_ctor_set(x_56, 5, x_43); +lean_ctor_set(x_56, 6, x_45); +lean_ctor_set(x_56, 7, x_46); +lean_ctor_set(x_56, 8, x_47); +lean_ctor_set(x_56, 9, x_48); +lean_ctor_set(x_56, 10, x_49); +lean_ctor_set(x_56, 11, x_50); +lean_ctor_set(x_56, 12, x_51); +lean_ctor_set(x_56, 13, x_52); +lean_ctor_set_uint8(x_56, sizeof(void*)*14, x_44); +x_57 = lean_st_ref_set(x_2, x_56, x_37); +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_51); - x_53 = lean_box(0); + lean_dec_ref(x_57); + x_59 = lean_box(0); } -x_54 = lean_box(0); -if (lean_is_scalar(x_53)) { - x_55 = lean_alloc_ctor(0, 2, 0); +x_60 = lean_box(0); +if (lean_is_scalar(x_59)) { + x_61 = lean_alloc_ctor(0, 2, 0); } else { - x_55 = x_53; + x_61 = x_59; } -lean_ctor_set(x_55, 0, x_54); -lean_ctor_set(x_55, 1, x_52); -return x_55; +lean_ctor_set(x_61, 0, x_60); +lean_ctor_set(x_61, 1, x_58); +return x_61; } 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; -x_56 = lean_ctor_get(x_47, 0); -lean_inc(x_56); -lean_dec(x_47); +lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; +x_62 = lean_ctor_get(x_53, 0); +lean_inc(x_62); +lean_dec(x_53); lean_ctor_set_tag(x_12, 1); -lean_ctor_set(x_12, 1, x_56); +lean_ctor_set(x_12, 1, x_62); lean_ctor_set(x_12, 0, x_1); -x_57 = l_Lean_PersistentHashMap_insert___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__4(x_42, x_11, x_12); -x_58 = lean_alloc_ctor(0, 8, 1); -lean_ctor_set(x_58, 0, x_38); -lean_ctor_set(x_58, 1, x_39); -lean_ctor_set(x_58, 2, x_40); -lean_ctor_set(x_58, 3, x_41); -lean_ctor_set(x_58, 4, x_57); -lean_ctor_set(x_58, 5, x_43); -lean_ctor_set(x_58, 6, x_45); -lean_ctor_set(x_58, 7, x_46); -lean_ctor_set_uint8(x_58, sizeof(void*)*8, x_44); -x_59 = lean_st_ref_set(x_2, x_58, x_37); -x_60 = lean_ctor_get(x_59, 1); -lean_inc(x_60); -if (lean_is_exclusive(x_59)) { - lean_ctor_release(x_59, 0); - lean_ctor_release(x_59, 1); - x_61 = x_59; +x_63 = l_Lean_PersistentHashMap_insert___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__4(x_42, x_11, x_12); +x_64 = lean_alloc_ctor(0, 14, 1); +lean_ctor_set(x_64, 0, x_38); +lean_ctor_set(x_64, 1, x_39); +lean_ctor_set(x_64, 2, x_40); +lean_ctor_set(x_64, 3, x_41); +lean_ctor_set(x_64, 4, x_63); +lean_ctor_set(x_64, 5, x_43); +lean_ctor_set(x_64, 6, x_45); +lean_ctor_set(x_64, 7, x_46); +lean_ctor_set(x_64, 8, x_47); +lean_ctor_set(x_64, 9, x_48); +lean_ctor_set(x_64, 10, x_49); +lean_ctor_set(x_64, 11, x_50); +lean_ctor_set(x_64, 12, x_51); +lean_ctor_set(x_64, 13, x_52); +lean_ctor_set_uint8(x_64, sizeof(void*)*14, x_44); +x_65 = lean_st_ref_set(x_2, x_64, x_37); +x_66 = lean_ctor_get(x_65, 1); +lean_inc(x_66); +if (lean_is_exclusive(x_65)) { + lean_ctor_release(x_65, 0); + lean_ctor_release(x_65, 1); + x_67 = x_65; } else { - lean_dec_ref(x_59); - x_61 = lean_box(0); + lean_dec_ref(x_65); + x_67 = lean_box(0); } -x_62 = lean_box(0); -if (lean_is_scalar(x_61)) { - x_63 = lean_alloc_ctor(0, 2, 0); +x_68 = lean_box(0); +if (lean_is_scalar(x_67)) { + x_69 = lean_alloc_ctor(0, 2, 0); } else { - x_63 = x_61; + x_69 = x_67; } -lean_ctor_set(x_63, 0, x_62); -lean_ctor_set(x_63, 1, x_60); -return x_63; +lean_ctor_set(x_69, 0, x_68); +lean_ctor_set(x_69, 1, x_66); +return x_69; } } } 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; uint8_t x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; -x_64 = lean_ctor_get(x_12, 0); -x_65 = lean_ctor_get(x_12, 1); -lean_inc(x_65); -lean_inc(x_64); -lean_dec(x_12); -x_66 = lean_ctor_get(x_64, 0); -lean_inc(x_66); -x_67 = lean_ctor_get(x_64, 1); -lean_inc(x_67); -x_68 = lean_ctor_get(x_64, 2); -lean_inc(x_68); -x_69 = lean_ctor_get(x_64, 3); -lean_inc(x_69); -x_70 = lean_ctor_get(x_64, 4); -lean_inc(x_70); -x_71 = lean_ctor_get(x_64, 5); +lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; uint8_t x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; +x_70 = lean_ctor_get(x_12, 0); +x_71 = lean_ctor_get(x_12, 1); lean_inc(x_71); -x_72 = lean_ctor_get_uint8(x_64, sizeof(void*)*8); -x_73 = lean_ctor_get(x_64, 6); +lean_inc(x_70); +lean_dec(x_12); +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_64, 7); +x_74 = lean_ctor_get(x_70, 2); lean_inc(x_74); -if (lean_is_exclusive(x_64)) { - lean_ctor_release(x_64, 0); - lean_ctor_release(x_64, 1); - lean_ctor_release(x_64, 2); - lean_ctor_release(x_64, 3); - lean_ctor_release(x_64, 4); - lean_ctor_release(x_64, 5); - lean_ctor_release(x_64, 6); - lean_ctor_release(x_64, 7); - x_75 = x_64; +x_75 = lean_ctor_get(x_70, 3); +lean_inc(x_75); +x_76 = lean_ctor_get(x_70, 4); +lean_inc(x_76); +x_77 = lean_ctor_get(x_70, 5); +lean_inc(x_77); +x_78 = lean_ctor_get_uint8(x_70, sizeof(void*)*14); +x_79 = lean_ctor_get(x_70, 6); +lean_inc(x_79); +x_80 = lean_ctor_get(x_70, 7); +lean_inc(x_80); +x_81 = lean_ctor_get(x_70, 8); +lean_inc(x_81); +x_82 = lean_ctor_get(x_70, 9); +lean_inc(x_82); +x_83 = lean_ctor_get(x_70, 10); +lean_inc(x_83); +x_84 = lean_ctor_get(x_70, 11); +lean_inc(x_84); +x_85 = lean_ctor_get(x_70, 12); +lean_inc(x_85); +x_86 = lean_ctor_get(x_70, 13); +lean_inc(x_86); +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); + lean_ctor_release(x_70, 7); + lean_ctor_release(x_70, 8); + lean_ctor_release(x_70, 9); + lean_ctor_release(x_70, 10); + lean_ctor_release(x_70, 11); + lean_ctor_release(x_70, 12); + lean_ctor_release(x_70, 13); + x_87 = x_70; } else { - lean_dec_ref(x_64); - x_75 = lean_box(0); + lean_dec_ref(x_70); + x_87 = lean_box(0); } -lean_inc(x_70); -x_76 = l_Lean_PersistentHashMap_find_x3f___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__1(x_70, x_11); -if (lean_obj_tag(x_76) == 0) +lean_inc(x_76); +x_88 = l_Lean_PersistentHashMap_find_x3f___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__1(x_76, x_11); +if (lean_obj_tag(x_88) == 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_box(0); -x_78 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_78, 0, x_1); -lean_ctor_set(x_78, 1, x_77); -x_79 = l_Lean_PersistentHashMap_insert___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__4(x_70, x_11, x_78); -if (lean_is_scalar(x_75)) { - x_80 = lean_alloc_ctor(0, 8, 1); +lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; +x_89 = lean_box(0); +x_90 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_90, 0, x_1); +lean_ctor_set(x_90, 1, x_89); +x_91 = l_Lean_PersistentHashMap_insert___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__4(x_76, x_11, x_90); +if (lean_is_scalar(x_87)) { + x_92 = lean_alloc_ctor(0, 14, 1); } else { - x_80 = x_75; -} -lean_ctor_set(x_80, 0, x_66); -lean_ctor_set(x_80, 1, x_67); -lean_ctor_set(x_80, 2, x_68); -lean_ctor_set(x_80, 3, x_69); -lean_ctor_set(x_80, 4, x_79); -lean_ctor_set(x_80, 5, x_71); -lean_ctor_set(x_80, 6, x_73); -lean_ctor_set(x_80, 7, x_74); -lean_ctor_set_uint8(x_80, sizeof(void*)*8, x_72); -x_81 = lean_st_ref_set(x_2, x_80, x_65); -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; + x_92 = x_87; +} +lean_ctor_set(x_92, 0, x_72); +lean_ctor_set(x_92, 1, x_73); +lean_ctor_set(x_92, 2, x_74); +lean_ctor_set(x_92, 3, x_75); +lean_ctor_set(x_92, 4, x_91); +lean_ctor_set(x_92, 5, x_77); +lean_ctor_set(x_92, 6, x_79); +lean_ctor_set(x_92, 7, x_80); +lean_ctor_set(x_92, 8, x_81); +lean_ctor_set(x_92, 9, x_82); +lean_ctor_set(x_92, 10, x_83); +lean_ctor_set(x_92, 11, x_84); +lean_ctor_set(x_92, 12, x_85); +lean_ctor_set(x_92, 13, x_86); +lean_ctor_set_uint8(x_92, sizeof(void*)*14, x_78); +x_93 = lean_st_ref_set(x_2, x_92, x_71); +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_81); - x_83 = lean_box(0); + lean_dec_ref(x_93); + x_95 = lean_box(0); } -x_84 = lean_box(0); -if (lean_is_scalar(x_83)) { - x_85 = lean_alloc_ctor(0, 2, 0); +x_96 = lean_box(0); +if (lean_is_scalar(x_95)) { + x_97 = lean_alloc_ctor(0, 2, 0); } else { - x_85 = x_83; + x_97 = x_95; } -lean_ctor_set(x_85, 0, x_84); -lean_ctor_set(x_85, 1, x_82); -return x_85; +lean_ctor_set(x_97, 0, x_96); +lean_ctor_set(x_97, 1, x_94); +return x_97; } 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; -x_86 = lean_ctor_get(x_76, 0); -lean_inc(x_86); -lean_dec(x_76); -x_87 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_87, 0, x_1); -lean_ctor_set(x_87, 1, x_86); -x_88 = l_Lean_PersistentHashMap_insert___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__4(x_70, x_11, x_87); -if (lean_is_scalar(x_75)) { - x_89 = lean_alloc_ctor(0, 8, 1); +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_98 = lean_ctor_get(x_88, 0); +lean_inc(x_98); +lean_dec(x_88); +x_99 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_99, 0, x_1); +lean_ctor_set(x_99, 1, x_98); +x_100 = l_Lean_PersistentHashMap_insert___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__4(x_76, x_11, x_99); +if (lean_is_scalar(x_87)) { + x_101 = lean_alloc_ctor(0, 14, 1); } else { - x_89 = x_75; -} -lean_ctor_set(x_89, 0, x_66); -lean_ctor_set(x_89, 1, x_67); -lean_ctor_set(x_89, 2, x_68); -lean_ctor_set(x_89, 3, x_69); -lean_ctor_set(x_89, 4, x_88); -lean_ctor_set(x_89, 5, x_71); -lean_ctor_set(x_89, 6, x_73); -lean_ctor_set(x_89, 7, x_74); -lean_ctor_set_uint8(x_89, sizeof(void*)*8, x_72); -x_90 = lean_st_ref_set(x_2, x_89, x_65); -x_91 = lean_ctor_get(x_90, 1); -lean_inc(x_91); -if (lean_is_exclusive(x_90)) { - lean_ctor_release(x_90, 0); - lean_ctor_release(x_90, 1); - x_92 = x_90; + x_101 = x_87; +} +lean_ctor_set(x_101, 0, x_72); +lean_ctor_set(x_101, 1, x_73); +lean_ctor_set(x_101, 2, x_74); +lean_ctor_set(x_101, 3, x_75); +lean_ctor_set(x_101, 4, x_100); +lean_ctor_set(x_101, 5, x_77); +lean_ctor_set(x_101, 6, x_79); +lean_ctor_set(x_101, 7, x_80); +lean_ctor_set(x_101, 8, x_81); +lean_ctor_set(x_101, 9, x_82); +lean_ctor_set(x_101, 10, x_83); +lean_ctor_set(x_101, 11, x_84); +lean_ctor_set(x_101, 12, x_85); +lean_ctor_set(x_101, 13, x_86); +lean_ctor_set_uint8(x_101, sizeof(void*)*14, x_78); +x_102 = lean_st_ref_set(x_2, x_101, x_71); +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_90); - x_92 = lean_box(0); + lean_dec_ref(x_102); + x_104 = lean_box(0); } -x_93 = lean_box(0); -if (lean_is_scalar(x_92)) { - x_94 = 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_94 = x_92; + x_106 = x_104; } -lean_ctor_set(x_94, 0, x_93); -lean_ctor_set(x_94, 1, x_91); -return x_94; +lean_ctor_set(x_106, 0, x_105); +lean_ctor_set(x_106, 1, x_103); +return x_106; } } } @@ -3799,541 +3989,441 @@ lean_dec(x_2); return x_11; } } -static lean_object* _init_l_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_instMonadMetaM; -x_2 = l_ReaderT_instMonad___rarg(x_1); -return x_2; -} -} -static lean_object* _init_l_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__1; -x_2 = l_ReaderT_instMonad___rarg(x_1); -return x_2; -} -} -static lean_object* _init_l_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__3() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__2; -x_2 = l_ReaderT_instMonad___rarg(x_1); -return x_2; -} -} -static lean_object* _init_l_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__4() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__3; -x_2 = l_ReaderT_instMonad___rarg(x_1); -return x_2; -} -} -static lean_object* _init_l_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__5() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__4; -x_2 = l_instInhabitedPUnit; -x_3 = l_instInhabitedOfMonad___rarg(x_1, x_2); -return x_3; -} -} -LEAN_EXPORT lean_object* l_panic___at_Lean_Meta_Grind_internalize___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_11 = l_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__5; -x_12 = lean_panic_fn(x_11, x_1); -x_13 = lean_apply_9(x_12, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_13; -} -} -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_internalize___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___spec__1(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { -lean_object* x_19; uint8_t x_20; -x_19 = lean_ctor_get(x_5, 1); -x_20 = lean_nat_dec_lt(x_7, x_19); -if (x_20 == 0) +uint8_t x_14; +x_14 = lean_usize_dec_lt(x_3, x_2); +if (x_14 == 0) { -lean_object* x_21; -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); +lean_object* 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_2); +lean_dec(x_6); +lean_dec(x_5); lean_dec(x_1); -x_21 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_21, 0, x_6); -lean_ctor_set(x_21, 1, x_18); -return x_21; +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_4); +lean_ctor_set(x_15, 1, x_13); +return x_15; } else { -lean_object* x_22; lean_object* x_23; -lean_dec(x_6); -x_22 = lean_array_fget(x_3, x_7); -lean_inc(x_17); -lean_inc(x_16); -lean_inc(x_15); -lean_inc(x_14); -lean_inc(x_13); +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_16 = lean_array_uget(x_4, x_3); +x_17 = lean_unsigned_to_nat(0u); +x_18 = lean_array_uset(x_4, x_3, x_17); lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); -lean_inc(x_2); -lean_inc(x_22); -x_23 = l_Lean_Meta_Grind_internalize(x_22, x_2, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); -if (lean_obj_tag(x_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, 1); -lean_inc(x_24); -lean_dec(x_23); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); lean_inc(x_1); -x_25 = l_Lean_Meta_Grind_registerParent(x_1, x_22, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_24); -lean_dec(x_22); -x_26 = lean_ctor_get(x_25, 1); -lean_inc(x_26); -lean_dec(x_25); -x_27 = lean_ctor_get(x_5, 2); -x_28 = lean_nat_add(x_7, x_27); -lean_dec(x_7); -x_29 = lean_box(0); -x_6 = x_29; -x_7 = x_28; -x_8 = lean_box(0); -x_9 = lean_box(0); -x_18 = x_26; +x_19 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern(x_16, x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_19) == 0) +{ +lean_object* x_20; lean_object* x_21; size_t x_22; size_t x_23; lean_object* x_24; +x_20 = lean_ctor_get(x_19, 0); +lean_inc(x_20); +x_21 = lean_ctor_get(x_19, 1); +lean_inc(x_21); +lean_dec(x_19); +x_22 = 1; +x_23 = lean_usize_add(x_3, x_22); +x_24 = lean_array_uset(x_18, x_3, x_20); +x_3 = x_23; +x_4 = x_24; +x_13 = x_21; goto _start; } else { -uint8_t x_31; -lean_dec(x_22); -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); +uint8_t x_26; +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_2); +lean_dec(x_6); +lean_dec(x_5); lean_dec(x_1); -x_31 = !lean_is_exclusive(x_23); -if (x_31 == 0) +x_26 = !lean_is_exclusive(x_19); +if (x_26 == 0) { -return x_23; +return x_19; } else { -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); -return x_34; +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_19, 0); +x_28 = lean_ctor_get(x_19, 1); +lean_inc(x_28); +lean_inc(x_27); +lean_dec(x_19); +x_29 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_29, 0, x_27); +lean_ctor_set(x_29, 1, x_28); +return x_29; } } } } } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___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_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___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) { _start: { -lean_object* x_13; -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_1); -x_13 = l_Lean_Meta_Grind_mkENode(x_1, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_13) == 0) +if (lean_obj_tag(x_2) == 5) { -lean_object* x_14; lean_object* x_15; -x_14 = lean_ctor_get(x_13, 1); +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_14 = lean_ctor_get(x_2, 0); lean_inc(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_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_1); -x_15 = l_Lean_Meta_Grind_addCongrTable(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_14); -if (lean_obj_tag(x_15) == 0) -{ -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_16 = lean_ctor_get(x_15, 1); -lean_inc(x_16); -lean_dec(x_15); -lean_inc(x_1); -x_17 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_16); -x_18 = lean_ctor_get(x_17, 1); -lean_inc(x_18); -lean_dec(x_17); -x_19 = l_Lean_Meta_Grind_propagateUp(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_18); -return x_19; +x_15 = lean_ctor_get(x_2, 1); +lean_inc(x_15); +lean_dec(x_2); +x_16 = lean_array_set(x_3, x_4, x_15); +x_17 = lean_unsigned_to_nat(1u); +x_18 = lean_nat_sub(x_4, x_17); +lean_dec(x_4); +x_2 = x_14; +x_3 = x_16; +x_4 = x_18; +goto _start; } else { -uint8_t x_20; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +size_t x_20; size_t x_21; lean_object* x_22; lean_dec(x_4); -lean_dec(x_1); -x_20 = !lean_is_exclusive(x_15); -if (x_20 == 0) +x_20 = lean_array_size(x_3); +x_21 = 0; +x_22 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___spec__1(x_1, x_20, x_21, x_3, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_22) == 0) { -return x_15; +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 = l_Lean_mkAppN(x_2, x_24); +lean_dec(x_24); +lean_ctor_set(x_22, 0, x_25); +return x_22; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_15, 0); -x_22 = lean_ctor_get(x_15, 1); -lean_inc(x_22); -lean_inc(x_21); -lean_dec(x_15); -x_23 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_23, 0, x_21); -lean_ctor_set(x_23, 1, x_22); -return x_23; -} +lean_object* x_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 = l_Lean_mkAppN(x_2, x_26); +lean_dec(x_26); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_27); +return x_29; } } else { -uint8_t x_24; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_1); -x_24 = !lean_is_exclusive(x_13); -if (x_24 == 0) +uint8_t x_30; +lean_dec(x_2); +x_30 = !lean_is_exclusive(x_22); +if (x_30 == 0) { -return x_13; +return x_22; } else { -lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_25 = lean_ctor_get(x_13, 0); -x_26 = lean_ctor_get(x_13, 1); -lean_inc(x_26); -lean_inc(x_25); -lean_dec(x_13); -x_27 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_27, 0, x_25); -lean_ctor_set(x_27, 1, x_26); -return x_27; -} +lean_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_Expr_withAppAux___at_Lean_Meta_Grind_internalize___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) { -_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; -x_15 = lean_array_get_size(x_1); -x_16 = lean_unsigned_to_nat(0u); -x_17 = lean_unsigned_to_nat(1u); -x_18 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_18, 0, x_16); -lean_ctor_set(x_18, 1, x_15); -lean_ctor_set(x_18, 2, x_17); -x_19 = 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); -lean_inc(x_7); -lean_inc(x_6); -x_20 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_internalize___spec__2(x_2, x_3, x_1, x_18, x_18, x_19, x_16, lean_box(0), lean_box(0), x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -lean_dec(x_18); -if (lean_obj_tag(x_20) == 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 = lean_apply_10(x_4, x_19, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_21); -return x_22; } -else -{ -uint8_t 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_6); -lean_dec(x_4); -x_23 = !lean_is_exclusive(x_20); -if (x_23 == 0) -{ -return x_20; } -else +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__1() { +_start: { -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_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_object* x_1; lean_object* x_2; +x_1 = l_Lean_levelZero; +x_2 = l_Lean_Expr_sort___override(x_1); +return x_2; } } -static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__1() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__2() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean", 4, 4); +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_unfoldReducible___lambda__2), 6, 0); return x_1; } } -static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__2() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__3() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("Grind", 5, 5); +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_unfoldReducible___lambda__3___boxed), 6, 0); return x_1; } } -static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__3() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__4() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("nestedProof", 11, 11); +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_normalizeLevels___lambda__1___boxed), 4, 0); return x_1; } } -static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__4() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__5() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__1; -x_2 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__2; -x_3 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__3; -x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); -return x_4; +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_eraseIrrelevantMData___lambda__2___boxed), 4, 0); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, 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: { -switch (lean_obj_tag(x_3)) { -case 0: -{ -lean_object* x_15; lean_object* x_16; lean_object* x_31; uint8_t x_32; -lean_dec(x_5); -lean_inc(x_2); -lean_inc(x_1); -x_15 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1___boxed), 12, 2); -lean_closure_set(x_15, 0, x_1); -lean_closure_set(x_15, 1, x_2); -x_31 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__4; -x_32 = l_Lean_Expr_isConstOf(x_3, x_31); -if (x_32 == 0) +uint8_t x_12; +x_12 = l_Lean_Expr_isBVar(x_1); +if (x_12 == 0) { -lean_object* x_33; -x_33 = lean_box(0); -x_16 = x_33; -goto block_30; -} -else +lean_object* x_13; uint8_t x_14; +x_13 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare; +x_14 = lean_expr_eqv(x_1, x_13); +if (x_14 == 0) { -lean_object* x_34; lean_object* x_35; uint8_t x_36; -x_34 = lean_array_get_size(x_4); -x_35 = lean_unsigned_to_nat(2u); -x_36 = lean_nat_dec_eq(x_34, x_35); -if (x_36 == 0) +lean_object* x_15; +x_15 = l_Lean_Meta_Grind_groundPattern_x3f(x_1); +if (lean_obj_tag(x_15) == 0) { -lean_object* x_37; -lean_dec(x_34); -x_37 = lean_box(0); -x_16 = x_37; -goto block_30; +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_16 = lean_unsigned_to_nat(0u); +x_17 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_16); +x_18 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__1; +lean_inc(x_17); +x_19 = lean_mk_array(x_17, x_18); +x_20 = lean_unsigned_to_nat(1u); +x_21 = lean_nat_sub(x_17, x_20); +lean_dec(x_17); +x_22 = l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___spec__2(x_2, x_1, x_19, x_21, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +return x_22; } else { -lean_object* x_38; uint8_t x_39; +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +lean_dec(x_1); +x_23 = lean_ctor_get(x_15, 0); +lean_inc(x_23); lean_dec(x_15); -lean_dec(x_3); -x_38 = lean_unsigned_to_nat(0u); -x_39 = lean_nat_dec_lt(x_38, x_34); -lean_dec(x_34); -if (x_39 == 0) +x_24 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__2; +x_25 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__3; +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +x_26 = l_Lean_Core_transform___at_Lean_Meta_Grind_unfoldReducible___spec__1(x_23, x_24, x_25, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_26) == 0) { -lean_object* x_40; lean_object* x_41; lean_object* x_42; -lean_dec(x_4); -x_40 = l_Lean_instInhabitedExpr; -x_41 = l_outOfBounds___rarg(x_40); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_27 = lean_ctor_get(x_26, 0); +lean_inc(x_27); +x_28 = lean_ctor_get(x_26, 1); +lean_inc(x_28); +lean_dec(x_26); +x_29 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__4; +x_30 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__5; +lean_inc(x_10); +lean_inc(x_9); +x_31 = l_Lean_Core_transform___at_Lean_Core_betaReduce___spec__1(x_27, x_29, x_30, x_9, x_10, x_28); +if (lean_obj_tag(x_31) == 0) +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_31, 0); +lean_inc(x_32); +x_33 = lean_ctor_get(x_31, 1); +lean_inc(x_33); +lean_dec(x_31); 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_41); -x_42 = l_Lean_Meta_Grind_internalize(x_41, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_42) == 0) +x_34 = l_Lean_Meta_Grind_canon(x_32, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_33); +if (lean_obj_tag(x_34) == 0) { -lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_43 = lean_ctor_get(x_42, 1); -lean_inc(x_43); +lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_35 = lean_ctor_get(x_34, 0); +lean_inc(x_35); +x_36 = lean_ctor_get(x_34, 1); +lean_inc(x_36); +lean_dec(x_34); +x_37 = l_Lean_Meta_Grind_shareCommon(x_35, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_36); +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); +lean_inc(x_38); +x_40 = l_Lean_Meta_Grind_internalize(x_38, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_39); +if (lean_obj_tag(x_40) == 0) +{ +uint8_t x_41; +x_41 = !lean_is_exclusive(x_40); +if (x_41 == 0) +{ +lean_object* x_42; lean_object* x_43; +x_42 = lean_ctor_get(x_40, 0); lean_dec(x_42); -lean_inc(x_1); -x_44 = l_Lean_Meta_Grind_registerParent(x_1, x_41, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_43); -lean_dec(x_41); -x_45 = lean_ctor_get(x_44, 0); -lean_inc(x_45); -x_46 = lean_ctor_get(x_44, 1); -lean_inc(x_46); -lean_dec(x_44); -x_47 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_45, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_46); -lean_dec(x_45); -return x_47; +x_43 = l_Lean_Meta_Grind_mkGroundPattern(x_38); +lean_ctor_set(x_40, 0, x_43); +return x_40; } else { -uint8_t x_48; -lean_dec(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_7); -lean_dec(x_6); -lean_dec(x_2); -lean_dec(x_1); -x_48 = !lean_is_exclusive(x_42); -if (x_48 == 0) -{ -return x_42; +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 = l_Lean_Meta_Grind_mkGroundPattern(x_38); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_44); +return x_46; +} } else { -lean_object* x_49; lean_object* x_50; lean_object* x_51; -x_49 = lean_ctor_get(x_42, 0); -x_50 = lean_ctor_get(x_42, 1); -lean_inc(x_50); +uint8_t x_47; +lean_dec(x_38); +x_47 = !lean_is_exclusive(x_40); +if (x_47 == 0) +{ +return x_40; +} +else +{ +lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_48 = lean_ctor_get(x_40, 0); +x_49 = lean_ctor_get(x_40, 1); lean_inc(x_49); -lean_dec(x_42); -x_51 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_51, 0, x_49); -lean_ctor_set(x_51, 1, x_50); -return x_51; +lean_inc(x_48); +lean_dec(x_40); +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_52; lean_object* x_53; -x_52 = lean_array_fget(x_4, x_38); +uint8_t x_51; +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_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_2); +lean_dec(x_3); +lean_dec(x_2); +x_51 = !lean_is_exclusive(x_34); +if (x_51 == 0) +{ +return x_34; +} +else +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; +x_52 = lean_ctor_get(x_34, 0); +x_53 = lean_ctor_get(x_34, 1); +lean_inc(x_53); lean_inc(x_52); -x_53 = l_Lean_Meta_Grind_internalize(x_52, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_53) == 0) +lean_dec(x_34); +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_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; -x_54 = lean_ctor_get(x_53, 1); -lean_inc(x_54); -lean_dec(x_53); -lean_inc(x_1); -x_55 = l_Lean_Meta_Grind_registerParent(x_1, x_52, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_54); -lean_dec(x_52); -x_56 = lean_ctor_get(x_55, 0); -lean_inc(x_56); -x_57 = lean_ctor_get(x_55, 1); +uint8_t x_55; +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_55 = !lean_is_exclusive(x_31); +if (x_55 == 0) +{ +return x_31; +} +else +{ +lean_object* x_56; lean_object* x_57; lean_object* x_58; +x_56 = lean_ctor_get(x_31, 0); +x_57 = lean_ctor_get(x_31, 1); lean_inc(x_57); -lean_dec(x_55); -x_58 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_56, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_57); -lean_dec(x_56); +lean_inc(x_56); +lean_dec(x_31); +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; -lean_dec(x_52); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -x_59 = !lean_is_exclusive(x_53); +x_59 = !lean_is_exclusive(x_26); if (x_59 == 0) { -return x_53; +return x_26; } else { lean_object* x_60; lean_object* x_61; lean_object* x_62; -x_60 = lean_ctor_get(x_53, 0); -x_61 = lean_ctor_get(x_53, 1); +x_60 = lean_ctor_get(x_26, 0); +x_61 = lean_ctor_get(x_26, 1); lean_inc(x_61); lean_inc(x_60); -lean_dec(x_53); +lean_dec(x_26); x_62 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_62, 0, x_60); lean_ctor_set(x_62, 1, x_61); @@ -4342,470 +4432,317 @@ return x_62; } } } -} -block_30: -{ -uint8_t x_17; -lean_dec(x_16); -x_17 = l_Lean_Expr_isConst(x_3); -if (x_17 == 0) -{ -lean_object* x_18; -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_2); -lean_inc(x_3); -x_18 = l_Lean_Meta_Grind_internalize(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_18) == 0) -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_19 = lean_ctor_get(x_18, 1); -lean_inc(x_19); -lean_dec(x_18); -lean_inc(x_1); -x_20 = l_Lean_Meta_Grind_registerParent(x_1, x_3, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_19); -lean_dec(x_3); -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_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(x_4, x_1, x_2, x_15, x_21, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_22); -lean_dec(x_21); -lean_dec(x_4); -return x_23; -} else { -uint8_t x_24; -lean_dec(x_15); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); +lean_object* x_63; 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_24 = !lean_is_exclusive(x_18); -if (x_24 == 0) -{ -return x_18; -} -else -{ -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_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; -} +x_63 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_63, 0, x_1); +lean_ctor_set(x_63, 1, x_11); +return x_63; } } else { -lean_object* x_28; lean_object* x_29; -lean_dec(x_3); -x_28 = lean_box(0); -x_29 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(x_4, x_1, x_2, x_15, x_28, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +lean_object* x_64; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); lean_dec(x_4); -return x_29; +lean_dec(x_3); +lean_dec(x_2); +x_64 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_64, 0, x_1); +lean_ctor_set(x_64, 1, x_11); +return x_64; } } } -case 1: -{ -lean_object* x_63; lean_object* x_64; lean_object* x_79; uint8_t x_80; -lean_dec(x_5); -lean_inc(x_2); -lean_inc(x_1); -x_63 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1___boxed), 12, 2); -lean_closure_set(x_63, 0, x_1); -lean_closure_set(x_63, 1, x_2); -x_79 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__4; -x_80 = l_Lean_Expr_isConstOf(x_3, x_79); -if (x_80 == 0) +static lean_object* _init_l_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__1() { +_start: { -lean_object* x_81; -x_81 = lean_box(0); -x_64 = x_81; -goto block_78; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_instMonadMetaM; +x_2 = l_ReaderT_instMonad___rarg(x_1); +return x_2; } -else +} +static lean_object* _init_l_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__2() { +_start: { -lean_object* x_82; lean_object* x_83; uint8_t x_84; -x_82 = lean_array_get_size(x_4); -x_83 = lean_unsigned_to_nat(2u); -x_84 = lean_nat_dec_eq(x_82, x_83); -if (x_84 == 0) +lean_object* x_1; lean_object* x_2; +x_1 = l_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__1; +x_2 = l_ReaderT_instMonad___rarg(x_1); +return x_2; +} +} +static lean_object* _init_l_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__3() { +_start: { -lean_object* x_85; -lean_dec(x_82); -x_85 = lean_box(0); -x_64 = x_85; -goto block_78; +lean_object* x_1; lean_object* x_2; +x_1 = l_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__2; +x_2 = l_ReaderT_instMonad___rarg(x_1); +return x_2; } -else +} +static lean_object* _init_l_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__4() { +_start: { -lean_object* x_86; uint8_t x_87; -lean_dec(x_63); -lean_dec(x_3); -x_86 = lean_unsigned_to_nat(0u); -x_87 = lean_nat_dec_lt(x_86, x_82); -lean_dec(x_82); -if (x_87 == 0) +lean_object* x_1; lean_object* x_2; +x_1 = l_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__3; +x_2 = l_ReaderT_instMonad___rarg(x_1); +return x_2; +} +} +static lean_object* _init_l_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__5() { +_start: { -lean_object* x_88; lean_object* x_89; lean_object* x_90; -lean_dec(x_4); -x_88 = l_Lean_instInhabitedExpr; -x_89 = l_outOfBounds___rarg(x_88); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_2); -lean_inc(x_89); -x_90 = l_Lean_Meta_Grind_internalize(x_89, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_90) == 0) +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__4; +x_2 = l_instInhabitedPUnit; +x_3 = l_instInhabitedOfMonad___rarg(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_panic___at_Lean_Meta_Grind_internalize___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_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; -x_91 = lean_ctor_get(x_90, 1); -lean_inc(x_91); -lean_dec(x_90); -lean_inc(x_1); -x_92 = l_Lean_Meta_Grind_registerParent(x_1, x_89, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_91); -lean_dec(x_89); -x_93 = lean_ctor_get(x_92, 0); -lean_inc(x_93); -x_94 = lean_ctor_get(x_92, 1); -lean_inc(x_94); -lean_dec(x_92); -x_95 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_93, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_94); -lean_dec(x_93); -return x_95; +lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_11 = l_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__5; +x_12 = lean_panic_fn(x_11, x_1); +x_13 = lean_apply_9(x_12, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_13; } -else +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_internalize___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18) { +_start: { -uint8_t x_96; -lean_dec(x_89); +lean_object* x_19; uint8_t x_20; +x_19 = lean_ctor_get(x_5, 1); +x_20 = lean_nat_dec_lt(x_7, x_19); +if (x_20 == 0) +{ +lean_object* x_21; +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); lean_dec(x_7); -lean_dec(x_6); lean_dec(x_2); lean_dec(x_1); -x_96 = !lean_is_exclusive(x_90); -if (x_96 == 0) -{ -return x_90; +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_6); +lean_ctor_set(x_21, 1, x_18); +return x_21; } else { -lean_object* x_97; lean_object* x_98; lean_object* x_99; -x_97 = lean_ctor_get(x_90, 0); -x_98 = lean_ctor_get(x_90, 1); -lean_inc(x_98); -lean_inc(x_97); -lean_dec(x_90); -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 -{ -lean_object* x_100; lean_object* x_101; -x_100 = lean_array_fget(x_4, x_86); -lean_dec(x_4); +lean_object* x_22; lean_object* x_23; +lean_dec(x_6); +x_22 = lean_array_fget(x_3, x_7); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); lean_inc(x_2); -lean_inc(x_100); -x_101 = l_Lean_Meta_Grind_internalize(x_100, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_101) == 0) +lean_inc(x_22); +x_23 = l_Lean_Meta_Grind_internalize(x_22, x_2, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +if (lean_obj_tag(x_23) == 0) { -lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; -x_102 = lean_ctor_get(x_101, 1); -lean_inc(x_102); -lean_dec(x_101); +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); lean_inc(x_1); -x_103 = l_Lean_Meta_Grind_registerParent(x_1, x_100, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_102); -lean_dec(x_100); -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_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_104, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_105); -lean_dec(x_104); -return x_106; +x_25 = l_Lean_Meta_Grind_registerParent(x_1, x_22, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_24); +lean_dec(x_22); +x_26 = lean_ctor_get(x_25, 1); +lean_inc(x_26); +lean_dec(x_25); +x_27 = lean_ctor_get(x_5, 2); +x_28 = lean_nat_add(x_7, x_27); +lean_dec(x_7); +x_29 = lean_box(0); +x_6 = x_29; +x_7 = x_28; +x_8 = lean_box(0); +x_9 = lean_box(0); +x_18 = x_26; +goto _start; } else { -uint8_t x_107; -lean_dec(x_100); +uint8_t x_31; +lean_dec(x_22); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); lean_dec(x_7); -lean_dec(x_6); lean_dec(x_2); lean_dec(x_1); -x_107 = !lean_is_exclusive(x_101); -if (x_107 == 0) +x_31 = !lean_is_exclusive(x_23); +if (x_31 == 0) { -return x_101; +return x_23; } else { -lean_object* x_108; lean_object* x_109; lean_object* x_110; -x_108 = lean_ctor_get(x_101, 0); -x_109 = lean_ctor_get(x_101, 1); -lean_inc(x_109); -lean_inc(x_108); -lean_dec(x_101); -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_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); +return x_34; } } } } } -block_78: +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___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) { +_start: { -uint8_t x_65; -lean_dec(x_64); -x_65 = l_Lean_Expr_isConst(x_3); -if (x_65 == 0) +lean_object* x_13; +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_1); +x_13 = l_Lean_Meta_Grind_mkENode(x_1, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_13) == 0) { -lean_object* x_66; -lean_inc(x_13); -lean_inc(x_12); +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_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); -lean_inc(x_3); -x_66 = l_Lean_Meta_Grind_internalize(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_66) == 0) +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_1); +x_15 = l_Lean_Meta_Grind_addCongrTable(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_14); +if (lean_obj_tag(x_15) == 0) { -lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; -x_67 = lean_ctor_get(x_66, 1); -lean_inc(x_67); -lean_dec(x_66); +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_16 = lean_ctor_get(x_15, 1); +lean_inc(x_16); +lean_dec(x_15); lean_inc(x_1); -x_68 = l_Lean_Meta_Grind_registerParent(x_1, x_3, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_67); -lean_dec(x_3); -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_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(x_4, x_1, x_2, x_63, x_69, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_70); -lean_dec(x_69); -lean_dec(x_4); -return x_71; +x_17 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_16); +x_18 = lean_ctor_get(x_17, 1); +lean_inc(x_18); +lean_dec(x_17); +x_19 = l_Lean_Meta_Grind_propagateUp(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_18); +return x_19; } else { -uint8_t x_72; -lean_dec(x_63); -lean_dec(x_13); -lean_dec(x_12); +uint8_t x_20; lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); +lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); lean_dec(x_1); -x_72 = !lean_is_exclusive(x_66); -if (x_72 == 0) -{ -return x_66; -} -else +x_20 = !lean_is_exclusive(x_15); +if (x_20 == 0) { -lean_object* x_73; lean_object* x_74; lean_object* x_75; -x_73 = lean_ctor_get(x_66, 0); -x_74 = lean_ctor_get(x_66, 1); -lean_inc(x_74); -lean_inc(x_73); -lean_dec(x_66); -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; -} -} +return x_15; } else { -lean_object* x_76; lean_object* x_77; -lean_dec(x_3); -x_76 = lean_box(0); -x_77 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(x_4, x_1, x_2, x_63, x_76, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -lean_dec(x_4); -return x_77; -} -} -} -case 2: -{ -lean_object* x_111; lean_object* x_112; lean_object* x_127; uint8_t x_128; -lean_dec(x_5); -lean_inc(x_2); -lean_inc(x_1); -x_111 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1___boxed), 12, 2); -lean_closure_set(x_111, 0, x_1); -lean_closure_set(x_111, 1, x_2); -x_127 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__4; -x_128 = l_Lean_Expr_isConstOf(x_3, x_127); -if (x_128 == 0) -{ -lean_object* x_129; -x_129 = lean_box(0); -x_112 = x_129; -goto block_126; +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_15, 0); +x_22 = lean_ctor_get(x_15, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_15); +x_23 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_23, 0, x_21); +lean_ctor_set(x_23, 1, x_22); +return x_23; } -else -{ -lean_object* x_130; lean_object* x_131; uint8_t x_132; -x_130 = lean_array_get_size(x_4); -x_131 = lean_unsigned_to_nat(2u); -x_132 = lean_nat_dec_eq(x_130, x_131); -if (x_132 == 0) -{ -lean_object* x_133; -lean_dec(x_130); -x_133 = lean_box(0); -x_112 = x_133; -goto block_126; } -else -{ -lean_object* x_134; uint8_t x_135; -lean_dec(x_111); -lean_dec(x_3); -x_134 = lean_unsigned_to_nat(0u); -x_135 = lean_nat_dec_lt(x_134, x_130); -lean_dec(x_130); -if (x_135 == 0) -{ -lean_object* x_136; lean_object* x_137; lean_object* x_138; -lean_dec(x_4); -x_136 = l_Lean_instInhabitedExpr; -x_137 = l_outOfBounds___rarg(x_136); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_2); -lean_inc(x_137); -x_138 = l_Lean_Meta_Grind_internalize(x_137, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_138) == 0) -{ -lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; -x_139 = lean_ctor_get(x_138, 1); -lean_inc(x_139); -lean_dec(x_138); -lean_inc(x_1); -x_140 = l_Lean_Meta_Grind_registerParent(x_1, x_137, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_139); -lean_dec(x_137); -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_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_141, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_142); -lean_dec(x_141); -return x_143; } else { -uint8_t x_144; -lean_dec(x_137); -lean_dec(x_13); -lean_dec(x_12); +uint8_t x_24; lean_dec(x_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_5); +lean_dec(x_4); lean_dec(x_1); -x_144 = !lean_is_exclusive(x_138); -if (x_144 == 0) +x_24 = !lean_is_exclusive(x_13); +if (x_24 == 0) { -return x_138; +return x_13; } else { -lean_object* x_145; lean_object* x_146; lean_object* x_147; -x_145 = lean_ctor_get(x_138, 0); -x_146 = lean_ctor_get(x_138, 1); -lean_inc(x_146); -lean_inc(x_145); -lean_dec(x_138); -x_147 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_147, 0, x_145); -lean_ctor_set(x_147, 1, x_146); -return x_147; +lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_25 = lean_ctor_get(x_13, 0); +x_26 = lean_ctor_get(x_13, 1); +lean_inc(x_26); +lean_inc(x_25); +lean_dec(x_13); +x_27 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +return x_27; } } } -else +} +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___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) { +_start: { -lean_object* x_148; lean_object* x_149; -x_148 = lean_array_fget(x_4, x_134); -lean_dec(x_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; +x_15 = lean_array_get_size(x_1); +x_16 = lean_unsigned_to_nat(0u); +x_17 = lean_unsigned_to_nat(1u); +x_18 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_18, 0, x_16); +lean_ctor_set(x_18, 1, x_15); +lean_ctor_set(x_18, 2, x_17); +x_19 = lean_box(0); lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); @@ -4814,31 +4751,20 @@ lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -lean_inc(x_2); -lean_inc(x_148); -x_149 = l_Lean_Meta_Grind_internalize(x_148, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_149) == 0) +x_20 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_internalize___spec__2(x_2, x_3, x_1, x_18, x_18, x_19, x_16, lean_box(0), lean_box(0), x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +lean_dec(x_18); +if (lean_obj_tag(x_20) == 0) { -lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; -x_150 = lean_ctor_get(x_149, 1); -lean_inc(x_150); -lean_dec(x_149); -lean_inc(x_1); -x_151 = l_Lean_Meta_Grind_registerParent(x_1, x_148, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_150); -lean_dec(x_148); -x_152 = lean_ctor_get(x_151, 0); -lean_inc(x_152); -x_153 = lean_ctor_get(x_151, 1); -lean_inc(x_153); -lean_dec(x_151); -x_154 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_152, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_153); -lean_dec(x_152); -return x_154; +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 = lean_apply_10(x_4, x_19, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_21); +return x_22; } else { -uint8_t x_155; -lean_dec(x_148); +uint8_t x_23; lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); @@ -4847,161 +4773,113 @@ 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_155 = !lean_is_exclusive(x_149); -if (x_155 == 0) +lean_dec(x_4); +x_23 = !lean_is_exclusive(x_20); +if (x_23 == 0) { -return x_149; +return x_20; } else { -lean_object* x_156; lean_object* x_157; lean_object* x_158; -x_156 = lean_ctor_get(x_149, 0); -x_157 = lean_ctor_get(x_149, 1); -lean_inc(x_157); -lean_inc(x_156); -lean_dec(x_149); -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_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_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; } } } } -block_126: -{ -uint8_t x_113; -lean_dec(x_112); -x_113 = l_Lean_Expr_isConst(x_3); -if (x_113 == 0) -{ -lean_object* x_114; -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_2); -lean_inc(x_3); -x_114 = l_Lean_Meta_Grind_internalize(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_114) == 0) +static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__1() { +_start: { -lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; -x_115 = lean_ctor_get(x_114, 1); -lean_inc(x_115); -lean_dec(x_114); -lean_inc(x_1); -x_116 = l_Lean_Meta_Grind_registerParent(x_1, x_3, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_115); -lean_dec(x_3); -x_117 = lean_ctor_get(x_116, 0); -lean_inc(x_117); -x_118 = lean_ctor_get(x_116, 1); -lean_inc(x_118); -lean_dec(x_116); -x_119 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(x_4, x_1, x_2, x_111, x_117, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_118); -lean_dec(x_117); -lean_dec(x_4); -return x_119; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean", 4, 4); +return x_1; } -else -{ -uint8_t x_120; -lean_dec(x_111); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_120 = !lean_is_exclusive(x_114); -if (x_120 == 0) -{ -return x_114; } -else +static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__2() { +_start: { -lean_object* x_121; lean_object* x_122; lean_object* x_123; -x_121 = lean_ctor_get(x_114, 0); -x_122 = lean_ctor_get(x_114, 1); -lean_inc(x_122); -lean_inc(x_121); -lean_dec(x_114); -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_1; +x_1 = lean_mk_string_unchecked("Grind", 5, 5); +return x_1; } } -else +static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__3() { +_start: { -lean_object* x_124; lean_object* x_125; -lean_dec(x_3); -x_124 = lean_box(0); -x_125 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(x_4, x_1, x_2, x_111, x_124, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -lean_dec(x_4); -return x_125; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("nestedProof", 11, 11); +return x_1; +} } +static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__1; +x_2 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__2; +x_3 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__3; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; } } -case 3: +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: { -lean_object* x_159; lean_object* x_160; lean_object* x_175; uint8_t x_176; +switch (lean_obj_tag(x_3)) { +case 0: +{ +lean_object* x_15; lean_object* x_16; lean_object* x_28; uint8_t x_29; lean_dec(x_5); lean_inc(x_2); lean_inc(x_1); -x_159 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1___boxed), 12, 2); -lean_closure_set(x_159, 0, x_1); -lean_closure_set(x_159, 1, x_2); -x_175 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__4; -x_176 = l_Lean_Expr_isConstOf(x_3, x_175); -if (x_176 == 0) +x_15 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1___boxed), 12, 2); +lean_closure_set(x_15, 0, x_1); +lean_closure_set(x_15, 1, x_2); +x_28 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__4; +x_29 = l_Lean_Expr_isConstOf(x_3, x_28); +if (x_29 == 0) { -lean_object* x_177; -x_177 = lean_box(0); -x_160 = x_177; -goto block_174; +lean_object* x_30; +x_30 = lean_box(0); +x_16 = x_30; +goto block_27; } else { -lean_object* x_178; lean_object* x_179; uint8_t x_180; -x_178 = lean_array_get_size(x_4); -x_179 = lean_unsigned_to_nat(2u); -x_180 = lean_nat_dec_eq(x_178, x_179); -if (x_180 == 0) +lean_object* x_31; lean_object* x_32; uint8_t x_33; +x_31 = lean_array_get_size(x_4); +x_32 = lean_unsigned_to_nat(2u); +x_33 = lean_nat_dec_eq(x_31, x_32); +if (x_33 == 0) { -lean_object* x_181; -lean_dec(x_178); -x_181 = lean_box(0); -x_160 = x_181; -goto block_174; +lean_object* x_34; +lean_dec(x_31); +x_34 = lean_box(0); +x_16 = x_34; +goto block_27; } else { -lean_object* x_182; uint8_t x_183; -lean_dec(x_159); +lean_object* x_35; uint8_t x_36; +lean_dec(x_15); lean_dec(x_3); -x_182 = lean_unsigned_to_nat(0u); -x_183 = lean_nat_dec_lt(x_182, x_178); -lean_dec(x_178); -if (x_183 == 0) +x_35 = lean_unsigned_to_nat(0u); +x_36 = lean_nat_dec_lt(x_35, x_31); +lean_dec(x_31); +if (x_36 == 0) { -lean_object* x_184; lean_object* x_185; lean_object* x_186; +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_dec(x_4); -x_184 = l_Lean_instInhabitedExpr; -x_185 = l_outOfBounds___rarg(x_184); +x_37 = l_Lean_instInhabitedExpr; +x_38 = l_outOfBounds___rarg(x_37); lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); @@ -5011,30 +4889,30 @@ lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_2); -lean_inc(x_185); -x_186 = l_Lean_Meta_Grind_internalize(x_185, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_186) == 0) +lean_inc(x_38); +x_39 = l_Lean_Meta_Grind_internalize(x_38, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_39) == 0) { -lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; -x_187 = lean_ctor_get(x_186, 1); -lean_inc(x_187); -lean_dec(x_186); +lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_40 = lean_ctor_get(x_39, 1); +lean_inc(x_40); +lean_dec(x_39); lean_inc(x_1); -x_188 = l_Lean_Meta_Grind_registerParent(x_1, x_185, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_187); -lean_dec(x_185); -x_189 = lean_ctor_get(x_188, 0); -lean_inc(x_189); -x_190 = lean_ctor_get(x_188, 1); -lean_inc(x_190); -lean_dec(x_188); -x_191 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_189, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_190); -lean_dec(x_189); -return x_191; +x_41 = l_Lean_Meta_Grind_registerParent(x_1, x_38, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_40); +lean_dec(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_43); +lean_dec(x_41); +x_44 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_42, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_43); +lean_dec(x_42); +return x_44; } else { -uint8_t x_192; -lean_dec(x_185); +uint8_t x_45; +lean_dec(x_38); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); @@ -5045,30 +4923,30 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_2); lean_dec(x_1); -x_192 = !lean_is_exclusive(x_186); -if (x_192 == 0) +x_45 = !lean_is_exclusive(x_39); +if (x_45 == 0) { -return x_186; +return x_39; } else { -lean_object* x_193; lean_object* x_194; lean_object* x_195; -x_193 = lean_ctor_get(x_186, 0); -x_194 = lean_ctor_get(x_186, 1); -lean_inc(x_194); -lean_inc(x_193); -lean_dec(x_186); -x_195 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_195, 0, x_193); -lean_ctor_set(x_195, 1, x_194); -return x_195; +lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_46 = lean_ctor_get(x_39, 0); +x_47 = lean_ctor_get(x_39, 1); +lean_inc(x_47); +lean_inc(x_46); +lean_dec(x_39); +x_48 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_48, 0, x_46); +lean_ctor_set(x_48, 1, x_47); +return x_48; } } } else { -lean_object* x_196; lean_object* x_197; -x_196 = lean_array_fget(x_4, x_182); +lean_object* x_49; lean_object* x_50; +x_49 = lean_array_fget(x_4, x_35); lean_dec(x_4); lean_inc(x_13); lean_inc(x_12); @@ -5079,30 +4957,30 @@ lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_2); -lean_inc(x_196); -x_197 = l_Lean_Meta_Grind_internalize(x_196, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_197) == 0) +lean_inc(x_49); +x_50 = l_Lean_Meta_Grind_internalize(x_49, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_50) == 0) { -lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; -x_198 = lean_ctor_get(x_197, 1); -lean_inc(x_198); -lean_dec(x_197); +lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_51 = lean_ctor_get(x_50, 1); +lean_inc(x_51); +lean_dec(x_50); lean_inc(x_1); -x_199 = l_Lean_Meta_Grind_registerParent(x_1, x_196, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_198); -lean_dec(x_196); -x_200 = lean_ctor_get(x_199, 0); -lean_inc(x_200); -x_201 = lean_ctor_get(x_199, 1); -lean_inc(x_201); -lean_dec(x_199); -x_202 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_200, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_201); -lean_dec(x_200); -return x_202; +x_52 = l_Lean_Meta_Grind_registerParent(x_1, x_49, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_51); +lean_dec(x_49); +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_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_53, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_54); +lean_dec(x_53); +return x_55; } else { -uint8_t x_203; -lean_dec(x_196); +uint8_t x_56; +lean_dec(x_49); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); @@ -5113,36 +4991,32 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_2); lean_dec(x_1); -x_203 = !lean_is_exclusive(x_197); -if (x_203 == 0) +x_56 = !lean_is_exclusive(x_50); +if (x_56 == 0) { -return x_197; +return x_50; } else { -lean_object* x_204; lean_object* x_205; lean_object* x_206; -x_204 = lean_ctor_get(x_197, 0); -x_205 = lean_ctor_get(x_197, 1); -lean_inc(x_205); -lean_inc(x_204); -lean_dec(x_197); -x_206 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_206, 0, x_204); -lean_ctor_set(x_206, 1, x_205); -return x_206; +lean_object* x_57; lean_object* x_58; lean_object* x_59; +x_57 = lean_ctor_get(x_50, 0); +x_58 = lean_ctor_get(x_50, 1); +lean_inc(x_58); +lean_inc(x_57); +lean_dec(x_50); +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; } } } } } -block_174: -{ -uint8_t x_161; -lean_dec(x_160); -x_161 = l_Lean_Expr_isConst(x_3); -if (x_161 == 0) +block_27: { -lean_object* x_162; +lean_object* x_17; +lean_dec(x_16); lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); @@ -5153,30 +5027,30 @@ lean_inc(x_7); lean_inc(x_6); lean_inc(x_2); lean_inc(x_3); -x_162 = l_Lean_Meta_Grind_internalize(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_162) == 0) +x_17 = l_Lean_Meta_Grind_internalize(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_17) == 0) { -lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; -x_163 = lean_ctor_get(x_162, 1); -lean_inc(x_163); -lean_dec(x_162); +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, 1); +lean_inc(x_18); +lean_dec(x_17); lean_inc(x_1); -x_164 = l_Lean_Meta_Grind_registerParent(x_1, x_3, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_163); +x_19 = l_Lean_Meta_Grind_registerParent(x_1, x_3, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_18); lean_dec(x_3); -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_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(x_4, x_1, x_2, x_159, x_165, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_166); -lean_dec(x_165); +x_20 = lean_ctor_get(x_19, 0); +lean_inc(x_20); +x_21 = lean_ctor_get(x_19, 1); +lean_inc(x_21); +lean_dec(x_19); +x_22 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(x_4, x_1, x_2, x_15, x_20, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_21); +lean_dec(x_20); lean_dec(x_4); -return x_167; +return x_22; } else { -uint8_t x_168; -lean_dec(x_159); +uint8_t x_23; +lean_dec(x_15); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); @@ -5189,83 +5063,73 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_168 = !lean_is_exclusive(x_162); -if (x_168 == 0) +x_23 = !lean_is_exclusive(x_17); +if (x_23 == 0) { -return x_162; +return x_17; } 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_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); -return x_171; -} -} +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_17, 0); +x_25 = lean_ctor_get(x_17, 1); +lean_inc(x_25); +lean_inc(x_24); +lean_dec(x_17); +x_26 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_26, 0, x_24); +lean_ctor_set(x_26, 1, x_25); +return x_26; } -else -{ -lean_object* x_172; lean_object* x_173; -lean_dec(x_3); -x_172 = lean_box(0); -x_173 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(x_4, x_1, x_2, x_159, x_172, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -lean_dec(x_4); -return x_173; } } } -case 4: +case 1: { -lean_object* x_207; lean_object* x_208; lean_object* x_223; uint8_t x_224; +lean_object* x_60; lean_object* x_61; lean_object* x_73; uint8_t x_74; lean_dec(x_5); lean_inc(x_2); lean_inc(x_1); -x_207 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1___boxed), 12, 2); -lean_closure_set(x_207, 0, x_1); -lean_closure_set(x_207, 1, x_2); -x_223 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__4; -x_224 = l_Lean_Expr_isConstOf(x_3, x_223); -if (x_224 == 0) +x_60 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1___boxed), 12, 2); +lean_closure_set(x_60, 0, x_1); +lean_closure_set(x_60, 1, x_2); +x_73 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__4; +x_74 = l_Lean_Expr_isConstOf(x_3, x_73); +if (x_74 == 0) { -lean_object* x_225; -x_225 = lean_box(0); -x_208 = x_225; -goto block_222; +lean_object* x_75; +x_75 = lean_box(0); +x_61 = x_75; +goto block_72; } else { -lean_object* x_226; lean_object* x_227; uint8_t x_228; -x_226 = lean_array_get_size(x_4); -x_227 = lean_unsigned_to_nat(2u); -x_228 = lean_nat_dec_eq(x_226, x_227); -if (x_228 == 0) +lean_object* x_76; lean_object* x_77; uint8_t x_78; +x_76 = lean_array_get_size(x_4); +x_77 = lean_unsigned_to_nat(2u); +x_78 = lean_nat_dec_eq(x_76, x_77); +if (x_78 == 0) { -lean_object* x_229; -lean_dec(x_226); -x_229 = lean_box(0); -x_208 = x_229; -goto block_222; +lean_object* x_79; +lean_dec(x_76); +x_79 = lean_box(0); +x_61 = x_79; +goto block_72; } else { -lean_object* x_230; uint8_t x_231; -lean_dec(x_207); +lean_object* x_80; uint8_t x_81; +lean_dec(x_60); lean_dec(x_3); -x_230 = lean_unsigned_to_nat(0u); -x_231 = lean_nat_dec_lt(x_230, x_226); -lean_dec(x_226); -if (x_231 == 0) +x_80 = lean_unsigned_to_nat(0u); +x_81 = lean_nat_dec_lt(x_80, x_76); +lean_dec(x_76); +if (x_81 == 0) { -lean_object* x_232; lean_object* x_233; lean_object* x_234; +lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_dec(x_4); -x_232 = l_Lean_instInhabitedExpr; -x_233 = l_outOfBounds___rarg(x_232); +x_82 = l_Lean_instInhabitedExpr; +x_83 = l_outOfBounds___rarg(x_82); lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); @@ -5275,30 +5139,30 @@ lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_2); -lean_inc(x_233); -x_234 = l_Lean_Meta_Grind_internalize(x_233, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_234) == 0) +lean_inc(x_83); +x_84 = l_Lean_Meta_Grind_internalize(x_83, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_84) == 0) { -lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; -x_235 = lean_ctor_get(x_234, 1); -lean_inc(x_235); -lean_dec(x_234); +lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; +x_85 = lean_ctor_get(x_84, 1); +lean_inc(x_85); +lean_dec(x_84); lean_inc(x_1); -x_236 = l_Lean_Meta_Grind_registerParent(x_1, x_233, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_235); -lean_dec(x_233); -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 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_237, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_238); -lean_dec(x_237); -return x_239; +x_86 = l_Lean_Meta_Grind_registerParent(x_1, x_83, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_85); +lean_dec(x_83); +x_87 = lean_ctor_get(x_86, 0); +lean_inc(x_87); +x_88 = lean_ctor_get(x_86, 1); +lean_inc(x_88); +lean_dec(x_86); +x_89 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_87, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_88); +lean_dec(x_87); +return x_89; } else { -uint8_t x_240; -lean_dec(x_233); +uint8_t x_90; +lean_dec(x_83); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); @@ -5309,30 +5173,30 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_2); lean_dec(x_1); -x_240 = !lean_is_exclusive(x_234); -if (x_240 == 0) +x_90 = !lean_is_exclusive(x_84); +if (x_90 == 0) { -return x_234; +return x_84; } else { -lean_object* x_241; lean_object* x_242; lean_object* x_243; -x_241 = lean_ctor_get(x_234, 0); -x_242 = lean_ctor_get(x_234, 1); -lean_inc(x_242); -lean_inc(x_241); -lean_dec(x_234); -x_243 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_243, 0, x_241); -lean_ctor_set(x_243, 1, x_242); -return x_243; +lean_object* x_91; lean_object* x_92; lean_object* x_93; +x_91 = lean_ctor_get(x_84, 0); +x_92 = lean_ctor_get(x_84, 1); +lean_inc(x_92); +lean_inc(x_91); +lean_dec(x_84); +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_244; lean_object* x_245; -x_244 = lean_array_fget(x_4, x_230); +lean_object* x_94; lean_object* x_95; +x_94 = lean_array_fget(x_4, x_80); lean_dec(x_4); lean_inc(x_13); lean_inc(x_12); @@ -5343,30 +5207,30 @@ lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_2); -lean_inc(x_244); -x_245 = l_Lean_Meta_Grind_internalize(x_244, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_245) == 0) +lean_inc(x_94); +x_95 = l_Lean_Meta_Grind_internalize(x_94, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_95) == 0) { -lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; -x_246 = lean_ctor_get(x_245, 1); -lean_inc(x_246); -lean_dec(x_245); +lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; +x_96 = lean_ctor_get(x_95, 1); +lean_inc(x_96); +lean_dec(x_95); lean_inc(x_1); -x_247 = l_Lean_Meta_Grind_registerParent(x_1, x_244, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_246); -lean_dec(x_244); -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_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_248, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_249); -lean_dec(x_248); -return x_250; +x_97 = l_Lean_Meta_Grind_registerParent(x_1, x_94, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_96); +lean_dec(x_94); +x_98 = lean_ctor_get(x_97, 0); +lean_inc(x_98); +x_99 = lean_ctor_get(x_97, 1); +lean_inc(x_99); +lean_dec(x_97); +x_100 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_98, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_99); +lean_dec(x_98); +return x_100; } else { -uint8_t x_251; -lean_dec(x_244); +uint8_t x_101; +lean_dec(x_94); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); @@ -5377,36 +5241,32 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_2); lean_dec(x_1); -x_251 = !lean_is_exclusive(x_245); -if (x_251 == 0) +x_101 = !lean_is_exclusive(x_95); +if (x_101 == 0) { -return x_245; +return x_95; } else { -lean_object* x_252; lean_object* x_253; lean_object* x_254; -x_252 = lean_ctor_get(x_245, 0); -x_253 = lean_ctor_get(x_245, 1); -lean_inc(x_253); -lean_inc(x_252); -lean_dec(x_245); -x_254 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_254, 0, x_252); -lean_ctor_set(x_254, 1, x_253); -return x_254; +lean_object* x_102; lean_object* x_103; lean_object* x_104; +x_102 = lean_ctor_get(x_95, 0); +x_103 = lean_ctor_get(x_95, 1); +lean_inc(x_103); +lean_inc(x_102); +lean_dec(x_95); +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; } } } } } -block_222: -{ -uint8_t x_209; -lean_dec(x_208); -x_209 = l_Lean_Expr_isConst(x_3); -if (x_209 == 0) +block_72: { -lean_object* x_210; +lean_object* x_62; +lean_dec(x_61); lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); @@ -5417,30 +5277,30 @@ lean_inc(x_7); lean_inc(x_6); lean_inc(x_2); lean_inc(x_3); -x_210 = l_Lean_Meta_Grind_internalize(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_210) == 0) +x_62 = l_Lean_Meta_Grind_internalize(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_62) == 0) { -lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; -x_211 = lean_ctor_get(x_210, 1); -lean_inc(x_211); -lean_dec(x_210); +lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; +x_63 = lean_ctor_get(x_62, 1); +lean_inc(x_63); +lean_dec(x_62); lean_inc(x_1); -x_212 = l_Lean_Meta_Grind_registerParent(x_1, x_3, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_211); +x_64 = l_Lean_Meta_Grind_registerParent(x_1, x_3, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_63); lean_dec(x_3); -x_213 = lean_ctor_get(x_212, 0); -lean_inc(x_213); -x_214 = lean_ctor_get(x_212, 1); -lean_inc(x_214); -lean_dec(x_212); -x_215 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(x_4, x_1, x_2, x_207, x_213, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_214); -lean_dec(x_213); +x_65 = lean_ctor_get(x_64, 0); +lean_inc(x_65); +x_66 = lean_ctor_get(x_64, 1); +lean_inc(x_66); +lean_dec(x_64); +x_67 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(x_4, x_1, x_2, x_60, x_65, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_66); +lean_dec(x_65); lean_dec(x_4); -return x_215; +return x_67; } else { -uint8_t x_216; -lean_dec(x_207); +uint8_t x_68; +lean_dec(x_60); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); @@ -5453,100 +5313,73 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_216 = !lean_is_exclusive(x_210); -if (x_216 == 0) +x_68 = !lean_is_exclusive(x_62); +if (x_68 == 0) { -return x_210; +return x_62; } else { -lean_object* x_217; lean_object* x_218; lean_object* x_219; -x_217 = lean_ctor_get(x_210, 0); -x_218 = lean_ctor_get(x_210, 1); -lean_inc(x_218); -lean_inc(x_217); -lean_dec(x_210); -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; +lean_object* x_69; lean_object* x_70; lean_object* x_71; +x_69 = lean_ctor_get(x_62, 0); +x_70 = lean_ctor_get(x_62, 1); +lean_inc(x_70); +lean_inc(x_69); +lean_dec(x_62); +x_71 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_71, 0, x_69); +lean_ctor_set(x_71, 1, x_70); +return x_71; } } } -else -{ -lean_object* x_220; lean_object* x_221; -lean_dec(x_3); -x_220 = lean_box(0); -x_221 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(x_4, x_1, x_2, x_207, x_220, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -lean_dec(x_4); -return x_221; -} -} -} -case 5: -{ -lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; -x_255 = lean_ctor_get(x_3, 0); -lean_inc(x_255); -x_256 = lean_ctor_get(x_3, 1); -lean_inc(x_256); -lean_dec(x_3); -x_257 = lean_array_set(x_4, x_5, x_256); -x_258 = lean_unsigned_to_nat(1u); -x_259 = lean_nat_sub(x_5, x_258); -lean_dec(x_5); -x_3 = x_255; -x_4 = x_257; -x_5 = x_259; -goto _start; } -case 6: +case 2: { -lean_object* x_261; lean_object* x_262; lean_object* x_277; uint8_t x_278; +lean_object* x_105; lean_object* x_106; lean_object* x_118; uint8_t x_119; lean_dec(x_5); lean_inc(x_2); lean_inc(x_1); -x_261 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1___boxed), 12, 2); -lean_closure_set(x_261, 0, x_1); -lean_closure_set(x_261, 1, x_2); -x_277 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__4; -x_278 = l_Lean_Expr_isConstOf(x_3, x_277); -if (x_278 == 0) +x_105 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1___boxed), 12, 2); +lean_closure_set(x_105, 0, x_1); +lean_closure_set(x_105, 1, x_2); +x_118 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__4; +x_119 = l_Lean_Expr_isConstOf(x_3, x_118); +if (x_119 == 0) { -lean_object* x_279; -x_279 = lean_box(0); -x_262 = x_279; -goto block_276; +lean_object* x_120; +x_120 = lean_box(0); +x_106 = x_120; +goto block_117; } else { -lean_object* x_280; lean_object* x_281; uint8_t x_282; -x_280 = lean_array_get_size(x_4); -x_281 = lean_unsigned_to_nat(2u); -x_282 = lean_nat_dec_eq(x_280, x_281); -if (x_282 == 0) +lean_object* x_121; lean_object* x_122; uint8_t x_123; +x_121 = lean_array_get_size(x_4); +x_122 = lean_unsigned_to_nat(2u); +x_123 = lean_nat_dec_eq(x_121, x_122); +if (x_123 == 0) { -lean_object* x_283; -lean_dec(x_280); -x_283 = lean_box(0); -x_262 = x_283; -goto block_276; +lean_object* x_124; +lean_dec(x_121); +x_124 = lean_box(0); +x_106 = x_124; +goto block_117; } else { -lean_object* x_284; uint8_t x_285; -lean_dec(x_261); +lean_object* x_125; uint8_t x_126; +lean_dec(x_105); lean_dec(x_3); -x_284 = lean_unsigned_to_nat(0u); -x_285 = lean_nat_dec_lt(x_284, x_280); -lean_dec(x_280); -if (x_285 == 0) +x_125 = lean_unsigned_to_nat(0u); +x_126 = lean_nat_dec_lt(x_125, x_121); +lean_dec(x_121); +if (x_126 == 0) { -lean_object* x_286; lean_object* x_287; lean_object* x_288; +lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_dec(x_4); -x_286 = l_Lean_instInhabitedExpr; -x_287 = l_outOfBounds___rarg(x_286); +x_127 = l_Lean_instInhabitedExpr; +x_128 = l_outOfBounds___rarg(x_127); lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); @@ -5556,30 +5389,30 @@ lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_2); -lean_inc(x_287); -x_288 = l_Lean_Meta_Grind_internalize(x_287, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_288) == 0) -{ -lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; -x_289 = lean_ctor_get(x_288, 1); -lean_inc(x_289); -lean_dec(x_288); +lean_inc(x_128); +x_129 = l_Lean_Meta_Grind_internalize(x_128, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +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; +x_130 = lean_ctor_get(x_129, 1); +lean_inc(x_130); +lean_dec(x_129); lean_inc(x_1); -x_290 = l_Lean_Meta_Grind_registerParent(x_1, x_287, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_289); -lean_dec(x_287); -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 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_291, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_292); -lean_dec(x_291); -return x_293; +x_131 = l_Lean_Meta_Grind_registerParent(x_1, x_128, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_130); +lean_dec(x_128); +x_132 = lean_ctor_get(x_131, 0); +lean_inc(x_132); +x_133 = lean_ctor_get(x_131, 1); +lean_inc(x_133); +lean_dec(x_131); +x_134 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_132, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_133); +lean_dec(x_132); +return x_134; } else { -uint8_t x_294; -lean_dec(x_287); +uint8_t x_135; +lean_dec(x_128); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); @@ -5590,30 +5423,30 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_2); lean_dec(x_1); -x_294 = !lean_is_exclusive(x_288); -if (x_294 == 0) +x_135 = !lean_is_exclusive(x_129); +if (x_135 == 0) { -return x_288; +return x_129; } else { -lean_object* x_295; lean_object* x_296; lean_object* x_297; -x_295 = lean_ctor_get(x_288, 0); -x_296 = lean_ctor_get(x_288, 1); -lean_inc(x_296); -lean_inc(x_295); -lean_dec(x_288); -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; +lean_object* x_136; lean_object* x_137; lean_object* x_138; +x_136 = lean_ctor_get(x_129, 0); +x_137 = lean_ctor_get(x_129, 1); +lean_inc(x_137); +lean_inc(x_136); +lean_dec(x_129); +x_138 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_138, 0, x_136); +lean_ctor_set(x_138, 1, x_137); +return x_138; } } } else { -lean_object* x_298; lean_object* x_299; -x_298 = lean_array_fget(x_4, x_284); +lean_object* x_139; lean_object* x_140; +x_139 = lean_array_fget(x_4, x_125); lean_dec(x_4); lean_inc(x_13); lean_inc(x_12); @@ -5624,30 +5457,30 @@ lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_2); -lean_inc(x_298); -x_299 = l_Lean_Meta_Grind_internalize(x_298, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_299) == 0) +lean_inc(x_139); +x_140 = l_Lean_Meta_Grind_internalize(x_139, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_140) == 0) { -lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; -x_300 = lean_ctor_get(x_299, 1); -lean_inc(x_300); -lean_dec(x_299); +lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; +x_141 = lean_ctor_get(x_140, 1); +lean_inc(x_141); +lean_dec(x_140); lean_inc(x_1); -x_301 = l_Lean_Meta_Grind_registerParent(x_1, x_298, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_300); -lean_dec(x_298); -x_302 = lean_ctor_get(x_301, 0); -lean_inc(x_302); -x_303 = lean_ctor_get(x_301, 1); -lean_inc(x_303); -lean_dec(x_301); -x_304 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_302, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_303); -lean_dec(x_302); -return x_304; +x_142 = l_Lean_Meta_Grind_registerParent(x_1, x_139, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_141); +lean_dec(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 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_143, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_144); +lean_dec(x_143); +return x_145; } else { -uint8_t x_305; -lean_dec(x_298); +uint8_t x_146; +lean_dec(x_139); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); @@ -5658,36 +5491,32 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_2); lean_dec(x_1); -x_305 = !lean_is_exclusive(x_299); -if (x_305 == 0) +x_146 = !lean_is_exclusive(x_140); +if (x_146 == 0) { -return x_299; +return x_140; } else { -lean_object* x_306; lean_object* x_307; lean_object* x_308; -x_306 = lean_ctor_get(x_299, 0); -x_307 = lean_ctor_get(x_299, 1); -lean_inc(x_307); -lean_inc(x_306); -lean_dec(x_299); -x_308 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_308, 0, x_306); -lean_ctor_set(x_308, 1, x_307); -return x_308; +lean_object* x_147; lean_object* x_148; lean_object* x_149; +x_147 = lean_ctor_get(x_140, 0); +x_148 = lean_ctor_get(x_140, 1); +lean_inc(x_148); +lean_inc(x_147); +lean_dec(x_140); +x_149 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_149, 0, x_147); +lean_ctor_set(x_149, 1, x_148); +return x_149; } } } } } -block_276: -{ -uint8_t x_263; -lean_dec(x_262); -x_263 = l_Lean_Expr_isConst(x_3); -if (x_263 == 0) +block_117: { -lean_object* x_264; +lean_object* x_107; +lean_dec(x_106); lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); @@ -5698,30 +5527,30 @@ lean_inc(x_7); lean_inc(x_6); lean_inc(x_2); lean_inc(x_3); -x_264 = l_Lean_Meta_Grind_internalize(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_264) == 0) +x_107 = l_Lean_Meta_Grind_internalize(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_107) == 0) { -lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; -x_265 = lean_ctor_get(x_264, 1); -lean_inc(x_265); -lean_dec(x_264); +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, 1); +lean_inc(x_108); +lean_dec(x_107); lean_inc(x_1); -x_266 = l_Lean_Meta_Grind_registerParent(x_1, x_3, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_265); +x_109 = l_Lean_Meta_Grind_registerParent(x_1, x_3, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_108); lean_dec(x_3); -x_267 = lean_ctor_get(x_266, 0); -lean_inc(x_267); -x_268 = lean_ctor_get(x_266, 1); -lean_inc(x_268); -lean_dec(x_266); -x_269 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(x_4, x_1, x_2, x_261, x_267, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_268); -lean_dec(x_267); +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_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(x_4, x_1, x_2, x_105, x_110, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_111); +lean_dec(x_110); lean_dec(x_4); -return x_269; +return x_112; } else { -uint8_t x_270; -lean_dec(x_261); +uint8_t x_113; +lean_dec(x_105); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); @@ -5734,83 +5563,73 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_270 = !lean_is_exclusive(x_264); -if (x_270 == 0) +x_113 = !lean_is_exclusive(x_107); +if (x_113 == 0) { -return x_264; +return x_107; } else { -lean_object* x_271; lean_object* x_272; lean_object* x_273; -x_271 = lean_ctor_get(x_264, 0); -x_272 = lean_ctor_get(x_264, 1); -lean_inc(x_272); -lean_inc(x_271); -lean_dec(x_264); -x_273 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_273, 0, x_271); -lean_ctor_set(x_273, 1, x_272); -return x_273; -} -} +lean_object* x_114; lean_object* x_115; lean_object* x_116; +x_114 = lean_ctor_get(x_107, 0); +x_115 = lean_ctor_get(x_107, 1); +lean_inc(x_115); +lean_inc(x_114); +lean_dec(x_107); +x_116 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_116, 0, x_114); +lean_ctor_set(x_116, 1, x_115); +return x_116; } -else -{ -lean_object* x_274; lean_object* x_275; -lean_dec(x_3); -x_274 = lean_box(0); -x_275 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(x_4, x_1, x_2, x_261, x_274, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -lean_dec(x_4); -return x_275; } } } -case 7: +case 3: { -lean_object* x_309; lean_object* x_310; lean_object* x_325; uint8_t x_326; +lean_object* x_150; lean_object* x_151; lean_object* x_163; uint8_t x_164; lean_dec(x_5); lean_inc(x_2); lean_inc(x_1); -x_309 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1___boxed), 12, 2); -lean_closure_set(x_309, 0, x_1); -lean_closure_set(x_309, 1, x_2); -x_325 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__4; -x_326 = l_Lean_Expr_isConstOf(x_3, x_325); -if (x_326 == 0) +x_150 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1___boxed), 12, 2); +lean_closure_set(x_150, 0, x_1); +lean_closure_set(x_150, 1, x_2); +x_163 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__4; +x_164 = l_Lean_Expr_isConstOf(x_3, x_163); +if (x_164 == 0) { -lean_object* x_327; -x_327 = lean_box(0); -x_310 = x_327; -goto block_324; +lean_object* x_165; +x_165 = lean_box(0); +x_151 = x_165; +goto block_162; } else { -lean_object* x_328; lean_object* x_329; uint8_t x_330; -x_328 = lean_array_get_size(x_4); -x_329 = lean_unsigned_to_nat(2u); -x_330 = lean_nat_dec_eq(x_328, x_329); -if (x_330 == 0) +lean_object* x_166; lean_object* x_167; uint8_t x_168; +x_166 = lean_array_get_size(x_4); +x_167 = lean_unsigned_to_nat(2u); +x_168 = lean_nat_dec_eq(x_166, x_167); +if (x_168 == 0) { -lean_object* x_331; -lean_dec(x_328); -x_331 = lean_box(0); -x_310 = x_331; -goto block_324; +lean_object* x_169; +lean_dec(x_166); +x_169 = lean_box(0); +x_151 = x_169; +goto block_162; } else { -lean_object* x_332; uint8_t x_333; -lean_dec(x_309); +lean_object* x_170; uint8_t x_171; +lean_dec(x_150); lean_dec(x_3); -x_332 = lean_unsigned_to_nat(0u); -x_333 = lean_nat_dec_lt(x_332, x_328); -lean_dec(x_328); -if (x_333 == 0) +x_170 = lean_unsigned_to_nat(0u); +x_171 = lean_nat_dec_lt(x_170, x_166); +lean_dec(x_166); +if (x_171 == 0) { -lean_object* x_334; lean_object* x_335; lean_object* x_336; +lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_dec(x_4); -x_334 = l_Lean_instInhabitedExpr; -x_335 = l_outOfBounds___rarg(x_334); +x_172 = l_Lean_instInhabitedExpr; +x_173 = l_outOfBounds___rarg(x_172); lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); @@ -5820,30 +5639,30 @@ lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_2); -lean_inc(x_335); -x_336 = l_Lean_Meta_Grind_internalize(x_335, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_336) == 0) +lean_inc(x_173); +x_174 = l_Lean_Meta_Grind_internalize(x_173, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_174) == 0) { -lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; -x_337 = lean_ctor_get(x_336, 1); -lean_inc(x_337); -lean_dec(x_336); +lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; +x_175 = lean_ctor_get(x_174, 1); +lean_inc(x_175); +lean_dec(x_174); lean_inc(x_1); -x_338 = l_Lean_Meta_Grind_registerParent(x_1, x_335, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_337); -lean_dec(x_335); -x_339 = lean_ctor_get(x_338, 0); -lean_inc(x_339); -x_340 = lean_ctor_get(x_338, 1); -lean_inc(x_340); -lean_dec(x_338); -x_341 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_339, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_340); -lean_dec(x_339); -return x_341; +x_176 = l_Lean_Meta_Grind_registerParent(x_1, x_173, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_175); +lean_dec(x_173); +x_177 = lean_ctor_get(x_176, 0); +lean_inc(x_177); +x_178 = lean_ctor_get(x_176, 1); +lean_inc(x_178); +lean_dec(x_176); +x_179 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_177, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_178); +lean_dec(x_177); +return x_179; } else { -uint8_t x_342; -lean_dec(x_335); +uint8_t x_180; +lean_dec(x_173); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); @@ -5854,30 +5673,30 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_2); lean_dec(x_1); -x_342 = !lean_is_exclusive(x_336); -if (x_342 == 0) +x_180 = !lean_is_exclusive(x_174); +if (x_180 == 0) { -return x_336; +return x_174; } else { -lean_object* x_343; lean_object* x_344; lean_object* x_345; -x_343 = lean_ctor_get(x_336, 0); -x_344 = lean_ctor_get(x_336, 1); -lean_inc(x_344); -lean_inc(x_343); -lean_dec(x_336); -x_345 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_345, 0, x_343); -lean_ctor_set(x_345, 1, x_344); -return x_345; +lean_object* x_181; lean_object* x_182; lean_object* x_183; +x_181 = lean_ctor_get(x_174, 0); +x_182 = lean_ctor_get(x_174, 1); +lean_inc(x_182); +lean_inc(x_181); +lean_dec(x_174); +x_183 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_183, 0, x_181); +lean_ctor_set(x_183, 1, x_182); +return x_183; } } } else { -lean_object* x_346; lean_object* x_347; -x_346 = lean_array_fget(x_4, x_332); +lean_object* x_184; lean_object* x_185; +x_184 = lean_array_fget(x_4, x_170); lean_dec(x_4); lean_inc(x_13); lean_inc(x_12); @@ -5888,30 +5707,30 @@ lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_2); -lean_inc(x_346); -x_347 = l_Lean_Meta_Grind_internalize(x_346, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_347) == 0) -{ -lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; -x_348 = lean_ctor_get(x_347, 1); -lean_inc(x_348); -lean_dec(x_347); +lean_inc(x_184); +x_185 = l_Lean_Meta_Grind_internalize(x_184, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +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; +x_186 = lean_ctor_get(x_185, 1); +lean_inc(x_186); +lean_dec(x_185); lean_inc(x_1); -x_349 = l_Lean_Meta_Grind_registerParent(x_1, x_346, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_348); -lean_dec(x_346); -x_350 = lean_ctor_get(x_349, 0); -lean_inc(x_350); -x_351 = lean_ctor_get(x_349, 1); -lean_inc(x_351); -lean_dec(x_349); -x_352 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_350, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_351); -lean_dec(x_350); -return x_352; +x_187 = l_Lean_Meta_Grind_registerParent(x_1, x_184, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_186); +lean_dec(x_184); +x_188 = lean_ctor_get(x_187, 0); +lean_inc(x_188); +x_189 = lean_ctor_get(x_187, 1); +lean_inc(x_189); +lean_dec(x_187); +x_190 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_188, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_189); +lean_dec(x_188); +return x_190; } else { -uint8_t x_353; -lean_dec(x_346); +uint8_t x_191; +lean_dec(x_184); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); @@ -5922,36 +5741,32 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_2); lean_dec(x_1); -x_353 = !lean_is_exclusive(x_347); -if (x_353 == 0) +x_191 = !lean_is_exclusive(x_185); +if (x_191 == 0) { -return x_347; +return x_185; } else { -lean_object* x_354; lean_object* x_355; lean_object* x_356; -x_354 = lean_ctor_get(x_347, 0); -x_355 = lean_ctor_get(x_347, 1); -lean_inc(x_355); -lean_inc(x_354); -lean_dec(x_347); -x_356 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_356, 0, x_354); -lean_ctor_set(x_356, 1, x_355); -return x_356; +lean_object* x_192; lean_object* x_193; lean_object* x_194; +x_192 = lean_ctor_get(x_185, 0); +x_193 = lean_ctor_get(x_185, 1); +lean_inc(x_193); +lean_inc(x_192); +lean_dec(x_185); +x_194 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_194, 0, x_192); +lean_ctor_set(x_194, 1, x_193); +return x_194; } } } } } -block_324: -{ -uint8_t x_311; -lean_dec(x_310); -x_311 = l_Lean_Expr_isConst(x_3); -if (x_311 == 0) +block_162: { -lean_object* x_312; +lean_object* x_152; +lean_dec(x_151); lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); @@ -5962,30 +5777,30 @@ lean_inc(x_7); lean_inc(x_6); lean_inc(x_2); lean_inc(x_3); -x_312 = l_Lean_Meta_Grind_internalize(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_312) == 0) +x_152 = l_Lean_Meta_Grind_internalize(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_152) == 0) { -lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; -x_313 = lean_ctor_get(x_312, 1); -lean_inc(x_313); -lean_dec(x_312); +lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; +x_153 = lean_ctor_get(x_152, 1); +lean_inc(x_153); +lean_dec(x_152); lean_inc(x_1); -x_314 = l_Lean_Meta_Grind_registerParent(x_1, x_3, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_313); +x_154 = l_Lean_Meta_Grind_registerParent(x_1, x_3, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_153); lean_dec(x_3); -x_315 = lean_ctor_get(x_314, 0); -lean_inc(x_315); -x_316 = lean_ctor_get(x_314, 1); -lean_inc(x_316); -lean_dec(x_314); -x_317 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(x_4, x_1, x_2, x_309, x_315, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_316); -lean_dec(x_315); +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 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(x_4, x_1, x_2, x_150, x_155, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_156); +lean_dec(x_155); lean_dec(x_4); -return x_317; +return x_157; } else { -uint8_t x_318; -lean_dec(x_309); +uint8_t x_158; +lean_dec(x_150); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); @@ -5998,83 +5813,76 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_318 = !lean_is_exclusive(x_312); -if (x_318 == 0) +x_158 = !lean_is_exclusive(x_152); +if (x_158 == 0) { -return x_312; +return x_152; } else { -lean_object* x_319; lean_object* x_320; lean_object* x_321; -x_319 = lean_ctor_get(x_312, 0); -x_320 = lean_ctor_get(x_312, 1); -lean_inc(x_320); -lean_inc(x_319); -lean_dec(x_312); -x_321 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_321, 0, x_319); -lean_ctor_set(x_321, 1, x_320); -return x_321; -} -} +lean_object* x_159; lean_object* x_160; lean_object* x_161; +x_159 = lean_ctor_get(x_152, 0); +x_160 = lean_ctor_get(x_152, 1); +lean_inc(x_160); +lean_inc(x_159); +lean_dec(x_152); +x_161 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_161, 0, x_159); +lean_ctor_set(x_161, 1, x_160); +return x_161; } -else -{ -lean_object* x_322; lean_object* x_323; -lean_dec(x_3); -x_322 = lean_box(0); -x_323 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(x_4, x_1, x_2, x_309, x_322, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -lean_dec(x_4); -return x_323; } } } -case 8: +case 4: { -lean_object* x_357; lean_object* x_358; lean_object* x_373; uint8_t x_374; +lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_207; uint8_t x_208; lean_dec(x_5); +x_195 = lean_ctor_get(x_3, 0); +lean_inc(x_195); lean_inc(x_2); lean_inc(x_1); -x_357 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1___boxed), 12, 2); -lean_closure_set(x_357, 0, x_1); -lean_closure_set(x_357, 1, x_2); -x_373 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__4; -x_374 = l_Lean_Expr_isConstOf(x_3, x_373); -if (x_374 == 0) +x_196 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1___boxed), 12, 2); +lean_closure_set(x_196, 0, x_1); +lean_closure_set(x_196, 1, x_2); +x_207 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__4; +x_208 = l_Lean_Expr_isConstOf(x_3, x_207); +lean_dec(x_3); +if (x_208 == 0) { -lean_object* x_375; -x_375 = lean_box(0); -x_358 = x_375; -goto block_372; +lean_object* x_209; +x_209 = lean_box(0); +x_197 = x_209; +goto block_206; } else { -lean_object* x_376; lean_object* x_377; uint8_t x_378; -x_376 = lean_array_get_size(x_4); -x_377 = lean_unsigned_to_nat(2u); -x_378 = lean_nat_dec_eq(x_376, x_377); -if (x_378 == 0) +lean_object* x_210; lean_object* x_211; uint8_t x_212; +x_210 = lean_array_get_size(x_4); +x_211 = lean_unsigned_to_nat(2u); +x_212 = lean_nat_dec_eq(x_210, x_211); +if (x_212 == 0) { -lean_object* x_379; -lean_dec(x_376); -x_379 = lean_box(0); -x_358 = x_379; -goto block_372; +lean_object* x_213; +lean_dec(x_210); +x_213 = lean_box(0); +x_197 = x_213; +goto block_206; } else { -lean_object* x_380; uint8_t x_381; -lean_dec(x_357); -lean_dec(x_3); -x_380 = lean_unsigned_to_nat(0u); -x_381 = lean_nat_dec_lt(x_380, x_376); -lean_dec(x_376); -if (x_381 == 0) +lean_object* x_214; uint8_t x_215; +lean_dec(x_196); +lean_dec(x_195); +x_214 = lean_unsigned_to_nat(0u); +x_215 = lean_nat_dec_lt(x_214, x_210); +lean_dec(x_210); +if (x_215 == 0) { -lean_object* x_382; lean_object* x_383; lean_object* x_384; +lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_dec(x_4); -x_382 = l_Lean_instInhabitedExpr; -x_383 = l_outOfBounds___rarg(x_382); +x_216 = l_Lean_instInhabitedExpr; +x_217 = l_outOfBounds___rarg(x_216); lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); @@ -6084,30 +5892,30 @@ lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_2); -lean_inc(x_383); -x_384 = l_Lean_Meta_Grind_internalize(x_383, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_384) == 0) +lean_inc(x_217); +x_218 = l_Lean_Meta_Grind_internalize(x_217, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_218) == 0) { -lean_object* x_385; lean_object* x_386; lean_object* x_387; lean_object* x_388; lean_object* x_389; -x_385 = lean_ctor_get(x_384, 1); -lean_inc(x_385); -lean_dec(x_384); +lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; +x_219 = lean_ctor_get(x_218, 1); +lean_inc(x_219); +lean_dec(x_218); lean_inc(x_1); -x_386 = l_Lean_Meta_Grind_registerParent(x_1, x_383, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_385); -lean_dec(x_383); -x_387 = lean_ctor_get(x_386, 0); -lean_inc(x_387); -x_388 = lean_ctor_get(x_386, 1); -lean_inc(x_388); -lean_dec(x_386); -x_389 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_387, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_388); -lean_dec(x_387); -return x_389; +x_220 = l_Lean_Meta_Grind_registerParent(x_1, x_217, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_219); +lean_dec(x_217); +x_221 = lean_ctor_get(x_220, 0); +lean_inc(x_221); +x_222 = lean_ctor_get(x_220, 1); +lean_inc(x_222); +lean_dec(x_220); +x_223 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_221, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_222); +lean_dec(x_221); +return x_223; } else { -uint8_t x_390; -lean_dec(x_383); +uint8_t x_224; +lean_dec(x_217); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); @@ -6118,30 +5926,30 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_2); lean_dec(x_1); -x_390 = !lean_is_exclusive(x_384); -if (x_390 == 0) +x_224 = !lean_is_exclusive(x_218); +if (x_224 == 0) { -return x_384; +return x_218; } else { -lean_object* x_391; lean_object* x_392; lean_object* x_393; -x_391 = lean_ctor_get(x_384, 0); -x_392 = lean_ctor_get(x_384, 1); -lean_inc(x_392); -lean_inc(x_391); -lean_dec(x_384); -x_393 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_393, 0, x_391); -lean_ctor_set(x_393, 1, x_392); -return x_393; +lean_object* x_225; lean_object* x_226; lean_object* x_227; +x_225 = lean_ctor_get(x_218, 0); +x_226 = lean_ctor_get(x_218, 1); +lean_inc(x_226); +lean_inc(x_225); +lean_dec(x_218); +x_227 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_227, 0, x_225); +lean_ctor_set(x_227, 1, x_226); +return x_227; } } } else { -lean_object* x_394; lean_object* x_395; -x_394 = lean_array_fget(x_4, x_380); +lean_object* x_228; lean_object* x_229; +x_228 = lean_array_fget(x_4, x_214); lean_dec(x_4); lean_inc(x_13); lean_inc(x_12); @@ -6152,30 +5960,30 @@ lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_2); -lean_inc(x_394); -x_395 = l_Lean_Meta_Grind_internalize(x_394, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_395) == 0) -{ -lean_object* x_396; lean_object* x_397; lean_object* x_398; lean_object* x_399; lean_object* x_400; -x_396 = lean_ctor_get(x_395, 1); -lean_inc(x_396); -lean_dec(x_395); +lean_inc(x_228); +x_229 = l_Lean_Meta_Grind_internalize(x_228, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_229) == 0) +{ +lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; +x_230 = lean_ctor_get(x_229, 1); +lean_inc(x_230); +lean_dec(x_229); lean_inc(x_1); -x_397 = l_Lean_Meta_Grind_registerParent(x_1, x_394, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_396); -lean_dec(x_394); -x_398 = lean_ctor_get(x_397, 0); -lean_inc(x_398); -x_399 = lean_ctor_get(x_397, 1); -lean_inc(x_399); -lean_dec(x_397); -x_400 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_398, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_399); -lean_dec(x_398); -return x_400; +x_231 = l_Lean_Meta_Grind_registerParent(x_1, x_228, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_230); +lean_dec(x_228); +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_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_232, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_233); +lean_dec(x_232); +return x_234; } else { -uint8_t x_401; -lean_dec(x_394); +uint8_t x_235; +lean_dec(x_228); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); @@ -6186,36 +5994,32 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_2); lean_dec(x_1); -x_401 = !lean_is_exclusive(x_395); -if (x_401 == 0) +x_235 = !lean_is_exclusive(x_229); +if (x_235 == 0) { -return x_395; +return x_229; } else { -lean_object* x_402; lean_object* x_403; lean_object* x_404; -x_402 = lean_ctor_get(x_395, 0); -x_403 = lean_ctor_get(x_395, 1); -lean_inc(x_403); -lean_inc(x_402); -lean_dec(x_395); -x_404 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_404, 0, x_402); -lean_ctor_set(x_404, 1, x_403); -return x_404; +lean_object* x_236; lean_object* x_237; lean_object* x_238; +x_236 = lean_ctor_get(x_229, 0); +x_237 = lean_ctor_get(x_229, 1); +lean_inc(x_237); +lean_inc(x_236); +lean_dec(x_229); +x_238 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_238, 0, x_236); +lean_ctor_set(x_238, 1, x_237); +return x_238; } } } } } -block_372: -{ -uint8_t x_359; -lean_dec(x_358); -x_359 = l_Lean_Expr_isConst(x_3); -if (x_359 == 0) +block_206: { -lean_object* x_360; +lean_object* x_198; +lean_dec(x_197); lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); @@ -6225,31 +6029,25 @@ lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_2); -lean_inc(x_3); -x_360 = l_Lean_Meta_Grind_internalize(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_360) == 0) -{ -lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; -x_361 = lean_ctor_get(x_360, 1); -lean_inc(x_361); -lean_dec(x_360); -lean_inc(x_1); -x_362 = l_Lean_Meta_Grind_registerParent(x_1, x_3, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_361); -lean_dec(x_3); -x_363 = lean_ctor_get(x_362, 0); -lean_inc(x_363); -x_364 = lean_ctor_get(x_362, 1); -lean_inc(x_364); -lean_dec(x_362); -x_365 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(x_4, x_1, x_2, x_357, x_363, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_364); -lean_dec(x_363); +x_198 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns(x_195, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +lean_dec(x_195); +if (lean_obj_tag(x_198) == 0) +{ +lean_object* x_199; lean_object* x_200; lean_object* x_201; +x_199 = lean_ctor_get(x_198, 0); +lean_inc(x_199); +x_200 = lean_ctor_get(x_198, 1); +lean_inc(x_200); +lean_dec(x_198); +x_201 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(x_4, x_1, x_2, x_196, x_199, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_200); +lean_dec(x_199); lean_dec(x_4); -return x_365; +return x_201; } else { -uint8_t x_366; -lean_dec(x_357); +uint8_t x_202; +lean_dec(x_196); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); @@ -6259,86 +6057,92 @@ 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_366 = !lean_is_exclusive(x_360); -if (x_366 == 0) +x_202 = !lean_is_exclusive(x_198); +if (x_202 == 0) { -return x_360; +return x_198; } else { -lean_object* x_367; lean_object* x_368; lean_object* x_369; -x_367 = lean_ctor_get(x_360, 0); -x_368 = lean_ctor_get(x_360, 1); -lean_inc(x_368); -lean_inc(x_367); -lean_dec(x_360); -x_369 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_369, 0, x_367); -lean_ctor_set(x_369, 1, x_368); -return x_369; +lean_object* x_203; lean_object* x_204; lean_object* x_205; +x_203 = lean_ctor_get(x_198, 0); +x_204 = lean_ctor_get(x_198, 1); +lean_inc(x_204); +lean_inc(x_203); +lean_dec(x_198); +x_205 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_205, 0, x_203); +lean_ctor_set(x_205, 1, x_204); +return x_205; } } } -else +} +case 5: { -lean_object* x_370; lean_object* x_371; +lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; +x_239 = lean_ctor_get(x_3, 0); +lean_inc(x_239); +x_240 = lean_ctor_get(x_3, 1); +lean_inc(x_240); lean_dec(x_3); -x_370 = lean_box(0); -x_371 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(x_4, x_1, x_2, x_357, x_370, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -lean_dec(x_4); -return x_371; -} -} +x_241 = lean_array_set(x_4, x_5, x_240); +x_242 = lean_unsigned_to_nat(1u); +x_243 = lean_nat_sub(x_5, x_242); +lean_dec(x_5); +x_3 = x_239; +x_4 = x_241; +x_5 = x_243; +goto _start; } -case 9: +case 6: { -lean_object* x_405; lean_object* x_406; lean_object* x_421; uint8_t x_422; +lean_object* x_245; lean_object* x_246; lean_object* x_258; uint8_t x_259; lean_dec(x_5); lean_inc(x_2); lean_inc(x_1); -x_405 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1___boxed), 12, 2); -lean_closure_set(x_405, 0, x_1); -lean_closure_set(x_405, 1, x_2); -x_421 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__4; -x_422 = l_Lean_Expr_isConstOf(x_3, x_421); -if (x_422 == 0) +x_245 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1___boxed), 12, 2); +lean_closure_set(x_245, 0, x_1); +lean_closure_set(x_245, 1, x_2); +x_258 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__4; +x_259 = l_Lean_Expr_isConstOf(x_3, x_258); +if (x_259 == 0) { -lean_object* x_423; -x_423 = lean_box(0); -x_406 = x_423; -goto block_420; +lean_object* x_260; +x_260 = lean_box(0); +x_246 = x_260; +goto block_257; } else { -lean_object* x_424; lean_object* x_425; uint8_t x_426; -x_424 = lean_array_get_size(x_4); -x_425 = lean_unsigned_to_nat(2u); -x_426 = lean_nat_dec_eq(x_424, x_425); -if (x_426 == 0) +lean_object* x_261; lean_object* x_262; uint8_t x_263; +x_261 = lean_array_get_size(x_4); +x_262 = lean_unsigned_to_nat(2u); +x_263 = lean_nat_dec_eq(x_261, x_262); +if (x_263 == 0) { -lean_object* x_427; -lean_dec(x_424); -x_427 = lean_box(0); -x_406 = x_427; -goto block_420; +lean_object* x_264; +lean_dec(x_261); +x_264 = lean_box(0); +x_246 = x_264; +goto block_257; } else { -lean_object* x_428; uint8_t x_429; -lean_dec(x_405); +lean_object* x_265; uint8_t x_266; +lean_dec(x_245); lean_dec(x_3); -x_428 = lean_unsigned_to_nat(0u); -x_429 = lean_nat_dec_lt(x_428, x_424); -lean_dec(x_424); -if (x_429 == 0) +x_265 = lean_unsigned_to_nat(0u); +x_266 = lean_nat_dec_lt(x_265, x_261); +lean_dec(x_261); +if (x_266 == 0) { -lean_object* x_430; lean_object* x_431; lean_object* x_432; +lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_dec(x_4); -x_430 = l_Lean_instInhabitedExpr; -x_431 = l_outOfBounds___rarg(x_430); +x_267 = l_Lean_instInhabitedExpr; +x_268 = l_outOfBounds___rarg(x_267); lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); @@ -6348,30 +6152,30 @@ lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_2); -lean_inc(x_431); -x_432 = l_Lean_Meta_Grind_internalize(x_431, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_432) == 0) +lean_inc(x_268); +x_269 = l_Lean_Meta_Grind_internalize(x_268, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_269) == 0) { -lean_object* x_433; lean_object* x_434; lean_object* x_435; lean_object* x_436; lean_object* x_437; -x_433 = lean_ctor_get(x_432, 1); -lean_inc(x_433); -lean_dec(x_432); +lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; +x_270 = lean_ctor_get(x_269, 1); +lean_inc(x_270); +lean_dec(x_269); lean_inc(x_1); -x_434 = l_Lean_Meta_Grind_registerParent(x_1, x_431, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_433); -lean_dec(x_431); -x_435 = lean_ctor_get(x_434, 0); -lean_inc(x_435); -x_436 = lean_ctor_get(x_434, 1); -lean_inc(x_436); -lean_dec(x_434); -x_437 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_435, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_436); -lean_dec(x_435); -return x_437; +x_271 = l_Lean_Meta_Grind_registerParent(x_1, x_268, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_270); +lean_dec(x_268); +x_272 = lean_ctor_get(x_271, 0); +lean_inc(x_272); +x_273 = lean_ctor_get(x_271, 1); +lean_inc(x_273); +lean_dec(x_271); +x_274 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_272, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_273); +lean_dec(x_272); +return x_274; } else { -uint8_t x_438; -lean_dec(x_431); +uint8_t x_275; +lean_dec(x_268); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); @@ -6382,30 +6186,30 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_2); lean_dec(x_1); -x_438 = !lean_is_exclusive(x_432); -if (x_438 == 0) +x_275 = !lean_is_exclusive(x_269); +if (x_275 == 0) { -return x_432; +return x_269; } else { -lean_object* x_439; lean_object* x_440; lean_object* x_441; -x_439 = lean_ctor_get(x_432, 0); -x_440 = lean_ctor_get(x_432, 1); -lean_inc(x_440); -lean_inc(x_439); -lean_dec(x_432); -x_441 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_441, 0, x_439); -lean_ctor_set(x_441, 1, x_440); -return x_441; +lean_object* x_276; lean_object* x_277; lean_object* x_278; +x_276 = lean_ctor_get(x_269, 0); +x_277 = lean_ctor_get(x_269, 1); +lean_inc(x_277); +lean_inc(x_276); +lean_dec(x_269); +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 { -lean_object* x_442; lean_object* x_443; -x_442 = lean_array_fget(x_4, x_428); +lean_object* x_279; lean_object* x_280; +x_279 = lean_array_fget(x_4, x_265); lean_dec(x_4); lean_inc(x_13); lean_inc(x_12); @@ -6416,30 +6220,30 @@ lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_2); -lean_inc(x_442); -x_443 = l_Lean_Meta_Grind_internalize(x_442, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_443) == 0) -{ -lean_object* x_444; lean_object* x_445; lean_object* x_446; lean_object* x_447; lean_object* x_448; -x_444 = lean_ctor_get(x_443, 1); -lean_inc(x_444); -lean_dec(x_443); +lean_inc(x_279); +x_280 = l_Lean_Meta_Grind_internalize(x_279, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_280) == 0) +{ +lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; +x_281 = lean_ctor_get(x_280, 1); +lean_inc(x_281); +lean_dec(x_280); lean_inc(x_1); -x_445 = l_Lean_Meta_Grind_registerParent(x_1, x_442, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_444); -lean_dec(x_442); -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_448 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_446, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_447); -lean_dec(x_446); -return x_448; +x_282 = l_Lean_Meta_Grind_registerParent(x_1, x_279, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_281); +lean_dec(x_279); +x_283 = lean_ctor_get(x_282, 0); +lean_inc(x_283); +x_284 = lean_ctor_get(x_282, 1); +lean_inc(x_284); +lean_dec(x_282); +x_285 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_283, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_284); +lean_dec(x_283); +return x_285; } else { -uint8_t x_449; -lean_dec(x_442); +uint8_t x_286; +lean_dec(x_279); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); @@ -6450,36 +6254,32 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_2); lean_dec(x_1); -x_449 = !lean_is_exclusive(x_443); -if (x_449 == 0) +x_286 = !lean_is_exclusive(x_280); +if (x_286 == 0) { -return x_443; +return x_280; } else { -lean_object* x_450; lean_object* x_451; lean_object* x_452; -x_450 = lean_ctor_get(x_443, 0); -x_451 = lean_ctor_get(x_443, 1); -lean_inc(x_451); -lean_inc(x_450); -lean_dec(x_443); -x_452 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_452, 0, x_450); -lean_ctor_set(x_452, 1, x_451); -return x_452; +lean_object* x_287; lean_object* x_288; lean_object* x_289; +x_287 = lean_ctor_get(x_280, 0); +x_288 = lean_ctor_get(x_280, 1); +lean_inc(x_288); +lean_inc(x_287); +lean_dec(x_280); +x_289 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_289, 0, x_287); +lean_ctor_set(x_289, 1, x_288); +return x_289; } } } } } -block_420: -{ -uint8_t x_407; -lean_dec(x_406); -x_407 = l_Lean_Expr_isConst(x_3); -if (x_407 == 0) +block_257: { -lean_object* x_408; +lean_object* x_247; +lean_dec(x_246); lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); @@ -6490,30 +6290,30 @@ lean_inc(x_7); lean_inc(x_6); lean_inc(x_2); lean_inc(x_3); -x_408 = l_Lean_Meta_Grind_internalize(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_408) == 0) +x_247 = l_Lean_Meta_Grind_internalize(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_247) == 0) { -lean_object* x_409; lean_object* x_410; lean_object* x_411; lean_object* x_412; lean_object* x_413; -x_409 = lean_ctor_get(x_408, 1); -lean_inc(x_409); -lean_dec(x_408); +lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; +x_248 = lean_ctor_get(x_247, 1); +lean_inc(x_248); +lean_dec(x_247); lean_inc(x_1); -x_410 = l_Lean_Meta_Grind_registerParent(x_1, x_3, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_409); +x_249 = l_Lean_Meta_Grind_registerParent(x_1, x_3, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_248); lean_dec(x_3); -x_411 = lean_ctor_get(x_410, 0); -lean_inc(x_411); -x_412 = lean_ctor_get(x_410, 1); -lean_inc(x_412); -lean_dec(x_410); -x_413 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(x_4, x_1, x_2, x_405, x_411, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_412); -lean_dec(x_411); +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_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(x_4, x_1, x_2, x_245, x_250, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_251); +lean_dec(x_250); lean_dec(x_4); -return x_413; +return x_252; } else { -uint8_t x_414; -lean_dec(x_405); +uint8_t x_253; +lean_dec(x_245); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); @@ -6526,83 +6326,73 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_414 = !lean_is_exclusive(x_408); -if (x_414 == 0) +x_253 = !lean_is_exclusive(x_247); +if (x_253 == 0) { -return x_408; +return x_247; } else { -lean_object* x_415; lean_object* x_416; lean_object* x_417; -x_415 = lean_ctor_get(x_408, 0); -x_416 = lean_ctor_get(x_408, 1); -lean_inc(x_416); -lean_inc(x_415); -lean_dec(x_408); -x_417 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_417, 0, x_415); -lean_ctor_set(x_417, 1, x_416); -return x_417; -} -} +lean_object* x_254; lean_object* x_255; lean_object* x_256; +x_254 = lean_ctor_get(x_247, 0); +x_255 = lean_ctor_get(x_247, 1); +lean_inc(x_255); +lean_inc(x_254); +lean_dec(x_247); +x_256 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_256, 0, x_254); +lean_ctor_set(x_256, 1, x_255); +return x_256; } -else -{ -lean_object* x_418; lean_object* x_419; -lean_dec(x_3); -x_418 = lean_box(0); -x_419 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(x_4, x_1, x_2, x_405, x_418, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -lean_dec(x_4); -return x_419; } } } -case 10: +case 7: { -lean_object* x_453; lean_object* x_454; lean_object* x_469; uint8_t x_470; +lean_object* x_290; lean_object* x_291; lean_object* x_303; uint8_t x_304; lean_dec(x_5); lean_inc(x_2); lean_inc(x_1); -x_453 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1___boxed), 12, 2); -lean_closure_set(x_453, 0, x_1); -lean_closure_set(x_453, 1, x_2); -x_469 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__4; -x_470 = l_Lean_Expr_isConstOf(x_3, x_469); -if (x_470 == 0) +x_290 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1___boxed), 12, 2); +lean_closure_set(x_290, 0, x_1); +lean_closure_set(x_290, 1, x_2); +x_303 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__4; +x_304 = l_Lean_Expr_isConstOf(x_3, x_303); +if (x_304 == 0) { -lean_object* x_471; -x_471 = lean_box(0); -x_454 = x_471; -goto block_468; +lean_object* x_305; +x_305 = lean_box(0); +x_291 = x_305; +goto block_302; } else { -lean_object* x_472; lean_object* x_473; uint8_t x_474; -x_472 = lean_array_get_size(x_4); -x_473 = lean_unsigned_to_nat(2u); -x_474 = lean_nat_dec_eq(x_472, x_473); -if (x_474 == 0) +lean_object* x_306; lean_object* x_307; uint8_t x_308; +x_306 = lean_array_get_size(x_4); +x_307 = lean_unsigned_to_nat(2u); +x_308 = lean_nat_dec_eq(x_306, x_307); +if (x_308 == 0) { -lean_object* x_475; -lean_dec(x_472); -x_475 = lean_box(0); -x_454 = x_475; -goto block_468; +lean_object* x_309; +lean_dec(x_306); +x_309 = lean_box(0); +x_291 = x_309; +goto block_302; } else { -lean_object* x_476; uint8_t x_477; -lean_dec(x_453); +lean_object* x_310; uint8_t x_311; +lean_dec(x_290); lean_dec(x_3); -x_476 = lean_unsigned_to_nat(0u); -x_477 = lean_nat_dec_lt(x_476, x_472); -lean_dec(x_472); -if (x_477 == 0) +x_310 = lean_unsigned_to_nat(0u); +x_311 = lean_nat_dec_lt(x_310, x_306); +lean_dec(x_306); +if (x_311 == 0) { -lean_object* x_478; lean_object* x_479; lean_object* x_480; +lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_dec(x_4); -x_478 = l_Lean_instInhabitedExpr; -x_479 = l_outOfBounds___rarg(x_478); +x_312 = l_Lean_instInhabitedExpr; +x_313 = l_outOfBounds___rarg(x_312); lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); @@ -6612,30 +6402,30 @@ lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_2); -lean_inc(x_479); -x_480 = l_Lean_Meta_Grind_internalize(x_479, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_480) == 0) +lean_inc(x_313); +x_314 = l_Lean_Meta_Grind_internalize(x_313, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_314) == 0) { -lean_object* x_481; lean_object* x_482; lean_object* x_483; lean_object* x_484; lean_object* x_485; -x_481 = lean_ctor_get(x_480, 1); -lean_inc(x_481); -lean_dec(x_480); +lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; +x_315 = lean_ctor_get(x_314, 1); +lean_inc(x_315); +lean_dec(x_314); lean_inc(x_1); -x_482 = l_Lean_Meta_Grind_registerParent(x_1, x_479, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_481); -lean_dec(x_479); -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 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_483, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_484); -lean_dec(x_483); -return x_485; +x_316 = l_Lean_Meta_Grind_registerParent(x_1, x_313, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_315); +lean_dec(x_313); +x_317 = lean_ctor_get(x_316, 0); +lean_inc(x_317); +x_318 = lean_ctor_get(x_316, 1); +lean_inc(x_318); +lean_dec(x_316); +x_319 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_317, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_318); +lean_dec(x_317); +return x_319; } else { -uint8_t x_486; -lean_dec(x_479); +uint8_t x_320; +lean_dec(x_313); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); @@ -6646,30 +6436,30 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_2); lean_dec(x_1); -x_486 = !lean_is_exclusive(x_480); -if (x_486 == 0) +x_320 = !lean_is_exclusive(x_314); +if (x_320 == 0) { -return x_480; +return x_314; } else { -lean_object* x_487; lean_object* x_488; lean_object* x_489; -x_487 = lean_ctor_get(x_480, 0); -x_488 = lean_ctor_get(x_480, 1); -lean_inc(x_488); -lean_inc(x_487); -lean_dec(x_480); -x_489 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_489, 0, x_487); -lean_ctor_set(x_489, 1, x_488); -return x_489; +lean_object* x_321; lean_object* x_322; lean_object* x_323; +x_321 = lean_ctor_get(x_314, 0); +x_322 = lean_ctor_get(x_314, 1); +lean_inc(x_322); +lean_inc(x_321); +lean_dec(x_314); +x_323 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_323, 0, x_321); +lean_ctor_set(x_323, 1, x_322); +return x_323; } } } else { -lean_object* x_490; lean_object* x_491; -x_490 = lean_array_fget(x_4, x_476); +lean_object* x_324; lean_object* x_325; +x_324 = lean_array_fget(x_4, x_310); lean_dec(x_4); lean_inc(x_13); lean_inc(x_12); @@ -6680,30 +6470,30 @@ lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_2); -lean_inc(x_490); -x_491 = l_Lean_Meta_Grind_internalize(x_490, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_491) == 0) -{ -lean_object* x_492; lean_object* x_493; lean_object* x_494; lean_object* x_495; lean_object* x_496; -x_492 = lean_ctor_get(x_491, 1); -lean_inc(x_492); -lean_dec(x_491); +lean_inc(x_324); +x_325 = l_Lean_Meta_Grind_internalize(x_324, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_325) == 0) +{ +lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; +x_326 = lean_ctor_get(x_325, 1); +lean_inc(x_326); +lean_dec(x_325); lean_inc(x_1); -x_493 = l_Lean_Meta_Grind_registerParent(x_1, x_490, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_492); -lean_dec(x_490); -x_494 = lean_ctor_get(x_493, 0); -lean_inc(x_494); -x_495 = lean_ctor_get(x_493, 1); -lean_inc(x_495); -lean_dec(x_493); -x_496 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_494, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_495); -lean_dec(x_494); -return x_496; +x_327 = l_Lean_Meta_Grind_registerParent(x_1, x_324, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_326); +lean_dec(x_324); +x_328 = lean_ctor_get(x_327, 0); +lean_inc(x_328); +x_329 = lean_ctor_get(x_327, 1); +lean_inc(x_329); +lean_dec(x_327); +x_330 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_328, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_329); +lean_dec(x_328); +return x_330; } else { -uint8_t x_497; -lean_dec(x_490); +uint8_t x_331; +lean_dec(x_324); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); @@ -6714,36 +6504,32 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_2); lean_dec(x_1); -x_497 = !lean_is_exclusive(x_491); -if (x_497 == 0) +x_331 = !lean_is_exclusive(x_325); +if (x_331 == 0) { -return x_491; +return x_325; } else { -lean_object* x_498; lean_object* x_499; lean_object* x_500; -x_498 = lean_ctor_get(x_491, 0); -x_499 = lean_ctor_get(x_491, 1); -lean_inc(x_499); -lean_inc(x_498); -lean_dec(x_491); -x_500 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_500, 0, x_498); -lean_ctor_set(x_500, 1, x_499); -return x_500; +lean_object* x_332; lean_object* x_333; lean_object* x_334; +x_332 = lean_ctor_get(x_325, 0); +x_333 = lean_ctor_get(x_325, 1); +lean_inc(x_333); +lean_inc(x_332); +lean_dec(x_325); +x_334 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_334, 0, x_332); +lean_ctor_set(x_334, 1, x_333); +return x_334; } } } } } -block_468: +block_302: { -uint8_t x_455; -lean_dec(x_454); -x_455 = l_Lean_Expr_isConst(x_3); -if (x_455 == 0) -{ -lean_object* x_456; +lean_object* x_292; +lean_dec(x_291); lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); @@ -6754,30 +6540,30 @@ lean_inc(x_7); lean_inc(x_6); lean_inc(x_2); lean_inc(x_3); -x_456 = l_Lean_Meta_Grind_internalize(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_456) == 0) +x_292 = l_Lean_Meta_Grind_internalize(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_292) == 0) { -lean_object* x_457; lean_object* x_458; lean_object* x_459; lean_object* x_460; lean_object* x_461; -x_457 = lean_ctor_get(x_456, 1); -lean_inc(x_457); -lean_dec(x_456); +lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; +x_293 = lean_ctor_get(x_292, 1); +lean_inc(x_293); +lean_dec(x_292); lean_inc(x_1); -x_458 = l_Lean_Meta_Grind_registerParent(x_1, x_3, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_457); +x_294 = l_Lean_Meta_Grind_registerParent(x_1, x_3, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_293); lean_dec(x_3); -x_459 = lean_ctor_get(x_458, 0); -lean_inc(x_459); -x_460 = lean_ctor_get(x_458, 1); -lean_inc(x_460); -lean_dec(x_458); -x_461 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(x_4, x_1, x_2, x_453, x_459, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_460); -lean_dec(x_459); +x_295 = lean_ctor_get(x_294, 0); +lean_inc(x_295); +x_296 = lean_ctor_get(x_294, 1); +lean_inc(x_296); +lean_dec(x_294); +x_297 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(x_4, x_1, x_2, x_290, x_295, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_296); +lean_dec(x_295); lean_dec(x_4); -return x_461; +return x_297; } else { -uint8_t x_462; -lean_dec(x_453); +uint8_t x_298; +lean_dec(x_290); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); @@ -6790,83 +6576,73 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_462 = !lean_is_exclusive(x_456); -if (x_462 == 0) +x_298 = !lean_is_exclusive(x_292); +if (x_298 == 0) { -return x_456; +return x_292; } else { -lean_object* x_463; lean_object* x_464; lean_object* x_465; -x_463 = lean_ctor_get(x_456, 0); -x_464 = lean_ctor_get(x_456, 1); -lean_inc(x_464); -lean_inc(x_463); -lean_dec(x_456); -x_465 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_465, 0, x_463); -lean_ctor_set(x_465, 1, x_464); -return x_465; -} -} +lean_object* x_299; lean_object* x_300; lean_object* x_301; +x_299 = lean_ctor_get(x_292, 0); +x_300 = lean_ctor_get(x_292, 1); +lean_inc(x_300); +lean_inc(x_299); +lean_dec(x_292); +x_301 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_301, 0, x_299); +lean_ctor_set(x_301, 1, x_300); +return x_301; } -else -{ -lean_object* x_466; lean_object* x_467; -lean_dec(x_3); -x_466 = lean_box(0); -x_467 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(x_4, x_1, x_2, x_453, x_466, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -lean_dec(x_4); -return x_467; } } } -default: +case 8: { -lean_object* x_501; lean_object* x_502; lean_object* x_517; uint8_t x_518; +lean_object* x_335; lean_object* x_336; lean_object* x_348; uint8_t x_349; lean_dec(x_5); lean_inc(x_2); lean_inc(x_1); -x_501 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1___boxed), 12, 2); -lean_closure_set(x_501, 0, x_1); -lean_closure_set(x_501, 1, x_2); -x_517 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__4; -x_518 = l_Lean_Expr_isConstOf(x_3, x_517); -if (x_518 == 0) +x_335 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1___boxed), 12, 2); +lean_closure_set(x_335, 0, x_1); +lean_closure_set(x_335, 1, x_2); +x_348 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__4; +x_349 = l_Lean_Expr_isConstOf(x_3, x_348); +if (x_349 == 0) { -lean_object* x_519; -x_519 = lean_box(0); -x_502 = x_519; -goto block_516; +lean_object* x_350; +x_350 = lean_box(0); +x_336 = x_350; +goto block_347; } else { -lean_object* x_520; lean_object* x_521; uint8_t x_522; -x_520 = lean_array_get_size(x_4); -x_521 = lean_unsigned_to_nat(2u); -x_522 = lean_nat_dec_eq(x_520, x_521); -if (x_522 == 0) +lean_object* x_351; lean_object* x_352; uint8_t x_353; +x_351 = lean_array_get_size(x_4); +x_352 = lean_unsigned_to_nat(2u); +x_353 = lean_nat_dec_eq(x_351, x_352); +if (x_353 == 0) { -lean_object* x_523; -lean_dec(x_520); -x_523 = lean_box(0); -x_502 = x_523; -goto block_516; +lean_object* x_354; +lean_dec(x_351); +x_354 = lean_box(0); +x_336 = x_354; +goto block_347; } else { -lean_object* x_524; uint8_t x_525; -lean_dec(x_501); +lean_object* x_355; uint8_t x_356; +lean_dec(x_335); lean_dec(x_3); -x_524 = lean_unsigned_to_nat(0u); -x_525 = lean_nat_dec_lt(x_524, x_520); -lean_dec(x_520); -if (x_525 == 0) +x_355 = lean_unsigned_to_nat(0u); +x_356 = lean_nat_dec_lt(x_355, x_351); +lean_dec(x_351); +if (x_356 == 0) { -lean_object* x_526; lean_object* x_527; lean_object* x_528; +lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_dec(x_4); -x_526 = l_Lean_instInhabitedExpr; -x_527 = l_outOfBounds___rarg(x_526); +x_357 = l_Lean_instInhabitedExpr; +x_358 = l_outOfBounds___rarg(x_357); lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); @@ -6876,30 +6652,30 @@ lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_2); -lean_inc(x_527); -x_528 = l_Lean_Meta_Grind_internalize(x_527, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_528) == 0) -{ -lean_object* x_529; lean_object* x_530; lean_object* x_531; lean_object* x_532; lean_object* x_533; -x_529 = lean_ctor_get(x_528, 1); -lean_inc(x_529); -lean_dec(x_528); +lean_inc(x_358); +x_359 = l_Lean_Meta_Grind_internalize(x_358, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_359) == 0) +{ +lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; +x_360 = lean_ctor_get(x_359, 1); +lean_inc(x_360); +lean_dec(x_359); lean_inc(x_1); -x_530 = l_Lean_Meta_Grind_registerParent(x_1, x_527, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_529); -lean_dec(x_527); -x_531 = lean_ctor_get(x_530, 0); -lean_inc(x_531); -x_532 = lean_ctor_get(x_530, 1); -lean_inc(x_532); -lean_dec(x_530); -x_533 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_531, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_532); -lean_dec(x_531); -return x_533; +x_361 = l_Lean_Meta_Grind_registerParent(x_1, x_358, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_360); +lean_dec(x_358); +x_362 = lean_ctor_get(x_361, 0); +lean_inc(x_362); +x_363 = lean_ctor_get(x_361, 1); +lean_inc(x_363); +lean_dec(x_361); +x_364 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_362, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_363); +lean_dec(x_362); +return x_364; } else { -uint8_t x_534; -lean_dec(x_527); +uint8_t x_365; +lean_dec(x_358); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); @@ -6910,30 +6686,30 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_2); lean_dec(x_1); -x_534 = !lean_is_exclusive(x_528); -if (x_534 == 0) +x_365 = !lean_is_exclusive(x_359); +if (x_365 == 0) { -return x_528; +return x_359; } else { -lean_object* x_535; lean_object* x_536; lean_object* x_537; -x_535 = lean_ctor_get(x_528, 0); -x_536 = lean_ctor_get(x_528, 1); -lean_inc(x_536); -lean_inc(x_535); -lean_dec(x_528); -x_537 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_537, 0, x_535); -lean_ctor_set(x_537, 1, x_536); -return x_537; +lean_object* x_366; lean_object* x_367; lean_object* x_368; +x_366 = lean_ctor_get(x_359, 0); +x_367 = lean_ctor_get(x_359, 1); +lean_inc(x_367); +lean_inc(x_366); +lean_dec(x_359); +x_368 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_368, 0, x_366); +lean_ctor_set(x_368, 1, x_367); +return x_368; } } } else { -lean_object* x_538; lean_object* x_539; -x_538 = lean_array_fget(x_4, x_524); +lean_object* x_369; lean_object* x_370; +x_369 = lean_array_fget(x_4, x_355); lean_dec(x_4); lean_inc(x_13); lean_inc(x_12); @@ -6944,30 +6720,30 @@ lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_2); -lean_inc(x_538); -x_539 = l_Lean_Meta_Grind_internalize(x_538, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_539) == 0) -{ -lean_object* x_540; lean_object* x_541; lean_object* x_542; lean_object* x_543; lean_object* x_544; -x_540 = lean_ctor_get(x_539, 1); -lean_inc(x_540); -lean_dec(x_539); +lean_inc(x_369); +x_370 = l_Lean_Meta_Grind_internalize(x_369, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_370) == 0) +{ +lean_object* x_371; lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; +x_371 = lean_ctor_get(x_370, 1); +lean_inc(x_371); +lean_dec(x_370); lean_inc(x_1); -x_541 = l_Lean_Meta_Grind_registerParent(x_1, x_538, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_540); -lean_dec(x_538); -x_542 = lean_ctor_get(x_541, 0); -lean_inc(x_542); -x_543 = lean_ctor_get(x_541, 1); -lean_inc(x_543); -lean_dec(x_541); -x_544 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_542, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_543); -lean_dec(x_542); -return x_544; +x_372 = l_Lean_Meta_Grind_registerParent(x_1, x_369, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_371); +lean_dec(x_369); +x_373 = lean_ctor_get(x_372, 0); +lean_inc(x_373); +x_374 = lean_ctor_get(x_372, 1); +lean_inc(x_374); +lean_dec(x_372); +x_375 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_373, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_374); +lean_dec(x_373); +return x_375; } else { -uint8_t x_545; -lean_dec(x_538); +uint8_t x_376; +lean_dec(x_369); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); @@ -6978,36 +6754,32 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_2); lean_dec(x_1); -x_545 = !lean_is_exclusive(x_539); -if (x_545 == 0) +x_376 = !lean_is_exclusive(x_370); +if (x_376 == 0) { -return x_539; +return x_370; } else { -lean_object* x_546; lean_object* x_547; lean_object* x_548; -x_546 = lean_ctor_get(x_539, 0); -x_547 = lean_ctor_get(x_539, 1); -lean_inc(x_547); -lean_inc(x_546); -lean_dec(x_539); -x_548 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_548, 0, x_546); -lean_ctor_set(x_548, 1, x_547); -return x_548; +lean_object* x_377; lean_object* x_378; lean_object* x_379; +x_377 = lean_ctor_get(x_370, 0); +x_378 = lean_ctor_get(x_370, 1); +lean_inc(x_378); +lean_inc(x_377); +lean_dec(x_370); +x_379 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_379, 0, x_377); +lean_ctor_set(x_379, 1, x_378); +return x_379; } } } } } -block_516: +block_347: { -uint8_t x_503; -lean_dec(x_502); -x_503 = l_Lean_Expr_isConst(x_3); -if (x_503 == 0) -{ -lean_object* x_504; +lean_object* x_337; +lean_dec(x_336); lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); @@ -7018,30 +6790,30 @@ lean_inc(x_7); lean_inc(x_6); lean_inc(x_2); lean_inc(x_3); -x_504 = l_Lean_Meta_Grind_internalize(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_504) == 0) +x_337 = l_Lean_Meta_Grind_internalize(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_337) == 0) { -lean_object* x_505; lean_object* x_506; lean_object* x_507; lean_object* x_508; lean_object* x_509; -x_505 = lean_ctor_get(x_504, 1); -lean_inc(x_505); -lean_dec(x_504); +lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; +x_338 = lean_ctor_get(x_337, 1); +lean_inc(x_338); +lean_dec(x_337); lean_inc(x_1); -x_506 = l_Lean_Meta_Grind_registerParent(x_1, x_3, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_505); +x_339 = l_Lean_Meta_Grind_registerParent(x_1, x_3, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_338); lean_dec(x_3); -x_507 = lean_ctor_get(x_506, 0); -lean_inc(x_507); -x_508 = lean_ctor_get(x_506, 1); -lean_inc(x_508); -lean_dec(x_506); -x_509 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(x_4, x_1, x_2, x_501, x_507, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_508); -lean_dec(x_507); +x_340 = lean_ctor_get(x_339, 0); +lean_inc(x_340); +x_341 = lean_ctor_get(x_339, 1); +lean_inc(x_341); +lean_dec(x_339); +x_342 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(x_4, x_1, x_2, x_335, x_340, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_341); +lean_dec(x_340); lean_dec(x_4); -return x_509; +return x_342; } else { -uint8_t x_510; -lean_dec(x_501); +uint8_t x_343; +lean_dec(x_335); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); @@ -7054,457 +6826,4146 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_510 = !lean_is_exclusive(x_504); -if (x_510 == 0) -{ -return x_504; -} -else +x_343 = !lean_is_exclusive(x_337); +if (x_343 == 0) { -lean_object* x_511; lean_object* x_512; lean_object* x_513; -x_511 = lean_ctor_get(x_504, 0); -x_512 = lean_ctor_get(x_504, 1); -lean_inc(x_512); -lean_inc(x_511); -lean_dec(x_504); -x_513 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_513, 0, x_511); -lean_ctor_set(x_513, 1, x_512); -return x_513; -} -} +return x_337; } else { -lean_object* x_514; lean_object* x_515; -lean_dec(x_3); -x_514 = lean_box(0); -x_515 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(x_4, x_1, x_2, x_501, x_514, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -lean_dec(x_4); -return x_515; -} -} -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_internalize___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -uint8_t x_13; lean_object* x_14; -x_13 = 0; -x_14 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_13, x_13, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -return x_14; -} -} -static lean_object* _init_l_Lean_Meta_Grind_internalize___lambda__2___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean.Meta.Tactic.Grind.Internalize", 34, 34); -return x_1; +lean_object* x_344; lean_object* x_345; lean_object* x_346; +x_344 = lean_ctor_get(x_337, 0); +x_345 = lean_ctor_get(x_337, 1); +lean_inc(x_345); +lean_inc(x_344); +lean_dec(x_337); +x_346 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_346, 0, x_344); +lean_ctor_set(x_346, 1, x_345); +return x_346; } } -static lean_object* _init_l_Lean_Meta_Grind_internalize___lambda__2___closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean.Meta.Grind.internalize", 27, 27); -return x_1; } } -static lean_object* _init_l_Lean_Meta_Grind_internalize___lambda__2___closed__3() { -_start: +case 9: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("unreachable code has been reached", 33, 33); -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_Grind_internalize___lambda__2___closed__4() { -_start: +lean_object* x_380; lean_object* x_381; lean_object* x_393; uint8_t x_394; +lean_dec(x_5); +lean_inc(x_2); +lean_inc(x_1); +x_380 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1___boxed), 12, 2); +lean_closure_set(x_380, 0, x_1); +lean_closure_set(x_380, 1, x_2); +x_393 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__4; +x_394 = l_Lean_Expr_isConstOf(x_3, x_393); +if (x_394 == 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_Lean_Meta_Grind_internalize___lambda__2___closed__1; -x_2 = l_Lean_Meta_Grind_internalize___lambda__2___closed__2; -x_3 = lean_unsigned_to_nat(43u); -x_4 = lean_unsigned_to_nat(16u); -x_5 = l_Lean_Meta_Grind_internalize___lambda__2___closed__3; -x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); -return x_6; -} +lean_object* x_395; +x_395 = lean_box(0); +x_381 = x_395; +goto block_392; } -static lean_object* _init_l_Lean_Meta_Grind_internalize___lambda__2___closed__5() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("unexpected term during internalization", 38, 38); -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_Grind_internalize___lambda__2___closed__6() { -_start: +lean_object* x_396; lean_object* x_397; uint8_t x_398; +x_396 = lean_array_get_size(x_4); +x_397 = lean_unsigned_to_nat(2u); +x_398 = lean_nat_dec_eq(x_396, x_397); +if (x_398 == 0) { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_Grind_internalize___lambda__2___closed__5; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; +lean_object* x_399; +lean_dec(x_396); +x_399 = lean_box(0); +x_381 = x_399; +goto block_392; } -} -static lean_object* _init_l_Lean_Meta_Grind_internalize___lambda__2___closed__7() { -_start: +else { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_levelZero; -x_2 = l_Lean_Expr_sort___override(x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_internalize___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: +lean_object* x_400; uint8_t x_401; +lean_dec(x_380); +lean_dec(x_3); +x_400 = lean_unsigned_to_nat(0u); +x_401 = lean_nat_dec_lt(x_400, x_396); +lean_dec(x_396); +if (x_401 == 0) { -switch (lean_obj_tag(x_1)) { -case 0: +lean_object* x_402; lean_object* x_403; lean_object* x_404; +lean_dec(x_4); +x_402 = l_Lean_instInhabitedExpr; +x_403 = l_outOfBounds___rarg(x_402); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_403); +x_404 = l_Lean_Meta_Grind_internalize(x_403, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_404) == 0) { -lean_object* x_13; lean_object* x_14; -lean_dec(x_2); -lean_dec(x_1); -x_13 = l_Lean_Meta_Grind_internalize___lambda__2___closed__4; -x_14 = l_panic___at_Lean_Meta_Grind_internalize___spec__1(x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -return x_14; +lean_object* x_405; lean_object* x_406; lean_object* x_407; lean_object* x_408; lean_object* x_409; +x_405 = lean_ctor_get(x_404, 1); +lean_inc(x_405); +lean_dec(x_404); +lean_inc(x_1); +x_406 = l_Lean_Meta_Grind_registerParent(x_1, x_403, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_405); +lean_dec(x_403); +x_407 = lean_ctor_get(x_406, 0); +lean_inc(x_407); +x_408 = lean_ctor_get(x_406, 1); +lean_inc(x_408); +lean_dec(x_406); +x_409 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_407, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_408); +lean_dec(x_407); +return x_409; } -case 2: -{ -lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; -x_15 = l_Lean_Meta_Grind_addCongrTable___closed__2; -x_16 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_15, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -x_18 = lean_unbox(x_17); -lean_dec(x_17); -if (x_18 == 0) +else { -lean_object* x_19; uint8_t x_20; lean_object* x_21; -x_19 = lean_ctor_get(x_16, 1); -lean_inc(x_19); -lean_dec(x_16); -x_20 = 0; -x_21 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_20, x_20, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_19); +uint8_t x_410; +lean_dec(x_403); +lean_dec(x_13); +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -return x_21; +lean_dec(x_2); +lean_dec(x_1); +x_410 = !lean_is_exclusive(x_404); +if (x_410 == 0) +{ +return x_404; } else { -uint8_t x_22; -x_22 = !lean_is_exclusive(x_16); -if (x_22 == 0) +lean_object* x_411; lean_object* x_412; lean_object* x_413; +x_411 = lean_ctor_get(x_404, 0); +x_412 = lean_ctor_get(x_404, 1); +lean_inc(x_412); +lean_inc(x_411); +lean_dec(x_404); +x_413 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_413, 0, x_411); +lean_ctor_set(x_413, 1, x_412); +return x_413; +} +} +} +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; uint8_t x_31; lean_object* x_32; -x_23 = lean_ctor_get(x_16, 1); -x_24 = lean_ctor_get(x_16, 0); -lean_dec(x_24); -lean_inc(x_1); -x_25 = l_Lean_indentExpr(x_1); -x_26 = l_Lean_Meta_Grind_internalize___lambda__2___closed__6; -lean_ctor_set_tag(x_16, 7); -lean_ctor_set(x_16, 1, x_25); -lean_ctor_set(x_16, 0, x_26); -x_27 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__4; -x_28 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_28, 0, x_16); -lean_ctor_set(x_28, 1, x_27); -x_29 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_15, x_28, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_23); -x_30 = lean_ctor_get(x_29, 1); -lean_inc(x_30); -lean_dec(x_29); -x_31 = 0; -x_32 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_31, x_31, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_30); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +lean_object* x_414; lean_object* x_415; +x_414 = lean_array_fget(x_4, x_400); lean_dec(x_4); -return x_32; +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_414); +x_415 = l_Lean_Meta_Grind_internalize(x_414, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_415) == 0) +{ +lean_object* x_416; lean_object* x_417; lean_object* x_418; lean_object* x_419; lean_object* x_420; +x_416 = lean_ctor_get(x_415, 1); +lean_inc(x_416); +lean_dec(x_415); +lean_inc(x_1); +x_417 = l_Lean_Meta_Grind_registerParent(x_1, x_414, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_416); +lean_dec(x_414); +x_418 = lean_ctor_get(x_417, 0); +lean_inc(x_418); +x_419 = lean_ctor_get(x_417, 1); +lean_inc(x_419); +lean_dec(x_417); +x_420 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_418, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_419); +lean_dec(x_418); +return x_420; } 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; uint8_t x_41; lean_object* x_42; -x_33 = lean_ctor_get(x_16, 1); -lean_inc(x_33); -lean_dec(x_16); -lean_inc(x_1); -x_34 = l_Lean_indentExpr(x_1); -x_35 = l_Lean_Meta_Grind_internalize___lambda__2___closed__6; -x_36 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_36, 0, x_35); -lean_ctor_set(x_36, 1, x_34); -x_37 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__4; -x_38 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_38, 0, x_36); -lean_ctor_set(x_38, 1, x_37); -x_39 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_15, x_38, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_33); -x_40 = lean_ctor_get(x_39, 1); -lean_inc(x_40); -lean_dec(x_39); -x_41 = 0; -x_42 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_41, x_41, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_40); +uint8_t x_421; +lean_dec(x_414); +lean_dec(x_13); +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -return x_42; +lean_dec(x_2); +lean_dec(x_1); +x_421 = !lean_is_exclusive(x_415); +if (x_421 == 0) +{ +return x_415; +} +else +{ +lean_object* x_422; lean_object* x_423; lean_object* x_424; +x_422 = lean_ctor_get(x_415, 0); +x_423 = lean_ctor_get(x_415, 1); +lean_inc(x_423); +lean_inc(x_422); +lean_dec(x_415); +x_424 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_424, 0, x_422); +lean_ctor_set(x_424, 1, x_423); +return x_424; } } } -case 3: +} +} +block_392: { -lean_object* x_43; lean_object* x_44; +lean_object* x_382; +lean_dec(x_381); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_3); +x_382 = l_Lean_Meta_Grind_internalize(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_382) == 0) +{ +lean_object* x_383; lean_object* x_384; lean_object* x_385; lean_object* x_386; lean_object* x_387; +x_383 = lean_ctor_get(x_382, 1); +lean_inc(x_383); +lean_dec(x_382); +lean_inc(x_1); +x_384 = l_Lean_Meta_Grind_registerParent(x_1, x_3, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_383); +lean_dec(x_3); +x_385 = lean_ctor_get(x_384, 0); +lean_inc(x_385); +x_386 = lean_ctor_get(x_384, 1); +lean_inc(x_386); +lean_dec(x_384); +x_387 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(x_4, x_1, x_2, x_380, x_385, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_386); +lean_dec(x_385); +lean_dec(x_4); +return x_387; +} +else +{ +uint8_t x_388; +lean_dec(x_380); +lean_dec(x_13); +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_4); +lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_43 = lean_box(0); -x_44 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_44, 0, x_43); -lean_ctor_set(x_44, 1, x_12); -return x_44; +x_388 = !lean_is_exclusive(x_382); +if (x_388 == 0) +{ +return x_382; } -case 4: +else { -lean_object* x_45; -x_45 = l_Lean_Meta_Grind_mkENode(x_1, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_7); -lean_dec(x_6); +lean_object* x_389; lean_object* x_390; lean_object* x_391; +x_389 = lean_ctor_get(x_382, 0); +x_390 = lean_ctor_get(x_382, 1); +lean_inc(x_390); +lean_inc(x_389); +lean_dec(x_382); +x_391 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_391, 0, x_389); +lean_ctor_set(x_391, 1, x_390); +return x_391; +} +} +} +} +case 10: +{ +lean_object* x_425; lean_object* x_426; lean_object* x_438; uint8_t x_439; lean_dec(x_5); -lean_dec(x_4); -return x_45; +lean_inc(x_2); +lean_inc(x_1); +x_425 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1___boxed), 12, 2); +lean_closure_set(x_425, 0, x_1); +lean_closure_set(x_425, 1, x_2); +x_438 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__4; +x_439 = l_Lean_Expr_isConstOf(x_3, x_438); +if (x_439 == 0) +{ +lean_object* x_440; +x_440 = lean_box(0); +x_426 = x_440; +goto block_437; } -case 5: +else { -lean_object* x_46; +lean_object* x_441; lean_object* x_442; uint8_t x_443; +x_441 = lean_array_get_size(x_4); +x_442 = lean_unsigned_to_nat(2u); +x_443 = lean_nat_dec_eq(x_441, x_442); +if (x_443 == 0) +{ +lean_object* x_444; +lean_dec(x_441); +x_444 = lean_box(0); +x_426 = x_444; +goto block_437; +} +else +{ +lean_object* x_445; uint8_t x_446; +lean_dec(x_425); +lean_dec(x_3); +x_445 = lean_unsigned_to_nat(0u); +x_446 = lean_nat_dec_lt(x_445, x_441); +lean_dec(x_441); +if (x_446 == 0) +{ +lean_object* x_447; lean_object* x_448; lean_object* x_449; +lean_dec(x_4); +x_447 = l_Lean_instInhabitedExpr; +x_448 = l_outOfBounds___rarg(x_447); +lean_inc(x_13); +lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); -lean_inc(x_1); -x_46 = l_Lean_Meta_isLitValue(x_1, x_8, x_9, x_10, x_11, x_12); -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); -lean_dec(x_47); -if (x_48 == 0) +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_448); +x_449 = l_Lean_Meta_Grind_internalize(x_448, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_449) == 0) { -lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; -x_49 = lean_ctor_get(x_46, 1); -lean_inc(x_49); -lean_dec(x_46); -x_50 = lean_unsigned_to_nat(0u); -x_51 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_50); -x_52 = l_Lean_Meta_Grind_internalize___lambda__2___closed__7; -lean_inc(x_51); -x_53 = lean_mk_array(x_51, x_52); -x_54 = lean_unsigned_to_nat(1u); -x_55 = lean_nat_sub(x_51, x_54); -lean_dec(x_51); +lean_object* x_450; lean_object* x_451; lean_object* x_452; lean_object* x_453; lean_object* x_454; +x_450 = lean_ctor_get(x_449, 1); +lean_inc(x_450); +lean_dec(x_449); lean_inc(x_1); -x_56 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3(x_1, x_2, x_1, x_53, x_55, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_49); -return x_56; +x_451 = l_Lean_Meta_Grind_registerParent(x_1, x_448, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_450); +lean_dec(x_448); +x_452 = lean_ctor_get(x_451, 0); +lean_inc(x_452); +x_453 = lean_ctor_get(x_451, 1); +lean_inc(x_453); +lean_dec(x_451); +x_454 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_452, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_453); +lean_dec(x_452); +return x_454; } else { -lean_object* x_57; lean_object* x_58; -x_57 = lean_ctor_get(x_46, 1); -lean_inc(x_57); -lean_dec(x_46); -x_58 = l_Lean_Meta_Grind_mkENode(x_1, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_57); +uint8_t x_455; +lean_dec(x_448); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -return x_58; +lean_dec(x_2); +lean_dec(x_1); +x_455 = !lean_is_exclusive(x_449); +if (x_455 == 0) +{ +return x_449; +} +else +{ +lean_object* x_456; lean_object* x_457; lean_object* x_458; +x_456 = lean_ctor_get(x_449, 0); +x_457 = lean_ctor_get(x_449, 1); +lean_inc(x_457); +lean_inc(x_456); +lean_dec(x_449); +x_458 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_458, 0, x_456); +lean_ctor_set(x_458, 1, x_457); +return x_458; +} } } else { -uint8_t x_59; +lean_object* x_459; lean_object* x_460; +x_459 = lean_array_fget(x_4, x_445); +lean_dec(x_4); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_459); +x_460 = l_Lean_Meta_Grind_internalize(x_459, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_460) == 0) +{ +lean_object* x_461; lean_object* x_462; lean_object* x_463; lean_object* x_464; lean_object* x_465; +x_461 = lean_ctor_get(x_460, 1); +lean_inc(x_461); +lean_dec(x_460); +lean_inc(x_1); +x_462 = l_Lean_Meta_Grind_registerParent(x_1, x_459, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_461); +lean_dec(x_459); +x_463 = lean_ctor_get(x_462, 0); +lean_inc(x_463); +x_464 = lean_ctor_get(x_462, 1); +lean_inc(x_464); +lean_dec(x_462); +x_465 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_463, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_464); +lean_dec(x_463); +return x_465; +} +else +{ +uint8_t x_466; +lean_dec(x_459); +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_2); lean_dec(x_1); -x_59 = !lean_is_exclusive(x_46); -if (x_59 == 0) +x_466 = !lean_is_exclusive(x_460); +if (x_466 == 0) { -return x_46; +return x_460; } else { -lean_object* x_60; lean_object* x_61; lean_object* x_62; -x_60 = lean_ctor_get(x_46, 0); -x_61 = lean_ctor_get(x_46, 1); -lean_inc(x_61); -lean_inc(x_60); -lean_dec(x_46); -x_62 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_62, 0, x_60); -lean_ctor_set(x_62, 1, x_61); -return x_62; +lean_object* x_467; lean_object* x_468; lean_object* x_469; +x_467 = lean_ctor_get(x_460, 0); +x_468 = lean_ctor_get(x_460, 1); +lean_inc(x_468); +lean_inc(x_467); +lean_dec(x_460); +x_469 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_469, 0, x_467); +lean_ctor_set(x_469, 1, x_468); +return x_469; } } } -case 9: +} +} +block_437: { -lean_object* x_63; -x_63 = l_Lean_Meta_Grind_mkENode(x_1, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +lean_object* x_427; +lean_dec(x_426); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_3); +x_427 = l_Lean_Meta_Grind_internalize(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_427) == 0) +{ +lean_object* x_428; lean_object* x_429; lean_object* x_430; lean_object* x_431; lean_object* x_432; +x_428 = lean_ctor_get(x_427, 1); +lean_inc(x_428); +lean_dec(x_427); +lean_inc(x_1); +x_429 = l_Lean_Meta_Grind_registerParent(x_1, x_3, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_428); +lean_dec(x_3); +x_430 = lean_ctor_get(x_429, 0); +lean_inc(x_430); +x_431 = lean_ctor_get(x_429, 1); +lean_inc(x_431); +lean_dec(x_429); +x_432 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(x_4, x_1, x_2, x_425, x_430, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_431); +lean_dec(x_430); lean_dec(x_4); -return x_63; +return x_432; } -case 10: -{ -lean_object* x_64; lean_object* x_65; lean_object* x_66; uint8_t x_67; -x_64 = l_Lean_Meta_Grind_addCongrTable___closed__2; -x_65 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_64, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -x_66 = lean_ctor_get(x_65, 0); -lean_inc(x_66); -x_67 = lean_unbox(x_66); -lean_dec(x_66); -if (x_67 == 0) +else { -lean_object* x_68; uint8_t x_69; lean_object* x_70; -x_68 = lean_ctor_get(x_65, 1); -lean_inc(x_68); -lean_dec(x_65); -x_69 = 0; -x_70 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_69, x_69, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_68); +uint8_t x_433; +lean_dec(x_425); +lean_dec(x_13); +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_4); -return x_70; +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_433 = !lean_is_exclusive(x_427); +if (x_433 == 0) +{ +return x_427; } else { -uint8_t x_71; -x_71 = !lean_is_exclusive(x_65); -if (x_71 == 0) +lean_object* x_434; lean_object* x_435; lean_object* x_436; +x_434 = lean_ctor_get(x_427, 0); +x_435 = lean_ctor_get(x_427, 1); +lean_inc(x_435); +lean_inc(x_434); +lean_dec(x_427); +x_436 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_436, 0, x_434); +lean_ctor_set(x_436, 1, x_435); +return x_436; +} +} +} +} +default: { -lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; uint8_t x_80; lean_object* x_81; -x_72 = lean_ctor_get(x_65, 1); -x_73 = lean_ctor_get(x_65, 0); -lean_dec(x_73); +lean_object* x_470; lean_object* x_471; lean_object* x_483; uint8_t x_484; +lean_dec(x_5); +lean_inc(x_2); lean_inc(x_1); -x_74 = l_Lean_indentExpr(x_1); -x_75 = l_Lean_Meta_Grind_internalize___lambda__2___closed__6; -lean_ctor_set_tag(x_65, 7); -lean_ctor_set(x_65, 1, x_74); -lean_ctor_set(x_65, 0, x_75); -x_76 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__4; -x_77 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_77, 0, x_65); -lean_ctor_set(x_77, 1, x_76); -x_78 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_64, x_77, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_72); -x_79 = lean_ctor_get(x_78, 1); -lean_inc(x_79); -lean_dec(x_78); -x_80 = 0; -x_81 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_80, x_80, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_79); +x_470 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1___boxed), 12, 2); +lean_closure_set(x_470, 0, x_1); +lean_closure_set(x_470, 1, x_2); +x_483 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__4; +x_484 = l_Lean_Expr_isConstOf(x_3, x_483); +if (x_484 == 0) +{ +lean_object* x_485; +x_485 = lean_box(0); +x_471 = x_485; +goto block_482; +} +else +{ +lean_object* x_486; lean_object* x_487; uint8_t x_488; +x_486 = lean_array_get_size(x_4); +x_487 = lean_unsigned_to_nat(2u); +x_488 = lean_nat_dec_eq(x_486, x_487); +if (x_488 == 0) +{ +lean_object* x_489; +lean_dec(x_486); +x_489 = lean_box(0); +x_471 = x_489; +goto block_482; +} +else +{ +lean_object* x_490; uint8_t x_491; +lean_dec(x_470); +lean_dec(x_3); +x_490 = lean_unsigned_to_nat(0u); +x_491 = lean_nat_dec_lt(x_490, x_486); +lean_dec(x_486); +if (x_491 == 0) +{ +lean_object* x_492; lean_object* x_493; lean_object* x_494; +lean_dec(x_4); +x_492 = l_Lean_instInhabitedExpr; +x_493 = l_outOfBounds___rarg(x_492); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_493); +x_494 = l_Lean_Meta_Grind_internalize(x_493, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_494) == 0) +{ +lean_object* x_495; lean_object* x_496; lean_object* x_497; lean_object* x_498; lean_object* x_499; +x_495 = lean_ctor_get(x_494, 1); +lean_inc(x_495); +lean_dec(x_494); +lean_inc(x_1); +x_496 = l_Lean_Meta_Grind_registerParent(x_1, x_493, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_495); +lean_dec(x_493); +x_497 = lean_ctor_get(x_496, 0); +lean_inc(x_497); +x_498 = lean_ctor_get(x_496, 1); +lean_inc(x_498); +lean_dec(x_496); +x_499 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_497, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_498); +lean_dec(x_497); +return x_499; +} +else +{ +uint8_t x_500; +lean_dec(x_493); +lean_dec(x_13); +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -return x_81; +lean_dec(x_2); +lean_dec(x_1); +x_500 = !lean_is_exclusive(x_494); +if (x_500 == 0) +{ +return x_494; } 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; uint8_t x_90; lean_object* x_91; -x_82 = lean_ctor_get(x_65, 1); -lean_inc(x_82); -lean_dec(x_65); +lean_object* x_501; lean_object* x_502; lean_object* x_503; +x_501 = lean_ctor_get(x_494, 0); +x_502 = lean_ctor_get(x_494, 1); +lean_inc(x_502); +lean_inc(x_501); +lean_dec(x_494); +x_503 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_503, 0, x_501); +lean_ctor_set(x_503, 1, x_502); +return x_503; +} +} +} +else +{ +lean_object* x_504; lean_object* x_505; +x_504 = lean_array_fget(x_4, x_490); +lean_dec(x_4); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_504); +x_505 = l_Lean_Meta_Grind_internalize(x_504, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_505) == 0) +{ +lean_object* x_506; lean_object* x_507; lean_object* x_508; lean_object* x_509; lean_object* x_510; +x_506 = lean_ctor_get(x_505, 1); +lean_inc(x_506); +lean_dec(x_505); lean_inc(x_1); -x_83 = l_Lean_indentExpr(x_1); -x_84 = l_Lean_Meta_Grind_internalize___lambda__2___closed__6; -x_85 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_85, 0, x_84); -lean_ctor_set(x_85, 1, x_83); -x_86 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__4; -x_87 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_87, 0, x_85); -lean_ctor_set(x_87, 1, x_86); -x_88 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_64, x_87, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_82); -x_89 = lean_ctor_get(x_88, 1); -lean_inc(x_89); -lean_dec(x_88); -x_90 = 0; -x_91 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_90, x_90, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_89); +x_507 = l_Lean_Meta_Grind_registerParent(x_1, x_504, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_506); +lean_dec(x_504); +x_508 = lean_ctor_get(x_507, 0); +lean_inc(x_508); +x_509 = lean_ctor_get(x_507, 1); +lean_inc(x_509); +lean_dec(x_507); +x_510 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_508, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_509); +lean_dec(x_508); +return x_510; +} +else +{ +uint8_t x_511; +lean_dec(x_504); +lean_dec(x_13); +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -return x_91; +lean_dec(x_2); +lean_dec(x_1); +x_511 = !lean_is_exclusive(x_505); +if (x_511 == 0) +{ +return x_505; +} +else +{ +lean_object* x_512; lean_object* x_513; lean_object* x_514; +x_512 = lean_ctor_get(x_505, 0); +x_513 = lean_ctor_get(x_505, 1); +lean_inc(x_513); +lean_inc(x_512); +lean_dec(x_505); +x_514 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_514, 0, x_512); +lean_ctor_set(x_514, 1, x_513); +return x_514; +} +} +} +} +} +block_482: +{ +lean_object* x_472; +lean_dec(x_471); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_3); +x_472 = l_Lean_Meta_Grind_internalize(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_472) == 0) +{ +lean_object* x_473; lean_object* x_474; lean_object* x_475; lean_object* x_476; lean_object* x_477; +x_473 = lean_ctor_get(x_472, 1); +lean_inc(x_473); +lean_dec(x_472); +lean_inc(x_1); +x_474 = l_Lean_Meta_Grind_registerParent(x_1, x_3, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_473); +lean_dec(x_3); +x_475 = lean_ctor_get(x_474, 0); +lean_inc(x_475); +x_476 = lean_ctor_get(x_474, 1); +lean_inc(x_476); +lean_dec(x_474); +x_477 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(x_4, x_1, x_2, x_470, x_475, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_476); +lean_dec(x_475); +lean_dec(x_4); +return x_477; +} +else +{ +uint8_t x_478; +lean_dec(x_470); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_478 = !lean_is_exclusive(x_472); +if (x_478 == 0) +{ +return x_472; +} +else +{ +lean_object* x_479; lean_object* x_480; lean_object* x_481; +x_479 = lean_ctor_get(x_472, 0); +x_480 = lean_ctor_get(x_472, 1); +lean_inc(x_480); +lean_inc(x_479); +lean_dec(x_472); +x_481 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_481, 0, x_479); +lean_ctor_set(x_481, 1, x_480); +return x_481; +} +} +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_internalize___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +uint8_t x_13; lean_object* x_14; +x_13 = 0; +x_14 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_13, x_13, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_14; +} +} +static lean_object* _init_l_Lean_Meta_Grind_internalize___lambda__2___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.Meta.Tactic.Grind.Internalize", 34, 34); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_internalize___lambda__2___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.Meta.Grind.internalize", 27, 27); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_internalize___lambda__2___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("unreachable code has been reached", 33, 33); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_internalize___lambda__2___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l_Lean_Meta_Grind_internalize___lambda__2___closed__1; +x_2 = l_Lean_Meta_Grind_internalize___lambda__2___closed__2; +x_3 = lean_unsigned_to_nat(76u); +x_4 = lean_unsigned_to_nat(16u); +x_5 = l_Lean_Meta_Grind_internalize___lambda__2___closed__3; +x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); +return x_6; +} +} +static lean_object* _init_l_Lean_Meta_Grind_internalize___lambda__2___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("unexpected term during internalization", 38, 38); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_internalize___lambda__2___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_internalize___lambda__2___closed__5; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_internalize___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +switch (lean_obj_tag(x_1)) { +case 0: +{ +lean_object* x_13; lean_object* x_14; +lean_dec(x_2); +lean_dec(x_1); +x_13 = l_Lean_Meta_Grind_internalize___lambda__2___closed__4; +x_14 = l_panic___at_Lean_Meta_Grind_internalize___spec__1(x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_14; +} +case 2: +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_15 = l_Lean_Meta_Grind_addCongrTable___closed__2; +x_16 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_15, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +x_18 = lean_unbox(x_17); +lean_dec(x_17); +if (x_18 == 0) +{ +lean_object* x_19; uint8_t x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_16, 1); +lean_inc(x_19); +lean_dec(x_16); +x_20 = 0; +x_21 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_20, x_20, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_19); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_21; +} +else +{ +uint8_t x_22; +x_22 = !lean_is_exclusive(x_16); +if (x_22 == 0) +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; lean_object* x_32; +x_23 = lean_ctor_get(x_16, 1); +x_24 = lean_ctor_get(x_16, 0); +lean_dec(x_24); +lean_inc(x_1); +x_25 = l_Lean_indentExpr(x_1); +x_26 = l_Lean_Meta_Grind_internalize___lambda__2___closed__6; +lean_ctor_set_tag(x_16, 7); +lean_ctor_set(x_16, 1, x_25); +lean_ctor_set(x_16, 0, x_26); +x_27 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__5; +x_28 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_28, 0, x_16); +lean_ctor_set(x_28, 1, x_27); +x_29 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_15, x_28, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_23); +x_30 = lean_ctor_get(x_29, 1); +lean_inc(x_30); +lean_dec(x_29); +x_31 = 0; +x_32 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_31, x_31, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_30); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +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; uint8_t x_41; lean_object* x_42; +x_33 = lean_ctor_get(x_16, 1); +lean_inc(x_33); +lean_dec(x_16); +lean_inc(x_1); +x_34 = l_Lean_indentExpr(x_1); +x_35 = l_Lean_Meta_Grind_internalize___lambda__2___closed__6; +x_36 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set(x_36, 1, x_34); +x_37 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__5; +x_38 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_38, 0, x_36); +lean_ctor_set(x_38, 1, x_37); +x_39 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_15, x_38, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_33); +x_40 = lean_ctor_get(x_39, 1); +lean_inc(x_40); +lean_dec(x_39); +x_41 = 0; +x_42 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_41, x_41, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_40); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_42; +} +} +} +case 3: +{ +lean_object* x_43; lean_object* x_44; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_43 = lean_box(0); +x_44 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_44, 0, x_43); +lean_ctor_set(x_44, 1, x_12); +return x_44; +} +case 4: +{ +lean_object* x_45; +x_45 = l_Lean_Meta_Grind_mkENode(x_1, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_45; +} +case 5: +{ +lean_object* x_46; +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_1); +x_46 = l_Lean_Meta_isLitValue(x_1, x_8, x_9, x_10, x_11, x_12); +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); +lean_dec(x_47); +if (x_48 == 0) +{ +lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_49 = lean_ctor_get(x_46, 1); +lean_inc(x_49); +lean_dec(x_46); +x_50 = lean_unsigned_to_nat(0u); +x_51 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_50); +x_52 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__1; +lean_inc(x_51); +x_53 = lean_mk_array(x_51, x_52); +x_54 = lean_unsigned_to_nat(1u); +x_55 = lean_nat_sub(x_51, x_54); +lean_dec(x_51); +lean_inc(x_1); +x_56 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3(x_1, x_2, x_1, x_53, x_55, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_49); +return x_56; +} +else +{ +lean_object* x_57; lean_object* x_58; +x_57 = lean_ctor_get(x_46, 1); +lean_inc(x_57); +lean_dec(x_46); +x_58 = l_Lean_Meta_Grind_mkENode(x_1, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_57); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_58; +} +} +else +{ +uint8_t x_59; +lean_dec(x_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_59 = !lean_is_exclusive(x_46); +if (x_59 == 0) +{ +return x_46; +} +else +{ +lean_object* x_60; lean_object* x_61; lean_object* x_62; +x_60 = lean_ctor_get(x_46, 0); +x_61 = lean_ctor_get(x_46, 1); +lean_inc(x_61); +lean_inc(x_60); +lean_dec(x_46); +x_62 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_62, 0, x_60); +lean_ctor_set(x_62, 1, x_61); +return x_62; +} +} +} +case 7: +{ +lean_object* x_63; uint8_t x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; +x_63 = lean_ctor_get(x_1, 1); +lean_inc(x_63); +x_64 = 0; +lean_inc(x_2); +lean_inc(x_1); +x_65 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_64, x_64, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +x_66 = lean_ctor_get(x_65, 1); +lean_inc(x_66); +lean_dec(x_65); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_63); +x_67 = l_Lean_Meta_isProp(x_63, x_8, x_9, x_10, x_11, 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; +lean_dec(x_63); +lean_dec(x_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_70 = !lean_is_exclusive(x_67); +if (x_70 == 0) +{ +lean_object* x_71; lean_object* x_72; +x_71 = lean_ctor_get(x_67, 0); +lean_dec(x_71); +x_72 = lean_box(0); +lean_ctor_set(x_67, 0, x_72); +return x_67; +} +else +{ +lean_object* x_73; lean_object* x_74; lean_object* x_75; +x_73 = lean_ctor_get(x_67, 1); +lean_inc(x_73); +lean_dec(x_67); +x_74 = lean_box(0); +x_75 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_75, 0, x_74); +lean_ctor_set(x_75, 1, x_73); +return x_75; +} +} +else +{ +lean_object* x_76; lean_object* x_77; +x_76 = lean_ctor_get(x_67, 1); +lean_inc(x_76); +lean_dec(x_67); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_1); +x_77 = l_Lean_Meta_isProp(x_1, x_8, x_9, x_10, x_11, x_76); +if (lean_obj_tag(x_77) == 0) +{ +lean_object* x_78; uint8_t x_79; +x_78 = lean_ctor_get(x_77, 0); +lean_inc(x_78); +x_79 = lean_unbox(x_78); +lean_dec(x_78); +if (x_79 == 0) +{ +uint8_t x_80; +lean_dec(x_63); +lean_dec(x_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_80 = !lean_is_exclusive(x_77); +if (x_80 == 0) +{ +lean_object* x_81; lean_object* x_82; +x_81 = lean_ctor_get(x_77, 0); +lean_dec(x_81); +x_82 = lean_box(0); +lean_ctor_set(x_77, 0, x_82); +return x_77; +} +else +{ +lean_object* x_83; lean_object* x_84; lean_object* x_85; +x_83 = lean_ctor_get(x_77, 1); +lean_inc(x_83); +lean_dec(x_77); +x_84 = lean_box(0); +x_85 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_85, 0, x_84); +lean_ctor_set(x_85, 1, x_83); +return x_85; +} +} +else +{ +lean_object* x_86; lean_object* x_87; +x_86 = lean_ctor_get(x_77, 1); +lean_inc(x_86); +lean_dec(x_77); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_63); +x_87 = l_Lean_Meta_Grind_internalize(x_63, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_86); +if (lean_obj_tag(x_87) == 0) +{ +lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; +x_88 = lean_ctor_get(x_87, 1); +lean_inc(x_88); +lean_dec(x_87); +lean_inc(x_1); +x_89 = l_Lean_Meta_Grind_registerParent(x_1, x_63, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_88); +lean_dec(x_63); +x_90 = lean_ctor_get(x_89, 1); +lean_inc(x_90); +lean_dec(x_89); +x_91 = l_Lean_Meta_Grind_propagateUp(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_90); +return x_91; +} +else +{ +uint8_t x_92; +lean_dec(x_63); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_92 = !lean_is_exclusive(x_87); +if (x_92 == 0) +{ +return x_87; +} +else +{ +lean_object* x_93; lean_object* x_94; lean_object* x_95; +x_93 = lean_ctor_get(x_87, 0); +x_94 = lean_ctor_get(x_87, 1); +lean_inc(x_94); +lean_inc(x_93); +lean_dec(x_87); +x_95 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_95, 0, x_93); +lean_ctor_set(x_95, 1, x_94); +return x_95; +} +} +} +} +else +{ +uint8_t x_96; +lean_dec(x_63); +lean_dec(x_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_96 = !lean_is_exclusive(x_77); +if (x_96 == 0) +{ +return x_77; +} +else +{ +lean_object* x_97; lean_object* x_98; lean_object* x_99; +x_97 = lean_ctor_get(x_77, 0); +x_98 = lean_ctor_get(x_77, 1); +lean_inc(x_98); +lean_inc(x_97); +lean_dec(x_77); +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_100; +lean_dec(x_63); +lean_dec(x_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_100 = !lean_is_exclusive(x_67); +if (x_100 == 0) +{ +return x_67; +} +else +{ +lean_object* x_101; lean_object* x_102; lean_object* x_103; +x_101 = lean_ctor_get(x_67, 0); +x_102 = lean_ctor_get(x_67, 1); +lean_inc(x_102); +lean_inc(x_101); +lean_dec(x_67); +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; +} +} +} +case 9: +{ +lean_object* x_104; +x_104 = l_Lean_Meta_Grind_mkENode(x_1, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_104; +} +case 10: +{ +lean_object* x_105; lean_object* x_106; lean_object* x_107; uint8_t x_108; +x_105 = l_Lean_Meta_Grind_addCongrTable___closed__2; +x_106 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_105, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +x_107 = lean_ctor_get(x_106, 0); +lean_inc(x_107); +x_108 = lean_unbox(x_107); +lean_dec(x_107); +if (x_108 == 0) +{ +lean_object* x_109; uint8_t x_110; lean_object* x_111; +x_109 = lean_ctor_get(x_106, 1); +lean_inc(x_109); +lean_dec(x_106); +x_110 = 0; +x_111 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_110, x_110, x_2, x_4, x_5, 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); +lean_dec(x_5); +lean_dec(x_4); +return x_111; +} +else +{ +uint8_t x_112; +x_112 = !lean_is_exclusive(x_106); +if (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; uint8_t x_121; lean_object* x_122; +x_113 = lean_ctor_get(x_106, 1); +x_114 = lean_ctor_get(x_106, 0); +lean_dec(x_114); +lean_inc(x_1); +x_115 = l_Lean_indentExpr(x_1); +x_116 = l_Lean_Meta_Grind_internalize___lambda__2___closed__6; +lean_ctor_set_tag(x_106, 7); +lean_ctor_set(x_106, 1, x_115); +lean_ctor_set(x_106, 0, x_116); +x_117 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__5; +x_118 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_118, 0, x_106); +lean_ctor_set(x_118, 1, x_117); +x_119 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_105, x_118, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_113); +x_120 = lean_ctor_get(x_119, 1); +lean_inc(x_120); +lean_dec(x_119); +x_121 = 0; +x_122 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_121, x_121, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_120); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_122; +} +else +{ +lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; uint8_t x_131; lean_object* x_132; +x_123 = lean_ctor_get(x_106, 1); +lean_inc(x_123); +lean_dec(x_106); +lean_inc(x_1); +x_124 = l_Lean_indentExpr(x_1); +x_125 = l_Lean_Meta_Grind_internalize___lambda__2___closed__6; +x_126 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_126, 0, x_125); +lean_ctor_set(x_126, 1, x_124); +x_127 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__5; +x_128 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_128, 0, x_126); +lean_ctor_set(x_128, 1, x_127); +x_129 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_105, x_128, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_123); +x_130 = lean_ctor_get(x_129, 1); +lean_inc(x_130); +lean_dec(x_129); +x_131 = 0; +x_132 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_131, x_131, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_130); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_132; +} +} +} +case 11: +{ +lean_object* x_133; lean_object* x_134; lean_object* x_135; uint8_t x_136; +x_133 = l_Lean_Meta_Grind_addCongrTable___closed__2; +x_134 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_133, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +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_137; uint8_t x_138; lean_object* x_139; +x_137 = lean_ctor_get(x_134, 1); +lean_inc(x_137); +lean_dec(x_134); +x_138 = 0; +x_139 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_138, x_138, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_137); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_139; +} +else +{ +uint8_t x_140; +x_140 = !lean_is_exclusive(x_134); +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; uint8_t x_149; lean_object* x_150; +x_141 = lean_ctor_get(x_134, 1); +x_142 = lean_ctor_get(x_134, 0); +lean_dec(x_142); +lean_inc(x_1); +x_143 = l_Lean_indentExpr(x_1); +x_144 = l_Lean_Meta_Grind_internalize___lambda__2___closed__6; +lean_ctor_set_tag(x_134, 7); +lean_ctor_set(x_134, 1, x_143); +lean_ctor_set(x_134, 0, x_144); +x_145 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__5; +x_146 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_146, 0, x_134); +lean_ctor_set(x_146, 1, x_145); +x_147 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_133, x_146, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_141); +x_148 = lean_ctor_get(x_147, 1); +lean_inc(x_148); +lean_dec(x_147); +x_149 = 0; +x_150 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_149, x_149, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_148); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_150; +} +else +{ +lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; uint8_t x_159; lean_object* x_160; +x_151 = lean_ctor_get(x_134, 1); +lean_inc(x_151); +lean_dec(x_134); +lean_inc(x_1); +x_152 = l_Lean_indentExpr(x_1); +x_153 = l_Lean_Meta_Grind_internalize___lambda__2___closed__6; +x_154 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_154, 0, x_153); +lean_ctor_set(x_154, 1, x_152); +x_155 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__5; +x_156 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_156, 0, x_154); +lean_ctor_set(x_156, 1, x_155); +x_157 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_133, x_156, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_151); +x_158 = lean_ctor_get(x_157, 1); +lean_inc(x_158); +lean_dec(x_157); +x_159 = 0; +x_160 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_159, x_159, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_158); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_160; +} +} +} +default: +{ +uint8_t x_161; lean_object* x_162; +x_161 = 0; +x_162 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_161, x_161, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_162; +} +} +} +} +static lean_object* _init_l_Lean_Meta_Grind_internalize___lambda__3___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("internalize", 11, 11); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_internalize___lambda__3___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__1; +x_2 = l_Lean_Meta_Grind_internalize___lambda__3___closed__1; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_internalize___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; lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_13 = l_Lean_Meta_Grind_internalize___lambda__3___closed__2; +x_14 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +x_16 = lean_unbox(x_15); +lean_dec(x_15); +if (x_16 == 0) +{ +lean_object* x_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_Grind_internalize___lambda__2(x_1, x_2, x_18, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_17); +return x_19; +} +else +{ +uint8_t x_20; +x_20 = !lean_is_exclusive(x_14); +if (x_20 == 0) +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_21 = lean_ctor_get(x_14, 1); +x_22 = lean_ctor_get(x_14, 0); +lean_dec(x_22); +lean_inc(x_1); +x_23 = l_Lean_MessageData_ofExpr(x_1); +x_24 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__5; +lean_ctor_set_tag(x_14, 7); +lean_ctor_set(x_14, 1, x_23); +lean_ctor_set(x_14, 0, x_24); +x_25 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_25, 0, x_14); +lean_ctor_set(x_25, 1, x_24); +x_26 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_13, x_25, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_21); +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_Grind_internalize___lambda__2(x_1, x_2, x_27, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_28); +lean_dec(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; +x_30 = lean_ctor_get(x_14, 1); +lean_inc(x_30); +lean_dec(x_14); +lean_inc(x_1); +x_31 = l_Lean_MessageData_ofExpr(x_1); +x_32 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__5; +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 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_34, 0, x_33); +lean_ctor_set(x_34, 1, x_32); +x_35 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_13, x_34, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, 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_Grind_internalize___lambda__2(x_1, x_2, x_36, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_37); +lean_dec(x_36); +return x_38; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_internalize(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, 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 = l_Lean_Meta_Grind_alreadyInternalized(x_1, 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_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_Meta_Grind_internalize___lambda__3(x_1, x_2, x_16, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_15); +return x_17; +} +else +{ +uint8_t x_18; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_18 = !lean_is_exclusive(x_12); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; +x_19 = lean_ctor_get(x_12, 0); +lean_dec(x_19); +x_20 = lean_box(0); +lean_ctor_set(x_12, 0, x_20); +return x_12; +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_12, 1); +lean_inc(x_21); +lean_dec(x_12); +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_EXPORT lean_object* l_Lean_PersistentHashMap_eraseAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___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); +x_5 = 5; +x_6 = l_Lean_PersistentHashMap_findEntryAux___at_Lean_Meta_Grind_addCongrTable___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); +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_name_eq(x_3, x_11); +lean_dec(x_11); +if (x_12 == 0) +{ +lean_dec(x_8); +lean_dec(x_4); +return x_1; +} +else +{ +uint8_t x_13; +x_13 = !lean_is_exclusive(x_1); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; +x_14 = lean_ctor_get(x_1, 0); +lean_dec(x_14); +x_15 = lean_array_set(x_4, x_8, x_9); +lean_dec(x_8); +lean_ctor_set(x_1, 0, x_15); +return x_1; +} +else +{ +lean_object* x_16; lean_object* x_17; +lean_dec(x_1); +x_16 = lean_array_set(x_4, x_8, x_9); +lean_dec(x_8); +x_17 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_17, 0, x_16); +return x_17; +} +} +} +case 1: +{ +uint8_t x_18; +x_18 = !lean_is_exclusive(x_1); +if (x_18 == 0) +{ +lean_object* x_19; uint8_t x_20; +x_19 = lean_ctor_get(x_1, 0); +lean_dec(x_19); +x_20 = !lean_is_exclusive(x_10); +if (x_20 == 0) +{ +lean_object* x_21; lean_object* x_22; size_t x_23; lean_object* x_24; lean_object* x_25; +x_21 = lean_ctor_get(x_10, 0); +x_22 = lean_array_set(x_4, x_8, x_9); +x_23 = lean_usize_shift_right(x_2, x_5); +x_24 = l_Lean_PersistentHashMap_eraseAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__2(x_21, x_23, x_3); +lean_inc(x_24); +x_25 = l_Lean_PersistentHashMap_isUnaryNode___rarg(x_24); +if (lean_obj_tag(x_25) == 0) +{ +lean_object* x_26; +lean_ctor_set(x_10, 0, x_24); +x_26 = lean_array_set(x_22, x_8, x_10); +lean_dec(x_8); +lean_ctor_set(x_1, 0, x_26); +return x_1; +} +else +{ +lean_object* x_27; uint8_t x_28; +lean_dec(x_24); +lean_free_object(x_10); +x_27 = lean_ctor_get(x_25, 0); +lean_inc(x_27); +lean_dec(x_25); +x_28 = !lean_is_exclusive(x_27); +if (x_28 == 0) +{ +lean_object* x_29; +x_29 = lean_array_set(x_22, x_8, x_27); +lean_dec(x_8); +lean_ctor_set(x_1, 0, x_29); +return x_1; +} +else +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_30 = lean_ctor_get(x_27, 0); +x_31 = lean_ctor_get(x_27, 1); +lean_inc(x_31); +lean_inc(x_30); +lean_dec(x_27); +x_32 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_32, 0, x_30); +lean_ctor_set(x_32, 1, x_31); +x_33 = lean_array_set(x_22, x_8, x_32); +lean_dec(x_8); +lean_ctor_set(x_1, 0, x_33); +return x_1; +} +} +} +else +{ +lean_object* x_34; lean_object* x_35; size_t x_36; lean_object* x_37; lean_object* x_38; +x_34 = lean_ctor_get(x_10, 0); +lean_inc(x_34); +lean_dec(x_10); +x_35 = lean_array_set(x_4, x_8, x_9); +x_36 = lean_usize_shift_right(x_2, x_5); +x_37 = l_Lean_PersistentHashMap_eraseAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__2(x_34, x_36, x_3); +lean_inc(x_37); +x_38 = l_Lean_PersistentHashMap_isUnaryNode___rarg(x_37); +if (lean_obj_tag(x_38) == 0) +{ +lean_object* x_39; lean_object* x_40; +x_39 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_39, 0, x_37); +x_40 = lean_array_set(x_35, x_8, x_39); +lean_dec(x_8); +lean_ctor_set(x_1, 0, x_40); +return x_1; +} +else +{ +lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; +lean_dec(x_37); +x_41 = lean_ctor_get(x_38, 0); +lean_inc(x_41); +lean_dec(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_43); +if (lean_is_exclusive(x_41)) { + lean_ctor_release(x_41, 0); + lean_ctor_release(x_41, 1); + x_44 = x_41; +} else { + lean_dec_ref(x_41); + x_44 = lean_box(0); +} +if (lean_is_scalar(x_44)) { + x_45 = lean_alloc_ctor(0, 2, 0); +} else { + x_45 = x_44; +} +lean_ctor_set(x_45, 0, x_42); +lean_ctor_set(x_45, 1, x_43); +x_46 = lean_array_set(x_35, x_8, x_45); +lean_dec(x_8); +lean_ctor_set(x_1, 0, x_46); +return x_1; +} +} +} +else +{ +lean_object* x_47; lean_object* x_48; lean_object* x_49; size_t x_50; lean_object* x_51; lean_object* x_52; +lean_dec(x_1); +x_47 = lean_ctor_get(x_10, 0); +lean_inc(x_47); +if (lean_is_exclusive(x_10)) { + lean_ctor_release(x_10, 0); + x_48 = x_10; +} else { + lean_dec_ref(x_10); + x_48 = lean_box(0); +} +x_49 = lean_array_set(x_4, x_8, x_9); +x_50 = lean_usize_shift_right(x_2, x_5); +x_51 = l_Lean_PersistentHashMap_eraseAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__2(x_47, x_50, x_3); +lean_inc(x_51); +x_52 = l_Lean_PersistentHashMap_isUnaryNode___rarg(x_51); +if (lean_obj_tag(x_52) == 0) +{ +lean_object* x_53; lean_object* x_54; lean_object* x_55; +if (lean_is_scalar(x_48)) { + x_53 = lean_alloc_ctor(1, 1, 0); +} else { + x_53 = x_48; +} +lean_ctor_set(x_53, 0, x_51); +x_54 = lean_array_set(x_49, x_8, x_53); +lean_dec(x_8); +x_55 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_55, 0, x_54); +return x_55; +} +else +{ +lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; +lean_dec(x_51); +lean_dec(x_48); +x_56 = lean_ctor_get(x_52, 0); +lean_inc(x_56); +lean_dec(x_52); +x_57 = lean_ctor_get(x_56, 0); +lean_inc(x_57); +x_58 = lean_ctor_get(x_56, 1); +lean_inc(x_58); +if (lean_is_exclusive(x_56)) { + lean_ctor_release(x_56, 0); + lean_ctor_release(x_56, 1); + x_59 = x_56; +} else { + lean_dec_ref(x_56); + x_59 = lean_box(0); +} +if (lean_is_scalar(x_59)) { + x_60 = lean_alloc_ctor(0, 2, 0); +} else { + x_60 = x_59; +} +lean_ctor_set(x_60, 0, x_57); +lean_ctor_set(x_60, 1, x_58); +x_61 = lean_array_set(x_49, x_8, x_60); +lean_dec(x_8); +x_62 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_62, 0, x_61); +return x_62; +} +} +} +default: +{ +lean_dec(x_8); +lean_dec(x_4); +return x_1; +} +} +} +else +{ +lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_63 = lean_ctor_get(x_1, 0); +lean_inc(x_63); +x_64 = lean_ctor_get(x_1, 1); +lean_inc(x_64); +x_65 = lean_unsigned_to_nat(0u); +x_66 = l_Array_indexOfAux___at_Lean_MetavarContext_setMVarUserName___spec__3(x_63, x_3, x_65); +if (lean_obj_tag(x_66) == 0) +{ +lean_dec(x_64); +lean_dec(x_63); +return x_1; +} +else +{ +uint8_t x_67; +x_67 = !lean_is_exclusive(x_1); +if (x_67 == 0) +{ +lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; +x_68 = lean_ctor_get(x_1, 1); +lean_dec(x_68); +x_69 = lean_ctor_get(x_1, 0); +lean_dec(x_69); +x_70 = lean_ctor_get(x_66, 0); +lean_inc(x_70); +lean_dec(x_66); +lean_inc(x_70); +x_71 = l_Array_eraseIdx___rarg(x_63, x_70, lean_box(0)); +x_72 = l_Array_eraseIdx___rarg(x_64, x_70, lean_box(0)); +lean_ctor_set(x_1, 1, x_72); +lean_ctor_set(x_1, 0, x_71); +return x_1; +} +else +{ +lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; +lean_dec(x_1); +x_73 = lean_ctor_get(x_66, 0); +lean_inc(x_73); +lean_dec(x_66); +lean_inc(x_73); +x_74 = l_Array_eraseIdx___rarg(x_63, x_73, lean_box(0)); +x_75 = l_Array_eraseIdx___rarg(x_64, x_73, lean_box(0)); +x_76 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_76, 0, x_74); +lean_ctor_set(x_76, 1, x_75); +return x_76; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_erase___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint64_t x_3; size_t x_4; lean_object* x_5; +x_3 = l_Lean_Name_hash___override(x_2); +x_4 = lean_uint64_to_usize(x_3); +x_5 = l_Lean_PersistentHashMap_eraseAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__2(x_1, x_4, x_2); +return x_5; +} +} +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAtAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; 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) +{ +uint8_t x_8; +lean_dec(x_4); +x_8 = 0; +return x_8; +} +else +{ +lean_object* x_9; uint8_t x_10; +x_9 = lean_array_fget(x_1, x_4); +x_10 = l___private_Lean_HeadIndex_0__Lean_beqHeadIndex____x40_Lean_HeadIndex___hyg_69_(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 +{ +uint8_t x_14; +lean_dec(x_4); +x_14 = 1; +return x_14; +} +} +} +} +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__4(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_findEntryAux___at_Lean_Meta_Grind_addCongrTable___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 = l___private_Lean_HeadIndex_0__Lean_beqHeadIndex____x40_Lean_HeadIndex___hyg_69_(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___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5(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___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__3(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint64_t x_3; size_t x_4; uint8_t x_5; +x_3 = l_Lean_HeadIndex_hash(x_2); +x_4 = lean_uint64_to_usize(x_3); +x_5 = l_Lean_PersistentHashMap_containsAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__4(x_1, x_4, x_2); +return x_5; +} +} +LEAN_EXPORT lean_object* l_List_filterTR_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__6(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_dec(x_1); +x_4 = l_List_reverse___rarg(x_3); +return x_4; +} +else +{ +uint8_t x_5; +x_5 = !lean_is_exclusive(x_2); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; uint8_t x_8; +x_6 = lean_ctor_get(x_2, 0); +x_7 = lean_ctor_get(x_2, 1); +lean_inc(x_1); +x_8 = l_Lean_PersistentHashMap_contains___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__3(x_1, x_6); +if (x_8 == 0) +{ +lean_ctor_set(x_2, 1, x_3); +{ +lean_object* _tmp_1 = x_7; +lean_object* _tmp_2 = x_2; +x_2 = _tmp_1; +x_3 = _tmp_2; +} +goto _start; +} +else +{ +lean_free_object(x_2); +lean_dec(x_6); +x_2 = x_7; +goto _start; +} +} +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); +lean_inc(x_1); +x_13 = l_Lean_PersistentHashMap_contains___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__3(x_1, x_11); +if (x_13 == 0) +{ +lean_object* x_14; +x_14 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_14, 0, x_11); +lean_ctor_set(x_14, 1, x_3); +x_2 = x_12; +x_3 = x_14; +goto _start; +} +else +{ +lean_dec(x_11); +x_2 = x_12; +goto _start; +} +} +} +} +} +LEAN_EXPORT lean_object* l_List_mapM_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_13; lean_object* x_14; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_13 = l_List_reverse___rarg(x_3); +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_2); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_2, 0); +x_17 = lean_ctor_get(x_2, 1); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_1); +x_18 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern(x_16, x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_18) == 0) +{ +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_2, 1, x_3); +lean_ctor_set(x_2, 0, x_19); +{ +lean_object* _tmp_1 = x_17; +lean_object* _tmp_2 = x_2; +lean_object* _tmp_11 = x_20; +x_2 = _tmp_1; +x_3 = _tmp_2; +x_12 = _tmp_11; +} +goto _start; +} +else +{ +uint8_t x_22; +lean_free_object(x_2); +lean_dec(x_17); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_22 = !lean_is_exclusive(x_18); +if (x_22 == 0) +{ +return x_18; +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_18, 0); +x_24 = lean_ctor_get(x_18, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_18); +x_25 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_24); +return x_25; +} +} +} +else +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_26 = lean_ctor_get(x_2, 0); +x_27 = lean_ctor_get(x_2, 1); +lean_inc(x_27); +lean_inc(x_26); +lean_dec(x_2); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_1); +x_28 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern(x_26, x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_28) == 0) +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_29 = lean_ctor_get(x_28, 0); +lean_inc(x_29); +x_30 = lean_ctor_get(x_28, 1); +lean_inc(x_30); +lean_dec(x_28); +x_31 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_31, 0, x_29); +lean_ctor_set(x_31, 1, x_3); +x_2 = x_27; +x_3 = x_31; +x_12 = x_30; +goto _start; +} +else +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; +lean_dec(x_27); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_33 = lean_ctor_get(x_28, 0); +lean_inc(x_33); +x_34 = lean_ctor_get(x_28, 1); +lean_inc(x_34); +if (lean_is_exclusive(x_28)) { + lean_ctor_release(x_28, 0); + lean_ctor_release(x_28, 1); + x_35 = x_28; +} else { + lean_dec_ref(x_28); + x_35 = lean_box(0); +} +if (lean_is_scalar(x_35)) { + x_36 = lean_alloc_ctor(1, 2, 0); +} else { + x_36 = x_35; +} +lean_ctor_set(x_36, 0, x_33); +lean_ctor_set(x_36, 1, x_34); +return x_36; +} +} +} +} +} +static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___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_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___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, 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; +x_12 = lean_st_ref_take(x_3, x_11); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_12, 1); +lean_inc(x_14); +lean_dec(x_12); +x_15 = !lean_is_exclusive(x_13); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; +x_16 = lean_ctor_get(x_13, 9); +x_17 = l_Lean_PersistentArray_push___rarg(x_16, x_1); +lean_ctor_set(x_13, 9, x_17); +x_18 = lean_st_ref_set(x_3, x_13, x_14); +x_19 = !lean_is_exclusive(x_18); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; +x_20 = lean_ctor_get(x_18, 0); +lean_dec(x_20); +x_21 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__1___closed__1; +lean_ctor_set(x_18, 0, x_21); +return x_18; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_22 = lean_ctor_get(x_18, 1); +lean_inc(x_22); +lean_dec(x_18); +x_23 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__1___closed__1; +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set(x_24, 1, x_22); +return x_24; +} +} +else +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; 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_25 = lean_ctor_get(x_13, 0); +x_26 = lean_ctor_get(x_13, 1); +x_27 = lean_ctor_get(x_13, 2); +x_28 = lean_ctor_get(x_13, 3); +x_29 = lean_ctor_get(x_13, 4); +x_30 = lean_ctor_get(x_13, 5); +x_31 = lean_ctor_get_uint8(x_13, sizeof(void*)*14); +x_32 = lean_ctor_get(x_13, 6); +x_33 = lean_ctor_get(x_13, 7); +x_34 = lean_ctor_get(x_13, 8); +x_35 = lean_ctor_get(x_13, 9); +x_36 = lean_ctor_get(x_13, 10); +x_37 = lean_ctor_get(x_13, 11); +x_38 = lean_ctor_get(x_13, 12); +x_39 = lean_ctor_get(x_13, 13); +lean_inc(x_39); +lean_inc(x_38); +lean_inc(x_37); +lean_inc(x_36); +lean_inc(x_35); +lean_inc(x_34); +lean_inc(x_33); +lean_inc(x_32); +lean_inc(x_30); +lean_inc(x_29); +lean_inc(x_28); +lean_inc(x_27); +lean_inc(x_26); +lean_inc(x_25); +lean_dec(x_13); +x_40 = l_Lean_PersistentArray_push___rarg(x_35, x_1); +x_41 = lean_alloc_ctor(0, 14, 1); +lean_ctor_set(x_41, 0, x_25); +lean_ctor_set(x_41, 1, x_26); +lean_ctor_set(x_41, 2, x_27); +lean_ctor_set(x_41, 3, x_28); +lean_ctor_set(x_41, 4, x_29); +lean_ctor_set(x_41, 5, x_30); +lean_ctor_set(x_41, 6, x_32); +lean_ctor_set(x_41, 7, x_33); +lean_ctor_set(x_41, 8, x_34); +lean_ctor_set(x_41, 9, x_40); +lean_ctor_set(x_41, 10, x_36); +lean_ctor_set(x_41, 11, x_37); +lean_ctor_set(x_41, 12, x_38); +lean_ctor_set(x_41, 13, x_39); +lean_ctor_set_uint8(x_41, sizeof(void*)*14, x_31); +x_42 = lean_st_ref_set(x_3, x_41, x_14); +x_43 = lean_ctor_get(x_42, 1); +lean_inc(x_43); +if (lean_is_exclusive(x_42)) { + lean_ctor_release(x_42, 0); + lean_ctor_release(x_42, 1); + x_44 = x_42; +} else { + lean_dec_ref(x_42); + x_44 = lean_box(0); +} +x_45 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__1___closed__1; +if (lean_is_scalar(x_44)) { + x_46 = lean_alloc_ctor(0, 2, 0); +} else { + x_46 = x_44; +} +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_43); +return x_46; +} +} +} +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___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; lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_12 = lean_st_ref_take(x_3, x_11); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_12, 1); +lean_inc(x_14); +lean_dec(x_12); +x_15 = !lean_is_exclusive(x_13); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; +x_16 = lean_ctor_get(x_13, 10); +x_17 = l_Lean_Meta_Grind_EMatchTheorems_insert(x_16, x_1); +lean_ctor_set(x_13, 10, x_17); +x_18 = lean_st_ref_set(x_3, x_13, x_14); +x_19 = !lean_is_exclusive(x_18); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; +x_20 = lean_ctor_get(x_18, 0); +lean_dec(x_20); +x_21 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__1___closed__1; +lean_ctor_set(x_18, 0, x_21); +return x_18; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_22 = lean_ctor_get(x_18, 1); +lean_inc(x_22); +lean_dec(x_18); +x_23 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__1___closed__1; +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set(x_24, 1, x_22); +return x_24; +} +} +else +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; 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_25 = lean_ctor_get(x_13, 0); +x_26 = lean_ctor_get(x_13, 1); +x_27 = lean_ctor_get(x_13, 2); +x_28 = lean_ctor_get(x_13, 3); +x_29 = lean_ctor_get(x_13, 4); +x_30 = lean_ctor_get(x_13, 5); +x_31 = lean_ctor_get_uint8(x_13, sizeof(void*)*14); +x_32 = lean_ctor_get(x_13, 6); +x_33 = lean_ctor_get(x_13, 7); +x_34 = lean_ctor_get(x_13, 8); +x_35 = lean_ctor_get(x_13, 9); +x_36 = lean_ctor_get(x_13, 10); +x_37 = lean_ctor_get(x_13, 11); +x_38 = lean_ctor_get(x_13, 12); +x_39 = lean_ctor_get(x_13, 13); +lean_inc(x_39); +lean_inc(x_38); +lean_inc(x_37); +lean_inc(x_36); +lean_inc(x_35); +lean_inc(x_34); +lean_inc(x_33); +lean_inc(x_32); +lean_inc(x_30); +lean_inc(x_29); +lean_inc(x_28); +lean_inc(x_27); +lean_inc(x_26); +lean_inc(x_25); +lean_dec(x_13); +x_40 = l_Lean_Meta_Grind_EMatchTheorems_insert(x_36, x_1); +x_41 = lean_alloc_ctor(0, 14, 1); +lean_ctor_set(x_41, 0, x_25); +lean_ctor_set(x_41, 1, x_26); +lean_ctor_set(x_41, 2, x_27); +lean_ctor_set(x_41, 3, x_28); +lean_ctor_set(x_41, 4, x_29); +lean_ctor_set(x_41, 5, x_30); +lean_ctor_set(x_41, 6, x_32); +lean_ctor_set(x_41, 7, x_33); +lean_ctor_set(x_41, 8, x_34); +lean_ctor_set(x_41, 9, x_35); +lean_ctor_set(x_41, 10, x_40); +lean_ctor_set(x_41, 11, x_37); +lean_ctor_set(x_41, 12, x_38); +lean_ctor_set(x_41, 13, x_39); +lean_ctor_set_uint8(x_41, sizeof(void*)*14, x_31); +x_42 = lean_st_ref_set(x_3, x_41, x_14); +x_43 = lean_ctor_get(x_42, 1); +lean_inc(x_43); +if (lean_is_exclusive(x_42)) { + lean_ctor_release(x_42, 0); + lean_ctor_release(x_42, 1); + x_44 = x_42; +} else { + lean_dec_ref(x_42); + x_44 = lean_box(0); +} +x_45 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__1___closed__1; +if (lean_is_scalar(x_44)) { + x_46 = lean_alloc_ctor(0, 2, 0); +} else { + x_46 = x_44; +} +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_43); +return x_46; +} +} +} +static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("ematch", 6, 6); +return x_1; +} +} +static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__1; +x_2 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__1; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("activated `", 11, 11); +return x_1; +} +} +static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__3; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("`, ", 3, 3); +return x_1; +} +} +static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__5; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("reinsert `", 10, 10); +return x_1; +} +} +static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__7; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__9() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("`", 1, 1); +return x_1; +} +} +static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__9; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17) { +_start: +{ +if (lean_obj_tag(x_6) == 0) +{ +lean_object* x_18; +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_3); +lean_dec(x_1); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_7); +lean_ctor_set(x_18, 1, x_17); +return x_18; +} +else +{ +uint8_t x_19; +lean_dec(x_7); +x_19 = !lean_is_exclusive(x_6); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_27; +x_20 = lean_ctor_get(x_6, 0); +x_21 = lean_ctor_get(x_6, 1); +x_27 = !lean_is_exclusive(x_20); +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; +x_28 = lean_ctor_get(x_20, 0); +x_29 = lean_ctor_get(x_20, 1); +x_30 = lean_ctor_get(x_20, 2); +x_31 = lean_ctor_get(x_20, 3); +x_32 = lean_ctor_get(x_20, 4); +x_33 = lean_ctor_get(x_20, 5); +x_34 = lean_box(0); +lean_inc(x_3); +x_35 = l_List_filterTR_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__6(x_3, x_32, x_34); +if (lean_obj_tag(x_35) == 0) +{ +lean_object* x_36; uint8_t x_37; +x_36 = l_Lean_Meta_Grind_shareCommon(x_29, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +x_37 = !lean_is_exclusive(x_36); +if (x_37 == 0) +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_38 = lean_ctor_get(x_36, 0); +x_39 = lean_ctor_get(x_36, 1); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_1); +x_40 = l_List_mapM_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__7(x_1, x_31, x_34, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_39); +if (lean_obj_tag(x_40) == 0) +{ +lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; uint8_t x_46; +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_33); +lean_inc(x_41); +lean_ctor_set(x_20, 4, x_35); +lean_ctor_set(x_20, 3, x_41); +lean_ctor_set(x_20, 1, x_38); +x_43 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__2; +x_44 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_43, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_42); +x_45 = lean_ctor_get(x_44, 0); +lean_inc(x_45); +x_46 = lean_unbox(x_45); +lean_dec(x_45); +if (x_46 == 0) +{ +lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; +lean_dec(x_41); +lean_free_object(x_36); +lean_dec(x_33); +lean_free_object(x_6); +x_47 = lean_ctor_get(x_44, 1); +lean_inc(x_47); +lean_dec(x_44); +x_48 = lean_box(0); +x_49 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__1(x_20, x_48, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_47); +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_22 = x_50; +x_23 = x_51; +goto block_26; +} +else +{ +uint8_t x_52; +x_52 = !lean_is_exclusive(x_44); +if (x_52 == 0) +{ +lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; +x_53 = lean_ctor_get(x_44, 1); +x_54 = lean_ctor_get(x_44, 0); +lean_dec(x_54); +x_55 = l_Lean_Meta_Grind_Origin_key(x_33); +lean_dec(x_33); +x_56 = l_Lean_MessageData_ofName(x_55); +x_57 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__4; +lean_ctor_set_tag(x_44, 7); +lean_ctor_set(x_44, 1, x_56); +lean_ctor_set(x_44, 0, x_57); +x_58 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__6; +lean_ctor_set_tag(x_36, 7); +lean_ctor_set(x_36, 1, x_58); +lean_ctor_set(x_36, 0, x_44); +x_59 = l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__3(x_41, x_34); +x_60 = l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__4(x_59, x_34); +x_61 = l_Lean_MessageData_ofList(x_60); +lean_ctor_set_tag(x_6, 7); +lean_ctor_set(x_6, 1, x_61); +lean_ctor_set(x_6, 0, x_36); +x_62 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__5; +x_63 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_63, 0, x_6); +lean_ctor_set(x_63, 1, x_62); +x_64 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_43, x_63, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_53); +x_65 = lean_ctor_get(x_64, 0); +lean_inc(x_65); +x_66 = lean_ctor_get(x_64, 1); +lean_inc(x_66); +lean_dec(x_64); +x_67 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__1(x_20, x_65, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_66); +lean_dec(x_65); +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_22 = x_68; +x_23 = x_69; +goto block_26; +} +else +{ +lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; +x_70 = lean_ctor_get(x_44, 1); +lean_inc(x_70); +lean_dec(x_44); +x_71 = l_Lean_Meta_Grind_Origin_key(x_33); +lean_dec(x_33); +x_72 = l_Lean_MessageData_ofName(x_71); +x_73 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__4; +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_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__6; +lean_ctor_set_tag(x_36, 7); +lean_ctor_set(x_36, 1, x_75); +lean_ctor_set(x_36, 0, x_74); +x_76 = l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__3(x_41, x_34); +x_77 = l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__4(x_76, x_34); +x_78 = l_Lean_MessageData_ofList(x_77); +lean_ctor_set_tag(x_6, 7); +lean_ctor_set(x_6, 1, x_78); +lean_ctor_set(x_6, 0, x_36); +x_79 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__5; +x_80 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_80, 0, x_6); +lean_ctor_set(x_80, 1, x_79); +x_81 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_43, x_80, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_70); +x_82 = lean_ctor_get(x_81, 0); +lean_inc(x_82); +x_83 = lean_ctor_get(x_81, 1); +lean_inc(x_83); +lean_dec(x_81); +x_84 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__1(x_20, x_82, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_83); +lean_dec(x_82); +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_26; +} +} +} +else +{ +uint8_t x_87; +lean_free_object(x_36); +lean_dec(x_38); +lean_free_object(x_20); +lean_dec(x_33); +lean_dec(x_30); +lean_dec(x_28); +lean_free_object(x_6); +lean_dec(x_21); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_3); +lean_dec(x_1); +x_87 = !lean_is_exclusive(x_40); +if (x_87 == 0) +{ +return x_40; +} +else +{ +lean_object* x_88; lean_object* x_89; lean_object* x_90; +x_88 = lean_ctor_get(x_40, 0); +x_89 = lean_ctor_get(x_40, 1); +lean_inc(x_89); +lean_inc(x_88); +lean_dec(x_40); +x_90 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_90, 0, x_88); +lean_ctor_set(x_90, 1, x_89); +return x_90; +} +} +} +else +{ +lean_object* x_91; lean_object* x_92; lean_object* x_93; +x_91 = lean_ctor_get(x_36, 0); +x_92 = lean_ctor_get(x_36, 1); +lean_inc(x_92); +lean_inc(x_91); +lean_dec(x_36); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_1); +x_93 = l_List_mapM_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__7(x_1, x_31, x_34, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_92); +if (lean_obj_tag(x_93) == 0) +{ +lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; uint8_t x_99; +x_94 = lean_ctor_get(x_93, 0); +lean_inc(x_94); +x_95 = lean_ctor_get(x_93, 1); +lean_inc(x_95); +lean_dec(x_93); +lean_inc(x_33); +lean_inc(x_94); +lean_ctor_set(x_20, 4, x_35); +lean_ctor_set(x_20, 3, x_94); +lean_ctor_set(x_20, 1, x_91); +x_96 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__2; +x_97 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_96, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_95); +x_98 = lean_ctor_get(x_97, 0); +lean_inc(x_98); +x_99 = lean_unbox(x_98); +lean_dec(x_98); +if (x_99 == 0) +{ +lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; +lean_dec(x_94); +lean_dec(x_33); +lean_free_object(x_6); +x_100 = lean_ctor_get(x_97, 1); +lean_inc(x_100); +lean_dec(x_97); +x_101 = lean_box(0); +x_102 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__1(x_20, x_101, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_100); +x_103 = lean_ctor_get(x_102, 0); +lean_inc(x_103); +x_104 = lean_ctor_get(x_102, 1); +lean_inc(x_104); +lean_dec(x_102); +x_22 = x_103; +x_23 = x_104; +goto block_26; +} +else +{ +lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; +x_105 = lean_ctor_get(x_97, 1); +lean_inc(x_105); +if (lean_is_exclusive(x_97)) { + lean_ctor_release(x_97, 0); + lean_ctor_release(x_97, 1); + x_106 = x_97; +} else { + lean_dec_ref(x_97); + x_106 = lean_box(0); +} +x_107 = l_Lean_Meta_Grind_Origin_key(x_33); +lean_dec(x_33); +x_108 = l_Lean_MessageData_ofName(x_107); +x_109 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__4; +if (lean_is_scalar(x_106)) { + x_110 = lean_alloc_ctor(7, 2, 0); +} else { + x_110 = x_106; + lean_ctor_set_tag(x_110, 7); +} +lean_ctor_set(x_110, 0, x_109); +lean_ctor_set(x_110, 1, x_108); +x_111 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__6; +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_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__3(x_94, x_34); +x_114 = l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__4(x_113, x_34); +x_115 = l_Lean_MessageData_ofList(x_114); +lean_ctor_set_tag(x_6, 7); +lean_ctor_set(x_6, 1, x_115); +lean_ctor_set(x_6, 0, x_112); +x_116 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__5; +x_117 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_117, 0, x_6); +lean_ctor_set(x_117, 1, x_116); +x_118 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_96, x_117, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_105); +x_119 = lean_ctor_get(x_118, 0); +lean_inc(x_119); +x_120 = lean_ctor_get(x_118, 1); +lean_inc(x_120); +lean_dec(x_118); +x_121 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__1(x_20, x_119, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_120); +lean_dec(x_119); +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_22 = x_122; +x_23 = x_123; +goto block_26; +} +} +else +{ +lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; +lean_dec(x_91); +lean_free_object(x_20); +lean_dec(x_33); +lean_dec(x_30); +lean_dec(x_28); +lean_free_object(x_6); +lean_dec(x_21); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_3); +lean_dec(x_1); +x_124 = lean_ctor_get(x_93, 0); +lean_inc(x_124); +x_125 = lean_ctor_get(x_93, 1); +lean_inc(x_125); +if (lean_is_exclusive(x_93)) { + lean_ctor_release(x_93, 0); + lean_ctor_release(x_93, 1); + x_126 = x_93; +} else { + lean_dec_ref(x_93); + x_126 = lean_box(0); +} +if (lean_is_scalar(x_126)) { + x_127 = lean_alloc_ctor(1, 2, 0); +} else { + x_127 = 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_128; +lean_free_object(x_6); +lean_inc(x_33); +lean_inc(x_35); +lean_ctor_set(x_20, 4, x_35); +x_128 = !lean_is_exclusive(x_35); +if (x_128 == 0) +{ +lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; uint8_t x_134; +x_129 = lean_ctor_get(x_35, 1); +lean_dec(x_129); +x_130 = lean_ctor_get(x_35, 0); +lean_dec(x_130); +x_131 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__2; +x_132 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_131, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +x_133 = lean_ctor_get(x_132, 0); +lean_inc(x_133); +x_134 = lean_unbox(x_133); +lean_dec(x_133); +if (x_134 == 0) +{ +lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; +lean_free_object(x_35); +lean_dec(x_33); +x_135 = lean_ctor_get(x_132, 1); +lean_inc(x_135); +lean_dec(x_132); +x_136 = lean_box(0); +x_137 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__2(x_20, x_136, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_135); +x_138 = lean_ctor_get(x_137, 0); +lean_inc(x_138); +x_139 = lean_ctor_get(x_137, 1); +lean_inc(x_139); +lean_dec(x_137); +x_22 = x_138; +x_23 = x_139; +goto block_26; +} +else +{ +uint8_t x_140; +x_140 = !lean_is_exclusive(x_132); +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; +x_141 = lean_ctor_get(x_132, 1); +x_142 = lean_ctor_get(x_132, 0); +lean_dec(x_142); +x_143 = l_Lean_Meta_Grind_Origin_key(x_33); +lean_dec(x_33); +x_144 = l_Lean_MessageData_ofName(x_143); +x_145 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__8; +lean_ctor_set_tag(x_132, 7); +lean_ctor_set(x_132, 1, x_144); +lean_ctor_set(x_132, 0, x_145); +x_146 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__10; +lean_ctor_set_tag(x_35, 7); +lean_ctor_set(x_35, 1, x_146); +lean_ctor_set(x_35, 0, x_132); +x_147 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_131, x_35, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_141); +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_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__2(x_20, x_148, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_149); +lean_dec(x_148); +x_151 = lean_ctor_get(x_150, 0); +lean_inc(x_151); +x_152 = lean_ctor_get(x_150, 1); +lean_inc(x_152); +lean_dec(x_150); +x_22 = x_151; +x_23 = x_152; +goto block_26; +} +else +{ +lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; +x_153 = lean_ctor_get(x_132, 1); +lean_inc(x_153); +lean_dec(x_132); +x_154 = l_Lean_Meta_Grind_Origin_key(x_33); +lean_dec(x_33); +x_155 = l_Lean_MessageData_ofName(x_154); +x_156 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__8; +x_157 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_157, 0, x_156); +lean_ctor_set(x_157, 1, x_155); +x_158 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__10; +lean_ctor_set_tag(x_35, 7); +lean_ctor_set(x_35, 1, x_158); +lean_ctor_set(x_35, 0, x_157); +x_159 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_131, x_35, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_153); +x_160 = lean_ctor_get(x_159, 0); +lean_inc(x_160); +x_161 = lean_ctor_get(x_159, 1); +lean_inc(x_161); +lean_dec(x_159); +x_162 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__2(x_20, x_160, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_161); +lean_dec(x_160); +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_22 = x_163; +x_23 = x_164; +goto block_26; +} +} +} +else +{ +lean_object* x_165; lean_object* x_166; lean_object* x_167; uint8_t x_168; +lean_dec(x_35); +x_165 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__2; +x_166 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_165, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +x_167 = lean_ctor_get(x_166, 0); +lean_inc(x_167); +x_168 = lean_unbox(x_167); +lean_dec(x_167); +if (x_168 == 0) +{ +lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; +lean_dec(x_33); +x_169 = lean_ctor_get(x_166, 1); +lean_inc(x_169); +lean_dec(x_166); +x_170 = lean_box(0); +x_171 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__2(x_20, x_170, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_169); +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_22 = x_172; +x_23 = x_173; +goto block_26; +} +else +{ +lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; +x_174 = lean_ctor_get(x_166, 1); +lean_inc(x_174); +if (lean_is_exclusive(x_166)) { + lean_ctor_release(x_166, 0); + lean_ctor_release(x_166, 1); + x_175 = x_166; +} else { + lean_dec_ref(x_166); + x_175 = lean_box(0); +} +x_176 = l_Lean_Meta_Grind_Origin_key(x_33); +lean_dec(x_33); +x_177 = l_Lean_MessageData_ofName(x_176); +x_178 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__8; +if (lean_is_scalar(x_175)) { + x_179 = lean_alloc_ctor(7, 2, 0); +} else { + x_179 = x_175; + lean_ctor_set_tag(x_179, 7); +} +lean_ctor_set(x_179, 0, x_178); +lean_ctor_set(x_179, 1, x_177); +x_180 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__10; +x_181 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_181, 0, x_179); +lean_ctor_set(x_181, 1, x_180); +x_182 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_165, x_181, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, 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_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__2(x_20, x_183, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_184); +lean_dec(x_183); +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_22 = x_186; +x_23 = x_187; +goto block_26; +} +} +} +} +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; +x_188 = lean_ctor_get(x_20, 0); +x_189 = lean_ctor_get(x_20, 1); +x_190 = lean_ctor_get(x_20, 2); +x_191 = lean_ctor_get(x_20, 3); +x_192 = lean_ctor_get(x_20, 4); +x_193 = lean_ctor_get(x_20, 5); +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_dec(x_20); +x_194 = lean_box(0); +lean_inc(x_3); +x_195 = l_List_filterTR_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__6(x_3, x_192, x_194); +if (lean_obj_tag(x_195) == 0) +{ +lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; +x_196 = l_Lean_Meta_Grind_shareCommon(x_189, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +x_197 = lean_ctor_get(x_196, 0); +lean_inc(x_197); +x_198 = lean_ctor_get(x_196, 1); +lean_inc(x_198); +if (lean_is_exclusive(x_196)) { + lean_ctor_release(x_196, 0); + lean_ctor_release(x_196, 1); + x_199 = x_196; +} else { + lean_dec_ref(x_196); + x_199 = lean_box(0); +} +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_1); +x_200 = l_List_mapM_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__7(x_1, x_191, x_194, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_198); +if (lean_obj_tag(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; uint8_t x_207; +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); +lean_inc(x_193); +lean_inc(x_201); +x_203 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_203, 0, x_188); +lean_ctor_set(x_203, 1, x_197); +lean_ctor_set(x_203, 2, x_190); +lean_ctor_set(x_203, 3, x_201); +lean_ctor_set(x_203, 4, x_195); +lean_ctor_set(x_203, 5, x_193); +x_204 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__2; +x_205 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_204, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_202); +x_206 = lean_ctor_get(x_205, 0); +lean_inc(x_206); +x_207 = lean_unbox(x_206); +lean_dec(x_206); +if (x_207 == 0) +{ +lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; +lean_dec(x_201); +lean_dec(x_199); +lean_dec(x_193); +lean_free_object(x_6); +x_208 = lean_ctor_get(x_205, 1); +lean_inc(x_208); +lean_dec(x_205); +x_209 = lean_box(0); +x_210 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__1(x_203, x_209, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, 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_22 = x_211; +x_23 = x_212; +goto block_26; +} +else +{ +lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; +x_213 = lean_ctor_get(x_205, 1); +lean_inc(x_213); +if (lean_is_exclusive(x_205)) { + lean_ctor_release(x_205, 0); + lean_ctor_release(x_205, 1); + x_214 = x_205; +} else { + lean_dec_ref(x_205); + x_214 = lean_box(0); +} +x_215 = l_Lean_Meta_Grind_Origin_key(x_193); +lean_dec(x_193); +x_216 = l_Lean_MessageData_ofName(x_215); +x_217 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__4; +if (lean_is_scalar(x_214)) { + x_218 = lean_alloc_ctor(7, 2, 0); +} else { + x_218 = x_214; + lean_ctor_set_tag(x_218, 7); +} +lean_ctor_set(x_218, 0, x_217); +lean_ctor_set(x_218, 1, x_216); +x_219 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__6; +if (lean_is_scalar(x_199)) { + x_220 = lean_alloc_ctor(7, 2, 0); +} else { + x_220 = x_199; + lean_ctor_set_tag(x_220, 7); +} +lean_ctor_set(x_220, 0, x_218); +lean_ctor_set(x_220, 1, x_219); +x_221 = l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__3(x_201, x_194); +x_222 = l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__4(x_221, x_194); +x_223 = l_Lean_MessageData_ofList(x_222); +lean_ctor_set_tag(x_6, 7); +lean_ctor_set(x_6, 1, x_223); +lean_ctor_set(x_6, 0, x_220); +x_224 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__5; +x_225 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_225, 0, x_6); +lean_ctor_set(x_225, 1, x_224); +x_226 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_204, x_225, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_213); +x_227 = lean_ctor_get(x_226, 0); +lean_inc(x_227); +x_228 = lean_ctor_get(x_226, 1); +lean_inc(x_228); +lean_dec(x_226); +x_229 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__1(x_203, x_227, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_228); +lean_dec(x_227); +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_22 = x_230; +x_23 = x_231; +goto block_26; +} +} +else +{ +lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; +lean_dec(x_199); +lean_dec(x_197); +lean_dec(x_193); +lean_dec(x_190); +lean_dec(x_188); +lean_free_object(x_6); +lean_dec(x_21); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_3); +lean_dec(x_1); +x_232 = lean_ctor_get(x_200, 0); +lean_inc(x_232); +x_233 = lean_ctor_get(x_200, 1); +lean_inc(x_233); +if (lean_is_exclusive(x_200)) { + lean_ctor_release(x_200, 0); + lean_ctor_release(x_200, 1); + x_234 = x_200; +} else { + lean_dec_ref(x_200); + 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; +} +} +else +{ +lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; uint8_t x_241; +lean_free_object(x_6); +lean_inc(x_193); +lean_inc(x_195); +x_236 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_236, 0, x_188); +lean_ctor_set(x_236, 1, x_189); +lean_ctor_set(x_236, 2, x_190); +lean_ctor_set(x_236, 3, x_191); +lean_ctor_set(x_236, 4, x_195); +lean_ctor_set(x_236, 5, x_193); +if (lean_is_exclusive(x_195)) { + lean_ctor_release(x_195, 0); + lean_ctor_release(x_195, 1); + x_237 = x_195; +} else { + lean_dec_ref(x_195); + x_237 = lean_box(0); +} +x_238 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__2; +x_239 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_238, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +x_240 = lean_ctor_get(x_239, 0); +lean_inc(x_240); +x_241 = lean_unbox(x_240); +lean_dec(x_240); +if (x_241 == 0) +{ +lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; +lean_dec(x_237); +lean_dec(x_193); +x_242 = lean_ctor_get(x_239, 1); +lean_inc(x_242); +lean_dec(x_239); +x_243 = lean_box(0); +x_244 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__2(x_236, x_243, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_242); +x_245 = lean_ctor_get(x_244, 0); +lean_inc(x_245); +x_246 = lean_ctor_get(x_244, 1); +lean_inc(x_246); +lean_dec(x_244); +x_22 = x_245; +x_23 = x_246; +goto block_26; +} +else +{ +lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; +x_247 = lean_ctor_get(x_239, 1); +lean_inc(x_247); +if (lean_is_exclusive(x_239)) { + lean_ctor_release(x_239, 0); + lean_ctor_release(x_239, 1); + x_248 = x_239; +} else { + lean_dec_ref(x_239); + x_248 = lean_box(0); +} +x_249 = l_Lean_Meta_Grind_Origin_key(x_193); +lean_dec(x_193); +x_250 = l_Lean_MessageData_ofName(x_249); +x_251 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__8; +if (lean_is_scalar(x_248)) { + x_252 = lean_alloc_ctor(7, 2, 0); +} else { + x_252 = x_248; + lean_ctor_set_tag(x_252, 7); +} +lean_ctor_set(x_252, 0, x_251); +lean_ctor_set(x_252, 1, x_250); +x_253 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__10; +if (lean_is_scalar(x_237)) { + x_254 = lean_alloc_ctor(7, 2, 0); +} else { + x_254 = x_237; + lean_ctor_set_tag(x_254, 7); +} +lean_ctor_set(x_254, 0, x_252); +lean_ctor_set(x_254, 1, x_253); +x_255 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_238, x_254, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_247); +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 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__2(x_236, x_256, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_257); +lean_dec(x_256); +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_22 = x_259; +x_23 = x_260; +goto block_26; +} +} +} +block_26: +{ +lean_object* x_24; +x_24 = lean_ctor_get(x_22, 0); +lean_inc(x_24); +lean_dec(x_22); +x_6 = x_21; +x_7 = x_24; +x_8 = lean_box(0); +x_17 = x_23; +goto _start; +} +} +else +{ +lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; +x_261 = lean_ctor_get(x_6, 0); +x_262 = lean_ctor_get(x_6, 1); +lean_inc(x_262); +lean_inc(x_261); +lean_dec(x_6); +x_268 = lean_ctor_get(x_261, 0); +lean_inc(x_268); +x_269 = lean_ctor_get(x_261, 1); +lean_inc(x_269); +x_270 = lean_ctor_get(x_261, 2); +lean_inc(x_270); +x_271 = lean_ctor_get(x_261, 3); +lean_inc(x_271); +x_272 = lean_ctor_get(x_261, 4); +lean_inc(x_272); +x_273 = lean_ctor_get(x_261, 5); +lean_inc(x_273); +if (lean_is_exclusive(x_261)) { + lean_ctor_release(x_261, 0); + lean_ctor_release(x_261, 1); + lean_ctor_release(x_261, 2); + lean_ctor_release(x_261, 3); + lean_ctor_release(x_261, 4); + lean_ctor_release(x_261, 5); + x_274 = x_261; +} else { + lean_dec_ref(x_261); + x_274 = lean_box(0); +} +x_275 = lean_box(0); +lean_inc(x_3); +x_276 = l_List_filterTR_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__6(x_3, x_272, x_275); +if (lean_obj_tag(x_276) == 0) +{ +lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; +x_277 = l_Lean_Meta_Grind_shareCommon(x_269, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +x_278 = lean_ctor_get(x_277, 0); +lean_inc(x_278); +x_279 = lean_ctor_get(x_277, 1); +lean_inc(x_279); +if (lean_is_exclusive(x_277)) { + lean_ctor_release(x_277, 0); + lean_ctor_release(x_277, 1); + x_280 = x_277; +} else { + lean_dec_ref(x_277); + x_280 = lean_box(0); +} +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_1); +x_281 = l_List_mapM_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__7(x_1, x_271, x_275, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_279); +if (lean_obj_tag(x_281) == 0) +{ +lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; uint8_t x_288; +x_282 = lean_ctor_get(x_281, 0); +lean_inc(x_282); +x_283 = lean_ctor_get(x_281, 1); +lean_inc(x_283); +lean_dec(x_281); +lean_inc(x_273); +lean_inc(x_282); +if (lean_is_scalar(x_274)) { + x_284 = lean_alloc_ctor(0, 6, 0); +} else { + x_284 = x_274; +} +lean_ctor_set(x_284, 0, x_268); +lean_ctor_set(x_284, 1, x_278); +lean_ctor_set(x_284, 2, x_270); +lean_ctor_set(x_284, 3, x_282); +lean_ctor_set(x_284, 4, x_276); +lean_ctor_set(x_284, 5, x_273); +x_285 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__2; +x_286 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_285, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_283); +x_287 = lean_ctor_get(x_286, 0); +lean_inc(x_287); +x_288 = lean_unbox(x_287); +lean_dec(x_287); +if (x_288 == 0) +{ +lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; +lean_dec(x_282); +lean_dec(x_280); +lean_dec(x_273); +x_289 = lean_ctor_get(x_286, 1); +lean_inc(x_289); +lean_dec(x_286); +x_290 = lean_box(0); +x_291 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__1(x_284, x_290, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_289); +x_292 = lean_ctor_get(x_291, 0); +lean_inc(x_292); +x_293 = lean_ctor_get(x_291, 1); +lean_inc(x_293); +lean_dec(x_291); +x_263 = x_292; +x_264 = x_293; +goto block_267; +} +else +{ +lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; +x_294 = lean_ctor_get(x_286, 1); +lean_inc(x_294); +if (lean_is_exclusive(x_286)) { + lean_ctor_release(x_286, 0); + lean_ctor_release(x_286, 1); + x_295 = x_286; +} else { + lean_dec_ref(x_286); + x_295 = lean_box(0); +} +x_296 = l_Lean_Meta_Grind_Origin_key(x_273); +lean_dec(x_273); +x_297 = l_Lean_MessageData_ofName(x_296); +x_298 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__4; +if (lean_is_scalar(x_295)) { + x_299 = lean_alloc_ctor(7, 2, 0); +} else { + x_299 = x_295; + lean_ctor_set_tag(x_299, 7); +} +lean_ctor_set(x_299, 0, x_298); +lean_ctor_set(x_299, 1, x_297); +x_300 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__6; +if (lean_is_scalar(x_280)) { + x_301 = lean_alloc_ctor(7, 2, 0); +} else { + x_301 = x_280; + lean_ctor_set_tag(x_301, 7); +} +lean_ctor_set(x_301, 0, x_299); +lean_ctor_set(x_301, 1, x_300); +x_302 = l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__3(x_282, x_275); +x_303 = l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__4(x_302, x_275); +x_304 = l_Lean_MessageData_ofList(x_303); +x_305 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_305, 0, x_301); +lean_ctor_set(x_305, 1, x_304); +x_306 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__5; +x_307 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_307, 0, x_305); +lean_ctor_set(x_307, 1, x_306); +x_308 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_285, x_307, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_294); +x_309 = lean_ctor_get(x_308, 0); +lean_inc(x_309); +x_310 = lean_ctor_get(x_308, 1); +lean_inc(x_310); +lean_dec(x_308); +x_311 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__1(x_284, x_309, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_310); +lean_dec(x_309); +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_263 = x_312; +x_264 = x_313; +goto block_267; +} +} +else +{ +lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; +lean_dec(x_280); +lean_dec(x_278); +lean_dec(x_274); +lean_dec(x_273); +lean_dec(x_270); +lean_dec(x_268); +lean_dec(x_262); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_3); +lean_dec(x_1); +x_314 = lean_ctor_get(x_281, 0); +lean_inc(x_314); +x_315 = lean_ctor_get(x_281, 1); +lean_inc(x_315); +if (lean_is_exclusive(x_281)) { + lean_ctor_release(x_281, 0); + lean_ctor_release(x_281, 1); + x_316 = x_281; +} else { + lean_dec_ref(x_281); + x_316 = lean_box(0); +} +if (lean_is_scalar(x_316)) { + x_317 = lean_alloc_ctor(1, 2, 0); +} else { + x_317 = x_316; +} +lean_ctor_set(x_317, 0, x_314); +lean_ctor_set(x_317, 1, x_315); +return x_317; +} +} +else +{ +lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; uint8_t x_323; +lean_inc(x_273); +lean_inc(x_276); +if (lean_is_scalar(x_274)) { + x_318 = lean_alloc_ctor(0, 6, 0); +} else { + x_318 = x_274; +} +lean_ctor_set(x_318, 0, x_268); +lean_ctor_set(x_318, 1, x_269); +lean_ctor_set(x_318, 2, x_270); +lean_ctor_set(x_318, 3, x_271); +lean_ctor_set(x_318, 4, x_276); +lean_ctor_set(x_318, 5, x_273); +if (lean_is_exclusive(x_276)) { + lean_ctor_release(x_276, 0); + lean_ctor_release(x_276, 1); + x_319 = x_276; +} else { + lean_dec_ref(x_276); + x_319 = lean_box(0); +} +x_320 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__2; +x_321 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_320, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +x_322 = lean_ctor_get(x_321, 0); +lean_inc(x_322); +x_323 = lean_unbox(x_322); +lean_dec(x_322); +if (x_323 == 0) +{ +lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; +lean_dec(x_319); +lean_dec(x_273); +x_324 = lean_ctor_get(x_321, 1); +lean_inc(x_324); +lean_dec(x_321); +x_325 = lean_box(0); +x_326 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__2(x_318, x_325, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_324); +x_327 = lean_ctor_get(x_326, 0); +lean_inc(x_327); +x_328 = lean_ctor_get(x_326, 1); +lean_inc(x_328); +lean_dec(x_326); +x_263 = x_327; +x_264 = x_328; +goto block_267; +} +else +{ +lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; +x_329 = lean_ctor_get(x_321, 1); +lean_inc(x_329); +if (lean_is_exclusive(x_321)) { + lean_ctor_release(x_321, 0); + lean_ctor_release(x_321, 1); + x_330 = x_321; +} else { + lean_dec_ref(x_321); + x_330 = lean_box(0); +} +x_331 = l_Lean_Meta_Grind_Origin_key(x_273); +lean_dec(x_273); +x_332 = l_Lean_MessageData_ofName(x_331); +x_333 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__8; +if (lean_is_scalar(x_330)) { + x_334 = lean_alloc_ctor(7, 2, 0); +} else { + x_334 = x_330; + lean_ctor_set_tag(x_334, 7); +} +lean_ctor_set(x_334, 0, x_333); +lean_ctor_set(x_334, 1, x_332); +x_335 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__10; +if (lean_is_scalar(x_319)) { + x_336 = lean_alloc_ctor(7, 2, 0); +} else { + x_336 = x_319; + lean_ctor_set_tag(x_336, 7); +} +lean_ctor_set(x_336, 0, x_334); +lean_ctor_set(x_336, 1, x_335); +x_337 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_320, x_336, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_329); +x_338 = lean_ctor_get(x_337, 0); +lean_inc(x_338); +x_339 = lean_ctor_get(x_337, 1); +lean_inc(x_339); +lean_dec(x_337); +x_340 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__2(x_318, x_338, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_339); +lean_dec(x_338); +x_341 = lean_ctor_get(x_340, 0); +lean_inc(x_341); +x_342 = lean_ctor_get(x_340, 1); +lean_inc(x_342); +lean_dec(x_340); +x_263 = x_341; +x_264 = x_342; +goto block_267; +} +} +block_267: +{ +lean_object* x_265; +x_265 = lean_ctor_get(x_263, 0); +lean_inc(x_265); +lean_dec(x_263); +x_6 = x_262; +x_7 = x_265; +x_8 = lean_box(0); +x_17 = x_264; +goto _start; +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, 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_st_ref_get(x_3, x_11); +x_13 = !lean_is_exclusive(x_12); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_14 = lean_ctor_get(x_12, 0); +x_15 = lean_ctor_get(x_12, 1); +x_16 = lean_ctor_get(x_14, 10); +lean_inc(x_16); +lean_dec(x_14); +x_17 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__2(x_16, x_1); +if (lean_obj_tag(x_17) == 0) +{ +lean_object* x_18; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_18 = lean_box(0); +lean_ctor_set(x_12, 0, x_18); +return x_12; +} +else +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; +lean_free_object(x_12); +x_19 = lean_ctor_get(x_17, 0); +lean_inc(x_19); +lean_dec(x_17); +x_20 = lean_st_ref_take(x_3, x_15); +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; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_24 = lean_ctor_get(x_21, 10); +x_25 = l_Lean_PersistentHashMap_erase___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__1(x_24, x_1); +lean_ctor_set(x_21, 10, x_25); +x_26 = lean_st_ref_set(x_3, x_21, x_22); +x_27 = lean_ctor_get(x_26, 1); +lean_inc(x_27); +lean_dec(x_26); +x_28 = lean_st_ref_get(x_3, 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_ctor_get(x_29, 4); +lean_inc(x_31); +lean_dec(x_29); +x_32 = lean_box(0); +x_33 = lean_box(0); +lean_inc(x_19); +x_34 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8(x_2, x_19, x_31, x_32, x_19, x_19, x_33, lean_box(0), x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_30); +lean_dec(x_19); +if (lean_obj_tag(x_34) == 0) +{ +uint8_t x_35; +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_33); +return x_34; +} +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_33); +lean_ctor_set(x_38, 1, x_37); +return x_38; +} +} +else +{ +uint8_t x_39; +x_39 = !lean_is_exclusive(x_34); +if (x_39 == 0) +{ +return x_34; +} +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_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; +} +} +} +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; 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; lean_object* x_62; 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_43 = lean_ctor_get(x_21, 0); +x_44 = lean_ctor_get(x_21, 1); +x_45 = lean_ctor_get(x_21, 2); +x_46 = lean_ctor_get(x_21, 3); +x_47 = lean_ctor_get(x_21, 4); +x_48 = lean_ctor_get(x_21, 5); +x_49 = lean_ctor_get_uint8(x_21, sizeof(void*)*14); +x_50 = lean_ctor_get(x_21, 6); +x_51 = lean_ctor_get(x_21, 7); +x_52 = lean_ctor_get(x_21, 8); +x_53 = lean_ctor_get(x_21, 9); +x_54 = lean_ctor_get(x_21, 10); +x_55 = lean_ctor_get(x_21, 11); +x_56 = lean_ctor_get(x_21, 12); +x_57 = lean_ctor_get(x_21, 13); +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_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_21); +x_58 = l_Lean_PersistentHashMap_erase___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__1(x_54, x_1); +x_59 = lean_alloc_ctor(0, 14, 1); +lean_ctor_set(x_59, 0, x_43); +lean_ctor_set(x_59, 1, x_44); +lean_ctor_set(x_59, 2, x_45); +lean_ctor_set(x_59, 3, x_46); +lean_ctor_set(x_59, 4, x_47); +lean_ctor_set(x_59, 5, x_48); +lean_ctor_set(x_59, 6, x_50); +lean_ctor_set(x_59, 7, x_51); +lean_ctor_set(x_59, 8, x_52); +lean_ctor_set(x_59, 9, x_53); +lean_ctor_set(x_59, 10, x_58); +lean_ctor_set(x_59, 11, x_55); +lean_ctor_set(x_59, 12, x_56); +lean_ctor_set(x_59, 13, x_57); +lean_ctor_set_uint8(x_59, sizeof(void*)*14, x_49); +x_60 = lean_st_ref_set(x_3, x_59, x_22); +x_61 = lean_ctor_get(x_60, 1); +lean_inc(x_61); +lean_dec(x_60); +x_62 = lean_st_ref_get(x_3, x_61); +x_63 = lean_ctor_get(x_62, 0); +lean_inc(x_63); +x_64 = lean_ctor_get(x_62, 1); +lean_inc(x_64); +lean_dec(x_62); +x_65 = lean_ctor_get(x_63, 4); +lean_inc(x_65); +lean_dec(x_63); +x_66 = lean_box(0); +x_67 = lean_box(0); +lean_inc(x_19); +x_68 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8(x_2, x_19, x_65, x_66, x_19, x_19, x_67, lean_box(0), x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_64); +lean_dec(x_19); +if (lean_obj_tag(x_68) == 0) +{ +lean_object* x_69; lean_object* x_70; lean_object* x_71; +x_69 = lean_ctor_get(x_68, 1); +lean_inc(x_69); +if (lean_is_exclusive(x_68)) { + lean_ctor_release(x_68, 0); + lean_ctor_release(x_68, 1); + x_70 = x_68; +} else { + lean_dec_ref(x_68); + x_70 = lean_box(0); +} +if (lean_is_scalar(x_70)) { + x_71 = lean_alloc_ctor(0, 2, 0); +} else { + x_71 = x_70; +} +lean_ctor_set(x_71, 0, x_67); +lean_ctor_set(x_71, 1, x_69); +return x_71; +} +else +{ +lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; +x_72 = lean_ctor_get(x_68, 0); +lean_inc(x_72); +x_73 = lean_ctor_get(x_68, 1); +lean_inc(x_73); +if (lean_is_exclusive(x_68)) { + lean_ctor_release(x_68, 0); + lean_ctor_release(x_68, 1); + x_74 = x_68; +} else { + lean_dec_ref(x_68); + 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; +} } } } -case 11: +else { -lean_object* x_92; lean_object* x_93; lean_object* x_94; uint8_t x_95; -x_92 = l_Lean_Meta_Grind_addCongrTable___closed__2; -x_93 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_92, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -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_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; +x_76 = lean_ctor_get(x_12, 0); +x_77 = lean_ctor_get(x_12, 1); +lean_inc(x_77); +lean_inc(x_76); +lean_dec(x_12); +x_78 = lean_ctor_get(x_76, 10); +lean_inc(x_78); +lean_dec(x_76); +x_79 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__2(x_78, x_1); +if (lean_obj_tag(x_79) == 0) { -lean_object* x_96; uint8_t x_97; lean_object* x_98; -x_96 = lean_ctor_get(x_93, 1); -lean_inc(x_96); -lean_dec(x_93); -x_97 = 0; -x_98 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_97, x_97, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_96); -lean_dec(x_11); +lean_object* x_80; lean_object* x_81; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -7512,151 +10973,174 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -return x_98; +lean_dec(x_3); +lean_dec(x_2); +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 { -uint8_t x_99; -x_99 = !lean_is_exclusive(x_93); -if (x_99 == 0) -{ -lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; uint8_t x_108; lean_object* x_109; -x_100 = lean_ctor_get(x_93, 1); -x_101 = lean_ctor_get(x_93, 0); -lean_dec(x_101); -lean_inc(x_1); -x_102 = l_Lean_indentExpr(x_1); -x_103 = l_Lean_Meta_Grind_internalize___lambda__2___closed__6; -lean_ctor_set_tag(x_93, 7); -lean_ctor_set(x_93, 1, x_102); -lean_ctor_set(x_93, 0, x_103); -x_104 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__4; -x_105 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_105, 0, x_93); -lean_ctor_set(x_105, 1, x_104); -x_106 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_92, x_105, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_100); -x_107 = lean_ctor_get(x_106, 1); +lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; uint8_t x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; 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_82 = lean_ctor_get(x_79, 0); +lean_inc(x_82); +lean_dec(x_79); +x_83 = lean_st_ref_take(x_3, x_77); +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_ctor_get(x_84, 0); +lean_inc(x_86); +x_87 = lean_ctor_get(x_84, 1); +lean_inc(x_87); +x_88 = lean_ctor_get(x_84, 2); +lean_inc(x_88); +x_89 = lean_ctor_get(x_84, 3); +lean_inc(x_89); +x_90 = lean_ctor_get(x_84, 4); +lean_inc(x_90); +x_91 = lean_ctor_get(x_84, 5); +lean_inc(x_91); +x_92 = lean_ctor_get_uint8(x_84, sizeof(void*)*14); +x_93 = lean_ctor_get(x_84, 6); +lean_inc(x_93); +x_94 = lean_ctor_get(x_84, 7); +lean_inc(x_94); +x_95 = lean_ctor_get(x_84, 8); +lean_inc(x_95); +x_96 = lean_ctor_get(x_84, 9); +lean_inc(x_96); +x_97 = lean_ctor_get(x_84, 10); +lean_inc(x_97); +x_98 = lean_ctor_get(x_84, 11); +lean_inc(x_98); +x_99 = lean_ctor_get(x_84, 12); +lean_inc(x_99); +x_100 = lean_ctor_get(x_84, 13); +lean_inc(x_100); +if (lean_is_exclusive(x_84)) { + lean_ctor_release(x_84, 0); + lean_ctor_release(x_84, 1); + lean_ctor_release(x_84, 2); + lean_ctor_release(x_84, 3); + lean_ctor_release(x_84, 4); + lean_ctor_release(x_84, 5); + lean_ctor_release(x_84, 6); + lean_ctor_release(x_84, 7); + lean_ctor_release(x_84, 8); + lean_ctor_release(x_84, 9); + lean_ctor_release(x_84, 10); + lean_ctor_release(x_84, 11); + lean_ctor_release(x_84, 12); + lean_ctor_release(x_84, 13); + x_101 = x_84; +} else { + lean_dec_ref(x_84); + x_101 = lean_box(0); +} +x_102 = l_Lean_PersistentHashMap_erase___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__1(x_97, x_1); +if (lean_is_scalar(x_101)) { + x_103 = lean_alloc_ctor(0, 14, 1); +} else { + x_103 = x_101; +} +lean_ctor_set(x_103, 0, x_86); +lean_ctor_set(x_103, 1, x_87); +lean_ctor_set(x_103, 2, x_88); +lean_ctor_set(x_103, 3, x_89); +lean_ctor_set(x_103, 4, x_90); +lean_ctor_set(x_103, 5, x_91); +lean_ctor_set(x_103, 6, x_93); +lean_ctor_set(x_103, 7, x_94); +lean_ctor_set(x_103, 8, x_95); +lean_ctor_set(x_103, 9, x_96); +lean_ctor_set(x_103, 10, x_102); +lean_ctor_set(x_103, 11, x_98); +lean_ctor_set(x_103, 12, x_99); +lean_ctor_set(x_103, 13, x_100); +lean_ctor_set_uint8(x_103, sizeof(void*)*14, x_92); +x_104 = lean_st_ref_set(x_3, x_103, x_85); +x_105 = lean_ctor_get(x_104, 1); +lean_inc(x_105); +lean_dec(x_104); +x_106 = lean_st_ref_get(x_3, x_105); +x_107 = lean_ctor_get(x_106, 0); lean_inc(x_107); +x_108 = lean_ctor_get(x_106, 1); +lean_inc(x_108); lean_dec(x_106); -x_108 = 0; -x_109 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_108, x_108, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_107); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -return x_109; +x_109 = lean_ctor_get(x_107, 4); +lean_inc(x_109); +lean_dec(x_107); +x_110 = lean_box(0); +x_111 = lean_box(0); +lean_inc(x_82); +x_112 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8(x_2, x_82, x_109, x_110, x_82, x_82, x_111, lean_box(0), x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_108); +lean_dec(x_82); +if (lean_obj_tag(x_112) == 0) +{ +lean_object* x_113; lean_object* x_114; lean_object* x_115; +x_113 = lean_ctor_get(x_112, 1); +lean_inc(x_113); +if (lean_is_exclusive(x_112)) { + lean_ctor_release(x_112, 0); + lean_ctor_release(x_112, 1); + x_114 = x_112; +} else { + lean_dec_ref(x_112); + x_114 = lean_box(0); +} +if (lean_is_scalar(x_114)) { + x_115 = lean_alloc_ctor(0, 2, 0); +} else { + x_115 = x_114; +} +lean_ctor_set(x_115, 0, x_111); +lean_ctor_set(x_115, 1, x_113); +return x_115; } else { -lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; uint8_t x_118; lean_object* x_119; -x_110 = lean_ctor_get(x_93, 1); -lean_inc(x_110); -lean_dec(x_93); -lean_inc(x_1); -x_111 = l_Lean_indentExpr(x_1); -x_112 = l_Lean_Meta_Grind_internalize___lambda__2___closed__6; -x_113 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_113, 0, x_112); -lean_ctor_set(x_113, 1, x_111); -x_114 = l_Lean_Meta_Grind_addCongrTable___lambda__2___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_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_92, x_115, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_110); -x_117 = lean_ctor_get(x_116, 1); +lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; +x_116 = lean_ctor_get(x_112, 0); +lean_inc(x_116); +x_117 = lean_ctor_get(x_112, 1); lean_inc(x_117); -lean_dec(x_116); -x_118 = 0; -x_119 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_118, x_118, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_117); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -return x_119; +if (lean_is_exclusive(x_112)) { + lean_ctor_release(x_112, 0); + lean_ctor_release(x_112, 1); + x_118 = x_112; +} else { + lean_dec_ref(x_112); + x_118 = lean_box(0); } +if (lean_is_scalar(x_118)) { + x_119 = lean_alloc_ctor(1, 2, 0); +} else { + x_119 = x_118; } +lean_ctor_set(x_119, 0, x_116); +lean_ctor_set(x_119, 1, x_117); +return x_119; } -default: -{ -uint8_t x_120; lean_object* x_121; -x_120 = 0; -x_121 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_120, x_120, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -return x_121; } } } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_internalize(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_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___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { -lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_12 = l_Lean_Meta_Grind_alreadyInternalized(x_1, 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_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_Meta_Grind_internalize___lambda__2(x_1, x_2, x_16, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_15); -return x_17; -} -else -{ -uint8_t x_18; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); +size_t x_14; size_t x_15; lean_object* x_16; +x_14 = lean_unbox_usize(x_2); lean_dec(x_2); -lean_dec(x_1); -x_18 = !lean_is_exclusive(x_12); -if (x_18 == 0) -{ -lean_object* x_19; lean_object* x_20; -x_19 = lean_ctor_get(x_12, 0); -lean_dec(x_19); -x_20 = lean_box(0); -lean_ctor_set(x_12, 0, x_20); -return x_12; -} -else -{ -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_12, 1); -lean_inc(x_21); -lean_dec(x_12); -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; -} -} +x_15 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_16 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___spec__1(x_1, x_14, x_15, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_16; } } LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_internalize___spec__2___boxed(lean_object** _args) { @@ -7733,9 +11217,144 @@ lean_dec(x_3); return x_13; } } +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_internalize___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +x_13 = l_Lean_Meta_Grind_internalize___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_3); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_eraseAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___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_eraseAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__2(x_1, x_4, x_3); +lean_dec(x_3); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_erase___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_PersistentHashMap_erase___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__1(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAtAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___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 = l_Lean_PersistentHashMap_containsAtAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_5); +lean_dec(x_2); +lean_dec(x_1); +x_7 = lean_box(x_6); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__4___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___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__4(x_1, x_4, x_3); +lean_dec(x_3); +x_6 = lean_box(x_5); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_contains___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__3___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___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__3(x_1, x_2); +lean_dec(x_2); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___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, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; +x_12 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_12; +} +} +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___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_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___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_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_12; +} +} +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +_start: +{ +lean_object* x_18; +x_18 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +return x_18; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; +x_12 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns(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_1); +return x_12; +} +} lean_object* initialize_Init_Grind_Util(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_LitValues(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_Tactic_Grind_Types(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Meta_Tactic_Grind_Util(uint8_t builtin, lean_object*); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Lean_Meta_Tactic_Grind_Internalize(uint8_t builtin, lean_object* w) { lean_object * res; @@ -7750,6 +11369,9 @@ lean_dec_ref(res); res = initialize_Lean_Meta_Tactic_Grind_Types(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +res = initialize_Lean_Meta_Tactic_Grind_Util(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); l_Lean_PersistentHashMap_findEntryAux___at_Lean_Meta_Grind_addCongrTable___spec__2___closed__1 = _init_l_Lean_PersistentHashMap_findEntryAux___at_Lean_Meta_Grind_addCongrTable___spec__2___closed__1(); l_Lean_PersistentHashMap_findEntryAux___at_Lean_Meta_Grind_addCongrTable___spec__2___closed__2 = _init_l_Lean_PersistentHashMap_findEntryAux___at_Lean_Meta_Grind_addCongrTable___spec__2___closed__2(); l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_addCongrTable___spec__5___closed__1 = _init_l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_addCongrTable___spec__5___closed__1(); @@ -7771,6 +11393,8 @@ l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__5 = _init_l_Lean_Meta_Grin lean_mark_persistent(l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__5); l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__6 = _init_l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__6(); lean_mark_persistent(l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__6); +l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__7 = _init_l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__7(); +lean_mark_persistent(l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__7); l_Lean_Meta_Grind_addCongrTable___closed__1 = _init_l_Lean_Meta_Grind_addCongrTable___closed__1(); lean_mark_persistent(l_Lean_Meta_Grind_addCongrTable___closed__1); l_Lean_Meta_Grind_addCongrTable___closed__2 = _init_l_Lean_Meta_Grind_addCongrTable___closed__2(); @@ -7789,6 +11413,16 @@ l_Lean_Meta_Grind_addCongrTable___closed__8 = _init_l_Lean_Meta_Grind_addCongrTa lean_mark_persistent(l_Lean_Meta_Grind_addCongrTable___closed__8); l_Lean_Meta_Grind_addCongrTable___closed__9 = _init_l_Lean_Meta_Grind_addCongrTable___closed__9(); lean_mark_persistent(l_Lean_Meta_Grind_addCongrTable___closed__9); +l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__1); +l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__2); +l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__3); +l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__4(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__4); +l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__5 = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__5(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__5); l_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__1 = _init_l_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__1(); lean_mark_persistent(l_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__1); l_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__2 = _init_l_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__2(); @@ -7819,8 +11453,32 @@ l_Lean_Meta_Grind_internalize___lambda__2___closed__5 = _init_l_Lean_Meta_Grind_ lean_mark_persistent(l_Lean_Meta_Grind_internalize___lambda__2___closed__5); l_Lean_Meta_Grind_internalize___lambda__2___closed__6 = _init_l_Lean_Meta_Grind_internalize___lambda__2___closed__6(); lean_mark_persistent(l_Lean_Meta_Grind_internalize___lambda__2___closed__6); -l_Lean_Meta_Grind_internalize___lambda__2___closed__7 = _init_l_Lean_Meta_Grind_internalize___lambda__2___closed__7(); -lean_mark_persistent(l_Lean_Meta_Grind_internalize___lambda__2___closed__7); +l_Lean_Meta_Grind_internalize___lambda__3___closed__1 = _init_l_Lean_Meta_Grind_internalize___lambda__3___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_internalize___lambda__3___closed__1); +l_Lean_Meta_Grind_internalize___lambda__3___closed__2 = _init_l_Lean_Meta_Grind_internalize___lambda__3___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_internalize___lambda__3___closed__2); +l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__1___closed__1 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__1___closed__1(); +lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__1___closed__1); +l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__1 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__1(); +lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__1); +l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__2 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__2(); +lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__2); +l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__3 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__3(); +lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__3); +l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__4 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__4(); +lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__4); +l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__5 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__5(); +lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__5); +l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__6 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__6(); +lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__6); +l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__7 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__7(); +lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__7); +l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__8 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__8(); +lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__8); +l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__9 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__9(); +lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__9); +l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__10 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__10(); +lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__10); return lean_io_result_mk_ok(lean_box(0)); } #ifdef __cplusplus diff --git a/stage0/stdlib/Lean/Meta/Tactic/Grind/Intro.c b/stage0/stdlib/Lean/Meta/Tactic/Grind/Intro.c new file mode 100644 index 000000000000..1ccf0922ead3 --- /dev/null +++ b/stage0/stdlib/Lean/Meta/Tactic/Grind/Intro.c @@ -0,0 +1,4860 @@ +// Lean compiler output +// Module: Lean.Meta.Tactic.Grind.Intro +// Imports: Init.Grind.Lemmas Lean.Meta.Tactic.Assert Lean.Meta.Tactic.Grind.Simp Lean.Meta.Tactic.Grind.Types Lean.Meta.Tactic.Grind.Cases Lean.Meta.Tactic.Grind.Injection Lean.Meta.Tactic.Grind.Core +#include +#if defined(__clang__) +#pragma clang diagnostic ignored "-Wunused-parameter" +#pragma clang diagnostic ignored "-Wunused-label" +#elif defined(__GNUC__) && !defined(__CLANG__) +#pragma GCC diagnostic ignored "-Wunused-parameter" +#pragma GCC diagnostic ignored "-Wunused-label" +#pragma GCC diagnostic ignored "-Wunused-but-set-variable" +#endif +#ifdef __cplusplus +extern "C" { +#endif +lean_object* l_Lean_Expr_bindingName_x21(lean_object*); +lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); +lean_object* l_Lean_Meta_Grind_injection_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instInhabitedIntroResult; +lean_object* l_Lean_mkAppN(lean_object*, lean_object*); +lean_object* l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___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___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_mkFreshExprMVarAt(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_isProp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Expr_isLet(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_iterate(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_array_push(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_applyInjection_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_forM___at_Lean_Meta_Grind_intros_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*); +lean_object* l_Lean_MVarId_getTag(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_assertAt(lean_object*, lean_object*, lean_object*, 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_Grind_assertNext_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_mkFreshId___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_Grind_simp(lean_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_Grind_Intro_0__Lean_Meta_Grind_isCasesCandidate(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_object* l_List_appendTR___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_forM___at_Lean_Meta_Grind_intros_go___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext(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___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_applyCases_x3f___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_forM___at_Lean_Meta_Grind_intros_go___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___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___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_MVarId_getType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_intros_go___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__2___boxed(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*); +static lean_object* l_Lean_Meta_Grind_assertAll___closed__1; +lean_object* l_Lean_FVarId_getDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__6___boxed(lean_object**); +static lean_object* l_Lean_Meta_Grind_intros___closed__1; +lean_object* lean_st_ref_get(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_intros_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_st_mk_ref(lean_object*, lean_object*); +lean_object* lean_array_to_list(lean_object*); +lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_assertAt___closed__1; +lean_object* l_Lean_Meta_Grind_isGrindCasesTarget(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_mkFreshId___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___spec__2___rarg___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_mapTR_loop___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_applyCases_x3f___spec__1(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_FVarId_getType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Std_Queue_dequeue_x3f___rarg(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_applyCases_x3f(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*); +uint8_t l_Lean_Expr_isArrow(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___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_Meta_Grind_assertAt___closed__2; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_intros_go___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_MVarId_assert(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_LocalContext_mkLocalDecl(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_intros_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_Lean_Meta_Grind_iterate_go(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_intros_go(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_isCasesCandidate___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_LocalDecl_userName(lean_object*); +lean_object* l_Lean_Meta_Grind_addHypothesis(lean_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*); +lean_object* l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___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_Meta_Grind_intros_go___lambda__4___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_intros(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_PersistentHashMap_insert___at_Lean_MVarId_assign___spec__1(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_getLocalInstances(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__2; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_intros_go___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___boxed(lean_object**); +lean_object* l_Lean_LocalDecl_type(lean_object*); +LEAN_EXPORT lean_object* l_List_forM___at_Lean_Meta_Grind_intros_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_EXPORT lean_object* l_Lean_Meta_Grind_intros_go___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Expr_getAppFn(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_intros_go___lambda__4___boxed(lean_object**); +lean_object* l_List_reverse___rarg(lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__3; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_intros_go___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_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_MVarId_byContra_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_mkFreshId___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___spec__2___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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_Grind_assertAll(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_array_mk(lean_object*); +lean_object* l_Lean_Expr_fvar___override(lean_object*); +lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Expr_bindingBody_x21(lean_object*); +LEAN_EXPORT lean_object* l_List_mapTR_loop___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_applyCases_x3f___spec__1___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_MVarId_assign___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_Grind_cases(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_intro1Core(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_mkLambdaFVars(lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, 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___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__4; +lean_object* l_Lean_Expr_bindingInfo_x21(lean_object*); +uint8_t l_Lean_Expr_isForall(lean_object*); +LEAN_EXPORT lean_object* l_Lean_mkFreshFVarId___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___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_mkFreshId___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___spec__2(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_mkFreshFVarId___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___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_Tactic_Grind_Intro_0__Lean_Meta_Grind_applyCases_x3f___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_MVarId_assign___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* _init_l_Lean_Meta_Grind_instInhabitedIntroResult() { +_start: +{ +lean_object* x_1; +x_1 = lean_box(0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_mkFreshId___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___spec__2___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_get(x_1, x_2); +x_4 = lean_ctor_get(x_3, 0); +lean_inc(x_4); +x_5 = lean_ctor_get(x_4, 2); +lean_inc(x_5); +lean_dec(x_4); +x_6 = lean_ctor_get(x_3, 1); +lean_inc(x_6); +lean_dec(x_3); +x_7 = !lean_is_exclusive(x_5); +if (x_7 == 0) +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_8 = lean_ctor_get(x_5, 0); +x_9 = lean_ctor_get(x_5, 1); +lean_inc(x_9); +lean_inc(x_8); +x_10 = l_Lean_Name_num___override(x_8, x_9); +x_11 = lean_unsigned_to_nat(1u); +x_12 = lean_nat_add(x_9, x_11); +lean_dec(x_9); +lean_ctor_set(x_5, 1, x_12); +x_13 = lean_st_ref_take(x_1, 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_is_exclusive(x_14); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; uint8_t x_19; +x_17 = lean_ctor_get(x_14, 2); +lean_dec(x_17); +lean_ctor_set(x_14, 2, x_5); +x_18 = lean_st_ref_set(x_1, x_14, x_15); +x_19 = !lean_is_exclusive(x_18); +if (x_19 == 0) +{ +lean_object* x_20; +x_20 = lean_ctor_get(x_18, 0); +lean_dec(x_20); +lean_ctor_set(x_18, 0, x_10); +return x_18; +} +else +{ +lean_object* x_21; lean_object* x_22; +x_21 = lean_ctor_get(x_18, 1); +lean_inc(x_21); +lean_dec(x_18); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_10); +lean_ctor_set(x_22, 1, x_21); +return x_22; +} +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_23 = lean_ctor_get(x_14, 0); +x_24 = lean_ctor_get(x_14, 1); +x_25 = lean_ctor_get(x_14, 3); +x_26 = lean_ctor_get(x_14, 4); +x_27 = lean_ctor_get(x_14, 5); +x_28 = lean_ctor_get(x_14, 6); +x_29 = lean_ctor_get(x_14, 7); +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_14); +x_30 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_30, 0, x_23); +lean_ctor_set(x_30, 1, x_24); +lean_ctor_set(x_30, 2, x_5); +lean_ctor_set(x_30, 3, x_25); +lean_ctor_set(x_30, 4, x_26); +lean_ctor_set(x_30, 5, x_27); +lean_ctor_set(x_30, 6, x_28); +lean_ctor_set(x_30, 7, x_29); +x_31 = lean_st_ref_set(x_1, x_30, x_15); +x_32 = lean_ctor_get(x_31, 1); +lean_inc(x_32); +if (lean_is_exclusive(x_31)) { + lean_ctor_release(x_31, 0); + lean_ctor_release(x_31, 1); + x_33 = x_31; +} else { + lean_dec_ref(x_31); + x_33 = lean_box(0); +} +if (lean_is_scalar(x_33)) { + x_34 = lean_alloc_ctor(0, 2, 0); +} else { + x_34 = x_33; +} +lean_ctor_set(x_34, 0, x_10); +lean_ctor_set(x_34, 1, x_32); +return x_34; +} +} +else +{ +lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; 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_35 = lean_ctor_get(x_5, 0); +x_36 = lean_ctor_get(x_5, 1); +lean_inc(x_36); +lean_inc(x_35); +lean_dec(x_5); +lean_inc(x_36); +lean_inc(x_35); +x_37 = l_Lean_Name_num___override(x_35, x_36); +x_38 = lean_unsigned_to_nat(1u); +x_39 = lean_nat_add(x_36, x_38); +lean_dec(x_36); +x_40 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_40, 0, x_35); +lean_ctor_set(x_40, 1, x_39); +x_41 = lean_st_ref_take(x_1, 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 = lean_ctor_get(x_42, 0); +lean_inc(x_44); +x_45 = lean_ctor_get(x_42, 1); +lean_inc(x_45); +x_46 = lean_ctor_get(x_42, 3); +lean_inc(x_46); +x_47 = lean_ctor_get(x_42, 4); +lean_inc(x_47); +x_48 = lean_ctor_get(x_42, 5); +lean_inc(x_48); +x_49 = lean_ctor_get(x_42, 6); +lean_inc(x_49); +x_50 = lean_ctor_get(x_42, 7); +lean_inc(x_50); +if (lean_is_exclusive(x_42)) { + lean_ctor_release(x_42, 0); + lean_ctor_release(x_42, 1); + lean_ctor_release(x_42, 2); + lean_ctor_release(x_42, 3); + lean_ctor_release(x_42, 4); + lean_ctor_release(x_42, 5); + lean_ctor_release(x_42, 6); + lean_ctor_release(x_42, 7); + x_51 = x_42; +} else { + lean_dec_ref(x_42); + x_51 = lean_box(0); +} +if (lean_is_scalar(x_51)) { + x_52 = lean_alloc_ctor(0, 8, 0); +} else { + x_52 = x_51; +} +lean_ctor_set(x_52, 0, x_44); +lean_ctor_set(x_52, 1, x_45); +lean_ctor_set(x_52, 2, x_40); +lean_ctor_set(x_52, 3, x_46); +lean_ctor_set(x_52, 4, x_47); +lean_ctor_set(x_52, 5, x_48); +lean_ctor_set(x_52, 6, x_49); +lean_ctor_set(x_52, 7, x_50); +x_53 = lean_st_ref_set(x_1, x_52, x_43); +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); +} +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_37); +lean_ctor_set(x_56, 1, x_54); +return x_56; +} +} +} +LEAN_EXPORT lean_object* l_Lean_mkFreshId___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; +x_7 = lean_alloc_closure((void*)(l_Lean_mkFreshId___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___spec__2___rarg___boxed), 2, 0); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_mkFreshFVarId___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___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 = l_Lean_mkFreshId___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___spec__2___rarg(x_7, x_8); +x_10 = !lean_is_exclusive(x_9); +if (x_10 == 0) +{ +return x_9; +} +else +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_11 = lean_ctor_get(x_9, 0); +x_12 = lean_ctor_get(x_9, 1); +lean_inc(x_12); +lean_inc(x_11); +lean_dec(x_9); +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_11); +lean_ctor_set(x_13, 1, x_12); +return x_13; +} +} +} +LEAN_EXPORT lean_object* l_Lean_MVarId_assign___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_11 = lean_st_ref_take(x_7, x_10); +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_11, 1); +lean_inc(x_14); +lean_dec(x_11); +x_15 = !lean_is_exclusive(x_12); +if (x_15 == 0) +{ +lean_object* x_16; uint8_t x_17; +x_16 = lean_ctor_get(x_12, 0); +lean_dec(x_16); +x_17 = !lean_is_exclusive(x_13); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_18 = lean_ctor_get(x_13, 7); +x_19 = l_Lean_PersistentHashMap_insert___at_Lean_MVarId_assign___spec__1(x_18, x_1, x_2); +lean_ctor_set(x_13, 7, x_19); +x_20 = lean_st_ref_set(x_7, x_12, x_14); +x_21 = !lean_is_exclusive(x_20); +if (x_21 == 0) +{ +lean_object* x_22; lean_object* x_23; +x_22 = lean_ctor_get(x_20, 0); +lean_dec(x_22); +x_23 = lean_box(0); +lean_ctor_set(x_20, 0, x_23); +return x_20; +} +else +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_20, 1); +lean_inc(x_24); +lean_dec(x_20); +x_25 = lean_box(0); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set(x_26, 1, x_24); +return x_26; +} +} +else +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_27 = lean_ctor_get(x_13, 0); +x_28 = lean_ctor_get(x_13, 1); +x_29 = lean_ctor_get(x_13, 2); +x_30 = lean_ctor_get(x_13, 3); +x_31 = lean_ctor_get(x_13, 4); +x_32 = lean_ctor_get(x_13, 5); +x_33 = lean_ctor_get(x_13, 6); +x_34 = lean_ctor_get(x_13, 7); +x_35 = lean_ctor_get(x_13, 8); +lean_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_dec(x_13); +x_36 = l_Lean_PersistentHashMap_insert___at_Lean_MVarId_assign___spec__1(x_34, x_1, x_2); +x_37 = lean_alloc_ctor(0, 9, 0); +lean_ctor_set(x_37, 0, x_27); +lean_ctor_set(x_37, 1, x_28); +lean_ctor_set(x_37, 2, x_29); +lean_ctor_set(x_37, 3, x_30); +lean_ctor_set(x_37, 4, x_31); +lean_ctor_set(x_37, 5, x_32); +lean_ctor_set(x_37, 6, x_33); +lean_ctor_set(x_37, 7, x_36); +lean_ctor_set(x_37, 8, x_35); +lean_ctor_set(x_12, 0, x_37); +x_38 = lean_st_ref_set(x_7, x_12, x_14); +x_39 = lean_ctor_get(x_38, 1); +lean_inc(x_39); +if (lean_is_exclusive(x_38)) { + lean_ctor_release(x_38, 0); + lean_ctor_release(x_38, 1); + x_40 = x_38; +} else { + lean_dec_ref(x_38); + x_40 = lean_box(0); +} +x_41 = lean_box(0); +if (lean_is_scalar(x_40)) { + x_42 = lean_alloc_ctor(0, 2, 0); +} else { + x_42 = x_40; +} +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_39); +return x_42; +} +} +else +{ +lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; 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_43 = lean_ctor_get(x_12, 1); +x_44 = lean_ctor_get(x_12, 2); +x_45 = lean_ctor_get(x_12, 3); +x_46 = lean_ctor_get(x_12, 4); +lean_inc(x_46); +lean_inc(x_45); +lean_inc(x_44); +lean_inc(x_43); +lean_dec(x_12); +x_47 = lean_ctor_get(x_13, 0); +lean_inc(x_47); +x_48 = lean_ctor_get(x_13, 1); +lean_inc(x_48); +x_49 = lean_ctor_get(x_13, 2); +lean_inc(x_49); +x_50 = lean_ctor_get(x_13, 3); +lean_inc(x_50); +x_51 = lean_ctor_get(x_13, 4); +lean_inc(x_51); +x_52 = lean_ctor_get(x_13, 5); +lean_inc(x_52); +x_53 = lean_ctor_get(x_13, 6); +lean_inc(x_53); +x_54 = lean_ctor_get(x_13, 7); +lean_inc(x_54); +x_55 = lean_ctor_get(x_13, 8); +lean_inc(x_55); +if (lean_is_exclusive(x_13)) { + lean_ctor_release(x_13, 0); + lean_ctor_release(x_13, 1); + lean_ctor_release(x_13, 2); + lean_ctor_release(x_13, 3); + lean_ctor_release(x_13, 4); + lean_ctor_release(x_13, 5); + lean_ctor_release(x_13, 6); + lean_ctor_release(x_13, 7); + lean_ctor_release(x_13, 8); + x_56 = x_13; +} else { + lean_dec_ref(x_13); + x_56 = lean_box(0); +} +x_57 = l_Lean_PersistentHashMap_insert___at_Lean_MVarId_assign___spec__1(x_54, x_1, x_2); +if (lean_is_scalar(x_56)) { + x_58 = lean_alloc_ctor(0, 9, 0); +} else { + x_58 = x_56; +} +lean_ctor_set(x_58, 0, x_47); +lean_ctor_set(x_58, 1, x_48); +lean_ctor_set(x_58, 2, x_49); +lean_ctor_set(x_58, 3, x_50); +lean_ctor_set(x_58, 4, x_51); +lean_ctor_set(x_58, 5, x_52); +lean_ctor_set(x_58, 6, x_53); +lean_ctor_set(x_58, 7, x_57); +lean_ctor_set(x_58, 8, x_55); +x_59 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_59, 0, x_58); +lean_ctor_set(x_59, 1, x_43); +lean_ctor_set(x_59, 2, x_44); +lean_ctor_set(x_59, 3, x_45); +lean_ctor_set(x_59, 4, x_46); +x_60 = lean_st_ref_set(x_7, x_59, x_14); +x_61 = lean_ctor_get(x_60, 1); +lean_inc(x_61); +if (lean_is_exclusive(x_60)) { + lean_ctor_release(x_60, 0); + lean_ctor_release(x_60, 1); + x_62 = x_60; +} else { + lean_dec_ref(x_60); + x_62 = lean_box(0); +} +x_63 = lean_box(0); +if (lean_is_scalar(x_62)) { + x_64 = lean_alloc_ctor(0, 2, 0); +} else { + x_64 = x_62; +} +lean_ctor_set(x_64, 0, x_63); +lean_ctor_set(x_64, 1, x_61); +return x_64; +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l_Lean_FVarId_getDecl(x_1, x_5, x_6, x_7, x_8, x_9); +return x_10; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, uint8_t x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19, lean_object* x_20, lean_object* x_21, lean_object* x_22, lean_object* x_23, lean_object* x_24, lean_object* x_25) { +_start: +{ +lean_object* x_26; lean_object* x_27; +x_26 = l_Lean_LocalDecl_type(x_17); +lean_inc(x_24); +lean_inc(x_23); +lean_inc(x_22); +lean_inc(x_21); +lean_inc(x_26); +x_27 = l_Lean_Meta_isProp(x_26, x_21, x_22, x_23, x_24, x_25); +if (lean_obj_tag(x_27) == 0) +{ +lean_object* x_28; uint8_t x_29; +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) +{ +uint8_t x_30; +lean_dec(x_26); +lean_dec(x_24); +lean_dec(x_23); +lean_dec(x_22); +lean_dec(x_21); +x_30 = !lean_is_exclusive(x_27); +if (x_30 == 0) +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_27, 0); +lean_dec(x_31); +x_32 = lean_alloc_ctor(0, 14, 1); +lean_ctor_set(x_32, 0, x_1); +lean_ctor_set(x_32, 1, x_2); +lean_ctor_set(x_32, 2, x_3); +lean_ctor_set(x_32, 3, x_4); +lean_ctor_set(x_32, 4, x_5); +lean_ctor_set(x_32, 5, x_6); +lean_ctor_set(x_32, 6, x_8); +lean_ctor_set(x_32, 7, x_9); +lean_ctor_set(x_32, 8, x_10); +lean_ctor_set(x_32, 9, x_11); +lean_ctor_set(x_32, 10, x_12); +lean_ctor_set(x_32, 11, x_13); +lean_ctor_set(x_32, 12, x_14); +lean_ctor_set(x_32, 13, x_15); +lean_ctor_set_uint8(x_32, sizeof(void*)*14, x_7); +x_33 = lean_alloc_ctor(3, 2, 0); +lean_ctor_set(x_33, 0, x_16); +lean_ctor_set(x_33, 1, x_32); +lean_ctor_set(x_27, 0, x_33); +return x_27; +} +else +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_34 = lean_ctor_get(x_27, 1); +lean_inc(x_34); +lean_dec(x_27); +x_35 = lean_alloc_ctor(0, 14, 1); +lean_ctor_set(x_35, 0, x_1); +lean_ctor_set(x_35, 1, x_2); +lean_ctor_set(x_35, 2, x_3); +lean_ctor_set(x_35, 3, x_4); +lean_ctor_set(x_35, 4, x_5); +lean_ctor_set(x_35, 5, x_6); +lean_ctor_set(x_35, 6, x_8); +lean_ctor_set(x_35, 7, x_9); +lean_ctor_set(x_35, 8, x_10); +lean_ctor_set(x_35, 9, x_11); +lean_ctor_set(x_35, 10, x_12); +lean_ctor_set(x_35, 11, x_13); +lean_ctor_set(x_35, 12, x_14); +lean_ctor_set(x_35, 13, x_15); +lean_ctor_set_uint8(x_35, sizeof(void*)*14, x_7); +x_36 = lean_alloc_ctor(3, 2, 0); +lean_ctor_set(x_36, 0, x_16); +lean_ctor_set(x_36, 1, x_35); +x_37 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_37, 0, x_36); +lean_ctor_set(x_37, 1, x_34); +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; +x_38 = lean_ctor_get(x_27, 1); +lean_inc(x_38); +lean_dec(x_27); +x_39 = l_Lean_LocalDecl_userName(x_17); +x_40 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_39, x_23, x_24, x_38); +x_41 = lean_ctor_get(x_40, 0); +lean_inc(x_41); +x_42 = lean_ctor_get(x_40, 1); +lean_inc(x_42); +lean_dec(x_40); +x_43 = l_Lean_Expr_fvar___override(x_16); +x_44 = l_Lean_MVarId_assert(x_1, x_41, x_26, x_43, x_21, x_22, x_23, x_24, x_42); +if (lean_obj_tag(x_44) == 0) +{ +uint8_t x_45; +x_45 = !lean_is_exclusive(x_44); +if (x_45 == 0) +{ +lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_46 = lean_ctor_get(x_44, 0); +x_47 = lean_alloc_ctor(0, 14, 1); +lean_ctor_set(x_47, 0, x_46); +lean_ctor_set(x_47, 1, x_2); +lean_ctor_set(x_47, 2, x_3); +lean_ctor_set(x_47, 3, x_4); +lean_ctor_set(x_47, 4, x_5); +lean_ctor_set(x_47, 5, x_6); +lean_ctor_set(x_47, 6, x_8); +lean_ctor_set(x_47, 7, x_9); +lean_ctor_set(x_47, 8, x_10); +lean_ctor_set(x_47, 9, x_11); +lean_ctor_set(x_47, 10, x_12); +lean_ctor_set(x_47, 11, x_13); +lean_ctor_set(x_47, 12, x_14); +lean_ctor_set(x_47, 13, x_15); +lean_ctor_set_uint8(x_47, sizeof(void*)*14, x_7); +x_48 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_48, 0, x_47); +lean_ctor_set(x_44, 0, x_48); +return x_44; +} +else +{ +lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_49 = lean_ctor_get(x_44, 0); +x_50 = lean_ctor_get(x_44, 1); +lean_inc(x_50); +lean_inc(x_49); +lean_dec(x_44); +x_51 = lean_alloc_ctor(0, 14, 1); +lean_ctor_set(x_51, 0, x_49); +lean_ctor_set(x_51, 1, x_2); +lean_ctor_set(x_51, 2, x_3); +lean_ctor_set(x_51, 3, x_4); +lean_ctor_set(x_51, 4, x_5); +lean_ctor_set(x_51, 5, x_6); +lean_ctor_set(x_51, 6, x_8); +lean_ctor_set(x_51, 7, x_9); +lean_ctor_set(x_51, 8, x_10); +lean_ctor_set(x_51, 9, x_11); +lean_ctor_set(x_51, 10, x_12); +lean_ctor_set(x_51, 11, x_13); +lean_ctor_set(x_51, 12, x_14); +lean_ctor_set(x_51, 13, x_15); +lean_ctor_set_uint8(x_51, sizeof(void*)*14, x_7); +x_52 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_52, 0, x_51); +x_53 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_53, 0, x_52); +lean_ctor_set(x_53, 1, x_50); +return x_53; +} +} +else +{ +uint8_t x_54; +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_54 = !lean_is_exclusive(x_44); +if (x_54 == 0) +{ +return x_44; +} +else +{ +lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_55 = lean_ctor_get(x_44, 0); +x_56 = lean_ctor_get(x_44, 1); +lean_inc(x_56); +lean_inc(x_55); +lean_dec(x_44); +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_58; +lean_dec(x_26); +lean_dec(x_24); +lean_dec(x_23); +lean_dec(x_22); +lean_dec(x_21); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_58 = !lean_is_exclusive(x_27); +if (x_58 == 0) +{ +return x_27; +} +else +{ +lean_object* x_59; lean_object* x_60; lean_object* x_61; +x_59 = lean_ctor_get(x_27, 0); +x_60 = lean_ctor_get(x_27, 1); +lean_inc(x_60); +lean_inc(x_59); +lean_dec(x_27); +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_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___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; +x_10 = l_Lean_Meta_isProp(x_1, x_5, x_6, x_7, x_8, x_9); +return x_10; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___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_11; uint8_t x_12; uint8_t x_13; lean_object* x_14; +x_11 = 0; +x_12 = 1; +x_13 = 1; +x_14 = l_Lean_Meta_mkLambdaFVars(x_1, x_2, x_11, x_12, x_11, x_13, x_6, x_7, x_8, x_9, x_10); +return x_14; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Grind", 5, 5); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("intro_with_eq", 13, 13); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__2; +x_3 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__3; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, uint8_t x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19, lean_object* x_20, lean_object* x_21, lean_object* x_22, lean_object* x_23, lean_object* x_24, lean_object* x_25, lean_object* x_26, lean_object* x_27, lean_object* x_28, lean_object* x_29, lean_object* x_30, lean_object* x_31) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_32; uint8_t x_33; +lean_dec(x_22); +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_19); +x_32 = l_Lean_MVarId_assign___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___spec__3(x_2, x_23, x_24, x_25, x_26, x_27, x_28, x_29, x_30, x_31); +x_33 = !lean_is_exclusive(x_32); +if (x_33 == 0) +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_32, 0); +lean_dec(x_34); +x_35 = lean_alloc_ctor(0, 14, 1); +lean_ctor_set(x_35, 0, x_3); +lean_ctor_set(x_35, 1, x_4); +lean_ctor_set(x_35, 2, x_5); +lean_ctor_set(x_35, 3, x_6); +lean_ctor_set(x_35, 4, x_7); +lean_ctor_set(x_35, 5, x_8); +lean_ctor_set(x_35, 6, x_10); +lean_ctor_set(x_35, 7, x_11); +lean_ctor_set(x_35, 8, x_12); +lean_ctor_set(x_35, 9, x_13); +lean_ctor_set(x_35, 10, x_14); +lean_ctor_set(x_35, 11, x_15); +lean_ctor_set(x_35, 12, x_16); +lean_ctor_set(x_35, 13, x_17); +lean_ctor_set_uint8(x_35, sizeof(void*)*14, x_9); +x_36 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_36, 0, x_18); +lean_ctor_set(x_36, 1, x_35); +lean_ctor_set(x_32, 0, x_36); +return x_32; +} +else +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_37 = lean_ctor_get(x_32, 1); +lean_inc(x_37); +lean_dec(x_32); +x_38 = lean_alloc_ctor(0, 14, 1); +lean_ctor_set(x_38, 0, x_3); +lean_ctor_set(x_38, 1, x_4); +lean_ctor_set(x_38, 2, x_5); +lean_ctor_set(x_38, 3, x_6); +lean_ctor_set(x_38, 4, x_7); +lean_ctor_set(x_38, 5, x_8); +lean_ctor_set(x_38, 6, x_10); +lean_ctor_set(x_38, 7, x_11); +lean_ctor_set(x_38, 8, x_12); +lean_ctor_set(x_38, 9, x_13); +lean_ctor_set(x_38, 10, x_14); +lean_ctor_set(x_38, 11, x_15); +lean_ctor_set(x_38, 12, x_16); +lean_ctor_set(x_38, 13, x_17); +lean_ctor_set_uint8(x_38, sizeof(void*)*14, x_9); +x_39 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_39, 0, x_18); +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_37); +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; uint8_t x_52; +x_41 = lean_ctor_get(x_1, 0); +x_42 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__4; +lean_inc(x_19); +x_43 = l_Lean_Expr_const___override(x_42, x_19); +x_44 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_44, 0, x_23); +lean_ctor_set(x_44, 1, x_19); +lean_inc(x_41); +x_45 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_45, 0, x_41); +lean_ctor_set(x_45, 1, x_44); +x_46 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_46, 0, x_20); +lean_ctor_set(x_46, 1, x_45); +x_47 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_47, 0, x_21); +lean_ctor_set(x_47, 1, x_46); +x_48 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_48, 0, x_22); +lean_ctor_set(x_48, 1, x_47); +x_49 = lean_array_mk(x_48); +x_50 = l_Lean_mkAppN(x_43, x_49); +lean_dec(x_49); +x_51 = l_Lean_MVarId_assign___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___spec__3(x_2, x_50, x_24, x_25, x_26, x_27, x_28, x_29, x_30, x_31); +x_52 = !lean_is_exclusive(x_51); +if (x_52 == 0) +{ +lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_53 = lean_ctor_get(x_51, 0); +lean_dec(x_53); +x_54 = lean_alloc_ctor(0, 14, 1); +lean_ctor_set(x_54, 0, x_3); +lean_ctor_set(x_54, 1, x_4); +lean_ctor_set(x_54, 2, x_5); +lean_ctor_set(x_54, 3, x_6); +lean_ctor_set(x_54, 4, x_7); +lean_ctor_set(x_54, 5, x_8); +lean_ctor_set(x_54, 6, x_10); +lean_ctor_set(x_54, 7, x_11); +lean_ctor_set(x_54, 8, x_12); +lean_ctor_set(x_54, 9, x_13); +lean_ctor_set(x_54, 10, x_14); +lean_ctor_set(x_54, 11, x_15); +lean_ctor_set(x_54, 12, x_16); +lean_ctor_set(x_54, 13, x_17); +lean_ctor_set_uint8(x_54, sizeof(void*)*14, x_9); +x_55 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_55, 0, x_18); +lean_ctor_set(x_55, 1, x_54); +lean_ctor_set(x_51, 0, x_55); +return x_51; +} +else +{ +lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; +x_56 = lean_ctor_get(x_51, 1); +lean_inc(x_56); +lean_dec(x_51); +x_57 = lean_alloc_ctor(0, 14, 1); +lean_ctor_set(x_57, 0, x_3); +lean_ctor_set(x_57, 1, x_4); +lean_ctor_set(x_57, 2, x_5); +lean_ctor_set(x_57, 3, x_6); +lean_ctor_set(x_57, 4, x_7); +lean_ctor_set(x_57, 5, x_8); +lean_ctor_set(x_57, 6, x_10); +lean_ctor_set(x_57, 7, x_11); +lean_ctor_set(x_57, 8, x_12); +lean_ctor_set(x_57, 9, x_13); +lean_ctor_set(x_57, 10, x_14); +lean_ctor_set(x_57, 11, x_15); +lean_ctor_set(x_57, 12, x_16); +lean_ctor_set(x_57, 13, x_17); +lean_ctor_set_uint8(x_57, sizeof(void*)*14, x_9); +x_58 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_58, 0, x_18); +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_56); +return x_59; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___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, 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, 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) { +_start: +{ +if (x_18 == 0) +{ +uint8_t x_27; lean_object* x_28; +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_17); +x_27 = 1; +x_28 = l_Lean_Meta_intro1Core(x_1, x_27, x_22, x_23, x_24, x_25, x_26); +if (lean_obj_tag(x_28) == 0) +{ +uint8_t x_29; +x_29 = !lean_is_exclusive(x_28); +if (x_29 == 0) +{ +lean_object* x_30; uint8_t x_31; +x_30 = lean_ctor_get(x_28, 0); +x_31 = !lean_is_exclusive(x_30); +if (x_31 == 0) +{ +lean_object* x_32; lean_object* x_33; +x_32 = lean_ctor_get(x_30, 1); +x_33 = lean_alloc_ctor(0, 14, 1); +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_2); +lean_ctor_set(x_33, 2, x_3); +lean_ctor_set(x_33, 3, x_4); +lean_ctor_set(x_33, 4, x_5); +lean_ctor_set(x_33, 5, x_6); +lean_ctor_set(x_33, 6, x_8); +lean_ctor_set(x_33, 7, x_9); +lean_ctor_set(x_33, 8, x_10); +lean_ctor_set(x_33, 9, x_11); +lean_ctor_set(x_33, 10, x_12); +lean_ctor_set(x_33, 11, x_13); +lean_ctor_set(x_33, 12, x_14); +lean_ctor_set(x_33, 13, x_15); +lean_ctor_set_uint8(x_33, sizeof(void*)*14, x_7); +lean_ctor_set_tag(x_30, 3); +lean_ctor_set(x_30, 1, x_33); +return x_28; +} +else +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_34 = lean_ctor_get(x_30, 0); +x_35 = lean_ctor_get(x_30, 1); +lean_inc(x_35); +lean_inc(x_34); +lean_dec(x_30); +x_36 = lean_alloc_ctor(0, 14, 1); +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set(x_36, 1, x_2); +lean_ctor_set(x_36, 2, x_3); +lean_ctor_set(x_36, 3, x_4); +lean_ctor_set(x_36, 4, x_5); +lean_ctor_set(x_36, 5, x_6); +lean_ctor_set(x_36, 6, x_8); +lean_ctor_set(x_36, 7, x_9); +lean_ctor_set(x_36, 8, x_10); +lean_ctor_set(x_36, 9, x_11); +lean_ctor_set(x_36, 10, x_12); +lean_ctor_set(x_36, 11, x_13); +lean_ctor_set(x_36, 12, x_14); +lean_ctor_set(x_36, 13, x_15); +lean_ctor_set_uint8(x_36, sizeof(void*)*14, x_7); +x_37 = lean_alloc_ctor(3, 2, 0); +lean_ctor_set(x_37, 0, x_34); +lean_ctor_set(x_37, 1, x_36); +lean_ctor_set(x_28, 0, x_37); +return x_28; +} +} +else +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_38 = lean_ctor_get(x_28, 0); +x_39 = lean_ctor_get(x_28, 1); +lean_inc(x_39); +lean_inc(x_38); +lean_dec(x_28); +x_40 = lean_ctor_get(x_38, 0); +lean_inc(x_40); +x_41 = lean_ctor_get(x_38, 1); +lean_inc(x_41); +if (lean_is_exclusive(x_38)) { + lean_ctor_release(x_38, 0); + lean_ctor_release(x_38, 1); + x_42 = x_38; +} else { + lean_dec_ref(x_38); + x_42 = lean_box(0); +} +x_43 = lean_alloc_ctor(0, 14, 1); +lean_ctor_set(x_43, 0, x_41); +lean_ctor_set(x_43, 1, x_2); +lean_ctor_set(x_43, 2, x_3); +lean_ctor_set(x_43, 3, x_4); +lean_ctor_set(x_43, 4, x_5); +lean_ctor_set(x_43, 5, x_6); +lean_ctor_set(x_43, 6, x_8); +lean_ctor_set(x_43, 7, x_9); +lean_ctor_set(x_43, 8, x_10); +lean_ctor_set(x_43, 9, x_11); +lean_ctor_set(x_43, 10, x_12); +lean_ctor_set(x_43, 11, x_13); +lean_ctor_set(x_43, 12, x_14); +lean_ctor_set(x_43, 13, x_15); +lean_ctor_set_uint8(x_43, sizeof(void*)*14, x_7); +if (lean_is_scalar(x_42)) { + x_44 = lean_alloc_ctor(3, 2, 0); +} else { + x_44 = x_42; + lean_ctor_set_tag(x_44, 3); +} +lean_ctor_set(x_44, 0, x_40); +lean_ctor_set(x_44, 1, x_43); +x_45 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_45, 0, x_44); +lean_ctor_set(x_45, 1, x_39); +return x_45; +} +} +else +{ +uint8_t x_46; +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_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_46 = !lean_is_exclusive(x_28); +if (x_46 == 0) +{ +return x_28; +} +else +{ +lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_47 = lean_ctor_get(x_28, 0); +x_48 = lean_ctor_get(x_28, 1); +lean_inc(x_48); +lean_inc(x_47); +lean_dec(x_28); +x_49 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_49, 0, x_47); +lean_ctor_set(x_49, 1, x_48); +return x_49; +} +} +} +else +{ +lean_object* x_50; +lean_inc(x_1); +x_50 = l_Lean_MVarId_getTag(x_1, x_22, x_23, x_24, x_25, x_26); +if (lean_obj_tag(x_50) == 0) +{ +lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; +x_51 = lean_ctor_get(x_50, 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_Expr_bindingBody_x21(x_16); +lean_inc(x_25); +lean_inc(x_24); +lean_inc(x_23); +lean_inc(x_22); +lean_inc(x_20); +lean_inc(x_17); +x_54 = l_Lean_Meta_Grind_simp(x_17, x_19, x_20, x_21, x_22, x_23, x_24, x_25, x_52); +if (lean_obj_tag(x_54) == 0) +{ +lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; uint8_t x_65; uint8_t x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; uint8_t x_71; lean_object* x_72; lean_object* x_73; uint8_t x_74; +x_55 = lean_ctor_get(x_54, 0); +lean_inc(x_55); +x_56 = lean_ctor_get(x_54, 1); +lean_inc(x_56); +lean_dec(x_54); +x_57 = l_Lean_mkFreshFVarId___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___spec__1(x_19, x_20, x_21, x_22, x_23, x_24, x_25, x_56); +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_ctor_get(x_22, 2); +lean_inc(x_60); +x_61 = l_Lean_Expr_bindingName_x21(x_16); +x_62 = lean_ctor_get(x_55, 0); +lean_inc(x_62); +x_63 = lean_ctor_get(x_55, 1); +lean_inc(x_63); +lean_dec(x_55); +x_64 = l_Lean_Expr_bindingInfo_x21(x_16); +x_65 = lean_unbox(x_64); +lean_dec(x_64); +x_66 = 0; +lean_inc(x_62); +lean_inc(x_58); +x_67 = l_Lean_LocalContext_mkLocalDecl(x_60, x_58, x_61, x_62, x_65, x_66); +x_68 = l_Lean_Meta_getLocalInstances(x_22, x_23, x_24, x_25, x_59); +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 = 2; +x_72 = lean_unsigned_to_nat(0u); +lean_inc(x_53); +x_73 = l_Lean_Meta_mkFreshExprMVarAt(x_67, x_69, x_53, x_71, x_51, x_72, x_22, x_23, x_24, x_25, x_70); +x_74 = !lean_is_exclusive(x_73); +if (x_74 == 0) +{ +lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; +x_75 = lean_ctor_get(x_73, 0); +x_76 = lean_ctor_get(x_73, 1); +x_77 = l_Lean_Expr_mvarId_x21(x_75); +lean_inc(x_58); +x_78 = l_Lean_Expr_fvar___override(x_58); +x_79 = lean_box(0); +lean_ctor_set_tag(x_73, 1); +lean_ctor_set(x_73, 1, x_79); +lean_ctor_set(x_73, 0, x_78); +x_80 = lean_array_mk(x_73); +x_81 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__4___boxed), 10, 2); +lean_closure_set(x_81, 0, x_80); +lean_closure_set(x_81, 1, x_75); +x_82 = lean_box(x_7); +lean_inc(x_77); +x_83 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___boxed), 31, 22); +lean_closure_set(x_83, 0, x_63); +lean_closure_set(x_83, 1, x_1); +lean_closure_set(x_83, 2, x_77); +lean_closure_set(x_83, 3, x_2); +lean_closure_set(x_83, 4, x_3); +lean_closure_set(x_83, 5, x_4); +lean_closure_set(x_83, 6, x_5); +lean_closure_set(x_83, 7, x_6); +lean_closure_set(x_83, 8, x_82); +lean_closure_set(x_83, 9, x_8); +lean_closure_set(x_83, 10, x_9); +lean_closure_set(x_83, 11, x_10); +lean_closure_set(x_83, 12, x_11); +lean_closure_set(x_83, 13, x_12); +lean_closure_set(x_83, 14, x_13); +lean_closure_set(x_83, 15, x_14); +lean_closure_set(x_83, 16, x_15); +lean_closure_set(x_83, 17, x_58); +lean_closure_set(x_83, 18, x_79); +lean_closure_set(x_83, 19, x_53); +lean_closure_set(x_83, 20, x_62); +lean_closure_set(x_83, 21, x_17); +x_84 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); +lean_closure_set(x_84, 0, x_81); +lean_closure_set(x_84, 1, x_83); +x_85 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(x_77, x_84, x_19, x_20, x_21, x_22, x_23, x_24, x_25, x_76); +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; lean_object* x_97; +x_86 = lean_ctor_get(x_73, 0); +x_87 = lean_ctor_get(x_73, 1); +lean_inc(x_87); +lean_inc(x_86); +lean_dec(x_73); +x_88 = l_Lean_Expr_mvarId_x21(x_86); +lean_inc(x_58); +x_89 = l_Lean_Expr_fvar___override(x_58); +x_90 = lean_box(0); +x_91 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_91, 0, x_89); +lean_ctor_set(x_91, 1, x_90); +x_92 = lean_array_mk(x_91); +x_93 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__4___boxed), 10, 2); +lean_closure_set(x_93, 0, x_92); +lean_closure_set(x_93, 1, x_86); +x_94 = lean_box(x_7); +lean_inc(x_88); +x_95 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___boxed), 31, 22); +lean_closure_set(x_95, 0, x_63); +lean_closure_set(x_95, 1, x_1); +lean_closure_set(x_95, 2, x_88); +lean_closure_set(x_95, 3, x_2); +lean_closure_set(x_95, 4, x_3); +lean_closure_set(x_95, 5, x_4); +lean_closure_set(x_95, 6, x_5); +lean_closure_set(x_95, 7, x_6); +lean_closure_set(x_95, 8, x_94); +lean_closure_set(x_95, 9, x_8); +lean_closure_set(x_95, 10, x_9); +lean_closure_set(x_95, 11, x_10); +lean_closure_set(x_95, 12, x_11); +lean_closure_set(x_95, 13, x_12); +lean_closure_set(x_95, 14, x_13); +lean_closure_set(x_95, 15, x_14); +lean_closure_set(x_95, 16, x_15); +lean_closure_set(x_95, 17, x_58); +lean_closure_set(x_95, 18, x_90); +lean_closure_set(x_95, 19, x_53); +lean_closure_set(x_95, 20, x_62); +lean_closure_set(x_95, 21, x_17); +x_96 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); +lean_closure_set(x_96, 0, x_93); +lean_closure_set(x_96, 1, x_95); +x_97 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(x_88, x_96, x_19, x_20, x_21, x_22, x_23, x_24, x_25, x_87); +return x_97; +} +} +else +{ +uint8_t x_98; +lean_dec(x_53); +lean_dec(x_51); +lean_dec(x_25); +lean_dec(x_24); +lean_dec(x_23); +lean_dec(x_22); +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_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_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_98 = !lean_is_exclusive(x_54); +if (x_98 == 0) +{ +return x_54; +} +else +{ +lean_object* x_99; lean_object* x_100; lean_object* x_101; +x_99 = lean_ctor_get(x_54, 0); +x_100 = lean_ctor_get(x_54, 1); +lean_inc(x_100); +lean_inc(x_99); +lean_dec(x_54); +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 +{ +uint8_t x_102; +lean_dec(x_25); +lean_dec(x_24); +lean_dec(x_23); +lean_dec(x_22); +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_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_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_102 = !lean_is_exclusive(x_50); +if (x_102 == 0) +{ +return x_50; +} +else +{ +lean_object* x_103; lean_object* x_104; lean_object* x_105; +x_103 = lean_ctor_get(x_50, 0); +x_104 = lean_ctor_get(x_50, 1); +lean_inc(x_104); +lean_inc(x_103); +lean_dec(x_50); +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_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; 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_1, 0); +lean_inc(x_10); +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 = lean_ctor_get(x_1, 3); +lean_inc(x_13); +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_uint8(x_1, sizeof(void*)*14); +x_17 = lean_ctor_get(x_1, 6); +lean_inc(x_17); +x_18 = lean_ctor_get(x_1, 7); +lean_inc(x_18); +x_19 = lean_ctor_get(x_1, 8); +lean_inc(x_19); +x_20 = lean_ctor_get(x_1, 9); +lean_inc(x_20); +x_21 = lean_ctor_get(x_1, 10); +lean_inc(x_21); +x_22 = lean_ctor_get(x_1, 11); +lean_inc(x_22); +x_23 = lean_ctor_get(x_1, 12); +lean_inc(x_23); +x_24 = lean_ctor_get(x_1, 13); +lean_inc(x_24); +lean_dec(x_1); +lean_inc(x_10); +x_25 = l_Lean_MVarId_getType(x_10, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_25) == 0) +{ +uint8_t x_26; +x_26 = !lean_is_exclusive(x_25); +if (x_26 == 0) +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; uint8_t x_46; +x_27 = lean_ctor_get(x_25, 0); +x_28 = lean_ctor_get(x_25, 1); +x_46 = l_Lean_Expr_isArrow(x_27); +if (x_46 == 0) +{ +uint8_t x_47; +x_47 = l_Lean_Expr_isLet(x_27); +if (x_47 == 0) +{ +uint8_t x_48; +x_48 = l_Lean_Expr_isForall(x_27); +lean_dec(x_27); +if (x_48 == 0) +{ +lean_object* x_49; +lean_dec(x_24); +lean_dec(x_23); +lean_dec(x_22); +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_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_3); +lean_dec(x_2); +x_49 = lean_box(0); +lean_ctor_set(x_25, 0, x_49); +return x_25; +} +else +{ +lean_object* x_50; +lean_free_object(x_25); +x_50 = lean_box(0); +x_29 = x_50; +goto block_45; +} +} +else +{ +lean_object* x_51; +lean_free_object(x_25); +lean_dec(x_27); +x_51 = lean_box(0); +x_29 = x_51; +goto block_45; +} +} +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_free_object(x_25); +x_52 = l_Lean_Expr_bindingDomain_x21(x_27); +lean_inc(x_52); +x_53 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__3___boxed), 9, 1); +lean_closure_set(x_53, 0, x_52); +x_54 = lean_box(x_16); +lean_inc(x_10); +x_55 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__6___boxed), 26, 17); +lean_closure_set(x_55, 0, x_10); +lean_closure_set(x_55, 1, x_11); +lean_closure_set(x_55, 2, x_12); +lean_closure_set(x_55, 3, x_13); +lean_closure_set(x_55, 4, x_14); +lean_closure_set(x_55, 5, x_15); +lean_closure_set(x_55, 6, x_54); +lean_closure_set(x_55, 7, x_17); +lean_closure_set(x_55, 8, x_18); +lean_closure_set(x_55, 9, x_19); +lean_closure_set(x_55, 10, x_20); +lean_closure_set(x_55, 11, x_21); +lean_closure_set(x_55, 12, x_22); +lean_closure_set(x_55, 13, x_23); +lean_closure_set(x_55, 14, x_24); +lean_closure_set(x_55, 15, x_27); +lean_closure_set(x_55, 16, x_52); +x_56 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); +lean_closure_set(x_56, 0, x_53); +lean_closure_set(x_56, 1, x_55); +x_57 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(x_10, x_56, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_28); +return x_57; +} +block_45: +{ +uint8_t x_30; lean_object* x_31; +lean_dec(x_29); +x_30 = 1; +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +x_31 = l_Lean_Meta_intro1Core(x_10, x_30, x_5, x_6, x_7, x_8, x_28); +if (lean_obj_tag(x_31) == 0) +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; +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_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); +lean_inc(x_34); +x_36 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__1___boxed), 9, 1); +lean_closure_set(x_36, 0, x_34); +x_37 = lean_box(x_16); +lean_inc(x_35); +x_38 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__2___boxed), 25, 16); +lean_closure_set(x_38, 0, x_35); +lean_closure_set(x_38, 1, x_11); +lean_closure_set(x_38, 2, x_12); +lean_closure_set(x_38, 3, x_13); +lean_closure_set(x_38, 4, x_14); +lean_closure_set(x_38, 5, x_15); +lean_closure_set(x_38, 6, x_37); +lean_closure_set(x_38, 7, x_17); +lean_closure_set(x_38, 8, x_18); +lean_closure_set(x_38, 9, x_19); +lean_closure_set(x_38, 10, x_20); +lean_closure_set(x_38, 11, x_21); +lean_closure_set(x_38, 12, x_22); +lean_closure_set(x_38, 13, x_23); +lean_closure_set(x_38, 14, x_24); +lean_closure_set(x_38, 15, x_34); +x_39 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); +lean_closure_set(x_39, 0, x_36); +lean_closure_set(x_39, 1, x_38); +x_40 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(x_35, x_39, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_33); +return x_40; +} +else +{ +uint8_t x_41; +lean_dec(x_24); +lean_dec(x_23); +lean_dec(x_22); +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_41 = !lean_is_exclusive(x_31); +if (x_41 == 0) +{ +return x_31; +} +else +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_42 = lean_ctor_get(x_31, 0); +x_43 = lean_ctor_get(x_31, 1); +lean_inc(x_43); +lean_inc(x_42); +lean_dec(x_31); +x_44 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_44, 0, x_42); +lean_ctor_set(x_44, 1, x_43); +return x_44; +} +} +} +} +else +{ +lean_object* x_58; lean_object* x_59; lean_object* x_60; uint8_t x_77; +x_58 = lean_ctor_get(x_25, 0); +x_59 = lean_ctor_get(x_25, 1); +lean_inc(x_59); +lean_inc(x_58); +lean_dec(x_25); +x_77 = l_Lean_Expr_isArrow(x_58); +if (x_77 == 0) +{ +uint8_t x_78; +x_78 = l_Lean_Expr_isLet(x_58); +if (x_78 == 0) +{ +uint8_t x_79; +x_79 = l_Lean_Expr_isForall(x_58); +lean_dec(x_58); +if (x_79 == 0) +{ +lean_object* x_80; lean_object* x_81; +lean_dec(x_24); +lean_dec(x_23); +lean_dec(x_22); +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_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_3); +lean_dec(x_2); +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_59); +return x_81; +} +else +{ +lean_object* x_82; +x_82 = lean_box(0); +x_60 = x_82; +goto block_76; +} +} +else +{ +lean_object* x_83; +lean_dec(x_58); +x_83 = lean_box(0); +x_60 = x_83; +goto block_76; +} +} +else +{ +lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; +x_84 = l_Lean_Expr_bindingDomain_x21(x_58); +lean_inc(x_84); +x_85 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__3___boxed), 9, 1); +lean_closure_set(x_85, 0, x_84); +x_86 = lean_box(x_16); +lean_inc(x_10); +x_87 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__6___boxed), 26, 17); +lean_closure_set(x_87, 0, x_10); +lean_closure_set(x_87, 1, x_11); +lean_closure_set(x_87, 2, x_12); +lean_closure_set(x_87, 3, x_13); +lean_closure_set(x_87, 4, x_14); +lean_closure_set(x_87, 5, x_15); +lean_closure_set(x_87, 6, x_86); +lean_closure_set(x_87, 7, x_17); +lean_closure_set(x_87, 8, x_18); +lean_closure_set(x_87, 9, x_19); +lean_closure_set(x_87, 10, x_20); +lean_closure_set(x_87, 11, x_21); +lean_closure_set(x_87, 12, x_22); +lean_closure_set(x_87, 13, x_23); +lean_closure_set(x_87, 14, x_24); +lean_closure_set(x_87, 15, x_58); +lean_closure_set(x_87, 16, x_84); +x_88 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); +lean_closure_set(x_88, 0, x_85); +lean_closure_set(x_88, 1, x_87); +x_89 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(x_10, x_88, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_59); +return x_89; +} +block_76: +{ +uint8_t x_61; lean_object* x_62; +lean_dec(x_60); +x_61 = 1; +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +x_62 = l_Lean_Meta_intro1Core(x_10, x_61, x_5, x_6, x_7, x_8, x_59); +if (lean_obj_tag(x_62) == 0) +{ +lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; +x_63 = lean_ctor_get(x_62, 0); +lean_inc(x_63); +x_64 = lean_ctor_get(x_62, 1); +lean_inc(x_64); +lean_dec(x_62); +x_65 = lean_ctor_get(x_63, 0); +lean_inc(x_65); +x_66 = lean_ctor_get(x_63, 1); +lean_inc(x_66); +lean_dec(x_63); +lean_inc(x_65); +x_67 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__1___boxed), 9, 1); +lean_closure_set(x_67, 0, x_65); +x_68 = lean_box(x_16); +lean_inc(x_66); +x_69 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__2___boxed), 25, 16); +lean_closure_set(x_69, 0, x_66); +lean_closure_set(x_69, 1, x_11); +lean_closure_set(x_69, 2, x_12); +lean_closure_set(x_69, 3, x_13); +lean_closure_set(x_69, 4, x_14); +lean_closure_set(x_69, 5, x_15); +lean_closure_set(x_69, 6, x_68); +lean_closure_set(x_69, 7, x_17); +lean_closure_set(x_69, 8, x_18); +lean_closure_set(x_69, 9, x_19); +lean_closure_set(x_69, 10, x_20); +lean_closure_set(x_69, 11, x_21); +lean_closure_set(x_69, 12, x_22); +lean_closure_set(x_69, 13, x_23); +lean_closure_set(x_69, 14, x_24); +lean_closure_set(x_69, 15, x_65); +x_70 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); +lean_closure_set(x_70, 0, x_67); +lean_closure_set(x_70, 1, x_69); +x_71 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(x_66, x_70, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_64); +return x_71; +} +else +{ +lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; +lean_dec(x_24); +lean_dec(x_23); +lean_dec(x_22); +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_72 = lean_ctor_get(x_62, 0); +lean_inc(x_72); +x_73 = lean_ctor_get(x_62, 1); +lean_inc(x_73); +if (lean_is_exclusive(x_62)) { + lean_ctor_release(x_62, 0); + lean_ctor_release(x_62, 1); + x_74 = x_62; +} else { + lean_dec_ref(x_62); + 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; +} +} +} +} +else +{ +uint8_t x_90; +lean_dec(x_24); +lean_dec(x_23); +lean_dec(x_22); +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_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_3); +lean_dec(x_2); +x_90 = !lean_is_exclusive(x_25); +if (x_90 == 0) +{ +return x_25; +} +else +{ +lean_object* x_91; lean_object* x_92; lean_object* x_93; +x_91 = lean_ctor_get(x_25, 0); +x_92 = lean_ctor_get(x_25, 1); +lean_inc(x_92); +lean_inc(x_91); +lean_dec(x_25); +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_EXPORT lean_object* l_Lean_mkFreshId___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___spec__2___rarg___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_mkFreshId___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___spec__2___rarg(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_mkFreshId___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; +x_7 = l_Lean_mkFreshId___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___spec__2(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_mkFreshFVarId___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___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_mkFreshFVarId___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___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); +lean_dec(x_1); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_MVarId_assign___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_MVarId_assign___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___spec__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_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___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___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_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___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_4); +lean_dec(x_3); +lean_dec(x_2); +return x_10; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__2___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +lean_object* x_18 = _args[17]; +lean_object* x_19 = _args[18]; +lean_object* x_20 = _args[19]; +lean_object* x_21 = _args[20]; +lean_object* x_22 = _args[21]; +lean_object* x_23 = _args[22]; +lean_object* x_24 = _args[23]; +lean_object* x_25 = _args[24]; +_start: +{ +uint8_t x_26; lean_object* x_27; +x_26 = lean_unbox(x_7); +lean_dec(x_7); +x_27 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_26, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_23, x_24, x_25); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +return x_27; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___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_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_10; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___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_11; +x_11 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___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); +lean_dec(x_3); +lean_dec(x_1); +return x_11; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +lean_object* x_18 = _args[17]; +lean_object* x_19 = _args[18]; +lean_object* x_20 = _args[19]; +lean_object* x_21 = _args[20]; +lean_object* x_22 = _args[21]; +lean_object* x_23 = _args[22]; +lean_object* x_24 = _args[23]; +lean_object* x_25 = _args[24]; +lean_object* x_26 = _args[25]; +lean_object* x_27 = _args[26]; +lean_object* x_28 = _args[27]; +lean_object* x_29 = _args[28]; +lean_object* x_30 = _args[29]; +lean_object* x_31 = _args[30]; +_start: +{ +uint8_t x_32; lean_object* x_33; +x_32 = lean_unbox(x_9); +lean_dec(x_9); +x_33 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_32, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_23, x_24, x_25, x_26, x_27, x_28, x_29, x_30, x_31); +lean_dec(x_30); +lean_dec(x_29); +lean_dec(x_28); +lean_dec(x_27); +lean_dec(x_26); +lean_dec(x_25); +lean_dec(x_24); +lean_dec(x_1); +return x_33; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__6___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +lean_object* x_18 = _args[17]; +lean_object* x_19 = _args[18]; +lean_object* x_20 = _args[19]; +lean_object* x_21 = _args[20]; +lean_object* x_22 = _args[21]; +lean_object* x_23 = _args[22]; +lean_object* x_24 = _args[23]; +lean_object* x_25 = _args[24]; +lean_object* x_26 = _args[25]; +_start: +{ +uint8_t x_27; uint8_t x_28; lean_object* x_29; +x_27 = lean_unbox(x_7); +lean_dec(x_7); +x_28 = lean_unbox(x_18); +lean_dec(x_18); +x_29 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__6(x_1, x_2, x_3, x_4, x_5, x_6, x_27, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_28, x_19, x_20, x_21, x_22, x_23, x_24, x_25, x_26); +lean_dec(x_16); +return x_29; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_isCasesCandidate(lean_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_FVarId_getType(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; lean_object* x_11; +x_9 = lean_ctor_get(x_7, 0); +x_10 = lean_ctor_get(x_7, 1); +x_11 = l_Lean_Expr_getAppFn(x_9); +lean_dec(x_9); +if (lean_obj_tag(x_11) == 4) +{ +lean_object* x_12; lean_object* x_13; +lean_free_object(x_7); +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +lean_dec(x_11); +x_13 = l_Lean_Meta_Grind_isGrindCasesTarget(x_12, x_4, x_5, x_10); +lean_dec(x_12); +return x_13; +} +else +{ +uint8_t x_14; lean_object* x_15; +lean_dec(x_11); +x_14 = 0; +x_15 = lean_box(x_14); +lean_ctor_set(x_7, 0, x_15); +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 = l_Lean_Expr_getAppFn(x_16); +lean_dec(x_16); +if (lean_obj_tag(x_18) == 4) +{ +lean_object* x_19; lean_object* x_20; +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +lean_dec(x_18); +x_20 = l_Lean_Meta_Grind_isGrindCasesTarget(x_19, x_4, x_5, x_17); +lean_dec(x_19); +return x_20; +} +else +{ +uint8_t x_21; lean_object* x_22; lean_object* x_23; +lean_dec(x_18); +x_21 = 0; +x_22 = lean_box(x_21); +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_17); +return x_23; +} +} +} +else +{ +uint8_t x_24; +x_24 = !lean_is_exclusive(x_7); +if (x_24 == 0) +{ +return x_7; +} +else +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_25 = lean_ctor_get(x_7, 0); +x_26 = lean_ctor_get(x_7, 1); +lean_inc(x_26); +lean_inc(x_25); +lean_dec(x_7); +x_27 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +return x_27; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_isCasesCandidate___boxed(lean_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_Tactic_Grind_Intro_0__Lean_Meta_Grind_isCasesCandidate(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_List_mapTR_loop___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_applyCases_x3f___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_4; +x_4 = l_List_reverse___rarg(x_3); +return x_4; +} +else +{ +uint8_t x_5; +x_5 = !lean_is_exclusive(x_2); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; 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_6 = lean_ctor_get(x_2, 0); +x_7 = lean_ctor_get(x_2, 1); +x_8 = lean_ctor_get(x_1, 1); +x_9 = lean_ctor_get(x_1, 2); +x_10 = lean_ctor_get(x_1, 3); +x_11 = lean_ctor_get(x_1, 4); +x_12 = lean_ctor_get(x_1, 5); +x_13 = lean_ctor_get_uint8(x_1, sizeof(void*)*14); +x_14 = lean_ctor_get(x_1, 6); +x_15 = lean_ctor_get(x_1, 7); +x_16 = lean_ctor_get(x_1, 8); +x_17 = lean_ctor_get(x_1, 9); +x_18 = lean_ctor_get(x_1, 10); +x_19 = lean_ctor_get(x_1, 11); +x_20 = lean_ctor_get(x_1, 12); +x_21 = lean_ctor_get(x_1, 13); +lean_inc(x_21); +lean_inc(x_20); +lean_inc(x_19); +lean_inc(x_18); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +x_22 = lean_alloc_ctor(0, 14, 1); +lean_ctor_set(x_22, 0, x_6); +lean_ctor_set(x_22, 1, x_8); +lean_ctor_set(x_22, 2, x_9); +lean_ctor_set(x_22, 3, x_10); +lean_ctor_set(x_22, 4, x_11); +lean_ctor_set(x_22, 5, x_12); +lean_ctor_set(x_22, 6, x_14); +lean_ctor_set(x_22, 7, x_15); +lean_ctor_set(x_22, 8, x_16); +lean_ctor_set(x_22, 9, x_17); +lean_ctor_set(x_22, 10, x_18); +lean_ctor_set(x_22, 11, x_19); +lean_ctor_set(x_22, 12, x_20); +lean_ctor_set(x_22, 13, x_21); +lean_ctor_set_uint8(x_22, sizeof(void*)*14, x_13); +lean_ctor_set(x_2, 1, x_3); +lean_ctor_set(x_2, 0, x_22); +{ +lean_object* _tmp_1 = x_7; +lean_object* _tmp_2 = x_2; +x_2 = _tmp_1; +x_3 = _tmp_2; +} +goto _start; +} +else +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_24 = lean_ctor_get(x_2, 0); +x_25 = lean_ctor_get(x_2, 1); +lean_inc(x_25); +lean_inc(x_24); +lean_dec(x_2); +x_26 = lean_ctor_get(x_1, 1); +x_27 = lean_ctor_get(x_1, 2); +x_28 = lean_ctor_get(x_1, 3); +x_29 = lean_ctor_get(x_1, 4); +x_30 = lean_ctor_get(x_1, 5); +x_31 = lean_ctor_get_uint8(x_1, sizeof(void*)*14); +x_32 = lean_ctor_get(x_1, 6); +x_33 = lean_ctor_get(x_1, 7); +x_34 = lean_ctor_get(x_1, 8); +x_35 = lean_ctor_get(x_1, 9); +x_36 = lean_ctor_get(x_1, 10); +x_37 = lean_ctor_get(x_1, 11); +x_38 = lean_ctor_get(x_1, 12); +x_39 = lean_ctor_get(x_1, 13); +lean_inc(x_39); +lean_inc(x_38); +lean_inc(x_37); +lean_inc(x_36); +lean_inc(x_35); +lean_inc(x_34); +lean_inc(x_33); +lean_inc(x_32); +lean_inc(x_30); +lean_inc(x_29); +lean_inc(x_28); +lean_inc(x_27); +lean_inc(x_26); +x_40 = lean_alloc_ctor(0, 14, 1); +lean_ctor_set(x_40, 0, x_24); +lean_ctor_set(x_40, 1, x_26); +lean_ctor_set(x_40, 2, x_27); +lean_ctor_set(x_40, 3, x_28); +lean_ctor_set(x_40, 4, x_29); +lean_ctor_set(x_40, 5, x_30); +lean_ctor_set(x_40, 6, x_32); +lean_ctor_set(x_40, 7, x_33); +lean_ctor_set(x_40, 8, x_34); +lean_ctor_set(x_40, 9, x_35); +lean_ctor_set(x_40, 10, x_36); +lean_ctor_set(x_40, 11, x_37); +lean_ctor_set(x_40, 12, x_38); +lean_ctor_set(x_40, 13, x_39); +lean_ctor_set_uint8(x_40, sizeof(void*)*14, x_31); +x_41 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_41, 0, x_40); +lean_ctor_set(x_41, 1, x_3); +x_2 = x_25; +x_3 = x_41; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_applyCases_x3f___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; +lean_inc(x_4); +lean_inc(x_1); +x_9 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_isCasesCandidate(x_1, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_9) == 0) +{ +lean_object* x_10; uint8_t x_11; +x_10 = lean_ctor_get(x_9, 0); +lean_inc(x_10); +x_11 = lean_unbox(x_10); +lean_dec(x_10); +if (x_11 == 0) +{ +uint8_t x_12; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_12 = !lean_is_exclusive(x_9); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; +x_13 = lean_ctor_get(x_9, 0); +lean_dec(x_13); +x_14 = lean_box(0); +lean_ctor_set(x_9, 0, x_14); +return x_9; +} +else +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_ctor_get(x_9, 1); +lean_inc(x_15); +lean_dec(x_9); +x_16 = lean_box(0); +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_16); +lean_ctor_set(x_17, 1, x_15); +return x_17; +} +} +else +{ +lean_object* x_18; lean_object* x_19; +x_18 = lean_ctor_get(x_9, 1); +lean_inc(x_18); +lean_dec(x_9); +x_19 = l_Lean_Meta_Grind_cases(x_2, x_1, x_4, x_5, x_6, x_7, x_18); +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; lean_object* x_23; lean_object* x_24; +x_21 = lean_ctor_get(x_19, 0); +x_22 = lean_box(0); +x_23 = l_List_mapTR_loop___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_applyCases_x3f___spec__1(x_3, x_21, x_22); +x_24 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set(x_19, 0, x_24); +return x_19; +} +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_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_box(0); +x_28 = l_List_mapTR_loop___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_applyCases_x3f___spec__1(x_3, x_25, x_27); +x_29 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_29, 0, 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_26); +return x_30; +} +} +else +{ +uint8_t x_31; +x_31 = !lean_is_exclusive(x_19); +if (x_31 == 0) +{ +return x_19; +} +else +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_19, 0); +x_33 = lean_ctor_get(x_19, 1); +lean_inc(x_33); +lean_inc(x_32); +lean_dec(x_19); +x_34 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +return x_34; +} +} +} +} +else +{ +uint8_t x_35; +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_35 = !lean_is_exclusive(x_9); +if (x_35 == 0) +{ +return x_9; +} +else +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_9, 0); +x_37 = lean_ctor_get(x_9, 1); +lean_inc(x_37); +lean_inc(x_36); +lean_dec(x_9); +x_38 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_38, 0, x_36); +lean_ctor_set(x_38, 1, x_37); +return x_38; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_applyCases_x3f(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_8 = lean_ctor_get(x_1, 0); +lean_inc(x_8); +lean_inc(x_8); +x_9 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_applyCases_x3f___lambda__1___boxed), 8, 3); +lean_closure_set(x_9, 0, x_2); +lean_closure_set(x_9, 1, x_8); +lean_closure_set(x_9, 2, x_1); +x_10 = l_Lean_MVarId_withContext___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___spec__2___rarg(x_8, x_9, x_3, x_4, x_5, x_6, x_7); +return x_10; +} +} +LEAN_EXPORT lean_object* l_List_mapTR_loop___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_applyCases_x3f___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_List_mapTR_loop___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_applyCases_x3f___spec__1(x_1, x_2, x_3); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_applyCases_x3f___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; +x_9 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_applyCases_x3f___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___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_applyInjection_x3f(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +uint8_t x_8; +x_8 = !lean_is_exclusive(x_1); +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; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_9 = lean_ctor_get(x_1, 0); +x_10 = lean_ctor_get(x_1, 1); +x_11 = lean_ctor_get(x_1, 2); +x_12 = lean_ctor_get(x_1, 3); +x_13 = lean_ctor_get(x_1, 4); +x_14 = lean_ctor_get(x_1, 5); +x_15 = lean_ctor_get(x_1, 6); +x_16 = lean_ctor_get(x_1, 7); +x_17 = lean_ctor_get(x_1, 8); +x_18 = lean_ctor_get(x_1, 9); +x_19 = lean_ctor_get(x_1, 10); +x_20 = lean_ctor_get(x_1, 11); +x_21 = lean_ctor_get(x_1, 12); +x_22 = lean_ctor_get(x_1, 13); +x_23 = l_Lean_Meta_Grind_injection_x3f(x_9, x_2, x_3, x_4, x_5, x_6, x_7); +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) +{ +uint8_t x_25; +lean_free_object(x_1); +lean_dec(x_22); +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +x_25 = !lean_is_exclusive(x_23); +if (x_25 == 0) +{ +lean_object* x_26; lean_object* x_27; +x_26 = lean_ctor_get(x_23, 0); +lean_dec(x_26); +x_27 = lean_box(0); +lean_ctor_set(x_23, 0, x_27); +return x_23; +} +else +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_28 = lean_ctor_get(x_23, 1); +lean_inc(x_28); +lean_dec(x_23); +x_29 = lean_box(0); +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_28); +return x_30; +} +} +else +{ +uint8_t x_31; +x_31 = !lean_is_exclusive(x_23); +if (x_31 == 0) +{ +lean_object* x_32; uint8_t x_33; +x_32 = lean_ctor_get(x_23, 0); +lean_dec(x_32); +x_33 = !lean_is_exclusive(x_24); +if (x_33 == 0) +{ +lean_object* x_34; +x_34 = lean_ctor_get(x_24, 0); +lean_ctor_set(x_1, 0, x_34); +lean_ctor_set(x_24, 0, x_1); +return x_23; +} +else +{ +lean_object* x_35; lean_object* x_36; +x_35 = lean_ctor_get(x_24, 0); +lean_inc(x_35); +lean_dec(x_24); +lean_ctor_set(x_1, 0, x_35); +x_36 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_36, 0, x_1); +lean_ctor_set(x_23, 0, x_36); +return x_23; +} +} +else +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_37 = lean_ctor_get(x_23, 1); +lean_inc(x_37); +lean_dec(x_23); +x_38 = lean_ctor_get(x_24, 0); +lean_inc(x_38); +if (lean_is_exclusive(x_24)) { + lean_ctor_release(x_24, 0); + x_39 = x_24; +} else { + lean_dec_ref(x_24); + x_39 = lean_box(0); +} +lean_ctor_set(x_1, 0, x_38); +if (lean_is_scalar(x_39)) { + x_40 = lean_alloc_ctor(1, 1, 0); +} else { + x_40 = x_39; +} +lean_ctor_set(x_40, 0, x_1); +x_41 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_41, 0, x_40); +lean_ctor_set(x_41, 1, x_37); +return x_41; +} +} +} +else +{ +uint8_t x_42; +lean_free_object(x_1); +lean_dec(x_22); +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +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; 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; +x_46 = lean_ctor_get(x_1, 0); +x_47 = lean_ctor_get(x_1, 1); +x_48 = lean_ctor_get(x_1, 2); +x_49 = lean_ctor_get(x_1, 3); +x_50 = lean_ctor_get(x_1, 4); +x_51 = lean_ctor_get(x_1, 5); +x_52 = lean_ctor_get_uint8(x_1, sizeof(void*)*14); +x_53 = lean_ctor_get(x_1, 6); +x_54 = lean_ctor_get(x_1, 7); +x_55 = lean_ctor_get(x_1, 8); +x_56 = lean_ctor_get(x_1, 9); +x_57 = lean_ctor_get(x_1, 10); +x_58 = lean_ctor_get(x_1, 11); +x_59 = lean_ctor_get(x_1, 12); +x_60 = lean_ctor_get(x_1, 13); +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_51); +lean_inc(x_50); +lean_inc(x_49); +lean_inc(x_48); +lean_inc(x_47); +lean_inc(x_46); +lean_dec(x_1); +x_61 = l_Lean_Meta_Grind_injection_x3f(x_46, x_2, x_3, x_4, x_5, x_6, x_7); +if (lean_obj_tag(x_61) == 0) +{ +lean_object* x_62; +x_62 = lean_ctor_get(x_61, 0); +lean_inc(x_62); +if (lean_obj_tag(x_62) == 0) +{ +lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; +lean_dec(x_60); +lean_dec(x_59); +lean_dec(x_58); +lean_dec(x_57); +lean_dec(x_56); +lean_dec(x_55); +lean_dec(x_54); +lean_dec(x_53); +lean_dec(x_51); +lean_dec(x_50); +lean_dec(x_49); +lean_dec(x_48); +lean_dec(x_47); +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); +} +x_65 = lean_box(0); +if (lean_is_scalar(x_64)) { + x_66 = lean_alloc_ctor(0, 2, 0); +} else { + x_66 = x_64; +} +lean_ctor_set(x_66, 0, x_65); +lean_ctor_set(x_66, 1, x_63); +return x_66; +} +else +{ +lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; +x_67 = lean_ctor_get(x_61, 1); +lean_inc(x_67); +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_61); + x_68 = lean_box(0); +} +x_69 = lean_ctor_get(x_62, 0); +lean_inc(x_69); +if (lean_is_exclusive(x_62)) { + lean_ctor_release(x_62, 0); + x_70 = x_62; +} else { + lean_dec_ref(x_62); + x_70 = lean_box(0); +} +x_71 = lean_alloc_ctor(0, 14, 1); +lean_ctor_set(x_71, 0, x_69); +lean_ctor_set(x_71, 1, x_47); +lean_ctor_set(x_71, 2, x_48); +lean_ctor_set(x_71, 3, x_49); +lean_ctor_set(x_71, 4, x_50); +lean_ctor_set(x_71, 5, x_51); +lean_ctor_set(x_71, 6, x_53); +lean_ctor_set(x_71, 7, x_54); +lean_ctor_set(x_71, 8, x_55); +lean_ctor_set(x_71, 9, x_56); +lean_ctor_set(x_71, 10, x_57); +lean_ctor_set(x_71, 11, x_58); +lean_ctor_set(x_71, 12, x_59); +lean_ctor_set(x_71, 13, x_60); +lean_ctor_set_uint8(x_71, sizeof(void*)*14, x_52); +if (lean_is_scalar(x_70)) { + x_72 = lean_alloc_ctor(1, 1, 0); +} else { + x_72 = x_70; +} +lean_ctor_set(x_72, 0, x_71); +if (lean_is_scalar(x_68)) { + x_73 = lean_alloc_ctor(0, 2, 0); +} else { + x_73 = x_68; +} +lean_ctor_set(x_73, 0, x_72); +lean_ctor_set(x_73, 1, x_67); +return x_73; +} +} +else +{ +lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; +lean_dec(x_60); +lean_dec(x_59); +lean_dec(x_58); +lean_dec(x_57); +lean_dec(x_56); +lean_dec(x_55); +lean_dec(x_54); +lean_dec(x_53); +lean_dec(x_51); +lean_dec(x_50); +lean_dec(x_49); +lean_dec(x_48); +lean_dec(x_47); +x_74 = lean_ctor_get(x_61, 0); +lean_inc(x_74); +x_75 = lean_ctor_get(x_61, 1); +lean_inc(x_75); +if (lean_is_exclusive(x_61)) { + lean_ctor_release(x_61, 0); + lean_ctor_release(x_61, 1); + x_76 = x_61; +} else { + lean_dec_ref(x_61); + x_76 = lean_box(0); +} +if (lean_is_scalar(x_76)) { + x_77 = lean_alloc_ctor(1, 2, 0); +} else { + x_77 = x_76; +} +lean_ctor_set(x_77, 0, x_74); +lean_ctor_set(x_77, 1, x_75); +return x_77; +} +} +} +} +LEAN_EXPORT lean_object* l_List_forM___at_Lean_Meta_Grind_intros_go___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_12; lean_object* x_13; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_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_11); +return x_13; +} +else +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; +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); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_1); +x_16 = l_Lean_Meta_Grind_intros_go(x_1, x_14, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_17; +x_17 = lean_ctor_get(x_16, 1); +lean_inc(x_17); +lean_dec(x_16); +x_2 = x_15; +x_11 = x_17; +goto _start; +} +else +{ +uint8_t x_19; +lean_dec(x_15); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_19 = !lean_is_exclusive(x_16); +if (x_19 == 0) +{ +return x_16; +} +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_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_EXPORT lean_object* l_List_forM___at_Lean_Meta_Grind_intros_go___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_12; lean_object* x_13; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_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_11); +return x_13; +} +else +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; +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); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_1); +x_16 = l_Lean_Meta_Grind_intros_go(x_1, x_14, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_17; +x_17 = lean_ctor_get(x_16, 1); +lean_inc(x_17); +lean_dec(x_16); +x_2 = x_15; +x_11 = x_17; +goto _start; +} +else +{ +uint8_t x_19; +lean_dec(x_15); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_19 = !lean_is_exclusive(x_16); +if (x_19 == 0) +{ +return x_16; +} +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_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_EXPORT lean_object* l_Lean_Meta_Grind_intros_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) { +_start: +{ +lean_object* x_10; uint8_t x_11; +x_10 = lean_st_mk_ref(x_1, x_9); +x_11 = !lean_is_exclusive(x_10); +if (x_11 == 0) +{ +return x_10; +} +else +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_12 = lean_ctor_get(x_10, 0); +x_13 = lean_ctor_get(x_10, 1); +lean_inc(x_13); +lean_inc(x_12); +lean_dec(x_10); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_12); +lean_ctor_set(x_14, 1, x_13); +return x_14; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_intros_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_object* x_11) { +_start: +{ +lean_object* x_12; +lean_inc(x_3); +x_12 = l_Lean_Meta_Grind_addHypothesis(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_12) == 0) +{ +lean_object* x_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_3, x_13); +x_15 = !lean_is_exclusive(x_14); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_16 = lean_ctor_get(x_14, 1); +x_17 = lean_st_ref_get(x_3, x_16); +lean_dec(x_3); +x_18 = !lean_is_exclusive(x_17); +if (x_18 == 0) +{ +lean_object* x_19; +x_19 = lean_ctor_get(x_17, 0); +lean_ctor_set(x_14, 1, x_19); +lean_ctor_set(x_17, 0, x_14); +return x_17; +} +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_17, 0); +x_21 = lean_ctor_get(x_17, 1); +lean_inc(x_21); +lean_inc(x_20); +lean_dec(x_17); +lean_ctor_set(x_14, 1, x_20); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_14); +lean_ctor_set(x_22, 1, x_21); +return x_22; +} +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_23 = lean_ctor_get(x_14, 0); +x_24 = lean_ctor_get(x_14, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_14); +x_25 = lean_st_ref_get(x_3, x_24); +lean_dec(x_3); +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); +} +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_23); +lean_ctor_set(x_29, 1, x_26); +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; +} +} +else +{ +uint8_t x_31; +lean_dec(x_3); +x_31 = !lean_is_exclusive(x_12); +if (x_31 == 0) +{ +return x_12; +} +else +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_12, 0); +x_33 = lean_ctor_get(x_12, 1); +lean_inc(x_33); +lean_inc(x_32); +lean_dec(x_12); +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_EXPORT lean_object* l_Lean_Meta_Grind_intros_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: +{ +uint8_t x_10; +x_10 = !lean_is_exclusive(x_1); +if (x_10 == 0) +{ +lean_object* x_11; +x_11 = lean_ctor_get(x_1, 1); +lean_dec(x_11); +lean_ctor_set(x_1, 1, x_9); +return x_1; +} +else +{ +lean_object* x_12; lean_object* x_13; +x_12 = lean_ctor_get(x_1, 0); +lean_inc(x_12); +lean_dec(x_1); +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_9); +return x_13; +} +} +} +static lean_object* _init_l_Lean_Meta_Grind_intros_go___lambda__4___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_intros_go___lambda__3___boxed), 9, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_intros_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, 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) { +_start: +{ +lean_object* x_28; +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_1); +x_28 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext(x_1, x_20, x_21, x_22, x_23, x_24, x_25, x_26, x_27); +if (lean_obj_tag(x_28) == 0) +{ +lean_object* x_29; +x_29 = lean_ctor_get(x_28, 0); +lean_inc(x_29); +switch (lean_obj_tag(x_29)) { +case 0: +{ +lean_object* x_30; lean_object* x_31; +x_30 = lean_ctor_get(x_28, 1); +lean_inc(x_30); +lean_dec(x_28); +lean_inc(x_26); +lean_inc(x_25); +lean_inc(x_24); +lean_inc(x_23); +x_31 = l_Lean_MVarId_byContra_x3f(x_2, x_23, x_24, x_25, x_26, x_30); +if (lean_obj_tag(x_31) == 0) +{ +lean_object* x_32; +x_32 = lean_ctor_get(x_31, 0); +lean_inc(x_32); +if (lean_obj_tag(x_32) == 0) +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; uint8_t x_39; +lean_dec(x_26); +lean_dec(x_25); +lean_dec(x_24); +lean_dec(x_23); +lean_dec(x_22); +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +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); +x_33 = lean_ctor_get(x_31, 1); +lean_inc(x_33); +lean_dec(x_31); +x_34 = lean_st_ref_take(x_19, x_33); +x_35 = lean_ctor_get(x_34, 0); +lean_inc(x_35); +x_36 = lean_ctor_get(x_34, 1); +lean_inc(x_36); +lean_dec(x_34); +x_37 = lean_array_push(x_35, x_1); +x_38 = lean_st_ref_set(x_19, x_37, x_36); +x_39 = !lean_is_exclusive(x_38); +if (x_39 == 0) +{ +lean_object* x_40; lean_object* x_41; +x_40 = lean_ctor_get(x_38, 0); +lean_dec(x_40); +x_41 = lean_box(0); +lean_ctor_set(x_38, 0, x_41); +return x_38; +} +else +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_42 = lean_ctor_get(x_38, 1); +lean_inc(x_42); +lean_dec(x_38); +x_43 = lean_box(0); +x_44 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_44, 0, x_43); +lean_ctor_set(x_44, 1, x_42); +return x_44; +} +} +else +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; +lean_dec(x_1); +x_45 = lean_ctor_get(x_31, 1); +lean_inc(x_45); +lean_dec(x_31); +x_46 = lean_ctor_get(x_32, 0); +lean_inc(x_46); +lean_dec(x_32); +x_47 = lean_alloc_ctor(0, 14, 1); +lean_ctor_set(x_47, 0, x_46); +lean_ctor_set(x_47, 1, x_3); +lean_ctor_set(x_47, 2, x_4); +lean_ctor_set(x_47, 3, x_5); +lean_ctor_set(x_47, 4, x_6); +lean_ctor_set(x_47, 5, x_7); +lean_ctor_set(x_47, 6, x_9); +lean_ctor_set(x_47, 7, x_10); +lean_ctor_set(x_47, 8, x_11); +lean_ctor_set(x_47, 9, x_12); +lean_ctor_set(x_47, 10, x_13); +lean_ctor_set(x_47, 11, x_14); +lean_ctor_set(x_47, 12, x_15); +lean_ctor_set(x_47, 13, x_16); +lean_ctor_set_uint8(x_47, sizeof(void*)*14, x_8); +x_48 = l_Lean_Meta_Grind_intros_go(x_17, x_47, x_19, x_20, x_21, x_22, x_23, x_24, x_25, x_26, x_45); +return x_48; +} +} +else +{ +uint8_t x_49; +lean_dec(x_26); +lean_dec(x_25); +lean_dec(x_24); +lean_dec(x_23); +lean_dec(x_22); +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +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_1); +x_49 = !lean_is_exclusive(x_31); +if (x_49 == 0) +{ +return x_31; +} +else +{ +lean_object* x_50; lean_object* x_51; lean_object* x_52; +x_50 = lean_ctor_get(x_31, 0); +x_51 = lean_ctor_get(x_31, 1); +lean_inc(x_51); +lean_inc(x_50); +lean_dec(x_31); +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; +} +} +} +case 1: +{ +lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_53 = lean_ctor_get(x_28, 1); +lean_inc(x_53); +lean_dec(x_28); +x_54 = lean_ctor_get(x_29, 0); +lean_inc(x_54); +x_55 = lean_ctor_get(x_29, 1); +lean_inc(x_55); +lean_dec(x_29); +lean_inc(x_26); +lean_inc(x_25); +lean_inc(x_24); +lean_inc(x_23); +lean_inc(x_54); +lean_inc(x_55); +x_56 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_applyCases_x3f(x_55, x_54, x_23, x_24, x_25, x_26, x_53); +if (lean_obj_tag(x_56) == 0) +{ +lean_object* x_57; +x_57 = lean_ctor_get(x_56, 0); +lean_inc(x_57); +if (lean_obj_tag(x_57) == 0) +{ +lean_object* x_58; lean_object* x_59; +x_58 = lean_ctor_get(x_56, 1); +lean_inc(x_58); +lean_dec(x_56); +lean_inc(x_26); +lean_inc(x_25); +lean_inc(x_24); +lean_inc(x_23); +lean_inc(x_54); +lean_inc(x_55); +x_59 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_applyInjection_x3f(x_55, x_54, x_23, x_24, x_25, x_26, x_58); +if (lean_obj_tag(x_59) == 0) +{ +lean_object* x_60; +x_60 = lean_ctor_get(x_59, 0); +lean_inc(x_60); +if (lean_obj_tag(x_60) == 0) +{ +lean_object* x_61; lean_object* x_62; 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(x_59, 1); +lean_inc(x_61); +lean_dec(x_59); +x_62 = lean_ctor_get(x_55, 0); +lean_inc(x_62); +x_63 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_intros_go___lambda__1___boxed), 9, 1); +lean_closure_set(x_63, 0, x_55); +lean_inc(x_17); +x_64 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_intros_go___lambda__2), 11, 2); +lean_closure_set(x_64, 0, x_54); +lean_closure_set(x_64, 1, x_17); +x_65 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); +lean_closure_set(x_65, 0, x_63); +lean_closure_set(x_65, 1, x_64); +x_66 = l_Lean_Meta_Grind_intros_go___lambda__4___closed__1; +x_67 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); +lean_closure_set(x_67, 0, x_65); +lean_closure_set(x_67, 1, x_66); +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); +x_68 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(x_62, x_67, x_20, x_21, x_22, x_23, x_24, x_25, x_26, x_61); +if (lean_obj_tag(x_68) == 0) +{ +lean_object* x_69; lean_object* x_70; lean_object* x_71; +x_69 = lean_ctor_get(x_68, 0); +lean_inc(x_69); +x_70 = lean_ctor_get(x_68, 1); +lean_inc(x_70); +lean_dec(x_68); +x_71 = l_Lean_Meta_Grind_intros_go(x_17, x_69, x_19, x_20, x_21, x_22, x_23, x_24, x_25, x_26, x_70); +return x_71; +} +else +{ +uint8_t x_72; +lean_dec(x_26); +lean_dec(x_25); +lean_dec(x_24); +lean_dec(x_23); +lean_dec(x_22); +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_17); +x_72 = !lean_is_exclusive(x_68); +if (x_72 == 0) +{ +return x_68; +} +else +{ +lean_object* x_73; lean_object* x_74; lean_object* x_75; +x_73 = lean_ctor_get(x_68, 0); +x_74 = lean_ctor_get(x_68, 1); +lean_inc(x_74); +lean_inc(x_73); +lean_dec(x_68); +x_75 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_75, 0, x_73); +lean_ctor_set(x_75, 1, x_74); +return x_75; +} +} +} +else +{ +lean_object* x_76; lean_object* x_77; lean_object* x_78; +lean_dec(x_55); +lean_dec(x_54); +x_76 = lean_ctor_get(x_59, 1); +lean_inc(x_76); +lean_dec(x_59); +x_77 = lean_ctor_get(x_60, 0); +lean_inc(x_77); +lean_dec(x_60); +x_78 = l_Lean_Meta_Grind_intros_go(x_17, x_77, x_19, x_20, x_21, x_22, x_23, x_24, x_25, x_26, x_76); +return x_78; +} +} +else +{ +uint8_t x_79; +lean_dec(x_55); +lean_dec(x_54); +lean_dec(x_26); +lean_dec(x_25); +lean_dec(x_24); +lean_dec(x_23); +lean_dec(x_22); +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_17); +x_79 = !lean_is_exclusive(x_59); +if (x_79 == 0) +{ +return x_59; +} +else +{ +lean_object* x_80; lean_object* x_81; lean_object* x_82; +x_80 = lean_ctor_get(x_59, 0); +x_81 = lean_ctor_get(x_59, 1); +lean_inc(x_81); +lean_inc(x_80); +lean_dec(x_59); +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 +{ +lean_object* x_83; lean_object* x_84; lean_object* x_85; +lean_dec(x_55); +lean_dec(x_54); +x_83 = lean_ctor_get(x_56, 1); +lean_inc(x_83); +lean_dec(x_56); +x_84 = lean_ctor_get(x_57, 0); +lean_inc(x_84); +lean_dec(x_57); +x_85 = l_List_forM___at_Lean_Meta_Grind_intros_go___spec__1(x_17, x_84, x_19, x_20, x_21, x_22, x_23, x_24, x_25, x_26, x_83); +return x_85; +} +} +else +{ +uint8_t x_86; +lean_dec(x_55); +lean_dec(x_54); +lean_dec(x_26); +lean_dec(x_25); +lean_dec(x_24); +lean_dec(x_23); +lean_dec(x_22); +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_17); +x_86 = !lean_is_exclusive(x_56); +if (x_86 == 0) +{ +return x_56; +} +else +{ +lean_object* x_87; lean_object* x_88; lean_object* x_89; +x_87 = lean_ctor_get(x_56, 0); +x_88 = lean_ctor_get(x_56, 1); +lean_inc(x_88); +lean_inc(x_87); +lean_dec(x_56); +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; +} +} +} +case 2: +{ +lean_object* x_90; lean_object* x_91; lean_object* x_92; +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_90 = lean_ctor_get(x_28, 1); +lean_inc(x_90); +lean_dec(x_28); +x_91 = lean_ctor_get(x_29, 0); +lean_inc(x_91); +lean_dec(x_29); +x_92 = l_Lean_Meta_Grind_intros_go(x_17, x_91, x_19, x_20, x_21, x_22, x_23, x_24, x_25, x_26, x_90); +return x_92; +} +default: +{ +lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_93 = lean_ctor_get(x_28, 1); +lean_inc(x_93); +lean_dec(x_28); +x_94 = lean_ctor_get(x_29, 0); +lean_inc(x_94); +x_95 = lean_ctor_get(x_29, 1); +lean_inc(x_95); +lean_dec(x_29); +lean_inc(x_26); +lean_inc(x_25); +lean_inc(x_24); +lean_inc(x_23); +lean_inc(x_95); +x_96 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_applyCases_x3f(x_95, x_94, x_23, x_24, x_25, x_26, x_93); +if (lean_obj_tag(x_96) == 0) +{ +lean_object* x_97; +x_97 = lean_ctor_get(x_96, 0); +lean_inc(x_97); +if (lean_obj_tag(x_97) == 0) +{ +lean_object* x_98; lean_object* x_99; +x_98 = lean_ctor_get(x_96, 1); +lean_inc(x_98); +lean_dec(x_96); +x_99 = l_Lean_Meta_Grind_intros_go(x_17, x_95, x_19, x_20, x_21, x_22, x_23, x_24, x_25, x_26, x_98); +return x_99; +} +else +{ +lean_object* x_100; lean_object* x_101; lean_object* x_102; +lean_dec(x_95); +x_100 = lean_ctor_get(x_96, 1); +lean_inc(x_100); +lean_dec(x_96); +x_101 = lean_ctor_get(x_97, 0); +lean_inc(x_101); +lean_dec(x_97); +x_102 = l_List_forM___at_Lean_Meta_Grind_intros_go___spec__2(x_17, x_101, x_19, x_20, x_21, x_22, x_23, x_24, x_25, x_26, x_100); +return x_102; +} +} +else +{ +uint8_t x_103; +lean_dec(x_95); +lean_dec(x_26); +lean_dec(x_25); +lean_dec(x_24); +lean_dec(x_23); +lean_dec(x_22); +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_17); +x_103 = !lean_is_exclusive(x_96); +if (x_103 == 0) +{ +return x_96; +} +else +{ +lean_object* x_104; lean_object* x_105; lean_object* x_106; +x_104 = lean_ctor_get(x_96, 0); +x_105 = lean_ctor_get(x_96, 1); +lean_inc(x_105); +lean_inc(x_104); +lean_dec(x_96); +x_106 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_106, 0, x_104); +lean_ctor_set(x_106, 1, x_105); +return x_106; +} +} +} +} +} +else +{ +uint8_t x_107; +lean_dec(x_26); +lean_dec(x_25); +lean_dec(x_24); +lean_dec(x_23); +lean_dec(x_22); +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +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); +lean_dec(x_1); +x_107 = !lean_is_exclusive(x_28); +if (x_107 == 0) +{ +return x_28; +} +else +{ +lean_object* x_108; lean_object* x_109; lean_object* x_110; +x_108 = lean_ctor_get(x_28, 0); +x_109 = lean_ctor_get(x_28, 1); +lean_inc(x_109); +lean_inc(x_108); +lean_dec(x_28); +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_EXPORT lean_object* l_Lean_Meta_Grind_intros_go(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +uint8_t x_12; +x_12 = lean_ctor_get_uint8(x_2, sizeof(void*)*14); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_13 = lean_ctor_get(x_2, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_2, 1); +lean_inc(x_14); +x_15 = lean_ctor_get(x_2, 2); +lean_inc(x_15); +x_16 = lean_ctor_get(x_2, 3); +lean_inc(x_16); +x_17 = lean_ctor_get(x_2, 4); +lean_inc(x_17); +x_18 = lean_ctor_get(x_2, 5); +lean_inc(x_18); +x_19 = lean_ctor_get(x_2, 6); +lean_inc(x_19); +x_20 = lean_ctor_get(x_2, 7); +lean_inc(x_20); +x_21 = lean_ctor_get(x_2, 8); +lean_inc(x_21); +x_22 = lean_ctor_get(x_2, 9); +lean_inc(x_22); +x_23 = lean_ctor_get(x_2, 10); +lean_inc(x_23); +x_24 = lean_ctor_get(x_2, 11); +lean_inc(x_24); +x_25 = lean_ctor_get(x_2, 12); +lean_inc(x_25); +x_26 = lean_ctor_get(x_2, 13); +lean_inc(x_26); +x_27 = lean_box(0); +x_28 = l_Lean_Meta_Grind_intros_go___lambda__4(x_2, x_13, x_14, x_15, x_16, x_17, x_18, x_12, x_19, x_20, x_21, x_22, x_23, x_24, x_25, x_26, x_1, x_27, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +return x_28; +} +else +{ +lean_object* x_29; lean_object* x_30; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_29 = lean_box(0); +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_11); +return x_30; +} +} +} +LEAN_EXPORT lean_object* l_List_forM___at_Lean_Meta_Grind_intros_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, 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_List_forM___at_Lean_Meta_Grind_intros_go___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_3); +return x_12; +} +} +LEAN_EXPORT lean_object* l_List_forM___at_Lean_Meta_Grind_intros_go___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; +x_12 = l_List_forM___at_Lean_Meta_Grind_intros_go___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_3); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_intros_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_Meta_Grind_intros_go___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_intros_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_Lean_Meta_Grind_intros_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); +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_Meta_Grind_intros_go___lambda__4___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_object* x_19 = _args[18]; +lean_object* x_20 = _args[19]; +lean_object* x_21 = _args[20]; +lean_object* x_22 = _args[21]; +lean_object* x_23 = _args[22]; +lean_object* x_24 = _args[23]; +lean_object* x_25 = _args[24]; +lean_object* x_26 = _args[25]; +lean_object* x_27 = _args[26]; +_start: +{ +uint8_t x_28; lean_object* x_29; +x_28 = lean_unbox(x_8); +lean_dec(x_8); +x_29 = l_Lean_Meta_Grind_intros_go___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_28, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_23, x_24, x_25, x_26, x_27); +lean_dec(x_19); +lean_dec(x_18); +return x_29; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_intros_go___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; +x_12 = l_Lean_Meta_Grind_intros_go(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_3); +return x_12; +} +} +static lean_object* _init_l_Lean_Meta_Grind_intros___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(0); +x_2 = lean_array_mk(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_intros(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_11 = l_Lean_Meta_Grind_intros___closed__1; +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_inc(x_14); +lean_dec(x_12); +x_15 = l_Lean_Meta_Grind_intros_go(x_2, x_1, x_13, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_14); +if (lean_obj_tag(x_15) == 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 = lean_st_ref_get(x_13, x_16); +lean_dec(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); +x_20 = lean_array_to_list(x_19); +lean_ctor_set(x_17, 0, x_20); +return x_17; +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_21 = lean_ctor_get(x_17, 0); +x_22 = lean_ctor_get(x_17, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_17); +x_23 = lean_array_to_list(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; +} +} +else +{ +uint8_t x_25; +lean_dec(x_13); +x_25 = !lean_is_exclusive(x_15); +if (x_25 == 0) +{ +return x_15; +} +else +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_26 = lean_ctor_get(x_15, 0); +x_27 = lean_ctor_get(x_15, 1); +lean_inc(x_27); +lean_inc(x_26); +lean_dec(x_15); +x_28 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_28, 0, x_26); +lean_ctor_set(x_28, 1, x_27); +return x_28; +} +} +} +} +static lean_object* _init_l_Lean_Meta_Grind_assertAt___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("h", 1, 1); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_assertAt___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_Grind_assertAt___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_assertAt(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, 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; +x_13 = l_Lean_Meta_Grind_assertAt___closed__2; +x_14 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_13, x_10, x_11, x_12); +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +x_16 = lean_ctor_get(x_14, 1); +lean_inc(x_16); +lean_dec(x_14); +x_17 = !lean_is_exclusive(x_1); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_18 = lean_ctor_get(x_1, 0); +x_19 = lean_ctor_get(x_1, 1); +x_20 = lean_ctor_get(x_1, 2); +x_21 = lean_ctor_get(x_1, 3); +x_22 = lean_ctor_get(x_1, 4); +x_23 = lean_ctor_get(x_1, 5); +x_24 = lean_ctor_get(x_1, 6); +x_25 = lean_ctor_get(x_1, 7); +x_26 = lean_ctor_get(x_1, 8); +x_27 = lean_ctor_get(x_1, 9); +x_28 = lean_ctor_get(x_1, 10); +x_29 = lean_ctor_get(x_1, 11); +x_30 = lean_ctor_get(x_1, 12); +x_31 = lean_ctor_get(x_1, 13); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +x_32 = l_Lean_MVarId_assert(x_18, x_15, x_3, x_2, x_8, x_9, x_10, x_11, x_16); +if (lean_obj_tag(x_32) == 0) +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_33 = lean_ctor_get(x_32, 0); +lean_inc(x_33); +x_34 = lean_ctor_get(x_32, 1); +lean_inc(x_34); +lean_dec(x_32); +lean_ctor_set(x_1, 0, x_33); +x_35 = l_Lean_Meta_Grind_intros(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_34); +return x_35; +} +else +{ +uint8_t x_36; +lean_free_object(x_1); +lean_dec(x_31); +lean_dec(x_30); +lean_dec(x_29); +lean_dec(x_28); +lean_dec(x_27); +lean_dec(x_26); +lean_dec(x_25); +lean_dec(x_24); +lean_dec(x_23); +lean_dec(x_22); +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); +lean_dec(x_5); +lean_dec(x_4); +x_36 = !lean_is_exclusive(x_32); +if (x_36 == 0) +{ +return x_32; +} +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_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; +} +} +} +else +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; uint8_t x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_40 = lean_ctor_get(x_1, 0); +x_41 = lean_ctor_get(x_1, 1); +x_42 = lean_ctor_get(x_1, 2); +x_43 = lean_ctor_get(x_1, 3); +x_44 = lean_ctor_get(x_1, 4); +x_45 = lean_ctor_get(x_1, 5); +x_46 = lean_ctor_get_uint8(x_1, sizeof(void*)*14); +x_47 = lean_ctor_get(x_1, 6); +x_48 = lean_ctor_get(x_1, 7); +x_49 = lean_ctor_get(x_1, 8); +x_50 = lean_ctor_get(x_1, 9); +x_51 = lean_ctor_get(x_1, 10); +x_52 = lean_ctor_get(x_1, 11); +x_53 = lean_ctor_get(x_1, 12); +x_54 = lean_ctor_get(x_1, 13); +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_45); +lean_inc(x_44); +lean_inc(x_43); +lean_inc(x_42); +lean_inc(x_41); +lean_inc(x_40); +lean_dec(x_1); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +x_55 = l_Lean_MVarId_assert(x_40, x_15, x_3, x_2, x_8, x_9, x_10, x_11, x_16); +if (lean_obj_tag(x_55) == 0) +{ +lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; +x_56 = lean_ctor_get(x_55, 0); +lean_inc(x_56); +x_57 = lean_ctor_get(x_55, 1); +lean_inc(x_57); +lean_dec(x_55); +x_58 = lean_alloc_ctor(0, 14, 1); +lean_ctor_set(x_58, 0, x_56); +lean_ctor_set(x_58, 1, x_41); +lean_ctor_set(x_58, 2, x_42); +lean_ctor_set(x_58, 3, x_43); +lean_ctor_set(x_58, 4, x_44); +lean_ctor_set(x_58, 5, x_45); +lean_ctor_set(x_58, 6, x_47); +lean_ctor_set(x_58, 7, x_48); +lean_ctor_set(x_58, 8, x_49); +lean_ctor_set(x_58, 9, x_50); +lean_ctor_set(x_58, 10, x_51); +lean_ctor_set(x_58, 11, x_52); +lean_ctor_set(x_58, 12, x_53); +lean_ctor_set(x_58, 13, x_54); +lean_ctor_set_uint8(x_58, sizeof(void*)*14, x_46); +x_59 = l_Lean_Meta_Grind_intros(x_58, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_57); +return x_59; +} +else +{ +lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; +lean_dec(x_54); +lean_dec(x_53); +lean_dec(x_52); +lean_dec(x_51); +lean_dec(x_50); +lean_dec(x_49); +lean_dec(x_48); +lean_dec(x_47); +lean_dec(x_45); +lean_dec(x_44); +lean_dec(x_43); +lean_dec(x_42); +lean_dec(x_41); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_60 = lean_ctor_get(x_55, 0); +lean_inc(x_60); +x_61 = lean_ctor_get(x_55, 1); +lean_inc(x_61); +if (lean_is_exclusive(x_55)) { + lean_ctor_release(x_55, 0); + lean_ctor_release(x_55, 1); + x_62 = x_55; +} else { + lean_dec_ref(x_55); + x_62 = lean_box(0); +} +if (lean_is_scalar(x_62)) { + x_63 = lean_alloc_ctor(1, 2, 0); +} else { + x_63 = x_62; +} +lean_ctor_set(x_63, 0, x_60); +lean_ctor_set(x_63, 1, x_61); +return x_63; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_assertNext_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) { +_start: +{ +uint8_t x_10; +x_10 = !lean_is_exclusive(x_1); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; 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_11 = lean_ctor_get(x_1, 0); +x_12 = lean_ctor_get(x_1, 1); +x_13 = lean_ctor_get(x_1, 2); +x_14 = lean_ctor_get(x_1, 3); +x_15 = lean_ctor_get(x_1, 4); +x_16 = lean_ctor_get(x_1, 5); +x_17 = lean_ctor_get(x_1, 6); +x_18 = lean_ctor_get(x_1, 7); +x_19 = lean_ctor_get(x_1, 8); +x_20 = lean_ctor_get(x_1, 9); +x_21 = lean_ctor_get(x_1, 10); +x_22 = lean_ctor_get(x_1, 11); +x_23 = lean_ctor_get(x_1, 12); +x_24 = lean_ctor_get(x_1, 13); +x_25 = l_Std_Queue_dequeue_x3f___rarg(x_24); +if (lean_obj_tag(x_25) == 0) +{ +lean_object* x_26; lean_object* x_27; +lean_free_object(x_1); +lean_dec(x_23); +lean_dec(x_22); +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +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_9); +return x_27; +} +else +{ +uint8_t x_28; +x_28 = !lean_is_exclusive(x_25); +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; +x_29 = lean_ctor_get(x_25, 0); +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_ctor_set(x_1, 13, x_31); +x_32 = lean_ctor_get(x_30, 0); +lean_inc(x_32); +x_33 = lean_ctor_get(x_30, 1); +lean_inc(x_33); +x_34 = lean_ctor_get(x_30, 2); +lean_inc(x_34); +lean_dec(x_30); +x_35 = l_Lean_Meta_Grind_assertAt(x_1, x_32, x_33, x_34, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +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; +x_37 = lean_ctor_get(x_35, 0); +lean_ctor_set(x_25, 0, x_37); +lean_ctor_set(x_35, 0, x_25); +return x_35; +} +else +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_38 = lean_ctor_get(x_35, 0); +x_39 = lean_ctor_get(x_35, 1); +lean_inc(x_39); +lean_inc(x_38); +lean_dec(x_35); +lean_ctor_set(x_25, 0, x_38); +x_40 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_40, 0, x_25); +lean_ctor_set(x_40, 1, x_39); +return x_40; +} +} +else +{ +uint8_t x_41; +lean_free_object(x_25); +x_41 = !lean_is_exclusive(x_35); +if (x_41 == 0) +{ +return x_35; +} +else +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_42 = lean_ctor_get(x_35, 0); +x_43 = lean_ctor_get(x_35, 1); +lean_inc(x_43); +lean_inc(x_42); +lean_dec(x_35); +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_45; 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_45 = lean_ctor_get(x_25, 0); +lean_inc(x_45); +lean_dec(x_25); +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); +lean_ctor_set(x_1, 13, x_47); +x_48 = lean_ctor_get(x_46, 0); +lean_inc(x_48); +x_49 = lean_ctor_get(x_46, 1); +lean_inc(x_49); +x_50 = lean_ctor_get(x_46, 2); +lean_inc(x_50); +lean_dec(x_46); +x_51 = l_Lean_Meta_Grind_assertAt(x_1, x_48, x_49, x_50, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_51) == 0) +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_52 = lean_ctor_get(x_51, 0); +lean_inc(x_52); +x_53 = lean_ctor_get(x_51, 1); +lean_inc(x_53); +if (lean_is_exclusive(x_51)) { + lean_ctor_release(x_51, 0); + lean_ctor_release(x_51, 1); + x_54 = x_51; +} else { + lean_dec_ref(x_51); + x_54 = lean_box(0); +} +x_55 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_55, 0, x_52); +if (lean_is_scalar(x_54)) { + x_56 = lean_alloc_ctor(0, 2, 0); +} else { + x_56 = x_54; +} +lean_ctor_set(x_56, 0, x_55); +lean_ctor_set(x_56, 1, x_53); +return x_56; +} +else +{ +lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_57 = lean_ctor_get(x_51, 0); +lean_inc(x_57); +x_58 = lean_ctor_get(x_51, 1); +lean_inc(x_58); +if (lean_is_exclusive(x_51)) { + lean_ctor_release(x_51, 0); + lean_ctor_release(x_51, 1); + x_59 = x_51; +} else { + lean_dec_ref(x_51); + 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; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; uint8_t x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; +x_61 = lean_ctor_get(x_1, 0); +x_62 = lean_ctor_get(x_1, 1); +x_63 = lean_ctor_get(x_1, 2); +x_64 = lean_ctor_get(x_1, 3); +x_65 = lean_ctor_get(x_1, 4); +x_66 = lean_ctor_get(x_1, 5); +x_67 = lean_ctor_get_uint8(x_1, sizeof(void*)*14); +x_68 = lean_ctor_get(x_1, 6); +x_69 = lean_ctor_get(x_1, 7); +x_70 = lean_ctor_get(x_1, 8); +x_71 = lean_ctor_get(x_1, 9); +x_72 = lean_ctor_get(x_1, 10); +x_73 = lean_ctor_get(x_1, 11); +x_74 = lean_ctor_get(x_1, 12); +x_75 = lean_ctor_get(x_1, 13); +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_66); +lean_inc(x_65); +lean_inc(x_64); +lean_inc(x_63); +lean_inc(x_62); +lean_inc(x_61); +lean_dec(x_1); +x_76 = l_Std_Queue_dequeue_x3f___rarg(x_75); +if (lean_obj_tag(x_76) == 0) +{ +lean_object* x_77; lean_object* x_78; +lean_dec(x_74); +lean_dec(x_73); +lean_dec(x_72); +lean_dec(x_71); +lean_dec(x_70); +lean_dec(x_69); +lean_dec(x_68); +lean_dec(x_66); +lean_dec(x_65); +lean_dec(x_64); +lean_dec(x_63); +lean_dec(x_62); +lean_dec(x_61); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +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_9); +return x_78; +} +else +{ +lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; +x_79 = lean_ctor_get(x_76, 0); +lean_inc(x_79); +if (lean_is_exclusive(x_76)) { + lean_ctor_release(x_76, 0); + x_80 = x_76; +} else { + lean_dec_ref(x_76); + x_80 = lean_box(0); +} +x_81 = lean_ctor_get(x_79, 0); +lean_inc(x_81); +x_82 = lean_ctor_get(x_79, 1); +lean_inc(x_82); +lean_dec(x_79); +x_83 = lean_alloc_ctor(0, 14, 1); +lean_ctor_set(x_83, 0, x_61); +lean_ctor_set(x_83, 1, x_62); +lean_ctor_set(x_83, 2, x_63); +lean_ctor_set(x_83, 3, x_64); +lean_ctor_set(x_83, 4, x_65); +lean_ctor_set(x_83, 5, x_66); +lean_ctor_set(x_83, 6, x_68); +lean_ctor_set(x_83, 7, x_69); +lean_ctor_set(x_83, 8, x_70); +lean_ctor_set(x_83, 9, x_71); +lean_ctor_set(x_83, 10, x_72); +lean_ctor_set(x_83, 11, x_73); +lean_ctor_set(x_83, 12, x_74); +lean_ctor_set(x_83, 13, x_82); +lean_ctor_set_uint8(x_83, sizeof(void*)*14, x_67); +x_84 = lean_ctor_get(x_81, 0); +lean_inc(x_84); +x_85 = lean_ctor_get(x_81, 1); +lean_inc(x_85); +x_86 = lean_ctor_get(x_81, 2); +lean_inc(x_86); +lean_dec(x_81); +x_87 = l_Lean_Meta_Grind_assertAt(x_83, x_84, x_85, x_86, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +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); +if (lean_is_exclusive(x_87)) { + lean_ctor_release(x_87, 0); + lean_ctor_release(x_87, 1); + x_90 = x_87; +} else { + lean_dec_ref(x_87); + x_90 = lean_box(0); +} +if (lean_is_scalar(x_80)) { + x_91 = lean_alloc_ctor(1, 1, 0); +} else { + x_91 = x_80; +} +lean_ctor_set(x_91, 0, x_88); +if (lean_is_scalar(x_90)) { + x_92 = lean_alloc_ctor(0, 2, 0); +} else { + x_92 = x_90; +} +lean_ctor_set(x_92, 0, x_91); +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_dec(x_80); +x_93 = lean_ctor_get(x_87, 0); +lean_inc(x_93); +x_94 = lean_ctor_get(x_87, 1); +lean_inc(x_94); +if (lean_is_exclusive(x_87)) { + lean_ctor_release(x_87, 0); + lean_ctor_release(x_87, 1); + x_95 = x_87; +} else { + lean_dec_ref(x_87); + 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; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_iterate_go(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_12; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_3); +lean_ctor_set(x_12, 1, x_11); +return x_12; +} +else +{ +uint8_t x_13; +x_13 = !lean_is_exclusive(x_2); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_14 = lean_ctor_get(x_2, 0); +x_15 = lean_ctor_get(x_2, 1); +lean_inc(x_1); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_14); +x_16 = lean_apply_9(x_1, x_14, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_16) == 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_18; +x_18 = lean_ctor_get(x_16, 1); +lean_inc(x_18); +lean_dec(x_16); +lean_ctor_set(x_2, 1, x_3); +{ +lean_object* _tmp_1 = x_15; +lean_object* _tmp_2 = x_2; +lean_object* _tmp_10 = x_18; +x_2 = _tmp_1; +x_3 = _tmp_2; +x_11 = _tmp_10; +} +goto _start; +} +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; +lean_free_object(x_2); +lean_dec(x_14); +x_20 = lean_ctor_get(x_16, 1); +lean_inc(x_20); +lean_dec(x_16); +x_21 = lean_ctor_get(x_17, 0); +lean_inc(x_21); +lean_dec(x_17); +x_22 = l_List_appendTR___rarg(x_21, x_15); +x_2 = x_22; +x_11 = x_20; +goto _start; +} +} +else +{ +uint8_t x_24; +lean_free_object(x_2); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_24 = !lean_is_exclusive(x_16); +if (x_24 == 0) +{ +return x_16; +} +else +{ +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_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; +} +} +} +else +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_28 = lean_ctor_get(x_2, 0); +x_29 = lean_ctor_get(x_2, 1); +lean_inc(x_29); +lean_inc(x_28); +lean_dec(x_2); +lean_inc(x_1); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_28); +x_30 = lean_apply_9(x_1, x_28, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_30) == 0) +{ +lean_object* x_31; +x_31 = lean_ctor_get(x_30, 0); +lean_inc(x_31); +if (lean_obj_tag(x_31) == 0) +{ +lean_object* x_32; lean_object* x_33; +x_32 = lean_ctor_get(x_30, 1); +lean_inc(x_32); +lean_dec(x_30); +x_33 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_33, 0, x_28); +lean_ctor_set(x_33, 1, x_3); +x_2 = x_29; +x_3 = x_33; +x_11 = x_32; +goto _start; +} +else +{ +lean_object* x_35; lean_object* x_36; lean_object* x_37; +lean_dec(x_28); +x_35 = lean_ctor_get(x_30, 1); +lean_inc(x_35); +lean_dec(x_30); +x_36 = lean_ctor_get(x_31, 0); +lean_inc(x_36); +lean_dec(x_31); +x_37 = l_List_appendTR___rarg(x_36, x_29); +x_2 = x_37; +x_11 = x_35; +goto _start; +} +} +else +{ +lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; +lean_dec(x_29); +lean_dec(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_3); +lean_dec(x_1); +x_39 = lean_ctor_get(x_30, 0); +lean_inc(x_39); +x_40 = lean_ctor_get(x_30, 1); +lean_inc(x_40); +if (lean_is_exclusive(x_30)) { + lean_ctor_release(x_30, 0); + lean_ctor_release(x_30, 1); + x_41 = x_30; +} else { + lean_dec_ref(x_30); + x_41 = lean_box(0); +} +if (lean_is_scalar(x_41)) { + x_42 = lean_alloc_ctor(1, 2, 0); +} else { + x_42 = x_41; +} +lean_ctor_set(x_42, 0, x_39); +lean_ctor_set(x_42, 1, x_40); +return x_42; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_iterate(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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_box(0); +x_12 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_12, 0, x_1); +lean_ctor_set(x_12, 1, x_11); +x_13 = l_Lean_Meta_Grind_iterate_go(x_2, x_12, x_11, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_13; +} +} +static lean_object* _init_l_Lean_Meta_Grind_assertAll___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_assertNext_x3f), 9, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_assertAll(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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_Meta_Grind_assertAll___closed__1; +x_11 = l_Lean_Meta_Grind_iterate(x_1, x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_11; +} +} +lean_object* initialize_Init_Grind_Lemmas(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Meta_Tactic_Assert(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Meta_Tactic_Grind_Simp(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Meta_Tactic_Grind_Types(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Meta_Tactic_Grind_Cases(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Meta_Tactic_Grind_Injection(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Meta_Tactic_Grind_Core(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Lean_Meta_Tactic_Grind_Intro(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_Grind_Lemmas(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Lean_Meta_Tactic_Assert(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Lean_Meta_Tactic_Grind_Simp(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Lean_Meta_Tactic_Grind_Types(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Lean_Meta_Tactic_Grind_Cases(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Lean_Meta_Tactic_Grind_Injection(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Lean_Meta_Tactic_Grind_Core(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l_Lean_Meta_Grind_instInhabitedIntroResult = _init_l_Lean_Meta_Grind_instInhabitedIntroResult(); +lean_mark_persistent(l_Lean_Meta_Grind_instInhabitedIntroResult); +l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__1); +l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__2); +l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__3); +l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__4(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__4); +l_Lean_Meta_Grind_intros_go___lambda__4___closed__1 = _init_l_Lean_Meta_Grind_intros_go___lambda__4___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_intros_go___lambda__4___closed__1); +l_Lean_Meta_Grind_intros___closed__1 = _init_l_Lean_Meta_Grind_intros___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_intros___closed__1); +l_Lean_Meta_Grind_assertAt___closed__1 = _init_l_Lean_Meta_Grind_assertAt___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_assertAt___closed__1); +l_Lean_Meta_Grind_assertAt___closed__2 = _init_l_Lean_Meta_Grind_assertAt___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_assertAt___closed__2); +l_Lean_Meta_Grind_assertAll___closed__1 = _init_l_Lean_Meta_Grind_assertAll___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_assertAll___closed__1); +return lean_io_result_mk_ok(lean_box(0)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Lean/Meta/Tactic/Grind/Inv.c b/stage0/stdlib/Lean/Meta/Tactic/Grind/Inv.c index 244a42840b2f..dc2a5be0cb31 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Grind/Inv.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Grind/Inv.c @@ -16,27 +16,29 @@ extern "C" { LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___closed__1; -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___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_panic___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__1___closed__6; -static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__3; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Goal_checkInvariants___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_PersistentHashMap_foldlM___at_Lean_Meta_Grind_checkInvariants___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___closed__2; LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___boxed(lean_object**); -static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__4; size_t lean_usize_shift_right(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_checkInvariants___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(lean_object*, lean_object*); +lean_object* l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___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_panic___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__2___closed__1; +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___closed__3; static lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__1___closed__3; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Goal_checkInvariants___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_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___closed__3; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux___at_Lean_Meta_Grind_checkInvariants___spec__3(lean_object*, lean_object*, lean_object*); size_t lean_uint64_to_usize(uint64_t); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Goal_checkInvariants___lambda__2(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_Grind_checkInvariants___spec__4___rarg(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isApp(lean_object*); -static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__3; uint8_t l_Lean_Meta_Grind_isCongruent(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_sort___override(lean_object*); +lean_object* l_Lean_PersistentArray_push___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findEntryAux___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__5(lean_object*, lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_Meta_Grind_checkInvariants___spec__5___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_mk_array(lean_object*, lean_object*); @@ -45,27 +47,30 @@ uint8_t lean_usize_dec_eq(size_t, size_t); static lean_object* l___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___closed__2; uint64_t l_Lean_Meta_Grind_congrHash(lean_object*, lean_object*); lean_object* lean_array_fget(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Goal_checkInvariants___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Meta_Grind_grind_debug_proofs; static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__1; static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__4; LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___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_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___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* l_Lean_Meta_Grind_getEqcs(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_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__5; static lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___lambda__1___closed__4; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Goal_checkInvariants___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_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_stringToMessageData(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__2; static lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___closed__2; +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___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*); +static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__4; extern lean_object* l_Lean_Meta_Grind_grind_debug; static lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___closed__8; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents(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_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__2; extern lean_object* l_instInhabitedNat; extern lean_object* l_instInhabitedPUnit; lean_object* l_Lean_Meta_Grind_getRoot(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_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__1; -lean_object* l_Lean_addTrace___at_Lean_Meta_Grind_mkEqProofImpl_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*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___closed__1; lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_checkInvariants___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -76,36 +81,41 @@ size_t lean_usize_of_nat(lean_object*); LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___closed__3; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forIn___at_Lean_Meta_Grind_checkInvariants___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__3; static lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___lambda__2___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Goal_checkInvariants(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_st_ref_take(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_checkInvariants___closed__2; static lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___lambda__2___closed__2; -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___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 lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___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_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__2; static lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___lambda__1___closed__2; static lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___closed__1; +lean_object* l_Lean_Meta_Grind_mkEqHEqProof(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___closed__5; LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___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_object*, lean_object*); +static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__4; static lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__1___closed__7; lean_object* l_Lean_Meta_check(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_get(lean_object*, lean_object*); +lean_object* lean_st_mk_ref(lean_object*, lean_object*); static lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___closed__4; -static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__1; static lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___closed__9; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_Grind_checkInvariants___spec__4___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findEntryAux___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___closed__5; extern lean_object* l_Lean_levelZero; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_checkInvariants(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_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___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*); extern lean_object* l_Lean_instInhabitedExpr; -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forIn___at_Lean_Meta_Grind_checkInvariants___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*); -static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__5; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_Meta_Grind_checkInvariants___spec__5___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_checkInvariants___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_Grind_checkInvariants___closed__1; +LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static size_t l_Lean_PersistentHashMap_findEntryAux___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__5___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static double l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__1; LEAN_EXPORT lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___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_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Meta_Grind_isSameExpr_unsafe__1(lean_object*, lean_object*); uint8_t lean_expr_equal(lean_object*, lean_object*); uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); @@ -113,7 +123,7 @@ static lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Me LEAN_EXPORT lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_getParents(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*); -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___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*); +static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__3; LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___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_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___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*); extern lean_object* l_Lean_Meta_instMonadMetaM; @@ -123,42 +133,45 @@ lean_object* lean_usize_to_nat(size_t); lean_object* l_Lean_Meta_Grind_getENodes(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*); LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__2___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*); lean_object* l_Lean_Meta_Grind_getNext(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___boxed(lean_object**); -lean_object* lean_grind_mk_eq_proof(lean_object*, 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_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__5; static lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___closed__7; -static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__6; +lean_object* l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___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_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__2; +static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__6; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__7; +LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___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_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__5; LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___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_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__1; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_Meta_Grind_checkInvariants___spec__5(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___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_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*); static lean_object* l_Lean_Meta_Grind_checkInvariants___lambda__1___closed__1; +static lean_object* l_Lean_Meta_Grind_Goal_checkInvariants___closed__1; lean_object* lean_nat_sub(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findEntry_x3f___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__4(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_getAppFn(lean_object*); -lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_mkEqProofImpl_go___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_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static size_t l_Lean_PersistentHashMap_findEntryAux___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__5___closed__1; lean_object* l_Lean_Meta_Grind_getENode(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_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__1___closed__4; static lean_object* l___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___closed__3; -static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__4; static lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___closed__2; static lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__2___closed__2; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findEntryAtAux___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___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_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_sub(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___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_array_mk(lean_object*); lean_object* l_Lean_Meta_Grind_isRoot(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_Grind_Inv_0__Lean_Meta_Grind_checkEqc___closed__1; -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forIn___at_Lean_Meta_Grind_checkInvariants___spec__1___lambda__1(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_EXPORT lean_object* l_Lean_PersistentHashMap_forIn___at_Lean_Meta_Grind_checkInvariants___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*); static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__6; +lean_object* l_Lean_isTracingEnabledForCore(lean_object*, lean_object*, lean_object*); size_t lean_usize_add(size_t, size_t); static lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___closed__6; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_checkInvariants___lambda__1(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -168,25 +181,34 @@ size_t lean_array_size(lean_object*); LEAN_EXPORT lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___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_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_instInhabitedOfMonad___rarg(lean_object*, lean_object*); +lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Goal_checkInvariants___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_shift_left(size_t, size_t); lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__3; +static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__1; lean_object* lean_string_append(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc(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_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__8___closed__1; lean_object* lean_array_get_size(lean_object*); +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_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_PersistentHashMap_findEntryAtAux___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___lambda__2___closed__3; uint8_t lean_usize_dec_lt(size_t, size_t); static lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___lambda__1___closed__5; -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___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_nat_add(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___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_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__3; +LEAN_EXPORT lean_object* l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___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_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__2; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Goal_checkInvariants___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_foldlMUnsafe_fold___at_Lean_Meta_Grind_checkInvariants___spec__4(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_getRoot_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_get(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___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*); +static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__2; lean_object* l_ReaderT_instMonad___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux___at_Lean_Meta_Grind_checkInvariants___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*, lean_object*); size_t lean_usize_land(size_t, size_t); @@ -3158,51 +3180,414 @@ lean_dec(x_1); return x_18; } } -static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__1() { +LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_11 = lean_ctor_get(x_8, 2); +x_12 = l_Lean_isTracingEnabledForCore(x_1, x_11, x_10); +x_13 = !lean_is_exclusive(x_12); +if (x_13 == 0) +{ +return x_12; +} +else +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_14 = lean_ctor_get(x_12, 0); +x_15 = lean_ctor_get(x_12, 1); +lean_inc(x_15); +lean_inc(x_14); +lean_dec(x_12); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_14); +lean_ctor_set(x_16, 1, x_15); +return x_16; +} +} +} +static double _init_l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__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_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("", 0, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__3() { _start: { lean_object* x_1; lean_object* x_2; x_1 = lean_box(0); -x_2 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_2, 0, x_1); +x_2 = lean_array_mk(x_1); return x_2; } } -static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__2() { +LEAN_EXPORT lean_object* l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("checked: ", 9, 9); -return x_1; +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; +x_12 = lean_ctor_get(x_9, 5); +x_13 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_2, x_7, x_8, x_9, x_10, x_11); +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +x_16 = lean_st_ref_take(x_10, x_15); +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +x_18 = lean_ctor_get(x_17, 3); +lean_inc(x_18); +x_19 = !lean_is_exclusive(x_16); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_20 = lean_ctor_get(x_16, 1); +x_21 = lean_ctor_get(x_16, 0); +lean_dec(x_21); +x_22 = !lean_is_exclusive(x_17); +if (x_22 == 0) +{ +lean_object* x_23; uint8_t x_24; +x_23 = lean_ctor_get(x_17, 3); +lean_dec(x_23); +x_24 = !lean_is_exclusive(x_18); +if (x_24 == 0) +{ +lean_object* x_25; double 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; uint8_t x_34; +x_25 = lean_ctor_get(x_18, 0); +x_26 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__1; +x_27 = 0; +x_28 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__2; +x_29 = lean_alloc_ctor(0, 2, 17); +lean_ctor_set(x_29, 0, x_1); +lean_ctor_set(x_29, 1, x_28); +lean_ctor_set_float(x_29, sizeof(void*)*2, x_26); +lean_ctor_set_float(x_29, sizeof(void*)*2 + 8, x_26); +lean_ctor_set_uint8(x_29, sizeof(void*)*2 + 16, x_27); +x_30 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__3; +x_31 = lean_alloc_ctor(9, 3, 0); +lean_ctor_set(x_31, 0, x_29); +lean_ctor_set(x_31, 1, x_14); +lean_ctor_set(x_31, 2, x_30); +lean_inc(x_12); +lean_ctor_set(x_16, 1, x_31); +lean_ctor_set(x_16, 0, x_12); +x_32 = l_Lean_PersistentArray_push___rarg(x_25, x_16); +lean_ctor_set(x_18, 0, x_32); +x_33 = lean_st_ref_set(x_10, x_17, x_20); +x_34 = !lean_is_exclusive(x_33); +if (x_34 == 0) +{ +lean_object* x_35; lean_object* x_36; +x_35 = lean_ctor_get(x_33, 0); +lean_dec(x_35); +x_36 = lean_box(0); +lean_ctor_set(x_33, 0, x_36); +return x_33; +} +else +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_37 = lean_ctor_get(x_33, 1); +lean_inc(x_37); +lean_dec(x_33); +x_38 = lean_box(0); +x_39 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_37); +return x_39; } } -static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__3() { +else +{ +uint64_t x_40; lean_object* x_41; double 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; +x_40 = lean_ctor_get_uint64(x_18, sizeof(void*)*1); +x_41 = lean_ctor_get(x_18, 0); +lean_inc(x_41); +lean_dec(x_18); +x_42 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__1; +x_43 = 0; +x_44 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__2; +x_45 = lean_alloc_ctor(0, 2, 17); +lean_ctor_set(x_45, 0, x_1); +lean_ctor_set(x_45, 1, x_44); +lean_ctor_set_float(x_45, sizeof(void*)*2, x_42); +lean_ctor_set_float(x_45, sizeof(void*)*2 + 8, x_42); +lean_ctor_set_uint8(x_45, sizeof(void*)*2 + 16, x_43); +x_46 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__3; +x_47 = lean_alloc_ctor(9, 3, 0); +lean_ctor_set(x_47, 0, x_45); +lean_ctor_set(x_47, 1, x_14); +lean_ctor_set(x_47, 2, x_46); +lean_inc(x_12); +lean_ctor_set(x_16, 1, x_47); +lean_ctor_set(x_16, 0, x_12); +x_48 = l_Lean_PersistentArray_push___rarg(x_41, x_16); +x_49 = lean_alloc_ctor(0, 1, 8); +lean_ctor_set(x_49, 0, x_48); +lean_ctor_set_uint64(x_49, sizeof(void*)*1, x_40); +lean_ctor_set(x_17, 3, x_49); +x_50 = lean_st_ref_set(x_10, x_17, x_20); +x_51 = lean_ctor_get(x_50, 1); +lean_inc(x_51); +if (lean_is_exclusive(x_50)) { + lean_ctor_release(x_50, 0); + lean_ctor_release(x_50, 1); + x_52 = x_50; +} else { + lean_dec_ref(x_50); + x_52 = lean_box(0); +} +x_53 = lean_box(0); +if (lean_is_scalar(x_52)) { + x_54 = lean_alloc_ctor(0, 2, 0); +} else { + x_54 = x_52; +} +lean_ctor_set(x_54, 0, x_53); +lean_ctor_set(x_54, 1, x_51); +return x_54; +} +} +else +{ +lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; uint64_t x_62; lean_object* x_63; lean_object* x_64; double x_65; uint8_t x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; +x_55 = lean_ctor_get(x_17, 0); +x_56 = lean_ctor_get(x_17, 1); +x_57 = lean_ctor_get(x_17, 2); +x_58 = lean_ctor_get(x_17, 4); +x_59 = lean_ctor_get(x_17, 5); +x_60 = lean_ctor_get(x_17, 6); +x_61 = lean_ctor_get(x_17, 7); +lean_inc(x_61); +lean_inc(x_60); +lean_inc(x_59); +lean_inc(x_58); +lean_inc(x_57); +lean_inc(x_56); +lean_inc(x_55); +lean_dec(x_17); +x_62 = lean_ctor_get_uint64(x_18, sizeof(void*)*1); +x_63 = lean_ctor_get(x_18, 0); +lean_inc(x_63); +if (lean_is_exclusive(x_18)) { + lean_ctor_release(x_18, 0); + x_64 = x_18; +} else { + lean_dec_ref(x_18); + x_64 = lean_box(0); +} +x_65 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__1; +x_66 = 0; +x_67 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__2; +x_68 = lean_alloc_ctor(0, 2, 17); +lean_ctor_set(x_68, 0, x_1); +lean_ctor_set(x_68, 1, x_67); +lean_ctor_set_float(x_68, sizeof(void*)*2, x_65); +lean_ctor_set_float(x_68, sizeof(void*)*2 + 8, x_65); +lean_ctor_set_uint8(x_68, sizeof(void*)*2 + 16, x_66); +x_69 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__3; +x_70 = lean_alloc_ctor(9, 3, 0); +lean_ctor_set(x_70, 0, x_68); +lean_ctor_set(x_70, 1, x_14); +lean_ctor_set(x_70, 2, x_69); +lean_inc(x_12); +lean_ctor_set(x_16, 1, x_70); +lean_ctor_set(x_16, 0, x_12); +x_71 = l_Lean_PersistentArray_push___rarg(x_63, x_16); +if (lean_is_scalar(x_64)) { + x_72 = lean_alloc_ctor(0, 1, 8); +} else { + x_72 = x_64; +} +lean_ctor_set(x_72, 0, x_71); +lean_ctor_set_uint64(x_72, sizeof(void*)*1, x_62); +x_73 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_73, 0, x_55); +lean_ctor_set(x_73, 1, x_56); +lean_ctor_set(x_73, 2, x_57); +lean_ctor_set(x_73, 3, x_72); +lean_ctor_set(x_73, 4, x_58); +lean_ctor_set(x_73, 5, x_59); +lean_ctor_set(x_73, 6, x_60); +lean_ctor_set(x_73, 7, x_61); +x_74 = lean_st_ref_set(x_10, x_73, x_20); +x_75 = lean_ctor_get(x_74, 1); +lean_inc(x_75); +if (lean_is_exclusive(x_74)) { + lean_ctor_release(x_74, 0); + lean_ctor_release(x_74, 1); + x_76 = x_74; +} else { + lean_dec_ref(x_74); + x_76 = lean_box(0); +} +x_77 = lean_box(0); +if (lean_is_scalar(x_76)) { + x_78 = lean_alloc_ctor(0, 2, 0); +} else { + x_78 = x_76; +} +lean_ctor_set(x_78, 0, x_77); +lean_ctor_set(x_78, 1, x_75); +return x_78; +} +} +else +{ +lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; uint64_t x_88; lean_object* x_89; lean_object* x_90; double x_91; uint8_t x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; +x_79 = lean_ctor_get(x_16, 1); +lean_inc(x_79); +lean_dec(x_16); +x_80 = lean_ctor_get(x_17, 0); +lean_inc(x_80); +x_81 = lean_ctor_get(x_17, 1); +lean_inc(x_81); +x_82 = lean_ctor_get(x_17, 2); +lean_inc(x_82); +x_83 = lean_ctor_get(x_17, 4); +lean_inc(x_83); +x_84 = lean_ctor_get(x_17, 5); +lean_inc(x_84); +x_85 = lean_ctor_get(x_17, 6); +lean_inc(x_85); +x_86 = lean_ctor_get(x_17, 7); +lean_inc(x_86); +if (lean_is_exclusive(x_17)) { + lean_ctor_release(x_17, 0); + lean_ctor_release(x_17, 1); + lean_ctor_release(x_17, 2); + lean_ctor_release(x_17, 3); + lean_ctor_release(x_17, 4); + lean_ctor_release(x_17, 5); + lean_ctor_release(x_17, 6); + lean_ctor_release(x_17, 7); + x_87 = x_17; +} else { + lean_dec_ref(x_17); + x_87 = lean_box(0); +} +x_88 = lean_ctor_get_uint64(x_18, sizeof(void*)*1); +x_89 = lean_ctor_get(x_18, 0); +lean_inc(x_89); +if (lean_is_exclusive(x_18)) { + lean_ctor_release(x_18, 0); + x_90 = x_18; +} else { + lean_dec_ref(x_18); + x_90 = lean_box(0); +} +x_91 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__1; +x_92 = 0; +x_93 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__2; +x_94 = lean_alloc_ctor(0, 2, 17); +lean_ctor_set(x_94, 0, x_1); +lean_ctor_set(x_94, 1, x_93); +lean_ctor_set_float(x_94, sizeof(void*)*2, x_91); +lean_ctor_set_float(x_94, sizeof(void*)*2 + 8, x_91); +lean_ctor_set_uint8(x_94, sizeof(void*)*2 + 16, x_92); +x_95 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__3; +x_96 = lean_alloc_ctor(9, 3, 0); +lean_ctor_set(x_96, 0, x_94); +lean_ctor_set(x_96, 1, x_14); +lean_ctor_set(x_96, 2, x_95); +lean_inc(x_12); +x_97 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_97, 0, x_12); +lean_ctor_set(x_97, 1, x_96); +x_98 = l_Lean_PersistentArray_push___rarg(x_89, x_97); +if (lean_is_scalar(x_90)) { + x_99 = lean_alloc_ctor(0, 1, 8); +} else { + x_99 = x_90; +} +lean_ctor_set(x_99, 0, x_98); +lean_ctor_set_uint64(x_99, sizeof(void*)*1, x_88); +if (lean_is_scalar(x_87)) { + x_100 = lean_alloc_ctor(0, 8, 0); +} else { + x_100 = x_87; +} +lean_ctor_set(x_100, 0, x_80); +lean_ctor_set(x_100, 1, x_81); +lean_ctor_set(x_100, 2, x_82); +lean_ctor_set(x_100, 3, x_99); +lean_ctor_set(x_100, 4, x_83); +lean_ctor_set(x_100, 5, x_84); +lean_ctor_set(x_100, 6, x_85); +lean_ctor_set(x_100, 7, x_86); +x_101 = lean_st_ref_set(x_10, x_100, x_79); +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; +} else { + lean_dec_ref(x_101); + x_103 = lean_box(0); +} +x_104 = lean_box(0); +if (lean_is_scalar(x_103)) { + x_105 = lean_alloc_ctor(0, 2, 0); +} else { + x_105 = x_103; +} +lean_ctor_set(x_105, 0, x_104); +lean_ctor_set(x_105, 1, x_102); +return x_105; +} +} +} +static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__1() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__2; -x_2 = l_Lean_stringToMessageData(x_1); +x_1 = lean_box(0); +x_2 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__4() { +static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__2() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("", 0, 0); +x_1 = lean_mk_string_unchecked("checked: ", 9, 9); return x_1; } } -static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__5() { +static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__3() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__4; +x_1 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__2; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__2; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___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) { _start: { lean_object* x_13; @@ -3219,7 +3604,7 @@ x_14 = lean_ctor_get(x_13, 1); lean_inc(x_14); lean_dec(x_13); lean_inc(x_2); -x_15 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_mkEqProofImpl_go___spec__1(x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_14); +x_15 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1(x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_14); x_16 = lean_ctor_get(x_15, 0); lean_inc(x_16); x_17 = lean_unbox(x_16); @@ -3239,7 +3624,7 @@ if (x_18 == 0) lean_object* x_19; lean_object* x_20; x_19 = lean_ctor_get(x_15, 0); lean_dec(x_19); -x_20 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__1; +x_20 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__1; lean_ctor_set(x_15, 0, x_20); return x_15; } @@ -3249,7 +3634,7 @@ lean_object* x_21; lean_object* x_22; lean_object* x_23; x_21 = lean_ctor_get(x_15, 1); lean_inc(x_21); lean_dec(x_15); -x_22 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__1; +x_22 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__1; x_23 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_23, 0, x_22); lean_ctor_set(x_23, 1, x_21); @@ -3276,15 +3661,15 @@ x_27 = lean_ctor_get(x_25, 1); lean_inc(x_27); lean_dec(x_25); x_28 = l_Lean_MessageData_ofExpr(x_26); -x_29 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__3; +x_29 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__3; x_30 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_30, 0, x_29); lean_ctor_set(x_30, 1, x_28); -x_31 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__5; +x_31 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__4; x_32 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_32, 0, x_30); lean_ctor_set(x_32, 1, x_31); -x_33 = l_Lean_addTrace___at_Lean_Meta_Grind_mkEqProofImpl_go___spec__2(x_2, x_32, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_27); +x_33 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2(x_2, x_32, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_27); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -3295,7 +3680,7 @@ if (x_34 == 0) lean_object* x_35; lean_object* x_36; x_35 = lean_ctor_get(x_33, 0); lean_dec(x_35); -x_36 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__1; +x_36 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__1; lean_ctor_set(x_33, 0, x_36); return x_33; } @@ -3305,7 +3690,7 @@ lean_object* x_37; lean_object* x_38; lean_object* x_39; x_37 = lean_ctor_get(x_33, 1); lean_inc(x_37); lean_dec(x_33); -x_38 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__1; +x_38 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__1; x_39 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_39, 0, x_38); lean_ctor_set(x_39, 1, x_37); @@ -3371,7 +3756,7 @@ return x_47; } } } -static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__1() { +static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__1() { _start: { lean_object* x_1; @@ -3379,7 +3764,7 @@ x_1 = lean_mk_string_unchecked("grind", 5, 5); return x_1; } } -static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__2() { +static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__2() { _start: { lean_object* x_1; @@ -3387,7 +3772,7 @@ x_1 = lean_mk_string_unchecked("debug", 5, 5); return x_1; } } -static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__3() { +static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__3() { _start: { lean_object* x_1; @@ -3395,18 +3780,18 @@ x_1 = lean_mk_string_unchecked("proofs", 6, 6); return x_1; } } -static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__4() { +static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__1; -x_2 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__2; -x_3 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__3; +x_1 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__1; +x_2 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__2; +x_3 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__3; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__5() { +static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__5() { _start: { lean_object* x_1; @@ -3414,16 +3799,16 @@ x_1 = lean_mk_string_unchecked(" = ", 3, 3); return x_1; } } -static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__6() { +static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__6() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__5; +x_1 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__5; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___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_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { _start: { if (lean_obj_tag(x_5) == 0) @@ -3467,7 +3852,7 @@ lean_inc(x_9); lean_inc(x_8); lean_inc(x_19); lean_inc(x_3); -x_22 = lean_grind_mk_eq_proof(x_3, x_19, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +x_22 = l_Lean_Meta_Grind_mkEqHEqProof(x_3, x_19, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); if (lean_obj_tag(x_22) == 0) { lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; @@ -3476,8 +3861,8 @@ lean_inc(x_23); x_24 = lean_ctor_get(x_22, 1); lean_inc(x_24); lean_dec(x_22); -x_25 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__4; -x_26 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_mkEqProofImpl_go___spec__1(x_25, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_24); +x_25 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__4; +x_26 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1(x_25, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_24); x_27 = lean_ctor_get(x_26, 0); lean_inc(x_27); x_28 = lean_unbox(x_27); @@ -3495,7 +3880,7 @@ lean_inc(x_15); lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); -x_31 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1(x_23, x_25, x_30, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_29); +x_31 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1(x_23, x_25, x_30, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_29); if (lean_obj_tag(x_31) == 0) { lean_object* x_32; @@ -3602,11 +3987,11 @@ x_48 = lean_ctor_get(x_26, 0); lean_dec(x_48); lean_inc(x_3); x_49 = l_Lean_MessageData_ofExpr(x_3); -x_50 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__5; +x_50 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__4; lean_ctor_set_tag(x_26, 7); lean_ctor_set(x_26, 1, x_49); lean_ctor_set(x_26, 0, x_50); -x_51 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__6; +x_51 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__6; lean_ctor_set_tag(x_5, 7); lean_ctor_set(x_5, 1, x_51); lean_ctor_set(x_5, 0, x_26); @@ -3617,7 +4002,7 @@ 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_50); -x_55 = l_Lean_addTrace___at_Lean_Meta_Grind_mkEqProofImpl_go___spec__2(x_25, x_54, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_47); +x_55 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2(x_25, x_54, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_47); x_56 = lean_ctor_get(x_55, 0); lean_inc(x_56); x_57 = lean_ctor_get(x_55, 1); @@ -3627,7 +4012,7 @@ lean_inc(x_15); lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); -x_58 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1(x_23, x_25, x_56, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_57); +x_58 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1(x_23, x_25, x_56, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_57); lean_dec(x_56); if (lean_obj_tag(x_58) == 0) { @@ -3731,11 +4116,11 @@ lean_inc(x_73); lean_dec(x_26); lean_inc(x_3); x_74 = l_Lean_MessageData_ofExpr(x_3); -x_75 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__5; +x_75 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__4; 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_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__6; +x_77 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__6; lean_ctor_set_tag(x_5, 7); lean_ctor_set(x_5, 1, x_77); lean_ctor_set(x_5, 0, x_76); @@ -3746,7 +4131,7 @@ lean_ctor_set(x_79, 1, x_78); x_80 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_80, 0, x_79); lean_ctor_set(x_80, 1, x_75); -x_81 = l_Lean_addTrace___at_Lean_Meta_Grind_mkEqProofImpl_go___spec__2(x_25, x_80, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_73); +x_81 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2(x_25, x_80, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_73); x_82 = lean_ctor_get(x_81, 0); lean_inc(x_82); x_83 = lean_ctor_get(x_81, 1); @@ -3756,7 +4141,7 @@ lean_inc(x_15); lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); -x_84 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1(x_23, x_25, x_82, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_83); +x_84 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1(x_23, x_25, x_82, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_83); lean_dec(x_82); if (lean_obj_tag(x_84) == 0) { @@ -3920,7 +4305,7 @@ lean_inc(x_9); lean_inc(x_8); lean_inc(x_103); lean_inc(x_3); -x_106 = lean_grind_mk_eq_proof(x_3, x_103, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +x_106 = l_Lean_Meta_Grind_mkEqHEqProof(x_3, x_103, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); 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; uint8_t x_112; @@ -3929,8 +4314,8 @@ lean_inc(x_107); x_108 = lean_ctor_get(x_106, 1); lean_inc(x_108); lean_dec(x_106); -x_109 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__4; -x_110 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_mkEqProofImpl_go___spec__1(x_109, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_108); +x_109 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__4; +x_110 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1(x_109, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_108); x_111 = lean_ctor_get(x_110, 0); lean_inc(x_111); x_112 = lean_unbox(x_111); @@ -3947,7 +4332,7 @@ lean_inc(x_15); lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); -x_115 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1(x_107, x_109, x_114, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_113); +x_115 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1(x_107, x_109, x_114, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_113); if (lean_obj_tag(x_115) == 0) { lean_object* x_116; @@ -4054,7 +4439,7 @@ if (lean_is_exclusive(x_110)) { } lean_inc(x_3); x_130 = l_Lean_MessageData_ofExpr(x_3); -x_131 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__5; +x_131 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__4; if (lean_is_scalar(x_129)) { x_132 = lean_alloc_ctor(7, 2, 0); } else { @@ -4063,7 +4448,7 @@ if (lean_is_scalar(x_129)) { } lean_ctor_set(x_132, 0, x_131); lean_ctor_set(x_132, 1, x_130); -x_133 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__6; +x_133 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__6; x_134 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_134, 0, x_132); lean_ctor_set(x_134, 1, x_133); @@ -4074,7 +4459,7 @@ lean_ctor_set(x_136, 1, x_135); x_137 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_137, 0, x_136); lean_ctor_set(x_137, 1, x_131); -x_138 = l_Lean_addTrace___at_Lean_Meta_Grind_mkEqProofImpl_go___spec__2(x_109, x_137, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_128); +x_138 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2(x_109, x_137, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_128); x_139 = lean_ctor_get(x_138, 0); lean_inc(x_139); x_140 = lean_ctor_get(x_138, 1); @@ -4084,7 +4469,7 @@ lean_inc(x_15); lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); -x_141 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1(x_107, x_109, x_139, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_140); +x_141 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1(x_107, x_109, x_139, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_140); lean_dec(x_139); if (lean_obj_tag(x_141) == 0) { @@ -4228,7 +4613,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___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_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { _start: { if (lean_obj_tag(x_4) == 0) @@ -4267,7 +4652,7 @@ lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_1); -x_20 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1(x_1, x_2, x_17, x_1, x_1, x_19, lean_box(0), x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +x_20 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3(x_1, x_2, x_17, x_1, x_1, x_19, lean_box(0), x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); if (lean_obj_tag(x_20) == 0) { lean_object* x_21; @@ -4315,7 +4700,7 @@ return x_26; } } } -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { _start: { if (lean_obj_tag(x_4) == 0) @@ -4354,7 +4739,7 @@ lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc_n(x_17, 2); -x_21 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2(x_17, x_19, x_17, x_17, x_20, lean_box(0), x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +x_21 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__4(x_17, x_19, x_17, x_17, x_20, lean_box(0), x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); lean_dec(x_17); if (lean_obj_tag(x_21) == 0) { @@ -4418,7 +4803,7 @@ lean_dec(x_10); x_13 = lean_box(0); x_14 = lean_box(0); lean_inc(x_11); -x_15 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3(x_11, x_13, x_11, x_11, x_14, lean_box(0), x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_12); +x_15 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__5(x_11, x_13, x_11, x_11, x_14, lean_box(0), x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_12); lean_dec(x_11); if (lean_obj_tag(x_15) == 0) { @@ -4499,11 +4884,43 @@ return x_27; } } } -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___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_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___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_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; +x_12 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_12; +} +} +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { lean_object* x_13; -x_13 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +x_13 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); @@ -4512,32 +4929,32 @@ lean_dec(x_3); return x_13; } } -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { _start: { lean_object* x_17; -x_17 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___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 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); return x_17; } } -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { _start: { lean_object* x_16; -x_16 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___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 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); lean_dec(x_3); lean_dec(x_2); return x_16; } } -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { _start: { lean_object* x_16; -x_16 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +x_16 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__5(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); @@ -5204,147 +5621,146 @@ x_13 = l_Lean_PersistentHashMap_foldlMAux___at_Lean_Meta_Grind_checkInvariants__ return x_13; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forIn___at_Lean_Meta_Grind_checkInvariants___spec__1___lambda__1(lean_object* x_1, lean_object* x_2, size_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forIn___at_Lean_Meta_Grind_checkInvariants___spec__1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { -lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_14 = lean_box_usize(x_3); -x_15 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_15, 0, x_14); -lean_ctor_set(x_15, 1, x_4); -x_16 = lean_apply_11(x_1, x_15, x_2, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -if (lean_obj_tag(x_16) == 0) +lean_object* x_14; lean_object* x_15; +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_3); +lean_ctor_set(x_14, 1, x_4); +x_15 = lean_apply_11(x_1, x_14, x_2, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_15) == 0) { -lean_object* x_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; -x_18 = !lean_is_exclusive(x_16); -if (x_18 == 0) +uint8_t x_17; +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_16); -if (x_28 == 0) +uint8_t x_27; +x_27 = !lean_is_exclusive(x_15); +if (x_27 == 0) { -lean_object* x_29; uint8_t x_30; -x_29 = lean_ctor_get(x_16, 0); -lean_dec(x_29); -x_30 = !lean_is_exclusive(x_17); -if (x_30 == 0) +lean_object* x_28; uint8_t x_29; +x_28 = lean_ctor_get(x_15, 0); +lean_dec(x_28); +x_29 = !lean_is_exclusive(x_16); +if (x_29 == 0) { -return x_16; +return x_15; } else { -lean_object* x_31; lean_object* x_32; -x_31 = lean_ctor_get(x_17, 0); -lean_inc(x_31); -lean_dec(x_17); -x_32 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_32, 0, x_31); -lean_ctor_set(x_16, 0, x_32); -return x_16; +lean_object* x_30; lean_object* x_31; +x_30 = lean_ctor_get(x_16, 0); +lean_inc(x_30); +lean_dec(x_16); +x_31 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_31, 0, x_30); +lean_ctor_set(x_15, 0, x_31); +return x_15; } } 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_16, 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_15, 1); +lean_inc(x_32); +lean_dec(x_15); +x_33 = lean_ctor_get(x_16, 0); lean_inc(x_33); -lean_dec(x_16); -x_34 = lean_ctor_get(x_17, 0); -lean_inc(x_34); -if (lean_is_exclusive(x_17)) { - lean_ctor_release(x_17, 0); - x_35 = x_17; +if (lean_is_exclusive(x_16)) { + lean_ctor_release(x_16, 0); + x_34 = x_16; } else { - lean_dec_ref(x_17); - x_35 = lean_box(0); + lean_dec_ref(x_16); + x_34 = lean_box(0); } -if (lean_is_scalar(x_35)) { - x_36 = lean_alloc_ctor(1, 1, 0); +if (lean_is_scalar(x_34)) { + x_35 = lean_alloc_ctor(1, 1, 0); } else { - x_36 = x_35; + x_35 = x_34; } -lean_ctor_set(x_36, 0, x_34); -x_37 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_37, 0, x_36); -lean_ctor_set(x_37, 1, x_33); -return x_37; +lean_ctor_set(x_35, 0, x_33); +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_38; -x_38 = !lean_is_exclusive(x_16); -if (x_38 == 0) +uint8_t x_37; +x_37 = !lean_is_exclusive(x_15); +if (x_37 == 0) { -return x_16; +return x_15; } else { -lean_object* x_39; lean_object* x_40; lean_object* x_41; -x_39 = lean_ctor_get(x_16, 0); -x_40 = lean_ctor_get(x_16, 1); -lean_inc(x_40); +lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_38 = lean_ctor_get(x_15, 0); +x_39 = lean_ctor_get(x_15, 1); lean_inc(x_39); -lean_dec(x_16); -x_41 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_41, 0, x_39); -lean_ctor_set(x_41, 1, x_40); -return x_41; +lean_inc(x_38); +lean_dec(x_15); +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; } } } @@ -5353,7 +5769,7 @@ LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forIn___at_Lean_Meta_Grind_che _start: { lean_object* x_13; lean_object* x_14; -x_13 = lean_alloc_closure((void*)(l_Lean_PersistentHashMap_forIn___at_Lean_Meta_Grind_checkInvariants___spec__1___lambda__1___boxed), 13, 1); +x_13 = lean_alloc_closure((void*)(l_Lean_PersistentHashMap_forIn___at_Lean_Meta_Grind_checkInvariants___spec__1___lambda__1), 13, 1); lean_closure_set(x_13, 0, x_3); x_14 = l_Lean_PersistentHashMap_foldlMAux___at_Lean_Meta_Grind_checkInvariants___spec__3___rarg(x_13, x_1, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); if (lean_obj_tag(x_14) == 0) @@ -5518,7 +5934,7 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_20 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__1; +x_20 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__1; lean_ctor_set(x_14, 0, x_20); return x_14; } @@ -5536,7 +5952,7 @@ if (x_22 == 0) lean_object* x_23; lean_object* x_24; x_23 = lean_ctor_get(x_21, 0); lean_dec(x_23); -x_24 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__1; +x_24 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__1; lean_ctor_set(x_21, 0, x_24); return x_21; } @@ -5546,7 +5962,7 @@ 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_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__1; +x_26 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__1; x_27 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_27, 0, x_26); lean_ctor_set(x_27, 1, x_25); @@ -5600,7 +6016,7 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_35 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__1; +x_35 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__1; x_36 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_36, 0, x_35); lean_ctor_set(x_36, 1, x_32); @@ -5623,7 +6039,7 @@ if (lean_is_exclusive(x_37)) { lean_dec_ref(x_37); x_39 = lean_box(0); } -x_40 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__1; +x_40 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__1; if (lean_is_scalar(x_39)) { x_41 = lean_alloc_ctor(0, 2, 0); } else { @@ -5875,16 +6291,6 @@ lean_dec(x_2); return x_16; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forIn___at_Lean_Meta_Grind_checkInvariants___spec__1___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { -_start: -{ -size_t x_14; lean_object* x_15; -x_14 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_15 = l_Lean_PersistentHashMap_forIn___at_Lean_Meta_Grind_checkInvariants___spec__1___lambda__1(x_1, x_2, x_14, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -return x_15; -} -} LEAN_EXPORT lean_object* l_Lean_Meta_Grind_checkInvariants___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: { @@ -5915,6 +6321,288 @@ x_12 = l_Lean_Meta_Grind_checkInvariants(x_11, x_2, x_3, x_4, x_5, x_6, x_7, x_8 return x_12; } } +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Goal_checkInvariants___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; uint8_t x_11; +x_10 = lean_st_mk_ref(x_1, x_9); +x_11 = !lean_is_exclusive(x_10); +if (x_11 == 0) +{ +return x_10; +} +else +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_12 = lean_ctor_get(x_10, 0); +x_13 = lean_ctor_get(x_10, 1); +lean_inc(x_13); +lean_inc(x_12); +lean_dec(x_10); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_12); +lean_ctor_set(x_14, 1, x_13); +return x_14; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Goal_checkInvariants___lambda__2(uint8_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +lean_inc(x_2); +x_11 = l_Lean_Meta_Grind_checkInvariants(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_11) == 0) +{ +lean_object* x_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 = lean_st_ref_get(x_2, x_12); +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_15 = lean_ctor_get(x_13, 1); +x_16 = lean_st_ref_get(x_2, x_15); +lean_dec(x_2); +x_17 = !lean_is_exclusive(x_16); +if (x_17 == 0) +{ +lean_object* x_18; +x_18 = lean_ctor_get(x_16, 0); +lean_ctor_set(x_13, 1, x_18); +lean_ctor_set(x_16, 0, x_13); +return x_16; +} +else +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_16, 0); +x_20 = lean_ctor_get(x_16, 1); +lean_inc(x_20); +lean_inc(x_19); +lean_dec(x_16); +lean_ctor_set(x_13, 1, x_19); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_13); +lean_ctor_set(x_21, 1, x_20); +return x_21; +} +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_22 = lean_ctor_get(x_13, 0); +x_23 = lean_ctor_get(x_13, 1); +lean_inc(x_23); +lean_inc(x_22); +lean_dec(x_13); +x_24 = lean_st_ref_get(x_2, x_23); +lean_dec(x_2); +x_25 = lean_ctor_get(x_24, 0); +lean_inc(x_25); +x_26 = lean_ctor_get(x_24, 1); +lean_inc(x_26); +if (lean_is_exclusive(x_24)) { + lean_ctor_release(x_24, 0); + lean_ctor_release(x_24, 1); + x_27 = x_24; +} else { + lean_dec_ref(x_24); + x_27 = lean_box(0); +} +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_22); +lean_ctor_set(x_28, 1, x_25); +if (lean_is_scalar(x_27)) { + x_29 = lean_alloc_ctor(0, 2, 0); +} else { + x_29 = x_27; +} +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_26); +return x_29; +} +} +else +{ +uint8_t x_30; +lean_dec(x_2); +x_30 = !lean_is_exclusive(x_11); +if (x_30 == 0) +{ +return x_11; +} +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_11, 0); +x_32 = lean_ctor_get(x_11, 1); +lean_inc(x_32); +lean_inc(x_31); +lean_dec(x_11); +x_33 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_33, 0, x_31); +lean_ctor_set(x_33, 1, x_32); +return x_33; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Goal_checkInvariants___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; +x_10 = !lean_is_exclusive(x_1); +if (x_10 == 0) +{ +lean_object* x_11; +x_11 = lean_ctor_get(x_1, 1); +lean_dec(x_11); +lean_ctor_set(x_1, 1, x_9); +return x_1; +} +else +{ +lean_object* x_12; lean_object* x_13; +x_12 = lean_ctor_get(x_1, 0); +lean_inc(x_12); +lean_dec(x_1); +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_9); +return x_13; +} +} +} +static lean_object* _init_l_Lean_Meta_Grind_Goal_checkInvariants___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_Goal_checkInvariants___lambda__3___boxed), 9, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Goal_checkInvariants(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; 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_1, 0); +lean_inc(x_11); +x_12 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_Goal_checkInvariants___lambda__1___boxed), 9, 1); +lean_closure_set(x_12, 0, x_1); +x_13 = lean_box(x_2); +x_14 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_Goal_checkInvariants___lambda__2___boxed), 10, 1); +lean_closure_set(x_14, 0, x_13); +x_15 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); +lean_closure_set(x_15, 0, x_12); +lean_closure_set(x_15, 1, x_14); +x_16 = l_Lean_Meta_Grind_Goal_checkInvariants___closed__1; +x_17 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); +lean_closure_set(x_17, 0, x_15); +lean_closure_set(x_17, 1, x_16); +x_18 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(x_11, x_17, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_18) == 0) +{ +uint8_t x_19; +x_19 = !lean_is_exclusive(x_18); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; +x_20 = lean_ctor_get(x_18, 0); +lean_dec(x_20); +x_21 = lean_box(0); +lean_ctor_set(x_18, 0, x_21); +return x_18; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_22 = lean_ctor_get(x_18, 1); +lean_inc(x_22); +lean_dec(x_18); +x_23 = lean_box(0); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set(x_24, 1, x_22); +return x_24; +} +} +else +{ +uint8_t x_25; +x_25 = !lean_is_exclusive(x_18); +if (x_25 == 0) +{ +return x_18; +} +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_EXPORT lean_object* l_Lean_Meta_Grind_Goal_checkInvariants___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l_Lean_Meta_Grind_Goal_checkInvariants___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Goal_checkInvariants___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +uint8_t x_11; lean_object* x_12; +x_11 = lean_unbox(x_1); +lean_dec(x_1); +x_12 = l_Lean_Meta_Grind_Goal_checkInvariants___lambda__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_Meta_Grind_Goal_checkInvariants___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l_Lean_Meta_Grind_Goal_checkInvariants___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Goal_checkInvariants___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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_Grind_Goal_checkInvariants(x_1, x_11, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_12; +} +} lean_object* initialize_Lean_Meta_Tactic_Grind_Types(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_Tactic_Grind_Proof(uint8_t builtin, lean_object*); static bool _G_initialized = false; @@ -6022,34 +6710,39 @@ l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Met lean_mark_persistent(l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__6); l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__7 = _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__7(); lean_mark_persistent(l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__7); -l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__1 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__1(); -lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__1); -l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__2 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__2(); -lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__2); -l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__3 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__3(); -lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__3); -l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__4 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__4(); -lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__4); -l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__5 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__5(); -lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__5); -l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__1 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__1(); -lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__1); -l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__2 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__2(); -lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__2); -l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__3 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__3(); -lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__3); -l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__4 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__4(); -lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__4); -l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__5 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__5(); -lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__5); -l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__6 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__6(); -lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__6); +l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__1 = _init_l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__1(); +l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__2 = _init_l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__2(); +lean_mark_persistent(l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__2); +l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__3 = _init_l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__3(); +lean_mark_persistent(l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__3); +l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__1 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__1(); +lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__1); +l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__2 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__2(); +lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__2); +l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__3 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__3(); +lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__3); +l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__4 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__4(); +lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__4); +l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__1 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__1(); +lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__1); +l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__2 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__2(); +lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__2); +l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__3 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__3(); +lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__3); +l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__4 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__4(); +lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__4); +l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__5 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__5(); +lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__5); +l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__6 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__6(); +lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__6); l_Lean_Meta_Grind_checkInvariants___lambda__1___closed__1 = _init_l_Lean_Meta_Grind_checkInvariants___lambda__1___closed__1(); lean_mark_persistent(l_Lean_Meta_Grind_checkInvariants___lambda__1___closed__1); l_Lean_Meta_Grind_checkInvariants___closed__1 = _init_l_Lean_Meta_Grind_checkInvariants___closed__1(); lean_mark_persistent(l_Lean_Meta_Grind_checkInvariants___closed__1); l_Lean_Meta_Grind_checkInvariants___closed__2 = _init_l_Lean_Meta_Grind_checkInvariants___closed__2(); lean_mark_persistent(l_Lean_Meta_Grind_checkInvariants___closed__2); +l_Lean_Meta_Grind_Goal_checkInvariants___closed__1 = _init_l_Lean_Meta_Grind_Goal_checkInvariants___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_Goal_checkInvariants___closed__1); return lean_io_result_mk_ok(lean_box(0)); } #ifdef __cplusplus diff --git a/stage0/stdlib/Lean/Meta/Tactic/Grind/Main.c b/stage0/stdlib/Lean/Meta/Tactic/Grind/Main.c new file mode 100644 index 000000000000..3b38da4bdebb --- /dev/null +++ b/stage0/stdlib/Lean/Meta/Tactic/Grind/Main.c @@ -0,0 +1,3559 @@ +// Lean compiler output +// Module: Lean.Meta.Tactic.Grind.Main +// Imports: Init.Grind.Lemmas Lean.Meta.Tactic.Util Lean.Meta.Tactic.Grind.RevertAll Lean.Meta.Tactic.Grind.PropagatorAttr Lean.Meta.Tactic.Grind.Proj Lean.Meta.Tactic.Grind.ForallProp Lean.Meta.Tactic.Grind.Util Lean.Meta.Tactic.Grind.Inv Lean.Meta.Tactic.Grind.Intro Lean.Meta.Tactic.Grind.EMatch +#include +#if defined(__clang__) +#pragma clang diagnostic ignored "-Wunused-parameter" +#pragma clang diagnostic ignored "-Wunused-label" +#elif defined(__GNUC__) && !defined(__CLANG__) +#pragma GCC diagnostic ignored "-Wunused-parameter" +#pragma GCC diagnostic ignored "-Wunused-label" +#pragma GCC diagnostic ignored "-Wunused-but-set-variable" +#endif +#ifdef __cplusplus +extern "C" { +#endif +lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkMethods(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_GrindM_run___rarg___closed__14; +lean_object* l_Lean_addTrace___at_Lean_Meta_Grind_simp___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple(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_Meta_Grind_ematchStar(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_Grind_preprocessAndProbe___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___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_Meta_Grind_GrindM_run___rarg___closed__20; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___closed__1; +static lean_object* l_Lean_Meta_Grind_GrindM_run___rarg___closed__11; +extern lean_object* l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedGoal___spec__2; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore___closed__2; +size_t lean_uint64_to_usize(uint64_t); +static lean_object* l_Lean_Meta_Grind_GrindM_run___rarg___closed__5; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___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_forM___at_Lean_Meta_Grind_preprocessAndProbe___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_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_mkMethods___spec__1___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_filterTR_loop___at___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore___spec__2(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_preprocessAndProbe(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_preprocessAndProbe___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_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore___closed__1; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___closed__3; +LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Meta_Grind_main___spec__1(lean_object*, lean_object*); +lean_object* l_Lean_stringToMessageData(lean_object*); +lean_object* l_Lean_MVarId_revertAll(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_getSimpCongrTheorems___rarg(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_main___lambda__2___closed__2; +lean_object* l_Std_Queue_empty(lean_object*); +lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_GrindM_run___rarg___closed__1; +lean_object* l_List_appendTR___rarg(lean_object*, lean_object*); +lean_object* l_Lean_Meta_Grind_getFalseExpr___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +size_t lean_usize_of_nat(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GrindM_run(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_GrindM_run___rarg___closed__10; +lean_object* l_Lean_Meta_Grind_Goal_checkInvariants(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_st_ref_take(lean_object*, lean_object*); +uint64_t lean_uint64_shift_right(uint64_t, uint64_t); +lean_object* l_Lean_Meta_Simp_SimprocExtension_getSimprocs(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___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_Meta_Grind_main___lambda__2___closed__3; +static lean_object* l_Lean_Meta_Grind_main___lambda__2___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkMethods___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* l_Lean_Meta_Grind_propagateProjEq(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_ofFormat(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_main___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_Grind_getTrueExpr___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___closed__2; +lean_object* lean_st_ref_get(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_GrindM_run___rarg___closed__13; +LEAN_EXPORT lean_object* l_Lean_Meta_withoutModifyingMCtx___at_Lean_Meta_Grind_preprocessAndProbe___spec__2(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___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_st_mk_ref(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_preprocessAndProbe___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_Grind_preprocessAndProbe___closed__1; +LEAN_EXPORT lean_object* l_List_forM___at___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore___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_Grind_main___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_GrindM_run___rarg___closed__15; +lean_object* l_ShareCommon_mkStateImpl(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkMethods___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 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_Grind_GrindM_run___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_Simp_mkContext(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +extern lean_object* l_Lean_Meta_Grind_builtinPropagatorsRef; +static lean_object* l_Lean_Meta_Grind_GrindM_run___rarg___closed__18; +static lean_object* l_Lean_Meta_Grind_main___lambda__2___closed__5; +lean_object* l_Lean_MVarId_clearAuxDecls(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_GrindM_run___rarg___closed__7; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_GrindM_run___rarg___closed__12; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___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_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkMethods___rarg(lean_object*); +extern lean_object* l_Lean_Meta_Grind_grindNormSimprocExt; +static lean_object* l_Lean_Meta_Grind_main___closed__1; +lean_object* l_Lean_Meta_SimpExtension_getTheorems(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_GrindM_run___rarg___closed__3; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_main___lambda__2___closed__4; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_main(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_GrindM_run___rarg___closed__6; +lean_object* l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___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_MVarId_ensureNoMVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_forM___at_Lean_Meta_Grind_preprocessAndProbe___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_Grind_mkMethods___boxed(lean_object*, lean_object*); +lean_object* l_Lean_Meta_Grind_intros(lean_object*, 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_Grind_main___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_PersistentHashMap_empty___at___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___spec__1; +static lean_object* l_Lean_Meta_Grind_GrindM_run___rarg___closed__19; +static lean_object* l_Lean_Meta_Grind_main___lambda__2___closed__6; +uint64_t l_Lean_Name_hash___override(lean_object*); +extern lean_object* l_Lean_Meta_Grind_grindNormExt; +uint64_t lean_uint64_xor(uint64_t, uint64_t); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_mkMethods___spec__1(lean_object*, lean_object*); +lean_object* l_Lean_MVarId_transformTarget(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_resetCache___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Expr_getAppFn(lean_object*); +lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_simp___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_Grind_GrindM_run___rarg___closed__17; +lean_object* l_Lean_Meta_Grind_mkENodeCore(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* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_withoutModifyingMCtx___at_Lean_Meta_Grind_preprocessAndProbe___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_List_reverse___rarg(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_all(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +size_t lean_usize_sub(size_t, size_t); +lean_object* lean_array_mk(lean_object*); +static lean_object* l_Lean_Meta_Grind_GrindM_run___rarg___closed__4; +LEAN_EXPORT lean_object* l_List_foldlM___at_Lean_Meta_Grind_all___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* l_Lean_MVarId_ensureProp(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_Meta_Grind_GrindM_run___rarg___closed__8; +lean_object* lean_array_uget(lean_object*, size_t); +lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_Grind_unfoldReducible(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_GrindM_run___rarg___closed__9; +lean_object* l_Lean_Meta_Grind_propagateForallProp(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_size(lean_object*); +extern lean_object* l_Lean_ShareCommon_objectFactory; +lean_object* lean_state_sharecommon(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_mkMethods___rarg___closed__1; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___closed__1; +static lean_object* l_Lean_Meta_Grind_GrindM_run___rarg___closed__16; +static lean_object* l_Lean_Meta_Grind_GrindM_run___rarg___closed__2; +lean_object* l_Lean_Meta_Grind_ppGoals(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_MVarId_betaReduce___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkMethods___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*); +size_t lean_usize_land(size_t, size_t); +lean_object* l_Lean_Meta_Grind_getEMatchTheorems___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_mkMethods___spec__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_3; +x_3 = lean_box(0); +return x_3; +} +else +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; +x_4 = lean_ctor_get(x_2, 0); +x_5 = lean_ctor_get(x_2, 1); +x_6 = lean_ctor_get(x_2, 2); +x_7 = lean_name_eq(x_4, x_1); +if (x_7 == 0) +{ +x_2 = x_6; +goto _start; +} +else +{ +lean_object* x_9; +lean_inc(x_5); +x_9 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_9, 0, x_5); +return x_9; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkMethods___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); +lean_inc(x_3); +lean_inc(x_2); +x_12 = l_Lean_Meta_Grind_propagateForallProp(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_12) == 0) +{ +uint8_t x_13; +x_13 = !lean_is_exclusive(x_12); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_14 = lean_ctor_get(x_12, 1); +x_15 = lean_ctor_get(x_12, 0); +lean_dec(x_15); +x_16 = l_Lean_Expr_getAppFn(x_2); +if (lean_obj_tag(x_16) == 4) +{ +lean_object* x_17; lean_object* x_18; +lean_free_object(x_12); +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +lean_dec(x_16); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +x_18 = l_Lean_Meta_Grind_propagateProjEq(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_14); +if (lean_obj_tag(x_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; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint64_t x_25; uint64_t x_26; uint64_t x_27; uint64_t x_28; uint64_t x_29; uint64_t x_30; uint64_t x_31; size_t x_32; size_t x_33; size_t x_34; size_t x_35; size_t x_36; lean_object* x_37; lean_object* x_38; +x_20 = lean_ctor_get(x_1, 0); +x_21 = lean_ctor_get(x_18, 1); +x_22 = lean_ctor_get(x_18, 0); +lean_dec(x_22); +x_23 = lean_ctor_get(x_20, 1); +x_24 = lean_array_get_size(x_23); +x_25 = l_Lean_Name_hash___override(x_17); +x_26 = 32; +x_27 = lean_uint64_shift_right(x_25, x_26); +x_28 = lean_uint64_xor(x_25, x_27); +x_29 = 16; +x_30 = lean_uint64_shift_right(x_28, x_29); +x_31 = lean_uint64_xor(x_28, x_30); +x_32 = lean_uint64_to_usize(x_31); +x_33 = lean_usize_of_nat(x_24); +lean_dec(x_24); +x_34 = 1; +x_35 = lean_usize_sub(x_33, x_34); +x_36 = lean_usize_land(x_32, x_35); +x_37 = lean_array_uget(x_23, x_36); +x_38 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_mkMethods___spec__1(x_17, x_37); +lean_dec(x_37); +lean_dec(x_17); +if (lean_obj_tag(x_38) == 0) +{ +lean_object* x_39; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_39 = lean_box(0); +lean_ctor_set(x_18, 0, x_39); +return x_18; +} +else +{ +lean_object* x_40; lean_object* x_41; +lean_free_object(x_18); +x_40 = lean_ctor_get(x_38, 0); +lean_inc(x_40); +lean_dec(x_38); +x_41 = lean_apply_10(x_40, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_21); +return x_41; +} +} +else +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; uint64_t x_46; uint64_t x_47; uint64_t x_48; uint64_t x_49; uint64_t x_50; uint64_t x_51; uint64_t x_52; size_t x_53; size_t x_54; size_t x_55; size_t x_56; size_t x_57; lean_object* x_58; lean_object* x_59; +x_42 = lean_ctor_get(x_1, 0); +x_43 = lean_ctor_get(x_18, 1); +lean_inc(x_43); +lean_dec(x_18); +x_44 = lean_ctor_get(x_42, 1); +x_45 = lean_array_get_size(x_44); +x_46 = l_Lean_Name_hash___override(x_17); +x_47 = 32; +x_48 = lean_uint64_shift_right(x_46, x_47); +x_49 = lean_uint64_xor(x_46, x_48); +x_50 = 16; +x_51 = lean_uint64_shift_right(x_49, x_50); +x_52 = lean_uint64_xor(x_49, x_51); +x_53 = lean_uint64_to_usize(x_52); +x_54 = lean_usize_of_nat(x_45); +lean_dec(x_45); +x_55 = 1; +x_56 = lean_usize_sub(x_54, x_55); +x_57 = lean_usize_land(x_53, x_56); +x_58 = lean_array_uget(x_44, x_57); +x_59 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_mkMethods___spec__1(x_17, x_58); +lean_dec(x_58); +lean_dec(x_17); +if (lean_obj_tag(x_59) == 0) +{ +lean_object* x_60; lean_object* x_61; +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_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_43); +return x_61; +} +else +{ +lean_object* x_62; lean_object* x_63; +x_62 = lean_ctor_get(x_59, 0); +lean_inc(x_62); +lean_dec(x_59); +x_63 = lean_apply_10(x_62, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_43); +return x_63; +} +} +} +else +{ +uint8_t x_64; +lean_dec(x_17); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_64 = !lean_is_exclusive(x_18); +if (x_64 == 0) +{ +return x_18; +} +else +{ +lean_object* x_65; lean_object* x_66; lean_object* x_67; +x_65 = lean_ctor_get(x_18, 0); +x_66 = lean_ctor_get(x_18, 1); +lean_inc(x_66); +lean_inc(x_65); +lean_dec(x_18); +x_67 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_67, 0, x_65); +lean_ctor_set(x_67, 1, x_66); +return x_67; +} +} +} +else +{ +lean_object* x_68; +lean_dec(x_16); +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_68 = lean_box(0); +lean_ctor_set(x_12, 0, x_68); +return x_12; +} +} +else +{ +lean_object* x_69; lean_object* x_70; +x_69 = lean_ctor_get(x_12, 1); +lean_inc(x_69); +lean_dec(x_12); +x_70 = l_Lean_Expr_getAppFn(x_2); +if (lean_obj_tag(x_70) == 4) +{ +lean_object* x_71; lean_object* x_72; +x_71 = lean_ctor_get(x_70, 0); +lean_inc(x_71); +lean_dec(x_70); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +x_72 = l_Lean_Meta_Grind_propagateProjEq(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_69); +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; uint64_t x_78; uint64_t x_79; uint64_t x_80; uint64_t x_81; uint64_t x_82; uint64_t x_83; uint64_t x_84; size_t x_85; size_t x_86; size_t x_87; size_t x_88; size_t x_89; lean_object* x_90; lean_object* x_91; +x_73 = lean_ctor_get(x_1, 0); +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_ctor_get(x_73, 1); +x_77 = lean_array_get_size(x_76); +x_78 = l_Lean_Name_hash___override(x_71); +x_79 = 32; +x_80 = lean_uint64_shift_right(x_78, x_79); +x_81 = lean_uint64_xor(x_78, x_80); +x_82 = 16; +x_83 = lean_uint64_shift_right(x_81, x_82); +x_84 = lean_uint64_xor(x_81, x_83); +x_85 = lean_uint64_to_usize(x_84); +x_86 = lean_usize_of_nat(x_77); +lean_dec(x_77); +x_87 = 1; +x_88 = lean_usize_sub(x_86, x_87); +x_89 = lean_usize_land(x_85, x_88); +x_90 = lean_array_uget(x_76, x_89); +x_91 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_mkMethods___spec__1(x_71, x_90); +lean_dec(x_90); +lean_dec(x_71); +if (lean_obj_tag(x_91) == 0) +{ +lean_object* x_92; lean_object* x_93; +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_92 = lean_box(0); +if (lean_is_scalar(x_75)) { + x_93 = lean_alloc_ctor(0, 2, 0); +} else { + x_93 = x_75; +} +lean_ctor_set(x_93, 0, x_92); +lean_ctor_set(x_93, 1, x_74); +return x_93; +} +else +{ +lean_object* x_94; lean_object* x_95; +lean_dec(x_75); +x_94 = lean_ctor_get(x_91, 0); +lean_inc(x_94); +lean_dec(x_91); +x_95 = lean_apply_10(x_94, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_74); +return x_95; +} +} +else +{ +lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; +lean_dec(x_71); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_96 = lean_ctor_get(x_72, 0); +lean_inc(x_96); +x_97 = lean_ctor_get(x_72, 1); +lean_inc(x_97); +if (lean_is_exclusive(x_72)) { + lean_ctor_release(x_72, 0); + lean_ctor_release(x_72, 1); + x_98 = x_72; +} else { + lean_dec_ref(x_72); + x_98 = lean_box(0); +} +if (lean_is_scalar(x_98)) { + x_99 = lean_alloc_ctor(1, 2, 0); +} else { + x_99 = x_98; +} +lean_ctor_set(x_99, 0, x_96); +lean_ctor_set(x_99, 1, x_97); +return x_99; +} +} +else +{ +lean_object* x_100; lean_object* x_101; +lean_dec(x_70); +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_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_69); +return x_101; +} +} +} +else +{ +uint8_t x_102; +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_102 = !lean_is_exclusive(x_12); +if (x_102 == 0) +{ +return x_12; +} +else +{ +lean_object* x_103; lean_object* x_104; lean_object* x_105; +x_103 = lean_ctor_get(x_12, 0); +x_104 = lean_ctor_get(x_12, 1); +lean_inc(x_104); +lean_inc(x_103); +lean_dec(x_12); +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_EXPORT lean_object* l_Lean_Meta_Grind_mkMethods___rarg___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; +x_12 = l_Lean_Expr_getAppFn(x_2); +if (lean_obj_tag(x_12) == 4) +{ +lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_13 = lean_ctor_get(x_1, 1); +lean_inc(x_13); +lean_dec(x_1); +x_14 = lean_ctor_get(x_12, 0); +lean_inc(x_14); +lean_dec(x_12); +x_15 = !lean_is_exclusive(x_13); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; uint64_t x_19; uint64_t x_20; uint64_t x_21; uint64_t x_22; uint64_t x_23; uint64_t x_24; uint64_t x_25; size_t x_26; size_t x_27; size_t x_28; size_t x_29; size_t x_30; lean_object* x_31; lean_object* x_32; +x_16 = lean_ctor_get(x_13, 1); +x_17 = lean_ctor_get(x_13, 0); +lean_dec(x_17); +x_18 = lean_array_get_size(x_16); +x_19 = l_Lean_Name_hash___override(x_14); +x_20 = 32; +x_21 = lean_uint64_shift_right(x_19, x_20); +x_22 = lean_uint64_xor(x_19, x_21); +x_23 = 16; +x_24 = lean_uint64_shift_right(x_22, x_23); +x_25 = lean_uint64_xor(x_22, x_24); +x_26 = lean_uint64_to_usize(x_25); +x_27 = lean_usize_of_nat(x_18); +lean_dec(x_18); +x_28 = 1; +x_29 = lean_usize_sub(x_27, x_28); +x_30 = lean_usize_land(x_26, x_29); +x_31 = lean_array_uget(x_16, x_30); +lean_dec(x_16); +x_32 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_mkMethods___spec__1(x_14, x_31); +lean_dec(x_31); +lean_dec(x_14); +if (lean_obj_tag(x_32) == 0) +{ +lean_object* x_33; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_33 = lean_box(0); +lean_ctor_set(x_13, 1, x_11); +lean_ctor_set(x_13, 0, x_33); +return x_13; +} +else +{ +lean_object* x_34; lean_object* x_35; +lean_free_object(x_13); +x_34 = lean_ctor_get(x_32, 0); +lean_inc(x_34); +lean_dec(x_32); +x_35 = lean_apply_10(x_34, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +return x_35; +} +} +else +{ +lean_object* x_36; lean_object* x_37; uint64_t x_38; uint64_t x_39; uint64_t x_40; uint64_t x_41; uint64_t x_42; uint64_t x_43; uint64_t x_44; size_t x_45; size_t x_46; size_t x_47; size_t x_48; size_t x_49; lean_object* x_50; lean_object* x_51; +x_36 = lean_ctor_get(x_13, 1); +lean_inc(x_36); +lean_dec(x_13); +x_37 = lean_array_get_size(x_36); +x_38 = l_Lean_Name_hash___override(x_14); +x_39 = 32; +x_40 = lean_uint64_shift_right(x_38, x_39); +x_41 = lean_uint64_xor(x_38, x_40); +x_42 = 16; +x_43 = lean_uint64_shift_right(x_41, x_42); +x_44 = lean_uint64_xor(x_41, x_43); +x_45 = lean_uint64_to_usize(x_44); +x_46 = lean_usize_of_nat(x_37); +lean_dec(x_37); +x_47 = 1; +x_48 = lean_usize_sub(x_46, x_47); +x_49 = lean_usize_land(x_45, x_48); +x_50 = lean_array_uget(x_36, x_49); +lean_dec(x_36); +x_51 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_mkMethods___spec__1(x_14, x_50); +lean_dec(x_50); +lean_dec(x_14); +if (lean_obj_tag(x_51) == 0) +{ +lean_object* x_52; lean_object* x_53; +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_52 = lean_box(0); +x_53 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_53, 0, x_52); +lean_ctor_set(x_53, 1, x_11); +return x_53; +} +else +{ +lean_object* x_54; lean_object* x_55; +x_54 = lean_ctor_get(x_51, 0); +lean_inc(x_54); +lean_dec(x_51); +x_55 = lean_apply_10(x_54, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +return x_55; +} +} +} +else +{ +lean_object* x_56; lean_object* x_57; +lean_dec(x_12); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_56 = lean_box(0); +x_57 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_57, 0, x_56); +lean_ctor_set(x_57, 1, x_11); +return x_57; +} +} +} +static lean_object* _init_l_Lean_Meta_Grind_mkMethods___rarg___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Meta_Grind_builtinPropagatorsRef; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkMethods___rarg(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; uint8_t x_4; +x_2 = l_Lean_Meta_Grind_mkMethods___rarg___closed__1; +x_3 = lean_st_ref_get(x_2, x_1); +x_4 = !lean_is_exclusive(x_3); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_5 = lean_ctor_get(x_3, 0); +lean_inc(x_5); +x_6 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_mkMethods___rarg___lambda__1___boxed), 11, 1); +lean_closure_set(x_6, 0, x_5); +x_7 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_mkMethods___rarg___lambda__2), 11, 1); +lean_closure_set(x_7, 0, x_5); +x_8 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_8, 0, x_6); +lean_ctor_set(x_8, 1, x_7); +lean_ctor_set(x_3, 0, x_8); +return x_3; +} +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 = lean_ctor_get(x_3, 0); +x_10 = lean_ctor_get(x_3, 1); +lean_inc(x_10); +lean_inc(x_9); +lean_dec(x_3); +lean_inc(x_9); +x_11 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_mkMethods___rarg___lambda__1___boxed), 11, 1); +lean_closure_set(x_11, 0, x_9); +x_12 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_mkMethods___rarg___lambda__2), 11, 1); +lean_closure_set(x_12, 0, x_9); +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_11); +lean_ctor_set(x_13, 1, x_12); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_10); +return x_14; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkMethods(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_mkMethods___rarg), 1, 0); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_mkMethods___spec__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_mkMethods___spec__1(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkMethods___rarg___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; +x_12 = l_Lean_Meta_Grind_mkMethods___rarg___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_1); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkMethods___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_Meta_Grind_mkMethods(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_ShareCommon_objectFactory; +x_2 = l_ShareCommon_mkStateImpl(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("False", 5, 5); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_GrindM_run___rarg___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_Grind_GrindM_run___rarg___closed__2; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind_GrindM_run___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_Meta_Grind_GrindM_run___rarg___closed__3; +x_3 = l_Lean_Expr_const___override(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_ShareCommon_objectFactory; +x_2 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__1; +x_3 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__4; +x_4 = lean_state_sharecommon(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("True", 4, 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__6; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind_GrindM_run___rarg___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_Grind_GrindM_run___rarg___closed__7; +x_3 = l_Lean_Expr_const___override(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__9() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Meta_Grind_grindNormExt; +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__10() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Meta_Grind_grindNormSimprocExt; +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; uint8_t x_3; uint8_t x_4; uint8_t x_5; lean_object* x_6; +x_1 = l_Lean_Meta_Simp_defaultMaxSteps; +x_2 = lean_unsigned_to_nat(2u); +x_3 = 0; +x_4 = 1; +x_5 = 0; +x_6 = lean_alloc_ctor(0, 2, 19); +lean_ctor_set(x_6, 0, x_1); +lean_ctor_set(x_6, 1, x_2); +lean_ctor_set_uint8(x_6, sizeof(void*)*2, x_3); +lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 1, x_4); +lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 2, x_3); +lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 3, x_4); +lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 4, x_4); +lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 5, x_4); +lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 6, x_5); +lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 7, x_4); +lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 8, x_4); +lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 9, x_3); +lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 10, x_4); +lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 11, x_3); +lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 12, x_4); +lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 13, x_4); +lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 14, x_3); +lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 15, x_3); +lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 16, x_3); +lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 17, x_4); +lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 18, x_4); +return x_6; +} +} +static lean_object* _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__12() { +_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_Grind_GrindM_run___rarg___closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__12; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__13; +x_2 = lean_alloc_ctor(0, 3, 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); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__15() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__13; +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_Grind_GrindM_run___rarg___closed__16() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(32u); +x_2 = lean_mk_empty_array_with_capacity(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__17() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__16; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__18() { +_start: +{ +size_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = 5; +x_2 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__17; +x_3 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__16; +x_4 = lean_unsigned_to_nat(0u); +x_5 = lean_alloc_ctor(0, 4, sizeof(size_t)*1); +lean_ctor_set(x_5, 0, x_2); +lean_ctor_set(x_5, 1, x_3); +lean_ctor_set(x_5, 2, x_4); +lean_ctor_set(x_5, 3, x_4); +lean_ctor_set_usize(x_5, 4, x_1); +return x_5; +} +} +static lean_object* _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__19() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__13; +x_2 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__18; +x_3 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_1); +lean_ctor_set(x_3, 2, x_1); +lean_ctor_set(x_3, 3, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__20() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__15; +x_2 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__19; +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_Grind_GrindM_run___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_9 = lean_box(0); +x_10 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__5; +x_11 = lean_ctor_get(x_10, 0); +lean_inc(x_11); +x_12 = lean_ctor_get(x_10, 1); +lean_inc(x_12); +x_13 = l_Lean_ShareCommon_objectFactory; +x_14 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__8; +x_15 = lean_state_sharecommon(x_13, x_12, x_14); +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_15, 1); +lean_inc(x_17); +lean_dec(x_15); +x_18 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__9; +x_19 = l_Lean_Meta_SimpExtension_getTheorems(x_18, x_6, x_7, x_8); +x_20 = lean_ctor_get(x_19, 0); +lean_inc(x_20); +x_21 = lean_ctor_get(x_19, 1); +lean_inc(x_21); +lean_dec(x_19); +x_22 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__10; +x_23 = l_Lean_Meta_Simp_SimprocExtension_getSimprocs(x_22, x_6, x_7, x_21); +x_24 = !lean_is_exclusive(x_23); +if (x_24 == 0) +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; +x_25 = lean_ctor_get(x_23, 1); +lean_ctor_set_tag(x_23, 1); +lean_ctor_set(x_23, 1, x_9); +x_26 = lean_array_mk(x_23); +x_27 = l_Lean_Meta_getSimpCongrTheorems___rarg(x_7, x_25); +x_28 = !lean_is_exclusive(x_27); +if (x_28 == 0) +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; 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_29 = lean_ctor_get(x_27, 0); +x_30 = lean_ctor_get(x_27, 1); +lean_ctor_set_tag(x_27, 1); +lean_ctor_set(x_27, 1, x_9); +lean_ctor_set(x_27, 0, x_20); +x_31 = lean_array_mk(x_27); +x_32 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__11; +x_33 = l_Lean_Meta_Simp_mkContext(x_32, x_31, x_29, x_4, x_5, x_6, x_7, x_30); +x_34 = lean_ctor_get(x_33, 0); +lean_inc(x_34); +x_35 = lean_ctor_get(x_33, 1); +lean_inc(x_35); +lean_dec(x_33); +x_36 = l_Lean_Meta_Grind_mkMethods___rarg(x_35); +x_37 = lean_ctor_get(x_36, 0); +lean_inc(x_37); +x_38 = lean_ctor_get(x_36, 1); +lean_inc(x_38); +lean_dec(x_36); +x_39 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_39, 0, x_34); +lean_ctor_set(x_39, 1, x_26); +lean_ctor_set(x_39, 2, x_2); +lean_ctor_set(x_39, 3, x_3); +x_40 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__14; +x_41 = lean_unsigned_to_nat(1u); +x_42 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__13; +x_43 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__20; +x_44 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_44, 0, x_40); +lean_ctor_set(x_44, 1, x_17); +lean_ctor_set(x_44, 2, x_41); +lean_ctor_set(x_44, 3, x_42); +lean_ctor_set(x_44, 4, x_43); +lean_ctor_set(x_44, 5, x_16); +lean_ctor_set(x_44, 6, x_11); +x_45 = lean_st_mk_ref(x_44, x_38); +x_46 = lean_ctor_get(x_45, 0); +lean_inc(x_46); +x_47 = lean_ctor_get(x_45, 1); +lean_inc(x_47); +lean_dec(x_45); +lean_inc(x_46); +x_48 = lean_apply_8(x_1, x_37, x_39, x_46, x_4, x_5, x_6, x_7, x_47); +if (lean_obj_tag(x_48) == 0) +{ +lean_object* x_49; lean_object* x_50; lean_object* x_51; uint8_t x_52; +x_49 = lean_ctor_get(x_48, 0); +lean_inc(x_49); +x_50 = lean_ctor_get(x_48, 1); +lean_inc(x_50); +lean_dec(x_48); +x_51 = lean_st_ref_get(x_46, x_50); +lean_dec(x_46); +x_52 = !lean_is_exclusive(x_51); +if (x_52 == 0) +{ +lean_object* x_53; +x_53 = lean_ctor_get(x_51, 0); +lean_dec(x_53); +lean_ctor_set(x_51, 0, x_49); +return x_51; +} +else +{ +lean_object* x_54; lean_object* x_55; +x_54 = lean_ctor_get(x_51, 1); +lean_inc(x_54); +lean_dec(x_51); +x_55 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_55, 0, x_49); +lean_ctor_set(x_55, 1, x_54); +return x_55; +} +} +else +{ +uint8_t x_56; +lean_dec(x_46); +x_56 = !lean_is_exclusive(x_48); +if (x_56 == 0) +{ +return x_48; +} +else +{ +lean_object* x_57; lean_object* x_58; lean_object* x_59; +x_57 = lean_ctor_get(x_48, 0); +x_58 = lean_ctor_get(x_48, 1); +lean_inc(x_58); +lean_inc(x_57); +lean_dec(x_48); +x_59 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_59, 0, x_57); +lean_ctor_set(x_59, 1, x_58); +return x_59; +} +} +} +else +{ +lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; +x_60 = lean_ctor_get(x_27, 0); +x_61 = lean_ctor_get(x_27, 1); +lean_inc(x_61); +lean_inc(x_60); +lean_dec(x_27); +x_62 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_62, 0, x_20); +lean_ctor_set(x_62, 1, x_9); +x_63 = lean_array_mk(x_62); +x_64 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__11; +x_65 = l_Lean_Meta_Simp_mkContext(x_64, x_63, x_60, x_4, x_5, x_6, x_7, x_61); +x_66 = lean_ctor_get(x_65, 0); +lean_inc(x_66); +x_67 = lean_ctor_get(x_65, 1); +lean_inc(x_67); +lean_dec(x_65); +x_68 = l_Lean_Meta_Grind_mkMethods___rarg(x_67); +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_alloc_ctor(0, 4, 0); +lean_ctor_set(x_71, 0, x_66); +lean_ctor_set(x_71, 1, x_26); +lean_ctor_set(x_71, 2, x_2); +lean_ctor_set(x_71, 3, x_3); +x_72 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__14; +x_73 = lean_unsigned_to_nat(1u); +x_74 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__13; +x_75 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__20; +x_76 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_76, 0, x_72); +lean_ctor_set(x_76, 1, x_17); +lean_ctor_set(x_76, 2, x_73); +lean_ctor_set(x_76, 3, x_74); +lean_ctor_set(x_76, 4, x_75); +lean_ctor_set(x_76, 5, x_16); +lean_ctor_set(x_76, 6, x_11); +x_77 = lean_st_mk_ref(x_76, x_70); +x_78 = lean_ctor_get(x_77, 0); +lean_inc(x_78); +x_79 = lean_ctor_get(x_77, 1); +lean_inc(x_79); +lean_dec(x_77); +lean_inc(x_78); +x_80 = lean_apply_8(x_1, x_69, x_71, x_78, x_4, x_5, x_6, x_7, x_79); +if (lean_obj_tag(x_80) == 0) +{ +lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; +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_st_ref_get(x_78, x_82); +lean_dec(x_78); +x_84 = lean_ctor_get(x_83, 1); +lean_inc(x_84); +if (lean_is_exclusive(x_83)) { + lean_ctor_release(x_83, 0); + lean_ctor_release(x_83, 1); + x_85 = x_83; +} else { + lean_dec_ref(x_83); + x_85 = lean_box(0); +} +if (lean_is_scalar(x_85)) { + x_86 = lean_alloc_ctor(0, 2, 0); +} else { + x_86 = x_85; +} +lean_ctor_set(x_86, 0, x_81); +lean_ctor_set(x_86, 1, x_84); +return x_86; +} +else +{ +lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; +lean_dec(x_78); +x_87 = lean_ctor_get(x_80, 0); +lean_inc(x_87); +x_88 = lean_ctor_get(x_80, 1); +lean_inc(x_88); +if (lean_is_exclusive(x_80)) { + lean_ctor_release(x_80, 0); + lean_ctor_release(x_80, 1); + x_89 = x_80; +} else { + lean_dec_ref(x_80); + x_89 = lean_box(0); +} +if (lean_is_scalar(x_89)) { + x_90 = lean_alloc_ctor(1, 2, 0); +} else { + x_90 = x_89; +} +lean_ctor_set(x_90, 0, x_87); +lean_ctor_set(x_90, 1, x_88); +return x_90; +} +} +} +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; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; +x_91 = lean_ctor_get(x_23, 0); +x_92 = lean_ctor_get(x_23, 1); +lean_inc(x_92); +lean_inc(x_91); +lean_dec(x_23); +x_93 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_93, 0, x_91); +lean_ctor_set(x_93, 1, x_9); +x_94 = lean_array_mk(x_93); +x_95 = l_Lean_Meta_getSimpCongrTheorems___rarg(x_7, x_92); +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_95); + x_98 = lean_box(0); +} +if (lean_is_scalar(x_98)) { + x_99 = lean_alloc_ctor(1, 2, 0); +} else { + x_99 = x_98; + lean_ctor_set_tag(x_99, 1); +} +lean_ctor_set(x_99, 0, x_20); +lean_ctor_set(x_99, 1, x_9); +x_100 = lean_array_mk(x_99); +x_101 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__11; +x_102 = l_Lean_Meta_Simp_mkContext(x_101, x_100, x_96, x_4, x_5, x_6, x_7, x_97); +x_103 = lean_ctor_get(x_102, 0); +lean_inc(x_103); +x_104 = lean_ctor_get(x_102, 1); +lean_inc(x_104); +lean_dec(x_102); +x_105 = l_Lean_Meta_Grind_mkMethods___rarg(x_104); +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(0, 4, 0); +lean_ctor_set(x_108, 0, x_103); +lean_ctor_set(x_108, 1, x_94); +lean_ctor_set(x_108, 2, x_2); +lean_ctor_set(x_108, 3, x_3); +x_109 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__14; +x_110 = lean_unsigned_to_nat(1u); +x_111 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__13; +x_112 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__20; +x_113 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_113, 0, x_109); +lean_ctor_set(x_113, 1, x_17); +lean_ctor_set(x_113, 2, x_110); +lean_ctor_set(x_113, 3, x_111); +lean_ctor_set(x_113, 4, x_112); +lean_ctor_set(x_113, 5, x_16); +lean_ctor_set(x_113, 6, x_11); +x_114 = lean_st_mk_ref(x_113, x_107); +x_115 = lean_ctor_get(x_114, 0); +lean_inc(x_115); +x_116 = lean_ctor_get(x_114, 1); +lean_inc(x_116); +lean_dec(x_114); +lean_inc(x_115); +x_117 = lean_apply_8(x_1, x_106, x_108, x_115, x_4, x_5, x_6, x_7, x_116); +if (lean_obj_tag(x_117) == 0) +{ +lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; +x_118 = lean_ctor_get(x_117, 0); +lean_inc(x_118); +x_119 = lean_ctor_get(x_117, 1); +lean_inc(x_119); +lean_dec(x_117); +x_120 = lean_st_ref_get(x_115, x_119); +lean_dec(x_115); +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); +} +if (lean_is_scalar(x_122)) { + x_123 = lean_alloc_ctor(0, 2, 0); +} else { + x_123 = x_122; +} +lean_ctor_set(x_123, 0, x_118); +lean_ctor_set(x_123, 1, x_121); +return x_123; +} +else +{ +lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; +lean_dec(x_115); +x_124 = lean_ctor_get(x_117, 0); +lean_inc(x_124); +x_125 = lean_ctor_get(x_117, 1); +lean_inc(x_125); +if (lean_is_exclusive(x_117)) { + lean_ctor_release(x_117, 0); + lean_ctor_release(x_117, 1); + x_126 = x_117; +} else { + lean_dec_ref(x_117); + x_126 = lean_box(0); +} +if (lean_is_scalar(x_126)) { + x_127 = lean_alloc_ctor(1, 2, 0); +} else { + x_127 = x_126; +} +lean_ctor_set(x_127, 0, x_124); +lean_ctor_set(x_127, 1, x_125); +return x_127; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GrindM_run(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_GrindM_run___rarg), 8, 0); +return x_2; +} +} +static lean_object* _init_l_Lean_PersistentHashMap_empty___at___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___spec__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__13; +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___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; uint8_t x_11; +x_10 = lean_st_mk_ref(x_1, x_9); +x_11 = !lean_is_exclusive(x_10); +if (x_11 == 0) +{ +return x_10; +} +else +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_12 = lean_ctor_get(x_10, 0); +x_13 = lean_ctor_get(x_10, 1); +lean_inc(x_13); +lean_inc(x_12); +lean_dec(x_10); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_12); +lean_ctor_set(x_14, 1, x_13); +return x_14; +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +uint8_t x_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; uint8_t x_20; +x_12 = 1; +x_13 = 0; +x_14 = lean_unsigned_to_nat(0u); +x_15 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_12, x_13, x_14, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_16 = lean_ctor_get(x_15, 1); +lean_inc(x_16); +lean_dec(x_15); +x_17 = l_Lean_Meta_Grind_mkENodeCore(x_2, x_12, x_13, x_14, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_16); +x_18 = lean_ctor_get(x_17, 1); +lean_inc(x_18); +lean_dec(x_17); +x_19 = lean_st_ref_get(x_3, x_18); +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, 1); +x_22 = lean_st_ref_get(x_3, x_21); +x_23 = !lean_is_exclusive(x_22); +if (x_23 == 0) +{ +lean_object* x_24; +x_24 = lean_ctor_get(x_22, 0); +lean_ctor_set(x_19, 1, x_24); +lean_ctor_set(x_22, 0, x_19); +return x_22; +} +else +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_25 = lean_ctor_get(x_22, 0); +x_26 = lean_ctor_get(x_22, 1); +lean_inc(x_26); +lean_inc(x_25); +lean_dec(x_22); +lean_ctor_set(x_19, 1, x_25); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_19); +lean_ctor_set(x_27, 1, x_26); +return x_27; +} +} +else +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_28 = lean_ctor_get(x_19, 0); +x_29 = lean_ctor_get(x_19, 1); +lean_inc(x_29); +lean_inc(x_28); +lean_dec(x_19); +x_30 = lean_st_ref_get(x_3, x_29); +x_31 = lean_ctor_get(x_30, 0); +lean_inc(x_31); +x_32 = lean_ctor_get(x_30, 1); +lean_inc(x_32); +if (lean_is_exclusive(x_30)) { + lean_ctor_release(x_30, 0); + lean_ctor_release(x_30, 1); + x_33 = x_30; +} else { + lean_dec_ref(x_30); + x_33 = lean_box(0); +} +x_34 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_34, 0, x_28); +lean_ctor_set(x_34, 1, x_31); +if (lean_is_scalar(x_33)) { + x_35 = lean_alloc_ctor(0, 2, 0); +} else { + x_35 = x_33; +} +lean_ctor_set(x_35, 0, x_34); +lean_ctor_set(x_35, 1, x_32); +return x_35; +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___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; +x_10 = !lean_is_exclusive(x_1); +if (x_10 == 0) +{ +lean_object* x_11; +x_11 = lean_ctor_get(x_1, 1); +lean_dec(x_11); +lean_ctor_set(x_1, 1, x_9); +return x_1; +} +else +{ +lean_object* x_12; lean_object* x_13; +x_12 = lean_ctor_get(x_1, 0); +lean_inc(x_12); +lean_dec(x_1); +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_9); +return x_13; +} +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(0); +x_2 = lean_array_mk(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Queue_empty(lean_box(0)); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___lambda__3___boxed), 9, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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; uint8_t x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_10 = l_Lean_Meta_Grind_getTrueExpr___rarg(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_Meta_Grind_getFalseExpr___rarg(x_4, x_5, x_6, x_7, x_8, x_12); +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +x_16 = l_Lean_Meta_Grind_getEMatchTheorems___rarg(x_8, 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 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__13; +x_20 = l_Lean_PersistentHashMap_empty___at___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___spec__1; +x_21 = l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___closed__1; +x_22 = 0; +x_23 = lean_unsigned_to_nat(0u); +x_24 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__18; +x_25 = l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedGoal___spec__2; +x_26 = l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___closed__2; +lean_inc(x_1); +x_27 = lean_alloc_ctor(0, 14, 1); +lean_ctor_set(x_27, 0, x_1); +lean_ctor_set(x_27, 1, x_19); +lean_ctor_set(x_27, 2, x_19); +lean_ctor_set(x_27, 3, x_20); +lean_ctor_set(x_27, 4, x_19); +lean_ctor_set(x_27, 5, x_21); +lean_ctor_set(x_27, 6, x_23); +lean_ctor_set(x_27, 7, x_23); +lean_ctor_set(x_27, 8, x_24); +lean_ctor_set(x_27, 9, x_24); +lean_ctor_set(x_27, 10, x_17); +lean_ctor_set(x_27, 11, x_23); +lean_ctor_set(x_27, 12, x_25); +lean_ctor_set(x_27, 13, x_26); +lean_ctor_set_uint8(x_27, sizeof(void*)*14, x_22); +x_28 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___lambda__1___boxed), 9, 1); +lean_closure_set(x_28, 0, x_27); +x_29 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___lambda__2___boxed), 11, 2); +lean_closure_set(x_29, 0, x_14); +lean_closure_set(x_29, 1, x_11); +x_30 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); +lean_closure_set(x_30, 0, x_28); +lean_closure_set(x_30, 1, x_29); +x_31 = l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___closed__3; +x_32 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); +lean_closure_set(x_32, 0, x_30); +lean_closure_set(x_32, 1, x_31); +x_33 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(x_1, x_32, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_18); +return x_33; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___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_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_10; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; +x_12 = l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_12; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___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_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_10; +} +} +LEAN_EXPORT lean_object* l_List_forM___at___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore___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: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_10; lean_object* 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_10 = lean_box(0); +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_10); +lean_ctor_set(x_11, 1, x_9); +return x_11; +} +else +{ +lean_object* x_12; lean_object* x_13; uint8_t x_14; lean_object* x_15; +x_12 = lean_ctor_get(x_1, 0); +lean_inc(x_12); +x_13 = lean_ctor_get(x_1, 1); +lean_inc(x_13); +lean_dec(x_1); +x_14 = 1; +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +x_15 = l_Lean_Meta_Grind_Goal_checkInvariants(x_12, x_14, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_15) == 0) +{ +lean_object* x_16; +x_16 = lean_ctor_get(x_15, 1); +lean_inc(x_16); +lean_dec(x_15); +x_1 = x_13; +x_9 = x_16; +goto _start; +} +else +{ +uint8_t x_18; +lean_dec(x_13); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_18 = !lean_is_exclusive(x_15); +if (x_18 == 0) +{ +return x_15; +} +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_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_EXPORT lean_object* l_List_filterTR_loop___at___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore___spec__2(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_3; +x_3 = l_List_reverse___rarg(x_2); +return x_3; +} +else +{ +lean_object* x_4; uint8_t x_5; +x_4 = lean_ctor_get(x_1, 0); +lean_inc(x_4); +x_5 = lean_ctor_get_uint8(x_4, sizeof(void*)*14); +if (x_5 == 0) +{ +uint8_t x_6; +x_6 = !lean_is_exclusive(x_1); +if (x_6 == 0) +{ +lean_object* x_7; lean_object* x_8; +x_7 = lean_ctor_get(x_1, 1); +x_8 = lean_ctor_get(x_1, 0); +lean_dec(x_8); +lean_ctor_set(x_1, 1, x_2); +{ +lean_object* _tmp_0 = x_7; +lean_object* _tmp_1 = x_1; +x_1 = _tmp_0; +x_2 = _tmp_1; +} +goto _start; +} +else +{ +lean_object* x_10; lean_object* x_11; +x_10 = lean_ctor_get(x_1, 1); +lean_inc(x_10); +lean_dec(x_1); +x_11 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_11, 0, x_4); +lean_ctor_set(x_11, 1, x_2); +x_1 = x_10; +x_2 = x_11; +goto _start; +} +} +else +{ +lean_object* x_13; +lean_dec(x_4); +x_13 = lean_ctor_get(x_1, 1); +lean_inc(x_13); +lean_dec(x_1); +x_1 = x_13; +goto _start; +} +} +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_unfoldReducible), 6, 0); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_MVarId_betaReduce___lambda__1___boxed), 6, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_1); +x_10 = l_Lean_MVarId_ensureProp(x_1, 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; +x_11 = lean_ctor_get(x_10, 1); +lean_inc(x_11); +lean_dec(x_10); +lean_inc(x_1); +x_12 = l_Lean_MVarId_ensureNoMVar(x_1, 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; +x_13 = lean_ctor_get(x_12, 1); +lean_inc(x_13); +lean_dec(x_12); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +x_14 = l_Lean_MVarId_clearAuxDecls(x_1, x_5, x_6, x_7, x_8, x_13); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +x_16 = lean_ctor_get(x_14, 1); +lean_inc(x_16); +lean_dec(x_14); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +x_17 = l_Lean_MVarId_revertAll(x_15, 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; 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___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore___closed__1; +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +x_21 = l_Lean_MVarId_transformTarget(x_18, x_20, x_5, x_6, x_7, x_8, x_19); +if (lean_obj_tag(x_21) == 0) +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +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___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore___closed__2; +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +x_25 = l_Lean_MVarId_transformTarget(x_22, x_24, x_5, x_6, x_7, x_8, x_23); +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); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +x_28 = l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal(x_26, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_27); +if (lean_obj_tag(x_28) == 0) +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_29 = lean_ctor_get(x_28, 0); +lean_inc(x_29); +x_30 = lean_ctor_get(x_28, 1); +lean_inc(x_30); +lean_dec(x_28); +x_31 = lean_unsigned_to_nat(0u); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +x_32 = l_Lean_Meta_Grind_intros(x_29, x_31, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_30); +if (lean_obj_tag(x_32) == 0) +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_33 = lean_ctor_get(x_32, 0); +lean_inc(x_33); +x_34 = lean_ctor_get(x_32, 1); +lean_inc(x_34); +lean_dec(x_32); +lean_inc(x_33); +x_35 = l_List_forM___at___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore___spec__1(x_33, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_34); +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; lean_object* x_39; +x_37 = lean_ctor_get(x_35, 0); +lean_dec(x_37); +x_38 = lean_box(0); +x_39 = l_List_filterTR_loop___at___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore___spec__2(x_33, x_38); +lean_ctor_set(x_35, 0, x_39); +return x_35; +} +else +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_40 = lean_ctor_get(x_35, 1); +lean_inc(x_40); +lean_dec(x_35); +x_41 = lean_box(0); +x_42 = l_List_filterTR_loop___at___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore___spec__2(x_33, x_41); +x_43 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_43, 0, x_42); +lean_ctor_set(x_43, 1, x_40); +return x_43; +} +} +else +{ +uint8_t x_44; +lean_dec(x_33); +x_44 = !lean_is_exclusive(x_35); +if (x_44 == 0) +{ +return x_35; +} +else +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_45 = lean_ctor_get(x_35, 0); +x_46 = lean_ctor_get(x_35, 1); +lean_inc(x_46); +lean_inc(x_45); +lean_dec(x_35); +x_47 = lean_alloc_ctor(1, 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_48; +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_48 = !lean_is_exclusive(x_32); +if (x_48 == 0) +{ +return x_32; +} +else +{ +lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_49 = lean_ctor_get(x_32, 0); +x_50 = lean_ctor_get(x_32, 1); +lean_inc(x_50); +lean_inc(x_49); +lean_dec(x_32); +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_52; +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_52 = !lean_is_exclusive(x_28); +if (x_52 == 0) +{ +return x_28; +} +else +{ +lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_53 = lean_ctor_get(x_28, 0); +x_54 = lean_ctor_get(x_28, 1); +lean_inc(x_54); +lean_inc(x_53); +lean_dec(x_28); +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_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); +lean_dec(x_2); +x_56 = !lean_is_exclusive(x_25); +if (x_56 == 0) +{ +return x_25; +} +else +{ +lean_object* x_57; lean_object* x_58; lean_object* x_59; +x_57 = lean_ctor_get(x_25, 0); +x_58 = lean_ctor_get(x_25, 1); +lean_inc(x_58); +lean_inc(x_57); +lean_dec(x_25); +x_59 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_59, 0, x_57); +lean_ctor_set(x_59, 1, x_58); +return x_59; +} +} +} +else +{ +uint8_t x_60; +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_60 = !lean_is_exclusive(x_21); +if (x_60 == 0) +{ +return x_21; +} +else +{ +lean_object* x_61; lean_object* x_62; lean_object* x_63; +x_61 = lean_ctor_get(x_21, 0); +x_62 = lean_ctor_get(x_21, 1); +lean_inc(x_62); +lean_inc(x_61); +lean_dec(x_21); +x_63 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_63, 0, x_61); +lean_ctor_set(x_63, 1, x_62); +return x_63; +} +} +} +else +{ +uint8_t x_64; +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_64 = !lean_is_exclusive(x_17); +if (x_64 == 0) +{ +return x_17; +} +else +{ +lean_object* x_65; lean_object* x_66; lean_object* x_67; +x_65 = lean_ctor_get(x_17, 0); +x_66 = lean_ctor_get(x_17, 1); +lean_inc(x_66); +lean_inc(x_65); +lean_dec(x_17); +x_67 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_67, 0, x_65); +lean_ctor_set(x_67, 1, x_66); +return x_67; +} +} +} +else +{ +uint8_t x_68; +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_68 = !lean_is_exclusive(x_14); +if (x_68 == 0) +{ +return x_14; +} +else +{ +lean_object* x_69; lean_object* x_70; lean_object* x_71; +x_69 = lean_ctor_get(x_14, 0); +x_70 = lean_ctor_get(x_14, 1); +lean_inc(x_70); +lean_inc(x_69); +lean_dec(x_14); +x_71 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_71, 0, x_69); +lean_ctor_set(x_71, 1, x_70); +return x_71; +} +} +} +else +{ +uint8_t x_72; +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_72 = !lean_is_exclusive(x_12); +if (x_72 == 0) +{ +return x_12; +} +else +{ +lean_object* x_73; lean_object* x_74; lean_object* x_75; +x_73 = lean_ctor_get(x_12, 0); +x_74 = lean_ctor_get(x_12, 1); +lean_inc(x_74); +lean_inc(x_73); +lean_dec(x_12); +x_75 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_75, 0, x_73); +lean_ctor_set(x_75, 1, x_74); +return x_75; +} +} +} +else +{ +uint8_t x_76; +lean_dec(x_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_76 = !lean_is_exclusive(x_10); +if (x_76 == 0) +{ +return x_10; +} +else +{ +lean_object* x_77; lean_object* x_78; lean_object* x_79; +x_77 = lean_ctor_get(x_10, 0); +x_78 = lean_ctor_get(x_10, 1); +lean_inc(x_78); +lean_inc(x_77); +lean_dec(x_10); +x_79 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_79, 0, x_77); +lean_ctor_set(x_79, 1, x_78); +return x_79; +} +} +} +} +LEAN_EXPORT lean_object* l_List_foldlM___at_Lean_Meta_Grind_all___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) { +_start: +{ +if (lean_obj_tag(x_3) == 0) +{ +lean_object* x_12; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_2); +lean_ctor_set(x_12, 1, x_11); +return x_12; +} +else +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_13 = lean_ctor_get(x_3, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_3, 1); +lean_inc(x_14); +lean_dec(x_3); +lean_inc(x_1); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +x_15 = lean_apply_9(x_1, x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_15) == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_15, 1); +lean_inc(x_17); +lean_dec(x_15); +x_18 = l_List_appendTR___rarg(x_2, x_16); +x_2 = x_18; +x_3 = x_14; +x_11 = x_17; +goto _start; +} +else +{ +uint8_t x_20; +lean_dec(x_14); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_20 = !lean_is_exclusive(x_15); +if (x_20 == 0) +{ +return x_15; +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_15, 0); +x_22 = lean_ctor_get(x_15, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_15); +x_23 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_23, 0, x_21); +lean_ctor_set(x_23, 1, x_22); +return x_23; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_all(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; lean_object* x_12; +x_11 = lean_box(0); +x_12 = l_List_foldlM___at_Lean_Meta_Grind_all___spec__1(x_2, x_11, x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_12; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_ematchStar), 9, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; lean_object* x_11; +x_10 = l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___closed__1; +x_11 = l_Lean_Meta_Grind_all(x_1, 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_List_mapTR_loop___at_Lean_Meta_Grind_main___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 = lean_ctor_get(x_5, 0); +lean_inc(x_7); +lean_dec(x_5); +lean_ctor_set(x_1, 1, x_2); +lean_ctor_set(x_1, 0, x_7); +{ +lean_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 = lean_ctor_get(x_9, 0); +lean_inc(x_11); +lean_dec(x_9); +x_12 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_2); +x_1 = x_10; +x_2 = x_12; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_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, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_11 = lean_box(0); +x_12 = l_List_mapTR_loop___at_Lean_Meta_Grind_main___spec__1(x_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_10); +return x_13; +} +} +static lean_object* _init_l_Lean_Meta_Grind_main___lambda__2___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("grind", 5, 5); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_main___lambda__2___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("debug", 5, 5); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_main___lambda__2___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("final", 5, 5); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_main___lambda__2___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Meta_Grind_main___lambda__2___closed__1; +x_2 = l_Lean_Meta_Grind_main___lambda__2___closed__2; +x_3 = l_Lean_Meta_Grind_main___lambda__2___closed__3; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Meta_Grind_main___lambda__2___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("", 0, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_main___lambda__2___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_main___lambda__2___closed__5; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_main___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; lean_object* x_11; +x_10 = l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___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); +lean_inc(x_2); +x_11 = l_Lean_Meta_Grind_all(x_1, x_10, x_2, 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; 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 = l_Lean_Meta_Grind_main___lambda__2___closed__4; +x_15 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_simp___spec__2(x_14, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_13); +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; 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_Meta_Grind_main___lambda__1(x_12, x_19, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_18); +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_20; +} +else +{ +uint8_t x_21; +x_21 = !lean_is_exclusive(x_15); +if (x_21 == 0) +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_22 = lean_ctor_get(x_15, 1); +x_23 = lean_ctor_get(x_15, 0); +lean_dec(x_23); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_12); +x_24 = l_Lean_Meta_Grind_ppGoals(x_12, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_22); +if (lean_obj_tag(x_24) == 0) +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; +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_MessageData_ofFormat(x_25); +x_28 = l_Lean_Meta_Grind_main___lambda__2___closed__6; +lean_ctor_set_tag(x_15, 7); +lean_ctor_set(x_15, 1, x_27); +lean_ctor_set(x_15, 0, x_28); +x_29 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_29, 0, x_15); +lean_ctor_set(x_29, 1, x_28); +x_30 = l_Lean_addTrace___at_Lean_Meta_Grind_simp___spec__3(x_14, x_29, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_26); +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_Grind_main___lambda__1(x_12, x_31, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_32); +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_31); +return x_33; +} +else +{ +uint8_t x_34; +lean_free_object(x_15); +lean_dec(x_12); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_34 = !lean_is_exclusive(x_24); +if (x_34 == 0) +{ +return x_24; +} +else +{ +lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_35 = lean_ctor_get(x_24, 0); +x_36 = lean_ctor_get(x_24, 1); +lean_inc(x_36); +lean_inc(x_35); +lean_dec(x_24); +x_37 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_37, 0, x_35); +lean_ctor_set(x_37, 1, x_36); +return x_37; +} +} +} +else +{ +lean_object* x_38; lean_object* x_39; +x_38 = lean_ctor_get(x_15, 1); +lean_inc(x_38); +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); +lean_inc(x_2); +lean_inc(x_12); +x_39 = l_Lean_Meta_Grind_ppGoals(x_12, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_38); +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; lean_object* x_48; lean_object* x_49; +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_MessageData_ofFormat(x_40); +x_43 = l_Lean_Meta_Grind_main___lambda__2___closed__6; +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 = 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_Lean_addTrace___at_Lean_Meta_Grind_simp___spec__3(x_14, x_45, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_41); +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_Meta_Grind_main___lambda__1(x_12, x_47, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_48); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_47); +return x_49; +} +else +{ +lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* 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); +lean_dec(x_3); +lean_dec(x_2); +x_50 = lean_ctor_get(x_39, 0); +lean_inc(x_50); +x_51 = lean_ctor_get(x_39, 1); +lean_inc(x_51); +if (lean_is_exclusive(x_39)) { + lean_ctor_release(x_39, 0); + lean_ctor_release(x_39, 1); + x_52 = x_39; +} else { + lean_dec_ref(x_39); + x_52 = lean_box(0); +} +if (lean_is_scalar(x_52)) { + x_53 = lean_alloc_ctor(1, 2, 0); +} else { + x_53 = x_52; +} +lean_ctor_set(x_53, 0, x_50); +lean_ctor_set(x_53, 1, x_51); +return x_53; +} +} +} +} +else +{ +uint8_t x_54; +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_54 = !lean_is_exclusive(x_11); +if (x_54 == 0) +{ +return x_11; +} +else +{ +lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_55 = lean_ctor_get(x_11, 0); +x_56 = lean_ctor_get(x_11, 1); +lean_inc(x_56); +lean_inc(x_55); +lean_dec(x_11); +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; +} +} +} +} +static lean_object* _init_l_Lean_Meta_Grind_main___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_main___lambda__2), 9, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_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) { +_start: +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_9 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore), 9, 1); +lean_closure_set(x_9, 0, x_1); +x_10 = l_Lean_Meta_Grind_main___closed__1; +x_11 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); +lean_closure_set(x_11, 0, x_9); +lean_closure_set(x_11, 1, x_10); +x_12 = l_Lean_Meta_Grind_GrindM_run___rarg(x_11, x_3, x_2, x_4, x_5, x_6, x_7, x_8); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_main___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_Meta_Grind_main___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_11; +} +} +LEAN_EXPORT lean_object* l_List_forM___at_Lean_Meta_Grind_preprocessAndProbe___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) { +_start: +{ +lean_object* x_11; +lean_inc(x_2); +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) +{ +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 = lean_st_ref_get(x_2, x_12); +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_15 = lean_ctor_get(x_13, 1); +x_16 = lean_st_ref_get(x_2, x_15); +lean_dec(x_2); +x_17 = !lean_is_exclusive(x_16); +if (x_17 == 0) +{ +lean_object* x_18; +x_18 = lean_ctor_get(x_16, 0); +lean_ctor_set(x_13, 1, x_18); +lean_ctor_set(x_16, 0, x_13); +return x_16; +} +else +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_16, 0); +x_20 = lean_ctor_get(x_16, 1); +lean_inc(x_20); +lean_inc(x_19); +lean_dec(x_16); +lean_ctor_set(x_13, 1, x_19); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_13); +lean_ctor_set(x_21, 1, x_20); +return x_21; +} +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_22 = lean_ctor_get(x_13, 0); +x_23 = lean_ctor_get(x_13, 1); +lean_inc(x_23); +lean_inc(x_22); +lean_dec(x_13); +x_24 = lean_st_ref_get(x_2, x_23); +lean_dec(x_2); +x_25 = lean_ctor_get(x_24, 0); +lean_inc(x_25); +x_26 = lean_ctor_get(x_24, 1); +lean_inc(x_26); +if (lean_is_exclusive(x_24)) { + lean_ctor_release(x_24, 0); + lean_ctor_release(x_24, 1); + x_27 = x_24; +} else { + lean_dec_ref(x_24); + x_27 = lean_box(0); +} +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_22); +lean_ctor_set(x_28, 1, x_25); +if (lean_is_scalar(x_27)) { + x_29 = lean_alloc_ctor(0, 2, 0); +} else { + x_29 = x_27; +} +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_26); +return x_29; +} +} +else +{ +uint8_t x_30; +lean_dec(x_2); +x_30 = !lean_is_exclusive(x_11); +if (x_30 == 0) +{ +return x_11; +} +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_11, 0); +x_32 = lean_ctor_get(x_11, 1); +lean_inc(x_32); +lean_inc(x_31); +lean_dec(x_11); +x_33 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_33, 0, x_31); +lean_ctor_set(x_33, 1, x_32); +return x_33; +} +} +} +} +LEAN_EXPORT lean_object* l_List_forM___at_Lean_Meta_Grind_preprocessAndProbe___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_2) == 0) +{ +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_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_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_13 = lean_ctor_get(x_2, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_2, 1); +lean_inc(x_14); +lean_dec(x_2); +x_15 = lean_ctor_get(x_13, 0); +lean_inc(x_15); +x_16 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___lambda__1___boxed), 9, 1); +lean_closure_set(x_16, 0, x_13); +lean_inc(x_1); +x_17 = lean_alloc_closure((void*)(l_List_forM___at_Lean_Meta_Grind_preprocessAndProbe___spec__1___lambda__1), 10, 1); +lean_closure_set(x_17, 0, x_1); +x_18 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); +lean_closure_set(x_18, 0, x_16); +lean_closure_set(x_18, 1, x_17); +x_19 = l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___closed__3; +x_20 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); +lean_closure_set(x_20, 0, x_18); +lean_closure_set(x_20, 1, x_19); +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_21 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(x_15, x_20, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_21) == 0) +{ +lean_object* x_22; +x_22 = lean_ctor_get(x_21, 1); +lean_inc(x_22); +lean_dec(x_21); +x_2 = x_14; +x_10 = x_22; +goto _start; +} +else +{ +uint8_t x_24; +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_24 = !lean_is_exclusive(x_21); +if (x_24 == 0) +{ +return x_21; +} +else +{ +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_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; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_withoutModifyingMCtx___at_Lean_Meta_Grind_preprocessAndProbe___spec__2___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_7 = lean_st_ref_get(x_3, x_6); +x_8 = lean_ctor_get(x_7, 0); +lean_inc(x_8); +x_9 = lean_ctor_get(x_7, 1); +lean_inc(x_9); +lean_dec(x_7); +x_10 = lean_ctor_get(x_8, 0); +lean_inc(x_10); +lean_dec(x_8); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_11 = lean_apply_5(x_1, x_2, x_3, x_4, x_5, x_9); +if (lean_obj_tag(x_11) == 0) +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +x_13 = lean_ctor_get(x_11, 1); +lean_inc(x_13); +lean_dec(x_11); +x_14 = l_Lean_Meta_resetCache___rarg(x_3, x_4, x_5, x_13); +lean_dec(x_5); +lean_dec(x_4); +x_15 = lean_ctor_get(x_14, 1); +lean_inc(x_15); +lean_dec(x_14); +x_16 = lean_st_ref_take(x_3, x_15); +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +x_18 = lean_ctor_get(x_16, 1); +lean_inc(x_18); +lean_dec(x_16); +x_19 = !lean_is_exclusive(x_17); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_20 = lean_ctor_get(x_17, 0); +lean_dec(x_20); +lean_ctor_set(x_17, 0, x_10); +x_21 = lean_st_ref_set(x_3, x_17, x_18); +lean_dec(x_3); +x_22 = !lean_is_exclusive(x_21); +if (x_22 == 0) +{ +lean_object* x_23; +x_23 = lean_ctor_get(x_21, 0); +lean_dec(x_23); +lean_ctor_set(x_21, 0, x_12); +return x_21; +} +else +{ +lean_object* x_24; lean_object* x_25; +x_24 = lean_ctor_get(x_21, 1); +lean_inc(x_24); +lean_dec(x_21); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_12); +lean_ctor_set(x_25, 1, x_24); +return x_25; +} +} +else +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_26 = lean_ctor_get(x_17, 1); +x_27 = lean_ctor_get(x_17, 2); +x_28 = lean_ctor_get(x_17, 3); +x_29 = lean_ctor_get(x_17, 4); +lean_inc(x_29); +lean_inc(x_28); +lean_inc(x_27); +lean_inc(x_26); +lean_dec(x_17); +x_30 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_30, 0, x_10); +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_18); +lean_dec(x_3); +x_32 = lean_ctor_get(x_31, 1); +lean_inc(x_32); +if (lean_is_exclusive(x_31)) { + lean_ctor_release(x_31, 0); + lean_ctor_release(x_31, 1); + x_33 = x_31; +} else { + lean_dec_ref(x_31); + x_33 = lean_box(0); +} +if (lean_is_scalar(x_33)) { + x_34 = lean_alloc_ctor(0, 2, 0); +} else { + x_34 = x_33; +} +lean_ctor_set(x_34, 0, x_12); +lean_ctor_set(x_34, 1, x_32); +return x_34; +} +} +else +{ +lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; uint8_t x_42; +x_35 = lean_ctor_get(x_11, 0); +lean_inc(x_35); +x_36 = lean_ctor_get(x_11, 1); +lean_inc(x_36); +lean_dec(x_11); +x_37 = l_Lean_Meta_resetCache___rarg(x_3, x_4, x_5, x_36); +lean_dec(x_5); +lean_dec(x_4); +x_38 = lean_ctor_get(x_37, 1); +lean_inc(x_38); +lean_dec(x_37); +x_39 = lean_st_ref_take(x_3, x_38); +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_is_exclusive(x_40); +if (x_42 == 0) +{ +lean_object* x_43; lean_object* x_44; uint8_t x_45; +x_43 = lean_ctor_get(x_40, 0); +lean_dec(x_43); +lean_ctor_set(x_40, 0, x_10); +x_44 = lean_st_ref_set(x_3, x_40, x_41); +lean_dec(x_3); +x_45 = !lean_is_exclusive(x_44); +if (x_45 == 0) +{ +lean_object* x_46; +x_46 = lean_ctor_get(x_44, 0); +lean_dec(x_46); +lean_ctor_set_tag(x_44, 1); +lean_ctor_set(x_44, 0, x_35); +return x_44; +} +else +{ +lean_object* x_47; lean_object* x_48; +x_47 = lean_ctor_get(x_44, 1); +lean_inc(x_47); +lean_dec(x_44); +x_48 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_48, 0, x_35); +lean_ctor_set(x_48, 1, x_47); +return x_48; +} +} +else +{ +lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_49 = lean_ctor_get(x_40, 1); +x_50 = lean_ctor_get(x_40, 2); +x_51 = lean_ctor_get(x_40, 3); +x_52 = lean_ctor_get(x_40, 4); +lean_inc(x_52); +lean_inc(x_51); +lean_inc(x_50); +lean_inc(x_49); +lean_dec(x_40); +x_53 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_53, 0, x_10); +lean_ctor_set(x_53, 1, x_49); +lean_ctor_set(x_53, 2, x_50); +lean_ctor_set(x_53, 3, x_51); +lean_ctor_set(x_53, 4, x_52); +x_54 = lean_st_ref_set(x_3, x_53, x_41); +lean_dec(x_3); +x_55 = lean_ctor_get(x_54, 1); +lean_inc(x_55); +if (lean_is_exclusive(x_54)) { + lean_ctor_release(x_54, 0); + lean_ctor_release(x_54, 1); + x_56 = x_54; +} else { + lean_dec_ref(x_54); + x_56 = lean_box(0); +} +if (lean_is_scalar(x_56)) { + x_57 = lean_alloc_ctor(1, 2, 0); +} else { + x_57 = x_56; + lean_ctor_set_tag(x_57, 1); +} +lean_ctor_set(x_57, 0, x_35); +lean_ctor_set(x_57, 1, x_55); +return x_57; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_withoutModifyingMCtx___at_Lean_Meta_Grind_preprocessAndProbe___spec__2(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_Meta_withoutModifyingMCtx___at_Lean_Meta_Grind_preprocessAndProbe___spec__2___rarg), 6, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_preprocessAndProbe___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; +x_12 = l_List_forM___at_Lean_Meta_Grind_preprocessAndProbe___spec__1(x_1, x_2, x_4, x_5, 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_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 +{ +uint8_t x_19; +x_19 = !lean_is_exclusive(x_12); +if (x_19 == 0) +{ +return x_12; +} +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_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_EXPORT lean_object* l_Lean_Meta_Grind_preprocessAndProbe___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_11 = l_Lean_Meta_Grind_main___lambda__2___closed__4; +x_12 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_simp___spec__2(x_11, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_unbox(x_13); +lean_dec(x_13); +if (x_14 == 0) +{ +lean_object* x_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_Meta_Grind_preprocessAndProbe___lambda__1(x_1, x_2, x_16, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_15); +return x_17; +} +else +{ +uint8_t x_18; +x_18 = !lean_is_exclusive(x_12); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_12, 1); +x_20 = lean_ctor_get(x_12, 0); +lean_dec(x_20); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +x_21 = l_Lean_Meta_Grind_ppGoals(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_19); +if (lean_obj_tag(x_21) == 0) +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_22 = lean_ctor_get(x_21, 0); +lean_inc(x_22); +x_23 = lean_ctor_get(x_21, 1); +lean_inc(x_23); +lean_dec(x_21); +x_24 = l_Lean_MessageData_ofFormat(x_22); +x_25 = l_Lean_Meta_Grind_main___lambda__2___closed__6; +lean_ctor_set_tag(x_12, 7); +lean_ctor_set(x_12, 1, x_24); +lean_ctor_set(x_12, 0, x_25); +x_26 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_26, 0, x_12); +lean_ctor_set(x_26, 1, x_25); +x_27 = l_Lean_addTrace___at_Lean_Meta_Grind_simp___spec__3(x_11, x_26, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_23); +x_28 = lean_ctor_get(x_27, 0); +lean_inc(x_28); +x_29 = lean_ctor_get(x_27, 1); +lean_inc(x_29); +lean_dec(x_27); +x_30 = l_Lean_Meta_Grind_preprocessAndProbe___lambda__1(x_1, x_2, x_28, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_29); +lean_dec(x_28); +return x_30; +} +else +{ +uint8_t x_31; +lean_free_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_2); +lean_dec(x_1); +x_31 = !lean_is_exclusive(x_21); +if (x_31 == 0) +{ +return x_21; +} +else +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_21, 0); +x_33 = lean_ctor_get(x_21, 1); +lean_inc(x_33); +lean_inc(x_32); +lean_dec(x_21); +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_35; lean_object* x_36; +x_35 = lean_ctor_get(x_12, 1); +lean_inc(x_35); +lean_dec(x_12); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +x_36 = l_Lean_Meta_Grind_ppGoals(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_35); +if (lean_obj_tag(x_36) == 0) +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_37 = lean_ctor_get(x_36, 0); +lean_inc(x_37); +x_38 = lean_ctor_get(x_36, 1); +lean_inc(x_38); +lean_dec(x_36); +x_39 = l_Lean_MessageData_ofFormat(x_37); +x_40 = l_Lean_Meta_Grind_main___lambda__2___closed__6; +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 = 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_addTrace___at_Lean_Meta_Grind_simp___spec__3(x_11, x_42, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_38); +x_44 = lean_ctor_get(x_43, 0); +lean_inc(x_44); +x_45 = lean_ctor_get(x_43, 1); +lean_inc(x_45); +lean_dec(x_43); +x_46 = l_Lean_Meta_Grind_preprocessAndProbe___lambda__1(x_1, x_2, x_44, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_45); +lean_dec(x_44); +return x_46; +} +else +{ +lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; +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_47 = lean_ctor_get(x_36, 0); +lean_inc(x_47); +x_48 = lean_ctor_get(x_36, 1); +lean_inc(x_48); +if (lean_is_exclusive(x_36)) { + lean_ctor_release(x_36, 0); + lean_ctor_release(x_36, 1); + x_49 = x_36; +} else { + lean_dec_ref(x_36); + x_49 = lean_box(0); +} +if (lean_is_scalar(x_49)) { + x_50 = lean_alloc_ctor(1, 2, 0); +} else { + x_50 = x_49; +} +lean_ctor_set(x_50, 0, x_47); +lean_ctor_set(x_50, 1, x_48); +return x_50; +} +} +} +} +} +static lean_object* _init_l_Lean_Meta_Grind_preprocessAndProbe___closed__1() { +_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 = lean_unsigned_to_nat(100u); +x_3 = lean_unsigned_to_nat(5u); +x_4 = lean_unsigned_to_nat(1000u); +x_5 = lean_alloc_ctor(0, 4, 1); +lean_ctor_set(x_5, 0, x_2); +lean_ctor_set(x_5, 1, x_3); +lean_ctor_set(x_5, 2, x_3); +lean_ctor_set(x_5, 3, x_4); +lean_ctor_set_uint8(x_5, sizeof(void*)*4, x_1); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_preprocessAndProbe(lean_object* x_1, lean_object* x_2, lean_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 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore), 9, 1); +lean_closure_set(x_9, 0, x_1); +x_10 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_preprocessAndProbe___lambda__2), 10, 1); +lean_closure_set(x_10, 0, x_3); +x_11 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); +lean_closure_set(x_11, 0, x_9); +lean_closure_set(x_11, 1, x_10); +x_12 = l_Lean_Meta_Grind_preprocessAndProbe___closed__1; +x_13 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_GrindM_run___rarg), 8, 3); +lean_closure_set(x_13, 0, x_11); +lean_closure_set(x_13, 1, x_2); +lean_closure_set(x_13, 2, x_12); +x_14 = l_Lean_Meta_withoutModifyingMCtx___at_Lean_Meta_Grind_preprocessAndProbe___spec__2___rarg(x_13, x_4, x_5, x_6, x_7, x_8); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_preprocessAndProbe___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_Meta_Grind_preprocessAndProbe___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_3); +return x_12; +} +} +lean_object* initialize_Init_Grind_Lemmas(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Meta_Tactic_Util(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Meta_Tactic_Grind_RevertAll(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Meta_Tactic_Grind_PropagatorAttr(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Meta_Tactic_Grind_Proj(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Meta_Tactic_Grind_ForallProp(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Meta_Tactic_Grind_Util(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Meta_Tactic_Grind_Inv(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Meta_Tactic_Grind_Intro(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Meta_Tactic_Grind_EMatch(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Lean_Meta_Tactic_Grind_Main(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_Grind_Lemmas(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Lean_Meta_Tactic_Util(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Lean_Meta_Tactic_Grind_RevertAll(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Lean_Meta_Tactic_Grind_PropagatorAttr(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Lean_Meta_Tactic_Grind_Proj(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Lean_Meta_Tactic_Grind_ForallProp(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Lean_Meta_Tactic_Grind_Util(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Lean_Meta_Tactic_Grind_Inv(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Lean_Meta_Tactic_Grind_Intro(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Lean_Meta_Tactic_Grind_EMatch(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l_Lean_Meta_Grind_mkMethods___rarg___closed__1 = _init_l_Lean_Meta_Grind_mkMethods___rarg___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_mkMethods___rarg___closed__1); +l_Lean_Meta_Grind_GrindM_run___rarg___closed__1 = _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_GrindM_run___rarg___closed__1); +l_Lean_Meta_Grind_GrindM_run___rarg___closed__2 = _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_GrindM_run___rarg___closed__2); +l_Lean_Meta_Grind_GrindM_run___rarg___closed__3 = _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind_GrindM_run___rarg___closed__3); +l_Lean_Meta_Grind_GrindM_run___rarg___closed__4 = _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__4(); +lean_mark_persistent(l_Lean_Meta_Grind_GrindM_run___rarg___closed__4); +l_Lean_Meta_Grind_GrindM_run___rarg___closed__5 = _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__5(); +lean_mark_persistent(l_Lean_Meta_Grind_GrindM_run___rarg___closed__5); +l_Lean_Meta_Grind_GrindM_run___rarg___closed__6 = _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__6(); +lean_mark_persistent(l_Lean_Meta_Grind_GrindM_run___rarg___closed__6); +l_Lean_Meta_Grind_GrindM_run___rarg___closed__7 = _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__7(); +lean_mark_persistent(l_Lean_Meta_Grind_GrindM_run___rarg___closed__7); +l_Lean_Meta_Grind_GrindM_run___rarg___closed__8 = _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__8(); +lean_mark_persistent(l_Lean_Meta_Grind_GrindM_run___rarg___closed__8); +l_Lean_Meta_Grind_GrindM_run___rarg___closed__9 = _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__9(); +lean_mark_persistent(l_Lean_Meta_Grind_GrindM_run___rarg___closed__9); +l_Lean_Meta_Grind_GrindM_run___rarg___closed__10 = _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__10(); +lean_mark_persistent(l_Lean_Meta_Grind_GrindM_run___rarg___closed__10); +l_Lean_Meta_Grind_GrindM_run___rarg___closed__11 = _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__11(); +lean_mark_persistent(l_Lean_Meta_Grind_GrindM_run___rarg___closed__11); +l_Lean_Meta_Grind_GrindM_run___rarg___closed__12 = _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__12(); +lean_mark_persistent(l_Lean_Meta_Grind_GrindM_run___rarg___closed__12); +l_Lean_Meta_Grind_GrindM_run___rarg___closed__13 = _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__13(); +lean_mark_persistent(l_Lean_Meta_Grind_GrindM_run___rarg___closed__13); +l_Lean_Meta_Grind_GrindM_run___rarg___closed__14 = _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__14(); +lean_mark_persistent(l_Lean_Meta_Grind_GrindM_run___rarg___closed__14); +l_Lean_Meta_Grind_GrindM_run___rarg___closed__15 = _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__15(); +lean_mark_persistent(l_Lean_Meta_Grind_GrindM_run___rarg___closed__15); +l_Lean_Meta_Grind_GrindM_run___rarg___closed__16 = _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__16(); +lean_mark_persistent(l_Lean_Meta_Grind_GrindM_run___rarg___closed__16); +l_Lean_Meta_Grind_GrindM_run___rarg___closed__17 = _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__17(); +lean_mark_persistent(l_Lean_Meta_Grind_GrindM_run___rarg___closed__17); +l_Lean_Meta_Grind_GrindM_run___rarg___closed__18 = _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__18(); +lean_mark_persistent(l_Lean_Meta_Grind_GrindM_run___rarg___closed__18); +l_Lean_Meta_Grind_GrindM_run___rarg___closed__19 = _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__19(); +lean_mark_persistent(l_Lean_Meta_Grind_GrindM_run___rarg___closed__19); +l_Lean_Meta_Grind_GrindM_run___rarg___closed__20 = _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__20(); +lean_mark_persistent(l_Lean_Meta_Grind_GrindM_run___rarg___closed__20); +l_Lean_PersistentHashMap_empty___at___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___spec__1 = _init_l_Lean_PersistentHashMap_empty___at___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___spec__1(); +lean_mark_persistent(l_Lean_PersistentHashMap_empty___at___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___spec__1); +l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___closed__1); +l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___closed__2); +l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___closed__3); +l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore___closed__1); +l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore___closed__2); +l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___closed__1); +l_Lean_Meta_Grind_main___lambda__2___closed__1 = _init_l_Lean_Meta_Grind_main___lambda__2___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_main___lambda__2___closed__1); +l_Lean_Meta_Grind_main___lambda__2___closed__2 = _init_l_Lean_Meta_Grind_main___lambda__2___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_main___lambda__2___closed__2); +l_Lean_Meta_Grind_main___lambda__2___closed__3 = _init_l_Lean_Meta_Grind_main___lambda__2___closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind_main___lambda__2___closed__3); +l_Lean_Meta_Grind_main___lambda__2___closed__4 = _init_l_Lean_Meta_Grind_main___lambda__2___closed__4(); +lean_mark_persistent(l_Lean_Meta_Grind_main___lambda__2___closed__4); +l_Lean_Meta_Grind_main___lambda__2___closed__5 = _init_l_Lean_Meta_Grind_main___lambda__2___closed__5(); +lean_mark_persistent(l_Lean_Meta_Grind_main___lambda__2___closed__5); +l_Lean_Meta_Grind_main___lambda__2___closed__6 = _init_l_Lean_Meta_Grind_main___lambda__2___closed__6(); +lean_mark_persistent(l_Lean_Meta_Grind_main___lambda__2___closed__6); +l_Lean_Meta_Grind_main___closed__1 = _init_l_Lean_Meta_Grind_main___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_main___closed__1); +l_Lean_Meta_Grind_preprocessAndProbe___closed__1 = _init_l_Lean_Meta_Grind_preprocessAndProbe___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_preprocessAndProbe___closed__1); +return lean_io_result_mk_ok(lean_box(0)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Lean/Meta/Tactic/Grind/MarkNestedProofs.c b/stage0/stdlib/Lean/Meta/Tactic/Grind/MarkNestedProofs.c index 99b9a11a5e7e..2f082515ee08 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Grind/MarkNestedProofs.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Grind/MarkNestedProofs.c @@ -24,7 +24,9 @@ size_t lean_uint64_to_usize(uint64_t); lean_object* l_Lean_Expr_sort___override(lean_object*); lean_object* lean_mk_array(lean_object*, lean_object*); uint8_t lean_usize_dec_eq(size_t, size_t); +lean_object* l_Lean_Expr_mdata___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markNestedProofs(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* lean_array_fget(lean_object*, lean_object*); lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_markNestedProofsImpl_visit___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*); @@ -5542,57 +5544,262 @@ return x_65; } } } +case 10: +{ +lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; +x_66 = lean_ctor_get(x_8, 1); +lean_inc(x_66); +lean_dec(x_8); +x_67 = lean_ctor_get(x_1, 0); +lean_inc(x_67); +x_68 = lean_ctor_get(x_1, 1); +lean_inc(x_68); +lean_inc(x_68); +x_69 = l_Lean_Meta_Grind_markNestedProofsImpl_visit(x_68, x_2, x_3, x_4, x_5, x_6, x_66); +if (lean_obj_tag(x_69) == 0) +{ +uint8_t x_70; +x_70 = !lean_is_exclusive(x_69); +if (x_70 == 0) +{ +lean_object* x_71; size_t x_72; size_t x_73; uint8_t x_74; +x_71 = lean_ctor_get(x_69, 0); +x_72 = lean_ptr_addr(x_68); +lean_dec(x_68); +x_73 = lean_ptr_addr(x_71); +x_74 = lean_usize_dec_eq(x_72, x_73); +if (x_74 == 0) +{ +lean_object* x_75; +lean_dec(x_1); +x_75 = l_Lean_Expr_mdata___override(x_67, x_71); +lean_ctor_set(x_69, 0, x_75); +return x_69; +} +else +{ +lean_dec(x_71); +lean_dec(x_67); +lean_ctor_set(x_69, 0, x_1); +return x_69; +} +} +else +{ +lean_object* x_76; lean_object* x_77; size_t x_78; size_t x_79; uint8_t x_80; +x_76 = lean_ctor_get(x_69, 0); +x_77 = lean_ctor_get(x_69, 1); +lean_inc(x_77); +lean_inc(x_76); +lean_dec(x_69); +x_78 = lean_ptr_addr(x_68); +lean_dec(x_68); +x_79 = lean_ptr_addr(x_76); +x_80 = lean_usize_dec_eq(x_78, x_79); +if (x_80 == 0) +{ +lean_object* x_81; lean_object* x_82; +lean_dec(x_1); +x_81 = l_Lean_Expr_mdata___override(x_67, x_76); +x_82 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_82, 0, x_81); +lean_ctor_set(x_82, 1, x_77); +return x_82; +} +else +{ +lean_object* x_83; +lean_dec(x_76); +lean_dec(x_67); +x_83 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_83, 0, x_1); +lean_ctor_set(x_83, 1, x_77); +return x_83; +} +} +} +else +{ +uint8_t x_84; +lean_dec(x_68); +lean_dec(x_67); +lean_dec(x_1); +x_84 = !lean_is_exclusive(x_69); +if (x_84 == 0) +{ +return x_69; +} +else +{ +lean_object* x_85; lean_object* x_86; lean_object* x_87; +x_85 = lean_ctor_get(x_69, 0); +x_86 = lean_ctor_get(x_69, 1); +lean_inc(x_86); +lean_inc(x_85); +lean_dec(x_69); +x_87 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_87, 0, x_85); +lean_ctor_set(x_87, 1, x_86); +return x_87; +} +} +} +case 11: +{ +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_8, 1); +lean_inc(x_88); +lean_dec(x_8); +x_89 = lean_ctor_get(x_1, 0); +lean_inc(x_89); +x_90 = lean_ctor_get(x_1, 1); +lean_inc(x_90); +x_91 = lean_ctor_get(x_1, 2); +lean_inc(x_91); +lean_inc(x_91); +x_92 = l_Lean_Meta_Grind_markNestedProofsImpl_visit(x_91, x_2, x_3, x_4, x_5, x_6, x_88); +if (lean_obj_tag(x_92) == 0) +{ +uint8_t x_93; +x_93 = !lean_is_exclusive(x_92); +if (x_93 == 0) +{ +lean_object* x_94; size_t x_95; size_t x_96; uint8_t x_97; +x_94 = lean_ctor_get(x_92, 0); +x_95 = lean_ptr_addr(x_91); +lean_dec(x_91); +x_96 = lean_ptr_addr(x_94); +x_97 = lean_usize_dec_eq(x_95, x_96); +if (x_97 == 0) +{ +lean_object* x_98; +lean_dec(x_1); +x_98 = l_Lean_Expr_proj___override(x_89, x_90, x_94); +lean_ctor_set(x_92, 0, x_98); +return x_92; +} +else +{ +lean_dec(x_94); +lean_dec(x_90); +lean_dec(x_89); +lean_ctor_set(x_92, 0, x_1); +return x_92; +} +} +else +{ +lean_object* x_99; lean_object* x_100; size_t x_101; size_t x_102; uint8_t x_103; +x_99 = lean_ctor_get(x_92, 0); +x_100 = lean_ctor_get(x_92, 1); +lean_inc(x_100); +lean_inc(x_99); +lean_dec(x_92); +x_101 = lean_ptr_addr(x_91); +lean_dec(x_91); +x_102 = lean_ptr_addr(x_99); +x_103 = lean_usize_dec_eq(x_101, x_102); +if (x_103 == 0) +{ +lean_object* x_104; lean_object* x_105; +lean_dec(x_1); +x_104 = l_Lean_Expr_proj___override(x_89, x_90, x_99); +x_105 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_105, 0, x_104); +lean_ctor_set(x_105, 1, x_100); +return x_105; +} +else +{ +lean_object* x_106; +lean_dec(x_99); +lean_dec(x_90); +lean_dec(x_89); +x_106 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_106, 0, x_1); +lean_ctor_set(x_106, 1, x_100); +return x_106; +} +} +} +else +{ +uint8_t x_107; +lean_dec(x_91); +lean_dec(x_90); +lean_dec(x_89); +lean_dec(x_1); +x_107 = !lean_is_exclusive(x_92); +if (x_107 == 0) +{ +return x_92; +} +else +{ +lean_object* x_108; lean_object* x_109; lean_object* x_110; +x_108 = lean_ctor_get(x_92, 0); +x_109 = lean_ctor_get(x_92, 1); +lean_inc(x_109); +lean_inc(x_108); +lean_dec(x_92); +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; +} +} +} default: { -uint8_t x_66; +uint8_t x_111; lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_66 = !lean_is_exclusive(x_8); -if (x_66 == 0) +x_111 = !lean_is_exclusive(x_8); +if (x_111 == 0) { -lean_object* x_67; -x_67 = lean_ctor_get(x_8, 0); -lean_dec(x_67); +lean_object* x_112; +x_112 = lean_ctor_get(x_8, 0); +lean_dec(x_112); lean_ctor_set(x_8, 0, x_1); return x_8; } else { -lean_object* x_68; lean_object* x_69; -x_68 = lean_ctor_get(x_8, 1); -lean_inc(x_68); +lean_object* x_113; lean_object* x_114; +x_113 = lean_ctor_get(x_8, 1); +lean_inc(x_113); lean_dec(x_8); -x_69 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_69, 0, x_1); -lean_ctor_set(x_69, 1, x_68); -return x_69; +x_114 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_114, 0, x_1); +lean_ctor_set(x_114, 1, x_113); +return x_114; } } } } else { -uint8_t x_70; -x_70 = !lean_is_exclusive(x_8); -if (x_70 == 0) +uint8_t x_115; +x_115 = !lean_is_exclusive(x_8); +if (x_115 == 0) { -lean_object* x_71; lean_object* x_72; lean_object* x_73; uint8_t x_74; -x_71 = lean_ctor_get(x_8, 1); -x_72 = lean_ctor_get(x_8, 0); -lean_dec(x_72); -x_73 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2___closed__4; -x_74 = l_Lean_Expr_isAppOf(x_1, x_73); -if (x_74 == 0) +lean_object* x_116; lean_object* x_117; lean_object* x_118; uint8_t x_119; +x_116 = lean_ctor_get(x_8, 1); +x_117 = lean_ctor_get(x_8, 0); +lean_dec(x_117); +x_118 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2___closed__4; +x_119 = l_Lean_Expr_isAppOf(x_1, x_118); +if (x_119 == 0) { -lean_object* x_75; lean_object* x_76; +lean_object* x_120; lean_object* x_121; lean_free_object(x_8); -x_75 = lean_box(0); -x_76 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3(x_1, x_75, x_2, x_3, x_4, x_5, x_6, x_71); +x_120 = lean_box(0); +x_121 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3(x_1, x_120, x_2, x_3, x_4, x_5, x_6, x_116); lean_dec(x_2); -return x_76; +return x_121; } else { @@ -5607,62 +5814,62 @@ return x_8; } else { -lean_object* x_77; lean_object* x_78; uint8_t x_79; -x_77 = lean_ctor_get(x_8, 1); -lean_inc(x_77); +lean_object* x_122; lean_object* x_123; uint8_t x_124; +x_122 = lean_ctor_get(x_8, 1); +lean_inc(x_122); lean_dec(x_8); -x_78 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2___closed__4; -x_79 = l_Lean_Expr_isAppOf(x_1, x_78); -if (x_79 == 0) +x_123 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2___closed__4; +x_124 = l_Lean_Expr_isAppOf(x_1, x_123); +if (x_124 == 0) { -lean_object* x_80; lean_object* x_81; -x_80 = lean_box(0); -x_81 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3(x_1, x_80, x_2, x_3, x_4, x_5, x_6, x_77); +lean_object* x_125; lean_object* x_126; +x_125 = lean_box(0); +x_126 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3(x_1, x_125, x_2, x_3, x_4, x_5, x_6, x_122); lean_dec(x_2); -return x_81; +return x_126; } else { -lean_object* x_82; +lean_object* x_127; lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_82 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_82, 0, x_1); -lean_ctor_set(x_82, 1, x_77); -return x_82; +x_127 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_127, 0, x_1); +lean_ctor_set(x_127, 1, x_122); +return x_127; } } } } else { -uint8_t x_83; +uint8_t x_128; 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_83 = !lean_is_exclusive(x_8); -if (x_83 == 0) +x_128 = !lean_is_exclusive(x_8); +if (x_128 == 0) { return x_8; } else { -lean_object* x_84; lean_object* x_85; lean_object* x_86; -x_84 = lean_ctor_get(x_8, 0); -x_85 = lean_ctor_get(x_8, 1); -lean_inc(x_85); -lean_inc(x_84); +lean_object* x_129; lean_object* x_130; lean_object* x_131; +x_129 = lean_ctor_get(x_8, 0); +x_130 = lean_ctor_get(x_8, 1); +lean_inc(x_130); +lean_inc(x_129); lean_dec(x_8); -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; +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; } } } diff --git a/stage0/stdlib/Lean/Meta/Tactic/Grind/PP.c b/stage0/stdlib/Lean/Meta/Tactic/Grind/PP.c index d1b03fc76b9a..0e511d3997ea 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Grind/PP.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Grind/PP.c @@ -14,12 +14,14 @@ extern "C" { #endif static lean_object* l_Lean_Meta_Grind_ppENodeRef___closed__3; +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Meta_Grind_ppGoals___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_Meta_Grind_ppENodeDecl___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_Meta_ppExpr(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppState___spec__1___closed__1; static lean_object* l_Lean_Meta_Grind_ppENodeDecl___closed__4; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ppENodeDecl___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(lean_object*, lean_object*); +lean_object* l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___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_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppState___spec__1(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_ppENodeDecl___lambda__2___closed__4; static lean_object* l_Lean_Meta_Grind_ppENodeDecl___closed__3; @@ -45,6 +47,8 @@ static lean_object* l_Lean_Meta_Grind_ppENodeDecl___lambda__3___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ppENodeDeclValue(lean_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_Grind_ppENodeDecl___lambda__2___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ppENodeDecl___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_st_ref_get(lean_object*, lean_object*); +lean_object* lean_st_mk_ref(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_ppENodeDecl___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ppENodeDecl___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ppENodeDecl___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*); @@ -53,6 +57,7 @@ static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppState___ LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ppState(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_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppState___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_List_forIn_x27_loop___at_Lean_Meta_Grind_ppGoals___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_Meta_Grind_ppENodeDecl___lambda__4___closed__2; static lean_object* l_Lean_Meta_Grind_ppENodeDecl___closed__2; uint8_t l_Lean_Meta_Grind_isSameExpr_unsafe__1(lean_object*, lean_object*); @@ -60,20 +65,24 @@ static lean_object* l_List_forIn_x27_loop___at_Lean_Meta_Grind_ppState___spec__3 LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ppENodeDecl___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_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_ppENodeDeclValue___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_List_forIn_x27_loop___at_Lean_Meta_Grind_ppGoals___spec__1___closed__1; LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Meta_Grind_ppState___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* l_Lean_Meta_Grind_getENodes(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppENodeDeclValue___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*); static lean_object* l_Lean_Meta_Grind_ppENodeDecl___lambda__3___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ppENodeDeclValue___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_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___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_Meta_Grind_ppENodeDecl___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ppENodeDecl___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); static lean_object* l_List_forIn_x27_loop___at_Lean_Meta_Grind_ppState___spec__3___closed__4; static lean_object* l_Lean_Meta_Grind_ppState___closed__2; static lean_object* l_Lean_Meta_Grind_ppENodeDeclValue___closed__1; +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Meta_Grind_ppGoals___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* l_Std_Format_joinSep___at_Prod_repr___spec__1(lean_object*, lean_object*); lean_object* lean_array_set(lean_object*, lean_object*, lean_object*); lean_object* lean_nat_sub(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Meta_Grind_ppGoals___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_Expr_withAppAux___at_Lean_Meta_Grind_ppENodeDeclValue___spec__2___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_List_forIn_x27_loop___at_Lean_Meta_Grind_ppState___spec__3___closed__5; lean_object* l_Lean_Meta_Grind_getENode_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -86,6 +95,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ppENodeDecl(lean_object*, lean_object LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ppENodeRef___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_add(size_t, size_t); LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Meta_Grind_ppState___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_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Meta_Grind_ppGoals___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_array_uget(lean_object*, size_t); size_t lean_array_size(lean_object*); uint8_t lean_usize_dec_lt(size_t, size_t); @@ -93,6 +103,7 @@ static lean_object* l_Lean_Meta_Grind_ppENodeRef___closed__4; LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_ppENodeDeclValue___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_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppENodeDeclValue___spec__1___closed__2; uint8_t l_Lean_Expr_isConst(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ppGoals(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_ppState___closed__1; lean_object* l___private_Init_Data_Repr_0__Nat_reprFast(lean_object*); lean_object* l_Lean_Meta_isLitValue(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1943,6 +1954,422 @@ lean_dec(x_1); return x_16; } } +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Meta_Grind_ppGoals___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) { +_start: +{ +lean_object* x_10; uint8_t x_11; +x_10 = lean_st_mk_ref(x_1, x_9); +x_11 = !lean_is_exclusive(x_10); +if (x_11 == 0) +{ +return x_10; +} +else +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_12 = lean_ctor_get(x_10, 0); +x_13 = lean_ctor_get(x_10, 1); +lean_inc(x_13); +lean_inc(x_12); +lean_dec(x_10); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_12); +lean_ctor_set(x_14, 1, x_13); +return x_14; +} +} +} +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Meta_Grind_ppGoals___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) { +_start: +{ +lean_object* x_10; +lean_inc(x_1); +x_10 = l_Lean_Meta_Grind_ppState(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_10) == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_11 = lean_ctor_get(x_10, 0); +lean_inc(x_11); +x_12 = lean_ctor_get(x_10, 1); +lean_inc(x_12); +lean_dec(x_10); +x_13 = lean_st_ref_get(x_1, x_12); +lean_dec(x_1); +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; +x_15 = lean_ctor_get(x_13, 0); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_11); +lean_ctor_set(x_16, 1, x_15); +lean_ctor_set(x_13, 0, x_16); +return x_13; +} +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_17 = lean_ctor_get(x_13, 0); +x_18 = lean_ctor_get(x_13, 1); +lean_inc(x_18); +lean_inc(x_17); +lean_dec(x_13); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_11); +lean_ctor_set(x_19, 1, x_17); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_19); +lean_ctor_set(x_20, 1, x_18); +return x_20; +} +} +else +{ +uint8_t x_21; +lean_dec(x_1); +x_21 = !lean_is_exclusive(x_10); +if (x_21 == 0) +{ +return x_10; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_22 = lean_ctor_get(x_10, 0); +x_23 = lean_ctor_get(x_10, 1); +lean_inc(x_23); +lean_inc(x_22); +lean_dec(x_10); +x_24 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_24, 0, x_22); +lean_ctor_set(x_24, 1, x_23); +return x_24; +} +} +} +} +static lean_object* _init_l_List_forIn_x27_loop___at_Lean_Meta_Grind_ppGoals___spec__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_List_forIn_x27_loop___at_Lean_Meta_Grind_ppGoals___spec__1___lambda__2), 9, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Meta_Grind_ppGoals___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +if (lean_obj_tag(x_4) == 0) +{ +lean_object* x_15; +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_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_5); +lean_ctor_set(x_15, 1, x_14); +return x_15; +} +else +{ +uint8_t x_16; +x_16 = !lean_is_exclusive(x_4); +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; +x_17 = lean_ctor_get(x_4, 0); +x_18 = lean_ctor_get(x_4, 1); +x_19 = lean_ctor_get(x_17, 0); +lean_inc(x_19); +x_20 = lean_alloc_closure((void*)(l_List_forIn_x27_loop___at_Lean_Meta_Grind_ppGoals___spec__1___lambda__1___boxed), 9, 1); +lean_closure_set(x_20, 0, x_17); +x_21 = l_List_forIn_x27_loop___at_Lean_Meta_Grind_ppGoals___spec__1___closed__1; +x_22 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); +lean_closure_set(x_22, 0, x_20); +lean_closure_set(x_22, 1, x_21); +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_23 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(x_19, x_22, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_23) == 0) +{ +lean_object* x_24; lean_object* x_25; uint8_t x_26; +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +x_25 = lean_ctor_get(x_23, 1); +lean_inc(x_25); +lean_dec(x_23); +x_26 = !lean_is_exclusive(x_24); +if (x_26 == 0) +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_24, 0); +x_28 = lean_ctor_get(x_24, 1); +lean_dec(x_28); +x_29 = lean_box(1); +lean_ctor_set_tag(x_24, 5); +lean_ctor_set(x_24, 1, x_29); +lean_ctor_set(x_24, 0, x_5); +lean_ctor_set_tag(x_4, 5); +lean_ctor_set(x_4, 1, x_27); +lean_ctor_set(x_4, 0, x_24); +{ +lean_object* _tmp_3 = x_18; +lean_object* _tmp_4 = x_4; +lean_object* _tmp_5 = lean_box(0); +lean_object* _tmp_13 = x_25; +x_4 = _tmp_3; +x_5 = _tmp_4; +x_6 = _tmp_5; +x_14 = _tmp_13; +} +goto _start; +} +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_24, 0); +lean_inc(x_31); +lean_dec(x_24); +x_32 = lean_box(1); +x_33 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_33, 0, x_5); +lean_ctor_set(x_33, 1, x_32); +lean_ctor_set_tag(x_4, 5); +lean_ctor_set(x_4, 1, x_31); +lean_ctor_set(x_4, 0, x_33); +{ +lean_object* _tmp_3 = x_18; +lean_object* _tmp_4 = x_4; +lean_object* _tmp_5 = lean_box(0); +lean_object* _tmp_13 = x_25; +x_4 = _tmp_3; +x_5 = _tmp_4; +x_6 = _tmp_5; +x_14 = _tmp_13; +} +goto _start; +} +} +else +{ +uint8_t x_35; +lean_free_object(x_4); +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_35 = !lean_is_exclusive(x_23); +if (x_35 == 0) +{ +return x_23; +} +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; +} +} +} +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; +x_39 = lean_ctor_get(x_4, 0); +x_40 = lean_ctor_get(x_4, 1); +lean_inc(x_40); +lean_inc(x_39); +lean_dec(x_4); +x_41 = lean_ctor_get(x_39, 0); +lean_inc(x_41); +x_42 = lean_alloc_closure((void*)(l_List_forIn_x27_loop___at_Lean_Meta_Grind_ppGoals___spec__1___lambda__1___boxed), 9, 1); +lean_closure_set(x_42, 0, x_39); +x_43 = l_List_forIn_x27_loop___at_Lean_Meta_Grind_ppGoals___spec__1___closed__1; +x_44 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); +lean_closure_set(x_44, 0, x_42); +lean_closure_set(x_44, 1, x_43); +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_45 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(x_41, x_44, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +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; +x_46 = lean_ctor_get(x_45, 0); +lean_inc(x_46); +x_47 = lean_ctor_get(x_45, 1); +lean_inc(x_47); +lean_dec(x_45); +x_48 = lean_ctor_get(x_46, 0); +lean_inc(x_48); +if (lean_is_exclusive(x_46)) { + lean_ctor_release(x_46, 0); + lean_ctor_release(x_46, 1); + x_49 = x_46; +} else { + lean_dec_ref(x_46); + x_49 = lean_box(0); +} +x_50 = lean_box(1); +if (lean_is_scalar(x_49)) { + x_51 = lean_alloc_ctor(5, 2, 0); +} else { + x_51 = x_49; + lean_ctor_set_tag(x_51, 5); +} +lean_ctor_set(x_51, 0, x_5); +lean_ctor_set(x_51, 1, x_50); +x_52 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_52, 0, x_51); +lean_ctor_set(x_52, 1, x_48); +x_4 = x_40; +x_5 = x_52; +x_6 = lean_box(0); +x_14 = x_47; +goto _start; +} +else +{ +lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; +lean_dec(x_40); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_5); +x_54 = lean_ctor_get(x_45, 0); +lean_inc(x_54); +x_55 = lean_ctor_get(x_45, 1); +lean_inc(x_55); +if (lean_is_exclusive(x_45)) { + lean_ctor_release(x_45, 0); + lean_ctor_release(x_45, 1); + x_56 = x_45; +} else { + lean_dec_ref(x_45); + x_56 = lean_box(0); +} +if (lean_is_scalar(x_56)) { + x_57 = lean_alloc_ctor(1, 2, 0); +} else { + x_57 = x_56; +} +lean_ctor_set(x_57, 0, x_54); +lean_ctor_set(x_57, 1, x_55); +return x_57; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ppGoals(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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 = l_Lean_Meta_Grind_ppENodeRef___closed__6; +lean_inc(x_1); +x_12 = l_List_forIn_x27_loop___at_Lean_Meta_Grind_ppGoals___spec__1(x_1, x_10, x_1, x_1, x_11, lean_box(0), x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_1); +if (lean_obj_tag(x_12) == 0) +{ +uint8_t x_13; +x_13 = !lean_is_exclusive(x_12); +if (x_13 == 0) +{ +return x_12; +} +else +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_14 = lean_ctor_get(x_12, 0); +x_15 = lean_ctor_get(x_12, 1); +lean_inc(x_15); +lean_inc(x_14); +lean_dec(x_12); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_14); +lean_ctor_set(x_16, 1, x_15); +return x_16; +} +} +else +{ +uint8_t x_17; +x_17 = !lean_is_exclusive(x_12); +if (x_17 == 0) +{ +return x_12; +} +else +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_18 = lean_ctor_get(x_12, 0); +x_19 = lean_ctor_get(x_12, 1); +lean_inc(x_19); +lean_inc(x_18); +lean_dec(x_12); +x_20 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_20, 0, x_18); +lean_ctor_set(x_20, 1, x_19); +return x_20; +} +} +} +} +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Meta_Grind_ppGoals___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) { +_start: +{ +lean_object* x_10; +x_10 = l_List_forIn_x27_loop___at_Lean_Meta_Grind_ppGoals___spec__1___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_10; +} +} +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Meta_Grind_ppGoals___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +lean_object* x_15; +x_15 = l_List_forIn_x27_loop___at_Lean_Meta_Grind_ppGoals___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_15; +} +} lean_object* initialize_Init_Grind_Util(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_Tactic_Grind_Types(uint8_t builtin, lean_object*); static bool _G_initialized = false; @@ -2018,6 +2445,8 @@ l_Lean_Meta_Grind_ppState___closed__1 = _init_l_Lean_Meta_Grind_ppState___closed lean_mark_persistent(l_Lean_Meta_Grind_ppState___closed__1); l_Lean_Meta_Grind_ppState___closed__2 = _init_l_Lean_Meta_Grind_ppState___closed__2(); lean_mark_persistent(l_Lean_Meta_Grind_ppState___closed__2); +l_List_forIn_x27_loop___at_Lean_Meta_Grind_ppGoals___spec__1___closed__1 = _init_l_List_forIn_x27_loop___at_Lean_Meta_Grind_ppGoals___spec__1___closed__1(); +lean_mark_persistent(l_List_forIn_x27_loop___at_Lean_Meta_Grind_ppGoals___spec__1___closed__1); return lean_io_result_mk_ok(lean_box(0)); } #ifdef __cplusplus diff --git a/stage0/stdlib/Lean/Meta/Tactic/Grind/Preprocessor.c b/stage0/stdlib/Lean/Meta/Tactic/Grind/Preprocessor.c deleted file mode 100644 index 3e7d841bf4ee..000000000000 --- a/stage0/stdlib/Lean/Meta/Tactic/Grind/Preprocessor.c +++ /dev/null @@ -1,8623 +0,0 @@ -// Lean compiler output -// Module: Lean.Meta.Tactic.Grind.Preprocessor -// Imports: Init.Grind.Lemmas Lean.Meta.Canonicalizer Lean.Meta.Tactic.Util Lean.Meta.Tactic.Intro Lean.Meta.Tactic.Simp.Main Lean.Meta.Tactic.Grind.Attr Lean.Meta.Tactic.Grind.RevertAll Lean.Meta.Tactic.Grind.Types Lean.Meta.Tactic.Grind.Util Lean.Meta.Tactic.Grind.Cases Lean.Meta.Tactic.Grind.Injection Lean.Meta.Tactic.Grind.Core Lean.Meta.Tactic.Grind.Simp Lean.Meta.Tactic.Grind.Run -#include -#if defined(__clang__) -#pragma clang diagnostic ignored "-Wunused-parameter" -#pragma clang diagnostic ignored "-Wunused-label" -#elif defined(__GNUC__) && !defined(__CLANG__) -#pragma GCC diagnostic ignored "-Wunused-parameter" -#pragma GCC diagnostic ignored "-Wunused-label" -#pragma GCC diagnostic ignored "-Wunused-but-set-variable" -#endif -#ifdef __cplusplus -extern "C" { -#endif -LEAN_EXPORT lean_object* l_Lean_mkFreshId___at_Lean_Meta_Grind_Preprocessor_introNext___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Expr_bindingName_x21(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Preprocessor_preprocess___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_const___override(lean_object*, lean_object*); -lean_object* l_Lean_Meta_Grind_injection_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Preprocessor_loop___lambda__4___boxed(lean_object**); -static lean_object* l_Lean_Meta_Grind_Preprocessor_preprocess___closed__3; -LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at_Lean_Meta_Grind_Preprocessor_ppGoals___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_mk_empty_array_with_capacity(lean_object*); -lean_object* l_Lean_mkAppN(lean_object*, lean_object*); -lean_object* l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___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_Meta_Grind_Preprocessor_introNext___lambda__5___closed__2; -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_PersistentArray_forIn___at_Lean_Meta_Grind_Preprocessor_ppGoals___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_Grind_Preprocessor_introNext___lambda__5___closed__4; -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Preprocessor_loop___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_mkFreshFVarId___at_Lean_Meta_Grind_Preprocessor_introNext___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Meta_isProp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_Preprocessor_ppGoals___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*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_Preprocessor_instInhabitedState___closed__4; -LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at_Lean_Meta_Grind_Preprocessor_preprocess___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Meta_Grind_Preprocessor_preprocess___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_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_Preprocessor_ppGoals___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_Meta_Grind_Preprocessor_preprocess___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_Expr_isLet(lean_object*); -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_Preprocessor_preprocess___spec__5___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_push___rarg(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_mkFreshId___at_Lean_Meta_Grind_Preprocessor_introNext___spec__4(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_List_filterTR_loop___at_Lean_Meta_Grind_main___spec__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at_Lean_Meta_Grind_Preprocessor_preprocess___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_Meta_Grind_preprocessAndProbe(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_MVarId_getTag(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Preprocessor_applyCases_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_withoutModifyingMCtx___at_Lean_Meta_Grind_preprocessAndProbe___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static size_t l_Lean_Meta_Grind_Preprocessor_instInhabitedState___closed__3; -size_t lean_usize_of_nat(lean_object*); -lean_object* l_Lean_Meta_Grind_getMainDeclName___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_PersistentArray_toList___rarg(lean_object*); -lean_object* l_Lean_MVarId_revertAll(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_Preprocessor_preprocess___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_Preprocessor_preprocess___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_Grind_Preprocessor_preprocessAndProbe___spec__5(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_Preprocessor_preprocess___closed__5; -LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Meta_Grind_Preprocessor_introNext___spec__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Preprocessor_introNext___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Preprocessor_loop___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_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_Preprocessor_instInhabitedState___closed__2; -LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Meta_Grind_Preprocessor_introNext___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*); -size_t lean_usize_of_nat(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Preprocessor_loop___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_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_Preprocessor_ppGoals___spec__3___boxed(lean_object**); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Preprocessor_introNext(lean_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_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_Preprocessor_preprocess___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* lean_st_ref_take(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_Preprocessor_ppGoals___closed__2; -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Preprocessor_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_isTracingEnabledFor___at_Lean_Meta_Grind_Preprocessor_preprocess___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_Grind_Preprocessor_preprocess___closed__1; -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_Grind_Preprocessor_preprocessAndProbe___spec__3(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Preprocessor_introNext___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_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*); -lean_object* l_Lean_MVarId_getType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_withoutModifyingMCtx___at_Lean_Meta_Grind_preprocessAndProbe___spec__1(lean_object*); -static lean_object* l_Lean_Meta_Grind_Preprocessor_PreM_run___rarg___closed__1; -LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at_Lean_Meta_Grind_Preprocessor_ppGoals___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_Preprocessor_ppGoals___spec__4___closed__1; -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_Preprocessor_preprocess___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_MessageData_ofFormat(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*); -static lean_object* l_Lean_Meta_Grind_Preprocessor_loop___lambda__4___closed__1; -static lean_object* l_Lean_Meta_Grind_Preprocessor_introNext___lambda__5___closed__3; -static lean_object* l_Lean_Meta_Grind_Preprocessor_PreM_run___rarg___closed__3; -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Preprocessor_preprocess(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_FVarId_getDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_Grind_Preprocessor_preprocessAndProbe___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Preprocessor_introNext___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* lean_st_ref_get(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_Preprocessor_ppGoals___closed__1; -lean_object* lean_st_mk_ref(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Preprocessor_pushResult___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withMVarContextImp___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); -lean_object* l_Lean_Meta_Grind_isGrindCasesTarget(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_main___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Meta_Grind_checkInvariants(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_addTrace___at_Lean_Meta_Grind_Preprocessor_preprocess___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Preprocessor_ppGoals___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Meta_Grind_ppState(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_FVarId_getType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_mkFreshId___at_Lean_Meta_Grind_Preprocessor_introNext___spec__4___rarg(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at_Lean_Meta_Grind_Preprocessor_ppGoals___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_Meta_Grind_Preprocessor_preprocess___closed__2; -LEAN_EXPORT lean_object* l_List_forM___at_Lean_Meta_Grind_Preprocessor_loop___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_foldlMUnsafe_fold___at_Lean_Meta_Grind_Preprocessor_preprocessAndProbe___spec__4___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Meta_Grind_GrindM_run___rarg(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_Grind_Preprocessor_preprocessAndProbe___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*); -static lean_object* l_Lean_Meta_Grind_Preprocessor_preprocess___closed__4; -uint8_t l_Lean_Expr_isArrow(lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at_Lean_Meta_Grind_Preprocessor_ppGoals___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_MVarId_clearAuxDecls(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Preprocessor_loop___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_MVarId_abstractNestedProofs(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentArray_forM___at_Lean_Meta_Grind_Preprocessor_preprocessAndProbe___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_Meta_Grind_Preprocessor_ppGoals(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_Grind_Preprocessor_preprocessAndProbe___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_Meta_Grind_Preprocessor_preprocessAndProbe(lean_object*, lean_object*, 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_Grind_Preprocessor_PreM_run___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_MVarId_assert(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_mkFreshId___at_Lean_Meta_Grind_Preprocessor_introNext___spec__4___rarg___boxed(lean_object*, lean_object*); -lean_object* l_Lean_LocalContext_mkLocalDecl(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t); -LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Meta_Grind_Preprocessor_applyCases_x3f___spec__1(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_main___closed__1; -LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Meta_Grind_main___spec__2(lean_object*, lean_object*); -double l_Float_ofScientific(lean_object*, uint8_t, lean_object*); -lean_object* l_Lean_LocalDecl_userName(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Preprocessor_PreM_run(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_main(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_Preprocessor_ppGoals___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Preprocessor_isCasesCandidate___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at_Lean_Meta_Grind_Preprocessor_preprocess___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_Lean_PersistentArray_forInAux___at_Lean_Meta_Grind_Preprocessor_ppGoals___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* l_Lean_Expr_bindingDomain_x21(lean_object*); -lean_object* l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___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_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_Preprocessor_ppGoals___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_MVarId_withContext___at_Lean_Meta_Grind_Preprocessor_introNext___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*); -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_Preprocessor_preprocess___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_MVarId_ensureNoMVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at_Lean_Meta_Grind_Preprocessor_ppGoals___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_Lean_Meta_Grind_Preprocessor_preprocess___closed__6; -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Preprocessor_applyCases_x3f___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -uint8_t lean_nat_dec_lt(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_main___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_PersistentHashMap_insert___at_Lean_MVarId_assign___spec__1(lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Meta_getLocalInstances(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_Grind_Preprocessor_introNext___lambda__6___boxed(lean_object**); -static double l_Lean_addTrace___at_Lean_Meta_Grind_Preprocessor_preprocess___spec__7___closed__1; -LEAN_EXPORT lean_object* l_Lean_PersistentArray_forMAux___at_Lean_Meta_Grind_Preprocessor_preprocessAndProbe___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_mkFreshFVarId___at_Lean_Meta_Grind_Preprocessor_introNext___spec__3(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_PersistentArray_forIn___at_Lean_Meta_Grind_Preprocessor_preprocess___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_LocalDecl_type(lean_object*); -static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_Preprocessor_preprocess___spec__5___closed__1; -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Preprocessor_introNext___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_Grind_Preprocessor_pushResult(lean_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_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_Preprocessor_ppGoals___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Preprocessor_applyInjection_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_MVarId_transformTarget(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Meta_resetCache___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Expr_getAppFn(lean_object*); -lean_object* l_List_reverse___rarg(lean_object*); -LEAN_EXPORT lean_object* l_Lean_MVarId_assign___at_Lean_Meta_Grind_Preprocessor_introNext___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Preprocessor_introNext___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* l_Lean_MVarId_byContra_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* lean_array_mk(lean_object*); -lean_object* l_Lean_MVarId_ensureProp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_isTracingEnabledForCore(lean_object*, lean_object*, lean_object*); -size_t lean_usize_add(size_t, size_t); -lean_object* l_Lean_Meta_Grind_pre(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* lean_array_uget(lean_object*, size_t); -lean_object* l_Lean_Expr_fvar___override(lean_object*); -size_t lean_array_size(lean_object*); -LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Meta_Grind_Preprocessor_applyCases_x3f___spec__1___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_MVarId_assign___at_Lean_Meta_Grind_Preprocessor_introNext___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_st_ref_set(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Preprocessor_introNext___lambda__5___boxed(lean_object**); -lean_object* l_Lean_Meta_Grind_unfoldReducible(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_Meta_Grind_Preprocessor_introNext___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at_Lean_Meta_Grind_Preprocessor_ppGoals___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_EXPORT lean_object* l_Lean_Meta_Grind_Preprocessor_introNext___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_Grind_Preprocessor_introNext___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_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_Meta_Grind_Preprocessor_introNext___spec__2(lean_object*); -lean_object* l_Lean_Meta_Grind_addHyp(lean_object*, lean_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_Grind_Preprocessor_instInhabitedState___closed__1; -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Preprocessor_introNext___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_Grind_Preprocessor_loop___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_isTracingEnabledFor___at_Lean_Meta_Grind_Preprocessor_preprocess___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_array_get_size(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_preprocess(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Preprocessor_instInhabitedState; -lean_object* l_Lean_Meta_Grind_mkGoal(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_PersistentArray_forM___at_Lean_Meta_Grind_Preprocessor_preprocessAndProbe___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -uint8_t lean_usize_dec_lt(size_t, size_t); -static lean_object* l_Lean_Meta_Grind_Preprocessor_introNext___lambda__5___closed__1; -lean_object* l_Lean_Meta_Grind_cases(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_Preprocessor_PreM_run___rarg___closed__2; -lean_object* l_Lean_Meta_intro1Core(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Meta_mkLambdaFVars(lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* lean_nat_add(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentArray_forMAux___at_Lean_Meta_Grind_Preprocessor_preprocessAndProbe___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_Meta_Grind_Preprocessor_isCasesCandidate(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Expr_bindingInfo_x21(lean_object*); -uint8_t l_Lean_Expr_isForall(lean_object*); -lean_object* l_Lean_Expr_mvarId_x21(lean_object*); -static lean_object* l_Lean_addTrace___at_Lean_Meta_Grind_Preprocessor_preprocess___spec__7___closed__2; -lean_object* l_Lean_MVarId_betaReduce___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Preprocessor_loop___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_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_Preprocessor_preprocess___spec__4___boxed(lean_object**); -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_Preprocessor_ppGoals___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Preprocessor_introNext___lambda__2___boxed(lean_object**); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Preprocessor_applyCases_x3f___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_Grind_Preprocessor_preprocessAndProbe___spec__4(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* _init_l_Lean_Meta_Grind_Preprocessor_instInhabitedState___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = lean_unsigned_to_nat(0u); -x_2 = lean_mk_empty_array_with_capacity(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Meta_Grind_Preprocessor_instInhabitedState___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_Grind_Preprocessor_instInhabitedState___closed__1; -x_2 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; -} -} -static size_t _init_l_Lean_Meta_Grind_Preprocessor_instInhabitedState___closed__3() { -_start: -{ -lean_object* x_1; size_t x_2; -x_1 = lean_unsigned_to_nat(0u); -x_2 = lean_usize_of_nat(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Meta_Grind_Preprocessor_instInhabitedState___closed__4() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; size_t x_4; lean_object* x_5; -x_1 = l_Lean_Meta_Grind_Preprocessor_instInhabitedState___closed__2; -x_2 = l_Lean_Meta_Grind_Preprocessor_instInhabitedState___closed__1; -x_3 = lean_unsigned_to_nat(0u); -x_4 = l_Lean_Meta_Grind_Preprocessor_instInhabitedState___closed__3; -x_5 = lean_alloc_ctor(0, 4, sizeof(size_t)*1); -lean_ctor_set(x_5, 0, x_1); -lean_ctor_set(x_5, 1, x_2); -lean_ctor_set(x_5, 2, x_3); -lean_ctor_set(x_5, 3, x_3); -lean_ctor_set_usize(x_5, 4, x_4); -return x_5; -} -} -static lean_object* _init_l_Lean_Meta_Grind_Preprocessor_instInhabitedState() { -_start: -{ -lean_object* x_1; -x_1 = l_Lean_Meta_Grind_Preprocessor_instInhabitedState___closed__4; -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_Grind_Preprocessor_PreM_run___rarg___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = lean_unsigned_to_nat(32u); -x_2 = lean_mk_empty_array_with_capacity(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Meta_Grind_Preprocessor_PreM_run___rarg___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_Grind_Preprocessor_PreM_run___rarg___closed__1; -x_2 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Meta_Grind_Preprocessor_PreM_run___rarg___closed__3() { -_start: -{ -size_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = 5; -x_2 = l_Lean_Meta_Grind_Preprocessor_PreM_run___rarg___closed__2; -x_3 = l_Lean_Meta_Grind_Preprocessor_PreM_run___rarg___closed__1; -x_4 = lean_unsigned_to_nat(0u); -x_5 = lean_alloc_ctor(0, 4, sizeof(size_t)*1); -lean_ctor_set(x_5, 0, x_2); -lean_ctor_set(x_5, 1, x_3); -lean_ctor_set(x_5, 2, x_4); -lean_ctor_set(x_5, 3, x_4); -lean_ctor_set_usize(x_5, 4, x_1); -return x_5; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Preprocessor_PreM_run___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, 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_Meta_Grind_Preprocessor_PreM_run___rarg___closed__3; -x_11 = lean_st_mk_ref(x_10, x_9); -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); -x_13 = lean_ctor_get(x_11, 1); -lean_inc(x_13); -lean_dec(x_11); -lean_inc(x_12); -x_14 = lean_apply_9(x_1, x_12, x_2, x_3, x_4, x_5, x_6, x_7, x_8, 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_st_ref_get(x_12, x_16); -lean_dec(x_12); -x_18 = !lean_is_exclusive(x_17); -if (x_18 == 0) -{ -lean_object* x_19; -x_19 = lean_ctor_get(x_17, 0); -lean_dec(x_19); -lean_ctor_set(x_17, 0, x_15); -return x_17; -} -else -{ -lean_object* x_20; lean_object* x_21; -x_20 = lean_ctor_get(x_17, 1); -lean_inc(x_20); -lean_dec(x_17); -x_21 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_21, 0, x_15); -lean_ctor_set(x_21, 1, x_20); -return x_21; -} -} -else -{ -uint8_t x_22; -lean_dec(x_12); -x_22 = !lean_is_exclusive(x_14); -if (x_22 == 0) -{ -return x_14; -} -else -{ -lean_object* x_23; lean_object* x_24; lean_object* x_25; -x_23 = lean_ctor_get(x_14, 0); -x_24 = lean_ctor_get(x_14, 1); -lean_inc(x_24); -lean_inc(x_23); -lean_dec(x_14); -x_25 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_25, 0, x_23); -lean_ctor_set(x_25, 1, x_24); -return x_25; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Preprocessor_PreM_run(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_Preprocessor_PreM_run___rarg), 9, 0); -return x_2; -} -} -LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Meta_Grind_Preprocessor_introNext___spec__1___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -lean_object* x_12; -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -x_12 = lean_apply_9(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_12) == 0) -{ -lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_13 = lean_ctor_get(x_12, 0); -lean_inc(x_13); -x_14 = lean_ctor_get(x_12, 1); -lean_inc(x_14); -lean_dec(x_12); -x_15 = lean_apply_10(x_2, x_13, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_14); -return x_15; -} -else -{ -uint8_t x_16; -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_16 = !lean_is_exclusive(x_12); -if (x_16 == 0) -{ -return x_12; -} -else -{ -lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_17 = lean_ctor_get(x_12, 0); -x_18 = lean_ctor_get(x_12, 1); -lean_inc(x_18); -lean_inc(x_17); -lean_dec(x_12); -x_19 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_19, 0, x_17); -lean_ctor_set(x_19, 1, x_18); -return x_19; -} -} -} -} -LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Meta_Grind_Preprocessor_introNext___spec__1(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_Preprocessor_introNext___spec__1___rarg), 11, 0); -return x_3; -} -} -LEAN_EXPORT lean_object* l_Lean_MVarId_withContext___at_Lean_Meta_Grind_Preprocessor_introNext___spec__2___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -lean_object* x_12; lean_object* x_13; -x_12 = lean_apply_4(x_2, x_3, x_4, x_5, x_6); -x_13 = l___private_Lean_Meta_Basic_0__Lean_Meta_withMVarContextImp___rarg(x_1, x_12, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_13) == 0) -{ -uint8_t x_14; -x_14 = !lean_is_exclusive(x_13); -if (x_14 == 0) -{ -return x_13; -} -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; -} -} -else -{ -uint8_t x_18; -x_18 = !lean_is_exclusive(x_13); -if (x_18 == 0) -{ -return x_13; -} -else -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_19 = lean_ctor_get(x_13, 0); -x_20 = lean_ctor_get(x_13, 1); -lean_inc(x_20); -lean_inc(x_19); -lean_dec(x_13); -x_21 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_21, 0, x_19); -lean_ctor_set(x_21, 1, x_20); -return x_21; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_MVarId_withContext___at_Lean_Meta_Grind_Preprocessor_introNext___spec__2(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_MVarId_withContext___at_Lean_Meta_Grind_Preprocessor_introNext___spec__2___rarg), 11, 0); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Lean_mkFreshId___at_Lean_Meta_Grind_Preprocessor_introNext___spec__4___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_get(x_1, x_2); -x_4 = lean_ctor_get(x_3, 0); -lean_inc(x_4); -x_5 = lean_ctor_get(x_4, 2); -lean_inc(x_5); -lean_dec(x_4); -x_6 = lean_ctor_get(x_3, 1); -lean_inc(x_6); -lean_dec(x_3); -x_7 = !lean_is_exclusive(x_5); -if (x_7 == 0) -{ -lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; -x_8 = lean_ctor_get(x_5, 0); -x_9 = lean_ctor_get(x_5, 1); -lean_inc(x_9); -lean_inc(x_8); -x_10 = l_Lean_Name_num___override(x_8, x_9); -x_11 = lean_unsigned_to_nat(1u); -x_12 = lean_nat_add(x_9, x_11); -lean_dec(x_9); -lean_ctor_set(x_5, 1, x_12); -x_13 = lean_st_ref_take(x_1, 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_is_exclusive(x_14); -if (x_16 == 0) -{ -lean_object* x_17; lean_object* x_18; uint8_t x_19; -x_17 = lean_ctor_get(x_14, 2); -lean_dec(x_17); -lean_ctor_set(x_14, 2, x_5); -x_18 = lean_st_ref_set(x_1, x_14, x_15); -x_19 = !lean_is_exclusive(x_18); -if (x_19 == 0) -{ -lean_object* x_20; -x_20 = lean_ctor_get(x_18, 0); -lean_dec(x_20); -lean_ctor_set(x_18, 0, x_10); -return x_18; -} -else -{ -lean_object* x_21; lean_object* x_22; -x_21 = lean_ctor_get(x_18, 1); -lean_inc(x_21); -lean_dec(x_18); -x_22 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_22, 0, x_10); -lean_ctor_set(x_22, 1, x_21); -return x_22; -} -} -else -{ -lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_23 = lean_ctor_get(x_14, 0); -x_24 = lean_ctor_get(x_14, 1); -x_25 = lean_ctor_get(x_14, 3); -x_26 = lean_ctor_get(x_14, 4); -x_27 = lean_ctor_get(x_14, 5); -x_28 = lean_ctor_get(x_14, 6); -x_29 = lean_ctor_get(x_14, 7); -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_14); -x_30 = lean_alloc_ctor(0, 8, 0); -lean_ctor_set(x_30, 0, x_23); -lean_ctor_set(x_30, 1, x_24); -lean_ctor_set(x_30, 2, x_5); -lean_ctor_set(x_30, 3, x_25); -lean_ctor_set(x_30, 4, x_26); -lean_ctor_set(x_30, 5, x_27); -lean_ctor_set(x_30, 6, x_28); -lean_ctor_set(x_30, 7, x_29); -x_31 = lean_st_ref_set(x_1, x_30, x_15); -x_32 = lean_ctor_get(x_31, 1); -lean_inc(x_32); -if (lean_is_exclusive(x_31)) { - lean_ctor_release(x_31, 0); - lean_ctor_release(x_31, 1); - x_33 = x_31; -} else { - lean_dec_ref(x_31); - x_33 = lean_box(0); -} -if (lean_is_scalar(x_33)) { - x_34 = lean_alloc_ctor(0, 2, 0); -} else { - x_34 = x_33; -} -lean_ctor_set(x_34, 0, x_10); -lean_ctor_set(x_34, 1, x_32); -return x_34; -} -} -else -{ -lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; 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_35 = lean_ctor_get(x_5, 0); -x_36 = lean_ctor_get(x_5, 1); -lean_inc(x_36); -lean_inc(x_35); -lean_dec(x_5); -lean_inc(x_36); -lean_inc(x_35); -x_37 = l_Lean_Name_num___override(x_35, x_36); -x_38 = lean_unsigned_to_nat(1u); -x_39 = lean_nat_add(x_36, x_38); -lean_dec(x_36); -x_40 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_40, 0, x_35); -lean_ctor_set(x_40, 1, x_39); -x_41 = lean_st_ref_take(x_1, 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 = lean_ctor_get(x_42, 0); -lean_inc(x_44); -x_45 = lean_ctor_get(x_42, 1); -lean_inc(x_45); -x_46 = lean_ctor_get(x_42, 3); -lean_inc(x_46); -x_47 = lean_ctor_get(x_42, 4); -lean_inc(x_47); -x_48 = lean_ctor_get(x_42, 5); -lean_inc(x_48); -x_49 = lean_ctor_get(x_42, 6); -lean_inc(x_49); -x_50 = lean_ctor_get(x_42, 7); -lean_inc(x_50); -if (lean_is_exclusive(x_42)) { - lean_ctor_release(x_42, 0); - lean_ctor_release(x_42, 1); - lean_ctor_release(x_42, 2); - lean_ctor_release(x_42, 3); - lean_ctor_release(x_42, 4); - lean_ctor_release(x_42, 5); - lean_ctor_release(x_42, 6); - lean_ctor_release(x_42, 7); - x_51 = x_42; -} else { - lean_dec_ref(x_42); - x_51 = lean_box(0); -} -if (lean_is_scalar(x_51)) { - x_52 = lean_alloc_ctor(0, 8, 0); -} else { - x_52 = x_51; -} -lean_ctor_set(x_52, 0, x_44); -lean_ctor_set(x_52, 1, x_45); -lean_ctor_set(x_52, 2, x_40); -lean_ctor_set(x_52, 3, x_46); -lean_ctor_set(x_52, 4, x_47); -lean_ctor_set(x_52, 5, x_48); -lean_ctor_set(x_52, 6, x_49); -lean_ctor_set(x_52, 7, x_50); -x_53 = lean_st_ref_set(x_1, x_52, x_43); -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); -} -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_37); -lean_ctor_set(x_56, 1, x_54); -return x_56; -} -} -} -LEAN_EXPORT lean_object* l_Lean_mkFreshId___at_Lean_Meta_Grind_Preprocessor_introNext___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: -{ -lean_object* x_8; -x_8 = lean_alloc_closure((void*)(l_Lean_mkFreshId___at_Lean_Meta_Grind_Preprocessor_introNext___spec__4___rarg___boxed), 2, 0); -return x_8; -} -} -LEAN_EXPORT lean_object* l_Lean_mkFreshFVarId___at_Lean_Meta_Grind_Preprocessor_introNext___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -lean_object* x_10; uint8_t x_11; -x_10 = l_Lean_mkFreshId___at_Lean_Meta_Grind_Preprocessor_introNext___spec__4___rarg(x_8, x_9); -x_11 = !lean_is_exclusive(x_10); -if (x_11 == 0) -{ -return x_10; -} -else -{ -lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_12 = lean_ctor_get(x_10, 0); -x_13 = lean_ctor_get(x_10, 1); -lean_inc(x_13); -lean_inc(x_12); -lean_dec(x_10); -x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_12); -lean_ctor_set(x_14, 1, x_13); -return x_14; -} -} -} -LEAN_EXPORT lean_object* l_Lean_MVarId_assign___at_Lean_Meta_Grind_Preprocessor_introNext___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; -x_12 = lean_st_ref_take(x_8, x_11); -x_13 = lean_ctor_get(x_12, 0); -lean_inc(x_13); -x_14 = lean_ctor_get(x_13, 0); -lean_inc(x_14); -x_15 = lean_ctor_get(x_12, 1); -lean_inc(x_15); -lean_dec(x_12); -x_16 = !lean_is_exclusive(x_13); -if (x_16 == 0) -{ -lean_object* x_17; uint8_t x_18; -x_17 = lean_ctor_get(x_13, 0); -lean_dec(x_17); -x_18 = !lean_is_exclusive(x_14); -if (x_18 == 0) -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; -x_19 = lean_ctor_get(x_14, 7); -x_20 = l_Lean_PersistentHashMap_insert___at_Lean_MVarId_assign___spec__1(x_19, x_1, x_2); -lean_ctor_set(x_14, 7, x_20); -x_21 = lean_st_ref_set(x_8, x_13, x_15); -x_22 = !lean_is_exclusive(x_21); -if (x_22 == 0) -{ -lean_object* x_23; lean_object* x_24; -x_23 = lean_ctor_get(x_21, 0); -lean_dec(x_23); -x_24 = lean_box(0); -lean_ctor_set(x_21, 0, x_24); -return x_21; -} -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 = lean_box(0); -x_27 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_27, 0, x_26); -lean_ctor_set(x_27, 1, x_25); -return x_27; -} -} -else -{ -lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; -x_28 = lean_ctor_get(x_14, 0); -x_29 = lean_ctor_get(x_14, 1); -x_30 = lean_ctor_get(x_14, 2); -x_31 = lean_ctor_get(x_14, 3); -x_32 = lean_ctor_get(x_14, 4); -x_33 = lean_ctor_get(x_14, 5); -x_34 = lean_ctor_get(x_14, 6); -x_35 = lean_ctor_get(x_14, 7); -x_36 = lean_ctor_get(x_14, 8); -lean_inc(x_36); -lean_inc(x_35); -lean_inc(x_34); -lean_inc(x_33); -lean_inc(x_32); -lean_inc(x_31); -lean_inc(x_30); -lean_inc(x_29); -lean_inc(x_28); -lean_dec(x_14); -x_37 = l_Lean_PersistentHashMap_insert___at_Lean_MVarId_assign___spec__1(x_35, x_1, x_2); -x_38 = lean_alloc_ctor(0, 9, 0); -lean_ctor_set(x_38, 0, x_28); -lean_ctor_set(x_38, 1, x_29); -lean_ctor_set(x_38, 2, x_30); -lean_ctor_set(x_38, 3, x_31); -lean_ctor_set(x_38, 4, x_32); -lean_ctor_set(x_38, 5, x_33); -lean_ctor_set(x_38, 6, x_34); -lean_ctor_set(x_38, 7, x_37); -lean_ctor_set(x_38, 8, x_36); -lean_ctor_set(x_13, 0, x_38); -x_39 = lean_st_ref_set(x_8, x_13, x_15); -x_40 = lean_ctor_get(x_39, 1); -lean_inc(x_40); -if (lean_is_exclusive(x_39)) { - lean_ctor_release(x_39, 0); - lean_ctor_release(x_39, 1); - x_41 = x_39; -} else { - lean_dec_ref(x_39); - x_41 = lean_box(0); -} -x_42 = lean_box(0); -if (lean_is_scalar(x_41)) { - x_43 = lean_alloc_ctor(0, 2, 0); -} else { - x_43 = x_41; -} -lean_ctor_set(x_43, 0, x_42); -lean_ctor_set(x_43, 1, x_40); -return x_43; -} -} -else -{ -lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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_44 = lean_ctor_get(x_13, 1); -x_45 = lean_ctor_get(x_13, 2); -x_46 = lean_ctor_get(x_13, 3); -x_47 = lean_ctor_get(x_13, 4); -lean_inc(x_47); -lean_inc(x_46); -lean_inc(x_45); -lean_inc(x_44); -lean_dec(x_13); -x_48 = lean_ctor_get(x_14, 0); -lean_inc(x_48); -x_49 = lean_ctor_get(x_14, 1); -lean_inc(x_49); -x_50 = lean_ctor_get(x_14, 2); -lean_inc(x_50); -x_51 = lean_ctor_get(x_14, 3); -lean_inc(x_51); -x_52 = lean_ctor_get(x_14, 4); -lean_inc(x_52); -x_53 = lean_ctor_get(x_14, 5); -lean_inc(x_53); -x_54 = lean_ctor_get(x_14, 6); -lean_inc(x_54); -x_55 = lean_ctor_get(x_14, 7); -lean_inc(x_55); -x_56 = lean_ctor_get(x_14, 8); -lean_inc(x_56); -if (lean_is_exclusive(x_14)) { - lean_ctor_release(x_14, 0); - lean_ctor_release(x_14, 1); - lean_ctor_release(x_14, 2); - lean_ctor_release(x_14, 3); - lean_ctor_release(x_14, 4); - lean_ctor_release(x_14, 5); - lean_ctor_release(x_14, 6); - lean_ctor_release(x_14, 7); - lean_ctor_release(x_14, 8); - x_57 = x_14; -} else { - lean_dec_ref(x_14); - x_57 = lean_box(0); -} -x_58 = l_Lean_PersistentHashMap_insert___at_Lean_MVarId_assign___spec__1(x_55, x_1, x_2); -if (lean_is_scalar(x_57)) { - x_59 = lean_alloc_ctor(0, 9, 0); -} else { - x_59 = x_57; -} -lean_ctor_set(x_59, 0, x_48); -lean_ctor_set(x_59, 1, x_49); -lean_ctor_set(x_59, 2, x_50); -lean_ctor_set(x_59, 3, x_51); -lean_ctor_set(x_59, 4, x_52); -lean_ctor_set(x_59, 5, x_53); -lean_ctor_set(x_59, 6, x_54); -lean_ctor_set(x_59, 7, x_58); -lean_ctor_set(x_59, 8, x_56); -x_60 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_60, 0, x_59); -lean_ctor_set(x_60, 1, x_44); -lean_ctor_set(x_60, 2, x_45); -lean_ctor_set(x_60, 3, x_46); -lean_ctor_set(x_60, 4, x_47); -x_61 = lean_st_ref_set(x_8, x_60, x_15); -x_62 = lean_ctor_get(x_61, 1); -lean_inc(x_62); -if (lean_is_exclusive(x_61)) { - lean_ctor_release(x_61, 0); - lean_ctor_release(x_61, 1); - x_63 = x_61; -} else { - lean_dec_ref(x_61); - x_63 = lean_box(0); -} -x_64 = lean_box(0); -if (lean_is_scalar(x_63)) { - x_65 = lean_alloc_ctor(0, 2, 0); -} else { - x_65 = x_63; -} -lean_ctor_set(x_65, 0, x_64); -lean_ctor_set(x_65, 1, x_62); -return x_65; -} -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Preprocessor_introNext___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -lean_object* x_10; -x_10 = l_Lean_FVarId_getDecl(x_1, x_5, x_6, x_7, x_8, x_9); -return x_10; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Preprocessor_introNext___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, uint8_t x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19, lean_object* x_20) { -_start: -{ -lean_object* x_21; lean_object* x_22; -x_21 = l_Lean_LocalDecl_type(x_11); -lean_inc(x_19); -lean_inc(x_18); -lean_inc(x_17); -lean_inc(x_16); -lean_inc(x_21); -x_22 = l_Lean_Meta_isProp(x_21, x_16, x_17, x_18, x_19, x_20); -if (lean_obj_tag(x_22) == 0) -{ -lean_object* x_23; uint8_t x_24; -x_23 = lean_ctor_get(x_22, 0); -lean_inc(x_23); -x_24 = lean_unbox(x_23); -lean_dec(x_23); -if (x_24 == 0) -{ -uint8_t x_25; -lean_dec(x_21); -lean_dec(x_19); -lean_dec(x_18); -lean_dec(x_17); -lean_dec(x_16); -x_25 = !lean_is_exclusive(x_22); -if (x_25 == 0) -{ -lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_26 = lean_ctor_get(x_22, 0); -lean_dec(x_26); -x_27 = lean_alloc_ctor(0, 8, 1); -lean_ctor_set(x_27, 0, x_1); -lean_ctor_set(x_27, 1, x_2); -lean_ctor_set(x_27, 2, x_3); -lean_ctor_set(x_27, 3, x_4); -lean_ctor_set(x_27, 4, x_5); -lean_ctor_set(x_27, 5, x_6); -lean_ctor_set(x_27, 6, x_8); -lean_ctor_set(x_27, 7, x_9); -lean_ctor_set_uint8(x_27, sizeof(void*)*8, x_7); -x_28 = lean_alloc_ctor(3, 2, 0); -lean_ctor_set(x_28, 0, x_10); -lean_ctor_set(x_28, 1, x_27); -lean_ctor_set(x_22, 0, x_28); -return x_22; -} -else -{ -lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; -x_29 = lean_ctor_get(x_22, 1); -lean_inc(x_29); -lean_dec(x_22); -x_30 = lean_alloc_ctor(0, 8, 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_4); -lean_ctor_set(x_30, 4, x_5); -lean_ctor_set(x_30, 5, x_6); -lean_ctor_set(x_30, 6, x_8); -lean_ctor_set(x_30, 7, x_9); -lean_ctor_set_uint8(x_30, sizeof(void*)*8, x_7); -x_31 = lean_alloc_ctor(3, 2, 0); -lean_ctor_set(x_31, 0, x_10); -lean_ctor_set(x_31, 1, x_30); -x_32 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_32, 0, x_31); -lean_ctor_set(x_32, 1, x_29); -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; -x_33 = lean_ctor_get(x_22, 1); -lean_inc(x_33); -lean_dec(x_22); -x_34 = l_Lean_LocalDecl_userName(x_11); -x_35 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_34, x_18, x_19, 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 = l_Lean_Expr_fvar___override(x_10); -x_39 = l_Lean_MVarId_assert(x_1, x_36, x_21, x_38, x_16, x_17, x_18, x_19, x_37); -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; lean_object* x_43; -x_41 = lean_ctor_get(x_39, 0); -x_42 = lean_alloc_ctor(0, 8, 1); -lean_ctor_set(x_42, 0, x_41); -lean_ctor_set(x_42, 1, x_2); -lean_ctor_set(x_42, 2, x_3); -lean_ctor_set(x_42, 3, x_4); -lean_ctor_set(x_42, 4, x_5); -lean_ctor_set(x_42, 5, x_6); -lean_ctor_set(x_42, 6, x_8); -lean_ctor_set(x_42, 7, x_9); -lean_ctor_set_uint8(x_42, sizeof(void*)*8, x_7); -x_43 = lean_alloc_ctor(2, 1, 0); -lean_ctor_set(x_43, 0, x_42); -lean_ctor_set(x_39, 0, x_43); -return x_39; -} -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_39, 0); -x_45 = lean_ctor_get(x_39, 1); -lean_inc(x_45); -lean_inc(x_44); -lean_dec(x_39); -x_46 = lean_alloc_ctor(0, 8, 1); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_2); -lean_ctor_set(x_46, 2, x_3); -lean_ctor_set(x_46, 3, x_4); -lean_ctor_set(x_46, 4, x_5); -lean_ctor_set(x_46, 5, x_6); -lean_ctor_set(x_46, 6, x_8); -lean_ctor_set(x_46, 7, x_9); -lean_ctor_set_uint8(x_46, sizeof(void*)*8, x_7); -x_47 = lean_alloc_ctor(2, 1, 0); -lean_ctor_set(x_47, 0, 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_49; -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -x_49 = !lean_is_exclusive(x_39); -if (x_49 == 0) -{ -return x_39; -} -else -{ -lean_object* x_50; lean_object* x_51; lean_object* x_52; -x_50 = lean_ctor_get(x_39, 0); -x_51 = lean_ctor_get(x_39, 1); -lean_inc(x_51); -lean_inc(x_50); -lean_dec(x_39); -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_53; -lean_dec(x_21); -lean_dec(x_19); -lean_dec(x_18); -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_53 = !lean_is_exclusive(x_22); -if (x_53 == 0) -{ -return x_22; -} -else -{ -lean_object* x_54; lean_object* x_55; lean_object* x_56; -x_54 = lean_ctor_get(x_22, 0); -x_55 = lean_ctor_get(x_22, 1); -lean_inc(x_55); -lean_inc(x_54); -lean_dec(x_22); -x_56 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_56, 0, x_54); -lean_ctor_set(x_56, 1, x_55); -return x_56; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Preprocessor_introNext___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; -x_10 = l_Lean_Meta_isProp(x_1, x_5, x_6, x_7, x_8, x_9); -return x_10; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Preprocessor_introNext___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_11; uint8_t x_12; uint8_t x_13; lean_object* x_14; -x_11 = 0; -x_12 = 1; -x_13 = 1; -x_14 = l_Lean_Meta_mkLambdaFVars(x_1, x_2, x_11, x_12, x_11, x_13, x_6, x_7, x_8, x_9, x_10); -return x_14; -} -} -static lean_object* _init_l_Lean_Meta_Grind_Preprocessor_introNext___lambda__5___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean", 4, 4); -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_Grind_Preprocessor_introNext___lambda__5___closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Grind", 5, 5); -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_Grind_Preprocessor_introNext___lambda__5___closed__3() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("intro_with_eq", 13, 13); -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_Grind_Preprocessor_introNext___lambda__5___closed__4() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Meta_Grind_Preprocessor_introNext___lambda__5___closed__1; -x_2 = l_Lean_Meta_Grind_Preprocessor_introNext___lambda__5___closed__2; -x_3 = l_Lean_Meta_Grind_Preprocessor_introNext___lambda__5___closed__3; -x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); -return x_4; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Preprocessor_introNext___lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, uint8_t x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19, lean_object* x_20, lean_object* x_21, lean_object* x_22, lean_object* x_23, lean_object* x_24, lean_object* x_25, lean_object* x_26) { -_start: -{ -if (lean_obj_tag(x_1) == 0) -{ -lean_object* x_27; uint8_t x_28; -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -x_27 = l_Lean_MVarId_assign___at_Lean_Meta_Grind_Preprocessor_introNext___spec__5(x_2, x_17, x_18, x_19, x_20, x_21, x_22, x_23, x_24, x_25, x_26); -x_28 = !lean_is_exclusive(x_27); -if (x_28 == 0) -{ -lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_29 = lean_ctor_get(x_27, 0); -lean_dec(x_29); -x_30 = lean_alloc_ctor(0, 8, 1); -lean_ctor_set(x_30, 0, x_3); -lean_ctor_set(x_30, 1, x_4); -lean_ctor_set(x_30, 2, x_5); -lean_ctor_set(x_30, 3, x_6); -lean_ctor_set(x_30, 4, x_7); -lean_ctor_set(x_30, 5, x_8); -lean_ctor_set(x_30, 6, x_10); -lean_ctor_set(x_30, 7, x_11); -lean_ctor_set_uint8(x_30, sizeof(void*)*8, x_9); -x_31 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_31, 0, x_12); -lean_ctor_set(x_31, 1, x_30); -lean_ctor_set(x_27, 0, x_31); -return x_27; -} -else -{ -lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_32 = lean_ctor_get(x_27, 1); -lean_inc(x_32); -lean_dec(x_27); -x_33 = lean_alloc_ctor(0, 8, 1); -lean_ctor_set(x_33, 0, x_3); -lean_ctor_set(x_33, 1, x_4); -lean_ctor_set(x_33, 2, x_5); -lean_ctor_set(x_33, 3, x_6); -lean_ctor_set(x_33, 4, x_7); -lean_ctor_set(x_33, 5, x_8); -lean_ctor_set(x_33, 6, x_10); -lean_ctor_set(x_33, 7, x_11); -lean_ctor_set_uint8(x_33, sizeof(void*)*8, x_9); -x_34 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_34, 0, x_12); -lean_ctor_set(x_34, 1, x_33); -x_35 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_35, 0, x_34); -lean_ctor_set(x_35, 1, x_32); -return x_35; -} -} -else -{ -lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; uint8_t x_47; -x_36 = lean_ctor_get(x_1, 0); -x_37 = l_Lean_Meta_Grind_Preprocessor_introNext___lambda__5___closed__4; -lean_inc(x_13); -x_38 = l_Lean_Expr_const___override(x_37, x_13); -x_39 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_39, 0, x_17); -lean_ctor_set(x_39, 1, x_13); -lean_inc(x_36); -x_40 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_40, 0, x_36); -lean_ctor_set(x_40, 1, x_39); -x_41 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_41, 0, x_14); -lean_ctor_set(x_41, 1, x_40); -x_42 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_42, 0, x_15); -lean_ctor_set(x_42, 1, x_41); -x_43 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_43, 0, x_16); -lean_ctor_set(x_43, 1, x_42); -x_44 = lean_array_mk(x_43); -x_45 = l_Lean_mkAppN(x_38, x_44); -lean_dec(x_44); -x_46 = l_Lean_MVarId_assign___at_Lean_Meta_Grind_Preprocessor_introNext___spec__5(x_2, x_45, x_18, x_19, x_20, x_21, x_22, x_23, x_24, x_25, x_26); -x_47 = !lean_is_exclusive(x_46); -if (x_47 == 0) -{ -lean_object* x_48; lean_object* x_49; lean_object* x_50; -x_48 = lean_ctor_get(x_46, 0); -lean_dec(x_48); -x_49 = lean_alloc_ctor(0, 8, 1); -lean_ctor_set(x_49, 0, x_3); -lean_ctor_set(x_49, 1, x_4); -lean_ctor_set(x_49, 2, x_5); -lean_ctor_set(x_49, 3, x_6); -lean_ctor_set(x_49, 4, x_7); -lean_ctor_set(x_49, 5, x_8); -lean_ctor_set(x_49, 6, x_10); -lean_ctor_set(x_49, 7, x_11); -lean_ctor_set_uint8(x_49, sizeof(void*)*8, x_9); -x_50 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_50, 0, x_12); -lean_ctor_set(x_50, 1, x_49); -lean_ctor_set(x_46, 0, x_50); -return x_46; -} -else -{ -lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; -x_51 = lean_ctor_get(x_46, 1); -lean_inc(x_51); -lean_dec(x_46); -x_52 = lean_alloc_ctor(0, 8, 1); -lean_ctor_set(x_52, 0, x_3); -lean_ctor_set(x_52, 1, x_4); -lean_ctor_set(x_52, 2, x_5); -lean_ctor_set(x_52, 3, x_6); -lean_ctor_set(x_52, 4, x_7); -lean_ctor_set(x_52, 5, x_8); -lean_ctor_set(x_52, 6, x_10); -lean_ctor_set(x_52, 7, x_11); -lean_ctor_set_uint8(x_52, sizeof(void*)*8, x_9); -x_53 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_53, 0, x_12); -lean_ctor_set(x_53, 1, x_52); -x_54 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_54, 0, x_53); -lean_ctor_set(x_54, 1, x_51); -return x_54; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Preprocessor_introNext___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, uint8_t x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, uint8_t x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19, lean_object* x_20, lean_object* x_21) { -_start: -{ -if (x_12 == 0) -{ -uint8_t x_22; lean_object* x_23; -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_11); -x_22 = 1; -x_23 = l_Lean_Meta_intro1Core(x_1, x_22, x_17, x_18, x_19, x_20, x_21); -if (lean_obj_tag(x_23) == 0) -{ -uint8_t x_24; -x_24 = !lean_is_exclusive(x_23); -if (x_24 == 0) -{ -lean_object* x_25; uint8_t x_26; -x_25 = lean_ctor_get(x_23, 0); -x_26 = !lean_is_exclusive(x_25); -if (x_26 == 0) -{ -lean_object* x_27; lean_object* x_28; -x_27 = lean_ctor_get(x_25, 1); -x_28 = lean_alloc_ctor(0, 8, 1); -lean_ctor_set(x_28, 0, x_27); -lean_ctor_set(x_28, 1, x_2); -lean_ctor_set(x_28, 2, x_3); -lean_ctor_set(x_28, 3, x_4); -lean_ctor_set(x_28, 4, x_5); -lean_ctor_set(x_28, 5, x_6); -lean_ctor_set(x_28, 6, x_8); -lean_ctor_set(x_28, 7, x_9); -lean_ctor_set_uint8(x_28, sizeof(void*)*8, x_7); -lean_ctor_set_tag(x_25, 3); -lean_ctor_set(x_25, 1, x_28); -return x_23; -} -else -{ -lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; -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_dec(x_25); -x_31 = lean_alloc_ctor(0, 8, 1); -lean_ctor_set(x_31, 0, x_30); -lean_ctor_set(x_31, 1, x_2); -lean_ctor_set(x_31, 2, x_3); -lean_ctor_set(x_31, 3, x_4); -lean_ctor_set(x_31, 4, x_5); -lean_ctor_set(x_31, 5, x_6); -lean_ctor_set(x_31, 6, x_8); -lean_ctor_set(x_31, 7, x_9); -lean_ctor_set_uint8(x_31, sizeof(void*)*8, x_7); -x_32 = lean_alloc_ctor(3, 2, 0); -lean_ctor_set(x_32, 0, x_29); -lean_ctor_set(x_32, 1, x_31); -lean_ctor_set(x_23, 0, x_32); -return x_23; -} -} -else -{ -lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_33 = lean_ctor_get(x_23, 0); -x_34 = lean_ctor_get(x_23, 1); -lean_inc(x_34); -lean_inc(x_33); -lean_dec(x_23); -x_35 = lean_ctor_get(x_33, 0); -lean_inc(x_35); -x_36 = lean_ctor_get(x_33, 1); -lean_inc(x_36); -if (lean_is_exclusive(x_33)) { - lean_ctor_release(x_33, 0); - lean_ctor_release(x_33, 1); - x_37 = x_33; -} else { - lean_dec_ref(x_33); - x_37 = lean_box(0); -} -x_38 = lean_alloc_ctor(0, 8, 1); -lean_ctor_set(x_38, 0, x_36); -lean_ctor_set(x_38, 1, x_2); -lean_ctor_set(x_38, 2, x_3); -lean_ctor_set(x_38, 3, x_4); -lean_ctor_set(x_38, 4, x_5); -lean_ctor_set(x_38, 5, x_6); -lean_ctor_set(x_38, 6, x_8); -lean_ctor_set(x_38, 7, x_9); -lean_ctor_set_uint8(x_38, sizeof(void*)*8, x_7); -if (lean_is_scalar(x_37)) { - x_39 = lean_alloc_ctor(3, 2, 0); -} else { - x_39 = x_37; - lean_ctor_set_tag(x_39, 3); -} -lean_ctor_set(x_39, 0, x_35); -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_34); -return x_40; -} -} -else -{ -uint8_t x_41; -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -x_41 = !lean_is_exclusive(x_23); -if (x_41 == 0) -{ -return x_23; -} -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_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; -} -} -} -else -{ -lean_object* x_45; -lean_inc(x_1); -x_45 = l_Lean_MVarId_getTag(x_1, x_17, x_18, x_19, x_20, x_21); -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 = l_Lean_Expr_bindingBody_x21(x_10); -lean_inc(x_20); -lean_inc(x_19); -lean_inc(x_18); -lean_inc(x_17); -lean_inc(x_15); -lean_inc(x_11); -x_49 = l_Lean_Meta_Grind_pre(x_11, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_47); -if (lean_obj_tag(x_49) == 0) -{ -lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; uint8_t x_60; uint8_t x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; uint8_t x_66; lean_object* x_67; lean_object* x_68; uint8_t x_69; -x_50 = lean_ctor_get(x_49, 0); -lean_inc(x_50); -x_51 = lean_ctor_get(x_49, 1); -lean_inc(x_51); -lean_dec(x_49); -x_52 = l_Lean_mkFreshFVarId___at_Lean_Meta_Grind_Preprocessor_introNext___spec__3(x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, 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_ctor_get(x_17, 2); -lean_inc(x_55); -x_56 = l_Lean_Expr_bindingName_x21(x_10); -x_57 = lean_ctor_get(x_50, 0); -lean_inc(x_57); -x_58 = lean_ctor_get(x_50, 1); -lean_inc(x_58); -lean_dec(x_50); -x_59 = l_Lean_Expr_bindingInfo_x21(x_10); -x_60 = lean_unbox(x_59); -lean_dec(x_59); -x_61 = 0; -lean_inc(x_57); -lean_inc(x_53); -x_62 = l_Lean_LocalContext_mkLocalDecl(x_55, x_53, x_56, x_57, x_60, x_61); -x_63 = l_Lean_Meta_getLocalInstances(x_17, x_18, x_19, x_20, 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 = 2; -x_67 = lean_unsigned_to_nat(0u); -lean_inc(x_48); -x_68 = l_Lean_Meta_mkFreshExprMVarAt(x_62, x_64, x_48, x_66, x_46, x_67, x_17, x_18, x_19, x_20, x_65); -x_69 = !lean_is_exclusive(x_68); -if (x_69 == 0) -{ -lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; -x_70 = lean_ctor_get(x_68, 0); -x_71 = lean_ctor_get(x_68, 1); -x_72 = l_Lean_Expr_mvarId_x21(x_70); -lean_inc(x_53); -x_73 = l_Lean_Expr_fvar___override(x_53); -x_74 = lean_box(0); -lean_ctor_set_tag(x_68, 1); -lean_ctor_set(x_68, 1, x_74); -lean_ctor_set(x_68, 0, x_73); -x_75 = lean_array_mk(x_68); -x_76 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_Preprocessor_introNext___lambda__4___boxed), 10, 2); -lean_closure_set(x_76, 0, x_75); -lean_closure_set(x_76, 1, x_70); -x_77 = lean_alloc_closure((void*)(l_StateRefT_x27_lift___rarg___boxed), 2, 1); -lean_closure_set(x_77, 0, x_76); -x_78 = lean_box(x_7); -lean_inc(x_72); -x_79 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_Preprocessor_introNext___lambda__5___boxed), 26, 16); -lean_closure_set(x_79, 0, x_58); -lean_closure_set(x_79, 1, x_1); -lean_closure_set(x_79, 2, x_72); -lean_closure_set(x_79, 3, x_2); -lean_closure_set(x_79, 4, x_3); -lean_closure_set(x_79, 5, x_4); -lean_closure_set(x_79, 6, x_5); -lean_closure_set(x_79, 7, x_6); -lean_closure_set(x_79, 8, x_78); -lean_closure_set(x_79, 9, x_8); -lean_closure_set(x_79, 10, x_9); -lean_closure_set(x_79, 11, x_53); -lean_closure_set(x_79, 12, x_74); -lean_closure_set(x_79, 13, x_48); -lean_closure_set(x_79, 14, x_57); -lean_closure_set(x_79, 15, x_11); -x_80 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_Preprocessor_introNext___spec__1___rarg), 11, 2); -lean_closure_set(x_80, 0, x_77); -lean_closure_set(x_80, 1, x_79); -x_81 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_Preprocessor_introNext___spec__2___rarg(x_72, x_80, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_71); -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; -x_82 = lean_ctor_get(x_68, 0); -x_83 = lean_ctor_get(x_68, 1); -lean_inc(x_83); -lean_inc(x_82); -lean_dec(x_68); -x_84 = l_Lean_Expr_mvarId_x21(x_82); -lean_inc(x_53); -x_85 = l_Lean_Expr_fvar___override(x_53); -x_86 = lean_box(0); -x_87 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_87, 0, x_85); -lean_ctor_set(x_87, 1, x_86); -x_88 = lean_array_mk(x_87); -x_89 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_Preprocessor_introNext___lambda__4___boxed), 10, 2); -lean_closure_set(x_89, 0, x_88); -lean_closure_set(x_89, 1, x_82); -x_90 = lean_alloc_closure((void*)(l_StateRefT_x27_lift___rarg___boxed), 2, 1); -lean_closure_set(x_90, 0, x_89); -x_91 = lean_box(x_7); -lean_inc(x_84); -x_92 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_Preprocessor_introNext___lambda__5___boxed), 26, 16); -lean_closure_set(x_92, 0, x_58); -lean_closure_set(x_92, 1, x_1); -lean_closure_set(x_92, 2, x_84); -lean_closure_set(x_92, 3, x_2); -lean_closure_set(x_92, 4, x_3); -lean_closure_set(x_92, 5, x_4); -lean_closure_set(x_92, 6, x_5); -lean_closure_set(x_92, 7, x_6); -lean_closure_set(x_92, 8, x_91); -lean_closure_set(x_92, 9, x_8); -lean_closure_set(x_92, 10, x_9); -lean_closure_set(x_92, 11, x_53); -lean_closure_set(x_92, 12, x_86); -lean_closure_set(x_92, 13, x_48); -lean_closure_set(x_92, 14, x_57); -lean_closure_set(x_92, 15, x_11); -x_93 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_Preprocessor_introNext___spec__1___rarg), 11, 2); -lean_closure_set(x_93, 0, x_90); -lean_closure_set(x_93, 1, x_92); -x_94 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_Preprocessor_introNext___spec__2___rarg(x_84, x_93, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_83); -return x_94; -} -} -else -{ -uint8_t x_95; -lean_dec(x_48); -lean_dec(x_46); -lean_dec(x_20); -lean_dec(x_19); -lean_dec(x_18); -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_11); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_95 = !lean_is_exclusive(x_49); -if (x_95 == 0) -{ -return x_49; -} -else -{ -lean_object* x_96; lean_object* x_97; lean_object* x_98; -x_96 = lean_ctor_get(x_49, 0); -x_97 = lean_ctor_get(x_49, 1); -lean_inc(x_97); -lean_inc(x_96); -lean_dec(x_49); -x_98 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_98, 0, x_96); -lean_ctor_set(x_98, 1, x_97); -return x_98; -} -} -} -else -{ -uint8_t x_99; -lean_dec(x_20); -lean_dec(x_19); -lean_dec(x_18); -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_11); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_99 = !lean_is_exclusive(x_45); -if (x_99 == 0) -{ -return x_45; -} -else -{ -lean_object* x_100; lean_object* x_101; lean_object* x_102; -x_100 = lean_ctor_get(x_45, 0); -x_101 = lean_ctor_get(x_45, 1); -lean_inc(x_101); -lean_inc(x_100); -lean_dec(x_45); -x_102 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_102, 0, x_100); -lean_ctor_set(x_102, 1, x_101); -return x_102; -} -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Preprocessor_introNext(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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; uint8_t x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_11 = lean_ctor_get(x_1, 0); -lean_inc(x_11); -x_12 = lean_ctor_get(x_1, 1); -lean_inc(x_12); -x_13 = lean_ctor_get(x_1, 2); -lean_inc(x_13); -x_14 = lean_ctor_get(x_1, 3); -lean_inc(x_14); -x_15 = lean_ctor_get(x_1, 4); -lean_inc(x_15); -x_16 = lean_ctor_get(x_1, 5); -lean_inc(x_16); -x_17 = lean_ctor_get_uint8(x_1, sizeof(void*)*8); -x_18 = lean_ctor_get(x_1, 6); -lean_inc(x_18); -x_19 = lean_ctor_get(x_1, 7); -lean_inc(x_19); -lean_dec(x_1); -lean_inc(x_11); -x_20 = l_Lean_MVarId_getType(x_11, x_6, x_7, x_8, x_9, x_10); -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; lean_object* x_24; uint8_t x_42; -x_22 = lean_ctor_get(x_20, 0); -x_23 = lean_ctor_get(x_20, 1); -x_42 = l_Lean_Expr_isArrow(x_22); -if (x_42 == 0) -{ -uint8_t x_43; -x_43 = l_Lean_Expr_isLet(x_22); -if (x_43 == 0) -{ -uint8_t x_44; -x_44 = l_Lean_Expr_isForall(x_22); -lean_dec(x_22); -if (x_44 == 0) -{ -lean_object* x_45; -lean_dec(x_19); -lean_dec(x_18); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_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_45 = lean_box(0); -lean_ctor_set(x_20, 0, x_45); -return x_20; -} -else -{ -lean_object* x_46; -lean_free_object(x_20); -x_46 = lean_box(0); -x_24 = x_46; -goto block_41; -} -} -else -{ -lean_object* x_47; -lean_free_object(x_20); -lean_dec(x_22); -x_47 = lean_box(0); -x_24 = x_47; -goto block_41; -} -} -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_free_object(x_20); -x_48 = l_Lean_Expr_bindingDomain_x21(x_22); -lean_inc(x_48); -x_49 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_Preprocessor_introNext___lambda__3___boxed), 9, 1); -lean_closure_set(x_49, 0, x_48); -x_50 = lean_alloc_closure((void*)(l_StateRefT_x27_lift___rarg___boxed), 2, 1); -lean_closure_set(x_50, 0, x_49); -x_51 = lean_box(x_17); -lean_inc(x_11); -x_52 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_Preprocessor_introNext___lambda__6___boxed), 21, 11); -lean_closure_set(x_52, 0, x_11); -lean_closure_set(x_52, 1, x_12); -lean_closure_set(x_52, 2, x_13); -lean_closure_set(x_52, 3, x_14); -lean_closure_set(x_52, 4, x_15); -lean_closure_set(x_52, 5, x_16); -lean_closure_set(x_52, 6, x_51); -lean_closure_set(x_52, 7, x_18); -lean_closure_set(x_52, 8, x_19); -lean_closure_set(x_52, 9, x_22); -lean_closure_set(x_52, 10, x_48); -x_53 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_Preprocessor_introNext___spec__1___rarg), 11, 2); -lean_closure_set(x_53, 0, x_50); -lean_closure_set(x_53, 1, x_52); -x_54 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_Preprocessor_introNext___spec__2___rarg(x_11, x_53, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_23); -return x_54; -} -block_41: -{ -uint8_t x_25; lean_object* x_26; -lean_dec(x_24); -x_25 = 1; -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -x_26 = l_Lean_Meta_intro1Core(x_11, x_25, x_6, x_7, x_8, x_9, x_23); -if (lean_obj_tag(x_26) == 0) -{ -lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_27 = lean_ctor_get(x_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); -x_30 = lean_ctor_get(x_27, 1); -lean_inc(x_30); -lean_dec(x_27); -lean_inc(x_29); -x_31 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_Preprocessor_introNext___lambda__1___boxed), 9, 1); -lean_closure_set(x_31, 0, x_29); -x_32 = lean_alloc_closure((void*)(l_StateRefT_x27_lift___rarg___boxed), 2, 1); -lean_closure_set(x_32, 0, x_31); -x_33 = lean_box(x_17); -lean_inc(x_30); -x_34 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_Preprocessor_introNext___lambda__2___boxed), 20, 10); -lean_closure_set(x_34, 0, x_30); -lean_closure_set(x_34, 1, x_12); -lean_closure_set(x_34, 2, x_13); -lean_closure_set(x_34, 3, x_14); -lean_closure_set(x_34, 4, x_15); -lean_closure_set(x_34, 5, x_16); -lean_closure_set(x_34, 6, x_33); -lean_closure_set(x_34, 7, x_18); -lean_closure_set(x_34, 8, x_19); -lean_closure_set(x_34, 9, x_29); -x_35 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_Preprocessor_introNext___spec__1___rarg), 11, 2); -lean_closure_set(x_35, 0, x_32); -lean_closure_set(x_35, 1, x_34); -x_36 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_Preprocessor_introNext___spec__2___rarg(x_30, x_35, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_28); -return x_36; -} -else -{ -uint8_t x_37; -lean_dec(x_19); -lean_dec(x_18); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -x_37 = !lean_is_exclusive(x_26); -if (x_37 == 0) -{ -return x_26; -} -else -{ -lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_38 = lean_ctor_get(x_26, 0); -x_39 = lean_ctor_get(x_26, 1); -lean_inc(x_39); -lean_inc(x_38); -lean_dec(x_26); -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_55; lean_object* x_56; lean_object* x_57; uint8_t x_75; -x_55 = lean_ctor_get(x_20, 0); -x_56 = lean_ctor_get(x_20, 1); -lean_inc(x_56); -lean_inc(x_55); -lean_dec(x_20); -x_75 = l_Lean_Expr_isArrow(x_55); -if (x_75 == 0) -{ -uint8_t x_76; -x_76 = l_Lean_Expr_isLet(x_55); -if (x_76 == 0) -{ -uint8_t x_77; -x_77 = l_Lean_Expr_isForall(x_55); -lean_dec(x_55); -if (x_77 == 0) -{ -lean_object* x_78; lean_object* x_79; -lean_dec(x_19); -lean_dec(x_18); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_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_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_56); -return x_79; -} -else -{ -lean_object* x_80; -x_80 = lean_box(0); -x_57 = x_80; -goto block_74; -} -} -else -{ -lean_object* x_81; -lean_dec(x_55); -x_81 = lean_box(0); -x_57 = x_81; -goto block_74; -} -} -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; -x_82 = l_Lean_Expr_bindingDomain_x21(x_55); -lean_inc(x_82); -x_83 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_Preprocessor_introNext___lambda__3___boxed), 9, 1); -lean_closure_set(x_83, 0, x_82); -x_84 = lean_alloc_closure((void*)(l_StateRefT_x27_lift___rarg___boxed), 2, 1); -lean_closure_set(x_84, 0, x_83); -x_85 = lean_box(x_17); -lean_inc(x_11); -x_86 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_Preprocessor_introNext___lambda__6___boxed), 21, 11); -lean_closure_set(x_86, 0, x_11); -lean_closure_set(x_86, 1, x_12); -lean_closure_set(x_86, 2, x_13); -lean_closure_set(x_86, 3, x_14); -lean_closure_set(x_86, 4, x_15); -lean_closure_set(x_86, 5, x_16); -lean_closure_set(x_86, 6, x_85); -lean_closure_set(x_86, 7, x_18); -lean_closure_set(x_86, 8, x_19); -lean_closure_set(x_86, 9, x_55); -lean_closure_set(x_86, 10, x_82); -x_87 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_Preprocessor_introNext___spec__1___rarg), 11, 2); -lean_closure_set(x_87, 0, x_84); -lean_closure_set(x_87, 1, x_86); -x_88 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_Preprocessor_introNext___spec__2___rarg(x_11, x_87, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_56); -return x_88; -} -block_74: -{ -uint8_t x_58; lean_object* x_59; -lean_dec(x_57); -x_58 = 1; -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -x_59 = l_Lean_Meta_intro1Core(x_11, x_58, x_6, x_7, x_8, x_9, x_56); -if (lean_obj_tag(x_59) == 0) -{ -lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; -x_60 = lean_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, 1); -lean_inc(x_63); -lean_dec(x_60); -lean_inc(x_62); -x_64 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_Preprocessor_introNext___lambda__1___boxed), 9, 1); -lean_closure_set(x_64, 0, x_62); -x_65 = lean_alloc_closure((void*)(l_StateRefT_x27_lift___rarg___boxed), 2, 1); -lean_closure_set(x_65, 0, x_64); -x_66 = lean_box(x_17); -lean_inc(x_63); -x_67 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_Preprocessor_introNext___lambda__2___boxed), 20, 10); -lean_closure_set(x_67, 0, x_63); -lean_closure_set(x_67, 1, x_12); -lean_closure_set(x_67, 2, x_13); -lean_closure_set(x_67, 3, x_14); -lean_closure_set(x_67, 4, x_15); -lean_closure_set(x_67, 5, x_16); -lean_closure_set(x_67, 6, x_66); -lean_closure_set(x_67, 7, x_18); -lean_closure_set(x_67, 8, x_19); -lean_closure_set(x_67, 9, x_62); -x_68 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_Preprocessor_introNext___spec__1___rarg), 11, 2); -lean_closure_set(x_68, 0, x_65); -lean_closure_set(x_68, 1, x_67); -x_69 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_Preprocessor_introNext___spec__2___rarg(x_63, x_68, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_61); -return x_69; -} -else -{ -lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; -lean_dec(x_19); -lean_dec(x_18); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -x_70 = lean_ctor_get(x_59, 0); -lean_inc(x_70); -x_71 = lean_ctor_get(x_59, 1); -lean_inc(x_71); -if (lean_is_exclusive(x_59)) { - lean_ctor_release(x_59, 0); - lean_ctor_release(x_59, 1); - x_72 = x_59; -} else { - lean_dec_ref(x_59); - x_72 = lean_box(0); -} -if (lean_is_scalar(x_72)) { - x_73 = lean_alloc_ctor(1, 2, 0); -} else { - x_73 = x_72; -} -lean_ctor_set(x_73, 0, x_70); -lean_ctor_set(x_73, 1, x_71); -return x_73; -} -} -} -} -else -{ -uint8_t x_89; -lean_dec(x_19); -lean_dec(x_18); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_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_89 = !lean_is_exclusive(x_20); -if (x_89 == 0) -{ -return x_20; -} -else -{ -lean_object* x_90; lean_object* x_91; lean_object* x_92; -x_90 = lean_ctor_get(x_20, 0); -x_91 = lean_ctor_get(x_20, 1); -lean_inc(x_91); -lean_inc(x_90); -lean_dec(x_20); -x_92 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_92, 0, x_90); -lean_ctor_set(x_92, 1, x_91); -return x_92; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_mkFreshId___at_Lean_Meta_Grind_Preprocessor_introNext___spec__4___rarg___boxed(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; -x_3 = l_Lean_mkFreshId___at_Lean_Meta_Grind_Preprocessor_introNext___spec__4___rarg(x_1, x_2); -lean_dec(x_1); -return x_3; -} -} -LEAN_EXPORT lean_object* l_Lean_mkFreshId___at_Lean_Meta_Grind_Preprocessor_introNext___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: -{ -lean_object* x_8; -x_8 = l_Lean_mkFreshId___at_Lean_Meta_Grind_Preprocessor_introNext___spec__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_8; -} -} -LEAN_EXPORT lean_object* l_Lean_mkFreshFVarId___at_Lean_Meta_Grind_Preprocessor_introNext___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -lean_object* x_10; -x_10 = l_Lean_mkFreshFVarId___at_Lean_Meta_Grind_Preprocessor_introNext___spec__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_10; -} -} -LEAN_EXPORT lean_object* l_Lean_MVarId_assign___at_Lean_Meta_Grind_Preprocessor_introNext___spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -lean_object* x_12; -x_12 = l_Lean_MVarId_assign___at_Lean_Meta_Grind_Preprocessor_introNext___spec__5(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -return x_12; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Preprocessor_introNext___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -lean_object* x_10; -x_10 = l_Lean_Meta_Grind_Preprocessor_introNext___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_4); -lean_dec(x_3); -lean_dec(x_2); -return x_10; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Preprocessor_introNext___lambda__2___boxed(lean_object** _args) { -lean_object* x_1 = _args[0]; -lean_object* x_2 = _args[1]; -lean_object* x_3 = _args[2]; -lean_object* x_4 = _args[3]; -lean_object* x_5 = _args[4]; -lean_object* x_6 = _args[5]; -lean_object* x_7 = _args[6]; -lean_object* x_8 = _args[7]; -lean_object* x_9 = _args[8]; -lean_object* x_10 = _args[9]; -lean_object* x_11 = _args[10]; -lean_object* x_12 = _args[11]; -lean_object* x_13 = _args[12]; -lean_object* x_14 = _args[13]; -lean_object* x_15 = _args[14]; -lean_object* x_16 = _args[15]; -lean_object* x_17 = _args[16]; -lean_object* x_18 = _args[17]; -lean_object* x_19 = _args[18]; -lean_object* x_20 = _args[19]; -_start: -{ -uint8_t x_21; lean_object* x_22; -x_21 = lean_unbox(x_7); -lean_dec(x_7); -x_22 = l_Lean_Meta_Grind_Preprocessor_introNext___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_21, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -return x_22; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Preprocessor_introNext___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -lean_object* x_10; -x_10 = l_Lean_Meta_Grind_Preprocessor_introNext___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_10; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Preprocessor_introNext___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_11; -x_11 = l_Lean_Meta_Grind_Preprocessor_introNext___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); -lean_dec(x_3); -lean_dec(x_1); -return x_11; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Preprocessor_introNext___lambda__5___boxed(lean_object** _args) { -lean_object* x_1 = _args[0]; -lean_object* x_2 = _args[1]; -lean_object* x_3 = _args[2]; -lean_object* x_4 = _args[3]; -lean_object* x_5 = _args[4]; -lean_object* x_6 = _args[5]; -lean_object* x_7 = _args[6]; -lean_object* x_8 = _args[7]; -lean_object* x_9 = _args[8]; -lean_object* x_10 = _args[9]; -lean_object* x_11 = _args[10]; -lean_object* x_12 = _args[11]; -lean_object* x_13 = _args[12]; -lean_object* x_14 = _args[13]; -lean_object* x_15 = _args[14]; -lean_object* x_16 = _args[15]; -lean_object* x_17 = _args[16]; -lean_object* x_18 = _args[17]; -lean_object* x_19 = _args[18]; -lean_object* x_20 = _args[19]; -lean_object* x_21 = _args[20]; -lean_object* x_22 = _args[21]; -lean_object* x_23 = _args[22]; -lean_object* x_24 = _args[23]; -lean_object* x_25 = _args[24]; -lean_object* x_26 = _args[25]; -_start: -{ -uint8_t x_27; lean_object* x_28; -x_27 = lean_unbox(x_9); -lean_dec(x_9); -x_28 = l_Lean_Meta_Grind_Preprocessor_introNext___lambda__5(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_27, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_23, x_24, x_25, x_26); -lean_dec(x_25); -lean_dec(x_24); -lean_dec(x_23); -lean_dec(x_22); -lean_dec(x_21); -lean_dec(x_20); -lean_dec(x_19); -lean_dec(x_18); -lean_dec(x_1); -return x_28; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Preprocessor_introNext___lambda__6___boxed(lean_object** _args) { -lean_object* x_1 = _args[0]; -lean_object* x_2 = _args[1]; -lean_object* x_3 = _args[2]; -lean_object* x_4 = _args[3]; -lean_object* x_5 = _args[4]; -lean_object* x_6 = _args[5]; -lean_object* x_7 = _args[6]; -lean_object* x_8 = _args[7]; -lean_object* x_9 = _args[8]; -lean_object* x_10 = _args[9]; -lean_object* x_11 = _args[10]; -lean_object* x_12 = _args[11]; -lean_object* x_13 = _args[12]; -lean_object* x_14 = _args[13]; -lean_object* x_15 = _args[14]; -lean_object* x_16 = _args[15]; -lean_object* x_17 = _args[16]; -lean_object* x_18 = _args[17]; -lean_object* x_19 = _args[18]; -lean_object* x_20 = _args[19]; -lean_object* x_21 = _args[20]; -_start: -{ -uint8_t x_22; uint8_t x_23; lean_object* x_24; -x_22 = lean_unbox(x_7); -lean_dec(x_7); -x_23 = lean_unbox(x_12); -lean_dec(x_12); -x_24 = l_Lean_Meta_Grind_Preprocessor_introNext___lambda__6(x_1, x_2, x_3, x_4, x_5, x_6, x_22, x_8, x_9, x_10, x_11, x_23, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21); -lean_dec(x_10); -return x_24; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Preprocessor_pushResult(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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 = lean_st_ref_take(x_2, 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_PersistentArray_push___rarg(x_12, x_1); -x_15 = lean_st_ref_set(x_2, x_14, x_13); -x_16 = !lean_is_exclusive(x_15); -if (x_16 == 0) -{ -lean_object* x_17; lean_object* x_18; -x_17 = lean_ctor_get(x_15, 0); -lean_dec(x_17); -x_18 = lean_box(0); -lean_ctor_set(x_15, 0, x_18); -return x_15; -} -else -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_19 = lean_ctor_get(x_15, 1); -lean_inc(x_19); -lean_dec(x_15); -x_20 = lean_box(0); -x_21 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_21, 0, x_20); -lean_ctor_set(x_21, 1, x_19); -return x_21; -} -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Preprocessor_pushResult___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_11; -x_11 = l_Lean_Meta_Grind_Preprocessor_pushResult(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_11; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Preprocessor_isCasesCandidate(lean_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_FVarId_getType(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; lean_object* x_11; -x_9 = lean_ctor_get(x_7, 0); -x_10 = lean_ctor_get(x_7, 1); -x_11 = l_Lean_Expr_getAppFn(x_9); -lean_dec(x_9); -if (lean_obj_tag(x_11) == 4) -{ -lean_object* x_12; lean_object* x_13; -lean_free_object(x_7); -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); -lean_dec(x_11); -x_13 = l_Lean_Meta_Grind_isGrindCasesTarget(x_12, x_4, x_5, x_10); -lean_dec(x_12); -return x_13; -} -else -{ -uint8_t x_14; lean_object* x_15; -lean_dec(x_11); -x_14 = 0; -x_15 = lean_box(x_14); -lean_ctor_set(x_7, 0, x_15); -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 = l_Lean_Expr_getAppFn(x_16); -lean_dec(x_16); -if (lean_obj_tag(x_18) == 4) -{ -lean_object* x_19; lean_object* x_20; -x_19 = lean_ctor_get(x_18, 0); -lean_inc(x_19); -lean_dec(x_18); -x_20 = l_Lean_Meta_Grind_isGrindCasesTarget(x_19, x_4, x_5, x_17); -lean_dec(x_19); -return x_20; -} -else -{ -uint8_t x_21; lean_object* x_22; lean_object* x_23; -lean_dec(x_18); -x_21 = 0; -x_22 = lean_box(x_21); -x_23 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_23, 0, x_22); -lean_ctor_set(x_23, 1, x_17); -return x_23; -} -} -} -else -{ -uint8_t x_24; -x_24 = !lean_is_exclusive(x_7); -if (x_24 == 0) -{ -return x_7; -} -else -{ -lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_25 = lean_ctor_get(x_7, 0); -x_26 = lean_ctor_get(x_7, 1); -lean_inc(x_26); -lean_inc(x_25); -lean_dec(x_7); -x_27 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_27, 0, x_25); -lean_ctor_set(x_27, 1, x_26); -return x_27; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Preprocessor_isCasesCandidate___boxed(lean_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_Grind_Preprocessor_isCasesCandidate(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_List_mapTR_loop___at_Lean_Meta_Grind_Preprocessor_applyCases_x3f___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -if (lean_obj_tag(x_2) == 0) -{ -lean_object* x_4; -x_4 = l_List_reverse___rarg(x_3); -return x_4; -} -else -{ -uint8_t x_5; -x_5 = !lean_is_exclusive(x_2); -if (x_5 == 0) -{ -lean_object* x_6; lean_object* x_7; lean_object* x_8; 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; -x_6 = lean_ctor_get(x_2, 0); -x_7 = lean_ctor_get(x_2, 1); -x_8 = lean_ctor_get(x_1, 1); -x_9 = lean_ctor_get(x_1, 2); -x_10 = lean_ctor_get(x_1, 3); -x_11 = lean_ctor_get(x_1, 4); -x_12 = lean_ctor_get(x_1, 5); -x_13 = lean_ctor_get_uint8(x_1, sizeof(void*)*8); -x_14 = lean_ctor_get(x_1, 6); -x_15 = lean_ctor_get(x_1, 7); -lean_inc(x_15); -lean_inc(x_14); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -x_16 = lean_alloc_ctor(0, 8, 1); -lean_ctor_set(x_16, 0, x_6); -lean_ctor_set(x_16, 1, x_8); -lean_ctor_set(x_16, 2, x_9); -lean_ctor_set(x_16, 3, x_10); -lean_ctor_set(x_16, 4, x_11); -lean_ctor_set(x_16, 5, x_12); -lean_ctor_set(x_16, 6, x_14); -lean_ctor_set(x_16, 7, x_15); -lean_ctor_set_uint8(x_16, sizeof(void*)*8, x_13); -lean_ctor_set(x_2, 1, x_3); -lean_ctor_set(x_2, 0, x_16); -{ -lean_object* _tmp_1 = x_7; -lean_object* _tmp_2 = x_2; -x_2 = _tmp_1; -x_3 = _tmp_2; -} -goto _start; -} -else -{ -lean_object* x_18; 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_18 = lean_ctor_get(x_2, 0); -x_19 = lean_ctor_get(x_2, 1); -lean_inc(x_19); -lean_inc(x_18); -lean_dec(x_2); -x_20 = lean_ctor_get(x_1, 1); -x_21 = lean_ctor_get(x_1, 2); -x_22 = lean_ctor_get(x_1, 3); -x_23 = lean_ctor_get(x_1, 4); -x_24 = lean_ctor_get(x_1, 5); -x_25 = lean_ctor_get_uint8(x_1, sizeof(void*)*8); -x_26 = lean_ctor_get(x_1, 6); -x_27 = lean_ctor_get(x_1, 7); -lean_inc(x_27); -lean_inc(x_26); -lean_inc(x_24); -lean_inc(x_23); -lean_inc(x_22); -lean_inc(x_21); -lean_inc(x_20); -x_28 = lean_alloc_ctor(0, 8, 1); -lean_ctor_set(x_28, 0, x_18); -lean_ctor_set(x_28, 1, x_20); -lean_ctor_set(x_28, 2, x_21); -lean_ctor_set(x_28, 3, x_22); -lean_ctor_set(x_28, 4, x_23); -lean_ctor_set(x_28, 5, x_24); -lean_ctor_set(x_28, 6, x_26); -lean_ctor_set(x_28, 7, x_27); -lean_ctor_set_uint8(x_28, sizeof(void*)*8, x_25); -x_29 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_29, 0, x_28); -lean_ctor_set(x_29, 1, x_3); -x_2 = x_19; -x_3 = x_29; -goto _start; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Preprocessor_applyCases_x3f___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -lean_object* x_9; -lean_inc(x_4); -lean_inc(x_1); -x_9 = l_Lean_Meta_Grind_Preprocessor_isCasesCandidate(x_1, x_4, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_9) == 0) -{ -lean_object* x_10; uint8_t x_11; -x_10 = lean_ctor_get(x_9, 0); -lean_inc(x_10); -x_11 = lean_unbox(x_10); -lean_dec(x_10); -if (x_11 == 0) -{ -uint8_t x_12; -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_2); -lean_dec(x_1); -x_12 = !lean_is_exclusive(x_9); -if (x_12 == 0) -{ -lean_object* x_13; lean_object* x_14; -x_13 = lean_ctor_get(x_9, 0); -lean_dec(x_13); -x_14 = lean_box(0); -lean_ctor_set(x_9, 0, x_14); -return x_9; -} -else -{ -lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = lean_ctor_get(x_9, 1); -lean_inc(x_15); -lean_dec(x_9); -x_16 = lean_box(0); -x_17 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_17, 0, x_16); -lean_ctor_set(x_17, 1, x_15); -return x_17; -} -} -else -{ -lean_object* x_18; lean_object* x_19; -x_18 = lean_ctor_get(x_9, 1); -lean_inc(x_18); -lean_dec(x_9); -x_19 = l_Lean_Meta_Grind_cases(x_2, x_1, x_4, x_5, x_6, x_7, x_18); -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; lean_object* x_23; lean_object* x_24; -x_21 = lean_ctor_get(x_19, 0); -x_22 = lean_box(0); -x_23 = l_List_mapTR_loop___at_Lean_Meta_Grind_Preprocessor_applyCases_x3f___spec__1(x_3, x_21, x_22); -x_24 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_24, 0, x_23); -lean_ctor_set(x_19, 0, x_24); -return x_19; -} -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_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_box(0); -x_28 = l_List_mapTR_loop___at_Lean_Meta_Grind_Preprocessor_applyCases_x3f___spec__1(x_3, x_25, x_27); -x_29 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_29, 0, 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_26); -return x_30; -} -} -else -{ -uint8_t x_31; -x_31 = !lean_is_exclusive(x_19); -if (x_31 == 0) -{ -return x_19; -} -else -{ -lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_32 = lean_ctor_get(x_19, 0); -x_33 = lean_ctor_get(x_19, 1); -lean_inc(x_33); -lean_inc(x_32); -lean_dec(x_19); -x_34 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_34, 0, x_32); -lean_ctor_set(x_34, 1, x_33); -return x_34; -} -} -} -} -else -{ -uint8_t x_35; -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_35 = !lean_is_exclusive(x_9); -if (x_35 == 0) -{ -return x_9; -} -else -{ -lean_object* x_36; lean_object* x_37; lean_object* x_38; -x_36 = lean_ctor_get(x_9, 0); -x_37 = lean_ctor_get(x_9, 1); -lean_inc(x_37); -lean_inc(x_36); -lean_dec(x_9); -x_38 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_38, 0, x_36); -lean_ctor_set(x_38, 1, x_37); -return x_38; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Preprocessor_applyCases_x3f(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: -{ -lean_object* x_8; lean_object* x_9; lean_object* x_10; -x_8 = lean_ctor_get(x_1, 0); -lean_inc(x_8); -lean_inc(x_8); -x_9 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_Preprocessor_applyCases_x3f___lambda__1___boxed), 8, 3); -lean_closure_set(x_9, 0, x_2); -lean_closure_set(x_9, 1, x_8); -lean_closure_set(x_9, 2, x_1); -x_10 = l_Lean_MVarId_withContext___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___spec__2___rarg(x_8, x_9, x_3, x_4, x_5, x_6, x_7); -return x_10; -} -} -LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Meta_Grind_Preprocessor_applyCases_x3f___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -lean_object* x_4; -x_4 = l_List_mapTR_loop___at_Lean_Meta_Grind_Preprocessor_applyCases_x3f___spec__1(x_1, x_2, x_3); -lean_dec(x_1); -return x_4; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Preprocessor_applyCases_x3f___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -lean_object* x_9; -x_9 = l_Lean_Meta_Grind_Preprocessor_applyCases_x3f___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_Meta_Grind_Preprocessor_applyInjection_x3f(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: -{ -uint8_t x_8; -x_8 = !lean_is_exclusive(x_1); -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; lean_object* x_17; -x_9 = lean_ctor_get(x_1, 0); -x_10 = lean_ctor_get(x_1, 1); -x_11 = lean_ctor_get(x_1, 2); -x_12 = lean_ctor_get(x_1, 3); -x_13 = lean_ctor_get(x_1, 4); -x_14 = lean_ctor_get(x_1, 5); -x_15 = lean_ctor_get(x_1, 6); -x_16 = lean_ctor_get(x_1, 7); -x_17 = l_Lean_Meta_Grind_injection_x3f(x_9, x_2, x_3, x_4, x_5, x_6, x_7); -if (lean_obj_tag(x_17) == 0) -{ -lean_object* x_18; -x_18 = lean_ctor_get(x_17, 0); -lean_inc(x_18); -if (lean_obj_tag(x_18) == 0) -{ -uint8_t x_19; -lean_free_object(x_1); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -x_19 = !lean_is_exclusive(x_17); -if (x_19 == 0) -{ -lean_object* x_20; lean_object* x_21; -x_20 = lean_ctor_get(x_17, 0); -lean_dec(x_20); -x_21 = lean_box(0); -lean_ctor_set(x_17, 0, x_21); -return x_17; -} -else -{ -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_17, 1); -lean_inc(x_22); -lean_dec(x_17); -x_23 = lean_box(0); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_23); -lean_ctor_set(x_24, 1, x_22); -return x_24; -} -} -else -{ -uint8_t x_25; -x_25 = !lean_is_exclusive(x_17); -if (x_25 == 0) -{ -lean_object* x_26; uint8_t x_27; -x_26 = lean_ctor_get(x_17, 0); -lean_dec(x_26); -x_27 = !lean_is_exclusive(x_18); -if (x_27 == 0) -{ -lean_object* x_28; -x_28 = lean_ctor_get(x_18, 0); -lean_ctor_set(x_1, 0, x_28); -lean_ctor_set(x_18, 0, x_1); -return x_17; -} -else -{ -lean_object* x_29; lean_object* x_30; -x_29 = lean_ctor_get(x_18, 0); -lean_inc(x_29); -lean_dec(x_18); -lean_ctor_set(x_1, 0, x_29); -x_30 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_30, 0, x_1); -lean_ctor_set(x_17, 0, x_30); -return x_17; -} -} -else -{ -lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_31 = lean_ctor_get(x_17, 1); -lean_inc(x_31); -lean_dec(x_17); -x_32 = lean_ctor_get(x_18, 0); -lean_inc(x_32); -if (lean_is_exclusive(x_18)) { - lean_ctor_release(x_18, 0); - x_33 = x_18; -} else { - lean_dec_ref(x_18); - x_33 = lean_box(0); -} -lean_ctor_set(x_1, 0, x_32); -if (lean_is_scalar(x_33)) { - x_34 = lean_alloc_ctor(1, 1, 0); -} else { - x_34 = x_33; -} -lean_ctor_set(x_34, 0, x_1); -x_35 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_35, 0, x_34); -lean_ctor_set(x_35, 1, x_31); -return x_35; -} -} -} -else -{ -uint8_t x_36; -lean_free_object(x_1); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -x_36 = !lean_is_exclusive(x_17); -if (x_36 == 0) -{ -return x_17; -} -else -{ -lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_37 = lean_ctor_get(x_17, 0); -x_38 = lean_ctor_get(x_17, 1); -lean_inc(x_38); -lean_inc(x_37); -lean_dec(x_17); -x_39 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_39, 0, x_37); -lean_ctor_set(x_39, 1, x_38); -return x_39; -} -} -} -else -{ -lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; uint8_t x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; -x_40 = lean_ctor_get(x_1, 0); -x_41 = lean_ctor_get(x_1, 1); -x_42 = lean_ctor_get(x_1, 2); -x_43 = lean_ctor_get(x_1, 3); -x_44 = lean_ctor_get(x_1, 4); -x_45 = lean_ctor_get(x_1, 5); -x_46 = lean_ctor_get_uint8(x_1, sizeof(void*)*8); -x_47 = lean_ctor_get(x_1, 6); -x_48 = lean_ctor_get(x_1, 7); -lean_inc(x_48); -lean_inc(x_47); -lean_inc(x_45); -lean_inc(x_44); -lean_inc(x_43); -lean_inc(x_42); -lean_inc(x_41); -lean_inc(x_40); -lean_dec(x_1); -x_49 = l_Lean_Meta_Grind_injection_x3f(x_40, x_2, x_3, x_4, x_5, x_6, x_7); -if (lean_obj_tag(x_49) == 0) -{ -lean_object* x_50; -x_50 = lean_ctor_get(x_49, 0); -lean_inc(x_50); -if (lean_obj_tag(x_50) == 0) -{ -lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; -lean_dec(x_48); -lean_dec(x_47); -lean_dec(x_45); -lean_dec(x_44); -lean_dec(x_43); -lean_dec(x_42); -lean_dec(x_41); -x_51 = lean_ctor_get(x_49, 1); -lean_inc(x_51); -if (lean_is_exclusive(x_49)) { - lean_ctor_release(x_49, 0); - lean_ctor_release(x_49, 1); - x_52 = x_49; -} else { - lean_dec_ref(x_49); - x_52 = lean_box(0); -} -x_53 = lean_box(0); -if (lean_is_scalar(x_52)) { - x_54 = lean_alloc_ctor(0, 2, 0); -} else { - x_54 = x_52; -} -lean_ctor_set(x_54, 0, x_53); -lean_ctor_set(x_54, 1, x_51); -return x_54; -} -else -{ -lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; -x_55 = lean_ctor_get(x_49, 1); -lean_inc(x_55); -if (lean_is_exclusive(x_49)) { - lean_ctor_release(x_49, 0); - lean_ctor_release(x_49, 1); - x_56 = x_49; -} else { - lean_dec_ref(x_49); - x_56 = lean_box(0); -} -x_57 = lean_ctor_get(x_50, 0); -lean_inc(x_57); -if (lean_is_exclusive(x_50)) { - lean_ctor_release(x_50, 0); - x_58 = x_50; -} else { - lean_dec_ref(x_50); - x_58 = lean_box(0); -} -x_59 = lean_alloc_ctor(0, 8, 1); -lean_ctor_set(x_59, 0, x_57); -lean_ctor_set(x_59, 1, x_41); -lean_ctor_set(x_59, 2, x_42); -lean_ctor_set(x_59, 3, x_43); -lean_ctor_set(x_59, 4, x_44); -lean_ctor_set(x_59, 5, x_45); -lean_ctor_set(x_59, 6, x_47); -lean_ctor_set(x_59, 7, x_48); -lean_ctor_set_uint8(x_59, sizeof(void*)*8, x_46); -if (lean_is_scalar(x_58)) { - x_60 = lean_alloc_ctor(1, 1, 0); -} else { - x_60 = x_58; -} -lean_ctor_set(x_60, 0, x_59); -if (lean_is_scalar(x_56)) { - x_61 = lean_alloc_ctor(0, 2, 0); -} else { - x_61 = x_56; -} -lean_ctor_set(x_61, 0, x_60); -lean_ctor_set(x_61, 1, x_55); -return x_61; -} -} -else -{ -lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; -lean_dec(x_48); -lean_dec(x_47); -lean_dec(x_45); -lean_dec(x_44); -lean_dec(x_43); -lean_dec(x_42); -lean_dec(x_41); -x_62 = lean_ctor_get(x_49, 0); -lean_inc(x_62); -x_63 = lean_ctor_get(x_49, 1); -lean_inc(x_63); -if (lean_is_exclusive(x_49)) { - lean_ctor_release(x_49, 0); - lean_ctor_release(x_49, 1); - x_64 = x_49; -} else { - lean_dec_ref(x_49); - x_64 = lean_box(0); -} -if (lean_is_scalar(x_64)) { - x_65 = lean_alloc_ctor(1, 2, 0); -} else { - x_65 = x_64; -} -lean_ctor_set(x_65, 0, x_62); -lean_ctor_set(x_65, 1, x_63); -return x_65; -} -} -} -} -LEAN_EXPORT lean_object* l_List_forM___at_Lean_Meta_Grind_Preprocessor_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) { -_start: -{ -if (lean_obj_tag(x_1) == 0) -{ -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_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_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_inc(x_14); -lean_dec(x_1); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -lean_inc(x_2); -x_15 = l_Lean_Meta_Grind_Preprocessor_loop(x_13, 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; -x_16 = lean_ctor_get(x_15, 1); -lean_inc(x_16); -lean_dec(x_15); -x_1 = x_14; -x_10 = x_16; -goto _start; -} -else -{ -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_2); -x_18 = !lean_is_exclusive(x_15); -if (x_18 == 0) -{ -return x_15; -} -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_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_EXPORT lean_object* l_Lean_Meta_Grind_Preprocessor_loop___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; uint8_t x_11; -x_10 = lean_st_mk_ref(x_1, x_9); -x_11 = !lean_is_exclusive(x_10); -if (x_11 == 0) -{ -return x_10; -} -else -{ -lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_12 = lean_ctor_get(x_10, 0); -x_13 = lean_ctor_get(x_10, 1); -lean_inc(x_13); -lean_inc(x_12); -lean_dec(x_10); -x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_12); -lean_ctor_set(x_14, 1, x_13); -return x_14; -} -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Preprocessor_loop___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_11; lean_object* x_12; -x_11 = lean_unsigned_to_nat(0u); -lean_inc(x_2); -x_12 = l_Lean_Meta_Grind_addHyp(x_1, x_11, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_12) == 0) -{ -lean_object* x_13; lean_object* x_14; 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_2, x_13); -x_15 = !lean_is_exclusive(x_14); -if (x_15 == 0) -{ -lean_object* x_16; lean_object* x_17; uint8_t x_18; -x_16 = lean_ctor_get(x_14, 1); -x_17 = lean_st_ref_get(x_2, x_16); -lean_dec(x_2); -x_18 = !lean_is_exclusive(x_17); -if (x_18 == 0) -{ -lean_object* x_19; -x_19 = lean_ctor_get(x_17, 0); -lean_ctor_set(x_14, 1, x_19); -lean_ctor_set(x_17, 0, x_14); -return x_17; -} -else -{ -lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_20 = lean_ctor_get(x_17, 0); -x_21 = lean_ctor_get(x_17, 1); -lean_inc(x_21); -lean_inc(x_20); -lean_dec(x_17); -lean_ctor_set(x_14, 1, x_20); -x_22 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_22, 0, x_14); -lean_ctor_set(x_22, 1, x_21); -return x_22; -} -} -else -{ -lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_23 = lean_ctor_get(x_14, 0); -x_24 = lean_ctor_get(x_14, 1); -lean_inc(x_24); -lean_inc(x_23); -lean_dec(x_14); -x_25 = lean_st_ref_get(x_2, x_24); -lean_dec(x_2); -x_26 = lean_ctor_get(x_25, 0); -lean_inc(x_26); -x_27 = lean_ctor_get(x_25, 1); -lean_inc(x_27); -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); -} -x_29 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_29, 0, x_23); -lean_ctor_set(x_29, 1, x_26); -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; -} -} -else -{ -uint8_t x_31; -lean_dec(x_2); -x_31 = !lean_is_exclusive(x_12); -if (x_31 == 0) -{ -return x_12; -} -else -{ -lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_32 = lean_ctor_get(x_12, 0); -x_33 = lean_ctor_get(x_12, 1); -lean_inc(x_33); -lean_inc(x_32); -lean_dec(x_12); -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_EXPORT lean_object* l_Lean_Meta_Grind_Preprocessor_loop___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; -x_10 = !lean_is_exclusive(x_1); -if (x_10 == 0) -{ -lean_object* x_11; -x_11 = lean_ctor_get(x_1, 1); -lean_dec(x_11); -lean_ctor_set(x_1, 1, x_9); -return x_1; -} -else -{ -lean_object* x_12; lean_object* x_13; -x_12 = lean_ctor_get(x_1, 0); -lean_inc(x_12); -lean_dec(x_1); -x_13 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_13, 0, x_12); -lean_ctor_set(x_13, 1, x_9); -return x_13; -} -} -} -static lean_object* _init_l_Lean_Meta_Grind_Preprocessor_loop___lambda__4___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_Preprocessor_loop___lambda__3___boxed), 9, 0); -return x_1; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Preprocessor_loop___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, uint8_t x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19, lean_object* x_20) { -_start: -{ -lean_object* x_21; -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_1); -x_21 = l_Lean_Meta_Grind_Preprocessor_introNext(x_1, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, 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); -switch (lean_obj_tag(x_22)) { -case 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); -lean_inc(x_19); -lean_inc(x_18); -lean_inc(x_17); -lean_inc(x_16); -x_24 = l_Lean_MVarId_byContra_x3f(x_2, x_16, x_17, x_18, x_19, x_23); -if (lean_obj_tag(x_24) == 0) -{ -lean_object* x_25; -x_25 = lean_ctor_get(x_24, 0); -lean_inc(x_25); -if (lean_obj_tag(x_25) == 0) -{ -lean_object* x_26; lean_object* x_27; -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); -x_26 = lean_ctor_get(x_24, 1); -lean_inc(x_26); -lean_dec(x_24); -x_27 = l_Lean_Meta_Grind_Preprocessor_pushResult(x_1, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_26); -lean_dec(x_19); -lean_dec(x_18); -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -return x_27; -} -else -{ -lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; -lean_dec(x_1); -x_28 = lean_ctor_get(x_24, 1); -lean_inc(x_28); -lean_dec(x_24); -x_29 = lean_ctor_get(x_25, 0); -lean_inc(x_29); -lean_dec(x_25); -x_30 = lean_alloc_ctor(0, 8, 1); -lean_ctor_set(x_30, 0, x_29); -lean_ctor_set(x_30, 1, x_3); -lean_ctor_set(x_30, 2, x_4); -lean_ctor_set(x_30, 3, x_5); -lean_ctor_set(x_30, 4, x_6); -lean_ctor_set(x_30, 5, x_7); -lean_ctor_set(x_30, 6, x_9); -lean_ctor_set(x_30, 7, x_10); -lean_ctor_set_uint8(x_30, sizeof(void*)*8, x_8); -x_31 = l_Lean_Meta_Grind_Preprocessor_loop(x_30, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_28); -return x_31; -} -} -else -{ -uint8_t x_32; -lean_dec(x_19); -lean_dec(x_18); -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -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_1); -x_32 = !lean_is_exclusive(x_24); -if (x_32 == 0) -{ -return x_24; -} -else -{ -lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_33 = lean_ctor_get(x_24, 0); -x_34 = lean_ctor_get(x_24, 1); -lean_inc(x_34); -lean_inc(x_33); -lean_dec(x_24); -x_35 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -return x_35; -} -} -} -case 1: -{ -lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; -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); -lean_dec(x_1); -x_36 = lean_ctor_get(x_21, 1); -lean_inc(x_36); -lean_dec(x_21); -x_37 = lean_ctor_get(x_22, 0); -lean_inc(x_37); -x_38 = lean_ctor_get(x_22, 1); -lean_inc(x_38); -lean_dec(x_22); -lean_inc(x_19); -lean_inc(x_18); -lean_inc(x_17); -lean_inc(x_16); -lean_inc(x_37); -lean_inc(x_38); -x_39 = l_Lean_Meta_Grind_Preprocessor_applyCases_x3f(x_38, x_37, x_16, x_17, x_18, x_19, x_36); -if (lean_obj_tag(x_39) == 0) -{ -lean_object* x_40; -x_40 = lean_ctor_get(x_39, 0); -lean_inc(x_40); -if (lean_obj_tag(x_40) == 0) -{ -lean_object* x_41; lean_object* x_42; -x_41 = lean_ctor_get(x_39, 1); -lean_inc(x_41); -lean_dec(x_39); -lean_inc(x_19); -lean_inc(x_18); -lean_inc(x_17); -lean_inc(x_16); -lean_inc(x_37); -lean_inc(x_38); -x_42 = l_Lean_Meta_Grind_Preprocessor_applyInjection_x3f(x_38, x_37, x_16, x_17, x_18, x_19, x_41); -if (lean_obj_tag(x_42) == 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_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; -x_44 = lean_ctor_get(x_42, 1); -lean_inc(x_44); -lean_dec(x_42); -x_45 = lean_ctor_get(x_38, 0); -lean_inc(x_45); -x_46 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_Preprocessor_loop___lambda__1___boxed), 9, 1); -lean_closure_set(x_46, 0, x_38); -x_47 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_Preprocessor_loop___lambda__2), 10, 1); -lean_closure_set(x_47, 0, x_37); -x_48 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); -lean_closure_set(x_48, 0, x_46); -lean_closure_set(x_48, 1, x_47); -x_49 = l_Lean_Meta_Grind_Preprocessor_loop___lambda__4___closed__1; -x_50 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); -lean_closure_set(x_50, 0, x_48); -lean_closure_set(x_50, 1, x_49); -lean_inc(x_19); -lean_inc(x_18); -lean_inc(x_17); -lean_inc(x_16); -lean_inc(x_15); -lean_inc(x_14); -lean_inc(x_13); -x_51 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(x_45, x_50, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_44); -if (lean_obj_tag(x_51) == 0) -{ -lean_object* x_52; lean_object* x_53; lean_object* x_54; -x_52 = lean_ctor_get(x_51, 0); -lean_inc(x_52); -x_53 = lean_ctor_get(x_51, 1); -lean_inc(x_53); -lean_dec(x_51); -x_54 = l_Lean_Meta_Grind_Preprocessor_loop(x_52, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_53); -return x_54; -} -else -{ -uint8_t x_55; -lean_dec(x_19); -lean_dec(x_18); -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -x_55 = !lean_is_exclusive(x_51); -if (x_55 == 0) -{ -return x_51; -} -else -{ -lean_object* x_56; lean_object* x_57; lean_object* x_58; -x_56 = lean_ctor_get(x_51, 0); -x_57 = lean_ctor_get(x_51, 1); -lean_inc(x_57); -lean_inc(x_56); -lean_dec(x_51); -x_58 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_58, 0, x_56); -lean_ctor_set(x_58, 1, x_57); -return x_58; -} -} -} -else -{ -lean_object* x_59; lean_object* x_60; lean_object* x_61; -lean_dec(x_38); -lean_dec(x_37); -x_59 = lean_ctor_get(x_42, 1); -lean_inc(x_59); -lean_dec(x_42); -x_60 = lean_ctor_get(x_43, 0); -lean_inc(x_60); -lean_dec(x_43); -x_61 = l_Lean_Meta_Grind_Preprocessor_loop(x_60, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_59); -return x_61; -} -} -else -{ -uint8_t x_62; -lean_dec(x_38); -lean_dec(x_37); -lean_dec(x_19); -lean_dec(x_18); -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -x_62 = !lean_is_exclusive(x_42); -if (x_62 == 0) -{ -return x_42; -} -else -{ -lean_object* x_63; lean_object* x_64; lean_object* x_65; -x_63 = lean_ctor_get(x_42, 0); -x_64 = lean_ctor_get(x_42, 1); -lean_inc(x_64); -lean_inc(x_63); -lean_dec(x_42); -x_65 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_65, 0, x_63); -lean_ctor_set(x_65, 1, x_64); -return x_65; -} -} -} -else -{ -lean_object* x_66; lean_object* x_67; lean_object* x_68; -lean_dec(x_38); -lean_dec(x_37); -x_66 = lean_ctor_get(x_39, 1); -lean_inc(x_66); -lean_dec(x_39); -x_67 = lean_ctor_get(x_40, 0); -lean_inc(x_67); -lean_dec(x_40); -x_68 = l_List_forM___at_Lean_Meta_Grind_Preprocessor_loop___spec__1(x_67, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_66); -return x_68; -} -} -else -{ -uint8_t x_69; -lean_dec(x_38); -lean_dec(x_37); -lean_dec(x_19); -lean_dec(x_18); -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -x_69 = !lean_is_exclusive(x_39); -if (x_69 == 0) -{ -return x_39; -} -else -{ -lean_object* x_70; lean_object* x_71; lean_object* x_72; -x_70 = lean_ctor_get(x_39, 0); -x_71 = lean_ctor_get(x_39, 1); -lean_inc(x_71); -lean_inc(x_70); -lean_dec(x_39); -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; -} -} -} -case 2: -{ -lean_object* x_73; lean_object* x_74; lean_object* x_75; -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); -lean_dec(x_1); -x_73 = lean_ctor_get(x_21, 1); -lean_inc(x_73); -lean_dec(x_21); -x_74 = lean_ctor_get(x_22, 0); -lean_inc(x_74); -lean_dec(x_22); -x_75 = l_Lean_Meta_Grind_Preprocessor_loop(x_74, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_73); -return x_75; -} -default: -{ -lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; -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); -lean_dec(x_1); -x_76 = lean_ctor_get(x_21, 1); -lean_inc(x_76); -lean_dec(x_21); -x_77 = lean_ctor_get(x_22, 0); -lean_inc(x_77); -x_78 = lean_ctor_get(x_22, 1); -lean_inc(x_78); -lean_dec(x_22); -lean_inc(x_19); -lean_inc(x_18); -lean_inc(x_17); -lean_inc(x_16); -lean_inc(x_78); -x_79 = l_Lean_Meta_Grind_Preprocessor_applyCases_x3f(x_78, x_77, x_16, x_17, x_18, x_19, x_76); -if (lean_obj_tag(x_79) == 0) -{ -lean_object* x_80; -x_80 = lean_ctor_get(x_79, 0); -lean_inc(x_80); -if (lean_obj_tag(x_80) == 0) -{ -lean_object* x_81; lean_object* x_82; -x_81 = lean_ctor_get(x_79, 1); -lean_inc(x_81); -lean_dec(x_79); -x_82 = l_Lean_Meta_Grind_Preprocessor_loop(x_78, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_81); -return x_82; -} -else -{ -lean_object* x_83; lean_object* x_84; lean_object* x_85; -lean_dec(x_78); -x_83 = lean_ctor_get(x_79, 1); -lean_inc(x_83); -lean_dec(x_79); -x_84 = lean_ctor_get(x_80, 0); -lean_inc(x_84); -lean_dec(x_80); -x_85 = l_List_forM___at_Lean_Meta_Grind_Preprocessor_loop___spec__1(x_84, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_83); -return x_85; -} -} -else -{ -uint8_t x_86; -lean_dec(x_78); -lean_dec(x_19); -lean_dec(x_18); -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -x_86 = !lean_is_exclusive(x_79); -if (x_86 == 0) -{ -return x_79; -} -else -{ -lean_object* x_87; lean_object* x_88; lean_object* x_89; -x_87 = lean_ctor_get(x_79, 0); -x_88 = lean_ctor_get(x_79, 1); -lean_inc(x_88); -lean_inc(x_87); -lean_dec(x_79); -x_89 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_89, 0, x_87); -lean_ctor_set(x_89, 1, x_88); -return x_89; -} -} -} -} -} -else -{ -uint8_t x_90; -lean_dec(x_19); -lean_dec(x_18); -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -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); -lean_dec(x_1); -x_90 = !lean_is_exclusive(x_21); -if (x_90 == 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_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; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Preprocessor_loop(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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_ctor_get_uint8(x_1, sizeof(void*)*8); -if (x_11 == 0) -{ -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_12 = lean_ctor_get(x_1, 0); -lean_inc(x_12); -x_13 = lean_ctor_get(x_1, 1); -lean_inc(x_13); -x_14 = lean_ctor_get(x_1, 2); -lean_inc(x_14); -x_15 = lean_ctor_get(x_1, 3); -lean_inc(x_15); -x_16 = lean_ctor_get(x_1, 4); -lean_inc(x_16); -x_17 = lean_ctor_get(x_1, 5); -lean_inc(x_17); -x_18 = lean_ctor_get(x_1, 6); -lean_inc(x_18); -x_19 = lean_ctor_get(x_1, 7); -lean_inc(x_19); -x_20 = lean_box(0); -x_21 = l_Lean_Meta_Grind_Preprocessor_loop___lambda__4(x_1, x_12, x_13, x_14, x_15, x_16, x_17, x_11, x_18, x_19, x_20, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_21; -} -else -{ -lean_object* x_22; lean_object* x_23; -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_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_10); -return x_23; -} -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Preprocessor_loop___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -lean_object* x_10; -x_10 = l_Lean_Meta_Grind_Preprocessor_loop___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_10; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Preprocessor_loop___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -lean_object* x_10; -x_10 = l_Lean_Meta_Grind_Preprocessor_loop___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_10; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Preprocessor_loop___lambda__4___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_object* x_19 = _args[18]; -lean_object* x_20 = _args[19]; -_start: -{ -uint8_t x_21; lean_object* x_22; -x_21 = lean_unbox(x_8); -lean_dec(x_8); -x_22 = l_Lean_Meta_Grind_Preprocessor_loop___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_21, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20); -lean_dec(x_11); -return x_22; -} -} -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_Preprocessor_ppGoals___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, size_t x_6, size_t x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17) { -_start: -{ -uint8_t x_18; -x_18 = lean_usize_dec_lt(x_7, x_6); -if (x_18 == 0) -{ -lean_object* x_19; -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_4); -x_19 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_19, 0, x_8); -lean_ctor_set(x_19, 1, x_17); -return x_19; -} -else -{ -lean_object* x_20; uint8_t x_21; -x_20 = lean_array_uget(x_5, x_7); -x_21 = !lean_is_exclusive(x_8); -if (x_21 == 0) -{ -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_8, 1); -x_23 = lean_ctor_get(x_8, 0); -lean_dec(x_23); -lean_inc(x_16); -lean_inc(x_15); -lean_inc(x_14); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_22); -x_24 = l_Lean_PersistentArray_forInAux___at_Lean_Meta_Grind_Preprocessor_ppGoals___spec__2(x_1, x_20, x_22, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); -lean_dec(x_20); -if (lean_obj_tag(x_24) == 0) -{ -lean_object* x_25; -x_25 = lean_ctor_get(x_24, 0); -lean_inc(x_25); -if (lean_obj_tag(x_25) == 0) -{ -uint8_t x_26; -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_4); -x_26 = !lean_is_exclusive(x_24); -if (x_26 == 0) -{ -lean_object* x_27; lean_object* x_28; -x_27 = lean_ctor_get(x_24, 0); -lean_dec(x_27); -x_28 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_28, 0, x_25); -lean_ctor_set(x_8, 0, x_28); -lean_ctor_set(x_24, 0, x_8); -return x_24; -} -else -{ -lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_29 = lean_ctor_get(x_24, 1); -lean_inc(x_29); -lean_dec(x_24); -x_30 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_30, 0, x_25); -lean_ctor_set(x_8, 0, x_30); -x_31 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_31, 0, x_8); -lean_ctor_set(x_31, 1, x_29); -return x_31; -} -} -else -{ -lean_object* x_32; lean_object* x_33; size_t x_34; size_t x_35; -lean_dec(x_22); -x_32 = lean_ctor_get(x_24, 1); -lean_inc(x_32); -lean_dec(x_24); -x_33 = lean_ctor_get(x_25, 0); -lean_inc(x_33); -lean_dec(x_25); -lean_inc(x_4); -lean_ctor_set(x_8, 1, x_33); -lean_ctor_set(x_8, 0, x_4); -x_34 = 1; -x_35 = lean_usize_add(x_7, x_34); -x_7 = x_35; -x_17 = x_32; -goto _start; -} -} -else -{ -uint8_t x_37; -lean_free_object(x_8); -lean_dec(x_22); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_4); -x_37 = !lean_is_exclusive(x_24); -if (x_37 == 0) -{ -return x_24; -} -else -{ -lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_38 = lean_ctor_get(x_24, 0); -x_39 = lean_ctor_get(x_24, 1); -lean_inc(x_39); -lean_inc(x_38); -lean_dec(x_24); -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; -x_41 = lean_ctor_get(x_8, 1); -lean_inc(x_41); -lean_dec(x_8); -lean_inc(x_16); -lean_inc(x_15); -lean_inc(x_14); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_41); -x_42 = l_Lean_PersistentArray_forInAux___at_Lean_Meta_Grind_Preprocessor_ppGoals___spec__2(x_1, x_20, x_41, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); -lean_dec(x_20); -if (lean_obj_tag(x_42) == 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_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_4); -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, 1, 0); -lean_ctor_set(x_46, 0, x_43); -x_47 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_47, 0, x_46); -lean_ctor_set(x_47, 1, x_41); -if (lean_is_scalar(x_45)) { - x_48 = lean_alloc_ctor(0, 2, 0); -} else { - x_48 = x_45; -} -lean_ctor_set(x_48, 0, x_47); -lean_ctor_set(x_48, 1, x_44); -return x_48; -} -else -{ -lean_object* x_49; lean_object* x_50; lean_object* x_51; size_t x_52; size_t x_53; -lean_dec(x_41); -x_49 = lean_ctor_get(x_42, 1); -lean_inc(x_49); -lean_dec(x_42); -x_50 = lean_ctor_get(x_43, 0); -lean_inc(x_50); -lean_dec(x_43); -lean_inc(x_4); -x_51 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_51, 0, x_4); -lean_ctor_set(x_51, 1, x_50); -x_52 = 1; -x_53 = lean_usize_add(x_7, x_52); -x_7 = x_53; -x_8 = x_51; -x_17 = x_49; -goto _start; -} -} -else -{ -lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; -lean_dec(x_41); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_4); -x_55 = lean_ctor_get(x_42, 0); -lean_inc(x_55); -x_56 = lean_ctor_get(x_42, 1); -lean_inc(x_56); -if (lean_is_exclusive(x_42)) { - lean_ctor_release(x_42, 0); - lean_ctor_release(x_42, 1); - x_57 = x_42; -} else { - lean_dec_ref(x_42); - x_57 = lean_box(0); -} -if (lean_is_scalar(x_57)) { - x_58 = lean_alloc_ctor(1, 2, 0); -} else { - x_58 = x_57; -} -lean_ctor_set(x_58, 0, x_55); -lean_ctor_set(x_58, 1, x_56); -return x_58; -} -} -} -} -} -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_Preprocessor_ppGoals___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_inc(x_1); -x_10 = l_Lean_Meta_Grind_ppState(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_10) == 0) -{ -lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_11 = lean_ctor_get(x_10, 0); -lean_inc(x_11); -x_12 = lean_ctor_get(x_10, 1); -lean_inc(x_12); -lean_dec(x_10); -x_13 = lean_st_ref_get(x_1, x_12); -lean_dec(x_1); -x_14 = !lean_is_exclusive(x_13); -if (x_14 == 0) -{ -lean_object* x_15; lean_object* x_16; -x_15 = lean_ctor_get(x_13, 0); -x_16 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_16, 0, x_11); -lean_ctor_set(x_16, 1, x_15); -lean_ctor_set(x_13, 0, x_16); -return x_13; -} -else -{ -lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_17 = lean_ctor_get(x_13, 0); -x_18 = lean_ctor_get(x_13, 1); -lean_inc(x_18); -lean_inc(x_17); -lean_dec(x_13); -x_19 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_19, 0, x_11); -lean_ctor_set(x_19, 1, x_17); -x_20 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_18); -return x_20; -} -} -else -{ -uint8_t x_21; -lean_dec(x_1); -x_21 = !lean_is_exclusive(x_10); -if (x_21 == 0) -{ -return x_10; -} -else -{ -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_10, 0); -x_23 = lean_ctor_get(x_10, 1); -lean_inc(x_23); -lean_inc(x_22); -lean_dec(x_10); -x_24 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_24, 0, x_22); -lean_ctor_set(x_24, 1, x_23); -return x_24; -} -} -} -} -static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_Preprocessor_ppGoals___spec__4___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_Preprocessor_ppGoals___spec__4___lambda__1), 9, 0); -return x_1; -} -} -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_Preprocessor_ppGoals___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, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { -_start: -{ -uint8_t x_17; -x_17 = lean_usize_dec_lt(x_6, x_5); -if (x_17 == 0) -{ -lean_object* x_18; -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_3); -x_18 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_18, 0, x_7); -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; -x_19 = lean_array_uget(x_4, x_6); -x_20 = lean_ctor_get(x_7, 1); -lean_inc(x_20); -lean_dec(x_7); -x_21 = lean_ctor_get(x_19, 0); -lean_inc(x_21); -x_22 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_Preprocessor_loop___lambda__1___boxed), 9, 1); -lean_closure_set(x_22, 0, x_19); -x_23 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_Preprocessor_ppGoals___spec__4___closed__1; -x_24 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); -lean_closure_set(x_24, 0, x_22); -lean_closure_set(x_24, 1, x_23); -lean_inc(x_15); -lean_inc(x_14); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -x_25 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(x_21, x_24, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); -if (lean_obj_tag(x_25) == 0) -{ -lean_object* x_26; lean_object* x_27; uint8_t x_28; -x_26 = lean_ctor_get(x_25, 0); -lean_inc(x_26); -x_27 = lean_ctor_get(x_25, 1); -lean_inc(x_27); -lean_dec(x_25); -x_28 = !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; size_t x_34; size_t x_35; -x_29 = lean_ctor_get(x_26, 0); -x_30 = lean_ctor_get(x_26, 1); -lean_dec(x_30); -x_31 = lean_box(1); -lean_ctor_set_tag(x_26, 5); -lean_ctor_set(x_26, 1, x_31); -lean_ctor_set(x_26, 0, x_20); -x_32 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_32, 0, x_26); -lean_ctor_set(x_32, 1, x_29); -lean_inc(x_3); -x_33 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_33, 0, x_3); -lean_ctor_set(x_33, 1, x_32); -x_34 = 1; -x_35 = lean_usize_add(x_6, x_34); -x_6 = x_35; -x_7 = x_33; -x_16 = x_27; -goto _start; -} -else -{ -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_37 = lean_ctor_get(x_26, 0); -lean_inc(x_37); -lean_dec(x_26); -x_38 = lean_box(1); -x_39 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_39, 0, x_20); -lean_ctor_set(x_39, 1, x_38); -x_40 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_40, 0, x_39); -lean_ctor_set(x_40, 1, x_37); -lean_inc(x_3); -x_41 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_41, 0, x_3); -lean_ctor_set(x_41, 1, x_40); -x_42 = 1; -x_43 = lean_usize_add(x_6, x_42); -x_6 = x_43; -x_7 = x_41; -x_16 = x_27; -goto _start; -} -} -else -{ -uint8_t x_45; -lean_dec(x_20); -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_3); -x_45 = !lean_is_exclusive(x_25); -if (x_45 == 0) -{ -return x_25; -} -else -{ -lean_object* x_46; lean_object* x_47; lean_object* x_48; -x_46 = lean_ctor_get(x_25, 0); -x_47 = lean_ctor_get(x_25, 1); -lean_inc(x_47); -lean_inc(x_46); -lean_dec(x_25); -x_48 = 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_PersistentArray_forInAux___at_Lean_Meta_Grind_Preprocessor_ppGoals___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; -x_12 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_12, 0, x_1); -x_13 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_13, 0, x_12); -lean_ctor_set(x_13, 1, x_11); -return x_13; -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at_Lean_Meta_Grind_Preprocessor_ppGoals___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -if (lean_obj_tag(x_2) == 0) -{ -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; size_t x_17; size_t x_18; lean_object* x_19; -x_13 = lean_ctor_get(x_2, 0); -x_14 = lean_box(0); -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_3); -x_17 = lean_array_size(x_13); -x_18 = 0; -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -x_19 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_Preprocessor_ppGoals___spec__3(x_1, x_13, x_14, x_15, x_13, x_17, x_18, x_16, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_19) == 0) -{ -lean_object* x_20; lean_object* x_21; -x_20 = lean_ctor_get(x_19, 0); -lean_inc(x_20); -x_21 = lean_ctor_get(x_20, 0); -lean_inc(x_21); -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_19, 1); -lean_inc(x_22); -lean_dec(x_19); -x_23 = lean_ctor_get(x_20, 1); -lean_inc(x_23); -lean_dec(x_20); -x_24 = lean_box(0); -x_25 = l_Lean_PersistentArray_forInAux___at_Lean_Meta_Grind_Preprocessor_ppGoals___spec__2___lambda__1(x_23, x_24, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_22); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -return x_25; -} -else -{ -uint8_t x_26; -lean_dec(x_20); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -x_26 = !lean_is_exclusive(x_19); -if (x_26 == 0) -{ -lean_object* x_27; lean_object* x_28; -x_27 = lean_ctor_get(x_19, 0); -lean_dec(x_27); -x_28 = lean_ctor_get(x_21, 0); -lean_inc(x_28); -lean_dec(x_21); -lean_ctor_set(x_19, 0, x_28); -return x_19; -} -else -{ -lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_29 = lean_ctor_get(x_19, 1); -lean_inc(x_29); -lean_dec(x_19); -x_30 = lean_ctor_get(x_21, 0); -lean_inc(x_30); -lean_dec(x_21); -x_31 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_31, 0, x_30); -lean_ctor_set(x_31, 1, x_29); -return x_31; -} -} -} -else -{ -uint8_t x_32; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -x_32 = !lean_is_exclusive(x_19); -if (x_32 == 0) -{ -return x_19; -} -else -{ -lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_33 = lean_ctor_get(x_19, 0); -x_34 = lean_ctor_get(x_19, 1); -lean_inc(x_34); -lean_inc(x_33); -lean_dec(x_19); -x_35 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -return x_35; -} -} -} -else -{ -lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; size_t x_40; size_t x_41; lean_object* x_42; -x_36 = lean_ctor_get(x_2, 0); -x_37 = lean_box(0); -x_38 = lean_box(0); -x_39 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_39, 0, x_38); -lean_ctor_set(x_39, 1, x_3); -x_40 = lean_array_size(x_36); -x_41 = 0; -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -x_42 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_Preprocessor_ppGoals___spec__4(x_36, x_37, x_38, x_36, x_40, x_41, x_39, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_42) == 0) -{ -lean_object* x_43; lean_object* x_44; -x_43 = lean_ctor_get(x_42, 0); -lean_inc(x_43); -x_44 = lean_ctor_get(x_43, 0); -lean_inc(x_44); -if (lean_obj_tag(x_44) == 0) -{ -lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; -x_45 = lean_ctor_get(x_42, 1); -lean_inc(x_45); -lean_dec(x_42); -x_46 = lean_ctor_get(x_43, 1); -lean_inc(x_46); -lean_dec(x_43); -x_47 = lean_box(0); -x_48 = l_Lean_PersistentArray_forInAux___at_Lean_Meta_Grind_Preprocessor_ppGoals___spec__2___lambda__1(x_46, x_47, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, 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); -return x_48; -} -else -{ -uint8_t x_49; -lean_dec(x_43); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -x_49 = !lean_is_exclusive(x_42); -if (x_49 == 0) -{ -lean_object* x_50; lean_object* x_51; -x_50 = lean_ctor_get(x_42, 0); -lean_dec(x_50); -x_51 = lean_ctor_get(x_44, 0); -lean_inc(x_51); -lean_dec(x_44); -lean_ctor_set(x_42, 0, x_51); -return x_42; -} -else -{ -lean_object* x_52; lean_object* x_53; lean_object* x_54; -x_52 = lean_ctor_get(x_42, 1); -lean_inc(x_52); -lean_dec(x_42); -x_53 = lean_ctor_get(x_44, 0); -lean_inc(x_53); -lean_dec(x_44); -x_54 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_54, 0, x_53); -lean_ctor_set(x_54, 1, x_52); -return x_54; -} -} -} -else -{ -uint8_t x_55; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -x_55 = !lean_is_exclusive(x_42); -if (x_55 == 0) -{ -return x_42; -} -else -{ -lean_object* x_56; lean_object* x_57; lean_object* x_58; -x_56 = lean_ctor_get(x_42, 0); -x_57 = lean_ctor_get(x_42, 1); -lean_inc(x_57); -lean_inc(x_56); -lean_dec(x_42); -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_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_Preprocessor_ppGoals___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, size_t x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { -_start: -{ -uint8_t x_17; -x_17 = lean_usize_dec_lt(x_6, x_5); -if (x_17 == 0) -{ -lean_object* x_18; -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_3); -x_18 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_18, 0, x_7); -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; -x_19 = lean_array_uget(x_4, x_6); -x_20 = lean_ctor_get(x_7, 1); -lean_inc(x_20); -lean_dec(x_7); -x_21 = lean_ctor_get(x_19, 0); -lean_inc(x_21); -x_22 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_Preprocessor_loop___lambda__1___boxed), 9, 1); -lean_closure_set(x_22, 0, x_19); -x_23 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_Preprocessor_ppGoals___spec__4___closed__1; -x_24 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); -lean_closure_set(x_24, 0, x_22); -lean_closure_set(x_24, 1, x_23); -lean_inc(x_15); -lean_inc(x_14); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -x_25 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(x_21, x_24, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); -if (lean_obj_tag(x_25) == 0) -{ -lean_object* x_26; lean_object* x_27; uint8_t x_28; -x_26 = lean_ctor_get(x_25, 0); -lean_inc(x_26); -x_27 = lean_ctor_get(x_25, 1); -lean_inc(x_27); -lean_dec(x_25); -x_28 = !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; size_t x_34; size_t x_35; -x_29 = lean_ctor_get(x_26, 0); -x_30 = lean_ctor_get(x_26, 1); -lean_dec(x_30); -x_31 = lean_box(1); -lean_ctor_set_tag(x_26, 5); -lean_ctor_set(x_26, 1, x_31); -lean_ctor_set(x_26, 0, x_20); -x_32 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_32, 0, x_26); -lean_ctor_set(x_32, 1, x_29); -lean_inc(x_3); -x_33 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_33, 0, x_3); -lean_ctor_set(x_33, 1, x_32); -x_34 = 1; -x_35 = lean_usize_add(x_6, x_34); -x_6 = x_35; -x_7 = x_33; -x_16 = x_27; -goto _start; -} -else -{ -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_37 = lean_ctor_get(x_26, 0); -lean_inc(x_37); -lean_dec(x_26); -x_38 = lean_box(1); -x_39 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_39, 0, x_20); -lean_ctor_set(x_39, 1, x_38); -x_40 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_40, 0, x_39); -lean_ctor_set(x_40, 1, x_37); -lean_inc(x_3); -x_41 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_41, 0, x_3); -lean_ctor_set(x_41, 1, x_40); -x_42 = 1; -x_43 = lean_usize_add(x_6, x_42); -x_6 = x_43; -x_7 = x_41; -x_16 = x_27; -goto _start; -} -} -else -{ -uint8_t x_45; -lean_dec(x_20); -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_3); -x_45 = !lean_is_exclusive(x_25); -if (x_45 == 0) -{ -return x_25; -} -else -{ -lean_object* x_46; lean_object* x_47; lean_object* x_48; -x_46 = lean_ctor_get(x_25, 0); -x_47 = lean_ctor_get(x_25, 1); -lean_inc(x_47); -lean_inc(x_46); -lean_dec(x_25); -x_48 = 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_PersistentArray_forIn___at_Lean_Meta_Grind_Preprocessor_ppGoals___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; -x_12 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_12, 0, x_1); -lean_ctor_set(x_12, 1, x_11); -return x_12; -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at_Lean_Meta_Grind_Preprocessor_ppGoals___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) { -_start: -{ -lean_object* x_12; lean_object* x_13; -x_12 = lean_ctor_get(x_1, 0); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_2); -x_13 = l_Lean_PersistentArray_forInAux___at_Lean_Meta_Grind_Preprocessor_ppGoals___spec__2(x_2, x_12, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -lean_dec(x_2); -if (lean_obj_tag(x_13) == 0) -{ -lean_object* x_14; -x_14 = lean_ctor_get(x_13, 0); -lean_inc(x_14); -if (lean_obj_tag(x_14) == 0) -{ -uint8_t x_15; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -x_15 = !lean_is_exclusive(x_13); -if (x_15 == 0) -{ -lean_object* x_16; lean_object* x_17; -x_16 = lean_ctor_get(x_13, 0); -lean_dec(x_16); -x_17 = lean_ctor_get(x_14, 0); -lean_inc(x_17); -lean_dec(x_14); -lean_ctor_set(x_13, 0, x_17); -return x_13; -} -else -{ -lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_18 = lean_ctor_get(x_13, 1); -lean_inc(x_18); -lean_dec(x_13); -x_19 = lean_ctor_get(x_14, 0); -lean_inc(x_19); -lean_dec(x_14); -x_20 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_18); -return x_20; -} -} -else -{ -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; size_t x_27; size_t x_28; lean_object* x_29; -x_21 = lean_ctor_get(x_13, 1); -lean_inc(x_21); -lean_dec(x_13); -x_22 = lean_ctor_get(x_14, 0); -lean_inc(x_22); -lean_dec(x_14); -x_23 = lean_ctor_get(x_1, 1); -x_24 = lean_box(0); -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_22); -x_27 = lean_array_size(x_23); -x_28 = 0; -x_29 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_Preprocessor_ppGoals___spec__5(x_23, x_24, x_25, x_23, x_27, x_28, x_26, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_21); -if (lean_obj_tag(x_29) == 0) -{ -lean_object* x_30; lean_object* x_31; -x_30 = lean_ctor_get(x_29, 0); -lean_inc(x_30); -x_31 = lean_ctor_get(x_30, 0); -lean_inc(x_31); -if (lean_obj_tag(x_31) == 0) -{ -uint8_t x_32; -x_32 = !lean_is_exclusive(x_29); -if (x_32 == 0) -{ -lean_object* x_33; lean_object* x_34; -x_33 = lean_ctor_get(x_29, 0); -lean_dec(x_33); -x_34 = lean_ctor_get(x_30, 1); -lean_inc(x_34); -lean_dec(x_30); -lean_ctor_set(x_29, 0, x_34); -return x_29; -} -else -{ -lean_object* x_35; lean_object* x_36; lean_object* x_37; -x_35 = lean_ctor_get(x_29, 1); -lean_inc(x_35); -lean_dec(x_29); -x_36 = lean_ctor_get(x_30, 1); -lean_inc(x_36); -lean_dec(x_30); -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; -lean_dec(x_30); -x_38 = !lean_is_exclusive(x_29); -if (x_38 == 0) -{ -lean_object* x_39; lean_object* x_40; -x_39 = lean_ctor_get(x_29, 0); -lean_dec(x_39); -x_40 = lean_ctor_get(x_31, 0); -lean_inc(x_40); -lean_dec(x_31); -lean_ctor_set(x_29, 0, x_40); -return x_29; -} -else -{ -lean_object* x_41; lean_object* x_42; lean_object* x_43; -x_41 = lean_ctor_get(x_29, 1); -lean_inc(x_41); -lean_dec(x_29); -x_42 = lean_ctor_get(x_31, 0); -lean_inc(x_42); -lean_dec(x_31); -x_43 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_43, 0, x_42); -lean_ctor_set(x_43, 1, x_41); -return x_43; -} -} -} -else -{ -uint8_t x_44; -x_44 = !lean_is_exclusive(x_29); -if (x_44 == 0) -{ -return x_29; -} -else -{ -lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_45 = lean_ctor_get(x_29, 0); -x_46 = lean_ctor_get(x_29, 1); -lean_inc(x_46); -lean_inc(x_45); -lean_dec(x_29); -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_48; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -x_48 = !lean_is_exclusive(x_13); -if (x_48 == 0) -{ -return x_13; -} -else -{ -lean_object* x_49; lean_object* x_50; lean_object* x_51; -x_49 = lean_ctor_get(x_13, 0); -x_50 = lean_ctor_get(x_13, 1); -lean_inc(x_50); -lean_inc(x_49); -lean_dec(x_13); -x_51 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_51, 0, x_49); -lean_ctor_set(x_51, 1, x_50); -return x_51; -} -} -} -} -static lean_object* _init_l_Lean_Meta_Grind_Preprocessor_ppGoals___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("", 0, 0); -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_Grind_Preprocessor_ppGoals___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_Grind_Preprocessor_ppGoals___closed__1; -x_2 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Preprocessor_ppGoals(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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 = lean_st_ref_get(x_1, 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_Meta_Grind_Preprocessor_ppGoals___closed__2; -x_14 = l_Lean_PersistentArray_forIn___at_Lean_Meta_Grind_Preprocessor_ppGoals___spec__1(x_11, x_13, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_12); -lean_dec(x_11); -if (lean_obj_tag(x_14) == 0) -{ -uint8_t x_15; -x_15 = !lean_is_exclusive(x_14); -if (x_15 == 0) -{ -return x_14; -} -else -{ -lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_16 = lean_ctor_get(x_14, 0); -x_17 = lean_ctor_get(x_14, 1); -lean_inc(x_17); -lean_inc(x_16); -lean_dec(x_14); -x_18 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_18, 0, x_16); -lean_ctor_set(x_18, 1, x_17); -return x_18; -} -} -else -{ -uint8_t x_19; -x_19 = !lean_is_exclusive(x_14); -if (x_19 == 0) -{ -return x_14; -} -else -{ -lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_20 = lean_ctor_get(x_14, 0); -x_21 = lean_ctor_get(x_14, 1); -lean_inc(x_21); -lean_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_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_Preprocessor_ppGoals___spec__3___boxed(lean_object** _args) { -lean_object* x_1 = _args[0]; -lean_object* x_2 = _args[1]; -lean_object* x_3 = _args[2]; -lean_object* x_4 = _args[3]; -lean_object* x_5 = _args[4]; -lean_object* x_6 = _args[5]; -lean_object* x_7 = _args[6]; -lean_object* x_8 = _args[7]; -lean_object* x_9 = _args[8]; -lean_object* x_10 = _args[9]; -lean_object* x_11 = _args[10]; -lean_object* x_12 = _args[11]; -lean_object* x_13 = _args[12]; -lean_object* x_14 = _args[13]; -lean_object* x_15 = _args[14]; -lean_object* x_16 = _args[15]; -lean_object* x_17 = _args[16]; -_start: -{ -size_t x_18; size_t x_19; lean_object* x_20; -x_18 = lean_unbox_usize(x_6); -lean_dec(x_6); -x_19 = lean_unbox_usize(x_7); -lean_dec(x_7); -x_20 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_Preprocessor_ppGoals___spec__3(x_1, x_2, x_3, x_4, x_5, x_18, x_19, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); -lean_dec(x_9); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_20; -} -} -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_Preprocessor_ppGoals___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { -_start: -{ -size_t x_17; size_t x_18; lean_object* x_19; -x_17 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_18 = lean_unbox_usize(x_6); -lean_dec(x_6); -x_19 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_Preprocessor_ppGoals___spec__4(x_1, x_2, x_3, x_4, x_17, x_18, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); -lean_dec(x_8); -lean_dec(x_4); -lean_dec(x_2); -lean_dec(x_1); -return x_19; -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at_Lean_Meta_Grind_Preprocessor_ppGoals___spec__2___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -lean_object* x_12; -x_12 = l_Lean_PersistentArray_forInAux___at_Lean_Meta_Grind_Preprocessor_ppGoals___spec__2___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_12; -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at_Lean_Meta_Grind_Preprocessor_ppGoals___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -lean_object* x_13; -x_13 = l_Lean_PersistentArray_forInAux___at_Lean_Meta_Grind_Preprocessor_ppGoals___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_4); -lean_dec(x_2); -lean_dec(x_1); -return x_13; -} -} -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_Preprocessor_ppGoals___spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { -_start: -{ -size_t x_17; size_t x_18; lean_object* x_19; -x_17 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_18 = lean_unbox_usize(x_6); -lean_dec(x_6); -x_19 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_Preprocessor_ppGoals___spec__5(x_1, x_2, x_3, x_4, x_17, x_18, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); -lean_dec(x_8); -lean_dec(x_4); -lean_dec(x_2); -lean_dec(x_1); -return x_19; -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at_Lean_Meta_Grind_Preprocessor_ppGoals___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_PersistentArray_forIn___at_Lean_Meta_Grind_Preprocessor_ppGoals___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_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_12; -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at_Lean_Meta_Grind_Preprocessor_ppGoals___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -lean_object* x_12; -x_12 = l_Lean_PersistentArray_forIn___at_Lean_Meta_Grind_Preprocessor_ppGoals___spec__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_3); -lean_dec(x_1); -return x_12; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Preprocessor_ppGoals___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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_Grind_Preprocessor_ppGoals(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_isTracingEnabledFor___at_Lean_Meta_Grind_Preprocessor_preprocess___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_11; lean_object* x_12; uint8_t x_13; -x_11 = lean_ctor_get(x_8, 2); -x_12 = l_Lean_isTracingEnabledForCore(x_1, x_11, x_10); -x_13 = !lean_is_exclusive(x_12); -if (x_13 == 0) -{ -return x_12; -} -else -{ -lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_14 = lean_ctor_get(x_12, 0); -x_15 = lean_ctor_get(x_12, 1); -lean_inc(x_15); -lean_inc(x_14); -lean_dec(x_12); -x_16 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_16, 0, x_14); -lean_ctor_set(x_16, 1, x_15); -return x_16; -} -} -} -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_Preprocessor_preprocess___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, size_t x_6, size_t x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17) { -_start: -{ -uint8_t x_18; -x_18 = lean_usize_dec_lt(x_7, x_6); -if (x_18 == 0) -{ -lean_object* x_19; -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_4); -x_19 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_19, 0, x_8); -lean_ctor_set(x_19, 1, x_17); -return x_19; -} -else -{ -lean_object* x_20; uint8_t x_21; -x_20 = lean_array_uget(x_5, x_7); -x_21 = !lean_is_exclusive(x_8); -if (x_21 == 0) -{ -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_8, 1); -x_23 = lean_ctor_get(x_8, 0); -lean_dec(x_23); -lean_inc(x_16); -lean_inc(x_15); -lean_inc(x_14); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_22); -x_24 = l_Lean_PersistentArray_forInAux___at_Lean_Meta_Grind_Preprocessor_preprocess___spec__3(x_1, x_20, x_22, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); -lean_dec(x_20); -if (lean_obj_tag(x_24) == 0) -{ -lean_object* x_25; -x_25 = lean_ctor_get(x_24, 0); -lean_inc(x_25); -if (lean_obj_tag(x_25) == 0) -{ -uint8_t x_26; -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_4); -x_26 = !lean_is_exclusive(x_24); -if (x_26 == 0) -{ -lean_object* x_27; lean_object* x_28; -x_27 = lean_ctor_get(x_24, 0); -lean_dec(x_27); -x_28 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_28, 0, x_25); -lean_ctor_set(x_8, 0, x_28); -lean_ctor_set(x_24, 0, x_8); -return x_24; -} -else -{ -lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_29 = lean_ctor_get(x_24, 1); -lean_inc(x_29); -lean_dec(x_24); -x_30 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_30, 0, x_25); -lean_ctor_set(x_8, 0, x_30); -x_31 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_31, 0, x_8); -lean_ctor_set(x_31, 1, x_29); -return x_31; -} -} -else -{ -lean_object* x_32; lean_object* x_33; size_t x_34; size_t x_35; -lean_dec(x_22); -x_32 = lean_ctor_get(x_24, 1); -lean_inc(x_32); -lean_dec(x_24); -x_33 = lean_ctor_get(x_25, 0); -lean_inc(x_33); -lean_dec(x_25); -lean_inc(x_4); -lean_ctor_set(x_8, 1, x_33); -lean_ctor_set(x_8, 0, x_4); -x_34 = 1; -x_35 = lean_usize_add(x_7, x_34); -x_7 = x_35; -x_17 = x_32; -goto _start; -} -} -else -{ -uint8_t x_37; -lean_free_object(x_8); -lean_dec(x_22); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_4); -x_37 = !lean_is_exclusive(x_24); -if (x_37 == 0) -{ -return x_24; -} -else -{ -lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_38 = lean_ctor_get(x_24, 0); -x_39 = lean_ctor_get(x_24, 1); -lean_inc(x_39); -lean_inc(x_38); -lean_dec(x_24); -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; -x_41 = lean_ctor_get(x_8, 1); -lean_inc(x_41); -lean_dec(x_8); -lean_inc(x_16); -lean_inc(x_15); -lean_inc(x_14); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_41); -x_42 = l_Lean_PersistentArray_forInAux___at_Lean_Meta_Grind_Preprocessor_preprocess___spec__3(x_1, x_20, x_41, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); -lean_dec(x_20); -if (lean_obj_tag(x_42) == 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_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_4); -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, 1, 0); -lean_ctor_set(x_46, 0, x_43); -x_47 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_47, 0, x_46); -lean_ctor_set(x_47, 1, x_41); -if (lean_is_scalar(x_45)) { - x_48 = lean_alloc_ctor(0, 2, 0); -} else { - x_48 = x_45; -} -lean_ctor_set(x_48, 0, x_47); -lean_ctor_set(x_48, 1, x_44); -return x_48; -} -else -{ -lean_object* x_49; lean_object* x_50; lean_object* x_51; size_t x_52; size_t x_53; -lean_dec(x_41); -x_49 = lean_ctor_get(x_42, 1); -lean_inc(x_49); -lean_dec(x_42); -x_50 = lean_ctor_get(x_43, 0); -lean_inc(x_50); -lean_dec(x_43); -lean_inc(x_4); -x_51 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_51, 0, x_4); -lean_ctor_set(x_51, 1, x_50); -x_52 = 1; -x_53 = lean_usize_add(x_7, x_52); -x_7 = x_53; -x_8 = x_51; -x_17 = x_49; -goto _start; -} -} -else -{ -lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; -lean_dec(x_41); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_4); -x_55 = lean_ctor_get(x_42, 0); -lean_inc(x_55); -x_56 = lean_ctor_get(x_42, 1); -lean_inc(x_56); -if (lean_is_exclusive(x_42)) { - lean_ctor_release(x_42, 0); - lean_ctor_release(x_42, 1); - x_57 = x_42; -} else { - lean_dec_ref(x_42); - x_57 = lean_box(0); -} -if (lean_is_scalar(x_57)) { - x_58 = lean_alloc_ctor(1, 2, 0); -} else { - x_58 = x_57; -} -lean_ctor_set(x_58, 0, x_55); -lean_ctor_set(x_58, 1, x_56); -return x_58; -} -} -} -} -} -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_Preprocessor_preprocess___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) { -_start: -{ -uint8_t x_10; lean_object* x_11; -x_10 = 1; -lean_inc(x_1); -x_11 = l_Lean_Meta_Grind_checkInvariants(x_10, x_1, x_2, 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; -x_12 = lean_ctor_get(x_11, 1); -lean_inc(x_12); -lean_dec(x_11); -x_13 = lean_st_ref_get(x_1, x_12); -x_14 = !lean_is_exclusive(x_13); -if (x_14 == 0) -{ -lean_object* x_15; lean_object* x_16; uint8_t x_17; -x_15 = lean_ctor_get(x_13, 1); -x_16 = lean_st_ref_get(x_1, x_15); -lean_dec(x_1); -x_17 = !lean_is_exclusive(x_16); -if (x_17 == 0) -{ -lean_object* x_18; -x_18 = lean_ctor_get(x_16, 0); -lean_ctor_set(x_13, 1, x_18); -lean_ctor_set(x_16, 0, x_13); -return x_16; -} -else -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_19 = lean_ctor_get(x_16, 0); -x_20 = lean_ctor_get(x_16, 1); -lean_inc(x_20); -lean_inc(x_19); -lean_dec(x_16); -lean_ctor_set(x_13, 1, x_19); -x_21 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_21, 0, x_13); -lean_ctor_set(x_21, 1, x_20); -return x_21; -} -} -else -{ -lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_22 = lean_ctor_get(x_13, 0); -x_23 = lean_ctor_get(x_13, 1); -lean_inc(x_23); -lean_inc(x_22); -lean_dec(x_13); -x_24 = lean_st_ref_get(x_1, x_23); -lean_dec(x_1); -x_25 = lean_ctor_get(x_24, 0); -lean_inc(x_25); -x_26 = lean_ctor_get(x_24, 1); -lean_inc(x_26); -if (lean_is_exclusive(x_24)) { - lean_ctor_release(x_24, 0); - lean_ctor_release(x_24, 1); - x_27 = x_24; -} else { - lean_dec_ref(x_24); - x_27 = lean_box(0); -} -x_28 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_28, 0, x_22); -lean_ctor_set(x_28, 1, x_25); -if (lean_is_scalar(x_27)) { - x_29 = lean_alloc_ctor(0, 2, 0); -} else { - x_29 = x_27; -} -lean_ctor_set(x_29, 0, x_28); -lean_ctor_set(x_29, 1, x_26); -return x_29; -} -} -else -{ -uint8_t x_30; -lean_dec(x_1); -x_30 = !lean_is_exclusive(x_11); -if (x_30 == 0) -{ -return x_11; -} -else -{ -lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_31 = lean_ctor_get(x_11, 0); -x_32 = lean_ctor_get(x_11, 1); -lean_inc(x_32); -lean_inc(x_31); -lean_dec(x_11); -x_33 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_33, 0, x_31); -lean_ctor_set(x_33, 1, x_32); -return x_33; -} -} -} -} -static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_Preprocessor_preprocess___spec__5___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_Preprocessor_preprocess___spec__5___lambda__1), 9, 0); -return x_1; -} -} -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_Preprocessor_preprocess___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, size_t x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { -_start: -{ -uint8_t x_17; -x_17 = lean_usize_dec_lt(x_6, x_5); -if (x_17 == 0) -{ -lean_object* x_18; -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_3); -x_18 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_18, 0, x_7); -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_dec(x_7); -x_19 = lean_array_uget(x_4, x_6); -x_20 = lean_ctor_get(x_19, 0); -lean_inc(x_20); -x_21 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_Preprocessor_loop___lambda__1___boxed), 9, 1); -lean_closure_set(x_21, 0, x_19); -x_22 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_Preprocessor_preprocess___spec__5___closed__1; -x_23 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); -lean_closure_set(x_23, 0, x_21); -lean_closure_set(x_23, 1, x_22); -x_24 = l_Lean_Meta_Grind_Preprocessor_loop___lambda__4___closed__1; -x_25 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); -lean_closure_set(x_25, 0, x_23); -lean_closure_set(x_25, 1, x_24); -lean_inc(x_15); -lean_inc(x_14); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -x_26 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(x_20, x_25, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); -if (lean_obj_tag(x_26) == 0) -{ -lean_object* x_27; lean_object* x_28; lean_object* x_29; size_t x_30; size_t x_31; -x_27 = lean_ctor_get(x_26, 1); -lean_inc(x_27); -lean_dec(x_26); -x_28 = lean_box(0); -lean_inc(x_3); -x_29 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_29, 0, x_3); -lean_ctor_set(x_29, 1, x_28); -x_30 = 1; -x_31 = lean_usize_add(x_6, x_30); -x_6 = x_31; -x_7 = x_29; -x_16 = x_27; -goto _start; -} -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_9); -lean_dec(x_3); -x_33 = !lean_is_exclusive(x_26); -if (x_33 == 0) -{ -return x_26; -} -else -{ -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_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; -} -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at_Lean_Meta_Grind_Preprocessor_preprocess___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -if (lean_obj_tag(x_2) == 0) -{ -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; size_t x_17; size_t x_18; lean_object* x_19; -x_13 = lean_ctor_get(x_2, 0); -x_14 = lean_box(0); -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_3); -x_17 = lean_array_size(x_13); -x_18 = 0; -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -x_19 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_Preprocessor_preprocess___spec__4(x_1, x_13, x_14, x_15, x_13, x_17, x_18, x_16, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_19) == 0) -{ -lean_object* x_20; lean_object* x_21; -x_20 = lean_ctor_get(x_19, 0); -lean_inc(x_20); -x_21 = lean_ctor_get(x_20, 0); -lean_inc(x_21); -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_19, 1); -lean_inc(x_22); -lean_dec(x_19); -x_23 = lean_ctor_get(x_20, 1); -lean_inc(x_23); -lean_dec(x_20); -x_24 = lean_box(0); -x_25 = l_Lean_PersistentArray_forInAux___at_Lean_Meta_Grind_Preprocessor_ppGoals___spec__2___lambda__1(x_23, x_24, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_22); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -return x_25; -} -else -{ -uint8_t x_26; -lean_dec(x_20); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -x_26 = !lean_is_exclusive(x_19); -if (x_26 == 0) -{ -lean_object* x_27; lean_object* x_28; -x_27 = lean_ctor_get(x_19, 0); -lean_dec(x_27); -x_28 = lean_ctor_get(x_21, 0); -lean_inc(x_28); -lean_dec(x_21); -lean_ctor_set(x_19, 0, x_28); -return x_19; -} -else -{ -lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_29 = lean_ctor_get(x_19, 1); -lean_inc(x_29); -lean_dec(x_19); -x_30 = lean_ctor_get(x_21, 0); -lean_inc(x_30); -lean_dec(x_21); -x_31 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_31, 0, x_30); -lean_ctor_set(x_31, 1, x_29); -return x_31; -} -} -} -else -{ -uint8_t x_32; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -x_32 = !lean_is_exclusive(x_19); -if (x_32 == 0) -{ -return x_19; -} -else -{ -lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_33 = lean_ctor_get(x_19, 0); -x_34 = lean_ctor_get(x_19, 1); -lean_inc(x_34); -lean_inc(x_33); -lean_dec(x_19); -x_35 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -return x_35; -} -} -} -else -{ -lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; size_t x_40; size_t x_41; lean_object* x_42; -x_36 = lean_ctor_get(x_2, 0); -x_37 = lean_box(0); -x_38 = lean_box(0); -x_39 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_39, 0, x_38); -lean_ctor_set(x_39, 1, x_3); -x_40 = lean_array_size(x_36); -x_41 = 0; -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -x_42 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_Preprocessor_preprocess___spec__5(x_36, x_37, x_38, x_36, x_40, x_41, x_39, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_42) == 0) -{ -lean_object* x_43; lean_object* x_44; -x_43 = lean_ctor_get(x_42, 0); -lean_inc(x_43); -x_44 = lean_ctor_get(x_43, 0); -lean_inc(x_44); -if (lean_obj_tag(x_44) == 0) -{ -lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; -x_45 = lean_ctor_get(x_42, 1); -lean_inc(x_45); -lean_dec(x_42); -x_46 = lean_ctor_get(x_43, 1); -lean_inc(x_46); -lean_dec(x_43); -x_47 = lean_box(0); -x_48 = l_Lean_PersistentArray_forInAux___at_Lean_Meta_Grind_Preprocessor_ppGoals___spec__2___lambda__1(x_46, x_47, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, 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); -return x_48; -} -else -{ -uint8_t x_49; -lean_dec(x_43); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -x_49 = !lean_is_exclusive(x_42); -if (x_49 == 0) -{ -lean_object* x_50; lean_object* x_51; -x_50 = lean_ctor_get(x_42, 0); -lean_dec(x_50); -x_51 = lean_ctor_get(x_44, 0); -lean_inc(x_51); -lean_dec(x_44); -lean_ctor_set(x_42, 0, x_51); -return x_42; -} -else -{ -lean_object* x_52; lean_object* x_53; lean_object* x_54; -x_52 = lean_ctor_get(x_42, 1); -lean_inc(x_52); -lean_dec(x_42); -x_53 = lean_ctor_get(x_44, 0); -lean_inc(x_53); -lean_dec(x_44); -x_54 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_54, 0, x_53); -lean_ctor_set(x_54, 1, x_52); -return x_54; -} -} -} -else -{ -uint8_t x_55; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -x_55 = !lean_is_exclusive(x_42); -if (x_55 == 0) -{ -return x_42; -} -else -{ -lean_object* x_56; lean_object* x_57; lean_object* x_58; -x_56 = lean_ctor_get(x_42, 0); -x_57 = lean_ctor_get(x_42, 1); -lean_inc(x_57); -lean_inc(x_56); -lean_dec(x_42); -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_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_Preprocessor_preprocess___spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, size_t x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { -_start: -{ -uint8_t x_17; -x_17 = lean_usize_dec_lt(x_6, x_5); -if (x_17 == 0) -{ -lean_object* x_18; -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_3); -x_18 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_18, 0, x_7); -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_dec(x_7); -x_19 = lean_array_uget(x_4, x_6); -x_20 = lean_ctor_get(x_19, 0); -lean_inc(x_20); -x_21 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_Preprocessor_loop___lambda__1___boxed), 9, 1); -lean_closure_set(x_21, 0, x_19); -x_22 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_Preprocessor_preprocess___spec__5___closed__1; -x_23 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); -lean_closure_set(x_23, 0, x_21); -lean_closure_set(x_23, 1, x_22); -x_24 = l_Lean_Meta_Grind_Preprocessor_loop___lambda__4___closed__1; -x_25 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); -lean_closure_set(x_25, 0, x_23); -lean_closure_set(x_25, 1, x_24); -lean_inc(x_15); -lean_inc(x_14); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -x_26 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(x_20, x_25, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); -if (lean_obj_tag(x_26) == 0) -{ -lean_object* x_27; lean_object* x_28; lean_object* x_29; size_t x_30; size_t x_31; -x_27 = lean_ctor_get(x_26, 1); -lean_inc(x_27); -lean_dec(x_26); -x_28 = lean_box(0); -lean_inc(x_3); -x_29 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_29, 0, x_3); -lean_ctor_set(x_29, 1, x_28); -x_30 = 1; -x_31 = lean_usize_add(x_6, x_30); -x_6 = x_31; -x_7 = x_29; -x_16 = x_27; -goto _start; -} -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_9); -lean_dec(x_3); -x_33 = !lean_is_exclusive(x_26); -if (x_33 == 0) -{ -return x_26; -} -else -{ -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_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; -} -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at_Lean_Meta_Grind_Preprocessor_preprocess___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -lean_object* x_12; lean_object* x_13; -x_12 = lean_ctor_get(x_1, 0); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_2); -x_13 = l_Lean_PersistentArray_forInAux___at_Lean_Meta_Grind_Preprocessor_preprocess___spec__3(x_2, x_12, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -lean_dec(x_2); -if (lean_obj_tag(x_13) == 0) -{ -lean_object* x_14; -x_14 = lean_ctor_get(x_13, 0); -lean_inc(x_14); -if (lean_obj_tag(x_14) == 0) -{ -uint8_t x_15; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -x_15 = !lean_is_exclusive(x_13); -if (x_15 == 0) -{ -lean_object* x_16; lean_object* x_17; -x_16 = lean_ctor_get(x_13, 0); -lean_dec(x_16); -x_17 = lean_ctor_get(x_14, 0); -lean_inc(x_17); -lean_dec(x_14); -lean_ctor_set(x_13, 0, x_17); -return x_13; -} -else -{ -lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_18 = lean_ctor_get(x_13, 1); -lean_inc(x_18); -lean_dec(x_13); -x_19 = lean_ctor_get(x_14, 0); -lean_inc(x_19); -lean_dec(x_14); -x_20 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_18); -return x_20; -} -} -else -{ -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; size_t x_27; size_t x_28; lean_object* x_29; -x_21 = lean_ctor_get(x_13, 1); -lean_inc(x_21); -lean_dec(x_13); -x_22 = lean_ctor_get(x_14, 0); -lean_inc(x_22); -lean_dec(x_14); -x_23 = lean_ctor_get(x_1, 1); -x_24 = lean_box(0); -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_22); -x_27 = lean_array_size(x_23); -x_28 = 0; -x_29 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_Preprocessor_preprocess___spec__6(x_23, x_24, x_25, x_23, x_27, x_28, x_26, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_21); -if (lean_obj_tag(x_29) == 0) -{ -lean_object* x_30; lean_object* x_31; -x_30 = lean_ctor_get(x_29, 0); -lean_inc(x_30); -x_31 = lean_ctor_get(x_30, 0); -lean_inc(x_31); -if (lean_obj_tag(x_31) == 0) -{ -uint8_t x_32; -x_32 = !lean_is_exclusive(x_29); -if (x_32 == 0) -{ -lean_object* x_33; lean_object* x_34; -x_33 = lean_ctor_get(x_29, 0); -lean_dec(x_33); -x_34 = lean_ctor_get(x_30, 1); -lean_inc(x_34); -lean_dec(x_30); -lean_ctor_set(x_29, 0, x_34); -return x_29; -} -else -{ -lean_object* x_35; lean_object* x_36; lean_object* x_37; -x_35 = lean_ctor_get(x_29, 1); -lean_inc(x_35); -lean_dec(x_29); -x_36 = lean_ctor_get(x_30, 1); -lean_inc(x_36); -lean_dec(x_30); -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; -lean_dec(x_30); -x_38 = !lean_is_exclusive(x_29); -if (x_38 == 0) -{ -lean_object* x_39; lean_object* x_40; -x_39 = lean_ctor_get(x_29, 0); -lean_dec(x_39); -x_40 = lean_ctor_get(x_31, 0); -lean_inc(x_40); -lean_dec(x_31); -lean_ctor_set(x_29, 0, x_40); -return x_29; -} -else -{ -lean_object* x_41; lean_object* x_42; lean_object* x_43; -x_41 = lean_ctor_get(x_29, 1); -lean_inc(x_41); -lean_dec(x_29); -x_42 = lean_ctor_get(x_31, 0); -lean_inc(x_42); -lean_dec(x_31); -x_43 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_43, 0, x_42); -lean_ctor_set(x_43, 1, x_41); -return x_43; -} -} -} -else -{ -uint8_t x_44; -x_44 = !lean_is_exclusive(x_29); -if (x_44 == 0) -{ -return x_29; -} -else -{ -lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_45 = lean_ctor_get(x_29, 0); -x_46 = lean_ctor_get(x_29, 1); -lean_inc(x_46); -lean_inc(x_45); -lean_dec(x_29); -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_48; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -x_48 = !lean_is_exclusive(x_13); -if (x_48 == 0) -{ -return x_13; -} -else -{ -lean_object* x_49; lean_object* x_50; lean_object* x_51; -x_49 = lean_ctor_get(x_13, 0); -x_50 = lean_ctor_get(x_13, 1); -lean_inc(x_50); -lean_inc(x_49); -lean_dec(x_13); -x_51 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_51, 0, x_49); -lean_ctor_set(x_51, 1, x_50); -return x_51; -} -} -} -} -static double _init_l_Lean_addTrace___at_Lean_Meta_Grind_Preprocessor_preprocess___spec__7___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_addTrace___at_Lean_Meta_Grind_Preprocessor_preprocess___spec__7___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = lean_box(0); -x_2 = lean_array_mk(x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Meta_Grind_Preprocessor_preprocess___spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; -x_12 = lean_ctor_get(x_9, 5); -x_13 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_2, x_7, x_8, x_9, x_10, x_11); -x_14 = lean_ctor_get(x_13, 0); -lean_inc(x_14); -x_15 = lean_ctor_get(x_13, 1); -lean_inc(x_15); -lean_dec(x_13); -x_16 = lean_st_ref_take(x_10, x_15); -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -x_18 = lean_ctor_get(x_17, 3); -lean_inc(x_18); -x_19 = !lean_is_exclusive(x_16); -if (x_19 == 0) -{ -lean_object* x_20; lean_object* x_21; uint8_t x_22; -x_20 = lean_ctor_get(x_16, 1); -x_21 = lean_ctor_get(x_16, 0); -lean_dec(x_21); -x_22 = !lean_is_exclusive(x_17); -if (x_22 == 0) -{ -lean_object* x_23; uint8_t x_24; -x_23 = lean_ctor_get(x_17, 3); -lean_dec(x_23); -x_24 = !lean_is_exclusive(x_18); -if (x_24 == 0) -{ -lean_object* x_25; double 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; uint8_t x_34; -x_25 = lean_ctor_get(x_18, 0); -x_26 = l_Lean_addTrace___at_Lean_Meta_Grind_Preprocessor_preprocess___spec__7___closed__1; -x_27 = 0; -x_28 = l_Lean_Meta_Grind_Preprocessor_ppGoals___closed__1; -x_29 = lean_alloc_ctor(0, 2, 17); -lean_ctor_set(x_29, 0, x_1); -lean_ctor_set(x_29, 1, x_28); -lean_ctor_set_float(x_29, sizeof(void*)*2, x_26); -lean_ctor_set_float(x_29, sizeof(void*)*2 + 8, x_26); -lean_ctor_set_uint8(x_29, sizeof(void*)*2 + 16, x_27); -x_30 = l_Lean_addTrace___at_Lean_Meta_Grind_Preprocessor_preprocess___spec__7___closed__2; -x_31 = lean_alloc_ctor(9, 3, 0); -lean_ctor_set(x_31, 0, x_29); -lean_ctor_set(x_31, 1, x_14); -lean_ctor_set(x_31, 2, x_30); -lean_inc(x_12); -lean_ctor_set(x_16, 1, x_31); -lean_ctor_set(x_16, 0, x_12); -x_32 = l_Lean_PersistentArray_push___rarg(x_25, x_16); -lean_ctor_set(x_18, 0, x_32); -x_33 = lean_st_ref_set(x_10, x_17, x_20); -x_34 = !lean_is_exclusive(x_33); -if (x_34 == 0) -{ -lean_object* x_35; lean_object* x_36; -x_35 = lean_ctor_get(x_33, 0); -lean_dec(x_35); -x_36 = lean_box(0); -lean_ctor_set(x_33, 0, x_36); -return x_33; -} -else -{ -lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_37 = lean_ctor_get(x_33, 1); -lean_inc(x_37); -lean_dec(x_33); -x_38 = lean_box(0); -x_39 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_39, 0, x_38); -lean_ctor_set(x_39, 1, x_37); -return x_39; -} -} -else -{ -uint64_t x_40; lean_object* x_41; double 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; -x_40 = lean_ctor_get_uint64(x_18, sizeof(void*)*1); -x_41 = lean_ctor_get(x_18, 0); -lean_inc(x_41); -lean_dec(x_18); -x_42 = l_Lean_addTrace___at_Lean_Meta_Grind_Preprocessor_preprocess___spec__7___closed__1; -x_43 = 0; -x_44 = l_Lean_Meta_Grind_Preprocessor_ppGoals___closed__1; -x_45 = lean_alloc_ctor(0, 2, 17); -lean_ctor_set(x_45, 0, x_1); -lean_ctor_set(x_45, 1, x_44); -lean_ctor_set_float(x_45, sizeof(void*)*2, x_42); -lean_ctor_set_float(x_45, sizeof(void*)*2 + 8, x_42); -lean_ctor_set_uint8(x_45, sizeof(void*)*2 + 16, x_43); -x_46 = l_Lean_addTrace___at_Lean_Meta_Grind_Preprocessor_preprocess___spec__7___closed__2; -x_47 = lean_alloc_ctor(9, 3, 0); -lean_ctor_set(x_47, 0, x_45); -lean_ctor_set(x_47, 1, x_14); -lean_ctor_set(x_47, 2, x_46); -lean_inc(x_12); -lean_ctor_set(x_16, 1, x_47); -lean_ctor_set(x_16, 0, x_12); -x_48 = l_Lean_PersistentArray_push___rarg(x_41, x_16); -x_49 = lean_alloc_ctor(0, 1, 8); -lean_ctor_set(x_49, 0, x_48); -lean_ctor_set_uint64(x_49, sizeof(void*)*1, x_40); -lean_ctor_set(x_17, 3, x_49); -x_50 = lean_st_ref_set(x_10, x_17, x_20); -x_51 = lean_ctor_get(x_50, 1); -lean_inc(x_51); -if (lean_is_exclusive(x_50)) { - lean_ctor_release(x_50, 0); - lean_ctor_release(x_50, 1); - x_52 = x_50; -} else { - lean_dec_ref(x_50); - x_52 = lean_box(0); -} -x_53 = lean_box(0); -if (lean_is_scalar(x_52)) { - x_54 = lean_alloc_ctor(0, 2, 0); -} else { - x_54 = x_52; -} -lean_ctor_set(x_54, 0, x_53); -lean_ctor_set(x_54, 1, x_51); -return x_54; -} -} -else -{ -lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; uint64_t x_62; lean_object* x_63; lean_object* x_64; double x_65; uint8_t x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; -x_55 = lean_ctor_get(x_17, 0); -x_56 = lean_ctor_get(x_17, 1); -x_57 = lean_ctor_get(x_17, 2); -x_58 = lean_ctor_get(x_17, 4); -x_59 = lean_ctor_get(x_17, 5); -x_60 = lean_ctor_get(x_17, 6); -x_61 = lean_ctor_get(x_17, 7); -lean_inc(x_61); -lean_inc(x_60); -lean_inc(x_59); -lean_inc(x_58); -lean_inc(x_57); -lean_inc(x_56); -lean_inc(x_55); -lean_dec(x_17); -x_62 = lean_ctor_get_uint64(x_18, sizeof(void*)*1); -x_63 = lean_ctor_get(x_18, 0); -lean_inc(x_63); -if (lean_is_exclusive(x_18)) { - lean_ctor_release(x_18, 0); - x_64 = x_18; -} else { - lean_dec_ref(x_18); - x_64 = lean_box(0); -} -x_65 = l_Lean_addTrace___at_Lean_Meta_Grind_Preprocessor_preprocess___spec__7___closed__1; -x_66 = 0; -x_67 = l_Lean_Meta_Grind_Preprocessor_ppGoals___closed__1; -x_68 = lean_alloc_ctor(0, 2, 17); -lean_ctor_set(x_68, 0, x_1); -lean_ctor_set(x_68, 1, x_67); -lean_ctor_set_float(x_68, sizeof(void*)*2, x_65); -lean_ctor_set_float(x_68, sizeof(void*)*2 + 8, x_65); -lean_ctor_set_uint8(x_68, sizeof(void*)*2 + 16, x_66); -x_69 = l_Lean_addTrace___at_Lean_Meta_Grind_Preprocessor_preprocess___spec__7___closed__2; -x_70 = lean_alloc_ctor(9, 3, 0); -lean_ctor_set(x_70, 0, x_68); -lean_ctor_set(x_70, 1, x_14); -lean_ctor_set(x_70, 2, x_69); -lean_inc(x_12); -lean_ctor_set(x_16, 1, x_70); -lean_ctor_set(x_16, 0, x_12); -x_71 = l_Lean_PersistentArray_push___rarg(x_63, x_16); -if (lean_is_scalar(x_64)) { - x_72 = lean_alloc_ctor(0, 1, 8); -} else { - x_72 = x_64; -} -lean_ctor_set(x_72, 0, x_71); -lean_ctor_set_uint64(x_72, sizeof(void*)*1, x_62); -x_73 = lean_alloc_ctor(0, 8, 0); -lean_ctor_set(x_73, 0, x_55); -lean_ctor_set(x_73, 1, x_56); -lean_ctor_set(x_73, 2, x_57); -lean_ctor_set(x_73, 3, x_72); -lean_ctor_set(x_73, 4, x_58); -lean_ctor_set(x_73, 5, x_59); -lean_ctor_set(x_73, 6, x_60); -lean_ctor_set(x_73, 7, x_61); -x_74 = lean_st_ref_set(x_10, x_73, x_20); -x_75 = lean_ctor_get(x_74, 1); -lean_inc(x_75); -if (lean_is_exclusive(x_74)) { - lean_ctor_release(x_74, 0); - lean_ctor_release(x_74, 1); - x_76 = x_74; -} else { - lean_dec_ref(x_74); - x_76 = lean_box(0); -} -x_77 = lean_box(0); -if (lean_is_scalar(x_76)) { - x_78 = lean_alloc_ctor(0, 2, 0); -} else { - x_78 = x_76; -} -lean_ctor_set(x_78, 0, x_77); -lean_ctor_set(x_78, 1, x_75); -return x_78; -} -} -else -{ -lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; uint64_t x_88; lean_object* x_89; lean_object* x_90; double x_91; uint8_t x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; -x_79 = lean_ctor_get(x_16, 1); -lean_inc(x_79); -lean_dec(x_16); -x_80 = lean_ctor_get(x_17, 0); -lean_inc(x_80); -x_81 = lean_ctor_get(x_17, 1); -lean_inc(x_81); -x_82 = lean_ctor_get(x_17, 2); -lean_inc(x_82); -x_83 = lean_ctor_get(x_17, 4); -lean_inc(x_83); -x_84 = lean_ctor_get(x_17, 5); -lean_inc(x_84); -x_85 = lean_ctor_get(x_17, 6); -lean_inc(x_85); -x_86 = lean_ctor_get(x_17, 7); -lean_inc(x_86); -if (lean_is_exclusive(x_17)) { - lean_ctor_release(x_17, 0); - lean_ctor_release(x_17, 1); - lean_ctor_release(x_17, 2); - lean_ctor_release(x_17, 3); - lean_ctor_release(x_17, 4); - lean_ctor_release(x_17, 5); - lean_ctor_release(x_17, 6); - lean_ctor_release(x_17, 7); - x_87 = x_17; -} else { - lean_dec_ref(x_17); - x_87 = lean_box(0); -} -x_88 = lean_ctor_get_uint64(x_18, sizeof(void*)*1); -x_89 = lean_ctor_get(x_18, 0); -lean_inc(x_89); -if (lean_is_exclusive(x_18)) { - lean_ctor_release(x_18, 0); - x_90 = x_18; -} else { - lean_dec_ref(x_18); - x_90 = lean_box(0); -} -x_91 = l_Lean_addTrace___at_Lean_Meta_Grind_Preprocessor_preprocess___spec__7___closed__1; -x_92 = 0; -x_93 = l_Lean_Meta_Grind_Preprocessor_ppGoals___closed__1; -x_94 = lean_alloc_ctor(0, 2, 17); -lean_ctor_set(x_94, 0, x_1); -lean_ctor_set(x_94, 1, x_93); -lean_ctor_set_float(x_94, sizeof(void*)*2, x_91); -lean_ctor_set_float(x_94, sizeof(void*)*2 + 8, x_91); -lean_ctor_set_uint8(x_94, sizeof(void*)*2 + 16, x_92); -x_95 = l_Lean_addTrace___at_Lean_Meta_Grind_Preprocessor_preprocess___spec__7___closed__2; -x_96 = lean_alloc_ctor(9, 3, 0); -lean_ctor_set(x_96, 0, x_94); -lean_ctor_set(x_96, 1, x_14); -lean_ctor_set(x_96, 2, x_95); -lean_inc(x_12); -x_97 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_97, 0, x_12); -lean_ctor_set(x_97, 1, x_96); -x_98 = l_Lean_PersistentArray_push___rarg(x_89, x_97); -if (lean_is_scalar(x_90)) { - x_99 = lean_alloc_ctor(0, 1, 8); -} else { - x_99 = x_90; -} -lean_ctor_set(x_99, 0, x_98); -lean_ctor_set_uint64(x_99, sizeof(void*)*1, x_88); -if (lean_is_scalar(x_87)) { - x_100 = lean_alloc_ctor(0, 8, 0); -} else { - x_100 = x_87; -} -lean_ctor_set(x_100, 0, x_80); -lean_ctor_set(x_100, 1, x_81); -lean_ctor_set(x_100, 2, x_82); -lean_ctor_set(x_100, 3, x_99); -lean_ctor_set(x_100, 4, x_83); -lean_ctor_set(x_100, 5, x_84); -lean_ctor_set(x_100, 6, x_85); -lean_ctor_set(x_100, 7, x_86); -x_101 = lean_st_ref_set(x_10, x_100, x_79); -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; -} else { - lean_dec_ref(x_101); - x_103 = lean_box(0); -} -x_104 = lean_box(0); -if (lean_is_scalar(x_103)) { - x_105 = lean_alloc_ctor(0, 2, 0); -} else { - x_105 = x_103; -} -lean_ctor_set(x_105, 0, x_104); -lean_ctor_set(x_105, 1, x_102); -return x_105; -} -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Preprocessor_preprocess___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_11 = lean_st_ref_get(x_2, x_10); -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); -x_13 = lean_ctor_get(x_11, 1); -lean_inc(x_13); -lean_dec(x_11); -x_14 = lean_box(0); -x_15 = l_Lean_PersistentArray_forIn___at_Lean_Meta_Grind_Preprocessor_preprocess___spec__2(x_12, x_14, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_13); -lean_dec(x_12); -if (lean_obj_tag(x_15) == 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 = lean_st_ref_get(x_2, x_16); -x_18 = !lean_is_exclusive(x_17); -if (x_18 == 0) -{ -return x_17; -} -else -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_19 = lean_ctor_get(x_17, 0); -x_20 = lean_ctor_get(x_17, 1); -lean_inc(x_20); -lean_inc(x_19); -lean_dec(x_17); -x_21 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_21, 0, x_19); -lean_ctor_set(x_21, 1, x_20); -return x_21; -} -} -else -{ -uint8_t x_22; -x_22 = !lean_is_exclusive(x_15); -if (x_22 == 0) -{ -return x_15; -} -else -{ -lean_object* x_23; lean_object* x_24; lean_object* x_25; -x_23 = lean_ctor_get(x_15, 0); -x_24 = lean_ctor_get(x_15, 1); -lean_inc(x_24); -lean_inc(x_23); -lean_dec(x_15); -x_25 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_25, 0, x_23); -lean_ctor_set(x_25, 1, x_24); -return x_25; -} -} -} -} -static lean_object* _init_l_Lean_Meta_Grind_Preprocessor_preprocess___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_unfoldReducible), 6, 0); -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_Grind_Preprocessor_preprocess___closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_MVarId_betaReduce___lambda__1___boxed), 6, 0); -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_Grind_Preprocessor_preprocess___closed__3() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("grind", 5, 5); -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_Grind_Preprocessor_preprocess___closed__4() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("pre", 3, 3); -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_Grind_Preprocessor_preprocess___closed__5() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_Grind_Preprocessor_preprocess___closed__3; -x_2 = l_Lean_Meta_Grind_Preprocessor_preprocess___closed__4; -x_3 = l_Lean_Name_mkStr2(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Meta_Grind_Preprocessor_preprocess___closed__6() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_Preprocessor_preprocess___lambda__1___boxed), 10, 0); -return x_1; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Preprocessor_preprocess(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, 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_1); -x_11 = l_Lean_MVarId_ensureProp(x_1, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_11) == 0) -{ -lean_object* x_12; lean_object* x_13; -x_12 = lean_ctor_get(x_11, 1); -lean_inc(x_12); -lean_dec(x_11); -lean_inc(x_1); -x_13 = l_Lean_MVarId_ensureNoMVar(x_1, x_6, x_7, x_8, x_9, x_12); -if (lean_obj_tag(x_13) == 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_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -x_15 = l_Lean_MVarId_clearAuxDecls(x_1, x_6, x_7, x_8, x_9, 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, 0); -lean_inc(x_16); -x_17 = lean_ctor_get(x_15, 1); -lean_inc(x_17); -lean_dec(x_15); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -x_18 = l_Lean_MVarId_revertAll(x_16, x_6, x_7, x_8, x_9, x_17); -if (lean_obj_tag(x_18) == 0) -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_19 = lean_ctor_get(x_18, 0); -lean_inc(x_19); -x_20 = lean_ctor_get(x_18, 1); -lean_inc(x_20); -lean_dec(x_18); -lean_inc(x_19); -x_21 = l_Lean_MVarId_ensureNoMVar(x_19, x_6, x_7, x_8, x_9, x_20); -if (lean_obj_tag(x_21) == 0) -{ -lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_22 = lean_ctor_get(x_21, 1); -lean_inc(x_22); -lean_dec(x_21); -x_23 = l_Lean_Meta_Grind_getMainDeclName___rarg(x_4, x_5, x_6, x_7, x_8, x_9, x_22); -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -x_25 = lean_ctor_get(x_23, 1); -lean_inc(x_25); -lean_dec(x_23); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -x_26 = l_Lean_MVarId_abstractNestedProofs(x_19, x_24, x_6, x_7, x_8, x_9, x_25); -if (lean_obj_tag(x_26) == 0) -{ -lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; -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_Grind_Preprocessor_preprocess___closed__1; -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -x_30 = l_Lean_MVarId_transformTarget(x_27, x_29, x_6, x_7, x_8, x_9, x_28); -if (lean_obj_tag(x_30) == 0) -{ -lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_31 = lean_ctor_get(x_30, 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_Grind_Preprocessor_preprocess___closed__2; -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -x_34 = l_Lean_MVarId_transformTarget(x_31, x_33, x_6, x_7, x_8, x_9, x_32); -if (lean_obj_tag(x_34) == 0) -{ -lean_object* x_35; lean_object* x_36; lean_object* x_37; -x_35 = lean_ctor_get(x_34, 0); -lean_inc(x_35); -x_36 = lean_ctor_get(x_34, 1); -lean_inc(x_36); -lean_dec(x_34); -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_37 = l_Lean_Meta_Grind_mkGoal(x_35, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_36); -if (lean_obj_tag(x_37) == 0) -{ -lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_38 = lean_ctor_get(x_37, 0); -lean_inc(x_38); -x_39 = lean_ctor_get(x_37, 1); -lean_inc(x_39); -lean_dec(x_37); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -lean_inc(x_2); -x_40 = l_Lean_Meta_Grind_Preprocessor_loop(x_38, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_39); -if (lean_obj_tag(x_40) == 0) -{ -lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; uint8_t x_47; -x_41 = lean_ctor_get(x_40, 1); -lean_inc(x_41); -lean_dec(x_40); -x_42 = l_Lean_Meta_Grind_Preprocessor_preprocess___closed__5; -x_43 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_Preprocessor_preprocess___spec__1(x_42, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_41); -x_44 = lean_ctor_get(x_43, 0); -lean_inc(x_44); -x_45 = lean_ctor_get(x_43, 1); -lean_inc(x_45); -lean_dec(x_43); -x_46 = l_Lean_Meta_Grind_Preprocessor_preprocess___closed__6; -x_47 = lean_unbox(x_44); -lean_dec(x_44); -if (x_47 == 0) -{ -lean_object* x_48; lean_object* x_49; -x_48 = lean_box(0); -x_49 = lean_apply_10(x_46, x_48, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_45); -return x_49; -} -else -{ -lean_object* x_50; lean_object* x_51; uint8_t x_52; -x_50 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_Preprocessor_preprocess___spec__1(x_42, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_45); -x_51 = lean_ctor_get(x_50, 0); -lean_inc(x_51); -x_52 = lean_unbox(x_51); -lean_dec(x_51); -if (x_52 == 0) -{ -lean_object* x_53; lean_object* x_54; lean_object* x_55; -x_53 = lean_ctor_get(x_50, 1); -lean_inc(x_53); -lean_dec(x_50); -x_54 = lean_box(0); -x_55 = lean_apply_10(x_46, x_54, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_53); -return x_55; -} -else -{ -lean_object* x_56; lean_object* x_57; -x_56 = lean_ctor_get(x_50, 1); -lean_inc(x_56); -lean_dec(x_50); -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_57 = l_Lean_Meta_Grind_Preprocessor_ppGoals(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_56); -if (lean_obj_tag(x_57) == 0) -{ -lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; -x_58 = lean_ctor_get(x_57, 0); -lean_inc(x_58); -x_59 = lean_ctor_get(x_57, 1); -lean_inc(x_59); -lean_dec(x_57); -x_60 = l_Lean_MessageData_ofFormat(x_58); -x_61 = l_Lean_addTrace___at_Lean_Meta_Grind_Preprocessor_preprocess___spec__7(x_42, x_60, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_59); -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_apply_10(x_46, x_62, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_63); -return x_64; -} -else -{ -uint8_t x_65; -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_57); -if (x_65 == 0) -{ -return x_57; -} -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; -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_69 = !lean_is_exclusive(x_40); -if (x_69 == 0) -{ -return x_40; -} -else -{ -lean_object* x_70; lean_object* x_71; lean_object* x_72; -x_70 = lean_ctor_get(x_40, 0); -x_71 = lean_ctor_get(x_40, 1); -lean_inc(x_71); -lean_inc(x_70); -lean_dec(x_40); -x_72 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_72, 0, x_70); -lean_ctor_set(x_72, 1, x_71); -return x_72; -} -} -} -else -{ -uint8_t x_73; -lean_dec(x_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_73 = !lean_is_exclusive(x_37); -if (x_73 == 0) -{ -return x_37; -} -else -{ -lean_object* x_74; lean_object* x_75; lean_object* x_76; -x_74 = lean_ctor_get(x_37, 0); -x_75 = lean_ctor_get(x_37, 1); -lean_inc(x_75); -lean_inc(x_74); -lean_dec(x_37); -x_76 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_76, 0, x_74); -lean_ctor_set(x_76, 1, x_75); -return x_76; -} -} -} -else -{ -uint8_t x_77; -lean_dec(x_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_77 = !lean_is_exclusive(x_34); -if (x_77 == 0) -{ -return x_34; -} -else -{ -lean_object* x_78; lean_object* x_79; lean_object* x_80; -x_78 = lean_ctor_get(x_34, 0); -x_79 = lean_ctor_get(x_34, 1); -lean_inc(x_79); -lean_inc(x_78); -lean_dec(x_34); -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 -{ -uint8_t x_81; -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_81 = !lean_is_exclusive(x_30); -if (x_81 == 0) -{ -return x_30; -} -else -{ -lean_object* x_82; lean_object* x_83; lean_object* x_84; -x_82 = lean_ctor_get(x_30, 0); -x_83 = lean_ctor_get(x_30, 1); -lean_inc(x_83); -lean_inc(x_82); -lean_dec(x_30); -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_85; -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -x_85 = !lean_is_exclusive(x_26); -if (x_85 == 0) -{ -return x_26; -} -else -{ -lean_object* x_86; lean_object* x_87; lean_object* x_88; -x_86 = lean_ctor_get(x_26, 0); -x_87 = lean_ctor_get(x_26, 1); -lean_inc(x_87); -lean_inc(x_86); -lean_dec(x_26); -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_89; -lean_dec(x_19); -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_89 = !lean_is_exclusive(x_21); -if (x_89 == 0) -{ -return x_21; -} -else -{ -lean_object* x_90; lean_object* x_91; lean_object* x_92; -x_90 = lean_ctor_get(x_21, 0); -x_91 = lean_ctor_get(x_21, 1); -lean_inc(x_91); -lean_inc(x_90); -lean_dec(x_21); -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_93; -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_93 = !lean_is_exclusive(x_18); -if (x_93 == 0) -{ -return x_18; -} -else -{ -lean_object* x_94; lean_object* x_95; lean_object* x_96; -x_94 = lean_ctor_get(x_18, 0); -x_95 = lean_ctor_get(x_18, 1); -lean_inc(x_95); -lean_inc(x_94); -lean_dec(x_18); -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_97; -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_97 = !lean_is_exclusive(x_15); -if (x_97 == 0) -{ -return x_15; -} -else -{ -lean_object* x_98; lean_object* x_99; lean_object* x_100; -x_98 = lean_ctor_get(x_15, 0); -x_99 = lean_ctor_get(x_15, 1); -lean_inc(x_99); -lean_inc(x_98); -lean_dec(x_15); -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_101; -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_101 = !lean_is_exclusive(x_13); -if (x_101 == 0) -{ -return x_13; -} -else -{ -lean_object* x_102; lean_object* x_103; lean_object* x_104; -x_102 = lean_ctor_get(x_13, 0); -x_103 = lean_ctor_get(x_13, 1); -lean_inc(x_103); -lean_inc(x_102); -lean_dec(x_13); -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_105; -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_105 = !lean_is_exclusive(x_11); -if (x_105 == 0) -{ -return x_11; -} -else -{ -lean_object* x_106; lean_object* x_107; lean_object* x_108; -x_106 = lean_ctor_get(x_11, 0); -x_107 = lean_ctor_get(x_11, 1); -lean_inc(x_107); -lean_inc(x_106); -lean_dec(x_11); -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_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_Preprocessor_preprocess___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_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_Preprocessor_preprocess___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_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_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_Preprocessor_preprocess___spec__4___boxed(lean_object** _args) { -lean_object* x_1 = _args[0]; -lean_object* x_2 = _args[1]; -lean_object* x_3 = _args[2]; -lean_object* x_4 = _args[3]; -lean_object* x_5 = _args[4]; -lean_object* x_6 = _args[5]; -lean_object* x_7 = _args[6]; -lean_object* x_8 = _args[7]; -lean_object* x_9 = _args[8]; -lean_object* x_10 = _args[9]; -lean_object* x_11 = _args[10]; -lean_object* x_12 = _args[11]; -lean_object* x_13 = _args[12]; -lean_object* x_14 = _args[13]; -lean_object* x_15 = _args[14]; -lean_object* x_16 = _args[15]; -lean_object* x_17 = _args[16]; -_start: -{ -size_t x_18; size_t x_19; lean_object* x_20; -x_18 = lean_unbox_usize(x_6); -lean_dec(x_6); -x_19 = lean_unbox_usize(x_7); -lean_dec(x_7); -x_20 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_Preprocessor_preprocess___spec__4(x_1, x_2, x_3, x_4, x_5, x_18, x_19, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); -lean_dec(x_9); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_20; -} -} -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_Preprocessor_preprocess___spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { -_start: -{ -size_t x_17; size_t x_18; lean_object* x_19; -x_17 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_18 = lean_unbox_usize(x_6); -lean_dec(x_6); -x_19 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_Preprocessor_preprocess___spec__5(x_1, x_2, x_3, x_4, x_17, x_18, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); -lean_dec(x_8); -lean_dec(x_4); -lean_dec(x_2); -lean_dec(x_1); -return x_19; -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at_Lean_Meta_Grind_Preprocessor_preprocess___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -lean_object* x_13; -x_13 = l_Lean_PersistentArray_forInAux___at_Lean_Meta_Grind_Preprocessor_preprocess___spec__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_4); -lean_dec(x_2); -lean_dec(x_1); -return x_13; -} -} -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_Preprocessor_preprocess___spec__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { -_start: -{ -size_t x_17; size_t x_18; lean_object* x_19; -x_17 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_18 = lean_unbox_usize(x_6); -lean_dec(x_6); -x_19 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_Preprocessor_preprocess___spec__6(x_1, x_2, x_3, x_4, x_17, x_18, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); -lean_dec(x_8); -lean_dec(x_4); -lean_dec(x_2); -lean_dec(x_1); -return x_19; -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at_Lean_Meta_Grind_Preprocessor_preprocess___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -lean_object* x_12; -x_12 = l_Lean_PersistentArray_forIn___at_Lean_Meta_Grind_Preprocessor_preprocess___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -lean_dec(x_3); -lean_dec(x_1); -return x_12; -} -} -LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Meta_Grind_Preprocessor_preprocess___spec__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -lean_object* x_12; -x_12 = l_Lean_addTrace___at_Lean_Meta_Grind_Preprocessor_preprocess___spec__7(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -return x_12; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Preprocessor_preprocess___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_11; -x_11 = l_Lean_Meta_Grind_Preprocessor_preprocess___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_2); -lean_dec(x_1); -return x_11; -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_Grind_Preprocessor_preprocessAndProbe___spec__3(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { -_start: -{ -uint8_t x_15; -x_15 = lean_usize_dec_eq(x_3, x_4); -if (x_15 == 0) -{ -lean_object* x_16; lean_object* x_17; -lean_dec(x_5); -x_16 = lean_array_uget(x_2, x_3); -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_17 = l_Lean_PersistentArray_forMAux___at_Lean_Meta_Grind_Preprocessor_preprocessAndProbe___spec__2(x_1, x_16, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -lean_dec(x_16); -if (lean_obj_tag(x_17) == 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_inc(x_19); -lean_dec(x_17); -x_20 = 1; -x_21 = lean_usize_add(x_3, x_20); -x_3 = x_21; -x_5 = x_18; -x_14 = x_19; -goto _start; -} -else -{ -uint8_t 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_1); -x_23 = !lean_is_exclusive(x_17); -if (x_23 == 0) -{ -return x_17; -} -else -{ -lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_24 = lean_ctor_get(x_17, 0); -x_25 = lean_ctor_get(x_17, 1); -lean_inc(x_25); -lean_inc(x_24); -lean_dec(x_17); -x_26 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_26, 0, x_24); -lean_ctor_set(x_26, 1, x_25); -return x_26; -} -} -} -else -{ -lean_object* x_27; -lean_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_1); -x_27 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_27, 0, x_5); -lean_ctor_set(x_27, 1, x_14); -return x_27; -} -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_Grind_Preprocessor_preprocessAndProbe___spec__4___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_11; -lean_inc(x_2); -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) -{ -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 = lean_st_ref_get(x_2, x_12); -x_14 = !lean_is_exclusive(x_13); -if (x_14 == 0) -{ -lean_object* x_15; lean_object* x_16; uint8_t x_17; -x_15 = lean_ctor_get(x_13, 1); -x_16 = lean_st_ref_get(x_2, x_15); -lean_dec(x_2); -x_17 = !lean_is_exclusive(x_16); -if (x_17 == 0) -{ -lean_object* x_18; -x_18 = lean_ctor_get(x_16, 0); -lean_ctor_set(x_13, 1, x_18); -lean_ctor_set(x_16, 0, x_13); -return x_16; -} -else -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_19 = lean_ctor_get(x_16, 0); -x_20 = lean_ctor_get(x_16, 1); -lean_inc(x_20); -lean_inc(x_19); -lean_dec(x_16); -lean_ctor_set(x_13, 1, x_19); -x_21 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_21, 0, x_13); -lean_ctor_set(x_21, 1, x_20); -return x_21; -} -} -else -{ -lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_22 = lean_ctor_get(x_13, 0); -x_23 = lean_ctor_get(x_13, 1); -lean_inc(x_23); -lean_inc(x_22); -lean_dec(x_13); -x_24 = lean_st_ref_get(x_2, x_23); -lean_dec(x_2); -x_25 = lean_ctor_get(x_24, 0); -lean_inc(x_25); -x_26 = lean_ctor_get(x_24, 1); -lean_inc(x_26); -if (lean_is_exclusive(x_24)) { - lean_ctor_release(x_24, 0); - lean_ctor_release(x_24, 1); - x_27 = x_24; -} else { - lean_dec_ref(x_24); - x_27 = lean_box(0); -} -x_28 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_28, 0, x_22); -lean_ctor_set(x_28, 1, x_25); -if (lean_is_scalar(x_27)) { - x_29 = lean_alloc_ctor(0, 2, 0); -} else { - x_29 = x_27; -} -lean_ctor_set(x_29, 0, x_28); -lean_ctor_set(x_29, 1, x_26); -return x_29; -} -} -else -{ -uint8_t x_30; -lean_dec(x_2); -x_30 = !lean_is_exclusive(x_11); -if (x_30 == 0) -{ -return x_11; -} -else -{ -lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_31 = lean_ctor_get(x_11, 0); -x_32 = lean_ctor_get(x_11, 1); -lean_inc(x_32); -lean_inc(x_31); -lean_dec(x_11); -x_33 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_33, 0, x_31); -lean_ctor_set(x_33, 1, x_32); -return x_33; -} -} -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_Grind_Preprocessor_preprocessAndProbe___spec__4(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { -_start: -{ -uint8_t x_15; -x_15 = lean_usize_dec_eq(x_3, x_4); -if (x_15 == 0) -{ -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; -lean_dec(x_5); -x_16 = lean_array_uget(x_2, x_3); -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -x_18 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_Preprocessor_loop___lambda__1___boxed), 9, 1); -lean_closure_set(x_18, 0, x_16); -lean_inc(x_1); -x_19 = lean_alloc_closure((void*)(l_Array_foldlMUnsafe_fold___at_Lean_Meta_Grind_Preprocessor_preprocessAndProbe___spec__4___lambda__1), 10, 1); -lean_closure_set(x_19, 0, x_1); -x_20 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); -lean_closure_set(x_20, 0, x_18); -lean_closure_set(x_20, 1, x_19); -x_21 = l_Lean_Meta_Grind_Preprocessor_loop___lambda__4___closed__1; -x_22 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); -lean_closure_set(x_22, 0, x_20); -lean_closure_set(x_22, 1, x_21); -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_23 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(x_17, x_22, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_23) == 0) -{ -lean_object* x_24; size_t x_25; size_t x_26; lean_object* x_27; -x_24 = lean_ctor_get(x_23, 1); -lean_inc(x_24); -lean_dec(x_23); -x_25 = 1; -x_26 = lean_usize_add(x_3, x_25); -x_27 = lean_box(0); -x_3 = x_26; -x_5 = x_27; -x_14 = x_24; -goto _start; -} -else -{ -uint8_t x_29; -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_1); -x_29 = !lean_is_exclusive(x_23); -if (x_29 == 0) -{ -return x_23; -} -else -{ -lean_object* x_30; lean_object* x_31; lean_object* x_32; -x_30 = lean_ctor_get(x_23, 0); -x_31 = lean_ctor_get(x_23, 1); -lean_inc(x_31); -lean_inc(x_30); -lean_dec(x_23); -x_32 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_32, 0, x_30); -lean_ctor_set(x_32, 1, x_31); -return x_32; -} -} -} -else -{ -lean_object* x_33; -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_1); -x_33 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_33, 0, x_5); -lean_ctor_set(x_33, 1, x_14); -return x_33; -} -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentArray_forMAux___at_Lean_Meta_Grind_Preprocessor_preprocessAndProbe___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -if (lean_obj_tag(x_2) == 0) -{ -lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; -x_12 = lean_ctor_get(x_2, 0); -x_13 = lean_array_get_size(x_12); -x_14 = lean_unsigned_to_nat(0u); -x_15 = lean_nat_dec_lt(x_14, x_13); -if (x_15 == 0) -{ -lean_object* x_16; lean_object* x_17; -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_16 = lean_box(0); -x_17 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_17, 0, x_16); -lean_ctor_set(x_17, 1, x_11); -return x_17; -} -else -{ -uint8_t x_18; -x_18 = lean_nat_dec_le(x_13, x_13); -if (x_18 == 0) -{ -lean_object* x_19; lean_object* x_20; -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_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_11); -return x_20; -} -else -{ -size_t x_21; size_t x_22; lean_object* x_23; lean_object* x_24; -x_21 = 0; -x_22 = lean_usize_of_nat(x_13); -lean_dec(x_13); -x_23 = lean_box(0); -x_24 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_Grind_Preprocessor_preprocessAndProbe___spec__3(x_1, x_12, x_21, x_22, x_23, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_24; -} -} -} -else -{ -lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; -x_25 = lean_ctor_get(x_2, 0); -x_26 = lean_array_get_size(x_25); -x_27 = lean_unsigned_to_nat(0u); -x_28 = lean_nat_dec_lt(x_27, x_26); -if (x_28 == 0) -{ -lean_object* x_29; lean_object* x_30; -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_5); -lean_dec(x_4); -lean_dec(x_1); -x_29 = lean_box(0); -x_30 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_30, 0, x_29); -lean_ctor_set(x_30, 1, x_11); -return x_30; -} -else -{ -uint8_t x_31; -x_31 = lean_nat_dec_le(x_26, x_26); -if (x_31 == 0) -{ -lean_object* x_32; lean_object* 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_5); -lean_dec(x_4); -lean_dec(x_1); -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_11); -return x_33; -} -else -{ -size_t x_34; size_t x_35; lean_object* x_36; lean_object* x_37; -x_34 = 0; -x_35 = lean_usize_of_nat(x_26); -lean_dec(x_26); -x_36 = lean_box(0); -x_37 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_Grind_Preprocessor_preprocessAndProbe___spec__4(x_1, x_25, x_34, x_35, x_36, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_37; -} -} -} -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_Grind_Preprocessor_preprocessAndProbe___spec__5(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { -_start: -{ -uint8_t x_15; -x_15 = lean_usize_dec_eq(x_3, x_4); -if (x_15 == 0) -{ -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; -lean_dec(x_5); -x_16 = lean_array_uget(x_2, x_3); -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -x_18 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_Preprocessor_loop___lambda__1___boxed), 9, 1); -lean_closure_set(x_18, 0, x_16); -lean_inc(x_1); -x_19 = lean_alloc_closure((void*)(l_Array_foldlMUnsafe_fold___at_Lean_Meta_Grind_Preprocessor_preprocessAndProbe___spec__4___lambda__1), 10, 1); -lean_closure_set(x_19, 0, x_1); -x_20 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); -lean_closure_set(x_20, 0, x_18); -lean_closure_set(x_20, 1, x_19); -x_21 = l_Lean_Meta_Grind_Preprocessor_loop___lambda__4___closed__1; -x_22 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); -lean_closure_set(x_22, 0, x_20); -lean_closure_set(x_22, 1, x_21); -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_23 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(x_17, x_22, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_23) == 0) -{ -lean_object* x_24; size_t x_25; size_t x_26; lean_object* x_27; -x_24 = lean_ctor_get(x_23, 1); -lean_inc(x_24); -lean_dec(x_23); -x_25 = 1; -x_26 = lean_usize_add(x_3, x_25); -x_27 = lean_box(0); -x_3 = x_26; -x_5 = x_27; -x_14 = x_24; -goto _start; -} -else -{ -uint8_t x_29; -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_1); -x_29 = !lean_is_exclusive(x_23); -if (x_29 == 0) -{ -return x_23; -} -else -{ -lean_object* x_30; lean_object* x_31; lean_object* x_32; -x_30 = lean_ctor_get(x_23, 0); -x_31 = lean_ctor_get(x_23, 1); -lean_inc(x_31); -lean_inc(x_30); -lean_dec(x_23); -x_32 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_32, 0, x_30); -lean_ctor_set(x_32, 1, x_31); -return x_32; -} -} -} -else -{ -lean_object* x_33; -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_1); -x_33 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_33, 0, x_5); -lean_ctor_set(x_33, 1, x_14); -return x_33; -} -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentArray_forM___at_Lean_Meta_Grind_Preprocessor_preprocessAndProbe___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) { -_start: -{ -lean_object* x_12; lean_object* x_13; -x_12 = lean_ctor_get(x_2, 0); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_1); -x_13 = l_Lean_PersistentArray_forMAux___at_Lean_Meta_Grind_Preprocessor_preprocessAndProbe___spec__2(x_1, x_12, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_13) == 0) -{ -uint8_t x_14; -x_14 = !lean_is_exclusive(x_13); -if (x_14 == 0) -{ -lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; -x_15 = lean_ctor_get(x_13, 1); -x_16 = lean_ctor_get(x_13, 0); -lean_dec(x_16); -x_17 = lean_ctor_get(x_2, 1); -x_18 = lean_array_get_size(x_17); -x_19 = lean_unsigned_to_nat(0u); -x_20 = lean_nat_dec_lt(x_19, x_18); -if (x_20 == 0) -{ -lean_object* x_21; -lean_dec(x_18); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_1); -x_21 = lean_box(0); -lean_ctor_set(x_13, 0, x_21); -return x_13; -} -else -{ -uint8_t x_22; -x_22 = lean_nat_dec_le(x_18, x_18); -if (x_22 == 0) -{ -lean_object* x_23; -lean_dec(x_18); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_1); -x_23 = lean_box(0); -lean_ctor_set(x_13, 0, x_23); -return x_13; -} -else -{ -size_t x_24; size_t x_25; lean_object* x_26; lean_object* x_27; -lean_free_object(x_13); -x_24 = 0; -x_25 = lean_usize_of_nat(x_18); -lean_dec(x_18); -x_26 = lean_box(0); -x_27 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_Grind_Preprocessor_preprocessAndProbe___spec__5(x_1, x_17, x_24, x_25, x_26, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_15); -return x_27; -} -} -} -else -{ -lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; uint8_t x_32; -x_28 = lean_ctor_get(x_13, 1); -lean_inc(x_28); -lean_dec(x_13); -x_29 = lean_ctor_get(x_2, 1); -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_33; lean_object* x_34; -lean_dec(x_30); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_1); -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_28); -return x_34; -} -else -{ -uint8_t x_35; -x_35 = lean_nat_dec_le(x_30, x_30); -if (x_35 == 0) -{ -lean_object* x_36; lean_object* x_37; -lean_dec(x_30); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_1); -x_36 = lean_box(0); -x_37 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_37, 0, x_36); -lean_ctor_set(x_37, 1, x_28); -return x_37; -} -else -{ -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_Meta_Grind_Preprocessor_preprocessAndProbe___spec__5(x_1, x_29, x_38, x_39, x_40, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_28); -return x_41; -} -} -} -} -else -{ -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_1); -x_42 = !lean_is_exclusive(x_13); -if (x_42 == 0) -{ -return x_13; -} -else -{ -lean_object* x_43; lean_object* x_44; lean_object* x_45; -x_43 = lean_ctor_get(x_13, 0); -x_44 = lean_ctor_get(x_13, 1); -lean_inc(x_44); -lean_inc(x_43); -lean_dec(x_13); -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_Meta_Grind_Preprocessor_preprocessAndProbe(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -lean_object* x_12; -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -x_12 = l_Lean_Meta_Grind_Preprocessor_preprocess(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_12) == 0) -{ -lean_object* x_13; lean_object* x_14; lean_object* x_15; -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_PersistentArray_forM___at_Lean_Meta_Grind_Preprocessor_preprocessAndProbe___spec__1(x_2, x_13, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_14); -lean_dec(x_3); -lean_dec(x_13); -return x_15; -} -else -{ -uint8_t x_16; -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_16 = !lean_is_exclusive(x_12); -if (x_16 == 0) -{ -return x_12; -} -else -{ -lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_17 = lean_ctor_get(x_12, 0); -x_18 = lean_ctor_get(x_12, 1); -lean_inc(x_18); -lean_inc(x_17); -lean_dec(x_12); -x_19 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_19, 0, x_17); -lean_ctor_set(x_19, 1, x_18); -return x_19; -} -} -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_Grind_Preprocessor_preprocessAndProbe___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { -_start: -{ -size_t x_15; size_t x_16; lean_object* x_17; -x_15 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_16 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_17 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_Grind_Preprocessor_preprocessAndProbe___spec__3(x_1, x_2, x_15, x_16, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -lean_dec(x_6); -lean_dec(x_2); -return x_17; -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_Grind_Preprocessor_preprocessAndProbe___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { -_start: -{ -size_t x_15; size_t x_16; lean_object* x_17; -x_15 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_16 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_17 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_Grind_Preprocessor_preprocessAndProbe___spec__4(x_1, x_2, x_15, x_16, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -lean_dec(x_6); -lean_dec(x_2); -return x_17; -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentArray_forMAux___at_Lean_Meta_Grind_Preprocessor_preprocessAndProbe___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -lean_object* x_12; -x_12 = l_Lean_PersistentArray_forMAux___at_Lean_Meta_Grind_Preprocessor_preprocessAndProbe___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -lean_dec(x_3); -lean_dec(x_2); -return x_12; -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_Grind_Preprocessor_preprocessAndProbe___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) { -_start: -{ -size_t x_15; size_t x_16; lean_object* x_17; -x_15 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_16 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_17 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_Grind_Preprocessor_preprocessAndProbe___spec__5(x_1, x_2, x_15, x_16, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -lean_dec(x_6); -lean_dec(x_2); -return x_17; -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentArray_forM___at_Lean_Meta_Grind_Preprocessor_preprocessAndProbe___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -lean_object* x_12; -x_12 = l_Lean_PersistentArray_forM___at_Lean_Meta_Grind_Preprocessor_preprocessAndProbe___spec__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_3); -lean_dec(x_2); -return x_12; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_withoutModifyingMCtx___at_Lean_Meta_Grind_preprocessAndProbe___spec__1___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -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_3, x_6); -x_8 = lean_ctor_get(x_7, 0); -lean_inc(x_8); -x_9 = lean_ctor_get(x_7, 1); -lean_inc(x_9); -lean_dec(x_7); -x_10 = lean_ctor_get(x_8, 0); -lean_inc(x_10); -lean_dec(x_8); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -x_11 = lean_apply_5(x_1, x_2, x_3, x_4, x_5, x_9); -if (lean_obj_tag(x_11) == 0) -{ -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); -x_13 = lean_ctor_get(x_11, 1); -lean_inc(x_13); -lean_dec(x_11); -x_14 = l_Lean_Meta_resetCache___rarg(x_3, x_4, x_5, x_13); -lean_dec(x_5); -lean_dec(x_4); -x_15 = lean_ctor_get(x_14, 1); -lean_inc(x_15); -lean_dec(x_14); -x_16 = lean_st_ref_take(x_3, x_15); -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -x_18 = lean_ctor_get(x_16, 1); -lean_inc(x_18); -lean_dec(x_16); -x_19 = !lean_is_exclusive(x_17); -if (x_19 == 0) -{ -lean_object* x_20; lean_object* x_21; uint8_t x_22; -x_20 = lean_ctor_get(x_17, 0); -lean_dec(x_20); -lean_ctor_set(x_17, 0, x_10); -x_21 = lean_st_ref_set(x_3, x_17, x_18); -lean_dec(x_3); -x_22 = !lean_is_exclusive(x_21); -if (x_22 == 0) -{ -lean_object* x_23; -x_23 = lean_ctor_get(x_21, 0); -lean_dec(x_23); -lean_ctor_set(x_21, 0, x_12); -return x_21; -} -else -{ -lean_object* x_24; lean_object* x_25; -x_24 = lean_ctor_get(x_21, 1); -lean_inc(x_24); -lean_dec(x_21); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_12); -lean_ctor_set(x_25, 1, x_24); -return x_25; -} -} -else -{ -lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_26 = lean_ctor_get(x_17, 1); -x_27 = lean_ctor_get(x_17, 2); -x_28 = lean_ctor_get(x_17, 3); -x_29 = lean_ctor_get(x_17, 4); -lean_inc(x_29); -lean_inc(x_28); -lean_inc(x_27); -lean_inc(x_26); -lean_dec(x_17); -x_30 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_30, 0, x_10); -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_18); -lean_dec(x_3); -x_32 = lean_ctor_get(x_31, 1); -lean_inc(x_32); -if (lean_is_exclusive(x_31)) { - lean_ctor_release(x_31, 0); - lean_ctor_release(x_31, 1); - x_33 = x_31; -} else { - lean_dec_ref(x_31); - x_33 = lean_box(0); -} -if (lean_is_scalar(x_33)) { - x_34 = lean_alloc_ctor(0, 2, 0); -} else { - x_34 = x_33; -} -lean_ctor_set(x_34, 0, x_12); -lean_ctor_set(x_34, 1, x_32); -return x_34; -} -} -else -{ -lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; uint8_t x_42; -x_35 = lean_ctor_get(x_11, 0); -lean_inc(x_35); -x_36 = lean_ctor_get(x_11, 1); -lean_inc(x_36); -lean_dec(x_11); -x_37 = l_Lean_Meta_resetCache___rarg(x_3, x_4, x_5, x_36); -lean_dec(x_5); -lean_dec(x_4); -x_38 = lean_ctor_get(x_37, 1); -lean_inc(x_38); -lean_dec(x_37); -x_39 = lean_st_ref_take(x_3, x_38); -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_is_exclusive(x_40); -if (x_42 == 0) -{ -lean_object* x_43; lean_object* x_44; uint8_t x_45; -x_43 = lean_ctor_get(x_40, 0); -lean_dec(x_43); -lean_ctor_set(x_40, 0, x_10); -x_44 = lean_st_ref_set(x_3, x_40, x_41); -lean_dec(x_3); -x_45 = !lean_is_exclusive(x_44); -if (x_45 == 0) -{ -lean_object* x_46; -x_46 = lean_ctor_get(x_44, 0); -lean_dec(x_46); -lean_ctor_set_tag(x_44, 1); -lean_ctor_set(x_44, 0, x_35); -return x_44; -} -else -{ -lean_object* x_47; lean_object* x_48; -x_47 = lean_ctor_get(x_44, 1); -lean_inc(x_47); -lean_dec(x_44); -x_48 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_48, 0, x_35); -lean_ctor_set(x_48, 1, x_47); -return x_48; -} -} -else -{ -lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; -x_49 = lean_ctor_get(x_40, 1); -x_50 = lean_ctor_get(x_40, 2); -x_51 = lean_ctor_get(x_40, 3); -x_52 = lean_ctor_get(x_40, 4); -lean_inc(x_52); -lean_inc(x_51); -lean_inc(x_50); -lean_inc(x_49); -lean_dec(x_40); -x_53 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_53, 0, x_10); -lean_ctor_set(x_53, 1, x_49); -lean_ctor_set(x_53, 2, x_50); -lean_ctor_set(x_53, 3, x_51); -lean_ctor_set(x_53, 4, x_52); -x_54 = lean_st_ref_set(x_3, x_53, x_41); -lean_dec(x_3); -x_55 = lean_ctor_get(x_54, 1); -lean_inc(x_55); -if (lean_is_exclusive(x_54)) { - lean_ctor_release(x_54, 0); - lean_ctor_release(x_54, 1); - x_56 = x_54; -} else { - lean_dec_ref(x_54); - x_56 = lean_box(0); -} -if (lean_is_scalar(x_56)) { - x_57 = lean_alloc_ctor(1, 2, 0); -} else { - x_57 = x_56; - lean_ctor_set_tag(x_57, 1); -} -lean_ctor_set(x_57, 0, x_35); -lean_ctor_set(x_57, 1, x_55); -return x_57; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_withoutModifyingMCtx___at_Lean_Meta_Grind_preprocessAndProbe___spec__1(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_Meta_withoutModifyingMCtx___at_Lean_Meta_Grind_preprocessAndProbe___spec__1___rarg), 6, 0); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_preprocessAndProbe(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_9 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_Preprocessor_preprocessAndProbe), 11, 2); -lean_closure_set(x_9, 0, x_1); -lean_closure_set(x_9, 1, x_3); -x_10 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_Preprocessor_PreM_run___rarg), 9, 1); -lean_closure_set(x_10, 0, x_9); -x_11 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_GrindM_run___rarg), 7, 2); -lean_closure_set(x_11, 0, x_10); -lean_closure_set(x_11, 1, x_2); -x_12 = l_Lean_Meta_withoutModifyingMCtx___at_Lean_Meta_Grind_preprocessAndProbe___spec__1___rarg(x_11, x_4, x_5, x_6, x_7, x_8); -return x_12; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_preprocess(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: -{ -lean_object* x_8; lean_object* x_9; lean_object* x_10; -x_8 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_Preprocessor_preprocess), 10, 1); -lean_closure_set(x_8, 0, x_1); -x_9 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_Preprocessor_PreM_run___rarg), 9, 1); -lean_closure_set(x_9, 0, x_8); -x_10 = l_Lean_Meta_Grind_GrindM_run___rarg(x_9, x_2, x_3, x_4, x_5, x_6, x_7); -return x_10; -} -} -LEAN_EXPORT lean_object* l_List_filterTR_loop___at_Lean_Meta_Grind_main___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 -{ -lean_object* x_4; uint8_t x_5; -x_4 = lean_ctor_get(x_1, 0); -lean_inc(x_4); -x_5 = lean_ctor_get_uint8(x_4, sizeof(void*)*8); -if (x_5 == 0) -{ -uint8_t x_6; -x_6 = !lean_is_exclusive(x_1); -if (x_6 == 0) -{ -lean_object* x_7; lean_object* x_8; -x_7 = lean_ctor_get(x_1, 1); -x_8 = lean_ctor_get(x_1, 0); -lean_dec(x_8); -lean_ctor_set(x_1, 1, x_2); -{ -lean_object* _tmp_0 = x_7; -lean_object* _tmp_1 = x_1; -x_1 = _tmp_0; -x_2 = _tmp_1; -} -goto _start; -} -else -{ -lean_object* x_10; lean_object* x_11; -x_10 = lean_ctor_get(x_1, 1); -lean_inc(x_10); -lean_dec(x_1); -x_11 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_11, 0, x_4); -lean_ctor_set(x_11, 1, x_2); -x_1 = x_10; -x_2 = x_11; -goto _start; -} -} -else -{ -lean_object* x_13; -lean_dec(x_4); -x_13 = lean_ctor_get(x_1, 1); -lean_inc(x_13); -lean_dec(x_1); -x_1 = x_13; -goto _start; -} -} -} -} -LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Meta_Grind_main___spec__2(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 = lean_ctor_get(x_5, 0); -lean_inc(x_7); -lean_dec(x_5); -lean_ctor_set(x_1, 1, x_2); -lean_ctor_set(x_1, 0, x_7); -{ -lean_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 = lean_ctor_get(x_9, 0); -lean_inc(x_11); -lean_dec(x_9); -x_12 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_12, 0, x_11); -lean_ctor_set(x_12, 1, x_2); -x_1 = x_10; -x_2 = x_12; -goto _start; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_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, 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_PersistentArray_toList___rarg(x_1); -x_11 = lean_box(0); -x_12 = l_List_filterTR_loop___at_Lean_Meta_Grind_main___spec__1(x_10, x_11); -x_13 = l_List_mapTR_loop___at_Lean_Meta_Grind_main___spec__2(x_12, 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_9); -return x_14; -} -} -static lean_object* _init_l_Lean_Meta_Grind_main___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_main___lambda__1___boxed), 9, 0); -return x_1; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_main(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: -{ -lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_8 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_Preprocessor_preprocess), 10, 1); -lean_closure_set(x_8, 0, x_1); -x_9 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_Preprocessor_PreM_run___rarg), 9, 1); -lean_closure_set(x_9, 0, x_8); -x_10 = l_Lean_Meta_Grind_main___closed__1; -x_11 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); -lean_closure_set(x_11, 0, x_9); -lean_closure_set(x_11, 1, x_10); -x_12 = l_Lean_Meta_Grind_GrindM_run___rarg(x_11, x_2, x_3, x_4, x_5, x_6, x_7); -return x_12; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_main___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -lean_object* x_10; -x_10 = l_Lean_Meta_Grind_main___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_10; -} -} -lean_object* initialize_Init_Grind_Lemmas(uint8_t builtin, lean_object*); -lean_object* initialize_Lean_Meta_Canonicalizer(uint8_t builtin, lean_object*); -lean_object* initialize_Lean_Meta_Tactic_Util(uint8_t builtin, lean_object*); -lean_object* initialize_Lean_Meta_Tactic_Intro(uint8_t builtin, lean_object*); -lean_object* initialize_Lean_Meta_Tactic_Simp_Main(uint8_t builtin, lean_object*); -lean_object* initialize_Lean_Meta_Tactic_Grind_Attr(uint8_t builtin, lean_object*); -lean_object* initialize_Lean_Meta_Tactic_Grind_RevertAll(uint8_t builtin, lean_object*); -lean_object* initialize_Lean_Meta_Tactic_Grind_Types(uint8_t builtin, lean_object*); -lean_object* initialize_Lean_Meta_Tactic_Grind_Util(uint8_t builtin, lean_object*); -lean_object* initialize_Lean_Meta_Tactic_Grind_Cases(uint8_t builtin, lean_object*); -lean_object* initialize_Lean_Meta_Tactic_Grind_Injection(uint8_t builtin, lean_object*); -lean_object* initialize_Lean_Meta_Tactic_Grind_Core(uint8_t builtin, lean_object*); -lean_object* initialize_Lean_Meta_Tactic_Grind_Simp(uint8_t builtin, lean_object*); -lean_object* initialize_Lean_Meta_Tactic_Grind_Run(uint8_t builtin, lean_object*); -static bool _G_initialized = false; -LEAN_EXPORT lean_object* initialize_Lean_Meta_Tactic_Grind_Preprocessor(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_Grind_Lemmas(builtin, lean_io_mk_world()); -if (lean_io_result_is_error(res)) return res; -lean_dec_ref(res); -res = initialize_Lean_Meta_Canonicalizer(builtin, lean_io_mk_world()); -if (lean_io_result_is_error(res)) return res; -lean_dec_ref(res); -res = initialize_Lean_Meta_Tactic_Util(builtin, lean_io_mk_world()); -if (lean_io_result_is_error(res)) return res; -lean_dec_ref(res); -res = initialize_Lean_Meta_Tactic_Intro(builtin, lean_io_mk_world()); -if (lean_io_result_is_error(res)) return res; -lean_dec_ref(res); -res = initialize_Lean_Meta_Tactic_Simp_Main(builtin, lean_io_mk_world()); -if (lean_io_result_is_error(res)) return res; -lean_dec_ref(res); -res = initialize_Lean_Meta_Tactic_Grind_Attr(builtin, lean_io_mk_world()); -if (lean_io_result_is_error(res)) return res; -lean_dec_ref(res); -res = initialize_Lean_Meta_Tactic_Grind_RevertAll(builtin, lean_io_mk_world()); -if (lean_io_result_is_error(res)) return res; -lean_dec_ref(res); -res = initialize_Lean_Meta_Tactic_Grind_Types(builtin, lean_io_mk_world()); -if (lean_io_result_is_error(res)) return res; -lean_dec_ref(res); -res = initialize_Lean_Meta_Tactic_Grind_Util(builtin, lean_io_mk_world()); -if (lean_io_result_is_error(res)) return res; -lean_dec_ref(res); -res = initialize_Lean_Meta_Tactic_Grind_Cases(builtin, lean_io_mk_world()); -if (lean_io_result_is_error(res)) return res; -lean_dec_ref(res); -res = initialize_Lean_Meta_Tactic_Grind_Injection(builtin, lean_io_mk_world()); -if (lean_io_result_is_error(res)) return res; -lean_dec_ref(res); -res = initialize_Lean_Meta_Tactic_Grind_Core(builtin, lean_io_mk_world()); -if (lean_io_result_is_error(res)) return res; -lean_dec_ref(res); -res = initialize_Lean_Meta_Tactic_Grind_Simp(builtin, lean_io_mk_world()); -if (lean_io_result_is_error(res)) return res; -lean_dec_ref(res); -res = initialize_Lean_Meta_Tactic_Grind_Run(builtin, lean_io_mk_world()); -if (lean_io_result_is_error(res)) return res; -lean_dec_ref(res); -l_Lean_Meta_Grind_Preprocessor_instInhabitedState___closed__1 = _init_l_Lean_Meta_Grind_Preprocessor_instInhabitedState___closed__1(); -lean_mark_persistent(l_Lean_Meta_Grind_Preprocessor_instInhabitedState___closed__1); -l_Lean_Meta_Grind_Preprocessor_instInhabitedState___closed__2 = _init_l_Lean_Meta_Grind_Preprocessor_instInhabitedState___closed__2(); -lean_mark_persistent(l_Lean_Meta_Grind_Preprocessor_instInhabitedState___closed__2); -l_Lean_Meta_Grind_Preprocessor_instInhabitedState___closed__3 = _init_l_Lean_Meta_Grind_Preprocessor_instInhabitedState___closed__3(); -l_Lean_Meta_Grind_Preprocessor_instInhabitedState___closed__4 = _init_l_Lean_Meta_Grind_Preprocessor_instInhabitedState___closed__4(); -lean_mark_persistent(l_Lean_Meta_Grind_Preprocessor_instInhabitedState___closed__4); -l_Lean_Meta_Grind_Preprocessor_instInhabitedState = _init_l_Lean_Meta_Grind_Preprocessor_instInhabitedState(); -lean_mark_persistent(l_Lean_Meta_Grind_Preprocessor_instInhabitedState); -l_Lean_Meta_Grind_Preprocessor_PreM_run___rarg___closed__1 = _init_l_Lean_Meta_Grind_Preprocessor_PreM_run___rarg___closed__1(); -lean_mark_persistent(l_Lean_Meta_Grind_Preprocessor_PreM_run___rarg___closed__1); -l_Lean_Meta_Grind_Preprocessor_PreM_run___rarg___closed__2 = _init_l_Lean_Meta_Grind_Preprocessor_PreM_run___rarg___closed__2(); -lean_mark_persistent(l_Lean_Meta_Grind_Preprocessor_PreM_run___rarg___closed__2); -l_Lean_Meta_Grind_Preprocessor_PreM_run___rarg___closed__3 = _init_l_Lean_Meta_Grind_Preprocessor_PreM_run___rarg___closed__3(); -lean_mark_persistent(l_Lean_Meta_Grind_Preprocessor_PreM_run___rarg___closed__3); -l_Lean_Meta_Grind_Preprocessor_introNext___lambda__5___closed__1 = _init_l_Lean_Meta_Grind_Preprocessor_introNext___lambda__5___closed__1(); -lean_mark_persistent(l_Lean_Meta_Grind_Preprocessor_introNext___lambda__5___closed__1); -l_Lean_Meta_Grind_Preprocessor_introNext___lambda__5___closed__2 = _init_l_Lean_Meta_Grind_Preprocessor_introNext___lambda__5___closed__2(); -lean_mark_persistent(l_Lean_Meta_Grind_Preprocessor_introNext___lambda__5___closed__2); -l_Lean_Meta_Grind_Preprocessor_introNext___lambda__5___closed__3 = _init_l_Lean_Meta_Grind_Preprocessor_introNext___lambda__5___closed__3(); -lean_mark_persistent(l_Lean_Meta_Grind_Preprocessor_introNext___lambda__5___closed__3); -l_Lean_Meta_Grind_Preprocessor_introNext___lambda__5___closed__4 = _init_l_Lean_Meta_Grind_Preprocessor_introNext___lambda__5___closed__4(); -lean_mark_persistent(l_Lean_Meta_Grind_Preprocessor_introNext___lambda__5___closed__4); -l_Lean_Meta_Grind_Preprocessor_loop___lambda__4___closed__1 = _init_l_Lean_Meta_Grind_Preprocessor_loop___lambda__4___closed__1(); -lean_mark_persistent(l_Lean_Meta_Grind_Preprocessor_loop___lambda__4___closed__1); -l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_Preprocessor_ppGoals___spec__4___closed__1 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_Preprocessor_ppGoals___spec__4___closed__1(); -lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_Preprocessor_ppGoals___spec__4___closed__1); -l_Lean_Meta_Grind_Preprocessor_ppGoals___closed__1 = _init_l_Lean_Meta_Grind_Preprocessor_ppGoals___closed__1(); -lean_mark_persistent(l_Lean_Meta_Grind_Preprocessor_ppGoals___closed__1); -l_Lean_Meta_Grind_Preprocessor_ppGoals___closed__2 = _init_l_Lean_Meta_Grind_Preprocessor_ppGoals___closed__2(); -lean_mark_persistent(l_Lean_Meta_Grind_Preprocessor_ppGoals___closed__2); -l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_Preprocessor_preprocess___spec__5___closed__1 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_Preprocessor_preprocess___spec__5___closed__1(); -lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_Preprocessor_preprocess___spec__5___closed__1); -l_Lean_addTrace___at_Lean_Meta_Grind_Preprocessor_preprocess___spec__7___closed__1 = _init_l_Lean_addTrace___at_Lean_Meta_Grind_Preprocessor_preprocess___spec__7___closed__1(); -l_Lean_addTrace___at_Lean_Meta_Grind_Preprocessor_preprocess___spec__7___closed__2 = _init_l_Lean_addTrace___at_Lean_Meta_Grind_Preprocessor_preprocess___spec__7___closed__2(); -lean_mark_persistent(l_Lean_addTrace___at_Lean_Meta_Grind_Preprocessor_preprocess___spec__7___closed__2); -l_Lean_Meta_Grind_Preprocessor_preprocess___closed__1 = _init_l_Lean_Meta_Grind_Preprocessor_preprocess___closed__1(); -lean_mark_persistent(l_Lean_Meta_Grind_Preprocessor_preprocess___closed__1); -l_Lean_Meta_Grind_Preprocessor_preprocess___closed__2 = _init_l_Lean_Meta_Grind_Preprocessor_preprocess___closed__2(); -lean_mark_persistent(l_Lean_Meta_Grind_Preprocessor_preprocess___closed__2); -l_Lean_Meta_Grind_Preprocessor_preprocess___closed__3 = _init_l_Lean_Meta_Grind_Preprocessor_preprocess___closed__3(); -lean_mark_persistent(l_Lean_Meta_Grind_Preprocessor_preprocess___closed__3); -l_Lean_Meta_Grind_Preprocessor_preprocess___closed__4 = _init_l_Lean_Meta_Grind_Preprocessor_preprocess___closed__4(); -lean_mark_persistent(l_Lean_Meta_Grind_Preprocessor_preprocess___closed__4); -l_Lean_Meta_Grind_Preprocessor_preprocess___closed__5 = _init_l_Lean_Meta_Grind_Preprocessor_preprocess___closed__5(); -lean_mark_persistent(l_Lean_Meta_Grind_Preprocessor_preprocess___closed__5); -l_Lean_Meta_Grind_Preprocessor_preprocess___closed__6 = _init_l_Lean_Meta_Grind_Preprocessor_preprocess___closed__6(); -lean_mark_persistent(l_Lean_Meta_Grind_Preprocessor_preprocess___closed__6); -l_Lean_Meta_Grind_main___closed__1 = _init_l_Lean_Meta_Grind_main___closed__1(); -lean_mark_persistent(l_Lean_Meta_Grind_main___closed__1); -return lean_io_result_mk_ok(lean_box(0)); -} -#ifdef __cplusplus -} -#endif diff --git a/stage0/stdlib/Lean/Meta/Tactic/Grind/Proj.c b/stage0/stdlib/Lean/Meta/Tactic/Grind/Proj.c index 6ca513b78b78..a7040611d752 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Grind/Proj.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Grind/Proj.c @@ -14,27 +14,42 @@ extern "C" { #endif LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateProjEq___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_Grind_propagateProjEq___lambda__3___closed__1; lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateProjEq___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_Grind_propagateProjEq___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_EXPORT lean_object* l_Lean_Meta_Grind_propagateProjEq___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateProjEq___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* l_Lean_Meta_Grind_pushEqCore(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isAppOf(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateProjEq___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* l_Lean_stringToMessageData(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateProjEq___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_Grind_isCongrRoot(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_appArg_x21(lean_object*); static lean_object* l_Lean_getProjectionFnInfo_x3f___at_Lean_Meta_Grind_propagateProjEq___spec__1___closed__1; extern lean_object* l_Lean_projectionFnInfoExt; lean_object* l_Lean_Meta_Grind_getRoot(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_getRevArg_x21(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateProjEq___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_Meta_Grind_propagateProjEq___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_Grind_propagateProjEq___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_Meta_Grind_propagateProjEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_get(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_propagateProjEq___lambda__3___closed__3; +static lean_object* l_Lean_Meta_Grind_propagateProjEq___lambda__3___closed__4; +lean_object* l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_propagateProjEq___lambda__3___closed__5; LEAN_EXPORT lean_object* l_Lean_getProjectionFnInfo_x3f___at_Lean_Meta_Grind_propagateProjEq___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_getProjectionFnInfo_x3f___at_Lean_Meta_Grind_propagateProjEq___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_Grind_getGeneration(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_Grind_isSameExpr_unsafe__1(lean_object*, lean_object*); lean_object* l_Lean_MapDeclarationExtension_find_x3f___rarg(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_appFn_x21(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateProjEq___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*); extern lean_object* l_Lean_instInhabitedProjectionFunctionInfo; +lean_object* l_Lean_MessageData_ofExpr(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateProjEq___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_propagateProjEq___lambda__3___closed__6; lean_object* l_Lean_Expr_app___override(lean_object*, lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_shareCommon(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -42,9 +57,12 @@ uint8_t lean_nat_dec_lt(lean_object*, lean_object*); lean_object* l_Lean_Meta_mkEqRefl(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*); +lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___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_Meta_Grind_propagateProjEq___lambda__3___closed__2; lean_object* lean_nat_add(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateProjEq___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_internalize(lean_object*, lean_object*, 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_Grind_propagateProjEq___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* _init_l_Lean_getProjectionFnInfo_x3f___at_Lean_Meta_Grind_propagateProjEq___spec__1___closed__1() { _start: { @@ -159,15 +177,181 @@ return x_28; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateProjEq___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_Meta_Grind_propagateProjEq___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; +x_14 = lean_ctor_get(x_1, 1); +x_15 = lean_ctor_get(x_1, 2); +x_16 = lean_nat_add(x_14, x_15); +x_17 = lean_unsigned_to_nat(0u); +x_18 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_2, x_17); +x_19 = lean_nat_dec_lt(x_16, x_18); +lean_dec(x_18); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; +lean_dec(x_16); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_3); +x_20 = lean_box(0); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_20); +lean_ctor_set(x_21, 1, x_13); +return x_21; +} +else +{ +lean_object* x_22; lean_object* x_23; +x_22 = lean_box(0); +x_23 = l_Lean_Meta_Grind_propagateProjEq___lambda__1(x_2, x_16, x_3, x_22, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_16); +return x_23; +} +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateProjEq___lambda__3___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("grind", 5, 5); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateProjEq___lambda__3___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("debug", 5, 5); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateProjEq___lambda__3___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("proj", 4, 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateProjEq___lambda__3___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Meta_Grind_propagateProjEq___lambda__3___closed__1; +x_2 = l_Lean_Meta_Grind_propagateProjEq___lambda__3___closed__2; +x_3 = l_Lean_Meta_Grind_propagateProjEq___lambda__3___closed__3; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateProjEq___lambda__3___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("", 0, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateProjEq___lambda__3___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_propagateProjEq___lambda__3___closed__5; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateProjEq___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; lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_13 = l_Lean_Meta_Grind_propagateProjEq___lambda__3___closed__4; +x_14 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +x_16 = lean_unbox(x_15); +lean_dec(x_15); +if (x_16 == 0) +{ +lean_object* x_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_Grind_propagateProjEq___lambda__2(x_1, x_2, x_3, x_18, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_17); +return x_19; +} +else +{ +uint8_t x_20; +x_20 = !lean_is_exclusive(x_14); +if (x_20 == 0) +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_21 = lean_ctor_get(x_14, 1); +x_22 = lean_ctor_get(x_14, 0); +lean_dec(x_22); +lean_inc(x_3); +x_23 = l_Lean_MessageData_ofExpr(x_3); +x_24 = l_Lean_Meta_Grind_propagateProjEq___lambda__3___closed__6; +lean_ctor_set_tag(x_14, 7); +lean_ctor_set(x_14, 1, x_23); +lean_ctor_set(x_14, 0, x_24); +x_25 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_25, 0, x_14); +lean_ctor_set(x_25, 1, x_24); +x_26 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_13, x_25, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_21); +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_Grind_propagateProjEq___lambda__2(x_1, x_2, x_3, x_27, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_28); +lean_dec(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; +x_30 = lean_ctor_get(x_14, 1); +lean_inc(x_30); +lean_dec(x_14); +lean_inc(x_3); +x_31 = l_Lean_MessageData_ofExpr(x_3); +x_32 = l_Lean_Meta_Grind_propagateProjEq___lambda__3___closed__6; +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 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_34, 0, x_33); +lean_ctor_set(x_34, 1, x_32); +x_35 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_13, x_34, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, 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_Grind_propagateProjEq___lambda__2(x_1, x_2, x_3, x_36, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_37); +lean_dec(x_36); +return x_38; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateProjEq___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; -x_15 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_1, x_2); +x_15 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_3, x_2); if (x_15 == 0) { lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_16 = l_Lean_Expr_appFn_x21(x_3); +x_16 = l_Lean_Expr_appFn_x21(x_4); +lean_inc(x_2); x_17 = l_Lean_Expr_app___override(x_16, x_2); x_18 = l_Lean_Meta_Grind_shareCommon(x_17, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); x_19 = lean_ctor_get(x_18, 0); @@ -175,7 +359,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_Meta_Grind_getGeneration(x_3, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_20); +x_21 = l_Lean_Meta_Grind_getGeneration(x_4, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_20); if (lean_obj_tag(x_21) == 0) { lean_object* x_22; lean_object* x_23; lean_object* x_24; @@ -184,12 +368,33 @@ lean_inc(x_22); x_23 = lean_ctor_get(x_21, 1); lean_inc(x_23); lean_dec(x_21); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_19); x_24 = l_Lean_Meta_Grind_internalize(x_19, x_22, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_23); -return x_24; +if (lean_obj_tag(x_24) == 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_Meta_Grind_propagateProjEq___lambda__3(x_1, x_2, x_19, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_25); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_2); +return x_26; } else { -uint8_t x_25; +uint8_t x_27; lean_dec(x_19); lean_dec(x_13); lean_dec(x_12); @@ -199,40 +404,31 @@ lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -x_25 = !lean_is_exclusive(x_21); -if (x_25 == 0) +lean_dec(x_2); +x_27 = !lean_is_exclusive(x_24); +if (x_27 == 0) { -return x_21; +return x_24; } else { -lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_26 = lean_ctor_get(x_21, 0); -x_27 = lean_ctor_get(x_21, 1); -lean_inc(x_27); -lean_inc(x_26); -lean_dec(x_21); -x_28 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set(x_28, 1, x_27); -return x_28; +lean_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_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; } } } else { -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_4, 1); -x_30 = lean_ctor_get(x_4, 2); -x_31 = lean_nat_add(x_29, x_30); -x_32 = lean_unsigned_to_nat(0u); -x_33 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_2, x_32); -x_34 = lean_nat_dec_lt(x_31, x_33); -lean_dec(x_33); -if (x_34 == 0) -{ -lean_object* x_35; lean_object* x_36; -lean_dec(x_31); +uint8_t x_31; +lean_dec(x_19); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); @@ -241,31 +437,41 @@ lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_3); lean_dec(x_2); -x_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_14); -return x_36; +x_31 = !lean_is_exclusive(x_21); +if (x_31 == 0) +{ +return x_21; } else { -lean_object* x_37; lean_object* x_38; -x_37 = lean_box(0); -x_38 = l_Lean_Meta_Grind_propagateProjEq___lambda__1(x_2, x_31, x_3, x_37, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_21, 0); +x_33 = lean_ctor_get(x_21, 1); +lean_inc(x_33); +lean_inc(x_32); +lean_dec(x_21); +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_35; +x_35 = l_Lean_Meta_Grind_propagateProjEq___lambda__3(x_1, x_2, x_4, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_31); lean_dec(x_2); -return x_38; +return x_35; } } } -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateProjEq___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_Lean_Meta_Grind_propagateProjEq___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) { _start: { lean_object* x_13; lean_object* x_14; @@ -306,7 +512,7 @@ else lean_object* x_21; lean_object* x_22; lean_free_object(x_14); x_21 = lean_box(0); -x_22 = l_Lean_Meta_Grind_propagateProjEq___lambda__2(x_13, x_16, x_1, x_2, x_21, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_17); +x_22 = l_Lean_Meta_Grind_propagateProjEq___lambda__4(x_2, x_16, x_13, x_1, x_21, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_17); lean_dec(x_13); return x_22; } @@ -345,7 +551,7 @@ else { lean_object* x_29; lean_object* x_30; x_29 = lean_box(0); -x_30 = l_Lean_Meta_Grind_propagateProjEq___lambda__2(x_13, x_23, x_1, x_2, x_29, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_24); +x_30 = l_Lean_Meta_Grind_propagateProjEq___lambda__4(x_2, x_23, x_13, x_1, x_29, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_24); lean_dec(x_13); return x_30; } @@ -385,6 +591,98 @@ return x_34; } } } +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateProjEq___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) { +_start: +{ +lean_object* x_13; +lean_inc(x_1); +x_13 = l_Lean_Meta_Grind_isCongrRoot(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_13) == 0) +{ +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) +{ +uint8_t 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_1); +x_16 = !lean_is_exclusive(x_13); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; +x_17 = lean_ctor_get(x_13, 0); +lean_dec(x_17); +x_18 = lean_box(0); +lean_ctor_set(x_13, 0, x_18); +return x_13; +} +else +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_13, 1); +lean_inc(x_19); +lean_dec(x_13); +x_20 = lean_box(0); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_20); +lean_ctor_set(x_21, 1, x_19); +return x_21; +} +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_22 = lean_ctor_get(x_13, 1); +lean_inc(x_22); +lean_dec(x_13); +x_23 = lean_box(0); +x_24 = l_Lean_Meta_Grind_propagateProjEq___lambda__5(x_1, x_2, x_23, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_22); +return x_24; +} +} +else +{ +uint8_t x_25; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_25 = !lean_is_exclusive(x_13); +if (x_25 == 0) +{ +return x_13; +} +else +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_26 = lean_ctor_get(x_13, 0); +x_27 = lean_ctor_get(x_13, 1); +lean_inc(x_27); +lean_inc(x_26); +lean_dec(x_13); +x_28 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_28, 0, x_26); +lean_ctor_set(x_28, 1, x_27); +return x_28; +} +} +} +} LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateProjEq(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { @@ -479,7 +777,7 @@ else lean_object* x_32; lean_object* x_33; lean_free_object(x_13); x_32 = lean_box(0); -x_33 = l_Lean_Meta_Grind_propagateProjEq___lambda__3(x_1, x_24, x_32, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_22); +x_33 = l_Lean_Meta_Grind_propagateProjEq___lambda__6(x_1, x_24, x_32, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_22); lean_dec(x_24); return x_33; } @@ -526,7 +824,7 @@ else { lean_object* x_44; lean_object* x_45; x_44 = lean_box(0); -x_45 = l_Lean_Meta_Grind_propagateProjEq___lambda__3(x_1, x_35, x_44, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_34); +x_45 = l_Lean_Meta_Grind_propagateProjEq___lambda__6(x_1, x_35, x_44, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_34); lean_dec(x_35); return x_45; } @@ -585,15 +883,19 @@ lean_dec(x_1); return x_14; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateProjEq___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateProjEq___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { -lean_object* x_15; -x_15 = l_Lean_Meta_Grind_propagateProjEq___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +lean_object* x_14; +x_14 = l_Lean_Meta_Grind_propagateProjEq___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_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); -return x_15; +return x_14; } } LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateProjEq___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) { @@ -601,6 +903,41 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateProjEq___lambda__3___boxed(l { lean_object* x_13; x_13 = l_Lean_Meta_Grind_propagateProjEq___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateProjEq___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_Meta_Grind_propagateProjEq___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_5); +lean_dec(x_3); +lean_dec(x_1); +return x_15; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateProjEq___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) { +_start: +{ +lean_object* x_13; +x_13 = l_Lean_Meta_Grind_propagateProjEq___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); +lean_dec(x_3); +lean_dec(x_2); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateProjEq___lambda__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +x_13 = l_Lean_Meta_Grind_propagateProjEq___lambda__6(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); lean_dec(x_3); lean_dec(x_2); return x_13; @@ -625,6 +962,18 @@ if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); l_Lean_getProjectionFnInfo_x3f___at_Lean_Meta_Grind_propagateProjEq___spec__1___closed__1 = _init_l_Lean_getProjectionFnInfo_x3f___at_Lean_Meta_Grind_propagateProjEq___spec__1___closed__1(); lean_mark_persistent(l_Lean_getProjectionFnInfo_x3f___at_Lean_Meta_Grind_propagateProjEq___spec__1___closed__1); +l_Lean_Meta_Grind_propagateProjEq___lambda__3___closed__1 = _init_l_Lean_Meta_Grind_propagateProjEq___lambda__3___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateProjEq___lambda__3___closed__1); +l_Lean_Meta_Grind_propagateProjEq___lambda__3___closed__2 = _init_l_Lean_Meta_Grind_propagateProjEq___lambda__3___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateProjEq___lambda__3___closed__2); +l_Lean_Meta_Grind_propagateProjEq___lambda__3___closed__3 = _init_l_Lean_Meta_Grind_propagateProjEq___lambda__3___closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateProjEq___lambda__3___closed__3); +l_Lean_Meta_Grind_propagateProjEq___lambda__3___closed__4 = _init_l_Lean_Meta_Grind_propagateProjEq___lambda__3___closed__4(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateProjEq___lambda__3___closed__4); +l_Lean_Meta_Grind_propagateProjEq___lambda__3___closed__5 = _init_l_Lean_Meta_Grind_propagateProjEq___lambda__3___closed__5(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateProjEq___lambda__3___closed__5); +l_Lean_Meta_Grind_propagateProjEq___lambda__3___closed__6 = _init_l_Lean_Meta_Grind_propagateProjEq___lambda__3___closed__6(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateProjEq___lambda__3___closed__6); return lean_io_result_mk_ok(lean_box(0)); } #ifdef __cplusplus diff --git a/stage0/stdlib/Lean/Meta/Tactic/Grind/Proof.c b/stage0/stdlib/Lean/Meta/Tactic/Grind/Proof.c index 9d09fa98773f..3378f014e3c2 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Grind/Proof.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Grind/Proof.c @@ -20,7 +20,6 @@ static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_ static lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__5___closed__4; static lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__5___closed__5; static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__6; -static lean_object* l_Lean_Meta_Grind_mkEqProofImpl_go___closed__3; lean_object* l_Lean_mkAppN(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_isCongrDefaultProofTarget_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_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_isEqProof(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -33,7 +32,6 @@ static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_ static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore___lambda__1___closed__5; static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__12; static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkNestedProofCongr___closed__2; -static lean_object* l_Lean_addTrace___at_Lean_Meta_Grind_mkEqProofImpl_go___spec__2___closed__2; lean_object* l_Lean_Meta_mkEqSymm(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkRefl___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore___lambda__1___closed__6; @@ -42,7 +40,6 @@ uint8_t l_Lean_Expr_isApp(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___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_object*, lean_object*); uint8_t l_Lean_RBNode_isRed___rarg(lean_object*); lean_object* l_Lean_Expr_sort___override(lean_object*); -lean_object* l_Lean_PersistentArray_push___rarg(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__4; LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__7___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkProofTo___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*); @@ -50,29 +47,25 @@ lean_object* lean_mk_array(lean_object*, lean_object*); static lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__5___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___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_array_fget(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_mkEqProofImpl_go___closed__7; LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___spec__1___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_mkEqProofImpl___closed__4; lean_object* l_Lean_Meta_mkHEqOfEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_mkEqProofImpl___closed__2; -static lean_object* l_Lean_Meta_Grind_mkEqProofImpl_go___closed__2; -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEqProofImpl_go___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__5___closed__2; -static lean_object* l_Lean_Meta_Grind_mkEqProofImpl_go___closed__5; uint8_t l_Lean_Expr_isAppOf(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkProofFrom___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__7___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_stringToMessageData(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore___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_Grind_mkEqProofImpl___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof_loop(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_appArg_x21(lean_object*); -static lean_object* l_Lean_Meta_Grind_mkEqProofImpl_go___closed__1; static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore___lambda__1___closed__3; LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___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*); extern lean_object* l_instInhabitedPUnit; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_isCongrDefaultProofTarget_loop___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Meta_Grind_mkEqProofImpl_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*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkProofFrom___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* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* lean_grind_mk_heq_proof(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkNestedProofCongr___boxed(lean_object*, lean_object*, lean_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_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkProofFrom___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_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___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*); @@ -83,10 +76,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_G LEAN_EXPORT lean_object* lean_grind_mk_eq_proof(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__5; -lean_object* lean_st_ref_take(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkProofTo___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_Grind_mkEqProofImpl_go___closed__6; -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEqProofImpl_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___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkTrans_x27___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_expr_eqv(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_isEqProof___closed__1; @@ -95,17 +85,13 @@ lean_object* l_Lean_Meta_getFunInfo(lean_object*, lean_object*, lean_object*, le LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___boxed(lean_object*, lean_object*, lean_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_Grind_Proof_0__Lean_Meta_Grind_mkProofFrom___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_mkEqNDRec(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEqProofImpl___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_Meta_mkHEqRefl(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_Meta_Grind_mkEqProofImpl_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_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_flipProof___lambda__1(uint8_t, uint8_t, 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___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_flipProof___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEqProofImpl___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkProofTo___lambda__1___closed__1; lean_object* l_outOfBounds___rarg(lean_object*); -static lean_object* l_Lean_Meta_Grind_mkEqProofImpl_go___closed__8; 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_RBNode_find___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__6___boxed(lean_object*, lean_object*); lean_object* l_Lean_Meta_mkHEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -118,7 +104,6 @@ static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_ extern lean_object* l_Lean_instInhabitedExpr; lean_object* l_Lean_Meta_whnfD(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkProofTo___lambda__1___closed__3; -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEqProofImpl_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*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_flipProof(lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_panic___at_Lean_Expr_appFn_x21___spec__1(lean_object*); lean_object* l_Lean_Meta_Grind_mkHCongrWithArity(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -141,18 +126,15 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_G static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__8; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___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_RBNode_ins___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__2(lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_MessageData_ofExpr(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkTrans___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___closed__1; LEAN_EXPORT lean_object* l_Lean_RBNode_insert___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__1(lean_object*, lean_object*, lean_object*); -double l_Float_ofScientific(lean_object*, uint8_t, lean_object*); lean_object* l_Lean_Meta_mkCongr(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_isCongrDefaultProofTarget___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_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__7___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___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___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkProofFrom___lambda__1___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkProofFrom(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_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkHEqProof(lean_object*, lean_object*, lean_object*, lean_object*, 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_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern uint8_t l_Lean_Meta_instInhabitedCongrArgKind; static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__10; @@ -163,7 +145,6 @@ uint8_t lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_isCongrDefaultProofTarget_loop___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_isEqProof___closed__2; static lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkProofFrom___spec__1___closed__1; -lean_object* l_Lean_Name_mkStr2(lean_object*, 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___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkNestedProofCongr(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___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_isCongrDefaultProofTarget(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -171,15 +152,11 @@ LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_Tact uint8_t l_Lean_Expr_isConstOf(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkProofFrom___boxed(lean_object*, lean_object*, lean_object*, 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_Grind_Proof_0__Lean_Meta_Grind_mkProofTo___lambda__1___closed__2; -lean_object* l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static double l_Lean_addTrace___at_Lean_Meta_Grind_mkEqProofImpl_go___spec__2___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof(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_panic_fn(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEqProofImpl_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* lean_nat_sub(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore___lambda__1___closed__1; lean_object* l_Lean_Expr_getAppFn(lean_object*); -LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_mkEqProofImpl_go___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_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqOfHEqIfNeeded___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof___closed__1; lean_object* l_Lean_Meta_Grind_getENode(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -188,23 +165,19 @@ LEAN_EXPORT lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_Proof_0__ lean_object* lean_array_mk(lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__2; LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__3___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___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_addTrace___at_Lean_Meta_Grind_mkEqProofImpl_go___spec__2___closed__3; lean_object* l_Lean_Meta_mkCongrFun(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkProofTo(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_isTracingEnabledForCore(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_mkEqOfHEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkTrans(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore___lambda__1___closed__4; lean_object* l_instInhabitedOfMonad___rarg(lean_object*, lean_object*); -lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__7___lambda__1___closed__1; static lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__9___closed__1; lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkProofTo___boxed(lean_object*, lean_object*, lean_object*, 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_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof___closed__3; lean_object* lean_string_append(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_mkEqProofImpl_go___closed__4; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkRefl(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_mkHEqTrans(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); @@ -217,20 +190,16 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_G LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof_loop___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_mkLambdaFVars(lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Meta_Grind_congrPlaceholderProof; -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEqProofImpl_go___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof___closed__4; static lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__7___lambda__1___closed__3; lean_object* l_Lean_mkApp5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___lambda__1___closed__1; -LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_mkEqProofImpl_go___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEqProofImpl_go(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___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_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__9; static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore___lambda__1___closed__2; lean_object* l_Lean_Meta_mkHEqSymm(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_ReaderT_instMonad___rarg(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Meta_Grind_mkEqProofImpl_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*); static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_isEqProof___closed__1() { _start: { @@ -5724,7 +5693,7 @@ static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__7___lambda__1___closed__1; x_2 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore___lambda__1___closed__4; -x_3 = lean_unsigned_to_nat(206u); +x_3 = lean_unsigned_to_nat(212u); x_4 = lean_unsigned_to_nat(4u); x_5 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore___lambda__1___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -5737,8 +5706,8 @@ static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__7___lambda__1___closed__1; x_2 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore___lambda__1___closed__4; -x_3 = lean_unsigned_to_nat(209u); -x_4 = lean_unsigned_to_nat(66u); +x_3 = lean_unsigned_to_nat(215u); +x_4 = lean_unsigned_to_nat(72u); x_5 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__7___lambda__1___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); return x_6; @@ -5749,46 +5718,58 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_G { lean_object* x_14; lean_inc(x_1); -x_14 = l_Lean_Meta_Grind_getENode(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +x_14 = l_Lean_Meta_Grind_getRootENode(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); if (lean_obj_tag(x_14) == 0) { -lean_object* x_15; lean_object* x_16; lean_object* x_17; +lean_object* x_15; lean_object* x_16; uint8_t x_17; lean_object* x_18; x_15 = lean_ctor_get(x_14, 0); lean_inc(x_15); x_16 = lean_ctor_get(x_14, 1); lean_inc(x_16); lean_dec(x_14); -lean_inc(x_2); -x_17 = l_Lean_Meta_Grind_getENode(x_2, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_16); -if (lean_obj_tag(x_17) == 0) +x_17 = lean_ctor_get_uint8(x_15, sizeof(void*)*10 + 4); +lean_dec(x_15); +lean_inc(x_1); +x_18 = l_Lean_Meta_Grind_getENode(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_16); +if (lean_obj_tag(x_18) == 0) { -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; -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_15, 2); +x_20 = lean_ctor_get(x_18, 1); lean_inc(x_20); -lean_dec(x_15); -x_21 = lean_ctor_get(x_18, 2); -lean_inc(x_21); lean_dec(x_18); -x_22 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_20, x_21); +lean_inc(x_2); +x_21 = l_Lean_Meta_Grind_getENode(x_2, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_20); +if (lean_obj_tag(x_21) == 0) +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; +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_dec(x_20); -if (x_22 == 0) +x_24 = lean_ctor_get(x_19, 2); +lean_inc(x_24); +lean_dec(x_19); +x_25 = lean_ctor_get(x_22, 2); +lean_inc(x_25); +lean_dec(x_22); +x_26 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_24, x_25); +lean_dec(x_25); +lean_dec(x_24); +if (x_26 == 0) { -lean_object* x_23; lean_object* x_24; +lean_object* x_27; lean_object* x_28; lean_dec(x_2); lean_dec(x_1); -x_23 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore___lambda__1___closed__5; -x_24 = l_panic___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__9(x_23, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_19); -return x_24; +x_27 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore___lambda__1___closed__5; +x_28 = l_panic___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__9(x_27, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_23); +return x_28; } else { -lean_object* x_25; +lean_object* x_29; lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); @@ -5799,16 +5780,16 @@ lean_inc(x_6); lean_inc(x_5); lean_inc(x_2); lean_inc(x_1); -x_25 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon(x_1, x_2, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_19); -if (lean_obj_tag(x_25) == 0) +x_29 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon(x_1, x_2, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_23); +if (lean_obj_tag(x_29) == 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 = lean_box(0); +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_30 = lean_ctor_get(x_29, 0); +lean_inc(x_30); +x_31 = lean_ctor_get(x_29, 1); +lean_inc(x_31); +lean_dec(x_29); +x_32 = lean_box(0); lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); @@ -5817,15 +5798,15 @@ lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -x_29 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkProofTo(x_1, x_26, x_28, x_3, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_27); -if (lean_obj_tag(x_29) == 0) +x_33 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkProofTo(x_1, x_30, x_32, x_17, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_31); +if (lean_obj_tag(x_33) == 0) { -lean_object* x_30; lean_object* x_31; lean_object* x_32; -x_30 = lean_ctor_get(x_29, 0); -lean_inc(x_30); -x_31 = lean_ctor_get(x_29, 1); -lean_inc(x_31); -lean_dec(x_29); +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); lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); @@ -5834,26 +5815,117 @@ lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -x_32 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkProofFrom(x_2, x_26, x_30, x_3, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_31); -lean_dec(x_26); -if (lean_obj_tag(x_32) == 0) +x_36 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkProofFrom(x_2, x_30, x_34, x_17, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_35); +lean_dec(x_30); +if (lean_obj_tag(x_36) == 0) { -lean_object* x_33; -x_33 = lean_ctor_get(x_32, 0); -lean_inc(x_33); -if (lean_obj_tag(x_33) == 0) +lean_object* x_37; +x_37 = lean_ctor_get(x_36, 0); +lean_inc(x_37); +if (lean_obj_tag(x_37) == 0) { -lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_34 = lean_ctor_get(x_32, 1); -lean_inc(x_34); -lean_dec(x_32); -x_35 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore___lambda__1___closed__6; -x_36 = l_panic___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__9(x_35, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_34); -return x_36; +lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_38 = lean_ctor_get(x_36, 1); +lean_inc(x_38); +lean_dec(x_36); +x_39 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore___lambda__1___closed__6; +x_40 = l_panic___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__9(x_39, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_38); +return x_40; } else { -uint8_t x_37; +lean_object* x_41; lean_object* x_42; lean_object* x_43; uint8_t x_44; +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +x_41 = lean_ctor_get(x_36, 1); +lean_inc(x_41); +if (lean_is_exclusive(x_36)) { + lean_ctor_release(x_36, 0); + lean_ctor_release(x_36, 1); + x_42 = x_36; +} else { + lean_dec_ref(x_36); + x_42 = lean_box(0); +} +x_43 = lean_ctor_get(x_37, 0); +lean_inc(x_43); +lean_dec(x_37); +if (x_3 == 0) +{ +if (x_17 == 0) +{ +uint8_t x_49; +x_49 = 1; +x_44 = x_49; +goto block_48; +} +else +{ +uint8_t x_50; +x_50 = 0; +x_44 = x_50; +goto block_48; +} +} +else +{ +if (x_17 == 0) +{ +uint8_t x_51; +x_51 = 0; +x_44 = x_51; +goto block_48; +} +else +{ +uint8_t x_52; +x_52 = 1; +x_44 = x_52; +goto block_48; +} +} +block_48: +{ +if (x_44 == 0) +{ +lean_dec(x_42); +if (x_3 == 0) +{ +lean_object* x_45; +x_45 = l_Lean_Meta_mkEqOfHEq(x_43, x_9, x_10, x_11, x_12, x_41); +return x_45; +} +else +{ +lean_object* x_46; +x_46 = l_Lean_Meta_mkHEqOfEq(x_43, x_9, x_10, x_11, x_12, x_41); +return x_46; +} +} +else +{ +lean_object* x_47; +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +if (lean_is_scalar(x_42)) { + x_47 = lean_alloc_ctor(0, 2, 0); +} else { + x_47 = x_42; +} +lean_ctor_set(x_47, 0, x_43); +lean_ctor_set(x_47, 1, x_41); +return x_47; +} +} +} +} +else +{ +uint8_t x_53; lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); @@ -5862,37 +5934,30 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -x_37 = !lean_is_exclusive(x_32); -if (x_37 == 0) +x_53 = !lean_is_exclusive(x_36); +if (x_53 == 0) { -lean_object* x_38; lean_object* x_39; -x_38 = lean_ctor_get(x_32, 0); -lean_dec(x_38); -x_39 = lean_ctor_get(x_33, 0); -lean_inc(x_39); -lean_dec(x_33); -lean_ctor_set(x_32, 0, x_39); -return x_32; +return x_36; } else { -lean_object* x_40; lean_object* x_41; lean_object* x_42; -x_40 = lean_ctor_get(x_32, 1); -lean_inc(x_40); -lean_dec(x_32); -x_41 = lean_ctor_get(x_33, 0); -lean_inc(x_41); -lean_dec(x_33); -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_54; lean_object* x_55; lean_object* x_56; +x_54 = lean_ctor_get(x_36, 0); +x_55 = lean_ctor_get(x_36, 1); +lean_inc(x_55); +lean_inc(x_54); +lean_dec(x_36); +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_43; +uint8_t x_57; +lean_dec(x_30); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); @@ -5901,30 +5966,30 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -x_43 = !lean_is_exclusive(x_32); -if (x_43 == 0) +lean_dec(x_2); +x_57 = !lean_is_exclusive(x_33); +if (x_57 == 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_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_object* x_58; lean_object* x_59; lean_object* x_60; +x_58 = lean_ctor_get(x_33, 0); +x_59 = lean_ctor_get(x_33, 1); +lean_inc(x_59); +lean_inc(x_58); +lean_dec(x_33); +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_47; -lean_dec(x_26); +uint8_t x_61; lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); @@ -5934,29 +5999,32 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_2); -x_47 = !lean_is_exclusive(x_29); -if (x_47 == 0) +lean_dec(x_1); +x_61 = !lean_is_exclusive(x_29); +if (x_61 == 0) { return x_29; } else { -lean_object* x_48; lean_object* x_49; lean_object* x_50; -x_48 = lean_ctor_get(x_29, 0); -x_49 = lean_ctor_get(x_29, 1); -lean_inc(x_49); -lean_inc(x_48); +lean_object* x_62; lean_object* x_63; lean_object* x_64; +x_62 = lean_ctor_get(x_29, 0); +x_63 = lean_ctor_get(x_29, 1); +lean_inc(x_63); +lean_inc(x_62); 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; +x_64 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_64, 0, x_62); +lean_ctor_set(x_64, 1, x_63); +return x_64; +} } } } else { -uint8_t x_51; +uint8_t x_65; +lean_dec(x_19); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); @@ -5967,31 +6035,29 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_2); lean_dec(x_1); -x_51 = !lean_is_exclusive(x_25); -if (x_51 == 0) +x_65 = !lean_is_exclusive(x_21); +if (x_65 == 0) { -return x_25; +return x_21; } else { -lean_object* x_52; lean_object* x_53; lean_object* x_54; -x_52 = lean_ctor_get(x_25, 0); -x_53 = lean_ctor_get(x_25, 1); -lean_inc(x_53); -lean_inc(x_52); -lean_dec(x_25); -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_66; lean_object* x_67; lean_object* x_68; +x_66 = lean_ctor_get(x_21, 0); +x_67 = lean_ctor_get(x_21, 1); +lean_inc(x_67); +lean_inc(x_66); +lean_dec(x_21); +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_55; -lean_dec(x_15); +uint8_t x_69; lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); @@ -6002,29 +6068,29 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_2); lean_dec(x_1); -x_55 = !lean_is_exclusive(x_17); -if (x_55 == 0) +x_69 = !lean_is_exclusive(x_18); +if (x_69 == 0) { -return x_17; +return x_18; } else { -lean_object* x_56; lean_object* x_57; lean_object* x_58; -x_56 = lean_ctor_get(x_17, 0); -x_57 = lean_ctor_get(x_17, 1); -lean_inc(x_57); -lean_inc(x_56); -lean_dec(x_17); -x_58 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_58, 0, x_56); -lean_ctor_set(x_58, 1, x_57); -return x_58; +lean_object* x_70; lean_object* x_71; lean_object* x_72; +x_70 = lean_ctor_get(x_18, 0); +x_71 = lean_ctor_get(x_18, 1); +lean_inc(x_71); +lean_inc(x_70); +lean_dec(x_18); +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_59; +uint8_t x_73; lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); @@ -6035,23 +6101,23 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_2); lean_dec(x_1); -x_59 = !lean_is_exclusive(x_14); -if (x_59 == 0) +x_73 = !lean_is_exclusive(x_14); +if (x_73 == 0) { return x_14; } else { -lean_object* x_60; lean_object* x_61; lean_object* x_62; -x_60 = lean_ctor_get(x_14, 0); -x_61 = lean_ctor_get(x_14, 1); -lean_inc(x_61); -lean_inc(x_60); +lean_object* x_74; lean_object* x_75; lean_object* x_76; +x_74 = lean_ctor_get(x_14, 0); +x_75 = lean_ctor_get(x_14, 1); +lean_inc(x_75); +lean_inc(x_74); lean_dec(x_14); -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_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; } } } @@ -9226,981 +9292,43 @@ lean_dec(x_2); return x_15; } } -LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_mkEqProofImpl_go___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +static lean_object* _init_l_Lean_Meta_Grind_mkEqProofImpl___closed__1() { _start: { -lean_object* x_11; lean_object* x_12; uint8_t x_13; -x_11 = lean_ctor_get(x_8, 2); -x_12 = l_Lean_isTracingEnabledForCore(x_1, x_11, x_10); -x_13 = !lean_is_exclusive(x_12); -if (x_13 == 0) -{ -return x_12; -} -else -{ -lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_14 = lean_ctor_get(x_12, 0); -x_15 = lean_ctor_get(x_12, 1); -lean_inc(x_15); -lean_inc(x_14); -lean_dec(x_12); -x_16 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_16, 0, x_14); -lean_ctor_set(x_16, 1, x_15); -return x_16; -} +lean_object* x_1; +x_1 = lean_mk_string_unchecked("( __do_lift._@.Lean.Meta.Tactic.Grind.Proof._hyg.2379.0 )\n ", 60, 60); +return x_1; } } -static double _init_l_Lean_addTrace___at_Lean_Meta_Grind_mkEqProofImpl_go___spec__2___closed__1() { +static lean_object* _init_l_Lean_Meta_Grind_mkEqProofImpl___closed__2() { _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); +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore___lambda__1___closed__1; +x_2 = l_Lean_Meta_Grind_mkEqProofImpl___closed__1; +x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_addTrace___at_Lean_Meta_Grind_mkEqProofImpl_go___spec__2___closed__2() { +static lean_object* _init_l_Lean_Meta_Grind_mkEqProofImpl___closed__3() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("", 0, 0); +x_1 = lean_mk_string_unchecked("Lean.Meta.Grind.mkEqProofImpl", 29, 29); return x_1; } } -static lean_object* _init_l_Lean_addTrace___at_Lean_Meta_Grind_mkEqProofImpl_go___spec__2___closed__3() { +static lean_object* _init_l_Lean_Meta_Grind_mkEqProofImpl___closed__4() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = lean_box(0); -x_2 = lean_array_mk(x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Meta_Grind_mkEqProofImpl_go___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -lean_object* x_12; 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_9, 5); -x_13 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_2, x_7, x_8, x_9, x_10, x_11); -x_14 = lean_ctor_get(x_13, 0); -lean_inc(x_14); -x_15 = lean_ctor_get(x_13, 1); -lean_inc(x_15); -lean_dec(x_13); -x_16 = lean_st_ref_take(x_10, x_15); -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -x_18 = lean_ctor_get(x_17, 3); -lean_inc(x_18); -x_19 = !lean_is_exclusive(x_16); -if (x_19 == 0) -{ -lean_object* x_20; lean_object* x_21; uint8_t x_22; -x_20 = lean_ctor_get(x_16, 1); -x_21 = lean_ctor_get(x_16, 0); -lean_dec(x_21); -x_22 = !lean_is_exclusive(x_17); -if (x_22 == 0) -{ -lean_object* x_23; uint8_t x_24; -x_23 = lean_ctor_get(x_17, 3); -lean_dec(x_23); -x_24 = !lean_is_exclusive(x_18); -if (x_24 == 0) -{ -lean_object* x_25; double 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; uint8_t x_34; -x_25 = lean_ctor_get(x_18, 0); -x_26 = l_Lean_addTrace___at_Lean_Meta_Grind_mkEqProofImpl_go___spec__2___closed__1; -x_27 = 0; -x_28 = l_Lean_addTrace___at_Lean_Meta_Grind_mkEqProofImpl_go___spec__2___closed__2; -x_29 = lean_alloc_ctor(0, 2, 17); -lean_ctor_set(x_29, 0, x_1); -lean_ctor_set(x_29, 1, x_28); -lean_ctor_set_float(x_29, sizeof(void*)*2, x_26); -lean_ctor_set_float(x_29, sizeof(void*)*2 + 8, x_26); -lean_ctor_set_uint8(x_29, sizeof(void*)*2 + 16, x_27); -x_30 = l_Lean_addTrace___at_Lean_Meta_Grind_mkEqProofImpl_go___spec__2___closed__3; -x_31 = lean_alloc_ctor(9, 3, 0); -lean_ctor_set(x_31, 0, x_29); -lean_ctor_set(x_31, 1, x_14); -lean_ctor_set(x_31, 2, x_30); -lean_inc(x_12); -lean_ctor_set(x_16, 1, x_31); -lean_ctor_set(x_16, 0, x_12); -x_32 = l_Lean_PersistentArray_push___rarg(x_25, x_16); -lean_ctor_set(x_18, 0, x_32); -x_33 = lean_st_ref_set(x_10, x_17, x_20); -x_34 = !lean_is_exclusive(x_33); -if (x_34 == 0) -{ -lean_object* x_35; lean_object* x_36; -x_35 = lean_ctor_get(x_33, 0); -lean_dec(x_35); -x_36 = lean_box(0); -lean_ctor_set(x_33, 0, x_36); -return x_33; -} -else -{ -lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_37 = lean_ctor_get(x_33, 1); -lean_inc(x_37); -lean_dec(x_33); -x_38 = lean_box(0); -x_39 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_39, 0, x_38); -lean_ctor_set(x_39, 1, x_37); -return x_39; -} -} -else -{ -uint64_t x_40; lean_object* x_41; double 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; -x_40 = lean_ctor_get_uint64(x_18, sizeof(void*)*1); -x_41 = lean_ctor_get(x_18, 0); -lean_inc(x_41); -lean_dec(x_18); -x_42 = l_Lean_addTrace___at_Lean_Meta_Grind_mkEqProofImpl_go___spec__2___closed__1; -x_43 = 0; -x_44 = l_Lean_addTrace___at_Lean_Meta_Grind_mkEqProofImpl_go___spec__2___closed__2; -x_45 = lean_alloc_ctor(0, 2, 17); -lean_ctor_set(x_45, 0, x_1); -lean_ctor_set(x_45, 1, x_44); -lean_ctor_set_float(x_45, sizeof(void*)*2, x_42); -lean_ctor_set_float(x_45, sizeof(void*)*2 + 8, x_42); -lean_ctor_set_uint8(x_45, sizeof(void*)*2 + 16, x_43); -x_46 = l_Lean_addTrace___at_Lean_Meta_Grind_mkEqProofImpl_go___spec__2___closed__3; -x_47 = lean_alloc_ctor(9, 3, 0); -lean_ctor_set(x_47, 0, x_45); -lean_ctor_set(x_47, 1, x_14); -lean_ctor_set(x_47, 2, x_46); -lean_inc(x_12); -lean_ctor_set(x_16, 1, x_47); -lean_ctor_set(x_16, 0, x_12); -x_48 = l_Lean_PersistentArray_push___rarg(x_41, x_16); -x_49 = lean_alloc_ctor(0, 1, 8); -lean_ctor_set(x_49, 0, x_48); -lean_ctor_set_uint64(x_49, sizeof(void*)*1, x_40); -lean_ctor_set(x_17, 3, x_49); -x_50 = lean_st_ref_set(x_10, x_17, x_20); -x_51 = lean_ctor_get(x_50, 1); -lean_inc(x_51); -if (lean_is_exclusive(x_50)) { - lean_ctor_release(x_50, 0); - lean_ctor_release(x_50, 1); - x_52 = x_50; -} else { - lean_dec_ref(x_50); - x_52 = lean_box(0); -} -x_53 = lean_box(0); -if (lean_is_scalar(x_52)) { - x_54 = lean_alloc_ctor(0, 2, 0); -} else { - x_54 = x_52; -} -lean_ctor_set(x_54, 0, x_53); -lean_ctor_set(x_54, 1, x_51); -return x_54; -} -} -else -{ -lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; uint64_t x_62; lean_object* x_63; lean_object* x_64; double x_65; uint8_t x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; -x_55 = lean_ctor_get(x_17, 0); -x_56 = lean_ctor_get(x_17, 1); -x_57 = lean_ctor_get(x_17, 2); -x_58 = lean_ctor_get(x_17, 4); -x_59 = lean_ctor_get(x_17, 5); -x_60 = lean_ctor_get(x_17, 6); -x_61 = lean_ctor_get(x_17, 7); -lean_inc(x_61); -lean_inc(x_60); -lean_inc(x_59); -lean_inc(x_58); -lean_inc(x_57); -lean_inc(x_56); -lean_inc(x_55); -lean_dec(x_17); -x_62 = lean_ctor_get_uint64(x_18, sizeof(void*)*1); -x_63 = lean_ctor_get(x_18, 0); -lean_inc(x_63); -if (lean_is_exclusive(x_18)) { - lean_ctor_release(x_18, 0); - x_64 = x_18; -} else { - lean_dec_ref(x_18); - x_64 = lean_box(0); -} -x_65 = l_Lean_addTrace___at_Lean_Meta_Grind_mkEqProofImpl_go___spec__2___closed__1; -x_66 = 0; -x_67 = l_Lean_addTrace___at_Lean_Meta_Grind_mkEqProofImpl_go___spec__2___closed__2; -x_68 = lean_alloc_ctor(0, 2, 17); -lean_ctor_set(x_68, 0, x_1); -lean_ctor_set(x_68, 1, x_67); -lean_ctor_set_float(x_68, sizeof(void*)*2, x_65); -lean_ctor_set_float(x_68, sizeof(void*)*2 + 8, x_65); -lean_ctor_set_uint8(x_68, sizeof(void*)*2 + 16, x_66); -x_69 = l_Lean_addTrace___at_Lean_Meta_Grind_mkEqProofImpl_go___spec__2___closed__3; -x_70 = lean_alloc_ctor(9, 3, 0); -lean_ctor_set(x_70, 0, x_68); -lean_ctor_set(x_70, 1, x_14); -lean_ctor_set(x_70, 2, x_69); -lean_inc(x_12); -lean_ctor_set(x_16, 1, x_70); -lean_ctor_set(x_16, 0, x_12); -x_71 = l_Lean_PersistentArray_push___rarg(x_63, x_16); -if (lean_is_scalar(x_64)) { - x_72 = lean_alloc_ctor(0, 1, 8); -} else { - x_72 = x_64; -} -lean_ctor_set(x_72, 0, x_71); -lean_ctor_set_uint64(x_72, sizeof(void*)*1, x_62); -x_73 = lean_alloc_ctor(0, 8, 0); -lean_ctor_set(x_73, 0, x_55); -lean_ctor_set(x_73, 1, x_56); -lean_ctor_set(x_73, 2, x_57); -lean_ctor_set(x_73, 3, x_72); -lean_ctor_set(x_73, 4, x_58); -lean_ctor_set(x_73, 5, x_59); -lean_ctor_set(x_73, 6, x_60); -lean_ctor_set(x_73, 7, x_61); -x_74 = lean_st_ref_set(x_10, x_73, x_20); -x_75 = lean_ctor_get(x_74, 1); -lean_inc(x_75); -if (lean_is_exclusive(x_74)) { - lean_ctor_release(x_74, 0); - lean_ctor_release(x_74, 1); - x_76 = x_74; -} else { - lean_dec_ref(x_74); - x_76 = lean_box(0); -} -x_77 = lean_box(0); -if (lean_is_scalar(x_76)) { - x_78 = lean_alloc_ctor(0, 2, 0); -} else { - x_78 = x_76; -} -lean_ctor_set(x_78, 0, x_77); -lean_ctor_set(x_78, 1, x_75); -return x_78; -} -} -else -{ -lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; uint64_t x_88; lean_object* x_89; lean_object* x_90; double x_91; uint8_t x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; -x_79 = lean_ctor_get(x_16, 1); -lean_inc(x_79); -lean_dec(x_16); -x_80 = lean_ctor_get(x_17, 0); -lean_inc(x_80); -x_81 = lean_ctor_get(x_17, 1); -lean_inc(x_81); -x_82 = lean_ctor_get(x_17, 2); -lean_inc(x_82); -x_83 = lean_ctor_get(x_17, 4); -lean_inc(x_83); -x_84 = lean_ctor_get(x_17, 5); -lean_inc(x_84); -x_85 = lean_ctor_get(x_17, 6); -lean_inc(x_85); -x_86 = lean_ctor_get(x_17, 7); -lean_inc(x_86); -if (lean_is_exclusive(x_17)) { - lean_ctor_release(x_17, 0); - lean_ctor_release(x_17, 1); - lean_ctor_release(x_17, 2); - lean_ctor_release(x_17, 3); - lean_ctor_release(x_17, 4); - lean_ctor_release(x_17, 5); - lean_ctor_release(x_17, 6); - lean_ctor_release(x_17, 7); - x_87 = x_17; -} else { - lean_dec_ref(x_17); - x_87 = lean_box(0); -} -x_88 = lean_ctor_get_uint64(x_18, sizeof(void*)*1); -x_89 = lean_ctor_get(x_18, 0); -lean_inc(x_89); -if (lean_is_exclusive(x_18)) { - lean_ctor_release(x_18, 0); - x_90 = x_18; -} else { - lean_dec_ref(x_18); - x_90 = lean_box(0); -} -x_91 = l_Lean_addTrace___at_Lean_Meta_Grind_mkEqProofImpl_go___spec__2___closed__1; -x_92 = 0; -x_93 = l_Lean_addTrace___at_Lean_Meta_Grind_mkEqProofImpl_go___spec__2___closed__2; -x_94 = lean_alloc_ctor(0, 2, 17); -lean_ctor_set(x_94, 0, x_1); -lean_ctor_set(x_94, 1, x_93); -lean_ctor_set_float(x_94, sizeof(void*)*2, x_91); -lean_ctor_set_float(x_94, sizeof(void*)*2 + 8, x_91); -lean_ctor_set_uint8(x_94, sizeof(void*)*2 + 16, x_92); -x_95 = l_Lean_addTrace___at_Lean_Meta_Grind_mkEqProofImpl_go___spec__2___closed__3; -x_96 = lean_alloc_ctor(9, 3, 0); -lean_ctor_set(x_96, 0, x_94); -lean_ctor_set(x_96, 1, x_14); -lean_ctor_set(x_96, 2, x_95); -lean_inc(x_12); -x_97 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_97, 0, x_12); -lean_ctor_set(x_97, 1, x_96); -x_98 = l_Lean_PersistentArray_push___rarg(x_89, x_97); -if (lean_is_scalar(x_90)) { - x_99 = lean_alloc_ctor(0, 1, 8); -} else { - x_99 = x_90; -} -lean_ctor_set(x_99, 0, x_98); -lean_ctor_set_uint64(x_99, sizeof(void*)*1, x_88); -if (lean_is_scalar(x_87)) { - x_100 = lean_alloc_ctor(0, 8, 0); -} else { - x_100 = x_87; -} -lean_ctor_set(x_100, 0, x_80); -lean_ctor_set(x_100, 1, x_81); -lean_ctor_set(x_100, 2, x_82); -lean_ctor_set(x_100, 3, x_99); -lean_ctor_set(x_100, 4, x_83); -lean_ctor_set(x_100, 5, x_84); -lean_ctor_set(x_100, 6, x_85); -lean_ctor_set(x_100, 7, x_86); -x_101 = lean_st_ref_set(x_10, x_100, x_79); -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; -} else { - lean_dec_ref(x_101); - x_103 = lean_box(0); -} -x_104 = lean_box(0); -if (lean_is_scalar(x_103)) { - x_105 = lean_alloc_ctor(0, 2, 0); -} else { - x_105 = x_103; -} -lean_ctor_set(x_105, 0, x_104); -lean_ctor_set(x_105, 1, x_102); -return x_105; -} -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEqProofImpl_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, lean_object* x_11, lean_object* x_12) { -_start: -{ -uint8_t x_13; lean_object* x_14; -x_13 = 0; -x_14 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore(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_Meta_Grind_mkEqProofImpl_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_object* x_11, lean_object* x_12) { -_start: -{ -uint8_t x_13; lean_object* x_14; -x_13 = 1; -x_14 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore(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_Meta_Grind_mkEqProofImpl_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_object* x_11, lean_object* x_12) { -_start: -{ -uint8_t x_13; lean_object* x_14; -x_13 = 1; -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -x_14 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore(x_1, x_2, x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_14) == 0) -{ -lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = lean_ctor_get(x_14, 0); -lean_inc(x_15); -x_16 = lean_ctor_get(x_14, 1); -lean_inc(x_16); -lean_dec(x_14); -x_17 = l_Lean_Meta_mkEqOfHEq(x_15, x_8, x_9, x_10, x_11, x_16); -return x_17; -} -else -{ -uint8_t x_18; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -x_18 = !lean_is_exclusive(x_14); -if (x_18 == 0) -{ -return x_14; -} -else -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_19 = lean_ctor_get(x_14, 0); -x_20 = lean_ctor_get(x_14, 1); -lean_inc(x_20); -lean_inc(x_19); -lean_dec(x_14); -x_21 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_21, 0, x_19); -lean_ctor_set(x_21, 1, x_20); -return x_21; -} -} -} -} -static lean_object* _init_l_Lean_Meta_Grind_mkEqProofImpl_go___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("grind", 5, 5); -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_Grind_mkEqProofImpl_go___closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("proof", 5, 5); -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_Grind_mkEqProofImpl_go___closed__3() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_Grind_mkEqProofImpl_go___closed__1; -x_2 = l_Lean_Meta_Grind_mkEqProofImpl_go___closed__2; -x_3 = l_Lean_Name_mkStr2(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Meta_Grind_mkEqProofImpl_go___closed__4() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_addTrace___at_Lean_Meta_Grind_mkEqProofImpl_go___spec__2___closed__2; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Meta_Grind_mkEqProofImpl_go___closed__5() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked(" = ", 3, 3); -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_Grind_mkEqProofImpl_go___closed__6() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_Grind_mkEqProofImpl_go___closed__5; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Meta_Grind_mkEqProofImpl_go___closed__7() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked(" ≡ ", 5, 3); -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_Grind_mkEqProofImpl_go___closed__8() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_Grind_mkEqProofImpl_go___closed__7; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEqProofImpl_go(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -lean_object* x_12; -lean_inc(x_1); -x_12 = l_Lean_Meta_Grind_getRootENode(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_12) == 0) -{ -lean_object* x_13; uint8_t x_14; -x_13 = lean_ctor_get(x_12, 0); -lean_inc(x_13); -x_14 = lean_ctor_get_uint8(x_13, sizeof(void*)*10 + 4); -lean_dec(x_13); -if (x_14 == 0) -{ -lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; -x_15 = lean_ctor_get(x_12, 1); -lean_inc(x_15); -lean_dec(x_12); -x_16 = l_Lean_Meta_Grind_mkEqProofImpl_go___closed__3; -x_17 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_mkEqProofImpl_go___spec__1(x_16, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_15); -x_18 = lean_ctor_get(x_17, 0); -lean_inc(x_18); -x_19 = lean_unbox(x_18); -lean_dec(x_18); -if (x_19 == 0) -{ -lean_object* x_20; uint8_t x_21; lean_object* x_22; -x_20 = lean_ctor_get(x_17, 1); -lean_inc(x_20); -lean_dec(x_17); -x_21 = 0; -x_22 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore(x_1, x_2, x_21, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_20); -return x_22; -} -else -{ -uint8_t x_23; -x_23 = !lean_is_exclusive(x_17); -if (x_23 == 0) -{ -lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; lean_object* x_36; -x_24 = lean_ctor_get(x_17, 1); -x_25 = lean_ctor_get(x_17, 0); -lean_dec(x_25); -lean_inc(x_1); -x_26 = l_Lean_MessageData_ofExpr(x_1); -x_27 = l_Lean_Meta_Grind_mkEqProofImpl_go___closed__4; -lean_ctor_set_tag(x_17, 7); -lean_ctor_set(x_17, 1, x_26); -lean_ctor_set(x_17, 0, x_27); -x_28 = l_Lean_Meta_Grind_mkEqProofImpl_go___closed__6; -x_29 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_29, 0, x_17); -lean_ctor_set(x_29, 1, x_28); -lean_inc(x_2); -x_30 = l_Lean_MessageData_ofExpr(x_2); -x_31 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_31, 0, x_29); -lean_ctor_set(x_31, 1, x_30); -x_32 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_32, 0, x_31); -lean_ctor_set(x_32, 1, x_27); -x_33 = l_Lean_addTrace___at_Lean_Meta_Grind_mkEqProofImpl_go___spec__2(x_16, x_32, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_24); -x_34 = lean_ctor_get(x_33, 1); -lean_inc(x_34); -lean_dec(x_33); -x_35 = 0; -x_36 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore(x_1, x_2, x_35, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, 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; uint8_t x_48; lean_object* x_49; -x_37 = lean_ctor_get(x_17, 1); -lean_inc(x_37); -lean_dec(x_17); -lean_inc(x_1); -x_38 = l_Lean_MessageData_ofExpr(x_1); -x_39 = l_Lean_Meta_Grind_mkEqProofImpl_go___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_Lean_Meta_Grind_mkEqProofImpl_go___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); -lean_inc(x_2); -x_43 = l_Lean_MessageData_ofExpr(x_2); -x_44 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_44, 0, x_42); -lean_ctor_set(x_44, 1, x_43); -x_45 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_45, 0, x_44); -lean_ctor_set(x_45, 1, x_39); -x_46 = l_Lean_addTrace___at_Lean_Meta_Grind_mkEqProofImpl_go___spec__2(x_16, x_45, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_37); -x_47 = lean_ctor_get(x_46, 1); -lean_inc(x_47); -lean_dec(x_46); -x_48 = 0; -x_49 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore(x_1, x_2, x_48, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_47); -return x_49; -} -} -} -else -{ -lean_object* x_50; lean_object* x_51; -x_50 = lean_ctor_get(x_12, 1); -lean_inc(x_50); -lean_dec(x_12); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_2); -lean_inc(x_1); -x_51 = l_Lean_Meta_Grind_hasSameType(x_1, x_2, x_7, x_8, x_9, x_10, x_50); -if (lean_obj_tag(x_51) == 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_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; uint8_t x_58; -x_54 = lean_ctor_get(x_51, 1); -lean_inc(x_54); -lean_dec(x_51); -x_55 = l_Lean_Meta_Grind_mkEqProofImpl_go___closed__3; -x_56 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_mkEqProofImpl_go___spec__1(x_55, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_54); -x_57 = lean_ctor_get(x_56, 0); -lean_inc(x_57); -x_58 = lean_unbox(x_57); -lean_dec(x_57); -if (x_58 == 0) -{ -lean_object* x_59; uint8_t x_60; lean_object* x_61; -x_59 = lean_ctor_get(x_56, 1); -lean_inc(x_59); -lean_dec(x_56); -x_60 = 1; -x_61 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore(x_1, x_2, x_60, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_59); -return x_61; -} -else -{ -uint8_t x_62; -x_62 = !lean_is_exclusive(x_56); -if (x_62 == 0) -{ -lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; uint8_t x_74; lean_object* x_75; -x_63 = lean_ctor_get(x_56, 1); -x_64 = lean_ctor_get(x_56, 0); -lean_dec(x_64); -lean_inc(x_1); -x_65 = l_Lean_MessageData_ofExpr(x_1); -x_66 = l_Lean_Meta_Grind_mkEqProofImpl_go___closed__4; -lean_ctor_set_tag(x_56, 7); -lean_ctor_set(x_56, 1, x_65); -lean_ctor_set(x_56, 0, x_66); -x_67 = l_Lean_Meta_Grind_mkEqProofImpl_go___closed__8; -x_68 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_68, 0, x_56); -lean_ctor_set(x_68, 1, x_67); -lean_inc(x_2); -x_69 = l_Lean_MessageData_ofExpr(x_2); -x_70 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_70, 0, x_68); -lean_ctor_set(x_70, 1, x_69); -x_71 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_71, 0, x_70); -lean_ctor_set(x_71, 1, x_66); -x_72 = l_Lean_addTrace___at_Lean_Meta_Grind_mkEqProofImpl_go___spec__2(x_55, x_71, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_63); -x_73 = lean_ctor_get(x_72, 1); -lean_inc(x_73); -lean_dec(x_72); -x_74 = 1; -x_75 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore(x_1, x_2, x_74, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_73); -return x_75; -} -else -{ -lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; uint8_t x_87; lean_object* x_88; -x_76 = lean_ctor_get(x_56, 1); -lean_inc(x_76); -lean_dec(x_56); -lean_inc(x_1); -x_77 = l_Lean_MessageData_ofExpr(x_1); -x_78 = l_Lean_Meta_Grind_mkEqProofImpl_go___closed__4; -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_Meta_Grind_mkEqProofImpl_go___closed__8; -x_81 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_81, 0, x_79); -lean_ctor_set(x_81, 1, x_80); -lean_inc(x_2); -x_82 = l_Lean_MessageData_ofExpr(x_2); -x_83 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_83, 0, x_81); -lean_ctor_set(x_83, 1, x_82); -x_84 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_84, 0, x_83); -lean_ctor_set(x_84, 1, x_78); -x_85 = l_Lean_addTrace___at_Lean_Meta_Grind_mkEqProofImpl_go___spec__2(x_55, x_84, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_76); -x_86 = lean_ctor_get(x_85, 1); -lean_inc(x_86); -lean_dec(x_85); -x_87 = 1; -x_88 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore(x_1, x_2, x_87, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_86); -return x_88; -} -} -} -else -{ -lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; uint8_t x_93; -x_89 = lean_ctor_get(x_51, 1); -lean_inc(x_89); -lean_dec(x_51); -x_90 = l_Lean_Meta_Grind_mkEqProofImpl_go___closed__3; -x_91 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_mkEqProofImpl_go___spec__1(x_90, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_89); -x_92 = lean_ctor_get(x_91, 0); -lean_inc(x_92); -x_93 = lean_unbox(x_92); -lean_dec(x_92); -if (x_93 == 0) -{ -lean_object* x_94; lean_object* x_95; lean_object* x_96; -x_94 = lean_ctor_get(x_91, 1); -lean_inc(x_94); -lean_dec(x_91); -x_95 = lean_box(0); -x_96 = l_Lean_Meta_Grind_mkEqProofImpl_go___lambda__3(x_1, x_2, x_95, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_94); -return x_96; -} -else -{ -uint8_t x_97; -x_97 = !lean_is_exclusive(x_91); -if (x_97 == 0) -{ -lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; 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_98 = lean_ctor_get(x_91, 1); -x_99 = lean_ctor_get(x_91, 0); -lean_dec(x_99); -lean_inc(x_1); -x_100 = l_Lean_MessageData_ofExpr(x_1); -x_101 = l_Lean_Meta_Grind_mkEqProofImpl_go___closed__4; -lean_ctor_set_tag(x_91, 7); -lean_ctor_set(x_91, 1, x_100); -lean_ctor_set(x_91, 0, x_101); -x_102 = l_Lean_Meta_Grind_mkEqProofImpl_go___closed__6; -x_103 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_103, 0, x_91); -lean_ctor_set(x_103, 1, x_102); -lean_inc(x_2); -x_104 = l_Lean_MessageData_ofExpr(x_2); -x_105 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_105, 0, x_103); -lean_ctor_set(x_105, 1, x_104); -x_106 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_106, 0, x_105); -lean_ctor_set(x_106, 1, x_101); -x_107 = l_Lean_addTrace___at_Lean_Meta_Grind_mkEqProofImpl_go___spec__2(x_90, x_106, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_98); -x_108 = lean_ctor_get(x_107, 0); -lean_inc(x_108); -x_109 = lean_ctor_get(x_107, 1); -lean_inc(x_109); -lean_dec(x_107); -x_110 = l_Lean_Meta_Grind_mkEqProofImpl_go___lambda__3(x_1, x_2, x_108, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_109); -lean_dec(x_108); -return x_110; -} -else -{ -lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; -x_111 = lean_ctor_get(x_91, 1); -lean_inc(x_111); -lean_dec(x_91); -lean_inc(x_1); -x_112 = l_Lean_MessageData_ofExpr(x_1); -x_113 = l_Lean_Meta_Grind_mkEqProofImpl_go___closed__4; -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_Meta_Grind_mkEqProofImpl_go___closed__6; -x_116 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_116, 0, x_114); -lean_ctor_set(x_116, 1, x_115); -lean_inc(x_2); -x_117 = l_Lean_MessageData_ofExpr(x_2); -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 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_119, 0, x_118); -lean_ctor_set(x_119, 1, x_113); -x_120 = l_Lean_addTrace___at_Lean_Meta_Grind_mkEqProofImpl_go___spec__2(x_90, x_119, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_111); -x_121 = lean_ctor_get(x_120, 0); -lean_inc(x_121); -x_122 = lean_ctor_get(x_120, 1); -lean_inc(x_122); -lean_dec(x_120); -x_123 = l_Lean_Meta_Grind_mkEqProofImpl_go___lambda__3(x_1, x_2, x_121, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_122); -lean_dec(x_121); -return x_123; -} -} -} -} -else -{ -uint8_t x_124; -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_124 = !lean_is_exclusive(x_51); -if (x_124 == 0) -{ -return x_51; -} -else -{ -lean_object* x_125; lean_object* x_126; lean_object* x_127; -x_125 = lean_ctor_get(x_51, 0); -x_126 = lean_ctor_get(x_51, 1); -lean_inc(x_126); -lean_inc(x_125); -lean_dec(x_51); -x_127 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_127, 0, x_125); -lean_ctor_set(x_127, 1, x_126); -return x_127; -} -} -} -} -else -{ -uint8_t x_128; -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_128 = !lean_is_exclusive(x_12); -if (x_128 == 0) -{ -return x_12; -} -else -{ -lean_object* x_129; lean_object* x_130; lean_object* x_131; -x_129 = lean_ctor_get(x_12, 0); -x_130 = lean_ctor_get(x_12, 1); -lean_inc(x_130); -lean_inc(x_129); -lean_dec(x_12); -x_131 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_131, 0, x_129); -lean_ctor_set(x_131, 1, x_130); -return x_131; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_mkEqProofImpl_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, 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_isTracingEnabledFor___at_Lean_Meta_Grind_mkEqProofImpl_go___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_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_11; -} -} -LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Meta_Grind_mkEqProofImpl_go___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -lean_object* x_12; -x_12 = l_Lean_addTrace___at_Lean_Meta_Grind_mkEqProofImpl_go___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -return x_12; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEqProofImpl_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, lean_object* x_11, lean_object* x_12) { -_start: -{ -lean_object* x_13; -x_13 = l_Lean_Meta_Grind_mkEqProofImpl_go___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_3); -return x_13; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEqProofImpl_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, lean_object* x_12) { -_start: -{ -lean_object* x_13; -x_13 = l_Lean_Meta_Grind_mkEqProofImpl_go___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_3); -return x_13; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEqProofImpl_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, lean_object* x_12) { -_start: -{ -lean_object* x_13; -x_13 = l_Lean_Meta_Grind_mkEqProofImpl_go___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_3); -return x_13; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEqProofImpl___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; -x_12 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_12, 0, x_1); -lean_ctor_set(x_12, 1, x_11); -return x_12; -} -} -static lean_object* _init_l_Lean_Meta_Grind_mkEqProofImpl___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("detail", 6, 6); -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_Grind_mkEqProofImpl___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Meta_Grind_mkEqProofImpl_go___closed__1; -x_2 = l_Lean_Meta_Grind_mkEqProofImpl_go___closed__2; -x_3 = l_Lean_Meta_Grind_mkEqProofImpl___closed__1; -x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); -return x_4; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__7___lambda__1___closed__1; +x_2 = l_Lean_Meta_Grind_mkEqProofImpl___closed__3; +x_3 = lean_unsigned_to_nat(231u); +x_4 = lean_unsigned_to_nat(2u); +x_5 = l_Lean_Meta_Grind_mkEqProofImpl___closed__2; +x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); +return x_6; } } LEAN_EXPORT lean_object* lean_grind_mk_eq_proof(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { @@ -10211,106 +9339,42 @@ lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -x_12 = l_Lean_Meta_Grind_mkEqProofImpl_go(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_inc(x_2); +lean_inc(x_1); +x_12 = l_Lean_Meta_Grind_hasSameType(x_1, x_2, x_7, x_8, x_9, x_10, x_11); if (lean_obj_tag(x_12) == 0) { -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +lean_object* x_13; uint8_t x_14; x_13 = lean_ctor_get(x_12, 0); lean_inc(x_13); -x_14 = lean_ctor_get(x_12, 1); -lean_inc(x_14); -lean_dec(x_12); -x_15 = l_Lean_Meta_Grind_mkEqProofImpl___closed__2; -x_16 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_mkEqProofImpl_go___spec__1(x_15, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_14); -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -x_18 = lean_unbox(x_17); -lean_dec(x_17); -if (x_18 == 0) -{ -uint8_t x_19; -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_19 = !lean_is_exclusive(x_16); -if (x_19 == 0) -{ -lean_object* x_20; -x_20 = lean_ctor_get(x_16, 0); -lean_dec(x_20); -lean_ctor_set(x_16, 0, x_13); -return x_16; -} -else -{ -lean_object* x_21; lean_object* x_22; -x_21 = lean_ctor_get(x_16, 1); -lean_inc(x_21); -lean_dec(x_16); -x_22 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_22, 0, x_13); -lean_ctor_set(x_22, 1, x_21); -return x_22; -} -} -else -{ -lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; -x_23 = lean_ctor_get(x_16, 1); -lean_inc(x_23); -lean_dec(x_16); -lean_inc(x_13); -x_24 = l_Lean_MessageData_ofExpr(x_13); -x_25 = l_Lean_Meta_Grind_mkEqProofImpl_go___closed__4; -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 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_27, 0, x_26); -lean_ctor_set(x_27, 1, x_25); -x_28 = l_Lean_addTrace___at_Lean_Meta_Grind_mkEqProofImpl_go___spec__2(x_15, x_27, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_23); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -x_29 = !lean_is_exclusive(x_28); -if (x_29 == 0) +x_14 = lean_unbox(x_13); +lean_dec(x_13); +if (x_14 == 0) { -lean_object* x_30; -x_30 = lean_ctor_get(x_28, 0); -lean_dec(x_30); -lean_ctor_set(x_28, 0, x_13); -return x_28; +lean_object* x_15; lean_object* x_16; lean_object* x_17; +lean_dec(x_2); +lean_dec(x_1); +x_15 = lean_ctor_get(x_12, 1); +lean_inc(x_15); +lean_dec(x_12); +x_16 = l_Lean_Meta_Grind_mkEqProofImpl___closed__4; +x_17 = l_panic___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__9(x_16, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_15); +return x_17; } else { -lean_object* x_31; lean_object* x_32; -x_31 = lean_ctor_get(x_28, 1); -lean_inc(x_31); -lean_dec(x_28); -x_32 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_32, 0, x_13); -lean_ctor_set(x_32, 1, x_31); -return x_32; -} +lean_object* x_18; uint8_t x_19; lean_object* x_20; +x_18 = lean_ctor_get(x_12, 1); +lean_inc(x_18); +lean_dec(x_12); +x_19 = 0; +x_20 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore(x_1, x_2, x_19, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_18); +return x_20; } } else { -uint8_t x_33; +uint8_t x_21; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -10319,45 +9383,30 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_33 = !lean_is_exclusive(x_12); -if (x_33 == 0) +lean_dec(x_2); +lean_dec(x_1); +x_21 = !lean_is_exclusive(x_12); +if (x_21 == 0) { return x_12; } else { -lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_34 = lean_ctor_get(x_12, 0); -x_35 = lean_ctor_get(x_12, 1); -lean_inc(x_35); -lean_inc(x_34); +lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_22 = lean_ctor_get(x_12, 0); +x_23 = lean_ctor_get(x_12, 1); +lean_inc(x_23); +lean_inc(x_22); lean_dec(x_12); -x_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_24 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_24, 0, x_22); +lean_ctor_set(x_24, 1, x_23); +return x_24; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEqProofImpl___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_Meta_Grind_mkEqProofImpl___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_12; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkHEqProof(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_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* lean_grind_mk_heq_proof(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { uint8_t x_12; lean_object* x_13; @@ -10475,31 +9524,14 @@ l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkProofTo___lambda__ lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkProofTo___lambda__1___closed__2); l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkProofTo___lambda__1___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkProofTo___lambda__1___closed__3(); lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkProofTo___lambda__1___closed__3); -l_Lean_addTrace___at_Lean_Meta_Grind_mkEqProofImpl_go___spec__2___closed__1 = _init_l_Lean_addTrace___at_Lean_Meta_Grind_mkEqProofImpl_go___spec__2___closed__1(); -l_Lean_addTrace___at_Lean_Meta_Grind_mkEqProofImpl_go___spec__2___closed__2 = _init_l_Lean_addTrace___at_Lean_Meta_Grind_mkEqProofImpl_go___spec__2___closed__2(); -lean_mark_persistent(l_Lean_addTrace___at_Lean_Meta_Grind_mkEqProofImpl_go___spec__2___closed__2); -l_Lean_addTrace___at_Lean_Meta_Grind_mkEqProofImpl_go___spec__2___closed__3 = _init_l_Lean_addTrace___at_Lean_Meta_Grind_mkEqProofImpl_go___spec__2___closed__3(); -lean_mark_persistent(l_Lean_addTrace___at_Lean_Meta_Grind_mkEqProofImpl_go___spec__2___closed__3); -l_Lean_Meta_Grind_mkEqProofImpl_go___closed__1 = _init_l_Lean_Meta_Grind_mkEqProofImpl_go___closed__1(); -lean_mark_persistent(l_Lean_Meta_Grind_mkEqProofImpl_go___closed__1); -l_Lean_Meta_Grind_mkEqProofImpl_go___closed__2 = _init_l_Lean_Meta_Grind_mkEqProofImpl_go___closed__2(); -lean_mark_persistent(l_Lean_Meta_Grind_mkEqProofImpl_go___closed__2); -l_Lean_Meta_Grind_mkEqProofImpl_go___closed__3 = _init_l_Lean_Meta_Grind_mkEqProofImpl_go___closed__3(); -lean_mark_persistent(l_Lean_Meta_Grind_mkEqProofImpl_go___closed__3); -l_Lean_Meta_Grind_mkEqProofImpl_go___closed__4 = _init_l_Lean_Meta_Grind_mkEqProofImpl_go___closed__4(); -lean_mark_persistent(l_Lean_Meta_Grind_mkEqProofImpl_go___closed__4); -l_Lean_Meta_Grind_mkEqProofImpl_go___closed__5 = _init_l_Lean_Meta_Grind_mkEqProofImpl_go___closed__5(); -lean_mark_persistent(l_Lean_Meta_Grind_mkEqProofImpl_go___closed__5); -l_Lean_Meta_Grind_mkEqProofImpl_go___closed__6 = _init_l_Lean_Meta_Grind_mkEqProofImpl_go___closed__6(); -lean_mark_persistent(l_Lean_Meta_Grind_mkEqProofImpl_go___closed__6); -l_Lean_Meta_Grind_mkEqProofImpl_go___closed__7 = _init_l_Lean_Meta_Grind_mkEqProofImpl_go___closed__7(); -lean_mark_persistent(l_Lean_Meta_Grind_mkEqProofImpl_go___closed__7); -l_Lean_Meta_Grind_mkEqProofImpl_go___closed__8 = _init_l_Lean_Meta_Grind_mkEqProofImpl_go___closed__8(); -lean_mark_persistent(l_Lean_Meta_Grind_mkEqProofImpl_go___closed__8); l_Lean_Meta_Grind_mkEqProofImpl___closed__1 = _init_l_Lean_Meta_Grind_mkEqProofImpl___closed__1(); lean_mark_persistent(l_Lean_Meta_Grind_mkEqProofImpl___closed__1); l_Lean_Meta_Grind_mkEqProofImpl___closed__2 = _init_l_Lean_Meta_Grind_mkEqProofImpl___closed__2(); lean_mark_persistent(l_Lean_Meta_Grind_mkEqProofImpl___closed__2); +l_Lean_Meta_Grind_mkEqProofImpl___closed__3 = _init_l_Lean_Meta_Grind_mkEqProofImpl___closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind_mkEqProofImpl___closed__3); +l_Lean_Meta_Grind_mkEqProofImpl___closed__4 = _init_l_Lean_Meta_Grind_mkEqProofImpl___closed__4(); +lean_mark_persistent(l_Lean_Meta_Grind_mkEqProofImpl___closed__4); return lean_io_result_mk_ok(lean_box(0)); } #ifdef __cplusplus diff --git a/stage0/stdlib/Lean/Meta/Tactic/Grind/Propagate.c b/stage0/stdlib/Lean/Meta/Tactic/Grind/Propagate.c index d4c75350fc17..80e456e2dbda 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Grind/Propagate.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Grind/Propagate.c @@ -29,7 +29,6 @@ static lean_object* l_Lean_Meta_Grind_propagateAndUp___lambda__1___closed__5; static lean_object* l___regBuiltin_Lean_Meta_Grind_propagateNotUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_1426____closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Meta_Grind_propagateAndUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_383_(lean_object*); static lean_object* l_Lean_Meta_Grind_propagateOrUp___lambda__1___closed__10; -lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_propagateAndDown___lambda__1___closed__5; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateEqUp(lean_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_Grind_propagateEqDown___lambda__1___closed__2; @@ -118,6 +117,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateOrDown___lambda__1(lean_obje static lean_object* l_Lean_Meta_Grind_propagateNotUp___lambda__1___closed__4; lean_object* lean_grind_mk_eq_proof(lean_object*, lean_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_Grind_propagateEqUp___closed__1; +lean_object* lean_grind_mk_heq_proof(lean_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_mkApp3(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_mkEqFalseProof(lean_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_Grind_propagateOrUp___closed__1; @@ -5137,11 +5137,10 @@ return x_21; } else { -lean_object* x_22; uint8_t x_23; lean_object* x_24; +lean_object* x_22; lean_object* x_23; x_22 = lean_ctor_get(x_13, 1); lean_inc(x_22); lean_dec(x_13); -x_23 = 1; lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); @@ -5150,19 +5149,19 @@ lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -x_24 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore(x_2, x_3, x_23, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_22); -if (lean_obj_tag(x_24) == 0) +x_23 = lean_grind_mk_heq_proof(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_22); +if (lean_obj_tag(x_23) == 0) { -lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_25 = lean_ctor_get(x_24, 0); +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +x_25 = lean_ctor_get(x_23, 1); lean_inc(x_25); -x_26 = lean_ctor_get(x_24, 1); -lean_inc(x_26); -lean_dec(x_24); -x_27 = l_Lean_Meta_Grind_propagateEqUp___lambda__1___closed__3; +lean_dec(x_23); +x_26 = l_Lean_Meta_Grind_propagateEqUp___lambda__1___closed__3; lean_inc(x_1); -x_28 = l_Lean_mkAppB(x_27, x_1, x_25); -x_29 = l_Lean_Meta_Grind_pushEqTrue(x_1, x_28, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_26); +x_27 = l_Lean_mkAppB(x_26, x_1, x_24); +x_28 = l_Lean_Meta_Grind_pushEqTrue(x_1, x_27, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_25); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -5171,11 +5170,11 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -return x_29; +return x_28; } else { -uint8_t x_30; +uint8_t x_29; lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -5185,30 +5184,30 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_1); -x_30 = !lean_is_exclusive(x_24); -if (x_30 == 0) +x_29 = !lean_is_exclusive(x_23); +if (x_29 == 0) { -return x_24; +return x_23; } else { -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_object* x_30; lean_object* x_31; lean_object* x_32; +x_30 = lean_ctor_get(x_23, 0); +x_31 = lean_ctor_get(x_23, 1); lean_inc(x_31); -lean_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; +lean_inc(x_30); +lean_dec(x_23); +x_32 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_32, 0, x_30); +lean_ctor_set(x_32, 1, x_31); +return x_32; } } } } else { -uint8_t x_34; +uint8_t x_33; lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -5220,23 +5219,23 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_34 = !lean_is_exclusive(x_13); -if (x_34 == 0) +x_33 = !lean_is_exclusive(x_13); +if (x_33 == 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_inc(x_36); +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_13, 0); +x_35 = lean_ctor_get(x_13, 1); lean_inc(x_35); +lean_inc(x_34); 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_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; } } } diff --git a/stage0/stdlib/Lean/Meta/Tactic/Grind/Run.c b/stage0/stdlib/Lean/Meta/Tactic/Grind/Run.c deleted file mode 100644 index 2aaf380b50aa..000000000000 --- a/stage0/stdlib/Lean/Meta/Tactic/Grind/Run.c +++ /dev/null @@ -1,1827 +0,0 @@ -// Lean compiler output -// Module: Lean.Meta.Tactic.Grind.Run -// Imports: Init.Grind.Lemmas Lean.Meta.Tactic.Grind.Types Lean.Meta.Tactic.Grind.PropagatorAttr Lean.Meta.Tactic.Grind.Proj -#include -#if defined(__clang__) -#pragma clang diagnostic ignored "-Wunused-parameter" -#pragma clang diagnostic ignored "-Wunused-label" -#elif defined(__GNUC__) && !defined(__CLANG__) -#pragma GCC diagnostic ignored "-Wunused-parameter" -#pragma GCC diagnostic ignored "-Wunused-label" -#pragma GCC diagnostic ignored "-Wunused-but-set-variable" -#endif -#ifdef __cplusplus -extern "C" { -#endif -lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GoalM_run___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_Grind_mkMethods(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_GrindM_run___rarg___closed__14; -lean_object* lean_mk_empty_array_with_capacity(lean_object*); -LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___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_Meta_Grind_GrindM_run___rarg___closed__20; -static lean_object* l_Lean_Meta_Grind_GrindM_run___rarg___closed__11; -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GoalM_run(lean_object*); -size_t lean_uint64_to_usize(uint64_t); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GoalM_run_x27(lean_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_Grind_GrindM_run___rarg___closed__5; -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_mkMethods___spec__1___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2(lean_object*); -lean_object* l_Lean_Meta_getSimpCongrTheorems___rarg(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_GrindM_run___rarg___closed__1; -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GoalM_run___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_Grind_getFalseExpr___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -size_t lean_usize_of_nat(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GrindM_run(lean_object*); -static lean_object* l_Lean_Meta_Grind_GrindM_run___rarg___closed__10; -uint64_t lean_uint64_shift_right(uint64_t, uint64_t); -lean_object* l_Lean_Meta_Simp_SimprocExtension_getSimprocs(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkMethods___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* l_Lean_Meta_Grind_propagateProjEq(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_Grind_getTrueExpr___rarg(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_Meta_Grind_GrindM_run___rarg___closed__13; -lean_object* lean_st_mk_ref(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_Meta_Grind_GrindM_run___rarg___closed__15; -lean_object* l_ShareCommon_mkStateImpl(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkMethods___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 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_Grind_GrindM_run___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkGoal___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Meta_Simp_mkContext(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -extern lean_object* l_Lean_Meta_Grind_builtinPropagatorsRef; -static lean_object* l_Lean_Meta_Grind_GrindM_run___rarg___closed__18; -static lean_object* l_Lean_Meta_Grind_GrindM_run___rarg___closed__7; -static lean_object* l_Lean_Meta_Grind_GrindM_run___rarg___closed__12; -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkMethods___rarg(lean_object*); -LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkGoal___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_Meta_Grind_grindNormSimprocExt; -lean_object* l_Lean_Meta_SimpExtension_getTheorems(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GoalM_run___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_GrindM_run___rarg___closed__3; -static lean_object* l_Lean_Meta_Grind_GoalM_run_x27___closed__1; -static lean_object* l_Lean_Meta_Grind_GrindM_run___rarg___closed__6; -LEAN_EXPORT lean_object* l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___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_Meta_Grind_mkMethods___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GoalM_run_x27___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_Grind_GrindM_run___rarg___closed__19; -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_mkGoal___spec__1; -uint64_t l_Lean_Name_hash___override(lean_object*); -extern lean_object* l_Lean_Meta_Grind_grindNormExt; -uint64_t lean_uint64_xor(uint64_t, uint64_t); -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_mkMethods___spec__1(lean_object*, lean_object*); -lean_object* l_Lean_Expr_getAppFn(lean_object*); -static lean_object* l_Lean_Meta_Grind_GrindM_run___rarg___closed__17; -lean_object* l_Lean_Meta_Grind_mkENodeCore(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* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GoalM_run_x27___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -size_t lean_usize_sub(size_t, size_t); -lean_object* lean_array_mk(lean_object*); -static lean_object* l_Lean_Meta_Grind_GrindM_run___rarg___closed__4; -extern lean_object* l_Lean_Meta_Simp_defaultMaxSteps; -static lean_object* l_Lean_Meta_Grind_GrindM_run___rarg___closed__8; -lean_object* lean_array_uget(lean_object*, size_t); -static lean_object* l_Lean_Meta_Grind_GrindM_run___rarg___closed__9; -lean_object* lean_array_get_size(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkGoal(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -extern lean_object* l_Lean_ShareCommon_objectFactory; -lean_object* lean_state_sharecommon(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_mkMethods___rarg___closed__1; -static lean_object* l_Lean_Meta_Grind_GrindM_run___rarg___closed__16; -static lean_object* l_Lean_Meta_Grind_GrindM_run___rarg___closed__2; -static lean_object* l_Lean_Meta_Grind_mkGoal___closed__1; -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GoalM_run_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_Lean_Meta_Grind_GoalM_run___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_Lean_Meta_Grind_mkMethods___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*); -size_t lean_usize_land(size_t, size_t); -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_mkMethods___spec__1(lean_object* x_1, lean_object* x_2) { -_start: -{ -if (lean_obj_tag(x_2) == 0) -{ -lean_object* x_3; -x_3 = lean_box(0); -return x_3; -} -else -{ -lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; -x_4 = lean_ctor_get(x_2, 0); -x_5 = lean_ctor_get(x_2, 1); -x_6 = lean_ctor_get(x_2, 2); -x_7 = lean_name_eq(x_4, x_1); -if (x_7 == 0) -{ -x_2 = x_6; -goto _start; -} -else -{ -lean_object* x_9; -lean_inc(x_5); -x_9 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_9, 0, x_5); -return x_9; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkMethods___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; -x_12 = l_Lean_Expr_getAppFn(x_2); -if (lean_obj_tag(x_12) == 4) -{ -lean_object* x_13; lean_object* x_14; -x_13 = lean_ctor_get(x_12, 0); -lean_inc(x_13); -lean_dec(x_12); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -lean_inc(x_2); -x_14 = l_Lean_Meta_Grind_propagateProjEq(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_14) == 0) -{ -uint8_t x_15; -x_15 = !lean_is_exclusive(x_14); -if (x_15 == 0) -{ -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint64_t x_21; uint64_t x_22; uint64_t x_23; uint64_t x_24; uint64_t x_25; uint64_t x_26; uint64_t x_27; size_t x_28; size_t x_29; size_t x_30; size_t x_31; size_t x_32; lean_object* x_33; lean_object* x_34; -x_16 = lean_ctor_get(x_1, 0); -x_17 = lean_ctor_get(x_14, 1); -x_18 = lean_ctor_get(x_14, 0); -lean_dec(x_18); -x_19 = lean_ctor_get(x_16, 1); -x_20 = lean_array_get_size(x_19); -x_21 = l_Lean_Name_hash___override(x_13); -x_22 = 32; -x_23 = lean_uint64_shift_right(x_21, x_22); -x_24 = lean_uint64_xor(x_21, x_23); -x_25 = 16; -x_26 = lean_uint64_shift_right(x_24, x_25); -x_27 = lean_uint64_xor(x_24, x_26); -x_28 = lean_uint64_to_usize(x_27); -x_29 = lean_usize_of_nat(x_20); -lean_dec(x_20); -x_30 = 1; -x_31 = lean_usize_sub(x_29, x_30); -x_32 = lean_usize_land(x_28, x_31); -x_33 = lean_array_uget(x_19, x_32); -x_34 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_mkMethods___spec__1(x_13, x_33); -lean_dec(x_33); -lean_dec(x_13); -if (lean_obj_tag(x_34) == 0) -{ -lean_object* x_35; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -x_35 = lean_box(0); -lean_ctor_set(x_14, 0, x_35); -return x_14; -} -else -{ -lean_object* x_36; lean_object* x_37; -lean_free_object(x_14); -x_36 = lean_ctor_get(x_34, 0); -lean_inc(x_36); -lean_dec(x_34); -x_37 = lean_apply_10(x_36, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_17); -return x_37; -} -} -else -{ -lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; uint64_t x_42; uint64_t x_43; uint64_t x_44; uint64_t x_45; uint64_t x_46; uint64_t x_47; uint64_t x_48; size_t x_49; size_t x_50; size_t x_51; size_t x_52; size_t x_53; lean_object* x_54; lean_object* x_55; -x_38 = lean_ctor_get(x_1, 0); -x_39 = lean_ctor_get(x_14, 1); -lean_inc(x_39); -lean_dec(x_14); -x_40 = lean_ctor_get(x_38, 1); -x_41 = lean_array_get_size(x_40); -x_42 = l_Lean_Name_hash___override(x_13); -x_43 = 32; -x_44 = lean_uint64_shift_right(x_42, x_43); -x_45 = lean_uint64_xor(x_42, x_44); -x_46 = 16; -x_47 = lean_uint64_shift_right(x_45, x_46); -x_48 = lean_uint64_xor(x_45, x_47); -x_49 = lean_uint64_to_usize(x_48); -x_50 = lean_usize_of_nat(x_41); -lean_dec(x_41); -x_51 = 1; -x_52 = lean_usize_sub(x_50, x_51); -x_53 = lean_usize_land(x_49, x_52); -x_54 = lean_array_uget(x_40, x_53); -x_55 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_mkMethods___spec__1(x_13, x_54); -lean_dec(x_54); -lean_dec(x_13); -if (lean_obj_tag(x_55) == 0) -{ -lean_object* x_56; lean_object* x_57; -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_box(0); -x_57 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_57, 0, x_56); -lean_ctor_set(x_57, 1, x_39); -return x_57; -} -else -{ -lean_object* x_58; lean_object* x_59; -x_58 = lean_ctor_get(x_55, 0); -lean_inc(x_58); -lean_dec(x_55); -x_59 = lean_apply_10(x_58, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_39); -return x_59; -} -} -} -else -{ -uint8_t x_60; -lean_dec(x_13); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -x_60 = !lean_is_exclusive(x_14); -if (x_60 == 0) -{ -return x_14; -} -else -{ -lean_object* x_61; lean_object* x_62; lean_object* x_63; -x_61 = lean_ctor_get(x_14, 0); -x_62 = lean_ctor_get(x_14, 1); -lean_inc(x_62); -lean_inc(x_61); -lean_dec(x_14); -x_63 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_63, 0, x_61); -lean_ctor_set(x_63, 1, x_62); -return x_63; -} -} -} -else -{ -lean_object* x_64; lean_object* x_65; -lean_dec(x_12); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -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_11); -return x_65; -} -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkMethods___rarg___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -lean_object* x_12; -x_12 = l_Lean_Expr_getAppFn(x_2); -if (lean_obj_tag(x_12) == 4) -{ -lean_object* x_13; lean_object* x_14; uint8_t x_15; -x_13 = lean_ctor_get(x_1, 1); -lean_inc(x_13); -lean_dec(x_1); -x_14 = lean_ctor_get(x_12, 0); -lean_inc(x_14); -lean_dec(x_12); -x_15 = !lean_is_exclusive(x_13); -if (x_15 == 0) -{ -lean_object* x_16; lean_object* x_17; lean_object* x_18; uint64_t x_19; uint64_t x_20; uint64_t x_21; uint64_t x_22; uint64_t x_23; uint64_t x_24; uint64_t x_25; size_t x_26; size_t x_27; size_t x_28; size_t x_29; size_t x_30; lean_object* x_31; lean_object* x_32; -x_16 = lean_ctor_get(x_13, 1); -x_17 = lean_ctor_get(x_13, 0); -lean_dec(x_17); -x_18 = lean_array_get_size(x_16); -x_19 = l_Lean_Name_hash___override(x_14); -x_20 = 32; -x_21 = lean_uint64_shift_right(x_19, x_20); -x_22 = lean_uint64_xor(x_19, x_21); -x_23 = 16; -x_24 = lean_uint64_shift_right(x_22, x_23); -x_25 = lean_uint64_xor(x_22, x_24); -x_26 = lean_uint64_to_usize(x_25); -x_27 = lean_usize_of_nat(x_18); -lean_dec(x_18); -x_28 = 1; -x_29 = lean_usize_sub(x_27, x_28); -x_30 = lean_usize_land(x_26, x_29); -x_31 = lean_array_uget(x_16, x_30); -lean_dec(x_16); -x_32 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_mkMethods___spec__1(x_14, x_31); -lean_dec(x_31); -lean_dec(x_14); -if (lean_obj_tag(x_32) == 0) -{ -lean_object* x_33; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -x_33 = lean_box(0); -lean_ctor_set(x_13, 1, x_11); -lean_ctor_set(x_13, 0, x_33); -return x_13; -} -else -{ -lean_object* x_34; lean_object* x_35; -lean_free_object(x_13); -x_34 = lean_ctor_get(x_32, 0); -lean_inc(x_34); -lean_dec(x_32); -x_35 = lean_apply_10(x_34, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_35; -} -} -else -{ -lean_object* x_36; lean_object* x_37; uint64_t x_38; uint64_t x_39; uint64_t x_40; uint64_t x_41; uint64_t x_42; uint64_t x_43; uint64_t x_44; size_t x_45; size_t x_46; size_t x_47; size_t x_48; size_t x_49; lean_object* x_50; lean_object* x_51; -x_36 = lean_ctor_get(x_13, 1); -lean_inc(x_36); -lean_dec(x_13); -x_37 = lean_array_get_size(x_36); -x_38 = l_Lean_Name_hash___override(x_14); -x_39 = 32; -x_40 = lean_uint64_shift_right(x_38, x_39); -x_41 = lean_uint64_xor(x_38, x_40); -x_42 = 16; -x_43 = lean_uint64_shift_right(x_41, x_42); -x_44 = lean_uint64_xor(x_41, x_43); -x_45 = lean_uint64_to_usize(x_44); -x_46 = lean_usize_of_nat(x_37); -lean_dec(x_37); -x_47 = 1; -x_48 = lean_usize_sub(x_46, x_47); -x_49 = lean_usize_land(x_45, x_48); -x_50 = lean_array_uget(x_36, x_49); -lean_dec(x_36); -x_51 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_mkMethods___spec__1(x_14, x_50); -lean_dec(x_50); -lean_dec(x_14); -if (lean_obj_tag(x_51) == 0) -{ -lean_object* x_52; lean_object* x_53; -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_52 = lean_box(0); -x_53 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_53, 0, x_52); -lean_ctor_set(x_53, 1, x_11); -return x_53; -} -else -{ -lean_object* x_54; lean_object* x_55; -x_54 = lean_ctor_get(x_51, 0); -lean_inc(x_54); -lean_dec(x_51); -x_55 = lean_apply_10(x_54, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_55; -} -} -} -else -{ -lean_object* x_56; lean_object* x_57; -lean_dec(x_12); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_56 = lean_box(0); -x_57 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_57, 0, x_56); -lean_ctor_set(x_57, 1, x_11); -return x_57; -} -} -} -static lean_object* _init_l_Lean_Meta_Grind_mkMethods___rarg___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = l_Lean_Meta_Grind_builtinPropagatorsRef; -return x_1; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkMethods___rarg(lean_object* x_1) { -_start: -{ -lean_object* x_2; lean_object* x_3; uint8_t x_4; -x_2 = l_Lean_Meta_Grind_mkMethods___rarg___closed__1; -x_3 = lean_st_ref_get(x_2, x_1); -x_4 = !lean_is_exclusive(x_3); -if (x_4 == 0) -{ -lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; -x_5 = lean_ctor_get(x_3, 0); -lean_inc(x_5); -x_6 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_mkMethods___rarg___lambda__1___boxed), 11, 1); -lean_closure_set(x_6, 0, x_5); -x_7 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_mkMethods___rarg___lambda__2), 11, 1); -lean_closure_set(x_7, 0, x_5); -x_8 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_8, 0, x_6); -lean_ctor_set(x_8, 1, x_7); -lean_ctor_set(x_3, 0, x_8); -return x_3; -} -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 = lean_ctor_get(x_3, 0); -x_10 = lean_ctor_get(x_3, 1); -lean_inc(x_10); -lean_inc(x_9); -lean_dec(x_3); -lean_inc(x_9); -x_11 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_mkMethods___rarg___lambda__1___boxed), 11, 1); -lean_closure_set(x_11, 0, x_9); -x_12 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_mkMethods___rarg___lambda__2), 11, 1); -lean_closure_set(x_12, 0, x_9); -x_13 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_13, 0, x_11); -lean_ctor_set(x_13, 1, x_12); -x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_13); -lean_ctor_set(x_14, 1, x_10); -return x_14; -} -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkMethods(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_mkMethods___rarg), 1, 0); -return x_3; -} -} -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_mkMethods___spec__1___boxed(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; -x_3 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_mkMethods___spec__1(x_1, x_2); -lean_dec(x_2); -lean_dec(x_1); -return x_3; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkMethods___rarg___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -lean_object* x_12; -x_12 = l_Lean_Meta_Grind_mkMethods___rarg___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_1); -return x_12; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkMethods___boxed(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; -x_3 = l_Lean_Meta_Grind_mkMethods(x_1, x_2); -lean_dec(x_2); -lean_dec(x_1); -return x_3; -} -} -static lean_object* _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_ShareCommon_objectFactory; -x_2 = l_ShareCommon_mkStateImpl(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("False", 5, 5); -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_Grind_GrindM_run___rarg___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_Grind_GrindM_run___rarg___closed__2; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Meta_Grind_GrindM_run___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_Meta_Grind_GrindM_run___rarg___closed__3; -x_3 = l_Lean_Expr_const___override(x_2, x_1); -return x_3; -} -} -static lean_object* _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__5() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_ShareCommon_objectFactory; -x_2 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__1; -x_3 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__4; -x_4 = lean_state_sharecommon(x_1, x_2, x_3); -return x_4; -} -} -static lean_object* _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__6() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("True", 4, 4); -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__7() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__6; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Meta_Grind_GrindM_run___rarg___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_Grind_GrindM_run___rarg___closed__7; -x_3 = l_Lean_Expr_const___override(x_2, x_1); -return x_3; -} -} -static lean_object* _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__9() { -_start: -{ -lean_object* x_1; -x_1 = l_Lean_Meta_Grind_grindNormExt; -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__10() { -_start: -{ -lean_object* x_1; -x_1 = l_Lean_Meta_Grind_grindNormSimprocExt; -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__11() { -_start: -{ -lean_object* x_1; lean_object* x_2; uint8_t x_3; uint8_t x_4; uint8_t x_5; lean_object* x_6; -x_1 = l_Lean_Meta_Simp_defaultMaxSteps; -x_2 = lean_unsigned_to_nat(2u); -x_3 = 0; -x_4 = 1; -x_5 = 0; -x_6 = lean_alloc_ctor(0, 2, 19); -lean_ctor_set(x_6, 0, x_1); -lean_ctor_set(x_6, 1, x_2); -lean_ctor_set_uint8(x_6, sizeof(void*)*2, x_3); -lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 1, x_4); -lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 2, x_3); -lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 3, x_4); -lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 4, x_4); -lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 5, x_4); -lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 6, x_5); -lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 7, x_4); -lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 8, x_4); -lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 9, x_3); -lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 10, x_4); -lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 11, x_3); -lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 12, x_4); -lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 13, x_4); -lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 14, x_3); -lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 15, x_3); -lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 16, x_3); -lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 17, x_4); -lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 18, x_4); -return x_6; -} -} -static lean_object* _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__12() { -_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_Grind_GrindM_run___rarg___closed__13() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__12; -x_2 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__14() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_Grind_GrindM_run___rarg___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; -} -} -static lean_object* _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__15() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__13; -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_Grind_GrindM_run___rarg___closed__16() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = lean_unsigned_to_nat(32u); -x_2 = lean_mk_empty_array_with_capacity(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__17() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__16; -x_2 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__18() { -_start: -{ -size_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = 5; -x_2 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__17; -x_3 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__16; -x_4 = lean_unsigned_to_nat(0u); -x_5 = lean_alloc_ctor(0, 4, sizeof(size_t)*1); -lean_ctor_set(x_5, 0, x_2); -lean_ctor_set(x_5, 1, x_3); -lean_ctor_set(x_5, 2, x_4); -lean_ctor_set(x_5, 3, x_4); -lean_ctor_set_usize(x_5, 4, x_1); -return x_5; -} -} -static lean_object* _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__19() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__13; -x_2 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__18; -x_3 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set(x_3, 1, x_1); -lean_ctor_set(x_3, 2, x_1); -lean_ctor_set(x_3, 3, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__20() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__15; -x_2 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__19; -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_Grind_GrindM_run___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: -{ -lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; 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_8 = lean_box(0); -x_9 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__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); -x_12 = l_Lean_ShareCommon_objectFactory; -x_13 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__8; -x_14 = lean_state_sharecommon(x_12, x_11, x_13); -x_15 = lean_ctor_get(x_14, 0); -lean_inc(x_15); -x_16 = lean_ctor_get(x_14, 1); -lean_inc(x_16); -lean_dec(x_14); -x_17 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__9; -x_18 = l_Lean_Meta_SimpExtension_getTheorems(x_17, x_5, x_6, x_7); -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_Meta_Grind_GrindM_run___rarg___closed__10; -x_22 = l_Lean_Meta_Simp_SimprocExtension_getSimprocs(x_21, x_5, x_6, x_20); -x_23 = !lean_is_exclusive(x_22); -if (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_22, 1); -lean_ctor_set_tag(x_22, 1); -lean_ctor_set(x_22, 1, x_8); -x_25 = lean_array_mk(x_22); -x_26 = l_Lean_Meta_getSimpCongrTheorems___rarg(x_6, x_24); -x_27 = !lean_is_exclusive(x_26); -if (x_27 == 0) -{ -lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_28 = lean_ctor_get(x_26, 0); -x_29 = lean_ctor_get(x_26, 1); -lean_ctor_set_tag(x_26, 1); -lean_ctor_set(x_26, 1, x_8); -lean_ctor_set(x_26, 0, x_19); -x_30 = lean_array_mk(x_26); -x_31 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__11; -x_32 = l_Lean_Meta_Simp_mkContext(x_31, x_30, x_28, x_3, x_4, x_5, x_6, x_29); -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_Grind_mkMethods___rarg(x_34); -x_36 = lean_ctor_get(x_35, 0); -lean_inc(x_36); -x_37 = lean_ctor_get(x_35, 1); -lean_inc(x_37); -lean_dec(x_35); -x_38 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_38, 0, x_33); -lean_ctor_set(x_38, 1, x_25); -lean_ctor_set(x_38, 2, x_2); -x_39 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__14; -x_40 = lean_unsigned_to_nat(1u); -x_41 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__13; -x_42 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__20; -x_43 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_43, 0, x_39); -lean_ctor_set(x_43, 1, x_16); -lean_ctor_set(x_43, 2, x_40); -lean_ctor_set(x_43, 3, x_41); -lean_ctor_set(x_43, 4, x_42); -lean_ctor_set(x_43, 5, x_15); -lean_ctor_set(x_43, 6, x_10); -x_44 = lean_st_mk_ref(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); -lean_dec(x_44); -lean_inc(x_45); -x_47 = lean_apply_8(x_1, x_36, x_38, x_45, x_3, x_4, x_5, x_6, x_46); -if (lean_obj_tag(x_47) == 0) -{ -lean_object* x_48; lean_object* x_49; lean_object* x_50; uint8_t x_51; -x_48 = lean_ctor_get(x_47, 0); -lean_inc(x_48); -x_49 = lean_ctor_get(x_47, 1); -lean_inc(x_49); -lean_dec(x_47); -x_50 = lean_st_ref_get(x_45, x_49); -lean_dec(x_45); -x_51 = !lean_is_exclusive(x_50); -if (x_51 == 0) -{ -lean_object* x_52; -x_52 = lean_ctor_get(x_50, 0); -lean_dec(x_52); -lean_ctor_set(x_50, 0, x_48); -return x_50; -} -else -{ -lean_object* x_53; lean_object* x_54; -x_53 = lean_ctor_get(x_50, 1); -lean_inc(x_53); -lean_dec(x_50); -x_54 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_54, 0, x_48); -lean_ctor_set(x_54, 1, x_53); -return x_54; -} -} -else -{ -uint8_t x_55; -lean_dec(x_45); -x_55 = !lean_is_exclusive(x_47); -if (x_55 == 0) -{ -return x_47; -} -else -{ -lean_object* x_56; lean_object* x_57; lean_object* x_58; -x_56 = lean_ctor_get(x_47, 0); -x_57 = lean_ctor_get(x_47, 1); -lean_inc(x_57); -lean_inc(x_56); -lean_dec(x_47); -x_58 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_58, 0, x_56); -lean_ctor_set(x_58, 1, x_57); -return x_58; -} -} -} -else -{ -lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; -x_59 = lean_ctor_get(x_26, 0); -x_60 = lean_ctor_get(x_26, 1); -lean_inc(x_60); -lean_inc(x_59); -lean_dec(x_26); -x_61 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_61, 0, x_19); -lean_ctor_set(x_61, 1, x_8); -x_62 = lean_array_mk(x_61); -x_63 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__11; -x_64 = l_Lean_Meta_Simp_mkContext(x_63, x_62, x_59, x_3, x_4, x_5, x_6, x_60); -x_65 = lean_ctor_get(x_64, 0); -lean_inc(x_65); -x_66 = lean_ctor_get(x_64, 1); -lean_inc(x_66); -lean_dec(x_64); -x_67 = l_Lean_Meta_Grind_mkMethods___rarg(x_66); -x_68 = lean_ctor_get(x_67, 0); -lean_inc(x_68); -x_69 = lean_ctor_get(x_67, 1); -lean_inc(x_69); -lean_dec(x_67); -x_70 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_70, 0, x_65); -lean_ctor_set(x_70, 1, x_25); -lean_ctor_set(x_70, 2, x_2); -x_71 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__14; -x_72 = lean_unsigned_to_nat(1u); -x_73 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__13; -x_74 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__20; -x_75 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_75, 0, x_71); -lean_ctor_set(x_75, 1, x_16); -lean_ctor_set(x_75, 2, x_72); -lean_ctor_set(x_75, 3, x_73); -lean_ctor_set(x_75, 4, x_74); -lean_ctor_set(x_75, 5, x_15); -lean_ctor_set(x_75, 6, x_10); -x_76 = lean_st_mk_ref(x_75, x_69); -x_77 = lean_ctor_get(x_76, 0); -lean_inc(x_77); -x_78 = lean_ctor_get(x_76, 1); -lean_inc(x_78); -lean_dec(x_76); -lean_inc(x_77); -x_79 = lean_apply_8(x_1, x_68, x_70, x_77, x_3, x_4, x_5, x_6, 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; -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_get(x_77, x_81); -lean_dec(x_77); -x_83 = lean_ctor_get(x_82, 1); -lean_inc(x_83); -if (lean_is_exclusive(x_82)) { - lean_ctor_release(x_82, 0); - lean_ctor_release(x_82, 1); - x_84 = x_82; -} else { - lean_dec_ref(x_82); - x_84 = lean_box(0); -} -if (lean_is_scalar(x_84)) { - x_85 = lean_alloc_ctor(0, 2, 0); -} else { - x_85 = x_84; -} -lean_ctor_set(x_85, 0, x_80); -lean_ctor_set(x_85, 1, x_83); -return x_85; -} -else -{ -lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; -lean_dec(x_77); -x_86 = lean_ctor_get(x_79, 0); -lean_inc(x_86); -x_87 = lean_ctor_get(x_79, 1); -lean_inc(x_87); -if (lean_is_exclusive(x_79)) { - lean_ctor_release(x_79, 0); - lean_ctor_release(x_79, 1); - x_88 = x_79; -} else { - lean_dec_ref(x_79); - x_88 = lean_box(0); -} -if (lean_is_scalar(x_88)) { - x_89 = lean_alloc_ctor(1, 2, 0); -} else { - x_89 = x_88; -} -lean_ctor_set(x_89, 0, x_86); -lean_ctor_set(x_89, 1, x_87); -return x_89; -} -} -} -else -{ -lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; -x_90 = lean_ctor_get(x_22, 0); -x_91 = lean_ctor_get(x_22, 1); -lean_inc(x_91); -lean_inc(x_90); -lean_dec(x_22); -x_92 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_92, 0, x_90); -lean_ctor_set(x_92, 1, x_8); -x_93 = lean_array_mk(x_92); -x_94 = l_Lean_Meta_getSimpCongrTheorems___rarg(x_6, x_91); -x_95 = lean_ctor_get(x_94, 0); -lean_inc(x_95); -x_96 = lean_ctor_get(x_94, 1); -lean_inc(x_96); -if (lean_is_exclusive(x_94)) { - lean_ctor_release(x_94, 0); - lean_ctor_release(x_94, 1); - x_97 = x_94; -} else { - lean_dec_ref(x_94); - x_97 = lean_box(0); -} -if (lean_is_scalar(x_97)) { - x_98 = lean_alloc_ctor(1, 2, 0); -} else { - x_98 = x_97; - lean_ctor_set_tag(x_98, 1); -} -lean_ctor_set(x_98, 0, x_19); -lean_ctor_set(x_98, 1, x_8); -x_99 = lean_array_mk(x_98); -x_100 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__11; -x_101 = l_Lean_Meta_Simp_mkContext(x_100, x_99, x_95, x_3, x_4, x_5, x_6, x_96); -x_102 = lean_ctor_get(x_101, 0); -lean_inc(x_102); -x_103 = lean_ctor_get(x_101, 1); -lean_inc(x_103); -lean_dec(x_101); -x_104 = l_Lean_Meta_Grind_mkMethods___rarg(x_103); -x_105 = lean_ctor_get(x_104, 0); -lean_inc(x_105); -x_106 = lean_ctor_get(x_104, 1); -lean_inc(x_106); -lean_dec(x_104); -x_107 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_107, 0, x_102); -lean_ctor_set(x_107, 1, x_93); -lean_ctor_set(x_107, 2, x_2); -x_108 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__14; -x_109 = lean_unsigned_to_nat(1u); -x_110 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__13; -x_111 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__20; -x_112 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_112, 0, x_108); -lean_ctor_set(x_112, 1, x_16); -lean_ctor_set(x_112, 2, x_109); -lean_ctor_set(x_112, 3, x_110); -lean_ctor_set(x_112, 4, x_111); -lean_ctor_set(x_112, 5, x_15); -lean_ctor_set(x_112, 6, x_10); -x_113 = lean_st_mk_ref(x_112, x_106); -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); -lean_inc(x_114); -x_116 = lean_apply_8(x_1, x_105, x_107, x_114, x_3, x_4, x_5, x_6, x_115); -if (lean_obj_tag(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; -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_st_ref_get(x_114, x_118); -lean_dec(x_114); -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); -} -if (lean_is_scalar(x_121)) { - x_122 = lean_alloc_ctor(0, 2, 0); -} else { - x_122 = x_121; -} -lean_ctor_set(x_122, 0, x_117); -lean_ctor_set(x_122, 1, x_120); -return x_122; -} -else -{ -lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; -lean_dec(x_114); -x_123 = lean_ctor_get(x_116, 0); -lean_inc(x_123); -x_124 = lean_ctor_get(x_116, 1); -lean_inc(x_124); -if (lean_is_exclusive(x_116)) { - lean_ctor_release(x_116, 0); - lean_ctor_release(x_116, 1); - x_125 = x_116; -} else { - lean_dec_ref(x_116); - x_125 = lean_box(0); -} -if (lean_is_scalar(x_125)) { - x_126 = lean_alloc_ctor(1, 2, 0); -} else { - x_126 = x_125; -} -lean_ctor_set(x_126, 0, x_123); -lean_ctor_set(x_126, 1, x_124); -return x_126; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GrindM_run(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_GrindM_run___rarg), 7, 0); -return x_2; -} -} -LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -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_Lean_Meta_Grind_GoalM_run___spec__1(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 0); -return x_3; -} -} -LEAN_EXPORT lean_object* l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___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; -x_11 = lean_apply_3(x_2, x_3, x_4, x_5); -x_12 = l___private_Lean_Meta_Basic_0__Lean_Meta_withMVarContextImp___rarg(x_1, x_11, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_12) == 0) -{ -uint8_t x_13; -x_13 = !lean_is_exclusive(x_12); -if (x_13 == 0) -{ -return x_12; -} -else -{ -lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_14 = lean_ctor_get(x_12, 0); -x_15 = lean_ctor_get(x_12, 1); -lean_inc(x_15); -lean_inc(x_14); -lean_dec(x_12); -x_16 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_16, 0, x_14); -lean_ctor_set(x_16, 1, x_15); -return x_16; -} -} -else -{ -uint8_t x_17; -x_17 = !lean_is_exclusive(x_12); -if (x_17 == 0) -{ -return x_12; -} -else -{ -lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_18 = lean_ctor_get(x_12, 0); -x_19 = lean_ctor_get(x_12, 1); -lean_inc(x_19); -lean_inc(x_18); -lean_dec(x_12); -x_20 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_20, 0, x_18); -lean_ctor_set(x_20, 1, x_19); -return x_20; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg), 10, 0); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GoalM_run___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; uint8_t x_11; -x_10 = lean_st_mk_ref(x_1, x_9); -x_11 = !lean_is_exclusive(x_10); -if (x_11 == 0) -{ -return x_10; -} -else -{ -lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_12 = lean_ctor_get(x_10, 0); -x_13 = lean_ctor_get(x_10, 1); -lean_inc(x_13); -lean_inc(x_12); -lean_dec(x_10); -x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_12); -lean_ctor_set(x_14, 1, x_13); -return x_14; -} -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GoalM_run___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: -{ -lean_object* x_11; -lean_inc(x_2); -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) -{ -lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); -x_13 = lean_ctor_get(x_11, 1); -lean_inc(x_13); -lean_dec(x_11); -x_14 = lean_st_ref_get(x_2, x_13); -lean_dec(x_2); -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); -x_17 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_17, 0, x_12); -lean_ctor_set(x_17, 1, x_16); -lean_ctor_set(x_14, 0, x_17); -return x_14; -} -else -{ -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_18 = lean_ctor_get(x_14, 0); -x_19 = lean_ctor_get(x_14, 1); -lean_inc(x_19); -lean_inc(x_18); -lean_dec(x_14); -x_20 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_20, 0, x_12); -lean_ctor_set(x_20, 1, x_18); -x_21 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_21, 0, x_20); -lean_ctor_set(x_21, 1, x_19); -return x_21; -} -} -else -{ -uint8_t x_22; -lean_dec(x_2); -x_22 = !lean_is_exclusive(x_11); -if (x_22 == 0) -{ -return x_11; -} -else -{ -lean_object* x_23; lean_object* x_24; lean_object* x_25; -x_23 = lean_ctor_get(x_11, 0); -x_24 = lean_ctor_get(x_11, 1); -lean_inc(x_24); -lean_inc(x_23); -lean_dec(x_11); -x_25 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_25, 0, x_23); -lean_ctor_set(x_25, 1, x_24); -return x_25; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GoalM_run___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_11 = lean_ctor_get(x_1, 0); -lean_inc(x_11); -x_12 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_GoalM_run___rarg___lambda__1___boxed), 9, 1); -lean_closure_set(x_12, 0, x_1); -x_13 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_GoalM_run___rarg___lambda__2), 10, 1); -lean_closure_set(x_13, 0, x_2); -x_14 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); -lean_closure_set(x_14, 0, x_12); -lean_closure_set(x_14, 1, x_13); -x_15 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(x_11, x_14, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_15; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GoalM_run(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_GoalM_run___rarg), 10, 0); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GoalM_run___rarg___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -lean_object* x_10; -x_10 = l_Lean_Meta_Grind_GoalM_run___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); -lean_dec(x_3); -lean_dec(x_2); -return x_10; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GoalM_run_x27___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_11; -lean_inc(x_2); -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) -{ -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 = lean_st_ref_get(x_2, x_12); -x_14 = !lean_is_exclusive(x_13); -if (x_14 == 0) -{ -lean_object* x_15; lean_object* x_16; uint8_t x_17; -x_15 = lean_ctor_get(x_13, 1); -x_16 = lean_st_ref_get(x_2, x_15); -lean_dec(x_2); -x_17 = !lean_is_exclusive(x_16); -if (x_17 == 0) -{ -lean_object* x_18; -x_18 = lean_ctor_get(x_16, 0); -lean_ctor_set(x_13, 1, x_18); -lean_ctor_set(x_16, 0, x_13); -return x_16; -} -else -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_19 = lean_ctor_get(x_16, 0); -x_20 = lean_ctor_get(x_16, 1); -lean_inc(x_20); -lean_inc(x_19); -lean_dec(x_16); -lean_ctor_set(x_13, 1, x_19); -x_21 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_21, 0, x_13); -lean_ctor_set(x_21, 1, x_20); -return x_21; -} -} -else -{ -lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_22 = lean_ctor_get(x_13, 0); -x_23 = lean_ctor_get(x_13, 1); -lean_inc(x_23); -lean_inc(x_22); -lean_dec(x_13); -x_24 = lean_st_ref_get(x_2, x_23); -lean_dec(x_2); -x_25 = lean_ctor_get(x_24, 0); -lean_inc(x_25); -x_26 = lean_ctor_get(x_24, 1); -lean_inc(x_26); -if (lean_is_exclusive(x_24)) { - lean_ctor_release(x_24, 0); - lean_ctor_release(x_24, 1); - x_27 = x_24; -} else { - lean_dec_ref(x_24); - x_27 = lean_box(0); -} -x_28 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_28, 0, x_22); -lean_ctor_set(x_28, 1, x_25); -if (lean_is_scalar(x_27)) { - x_29 = lean_alloc_ctor(0, 2, 0); -} else { - x_29 = x_27; -} -lean_ctor_set(x_29, 0, x_28); -lean_ctor_set(x_29, 1, x_26); -return x_29; -} -} -else -{ -uint8_t x_30; -lean_dec(x_2); -x_30 = !lean_is_exclusive(x_11); -if (x_30 == 0) -{ -return x_11; -} -else -{ -lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_31 = lean_ctor_get(x_11, 0); -x_32 = lean_ctor_get(x_11, 1); -lean_inc(x_32); -lean_inc(x_31); -lean_dec(x_11); -x_33 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_33, 0, x_31); -lean_ctor_set(x_33, 1, x_32); -return x_33; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GoalM_run_x27___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; -x_10 = !lean_is_exclusive(x_1); -if (x_10 == 0) -{ -lean_object* x_11; -x_11 = lean_ctor_get(x_1, 1); -lean_dec(x_11); -lean_ctor_set(x_1, 1, x_9); -return x_1; -} -else -{ -lean_object* x_12; lean_object* x_13; -x_12 = lean_ctor_get(x_1, 0); -lean_inc(x_12); -lean_dec(x_1); -x_13 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_13, 0, x_12); -lean_ctor_set(x_13, 1, x_9); -return x_13; -} -} -} -static lean_object* _init_l_Lean_Meta_Grind_GoalM_run_x27___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_GoalM_run_x27___lambda__2___boxed), 9, 0); -return x_1; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GoalM_run_x27(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_11 = lean_ctor_get(x_1, 0); -lean_inc(x_11); -x_12 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_GoalM_run___rarg___lambda__1___boxed), 9, 1); -lean_closure_set(x_12, 0, x_1); -x_13 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_GoalM_run_x27___lambda__1), 10, 1); -lean_closure_set(x_13, 0, x_2); -x_14 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); -lean_closure_set(x_14, 0, x_12); -lean_closure_set(x_14, 1, x_13); -x_15 = l_Lean_Meta_Grind_GoalM_run_x27___closed__1; -x_16 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); -lean_closure_set(x_16, 0, x_14); -lean_closure_set(x_16, 1, x_15); -x_17 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(x_11, x_16, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_17; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GoalM_run_x27___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -lean_object* x_10; -x_10 = l_Lean_Meta_Grind_GoalM_run_x27___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_10; -} -} -static lean_object* _init_l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_mkGoal___spec__1() { -_start: -{ -lean_object* x_1; -x_1 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__13; -return x_1; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkGoal___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_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; uint8_t x_20; -x_12 = 1; -x_13 = 0; -x_14 = lean_unsigned_to_nat(0u); -x_15 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_12, x_13, x_14, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -x_16 = lean_ctor_get(x_15, 1); -lean_inc(x_16); -lean_dec(x_15); -x_17 = l_Lean_Meta_Grind_mkENodeCore(x_2, x_12, x_13, x_14, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_16); -x_18 = lean_ctor_get(x_17, 1); -lean_inc(x_18); -lean_dec(x_17); -x_19 = lean_st_ref_get(x_3, x_18); -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, 1); -x_22 = lean_st_ref_get(x_3, x_21); -x_23 = !lean_is_exclusive(x_22); -if (x_23 == 0) -{ -lean_object* x_24; -x_24 = lean_ctor_get(x_22, 0); -lean_ctor_set(x_19, 1, x_24); -lean_ctor_set(x_22, 0, x_19); -return x_22; -} -else -{ -lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_25 = lean_ctor_get(x_22, 0); -x_26 = lean_ctor_get(x_22, 1); -lean_inc(x_26); -lean_inc(x_25); -lean_dec(x_22); -lean_ctor_set(x_19, 1, x_25); -x_27 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_27, 0, x_19); -lean_ctor_set(x_27, 1, x_26); -return x_27; -} -} -else -{ -lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_28 = lean_ctor_get(x_19, 0); -x_29 = lean_ctor_get(x_19, 1); -lean_inc(x_29); -lean_inc(x_28); -lean_dec(x_19); -x_30 = lean_st_ref_get(x_3, x_29); -x_31 = lean_ctor_get(x_30, 0); -lean_inc(x_31); -x_32 = lean_ctor_get(x_30, 1); -lean_inc(x_32); -if (lean_is_exclusive(x_30)) { - lean_ctor_release(x_30, 0); - lean_ctor_release(x_30, 1); - x_33 = x_30; -} else { - lean_dec_ref(x_30); - x_33 = lean_box(0); -} -x_34 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_34, 0, x_28); -lean_ctor_set(x_34, 1, x_31); -if (lean_is_scalar(x_33)) { - x_35 = lean_alloc_ctor(0, 2, 0); -} else { - x_35 = x_33; -} -lean_ctor_set(x_35, 0, x_34); -lean_ctor_set(x_35, 1, x_32); -return x_35; -} -} -} -static lean_object* _init_l_Lean_Meta_Grind_mkGoal___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = lean_box(0); -x_2 = lean_array_mk(x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkGoal(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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; 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; -x_10 = l_Lean_Meta_Grind_getTrueExpr___rarg(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_Meta_Grind_getFalseExpr___rarg(x_4, x_5, x_6, x_7, x_8, x_12); -x_14 = lean_ctor_get(x_13, 0); -lean_inc(x_14); -x_15 = lean_ctor_get(x_13, 1); -lean_inc(x_15); -lean_dec(x_13); -x_16 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__13; -x_17 = l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_mkGoal___spec__1; -x_18 = l_Lean_Meta_Grind_mkGoal___closed__1; -x_19 = 0; -x_20 = lean_unsigned_to_nat(0u); -lean_inc(x_1); -x_21 = lean_alloc_ctor(0, 8, 1); -lean_ctor_set(x_21, 0, x_1); -lean_ctor_set(x_21, 1, x_16); -lean_ctor_set(x_21, 2, x_16); -lean_ctor_set(x_21, 3, x_17); -lean_ctor_set(x_21, 4, x_16); -lean_ctor_set(x_21, 5, x_18); -lean_ctor_set(x_21, 6, x_20); -lean_ctor_set(x_21, 7, x_20); -lean_ctor_set_uint8(x_21, sizeof(void*)*8, x_19); -x_22 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_GoalM_run___rarg___lambda__1___boxed), 9, 1); -lean_closure_set(x_22, 0, x_21); -x_23 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_mkGoal___lambda__1___boxed), 11, 2); -lean_closure_set(x_23, 0, x_14); -lean_closure_set(x_23, 1, x_11); -x_24 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); -lean_closure_set(x_24, 0, x_22); -lean_closure_set(x_24, 1, x_23); -x_25 = l_Lean_Meta_Grind_GoalM_run_x27___closed__1; -x_26 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); -lean_closure_set(x_26, 0, x_24); -lean_closure_set(x_26, 1, x_25); -x_27 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(x_1, x_26, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_15); -return x_27; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkGoal___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_Meta_Grind_mkGoal___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -return x_12; -} -} -lean_object* initialize_Init_Grind_Lemmas(uint8_t builtin, lean_object*); -lean_object* initialize_Lean_Meta_Tactic_Grind_Types(uint8_t builtin, lean_object*); -lean_object* initialize_Lean_Meta_Tactic_Grind_PropagatorAttr(uint8_t builtin, lean_object*); -lean_object* initialize_Lean_Meta_Tactic_Grind_Proj(uint8_t builtin, lean_object*); -static bool _G_initialized = false; -LEAN_EXPORT lean_object* initialize_Lean_Meta_Tactic_Grind_Run(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_Grind_Lemmas(builtin, lean_io_mk_world()); -if (lean_io_result_is_error(res)) return res; -lean_dec_ref(res); -res = initialize_Lean_Meta_Tactic_Grind_Types(builtin, lean_io_mk_world()); -if (lean_io_result_is_error(res)) return res; -lean_dec_ref(res); -res = initialize_Lean_Meta_Tactic_Grind_PropagatorAttr(builtin, lean_io_mk_world()); -if (lean_io_result_is_error(res)) return res; -lean_dec_ref(res); -res = initialize_Lean_Meta_Tactic_Grind_Proj(builtin, lean_io_mk_world()); -if (lean_io_result_is_error(res)) return res; -lean_dec_ref(res); -l_Lean_Meta_Grind_mkMethods___rarg___closed__1 = _init_l_Lean_Meta_Grind_mkMethods___rarg___closed__1(); -lean_mark_persistent(l_Lean_Meta_Grind_mkMethods___rarg___closed__1); -l_Lean_Meta_Grind_GrindM_run___rarg___closed__1 = _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__1(); -lean_mark_persistent(l_Lean_Meta_Grind_GrindM_run___rarg___closed__1); -l_Lean_Meta_Grind_GrindM_run___rarg___closed__2 = _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__2(); -lean_mark_persistent(l_Lean_Meta_Grind_GrindM_run___rarg___closed__2); -l_Lean_Meta_Grind_GrindM_run___rarg___closed__3 = _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__3(); -lean_mark_persistent(l_Lean_Meta_Grind_GrindM_run___rarg___closed__3); -l_Lean_Meta_Grind_GrindM_run___rarg___closed__4 = _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__4(); -lean_mark_persistent(l_Lean_Meta_Grind_GrindM_run___rarg___closed__4); -l_Lean_Meta_Grind_GrindM_run___rarg___closed__5 = _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__5(); -lean_mark_persistent(l_Lean_Meta_Grind_GrindM_run___rarg___closed__5); -l_Lean_Meta_Grind_GrindM_run___rarg___closed__6 = _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__6(); -lean_mark_persistent(l_Lean_Meta_Grind_GrindM_run___rarg___closed__6); -l_Lean_Meta_Grind_GrindM_run___rarg___closed__7 = _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__7(); -lean_mark_persistent(l_Lean_Meta_Grind_GrindM_run___rarg___closed__7); -l_Lean_Meta_Grind_GrindM_run___rarg___closed__8 = _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__8(); -lean_mark_persistent(l_Lean_Meta_Grind_GrindM_run___rarg___closed__8); -l_Lean_Meta_Grind_GrindM_run___rarg___closed__9 = _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__9(); -lean_mark_persistent(l_Lean_Meta_Grind_GrindM_run___rarg___closed__9); -l_Lean_Meta_Grind_GrindM_run___rarg___closed__10 = _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__10(); -lean_mark_persistent(l_Lean_Meta_Grind_GrindM_run___rarg___closed__10); -l_Lean_Meta_Grind_GrindM_run___rarg___closed__11 = _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__11(); -lean_mark_persistent(l_Lean_Meta_Grind_GrindM_run___rarg___closed__11); -l_Lean_Meta_Grind_GrindM_run___rarg___closed__12 = _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__12(); -lean_mark_persistent(l_Lean_Meta_Grind_GrindM_run___rarg___closed__12); -l_Lean_Meta_Grind_GrindM_run___rarg___closed__13 = _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__13(); -lean_mark_persistent(l_Lean_Meta_Grind_GrindM_run___rarg___closed__13); -l_Lean_Meta_Grind_GrindM_run___rarg___closed__14 = _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__14(); -lean_mark_persistent(l_Lean_Meta_Grind_GrindM_run___rarg___closed__14); -l_Lean_Meta_Grind_GrindM_run___rarg___closed__15 = _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__15(); -lean_mark_persistent(l_Lean_Meta_Grind_GrindM_run___rarg___closed__15); -l_Lean_Meta_Grind_GrindM_run___rarg___closed__16 = _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__16(); -lean_mark_persistent(l_Lean_Meta_Grind_GrindM_run___rarg___closed__16); -l_Lean_Meta_Grind_GrindM_run___rarg___closed__17 = _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__17(); -lean_mark_persistent(l_Lean_Meta_Grind_GrindM_run___rarg___closed__17); -l_Lean_Meta_Grind_GrindM_run___rarg___closed__18 = _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__18(); -lean_mark_persistent(l_Lean_Meta_Grind_GrindM_run___rarg___closed__18); -l_Lean_Meta_Grind_GrindM_run___rarg___closed__19 = _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__19(); -lean_mark_persistent(l_Lean_Meta_Grind_GrindM_run___rarg___closed__19); -l_Lean_Meta_Grind_GrindM_run___rarg___closed__20 = _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__20(); -lean_mark_persistent(l_Lean_Meta_Grind_GrindM_run___rarg___closed__20); -l_Lean_Meta_Grind_GoalM_run_x27___closed__1 = _init_l_Lean_Meta_Grind_GoalM_run_x27___closed__1(); -lean_mark_persistent(l_Lean_Meta_Grind_GoalM_run_x27___closed__1); -l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_mkGoal___spec__1 = _init_l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_mkGoal___spec__1(); -lean_mark_persistent(l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_mkGoal___spec__1); -l_Lean_Meta_Grind_mkGoal___closed__1 = _init_l_Lean_Meta_Grind_mkGoal___closed__1(); -lean_mark_persistent(l_Lean_Meta_Grind_mkGoal___closed__1); -return lean_io_result_mk_ok(lean_box(0)); -} -#ifdef __cplusplus -} -#endif diff --git a/stage0/stdlib/Lean/Meta/Tactic/Grind/Simp.c b/stage0/stdlib/Lean/Meta/Tactic/Grind/Simp.c index 1d15456826b3..3e5d9b326863 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Grind/Simp.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Grind/Simp.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Lean.Meta.Tactic.Grind.Simp -// Imports: Lean.Meta.Tactic.Simp.Main Lean.Meta.Tactic.Grind.Util Lean.Meta.Tactic.Grind.Types Lean.Meta.Tactic.Grind.MarkNestedProofs +// Imports: Init.Grind.Lemmas Lean.Meta.Tactic.Assert Lean.Meta.Tactic.Simp.Main Lean.Meta.Tactic.Grind.Util Lean.Meta.Tactic.Grind.Types Lean.Meta.Tactic.Grind.MarkNestedProofs #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -13,58 +13,63 @@ #ifdef __cplusplus extern "C" { #endif -static lean_object* l_Lean_Meta_Grind_pre___closed__7; +static lean_object* l_Lean_Meta_Grind_simp___closed__7; +LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Meta_Grind_simp___spec__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_Meta_Grind_abstractNestedProofs(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_Meta_Grind_simp___spec__3___closed__3; lean_object* l_Lean_Meta_Grind_eraseIrrelevantMData___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_unfoldReducible___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PersistentArray_push___rarg(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_pre___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_simp___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* l_Lean_Meta_Grind_normalizeLevels___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_pre___closed__11; -static lean_object* l_Lean_addTrace___at_Lean_Meta_Grind_pre___spec__2___closed__3; -LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Meta_Grind_pre___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_simp___closed__12; lean_object* l_Lean_stringToMessageData(lean_object*); -static lean_object* l_Lean_addTrace___at_Lean_Meta_Grind_pre___spec__2___closed__2; -static lean_object* l_Lean_Meta_Grind_pre___closed__4; +static lean_object* l_Lean_Meta_Grind_simp___closed__1; +static lean_object* l_Lean_Meta_Grind_simp___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_simp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_pre___closed__3; -static lean_object* l_Lean_Meta_Grind_pre___closed__10; +uint8_t l_Lean_Expr_hasMVar(lean_object*); +LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Meta_Grind_simp___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_addTrace___at_Lean_Meta_Grind_simp___spec__3___closed__2; +static lean_object* l_Lean_Meta_Grind_simp___closed__4; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_simp___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Core_transform___at_Lean_Meta_Grind_unfoldReducible___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_take(lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_eraseIrrelevantMData___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_pre___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_Grind_pre___closed__5; +static lean_object* l_Lean_Meta_Grind_simp___closed__9; 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*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_simp___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_Grind_markNestedProofsImpl(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_Meta_Grind_pre___closed__9; -static lean_object* l_Lean_Meta_Grind_pre___closed__8; +LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_simp___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofExpr(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pre___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_Grind_simp___closed__5; double l_Float_ofScientific(lean_object*, uint8_t, lean_object*); lean_object* l_Lean_Core_transform___at_Lean_Core_betaReduce___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_shareCommon(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_simp___closed__3; lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at_Lean_Meta_Grind_simp___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_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_addTrace___at_Lean_Meta_Grind_pre___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static double l_Lean_addTrace___at_Lean_Meta_Grind_pre___spec__2___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_simpCore(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_Meta_Grind_simp___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_Grind_simpCore___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_Grind_simp___closed__11; lean_object* lean_array_mk(lean_object*); +lean_object* l_Lean_instantiateMVarsCore(lean_object*, lean_object*); lean_object* l_Lean_isTracingEnabledForCore(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pre(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_pre___closed__6; +static double l_Lean_addTrace___at_Lean_Meta_Grind_simp___spec__3___closed__1; +static lean_object* l_Lean_Meta_Grind_simp___closed__8; lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_foldProjs___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_simp___closed__10; +static lean_object* l_Lean_Meta_Grind_simp___closed__6; +static lean_object* l_Lean_Meta_Grind_simp___closed__13; lean_object* l_Lean_Meta_Grind_canon(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_pre___closed__12; -LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_pre___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at_Lean_Meta_Grind_simp___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_unfoldReducible___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pre___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_foldProjs___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pre___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*); -static lean_object* l_Lean_Meta_Grind_pre___closed__2; -static lean_object* l_Lean_Meta_Grind_pre___closed__13; -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_simp(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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_Grind_simpCore(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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; @@ -201,17 +206,124 @@ return x_46; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_simp___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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_Grind_simpCore___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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_Grind_simp(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_10 = l_Lean_Meta_Grind_simpCore(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); lean_dec(x_4); lean_dec(x_2); return x_10; } } -LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_pre___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at_Lean_Meta_Grind_simp___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 = l_Lean_Expr_hasMVar(x_1); +if (x_10 == 0) +{ +lean_object* x_11; +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_1); +lean_ctor_set(x_11, 1, x_9); +return x_11; +} +else +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_12 = lean_st_ref_get(x_6, x_9); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_12, 1); +lean_inc(x_14); +lean_dec(x_12); +x_15 = lean_ctor_get(x_13, 0); +lean_inc(x_15); +lean_dec(x_13); +x_16 = l_Lean_instantiateMVarsCore(x_15, x_1); +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +x_18 = lean_ctor_get(x_16, 1); +lean_inc(x_18); +lean_dec(x_16); +x_19 = lean_st_ref_take(x_6, x_14); +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; uint8_t x_25; +x_23 = lean_ctor_get(x_20, 0); +lean_dec(x_23); +lean_ctor_set(x_20, 0, x_18); +x_24 = lean_st_ref_set(x_6, x_20, x_21); +x_25 = !lean_is_exclusive(x_24); +if (x_25 == 0) +{ +lean_object* x_26; +x_26 = lean_ctor_get(x_24, 0); +lean_dec(x_26); +lean_ctor_set(x_24, 0, x_17); +return x_24; +} +else +{ +lean_object* x_27; lean_object* x_28; +x_27 = lean_ctor_get(x_24, 1); +lean_inc(x_27); +lean_dec(x_24); +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_17); +lean_ctor_set(x_28, 1, x_27); +return x_28; +} +} +else +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_29 = lean_ctor_get(x_20, 1); +x_30 = lean_ctor_get(x_20, 2); +x_31 = lean_ctor_get(x_20, 3); +x_32 = lean_ctor_get(x_20, 4); +lean_inc(x_32); +lean_inc(x_31); +lean_inc(x_30); +lean_inc(x_29); +lean_dec(x_20); +x_33 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_33, 0, x_18); +lean_ctor_set(x_33, 1, x_29); +lean_ctor_set(x_33, 2, x_30); +lean_ctor_set(x_33, 3, x_31); +lean_ctor_set(x_33, 4, x_32); +x_34 = lean_st_ref_set(x_6, x_33, x_21); +x_35 = lean_ctor_get(x_34, 1); +lean_inc(x_35); +if (lean_is_exclusive(x_34)) { + lean_ctor_release(x_34, 0); + lean_ctor_release(x_34, 1); + x_36 = x_34; +} else { + lean_dec_ref(x_34); + x_36 = lean_box(0); +} +if (lean_is_scalar(x_36)) { + x_37 = lean_alloc_ctor(0, 2, 0); +} else { + x_37 = x_36; +} +lean_ctor_set(x_37, 0, x_17); +lean_ctor_set(x_37, 1, x_35); +return x_37; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_simp___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { lean_object* x_10; lean_object* x_11; uint8_t x_12; @@ -237,7 +349,7 @@ return x_15; } } } -static double _init_l_Lean_addTrace___at_Lean_Meta_Grind_pre___spec__2___closed__1() { +static double _init_l_Lean_addTrace___at_Lean_Meta_Grind_simp___spec__3___closed__1() { _start: { lean_object* x_1; uint8_t x_2; double x_3; @@ -247,7 +359,7 @@ x_3 = l_Float_ofScientific(x_1, x_2, x_1); return x_3; } } -static lean_object* _init_l_Lean_addTrace___at_Lean_Meta_Grind_pre___spec__2___closed__2() { +static lean_object* _init_l_Lean_addTrace___at_Lean_Meta_Grind_simp___spec__3___closed__2() { _start: { lean_object* x_1; @@ -255,7 +367,7 @@ x_1 = lean_mk_string_unchecked("", 0, 0); return x_1; } } -static lean_object* _init_l_Lean_addTrace___at_Lean_Meta_Grind_pre___spec__2___closed__3() { +static lean_object* _init_l_Lean_addTrace___at_Lean_Meta_Grind_simp___spec__3___closed__3() { _start: { lean_object* x_1; lean_object* x_2; @@ -264,7 +376,7 @@ x_2 = lean_array_mk(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Meta_Grind_pre___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_EXPORT lean_object* l_Lean_addTrace___at_Lean_Meta_Grind_simp___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { lean_object* x_11; 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; @@ -298,16 +410,16 @@ if (x_23 == 0) { lean_object* x_24; double 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; x_24 = lean_ctor_get(x_17, 0); -x_25 = l_Lean_addTrace___at_Lean_Meta_Grind_pre___spec__2___closed__1; +x_25 = l_Lean_addTrace___at_Lean_Meta_Grind_simp___spec__3___closed__1; x_26 = 0; -x_27 = l_Lean_addTrace___at_Lean_Meta_Grind_pre___spec__2___closed__2; +x_27 = l_Lean_addTrace___at_Lean_Meta_Grind_simp___spec__3___closed__2; x_28 = lean_alloc_ctor(0, 2, 17); lean_ctor_set(x_28, 0, x_1); lean_ctor_set(x_28, 1, x_27); lean_ctor_set_float(x_28, sizeof(void*)*2, x_25); lean_ctor_set_float(x_28, sizeof(void*)*2 + 8, x_25); lean_ctor_set_uint8(x_28, sizeof(void*)*2 + 16, x_26); -x_29 = l_Lean_addTrace___at_Lean_Meta_Grind_pre___spec__2___closed__3; +x_29 = l_Lean_addTrace___at_Lean_Meta_Grind_simp___spec__3___closed__3; x_30 = lean_alloc_ctor(9, 3, 0); lean_ctor_set(x_30, 0, x_28); lean_ctor_set(x_30, 1, x_13); @@ -348,16 +460,16 @@ x_39 = lean_ctor_get_uint64(x_17, sizeof(void*)*1); x_40 = lean_ctor_get(x_17, 0); lean_inc(x_40); lean_dec(x_17); -x_41 = l_Lean_addTrace___at_Lean_Meta_Grind_pre___spec__2___closed__1; +x_41 = l_Lean_addTrace___at_Lean_Meta_Grind_simp___spec__3___closed__1; x_42 = 0; -x_43 = l_Lean_addTrace___at_Lean_Meta_Grind_pre___spec__2___closed__2; +x_43 = l_Lean_addTrace___at_Lean_Meta_Grind_simp___spec__3___closed__2; x_44 = lean_alloc_ctor(0, 2, 17); lean_ctor_set(x_44, 0, x_1); lean_ctor_set(x_44, 1, x_43); lean_ctor_set_float(x_44, sizeof(void*)*2, x_41); lean_ctor_set_float(x_44, sizeof(void*)*2 + 8, x_41); lean_ctor_set_uint8(x_44, sizeof(void*)*2 + 16, x_42); -x_45 = l_Lean_addTrace___at_Lean_Meta_Grind_pre___spec__2___closed__3; +x_45 = l_Lean_addTrace___at_Lean_Meta_Grind_simp___spec__3___closed__3; x_46 = lean_alloc_ctor(9, 3, 0); lean_ctor_set(x_46, 0, x_44); lean_ctor_set(x_46, 1, x_13); @@ -420,16 +532,16 @@ if (lean_is_exclusive(x_17)) { lean_dec_ref(x_17); x_63 = lean_box(0); } -x_64 = l_Lean_addTrace___at_Lean_Meta_Grind_pre___spec__2___closed__1; +x_64 = l_Lean_addTrace___at_Lean_Meta_Grind_simp___spec__3___closed__1; x_65 = 0; -x_66 = l_Lean_addTrace___at_Lean_Meta_Grind_pre___spec__2___closed__2; +x_66 = l_Lean_addTrace___at_Lean_Meta_Grind_simp___spec__3___closed__2; x_67 = lean_alloc_ctor(0, 2, 17); lean_ctor_set(x_67, 0, x_1); lean_ctor_set(x_67, 1, x_66); lean_ctor_set_float(x_67, sizeof(void*)*2, x_64); lean_ctor_set_float(x_67, sizeof(void*)*2 + 8, x_64); lean_ctor_set_uint8(x_67, sizeof(void*)*2 + 16, x_65); -x_68 = l_Lean_addTrace___at_Lean_Meta_Grind_pre___spec__2___closed__3; +x_68 = l_Lean_addTrace___at_Lean_Meta_Grind_simp___spec__3___closed__3; x_69 = lean_alloc_ctor(9, 3, 0); lean_ctor_set(x_69, 0, x_67); lean_ctor_set(x_69, 1, x_13); @@ -520,16 +632,16 @@ if (lean_is_exclusive(x_17)) { lean_dec_ref(x_17); x_89 = lean_box(0); } -x_90 = l_Lean_addTrace___at_Lean_Meta_Grind_pre___spec__2___closed__1; +x_90 = l_Lean_addTrace___at_Lean_Meta_Grind_simp___spec__3___closed__1; x_91 = 0; -x_92 = l_Lean_addTrace___at_Lean_Meta_Grind_pre___spec__2___closed__2; +x_92 = l_Lean_addTrace___at_Lean_Meta_Grind_simp___spec__3___closed__2; x_93 = lean_alloc_ctor(0, 2, 17); lean_ctor_set(x_93, 0, x_1); lean_ctor_set(x_93, 1, x_92); lean_ctor_set_float(x_93, sizeof(void*)*2, x_90); lean_ctor_set_float(x_93, sizeof(void*)*2 + 8, x_90); lean_ctor_set_uint8(x_93, sizeof(void*)*2 + 16, x_91); -x_94 = l_Lean_addTrace___at_Lean_Meta_Grind_pre___spec__2___closed__3; +x_94 = l_Lean_addTrace___at_Lean_Meta_Grind_simp___spec__3___closed__3; x_95 = lean_alloc_ctor(9, 3, 0); lean_ctor_set(x_95, 0, x_93); lean_ctor_set(x_95, 1, x_13); @@ -582,7 +694,7 @@ return x_104; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pre___lambda__1(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_simp___lambda__1(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { lean_object* x_13; lean_object* x_14; @@ -596,7 +708,7 @@ lean_ctor_set(x_14, 1, x_12); return x_14; } } -static lean_object* _init_l_Lean_Meta_Grind_pre___closed__1() { +static lean_object* _init_l_Lean_Meta_Grind_simp___closed__1() { _start: { lean_object* x_1; @@ -604,7 +716,7 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_unfoldReducible___lambda__2), return x_1; } } -static lean_object* _init_l_Lean_Meta_Grind_pre___closed__2() { +static lean_object* _init_l_Lean_Meta_Grind_simp___closed__2() { _start: { lean_object* x_1; @@ -612,7 +724,7 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_unfoldReducible___lambda__3__ return x_1; } } -static lean_object* _init_l_Lean_Meta_Grind_pre___closed__3() { +static lean_object* _init_l_Lean_Meta_Grind_simp___closed__3() { _start: { lean_object* x_1; @@ -620,7 +732,7 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_eraseIrrelevantMData___lambda return x_1; } } -static lean_object* _init_l_Lean_Meta_Grind_pre___closed__4() { +static lean_object* _init_l_Lean_Meta_Grind_simp___closed__4() { _start: { lean_object* x_1; @@ -628,7 +740,7 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_eraseIrrelevantMData___lambda return x_1; } } -static lean_object* _init_l_Lean_Meta_Grind_pre___closed__5() { +static lean_object* _init_l_Lean_Meta_Grind_simp___closed__5() { _start: { lean_object* x_1; @@ -636,7 +748,7 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_foldProjs___lambda__2___boxed return x_1; } } -static lean_object* _init_l_Lean_Meta_Grind_pre___closed__6() { +static lean_object* _init_l_Lean_Meta_Grind_simp___closed__6() { _start: { lean_object* x_1; @@ -644,7 +756,7 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_foldProjs___lambda__1), 6, 0) return x_1; } } -static lean_object* _init_l_Lean_Meta_Grind_pre___closed__7() { +static lean_object* _init_l_Lean_Meta_Grind_simp___closed__7() { _start: { lean_object* x_1; @@ -652,7 +764,7 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_normalizeLevels___lambda__1__ return x_1; } } -static lean_object* _init_l_Lean_Meta_Grind_pre___closed__8() { +static lean_object* _init_l_Lean_Meta_Grind_simp___closed__8() { _start: { lean_object* x_1; @@ -660,7 +772,7 @@ x_1 = lean_mk_string_unchecked("grind", 5, 5); return x_1; } } -static lean_object* _init_l_Lean_Meta_Grind_pre___closed__9() { +static lean_object* _init_l_Lean_Meta_Grind_simp___closed__9() { _start: { lean_object* x_1; @@ -668,26 +780,26 @@ x_1 = lean_mk_string_unchecked("simp", 4, 4); return x_1; } } -static lean_object* _init_l_Lean_Meta_Grind_pre___closed__10() { +static lean_object* _init_l_Lean_Meta_Grind_simp___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_Grind_pre___closed__8; -x_2 = l_Lean_Meta_Grind_pre___closed__9; +x_1 = l_Lean_Meta_Grind_simp___closed__8; +x_2 = l_Lean_Meta_Grind_simp___closed__9; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_Grind_pre___closed__11() { +static lean_object* _init_l_Lean_Meta_Grind_simp___closed__11() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_addTrace___at_Lean_Meta_Grind_pre___spec__2___closed__2; +x_1 = l_Lean_addTrace___at_Lean_Meta_Grind_simp___spec__3___closed__2; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Lean_Meta_Grind_pre___closed__12() { +static lean_object* _init_l_Lean_Meta_Grind_simp___closed__12() { _start: { lean_object* x_1; @@ -695,342 +807,334 @@ x_1 = lean_mk_string_unchecked("\n===>\n", 6, 6); return x_1; } } -static lean_object* _init_l_Lean_Meta_Grind_pre___closed__13() { +static lean_object* _init_l_Lean_Meta_Grind_simp___closed__13() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_Grind_pre___closed__12; +x_1 = l_Lean_Meta_Grind_simp___closed__12; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pre(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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_Grind_simp(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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_10; uint8_t x_11; +x_10 = l_Lean_instantiateMVars___at_Lean_Meta_Grind_simp___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_11 = !lean_is_exclusive(x_10); +if (x_11 == 0) +{ +lean_object* x_12; 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_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_3); -lean_inc(x_1); -x_10 = l_Lean_Meta_Grind_simp(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_10) == 0) -{ -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; -x_11 = lean_ctor_get(x_10, 0); -lean_inc(x_11); -x_12 = lean_ctor_get(x_10, 1); lean_inc(x_12); -lean_dec(x_10); -x_13 = lean_ctor_get(x_11, 0); -lean_inc(x_13); -x_14 = lean_ctor_get(x_11, 1); -lean_inc(x_14); -x_15 = lean_ctor_get_uint8(x_11, sizeof(void*)*2); -lean_dec(x_11); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -x_16 = l_Lean_Meta_Grind_markNestedProofsImpl(x_13, x_5, x_6, x_7, x_8, x_12); -if (lean_obj_tag(x_16) == 0) +x_14 = l_Lean_Meta_Grind_simpCore(x_12, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_13); +if (lean_obj_tag(x_14) == 0) { -lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_17 = lean_ctor_get(x_16, 0); +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; lean_object* x_20; +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +x_16 = lean_ctor_get(x_14, 1); +lean_inc(x_16); +lean_dec(x_14); +x_17 = lean_ctor_get(x_15, 0); lean_inc(x_17); -x_18 = lean_ctor_get(x_16, 1); +x_18 = lean_ctor_get(x_15, 1); lean_inc(x_18); -lean_dec(x_16); -x_19 = l_Lean_Meta_Grind_pre___closed__1; -x_20 = l_Lean_Meta_Grind_pre___closed__2; +x_19 = lean_ctor_get_uint8(x_15, sizeof(void*)*2); +lean_dec(x_15); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -x_21 = l_Lean_Core_transform___at_Lean_Meta_Grind_unfoldReducible___spec__1(x_17, x_19, x_20, x_5, x_6, x_7, x_8, x_18); -if (lean_obj_tag(x_21) == 0) +x_20 = l_Lean_Meta_Grind_abstractNestedProofs(x_17, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_16); +if (lean_obj_tag(x_20) == 0) { -lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_22 = lean_ctor_get(x_21, 0); +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_20, 0); +lean_inc(x_21); +x_22 = lean_ctor_get(x_20, 1); lean_inc(x_22); -x_23 = lean_ctor_get(x_21, 1); -lean_inc(x_23); -lean_dec(x_21); -x_24 = l_Lean_Meta_Grind_pre___closed__3; -x_25 = l_Lean_Meta_Grind_pre___closed__4; +lean_dec(x_20); lean_inc(x_8); lean_inc(x_7); -x_26 = l_Lean_Core_transform___at_Lean_Core_betaReduce___spec__1(x_22, x_24, x_25, x_7, x_8, x_23); -if (lean_obj_tag(x_26) == 0) +lean_inc(x_6); +lean_inc(x_5); +x_23 = l_Lean_Meta_Grind_markNestedProofsImpl(x_21, x_5, x_6, x_7, x_8, x_22); +if (lean_obj_tag(x_23) == 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; -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_Grind_pre___closed__5; -x_30 = l_Lean_Meta_Grind_pre___closed__6; -x_31 = 0; +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +x_25 = lean_ctor_get(x_23, 1); +lean_inc(x_25); +lean_dec(x_23); +x_26 = l_Lean_Meta_Grind_simp___closed__1; +x_27 = l_Lean_Meta_Grind_simp___closed__2; lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -x_32 = l_Lean_Meta_transform___at_Lean_Meta_zetaReduce___spec__1(x_27, x_29, x_30, x_31, x_31, x_5, x_6, x_7, x_8, x_28); -if (lean_obj_tag(x_32) == 0) +x_28 = l_Lean_Core_transform___at_Lean_Meta_Grind_unfoldReducible___spec__1(x_24, x_26, x_27, x_5, x_6, x_7, x_8, x_25); +if (lean_obj_tag(x_28) == 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 = l_Lean_Meta_Grind_pre___closed__7; +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); +x_30 = lean_ctor_get(x_28, 1); +lean_inc(x_30); +lean_dec(x_28); +x_31 = l_Lean_Meta_Grind_simp___closed__3; +x_32 = l_Lean_Meta_Grind_simp___closed__4; lean_inc(x_8); lean_inc(x_7); -x_36 = l_Lean_Core_transform___at_Lean_Core_betaReduce___spec__1(x_33, x_35, x_25, x_7, x_8, x_34); -if (lean_obj_tag(x_36) == 0) +x_33 = l_Lean_Core_transform___at_Lean_Core_betaReduce___spec__1(x_29, x_31, x_32, x_7, x_8, x_30); +if (lean_obj_tag(x_33) == 0) { -lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_37 = lean_ctor_get(x_36, 0); -lean_inc(x_37); -x_38 = lean_ctor_get(x_36, 1); -lean_inc(x_38); -lean_dec(x_36); +lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; uint8_t x_38; lean_object* x_39; +x_34 = lean_ctor_get(x_33, 0); +lean_inc(x_34); +x_35 = lean_ctor_get(x_33, 1); +lean_inc(x_35); +lean_dec(x_33); +x_36 = l_Lean_Meta_Grind_simp___closed__5; +x_37 = l_Lean_Meta_Grind_simp___closed__6; +x_38 = 0; lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -x_39 = l_Lean_Meta_Grind_canon(x_37, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_38); +x_39 = l_Lean_Meta_transform___at_Lean_Meta_zetaReduce___spec__1(x_34, x_36, x_37, x_38, x_38, x_5, x_6, x_7, x_8, x_35); if (lean_obj_tag(x_39) == 0) { -lean_object* x_40; lean_object* x_41; lean_object* x_42; uint8_t x_43; +lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; x_40 = lean_ctor_get(x_39, 0); lean_inc(x_40); x_41 = lean_ctor_get(x_39, 1); lean_inc(x_41); lean_dec(x_39); -x_42 = l_Lean_Meta_Grind_shareCommon(x_40, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_41); -x_43 = !lean_is_exclusive(x_42); -if (x_43 == 0) +x_42 = l_Lean_Meta_Grind_simp___closed__7; +lean_inc(x_8); +lean_inc(x_7); +x_43 = l_Lean_Core_transform___at_Lean_Core_betaReduce___spec__1(x_40, x_42, x_32, x_7, x_8, x_41); +if (lean_obj_tag(x_43) == 0) { -lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; uint8_t x_49; -x_44 = lean_ctor_get(x_42, 0); -x_45 = lean_ctor_get(x_42, 1); -x_46 = l_Lean_Meta_Grind_pre___closed__10; -x_47 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_pre___spec__1(x_46, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_45); -x_48 = lean_ctor_get(x_47, 0); +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_43, 0); +lean_inc(x_44); +x_45 = lean_ctor_get(x_43, 1); +lean_inc(x_45); +lean_dec(x_43); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +x_46 = l_Lean_Meta_Grind_canon(x_44, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_45); +if (lean_obj_tag(x_46) == 0) +{ +lean_object* x_47; lean_object* x_48; lean_object* x_49; uint8_t x_50; +x_47 = lean_ctor_get(x_46, 0); +lean_inc(x_47); +x_48 = lean_ctor_get(x_46, 1); 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; lean_object* x_52; -lean_free_object(x_42); -lean_dec(x_1); -x_50 = lean_ctor_get(x_47, 1); -lean_inc(x_50); -lean_dec(x_47); -x_51 = lean_box(0); -x_52 = l_Lean_Meta_Grind_pre___lambda__1(x_44, x_14, x_15, x_51, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_50); +lean_dec(x_46); +x_49 = l_Lean_Meta_Grind_shareCommon(x_47, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_48); +x_50 = !lean_is_exclusive(x_49); +if (x_50 == 0) +{ +lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; uint8_t x_56; +x_51 = lean_ctor_get(x_49, 0); +x_52 = lean_ctor_get(x_49, 1); +x_53 = l_Lean_Meta_Grind_simp___closed__10; +x_54 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_simp___spec__2(x_53, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_52); +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_free_object(x_49); +lean_free_object(x_10); +lean_dec(x_12); +x_57 = lean_ctor_get(x_54, 1); +lean_inc(x_57); +lean_dec(x_54); +x_58 = lean_box(0); +x_59 = l_Lean_Meta_Grind_simp___lambda__1(x_51, x_18, x_19, x_58, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_57); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -return x_52; +return x_59; } else { -uint8_t x_53; -x_53 = !lean_is_exclusive(x_47); -if (x_53 == 0) +uint8_t x_60; +x_60 = !lean_is_exclusive(x_54); +if (x_60 == 0) { -lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; -x_54 = lean_ctor_get(x_47, 1); -x_55 = lean_ctor_get(x_47, 0); -lean_dec(x_55); -x_56 = l_Lean_MessageData_ofExpr(x_1); -x_57 = l_Lean_Meta_Grind_pre___closed__11; -lean_ctor_set_tag(x_47, 7); -lean_ctor_set(x_47, 1, x_56); -lean_ctor_set(x_47, 0, x_57); -x_58 = l_Lean_Meta_Grind_pre___closed__13; -lean_ctor_set_tag(x_42, 7); -lean_ctor_set(x_42, 1, x_58); -lean_ctor_set(x_42, 0, x_47); -lean_inc(x_44); -x_59 = l_Lean_MessageData_ofExpr(x_44); -x_60 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_60, 0, x_42); -lean_ctor_set(x_60, 1, x_59); -x_61 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_61, 0, x_60); -lean_ctor_set(x_61, 1, x_57); -x_62 = l_Lean_addTrace___at_Lean_Meta_Grind_pre___spec__2(x_46, x_61, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_54); -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_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; +x_61 = lean_ctor_get(x_54, 1); +x_62 = lean_ctor_get(x_54, 0); lean_dec(x_62); -x_65 = l_Lean_Meta_Grind_pre___lambda__1(x_44, x_14, x_15, x_63, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_64); +x_63 = l_Lean_MessageData_ofExpr(x_12); +x_64 = l_Lean_Meta_Grind_simp___closed__11; +lean_ctor_set_tag(x_54, 7); +lean_ctor_set(x_54, 1, x_63); +lean_ctor_set(x_54, 0, x_64); +x_65 = l_Lean_Meta_Grind_simp___closed__13; +lean_ctor_set_tag(x_49, 7); +lean_ctor_set(x_49, 1, x_65); +lean_ctor_set(x_49, 0, x_54); +lean_inc(x_51); +x_66 = l_Lean_MessageData_ofExpr(x_51); +lean_ctor_set_tag(x_10, 7); +lean_ctor_set(x_10, 1, x_66); +lean_ctor_set(x_10, 0, x_49); +x_67 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_67, 0, x_10); +lean_ctor_set(x_67, 1, x_64); +x_68 = l_Lean_addTrace___at_Lean_Meta_Grind_simp___spec__3(x_53, x_67, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_61); +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_Grind_simp___lambda__1(x_51, x_18, x_19, x_69, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_70); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -lean_dec(x_63); -return x_65; +lean_dec(x_69); +return x_71; } else { -lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; -x_66 = lean_ctor_get(x_47, 1); -lean_inc(x_66); -lean_dec(x_47); -x_67 = l_Lean_MessageData_ofExpr(x_1); -x_68 = l_Lean_Meta_Grind_pre___closed__11; -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_Meta_Grind_pre___closed__13; -lean_ctor_set_tag(x_42, 7); -lean_ctor_set(x_42, 1, x_70); -lean_ctor_set(x_42, 0, x_69); -lean_inc(x_44); -x_71 = l_Lean_MessageData_ofExpr(x_44); -x_72 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_72, 0, x_42); -lean_ctor_set(x_72, 1, x_71); -x_73 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_73, 0, x_72); -lean_ctor_set(x_73, 1, x_68); -x_74 = l_Lean_addTrace___at_Lean_Meta_Grind_pre___spec__2(x_46, x_73, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_66); -x_75 = lean_ctor_get(x_74, 0); -lean_inc(x_75); -x_76 = lean_ctor_get(x_74, 1); -lean_inc(x_76); -lean_dec(x_74); -x_77 = l_Lean_Meta_Grind_pre___lambda__1(x_44, x_14, x_15, x_75, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_76); +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_72 = lean_ctor_get(x_54, 1); +lean_inc(x_72); +lean_dec(x_54); +x_73 = l_Lean_MessageData_ofExpr(x_12); +x_74 = l_Lean_Meta_Grind_simp___closed__11; +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_Meta_Grind_simp___closed__13; +lean_ctor_set_tag(x_49, 7); +lean_ctor_set(x_49, 1, x_76); +lean_ctor_set(x_49, 0, x_75); +lean_inc(x_51); +x_77 = l_Lean_MessageData_ofExpr(x_51); +lean_ctor_set_tag(x_10, 7); +lean_ctor_set(x_10, 1, x_77); +lean_ctor_set(x_10, 0, x_49); +x_78 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_78, 0, x_10); +lean_ctor_set(x_78, 1, x_74); +x_79 = l_Lean_addTrace___at_Lean_Meta_Grind_simp___spec__3(x_53, x_78, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_72); +x_80 = lean_ctor_get(x_79, 0); +lean_inc(x_80); +x_81 = lean_ctor_get(x_79, 1); +lean_inc(x_81); +lean_dec(x_79); +x_82 = l_Lean_Meta_Grind_simp___lambda__1(x_51, x_18, x_19, x_80, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_81); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -lean_dec(x_75); -return x_77; +lean_dec(x_80); +return x_82; } } } else { -lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; uint8_t x_83; -x_78 = lean_ctor_get(x_42, 0); -x_79 = lean_ctor_get(x_42, 1); -lean_inc(x_79); -lean_inc(x_78); -lean_dec(x_42); -x_80 = l_Lean_Meta_Grind_pre___closed__10; -x_81 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_pre___spec__1(x_80, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_79); -x_82 = lean_ctor_get(x_81, 0); -lean_inc(x_82); -x_83 = lean_unbox(x_82); -lean_dec(x_82); -if (x_83 == 0) -{ -lean_object* x_84; lean_object* x_85; lean_object* x_86; -lean_dec(x_1); -x_84 = lean_ctor_get(x_81, 1); +lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; uint8_t x_88; +x_83 = lean_ctor_get(x_49, 0); +x_84 = lean_ctor_get(x_49, 1); lean_inc(x_84); -lean_dec(x_81); -x_85 = lean_box(0); -x_86 = l_Lean_Meta_Grind_pre___lambda__1(x_78, x_14, x_15, x_85, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_84); +lean_inc(x_83); +lean_dec(x_49); +x_85 = l_Lean_Meta_Grind_simp___closed__10; +x_86 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_simp___spec__2(x_85, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_84); +x_87 = lean_ctor_get(x_86, 0); +lean_inc(x_87); +x_88 = lean_unbox(x_87); +lean_dec(x_87); +if (x_88 == 0) +{ +lean_object* x_89; lean_object* x_90; lean_object* x_91; +lean_free_object(x_10); +lean_dec(x_12); +x_89 = lean_ctor_get(x_86, 1); +lean_inc(x_89); +lean_dec(x_86); +x_90 = lean_box(0); +x_91 = l_Lean_Meta_Grind_simp___lambda__1(x_83, x_18, x_19, x_90, x_2, 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); -return x_86; +return x_91; } else { -lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; -x_87 = lean_ctor_get(x_81, 1); -lean_inc(x_87); -if (lean_is_exclusive(x_81)) { - lean_ctor_release(x_81, 0); - lean_ctor_release(x_81, 1); - x_88 = x_81; +lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; +x_92 = lean_ctor_get(x_86, 1); +lean_inc(x_92); +if (lean_is_exclusive(x_86)) { + lean_ctor_release(x_86, 0); + lean_ctor_release(x_86, 1); + x_93 = x_86; } else { - lean_dec_ref(x_81); - x_88 = lean_box(0); + lean_dec_ref(x_86); + x_93 = lean_box(0); } -x_89 = l_Lean_MessageData_ofExpr(x_1); -x_90 = l_Lean_Meta_Grind_pre___closed__11; -if (lean_is_scalar(x_88)) { - x_91 = lean_alloc_ctor(7, 2, 0); +x_94 = l_Lean_MessageData_ofExpr(x_12); +x_95 = l_Lean_Meta_Grind_simp___closed__11; +if (lean_is_scalar(x_93)) { + x_96 = lean_alloc_ctor(7, 2, 0); } else { - x_91 = x_88; - lean_ctor_set_tag(x_91, 7); -} -lean_ctor_set(x_91, 0, x_90); -lean_ctor_set(x_91, 1, x_89); -x_92 = l_Lean_Meta_Grind_pre___closed__13; -x_93 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_93, 0, x_91); -lean_ctor_set(x_93, 1, x_92); -lean_inc(x_78); -x_94 = l_Lean_MessageData_ofExpr(x_78); -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 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_96, 0, x_95); -lean_ctor_set(x_96, 1, x_90); -x_97 = l_Lean_addTrace___at_Lean_Meta_Grind_pre___spec__2(x_80, x_96, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_87); -x_98 = lean_ctor_get(x_97, 0); -lean_inc(x_98); -x_99 = lean_ctor_get(x_97, 1); -lean_inc(x_99); -lean_dec(x_97); -x_100 = l_Lean_Meta_Grind_pre___lambda__1(x_78, x_14, x_15, x_98, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_99); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_98); -return x_100; -} + x_96 = x_93; + lean_ctor_set_tag(x_96, 7); } -} -else -{ -uint8_t x_101; -lean_dec(x_14); +lean_ctor_set(x_96, 0, x_95); +lean_ctor_set(x_96, 1, x_94); +x_97 = l_Lean_Meta_Grind_simp___closed__13; +x_98 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_98, 0, x_96); +lean_ctor_set(x_98, 1, x_97); +lean_inc(x_83); +x_99 = l_Lean_MessageData_ofExpr(x_83); +lean_ctor_set_tag(x_10, 7); +lean_ctor_set(x_10, 1, x_99); +lean_ctor_set(x_10, 0, x_98); +x_100 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_100, 0, x_10); +lean_ctor_set(x_100, 1, x_95); +x_101 = l_Lean_addTrace___at_Lean_Meta_Grind_simp___spec__3(x_85, x_100, x_2, x_3, x_4, x_5, x_6, x_7, x_8, 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_104 = l_Lean_Meta_Grind_simp___lambda__1(x_83, x_18, x_19, x_102, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_103); 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_101 = !lean_is_exclusive(x_39); -if (x_101 == 0) -{ -return x_39; -} -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); +lean_dec(x_102); return x_104; } } @@ -1038,26 +1142,27 @@ return x_104; else { uint8_t x_105; -lean_dec(x_14); +lean_dec(x_18); +lean_free_object(x_10); +lean_dec(x_12); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -lean_dec(x_1); -x_105 = !lean_is_exclusive(x_36); +x_105 = !lean_is_exclusive(x_46); if (x_105 == 0) { -return x_36; +return x_46; } else { lean_object* x_106; lean_object* x_107; lean_object* x_108; -x_106 = lean_ctor_get(x_36, 0); -x_107 = lean_ctor_get(x_36, 1); +x_106 = lean_ctor_get(x_46, 0); +x_107 = lean_ctor_get(x_46, 1); lean_inc(x_107); lean_inc(x_106); -lean_dec(x_36); +lean_dec(x_46); x_108 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_108, 0, x_106); lean_ctor_set(x_108, 1, x_107); @@ -1068,26 +1173,27 @@ return x_108; else { uint8_t x_109; -lean_dec(x_14); +lean_dec(x_18); +lean_free_object(x_10); +lean_dec(x_12); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -lean_dec(x_1); -x_109 = !lean_is_exclusive(x_32); +x_109 = !lean_is_exclusive(x_43); if (x_109 == 0) { -return x_32; +return x_43; } else { lean_object* x_110; lean_object* x_111; lean_object* x_112; -x_110 = lean_ctor_get(x_32, 0); -x_111 = lean_ctor_get(x_32, 1); +x_110 = lean_ctor_get(x_43, 0); +x_111 = lean_ctor_get(x_43, 1); lean_inc(x_111); lean_inc(x_110); -lean_dec(x_32); +lean_dec(x_43); x_112 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_112, 0, x_110); lean_ctor_set(x_112, 1, x_111); @@ -1098,26 +1204,27 @@ return x_112; else { uint8_t x_113; -lean_dec(x_14); +lean_dec(x_18); +lean_free_object(x_10); +lean_dec(x_12); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -lean_dec(x_1); -x_113 = !lean_is_exclusive(x_26); +x_113 = !lean_is_exclusive(x_39); if (x_113 == 0) { -return x_26; +return x_39; } else { lean_object* x_114; lean_object* x_115; lean_object* x_116; -x_114 = lean_ctor_get(x_26, 0); -x_115 = lean_ctor_get(x_26, 1); +x_114 = lean_ctor_get(x_39, 0); +x_115 = lean_ctor_get(x_39, 1); lean_inc(x_115); lean_inc(x_114); -lean_dec(x_26); +lean_dec(x_39); x_116 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_116, 0, x_114); lean_ctor_set(x_116, 1, x_115); @@ -1128,26 +1235,27 @@ return x_116; else { uint8_t x_117; -lean_dec(x_14); +lean_dec(x_18); +lean_free_object(x_10); +lean_dec(x_12); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -lean_dec(x_1); -x_117 = !lean_is_exclusive(x_21); +x_117 = !lean_is_exclusive(x_33); if (x_117 == 0) { -return x_21; +return x_33; } else { lean_object* x_118; lean_object* x_119; lean_object* x_120; -x_118 = lean_ctor_get(x_21, 0); -x_119 = lean_ctor_get(x_21, 1); +x_118 = lean_ctor_get(x_33, 0); +x_119 = lean_ctor_get(x_33, 1); lean_inc(x_119); lean_inc(x_118); -lean_dec(x_21); +lean_dec(x_33); x_120 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_120, 0, x_118); lean_ctor_set(x_120, 1, x_119); @@ -1158,26 +1266,27 @@ return x_120; else { uint8_t x_121; -lean_dec(x_14); +lean_dec(x_18); +lean_free_object(x_10); +lean_dec(x_12); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -lean_dec(x_1); -x_121 = !lean_is_exclusive(x_16); +x_121 = !lean_is_exclusive(x_28); if (x_121 == 0) { -return x_16; +return x_28; } else { lean_object* x_122; lean_object* x_123; lean_object* x_124; -x_122 = lean_ctor_get(x_16, 0); -x_123 = lean_ctor_get(x_16, 1); +x_122 = lean_ctor_get(x_28, 0); +x_123 = lean_ctor_get(x_28, 1); lean_inc(x_123); lean_inc(x_122); -lean_dec(x_16); +lean_dec(x_28); x_124 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_124, 0, x_122); lean_ctor_set(x_124, 1, x_123); @@ -1188,25 +1297,27 @@ return x_124; else { uint8_t x_125; +lean_dec(x_18); +lean_free_object(x_10); +lean_dec(x_12); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -lean_dec(x_1); -x_125 = !lean_is_exclusive(x_10); +x_125 = !lean_is_exclusive(x_23); if (x_125 == 0) { -return x_10; +return x_23; } else { lean_object* x_126; lean_object* x_127; lean_object* x_128; -x_126 = lean_ctor_get(x_10, 0); -x_127 = lean_ctor_get(x_10, 1); +x_126 = lean_ctor_get(x_23, 0); +x_127 = lean_ctor_get(x_23, 1); lean_inc(x_127); lean_inc(x_126); -lean_dec(x_10); +lean_dec(x_23); x_128 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_128, 0, x_126); lean_ctor_set(x_128, 1, x_127); @@ -1214,12 +1325,560 @@ return x_128; } } } +else +{ +uint8_t x_129; +lean_dec(x_18); +lean_free_object(x_10); +lean_dec(x_12); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +x_129 = !lean_is_exclusive(x_20); +if (x_129 == 0) +{ +return x_20; +} +else +{ +lean_object* x_130; lean_object* x_131; lean_object* x_132; +x_130 = lean_ctor_get(x_20, 0); +x_131 = lean_ctor_get(x_20, 1); +lean_inc(x_131); +lean_inc(x_130); +lean_dec(x_20); +x_132 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_132, 0, x_130); +lean_ctor_set(x_132, 1, x_131); +return x_132; +} +} +} +else +{ +uint8_t x_133; +lean_free_object(x_10); +lean_dec(x_12); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +x_133 = !lean_is_exclusive(x_14); +if (x_133 == 0) +{ +return x_14; } -LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_pre___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) { +else +{ +lean_object* x_134; lean_object* x_135; lean_object* x_136; +x_134 = lean_ctor_get(x_14, 0); +x_135 = lean_ctor_get(x_14, 1); +lean_inc(x_135); +lean_inc(x_134); +lean_dec(x_14); +x_136 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_136, 0, x_134); +lean_ctor_set(x_136, 1, x_135); +return x_136; +} +} +} +else +{ +lean_object* x_137; lean_object* x_138; lean_object* x_139; +x_137 = lean_ctor_get(x_10, 0); +x_138 = lean_ctor_get(x_10, 1); +lean_inc(x_138); +lean_inc(x_137); +lean_dec(x_10); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_3); +lean_inc(x_137); +x_139 = l_Lean_Meta_Grind_simpCore(x_137, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_138); +if (lean_obj_tag(x_139) == 0) +{ +lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; uint8_t x_144; lean_object* x_145; +x_140 = lean_ctor_get(x_139, 0); +lean_inc(x_140); +x_141 = lean_ctor_get(x_139, 1); +lean_inc(x_141); +lean_dec(x_139); +x_142 = lean_ctor_get(x_140, 0); +lean_inc(x_142); +x_143 = lean_ctor_get(x_140, 1); +lean_inc(x_143); +x_144 = lean_ctor_get_uint8(x_140, sizeof(void*)*2); +lean_dec(x_140); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +x_145 = l_Lean_Meta_Grind_abstractNestedProofs(x_142, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_141); +if (lean_obj_tag(x_145) == 0) +{ +lean_object* x_146; lean_object* x_147; lean_object* x_148; +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); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +x_148 = l_Lean_Meta_Grind_markNestedProofsImpl(x_146, x_5, x_6, x_7, x_8, x_147); +if (lean_obj_tag(x_148) == 0) +{ +lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; +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 = l_Lean_Meta_Grind_simp___closed__1; +x_152 = l_Lean_Meta_Grind_simp___closed__2; +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +x_153 = l_Lean_Core_transform___at_Lean_Meta_Grind_unfoldReducible___spec__1(x_149, x_151, x_152, x_5, x_6, x_7, x_8, x_150); +if (lean_obj_tag(x_153) == 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, 0); +lean_inc(x_154); +x_155 = lean_ctor_get(x_153, 1); +lean_inc(x_155); +lean_dec(x_153); +x_156 = l_Lean_Meta_Grind_simp___closed__3; +x_157 = l_Lean_Meta_Grind_simp___closed__4; +lean_inc(x_8); +lean_inc(x_7); +x_158 = l_Lean_Core_transform___at_Lean_Core_betaReduce___spec__1(x_154, x_156, x_157, x_7, x_8, x_155); +if (lean_obj_tag(x_158) == 0) +{ +lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; uint8_t x_163; lean_object* x_164; +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_Grind_simp___closed__5; +x_162 = l_Lean_Meta_Grind_simp___closed__6; +x_163 = 0; +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +x_164 = l_Lean_Meta_transform___at_Lean_Meta_zetaReduce___spec__1(x_159, x_161, x_162, x_163, x_163, x_5, x_6, x_7, x_8, x_160); +if (lean_obj_tag(x_164) == 0) +{ +lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; +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_Meta_Grind_simp___closed__7; +lean_inc(x_8); +lean_inc(x_7); +x_168 = l_Lean_Core_transform___at_Lean_Core_betaReduce___spec__1(x_165, x_167, x_157, x_7, x_8, x_166); +if (lean_obj_tag(x_168) == 0) +{ +lean_object* x_169; lean_object* x_170; lean_object* x_171; +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_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +x_171 = l_Lean_Meta_Grind_canon(x_169, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_170); +if (lean_obj_tag(x_171) == 0) +{ +lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; uint8_t x_181; +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_Grind_shareCommon(x_172, x_2, x_3, x_4, x_5, x_6, x_7, x_8, 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); +if (lean_is_exclusive(x_174)) { + lean_ctor_release(x_174, 0); + lean_ctor_release(x_174, 1); + x_177 = x_174; +} else { + lean_dec_ref(x_174); + x_177 = lean_box(0); +} +x_178 = l_Lean_Meta_Grind_simp___closed__10; +x_179 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_simp___spec__2(x_178, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_176); +x_180 = lean_ctor_get(x_179, 0); +lean_inc(x_180); +x_181 = lean_unbox(x_180); +lean_dec(x_180); +if (x_181 == 0) +{ +lean_object* x_182; lean_object* x_183; lean_object* x_184; +lean_dec(x_177); +lean_dec(x_137); +x_182 = lean_ctor_get(x_179, 1); +lean_inc(x_182); +lean_dec(x_179); +x_183 = lean_box(0); +x_184 = l_Lean_Meta_Grind_simp___lambda__1(x_175, x_143, x_144, x_183, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_182); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +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; lean_object* x_196; lean_object* x_197; lean_object* x_198; +x_185 = lean_ctor_get(x_179, 1); +lean_inc(x_185); +if (lean_is_exclusive(x_179)) { + lean_ctor_release(x_179, 0); + lean_ctor_release(x_179, 1); + x_186 = x_179; +} else { + lean_dec_ref(x_179); + x_186 = lean_box(0); +} +x_187 = l_Lean_MessageData_ofExpr(x_137); +x_188 = l_Lean_Meta_Grind_simp___closed__11; +if (lean_is_scalar(x_186)) { + x_189 = lean_alloc_ctor(7, 2, 0); +} else { + x_189 = x_186; + lean_ctor_set_tag(x_189, 7); +} +lean_ctor_set(x_189, 0, x_188); +lean_ctor_set(x_189, 1, x_187); +x_190 = l_Lean_Meta_Grind_simp___closed__13; +if (lean_is_scalar(x_177)) { + x_191 = lean_alloc_ctor(7, 2, 0); +} else { + x_191 = x_177; + lean_ctor_set_tag(x_191, 7); +} +lean_ctor_set(x_191, 0, x_189); +lean_ctor_set(x_191, 1, x_190); +lean_inc(x_175); +x_192 = l_Lean_MessageData_ofExpr(x_175); +x_193 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_193, 0, x_191); +lean_ctor_set(x_193, 1, x_192); +x_194 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_194, 0, x_193); +lean_ctor_set(x_194, 1, x_188); +x_195 = l_Lean_addTrace___at_Lean_Meta_Grind_simp___spec__3(x_178, x_194, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_185); +x_196 = lean_ctor_get(x_195, 0); +lean_inc(x_196); +x_197 = lean_ctor_get(x_195, 1); +lean_inc(x_197); +lean_dec(x_195); +x_198 = l_Lean_Meta_Grind_simp___lambda__1(x_175, x_143, x_144, x_196, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_197); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_196); +return x_198; +} +} +else +{ +lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; +lean_dec(x_143); +lean_dec(x_137); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +x_199 = lean_ctor_get(x_171, 0); +lean_inc(x_199); +x_200 = lean_ctor_get(x_171, 1); +lean_inc(x_200); +if (lean_is_exclusive(x_171)) { + lean_ctor_release(x_171, 0); + lean_ctor_release(x_171, 1); + x_201 = x_171; +} else { + lean_dec_ref(x_171); + x_201 = lean_box(0); +} +if (lean_is_scalar(x_201)) { + x_202 = lean_alloc_ctor(1, 2, 0); +} else { + x_202 = x_201; +} +lean_ctor_set(x_202, 0, x_199); +lean_ctor_set(x_202, 1, x_200); +return x_202; +} +} +else +{ +lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; +lean_dec(x_143); +lean_dec(x_137); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +x_203 = lean_ctor_get(x_168, 0); +lean_inc(x_203); +x_204 = lean_ctor_get(x_168, 1); +lean_inc(x_204); +if (lean_is_exclusive(x_168)) { + lean_ctor_release(x_168, 0); + lean_ctor_release(x_168, 1); + x_205 = x_168; +} else { + lean_dec_ref(x_168); + x_205 = lean_box(0); +} +if (lean_is_scalar(x_205)) { + x_206 = lean_alloc_ctor(1, 2, 0); +} else { + x_206 = x_205; +} +lean_ctor_set(x_206, 0, x_203); +lean_ctor_set(x_206, 1, x_204); +return x_206; +} +} +else +{ +lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; +lean_dec(x_143); +lean_dec(x_137); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +x_207 = lean_ctor_get(x_164, 0); +lean_inc(x_207); +x_208 = lean_ctor_get(x_164, 1); +lean_inc(x_208); +if (lean_is_exclusive(x_164)) { + lean_ctor_release(x_164, 0); + lean_ctor_release(x_164, 1); + x_209 = x_164; +} else { + lean_dec_ref(x_164); + 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_143); +lean_dec(x_137); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +x_211 = lean_ctor_get(x_158, 0); +lean_inc(x_211); +x_212 = lean_ctor_get(x_158, 1); +lean_inc(x_212); +if (lean_is_exclusive(x_158)) { + lean_ctor_release(x_158, 0); + lean_ctor_release(x_158, 1); + x_213 = x_158; +} else { + lean_dec_ref(x_158); + 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_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; +lean_dec(x_143); +lean_dec(x_137); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +x_215 = lean_ctor_get(x_153, 0); +lean_inc(x_215); +x_216 = lean_ctor_get(x_153, 1); +lean_inc(x_216); +if (lean_is_exclusive(x_153)) { + lean_ctor_release(x_153, 0); + lean_ctor_release(x_153, 1); + x_217 = x_153; +} else { + lean_dec_ref(x_153); + x_217 = lean_box(0); +} +if (lean_is_scalar(x_217)) { + x_218 = lean_alloc_ctor(1, 2, 0); +} else { + x_218 = x_217; +} +lean_ctor_set(x_218, 0, x_215); +lean_ctor_set(x_218, 1, x_216); +return x_218; +} +} +else +{ +lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; +lean_dec(x_143); +lean_dec(x_137); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +x_219 = lean_ctor_get(x_148, 0); +lean_inc(x_219); +x_220 = lean_ctor_get(x_148, 1); +lean_inc(x_220); +if (lean_is_exclusive(x_148)) { + lean_ctor_release(x_148, 0); + lean_ctor_release(x_148, 1); + x_221 = x_148; +} else { + lean_dec_ref(x_148); + x_221 = lean_box(0); +} +if (lean_is_scalar(x_221)) { + x_222 = lean_alloc_ctor(1, 2, 0); +} else { + x_222 = x_221; +} +lean_ctor_set(x_222, 0, x_219); +lean_ctor_set(x_222, 1, x_220); +return x_222; +} +} +else +{ +lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; +lean_dec(x_143); +lean_dec(x_137); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +x_223 = lean_ctor_get(x_145, 0); +lean_inc(x_223); +x_224 = lean_ctor_get(x_145, 1); +lean_inc(x_224); +if (lean_is_exclusive(x_145)) { + lean_ctor_release(x_145, 0); + lean_ctor_release(x_145, 1); + x_225 = x_145; +} else { + lean_dec_ref(x_145); + x_225 = lean_box(0); +} +if (lean_is_scalar(x_225)) { + x_226 = lean_alloc_ctor(1, 2, 0); +} else { + x_226 = x_225; +} +lean_ctor_set(x_226, 0, x_223); +lean_ctor_set(x_226, 1, x_224); +return x_226; +} +} +else +{ +lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; +lean_dec(x_137); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +x_227 = lean_ctor_get(x_139, 0); +lean_inc(x_227); +x_228 = lean_ctor_get(x_139, 1); +lean_inc(x_228); +if (lean_is_exclusive(x_139)) { + lean_ctor_release(x_139, 0); + lean_ctor_release(x_139, 1); + x_229 = x_139; +} else { + lean_dec_ref(x_139); + x_229 = lean_box(0); +} +if (lean_is_scalar(x_229)) { + x_230 = lean_alloc_ctor(1, 2, 0); +} else { + x_230 = x_229; +} +lean_ctor_set(x_230, 0, x_227); +lean_ctor_set(x_230, 1, x_228); +return x_230; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at_Lean_Meta_Grind_simp___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_instantiateMVars___at_Lean_Meta_Grind_simp___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_isTracingEnabledFor___at_Lean_Meta_Grind_simp___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { lean_object* x_10; -x_10 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_pre___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_10 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_simp___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -1230,11 +1889,11 @@ lean_dec(x_2); return x_10; } } -LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Meta_Grind_pre___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Meta_Grind_simp___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { lean_object* x_11; -x_11 = l_Lean_addTrace___at_Lean_Meta_Grind_pre___spec__2(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_Lean_Meta_Grind_simp___spec__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -1245,13 +1904,13 @@ lean_dec(x_3); return x_11; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pre___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 lean_object* l_Lean_Meta_Grind_simp___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { uint8_t x_13; lean_object* x_14; x_13 = lean_unbox(x_3); lean_dec(x_3); -x_14 = l_Lean_Meta_Grind_pre___lambda__1(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_14 = l_Lean_Meta_Grind_simp___lambda__1(x_1, x_2, x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -1263,16 +1922,18 @@ lean_dec(x_4); return x_14; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pre___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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_Grind_simp___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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_Grind_pre(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_10 = l_Lean_Meta_Grind_simp(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); lean_dec(x_4); lean_dec(x_2); return x_10; } } +lean_object* initialize_Init_Grind_Lemmas(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Meta_Tactic_Assert(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_Tactic_Simp_Main(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_Tactic_Grind_Util(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_Tactic_Grind_Types(uint8_t builtin, lean_object*); @@ -1282,6 +1943,12 @@ LEAN_EXPORT lean_object* initialize_Lean_Meta_Tactic_Grind_Simp(uint8_t builtin, lean_object * res; if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); _G_initialized = true; +res = initialize_Init_Grind_Lemmas(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Lean_Meta_Tactic_Assert(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); res = initialize_Lean_Meta_Tactic_Simp_Main(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); @@ -1294,37 +1961,37 @@ lean_dec_ref(res); res = initialize_Lean_Meta_Tactic_Grind_MarkNestedProofs(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l_Lean_addTrace___at_Lean_Meta_Grind_pre___spec__2___closed__1 = _init_l_Lean_addTrace___at_Lean_Meta_Grind_pre___spec__2___closed__1(); -l_Lean_addTrace___at_Lean_Meta_Grind_pre___spec__2___closed__2 = _init_l_Lean_addTrace___at_Lean_Meta_Grind_pre___spec__2___closed__2(); -lean_mark_persistent(l_Lean_addTrace___at_Lean_Meta_Grind_pre___spec__2___closed__2); -l_Lean_addTrace___at_Lean_Meta_Grind_pre___spec__2___closed__3 = _init_l_Lean_addTrace___at_Lean_Meta_Grind_pre___spec__2___closed__3(); -lean_mark_persistent(l_Lean_addTrace___at_Lean_Meta_Grind_pre___spec__2___closed__3); -l_Lean_Meta_Grind_pre___closed__1 = _init_l_Lean_Meta_Grind_pre___closed__1(); -lean_mark_persistent(l_Lean_Meta_Grind_pre___closed__1); -l_Lean_Meta_Grind_pre___closed__2 = _init_l_Lean_Meta_Grind_pre___closed__2(); -lean_mark_persistent(l_Lean_Meta_Grind_pre___closed__2); -l_Lean_Meta_Grind_pre___closed__3 = _init_l_Lean_Meta_Grind_pre___closed__3(); -lean_mark_persistent(l_Lean_Meta_Grind_pre___closed__3); -l_Lean_Meta_Grind_pre___closed__4 = _init_l_Lean_Meta_Grind_pre___closed__4(); -lean_mark_persistent(l_Lean_Meta_Grind_pre___closed__4); -l_Lean_Meta_Grind_pre___closed__5 = _init_l_Lean_Meta_Grind_pre___closed__5(); -lean_mark_persistent(l_Lean_Meta_Grind_pre___closed__5); -l_Lean_Meta_Grind_pre___closed__6 = _init_l_Lean_Meta_Grind_pre___closed__6(); -lean_mark_persistent(l_Lean_Meta_Grind_pre___closed__6); -l_Lean_Meta_Grind_pre___closed__7 = _init_l_Lean_Meta_Grind_pre___closed__7(); -lean_mark_persistent(l_Lean_Meta_Grind_pre___closed__7); -l_Lean_Meta_Grind_pre___closed__8 = _init_l_Lean_Meta_Grind_pre___closed__8(); -lean_mark_persistent(l_Lean_Meta_Grind_pre___closed__8); -l_Lean_Meta_Grind_pre___closed__9 = _init_l_Lean_Meta_Grind_pre___closed__9(); -lean_mark_persistent(l_Lean_Meta_Grind_pre___closed__9); -l_Lean_Meta_Grind_pre___closed__10 = _init_l_Lean_Meta_Grind_pre___closed__10(); -lean_mark_persistent(l_Lean_Meta_Grind_pre___closed__10); -l_Lean_Meta_Grind_pre___closed__11 = _init_l_Lean_Meta_Grind_pre___closed__11(); -lean_mark_persistent(l_Lean_Meta_Grind_pre___closed__11); -l_Lean_Meta_Grind_pre___closed__12 = _init_l_Lean_Meta_Grind_pre___closed__12(); -lean_mark_persistent(l_Lean_Meta_Grind_pre___closed__12); -l_Lean_Meta_Grind_pre___closed__13 = _init_l_Lean_Meta_Grind_pre___closed__13(); -lean_mark_persistent(l_Lean_Meta_Grind_pre___closed__13); +l_Lean_addTrace___at_Lean_Meta_Grind_simp___spec__3___closed__1 = _init_l_Lean_addTrace___at_Lean_Meta_Grind_simp___spec__3___closed__1(); +l_Lean_addTrace___at_Lean_Meta_Grind_simp___spec__3___closed__2 = _init_l_Lean_addTrace___at_Lean_Meta_Grind_simp___spec__3___closed__2(); +lean_mark_persistent(l_Lean_addTrace___at_Lean_Meta_Grind_simp___spec__3___closed__2); +l_Lean_addTrace___at_Lean_Meta_Grind_simp___spec__3___closed__3 = _init_l_Lean_addTrace___at_Lean_Meta_Grind_simp___spec__3___closed__3(); +lean_mark_persistent(l_Lean_addTrace___at_Lean_Meta_Grind_simp___spec__3___closed__3); +l_Lean_Meta_Grind_simp___closed__1 = _init_l_Lean_Meta_Grind_simp___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_simp___closed__1); +l_Lean_Meta_Grind_simp___closed__2 = _init_l_Lean_Meta_Grind_simp___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_simp___closed__2); +l_Lean_Meta_Grind_simp___closed__3 = _init_l_Lean_Meta_Grind_simp___closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind_simp___closed__3); +l_Lean_Meta_Grind_simp___closed__4 = _init_l_Lean_Meta_Grind_simp___closed__4(); +lean_mark_persistent(l_Lean_Meta_Grind_simp___closed__4); +l_Lean_Meta_Grind_simp___closed__5 = _init_l_Lean_Meta_Grind_simp___closed__5(); +lean_mark_persistent(l_Lean_Meta_Grind_simp___closed__5); +l_Lean_Meta_Grind_simp___closed__6 = _init_l_Lean_Meta_Grind_simp___closed__6(); +lean_mark_persistent(l_Lean_Meta_Grind_simp___closed__6); +l_Lean_Meta_Grind_simp___closed__7 = _init_l_Lean_Meta_Grind_simp___closed__7(); +lean_mark_persistent(l_Lean_Meta_Grind_simp___closed__7); +l_Lean_Meta_Grind_simp___closed__8 = _init_l_Lean_Meta_Grind_simp___closed__8(); +lean_mark_persistent(l_Lean_Meta_Grind_simp___closed__8); +l_Lean_Meta_Grind_simp___closed__9 = _init_l_Lean_Meta_Grind_simp___closed__9(); +lean_mark_persistent(l_Lean_Meta_Grind_simp___closed__9); +l_Lean_Meta_Grind_simp___closed__10 = _init_l_Lean_Meta_Grind_simp___closed__10(); +lean_mark_persistent(l_Lean_Meta_Grind_simp___closed__10); +l_Lean_Meta_Grind_simp___closed__11 = _init_l_Lean_Meta_Grind_simp___closed__11(); +lean_mark_persistent(l_Lean_Meta_Grind_simp___closed__11); +l_Lean_Meta_Grind_simp___closed__12 = _init_l_Lean_Meta_Grind_simp___closed__12(); +lean_mark_persistent(l_Lean_Meta_Grind_simp___closed__12); +l_Lean_Meta_Grind_simp___closed__13 = _init_l_Lean_Meta_Grind_simp___closed__13(); +lean_mark_persistent(l_Lean_Meta_Grind_simp___closed__13); return lean_io_result_mk_ok(lean_box(0)); } #ifdef __cplusplus diff --git a/stage0/stdlib/Lean/Meta/Tactic/Grind/TheoremPatterns.c b/stage0/stdlib/Lean/Meta/Tactic/Grind/TheoremPatterns.c deleted file mode 100644 index 3b7c5ad9dff2..000000000000 --- a/stage0/stdlib/Lean/Meta/Tactic/Grind/TheoremPatterns.c +++ /dev/null @@ -1,8928 +0,0 @@ -// Lean compiler output -// Module: Lean.Meta.Tactic.Grind.TheoremPatterns -// Imports: Lean.HeadIndex Lean.Util.FoldConsts Lean.Meta.Basic Lean.Meta.InferType -#include -#if defined(__clang__) -#pragma clang diagnostic ignored "-Wunused-parameter" -#pragma clang diagnostic ignored "-Wunused-label" -#elif defined(__GNUC__) && !defined(__CLANG__) -#pragma GCC diagnostic ignored "-Wunused-parameter" -#pragma GCC diagnostic ignored "-Wunused-label" -#pragma GCC diagnostic ignored "-Wunused-but-set-variable" -#endif -#ifdef __cplusplus -extern "C" { -#endif -lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___closed__1; -lean_object* l___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2170_(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_addTheoremPattern___closed__2; -lean_object* l_Lean_Name_reprPrec(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_ppPattern___closed__4; -static lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__3; -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_isForbidden___boxed(lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_getPatternFunInfo___spec__2(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__12___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____lambda__1___closed__3; -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__5(lean_object*, lean_object*, lean_object*, size_t, size_t, 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_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__6(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__15; -LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_mkAppN(lean_object*, lean_object*); -size_t lean_usize_shift_right(size_t, size_t); -lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__2(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_getPatternFunInfo___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_Meta_Grind_NormalizePattern_main___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__4; -static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___closed__2; -static lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__5; -static lean_object* l_Lean_SMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__17___closed__5; -static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____closed__1; -uint8_t lean_usize_dec_le(size_t, size_t); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__3___boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__5; -static lean_object* l_Lean_Meta_Grind_addTheoremPattern___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_dontCare; -uint64_t lean_uint64_of_nat(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instInhabitedOrigin; -size_t lean_uint64_to_usize(uint64_t); -static lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__8; -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addTheoremPattern(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -uint8_t l_Lean_Expr_isApp(lean_object*); -lean_object* l_Lean_Expr_sort___override(lean_object*); -lean_object* l_Lean_MessageData_ofList(lean_object*); -lean_object* lean_array_push(lean_object*, lean_object*); -static lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___spec__3___closed__1; -size_t lean_usize_mul(size_t, size_t); -static lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__2; -lean_object* lean_mk_array(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__3(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_ppPattern___closed__8; -static lean_object* l_Lean_Meta_Grind_instInhabitedOrigin___closed__1; -lean_object* l_Lean_Expr_bvar___override(lean_object*); -lean_object* lean_array_fget(lean_object*, lean_object*); -static lean_object* l_panic___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__1___closed__1; -lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__3___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__4(lean_object*, lean_object*); -lean_object* l_Lean_Expr_fvarId_x21(lean_object*); -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__9(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_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*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__13; -lean_object* l_Nat_nextPowerOfTwo_go(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_SMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__17___closed__6; -lean_object* l_Lean_stringToMessageData(lean_object*); -static lean_object* l_Lean_Meta_Grind_NormalizePattern_main___closed__2; -static lean_object* l_Lean_Meta_Grind_instInhabitedTheoremPattern___closed__1; -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instReprOrigin; -static lean_object* l_Lean_SMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__17___closed__1; -static lean_object* l_Lean_Meta_Grind_instReprOrigin___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand_go___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__2(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__4(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Meta_Grind_addTheoremPattern___spec__3(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_panic___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__1(lean_object*); -lean_object* l_Lean_throwError___at_Lean_Meta_setInlineAttribute___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addTheoremPattern___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_throwError___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_ppPattern___closed__3; -static lean_object* l_Lean_Meta_Grind_addTheoremPattern___closed__6; -uint8_t l_Lean_Expr_isBVar(lean_object*); -lean_object* l_List_mapTR_loop___at_Lean_mkConstWithLevelParams___spec__1(lean_object*, lean_object*); -uint8_t l_Lean_Expr_hasMVar(lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__12; -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____lambda__1(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_mkGroundPattern___closed__1; -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__10(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); -size_t lean_usize_of_nat(lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__4(lean_object*, size_t, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__9(lean_object*, size_t, size_t, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_ppPattern___closed__7; -static lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__19; -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_saveBVar___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*); -static lean_object* l_Lean_Meta_Grind_addTheoremPattern___closed__4; -uint8_t lean_expr_eqv(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_dontCare___closed__2; -lean_object* l_Lean_registerSimpleScopedEnvExtension___rarg(lean_object*, lean_object*); -lean_object* l_Lean_Meta_forallBoundedTelescope___at_Lean_Meta_arrowDomainsN___spec__6___rarg(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___closed__3; -uint64_t lean_uint64_shift_right(uint64_t, uint64_t); -lean_object* lean_nat_to_int(lean_object*); -lean_object* lean_nat_div(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__6; -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addTheoremPattern___spec__1___closed__1; -static lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__16; -lean_object* l_Lean_MessageData_ofFormat(lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___closed__2; -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_FVarId_getDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_outOfBounds___rarg(lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__13; -lean_object* lean_st_ref_get(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____closed__6; -lean_object* lean_st_mk_ref(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames; -lean_object* lean_array_to_list(lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_getPatternFunInfo___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____closed__4; -uint8_t l_Lean_Expr_hasLooseBVars(lean_object*); -static lean_object* l_Lean_Meta_Grind_ppPattern___closed__2; -lean_object* l_Lean_Expr_toHeadIndex(lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_getPatternFunInfo___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_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__1(lean_object*); -extern lean_object* l_Lean_levelZero; -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_NormalizePattern_main(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_ppPattern___closed__11; -LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_isForbidden(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_theoremPatternsExt; -extern lean_object* l_Lean_instInhabitedExpr; -static lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__9___closed__1; -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_mkAnnotation(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_addTheoremPattern___closed__8; -LEAN_EXPORT lean_object* l_Lean_SMap_find_x3f___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__2___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_SMap_insert___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__7(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_SMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__17___closed__3; -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_getPatternFn_x3f(lean_object*); -LEAN_EXPORT lean_object* l_Lean_SMap_find_x3f___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__2(lean_object*, lean_object*); -uint8_t lean_name_eq(lean_object*, lean_object*); -lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__4(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); -LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____lambda__1___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_foundBVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____closed__2; -static lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__2; -static lean_object* l_Lean_Meta_Grind_addTheoremPattern___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_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand_go___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__14(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_foundBVar___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT uint8_t l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_foundBVar___spec__1(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__11; -static lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_mkGroundPattern___closed__2; -static lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__14; -LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_isGroundPattern(lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__4___boxed(lean_object*, lean_object*, lean_object*); -extern lean_object* l_Lean_Meta_instMonadMetaM; -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__3; -static lean_object* l_Lean_Meta_Grind_addTheoremPattern___lambda__1___closed__1; -static lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__10; -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand_go___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__3(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____lambda__1___closed__4; -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_groundPattern_x3f___boxed(lean_object*); -lean_object* lean_usize_to_nat(size_t); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_mkGroundPattern(lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__4; -lean_object* l_Lean_MessageData_ofExpr(lean_object*); -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__3(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); -LEAN_EXPORT lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addTheoremPattern___spec__1(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__9; -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addTheoremPattern___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_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Meta_Grind_addTheoremPattern___spec__2(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____lambda__1___closed__2; -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__3(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__16(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_mkGroundPattern___closed__3; -static lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__17; -static lean_object* l_Lean_Meta_Grind_ppPattern___closed__9; -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__7(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__4; -static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____closed__5; -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_isAtomicPattern___boxed(lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__12; -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__6___boxed(lean_object*, lean_object*); -uint8_t lean_nat_dec_eq(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_instInhabitedTheoremPattern___closed__2; -uint8_t lean_nat_dec_lt(lean_object*, lean_object*); -lean_object* l_id___rarg___boxed(lean_object*); -lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); -lean_object* l_Lean_PersistentHashMap_mkEmptyEntries(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ppPattern(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_getPatternFunInfo___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_isAtomicPattern(lean_object*); -lean_object* lean_array_set(lean_object*, lean_object*, lean_object*); -static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3; -static size_t l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__4___closed__2; -uint64_t l_Lean_Name_hash___override(lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_dontCare___closed__1; -lean_object* l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -uint8_t l_Array_contains___at_Lean_registerInternalExceptionId___spec__1(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__1; -uint64_t lean_uint64_xor(uint64_t, uint64_t); -extern lean_object* l_Id_instMonad; -lean_object* l_Repr_addAppParen(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_SMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__17; -lean_object* lean_panic_fn(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__7; -lean_object* l_Lean_annotation_x3f(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__3; -static lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__7; -lean_object* lean_nat_sub(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Expr_getAppFn(lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__5; -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* lean_nat_mul(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__11(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); -static lean_object* l_Lean_Meta_Grind_addTheoremPattern___closed__7; -lean_object* l_Lean_Meta_isTypeFormer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_addTheoremPattern___closed__3; -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__1___boxed(lean_object*, lean_object*); -lean_object* l_Lean_ScopedEnvExtension_addCore___rarg(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*); -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__13(lean_object*); -lean_object* l_Lean_PersistentHashMap_mkCollisionNode___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__2; -static lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__7; -static lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__1; -lean_object* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); -uint8_t l_Lean_LocalDecl_binderInfo(lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__8; -static lean_object* l_Lean_Meta_Grind_NormalizePattern_main___closed__1; -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instInhabitedTheoremPattern; -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_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* l_List_reverse___rarg(lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__11; -static lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__1; -size_t lean_usize_sub(size_t, size_t); -lean_object* lean_array_mk(lean_object*); -static lean_object* l_Lean_Meta_Grind_ppPattern___closed__1; -LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___spec__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); -static lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__14; -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__8(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____closed__7; -lean_object* lean_array_uget(lean_object*, size_t); -static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___closed__1; -size_t lean_array_size(lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__15; -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_saveSymbol(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__17; -lean_object* l_Lean_SMap_instInhabited___rarg(lean_object*, lean_object*); -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_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_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_getPatternFn_x3f___boxed(lean_object*); -size_t lean_usize_shift_left(size_t, size_t); -lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__16; -uint8_t l___private_Lean_HeadIndex_0__Lean_beqHeadIndex____x40_Lean_HeadIndex___hyg_69_(lean_object*, lean_object*); -lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__6; -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static size_t l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__4___closed__1; -static lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___spec__3___closed__2; -lean_object* lean_string_append(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_ppPattern___closed__10; -lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__9; -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__2(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); -static lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addTheoremPattern___spec__1___closed__2; -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307_(lean_object*); -extern lean_object* l_Lean_Name_instBEq; -lean_object* lean_array_get_size(lean_object*); -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_foundBVar___spec__1___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_groundPattern_x3f(lean_object*); -static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1; -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_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__10; -static lean_object* l_Lean_SMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__17___closed__2; -uint8_t lean_usize_dec_lt(size_t, size_t); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_saveBVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT uint8_t l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__1(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__6; -static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____closed__3; -LEAN_EXPORT lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addTheoremPattern___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_ppPattern___closed__5; -lean_object* lean_nat_add(lean_object*, lean_object*); -lean_object* l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(lean_object*); -lean_object* l_Lean_MessageData_bracket(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__6(lean_object*, lean_object*); -uint8_t l_Lean_Expr_isConst(lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_dontCare___closed__3; -static lean_object* l_Lean_SMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__17___closed__4; -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__11(lean_object*, lean_object*, lean_object*, lean_object*); -uint64_t l_Lean_HeadIndex_hash(lean_object*); -static lean_object* l_panic___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__1___closed__2; -uint8_t l_Lean_Expr_isFVar(lean_object*); -static lean_object* l_Lean_Meta_Grind_ppPattern___closed__6; -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_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_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_getPatternFunInfo(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_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__10(size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___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_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__8(lean_object*, lean_object*, lean_object*, size_t, size_t, 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___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__18; -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_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2; -LEAN_EXPORT uint8_t l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__12(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51_(lean_object*, lean_object*); -lean_object* l_ReaderT_instMonad___rarg(lean_object*); -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__15(lean_object*, lean_object*); -size_t lean_usize_land(size_t, size_t); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_isGroundPattern___boxed(lean_object*); -extern lean_object* l_Lean_instHashableName; -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_getPatternFunInfo___spec__1(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_getPatternFunInfo___closed__1; -static lean_object* _init_l_Lean_Meta_Grind_instInhabitedOrigin___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = lean_box(0); -x_2 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Meta_Grind_instInhabitedOrigin() { -_start: -{ -lean_object* x_1; -x_1 = l_Lean_Meta_Grind_instInhabitedOrigin___closed__1; -return x_1; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean.Meta.Grind.Origin.decl", 27, 27); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__1; -x_2 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__3() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__2; -x_2 = lean_box(1); -x_3 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set(x_3, 1, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__4() { -_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_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__5() { -_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_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__6() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean.Meta.Grind.Origin.fvar", 27, 27); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__7() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__6; -x_2 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__8() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__7; -x_2 = lean_box(1); -x_3 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set(x_3, 1, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__9() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean.Meta.Grind.Origin.stx", 26, 26); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__10() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____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_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__11() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__10; -x_2 = lean_box(1); -x_3 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set(x_3, 1, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__12() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean.Meta.Grind.Origin.other", 28, 28); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__13() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____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_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__14() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__4; -x_2 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__13; -x_3 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set(x_3, 1, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__15() { -_start: -{ -lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__14; -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_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__16() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__5; -x_2 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__13; -x_3 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set(x_3, 1, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__17() { -_start: -{ -lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__16; -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_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51_(lean_object* x_1, lean_object* x_2) { -_start: -{ -switch (lean_obj_tag(x_1)) { -case 0: -{ -lean_object* x_3; lean_object* x_4; uint8_t x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; -x_3 = lean_ctor_get(x_1, 0); -lean_inc(x_3); -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_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____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_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____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); -x_11 = 0; -x_12 = lean_alloc_ctor(6, 1, 1); -lean_ctor_set(x_12, 0, x_10); -lean_ctor_set_uint8(x_12, sizeof(void*)*1, x_11); -x_13 = l_Repr_addAppParen(x_12, x_2); -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_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____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); -x_16 = 0; -x_17 = lean_alloc_ctor(6, 1, 1); -lean_ctor_set(x_17, 0, x_15); -lean_ctor_set_uint8(x_17, sizeof(void*)*1, x_16); -x_18 = l_Repr_addAppParen(x_17, x_2); -return x_18; -} -} -case 1: -{ -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_1, 0); -lean_inc(x_19); -lean_dec(x_1); -x_20 = lean_unsigned_to_nat(1024u); -x_21 = lean_nat_dec_le(x_20, x_2); -x_22 = l_Lean_Name_reprPrec(x_19, x_20); -x_23 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__8; -x_24 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_24, 0, x_23); -lean_ctor_set(x_24, 1, x_22); -if (x_21 == 0) -{ -lean_object* x_25; lean_object* x_26; uint8_t x_27; lean_object* x_28; lean_object* x_29; -x_25 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__4; -x_26 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_26, 0, x_25); -lean_ctor_set(x_26, 1, x_24); -x_27 = 0; -x_28 = lean_alloc_ctor(6, 1, 1); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set_uint8(x_28, sizeof(void*)*1, x_27); -x_29 = l_Repr_addAppParen(x_28, x_2); -return x_29; -} -else -{ -lean_object* x_30; lean_object* x_31; uint8_t x_32; lean_object* x_33; lean_object* x_34; -x_30 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__5; -x_31 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_31, 0, x_30); -lean_ctor_set(x_31, 1, x_24); -x_32 = 0; -x_33 = lean_alloc_ctor(6, 1, 1); -lean_ctor_set(x_33, 0, x_31); -lean_ctor_set_uint8(x_33, sizeof(void*)*1, x_32); -x_34 = l_Repr_addAppParen(x_33, x_2); -return x_34; -} -} -case 2: -{ -uint8_t x_35; -x_35 = !lean_is_exclusive(x_1); -if (x_35 == 0) -{ -lean_object* x_36; lean_object* x_37; lean_object* x_38; uint8_t x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; -x_36 = lean_ctor_get(x_1, 0); -x_37 = lean_ctor_get(x_1, 1); -x_38 = lean_unsigned_to_nat(1024u); -x_39 = lean_nat_dec_le(x_38, x_2); -x_40 = l_Lean_Name_reprPrec(x_36, x_38); -x_41 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__11; -lean_ctor_set_tag(x_1, 5); -lean_ctor_set(x_1, 1, x_40); -lean_ctor_set(x_1, 0, x_41); -x_42 = lean_box(1); -x_43 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_43, 0, x_1); -lean_ctor_set(x_43, 1, x_42); -x_44 = l___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2170_(x_37, x_38); -x_45 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_45, 0, x_43); -lean_ctor_set(x_45, 1, x_44); -if (x_39 == 0) -{ -lean_object* x_46; lean_object* x_47; uint8_t x_48; lean_object* x_49; lean_object* x_50; -x_46 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__4; -x_47 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_47, 0, x_46); -lean_ctor_set(x_47, 1, x_45); -x_48 = 0; -x_49 = lean_alloc_ctor(6, 1, 1); -lean_ctor_set(x_49, 0, x_47); -lean_ctor_set_uint8(x_49, sizeof(void*)*1, x_48); -x_50 = l_Repr_addAppParen(x_49, x_2); -return x_50; -} -else -{ -lean_object* x_51; lean_object* x_52; uint8_t x_53; lean_object* x_54; lean_object* x_55; -x_51 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____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_45); -x_53 = 0; -x_54 = lean_alloc_ctor(6, 1, 1); -lean_ctor_set(x_54, 0, x_52); -lean_ctor_set_uint8(x_54, sizeof(void*)*1, x_53); -x_55 = l_Repr_addAppParen(x_54, x_2); -return x_55; -} -} -else -{ -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; -x_56 = lean_ctor_get(x_1, 0); -x_57 = lean_ctor_get(x_1, 1); -lean_inc(x_57); -lean_inc(x_56); -lean_dec(x_1); -x_58 = lean_unsigned_to_nat(1024u); -x_59 = lean_nat_dec_le(x_58, x_2); -x_60 = l_Lean_Name_reprPrec(x_56, x_58); -x_61 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__11; -x_62 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_62, 0, x_61); -lean_ctor_set(x_62, 1, x_60); -x_63 = lean_box(1); -x_64 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_64, 0, x_62); -lean_ctor_set(x_64, 1, x_63); -x_65 = l___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2170_(x_57, x_58); -x_66 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_66, 0, x_64); -lean_ctor_set(x_66, 1, x_65); -if (x_59 == 0) -{ -lean_object* x_67; lean_object* x_68; uint8_t x_69; lean_object* x_70; lean_object* x_71; -x_67 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__4; -x_68 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_68, 0, x_67); -lean_ctor_set(x_68, 1, x_66); -x_69 = 0; -x_70 = lean_alloc_ctor(6, 1, 1); -lean_ctor_set(x_70, 0, x_68); -lean_ctor_set_uint8(x_70, sizeof(void*)*1, x_69); -x_71 = l_Repr_addAppParen(x_70, x_2); -return x_71; -} -else -{ -lean_object* x_72; lean_object* x_73; uint8_t x_74; lean_object* x_75; lean_object* x_76; -x_72 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__5; -x_73 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_73, 0, x_72); -lean_ctor_set(x_73, 1, x_66); -x_74 = 0; -x_75 = lean_alloc_ctor(6, 1, 1); -lean_ctor_set(x_75, 0, x_73); -lean_ctor_set_uint8(x_75, sizeof(void*)*1, x_74); -x_76 = l_Repr_addAppParen(x_75, x_2); -return x_76; -} -} -} -default: -{ -lean_object* x_77; uint8_t x_78; -x_77 = lean_unsigned_to_nat(1024u); -x_78 = lean_nat_dec_le(x_77, x_2); -if (x_78 == 0) -{ -lean_object* x_79; lean_object* x_80; -x_79 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__15; -x_80 = l_Repr_addAppParen(x_79, x_2); -return x_80; -} -else -{ -lean_object* x_81; lean_object* x_82; -x_81 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__17; -x_82 = l_Repr_addAppParen(x_81, x_2); -return x_82; -} -} -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____boxed(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; -x_3 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51_(x_1, x_2); -lean_dec(x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Meta_Grind_instReprOrigin___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____boxed), 2, 0); -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_Grind_instReprOrigin() { -_start: -{ -lean_object* x_1; -x_1 = l_Lean_Meta_Grind_instReprOrigin___closed__1; -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_Grind_instInhabitedTheoremPattern___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = lean_unsigned_to_nat(0u); -x_2 = l_Lean_Expr_bvar___override(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Meta_Grind_instInhabitedTheoremPattern___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 = lean_box(0); -x_2 = l_Lean_Meta_Grind_instInhabitedTheoremPattern___closed__1; -x_3 = lean_unsigned_to_nat(0u); -x_4 = l_Lean_Meta_Grind_instInhabitedOrigin___closed__1; -x_5 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_5, 0, x_2); -lean_ctor_set(x_5, 1, x_3); -lean_ctor_set(x_5, 2, x_1); -lean_ctor_set(x_5, 3, x_1); -lean_ctor_set(x_5, 4, x_4); -return x_5; -} -} -static lean_object* _init_l_Lean_Meta_Grind_instInhabitedTheoremPattern() { -_start: -{ -lean_object* x_1; -x_1 = l_Lean_Meta_Grind_instInhabitedTheoremPattern___closed__2; -return x_1; -} -} -static lean_object* _init_l_panic___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__1___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Name_instBEq; -x_2 = l_Lean_instHashableName; -x_3 = l_Lean_SMap_instInhabited___rarg(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_panic___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__1___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Id_instMonad; -x_2 = l_panic___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__1___closed__1; -x_3 = l_instInhabitedOfMonad___rarg(x_1, x_2); -return x_3; -} -} -LEAN_EXPORT lean_object* l_panic___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__1(lean_object* x_1) { -_start: -{ -lean_object* x_2; lean_object* x_3; -x_2 = l_panic___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__1___closed__2; -x_3 = lean_panic_fn(x_2, x_1); -return x_3; -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: -{ -lean_object* x_6; 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_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__4___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_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__4___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_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__4___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_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__4(lean_object* x_1, size_t x_2, lean_object* x_3) { -_start: -{ -if (lean_obj_tag(x_1) == 0) -{ -uint8_t x_4; -x_4 = !lean_is_exclusive(x_1); -if (x_4 == 0) -{ -lean_object* x_5; size_t x_6; size_t x_7; size_t x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; -x_5 = lean_ctor_get(x_1, 0); -x_6 = 5; -x_7 = l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__4___closed__2; -x_8 = lean_usize_land(x_2, x_7); -x_9 = lean_usize_to_nat(x_8); -x_10 = lean_box(2); -x_11 = lean_array_get(x_10, x_5, x_9); -lean_dec(x_9); -lean_dec(x_5); -switch (lean_obj_tag(x_11)) { -case 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_name_eq(x_3, x_12); -lean_dec(x_12); -if (x_14 == 0) -{ -lean_object* x_15; -lean_dec(x_13); -lean_free_object(x_1); -x_15 = lean_box(0); -return x_15; -} -else -{ -lean_ctor_set_tag(x_1, 1); -lean_ctor_set(x_1, 0, x_13); -return x_1; -} -} -case 1: -{ -lean_object* x_16; size_t x_17; -lean_free_object(x_1); -x_16 = lean_ctor_get(x_11, 0); -lean_inc(x_16); -lean_dec(x_11); -x_17 = lean_usize_shift_right(x_2, x_6); -x_1 = x_16; -x_2 = x_17; -goto _start; -} -default: -{ -lean_object* x_19; -lean_free_object(x_1); -x_19 = lean_box(0); -return x_19; -} -} -} -else -{ -lean_object* x_20; size_t x_21; size_t x_22; size_t x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_20 = lean_ctor_get(x_1, 0); -lean_inc(x_20); -lean_dec(x_1); -x_21 = 5; -x_22 = l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__4___closed__2; -x_23 = lean_usize_land(x_2, x_22); -x_24 = lean_usize_to_nat(x_23); -x_25 = lean_box(2); -x_26 = lean_array_get(x_25, x_20, x_24); -lean_dec(x_24); -lean_dec(x_20); -switch (lean_obj_tag(x_26)) { -case 0: -{ -lean_object* x_27; lean_object* x_28; uint8_t x_29; -x_27 = lean_ctor_get(x_26, 0); -lean_inc(x_27); -x_28 = lean_ctor_get(x_26, 1); -lean_inc(x_28); -lean_dec(x_26); -x_29 = lean_name_eq(x_3, x_27); -lean_dec(x_27); -if (x_29 == 0) -{ -lean_object* x_30; -lean_dec(x_28); -x_30 = lean_box(0); -return x_30; -} -else -{ -lean_object* x_31; -x_31 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_31, 0, x_28); -return x_31; -} -} -case 1: -{ -lean_object* x_32; size_t x_33; -x_32 = lean_ctor_get(x_26, 0); -lean_inc(x_32); -lean_dec(x_26); -x_33 = lean_usize_shift_right(x_2, x_21); -x_1 = x_32; -x_2 = x_33; -goto _start; -} -default: -{ -lean_object* x_35; -x_35 = lean_box(0); -return x_35; -} -} -} -} -else -{ -lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_36 = lean_ctor_get(x_1, 0); -lean_inc(x_36); -x_37 = lean_ctor_get(x_1, 1); -lean_inc(x_37); -lean_dec(x_1); -x_38 = lean_unsigned_to_nat(0u); -x_39 = l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__5(x_36, x_37, lean_box(0), x_38, x_3); -lean_dec(x_37); -lean_dec(x_36); -return x_39; -} -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__3(lean_object* x_1, lean_object* x_2) { -_start: -{ -uint64_t x_3; size_t x_4; lean_object* x_5; -x_3 = l_Lean_Name_hash___override(x_2); -x_4 = lean_uint64_to_usize(x_3); -x_5 = l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__4(x_1, x_4, x_2); -return x_5; -} -} -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__6(lean_object* x_1, lean_object* x_2) { -_start: -{ -if (lean_obj_tag(x_2) == 0) -{ -lean_object* x_3; -x_3 = lean_box(0); -return x_3; -} -else -{ -lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; -x_4 = lean_ctor_get(x_2, 0); -x_5 = lean_ctor_get(x_2, 1); -x_6 = lean_ctor_get(x_2, 2); -x_7 = lean_name_eq(x_4, x_1); -if (x_7 == 0) -{ -x_2 = x_6; -goto _start; -} -else -{ -lean_object* x_9; -lean_inc(x_5); -x_9 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_9, 0, x_5); -return x_9; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_SMap_find_x3f___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__2(lean_object* x_1, lean_object* x_2) { -_start: -{ -uint8_t x_3; -x_3 = lean_ctor_get_uint8(x_1, sizeof(void*)*2); -if (x_3 == 0) -{ -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 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__3(x_5, x_2); -if (lean_obj_tag(x_6) == 0) -{ -lean_object* x_7; lean_object* x_8; uint64_t 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; size_t x_16; size_t x_17; size_t x_18; size_t x_19; size_t x_20; lean_object* x_21; lean_object* x_22; -x_7 = lean_ctor_get(x_4, 1); -lean_inc(x_7); -lean_dec(x_4); -x_8 = lean_array_get_size(x_7); -x_9 = l_Lean_Name_hash___override(x_2); -x_10 = 32; -x_11 = lean_uint64_shift_right(x_9, x_10); -x_12 = lean_uint64_xor(x_9, x_11); -x_13 = 16; -x_14 = lean_uint64_shift_right(x_12, x_13); -x_15 = lean_uint64_xor(x_12, x_14); -x_16 = lean_uint64_to_usize(x_15); -x_17 = lean_usize_of_nat(x_8); -lean_dec(x_8); -x_18 = 1; -x_19 = lean_usize_sub(x_17, x_18); -x_20 = lean_usize_land(x_16, x_19); -x_21 = lean_array_uget(x_7, x_20); -lean_dec(x_7); -x_22 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__6(x_2, x_21); -lean_dec(x_21); -return x_22; -} -else -{ -uint8_t x_23; -lean_dec(x_4); -x_23 = !lean_is_exclusive(x_6); -if (x_23 == 0) -{ -return x_6; -} -else -{ -lean_object* x_24; lean_object* x_25; -x_24 = lean_ctor_get(x_6, 0); -lean_inc(x_24); -lean_dec(x_6); -x_25 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_25, 0, x_24); -return x_25; -} -} -} -else -{ -lean_object* x_26; lean_object* x_27; lean_object* x_28; uint64_t x_29; uint64_t x_30; uint64_t x_31; uint64_t x_32; uint64_t x_33; uint64_t x_34; uint64_t x_35; size_t x_36; size_t x_37; size_t x_38; size_t x_39; size_t x_40; lean_object* x_41; lean_object* x_42; -x_26 = lean_ctor_get(x_1, 0); -lean_inc(x_26); -lean_dec(x_1); -x_27 = lean_ctor_get(x_26, 1); -lean_inc(x_27); -lean_dec(x_26); -x_28 = lean_array_get_size(x_27); -x_29 = l_Lean_Name_hash___override(x_2); -x_30 = 32; -x_31 = lean_uint64_shift_right(x_29, x_30); -x_32 = lean_uint64_xor(x_29, x_31); -x_33 = 16; -x_34 = lean_uint64_shift_right(x_32, x_33); -x_35 = lean_uint64_xor(x_32, x_34); -x_36 = lean_uint64_to_usize(x_35); -x_37 = lean_usize_of_nat(x_28); -lean_dec(x_28); -x_38 = 1; -x_39 = lean_usize_sub(x_37, x_38); -x_40 = lean_usize_land(x_36, x_39); -x_41 = lean_array_uget(x_27, x_40); -lean_dec(x_27); -x_42 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__6(x_2, x_41); -lean_dec(x_41); -return x_42; -} -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__10(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_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__9(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_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__11(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; -x_5 = lean_ctor_get(x_1, 0); -lean_inc(x_5); -x_6 = lean_ctor_get(x_1, 1); -lean_inc(x_6); -x_7 = lean_array_get_size(x_5); -x_8 = lean_nat_dec_lt(x_2, x_7); -lean_dec(x_7); -if (x_8 == 0) -{ -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_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__9___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_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__9(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_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__4___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_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__9(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_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__9(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_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__4___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_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__9(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_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__11(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_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__9___closed__1; -x_93 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__10(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_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__11(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_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__9___closed__1; -x_107 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__10(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_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__8(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -uint64_t x_4; size_t x_5; size_t x_6; lean_object* x_7; -x_4 = l_Lean_Name_hash___override(x_2); -x_5 = lean_uint64_to_usize(x_4); -x_6 = 1; -x_7 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__9(x_1, x_5, x_6, x_2, x_3); -return x_7; -} -} -LEAN_EXPORT uint8_t l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__12(lean_object* x_1, lean_object* x_2) { -_start: -{ -if (lean_obj_tag(x_2) == 0) -{ -uint8_t x_3; -x_3 = 0; -return x_3; -} -else -{ -lean_object* x_4; lean_object* x_5; uint8_t x_6; -x_4 = lean_ctor_get(x_2, 0); -x_5 = lean_ctor_get(x_2, 2); -x_6 = lean_name_eq(x_4, x_1); -if (x_6 == 0) -{ -x_2 = x_5; -goto _start; -} -else -{ -uint8_t x_8; -x_8 = 1; -return x_8; -} -} -} -} -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__15(lean_object* x_1, lean_object* x_2) { -_start: -{ -if (lean_obj_tag(x_2) == 0) -{ -return x_1; -} -else -{ -uint8_t x_3; -x_3 = !lean_is_exclusive(x_2); -if (x_3 == 0) -{ -lean_object* x_4; lean_object* x_5; lean_object* x_6; uint64_t x_7; uint64_t x_8; uint64_t x_9; uint64_t x_10; uint64_t x_11; uint64_t x_12; uint64_t x_13; size_t x_14; size_t x_15; size_t x_16; size_t x_17; size_t x_18; lean_object* x_19; lean_object* x_20; -x_4 = lean_ctor_get(x_2, 0); -x_5 = lean_ctor_get(x_2, 2); -x_6 = lean_array_get_size(x_1); -x_7 = l_Lean_Name_hash___override(x_4); -x_8 = 32; -x_9 = lean_uint64_shift_right(x_7, x_8); -x_10 = lean_uint64_xor(x_7, x_9); -x_11 = 16; -x_12 = lean_uint64_shift_right(x_10, x_11); -x_13 = lean_uint64_xor(x_10, x_12); -x_14 = lean_uint64_to_usize(x_13); -x_15 = lean_usize_of_nat(x_6); -lean_dec(x_6); -x_16 = 1; -x_17 = lean_usize_sub(x_15, x_16); -x_18 = lean_usize_land(x_14, x_17); -x_19 = lean_array_uget(x_1, x_18); -lean_ctor_set(x_2, 2, x_19); -x_20 = lean_array_uset(x_1, x_18, x_2); -x_1 = x_20; -x_2 = x_5; -goto _start; -} -else -{ -lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint64_t x_26; uint64_t x_27; uint64_t x_28; uint64_t x_29; uint64_t x_30; uint64_t x_31; uint64_t x_32; size_t x_33; size_t x_34; size_t x_35; size_t x_36; size_t x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_22 = lean_ctor_get(x_2, 0); -x_23 = lean_ctor_get(x_2, 1); -x_24 = lean_ctor_get(x_2, 2); -lean_inc(x_24); -lean_inc(x_23); -lean_inc(x_22); -lean_dec(x_2); -x_25 = lean_array_get_size(x_1); -x_26 = l_Lean_Name_hash___override(x_22); -x_27 = 32; -x_28 = lean_uint64_shift_right(x_26, x_27); -x_29 = lean_uint64_xor(x_26, x_28); -x_30 = 16; -x_31 = lean_uint64_shift_right(x_29, x_30); -x_32 = lean_uint64_xor(x_29, x_31); -x_33 = lean_uint64_to_usize(x_32); -x_34 = lean_usize_of_nat(x_25); -lean_dec(x_25); -x_35 = 1; -x_36 = lean_usize_sub(x_34, x_35); -x_37 = lean_usize_land(x_33, x_36); -x_38 = lean_array_uget(x_1, x_37); -x_39 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_39, 0, x_22); -lean_ctor_set(x_39, 1, x_23); -lean_ctor_set(x_39, 2, x_38); -x_40 = lean_array_uset(x_1, x_37, x_39); -x_1 = x_40; -x_2 = x_24; -goto _start; -} -} -} -} -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand_go___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__14(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -lean_object* x_4; uint8_t x_5; -x_4 = lean_array_get_size(x_2); -x_5 = lean_nat_dec_lt(x_1, x_4); -lean_dec(x_4); -if (x_5 == 0) -{ -lean_dec(x_2); -lean_dec(x_1); -return x_3; -} -else -{ -lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; -x_6 = lean_array_fget(x_2, x_1); -x_7 = lean_box(0); -x_8 = lean_array_fset(x_2, x_1, x_7); -x_9 = l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__15(x_3, x_6); -x_10 = lean_unsigned_to_nat(1u); -x_11 = lean_nat_add(x_1, x_10); -lean_dec(x_1); -x_1 = x_11; -x_2 = x_8; -x_3 = x_9; -goto _start; -} -} -} -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__13(lean_object* x_1) { -_start: -{ -lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; -x_2 = lean_array_get_size(x_1); -x_3 = lean_unsigned_to_nat(2u); -x_4 = lean_nat_mul(x_2, x_3); -lean_dec(x_2); -x_5 = lean_box(0); -x_6 = lean_mk_array(x_4, x_5); -x_7 = lean_unsigned_to_nat(0u); -x_8 = l_Std_DHashMap_Internal_Raw_u2080_expand_go___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__14(x_7, x_1, x_6); -return x_8; -} -} -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__16(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -if (lean_obj_tag(x_3) == 0) -{ -lean_object* x_4; -lean_dec(x_2); -lean_dec(x_1); -x_4 = lean_box(0); -return x_4; -} -else -{ -uint8_t x_5; -x_5 = !lean_is_exclusive(x_3); -if (x_5 == 0) -{ -lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; -x_6 = lean_ctor_get(x_3, 0); -x_7 = lean_ctor_get(x_3, 1); -x_8 = lean_ctor_get(x_3, 2); -x_9 = lean_name_eq(x_6, x_1); -if (x_9 == 0) -{ -lean_object* x_10; -x_10 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__16(x_1, x_2, x_8); -lean_ctor_set(x_3, 2, x_10); -return x_3; -} -else -{ -lean_dec(x_7); -lean_dec(x_6); -lean_ctor_set(x_3, 1, x_2); -lean_ctor_set(x_3, 0, x_1); -return x_3; -} -} -else -{ -lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_11 = lean_ctor_get(x_3, 0); -x_12 = lean_ctor_get(x_3, 1); -x_13 = lean_ctor_get(x_3, 2); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_dec(x_3); -x_14 = lean_name_eq(x_11, x_1); -if (x_14 == 0) -{ -lean_object* x_15; lean_object* x_16; -x_15 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__16(x_1, x_2, x_13); -x_16 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_16, 0, x_11); -lean_ctor_set(x_16, 1, x_12); -lean_ctor_set(x_16, 2, x_15); -return x_16; -} -else -{ -lean_object* x_17; -lean_dec(x_12); -lean_dec(x_11); -x_17 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_17, 0, x_1); -lean_ctor_set(x_17, 1, x_2); -lean_ctor_set(x_17, 2, x_13); -return x_17; -} -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_SMap_insert___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -uint8_t x_4; -x_4 = lean_ctor_get_uint8(x_1, sizeof(void*)*2); -if (x_4 == 0) -{ -uint8_t x_5; -x_5 = !lean_is_exclusive(x_1); -if (x_5 == 0) -{ -lean_object* x_6; lean_object* x_7; uint8_t x_8; -x_6 = lean_ctor_get(x_1, 1); -x_7 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__8(x_6, x_2, x_3); -x_8 = 0; -lean_ctor_set(x_1, 1, x_7); -lean_ctor_set_uint8(x_1, sizeof(void*)*2, x_8); -return x_1; -} -else -{ -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, 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_PersistentHashMap_insert___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__8(x_10, x_2, x_3); -x_12 = 0; -x_13 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_13, 0, x_9); -lean_ctor_set(x_13, 1, x_11); -lean_ctor_set_uint8(x_13, sizeof(void*)*2, x_12); -return x_13; -} -} -else -{ -uint8_t x_14; -x_14 = !lean_is_exclusive(x_1); -if (x_14 == 0) -{ -lean_object* x_15; uint8_t x_16; -x_15 = lean_ctor_get(x_1, 0); -x_16 = !lean_is_exclusive(x_15); -if (x_16 == 0) -{ -lean_object* x_17; lean_object* x_18; lean_object* x_19; uint64_t x_20; uint64_t x_21; uint64_t x_22; uint64_t x_23; uint64_t x_24; uint64_t x_25; uint64_t x_26; size_t x_27; size_t x_28; size_t x_29; size_t x_30; size_t x_31; lean_object* x_32; uint8_t x_33; -x_17 = lean_ctor_get(x_15, 0); -x_18 = lean_ctor_get(x_15, 1); -x_19 = lean_array_get_size(x_18); -x_20 = l_Lean_Name_hash___override(x_2); -x_21 = 32; -x_22 = lean_uint64_shift_right(x_20, x_21); -x_23 = lean_uint64_xor(x_20, x_22); -x_24 = 16; -x_25 = lean_uint64_shift_right(x_23, x_24); -x_26 = lean_uint64_xor(x_23, x_25); -x_27 = lean_uint64_to_usize(x_26); -x_28 = lean_usize_of_nat(x_19); -lean_dec(x_19); -x_29 = 1; -x_30 = lean_usize_sub(x_28, x_29); -x_31 = lean_usize_land(x_27, x_30); -x_32 = lean_array_uget(x_18, x_31); -x_33 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__12(x_2, x_32); -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; uint8_t x_43; -x_34 = lean_unsigned_to_nat(1u); -x_35 = lean_nat_add(x_17, x_34); -lean_dec(x_17); -x_36 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_36, 0, x_2); -lean_ctor_set(x_36, 1, x_3); -lean_ctor_set(x_36, 2, x_32); -x_37 = lean_array_uset(x_18, x_31, x_36); -x_38 = lean_unsigned_to_nat(4u); -x_39 = lean_nat_mul(x_35, x_38); -x_40 = lean_unsigned_to_nat(3u); -x_41 = lean_nat_div(x_39, x_40); -lean_dec(x_39); -x_42 = lean_array_get_size(x_37); -x_43 = lean_nat_dec_le(x_41, x_42); -lean_dec(x_42); -lean_dec(x_41); -if (x_43 == 0) -{ -lean_object* x_44; uint8_t x_45; -x_44 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__13(x_37); -lean_ctor_set(x_15, 1, x_44); -lean_ctor_set(x_15, 0, x_35); -x_45 = 1; -lean_ctor_set_uint8(x_1, sizeof(void*)*2, x_45); -return x_1; -} -else -{ -uint8_t x_46; -lean_ctor_set(x_15, 1, x_37); -lean_ctor_set(x_15, 0, x_35); -x_46 = 1; -lean_ctor_set_uint8(x_1, sizeof(void*)*2, x_46); -return x_1; -} -} -else -{ -lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; uint8_t x_51; -x_47 = lean_box(0); -x_48 = lean_array_uset(x_18, x_31, x_47); -x_49 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__16(x_2, x_3, x_32); -x_50 = lean_array_uset(x_48, x_31, x_49); -lean_ctor_set(x_15, 1, x_50); -x_51 = 1; -lean_ctor_set_uint8(x_1, sizeof(void*)*2, x_51); -return x_1; -} -} -else -{ -lean_object* x_52; lean_object* x_53; lean_object* x_54; uint64_t x_55; uint64_t x_56; uint64_t x_57; uint64_t x_58; uint64_t x_59; uint64_t x_60; uint64_t x_61; size_t x_62; size_t x_63; size_t x_64; size_t x_65; size_t x_66; lean_object* x_67; uint8_t x_68; -x_52 = lean_ctor_get(x_15, 0); -x_53 = lean_ctor_get(x_15, 1); -lean_inc(x_53); -lean_inc(x_52); -lean_dec(x_15); -x_54 = lean_array_get_size(x_53); -x_55 = l_Lean_Name_hash___override(x_2); -x_56 = 32; -x_57 = lean_uint64_shift_right(x_55, x_56); -x_58 = lean_uint64_xor(x_55, x_57); -x_59 = 16; -x_60 = lean_uint64_shift_right(x_58, x_59); -x_61 = lean_uint64_xor(x_58, x_60); -x_62 = lean_uint64_to_usize(x_61); -x_63 = lean_usize_of_nat(x_54); -lean_dec(x_54); -x_64 = 1; -x_65 = lean_usize_sub(x_63, x_64); -x_66 = lean_usize_land(x_62, x_65); -x_67 = lean_array_uget(x_53, x_66); -x_68 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__12(x_2, x_67); -if (x_68 == 0) -{ -lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; uint8_t x_78; -x_69 = lean_unsigned_to_nat(1u); -x_70 = lean_nat_add(x_52, x_69); -lean_dec(x_52); -x_71 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_71, 0, x_2); -lean_ctor_set(x_71, 1, x_3); -lean_ctor_set(x_71, 2, x_67); -x_72 = lean_array_uset(x_53, x_66, x_71); -x_73 = lean_unsigned_to_nat(4u); -x_74 = lean_nat_mul(x_70, x_73); -x_75 = lean_unsigned_to_nat(3u); -x_76 = lean_nat_div(x_74, x_75); -lean_dec(x_74); -x_77 = lean_array_get_size(x_72); -x_78 = lean_nat_dec_le(x_76, x_77); -lean_dec(x_77); -lean_dec(x_76); -if (x_78 == 0) -{ -lean_object* x_79; lean_object* x_80; uint8_t x_81; -x_79 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__13(x_72); -x_80 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_80, 0, x_70); -lean_ctor_set(x_80, 1, x_79); -x_81 = 1; -lean_ctor_set(x_1, 0, x_80); -lean_ctor_set_uint8(x_1, sizeof(void*)*2, x_81); -return x_1; -} -else -{ -lean_object* x_82; uint8_t x_83; -x_82 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_82, 0, x_70); -lean_ctor_set(x_82, 1, x_72); -x_83 = 1; -lean_ctor_set(x_1, 0, x_82); -lean_ctor_set_uint8(x_1, sizeof(void*)*2, x_83); -return x_1; -} -} -else -{ -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_84 = lean_box(0); -x_85 = lean_array_uset(x_53, x_66, x_84); -x_86 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__16(x_2, x_3, x_67); -x_87 = lean_array_uset(x_85, x_66, x_86); -x_88 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_88, 0, x_52); -lean_ctor_set(x_88, 1, x_87); -x_89 = 1; -lean_ctor_set(x_1, 0, x_88); -lean_ctor_set_uint8(x_1, sizeof(void*)*2, x_89); -return x_1; -} -} -} -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; uint64_t x_96; uint64_t x_97; uint64_t x_98; uint64_t x_99; uint64_t x_100; uint64_t x_101; uint64_t x_102; size_t x_103; size_t x_104; size_t x_105; size_t x_106; size_t x_107; lean_object* x_108; uint8_t x_109; -x_90 = lean_ctor_get(x_1, 0); -x_91 = lean_ctor_get(x_1, 1); -lean_inc(x_91); -lean_inc(x_90); -lean_dec(x_1); -x_92 = lean_ctor_get(x_90, 0); -lean_inc(x_92); -x_93 = lean_ctor_get(x_90, 1); -lean_inc(x_93); -if (lean_is_exclusive(x_90)) { - lean_ctor_release(x_90, 0); - lean_ctor_release(x_90, 1); - x_94 = x_90; -} else { - lean_dec_ref(x_90); - x_94 = lean_box(0); -} -x_95 = lean_array_get_size(x_93); -x_96 = l_Lean_Name_hash___override(x_2); -x_97 = 32; -x_98 = lean_uint64_shift_right(x_96, x_97); -x_99 = lean_uint64_xor(x_96, x_98); -x_100 = 16; -x_101 = lean_uint64_shift_right(x_99, x_100); -x_102 = lean_uint64_xor(x_99, x_101); -x_103 = lean_uint64_to_usize(x_102); -x_104 = lean_usize_of_nat(x_95); -lean_dec(x_95); -x_105 = 1; -x_106 = lean_usize_sub(x_104, x_105); -x_107 = lean_usize_land(x_103, x_106); -x_108 = lean_array_uget(x_93, x_107); -x_109 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__12(x_2, x_108); -if (x_109 == 0) -{ -lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; uint8_t x_119; -x_110 = lean_unsigned_to_nat(1u); -x_111 = lean_nat_add(x_92, x_110); -lean_dec(x_92); -x_112 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_112, 0, x_2); -lean_ctor_set(x_112, 1, x_3); -lean_ctor_set(x_112, 2, x_108); -x_113 = lean_array_uset(x_93, x_107, x_112); -x_114 = lean_unsigned_to_nat(4u); -x_115 = lean_nat_mul(x_111, x_114); -x_116 = lean_unsigned_to_nat(3u); -x_117 = lean_nat_div(x_115, x_116); -lean_dec(x_115); -x_118 = lean_array_get_size(x_113); -x_119 = lean_nat_dec_le(x_117, x_118); -lean_dec(x_118); -lean_dec(x_117); -if (x_119 == 0) -{ -lean_object* x_120; lean_object* x_121; uint8_t x_122; lean_object* x_123; -x_120 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__13(x_113); -if (lean_is_scalar(x_94)) { - x_121 = lean_alloc_ctor(0, 2, 0); -} else { - x_121 = x_94; -} -lean_ctor_set(x_121, 0, x_111); -lean_ctor_set(x_121, 1, x_120); -x_122 = 1; -x_123 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_123, 0, x_121); -lean_ctor_set(x_123, 1, x_91); -lean_ctor_set_uint8(x_123, sizeof(void*)*2, x_122); -return x_123; -} -else -{ -lean_object* x_124; uint8_t x_125; lean_object* x_126; -if (lean_is_scalar(x_94)) { - x_124 = lean_alloc_ctor(0, 2, 0); -} else { - x_124 = x_94; -} -lean_ctor_set(x_124, 0, x_111); -lean_ctor_set(x_124, 1, x_113); -x_125 = 1; -x_126 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_126, 0, x_124); -lean_ctor_set(x_126, 1, x_91); -lean_ctor_set_uint8(x_126, sizeof(void*)*2, x_125); -return x_126; -} -} -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; -x_127 = lean_box(0); -x_128 = lean_array_uset(x_93, x_107, x_127); -x_129 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__16(x_2, x_3, x_108); -x_130 = lean_array_uset(x_128, x_107, x_129); -if (lean_is_scalar(x_94)) { - x_131 = lean_alloc_ctor(0, 2, 0); -} else { - x_131 = x_94; -} -lean_ctor_set(x_131, 0, x_92); -lean_ctor_set(x_131, 1, x_130); -x_132 = 1; -x_133 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_133, 0, x_131); -lean_ctor_set(x_133, 1, x_91); -lean_ctor_set_uint8(x_133, sizeof(void*)*2, x_132); -return x_133; -} -} -} -} -} -static lean_object* _init_l_Lean_SMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__17___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(10u); -x_2 = lean_unsigned_to_nat(1u); -x_3 = l_Nat_nextPowerOfTwo_go(x_1, x_2, lean_box(0)); -return x_3; -} -} -static lean_object* _init_l_Lean_SMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__17___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_SMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__17___closed__1; -x_3 = lean_mk_array(x_2, x_1); -return x_3; -} -} -static lean_object* _init_l_Lean_SMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__17___closed__3() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(0u); -x_2 = l_Lean_SMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__17___closed__2; -x_3 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set(x_3, 1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_SMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__17___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_SMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__17___closed__5() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_SMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__17___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_SMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__17___closed__6() { -_start: -{ -uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = 1; -x_2 = l_Lean_SMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__17___closed__3; -x_3 = l_Lean_SMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__17___closed__5; -x_4 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_4, 0, x_2); -lean_ctor_set(x_4, 1, x_3); -lean_ctor_set_uint8(x_4, sizeof(void*)*2, x_1); -return x_4; -} -} -static lean_object* _init_l_Lean_SMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__17() { -_start: -{ -lean_object* x_1; -x_1 = l_Lean_SMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__17___closed__6; -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____lambda__1___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean.Meta.Tactic.Grind.TheoremPatterns", 38, 38); -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____lambda__1___closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean.Meta.Grind.theoremPatternsExt", 34, 34); -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____lambda__1___closed__3() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("unreachable code has been reached", 33, 33); -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____lambda__1___closed__4() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____lambda__1___closed__1; -x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____lambda__1___closed__2; -x_3 = lean_unsigned_to_nat(41u); -x_4 = lean_unsigned_to_nat(46u); -x_5 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____lambda__1___closed__3; -x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); -return x_6; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____lambda__1(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; -x_3 = lean_ctor_get(x_2, 3); -lean_inc(x_3); -if (lean_obj_tag(x_3) == 0) -{ -lean_object* x_4; lean_object* x_5; -lean_dec(x_2); -lean_dec(x_1); -x_4 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____lambda__1___closed__4; -x_5 = l_panic___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__1(x_4); -return x_5; -} -else -{ -uint8_t x_6; -x_6 = !lean_is_exclusive(x_3); -if (x_6 == 0) -{ -lean_object* x_7; lean_object* x_8; -x_7 = lean_ctor_get(x_3, 0); -x_8 = lean_ctor_get(x_3, 1); -lean_dec(x_8); -if (lean_obj_tag(x_7) == 2) -{ -lean_object* x_9; lean_object* x_10; -x_9 = lean_ctor_get(x_7, 0); -lean_inc(x_9); -lean_dec(x_7); -lean_inc(x_1); -x_10 = l_Lean_SMap_find_x3f___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__2(x_1, x_9); -if (lean_obj_tag(x_10) == 0) -{ -lean_object* x_11; lean_object* x_12; -x_11 = lean_box(0); -lean_ctor_set(x_3, 1, x_11); -lean_ctor_set(x_3, 0, x_2); -x_12 = l_Lean_SMap_insert___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__7(x_1, x_9, x_3); -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); -lean_ctor_set(x_3, 1, x_13); -lean_ctor_set(x_3, 0, x_2); -x_14 = l_Lean_SMap_insert___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__7(x_1, x_9, x_3); -return x_14; -} -} -else -{ -lean_object* x_15; lean_object* x_16; -lean_free_object(x_3); -lean_dec(x_7); -lean_dec(x_2); -lean_dec(x_1); -x_15 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____lambda__1___closed__4; -x_16 = l_panic___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__1(x_15); -return x_16; -} -} -else -{ -lean_object* x_17; -x_17 = lean_ctor_get(x_3, 0); -lean_inc(x_17); -lean_dec(x_3); -if (lean_obj_tag(x_17) == 2) -{ -lean_object* x_18; lean_object* x_19; -x_18 = lean_ctor_get(x_17, 0); -lean_inc(x_18); -lean_dec(x_17); -lean_inc(x_1); -x_19 = l_Lean_SMap_find_x3f___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__2(x_1, x_18); -if (lean_obj_tag(x_19) == 0) -{ -lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_20 = lean_box(0); -x_21 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_21, 0, x_2); -lean_ctor_set(x_21, 1, x_20); -x_22 = l_Lean_SMap_insert___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__7(x_1, x_18, x_21); -return x_22; -} -else -{ -lean_object* x_23; lean_object* x_24; lean_object* x_25; -x_23 = lean_ctor_get(x_19, 0); -lean_inc(x_23); -lean_dec(x_19); -x_24 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_24, 0, x_2); -lean_ctor_set(x_24, 1, x_23); -x_25 = l_Lean_SMap_insert___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__7(x_1, x_18, x_24); -return x_25; -} -} -else -{ -lean_object* x_26; lean_object* x_27; -lean_dec(x_17); -lean_dec(x_2); -lean_dec(x_1); -x_26 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____lambda__1___closed__4; -x_27 = l_panic___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__1(x_26); -return x_27; -} -} -} -} -} -static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean", 4, 4); -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Meta", 4, 4); -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____closed__3() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Grind", 5, 5); -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____closed__4() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("theoremPatternsExt", 18, 18); -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____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_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____closed__1; -x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____closed__2; -x_3 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____closed__3; -x_4 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____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_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____closed__6() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____lambda__1), 2, 0); -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____closed__7() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_id___rarg___boxed), 1, 0); -return x_1; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307_(lean_object* x_1) { -_start: -{ -lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____closed__5; -x_3 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____closed__6; -x_4 = l_Lean_SMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__17; -x_5 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____closed__7; -x_6 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_6, 0, x_2); -lean_ctor_set(x_6, 1, x_3); -lean_ctor_set(x_6, 2, x_4); -lean_ctor_set(x_6, 3, x_5); -x_7 = l_Lean_registerSimpleScopedEnvExtension___rarg(x_6, x_1); -return x_7; -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__5___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_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__5(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_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__4___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_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__4(x_1, x_4, x_3); -lean_dec(x_3); -return x_5; -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__3___boxed(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; -x_3 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__3(x_1, x_2); -lean_dec(x_2); -return x_3; -} -} -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__6___boxed(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; -x_3 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__6(x_1, x_2); -lean_dec(x_2); -lean_dec(x_1); -return x_3; -} -} -LEAN_EXPORT lean_object* l_Lean_SMap_find_x3f___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__2___boxed(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; -x_3 = l_Lean_SMap_find_x3f___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__2(x_1, x_2); -lean_dec(x_2); -return x_3; -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____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: -{ -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_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__10(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_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__9___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: -{ -size_t x_6; size_t x_7; lean_object* x_8; -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_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__9(x_1, x_6, x_7, x_4, x_5); -return x_8; -} -} -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__12___boxed(lean_object* x_1, lean_object* x_2) { -_start: -{ -uint8_t x_3; lean_object* x_4; -x_3 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__12(x_1, x_2); -lean_dec(x_2); -lean_dec(x_1); -x_4 = lean_box(x_3); -return x_4; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Eq", 2, 2); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__1; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__3() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("HEq", 3, 3); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___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_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__3; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__5() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Iff", 3, 3); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___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_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__5; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__7() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("And", 3, 3); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___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_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__7; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__9() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Or", 2, 2); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___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_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__9; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__11() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Not", 3, 3); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___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_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__11; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___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_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__12; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set(x_3, 1, x_1); -return x_3; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__14() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__10; -x_2 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__13; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set(x_3, 1, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__15() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__8; -x_2 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__14; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set(x_3, 1, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__16() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__6; -x_2 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__15; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set(x_3, 1, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__17() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__4; -x_2 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__16; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set(x_3, 1, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__18() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__2; -x_2 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__17; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set(x_3, 1, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__19() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__18; -x_2 = lean_array_mk(x_1); -return x_2; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames() { -_start: -{ -lean_object* x_1; -x_1 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__19; -return x_1; -} -} -LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_isForbidden(lean_object* x_1) { -_start: -{ -lean_object* x_2; uint8_t x_3; -x_2 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames; -x_3 = l_Array_contains___at_Lean_registerInternalExceptionId___spec__1(x_2, x_1); -return x_3; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_isForbidden___boxed(lean_object* x_1) { -_start: -{ -uint8_t x_2; lean_object* x_3; -x_2 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_isForbidden(x_1); -lean_dec(x_1); -x_3 = lean_box(x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_dontCare___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("[grind_dontcare]", 16, 16); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_dontCare___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_dontCare___closed__1; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_dontCare___closed__3() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_dontCare___closed__2; -x_3 = l_Lean_Expr_const___override(x_2, x_1); -return x_3; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_dontCare() { -_start: -{ -lean_object* x_1; -x_1 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_dontCare___closed__3; -return x_1; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_mkGroundPattern___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("grind", 5, 5); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_mkGroundPattern___closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("ground_pat", 10, 10); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_mkGroundPattern___closed__3() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_mkGroundPattern___closed__1; -x_2 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_mkGroundPattern___closed__2; -x_3 = l_Lean_Name_mkStr2(x_1, x_2); -return x_3; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_mkGroundPattern(lean_object* x_1) { -_start: -{ -lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_mkGroundPattern___closed__3; -x_3 = l_Lean_mkAnnotation(x_2, x_1); -return x_3; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_groundPattern_x3f(lean_object* x_1) { -_start: -{ -lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_mkGroundPattern___closed__3; -x_3 = l_Lean_annotation_x3f(x_2, x_1); -return x_3; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_groundPattern_x3f___boxed(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_groundPattern_x3f(x_1); -lean_dec(x_1); -return x_2; -} -} -LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_isGroundPattern(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_groundPattern_x3f(x_1); -if (lean_obj_tag(x_2) == 0) -{ -uint8_t x_3; -x_3 = 0; -return x_3; -} -else -{ -uint8_t x_4; -lean_dec(x_2); -x_4 = 1; -return x_4; -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_isGroundPattern___boxed(lean_object* x_1) { -_start: -{ -uint8_t x_2; lean_object* x_3; -x_2 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_isGroundPattern(x_1); -lean_dec(x_1); -x_3 = lean_box(x_2); -return x_3; -} -} -LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_isAtomicPattern(lean_object* x_1) { -_start: -{ -uint8_t x_2; -x_2 = l_Lean_Expr_isBVar(x_1); -if (x_2 == 0) -{ -lean_object* x_3; uint8_t x_4; -x_3 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_dontCare; -x_4 = lean_expr_eqv(x_1, x_3); -if (x_4 == 0) -{ -uint8_t x_5; -x_5 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_isGroundPattern(x_1); -return x_5; -} -else -{ -uint8_t x_6; -x_6 = 1; -return x_6; -} -} -else -{ -uint8_t x_7; -x_7 = 1; -return x_7; -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_isAtomicPattern___boxed(lean_object* x_1) { -_start: -{ -uint8_t x_2; lean_object* x_3; -x_2 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_isAtomicPattern(x_1); -lean_dec(x_1); -x_3 = lean_box(x_2); -return x_3; -} -} -static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked(" ", 1, 1); -return x_1; -} -} -static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___closed__1; -x_2 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; -} -} -static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___closed__3() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___closed__2; -x_2 = l_Lean_MessageData_ofFormat(x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_4 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___closed__3; -x_5 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_5, 0, x_2); -lean_ctor_set(x_5, 1, x_4); -x_6 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_6, 0, x_5); -lean_ctor_set(x_6, 1, x_1); -x_7 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_7, 0, x_6); -return x_7; -} -} -static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___boxed), 3, 0); -return x_1; -} -} -static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("(", 1, 1); -return x_1; -} -} -static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked(")", 1, 1); -return x_1; -} -} -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___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) { -_start: -{ -uint8_t x_7; -x_7 = lean_usize_dec_lt(x_5, x_4); -if (x_7 == 0) -{ -return x_6; -} -else -{ -lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; -x_8 = lean_array_uget(x_3, x_5); -lean_inc(x_8); -x_9 = l_Lean_Meta_Grind_ppPattern(x_8); -x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1; -x_11 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_isAtomicPattern(x_8); -lean_dec(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 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2; -x_13 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3; -x_14 = l_Lean_MessageData_bracket(x_12, x_9, x_13); -x_15 = lean_box(0); -x_16 = lean_apply_3(x_10, x_14, x_6, x_15); -if (lean_obj_tag(x_16) == 0) -{ -lean_object* x_17; -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -lean_dec(x_16); -return x_17; -} -else -{ -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_5, x_19); -x_5 = x_20; -x_6 = x_18; -goto _start; -} -} -else -{ -lean_object* x_22; lean_object* x_23; -x_22 = lean_box(0); -x_23 = lean_apply_3(x_10, x_9, x_6, x_22); -if (lean_obj_tag(x_23) == 0) -{ -lean_object* x_24; -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -lean_dec(x_23); -return x_24; -} -else -{ -lean_object* x_25; size_t x_26; size_t x_27; -x_25 = lean_ctor_get(x_23, 0); -lean_inc(x_25); -lean_dec(x_23); -x_26 = 1; -x_27 = lean_usize_add(x_5, x_26); -x_5 = x_27; -x_6 = x_25; -goto _start; -} -} -} -} -} -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { -_start: -{ -uint8_t x_7; -x_7 = lean_usize_dec_lt(x_5, x_4); -if (x_7 == 0) -{ -return x_6; -} -else -{ -lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; -x_8 = lean_array_uget(x_3, x_5); -lean_inc(x_8); -x_9 = l_Lean_Meta_Grind_ppPattern(x_8); -x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1; -x_11 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_isAtomicPattern(x_8); -lean_dec(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 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2; -x_13 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3; -x_14 = l_Lean_MessageData_bracket(x_12, x_9, x_13); -x_15 = lean_box(0); -x_16 = lean_apply_3(x_10, x_14, x_6, x_15); -if (lean_obj_tag(x_16) == 0) -{ -lean_object* x_17; -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -lean_dec(x_16); -return x_17; -} -else -{ -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_5, x_19); -x_5 = x_20; -x_6 = x_18; -goto _start; -} -} -else -{ -lean_object* x_22; lean_object* x_23; -x_22 = lean_box(0); -x_23 = lean_apply_3(x_10, x_9, x_6, x_22); -if (lean_obj_tag(x_23) == 0) -{ -lean_object* x_24; -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -lean_dec(x_23); -return x_24; -} -else -{ -lean_object* x_25; size_t x_26; size_t x_27; -x_25 = lean_ctor_get(x_23, 0); -lean_inc(x_25); -lean_dec(x_23); -x_26 = 1; -x_27 = lean_usize_add(x_5, x_26); -x_5 = x_27; -x_6 = x_25; -goto _start; -} -} -} -} -} -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { -_start: -{ -uint8_t x_7; -x_7 = lean_usize_dec_lt(x_5, x_4); -if (x_7 == 0) -{ -return x_6; -} -else -{ -lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; -x_8 = lean_array_uget(x_3, x_5); -lean_inc(x_8); -x_9 = l_Lean_Meta_Grind_ppPattern(x_8); -x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1; -x_11 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_isAtomicPattern(x_8); -lean_dec(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 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2; -x_13 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3; -x_14 = l_Lean_MessageData_bracket(x_12, x_9, x_13); -x_15 = lean_box(0); -x_16 = lean_apply_3(x_10, x_14, x_6, x_15); -if (lean_obj_tag(x_16) == 0) -{ -lean_object* x_17; -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -lean_dec(x_16); -return x_17; -} -else -{ -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_5, x_19); -x_5 = x_20; -x_6 = x_18; -goto _start; -} -} -else -{ -lean_object* x_22; lean_object* x_23; -x_22 = lean_box(0); -x_23 = lean_apply_3(x_10, x_9, x_6, x_22); -if (lean_obj_tag(x_23) == 0) -{ -lean_object* x_24; -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -lean_dec(x_23); -return x_24; -} -else -{ -lean_object* x_25; size_t x_26; size_t x_27; -x_25 = lean_ctor_get(x_23, 0); -lean_inc(x_25); -lean_dec(x_23); -x_26 = 1; -x_27 = lean_usize_add(x_5, x_26); -x_5 = x_27; -x_6 = x_25; -goto _start; -} -} -} -} -} -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { -_start: -{ -uint8_t x_7; -x_7 = lean_usize_dec_lt(x_5, x_4); -if (x_7 == 0) -{ -return x_6; -} -else -{ -lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; -x_8 = lean_array_uget(x_3, x_5); -lean_inc(x_8); -x_9 = l_Lean_Meta_Grind_ppPattern(x_8); -x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1; -x_11 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_isAtomicPattern(x_8); -lean_dec(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 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2; -x_13 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3; -x_14 = l_Lean_MessageData_bracket(x_12, x_9, x_13); -x_15 = lean_box(0); -x_16 = lean_apply_3(x_10, x_14, x_6, x_15); -if (lean_obj_tag(x_16) == 0) -{ -lean_object* x_17; -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -lean_dec(x_16); -return x_17; -} -else -{ -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_5, x_19); -x_5 = x_20; -x_6 = x_18; -goto _start; -} -} -else -{ -lean_object* x_22; lean_object* x_23; -x_22 = lean_box(0); -x_23 = lean_apply_3(x_10, x_9, x_6, x_22); -if (lean_obj_tag(x_23) == 0) -{ -lean_object* x_24; -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -lean_dec(x_23); -return x_24; -} -else -{ -lean_object* x_25; size_t x_26; size_t x_27; -x_25 = lean_ctor_get(x_23, 0); -lean_inc(x_25); -lean_dec(x_23); -x_26 = 1; -x_27 = lean_usize_add(x_5, x_26); -x_5 = x_27; -x_6 = x_25; -goto _start; -} -} -} -} -} -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { -_start: -{ -uint8_t x_7; -x_7 = lean_usize_dec_lt(x_5, x_4); -if (x_7 == 0) -{ -return x_6; -} -else -{ -lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; -x_8 = lean_array_uget(x_3, x_5); -lean_inc(x_8); -x_9 = l_Lean_Meta_Grind_ppPattern(x_8); -x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1; -x_11 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_isAtomicPattern(x_8); -lean_dec(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 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2; -x_13 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3; -x_14 = l_Lean_MessageData_bracket(x_12, x_9, x_13); -x_15 = lean_box(0); -x_16 = lean_apply_3(x_10, x_14, x_6, x_15); -if (lean_obj_tag(x_16) == 0) -{ -lean_object* x_17; -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -lean_dec(x_16); -return x_17; -} -else -{ -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_5, x_19); -x_5 = x_20; -x_6 = x_18; -goto _start; -} -} -else -{ -lean_object* x_22; lean_object* x_23; -x_22 = lean_box(0); -x_23 = lean_apply_3(x_10, x_9, x_6, x_22); -if (lean_obj_tag(x_23) == 0) -{ -lean_object* x_24; -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -lean_dec(x_23); -return x_24; -} -else -{ -lean_object* x_25; size_t x_26; size_t x_27; -x_25 = lean_ctor_get(x_23, 0); -lean_inc(x_25); -lean_dec(x_23); -x_26 = 1; -x_27 = lean_usize_add(x_5, x_26); -x_5 = x_27; -x_6 = x_25; -goto _start; -} -} -} -} -} -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { -_start: -{ -uint8_t x_7; -x_7 = lean_usize_dec_lt(x_5, x_4); -if (x_7 == 0) -{ -return x_6; -} -else -{ -lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; -x_8 = lean_array_uget(x_3, x_5); -lean_inc(x_8); -x_9 = l_Lean_Meta_Grind_ppPattern(x_8); -x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1; -x_11 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_isAtomicPattern(x_8); -lean_dec(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 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2; -x_13 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3; -x_14 = l_Lean_MessageData_bracket(x_12, x_9, x_13); -x_15 = lean_box(0); -x_16 = lean_apply_3(x_10, x_14, x_6, x_15); -if (lean_obj_tag(x_16) == 0) -{ -lean_object* x_17; -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -lean_dec(x_16); -return x_17; -} -else -{ -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_5, x_19); -x_5 = x_20; -x_6 = x_18; -goto _start; -} -} -else -{ -lean_object* x_22; lean_object* x_23; -x_22 = lean_box(0); -x_23 = lean_apply_3(x_10, x_9, x_6, x_22); -if (lean_obj_tag(x_23) == 0) -{ -lean_object* x_24; -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -lean_dec(x_23); -return x_24; -} -else -{ -lean_object* x_25; size_t x_26; size_t x_27; -x_25 = lean_ctor_get(x_23, 0); -lean_inc(x_25); -lean_dec(x_23); -x_26 = 1; -x_27 = lean_usize_add(x_5, x_26); -x_5 = x_27; -x_6 = x_25; -goto _start; -} -} -} -} -} -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { -_start: -{ -uint8_t x_7; -x_7 = lean_usize_dec_lt(x_5, x_4); -if (x_7 == 0) -{ -return x_6; -} -else -{ -lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; -x_8 = lean_array_uget(x_3, x_5); -lean_inc(x_8); -x_9 = l_Lean_Meta_Grind_ppPattern(x_8); -x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1; -x_11 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_isAtomicPattern(x_8); -lean_dec(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 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2; -x_13 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3; -x_14 = l_Lean_MessageData_bracket(x_12, x_9, x_13); -x_15 = lean_box(0); -x_16 = lean_apply_3(x_10, x_14, x_6, x_15); -if (lean_obj_tag(x_16) == 0) -{ -lean_object* x_17; -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -lean_dec(x_16); -return x_17; -} -else -{ -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_5, x_19); -x_5 = x_20; -x_6 = x_18; -goto _start; -} -} -else -{ -lean_object* x_22; lean_object* x_23; -x_22 = lean_box(0); -x_23 = lean_apply_3(x_10, x_9, x_6, x_22); -if (lean_obj_tag(x_23) == 0) -{ -lean_object* x_24; -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -lean_dec(x_23); -return x_24; -} -else -{ -lean_object* x_25; size_t x_26; size_t x_27; -x_25 = lean_ctor_get(x_23, 0); -lean_inc(x_25); -lean_dec(x_23); -x_26 = 1; -x_27 = lean_usize_add(x_5, x_26); -x_5 = x_27; -x_6 = x_25; -goto _start; -} -} -} -} -} -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__8(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { -_start: -{ -uint8_t x_7; -x_7 = lean_usize_dec_lt(x_5, x_4); -if (x_7 == 0) -{ -return x_6; -} -else -{ -lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; -x_8 = lean_array_uget(x_3, x_5); -lean_inc(x_8); -x_9 = l_Lean_Meta_Grind_ppPattern(x_8); -x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1; -x_11 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_isAtomicPattern(x_8); -lean_dec(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 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2; -x_13 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3; -x_14 = l_Lean_MessageData_bracket(x_12, x_9, x_13); -x_15 = lean_box(0); -x_16 = lean_apply_3(x_10, x_14, x_6, x_15); -if (lean_obj_tag(x_16) == 0) -{ -lean_object* x_17; -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -lean_dec(x_16); -return x_17; -} -else -{ -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_5, x_19); -x_5 = x_20; -x_6 = x_18; -goto _start; -} -} -else -{ -lean_object* x_22; lean_object* x_23; -x_22 = lean_box(0); -x_23 = lean_apply_3(x_10, x_9, x_6, x_22); -if (lean_obj_tag(x_23) == 0) -{ -lean_object* x_24; -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -lean_dec(x_23); -return x_24; -} -else -{ -lean_object* x_25; size_t x_26; size_t x_27; -x_25 = lean_ctor_get(x_23, 0); -lean_inc(x_25); -lean_dec(x_23); -x_26 = 1; -x_27 = lean_usize_add(x_5, x_26); -x_5 = x_27; -x_6 = x_25; -goto _start; -} -} -} -} -} -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__9(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { -_start: -{ -uint8_t x_7; -x_7 = lean_usize_dec_lt(x_5, x_4); -if (x_7 == 0) -{ -return x_6; -} -else -{ -lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; -x_8 = lean_array_uget(x_3, x_5); -lean_inc(x_8); -x_9 = l_Lean_Meta_Grind_ppPattern(x_8); -x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1; -x_11 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_isAtomicPattern(x_8); -lean_dec(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 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2; -x_13 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3; -x_14 = l_Lean_MessageData_bracket(x_12, x_9, x_13); -x_15 = lean_box(0); -x_16 = lean_apply_3(x_10, x_14, x_6, x_15); -if (lean_obj_tag(x_16) == 0) -{ -lean_object* x_17; -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -lean_dec(x_16); -return x_17; -} -else -{ -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_5, x_19); -x_5 = x_20; -x_6 = x_18; -goto _start; -} -} -else -{ -lean_object* x_22; lean_object* x_23; -x_22 = lean_box(0); -x_23 = lean_apply_3(x_10, x_9, x_6, x_22); -if (lean_obj_tag(x_23) == 0) -{ -lean_object* x_24; -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -lean_dec(x_23); -return x_24; -} -else -{ -lean_object* x_25; size_t x_26; size_t x_27; -x_25 = lean_ctor_get(x_23, 0); -lean_inc(x_25); -lean_dec(x_23); -x_26 = 1; -x_27 = lean_usize_add(x_5, x_26); -x_5 = x_27; -x_6 = x_25; -goto _start; -} -} -} -} -} -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__10(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { -_start: -{ -uint8_t x_7; -x_7 = lean_usize_dec_lt(x_5, x_4); -if (x_7 == 0) -{ -return x_6; -} -else -{ -lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; -x_8 = lean_array_uget(x_3, x_5); -lean_inc(x_8); -x_9 = l_Lean_Meta_Grind_ppPattern(x_8); -x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1; -x_11 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_isAtomicPattern(x_8); -lean_dec(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 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2; -x_13 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3; -x_14 = l_Lean_MessageData_bracket(x_12, x_9, x_13); -x_15 = lean_box(0); -x_16 = lean_apply_3(x_10, x_14, x_6, x_15); -if (lean_obj_tag(x_16) == 0) -{ -lean_object* x_17; -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -lean_dec(x_16); -return x_17; -} -else -{ -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_5, x_19); -x_5 = x_20; -x_6 = x_18; -goto _start; -} -} -else -{ -lean_object* x_22; lean_object* x_23; -x_22 = lean_box(0); -x_23 = lean_apply_3(x_10, x_9, x_6, x_22); -if (lean_obj_tag(x_23) == 0) -{ -lean_object* x_24; -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -lean_dec(x_23); -return x_24; -} -else -{ -lean_object* x_25; size_t x_26; size_t x_27; -x_25 = lean_ctor_get(x_23, 0); -lean_inc(x_25); -lean_dec(x_23); -x_26 = 1; -x_27 = lean_usize_add(x_5, x_26); -x_5 = x_27; -x_6 = x_25; -goto _start; -} -} -} -} -} -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__11(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { -_start: -{ -uint8_t x_7; -x_7 = lean_usize_dec_lt(x_5, x_4); -if (x_7 == 0) -{ -return x_6; -} -else -{ -lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; -x_8 = lean_array_uget(x_3, x_5); -lean_inc(x_8); -x_9 = l_Lean_Meta_Grind_ppPattern(x_8); -x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1; -x_11 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_isAtomicPattern(x_8); -lean_dec(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 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2; -x_13 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3; -x_14 = l_Lean_MessageData_bracket(x_12, x_9, x_13); -x_15 = lean_box(0); -x_16 = lean_apply_3(x_10, x_14, x_6, x_15); -if (lean_obj_tag(x_16) == 0) -{ -lean_object* x_17; -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -lean_dec(x_16); -return x_17; -} -else -{ -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_5, x_19); -x_5 = x_20; -x_6 = x_18; -goto _start; -} -} -else -{ -lean_object* x_22; lean_object* x_23; -x_22 = lean_box(0); -x_23 = lean_apply_3(x_10, x_9, x_6, x_22); -if (lean_obj_tag(x_23) == 0) -{ -lean_object* x_24; -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -lean_dec(x_23); -return x_24; -} -else -{ -lean_object* x_25; size_t x_26; size_t x_27; -x_25 = lean_ctor_get(x_23, 0); -lean_inc(x_25); -lean_dec(x_23); -x_26 = 1; -x_27 = lean_usize_add(x_5, x_26); -x_5 = x_27; -x_6 = x_25; -goto _start; -} -} -} -} -} -static lean_object* _init_l_Lean_Meta_Grind_ppPattern___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("#", 1, 1); -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_Grind_ppPattern___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_Grind_ppPattern___closed__1; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Meta_Grind_ppPattern___closed__3() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("", 0, 0); -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_Grind_ppPattern___closed__4() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_Grind_ppPattern___closed__3; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Meta_Grind_ppPattern___closed__5() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_levelZero; -x_2 = l_Lean_Expr_sort___override(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Meta_Grind_ppPattern___closed__6() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("\?", 1, 1); -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_Grind_ppPattern___closed__7() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_Grind_ppPattern___closed__6; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Meta_Grind_ppPattern___closed__8() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("`[", 2, 2); -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_Grind_ppPattern___closed__9() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_Grind_ppPattern___closed__8; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Meta_Grind_ppPattern___closed__10() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("]", 1, 1); -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_Grind_ppPattern___closed__11() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_Grind_ppPattern___closed__10; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ppPattern(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_groundPattern_x3f(x_1); -if (lean_obj_tag(x_2) == 0) -{ -lean_object* x_3; uint8_t x_4; -x_3 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_dontCare; -x_4 = lean_expr_eqv(x_1, x_3); -if (x_4 == 0) -{ -switch (lean_obj_tag(x_1)) { -case 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; -x_5 = lean_ctor_get(x_1, 0); -lean_inc(x_5); -lean_dec(x_1); -x_6 = l___private_Init_Data_Repr_0__Nat_reprFast(x_5); -x_7 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_7, 0, x_6); -x_8 = l_Lean_MessageData_ofFormat(x_7); -x_9 = l_Lean_Meta_Grind_ppPattern___closed__2; -x_10 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_10, 0, x_9); -lean_ctor_set(x_10, 1, x_8); -x_11 = l_Lean_Meta_Grind_ppPattern___closed__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; -} -case 1: -{ -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; size_t x_26; size_t x_27; lean_object* x_28; -x_13 = l_Lean_Expr_getAppFn(x_1); -x_14 = l_Lean_MessageData_ofExpr(x_13); -x_15 = l_Lean_Meta_Grind_ppPattern___closed__4; -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 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_17, 0, x_16); -lean_ctor_set(x_17, 1, x_15); -x_18 = lean_unsigned_to_nat(0u); -x_19 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_18); -x_20 = l_Lean_Meta_Grind_ppPattern___closed__5; -lean_inc(x_19); -x_21 = lean_mk_array(x_19, x_20); -x_22 = lean_unsigned_to_nat(1u); -x_23 = lean_nat_sub(x_19, x_22); -lean_dec(x_19); -x_24 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_21, x_23); -x_25 = lean_box(0); -x_26 = lean_array_size(x_24); -x_27 = 0; -x_28 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1(x_24, x_25, x_24, x_26, x_27, x_17); -lean_dec(x_24); -return x_28; -} -case 2: -{ -lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; 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_29 = l_Lean_Expr_getAppFn(x_1); -x_30 = l_Lean_MessageData_ofExpr(x_29); -x_31 = l_Lean_Meta_Grind_ppPattern___closed__4; -x_32 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_32, 0, x_31); -lean_ctor_set(x_32, 1, x_30); -x_33 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_33, 0, x_32); -lean_ctor_set(x_33, 1, x_31); -x_34 = lean_unsigned_to_nat(0u); -x_35 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_34); -x_36 = l_Lean_Meta_Grind_ppPattern___closed__5; -lean_inc(x_35); -x_37 = lean_mk_array(x_35, x_36); -x_38 = lean_unsigned_to_nat(1u); -x_39 = lean_nat_sub(x_35, x_38); -lean_dec(x_35); -x_40 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_37, x_39); -x_41 = lean_box(0); -x_42 = lean_array_size(x_40); -x_43 = 0; -x_44 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__2(x_40, x_41, x_40, x_42, x_43, x_33); -lean_dec(x_40); -return x_44; -} -case 3: -{ -lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; size_t x_58; size_t x_59; lean_object* x_60; -x_45 = l_Lean_Expr_getAppFn(x_1); -x_46 = l_Lean_MessageData_ofExpr(x_45); -x_47 = l_Lean_Meta_Grind_ppPattern___closed__4; -x_48 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_48, 0, x_47); -lean_ctor_set(x_48, 1, x_46); -x_49 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_49, 0, x_48); -lean_ctor_set(x_49, 1, x_47); -x_50 = lean_unsigned_to_nat(0u); -x_51 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_50); -x_52 = l_Lean_Meta_Grind_ppPattern___closed__5; -lean_inc(x_51); -x_53 = lean_mk_array(x_51, x_52); -x_54 = lean_unsigned_to_nat(1u); -x_55 = lean_nat_sub(x_51, x_54); -lean_dec(x_51); -x_56 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_53, x_55); -x_57 = lean_box(0); -x_58 = lean_array_size(x_56); -x_59 = 0; -x_60 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__3(x_56, x_57, x_56, x_58, x_59, x_49); -lean_dec(x_56); -return x_60; -} -case 4: -{ -lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; size_t x_74; size_t x_75; lean_object* x_76; -x_61 = l_Lean_Expr_getAppFn(x_1); -x_62 = l_Lean_MessageData_ofExpr(x_61); -x_63 = l_Lean_Meta_Grind_ppPattern___closed__4; -x_64 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_64, 0, x_63); -lean_ctor_set(x_64, 1, x_62); -x_65 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_65, 0, x_64); -lean_ctor_set(x_65, 1, x_63); -x_66 = lean_unsigned_to_nat(0u); -x_67 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_66); -x_68 = l_Lean_Meta_Grind_ppPattern___closed__5; -lean_inc(x_67); -x_69 = lean_mk_array(x_67, x_68); -x_70 = lean_unsigned_to_nat(1u); -x_71 = lean_nat_sub(x_67, x_70); -lean_dec(x_67); -x_72 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_69, x_71); -x_73 = lean_box(0); -x_74 = lean_array_size(x_72); -x_75 = 0; -x_76 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__4(x_72, x_73, x_72, x_74, x_75, x_65); -lean_dec(x_72); -return x_76; -} -case 5: -{ -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; size_t x_90; size_t x_91; lean_object* x_92; -x_77 = l_Lean_Expr_getAppFn(x_1); -x_78 = l_Lean_MessageData_ofExpr(x_77); -x_79 = l_Lean_Meta_Grind_ppPattern___closed__4; -x_80 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_80, 0, x_79); -lean_ctor_set(x_80, 1, x_78); -x_81 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_81, 0, x_80); -lean_ctor_set(x_81, 1, x_79); -x_82 = lean_unsigned_to_nat(0u); -x_83 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_82); -x_84 = l_Lean_Meta_Grind_ppPattern___closed__5; -lean_inc(x_83); -x_85 = lean_mk_array(x_83, x_84); -x_86 = lean_unsigned_to_nat(1u); -x_87 = lean_nat_sub(x_83, x_86); -lean_dec(x_83); -x_88 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_85, x_87); -x_89 = lean_box(0); -x_90 = lean_array_size(x_88); -x_91 = 0; -x_92 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__5(x_88, x_89, x_88, x_90, x_91, x_81); -lean_dec(x_88); -return x_92; -} -case 6: -{ -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; size_t x_106; size_t x_107; lean_object* x_108; -x_93 = l_Lean_Expr_getAppFn(x_1); -x_94 = l_Lean_MessageData_ofExpr(x_93); -x_95 = l_Lean_Meta_Grind_ppPattern___closed__4; -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_alloc_ctor(7, 2, 0); -lean_ctor_set(x_97, 0, x_96); -lean_ctor_set(x_97, 1, x_95); -x_98 = lean_unsigned_to_nat(0u); -x_99 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_98); -x_100 = l_Lean_Meta_Grind_ppPattern___closed__5; -lean_inc(x_99); -x_101 = lean_mk_array(x_99, x_100); -x_102 = lean_unsigned_to_nat(1u); -x_103 = lean_nat_sub(x_99, x_102); -lean_dec(x_99); -x_104 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_101, x_103); -x_105 = lean_box(0); -x_106 = lean_array_size(x_104); -x_107 = 0; -x_108 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__6(x_104, x_105, x_104, x_106, x_107, x_97); -lean_dec(x_104); -return x_108; -} -case 7: -{ -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; size_t x_122; size_t x_123; lean_object* x_124; -x_109 = l_Lean_Expr_getAppFn(x_1); -x_110 = l_Lean_MessageData_ofExpr(x_109); -x_111 = l_Lean_Meta_Grind_ppPattern___closed__4; -x_112 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_112, 0, x_111); -lean_ctor_set(x_112, 1, x_110); -x_113 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_113, 0, x_112); -lean_ctor_set(x_113, 1, x_111); -x_114 = lean_unsigned_to_nat(0u); -x_115 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_114); -x_116 = l_Lean_Meta_Grind_ppPattern___closed__5; -lean_inc(x_115); -x_117 = lean_mk_array(x_115, x_116); -x_118 = lean_unsigned_to_nat(1u); -x_119 = lean_nat_sub(x_115, x_118); -lean_dec(x_115); -x_120 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_117, x_119); -x_121 = lean_box(0); -x_122 = lean_array_size(x_120); -x_123 = 0; -x_124 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__7(x_120, x_121, x_120, x_122, x_123, x_113); -lean_dec(x_120); -return x_124; -} -case 8: -{ -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; size_t x_138; size_t x_139; lean_object* x_140; -x_125 = l_Lean_Expr_getAppFn(x_1); -x_126 = l_Lean_MessageData_ofExpr(x_125); -x_127 = l_Lean_Meta_Grind_ppPattern___closed__4; -x_128 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_128, 0, x_127); -lean_ctor_set(x_128, 1, x_126); -x_129 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_129, 0, x_128); -lean_ctor_set(x_129, 1, x_127); -x_130 = lean_unsigned_to_nat(0u); -x_131 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_130); -x_132 = l_Lean_Meta_Grind_ppPattern___closed__5; -lean_inc(x_131); -x_133 = lean_mk_array(x_131, x_132); -x_134 = lean_unsigned_to_nat(1u); -x_135 = lean_nat_sub(x_131, x_134); -lean_dec(x_131); -x_136 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_133, x_135); -x_137 = lean_box(0); -x_138 = lean_array_size(x_136); -x_139 = 0; -x_140 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__8(x_136, x_137, x_136, x_138, x_139, x_129); -lean_dec(x_136); -return x_140; -} -case 9: -{ -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; size_t x_154; size_t x_155; lean_object* x_156; -x_141 = l_Lean_Expr_getAppFn(x_1); -x_142 = l_Lean_MessageData_ofExpr(x_141); -x_143 = l_Lean_Meta_Grind_ppPattern___closed__4; -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 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_145, 0, x_144); -lean_ctor_set(x_145, 1, x_143); -x_146 = lean_unsigned_to_nat(0u); -x_147 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_146); -x_148 = l_Lean_Meta_Grind_ppPattern___closed__5; -lean_inc(x_147); -x_149 = lean_mk_array(x_147, x_148); -x_150 = lean_unsigned_to_nat(1u); -x_151 = lean_nat_sub(x_147, x_150); -lean_dec(x_147); -x_152 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_149, x_151); -x_153 = lean_box(0); -x_154 = lean_array_size(x_152); -x_155 = 0; -x_156 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__9(x_152, x_153, x_152, x_154, x_155, x_145); -lean_dec(x_152); -return x_156; -} -case 10: -{ -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; size_t x_170; size_t x_171; lean_object* x_172; -x_157 = l_Lean_Expr_getAppFn(x_1); -x_158 = l_Lean_MessageData_ofExpr(x_157); -x_159 = l_Lean_Meta_Grind_ppPattern___closed__4; -x_160 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_160, 0, x_159); -lean_ctor_set(x_160, 1, x_158); -x_161 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_161, 0, x_160); -lean_ctor_set(x_161, 1, x_159); -x_162 = lean_unsigned_to_nat(0u); -x_163 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_162); -x_164 = l_Lean_Meta_Grind_ppPattern___closed__5; -lean_inc(x_163); -x_165 = lean_mk_array(x_163, x_164); -x_166 = lean_unsigned_to_nat(1u); -x_167 = lean_nat_sub(x_163, x_166); -lean_dec(x_163); -x_168 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_165, x_167); -x_169 = lean_box(0); -x_170 = lean_array_size(x_168); -x_171 = 0; -x_172 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__10(x_168, x_169, x_168, x_170, x_171, x_161); -lean_dec(x_168); -return x_172; -} -default: -{ -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; size_t x_186; size_t x_187; lean_object* x_188; -x_173 = l_Lean_Expr_getAppFn(x_1); -x_174 = l_Lean_MessageData_ofExpr(x_173); -x_175 = l_Lean_Meta_Grind_ppPattern___closed__4; -x_176 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_176, 0, x_175); -lean_ctor_set(x_176, 1, x_174); -x_177 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_177, 0, x_176); -lean_ctor_set(x_177, 1, x_175); -x_178 = lean_unsigned_to_nat(0u); -x_179 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_178); -x_180 = l_Lean_Meta_Grind_ppPattern___closed__5; -lean_inc(x_179); -x_181 = lean_mk_array(x_179, x_180); -x_182 = lean_unsigned_to_nat(1u); -x_183 = lean_nat_sub(x_179, x_182); -lean_dec(x_179); -x_184 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_181, x_183); -x_185 = lean_box(0); -x_186 = lean_array_size(x_184); -x_187 = 0; -x_188 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__11(x_184, x_185, x_184, x_186, x_187, x_177); -lean_dec(x_184); -return x_188; -} -} -} -else -{ -lean_object* x_189; -lean_dec(x_1); -x_189 = l_Lean_Meta_Grind_ppPattern___closed__7; -return x_189; -} -} -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_dec(x_1); -x_190 = lean_ctor_get(x_2, 0); -lean_inc(x_190); -lean_dec(x_2); -x_191 = l_Lean_MessageData_ofExpr(x_190); -x_192 = l_Lean_Meta_Grind_ppPattern___closed__9; -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_Meta_Grind_ppPattern___closed__11; -x_195 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_195, 0, x_193); -lean_ctor_set(x_195, 1, x_194); -return x_195; -} -} -} -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -lean_object* x_4; -x_4 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1(x_1, x_2, x_3); -lean_dec(x_3); -return x_4; -} -} -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___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_7; size_t x_8; lean_object* x_9; -x_7 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_8 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_9 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1(x_1, x_2, x_3, x_7, x_8, x_6); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_9; -} -} -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -size_t x_7; size_t x_8; lean_object* x_9; -x_7 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_8 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_9 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__2(x_1, x_2, x_3, x_7, x_8, x_6); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_9; -} -} -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -size_t x_7; size_t x_8; lean_object* x_9; -x_7 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_8 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_9 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__3(x_1, x_2, x_3, x_7, x_8, x_6); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_9; -} -} -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -size_t x_7; size_t x_8; lean_object* x_9; -x_7 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_8 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_9 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__4(x_1, x_2, x_3, x_7, x_8, x_6); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_9; -} -} -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -size_t x_7; size_t x_8; lean_object* x_9; -x_7 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_8 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_9 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__5(x_1, x_2, x_3, x_7, x_8, x_6); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_9; -} -} -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___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; size_t x_8; lean_object* x_9; -x_7 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_8 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_9 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__6(x_1, x_2, x_3, x_7, x_8, x_6); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_9; -} -} -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -size_t x_7; size_t x_8; lean_object* x_9; -x_7 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_8 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_9 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__7(x_1, x_2, x_3, x_7, x_8, x_6); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_9; -} -} -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -size_t x_7; size_t x_8; lean_object* x_9; -x_7 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_8 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_9 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__8(x_1, x_2, x_3, x_7, x_8, x_6); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_9; -} -} -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___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) { -_start: -{ -size_t x_7; size_t x_8; lean_object* x_9; -x_7 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_8 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_9 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__9(x_1, x_2, x_3, x_7, x_8, x_6); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_9; -} -} -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___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: -{ -size_t x_7; size_t x_8; lean_object* x_9; -x_7 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_8 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_9 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__10(x_1, x_2, x_3, x_7, x_8, x_6); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_9; -} -} -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___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: -{ -size_t x_7; size_t x_8; lean_object* x_9; -x_7 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_8 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_9 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__11(x_1, x_2, x_3, x_7, x_8, x_6); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_9; -} -} -LEAN_EXPORT uint8_t l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__1(lean_object* x_1, lean_object* x_2) { -_start: -{ -if (lean_obj_tag(x_2) == 0) -{ -uint8_t x_3; -x_3 = 0; -return x_3; -} -else -{ -lean_object* x_4; lean_object* x_5; uint8_t x_6; -x_4 = lean_ctor_get(x_2, 0); -x_5 = lean_ctor_get(x_2, 2); -x_6 = l___private_Lean_HeadIndex_0__Lean_beqHeadIndex____x40_Lean_HeadIndex___hyg_69_(x_4, x_1); -if (x_6 == 0) -{ -x_2 = x_5; -goto _start; -} -else -{ -uint8_t x_8; -x_8 = 1; -return x_8; -} -} -} -} -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__4(lean_object* x_1, lean_object* x_2) { -_start: -{ -if (lean_obj_tag(x_2) == 0) -{ -return x_1; -} -else -{ -uint8_t x_3; -x_3 = !lean_is_exclusive(x_2); -if (x_3 == 0) -{ -lean_object* x_4; lean_object* x_5; lean_object* x_6; uint64_t x_7; uint64_t x_8; uint64_t x_9; uint64_t x_10; uint64_t x_11; uint64_t x_12; uint64_t x_13; size_t x_14; size_t x_15; size_t x_16; size_t x_17; size_t x_18; lean_object* x_19; lean_object* x_20; -x_4 = lean_ctor_get(x_2, 0); -x_5 = lean_ctor_get(x_2, 2); -x_6 = lean_array_get_size(x_1); -x_7 = l_Lean_HeadIndex_hash(x_4); -x_8 = 32; -x_9 = lean_uint64_shift_right(x_7, x_8); -x_10 = lean_uint64_xor(x_7, x_9); -x_11 = 16; -x_12 = lean_uint64_shift_right(x_10, x_11); -x_13 = lean_uint64_xor(x_10, x_12); -x_14 = lean_uint64_to_usize(x_13); -x_15 = lean_usize_of_nat(x_6); -lean_dec(x_6); -x_16 = 1; -x_17 = lean_usize_sub(x_15, x_16); -x_18 = lean_usize_land(x_14, x_17); -x_19 = lean_array_uget(x_1, x_18); -lean_ctor_set(x_2, 2, x_19); -x_20 = lean_array_uset(x_1, x_18, x_2); -x_1 = x_20; -x_2 = x_5; -goto _start; -} -else -{ -lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint64_t x_26; uint64_t x_27; uint64_t x_28; uint64_t x_29; uint64_t x_30; uint64_t x_31; uint64_t x_32; size_t x_33; size_t x_34; size_t x_35; size_t x_36; size_t x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_22 = lean_ctor_get(x_2, 0); -x_23 = lean_ctor_get(x_2, 1); -x_24 = lean_ctor_get(x_2, 2); -lean_inc(x_24); -lean_inc(x_23); -lean_inc(x_22); -lean_dec(x_2); -x_25 = lean_array_get_size(x_1); -x_26 = l_Lean_HeadIndex_hash(x_22); -x_27 = 32; -x_28 = lean_uint64_shift_right(x_26, x_27); -x_29 = lean_uint64_xor(x_26, x_28); -x_30 = 16; -x_31 = lean_uint64_shift_right(x_29, x_30); -x_32 = lean_uint64_xor(x_29, x_31); -x_33 = lean_uint64_to_usize(x_32); -x_34 = lean_usize_of_nat(x_25); -lean_dec(x_25); -x_35 = 1; -x_36 = lean_usize_sub(x_34, x_35); -x_37 = lean_usize_land(x_33, x_36); -x_38 = lean_array_uget(x_1, x_37); -x_39 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_39, 0, x_22); -lean_ctor_set(x_39, 1, x_23); -lean_ctor_set(x_39, 2, x_38); -x_40 = lean_array_uset(x_1, x_37, x_39); -x_1 = x_40; -x_2 = x_24; -goto _start; -} -} -} -} -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand_go___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -lean_object* x_4; uint8_t x_5; -x_4 = lean_array_get_size(x_2); -x_5 = lean_nat_dec_lt(x_1, x_4); -lean_dec(x_4); -if (x_5 == 0) -{ -lean_dec(x_2); -lean_dec(x_1); -return x_3; -} -else -{ -lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; -x_6 = lean_array_fget(x_2, x_1); -x_7 = lean_box(0); -x_8 = lean_array_fset(x_2, x_1, x_7); -x_9 = l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__4(x_3, x_6); -x_10 = lean_unsigned_to_nat(1u); -x_11 = lean_nat_add(x_1, x_10); -lean_dec(x_1); -x_1 = x_11; -x_2 = x_8; -x_3 = x_9; -goto _start; -} -} -} -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__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; -x_2 = lean_array_get_size(x_1); -x_3 = lean_unsigned_to_nat(2u); -x_4 = lean_nat_mul(x_2, x_3); -lean_dec(x_2); -x_5 = lean_box(0); -x_6 = lean_mk_array(x_4, x_5); -x_7 = lean_unsigned_to_nat(0u); -x_8 = l_Std_DHashMap_Internal_Raw_u2080_expand_go___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__3(x_7, x_1, x_6); -return x_8; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_saveSymbol(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: -{ -lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; -x_8 = lean_st_ref_get(x_2, x_7); -x_9 = lean_ctor_get(x_8, 0); -lean_inc(x_9); -x_10 = lean_ctor_get(x_9, 1); -lean_inc(x_10); -lean_dec(x_9); -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; uint64_t x_16; uint64_t x_17; uint64_t x_18; uint64_t x_19; uint64_t x_20; uint64_t x_21; uint64_t x_22; size_t x_23; size_t x_24; size_t x_25; size_t x_26; size_t x_27; lean_object* x_28; uint8_t x_29; -x_12 = lean_ctor_get(x_8, 1); -x_13 = lean_ctor_get(x_8, 0); -lean_dec(x_13); -x_14 = lean_ctor_get(x_10, 1); -lean_inc(x_14); -lean_dec(x_10); -x_15 = lean_array_get_size(x_14); -x_16 = l_Lean_HeadIndex_hash(x_1); -x_17 = 32; -x_18 = lean_uint64_shift_right(x_16, x_17); -x_19 = lean_uint64_xor(x_16, x_18); -x_20 = 16; -x_21 = lean_uint64_shift_right(x_19, x_20); -x_22 = lean_uint64_xor(x_19, x_21); -x_23 = lean_uint64_to_usize(x_22); -x_24 = lean_usize_of_nat(x_15); -lean_dec(x_15); -x_25 = 1; -x_26 = lean_usize_sub(x_24, x_25); -x_27 = lean_usize_land(x_23, x_26); -x_28 = lean_array_uget(x_14, x_27); -lean_dec(x_14); -x_29 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__1(x_1, x_28); -lean_dec(x_28); -if (x_29 == 0) -{ -lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; -lean_free_object(x_8); -x_30 = lean_st_ref_take(x_2, x_12); -x_31 = lean_ctor_get(x_30, 0); -lean_inc(x_31); -x_32 = lean_ctor_get(x_30, 1); -lean_inc(x_32); -lean_dec(x_30); -x_33 = !lean_is_exclusive(x_31); -if (x_33 == 0) -{ -lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; -x_34 = lean_ctor_get(x_31, 0); -x_35 = lean_ctor_get(x_31, 1); -lean_inc(x_1); -x_36 = lean_array_push(x_34, x_1); -x_37 = !lean_is_exclusive(x_35); -if (x_37 == 0) -{ -lean_object* x_38; lean_object* x_39; lean_object* x_40; size_t x_41; size_t x_42; size_t x_43; lean_object* x_44; uint8_t x_45; -x_38 = lean_ctor_get(x_35, 0); -x_39 = lean_ctor_get(x_35, 1); -x_40 = lean_array_get_size(x_39); -x_41 = lean_usize_of_nat(x_40); -lean_dec(x_40); -x_42 = lean_usize_sub(x_41, x_25); -x_43 = lean_usize_land(x_23, x_42); -x_44 = lean_array_uget(x_39, x_43); -x_45 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__1(x_1, x_44); -if (x_45 == 0) -{ -lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; uint8_t x_56; -x_46 = lean_unsigned_to_nat(1u); -x_47 = lean_nat_add(x_38, x_46); -lean_dec(x_38); -x_48 = lean_box(0); -x_49 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_49, 0, x_1); -lean_ctor_set(x_49, 1, x_48); -lean_ctor_set(x_49, 2, x_44); -x_50 = lean_array_uset(x_39, x_43, x_49); -x_51 = lean_unsigned_to_nat(4u); -x_52 = lean_nat_mul(x_47, x_51); -x_53 = lean_unsigned_to_nat(3u); -x_54 = lean_nat_div(x_52, x_53); -lean_dec(x_52); -x_55 = lean_array_get_size(x_50); -x_56 = lean_nat_dec_le(x_54, x_55); -lean_dec(x_55); -lean_dec(x_54); -if (x_56 == 0) -{ -lean_object* x_57; lean_object* x_58; uint8_t x_59; -x_57 = l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__2(x_50); -lean_ctor_set(x_35, 1, x_57); -lean_ctor_set(x_35, 0, x_47); -lean_ctor_set(x_31, 0, x_36); -x_58 = lean_st_ref_set(x_2, x_31, x_32); -x_59 = !lean_is_exclusive(x_58); -if (x_59 == 0) -{ -lean_object* x_60; -x_60 = lean_ctor_get(x_58, 0); -lean_dec(x_60); -lean_ctor_set(x_58, 0, x_48); -return x_58; -} -else -{ -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(0, 2, 0); -lean_ctor_set(x_62, 0, x_48); -lean_ctor_set(x_62, 1, x_61); -return x_62; -} -} -else -{ -lean_object* x_63; uint8_t x_64; -lean_ctor_set(x_35, 1, x_50); -lean_ctor_set(x_35, 0, x_47); -lean_ctor_set(x_31, 0, x_36); -x_63 = lean_st_ref_set(x_2, x_31, x_32); -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(x_63, 0, x_48); -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(0, 2, 0); -lean_ctor_set(x_67, 0, x_48); -lean_ctor_set(x_67, 1, x_66); -return x_67; -} -} -} -else -{ -lean_object* x_68; uint8_t x_69; -lean_dec(x_44); -lean_dec(x_1); -lean_ctor_set(x_31, 0, x_36); -x_68 = lean_st_ref_set(x_2, x_31, x_32); -x_69 = !lean_is_exclusive(x_68); -if (x_69 == 0) -{ -lean_object* x_70; lean_object* x_71; -x_70 = lean_ctor_get(x_68, 0); -lean_dec(x_70); -x_71 = lean_box(0); -lean_ctor_set(x_68, 0, x_71); -return x_68; -} -else -{ -lean_object* x_72; lean_object* x_73; lean_object* x_74; -x_72 = lean_ctor_get(x_68, 1); -lean_inc(x_72); -lean_dec(x_68); -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_72); -return x_74; -} -} -} -else -{ -lean_object* x_75; lean_object* x_76; lean_object* x_77; size_t x_78; size_t x_79; size_t x_80; lean_object* x_81; uint8_t x_82; -x_75 = lean_ctor_get(x_35, 0); -x_76 = lean_ctor_get(x_35, 1); -lean_inc(x_76); -lean_inc(x_75); -lean_dec(x_35); -x_77 = lean_array_get_size(x_76); -x_78 = lean_usize_of_nat(x_77); -lean_dec(x_77); -x_79 = lean_usize_sub(x_78, x_25); -x_80 = lean_usize_land(x_23, x_79); -x_81 = lean_array_uget(x_76, x_80); -x_82 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__1(x_1, x_81); -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; uint8_t x_93; -x_83 = lean_unsigned_to_nat(1u); -x_84 = lean_nat_add(x_75, x_83); -lean_dec(x_75); -x_85 = lean_box(0); -x_86 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_86, 0, x_1); -lean_ctor_set(x_86, 1, x_85); -lean_ctor_set(x_86, 2, x_81); -x_87 = lean_array_uset(x_76, x_80, x_86); -x_88 = lean_unsigned_to_nat(4u); -x_89 = lean_nat_mul(x_84, x_88); -x_90 = lean_unsigned_to_nat(3u); -x_91 = lean_nat_div(x_89, x_90); -lean_dec(x_89); -x_92 = lean_array_get_size(x_87); -x_93 = lean_nat_dec_le(x_91, x_92); -lean_dec(x_92); -lean_dec(x_91); -if (x_93 == 0) -{ -lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; -x_94 = l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__2(x_87); -x_95 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_95, 0, x_84); -lean_ctor_set(x_95, 1, x_94); -lean_ctor_set(x_31, 1, x_95); -lean_ctor_set(x_31, 0, x_36); -x_96 = lean_st_ref_set(x_2, x_31, x_32); -x_97 = lean_ctor_get(x_96, 1); -lean_inc(x_97); -if (lean_is_exclusive(x_96)) { - lean_ctor_release(x_96, 0); - lean_ctor_release(x_96, 1); - x_98 = x_96; -} else { - lean_dec_ref(x_96); - x_98 = lean_box(0); -} -if (lean_is_scalar(x_98)) { - x_99 = lean_alloc_ctor(0, 2, 0); -} else { - x_99 = x_98; -} -lean_ctor_set(x_99, 0, x_85); -lean_ctor_set(x_99, 1, x_97); -return x_99; -} -else -{ -lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; -x_100 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_100, 0, x_84); -lean_ctor_set(x_100, 1, x_87); -lean_ctor_set(x_31, 1, x_100); -lean_ctor_set(x_31, 0, x_36); -x_101 = lean_st_ref_set(x_2, x_31, x_32); -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; -} else { - lean_dec_ref(x_101); - x_103 = lean_box(0); -} -if (lean_is_scalar(x_103)) { - x_104 = lean_alloc_ctor(0, 2, 0); -} else { - x_104 = x_103; -} -lean_ctor_set(x_104, 0, x_85); -lean_ctor_set(x_104, 1, x_102); -return x_104; -} -} -else -{ -lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; -lean_dec(x_81); -lean_dec(x_1); -x_105 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_105, 0, x_75); -lean_ctor_set(x_105, 1, x_76); -lean_ctor_set(x_31, 1, x_105); -lean_ctor_set(x_31, 0, x_36); -x_106 = lean_st_ref_set(x_2, x_31, x_32); -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); -} -x_109 = lean_box(0); -if (lean_is_scalar(x_108)) { - x_110 = lean_alloc_ctor(0, 2, 0); -} else { - x_110 = x_108; -} -lean_ctor_set(x_110, 0, x_109); -lean_ctor_set(x_110, 1, x_107); -return x_110; -} -} -} -else -{ -lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; size_t x_119; size_t x_120; size_t x_121; lean_object* x_122; uint8_t x_123; -x_111 = lean_ctor_get(x_31, 0); -x_112 = lean_ctor_get(x_31, 1); -x_113 = lean_ctor_get(x_31, 2); -lean_inc(x_113); -lean_inc(x_112); -lean_inc(x_111); -lean_dec(x_31); -lean_inc(x_1); -x_114 = lean_array_push(x_111, x_1); -x_115 = lean_ctor_get(x_112, 0); -lean_inc(x_115); -x_116 = lean_ctor_get(x_112, 1); -lean_inc(x_116); -if (lean_is_exclusive(x_112)) { - lean_ctor_release(x_112, 0); - lean_ctor_release(x_112, 1); - x_117 = x_112; -} else { - lean_dec_ref(x_112); - x_117 = lean_box(0); -} -x_118 = lean_array_get_size(x_116); -x_119 = lean_usize_of_nat(x_118); -lean_dec(x_118); -x_120 = lean_usize_sub(x_119, x_25); -x_121 = lean_usize_land(x_23, x_120); -x_122 = lean_array_uget(x_116, x_121); -x_123 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__1(x_1, x_122); -if (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; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; uint8_t x_134; -x_124 = lean_unsigned_to_nat(1u); -x_125 = lean_nat_add(x_115, x_124); -lean_dec(x_115); -x_126 = lean_box(0); -x_127 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_127, 0, x_1); -lean_ctor_set(x_127, 1, x_126); -lean_ctor_set(x_127, 2, x_122); -x_128 = lean_array_uset(x_116, x_121, x_127); -x_129 = lean_unsigned_to_nat(4u); -x_130 = lean_nat_mul(x_125, x_129); -x_131 = lean_unsigned_to_nat(3u); -x_132 = lean_nat_div(x_130, x_131); -lean_dec(x_130); -x_133 = lean_array_get_size(x_128); -x_134 = lean_nat_dec_le(x_132, x_133); -lean_dec(x_133); -lean_dec(x_132); -if (x_134 == 0) -{ -lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; -x_135 = l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__2(x_128); -if (lean_is_scalar(x_117)) { - x_136 = lean_alloc_ctor(0, 2, 0); -} else { - x_136 = x_117; -} -lean_ctor_set(x_136, 0, x_125); -lean_ctor_set(x_136, 1, x_135); -x_137 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_137, 0, x_114); -lean_ctor_set(x_137, 1, x_136); -lean_ctor_set(x_137, 2, x_113); -x_138 = lean_st_ref_set(x_2, x_137, x_32); -x_139 = lean_ctor_get(x_138, 1); -lean_inc(x_139); -if (lean_is_exclusive(x_138)) { - lean_ctor_release(x_138, 0); - lean_ctor_release(x_138, 1); - x_140 = x_138; -} else { - lean_dec_ref(x_138); - x_140 = lean_box(0); -} -if (lean_is_scalar(x_140)) { - x_141 = lean_alloc_ctor(0, 2, 0); -} else { - x_141 = x_140; -} -lean_ctor_set(x_141, 0, x_126); -lean_ctor_set(x_141, 1, x_139); -return x_141; -} -else -{ -lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; -if (lean_is_scalar(x_117)) { - x_142 = lean_alloc_ctor(0, 2, 0); -} else { - x_142 = x_117; -} -lean_ctor_set(x_142, 0, x_125); -lean_ctor_set(x_142, 1, x_128); -x_143 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_143, 0, x_114); -lean_ctor_set(x_143, 1, x_142); -lean_ctor_set(x_143, 2, x_113); -x_144 = lean_st_ref_set(x_2, x_143, x_32); -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_126); -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_object* x_152; lean_object* x_153; lean_object* x_154; -lean_dec(x_122); -lean_dec(x_1); -if (lean_is_scalar(x_117)) { - x_148 = lean_alloc_ctor(0, 2, 0); -} else { - x_148 = x_117; -} -lean_ctor_set(x_148, 0, x_115); -lean_ctor_set(x_148, 1, x_116); -x_149 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_149, 0, x_114); -lean_ctor_set(x_149, 1, x_148); -lean_ctor_set(x_149, 2, x_113); -x_150 = lean_st_ref_set(x_2, x_149, x_32); -x_151 = lean_ctor_get(x_150, 1); -lean_inc(x_151); -if (lean_is_exclusive(x_150)) { - lean_ctor_release(x_150, 0); - lean_ctor_release(x_150, 1); - x_152 = x_150; -} else { - lean_dec_ref(x_150); - x_152 = lean_box(0); -} -x_153 = lean_box(0); -if (lean_is_scalar(x_152)) { - x_154 = lean_alloc_ctor(0, 2, 0); -} else { - x_154 = x_152; -} -lean_ctor_set(x_154, 0, x_153); -lean_ctor_set(x_154, 1, x_151); -return x_154; -} -} -} -else -{ -lean_object* x_155; -lean_dec(x_1); -x_155 = lean_box(0); -lean_ctor_set(x_8, 0, x_155); -return x_8; -} -} -else -{ -lean_object* x_156; lean_object* x_157; lean_object* x_158; uint64_t x_159; uint64_t x_160; uint64_t x_161; uint64_t x_162; uint64_t x_163; uint64_t x_164; uint64_t x_165; size_t x_166; size_t x_167; size_t x_168; size_t x_169; size_t x_170; lean_object* x_171; uint8_t x_172; -x_156 = lean_ctor_get(x_8, 1); -lean_inc(x_156); -lean_dec(x_8); -x_157 = lean_ctor_get(x_10, 1); -lean_inc(x_157); -lean_dec(x_10); -x_158 = lean_array_get_size(x_157); -x_159 = l_Lean_HeadIndex_hash(x_1); -x_160 = 32; -x_161 = lean_uint64_shift_right(x_159, x_160); -x_162 = lean_uint64_xor(x_159, x_161); -x_163 = 16; -x_164 = lean_uint64_shift_right(x_162, x_163); -x_165 = lean_uint64_xor(x_162, x_164); -x_166 = lean_uint64_to_usize(x_165); -x_167 = lean_usize_of_nat(x_158); -lean_dec(x_158); -x_168 = 1; -x_169 = lean_usize_sub(x_167, x_168); -x_170 = lean_usize_land(x_166, x_169); -x_171 = lean_array_uget(x_157, x_170); -lean_dec(x_157); -x_172 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__1(x_1, x_171); -lean_dec(x_171); -if (x_172 == 0) -{ -lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; size_t x_185; size_t x_186; size_t x_187; lean_object* x_188; uint8_t x_189; -x_173 = lean_st_ref_take(x_2, x_156); -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); -if (lean_is_exclusive(x_174)) { - lean_ctor_release(x_174, 0); - lean_ctor_release(x_174, 1); - lean_ctor_release(x_174, 2); - x_179 = x_174; -} else { - lean_dec_ref(x_174); - x_179 = lean_box(0); -} -lean_inc(x_1); -x_180 = lean_array_push(x_176, x_1); -x_181 = lean_ctor_get(x_177, 0); -lean_inc(x_181); -x_182 = lean_ctor_get(x_177, 1); -lean_inc(x_182); -if (lean_is_exclusive(x_177)) { - lean_ctor_release(x_177, 0); - lean_ctor_release(x_177, 1); - x_183 = x_177; -} else { - lean_dec_ref(x_177); - x_183 = lean_box(0); -} -x_184 = lean_array_get_size(x_182); -x_185 = lean_usize_of_nat(x_184); -lean_dec(x_184); -x_186 = lean_usize_sub(x_185, x_168); -x_187 = lean_usize_land(x_166, x_186); -x_188 = lean_array_uget(x_182, x_187); -x_189 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__1(x_1, x_188); -if (x_189 == 0) -{ -lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; uint8_t x_200; -x_190 = lean_unsigned_to_nat(1u); -x_191 = lean_nat_add(x_181, x_190); -lean_dec(x_181); -x_192 = lean_box(0); -x_193 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_193, 0, x_1); -lean_ctor_set(x_193, 1, x_192); -lean_ctor_set(x_193, 2, x_188); -x_194 = lean_array_uset(x_182, x_187, x_193); -x_195 = lean_unsigned_to_nat(4u); -x_196 = lean_nat_mul(x_191, x_195); -x_197 = lean_unsigned_to_nat(3u); -x_198 = lean_nat_div(x_196, x_197); -lean_dec(x_196); -x_199 = lean_array_get_size(x_194); -x_200 = lean_nat_dec_le(x_198, x_199); -lean_dec(x_199); -lean_dec(x_198); -if (x_200 == 0) -{ -lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; -x_201 = l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__2(x_194); -if (lean_is_scalar(x_183)) { - x_202 = lean_alloc_ctor(0, 2, 0); -} else { - x_202 = x_183; -} -lean_ctor_set(x_202, 0, x_191); -lean_ctor_set(x_202, 1, x_201); -if (lean_is_scalar(x_179)) { - x_203 = lean_alloc_ctor(0, 3, 0); -} else { - x_203 = x_179; -} -lean_ctor_set(x_203, 0, x_180); -lean_ctor_set(x_203, 1, x_202); -lean_ctor_set(x_203, 2, x_178); -x_204 = lean_st_ref_set(x_2, x_203, x_175); -x_205 = lean_ctor_get(x_204, 1); -lean_inc(x_205); -if (lean_is_exclusive(x_204)) { - lean_ctor_release(x_204, 0); - lean_ctor_release(x_204, 1); - x_206 = x_204; -} else { - lean_dec_ref(x_204); - x_206 = lean_box(0); -} -if (lean_is_scalar(x_206)) { - x_207 = lean_alloc_ctor(0, 2, 0); -} else { - x_207 = x_206; -} -lean_ctor_set(x_207, 0, x_192); -lean_ctor_set(x_207, 1, x_205); -return x_207; -} -else -{ -lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; -if (lean_is_scalar(x_183)) { - x_208 = lean_alloc_ctor(0, 2, 0); -} else { - x_208 = x_183; -} -lean_ctor_set(x_208, 0, x_191); -lean_ctor_set(x_208, 1, x_194); -if (lean_is_scalar(x_179)) { - x_209 = lean_alloc_ctor(0, 3, 0); -} else { - x_209 = x_179; -} -lean_ctor_set(x_209, 0, x_180); -lean_ctor_set(x_209, 1, x_208); -lean_ctor_set(x_209, 2, x_178); -x_210 = lean_st_ref_set(x_2, x_209, x_175); -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_210); - x_212 = lean_box(0); -} -if (lean_is_scalar(x_212)) { - x_213 = lean_alloc_ctor(0, 2, 0); -} else { - x_213 = x_212; -} -lean_ctor_set(x_213, 0, x_192); -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_dec(x_188); -lean_dec(x_1); -if (lean_is_scalar(x_183)) { - x_214 = lean_alloc_ctor(0, 2, 0); -} else { - x_214 = x_183; -} -lean_ctor_set(x_214, 0, x_181); -lean_ctor_set(x_214, 1, x_182); -if (lean_is_scalar(x_179)) { - x_215 = lean_alloc_ctor(0, 3, 0); -} else { - x_215 = x_179; -} -lean_ctor_set(x_215, 0, x_180); -lean_ctor_set(x_215, 1, x_214); -lean_ctor_set(x_215, 2, x_178); -x_216 = lean_st_ref_set(x_2, x_215, x_175); -x_217 = lean_ctor_get(x_216, 1); -lean_inc(x_217); -if (lean_is_exclusive(x_216)) { - lean_ctor_release(x_216, 0); - lean_ctor_release(x_216, 1); - x_218 = x_216; -} else { - lean_dec_ref(x_216); - x_218 = lean_box(0); -} -x_219 = lean_box(0); -if (lean_is_scalar(x_218)) { - x_220 = lean_alloc_ctor(0, 2, 0); -} else { - x_220 = x_218; -} -lean_ctor_set(x_220, 0, x_219); -lean_ctor_set(x_220, 1, x_217); -return x_220; -} -} -else -{ -lean_object* x_221; lean_object* x_222; -lean_dec(x_1); -x_221 = lean_box(0); -x_222 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_222, 0, x_221); -lean_ctor_set(x_222, 1, x_156); -return x_222; -} -} -} -} -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__1___boxed(lean_object* x_1, lean_object* x_2) { -_start: -{ -uint8_t x_3; lean_object* x_4; -x_3 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__1(x_1, x_2); -lean_dec(x_2); -lean_dec(x_1); -x_4 = lean_box(x_3); -return x_4; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: -{ -lean_object* x_8; -x_8 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_saveSymbol(x_1, x_2, x_3, x_4, x_5, x_6, x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_8; -} -} -LEAN_EXPORT uint8_t l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_foundBVar___spec__1(lean_object* x_1, lean_object* x_2) { -_start: -{ -if (lean_obj_tag(x_2) == 0) -{ -uint8_t x_3; -x_3 = 0; -return x_3; -} -else -{ -lean_object* x_4; lean_object* x_5; uint8_t x_6; -x_4 = lean_ctor_get(x_2, 0); -x_5 = lean_ctor_get(x_2, 2); -x_6 = lean_nat_dec_eq(x_4, x_1); -if (x_6 == 0) -{ -x_2 = x_5; -goto _start; -} -else -{ -uint8_t x_8; -x_8 = 1; -return x_8; -} -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_foundBVar(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: -{ -lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; -x_8 = lean_st_ref_get(x_2, x_7); -x_9 = lean_ctor_get(x_8, 0); -lean_inc(x_9); -x_10 = lean_ctor_get(x_9, 2); -lean_inc(x_10); -lean_dec(x_9); -x_11 = !lean_is_exclusive(x_8); -if (x_11 == 0) -{ -lean_object* x_12; lean_object* x_13; lean_object* x_14; uint64_t x_15; uint64_t x_16; uint64_t x_17; uint64_t x_18; uint64_t x_19; uint64_t x_20; uint64_t x_21; size_t x_22; size_t x_23; size_t x_24; size_t x_25; size_t x_26; lean_object* x_27; uint8_t x_28; lean_object* x_29; -x_12 = lean_ctor_get(x_8, 0); -lean_dec(x_12); -x_13 = lean_ctor_get(x_10, 1); -lean_inc(x_13); -lean_dec(x_10); -x_14 = lean_array_get_size(x_13); -x_15 = lean_uint64_of_nat(x_1); -x_16 = 32; -x_17 = lean_uint64_shift_right(x_15, x_16); -x_18 = lean_uint64_xor(x_15, x_17); -x_19 = 16; -x_20 = lean_uint64_shift_right(x_18, x_19); -x_21 = lean_uint64_xor(x_18, x_20); -x_22 = lean_uint64_to_usize(x_21); -x_23 = lean_usize_of_nat(x_14); -lean_dec(x_14); -x_24 = 1; -x_25 = lean_usize_sub(x_23, x_24); -x_26 = lean_usize_land(x_22, x_25); -x_27 = lean_array_uget(x_13, x_26); -lean_dec(x_13); -x_28 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_foundBVar___spec__1(x_1, x_27); -lean_dec(x_27); -x_29 = lean_box(x_28); -lean_ctor_set(x_8, 0, x_29); -return x_8; -} -else -{ -lean_object* x_30; lean_object* x_31; lean_object* x_32; uint64_t x_33; uint64_t x_34; uint64_t x_35; uint64_t x_36; uint64_t x_37; uint64_t x_38; uint64_t x_39; size_t x_40; size_t x_41; size_t x_42; size_t x_43; size_t x_44; lean_object* x_45; uint8_t x_46; lean_object* x_47; lean_object* x_48; -x_30 = lean_ctor_get(x_8, 1); -lean_inc(x_30); -lean_dec(x_8); -x_31 = lean_ctor_get(x_10, 1); -lean_inc(x_31); -lean_dec(x_10); -x_32 = lean_array_get_size(x_31); -x_33 = lean_uint64_of_nat(x_1); -x_34 = 32; -x_35 = lean_uint64_shift_right(x_33, x_34); -x_36 = lean_uint64_xor(x_33, x_35); -x_37 = 16; -x_38 = lean_uint64_shift_right(x_36, x_37); -x_39 = lean_uint64_xor(x_36, x_38); -x_40 = lean_uint64_to_usize(x_39); -x_41 = lean_usize_of_nat(x_32); -lean_dec(x_32); -x_42 = 1; -x_43 = lean_usize_sub(x_41, x_42); -x_44 = lean_usize_land(x_40, x_43); -x_45 = lean_array_uget(x_31, x_44); -lean_dec(x_31); -x_46 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_foundBVar___spec__1(x_1, x_45); -lean_dec(x_45); -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_30); -return x_48; -} -} -} -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_foundBVar___spec__1___boxed(lean_object* x_1, lean_object* x_2) { -_start: -{ -uint8_t x_3; lean_object* x_4; -x_3 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_foundBVar___spec__1(x_1, x_2); -lean_dec(x_2); -lean_dec(x_1); -x_4 = lean_box(x_3); -return x_4; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_foundBVar___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: -{ -lean_object* x_8; -x_8 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_foundBVar(x_1, x_2, x_3, x_4, x_5, x_6, x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_8; -} -} -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -if (lean_obj_tag(x_3) == 0) -{ -lean_dec(x_1); -return x_2; -} -else -{ -uint8_t x_4; -x_4 = !lean_is_exclusive(x_3); -if (x_4 == 0) -{ -lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint64_t x_9; uint64_t x_10; uint64_t x_11; uint64_t x_12; uint64_t x_13; uint64_t x_14; uint64_t x_15; size_t x_16; size_t x_17; size_t x_18; size_t x_19; size_t x_20; lean_object* x_21; lean_object* x_22; -x_5 = lean_ctor_get(x_3, 0); -x_6 = lean_ctor_get(x_3, 2); -x_7 = lean_array_get_size(x_2); -lean_inc(x_1); -lean_inc(x_5); -x_8 = lean_apply_1(x_1, x_5); -x_9 = lean_unbox_uint64(x_8); -lean_dec(x_8); -x_10 = 32; -x_11 = lean_uint64_shift_right(x_9, x_10); -x_12 = lean_uint64_xor(x_9, x_11); -x_13 = 16; -x_14 = lean_uint64_shift_right(x_12, x_13); -x_15 = lean_uint64_xor(x_12, x_14); -x_16 = lean_uint64_to_usize(x_15); -x_17 = lean_usize_of_nat(x_7); -lean_dec(x_7); -x_18 = 1; -x_19 = lean_usize_sub(x_17, x_18); -x_20 = lean_usize_land(x_16, x_19); -x_21 = lean_array_uget(x_2, x_20); -lean_ctor_set(x_3, 2, x_21); -x_22 = lean_array_uset(x_2, x_20, x_3); -x_2 = x_22; -x_3 = x_6; -goto _start; -} -else -{ -lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; uint64_t x_29; uint64_t x_30; uint64_t x_31; uint64_t x_32; uint64_t x_33; uint64_t x_34; uint64_t x_35; size_t x_36; size_t x_37; size_t x_38; size_t x_39; size_t x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; -x_24 = lean_ctor_get(x_3, 0); -x_25 = lean_ctor_get(x_3, 1); -x_26 = lean_ctor_get(x_3, 2); -lean_inc(x_26); -lean_inc(x_25); -lean_inc(x_24); -lean_dec(x_3); -x_27 = lean_array_get_size(x_2); -lean_inc(x_1); -lean_inc(x_24); -x_28 = lean_apply_1(x_1, x_24); -x_29 = lean_unbox_uint64(x_28); -lean_dec(x_28); -x_30 = 32; -x_31 = lean_uint64_shift_right(x_29, x_30); -x_32 = lean_uint64_xor(x_29, x_31); -x_33 = 16; -x_34 = lean_uint64_shift_right(x_32, x_33); -x_35 = lean_uint64_xor(x_32, x_34); -x_36 = lean_uint64_to_usize(x_35); -x_37 = lean_usize_of_nat(x_27); -lean_dec(x_27); -x_38 = 1; -x_39 = lean_usize_sub(x_37, x_38); -x_40 = lean_usize_land(x_36, x_39); -x_41 = lean_array_uget(x_2, x_40); -x_42 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_42, 0, x_24); -lean_ctor_set(x_42, 1, x_25); -lean_ctor_set(x_42, 2, x_41); -x_43 = lean_array_uset(x_2, x_40, x_42); -x_2 = x_43; -x_3 = x_26; -goto _start; -} -} -} -} -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__3___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__4(lean_object* x_1, lean_object* x_2) { -_start: -{ -if (lean_obj_tag(x_2) == 0) -{ -return x_1; -} -else -{ -uint8_t x_3; -x_3 = !lean_is_exclusive(x_2); -if (x_3 == 0) -{ -lean_object* x_4; lean_object* x_5; lean_object* x_6; uint64_t x_7; uint64_t x_8; uint64_t x_9; uint64_t x_10; uint64_t x_11; uint64_t x_12; uint64_t x_13; size_t x_14; size_t x_15; size_t x_16; size_t x_17; size_t x_18; lean_object* x_19; lean_object* x_20; -x_4 = lean_ctor_get(x_2, 0); -x_5 = lean_ctor_get(x_2, 2); -x_6 = lean_array_get_size(x_1); -x_7 = lean_uint64_of_nat(x_4); -x_8 = 32; -x_9 = lean_uint64_shift_right(x_7, x_8); -x_10 = lean_uint64_xor(x_7, x_9); -x_11 = 16; -x_12 = lean_uint64_shift_right(x_10, x_11); -x_13 = lean_uint64_xor(x_10, x_12); -x_14 = lean_uint64_to_usize(x_13); -x_15 = lean_usize_of_nat(x_6); -lean_dec(x_6); -x_16 = 1; -x_17 = lean_usize_sub(x_15, x_16); -x_18 = lean_usize_land(x_14, x_17); -x_19 = lean_array_uget(x_1, x_18); -lean_ctor_set(x_2, 2, x_19); -x_20 = lean_array_uset(x_1, x_18, x_2); -x_1 = x_20; -x_2 = x_5; -goto _start; -} -else -{ -lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint64_t x_26; uint64_t x_27; uint64_t x_28; uint64_t x_29; uint64_t x_30; uint64_t x_31; uint64_t x_32; size_t x_33; size_t x_34; size_t x_35; size_t x_36; size_t x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_22 = lean_ctor_get(x_2, 0); -x_23 = lean_ctor_get(x_2, 1); -x_24 = lean_ctor_get(x_2, 2); -lean_inc(x_24); -lean_inc(x_23); -lean_inc(x_22); -lean_dec(x_2); -x_25 = lean_array_get_size(x_1); -x_26 = lean_uint64_of_nat(x_22); -x_27 = 32; -x_28 = lean_uint64_shift_right(x_26, x_27); -x_29 = lean_uint64_xor(x_26, x_28); -x_30 = 16; -x_31 = lean_uint64_shift_right(x_29, x_30); -x_32 = lean_uint64_xor(x_29, x_31); -x_33 = lean_uint64_to_usize(x_32); -x_34 = lean_usize_of_nat(x_25); -lean_dec(x_25); -x_35 = 1; -x_36 = lean_usize_sub(x_34, x_35); -x_37 = lean_usize_land(x_33, x_36); -x_38 = lean_array_uget(x_1, x_37); -x_39 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_39, 0, x_22); -lean_ctor_set(x_39, 1, x_23); -lean_ctor_set(x_39, 2, x_38); -x_40 = lean_array_uset(x_1, x_37, x_39); -x_1 = x_40; -x_2 = x_24; -goto _start; -} -} -} -} -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand_go___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -lean_object* x_4; uint8_t x_5; -x_4 = lean_array_get_size(x_2); -x_5 = lean_nat_dec_lt(x_1, x_4); -lean_dec(x_4); -if (x_5 == 0) -{ -lean_dec(x_2); -lean_dec(x_1); -return x_3; -} -else -{ -lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; -x_6 = lean_array_fget(x_2, x_1); -x_7 = lean_box(0); -x_8 = lean_array_fset(x_2, x_1, x_7); -x_9 = l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__3___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__4(x_3, x_6); -x_10 = lean_unsigned_to_nat(1u); -x_11 = lean_nat_add(x_1, x_10); -lean_dec(x_1); -x_1 = x_11; -x_2 = x_8; -x_3 = x_9; -goto _start; -} -} -} -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__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; lean_object* x_7; lean_object* x_8; -x_2 = lean_array_get_size(x_1); -x_3 = lean_unsigned_to_nat(2u); -x_4 = lean_nat_mul(x_2, x_3); -lean_dec(x_2); -x_5 = lean_box(0); -x_6 = lean_mk_array(x_4, x_5); -x_7 = lean_unsigned_to_nat(0u); -x_8 = l_Std_DHashMap_Internal_Raw_u2080_expand_go___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__2(x_7, x_1, x_6); -return x_8; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_saveBVar(lean_object* x_1, 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; -x_8 = lean_st_ref_take(x_2, x_7); -x_9 = lean_ctor_get(x_8, 0); -lean_inc(x_9); -x_10 = lean_ctor_get(x_9, 2); -lean_inc(x_10); -x_11 = lean_ctor_get(x_8, 1); -lean_inc(x_11); -lean_dec(x_8); -x_12 = !lean_is_exclusive(x_9); -if (x_12 == 0) -{ -lean_object* x_13; uint8_t x_14; -x_13 = lean_ctor_get(x_9, 2); -lean_dec(x_13); -x_14 = !lean_is_exclusive(x_10); -if (x_14 == 0) -{ -lean_object* x_15; lean_object* x_16; lean_object* x_17; uint64_t x_18; uint64_t x_19; uint64_t x_20; uint64_t x_21; uint64_t x_22; uint64_t x_23; uint64_t x_24; size_t x_25; size_t x_26; size_t x_27; size_t x_28; size_t x_29; lean_object* x_30; uint8_t x_31; -x_15 = lean_ctor_get(x_10, 0); -x_16 = lean_ctor_get(x_10, 1); -x_17 = lean_array_get_size(x_16); -x_18 = lean_uint64_of_nat(x_1); -x_19 = 32; -x_20 = lean_uint64_shift_right(x_18, x_19); -x_21 = lean_uint64_xor(x_18, x_20); -x_22 = 16; -x_23 = lean_uint64_shift_right(x_21, x_22); -x_24 = lean_uint64_xor(x_21, x_23); -x_25 = lean_uint64_to_usize(x_24); -x_26 = lean_usize_of_nat(x_17); -lean_dec(x_17); -x_27 = 1; -x_28 = lean_usize_sub(x_26, x_27); -x_29 = lean_usize_land(x_25, x_28); -x_30 = lean_array_uget(x_16, x_29); -x_31 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_foundBVar___spec__1(x_1, x_30); -if (x_31 == 0) -{ -lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; uint8_t x_42; -x_32 = lean_unsigned_to_nat(1u); -x_33 = lean_nat_add(x_15, x_32); -lean_dec(x_15); -x_34 = lean_box(0); -x_35 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_35, 0, x_1); -lean_ctor_set(x_35, 1, x_34); -lean_ctor_set(x_35, 2, x_30); -x_36 = lean_array_uset(x_16, x_29, x_35); -x_37 = lean_unsigned_to_nat(4u); -x_38 = lean_nat_mul(x_33, x_37); -x_39 = lean_unsigned_to_nat(3u); -x_40 = lean_nat_div(x_38, x_39); -lean_dec(x_38); -x_41 = lean_array_get_size(x_36); -x_42 = lean_nat_dec_le(x_40, x_41); -lean_dec(x_41); -lean_dec(x_40); -if (x_42 == 0) -{ -lean_object* x_43; lean_object* x_44; uint8_t x_45; -x_43 = l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__1(x_36); -lean_ctor_set(x_10, 1, x_43); -lean_ctor_set(x_10, 0, x_33); -x_44 = lean_st_ref_set(x_2, x_9, x_11); -x_45 = !lean_is_exclusive(x_44); -if (x_45 == 0) -{ -lean_object* x_46; -x_46 = lean_ctor_get(x_44, 0); -lean_dec(x_46); -lean_ctor_set(x_44, 0, x_34); -return x_44; -} -else -{ -lean_object* x_47; lean_object* x_48; -x_47 = lean_ctor_get(x_44, 1); -lean_inc(x_47); -lean_dec(x_44); -x_48 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_48, 0, x_34); -lean_ctor_set(x_48, 1, x_47); -return x_48; -} -} -else -{ -lean_object* x_49; uint8_t x_50; -lean_ctor_set(x_10, 1, x_36); -lean_ctor_set(x_10, 0, x_33); -x_49 = lean_st_ref_set(x_2, x_9, x_11); -x_50 = !lean_is_exclusive(x_49); -if (x_50 == 0) -{ -lean_object* x_51; -x_51 = lean_ctor_get(x_49, 0); -lean_dec(x_51); -lean_ctor_set(x_49, 0, x_34); -return x_49; -} -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_34); -lean_ctor_set(x_53, 1, x_52); -return x_53; -} -} -} -else -{ -lean_object* x_54; uint8_t x_55; -lean_dec(x_30); -lean_dec(x_1); -x_54 = lean_st_ref_set(x_2, x_9, x_11); -x_55 = !lean_is_exclusive(x_54); -if (x_55 == 0) -{ -lean_object* x_56; lean_object* x_57; -x_56 = lean_ctor_get(x_54, 0); -lean_dec(x_56); -x_57 = lean_box(0); -lean_ctor_set(x_54, 0, x_57); -return x_54; -} -else -{ -lean_object* x_58; lean_object* x_59; lean_object* x_60; -x_58 = lean_ctor_get(x_54, 1); -lean_inc(x_58); -lean_dec(x_54); -x_59 = lean_box(0); -x_60 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_60, 0, x_59); -lean_ctor_set(x_60, 1, x_58); -return x_60; -} -} -} -else -{ -lean_object* x_61; lean_object* x_62; lean_object* x_63; uint64_t x_64; uint64_t x_65; uint64_t x_66; uint64_t x_67; uint64_t x_68; uint64_t x_69; uint64_t x_70; size_t x_71; size_t x_72; size_t x_73; size_t x_74; size_t x_75; lean_object* x_76; uint8_t x_77; -x_61 = lean_ctor_get(x_10, 0); -x_62 = lean_ctor_get(x_10, 1); -lean_inc(x_62); -lean_inc(x_61); -lean_dec(x_10); -x_63 = lean_array_get_size(x_62); -x_64 = lean_uint64_of_nat(x_1); -x_65 = 32; -x_66 = lean_uint64_shift_right(x_64, x_65); -x_67 = lean_uint64_xor(x_64, x_66); -x_68 = 16; -x_69 = lean_uint64_shift_right(x_67, x_68); -x_70 = lean_uint64_xor(x_67, x_69); -x_71 = lean_uint64_to_usize(x_70); -x_72 = lean_usize_of_nat(x_63); -lean_dec(x_63); -x_73 = 1; -x_74 = lean_usize_sub(x_72, x_73); -x_75 = lean_usize_land(x_71, x_74); -x_76 = lean_array_uget(x_62, x_75); -x_77 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_foundBVar___spec__1(x_1, x_76); -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; uint8_t x_88; -x_78 = lean_unsigned_to_nat(1u); -x_79 = lean_nat_add(x_61, x_78); -lean_dec(x_61); -x_80 = lean_box(0); -x_81 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_81, 0, x_1); -lean_ctor_set(x_81, 1, x_80); -lean_ctor_set(x_81, 2, x_76); -x_82 = lean_array_uset(x_62, x_75, x_81); -x_83 = lean_unsigned_to_nat(4u); -x_84 = lean_nat_mul(x_79, x_83); -x_85 = lean_unsigned_to_nat(3u); -x_86 = lean_nat_div(x_84, x_85); -lean_dec(x_84); -x_87 = lean_array_get_size(x_82); -x_88 = lean_nat_dec_le(x_86, x_87); -lean_dec(x_87); -lean_dec(x_86); -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; -x_89 = l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__1(x_82); -x_90 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_90, 0, x_79); -lean_ctor_set(x_90, 1, x_89); -lean_ctor_set(x_9, 2, x_90); -x_91 = lean_st_ref_set(x_2, x_9, x_11); -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_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; -x_95 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_95, 0, x_79); -lean_ctor_set(x_95, 1, x_82); -lean_ctor_set(x_9, 2, x_95); -x_96 = lean_st_ref_set(x_2, x_9, x_11); -x_97 = lean_ctor_get(x_96, 1); -lean_inc(x_97); -if (lean_is_exclusive(x_96)) { - lean_ctor_release(x_96, 0); - lean_ctor_release(x_96, 1); - x_98 = x_96; -} else { - lean_dec_ref(x_96); - x_98 = lean_box(0); -} -if (lean_is_scalar(x_98)) { - x_99 = lean_alloc_ctor(0, 2, 0); -} else { - x_99 = x_98; -} -lean_ctor_set(x_99, 0, x_80); -lean_ctor_set(x_99, 1, x_97); -return x_99; -} -} -else -{ -lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; -lean_dec(x_76); -lean_dec(x_1); -x_100 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_100, 0, x_61); -lean_ctor_set(x_100, 1, x_62); -lean_ctor_set(x_9, 2, x_100); -x_101 = lean_st_ref_set(x_2, x_9, x_11); -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; -} else { - lean_dec_ref(x_101); - x_103 = lean_box(0); -} -x_104 = lean_box(0); -if (lean_is_scalar(x_103)) { - x_105 = lean_alloc_ctor(0, 2, 0); -} else { - x_105 = x_103; -} -lean_ctor_set(x_105, 0, x_104); -lean_ctor_set(x_105, 1, x_102); -return x_105; -} -} -} -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; uint64_t x_112; uint64_t x_113; uint64_t x_114; uint64_t x_115; uint64_t x_116; uint64_t x_117; uint64_t x_118; size_t x_119; size_t x_120; size_t x_121; size_t x_122; size_t x_123; lean_object* x_124; uint8_t x_125; -x_106 = lean_ctor_get(x_9, 0); -x_107 = lean_ctor_get(x_9, 1); -lean_inc(x_107); -lean_inc(x_106); -lean_dec(x_9); -x_108 = lean_ctor_get(x_10, 0); -lean_inc(x_108); -x_109 = lean_ctor_get(x_10, 1); -lean_inc(x_109); -if (lean_is_exclusive(x_10)) { - lean_ctor_release(x_10, 0); - lean_ctor_release(x_10, 1); - x_110 = x_10; -} else { - lean_dec_ref(x_10); - x_110 = lean_box(0); -} -x_111 = lean_array_get_size(x_109); -x_112 = lean_uint64_of_nat(x_1); -x_113 = 32; -x_114 = lean_uint64_shift_right(x_112, x_113); -x_115 = lean_uint64_xor(x_112, x_114); -x_116 = 16; -x_117 = lean_uint64_shift_right(x_115, x_116); -x_118 = lean_uint64_xor(x_115, x_117); -x_119 = lean_uint64_to_usize(x_118); -x_120 = lean_usize_of_nat(x_111); -lean_dec(x_111); -x_121 = 1; -x_122 = lean_usize_sub(x_120, x_121); -x_123 = lean_usize_land(x_119, x_122); -x_124 = lean_array_uget(x_109, x_123); -x_125 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_foundBVar___spec__1(x_1, x_124); -if (x_125 == 0) -{ -lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; uint8_t x_136; -x_126 = lean_unsigned_to_nat(1u); -x_127 = lean_nat_add(x_108, x_126); -lean_dec(x_108); -x_128 = lean_box(0); -x_129 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_129, 0, x_1); -lean_ctor_set(x_129, 1, x_128); -lean_ctor_set(x_129, 2, x_124); -x_130 = lean_array_uset(x_109, x_123, x_129); -x_131 = lean_unsigned_to_nat(4u); -x_132 = lean_nat_mul(x_127, x_131); -x_133 = lean_unsigned_to_nat(3u); -x_134 = lean_nat_div(x_132, x_133); -lean_dec(x_132); -x_135 = lean_array_get_size(x_130); -x_136 = lean_nat_dec_le(x_134, x_135); -lean_dec(x_135); -lean_dec(x_134); -if (x_136 == 0) -{ -lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; -x_137 = l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__1(x_130); -if (lean_is_scalar(x_110)) { - x_138 = lean_alloc_ctor(0, 2, 0); -} else { - x_138 = x_110; -} -lean_ctor_set(x_138, 0, x_127); -lean_ctor_set(x_138, 1, x_137); -x_139 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_139, 0, x_106); -lean_ctor_set(x_139, 1, x_107); -lean_ctor_set(x_139, 2, x_138); -x_140 = lean_st_ref_set(x_2, x_139, x_11); -x_141 = lean_ctor_get(x_140, 1); -lean_inc(x_141); -if (lean_is_exclusive(x_140)) { - lean_ctor_release(x_140, 0); - lean_ctor_release(x_140, 1); - x_142 = x_140; -} else { - lean_dec_ref(x_140); - x_142 = lean_box(0); -} -if (lean_is_scalar(x_142)) { - x_143 = lean_alloc_ctor(0, 2, 0); -} else { - x_143 = x_142; -} -lean_ctor_set(x_143, 0, x_128); -lean_ctor_set(x_143, 1, x_141); -return x_143; -} -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; -if (lean_is_scalar(x_110)) { - x_144 = lean_alloc_ctor(0, 2, 0); -} else { - x_144 = x_110; -} -lean_ctor_set(x_144, 0, x_127); -lean_ctor_set(x_144, 1, x_130); -x_145 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_145, 0, x_106); -lean_ctor_set(x_145, 1, x_107); -lean_ctor_set(x_145, 2, x_144); -x_146 = lean_st_ref_set(x_2, x_145, x_11); -x_147 = lean_ctor_get(x_146, 1); -lean_inc(x_147); -if (lean_is_exclusive(x_146)) { - lean_ctor_release(x_146, 0); - lean_ctor_release(x_146, 1); - x_148 = x_146; -} else { - lean_dec_ref(x_146); - x_148 = lean_box(0); -} -if (lean_is_scalar(x_148)) { - x_149 = lean_alloc_ctor(0, 2, 0); -} else { - x_149 = x_148; -} -lean_ctor_set(x_149, 0, x_128); -lean_ctor_set(x_149, 1, x_147); -return x_149; -} -} -else -{ -lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; -lean_dec(x_124); -lean_dec(x_1); -if (lean_is_scalar(x_110)) { - x_150 = lean_alloc_ctor(0, 2, 0); -} else { - x_150 = x_110; -} -lean_ctor_set(x_150, 0, x_108); -lean_ctor_set(x_150, 1, x_109); -x_151 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_151, 0, x_106); -lean_ctor_set(x_151, 1, x_107); -lean_ctor_set(x_151, 2, x_150); -x_152 = lean_st_ref_set(x_2, x_151, x_11); -x_153 = lean_ctor_get(x_152, 1); -lean_inc(x_153); -if (lean_is_exclusive(x_152)) { - lean_ctor_release(x_152, 0); - lean_ctor_release(x_152, 1); - x_154 = x_152; -} else { - lean_dec_ref(x_152); - x_154 = lean_box(0); -} -x_155 = lean_box(0); -if (lean_is_scalar(x_154)) { - x_156 = lean_alloc_ctor(0, 2, 0); -} else { - x_156 = x_154; -} -lean_ctor_set(x_156, 0, x_155); -lean_ctor_set(x_156, 1, x_153); -return x_156; -} -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_saveBVar___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: -{ -lean_object* x_8; -x_8 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_saveBVar(x_1, x_2, x_3, x_4, x_5, x_6, x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_8; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_getPatternFn_x3f(lean_object* x_1) { -_start: -{ -uint8_t x_2; -x_2 = l_Lean_Expr_isApp(x_1); -if (x_2 == 0) -{ -lean_object* x_3; -x_3 = lean_box(0); -return x_3; -} -else -{ -lean_object* x_4; -x_4 = l_Lean_Expr_getAppFn(x_1); -switch (lean_obj_tag(x_4)) { -case 1: -{ -lean_object* x_5; -x_5 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_5, 0, x_4); -return x_5; -} -case 4: -{ -lean_object* x_6; lean_object* x_7; uint8_t x_8; -x_6 = lean_ctor_get(x_4, 0); -lean_inc(x_6); -x_7 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames; -x_8 = l_Array_contains___at_Lean_registerInternalExceptionId___spec__1(x_7, x_6); -lean_dec(x_6); -if (x_8 == 0) -{ -lean_object* x_9; -x_9 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_9, 0, x_4); -return x_9; -} -else -{ -lean_object* x_10; -lean_dec(x_4); -x_10 = lean_box(0); -return x_10; -} -} -default: -{ -lean_object* x_11; -lean_dec(x_4); -x_11 = lean_box(0); -return x_11; -} -} -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_getPatternFn_x3f___boxed(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_getPatternFn_x3f(x_1); -lean_dec(x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_getPatternFunInfo___spec__1(size_t x_1, size_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -uint8_t x_9; -x_9 = lean_usize_dec_lt(x_2, x_1); -if (x_9 == 0) -{ -lean_object* x_10; -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -x_10 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_10, 0, x_3); -lean_ctor_set(x_10, 1, x_8); -return x_10; -} -else -{ -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_11 = lean_array_uget(x_3, x_2); -x_12 = lean_unsigned_to_nat(0u); -x_13 = lean_array_uset(x_3, x_2, x_12); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -x_14 = l_Lean_Meta_isTypeFormer(x_11, x_4, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_14) == 0) -{ -lean_object* x_15; lean_object* x_16; size_t x_17; size_t x_18; lean_object* x_19; -x_15 = lean_ctor_get(x_14, 0); -lean_inc(x_15); -x_16 = lean_ctor_get(x_14, 1); -lean_inc(x_16); -lean_dec(x_14); -x_17 = 1; -x_18 = lean_usize_add(x_2, x_17); -x_19 = lean_array_uset(x_13, x_2, x_15); -x_2 = x_18; -x_3 = x_19; -x_8 = x_16; -goto _start; -} -else -{ -uint8_t x_21; -lean_dec(x_13); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -x_21 = !lean_is_exclusive(x_14); -if (x_21 == 0) -{ -return x_14; -} -else -{ -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_14, 0); -x_23 = lean_ctor_get(x_14, 1); -lean_inc(x_23); -lean_inc(x_22); -lean_dec(x_14); -x_24 = lean_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___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_getPatternFunInfo___spec__2(size_t x_1, size_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -uint8_t x_9; -x_9 = lean_usize_dec_lt(x_2, x_1); -if (x_9 == 0) -{ -lean_object* x_10; -lean_dec(x_4); -x_10 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_10, 0, x_3); -lean_ctor_set(x_10, 1, x_8); -return x_10; -} -else -{ -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_11 = lean_array_uget(x_3, x_2); -x_12 = lean_unsigned_to_nat(0u); -x_13 = lean_array_uset(x_3, x_2, x_12); -x_14 = l_Lean_Expr_fvarId_x21(x_11); -lean_dec(x_11); -lean_inc(x_4); -x_15 = l_Lean_FVarId_getDecl(x_14, 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; uint8_t x_18; lean_object* x_19; -x_16 = lean_ctor_get(x_15, 0); -lean_inc(x_16); -x_17 = lean_ctor_get(x_15, 1); -lean_inc(x_17); -lean_dec(x_15); -x_18 = l_Lean_LocalDecl_binderInfo(x_16); -lean_dec(x_16); -x_19 = lean_box(x_18); -if (lean_obj_tag(x_19) == 3) -{ -size_t x_20; size_t x_21; uint8_t x_22; lean_object* x_23; lean_object* x_24; -x_20 = 1; -x_21 = lean_usize_add(x_2, x_20); -x_22 = 1; -x_23 = lean_box(x_22); -x_24 = lean_array_uset(x_13, x_2, x_23); -x_2 = x_21; -x_3 = x_24; -x_8 = x_17; -goto _start; -} -else -{ -size_t x_26; size_t x_27; uint8_t x_28; lean_object* x_29; lean_object* x_30; -lean_dec(x_19); -x_26 = 1; -x_27 = lean_usize_add(x_2, x_26); -x_28 = 0; -x_29 = lean_box(x_28); -x_30 = lean_array_uset(x_13, x_2, x_29); -x_2 = x_27; -x_3 = x_30; -x_8 = x_17; -goto _start; -} -} -else -{ -uint8_t x_32; -lean_dec(x_13); -lean_dec(x_4); -x_32 = !lean_is_exclusive(x_15); -if (x_32 == 0) -{ -return x_15; -} -else -{ -lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_33 = lean_ctor_get(x_15, 0); -x_34 = lean_ctor_get(x_15, 1); -lean_inc(x_34); -lean_inc(x_33); -lean_dec(x_15); -x_35 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -return x_35; -} -} -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_getPatternFunInfo___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: -{ -size_t x_8; size_t x_9; lean_object* x_10; -x_8 = lean_array_size(x_1); -x_9 = 0; -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -lean_inc(x_1); -x_10 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_getPatternFunInfo___spec__1(x_8, x_9, x_1, x_3, x_4, x_5, x_6, x_7); -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_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_getPatternFunInfo___spec__2(x_8, x_9, x_1, x_3, x_4, x_5, x_6, x_12); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -if (lean_obj_tag(x_13) == 0) -{ -uint8_t x_14; -x_14 = !lean_is_exclusive(x_13); -if (x_14 == 0) -{ -lean_object* x_15; lean_object* x_16; -x_15 = lean_ctor_get(x_13, 0); -x_16 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_16, 0, x_15); -lean_ctor_set(x_16, 1, x_11); -lean_ctor_set(x_13, 0, x_16); -return x_13; -} -else -{ -lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_17 = lean_ctor_get(x_13, 0); -x_18 = lean_ctor_get(x_13, 1); -lean_inc(x_18); -lean_inc(x_17); -lean_dec(x_13); -x_19 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_19, 0, x_17); -lean_ctor_set(x_19, 1, x_11); -x_20 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_18); -return x_20; -} -} -else -{ -uint8_t x_21; -lean_dec(x_11); -x_21 = !lean_is_exclusive(x_13); -if (x_21 == 0) -{ -return x_13; -} -else -{ -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_13, 0); -x_23 = lean_ctor_get(x_13, 1); -lean_inc(x_23); -lean_inc(x_22); -lean_dec(x_13); -x_24 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_24, 0, x_22); -lean_ctor_set(x_24, 1, x_23); -return x_24; -} -} -} -else -{ -uint8_t x_25; -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_1); -x_25 = !lean_is_exclusive(x_10); -if (x_25 == 0) -{ -return x_10; -} -else -{ -lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_26 = lean_ctor_get(x_10, 0); -x_27 = lean_ctor_get(x_10, 1); -lean_inc(x_27); -lean_inc(x_26); -lean_dec(x_10); -x_28 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set(x_28, 1, x_27); -return x_28; -} -} -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_getPatternFunInfo___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_getPatternFunInfo___lambda__1___boxed), 7, 0); -return x_1; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_getPatternFunInfo(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: -{ -lean_object* x_8; -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -x_8 = lean_infer_type(x_1, x_3, x_4, x_5, x_6, x_7); -if (lean_obj_tag(x_8) == 0) -{ -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; -x_9 = lean_ctor_get(x_8, 0); -lean_inc(x_9); -x_10 = lean_ctor_get(x_8, 1); -lean_inc(x_10); -lean_dec(x_8); -x_11 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_11, 0, x_2); -x_12 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_getPatternFunInfo___closed__1; -x_13 = 0; -x_14 = l_Lean_Meta_forallBoundedTelescope___at_Lean_Meta_arrowDomainsN___spec__6___rarg(x_9, x_11, x_12, x_13, x_3, x_4, x_5, x_6, x_10); -return x_14; -} -else -{ -uint8_t x_15; -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_8); -if (x_15 == 0) -{ -return x_8; -} -else -{ -lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_16 = lean_ctor_get(x_8, 0); -x_17 = lean_ctor_get(x_8, 1); -lean_inc(x_17); -lean_inc(x_16); -lean_dec(x_8); -x_18 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_18, 0, x_16); -lean_ctor_set(x_18, 1, x_17); -return x_18; -} -} -} -} -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_getPatternFunInfo___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -size_t x_9; size_t x_10; lean_object* x_11; -x_9 = lean_unbox_usize(x_1); -lean_dec(x_1); -x_10 = lean_unbox_usize(x_2); -lean_dec(x_2); -x_11 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_getPatternFunInfo___spec__1(x_9, x_10, x_3, x_4, x_5, x_6, x_7, x_8); -return x_11; -} -} -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_getPatternFunInfo___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -size_t x_9; size_t x_10; lean_object* x_11; -x_9 = lean_unbox_usize(x_1); -lean_dec(x_1); -x_10 = lean_unbox_usize(x_2); -lean_dec(x_2); -x_11 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_getPatternFunInfo___spec__2(x_9, x_10, x_3, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -return x_11; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_getPatternFunInfo___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: -{ -lean_object* x_8; -x_8 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_getPatternFunInfo___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7); -lean_dec(x_2); -return x_8; -} -} -LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: -{ -lean_object* x_8; lean_object* x_9; uint8_t x_10; -x_8 = lean_ctor_get(x_5, 5); -x_9 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_1, x_3, x_4, x_5, x_6, x_7); -x_10 = !lean_is_exclusive(x_9); -if (x_10 == 0) -{ -lean_object* x_11; lean_object* x_12; -x_11 = lean_ctor_get(x_9, 0); -lean_inc(x_8); -x_12 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_12, 0, x_8); -lean_ctor_set(x_12, 1, x_11); -lean_ctor_set_tag(x_9, 1); -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); -lean_inc(x_8); -x_15 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_15, 0, x_8); -lean_ctor_set(x_15, 1, x_13); -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_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___spec__2___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_10 = lean_array_set(x_2, x_1, x_3); -x_11 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_11, 0, 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_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { -_start: -{ -lean_object* x_15; uint8_t x_16; -x_15 = lean_ctor_get(x_4, 1); -x_16 = lean_nat_dec_lt(x_6, x_15); -if (x_16 == 0) -{ -lean_object* x_17; -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_6); -x_17 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_17, 0, x_5); -lean_ctor_set(x_17, 1, x_14); -return x_17; -} -else -{ -lean_object* x_18; lean_object* x_36; uint8_t x_37; lean_object* x_38; uint8_t x_39; lean_object* x_40; uint8_t x_41; lean_object* x_42; -x_36 = lean_array_get_size(x_5); -x_37 = lean_nat_dec_lt(x_6, x_36); -lean_dec(x_36); -x_38 = lean_array_get_size(x_2); -x_39 = lean_nat_dec_lt(x_6, x_38); -lean_dec(x_38); -x_40 = lean_array_get_size(x_1); -x_41 = lean_nat_dec_lt(x_6, x_40); -lean_dec(x_40); -if (x_37 == 0) -{ -lean_object* x_112; lean_object* x_113; -x_112 = l_Lean_instInhabitedExpr; -x_113 = l_outOfBounds___rarg(x_112); -x_42 = x_113; -goto block_111; -} -else -{ -lean_object* x_114; -x_114 = lean_array_fget(x_5, x_6); -x_42 = x_114; -goto block_111; -} -block_35: -{ -if (lean_obj_tag(x_18) == 0) -{ -lean_object* x_19; -x_19 = lean_ctor_get(x_18, 0); -lean_inc(x_19); -if (lean_obj_tag(x_19) == 0) -{ -uint8_t x_20; -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_6); -x_20 = !lean_is_exclusive(x_18); -if (x_20 == 0) -{ -lean_object* x_21; lean_object* x_22; -x_21 = lean_ctor_get(x_18, 0); -lean_dec(x_21); -x_22 = lean_ctor_get(x_19, 0); -lean_inc(x_22); -lean_dec(x_19); -lean_ctor_set(x_18, 0, x_22); -return x_18; -} -else -{ -lean_object* x_23; lean_object* x_24; lean_object* x_25; -x_23 = lean_ctor_get(x_18, 1); -lean_inc(x_23); -lean_dec(x_18); -x_24 = lean_ctor_get(x_19, 0); -lean_inc(x_24); -lean_dec(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_23); -return x_25; -} -} -else -{ -lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_26 = lean_ctor_get(x_18, 1); -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_4, 2); -x_29 = lean_nat_add(x_6, x_28); -lean_dec(x_6); -x_5 = x_27; -x_6 = x_29; -x_7 = lean_box(0); -x_8 = lean_box(0); -x_14 = x_26; -goto _start; -} -} -else -{ -uint8_t x_31; -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_6); -x_31 = !lean_is_exclusive(x_18); -if (x_31 == 0) -{ -return x_18; -} -else -{ -lean_object* x_32; lean_object* x_33; lean_object* x_34; -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_dec(x_18); -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; -} -} -} -block_111: -{ -uint8_t x_43; uint8_t x_44; -x_43 = l_Lean_Expr_hasLooseBVars(x_42); -if (x_39 == 0) -{ -uint8_t x_108; -x_108 = 0; -x_44 = x_108; -goto block_107; -} -else -{ -lean_object* x_109; uint8_t x_110; -x_109 = lean_array_fget(x_2, x_6); -x_110 = lean_unbox(x_109); -lean_dec(x_109); -x_44 = x_110; -goto block_107; -} -block_107: -{ -uint8_t x_45; -if (x_41 == 0) -{ -uint8_t x_104; -x_104 = 0; -x_45 = x_104; -goto block_103; -} -else -{ -lean_object* x_105; uint8_t x_106; -x_105 = lean_array_fget(x_1, x_6); -x_106 = lean_unbox(x_105); -lean_dec(x_105); -x_45 = x_106; -goto block_103; -} -block_103: -{ -if (x_43 == 0) -{ -uint8_t x_46; -x_46 = l_Lean_Expr_hasMVar(x_42); -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; -x_47 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_mkGroundPattern(x_42); -x_48 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___spec__2___lambda__1(x_6, x_5, x_47, x_9, x_10, x_11, x_12, x_13, x_14); -x_49 = lean_ctor_get(x_48, 0); -lean_inc(x_49); -x_50 = lean_ctor_get(x_48, 1); -lean_inc(x_50); -lean_dec(x_48); -x_51 = lean_ctor_get(x_49, 0); -lean_inc(x_51); -lean_dec(x_49); -x_52 = lean_ctor_get(x_4, 2); -x_53 = lean_nat_add(x_6, x_52); -lean_dec(x_6); -x_5 = x_51; -x_6 = x_53; -x_7 = lean_box(0); -x_8 = lean_box(0); -x_14 = x_50; -goto _start; -} -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_dec(x_42); -x_55 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_dontCare; -x_56 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___spec__2___lambda__1(x_6, x_5, x_55, x_9, x_10, x_11, x_12, x_13, x_14); -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_inc(x_59); -lean_dec(x_57); -x_60 = lean_ctor_get(x_4, 2); -x_61 = lean_nat_add(x_6, x_60); -lean_dec(x_6); -x_5 = x_59; -x_6 = x_61; -x_7 = lean_box(0); -x_8 = lean_box(0); -x_14 = x_58; -goto _start; -} -} -else -{ -if (lean_obj_tag(x_42) == 0) -{ -lean_object* x_63; lean_object* x_64; -x_63 = lean_ctor_get(x_42, 0); -lean_inc(x_63); -x_64 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_foundBVar(x_63, x_9, x_10, x_11, x_12, x_13, x_14); -if (x_44 == 0) -{ -if (x_45 == 0) -{ -lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; -x_65 = lean_ctor_get(x_64, 1); -lean_inc(x_65); -lean_dec(x_64); -x_66 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_saveBVar(x_63, x_9, x_10, x_11, x_12, x_13, x_65); -x_67 = lean_ctor_get(x_66, 1); -lean_inc(x_67); -lean_dec(x_66); -x_68 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___spec__2___lambda__1(x_6, x_5, x_42, x_9, x_10, x_11, x_12, x_13, x_67); -x_18 = x_68; -goto block_35; -} -else -{ -lean_object* x_69; uint8_t x_70; -x_69 = lean_ctor_get(x_64, 0); -lean_inc(x_69); -x_70 = lean_unbox(x_69); -lean_dec(x_69); -if (x_70 == 0) -{ -lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; -x_71 = lean_ctor_get(x_64, 1); -lean_inc(x_71); -lean_dec(x_64); -x_72 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_saveBVar(x_63, x_9, x_10, x_11, x_12, x_13, x_71); -x_73 = lean_ctor_get(x_72, 1); -lean_inc(x_73); -lean_dec(x_72); -x_74 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___spec__2___lambda__1(x_6, x_5, x_42, x_9, x_10, x_11, x_12, x_13, x_73); -x_18 = x_74; -goto block_35; -} -else -{ -lean_object* x_75; lean_object* x_76; lean_object* x_77; -lean_dec(x_63); -lean_dec(x_42); -x_75 = lean_ctor_get(x_64, 1); -lean_inc(x_75); -lean_dec(x_64); -x_76 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_dontCare; -x_77 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___spec__2___lambda__1(x_6, x_5, x_76, x_9, x_10, x_11, x_12, x_13, x_75); -x_18 = x_77; -goto block_35; -} -} -} -else -{ -lean_object* x_78; uint8_t x_79; -x_78 = lean_ctor_get(x_64, 0); -lean_inc(x_78); -x_79 = lean_unbox(x_78); -lean_dec(x_78); -if (x_79 == 0) -{ -lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; -x_80 = lean_ctor_get(x_64, 1); -lean_inc(x_80); -lean_dec(x_64); -x_81 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_saveBVar(x_63, x_9, x_10, x_11, x_12, x_13, x_80); -x_82 = lean_ctor_get(x_81, 1); -lean_inc(x_82); -lean_dec(x_81); -x_83 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___spec__2___lambda__1(x_6, x_5, x_42, x_9, x_10, x_11, x_12, x_13, x_82); -x_18 = x_83; -goto block_35; -} -else -{ -lean_object* x_84; lean_object* x_85; lean_object* x_86; -lean_dec(x_63); -lean_dec(x_42); -x_84 = lean_ctor_get(x_64, 1); -lean_inc(x_84); -lean_dec(x_64); -x_85 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_dontCare; -x_86 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___spec__2___lambda__1(x_6, x_5, x_85, x_9, x_10, x_11, x_12, x_13, x_84); -x_18 = x_86; -goto block_35; -} -} -} -else -{ -if (x_44 == 0) -{ -if (x_45 == 0) -{ -lean_object* x_87; -x_87 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_getPatternFn_x3f(x_42); -if (lean_obj_tag(x_87) == 0) -{ -lean_object* x_88; lean_object* x_89; -lean_dec(x_42); -x_88 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_dontCare; -x_89 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___spec__2___lambda__1(x_6, x_5, x_88, x_9, x_10, x_11, x_12, x_13, x_14); -x_18 = x_89; -goto block_35; -} -else -{ -uint8_t x_90; lean_object* x_91; -lean_dec(x_87); -x_90 = 0; -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -x_91 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go(x_42, x_90, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_91) == 0) -{ -lean_object* x_92; lean_object* x_93; lean_object* x_94; -x_92 = lean_ctor_get(x_91, 0); -lean_inc(x_92); -x_93 = lean_ctor_get(x_91, 1); -lean_inc(x_93); -lean_dec(x_91); -x_94 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___spec__2___lambda__1(x_6, x_5, x_92, x_9, x_10, x_11, x_12, x_13, x_93); -x_18 = x_94; -goto block_35; -} -else -{ -uint8_t x_95; -lean_dec(x_5); -x_95 = !lean_is_exclusive(x_91); -if (x_95 == 0) -{ -x_18 = x_91; -goto block_35; -} -else -{ -lean_object* x_96; lean_object* x_97; lean_object* x_98; -x_96 = lean_ctor_get(x_91, 0); -x_97 = lean_ctor_get(x_91, 1); -lean_inc(x_97); -lean_inc(x_96); -lean_dec(x_91); -x_98 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_98, 0, x_96); -lean_ctor_set(x_98, 1, x_97); -x_18 = x_98; -goto block_35; -} -} -} -} -else -{ -lean_object* x_99; lean_object* x_100; -lean_dec(x_42); -x_99 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_dontCare; -x_100 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___spec__2___lambda__1(x_6, x_5, x_99, x_9, x_10, x_11, x_12, x_13, x_14); -x_18 = x_100; -goto block_35; -} -} -else -{ -lean_object* x_101; lean_object* x_102; -lean_dec(x_42); -x_101 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_dontCare; -x_102 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___spec__2___lambda__1(x_6, x_5, x_101, x_9, x_10, x_11, x_12, x_13, x_14); -x_18 = x_102; -goto block_35; -} -} -} -} -} -} -} -} -} -static lean_object* _init_l_panic___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___spec__3___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_instMonadMetaM; -x_2 = l_ReaderT_instMonad___rarg(x_1); -return x_2; -} -} -static lean_object* _init_l_panic___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___spec__3___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_panic___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___spec__3___closed__1; -x_2 = l_Lean_instInhabitedExpr; -x_3 = l_instInhabitedOfMonad___rarg(x_1, x_2); -return x_3; -} -} -LEAN_EXPORT lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_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) { -_start: -{ -lean_object* x_8; lean_object* x_9; lean_object* x_10; -x_8 = l_panic___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___spec__3___closed__2; -x_9 = lean_panic_fn(x_8, x_1); -x_10 = lean_apply_6(x_9, x_2, x_3, x_4, x_5, x_6, x_7); -return x_10; -} -} -LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_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) { -_start: -{ -lean_object* x_8; lean_object* x_9; uint8_t x_10; -x_8 = lean_ctor_get(x_5, 5); -x_9 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_1, x_3, x_4, x_5, x_6, x_7); -x_10 = !lean_is_exclusive(x_9); -if (x_10 == 0) -{ -lean_object* x_11; lean_object* x_12; -x_11 = lean_ctor_get(x_9, 0); -lean_inc(x_8); -x_12 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_12, 0, x_8); -lean_ctor_set(x_12, 1, x_11); -lean_ctor_set_tag(x_9, 1); -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); -lean_inc(x_8); -x_15 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_15, 0, x_8); -lean_ctor_set(x_15, 1, x_13); -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; -} -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("invalid pattern, (non-forbidden) application expected", 53, 53); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__1; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__3() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("assertion violation: ", 21, 21); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__4() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("f.isConst || f.isFVar\n ", 24, 24); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__5() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__3; -x_2 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__4; -x_3 = lean_string_append(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__6() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("_private.Lean.Meta.Tactic.Grind.TheoremPatterns.0.Lean.Meta.Grind.NormalizePattern.go", 85, 85); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__7() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____lambda__1___closed__1; -x_2 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__6; -x_3 = lean_unsigned_to_nat(126u); -x_4 = lean_unsigned_to_nat(2u); -x_5 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__5; -x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); -return x_6; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_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; -x_9 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_getPatternFn_x3f(x_1); -if (lean_obj_tag(x_9) == 0) -{ -lean_object* x_10; lean_object* x_11; -lean_dec(x_1); -x_10 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__2; -x_11 = l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___spec__1(x_10, 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_11; -} -else -{ -lean_object* x_12; lean_object* x_13; uint8_t x_48; -x_12 = lean_ctor_get(x_9, 0); -lean_inc(x_12); -lean_dec(x_9); -x_48 = l_Lean_Expr_isConst(x_12); -if (x_48 == 0) -{ -uint8_t x_49; -x_49 = l_Lean_Expr_isFVar(x_12); -if (x_49 == 0) -{ -lean_object* x_50; lean_object* x_51; -lean_dec(x_12); -lean_dec(x_1); -x_50 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__7; -x_51 = l_panic___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___spec__3(x_50, x_3, x_4, x_5, x_6, x_7, x_8); -return x_51; -} -else -{ -lean_object* x_52; -x_52 = lean_box(0); -x_13 = x_52; -goto block_47; -} -} -else -{ -lean_object* x_53; -x_53 = lean_box(0); -x_13 = x_53; -goto block_47; -} -block_47: -{ -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; -lean_dec(x_13); -lean_inc(x_12); -x_14 = l_Lean_Expr_toHeadIndex(x_12); -x_15 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_saveSymbol(x_14, x_3, x_4, x_5, x_6, x_7, x_8); -x_16 = lean_ctor_get(x_15, 1); -lean_inc(x_16); -lean_dec(x_15); -x_17 = lean_unsigned_to_nat(0u); -x_18 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_17); -x_19 = l_Lean_Meta_Grind_ppPattern___closed__5; -lean_inc(x_18); -x_20 = lean_mk_array(x_18, x_19); -x_21 = lean_unsigned_to_nat(1u); -x_22 = lean_nat_sub(x_18, x_21); -lean_dec(x_18); -x_23 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_20, x_22); -x_24 = lean_array_get_size(x_23); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_24); -lean_inc(x_12); -x_25 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_getPatternFunInfo(x_12, x_24, x_4, x_5, x_6, x_7, x_16); -if (lean_obj_tag(x_25) == 0) -{ -lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_26 = lean_ctor_get(x_25, 0); -lean_inc(x_26); -x_27 = lean_ctor_get(x_25, 1); -lean_inc(x_27); -lean_dec(x_25); -x_28 = lean_ctor_get(x_26, 0); -lean_inc(x_28); -x_29 = lean_ctor_get(x_26, 1); -lean_inc(x_29); -lean_dec(x_26); -x_30 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_30, 0, x_17); -lean_ctor_set(x_30, 1, x_24); -lean_ctor_set(x_30, 2, x_21); -x_31 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___spec__2(x_28, x_29, x_30, x_30, x_23, x_17, lean_box(0), lean_box(0), x_3, x_4, x_5, x_6, x_7, x_27); -lean_dec(x_30); -lean_dec(x_29); -lean_dec(x_28); -if (lean_obj_tag(x_31) == 0) -{ -uint8_t x_32; -x_32 = !lean_is_exclusive(x_31); -if (x_32 == 0) -{ -lean_object* x_33; lean_object* x_34; -x_33 = lean_ctor_get(x_31, 0); -x_34 = l_Lean_mkAppN(x_12, x_33); -lean_dec(x_33); -lean_ctor_set(x_31, 0, x_34); -return x_31; -} -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_inc(x_36); -lean_inc(x_35); -lean_dec(x_31); -x_37 = l_Lean_mkAppN(x_12, x_35); -lean_dec(x_35); -x_38 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_38, 0, x_37); -lean_ctor_set(x_38, 1, x_36); -return x_38; -} -} -else -{ -uint8_t x_39; -lean_dec(x_12); -x_39 = !lean_is_exclusive(x_31); -if (x_39 == 0) -{ -return x_31; -} -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_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; -} -} -} -else -{ -uint8_t x_43; -lean_dec(x_24); -lean_dec(x_23); -lean_dec(x_12); -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_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); -return x_46; -} -} -} -} -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("invalid pattern, it does not have pattern variables", 51, 51); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___closed__1; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -if (x_2 == 0) -{ -lean_object* x_9; lean_object* x_10; -x_9 = lean_box(0); -x_10 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1(x_1, x_9, x_3, x_4, x_5, x_6, x_7, x_8); -return x_10; -} -else -{ -uint8_t x_11; -x_11 = l_Lean_Expr_hasLooseBVars(x_1); -if (x_11 == 0) -{ -lean_object* x_12; lean_object* x_13; uint8_t x_14; -lean_dec(x_1); -x_12 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___closed__2; -x_13 = l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___spec__4(x_12, x_3, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -x_14 = !lean_is_exclusive(x_13); -if (x_14 == 0) -{ -return x_13; -} -else -{ -lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = lean_ctor_get(x_13, 0); -x_16 = lean_ctor_get(x_13, 1); -lean_inc(x_16); -lean_inc(x_15); -lean_dec(x_13); -x_17 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_17, 0, x_15); -lean_ctor_set(x_17, 1, x_16); -return x_17; -} -} -else -{ -lean_object* x_18; lean_object* x_19; -x_18 = lean_box(0); -x_19 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1(x_1, x_18, x_3, x_4, x_5, x_6, x_7, x_8); -return x_19; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_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, lean_object* x_7) { -_start: -{ -lean_object* x_8; -x_8 = l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_8; -} -} -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___spec__2___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -lean_object* x_10; -x_10 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___spec__2___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_1); -return x_10; -} -} -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { -_start: -{ -lean_object* x_15; -x_15 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_15; -} -} -LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_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) { -_start: -{ -lean_object* x_8; -x_8 = l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___spec__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_8; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_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) { -_start: -{ -lean_object* x_9; -x_9 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_2); -return x_9; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_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) { -_start: -{ -uint8_t x_9; lean_object* x_10; -x_9 = lean_unbox(x_2); -lean_dec(x_2); -x_10 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go(x_1, x_9, x_3, x_4, x_5, x_6, x_7, x_8); -return x_10; -} -} -LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_Meta_Grind_NormalizePattern_main___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -if (lean_obj_tag(x_1) == 0) -{ -lean_object* x_9; lean_object* x_10; -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -x_9 = l_List_reverse___rarg(x_2); -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_1); -if (x_11 == 0) -{ -lean_object* x_12; lean_object* x_13; uint8_t x_14; lean_object* x_15; -x_12 = lean_ctor_get(x_1, 0); -x_13 = lean_ctor_get(x_1, 1); -x_14 = 0; -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -x_15 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go(x_12, 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; -x_16 = lean_ctor_get(x_15, 0); -lean_inc(x_16); -x_17 = lean_ctor_get(x_15, 1); -lean_inc(x_17); -lean_dec(x_15); -lean_ctor_set(x_1, 1, x_2); -lean_ctor_set(x_1, 0, x_16); -{ -lean_object* _tmp_0 = x_13; -lean_object* _tmp_1 = x_1; -lean_object* _tmp_7 = x_17; -x_1 = _tmp_0; -x_2 = _tmp_1; -x_8 = _tmp_7; -} -goto _start; -} -else -{ -uint8_t x_19; -lean_free_object(x_1); -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_19 = !lean_is_exclusive(x_15); -if (x_19 == 0) -{ -return x_15; -} -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_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; -} -} -} -else -{ -lean_object* x_23; lean_object* x_24; uint8_t x_25; lean_object* x_26; -x_23 = lean_ctor_get(x_1, 0); -x_24 = lean_ctor_get(x_1, 1); -lean_inc(x_24); -lean_inc(x_23); -lean_dec(x_1); -x_25 = 0; -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -x_26 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go(x_23, x_25, x_3, x_4, x_5, x_6, x_7, x_8); -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 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_29, 0, x_27); -lean_ctor_set(x_29, 1, x_2); -x_1 = x_24; -x_2 = x_29; -x_8 = x_28; -goto _start; -} -else -{ -lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; -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_2); -x_31 = lean_ctor_get(x_26, 0); -lean_inc(x_31); -x_32 = lean_ctor_get(x_26, 1); -lean_inc(x_32); -if (lean_is_exclusive(x_26)) { - lean_ctor_release(x_26, 0); - lean_ctor_release(x_26, 1); - x_33 = x_26; -} else { - lean_dec_ref(x_26); - x_33 = lean_box(0); -} -if (lean_is_scalar(x_33)) { - x_34 = lean_alloc_ctor(1, 2, 0); -} else { - x_34 = x_33; -} -lean_ctor_set(x_34, 0, x_31); -lean_ctor_set(x_34, 1, x_32); -return x_34; -} -} -} -} -} -static lean_object* _init_l_Lean_Meta_Grind_NormalizePattern_main___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = lean_box(0); -x_2 = lean_array_mk(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Meta_Grind_NormalizePattern_main___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_Grind_NormalizePattern_main___closed__1; -x_2 = l_Lean_SMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__17___closed__3; -x_3 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set(x_3, 1, x_2); -lean_ctor_set(x_3, 2, x_2); -return x_3; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_NormalizePattern_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) { -_start: -{ -lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; -x_7 = lean_box(0); -x_8 = l_Lean_Meta_Grind_NormalizePattern_main___closed__2; -x_9 = lean_st_mk_ref(x_8, 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; -x_11 = lean_ctor_get(x_9, 0); -x_12 = lean_ctor_get(x_9, 1); -lean_inc(x_11); -x_13 = l_List_mapM_loop___at_Lean_Meta_Grind_NormalizePattern_main___spec__1(x_1, x_7, x_11, x_2, x_3, x_4, x_5, 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_st_ref_get(x_11, x_15); -lean_dec(x_11); -x_17 = !lean_is_exclusive(x_16); -if (x_17 == 0) -{ -lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_18 = lean_ctor_get(x_16, 0); -x_19 = lean_ctor_get(x_18, 0); -lean_inc(x_19); -lean_dec(x_18); -x_20 = lean_array_to_list(x_19); -lean_ctor_set(x_9, 1, x_20); -lean_ctor_set(x_9, 0, x_14); -lean_ctor_set(x_16, 0, x_9); -return x_16; -} -else -{ -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; -x_21 = lean_ctor_get(x_16, 0); -x_22 = lean_ctor_get(x_16, 1); -lean_inc(x_22); -lean_inc(x_21); -lean_dec(x_16); -x_23 = lean_ctor_get(x_21, 0); -lean_inc(x_23); -lean_dec(x_21); -x_24 = lean_array_to_list(x_23); -lean_ctor_set(x_9, 1, x_24); -lean_ctor_set(x_9, 0, x_14); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_9); -lean_ctor_set(x_25, 1, x_22); -return x_25; -} -} -else -{ -uint8_t x_26; -lean_free_object(x_9); -lean_dec(x_11); -x_26 = !lean_is_exclusive(x_13); -if (x_26 == 0) -{ -return x_13; -} -else -{ -lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_27 = lean_ctor_get(x_13, 0); -x_28 = lean_ctor_get(x_13, 1); -lean_inc(x_28); -lean_inc(x_27); -lean_dec(x_13); -x_29 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_29, 0, x_27); -lean_ctor_set(x_29, 1, x_28); -return x_29; -} -} -} -else -{ -lean_object* x_30; lean_object* x_31; lean_object* x_32; -x_30 = lean_ctor_get(x_9, 0); -x_31 = lean_ctor_get(x_9, 1); -lean_inc(x_31); -lean_inc(x_30); -lean_dec(x_9); -lean_inc(x_30); -x_32 = l_List_mapM_loop___at_Lean_Meta_Grind_NormalizePattern_main___spec__1(x_1, x_7, x_30, x_2, x_3, x_4, x_5, x_31); -if (lean_obj_tag(x_32) == 0) -{ -lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; -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); -if (lean_is_exclusive(x_35)) { - lean_ctor_release(x_35, 0); - lean_ctor_release(x_35, 1); - x_38 = x_35; -} else { - lean_dec_ref(x_35); - x_38 = lean_box(0); -} -x_39 = lean_ctor_get(x_36, 0); -lean_inc(x_39); -lean_dec(x_36); -x_40 = lean_array_to_list(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_40); -if (lean_is_scalar(x_38)) { - x_42 = lean_alloc_ctor(0, 2, 0); -} else { - x_42 = x_38; -} -lean_ctor_set(x_42, 0, x_41); -lean_ctor_set(x_42, 1, x_37); -return x_42; -} -else -{ -lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; -lean_dec(x_30); -x_43 = lean_ctor_get(x_32, 0); -lean_inc(x_43); -x_44 = lean_ctor_get(x_32, 1); -lean_inc(x_44); -if (lean_is_exclusive(x_32)) { - lean_ctor_release(x_32, 0); - lean_ctor_release(x_32, 1); - x_45 = x_32; -} else { - lean_dec_ref(x_32); - x_45 = lean_box(0); -} -if (lean_is_scalar(x_45)) { - x_46 = lean_alloc_ctor(1, 2, 0); -} else { - x_46 = x_45; -} -lean_ctor_set(x_46, 0, x_43); -lean_ctor_set(x_46, 1, x_44); -return x_46; -} -} -} -} -static lean_object* _init_l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addTheoremPattern___spec__1___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_SMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__17___closed__5; -x_2 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_2, 0, x_1); -lean_ctor_set(x_2, 1, x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addTheoremPattern___spec__1___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_SMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__17___closed__5; -x_2 = lean_alloc_ctor(0, 6, 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); -lean_ctor_set(x_2, 4, x_1); -lean_ctor_set(x_2, 5, x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addTheoremPattern___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, 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_6, 6); -lean_inc(x_9); -lean_dec(x_6); -x_10 = lean_st_ref_take(x_7, 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_13 = !lean_is_exclusive(x_11); -if (x_13 == 0) -{ -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; -x_14 = lean_ctor_get(x_11, 0); -x_15 = lean_ctor_get(x_11, 4); -lean_dec(x_15); -x_16 = l_Lean_ScopedEnvExtension_addCore___rarg(x_14, x_1, x_2, x_3, x_9); -x_17 = l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addTheoremPattern___spec__1___closed__1; -lean_ctor_set(x_11, 4, x_17); -lean_ctor_set(x_11, 0, x_16); -x_18 = lean_st_ref_set(x_7, x_11, x_12); -x_19 = lean_ctor_get(x_18, 1); -lean_inc(x_19); -lean_dec(x_18); -x_20 = lean_st_ref_take(x_5, x_19); -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; uint8_t x_27; -x_24 = lean_ctor_get(x_21, 1); -lean_dec(x_24); -x_25 = l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addTheoremPattern___spec__1___closed__2; -lean_ctor_set(x_21, 1, x_25); -x_26 = lean_st_ref_set(x_5, x_21, x_22); -x_27 = !lean_is_exclusive(x_26); -if (x_27 == 0) -{ -lean_object* x_28; lean_object* x_29; -x_28 = lean_ctor_get(x_26, 0); -lean_dec(x_28); -x_29 = lean_box(0); -lean_ctor_set(x_26, 0, x_29); -return x_26; -} -else -{ -lean_object* x_30; lean_object* x_31; lean_object* x_32; -x_30 = lean_ctor_get(x_26, 1); -lean_inc(x_30); -lean_dec(x_26); -x_31 = lean_box(0); -x_32 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_32, 0, x_31); -lean_ctor_set(x_32, 1, x_30); -return x_32; -} -} -else -{ -lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; -x_33 = lean_ctor_get(x_21, 0); -x_34 = lean_ctor_get(x_21, 2); -x_35 = lean_ctor_get(x_21, 3); -x_36 = lean_ctor_get(x_21, 4); -lean_inc(x_36); -lean_inc(x_35); -lean_inc(x_34); -lean_inc(x_33); -lean_dec(x_21); -x_37 = l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addTheoremPattern___spec__1___closed__2; -x_38 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_38, 0, x_33); -lean_ctor_set(x_38, 1, x_37); -lean_ctor_set(x_38, 2, x_34); -lean_ctor_set(x_38, 3, x_35); -lean_ctor_set(x_38, 4, x_36); -x_39 = lean_st_ref_set(x_5, x_38, x_22); -x_40 = lean_ctor_get(x_39, 1); -lean_inc(x_40); -if (lean_is_exclusive(x_39)) { - lean_ctor_release(x_39, 0); - lean_ctor_release(x_39, 1); - x_41 = x_39; -} else { - lean_dec_ref(x_39); - x_41 = lean_box(0); -} -x_42 = lean_box(0); -if (lean_is_scalar(x_41)) { - x_43 = lean_alloc_ctor(0, 2, 0); -} else { - x_43 = x_41; -} -lean_ctor_set(x_43, 0, x_42); -lean_ctor_set(x_43, 1, x_40); -return x_43; -} -} -else -{ -lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; 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_44 = lean_ctor_get(x_11, 0); -x_45 = lean_ctor_get(x_11, 1); -x_46 = lean_ctor_get(x_11, 2); -x_47 = lean_ctor_get(x_11, 3); -x_48 = lean_ctor_get(x_11, 5); -x_49 = lean_ctor_get(x_11, 6); -x_50 = lean_ctor_get(x_11, 7); -lean_inc(x_50); -lean_inc(x_49); -lean_inc(x_48); -lean_inc(x_47); -lean_inc(x_46); -lean_inc(x_45); -lean_inc(x_44); -lean_dec(x_11); -x_51 = l_Lean_ScopedEnvExtension_addCore___rarg(x_44, x_1, x_2, x_3, x_9); -x_52 = l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addTheoremPattern___spec__1___closed__1; -x_53 = lean_alloc_ctor(0, 8, 0); -lean_ctor_set(x_53, 0, x_51); -lean_ctor_set(x_53, 1, x_45); -lean_ctor_set(x_53, 2, x_46); -lean_ctor_set(x_53, 3, x_47); -lean_ctor_set(x_53, 4, x_52); -lean_ctor_set(x_53, 5, x_48); -lean_ctor_set(x_53, 6, x_49); -lean_ctor_set(x_53, 7, x_50); -x_54 = lean_st_ref_set(x_7, x_53, x_12); -x_55 = lean_ctor_get(x_54, 1); -lean_inc(x_55); -lean_dec(x_54); -x_56 = lean_st_ref_take(x_5, x_55); -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_inc(x_59); -x_60 = lean_ctor_get(x_57, 2); -lean_inc(x_60); -x_61 = lean_ctor_get(x_57, 3); -lean_inc(x_61); -x_62 = lean_ctor_get(x_57, 4); -lean_inc(x_62); -if (lean_is_exclusive(x_57)) { - lean_ctor_release(x_57, 0); - lean_ctor_release(x_57, 1); - lean_ctor_release(x_57, 2); - lean_ctor_release(x_57, 3); - lean_ctor_release(x_57, 4); - x_63 = x_57; -} else { - lean_dec_ref(x_57); - x_63 = lean_box(0); -} -x_64 = l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addTheoremPattern___spec__1___closed__2; -if (lean_is_scalar(x_63)) { - x_65 = lean_alloc_ctor(0, 5, 0); -} else { - x_65 = x_63; -} -lean_ctor_set(x_65, 0, x_59); -lean_ctor_set(x_65, 1, x_64); -lean_ctor_set(x_65, 2, x_60); -lean_ctor_set(x_65, 3, x_61); -lean_ctor_set(x_65, 4, x_62); -x_66 = lean_st_ref_set(x_5, x_65, x_58); -x_67 = lean_ctor_get(x_66, 1); -lean_inc(x_67); -if (lean_is_exclusive(x_66)) { - lean_ctor_release(x_66, 0); - lean_ctor_release(x_66, 1); - x_68 = x_66; -} else { - lean_dec_ref(x_66); - x_68 = lean_box(0); -} -x_69 = lean_box(0); -if (lean_is_scalar(x_68)) { - x_70 = lean_alloc_ctor(0, 2, 0); -} else { - x_70 = x_68; -} -lean_ctor_set(x_70, 0, x_69); -lean_ctor_set(x_70, 1, x_67); -return x_70; -} -} -} -LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Meta_Grind_addTheoremPattern___spec__2(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_Meta_Grind_ppPattern(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_Meta_Grind_ppPattern(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_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Meta_Grind_addTheoremPattern___spec__3(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; -x_5 = lean_ctor_get(x_1, 1); -lean_ctor_set(x_1, 1, x_2); -{ -lean_object* _tmp_0 = x_5; -lean_object* _tmp_1 = x_1; -x_1 = _tmp_0; -x_2 = _tmp_1; -} -goto _start; -} -else -{ -lean_object* x_7; lean_object* x_8; lean_object* x_9; -x_7 = lean_ctor_get(x_1, 0); -x_8 = lean_ctor_get(x_1, 1); -lean_inc(x_8); -lean_inc(x_7); -lean_dec(x_1); -x_9 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_9, 0, x_7); -lean_ctor_set(x_9, 1, x_2); -x_1 = x_8; -x_2 = x_9; -goto _start; -} -} -} -} -static lean_object* _init_l_Lean_Meta_Grind_addTheoremPattern___lambda__1___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = l_Lean_Meta_Grind_theoremPatternsExt; -return x_1; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addTheoremPattern___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; uint8_t x_15; lean_object* x_16; -x_12 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_12, 0, x_1); -x_13 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_13, 0, x_2); -lean_ctor_set(x_13, 1, x_3); -lean_ctor_set(x_13, 2, x_4); -lean_ctor_set(x_13, 3, x_5); -lean_ctor_set(x_13, 4, x_12); -x_14 = l_Lean_Meta_Grind_addTheoremPattern___lambda__1___closed__1; -x_15 = 0; -x_16 = l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addTheoremPattern___spec__1(x_14, x_13, x_15, x_7, x_8, x_9, x_10, x_11); -return x_16; -} -} -static lean_object* _init_l_Lean_Meta_Grind_addTheoremPattern___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("`", 1, 1); -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_Grind_addTheoremPattern___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_Grind_addTheoremPattern___closed__1; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Meta_Grind_addTheoremPattern___closed__3() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("` is not a theorem, you cannot assign patterns to non-theorems for the `grind` tactic", 85, 85); -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_Grind_addTheoremPattern___closed__4() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_Grind_addTheoremPattern___closed__3; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Meta_Grind_addTheoremPattern___closed__5() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("pattern", 7, 7); -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_Grind_addTheoremPattern___closed__6() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_mkGroundPattern___closed__1; -x_2 = l_Lean_Meta_Grind_addTheoremPattern___closed__5; -x_3 = l_Lean_Name_mkStr2(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Meta_Grind_addTheoremPattern___closed__7() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked(": ", 2, 2); -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_Grind_addTheoremPattern___closed__8() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_Grind_addTheoremPattern___closed__7; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addTheoremPattern(lean_object* x_1, lean_object* x_2, lean_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 = l_Lean_getConstInfo___at_Lean_Meta_mkConstWithFreshMVarLevels___spec__1(x_1, x_4, x_5, x_6, x_7, x_8); -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) == 2) -{ -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_11 = lean_ctor_get(x_9, 1); -lean_inc(x_11); -lean_dec(x_9); -x_12 = lean_ctor_get(x_10, 0); -lean_inc(x_12); -lean_dec(x_10); -x_13 = lean_ctor_get(x_12, 0); -lean_inc(x_13); -lean_dec(x_12); -x_14 = lean_ctor_get(x_13, 1); -lean_inc(x_14); -lean_dec(x_13); -x_15 = lean_box(0); -x_16 = l_List_mapTR_loop___at_Lean_mkConstWithLevelParams___spec__1(x_14, x_15); -lean_inc(x_1); -x_17 = l_Lean_Expr_const___override(x_1, x_16); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -x_18 = l_Lean_Meta_Grind_NormalizePattern_main(x_3, x_4, x_5, x_6, x_7, x_11); -if (lean_obj_tag(x_18) == 0) -{ -lean_object* x_19; lean_object* x_20; uint8_t x_21; -x_19 = lean_ctor_get(x_18, 0); -lean_inc(x_19); -x_20 = lean_ctor_get(x_18, 1); -lean_inc(x_20); -lean_dec(x_18); -x_21 = !lean_is_exclusive(x_19); -if (x_21 == 0) -{ -lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; -x_22 = lean_ctor_get(x_19, 0); -x_23 = lean_ctor_get(x_19, 1); -x_24 = l_Lean_Meta_Grind_addTheoremPattern___closed__6; -x_25 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_24, x_4, x_5, x_6, x_7, x_20); -x_26 = lean_ctor_get(x_25, 0); -lean_inc(x_26); -x_27 = lean_unbox(x_26); -lean_dec(x_26); -if (x_27 == 0) -{ -lean_object* x_28; lean_object* x_29; lean_object* x_30; -lean_free_object(x_19); -x_28 = lean_ctor_get(x_25, 1); -lean_inc(x_28); -lean_dec(x_25); -x_29 = lean_box(0); -x_30 = l_Lean_Meta_Grind_addTheoremPattern___lambda__1(x_1, x_17, x_2, x_22, x_23, x_29, x_4, x_5, x_6, x_7, x_28); -lean_dec(x_7); -lean_dec(x_5); -lean_dec(x_4); -return x_30; -} -else -{ -uint8_t x_31; -x_31 = !lean_is_exclusive(x_25); -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; lean_object* x_39; 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_32 = lean_ctor_get(x_25, 1); -x_33 = lean_ctor_get(x_25, 0); -lean_dec(x_33); -lean_inc(x_1); -x_34 = l_Lean_MessageData_ofName(x_1); -x_35 = l_Lean_Meta_Grind_ppPattern___closed__4; -lean_ctor_set_tag(x_25, 7); -lean_ctor_set(x_25, 1, x_34); -lean_ctor_set(x_25, 0, x_35); -x_36 = l_Lean_Meta_Grind_addTheoremPattern___closed__8; -lean_ctor_set_tag(x_19, 7); -lean_ctor_set(x_19, 1, x_36); -lean_ctor_set(x_19, 0, x_25); -lean_inc(x_22); -x_37 = l_List_mapTR_loop___at_Lean_Meta_Grind_addTheoremPattern___spec__2(x_22, x_15); -x_38 = l_List_mapTR_loop___at_Lean_Meta_Grind_addTheoremPattern___spec__3(x_37, x_15); -x_39 = l_Lean_MessageData_ofList(x_38); -x_40 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_40, 0, x_19); -lean_ctor_set(x_40, 1, x_39); -x_41 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_41, 0, x_40); -lean_ctor_set(x_41, 1, x_35); -x_42 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_24, x_41, x_4, x_5, x_6, x_7, x_32); -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_Meta_Grind_addTheoremPattern___lambda__1(x_1, x_17, x_2, x_22, x_23, x_43, x_4, x_5, x_6, x_7, x_44); -lean_dec(x_7); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_43); -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; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; -x_46 = lean_ctor_get(x_25, 1); -lean_inc(x_46); -lean_dec(x_25); -lean_inc(x_1); -x_47 = l_Lean_MessageData_ofName(x_1); -x_48 = l_Lean_Meta_Grind_ppPattern___closed__4; -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_Meta_Grind_addTheoremPattern___closed__8; -lean_ctor_set_tag(x_19, 7); -lean_ctor_set(x_19, 1, x_50); -lean_ctor_set(x_19, 0, x_49); -lean_inc(x_22); -x_51 = l_List_mapTR_loop___at_Lean_Meta_Grind_addTheoremPattern___spec__2(x_22, x_15); -x_52 = l_List_mapTR_loop___at_Lean_Meta_Grind_addTheoremPattern___spec__3(x_51, x_15); -x_53 = l_Lean_MessageData_ofList(x_52); -x_54 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_54, 0, x_19); -lean_ctor_set(x_54, 1, x_53); -x_55 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_55, 0, x_54); -lean_ctor_set(x_55, 1, x_48); -x_56 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_24, x_55, x_4, x_5, x_6, x_7, x_46); -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_Grind_addTheoremPattern___lambda__1(x_1, x_17, x_2, x_22, x_23, x_57, x_4, x_5, x_6, x_7, x_58); -lean_dec(x_7); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_57); -return x_59; -} -} -} -else -{ -lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; uint8_t x_65; -x_60 = lean_ctor_get(x_19, 0); -x_61 = lean_ctor_get(x_19, 1); -lean_inc(x_61); -lean_inc(x_60); -lean_dec(x_19); -x_62 = l_Lean_Meta_Grind_addTheoremPattern___closed__6; -x_63 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_62, x_4, x_5, x_6, x_7, x_20); -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; -x_66 = lean_ctor_get(x_63, 1); -lean_inc(x_66); -lean_dec(x_63); -x_67 = lean_box(0); -x_68 = l_Lean_Meta_Grind_addTheoremPattern___lambda__1(x_1, x_17, x_2, x_60, x_61, x_67, x_4, x_5, x_6, x_7, x_66); -lean_dec(x_7); -lean_dec(x_5); -lean_dec(x_4); -return x_68; -} -else -{ -lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_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_ctor_get(x_63, 1); -lean_inc(x_69); -if (lean_is_exclusive(x_63)) { - lean_ctor_release(x_63, 0); - lean_ctor_release(x_63, 1); - x_70 = x_63; -} else { - lean_dec_ref(x_63); - x_70 = lean_box(0); -} -lean_inc(x_1); -x_71 = l_Lean_MessageData_ofName(x_1); -x_72 = l_Lean_Meta_Grind_ppPattern___closed__4; -if (lean_is_scalar(x_70)) { - x_73 = lean_alloc_ctor(7, 2, 0); -} else { - x_73 = x_70; - lean_ctor_set_tag(x_73, 7); -} -lean_ctor_set(x_73, 0, x_72); -lean_ctor_set(x_73, 1, x_71); -x_74 = l_Lean_Meta_Grind_addTheoremPattern___closed__8; -x_75 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_75, 0, x_73); -lean_ctor_set(x_75, 1, x_74); -lean_inc(x_60); -x_76 = l_List_mapTR_loop___at_Lean_Meta_Grind_addTheoremPattern___spec__2(x_60, x_15); -x_77 = l_List_mapTR_loop___at_Lean_Meta_Grind_addTheoremPattern___spec__3(x_76, x_15); -x_78 = l_Lean_MessageData_ofList(x_77); -x_79 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_79, 0, x_75); -lean_ctor_set(x_79, 1, x_78); -x_80 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_80, 0, x_79); -lean_ctor_set(x_80, 1, x_72); -x_81 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_62, x_80, x_4, x_5, x_6, x_7, x_69); -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_Meta_Grind_addTheoremPattern___lambda__1(x_1, x_17, x_2, x_60, x_61, x_82, x_4, x_5, x_6, x_7, x_83); -lean_dec(x_7); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_82); -return x_84; -} -} -} -else -{ -uint8_t x_85; -lean_dec(x_17); -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_85 = !lean_is_exclusive(x_18); -if (x_85 == 0) -{ -return x_18; -} -else -{ -lean_object* x_86; lean_object* x_87; lean_object* x_88; -x_86 = lean_ctor_get(x_18, 0); -x_87 = lean_ctor_get(x_18, 1); -lean_inc(x_87); -lean_inc(x_86); -lean_dec(x_18); -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 -{ -lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; -lean_dec(x_10); -lean_dec(x_3); -lean_dec(x_2); -x_89 = lean_ctor_get(x_9, 1); -lean_inc(x_89); -lean_dec(x_9); -x_90 = l_Lean_MessageData_ofName(x_1); -x_91 = l_Lean_Meta_Grind_addTheoremPattern___closed__2; -x_92 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_92, 0, x_91); -lean_ctor_set(x_92, 1, x_90); -x_93 = l_Lean_Meta_Grind_addTheoremPattern___closed__4; -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_Lean_throwError___at_Lean_Meta_setInlineAttribute___spec__1(x_94, x_4, x_5, x_6, x_7, x_89); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -return x_95; -} -} -else -{ -uint8_t x_96; -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_96 = !lean_is_exclusive(x_9); -if (x_96 == 0) -{ -return x_9; -} -else -{ -lean_object* x_97; lean_object* x_98; lean_object* x_99; -x_97 = lean_ctor_get(x_9, 0); -x_98 = lean_ctor_get(x_9, 1); -lean_inc(x_98); -lean_inc(x_97); -lean_dec(x_9); -x_99 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_99, 0, x_97); -lean_ctor_set(x_99, 1, x_98); -return x_99; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addTheoremPattern___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -uint8_t x_9; lean_object* x_10; -x_9 = lean_unbox(x_3); -lean_dec(x_3); -x_10 = l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addTheoremPattern___spec__1(x_1, x_2, x_9, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_7); -lean_dec(x_5); -lean_dec(x_4); -return x_10; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addTheoremPattern___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_Meta_Grind_addTheoremPattern___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -lean_dec(x_10); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -return x_12; -} -} -lean_object* initialize_Lean_HeadIndex(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_InferType(uint8_t builtin, lean_object*); -static bool _G_initialized = false; -LEAN_EXPORT lean_object* initialize_Lean_Meta_Tactic_Grind_TheoremPatterns(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_HeadIndex(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); -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_InferType(builtin, lean_io_mk_world()); -if (lean_io_result_is_error(res)) return res; -lean_dec_ref(res); -l_Lean_Meta_Grind_instInhabitedOrigin___closed__1 = _init_l_Lean_Meta_Grind_instInhabitedOrigin___closed__1(); -lean_mark_persistent(l_Lean_Meta_Grind_instInhabitedOrigin___closed__1); -l_Lean_Meta_Grind_instInhabitedOrigin = _init_l_Lean_Meta_Grind_instInhabitedOrigin(); -lean_mark_persistent(l_Lean_Meta_Grind_instInhabitedOrigin); -l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__1(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__1); -l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__2(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__2); -l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__3(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__3); -l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__4(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__4); -l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__5 = _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__5(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__5); -l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__6 = _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__6(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__6); -l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__7 = _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__7(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__7); -l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__8 = _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__8(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__8); -l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__9 = _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__9(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__9); -l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__10 = _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__10(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__10); -l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__11 = _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__11(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__11); -l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__12 = _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__12(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__12); -l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__13 = _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__13(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__13); -l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__14 = _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__14(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__14); -l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__15 = _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__15(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__15); -l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__16 = _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__16(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__16); -l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__17 = _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__17(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_51____closed__17); -l_Lean_Meta_Grind_instReprOrigin___closed__1 = _init_l_Lean_Meta_Grind_instReprOrigin___closed__1(); -lean_mark_persistent(l_Lean_Meta_Grind_instReprOrigin___closed__1); -l_Lean_Meta_Grind_instReprOrigin = _init_l_Lean_Meta_Grind_instReprOrigin(); -lean_mark_persistent(l_Lean_Meta_Grind_instReprOrigin); -l_Lean_Meta_Grind_instInhabitedTheoremPattern___closed__1 = _init_l_Lean_Meta_Grind_instInhabitedTheoremPattern___closed__1(); -lean_mark_persistent(l_Lean_Meta_Grind_instInhabitedTheoremPattern___closed__1); -l_Lean_Meta_Grind_instInhabitedTheoremPattern___closed__2 = _init_l_Lean_Meta_Grind_instInhabitedTheoremPattern___closed__2(); -lean_mark_persistent(l_Lean_Meta_Grind_instInhabitedTheoremPattern___closed__2); -l_Lean_Meta_Grind_instInhabitedTheoremPattern = _init_l_Lean_Meta_Grind_instInhabitedTheoremPattern(); -lean_mark_persistent(l_Lean_Meta_Grind_instInhabitedTheoremPattern); -l_panic___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__1___closed__1 = _init_l_panic___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__1___closed__1(); -lean_mark_persistent(l_panic___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__1___closed__1); -l_panic___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__1___closed__2 = _init_l_panic___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__1___closed__2(); -lean_mark_persistent(l_panic___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__1___closed__2); -l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__4___closed__1 = _init_l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__4___closed__1(); -l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__4___closed__2 = _init_l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__4___closed__2(); -l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__9___closed__1 = _init_l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__9___closed__1(); -lean_mark_persistent(l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__9___closed__1); -l_Lean_SMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__17___closed__1 = _init_l_Lean_SMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__17___closed__1(); -lean_mark_persistent(l_Lean_SMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__17___closed__1); -l_Lean_SMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__17___closed__2 = _init_l_Lean_SMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__17___closed__2(); -lean_mark_persistent(l_Lean_SMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__17___closed__2); -l_Lean_SMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__17___closed__3 = _init_l_Lean_SMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__17___closed__3(); -lean_mark_persistent(l_Lean_SMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__17___closed__3); -l_Lean_SMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__17___closed__4 = _init_l_Lean_SMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__17___closed__4(); -lean_mark_persistent(l_Lean_SMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__17___closed__4); -l_Lean_SMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__17___closed__5 = _init_l_Lean_SMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__17___closed__5(); -lean_mark_persistent(l_Lean_SMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__17___closed__5); -l_Lean_SMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__17___closed__6 = _init_l_Lean_SMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__17___closed__6(); -lean_mark_persistent(l_Lean_SMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__17___closed__6); -l_Lean_SMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__17 = _init_l_Lean_SMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__17(); -lean_mark_persistent(l_Lean_SMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____spec__17); -l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____lambda__1___closed__1 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____lambda__1___closed__1(); -lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____lambda__1___closed__1); -l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____lambda__1___closed__2 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____lambda__1___closed__2(); -lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____lambda__1___closed__2); -l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____lambda__1___closed__3 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____lambda__1___closed__3(); -lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____lambda__1___closed__3); -l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____lambda__1___closed__4 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____lambda__1___closed__4(); -lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____lambda__1___closed__4); -l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____closed__1 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____closed__1(); -lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____closed__1); -l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____closed__2 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____closed__2(); -lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____closed__2); -l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____closed__3 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____closed__3(); -lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____closed__3); -l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____closed__4 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____closed__4(); -lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____closed__4); -l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____closed__5 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____closed__5(); -lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____closed__5); -l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____closed__6 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____closed__6(); -lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____closed__6); -l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____closed__7 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____closed__7(); -lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307____closed__7); -if (builtin) {res = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_TheoremPatterns___hyg_307_(lean_io_mk_world()); -if (lean_io_result_is_error(res)) return res; -l_Lean_Meta_Grind_theoremPatternsExt = lean_io_result_get_value(res); -lean_mark_persistent(l_Lean_Meta_Grind_theoremPatternsExt); -lean_dec_ref(res); -}l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__1(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__1); -l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__2(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__2); -l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__3(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__3); -l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__4(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__4); -l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__5 = _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__5(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__5); -l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__6 = _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__6(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__6); -l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__7 = _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__7(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__7); -l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__8 = _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__8(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__8); -l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__9 = _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__9(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__9); -l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__10 = _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__10(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__10); -l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__11 = _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__11(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__11); -l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__12 = _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__12(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__12); -l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__13 = _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__13(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__13); -l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__14 = _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__14(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__14); -l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__15 = _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__15(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__15); -l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__16 = _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__16(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__16); -l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__17 = _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__17(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__17); -l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__18 = _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__18(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__18); -l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__19 = _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__19(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames___closed__19); -l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames = _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_forbiddenDeclNames); -l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_dontCare___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_dontCare___closed__1(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_dontCare___closed__1); -l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_dontCare___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_dontCare___closed__2(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_dontCare___closed__2); -l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_dontCare___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_dontCare___closed__3(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_dontCare___closed__3); -l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_dontCare = _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_dontCare(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_dontCare); -l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_mkGroundPattern___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_mkGroundPattern___closed__1(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_mkGroundPattern___closed__1); -l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_mkGroundPattern___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_mkGroundPattern___closed__2(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_mkGroundPattern___closed__2); -l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_mkGroundPattern___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_mkGroundPattern___closed__3(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_mkGroundPattern___closed__3); -l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___closed__1 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___closed__1(); -lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___closed__1); -l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___closed__2 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___closed__2(); -lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___closed__2); -l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___closed__3 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___closed__3(); -lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___closed__3); -l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1(); -lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1); -l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2(); -lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2); -l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3(); -lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3); -l_Lean_Meta_Grind_ppPattern___closed__1 = _init_l_Lean_Meta_Grind_ppPattern___closed__1(); -lean_mark_persistent(l_Lean_Meta_Grind_ppPattern___closed__1); -l_Lean_Meta_Grind_ppPattern___closed__2 = _init_l_Lean_Meta_Grind_ppPattern___closed__2(); -lean_mark_persistent(l_Lean_Meta_Grind_ppPattern___closed__2); -l_Lean_Meta_Grind_ppPattern___closed__3 = _init_l_Lean_Meta_Grind_ppPattern___closed__3(); -lean_mark_persistent(l_Lean_Meta_Grind_ppPattern___closed__3); -l_Lean_Meta_Grind_ppPattern___closed__4 = _init_l_Lean_Meta_Grind_ppPattern___closed__4(); -lean_mark_persistent(l_Lean_Meta_Grind_ppPattern___closed__4); -l_Lean_Meta_Grind_ppPattern___closed__5 = _init_l_Lean_Meta_Grind_ppPattern___closed__5(); -lean_mark_persistent(l_Lean_Meta_Grind_ppPattern___closed__5); -l_Lean_Meta_Grind_ppPattern___closed__6 = _init_l_Lean_Meta_Grind_ppPattern___closed__6(); -lean_mark_persistent(l_Lean_Meta_Grind_ppPattern___closed__6); -l_Lean_Meta_Grind_ppPattern___closed__7 = _init_l_Lean_Meta_Grind_ppPattern___closed__7(); -lean_mark_persistent(l_Lean_Meta_Grind_ppPattern___closed__7); -l_Lean_Meta_Grind_ppPattern___closed__8 = _init_l_Lean_Meta_Grind_ppPattern___closed__8(); -lean_mark_persistent(l_Lean_Meta_Grind_ppPattern___closed__8); -l_Lean_Meta_Grind_ppPattern___closed__9 = _init_l_Lean_Meta_Grind_ppPattern___closed__9(); -lean_mark_persistent(l_Lean_Meta_Grind_ppPattern___closed__9); -l_Lean_Meta_Grind_ppPattern___closed__10 = _init_l_Lean_Meta_Grind_ppPattern___closed__10(); -lean_mark_persistent(l_Lean_Meta_Grind_ppPattern___closed__10); -l_Lean_Meta_Grind_ppPattern___closed__11 = _init_l_Lean_Meta_Grind_ppPattern___closed__11(); -lean_mark_persistent(l_Lean_Meta_Grind_ppPattern___closed__11); -l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_getPatternFunInfo___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_getPatternFunInfo___closed__1(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_getPatternFunInfo___closed__1); -l_panic___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___spec__3___closed__1 = _init_l_panic___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___spec__3___closed__1(); -lean_mark_persistent(l_panic___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___spec__3___closed__1); -l_panic___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___spec__3___closed__2 = _init_l_panic___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___spec__3___closed__2(); -lean_mark_persistent(l_panic___at___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___spec__3___closed__2); -l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__1(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__1); -l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__2(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__2); -l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__3(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__3); -l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__4(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__4); -l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__5 = _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__5(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__5); -l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__6 = _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__6(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__6); -l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__7 = _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__7(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__7); -l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___closed__1(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___closed__1); -l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___closed__2(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_TheoremPatterns_0__Lean_Meta_Grind_NormalizePattern_go___closed__2); -l_Lean_Meta_Grind_NormalizePattern_main___closed__1 = _init_l_Lean_Meta_Grind_NormalizePattern_main___closed__1(); -lean_mark_persistent(l_Lean_Meta_Grind_NormalizePattern_main___closed__1); -l_Lean_Meta_Grind_NormalizePattern_main___closed__2 = _init_l_Lean_Meta_Grind_NormalizePattern_main___closed__2(); -lean_mark_persistent(l_Lean_Meta_Grind_NormalizePattern_main___closed__2); -l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addTheoremPattern___spec__1___closed__1 = _init_l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addTheoremPattern___spec__1___closed__1(); -lean_mark_persistent(l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addTheoremPattern___spec__1___closed__1); -l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addTheoremPattern___spec__1___closed__2 = _init_l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addTheoremPattern___spec__1___closed__2(); -lean_mark_persistent(l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addTheoremPattern___spec__1___closed__2); -l_Lean_Meta_Grind_addTheoremPattern___lambda__1___closed__1 = _init_l_Lean_Meta_Grind_addTheoremPattern___lambda__1___closed__1(); -lean_mark_persistent(l_Lean_Meta_Grind_addTheoremPattern___lambda__1___closed__1); -l_Lean_Meta_Grind_addTheoremPattern___closed__1 = _init_l_Lean_Meta_Grind_addTheoremPattern___closed__1(); -lean_mark_persistent(l_Lean_Meta_Grind_addTheoremPattern___closed__1); -l_Lean_Meta_Grind_addTheoremPattern___closed__2 = _init_l_Lean_Meta_Grind_addTheoremPattern___closed__2(); -lean_mark_persistent(l_Lean_Meta_Grind_addTheoremPattern___closed__2); -l_Lean_Meta_Grind_addTheoremPattern___closed__3 = _init_l_Lean_Meta_Grind_addTheoremPattern___closed__3(); -lean_mark_persistent(l_Lean_Meta_Grind_addTheoremPattern___closed__3); -l_Lean_Meta_Grind_addTheoremPattern___closed__4 = _init_l_Lean_Meta_Grind_addTheoremPattern___closed__4(); -lean_mark_persistent(l_Lean_Meta_Grind_addTheoremPattern___closed__4); -l_Lean_Meta_Grind_addTheoremPattern___closed__5 = _init_l_Lean_Meta_Grind_addTheoremPattern___closed__5(); -lean_mark_persistent(l_Lean_Meta_Grind_addTheoremPattern___closed__5); -l_Lean_Meta_Grind_addTheoremPattern___closed__6 = _init_l_Lean_Meta_Grind_addTheoremPattern___closed__6(); -lean_mark_persistent(l_Lean_Meta_Grind_addTheoremPattern___closed__6); -l_Lean_Meta_Grind_addTheoremPattern___closed__7 = _init_l_Lean_Meta_Grind_addTheoremPattern___closed__7(); -lean_mark_persistent(l_Lean_Meta_Grind_addTheoremPattern___closed__7); -l_Lean_Meta_Grind_addTheoremPattern___closed__8 = _init_l_Lean_Meta_Grind_addTheoremPattern___closed__8(); -lean_mark_persistent(l_Lean_Meta_Grind_addTheoremPattern___closed__8); -return lean_io_result_mk_ok(lean_box(0)); -} -#ifdef __cplusplus -} -#endif diff --git a/stage0/stdlib/Lean/Meta/Tactic/Grind/Types.c b/stage0/stdlib/Lean/Meta/Tactic/Grind/Types.c index b303c03e7ddf..89336ad917da 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Grind/Types.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Grind/Types.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Lean.Meta.Tactic.Grind.Types -// Imports: Lean.Util.ShareCommon Lean.HeadIndex Lean.Meta.Basic Lean.Meta.CongrTheorems Lean.Meta.AbstractNestedProofs Lean.Meta.Tactic.Simp.Types Lean.Meta.Tactic.Util Lean.Meta.Tactic.Grind.Canon Lean.Meta.Tactic.Grind.Attr +// Imports: Init.Grind.Tactics Init.Data.Queue Lean.Util.ShareCommon Lean.HeadIndex Lean.Meta.Basic Lean.Meta.CongrTheorems Lean.Meta.AbstractNestedProofs Lean.Meta.Tactic.Simp.Types Lean.Meta.Tactic.Util Lean.Meta.Tactic.Grind.Canon Lean.Meta.Tactic.Grind.Attr Lean.Meta.Tactic.Grind.EMatchTheorem #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -13,250 +13,313 @@ #ifdef __cplusplus extern "C" { #endif -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__19; +static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1___closed__1; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_registerParent___spec__7(lean_object*, lean_object*); lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_instInhabitedMethods___closed__2; +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GoalM_run___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_Grind_instInhabitedMethods; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__27; +LEAN_EXPORT lean_object* l_Lean_RBNode_ins___at_Lean_Meta_Grind_registerParent___spec__2(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_instInhabitedENode___closed__1; -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__39; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instHashablePreInstance_unsafe__1___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getEqc_go___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pushEq(lean_object*, lean_object*, lean_object*, 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_insert___at_Lean_Meta_Grind_setENode___spec__1(lean_object*, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_setENode___spec__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instReprENode; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_forEachENode___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_markTheoremInstance___spec__4(size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isEqv___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_Meta_Grind_getENode___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_mk_empty_array_with_capacity(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addNewFact___boxed(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_shift_right(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pushEqHEq___boxed(lean_object*, lean_object*, 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_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__41; lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_abstractNestedProofs(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_Meta_Grind_GoalM_run___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_Meta_Grind_getTrueExpr___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_registerParent___spec__6___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__22; -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__45; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__47; static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_118____closed__4; +static size_t l_Lean_Meta_Grind_instInhabitedGoal___closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_canon___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_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__54; -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_setENode_unsafe__2___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getTarget_x3f___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_Grind_abstractNestedProofs___closed__2; -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__15; -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__35; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getConfig(lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_registerParent___spec__4(lean_object*, size_t, size_t, lean_object*, lean_object*); uint8_t lean_usize_dec_le(size_t, size_t); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedGoal___spec__2; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getMethodsRef(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_erase___at_Lean_Meta_Grind_getParentsAndReset___spec__1(lean_object*, size_t); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_erase___at_Lean_Meta_Grind_getParentsAndReset___spec__1(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Lean_Meta_Grind_instBEqENodeKey(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instHashablePreInstance___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_filterENodes___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_instBEqPreInstance___lambda__2___closed__1; uint64_t lean_uint64_of_nat(lean_object*); uint64_t lean_uint64_mix_hash(uint64_t, uint64_t); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GoalM_run(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_Array_qsort_sort___at_Lean_Meta_Grind_getENodes___spec__4___lambda__1___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_setENode_unsafe__1___rarg(lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__46; size_t lean_uint64_to_usize(uint64_t); uint64_t lean_uint64_lor(uint64_t, uint64_t); -LEAN_EXPORT size_t l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_toENodeKey_unsafe__1(lean_object*); uint8_t l_Lean_Expr_isAppOfArity(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isTrueExpr___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_Grind_instInhabitedGoal___closed__5; static lean_object* l_Lean_Meta_Grind_getENode___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkENode(lean_object*, lean_object*, 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_Grind_getMainDeclName___boxed(lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__51; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GoalM_run_x27(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_isApp(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_118_(lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__24; LEAN_EXPORT uint8_t l_Lean_Meta_Grind_isCongruent(lean_object*, lean_object*, lean_object*); uint8_t l_Lean_RBNode_isRed___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isInterpreted(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*); size_t lean_usize_mul(size_t, size_t); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__7; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_alreadyInternalized___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__53; +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); lean_object* lean_mk_array(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkHCongrWithArity___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_usize_dec_eq(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pushEq___boxed(lean_object*, lean_object*, lean_object*, 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_insertAtCollisionNodeAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__4(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isSameExpr_unsafe__1___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__44; LEAN_EXPORT uint64_t l_Lean_Meta_Grind_congrHash(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__11; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__3; lean_object* l_Lean_Expr_bvar___override(lean_object*); lean_object* lean_array_fget(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getMaxGeneration___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getFalseExpr___boxed(lean_object*, lean_object*); lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_grind_debug_proofs; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_registerParent___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_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__8; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pushEqCore(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___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__7; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_forEachENode(lean_object*, 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_Grind_propagateDown(lean_object*, 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___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot___spec__2(lean_object*, size_t, size_t); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot___spec__2(lean_object*, size_t, lean_object*); static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78____closed__8; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_forEachEqc(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_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__37; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_alreadyInternalized___spec__2___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__38; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getEqcs(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Std_Queue_enqueue___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instBEqPreInstance___lambda__2(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint64_t l_Lean_Meta_Grind_instHashableCongrTheoremCacheKey(lean_object*); -LEAN_EXPORT size_t l_Lean_Meta_Grind_alreadyInternalized_unsafe__1(lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Grind_getENodes___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_registerParent___spec__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_setENode(lean_object*, lean_object*, 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_insertAux_traverse___at_Lean_Meta_Grind_registerParent___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__19; lean_object* l_Lean_Name_mkStr5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__17; +size_t lean_usize_of_nat(lean_object*); lean_object* l_Nat_nextPowerOfTwo_go(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getMethods___boxed(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_Meta_Grind_instInhabitedNewFact___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getMainDeclName___rarg(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_Grind_mkHCongrWithArity___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkHCongrWithArity___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getFalseExpr(lean_object*, lean_object*); -LEAN_EXPORT size_t l_Lean_Meta_Grind_getENode_x3f_unsafe__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instInhabitedMethods___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_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__33; LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at_Lean_Meta_Grind_copyParentsTo___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_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_registerParent___spec__6(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_registerParent___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkENodeCore___boxed(lean_object*, lean_object*, lean_object*, lean_object*, 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_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3(lean_object*, size_t, size_t, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_markTheoremInstance___spec__13___lambda__1(lean_object*, lean_object*, lean_object*); static uint64_t l_Lean_Meta_Grind_hasSameType___closed__1; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__9; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instBEqCongrKey___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instBEqCongrTheoremCacheKey___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Methods_toMethodsRef___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_getParentsAndReset___spec__2___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__8; -LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_alreadyInternalized___spec__1(lean_object*, size_t); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__50; -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_setENode___spec__4(lean_object*, lean_object*, size_t, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__38; +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_alreadyInternalized___spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_setENode___spec__4(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getEqc___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__46; -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__3; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isCongrRoot(lean_object*, 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_Grind_grind_debug; LEAN_EXPORT lean_object* l_Lean_MVarId_assign___at_Lean_Meta_Grind_closeGoal___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_MethodsRef_toMethods(lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__2; lean_object* l_Lean_Expr_appArg_x21(lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_AbstractNestedProofs_visit(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_Grind_propagateUp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_registerParent___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, size_t); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__15; static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_118____closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isEqTrue___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint64_t l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instHashablePreInstance___spec__1(lean_object*, lean_object*, lean_object*, size_t, size_t, uint64_t); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_markTheoremInstance___spec__11___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instHashablePreInstance___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_registerParent___spec__5(size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getMaxGeneration___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getRoot(lean_object*, 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_Grind_mkEqProof___boxed(lean_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_Queue_empty(lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_markTheoremInstance___spec__1(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__52; uint8_t l_Lean_Expr_isLambda(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instBEqPreInstance___lambda__3___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Meta_Grind_getENodes___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instBEqPreInstance___boxed(lean_object*, lean_object*); size_t lean_ptr_addr(lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__18; lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_isAssigned___at_Lean_Meta_Grind_closeGoal___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___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__16; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pushEqCore___boxed(lean_object*, lean_object*, lean_object*, lean_object*, 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_Grind_GoalM_run___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_Grind_getFalseExpr___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Methods_toMethodsRef_unsafe__1___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getTarget_x3f(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_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__33; lean_object* l_Lean_PersistentHashMap_isUnaryNode___rarg(lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_registerParent___spec__8(lean_object*, size_t, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__1; static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78____closed__1; -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_setENode_unsafe__1___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_setENode_unsafe__1___boxed(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Lean_Meta_Grind_instBEqPreInstance___lambda__1(lean_object*); static lean_object* l_Lean_Meta_Grind_getENode___closed__4; -LEAN_EXPORT lean_object* l_Lean_RBNode_insert___at_Lean_Meta_Grind_registerParent___spec__4(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markTheoremInstance___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_Grind_closeGoal___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_Grind_getRootENode___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_take(lean_object*, lean_object*); lean_object* l_Lean_Expr_getRevArg_x21(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isEqFalse___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__25; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getParentsAndReset___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_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78____closed__2; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isCongrRoot___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_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_Grind_getENode___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_Grind_filterENodes___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 size_t l_Lean_Meta_Grind_getENode_unsafe__1(lean_object*); +static lean_object* l_Lean_Meta_Grind_instBEqPreInstance___lambda__2___closed__2; LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hasSameRoot(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instHashablePreInstance_unsafe__2___boxed(lean_object*); uint64_t lean_uint64_shift_right(uint64_t, uint64_t); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getGeneration___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_RBNode_setBlack___rarg(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addTheoremInstance___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint64_t lean_usize_to_uint64(size_t); lean_object* lean_nat_to_int(lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_registerParent___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_registerParent___spec__8___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEqHEqProof(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MVarId_getType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static size_t l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__2___closed__1; static lean_object* l_Lean_Meta_Grind_abstractNestedProofs___closed__1; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_forEachEqc___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_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pushEqHEq(lean_object*, lean_object*, lean_object*, 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_Grind_getENodes___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_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__10; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__16; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_checkMaxInstancesExceeded(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_Grind_closeGoal(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MVarId_admit(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_qsort_sort___at_Lean_Meta_Grind_getENodes___spec__4___closed__1; LEAN_EXPORT lean_object* l_Lean_MVarId_assign___at_Lean_Meta_Grind_closeGoal___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___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__13; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__18; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getENode_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getConfig___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkHCongrWithArity___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_Grind_checkMaxInstancesExceeded___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_mkHCongrWithArityForConst_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_setENode___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__2(lean_object*, size_t, size_t, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__9; -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_registerParent___spec__8(size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3___lambda__2___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78_(lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__30; +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_markTheoremInstance___spec__11___lambda__2(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__14(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78____closed__6; +LEAN_EXPORT uint64_t l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__5(lean_object*, lean_object*, lean_object*, size_t, size_t, uint64_t); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getENode___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_Grind_getTrueExpr___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__25; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkENode___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_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_markTheoremInstance___spec__13___lambda__2___boxed(lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_get(lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_foldlMAux___at_Lean_MetavarContext_getExprAssignmentDomain___spec__2___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_congrPlaceholderProof___closed__1; -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__5; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__48; +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__8___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_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_markTheoremInstance___spec__7(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isCongruent___boxed(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*); LEAN_EXPORT lean_object* l_Lean_MVarId_isAssigned___at_Lean_Meta_Grind_closeGoal___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_Grind_getMainDeclName(lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__36; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__21; static lean_object* l_Lean_Meta_Grind_canon___closed__3; -LEAN_EXPORT lean_object* l_Lean_RBNode_ins___at_Lean_Meta_Grind_registerParent___spec__5(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__47; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__20; +LEAN_EXPORT uint64_t l_Lean_Meta_Grind_instHashablePreInstance(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_Grind_getENodes___spec__3(size_t, size_t, lean_object*); +LEAN_EXPORT uint64_t l_Lean_Meta_Grind_instHashableENodeKey(lean_object*); LEAN_EXPORT uint64_t l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_mkHCongrWithArity___spec__5(lean_object*, lean_object*); -LEAN_EXPORT size_t l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_toENodeKey(lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__39; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_MethodsRef_toMethods_unsafe__1(lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__31; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__34; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isSameExpr___boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__13; static size_t l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__2___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getEqc(lean_object*, 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_Grind_getMethodsRef___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_congrHash___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pushEqTrue(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_Meta_Grind_instBEqCongrTheoremCacheKey(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_toENodeKey___boxed(lean_object*); +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_markTheoremInstance___spec__9(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_setENode_unsafe__1___rarg___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkHCongrWithArity(lean_object*, 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_Grind_getEqcs___boxed(lean_object*, 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_Grind_mkHCongrWithArity___spec__2___closed__3; LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at_Lean_Meta_Grind_copyParentsTo___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_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_registerParent___spec__7___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_canon___closed__1; static lean_object* l_Lean_Meta_Grind_congrHash___closed__1; lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getGeneration(lean_object*, 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_Grind_instBEqENodeKey___boxed(lean_object*, lean_object*); lean_object* l_Lean_Meta_mkHCongrWithArity(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Meta_Grind_isSameExpr_unsafe__1(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_118____closed__5; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_alreadyInternalized(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_PersistentHashMap_containsAux___at_Lean_Meta_Grind_alreadyInternalized___spec__2(lean_object*, size_t, size_t); +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_alreadyInternalized___spec__2(lean_object*, size_t, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__41; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isCongruent_go___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getNext___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_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_registerParent___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_Grind_getENodes___spec__3___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instBEqPreInstance___lambda__2___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getMethods(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getMainDeclName___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_filterENodes(lean_object*, 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_containsAtAux___at_Lean_Meta_Grind_markTheoremInstance___spec__13___lambda__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getParents(lean_object*, 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_insertAux_traverse___at_Lean_Meta_Grind_markTheoremInstance___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getRoot___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_PersistentHashMap_findAux___at_Lean_Meta_Grind_registerParent___spec__2(lean_object*, size_t, size_t); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__30; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isInterpreted___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Grind_getENodes___spec__1___closed__2; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markTheoremInstance___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_PersistentHashMap_foldlM___at_Lean_Meta_Grind_getENodes___spec__2(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot___spec__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Grind_getENodes___spec__1___closed__1; static lean_object* l_Lean_Meta_Grind_instInhabitedGoal___closed__1; lean_object* l_Lean_Expr_appFn_x21(lean_object*); +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_markTheoremInstance___spec__13(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1(lean_object*, lean_object*); lean_object* l_Array_eraseIdx___rarg(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__36; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_hasSameType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkHEqProof___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_Grind_instHashableCongrKey___boxed(lean_object*, lean_object*); +LEAN_EXPORT uint64_t l_Lean_Meta_Grind_instHashablePreInstance_unsafe__1(lean_object*); lean_object* l_Lean_Meta_Grind_Canon_canonImpl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_registerParent___spec__9(lean_object*, lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markAsInconsistent___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_PersistentHashMap_contains___at_Lean_MVarId_isAssigned___spec__1(lean_object*, lean_object*); lean_object* l_Lean_Meta_mkFalseElim(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -270,106 +333,133 @@ LEAN_EXPORT uint64_t l_Lean_Meta_Grind_instHashableCongrKey(lean_object*, lean_o LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getENodes(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_Grind_copyParentsTo(lean_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_register___at_Lean_Elab_initFn____x40_Lean_Elab_AutoBound___hyg_6____spec__1(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__40; -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_registerParent___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__6; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__14; +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__14___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_instReprENode___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_shareCommon___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_PersistentHashMap_find_x3f___at___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot___spec__1(lean_object*, size_t); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GoalM_run___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_PersistentHashMap_find_x3f___at___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot___spec__1(lean_object*, lean_object*); +LEAN_EXPORT uint64_t l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__10(lean_object*, lean_object*, lean_object*, size_t, size_t, uint64_t); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_registerParent___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_118____closed__3; -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__21; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__11; +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_markTheoremInstance___spec__7___lambda__2(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_GoalM_run_x27___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Goal_admit(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__32; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getNext(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT size_t l_Lean_Meta_Grind_setENode_unsafe__1(lean_object*); +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3___lambda__2(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_setENode_unsafe__1(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_abstractNestedProofs___closed__3; -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getENode_unsafe__1___boxed(lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_registerParent___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__23; +LEAN_EXPORT uint64_t l_Lean_Meta_Grind_instHashablePreInstance_unsafe__2(lean_object*); lean_object* lean_grind_mk_eq_proof(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*); +lean_object* lean_grind_mk_heq_proof(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint64_t l_Lean_Meta_Grind_congrHash_go(lean_object*, lean_object*, uint64_t); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__35; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instInhabitedMethods___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_Grind_setENode_unsafe__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instBEqPreInstance___lambda__1___boxed(lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__54; +LEAN_EXPORT lean_object* l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___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_Meta_Grind_addNewFact(lean_object*, lean_object*, lean_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_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_markTheoremInstance___spec__11___lambda__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_congrPlaceholderProof___closed__3; +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_markTheoremInstance___spec__7___lambda__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_alreadyInternalized___spec__1___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_registerParent___spec__6(lean_object*, size_t, lean_object*); lean_object* lean_string_length(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_abstractNestedProofs___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_Grind_setENode_unsafe__2___rarg(lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEqFalseProof(lean_object*, 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_Grind_shareCommon(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_Grind_instBEqPreInstance(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_markTheoremInstance___spec__7___lambda__2___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78____closed__4; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_mkHCongrWithArity___spec__3(size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addTheoremInstance(lean_object*, lean_object*, lean_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_Grind_Types_0__Lean_Meta_Grind_MethodsRef_toMethods___boxed(lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__29; +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_markTheoremInstance___spec__13___lambda__2(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__5; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__22; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GoalM_run_x27___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_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_setENode___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_insert___at_Lean_MVarId_assign___spec__1(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot_unsafe__1___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isInconsistent___boxed(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*); static lean_object* l_Lean_Meta_Grind_instInhabitedMethods___closed__1; lean_object* l_Lean_Meta_isConstructorAppCore_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getEqc_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_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_registerParent(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_Meta_Grind_isSameExpr(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_congrPlaceholderProof___closed__2; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__45; lean_object* l_Lean_indentExpr(lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__32; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getMaxGeneration___rarg___boxed(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_Lean_Meta_Grind_instInhabitedENode; -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__44; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_registerParent___spec__3(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markAsInconsistent(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_find_x3f___at_Lean_Meta_Grind_registerParent___spec__1___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_markTheoremInstance___spec__13___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_instInhabitedGoal___closed__2; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markTheoremInstance(lean_object*, lean_object*, 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_Grind_mkHCongrWithArity___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_Grind_instBEqCongrKey(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Grind_getENodes___spec__1___boxed(lean_object*); +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_markTheoremInstance___spec__7___lambda__1(lean_object*, lean_object*, lean_object*); lean_object* lean_array_set(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78____closed__3; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_markTheoremInstance___spec__9___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Grind_getENodes___spec__1(lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__40; 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_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_getParentsAndReset___spec__2(lean_object*, size_t, size_t); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__50; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_getParentsAndReset___spec__2(lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isEqv(lean_object*, lean_object*, 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_Grind_pushEqFalse___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_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__43; -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__48; lean_object* l___private_Lean_Expr_0__Lean_reprExpr____x40_Lean_Expr___hyg_3036_(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__12(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__2___boxed(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_PersistentHashMap_findAux___at_Lean_Meta_Grind_registerParent___spec__2___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Methods_toMethodsRef(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEqTrueProof(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_indexOfAux___at_Lean_Meta_Grind_getParentsAndReset___spec__3(lean_object*, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Array_indexOfAux___at_Lean_Meta_Grind_getParentsAndReset___spec__3(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_indexOfAux___at_Lean_Meta_Grind_getParentsAndReset___spec__3___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__28; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isRoot___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_Grind_setENode___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_Grind_copyParentsTo___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint64_t lean_uint64_shift_left(uint64_t, uint64_t); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__6; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Methods_toMethodsRef_unsafe__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pushHEq(lean_object*, lean_object*, 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_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__20; -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__4; lean_object* l___private_Init_Data_Array_QSort_0__Array_qpartition___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getENode_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_alreadyInternalized_unsafe__1___boxed(lean_object*); lean_object* l_Lean_PersistentHashMap_mkCollisionNode___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__24; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getTrueExpr___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getENode(lean_object*, 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_Grind_getTrueExpr(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkENodeCore(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_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__12; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GoalM_run_x27___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_forEachENode___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_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isTrueExpr(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_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__23; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot___spec__2___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__10; LEAN_EXPORT uint8_t l_Array_qsort_sort___at_Lean_Meta_Grind_getENodes___spec__4___lambda__1(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__49; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__42; size_t lean_usize_sub(size_t, size_t); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instHashableENodeKey_unsafe__1___boxed(lean_object*); lean_object* lean_array_mk(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getParentsAndReset(lean_object*, 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_Grind_isRoot(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -379,67 +469,78 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getRoot_x3f___boxed(lean_object*, lea LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isEqFalse(lean_object*, 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_Grind_pushEqFalse(lean_object*, lean_object*, 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_Grind_congrHash_go___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3___lambda__1___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__17; +LEAN_EXPORT uint64_t l_Lean_Meta_Grind_instHashableENodeKey_unsafe__1(lean_object*); size_t lean_usize_add(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_erase___at_Lean_Meta_Grind_getParentsAndReset___spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instHashableCongrTheoremCacheKey___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instInhabitedNewFact; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__26; static lean_object* l_Lean_Meta_Grind_getENode___closed__1; static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78____closed__9; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getMaxGeneration(lean_object*); lean_object* lean_array_uget(lean_object*, size_t); size_t lean_array_size(lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__51; +static lean_object* l_Lean_Meta_Grind_instInhabitedGoal___closed__4; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pushHEq___boxed(lean_object*, lean_object*, lean_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_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_forEachEqc___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_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at_Lean_Meta_Grind_getENodes___spec__2___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__6(lean_object*, size_t, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__12; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkENode___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hasSameRoot___boxed(lean_object*, lean_object*, lean_object*); size_t lean_usize_shift_left(size_t, size_t); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_MethodsRef_toMethods_unsafe__1___boxed(lean_object*); static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_118____closed__1; -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Grind_getENodes___spec__1___lambda__1(lean_object*, size_t, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__26; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Grind_getENodes___spec__1___lambda__1(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint64_t l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot_unsafe__1(lean_object*); static lean_object* l_Lean_Meta_Grind_congrHash___closed__2; -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__42; -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__27; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkHCongrWithArity___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_RBNode_insert___at_Lean_Meta_Grind_registerParent___spec__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isEqTrue(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_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3___lambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_getEqcs___spec__1(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Meta_Grind_getENodes___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pushEqTrue___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_Grind_instBEqPreInstance___lambda__3(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_canon(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78____closed__5; lean_object* lean_array_get_size(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isFalseExpr___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_ShareCommon_objectFactory; -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__52; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isFalseExpr(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_Grind_getRootENode(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_state_sharecommon(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_mkHCongrWithArity___spec__1(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__28; lean_object* lean_infer_type(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__14; 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_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78____closed__7; -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_setENode___spec__1___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__29; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markTheoremInstance___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint64_t l_Lean_Meta_TransparencyMode_toUInt64(uint8_t); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____boxed(lean_object*, lean_object*); lean_object* lean_nat_add(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_congrPlaceholderProof; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_markTheoremInstance___spec__11___lambda__1___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(lean_object*); -LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_alreadyInternalized___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, size_t); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__1; -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_registerParent___spec__7(lean_object*, size_t, size_t, size_t, lean_object*); +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_markTheoremInstance___spec__11(lean_object*, size_t, lean_object*); +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_alreadyInternalized___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__43; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__4; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__31; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_MethodsRefPointed; -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getENode_x3f_unsafe__1___boxed(lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_registerParent___spec__1(lean_object*, size_t); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389_(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_canon___closed__2; -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_setENode___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_setENode_unsafe__2___rarg___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GoalM_run_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_Lean_Meta_Grind_getConfig___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GoalM_run___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_Lean_Meta_Grind_getFalseExpr___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_uset(lean_object*, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instHashableENodeKey___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getRoot_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT size_t l_Lean_Meta_Grind_instHashableCongrTheoremCacheKey_unsafe__1(lean_object*); lean_object* lean_array_get(lean_object*, lean_object*, lean_object*); @@ -449,19 +550,21 @@ LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_ins LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isInterpreted___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_getEqcs___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_PersistentHashMap_findAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__6___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__53; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_markTheoremInstance___spec__11___lambda__2___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isInconsistent(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_Grind_instInhabitedGoal; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instHashableCongrTheoremCacheKey_unsafe__1___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkENode___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_PersistentHashMap_insertAux___at_Lean_Meta_Grind_setENode___spec__2(lean_object*, size_t, size_t, size_t, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_toENodeKey_unsafe__1___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_setENode___spec__2(lean_object*, size_t, size_t, lean_object*, lean_object*); size_t lean_usize_land(size_t, size_t); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, size_t); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__37; -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__34; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getConfig___boxed(lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__49; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_setENode___spec__3(size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Option_repr___at___private_Lean_Meta_Transform_0__Lean_reprTransformStep____x40_Lean_Meta_Transform___hyg_46____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_mkHCongrWithArity___spec__5___boxed(lean_object*, lean_object*); +LEAN_EXPORT uint64_t l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__2(lean_object*, lean_object*, lean_object*, size_t, size_t, uint64_t); LEAN_EXPORT uint8_t l_Lean_Meta_Grind_isSameExpr_unsafe__1(lean_object* x_1, lean_object* x_2) { _start: { @@ -850,6 +953,49 @@ static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_ return lean_box(0); } } +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getConfig___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_ctor_get(x_1, 3); +lean_inc(x_8); +x_9 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_9, 0, x_8); +lean_ctor_set(x_9, 1, x_7); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getConfig(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_getConfig___rarg___boxed), 7, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getConfig___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_Grind_getConfig___rarg(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getConfig___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Lean_Meta_Grind_getConfig(x_1); +lean_dec(x_1); +return x_2; +} +} LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getTrueExpr___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: { @@ -1047,6 +1193,71 @@ lean_dec(x_2); return x_9; } } +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getMaxGeneration___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; uint8_t x_9; +x_8 = l_Lean_Meta_Grind_getConfig___rarg(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +x_9 = !lean_is_exclusive(x_8); +if (x_9 == 0) +{ +lean_object* x_10; lean_object* x_11; +x_10 = lean_ctor_get(x_8, 0); +x_11 = lean_ctor_get(x_10, 2); +lean_inc(x_11); +lean_dec(x_10); +lean_ctor_set(x_8, 0, x_11); +return x_8; +} +else +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_12 = lean_ctor_get(x_8, 0); +x_13 = lean_ctor_get(x_8, 1); +lean_inc(x_13); +lean_inc(x_12); +lean_dec(x_8); +x_14 = lean_ctor_get(x_12, 2); +lean_inc(x_14); +lean_dec(x_12); +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; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getMaxGeneration(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_getMaxGeneration___rarg___boxed), 7, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getMaxGeneration___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_Grind_getMaxGeneration___rarg(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getMaxGeneration___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Lean_Meta_Grind_getMaxGeneration(x_1); +lean_dec(x_1); +return x_2; +} +} static lean_object* _init_l_Lean_Meta_Grind_abstractNestedProofs___closed__1() { _start: { @@ -1389,9 +1600,10 @@ static lean_object* _init_l_Lean_Meta_Grind_canon___closed__3() { { lean_object* x_1; lean_object* x_2; x_1 = l_Lean_Meta_Grind_canon___closed__2; -x_2 = lean_alloc_ctor(0, 2, 0); +x_2 = lean_alloc_ctor(0, 3, 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); return x_2; } } @@ -1679,6 +1891,106 @@ lean_dec(x_2); return x_10; } } +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isTrueExpr(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; uint8_t x_11; +x_10 = l_Lean_Meta_Grind_getTrueExpr___rarg(x_4, x_5, x_6, x_7, x_8, x_9); +x_11 = !lean_is_exclusive(x_10); +if (x_11 == 0) +{ +lean_object* x_12; uint8_t x_13; lean_object* x_14; +x_12 = lean_ctor_get(x_10, 0); +x_13 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_1, x_12); +lean_dec(x_12); +x_14 = lean_box(x_13); +lean_ctor_set(x_10, 0, x_14); +return x_10; +} +else +{ +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_10, 0); +x_16 = lean_ctor_get(x_10, 1); +lean_inc(x_16); +lean_inc(x_15); +lean_dec(x_10); +x_17 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_1, x_15); +lean_dec(x_15); +x_18 = lean_box(x_17); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_16); +return x_19; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isTrueExpr___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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_Grind_isTrueExpr(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); +lean_dec(x_1); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isFalseExpr(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; uint8_t x_11; +x_10 = l_Lean_Meta_Grind_getFalseExpr___rarg(x_4, x_5, x_6, x_7, x_8, x_9); +x_11 = !lean_is_exclusive(x_10); +if (x_11 == 0) +{ +lean_object* x_12; uint8_t x_13; lean_object* x_14; +x_12 = lean_ctor_get(x_10, 0); +x_13 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_1, x_12); +lean_dec(x_12); +x_14 = lean_box(x_13); +lean_ctor_set(x_10, 0, x_14); +return x_10; +} +else +{ +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_10, 0); +x_16 = lean_ctor_get(x_10, 1); +lean_inc(x_16); +lean_inc(x_15); +lean_dec(x_10); +x_17 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_1, x_15); +lean_dec(x_15); +x_18 = lean_box(x_17); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_16); +return x_19; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isFalseExpr___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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_Grind_isFalseExpr(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); +lean_dec(x_1); +return x_10; +} +} LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_mkHCongrWithArity___spec__3(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: { @@ -3083,7 +3395,7 @@ x_1 = l_Lean_Meta_Grind_instInhabitedENode___closed__2; return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__1() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__1() { _start: { lean_object* x_1; @@ -3091,29 +3403,29 @@ x_1 = lean_mk_string_unchecked("self", 4, 4); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__2() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__1; +x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__1; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__3() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__2; +x_2 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__2; x_3 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__4() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__4() { _start: { lean_object* x_1; @@ -3121,29 +3433,29 @@ x_1 = lean_mk_string_unchecked(" := ", 4, 4); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__5() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__5() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__4; +x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__4; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__6() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__3; -x_2 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__5; +x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__3; +x_2 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____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_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__7() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__7() { _start: { lean_object* x_1; lean_object* x_2; @@ -3152,7 +3464,7 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__8() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__8() { _start: { lean_object* x_1; @@ -3160,17 +3472,17 @@ x_1 = lean_mk_string_unchecked(",", 1, 1); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__9() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__9() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__8; +x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____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_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__10() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__10() { _start: { lean_object* x_1; @@ -3178,17 +3490,17 @@ x_1 = lean_mk_string_unchecked("next", 4, 4); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__11() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__11() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__10; +x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____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_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__12() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__12() { _start: { lean_object* x_1; @@ -3196,17 +3508,17 @@ x_1 = lean_mk_string_unchecked("root", 4, 4); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__13() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__13() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__12; +x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____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_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__14() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__14() { _start: { lean_object* x_1; @@ -3214,17 +3526,17 @@ x_1 = lean_mk_string_unchecked("cgRoot", 6, 6); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__15() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__15() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__14; +x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__14; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__16() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__16() { _start: { lean_object* x_1; lean_object* x_2; @@ -3233,7 +3545,7 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__17() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__17() { _start: { lean_object* x_1; @@ -3241,17 +3553,17 @@ x_1 = lean_mk_string_unchecked("target\?", 7, 7); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__18() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__18() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__17; +x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____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_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__19() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__19() { _start: { lean_object* x_1; lean_object* x_2; @@ -3260,7 +3572,7 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__20() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__20() { _start: { lean_object* x_1; @@ -3268,17 +3580,17 @@ x_1 = lean_mk_string_unchecked("proof\?", 6, 6); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__21() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__21() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__20; +x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____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_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__22() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__22() { _start: { lean_object* x_1; @@ -3286,17 +3598,17 @@ x_1 = lean_mk_string_unchecked("flipped", 7, 7); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__23() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__23() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__22; +x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__22; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__24() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__24() { _start: { lean_object* x_1; @@ -3304,17 +3616,17 @@ x_1 = lean_mk_string_unchecked("size", 4, 4); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__25() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__25() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__24; +x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__24; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__26() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__26() { _start: { lean_object* x_1; @@ -3322,17 +3634,17 @@ x_1 = lean_mk_string_unchecked("interpreted", 11, 11); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__27() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__27() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__26; +x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____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_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__28() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__28() { _start: { lean_object* x_1; lean_object* x_2; @@ -3341,7 +3653,7 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__29() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__29() { _start: { lean_object* x_1; @@ -3349,17 +3661,17 @@ x_1 = lean_mk_string_unchecked("ctor", 4, 4); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__30() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__30() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__29; +x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____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_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__31() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__31() { _start: { lean_object* x_1; @@ -3367,17 +3679,17 @@ x_1 = lean_mk_string_unchecked("hasLambdas", 10, 10); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__32() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__32() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__31; +x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____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_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__33() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__33() { _start: { lean_object* x_1; lean_object* x_2; @@ -3386,7 +3698,7 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__34() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__34() { _start: { lean_object* x_1; @@ -3394,17 +3706,17 @@ x_1 = lean_mk_string_unchecked("heqProofs", 9, 9); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__35() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__35() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__34; +x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__34; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__36() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__36() { _start: { lean_object* x_1; lean_object* x_2; @@ -3413,7 +3725,7 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__37() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__37() { _start: { lean_object* x_1; @@ -3421,17 +3733,17 @@ x_1 = lean_mk_string_unchecked("idx", 3, 3); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__38() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__38() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__37; +x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____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_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__39() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__39() { _start: { lean_object* x_1; lean_object* x_2; @@ -3440,7 +3752,7 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__40() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__40() { _start: { lean_object* x_1; @@ -3448,17 +3760,17 @@ x_1 = lean_mk_string_unchecked("generation", 10, 10); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__41() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__41() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__40; +x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__40; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__42() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__42() { _start: { lean_object* x_1; @@ -3466,17 +3778,17 @@ x_1 = lean_mk_string_unchecked("mt", 2, 2); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__43() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__43() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__42; +x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__42; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__44() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__44() { _start: { lean_object* x_1; lean_object* x_2; @@ -3485,7 +3797,7 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__45() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__45() { _start: { lean_object* x_1; @@ -3493,35 +3805,35 @@ x_1 = lean_mk_string_unchecked("{ ", 2, 2); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__46() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__46() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__45; +x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__45; x_2 = lean_string_length(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__47() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__47() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__46; +x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__46; x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__48() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__48() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__45; +x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____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_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__49() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__49() { _start: { lean_object* x_1; @@ -3529,17 +3841,17 @@ x_1 = lean_mk_string_unchecked(" }", 2, 2); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__50() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__50() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__49; +x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__49; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__51() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__51() { _start: { lean_object* x_1; @@ -3547,17 +3859,17 @@ x_1 = lean_mk_string_unchecked("false", 5, 5); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__52() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__52() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__51; +x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____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_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__53() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__53() { _start: { lean_object* x_1; @@ -3565,17 +3877,17 @@ x_1 = lean_mk_string_unchecked("true", 4, 4); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__54() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__54() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__53; +x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__53; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532_(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; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; 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; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* 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; 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; @@ -3583,7 +3895,7 @@ x_3 = lean_ctor_get(x_1, 0); lean_inc(x_3); x_4 = lean_unsigned_to_nat(0u); x_5 = l___private_Lean_Expr_0__Lean_reprExpr____x40_Lean_Expr___hyg_3036_(x_3, x_4); -x_6 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__7; +x_6 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____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); @@ -3591,11 +3903,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_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__6; +x_10 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____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_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__9; +x_12 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____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); @@ -3603,11 +3915,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_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__11; +x_16 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____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_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__5; +x_18 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____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); @@ -3629,7 +3941,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_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__13; +x_27 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__13; x_28 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_28, 0, x_26); lean_ctor_set(x_28, 1, x_27); @@ -3654,7 +3966,7 @@ lean_ctor_set(x_35, 1, x_12); x_36 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_36, 0, x_35); lean_ctor_set(x_36, 1, x_14); -x_37 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__15; +x_37 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__15; x_38 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_38, 0, x_36); lean_ctor_set(x_38, 1, x_37); @@ -3664,7 +3976,7 @@ lean_ctor_set(x_39, 1, x_18); x_40 = lean_ctor_get(x_1, 3); lean_inc(x_40); x_41 = l___private_Lean_Expr_0__Lean_reprExpr____x40_Lean_Expr___hyg_3036_(x_40, x_4); -x_42 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__16; +x_42 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__16; x_43 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_43, 0, x_42); lean_ctor_set(x_43, 1, x_41); @@ -3680,7 +3992,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_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__18; +x_48 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____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); @@ -3690,7 +4002,7 @@ lean_ctor_set(x_50, 1, x_18); x_51 = lean_ctor_get(x_1, 4); lean_inc(x_51); x_52 = l_Option_repr___at___private_Lean_Meta_Transform_0__Lean_reprTransformStep____x40_Lean_Meta_Transform___hyg_46____spec__1(x_51, x_4); -x_53 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__19; +x_53 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____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); @@ -3706,7 +4018,7 @@ lean_ctor_set(x_57, 1, x_12); x_58 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_58, 0, x_57); lean_ctor_set(x_58, 1, x_14); -x_59 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__21; +x_59 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__21; x_60 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_60, 0, x_58); lean_ctor_set(x_60, 1, x_59); @@ -3731,7 +4043,7 @@ lean_ctor_set(x_67, 1, x_12); x_68 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_68, 0, x_67); lean_ctor_set(x_68, 1, x_14); -x_69 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__23; +x_69 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__23; x_70 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_70, 0, x_68); lean_ctor_set(x_70, 1, x_69); @@ -3759,7 +4071,7 @@ lean_inc(x_82); x_83 = l___private_Init_Data_Repr_0__Nat_reprFast(x_82); x_84 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_84, 0, x_83); -x_85 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__39; +x_85 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__39; x_86 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_86, 0, x_85); lean_ctor_set(x_86, 1, x_84); @@ -3771,7 +4083,7 @@ lean_inc(x_88); x_89 = l___private_Init_Data_Repr_0__Nat_reprFast(x_88); x_90 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_90, 0, x_89); -x_91 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__33; +x_91 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__33; x_92 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_92, 0, x_91); lean_ctor_set(x_92, 1, x_90); @@ -3784,7 +4096,7 @@ lean_dec(x_1); x_95 = l___private_Init_Data_Repr_0__Nat_reprFast(x_94); x_96 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_96, 0, x_95); -x_97 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__44; +x_97 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__44; x_98 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_98, 0, x_97); lean_ctor_set(x_98, 1, x_96); @@ -3794,14 +4106,14 @@ lean_ctor_set_uint8(x_99, sizeof(void*)*1, x_8); if (x_72 == 0) { lean_object* x_186; -x_186 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__52; +x_186 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__52; x_100 = x_186; goto block_185; } else { lean_object* x_187; -x_187 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__54; +x_187 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__54; x_100 = x_187; goto block_185; } @@ -3823,7 +4135,7 @@ lean_ctor_set(x_104, 1, x_12); x_105 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_105, 0, x_104); lean_ctor_set(x_105, 1, x_14); -x_106 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__25; +x_106 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__25; x_107 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_107, 0, x_105); lean_ctor_set(x_107, 1, x_106); @@ -3839,7 +4151,7 @@ lean_ctor_set(x_110, 1, x_12); x_111 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_111, 0, x_110); lean_ctor_set(x_111, 1, x_14); -x_112 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__27; +x_112 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__27; x_113 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_113, 0, x_111); lean_ctor_set(x_113, 1, x_112); @@ -3849,21 +4161,21 @@ lean_ctor_set(x_114, 1, x_18); if (x_78 == 0) { lean_object* x_183; -x_183 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__52; +x_183 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__52; x_115 = x_183; goto block_182; } else { lean_object* x_184; -x_184 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__54; +x_184 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__54; x_115 = x_184; goto block_182; } block_182: { 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_116 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__28; +x_116 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__28; x_117 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_117, 0, x_116); lean_ctor_set(x_117, 1, x_115); @@ -3879,7 +4191,7 @@ lean_ctor_set(x_120, 1, x_12); x_121 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_121, 0, x_120); lean_ctor_set(x_121, 1, x_14); -x_122 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__30; +x_122 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__30; x_123 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_123, 0, x_121); lean_ctor_set(x_123, 1, x_122); @@ -3889,14 +4201,14 @@ lean_ctor_set(x_124, 1, x_18); if (x_79 == 0) { lean_object* x_180; -x_180 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__52; +x_180 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__52; x_125 = x_180; goto block_179; } else { lean_object* x_181; -x_181 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__54; +x_181 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__54; x_125 = x_181; goto block_179; } @@ -3918,7 +4230,7 @@ lean_ctor_set(x_129, 1, x_12); x_130 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_130, 0, x_129); lean_ctor_set(x_130, 1, x_14); -x_131 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__32; +x_131 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__32; x_132 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_132, 0, x_130); lean_ctor_set(x_132, 1, x_131); @@ -3928,14 +4240,14 @@ lean_ctor_set(x_133, 1, x_18); if (x_80 == 0) { lean_object* x_177; -x_177 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__52; +x_177 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__52; x_134 = x_177; goto block_176; } else { lean_object* x_178; -x_178 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__54; +x_178 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__54; x_134 = x_178; goto block_176; } @@ -3957,7 +4269,7 @@ lean_ctor_set(x_138, 1, x_12); x_139 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_139, 0, x_138); lean_ctor_set(x_139, 1, x_14); -x_140 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__35; +x_140 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__35; x_141 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_141, 0, x_139); lean_ctor_set(x_141, 1, x_140); @@ -3967,21 +4279,21 @@ lean_ctor_set(x_142, 1, x_18); if (x_81 == 0) { lean_object* x_174; -x_174 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__52; +x_174 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__52; x_143 = x_174; goto block_173; } else { lean_object* x_175; -x_175 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__54; +x_175 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__54; x_143 = x_175; goto block_173; } block_173: { lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; -x_144 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__36; +x_144 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__36; x_145 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_145, 0, x_144); lean_ctor_set(x_145, 1, x_143); @@ -3997,7 +4309,7 @@ lean_ctor_set(x_148, 1, x_12); x_149 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_149, 0, x_148); lean_ctor_set(x_149, 1, x_14); -x_150 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__38; +x_150 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__38; x_151 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_151, 0, x_149); lean_ctor_set(x_151, 1, x_150); @@ -4013,7 +4325,7 @@ lean_ctor_set(x_154, 1, x_12); x_155 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_155, 0, x_154); lean_ctor_set(x_155, 1, x_14); -x_156 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__41; +x_156 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__41; x_157 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_157, 0, x_155); lean_ctor_set(x_157, 1, x_156); @@ -4029,7 +4341,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_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__43; +x_162 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__43; x_163 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_163, 0, x_161); lean_ctor_set(x_163, 1, x_162); @@ -4039,15 +4351,15 @@ lean_ctor_set(x_164, 1, x_18); x_165 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_165, 0, x_164); lean_ctor_set(x_165, 1, x_99); -x_166 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__48; +x_166 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__48; x_167 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_167, 0, x_166); lean_ctor_set(x_167, 1, x_165); -x_168 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__50; +x_168 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__50; x_169 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_169, 0, x_167); lean_ctor_set(x_169, 1, x_168); -x_170 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__47; +x_170 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__47; x_171 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_171, 0, x_170); lean_ctor_set(x_171, 1, x_169); @@ -4062,11 +4374,11 @@ return x_172; } } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389_(x_1, x_2); +x_3 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532_(x_1, x_2); lean_dec(x_2); return x_3; } @@ -4075,7 +4387,7 @@ static lean_object* _init_l_Lean_Meta_Grind_instReprENode___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____boxed), 2, 0); return x_1; } } @@ -4087,43 +4399,83 @@ x_1 = l_Lean_Meta_Grind_instReprENode___closed__1; return x_1; } } -LEAN_EXPORT size_t l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_toENodeKey_unsafe__1(lean_object* x_1) { +LEAN_EXPORT uint64_t l_Lean_Meta_Grind_instHashableENodeKey_unsafe__1(lean_object* x_1) { _start: { -size_t x_2; +size_t x_2; uint64_t x_3; x_2 = lean_ptr_addr(x_1); -return x_2; +x_3 = lean_usize_to_uint64(x_2); +return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_toENodeKey_unsafe__1___boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instHashableENodeKey_unsafe__1___boxed(lean_object* x_1) { _start: { -size_t x_2; lean_object* x_3; -x_2 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_toENodeKey_unsafe__1(x_1); +uint64_t x_2; lean_object* x_3; +x_2 = l_Lean_Meta_Grind_instHashableENodeKey_unsafe__1(x_1); lean_dec(x_1); -x_3 = lean_box_usize(x_2); +x_3 = lean_box_uint64(x_2); return x_3; } } -LEAN_EXPORT size_t l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_toENodeKey(lean_object* x_1) { +LEAN_EXPORT uint64_t l_Lean_Meta_Grind_instHashableENodeKey(lean_object* x_1) { _start: { -size_t x_2; -x_2 = lean_ptr_addr(x_1); +uint64_t x_2; +x_2 = l_Lean_Meta_Grind_instHashableENodeKey_unsafe__1(x_1); return x_2; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_toENodeKey___boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instHashableENodeKey___boxed(lean_object* x_1) { _start: { -size_t x_2; lean_object* x_3; -x_2 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_toENodeKey(x_1); +uint64_t x_2; lean_object* x_3; +x_2 = l_Lean_Meta_Grind_instHashableENodeKey(x_1); lean_dec(x_1); -x_3 = lean_box_usize(x_2); +x_3 = lean_box_uint64(x_2); +return x_3; +} +} +LEAN_EXPORT uint8_t l_Lean_Meta_Grind_instBEqENodeKey(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instBEqENodeKey___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Lean_Meta_Grind_instBEqENodeKey(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT uint64_t l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot_unsafe__1(lean_object* x_1) { +_start: +{ +lean_object* x_2; size_t x_3; uint64_t x_4; +x_2 = lean_ctor_get(x_1, 2); +x_3 = lean_ptr_addr(x_2); +x_4 = lean_usize_to_uint64(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot_unsafe__1___boxed(lean_object* x_1) { +_start: +{ +uint64_t x_2; lean_object* x_3; +x_2 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot_unsafe__1(x_1); +lean_dec(x_1); +x_3 = lean_box_uint64(x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot___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; @@ -4139,34 +4491,33 @@ return x_8; } else { -lean_object* x_9; size_t x_10; uint8_t x_11; +lean_object* x_9; uint8_t x_10; x_9 = lean_array_fget(x_1, x_4); -x_10 = lean_unbox_usize(x_9); +x_10 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_5, x_9); lean_dec(x_9); -x_11 = lean_usize_dec_eq(x_5, x_10); -if (x_11 == 0) +if (x_10 == 0) { -lean_object* x_12; lean_object* x_13; -x_12 = lean_unsigned_to_nat(1u); -x_13 = lean_nat_add(x_4, x_12); +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_13; +x_4 = x_12; goto _start; } else { -lean_object* x_15; lean_object* x_16; -x_15 = lean_array_fget(x_2, x_4); +lean_object* x_14; lean_object* x_15; +x_14 = lean_array_fget(x_2, x_4); lean_dec(x_4); -x_16 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_16, 0, x_15); -return x_16; +x_15 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_15, 0, x_14); +return x_15; } } } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot___spec__2(lean_object* x_1, size_t x_2, size_t x_3) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot___spec__2(lean_object* x_1, size_t x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_1) == 0) @@ -4188,22 +4539,21 @@ lean_dec(x_5); switch (lean_obj_tag(x_11)) { case 0: { -lean_object* x_12; lean_object* x_13; size_t x_14; uint8_t x_15; +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_unbox_usize(x_12); +x_14 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_3, x_12); lean_dec(x_12); -x_15 = lean_usize_dec_eq(x_3, x_14); -if (x_15 == 0) +if (x_14 == 0) { -lean_object* x_16; +lean_object* x_15; lean_dec(x_13); lean_free_object(x_1); -x_16 = lean_box(0); -return x_16; +x_15 = lean_box(0); +return x_15; } else { @@ -4214,107 +4564,106 @@ return x_1; } case 1: { -lean_object* x_17; size_t x_18; +lean_object* x_16; size_t x_17; lean_free_object(x_1); -x_17 = lean_ctor_get(x_11, 0); -lean_inc(x_17); +x_16 = lean_ctor_get(x_11, 0); +lean_inc(x_16); lean_dec(x_11); -x_18 = lean_usize_shift_right(x_2, x_6); -x_1 = x_17; -x_2 = x_18; +x_17 = lean_usize_shift_right(x_2, x_6); +x_1 = x_16; +x_2 = x_17; goto _start; } default: { -lean_object* x_20; +lean_object* x_19; lean_free_object(x_1); -x_20 = lean_box(0); -return x_20; +x_19 = lean_box(0); +return x_19; } } } else { -lean_object* x_21; size_t x_22; size_t x_23; size_t x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_21 = lean_ctor_get(x_1, 0); -lean_inc(x_21); -lean_dec(x_1); -x_22 = 5; -x_23 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__2___closed__2; -x_24 = lean_usize_land(x_2, x_23); -x_25 = lean_usize_to_nat(x_24); -x_26 = lean_box(2); -x_27 = lean_array_get(x_26, x_21, x_25); -lean_dec(x_25); -lean_dec(x_21); -switch (lean_obj_tag(x_27)) { +lean_object* x_20; size_t x_21; size_t x_22; size_t x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_20 = lean_ctor_get(x_1, 0); +lean_inc(x_20); +lean_dec(x_1); +x_21 = 5; +x_22 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__2___closed__2; +x_23 = lean_usize_land(x_2, x_22); +x_24 = lean_usize_to_nat(x_23); +x_25 = lean_box(2); +x_26 = lean_array_get(x_25, x_20, x_24); +lean_dec(x_24); +lean_dec(x_20); +switch (lean_obj_tag(x_26)) { case 0: { -lean_object* x_28; lean_object* x_29; size_t x_30; uint8_t x_31; -x_28 = lean_ctor_get(x_27, 0); +lean_object* x_27; lean_object* x_28; uint8_t x_29; +x_27 = lean_ctor_get(x_26, 0); +lean_inc(x_27); +x_28 = lean_ctor_get(x_26, 1); lean_inc(x_28); -x_29 = lean_ctor_get(x_27, 1); -lean_inc(x_29); +lean_dec(x_26); +x_29 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_3, x_27); lean_dec(x_27); -x_30 = lean_unbox_usize(x_28); -lean_dec(x_28); -x_31 = lean_usize_dec_eq(x_3, x_30); -if (x_31 == 0) +if (x_29 == 0) { -lean_object* x_32; -lean_dec(x_29); -x_32 = lean_box(0); -return x_32; +lean_object* x_30; +lean_dec(x_28); +x_30 = lean_box(0); +return x_30; } else { -lean_object* x_33; -x_33 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_33, 0, x_29); -return x_33; +lean_object* x_31; +x_31 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_31, 0, x_28); +return x_31; } } case 1: { -lean_object* x_34; size_t x_35; -x_34 = lean_ctor_get(x_27, 0); -lean_inc(x_34); -lean_dec(x_27); -x_35 = lean_usize_shift_right(x_2, x_22); -x_1 = x_34; -x_2 = x_35; +lean_object* x_32; size_t x_33; +x_32 = lean_ctor_get(x_26, 0); +lean_inc(x_32); +lean_dec(x_26); +x_33 = lean_usize_shift_right(x_2, x_21); +x_1 = x_32; +x_2 = x_33; goto _start; } default: { -lean_object* x_37; -x_37 = lean_box(0); -return x_37; +lean_object* x_35; +x_35 = lean_box(0); +return x_35; } } } } else { -lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; -x_38 = lean_ctor_get(x_1, 0); -lean_inc(x_38); -x_39 = lean_ctor_get(x_1, 1); -lean_inc(x_39); +lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_36 = lean_ctor_get(x_1, 0); +lean_inc(x_36); +x_37 = lean_ctor_get(x_1, 1); +lean_inc(x_37); lean_dec(x_1); -x_40 = lean_unsigned_to_nat(0u); -x_41 = l_Lean_PersistentHashMap_findAtAux___at___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot___spec__3(x_38, x_39, lean_box(0), x_40, x_3); -lean_dec(x_39); -lean_dec(x_38); -return x_41; +x_38 = lean_unsigned_to_nat(0u); +x_39 = l_Lean_PersistentHashMap_findAtAux___at___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot___spec__3(x_36, x_37, lean_box(0), x_38, x_3); +lean_dec(x_37); +lean_dec(x_36); +return x_39; } } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot___spec__1(lean_object* x_1, size_t x_2) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot___spec__1(lean_object* x_1, lean_object* x_2) { _start: { uint64_t x_3; size_t x_4; lean_object* x_5; -x_3 = lean_usize_to_uint64(x_2); +x_3 = l_Lean_Meta_Grind_instHashableENodeKey_unsafe__1(x_2); x_4 = lean_uint64_to_usize(x_3); x_5 = l_Lean_PersistentHashMap_findAux___at___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot___spec__2(x_1, x_4, x_2); return x_5; @@ -4323,63 +4672,55 @@ return x_5; LEAN_EXPORT uint64_t l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot(lean_object* x_1, lean_object* x_2) { _start: { -size_t x_3; lean_object* x_4; -x_3 = lean_ptr_addr(x_2); -x_4 = l_Lean_PersistentHashMap_find_x3f___at___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot___spec__1(x_1, x_3); -if (lean_obj_tag(x_4) == 0) +lean_object* x_3; +x_3 = l_Lean_PersistentHashMap_find_x3f___at___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot___spec__1(x_1, x_2); +if (lean_obj_tag(x_3) == 0) { -uint64_t x_5; -x_5 = 13; -return x_5; +uint64_t x_4; +x_4 = 13; +return x_4; } else { -lean_object* x_6; lean_object* x_7; size_t x_8; uint64_t x_9; -x_6 = lean_ctor_get(x_4, 0); -lean_inc(x_6); -lean_dec(x_4); -x_7 = lean_ctor_get(x_6, 2); -lean_inc(x_7); -lean_dec(x_6); -x_8 = lean_ptr_addr(x_7); -lean_dec(x_7); -x_9 = lean_usize_to_uint64(x_8); -return x_9; +lean_object* x_5; uint64_t x_6; +x_5 = lean_ctor_get(x_3, 0); +lean_inc(x_5); +lean_dec(x_3); +x_6 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot_unsafe__1(x_5); +lean_dec(x_5); +return x_6; } } } LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -size_t x_6; lean_object* x_7; -x_6 = lean_unbox_usize(x_5); +lean_object* x_6; +x_6 = l_Lean_PersistentHashMap_findAtAux___at___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot___spec__3(x_1, x_2, x_3, x_4, x_5); lean_dec(x_5); -x_7 = l_Lean_PersistentHashMap_findAtAux___at___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot___spec__3(x_1, x_2, x_3, x_4, x_6); lean_dec(x_2); lean_dec(x_1); -return x_7; +return x_6; } } LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot___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; +size_t x_4; lean_object* x_5; x_4 = lean_unbox_usize(x_2); lean_dec(x_2); -x_5 = lean_unbox_usize(x_3); +x_5 = l_Lean_PersistentHashMap_findAux___at___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot___spec__2(x_1, x_4, x_3); lean_dec(x_3); -x_6 = l_Lean_PersistentHashMap_findAux___at___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot___spec__2(x_1, x_4, x_5); -return x_6; +return x_5; } } LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot___spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { -size_t x_3; lean_object* x_4; -x_3 = lean_unbox_usize(x_2); +lean_object* x_3; +x_3 = l_Lean_PersistentHashMap_find_x3f___at___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot___spec__1(x_1, x_2); lean_dec(x_2); -x_4 = l_Lean_PersistentHashMap_find_x3f___at___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot___spec__1(x_1, x_3); -return x_4; +return x_3; } } LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot___boxed(lean_object* x_1, lean_object* x_2) { @@ -4395,63 +4736,59 @@ return x_4; LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hasSameRoot(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; -x_4 = lean_ptr_addr(x_2); -x_5 = lean_ptr_addr(x_3); -x_6 = lean_usize_dec_eq(x_4, x_5); -if (x_6 == 0) +uint8_t x_4; +x_4 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_2, x_3); +if (x_4 == 0) { -lean_object* x_7; +lean_object* x_5; lean_inc(x_1); -x_7 = l_Lean_PersistentHashMap_find_x3f___at___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot___spec__1(x_1, x_4); -if (lean_obj_tag(x_7) == 0) +x_5 = l_Lean_PersistentHashMap_find_x3f___at___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot___spec__1(x_1, x_2); +if (lean_obj_tag(x_5) == 0) { -uint8_t x_8; +uint8_t x_6; lean_dec(x_1); -x_8 = 0; -return x_8; +x_6 = 0; +return x_6; } else { -lean_object* x_9; lean_object* x_10; -x_9 = lean_ctor_get(x_7, 0); -lean_inc(x_9); -lean_dec(x_7); -x_10 = l_Lean_PersistentHashMap_find_x3f___at___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot___spec__1(x_1, x_5); -if (lean_obj_tag(x_10) == 0) +lean_object* x_7; lean_object* x_8; +x_7 = lean_ctor_get(x_5, 0); +lean_inc(x_7); +lean_dec(x_5); +x_8 = l_Lean_PersistentHashMap_find_x3f___at___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot___spec__1(x_1, x_3); +if (lean_obj_tag(x_8) == 0) { -uint8_t x_11; -lean_dec(x_9); -x_11 = 0; -return x_11; +uint8_t x_9; +lean_dec(x_7); +x_9 = 0; +return x_9; } else { -lean_object* x_12; lean_object* x_13; size_t x_14; lean_object* x_15; size_t x_16; uint8_t x_17; -x_12 = lean_ctor_get(x_10, 0); +lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_10 = lean_ctor_get(x_8, 0); +lean_inc(x_10); +lean_dec(x_8); +x_11 = lean_ctor_get(x_7, 2); +lean_inc(x_11); +lean_dec(x_7); +x_12 = lean_ctor_get(x_10, 2); lean_inc(x_12); lean_dec(x_10); -x_13 = lean_ctor_get(x_9, 2); -lean_inc(x_13); -lean_dec(x_9); -x_14 = lean_ptr_addr(x_13); -lean_dec(x_13); -x_15 = lean_ctor_get(x_12, 2); -lean_inc(x_15); +x_13 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_11, x_12); lean_dec(x_12); -x_16 = lean_ptr_addr(x_15); -lean_dec(x_15); -x_17 = lean_usize_dec_eq(x_14, x_16); -return x_17; +lean_dec(x_11); +return x_13; } } } else { -uint8_t x_18; +uint8_t x_14; lean_dec(x_1); -x_18 = 1; -return x_18; +x_14 = 1; +return x_14; } } } @@ -4723,211 +5060,3812 @@ x_5 = lean_box(x_4); return x_5; } } -static lean_object* _init_l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedGoal___spec__1() { +LEAN_EXPORT uint64_t l_Lean_Meta_Grind_instHashablePreInstance_unsafe__1(lean_object* x_1) { _start: { -lean_object* x_1; -x_1 = l_Lean_Meta_Grind_canon___closed__2; -return x_1; +lean_object* x_2; size_t x_3; size_t x_4; size_t x_5; uint64_t x_6; +x_2 = lean_ctor_get(x_1, 0); +x_3 = lean_ptr_addr(x_2); +x_4 = 3; +x_5 = lean_usize_shift_right(x_3, x_4); +x_6 = lean_usize_to_uint64(x_5); +return x_6; } } -static lean_object* _init_l_Lean_Meta_Grind_instInhabitedGoal___closed__1() { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instHashablePreInstance_unsafe__1___boxed(lean_object* x_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; +uint64_t x_2; lean_object* x_3; +x_2 = l_Lean_Meta_Grind_instHashablePreInstance_unsafe__1(x_1); +lean_dec(x_1); +x_3 = lean_box_uint64(x_2); +return x_3; } } -static lean_object* _init_l_Lean_Meta_Grind_instInhabitedGoal() { +LEAN_EXPORT uint64_t l_Lean_Meta_Grind_instHashablePreInstance_unsafe__2(lean_object* x_1) { _start: { -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; -x_1 = lean_box(0); -x_2 = l_Lean_Meta_Grind_canon___closed__2; -x_3 = l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedGoal___spec__1; -x_4 = l_Lean_Meta_Grind_instInhabitedGoal___closed__1; -x_5 = 0; -x_6 = lean_unsigned_to_nat(0u); -x_7 = lean_alloc_ctor(0, 8, 1); -lean_ctor_set(x_7, 0, x_1); -lean_ctor_set(x_7, 1, x_2); -lean_ctor_set(x_7, 2, x_2); -lean_ctor_set(x_7, 3, x_3); -lean_ctor_set(x_7, 4, x_2); -lean_ctor_set(x_7, 5, x_4); -lean_ctor_set(x_7, 6, x_6); -lean_ctor_set(x_7, 7, x_6); -lean_ctor_set_uint8(x_7, sizeof(void*)*8, x_5); -return x_7; +size_t x_2; size_t x_3; size_t x_4; uint64_t x_5; +x_2 = lean_ptr_addr(x_1); +x_3 = 3; +x_4 = lean_usize_shift_right(x_2, x_3); +x_5 = lean_usize_to_uint64(x_4); +return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Goal_admit(lean_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_Grind_instHashablePreInstance_unsafe__2___boxed(lean_object* x_1) { _start: { -lean_object* x_7; uint8_t x_8; lean_object* x_9; -x_7 = lean_ctor_get(x_1, 0); -lean_inc(x_7); +uint64_t x_2; lean_object* x_3; +x_2 = l_Lean_Meta_Grind_instHashablePreInstance_unsafe__2(x_1); lean_dec(x_1); -x_8 = 1; -x_9 = l_Lean_MVarId_admit(x_7, x_8, x_2, x_3, x_4, x_5, x_6); -return x_9; +x_3 = lean_box_uint64(x_2); +return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isTrueExpr(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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 uint64_t l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instHashablePreInstance___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, uint64_t x_6) { _start: { -lean_object* x_10; uint8_t x_11; -x_10 = l_Lean_Meta_Grind_getTrueExpr___rarg(x_4, x_5, x_6, x_7, x_8, x_9); -x_11 = !lean_is_exclusive(x_10); -if (x_11 == 0) +uint8_t x_7; +x_7 = lean_usize_dec_lt(x_5, x_4); +if (x_7 == 0) { -lean_object* x_12; uint8_t x_13; lean_object* x_14; -x_12 = lean_ctor_get(x_10, 0); -x_13 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_1, x_12); -lean_dec(x_12); -x_14 = lean_box(x_13); -lean_ctor_set(x_10, 0, x_14); -return x_10; +return x_6; } else { -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_10, 0); -x_16 = lean_ctor_get(x_10, 1); -lean_inc(x_16); -lean_inc(x_15); -lean_dec(x_10); -x_17 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_1, x_15); -lean_dec(x_15); -x_18 = lean_box(x_17); -x_19 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_19, 0, x_18); -lean_ctor_set(x_19, 1, x_16); -return x_19; -} -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isTrueExpr___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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_Grind_isTrueExpr(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_object* x_8; uint64_t x_9; uint64_t x_10; size_t x_11; size_t x_12; +x_8 = lean_array_uget(x_3, x_5); +x_9 = l_Lean_Meta_Grind_instHashablePreInstance_unsafe__2(x_8); 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_10; +x_10 = lean_uint64_mix_hash(x_6, x_9); +x_11 = 1; +x_12 = lean_usize_add(x_5, x_11); +x_5 = x_12; +x_6 = x_10; +goto _start; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isFalseExpr(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -lean_object* x_10; uint8_t x_11; -x_10 = l_Lean_Meta_Grind_getFalseExpr___rarg(x_4, x_5, x_6, x_7, x_8, x_9); -x_11 = !lean_is_exclusive(x_10); -if (x_11 == 0) -{ -lean_object* x_12; uint8_t x_13; lean_object* x_14; -x_12 = lean_ctor_get(x_10, 0); -x_13 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_1, x_12); -lean_dec(x_12); -x_14 = lean_box(x_13); -lean_ctor_set(x_10, 0, x_14); -return x_10; } -else +LEAN_EXPORT uint64_t l_Lean_Meta_Grind_instHashablePreInstance(lean_object* x_1) { +_start: { -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_10, 0); -x_16 = lean_ctor_get(x_10, 1); -lean_inc(x_16); -lean_inc(x_15); -lean_dec(x_10); -x_17 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_1, x_15); -lean_dec(x_15); -x_18 = lean_box(x_17); -x_19 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_19, 0, x_18); -lean_ctor_set(x_19, 1, x_16); -return x_19; -} +uint64_t x_2; lean_object* x_3; lean_object* x_4; size_t x_5; size_t x_6; uint64_t x_7; +x_2 = l_Lean_Meta_Grind_instHashablePreInstance_unsafe__1(x_1); +x_3 = lean_ctor_get(x_1, 1); +x_4 = lean_box(0); +x_5 = lean_array_size(x_3); +x_6 = 0; +x_7 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instHashablePreInstance___spec__1(x_3, x_4, x_3, x_5, x_6, x_2); +return x_7; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isFalseExpr___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instHashablePreInstance___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_10; -x_10 = l_Lean_Meta_Grind_isFalseExpr(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); +size_t x_7; size_t x_8; uint64_t x_9; uint64_t x_10; lean_object* x_11; +x_7 = lean_unbox_usize(x_4); lean_dec(x_4); +x_8 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_9 = lean_unbox_uint64(x_6); +lean_dec(x_6); +x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instHashablePreInstance___spec__1(x_1, x_2, x_3, x_7, x_8, x_9); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -return x_10; +x_11 = lean_box_uint64(x_10); +return x_11; } } -LEAN_EXPORT size_t l_Lean_Meta_Grind_getENode_x3f_unsafe__1(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instHashablePreInstance___boxed(lean_object* x_1) { _start: { -size_t x_2; -x_2 = lean_ptr_addr(x_1); -return x_2; +uint64_t x_2; lean_object* x_3; +x_2 = l_Lean_Meta_Grind_instHashablePreInstance(x_1); +lean_dec(x_1); +x_3 = lean_box_uint64(x_2); +return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getENode_x3f_unsafe__1___boxed(lean_object* x_1) { +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1___closed__1() { _start: { -size_t x_2; lean_object* x_3; -x_2 = l_Lean_Meta_Grind_getENode_x3f_unsafe__1(x_1); -lean_dec(x_1); -x_3 = lean_box_usize(x_2); +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_Meta_Grind_getENode_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) { +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___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) { _start: { -lean_object* x_11; uint8_t x_12; -x_11 = lean_st_ref_get(x_2, x_10); -x_12 = !lean_is_exclusive(x_11); -if (x_12 == 0) +uint8_t x_8; +x_8 = lean_usize_dec_lt(x_6, x_5); +if (x_8 == 0) { -lean_object* x_13; lean_object* x_14; size_t x_15; lean_object* x_16; +lean_dec(x_3); +return x_7; +} +else +{ +lean_object* x_9; uint8_t x_10; +x_9 = lean_array_uget(x_4, x_6); +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; uint8_t x_16; +x_11 = lean_ctor_get(x_7, 1); +x_12 = lean_ctor_get(x_7, 0); +lean_dec(x_12); +x_13 = lean_ctor_get(x_11, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_11, 1); +lean_inc(x_14); +x_15 = lean_ctor_get(x_11, 2); +lean_inc(x_15); +x_16 = lean_nat_dec_lt(x_14, x_15); +if (x_16 == 0) +{ +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_9); +lean_ctor_set(x_7, 0, x_3); +return x_7; +} +else +{ +uint8_t x_17; +x_17 = !lean_is_exclusive(x_11); +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; uint8_t x_24; +x_18 = lean_ctor_get(x_11, 2); +lean_dec(x_18); +x_19 = lean_ctor_get(x_11, 1); +lean_dec(x_19); +x_20 = lean_ctor_get(x_11, 0); +lean_dec(x_20); +x_21 = lean_array_fget(x_13, x_14); +x_22 = lean_unsigned_to_nat(1u); +x_23 = lean_nat_add(x_14, x_22); +lean_dec(x_14); +lean_ctor_set(x_11, 1, x_23); +x_24 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_9, x_21); +lean_dec(x_21); +lean_dec(x_9); +if (x_24 == 0) +{ +lean_object* x_25; +lean_dec(x_3); +x_25 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1___closed__1; +lean_ctor_set(x_7, 0, x_25); +return x_7; +} +else +{ +size_t x_26; size_t x_27; +lean_inc(x_3); +lean_ctor_set(x_7, 0, x_3); +x_26 = 1; +x_27 = lean_usize_add(x_6, x_26); +x_6 = x_27; +goto _start; +} +} +else +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; +lean_dec(x_11); +x_29 = lean_array_fget(x_13, x_14); +x_30 = lean_unsigned_to_nat(1u); +x_31 = lean_nat_add(x_14, x_30); +lean_dec(x_14); +x_32 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_32, 0, x_13); +lean_ctor_set(x_32, 1, x_31); +lean_ctor_set(x_32, 2, x_15); +x_33 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_9, x_29); +lean_dec(x_29); +lean_dec(x_9); +if (x_33 == 0) +{ +lean_object* x_34; +lean_dec(x_3); +x_34 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1___closed__1; +lean_ctor_set(x_7, 1, x_32); +lean_ctor_set(x_7, 0, x_34); +return x_7; +} +else +{ +size_t x_35; size_t x_36; +lean_inc(x_3); +lean_ctor_set(x_7, 1, x_32); +lean_ctor_set(x_7, 0, x_3); +x_35 = 1; +x_36 = lean_usize_add(x_6, x_35); +x_6 = x_36; +goto _start; +} +} +} +} +else +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; uint8_t x_42; +x_38 = lean_ctor_get(x_7, 1); +lean_inc(x_38); +lean_dec(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); +x_41 = lean_ctor_get(x_38, 2); +lean_inc(x_41); +x_42 = lean_nat_dec_lt(x_40, x_41); +if (x_42 == 0) +{ +lean_object* x_43; +lean_dec(x_41); +lean_dec(x_40); +lean_dec(x_39); +lean_dec(x_9); +x_43 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_43, 0, x_3); +lean_ctor_set(x_43, 1, x_38); +return x_43; +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; uint8_t x_49; +if (lean_is_exclusive(x_38)) { + lean_ctor_release(x_38, 0); + lean_ctor_release(x_38, 1); + lean_ctor_release(x_38, 2); + x_44 = x_38; +} else { + lean_dec_ref(x_38); + x_44 = lean_box(0); +} +x_45 = lean_array_fget(x_39, x_40); +x_46 = lean_unsigned_to_nat(1u); +x_47 = lean_nat_add(x_40, x_46); +lean_dec(x_40); +if (lean_is_scalar(x_44)) { + x_48 = lean_alloc_ctor(0, 3, 0); +} else { + x_48 = x_44; +} +lean_ctor_set(x_48, 0, x_39); +lean_ctor_set(x_48, 1, x_47); +lean_ctor_set(x_48, 2, x_41); +x_49 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_9, x_45); +lean_dec(x_45); +lean_dec(x_9); +if (x_49 == 0) +{ +lean_object* x_50; lean_object* x_51; +lean_dec(x_3); +x_50 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1___closed__1; +x_51 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_51, 0, x_50); +lean_ctor_set(x_51, 1, x_48); +return x_51; +} +else +{ +lean_object* x_52; size_t x_53; size_t x_54; +lean_inc(x_3); +x_52 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_52, 0, x_3); +lean_ctor_set(x_52, 1, x_48); +x_53 = 1; +x_54 = lean_usize_add(x_6, x_53); +x_6 = x_54; +x_7 = x_52; +goto _start; +} +} +} +} +} +} +LEAN_EXPORT uint8_t l_Lean_Meta_Grind_instBEqPreInstance___lambda__1(lean_object* x_1) { +_start: +{ +uint8_t x_2; +x_2 = 1; +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Grind_instBEqPreInstance___lambda__2___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_instBEqPreInstance___lambda__1___boxed), 1, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_instBEqPreInstance___lambda__2___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_instBEqPreInstance___lambda__2___closed__1; +x_2 = lean_box(0); +x_3 = lean_apply_1(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instBEqPreInstance___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; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; size_t x_12; size_t x_13; lean_object* x_14; lean_object* x_15; +x_4 = lean_ctor_get(x_1, 1); +lean_inc(x_4); +lean_dec(x_1); +x_5 = lean_array_get_size(x_4); +x_6 = lean_unsigned_to_nat(0u); +x_7 = l_Array_toSubarray___rarg(x_4, x_6, x_5); +x_8 = lean_ctor_get(x_2, 1); +x_9 = lean_box(0); +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_7); +x_12 = lean_array_size(x_8); +x_13 = 0; +x_14 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1(x_8, x_9, x_10, x_8, x_12, x_13, x_11); +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +lean_dec(x_14); +if (lean_obj_tag(x_15) == 0) +{ +lean_object* x_16; +x_16 = l_Lean_Meta_Grind_instBEqPreInstance___lambda__2___closed__2; +return x_16; +} +else +{ +lean_object* x_17; +x_17 = lean_ctor_get(x_15, 0); +lean_inc(x_17); +lean_dec(x_15); +return x_17; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instBEqPreInstance___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; uint8_t x_8; +x_4 = lean_ctor_get(x_2, 1); +x_5 = lean_array_get_size(x_4); +x_6 = lean_ctor_get(x_1, 1); +lean_inc(x_6); +x_7 = lean_array_get_size(x_6); +lean_dec(x_6); +x_8 = lean_nat_dec_eq(x_5, x_7); +lean_dec(x_7); +lean_dec(x_5); +if (x_8 == 0) +{ +uint8_t x_9; lean_object* x_10; +lean_dec(x_1); +x_9 = 0; +x_10 = lean_box(x_9); +return x_10; +} +else +{ +lean_object* x_11; lean_object* x_12; +x_11 = lean_box(0); +x_12 = l_Lean_Meta_Grind_instBEqPreInstance___lambda__2(x_1, x_2, x_11); +return x_12; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instBEqPreInstance(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; uint8_t x_5; +x_3 = lean_ctor_get(x_1, 0); +x_4 = lean_ctor_get(x_2, 0); +lean_inc(x_4); +x_5 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_3, x_4); +lean_dec(x_4); +if (x_5 == 0) +{ +uint8_t x_6; lean_object* x_7; +lean_dec(x_2); +x_6 = 0; +x_7 = lean_box(x_6); +return x_7; +} +else +{ +lean_object* x_8; lean_object* x_9; +x_8 = lean_box(0); +x_9 = l_Lean_Meta_Grind_instBEqPreInstance___lambda__3(x_2, x_1, x_8); +return x_9; +} +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___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: +{ +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_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1(x_1, x_2, x_3, x_4, x_8, x_9, x_7); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instBEqPreInstance___lambda__1___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = l_Lean_Meta_Grind_instBEqPreInstance___lambda__1(x_1); +lean_dec(x_1); +x_3 = lean_box(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instBEqPreInstance___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Lean_Meta_Grind_instBEqPreInstance___lambda__2(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instBEqPreInstance___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Lean_Meta_Grind_instBEqPreInstance___lambda__3(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instBEqPreInstance___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_Meta_Grind_instBEqPreInstance(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind_instInhabitedNewFact___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_instInhabitedENode___closed__1; +x_2 = lean_unsigned_to_nat(0u); +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_Meta_Grind_instInhabitedNewFact() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Meta_Grind_instInhabitedNewFact___closed__1; +return x_1; +} +} +static lean_object* _init_l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedGoal___spec__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Meta_Grind_canon___closed__2; +return x_1; +} +} +static lean_object* _init_l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedGoal___spec__2() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Meta_Grind_canon___closed__2; +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_instInhabitedGoal___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = lean_mk_empty_array_with_capacity(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Grind_instInhabitedGoal___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_instInhabitedGoal___closed__1; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static size_t _init_l_Lean_Meta_Grind_instInhabitedGoal___closed__3() { +_start: +{ +lean_object* x_1; size_t x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = lean_usize_of_nat(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Grind_instInhabitedGoal___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; size_t x_4; lean_object* x_5; +x_1 = l_Lean_Meta_Grind_instInhabitedGoal___closed__2; +x_2 = l_Lean_Meta_Grind_instInhabitedGoal___closed__1; +x_3 = lean_unsigned_to_nat(0u); +x_4 = l_Lean_Meta_Grind_instInhabitedGoal___closed__3; +x_5 = lean_alloc_ctor(0, 4, sizeof(size_t)*1); +lean_ctor_set(x_5, 0, x_1); +lean_ctor_set(x_5, 1, x_2); +lean_ctor_set(x_5, 2, x_3); +lean_ctor_set(x_5, 3, x_3); +lean_ctor_set_usize(x_5, 4, x_4); +return x_5; +} +} +static lean_object* _init_l_Lean_Meta_Grind_instInhabitedGoal___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Queue_empty(lean_box(0)); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_instInhabitedGoal() { +_start: +{ +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; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind_canon___closed__2; +x_3 = l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedGoal___spec__1; +x_4 = l_Lean_Meta_Grind_instInhabitedGoal___closed__1; +x_5 = 0; +x_6 = lean_unsigned_to_nat(0u); +x_7 = l_Lean_Meta_Grind_instInhabitedGoal___closed__4; +x_8 = l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedGoal___spec__2; +x_9 = l_Lean_Meta_Grind_instInhabitedGoal___closed__5; +x_10 = lean_alloc_ctor(0, 14, 1); +lean_ctor_set(x_10, 0, x_1); +lean_ctor_set(x_10, 1, x_2); +lean_ctor_set(x_10, 2, x_2); +lean_ctor_set(x_10, 3, x_3); +lean_ctor_set(x_10, 4, x_2); +lean_ctor_set(x_10, 5, x_4); +lean_ctor_set(x_10, 6, x_6); +lean_ctor_set(x_10, 7, x_6); +lean_ctor_set(x_10, 8, x_7); +lean_ctor_set(x_10, 9, x_7); +lean_ctor_set(x_10, 10, x_2); +lean_ctor_set(x_10, 11, x_6); +lean_ctor_set(x_10, 12, x_8); +lean_ctor_set(x_10, 13, x_9); +lean_ctor_set_uint8(x_10, sizeof(void*)*14, x_5); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Goal_admit(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; uint8_t x_8; lean_object* x_9; +x_7 = lean_ctor_get(x_1, 0); +lean_inc(x_7); +lean_dec(x_1); +x_8 = 1; +x_9 = l_Lean_MVarId_admit(x_7, x_8, x_2, x_3, x_4, x_5, x_6); +return x_9; +} +} +LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +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_Lean_Meta_Grind_GoalM_run___spec__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 0); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___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; +x_11 = lean_apply_3(x_2, x_3, x_4, x_5); +x_12 = l___private_Lean_Meta_Basic_0__Lean_Meta_withMVarContextImp___rarg(x_1, x_11, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_12) == 0) +{ +uint8_t x_13; +x_13 = !lean_is_exclusive(x_12); +if (x_13 == 0) +{ +return x_12; +} +else +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_14 = lean_ctor_get(x_12, 0); +x_15 = lean_ctor_get(x_12, 1); +lean_inc(x_15); +lean_inc(x_14); +lean_dec(x_12); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_14); +lean_ctor_set(x_16, 1, x_15); +return x_16; +} +} +else +{ +uint8_t x_17; +x_17 = !lean_is_exclusive(x_12); +if (x_17 == 0) +{ +return x_12; +} +else +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_18 = lean_ctor_get(x_12, 0); +x_19 = lean_ctor_get(x_12, 1); +lean_inc(x_19); +lean_inc(x_18); +lean_dec(x_12); +x_20 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_20, 0, x_18); +lean_ctor_set(x_20, 1, x_19); +return x_20; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg), 10, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GoalM_run___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; uint8_t x_11; +x_10 = lean_st_mk_ref(x_1, x_9); +x_11 = !lean_is_exclusive(x_10); +if (x_11 == 0) +{ +return x_10; +} +else +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_12 = lean_ctor_get(x_10, 0); +x_13 = lean_ctor_get(x_10, 1); +lean_inc(x_13); +lean_inc(x_12); +lean_dec(x_10); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_12); +lean_ctor_set(x_14, 1, x_13); +return x_14; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GoalM_run___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: +{ +lean_object* x_11; +lean_inc(x_2); +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) +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +x_13 = lean_ctor_get(x_11, 1); +lean_inc(x_13); +lean_dec(x_11); +x_14 = lean_st_ref_get(x_2, x_13); +lean_dec(x_2); +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); +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_12); +lean_ctor_set(x_17, 1, x_16); +lean_ctor_set(x_14, 0, x_17); +return x_14; +} +else +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_18 = lean_ctor_get(x_14, 0); +x_19 = lean_ctor_get(x_14, 1); +lean_inc(x_19); +lean_inc(x_18); +lean_dec(x_14); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_12); +lean_ctor_set(x_20, 1, x_18); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_20); +lean_ctor_set(x_21, 1, x_19); +return x_21; +} +} +else +{ +uint8_t x_22; +lean_dec(x_2); +x_22 = !lean_is_exclusive(x_11); +if (x_22 == 0) +{ +return x_11; +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_11, 0); +x_24 = lean_ctor_get(x_11, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_11); +x_25 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_24); +return x_25; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GoalM_run___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_11 = lean_ctor_get(x_1, 0); +lean_inc(x_11); +x_12 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_GoalM_run___rarg___lambda__1___boxed), 9, 1); +lean_closure_set(x_12, 0, x_1); +x_13 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_GoalM_run___rarg___lambda__2), 10, 1); +lean_closure_set(x_13, 0, x_2); +x_14 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); +lean_closure_set(x_14, 0, x_12); +lean_closure_set(x_14, 1, x_13); +x_15 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(x_11, x_14, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_15; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GoalM_run(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_GoalM_run___rarg), 10, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GoalM_run___rarg___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l_Lean_Meta_Grind_GoalM_run___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); +lean_dec(x_3); +lean_dec(x_2); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GoalM_run_x27___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +lean_inc(x_2); +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) +{ +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 = lean_st_ref_get(x_2, x_12); +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_15 = lean_ctor_get(x_13, 1); +x_16 = lean_st_ref_get(x_2, x_15); +lean_dec(x_2); +x_17 = !lean_is_exclusive(x_16); +if (x_17 == 0) +{ +lean_object* x_18; +x_18 = lean_ctor_get(x_16, 0); +lean_ctor_set(x_13, 1, x_18); +lean_ctor_set(x_16, 0, x_13); +return x_16; +} +else +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_16, 0); +x_20 = lean_ctor_get(x_16, 1); +lean_inc(x_20); +lean_inc(x_19); +lean_dec(x_16); +lean_ctor_set(x_13, 1, x_19); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_13); +lean_ctor_set(x_21, 1, x_20); +return x_21; +} +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_22 = lean_ctor_get(x_13, 0); +x_23 = lean_ctor_get(x_13, 1); +lean_inc(x_23); +lean_inc(x_22); +lean_dec(x_13); +x_24 = lean_st_ref_get(x_2, x_23); +lean_dec(x_2); +x_25 = lean_ctor_get(x_24, 0); +lean_inc(x_25); +x_26 = lean_ctor_get(x_24, 1); +lean_inc(x_26); +if (lean_is_exclusive(x_24)) { + lean_ctor_release(x_24, 0); + lean_ctor_release(x_24, 1); + x_27 = x_24; +} else { + lean_dec_ref(x_24); + x_27 = lean_box(0); +} +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_22); +lean_ctor_set(x_28, 1, x_25); +if (lean_is_scalar(x_27)) { + x_29 = lean_alloc_ctor(0, 2, 0); +} else { + x_29 = x_27; +} +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_26); +return x_29; +} +} +else +{ +uint8_t x_30; +lean_dec(x_2); +x_30 = !lean_is_exclusive(x_11); +if (x_30 == 0) +{ +return x_11; +} +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_11, 0); +x_32 = lean_ctor_get(x_11, 1); +lean_inc(x_32); +lean_inc(x_31); +lean_dec(x_11); +x_33 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_33, 0, x_31); +lean_ctor_set(x_33, 1, x_32); +return x_33; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GoalM_run_x27___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; +x_10 = !lean_is_exclusive(x_1); +if (x_10 == 0) +{ +lean_object* x_11; +x_11 = lean_ctor_get(x_1, 1); +lean_dec(x_11); +lean_ctor_set(x_1, 1, x_9); +return x_1; +} +else +{ +lean_object* x_12; lean_object* x_13; +x_12 = lean_ctor_get(x_1, 0); +lean_inc(x_12); +lean_dec(x_1); +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_9); +return x_13; +} +} +} +static lean_object* _init_l_Lean_Meta_Grind_GoalM_run_x27___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_GoalM_run_x27___lambda__2___boxed), 9, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GoalM_run_x27(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_11 = lean_ctor_get(x_1, 0); +lean_inc(x_11); +x_12 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_GoalM_run___rarg___lambda__1___boxed), 9, 1); +lean_closure_set(x_12, 0, x_1); +x_13 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_GoalM_run_x27___lambda__1), 10, 1); +lean_closure_set(x_13, 0, x_2); +x_14 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); +lean_closure_set(x_14, 0, x_12); +lean_closure_set(x_14, 1, x_13); +x_15 = l_Lean_Meta_Grind_GoalM_run_x27___closed__1; +x_16 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); +lean_closure_set(x_16, 0, x_14); +lean_closure_set(x_16, 1, x_15); +x_17 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(x_11, x_16, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_17; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GoalM_run_x27___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l_Lean_Meta_Grind_GoalM_run_x27___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_10; +} +} +LEAN_EXPORT uint64_t l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, uint64_t x_6) { +_start: +{ +uint8_t x_7; +x_7 = lean_usize_dec_lt(x_5, x_4); +if (x_7 == 0) +{ +return x_6; +} +else +{ +lean_object* x_8; uint64_t x_9; uint64_t x_10; size_t x_11; size_t x_12; +x_8 = lean_array_uget(x_3, x_5); +x_9 = l_Lean_Meta_Grind_instHashablePreInstance_unsafe__2(x_8); +lean_dec(x_8); +x_10 = lean_uint64_mix_hash(x_6, x_9); +x_11 = 1; +x_12 = lean_usize_add(x_5, x_11); +x_5 = x_12; +x_6 = x_10; +goto _start; +} +} +} +LEAN_EXPORT uint64_t l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, uint64_t x_6) { +_start: +{ +uint8_t x_7; +x_7 = lean_usize_dec_lt(x_5, x_4); +if (x_7 == 0) +{ +return x_6; +} +else +{ +lean_object* x_8; uint64_t x_9; uint64_t x_10; size_t x_11; size_t x_12; +x_8 = lean_array_uget(x_3, x_5); +x_9 = l_Lean_Meta_Grind_instHashablePreInstance_unsafe__2(x_8); +lean_dec(x_8); +x_10 = lean_uint64_mix_hash(x_6, x_9); +x_11 = 1; +x_12 = lean_usize_add(x_5, x_11); +x_5 = x_12; +x_6 = x_10; +goto _start; +} +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_markTheoremInstance___spec__4(size_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_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; lean_object* x_12; lean_object* x_13; size_t x_14; size_t x_15; uint64_t x_16; size_t x_17; size_t x_18; size_t x_19; size_t x_20; size_t x_21; size_t x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_9 = lean_array_fget(x_2, x_5); +x_10 = lean_array_fget(x_3, x_5); +x_11 = l_Lean_Meta_Grind_instHashablePreInstance_unsafe__1(x_9); +x_12 = lean_ctor_get(x_9, 1); +lean_inc(x_12); +x_13 = lean_box(0); +x_14 = lean_array_size(x_12); +x_15 = 0; +x_16 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__5(x_12, x_13, x_12, x_14, x_15, x_11); +lean_dec(x_12); +x_17 = lean_uint64_to_usize(x_16); +x_18 = 1; +x_19 = lean_usize_sub(x_1, x_18); +x_20 = 5; +x_21 = lean_usize_mul(x_20, x_19); +x_22 = lean_usize_shift_right(x_17, x_21); +x_23 = lean_unsigned_to_nat(1u); +x_24 = lean_nat_add(x_5, x_23); +lean_dec(x_5); +x_25 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3(x_6, x_22, x_1, x_9, x_10); +x_4 = lean_box(0); +x_5 = x_24; +x_6 = x_25; +goto _start; +} +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, size_t x_6, lean_object* x_7) { +_start: +{ +uint8_t x_8; +x_8 = lean_usize_dec_lt(x_6, x_5); +if (x_8 == 0) +{ +lean_dec(x_3); +return x_7; +} +else +{ +lean_object* x_9; uint8_t x_10; +x_9 = lean_array_uget(x_4, x_6); +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; uint8_t x_16; +x_11 = lean_ctor_get(x_7, 1); +x_12 = lean_ctor_get(x_7, 0); +lean_dec(x_12); +x_13 = lean_ctor_get(x_11, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_11, 1); +lean_inc(x_14); +x_15 = lean_ctor_get(x_11, 2); +lean_inc(x_15); +x_16 = lean_nat_dec_lt(x_14, x_15); +if (x_16 == 0) +{ +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_9); +lean_ctor_set(x_7, 0, x_3); +return x_7; +} +else +{ +uint8_t x_17; +x_17 = !lean_is_exclusive(x_11); +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; uint8_t x_24; +x_18 = lean_ctor_get(x_11, 2); +lean_dec(x_18); +x_19 = lean_ctor_get(x_11, 1); +lean_dec(x_19); +x_20 = lean_ctor_get(x_11, 0); +lean_dec(x_20); +x_21 = lean_array_fget(x_13, x_14); +x_22 = lean_unsigned_to_nat(1u); +x_23 = lean_nat_add(x_14, x_22); +lean_dec(x_14); +lean_ctor_set(x_11, 1, x_23); +x_24 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_9, x_21); +lean_dec(x_21); +lean_dec(x_9); +if (x_24 == 0) +{ +lean_object* x_25; +lean_dec(x_3); +x_25 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1___closed__1; +lean_ctor_set(x_7, 0, x_25); +return x_7; +} +else +{ +size_t x_26; size_t x_27; +lean_inc(x_3); +lean_ctor_set(x_7, 0, x_3); +x_26 = 1; +x_27 = lean_usize_add(x_6, x_26); +x_6 = x_27; +goto _start; +} +} +else +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; +lean_dec(x_11); +x_29 = lean_array_fget(x_13, x_14); +x_30 = lean_unsigned_to_nat(1u); +x_31 = lean_nat_add(x_14, x_30); +lean_dec(x_14); +x_32 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_32, 0, x_13); +lean_ctor_set(x_32, 1, x_31); +lean_ctor_set(x_32, 2, x_15); +x_33 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_9, x_29); +lean_dec(x_29); +lean_dec(x_9); +if (x_33 == 0) +{ +lean_object* x_34; +lean_dec(x_3); +x_34 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1___closed__1; +lean_ctor_set(x_7, 1, x_32); +lean_ctor_set(x_7, 0, x_34); +return x_7; +} +else +{ +size_t x_35; size_t x_36; +lean_inc(x_3); +lean_ctor_set(x_7, 1, x_32); +lean_ctor_set(x_7, 0, x_3); +x_35 = 1; +x_36 = lean_usize_add(x_6, x_35); +x_6 = x_36; +goto _start; +} +} +} +} +else +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; uint8_t x_42; +x_38 = lean_ctor_get(x_7, 1); +lean_inc(x_38); +lean_dec(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); +x_41 = lean_ctor_get(x_38, 2); +lean_inc(x_41); +x_42 = lean_nat_dec_lt(x_40, x_41); +if (x_42 == 0) +{ +lean_object* x_43; +lean_dec(x_41); +lean_dec(x_40); +lean_dec(x_39); +lean_dec(x_9); +x_43 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_43, 0, x_3); +lean_ctor_set(x_43, 1, x_38); +return x_43; +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; uint8_t x_49; +if (lean_is_exclusive(x_38)) { + lean_ctor_release(x_38, 0); + lean_ctor_release(x_38, 1); + lean_ctor_release(x_38, 2); + x_44 = x_38; +} else { + lean_dec_ref(x_38); + x_44 = lean_box(0); +} +x_45 = lean_array_fget(x_39, x_40); +x_46 = lean_unsigned_to_nat(1u); +x_47 = lean_nat_add(x_40, x_46); +lean_dec(x_40); +if (lean_is_scalar(x_44)) { + x_48 = lean_alloc_ctor(0, 3, 0); +} else { + x_48 = x_44; +} +lean_ctor_set(x_48, 0, x_39); +lean_ctor_set(x_48, 1, x_47); +lean_ctor_set(x_48, 2, x_41); +x_49 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_9, x_45); +lean_dec(x_45); +lean_dec(x_9); +if (x_49 == 0) +{ +lean_object* x_50; lean_object* x_51; +lean_dec(x_3); +x_50 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1___closed__1; +x_51 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_51, 0, x_50); +lean_ctor_set(x_51, 1, x_48); +return x_51; +} +else +{ +lean_object* x_52; size_t x_53; size_t x_54; +lean_inc(x_3); +x_52 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_52, 0, x_3); +lean_ctor_set(x_52, 1, x_48); +x_53 = 1; +x_54 = lean_usize_add(x_6, x_53); +x_6 = x_54; +x_7 = x_52; +goto _start; +} +} +} +} +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___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) { +_start: +{ +uint8_t x_8; +x_8 = lean_usize_dec_lt(x_6, x_5); +if (x_8 == 0) +{ +lean_dec(x_3); +return x_7; +} +else +{ +lean_object* x_9; uint8_t x_10; +x_9 = lean_array_uget(x_4, x_6); +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; uint8_t x_16; +x_11 = lean_ctor_get(x_7, 1); +x_12 = lean_ctor_get(x_7, 0); +lean_dec(x_12); +x_13 = lean_ctor_get(x_11, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_11, 1); +lean_inc(x_14); +x_15 = lean_ctor_get(x_11, 2); +lean_inc(x_15); +x_16 = lean_nat_dec_lt(x_14, x_15); +if (x_16 == 0) +{ +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_9); +lean_ctor_set(x_7, 0, x_3); +return x_7; +} +else +{ +uint8_t x_17; +x_17 = !lean_is_exclusive(x_11); +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; uint8_t x_24; +x_18 = lean_ctor_get(x_11, 2); +lean_dec(x_18); +x_19 = lean_ctor_get(x_11, 1); +lean_dec(x_19); +x_20 = lean_ctor_get(x_11, 0); +lean_dec(x_20); +x_21 = lean_array_fget(x_13, x_14); +x_22 = lean_unsigned_to_nat(1u); +x_23 = lean_nat_add(x_14, x_22); +lean_dec(x_14); +lean_ctor_set(x_11, 1, x_23); +x_24 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_9, x_21); +lean_dec(x_21); +lean_dec(x_9); +if (x_24 == 0) +{ +lean_object* x_25; +lean_dec(x_3); +x_25 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1___closed__1; +lean_ctor_set(x_7, 0, x_25); +return x_7; +} +else +{ +size_t x_26; size_t x_27; +lean_inc(x_3); +lean_ctor_set(x_7, 0, x_3); +x_26 = 1; +x_27 = lean_usize_add(x_6, x_26); +x_6 = x_27; +goto _start; +} +} +else +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; +lean_dec(x_11); +x_29 = lean_array_fget(x_13, x_14); +x_30 = lean_unsigned_to_nat(1u); +x_31 = lean_nat_add(x_14, x_30); +lean_dec(x_14); +x_32 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_32, 0, x_13); +lean_ctor_set(x_32, 1, x_31); +lean_ctor_set(x_32, 2, x_15); +x_33 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_9, x_29); +lean_dec(x_29); +lean_dec(x_9); +if (x_33 == 0) +{ +lean_object* x_34; +lean_dec(x_3); +x_34 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1___closed__1; +lean_ctor_set(x_7, 1, x_32); +lean_ctor_set(x_7, 0, x_34); +return x_7; +} +else +{ +size_t x_35; size_t x_36; +lean_inc(x_3); +lean_ctor_set(x_7, 1, x_32); +lean_ctor_set(x_7, 0, x_3); +x_35 = 1; +x_36 = lean_usize_add(x_6, x_35); +x_6 = x_36; +goto _start; +} +} +} +} +else +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; uint8_t x_42; +x_38 = lean_ctor_get(x_7, 1); +lean_inc(x_38); +lean_dec(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); +x_41 = lean_ctor_get(x_38, 2); +lean_inc(x_41); +x_42 = lean_nat_dec_lt(x_40, x_41); +if (x_42 == 0) +{ +lean_object* x_43; +lean_dec(x_41); +lean_dec(x_40); +lean_dec(x_39); +lean_dec(x_9); +x_43 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_43, 0, x_3); +lean_ctor_set(x_43, 1, x_38); +return x_43; +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; uint8_t x_49; +if (lean_is_exclusive(x_38)) { + lean_ctor_release(x_38, 0); + lean_ctor_release(x_38, 1); + lean_ctor_release(x_38, 2); + x_44 = x_38; +} else { + lean_dec_ref(x_38); + x_44 = lean_box(0); +} +x_45 = lean_array_fget(x_39, x_40); +x_46 = lean_unsigned_to_nat(1u); +x_47 = lean_nat_add(x_40, x_46); +lean_dec(x_40); +if (lean_is_scalar(x_44)) { + x_48 = lean_alloc_ctor(0, 3, 0); +} else { + x_48 = x_44; +} +lean_ctor_set(x_48, 0, x_39); +lean_ctor_set(x_48, 1, x_47); +lean_ctor_set(x_48, 2, x_41); +x_49 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_9, x_45); +lean_dec(x_45); +lean_dec(x_9); +if (x_49 == 0) +{ +lean_object* x_50; lean_object* x_51; +lean_dec(x_3); +x_50 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1___closed__1; +x_51 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_51, 0, x_50); +lean_ctor_set(x_51, 1, x_48); +return x_51; +} +else +{ +lean_object* x_52; size_t x_53; size_t x_54; +lean_inc(x_3); +x_52 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_52, 0, x_3); +lean_ctor_set(x_52, 1, x_48); +x_53 = 1; +x_54 = lean_usize_add(x_6, x_53); +x_6 = x_54; +x_7 = x_52; +goto _start; +} +} +} +} +} +} +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_markTheoremInstance___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; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; size_t x_12; size_t x_13; lean_object* x_14; lean_object* x_15; +x_4 = lean_ctor_get(x_1, 1); +lean_inc(x_4); +lean_dec(x_1); +x_5 = lean_array_get_size(x_4); +x_6 = lean_unsigned_to_nat(0u); +x_7 = l_Array_toSubarray___rarg(x_4, x_6, x_5); +x_8 = lean_ctor_get(x_2, 1); +x_9 = lean_box(0); +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_7); +x_12 = lean_array_size(x_8); +x_13 = 0; +x_14 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__8(x_8, x_9, x_10, x_8, x_12, x_13, x_11); +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +lean_dec(x_14); +if (lean_obj_tag(x_15) == 0) +{ +uint8_t x_16; +x_16 = 1; +return x_16; +} +else +{ +lean_object* x_17; uint8_t x_18; +x_17 = lean_ctor_get(x_15, 0); +lean_inc(x_17); +lean_dec(x_15); +x_18 = lean_unbox(x_17); +lean_dec(x_17); +return x_18; +} +} +} +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_markTheoremInstance___spec__7___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; uint8_t x_8; +x_4 = lean_ctor_get(x_2, 1); +x_5 = lean_array_get_size(x_4); +x_6 = lean_ctor_get(x_1, 1); +lean_inc(x_6); +x_7 = lean_array_get_size(x_6); +lean_dec(x_6); +x_8 = lean_nat_dec_eq(x_5, x_7); +lean_dec(x_7); +lean_dec(x_5); +if (x_8 == 0) +{ +uint8_t x_9; +lean_dec(x_1); +x_9 = 0; +return x_9; +} +else +{ +lean_object* x_10; uint8_t x_11; +x_10 = lean_box(0); +x_11 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_markTheoremInstance___spec__7___lambda__1(x_1, x_2, x_10); +return x_11; +} +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_markTheoremInstance___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; lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_17 = lean_array_fget(x_5, x_2); +x_18 = lean_ctor_get(x_3, 0); +lean_inc(x_18); +x_19 = lean_ctor_get(x_17, 0); +lean_inc(x_19); +x_20 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_18, x_19); +lean_dec(x_19); +lean_dec(x_18); +if (x_20 == 0) +{ +lean_object* x_21; lean_object* x_22; +lean_dec(x_17); +lean_dec(x_6); +lean_dec(x_5); +x_21 = lean_unsigned_to_nat(1u); +x_22 = lean_nat_add(x_2, x_21); +lean_dec(x_2); +x_2 = x_22; +goto _start; +} +else +{ +lean_object* x_24; uint8_t x_25; +x_24 = lean_box(0); +x_25 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_markTheoremInstance___spec__7___lambda__2(x_17, x_3, x_24); +if (x_25 == 0) +{ +lean_object* x_26; lean_object* x_27; +lean_dec(x_6); +lean_dec(x_5); +x_26 = lean_unsigned_to_nat(1u); +x_27 = lean_nat_add(x_2, x_26); +lean_dec(x_2); +x_2 = x_27; +goto _start; +} +else +{ +uint8_t x_29; +x_29 = !lean_is_exclusive(x_1); +if (x_29 == 0) +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_30 = lean_ctor_get(x_1, 1); +lean_dec(x_30); +x_31 = lean_ctor_get(x_1, 0); +lean_dec(x_31); +x_32 = lean_array_fset(x_5, x_2, x_3); +x_33 = lean_array_fset(x_6, x_2, x_4); +lean_dec(x_2); +lean_ctor_set(x_1, 1, x_33); +lean_ctor_set(x_1, 0, x_32); +return x_1; +} +else +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; +lean_dec(x_1); +x_34 = lean_array_fset(x_5, x_2, x_3); +x_35 = lean_array_fset(x_6, x_2, x_4); +lean_dec(x_2); +x_36 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_36, 0, x_34); +lean_ctor_set(x_36, 1, x_35); +return x_36; +} +} +} +} +} +} +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3___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; 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; +x_4 = lean_ctor_get(x_1, 1); +lean_inc(x_4); +lean_dec(x_1); +x_5 = lean_array_get_size(x_4); +x_6 = lean_unsigned_to_nat(0u); +x_7 = l_Array_toSubarray___rarg(x_4, x_6, x_5); +x_8 = lean_ctor_get(x_2, 1); +x_9 = lean_box(0); +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_7); +x_12 = lean_array_size(x_8); +x_13 = 0; +x_14 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__6(x_8, x_9, x_10, x_8, x_12, x_13, x_11); +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +lean_dec(x_14); +if (lean_obj_tag(x_15) == 0) +{ +uint8_t x_16; +x_16 = 1; +return x_16; +} +else +{ +lean_object* x_17; uint8_t x_18; +x_17 = lean_ctor_get(x_15, 0); +lean_inc(x_17); +lean_dec(x_15); +x_18 = lean_unbox(x_17); +lean_dec(x_17); +return x_18; +} +} +} +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3___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; uint8_t x_8; +x_4 = lean_ctor_get(x_2, 1); +x_5 = lean_array_get_size(x_4); +x_6 = lean_ctor_get(x_1, 1); +lean_inc(x_6); +x_7 = lean_array_get_size(x_6); +lean_dec(x_6); +x_8 = lean_nat_dec_eq(x_5, x_7); +lean_dec(x_7); +lean_dec(x_5); +if (x_8 == 0) +{ +uint8_t x_9; +lean_dec(x_1); +x_9 = 0; +return x_9; +} +else +{ +lean_object* x_10; uint8_t x_11; +x_10 = lean_box(0); +x_11 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3___lambda__1(x_1, x_2, x_10); +return x_11; +} +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3(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_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___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; lean_object* x_21; lean_object* x_22; uint8_t x_23; +x_19 = lean_ctor_get(x_15, 0); +x_20 = lean_ctor_get(x_15, 1); +x_21 = lean_ctor_get(x_4, 0); +lean_inc(x_21); +x_22 = lean_ctor_get(x_19, 0); +lean_inc(x_22); +x_23 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_21, x_22); +lean_dec(x_22); +lean_dec(x_21); +if (x_23 == 0) +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +lean_free_object(x_15); +x_24 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_19, x_20, x_4, x_5); +x_25 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_25, 0, x_24); +x_26 = lean_array_fset(x_17, x_12, x_25); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_26); +return x_1; +} +else +{ +uint8_t x_27; +lean_inc(x_19); +x_27 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3___lambda__2(x_19, x_4, x_16); +if (x_27 == 0) +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; +lean_free_object(x_15); +x_28 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_19, x_20, x_4, x_5); +x_29 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_29, 0, x_28); +x_30 = lean_array_fset(x_17, x_12, x_29); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_30); +return x_1; +} +else +{ +lean_object* x_31; +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_31 = lean_array_fset(x_17, x_12, x_15); +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_object* x_34; lean_object* x_35; uint8_t x_36; +x_32 = lean_ctor_get(x_15, 0); +x_33 = lean_ctor_get(x_15, 1); +lean_inc(x_33); +lean_inc(x_32); +lean_dec(x_15); +x_34 = lean_ctor_get(x_4, 0); +lean_inc(x_34); +x_35 = lean_ctor_get(x_32, 0); +lean_inc(x_35); +x_36 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_34, x_35); +lean_dec(x_35); +lean_dec(x_34); +if (x_36 == 0) +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_37 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_32, x_33, x_4, x_5); +x_38 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_38, 0, x_37); +x_39 = lean_array_fset(x_17, x_12, x_38); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_39); +return x_1; +} +else +{ +uint8_t x_40; +lean_inc(x_32); +x_40 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3___lambda__2(x_32, x_4, x_16); +if (x_40 == 0) +{ +lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_41 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_32, x_33, x_4, x_5); +x_42 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_42, 0, x_41); +x_43 = lean_array_fset(x_17, x_12, x_42); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_43); +return x_1; +} +else +{ +lean_object* x_44; lean_object* x_45; +lean_dec(x_33); +lean_dec(x_32); +x_44 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_44, 0, x_4); +lean_ctor_set(x_44, 1, x_5); +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; +} +} +} +} +case 1: +{ +uint8_t x_46; +x_46 = !lean_is_exclusive(x_15); +if (x_46 == 0) +{ +lean_object* x_47; size_t x_48; size_t x_49; lean_object* x_50; lean_object* x_51; +x_47 = lean_ctor_get(x_15, 0); +x_48 = lean_usize_shift_right(x_2, x_9); +x_49 = lean_usize_add(x_3, x_8); +x_50 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3(x_47, x_48, x_49, x_4, x_5); +lean_ctor_set(x_15, 0, x_50); +x_51 = lean_array_fset(x_17, x_12, x_15); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_51); +return x_1; +} +else +{ +lean_object* x_52; size_t x_53; size_t x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_52 = lean_ctor_get(x_15, 0); +lean_inc(x_52); +lean_dec(x_15); +x_53 = lean_usize_shift_right(x_2, x_9); +x_54 = lean_usize_add(x_3, x_8); +x_55 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3(x_52, x_53, x_54, x_4, x_5); +x_56 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_56, 0, x_55); +x_57 = lean_array_fset(x_17, x_12, x_56); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_57); +return x_1; +} +} +default: +{ +lean_object* x_58; lean_object* x_59; +x_58 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_58, 0, x_4); +lean_ctor_set(x_58, 1, x_5); +x_59 = lean_array_fset(x_17, x_12, x_58); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_59); +return x_1; +} +} +} +} +else +{ +lean_object* x_60; size_t x_61; size_t x_62; size_t x_63; size_t x_64; lean_object* x_65; lean_object* x_66; uint8_t x_67; +x_60 = lean_ctor_get(x_1, 0); +lean_inc(x_60); +lean_dec(x_1); +x_61 = 1; +x_62 = 5; +x_63 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__2___closed__2; +x_64 = lean_usize_land(x_2, x_63); +x_65 = lean_usize_to_nat(x_64); +x_66 = lean_array_get_size(x_60); +x_67 = lean_nat_dec_lt(x_65, x_66); +lean_dec(x_66); +if (x_67 == 0) +{ +lean_object* x_68; +lean_dec(x_65); +lean_dec(x_5); +lean_dec(x_4); +x_68 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_68, 0, x_60); +return x_68; +} +else +{ +lean_object* x_69; lean_object* x_70; lean_object* x_71; +x_69 = lean_array_fget(x_60, x_65); +x_70 = lean_box(0); +x_71 = lean_array_fset(x_60, x_65, x_70); +switch (lean_obj_tag(x_69)) { +case 0: +{ +lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; uint8_t x_77; +x_72 = lean_ctor_get(x_69, 0); +lean_inc(x_72); +x_73 = lean_ctor_get(x_69, 1); +lean_inc(x_73); +if (lean_is_exclusive(x_69)) { + lean_ctor_release(x_69, 0); + lean_ctor_release(x_69, 1); + x_74 = x_69; +} else { + lean_dec_ref(x_69); + x_74 = lean_box(0); +} +x_75 = lean_ctor_get(x_4, 0); +lean_inc(x_75); +x_76 = lean_ctor_get(x_72, 0); +lean_inc(x_76); +x_77 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_75, x_76); +lean_dec(x_76); +lean_dec(x_75); +if (x_77 == 0) +{ +lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; +lean_dec(x_74); +x_78 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_72, x_73, x_4, x_5); +x_79 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_79, 0, x_78); +x_80 = lean_array_fset(x_71, x_65, x_79); +lean_dec(x_65); +x_81 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_81, 0, x_80); +return x_81; +} +else +{ +uint8_t x_82; +lean_inc(x_72); +x_82 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3___lambda__2(x_72, x_4, x_70); +if (x_82 == 0) +{ +lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; +lean_dec(x_74); +x_83 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_72, x_73, x_4, x_5); +x_84 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_84, 0, x_83); +x_85 = lean_array_fset(x_71, x_65, x_84); +lean_dec(x_65); +x_86 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_86, 0, x_85); +return x_86; +} +else +{ +lean_object* x_87; lean_object* x_88; lean_object* x_89; +lean_dec(x_73); +lean_dec(x_72); +if (lean_is_scalar(x_74)) { + x_87 = lean_alloc_ctor(0, 2, 0); +} else { + x_87 = x_74; +} +lean_ctor_set(x_87, 0, x_4); +lean_ctor_set(x_87, 1, x_5); +x_88 = lean_array_fset(x_71, x_65, x_87); +lean_dec(x_65); +x_89 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_89, 0, x_88); +return x_89; +} +} +} +case 1: +{ +lean_object* x_90; lean_object* x_91; size_t x_92; size_t x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; +x_90 = lean_ctor_get(x_69, 0); +lean_inc(x_90); +if (lean_is_exclusive(x_69)) { + lean_ctor_release(x_69, 0); + x_91 = x_69; +} else { + lean_dec_ref(x_69); + x_91 = lean_box(0); +} +x_92 = lean_usize_shift_right(x_2, x_62); +x_93 = lean_usize_add(x_3, x_61); +x_94 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3(x_90, x_92, x_93, x_4, x_5); +if (lean_is_scalar(x_91)) { + x_95 = lean_alloc_ctor(1, 1, 0); +} else { + x_95 = x_91; +} +lean_ctor_set(x_95, 0, x_94); +x_96 = lean_array_fset(x_71, x_65, x_95); +lean_dec(x_65); +x_97 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_97, 0, x_96); +return x_97; +} +default: +{ +lean_object* x_98; lean_object* x_99; lean_object* x_100; +x_98 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_98, 0, x_4); +lean_ctor_set(x_98, 1, x_5); +x_99 = lean_array_fset(x_71, x_65, x_98); +lean_dec(x_65); +x_100 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_100, 0, x_99); +return x_100; +} +} +} +} +} +else +{ +uint8_t x_101; +x_101 = !lean_is_exclusive(x_1); +if (x_101 == 0) +{ +lean_object* x_102; lean_object* x_103; size_t x_104; uint8_t x_105; +x_102 = lean_unsigned_to_nat(0u); +x_103 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_markTheoremInstance___spec__7(x_1, x_102, x_4, x_5); +x_104 = 7; +x_105 = lean_usize_dec_le(x_104, x_3); +if (x_105 == 0) +{ +lean_object* x_106; lean_object* x_107; uint8_t x_108; +x_106 = l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(x_103); +x_107 = lean_unsigned_to_nat(4u); +x_108 = lean_nat_dec_lt(x_106, x_107); +lean_dec(x_106); +if (x_108 == 0) +{ +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_inc(x_109); +x_110 = lean_ctor_get(x_103, 1); +lean_inc(x_110); +lean_dec(x_103); +x_111 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__2___closed__3; +x_112 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_markTheoremInstance___spec__4(x_3, x_109, x_110, lean_box(0), x_102, x_111); +lean_dec(x_110); +lean_dec(x_109); +return x_112; +} +else +{ +return x_103; +} +} +else +{ +return x_103; +} +} +else +{ +lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; size_t x_118; uint8_t x_119; +x_113 = lean_ctor_get(x_1, 0); +x_114 = lean_ctor_get(x_1, 1); +lean_inc(x_114); +lean_inc(x_113); +lean_dec(x_1); +x_115 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_115, 0, x_113); +lean_ctor_set(x_115, 1, x_114); +x_116 = lean_unsigned_to_nat(0u); +x_117 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_markTheoremInstance___spec__7(x_115, x_116, x_4, x_5); +x_118 = 7; +x_119 = lean_usize_dec_le(x_118, x_3); +if (x_119 == 0) +{ +lean_object* x_120; lean_object* x_121; uint8_t x_122; +x_120 = l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(x_117); +x_121 = lean_unsigned_to_nat(4u); +x_122 = lean_nat_dec_lt(x_120, x_121); +lean_dec(x_120); +if (x_122 == 0) +{ +lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; +x_123 = lean_ctor_get(x_117, 0); +lean_inc(x_123); +x_124 = lean_ctor_get(x_117, 1); +lean_inc(x_124); +lean_dec(x_117); +x_125 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__2___closed__3; +x_126 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_markTheoremInstance___spec__4(x_3, x_123, x_124, lean_box(0), x_116, x_125); +lean_dec(x_124); +lean_dec(x_123); +return x_126; +} +else +{ +return x_117; +} +} +else +{ +return x_117; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_markTheoremInstance___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint64_t x_4; lean_object* x_5; lean_object* x_6; size_t x_7; size_t x_8; uint64_t x_9; size_t x_10; size_t x_11; lean_object* x_12; +x_4 = l_Lean_Meta_Grind_instHashablePreInstance_unsafe__1(x_2); +x_5 = lean_ctor_get(x_2, 1); +lean_inc(x_5); +x_6 = lean_box(0); +x_7 = lean_array_size(x_5); +x_8 = 0; +x_9 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__2(x_5, x_6, x_5, x_7, x_8, x_4); +lean_dec(x_5); +x_10 = lean_uint64_to_usize(x_9); +x_11 = 1; +x_12 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3(x_1, x_10, x_11, x_2, x_3); +return x_12; +} +} +LEAN_EXPORT uint64_t l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__10(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, uint64_t x_6) { +_start: +{ +uint8_t x_7; +x_7 = lean_usize_dec_lt(x_5, x_4); +if (x_7 == 0) +{ +return x_6; +} +else +{ +lean_object* x_8; uint64_t x_9; uint64_t x_10; size_t x_11; size_t x_12; +x_8 = lean_array_uget(x_3, x_5); +x_9 = l_Lean_Meta_Grind_instHashablePreInstance_unsafe__2(x_8); +lean_dec(x_8); +x_10 = lean_uint64_mix_hash(x_6, x_9); +x_11 = 1; +x_12 = lean_usize_add(x_5, x_11); +x_5 = x_12; +x_6 = x_10; +goto _start; +} +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__12(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, size_t x_6, lean_object* x_7) { +_start: +{ +uint8_t x_8; +x_8 = lean_usize_dec_lt(x_6, x_5); +if (x_8 == 0) +{ +lean_dec(x_3); +return x_7; +} +else +{ +lean_object* x_9; uint8_t x_10; +x_9 = lean_array_uget(x_4, x_6); +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; uint8_t x_16; +x_11 = lean_ctor_get(x_7, 1); +x_12 = lean_ctor_get(x_7, 0); +lean_dec(x_12); +x_13 = lean_ctor_get(x_11, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_11, 1); +lean_inc(x_14); +x_15 = lean_ctor_get(x_11, 2); +lean_inc(x_15); +x_16 = lean_nat_dec_lt(x_14, x_15); +if (x_16 == 0) +{ +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_9); +lean_ctor_set(x_7, 0, x_3); +return x_7; +} +else +{ +uint8_t x_17; +x_17 = !lean_is_exclusive(x_11); +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; uint8_t x_24; +x_18 = lean_ctor_get(x_11, 2); +lean_dec(x_18); +x_19 = lean_ctor_get(x_11, 1); +lean_dec(x_19); +x_20 = lean_ctor_get(x_11, 0); +lean_dec(x_20); +x_21 = lean_array_fget(x_13, x_14); +x_22 = lean_unsigned_to_nat(1u); +x_23 = lean_nat_add(x_14, x_22); +lean_dec(x_14); +lean_ctor_set(x_11, 1, x_23); +x_24 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_9, x_21); +lean_dec(x_21); +lean_dec(x_9); +if (x_24 == 0) +{ +lean_object* x_25; +lean_dec(x_3); +x_25 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1___closed__1; +lean_ctor_set(x_7, 0, x_25); +return x_7; +} +else +{ +size_t x_26; size_t x_27; +lean_inc(x_3); +lean_ctor_set(x_7, 0, x_3); +x_26 = 1; +x_27 = lean_usize_add(x_6, x_26); +x_6 = x_27; +goto _start; +} +} +else +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; +lean_dec(x_11); +x_29 = lean_array_fget(x_13, x_14); +x_30 = lean_unsigned_to_nat(1u); +x_31 = lean_nat_add(x_14, x_30); +lean_dec(x_14); +x_32 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_32, 0, x_13); +lean_ctor_set(x_32, 1, x_31); +lean_ctor_set(x_32, 2, x_15); +x_33 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_9, x_29); +lean_dec(x_29); +lean_dec(x_9); +if (x_33 == 0) +{ +lean_object* x_34; +lean_dec(x_3); +x_34 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1___closed__1; +lean_ctor_set(x_7, 1, x_32); +lean_ctor_set(x_7, 0, x_34); +return x_7; +} +else +{ +size_t x_35; size_t x_36; +lean_inc(x_3); +lean_ctor_set(x_7, 1, x_32); +lean_ctor_set(x_7, 0, x_3); +x_35 = 1; +x_36 = lean_usize_add(x_6, x_35); +x_6 = x_36; +goto _start; +} +} +} +} +else +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; uint8_t x_42; +x_38 = lean_ctor_get(x_7, 1); +lean_inc(x_38); +lean_dec(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); +x_41 = lean_ctor_get(x_38, 2); +lean_inc(x_41); +x_42 = lean_nat_dec_lt(x_40, x_41); +if (x_42 == 0) +{ +lean_object* x_43; +lean_dec(x_41); +lean_dec(x_40); +lean_dec(x_39); +lean_dec(x_9); +x_43 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_43, 0, x_3); +lean_ctor_set(x_43, 1, x_38); +return x_43; +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; uint8_t x_49; +if (lean_is_exclusive(x_38)) { + lean_ctor_release(x_38, 0); + lean_ctor_release(x_38, 1); + lean_ctor_release(x_38, 2); + x_44 = x_38; +} else { + lean_dec_ref(x_38); + x_44 = lean_box(0); +} +x_45 = lean_array_fget(x_39, x_40); +x_46 = lean_unsigned_to_nat(1u); +x_47 = lean_nat_add(x_40, x_46); +lean_dec(x_40); +if (lean_is_scalar(x_44)) { + x_48 = lean_alloc_ctor(0, 3, 0); +} else { + x_48 = x_44; +} +lean_ctor_set(x_48, 0, x_39); +lean_ctor_set(x_48, 1, x_47); +lean_ctor_set(x_48, 2, x_41); +x_49 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_9, x_45); +lean_dec(x_45); +lean_dec(x_9); +if (x_49 == 0) +{ +lean_object* x_50; lean_object* x_51; +lean_dec(x_3); +x_50 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1___closed__1; +x_51 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_51, 0, x_50); +lean_ctor_set(x_51, 1, x_48); +return x_51; +} +else +{ +lean_object* x_52; size_t x_53; size_t x_54; +lean_inc(x_3); +x_52 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_52, 0, x_3); +lean_ctor_set(x_52, 1, x_48); +x_53 = 1; +x_54 = lean_usize_add(x_6, x_53); +x_6 = x_54; +x_7 = x_52; +goto _start; +} +} +} +} +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___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: +{ +uint8_t x_8; +x_8 = lean_usize_dec_lt(x_6, x_5); +if (x_8 == 0) +{ +lean_dec(x_3); +return x_7; +} +else +{ +lean_object* x_9; uint8_t x_10; +x_9 = lean_array_uget(x_4, x_6); +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; uint8_t x_16; +x_11 = lean_ctor_get(x_7, 1); +x_12 = lean_ctor_get(x_7, 0); +lean_dec(x_12); +x_13 = lean_ctor_get(x_11, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_11, 1); +lean_inc(x_14); +x_15 = lean_ctor_get(x_11, 2); +lean_inc(x_15); +x_16 = lean_nat_dec_lt(x_14, x_15); +if (x_16 == 0) +{ +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_9); +lean_ctor_set(x_7, 0, x_3); +return x_7; +} +else +{ +uint8_t x_17; +x_17 = !lean_is_exclusive(x_11); +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; uint8_t x_24; +x_18 = lean_ctor_get(x_11, 2); +lean_dec(x_18); +x_19 = lean_ctor_get(x_11, 1); +lean_dec(x_19); +x_20 = lean_ctor_get(x_11, 0); +lean_dec(x_20); +x_21 = lean_array_fget(x_13, x_14); +x_22 = lean_unsigned_to_nat(1u); +x_23 = lean_nat_add(x_14, x_22); +lean_dec(x_14); +lean_ctor_set(x_11, 1, x_23); +x_24 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_9, x_21); +lean_dec(x_21); +lean_dec(x_9); +if (x_24 == 0) +{ +lean_object* x_25; +lean_dec(x_3); +x_25 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1___closed__1; +lean_ctor_set(x_7, 0, x_25); +return x_7; +} +else +{ +size_t x_26; size_t x_27; +lean_inc(x_3); +lean_ctor_set(x_7, 0, x_3); +x_26 = 1; +x_27 = lean_usize_add(x_6, x_26); +x_6 = x_27; +goto _start; +} +} +else +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; +lean_dec(x_11); +x_29 = lean_array_fget(x_13, x_14); +x_30 = lean_unsigned_to_nat(1u); +x_31 = lean_nat_add(x_14, x_30); +lean_dec(x_14); +x_32 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_32, 0, x_13); +lean_ctor_set(x_32, 1, x_31); +lean_ctor_set(x_32, 2, x_15); +x_33 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_9, x_29); +lean_dec(x_29); +lean_dec(x_9); +if (x_33 == 0) +{ +lean_object* x_34; +lean_dec(x_3); +x_34 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1___closed__1; +lean_ctor_set(x_7, 1, x_32); +lean_ctor_set(x_7, 0, x_34); +return x_7; +} +else +{ +size_t x_35; size_t x_36; +lean_inc(x_3); +lean_ctor_set(x_7, 1, x_32); +lean_ctor_set(x_7, 0, x_3); +x_35 = 1; +x_36 = lean_usize_add(x_6, x_35); +x_6 = x_36; +goto _start; +} +} +} +} +else +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; uint8_t x_42; +x_38 = lean_ctor_get(x_7, 1); +lean_inc(x_38); +lean_dec(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); +x_41 = lean_ctor_get(x_38, 2); +lean_inc(x_41); +x_42 = lean_nat_dec_lt(x_40, x_41); +if (x_42 == 0) +{ +lean_object* x_43; +lean_dec(x_41); +lean_dec(x_40); +lean_dec(x_39); +lean_dec(x_9); +x_43 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_43, 0, x_3); +lean_ctor_set(x_43, 1, x_38); +return x_43; +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; uint8_t x_49; +if (lean_is_exclusive(x_38)) { + lean_ctor_release(x_38, 0); + lean_ctor_release(x_38, 1); + lean_ctor_release(x_38, 2); + x_44 = x_38; +} else { + lean_dec_ref(x_38); + x_44 = lean_box(0); +} +x_45 = lean_array_fget(x_39, x_40); +x_46 = lean_unsigned_to_nat(1u); +x_47 = lean_nat_add(x_40, x_46); +lean_dec(x_40); +if (lean_is_scalar(x_44)) { + x_48 = lean_alloc_ctor(0, 3, 0); +} else { + x_48 = x_44; +} +lean_ctor_set(x_48, 0, x_39); +lean_ctor_set(x_48, 1, x_47); +lean_ctor_set(x_48, 2, x_41); +x_49 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_9, x_45); +lean_dec(x_45); +lean_dec(x_9); +if (x_49 == 0) +{ +lean_object* x_50; lean_object* x_51; +lean_dec(x_3); +x_50 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1___closed__1; +x_51 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_51, 0, x_50); +lean_ctor_set(x_51, 1, x_48); +return x_51; +} +else +{ +lean_object* x_52; size_t x_53; size_t x_54; +lean_inc(x_3); +x_52 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_52, 0, x_3); +lean_ctor_set(x_52, 1, x_48); +x_53 = 1; +x_54 = lean_usize_add(x_6, x_53); +x_6 = x_54; +x_7 = x_52; +goto _start; +} +} +} +} +} +} +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_markTheoremInstance___spec__13___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; 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; +x_4 = lean_ctor_get(x_1, 1); +lean_inc(x_4); +lean_dec(x_1); +x_5 = lean_array_get_size(x_4); +x_6 = lean_unsigned_to_nat(0u); +x_7 = l_Array_toSubarray___rarg(x_4, x_6, x_5); +x_8 = lean_ctor_get(x_2, 1); +x_9 = lean_box(0); +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_7); +x_12 = lean_array_size(x_8); +x_13 = 0; +x_14 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__14(x_8, x_9, x_10, x_8, x_12, x_13, x_11); +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +lean_dec(x_14); +if (lean_obj_tag(x_15) == 0) +{ +uint8_t x_16; +x_16 = 1; +return x_16; +} +else +{ +lean_object* x_17; uint8_t x_18; +x_17 = lean_ctor_get(x_15, 0); +lean_inc(x_17); +lean_dec(x_15); +x_18 = lean_unbox(x_17); +lean_dec(x_17); +return x_18; +} +} +} +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_markTheoremInstance___spec__13___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; uint8_t x_8; +x_4 = lean_ctor_get(x_2, 1); +x_5 = lean_array_get_size(x_4); +x_6 = lean_ctor_get(x_1, 1); +lean_inc(x_6); +x_7 = lean_array_get_size(x_6); +lean_dec(x_6); +x_8 = lean_nat_dec_eq(x_5, x_7); +lean_dec(x_7); +lean_dec(x_5); +if (x_8 == 0) +{ +uint8_t x_9; +lean_dec(x_1); +x_9 = 0; +return x_9; +} +else +{ +lean_object* x_10; uint8_t x_11; +x_10 = lean_box(0); +x_11 = l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_markTheoremInstance___spec__13___lambda__1(x_1, x_2, x_10); +return x_11; +} +} +} +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_markTheoremInstance___spec__13(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; 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) +{ +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; uint8_t x_12; +x_9 = lean_array_fget(x_1, x_4); +x_10 = lean_ctor_get(x_5, 0); +x_11 = lean_ctor_get(x_9, 0); +lean_inc(x_11); +x_12 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_10, x_11); +lean_dec(x_11); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; +lean_dec(x_9); +x_13 = lean_unsigned_to_nat(1u); +x_14 = lean_nat_add(x_4, x_13); +lean_dec(x_4); +x_3 = lean_box(0); +x_4 = x_14; +goto _start; +} +else +{ +lean_object* x_16; uint8_t x_17; +x_16 = lean_box(0); +x_17 = l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_markTheoremInstance___spec__13___lambda__2(x_9, x_5, x_16); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; +x_18 = lean_unsigned_to_nat(1u); +x_19 = lean_nat_add(x_4, x_18); +lean_dec(x_4); +x_3 = lean_box(0); +x_4 = x_19; +goto _start; +} +else +{ +uint8_t x_21; +lean_dec(x_4); +x_21 = 1; +return x_21; +} +} +} +} +} +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_markTheoremInstance___spec__11___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; 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; +x_4 = lean_ctor_get(x_1, 1); +lean_inc(x_4); +lean_dec(x_1); +x_5 = lean_array_get_size(x_4); +x_6 = lean_unsigned_to_nat(0u); +x_7 = l_Array_toSubarray___rarg(x_4, x_6, x_5); +x_8 = lean_ctor_get(x_2, 1); +x_9 = lean_box(0); +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_7); +x_12 = lean_array_size(x_8); +x_13 = 0; +x_14 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__12(x_8, x_9, x_10, x_8, x_12, x_13, x_11); +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +lean_dec(x_14); +if (lean_obj_tag(x_15) == 0) +{ +uint8_t x_16; +x_16 = 1; +return x_16; +} +else +{ +lean_object* x_17; uint8_t x_18; +x_17 = lean_ctor_get(x_15, 0); +lean_inc(x_17); +lean_dec(x_15); +x_18 = lean_unbox(x_17); +lean_dec(x_17); +return x_18; +} +} +} +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_markTheoremInstance___spec__11___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; uint8_t x_8; +x_4 = lean_ctor_get(x_2, 1); +x_5 = lean_array_get_size(x_4); +x_6 = lean_ctor_get(x_1, 1); +lean_inc(x_6); +x_7 = lean_array_get_size(x_6); +lean_dec(x_6); +x_8 = lean_nat_dec_eq(x_5, x_7); +lean_dec(x_7); +lean_dec(x_5); +if (x_8 == 0) +{ +uint8_t x_9; +lean_dec(x_1); +x_9 = 0; +return x_9; +} +else +{ +lean_object* x_10; uint8_t x_11; +x_10 = lean_box(0); +x_11 = l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_markTheoremInstance___spec__11___lambda__1(x_1, x_2, x_10); +return x_11; +} +} +} +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_markTheoremInstance___spec__11(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_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___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; lean_object* x_13; uint8_t x_14; +x_11 = lean_ctor_get(x_10, 0); +lean_inc(x_11); +lean_dec(x_10); +x_12 = lean_ctor_get(x_3, 0); +x_13 = lean_ctor_get(x_11, 0); +lean_inc(x_13); +x_14 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_12, x_13); +lean_dec(x_13); +if (x_14 == 0) +{ +uint8_t x_15; +lean_dec(x_11); +x_15 = 0; +return x_15; +} +else +{ +lean_object* x_16; uint8_t x_17; +x_16 = lean_box(0); +x_17 = l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_markTheoremInstance___spec__11___lambda__2(x_11, x_3, x_16); +return x_17; +} +} +case 1: +{ +lean_object* x_18; size_t x_19; +x_18 = lean_ctor_get(x_10, 0); +lean_inc(x_18); +lean_dec(x_10); +x_19 = lean_usize_shift_right(x_2, x_5); +x_1 = x_18; +x_2 = x_19; +goto _start; +} +default: +{ +uint8_t x_21; +x_21 = 0; +return x_21; +} +} +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; +x_22 = lean_ctor_get(x_1, 0); +lean_inc(x_22); +x_23 = lean_ctor_get(x_1, 1); +lean_inc(x_23); +lean_dec(x_1); +x_24 = lean_unsigned_to_nat(0u); +x_25 = l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_markTheoremInstance___spec__13(x_22, x_23, lean_box(0), x_24, x_3); +lean_dec(x_23); +lean_dec(x_22); +return x_25; +} +} +} +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_markTheoremInstance___spec__9(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint64_t x_3; lean_object* x_4; lean_object* x_5; size_t x_6; size_t x_7; uint64_t x_8; size_t x_9; uint8_t x_10; +x_3 = l_Lean_Meta_Grind_instHashablePreInstance_unsafe__1(x_2); +x_4 = lean_ctor_get(x_2, 1); +x_5 = lean_box(0); +x_6 = lean_array_size(x_4); +x_7 = 0; +x_8 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__10(x_4, x_5, x_4, x_6, x_7, x_3); +x_9 = lean_uint64_to_usize(x_8); +x_10 = l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_markTheoremInstance___spec__11(x_1, x_9, x_2); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markTheoremInstance___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; uint8_t x_15; +x_12 = lean_st_ref_take(x_3, x_11); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_12, 1); +lean_inc(x_14); +lean_dec(x_12); +x_15 = !lean_is_exclusive(x_13); +if (x_15 == 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_13, 12); +x_17 = lean_box(0); +x_18 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_markTheoremInstance___spec__1(x_16, x_1, x_17); +lean_ctor_set(x_13, 12, x_18); +x_19 = lean_st_ref_set(x_3, x_13, x_14); +x_20 = !lean_is_exclusive(x_19); +if (x_20 == 0) +{ +lean_object* x_21; uint8_t x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_19, 0); +lean_dec(x_21); +x_22 = 1; +x_23 = lean_box(x_22); +lean_ctor_set(x_19, 0, x_23); +return x_19; +} +else +{ +lean_object* x_24; uint8_t x_25; lean_object* x_26; lean_object* x_27; +x_24 = lean_ctor_get(x_19, 1); +lean_inc(x_24); +lean_dec(x_19); +x_25 = 1; +x_26 = lean_box(x_25); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_24); +return x_27; +} +} +else +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; 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; uint8_t x_49; lean_object* x_50; lean_object* x_51; +x_28 = lean_ctor_get(x_13, 0); +x_29 = lean_ctor_get(x_13, 1); +x_30 = lean_ctor_get(x_13, 2); +x_31 = lean_ctor_get(x_13, 3); +x_32 = lean_ctor_get(x_13, 4); +x_33 = lean_ctor_get(x_13, 5); +x_34 = lean_ctor_get_uint8(x_13, sizeof(void*)*14); +x_35 = lean_ctor_get(x_13, 6); +x_36 = lean_ctor_get(x_13, 7); +x_37 = lean_ctor_get(x_13, 8); +x_38 = lean_ctor_get(x_13, 9); +x_39 = lean_ctor_get(x_13, 10); +x_40 = lean_ctor_get(x_13, 11); +x_41 = lean_ctor_get(x_13, 12); +x_42 = lean_ctor_get(x_13, 13); +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_inc(x_33); +lean_inc(x_32); +lean_inc(x_31); +lean_inc(x_30); +lean_inc(x_29); +lean_inc(x_28); +lean_dec(x_13); +x_43 = lean_box(0); +x_44 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_markTheoremInstance___spec__1(x_41, x_1, x_43); +x_45 = lean_alloc_ctor(0, 14, 1); +lean_ctor_set(x_45, 0, x_28); +lean_ctor_set(x_45, 1, x_29); +lean_ctor_set(x_45, 2, x_30); +lean_ctor_set(x_45, 3, x_31); +lean_ctor_set(x_45, 4, x_32); +lean_ctor_set(x_45, 5, x_33); +lean_ctor_set(x_45, 6, x_35); +lean_ctor_set(x_45, 7, x_36); +lean_ctor_set(x_45, 8, x_37); +lean_ctor_set(x_45, 9, x_38); +lean_ctor_set(x_45, 10, x_39); +lean_ctor_set(x_45, 11, x_40); +lean_ctor_set(x_45, 12, x_44); +lean_ctor_set(x_45, 13, x_42); +lean_ctor_set_uint8(x_45, sizeof(void*)*14, x_34); +x_46 = lean_st_ref_set(x_3, x_45, x_14); +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); +} +x_49 = 1; +x_50 = lean_box(x_49); +if (lean_is_scalar(x_48)) { + x_51 = lean_alloc_ctor(0, 2, 0); +} else { + x_51 = x_48; +} +lean_ctor_set(x_51, 0, x_50); +lean_ctor_set(x_51, 1, x_47); +return x_51; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markTheoremInstance(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, 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_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_1); +lean_ctor_set(x_12, 1, x_2); +x_13 = lean_st_ref_get(x_3, x_11); +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_15 = lean_ctor_get(x_13, 0); +x_16 = lean_ctor_get(x_13, 1); +x_17 = lean_ctor_get(x_15, 12); +lean_inc(x_17); +lean_dec(x_15); +x_18 = l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_markTheoremInstance___spec__9(x_17, x_12); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; +lean_free_object(x_13); +x_19 = lean_box(0); +x_20 = l_Lean_Meta_Grind_markTheoremInstance___lambda__1(x_12, x_19, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_16); +return x_20; +} +else +{ +uint8_t x_21; lean_object* x_22; +lean_dec(x_12); +x_21 = 0; +x_22 = lean_box(x_21); +lean_ctor_set(x_13, 0, x_22); +return x_13; +} +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; +x_23 = lean_ctor_get(x_13, 0); +x_24 = lean_ctor_get(x_13, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_13); +x_25 = lean_ctor_get(x_23, 12); +lean_inc(x_25); +lean_dec(x_23); +x_26 = l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_markTheoremInstance___spec__9(x_25, x_12); +if (x_26 == 0) +{ +lean_object* x_27; lean_object* x_28; +x_27 = lean_box(0); +x_28 = l_Lean_Meta_Grind_markTheoremInstance___lambda__1(x_12, 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_29; lean_object* x_30; lean_object* x_31; +lean_dec(x_12); +x_29 = 0; +x_30 = lean_box(x_29); +x_31 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_31, 0, x_30); +lean_ctor_set(x_31, 1, x_24); +return x_31; +} +} +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +size_t x_7; size_t x_8; uint64_t x_9; uint64_t x_10; lean_object* x_11; +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_9 = lean_unbox_uint64(x_6); +lean_dec(x_6); +x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__2(x_1, x_2, x_3, x_7, x_8, x_9); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_11 = lean_box_uint64(x_10); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +size_t x_7; size_t x_8; uint64_t x_9; uint64_t x_10; lean_object* x_11; +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_9 = lean_unbox_uint64(x_6); +lean_dec(x_6); +x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__5(x_1, x_2, x_3, x_7, x_8, x_9); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_11 = lean_box_uint64(x_10); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_markTheoremInstance___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +size_t x_7; 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_Grind_markTheoremInstance___spec__4(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_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___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) { +_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_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__6(x_1, x_2, x_3, x_4, x_8, x_9, x_7); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___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: +{ +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_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__8(x_1, x_2, x_3, x_4, x_8, x_9, x_7); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_markTheoremInstance___spec__7___lambda__1___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_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_markTheoremInstance___spec__7___lambda__1(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +x_5 = lean_box(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_markTheoremInstance___spec__7___lambda__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 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_markTheoremInstance___spec__7___lambda__2(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +x_5 = lean_box(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3___lambda__1___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_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3___lambda__1(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +x_5 = lean_box(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3___lambda__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 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3___lambda__2(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +x_5 = lean_box(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +size_t x_6; size_t x_7; lean_object* x_8; +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_Grind_markTheoremInstance___spec__3(x_1, x_6, x_7, x_4, x_5); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___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: +{ +size_t x_7; size_t x_8; uint64_t x_9; uint64_t x_10; lean_object* x_11; +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_9 = lean_unbox_uint64(x_6); +lean_dec(x_6); +x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__10(x_1, x_2, x_3, x_7, x_8, x_9); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_11 = lean_box_uint64(x_10); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__12___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +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_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__12(x_1, x_2, x_3, x_4, x_8, x_9, x_7); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___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) { +_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_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__14(x_1, x_2, x_3, x_4, x_8, x_9, x_7); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_markTheoremInstance___spec__13___lambda__1___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_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_markTheoremInstance___spec__13___lambda__1(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +x_5 = lean_box(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_markTheoremInstance___spec__13___lambda__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 = l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_markTheoremInstance___spec__13___lambda__2(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +x_5 = lean_box(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_markTheoremInstance___spec__13___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 = l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_markTheoremInstance___spec__13(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_5); +lean_dec(x_2); +lean_dec(x_1); +x_7 = lean_box(x_6); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_markTheoremInstance___spec__11___lambda__1___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_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_markTheoremInstance___spec__11___lambda__1(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +x_5 = lean_box(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_markTheoremInstance___spec__11___lambda__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 = l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_markTheoremInstance___spec__11___lambda__2(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +x_5 = lean_box(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_markTheoremInstance___spec__11___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_Meta_Grind_markTheoremInstance___spec__11(x_1, x_4, x_3); +lean_dec(x_3); +x_6 = lean_box(x_5); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_markTheoremInstance___spec__9___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_Meta_Grind_markTheoremInstance___spec__9(x_1, x_2); +lean_dec(x_2); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markTheoremInstance___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_Meta_Grind_markTheoremInstance___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markTheoremInstance___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, 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_Meta_Grind_markTheoremInstance(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addNewFact(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_13 = lean_st_ref_take(x_4, x_12); +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +x_16 = !lean_is_exclusive(x_14); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_17 = lean_ctor_get(x_14, 13); +x_18 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_18, 0, x_1); +lean_ctor_set(x_18, 1, x_2); +lean_ctor_set(x_18, 2, x_3); +x_19 = l_Std_Queue_enqueue___rarg(x_18, x_17); +lean_ctor_set(x_14, 13, x_19); +x_20 = lean_st_ref_set(x_4, x_14, x_15); +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); +lean_dec(x_22); +x_23 = lean_box(0); +lean_ctor_set(x_20, 0, x_23); +return x_20; +} +else +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_20, 1); +lean_inc(x_24); +lean_dec(x_20); +x_25 = lean_box(0); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set(x_26, 1, x_24); +return x_26; +} +} +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; 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; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_27 = lean_ctor_get(x_14, 0); +x_28 = lean_ctor_get(x_14, 1); +x_29 = lean_ctor_get(x_14, 2); +x_30 = lean_ctor_get(x_14, 3); +x_31 = lean_ctor_get(x_14, 4); +x_32 = lean_ctor_get(x_14, 5); +x_33 = lean_ctor_get_uint8(x_14, sizeof(void*)*14); +x_34 = lean_ctor_get(x_14, 6); +x_35 = lean_ctor_get(x_14, 7); +x_36 = lean_ctor_get(x_14, 8); +x_37 = lean_ctor_get(x_14, 9); +x_38 = lean_ctor_get(x_14, 10); +x_39 = lean_ctor_get(x_14, 11); +x_40 = lean_ctor_get(x_14, 12); +x_41 = lean_ctor_get(x_14, 13); +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); +lean_inc(x_32); +lean_inc(x_31); +lean_inc(x_30); +lean_inc(x_29); +lean_inc(x_28); +lean_inc(x_27); +lean_dec(x_14); +x_42 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_42, 0, x_1); +lean_ctor_set(x_42, 1, x_2); +lean_ctor_set(x_42, 2, x_3); +x_43 = l_Std_Queue_enqueue___rarg(x_42, x_41); +x_44 = lean_alloc_ctor(0, 14, 1); +lean_ctor_set(x_44, 0, x_27); +lean_ctor_set(x_44, 1, x_28); +lean_ctor_set(x_44, 2, x_29); +lean_ctor_set(x_44, 3, x_30); +lean_ctor_set(x_44, 4, x_31); +lean_ctor_set(x_44, 5, x_32); +lean_ctor_set(x_44, 6, x_34); +lean_ctor_set(x_44, 7, x_35); +lean_ctor_set(x_44, 8, x_36); +lean_ctor_set(x_44, 9, x_37); +lean_ctor_set(x_44, 10, x_38); +lean_ctor_set(x_44, 11, x_39); +lean_ctor_set(x_44, 12, x_40); +lean_ctor_set(x_44, 13, x_43); +lean_ctor_set_uint8(x_44, sizeof(void*)*14, x_33); +x_45 = lean_st_ref_set(x_4, x_44, x_15); +x_46 = lean_ctor_get(x_45, 1); +lean_inc(x_46); +if (lean_is_exclusive(x_45)) { + lean_ctor_release(x_45, 0); + lean_ctor_release(x_45, 1); + x_47 = x_45; +} else { + lean_dec_ref(x_45); + x_47 = lean_box(0); +} +x_48 = lean_box(0); +if (lean_is_scalar(x_47)) { + x_49 = lean_alloc_ctor(0, 2, 0); +} else { + x_49 = x_47; +} +lean_ctor_set(x_49, 0, x_48); +lean_ctor_set(x_49, 1, x_46); +return x_49; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addNewFact___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +x_13 = l_Lean_Meta_Grind_addNewFact(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addTheoremInstance(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, 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; uint8_t x_18; +x_13 = l_Lean_Meta_Grind_addNewFact(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_14 = lean_ctor_get(x_13, 1); +lean_inc(x_14); +lean_dec(x_13); +x_15 = lean_st_ref_take(x_4, x_14); +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_15, 1); +lean_inc(x_17); +lean_dec(x_15); +x_18 = !lean_is_exclusive(x_16); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; +x_19 = lean_ctor_get(x_16, 11); +x_20 = lean_unsigned_to_nat(1u); +x_21 = lean_nat_add(x_19, x_20); +lean_dec(x_19); +lean_ctor_set(x_16, 11, x_21); +x_22 = lean_st_ref_set(x_4, x_16, x_17); +x_23 = !lean_is_exclusive(x_22); +if (x_23 == 0) +{ +lean_object* x_24; lean_object* x_25; +x_24 = lean_ctor_get(x_22, 0); +lean_dec(x_24); +x_25 = lean_box(0); +lean_ctor_set(x_22, 0, x_25); +return x_22; +} +else +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_26 = lean_ctor_get(x_22, 1); +lean_inc(x_26); +lean_dec(x_22); +x_27 = lean_box(0); +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_27); +lean_ctor_set(x_28, 1, x_26); +return x_28; +} +} +else +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_29 = lean_ctor_get(x_16, 0); +x_30 = lean_ctor_get(x_16, 1); +x_31 = lean_ctor_get(x_16, 2); +x_32 = lean_ctor_get(x_16, 3); +x_33 = lean_ctor_get(x_16, 4); +x_34 = lean_ctor_get(x_16, 5); +x_35 = lean_ctor_get_uint8(x_16, sizeof(void*)*14); +x_36 = lean_ctor_get(x_16, 6); +x_37 = lean_ctor_get(x_16, 7); +x_38 = lean_ctor_get(x_16, 8); +x_39 = lean_ctor_get(x_16, 9); +x_40 = lean_ctor_get(x_16, 10); +x_41 = lean_ctor_get(x_16, 11); +x_42 = lean_ctor_get(x_16, 12); +x_43 = lean_ctor_get(x_16, 13); +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_34); +lean_inc(x_33); +lean_inc(x_32); +lean_inc(x_31); +lean_inc(x_30); +lean_inc(x_29); +lean_dec(x_16); +x_44 = lean_unsigned_to_nat(1u); +x_45 = lean_nat_add(x_41, x_44); +lean_dec(x_41); +x_46 = lean_alloc_ctor(0, 14, 1); +lean_ctor_set(x_46, 0, x_29); +lean_ctor_set(x_46, 1, x_30); +lean_ctor_set(x_46, 2, x_31); +lean_ctor_set(x_46, 3, x_32); +lean_ctor_set(x_46, 4, x_33); +lean_ctor_set(x_46, 5, x_34); +lean_ctor_set(x_46, 6, x_36); +lean_ctor_set(x_46, 7, x_37); +lean_ctor_set(x_46, 8, x_38); +lean_ctor_set(x_46, 9, x_39); +lean_ctor_set(x_46, 10, x_40); +lean_ctor_set(x_46, 11, x_45); +lean_ctor_set(x_46, 12, x_42); +lean_ctor_set(x_46, 13, x_43); +lean_ctor_set_uint8(x_46, sizeof(void*)*14, x_35); +x_47 = lean_st_ref_set(x_4, x_46, x_17); +x_48 = lean_ctor_get(x_47, 1); +lean_inc(x_48); +if (lean_is_exclusive(x_47)) { + lean_ctor_release(x_47, 0); + lean_ctor_release(x_47, 1); + x_49 = x_47; +} else { + lean_dec_ref(x_47); + x_49 = lean_box(0); +} +x_50 = lean_box(0); +if (lean_is_scalar(x_49)) { + x_51 = lean_alloc_ctor(0, 2, 0); +} else { + x_51 = x_49; +} +lean_ctor_set(x_51, 0, x_50); +lean_ctor_set(x_51, 1, x_48); +return x_51; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addTheoremInstance___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +x_13 = l_Lean_Meta_Grind_addTheoremInstance(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_checkMaxInstancesExceeded(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_10 = lean_st_ref_get(x_1, 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_Meta_Grind_getConfig___rarg(x_3, x_4, x_5, x_6, x_7, x_8, x_12); +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; lean_object* x_19; +x_15 = lean_ctor_get(x_13, 0); +x_16 = lean_ctor_get(x_15, 3); +lean_inc(x_16); +lean_dec(x_15); +x_17 = lean_ctor_get(x_11, 11); +lean_inc(x_17); +lean_dec(x_11); +x_18 = lean_nat_dec_le(x_16, x_17); +lean_dec(x_17); +lean_dec(x_16); +x_19 = lean_box(x_18); +lean_ctor_set(x_13, 0, x_19); +return x_13; +} +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; lean_object* x_25; lean_object* x_26; +x_20 = lean_ctor_get(x_13, 0); +x_21 = lean_ctor_get(x_13, 1); +lean_inc(x_21); +lean_inc(x_20); +lean_dec(x_13); +x_22 = lean_ctor_get(x_20, 3); +lean_inc(x_22); +lean_dec(x_20); +x_23 = lean_ctor_get(x_11, 11); +lean_inc(x_23); +lean_dec(x_11); +x_24 = lean_nat_dec_le(x_22, x_23); +lean_dec(x_23); +lean_dec(x_22); +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_21); +return x_26; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_checkMaxInstancesExceeded___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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_Grind_checkMaxInstancesExceeded(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); +lean_dec(x_1); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getENode_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; uint8_t x_12; +x_11 = lean_st_ref_get(x_2, x_10); +x_12 = !lean_is_exclusive(x_11); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; x_13 = lean_ctor_get(x_11, 0); x_14 = lean_ctor_get(x_13, 1); lean_inc(x_14); lean_dec(x_13); -x_15 = lean_ptr_addr(x_1); -x_16 = l_Lean_PersistentHashMap_find_x3f___at___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot___spec__1(x_14, x_15); -lean_ctor_set(x_11, 0, x_16); +x_15 = l_Lean_PersistentHashMap_find_x3f___at___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot___spec__1(x_14, x_1); +lean_ctor_set(x_11, 0, x_15); return x_11; } else { -lean_object* x_17; lean_object* x_18; lean_object* x_19; size_t x_20; lean_object* x_21; lean_object* x_22; -x_17 = lean_ctor_get(x_11, 0); -x_18 = lean_ctor_get(x_11, 1); -lean_inc(x_18); +lean_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_11, 0); +x_17 = lean_ctor_get(x_11, 1); lean_inc(x_17); +lean_inc(x_16); lean_dec(x_11); -x_19 = lean_ctor_get(x_17, 1); -lean_inc(x_19); -lean_dec(x_17); -x_20 = lean_ptr_addr(x_1); -x_21 = l_Lean_PersistentHashMap_find_x3f___at___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot___spec__1(x_19, x_20); -x_22 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_22, 0, x_21); -lean_ctor_set(x_22, 1, x_18); -return x_22; +x_18 = lean_ctor_get(x_16, 1); +lean_inc(x_18); +lean_dec(x_16); +x_19 = l_Lean_PersistentHashMap_find_x3f___at___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot___spec__1(x_18, x_1); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_19); +lean_ctor_set(x_20, 1, x_17); +return x_20; } } } @@ -4948,24 +8886,6 @@ lean_dec(x_1); return x_11; } } -LEAN_EXPORT size_t l_Lean_Meta_Grind_getENode_unsafe__1(lean_object* x_1) { -_start: -{ -size_t x_2; -x_2 = lean_ptr_addr(x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getENode_unsafe__1___boxed(lean_object* x_1) { -_start: -{ -size_t x_2; lean_object* x_3; -x_2 = l_Lean_Meta_Grind_getENode_unsafe__1(x_1); -lean_dec(x_1); -x_3 = lean_box_usize(x_2); -return x_3; -} -} LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_Grind_getENode___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: { @@ -5046,80 +8966,78 @@ x_11 = lean_st_ref_get(x_2, x_10); x_12 = !lean_is_exclusive(x_11); if (x_12 == 0) { -lean_object* x_13; lean_object* x_14; lean_object* x_15; size_t x_16; lean_object* x_17; +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; x_13 = lean_ctor_get(x_11, 0); x_14 = lean_ctor_get(x_11, 1); x_15 = lean_ctor_get(x_13, 1); lean_inc(x_15); lean_dec(x_13); -x_16 = lean_ptr_addr(x_1); -x_17 = l_Lean_PersistentHashMap_find_x3f___at___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot___spec__1(x_15, x_16); -if (lean_obj_tag(x_17) == 0) +x_16 = l_Lean_PersistentHashMap_find_x3f___at___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot___spec__1(x_15, x_1); +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_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_11); -x_18 = l_Lean_indentExpr(x_1); -x_19 = l_Lean_Meta_Grind_getENode___closed__2; -x_20 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_18); -x_21 = l_Lean_Meta_Grind_getENode___closed__4; -x_22 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_22, 0, x_20); -lean_ctor_set(x_22, 1, x_21); -x_23 = l_Lean_throwError___at_Lean_Meta_Grind_getENode___spec__1(x_22, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_14); -return x_23; +x_17 = l_Lean_indentExpr(x_1); +x_18 = l_Lean_Meta_Grind_getENode___closed__2; +x_19 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_17); +x_20 = l_Lean_Meta_Grind_getENode___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_Meta_Grind_getENode___spec__1(x_21, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_14); +return x_22; } else { -lean_object* x_24; +lean_object* x_23; lean_dec(x_1); -x_24 = lean_ctor_get(x_17, 0); -lean_inc(x_24); -lean_dec(x_17); -lean_ctor_set(x_11, 0, x_24); +x_23 = lean_ctor_get(x_16, 0); +lean_inc(x_23); +lean_dec(x_16); +lean_ctor_set(x_11, 0, x_23); return x_11; } } else { -lean_object* x_25; lean_object* x_26; lean_object* x_27; size_t x_28; lean_object* x_29; -x_25 = lean_ctor_get(x_11, 0); -x_26 = lean_ctor_get(x_11, 1); -lean_inc(x_26); +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +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_27 = lean_ctor_get(x_25, 1); -lean_inc(x_27); -lean_dec(x_25); -x_28 = lean_ptr_addr(x_1); -x_29 = l_Lean_PersistentHashMap_find_x3f___at___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot___spec__1(x_27, x_28); -if (lean_obj_tag(x_29) == 0) +x_26 = lean_ctor_get(x_24, 1); +lean_inc(x_26); +lean_dec(x_24); +x_27 = l_Lean_PersistentHashMap_find_x3f___at___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot___spec__1(x_26, x_1); +if (lean_obj_tag(x_27) == 0) { -lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_30 = l_Lean_indentExpr(x_1); -x_31 = l_Lean_Meta_Grind_getENode___closed__2; +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 = l_Lean_indentExpr(x_1); +x_29 = l_Lean_Meta_Grind_getENode___closed__2; +x_30 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_28); +x_31 = l_Lean_Meta_Grind_getENode___closed__4; x_32 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_32, 0, x_31); -lean_ctor_set(x_32, 1, x_30); -x_33 = l_Lean_Meta_Grind_getENode___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_Meta_Grind_getENode___spec__1(x_34, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_26); -return x_35; +lean_ctor_set(x_32, 0, x_30); +lean_ctor_set(x_32, 1, x_31); +x_33 = l_Lean_throwError___at_Lean_Meta_Grind_getENode___spec__1(x_32, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_25); +return x_33; } else { -lean_object* x_36; lean_object* x_37; +lean_object* x_34; lean_object* x_35; lean_dec(x_1); -x_36 = lean_ctor_get(x_29, 0); -lean_inc(x_36); -lean_dec(x_29); -x_37 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_37, 0, x_36); -lean_ctor_set(x_37, 1, x_26); -return x_37; +x_34 = lean_ctor_get(x_27, 0); +lean_inc(x_34); +lean_dec(x_27); +x_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_34); +lean_ctor_set(x_35, 1, x_25); +return x_35; } } } @@ -5974,25 +9892,7 @@ lean_dec(x_2); return x_11; } } -LEAN_EXPORT size_t l_Lean_Meta_Grind_alreadyInternalized_unsafe__1(lean_object* x_1) { -_start: -{ -size_t x_2; -x_2 = lean_ptr_addr(x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_alreadyInternalized_unsafe__1___boxed(lean_object* x_1) { -_start: -{ -size_t x_2; lean_object* x_3; -x_2 = l_Lean_Meta_Grind_alreadyInternalized_unsafe__1(x_1); -lean_dec(x_1); -x_3 = lean_box_usize(x_2); -return x_3; -} -} -LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_alreadyInternalized___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5) { +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_alreadyInternalized___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; @@ -6008,32 +9908,31 @@ return x_8; } else { -lean_object* x_9; size_t x_10; uint8_t x_11; +lean_object* x_9; uint8_t x_10; x_9 = lean_array_fget(x_1, x_4); -x_10 = lean_unbox_usize(x_9); +x_10 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_5, x_9); lean_dec(x_9); -x_11 = lean_usize_dec_eq(x_5, x_10); -if (x_11 == 0) +if (x_10 == 0) { -lean_object* x_12; lean_object* x_13; -x_12 = lean_unsigned_to_nat(1u); -x_13 = lean_nat_add(x_4, x_12); +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_13; +x_4 = x_12; goto _start; } else { -uint8_t x_15; +uint8_t x_14; lean_dec(x_4); -x_15 = 1; -return x_15; +x_14 = 1; +return x_14; } } } } -LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_alreadyInternalized___spec__2(lean_object* x_1, size_t x_2, size_t x_3) { +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_alreadyInternalized___spec__2(lean_object* x_1, size_t x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_1) == 0) @@ -6053,55 +9952,54 @@ lean_dec(x_4); switch (lean_obj_tag(x_10)) { case 0: { -lean_object* x_11; size_t x_12; uint8_t x_13; +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_unbox_usize(x_11); +x_12 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_3, x_11); lean_dec(x_11); -x_13 = lean_usize_dec_eq(x_3, x_12); -return x_13; +return x_12; } case 1: { -lean_object* x_14; size_t x_15; -x_14 = lean_ctor_get(x_10, 0); -lean_inc(x_14); +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_15 = lean_usize_shift_right(x_2, x_5); -x_1 = x_14; -x_2 = x_15; +x_14 = lean_usize_shift_right(x_2, x_5); +x_1 = x_13; +x_2 = x_14; goto _start; } default: { -uint8_t x_17; -x_17 = 0; -return x_17; +uint8_t x_16; +x_16 = 0; +return x_16; } } } else { -lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; -x_18 = lean_ctor_get(x_1, 0); +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); -x_19 = lean_ctor_get(x_1, 1); -lean_inc(x_19); lean_dec(x_1); -x_20 = lean_unsigned_to_nat(0u); -x_21 = l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_alreadyInternalized___spec__3(x_18, x_19, lean_box(0), x_20, x_3); -lean_dec(x_19); +x_19 = lean_unsigned_to_nat(0u); +x_20 = l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_alreadyInternalized___spec__3(x_17, x_18, lean_box(0), x_19, x_3); lean_dec(x_18); -return x_21; +lean_dec(x_17); +return x_20; } } } -LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_alreadyInternalized___spec__1(lean_object* x_1, size_t x_2) { +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_alreadyInternalized___spec__1(lean_object* x_1, lean_object* x_2) { _start: { uint64_t x_3; size_t x_4; uint8_t x_5; -x_3 = lean_usize_to_uint64(x_2); +x_3 = l_Lean_Meta_Grind_instHashableENodeKey_unsafe__1(x_2); x_4 = lean_uint64_to_usize(x_3); x_5 = l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_alreadyInternalized___spec__2(x_1, x_4, x_2); return x_5; @@ -6115,73 +10013,68 @@ x_11 = lean_st_ref_get(x_2, x_10); x_12 = !lean_is_exclusive(x_11); if (x_12 == 0) { -lean_object* x_13; lean_object* x_14; size_t x_15; uint8_t x_16; lean_object* x_17; +lean_object* x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; x_13 = lean_ctor_get(x_11, 0); x_14 = lean_ctor_get(x_13, 1); lean_inc(x_14); lean_dec(x_13); -x_15 = lean_ptr_addr(x_1); -x_16 = l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_alreadyInternalized___spec__1(x_14, x_15); -x_17 = lean_box(x_16); -lean_ctor_set(x_11, 0, x_17); +x_15 = l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_alreadyInternalized___spec__1(x_14, x_1); +x_16 = lean_box(x_15); +lean_ctor_set(x_11, 0, x_16); return x_11; } else { -lean_object* x_18; lean_object* x_19; lean_object* x_20; size_t x_21; uint8_t x_22; lean_object* x_23; lean_object* x_24; -x_18 = lean_ctor_get(x_11, 0); -x_19 = lean_ctor_get(x_11, 1); -lean_inc(x_19); +lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; lean_object* x_21; lean_object* x_22; +x_17 = lean_ctor_get(x_11, 0); +x_18 = lean_ctor_get(x_11, 1); lean_inc(x_18); +lean_inc(x_17); lean_dec(x_11); -x_20 = lean_ctor_get(x_18, 1); -lean_inc(x_20); -lean_dec(x_18); -x_21 = lean_ptr_addr(x_1); -x_22 = l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_alreadyInternalized___spec__1(x_20, x_21); -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_19); -return x_24; +x_19 = lean_ctor_get(x_17, 1); +lean_inc(x_19); +lean_dec(x_17); +x_20 = l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_alreadyInternalized___spec__1(x_19, x_1); +x_21 = lean_box(x_20); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_18); +return x_22; } } } LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_alreadyInternalized___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -size_t x_6; uint8_t x_7; lean_object* x_8; -x_6 = lean_unbox_usize(x_5); +uint8_t x_6; lean_object* x_7; +x_6 = l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_alreadyInternalized___spec__3(x_1, x_2, x_3, x_4, x_5); lean_dec(x_5); -x_7 = l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_alreadyInternalized___spec__3(x_1, x_2, x_3, x_4, x_6); lean_dec(x_2); lean_dec(x_1); -x_8 = lean_box(x_7); -return x_8; +x_7 = lean_box(x_6); +return x_7; } } LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_alreadyInternalized___spec__2___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; +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 = lean_unbox_usize(x_3); +x_5 = l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_alreadyInternalized___spec__2(x_1, x_4, x_3); lean_dec(x_3); -x_6 = l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_alreadyInternalized___spec__2(x_1, x_4, x_5); -x_7 = lean_box(x_6); -return x_7; +x_6 = lean_box(x_5); +return x_6; } } LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_alreadyInternalized___spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { -size_t x_3; uint8_t x_4; lean_object* x_5; -x_3 = lean_unbox_usize(x_2); +uint8_t x_3; lean_object* x_4; +x_3 = l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_alreadyInternalized___spec__1(x_1, x_2); lean_dec(x_2); -x_4 = l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_alreadyInternalized___spec__1(x_1, x_3); -x_5 = lean_box(x_4); -return x_5; +x_4 = lean_box(x_3); +return x_4; } } LEAN_EXPORT lean_object* l_Lean_Meta_Grind_alreadyInternalized___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { @@ -6337,16 +10230,28 @@ return x_27; } else { -lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; 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_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; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; 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_28 = lean_ctor_get(x_15, 0); x_29 = lean_ctor_get(x_15, 1); x_30 = lean_ctor_get(x_15, 2); x_31 = lean_ctor_get(x_15, 3); x_32 = lean_ctor_get(x_15, 4); x_33 = lean_ctor_get(x_15, 5); -x_34 = lean_ctor_get_uint8(x_15, sizeof(void*)*8); +x_34 = lean_ctor_get_uint8(x_15, sizeof(void*)*14); x_35 = lean_ctor_get(x_15, 6); x_36 = lean_ctor_get(x_15, 7); +x_37 = lean_ctor_get(x_15, 8); +x_38 = lean_ctor_get(x_15, 9); +x_39 = lean_ctor_get(x_15, 10); +x_40 = lean_ctor_get(x_15, 11); +x_41 = lean_ctor_get(x_15, 12); +x_42 = lean_ctor_get(x_15, 13); +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_inc(x_33); @@ -6356,42 +10261,48 @@ lean_inc(x_30); lean_inc(x_29); lean_inc(x_28); lean_dec(x_15); -x_37 = lean_alloc_ctor(0, 3, 1); -lean_ctor_set(x_37, 0, x_1); -lean_ctor_set(x_37, 1, x_2); -lean_ctor_set(x_37, 2, x_3); -lean_ctor_set_uint8(x_37, sizeof(void*)*3, x_4); -x_38 = lean_array_push(x_33, x_37); -x_39 = lean_alloc_ctor(0, 8, 1); -lean_ctor_set(x_39, 0, x_28); -lean_ctor_set(x_39, 1, x_29); -lean_ctor_set(x_39, 2, x_30); -lean_ctor_set(x_39, 3, x_31); -lean_ctor_set(x_39, 4, x_32); -lean_ctor_set(x_39, 5, x_38); -lean_ctor_set(x_39, 6, x_35); -lean_ctor_set(x_39, 7, x_36); -lean_ctor_set_uint8(x_39, sizeof(void*)*8, x_34); -x_40 = lean_st_ref_set(x_5, x_39, x_16); -x_41 = lean_ctor_get(x_40, 1); -lean_inc(x_41); -if (lean_is_exclusive(x_40)) { - lean_ctor_release(x_40, 0); - lean_ctor_release(x_40, 1); - x_42 = x_40; +x_43 = lean_alloc_ctor(0, 3, 1); +lean_ctor_set(x_43, 0, x_1); +lean_ctor_set(x_43, 1, x_2); +lean_ctor_set(x_43, 2, x_3); +lean_ctor_set_uint8(x_43, sizeof(void*)*3, x_4); +x_44 = lean_array_push(x_33, x_43); +x_45 = lean_alloc_ctor(0, 14, 1); +lean_ctor_set(x_45, 0, x_28); +lean_ctor_set(x_45, 1, x_29); +lean_ctor_set(x_45, 2, x_30); +lean_ctor_set(x_45, 3, x_31); +lean_ctor_set(x_45, 4, x_32); +lean_ctor_set(x_45, 5, x_44); +lean_ctor_set(x_45, 6, x_35); +lean_ctor_set(x_45, 7, x_36); +lean_ctor_set(x_45, 8, x_37); +lean_ctor_set(x_45, 9, x_38); +lean_ctor_set(x_45, 10, x_39); +lean_ctor_set(x_45, 11, x_40); +lean_ctor_set(x_45, 12, x_41); +lean_ctor_set(x_45, 13, x_42); +lean_ctor_set_uint8(x_45, sizeof(void*)*14, x_34); +x_46 = lean_st_ref_set(x_5, x_45, x_16); +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_40); - x_42 = lean_box(0); + lean_dec_ref(x_46); + x_48 = lean_box(0); } -x_43 = lean_box(0); -if (lean_is_scalar(x_42)) { - x_44 = lean_alloc_ctor(0, 2, 0); +x_49 = lean_box(0); +if (lean_is_scalar(x_48)) { + x_50 = lean_alloc_ctor(0, 2, 0); } else { - x_44 = x_42; + x_50 = x_48; } -lean_ctor_set(x_44, 0, x_43); -lean_ctor_set(x_44, 1, x_41); -return x_44; +lean_ctor_set(x_50, 0, x_49); +lean_ctor_set(x_50, 1, x_47); +return x_50; } } } @@ -7180,204 +11091,7 @@ lean_dec(x_3); return x_12; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_registerParent___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t 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; size_t x_10; uint8_t x_11; -x_9 = lean_array_fget(x_1, x_4); -x_10 = lean_unbox_usize(x_9); -lean_dec(x_9); -x_11 = lean_usize_dec_eq(x_5, x_10); -if (x_11 == 0) -{ -lean_object* x_12; lean_object* x_13; -x_12 = lean_unsigned_to_nat(1u); -x_13 = lean_nat_add(x_4, x_12); -lean_dec(x_4); -x_3 = lean_box(0); -x_4 = x_13; -goto _start; -} -else -{ -lean_object* x_15; lean_object* x_16; -x_15 = lean_array_fget(x_2, x_4); -lean_dec(x_4); -x_16 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_16, 0, x_15); -return x_16; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_registerParent___spec__2(lean_object* x_1, size_t x_2, size_t x_3) { -_start: -{ -if (lean_obj_tag(x_1) == 0) -{ -uint8_t x_4; -x_4 = !lean_is_exclusive(x_1); -if (x_4 == 0) -{ -lean_object* x_5; size_t x_6; size_t x_7; size_t x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; -x_5 = lean_ctor_get(x_1, 0); -x_6 = 5; -x_7 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__2___closed__2; -x_8 = lean_usize_land(x_2, x_7); -x_9 = lean_usize_to_nat(x_8); -x_10 = lean_box(2); -x_11 = lean_array_get(x_10, x_5, x_9); -lean_dec(x_9); -lean_dec(x_5); -switch (lean_obj_tag(x_11)) { -case 0: -{ -lean_object* x_12; lean_object* x_13; size_t x_14; uint8_t x_15; -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); -x_13 = lean_ctor_get(x_11, 1); -lean_inc(x_13); -lean_dec(x_11); -x_14 = lean_unbox_usize(x_12); -lean_dec(x_12); -x_15 = lean_usize_dec_eq(x_3, x_14); -if (x_15 == 0) -{ -lean_object* x_16; -lean_dec(x_13); -lean_free_object(x_1); -x_16 = lean_box(0); -return x_16; -} -else -{ -lean_ctor_set_tag(x_1, 1); -lean_ctor_set(x_1, 0, x_13); -return x_1; -} -} -case 1: -{ -lean_object* x_17; size_t x_18; -lean_free_object(x_1); -x_17 = lean_ctor_get(x_11, 0); -lean_inc(x_17); -lean_dec(x_11); -x_18 = lean_usize_shift_right(x_2, x_6); -x_1 = x_17; -x_2 = x_18; -goto _start; -} -default: -{ -lean_object* x_20; -lean_free_object(x_1); -x_20 = lean_box(0); -return x_20; -} -} -} -else -{ -lean_object* x_21; size_t x_22; size_t x_23; size_t x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_21 = lean_ctor_get(x_1, 0); -lean_inc(x_21); -lean_dec(x_1); -x_22 = 5; -x_23 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__2___closed__2; -x_24 = lean_usize_land(x_2, x_23); -x_25 = lean_usize_to_nat(x_24); -x_26 = lean_box(2); -x_27 = lean_array_get(x_26, x_21, x_25); -lean_dec(x_25); -lean_dec(x_21); -switch (lean_obj_tag(x_27)) { -case 0: -{ -lean_object* x_28; lean_object* x_29; size_t 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_unbox_usize(x_28); -lean_dec(x_28); -x_31 = lean_usize_dec_eq(x_3, x_30); -if (x_31 == 0) -{ -lean_object* x_32; -lean_dec(x_29); -x_32 = lean_box(0); -return x_32; -} -else -{ -lean_object* x_33; -x_33 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_33, 0, x_29); -return x_33; -} -} -case 1: -{ -lean_object* x_34; size_t x_35; -x_34 = lean_ctor_get(x_27, 0); -lean_inc(x_34); -lean_dec(x_27); -x_35 = lean_usize_shift_right(x_2, x_22); -x_1 = x_34; -x_2 = x_35; -goto _start; -} -default: -{ -lean_object* x_37; -x_37 = lean_box(0); -return x_37; -} -} -} -} -else -{ -lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; -x_38 = lean_ctor_get(x_1, 0); -lean_inc(x_38); -x_39 = lean_ctor_get(x_1, 1); -lean_inc(x_39); -lean_dec(x_1); -x_40 = lean_unsigned_to_nat(0u); -x_41 = l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_registerParent___spec__3(x_38, x_39, lean_box(0), x_40, x_3); -lean_dec(x_39); -lean_dec(x_38); -return x_41; -} -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_registerParent___spec__1(lean_object* x_1, size_t x_2) { -_start: -{ -uint64_t x_3; size_t x_4; lean_object* x_5; -x_3 = lean_usize_to_uint64(x_2); -x_4 = lean_uint64_to_usize(x_3); -x_5 = l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_registerParent___spec__2(x_1, x_4, x_2); -return x_5; -} -} -LEAN_EXPORT lean_object* l_Lean_RBNode_ins___at_Lean_Meta_Grind_registerParent___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_RBNode_ins___at_Lean_Meta_Grind_registerParent___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_1) == 0) @@ -7413,7 +11127,7 @@ switch (x_13) { case 0: { lean_object* x_14; uint8_t x_15; -x_14 = l_Lean_RBNode_ins___at_Lean_Meta_Grind_registerParent___spec__5(x_9, x_2, x_3); +x_14 = l_Lean_RBNode_ins___at_Lean_Meta_Grind_registerParent___spec__2(x_9, x_2, x_3); x_15 = 0; lean_ctor_set(x_1, 0, x_14); lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_15); @@ -7433,7 +11147,7 @@ return x_1; default: { lean_object* x_17; uint8_t x_18; -x_17 = l_Lean_RBNode_ins___at_Lean_Meta_Grind_registerParent___spec__5(x_12, x_2, x_3); +x_17 = l_Lean_RBNode_ins___at_Lean_Meta_Grind_registerParent___spec__2(x_12, x_2, x_3); x_18 = 0; lean_ctor_set(x_1, 3, x_17); lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_18); @@ -7458,7 +11172,7 @@ switch (x_23) { case 0: { lean_object* x_24; uint8_t x_25; lean_object* x_26; -x_24 = l_Lean_RBNode_ins___at_Lean_Meta_Grind_registerParent___spec__5(x_19, x_2, x_3); +x_24 = l_Lean_RBNode_ins___at_Lean_Meta_Grind_registerParent___spec__2(x_19, x_2, x_3); x_25 = 0; x_26 = lean_alloc_ctor(1, 4, 1); lean_ctor_set(x_26, 0, x_24); @@ -7485,7 +11199,7 @@ return x_28; default: { lean_object* x_29; uint8_t x_30; lean_object* x_31; -x_29 = l_Lean_RBNode_ins___at_Lean_Meta_Grind_registerParent___spec__5(x_22, x_2, x_3); +x_29 = l_Lean_RBNode_ins___at_Lean_Meta_Grind_registerParent___spec__2(x_22, x_2, x_3); x_30 = 0; x_31 = lean_alloc_ctor(1, 4, 1); lean_ctor_set(x_31, 0, x_19); @@ -7514,7 +11228,7 @@ switch (x_37) { case 0: { lean_object* x_38; uint8_t x_39; -x_38 = l_Lean_RBNode_ins___at_Lean_Meta_Grind_registerParent___spec__5(x_33, x_2, x_3); +x_38 = l_Lean_RBNode_ins___at_Lean_Meta_Grind_registerParent___spec__2(x_33, x_2, x_3); x_39 = lean_ctor_get_uint8(x_38, sizeof(void*)*4); if (x_39 == 0) { @@ -8213,7 +11927,7 @@ return x_1; default: { lean_object* x_191; uint8_t x_192; -x_191 = l_Lean_RBNode_ins___at_Lean_Meta_Grind_registerParent___spec__5(x_36, x_2, x_3); +x_191 = l_Lean_RBNode_ins___at_Lean_Meta_Grind_registerParent___spec__2(x_36, x_2, x_3); x_192 = lean_ctor_get_uint8(x_191, sizeof(void*)*4); if (x_192 == 0) { @@ -8908,7 +12622,7 @@ switch (x_346) { case 0: { lean_object* x_347; uint8_t x_348; -x_347 = l_Lean_RBNode_ins___at_Lean_Meta_Grind_registerParent___spec__5(x_342, x_2, x_3); +x_347 = l_Lean_RBNode_ins___at_Lean_Meta_Grind_registerParent___spec__2(x_342, x_2, x_3); x_348 = lean_ctor_get_uint8(x_347, sizeof(void*)*4); if (x_348 == 0) { @@ -9337,7 +13051,7 @@ return x_423; default: { lean_object* x_424; uint8_t x_425; -x_424 = l_Lean_RBNode_ins___at_Lean_Meta_Grind_registerParent___spec__5(x_345, x_2, x_3); +x_424 = l_Lean_RBNode_ins___at_Lean_Meta_Grind_registerParent___spec__2(x_345, x_2, x_3); x_425 = lean_ctor_get_uint8(x_424, sizeof(void*)*4); if (x_425 == 0) { @@ -9755,7 +13469,7 @@ return x_498; } } } -LEAN_EXPORT lean_object* l_Lean_RBNode_insert___at_Lean_Meta_Grind_registerParent___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_RBNode_insert___at_Lean_Meta_Grind_registerParent___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { uint8_t x_4; @@ -9763,19 +13477,19 @@ x_4 = l_Lean_RBNode_isRed___rarg(x_1); if (x_4 == 0) { lean_object* x_5; -x_5 = l_Lean_RBNode_ins___at_Lean_Meta_Grind_registerParent___spec__5(x_1, x_2, x_3); +x_5 = l_Lean_RBNode_ins___at_Lean_Meta_Grind_registerParent___spec__2(x_1, x_2, x_3); return x_5; } else { lean_object* x_6; lean_object* x_7; -x_6 = l_Lean_RBNode_ins___at_Lean_Meta_Grind_registerParent___spec__5(x_1, x_2, x_3); +x_6 = l_Lean_RBNode_ins___at_Lean_Meta_Grind_registerParent___spec__2(x_1, x_2, x_3); x_7 = l_Lean_RBNode_setBlack___rarg(x_6); return x_7; } } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_registerParent___spec__8(size_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_registerParent___spec__5(size_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { lean_object* x_7; uint8_t x_8; @@ -9789,30 +13503,28 @@ return x_6; } else { -lean_object* x_9; size_t x_10; lean_object* x_11; uint64_t x_12; size_t x_13; size_t x_14; size_t x_15; size_t x_16; size_t x_17; size_t x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +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_unbox_usize(x_9); -lean_dec(x_9); -x_11 = lean_array_fget(x_3, x_5); -x_12 = lean_usize_to_uint64(x_10); -x_13 = lean_uint64_to_usize(x_12); -x_14 = 1; -x_15 = lean_usize_sub(x_1, x_14); -x_16 = 5; -x_17 = lean_usize_mul(x_16, x_15); -x_18 = lean_usize_shift_right(x_13, x_17); -x_19 = lean_unsigned_to_nat(1u); -x_20 = lean_nat_add(x_5, x_19); +x_10 = lean_array_fget(x_3, x_5); +x_11 = l_Lean_Meta_Grind_instHashableENodeKey_unsafe__1(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_21 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_registerParent___spec__7(x_6, x_18, x_1, x_10, x_11); +x_20 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_registerParent___spec__4(x_6, x_17, x_1, x_9, x_10); x_4 = lean_box(0); -x_5 = x_20; -x_6 = x_21; +x_5 = x_19; +x_6 = x_20; goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_registerParent___spec__9(lean_object* x_1, lean_object* x_2, size_t x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_registerParent___spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; @@ -9830,86 +13542,81 @@ 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; lean_object* x_14; +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_box_usize(x_3); -x_13 = lean_array_push(x_5, x_12); -x_14 = lean_array_push(x_6, x_4); -lean_ctor_set(x_1, 1, x_14); -lean_ctor_set(x_1, 0, x_13); +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_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_dec(x_1); -x_15 = lean_box_usize(x_3); -x_16 = lean_array_push(x_5, x_15); -x_17 = lean_array_push(x_6, x_4); -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; +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_19; size_t x_20; uint8_t x_21; -x_19 = lean_array_fget(x_5, x_2); -x_20 = lean_unbox_usize(x_19); -lean_dec(x_19); -x_21 = lean_usize_dec_eq(x_3, x_20); -if (x_21 == 0) +lean_object* x_17; uint8_t x_18; +x_17 = lean_array_fget(x_5, x_2); +x_18 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_3, x_17); +lean_dec(x_17); +if (x_18 == 0) { -lean_object* x_22; lean_object* x_23; +lean_object* x_19; lean_object* x_20; lean_dec(x_6); lean_dec(x_5); -x_22 = lean_unsigned_to_nat(1u); -x_23 = lean_nat_add(x_2, x_22); +x_19 = lean_unsigned_to_nat(1u); +x_20 = lean_nat_add(x_2, x_19); lean_dec(x_2); -x_2 = x_23; +x_2 = x_20; goto _start; } else { -uint8_t x_25; -x_25 = !lean_is_exclusive(x_1); -if (x_25 == 0) +uint8_t x_22; +x_22 = !lean_is_exclusive(x_1); +if (x_22 == 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_1, 1); -lean_dec(x_26); -x_27 = lean_ctor_get(x_1, 0); -lean_dec(x_27); -x_28 = lean_box_usize(x_3); -x_29 = lean_array_fset(x_5, x_2, x_28); -x_30 = lean_array_fset(x_6, x_2, x_4); +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_30); -lean_ctor_set(x_1, 0, x_29); +lean_ctor_set(x_1, 1, x_26); +lean_ctor_set(x_1, 0, x_25); return x_1; } else { -lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_dec(x_1); -x_31 = lean_box_usize(x_3); -x_32 = lean_array_fset(x_5, x_2, x_31); -x_33 = lean_array_fset(x_6, x_2, x_4); +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_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_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_PersistentHashMap_insertAux___at_Lean_Meta_Grind_registerParent___spec__7(lean_object* x_1, size_t x_2, size_t x_3, size_t x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_registerParent___spec__4(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) @@ -9932,6 +13639,7 @@ if (x_14 == 0) { lean_dec(x_12); lean_dec(x_5); +lean_dec(x_4); return x_1; } else @@ -9947,120 +13655,113 @@ uint8_t x_18; x_18 = !lean_is_exclusive(x_15); if (x_18 == 0) { -lean_object* x_19; lean_object* x_20; size_t x_21; uint8_t x_22; +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_unbox_usize(x_19); -x_22 = lean_usize_dec_eq(x_4, x_21); -if (x_22 == 0) +x_21 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_4, x_19); +if (x_21 == 0) { -lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_free_object(x_15); -x_23 = lean_box_usize(x_4); -x_24 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_19, x_20, x_23, x_5); -x_25 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_25, 0, x_24); -x_26 = lean_array_fset(x_17, x_12, x_25); +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_26); +lean_ctor_set(x_1, 0, x_24); return x_1; } else { -lean_object* x_27; lean_object* x_28; +lean_object* x_25; lean_dec(x_20); lean_dec(x_19); -x_27 = lean_box_usize(x_4); lean_ctor_set(x_15, 1, x_5); -lean_ctor_set(x_15, 0, x_27); -x_28 = lean_array_fset(x_17, x_12, x_15); +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_28); +lean_ctor_set(x_1, 0, x_25); return x_1; } } else { -lean_object* x_29; lean_object* x_30; size_t x_31; uint8_t x_32; -x_29 = lean_ctor_get(x_15, 0); -x_30 = lean_ctor_get(x_15, 1); -lean_inc(x_30); -lean_inc(x_29); +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_31 = lean_unbox_usize(x_29); -x_32 = lean_usize_dec_eq(x_4, x_31); -if (x_32 == 0) +x_28 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_4, x_26); +if (x_28 == 0) { -lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_33 = lean_box_usize(x_4); -x_34 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_29, x_30, x_33, x_5); -x_35 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_35, 0, x_34); -x_36 = lean_array_fset(x_17, x_12, x_35); +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_36); +lean_ctor_set(x_1, 0, x_31); return x_1; } else { -lean_object* x_37; lean_object* x_38; lean_object* x_39; -lean_dec(x_30); -lean_dec(x_29); -x_37 = lean_box_usize(x_4); -x_38 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_38, 0, x_37); -lean_ctor_set(x_38, 1, x_5); -x_39 = lean_array_fset(x_17, x_12, x_38); +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_39); +lean_ctor_set(x_1, 0, x_33); return x_1; } } } case 1: { -uint8_t x_40; -x_40 = !lean_is_exclusive(x_15); -if (x_40 == 0) +uint8_t x_34; +x_34 = !lean_is_exclusive(x_15); +if (x_34 == 0) { -lean_object* x_41; size_t x_42; size_t x_43; lean_object* x_44; lean_object* x_45; -x_41 = lean_ctor_get(x_15, 0); -x_42 = lean_usize_shift_right(x_2, x_9); -x_43 = lean_usize_add(x_3, x_8); -x_44 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_registerParent___spec__7(x_41, x_42, x_43, x_4, x_5); -lean_ctor_set(x_15, 0, x_44); -x_45 = lean_array_fset(x_17, x_12, x_15); +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_Grind_registerParent___spec__4(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_45); +lean_ctor_set(x_1, 0, x_39); return x_1; } else { -lean_object* x_46; size_t x_47; size_t x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; -x_46 = lean_ctor_get(x_15, 0); -lean_inc(x_46); +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_47 = lean_usize_shift_right(x_2, x_9); -x_48 = lean_usize_add(x_3, x_8); -x_49 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_registerParent___spec__7(x_46, x_47, x_48, x_4, x_5); -x_50 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_50, 0, x_49); -x_51 = lean_array_fset(x_17, x_12, x_50); +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_Grind_registerParent___spec__4(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_51); +lean_ctor_set(x_1, 0, x_45); return x_1; } } default: { -lean_object* x_52; lean_object* x_53; lean_object* x_54; -x_52 = lean_box_usize(x_4); -x_53 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_53, 0, x_52); -lean_ctor_set(x_53, 1, x_5); -x_54 = lean_array_fset(x_17, x_12, x_53); +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_54); +lean_ctor_set(x_1, 0, x_47); return x_1; } } @@ -10068,124 +13769,121 @@ return x_1; } else { -lean_object* x_55; size_t x_56; size_t x_57; size_t x_58; size_t x_59; lean_object* x_60; lean_object* x_61; uint8_t x_62; -x_55 = lean_ctor_get(x_1, 0); -lean_inc(x_55); +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_56 = 1; -x_57 = 5; -x_58 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__2___closed__2; -x_59 = lean_usize_land(x_2, x_58); -x_60 = lean_usize_to_nat(x_59); -x_61 = lean_array_get_size(x_55); -x_62 = lean_nat_dec_lt(x_60, x_61); -lean_dec(x_61); -if (x_62 == 0) +x_49 = 1; +x_50 = 5; +x_51 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___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_63; -lean_dec(x_60); +lean_object* x_56; +lean_dec(x_53); lean_dec(x_5); -x_63 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_63, 0, x_55); -return x_63; +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_64; lean_object* x_65; lean_object* x_66; -x_64 = lean_array_fget(x_55, x_60); -x_65 = lean_box(0); -x_66 = lean_array_fset(x_55, x_60, x_65); -switch (lean_obj_tag(x_64)) { +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_67; lean_object* x_68; lean_object* x_69; size_t x_70; uint8_t x_71; -x_67 = lean_ctor_get(x_64, 0); -lean_inc(x_67); -x_68 = lean_ctor_get(x_64, 1); -lean_inc(x_68); -if (lean_is_exclusive(x_64)) { - lean_ctor_release(x_64, 0); - lean_ctor_release(x_64, 1); - x_69 = x_64; +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_64); - x_69 = lean_box(0); + lean_dec_ref(x_57); + x_62 = lean_box(0); } -x_70 = lean_unbox_usize(x_67); -x_71 = lean_usize_dec_eq(x_4, x_70); -if (x_71 == 0) +x_63 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_4, x_60); +if (x_63 == 0) { -lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; -lean_dec(x_69); -x_72 = lean_box_usize(x_4); -x_73 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_67, x_68, x_72, x_5); -x_74 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_74, 0, x_73); -x_75 = lean_array_fset(x_66, x_60, x_74); -lean_dec(x_60); -x_76 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_76, 0, x_75); -return x_76; +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_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; -lean_dec(x_68); -lean_dec(x_67); -x_77 = lean_box_usize(x_4); -if (lean_is_scalar(x_69)) { - x_78 = lean_alloc_ctor(0, 2, 0); +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_78 = x_69; + x_68 = x_62; } -lean_ctor_set(x_78, 0, x_77); -lean_ctor_set(x_78, 1, x_5); -x_79 = lean_array_fset(x_66, x_60, x_78); -lean_dec(x_60); -x_80 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_80, 0, x_79); -return x_80; +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_81; lean_object* x_82; size_t x_83; size_t x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; -x_81 = lean_ctor_get(x_64, 0); -lean_inc(x_81); -if (lean_is_exclusive(x_64)) { - lean_ctor_release(x_64, 0); - x_82 = x_64; +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_64); - x_82 = lean_box(0); + lean_dec_ref(x_57); + x_72 = lean_box(0); } -x_83 = lean_usize_shift_right(x_2, x_57); -x_84 = lean_usize_add(x_3, x_56); -x_85 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_registerParent___spec__7(x_81, x_83, x_84, x_4, x_5); -if (lean_is_scalar(x_82)) { - x_86 = lean_alloc_ctor(1, 1, 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_Grind_registerParent___spec__4(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_86 = x_82; + x_76 = x_72; } -lean_ctor_set(x_86, 0, x_85); -x_87 = lean_array_fset(x_66, x_60, x_86); -lean_dec(x_60); -x_88 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_88, 0, x_87); -return x_88; +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_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; -x_89 = lean_box_usize(x_4); -x_90 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_90, 0, x_89); -lean_ctor_set(x_90, 1, x_5); -x_91 = lean_array_fset(x_66, x_60, x_90); -lean_dec(x_60); -x_92 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_92, 0, x_91); -return x_92; +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; } } } @@ -10193,104 +13891,298 @@ return x_92; } else { -uint8_t x_93; -x_93 = !lean_is_exclusive(x_1); -if (x_93 == 0) +uint8_t x_82; +x_82 = !lean_is_exclusive(x_1); +if (x_82 == 0) { -lean_object* x_94; lean_object* x_95; size_t x_96; uint8_t x_97; -x_94 = lean_unsigned_to_nat(0u); -x_95 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_registerParent___spec__9(x_1, x_94, x_4, x_5); -x_96 = 7; -x_97 = lean_usize_dec_le(x_96, x_3); -if (x_97 == 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_Grind_registerParent___spec__6(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_98; lean_object* x_99; uint8_t x_100; -x_98 = l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(x_95); -x_99 = lean_unsigned_to_nat(4u); -x_100 = lean_nat_dec_lt(x_98, x_99); -lean_dec(x_98); +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_Grind_mkHCongrWithArity___spec__2___closed__3; +x_93 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_registerParent___spec__5(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_Grind_registerParent___spec__6(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; 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_inc(x_102); -lean_dec(x_95); -x_103 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__2___closed__3; -x_104 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_registerParent___spec__8(x_3, x_101, x_102, lean_box(0), x_94, x_103); -lean_dec(x_102); +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); -return x_104; +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_Grind_mkHCongrWithArity___spec__2___closed__3; +x_107 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_registerParent___spec__5(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_95; +return x_98; } } else { -return x_95; +return x_98; +} +} +} +} } +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_registerParent___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint64_t x_4; size_t x_5; size_t x_6; lean_object* x_7; +x_4 = l_Lean_Meta_Grind_instHashableENodeKey_unsafe__1(x_2); +x_5 = lean_uint64_to_usize(x_4); +x_6 = 1; +x_7 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_registerParent___spec__4(x_1, x_5, x_6, x_2, x_3); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_registerParent___spec__9(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; 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_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; size_t x_110; uint8_t x_111; -x_105 = lean_ctor_get(x_1, 0); -x_106 = lean_ctor_get(x_1, 1); -lean_inc(x_106); -lean_inc(x_105); -lean_dec(x_1); -x_107 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_107, 0, x_105); -lean_ctor_set(x_107, 1, x_106); -x_108 = lean_unsigned_to_nat(0u); -x_109 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_registerParent___spec__9(x_107, x_108, x_4, x_5); -x_110 = 7; -x_111 = lean_usize_dec_le(x_110, x_3); -if (x_111 == 0) +lean_object* x_9; uint8_t x_10; +x_9 = lean_array_fget(x_1, x_4); +x_10 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_5, x_9); +lean_dec(x_9); +if (x_10 == 0) { -lean_object* x_112; lean_object* x_113; uint8_t x_114; -x_112 = l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(x_109); -x_113 = lean_unsigned_to_nat(4u); -x_114 = lean_nat_dec_lt(x_112, x_113); -lean_dec(x_112); -if (x_114 == 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_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; -x_115 = lean_ctor_get(x_109, 0); -lean_inc(x_115); -x_116 = lean_ctor_get(x_109, 1); -lean_inc(x_116); -lean_dec(x_109); -x_117 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__2___closed__3; -x_118 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_registerParent___spec__8(x_3, x_115, x_116, lean_box(0), x_108, x_117); -lean_dec(x_116); -lean_dec(x_115); -return x_118; +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; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_registerParent___spec__8(lean_object* x_1, size_t x_2, lean_object* x_3) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +uint8_t x_4; +x_4 = !lean_is_exclusive(x_1); +if (x_4 == 0) +{ +lean_object* x_5; size_t x_6; size_t x_7; size_t x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_5 = lean_ctor_get(x_1, 0); +x_6 = 5; +x_7 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__2___closed__2; +x_8 = lean_usize_land(x_2, x_7); +x_9 = lean_usize_to_nat(x_8); +x_10 = lean_box(2); +x_11 = lean_array_get(x_10, x_5, x_9); +lean_dec(x_9); +lean_dec(x_5); +switch (lean_obj_tag(x_11)) { +case 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 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_3, x_12); +lean_dec(x_12); +if (x_14 == 0) +{ +lean_object* x_15; +lean_dec(x_13); +lean_free_object(x_1); +x_15 = lean_box(0); +return x_15; } else { -return x_109; +lean_ctor_set_tag(x_1, 1); +lean_ctor_set(x_1, 0, x_13); +return x_1; +} +} +case 1: +{ +lean_object* x_16; size_t x_17; +lean_free_object(x_1); +x_16 = lean_ctor_get(x_11, 0); +lean_inc(x_16); +lean_dec(x_11); +x_17 = lean_usize_shift_right(x_2, x_6); +x_1 = x_16; +x_2 = x_17; +goto _start; +} +default: +{ +lean_object* x_19; +lean_free_object(x_1); +x_19 = lean_box(0); +return x_19; +} } } else { -return x_109; +lean_object* x_20; size_t x_21; size_t x_22; size_t x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_20 = lean_ctor_get(x_1, 0); +lean_inc(x_20); +lean_dec(x_1); +x_21 = 5; +x_22 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__2___closed__2; +x_23 = lean_usize_land(x_2, x_22); +x_24 = lean_usize_to_nat(x_23); +x_25 = lean_box(2); +x_26 = lean_array_get(x_25, x_20, x_24); +lean_dec(x_24); +lean_dec(x_20); +switch (lean_obj_tag(x_26)) { +case 0: +{ +lean_object* x_27; lean_object* x_28; uint8_t 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_Grind_isSameExpr_unsafe__1(x_3, x_27); +lean_dec(x_27); +if (x_29 == 0) +{ +lean_object* x_30; +lean_dec(x_28); +x_30 = lean_box(0); +return x_30; +} +else +{ +lean_object* x_31; +x_31 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_31, 0, x_28); +return x_31; +} +} +case 1: +{ +lean_object* x_32; size_t x_33; +x_32 = lean_ctor_get(x_26, 0); +lean_inc(x_32); +lean_dec(x_26); +x_33 = lean_usize_shift_right(x_2, x_21); +x_1 = x_32; +x_2 = x_33; +goto _start; +} +default: +{ +lean_object* x_35; +x_35 = lean_box(0); +return x_35; +} +} } } +else +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_36 = lean_ctor_get(x_1, 0); +lean_inc(x_36); +x_37 = lean_ctor_get(x_1, 1); +lean_inc(x_37); +lean_dec(x_1); +x_38 = lean_unsigned_to_nat(0u); +x_39 = l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_registerParent___spec__9(x_36, x_37, lean_box(0), x_38, x_3); +lean_dec(x_37); +lean_dec(x_36); +return x_39; } } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_registerParent___spec__6(lean_object* x_1, size_t x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_registerParent___spec__7(lean_object* x_1, lean_object* x_2) { _start: { -uint64_t x_4; size_t x_5; size_t x_6; lean_object* x_7; -x_4 = lean_usize_to_uint64(x_2); -x_5 = lean_uint64_to_usize(x_4); -x_6 = 1; -x_7 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_registerParent___spec__7(x_1, x_5, x_6, x_2, x_3); -return x_7; +uint64_t x_3; size_t x_4; lean_object* x_5; +x_3 = l_Lean_Meta_Grind_instHashableENodeKey_unsafe__1(x_2); +x_4 = lean_uint64_to_usize(x_3); +x_5 = l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_registerParent___spec__8(x_1, x_4, x_2); +return x_5; } } LEAN_EXPORT lean_object* l_Lean_Meta_Grind_registerParent(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { @@ -10329,299 +14221,217 @@ return x_19; } else { -lean_object* x_20; lean_object* x_21; size_t x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_63; lean_object* x_64; x_20 = lean_ctor_get(x_12, 1); lean_inc(x_20); lean_dec(x_12); x_21 = lean_ctor_get(x_13, 0); lean_inc(x_21); lean_dec(x_13); -x_22 = lean_ptr_addr(x_21); -lean_dec(x_21); -x_23 = lean_st_ref_get(x_3, x_20); -x_24 = lean_ctor_get(x_23, 0); +x_22 = lean_st_ref_get(x_3, x_20); +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +x_24 = lean_ctor_get(x_22, 1); lean_inc(x_24); -x_25 = lean_ctor_get(x_23, 1); +lean_dec(x_22); +x_25 = lean_ctor_get(x_23, 2); lean_inc(x_25); lean_dec(x_23); -x_26 = lean_ctor_get(x_24, 2); -lean_inc(x_26); -lean_dec(x_24); -x_27 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_registerParent___spec__1(x_26, x_22); -x_28 = lean_st_ref_take(x_3, x_25); -if (lean_obj_tag(x_27) == 0) +x_63 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_registerParent___spec__7(x_25, x_21); +x_64 = lean_st_ref_take(x_3, x_24); +if (lean_obj_tag(x_63) == 0) { -lean_object* x_29; lean_object* x_30; lean_object* x_31; uint8_t x_32; -x_29 = lean_ctor_get(x_28, 0); -lean_inc(x_29); -x_30 = lean_ctor_get(x_28, 1); -lean_inc(x_30); -lean_dec(x_28); -x_31 = lean_box(0); -x_32 = !lean_is_exclusive(x_29); -if (x_32 == 0) +lean_object* x_65; lean_object* x_66; lean_object* x_67; +x_65 = lean_ctor_get(x_64, 0); +lean_inc(x_65); +x_66 = lean_ctor_get(x_64, 1); +lean_inc(x_66); +lean_dec(x_64); +x_67 = lean_box(0); +x_26 = x_67; +x_27 = x_65; +x_28 = x_66; +goto block_62; +} +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; -x_33 = lean_ctor_get(x_29, 2); -x_34 = lean_box(0); -x_35 = l_Lean_RBNode_insert___at_Lean_Meta_Grind_registerParent___spec__4(x_31, x_1, x_34); -x_36 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_registerParent___spec__6(x_33, x_22, x_35); -lean_ctor_set(x_29, 2, x_36); -x_37 = lean_st_ref_set(x_3, x_29, x_30); -x_38 = !lean_is_exclusive(x_37); -if (x_38 == 0) +lean_object* x_68; lean_object* x_69; lean_object* x_70; +x_68 = lean_ctor_get(x_63, 0); +lean_inc(x_68); +lean_dec(x_63); +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_26 = x_68; +x_27 = x_69; +x_28 = x_70; +goto block_62; +} +block_62: { -lean_object* x_39; -x_39 = lean_ctor_get(x_37, 0); -lean_dec(x_39); -lean_ctor_set(x_37, 0, x_34); -return x_37; +uint8_t x_29; +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; uint8_t x_35; +x_30 = lean_ctor_get(x_27, 2); +x_31 = lean_box(0); +x_32 = l_Lean_RBNode_insert___at_Lean_Meta_Grind_registerParent___spec__1(x_26, x_1, x_31); +x_33 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_registerParent___spec__3(x_30, x_21, x_32); +lean_ctor_set(x_27, 2, x_33); +x_34 = lean_st_ref_set(x_3, x_27, x_28); +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_31); +return x_34; } else { -lean_object* x_40; lean_object* x_41; -x_40 = lean_ctor_get(x_37, 1); -lean_inc(x_40); -lean_dec(x_37); -x_41 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_41, 0, x_34); -lean_ctor_set(x_41, 1, x_40); -return x_41; +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_31); +lean_ctor_set(x_38, 1, x_37); +return x_38; } } 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; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; -x_42 = lean_ctor_get(x_29, 0); -x_43 = lean_ctor_get(x_29, 1); -x_44 = lean_ctor_get(x_29, 2); -x_45 = lean_ctor_get(x_29, 3); -x_46 = lean_ctor_get(x_29, 4); -x_47 = lean_ctor_get(x_29, 5); -x_48 = lean_ctor_get_uint8(x_29, sizeof(void*)*8); -x_49 = lean_ctor_get(x_29, 6); -x_50 = lean_ctor_get(x_29, 7); +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; lean_object* x_61; +x_39 = lean_ctor_get(x_27, 0); +x_40 = lean_ctor_get(x_27, 1); +x_41 = lean_ctor_get(x_27, 2); +x_42 = lean_ctor_get(x_27, 3); +x_43 = lean_ctor_get(x_27, 4); +x_44 = lean_ctor_get(x_27, 5); +x_45 = lean_ctor_get_uint8(x_27, sizeof(void*)*14); +x_46 = lean_ctor_get(x_27, 6); +x_47 = lean_ctor_get(x_27, 7); +x_48 = lean_ctor_get(x_27, 8); +x_49 = lean_ctor_get(x_27, 9); +x_50 = lean_ctor_get(x_27, 10); +x_51 = lean_ctor_get(x_27, 11); +x_52 = lean_ctor_get(x_27, 12); +x_53 = lean_ctor_get(x_27, 13); +lean_inc(x_53); +lean_inc(x_52); +lean_inc(x_51); lean_inc(x_50); lean_inc(x_49); +lean_inc(x_48); lean_inc(x_47); lean_inc(x_46); -lean_inc(x_45); lean_inc(x_44); lean_inc(x_43); lean_inc(x_42); -lean_dec(x_29); -x_51 = lean_box(0); -x_52 = l_Lean_RBNode_insert___at_Lean_Meta_Grind_registerParent___spec__4(x_31, x_1, x_51); -x_53 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_registerParent___spec__6(x_44, x_22, x_52); -x_54 = lean_alloc_ctor(0, 8, 1); -lean_ctor_set(x_54, 0, x_42); -lean_ctor_set(x_54, 1, x_43); -lean_ctor_set(x_54, 2, x_53); -lean_ctor_set(x_54, 3, x_45); -lean_ctor_set(x_54, 4, x_46); -lean_ctor_set(x_54, 5, x_47); -lean_ctor_set(x_54, 6, x_49); -lean_ctor_set(x_54, 7, x_50); -lean_ctor_set_uint8(x_54, sizeof(void*)*8, x_48); -x_55 = lean_st_ref_set(x_3, x_54, x_30); -x_56 = lean_ctor_get(x_55, 1); -lean_inc(x_56); -if (lean_is_exclusive(x_55)) { - lean_ctor_release(x_55, 0); - lean_ctor_release(x_55, 1); - x_57 = x_55; -} else { - lean_dec_ref(x_55); - x_57 = lean_box(0); -} -if (lean_is_scalar(x_57)) { - x_58 = lean_alloc_ctor(0, 2, 0); -} else { - x_58 = x_57; -} -lean_ctor_set(x_58, 0, x_51); -lean_ctor_set(x_58, 1, x_56); -return x_58; -} -} -else -{ -lean_object* x_59; lean_object* x_60; lean_object* x_61; uint8_t x_62; -x_59 = lean_ctor_get(x_28, 0); -lean_inc(x_59); -x_60 = lean_ctor_get(x_27, 0); -lean_inc(x_60); +lean_inc(x_41); +lean_inc(x_40); +lean_inc(x_39); lean_dec(x_27); -x_61 = lean_ctor_get(x_28, 1); -lean_inc(x_61); -lean_dec(x_28); -x_62 = !lean_is_exclusive(x_59); -if (x_62 == 0) -{ -lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; uint8_t x_68; -x_63 = lean_ctor_get(x_59, 2); -x_64 = lean_box(0); -x_65 = l_Lean_RBNode_insert___at_Lean_Meta_Grind_registerParent___spec__4(x_60, x_1, x_64); -x_66 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_registerParent___spec__6(x_63, x_22, x_65); -lean_ctor_set(x_59, 2, x_66); -x_67 = lean_st_ref_set(x_3, x_59, x_61); -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_64); -return x_67; -} -else -{ -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_64); -lean_ctor_set(x_71, 1, x_70); -return x_71; -} -} -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; uint8_t x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; -x_72 = lean_ctor_get(x_59, 0); -x_73 = lean_ctor_get(x_59, 1); -x_74 = lean_ctor_get(x_59, 2); -x_75 = lean_ctor_get(x_59, 3); -x_76 = lean_ctor_get(x_59, 4); -x_77 = lean_ctor_get(x_59, 5); -x_78 = lean_ctor_get_uint8(x_59, sizeof(void*)*8); -x_79 = lean_ctor_get(x_59, 6); -x_80 = lean_ctor_get(x_59, 7); -lean_inc(x_80); -lean_inc(x_79); -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_59); -x_81 = lean_box(0); -x_82 = l_Lean_RBNode_insert___at_Lean_Meta_Grind_registerParent___spec__4(x_60, x_1, x_81); -x_83 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_registerParent___spec__6(x_74, x_22, x_82); -x_84 = lean_alloc_ctor(0, 8, 1); -lean_ctor_set(x_84, 0, x_72); -lean_ctor_set(x_84, 1, x_73); -lean_ctor_set(x_84, 2, x_83); -lean_ctor_set(x_84, 3, x_75); -lean_ctor_set(x_84, 4, x_76); -lean_ctor_set(x_84, 5, x_77); -lean_ctor_set(x_84, 6, x_79); -lean_ctor_set(x_84, 7, x_80); -lean_ctor_set_uint8(x_84, sizeof(void*)*8, x_78); -x_85 = lean_st_ref_set(x_3, x_84, x_61); -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; +x_54 = lean_box(0); +x_55 = l_Lean_RBNode_insert___at_Lean_Meta_Grind_registerParent___spec__1(x_26, x_1, x_54); +x_56 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_registerParent___spec__3(x_41, x_21, x_55); +x_57 = lean_alloc_ctor(0, 14, 1); +lean_ctor_set(x_57, 0, x_39); +lean_ctor_set(x_57, 1, x_40); +lean_ctor_set(x_57, 2, x_56); +lean_ctor_set(x_57, 3, x_42); +lean_ctor_set(x_57, 4, x_43); +lean_ctor_set(x_57, 5, x_44); +lean_ctor_set(x_57, 6, x_46); +lean_ctor_set(x_57, 7, x_47); +lean_ctor_set(x_57, 8, x_48); +lean_ctor_set(x_57, 9, x_49); +lean_ctor_set(x_57, 10, x_50); +lean_ctor_set(x_57, 11, x_51); +lean_ctor_set(x_57, 12, x_52); +lean_ctor_set(x_57, 13, x_53); +lean_ctor_set_uint8(x_57, sizeof(void*)*14, x_45); +x_58 = lean_st_ref_set(x_3, x_57, x_28); +x_59 = lean_ctor_get(x_58, 1); +lean_inc(x_59); +if (lean_is_exclusive(x_58)) { + lean_ctor_release(x_58, 0); + lean_ctor_release(x_58, 1); + x_60 = x_58; } else { - lean_dec_ref(x_85); - x_87 = lean_box(0); + lean_dec_ref(x_58); + x_60 = lean_box(0); } -if (lean_is_scalar(x_87)) { - x_88 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_60)) { + x_61 = lean_alloc_ctor(0, 2, 0); } else { - x_88 = x_87; + x_61 = x_60; } -lean_ctor_set(x_88, 0, x_81); -lean_ctor_set(x_88, 1, x_86); -return x_88; +lean_ctor_set(x_61, 0, x_54); +lean_ctor_set(x_61, 1, x_59); +return x_61; } } } } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_registerParent___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_registerParent___spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { -size_t x_6; lean_object* x_7; -x_6 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_7 = l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_registerParent___spec__3(x_1, x_2, x_3, x_4, x_6); -lean_dec(x_2); +size_t x_7; lean_object* x_8; +x_7 = lean_unbox_usize(x_1); lean_dec(x_1); -return x_7; -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_registerParent___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -size_t x_4; size_t x_5; lean_object* x_6; -x_4 = lean_unbox_usize(x_2); -lean_dec(x_2); -x_5 = lean_unbox_usize(x_3); +x_8 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_registerParent___spec__5(x_7, x_2, x_3, x_4, x_5, x_6); lean_dec(x_3); -x_6 = l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_registerParent___spec__2(x_1, x_4, x_5); -return x_6; -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_registerParent___spec__1___boxed(lean_object* x_1, lean_object* x_2) { -_start: -{ -size_t x_3; lean_object* x_4; -x_3 = lean_unbox_usize(x_2); lean_dec(x_2); -x_4 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_registerParent___spec__1(x_1, x_3); -return x_4; +return x_8; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_registerParent___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_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_registerParent___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -size_t x_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_Grind_registerParent___spec__8(x_7, x_2, x_3, x_4, x_5, x_6); -lean_dec(x_3); +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_Grind_registerParent___spec__4(x_1, x_6, x_7, x_4, x_5); return x_8; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_registerParent___spec__9___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_registerParent___spec__9___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -size_t x_5; lean_object* x_6; -x_5 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_6 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_registerParent___spec__9(x_1, x_2, x_5, x_4); +lean_object* x_6; +x_6 = l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_registerParent___spec__9(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_insertAux___at_Lean_Meta_Grind_registerParent___spec__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_registerParent___spec__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -size_t x_6; size_t x_7; size_t x_8; lean_object* x_9; -x_6 = lean_unbox_usize(x_2); +size_t x_4; lean_object* x_5; +x_4 = lean_unbox_usize(x_2); lean_dec(x_2); -x_7 = lean_unbox_usize(x_3); +x_5 = l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_registerParent___spec__8(x_1, x_4, x_3); lean_dec(x_3); -x_8 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_9 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_registerParent___spec__7(x_1, x_6, x_7, x_8, x_5); -return x_9; +return x_5; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_registerParent___spec__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_registerParent___spec__7___boxed(lean_object* x_1, lean_object* x_2) { _start: { -size_t x_4; lean_object* x_5; -x_4 = lean_unbox_usize(x_2); +lean_object* x_3; +x_3 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_registerParent___spec__7(x_1, x_2); lean_dec(x_2); -x_5 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_registerParent___spec__6(x_1, x_4, x_3); -return x_5; +return x_3; } } LEAN_EXPORT lean_object* l_Lean_Meta_Grind_registerParent___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { @@ -10649,63 +14459,61 @@ x_11 = lean_st_ref_get(x_2, x_10); x_12 = !lean_is_exclusive(x_11); if (x_12 == 0) { -lean_object* x_13; lean_object* x_14; size_t x_15; lean_object* x_16; +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_13, 2); lean_inc(x_14); lean_dec(x_13); -x_15 = lean_ptr_addr(x_1); -x_16 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_registerParent___spec__1(x_14, x_15); -if (lean_obj_tag(x_16) == 0) +x_15 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_registerParent___spec__7(x_14, x_1); +if (lean_obj_tag(x_15) == 0) { -lean_object* x_17; -x_17 = lean_box(0); -lean_ctor_set(x_11, 0, x_17); +lean_object* x_16; +x_16 = lean_box(0); +lean_ctor_set(x_11, 0, x_16); return x_11; } else { -lean_object* x_18; -x_18 = lean_ctor_get(x_16, 0); -lean_inc(x_18); -lean_dec(x_16); -lean_ctor_set(x_11, 0, x_18); +lean_object* x_17; +x_17 = lean_ctor_get(x_15, 0); +lean_inc(x_17); +lean_dec(x_15); +lean_ctor_set(x_11, 0, x_17); return x_11; } } else { -lean_object* x_19; lean_object* x_20; lean_object* x_21; size_t x_22; lean_object* x_23; -x_19 = lean_ctor_get(x_11, 0); -x_20 = lean_ctor_get(x_11, 1); -lean_inc(x_20); +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_18 = lean_ctor_get(x_11, 0); +x_19 = lean_ctor_get(x_11, 1); lean_inc(x_19); +lean_inc(x_18); lean_dec(x_11); -x_21 = lean_ctor_get(x_19, 2); -lean_inc(x_21); -lean_dec(x_19); -x_22 = lean_ptr_addr(x_1); -x_23 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_registerParent___spec__1(x_21, x_22); -if (lean_obj_tag(x_23) == 0) +x_20 = lean_ctor_get(x_18, 2); +lean_inc(x_20); +lean_dec(x_18); +x_21 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_registerParent___spec__7(x_20, x_1); +if (lean_obj_tag(x_21) == 0) +{ +lean_object* x_22; lean_object* x_23; +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_19); +return x_23; +} +else { lean_object* x_24; lean_object* x_25; -x_24 = lean_box(0); +x_24 = lean_ctor_get(x_21, 0); +lean_inc(x_24); +lean_dec(x_21); x_25 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_25, 0, x_24); -lean_ctor_set(x_25, 1, x_20); +lean_ctor_set(x_25, 1, x_19); return x_25; } -else -{ -lean_object* x_26; lean_object* x_27; -x_26 = lean_ctor_get(x_23, 0); -lean_inc(x_26); -lean_dec(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_20); -return x_27; -} } } } @@ -10726,7 +14534,7 @@ lean_dec(x_1); return x_11; } } -LEAN_EXPORT lean_object* l_Array_indexOfAux___at_Lean_Meta_Grind_getParentsAndReset___spec__3(lean_object* x_1, size_t x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_indexOfAux___at_Lean_Meta_Grind_getParentsAndReset___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; uint8_t x_5; @@ -10742,31 +14550,30 @@ return x_6; } else { -lean_object* x_7; size_t x_8; uint8_t x_9; +lean_object* x_7; uint8_t x_8; x_7 = lean_array_fget(x_1, x_3); -x_8 = lean_unbox_usize(x_7); +x_8 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_7, x_2); lean_dec(x_7); -x_9 = lean_usize_dec_eq(x_8, x_2); -if (x_9 == 0) +if (x_8 == 0) { -lean_object* x_10; lean_object* x_11; -x_10 = lean_unsigned_to_nat(1u); -x_11 = lean_nat_add(x_3, x_10); +lean_object* x_9; lean_object* x_10; +x_9 = lean_unsigned_to_nat(1u); +x_10 = lean_nat_add(x_3, x_9); lean_dec(x_3); -x_3 = x_11; +x_3 = x_10; goto _start; } else { -lean_object* x_13; -x_13 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_13, 0, x_3); -return x_13; +lean_object* x_12; +x_12 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_12, 0, x_3); +return x_12; } } } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_getParentsAndReset___spec__2(lean_object* x_1, size_t x_2, size_t x_3) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_getParentsAndReset___spec__2(lean_object* x_1, size_t x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_1) == 0) @@ -10783,14 +14590,13 @@ x_10 = lean_array_get(x_9, x_4, x_8); switch (lean_obj_tag(x_10)) { case 0: { -lean_object* x_11; size_t x_12; uint8_t x_13; +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_unbox_usize(x_11); +x_12 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_3, x_11); lean_dec(x_11); -x_13 = lean_usize_dec_eq(x_3, x_12); -if (x_13 == 0) +if (x_12 == 0) { lean_dec(x_8); lean_dec(x_4); @@ -10798,212 +14604,212 @@ return x_1; } else { -uint8_t x_14; -x_14 = !lean_is_exclusive(x_1); -if (x_14 == 0) +uint8_t x_13; +x_13 = !lean_is_exclusive(x_1); +if (x_13 == 0) { -lean_object* x_15; lean_object* x_16; -x_15 = lean_ctor_get(x_1, 0); -lean_dec(x_15); -x_16 = lean_array_set(x_4, x_8, x_9); +lean_object* x_14; lean_object* x_15; +x_14 = lean_ctor_get(x_1, 0); +lean_dec(x_14); +x_15 = lean_array_set(x_4, x_8, x_9); lean_dec(x_8); -lean_ctor_set(x_1, 0, x_16); +lean_ctor_set(x_1, 0, x_15); return x_1; } else { -lean_object* x_17; lean_object* x_18; +lean_object* x_16; lean_object* x_17; lean_dec(x_1); -x_17 = lean_array_set(x_4, x_8, x_9); +x_16 = lean_array_set(x_4, x_8, x_9); lean_dec(x_8); -x_18 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_18, 0, x_17); -return x_18; +x_17 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_17, 0, x_16); +return x_17; } } } case 1: { -uint8_t x_19; -x_19 = !lean_is_exclusive(x_1); -if (x_19 == 0) +uint8_t x_18; +x_18 = !lean_is_exclusive(x_1); +if (x_18 == 0) { -lean_object* x_20; uint8_t x_21; -x_20 = lean_ctor_get(x_1, 0); -lean_dec(x_20); -x_21 = !lean_is_exclusive(x_10); -if (x_21 == 0) +lean_object* x_19; uint8_t x_20; +x_19 = lean_ctor_get(x_1, 0); +lean_dec(x_19); +x_20 = !lean_is_exclusive(x_10); +if (x_20 == 0) { -lean_object* x_22; lean_object* x_23; size_t x_24; lean_object* x_25; lean_object* x_26; -x_22 = lean_ctor_get(x_10, 0); -x_23 = lean_array_set(x_4, x_8, x_9); -x_24 = lean_usize_shift_right(x_2, x_5); -x_25 = l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_getParentsAndReset___spec__2(x_22, x_24, x_3); -lean_inc(x_25); -x_26 = l_Lean_PersistentHashMap_isUnaryNode___rarg(x_25); -if (lean_obj_tag(x_26) == 0) +lean_object* x_21; lean_object* x_22; size_t x_23; lean_object* x_24; lean_object* x_25; +x_21 = lean_ctor_get(x_10, 0); +x_22 = lean_array_set(x_4, x_8, x_9); +x_23 = lean_usize_shift_right(x_2, x_5); +x_24 = l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_getParentsAndReset___spec__2(x_21, x_23, x_3); +lean_inc(x_24); +x_25 = l_Lean_PersistentHashMap_isUnaryNode___rarg(x_24); +if (lean_obj_tag(x_25) == 0) { -lean_object* x_27; -lean_ctor_set(x_10, 0, x_25); -x_27 = lean_array_set(x_23, x_8, x_10); +lean_object* x_26; +lean_ctor_set(x_10, 0, x_24); +x_26 = lean_array_set(x_22, x_8, x_10); lean_dec(x_8); -lean_ctor_set(x_1, 0, x_27); +lean_ctor_set(x_1, 0, x_26); return x_1; } else { -lean_object* x_28; uint8_t x_29; -lean_dec(x_25); +lean_object* x_27; uint8_t x_28; +lean_dec(x_24); lean_free_object(x_10); -x_28 = lean_ctor_get(x_26, 0); -lean_inc(x_28); -lean_dec(x_26); -x_29 = !lean_is_exclusive(x_28); -if (x_29 == 0) +x_27 = lean_ctor_get(x_25, 0); +lean_inc(x_27); +lean_dec(x_25); +x_28 = !lean_is_exclusive(x_27); +if (x_28 == 0) { -lean_object* x_30; -x_30 = lean_array_set(x_23, x_8, x_28); +lean_object* x_29; +x_29 = lean_array_set(x_22, x_8, x_27); lean_dec(x_8); -lean_ctor_set(x_1, 0, x_30); +lean_ctor_set(x_1, 0, x_29); return x_1; } else { -lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_31 = lean_ctor_get(x_28, 0); -x_32 = lean_ctor_get(x_28, 1); -lean_inc(x_32); +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_30 = lean_ctor_get(x_27, 0); +x_31 = lean_ctor_get(x_27, 1); lean_inc(x_31); -lean_dec(x_28); -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_array_set(x_23, x_8, x_33); +lean_inc(x_30); +lean_dec(x_27); +x_32 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_32, 0, x_30); +lean_ctor_set(x_32, 1, x_31); +x_33 = lean_array_set(x_22, x_8, x_32); lean_dec(x_8); -lean_ctor_set(x_1, 0, x_34); +lean_ctor_set(x_1, 0, x_33); return x_1; } } } else { -lean_object* x_35; lean_object* x_36; size_t x_37; lean_object* x_38; lean_object* x_39; -x_35 = lean_ctor_get(x_10, 0); -lean_inc(x_35); +lean_object* x_34; lean_object* x_35; size_t x_36; lean_object* x_37; lean_object* x_38; +x_34 = lean_ctor_get(x_10, 0); +lean_inc(x_34); lean_dec(x_10); -x_36 = lean_array_set(x_4, x_8, x_9); -x_37 = lean_usize_shift_right(x_2, x_5); -x_38 = l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_getParentsAndReset___spec__2(x_35, x_37, x_3); -lean_inc(x_38); -x_39 = l_Lean_PersistentHashMap_isUnaryNode___rarg(x_38); -if (lean_obj_tag(x_39) == 0) +x_35 = lean_array_set(x_4, x_8, x_9); +x_36 = lean_usize_shift_right(x_2, x_5); +x_37 = l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_getParentsAndReset___spec__2(x_34, x_36, x_3); +lean_inc(x_37); +x_38 = l_Lean_PersistentHashMap_isUnaryNode___rarg(x_37); +if (lean_obj_tag(x_38) == 0) { -lean_object* x_40; lean_object* x_41; -x_40 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_40, 0, x_38); -x_41 = lean_array_set(x_36, x_8, x_40); +lean_object* x_39; lean_object* x_40; +x_39 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_39, 0, x_37); +x_40 = lean_array_set(x_35, x_8, x_39); lean_dec(x_8); -lean_ctor_set(x_1, 0, x_41); +lean_ctor_set(x_1, 0, x_40); return x_1; } else { -lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; +lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; +lean_dec(x_37); +x_41 = lean_ctor_get(x_38, 0); +lean_inc(x_41); lean_dec(x_38); -x_42 = lean_ctor_get(x_39, 0); +x_42 = lean_ctor_get(x_41, 0); lean_inc(x_42); -lean_dec(x_39); -x_43 = lean_ctor_get(x_42, 0); +x_43 = lean_ctor_get(x_41, 1); 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; +if (lean_is_exclusive(x_41)) { + lean_ctor_release(x_41, 0); + lean_ctor_release(x_41, 1); + x_44 = x_41; } else { - lean_dec_ref(x_42); - x_45 = lean_box(0); + lean_dec_ref(x_41); + x_44 = lean_box(0); } -if (lean_is_scalar(x_45)) { - x_46 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_44)) { + x_45 = lean_alloc_ctor(0, 2, 0); } else { - x_46 = x_45; + x_45 = x_44; } -lean_ctor_set(x_46, 0, x_43); -lean_ctor_set(x_46, 1, x_44); -x_47 = lean_array_set(x_36, x_8, x_46); +lean_ctor_set(x_45, 0, x_42); +lean_ctor_set(x_45, 1, x_43); +x_46 = lean_array_set(x_35, x_8, x_45); lean_dec(x_8); -lean_ctor_set(x_1, 0, x_47); +lean_ctor_set(x_1, 0, x_46); return x_1; } } } else { -lean_object* x_48; lean_object* x_49; lean_object* x_50; size_t x_51; lean_object* x_52; lean_object* x_53; +lean_object* x_47; lean_object* x_48; lean_object* x_49; size_t x_50; lean_object* x_51; lean_object* x_52; lean_dec(x_1); -x_48 = lean_ctor_get(x_10, 0); -lean_inc(x_48); +x_47 = lean_ctor_get(x_10, 0); +lean_inc(x_47); if (lean_is_exclusive(x_10)) { lean_ctor_release(x_10, 0); - x_49 = x_10; + x_48 = x_10; } else { lean_dec_ref(x_10); - x_49 = lean_box(0); + x_48 = lean_box(0); } -x_50 = lean_array_set(x_4, x_8, x_9); -x_51 = lean_usize_shift_right(x_2, x_5); -x_52 = l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_getParentsAndReset___spec__2(x_48, x_51, x_3); -lean_inc(x_52); -x_53 = l_Lean_PersistentHashMap_isUnaryNode___rarg(x_52); -if (lean_obj_tag(x_53) == 0) +x_49 = lean_array_set(x_4, x_8, x_9); +x_50 = lean_usize_shift_right(x_2, x_5); +x_51 = l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_getParentsAndReset___spec__2(x_47, x_50, x_3); +lean_inc(x_51); +x_52 = l_Lean_PersistentHashMap_isUnaryNode___rarg(x_51); +if (lean_obj_tag(x_52) == 0) { -lean_object* x_54; lean_object* x_55; lean_object* x_56; -if (lean_is_scalar(x_49)) { - x_54 = lean_alloc_ctor(1, 1, 0); +lean_object* x_53; lean_object* x_54; lean_object* x_55; +if (lean_is_scalar(x_48)) { + x_53 = lean_alloc_ctor(1, 1, 0); } else { - x_54 = x_49; + x_53 = x_48; } -lean_ctor_set(x_54, 0, x_52); -x_55 = lean_array_set(x_50, x_8, x_54); +lean_ctor_set(x_53, 0, x_51); +x_54 = lean_array_set(x_49, x_8, x_53); lean_dec(x_8); -x_56 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_56, 0, x_55); -return x_56; +x_55 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_55, 0, x_54); +return x_55; } else { -lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; +lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; +lean_dec(x_51); +lean_dec(x_48); +x_56 = lean_ctor_get(x_52, 0); +lean_inc(x_56); lean_dec(x_52); -lean_dec(x_49); -x_57 = lean_ctor_get(x_53, 0); +x_57 = lean_ctor_get(x_56, 0); lean_inc(x_57); -lean_dec(x_53); -x_58 = lean_ctor_get(x_57, 0); +x_58 = lean_ctor_get(x_56, 1); 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; +if (lean_is_exclusive(x_56)) { + lean_ctor_release(x_56, 0); + lean_ctor_release(x_56, 1); + x_59 = x_56; } else { - lean_dec_ref(x_57); - x_60 = lean_box(0); + lean_dec_ref(x_56); + x_59 = lean_box(0); } -if (lean_is_scalar(x_60)) { - x_61 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_59)) { + x_60 = lean_alloc_ctor(0, 2, 0); } else { - x_61 = x_60; + x_60 = x_59; } -lean_ctor_set(x_61, 0, x_58); -lean_ctor_set(x_61, 1, x_59); -x_62 = lean_array_set(x_50, x_8, x_61); +lean_ctor_set(x_60, 0, x_57); +lean_ctor_set(x_60, 1, x_58); +x_61 = lean_array_set(x_49, x_8, x_60); lean_dec(x_8); -x_63 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_63, 0, x_62); -return x_63; +x_62 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_62, 0, x_61); +return x_62; } } } @@ -11017,64 +14823,64 @@ return x_1; } else { -lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; -x_64 = lean_ctor_get(x_1, 0); +lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_63 = lean_ctor_get(x_1, 0); +lean_inc(x_63); +x_64 = lean_ctor_get(x_1, 1); lean_inc(x_64); -x_65 = lean_ctor_get(x_1, 1); -lean_inc(x_65); -x_66 = lean_unsigned_to_nat(0u); -x_67 = l_Array_indexOfAux___at_Lean_Meta_Grind_getParentsAndReset___spec__3(x_64, x_3, x_66); -if (lean_obj_tag(x_67) == 0) +x_65 = lean_unsigned_to_nat(0u); +x_66 = l_Array_indexOfAux___at_Lean_Meta_Grind_getParentsAndReset___spec__3(x_63, x_3, x_65); +if (lean_obj_tag(x_66) == 0) { -lean_dec(x_65); lean_dec(x_64); +lean_dec(x_63); return x_1; } else { -uint8_t x_68; -x_68 = !lean_is_exclusive(x_1); -if (x_68 == 0) +uint8_t x_67; +x_67 = !lean_is_exclusive(x_1); +if (x_67 == 0) { -lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; -x_69 = lean_ctor_get(x_1, 1); +lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; +x_68 = lean_ctor_get(x_1, 1); +lean_dec(x_68); +x_69 = lean_ctor_get(x_1, 0); lean_dec(x_69); -x_70 = lean_ctor_get(x_1, 0); -lean_dec(x_70); -x_71 = lean_ctor_get(x_67, 0); -lean_inc(x_71); -lean_dec(x_67); -lean_inc(x_71); -x_72 = l_Array_eraseIdx___rarg(x_64, x_71, lean_box(0)); -x_73 = l_Array_eraseIdx___rarg(x_65, x_71, lean_box(0)); -lean_ctor_set(x_1, 1, x_73); -lean_ctor_set(x_1, 0, x_72); +x_70 = lean_ctor_get(x_66, 0); +lean_inc(x_70); +lean_dec(x_66); +lean_inc(x_70); +x_71 = l_Array_eraseIdx___rarg(x_63, x_70, lean_box(0)); +x_72 = l_Array_eraseIdx___rarg(x_64, x_70, lean_box(0)); +lean_ctor_set(x_1, 1, x_72); +lean_ctor_set(x_1, 0, x_71); return x_1; } else { -lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; +lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_dec(x_1); -x_74 = lean_ctor_get(x_67, 0); -lean_inc(x_74); -lean_dec(x_67); -lean_inc(x_74); -x_75 = l_Array_eraseIdx___rarg(x_64, x_74, lean_box(0)); -x_76 = l_Array_eraseIdx___rarg(x_65, x_74, lean_box(0)); -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; +x_73 = lean_ctor_get(x_66, 0); +lean_inc(x_73); +lean_dec(x_66); +lean_inc(x_73); +x_74 = l_Array_eraseIdx___rarg(x_63, x_73, lean_box(0)); +x_75 = l_Array_eraseIdx___rarg(x_64, x_73, lean_box(0)); +x_76 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_76, 0, x_74); +lean_ctor_set(x_76, 1, x_75); +return x_76; } } } } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_erase___at_Lean_Meta_Grind_getParentsAndReset___spec__1(lean_object* x_1, size_t x_2) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_erase___at_Lean_Meta_Grind_getParentsAndReset___spec__1(lean_object* x_1, lean_object* x_2) { _start: { uint64_t x_3; size_t x_4; lean_object* x_5; -x_3 = lean_usize_to_uint64(x_2); +x_3 = l_Lean_Meta_Grind_instHashableENodeKey_unsafe__1(x_2); x_4 = lean_uint64_to_usize(x_3); x_5 = l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_getParentsAndReset___spec__2(x_1, x_4, x_2); return x_5; @@ -11099,119 +14905,132 @@ lean_dec(x_14); x_17 = !lean_is_exclusive(x_15); if (x_17 == 0) { -lean_object* x_18; size_t x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; +lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; x_18 = lean_ctor_get(x_15, 2); -x_19 = lean_ptr_addr(x_1); -x_20 = l_Lean_PersistentHashMap_erase___at_Lean_Meta_Grind_getParentsAndReset___spec__1(x_18, x_19); -lean_ctor_set(x_15, 2, x_20); -x_21 = lean_st_ref_set(x_2, x_15, x_16); -x_22 = !lean_is_exclusive(x_21); -if (x_22 == 0) +x_19 = l_Lean_PersistentHashMap_erase___at_Lean_Meta_Grind_getParentsAndReset___spec__1(x_18, x_1); +lean_ctor_set(x_15, 2, x_19); +x_20 = lean_st_ref_set(x_2, x_15, x_16); +x_21 = !lean_is_exclusive(x_20); +if (x_21 == 0) { -lean_object* x_23; -x_23 = lean_ctor_get(x_21, 0); -lean_dec(x_23); -lean_ctor_set(x_21, 0, x_12); -return x_21; +lean_object* x_22; +x_22 = lean_ctor_get(x_20, 0); +lean_dec(x_22); +lean_ctor_set(x_20, 0, x_12); +return x_20; } else { -lean_object* x_24; lean_object* x_25; -x_24 = lean_ctor_get(x_21, 1); -lean_inc(x_24); -lean_dec(x_21); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_12); -lean_ctor_set(x_25, 1, x_24); -return x_25; +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_12); +lean_ctor_set(x_24, 1, x_23); +return x_24; } } else { -lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; uint8_t x_32; lean_object* x_33; lean_object* 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; lean_object* x_41; -x_26 = lean_ctor_get(x_15, 0); -x_27 = lean_ctor_get(x_15, 1); -x_28 = lean_ctor_get(x_15, 2); -x_29 = lean_ctor_get(x_15, 3); -x_30 = lean_ctor_get(x_15, 4); -x_31 = lean_ctor_get(x_15, 5); -x_32 = lean_ctor_get_uint8(x_15, sizeof(void*)*8); -x_33 = lean_ctor_get(x_15, 6); -x_34 = lean_ctor_get(x_15, 7); +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; 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_25 = lean_ctor_get(x_15, 0); +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); +x_30 = lean_ctor_get(x_15, 5); +x_31 = lean_ctor_get_uint8(x_15, sizeof(void*)*14); +x_32 = lean_ctor_get(x_15, 6); +x_33 = lean_ctor_get(x_15, 7); +x_34 = lean_ctor_get(x_15, 8); +x_35 = lean_ctor_get(x_15, 9); +x_36 = lean_ctor_get(x_15, 10); +x_37 = lean_ctor_get(x_15, 11); +x_38 = lean_ctor_get(x_15, 12); +x_39 = lean_ctor_get(x_15, 13); +lean_inc(x_39); +lean_inc(x_38); +lean_inc(x_37); +lean_inc(x_36); +lean_inc(x_35); lean_inc(x_34); lean_inc(x_33); -lean_inc(x_31); +lean_inc(x_32); 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_dec(x_15); -x_35 = lean_ptr_addr(x_1); -x_36 = l_Lean_PersistentHashMap_erase___at_Lean_Meta_Grind_getParentsAndReset___spec__1(x_28, x_35); -x_37 = lean_alloc_ctor(0, 8, 1); -lean_ctor_set(x_37, 0, x_26); -lean_ctor_set(x_37, 1, x_27); -lean_ctor_set(x_37, 2, x_36); -lean_ctor_set(x_37, 3, x_29); -lean_ctor_set(x_37, 4, x_30); -lean_ctor_set(x_37, 5, x_31); -lean_ctor_set(x_37, 6, x_33); -lean_ctor_set(x_37, 7, x_34); -lean_ctor_set_uint8(x_37, sizeof(void*)*8, x_32); -x_38 = lean_st_ref_set(x_2, x_37, x_16); -x_39 = lean_ctor_get(x_38, 1); -lean_inc(x_39); -if (lean_is_exclusive(x_38)) { - lean_ctor_release(x_38, 0); - lean_ctor_release(x_38, 1); - x_40 = x_38; +x_40 = l_Lean_PersistentHashMap_erase___at_Lean_Meta_Grind_getParentsAndReset___spec__1(x_27, x_1); +x_41 = lean_alloc_ctor(0, 14, 1); +lean_ctor_set(x_41, 0, x_25); +lean_ctor_set(x_41, 1, x_26); +lean_ctor_set(x_41, 2, x_40); +lean_ctor_set(x_41, 3, x_28); +lean_ctor_set(x_41, 4, x_29); +lean_ctor_set(x_41, 5, x_30); +lean_ctor_set(x_41, 6, x_32); +lean_ctor_set(x_41, 7, x_33); +lean_ctor_set(x_41, 8, x_34); +lean_ctor_set(x_41, 9, x_35); +lean_ctor_set(x_41, 10, x_36); +lean_ctor_set(x_41, 11, x_37); +lean_ctor_set(x_41, 12, x_38); +lean_ctor_set(x_41, 13, x_39); +lean_ctor_set_uint8(x_41, sizeof(void*)*14, x_31); +x_42 = lean_st_ref_set(x_2, x_41, x_16); +x_43 = lean_ctor_get(x_42, 1); +lean_inc(x_43); +if (lean_is_exclusive(x_42)) { + lean_ctor_release(x_42, 0); + lean_ctor_release(x_42, 1); + x_44 = x_42; } else { - lean_dec_ref(x_38); - x_40 = lean_box(0); + lean_dec_ref(x_42); + x_44 = lean_box(0); } -if (lean_is_scalar(x_40)) { - x_41 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_44)) { + x_45 = lean_alloc_ctor(0, 2, 0); } else { - x_41 = x_40; + x_45 = x_44; } -lean_ctor_set(x_41, 0, x_12); -lean_ctor_set(x_41, 1, x_39); -return x_41; +lean_ctor_set(x_45, 0, x_12); +lean_ctor_set(x_45, 1, x_43); +return x_45; } } } LEAN_EXPORT lean_object* l_Array_indexOfAux___at_Lean_Meta_Grind_getParentsAndReset___spec__3___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_object* x_4; +x_4 = l_Array_indexOfAux___at_Lean_Meta_Grind_getParentsAndReset___spec__3(x_1, x_2, x_3); lean_dec(x_2); -x_5 = l_Array_indexOfAux___at_Lean_Meta_Grind_getParentsAndReset___spec__3(x_1, x_4, x_3); lean_dec(x_1); -return x_5; +return x_4; } } LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_getParentsAndReset___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; +size_t x_4; lean_object* x_5; x_4 = lean_unbox_usize(x_2); lean_dec(x_2); -x_5 = lean_unbox_usize(x_3); +x_5 = l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_getParentsAndReset___spec__2(x_1, x_4, x_3); lean_dec(x_3); -x_6 = l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_getParentsAndReset___spec__2(x_1, x_4, x_5); -return x_6; +return x_5; } } LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_erase___at_Lean_Meta_Grind_getParentsAndReset___spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { -size_t x_3; lean_object* x_4; -x_3 = lean_unbox_usize(x_2); +lean_object* x_3; +x_3 = l_Lean_PersistentHashMap_erase___at_Lean_Meta_Grind_getParentsAndReset___spec__1(x_1, x_2); lean_dec(x_2); -x_4 = l_Lean_PersistentHashMap_erase___at_Lean_Meta_Grind_getParentsAndReset___spec__1(x_1, x_3); -return x_4; +return x_3; } } LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getParentsAndReset___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { @@ -11264,7 +15083,7 @@ x_20 = lean_ctor_get(x_18, 0); lean_inc(x_20); lean_dec(x_18); x_21 = lean_box(0); -x_22 = l_Lean_RBNode_insert___at_Lean_Meta_Grind_registerParent___spec__4(x_20, x_15, x_21); +x_22 = l_Lean_RBNode_insert___at_Lean_Meta_Grind_registerParent___spec__1(x_20, x_15, x_21); x_1 = x_16; x_2 = x_22; x_11 = x_19; @@ -11275,42 +15094,22 @@ goto _start; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_copyParentsTo(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, 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; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_48; lean_object* x_49; -x_12 = lean_ptr_addr(x_2); -x_13 = lean_st_ref_get(x_3, x_11); -x_14 = lean_ctor_get(x_13, 0); +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_54; +x_12 = lean_st_ref_get(x_3, x_11); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_12, 1); lean_inc(x_14); -x_15 = lean_ctor_get(x_13, 1); +lean_dec(x_12); +x_15 = lean_ctor_get(x_13, 2); lean_inc(x_15); lean_dec(x_13); -x_48 = lean_ctor_get(x_14, 2); -lean_inc(x_48); -lean_dec(x_14); -x_49 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_registerParent___spec__1(x_48, x_12); -if (lean_obj_tag(x_49) == 0) -{ -lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; -x_50 = lean_box(0); -x_51 = l_Lean_RBNode_forIn_visit___at_Lean_Meta_Grind_copyParentsTo___spec__1(x_1, x_50, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_15); -x_52 = lean_ctor_get(x_51, 0); -lean_inc(x_52); -x_53 = lean_ctor_get(x_51, 1); -lean_inc(x_53); -lean_dec(x_51); -x_54 = lean_ctor_get(x_52, 0); -lean_inc(x_54); -lean_dec(x_52); -x_16 = x_54; -x_17 = x_53; -goto block_47; -} -else +x_54 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_registerParent___spec__7(x_15, x_2); +if (lean_obj_tag(x_54) == 0) { lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; -x_55 = lean_ctor_get(x_49, 0); -lean_inc(x_55); -lean_dec(x_49); -x_56 = l_Lean_RBNode_forIn_visit___at_Lean_Meta_Grind_copyParentsTo___spec__1(x_1, x_55, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_15); +x_55 = lean_box(0); +x_56 = l_Lean_RBNode_forIn_visit___at_Lean_Meta_Grind_copyParentsTo___spec__1(x_1, x_55, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_14); x_57 = lean_ctor_get(x_56, 0); lean_inc(x_57); x_58 = lean_ctor_get(x_56, 1); @@ -11321,9 +15120,28 @@ lean_inc(x_59); lean_dec(x_57); x_16 = x_59; x_17 = x_58; -goto block_47; +goto block_53; +} +else +{ +lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; +x_60 = lean_ctor_get(x_54, 0); +lean_inc(x_60); +lean_dec(x_54); +x_61 = l_Lean_RBNode_forIn_visit___at_Lean_Meta_Grind_copyParentsTo___spec__1(x_1, x_60, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_14); +x_62 = lean_ctor_get(x_61, 0); +lean_inc(x_62); +x_63 = lean_ctor_get(x_61, 1); +lean_inc(x_63); +lean_dec(x_61); +x_64 = lean_ctor_get(x_62, 0); +lean_inc(x_64); +lean_dec(x_62); +x_16 = x_64; +x_17 = x_63; +goto block_53; } -block_47: +block_53: { lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; x_18 = lean_st_ref_take(x_3, x_17); @@ -11337,7 +15155,7 @@ if (x_21 == 0) { lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; x_22 = lean_ctor_get(x_19, 2); -x_23 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_registerParent___spec__6(x_22, x_12, x_16); +x_23 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_registerParent___spec__3(x_22, x_2, x_16); lean_ctor_set(x_19, 2, x_23); x_24 = lean_st_ref_set(x_3, x_19, x_20); x_25 = !lean_is_exclusive(x_24); @@ -11365,16 +15183,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; 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_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; x_31 = lean_ctor_get(x_19, 0); x_32 = lean_ctor_get(x_19, 1); x_33 = lean_ctor_get(x_19, 2); x_34 = lean_ctor_get(x_19, 3); x_35 = lean_ctor_get(x_19, 4); x_36 = lean_ctor_get(x_19, 5); -x_37 = lean_ctor_get_uint8(x_19, sizeof(void*)*8); +x_37 = lean_ctor_get_uint8(x_19, sizeof(void*)*14); x_38 = lean_ctor_get(x_19, 6); x_39 = lean_ctor_get(x_19, 7); +x_40 = lean_ctor_get(x_19, 8); +x_41 = lean_ctor_get(x_19, 9); +x_42 = lean_ctor_get(x_19, 10); +x_43 = lean_ctor_get(x_19, 11); +x_44 = lean_ctor_get(x_19, 12); +x_45 = lean_ctor_get(x_19, 13); +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_36); @@ -11384,37 +15214,43 @@ lean_inc(x_33); lean_inc(x_32); lean_inc(x_31); lean_dec(x_19); -x_40 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_registerParent___spec__6(x_33, x_12, x_16); -x_41 = lean_alloc_ctor(0, 8, 1); -lean_ctor_set(x_41, 0, x_31); -lean_ctor_set(x_41, 1, x_32); -lean_ctor_set(x_41, 2, x_40); -lean_ctor_set(x_41, 3, x_34); -lean_ctor_set(x_41, 4, x_35); -lean_ctor_set(x_41, 5, x_36); -lean_ctor_set(x_41, 6, x_38); -lean_ctor_set(x_41, 7, x_39); -lean_ctor_set_uint8(x_41, sizeof(void*)*8, x_37); -x_42 = lean_st_ref_set(x_3, x_41, x_20); -x_43 = lean_ctor_get(x_42, 1); -lean_inc(x_43); -if (lean_is_exclusive(x_42)) { - lean_ctor_release(x_42, 0); - lean_ctor_release(x_42, 1); - x_44 = x_42; +x_46 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_registerParent___spec__3(x_33, x_2, x_16); +x_47 = lean_alloc_ctor(0, 14, 1); +lean_ctor_set(x_47, 0, x_31); +lean_ctor_set(x_47, 1, x_32); +lean_ctor_set(x_47, 2, x_46); +lean_ctor_set(x_47, 3, x_34); +lean_ctor_set(x_47, 4, x_35); +lean_ctor_set(x_47, 5, x_36); +lean_ctor_set(x_47, 6, x_38); +lean_ctor_set(x_47, 7, x_39); +lean_ctor_set(x_47, 8, x_40); +lean_ctor_set(x_47, 9, x_41); +lean_ctor_set(x_47, 10, x_42); +lean_ctor_set(x_47, 11, x_43); +lean_ctor_set(x_47, 12, x_44); +lean_ctor_set(x_47, 13, x_45); +lean_ctor_set_uint8(x_47, sizeof(void*)*14, x_37); +x_48 = lean_st_ref_set(x_3, x_47, x_20); +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_42); - x_44 = lean_box(0); + lean_dec_ref(x_48); + x_50 = lean_box(0); } -x_45 = lean_box(0); -if (lean_is_scalar(x_44)) { - x_46 = lean_alloc_ctor(0, 2, 0); +x_51 = lean_box(0); +if (lean_is_scalar(x_50)) { + x_52 = lean_alloc_ctor(0, 2, 0); } else { - x_46 = x_44; + x_52 = x_50; } -lean_ctor_set(x_46, 0, x_45); -lean_ctor_set(x_46, 1, x_43); -return x_46; +lean_ctor_set(x_52, 0, x_51); +lean_ctor_set(x_52, 1, x_49); +return x_52; } } } @@ -11437,40 +15273,21 @@ return x_12; } LEAN_EXPORT lean_object* l_Lean_Meta_Grind_copyParentsTo___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, 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_Meta_Grind_copyParentsTo(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_12; -} -} -LEAN_EXPORT size_t l_Lean_Meta_Grind_setENode_unsafe__1(lean_object* x_1) { -_start: -{ -size_t x_2; -x_2 = lean_ptr_addr(x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_setENode_unsafe__1___boxed(lean_object* x_1) { -_start: -{ -size_t x_2; lean_object* x_3; -x_2 = l_Lean_Meta_Grind_setENode_unsafe__1(x_1); -lean_dec(x_1); -x_3 = lean_box_usize(x_2); -return x_3; +{ +lean_object* x_12; +x_12 = l_Lean_Meta_Grind_copyParentsTo(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_12; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_setENode_unsafe__2___rarg(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_setENode_unsafe__1___rarg(lean_object* x_1) { _start: { lean_object* x_2; @@ -11479,28 +15296,28 @@ lean_inc(x_2); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_setENode_unsafe__2(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_setENode_unsafe__1(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_setENode_unsafe__2___rarg___boxed), 1, 0); +x_3 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_setENode_unsafe__1___rarg___boxed), 1, 0); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_setENode_unsafe__2___rarg___boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_setENode_unsafe__1___rarg___boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l_Lean_Meta_Grind_setENode_unsafe__2___rarg(x_1); +x_2 = l_Lean_Meta_Grind_setENode_unsafe__1___rarg(x_1); lean_dec(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_setENode_unsafe__2___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_setENode_unsafe__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Meta_Grind_setENode_unsafe__2(x_1, x_2); +x_3 = l_Lean_Meta_Grind_setENode_unsafe__1(x_1, x_2); lean_dec(x_2); lean_dec(x_1); return x_3; @@ -11520,30 +15337,28 @@ return x_6; } else { -lean_object* x_9; size_t x_10; lean_object* x_11; uint64_t x_12; size_t x_13; size_t x_14; size_t x_15; size_t x_16; size_t x_17; size_t x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +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_unbox_usize(x_9); -lean_dec(x_9); -x_11 = lean_array_fget(x_3, x_5); -x_12 = lean_usize_to_uint64(x_10); -x_13 = lean_uint64_to_usize(x_12); -x_14 = 1; -x_15 = lean_usize_sub(x_1, x_14); -x_16 = 5; -x_17 = lean_usize_mul(x_16, x_15); -x_18 = lean_usize_shift_right(x_13, x_17); -x_19 = lean_unsigned_to_nat(1u); -x_20 = lean_nat_add(x_5, x_19); +x_10 = lean_array_fget(x_3, x_5); +x_11 = l_Lean_Meta_Grind_instHashableENodeKey_unsafe__1(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_21 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_setENode___spec__2(x_6, x_18, x_1, x_10, x_11); +x_20 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_setENode___spec__2(x_6, x_17, x_1, x_9, x_10); x_4 = lean_box(0); -x_5 = x_20; -x_6 = x_21; +x_5 = x_19; +x_6 = x_20; goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_setENode___spec__4(lean_object* x_1, lean_object* x_2, size_t x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_setENode___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; @@ -11561,86 +15376,81 @@ 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; lean_object* x_14; +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_box_usize(x_3); -x_13 = lean_array_push(x_5, x_12); -x_14 = lean_array_push(x_6, x_4); -lean_ctor_set(x_1, 1, x_14); -lean_ctor_set(x_1, 0, x_13); +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_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_dec(x_1); -x_15 = lean_box_usize(x_3); -x_16 = lean_array_push(x_5, x_15); -x_17 = lean_array_push(x_6, x_4); -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; +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_19; size_t x_20; uint8_t x_21; -x_19 = lean_array_fget(x_5, x_2); -x_20 = lean_unbox_usize(x_19); -lean_dec(x_19); -x_21 = lean_usize_dec_eq(x_3, x_20); -if (x_21 == 0) +lean_object* x_17; uint8_t x_18; +x_17 = lean_array_fget(x_5, x_2); +x_18 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_3, x_17); +lean_dec(x_17); +if (x_18 == 0) { -lean_object* x_22; lean_object* x_23; +lean_object* x_19; lean_object* x_20; lean_dec(x_6); lean_dec(x_5); -x_22 = lean_unsigned_to_nat(1u); -x_23 = lean_nat_add(x_2, x_22); +x_19 = lean_unsigned_to_nat(1u); +x_20 = lean_nat_add(x_2, x_19); lean_dec(x_2); -x_2 = x_23; +x_2 = x_20; goto _start; } else { -uint8_t x_25; -x_25 = !lean_is_exclusive(x_1); -if (x_25 == 0) +uint8_t x_22; +x_22 = !lean_is_exclusive(x_1); +if (x_22 == 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_1, 1); -lean_dec(x_26); -x_27 = lean_ctor_get(x_1, 0); -lean_dec(x_27); -x_28 = lean_box_usize(x_3); -x_29 = lean_array_fset(x_5, x_2, x_28); -x_30 = lean_array_fset(x_6, x_2, x_4); +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_30); -lean_ctor_set(x_1, 0, x_29); +lean_ctor_set(x_1, 1, x_26); +lean_ctor_set(x_1, 0, x_25); return x_1; } else { -lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_dec(x_1); -x_31 = lean_box_usize(x_3); -x_32 = lean_array_fset(x_5, x_2, x_31); -x_33 = lean_array_fset(x_6, x_2, x_4); +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_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_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_PersistentHashMap_insertAux___at_Lean_Meta_Grind_setENode___spec__2(lean_object* x_1, size_t x_2, size_t x_3, size_t x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_setENode___spec__2(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) @@ -11663,6 +15473,7 @@ if (x_14 == 0) { lean_dec(x_12); lean_dec(x_5); +lean_dec(x_4); return x_1; } else @@ -11678,120 +15489,113 @@ uint8_t x_18; x_18 = !lean_is_exclusive(x_15); if (x_18 == 0) { -lean_object* x_19; lean_object* x_20; size_t x_21; uint8_t x_22; +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_unbox_usize(x_19); -x_22 = lean_usize_dec_eq(x_4, x_21); -if (x_22 == 0) +x_21 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_4, x_19); +if (x_21 == 0) { -lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_free_object(x_15); -x_23 = lean_box_usize(x_4); -x_24 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_19, x_20, x_23, x_5); -x_25 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_25, 0, x_24); -x_26 = lean_array_fset(x_17, x_12, x_25); +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_26); +lean_ctor_set(x_1, 0, x_24); return x_1; } else { -lean_object* x_27; lean_object* x_28; +lean_object* x_25; lean_dec(x_20); lean_dec(x_19); -x_27 = lean_box_usize(x_4); lean_ctor_set(x_15, 1, x_5); -lean_ctor_set(x_15, 0, x_27); -x_28 = lean_array_fset(x_17, x_12, x_15); +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_28); +lean_ctor_set(x_1, 0, x_25); return x_1; } } else { -lean_object* x_29; lean_object* x_30; size_t x_31; uint8_t x_32; -x_29 = lean_ctor_get(x_15, 0); -x_30 = lean_ctor_get(x_15, 1); -lean_inc(x_30); -lean_inc(x_29); +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_31 = lean_unbox_usize(x_29); -x_32 = lean_usize_dec_eq(x_4, x_31); -if (x_32 == 0) +x_28 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_4, x_26); +if (x_28 == 0) { -lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_33 = lean_box_usize(x_4); -x_34 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_29, x_30, x_33, x_5); -x_35 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_35, 0, x_34); -x_36 = lean_array_fset(x_17, x_12, x_35); +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_36); +lean_ctor_set(x_1, 0, x_31); return x_1; } else { -lean_object* x_37; lean_object* x_38; lean_object* x_39; -lean_dec(x_30); -lean_dec(x_29); -x_37 = lean_box_usize(x_4); -x_38 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_38, 0, x_37); -lean_ctor_set(x_38, 1, x_5); -x_39 = lean_array_fset(x_17, x_12, x_38); +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_39); +lean_ctor_set(x_1, 0, x_33); return x_1; } } } case 1: { -uint8_t x_40; -x_40 = !lean_is_exclusive(x_15); -if (x_40 == 0) +uint8_t x_34; +x_34 = !lean_is_exclusive(x_15); +if (x_34 == 0) { -lean_object* x_41; size_t x_42; size_t x_43; lean_object* x_44; lean_object* x_45; -x_41 = lean_ctor_get(x_15, 0); -x_42 = lean_usize_shift_right(x_2, x_9); -x_43 = lean_usize_add(x_3, x_8); -x_44 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_setENode___spec__2(x_41, x_42, x_43, x_4, x_5); -lean_ctor_set(x_15, 0, x_44); -x_45 = lean_array_fset(x_17, x_12, x_15); +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_Grind_setENode___spec__2(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_45); +lean_ctor_set(x_1, 0, x_39); return x_1; } else { -lean_object* x_46; size_t x_47; size_t x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; -x_46 = lean_ctor_get(x_15, 0); -lean_inc(x_46); +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_47 = lean_usize_shift_right(x_2, x_9); -x_48 = lean_usize_add(x_3, x_8); -x_49 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_setENode___spec__2(x_46, x_47, x_48, x_4, x_5); -x_50 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_50, 0, x_49); -x_51 = lean_array_fset(x_17, x_12, x_50); +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_Grind_setENode___spec__2(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_51); +lean_ctor_set(x_1, 0, x_45); return x_1; } } default: { -lean_object* x_52; lean_object* x_53; lean_object* x_54; -x_52 = lean_box_usize(x_4); -x_53 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_53, 0, x_52); -lean_ctor_set(x_53, 1, x_5); -x_54 = lean_array_fset(x_17, x_12, x_53); +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_54); +lean_ctor_set(x_1, 0, x_47); return x_1; } } @@ -11799,124 +15603,121 @@ return x_1; } else { -lean_object* x_55; size_t x_56; size_t x_57; size_t x_58; size_t x_59; lean_object* x_60; lean_object* x_61; uint8_t x_62; -x_55 = lean_ctor_get(x_1, 0); -lean_inc(x_55); +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_56 = 1; -x_57 = 5; -x_58 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__2___closed__2; -x_59 = lean_usize_land(x_2, x_58); -x_60 = lean_usize_to_nat(x_59); -x_61 = lean_array_get_size(x_55); -x_62 = lean_nat_dec_lt(x_60, x_61); -lean_dec(x_61); -if (x_62 == 0) +x_49 = 1; +x_50 = 5; +x_51 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___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_63; -lean_dec(x_60); +lean_object* x_56; +lean_dec(x_53); lean_dec(x_5); -x_63 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_63, 0, x_55); -return x_63; +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_64; lean_object* x_65; lean_object* x_66; -x_64 = lean_array_fget(x_55, x_60); -x_65 = lean_box(0); -x_66 = lean_array_fset(x_55, x_60, x_65); -switch (lean_obj_tag(x_64)) { +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_67; lean_object* x_68; lean_object* x_69; size_t x_70; uint8_t x_71; -x_67 = lean_ctor_get(x_64, 0); -lean_inc(x_67); -x_68 = lean_ctor_get(x_64, 1); -lean_inc(x_68); -if (lean_is_exclusive(x_64)) { - lean_ctor_release(x_64, 0); - lean_ctor_release(x_64, 1); - x_69 = x_64; +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_64); - x_69 = lean_box(0); + lean_dec_ref(x_57); + x_62 = lean_box(0); } -x_70 = lean_unbox_usize(x_67); -x_71 = lean_usize_dec_eq(x_4, x_70); -if (x_71 == 0) +x_63 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_4, x_60); +if (x_63 == 0) { -lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; -lean_dec(x_69); -x_72 = lean_box_usize(x_4); -x_73 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_67, x_68, x_72, x_5); -x_74 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_74, 0, x_73); -x_75 = lean_array_fset(x_66, x_60, x_74); -lean_dec(x_60); -x_76 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_76, 0, x_75); -return x_76; +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_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; -lean_dec(x_68); -lean_dec(x_67); -x_77 = lean_box_usize(x_4); -if (lean_is_scalar(x_69)) { - x_78 = lean_alloc_ctor(0, 2, 0); +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_78 = x_69; + x_68 = x_62; } -lean_ctor_set(x_78, 0, x_77); -lean_ctor_set(x_78, 1, x_5); -x_79 = lean_array_fset(x_66, x_60, x_78); -lean_dec(x_60); -x_80 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_80, 0, x_79); -return x_80; +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_81; lean_object* x_82; size_t x_83; size_t x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; -x_81 = lean_ctor_get(x_64, 0); -lean_inc(x_81); -if (lean_is_exclusive(x_64)) { - lean_ctor_release(x_64, 0); - x_82 = x_64; +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_64); - x_82 = lean_box(0); + lean_dec_ref(x_57); + x_72 = lean_box(0); } -x_83 = lean_usize_shift_right(x_2, x_57); -x_84 = lean_usize_add(x_3, x_56); -x_85 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_setENode___spec__2(x_81, x_83, x_84, x_4, x_5); -if (lean_is_scalar(x_82)) { - x_86 = lean_alloc_ctor(1, 1, 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_Grind_setENode___spec__2(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_86 = x_82; + x_76 = x_72; } -lean_ctor_set(x_86, 0, x_85); -x_87 = lean_array_fset(x_66, x_60, x_86); -lean_dec(x_60); -x_88 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_88, 0, x_87); -return x_88; +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_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; -x_89 = lean_box_usize(x_4); -x_90 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_90, 0, x_89); -lean_ctor_set(x_90, 1, x_5); -x_91 = lean_array_fset(x_66, x_60, x_90); -lean_dec(x_60); -x_92 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_92, 0, x_91); -return x_92; +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; } } } @@ -11924,100 +15725,100 @@ return x_92; } else { -uint8_t x_93; -x_93 = !lean_is_exclusive(x_1); -if (x_93 == 0) +uint8_t x_82; +x_82 = !lean_is_exclusive(x_1); +if (x_82 == 0) { -lean_object* x_94; lean_object* x_95; size_t x_96; uint8_t x_97; -x_94 = lean_unsigned_to_nat(0u); -x_95 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_setENode___spec__4(x_1, x_94, x_4, x_5); -x_96 = 7; -x_97 = lean_usize_dec_le(x_96, x_3); -if (x_97 == 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_Grind_setENode___spec__4(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_98; lean_object* x_99; uint8_t x_100; -x_98 = l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(x_95); -x_99 = lean_unsigned_to_nat(4u); -x_100 = lean_nat_dec_lt(x_98, x_99); -lean_dec(x_98); -if (x_100 == 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_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_inc(x_102); -lean_dec(x_95); -x_103 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__2___closed__3; -x_104 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_setENode___spec__3(x_3, x_101, x_102, lean_box(0), x_94, x_103); -lean_dec(x_102); -lean_dec(x_101); -return x_104; +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_Grind_mkHCongrWithArity___spec__2___closed__3; +x_93 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_setENode___spec__3(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_95; +return x_84; } } else { -return x_95; +return x_84; } } else { -lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; size_t x_110; uint8_t x_111; -x_105 = lean_ctor_get(x_1, 0); -x_106 = lean_ctor_get(x_1, 1); -lean_inc(x_106); -lean_inc(x_105); +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_107 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_107, 0, x_105); -lean_ctor_set(x_107, 1, x_106); -x_108 = lean_unsigned_to_nat(0u); -x_109 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_setENode___spec__4(x_107, x_108, x_4, x_5); -x_110 = 7; -x_111 = lean_usize_dec_le(x_110, x_3); -if (x_111 == 0) +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_Grind_setENode___spec__4(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_112; lean_object* x_113; uint8_t x_114; -x_112 = l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(x_109); -x_113 = lean_unsigned_to_nat(4u); -x_114 = lean_nat_dec_lt(x_112, x_113); -lean_dec(x_112); -if (x_114 == 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_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; -x_115 = lean_ctor_get(x_109, 0); -lean_inc(x_115); -x_116 = lean_ctor_get(x_109, 1); -lean_inc(x_116); -lean_dec(x_109); -x_117 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__2___closed__3; -x_118 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_setENode___spec__3(x_3, x_115, x_116, lean_box(0), x_108, x_117); -lean_dec(x_116); -lean_dec(x_115); -return x_118; +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_Grind_mkHCongrWithArity___spec__2___closed__3; +x_107 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_setENode___spec__3(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_109; +return x_98; } } else { -return x_109; +return x_98; } } } } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_setENode___spec__1(lean_object* x_1, size_t x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_setENode___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { uint64_t x_4; size_t x_5; size_t x_6; lean_object* x_7; -x_4 = lean_usize_to_uint64(x_2); +x_4 = l_Lean_Meta_Grind_instHashableENodeKey_unsafe__1(x_2); x_5 = lean_uint64_to_usize(x_4); x_6 = 1; x_7 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_setENode___spec__2(x_1, x_5, x_6, x_2, x_3); @@ -12037,88 +15838,104 @@ lean_dec(x_12); x_15 = !lean_is_exclusive(x_13); if (x_15 == 0) { -lean_object* x_16; size_t 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; uint8_t x_19; x_16 = lean_ctor_get(x_13, 1); -x_17 = lean_ptr_addr(x_1); -x_18 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_setENode___spec__1(x_16, x_17, x_2); -lean_ctor_set(x_13, 1, x_18); -x_19 = lean_st_ref_set(x_3, x_13, x_14); -x_20 = !lean_is_exclusive(x_19); -if (x_20 == 0) +x_17 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_setENode___spec__1(x_16, x_1, x_2); +lean_ctor_set(x_13, 1, x_17); +x_18 = lean_st_ref_set(x_3, x_13, x_14); +x_19 = !lean_is_exclusive(x_18); +if (x_19 == 0) { -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; +lean_object* x_20; lean_object* x_21; +x_20 = lean_ctor_get(x_18, 0); +lean_dec(x_20); +x_21 = lean_box(0); +lean_ctor_set(x_18, 0, x_21); +return x_18; } else { -lean_object* x_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; +lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_22 = lean_ctor_get(x_18, 1); +lean_inc(x_22); +lean_dec(x_18); +x_23 = lean_box(0); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set(x_24, 1, x_22); +return x_24; } } else { -lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; uint8_t x_32; lean_object* x_33; lean_object* x_34; size_t x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; -x_26 = lean_ctor_get(x_13, 0); -x_27 = lean_ctor_get(x_13, 1); -x_28 = lean_ctor_get(x_13, 2); -x_29 = lean_ctor_get(x_13, 3); -x_30 = lean_ctor_get(x_13, 4); -x_31 = lean_ctor_get(x_13, 5); -x_32 = lean_ctor_get_uint8(x_13, sizeof(void*)*8); -x_33 = lean_ctor_get(x_13, 6); -x_34 = lean_ctor_get(x_13, 7); +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; 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_25 = lean_ctor_get(x_13, 0); +x_26 = lean_ctor_get(x_13, 1); +x_27 = lean_ctor_get(x_13, 2); +x_28 = lean_ctor_get(x_13, 3); +x_29 = lean_ctor_get(x_13, 4); +x_30 = lean_ctor_get(x_13, 5); +x_31 = lean_ctor_get_uint8(x_13, sizeof(void*)*14); +x_32 = lean_ctor_get(x_13, 6); +x_33 = lean_ctor_get(x_13, 7); +x_34 = lean_ctor_get(x_13, 8); +x_35 = lean_ctor_get(x_13, 9); +x_36 = lean_ctor_get(x_13, 10); +x_37 = lean_ctor_get(x_13, 11); +x_38 = lean_ctor_get(x_13, 12); +x_39 = lean_ctor_get(x_13, 13); +lean_inc(x_39); +lean_inc(x_38); +lean_inc(x_37); +lean_inc(x_36); +lean_inc(x_35); lean_inc(x_34); lean_inc(x_33); -lean_inc(x_31); +lean_inc(x_32); 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_dec(x_13); -x_35 = lean_ptr_addr(x_1); -x_36 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_setENode___spec__1(x_27, x_35, x_2); -x_37 = lean_alloc_ctor(0, 8, 1); -lean_ctor_set(x_37, 0, x_26); -lean_ctor_set(x_37, 1, x_36); -lean_ctor_set(x_37, 2, x_28); -lean_ctor_set(x_37, 3, x_29); -lean_ctor_set(x_37, 4, x_30); -lean_ctor_set(x_37, 5, x_31); -lean_ctor_set(x_37, 6, x_33); -lean_ctor_set(x_37, 7, x_34); -lean_ctor_set_uint8(x_37, sizeof(void*)*8, x_32); -x_38 = lean_st_ref_set(x_3, x_37, x_14); -x_39 = lean_ctor_get(x_38, 1); -lean_inc(x_39); -if (lean_is_exclusive(x_38)) { - lean_ctor_release(x_38, 0); - lean_ctor_release(x_38, 1); - x_40 = x_38; +x_40 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_setENode___spec__1(x_26, x_1, x_2); +x_41 = lean_alloc_ctor(0, 14, 1); +lean_ctor_set(x_41, 0, x_25); +lean_ctor_set(x_41, 1, x_40); +lean_ctor_set(x_41, 2, x_27); +lean_ctor_set(x_41, 3, x_28); +lean_ctor_set(x_41, 4, x_29); +lean_ctor_set(x_41, 5, x_30); +lean_ctor_set(x_41, 6, x_32); +lean_ctor_set(x_41, 7, x_33); +lean_ctor_set(x_41, 8, x_34); +lean_ctor_set(x_41, 9, x_35); +lean_ctor_set(x_41, 10, x_36); +lean_ctor_set(x_41, 11, x_37); +lean_ctor_set(x_41, 12, x_38); +lean_ctor_set(x_41, 13, x_39); +lean_ctor_set_uint8(x_41, sizeof(void*)*14, x_31); +x_42 = lean_st_ref_set(x_3, x_41, x_14); +x_43 = lean_ctor_get(x_42, 1); +lean_inc(x_43); +if (lean_is_exclusive(x_42)) { + lean_ctor_release(x_42, 0); + lean_ctor_release(x_42, 1); + x_44 = x_42; } else { - lean_dec_ref(x_38); - x_40 = lean_box(0); + lean_dec_ref(x_42); + x_44 = lean_box(0); } -x_41 = lean_box(0); -if (lean_is_scalar(x_40)) { - x_42 = lean_alloc_ctor(0, 2, 0); +x_45 = lean_box(0); +if (lean_is_scalar(x_44)) { + x_46 = lean_alloc_ctor(0, 2, 0); } else { - x_42 = x_40; + x_46 = x_44; } -lean_ctor_set(x_42, 0, x_41); -lean_ctor_set(x_42, 1, x_39); -return x_42; +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_43); +return x_46; } } } @@ -12134,38 +15951,16 @@ lean_dec(x_2); return x_8; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_setENode___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -size_t x_5; lean_object* x_6; -x_5 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_6 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_setENode___spec__4(x_1, x_2, x_5, x_4); -return x_6; -} -} LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_setENode___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -size_t x_6; size_t x_7; size_t x_8; lean_object* x_9; +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 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_9 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_setENode___spec__2(x_1, x_6, x_7, x_8, x_5); -return x_9; -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_setENode___spec__1___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_insert___at_Lean_Meta_Grind_setENode___spec__1(x_1, x_4, x_3); -return x_5; +x_8 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_setENode___spec__2(x_1, x_6, x_7, x_4, x_5); +return x_8; } } LEAN_EXPORT lean_object* l_Lean_Meta_Grind_setENode___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { @@ -12181,7 +15976,6 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -lean_dec(x_1); return x_12; } } @@ -12229,7 +16023,6 @@ lean_ctor_set_uint8(x_26, sizeof(void*)*10 + 2, x_3); lean_ctor_set_uint8(x_26, sizeof(void*)*10 + 3, x_21); lean_ctor_set_uint8(x_26, sizeof(void*)*10 + 4, x_24); x_27 = l_Lean_Meta_Grind_setENode(x_1, x_26, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_19); -lean_dec(x_1); x_28 = lean_ctor_get(x_27, 1); lean_inc(x_28); lean_dec(x_27); @@ -12273,16 +16066,28 @@ return x_41; } else { -lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; 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_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; lean_object* x_62; lean_object* x_63; x_42 = lean_ctor_get(x_30, 0); x_43 = lean_ctor_get(x_30, 1); x_44 = lean_ctor_get(x_30, 2); x_45 = lean_ctor_get(x_30, 3); x_46 = lean_ctor_get(x_30, 4); x_47 = lean_ctor_get(x_30, 5); -x_48 = lean_ctor_get_uint8(x_30, sizeof(void*)*8); +x_48 = lean_ctor_get_uint8(x_30, sizeof(void*)*14); x_49 = lean_ctor_get(x_30, 6); x_50 = lean_ctor_get(x_30, 7); +x_51 = lean_ctor_get(x_30, 8); +x_52 = lean_ctor_get(x_30, 9); +x_53 = lean_ctor_get(x_30, 10); +x_54 = lean_ctor_get(x_30, 11); +x_55 = lean_ctor_get(x_30, 12); +x_56 = lean_ctor_get(x_30, 13); +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_inc(x_49); lean_inc(x_47); @@ -12292,38 +16097,44 @@ lean_inc(x_44); lean_inc(x_43); lean_inc(x_42); lean_dec(x_30); -x_51 = lean_nat_add(x_50, x_25); +x_57 = lean_nat_add(x_50, x_25); lean_dec(x_50); -x_52 = lean_alloc_ctor(0, 8, 1); -lean_ctor_set(x_52, 0, x_42); -lean_ctor_set(x_52, 1, x_43); -lean_ctor_set(x_52, 2, x_44); -lean_ctor_set(x_52, 3, x_45); -lean_ctor_set(x_52, 4, x_46); -lean_ctor_set(x_52, 5, x_47); -lean_ctor_set(x_52, 6, x_49); -lean_ctor_set(x_52, 7, x_51); -lean_ctor_set_uint8(x_52, sizeof(void*)*8, x_48); -x_53 = lean_st_ref_set(x_5, x_52, x_31); -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; +x_58 = lean_alloc_ctor(0, 14, 1); +lean_ctor_set(x_58, 0, x_42); +lean_ctor_set(x_58, 1, x_43); +lean_ctor_set(x_58, 2, x_44); +lean_ctor_set(x_58, 3, x_45); +lean_ctor_set(x_58, 4, x_46); +lean_ctor_set(x_58, 5, x_47); +lean_ctor_set(x_58, 6, x_49); +lean_ctor_set(x_58, 7, x_57); +lean_ctor_set(x_58, 8, x_51); +lean_ctor_set(x_58, 9, x_52); +lean_ctor_set(x_58, 10, x_53); +lean_ctor_set(x_58, 11, x_54); +lean_ctor_set(x_58, 12, x_55); +lean_ctor_set(x_58, 13, x_56); +lean_ctor_set_uint8(x_58, sizeof(void*)*14, x_48); +x_59 = lean_st_ref_set(x_5, x_58, x_31); +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_53); - x_55 = lean_box(0); + lean_dec_ref(x_59); + x_61 = lean_box(0); } -x_56 = lean_box(0); -if (lean_is_scalar(x_55)) { - x_57 = lean_alloc_ctor(0, 2, 0); +x_62 = lean_box(0); +if (lean_is_scalar(x_61)) { + x_63 = lean_alloc_ctor(0, 2, 0); } else { - x_57 = x_55; + x_63 = x_61; } -lean_ctor_set(x_57, 0, x_56); -lean_ctor_set(x_57, 1, x_54); -return x_57; +lean_ctor_set(x_63, 0, x_62); +lean_ctor_set(x_63, 1, x_60); +return x_63; } } } @@ -12542,6 +16353,92 @@ lean_dec(x_3); return x_12; } } +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isCongrRoot(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +lean_inc(x_1); +x_11 = l_Lean_Meta_Grind_getENode(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; uint8_t x_15; lean_object* x_16; +x_13 = lean_ctor_get(x_11, 0); +x_14 = lean_ctor_get(x_13, 3); +lean_inc(x_14); +lean_dec(x_13); +x_15 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_1, x_14); +lean_dec(x_14); +lean_dec(x_1); +x_16 = lean_box(x_15); +lean_ctor_set(x_11, 0, x_16); +return x_11; +} +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; lean_object* x_21; lean_object* x_22; +x_17 = lean_ctor_get(x_11, 0); +x_18 = lean_ctor_get(x_11, 1); +lean_inc(x_18); +lean_inc(x_17); +lean_dec(x_11); +x_19 = lean_ctor_get(x_17, 3); +lean_inc(x_19); +lean_dec(x_17); +x_20 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_1, x_19); +lean_dec(x_19); +lean_dec(x_1); +x_21 = lean_box(x_20); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_18); +return x_22; +} +} +else +{ +uint8_t x_23; +lean_dec(x_1); +x_23 = !lean_is_exclusive(x_11); +if (x_23 == 0) +{ +return x_11; +} +else +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_11, 0); +x_25 = lean_ctor_get(x_11, 1); +lean_inc(x_25); +lean_inc(x_24); +lean_dec(x_11); +x_26 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_26, 0, x_24); +lean_ctor_set(x_26, 1, x_25); +return x_26; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isCongrRoot___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_Meta_Grind_isCongrRoot(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_11; +} +} LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isInconsistent(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { @@ -12552,7 +16449,7 @@ if (x_11 == 0) { lean_object* x_12; uint8_t x_13; lean_object* x_14; x_12 = lean_ctor_get(x_10, 0); -x_13 = lean_ctor_get_uint8(x_12, sizeof(void*)*8); +x_13 = lean_ctor_get_uint8(x_12, sizeof(void*)*14); lean_dec(x_12); x_14 = lean_box(x_13); lean_ctor_set(x_10, 0, x_14); @@ -12566,7 +16463,7 @@ x_16 = lean_ctor_get(x_10, 1); lean_inc(x_16); lean_inc(x_15); lean_dec(x_10); -x_17 = lean_ctor_get_uint8(x_15, sizeof(void*)*8); +x_17 = lean_ctor_get_uint8(x_15, sizeof(void*)*14); lean_dec(x_15); x_18 = lean_box(x_17); x_19 = lean_alloc_ctor(0, 2, 0); @@ -12600,6 +16497,85 @@ x_12 = lean_grind_mk_eq_proof(x_1, 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_Meta_Grind_mkHEqProof___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; +x_12 = lean_grind_mk_heq_proof(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEqHEqProof(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, 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_2); +lean_inc(x_1); +x_12 = l_Lean_Meta_Grind_hasSameType(x_1, x_2, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_12) == 0) +{ +lean_object* x_13; uint8_t x_14; +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_unbox(x_13); +lean_dec(x_13); +if (x_14 == 0) +{ +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_grind_mk_heq_proof(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_15); +return x_16; +} +else +{ +lean_object* x_17; lean_object* x_18; +x_17 = lean_ctor_get(x_12, 1); +lean_inc(x_17); +lean_dec(x_12); +x_18 = lean_grind_mk_eq_proof(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_17); +return x_18; +} +} +else +{ +uint8_t x_19; +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_19 = !lean_is_exclusive(x_12); +if (x_19 == 0) +{ +return x_12; +} +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_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_EXPORT lean_object* l_Lean_Meta_Grind_mkEqTrueProof(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { @@ -12643,7 +16619,7 @@ if (x_13 == 0) { uint8_t x_14; lean_object* x_15; uint8_t x_16; x_14 = 1; -lean_ctor_set_uint8(x_11, sizeof(void*)*8, x_14); +lean_ctor_set_uint8(x_11, sizeof(void*)*14, x_14); x_15 = lean_st_ref_set(x_1, x_11, x_12); x_16 = !lean_is_exclusive(x_15); if (x_16 == 0) @@ -12670,7 +16646,7 @@ return x_21; } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; uint8_t x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; x_22 = lean_ctor_get(x_11, 0); x_23 = lean_ctor_get(x_11, 1); x_24 = lean_ctor_get(x_11, 2); @@ -12679,6 +16655,18 @@ x_26 = lean_ctor_get(x_11, 4); x_27 = lean_ctor_get(x_11, 5); x_28 = lean_ctor_get(x_11, 6); x_29 = lean_ctor_get(x_11, 7); +x_30 = lean_ctor_get(x_11, 8); +x_31 = lean_ctor_get(x_11, 9); +x_32 = lean_ctor_get(x_11, 10); +x_33 = lean_ctor_get(x_11, 11); +x_34 = lean_ctor_get(x_11, 12); +x_35 = lean_ctor_get(x_11, 13); +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); @@ -12688,37 +16676,43 @@ lean_inc(x_24); lean_inc(x_23); lean_inc(x_22); lean_dec(x_11); -x_30 = 1; -x_31 = lean_alloc_ctor(0, 8, 1); -lean_ctor_set(x_31, 0, x_22); -lean_ctor_set(x_31, 1, x_23); -lean_ctor_set(x_31, 2, x_24); -lean_ctor_set(x_31, 3, x_25); -lean_ctor_set(x_31, 4, x_26); -lean_ctor_set(x_31, 5, x_27); -lean_ctor_set(x_31, 6, x_28); -lean_ctor_set(x_31, 7, x_29); -lean_ctor_set_uint8(x_31, sizeof(void*)*8, x_30); -x_32 = lean_st_ref_set(x_1, x_31, x_12); -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; +x_36 = 1; +x_37 = lean_alloc_ctor(0, 14, 1); +lean_ctor_set(x_37, 0, x_22); +lean_ctor_set(x_37, 1, x_23); +lean_ctor_set(x_37, 2, x_24); +lean_ctor_set(x_37, 3, x_25); +lean_ctor_set(x_37, 4, x_26); +lean_ctor_set(x_37, 5, x_27); +lean_ctor_set(x_37, 6, x_28); +lean_ctor_set(x_37, 7, x_29); +lean_ctor_set(x_37, 8, x_30); +lean_ctor_set(x_37, 9, x_31); +lean_ctor_set(x_37, 10, x_32); +lean_ctor_set(x_37, 11, x_33); +lean_ctor_set(x_37, 12, x_34); +lean_ctor_set(x_37, 13, x_35); +lean_ctor_set_uint8(x_37, sizeof(void*)*14, x_36); +x_38 = lean_st_ref_set(x_1, x_37, x_12); +x_39 = lean_ctor_get(x_38, 1); +lean_inc(x_39); +if (lean_is_exclusive(x_38)) { + lean_ctor_release(x_38, 0); + lean_ctor_release(x_38, 1); + x_40 = x_38; } else { - lean_dec_ref(x_32); - x_34 = lean_box(0); + lean_dec_ref(x_38); + x_40 = lean_box(0); } -x_35 = lean_box(0); -if (lean_is_scalar(x_34)) { - x_36 = lean_alloc_ctor(0, 2, 0); +x_41 = lean_box(0); +if (lean_is_scalar(x_40)) { + x_42 = lean_alloc_ctor(0, 2, 0); } else { - x_36 = x_34; + x_42 = x_40; } -lean_ctor_set(x_36, 0, x_35); -lean_ctor_set(x_36, 1, x_33); -return x_36; +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_39); +return x_42; } } } @@ -13195,16 +17189,15 @@ x_4 = l_Lean_PersistentHashMap_foldlMAux___at_Lean_MetavarContext_getExprAssignm return x_4; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Grind_getENodes___spec__1___lambda__1(lean_object* x_1, size_t x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Grind_getENodes___spec__1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_4 = lean_box_usize(x_2); -x_5 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_5, 0, x_4); -lean_ctor_set(x_5, 1, x_3); -x_6 = lean_array_push(x_1, x_5); -return x_6; +lean_object* x_4; lean_object* x_5; +x_4 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_4, 0, x_2); +lean_ctor_set(x_4, 1, x_3); +x_5 = lean_array_push(x_1, x_4); +return x_5; } } static lean_object* _init_l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Grind_getENodes___spec__1___closed__1() { @@ -13220,7 +17213,7 @@ static lean_object* _init_l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Grind_ _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Grind_getENodes___spec__1___lambda__1___boxed), 3, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Grind_getENodes___spec__1___lambda__1), 3, 0); return x_1; } } @@ -13449,16 +17442,6 @@ lean_dec(x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Grind_getENodes___spec__1___lambda__1___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_toArray___at_Lean_Meta_Grind_getENodes___spec__1___lambda__1(x_1, x_4, x_3); -return x_5; -} -} LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Grind_getENodes___spec__1___boxed(lean_object* x_1) { _start: { @@ -14523,6 +18506,8 @@ lean_dec(x_1); return x_10; } } +lean_object* initialize_Init_Grind_Tactics(uint8_t builtin, lean_object*); +lean_object* initialize_Init_Data_Queue(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Util_ShareCommon(uint8_t builtin, lean_object*); lean_object* initialize_Lean_HeadIndex(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_Basic(uint8_t builtin, lean_object*); @@ -14532,11 +18517,18 @@ lean_object* initialize_Lean_Meta_Tactic_Simp_Types(uint8_t builtin, lean_object lean_object* initialize_Lean_Meta_Tactic_Util(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_Tactic_Grind_Canon(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_Tactic_Grind_Attr(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Meta_Tactic_Grind_EMatchTheorem(uint8_t builtin, lean_object*); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Lean_Meta_Tactic_Grind_Types(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_Grind_Tactics(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Init_Data_Queue(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); res = initialize_Lean_Util_ShareCommon(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); @@ -14564,6 +18556,9 @@ lean_dec_ref(res); res = initialize_Lean_Meta_Tactic_Grind_Attr(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +res = initialize_Lean_Meta_Tactic_Grind_EMatchTheorem(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); l_Lean_Meta_Grind_congrPlaceholderProof___closed__1 = _init_l_Lean_Meta_Grind_congrPlaceholderProof___closed__1(); lean_mark_persistent(l_Lean_Meta_Grind_congrPlaceholderProof___closed__1); l_Lean_Meta_Grind_congrPlaceholderProof___closed__2 = _init_l_Lean_Meta_Grind_congrPlaceholderProof___closed__2(); @@ -14633,114 +18628,114 @@ l_Lean_Meta_Grind_instInhabitedENode___closed__2 = _init_l_Lean_Meta_Grind_instI lean_mark_persistent(l_Lean_Meta_Grind_instInhabitedENode___closed__2); l_Lean_Meta_Grind_instInhabitedENode = _init_l_Lean_Meta_Grind_instInhabitedENode(); lean_mark_persistent(l_Lean_Meta_Grind_instInhabitedENode); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__1(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__1); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__2(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__2); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__3(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__3); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__4(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__4); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__5 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__5(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__5); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__6 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__6(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__6); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__7 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__7(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__7); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__8 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__8(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__8); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__9 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__9(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__9); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__10 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__10(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__10); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__11 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__11(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__11); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__12 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__12(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__12); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__13 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__13(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__13); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__14 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__14(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__14); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__15 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__15(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__15); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__16 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__16(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__16); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__17 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__17(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__17); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__18 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__18(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__18); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__19 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__19(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__19); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__20 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__20(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__20); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__21 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__21(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__21); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__22 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__22(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__22); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__23 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__23(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__23); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__24 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__24(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__24); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__25 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__25(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__25); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__26 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__26(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__26); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__27 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__27(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__27); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__28 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__28(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__28); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__29 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__29(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__29); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__30 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__30(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__30); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__31 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__31(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__31); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__32 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__32(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__32); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__33 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__33(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__33); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__34 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__34(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__34); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__35 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__35(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__35); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__36 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__36(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__36); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__37 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__37(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__37); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__38 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__38(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__38); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__39 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__39(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__39); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__40 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__40(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__40); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__41 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__41(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__41); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__42 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__42(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__42); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__43 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__43(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__43); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__44 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__44(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__44); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__45 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__45(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__45); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__46 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__46(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__46); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__47 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__47(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__47); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__48 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__48(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__48); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__49 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__49(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__49); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__50 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__50(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__50); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__51 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__51(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__51); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__52 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__52(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__52); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__53 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__53(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__53); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__54 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__54(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1389____closed__54); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__1); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__2); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__3); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__4(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__4); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__5 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__5(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__5); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__6 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__6(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__6); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__7 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__7(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__7); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__8 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__8(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__8); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__9 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__9(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__9); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__10 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__10(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__10); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__11 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__11(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__11); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__12 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__12(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__12); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__13 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__13(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__13); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__14 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__14(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__14); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__15 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__15(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__15); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__16 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__16(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__16); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__17 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__17(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__17); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__18 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__18(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__18); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__19 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__19(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__19); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__20 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__20(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__20); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__21 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__21(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__21); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__22 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__22(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__22); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__23 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__23(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__23); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__24 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__24(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__24); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__25 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__25(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__25); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__26 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__26(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__26); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__27 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__27(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__27); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__28 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__28(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__28); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__29 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__29(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__29); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__30 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__30(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__30); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__31 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__31(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__31); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__32 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__32(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__32); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__33 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__33(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__33); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__34 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__34(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__34); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__35 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__35(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__35); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__36 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__36(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__36); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__37 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__37(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__37); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__38 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__38(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__38); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__39 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__39(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__39); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__40 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__40(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__40); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__41 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__41(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__41); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__42 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__42(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__42); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__43 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__43(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__43); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__44 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__44(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__44); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__45 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__45(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__45); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__46 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__46(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__46); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__47 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__47(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__47); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__48 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__48(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__48); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__49 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__49(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__49); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__50 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__50(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__50); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__51 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__51(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__51); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__52 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__52(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__52); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__53 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__53(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__53); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__54 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__54(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__54); l_Lean_Meta_Grind_instReprENode___closed__1 = _init_l_Lean_Meta_Grind_instReprENode___closed__1(); lean_mark_persistent(l_Lean_Meta_Grind_instReprENode___closed__1); l_Lean_Meta_Grind_instReprENode = _init_l_Lean_Meta_Grind_instReprENode(); @@ -14749,12 +18744,33 @@ l_Lean_Meta_Grind_congrHash___closed__1 = _init_l_Lean_Meta_Grind_congrHash___cl lean_mark_persistent(l_Lean_Meta_Grind_congrHash___closed__1); l_Lean_Meta_Grind_congrHash___closed__2 = _init_l_Lean_Meta_Grind_congrHash___closed__2(); lean_mark_persistent(l_Lean_Meta_Grind_congrHash___closed__2); +l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1___closed__1 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1___closed__1(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1___closed__1); +l_Lean_Meta_Grind_instBEqPreInstance___lambda__2___closed__1 = _init_l_Lean_Meta_Grind_instBEqPreInstance___lambda__2___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_instBEqPreInstance___lambda__2___closed__1); +l_Lean_Meta_Grind_instBEqPreInstance___lambda__2___closed__2 = _init_l_Lean_Meta_Grind_instBEqPreInstance___lambda__2___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_instBEqPreInstance___lambda__2___closed__2); +l_Lean_Meta_Grind_instInhabitedNewFact___closed__1 = _init_l_Lean_Meta_Grind_instInhabitedNewFact___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_instInhabitedNewFact___closed__1); +l_Lean_Meta_Grind_instInhabitedNewFact = _init_l_Lean_Meta_Grind_instInhabitedNewFact(); +lean_mark_persistent(l_Lean_Meta_Grind_instInhabitedNewFact); l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedGoal___spec__1 = _init_l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedGoal___spec__1(); lean_mark_persistent(l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedGoal___spec__1); +l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedGoal___spec__2 = _init_l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedGoal___spec__2(); +lean_mark_persistent(l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedGoal___spec__2); l_Lean_Meta_Grind_instInhabitedGoal___closed__1 = _init_l_Lean_Meta_Grind_instInhabitedGoal___closed__1(); lean_mark_persistent(l_Lean_Meta_Grind_instInhabitedGoal___closed__1); +l_Lean_Meta_Grind_instInhabitedGoal___closed__2 = _init_l_Lean_Meta_Grind_instInhabitedGoal___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_instInhabitedGoal___closed__2); +l_Lean_Meta_Grind_instInhabitedGoal___closed__3 = _init_l_Lean_Meta_Grind_instInhabitedGoal___closed__3(); +l_Lean_Meta_Grind_instInhabitedGoal___closed__4 = _init_l_Lean_Meta_Grind_instInhabitedGoal___closed__4(); +lean_mark_persistent(l_Lean_Meta_Grind_instInhabitedGoal___closed__4); +l_Lean_Meta_Grind_instInhabitedGoal___closed__5 = _init_l_Lean_Meta_Grind_instInhabitedGoal___closed__5(); +lean_mark_persistent(l_Lean_Meta_Grind_instInhabitedGoal___closed__5); l_Lean_Meta_Grind_instInhabitedGoal = _init_l_Lean_Meta_Grind_instInhabitedGoal(); lean_mark_persistent(l_Lean_Meta_Grind_instInhabitedGoal); +l_Lean_Meta_Grind_GoalM_run_x27___closed__1 = _init_l_Lean_Meta_Grind_GoalM_run_x27___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_GoalM_run_x27___closed__1); l_Lean_Meta_Grind_getENode___closed__1 = _init_l_Lean_Meta_Grind_getENode___closed__1(); lean_mark_persistent(l_Lean_Meta_Grind_getENode___closed__1); l_Lean_Meta_Grind_getENode___closed__2 = _init_l_Lean_Meta_Grind_getENode___closed__2(); diff --git a/stage0/stdlib/Lean/ToExpr.c b/stage0/stdlib/Lean/ToExpr.c index d11f41c1a7d0..86b7e9a0c415 100644 --- a/stage0/stdlib/Lean/ToExpr.c +++ b/stage0/stdlib/Lean/ToExpr.c @@ -2030,7 +2030,7 @@ static lean_object* _init_l___private_Lean_ToExpr_0__Lean_Name_toExprAux_mkStr__ 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_ToExpr_0__Lean_Name_toExprAux_mkStr___closed__5; x_2 = l___private_Lean_ToExpr_0__Lean_Name_toExprAux_mkStr___closed__6; -x_3 = lean_unsigned_to_nat(131u); +x_3 = lean_unsigned_to_nat(136u); x_4 = lean_unsigned_to_nat(11u); x_5 = l___private_Lean_ToExpr_0__Lean_Name_toExprAux_mkStr___closed__7; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); From 3cba17140f85f1544b48e68941918c58244851cf Mon Sep 17 00:00:00 2001 From: Leonardo de Moura Date: Thu, 2 Jan 2025 17:06:52 +0100 Subject: [PATCH 005/100] fix: missing case in `checkParents` (#6508) This PR fixes a bug in the sanity checkers for the `grind` tactic. See the new test for an example of a case where it was panicking. --- src/Lean/Meta/Tactic/Grind/Inv.lean | 4 ++++ tests/lean/run/grind_ematch2.lean | 19 +++++++++++++++++++ 2 files changed, 23 insertions(+) create mode 100644 tests/lean/run/grind_ematch2.lean diff --git a/src/Lean/Meta/Tactic/Grind/Inv.lean b/src/Lean/Meta/Tactic/Grind/Inv.lean index 4f49167a3f14..b0e77d7a3069 100644 --- a/src/Lean/Meta/Tactic/Grind/Inv.lean +++ b/src/Lean/Meta/Tactic/Grind/Inv.lean @@ -57,6 +57,10 @@ private def checkParents (e : Expr) : GoalM Unit := do if (← checkChild arg) then found := true break + -- Recall that we have support for `Expr.forallE` propagation. See `ForallProp.lean`. + if let .forallE _ d _ _ := parent then + if (← checkChild d) then + found := true unless found do assert! (← checkChild parent.getAppFn) else diff --git a/tests/lean/run/grind_ematch2.lean b/tests/lean/run/grind_ematch2.lean new file mode 100644 index 000000000000..08d2238d17c8 --- /dev/null +++ b/tests/lean/run/grind_ematch2.lean @@ -0,0 +1,19 @@ +grind_pattern Array.size_set => Array.set a i v h +grind_pattern Array.get_set_eq => a.set i v h +grind_pattern Array.get_set_ne => (a.set i v hi)[j] + +set_option grind.debug true +set_option trace.grind.ematch.pattern true +set_option trace.grind.ematch.instance true + +example (as bs cs : Array α) (v₁ v₂ : α) + (i₁ i₂ j : Nat) + (h₁ : i₁ < as.size) + (h₂ : bs = as.set i₁ v₁) + (h₃ : i₂ < bs.size) + (h₃ : cs = bs.set i₂ v₂) + (h₄ : i₁ ≠ j ∧ i₂ ≠ j) + (h₅ : j < cs.size) + (h₆ : j < as.size) + : cs[j] = as[j] := by + grind From e46b5f39bf77dffed3af8bfe154c73c2b772107f Mon Sep 17 00:00:00 2001 From: Leonardo de Moura Date: Thu, 2 Jan 2025 20:08:02 +0100 Subject: [PATCH 006/100] fix: congruence closure in the `grind` tactic (#6509) This PR fixes a bug in the congruence closure data structure used in the `grind` tactic. The new test includes an example that previously caused a panic. A similar panic was also occurring in the test `grind_nested_proofs.lean`. --- src/Lean/Meta/Tactic/Grind/Core.lean | 10 ++++++---- src/Lean/Meta/Tactic/Grind/EMatch.lean | 6 +++--- src/Lean/Meta/Tactic/Grind/Internalize.lean | 2 +- src/Lean/Meta/Tactic/Grind/Inv.lean | 4 ++-- src/Lean/Meta/Tactic/Grind/Types.lean | 21 +++++++++++++++++---- tests/lean/run/grind_ematch2.lean | 12 ++++++++++++ tests/lean/run/grind_nested_proofs.lean | 5 ++--- 7 files changed, 43 insertions(+), 17 deletions(-) diff --git a/src/Lean/Meta/Tactic/Grind/Core.lean b/src/Lean/Meta/Tactic/Grind/Core.lean index 0a80f7e75338..747e9eb9875f 100644 --- a/src/Lean/Meta/Tactic/Grind/Core.lean +++ b/src/Lean/Meta/Tactic/Grind/Core.lean @@ -41,8 +41,9 @@ This is an auxiliary function performed while merging equivalence classes. private def removeParents (root : Expr) : GoalM ParentSet := do let parents ← getParentsAndReset root for parent in parents do - trace[grind.debug.parent] "remove: {parent}" - modify fun s => { s with congrTable := s.congrTable.erase { e := parent } } + if (← isCongrRoot parent) then + trace[grind.debug.parent] "remove: {parent}" + modify fun s => { s with congrTable := s.congrTable.erase { e := parent } } return parents /-- @@ -51,8 +52,9 @@ This is an auxiliary function performed while merging equivalence classes. -/ private def reinsertParents (parents : ParentSet) : GoalM Unit := do for parent in parents do - trace[grind.debug.parent] "reinsert: {parent}" - addCongrTable parent + if (← isCongrRoot parent) then + trace[grind.debug.parent] "reinsert: {parent}" + addCongrTable parent /-- Closes the goal when `True` and `False` are in the same equivalence class. -/ private def closeGoalWithTrueEqFalse : GoalM Unit := do diff --git a/src/Lean/Meta/Tactic/Grind/EMatch.lean b/src/Lean/Meta/Tactic/Grind/EMatch.lean index 762048e1202c..1d19b860ba05 100644 --- a/src/Lean/Meta/Tactic/Grind/EMatch.lean +++ b/src/Lean/Meta/Tactic/Grind/EMatch.lean @@ -122,7 +122,7 @@ private partial def processMatch (c : Choice) (p : Expr) (e : Expr) : M Unit := let n ← getENode curr if n.generation <= maxGeneration -- uses heterogeneous equality or is the root of its congruence class - && (n.heqProofs || isSameExpr curr n.cgRoot) + && (n.heqProofs || n.isCongrRoot) && eqvFunctions pFn curr.getAppFn && curr.getAppNumArgs == numArgs then if let some c ← matchArgs? c p curr |>.run then @@ -140,7 +140,7 @@ private def processContinue (c : Choice) (p : Expr) : M Unit := do for app in apps do let n ← getENode app if n.generation <= maxGeneration - && (n.heqProofs || isSameExpr n.cgRoot app) then + && (n.heqProofs || n.isCongrRoot) then if let some c ← matchArgs? c p app |>.run then let gen := n.generation let c := { c with gen := Nat.max gen c.gen } @@ -225,7 +225,7 @@ private def main (p : Expr) (cnstrs : List Cnstr) : M Unit := do for app in apps do if (← checkMaxInstancesExceeded) then return () let n ← getENode app - if (n.heqProofs || isSameExpr n.cgRoot app) && + if (n.heqProofs || n.isCongrRoot) && (!useMT || n.mt == gmt) then if let some c ← matchArgs? { cnstrs, assignment, gen := n.generation } p app |>.run then modify fun s => { s with choiceStack := [c] } diff --git a/src/Lean/Meta/Tactic/Grind/Internalize.lean b/src/Lean/Meta/Tactic/Grind/Internalize.lean index fb4986404eb9..6231923227bf 100644 --- a/src/Lean/Meta/Tactic/Grind/Internalize.lean +++ b/src/Lean/Meta/Tactic/Grind/Internalize.lean @@ -25,7 +25,7 @@ def addCongrTable (e : Expr) : GoalM Unit := do trace[grind.debug.congr] "{e} = {e'}" pushEqHEq e e' congrPlaceholderProof let node ← getENode e - setENode e { node with cgRoot := e' } + setENode e { node with congr := e' } else modify fun s => { s with congrTable := s.congrTable.insert { e } } diff --git a/src/Lean/Meta/Tactic/Grind/Inv.lean b/src/Lean/Meta/Tactic/Grind/Inv.lean index b0e77d7a3069..c09ad807e7d8 100644 --- a/src/Lean/Meta/Tactic/Grind/Inv.lean +++ b/src/Lean/Meta/Tactic/Grind/Inv.lean @@ -24,9 +24,9 @@ private def checkEqc (root : ENode) : GoalM Unit := do if curr.isApp then if let some { e } := (← get).congrTable.find? { e := curr } then if (← hasSameType e.getAppFn curr.getAppFn) then - assert! isSameExpr e (← getENode curr).cgRoot + assert! isSameExpr e (← getCongrRoot curr) else - assert! isSameExpr curr (← getENode curr).cgRoot + assert! (← isCongrRoot curr) -- If the equivalence class does not have HEq proofs, then the types must be definitionally equal. unless root.heqProofs do assert! (← hasSameType curr root.self) diff --git a/src/Lean/Meta/Tactic/Grind/Types.lean b/src/Lean/Meta/Tactic/Grind/Types.lean index 2f62f59b72fc..1065c7b4d344 100644 --- a/src/Lean/Meta/Tactic/Grind/Types.lean +++ b/src/Lean/Meta/Tactic/Grind/Types.lean @@ -173,8 +173,12 @@ structure ENode where next : Expr /-- Root (aka canonical representative) of the equivalence class -/ root : Expr - /-- Root of the congruence class. This is field is a don't care if `e` is not an application. -/ - cgRoot : Expr + /-- + `congr` is the term `self` is congruent to. + We say `self` is the congruence class root if `isSameExpr congr self`. + This field is initialized to `self` even if `e` is not an application. + -/ + congr : Expr /-- When `e` was added to this equivalence class because of an equality `h : e = target`, then we store `target` here, and `h` at `proof?`. @@ -206,6 +210,9 @@ structure ENode where -- TODO: see Lean 3 implementation deriving Inhabited, Repr +def ENode.isCongrRoot (n : ENode) := + isSameExpr n.self n.congr + /-- New equality to be processed. -/ structure NewEq where lhs : Expr @@ -540,7 +547,7 @@ def setENode (e : Expr) (n : ENode) : GoalM Unit := def mkENodeCore (e : Expr) (interpreted ctor : Bool) (generation : Nat) : GoalM Unit := do setENode e { - self := e, next := e, root := e, cgRoot := e, size := 1 + self := e, next := e, root := e, congr := e, size := 1 flipped := false heqProofs := false hasLambdas := e.isLambda @@ -562,7 +569,13 @@ def mkENode (e : Expr) (generation : Nat) : GoalM Unit := do /-- Returns `true` is `e` is the root of its congruence class. -/ def isCongrRoot (e : Expr) : GoalM Bool := do - return isSameExpr e (← getENode e).cgRoot + return (← getENode e).isCongrRoot + +/-- Returns the root of the congruence class containing `e`. -/ +partial def getCongrRoot (e : Expr) : GoalM Expr := do + let n ← getENode e + if isSameExpr n.congr e then return e + getCongrRoot n.congr /-- Return `true` if the goal is inconsistent. -/ def isInconsistent : GoalM Bool := diff --git a/tests/lean/run/grind_ematch2.lean b/tests/lean/run/grind_ematch2.lean index 08d2238d17c8..2d4f1c850b98 100644 --- a/tests/lean/run/grind_ematch2.lean +++ b/tests/lean/run/grind_ematch2.lean @@ -17,3 +17,15 @@ example (as bs cs : Array α) (v₁ v₂ : α) (h₆ : j < as.size) : cs[j] = as[j] := by grind + +example (as bs cs : Array α) (v₁ v₂ : α) + (i₁ i₂ j : Nat) + (h₁ : i₁ < as.size) + (h₂ : as.set i₁ v₁ = bs) + (h₃ : i₂ < bs.size) + (h₃ : bs.set i₂ v₂ = cs) + (h₄ : i₁ ≠ j ∧ j ≠ i₂) + (h₅ : j < cs.size) + (h₆ : j < as.size) + : cs[j] = as[j] := by + grind diff --git a/tests/lean/run/grind_nested_proofs.lean b/tests/lean/run/grind_nested_proofs.lean index 12bc224f99f9..bee9020b60e5 100644 --- a/tests/lean/run/grind_nested_proofs.lean +++ b/tests/lean/run/grind_nested_proofs.lean @@ -12,9 +12,8 @@ elab "grind_test" : tactic => withMainContext do let [_, n, _] := nodes.toList | unreachable! logInfo (← getEqc n.self) --- TODO: fix --- set_option grind.debug true --- set_option grind.debug.proofs true +set_option grind.debug true +set_option grind.debug.proofs true /- Recall that array access terms, such as `a[i]`, have nested proofs. From 9d622270a16070d4a7a21f058c1a9ad91dce65e2 Mon Sep 17 00:00:00 2001 From: Leonardo de Moura Date: Thu, 2 Jan 2025 23:08:19 +0100 Subject: [PATCH 007/100] feat: custom congruence rule for equality in `grind` (#6510) This PR adds a custom congruence rule for equality in `grind`. The new rule takes into account that `Eq` is a symmetric relation. In the future, we will add support for arbitrary symmetric relations. The current rule is important for propagating disequalities effectively in `grind`. --- src/Init/Grind/Lemmas.lean | 3 ++ src/Lean/Meta/Tactic/Grind.lean | 1 + src/Lean/Meta/Tactic/Grind/Proof.lean | 78 +++++++++++++++++---------- src/Lean/Meta/Tactic/Grind/Types.lean | 34 ++++++++---- tests/lean/run/grind_diseq.lean | 5 ++ tests/lean/run/grind_ematch2.lean | 12 +++++ 6 files changed, 95 insertions(+), 38 deletions(-) create mode 100644 tests/lean/run/grind_diseq.lean diff --git a/src/Init/Grind/Lemmas.lean b/src/Init/Grind/Lemmas.lean index 4372e8c91848..d76e243ecf8d 100644 --- a/src/Init/Grind/Lemmas.lean +++ b/src/Init/Grind/Lemmas.lean @@ -51,6 +51,9 @@ theorem false_of_not_eq_self {a : Prop} (h : (Not a) = a) : False := by theorem eq_eq_of_eq_true_left {a b : Prop} (h : a = True) : (a = b) = b := by simp [h] theorem eq_eq_of_eq_true_right {a b : Prop} (h : b = True) : (a = b) = a := by simp [h] +theorem eq_congr {α : Sort u} {a₁ b₁ a₂ b₂ : α} (h₁ : a₁ = a₂) (h₂ : b₁ = b₂) : (a₁ = b₁) = (a₂ = b₂) := by simp [*] +theorem eq_congr' {α : Sort u} {a₁ b₁ a₂ b₂ : α} (h₁ : a₁ = b₂) (h₂ : b₁ = a₂) : (a₁ = b₁) = (a₂ = b₂) := by rw [h₁, h₂, Eq.comm (a := a₂)] + /-! Forall -/ theorem forall_propagator (p : Prop) (q : p → Prop) (q' : Prop) (h₁ : p = True) (h₂ : q (of_eq_true h₁) = q') : (∀ hp : p, q hp) = q' := by diff --git a/src/Lean/Meta/Tactic/Grind.lean b/src/Lean/Meta/Tactic/Grind.lean index 938341c3ec90..11b15243e263 100644 --- a/src/Lean/Meta/Tactic/Grind.lean +++ b/src/Lean/Meta/Tactic/Grind.lean @@ -46,5 +46,6 @@ builtin_initialize registerTraceClass `grind.debug.congr builtin_initialize registerTraceClass `grind.debug.proof builtin_initialize registerTraceClass `grind.debug.proj builtin_initialize registerTraceClass `grind.debug.parent +builtin_initialize registerTraceClass `grind.debug.final end Lean diff --git a/src/Lean/Meta/Tactic/Grind/Proof.lean b/src/Lean/Meta/Tactic/Grind/Proof.lean index 580cf2362dde..f395c3fe89dc 100644 --- a/src/Lean/Meta/Tactic/Grind/Proof.lean +++ b/src/Lean/Meta/Tactic/Grind/Proof.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Leonardo de Moura -/ prelude -import Lean.Meta.Sorry -- TODO: remove +import Init.Grind.Lemmas import Lean.Meta.Tactic.Grind.Types namespace Lean.Meta.Grind @@ -128,6 +128,52 @@ mutual let r := (← loop lhs rhs).get! if heq then mkHEqOfEq r else return r + private partial def mkHCongrProof (lhs rhs : Expr) (heq : Bool) : GoalM Expr := do + let f := lhs.getAppFn + let g := rhs.getAppFn + let numArgs := lhs.getAppNumArgs + assert! rhs.getAppNumArgs == numArgs + let thm ← mkHCongrWithArity f numArgs + assert! thm.argKinds.size == numArgs + let rec loop (lhs rhs : Expr) (i : Nat) : GoalM Expr := do + let i := i - 1 + if lhs.isApp then + let proof ← loop lhs.appFn! rhs.appFn! i + let a₁ := lhs.appArg! + let a₂ := rhs.appArg! + let k := thm.argKinds[i]! + return mkApp3 proof a₁ a₂ (← mkEqProofCore a₁ a₂ (k matches .heq)) + else + return thm.proof + let proof ← loop lhs rhs numArgs + if isSameExpr f g then + mkEqOfHEqIfNeeded proof heq + else + /- + `lhs` is of the form `f a_1 ... a_n` + `rhs` is of the form `g b_1 ... b_n` + `proof : HEq (f a_1 ... a_n) (f b_1 ... b_n)` + We construct a proof for `HEq (f a_1 ... a_n) (g b_1 ... b_n)` using `Eq.ndrec` + -/ + let motive ← withLocalDeclD (← mkFreshUserName `x) (← inferType f) fun x => do + mkLambdaFVars #[x] (← mkHEq lhs (mkAppN x rhs.getAppArgs)) + let fEq ← mkEqProofCore f g false + let proof ← mkEqNDRec motive proof fEq + mkEqOfHEqIfNeeded proof heq + + private partial def mkEqCongrProof (lhs rhs : Expr) (heq : Bool) : GoalM Expr := do + let_expr f@Eq α₁ a₁ b₁ := lhs | unreachable! + let_expr Eq α₂ a₂ b₂ := rhs | unreachable! + let enodes := (← get).enodes + let us := f.constLevels! + if !isSameExpr α₁ α₂ then + mkHCongrProof lhs rhs heq + else if hasSameRoot enodes a₁ a₂ && hasSameRoot enodes b₁ b₂ then + return mkApp7 (mkConst ``Grind.eq_congr us) α₁ a₁ b₁ a₂ b₂ (← mkEqProofCore a₁ a₂ false) (← mkEqProofCore b₁ b₂ false) + else + assert! hasSameRoot enodes a₁ b₂ && hasSameRoot enodes b₁ a₂ + return mkApp7 (mkConst ``Grind.eq_congr' us) α₁ a₁ b₁ a₂ b₂ (← mkEqProofCore a₁ b₂ false) (← mkEqProofCore b₁ a₂ false) + /-- Constructs a congruence proof for `lhs` and `rhs`. -/ private partial def mkCongrProof (lhs rhs : Expr) (heq : Bool) : GoalM Expr := do let f := lhs.getAppFn @@ -136,36 +182,12 @@ mutual assert! rhs.getAppNumArgs == numArgs if f.isConstOf ``Lean.Grind.nestedProof && g.isConstOf ``Lean.Grind.nestedProof && numArgs == 2 then mkNestedProofCongr lhs rhs heq + else if f.isConstOf ``Eq && g.isConstOf ``Eq && numArgs == 3 then + mkEqCongrProof lhs rhs heq else if (← isCongrDefaultProofTarget lhs rhs f g numArgs) then mkCongrDefaultProof lhs rhs heq else - let thm ← mkHCongrWithArity f numArgs - assert! thm.argKinds.size == numArgs - let rec loop (lhs rhs : Expr) (i : Nat) : GoalM Expr := do - let i := i - 1 - if lhs.isApp then - let proof ← loop lhs.appFn! rhs.appFn! i - let a₁ := lhs.appArg! - let a₂ := rhs.appArg! - let k := thm.argKinds[i]! - return mkApp3 proof a₁ a₂ (← mkEqProofCore a₁ a₂ (k matches .heq)) - else - return thm.proof - let proof ← loop lhs rhs numArgs - if isSameExpr f g then - mkEqOfHEqIfNeeded proof heq - else - /- - `lhs` is of the form `f a_1 ... a_n` - `rhs` is of the form `g b_1 ... b_n` - `proof : HEq (f a_1 ... a_n) (f b_1 ... b_n)` - We construct a proof for `HEq (f a_1 ... a_n) (g b_1 ... b_n)` using `Eq.ndrec` - -/ - let motive ← withLocalDeclD (← mkFreshUserName `x) (← inferType f) fun x => do - mkLambdaFVars #[x] (← mkHEq lhs (mkAppN x rhs.getAppArgs)) - let fEq ← mkEqProofCore f g false - let proof ← mkEqNDRec motive proof fEq - mkEqOfHEqIfNeeded proof heq + mkHCongrProof lhs rhs heq private partial def realizeEqProof (lhs rhs : Expr) (h : Expr) (flipped : Bool) (heq : Bool) : GoalM Expr := do let h ← if h == congrPlaceholderProof then diff --git a/src/Lean/Meta/Tactic/Grind/Types.lean b/src/Lean/Meta/Tactic/Grind/Types.lean index 1065c7b4d344..0a3329915afd 100644 --- a/src/Lean/Meta/Tactic/Grind/Types.lean +++ b/src/Lean/Meta/Tactic/Grind/Types.lean @@ -249,7 +249,7 @@ private def hashRoot (enodes : ENodeMap) (e : Expr) : UInt64 := else 13 -private def hasSameRoot (enodes : ENodeMap) (a b : Expr) : Bool := Id.run do +def hasSameRoot (enodes : ENodeMap) (a b : Expr) : Bool := Id.run do if isSameExpr a b then return true else @@ -258,12 +258,15 @@ private def hasSameRoot (enodes : ENodeMap) (a b : Expr) : Bool := Id.run do isSameExpr n1.root n2.root def congrHash (enodes : ENodeMap) (e : Expr) : UInt64 := - if e.isAppOfArity ``Lean.Grind.nestedProof 2 then - -- We only hash the proposition - hashRoot enodes (e.getArg! 0) - else - go e 17 + match_expr e with + | Grind.nestedProof p _ => hashRoot enodes p + | Eq _ lhs rhs => goEq lhs rhs + | _ => go e 17 where + goEq (lhs rhs : Expr) : UInt64 := + let h₁ := hashRoot enodes lhs + let h₂ := hashRoot enodes rhs + if h₁ > h₂ then mixHash h₂ h₁ else mixHash h₁ h₂ go (e : Expr) (r : UInt64) : UInt64 := match e with | .app f a => go f (mixHash r (hashRoot enodes a)) @@ -271,11 +274,22 @@ where /-- Returns `true` if `a` and `b` are congruent modulo the equivalence classes in `enodes`. -/ partial def isCongruent (enodes : ENodeMap) (a b : Expr) : Bool := - if a.isAppOfArity ``Lean.Grind.nestedProof 2 && b.isAppOfArity ``Lean.Grind.nestedProof 2 then - hasSameRoot enodes (a.getArg! 0) (b.getArg! 0) - else - go a b + match_expr a with + | Grind.nestedProof p₁ _ => + let_expr Grind.nestedProof p₂ _ := b | false + hasSameRoot enodes p₁ p₂ + | Eq α₁ lhs₁ rhs₁ => + let_expr Eq α₂ lhs₂ rhs₂ := b | false + if isSameExpr α₁ α₂ then + goEq lhs₁ rhs₁ lhs₂ rhs₂ + else + go a b + | _ => go a b where + goEq (lhs₁ rhs₁ lhs₂ rhs₂ : Expr) : Bool := + (hasSameRoot enodes lhs₁ lhs₂ && hasSameRoot enodes rhs₁ rhs₂) + || + (hasSameRoot enodes lhs₁ rhs₂ && hasSameRoot enodes rhs₁ lhs₂) go (a b : Expr) : Bool := if a.isApp && b.isApp then hasSameRoot enodes a.appArg! b.appArg! && go a.appFn! b.appFn! diff --git a/tests/lean/run/grind_diseq.lean b/tests/lean/run/grind_diseq.lean new file mode 100644 index 000000000000..724272154ede --- /dev/null +++ b/tests/lean/run/grind_diseq.lean @@ -0,0 +1,5 @@ +set_option grind.debug true + +example (p q : Prop) (a b c d : Nat) : + a = b → c = d → a ≠ c → (d ≠ b → p) → (d ≠ b → q) → p ∧ q := by + grind (splits:=0) diff --git a/tests/lean/run/grind_ematch2.lean b/tests/lean/run/grind_ematch2.lean index 2d4f1c850b98..06b6256ba1de 100644 --- a/tests/lean/run/grind_ematch2.lean +++ b/tests/lean/run/grind_ematch2.lean @@ -29,3 +29,15 @@ example (as bs cs : Array α) (v₁ v₂ : α) (h₆ : j < as.size) : cs[j] = as[j] := by grind + +example (as bs cs : Array α) (v₁ v₂ : α) + (i₁ i₂ j : Nat) + (h₁ : i₁ < as.size) + (h₂ : as.set i₁ v₁ = bs) + (h₃ : i₂ < bs.size) + (h₃ : bs.set i₂ v₂ = cs) + (h₄ : j ≠ i₁ ∧ j ≠ i₂) + (h₅ : j < cs.size) + (h₆ : j < as.size) + : cs[j] = as[j] := by + grind From 3e2f1faebf4b525a5e0be7186b9265f418602d38 Mon Sep 17 00:00:00 2001 From: Leonardo de Moura Date: Fri, 3 Jan 2025 00:56:51 +0100 Subject: [PATCH 008/100] feat: add user-defined fallback procedure for the `grind` tactic (#6512) This PR introduces support for user-defined fallback code in the `grind` tactic. The fallback code can be utilized to inspect the state of failing `grind` subgoals and/or invoke user-defined automation. Users can now write `grind on_failure `, where `` should have the type `GoalM Unit`. See the modified tests in this PR for examples. --- src/Init/Grind/Tactics.lean | 2 +- src/Lean/Elab/Tactic/Grind.lean | 26 ++++- src/Lean/Meta/Tactic/Grind/Main.lean | 28 +++--- src/Lean/Meta/Tactic/Grind/Types.lean | 10 +- tests/lean/run/grind_canon_insts.lean | 29 +++--- tests/lean/run/grind_canon_types.lean | 21 ++-- tests/lean/run/grind_congr.lean | 41 +++----- tests/lean/run/grind_ematch2.lean | 25 +++++ tests/lean/run/grind_many_eqs.lean | 30 +++--- tests/lean/run/grind_nested_proofs.lean | 33 +++---- tests/lean/run/grind_norm_levels.lean | 20 ++-- .../lean/run/grind_propagate_connectives.lean | 96 +++++++------------ 12 files changed, 170 insertions(+), 191 deletions(-) diff --git a/src/Init/Grind/Tactics.lean b/src/Init/Grind/Tactics.lean index bdc6bea97107..1554e8b0a466 100644 --- a/src/Init/Grind/Tactics.lean +++ b/src/Init/Grind/Tactics.lean @@ -39,6 +39,6 @@ namespace Lean.Parser.Tactic -/ -- TODO: parameters -syntax (name := grind) "grind" optConfig : tactic +syntax (name := grind) "grind" optConfig ("on_failure " term)? : tactic end Lean.Parser.Tactic diff --git a/src/Lean/Elab/Tactic/Grind.lean b/src/Lean/Elab/Tactic/Grind.lean index 774c6c055185..b9a62e4baed7 100644 --- a/src/Lean/Elab/Tactic/Grind.lean +++ b/src/Lean/Elab/Tactic/Grind.lean @@ -32,18 +32,36 @@ def elabGrindPattern : CommandElab := fun stx => do Grind.addEMatchTheorem declName xs.size patterns.toList | _ => throwUnsupportedSyntax -def grind (mvarId : MVarId) (config : Grind.Config) (mainDeclName : Name) : MetaM Unit := do - let mvarIds ← Grind.main mvarId config mainDeclName +def grind (mvarId : MVarId) (config : Grind.Config) (mainDeclName : Name) (fallback : Grind.Fallback) : MetaM Unit := do + let mvarIds ← Grind.main mvarId config mainDeclName fallback unless mvarIds.isEmpty do throwError "`grind` failed\n{goalsToMessageData mvarIds}" +private def elabFallback (fallback? : Option Term) : TermElabM (Grind.GoalM Unit) := do + let some fallback := fallback? | return (pure ()) + let type := mkApp (mkConst ``Grind.GoalM) (mkConst ``Unit) + let value ← withLCtx {} {} do Term.elabTermAndSynthesize fallback type + let auxDeclName ← if let .const declName _ := value then + pure declName + else + let auxDeclName ← Term.mkAuxName `_grind_fallback + let decl := Declaration.defnDecl { + name := auxDeclName + levelParams := [] + type, value, hints := .opaque, safety := .safe + } + addAndCompile decl + pure auxDeclName + unsafe evalConst (Grind.GoalM Unit) auxDeclName + @[builtin_tactic Lean.Parser.Tactic.grind] def evalApplyRfl : Tactic := fun stx => do match stx with - | `(tactic| grind $config:optConfig) => + | `(tactic| grind $config:optConfig $[on_failure $fallback?]?) => + let fallback ← elabFallback fallback? logWarningAt stx "The `grind` tactic is experimental and still under development. Avoid using it in production projects" let declName := (← Term.getDeclName?).getD `_grind let config ← elabGrindConfig config - withMainContext do liftMetaFinishingTactic (grind · config declName) + withMainContext do liftMetaFinishingTactic (grind · config declName fallback) | _ => throwUnsupportedSyntax end Lean.Elab.Tactic diff --git a/src/Lean/Meta/Tactic/Grind/Main.lean b/src/Lean/Meta/Tactic/Grind/Main.lean index da2f202c8fd4..77a7b4ffa7bf 100644 --- a/src/Lean/Meta/Tactic/Grind/Main.lean +++ b/src/Lean/Meta/Tactic/Grind/Main.lean @@ -17,9 +17,10 @@ import Lean.Meta.Tactic.Grind.EMatch namespace Lean.Meta.Grind -def mkMethods : CoreM Methods := do +def mkMethods (fallback : Fallback) : CoreM Methods := do let builtinPropagators ← builtinPropagatorsRef.get return { + fallback propagateUp := fun e => do propagateForallProp e let .const declName _ := e.getAppFn | return () @@ -32,7 +33,7 @@ def mkMethods : CoreM Methods := do prop e } -def GrindM.run (x : GrindM α) (mainDeclName : Name) (config : Grind.Config) : MetaM α := do +def GrindM.run (x : GrindM α) (mainDeclName : Name) (config : Grind.Config) (fallback : Fallback) : MetaM α := do let scState := ShareCommon.State.mk _ let (falseExpr, scState) := ShareCommon.State.shareCommon scState (mkConst ``False) let (trueExpr, scState) := ShareCommon.State.shareCommon scState (mkConst ``True) @@ -42,7 +43,7 @@ def GrindM.run (x : GrindM α) (mainDeclName : Name) (config : Grind.Config) : M (config := { arith := true }) (simpTheorems := #[thms]) (congrTheorems := (← getSimpCongrTheorems)) - x (← mkMethods).toMethodsRef { mainDeclName, config, simprocs, simp } |>.run' { scState, trueExpr, falseExpr } + x (← mkMethods fallback).toMethodsRef { mainDeclName, config, simprocs, simp } |>.run' { scState, trueExpr, falseExpr } private def mkGoal (mvarId : MVarId) : GrindM Goal := do let trueExpr ← getTrueExpr @@ -71,23 +72,18 @@ def all (goals : List Goal) (f : Goal → GrindM (List Goal)) : GrindM (List Goa private def simple (goals : List Goal) : GrindM (List Goal) := do all goals ematchStar -def main (mvarId : MVarId) (config : Grind.Config) (mainDeclName : Name) : MetaM (List MVarId) := do +def main (mvarId : MVarId) (config : Grind.Config) (mainDeclName : Name) (fallback : Fallback) : MetaM (List MVarId) := do let go : GrindM (List MVarId) := do let goals ← initCore mvarId let goals ← simple goals + let goals ← goals.filterMapM fun goal => do + if goal.inconsistent then return none + let goal ← GoalM.run' goal fallback + if goal.inconsistent then return none + if (← goal.mvarId.isAssigned) then return none + return some goal trace[grind.debug.final] "{← ppGoals goals}" return goals.map (·.mvarId) - go.run mainDeclName config - -/-- Helper function for debugging purposes -/ -def preprocessAndProbe (mvarId : MVarId) (mainDeclName : Name) (p : GoalM Unit) : MetaM Unit := - let go : GrindM Unit := do - let goals ← initCore mvarId - trace[grind.debug.final] "{← ppGoals goals}" - goals.forM fun goal => - discard <| GoalM.run' goal p - return () - withoutModifyingMCtx do - go.run mainDeclName {} + go.run mainDeclName config fallback end Lean.Meta.Grind diff --git a/src/Lean/Meta/Tactic/Grind/Types.lean b/src/Lean/Meta/Tactic/Grind/Types.lean index 0a3329915afd..5ac73e6c4dff 100644 --- a/src/Lean/Meta/Tactic/Grind/Types.lean +++ b/src/Lean/Meta/Tactic/Grind/Types.lean @@ -392,8 +392,6 @@ abbrev GoalM := StateRefT Goal GrindM @[inline] def GoalM.run' (goal : Goal) (x : GoalM Unit) : GrindM Goal := goal.mvarId.withContext do StateRefT'.run' (x *> get) goal -abbrev Propagator := Expr → GoalM Unit - /-- A helper function used to mark a theorem instance found by the E-matching module. It returns `true` if it is a new instance and `false` otherwise. @@ -677,9 +675,13 @@ def forEachEqc (f : ENode → GoalM Unit) : GoalM Unit := do if isSameExpr n.self n.root then f n +abbrev Propagator := Expr → GoalM Unit +abbrev Fallback := GoalM Unit + structure Methods where propagateUp : Propagator := fun _ => return () propagateDown : Propagator := fun _ => return () + fallback : Fallback := pure () deriving Inhabited def Methods.toMethodsRef (m : Methods) : MethodsRef := @@ -697,6 +699,10 @@ def propagateUp (e : Expr) : GoalM Unit := do def propagateDown (e : Expr) : GoalM Unit := do (← getMethods).propagateDown e +def applyFallback : GoalM Unit := do + let fallback : GoalM Unit := (← getMethods).fallback + fallback + /-- Returns expressions in the given expression equivalence class. -/ partial def getEqc (e : Expr) : GoalM (List Expr) := go e e [] diff --git a/tests/lean/run/grind_canon_insts.lean b/tests/lean/run/grind_canon_insts.lean index ddbadb62f5e9..c38feabfc8fa 100644 --- a/tests/lean/run/grind_canon_insts.lean +++ b/tests/lean/run/grind_canon_insts.lean @@ -1,11 +1,4 @@ -import Lean - -open Lean Meta Elab Tactic Grind in -elab "grind_test" : tactic => withMainContext do - let declName := (← Term.getDeclName?).getD `_main - Meta.Grind.preprocessAndProbe (← getMainGoal) declName do - let nodes ← filterENodes fun e => return e.self.isAppOf ``HMul.hMul - logInfo (nodes.toList.map (·.self)) +import Lean.Meta.Tactic.Grind set_option grind.debug true @@ -57,26 +50,27 @@ instance : CommMonoid Nat where theorem left_comm [CommMonoid α] (a b c : α) : a * (b * c) = b * (a * c) := by rw [← Semigroup.mul_assoc, CommMonoid.mul_comm a b, Semigroup.mul_assoc] +open Lean Meta Elab Tactic Grind in +def fallback : Fallback := do + let nodes ← filterENodes fun e => return e.self.isAppOf ``HMul.hMul + logInfo (nodes.toList.map (·.self)) + (← get).mvarId.admit + /-- info: [b * c, a * (b * c), d * (b * c)] ---- -warning: declaration uses 'sorry' -/ -#guard_msgs in +#guard_msgs (info) in example (a b c d : Nat) : b * (a * c) = d * (b * c) → False := by rw [left_comm] -- Introduces a new (non-canonical) instance for `Mul Nat` - grind_test -- State should have only 3 `*`-applications - sorry + grind on_failure fallback -- State should have only 3 `*`-applications set_option pp.notation false in set_option pp.explicit true in /-- info: [@HMul.hMul Nat Nat Nat (@instHMul Nat instMulNat) b a, @HMul.hMul Nat Nat Nat (@instHMul Nat instMulNat) b d] ---- -warning: declaration uses 'sorry' -/ -#guard_msgs in +#guard_msgs (info) in example (a b c d : Nat) : b * a = d * b → False := by rw [CommMonoid.mul_comm d b] -- Introduces a new (non-canonical) instance for `Mul Nat` -- See target here @@ -85,5 +79,4 @@ example (a b c d : Nat) : b * a = d * b → False := by = @HMul.hMul Nat Nat Nat (@instHMul Nat (@Semigroup.toMul Nat (@Monoid.toSemigroup Nat (@CommMonoid.toMonoid Nat instCommMonoidNat)))) b d → False - grind_test -- State should have only 2 `*`-applications, and they use the same instance - sorry + grind on_failure fallback -- State should have only 2 `*`-applications, and they use the same instance diff --git a/tests/lean/run/grind_canon_types.lean b/tests/lean/run/grind_canon_types.lean index ec956e8f61c0..08a1e17276a7 100644 --- a/tests/lean/run/grind_canon_types.lean +++ b/tests/lean/run/grind_canon_types.lean @@ -1,27 +1,22 @@ -import Lean +import Lean.Meta.Tactic.Grind def g (s : Type) := s def f (a : α) := a -open Lean Meta Elab Tactic Grind in -elab "grind_test" : tactic => withMainContext do - let declName := (← Term.getDeclName?).getD `_main - Meta.Grind.preprocessAndProbe (← getMainGoal) declName do - let nodes ← filterENodes fun e => return e.self.isAppOf ``f - logInfo (nodes.toList.map (·.self)) - +open Lean Meta Grind in +def fallback : Fallback := do + let nodes ← filterENodes fun e => return e.self.isAppOf ``f + logInfo (nodes.toList.map (·.self)) + (← get).mvarId.admit set_option pp.explicit true /-- info: [@f Nat a, @f Nat b] ---- -warning: declaration uses 'sorry' -/ -#guard_msgs in +#guard_msgs (info) in example (a b c d : Nat) : @f Nat a = b → @f (g Nat) a = c → @f (g Nat) b = d → a = b → False := by -- State should have only two `f`-applications: `@f Nat a`, `@f Nat b` -- Note that `@f (g Nat) b` has been canonicalized to `@f Nat b`. -- Thus, if `a` and `b` equivalence classes are merged, `grind` can still detect that -- `@f Nat a` and `@f Nat b` are equal too. - grind_test - sorry + grind on_failure fallback diff --git a/tests/lean/run/grind_congr.lean b/tests/lean/run/grind_congr.lean index f981018a26a3..1589785ef00f 100644 --- a/tests/lean/run/grind_congr.lean +++ b/tests/lean/run/grind_congr.lean @@ -4,53 +4,40 @@ def f (a : Nat) := a + a + a def g (a : Nat) := a + a -- Prints the equivalence class containing a `f` application -open Lean Meta Elab Tactic Grind in -elab "grind_test" : tactic => withMainContext do - let declName := (← Term.getDeclName?).getD `_main - Meta.Grind.preprocessAndProbe (← getMainGoal) declName do - let #[n, _] ← filterENodes fun e => return e.self.isAppOf ``f | unreachable! - let eqc ← getEqc n.self - logInfo eqc +open Lean Meta Grind in +def fallback : Fallback := do + let #[n, _] ← filterENodes fun e => return e.self.isAppOf ``f | unreachable! + let eqc ← getEqc n.self + logInfo eqc + (← get).mvarId.admit set_option grind.debug true set_option grind.debug.proofs true /-- info: [d, f b, c, f a] ---- -warning: declaration uses 'sorry' -/ -#guard_msgs in +#guard_msgs (info) in example (a b c d : Nat) : a = b → f a = c → f b = d → False := by - grind_test - sorry + grind on_failure fallback /-- info: [d, f b, c, f a] ---- -warning: declaration uses 'sorry' -/ -#guard_msgs in +#guard_msgs (info) in example (a b c d : Nat) : f a = c → f b = d → a = b → False := by - grind_test - sorry + grind on_failure fallback /-- info: [d, f (g b), c, f (g a)] ---- -warning: declaration uses 'sorry' -/ -#guard_msgs in +#guard_msgs (info) in example (a b c d e : Nat) : f (g a) = c → f (g b) = d → a = e → b = e → False := by - grind_test - sorry + grind on_failure fallback /-- info: [d, f (g b), c, f v] ---- -warning: declaration uses 'sorry' -/ -#guard_msgs in +#guard_msgs (info) in example (a b c d e v : Nat) : f v = c → f (g b) = d → a = e → b = e → v = g a → False := by - grind_test - sorry + grind on_failure fallback diff --git a/tests/lean/run/grind_ematch2.lean b/tests/lean/run/grind_ematch2.lean index 06b6256ba1de..613972c80e11 100644 --- a/tests/lean/run/grind_ematch2.lean +++ b/tests/lean/run/grind_ematch2.lean @@ -41,3 +41,28 @@ example (as bs cs : Array α) (v₁ v₂ : α) (h₆ : j < as.size) : cs[j] = as[j] := by grind + +example (as bs cs ds : Array α) (v₁ v₂ v₃ : α) + (i₁ i₂ i₃ j : Nat) + (h₁ : i₁ < as.size) + (h₂ : as.set i₁ v₁ = bs) + (h₃ : i₂ < bs.size) + (h₃ : bs.set i₂ v₂ = cs) + (h₄ : i₃ < cs.size) + (h₅ : ds = cs.set i₃ v₃) + (h₆ : j ≠ i₁ ∧ j ≠ i₂ ∧ i₃ ≠ j) + (h₇ : j < ds.size) + (h₈ : j < as.size) + : ds[j] = as[j] := by + grind + +opaque f (a b : α) : α := a +theorem fx : f x (f x x) = x := sorry +grind_pattern fx => f x (f x x) + +/-- +info: [grind.ematch.instance] fx: f a (f a a) = a +-/ +#guard_msgs (info) in +example : a = b₁ → c = f b₁ b₂ → f a c ≠ a → a = b₂ → False := by + grind diff --git a/tests/lean/run/grind_many_eqs.lean b/tests/lean/run/grind_many_eqs.lean index 5f9825eadab2..25c254917c1b 100644 --- a/tests/lean/run/grind_many_eqs.lean +++ b/tests/lean/run/grind_many_eqs.lean @@ -1,4 +1,4 @@ -import Lean +import Lean.Meta.Tactic.Grind def f (a : Nat) := a + a + a def g (a : Nat) := a + a @@ -8,27 +8,23 @@ def h (n : Nat) : Prop := | n+1 => f (n+1) = f n ∧ g (2*n + 1) = g (2*n) ∧ h n -- Prints the equivalence class containing a `f` application -open Lean Meta Elab Tactic Grind in -elab "grind_test" n:num : tactic => withMainContext do - let n := n.getNat - let declName := (← Term.getDeclName?).getD `_main - Meta.Grind.preprocessAndProbe (← getMainGoal) declName do - let f0 ← Grind.shareCommon (mkApp (mkConst ``f) (mkNatLit 0)) - -- The `f 0` equivalence class contains `n+1` elements - assert! (← getEqc f0).length == n + 1 - forEachENode fun node => do - if node.self.isAppOf ``g then - -- Any equivalence class containing a `g`-application contains 2 elements - assert! (← getEqc (← getRoot node.self)).length == 2 +open Lean Meta Grind in +def fallback (n : Nat) : Fallback := do + let f0 ← Grind.shareCommon (mkApp (mkConst ``f) (mkNatLit 0)) + -- The `f 0` equivalence class contains `n+1` elements + assert! (← getEqc f0).length == n + 1 + forEachENode fun node => do + if node.self.isAppOf ``g then + -- Any equivalence class containing a `g`-application contains 2 elements + assert! (← getEqc (← getRoot node.self)).length == 2 + (← get).mvarId.admit set_option grind.debug true in example : h 5 → False := by simp [h] - grind_test 5 - sorry + grind on_failure fallback 5 set_option maxRecDepth 2048 example : h 100 → False := by simp [h] - grind_test 100 - sorry + grind on_failure fallback 100 diff --git a/tests/lean/run/grind_nested_proofs.lean b/tests/lean/run/grind_nested_proofs.lean index bee9020b60e5..04b32f7fe742 100644 --- a/tests/lean/run/grind_nested_proofs.lean +++ b/tests/lean/run/grind_nested_proofs.lean @@ -1,16 +1,15 @@ -import Lean +import Lean.Meta.Tactic.Grind def f (α : Type) [Add α] (a : α) := a + a + a -open Lean Meta Elab Tactic Grind in -elab "grind_test" : tactic => withMainContext do - let declName := (← Term.getDeclName?).getD `_main - Meta.Grind.preprocessAndProbe (← getMainGoal) declName do - let nodes ← filterENodes fun e => return e.self.isAppOf ``Lean.Grind.nestedProof - logInfo (nodes.toList.map (·.self)) - let nodes ← filterENodes fun e => return e.self.isAppOf ``GetElem.getElem - let [_, n, _] := nodes.toList | unreachable! - logInfo (← getEqc n.self) +open Lean Meta Grind in +def fallback : Fallback := do + let nodes ← filterENodes fun e => return e.self.isAppOf ``Lean.Grind.nestedProof + logInfo (nodes.toList.map (·.self)) + let nodes ← filterENodes fun e => return e.self.isAppOf ``GetElem.getElem + let [_, n, _] := nodes.toList | unreachable! + logInfo (← getEqc n.self) + (← get).mvarId.admit set_option grind.debug true set_option grind.debug.proofs true @@ -32,13 +31,8 @@ warning: declaration uses 'sorry' -/ -- #guard_msgs in -set_option trace.Meta.debug true - example (i j : Nat) (a b : Array Nat) (h1 : j < a.size) (h : j < b.size) (h2 : i ≤ j) : a[i] < a[j] + b[j] → i = j → a = b → False := by - grind_test - sorry - -#exit + grind on_failure fallback /-- info: [Lean.Grind.nestedProof (i < a.toList.length) (_example.proof_1 i j a b h1 h2), @@ -46,10 +40,7 @@ info: [Lean.Grind.nestedProof (i < a.toList.length) (_example.proof_1 i j a b h1 Lean.Grind.nestedProof (j < b.toList.length) h] --- info: [a[i], a[j]] ---- -warning: declaration uses 'sorry' -/ -#guard_msgs in +#guard_msgs (info) in example (i j : Nat) (a b : Array Nat) (h1 : j < a.size) (h : j < b.size) (h2 : i ≤ j) : a[i] < a[j] + b[j] → i = j → False := by - grind_test - sorry + grind on_failure fallback diff --git a/tests/lean/run/grind_norm_levels.lean b/tests/lean/run/grind_norm_levels.lean index aa29284c8fa3..354b56319dd2 100644 --- a/tests/lean/run/grind_norm_levels.lean +++ b/tests/lean/run/grind_norm_levels.lean @@ -1,21 +1,17 @@ -import Lean +import Lean.Meta.Tactic.Grind def g {α : Sort u} (a : α) := a -open Lean Meta Elab Tactic Grind in -elab "grind_test" : tactic => withMainContext do - let declName := (← Term.getDeclName?).getD `_main - Meta.Grind.preprocessAndProbe (← getMainGoal) declName do - let nodes ← filterENodes fun e => return e.self.isAppOf ``g - logInfo (nodes.toList.map (·.self)) +open Lean Meta Grind in +def fallback : Fallback := do + let nodes ← filterENodes fun e => return e.self.isAppOf ``g + logInfo (nodes.toList.map (·.self)) + (← get).mvarId.admit -- `grind` final state must contain only two `g`-applications /-- info: [g (a, b), g (g (a, b))] ---- -warning: declaration uses 'sorry' -/ -#guard_msgs in +#guard_msgs (info) in example {β : Type v} {α : Type u} (a c : α) (b d : β) : g.{max u v + 1} (a, b) = (c, d) → g (g.{max (u+1) (v+1)} (a, b)) = (c, d) → False := by - grind_test - sorry + grind on_failure fallback diff --git a/tests/lean/run/grind_propagate_connectives.lean b/tests/lean/run/grind_propagate_connectives.lean index 66cc3add7adf..e4f4a020dc68 100644 --- a/tests/lean/run/grind_propagate_connectives.lean +++ b/tests/lean/run/grind_propagate_connectives.lean @@ -1,18 +1,16 @@ -import Lean +import Lean.Meta.Tactic.Grind --- Prints the equivalence class containing a `f` application -open Lean Meta Elab Tactic Grind in -elab "grind_test" : tactic => withMainContext do - let declName := (← Term.getDeclName?).getD `_main - Meta.Grind.preprocessAndProbe (← getMainGoal) declName do - let trueExprs := (← filterENodes fun e => return e.self.isFVar && (← isEqTrue e.self)).toList.map (·.self) - let falseExprs := (← filterENodes fun e => return e.self.isFVar && (← isEqFalse e.self)).toList.map (·.self) - logInfo m!"true: {trueExprs}" - logInfo m!"false: {falseExprs}" - forEachEqc fun n => do - unless (← isProp n.self) || (← isType n.self) || n.size == 1 do - let eqc ← getEqc n.self - logInfo eqc +open Lean Meta Grind in +def fallback : Fallback := do + let trueExprs := (← filterENodes fun e => return e.self.isFVar && (← isEqTrue e.self)).toList.map (·.self) + let falseExprs := (← filterENodes fun e => return e.self.isFVar && (← isEqFalse e.self)).toList.map (·.self) + logInfo m!"true: {trueExprs}" + logInfo m!"false: {falseExprs}" + forEachEqc fun n => do + unless (← isProp n.self) || (← isType n.self) || n.size == 1 do + let eqc ← getEqc n.self + logInfo eqc + (← get).mvarId.admit set_option grind.debug true set_option grind.debug.proofs true @@ -21,73 +19,57 @@ set_option grind.debug.proofs true info: true: [q, w] --- info: false: [p, r] ---- -warning: declaration uses 'sorry' -/ -#guard_msgs in +#guard_msgs (info) in example : (p ∨ (q ∧ ¬r ∧ w)) → ¬p → False := by - grind_test - sorry + grind on_failure fallback + /-- info: true: [r] --- info: false: [p, q] ---- -warning: declaration uses 'sorry' -/ -#guard_msgs in +#guard_msgs (info) in example : (p ∨ q ∨ r) → (p ∨ ¬q) → ¬p → False := by - grind_test - sorry + grind on_failure fallback + /-- info: true: [r] --- info: false: [p₁, q] ---- -warning: declaration uses 'sorry' -/ -#guard_msgs in +#guard_msgs (info) in example : ((p₁ ∧ p₂) ∨ q ∨ r) → (p₁ ∨ ¬q) → p₁ = False → False := by - grind_test - sorry + grind on_failure fallback /-- info: true: [r] --- info: false: [p₂, q] ---- -warning: declaration uses 'sorry' -/ -#guard_msgs in +#guard_msgs (info) in example : ((p₁ ∧ p₂) ∨ q ∨ r) → ((p₂ ∧ p₁) ∨ ¬q) → p₂ = False → False := by - grind_test - sorry + grind on_failure fallback /-- info: true: [q, r] --- info: false: [p] ---- -warning: declaration uses 'sorry' -/ -#guard_msgs in +#guard_msgs (info) in example (p q r : Prop) : p ∨ (q ↔ r) → p = False → q → False := by - grind_test - sorry + grind on_failure fallback /-- info: true: [r] --- info: false: [p, s] ---- -warning: declaration uses 'sorry' -/ -#guard_msgs in +#guard_msgs (info) in example (p q r : Prop) : p ∨ ¬(s ∨ (p ↔ r)) → p = False → False := by - grind_test - sorry + grind on_failure fallback /-- info: true: [p] @@ -95,35 +77,29 @@ info: true: [p] info: false: [] --- info: [a, b] ---- -warning: declaration uses 'sorry' -/ -#guard_msgs in +#guard_msgs (info) in example (p : Prop) (a : Vector Nat 5) (b : Vector Nat 6) : (p → HEq a b) → p → False := by - grind_test - sorry - + grind on_failure fallback /-- info: true: [p, q] --- info: false: [r] ---- -warning: declaration uses 'sorry' -/ -#guard_msgs in +#guard_msgs (info) in example (p q r : Prop) : p ∨ (q ↔ r) → q → ¬r → False := by - grind_test - sorry + grind on_failure fallback /-- +info: hello world +--- info: true: [p, q] --- info: false: [r] ---- -warning: declaration uses 'sorry' -/ -#guard_msgs in +#guard_msgs (info) in example (p q r : Prop) : p ∨ (q ↔ r) → ¬r → q → False := by - grind_test - sorry + grind on_failure do + Lean.logInfo "hello world" + fallback From df9ed20385dec2f616533e6d61956b5e53ed3da2 Mon Sep 17 00:00:00 2001 From: Leonardo de Moura Date: Fri, 3 Jan 2025 02:05:44 +0100 Subject: [PATCH 009/100] feat: `ite` and `dite` support in `grind` (#6513) This PR adds support for (dependent) if-then-else terms (i.e., `ite` and `dite` applications) in the `grind` tactic. --- src/Init/Grind/Lemmas.lean | 5 ++++ src/Lean/Meta/Tactic/Grind/Propagate.lean | 30 +++++++++++++++++++++++ tests/lean/run/grind_t1.lean | 28 +++++++++++++++++++++ 3 files changed, 63 insertions(+) diff --git a/src/Init/Grind/Lemmas.lean b/src/Init/Grind/Lemmas.lean index d76e243ecf8d..e89e66b8725b 100644 --- a/src/Init/Grind/Lemmas.lean +++ b/src/Init/Grind/Lemmas.lean @@ -61,4 +61,9 @@ theorem forall_propagator (p : Prop) (q : p → Prop) (q' : Prop) (h₁ : p = Tr · intro h'; exact Eq.mp h₂ (h' (of_eq_true h₁)) · intro h'; intros; exact Eq.mpr h₂ h' +/-! dite -/ + +theorem dite_cond_eq_true' {α : Sort u} {c : Prop} {_ : Decidable c} {a : c → α} {b : ¬ c → α} {r : α} (h₁ : c = True) (h₂ : a (of_eq_true h₁) = r) : (dite c a b) = r := by simp [h₁, h₂] +theorem dite_cond_eq_false' {α : Sort u} {c : Prop} {_ : Decidable c} {a : c → α} {b : ¬ c → α} {r : α} (h₁ : c = False) (h₂ : b (of_eq_false h₁) = r) : (dite c a b) = r := by simp [h₁, h₂] + end Lean.Grind diff --git a/src/Lean/Meta/Tactic/Grind/Propagate.lean b/src/Lean/Meta/Tactic/Grind/Propagate.lean index 00dadb8f1c13..14a3a786ef84 100644 --- a/src/Lean/Meta/Tactic/Grind/Propagate.lean +++ b/src/Lean/Meta/Tactic/Grind/Propagate.lean @@ -7,6 +7,8 @@ prelude import Init.Grind import Lean.Meta.Tactic.Grind.Proof import Lean.Meta.Tactic.Grind.PropagatorAttr +import Lean.Meta.Tactic.Grind.Simp +import Lean.Meta.Tactic.Grind.Internalize namespace Lean.Meta.Grind @@ -142,4 +144,32 @@ builtin_grind_propagator propagateHEqUp ↑HEq := fun e => do if (← isEqv a b) then pushEqTrue e <| mkApp2 (mkConst ``eq_true) e (← mkHEqProof a b) +/-- Propagates `ite` upwards -/ +builtin_grind_propagator propagateIte ↑ite := fun e => do + let_expr f@ite α c h a b := e | return () + if (← isEqTrue c) then + pushEq e a <| mkApp6 (mkConst ``ite_cond_eq_true f.constLevels!) α c h a b (← mkEqTrueProof c) + else if (← isEqFalse c) then + pushEq e b <| mkApp6 (mkConst ``ite_cond_eq_false f.constLevels!) α c h a b (← mkEqFalseProof c) + +/-- Propagates `dite` upwards -/ +builtin_grind_propagator propagateDIte ↑dite := fun e => do + let_expr f@dite α c h a b := e | return () + if (← isEqTrue c) then + let h₁ ← mkEqTrueProof c + let ah₁ := mkApp a (mkApp2 (mkConst ``of_eq_true) c h₁) + let p ← simp ah₁ + let r := p.expr + let h₂ ← p.getProof + internalize r (← getGeneration e) + pushEq e r <| mkApp8 (mkConst ``Grind.dite_cond_eq_true' f.constLevels!) α c h a b r h₁ h₂ + else if (← isEqFalse c) then + let h₁ ← mkEqFalseProof c + let bh₁ := mkApp b (mkApp2 (mkConst ``of_eq_false) c h₁) + let p ← simp bh₁ + let r := p.expr + let h₂ ← p.getProof + internalize r (← getGeneration e) + pushEq e r <| mkApp8 (mkConst ``Grind.dite_cond_eq_false' f.constLevels!) α c h a b r h₁ h₂ + end Lean.Meta.Grind diff --git a/tests/lean/run/grind_t1.lean b/tests/lean/run/grind_t1.lean index 4bf739a213bc..168bdc895fff 100644 --- a/tests/lean/run/grind_t1.lean +++ b/tests/lean/run/grind_t1.lean @@ -83,3 +83,31 @@ info: [grind.debug.proj] { a := b, b := v₁, c := v₂ }.a set_option trace.grind.debug.proj true in example (a b d e : Nat) (x y z : Boo Nat) (f : Nat → Boo Nat) : (f d).1 ≠ a → f d = ⟨b, v₁, v₂⟩ → x.1 = e → y.1 = e → z.1 = e → f d = x → f d = y → f d = z → b = a → False := by grind + +example (f : Nat → Nat) (a b c : Nat) : f (if a = b then x else y) = z → a = c → c = b → f x = z := by + grind + +example (f : Nat → Nat) (a b c : Nat) : f (if a = b then x else y) = z → a = c → b ≠ c → f y = z := by + grind + +namespace dite_propagator_test + +opaque R : Nat → Nat → Prop +opaque f (a : Nat) (b : Nat) (_ : R a b) : Nat +opaque g (a : Nat) (b : Nat) (_ : ¬ R a b) : Nat +open Classical + +example (foo : Nat → Nat) + (_ : foo (if h : R a c then f a c h else g a c h) = x) + (_ : R a b) + (_ : c = b) : foo (f a c (by grind)) = x := by + grind + +example (foo : Nat → Nat) + (_ : foo (if h : R a c then f a c h else g a c h) = x) + (_ : ¬ R a b) + (_ : c = b) + : foo (g a c (by grind)) = x := by + grind + +end dite_propagator_test From 19078655bcfdf30f8309541309456055326dea6c Mon Sep 17 00:00:00 2001 From: Leonardo de Moura Date: Fri, 3 Jan 2025 03:16:08 +0100 Subject: [PATCH 010/100] perf: avoid unnecessary assert/intro pairs in `grind` (#6514) This PR enhances the assertion of new facts in `grind` by avoiding the creation of unnecessary metavariables. --- src/Lean/Meta/Tactic/Grind/Intro.lean | 21 ++++++++++++++------- 1 file changed, 14 insertions(+), 7 deletions(-) diff --git a/src/Lean/Meta/Tactic/Grind/Intro.lean b/src/Lean/Meta/Tactic/Grind/Intro.lean index 00438e10424a..5c3c9da4e115 100644 --- a/src/Lean/Meta/Tactic/Grind/Intro.lean +++ b/src/Lean/Meta/Tactic/Grind/Intro.lean @@ -62,12 +62,12 @@ private def introNext (goal : Goal) : GrindM IntroResult := do else return .done -private def isCasesCandidate (fvarId : FVarId) : MetaM Bool := do - let .const declName _ := (← fvarId.getType).getAppFn | return false +private def isCasesCandidate (type : Expr) : MetaM Bool := do + let .const declName _ := type.getAppFn | return false isGrindCasesTarget declName private def applyCases? (goal : Goal) (fvarId : FVarId) : MetaM (Option (List Goal)) := goal.mvarId.withContext do - if (← isCasesCandidate fvarId) then + if (← isCasesCandidate (← fvarId.getType)) then let mvarIds ← cases goal.mvarId fvarId return mvarIds.map fun mvarId => { goal with mvarId } else @@ -109,10 +109,17 @@ partial def intros (goal : Goal) (generation : Nat) : GrindM (List Goal) := do /-- Asserts a new fact `prop` with proof `proof` to the given `goal`. -/ def assertAt (goal : Goal) (proof : Expr) (prop : Expr) (generation : Nat) : GrindM (List Goal) := do - -- TODO: check whether `prop` may benefit from `intros` or not. If not, we should avoid the `assert`+`intros` step and use `Grind.add` - let mvarId ← goal.mvarId.assert (← mkFreshUserName `h) prop proof - let goal := { goal with mvarId } - intros goal generation + if (← isCasesCandidate prop) then + let mvarId ← goal.mvarId.assert (← mkFreshUserName `h) prop proof + let goal := { goal with mvarId } + intros goal generation + else + let goal ← GoalM.run' goal do + let r ← simp prop + let prop' := r.expr + let proof' ← mkEqMP (← r.getProof) proof + add prop' proof' generation + if goal.inconsistent then return [] else return [goal] /-- Asserts next fact in the `goal` fact queue. -/ def assertNext? (goal : Goal) : GrindM (Option (List Goal)) := do From 10b2f6b27e79e2c38d4d613f18ead3323a58ba4b Mon Sep 17 00:00:00 2001 From: Mitchell Lee <130885647+trivial1711@users.noreply.github.com> Date: Fri, 3 Jan 2025 04:37:02 -0500 Subject: [PATCH 011/100] feat: bdiv and bmod lemmas (#6494) This PR proves the basic theorems about the functions `Int.bdiv` and `Int.bmod`. For all integers `x` and all natural numbers `m`, we have: - `Int.bdiv_add_bmod`: `m * bdiv x m + bmod x m = x` (which is stated in the docstring for docs#Int.bdiv) - `Int.bmod_add_bdiv`: `bmod x m + m * bdiv x m = x` - `Int.bdiv_add_bmod'`: `bdiv x m * m + bmod x m = x` - `Int.bmod_add_bdiv'`: `bmod x m + bdiv x m * m = x` - `Int.bmod_eq_self_sub_mul_bdiv`: `bmod x m = x - m * bdiv x m` - `Int.bmod_eq_self_sub_bdiv_mul`: `bmod x m = x - bdiv x m * m` These theorems are all equivalent to each other by the basic properties of addition, multiplication, and subtraction of integers. The names `Int.bdiv_add_bmod`, `Int.bmod_add_bdiv`, `Int.bdiv_add_bmod'`, and `Int.bmod_add_bdiv'` are meant to parallel the names of the existing theorems docs#Int.tmod_add_tdiv, docs#Int.tdiv_add_tmod, docs#Int.tmod_add_tdiv', and docs#Int.tdiv_add_tmod'. The names `Int.bmod_eq_self_sub_mul_bdiv` and `Int.bmod_eq_self_sub_bdiv_mul` follow mathlib's naming conventions. Note that there is already a theorem called docs#Int.bmod_def, so it would not have been possible to parallel the name of the existing theorem docs#Int.tmod_def. See https://leanprover.zulipchat.com/#narrow/channel/217875-Is-there-code-for-X.3F/topic/bdiv.20and.20bmod. Closes #6493. --- src/Init/Data/Int/DivModLemmas.lean | 26 ++++++++++++++++++++++++++ 1 file changed, 26 insertions(+) diff --git a/src/Init/Data/Int/DivModLemmas.lean b/src/Init/Data/Int/DivModLemmas.lean index c211ed8f84e2..d0c86c5c98d0 100644 --- a/src/Init/Data/Int/DivModLemmas.lean +++ b/src/Init/Data/Int/DivModLemmas.lean @@ -1098,6 +1098,32 @@ theorem bmod_def (x : Int) (m : Nat) : bmod x m = (x % m) - m := rfl +theorem bdiv_add_bmod (x : Int) (m : Nat) : m * bdiv x m + bmod x m = x := by + unfold bdiv bmod + split + · simp_all only [Nat.cast_ofNat_Int, Int.mul_zero, emod_zero, Int.zero_add, Int.sub_zero, + ite_self] + · dsimp only + split + · exact ediv_add_emod x m + · rw [Int.mul_add, Int.mul_one, Int.add_assoc, Int.add_comm m, Int.sub_add_cancel] + exact ediv_add_emod x m + +theorem bmod_add_bdiv (x : Int) (m : Nat) : bmod x m + m * bdiv x m = x := by + rw [Int.add_comm]; exact bdiv_add_bmod x m + +theorem bdiv_add_bmod' (x : Int) (m : Nat) : bdiv x m * m + bmod x m = x := by + rw [Int.mul_comm]; exact bdiv_add_bmod x m + +theorem bmod_add_bdiv' (x : Int) (m : Nat) : bmod x m + bdiv x m * m = x := by + rw [Int.add_comm]; exact bdiv_add_bmod' x m + +theorem bmod_eq_self_sub_mul_bdiv (x : Int) (m : Nat) : bmod x m = x - m * bdiv x m := by + rw [← Int.add_sub_cancel (bmod x m), bmod_add_bdiv] + +theorem bmod_eq_self_sub_bdiv_mul (x : Int) (m : Nat) : bmod x m = x - bdiv x m * m := by + rw [← Int.add_sub_cancel (bmod x m), bmod_add_bdiv'] + theorem bmod_pos (x : Int) (m : Nat) (p : x % m < (m + 1) / 2) : bmod x m = x % m := by simp [bmod_def, p] From 7b496bf44b990a250ff5844fda41ecd9ddff2881 Mon Sep 17 00:00:00 2001 From: Leonardo de Moura Date: Fri, 3 Jan 2025 16:54:55 +0100 Subject: [PATCH 012/100] feat: improve `cases` tactic used in `grind` (#6516) This PR enhances the `cases` tactic used in the `grind` tactic and ensures that it can be applied to arbitrary expressions. --- src/Lean/Meta/Tactic/Cases.lean | 71 ++++++++++++++++++--------- src/Lean/Meta/Tactic/Grind/Cases.lean | 33 +++++++------ src/Lean/Meta/Tactic/Grind/Intro.lean | 2 +- tests/lean/run/grind_cases_tac.lean | 43 ++++++++++++++++ 4 files changed, 109 insertions(+), 40 deletions(-) create mode 100644 tests/lean/run/grind_cases_tac.lean diff --git a/src/Lean/Meta/Tactic/Cases.lean b/src/Lean/Meta/Tactic/Cases.lean index ecc00556bd19..56e8b8514a27 100644 --- a/src/Lean/Meta/Tactic/Cases.lean +++ b/src/Lean/Meta/Tactic/Cases.lean @@ -66,30 +66,31 @@ structure GeneralizeIndicesSubgoal where numEqs : Nat /-- - Similar to `generalizeTargets` but customized for the `casesOn` motive. - Given a metavariable `mvarId` representing the - ``` - Ctx, h : I A j, D |- T - ``` - where `fvarId` is `h`s id, and the type `I A j` is an inductive datatype where `A` are parameters, - and `j` the indices. Generate the goal - ``` - Ctx, h : I A j, D, j' : J, h' : I A j' |- j == j' -> h == h' -> T - ``` - Remark: `(j == j' -> h == h')` is a "telescopic" equality. - Remark: `j` is sequence of terms, and `j'` a sequence of free variables. - The result contains the fields - - `mvarId`: the new goal - - `indicesFVarIds`: `j'` ids - - `fvarId`: `h'` id - - `numEqs`: number of equations in the target -/ -def generalizeIndices (mvarId : MVarId) (fvarId : FVarId) : MetaM GeneralizeIndicesSubgoal := +Given a metavariable `mvarId` representing the goal +``` +Ctx |- T +``` +and an expression `e : I A j`, where `I A j` is an inductive datatype where `A` are parameters, +and `j` the indices. Generate the goal +``` +Ctx, j' : J, h' : I A j' |- j == j' -> e == h' -> T +``` +Remark: `(j == j' -> e == h')` is a "telescopic" equality. +Remark: `j` is sequence of terms, and `j'` a sequence of free variables. +The result contains the fields +- `mvarId`: the new goal +- `indicesFVarIds`: `j'` ids +- `fvarId`: `h'` id +- `numEqs`: number of equations in the target + +If `varName?` is not none, it is used to name `h'`. +-/ +def generalizeIndices' (mvarId : MVarId) (e : Expr) (varName? : Option Name := none) : MetaM GeneralizeIndicesSubgoal := mvarId.withContext do let lctx ← getLCtx let localInsts ← getLocalInstances mvarId.checkNotAssigned `generalizeIndices - let fvarDecl ← fvarId.getDecl - let type ← whnf fvarDecl.type + let type ← whnfD (← inferType e) type.withApp fun f args => matchConstInduct f (fun _ => throwTacticEx `generalizeIndices mvarId "inductive type expected") fun val _ => do unless val.numIndices > 0 do throwTacticEx `generalizeIndices mvarId "indexed inductive type expected" unless args.size == val.numIndices + val.numParams do throwTacticEx `generalizeIndices mvarId "ill-formed inductive datatype" @@ -98,9 +99,10 @@ def generalizeIndices (mvarId : MVarId) (fvarId : FVarId) : MetaM GeneralizeIndi let IAType ← inferType IA forallTelescopeReducing IAType fun newIndices _ => do let newType := mkAppN IA newIndices - withLocalDeclD fvarDecl.userName newType fun h' => + let varName ← if let some varName := varName? then pure varName else mkFreshUserName `x + withLocalDeclD varName newType fun h' => withNewEqs indices newIndices fun newEqs newRefls => do - let (newEqType, newRefl) ← mkEqAndProof fvarDecl.toExpr h' + let (newEqType, newRefl) ← mkEqAndProof e h' let newRefls := newRefls.push newRefl withLocalDeclD `h newEqType fun newEq => do let newEqs := newEqs.push newEq @@ -112,7 +114,7 @@ def generalizeIndices (mvarId : MVarId) (fvarId : FVarId) : MetaM GeneralizeIndi let auxType ← mkForallFVars newIndices auxType let newMVar ← mkFreshExprMVarAt lctx localInsts auxType MetavarKind.syntheticOpaque tag /- assign mvarId := newMVar indices h refls -/ - mvarId.assign (mkAppN (mkApp (mkAppN newMVar indices) fvarDecl.toExpr) newRefls) + mvarId.assign (mkAppN (mkApp (mkAppN newMVar indices) e) newRefls) let (indicesFVarIds, newMVarId) ← newMVar.mvarId!.introNP newIndices.size let (fvarId, newMVarId) ← newMVarId.intro1P return { @@ -122,6 +124,29 @@ def generalizeIndices (mvarId : MVarId) (fvarId : FVarId) : MetaM GeneralizeIndi numEqs := newEqs.size } +/-- +Similar to `generalizeTargets` but customized for the `casesOn` motive. +Given a metavariable `mvarId` representing the +``` +Ctx, h : I A j, D |- T +``` +where `fvarId` is `h`s id, and the type `I A j` is an inductive datatype where `A` are parameters, +and `j` the indices. Generate the goal +``` +Ctx, h : I A j, D, j' : J, h' : I A j' |- j == j' -> h == h' -> T +``` +Remark: `(j == j' -> h == h')` is a "telescopic" equality. +Remark: `j` is sequence of terms, and `j'` a sequence of free variables. +The result contains the fields +- `mvarId`: the new goal +- `indicesFVarIds`: `j'` ids +- `fvarId`: `h'` id +- `numEqs`: number of equations in the target -/ +def generalizeIndices (mvarId : MVarId) (fvarId : FVarId) : MetaM GeneralizeIndicesSubgoal := + mvarId.withContext do + let fvarDecl ← fvarId.getDecl + generalizeIndices' mvarId fvarDecl.toExpr fvarDecl.userName + structure CasesSubgoal extends InductionSubgoal where ctorName : Name diff --git a/src/Lean/Meta/Tactic/Grind/Cases.lean b/src/Lean/Meta/Tactic/Grind/Cases.lean index 9dd44b4f270e..d7e620f4497c 100644 --- a/src/Lean/Meta/Tactic/Grind/Cases.lean +++ b/src/Lean/Meta/Tactic/Grind/Cases.lean @@ -11,28 +11,29 @@ namespace Lean.Meta.Grind The `grind` tactic includes an auxiliary `cases` tactic that is not intended for direct use by users. This method implements it. This tactic is automatically applied when introducing local declarations with a type tagged with `[grind_cases]`. +It is also used for "case-splitting" on terms during the search. + It differs from the user-facing Lean `cases` tactic in the following ways: - It avoids unnecessary `revert` and `intro` operations. - It does not introduce new local declarations for each minor premise. Instead, the `grind` tactic preprocessor is responsible for introducing them. -- It assumes that the major premise (i.e., the parameter `fvarId`) is the latest local declaration in the current goal. - - If the major premise type is an indexed family, auxiliary declarations and (heterogeneous) equalities are introduced. However, these equalities are not resolved using `unifyEqs`. Instead, the `grind` tactic employs union-find and congruence closure to process these auxiliary equalities. This approach avoids applying substitution to propositions that have already been internalized by `grind`. -/ -def cases (mvarId : MVarId) (fvarId : FVarId) : MetaM (List MVarId) := mvarId.withContext do +def cases (mvarId : MVarId) (e : Expr) : MetaM (List MVarId) := mvarId.withContext do let tag ← mvarId.getTag - let type ← whnf (← fvarId.getType) + let type ← whnf (← inferType e) let .const declName _ := type.getAppFn | throwInductiveExpected type let .inductInfo _ ← getConstInfo declName | throwInductiveExpected type let recursorInfo ← mkRecursorInfo (mkCasesOnName declName) - let k (mvarId : MVarId) (fvarId : FVarId) (indices : Array Expr) (clearMajor : Bool) : MetaM (List MVarId) := do - let recursor ← mkRecursorAppPrefix mvarId `grind.cases fvarId recursorInfo indices - let mut recursor := mkApp (mkAppN recursor indices) (mkFVar fvarId) + let k (mvarId : MVarId) (fvarId : FVarId) (indices : Array FVarId) : MetaM (List MVarId) := do + let indicesExpr := indices.map mkFVar + let recursor ← mkRecursorAppPrefix mvarId `grind.cases fvarId recursorInfo indicesExpr + let mut recursor := mkApp (mkAppN recursor indicesExpr) (mkFVar fvarId) let mut recursorType ← inferType recursor let mut mvarIdsNew := #[] for _ in [:recursorInfo.numMinors] do @@ -41,22 +42,22 @@ def cases (mvarId : MVarId) (fvarId : FVarId) : MetaM (List MVarId) := mvarId.wi recursorType := recursorTypeNew let mvar ← mkFreshExprSyntheticOpaqueMVar targetNew tag recursor := mkApp recursor mvar - let mvarIdNew ← if clearMajor then - mvar.mvarId!.clear fvarId - else - pure mvar.mvarId! + let mvarIdNew ← mvar.mvarId!.tryClearMany (indices.push fvarId) mvarIdsNew := mvarIdsNew.push mvarIdNew mvarId.assign recursor return mvarIdsNew.toList if recursorInfo.numIndices > 0 then - let s ← generalizeIndices mvarId fvarId + let s ← generalizeIndices' mvarId e s.mvarId.withContext do - k s.mvarId s.fvarId (s.indicesFVarIds.map mkFVar) (clearMajor := false) + k s.mvarId s.fvarId s.indicesFVarIds + else if let .fvar fvarId := e then + k mvarId fvarId #[] else - let indices ← getMajorTypeIndices mvarId `grind.cases recursorInfo type - k mvarId fvarId indices (clearMajor := true) + let mvarId ← mvarId.assert (← mkFreshUserName `x) type e + let (fvarId, mvarId) ← mvarId.intro1 + mvarId.withContext do k mvarId fvarId #[] where throwInductiveExpected {α} (type : Expr) : MetaM α := do - throwTacticEx `grind.cases mvarId m!"(non-recursive) inductive type expected at {mkFVar fvarId}{indentExpr type}" + throwTacticEx `grind.cases mvarId m!"(non-recursive) inductive type expected at {e}{indentExpr type}" end Lean.Meta.Grind diff --git a/src/Lean/Meta/Tactic/Grind/Intro.lean b/src/Lean/Meta/Tactic/Grind/Intro.lean index 5c3c9da4e115..b0b60282b922 100644 --- a/src/Lean/Meta/Tactic/Grind/Intro.lean +++ b/src/Lean/Meta/Tactic/Grind/Intro.lean @@ -68,7 +68,7 @@ private def isCasesCandidate (type : Expr) : MetaM Bool := do private def applyCases? (goal : Goal) (fvarId : FVarId) : MetaM (Option (List Goal)) := goal.mvarId.withContext do if (← isCasesCandidate (← fvarId.getType)) then - let mvarIds ← cases goal.mvarId fvarId + let mvarIds ← cases goal.mvarId (mkFVar fvarId) return mvarIds.map fun mvarId => { goal with mvarId } else return none diff --git a/tests/lean/run/grind_cases_tac.lean b/tests/lean/run/grind_cases_tac.lean new file mode 100644 index 000000000000..c63200ab146b --- /dev/null +++ b/tests/lean/run/grind_cases_tac.lean @@ -0,0 +1,43 @@ +import Lean + +open Lean Meta Grind Elab Tactic in +elab "cases' " e:term : tactic => withMainContext do + let e ← elabTerm e none + setGoals (← Grind.cases (← getMainGoal) e) + +inductive Vec (α : Type u) : Nat → Type u + | nil : Vec α 0 + | cons : α → Vec α n → Vec α (n+1) + +def f (v : Vec α n) : Bool := + match v with + | .nil => true + | .cons .. => false + +/-- +info: n : Nat +v : Vec Nat n +h : f v ≠ false +⊢ n + 1 = 0 → HEq (Vec.cons 10 v) Vec.nil → False +--- +info: n : Nat +v : Vec Nat n +h : f v ≠ false +⊢ ∀ {n_1 : Nat} (a : Nat) (a_1 : Vec Nat n_1), n + 1 = n_1 + 1 → HEq (Vec.cons 10 v) (Vec.cons a a_1) → False +-/ +#guard_msgs (info) in +example (v : Vec Nat n) (h : f v ≠ false) : False := by + cases' (Vec.cons 10 v) + next => trace_state; sorry + next => trace_state; sorry + +/-- +info: ⊢ False → False +--- +info: ⊢ True → False +-/ +#guard_msgs (info) in +example : False := by + cases' (Or.inr (a := False) True.intro) + next => trace_state; sorry + next => trace_state; sorry From 58d178e68f47897ed804000be7abbd27cf389d4c Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Henrik=20B=C3=B6ving?= Date: Fri, 3 Jan 2025 20:35:58 +0100 Subject: [PATCH 013/100] fix: cond reflection bug in bv_decide (#6517) This PR fixes a slight bug that was created in the reflection of `bif` in `bv_decide`. Tagged as changelog-no as the code in question isn't in an RC yet. --- src/Std/Tactic/BVDecide/Reflect.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/Std/Tactic/BVDecide/Reflect.lean b/src/Std/Tactic/BVDecide/Reflect.lean index 89253715472e..070d10aa9f01 100644 --- a/src/Std/Tactic/BVDecide/Reflect.lean +++ b/src/Std/Tactic/BVDecide/Reflect.lean @@ -156,7 +156,7 @@ theorem or_congr (lhs rhs lhs' rhs' : Bool) (h1 : lhs' = lhs) (h2 : rhs' = rhs) theorem cond_congr (discr lhs rhs discr' lhs' rhs' : Bool) (h1 : discr' = discr) (h2 : lhs' = lhs) (h3 : rhs' = rhs) : - (bif discr' = true then lhs' else rhs') = (bif discr = true then lhs else rhs) := by + (bif discr' then lhs' else rhs') = (bif discr then lhs else rhs) := by simp[*] theorem false_of_eq_true_of_eq_false (h₁ : x = true) (h₂ : x = false) : False := by From d991feddadc9a1be4d385f06ef1f77ddebde4f42 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Sat, 4 Jan 2025 12:25:33 +1100 Subject: [PATCH 014/100] chore: cherry-pick release notes from releases/v4.15.0 and releases/v4.16.0 (#6520) These release notes were automatically generated by the script in #6519. --- RELEASES.md | 1187 ++++++++++++++++++++++++++++++++++- releases_drafts/list_lex.md | 16 - 2 files changed, 1185 insertions(+), 18 deletions(-) delete mode 100644 releases_drafts/list_lex.md diff --git a/RELEASES.md b/RELEASES.md index 7360725db20b..a657a38bd898 100644 --- a/RELEASES.md +++ b/RELEASES.md @@ -11,12 +11,1195 @@ of each version. v4.16.0 ---------- -Development in progress. +## Language + +* [#3696](https://github.com/leanprover/lean4/pull/3696) makes all message constructors handle pretty printer errors. + +* [#4460](https://github.com/leanprover/lean4/pull/4460) runs all linters for a single command (together) on a separate +thread from further elaboration, making a first step towards +parallelizing the elaborator. + +* [#5757](https://github.com/leanprover/lean4/pull/5757) makes it harder to create "fake" theorems about definitions that +are stubbed-out with `sorry` by ensuring that each `sorry` is not +definitionally equal to any other. For example, this now fails: +```lean +example : (sorry : Nat) = sorry := rfl -- fails +``` +However, this still succeeds, since the `sorry` is a single +indeterminate `Nat`: +```lean +def f (n : Nat) : Nat := sorry +example : f 0 = f 1 := rfl -- succeeds +``` +One can be more careful by putting parameters to the right of the colon: +```lean +def f : (n : Nat) → Nat := sorry +example : f 0 = f 1 := rfl -- fails +``` +Most sources of synthetic sorries (recall: a sorry that originates from +the elaborator) are now unique, except for elaboration errors, since +making these unique tends to cause a confusing cascade of errors. In +general, however, such sorries are labeled. This enables "go to +definition" on `sorry` in the Infoview, which brings you to its origin. +The option `set_option pp.sorrySource true` causes the pretty printer to +show source position information on sorries. + +* [#6123](https://github.com/leanprover/lean4/pull/6123) ensures that the configuration in `Simp.Config` is used when +reducing terms and checking definitional equality in `simp`. + +* [#6204](https://github.com/leanprover/lean4/pull/6204) lets `_` be used in numeric literals as a separator. For +example, `1_000_000`, `0xff_ff` or `0b_10_11_01_00`. New lexical syntax: +```text +numeral10 : [0-9]+ ("_"+ [0-9]+)* +numeral2 : "0" [bB] ("_"* [0-1]+)+ +numeral8 : "0" [oO] ("_"* [0-7]+)+ +numeral16 : "0" [xX] ("_"* hex_char+)+ +float : numeral10 "." numeral10? [eE[+-]numeral10] +``` + +* [#6270](https://github.com/leanprover/lean4/pull/6270) fixes a bug that could cause the `injectivity` tactic to fail in +reducible mode, which could cause unfolding lemma generation to fail +(used by tactics such as `unfold`). In particular, +`Lean.Meta.isConstructorApp'?` was not aware that `n + 1` is equivalent +to `Nat.succ n`. + +* [#6273](https://github.com/leanprover/lean4/pull/6273) modifies the "foo has been deprecated: use betterFoo instead" +warning so that foo and betterFoo are hoverable. + +* [#6278](https://github.com/leanprover/lean4/pull/6278) enables simp configuration options to be passed to `norm_cast`. + +* [#6286](https://github.com/leanprover/lean4/pull/6286) ensure `bv_decide` uses definitional equality in its reflection +procedure as much as possible. Previously it would build up explicit +congruence proofs for the kernel to check. This reduces the size of +proof terms passed to kernel speeds up checking of large reflection +proofs. + +* [#6288](https://github.com/leanprover/lean4/pull/6288) uses Lean.RArray in bv_decide's reflection proofs. Giving +speedups on problems with lots of variables. + +* [#6295](https://github.com/leanprover/lean4/pull/6295) sets up simprocs for all the remaining operations defined in +`Init.Data.Fin.Basic` + +* [#6300](https://github.com/leanprover/lean4/pull/6300) adds the `debug.proofAsSorry` option. When enabled, the proofs +of theorems are ignored and replaced with `sorry`. + +* [#6330](https://github.com/leanprover/lean4/pull/6330) removes unnecessary parameters from the funcion induction +principles. This is a breaking change; broken code can typically be adjusted +simply by passing fewer parameters. + +* [#6330](https://github.com/leanprover/lean4/pull/6330) removes unnecessary parameters from the funcion induction +principles. This is a breaking change; broken code can typically be adjusted +simply by passing fewer parameters. + +* [#6362](https://github.com/leanprover/lean4/pull/6362) adds the `--error=kind` option (shorthand: `-Ekind`) to the +`lean` CLI. When set, messages of `kind` (e.g., +`linter.unusedVariables`) will be reported as errors. This setting does +nothing in interactive contexts (e.g., the server). + +* [#6366](https://github.com/leanprover/lean4/pull/6366) adds support for `Float32` and fixes a bug in the runtime. + +* [#6375](https://github.com/leanprover/lean4/pull/6375) fixes a bug in the simplifier. It was producing terms with loose +bound variables when eliminating unused `let_fun` expressions. + +* [#6378](https://github.com/leanprover/lean4/pull/6378) adds an explanation to the error message when `cases` and +`induction` are applied to a term whose type is not an inductive type. +For `Prop`, these tactics now suggest the `by_cases` tactic. Example: +``` +tactic 'cases' failed, major premise type is not an inductive type + Prop + +* [#6381](https://github.com/leanprover/lean4/pull/6381) fixes a bug in `withTrackingZetaDelta` and +`withTrackingZetaDeltaSet`. The `MetaM` caches need to be reset. See new +test. + +* [#6385](https://github.com/leanprover/lean4/pull/6385) fixes a bug in `simp_all?` that caused some local declarations +to be omitted from the `Try this:` suggestions. + +* [#6386](https://github.com/leanprover/lean4/pull/6386) ensures that `revertAll` clears auxiliary declarations when +invoked directly by users. + +* [#6387](https://github.com/leanprover/lean4/pull/6387) fixes a type error in the proof generated by the `contradiction` +tactic. + +* [#6397](https://github.com/leanprover/lean4/pull/6397) ensures that `simp` and `dsimp` do not unfold definitions that +are not intended to be unfolded by the user. See issue #5755 for an +example affected by this issue. + +* [#6398](https://github.com/leanprover/lean4/pull/6398) ensures `Meta.check` check projections. + +* [#6412](https://github.com/leanprover/lean4/pull/6412) adds reserved names for congruence theorems used in the +simplifier and `grind` tactics. The idea is prevent the same congruence +theorems to be generated over and over again. + +* [#6413](https://github.com/leanprover/lean4/pull/6413) introduces the following features to the WIP `grind` tactic: +- `Expr` internalization. +- Congruence theorem cache. +- Procedure for adding new facts +- New tracing options +- New preprocessing steps: fold projections and eliminate dangling +`Expr.mdata` + +* [#6414](https://github.com/leanprover/lean4/pull/6414) fixes a bug in `Lean.Meta.Closure` that would introduce +under-applied delayed assignment metavariables, which would keep them +from ever getting instantiated. This bug affected `match` elaboration +when the expected type contained postponed elaboration problems, for +example tactic blocks. + +* [#6419](https://github.com/leanprover/lean4/pull/6419) fixes multiple bugs in the WIP `grind` tactic. It also adds +support for printing the `grind` internal state. + +* [#6428](https://github.com/leanprover/lean4/pull/6428) adds a new preprocessing step to the `grind` tactic: +universe-level normalization. The goal is to avoid missing equalities in +the congruence closure module. + +* [#6430](https://github.com/leanprover/lean4/pull/6430) adds the predicate `Expr.fvarsSet a b`, which returns `true` if +and only if the free variables in `a` are a subset of the free variables +in `b`. + +* [#6433](https://github.com/leanprover/lean4/pull/6433) adds a custom type and instance canonicalizer for the (WIP) +`grind` tactic. The `grind` tactic uses congruence closure but +disregards types, type formers, instances, and proofs. Proofs are +ignored due to proof irrelevance. Types, type formers, and instances are +considered supporting elements and are not factored into congruence +detection. Instead, `grind` only checks whether elements are +structurally equal, which, in the context of the `grind` tactic, is +equivalent to pointer equality. See new tests for examples where the +canonicalizer is important. + +* [#6435](https://github.com/leanprover/lean4/pull/6435) implements the congruence table for the (WIP) `grind` tactic. It +also fixes several bugs, and adds a new preprocessing step. + +* [#6437](https://github.com/leanprover/lean4/pull/6437) adds support for detecting congruent terms in the (WIP) `grind` +tactic. It also introduces the `grind.debug` option, which, when set to +`true`, checks many invariants after each equivalence class is merged. +This option is intended solely for debugging purposes. + +* [#6438](https://github.com/leanprover/lean4/pull/6438) ensures `norm_cast` doesn't fail to act in the presence of +`no_index` annotations + +* [#6441](https://github.com/leanprover/lean4/pull/6441) adds basic truth value propagation rules to the (WIP) `grind` +tactic. + +* [#6442](https://github.com/leanprover/lean4/pull/6442) fixes the `checkParents` sanity check in `grind`. + +* [#6443](https://github.com/leanprover/lean4/pull/6443) adds support for propagating the truth value of equalities in +the (WIP) `grind` tactic. + +* [#6447](https://github.com/leanprover/lean4/pull/6447) refactors `grind` and adds support for invoking the simplifier +using the `GrindM` monad. + +* [#6448](https://github.com/leanprover/lean4/pull/6448) declares the command `builtin_grind_propagator` for registering +equation propagator for `grind`. It also declares the auxiliary the +attribute. + +* [#6449](https://github.com/leanprover/lean4/pull/6449) completes the implementation of the command +`builtin_grind_propagator`. + +* [#6452](https://github.com/leanprover/lean4/pull/6452) adds support for generating (small) proofs for any two +expressions that belong to the same equivalence class in the `grind` +tactic state. + +* [#6453](https://github.com/leanprover/lean4/pull/6453) improves bv_decide's performance in the presence of large +literals. + +* [#6455](https://github.com/leanprover/lean4/pull/6455) fixes a bug in the equality proof generator in the (WIP) `grind` +tactic. + +* [#6456](https://github.com/leanprover/lean4/pull/6456) fixes another bug in the equality proof generator in the (WIP) +`grind` tactic. + +* [#6457](https://github.com/leanprover/lean4/pull/6457) adds support for generating congruence proofs for congruences +detected by the `grind` tactic. + +* [#6458](https://github.com/leanprover/lean4/pull/6458) adds support for compact congruence proofs in the (WIP) `grind` +tactic. The `mkCongrProof` function now verifies whether the congruence +proof can be constructed using only `congr`, `congrFun`, and `congrArg`, +avoiding the need to generate the more complex `hcongr` auxiliary +theorems. + +* [#6459](https://github.com/leanprover/lean4/pull/6459) adds the (WIP) `grind` tactic. It currently generates a warning +message to make it clear that the tactic is not ready for production. + +* [#6461](https://github.com/leanprover/lean4/pull/6461) adds a new propagation rule for negation to the (WIP) `grind` +tactic. + +* [#6463](https://github.com/leanprover/lean4/pull/6463) adds support for constructors to the (WIP) `grind` tactic. When +merging equivalence classes, `grind` checks for equalities between +constructors. If they are distinct, it closes the goal; if they are the +same, it applies injectivity. + +* [#6464](https://github.com/leanprover/lean4/pull/6464) completes support for literal values in the (WIP) `grind` +tactic. `grind` now closes the goal whenever it merges two equivalence +classes with distinct literal values. + +* [#6465](https://github.com/leanprover/lean4/pull/6465) adds support for projection functions to the (WIP) `grind` +tactic. + +* [#6466](https://github.com/leanprover/lean4/pull/6466) completes the implementation of `addCongrTable` in the (WIP) +`grind` tactic. It also adds a new test to demonstrate why the extra +check is needed. It also updates the field `cgRoot` (congruence root). + +* [#6468](https://github.com/leanprover/lean4/pull/6468) fixes issue #6467 + +* [#6469](https://github.com/leanprover/lean4/pull/6469) adds support code for implementing e-match in the (WIP) `grind` +tactic. + +* [#6470](https://github.com/leanprover/lean4/pull/6470) introduces a command for specifying patterns used in the +heuristic instantiation of global theorems in the `grind` tactic. Note +that this PR only adds the parser. + +* [#6472](https://github.com/leanprover/lean4/pull/6472) implements the command `grind_pattern`. The new command allows +users to associate patterns with theorems. These patterns are used for +performing heuristic instantiation with e-matching. In the future, we +will add the attributes `@[grind_eq]`, `@[grind_fwd]`, and +`@[grind_bwd]` to compute the patterns automatically for theorems. + +* [#6473](https://github.com/leanprover/lean4/pull/6473) adds a deriving handler for the `ToExpr` class. It can handle +mutual and nested inductive types, however it falls back to creating +`partial` instances in such cases. This is upstreamed from the Mathlib +deriving handler written by @kmill, but has fixes to handle autoimplicit +universe level variables. + +* [#6474](https://github.com/leanprover/lean4/pull/6474) adds pattern validation to the `grind_pattern` command. The new +`checkCoverage` function will also be used to implement the attributes +`@[grind_eq]`, `@[grind_fwd]`, and `@[grind_bwd]`. + +* [#6475](https://github.com/leanprover/lean4/pull/6475) adds support for activating relevant theorems for the (WIP) +`grind` tactic. We say a theorem is relevant to a `grind` goal if the +symbols occurring in its patterns also occur in the goal. + +* [#6478](https://github.com/leanprover/lean4/pull/6478) internalize nested ground patterns when activating ematch +theorems in the (WIP) `grind` tactic. + +* [#6481](https://github.com/leanprover/lean4/pull/6481) implements E-matching for the (WIP) `grind` tactic. We still +need to finalize and internalize the new instances. + +* [#6484](https://github.com/leanprover/lean4/pull/6484) addresses a few error messages where diffs weren't being +exposed. + +* [#6485](https://github.com/leanprover/lean4/pull/6485) implements `Grind.EMatch.instantiateTheorem` in the (WIP) +`grind` tactic. + +* [#6487](https://github.com/leanprover/lean4/pull/6487) adds source position information for `structure` parent +projections, supporting "go to definition". Closes #3063. + +* [#6488](https://github.com/leanprover/lean4/pull/6488) fixes and refactors the E-matching module for the (WIP) `grind` +tactic. + +* [#6490](https://github.com/leanprover/lean4/pull/6490) adds basic configuration options for the `grind` tactic. + +* [#6492](https://github.com/leanprover/lean4/pull/6492) fixes a bug in the theorem instantiation procedure in the (WIP) +`grind` tactic. For example, it was missing the following instance in +one of the tests: + +* [#6497](https://github.com/leanprover/lean4/pull/6497) fixes another theorem instantiation bug in the `grind` tactic. +It also moves new instances to be processed to `Goal`. + +* [#6498](https://github.com/leanprover/lean4/pull/6498) adds support in the `grind` tactic for propagating dependent +forall terms `forall (h : p), q[h]` where `p` is a proposition. + +* [#6499](https://github.com/leanprover/lean4/pull/6499) fixes the proof canonicalizer for `grind`. + +* [#6500](https://github.com/leanprover/lean4/pull/6500) fixes a bug in the `markNestedProofs` used in `grind`. See new +test. + +* [#6502](https://github.com/leanprover/lean4/pull/6502) fixes a bug in the proof assembly procedure utilized by the +`grind` tactic. + +* [#6503](https://github.com/leanprover/lean4/pull/6503) adds a simple strategy to the (WIP) `grind` tactic. It just +keeps internalizing new theorem instances found by E-matching. The +simple strategy can solve examples such as: + +* [#6506](https://github.com/leanprover/lean4/pull/6506) adds the `monotonicity` tactic, intended to be used inside the +`partial_fixpoint` feature. + +* [#6508](https://github.com/leanprover/lean4/pull/6508) fixes a bug in the sanity checkers for the `grind` tactic. See +the new test for an example of a case where it was panicking. + +* [#6509](https://github.com/leanprover/lean4/pull/6509) fixes a bug in the congruence closure data structure used in the +`grind` tactic. The new test includes an example that previously caused +a panic. A similar panic was also occurring in the test +`grind_nested_proofs.lean`. + +* [#6510](https://github.com/leanprover/lean4/pull/6510) adds a custom congruence rule for equality in `grind`. The new +rule takes into account that `Eq` is a symmetric relation. In the +future, we will add support for arbitrary symmetric relations. The +current rule is important for propagating disequalities effectively in +`grind`. + +* [#6512](https://github.com/leanprover/lean4/pull/6512) introduces support for user-defined fallback code in the `grind` +tactic. The fallback code can be utilized to inspect the state of +failing `grind` subgoals and/or invoke user-defined automation. Users +can now write `grind on_failure `, where `` should have the +type `GoalM Unit`. See the modified tests in this PR for examples. + +* [#6513](https://github.com/leanprover/lean4/pull/6513) adds support for (dependent) if-then-else terms (i.e., `ite` and +`dite` applications) in the `grind` tactic. + +* [#6514](https://github.com/leanprover/lean4/pull/6514) enhances the assertion of new facts in `grind` by avoiding the +creation of unnecessary metavariables. + +## Library + +* [#6182](https://github.com/leanprover/lean4/pull/6182) adds `BitVec.[toInt|toFin]_concat` and moves a couple of +theorems into the concat section, as `BitVec.msb_concat` is needed for +the `toInt_concat` proof. + +* [#6188](https://github.com/leanprover/lean4/pull/6188) completes the `toNat` theorems for the bitwise operations +(`and`, `or`, `xor`, `shiftLeft`, `shiftRight`) of the UInt types and +adds `toBitVec` theorems as well. It also renames `and_toNat` to +`toNat_and` to fit with the current naming convention. + +* [#6238](https://github.com/leanprover/lean4/pull/6238) adds theorems characterizing the value of the unsigned shift +right of a bitvector in terms of its 2s complement interpretation as an +integer. +Unsigned shift right by at least one bit makes the value of the +bitvector less than or equal to `2^(w-1)`, +makes the interpretation of the bitvector `Int` and `Nat` agree. +In the case when `n = 0`, then the shift right value equals the integer +interpretation. + +* [#6244](https://github.com/leanprover/lean4/pull/6244) changes the implementation of `HashMap.toList`, so the ordering +agrees with `HashMap.toArray`. + +* [#6272](https://github.com/leanprover/lean4/pull/6272) introduces the basic theory of permutations of `Array`s and +proves `Array.swap_perm`. + +* [#6282](https://github.com/leanprover/lean4/pull/6282) moves `IO.Channel` and `IO.Mutex` from `Init` to `Std.Sync` and +renames them to `Std.Channel` and `Std.Mutex`. + +* [#6294](https://github.com/leanprover/lean4/pull/6294) upstreams `List.length_flatMap`, `countP_flatMap` and +`count_flatMap` from Mathlib. These were not possible to state before we +upstreamed `List.sum`. + +* [#6315](https://github.com/leanprover/lean4/pull/6315) adds `protected` to `Fin.cast` and `BitVec.cast`, to avoid +confusion with `_root_.cast`. These should mostly be used via +dot-notation in any case. + +* [#6316](https://github.com/leanprover/lean4/pull/6316) adds lemmas simplifying `for` loops over `Option` into +`Option.pelim`, giving parity with lemmas simplifying `for` loops of +`List` into `List.fold`. + +* [#6317](https://github.com/leanprover/lean4/pull/6317) completes the basic API for BitVec.ofBool. + +* [#6318](https://github.com/leanprover/lean4/pull/6318) generalizes the universe level for `Array.find?`, by giving it a +separate implementation from `Array.findM?`. + +* [#6324](https://github.com/leanprover/lean4/pull/6324) adds `GetElem` lemmas for the basic `Vector` operations. + +* [#6333](https://github.com/leanprover/lean4/pull/6333) generalizes the panic functions to a type of `Sort u` rather +than `Type u`. This better supports universe polymorphic types and +avoids confusing errors. + +* [#6334](https://github.com/leanprover/lean4/pull/6334) adds `Nat` theorems for distributing `>>>` over bitwise +operations, paralleling those of `BitVec`. + +* [#6338](https://github.com/leanprover/lean4/pull/6338) adds `BitVec.[toFin|getMsbD]_setWidth` and +`[getMsb|msb]_signExtend` as well as `ofInt_toInt`. + +* [#6341](https://github.com/leanprover/lean4/pull/6341) generalizes `DecidableRel` to allow a heterogeneous relation. + +* [#6353](https://github.com/leanprover/lean4/pull/6353) reproduces the API around `List.any/all` for `Array.any/all`. + +* [#6364](https://github.com/leanprover/lean4/pull/6364) makes fixes suggested by the Batteries environment linters, +particularly `simpNF`, and `unusedHavesSuffices`. + +* [#6365](https://github.com/leanprover/lean4/pull/6365) expands the `Array.set` and `Array.setIfInBounds` lemmas to +match existing lemmas for `List.set`. + +* [#6367](https://github.com/leanprover/lean4/pull/6367) brings Vector lemmas about membership and indexing to parity +with List and Array. + +* [#6369](https://github.com/leanprover/lean4/pull/6369) adds lemmas about `Vector.set`, `anyM`, `any`, `allM`, and +`all`. + +* [#6376](https://github.com/leanprover/lean4/pull/6376) adds theorems about `==` on `Vector`, reproducing those already +on `List` and `Array`. + +* [#6379](https://github.com/leanprover/lean4/pull/6379) replaces the inductive predicate `List.lt` with an upstreamed version of `List.Lex` from Mathlib. +(Previously `Lex.lt` was defined in terms of `<`; now it is generalized to take an arbitrary relation.) +This subtly changes the notion of ordering on `List α`. + + `List.lt` was a weaker relation: in particular if `l₁ < l₂`, then `a :: l₁ < b :: l₂` may hold according to `List.lt` even if `a` and `b` are merely incomparable (either neither `a < b` nor `b < a`), whereas according to `List.Lex` this would require `a = b`. + + When `<` is total, in the sense that `¬ · < ·` is antisymmetric, then the two relations coincide. + + Mathlib was already overriding the order instances for `List α`, so this change should not be noticed by anyone already using Mathlib. + + We simultaneously add the boolean valued `List.lex` function, parameterised by a `BEq` typeclass and an arbitrary `lt` function. This will support the flexibility previously provided for `List.lt`, via a `==` function which is weaker than strict equality. + +* [#6390](https://github.com/leanprover/lean4/pull/6390) redefines `Range.forIn'` and `Range.forM`, in preparation for +writing lemmas about them. + +* [#6391](https://github.com/leanprover/lean4/pull/6391) requires that the step size in `Std.Range` is positive, to avoid +ill-specified behaviour. + +* [#6396](https://github.com/leanprover/lean4/pull/6396) adds lemmas reducing for loops over `Std.Range` to for loops +over `List.range'`. + +* [#6399](https://github.com/leanprover/lean4/pull/6399) adds basic lemmas about lexicographic order on Array and Vector, +achieving parity with List. + +* [#6423](https://github.com/leanprover/lean4/pull/6423) adds missing lemmas about lexicographic order on +List/Array/Vector. + +* [#6477](https://github.com/leanprover/lean4/pull/6477) adds the necessary domain theory that backs the +`partial_fixpoint` feature. + +## Compiler + +* [#6311](https://github.com/leanprover/lean4/pull/6311) adds support for `HEq` to the new code generator. + +* [#6348](https://github.com/leanprover/lean4/pull/6348) adds support for `Float32` to the Lean runtime. + +* [#6350](https://github.com/leanprover/lean4/pull/6350) adds missing features and fixes bugs in the `Float32` support + +* [#6383](https://github.com/leanprover/lean4/pull/6383) ensures the new code generator produces code for `opaque` +definitions that are not tagged as `@[extern]`. +Remark: This is the behavior of the old code generator. + +* [#6405](https://github.com/leanprover/lean4/pull/6405) adds support for erasure of `Decidable.decide` to the new code +generator. It also adds a new `Probe.runOnDeclsNamed` function, which is +helpful for writing targeted single-file tests of compiler internals. + +* [#6415](https://github.com/leanprover/lean4/pull/6415) fixes a bug in the `sharecommon` module, which was returning +incorrect results for objects that had already been processed by +`sharecommon`. See the new test for an example that triggered the bug. + +* [#6429](https://github.com/leanprover/lean4/pull/6429) adds support for extern LCNF decls, which is required for parity +with the existing code generator. + +## Pretty Printing + +* [#5689](https://github.com/leanprover/lean4/pull/5689) adjusts the way the pretty printer unresolves names. It used to +make use of all `export`s when pretty printing, but now it only uses +`export`s that put names into parent namespaces (heuristic: these are +"API exports" that are intended by the library author), rather than +"horizontal exports" that put the names into an unrelated namespace, +which the dot notation feature in #6189 now incentivizes. + +* [#5757](https://github.com/leanprover/lean4/pull/5757) makes it harder to create "fake" theorems about definitions that +are stubbed-out with `sorry` by ensuring that each `sorry` is not +definitionally equal to any other. For example, this now fails: +```lean +example : (sorry : Nat) = sorry := rfl -- fails +``` +However, this still succeeds, since the `sorry` is a single +indeterminate `Nat`: +```lean +def f (n : Nat) : Nat := sorry +example : f 0 = f 1 := rfl -- succeeds +``` +One can be more careful by putting parameters to the right of the colon: +```lean +def f : (n : Nat) → Nat := sorry +example : f 0 = f 1 := rfl -- fails +``` +Most sources of synthetic sorries (recall: a sorry that originates from +the elaborator) are now unique, except for elaboration errors, since +making these unique tends to cause a confusing cascade of errors. In +general, however, such sorries are labeled. This enables "go to +definition" on `sorry` in the Infoview, which brings you to its origin. +The option `set_option pp.sorrySource true` causes the pretty printer to +show source position information on sorries. + +## Documentation + +* [#6450](https://github.com/leanprover/lean4/pull/6450) adds a docstring to the `@[app_delab]` attribute. + +## Server + +* [#6279](https://github.com/leanprover/lean4/pull/6279) fixes a bug in structure instance field completion that caused +it to not function correctly for bracketed structure instances written +in Mathlib style. + +* [#6408](https://github.com/leanprover/lean4/pull/6408) fixes a regression where goals that don't exist were being +displayed. The regression was triggered by #5835 and originally caused +by #4926. + +## Lake + +* [#6176](https://github.com/leanprover/lean4/pull/6176) changes Lake's build process to no longer use `leanc` for +compiling C files or linking shared libraries and executables. Instead, +it directly invokes the bundled compiler (or the native compiler if +none) using the necessary flags. + +* [#6289](https://github.com/leanprover/lean4/pull/6289) adapts Lake modules to use `prelude` and includes them in the +`check-prelude` CI. + +* [#6291](https://github.com/leanprover/lean4/pull/6291) ensures the the log error position is properly preserved when +prepending stray log entries to the job log. It also adds comparison +support for `Log.Pos`. + +* [#6388](https://github.com/leanprover/lean4/pull/6388) merges `BuildJob` and `Job`, deprecating the former. `Job` now +contains a trace as part of its state which can be interacted with +monadically. also simplifies the implementation of `OpaqueJob`. + +* [#6411](https://github.com/leanprover/lean4/pull/6411) adds the ability to override package entries in a Lake manifest +via a separate JSON file. This file can be specified on the command line +with `--packages` or applied persistently by placing it at +`.lake/package-overrides.json`. + +* [#6422](https://github.com/leanprover/lean4/pull/6422) fixes a bug in #6388 where the `Package.afterBuildCahe*` +functions would produce different traces depending on whether the cache +was fetched. + +## Other + +* [#6285](https://github.com/leanprover/lean4/pull/6285) upstreams the `ToLevel` typeclass from mathlib and uses it to +fix the existing `ToExpr` instances so that they are truly universe +polymorphic (previously it generated malformed expressions when the +universe level was nonzero). We improve on the mathlib definition of +`ToLevel` to ensure the class always lives in `Type`, irrespective of +the universe parameter. + +* [#6363](https://github.com/leanprover/lean4/pull/6363) fixes errors at load time in the comparison mode of the Firefox +profiler. v4.15.0 ---------- -Release candidate, release notes will be copied from the branch `releases/v4.15.0` once completed. +## Language + +- [#4595](https://github.com/leanprover/lean4/pull/4595) implements `Simp.Config.implicitDefEqsProofs`. When `true` +(default: `true`), `simp` will **not** create a proof term for a +rewriting rule associated with an `rfl`-theorem. Rewriting rules are +provided by users by annotating theorems with the attribute `@[simp]`. +If the proof of the theorem is just `rfl` (reflexivity), and +`implicitDefEqProofs := true`, `simp` will **not** create a proof term +which is an application of the annotated theorem. + +- [#5429](https://github.com/leanprover/lean4/pull/5429) avoid negative environment lookup + +- [#5501](https://github.com/leanprover/lean4/pull/5501) ensure `instantiateMVarsProfiling` adds a trace node + +- [#5856](https://github.com/leanprover/lean4/pull/5856) adds a feature to the the mutual def elaborator where the +`instance` command yields theorems instead of definitions when the class +is a `Prop`. + +- [#5907](https://github.com/leanprover/lean4/pull/5907) unset trailing for `simpa?` "try this" suggestion + +- [#5920](https://github.com/leanprover/lean4/pull/5920) changes the rule for which projections become instances. Before, +all parents along with all indirect ancestors that were represented as +subobject fields would have their projections become instances. Now only +projections for direct parents become instances. + +- [#5934](https://github.com/leanprover/lean4/pull/5934) make `all_goals` admit goals on failure + +- [#5942](https://github.com/leanprover/lean4/pull/5942) introduce synthetic atoms in bv_decide + +- [#5945](https://github.com/leanprover/lean4/pull/5945) adds a new definition `Message.kind` which returns the top-level +tag of a message. This is serialized as the new field `kind` in +`SerialMessaege` so that i can be used by external consumers (e.g., +Lake) to identify messages via `lean --json`. + +- [#5968](https://github.com/leanprover/lean4/pull/5968) `arg` conv tactic misreported number of arguments on error + +- [#5979](https://github.com/leanprover/lean4/pull/5979) BitVec.twoPow in bv_decide + +- [#5991](https://github.com/leanprover/lean4/pull/5991) simplifies the implementation of `omega`. + +- [#5992](https://github.com/leanprover/lean4/pull/5992) fix style in bv_decide normalizer + +- [#5999](https://github.com/leanprover/lean4/pull/5999) adds configuration options for +`decide`/`decide!`/`native_decide` and refactors the tactics to be +frontends to the same backend. Adds a `+revert` option that cleans up +the local context and reverts all local variables the goal depends on, +along with indirect propositional hypotheses. Makes `native_decide` fail +at elaboration time on failure without sacrificing performance (the +decision procedure is still evaluated just once). Now `native_decide` +supports universe polymorphism. + +- [#6010](https://github.com/leanprover/lean4/pull/6010) changes `bv_decide`'s configuration from lots of `set_option` to +an elaborated config like `simp` or `omega`. The notable exception is +`sat.solver` which is still a `set_option` such that users can configure +a custom SAT solver globally for an entire project or file. Additionally +it introduces the ability to set `maxSteps` for the simp preprocessing +run through the new config. + +- [#6012](https://github.com/leanprover/lean4/pull/6012) improves the validation of new syntactic tokens. Previously, the +validation code had inconsistencies: some atoms would be accepted only +if they had a leading space as a pretty printer hint. Additionally, +atoms with internal whitespace are no longer allowed. + +- [#6016](https://github.com/leanprover/lean4/pull/6016) removes the `decide!` tactic in favor of `decide +kernel` +(breaking change). + +- [#6019](https://github.com/leanprover/lean4/pull/6019) removes @[specilize] from `MkBinding.mkBinding`, which is a +function that cannot be specialized (as none of its arguments are +functions). As a result, the specializable function `Nat.foldRevM.loop` +doesn't get specialized, which leads to worse performing code. + +- [#6022](https://github.com/leanprover/lean4/pull/6022) makes the `change` tactic and conv tactic use the same +elaboration strategy. It works uniformly for both the target and local +hypotheses. Now `change` can assign metavariables, for example: +```lean +example (x y z : Nat) : x + y = z := by + change ?a = _ + let w := ?a + -- now `w : Nat := x + y` +``` + +- [#6024](https://github.com/leanprover/lean4/pull/6024) fixes a bug where the monad lift coercion elaborator would +partially unify expressions even if they were not monads. This could be +taken advantage of to propagate information that could help elaboration +make progress, for example the first `change` worked because the monad +lift coercion elaborator was unifying `@Eq _ _` with `@Eq (Nat × Nat) +p`: +```lean +example (p : Nat × Nat) : p = p := by + change _ = ⟨_, _⟩ -- used to work (yielding `p = (p.fst, p.snd)`), now it doesn't + change ⟨_, _⟩ = _ -- never worked +``` +As such, this is a breaking change; you may need to adjust expressions +to include additional implicit arguments. + +- [#6029](https://github.com/leanprover/lean4/pull/6029) adds a normalization rule to `bv_normalize` (which is used by +`bv_decide`) that converts `x / 2^k` into `x >>> k` under suitable +conditions. This allows us to simplify the expensive division circuits +that are used for bitblasting into much cheaper shifting circuits. +Concretely, it allows for the following canonicalization: + +- [#6030](https://github.com/leanprover/lean4/pull/6030) fixes `simp only [· ∈ ·]` after #5020. + +- [#6035](https://github.com/leanprover/lean4/pull/6035) introduces the and flattening pre processing pass from Bitwuzla +to `bv_decide`. It splits hypotheses of the form `(a && b) = true` into +`a = true` and `b = true` which has synergy potential with the already +existing embedded constraint substitution pass. + +- [#6037](https://github.com/leanprover/lean4/pull/6037) fixes `bv_decide`'s embedded constraint substitution to generate +correct counter examples in the corner case where duplicate theorems are +in the local context. + +- [#6045](https://github.com/leanprover/lean4/pull/6045) add `LEAN_ALWAYS_INLINE` to some functions + +- [#6048](https://github.com/leanprover/lean4/pull/6048) fixes `simp?` suggesting output with invalid indentation + +- [#6051](https://github.com/leanprover/lean4/pull/6051) mark `Meta.Context.config` as private + +- [#6053](https://github.com/leanprover/lean4/pull/6053) fixes the caching infrastructure for `whnf` and `isDefEq`, +ensuring the cache accounts for all relevant configuration flags. It +also cleans up the `WHNF.lean` module and improves the configuration of +`whnf`. + +- [#6061](https://github.com/leanprover/lean4/pull/6061) adds a simp_arith benchmark. + +- [#6062](https://github.com/leanprover/lean4/pull/6062) optimize Nat.Linear.Expr.toPoly + +- [#6064](https://github.com/leanprover/lean4/pull/6064) optimize Nat.Linear.Poly.norm + +- [#6068](https://github.com/leanprover/lean4/pull/6068) improves the asymptotic performance of `simp_arith` when there are many variables to consider. + +- [#6077](https://github.com/leanprover/lean4/pull/6077) adds options to `bv_decide`'s configuration structure such that +all non mandatory preprocessing passes can be disabled. + +- [#6082](https://github.com/leanprover/lean4/pull/6082) changes how the canonicalizer handles `forall` and `lambda`, +replacing bvars with temporary fvars. Fixes a bug reported by @hrmacbeth +on +[zulip](https://leanprover.zulipchat.com/#narrow/channel/270676-lean4/topic/Quantifiers.20in.20CanonM/near/482483448). + +- [#6093](https://github.com/leanprover/lean4/pull/6093) use mkFreshUserName in ArgsPacker + +- [#6096](https://github.com/leanprover/lean4/pull/6096) improves the `#print` command for structures to show all fields +and which parents the fields were inherited from, hiding internal +details such as which parents are represented as subobjects. This +information is still present in the constructor if needed. The pretty +printer for private constants is also improved, and it now handles +private names from the current module like any other name; private names +from other modules are made hygienic. + +- [#6098](https://github.com/leanprover/lean4/pull/6098) modifies `Lean.MVarId.replaceTargetDefEq` and +`Lean.MVarId.replaceLocalDeclDefEq` to use `Expr.equal` instead of +`Expr.eqv` when determining whether the expression has changed. This is +justified on the grounds that binder names and binder infos are +user-visible and affect elaboration. + +- [#6105](https://github.com/leanprover/lean4/pull/6105) fixes a stack overflow caused by a cyclic assignment in the +metavariable context. The cycle is unintentionally introduced by the +structure instance elaborator. + +- [#6108](https://github.com/leanprover/lean4/pull/6108) turn off pp.mvars in apply? results + +- [#6109](https://github.com/leanprover/lean4/pull/6109) fixes an issue in the `injection` tactic. This tactic may +execute multiple sub-tactics. If any of them fail, we must backtrack the +partial assignment. This issue was causing the error: "`mvarId` is +already assigned" in issue #6066. The issue is not yet resolved, as the +equation generator for the match expressions is failing in the example +provided in this issue. + +- [#6112](https://github.com/leanprover/lean4/pull/6112) makes stricter requirements for the `@[deprecated]` attribute, +requiring either a replacement identifier as `@[deprecated bar]` or +suggestion text `@[deprecated "Past its use by date"]`, and also +requires a `since := "..."` field. + +- [#6114](https://github.com/leanprover/lean4/pull/6114) liberalizes atom rules by allowing `''` to be a prefix of an +atom, after #6012 only added an exception for `''` alone, and also adds +some unit tests for atom validation. + +- [#6116](https://github.com/leanprover/lean4/pull/6116) fixes a bug where structural recursion did not work when indices +of the recursive argument appeared as function parameters in a different +order than in the argument's type's definition. + +- [#6125](https://github.com/leanprover/lean4/pull/6125) adds support for `structure` in `mutual` blocks, allowing +inductive types defined by `inductive` and `structure` to be mutually +recursive. The limitations are (1) that the parents in the `extends` +clause must be defined before the `mutual` block and (2) mutually +recursive classes are not allowed (a limitation shared by `class +inductive`). There are also improvements to universe level inference for +inductive types and structures. Breaking change: structure parents now +elaborate with the structure in scope (fix: use qualified names or +rename the structure to avoid shadowing), and structure parents no +longer elaborate with autoimplicits enabled. + +- [#6128](https://github.com/leanprover/lean4/pull/6128) does the same fix as #6104, but such that it doesn't break the +test/the file in `Plausible`. This is done by not creating unused let +binders in metavariable types that are made by `elimMVar`. (This is also +a positive thing for users looking at metavariable types, for example in +error messages) + +- [#6129](https://github.com/leanprover/lean4/pull/6129) fixes a bug at `isDefEq` when `zetaDelta := false`. See new test +for a small example that exposes the issue. + +- [#6131](https://github.com/leanprover/lean4/pull/6131) fixes a bug at the definitional equality test (`isDefEq`). At +unification constraints of the form `c.{u} =?= c.{v}`, it was not trying +to unfold `c`. This bug did not affect the kernel. + +- [#6141](https://github.com/leanprover/lean4/pull/6141) make use of recursive structures in snapshot types + +- [#6145](https://github.com/leanprover/lean4/pull/6145) fixes the `revert` tactic so that it creates a `syntheticOpaque` +metavariable as the new goal, instead of a `natural` metavariable + +- [#6146](https://github.com/leanprover/lean4/pull/6146) fixes a non-termination bug that occurred when generating the +match-expression splitter theorem. The bug was triggered when the proof +automation for the splitter theorem repeatedly applied `injection` to +the same local declaration, as it could not be removed due to forward +dependencies. See issue #6065 for an example that reproduces this issue. + +- [#6165](https://github.com/leanprover/lean4/pull/6165) modifies structure instance notation and `where` notation to use +the same notation for fields. Structure instance notation now admits +binders, type ascriptions, and equations, and `where` notation admits +full structure lvals. Examples of these for structure instance notation: +```lean +structure PosFun where + f : Nat → Nat + pos : ∀ n, 0 < f n + +- [#6168](https://github.com/leanprover/lean4/pull/6168) extends the "motive is not type correct" error message for the +rewrite tactic to explain what it means. It also pretty prints the +type-incorrect motive and reports the type error. + +- [#6170](https://github.com/leanprover/lean4/pull/6170) adds core metaprogramming functions for forking off background +tasks from elaboration such that their results are visible to reporting +and the language server + +- [#6175](https://github.com/leanprover/lean4/pull/6175) fixes a bug with the `structure`/`class` command where if there +are parents that are not represented as subobjects but which used other +parents as instances, then there would be a kernel error. Closes #2611. + +- [#6180](https://github.com/leanprover/lean4/pull/6180) fixes a non-termination bug that occurred when generating the +match-expression equation theorems. The bug was triggered when the proof +automation for the equation theorem repeatedly applied `injection(` to +the same local declaration, as it could not be removed due to forward +dependencies. See issue #6067 for an example that reproduces this issue. + +- [#6189](https://github.com/leanprover/lean4/pull/6189) changes how generalized field notation ("dot notation") resolves +the function. The new resolution rule is that if `x : S`, then `x.f` +resolves the name `S.f` relative to the root namespace (hence it now +affected by `export` and `open`). Breaking change: aliases now resolve +differently. Before, if `x : S`, and if `S.f` is an alias for `S'.f`, +then `x.f` would use `S'.f` and look for an argument of type `S'`. Now, +it looks for an argument of type `S`, which is more generally useful +behavior. Code making use of the old behavior should consider defining +`S` or `S'` in terms of the other, since dot notation can unfold +definitions during resolution. + +- [#6206](https://github.com/leanprover/lean4/pull/6206) makes it possible to write `rw (occs := [1,2]) ...` instead of +`rw (occs := .pos [1,2]) ...` by adding a coercion from `List.Nat` to +`Lean.Meta.Occurrences`. + +- [#6220](https://github.com/leanprover/lean4/pull/6220) adds proper support for `let_fun` in `simp`. + +- [#6236](https://github.com/leanprover/lean4/pull/6236) fixes an issue where edits to a command containing a nested +docstring fail to reparse the entire command. + +## Library + +- [#4904](https://github.com/leanprover/lean4/pull/4904) introduces date and time functionality to the Lean 4 Std. + +- [#5616](https://github.com/leanprover/lean4/pull/5616) is a follow-up to https://github.com/leanprover/lean4/pull/5609, +where we add lemmas characterizing `smtUDiv` and `smtSDiv`'s behavior +when the denominator is zero. + +- [#5866](https://github.com/leanprover/lean4/pull/5866) verifies the `keys` function on `Std.HashMap`. + +- [#5885](https://github.com/leanprover/lean4/pull/5885) add Int16/Int32/Int64 + +- [#5926](https://github.com/leanprover/lean4/pull/5926) add `Option.or_some'` + +- [#5927](https://github.com/leanprover/lean4/pull/5927) `List.pmap_eq_self` + +- [#5937](https://github.com/leanprover/lean4/pull/5937) upstream lemmas about Fin.foldX + +- [#5938](https://github.com/leanprover/lean4/pull/5938) upstream List.ofFn and relate to Array.ofFn + +- [#5941](https://github.com/leanprover/lean4/pull/5941) List.mapFinIdx, lemmas, relate to Array version + +- [#5949](https://github.com/leanprover/lean4/pull/5949) consolidate `decide_True` and `decide_true_eq_true` + +- [#5950](https://github.com/leanprover/lean4/pull/5950) relate Array.takeWhile with List.takeWhile + +- [#5951](https://github.com/leanprover/lean4/pull/5951) remove @[simp] from BitVec.ofFin_sub and sub_ofFin + +- [#5952](https://github.com/leanprover/lean4/pull/5952) relate Array.eraseIdx with List.eraseIdx + +- [#5961](https://github.com/leanprover/lean4/pull/5961) define ISize and basic operations on it + +- [#5969](https://github.com/leanprover/lean4/pull/5969) upstream List.insertIdx from Batteries, lemmas from Mathlib, and revise lemmas + +- [#5970](https://github.com/leanprover/lean4/pull/5970) deprecate Array.split in favour of identical Array.partition + +- [#5971](https://github.com/leanprover/lean4/pull/5971) relate Array.isPrefixOf with List.isPrefixOf + +- [#5972](https://github.com/leanprover/lean4/pull/5972) relate Array.zipWith/zip/unzip with List versions + +- [#5974](https://github.com/leanprover/lean4/pull/5974) add another List.find?_eq_some lemma + +- [#5981](https://github.com/leanprover/lean4/pull/5981) names the default SizeOf instance `instSizeOfDefault` + +- [#5982](https://github.com/leanprover/lean4/pull/5982) minor lemmas about List.ofFn + +- [#5984](https://github.com/leanprover/lean4/pull/5984) adds lemmas for `List` for the interactions between {`foldl`, +`foldr`, `foldlM`, `foldlrM`} and {`filter`, `filterMap`}. + +- [#5985](https://github.com/leanprover/lean4/pull/5985) relates the operations `findSomeM?`, `findM?`, `findSome?`, and +`find?` on `Array` with the corresponding operations on `List`, and also +provides simp lemmas for the `Array` operations `findSomeRevM?`, +`findRevM?`, `findSomeRev?`, `findRev?` (in terms of `reverse` and the +usual forward find operations). + +- [#5987](https://github.com/leanprover/lean4/pull/5987) BitVec.getMsbD in bv_decide + +- [#5988](https://github.com/leanprover/lean4/pull/5988) changes the signature of `Array.set` to take a `Nat`, and a +tactic-provided bound, rather than a `Fin`. + +- [#5995](https://github.com/leanprover/lean4/pull/5995) BitVec.sshiftRight' in bv_decide + +- [#6007](https://github.com/leanprover/lean4/pull/6007) List.modifyTailIdx naming fix + +- [#6008](https://github.com/leanprover/lean4/pull/6008) missing @[ext] attribute on monad transformer ext lemmas + +- [#6023](https://github.com/leanprover/lean4/pull/6023) variants of List.forIn_eq_foldlM + +- [#6025](https://github.com/leanprover/lean4/pull/6025) deprecate duplicated Fin.size_pos + +- [#6032](https://github.com/leanprover/lean4/pull/6032) changes the signature of `Array.get` to take a Nat and a proof, +rather than a `Fin`, for consistency with the rest of the (planned) +Array API. Note that because of bootstrapping issues we can't provide +`get_elem_tactic` as an autoparameter for the proof. As users will +mostly use the `xs[i]` notation provided by `GetElem`, this hopefully +isn't a problem. + +- [#6041](https://github.com/leanprover/lean4/pull/6041) modifies the order of arguments for higher-order `Array` +functions, preferring to put the `Array` last (besides positional +arguments with defaults). This is more consistent with the `List` API, +and is more flexible, as dot notation allows two different partially +applied versions. + +- [#6049](https://github.com/leanprover/lean4/pull/6049) adds a primitive for accessing the current thread ID + +- [#6052](https://github.com/leanprover/lean4/pull/6052) adds `Array.pmap`, as well as a `@[csimp]` lemma in terms of the +no-copy `Array.attachWith`. + +- [#6055](https://github.com/leanprover/lean4/pull/6055) adds lemmas about for loops over `Array`, following the existing +lemmas for `List`. + +- [#6056](https://github.com/leanprover/lean4/pull/6056) upstream some NameMap functions + +- [#6060](https://github.com/leanprover/lean4/pull/6060) implements conversion functions from `Bool` to all `UIntX` and +`IntX` types. + +- [#6070](https://github.com/leanprover/lean4/pull/6070) adds the Lean.RArray data structure. + +- [#6074](https://github.com/leanprover/lean4/pull/6074) allow `Sort u` in `Squash` + +- [#6094](https://github.com/leanprover/lean4/pull/6094) adds raw transmutation of floating-point numbers to and from +`UInt64`. Floats and UInts share the same endianness across all +supported platforms. The IEEE 754 standard precisely specifies the bit +layout of floats. Note that `Float.toBits` is distinct from +`Float.toUInt64`, which attempts to preserve the numeric value rather +than the bitwise value. + +- [#6095](https://github.com/leanprover/lean4/pull/6095) generalize `List.get_mem` + +- [#6097](https://github.com/leanprover/lean4/pull/6097) naming convention and `NaN` normalization + +- [#6102](https://github.com/leanprover/lean4/pull/6102) moves `IO.rand` and `IO.setRandSeed` to be in the `BaseIO` +monad. + +- [#6106](https://github.com/leanprover/lean4/pull/6106) fix naming of left/right injectivity lemmas + +- [#6111](https://github.com/leanprover/lean4/pull/6111) fills in the API for `Array.findSome?` and `Array.find?`, +transferring proofs from the corresponding List statements. + +- [#6120](https://github.com/leanprover/lean4/pull/6120) adds theorems `BitVec.(getMsbD, msb)_(rotateLeft, rotateRight)`. + +- [#6126](https://github.com/leanprover/lean4/pull/6126) adds lemmas for extracting a given bit of a `BitVec` obtained +via `sub`/`neg`/`sshiftRight'`/`abs`. + +- [#6130](https://github.com/leanprover/lean4/pull/6130) adds `Lean.loadPlugin` which exposes functionality similar to +the `lean` executable's `--plugin` option to Lean code. + +- [#6132](https://github.com/leanprover/lean4/pull/6132) duplicates the verification API for +`List.attach`/`attachWith`/`pmap` over to `Array`. + +- [#6133](https://github.com/leanprover/lean4/pull/6133) replaces `Array.feraseIdx` and `Array.insertAt` with +`Array.eraseIdx` and `Array.insertIdx`, both of which take a `Nat` +argument and a tactic-provided proof that it is in bounds. We also have +`eraseIdxIfInBounds` and `insertIdxIfInBounds` which are noops if the +index is out of bounds. We also provide a `Fin` valued version of +`Array.findIdx?`. Together, these quite ergonomically improve the array +indexing safety at a number of places in the compiler/elaborator. + +- [#6136](https://github.com/leanprover/lean4/pull/6136) fixes the run-time evaluation of `(default : Float)`. + +- [#6139](https://github.com/leanprover/lean4/pull/6139) modifies the signature of the functions `Nat.fold`, +`Nat.foldRev`, `Nat.any`, `Nat.all`, so that the function is passed the +upper bound. This allows us to change runtime array bounds checks to +compile time checks in many places. + +- [#6148](https://github.com/leanprover/lean4/pull/6148) adds a primitive for creating temporary directories, akin to the +existing functionality for creating temporary files. + +- [#6149](https://github.com/leanprover/lean4/pull/6149) completes the elementwise accessors for `ofNatLt`, `allOnes`, +and `not` by adding their implementations of `getMsbD`. + +- [#6151](https://github.com/leanprover/lean4/pull/6151) completes the `toInt` interface for `BitVec` bitwise operations. + +- [#6154](https://github.com/leanprover/lean4/pull/6154) implements `BitVec.toInt_abs`. + +- [#6155](https://github.com/leanprover/lean4/pull/6155) adds `toNat` theorems for `BitVec.signExtend.` + +- [#6157](https://github.com/leanprover/lean4/pull/6157) adds toInt theorems for BitVec.signExtend. + +- [#6160](https://github.com/leanprover/lean4/pull/6160) adds theorem `mod_eq_sub`, makes theorem +`sub_mul_eq_mod_of_lt_of_le` not private anymore and moves its location +within the `rotate*` section to use it in other proofs. + +- [#6184](https://github.com/leanprover/lean4/pull/6184) uses `Array.findFinIdx?` in preference to `Array.findIdx?` where +it allows converting a runtime bounds check to a compile time bounds +check. + +- [#6188](https://github.com/leanprover/lean4/pull/6188) completes the `toNat` theorems for the bitwise operations +(`and`, `or`, `xor`, `shiftLeft`, `shiftRight`) of the UInt types and +adds `toBitVec` theorems as well. It also renames `and_toNat` to +`toNat_and` to fit with the current naming convention. + +- [#6190](https://github.com/leanprover/lean4/pull/6190) adds the builtin simproc `USize.reduceToNat` which reduces the +`USize.toNat` operation on literals less than `UInt32.size` (i.e., +`4294967296`). + +- [#6191](https://github.com/leanprover/lean4/pull/6191) adds `Array.zipWithAll`, and the basic lemmas relating it to +`List.zipWithAll`. + +- [#6192](https://github.com/leanprover/lean4/pull/6192) adds deprecations for `Lean.HashMap` functions which did not +receive deprecation attributes initially. + +- [#6193](https://github.com/leanprover/lean4/pull/6193) completes the TODO in `Init.Data.Array.BinSearch`, removing the +`partial` keyword and converting runtime bounds checks to compile time +bounds checks. + +- [#6194](https://github.com/leanprover/lean4/pull/6194) changes the signature of `Array.swap`, so it takes `Nat` +arguments with tactic provided bounds checking. It also renames +`Array.swap!` to `Array.swapIfInBounds`. + +- [#6195](https://github.com/leanprover/lean4/pull/6195) renames `Array.setD` to `Array.setIfInBounds`. + +- [#6197](https://github.com/leanprover/lean4/pull/6197) upstreams the definition of `Vector` from Batteries, along with +the basic functions. + +- [#6200](https://github.com/leanprover/lean4/pull/6200) upstreams `Nat.lt_pow_self` and `Nat.lt_two_pow` from Mathlib +and uses them to prove the simp theorem `Nat.mod_two_pow`. + +- [#6202](https://github.com/leanprover/lean4/pull/6202) makes `USize.toUInt64` a regular non-opaque definition. + +- [#6203](https://github.com/leanprover/lean4/pull/6203) adds the theorems `le_usize_size` and `usize_size_le`, which +make proving inequalities about `USize.size` easier. + +- [#6205](https://github.com/leanprover/lean4/pull/6205) upstreams some UInt theorems from Batteries and adds more +`toNat`-related theorems. It also adds the missing `UInt8` and `UInt16` +to/from `USize` conversions so that the the interface is uniform across +the UInt types. + +- [#6207](https://github.com/leanprover/lean4/pull/6207) ensures the `Fin.foldl` and `Fin.foldr` are semireducible. +Without this the defeq `example (f : Fin 3 → ℕ) : List.ofFn f = [f 0, f +1, f 2] := rfl` was failing. + +- [#6208](https://github.com/leanprover/lean4/pull/6208) fix Vector.indexOf? + +- [#6217](https://github.com/leanprover/lean4/pull/6217) adds `simp` lemmas about `List`'s `==` operation. + +- [#6221](https://github.com/leanprover/lean4/pull/6221) fixes: +- Problems in other linux distributions that the default `tzdata` +directory is not the same as previously defined by ensuring it with a +fallback behavior when directory is missing. +- Trim unnecessary characters from local time identifier. + +- [#6222](https://github.com/leanprover/lean4/pull/6222) changes the definition of `HashSet.insertMany` and +`HashSet.Raw.insertMany` so that it is equivalent to repeatedly calling +`HashSet.insert`/`HashSet.Raw.insert`. It also clarifies the docstrings +of all the `insert` and `insertMany` functions. + +- [#6230](https://github.com/leanprover/lean4/pull/6230) copies some lemmas about `List.foldX` to `Array`. + +- [#6233](https://github.com/leanprover/lean4/pull/6233) upstreams lemmas about `Vector` from Batteries. + +- [#6234](https://github.com/leanprover/lean4/pull/6234) upstreams the definition and basic lemmas about `List.finRange` +from Batteries. + +- [#6235](https://github.com/leanprover/lean4/pull/6235) relates that operations `Nat.fold`/`foldRev`/`any`/`all` to the +corresponding List operations over `List.finRange`. + +- [#6241](https://github.com/leanprover/lean4/pull/6241) refactors `Array.qsort` to remove runtime array bounds checks, +and avoids the use of `partial`. We use the `Vector` API, along with +auto_params, to avoid having to write any proofs. The new code +benchmarks indistinguishably from the old. + +- [#6242](https://github.com/leanprover/lean4/pull/6242) deprecates `Fin.ofNat` in favour of `Fin.ofNat'` (which takes an +`[NeZero]` instance, rather than returning an element of `Fin (n+1)`). + +- [#6247](https://github.com/leanprover/lean4/pull/6247) adds the theorems `numBits_pos`, `le_numBits`, `numBits_le` , +which make proving inequalities about `System.Platform.numBits` easier. + +## Compiler + +- [#5840](https://github.com/leanprover/lean4/pull/5840) changes `lean_sharecommon_{eq,hash}` to only consider the +salient bytes of an object, and not any bytes of any +unspecified/uninitialized unused capacity. + +- [#6087](https://github.com/leanprover/lean4/pull/6087) fixes a bug in the constant folding for the `Nat.ble` and +`Nat.blt` function in the old code generator, leading to a +miscompilation. + +- [#6143](https://github.com/leanprover/lean4/pull/6143) should make lean better-behaved around sanitizers, per +https://github.com/google/sanitizers/issues/1688. +As far as I can tell, +https://github.com/google/sanitizers/wiki/AddressSanitizerUseAfterReturn#algorithm +replaces local variables with heap allocations, and so taking the +address of a local is not effective at producing a monotonic measure of +stack usage. + +- [#6209](https://github.com/leanprover/lean4/pull/6209) documents under which conditions `Runtime.markPersistent` is +unsafe and adjusts the elaborator accordingly + +- [#6257](https://github.com/leanprover/lean4/pull/6257) harden `markPersistent` uses + +## Pretty Printing + +- [#2934](https://github.com/leanprover/lean4/pull/2934) adds the option `pp.parens` (default: false) that causes the +pretty printer to eagerly insert parentheses, which can be useful for +teaching and for understanding the structure of expressions. For +example, it causes `p → q → r` to pretty print as `p → (q → r)`. + +- [#6014](https://github.com/leanprover/lean4/pull/6014) prevents `Nat.succ ?_` from pretty printing as `?_.succ`, which +should make `apply?` be more usable. + +- [#6085](https://github.com/leanprover/lean4/pull/6085) improves the term info for coercions marked with +`CoeFnType.coeFun` (such as `DFunLike.coe` in Mathlib), making "go to +definition" on the function name work. Hovering over such a coerced +function will show the coercee rather than the coercion expression. The +coercion expression can still be seen by hovering over the whitespace in +the function application. + +- [#6096](https://github.com/leanprover/lean4/pull/6096) improves the `#print` command for structures to show all fields +and which parents the fields were inherited from, hiding internal +details such as which parents are represented as subobjects. This +information is still present in the constructor if needed. The pretty +printer for private constants is also improved, and it now handles +private names from the current module like any other name; private names +from other modules are made hygienic. + +- [#6119](https://github.com/leanprover/lean4/pull/6119) adds a new delab option `pp.coercions.types` which, when +enabled, will display all coercions with an explicit type ascription. + +- [#6161](https://github.com/leanprover/lean4/pull/6161) ensures whitespace is printed before `+opt` and `-opt` +configuration options when pretty printing, improving the experience of +tactics such as `simp?`. + +- [#6181](https://github.com/leanprover/lean4/pull/6181) fixes a bug where the signature pretty printer would ignore the +current setting of `pp.raw`. This fixes an issue where `#check ident` +would not heed `pp.raw`. Closes #6090. + +- [#6213](https://github.com/leanprover/lean4/pull/6213) exposes the difference in "synthesized type class instance is +not definitionally equal" errors. + +## Documentation + +- [#6009](https://github.com/leanprover/lean4/pull/6009) fixes a typo in the docstring for prec and makes the text +slightly more precise. + +- [#6040](https://github.com/leanprover/lean4/pull/6040) join → flatten in docstring + +- [#6110](https://github.com/leanprover/lean4/pull/6110) does some mild refactoring of the `Lean.Elab.StructInst` module +while adding documentation. + +- [#6144](https://github.com/leanprover/lean4/pull/6144) converts 3 doc-string to module docs since it seems that this is +what they were intended to be! + +- [#6150](https://github.com/leanprover/lean4/pull/6150) refine kernel code comments + +- [#6158](https://github.com/leanprover/lean4/pull/6158) adjust file reference in Data.Sum + +- [#6239](https://github.com/leanprover/lean4/pull/6239) explains the order in which `Expr.abstract` introduces de Bruijn +indices. + +## Server + +- [#5835](https://github.com/leanprover/lean4/pull/5835) adds auto-completion for the fields of structure instance notation. Specifically, querying the completions via `Ctrl+Space` in the whitespace of a structure instance notation will now bring up the full list of fields. Whitespace structure completion can be enabled for custom syntax by wrapping the parser for the list of fields in a `structInstFields` parser. + +- [#5837](https://github.com/leanprover/lean4/pull/5837) fixes an old auto-completion bug where `x.` would issue +nonsensical completions when `x.` could not be elaborated as a dot +completion. + +- [#5996](https://github.com/leanprover/lean4/pull/5996) avoid max heartbeat error in completion + +- [#6031](https://github.com/leanprover/lean4/pull/6031) fixes a regression with go-to-definition and document highlight +misbehaving on tactic blocks. + +- [#6246](https://github.com/leanprover/lean4/pull/6246) fixes a performance issue where the Lean language server would +walk the full project file tree every time a file was saved, blocking +the processing of all other requests and notifications and significantly +increasing overall language server latency after saving. + +## Lake + +- [#5684](https://github.com/leanprover/lean4/pull/5684) update toolchain on `lake update` + +- [#6026](https://github.com/leanprover/lean4/pull/6026) adds a newline at end of each Lean file generated by `lake new` +templates. + +- [#6218](https://github.com/leanprover/lean4/pull/6218) makes Lake no longer automatically fetch GitHub cloud releases +if the package build directory is already present (mirroring the +behavior of the Reservoir cache). This prevents the cache from +clobbering existing prebuilt artifacts. Users can still manually fetch +the cache and clobber the build directory by running `lake build +:release`. + +- [#6225](https://github.com/leanprover/lean4/pull/6225) makes `lake build` also eagerly print package materialization +log lines. Previously, only a `lake update` performed eager logging. + +- [#6231](https://github.com/leanprover/lean4/pull/6231) improves the errors Lake produces when it fails to fetch a +dependency from Reservoir. If the package is not indexed, it will +produce a suggestion about how to require it from GitHub. + +## Other + +- [#6137](https://github.com/leanprover/lean4/pull/6137) adds support for displaying multiple threads in the trace +profiler output. + +- [#6138](https://github.com/leanprover/lean4/pull/6138) fixes `trace.profiler.pp` not using the term pretty printer. + +- [#6259](https://github.com/leanprover/lean4/pull/6259) ensures that nesting trace nodes are annotated with timing +information iff `trace.profiler` is active. v4.14.0 ---------- diff --git a/releases_drafts/list_lex.md b/releases_drafts/list_lex.md deleted file mode 100644 index f1e0511de7d4..000000000000 --- a/releases_drafts/list_lex.md +++ /dev/null @@ -1,16 +0,0 @@ -We replace the inductive predicate `List.lt` with an upstreamed version of `List.Lex` from Mathlib. -(Previously `Lex.lt` was defined in terms of `<`; now it is generalized to take an arbitrary relation.) -This subtely changes the notion of ordering on `List α`. - -`List.lt` was a weaker relation: in particular if `l₁ < l₂`, then -`a :: l₁ < b :: l₂` may hold according to `List.lt` even if `a` and `b` are merely incomparable -(either neither `a < b` nor `b < a`), whereas according to `List.Lex` this would require `a = b`. - -When `<` is total, in the sense that `¬ · < ·` is antisymmetric, then the two relations coincide. - -Mathlib was already overriding the order instances for `List α`, -so this change should not be noticed by anyone already using Mathlib. - -We simultaneously add the boolean valued `List.lex` function, parameterised by a `BEq` typeclass -and an arbitrary `lt` function. This will support the flexibility previously provided for `List.lt`, -via a `==` function which is weaker than strict equality. From 28a70987284f403fd2d17812d39ee3fbecaa5418 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Sat, 4 Jan 2025 12:31:02 +1100 Subject: [PATCH 015/100] feat: add script for generating release notes (#6519) This PR adds a script to automatically generate release notes using the new `changelog-*` labels and "This PR ..." conventions. Usage: ``` script/release_notes.py v4.X.0 ``` where `v4.X.0` is the **previous** release, i.e. the script will process all commits *since* that tag. --- doc/dev/release_checklist.md | 27 ++----- script/release_notes.py | 145 +++++++++++++++++++++++++++++++++++ 2 files changed, 153 insertions(+), 19 deletions(-) create mode 100755 script/release_notes.py diff --git a/doc/dev/release_checklist.md b/doc/dev/release_checklist.md index 5541d5e2d51e..fc2ff4ba4ee9 100644 --- a/doc/dev/release_checklist.md +++ b/doc/dev/release_checklist.md @@ -5,11 +5,6 @@ See below for the checklist for release candidates. We'll use `v4.6.0` as the intended release version as a running example. -- One week before the planned release, ensure that - (1) someone has written the release notes and - (2) someone has written the first draft of the release blog post. - If there is any material in `./releases_drafts/` on the `releases/v4.6.0` branch, then the release notes are not done. - (See the section "Writing the release notes".) - `git checkout releases/v4.6.0` (This branch should already exist, from the release candidates.) - `git pull` @@ -86,7 +81,7 @@ We'll use `v4.6.0` as the intended release version as a running example. - Toolchain bump PR notes: - In addition to updating the `lean-toolchain` and `lakefile.lean`, in `.github/workflows/lean4checker.yml` update the line - `git checkout v4.6.0` to the appropriate tag. + `git checkout v4.6.0` to the appropriate tag. - Push the PR branch to the main Mathlib repository rather than a fork, or CI may not work reliably - Create and push the tag - Create a new branch from the tag, push it, and open a pull request against `stable`. @@ -139,16 +134,13 @@ We'll use `v4.7.0-rc1` as the intended release version in this example. git checkout -b releases/v4.7.0 ``` - In `RELEASES.md` replace `Development in progress` in the `v4.7.0` section with `Release notes to be written.` -- We will rely on automatically generated release notes for release candidates, - and the written release notes will be used for stable versions only. - It is essential to choose the nightly that will become the release candidate as early as possible, to avoid confusion. +- It is essential to choose the nightly that will become the release candidate as early as possible, to avoid confusion. - In `src/CMakeLists.txt`, - verify that you see `set(LEAN_VERSION_MINOR 7)` (for whichever `7` is appropriate); this should already have been updated when the development cycle began. - `set(LEAN_VERSION_IS_RELEASE 1)` (this should be a change; on `master` and nightly releases it is always `0`). - Commit your changes to `src/CMakeLists.txt`, and push. - `git tag v4.7.0-rc1` - `git push origin v4.7.0-rc1` -- Ping the FRO Zulip that release notes need to be written. The release notes do not block completing the rest of this checklist. - Now wait, while CI runs. - You can monitor this at `https://github.com/leanprover/lean4/actions/workflows/ci.yml`, looking for the `v4.7.0-rc1` tag. - This step can take up to an hour. @@ -248,15 +240,12 @@ Please read https://leanprover-community.github.io/contribute/tags_and_branches. # Writing the release notes -We are currently trying a system where release notes are compiled all at once from someone looking through the commit history. -The exact steps are a work in progress. -Here is the general idea: +Release notes are automatically generated from the commit history, using `script/release_notes.py`. -* The work is done right on the `releases/v4.6.0` branch sometime after it is created but before the stable release is made. - The release notes for `v4.6.0` will later be copied to `master` when we begin a new development cycle. -* There can be material for release notes entries in commit messages. -* There can also be pre-written entries in `./releases_drafts`, which should be all incorporated in the release notes and then deleted from the branch. +Run this as `script/release_notes.py v4.6.0`, where `v4.6.0` is the *previous* release version. This will generate output +for all commits since that tag. Note that there is output on both stderr, which should be manually reviewed, +and on stdout, which should be manually copied to `RELEASES.md`. + +There can also be pre-written entries in `./releases_drafts`, which should be all incorporated in the release notes and then deleted from the branch. See `./releases_drafts/README.md` for more information. -* The release notes should be written from a downstream expert user's point of view. -This section will be updated when the next release notes are written (for `v4.10.0`). diff --git a/script/release_notes.py b/script/release_notes.py new file mode 100755 index 000000000000..d8bb03ca5d92 --- /dev/null +++ b/script/release_notes.py @@ -0,0 +1,145 @@ +#!/usr/bin/env python3 + +import sys +import re +import json +import requests +import subprocess +from collections import defaultdict +from git import Repo + +def get_commits_since_tag(repo, tag): + try: + tag_commit = repo.commit(tag) + commits = list(repo.iter_commits(f"{tag_commit.hexsha}..HEAD")) + return [ + (commit.hexsha, commit.message.splitlines()[0], commit.message) + for commit in commits + ] + except Exception as e: + sys.stderr.write(f"Error retrieving commits: {e}\n") + sys.exit(1) + +def check_pr_number(first_line): + match = re.search(r"\(\#(\d+)\)$", first_line) + if match: + return int(match.group(1)) + return None + +def fetch_pr_labels(pr_number): + try: + # Use gh CLI to fetch PR details + result = subprocess.run([ + "gh", "api", f"repos/leanprover/lean4/pulls/{pr_number}" + ], capture_output=True, text=True, check=True) + pr_data = result.stdout + pr_json = json.loads(pr_data) + return [label["name"] for label in pr_json.get("labels", [])] + except subprocess.CalledProcessError as e: + sys.stderr.write(f"Failed to fetch PR #{pr_number} using gh: {e.stderr}\n") + return [] + +def format_section_title(label): + title = label.replace("changelog-", "").capitalize() + if title == "Doc": + return "Documentation" + elif title == "Pp": + return "Pretty Printing" + return title + +def sort_sections_order(): + return [ + "Language", + "Library", + "Compiler", + "Pretty Printing", + "Documentation", + "Server", + "Lake", + "Other", + "Uncategorised" + ] + +def format_markdown_description(pr_number, description): + link = f"[#{pr_number}](https://github.com/leanprover/lean4/pull/{pr_number})" + return f"{link} {description}" + +def main(): + if len(sys.argv) != 2: + sys.stderr.write("Usage: script.py \n") + sys.exit(1) + + tag = sys.argv[1] + try: + repo = Repo(".") + except Exception as e: + sys.stderr.write(f"Error opening Git repository: {e}\n") + sys.exit(1) + + commits = get_commits_since_tag(repo, tag) + + sys.stderr.write(f"Found {len(commits)} commits since tag {tag}:\n") + for commit_hash, first_line, _ in commits: + sys.stderr.write(f"- {commit_hash}: {first_line}\n") + + changelog = defaultdict(list) + + for commit_hash, first_line, full_message in commits: + # Skip commits with the specific first lines + if first_line == "chore: update stage0" or first_line.startswith("chore: CI: bump "): + continue + + pr_number = check_pr_number(first_line) + + if not pr_number: + sys.stderr.write(f"No PR number found in {first_line}\n") + continue + + # Remove the first line from the full_message for further processing + body = full_message[len(first_line):].strip() + + paragraphs = body.split('\n\n') + second_paragraph = paragraphs[0] if len(paragraphs) > 0 else "" + + labels = fetch_pr_labels(pr_number) + + # Skip entries with the "changelog-no" label + if "changelog-no" in labels: + continue + + report_errors = first_line.startswith("feat:") or first_line.startswith("fix:") + + if not second_paragraph.startswith("This PR "): + if report_errors: + sys.stderr.write(f"No PR description found in commit:\n{commit_hash}\n{first_line}\n{body}\n\n") + fallback_description = re.sub(r":$", "", first_line.split(" ", 1)[1]).rsplit(" (#", 1)[0] + markdown_description = format_markdown_description(pr_number, fallback_description) + else: + continue + else: + markdown_description = format_markdown_description(pr_number, second_paragraph.replace("This PR ", "")) + + changelog_labels = [label for label in labels if label.startswith("changelog-")] + if len(changelog_labels) > 1: + sys.stderr.write(f"Warning: Multiple changelog-* labels found for PR #{pr_number}: {changelog_labels}\n") + + if not changelog_labels: + if report_errors: + sys.stderr.write(f"Warning: No changelog-* label found for PR #{pr_number}\n") + else: + continue + + for label in changelog_labels: + changelog[label].append((pr_number, markdown_description)) + + section_order = sort_sections_order() + sorted_changelog = sorted(changelog.items(), key=lambda item: section_order.index(format_section_title(item[0])) if format_section_title(item[0]) in section_order else len(section_order)) + + for label, entries in sorted_changelog: + section_title = format_section_title(label) if label != "Uncategorised" else "Uncategorised" + print(f"## {section_title}\n") + for _, entry in sorted(entries, key=lambda x: x[0]): + print(f"* {entry}\n") + +if __name__ == "__main__": + main() From ad593b36d9eeac954a6b1810385d3707bada87a7 Mon Sep 17 00:00:00 2001 From: Leonardo de Moura Date: Sat, 4 Jan 2025 03:18:43 +0100 Subject: [PATCH 016/100] feat: add support for `match`-expressions to `grind` (#6521) This PR adds support for activating relevant `match`-equations as E-matching theorems. It uses the `match`-equation lhs as the pattern. --- src/Init/Grind/Norm.lean | 3 +- src/Init/Grind/Tactics.lean | 2 + src/Init/Grind/Util.lean | 7 ++ src/Lean/Meta/Tactic/Grind/DoNotSimp.lean | 35 ++++++++ src/Lean/Meta/Tactic/Grind/EMatch.lean | 18 +++- src/Lean/Meta/Tactic/Grind/EMatchTheorem.lean | 88 ++++++++++++++++--- src/Lean/Meta/Tactic/Grind/Internalize.lean | 36 ++++++-- src/Lean/Meta/Tactic/Grind/Main.lean | 3 +- src/Lean/Meta/Tactic/Grind/Simp.lean | 2 + src/Lean/Meta/Tactic/Grind/Types.lean | 2 + tests/lean/run/grind_match1.lean | 56 ++++++++++++ 11 files changed, 224 insertions(+), 28 deletions(-) create mode 100644 src/Lean/Meta/Tactic/Grind/DoNotSimp.lean create mode 100644 tests/lean/run/grind_match1.lean diff --git a/src/Init/Grind/Norm.lean b/src/Init/Grind/Norm.lean index c25e02db1df4..4bf087bdb84e 100644 --- a/src/Init/Grind/Norm.lean +++ b/src/Init/Grind/Norm.lean @@ -5,6 +5,7 @@ Authors: Leonardo de Moura -/ prelude import Init.SimpLemmas +import Init.PropLemmas import Init.Classical import Init.ByCases @@ -64,7 +65,7 @@ attribute [grind_norm] forall_and -- Exists @[grind_norm↓] theorem not_exists (p : α → Prop) : (¬∃ x, p x) = ∀ x, ¬p x := by simp -attribute [grind_norm] exists_const exists_or +attribute [grind_norm] exists_const exists_or exists_prop exists_and_left exists_and_right -- Bool cond @[grind_norm] theorem cond_eq_ite (c : Bool) (a b : α) : cond c a b = ite c a b := by diff --git a/src/Init/Grind/Tactics.lean b/src/Init/Grind/Tactics.lean index 1554e8b0a466..71b9455a93a2 100644 --- a/src/Init/Grind/Tactics.lean +++ b/src/Init/Grind/Tactics.lean @@ -28,6 +28,8 @@ structure Config where gen : Nat := 5 /-- Maximum number of theorem instances generated using E-matching in a proof search tree branch. -/ instances : Nat := 1000 + /-- If `matchEqs` is `true`, `grind` uses `match`-equations as E-matching theorems. -/ + matchEqs : Bool := true deriving Inhabited, BEq end Lean.Grind diff --git a/src/Init/Grind/Util.lean b/src/Init/Grind/Util.lean index f445309de85a..9c7523d992cb 100644 --- a/src/Init/Grind/Util.lean +++ b/src/Init/Grind/Util.lean @@ -11,6 +11,13 @@ namespace Lean.Grind /-- A helper gadget for annotating nested proofs in goals. -/ def nestedProof (p : Prop) (h : p) : p := h +/-- +Gadget for marking terms that should not be normalized by `grind`s simplifier. +`grind` uses a simproc to implement this feature. +We use it when adding instances of `match`-equations to prevent them from being simplified to true. +-/ +def doNotSimp {α : Sort u} (a : α) : α := a + set_option pp.proofs true theorem nestedProof_congr (p q : Prop) (h : p = q) (hp : p) (hq : q) : HEq (nestedProof p hp) (nestedProof q hq) := by diff --git a/src/Lean/Meta/Tactic/Grind/DoNotSimp.lean b/src/Lean/Meta/Tactic/Grind/DoNotSimp.lean new file mode 100644 index 000000000000..9c647aa37ff2 --- /dev/null +++ b/src/Lean/Meta/Tactic/Grind/DoNotSimp.lean @@ -0,0 +1,35 @@ +/- +Copyright (c) 2025 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 Init.Grind.Util +import Init.Simproc +import Lean.Meta.Tactic.Simp.Simproc + +namespace Lean.Meta.Grind + +/-- +Returns `Grind.doNotSimp e`. +Recall that `Grind.doNotSimp` is an identity function, but the following simproc is used to prevent the term `e` from being simplified. +-/ +def markAsDoNotSimp (e : Expr) : MetaM Expr := + mkAppM ``Grind.doNotSimp #[e] + +builtin_dsimproc_decl reduceDoNotSimp (Grind.doNotSimp _) := fun e => do + let_expr Grind.doNotSimp _ _ ← e | return .continue + return .done e + +/-- Adds `reduceDoNotSimp` to `s` -/ +def addDoNotSimp (s : Simprocs) : CoreM Simprocs := do + s.add ``reduceDoNotSimp (post := false) + +/-- Erases `Grind.doNotSimp` annotations. -/ +def eraseDoNotSimp (e : Expr) : CoreM Expr := do + let pre (e : Expr) := do + let_expr Grind.doNotSimp _ a := e | return .continue e + return .continue a + Core.transform e (pre := pre) + +end Lean.Meta.Grind diff --git a/src/Lean/Meta/Tactic/Grind/EMatch.lean b/src/Lean/Meta/Tactic/Grind/EMatch.lean index 1d19b860ba05..42031260a0a7 100644 --- a/src/Lean/Meta/Tactic/Grind/EMatch.lean +++ b/src/Lean/Meta/Tactic/Grind/EMatch.lean @@ -6,6 +6,7 @@ Authors: Leonardo de Moura prelude import Lean.Meta.Tactic.Grind.Types import Lean.Meta.Tactic.Grind.Intro +import Lean.Meta.Tactic.Grind.DoNotSimp namespace Lean.Meta.Grind namespace EMatch @@ -146,6 +147,15 @@ private def processContinue (c : Choice) (p : Expr) : M Unit := do let c := { c with gen := Nat.max gen c.gen } modify fun s => { s with choiceStack := c :: s.choiceStack } +/-- Helper function for marking parts of `match`-equation theorem as "do-not-simplify" -/ +private partial def annotateMatchEqnType (prop : Expr) : M Expr := do + if let .forallE n d b bi := prop then + withLocalDecl n bi (← markAsDoNotSimp d) fun x => do + mkForallFVars #[x] (← annotateMatchEqnType (b.instantiate1 x)) + else + let_expr f@Eq α lhs rhs := prop | return prop + return mkApp3 f α (← markAsDoNotSimp lhs) rhs + /-- Stores new theorem instance in the state. Recall that new instances are internalized later, after a full round of ematching. @@ -154,7 +164,9 @@ private def addNewInstance (origin : Origin) (proof : Expr) (generation : Nat) : let proof ← instantiateMVars proof if grind.debug.proofs.get (← getOptions) then check proof - let prop ← inferType proof + let mut prop ← inferType proof + if Match.isMatchEqnTheorem (← getEnv) origin.key then + prop ← annotateMatchEqnType prop trace[grind.ematch.instance] "{← origin.pp}: {prop}" addTheoremInstance proof prop generation @@ -189,10 +201,10 @@ private partial def instantiateTheorem (c : Choice) : M Unit := withDefault do w unless (← synthesizeInstance mvar type) do trace[grind.issues] "failed to synthesize instance when instantiating {← thm.origin.pp}{indentExpr type}" return () + let proof := mkAppN proof mvars if (← mvars.allM (·.mvarId!.isAssigned)) then - addNewInstance thm.origin (mkAppN proof mvars) c.gen + addNewInstance thm.origin proof c.gen else - let proof := mkAppN proof mvars let mvars ← mvars.filterM fun mvar => return !(← mvar.mvarId!.isAssigned) if let some mvarBad ← mvars.findM? fun mvar => return !(← isProof mvar) then trace[grind.issues] "failed to instantiate {← thm.origin.pp}, failed to instantiate non propositional argument with type{indentExpr (← inferType mvarBad)}" diff --git a/src/Lean/Meta/Tactic/Grind/EMatchTheorem.lean b/src/Lean/Meta/Tactic/Grind/EMatchTheorem.lean index f8626bef0a7c..4b03d9021fe5 100644 --- a/src/Lean/Meta/Tactic/Grind/EMatchTheorem.lean +++ b/src/Lean/Meta/Tactic/Grind/EMatchTheorem.lean @@ -56,17 +56,47 @@ structure EMatchTheorem where origin : Origin deriving Inhabited -/-- The key is a symbol from `EMatchTheorem.symbols`. -/ -abbrev EMatchTheorems := PHashMap Name (List EMatchTheorem) +/-- Set of E-matching theorems. -/ +structure EMatchTheorems where + /-- The key is a symbol from `EMatchTheorem.symbols`. -/ + private map : PHashMap Name (List EMatchTheorem) := {} + /-- Set of theorem names that have been inserted using `insert`. -/ + private thmNames : PHashSet Name := {} + deriving Inhabited +/-- +Inserts a `thm` with symbols `[s_1, ..., s_n]` to `s`. +We add `s_1 -> { thm with symbols := [s_2, ..., s_n] }`. +When `grind` internalizes a term containing symbol `s`, we +process all theorems `thm` associated with key `s`. +If their `thm.symbols` is empty, we say they are activated. +Otherwise, we reinsert into `map`. +-/ def EMatchTheorems.insert (s : EMatchTheorems) (thm : EMatchTheorem) : EMatchTheorems := Id.run do let .const declName :: syms := thm.symbols | unreachable! let thm := { thm with symbols := syms } - if let some thms := s.find? declName then - return PersistentHashMap.insert s declName (thm::thms) + let { map, thmNames } := s + let thmNames := thmNames.insert thm.origin.key + if let some thms := map.find? declName then + return { map := map.insert declName (thm::thms), thmNames } else - return PersistentHashMap.insert s declName [thm] + return { map := map.insert declName [thm], thmNames } + +/-- +Retrieves theorems from `s` associated with the given symbol. See `EMatchTheorem.insert`. +The theorems are removed from `s`. +-/ +@[inline] +def EMatchTheorems.retrieve? (s : EMatchTheorems) (sym : Name) : Option (List EMatchTheorem × EMatchTheorems) := + if let some thms := s.map.find? sym then + some (thms, { s with map := s.map.erase sym }) + else + none + +/-- Returns `true` if `declName` is the name of a theorem that was inserted using `insert`. -/ +def EMatchTheorems.containsTheoremName (s : EMatchTheorems) (declName : Name) : Bool := + s.thmNames.contains declName def EMatchTheorem.getProofWithFreshMVarLevels (thm : EMatchTheorem) : MetaM Expr := do if thm.proof.isConst && thm.levelParams.isEmpty then @@ -85,7 +115,7 @@ def EMatchTheorem.getProofWithFreshMVarLevels (thm : EMatchTheorem) : MetaM Expr private builtin_initialize ematchTheoremsExt : SimpleScopedEnvExtension EMatchTheorem EMatchTheorems ← registerSimpleScopedEnvExtension { addEntry := EMatchTheorems.insert - initial := .empty + initial := {} } -- TODO: create attribute? @@ -320,8 +350,8 @@ private def checkCoverage (thmProof : Expr) (numParams : Nat) (bvarsFound : Std. Given a theorem with proof `proof` and `numParams` parameters, returns a message containing the parameters at positions `paramPos`. -/ -private def ppParamsAt (proof : Expr) (numParms : Nat) (paramPos : List Nat) : MetaM MessageData := do - forallBoundedTelescope (← inferType proof) numParms fun xs _ => do +private def ppParamsAt (proof : Expr) (numParams : Nat) (paramPos : List Nat) : MetaM MessageData := do + forallBoundedTelescope (← inferType proof) numParams fun xs _ => do let mut msg := m!"" let mut first := true for h : i in [:xs.size] do @@ -331,23 +361,53 @@ private def ppParamsAt (proof : Expr) (numParms : Nat) (paramPos : List Nat) : M msg := msg ++ m!"{x} : {← inferType x}" addMessageContextFull msg -def addEMatchTheorem (declName : Name) (numParams : Nat) (patterns : List Expr) : MetaM Unit := do +/-- +Creates an E-matching theorem for `declName` with `numParams` parameters, and the given set of patterns. +Pattern variables are represented using de Bruijn indices. +-/ +def mkEMatchTheorem (declName : Name) (numParams : Nat) (patterns : List Expr) : MetaM EMatchTheorem := do let .thmInfo info ← getConstInfo declName | throwError "`{declName}` is not a theorem, you cannot assign patterns to non-theorems for the `grind` tactic" let us := info.levelParams.map mkLevelParam let proof := mkConst declName us let (patterns, symbols, bvarFound) ← NormalizePattern.main patterns assert! symbols.all fun s => s matches .const _ - trace[grind.ematch.pattern] "{declName}: {patterns.map ppPattern}" + trace[grind.ematch.pattern] "{MessageData.ofConst proof}: {patterns.map ppPattern}" if let .missing pos ← checkCoverage proof numParams bvarFound then let pats : MessageData := m!"{patterns.map ppPattern}" throwError "invalid pattern(s) for `{declName}`{indentD pats}\nthe following theorem parameters cannot be instantiated:{indentD (← ppParamsAt proof numParams pos)}" - ematchTheoremsExt.add { - proof, patterns, numParams, symbols - levelParams := #[] - origin := .decl declName + return { + proof, patterns, numParams, symbols + levelParams := #[] + origin := .decl declName } +/-- +Given theorem with name `declName` and type of the form `∀ (a_1 ... a_n), lhs = rhs`, +creates an E-matching pattern for it using `addEMatchTheorem n [lhs]` +-/ +def mkEMatchEqTheorem (declName : Name) : MetaM EMatchTheorem := do + let info ← getConstInfo declName + let (numParams, patterns) ← forallTelescopeReducing info.type fun xs type => do + let_expr Eq _ lhs _ := type | throwError "invalid E-matching equality theorem, conclusion must be an equality{indentExpr type}" + return (xs.size, [lhs.abstract xs]) + mkEMatchTheorem declName numParams patterns + +/-- +Adds an E-matching theorem to the environment. +See `mkEMatchTheorem`. +-/ +def addEMatchTheorem (declName : Name) (numParams : Nat) (patterns : List Expr) : MetaM Unit := do + ematchTheoremsExt.add (← mkEMatchTheorem declName numParams patterns) + +/-- +Adds an E-matching equality theorem to the environment. +See `mkEMatchEqTheorem`. +-/ +def addEMatchEqTheorem (declName : Name) : MetaM Unit := do + ematchTheoremsExt.add (← mkEMatchEqTheorem declName) + +/-- Returns the E-matching theorems registered in the environment. -/ def getEMatchTheorems : CoreM EMatchTheorems := return ematchTheoremsExt.getState (← getEnv) diff --git a/src/Lean/Meta/Tactic/Grind/Internalize.lean b/src/Lean/Meta/Tactic/Grind/Internalize.lean index 6231923227bf..63c2fc81ccae 100644 --- a/src/Lean/Meta/Tactic/Grind/Internalize.lean +++ b/src/Lean/Meta/Tactic/Grind/Internalize.lean @@ -6,6 +6,8 @@ Authors: Leonardo de Moura prelude import Init.Grind.Util import Lean.Meta.LitValues +import Lean.Meta.Match.MatcherInfo +import Lean.Meta.Match.MatchEqsExt import Lean.Meta.Tactic.Grind.Types import Lean.Meta.Tactic.Grind.Util @@ -50,21 +52,36 @@ private partial def internalizePattern (pattern : Expr) (generation : Nat) : Goa else pattern.withApp fun f args => do return mkAppN f (← args.mapM (internalizePattern · generation)) +private partial def activateTheorem (thm : EMatchTheorem) (generation : Nat) : GoalM Unit := do + -- Recall that we use the proof as part of the key for a set of instances found so far. + -- We don't want to use structural equality when comparing keys. + let proof ← shareCommon thm.proof + let thm := { thm with proof, patterns := (← thm.patterns.mapM (internalizePattern · generation)) } + trace[grind.ematch] "activated `{thm.origin.key}`, {thm.patterns.map ppPattern}" + modify fun s => { s with newThms := s.newThms.push thm } + +/-- +If `Config.matchEqs` is set to `true`, and `f` is `match`-auxiliary function, +adds its equations to `newThms`. +-/ +private partial def addMatchEqns (f : Expr) (generation : Nat) : GoalM Unit := do + if !(← getConfig).matchEqs then return () + let .const declName _ := f | return () + if !(← isMatcher declName) then return () + if (← get).matchEqNames.contains declName then return () + modify fun s => { s with matchEqNames := s.matchEqNames.insert declName } + for eqn in (← Match.getEquationsFor declName).eqnNames do + activateTheorem (← mkEMatchEqTheorem eqn) generation + private partial def activateTheoremPatterns (fName : Name) (generation : Nat) : GoalM Unit := do - if let some thms := (← get).thmMap.find? fName then - modify fun s => { s with thmMap := s.thmMap.erase fName } + if let some (thms, thmMap) := (← get).thmMap.retrieve? fName then + modify fun s => { s with thmMap } let appMap := (← get).appMap for thm in thms do let symbols := thm.symbols.filter fun sym => !appMap.contains sym let thm := { thm with symbols } match symbols with - | [] => - -- Recall that we use the proof as part of the key for a set of instances found so far. - -- We don't want to use structural equality when comparing keys. - let proof ← shareCommon thm.proof - let thm := { thm with proof, patterns := (← thm.patterns.mapM (internalizePattern · generation)) } - trace[grind.ematch] "activated `{thm.origin.key}`, {thm.patterns.map ppPattern}" - modify fun s => { s with newThms := s.newThms.push thm } + | [] => activateTheorem thm generation | _ => trace[grind.ematch] "reinsert `{thm.origin.key}`" modify fun s => { s with thmMap := s.thmMap.insert thm } @@ -95,6 +112,7 @@ partial def internalize (e : Expr) (generation : Nat) : GoalM Unit := do -- We do not want to internalize the components of a literal value. mkENode e generation else e.withApp fun f args => do + addMatchEqns f generation if f.isConstOf ``Lean.Grind.nestedProof && args.size == 2 then -- We only internalize the proposition. We can skip the proof because of -- proof irrelevance diff --git a/src/Lean/Meta/Tactic/Grind/Main.lean b/src/Lean/Meta/Tactic/Grind/Main.lean index 77a7b4ffa7bf..98214895ee8d 100644 --- a/src/Lean/Meta/Tactic/Grind/Main.lean +++ b/src/Lean/Meta/Tactic/Grind/Main.lean @@ -14,6 +14,7 @@ import Lean.Meta.Tactic.Grind.Util import Lean.Meta.Tactic.Grind.Inv import Lean.Meta.Tactic.Grind.Intro import Lean.Meta.Tactic.Grind.EMatch +import Lean.Meta.Tactic.Grind.DoNotSimp namespace Lean.Meta.Grind @@ -38,7 +39,7 @@ def GrindM.run (x : GrindM α) (mainDeclName : Name) (config : Grind.Config) (fa let (falseExpr, scState) := ShareCommon.State.shareCommon scState (mkConst ``False) let (trueExpr, scState) := ShareCommon.State.shareCommon scState (mkConst ``True) let thms ← grindNormExt.getTheorems - let simprocs := #[(← grindNormSimprocExt.getSimprocs)] + let simprocs := #[(← addDoNotSimp (← grindNormSimprocExt.getSimprocs))] let simp ← Simp.mkContext (config := { arith := true }) (simpTheorems := #[thms]) diff --git a/src/Lean/Meta/Tactic/Grind/Simp.lean b/src/Lean/Meta/Tactic/Grind/Simp.lean index 1a71d3831814..7a4d7c2d558d 100644 --- a/src/Lean/Meta/Tactic/Grind/Simp.lean +++ b/src/Lean/Meta/Tactic/Grind/Simp.lean @@ -9,6 +9,7 @@ import Lean.Meta.Tactic.Assert import Lean.Meta.Tactic.Simp.Main import Lean.Meta.Tactic.Grind.Util import Lean.Meta.Tactic.Grind.Types +import Lean.Meta.Tactic.Grind.DoNotSimp import Lean.Meta.Tactic.Grind.MarkNestedProofs namespace Lean.Meta.Grind @@ -33,6 +34,7 @@ def simp (e : Expr) : GrindM Simp.Result := do let e' ← eraseIrrelevantMData e' let e' ← foldProjs e' let e' ← normalizeLevels e' + let e' ← eraseDoNotSimp e' let e' ← canon e' let e' ← shareCommon e' trace[grind.simp] "{e}\n===>\n{e'}" diff --git a/src/Lean/Meta/Tactic/Grind/Types.lean b/src/Lean/Meta/Tactic/Grind/Types.lean index 5ac73e6c4dff..92f7cd363293 100644 --- a/src/Lean/Meta/Tactic/Grind/Types.lean +++ b/src/Lean/Meta/Tactic/Grind/Types.lean @@ -379,6 +379,8 @@ structure Goal where preInstances : PreInstanceSet := {} /-- new facts to be processed. -/ newFacts : Std.Queue NewFact := ∅ + /-- `match` auxiliary functions whose equations have already been created and activated. -/ + matchEqNames : PHashSet Name := {} deriving Inhabited def Goal.admit (goal : Goal) : MetaM Unit := diff --git a/tests/lean/run/grind_match1.lean b/tests/lean/run/grind_match1.lean new file mode 100644 index 000000000000..483b8600c05c --- /dev/null +++ b/tests/lean/run/grind_match1.lean @@ -0,0 +1,56 @@ +def g (xs : List α) (ys : List α) := + match xs, ys with + | [], _ => ys + | _::_::_, [ ] => [] + | x::xs, ys => x :: g xs ys + +attribute [simp] g + +set_option trace.grind.assert true + +/-- +info: [grind.assert] (match as, bs with + | [], x => bs + | head :: head_1 :: tail, [] => [] + | x :: xs, ys => x :: g xs ys) = + d +[grind.assert] bs = [] +[grind.assert] a₁ :: f 0 = as +[grind.assert] f 0 = a₂ :: f 1 +[grind.assert] ¬d = [] +[grind.assert] (match a₁ :: a₂ :: f 1, [] with + | [], x => bs + | head :: head_1 :: tail, [] => [] + | x :: xs, ys => x :: g xs ys) = + [] +-/ +#guard_msgs (info) in +example (f : Nat → List Nat) : g as bs = d → bs = [] → a₁ :: f 0 = as → f 0 = a₂ :: f 1 → d = [] := by + unfold g + grind + + +example : g as bs = d → as = [] → d = bs := by + unfold g + grind + +def f (x : List α) : Bool := + match x with + | [] => true + | _::_ => false + +example : f a = b → a = [] → b = true := by + unfold f + grind + +def f' (x : List Nat) : Bool := + match x with + | [] => true + | _::_ => false + +attribute [simp] f' +#check f'.match_1.eq_1 + +example : f' a = b → a = [] → b = true := by + unfold f' + grind From d2189542b5a29393fe5fc244b86fb64d1a43baa5 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Sat, 4 Jan 2025 15:04:13 +1100 Subject: [PATCH 017/100] chore: upstream some List.Perm lemmas (#6524) This PR upstreams some remaining `List.Perm` lemmas from Batteries. --- src/Init/Data/List/Perm.lean | 14 ++++++++++++++ src/Init/Data/List/Sort/Lemmas.lean | 4 ++++ 2 files changed, 18 insertions(+) diff --git a/src/Init/Data/List/Perm.lean b/src/Init/Data/List/Perm.lean index 501d02daf5ae..5eef4cc0367e 100644 --- a/src/Init/Data/List/Perm.lean +++ b/src/Init/Data/List/Perm.lean @@ -510,4 +510,18 @@ theorem Perm.eraseP (f : α → Bool) {l₁ l₂ : List α} refine (IH₁ H).trans (IH₂ ((p₁.pairwise_iff ?_).1 H)) exact fun h h₁ h₂ => h h₂ h₁ +theorem perm_insertIdx [DecidableEq α] {α} (x : α) (l : List α) {n} (h : n ≤ l.length) : + insertIdx n x l ~ x :: l := by + induction l generalizing n with + | nil => + cases n with + | zero => rfl + | succ => cases h + | cons _ _ ih => + cases n with + | zero => simp [insertIdx] + | succ => + simp only [insertIdx, modifyTailIdx] + refine .trans (.cons _ (ih (Nat.le_of_succ_le_succ h))) (.swap ..) + end List diff --git a/src/Init/Data/List/Sort/Lemmas.lean b/src/Init/Data/List/Sort/Lemmas.lean index 9bb216e5bfa9..a8e160fe5164 100644 --- a/src/Init/Data/List/Sort/Lemmas.lean +++ b/src/Init/Data/List/Sort/Lemmas.lean @@ -253,6 +253,10 @@ theorem merge_perm_append : ∀ {xs ys : List α}, merge xs ys le ~ xs ++ ys · exact (merge_perm_append.cons y).trans ((Perm.swap x y _).trans (perm_middle.symm.cons x)) +theorem Perm.merge (s₁ s₂ : α → α → Bool) (hl : l₁ ~ l₂) (hr : r₁ ~ r₂) : + merge l₁ r₁ s₁ ~ merge l₂ r₂ s₂ := + Perm.trans (merge_perm_append ..) <| Perm.trans (Perm.append hl hr) <| Perm.symm (merge_perm_append ..) + /-! ### mergeSort -/ @[simp] theorem mergeSort_nil : [].mergeSort r = [] := by rw [List.mergeSort] From cdeb958afd4ccf1f87941f32c0e5291bed1495a4 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Sat, 4 Jan 2025 15:08:21 +1100 Subject: [PATCH 018/100] chore: add plausible to release checklist (#6525) --- doc/dev/release_checklist.md | 4 ++++ 1 file changed, 4 insertions(+) diff --git a/doc/dev/release_checklist.md b/doc/dev/release_checklist.md index fc2ff4ba4ee9..e5b27e203012 100644 --- a/doc/dev/release_checklist.md +++ b/doc/dev/release_checklist.md @@ -76,6 +76,10 @@ We'll use `v4.6.0` as the intended release version as a running example. - Toolchain bump PR including updated Lake manifest - Create and push the tag - There is no `stable` branch; skip this step + - [plausible](https://github.com/leanprover-community/plausible) + - Toolchain bump PR including updated Lake manifest + - Create and push the tag + - There is no `stable` branch; skip this step - [Mathlib](https://github.com/leanprover-community/mathlib4) - Dependencies: `Aesop`, `ProofWidgets4`, `lean4checker`, `Batteries`, `doc-gen4`, `import-graph` - Toolchain bump PR notes: From 9080df3110bd680956237c82dbbe96e1c818992f Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Sat, 4 Jan 2025 15:13:13 +1100 Subject: [PATCH 019/100] chore: import cleanup in Init (#6522) This PR avoids unnecessarily importing "kitchen sink" files. --- src/Lean/Data/Lsp/Utf16.lean | 1 - src/Lean/Linter/MissingDocs.lean | 1 + src/Lean/PrettyPrinter/Delaborator/Builtins.lean | 1 - src/Lean/Server/Completion.lean | 1 + src/Lean/Util/ShareCommon.lean | 4 ++-- src/Std/Sat/AIG/Basic.lean | 1 - .../Tactic/BVDecide/LRAT/Internal/Formula/Implementation.lean | 1 - src/Std/Time/DateTime/Timestamp.lean | 1 - src/Std/Time/Internal/Bounded.lean | 2 +- 9 files changed, 5 insertions(+), 8 deletions(-) diff --git a/src/Lean/Data/Lsp/Utf16.lean b/src/Lean/Data/Lsp/Utf16.lean index b9d5bea8ba80..30d36e8a478a 100644 --- a/src/Lean/Data/Lsp/Utf16.lean +++ b/src/Lean/Data/Lsp/Utf16.lean @@ -6,7 +6,6 @@ Authors: Marc Huisinga, Wojciech Nawrocki -/ prelude import Init.Data.String -import Init.Data.Array import Lean.Data.Lsp.Basic import Lean.Data.Position import Lean.DeclarationRange diff --git a/src/Lean/Linter/MissingDocs.lean b/src/Lean/Linter/MissingDocs.lean index 6df6d7473f90..742ad05f8b96 100644 --- a/src/Lean/Linter/MissingDocs.lean +++ b/src/Lean/Linter/MissingDocs.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ prelude +import Lean.Parser.Syntax import Lean.Meta.Tactic.Simp.RegisterCommand import Lean.Elab.Command import Lean.Elab.SetOption diff --git a/src/Lean/PrettyPrinter/Delaborator/Builtins.lean b/src/Lean/PrettyPrinter/Delaborator/Builtins.lean index 5dc3aa3ee97f..61d9e209f42e 100644 --- a/src/Lean/PrettyPrinter/Delaborator/Builtins.lean +++ b/src/Lean/PrettyPrinter/Delaborator/Builtins.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sebastian Ullrich, Leonardo de Moura, Gabriel Ebner, Mario Carneiro -/ prelude -import Lean.Parser import Lean.PrettyPrinter.Delaborator.Attributes import Lean.PrettyPrinter.Delaborator.Basic import Lean.PrettyPrinter.Delaborator.SubExpr diff --git a/src/Lean/Server/Completion.lean b/src/Lean/Server/Completion.lean index aec3001948f7..9ffb290958a4 100644 --- a/src/Lean/Server/Completion.lean +++ b/src/Lean/Server/Completion.lean @@ -5,6 +5,7 @@ Authors: Leonardo de Moura, Marc Huisinga -/ prelude import Lean.Server.Completion.CompletionCollectors +import Std.Data.HashMap namespace Lean.Server.Completion open Lsp diff --git a/src/Lean/Util/ShareCommon.lean b/src/Lean/Util/ShareCommon.lean index 4bc3bc7b694d..3a508533fd5b 100644 --- a/src/Lean/Util/ShareCommon.lean +++ b/src/Lean/Util/ShareCommon.lean @@ -5,8 +5,8 @@ Authors: Leonardo de Moura -/ prelude import Init.ShareCommon -import Std.Data.HashSet -import Std.Data.HashMap +import Std.Data.HashSet.Basic +import Std.Data.HashMap.Basic import Lean.Data.PersistentHashMap import Lean.Data.PersistentHashSet diff --git a/src/Std/Sat/AIG/Basic.lean b/src/Std/Sat/AIG/Basic.lean index 4a12ef54dce9..f68ce1a3c9e4 100644 --- a/src/Std/Sat/AIG/Basic.lean +++ b/src/Std/Sat/AIG/Basic.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Henrik Böving -/ prelude -import Std.Data.HashMap import Std.Data.HashSet namespace Std diff --git a/src/Std/Tactic/BVDecide/LRAT/Internal/Formula/Implementation.lean b/src/Std/Tactic/BVDecide/LRAT/Internal/Formula/Implementation.lean index 7027a2c52b06..108fc7114ed0 100644 --- a/src/Std/Tactic/BVDecide/LRAT/Internal/Formula/Implementation.lean +++ b/src/Std/Tactic/BVDecide/LRAT/Internal/Formula/Implementation.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Josh Clune -/ prelude -import Init.Data.Array import Std.Tactic.BVDecide.LRAT.Internal.Formula.Class import Std.Tactic.BVDecide.LRAT.Internal.Assignment import Std.Sat.CNF.Basic diff --git a/src/Std/Time/DateTime/Timestamp.lean b/src/Std/Time/DateTime/Timestamp.lean index 8ba43cc11892..3d4527a00a4f 100644 --- a/src/Std/Time/DateTime/Timestamp.lean +++ b/src/Std/Time/DateTime/Timestamp.lean @@ -5,7 +5,6 @@ Authors: Sofia Rodrigues -/ prelude import Std.Time.Internal -import Init.Data.Int import Init.System.IO import Std.Time.Time import Std.Time.Date diff --git a/src/Std/Time/Internal/Bounded.lean b/src/Std/Time/Internal/Bounded.lean index 897636c9a077..951dc73b43ce 100644 --- a/src/Std/Time/Internal/Bounded.lean +++ b/src/Std/Time/Internal/Bounded.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Init.Data.Int +import Init.Omega namespace Std namespace Time From 639e6e92a47df90000b95d7f092ae42acd7aabea Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Sat, 4 Jan 2025 15:33:24 +1100 Subject: [PATCH 020/100] chore: cleanup imports in Lean.Lsp (#6523) This PR splits a definition out of `Lean.Lsp.Basic`, with the effect that material about JSON is not needed for `Lean.Meta.Sorry` and its dependencies. --- src/Lean/Data/Lsp.lean | 1 + src/Lean/Data/Lsp/Basic.lean | 5 ----- src/Lean/Data/Lsp/CancelParams.lean | 25 +++++++++++++++++++++++++ 3 files changed, 26 insertions(+), 5 deletions(-) create mode 100644 src/Lean/Data/Lsp/CancelParams.lean diff --git a/src/Lean/Data/Lsp.lean b/src/Lean/Data/Lsp.lean index 441d4e82fc03..8d3c7daf260f 100644 --- a/src/Lean/Data/Lsp.lean +++ b/src/Lean/Data/Lsp.lean @@ -6,6 +6,7 @@ Authors: Marc Huisinga, Wojciech Nawrocki -/ prelude import Lean.Data.Lsp.Basic +import Lean.Data.Lsp.CancelParams import Lean.Data.Lsp.Capabilities import Lean.Data.Lsp.Client import Lean.Data.Lsp.Communication diff --git a/src/Lean/Data/Lsp/Basic.lean b/src/Lean/Data/Lsp/Basic.lean index 55fb5dd8ab85..a9aaa93c1e60 100644 --- a/src/Lean/Data/Lsp/Basic.lean +++ b/src/Lean/Data/Lsp/Basic.lean @@ -6,7 +6,6 @@ Authors: Marc Huisinga, Wojciech Nawrocki -/ prelude import Lean.Data.Json -import Lean.Data.JsonRpc /-! Defines most of the 'Basic Structures' in the LSP specification (https://microsoft.github.io/language-server-protocol/specifications/specification-current/), @@ -19,10 +18,6 @@ namespace Lsp open Json -structure CancelParams where - id : JsonRpc.RequestID - deriving Inhabited, BEq, ToJson, FromJson - abbrev DocumentUri := String /-- We adopt the convention that zero-based UTF-16 positions as sent by LSP clients diff --git a/src/Lean/Data/Lsp/CancelParams.lean b/src/Lean/Data/Lsp/CancelParams.lean new file mode 100644 index 000000000000..84e95838b41d --- /dev/null +++ b/src/Lean/Data/Lsp/CancelParams.lean @@ -0,0 +1,25 @@ +/- +Copyright (c) 2020 Marc Huisinga. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. + +Authors: Marc Huisinga, Wojciech Nawrocki +-/ +prelude +import Lean.Data.JsonRpc + +/-! # Defines `Lean.Lsp.CancelParams`. + +This is separate from `Lean.Data.Lsp.Basic` to reduce transitive dependencies. +-/ + +namespace Lean +namespace Lsp + +open Json + +structure CancelParams where + id : JsonRpc.RequestID + deriving Inhabited, BEq, ToJson, FromJson + +end Lsp +end Lean From 31435e9cd156ecc53c7070b5dd061b98d95187ab Mon Sep 17 00:00:00 2001 From: Kitamado <47292598+Seasawher@users.noreply.github.com> Date: Sat, 4 Jan 2025 17:08:12 +0900 Subject: [PATCH 021/100] doc: fix broken code blocks in RELEASES.md (#6527) just fix markdown --- RELEASES.md | 2 ++ 1 file changed, 2 insertions(+) diff --git a/RELEASES.md b/RELEASES.md index a657a38bd898..d1958e403f8f 100644 --- a/RELEASES.md +++ b/RELEASES.md @@ -107,6 +107,7 @@ For `Prop`, these tactics now suggest the `by_cases` tactic. Example: ``` tactic 'cases' failed, major premise type is not an inductive type Prop +``` * [#6381](https://github.com/leanprover/lean4/pull/6381) fixes a bug in `withTrackingZetaDelta` and `withTrackingZetaDeltaSet`. The `MetaM` caches need to be reset. See new @@ -783,6 +784,7 @@ full structure lvals. Examples of these for structure instance notation: structure PosFun where f : Nat → Nat pos : ∀ n, 0 < f n +``` - [#6168](https://github.com/leanprover/lean4/pull/6168) extends the "motive is not type correct" error message for the rewrite tactic to explain what it means. It also pretty prints the From 37127ead07980694981d21a5866978ff6828403d Mon Sep 17 00:00:00 2001 From: Leonardo de Moura Date: Sat, 4 Jan 2025 19:45:55 +0100 Subject: [PATCH 022/100] fix: missing propagation in `grind` (#6528) This PR adds a missing propagation rule to the (WIP) `grind` tactic. --- src/Lean/Meta/Tactic/Grind/Propagate.lean | 2 ++ tests/lean/run/grind_propagate_connectives.lean | 3 +++ 2 files changed, 5 insertions(+) diff --git a/src/Lean/Meta/Tactic/Grind/Propagate.lean b/src/Lean/Meta/Tactic/Grind/Propagate.lean index 14a3a786ef84..42d1135a14f0 100644 --- a/src/Lean/Meta/Tactic/Grind/Propagate.lean +++ b/src/Lean/Meta/Tactic/Grind/Propagate.lean @@ -99,6 +99,8 @@ builtin_grind_propagator propagateNotUp ↑Not := fun e => do else if (← isEqTrue a) then -- a = True → (Not a) = False pushEqFalse e <| mkApp2 (mkConst ``Lean.Grind.not_eq_of_eq_true) a (← mkEqTrueProof a) + else if (← isEqv e a) then + closeGoal <| mkApp2 (mkConst ``Lean.Grind.false_of_not_eq_self) a (← mkEqProof e a) /-- Propagates truth values downwards for a negation expression `Not a` based on the truth value of `Not a`. diff --git a/tests/lean/run/grind_propagate_connectives.lean b/tests/lean/run/grind_propagate_connectives.lean index e4f4a020dc68..6cb042ae1bba 100644 --- a/tests/lean/run/grind_propagate_connectives.lean +++ b/tests/lean/run/grind_propagate_connectives.lean @@ -103,3 +103,6 @@ example (p q r : Prop) : p ∨ (q ↔ r) → ¬r → q → False := by grind on_failure do Lean.logInfo "hello world" fallback + +example (a b : Nat) : (a = b) = (b = a) := by + grind From ad2c16daded621cb6bd1017515d1352df19390c3 Mon Sep 17 00:00:00 2001 From: Leonardo de Moura Date: Sat, 4 Jan 2025 20:24:56 +0100 Subject: [PATCH 023/100] feat: add support for `let`-declarations to `grind` (#6529) This PR adds support for `let`-declarations to the (WIP) `grind` tactic. --- src/Init/Grind/Norm.lean | 6 ++++++ src/Lean/Meta/Tactic/Grind/Core.lean | 13 ++++++++++--- src/Lean/Meta/Tactic/Grind/Intro.lean | 14 +++++++++++--- tests/lean/run/grind_t1.lean | 3 +++ 4 files changed, 30 insertions(+), 6 deletions(-) diff --git a/src/Init/Grind/Norm.lean b/src/Init/Grind/Norm.lean index 4bf087bdb84e..72bc3021ec1a 100644 --- a/src/Init/Grind/Norm.lean +++ b/src/Init/Grind/Norm.lean @@ -59,6 +59,12 @@ attribute [grind_norm] ite_true ite_false @[grind_norm↓] theorem not_ite {_ : Decidable p} (q r : Prop) : (¬ite p q r) = ite p (¬q) (¬r) := by by_cases p <;> simp [*] +@[grind_norm] theorem ite_true_false {_ : Decidable p} : (ite p True False) = p := by + by_cases p <;> simp + +@[grind_norm] theorem ite_false_true {_ : Decidable p} : (ite p False True) = ¬p := by + by_cases p <;> simp + -- Forall @[grind_norm↓] theorem not_forall (p : α → Prop) : (¬∀ x, p x) = ∃ x, ¬p x := by simp attribute [grind_norm] forall_and diff --git a/src/Lean/Meta/Tactic/Grind/Core.lean b/src/Lean/Meta/Tactic/Grind/Core.lean index 747e9eb9875f..ce63088c7fbb 100644 --- a/src/Lean/Meta/Tactic/Grind/Core.lean +++ b/src/Lean/Meta/Tactic/Grind/Core.lean @@ -190,15 +190,22 @@ where addEqStep lhs rhs proof isHEq processTodo +/-- Adds a new equality `lhs = rhs`. It assumes `lhs` and `rhs` have already been internalized. -/ def addEq (lhs rhs proof : Expr) : GoalM Unit := do addEqCore lhs rhs proof false + +/-- Adds a new heterogeneous equality `HEq lhs rhs`. It assumes `lhs` and `rhs` have already been internalized. -/ def addHEq (lhs rhs proof : Expr) : GoalM Unit := do addEqCore lhs rhs proof true -/-- -Adds a new `fact` justified by the given proof and using the given generation. --/ +/-- Internalizes `lhs` and `rhs`, and then adds equality `lhs = rhs`. -/ +def addNewEq (lhs rhs proof : Expr) (generation : Nat) : GoalM Unit := do + internalize lhs generation + internalize rhs generation + addEq lhs rhs proof + +/-- Adds a new `fact` justified by the given proof and using the given generation. -/ def add (fact : Expr) (proof : Expr) (generation := 0) : GoalM Unit := do trace[grind.assert] "{fact}" if (← isInconsistent) then return () diff --git a/src/Lean/Meta/Tactic/Grind/Intro.lean b/src/Lean/Meta/Tactic/Grind/Intro.lean index b0b60282b922..e870a0931c47 100644 --- a/src/Lean/Meta/Tactic/Grind/Intro.lean +++ b/src/Lean/Meta/Tactic/Grind/Intro.lean @@ -21,7 +21,7 @@ private inductive IntroResult where | newLocal (fvarId : FVarId) (goal : Goal) deriving Inhabited -private def introNext (goal : Goal) : GrindM IntroResult := do +private def introNext (goal : Goal) (generation : Nat) : GrindM IntroResult := do let target ← goal.mvarId.getType if target.isArrow then goal.mvarId.withContext do @@ -58,7 +58,15 @@ private def introNext (goal : Goal) : GrindM IntroResult := do let mvarId ← mvarId.assert (← mkFreshUserName localDecl.userName) localDecl.type (mkFVar fvarId) return .newDepHyp { goal with mvarId } else - return .newLocal fvarId { goal with mvarId } + let goal := { goal with mvarId } + if target.isLet then + let v := target.letValue! + let r ← simp v + let x ← shareCommon (mkFVar fvarId) + let goal ← GoalM.run' goal <| addNewEq x r.expr (← r.getProof) generation + return .newLocal fvarId goal + else + return .newLocal fvarId goal else return .done @@ -84,7 +92,7 @@ partial def intros (goal : Goal) (generation : Nat) : GrindM (List Goal) := do let rec go (goal : Goal) : StateRefT (Array Goal) GrindM Unit := do if goal.inconsistent then return () - match (← introNext goal) with + match (← introNext goal generation) with | .done => if let some mvarId ← goal.mvarId.byContra? then go { goal with mvarId } diff --git a/tests/lean/run/grind_t1.lean b/tests/lean/run/grind_t1.lean index 168bdc895fff..3f051264842d 100644 --- a/tests/lean/run/grind_t1.lean +++ b/tests/lean/run/grind_t1.lean @@ -111,3 +111,6 @@ example (foo : Nat → Nat) grind end dite_propagator_test + +example (a : Nat) : let x := a + a; y = x → y = a + a := by + grind From a5b1ed906c867d56b6e595fd74e2aff6852c0463 Mon Sep 17 00:00:00 2001 From: Leonardo de Moura Date: Sat, 4 Jan 2025 22:40:32 +0100 Subject: [PATCH 024/100] fix: nondeterministic failure in `grind` (#6530) This PR fixes nondeterministic failures in the (WIP) `grind` tactic. --- src/Lean/Meta/Tactic/Grind.lean | 1 + src/Lean/Meta/Tactic/Grind/Canon.lean | 40 ++++++-------- src/Lean/Meta/Tactic/Grind/Core.lean | 5 +- src/Lean/Meta/Tactic/Grind/ForallProp.lean | 3 + .../Meta/Tactic/Grind/MarkNestedProofs.lean | 55 +++++++++++-------- src/Lean/Meta/Tactic/Grind/Util.lean | 11 +++- 6 files changed, 64 insertions(+), 51 deletions(-) diff --git a/src/Lean/Meta/Tactic/Grind.lean b/src/Lean/Meta/Tactic/Grind.lean index 11b15243e263..01ee792bbec6 100644 --- a/src/Lean/Meta/Tactic/Grind.lean +++ b/src/Lean/Meta/Tactic/Grind.lean @@ -47,5 +47,6 @@ builtin_initialize registerTraceClass `grind.debug.proof builtin_initialize registerTraceClass `grind.debug.proj builtin_initialize registerTraceClass `grind.debug.parent builtin_initialize registerTraceClass `grind.debug.final +builtin_initialize registerTraceClass `grind.debug.forallPropagator end Lean diff --git a/src/Lean/Meta/Tactic/Grind/Canon.lean b/src/Lean/Meta/Tactic/Grind/Canon.lean index c895dd639a73..05057121709c 100644 --- a/src/Lean/Meta/Tactic/Grind/Canon.lean +++ b/src/Lean/Meta/Tactic/Grind/Canon.lean @@ -111,22 +111,13 @@ unsafe def canonImpl (e : Expr) : StateT State MetaM Expr := do visit e |>.run' mkPtrMap where visit (e : Expr) : StateRefT (PtrMap Expr Expr) (StateT State MetaM) Expr := do - match e with - | .bvar .. => unreachable! - -- Recall that `grind` treats `let`, `forall`, and `lambda` as atomic terms. - | .letE .. | .forallE .. | .lam .. - | .const .. | .lit .. | .mvar .. | .sort .. | .fvar .. - -- Recall that the `grind` preprocessor uses the `foldProjs` preprocessing step. - | .proj .. - -- Recall that the `grind` preprocessor uses the `eraseIrrelevantMData` preprocessing step. - | .mdata .. => return e - -- We only visit applications - | .app .. => - -- Check whether it is cached - if let some r := (← get).find? e then - return r - e.withApp fun f args => do - let e' ← if f.isConstOf ``Lean.Grind.nestedProof && args.size == 2 then + unless e.isApp || e.isForall do return e + -- Check whether it is cached + if let some r := (← get).find? e then + return r + let e' ← match e with + | .app .. => e.withApp fun f args => do + if f.isConstOf ``Lean.Grind.nestedProof && args.size == 2 then let prop := args[0]! let prop' ← visit prop if let some r := (← getThe State).proofCanon.find? prop' then @@ -149,12 +140,17 @@ where args := args.set! i arg' modified := true pure <| if modified then mkAppN f args else e - modify fun s => s.insert e e' - return e' - -/-- -Canonicalizes nested types, type formers, and instances in `e`. --/ + | .forallE _ d b _ => + -- Recall that we have `ForallProp.lean`. + let d' ← visit d + -- Remark: users may not want to convert `p → q` into `¬p ∨ q` + let b' ← if b.hasLooseBVars then pure b else visit b + pure <| e.updateForallE! d' b' + | _ => unreachable! + modify fun s => s.insert e e' + return e' + +/-- Canonicalizes nested types, type formers, and instances in `e`. -/ def canon (e : Expr) : StateT State MetaM Expr := unsafe canonImpl e diff --git a/src/Lean/Meta/Tactic/Grind/Core.lean b/src/Lean/Meta/Tactic/Grind/Core.lean index ce63088c7fbb..635e7aef85e9 100644 --- a/src/Lean/Meta/Tactic/Grind/Core.lean +++ b/src/Lean/Meta/Tactic/Grind/Core.lean @@ -41,7 +41,8 @@ This is an auxiliary function performed while merging equivalence classes. private def removeParents (root : Expr) : GoalM ParentSet := do let parents ← getParentsAndReset root for parent in parents do - if (← isCongrRoot parent) then + -- Recall that we may have `Expr.forallE` in `parents` because of `ForallProp.lean` + if (← pure parent.isApp <&&> isCongrRoot parent) then trace[grind.debug.parent] "remove: {parent}" modify fun s => { s with congrTable := s.congrTable.erase { e := parent } } return parents @@ -52,7 +53,7 @@ This is an auxiliary function performed while merging equivalence classes. -/ private def reinsertParents (parents : ParentSet) : GoalM Unit := do for parent in parents do - if (← isCongrRoot parent) then + if (← pure parent.isApp <&&> isCongrRoot parent) then trace[grind.debug.parent] "reinsert: {parent}" addCongrTable parent diff --git a/src/Lean/Meta/Tactic/Grind/ForallProp.lean b/src/Lean/Meta/Tactic/Grind/ForallProp.lean index 9228fe671309..9993f0e50e2d 100644 --- a/src/Lean/Meta/Tactic/Grind/ForallProp.lean +++ b/src/Lean/Meta/Tactic/Grind/ForallProp.lean @@ -17,13 +17,16 @@ If so, internalize the term `proj_i (ctor ... a_i ...)` and add the equality `pr -/ def propagateForallProp (parent : Expr) : GoalM Unit := do let .forallE n p q bi := parent | return () + trace[grind.debug.forallPropagator] "{parent}" unless (← isEqTrue p) do return () + trace[grind.debug.forallPropagator] "isEqTrue, {parent}" let h₁ ← mkEqTrueProof p let qh₁ := q.instantiate1 (mkApp2 (mkConst ``of_eq_true) p h₁) let r ← simp qh₁ let q := mkLambda n bi p q let q' := r.expr internalize q' (← getGeneration parent) + trace[grind.debug.forallPropagator] "q': {q'} for{indentExpr parent}" let h₂ ← r.getProof let h := mkApp5 (mkConst ``Lean.Grind.forall_propagator) p q q' h₁ h₂ pushEq parent q' h diff --git a/src/Lean/Meta/Tactic/Grind/MarkNestedProofs.lean b/src/Lean/Meta/Tactic/Grind/MarkNestedProofs.lean index 485bd0a35c7d..85d2d8fa8cb4 100644 --- a/src/Lean/Meta/Tactic/Grind/MarkNestedProofs.lean +++ b/src/Lean/Meta/Tactic/Grind/MarkNestedProofs.lean @@ -24,30 +24,37 @@ where let e' := mkApp2 (mkConst ``Lean.Grind.nestedProof) prop e modify fun s => s.insert e e' return e' - else match e with - | .bvar .. => unreachable! - -- See comments on `Canon.lean` for why we do not visit these cases. - | .letE .. | .forallE .. | .lam .. - | .const .. | .lit .. | .mvar .. | .sort .. | .fvar .. => return e - | .proj _ _ b => return e.updateProj! (← visit b) - | .mdata _ b => return e.updateMData! (← visit b) - -- We only visit applications - | .app .. => - -- Check whether it is cached - if let some r := (← get).find? e then - return r - e.withApp fun f args => do - let mut modified := false - let mut args := args - for i in [:args.size] do - let arg := args[i]! - let arg' ← visit arg - unless ptrEq arg arg' do - args := args.set! i arg' - modified := true - let e' := if modified then mkAppN f args else e - modify fun s => s.insert e e' - return e' + -- Remark: we have to process `Expr.proj` since we only + -- fold projections later during term internalization + unless e.isApp || e.isForall || e.isProj do + return e + -- Check whether it is cached + if let some r := (← get).find? e then + return r + let e' ← match e with + | .app .. => e.withApp fun f args => do + let mut modified := false + let mut args := args + for i in [:args.size] do + let arg := args[i]! + let arg' ← visit arg + unless ptrEq arg arg' do + args := args.set! i arg' + modified := true + if modified then + pure <| mkAppN f args + else + pure e + | .proj _ _ b => + pure <| e.updateProj! (← visit b) + | .forallE _ d b _ => + -- Recall that we have `ForallProp.lean`. + let d' ← visit d + let b' ← if b.hasLooseBVars then pure b else visit b + pure <| e.updateForallE! d' b' + | _ => unreachable! + modify fun s => s.insert e e' + return e' /-- Wrap nested proofs `e` with `Lean.Grind.nestedProof`-applications. diff --git a/src/Lean/Meta/Tactic/Grind/Util.lean b/src/Lean/Meta/Tactic/Grind/Util.lean index bf7bdabc5817..0805b4dd7fb8 100644 --- a/src/Lean/Meta/Tactic/Grind/Util.lean +++ b/src/Lean/Meta/Tactic/Grind/Util.lean @@ -100,12 +100,14 @@ def _root_.Lean.MVarId.clearAuxDecls (mvarId : MVarId) : MetaM MVarId := mvarId. /-- In the `grind` tactic, during `Expr` internalization, we don't expect to find `Expr.mdata`. This function ensures `Expr.mdata` is not found during internalization. -Recall that we do not internalize `Expr.forallE` and `Expr.lam` components. +Recall that we do not internalize `Expr.lam` children. +Recall that we still have to process `Expr.forallE` because of `ForallProp.lean`. +Moreover, we may not want to reduce `p → q` to `¬p ∨ q` when `(p q : Prop)`. -/ def eraseIrrelevantMData (e : Expr) : CoreM Expr := do let pre (e : Expr) := do match e with - | .letE .. | .lam .. | .forallE .. => return .done e + | .letE .. | .lam .. => return .done e | .mdata _ e => return .continue e | _ => return .continue e Core.transform e (pre := pre) @@ -116,11 +118,14 @@ Converts nested `Expr.proj`s into projection applications if possible. def foldProjs (e : Expr) : MetaM Expr := do let post (e : Expr) := do let .proj structName idx s := e | return .done e - let some info := getStructureInfo? (← getEnv) structName | return .done e + let some info := getStructureInfo? (← getEnv) structName | + trace[grind.issues] "found `Expr.proj` but `{structName}` is not marked as structure{indentExpr e}" + return .done e if h : idx < info.fieldNames.size then let fieldName := info.fieldNames[idx] return .done (← mkProjection s fieldName) else + trace[grind.issues] "found `Expr.proj` with invalid field index `{idx}`{indentExpr e}" return .done e Meta.transform e (post := post) From d22233fc7b9b865905cdd55d40498b983d765543 Mon Sep 17 00:00:00 2001 From: Leonardo de Moura Date: Sat, 4 Jan 2025 23:22:55 +0100 Subject: [PATCH 025/100] fix: `let_fun` support in `grind` (#6531) This PR fixes the support for `let_fun` in `grind`. --- src/Lean/Meta/Tactic/Grind/Intro.lean | 6 +++--- tests/lean/run/grind_t1.lean | 19 +++++++++++++++++++ 2 files changed, 22 insertions(+), 3 deletions(-) diff --git a/src/Lean/Meta/Tactic/Grind/Intro.lean b/src/Lean/Meta/Tactic/Grind/Intro.lean index e870a0931c47..3923b4d1abbe 100644 --- a/src/Lean/Meta/Tactic/Grind/Intro.lean +++ b/src/Lean/Meta/Tactic/Grind/Intro.lean @@ -49,7 +49,7 @@ private def introNext (goal : Goal) (generation : Nat) : GrindM IntroResult := d -- `p` and `p'` are definitionally equal goal.mvarId.assign h return .newHyp fvarId { goal with mvarId := mvarIdNew } - else if target.isLet || target.isForall then + else if target.isLet || target.isForall || target.isLetFun then let (fvarId, mvarId) ← goal.mvarId.intro1P mvarId.withContext do let localDecl ← fvarId.getDecl @@ -59,8 +59,8 @@ private def introNext (goal : Goal) (generation : Nat) : GrindM IntroResult := d return .newDepHyp { goal with mvarId } else let goal := { goal with mvarId } - if target.isLet then - let v := target.letValue! + if target.isLet || target.isLetFun then + let v := (← fvarId.getDecl).value let r ← simp v let x ← shareCommon (mkFVar fvarId) let goal ← GoalM.run' goal <| addNewEq x r.expr (← r.getProof) generation diff --git a/tests/lean/run/grind_t1.lean b/tests/lean/run/grind_t1.lean index 3f051264842d..7686d4c0d832 100644 --- a/tests/lean/run/grind_t1.lean +++ b/tests/lean/run/grind_t1.lean @@ -112,5 +112,24 @@ example (foo : Nat → Nat) end dite_propagator_test +/-- +info: [grind.eqc] x = 2 * a +[grind.eqc] y = x +[grind.eqc] (y = 2 * a) = False +[grind.eqc] (y = 2 * a) = True +-/ +#guard_msgs (info) in +set_option trace.grind.eqc true in example (a : Nat) : let x := a + a; y = x → y = a + a := by grind + +/-- +info: [grind.eqc] x = 2 * a +[grind.eqc] y = x +[grind.eqc] (y = 2 * a) = False +[grind.eqc] (y = 2 * a) = True +-/ +#guard_msgs (info) in +set_option trace.grind.eqc true in +example (a : Nat) : let_fun x := a + a; y = x → y = a + a := by + grind From 9dcbc330fd08a5ae455a6e5a5c6b30b877bb2917 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Sun, 5 Jan 2025 10:43:23 +1100 Subject: [PATCH 026/100] chore: fix signature of perm_insertIdx (#6532) --- src/Init/Data/List/Perm.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/Init/Data/List/Perm.lean b/src/Init/Data/List/Perm.lean index 5eef4cc0367e..57fc5b45e77c 100644 --- a/src/Init/Data/List/Perm.lean +++ b/src/Init/Data/List/Perm.lean @@ -510,7 +510,7 @@ theorem Perm.eraseP (f : α → Bool) {l₁ l₂ : List α} refine (IH₁ H).trans (IH₂ ((p₁.pairwise_iff ?_).1 H)) exact fun h h₁ h₂ => h h₂ h₁ -theorem perm_insertIdx [DecidableEq α] {α} (x : α) (l : List α) {n} (h : n ≤ l.length) : +theorem perm_insertIdx {α} (x : α) (l : List α) {n} (h : n ≤ l.length) : insertIdx n x l ~ x :: l := by induction l generalizing n with | nil => From dc5c8097b5ca053dd49b5de246add54c3e999734 Mon Sep 17 00:00:00 2001 From: Leonardo de Moura Date: Sun, 5 Jan 2025 03:20:17 +0100 Subject: [PATCH 027/100] feat: add term offset support to the `grind` E-matching modulo (#6533) This PR adds support to E-matching offset patterns. For example, we want to be able to E-match the pattern `f (#0 + 1)` with term `f (a + 2)`. --- src/Init/Grind/Norm.lean | 3 + src/Init/Grind/Util.lean | 3 + src/Lean/Elab/Tactic/Grind.lean | 6 +- src/Lean/Meta/Tactic/Grind/EMatch.lean | 73 +++++++++++++--- src/Lean/Meta/Tactic/Grind/EMatchTheorem.lean | 45 +++++++++- src/Lean/Meta/Tactic/Grind/Main.lean | 8 +- tests/lean/run/grind_offset.lean | 87 +++++++++++++++++++ 7 files changed, 206 insertions(+), 19 deletions(-) create mode 100644 tests/lean/run/grind_offset.lean diff --git a/src/Init/Grind/Norm.lean b/src/Init/Grind/Norm.lean index 72bc3021ec1a..d911cf634703 100644 --- a/src/Init/Grind/Norm.lean +++ b/src/Init/Grind/Norm.lean @@ -114,4 +114,7 @@ attribute [grind_norm] Nat.le_zero_eq -- GT GE attribute [grind_norm] GT.gt GE.ge +-- Succ +attribute [grind_norm] Nat.succ_eq_add_one + end Lean.Grind diff --git a/src/Init/Grind/Util.lean b/src/Init/Grind/Util.lean index 9c7523d992cb..9a231e318162 100644 --- a/src/Init/Grind/Util.lean +++ b/src/Init/Grind/Util.lean @@ -18,6 +18,9 @@ We use it when adding instances of `match`-equations to prevent them from being -/ def doNotSimp {α : Sort u} (a : α) : α := a +/-- Gadget for representing offsets `t+k` in patterns. -/ +def offset (a b : Nat) : Nat := a + b + set_option pp.proofs true theorem nestedProof_congr (p q : Prop) (h : p = q) (hp : p) (hq : q) : HEq (nestedProof p hp) (nestedProof q hq) := by diff --git a/src/Lean/Elab/Tactic/Grind.lean b/src/Lean/Elab/Tactic/Grind.lean index b9a62e4baed7..d094ae98e4bb 100644 --- a/src/Lean/Elab/Tactic/Grind.lean +++ b/src/Lean/Elab/Tactic/Grind.lean @@ -26,8 +26,10 @@ def elabGrindPattern : CommandElab := fun stx => do let info ← getConstInfo declName forallTelescope info.type fun xs _ => do let patterns ← terms.getElems.mapM fun term => do - let pattern ← instantiateMVars (← elabTerm term none) - let pattern ← Grind.unfoldReducible pattern + let pattern ← elabTerm term none + synthesizeSyntheticMVarsUsingDefault + let pattern ← instantiateMVars pattern + let pattern ← Grind.preprocessPattern pattern return pattern.abstract xs Grind.addEMatchTheorem declName xs.size patterns.toList | _ => throwUnsupportedSyntax diff --git a/src/Lean/Meta/Tactic/Grind/EMatch.lean b/src/Lean/Meta/Tactic/Grind/EMatch.lean index 42031260a0a7..ca6df2b7a07c 100644 --- a/src/Lean/Meta/Tactic/Grind/EMatch.lean +++ b/src/Lean/Meta/Tactic/Grind/EMatch.lean @@ -16,6 +16,8 @@ namespace EMatch inductive Cnstr where | /-- Matches pattern `pat` with term `e` -/ «match» (pat : Expr) (e : Expr) + | /-- Matches offset pattern `pat+k` with term `e` -/ + offset (pat : Expr) (k : Nat) (e : Expr) | /-- This constraint is used to encode multi-patterns. -/ «continue» (pat : Expr) deriving Inhabited @@ -88,6 +90,28 @@ private def eqvFunctions (pFn eFn : Expr) : Bool := (pFn.isFVar && pFn == eFn) || (pFn.isConst && eFn.isConstOf pFn.constName!) +/-- Matches a pattern argument. See `matchArgs?`. -/ +private def matchArg? (c : Choice) (pArg : Expr) (eArg : Expr) : OptionT GoalM Choice := do + if isPatternDontCare pArg then + return c + else if pArg.isBVar then + assign? c pArg.bvarIdx! eArg + else if let some pArg := groundPattern? pArg then + guard (← isEqv pArg eArg) + return c + else if let some (pArg, k) := isOffsetPattern? pArg then + assert! Option.isNone <| isOffsetPattern? pArg + assert! !isPatternDontCare pArg + return { c with cnstrs := .offset pArg k eArg :: c.cnstrs } + else + return { c with cnstrs := .match pArg eArg :: c.cnstrs } + +private def Choice.updateGen (c : Choice) (gen : Nat) : Choice := + { c with gen := Nat.max gen c.gen } + +private def pushChoice (c : Choice) : M Unit := + modify fun s => { s with choiceStack := c :: s.choiceStack } + /-- Matches arguments of pattern `p` with term `e`. Returns `some` if successful, and `none` otherwise. It may update `c`s assignment and list of contraints to be @@ -97,16 +121,8 @@ private partial def matchArgs? (c : Choice) (p : Expr) (e : Expr) : OptionT Goal if !p.isApp then return c -- Done let pArg := p.appArg! let eArg := e.appArg! - let goFn c := matchArgs? c p.appFn! e.appFn! - if isPatternDontCare pArg then - goFn c - else if pArg.isBVar then - goFn (← assign? c pArg.bvarIdx! eArg) - else if let some pArg := groundPattern? pArg then - guard (← isEqv pArg eArg) - goFn c - else - goFn { c with cnstrs := .match pArg eArg :: c.cnstrs } + let c ← matchArg? c pArg eArg + matchArgs? c p.appFn! e.appFn! /-- Matches pattern `p` with term `e` with respect to choice `c`. @@ -127,9 +143,39 @@ private partial def processMatch (c : Choice) (p : Expr) (e : Expr) : M Unit := && eqvFunctions pFn curr.getAppFn && curr.getAppNumArgs == numArgs then if let some c ← matchArgs? c p curr |>.run then - let gen := n.generation - let c := { c with gen := Nat.max gen c.gen } - modify fun s => { s with choiceStack := c :: s.choiceStack } + pushChoice (c.updateGen n.generation) + curr ← getNext curr + if isSameExpr curr e then break + +/-- +Matches offset pattern `pArg+k` with term `e` with respect to choice `c`. +-/ +private partial def processOffset (c : Choice) (pArg : Expr) (k : Nat) (e : Expr) : M Unit := do + let maxGeneration ← getMaxGeneration + let mut curr := e + repeat + let n ← getENode curr + if n.generation <= maxGeneration then + if let some (eArg, k') ← isOffset? curr |>.run then + if k' < k then + let c := c.updateGen n.generation + pushChoice { c with cnstrs := .offset pArg (k - k') eArg :: c.cnstrs } + else if k' == k then + if let some c ← matchArg? c pArg eArg |>.run then + pushChoice (c.updateGen n.generation) + else if k' > k then + let eArg' := mkNatAdd eArg (mkNatLit (k' - k)) + let eArg' ← shareCommon (← canon eArg') + internalize eArg' n.generation + if let some c ← matchArg? c pArg eArg' |>.run then + pushChoice (c.updateGen n.generation) + else if let some k' ← evalNat curr |>.run then + if k' >= k then + let eArg' := mkNatLit (k' - k) + let eArg' ← shareCommon (← canon eArg') + internalize eArg' n.generation + if let some c ← matchArg? c pArg eArg' |>.run then + pushChoice (c.updateGen n.generation) curr ← getNext curr if isSameExpr curr e then break @@ -224,6 +270,7 @@ private partial def processChoices : M Unit := do match c.cnstrs with | [] => instantiateTheorem c | .match p e :: cnstrs => processMatch { c with cnstrs } p e + | .offset p k e :: cnstrs => processOffset { c with cnstrs } p k e | .continue p :: cnstrs => processContinue { c with cnstrs } p processChoices diff --git a/src/Lean/Meta/Tactic/Grind/EMatchTheorem.lean b/src/Lean/Meta/Tactic/Grind/EMatchTheorem.lean index 4b03d9021fe5..8b50693e1523 100644 --- a/src/Lean/Meta/Tactic/Grind/EMatchTheorem.lean +++ b/src/Lean/Meta/Tactic/Grind/EMatchTheorem.lean @@ -4,15 +4,45 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Leonardo de Moura -/ prelude +import Init.Grind.Util import Lean.HeadIndex import Lean.PrettyPrinter import Lean.Util.FoldConsts import Lean.Util.CollectFVars import Lean.Meta.Basic import Lean.Meta.InferType +import Lean.Meta.Tactic.Grind.Util namespace Lean.Meta.Grind +def mkOffsetPattern (pat : Expr) (k : Nat) : Expr := + mkApp2 (mkConst ``Grind.offset) pat (mkRawNatLit k) + +private def detectOffsets (pat : Expr) : MetaM Expr := do + let pre (e : Expr) := do + if e == pat then + -- We only consider nested offset patterns + return .continue e + else match e with + | .letE .. | .lam .. | .forallE .. => return .done e + | _ => + let some (e, k) ← isOffset? e + | return .continue e + if k == 0 then return .continue e + return .continue <| mkOffsetPattern e k + Core.transform pat (pre := pre) + +def isOffsetPattern? (pat : Expr) : Option (Expr × Nat) := Id.run do + let_expr Grind.offset pat k := pat | none + let .lit (.natVal k) := k | none + return some (pat, k) + +def preprocessPattern (pat : Expr) : MetaM Expr := do + let pat ← instantiateMVars pat + let pat ← unfoldReducible pat + let pat ← detectOffsets pat + return pat + inductive Origin where /-- A global declaration in the environment. -/ | decl (declName : Name) @@ -202,6 +232,12 @@ private def getPatternFunMask (f : Expr) (numArgs : Nat) : MetaM (Array Bool) := private partial def go (pattern : Expr) (root := false) : M Expr := do if root && !pattern.hasLooseBVars then throwError "invalid pattern, it does not have pattern variables" + if let some (e, k) := isOffsetPattern? pattern then + let e ← goArg e (isSupport := false) + if e == dontCare then + return dontCare + else + return mkOffsetPattern e k let some f := getPatternFn? pattern | throwError "invalid pattern, (non-forbidden) application expected" assert! f.isConst || f.isFVar @@ -211,7 +247,11 @@ private partial def go (pattern : Expr) (root := false) : M Expr := do for i in [:args.size] do let arg := args[i]! let isSupport := supportMask[i]?.getD false - let arg ← if !arg.hasLooseBVars then + args := args.set! i (← goArg arg isSupport) + return mkAppN f args +where + goArg (arg : Expr) (isSupport : Bool) : M Expr := do + if !arg.hasLooseBVars then if arg.hasMVar then pure dontCare else @@ -230,8 +270,6 @@ private partial def go (pattern : Expr) (root := false) : M Expr := do go arg else pure dontCare - args := args.set! i arg - return mkAppN f args def main (patterns : List Expr) : MetaM (List Expr × List HeadIndex × Std.HashSet Nat) := do let (patterns, s) ← patterns.mapM go |>.run {} @@ -390,6 +428,7 @@ def mkEMatchEqTheorem (declName : Name) : MetaM EMatchTheorem := do let info ← getConstInfo declName let (numParams, patterns) ← forallTelescopeReducing info.type fun xs type => do let_expr Eq _ lhs _ := type | throwError "invalid E-matching equality theorem, conclusion must be an equality{indentExpr type}" + let lhs ← preprocessPattern lhs return (xs.size, [lhs.abstract xs]) mkEMatchTheorem declName numParams patterns diff --git a/src/Lean/Meta/Tactic/Grind/Main.lean b/src/Lean/Meta/Tactic/Grind/Main.lean index 98214895ee8d..8ee85469bdc3 100644 --- a/src/Lean/Meta/Tactic/Grind/Main.lean +++ b/src/Lean/Meta/Tactic/Grind/Main.lean @@ -6,6 +6,7 @@ Authors: Leonardo de Moura prelude import Init.Grind.Lemmas import Lean.Meta.Tactic.Util +import Lean.Meta.Tactic.Simp.Simproc import Lean.Meta.Tactic.Grind.RevertAll import Lean.Meta.Tactic.Grind.PropagatorAttr import Lean.Meta.Tactic.Grind.Proj @@ -34,12 +35,17 @@ def mkMethods (fallback : Fallback) : CoreM Methods := do prop e } +private def getGrindSimprocs : MetaM Simprocs := do + let s ← grindNormSimprocExt.getSimprocs + let s ← addDoNotSimp s + return s + def GrindM.run (x : GrindM α) (mainDeclName : Name) (config : Grind.Config) (fallback : Fallback) : MetaM α := do let scState := ShareCommon.State.mk _ let (falseExpr, scState) := ShareCommon.State.shareCommon scState (mkConst ``False) let (trueExpr, scState) := ShareCommon.State.shareCommon scState (mkConst ``True) let thms ← grindNormExt.getTheorems - let simprocs := #[(← addDoNotSimp (← grindNormSimprocExt.getSimprocs))] + let simprocs := #[(← getGrindSimprocs), (← Simp.getSEvalSimprocs)] let simp ← Simp.mkContext (config := { arith := true }) (simpTheorems := #[thms]) diff --git a/tests/lean/run/grind_offset.lean b/tests/lean/run/grind_offset.lean new file mode 100644 index 000000000000..7711ece39be1 --- /dev/null +++ b/tests/lean/run/grind_offset.lean @@ -0,0 +1,87 @@ +opaque g : Nat → Nat + +@[simp] def f (a : Nat) := + match a with + | 0 => 10 + | x+1 => g (f x) + +set_option trace.grind.ematch.pattern true +set_option trace.grind.ematch.instance true +set_option trace.grind.assert true + +/-- +info: [grind.ematch.pattern] f.eq_2: [f (Lean.Grind.offset #0 (1))] +-/ +#guard_msgs in +grind_pattern f.eq_2 => f (x + 1) + + +/-- +info: [grind.assert] f (y + 1) = a +[grind.assert] ¬a = g (f y) +[grind.ematch.instance] f.eq_2: f y.succ = g (f y) +[grind.assert] f (y + 1) = g (f y) +-/ +#guard_msgs (info) in +example : f (y + 1) = a → a = g (f y):= by + grind + +/-- +info: [grind.assert] f 1 = a +[grind.assert] ¬a = g (f 0) +[grind.ematch.instance] f.eq_2: f (Nat.succ 0) = g (f 0) +[grind.assert] f 1 = g (f 0) +-/ +#guard_msgs (info) in +example : f 1 = a → a = g (f 0) := by + grind + +/-- +info: [grind.assert] f 10 = a +[grind.assert] ¬a = g (f 9) +[grind.ematch.instance] f.eq_2: f (Nat.succ 8) = g (f 8) +[grind.ematch.instance] f.eq_2: f (Nat.succ 9) = g (f 9) +[grind.assert] f 9 = g (f 8) +[grind.assert] f 10 = g (f 9) +-/ +#guard_msgs (info) in +example : f 10 = a → a = g (f 9) := by + grind + +/-- +info: [grind.assert] f (c + 2) = a +[grind.assert] ¬a = g (g (f c)) +[grind.ematch.instance] f.eq_2: f (c + 1).succ = g (f (c + 1)) +[grind.assert] f (c + 2) = g (f (c + 1)) +[grind.ematch.instance] f.eq_2: f c.succ = g (f c) +[grind.assert] f (c + 1) = g (f c) +-/ +#guard_msgs (info) in +example : f (c + 2) = a → a = g (g (f c)) := by + grind + +@[simp] def foo (a : Nat) := + match a with + | 0 => 10 + | 1 => 10 + | a+2 => g (foo a) + +/-- +info: [grind.ematch.pattern] foo.eq_3: [foo (Lean.Grind.offset #0 (2))] +-/ +#guard_msgs in +grind_pattern foo.eq_3 => foo (a_2 + 2) + +-- The instance is correctly found in the following example. +-- TODO: to complete the proof, we need linear arithmetic support to prove that `b + 2 = c + 1`. +/-- +info: [grind.assert] foo (c + 1) = a +[grind.assert] c = b + 1 +[grind.assert] ¬a = g (foo b) +[grind.ematch.instance] foo.eq_3: foo b.succ.succ = g (foo b) +[grind.assert] foo (b + 2) = g (foo b) +-/ +#guard_msgs (info) in +example : foo (c + 1) = a → c = b + 1 → a = g (foo b) := by + fail_if_success grind + sorry From fb506b957c377134ee90bbb144598aa7a782b72d Mon Sep 17 00:00:00 2001 From: Leonardo de Moura Date: Sun, 5 Jan 2025 17:20:54 +0100 Subject: [PATCH 028/100] fix: allow projections in E-matching patterns (#6534) This PR ensures that users can utilize projections in E-matching patterns within the `grind` tactic. --- src/Lean/Meta/Tactic/Grind/EMatchTheorem.lean | 1 + tests/lean/run/grind_pattern_proj.lean | 47 +++++++++++++++++++ 2 files changed, 48 insertions(+) create mode 100644 tests/lean/run/grind_pattern_proj.lean diff --git a/src/Lean/Meta/Tactic/Grind/EMatchTheorem.lean b/src/Lean/Meta/Tactic/Grind/EMatchTheorem.lean index 8b50693e1523..fd175a908c7b 100644 --- a/src/Lean/Meta/Tactic/Grind/EMatchTheorem.lean +++ b/src/Lean/Meta/Tactic/Grind/EMatchTheorem.lean @@ -41,6 +41,7 @@ def preprocessPattern (pat : Expr) : MetaM Expr := do let pat ← instantiateMVars pat let pat ← unfoldReducible pat let pat ← detectOffsets pat + let pat ← foldProjs pat return pat inductive Origin where diff --git a/tests/lean/run/grind_pattern_proj.lean b/tests/lean/run/grind_pattern_proj.lean new file mode 100644 index 000000000000..8bf40291bdd7 --- /dev/null +++ b/tests/lean/run/grind_pattern_proj.lean @@ -0,0 +1,47 @@ +universe v v₁ v₂ u u₁ u₂ + +namespace CategoryTheory + +class Quiver (V : Type u) where + Hom : V → V → Sort v + +infixr:10 " ⟶ " => Quiver.Hom + +class CategoryStruct (obj : Type u) extends Quiver.{v + 1} obj : Type max u (v + 1) where + /-- The identity morphism on an object. -/ + id : ∀ X : obj, Hom X X + /-- Composition of morphisms in a category, written `f ≫ g`. -/ + comp : ∀ {X Y Z : obj}, (X ⟶ Y) → (Y ⟶ Z) → (X ⟶ Z) + +notation "𝟙" => CategoryStruct.id -- type as \b1 + +infixr:80 " ≫ " => CategoryStruct.comp -- type as \gg + +class Category (obj : Type u) extends CategoryStruct.{v} obj : Type max u (v + 1) where + id_comp : ∀ {X Y : obj} (f : X ⟶ Y), 𝟙 X ≫ f = f + comp_id : ∀ {X Y : obj} (f : X ⟶ Y), f ≫ 𝟙 Y = f + assoc : ∀ {W X Y Z : obj} (f : W ⟶ X) (g : X ⟶ Y) (h : Y ⟶ Z), (f ≫ g) ≫ h = f ≫ g ≫ h + +structure Prefunctor (V : Type u₁) [Quiver.{v₁} V] (W : Type u₂) [Quiver.{v₂} W] where + obj : V → W + map : ∀ {X Y : V}, (X ⟶ Y) → (obj X ⟶ obj Y) + +structure Functor (C : Type u₁) [Category.{v₁} C] (D : Type u₂) [Category.{v₂} D] + extends Prefunctor C D : Type max v₁ v₂ u₁ u₂ where + map_id : ∀ X : C, map (𝟙 X) = 𝟙 (obj X) + map_comp : ∀ {X Y Z : C} (f : X ⟶ Y) (g : Y ⟶ Z), map (f ≫ g) = map f ≫ map g + + +set_option trace.grind.ematch.pattern true + +/-- +info: [grind.ematch.pattern] Functor.map_id: [@Prefunctor.map #5 ? #3 ? (@Functor.toPrefunctor ? #4 ? #2 #1) #0 #0 (@CategoryStruct.id ? ? #0)] +-/ +#guard_msgs in +grind_pattern Functor.map_id => self.map (𝟙 X) + +/-- +info: [grind.ematch.pattern] Functor.map_comp: [@Prefunctor.map #9 ? #7 ? (@Functor.toPrefunctor ? #8 ? #6 #5) #4 #2 (@CategoryStruct.comp ? ? #4 #3 #2 #1 #0)] +-/ +#guard_msgs in +grind_pattern Functor.map_comp => self.map (f ≫ g) From 7b29f488df307ca73b3fc911bf2d2ddc03eb405d Mon Sep 17 00:00:00 2001 From: Leonardo de Moura Date: Sun, 5 Jan 2025 18:34:56 +0100 Subject: [PATCH 029/100] fix: E-matching thresholds in the `grind` tactic (#6536) This PR fixes different thresholds for controlling E-matching in the `grind` tactic. --- src/Lean/Meta/Tactic/Grind/EMatch.lean | 33 +++++++------ src/Lean/Meta/Tactic/Grind/Types.lean | 12 +++-- tests/lean/run/grind_offset.lean | 67 ++++++++++++++++++++++++++ 3 files changed, 94 insertions(+), 18 deletions(-) diff --git a/src/Lean/Meta/Tactic/Grind/EMatch.lean b/src/Lean/Meta/Tactic/Grind/EMatch.lean index ca6df2b7a07c..64b716c698bc 100644 --- a/src/Lean/Meta/Tactic/Grind/EMatch.lean +++ b/src/Lean/Meta/Tactic/Grind/EMatch.lean @@ -137,7 +137,8 @@ private partial def processMatch (c : Choice) (p : Expr) (e : Expr) : M Unit := let mut curr := e repeat let n ← getENode curr - if n.generation <= maxGeneration + -- Remark: we use `<` because the instance generation is the maximum term generation + 1 + if n.generation < maxGeneration -- uses heterogeneous equality or is the root of its congruence class && (n.heqProofs || n.isCongrRoot) && eqvFunctions pFn curr.getAppFn @@ -155,7 +156,7 @@ private partial def processOffset (c : Choice) (pArg : Expr) (k : Nat) (e : Expr let mut curr := e repeat let n ← getENode curr - if n.generation <= maxGeneration then + if n.generation < maxGeneration then if let some (eArg, k') ← isOffset? curr |>.run then if k' < k then let c := c.updateGen n.generation @@ -166,14 +167,14 @@ private partial def processOffset (c : Choice) (pArg : Expr) (k : Nat) (e : Expr else if k' > k then let eArg' := mkNatAdd eArg (mkNatLit (k' - k)) let eArg' ← shareCommon (← canon eArg') - internalize eArg' n.generation + internalize eArg' (n.generation+1) if let some c ← matchArg? c pArg eArg' |>.run then pushChoice (c.updateGen n.generation) else if let some k' ← evalNat curr |>.run then if k' >= k then let eArg' := mkNatLit (k' - k) let eArg' ← shareCommon (← canon eArg') - internalize eArg' n.generation + internalize eArg' (n.generation+1) if let some c ← matchArg? c pArg eArg' |>.run then pushChoice (c.updateGen n.generation) curr ← getNext curr @@ -186,7 +187,7 @@ private def processContinue (c : Choice) (p : Expr) : M Unit := do let maxGeneration ← getMaxGeneration for app in apps do let n ← getENode app - if n.generation <= maxGeneration + if n.generation < maxGeneration && (n.heqProofs || n.isCongrRoot) then if let some c ← matchArgs? c p app |>.run then let gen := n.generation @@ -214,7 +215,7 @@ private def addNewInstance (origin : Origin) (proof : Expr) (generation : Nat) : if Match.isMatchEqnTheorem (← getEnv) origin.key then prop ← annotateMatchEqnType prop trace[grind.ematch.instance] "{← origin.pp}: {prop}" - addTheoremInstance proof prop generation + addTheoremInstance proof prop (generation+1) /-- After processing a (multi-)pattern, use the choice assignment to instantiate the proof. @@ -262,17 +263,18 @@ where isDefEq x val /-- Process choice stack until we don't have more choices to be processed. -/ -private partial def processChoices : M Unit := do - unless (← get).choiceStack.isEmpty do +private def processChoices : M Unit := do + let maxGeneration ← getMaxGeneration + while !(← get).choiceStack.isEmpty do checkSystem "ematch" if (← checkMaxInstancesExceeded) then return () let c ← modifyGet fun s : State => (s.choiceStack.head!, { s with choiceStack := s.choiceStack.tail! }) - match c.cnstrs with - | [] => instantiateTheorem c - | .match p e :: cnstrs => processMatch { c with cnstrs } p e - | .offset p k e :: cnstrs => processOffset { c with cnstrs } p k e - | .continue p :: cnstrs => processContinue { c with cnstrs } p - processChoices + if c.gen < maxGeneration then + match c.cnstrs with + | [] => instantiateTheorem c + | .match p e :: cnstrs => processMatch { c with cnstrs } p e + | .offset p k e :: cnstrs => processOffset { c with cnstrs } p k e + | .continue p :: cnstrs => processContinue { c with cnstrs } p private def main (p : Expr) (cnstrs : List Cnstr) : M Unit := do let some apps := (← getThe Goal).appMap.find? p.toHeadIndex @@ -327,7 +329,7 @@ def ematch : GoalM Unit := do let go (thms newThms : PArray EMatchTheorem) : EMatch.M Unit := do withReader (fun ctx => { ctx with useMT := true }) <| ematchTheorems thms withReader (fun ctx => { ctx with useMT := false }) <| ematchTheorems newThms - if (← checkMaxInstancesExceeded) then + if (← checkMaxInstancesExceeded <||> checkMaxEmatchExceeded) then return () else go (← get).thms (← get).newThms |>.run' @@ -335,6 +337,7 @@ def ematch : GoalM Unit := do thms := s.thms ++ s.newThms newThms := {} gmt := s.gmt + 1 + numEmatch := s.numEmatch + 1 } /-- Performs one round of E-matching, and assert new instances. -/ diff --git a/src/Lean/Meta/Tactic/Grind/Types.lean b/src/Lean/Meta/Tactic/Grind/Types.lean index 92f7cd363293..80b934fc186d 100644 --- a/src/Lean/Meta/Tactic/Grind/Types.lean +++ b/src/Lean/Meta/Tactic/Grind/Types.lean @@ -68,7 +68,7 @@ instance : Hashable CongrTheoremCacheKey where hash a := mixHash (unsafe ptrAddrUnsafe a.f).toUInt64 (hash a.numArgs) /-- State for the `GrindM` monad. -/ -structure CoreState where +structure State where canon : Canon.State := {} /-- `ShareCommon` (aka `Hashconsing`) state. -/ scState : ShareCommon.State.{0} ShareCommon.objectFactory := ShareCommon.State.mk _ @@ -88,7 +88,7 @@ private opaque MethodsRefPointed : NonemptyType.{0} private def MethodsRef : Type := MethodsRefPointed.type instance : Nonempty MethodsRef := MethodsRefPointed.property -abbrev GrindM := ReaderT MethodsRef $ ReaderT Context $ StateRefT CoreState MetaM +abbrev GrindM := ReaderT MethodsRef $ ReaderT Context $ StateRefT State MetaM /-- Returns the user-defined configuration options -/ def getConfig : GrindM Grind.Config := @@ -207,7 +207,6 @@ structure ENode where generation : Nat := 0 /-- Modification time -/ mt : Nat := 0 - -- TODO: see Lean 3 implementation deriving Inhabited, Repr def ENode.isCongrRoot (n : ENode) := @@ -375,6 +374,9 @@ structure Goal where thmMap : EMatchTheorems /-- Number of theorem instances generated so far -/ numInstances : Nat := 0 + /-- Number of E-matching rounds performed in this goal so far. -/ + -- Remark: it is always equal to `gmt` in the current implementation. + numEmatch : Nat := 0 /-- (pre-)instances found so far. It includes instances that failed to be instantiated. -/ preInstances : PreInstanceSet := {} /-- new facts to be processed. -/ @@ -418,6 +420,10 @@ def addTheoremInstance (proof : Expr) (prop : Expr) (generation : Nat) : GoalM U def checkMaxInstancesExceeded : GoalM Bool := do return (← get).numInstances >= (← getConfig).instances +/-- Returns `true` if the maximum number of E-matching rounds has been reached. -/ +def checkMaxEmatchExceeded : GoalM Bool := do + return (← get).numEmatch >= (← getConfig).ematch + /-- Returns `some n` if `e` has already been "internalized" into the Otherwise, returns `none`s. diff --git a/tests/lean/run/grind_offset.lean b/tests/lean/run/grind_offset.lean index 7711ece39be1..946f4759a867 100644 --- a/tests/lean/run/grind_offset.lean +++ b/tests/lean/run/grind_offset.lean @@ -85,3 +85,70 @@ info: [grind.assert] foo (c + 1) = a example : foo (c + 1) = a → c = b + 1 → a = g (foo b) := by fail_if_success grind sorry + +set_option trace.grind.assert false + +/-- +info: [grind.ematch.instance] f.eq_2: f (x + 99).succ = g (f (x + 99)) +[grind.ematch.instance] f.eq_2: f (x + 98).succ = g (f (x + 98)) +-/ +#guard_msgs (info) in +example : f (x + 100) = a → a = b := by + fail_if_success grind (ematch := 2) + sorry + +/-- +info: [grind.ematch.instance] f.eq_2: f (x + 99).succ = g (f (x + 99)) +[grind.ematch.instance] f.eq_2: f (x + 98).succ = g (f (x + 98)) +[grind.ematch.instance] f.eq_2: f (x + 97).succ = g (f (x + 97)) +[grind.ematch.instance] f.eq_2: f (x + 96).succ = g (f (x + 96)) +[grind.ematch.instance] f.eq_2: f (x + 95).succ = g (f (x + 95)) +-/ +#guard_msgs (info) in +example : f (x + 100) = a → a = b := by + fail_if_success grind (ematch := 5) + sorry + +/-- +info: [grind.ematch.instance] f.eq_2: f (x + 99).succ = g (f (x + 99)) +[grind.ematch.instance] f.eq_2: f (x + 98).succ = g (f (x + 98)) +[grind.ematch.instance] f.eq_2: f (x + 97).succ = g (f (x + 97)) +[grind.ematch.instance] f.eq_2: f (x + 96).succ = g (f (x + 96)) +-/ +#guard_msgs (info) in +example : f (x + 100) = a → a = b := by + fail_if_success grind (ematch := 100) (instances := 4) + sorry + +/-- +info: [grind.ematch.instance] f.eq_2: f (y + 9).succ = g (f (y + 9)) +[grind.ematch.instance] f.eq_2: f (x + 99).succ = g (f (x + 99)) +[grind.ematch.instance] f.eq_2: f (x + 98).succ = g (f (x + 98)) +[grind.ematch.instance] f.eq_2: f (y + 8).succ = g (f (y + 8)) +[grind.ematch.instance] f.eq_2: f (y + 7).succ = g (f (y + 7)) +-/ +#guard_msgs (info) in +example : f (x + 100) = a → f (y + 10) = c → a = b := by + fail_if_success grind (ematch := 100) (instances := 5) + sorry + +/-- +info: [grind.ematch.instance] f.eq_2: f (x + 99).succ = g (f (x + 99)) +[grind.ematch.instance] f.eq_2: f (x + 98).succ = g (f (x + 98)) +-/ +#guard_msgs (info) in +example : f (x + 100) = a → a = b := by + fail_if_success grind (gen := 2) + sorry + +/-- +info: [grind.ematch.instance] f.eq_2: f (x + 99).succ = g (f (x + 99)) +[grind.ematch.instance] f.eq_2: f (x + 98).succ = g (f (x + 98)) +[grind.ematch.instance] f.eq_2: f (x + 97).succ = g (f (x + 97)) +[grind.ematch.instance] f.eq_2: f (x + 96).succ = g (f (x + 96)) +[grind.ematch.instance] f.eq_2: f (x + 95).succ = g (f (x + 95)) +-/ +#guard_msgs (info) in +example : f (x + 100) = a → a = b := by + fail_if_success grind (gen := 5) + sorry From fd091d1dfe03b821ac2888921dfe7d13b6096b4a Mon Sep 17 00:00:00 2001 From: Leonardo de Moura Date: Sun, 5 Jan 2025 20:35:52 +0100 Subject: [PATCH 030/100] feat: pattern normalization in the `grind` tactic (#6538) This PR ensures patterns provided by users are normalized. See new test to understand why this is needed. --- src/Lean/Meta/Tactic/Grind/EMatchTheorem.lean | 10 ++++-- src/Lean/Meta/Tactic/Grind/Internalize.lean | 3 +- src/Lean/Meta/Tactic/Grind/Main.lean | 16 ++-------- src/Lean/Meta/Tactic/Grind/SimpUtil.lean | 32 +++++++++++++++++++ src/Lean/Meta/Tactic/Grind/Util.lean | 7 ++++ tests/lean/run/grind_ematch2.lean | 27 ++++++++++++++++ 6 files changed, 78 insertions(+), 17 deletions(-) create mode 100644 src/Lean/Meta/Tactic/Grind/SimpUtil.lean diff --git a/src/Lean/Meta/Tactic/Grind/EMatchTheorem.lean b/src/Lean/Meta/Tactic/Grind/EMatchTheorem.lean index fd175a908c7b..01be831b1f53 100644 --- a/src/Lean/Meta/Tactic/Grind/EMatchTheorem.lean +++ b/src/Lean/Meta/Tactic/Grind/EMatchTheorem.lean @@ -37,9 +37,10 @@ def isOffsetPattern? (pat : Expr) : Option (Expr × Nat) := Id.run do let .lit (.natVal k) := k | none return some (pat, k) -def preprocessPattern (pat : Expr) : MetaM Expr := do +def preprocessPattern (pat : Expr) (normalizePattern := true) : MetaM Expr := do let pat ← instantiateMVars pat let pat ← unfoldReducible pat + let pat ← if normalizePattern then normalize pat else pure pat let pat ← detectOffsets pat let pat ← foldProjs pat return pat @@ -424,12 +425,15 @@ def mkEMatchTheorem (declName : Name) (numParams : Nat) (patterns : List Expr) : /-- Given theorem with name `declName` and type of the form `∀ (a_1 ... a_n), lhs = rhs`, creates an E-matching pattern for it using `addEMatchTheorem n [lhs]` + +If `normalizePattern` is true, it applies the `grind` simplification theorems and simprocs to the +pattern. -/ -def mkEMatchEqTheorem (declName : Name) : MetaM EMatchTheorem := do +def mkEMatchEqTheorem (declName : Name) (normalizePattern := true) : MetaM EMatchTheorem := do let info ← getConstInfo declName let (numParams, patterns) ← forallTelescopeReducing info.type fun xs type => do let_expr Eq _ lhs _ := type | throwError "invalid E-matching equality theorem, conclusion must be an equality{indentExpr type}" - let lhs ← preprocessPattern lhs + let lhs ← preprocessPattern lhs normalizePattern return (xs.size, [lhs.abstract xs]) mkEMatchTheorem declName numParams patterns diff --git a/src/Lean/Meta/Tactic/Grind/Internalize.lean b/src/Lean/Meta/Tactic/Grind/Internalize.lean index 63c2fc81ccae..b0b93b6755e7 100644 --- a/src/Lean/Meta/Tactic/Grind/Internalize.lean +++ b/src/Lean/Meta/Tactic/Grind/Internalize.lean @@ -71,7 +71,8 @@ private partial def addMatchEqns (f : Expr) (generation : Nat) : GoalM Unit := d if (← get).matchEqNames.contains declName then return () modify fun s => { s with matchEqNames := s.matchEqNames.insert declName } for eqn in (← Match.getEquationsFor declName).eqnNames do - activateTheorem (← mkEMatchEqTheorem eqn) generation + -- We disable pattern normalization to prevent the `match`-expression to be reduced. + activateTheorem (← mkEMatchEqTheorem eqn (normalizePattern := false)) generation private partial def activateTheoremPatterns (fName : Name) (generation : Nat) : GoalM Unit := do if let some (thms, thmMap) := (← get).thmMap.retrieve? fName then diff --git a/src/Lean/Meta/Tactic/Grind/Main.lean b/src/Lean/Meta/Tactic/Grind/Main.lean index 8ee85469bdc3..480301d5bb52 100644 --- a/src/Lean/Meta/Tactic/Grind/Main.lean +++ b/src/Lean/Meta/Tactic/Grind/Main.lean @@ -6,7 +6,6 @@ Authors: Leonardo de Moura prelude import Init.Grind.Lemmas import Lean.Meta.Tactic.Util -import Lean.Meta.Tactic.Simp.Simproc import Lean.Meta.Tactic.Grind.RevertAll import Lean.Meta.Tactic.Grind.PropagatorAttr import Lean.Meta.Tactic.Grind.Proj @@ -15,7 +14,7 @@ import Lean.Meta.Tactic.Grind.Util import Lean.Meta.Tactic.Grind.Inv import Lean.Meta.Tactic.Grind.Intro import Lean.Meta.Tactic.Grind.EMatch -import Lean.Meta.Tactic.Grind.DoNotSimp +import Lean.Meta.Tactic.Grind.SimpUtil namespace Lean.Meta.Grind @@ -35,21 +34,12 @@ def mkMethods (fallback : Fallback) : CoreM Methods := do prop e } -private def getGrindSimprocs : MetaM Simprocs := do - let s ← grindNormSimprocExt.getSimprocs - let s ← addDoNotSimp s - return s - def GrindM.run (x : GrindM α) (mainDeclName : Name) (config : Grind.Config) (fallback : Fallback) : MetaM α := do let scState := ShareCommon.State.mk _ let (falseExpr, scState) := ShareCommon.State.shareCommon scState (mkConst ``False) let (trueExpr, scState) := ShareCommon.State.shareCommon scState (mkConst ``True) - let thms ← grindNormExt.getTheorems - let simprocs := #[(← getGrindSimprocs), (← Simp.getSEvalSimprocs)] - let simp ← Simp.mkContext - (config := { arith := true }) - (simpTheorems := #[thms]) - (congrTheorems := (← getSimpCongrTheorems)) + let simprocs ← Grind.getSimprocs + let simp ← Grind.getSimpContext x (← mkMethods fallback).toMethodsRef { mainDeclName, config, simprocs, simp } |>.run' { scState, trueExpr, falseExpr } private def mkGoal (mvarId : MVarId) : GrindM Goal := do diff --git a/src/Lean/Meta/Tactic/Grind/SimpUtil.lean b/src/Lean/Meta/Tactic/Grind/SimpUtil.lean new file mode 100644 index 000000000000..878b0fc59356 --- /dev/null +++ b/src/Lean/Meta/Tactic/Grind/SimpUtil.lean @@ -0,0 +1,32 @@ +/- +Copyright (c) 2025 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.Meta.Tactic.Simp.Simproc +import Lean.Meta.Tactic.Grind.Simp +import Lean.Meta.Tactic.Grind.DoNotSimp + +namespace Lean.Meta.Grind + +/-- Returns the array of simprocs used by `grind`. -/ +protected def getSimprocs : MetaM (Array Simprocs) := do + let s ← grindNormSimprocExt.getSimprocs + let s ← addDoNotSimp s + return #[s, (← Simp.getSEvalSimprocs)] + +/-- Returns the simplification context used by `grind`. -/ +protected def getSimpContext : MetaM Simp.Context := do + let thms ← grindNormExt.getTheorems + Simp.mkContext + (config := { arith := true }) + (simpTheorems := #[thms]) + (congrTheorems := (← getSimpCongrTheorems)) + +@[export lean_grind_normalize] +def normalizeImp (e : Expr) : MetaM Expr := do + let (r, _) ← Meta.simp e (← Grind.getSimpContext) (← Grind.getSimprocs) + return r.expr + +end Lean.Meta.Grind diff --git a/src/Lean/Meta/Tactic/Grind/Util.lean b/src/Lean/Meta/Tactic/Grind/Util.lean index 0805b4dd7fb8..4fdb42055d1b 100644 --- a/src/Lean/Meta/Tactic/Grind/Util.lean +++ b/src/Lean/Meta/Tactic/Grind/Util.lean @@ -140,4 +140,11 @@ def normalizeLevels (e : Expr) : CoreM Expr := do | _ => return .continue Core.transform e (pre := pre) +/-- +Normalizes the given expression using the `grind` simplification theorems and simprocs. +This function is used for normalzing E-matching patterns. Note that it does not return a proof. +-/ +@[extern "lean_grind_normalize"] -- forward definition +opaque normalize (e : Expr) : MetaM Expr + end Lean.Meta.Grind diff --git a/tests/lean/run/grind_ematch2.lean b/tests/lean/run/grind_ematch2.lean index 613972c80e11..62a4ba2161f1 100644 --- a/tests/lean/run/grind_ematch2.lean +++ b/tests/lean/run/grind_ematch2.lean @@ -66,3 +66,30 @@ info: [grind.ematch.instance] fx: f a (f a a) = a #guard_msgs (info) in example : a = b₁ → c = f b₁ b₂ → f a c ≠ a → a = b₂ → False := by grind + + +namespace pattern_normalization +opaque g : Nat → Nat +@[grind_norm] theorem gthm : g x = x := sorry +opaque f : Nat → Nat → Nat +theorem fthm : f (g x) x = x := sorry +-- The following pattern should be normalized by `grind`. Otherwise, we will not find any instance during E-matching. +/-- +info: [grind.ematch.pattern] fthm: [f #0 #0] +-/ +#guard_msgs (info) in +grind_pattern fthm => f (g x) x + +/-- +info: [grind.assert] f x y = b +[grind.assert] y = x +[grind.assert] ¬b = x +[grind.ematch.instance] fthm: f (g y) y = y +[grind.assert] f y y = y +-/ +#guard_msgs (info) in +set_option trace.grind.assert true in +example : f (g x) y = b → y = x → b = x := by + grind + +end pattern_normalization From 675244de76633aa3ca93e83a17729c403648bb86 Mon Sep 17 00:00:00 2001 From: Leonardo de Moura Date: Sun, 5 Jan 2025 13:38:05 -0800 Subject: [PATCH 031/100] feat: `[grind_eq]` attribute for the `grind` tactic (#6539) This PR introduces the `[grind_eq]` attribute, designed to annotate equational theorems and functions for heuristic instantiations in the `grind` tactic. When applied to an equational theorem, the `[grind_eq]` attribute instructs the `grind` tactic to automatically use the annotated theorem to instantiate patterns during proof search. If applied to a function, it marks all equational theorems associated with that function. ```lean @[grind_eq] theorem foo_idempotent : foo (foo x) = foo x := ... @[grind_eq] def f (a : Nat) := match a with | 0 => 10 | x+1 => g (f x) ``` In the example above, the `grind` tactic will add instances of the theorem `foo_idempotent` to the local context whenever it encounters the pattern `foo (foo x)`. Similarly, functions annotated with `[grind_eq]` will propagate this annotation to their associated equational theorems. --- src/Lean/Meta/Tactic/Grind/Ctor.lean | 8 +- src/Lean/Meta/Tactic/Grind/EMatchTheorem.lean | 33 +++++++- tests/lean/run/grind_eq.lean | 76 +++++++++++++++++++ tests/lean/run/grind_pattern1.lean | 7 ++ 4 files changed, 120 insertions(+), 4 deletions(-) create mode 100644 tests/lean/run/grind_eq.lean diff --git a/src/Lean/Meta/Tactic/Grind/Ctor.lean b/src/Lean/Meta/Tactic/Grind/Ctor.lean index 027218e5b6c9..d7971d751ba6 100644 --- a/src/Lean/Meta/Tactic/Grind/Ctor.lean +++ b/src/Lean/Meta/Tactic/Grind/Ctor.lean @@ -9,12 +9,16 @@ import Lean.Meta.Tactic.Grind.Types namespace Lean.Meta.Grind private partial def propagateInjEqs (eqs : Expr) (proof : Expr) : GoalM Unit := do + -- Remark: we must use `shareCommon` before using `pushEq` and `pushHEq`. + -- This is needed because the result type of the injection theorem may allocate match_expr eqs with | And left right => propagateInjEqs left (.proj ``And 0 proof) propagateInjEqs right (.proj ``And 1 proof) - | Eq _ lhs rhs => pushEq lhs rhs proof - | HEq _ lhs _ rhs => pushHEq lhs rhs proof + | Eq _ lhs rhs => + pushEq (← shareCommon lhs) (← shareCommon rhs) proof + | HEq _ lhs _ rhs => + pushHEq (← shareCommon lhs) (← shareCommon rhs) proof | _ => trace[grind.issues] "unexpected injectivity theorem result type{indentExpr eqs}" return () diff --git a/src/Lean/Meta/Tactic/Grind/EMatchTheorem.lean b/src/Lean/Meta/Tactic/Grind/EMatchTheorem.lean index 01be831b1f53..99e8a0a5182a 100644 --- a/src/Lean/Meta/Tactic/Grind/EMatchTheorem.lean +++ b/src/Lean/Meta/Tactic/Grind/EMatchTheorem.lean @@ -11,6 +11,7 @@ import Lean.Util.FoldConsts import Lean.Util.CollectFVars import Lean.Meta.Basic import Lean.Meta.InferType +import Lean.Meta.Eqns import Lean.Meta.Tactic.Grind.Util namespace Lean.Meta.Grind @@ -241,7 +242,7 @@ private partial def go (pattern : Expr) (root := false) : M Expr := do else return mkOffsetPattern e k let some f := getPatternFn? pattern - | throwError "invalid pattern, (non-forbidden) application expected" + | throwError "invalid pattern, (non-forbidden) application expected{indentExpr pattern}" assert! f.isConst || f.isFVar saveSymbol f.toHeadIndex let mut args := pattern.getAppArgs @@ -432,7 +433,11 @@ pattern. def mkEMatchEqTheorem (declName : Name) (normalizePattern := true) : MetaM EMatchTheorem := do let info ← getConstInfo declName let (numParams, patterns) ← forallTelescopeReducing info.type fun xs type => do - let_expr Eq _ lhs _ := type | throwError "invalid E-matching equality theorem, conclusion must be an equality{indentExpr type}" + let lhs ← match_expr type with + | Eq _ lhs _ => pure lhs + | Iff lhs _ => pure lhs + | HEq _ lhs _ _ => pure lhs + | _ => throwError "invalid E-matching equality theorem, conclusion must be an equality{indentExpr type}" let lhs ← preprocessPattern lhs normalizePattern return (xs.size, [lhs.abstract xs]) mkEMatchTheorem declName numParams patterns @@ -455,4 +460,28 @@ def addEMatchEqTheorem (declName : Name) : MetaM Unit := do def getEMatchTheorems : CoreM EMatchTheorems := return ematchTheoremsExt.getState (← getEnv) +private def addGrindEqAttr (declName : Name) (attrKind : AttributeKind) : MetaM Unit := do + if (← getConstInfo declName).isTheorem then + ematchTheoremsExt.add (← mkEMatchEqTheorem declName) attrKind + else if let some eqns ← getEqnsFor? declName then + for eqn in eqns do + ematchTheoremsExt.add (← mkEMatchEqTheorem eqn) attrKind + else + throwError "`[grind_eq]` attribute can only be applied to equational theorems or function definitions" + +builtin_initialize + registerBuiltinAttribute { + name := `grind_eq + descr := + "The `[grind_eq]` attribute is used to annotate equational theorems and functions.\ + When applied to an equational theorem, it marks the theorem for use in heuristic instantiations by the `grind` tactic.\ + When applied to a function, it automatically annotates the equational theorems associated with that function.\ + The `grind` tactic utilizes annotated theorems to add instances of matching patterns into the local context during proof search.\ + For example, if a theorem `@[grind_eq] theorem foo_idempotent : foo (foo x) = foo x` is annotated,\ + `grind` will add an instance of this theorem to the local context whenever it encounters the pattern `foo (foo x)`." + applicationTime := .afterCompilation + add := fun declName _ attrKind => + addGrindEqAttr declName attrKind |>.run' {} + } + end Lean.Meta.Grind diff --git a/tests/lean/run/grind_eq.lean b/tests/lean/run/grind_eq.lean new file mode 100644 index 000000000000..b8d7c22d624b --- /dev/null +++ b/tests/lean/run/grind_eq.lean @@ -0,0 +1,76 @@ +opaque g : Nat → Nat + +@[grind_eq] def f (a : Nat) := + match a with + | 0 => 10 + | x+1 => g (f x) + +set_option grind.debug true +set_option grind.debug.proofs true + +set_option trace.grind.ematch.instance true +set_option trace.grind.assert true + +/-- +info: [grind.assert] f (y + 1) = a +[grind.assert] ¬a = g (f y) +[grind.ematch.instance] f.eq_2: f y.succ = g (f y) +[grind.assert] f (y + 1) = g (f y) +-/ +#guard_msgs (info) in +example : f (y + 1) = a → a = g (f y):= by + grind + +@[grind_eq] def app (xs ys : List α) := + match xs with + | [] => ys + | x::xs => x :: app xs ys + +/-- +info: [grind.assert] app [1, 2] ys = xs +[grind.assert] ¬xs = 1 :: 2 :: ys +[grind.ematch.instance] app.eq_2: app [1, 2] ys = 1 :: app [2] ys +[grind.assert] app [1, 2] ys = 1 :: app [2] ys +[grind.ematch.instance] app.eq_2: app [2] ys = 2 :: app [] ys +[grind.assert] app [2] ys = 2 :: app [] ys +[grind.ematch.instance] app.eq_1: app [] ys = ys +[grind.assert] app [] ys = ys +-/ +#guard_msgs (info) in +example : app [1, 2] ys = xs → xs = 1::2::ys := by + grind + +opaque p : Nat → Nat → Prop +opaque q : Nat → Prop + +@[grind_eq] theorem pq : p x x ↔ q x := by sorry + +/-- +info: [grind.assert] p a a +[grind.assert] ¬q a +[grind.ematch.instance] pq: p a a ↔ q a +[grind.assert] p a a = q a +-/ +#guard_msgs (info) in +example : p a a → q a := by + grind + +opaque appV (xs : Vector α n) (ys : Vector α m) : Vector α (n + m) := + Vector.append xs ys + +@[grind_eq] +theorem appV_assoc (a : Vector α n) (b : Vector α m) (c : Vector α n') : + HEq (appV a (appV b c)) (appV (appV a b) c) := sorry + +/-- +info: [grind.assert] x1 = appV a b +[grind.assert] x2 = appV x1 c +[grind.assert] x3 = appV b c +[grind.assert] x4 = appV a x3 +[grind.assert] ¬HEq x2 x4 +[grind.ematch.instance] appV_assoc: HEq (appV a (appV b c)) (appV (appV a b) c) +[grind.assert] HEq (appV a (appV b c)) (appV (appV a b) c) +-/ +#guard_msgs (info) in +example : x1 = appV a b → x2 = appV x1 c → x3 = appV b c → x4 = appV a x3 → HEq x2 x4 := by + grind diff --git a/tests/lean/run/grind_pattern1.lean b/tests/lean/run/grind_pattern1.lean index 5ca785dac106..f41e3bb96f74 100644 --- a/tests/lean/run/grind_pattern1.lean +++ b/tests/lean/run/grind_pattern1.lean @@ -112,3 +112,10 @@ grind_pattern hThm1 => plus a c /-- info: [grind.ematch.pattern] hThm1: [plus #2 #1, plus #2 #3] -/ #guard_msgs in grind_pattern hThm1 => plus a c, plus a b + +/-- +error: invalid pattern, (non-forbidden) application expected + #4 ∧ #3 +-/ +#guard_msgs in +grind_pattern And.imp_left => a ∧ b From 76f883b9995f684663e3106b08defec46f7a7f4c Mon Sep 17 00:00:00 2001 From: Sebastian Ullrich Date: Sun, 5 Jan 2025 21:57:56 +0000 Subject: [PATCH 032/100] fix: remove unused `-static-libgcc` MinGW linker arg (#6535) This PR avoids a linker warning on Windows. The argument may have been superfluous ever since the initial implementation. --- script/prepare-llvm-mingw.sh | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/script/prepare-llvm-mingw.sh b/script/prepare-llvm-mingw.sh index f3e1b0a65a20..0e2745971801 100644 --- a/script/prepare-llvm-mingw.sh +++ b/script/prepare-llvm-mingw.sh @@ -43,7 +43,7 @@ echo -n " -DCMAKE_C_COMPILER=$PWD/stage1/bin/clang.exe -DCMAKE_C_COMPILER_WORKS= echo -n " -DSTAGE0_CMAKE_C_COMPILER=clang -DSTAGE0_CMAKE_CXX_COMPILER=clang++" echo -n " -DLEAN_EXTRA_CXX_FLAGS='--sysroot $PWD/llvm -idirafter /clang64/include/'" echo -n " -DLEANC_INTERNAL_FLAGS='--sysroot ROOT -nostdinc -isystem ROOT/include/clang' -DLEANC_CC=ROOT/bin/clang.exe" -echo -n " -DLEANC_INTERNAL_LINKER_FLAGS='-L ROOT/lib -static-libgcc -Wl,-Bstatic -lgmp $(pkg-config --static --libs libuv) -lunwind -Wl,-Bdynamic -fuse-ld=lld'" +echo -n " -DLEANC_INTERNAL_LINKER_FLAGS='-L ROOT/lib -Wl,-Bstatic -lgmp $(pkg-config --static --libs libuv) -lunwind -Wl,-Bdynamic -fuse-ld=lld'" # when not using the above flags, link GMP dynamically/as usual. Always link ICU dynamically. echo -n " -DLEAN_EXTRA_LINKER_FLAGS='-lgmp $(pkg-config --libs libuv) -lucrtbase'" # do not set `LEAN_CC` for tests From 2ed77f3b262bb2b2c509e95ac01e6fe161ed40fd Mon Sep 17 00:00:00 2001 From: Leonardo de Moura Date: Sun, 5 Jan 2025 19:05:20 -0800 Subject: [PATCH 033/100] feat: attribute `[grind]` (#6545) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR introduces the parametric attribute `[grind]` for annotating theorems and definitions. It also replaces `[grind_eq]` with `[grind =]`. For definitions, `[grind]` is equivalent to `[grind =]`. The new attribute supports the following variants: - **`[grind =]`**: Uses the left-hand side of the theorem's conclusion as the pattern for E-matching. - **`[grind =_]`**: Uses the right-hand side of the theorem's conclusion as the pattern for E-matching. - **`[grind _=_]`**: Creates two patterns. One for the left-hand side and one for the right-hand side. - **`[grind →]`**: Searches for (multi-)patterns in the theorem's antecedents, stopping once a usable multi-pattern is found. - **`[grind ←]`**: Searches for (multi-)patterns in the theorem's conclusion, stopping once a usable multi-pattern is found. - **`[grind]`**: Searches for (multi-)patterns in both the theorem's conclusion and antecedents. It starts with the conclusion and stops once a usable multi-pattern is found. The `grind_pattern` command remains available for cases where these attributes do not yield the desired result. --- src/Init/Grind/Tactics.lean | 12 + src/Lean/Meta/Tactic/Grind.lean | 1 + src/Lean/Meta/Tactic/Grind/EMatchTheorem.lean | 240 +++++++++++++++--- tests/lean/run/grind_ematch1.lean | 145 +++++++++++ tests/lean/run/grind_eq.lean | 10 +- tests/lean/run/grind_pattern1.lean | 2 +- 6 files changed, 375 insertions(+), 35 deletions(-) diff --git a/src/Init/Grind/Tactics.lean b/src/Init/Grind/Tactics.lean index 71b9455a93a2..ea7e0b657c6c 100644 --- a/src/Init/Grind/Tactics.lean +++ b/src/Init/Grind/Tactics.lean @@ -6,6 +6,18 @@ Authors: Leonardo de Moura prelude import Init.Tactics +namespace Lean.Parser.Attr + +syntax grindEq := "=" +syntax grindEqBoth := "_=_" +syntax grindEqRhs := "=_" +syntax grindBwd := "←" +syntax grindFwd := "→" + +syntax (name := grind) "grind" (grindEq <|> grindBwd <|> grindFwd <|> grindEqBoth <|> grindEqRhs)? : attr + +end Lean.Parser.Attr + namespace Lean.Grind /-- The configuration for `grind`. diff --git a/src/Lean/Meta/Tactic/Grind.lean b/src/Lean/Meta/Tactic/Grind.lean index 01ee792bbec6..022695fb35ba 100644 --- a/src/Lean/Meta/Tactic/Grind.lean +++ b/src/Lean/Meta/Tactic/Grind.lean @@ -34,6 +34,7 @@ builtin_initialize registerTraceClass `grind.eqc builtin_initialize registerTraceClass `grind.internalize builtin_initialize registerTraceClass `grind.ematch builtin_initialize registerTraceClass `grind.ematch.pattern +builtin_initialize registerTraceClass `grind.ematch.pattern.search builtin_initialize registerTraceClass `grind.ematch.instance builtin_initialize registerTraceClass `grind.ematch.instance.assignment builtin_initialize registerTraceClass `grind.issues diff --git a/src/Lean/Meta/Tactic/Grind/EMatchTheorem.lean b/src/Lean/Meta/Tactic/Grind/EMatchTheorem.lean index 99e8a0a5182a..9605c8e116e7 100644 --- a/src/Lean/Meta/Tactic/Grind/EMatchTheorem.lean +++ b/src/Lean/Meta/Tactic/Grind/EMatchTheorem.lean @@ -5,6 +5,7 @@ Authors: Leonardo de Moura -/ prelude import Init.Grind.Util +import Init.Grind.Tactics import Lean.HeadIndex import Lean.PrettyPrinter import Lean.Util.FoldConsts @@ -218,16 +219,18 @@ private def getPatternFn? (pattern : Expr) : Option Expr := /-- Returns a bit-mask `mask` s.t. `mask[i]` is true if the the corresponding argument is -- a type or type former, or +- a type (that is not a proposition) or type former, or - a proof, or - an instance implicit argument When `mask[i]`, we say the corresponding argument is a "support" argument. -/ -private def getPatternFunMask (f : Expr) (numArgs : Nat) : MetaM (Array Bool) := do +def getPatternSupportMask (f : Expr) (numArgs : Nat) : MetaM (Array Bool) := do forallBoundedTelescope (← inferType f) numArgs fun xs _ => do xs.mapM fun x => do - if (← isTypeFormer x <||> isProof x) then + if (← isProp x) then + return false + else if (← isTypeFormer x <||> isProof x) then return true else return (← x.fvarId!.getDecl).binderInfo matches .instImplicit @@ -246,7 +249,7 @@ private partial def go (pattern : Expr) (root := false) : M Expr := do assert! f.isConst || f.isFVar saveSymbol f.toHeadIndex let mut args := pattern.getAppArgs - let supportMask ← getPatternFunMask f args.size + let supportMask ← getPatternSupportMask f args.size for i in [:args.size] do let arg := args[i]! let isSupport := supportMask[i]?.getD false @@ -278,6 +281,9 @@ def main (patterns : List Expr) : MetaM (List Expr × List HeadIndex × Std.Hash let (patterns, s) ← patterns.mapM go |>.run {} return (patterns, s.symbols.toList, s.bvarsFound) +def normalizePattern (e : Expr) : M Expr := do + go e + end NormalizePattern /-- @@ -403,26 +409,50 @@ private def ppParamsAt (proof : Expr) (numParams : Nat) (paramPos : List Nat) : addMessageContextFull msg /-- -Creates an E-matching theorem for `declName` with `numParams` parameters, and the given set of patterns. +Creates an E-matching theorem for a theorem with proof `proof`, `numParams` parameters, and the given set of patterns. Pattern variables are represented using de Bruijn indices. -/ -def mkEMatchTheorem (declName : Name) (numParams : Nat) (patterns : List Expr) : MetaM EMatchTheorem := do - let .thmInfo info ← getConstInfo declName - | throwError "`{declName}` is not a theorem, you cannot assign patterns to non-theorems for the `grind` tactic" - let us := info.levelParams.map mkLevelParam - let proof := mkConst declName us +def mkEMatchTheoremCore (origin : Origin) (levelParams : Array Name) (numParams : Nat) (proof : Expr) (patterns : List Expr) : MetaM EMatchTheorem := do let (patterns, symbols, bvarFound) ← NormalizePattern.main patterns - assert! symbols.all fun s => s matches .const _ trace[grind.ematch.pattern] "{MessageData.ofConst proof}: {patterns.map ppPattern}" if let .missing pos ← checkCoverage proof numParams bvarFound then let pats : MessageData := m!"{patterns.map ppPattern}" - throwError "invalid pattern(s) for `{declName}`{indentD pats}\nthe following theorem parameters cannot be instantiated:{indentD (← ppParamsAt proof numParams pos)}" + throwError "invalid pattern(s) for `{← origin.pp}`{indentD pats}\nthe following theorem parameters cannot be instantiated:{indentD (← ppParamsAt proof numParams pos)}" return { proof, patterns, numParams, symbols - levelParams := #[] - origin := .decl declName + levelParams, origin } +private def getProofFor (declName : Name) : CoreM Expr := do + let .thmInfo info ← getConstInfo declName + | throwError "`{declName}` is not a theorem" + let us := info.levelParams.map mkLevelParam + return mkConst declName us + +/-- +Creates an E-matching theorem for `declName` with `numParams` parameters, and the given set of patterns. +Pattern variables are represented using de Bruijn indices. +-/ +def mkEMatchTheorem (declName : Name) (numParams : Nat) (patterns : List Expr) : MetaM EMatchTheorem := do + mkEMatchTheoremCore (.decl declName) #[] numParams (← getProofFor declName) patterns + +/-- +Given a theorem with proof `proof` and type of the form `∀ (a_1 ... a_n), lhs = rhs`, +creates an E-matching pattern for it using `addEMatchTheorem n [lhs]` +If `normalizePattern` is true, it applies the `grind` simplification theorems and simprocs to the pattern. +-/ +def mkEMatchEqTheoremCore (origin : Origin) (levelParams : Array Name) (proof : Expr) (normalizePattern : Bool) (useLhs : Bool) : MetaM EMatchTheorem := do + let (numParams, patterns) ← forallTelescopeReducing (← inferType proof) fun xs type => do + let (lhs, rhs) ← match_expr type with + | Eq _ lhs rhs => pure (lhs, rhs) + | Iff lhs rhs => pure (lhs, rhs) + | HEq _ lhs _ rhs => pure (lhs, rhs) + | _ => throwError "invalid E-matching equality theorem, conclusion must be an equality{indentExpr type}" + let pat := if useLhs then lhs else rhs + let pat ← preprocessPattern pat normalizePattern + return (xs.size, [pat.abstract xs]) + mkEMatchTheoremCore origin levelParams numParams proof patterns + /-- Given theorem with name `declName` and type of the form `∀ (a_1 ... a_n), lhs = rhs`, creates an E-matching pattern for it using `addEMatchTheorem n [lhs]` @@ -430,17 +460,8 @@ creates an E-matching pattern for it using `addEMatchTheorem n [lhs]` If `normalizePattern` is true, it applies the `grind` simplification theorems and simprocs to the pattern. -/ -def mkEMatchEqTheorem (declName : Name) (normalizePattern := true) : MetaM EMatchTheorem := do - let info ← getConstInfo declName - let (numParams, patterns) ← forallTelescopeReducing info.type fun xs type => do - let lhs ← match_expr type with - | Eq _ lhs _ => pure lhs - | Iff lhs _ => pure lhs - | HEq _ lhs _ _ => pure lhs - | _ => throwError "invalid E-matching equality theorem, conclusion must be an equality{indentExpr type}" - let lhs ← preprocessPattern lhs normalizePattern - return (xs.size, [lhs.abstract xs]) - mkEMatchTheorem declName numParams patterns +def mkEMatchEqTheorem (declName : Name) (normalizePattern := true) (useLhs : Bool := true) : MetaM EMatchTheorem := do + mkEMatchEqTheoremCore (.decl declName) #[] (← getProofFor declName) normalizePattern useLhs /-- Adds an E-matching theorem to the environment. @@ -460,18 +481,177 @@ def addEMatchEqTheorem (declName : Name) : MetaM Unit := do def getEMatchTheorems : CoreM EMatchTheorems := return ematchTheoremsExt.getState (← getEnv) -private def addGrindEqAttr (declName : Name) (attrKind : AttributeKind) : MetaM Unit := do +private inductive TheoremKind where + | eqLhs | eqRhs | eqBoth | fwd | bwd | default + deriving Inhabited, BEq + +private def TheoremKind.toAttribute : TheoremKind → String + | .eqLhs => "[grind =]" + | .eqRhs => "[grind =_]" + | .eqBoth => "[grind _=_]" + | .fwd => "[grind →]" + | .bwd => "[grind ←]" + | .default => "[grind]" + +private def TheoremKind.explainFailure : TheoremKind → String + | .eqLhs => "failed to find pattern in the left-hand side of the theorem's conclusion" + | .eqRhs => "failed to find pattern in the right-hand side of the theorem's conclusion" + | .eqBoth => unreachable! -- eqBoth is a macro + | .fwd => "failed to find patterns in the antecedents of the theorem" + | .bwd => "failed to find patterns in the theorem's conclusion" + | .default => "failed to find patterns" + +/-- Returns the types of `xs` that are propositions. -/ +private def getPropTypes (xs : Array Expr) : MetaM (Array Expr) := + xs.filterMapM fun x => do + let type ← inferType x + if (← isProp type) then return some type else return none + +/-- State for the (pattern) `CollectorM` monad -/ +private structure Collector.State where + /-- Pattern found so far. -/ + patterns : Array Expr := #[] + done : Bool := false + +private structure Collector.Context where + proof : Expr + xs : Array Expr + +/-- Monad for collecting patterns for a theorem. -/ +private abbrev CollectorM := ReaderT Collector.Context $ StateRefT Collector.State NormalizePattern.M + +/-- Similar to `getPatternFn?`, but operates on expressions that do not contain loose de Bruijn variables. -/ +private def isPatternFnCandidate (f : Expr) : CollectorM Bool := do + match f with + | .const declName _ => return !isForbidden declName + | .fvar .. => return !(← read).xs.contains f + | _ => return false + +private def addNewPattern (p : Expr) : CollectorM Unit := do + trace[grind.ematch.pattern.search] "found pattern: {ppPattern p}" + let bvarsFound := (← getThe NormalizePattern.State).bvarsFound + let done := (← checkCoverage (← read).proof (← read).xs.size bvarsFound) matches .ok + if done then + trace[grind.ematch.pattern.search] "found full coverage" + modify fun s => { s with patterns := s.patterns.push p, done } + +private partial def collect (e : Expr) : CollectorM Unit := do + if (← get).done then return () + match e with + | .app .. => + let f := e.getAppFn + if (← isPatternFnCandidate f) then + let saved ← getThe NormalizePattern.State + try + trace[grind.ematch.pattern.search] "candidate: {e}" + let p := e.abstract (← read).xs + unless p.hasLooseBVars do + trace[grind.ematch.pattern.search] "skip, does not contain pattern variables" + return () + let p ← NormalizePattern.normalizePattern p + if saved.bvarsFound.size < (← getThe NormalizePattern.State).bvarsFound.size then + addNewPattern p + return () + trace[grind.ematch.pattern.search] "skip, no new variables covered" + -- restore state and continue search + set saved + catch _ => + -- restore state and continue search + trace[grind.ematch.pattern.search] "skip, exception during normalization" + set saved + let args := e.getAppArgs + for arg in args, flag in (← NormalizePattern.getPatternSupportMask f args.size) do + unless flag do + collect arg + | .forallE _ d b _ => + if (← pure e.isArrow <&&> isProp d <&&> isProp b) then + collect d + collect b + | _ => return () + +private def collectPatterns? (proof : Expr) (xs : Array Expr) (searchPlaces : Array Expr) : MetaM (Option (List Expr × List HeadIndex)) := do + let go : CollectorM (Option (List Expr)) := do + for place in searchPlaces do + let place ← preprocessPattern place + collect place + if (← get).done then + return some ((← get).patterns.toList) + return none + let (some ps, s) ← go { proof, xs } |>.run' {} |>.run {} + | return none + return some (ps, s.symbols.toList) + +private def mkEMatchTheoremWithKind? (origin : Origin) (levelParams : Array Name) (proof : Expr) (kind : TheoremKind) : MetaM (Option EMatchTheorem) := do + if kind == .eqLhs then + return (← mkEMatchEqTheoremCore origin levelParams proof (normalizePattern := false) (useLhs := true)) + else if kind == .eqRhs then + return (← mkEMatchEqTheoremCore origin levelParams proof (normalizePattern := false) (useLhs := false)) + let type ← inferType proof + forallTelescopeReducing type fun xs type => do + let searchPlaces ← match kind with + | .fwd => + let ps ← getPropTypes xs + if ps.isEmpty then + throwError "invalid `grind` forward theorem, theorem `{← origin.pp}` does not have proposional hypotheses" + pure ps + | .bwd => pure #[type] + | .default => pure <| #[type] ++ (← getPropTypes xs) + | _ => unreachable! + go xs searchPlaces +where + go (xs : Array Expr) (searchPlaces : Array Expr) : MetaM (Option EMatchTheorem) := do + let some (patterns, symbols) ← collectPatterns? proof xs searchPlaces + | return none + let numParams := xs.size + trace[grind.ematch.pattern] "{← origin.pp}: {patterns.map ppPattern}" + return some { + proof, patterns, numParams, symbols + levelParams, origin + } + +private def getKind (stx : Syntax) : TheoremKind := + if stx[1].isNone then + .default + else if stx[1][0].getKind == ``Parser.Attr.grindEq then + .eqLhs + else if stx[1][0].getKind == ``Parser.Attr.grindFwd then + .fwd + else if stx[1][0].getKind == ``Parser.Attr.grindEqRhs then + .eqRhs + else if stx[1][0].getKind == ``Parser.Attr.grindEqBoth then + .eqBoth + else + .bwd + +private def addGrindEqAttr (declName : Name) (attrKind : AttributeKind) (useLhs := true) : MetaM Unit := do if (← getConstInfo declName).isTheorem then - ematchTheoremsExt.add (← mkEMatchEqTheorem declName) attrKind + ematchTheoremsExt.add (← mkEMatchEqTheorem declName (normalizePattern := true) (useLhs := useLhs)) attrKind else if let some eqns ← getEqnsFor? declName then + unless useLhs do + throwError "`{declName}` is a definition, you must only use the left-hand side for extracting patterns" for eqn in eqns do ematchTheoremsExt.add (← mkEMatchEqTheorem eqn) attrKind else throwError "`[grind_eq]` attribute can only be applied to equational theorems or function definitions" +private def addGrindAttr (declName : Name) (attrKind : AttributeKind) (thmKind : TheoremKind) : MetaM Unit := do + if thmKind == .eqLhs then + addGrindEqAttr declName attrKind (useLhs := true) + else if thmKind == .eqRhs then + addGrindEqAttr declName attrKind (useLhs := false) + else if thmKind == .eqBoth then + addGrindEqAttr declName attrKind (useLhs := true) + addGrindEqAttr declName attrKind (useLhs := false) + else if !(← getConstInfo declName).isTheorem then + addGrindEqAttr declName attrKind + else + let some thm ← mkEMatchTheoremWithKind? (.decl declName) #[] (← getProofFor declName) thmKind + | throwError "`@{thmKind.toAttribute} theorem {declName}` {thmKind.explainFailure}, consider using different options or the `grind_pattern` command" + ematchTheoremsExt.add thm attrKind + builtin_initialize registerBuiltinAttribute { - name := `grind_eq + name := `grind descr := "The `[grind_eq]` attribute is used to annotate equational theorems and functions.\ When applied to an equational theorem, it marks the theorem for use in heuristic instantiations by the `grind` tactic.\ @@ -480,8 +660,8 @@ builtin_initialize For example, if a theorem `@[grind_eq] theorem foo_idempotent : foo (foo x) = foo x` is annotated,\ `grind` will add an instance of this theorem to the local context whenever it encounters the pattern `foo (foo x)`." applicationTime := .afterCompilation - add := fun declName _ attrKind => - addGrindEqAttr declName attrKind |>.run' {} + add := fun declName stx attrKind => do + addGrindAttr declName attrKind (getKind stx) |>.run' {} } end Lean.Meta.Grind diff --git a/tests/lean/run/grind_ematch1.lean b/tests/lean/run/grind_ematch1.lean index 5bce28cc44a3..9523882c486b 100644 --- a/tests/lean/run/grind_ematch1.lean +++ b/tests/lean/run/grind_ematch1.lean @@ -69,3 +69,148 @@ info: [grind.ematch.instance] Rtrans: R a d → R d e → R a e #guard_msgs (info) in example : R a b → R b c → R c d → R d e → R a d := by grind + + +namespace using_grind_fwd + +opaque S : Nat → Nat → Prop + +/-- +error: `@[grind →] theorem using_grind_fwd.StransBad` failed to find patterns in the antecedents of the theorem, consider using different options or the `grind_pattern` command +-/ +#guard_msgs (error) in +@[grind→] theorem StransBad (a b c d : Nat) : S a b ∨ R a b → S b c → S a c ∧ S b d := sorry + + +set_option trace.grind.ematch.pattern.search true in +/-- +info: [grind.ematch.pattern.search] candidate: S a b +[grind.ematch.pattern.search] found pattern: S #4 #3 +[grind.ematch.pattern.search] candidate: R a b +[grind.ematch.pattern.search] skip, no new variables covered +[grind.ematch.pattern.search] candidate: S b c +[grind.ematch.pattern.search] found pattern: S #3 #2 +[grind.ematch.pattern.search] found full coverage +[grind.ematch.pattern] Strans: [S #4 #3, S #3 #2] +-/ +#guard_msgs (info) in +@[grind→] theorem Strans (a b c : Nat) : S a b ∨ R a b → S b c → S a c := sorry + +/-- +info: [grind.ematch.instance] Strans: S a b ∨ R a b → S b c → S a c +-/ +#guard_msgs (info) in +example : S a b → S b c → S a c := by + grind + +end using_grind_fwd + +namespace using_grind_bwd + +opaque P : Nat → Prop +opaque Q : Nat → Prop +opaque f : Nat → Nat → Nat + +/-- +info: [grind.ematch.pattern] pqf: [P (f #2 #1)] +-/ +#guard_msgs (info) in +@[grind←] theorem pqf : Q x → P (f x y) := sorry + +/-- +info: [grind.ematch.instance] pqf: Q a → P (f a b) +-/ +#guard_msgs (info) in +example : Q 0 → Q 1 → Q 2 → Q 3 → ¬ P (f a b) → a = 1 → False := by + grind + +end using_grind_bwd + +namespace using_grind_fwd2 + +opaque P : Nat → Prop +opaque Q : Nat → Prop +opaque f : Nat → Nat → Nat + +/-- +error: `@[grind →] theorem using_grind_fwd2.pqfBad` failed to find patterns in the antecedents of the theorem, consider using different options or the `grind_pattern` command +-/ +#guard_msgs (error) in +@[grind→] theorem pqfBad : Q x → P (f x y) := sorry + +/-- +info: [grind.ematch.pattern] pqf: [Q #1] +-/ +#guard_msgs (info) in +@[grind→] theorem pqf : Q x → P (f x x) := sorry + +/-- +info: [grind.ematch.instance] pqf: Q 3 → P (f 3 3) +[grind.ematch.instance] pqf: Q 2 → P (f 2 2) +[grind.ematch.instance] pqf: Q 1 → P (f 1 1) +[grind.ematch.instance] pqf: Q 0 → P (f 0 0) +-/ +#guard_msgs (info) in +example : Q 0 → Q 1 → Q 2 → Q 3 → ¬ P (f a a) → a = 1 → False := by + grind + +end using_grind_fwd2 + +namespace using_grind_mixed + +opaque P : Nat → Nat → Prop +opaque Q : Nat → Nat → Prop + +/-- +error: `@[grind →] theorem using_grind_mixed.pqBad1` failed to find patterns in the antecedents of the theorem, consider using different options or the `grind_pattern` command +-/ +#guard_msgs (error) in +@[grind→] theorem pqBad1 : P x y → Q x z := sorry + +/-- +error: `@[grind ←] theorem using_grind_mixed.pqBad2` failed to find patterns in the theorem's conclusion, consider using different options or the `grind_pattern` command +-/ +#guard_msgs (error) in +@[grind←] theorem pqBad2 : P x y → Q x z := sorry + + +/-- +info: [grind.ematch.pattern] pqBad: [Q #3 #1, P #3 #2] +-/ +#guard_msgs (info) in +@[grind] theorem pqBad : P x y → Q x z := sorry + +example : P a b → Q a c := by + grind + +end using_grind_mixed + + +namespace using_grind_rhs + +opaque f : Nat → Nat +opaque g : Nat → Nat → Nat + +/-- +info: [grind.ematch.pattern] fq: [g #0 (f #0)] +-/ +#guard_msgs (info) in +@[grind =_] +theorem fq : f x = g x (f x) := sorry + +end using_grind_rhs + +namespace using_grind_lhs_rhs + +opaque f : Nat → Nat +opaque g : Nat → Nat → Nat + +/-- +info: [grind.ematch.pattern] fq: [f #0] +[grind.ematch.pattern] fq: [g #0 (g #0 #0)] +-/ +#guard_msgs (info) in +@[grind _=_] +theorem fq : f x = g x (g x x) := sorry + +end using_grind_lhs_rhs diff --git a/tests/lean/run/grind_eq.lean b/tests/lean/run/grind_eq.lean index b8d7c22d624b..333aab7593bf 100644 --- a/tests/lean/run/grind_eq.lean +++ b/tests/lean/run/grind_eq.lean @@ -1,6 +1,8 @@ opaque g : Nat → Nat -@[grind_eq] def f (a : Nat) := +set_option trace.Meta.debug true + +@[grind] def f (a : Nat) := match a with | 0 => 10 | x+1 => g (f x) @@ -21,7 +23,7 @@ info: [grind.assert] f (y + 1) = a example : f (y + 1) = a → a = g (f y):= by grind -@[grind_eq] def app (xs ys : List α) := +@[grind] def app (xs ys : List α) := match xs with | [] => ys | x::xs => x :: app xs ys @@ -43,7 +45,7 @@ example : app [1, 2] ys = xs → xs = 1::2::ys := by opaque p : Nat → Nat → Prop opaque q : Nat → Prop -@[grind_eq] theorem pq : p x x ↔ q x := by sorry +@[grind =] theorem pq : p x x ↔ q x := by sorry /-- info: [grind.assert] p a a @@ -58,7 +60,7 @@ example : p a a → q a := by opaque appV (xs : Vector α n) (ys : Vector α m) : Vector α (n + m) := Vector.append xs ys -@[grind_eq] +@[grind =] theorem appV_assoc (a : Vector α n) (b : Vector α m) (c : Vector α n') : HEq (appV a (appV b c)) (appV (appV a b) c) := sorry diff --git a/tests/lean/run/grind_pattern1.lean b/tests/lean/run/grind_pattern1.lean index f41e3bb96f74..d827e1efa0d5 100644 --- a/tests/lean/run/grind_pattern1.lean +++ b/tests/lean/run/grind_pattern1.lean @@ -20,7 +20,7 @@ grind_pattern List.mem_concat_self => a ∈ xs ++ [a] def foo (x : Nat) := x + x /-- -error: `foo` is not a theorem, you cannot assign patterns to non-theorems for the `grind` tactic +error: `foo` is not a theorem -/ #guard_msgs in grind_pattern foo => x + x From 78ddee911287bf7b0069a695b01001db2654a765 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Mon, 6 Jan 2025 17:37:01 +1100 Subject: [PATCH 034/100] feat: release checklist script (#6542) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR introduces a script that automates checking whether major downstream repositories have been updated for a new toolchain release. Sample output: ``` % ./release_checklist.py v4.16.0-rc1 Repository: Batteries ✅ On compatible toolchain (>= v4.16.0-rc1) ✅ Tag v4.16.0-rc1 exists Repository: lean4checker ✅ On compatible toolchain (>= v4.16.0-rc1) ✅ Tag v4.16.0-rc1 exists Repository: doc-gen4 ✅ On compatible toolchain (>= v4.16.0-rc1) ✅ Tag v4.16.0-rc1 exists Repository: Verso ❌ Not on target toolchain (needs ≥ v4.16.0-rc1, but main is on leanprover/lean4:v4.15.0) Repository: ProofWidgets4 ✅ On compatible toolchain (>= v4.16.0-rc1) Repository: Aesop ✅ On compatible toolchain (>= v4.16.0-rc1) ✅ Tag v4.16.0-rc1 exists Repository: import-graph ✅ On compatible toolchain (>= v4.16.0-rc1) ✅ Tag v4.16.0-rc1 exists Repository: plausible ✅ On compatible toolchain (>= v4.16.0-rc1) ✅ Tag v4.16.0-rc1 exists Repository: Mathlib ✅ On compatible toolchain (>= v4.16.0-rc1) ✅ Tag v4.16.0-rc1 exists Repository: REPL ❌ Not on target toolchain (needs ≥ v4.16.0-rc1, but master is on leanprover/lean4:v4.14.0) ``` --- doc/dev/release_checklist.md | 1 + script/release_checklist.py | 132 +++++++++++++++++++++++++++++++++++ script/release_repos.yml | 79 +++++++++++++++++++++ 3 files changed, 212 insertions(+) create mode 100755 script/release_checklist.py create mode 100644 script/release_repos.yml diff --git a/doc/dev/release_checklist.md b/doc/dev/release_checklist.md index e5b27e203012..82b39d029f01 100644 --- a/doc/dev/release_checklist.md +++ b/doc/dev/release_checklist.md @@ -97,6 +97,7 @@ We'll use `v4.6.0` as the intended release version as a running example. - Toolchain bump PR including updated Lake manifest - Create and push the tag - Merge the tag into `stable` +- Run `scripts/release_checklist.py v4.6.0` to check that everything is in order. - The `v4.6.0` section of `RELEASES.md` is out of sync between `releases/v4.6.0` and `master`. This should be reconciled: - Replace the `v4.6.0` section on `master` with the `v4.6.0` section on `releases/v4.6.0` diff --git a/script/release_checklist.py b/script/release_checklist.py new file mode 100755 index 000000000000..53f0b340b87d --- /dev/null +++ b/script/release_checklist.py @@ -0,0 +1,132 @@ +#!/usr/bin/env python3 + +import argparse +import yaml +import requests +import base64 +import subprocess +import sys +import os + +def parse_repos_config(file_path): + with open(file_path, "r") as f: + return yaml.safe_load(f)["repositories"] + +def get_github_token(): + try: + import subprocess + result = subprocess.run(['gh', 'auth', 'token'], capture_output=True, text=True) + if result.returncode == 0: + return result.stdout.strip() + except FileNotFoundError: + print("Warning: 'gh' CLI not found. Some API calls may be rate-limited.") + return None + +def get_branch_content(repo_url, branch, file_path, github_token): + api_url = repo_url.replace("https://github.com/", "https://api.github.com/repos/") + f"/contents/{file_path}?ref={branch}" + headers = {'Authorization': f'token {github_token}'} if github_token else {} + response = requests.get(api_url, headers=headers) + if response.status_code == 200: + content = response.json().get("content", "") + content = content.replace("\n", "") + try: + return base64.b64decode(content).decode('utf-8').strip() + except Exception: + return None + return None + +def tag_exists(repo_url, tag_name, github_token): + api_url = repo_url.replace("https://github.com/", "https://api.github.com/repos/") + f"/git/refs/tags/{tag_name}" + headers = {'Authorization': f'token {github_token}'} if github_token else {} + response = requests.get(api_url, headers=headers) + return response.status_code == 200 + +def is_merged_into_stable(repo_url, tag_name, stable_branch, github_token): + # First get the commit SHA for the tag + api_base = repo_url.replace("https://github.com/", "https://api.github.com/repos/") + headers = {'Authorization': f'token {github_token}'} if github_token else {} + + # Get tag's commit SHA + tag_response = requests.get(f"{api_base}/git/refs/tags/{tag_name}", headers=headers) + if tag_response.status_code != 200: + return False + tag_sha = tag_response.json()['object']['sha'] + + # Get commits on stable branch containing this SHA + commits_response = requests.get( + f"{api_base}/commits?sha={stable_branch}&per_page=100", + headers=headers + ) + if commits_response.status_code != 200: + return False + + # Check if any commit in stable's history matches our tag's SHA + stable_commits = [commit['sha'] for commit in commits_response.json()] + return tag_sha in stable_commits + +def parse_version(version_str): + # Remove 'v' prefix and split into components + # Handle Lean toolchain format (leanprover/lean4:v4.x.y) + if ':' in version_str: + version_str = version_str.split(':')[1] + version = version_str.lstrip('v') + # Handle release candidates by removing -rc part for comparison + version = version.split('-')[0] + return tuple(map(int, version.split('.'))) + +def is_version_gte(version1, version2): + """Check if version1 >= version2""" + return parse_version(version1) >= parse_version(version2) + +def is_release_candidate(version): + return "-rc" in version + +def main(): + github_token = get_github_token() + + if len(sys.argv) != 2: + print("Usage: python3 release_checklist.py ") + sys.exit(1) + + toolchain = sys.argv[1] + + with open(os.path.join(os.path.dirname(__file__), "release_repos.yml")) as f: + repos = yaml.safe_load(f)["repositories"] + + for repo in repos: + name = repo["name"] + url = repo["url"] + branch = repo["branch"] + check_stable = repo["stable-branch"] + check_tag = repo.get("toolchain-tag", True) + + print(f"\nRepository: {name}") + + # Check if branch is on at least the target toolchain + lean_toolchain_content = get_branch_content(url, branch, "lean-toolchain", github_token) + if lean_toolchain_content is None: + print(f" ❌ No lean-toolchain file found in {branch} branch") + continue + + on_target_toolchain = is_version_gte(lean_toolchain_content.strip(), toolchain) + if not on_target_toolchain: + print(f" ❌ Not on target toolchain (needs ≥ {toolchain}, but {branch} is on {lean_toolchain_content.strip()})") + continue + print(f" ✅ On compatible toolchain (>= {toolchain})") + + # Only check for tag if toolchain-tag is true + if check_tag: + if not tag_exists(url, toolchain, github_token): + print(f" ❌ Tag {toolchain} does not exist") + continue + print(f" ✅ Tag {toolchain} exists") + + # Only check merging into stable if stable-branch is true and not a release candidate + if check_stable and not is_release_candidate(toolchain): + if not is_merged_into_stable(url, toolchain, "stable", github_token): + print(f" ❌ Tag {toolchain} is not merged into stable") + continue + print(f" ✅ Tag {toolchain} is merged into stable") + +if __name__ == "__main__": + main() diff --git a/script/release_repos.yml b/script/release_repos.yml new file mode 100644 index 000000000000..10bc53e85e47 --- /dev/null +++ b/script/release_repos.yml @@ -0,0 +1,79 @@ +repositories: + - name: Batteries + url: https://github.com/leanprover-community/batteries + toolchain-tag: true + stable-branch: true + branch: main + dependencies: [] + + - name: lean4checker + url: https://github.com/leanprover/lean4checker + toolchain-tag: true + stable-branch: true + branch: master + dependencies: [] + + - name: doc-gen4 + url: https://github.com/leanprover/doc-gen4 + toolchain-tag: true + stable-branch: false + branch: main + dependencies: [] + + - name: Verso + url: https://github.com/leanprover/verso + toolchain-tag: true + stable-branch: false + branch: main + dependencies: [] + + - name: ProofWidgets4 + url: https://github.com/leanprover-community/ProofWidgets4 + toolchain-tag: false + stable-branch: false + branch: main + dependencies: + - Batteries + + - name: Aesop + url: https://github.com/leanprover-community/aesop + toolchain-tag: true + stable-branch: true + branch: master + dependencies: + - Batteries + + - name: import-graph + url: https://github.com/leanprover-community/import-graph + toolchain-tag: true + stable-branch: false + branch: main + dependencies: [] + + - name: plausible + url: https://github.com/leanprover-community/plausible + toolchain-tag: true + stable-branch: false + branch: main + dependencies: [] + + - name: Mathlib + url: https://github.com/leanprover-community/mathlib4 + toolchain-tag: true + stable-branch: true + branch: master + dependencies: + - Aesop + - ProofWidgets4 + - lean4checker + - Batteries + - doc-gen4 + - import-graph + + - name: REPL + url: https://github.com/leanprover-community/repl + toolchain-tag: true + stable-branch: true + branch: master + dependencies: + - Mathlib From 2c9641f621f8aecf69d8644d5b7f99c96cb947fa Mon Sep 17 00:00:00 2001 From: Bhavik Mehta Date: Mon, 6 Jan 2025 13:13:16 +0000 Subject: [PATCH 035/100] doc: modify aesop usage example of `omegaDefault` (#6549) This PR fixes #6548. --- src/Lean/Elab/Tactic/Omega/Frontend.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/Lean/Elab/Tactic/Omega/Frontend.lean b/src/Lean/Elab/Tactic/Omega/Frontend.lean index 9c407fda6186..e98390e07f1e 100644 --- a/src/Lean/Elab/Tactic/Omega/Frontend.lean +++ b/src/Lean/Elab/Tactic/Omega/Frontend.lean @@ -680,7 +680,7 @@ def omegaTactic (cfg : OmegaConfig) : TacticM Unit := do /-- The `omega` tactic, for resolving integer and natural linear arithmetic problems. This `TacticM Unit` frontend with default configuration can be used as an Aesop rule, for example via -the tactic call `aesop (add 50% tactic Lean.Omega.omegaDefault)`. -/ +the tactic call `aesop (add 50% tactic Lean.Elab.Tactic.Omega.omegaDefault)`. -/ def omegaDefault : TacticM Unit := omegaTactic {} @[builtin_tactic Lean.Parser.Tactic.omega] From 3ca3f848a8358d8bae0d1a27d53974518c850ec7 Mon Sep 17 00:00:00 2001 From: Leonardo de Moura Date: Mon, 6 Jan 2025 08:18:44 -0800 Subject: [PATCH 036/100] fix: avoid new tokens `_=_` and `=_` (#6554) This PR an issue introduced by the `[grind _=_]` attribute. --- src/Init/Grind/Tactics.lean | 6 +++--- tests/lean/run/grind_ematch1.lean | 3 +++ 2 files changed, 6 insertions(+), 3 deletions(-) diff --git a/src/Init/Grind/Tactics.lean b/src/Init/Grind/Tactics.lean index ea7e0b657c6c..7def9222d434 100644 --- a/src/Init/Grind/Tactics.lean +++ b/src/Init/Grind/Tactics.lean @@ -9,12 +9,12 @@ import Init.Tactics namespace Lean.Parser.Attr syntax grindEq := "=" -syntax grindEqBoth := "_=_" -syntax grindEqRhs := "=_" +syntax grindEqBoth := atomic("_" "=" "_") +syntax grindEqRhs := atomic("=" "_") syntax grindBwd := "←" syntax grindFwd := "→" -syntax (name := grind) "grind" (grindEq <|> grindBwd <|> grindFwd <|> grindEqBoth <|> grindEqRhs)? : attr +syntax (name := grind) "grind" (grindEqBoth <|> grindEqRhs <|> grindEq <|> grindBwd <|> grindFwd)? : attr end Lean.Parser.Attr diff --git a/tests/lean/run/grind_ematch1.lean b/tests/lean/run/grind_ematch1.lean index 9523882c486b..7f8a9c053b83 100644 --- a/tests/lean/run/grind_ematch1.lean +++ b/tests/lean/run/grind_ematch1.lean @@ -214,3 +214,6 @@ info: [grind.ematch.pattern] fq: [f #0] theorem fq : f x = g x (g x x) := sorry end using_grind_lhs_rhs + +-- the following should still work +#check _=_ From 8dec57987ab13244c67fe03edf004d9e288b4df1 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Tue, 7 Jan 2025 03:29:50 +1100 Subject: [PATCH 037/100] feat: `grind` tests for basic category theory (#6543) This PR adds additional tests for `grind`, demonstrating that we can automate some manual proofs from Mathlib's basic category theory library, with less reliance on Mathlib's `@[reassoc]` trick. In several places I've added bidirectional patterns for equational lemmas. I've updated some other files to use the new `@[grind_eq]` attribute (but left as is all cases where we are inspecting the info messages from `grind_pattern`). --------- Co-authored-by: Leonardo de Moura --- tests/lean/run/grind_cat.lean | 125 ++++++++++++++++++++++++++++++ tests/lean/run/grind_ematch1.lean | 9 +-- tests/lean/run/grind_ematch2.lean | 7 +- tests/lean/run/grind_shelf.lean | 30 +++++++ 4 files changed, 161 insertions(+), 10 deletions(-) create mode 100644 tests/lean/run/grind_cat.lean create mode 100644 tests/lean/run/grind_shelf.lean diff --git a/tests/lean/run/grind_cat.lean b/tests/lean/run/grind_cat.lean new file mode 100644 index 000000000000..9e501f97f97b --- /dev/null +++ b/tests/lean/run/grind_cat.lean @@ -0,0 +1,125 @@ +universe v v₁ v₂ v₃ u u₁ u₂ u₃ + +namespace CategoryTheory + +macro "cat_tac" : tactic => `(tactic| (intros; (try ext); grind)) + +class Category (obj : Type u) : Type max u (v + 1) where + Hom : obj → obj → Type v + /-- The identity morphism on an object. -/ + id : ∀ X : obj, Hom X X + /-- Composition of morphisms in a category, written `f ≫ g`. -/ + comp : ∀ {X Y Z : obj}, (Hom X Y) → (Hom Y Z) → (Hom X Z) + /-- Identity morphisms are left identities for composition. -/ + id_comp : ∀ {X Y : obj} (f : Hom X Y), comp (id X) f = f := by cat_tac + /-- Identity morphisms are right identities for composition. -/ + comp_id : ∀ {X Y : obj} (f : Hom X Y), comp f (id Y) = f := by cat_tac + /-- Composition in a category is associative. -/ + assoc : ∀ {W X Y Z : obj} (f : Hom W X) (g : Hom X Y) (h : Hom Y Z), comp (comp f g) h = comp f (comp g h) := by cat_tac + +infixr:10 " ⟶ " => Category.Hom +scoped notation "𝟙" => Category.id -- type as \b1 +scoped infixr:80 " ≫ " => Category.comp + +attribute [simp] Category.id_comp Category.comp_id Category.assoc + +attribute [grind =] Category.id_comp Category.comp_id +attribute [grind _=_] Category.assoc + +structure Functor (C : Type u₁) [Category.{v₁} C] (D : Type u₂) [Category.{v₂} D] : Type max v₁ v₂ u₁ u₂ where + /-- The action of a functor on objects. -/ + obj : C → D + /-- The action of a functor on morphisms. -/ + map : ∀ {X Y : C}, (X ⟶ Y) → ((obj X) ⟶ (obj Y)) + /-- A functor preserves identity morphisms. -/ + map_id : ∀ X : C, map (𝟙 X) = 𝟙 (obj X) := by cat_tac + /-- A functor preserves composition. -/ + map_comp : ∀ {X Y Z : C} (f : X ⟶ Y) (g : Y ⟶ Z), map (f ≫ g) = (map f) ≫ (map g) := by cat_tac + +attribute [simp] Functor.map_id Functor.map_comp + +attribute [grind =] Functor.map_id +attribute [grind _=_] Functor.map_comp + +variable {C : Type u₁} [Category.{v₁} C] {D : Type u₂} [Category.{v₂} D] {E : Type u₃} [Category.{v₃} E] +variable {F G H : Functor C D} + +namespace Functor + +def comp (F : Functor C D) (G : Functor D E) : Functor C E where + obj X := G.obj (F.obj X) + map f := G.map (F.map f) + -- Note `map_id` and `map_comp` are handled by `cat_tac`. + +variable {X Y : C} {G : Functor D E} + +@[simp, grind =] theorem comp_obj : (F.comp G).obj X = G.obj (F.obj X) := rfl +@[simp, grind =] theorem comp_map (f : X ⟶ Y) : (F.comp G).map f = G.map (F.map f) := rfl + +end Functor + +@[ext] +structure NatTrans [Category.{v₁, u₁} C] [Category.{v₂, u₂} D] (F G : Functor C D) : Type max u₁ v₂ where + /-- The component of a natural transformation. -/ + app : ∀ X : C, F.obj X ⟶ G.obj X + /-- The naturality square for a given morphism. -/ + naturality : ∀ ⦃X Y : C⦄ (f : X ⟶ Y), F.map f ≫ app Y = app X ≫ G.map f := by cat_tac + +attribute [simp, grind =] NatTrans.naturality + +namespace NatTrans + +variable {X : C} + +protected def id (F : Functor C D) : NatTrans F F where app X := 𝟙 (F.obj X) + +@[simp, grind =] theorem id_app : (NatTrans.id F).app X = 𝟙 (F.obj X) := rfl + +protected def vcomp (α : NatTrans F G) (β : NatTrans G H) : NatTrans F H where + app X := α.app X ≫ β.app X + -- `naturality` is now handled by `grind`; in Mathlib this relies on `@[reassoc]` attributes. + -- Manual proof: + -- rw [← Category.assoc] + -- rw [α.naturality f] + -- rw [Category.assoc] + -- rw [β.naturality f] + -- rw [← Category.assoc] + +@[simp, grind =] theorem vcomp_app (α : NatTrans F G) (β : NatTrans G H) (X : C) : + (α.vcomp β).app X = α.app X ≫ β.app X := rfl + +end NatTrans + +instance Functor.category : Category.{max u₁ v₂} (Functor C D) where + Hom F G := NatTrans F G + id F := NatTrans.id F + comp α β := NatTrans.vcomp α β + -- Here we're okay: all the proofs are handled by `cat_tac`. + +@[simp, grind =] +theorem id_app (F : Functor C D) (X : C) : (𝟙 F : F ⟶ F).app X = 𝟙 (F.obj X) := rfl + +@[simp, grind _=_] +theorem comp_app {F G H : Functor C D} (α : F ⟶ G) (β : G ⟶ H) (X : C) : + (α ≫ β).app X = α.app X ≫ β.app X := rfl + +theorem app_naturality {F G : Functor C (Functor D E)} (T : F ⟶ G) (X : C) {Y Z : D} (f : Y ⟶ Z) : + (F.obj X).map f ≫ (T.app X).app Z = (T.app X).app Y ≫ (G.obj X).map f := by + cat_tac + +theorem naturality_app {F G : Functor C (Functor D E)} (T : F ⟶ G) (Z : D) {X Y : C} (f : X ⟶ Y) : + (F.map f).app Z ≫ (T.app Y).app Z = (T.app X).app Z ≫ (G.map f).app Z := by + cat_tac -- this is done manually in Mathlib! + -- rw [← comp_app] + -- rw [T.naturality f] + -- rw [comp_app] + +open Category Functor NatTrans + +def hcomp {H I : Functor D E} (α : F ⟶ G) (β : H ⟶ I) : F.comp H ⟶ G.comp I where + app := fun X : C => β.app (F.obj X) ≫ I.map (α.app X) + -- `grind` can now handle `naturality`, while Mathlib does this manually: + -- rw [Functor.comp_map, Functor.comp_map, ← assoc, naturality, assoc, ← I.map_comp, naturality, + -- map_comp, assoc] + +end CategoryTheory diff --git a/tests/lean/run/grind_ematch1.lean b/tests/lean/run/grind_ematch1.lean index 7f8a9c053b83..b2a79778f8cd 100644 --- a/tests/lean/run/grind_ematch1.lean +++ b/tests/lean/run/grind_ematch1.lean @@ -1,5 +1,6 @@ set_option trace.grind.ematch.pattern true -grind_pattern Array.size_set => Array.set a i v h + +attribute [grind =] Array.size_set set_option grind.debug true @@ -21,12 +22,10 @@ example (as bs : Array α) (v : α) set_option trace.grind.ematch.instance true -grind_pattern Array.get_set_eq => a.set i v h -grind_pattern Array.get_set_ne => (a.set i v hi)[j] +attribute [grind =] Array.get_set_ne /-- -info: [grind.ematch.instance] Array.get_set_eq: (as.set i v ⋯)[i] = v -[grind.ematch.instance] Array.size_set: (as.set i v ⋯).size = as.size +info: [grind.ematch.instance] Array.size_set: (as.set i v ⋯).size = as.size [grind.ematch.instance] Array.get_set_ne: ∀ (hj : j < as.size), i ≠ j → (as.set i v ⋯)[j] = as[j] -/ #guard_msgs (info) in diff --git a/tests/lean/run/grind_ematch2.lean b/tests/lean/run/grind_ematch2.lean index 62a4ba2161f1..3341d3266af0 100644 --- a/tests/lean/run/grind_ematch2.lean +++ b/tests/lean/run/grind_ematch2.lean @@ -1,6 +1,4 @@ -grind_pattern Array.size_set => Array.set a i v h -grind_pattern Array.get_set_eq => a.set i v h -grind_pattern Array.get_set_ne => (a.set i v hi)[j] +attribute [grind =] Array.size_set Array.get_set_eq Array.get_set_ne set_option grind.debug true set_option trace.grind.ematch.pattern true @@ -57,8 +55,7 @@ example (as bs cs ds : Array α) (v₁ v₂ v₃ : α) grind opaque f (a b : α) : α := a -theorem fx : f x (f x x) = x := sorry -grind_pattern fx => f x (f x x) +@[grind =] theorem fx : f x (f x x) = x := sorry /-- info: [grind.ematch.instance] fx: f a (f a a) = a diff --git a/tests/lean/run/grind_shelf.lean b/tests/lean/run/grind_shelf.lean new file mode 100644 index 000000000000..8c5f6331764c --- /dev/null +++ b/tests/lean/run/grind_shelf.lean @@ -0,0 +1,30 @@ +class One (α : Type u) where + one : α + +instance (priority := 300) One.toOfNat1 {α} [One α] : OfNat α (nat_lit 1) where + ofNat := ‹One α›.1 + +class Shelf (α : Type u) where + act : α → α → α + self_distrib : ∀ {x y z : α}, act x (act y z) = act (act x y) (act x z) + +class UnitalShelf (α : Type u) extends Shelf α, One α where + one_act : ∀ a : α, act 1 a = a + act_one : ∀ a : α, act a 1 = a + +infixr:65 " ◃ " => Shelf.act + +-- Mathlib proof from UnitalShelf.act_act_self_eq +example {S} [UnitalShelf S] (x y : S) : (x ◃ y) ◃ x = x ◃ y := by + have h : (x ◃ y) ◃ x = (x ◃ y) ◃ (x ◃ 1) := by rw [UnitalShelf.act_one] + rw [h, ← Shelf.self_distrib, UnitalShelf.act_one] + +attribute [grind =] UnitalShelf.one_act UnitalShelf.act_one + +-- We actually want the reverse direction of `Shelf.self_distrib`, so don't use the `grind_eq` attribute. +grind_pattern Shelf.self_distrib => self.act (self.act x y) (self.act x z) + +-- Proof using `grind`: +example {S} [UnitalShelf S] (x y : S) : (x ◃ y) ◃ x = x ◃ y := by + have h : (x ◃ y) ◃ x = (x ◃ y) ◃ (x ◃ 1) := by grind + grind From db3ab39e05cc7ba343dae258b325690df89a985d Mon Sep 17 00:00:00 2001 From: Leonardo de Moura Date: Mon, 6 Jan 2025 13:31:12 -0800 Subject: [PATCH 038/100] feat: propagate implication in the `grind` tactic (#6556) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR adds propagators for implication to the `grind` tactic. It also disables the normalization rule: `(p → q) = (¬ p ∨ q)` --- src/Init/Grind/Lemmas.lean | 9 ++++ src/Init/Grind/Norm.lean | 5 +- src/Lean/Meta/Tactic/Grind/ForallProp.lean | 52 +++++++++++++++------ src/Lean/Meta/Tactic/Grind/Internalize.lean | 5 +- src/Lean/Meta/Tactic/Grind/Main.lean | 3 +- 5 files changed, 55 insertions(+), 19 deletions(-) diff --git a/src/Init/Grind/Lemmas.lean b/src/Init/Grind/Lemmas.lean index e89e66b8725b..5672b5618dae 100644 --- a/src/Init/Grind/Lemmas.lean +++ b/src/Init/Grind/Lemmas.lean @@ -35,6 +35,15 @@ theorem or_eq_of_eq_false_right {a b : Prop} (h : b = False) : (a ∨ b) = a := theorem eq_false_of_or_eq_false_left {a b : Prop} (h : (a ∨ b) = False) : a = False := by simp_all theorem eq_false_of_or_eq_false_right {a b : Prop} (h : (a ∨ b) = False) : b = False := by simp_all +/-! Implies -/ + +theorem imp_eq_of_eq_false_left {a b : Prop} (h : a = False) : (a → b) = True := by simp [h] +theorem imp_eq_of_eq_true_right {a b : Prop} (h : b = True) : (a → b) = True := by simp [h] +theorem imp_eq_of_eq_true_left {a b : Prop} (h : a = True) : (a → b) = b := by simp [h] + +theorem eq_true_of_imp_eq_false {a b : Prop} (h : (a → b) = False) : a = True := by simp_all +theorem eq_false_of_imp_eq_false {a b : Prop} (h : (a → b) = False) : b = False := by simp_all + /-! Not -/ theorem not_eq_of_eq_true {a : Prop} (h : a = True) : (Not a) = False := by simp [h] diff --git a/src/Init/Grind/Norm.lean b/src/Init/Grind/Norm.lean index d911cf634703..c3d8b62a5f8d 100644 --- a/src/Init/Grind/Norm.lean +++ b/src/Init/Grind/Norm.lean @@ -41,9 +41,10 @@ attribute [grind_norm] not_true -- False attribute [grind_norm] not_false_eq_true +-- Remark: we disabled the following normalization rule because we want this information when implementing splitting heuristics -- Implication as a clause -@[grind_norm↓] theorem imp_eq (p q : Prop) : (p → q) = (¬ p ∨ q) := by - by_cases p <;> by_cases q <;> simp [*] +-- @[grind_norm↓] theorem imp_eq (p q : Prop) : (p → q) = (¬ p ∨ q) := by +-- by_cases p <;> by_cases q <;> simp [*] -- And @[grind_norm↓] theorem not_and (p q : Prop) : (¬(p ∧ q)) = (¬p ∨ ¬q) := by diff --git a/src/Lean/Meta/Tactic/Grind/ForallProp.lean b/src/Lean/Meta/Tactic/Grind/ForallProp.lean index 9993f0e50e2d..ee46ee6eb915 100644 --- a/src/Lean/Meta/Tactic/Grind/ForallProp.lean +++ b/src/Lean/Meta/Tactic/Grind/ForallProp.lean @@ -15,20 +15,42 @@ If `parent` is a projection-application `proj_i c`, check whether the root of the equivalence class containing `c` is a constructor-application `ctor ... a_i ...`. If so, internalize the term `proj_i (ctor ... a_i ...)` and add the equality `proj_i (ctor ... a_i ...) = a_i`. -/ -def propagateForallProp (parent : Expr) : GoalM Unit := do - let .forallE n p q bi := parent | return () - trace[grind.debug.forallPropagator] "{parent}" - unless (← isEqTrue p) do return () - trace[grind.debug.forallPropagator] "isEqTrue, {parent}" - let h₁ ← mkEqTrueProof p - let qh₁ := q.instantiate1 (mkApp2 (mkConst ``of_eq_true) p h₁) - let r ← simp qh₁ - let q := mkLambda n bi p q - let q' := r.expr - internalize q' (← getGeneration parent) - trace[grind.debug.forallPropagator] "q': {q'} for{indentExpr parent}" - let h₂ ← r.getProof - let h := mkApp5 (mkConst ``Lean.Grind.forall_propagator) p q q' h₁ h₂ - pushEq parent q' h +def propagateForallPropUp (e : Expr) : GoalM Unit := do + let .forallE n p q bi := e | return () + trace[grind.debug.forallPropagator] "{e}" + if !q.hasLooseBVars then + propagateImpliesUp p q + else + unless (← isEqTrue p) do return + trace[grind.debug.forallPropagator] "isEqTrue, {e}" + let h₁ ← mkEqTrueProof p + let qh₁ := q.instantiate1 (mkApp2 (mkConst ``of_eq_true) p h₁) + let r ← simp qh₁ + let q := mkLambda n bi p q + let q' := r.expr + internalize q' (← getGeneration e) + trace[grind.debug.forallPropagator] "q': {q'} for{indentExpr e}" + let h₂ ← r.getProof + let h := mkApp5 (mkConst ``Lean.Grind.forall_propagator) p q q' h₁ h₂ + pushEq e q' h +where + propagateImpliesUp (a b : Expr) : GoalM Unit := do + unless (← alreadyInternalized b) do return () + if (← isEqFalse a) then + -- a = False → (a → b) = True + pushEqTrue e <| mkApp3 (mkConst ``Grind.imp_eq_of_eq_false_left) a b (← mkEqFalseProof a) + else if (← isEqTrue a) then + -- a = True → (a → b) = b + pushEq e b <| mkApp3 (mkConst ``Grind.imp_eq_of_eq_true_left) a b (← mkEqTrueProof a) + else if (← isEqTrue b) then + -- b = True → (a → b) = True + pushEqTrue e <| mkApp3 (mkConst ``Grind.imp_eq_of_eq_true_right) a b (← mkEqTrueProof b) + +def propagateImpliesDown (e : Expr) : GoalM Unit := do + if (← isEqFalse e) then + let .forallE _ a b _ := e | return () + let h ← mkEqFalseProof e + pushEqTrue a <| mkApp3 (mkConst ``Grind.eq_true_of_imp_eq_false) a b h + pushEqFalse b <| mkApp3 (mkConst ``Grind.eq_false_of_imp_eq_false) a b h end Lean.Meta.Grind diff --git a/src/Lean/Meta/Tactic/Grind/Internalize.lean b/src/Lean/Meta/Tactic/Grind/Internalize.lean index b0b93b6755e7..fe801763f326 100644 --- a/src/Lean/Meta/Tactic/Grind/Internalize.lean +++ b/src/Lean/Meta/Tactic/Grind/Internalize.lean @@ -95,11 +95,14 @@ partial def internalize (e : Expr) (generation : Nat) : GoalM Unit := do | .sort .. => return () | .fvar .. | .letE .. | .lam .. => mkENodeCore e (ctor := false) (interpreted := false) (generation := generation) - | .forallE _ d _ _ => + | .forallE _ d b _ => mkENodeCore e (ctor := false) (interpreted := false) (generation := generation) if (← isProp d <&&> isProp e) then internalize d generation registerParent e d + unless b.hasLooseBVars do + internalize b generation + registerParent e b propagateUp e | .lit .. | .const .. => mkENode e generation diff --git a/src/Lean/Meta/Tactic/Grind/Main.lean b/src/Lean/Meta/Tactic/Grind/Main.lean index 480301d5bb52..139b32ecebb7 100644 --- a/src/Lean/Meta/Tactic/Grind/Main.lean +++ b/src/Lean/Meta/Tactic/Grind/Main.lean @@ -23,12 +23,13 @@ def mkMethods (fallback : Fallback) : CoreM Methods := do return { fallback propagateUp := fun e => do - propagateForallProp e + propagateForallPropUp e let .const declName _ := e.getAppFn | return () propagateProjEq e if let some prop := builtinPropagators.up[declName]? then prop e propagateDown := fun e => do + propagateImpliesDown e let .const declName _ := e.getAppFn | return () if let some prop := builtinPropagators.down[declName]? then prop e From a4240294757960441c5dbad03e4718456cfbe967 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Tue, 7 Jan 2025 09:20:09 +1100 Subject: [PATCH 039/100] feat: Array lemma alignment; fold and map (#6546) This PR continues aligning `Array` and `Vector` lemmas with `List`, working on `fold` and `map` operations. --- src/Init/Data/Array/Lemmas.lean | 385 ++++++++++++++---------- src/Init/Data/List/Lemmas.lean | 490 ++++++++++++++++--------------- src/Init/Data/Vector/Basic.lean | 12 + src/Init/Data/Vector/Lemmas.lean | 57 +++- 4 files changed, 532 insertions(+), 412 deletions(-) diff --git a/src/Init/Data/Array/Lemmas.lean b/src/Init/Data/Array/Lemmas.lean index 2e6db2b2aa32..8bb29aae6fae 100644 --- a/src/Init/Data/Array/Lemmas.lean +++ b/src/Init/Data/Array/Lemmas.lean @@ -36,6 +36,8 @@ namespace Array @[simp] theorem mem_toList_iff (a : α) (l : Array α) : a ∈ l.toList ↔ a ∈ l := by cases l <;> simp +@[simp] theorem length_toList {l : Array α} : l.toList.length = l.size := rfl + /-! ### empty -/ @[simp] theorem empty_eq {xs : Array α} : #[] = xs ↔ xs = #[] := by @@ -1001,52 +1003,7 @@ private theorem beq_of_beq_singleton [BEq α] {a b : α} : #[a] == #[b] → a == /-! Content below this point has not yet been aligned with `List`. -/ -theorem singleton_eq_toArray_singleton (a : α) : #[a] = [a].toArray := rfl - -@[simp] theorem singleton_def (v : α) : singleton v = #[v] := rfl - --- This is a duplicate of `List.toArray_toList`. --- It's confusing to guess which namespace this theorem should live in, --- so we provide both. -@[simp] theorem toArray_toList (a : Array α) : a.toList.toArray = a := rfl - -@[simp] theorem length_toList {l : Array α} : l.toList.length = l.size := rfl - -@[simp] theorem mkEmpty_eq (α n) : @mkEmpty α n = #[] := rfl - -@[deprecated size_toArray (since := "2024-12-11")] -theorem size_mk (as : List α) : (Array.mk as).size = as.length := by simp [size] - -theorem foldrM_push [Monad m] (f : α → β → m β) (init : β) (arr : Array α) (a : α) : - (arr.push a).foldrM f init = f a init >>= arr.foldrM f := by - simp only [foldrM_eq_reverse_foldlM_toList, push_toList, List.reverse_append, List.reverse_cons, - List.reverse_nil, List.nil_append, List.singleton_append, List.foldlM_cons, List.foldlM_reverse] - -/-- -Variant of `foldrM_push` with `h : start = arr.size + 1` -rather than `(arr.push a).size` as the argument. --/ -@[simp] theorem foldrM_push' [Monad m] (f : α → β → m β) (init : β) (arr : Array α) (a : α) - {start} (h : start = arr.size + 1) : - (arr.push a).foldrM f init start = f a init >>= arr.foldrM f := by - simp [← foldrM_push, h] - -theorem foldr_push (f : α → β → β) (init : β) (arr : Array α) (a : α) : - (arr.push a).foldr f init = arr.foldr f (f a init) := foldrM_push .. - -/-- -Variant of `foldr_push` with the `h : start = arr.size + 1` -rather than `(arr.push a).size` as the argument. --/ -@[simp] theorem foldr_push' (f : α → β → β) (init : β) (arr : Array α) (a : α) {start} - (h : start = arr.size + 1) : (arr.push a).foldr f init start = arr.foldr f (f a init) := - foldrM_push' _ _ _ _ h - -/-- A more efficient version of `arr.toList.reverse`. -/ -@[inline] def toListRev (arr : Array α) : List α := arr.foldl (fun l t => t :: l) [] - -@[simp] theorem toListRev_eq (arr : Array α) : arr.toListRev = arr.toList.reverse := by - rw [toListRev, ← foldl_toList, ← List.foldr_reverse, List.foldr_cons_nil] +/-! ### map -/ theorem mapM_eq_foldlM [Monad m] [LawfulMonad m] (f : α → m β) (arr : Array α) : arr.mapM f = arr.foldlM (fun bs a => bs.push <$> f a) #[] := by @@ -1070,6 +1027,11 @@ where induction l generalizing arr <;> simp [*] simp [H] +@[simp] theorem _root_.List.map_toArray (f : α → β) (l : List α) : + l.toArray.map f = (l.map f).toArray := by + apply ext' + simp + @[simp] theorem size_map (f : α → β) (arr : Array α) : (arr.map f).size = arr.size := by simp only [← length_toList] simp @@ -1079,6 +1041,128 @@ where @[simp] theorem map_empty (f : α → β) : map f #[] = #[] := mapM_empty f +@[simp] theorem getElem_map (f : α → β) (a : Array α) (i : Nat) (hi : i < (a.map f).size) : + (a.map f)[i] = f (a[i]'(by simpa using hi)) := by + cases a + simp + +@[simp] theorem getElem?_map (f : α → β) (as : Array α) (i : Nat) : + (as.map f)[i]? = as[i]?.map f := by + simp [getElem?_def] + +@[simp] theorem map_id_fun : map (id : α → α) = id := by + funext l + induction l <;> simp_all + +/-- `map_id_fun'` differs from `map_id_fun` by representing the identity function as a lambda, rather than `id`. -/ +@[simp] theorem map_id_fun' : map (fun (a : α) => a) = id := map_id_fun + +-- This is not a `@[simp]` lemma because `map_id_fun` will apply. +theorem map_id (l : Array α) : map (id : α → α) l = l := by + cases l <;> simp_all + +/-- `map_id'` differs from `map_id` by representing the identity function as a lambda, rather than `id`. -/ +-- This is not a `@[simp]` lemma because `map_id_fun'` will apply. +theorem map_id' (l : Array α) : map (fun (a : α) => a) l = l := map_id l + +/-- Variant of `map_id`, with a side condition that the function is pointwise the identity. -/ +theorem map_id'' {f : α → α} (h : ∀ x, f x = x) (l : Array α) : map f l = l := by + simp [show f = id from funext h] + +theorem map_singleton (f : α → β) (a : α) : map f #[a] = #[f a] := rfl + +@[simp] theorem mem_map {f : α → β} {l : Array α} : b ∈ l.map f ↔ ∃ a, a ∈ l ∧ f a = b := by + simp only [mem_def, toList_map, List.mem_map] + +theorem exists_of_mem_map (h : b ∈ map f l) : ∃ a, a ∈ l ∧ f a = b := mem_map.1 h + +theorem mem_map_of_mem (f : α → β) (h : a ∈ l) : f a ∈ map f l := mem_map.2 ⟨_, h, rfl⟩ + +theorem forall_mem_map {f : α → β} {l : Array α} {P : β → Prop} : + (∀ (i) (_ : i ∈ l.map f), P i) ↔ ∀ (j) (_ : j ∈ l), P (f j) := by + simp + +@[simp] theorem map_inj_left {f g : α → β} : map f l = map g l ↔ ∀ a ∈ l, f a = g a := by + cases l <;> simp_all + +theorem map_congr_left (h : ∀ a ∈ l, f a = g a) : map f l = map g l := + map_inj_left.2 h + +theorem map_inj : map f = map g ↔ f = g := by + constructor + · intro h; ext a; replace h := congrFun h #[a]; simpa using h + · intro h; subst h; rfl + +@[simp] theorem map_eq_empty_iff {f : α → β} {l : Array α} : map f l = #[] ↔ l = #[] := by + cases l + simp + +theorem eq_empty_of_map_eq_empty {f : α → β} {l : Array α} (h : map f l = #[]) : l = #[] := + map_eq_empty_iff.mp h + +@[simp] theorem map_map {f : α → β} {g : β → γ} {as : Array α} : + (as.map f).map g = as.map (g ∘ f) := by + cases as; simp + +@[simp] theorem map_push {f : α → β} {as : Array α} {x : α} : + (as.push x).map f = (as.map f).push (f x) := by + ext + · simp + · simp only [getElem_map, getElem_push, size_map] + split <;> rfl + +theorem mapM_eq_mapM_toList [Monad m] [LawfulMonad m] (f : α → m β) (arr : Array α) : + arr.mapM f = List.toArray <$> (arr.toList.mapM f) := by + rw [mapM_eq_foldlM, ← foldlM_toList, ← List.foldrM_reverse] + conv => rhs; rw [← List.reverse_reverse arr.toList] + induction arr.toList.reverse with + | nil => simp + | cons a l ih => simp [ih] + +@[simp] theorem toList_mapM [Monad m] [LawfulMonad m] (f : α → m β) (arr : Array α) : + toList <$> arr.mapM f = arr.toList.mapM f := by + simp [mapM_eq_mapM_toList] + +theorem mapM_map_eq_foldl (as : Array α) (f : α → β) (i) : + mapM.map (m := Id) f as i b = as.foldl (start := i) (fun r a => r.push (f a)) b := by + unfold mapM.map + split <;> rename_i h + · simp only [Id.bind_eq] + dsimp [foldl, Id.run, foldlM] + rw [mapM_map_eq_foldl, dif_pos (by omega), foldlM.loop, dif_pos h] + -- Calling `split` here gives a bad goal. + have : size as - i = Nat.succ (size as - i - 1) := by omega + rw [this] + simp [foldl, foldlM, Id.run, Nat.sub_add_eq] + · dsimp [foldl, Id.run, foldlM] + rw [dif_pos (by omega), foldlM.loop, dif_neg h] + rfl +termination_by as.size - i + +theorem map_eq_foldl (as : Array α) (f : α → β) : + as.map f = as.foldl (fun r a => r.push (f a)) #[] := + mapM_map_eq_foldl _ _ _ + +theorem singleton_eq_toArray_singleton (a : α) : #[a] = [a].toArray := rfl + +@[simp] theorem singleton_def (v : α) : singleton v = #[v] := rfl + +-- This is a duplicate of `List.toArray_toList`. +-- It's confusing to guess which namespace this theorem should live in, +-- so we provide both. +@[simp] theorem toArray_toList (a : Array α) : a.toList.toArray = a := rfl + +@[simp] theorem mkEmpty_eq (α n) : @mkEmpty α n = #[] := rfl + +@[deprecated size_toArray (since := "2024-12-11")] +theorem size_mk (as : List α) : (Array.mk as).size = as.length := by simp [size] + +/-- A more efficient version of `arr.toList.reverse`. -/ +@[inline] def toListRev (arr : Array α) : List α := arr.foldl (fun l t => t :: l) [] + +@[simp] theorem toListRev_eq (arr : Array α) : arr.toListRev = arr.toList.reverse := by + rw [toListRev, ← foldl_toList, ← List.foldr_reverse, List.foldr_cons_nil] + @[simp] theorem appendList_nil (arr : Array α) : arr ++ ([] : List α) = arr := Array.ext' (by simp) @[simp] theorem appendList_cons (arr : Array α) (a : α) (l : List α) : @@ -1094,7 +1178,6 @@ theorem foldl_toList_eq_map (l : List α) (acc : Array β) (G : α → β) : (l.foldl (fun acc a => acc.push (G a)) acc).toList = acc.toList ++ l.map G := by induction l generalizing acc <;> simp [*] - /-! # uset -/ attribute [simp] uset @@ -1270,18 +1353,16 @@ theorem get_set (a : Array α) (i : Nat) (hi : i < a.size) (j : Nat) (hj : j < a (h : i ≠ j) : (a.set i v)[j]'(by simp [*]) = a[j] := by simp only [set, ← getElem_toList, List.getElem_set_ne h] -private theorem fin_cast_val (e : n = n') (i : Fin n) : e ▸ i = ⟨i.1, e ▸ i.2⟩ := by cases e; rfl - theorem swap_def (a : Array α) (i j : Nat) (hi hj) : a.swap i j hi hj = (a.set i a[j]).set j a[i] (by simpa using hj) := by - simp [swap, fin_cast_val] + simp [swap] @[simp] theorem toList_swap (a : Array α) (i j : Nat) (hi hj) : (a.swap i j hi hj).toList = (a.toList.set i a[j]).set j a[i] := by simp [swap_def] theorem getElem?_swap (a : Array α) (i j : Nat) (hi hj) (k : Nat) : (a.swap i j hi hj)[k]? = if j = k then some a[i] else if i = k then some a[j] else a[k]? := by - simp [swap_def, get?_set, ← getElem_fin_eq_getElem_toList] + simp [swap_def, get?_set] @[simp] theorem swapAt_def (a : Array α) (i : Nat) (v : α) (hi) : a.swapAt i v hi = (a[i], a.set i v) := rfl @@ -1396,6 +1477,90 @@ theorem getElem_range {n : Nat} {x : Nat} (h : x < (Array.range n).size) : (Arra true_and, Nat.not_lt] at h rw [List.getElem?_eq_none_iff.2 ‹_›, List.getElem?_eq_none_iff.2 (a.toList.length_reverse ▸ ‹_›)] +end Array + +open Array + +namespace List + +@[simp] theorem reverse_toArray (l : List α) : l.toArray.reverse = l.reverse.toArray := by + apply ext' + simp only [toList_reverse] + +end List + +namespace Array + +/-! ### foldlM and foldrM -/ + +theorem foldlM_append [Monad m] [LawfulMonad m] (f : β → α → m β) (b) (l l' : Array α) : + (l ++ l').foldlM f b = l.foldlM f b >>= l'.foldlM f := by + cases l + cases l' + rw [List.append_toArray] + simp + +/-- Variant of `foldM_append` with `h : stop = (l ++ l').size`. -/ +@[simp] theorem foldlM_append' [Monad m] [LawfulMonad m] (f : β → α → m β) (b) (l l' : Array α) + (h : stop = (l ++ l').size) : + (l ++ l').foldlM f b 0 stop = l.foldlM f b >>= l'.foldlM f := by + subst h + rw [foldlM_append] + +theorem foldrM_push [Monad m] (f : α → β → m β) (init : β) (arr : Array α) (a : α) : + (arr.push a).foldrM f init = f a init >>= arr.foldrM f := by + simp only [foldrM_eq_reverse_foldlM_toList, push_toList, List.reverse_append, List.reverse_cons, + List.reverse_nil, List.nil_append, List.singleton_append, List.foldlM_cons, List.foldlM_reverse] + +/-- +Variant of `foldrM_push` with `h : start = arr.size + 1` +rather than `(arr.push a).size` as the argument. +-/ +@[simp] theorem foldrM_push' [Monad m] (f : α → β → m β) (init : β) (arr : Array α) (a : α) + {start} (h : start = arr.size + 1) : + (arr.push a).foldrM f init start = f a init >>= arr.foldrM f := by + simp [← foldrM_push, h] + +theorem foldl_eq_foldlM (f : β → α → β) (b) (l : Array α) : + l.foldl f b = l.foldlM (m := Id) f b := by + cases l + simp [List.foldl_eq_foldlM] + +theorem foldr_eq_foldrM (f : α → β → β) (b) (l : Array α) : + l.foldr f b = l.foldrM (m := Id) f b := by + cases l + simp [List.foldr_eq_foldrM] + +@[simp] theorem id_run_foldlM (f : β → α → Id β) (b) (l : Array α) : + Id.run (l.foldlM f b) = l.foldl f b := (foldl_eq_foldlM f b l).symm + +@[simp] theorem id_run_foldrM (f : α → β → Id β) (b) (l : Array α) : + Id.run (l.foldrM f b) = l.foldr f b := (foldr_eq_foldrM f b l).symm + +/-! ### foldl and foldr -/ + +theorem foldr_push (f : α → β → β) (init : β) (arr : Array α) (a : α) : + (arr.push a).foldr f init = arr.foldr f (f a init) := foldrM_push .. + +/-- +Variant of `foldr_push` with the `h : start = arr.size + 1` +rather than `(arr.push a).size` as the argument. +-/ +@[simp] theorem foldr_push' (f : α → β → β) (init : β) (arr : Array α) (a : α) {start} + (h : start = arr.size + 1) : (arr.push a).foldr f init start = arr.foldr f (f a init) := + foldrM_push' _ _ _ _ h + +@[simp] theorem foldl_push_eq_append (l l' : Array α) : l.foldl push l' = l' ++ l := by + cases l + cases l' + simp + +@[simp] theorem foldr_flip_push_eq_append (l l' : Array α) : + l.foldr (fun x y => push y x) l' = l' ++ l.reverse := by + cases l + cases l' + simp + /-! ### take -/ @[simp] theorem size_take_loop (a : Array α) (n : Nat) : (take.loop n a).size = a.size - n := by @@ -1508,22 +1673,6 @@ theorem foldr_congr {as bs : Array α} (h₀ : as = bs) {f g : α → β → β} as.foldr f a start stop = bs.foldr g b start' stop' := by congr -theorem foldl_eq_foldlM (f : β → α → β) (b) (l : Array α) : - l.foldl f b = l.foldlM (m := Id) f b := by - cases l - simp [List.foldl_eq_foldlM] - -theorem foldr_eq_foldrM (f : α → β → β) (b) (l : Array α) : - l.foldr f b = l.foldrM (m := Id) f b := by - cases l - simp [List.foldr_eq_foldrM] - -@[simp] theorem id_run_foldlM (f : β → α → Id β) (b) (l : Array α) : - Id.run (l.foldlM f b) = l.foldl f b := (foldl_eq_foldlM f b l).symm - -@[simp] theorem id_run_foldrM (f : α → β → Id β) (b) (l : Array α) : - Id.run (l.foldrM f b) = l.foldr f b := (foldr_eq_foldrM f b l).symm - theorem foldl_hom (f : α₁ → α₂) (g₁ : α₁ → β → α₁) (g₂ : α₂ → β → α₂) (l : Array β) (init : α₁) (H : ∀ x y, g₂ (f x) y = f (g₁ x y)) : l.foldl g₂ (f init) = f (l.foldl g₁ init) := by cases l @@ -1538,45 +1687,13 @@ theorem foldr_hom (f : β₁ → β₂) (g₁ : α → β₁ → β₁) (g₂ : /-! ### map -/ -@[simp] theorem mem_map {f : α → β} {l : Array α} : b ∈ l.map f ↔ ∃ a, a ∈ l ∧ f a = b := by - simp only [mem_def, toList_map, List.mem_map] - -theorem exists_of_mem_map (h : b ∈ map f l) : ∃ a, a ∈ l ∧ f a = b := mem_map.1 h - -theorem mem_map_of_mem (f : α → β) (h : a ∈ l) : f a ∈ map f l := mem_map.2 ⟨_, h, rfl⟩ - -theorem mapM_eq_mapM_toList [Monad m] [LawfulMonad m] (f : α → m β) (arr : Array α) : - arr.mapM f = List.toArray <$> (arr.toList.mapM f) := by - rw [mapM_eq_foldlM, ← foldlM_toList, ← List.foldrM_reverse] - conv => rhs; rw [← List.reverse_reverse arr.toList] - induction arr.toList.reverse with - | nil => simp - | cons a l ih => simp [ih] - -@[simp] theorem toList_mapM [Monad m] [LawfulMonad m] (f : α → m β) (arr : Array α) : - toList <$> arr.mapM f = arr.toList.mapM f := by - simp [mapM_eq_mapM_toList] - -theorem mapM_map_eq_foldl (as : Array α) (f : α → β) (i) : - mapM.map (m := Id) f as i b = as.foldl (start := i) (fun r a => r.push (f a)) b := by - unfold mapM.map - split <;> rename_i h - · simp only [Id.bind_eq] - dsimp [foldl, Id.run, foldlM] - rw [mapM_map_eq_foldl, dif_pos (by omega), foldlM.loop, dif_pos h] - -- Calling `split` here gives a bad goal. - have : size as - i = Nat.succ (size as - i - 1) := by omega - rw [this] - simp [foldl, foldlM, Id.run, Nat.sub_add_eq] - · dsimp [foldl, Id.run, foldlM] - rw [dif_pos (by omega), foldlM.loop, dif_neg h] - rfl -termination_by as.size - i - -theorem map_eq_foldl (as : Array α) (f : α → β) : - as.map f = as.foldl (fun r a => r.push (f a)) #[] := - mapM_map_eq_foldl _ _ _ +@[simp] theorem map_pop {f : α → β} {as : Array α} : + as.pop.map f = (as.map f).pop := by + ext + · simp + · simp only [getElem_map, getElem_pop, size_map] +@[deprecated "Use `toList_map` or `List.map_toArray` to characterize `Array.map`." (since := "2025-01-06")] theorem map_induction (as : Array α) (f : α → β) (motive : Nat → Prop) (h0 : motive 0) (p : Fin as.size → β → Prop) (hs : ∀ i, motive i.1 → p i (f as[i]) ∧ motive (i+1)) : motive as.size ∧ @@ -1604,36 +1721,13 @@ theorem map_induction (as : Array α) (f : α → β) (motive : Nat → Prop) (h simp only [show j = i by omega] exact (hs _ m).1 +set_option linter.deprecated false in +@[deprecated "Use `toList_map` or `List.map_toArray` to characterize `Array.map`." (since := "2025-01-06")] theorem map_spec (as : Array α) (f : α → β) (p : Fin as.size → β → Prop) (hs : ∀ i, p i (f as[i])) : ∃ eq : (as.map f).size = as.size, ∀ i h, p ⟨i, h⟩ ((as.map f)[i]) := by simpa using map_induction as f (fun _ => True) trivial p (by simp_all) -@[simp] theorem getElem_map (f : α → β) (as : Array α) (i : Nat) (h) : - (as.map f)[i] = f (as[i]'(size_map .. ▸ h)) := by - have := map_spec as f (fun i b => b = f (as[i])) - simp only [implies_true, true_implies] at this - obtain ⟨eq, w⟩ := this - apply w - simp_all - -@[simp] theorem getElem?_map (f : α → β) (as : Array α) (i : Nat) : - (as.map f)[i]? = as[i]?.map f := by - simp [getElem?_def] - -@[simp] theorem map_push {f : α → β} {as : Array α} {x : α} : - (as.push x).map f = (as.map f).push (f x) := by - ext - · simp - · simp only [getElem_map, getElem_push, size_map] - split <;> rfl - -@[simp] theorem map_pop {f : α → β} {as : Array α} : - as.pop.map f = (as.map f).pop := by - ext - · simp - · simp only [getElem_map, getElem_pop, size_map] - /-! ### modify -/ @[simp] theorem size_modify (a : Array α) (i : Nat) (f : α → α) : (a.modify i f).size = a.size := by @@ -2120,10 +2214,6 @@ theorem toListRev_toArray (l : List α) : l.toArray.toListRev = l.reverse := by rw [size_toArray, mapM'_cons, foldlM_toArray] simp [ih] -@[simp] theorem map_toArray (f : α → β) (l : List α) : l.toArray.map f = (l.map f).toArray := by - apply ext' - simp - theorem uset_toArray (l : List α) (i : USize) (a : α) (h : i.toNat < l.toArray.size) : l.toArray.uset i a h = (l.set i.toNat a).toArray := by simp @@ -2132,10 +2222,6 @@ theorem uset_toArray (l : List α) (i : USize) (a : α) (h : i.toNat < l.toArray apply ext' simp -@[simp] theorem reverse_toArray (l : List α) : l.toArray.reverse = l.reverse.toArray := by - apply ext' - simp - @[simp] theorem modify_toArray (f : α → α) (l : List α) : l.toArray.modify i f = (l.modify f i).toArray := by apply ext' @@ -2229,29 +2315,6 @@ namespace Array /-! ### map -/ -@[simp] theorem map_map {f : α → β} {g : β → γ} {as : Array α} : - (as.map f).map g = as.map (g ∘ f) := by - cases as; simp - -@[simp] theorem map_id_fun : map (id : α → α) = id := by - funext l - induction l <;> simp_all - -/-- `map_id_fun'` differs from `map_id_fun` by representing the identity function as a lambda, rather than `id`. -/ -@[simp] theorem map_id_fun' : map (fun (a : α) => a) = id := map_id_fun - --- This is not a `@[simp]` lemma because `map_id_fun` will apply. -theorem map_id (as : Array α) : map (id : α → α) as = as := by - cases as <;> simp_all - -/-- `map_id'` differs from `map_id` by representing the identity function as a lambda, rather than `id`. -/ --- This is not a `@[simp]` lemma because `map_id_fun'` will apply. -theorem map_id' (as : Array α) : map (fun (a : α) => a) as = as := map_id as - -/-- Variant of `map_id`, with a side condition that the function is pointwise the identity. -/ -theorem map_id'' {f : α → α} (h : ∀ x, f x = x) (as : Array α) : map f as = as := by - simp [show f = id from funext h] - theorem array_array_induction (P : Array (Array α) → Prop) (h : ∀ (xss : List (List α)), P (xss.map List.toArray).toArray) (ass : Array (Array α)) : P ass := by specialize h (ass.toList.map toList) diff --git a/src/Init/Data/List/Lemmas.lean b/src/Init/Data/List/Lemmas.lean index a06c12d4d9a9..e29c161a31f0 100644 --- a/src/Init/Data/List/Lemmas.lean +++ b/src/Init/Data/List/Lemmas.lean @@ -757,207 +757,6 @@ theorem length_eq_of_beq [BEq α] {l₁ l₂ : List α} (h : l₁ == l₂) : l | nil => simp | cons b l₂ => simp [isEqv, ih] -/-! ### foldlM and foldrM -/ - -@[simp] theorem foldlM_reverse [Monad m] (l : List α) (f : β → α → m β) (b) : - l.reverse.foldlM f b = l.foldrM (fun x y => f y x) b := rfl - -@[simp] theorem foldlM_append [Monad m] [LawfulMonad m] (f : β → α → m β) (b) (l l' : List α) : - (l ++ l').foldlM f b = l.foldlM f b >>= l'.foldlM f := by - induction l generalizing b <;> simp [*] - -@[simp] theorem foldrM_cons [Monad m] [LawfulMonad m] (a : α) (l) (f : α → β → m β) (b) : - (a :: l).foldrM f b = l.foldrM f b >>= f a := by - simp only [foldrM] - induction l <;> simp_all - -theorem foldl_eq_foldlM (f : β → α → β) (b) (l : List α) : - l.foldl f b = l.foldlM (m := Id) f b := by - induction l generalizing b <;> simp [*, foldl] - -theorem foldr_eq_foldrM (f : α → β → β) (b) (l : List α) : - l.foldr f b = l.foldrM (m := Id) f b := by - induction l <;> simp [*, foldr] - -@[simp] theorem id_run_foldlM (f : β → α → Id β) (b) (l : List α) : - Id.run (l.foldlM f b) = l.foldl f b := (foldl_eq_foldlM f b l).symm - -@[simp] theorem id_run_foldrM (f : α → β → Id β) (b) (l : List α) : - Id.run (l.foldrM f b) = l.foldr f b := (foldr_eq_foldrM f b l).symm - -/-! ### foldl and foldr -/ - -@[simp] theorem foldr_cons_eq_append (l : List α) : l.foldr cons l' = l ++ l' := by - induction l <;> simp [*] - -@[deprecated foldr_cons_eq_append (since := "2024-08-22")] abbrev foldr_self_append := @foldr_cons_eq_append - -@[simp] theorem foldl_flip_cons_eq_append (l : List α) : l.foldl (fun x y => y :: x) l' = l.reverse ++ l' := by - induction l generalizing l' <;> simp [*] - -theorem foldr_cons_nil (l : List α) : l.foldr cons [] = l := by simp - -@[deprecated foldr_cons_nil (since := "2024-09-04")] abbrev foldr_self := @foldr_cons_nil - -theorem foldl_map (f : β₁ → β₂) (g : α → β₂ → α) (l : List β₁) (init : α) : - (l.map f).foldl g init = l.foldl (fun x y => g x (f y)) init := by - induction l generalizing init <;> simp [*] - -theorem foldr_map (f : α₁ → α₂) (g : α₂ → β → β) (l : List α₁) (init : β) : - (l.map f).foldr g init = l.foldr (fun x y => g (f x) y) init := by - induction l generalizing init <;> simp [*] - -theorem foldl_filterMap (f : α → Option β) (g : γ → β → γ) (l : List α) (init : γ) : - (l.filterMap f).foldl g init = l.foldl (fun x y => match f y with | some b => g x b | none => x) init := by - induction l generalizing init with - | nil => rfl - | cons a l ih => - simp only [filterMap_cons, foldl_cons] - cases f a <;> simp [ih] - -theorem foldr_filterMap (f : α → Option β) (g : β → γ → γ) (l : List α) (init : γ) : - (l.filterMap f).foldr g init = l.foldr (fun x y => match f x with | some b => g b y | none => y) init := by - induction l generalizing init with - | nil => rfl - | cons a l ih => - simp only [filterMap_cons, foldr_cons] - cases f a <;> simp [ih] - -theorem foldl_map' (g : α → β) (f : α → α → α) (f' : β → β → β) (a : α) (l : List α) - (h : ∀ x y, f' (g x) (g y) = g (f x y)) : - (l.map g).foldl f' (g a) = g (l.foldl f a) := by - induction l generalizing a - · simp - · simp [*, h] - -theorem foldr_map' (g : α → β) (f : α → α → α) (f' : β → β → β) (a : α) (l : List α) - (h : ∀ x y, f' (g x) (g y) = g (f x y)) : - (l.map g).foldr f' (g a) = g (l.foldr f a) := by - induction l generalizing a - · simp - · simp [*, h] - -theorem foldl_assoc {op : α → α → α} [ha : Std.Associative op] : - ∀ {l : List α} {a₁ a₂}, l.foldl op (op a₁ a₂) = op a₁ (l.foldl op a₂) - | [], a₁, a₂ => rfl - | a :: l, a₁, a₂ => by - simp only [foldl_cons, ha.assoc] - rw [foldl_assoc] - -theorem foldr_assoc {op : α → α → α} [ha : Std.Associative op] : - ∀ {l : List α} {a₁ a₂}, l.foldr op (op a₁ a₂) = op (l.foldr op a₁) a₂ - | [], a₁, a₂ => rfl - | a :: l, a₁, a₂ => by - simp only [foldr_cons, ha.assoc] - rw [foldr_assoc] - -theorem foldl_hom (f : α₁ → α₂) (g₁ : α₁ → β → α₁) (g₂ : α₂ → β → α₂) (l : List β) (init : α₁) - (H : ∀ x y, g₂ (f x) y = f (g₁ x y)) : l.foldl g₂ (f init) = f (l.foldl g₁ init) := by - induction l generalizing init <;> simp [*, H] - -theorem foldr_hom (f : β₁ → β₂) (g₁ : α → β₁ → β₁) (g₂ : α → β₂ → β₂) (l : List α) (init : β₁) - (H : ∀ x y, g₂ x (f y) = f (g₁ x y)) : l.foldr g₂ (f init) = f (l.foldr g₁ init) := by - induction l <;> simp [*, H] - -/-- -Prove a proposition about the result of `List.foldl`, -by proving it for the initial data, -and the implication that the operation applied to any element of the list preserves the property. - -The motive can take values in `Sort _`, so this may be used to construct data, -as well as to prove propositions. --/ -def foldlRecOn {motive : β → Sort _} : ∀ (l : List α) (op : β → α → β) (b : β) (_ : motive b) - (_ : ∀ (b : β) (_ : motive b) (a : α) (_ : a ∈ l), motive (op b a)), motive (List.foldl op b l) - | [], _, _, hb, _ => hb - | hd :: tl, op, b, hb, hl => - foldlRecOn tl op (op b hd) (hl b hb hd (mem_cons_self hd tl)) - fun y hy x hx => hl y hy x (mem_cons_of_mem hd hx) - -@[simp] theorem foldlRecOn_nil {motive : β → Sort _} (hb : motive b) - (hl : ∀ (b : β) (_ : motive b) (a : α) (_ : a ∈ []), motive (op b a)) : - foldlRecOn [] op b hb hl = hb := rfl - -@[simp] theorem foldlRecOn_cons {motive : β → Sort _} (hb : motive b) - (hl : ∀ (b : β) (_ : motive b) (a : α) (_ : a ∈ x :: l), motive (op b a)) : - foldlRecOn (x :: l) op b hb hl = - foldlRecOn l op (op b x) (hl b hb x (mem_cons_self x l)) - (fun b c a m => hl b c a (mem_cons_of_mem x m)) := - rfl - -/-- -Prove a proposition about the result of `List.foldr`, -by proving it for the initial data, -and the implication that the operation applied to any element of the list preserves the property. - -The motive can take values in `Sort _`, so this may be used to construct data, -as well as to prove propositions. --/ -def foldrRecOn {motive : β → Sort _} : ∀ (l : List α) (op : α → β → β) (b : β) (_ : motive b) - (_ : ∀ (b : β) (_ : motive b) (a : α) (_ : a ∈ l), motive (op a b)), motive (List.foldr op b l) - | nil, _, _, hb, _ => hb - | x :: l, op, b, hb, hl => - hl (foldr op b l) - (foldrRecOn l op b hb fun b c a m => hl b c a (mem_cons_of_mem x m)) x (mem_cons_self x l) - -@[simp] theorem foldrRecOn_nil {motive : β → Sort _} (hb : motive b) - (hl : ∀ (b : β) (_ : motive b) (a : α) (_ : a ∈ []), motive (op a b)) : - foldrRecOn [] op b hb hl = hb := rfl - -@[simp] theorem foldrRecOn_cons {motive : β → Sort _} (hb : motive b) - (hl : ∀ (b : β) (_ : motive b) (a : α) (_ : a ∈ x :: l), motive (op a b)) : - foldrRecOn (x :: l) op b hb hl = - hl _ (foldrRecOn l op b hb fun b c a m => hl b c a (mem_cons_of_mem x m)) - x (mem_cons_self x l) := - rfl - -/-- -We can prove that two folds over the same list are related (by some arbitrary relation) -if we know that the initial elements are related and the folding function, for each element of the list, -preserves the relation. --/ -theorem foldl_rel {l : List α} {f g : β → α → β} {a b : β} (r : β → β → Prop) - (h : r a b) (h' : ∀ (a : α), a ∈ l → ∀ (c c' : β), r c c' → r (f c a) (g c' a)) : - r (l.foldl (fun acc a => f acc a) a) (l.foldl (fun acc a => g acc a) b) := by - induction l generalizing a b with - | nil => simp_all - | cons a l ih => - simp only [foldl_cons] - apply ih - · simp_all - · exact fun a m c c' h => h' _ (by simp_all) _ _ h - -/-- -We can prove that two folds over the same list are related (by some arbitrary relation) -if we know that the initial elements are related and the folding function, for each element of the list, -preserves the relation. --/ -theorem foldr_rel {l : List α} {f g : α → β → β} {a b : β} (r : β → β → Prop) - (h : r a b) (h' : ∀ (a : α), a ∈ l → ∀ (c c' : β), r c c' → r (f a c) (g a c')) : - r (l.foldr (fun a acc => f a acc) a) (l.foldr (fun a acc => g a acc) b) := by - induction l generalizing a b with - | nil => simp_all - | cons a l ih => - simp only [foldr_cons] - apply h' - · simp - · exact ih h fun a m c c' h => h' _ (by simp_all) _ _ h - -@[simp] theorem foldl_add_const (l : List α) (a b : Nat) : - l.foldl (fun x _ => x + a) b = b + a * l.length := by - induction l generalizing b with - | nil => simp - | cons y l ih => - simp only [foldl_cons, ih, length_cons, Nat.mul_add, Nat.mul_one, Nat.add_assoc, - Nat.add_comm a] - -@[simp] theorem foldr_add_const (l : List α) (a b : Nat) : - l.foldr (fun _ x => x + a) b = b + a * l.length := by - induction l generalizing b with - | nil => simp - | cons y l ih => - simp only [foldr_cons, ih, length_cons, Nat.mul_add, Nat.mul_one, Nat.add_assoc] - /-! ### getLast -/ theorem getLast_eq_getElem : ∀ (l : List α) (h : l ≠ []), @@ -1216,27 +1015,6 @@ theorem getLast?_tail (l : List α) : (tail l).getLast? = if l.length = 1 then n /-! ### map -/ -@[simp] theorem map_id_fun : map (id : α → α) = id := by - funext l - induction l <;> simp_all - -/-- `map_id_fun'` differs from `map_id_fun` by representing the identity function as a lambda, rather than `id`. -/ -@[simp] theorem map_id_fun' : map (fun (a : α) => a) = id := map_id_fun - --- This is not a `@[simp]` lemma because `map_id_fun` will apply. -theorem map_id (l : List α) : map (id : α → α) l = l := by - induction l <;> simp_all - -/-- `map_id'` differs from `map_id` by representing the identity function as a lambda, rather than `id`. -/ --- This is not a `@[simp]` lemma because `map_id_fun'` will apply. -theorem map_id' (l : List α) : map (fun (a : α) => a) l = l := map_id l - -/-- Variant of `map_id`, with a side condition that the function is pointwise the identity. -/ -theorem map_id'' {f : α → α} (h : ∀ x, f x = x) (l : List α) : map f l = l := by - simp [show f = id from funext h] - -theorem map_singleton (f : α → β) (a : α) : map f [a] = [f a] := rfl - @[simp] theorem length_map (as : List α) (f : α → β) : (as.map f).length = as.length := by induction as with | nil => simp [List.map] @@ -1262,6 +1040,27 @@ theorem get_map (f : α → β) {l i} : get (map f l) i = f (get l ⟨i, length_map l f ▸ i.2⟩) := by simp +@[simp] theorem map_id_fun : map (id : α → α) = id := by + funext l + induction l <;> simp_all + +/-- `map_id_fun'` differs from `map_id_fun` by representing the identity function as a lambda, rather than `id`. -/ +@[simp] theorem map_id_fun' : map (fun (a : α) => a) = id := map_id_fun + +-- This is not a `@[simp]` lemma because `map_id_fun` will apply. +theorem map_id (l : List α) : map (id : α → α) l = l := by + induction l <;> simp_all + +/-- `map_id'` differs from `map_id` by representing the identity function as a lambda, rather than `id`. -/ +-- This is not a `@[simp]` lemma because `map_id_fun'` will apply. +theorem map_id' (l : List α) : map (fun (a : α) => a) l = l := map_id l + +/-- Variant of `map_id`, with a side condition that the function is pointwise the identity. -/ +theorem map_id'' {f : α → α} (h : ∀ x, f x = x) (l : List α) : map f l = l := by + simp [show f = id from funext h] + +theorem map_singleton (f : α → β) (a : α) : map f [a] = [f a] := rfl + @[simp] theorem mem_map {f : α → β} : ∀ {l : List α}, b ∈ l.map f ↔ ∃ a, a ∈ l ∧ f a = b | [] => by simp | _ :: l => by simp [mem_map (l := l), eq_comm (a := b)] @@ -1961,16 +1760,6 @@ theorem set_append {s t : List α} : (s ++ t).set i x = s ++ t.set (i - s.length) x := by rw [set_append, if_neg (by simp_all)] -@[simp] theorem foldrM_append [Monad m] [LawfulMonad m] (f : α → β → m β) (b) (l l' : List α) : - (l ++ l').foldrM f b = l'.foldrM f b >>= l.foldrM f := by - induction l <;> simp [*] - -@[simp] theorem foldl_append {β : Type _} (f : β → α → β) (b) (l l' : List α) : - (l ++ l').foldl f b = l'.foldl f (l.foldl f b) := by simp [foldl_eq_foldlM] - -@[simp] theorem foldr_append (f : α → β → β) (b) (l l' : List α) : - (l ++ l').foldr f b = l.foldr f (l'.foldr f b) := by simp [foldr_eq_foldrM] - theorem filterMap_eq_append_iff {f : α → Option β} : filterMap f l = L₁ ++ L₂ ↔ ∃ l₁ l₂, l = l₁ ++ l₂ ∧ filterMap f l₁ = L₁ ∧ filterMap f l₂ = L₂ := by constructor @@ -2119,14 +1908,6 @@ theorem head?_flatten {L : List (List α)} : (flatten L).head? = L.findSome? fun -- `getLast?_flatten` is proved later, after the `reverse` section. -- `head_flatten` and `getLast_flatten` are proved in `Init.Data.List.Find`. -theorem foldl_flatten (f : β → α → β) (b : β) (L : List (List α)) : - (flatten L).foldl f b = L.foldl (fun b l => l.foldl f b) b := by - induction L generalizing b <;> simp_all - -theorem foldr_flatten (f : α → β → β) (b : β) (L : List (List α)) : - (flatten L).foldr f b = L.foldr (fun l b => l.foldr f b) b := by - induction L <;> simp_all - @[simp] theorem map_flatten (f : α → β) (L : List (List α)) : map f (flatten L) = flatten (map (map f) L) := by induction L <;> simp_all @@ -2699,10 +2480,114 @@ theorem flatMap_reverse {β} (l : List α) (f : α → List β) : (l.reverse.fla @[simp] theorem reverseAux_eq (as bs : List α) : reverseAux as bs = reverse as ++ bs := reverseAux_eq_append .. +@[simp] theorem reverse_replicate (n) (a : α) : reverse (replicate n a) = replicate n a := + eq_replicate_iff.2 + ⟨by rw [length_reverse, length_replicate], + fun _ h => eq_of_mem_replicate (mem_reverse.1 h)⟩ + + +/-! ### foldlM and foldrM -/ + +@[simp] theorem foldlM_append [Monad m] [LawfulMonad m] (f : β → α → m β) (b) (l l' : List α) : + (l ++ l').foldlM f b = l.foldlM f b >>= l'.foldlM f := by + induction l generalizing b <;> simp [*] + +@[simp] theorem foldrM_cons [Monad m] [LawfulMonad m] (a : α) (l) (f : α → β → m β) (b) : + (a :: l).foldrM f b = l.foldrM f b >>= f a := by + simp only [foldrM] + induction l <;> simp_all + +theorem foldl_eq_foldlM (f : β → α → β) (b) (l : List α) : + l.foldl f b = l.foldlM (m := Id) f b := by + induction l generalizing b <;> simp [*, foldl] + +theorem foldr_eq_foldrM (f : α → β → β) (b) (l : List α) : + l.foldr f b = l.foldrM (m := Id) f b := by + induction l <;> simp [*, foldr] + +@[simp] theorem id_run_foldlM (f : β → α → Id β) (b) (l : List α) : + Id.run (l.foldlM f b) = l.foldl f b := (foldl_eq_foldlM f b l).symm + +@[simp] theorem id_run_foldrM (f : α → β → Id β) (b) (l : List α) : + Id.run (l.foldrM f b) = l.foldr f b := (foldr_eq_foldrM f b l).symm + +@[simp] theorem foldlM_reverse [Monad m] (l : List α) (f : β → α → m β) (b) : + l.reverse.foldlM f b = l.foldrM (fun x y => f y x) b := rfl + @[simp] theorem foldrM_reverse [Monad m] (l : List α) (f : α → β → m β) (b) : l.reverse.foldrM f b = l.foldlM (fun x y => f y x) b := (foldlM_reverse ..).symm.trans <| by simp +/-! ### foldl and foldr -/ + +@[simp] theorem foldr_cons_eq_append (l : List α) : l.foldr cons l' = l ++ l' := by + induction l <;> simp [*] + +@[deprecated foldr_cons_eq_append (since := "2024-08-22")] abbrev foldr_self_append := @foldr_cons_eq_append + +@[simp] theorem foldl_flip_cons_eq_append (l : List α) : l.foldl (fun x y => y :: x) l' = l.reverse ++ l' := by + induction l generalizing l' <;> simp [*] + +theorem foldr_cons_nil (l : List α) : l.foldr cons [] = l := by simp + +@[deprecated foldr_cons_nil (since := "2024-09-04")] abbrev foldr_self := @foldr_cons_nil + +theorem foldl_map (f : β₁ → β₂) (g : α → β₂ → α) (l : List β₁) (init : α) : + (l.map f).foldl g init = l.foldl (fun x y => g x (f y)) init := by + induction l generalizing init <;> simp [*] + +theorem foldr_map (f : α₁ → α₂) (g : α₂ → β → β) (l : List α₁) (init : β) : + (l.map f).foldr g init = l.foldr (fun x y => g (f x) y) init := by + induction l generalizing init <;> simp [*] + +theorem foldl_filterMap (f : α → Option β) (g : γ → β → γ) (l : List α) (init : γ) : + (l.filterMap f).foldl g init = l.foldl (fun x y => match f y with | some b => g x b | none => x) init := by + induction l generalizing init with + | nil => rfl + | cons a l ih => + simp only [filterMap_cons, foldl_cons] + cases f a <;> simp [ih] + +theorem foldr_filterMap (f : α → Option β) (g : β → γ → γ) (l : List α) (init : γ) : + (l.filterMap f).foldr g init = l.foldr (fun x y => match f x with | some b => g b y | none => y) init := by + induction l generalizing init with + | nil => rfl + | cons a l ih => + simp only [filterMap_cons, foldr_cons] + cases f a <;> simp [ih] + +theorem foldl_map' (g : α → β) (f : α → α → α) (f' : β → β → β) (a : α) (l : List α) + (h : ∀ x y, f' (g x) (g y) = g (f x y)) : + (l.map g).foldl f' (g a) = g (l.foldl f a) := by + induction l generalizing a + · simp + · simp [*, h] + +theorem foldr_map' (g : α → β) (f : α → α → α) (f' : β → β → β) (a : α) (l : List α) + (h : ∀ x y, f' (g x) (g y) = g (f x y)) : + (l.map g).foldr f' (g a) = g (l.foldr f a) := by + induction l generalizing a + · simp + · simp [*, h] + +@[simp] theorem foldrM_append [Monad m] [LawfulMonad m] (f : α → β → m β) (b) (l l' : List α) : + (l ++ l').foldrM f b = l'.foldrM f b >>= l.foldrM f := by + induction l <;> simp [*] + +@[simp] theorem foldl_append {β : Type _} (f : β → α → β) (b) (l l' : List α) : + (l ++ l').foldl f b = l'.foldl f (l.foldl f b) := by simp [foldl_eq_foldlM] + +@[simp] theorem foldr_append (f : α → β → β) (b) (l l' : List α) : + (l ++ l').foldr f b = l.foldr f (l'.foldr f b) := by simp [foldr_eq_foldrM] + +theorem foldl_flatten (f : β → α → β) (b : β) (L : List (List α)) : + (flatten L).foldl f b = L.foldl (fun b l => l.foldl f b) b := by + induction L generalizing b <;> simp_all + +theorem foldr_flatten (f : α → β → β) (b : β) (L : List (List α)) : + (flatten L).foldr f b = L.foldr (fun l b => l.foldr f b) b := by + induction L <;> simp_all + @[simp] theorem foldl_reverse (l : List α) (f : β → α → β) (b) : l.reverse.foldl f b = l.foldr (fun x y => f y x) b := by simp [foldl_eq_foldlM, foldr_eq_foldrM] @@ -2716,10 +2601,127 @@ theorem foldl_eq_foldr_reverse (l : List α) (f : β → α → β) (b) : theorem foldr_eq_foldl_reverse (l : List α) (f : α → β → β) (b) : l.foldr f b = l.reverse.foldl (fun x y => f y x) b := by simp -@[simp] theorem reverse_replicate (n) (a : α) : reverse (replicate n a) = replicate n a := - eq_replicate_iff.2 - ⟨by rw [length_reverse, length_replicate], - fun _ h => eq_of_mem_replicate (mem_reverse.1 h)⟩ +theorem foldl_assoc {op : α → α → α} [ha : Std.Associative op] : + ∀ {l : List α} {a₁ a₂}, l.foldl op (op a₁ a₂) = op a₁ (l.foldl op a₂) + | [], a₁, a₂ => rfl + | a :: l, a₁, a₂ => by + simp only [foldl_cons, ha.assoc] + rw [foldl_assoc] + +theorem foldr_assoc {op : α → α → α} [ha : Std.Associative op] : + ∀ {l : List α} {a₁ a₂}, l.foldr op (op a₁ a₂) = op (l.foldr op a₁) a₂ + | [], a₁, a₂ => rfl + | a :: l, a₁, a₂ => by + simp only [foldr_cons, ha.assoc] + rw [foldr_assoc] + +theorem foldl_hom (f : α₁ → α₂) (g₁ : α₁ → β → α₁) (g₂ : α₂ → β → α₂) (l : List β) (init : α₁) + (H : ∀ x y, g₂ (f x) y = f (g₁ x y)) : l.foldl g₂ (f init) = f (l.foldl g₁ init) := by + induction l generalizing init <;> simp [*, H] + +theorem foldr_hom (f : β₁ → β₂) (g₁ : α → β₁ → β₁) (g₂ : α → β₂ → β₂) (l : List α) (init : β₁) + (H : ∀ x y, g₂ x (f y) = f (g₁ x y)) : l.foldr g₂ (f init) = f (l.foldr g₁ init) := by + induction l <;> simp [*, H] + +/-- +Prove a proposition about the result of `List.foldl`, +by proving it for the initial data, +and the implication that the operation applied to any element of the list preserves the property. + +The motive can take values in `Sort _`, so this may be used to construct data, +as well as to prove propositions. +-/ +def foldlRecOn {motive : β → Sort _} : ∀ (l : List α) (op : β → α → β) (b : β) (_ : motive b) + (_ : ∀ (b : β) (_ : motive b) (a : α) (_ : a ∈ l), motive (op b a)), motive (List.foldl op b l) + | [], _, _, hb, _ => hb + | hd :: tl, op, b, hb, hl => + foldlRecOn tl op (op b hd) (hl b hb hd (mem_cons_self hd tl)) + fun y hy x hx => hl y hy x (mem_cons_of_mem hd hx) + +@[simp] theorem foldlRecOn_nil {motive : β → Sort _} (hb : motive b) + (hl : ∀ (b : β) (_ : motive b) (a : α) (_ : a ∈ []), motive (op b a)) : + foldlRecOn [] op b hb hl = hb := rfl + +@[simp] theorem foldlRecOn_cons {motive : β → Sort _} (hb : motive b) + (hl : ∀ (b : β) (_ : motive b) (a : α) (_ : a ∈ x :: l), motive (op b a)) : + foldlRecOn (x :: l) op b hb hl = + foldlRecOn l op (op b x) (hl b hb x (mem_cons_self x l)) + (fun b c a m => hl b c a (mem_cons_of_mem x m)) := + rfl + +/-- +Prove a proposition about the result of `List.foldr`, +by proving it for the initial data, +and the implication that the operation applied to any element of the list preserves the property. + +The motive can take values in `Sort _`, so this may be used to construct data, +as well as to prove propositions. +-/ +def foldrRecOn {motive : β → Sort _} : ∀ (l : List α) (op : α → β → β) (b : β) (_ : motive b) + (_ : ∀ (b : β) (_ : motive b) (a : α) (_ : a ∈ l), motive (op a b)), motive (List.foldr op b l) + | nil, _, _, hb, _ => hb + | x :: l, op, b, hb, hl => + hl (foldr op b l) + (foldrRecOn l op b hb fun b c a m => hl b c a (mem_cons_of_mem x m)) x (mem_cons_self x l) + +@[simp] theorem foldrRecOn_nil {motive : β → Sort _} (hb : motive b) + (hl : ∀ (b : β) (_ : motive b) (a : α) (_ : a ∈ []), motive (op a b)) : + foldrRecOn [] op b hb hl = hb := rfl + +@[simp] theorem foldrRecOn_cons {motive : β → Sort _} (hb : motive b) + (hl : ∀ (b : β) (_ : motive b) (a : α) (_ : a ∈ x :: l), motive (op a b)) : + foldrRecOn (x :: l) op b hb hl = + hl _ (foldrRecOn l op b hb fun b c a m => hl b c a (mem_cons_of_mem x m)) + x (mem_cons_self x l) := + rfl + +/-- +We can prove that two folds over the same list are related (by some arbitrary relation) +if we know that the initial elements are related and the folding function, for each element of the list, +preserves the relation. +-/ +theorem foldl_rel {l : List α} {f g : β → α → β} {a b : β} (r : β → β → Prop) + (h : r a b) (h' : ∀ (a : α), a ∈ l → ∀ (c c' : β), r c c' → r (f c a) (g c' a)) : + r (l.foldl (fun acc a => f acc a) a) (l.foldl (fun acc a => g acc a) b) := by + induction l generalizing a b with + | nil => simp_all + | cons a l ih => + simp only [foldl_cons] + apply ih + · simp_all + · exact fun a m c c' h => h' _ (by simp_all) _ _ h + +/-- +We can prove that two folds over the same list are related (by some arbitrary relation) +if we know that the initial elements are related and the folding function, for each element of the list, +preserves the relation. +-/ +theorem foldr_rel {l : List α} {f g : α → β → β} {a b : β} (r : β → β → Prop) + (h : r a b) (h' : ∀ (a : α), a ∈ l → ∀ (c c' : β), r c c' → r (f a c) (g a c')) : + r (l.foldr (fun a acc => f a acc) a) (l.foldr (fun a acc => g a acc) b) := by + induction l generalizing a b with + | nil => simp_all + | cons a l ih => + simp only [foldr_cons] + apply h' + · simp + · exact ih h fun a m c c' h => h' _ (by simp_all) _ _ h + +@[simp] theorem foldl_add_const (l : List α) (a b : Nat) : + l.foldl (fun x _ => x + a) b = b + a * l.length := by + induction l generalizing b with + | nil => simp + | cons y l ih => + simp only [foldl_cons, ih, length_cons, Nat.mul_add, Nat.mul_one, Nat.add_assoc, + Nat.add_comm a] + +@[simp] theorem foldr_add_const (l : List α) (a b : Nat) : + l.foldr (fun _ x => x + a) b = b + a * l.length := by + induction l generalizing b with + | nil => simp + | cons y l ih => + simp only [foldr_cons, ih, length_cons, Nat.mul_add, Nat.mul_one, Nat.add_assoc] + /-! #### Further results about `getLast` and `getLast?` -/ diff --git a/src/Init/Data/Vector/Basic.lean b/src/Init/Data/Vector/Basic.lean index 44553e763419..31e3e430bc69 100644 --- a/src/Init/Data/Vector/Basic.lean +++ b/src/Init/Data/Vector/Basic.lean @@ -136,6 +136,18 @@ This will perform the update destructively provided that the vector has a refere @[inline] def set! (v : Vector α n) (i : Nat) (x : α) : Vector α n := ⟨v.toArray.set! i x, by simp⟩ +@[inline] def foldlM [Monad m] (f : β → α → m β) (b : β) (v : Vector α n) : m β := + v.toArray.foldlM f b + +@[inline] def foldrM [Monad m] (f : α → β → m β) (b : β) (v : Vector α n) : m β := + v.toArray.foldrM f b + +@[inline] def foldl (f : β → α → β) (b : β) (v : Vector α n) : β := + v.toArray.foldl f b + +@[inline] def foldr (f : α → β → β) (b : β) (v : Vector α n) : β := + v.toArray.foldr f b + /-- Append two vectors. -/ @[inline] def append (v : Vector α n) (w : Vector α m) : Vector α (n + m) := ⟨v.toArray ++ w.toArray, by simp⟩ diff --git a/src/Init/Data/Vector/Lemmas.lean b/src/Init/Data/Vector/Lemmas.lean index 0f91c8ba9dae..88ccbc6dcb63 100644 --- a/src/Init/Data/Vector/Lemmas.lean +++ b/src/Init/Data/Vector/Lemmas.lean @@ -66,6 +66,18 @@ theorem toArray_mk (a : Array α) (h : a.size = n) : (Vector.mk a h).toArray = a @[simp] theorem back?_mk (a : Array α) (h : a.size = n) : (Vector.mk a h).back? = a.back? := rfl +@[simp] theorem foldlM_mk [Monad m] (f : β → α → m β) (b : β) (a : Array α) (h : a.size = n) : + (Vector.mk a h).foldlM f b = a.foldlM f b := rfl + +@[simp] theorem foldrM_mk [Monad m] (f : α → β → m β) (b : β) (a : Array α) (h : a.size = n) : + (Vector.mk a h).foldrM f b = a.foldrM f b := rfl + +@[simp] theorem foldl_mk (f : β → α → β) (b : β) (a : Array α) (h : a.size = n) : + (Vector.mk a h).foldl f b = a.foldl f b := rfl + +@[simp] theorem foldr_mk (f : α → β → β) (b : β) (a : Array α) (h : a.size = n) : + (Vector.mk a h).foldr f b = a.foldr f b := rfl + @[simp] theorem drop_mk (a : Array α) (h : a.size = n) (m) : (Vector.mk a h).drop m = Vector.mk (a.extract m a.size) (by simp [h]) := rfl @@ -1025,6 +1037,13 @@ theorem mem_setIfInBounds (v : Vector α n) (i : Nat) (hi : i < n) (a : α) : /-! Content below this point has not yet been aligned with `List` and `Array`. -/ +/-! ### map -/ + +@[simp] theorem getElem_map (f : α → β) (a : Vector α n) (i : Nat) (hi : i < n) : + (a.map f)[i] = f a[i] := by + cases a + simp + @[simp] theorem getElem_ofFn {α n} (f : Fin n → α) (i : Nat) (h : i < n) : (Vector.ofFn f)[i] = f ⟨i, by simpa using h⟩ := by simp [ofFn] @@ -1088,13 +1107,6 @@ theorem getElem_append_right {a : Vector α n} {b : Vector α m} {i : Nat} (h : cases a simp -/-! ### map -/ - -@[simp] theorem getElem_map (f : α → β) (a : Vector α n) (i : Nat) (hi : i < n) : - (a.map f)[i] = f a[i] := by - cases a - simp - /-! ### zipWith -/ @[simp] theorem getElem_zipWith (f : α → β → γ) (a : Vector α n) (b : Vector β n) (i : Nat) @@ -1103,6 +1115,37 @@ theorem getElem_append_right {a : Vector α n} {b : Vector α m} {i : Nat} (h : cases b simp +/-! ### foldlM and foldrM -/ + +@[simp] theorem foldlM_append [Monad m] [LawfulMonad m] (f : β → α → m β) (b) (l : Vector α n) (l' : Vector α n') : + (l ++ l').foldlM f b = l.foldlM f b >>= l'.foldlM f := by + cases l + cases l' + simp + +@[simp] theorem foldrM_push [Monad m] (f : α → β → m β) (init : β) (l : Vector α n) (a : α) : + (l.push a).foldrM f init = f a init >>= l.foldrM f := by + cases l + simp + +theorem foldl_eq_foldlM (f : β → α → β) (b) (l : Vector α n) : + l.foldl f b = l.foldlM (m := Id) f b := by + cases l + simp [Array.foldl_eq_foldlM] + +theorem foldr_eq_foldrM (f : α → β → β) (b) (l : Vector α n) : + l.foldr f b = l.foldrM (m := Id) f b := by + cases l + simp [Array.foldr_eq_foldrM] + +@[simp] theorem id_run_foldlM (f : β → α → Id β) (b) (l : Vector α n) : + Id.run (l.foldlM f b) = l.foldl f b := (foldl_eq_foldlM f b l).symm + +@[simp] theorem id_run_foldrM (f : α → β → Id β) (b) (l : Vector α n) : + Id.run (l.foldrM f b) = l.foldr f b := (foldr_eq_foldrM f b l).symm + +/-! ### foldl and foldr -/ + /-! ### take -/ @[simp] theorem take_size (a : Vector α n) : a.take n = a.cast (by simp) := by From 97d07a54a353cc467e34470ca5417243cf3f9d9f Mon Sep 17 00:00:00 2001 From: Leonardo de Moura Date: Mon, 6 Jan 2025 17:53:04 -0800 Subject: [PATCH 040/100] feat: basic case-split for `grind` (#6559) This PR adds a basic case-splitting strategy for the `grind` tactic. We still need to add support for user customization. --- src/Init/Grind/Lemmas.lean | 3 + src/Init/Grind/Tactics.lean | 11 +- src/Lean/Meta/Tactic/Grind.lean | 3 + src/Lean/Meta/Tactic/Grind/Combinators.lean | 71 +++++++++++ src/Lean/Meta/Tactic/Grind/EMatch.lean | 20 +++- src/Lean/Meta/Tactic/Grind/Internalize.lean | 33 ++++++ src/Lean/Meta/Tactic/Grind/Intro.lean | 27 ++--- src/Lean/Meta/Tactic/Grind/Main.lean | 3 +- src/Lean/Meta/Tactic/Grind/Split.lean | 123 ++++++++++++++++++++ src/Lean/Meta/Tactic/Grind/Types.lean | 26 ++++- tests/lean/run/grind_implies.lean | 87 ++++++++++++++ tests/lean/run/grind_match1.lean | 10 ++ tests/lean/run/grind_pre.lean | 2 + tests/lean/run/grind_split.lean | 35 ++++++ 14 files changed, 418 insertions(+), 36 deletions(-) create mode 100644 src/Lean/Meta/Tactic/Grind/Combinators.lean create mode 100644 src/Lean/Meta/Tactic/Grind/Split.lean create mode 100644 tests/lean/run/grind_implies.lean create mode 100644 tests/lean/run/grind_split.lean diff --git a/src/Init/Grind/Lemmas.lean b/src/Init/Grind/Lemmas.lean index 5672b5618dae..b19515b7e6c9 100644 --- a/src/Init/Grind/Lemmas.lean +++ b/src/Init/Grind/Lemmas.lean @@ -25,6 +25,9 @@ theorem and_eq_of_eq_false_right {a b : Prop} (h : b = False) : (a ∧ b) = Fals theorem eq_true_of_and_eq_true_left {a b : Prop} (h : (a ∧ b) = True) : a = True := by simp_all theorem eq_true_of_and_eq_true_right {a b : Prop} (h : (a ∧ b) = True) : b = True := by simp_all +theorem or_of_and_eq_false {a b : Prop} (h : (a ∧ b) = False) : (¬a ∨ ¬b) := by + by_cases a <;> by_cases b <;> simp_all + /-! Or -/ theorem or_eq_of_eq_true_left {a b : Prop} (h : a = True) : (a ∨ b) = True := by simp [h] diff --git a/src/Init/Grind/Tactics.lean b/src/Init/Grind/Tactics.lean index 7def9222d434..3696eca31cbd 100644 --- a/src/Init/Grind/Tactics.lean +++ b/src/Init/Grind/Tactics.lean @@ -24,14 +24,9 @@ The configuration for `grind`. Passed to `grind` using, for example, the `grind (config := { eager := true })` syntax. -/ structure Config where - /-- When `eager` is true (default: `false`), `grind` eagerly splits `if-then-else` and `match` expressions during internalization. -/ - eager : Bool := false - /-- Maximum number of branches (i.e., case-splits) in the proof search tree. -/ - splits : Nat := 100 - /-- - Maximum number of E-matching (aka heuristic theorem instantiation) - in a proof search tree branch. - -/ + /-- Maximum number of case-splits in a proof search branch. It does not include splits performed during normalization. -/ + splits : Nat := 5 + /-- Maximum number of E-matching (aka heuristic theorem instantiation) rounds before each case split. -/ ematch : Nat := 5 /-- Maximum term generation. diff --git a/src/Lean/Meta/Tactic/Grind.lean b/src/Lean/Meta/Tactic/Grind.lean index 022695fb35ba..593bb147c6ab 100644 --- a/src/Lean/Meta/Tactic/Grind.lean +++ b/src/Lean/Meta/Tactic/Grind.lean @@ -39,6 +39,9 @@ builtin_initialize registerTraceClass `grind.ematch.instance builtin_initialize registerTraceClass `grind.ematch.instance.assignment builtin_initialize registerTraceClass `grind.issues builtin_initialize registerTraceClass `grind.simp +builtin_initialize registerTraceClass `grind.split +builtin_initialize registerTraceClass `grind.split.candidate +builtin_initialize registerTraceClass `grind.split.resolved /-! Trace options for `grind` developers -/ builtin_initialize registerTraceClass `grind.debug diff --git a/src/Lean/Meta/Tactic/Grind/Combinators.lean b/src/Lean/Meta/Tactic/Grind/Combinators.lean new file mode 100644 index 000000000000..cfeb0efb2a74 --- /dev/null +++ b/src/Lean/Meta/Tactic/Grind/Combinators.lean @@ -0,0 +1,71 @@ +/- +Copyright (c) 2025 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.Meta.Tactic.Grind.Types + +namespace Lean.Meta.Grind + +/-! +Combinators for manipulating `GrindTactic`s. +TODO: a proper tactic language for `grind`. +-/ + +def GrindTactic := Goal → GrindM (Option (List Goal)) + +def GrindTactic.try (x : GrindTactic) : GrindTactic := fun g => do + let some gs ← x g | return some [g] + return some gs + +def applyToAll (x : GrindTactic) (goals : List Goal) : GrindM (List Goal) := do + go goals [] +where + go (goals : List Goal) (acc : List Goal) : GrindM (List Goal) := do + match goals with + | [] => return acc.reverse + | goal :: goals => match (← x goal) with + | none => go goals (goal :: acc) + | some goals' => go goals (goals' ++ acc) + +partial def GrindTactic.andThen (x y : GrindTactic) : GrindTactic := fun goal => do + let some goals ← x goal | return none + applyToAll y goals + +instance : AndThen GrindTactic where + andThen a b := GrindTactic.andThen a (b ()) + +partial def GrindTactic.iterate (x : GrindTactic) : GrindTactic := fun goal => do + go [goal] [] +where + go (todo : List Goal) (result : List Goal) : GrindM (List Goal) := do + match todo with + | [] => return result + | goal :: todo => + if let some goalsNew ← x goal then + go (goalsNew ++ todo) result + else + go todo (goal :: result) + +partial def GrindTactic.orElse (x y : GrindTactic) : GrindTactic := fun goal => do + let some goals ← x goal | y goal + return goals + +instance : OrElse GrindTactic where + orElse a b := GrindTactic.andThen a (b ()) + +def toGrindTactic (f : GoalM Unit) : GrindTactic := fun goal => do + let goal ← GoalM.run' goal f + if goal.inconsistent then + return some [] + else + return some [goal] + +def GrindTactic' := Goal → GrindM (List Goal) + +def GrindTactic'.toGrindTactic (x : GrindTactic') : GrindTactic := fun goal => do + let goals ← x goal + return some goals + +end Lean.Meta.Grind diff --git a/src/Lean/Meta/Tactic/Grind/EMatch.lean b/src/Lean/Meta/Tactic/Grind/EMatch.lean index 64b716c698bc..4ba4d6eff980 100644 --- a/src/Lean/Meta/Tactic/Grind/EMatch.lean +++ b/src/Lean/Meta/Tactic/Grind/EMatch.lean @@ -51,6 +51,8 @@ structure Context where useMT : Bool := true /-- `EMatchTheorem` being processed. -/ thm : EMatchTheorem := default + /-- Initial application used to start E-matching -/ + initApp : Expr := default deriving Inhabited /-- State for the E-matching monad -/ @@ -64,6 +66,9 @@ abbrev M := ReaderT Context $ StateRefT State GoalM def M.run' (x : M α) : GoalM α := x {} |>.run' {} +@[inline] private abbrev withInitApp (e : Expr) (x : M α) : M α := + withReader (fun ctx => { ctx with initApp := e }) x + /-- Assigns `bidx := e` in `c`. If `bidx` is already assigned in `c`, we check whether `e` and `c.assignment[bidx]` are in the same equivalence class. @@ -213,6 +218,8 @@ private def addNewInstance (origin : Origin) (proof : Expr) (generation : Nat) : check proof let mut prop ← inferType proof if Match.isMatchEqnTheorem (← getEnv) origin.key then + -- `initApp` is a match-application that we don't need to split at anymore. + markCaseSplitAsResolved (← read).initApp prop ← annotateMatchEqnType prop trace[grind.ematch.instance] "{← origin.pp}: {prop}" addTheoremInstance proof prop (generation+1) @@ -288,9 +295,10 @@ private def main (p : Expr) (cnstrs : List Cnstr) : M Unit := do let n ← getENode app if (n.heqProofs || n.isCongrRoot) && (!useMT || n.mt == gmt) then - if let some c ← matchArgs? { cnstrs, assignment, gen := n.generation } p app |>.run then - modify fun s => { s with choiceStack := [c] } - processChoices + withInitApp app do + if let some c ← matchArgs? { cnstrs, assignment, gen := n.generation } p app |>.run then + modify fun s => { s with choiceStack := [c] } + processChoices def ematchTheorem (thm : EMatchTheorem) : M Unit := do if (← checkMaxInstancesExceeded) then return () @@ -341,14 +349,14 @@ def ematch : GoalM Unit := do } /-- Performs one round of E-matching, and assert new instances. -/ -def ematchAndAssert? (goal : Goal) : GrindM (Option (List Goal)) := do +def ematchAndAssert : GrindTactic := fun goal => do let numInstances := goal.numInstances let goal ← GoalM.run' goal ematch if goal.numInstances == numInstances then return none assertAll goal -def ematchStar (goal : Goal) : GrindM (List Goal) := do - iterate goal ematchAndAssert? +def ematchStar : GrindTactic := + ematchAndAssert.iterate end Lean.Meta.Grind diff --git a/src/Lean/Meta/Tactic/Grind/Internalize.lean b/src/Lean/Meta/Tactic/Grind/Internalize.lean index fe801763f326..75b474303380 100644 --- a/src/Lean/Meta/Tactic/Grind/Internalize.lean +++ b/src/Lean/Meta/Tactic/Grind/Internalize.lean @@ -31,6 +31,10 @@ def addCongrTable (e : Expr) : GoalM Unit := do else modify fun s => { s with congrTable := s.congrTable.insert { e } } +/-- +Given an application `e` of the form `f a_1 ... a_n`, +adds entry `f ↦ e` to `appMap`. Recall that `appMap` is a multi-map. +-/ private def updateAppMap (e : Expr) : GoalM Unit := do let key := e.toHeadIndex modify fun s => { s with @@ -40,6 +44,34 @@ private def updateAppMap (e : Expr) : GoalM Unit := do s.appMap.insert key [e] } +/-- Inserts `e` into the list of case-split candidates. -/ +private def addSplitCandidate (e : Expr) : GoalM Unit := do + trace[grind.split.candidate] "{e}" + modify fun s => { s with splitCadidates := e :: s.splitCadidates } + +-- TODO: add attribute to make this extensible +private def forbiddenSplitTypes := [``Eq, ``HEq, ``True, ``False] + +/-- Inserts `e` into the list of case-split candidates if applicable. -/ +private def checkAndaddSplitCandidate (e : Expr) : GoalM Unit := do + unless e.isApp do return () + if e.isIte || e.isDIte then + addSplitCandidate e + else if (← isMatcherApp e) then + if let .reduced _ ← reduceMatcher? e then + -- When instantiating `match`-equations, we add `match`-applications that can be reduced, + -- and consequently don't need to be splitted. + return () + else + addSplitCandidate e + else + let .const declName _ := e.getAppFn | return () + if forbiddenSplitTypes.contains declName then return () + -- We should have a mechanism for letting users to select types to case-split. + -- Right now, we just consider inductive predicates that are not in the forbidden list + if (← isInductivePredicate declName) then + addSplitCandidate e + mutual /-- Internalizes the nested ground terms in the given pattern. -/ private partial def internalizePattern (pattern : Expr) (generation : Nat) : GoalM Expr := do @@ -116,6 +148,7 @@ partial def internalize (e : Expr) (generation : Nat) : GoalM Unit := do -- We do not want to internalize the components of a literal value. mkENode e generation else e.withApp fun f args => do + checkAndaddSplitCandidate e addMatchEqns f generation if f.isConstOf ``Lean.Grind.nestedProof && args.size == 2 then -- We only internalize the proposition. We can skip the proof because of diff --git a/src/Lean/Meta/Tactic/Grind/Intro.lean b/src/Lean/Meta/Tactic/Grind/Intro.lean index 3923b4d1abbe..1fa58b4a60f3 100644 --- a/src/Lean/Meta/Tactic/Grind/Intro.lean +++ b/src/Lean/Meta/Tactic/Grind/Intro.lean @@ -11,6 +11,7 @@ import Lean.Meta.Tactic.Grind.Types import Lean.Meta.Tactic.Grind.Cases import Lean.Meta.Tactic.Grind.Injection import Lean.Meta.Tactic.Grind.Core +import Lean.Meta.Tactic.Grind.Combinators namespace Lean.Meta.Grind @@ -88,7 +89,7 @@ private def applyInjection? (goal : Goal) (fvarId : FVarId) : MetaM (Option Goal return none /-- Introduce new hypotheses (and apply `by_contra`) until goal is of the form `... ⊢ False` -/ -partial def intros (goal : Goal) (generation : Nat) : GrindM (List Goal) := do +partial def intros (generation : Nat) : GrindTactic' := fun goal => do let rec go (goal : Goal) : StateRefT (Array Goal) GrindM Unit := do if goal.inconsistent then return () @@ -116,11 +117,11 @@ partial def intros (goal : Goal) (generation : Nat) : GrindM (List Goal) := do return goals.toList /-- Asserts a new fact `prop` with proof `proof` to the given `goal`. -/ -def assertAt (goal : Goal) (proof : Expr) (prop : Expr) (generation : Nat) : GrindM (List Goal) := do +def assertAt (proof : Expr) (prop : Expr) (generation : Nat) : GrindTactic' := fun goal => do if (← isCasesCandidate prop) then let mvarId ← goal.mvarId.assert (← mkFreshUserName `h) prop proof let goal := { goal with mvarId } - intros goal generation + intros generation goal else let goal ← GoalM.run' goal do let r ← simp prop @@ -130,25 +131,13 @@ def assertAt (goal : Goal) (proof : Expr) (prop : Expr) (generation : Nat) : Gri if goal.inconsistent then return [] else return [goal] /-- Asserts next fact in the `goal` fact queue. -/ -def assertNext? (goal : Goal) : GrindM (Option (List Goal)) := do +def assertNext : GrindTactic := fun goal => do let some (fact, newFacts) := goal.newFacts.dequeue? | return none - assertAt { goal with newFacts } fact.proof fact.prop fact.generation - -partial def iterate (goal : Goal) (f : Goal → GrindM (Option (List Goal))) : GrindM (List Goal) := do - go [goal] [] -where - go (todo : List Goal) (result : List Goal) : GrindM (List Goal) := do - match todo with - | [] => return result - | goal :: todo => - if let some goalsNew ← f goal then - go (goalsNew ++ todo) result - else - go todo (goal :: result) + assertAt fact.proof fact.prop fact.generation { goal with newFacts } /-- Asserts all facts in the `goal` fact queue. -/ -partial def assertAll (goal : Goal) : GrindM (List Goal) := do - iterate goal assertNext? +partial def assertAll : GrindTactic := + assertNext.iterate end Lean.Meta.Grind diff --git a/src/Lean/Meta/Tactic/Grind/Main.lean b/src/Lean/Meta/Tactic/Grind/Main.lean index 139b32ecebb7..60a8fa0e95c3 100644 --- a/src/Lean/Meta/Tactic/Grind/Main.lean +++ b/src/Lean/Meta/Tactic/Grind/Main.lean @@ -14,6 +14,7 @@ import Lean.Meta.Tactic.Grind.Util import Lean.Meta.Tactic.Grind.Inv import Lean.Meta.Tactic.Grind.Intro import Lean.Meta.Tactic.Grind.EMatch +import Lean.Meta.Tactic.Grind.Split import Lean.Meta.Tactic.Grind.SimpUtil namespace Lean.Meta.Grind @@ -68,7 +69,7 @@ def all (goals : List Goal) (f : Goal → GrindM (List Goal)) : GrindM (List Goa /-- A very simple strategy -/ private def simple (goals : List Goal) : GrindM (List Goal) := do - all goals ematchStar + applyToAll (ematchStar >> (splitNext >> ematchStar).iterate) goals def main (mvarId : MVarId) (config : Grind.Config) (mainDeclName : Name) (fallback : Fallback) : MetaM (List MVarId) := do let go : GrindM (List MVarId) := do diff --git a/src/Lean/Meta/Tactic/Grind/Split.lean b/src/Lean/Meta/Tactic/Grind/Split.lean new file mode 100644 index 000000000000..7c98474232be --- /dev/null +++ b/src/Lean/Meta/Tactic/Grind/Split.lean @@ -0,0 +1,123 @@ +/- +Copyright (c) 2025 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.Meta.Tactic.Grind.Types +import Lean.Meta.Tactic.Grind.Intro +import Lean.Meta.Tactic.Grind.Cases + +namespace Lean.Meta.Grind + +inductive CaseSplitStatus where + | resolved + | notReady + | ready + deriving Inhabited, BEq + +private def checkCaseSplitStatus (e : Expr) : GoalM CaseSplitStatus := do + match_expr e with + | Or a b => + if (← isEqTrue e) then + if (← isEqTrue a <||> isEqTrue b) then + return .resolved + else + return .ready + else if (← isEqFalse e) then + return .resolved + else + return .notReady + | And a b => + if (← isEqTrue e) then + return .resolved + else if (← isEqFalse e) then + if (← isEqFalse a <||> isEqFalse b) then + return .resolved + else + return .ready + else + return .notReady + | ite _ c _ _ _ => + if (← isEqTrue c <||> isEqFalse c) then + return .resolved + else + return .ready + | dite _ c _ _ _ => + if (← isEqTrue c <||> isEqFalse c) then + return .resolved + else + return .ready + | _ => + if (← isResolvedCaseSplit e) then + return .resolved + if (← isMatcherApp e) then + return .notReady -- TODO: implement splittes for `match` + -- return .ready + let .const declName .. := e.getAppFn | unreachable! + if (← isInductivePredicate declName <&&> isEqTrue e) then + return .ready + return .notReady + +/-- Returns the next case-split to be peformed. It uses a very simple heuristic. -/ +private def selectNextSplit? : GoalM (Option Expr) := do + if (← isInconsistent) then return none + if (← checkMaxCaseSplit) then return none + go (← get).splitCadidates none [] +where + go (cs : List Expr) (c? : Option Expr) (cs' : List Expr) : GoalM (Option Expr) := do + match cs with + | [] => + modify fun s => { s with splitCadidates := cs'.reverse } + if c?.isSome then + -- Remark: we reset `numEmatch` after each case split. + -- We should consider other strategies in the future. + modify fun s => { s with numSplits := s.numSplits + 1, numEmatch := 0 } + return c? + | c::cs => + match (← checkCaseSplitStatus c) with + | .notReady => go cs c? (c::cs') + | .resolved => go cs c? cs' + | .ready => + match c? with + | none => go cs (some c) cs' + | some c' => + if (← getGeneration c) < (← getGeneration c') then + go cs (some c) (c'::cs') + else + go cs c? (c::cs') + +/-- Constructs a major premise for the `cases` tactic used by `grind`. -/ +private def mkCasesMajor (c : Expr) : GoalM Expr := do + match_expr c with + | And a b => return mkApp3 (mkConst ``Grind.or_of_and_eq_false) a b (← mkEqFalseProof c) + | ite _ c _ _ _ => return mkEM c + | dite _ c _ _ _ => return mkEM c + | _ => return mkApp2 (mkConst ``of_eq_true) c (← mkEqTrueProof c) + +/-- Introduces new hypotheses in each goal. -/ +private def introNewHyp (goals : List Goal) (acc : List Goal) (generation : Nat) : GrindM (List Goal) := do + match goals with + | [] => return acc.reverse + | goal::goals => introNewHyp goals ((← intros generation goal) ++ acc) generation + +/-- +Selects a case-split from the list of candidates, +and returns a new list of goals if successful. +-/ +def splitNext : GrindTactic := fun goal => do + let (goals?, _) ← GoalM.run goal do + let some c ← selectNextSplit? + | return none + let gen ← getGeneration c + trace[grind.split] "{c}, generation: {gen}" + -- TODO: `match` + let major ← mkCasesMajor c + let mvarIds ← cases (← get).mvarId major + let goal ← get + let goals := mvarIds.map fun mvarId => { goal with mvarId } + let goals ← introNewHyp goals [] (gen+1) + return some goals + return goals? + +end Lean.Meta.Grind diff --git a/src/Lean/Meta/Tactic/Grind/Types.lean b/src/Lean/Meta/Tactic/Grind/Types.lean index 80b934fc186d..4544196da57f 100644 --- a/src/Lean/Meta/Tactic/Grind/Types.lean +++ b/src/Lean/Meta/Tactic/Grind/Types.lean @@ -374,8 +374,7 @@ structure Goal where thmMap : EMatchTheorems /-- Number of theorem instances generated so far -/ numInstances : Nat := 0 - /-- Number of E-matching rounds performed in this goal so far. -/ - -- Remark: it is always equal to `gmt` in the current implementation. + /-- Number of E-matching rounds performed in this goal since the last case-split. -/ numEmatch : Nat := 0 /-- (pre-)instances found so far. It includes instances that failed to be instantiated. -/ preInstances : PreInstanceSet := {} @@ -383,6 +382,12 @@ structure Goal where newFacts : Std.Queue NewFact := ∅ /-- `match` auxiliary functions whose equations have already been created and activated. -/ matchEqNames : PHashSet Name := {} + /-- Case-split candidates. -/ + splitCadidates : List Expr := [] + /-- Number of splits performed to get to this goal. -/ + numSplits : Nat := 0 + /-- Case-splits that do not have to be performed anymore. -/ + resolvedSplits : PHashSet ENodeKey := {} deriving Inhabited def Goal.admit (goal : Goal) : MetaM Unit := @@ -420,6 +425,10 @@ def addTheoremInstance (proof : Expr) (prop : Expr) (generation : Nat) : GoalM U def checkMaxInstancesExceeded : GoalM Bool := do return (← get).numInstances >= (← getConfig).instances +/-- Returns `true` if the maximum number of case-splits has been reached. -/ +def checkMaxCaseSplit : GoalM Bool := do + return (← get).numSplits >= (← getConfig).splits + /-- Returns `true` if the maximum number of E-matching rounds has been reached. -/ def checkMaxEmatchExceeded : GoalM Bool := do return (← get).numEmatch >= (← getConfig).ematch @@ -732,4 +741,17 @@ partial def getEqcs : GoalM (List (List Expr)) := do r := (← getEqc node.self) :: r return r +/-- Returns `true` if `e` is a case-split that does not need to be performed anymore. -/ +def isResolvedCaseSplit (e : Expr) : GoalM Bool := + return (← get).resolvedSplits.contains { expr := e } + +/-- +Mark `e` as a case-split that does not need to be performed anymore. +Remark: we currently use this feature to disable `match`-case-splits +-/ +def markCaseSplitAsResolved (e : Expr) : GoalM Unit := do + unless (← isResolvedCaseSplit e) do + trace[grind.split.resolved] "{e}" + modify fun s => { s with resolvedSplits := s.resolvedSplits.insert { expr := e } } + end Lean.Meta.Grind diff --git a/tests/lean/run/grind_implies.lean b/tests/lean/run/grind_implies.lean new file mode 100644 index 000000000000..9e68cc2e5572 --- /dev/null +++ b/tests/lean/run/grind_implies.lean @@ -0,0 +1,87 @@ +set_option trace.grind.eqc true +set_option trace.grind.internalize true + +/-- +info: [grind.internalize] p → q +[grind.internalize] p +[grind.internalize] q +[grind.eqc] (p → q) = True +[grind.eqc] p = True +[grind.eqc] (p → q) = q +[grind.eqc] q = False +-/ +#guard_msgs (info) in +example (p q : Prop) : (p → q) → p → q := by + grind + +/-- +info: [grind.internalize] p → q +[grind.internalize] p +[grind.internalize] q +[grind.eqc] (p → q) = True +[grind.eqc] q = False +[grind.eqc] p = True +[grind.eqc] (p → q) = q +-/ +#guard_msgs (info) in +example (p q : Prop) : (p → q) → ¬q → ¬p := by + grind + +/-- +info: [grind.internalize] p → q +[grind.internalize] p +[grind.internalize] q +[grind.internalize] r +[grind.eqc] (p → q) = r +[grind.eqc] p = False +[grind.eqc] (p → q) = True +[grind.eqc] r = False +-/ +#guard_msgs (info) in +example (p q : Prop) : (p → q) = r → ¬p → r := by + grind + + +/-- +info: [grind.internalize] p → q +[grind.internalize] p +[grind.internalize] q +[grind.internalize] r +[grind.eqc] (p → q) = r +[grind.eqc] q = True +[grind.eqc] (p → q) = True +[grind.eqc] r = False +-/ +#guard_msgs (info) in +example (p q : Prop) : (p → q) = r → q → r := by + grind + +/-- +info: [grind.internalize] p → q +[grind.internalize] p +[grind.internalize] q +[grind.internalize] r +[grind.eqc] (p → q) = r +[grind.eqc] q = False +[grind.eqc] r = True +[grind.eqc] p = True +[grind.eqc] (p → q) = q +-/ +#guard_msgs (info) in +example (p q : Prop) : (p → q) = r → ¬q → r → ¬p := by + grind + +/-- +info: [grind.internalize] p → q +[grind.internalize] p +[grind.internalize] q +[grind.internalize] r +[grind.eqc] (p → q) = r +[grind.eqc] q = False +[grind.eqc] r = False +[grind.eqc] p = True +[grind.eqc] p = False +-/ +#guard_msgs (info) in +example (p q : Prop) : (p → q) = r → ¬q → ¬r → p := by + grind diff --git a/tests/lean/run/grind_match1.lean b/tests/lean/run/grind_match1.lean index 483b8600c05c..c50bfcf6e618 100644 --- a/tests/lean/run/grind_match1.lean +++ b/tests/lean/run/grind_match1.lean @@ -7,6 +7,8 @@ def g (xs : List α) (ys : List α) := attribute [simp] g set_option trace.grind.assert true +set_option trace.grind.split.candidate true +set_option trace.grind.split.resolved true /-- info: [grind.assert] (match as, bs with @@ -14,10 +16,18 @@ info: [grind.assert] (match as, bs with | head :: head_1 :: tail, [] => [] | x :: xs, ys => x :: g xs ys) = d +[grind.split.candidate] match as, bs with + | [], x => bs + | head :: head_1 :: tail, [] => [] + | x :: xs, ys => x :: g xs ys [grind.assert] bs = [] [grind.assert] a₁ :: f 0 = as [grind.assert] f 0 = a₂ :: f 1 [grind.assert] ¬d = [] +[grind.split.resolved] match as, bs with + | [], x => bs + | head :: head_1 :: tail, [] => [] + | x :: xs, ys => x :: g xs ys [grind.assert] (match a₁ :: a₂ :: f 1, [] with | [], x => bs | head :: head_1 :: tail, [] => [] diff --git a/tests/lean/run/grind_pre.lean b/tests/lean/run/grind_pre.lean index 63c247088be1..be879cd20222 100644 --- a/tests/lean/run/grind_pre.lean +++ b/tests/lean/run/grind_pre.lean @@ -12,6 +12,8 @@ right✝ : b = true ∨ c = true left : p right : q x✝ : b = false ∨ a = false +h✝ : b = false +h : c = true ⊢ False -/ #guard_msgs (error) in diff --git a/tests/lean/run/grind_split.lean b/tests/lean/run/grind_split.lean new file mode 100644 index 000000000000..917e62e6b7f6 --- /dev/null +++ b/tests/lean/run/grind_split.lean @@ -0,0 +1,35 @@ +set_option trace.grind.split true +set_option trace.grind.eqc true +example (p q : Prop) : p ∨ q → p ∨ ¬q → ¬p ∨ q → ¬p ∨ ¬q → False := by + grind + +opaque R : Nat → Prop + +example (p : Prop) [Decidable p] (a b c : Nat) : (if p then a else b) = c → R a → R b → R c := by + grind + + +namespace grind_test_induct_pred + +inductive Expr where + | nat : Nat → Expr + | plus : Expr → Expr → Expr + | bool : Bool → Expr + | and : Expr → Expr → Expr + deriving DecidableEq + +inductive Ty where + | nat + | bool + deriving DecidableEq + +inductive HasType : Expr → Ty → Prop + | nat : HasType (.nat v) .nat + | plus : HasType a .nat → HasType b .nat → HasType (.plus a b) .nat + | bool : HasType (.bool v) .bool + | and : HasType a .bool → HasType b .bool → HasType (.and a b) .bool + +theorem HasType.det (h₁ : HasType e t₁) (h₂ : HasType e t₂) : t₁ = t₂ := by + grind + +end grind_test_induct_pred From a2a525f5c7efa2b4604c4bcd1c655c17ab489944 Mon Sep 17 00:00:00 2001 From: Sebastian Ullrich Date: Tue, 7 Jan 2025 14:06:24 +0000 Subject: [PATCH 041/100] fix: set absolute linker path (#6547) This PR should prevent Lake from accidentally picking up other linkers installed on the machine. --- script/prepare-llvm-linux.sh | 2 +- script/prepare-llvm-macos.sh | 9 +++++++-- script/prepare-llvm-mingw.sh | 2 +- 3 files changed, 9 insertions(+), 4 deletions(-) diff --git a/script/prepare-llvm-linux.sh b/script/prepare-llvm-linux.sh index ff4b11f43180..2ebb7a41b88e 100755 --- a/script/prepare-llvm-linux.sh +++ b/script/prepare-llvm-linux.sh @@ -64,7 +64,7 @@ fi # use `-nostdinc` to make sure headers are not visible by default (in particular, not to `#include_next` in the clang headers), # but do not change sysroot so users can still link against system libs echo -n " -DLEANC_INTERNAL_FLAGS='-nostdinc -isystem ROOT/include/clang' -DLEANC_CC=ROOT/bin/clang" -echo -n " -DLEANC_INTERNAL_LINKER_FLAGS='-L ROOT/lib -L ROOT/lib/glibc ROOT/lib/glibc/libc_nonshared.a ROOT/lib/glibc/libpthread_nonshared.a -Wl,--as-needed -Wl,-Bstatic -lgmp -lunwind -luv -Wl,-Bdynamic -Wl,--no-as-needed -fuse-ld=lld'" +echo -n " -DLEANC_INTERNAL_LINKER_FLAGS='-L ROOT/lib -L ROOT/lib/glibc ROOT/lib/glibc/libc_nonshared.a ROOT/lib/glibc/libpthread_nonshared.a -Wl,--as-needed -Wl,-Bstatic -lgmp -lunwind -luv -Wl,-Bdynamic -Wl,--no-as-needed -fuse-ld=lld --ld-path=ROOT/bin/ld.lld'" # when not using the above flags, link GMP dynamically/as usual echo -n " -DLEAN_EXTRA_LINKER_FLAGS='-Wl,--as-needed -lgmp -luv -lpthread -ldl -lrt -Wl,--no-as-needed'" # do not set `LEAN_CC` for tests diff --git a/script/prepare-llvm-macos.sh b/script/prepare-llvm-macos.sh index d3cdc61b3f2e..8c15fb115db4 100755 --- a/script/prepare-llvm-macos.sh +++ b/script/prepare-llvm-macos.sh @@ -48,12 +48,17 @@ if [[ -L llvm-host ]]; then echo -n " -DCMAKE_C_COMPILER=$PWD/stage1/bin/clang" gcp $GMP/lib/libgmp.a stage1/lib/ gcp $LIBUV/lib/libuv.a stage1/lib/ - echo -n " -DLEANC_INTERNAL_LINKER_FLAGS='-L ROOT/lib -L ROOT/lib/libc -fuse-ld=lld'" echo -n " -DLEAN_EXTRA_LINKER_FLAGS='-lgmp -luv'" else echo -n " -DCMAKE_C_COMPILER=$PWD/llvm-host/bin/clang -DLEANC_OPTS='--sysroot $PWD/stage1 -resource-dir $PWD/stage1/lib/clang/15.0.1 ${EXTRA_FLAGS:-}'" - echo -n " -DLEANC_INTERNAL_LINKER_FLAGS='-L ROOT/lib -L ROOT/lib/libc -fuse-ld=lld'" fi echo -n " -DLEANC_INTERNAL_FLAGS='-nostdinc -isystem ROOT/include/clang' -DLEANC_CC=ROOT/bin/clang" +if [[ "$(uname -p)" == "i386" ]]; then + # `--ld-path` creates some platform_version troubles on macOS x64 but this is not a high-prio platform anymore and + # its users are not that likely to have conflicting `lld`s in their `PATH` + echo -n " -DLEANC_INTERNAL_LINKER_FLAGS='-L ROOT/lib -L ROOT/lib/libc -fuse-ld=lld'" +else + echo -n " -DLEANC_INTERNAL_LINKER_FLAGS='-L ROOT/lib -L ROOT/lib/libc -fuse-ld=lld --ld-path=ROOT/bin/ld64.lld'" +fi # do not set `LEAN_CC` for tests echo -n " -DLEAN_TEST_VARS=''" diff --git a/script/prepare-llvm-mingw.sh b/script/prepare-llvm-mingw.sh index 0e2745971801..0bd255f9edf3 100644 --- a/script/prepare-llvm-mingw.sh +++ b/script/prepare-llvm-mingw.sh @@ -43,7 +43,7 @@ echo -n " -DCMAKE_C_COMPILER=$PWD/stage1/bin/clang.exe -DCMAKE_C_COMPILER_WORKS= echo -n " -DSTAGE0_CMAKE_C_COMPILER=clang -DSTAGE0_CMAKE_CXX_COMPILER=clang++" echo -n " -DLEAN_EXTRA_CXX_FLAGS='--sysroot $PWD/llvm -idirafter /clang64/include/'" echo -n " -DLEANC_INTERNAL_FLAGS='--sysroot ROOT -nostdinc -isystem ROOT/include/clang' -DLEANC_CC=ROOT/bin/clang.exe" -echo -n " -DLEANC_INTERNAL_LINKER_FLAGS='-L ROOT/lib -Wl,-Bstatic -lgmp $(pkg-config --static --libs libuv) -lunwind -Wl,-Bdynamic -fuse-ld=lld'" +echo -n " -DLEANC_INTERNAL_LINKER_FLAGS='-L ROOT/lib -Wl,-Bstatic -lgmp $(pkg-config --static --libs libuv) -lunwind -Wl,-Bdynamic -fuse-ld=lld --ld-path=ROOT/bin/ld.lld'" # when not using the above flags, link GMP dynamically/as usual. Always link ICU dynamically. echo -n " -DLEAN_EXTRA_LINKER_FLAGS='-lgmp $(pkg-config --libs libuv) -lucrtbase'" # do not set `LEAN_CC` for tests From 83098cdaec01dc982a4b5f81ad621e97ee413f42 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Wed, 8 Jan 2025 01:25:01 +1100 Subject: [PATCH 042/100] chore: typos / improvements to grind messages (#6561) This PR fixes some typos and makes minor improvements to grind doc-strings and messages. --- src/Init/Grind/Tactics.lean | 2 +- src/Init/Grind/Util.lean | 2 - src/Lean/Meta/Tactic/Grind/Canon.lean | 12 ++--- src/Lean/Meta/Tactic/Grind/EMatchTheorem.lean | 48 +++++++++++-------- src/Lean/Meta/Tactic/Grind/Internalize.lean | 6 +-- src/Lean/Meta/Tactic/Grind/Split.lean | 8 ++-- src/Lean/Meta/Tactic/Grind/Types.lean | 12 ++--- 7 files changed, 49 insertions(+), 41 deletions(-) diff --git a/src/Init/Grind/Tactics.lean b/src/Init/Grind/Tactics.lean index 3696eca31cbd..f4678adfc688 100644 --- a/src/Init/Grind/Tactics.lean +++ b/src/Init/Grind/Tactics.lean @@ -21,7 +21,7 @@ end Lean.Parser.Attr namespace Lean.Grind /-- The configuration for `grind`. -Passed to `grind` using, for example, the `grind (config := { eager := true })` syntax. +Passed to `grind` using, for example, the `grind (config := { matchEqs := true })` syntax. -/ structure Config where /-- Maximum number of case-splits in a proof search branch. It does not include splits performed during normalization. -/ diff --git a/src/Init/Grind/Util.lean b/src/Init/Grind/Util.lean index 9a231e318162..9ef005388f82 100644 --- a/src/Init/Grind/Util.lean +++ b/src/Init/Grind/Util.lean @@ -21,8 +21,6 @@ def doNotSimp {α : Sort u} (a : α) : α := a /-- Gadget for representing offsets `t+k` in patterns. -/ def offset (a b : Nat) : Nat := a + b -set_option pp.proofs true - theorem nestedProof_congr (p q : Prop) (h : p = q) (hp : p) (hq : q) : HEq (nestedProof p hp) (nestedProof q hq) := by subst h; apply HEq.refl diff --git a/src/Lean/Meta/Tactic/Grind/Canon.lean b/src/Lean/Meta/Tactic/Grind/Canon.lean index 05057121709c..6ac02c336ebd 100644 --- a/src/Lean/Meta/Tactic/Grind/Canon.lean +++ b/src/Lean/Meta/Tactic/Grind/Canon.lean @@ -22,7 +22,7 @@ to detect when two structurally different atoms are definitionally equal. The `grind` tactic, on the other hand, uses congruence closure. Moreover, types, type formers, proofs, and instances are considered supporting elements and are not factored into congruence detection. -This module minimizes the number of `isDefEq` checks by comparing two terms `a` and `b` only if they instances, +This module minimizes the number of `isDefEq` checks by comparing two terms `a` and `b` only if they are instances, types, or type formers and are the `i`-th arguments of two different `f`-applications. This approach is sufficient for the congruence closure procedure used by the `grind` tactic. @@ -129,17 +129,17 @@ where else let pinfos := (← getFunInfo f).paramInfo let mut modified := false - let mut args := args - for i in [:args.size] do - let arg := args[i]! + let mut args := args.toVector + for h : i in [:args.size] do + let arg := args[i] let arg' ← match (← shouldCanon pinfos i arg) with | .canonType => canonType f i arg | .canonInst => canonInst f i arg | .visit => visit arg unless ptrEq arg arg' do - args := args.set! i arg' + args := args.set i arg' modified := true - pure <| if modified then mkAppN f args else e + pure <| if modified then mkAppN f args.toArray else e | .forallE _ d b _ => -- Recall that we have `ForallProp.lean`. let d' ← visit d diff --git a/src/Lean/Meta/Tactic/Grind/EMatchTheorem.lean b/src/Lean/Meta/Tactic/Grind/EMatchTheorem.lean index 9605c8e116e7..c8299413f00f 100644 --- a/src/Lean/Meta/Tactic/Grind/EMatchTheorem.lean +++ b/src/Lean/Meta/Tactic/Grind/EMatchTheorem.lean @@ -218,7 +218,7 @@ private def getPatternFn? (pattern : Expr) : Option Expr := | _ => none /-- -Returns a bit-mask `mask` s.t. `mask[i]` is true if the the corresponding argument is +Returns a bit-mask `mask` s.t. `mask[i]` is true if the corresponding argument is - a type (that is not a proposition) or type former, or - a proof, or - an instance implicit argument @@ -248,13 +248,13 @@ private partial def go (pattern : Expr) (root := false) : M Expr := do | throwError "invalid pattern, (non-forbidden) application expected{indentExpr pattern}" assert! f.isConst || f.isFVar saveSymbol f.toHeadIndex - let mut args := pattern.getAppArgs + let mut args := pattern.getAppArgs.toVector let supportMask ← getPatternSupportMask f args.size - for i in [:args.size] do - let arg := args[i]! + for h : i in [:args.size] do + let arg := args[i] let isSupport := supportMask[i]?.getD false - args := args.set! i (← goArg arg isSupport) - return mkAppN f args + args := args.set i (← goArg arg isSupport) + return mkAppN f args.toArray where goArg (arg : Expr) (isSupport : Bool) : M Expr := do if !arg.hasLooseBVars then @@ -556,8 +556,8 @@ private partial def collect (e : Expr) : CollectorM Unit := do -- restore state and continue search set saved catch _ => - -- restore state and continue search trace[grind.ematch.pattern.search] "skip, exception during normalization" + -- restore state and continue search set saved let args := e.getAppArgs for arg in args, flag in (← NormalizePattern.getPatternSupportMask f args.size) do @@ -592,7 +592,7 @@ private def mkEMatchTheoremWithKind? (origin : Origin) (levelParams : Array Name | .fwd => let ps ← getPropTypes xs if ps.isEmpty then - throwError "invalid `grind` forward theorem, theorem `{← origin.pp}` does not have proposional hypotheses" + throwError "invalid `grind` forward theorem, theorem `{← origin.pp}` does not have propositional hypotheses" pure ps | .bwd => pure #[type] | .default => pure <| #[type] ++ (← getPropTypes xs) @@ -623,7 +623,7 @@ private def getKind (stx : Syntax) : TheoremKind := else .bwd -private def addGrindEqAttr (declName : Name) (attrKind : AttributeKind) (useLhs := true) : MetaM Unit := do +private def addGrindEqAttr (declName : Name) (attrKind : AttributeKind) (thmKind : TheoremKind) (useLhs := true) : MetaM Unit := do if (← getConstInfo declName).isTheorem then ematchTheoremsExt.add (← mkEMatchEqTheorem declName (normalizePattern := true) (useLhs := useLhs)) attrKind else if let some eqns ← getEqnsFor? declName then @@ -632,18 +632,18 @@ private def addGrindEqAttr (declName : Name) (attrKind : AttributeKind) (useLhs for eqn in eqns do ematchTheoremsExt.add (← mkEMatchEqTheorem eqn) attrKind else - throwError "`[grind_eq]` attribute can only be applied to equational theorems or function definitions" + throwError s!"`{thmKind.toAttribute}` attribute can only be applied to equational theorems or function definitions" private def addGrindAttr (declName : Name) (attrKind : AttributeKind) (thmKind : TheoremKind) : MetaM Unit := do if thmKind == .eqLhs then - addGrindEqAttr declName attrKind (useLhs := true) + addGrindEqAttr declName attrKind thmKind (useLhs := true) else if thmKind == .eqRhs then - addGrindEqAttr declName attrKind (useLhs := false) + addGrindEqAttr declName attrKind thmKind (useLhs := false) else if thmKind == .eqBoth then - addGrindEqAttr declName attrKind (useLhs := true) - addGrindEqAttr declName attrKind (useLhs := false) + addGrindEqAttr declName attrKind thmKind (useLhs := true) + addGrindEqAttr declName attrKind thmKind (useLhs := false) else if !(← getConstInfo declName).isTheorem then - addGrindEqAttr declName attrKind + addGrindEqAttr declName attrKind thmKind else let some thm ← mkEMatchTheoremWithKind? (.decl declName) #[] (← getProofFor declName) thmKind | throwError "`@{thmKind.toAttribute} theorem {declName}` {thmKind.explainFailure}, consider using different options or the `grind_pattern` command" @@ -653,11 +653,21 @@ builtin_initialize registerBuiltinAttribute { name := `grind descr := - "The `[grind_eq]` attribute is used to annotate equational theorems and functions.\ - When applied to an equational theorem, it marks the theorem for use in heuristic instantiations by the `grind` tactic.\ - When applied to a function, it automatically annotates the equational theorems associated with that function.\ + "The `[grind]` attribute is used to annotate declarations.\ + \ + When applied to an equational theorem, `[grind =]`, `[grind =_]`, or `[grind _=_]`\ + will mark the theorem for use in heuristic instantiations by the `grind` tactic, + using respectively the left-hand side, the right-hand side, or both sides of the theorem.\ + When applied to a function, `[grind =]` automatically annotates the equational theorems associated with that function.\ + When applied to a theorem `[grind ←]` will instantiate the theorem whenever it encounters the conclusion of the theorem + (that is, it will use the theorem for backwards reasoning).\ + When applied to a theorem `[grind →]` will instantiate the theorem whenever it encounters sufficiently many of the propositional hypotheses + (that is, it will use the theorem for forwards reasoning).\ + \ + The attribute `[grind]` by itself will effectively try `[grind ←]` (if the conclusion is sufficient for instantiation) and then `[grind →]`.\ + \ The `grind` tactic utilizes annotated theorems to add instances of matching patterns into the local context during proof search.\ - For example, if a theorem `@[grind_eq] theorem foo_idempotent : foo (foo x) = foo x` is annotated,\ + For example, if a theorem `@[grind =] theorem foo_idempotent : foo (foo x) = foo x` is annotated,\ `grind` will add an instance of this theorem to the local context whenever it encounters the pattern `foo (foo x)`." applicationTime := .afterCompilation add := fun declName stx attrKind => do diff --git a/src/Lean/Meta/Tactic/Grind/Internalize.lean b/src/Lean/Meta/Tactic/Grind/Internalize.lean index 75b474303380..b0ee4931e2eb 100644 --- a/src/Lean/Meta/Tactic/Grind/Internalize.lean +++ b/src/Lean/Meta/Tactic/Grind/Internalize.lean @@ -47,13 +47,13 @@ private def updateAppMap (e : Expr) : GoalM Unit := do /-- Inserts `e` into the list of case-split candidates. -/ private def addSplitCandidate (e : Expr) : GoalM Unit := do trace[grind.split.candidate] "{e}" - modify fun s => { s with splitCadidates := e :: s.splitCadidates } + modify fun s => { s with splitCandidates := e :: s.splitCandidates } -- TODO: add attribute to make this extensible private def forbiddenSplitTypes := [``Eq, ``HEq, ``True, ``False] /-- Inserts `e` into the list of case-split candidates if applicable. -/ -private def checkAndaddSplitCandidate (e : Expr) : GoalM Unit := do +private def checkAndAddSplitCandidate (e : Expr) : GoalM Unit := do unless e.isApp do return () if e.isIte || e.isDIte then addSplitCandidate e @@ -148,7 +148,7 @@ partial def internalize (e : Expr) (generation : Nat) : GoalM Unit := do -- We do not want to internalize the components of a literal value. mkENode e generation else e.withApp fun f args => do - checkAndaddSplitCandidate e + checkAndAddSplitCandidate e addMatchEqns f generation if f.isConstOf ``Lean.Grind.nestedProof && args.size == 2 then -- We only internalize the proposition. We can skip the proof because of diff --git a/src/Lean/Meta/Tactic/Grind/Split.lean b/src/Lean/Meta/Tactic/Grind/Split.lean index 7c98474232be..df1ff3f2852e 100644 --- a/src/Lean/Meta/Tactic/Grind/Split.lean +++ b/src/Lean/Meta/Tactic/Grind/Split.lean @@ -52,23 +52,23 @@ private def checkCaseSplitStatus (e : Expr) : GoalM CaseSplitStatus := do if (← isResolvedCaseSplit e) then return .resolved if (← isMatcherApp e) then - return .notReady -- TODO: implement splittes for `match` + return .notReady -- TODO: implement splitters for `match` -- return .ready let .const declName .. := e.getAppFn | unreachable! if (← isInductivePredicate declName <&&> isEqTrue e) then return .ready return .notReady -/-- Returns the next case-split to be peformed. It uses a very simple heuristic. -/ +/-- Returns the next case-split to be performed. It uses a very simple heuristic. -/ private def selectNextSplit? : GoalM (Option Expr) := do if (← isInconsistent) then return none if (← checkMaxCaseSplit) then return none - go (← get).splitCadidates none [] + go (← get).splitCandidates none [] where go (cs : List Expr) (c? : Option Expr) (cs' : List Expr) : GoalM (Option Expr) := do match cs with | [] => - modify fun s => { s with splitCadidates := cs'.reverse } + modify fun s => { s with splitCandidates := cs'.reverse } if c?.isSome then -- Remark: we reset `numEmatch` after each case split. -- We should consider other strategies in the future. diff --git a/src/Lean/Meta/Tactic/Grind/Types.lean b/src/Lean/Meta/Tactic/Grind/Types.lean index 4544196da57f..e76d726275d6 100644 --- a/src/Lean/Meta/Tactic/Grind/Types.lean +++ b/src/Lean/Meta/Tactic/Grind/Types.lean @@ -29,7 +29,7 @@ def congrPlaceholderProof := mkConst (Name.mkSimple "[congruence]") /-- Returns `true` if `e` is `True`, `False`, or a literal value. -See `LitValues` for supported literals. +See `Lean.Meta.LitValues` for supported literals. -/ def isInterpreted (e : Expr) : MetaM Bool := do if e.isTrue || e.isFalse then return true @@ -59,11 +59,11 @@ structure CongrTheoremCacheKey where f : Expr numArgs : Nat --- We manually define `BEq` because we wannt to use pointer equality. +-- We manually define `BEq` because we want to use pointer equality. instance : BEq CongrTheoremCacheKey where beq a b := isSameExpr a.f b.f && a.numArgs == b.numArgs --- We manually define `Hashable` because we wannt to use pointer equality. +-- We manually define `Hashable` because we want to use pointer equality. instance : Hashable CongrTheoremCacheKey where hash a := mixHash (unsafe ptrAddrUnsafe a.f).toUInt64 (hash a.numArgs) @@ -123,7 +123,7 @@ def abstractNestedProofs (e : Expr) : GrindM Expr := do /-- Applies hash-consing to `e`. Recall that all expressions in a `grind` goal have -been hash-consing. We perform this step before we internalize expressions. +been hash-consed. We perform this step before we internalize expressions. -/ def shareCommon (e : Expr) : GrindM Expr := do modifyGet fun { canon, scState, nextThmIdx, congrThms, trueExpr, falseExpr, simpStats } => @@ -193,7 +193,7 @@ structure ENode where interpreted : Bool := false /-- `ctor := true` if the head symbol is a constructor application. -/ ctor : Bool := false - /-- `hasLambdas := true` if equivalence class contains lambda expressions. -/ + /-- `hasLambdas := true` if the equivalence class contains lambda expressions. -/ hasLambdas : Bool := false /-- If `heqProofs := true`, then some proofs in the equivalence class are based @@ -383,7 +383,7 @@ structure Goal where /-- `match` auxiliary functions whose equations have already been created and activated. -/ matchEqNames : PHashSet Name := {} /-- Case-split candidates. -/ - splitCadidates : List Expr := [] + splitCandidates : List Expr := [] /-- Number of splits performed to get to this goal. -/ numSplits : Nat := 0 /-- Case-splits that do not have to be performed anymore. -/ From 0da5be1ba18c4d5ff7dd471903d9ee9d59b1b622 Mon Sep 17 00:00:00 2001 From: Leonardo de Moura Date: Tue, 7 Jan 2025 11:35:31 -0800 Subject: [PATCH 043/100] feat: add support for erasing the `[grind]` attribute (#6566) This PR adds support for erasing the `[grind]` attribute used to mark theorems for heuristic instantiation in the `grind` tactic. --- src/Lean/Meta/Tactic/Grind/EMatchTheorem.lean | 75 +++++++++++++---- src/Lean/Meta/Tactic/Grind/Internalize.lean | 15 ++-- tests/lean/run/grind_erase_attr.lean | 81 +++++++++++++++++++ 3 files changed, 150 insertions(+), 21 deletions(-) create mode 100644 tests/lean/run/grind_erase_attr.lean diff --git a/src/Lean/Meta/Tactic/Grind/EMatchTheorem.lean b/src/Lean/Meta/Tactic/Grind/EMatchTheorem.lean index c8299413f00f..96e6b0690186 100644 --- a/src/Lean/Meta/Tactic/Grind/EMatchTheorem.lean +++ b/src/Lean/Meta/Tactic/Grind/EMatchTheorem.lean @@ -57,22 +57,28 @@ inductive Origin where is the provided grind argument. The `id` is a unique identifier for the call. -/ | stx (id : Name) (ref : Syntax) - | other - deriving Inhabited, Repr + | other (id : Name) + deriving Inhabited, Repr, BEq /-- A unique identifier corresponding to the origin. -/ def Origin.key : Origin → Name | .decl declName => declName | .fvar fvarId => fvarId.name | .stx id _ => id - | .other => `other + | .other id => id def Origin.pp [Monad m] [MonadEnv m] [MonadError m] (o : Origin) : m MessageData := do match o with | .decl declName => return MessageData.ofConst (← mkConstWithLevelParams declName) | .fvar fvarId => return mkFVar fvarId | .stx _ ref => return ref - | .other => return "[unknown]" + | .other id => return id + +instance : BEq Origin where + beq a b := a.key == b.key + +instance : Hashable Origin where + hash a := hash a.key /-- A theorem for heuristic instantiation based on E-matching. -/ structure EMatchTheorem where @@ -94,8 +100,10 @@ structure EMatchTheorem where structure EMatchTheorems where /-- The key is a symbol from `EMatchTheorem.symbols`. -/ private map : PHashMap Name (List EMatchTheorem) := {} - /-- Set of theorem names that have been inserted using `insert`. -/ - private thmNames : PHashSet Name := {} + /-- Set of theorem ids that have been inserted using `insert`. -/ + private origins : PHashSet Origin := {} + /-- Theorems that have been marked as erased -/ + private erased : PHashSet Origin := {} deriving Inhabited /-- @@ -110,12 +118,25 @@ def EMatchTheorems.insert (s : EMatchTheorems) (thm : EMatchTheorem) : EMatchThe let .const declName :: syms := thm.symbols | unreachable! let thm := { thm with symbols := syms } - let { map, thmNames } := s - let thmNames := thmNames.insert thm.origin.key + let { map, origins, erased } := s + let origins := origins.insert thm.origin + let erased := erased.erase thm.origin if let some thms := map.find? declName then - return { map := map.insert declName (thm::thms), thmNames } + return { map := map.insert declName (thm::thms), origins, erased } else - return { map := map.insert declName [thm], thmNames } + return { map := map.insert declName [thm], origins, erased } + +/-- Returns `true` if `s` contains a theorem with the given origin. -/ +def EMatchTheorems.contains (s : EMatchTheorems) (origin : Origin) : Bool := + s.origins.contains origin + +/-- Mark the theorm with the given origin as `erased` -/ +def EMatchTheorems.erase (s : EMatchTheorems) (origin : Origin) : EMatchTheorems := + { s with erased := s.erased.insert origin, origins := s.origins.erase origin } + +/-- Returns true if the theorem has been marked as erased. -/ +def EMatchTheorems.isErased (s : EMatchTheorems) (origin : Origin) : Bool := + s.erased.contains origin /-- Retrieves theorems from `s` associated with the given symbol. See `EMatchTheorem.insert`. @@ -128,10 +149,6 @@ def EMatchTheorems.retrieve? (s : EMatchTheorems) (sym : Name) : Option (List EM else none -/-- Returns `true` if `declName` is the name of a theorem that was inserted using `insert`. -/ -def EMatchTheorems.containsTheoremName (s : EMatchTheorems) (declName : Name) : Bool := - s.thmNames.contains declName - def EMatchTheorem.getProofWithFreshMVarLevels (thm : EMatchTheorem) : MetaM Expr := do if thm.proof.isConst && thm.levelParams.isEmpty then let declName := thm.proof.constName! @@ -672,6 +689,36 @@ builtin_initialize applicationTime := .afterCompilation add := fun declName stx attrKind => do addGrindAttr declName attrKind (getKind stx) |>.run' {} + erase := fun declName => MetaM.run' do + /- + Remark: consider the following example + ``` + attribute [grind] foo -- ok + attribute [-grind] foo.eqn_2 -- ok + attribute [-grind] foo -- error + ``` + One may argue that the correct behavior should be + ``` + attribute [grind] foo -- ok + attribute [-grind] foo.eqn_2 -- error + attribute [-grind] foo -- ok + ``` + -/ + let throwErr := throwError "`{declName}` is not marked with the `[grind]` attribute" + let info ← getConstInfo declName + if !info.isTheorem then + if let some eqns ← getEqnsFor? declName then + let s := ematchTheoremsExt.getState (← getEnv) + unless eqns.all fun eqn => s.contains (.decl eqn) do + throwErr + modifyEnv fun env => ematchTheoremsExt.modifyState env fun s => + eqns.foldl (init := s) fun s eqn => s.erase (.decl eqn) + else + throwErr + else + unless ematchTheoremsExt.getState (← getEnv) |>.contains (.decl declName) do + throwErr + modifyEnv fun env => ematchTheoremsExt.modifyState env fun s => s.erase (.decl declName) } end Lean.Meta.Grind diff --git a/src/Lean/Meta/Tactic/Grind/Internalize.lean b/src/Lean/Meta/Tactic/Grind/Internalize.lean index b0ee4931e2eb..86b582423526 100644 --- a/src/Lean/Meta/Tactic/Grind/Internalize.lean +++ b/src/Lean/Meta/Tactic/Grind/Internalize.lean @@ -111,13 +111,14 @@ private partial def activateTheoremPatterns (fName : Name) (generation : Nat) : modify fun s => { s with thmMap } let appMap := (← get).appMap for thm in thms do - let symbols := thm.symbols.filter fun sym => !appMap.contains sym - let thm := { thm with symbols } - match symbols with - | [] => activateTheorem thm generation - | _ => - trace[grind.ematch] "reinsert `{thm.origin.key}`" - modify fun s => { s with thmMap := s.thmMap.insert thm } + unless (← get).thmMap.isErased thm.origin do + let symbols := thm.symbols.filter fun sym => !appMap.contains sym + let thm := { thm with symbols } + match symbols with + | [] => activateTheorem thm generation + | _ => + trace[grind.ematch] "reinsert `{thm.origin.key}`" + modify fun s => { s with thmMap := s.thmMap.insert thm } partial def internalize (e : Expr) (generation : Nat) : GoalM Unit := do if (← alreadyInternalized e) then return () diff --git a/tests/lean/run/grind_erase_attr.lean b/tests/lean/run/grind_erase_attr.lean new file mode 100644 index 000000000000..427242d796cc --- /dev/null +++ b/tests/lean/run/grind_erase_attr.lean @@ -0,0 +1,81 @@ +opaque f : Nat → Nat + +@[grind] theorem fthm : f (f x) = f x := sorry + +theorem fthm' : f (f x) = x := sorry + +/-- +error: `fthm'` is not marked with the `[grind]` attribute +-/ +#guard_msgs in +attribute [-grind] fthm' + +set_option trace.grind.assert true + +/-- +info: [grind.assert] ¬f (f (f a)) = f a +[grind.assert] f (f (f a)) = f (f a) +[grind.assert] f (f a) = f a +-/ +#guard_msgs (info) in +example : f (f (f a)) = f a := by + grind + +attribute [-grind] fthm + +/-- +error: `grind` failed +a : Nat +x✝ : ¬f (f (f a)) = f a +⊢ False +--- +info: [grind.assert] ¬f (f (f a)) = f a +-/ +#guard_msgs (info, error) in +example : f (f (f a)) = f a := by + grind + +/-- +error: `fthm` is not marked with the `[grind]` attribute +-/ +#guard_msgs in +attribute [-grind] fthm + +attribute [grind] fthm + +example : f (f (f a)) = f a := by + grind + +def g (x : Nat) := + match x with + | 0 => 1 + | x+1 => 2 * g x + +attribute [grind] g + +example : g a = b → a = 0 → b = 1 := by + grind + +attribute [-grind] g + +/-- +error: `grind` failed +a b : Nat +a✝¹ : g a = b +a✝ : a = 0 +x✝ : ¬b = 1 +⊢ False +--- +info: [grind.assert] g a = b +[grind.assert] a = 0 +[grind.assert] ¬b = 1 +-/ +#guard_msgs (info, error) in +example : g a = b → a = 0 → b = 1 := by + grind + +/-- +error: `g` is not marked with the `[grind]` attribute +-/ +#guard_msgs in +attribute [-grind] g From 5decd2ce201e47e43f2ff4b5f78f27370c265df2 Mon Sep 17 00:00:00 2001 From: Leonardo de Moura Date: Tue, 7 Jan 2025 15:27:36 -0800 Subject: [PATCH 044/100] feat: trace messages for working and closing goals in the `grind` tactic (#6567) This PR adds support for erasing the `[grind]` attribute used to mark theorems for heuristic instantiation in the `grind` tactic. --- src/Lean/Meta/Tactic/Grind/Cases.lean | 5 ++- src/Lean/Meta/Tactic/Grind/Core.lean | 16 +++++----- src/Lean/Meta/Tactic/Grind/Ctor.lean | 2 +- src/Lean/Meta/Tactic/Grind/EMatch.lean | 12 ++++---- src/Lean/Meta/Tactic/Grind/ForallProp.lean | 6 ++-- src/Lean/Meta/Tactic/Grind/Internalize.lean | 14 ++++----- src/Lean/Meta/Tactic/Grind/Inv.lean | 4 +-- src/Lean/Meta/Tactic/Grind/Main.lean | 1 + src/Lean/Meta/Tactic/Grind/Proj.lean | 2 +- src/Lean/Meta/Tactic/Grind/Split.lean | 2 +- src/Lean/Meta/Tactic/Grind/Types.lean | 34 ++++++++++++++++++--- src/Lean/Util/Trace.lean | 7 +++-- tests/lean/run/grind_erase_attr.lean | 2 ++ tests/lean/run/grind_pre.lean | 5 +++ tests/lean/run/grind_split.lean | 22 ++++++++++++- 15 files changed, 97 insertions(+), 37 deletions(-) diff --git a/src/Lean/Meta/Tactic/Grind/Cases.lean b/src/Lean/Meta/Tactic/Grind/Cases.lean index d7e620f4497c..ea81d88ddb13 100644 --- a/src/Lean/Meta/Tactic/Grind/Cases.lean +++ b/src/Lean/Meta/Tactic/Grind/Cases.lean @@ -36,14 +36,17 @@ def cases (mvarId : MVarId) (e : Expr) : MetaM (List MVarId) := mvarId.withConte let mut recursor := mkApp (mkAppN recursor indicesExpr) (mkFVar fvarId) let mut recursorType ← inferType recursor let mut mvarIdsNew := #[] + let mut idx := 1 for _ in [:recursorInfo.numMinors] do let .forallE _ targetNew recursorTypeNew _ ← whnf recursorType | throwTacticEx `grind.cases mvarId "unexpected recursor type" recursorType := recursorTypeNew - let mvar ← mkFreshExprSyntheticOpaqueMVar targetNew tag + let tagNew := if recursorInfo.numMinors > 1 then Name.num tag idx else tag + let mvar ← mkFreshExprSyntheticOpaqueMVar targetNew tagNew recursor := mkApp recursor mvar let mvarIdNew ← mvar.mvarId!.tryClearMany (indices.push fvarId) mvarIdsNew := mvarIdsNew.push mvarIdNew + idx := idx + 1 mvarId.assign recursor return mvarIdsNew.toList if recursorInfo.numIndices > 0 then diff --git a/src/Lean/Meta/Tactic/Grind/Core.lean b/src/Lean/Meta/Tactic/Grind/Core.lean index 635e7aef85e9..4aab792f849d 100644 --- a/src/Lean/Meta/Tactic/Grind/Core.lean +++ b/src/Lean/Meta/Tactic/Grind/Core.lean @@ -43,7 +43,7 @@ private def removeParents (root : Expr) : GoalM ParentSet := do for parent in parents do -- Recall that we may have `Expr.forallE` in `parents` because of `ForallProp.lean` if (← pure parent.isApp <&&> isCongrRoot parent) then - trace[grind.debug.parent] "remove: {parent}" + trace_goal[grind.debug.parent] "remove: {parent}" modify fun s => { s with congrTable := s.congrTable.erase { e := parent } } return parents @@ -54,7 +54,7 @@ This is an auxiliary function performed while merging equivalence classes. private def reinsertParents (parents : ParentSet) : GoalM Unit := do for parent in parents do if (← pure parent.isApp <&&> isCongrRoot parent) then - trace[grind.debug.parent] "reinsert: {parent}" + trace_goal[grind.debug.parent] "reinsert: {parent}" addCongrTable parent /-- Closes the goal when `True` and `False` are in the same equivalence class. -/ @@ -91,9 +91,9 @@ private partial def addEqStep (lhs rhs proof : Expr) (isHEq : Bool) : GoalM Unit let rhsNode ← getENode rhs if isSameExpr lhsNode.root rhsNode.root then -- `lhs` and `rhs` are already in the same equivalence class. - trace[grind.debug] "{← ppENodeRef lhs} and {← ppENodeRef rhs} are already in the same equivalence class" + trace_goal[grind.debug] "{← ppENodeRef lhs} and {← ppENodeRef rhs} are already in the same equivalence class" return () - trace[grind.eqc] "{← if isHEq then mkHEq lhs rhs else mkEq lhs rhs}" + trace_goal[grind.eqc] "{← if isHEq then mkHEq lhs rhs else mkEq lhs rhs}" let lhsRoot ← getENode lhsNode.root let rhsRoot ← getENode rhsNode.root let mut valueInconsistency := false @@ -118,11 +118,11 @@ private partial def addEqStep (lhs rhs proof : Expr) (isHEq : Bool) : GoalM Unit unless (← isInconsistent) do if valueInconsistency then closeGoalWithValuesEq lhsRoot.self rhsRoot.self - trace[grind.debug] "after addEqStep, {← ppState}" + trace_goal[grind.debug] "after addEqStep, {← ppState}" checkInvariants where go (lhs rhs : Expr) (lhsNode rhsNode lhsRoot rhsRoot : ENode) (flipped : Bool) : GoalM Unit := do - trace[grind.debug] "adding {← ppENodeRef lhs} ↦ {← ppENodeRef rhs}" + trace_goal[grind.debug] "adding {← ppENodeRef lhs} ↦ {← ppENodeRef rhs}" /- We have the following `target?/proof?` `lhs -> ... -> lhsNode.root` @@ -139,7 +139,7 @@ where } let parents ← removeParents lhsRoot.self updateRoots lhs rhsNode.root - trace[grind.debug] "{← ppENodeRef lhs} new root {← ppENodeRef rhsNode.root}, {← ppENodeRef (← getRoot lhs)}" + trace_goal[grind.debug] "{← ppENodeRef lhs} new root {← ppENodeRef rhsNode.root}, {← ppENodeRef (← getRoot lhs)}" reinsertParents parents setENode lhsNode.root { (← getENode lhsRoot.self) with -- We must retrieve `lhsRoot` since it was updated. next := rhsRoot.next @@ -208,7 +208,7 @@ def addNewEq (lhs rhs proof : Expr) (generation : Nat) : GoalM Unit := do /-- Adds a new `fact` justified by the given proof and using the given generation. -/ def add (fact : Expr) (proof : Expr) (generation := 0) : GoalM Unit := do - trace[grind.assert] "{fact}" + trace_goal[grind.assert] "{fact}" if (← isInconsistent) then return () resetNewEqs let_expr Not p := fact diff --git a/src/Lean/Meta/Tactic/Grind/Ctor.lean b/src/Lean/Meta/Tactic/Grind/Ctor.lean index d7971d751ba6..754fe27ad996 100644 --- a/src/Lean/Meta/Tactic/Grind/Ctor.lean +++ b/src/Lean/Meta/Tactic/Grind/Ctor.lean @@ -20,7 +20,7 @@ private partial def propagateInjEqs (eqs : Expr) (proof : Expr) : GoalM Unit := | HEq _ lhs _ rhs => pushHEq (← shareCommon lhs) (← shareCommon rhs) proof | _ => - trace[grind.issues] "unexpected injectivity theorem result type{indentExpr eqs}" + trace_goal[grind.issues] "unexpected injectivity theorem result type{indentExpr eqs}" return () /-- diff --git a/src/Lean/Meta/Tactic/Grind/EMatch.lean b/src/Lean/Meta/Tactic/Grind/EMatch.lean index 4ba4d6eff980..67411cb8ed84 100644 --- a/src/Lean/Meta/Tactic/Grind/EMatch.lean +++ b/src/Lean/Meta/Tactic/Grind/EMatch.lean @@ -221,7 +221,7 @@ private def addNewInstance (origin : Origin) (proof : Expr) (generation : Nat) : -- `initApp` is a match-application that we don't need to split at anymore. markCaseSplitAsResolved (← read).initApp prop ← annotateMatchEqnType prop - trace[grind.ematch.instance] "{← origin.pp}: {prop}" + trace_goal[grind.ematch.instance] "{← origin.pp}: {prop}" addTheoremInstance proof prop (generation+1) /-- @@ -232,13 +232,13 @@ private partial def instantiateTheorem (c : Choice) : M Unit := withDefault do w let thm := (← read).thm unless (← markTheoremInstance thm.proof c.assignment) do return () - trace[grind.ematch.instance.assignment] "{← thm.origin.pp}: {assignmentToMessageData c.assignment}" + trace_goal[grind.ematch.instance.assignment] "{← thm.origin.pp}: {assignmentToMessageData c.assignment}" let proof ← thm.getProofWithFreshMVarLevels let numParams := thm.numParams assert! c.assignment.size == numParams let (mvars, bis, _) ← forallMetaBoundedTelescope (← inferType proof) numParams if mvars.size != thm.numParams then - trace[grind.issues] "unexpected number of parameters at {← thm.origin.pp}" + trace_goal[grind.issues] "unexpected number of parameters at {← thm.origin.pp}" return () -- Apply assignment for h : i in [:mvars.size] do @@ -246,14 +246,14 @@ private partial def instantiateTheorem (c : Choice) : M Unit := withDefault do w unless isSameExpr v unassigned do let mvarId := mvars[i].mvarId! unless (← isDefEq (← mvarId.getType) (← inferType v) <&&> mvarId.checkedAssign v) do - trace[grind.issues] "type error constructing proof for {← thm.origin.pp}\nwhen assigning metavariable {mvars[i]} with {indentExpr v}" + trace_goal[grind.issues] "type error constructing proof for {← thm.origin.pp}\nwhen assigning metavariable {mvars[i]} with {indentExpr v}" return () -- Synthesize instances for mvar in mvars, bi in bis do if bi.isInstImplicit && !(← mvar.mvarId!.isAssigned) then let type ← inferType mvar unless (← synthesizeInstance mvar type) do - trace[grind.issues] "failed to synthesize instance when instantiating {← thm.origin.pp}{indentExpr type}" + trace_goal[grind.issues] "failed to synthesize instance when instantiating {← thm.origin.pp}{indentExpr type}" return () let proof := mkAppN proof mvars if (← mvars.allM (·.mvarId!.isAssigned)) then @@ -261,7 +261,7 @@ private partial def instantiateTheorem (c : Choice) : M Unit := withDefault do w else let mvars ← mvars.filterM fun mvar => return !(← mvar.mvarId!.isAssigned) if let some mvarBad ← mvars.findM? fun mvar => return !(← isProof mvar) then - trace[grind.issues] "failed to instantiate {← thm.origin.pp}, failed to instantiate non propositional argument with type{indentExpr (← inferType mvarBad)}" + trace_goal[grind.issues] "failed to instantiate {← thm.origin.pp}, failed to instantiate non propositional argument with type{indentExpr (← inferType mvarBad)}" let proof ← mkLambdaFVars (binderInfoForMVars := .default) mvars (← instantiateMVars proof) addNewInstance thm.origin proof c.gen where diff --git a/src/Lean/Meta/Tactic/Grind/ForallProp.lean b/src/Lean/Meta/Tactic/Grind/ForallProp.lean index ee46ee6eb915..e89697fa803e 100644 --- a/src/Lean/Meta/Tactic/Grind/ForallProp.lean +++ b/src/Lean/Meta/Tactic/Grind/ForallProp.lean @@ -17,19 +17,19 @@ If so, internalize the term `proj_i (ctor ... a_i ...)` and add the equality `pr -/ def propagateForallPropUp (e : Expr) : GoalM Unit := do let .forallE n p q bi := e | return () - trace[grind.debug.forallPropagator] "{e}" + trace_goal[grind.debug.forallPropagator] "{e}" if !q.hasLooseBVars then propagateImpliesUp p q else unless (← isEqTrue p) do return - trace[grind.debug.forallPropagator] "isEqTrue, {e}" + trace_goal[grind.debug.forallPropagator] "isEqTrue, {e}" let h₁ ← mkEqTrueProof p let qh₁ := q.instantiate1 (mkApp2 (mkConst ``of_eq_true) p h₁) let r ← simp qh₁ let q := mkLambda n bi p q let q' := r.expr internalize q' (← getGeneration e) - trace[grind.debug.forallPropagator] "q': {q'} for{indentExpr e}" + trace_goal[grind.debug.forallPropagator] "q': {q'} for{indentExpr e}" let h₂ ← r.getProof let h := mkApp5 (mkConst ``Lean.Grind.forall_propagator) p q q' h₁ h₂ pushEq e q' h diff --git a/src/Lean/Meta/Tactic/Grind/Internalize.lean b/src/Lean/Meta/Tactic/Grind/Internalize.lean index 86b582423526..9ce47e2583c6 100644 --- a/src/Lean/Meta/Tactic/Grind/Internalize.lean +++ b/src/Lean/Meta/Tactic/Grind/Internalize.lean @@ -22,9 +22,9 @@ def addCongrTable (e : Expr) : GoalM Unit := do let g := e'.getAppFn unless isSameExpr f g do unless (← hasSameType f g) do - trace[grind.issues] "found congruence between{indentExpr e}\nand{indentExpr e'}\nbut functions have different types" + trace_goal[grind.issues] "found congruence between{indentExpr e}\nand{indentExpr e'}\nbut functions have different types" return () - trace[grind.debug.congr] "{e} = {e'}" + trace_goal[grind.debug.congr] "{e} = {e'}" pushEqHEq e e' congrPlaceholderProof let node ← getENode e setENode e { node with congr := e' } @@ -46,7 +46,7 @@ private def updateAppMap (e : Expr) : GoalM Unit := do /-- Inserts `e` into the list of case-split candidates. -/ private def addSplitCandidate (e : Expr) : GoalM Unit := do - trace[grind.split.candidate] "{e}" + trace_goal[grind.split.candidate] "{e}" modify fun s => { s with splitCandidates := e :: s.splitCandidates } -- TODO: add attribute to make this extensible @@ -89,7 +89,7 @@ private partial def activateTheorem (thm : EMatchTheorem) (generation : Nat) : G -- We don't want to use structural equality when comparing keys. let proof ← shareCommon thm.proof let thm := { thm with proof, patterns := (← thm.patterns.mapM (internalizePattern · generation)) } - trace[grind.ematch] "activated `{thm.origin.key}`, {thm.patterns.map ppPattern}" + trace_goal[grind.ematch] "activated `{thm.origin.key}`, {thm.patterns.map ppPattern}" modify fun s => { s with newThms := s.newThms.push thm } /-- @@ -117,12 +117,12 @@ private partial def activateTheoremPatterns (fName : Name) (generation : Nat) : match symbols with | [] => activateTheorem thm generation | _ => - trace[grind.ematch] "reinsert `{thm.origin.key}`" + trace_goal[grind.ematch] "reinsert `{thm.origin.key}`" modify fun s => { s with thmMap := s.thmMap.insert thm } partial def internalize (e : Expr) (generation : Nat) : GoalM Unit := do if (← alreadyInternalized e) then return () - trace[grind.internalize] "{e}" + trace_goal[grind.internalize] "{e}" match e with | .bvar .. => unreachable! | .sort .. => return () @@ -142,7 +142,7 @@ partial def internalize (e : Expr) (generation : Nat) : GoalM Unit := do | .mvar .. | .mdata .. | .proj .. => - trace[grind.issues] "unexpected term during internalization{indentExpr e}" + trace_goal[grind.issues] "unexpected term during internalization{indentExpr e}" mkENodeCore e (ctor := false) (interpreted := false) (generation := generation) | .app .. => if (← isLitValue e) then diff --git a/src/Lean/Meta/Tactic/Grind/Inv.lean b/src/Lean/Meta/Tactic/Grind/Inv.lean index c09ad807e7d8..0b374a441f3a 100644 --- a/src/Lean/Meta/Tactic/Grind/Inv.lean +++ b/src/Lean/Meta/Tactic/Grind/Inv.lean @@ -85,9 +85,9 @@ private def checkProofs : GoalM Unit := do for b in eqc do unless isSameExpr a b do let p ← mkEqHEqProof a b - trace[grind.debug.proofs] "{a} = {b}" + trace_goal[grind.debug.proofs] "{a} = {b}" check p - trace[grind.debug.proofs] "checked: {← inferType p}" + trace_goal[grind.debug.proofs] "checked: {← inferType p}" /-- Checks basic invariants if `grind.debug` is enabled. diff --git a/src/Lean/Meta/Tactic/Grind/Main.lean b/src/Lean/Meta/Tactic/Grind/Main.lean index 60a8fa0e95c3..715191ed68f2 100644 --- a/src/Lean/Meta/Tactic/Grind/Main.lean +++ b/src/Lean/Meta/Tactic/Grind/Main.lean @@ -60,6 +60,7 @@ private def initCore (mvarId : MVarId) : GrindM (List Goal) := do let mvarId ← mvarId.revertAll let mvarId ← mvarId.unfoldReducible let mvarId ← mvarId.betaReduce + appendTagSuffix mvarId `grind let goals ← intros (← mkGoal mvarId) (generation := 0) goals.forM (·.checkInvariants (expensive := true)) return goals.filter fun goal => !goal.inconsistent diff --git a/src/Lean/Meta/Tactic/Grind/Proj.lean b/src/Lean/Meta/Tactic/Grind/Proj.lean index d9bcb2de28ee..f587b78eace5 100644 --- a/src/Lean/Meta/Tactic/Grind/Proj.lean +++ b/src/Lean/Meta/Tactic/Grind/Proj.lean @@ -30,7 +30,7 @@ def propagateProjEq (parent : Expr) : GoalM Unit := do let parentNew ← shareCommon (mkApp parent.appFn! ctor) internalize parentNew (← getGeneration parent) pure parentNew - trace[grind.debug.proj] "{parentNew}" + trace_goal[grind.debug.proj] "{parentNew}" let idx := info.numParams + info.i unless idx < ctor.getAppNumArgs do return () let v := ctor.getArg! idx diff --git a/src/Lean/Meta/Tactic/Grind/Split.lean b/src/Lean/Meta/Tactic/Grind/Split.lean index df1ff3f2852e..96ae7ce4160a 100644 --- a/src/Lean/Meta/Tactic/Grind/Split.lean +++ b/src/Lean/Meta/Tactic/Grind/Split.lean @@ -110,7 +110,7 @@ def splitNext : GrindTactic := fun goal => do let some c ← selectNextSplit? | return none let gen ← getGeneration c - trace[grind.split] "{c}, generation: {gen}" + trace_goal[grind.split] "{c}, generation: {gen}" -- TODO: `match` let major ← mkCasesMajor c let mvarIds ← cases (← get).mvarId major diff --git a/src/Lean/Meta/Tactic/Grind/Types.lean b/src/Lean/Meta/Tactic/Grind/Types.lean index e76d726275d6..3665bece75a7 100644 --- a/src/Lean/Meta/Tactic/Grind/Types.lean +++ b/src/Lean/Meta/Tactic/Grind/Types.lean @@ -83,6 +83,11 @@ structure State where simpStats : Simp.Stats := {} trueExpr : Expr falseExpr : Expr + /-- + Used to generate trace messages of the for `[grind] working on `, + and implement the macro `trace_goal`. + -/ + lastTag : Name := .anonymous private opaque MethodsRefPointed : NonemptyType.{0} private def MethodsRef : Type := MethodsRefPointed.type @@ -126,9 +131,9 @@ Applies hash-consing to `e`. Recall that all expressions in a `grind` goal have been hash-consed. We perform this step before we internalize expressions. -/ def shareCommon (e : Expr) : GrindM Expr := do - modifyGet fun { canon, scState, nextThmIdx, congrThms, trueExpr, falseExpr, simpStats } => + modifyGet fun { canon, scState, nextThmIdx, congrThms, trueExpr, falseExpr, simpStats, lastTag } => let (e, scState) := ShareCommon.State.shareCommon scState e - (e, { canon, scState, nextThmIdx, congrThms, trueExpr, falseExpr, simpStats }) + (e, { canon, scState, nextThmIdx, congrThms, trueExpr, falseExpr, simpStats, lastTag }) /-- Canonicalizes nested types, type formers, and instances in `e`. @@ -401,6 +406,25 @@ abbrev GoalM := StateRefT Goal GrindM @[inline] def GoalM.run' (goal : Goal) (x : GoalM Unit) : GrindM Goal := goal.mvarId.withContext do StateRefT'.run' (x *> get) goal +def updateLastTag : GoalM Unit := do + if (← isTracingEnabledFor `grind) then + let currTag ← (← get).mvarId.getTag + if currTag != (← getThe Grind.State).lastTag then + trace[grind] "working on goal `{currTag}`" + modifyThe Grind.State fun s => { s with lastTag := currTag } + +/-- +Macro similar to `trace[...]`, but it includes the trace message `trace[grind] "working on "` +if the tag has changed since the last trace message. +-/ +macro "trace_goal[" id:ident "]" s:(interpolatedStr(term) <|> term) : doElem => do + let msg ← if s.raw.getKind == interpolatedStrKind then `(m! $(⟨s⟩)) else `(($(⟨s⟩) : MessageData)) + `(doElem| do + let cls := $(quote id.getId.eraseMacroScopes) + if (← Lean.isTracingEnabledFor cls) then + updateLastTag + Lean.addTrace cls $msg) + /-- A helper function used to mark a theorem instance found by the E-matching module. It returns `true` if it is a new instance and `false` otherwise. @@ -652,7 +676,9 @@ def mkEqFalseProof (a : Expr) : GoalM Expr := do /-- Marks current goal as inconsistent without assigning `mvarId`. -/ def markAsInconsistent : GoalM Unit := do - modify fun s => { s with inconsistent := true } + unless (← get).inconsistent do + trace[grind] "closed `{← (← get).mvarId.getTag}`" + modify fun s => { s with inconsistent := true } /-- Closes the current goal using the given proof of `False` and @@ -751,7 +777,7 @@ Remark: we currently use this feature to disable `match`-case-splits -/ def markCaseSplitAsResolved (e : Expr) : GoalM Unit := do unless (← isResolvedCaseSplit e) do - trace[grind.split.resolved] "{e}" + trace_goal[grind.split.resolved] "{e}" modify fun s => { s with resolvedSplits := s.resolvedSplits.insert { expr := e } } end Lean.Meta.Grind diff --git a/src/Lean/Util/Trace.lean b/src/Lean/Util/Trace.lean index 66eaf9decc1b..22db50dd6a99 100644 --- a/src/Lean/Util/Trace.lean +++ b/src/Lean/Util/Trace.lean @@ -293,13 +293,16 @@ def registerTraceClass (traceClassName : Name) (inherited := false) (ref : Name if inherited then inheritedTraceOptions.modify (·.insert optionName) -macro "trace[" id:ident "]" s:(interpolatedStr(term) <|> term) : doElem => do - let msg ← if s.raw.getKind == interpolatedStrKind then `(m! $(⟨s⟩)) else `(($(⟨s⟩) : MessageData)) +def expandTraceMacro (id : Syntax) (s : Syntax) : MacroM (TSyntax `doElem) := do + let msg ← if s.getKind == interpolatedStrKind then `(m! $(⟨s⟩)) else `(($(⟨s⟩) : MessageData)) `(doElem| do let cls := $(quote id.getId.eraseMacroScopes) if (← Lean.isTracingEnabledFor cls) then Lean.addTrace cls $msg) +macro "trace[" id:ident "]" s:(interpolatedStr(term) <|> term) : doElem => do + expandTraceMacro id s.raw + def bombEmoji := "💥️" def checkEmoji := "✅️" def crossEmoji := "❌️" diff --git a/tests/lean/run/grind_erase_attr.lean b/tests/lean/run/grind_erase_attr.lean index 427242d796cc..467b18d34bfd 100644 --- a/tests/lean/run/grind_erase_attr.lean +++ b/tests/lean/run/grind_erase_attr.lean @@ -25,6 +25,7 @@ attribute [-grind] fthm /-- error: `grind` failed +case grind a : Nat x✝ : ¬f (f (f a)) = f a ⊢ False @@ -60,6 +61,7 @@ attribute [-grind] g /-- error: `grind` failed +case grind a b : Nat a✝¹ : g a = b a✝ : a = 0 diff --git a/tests/lean/run/grind_pre.lean b/tests/lean/run/grind_pre.lean index be879cd20222..1a2bbb378c0a 100644 --- a/tests/lean/run/grind_pre.lean +++ b/tests/lean/run/grind_pre.lean @@ -5,6 +5,7 @@ set_option grind.debug.proofs true /-- error: `grind` failed +case grind.1.2 a b c : Bool p q : Prop left✝ : a = true @@ -23,6 +24,7 @@ theorem ex (h : (f a && (b || f (f c))) = true) (h' : p ∧ q) : b && a := by open Lean.Grind.Eager in /-- error: `grind` failed +case grind.2.1 a b c : Bool p q : Prop left✝ : a = true @@ -40,6 +42,7 @@ def g (i : Nat) (j : Nat) (_ : i > j := by omega) := i + j /-- error: `grind` failed +case grind i j : Nat h : j + 1 < i + 1 h✝ : j + 1 ≤ i @@ -56,6 +59,7 @@ structure Point where /-- error: `grind` failed +case grind a₁ : Point a₂ : Nat a₃ : Int @@ -82,6 +86,7 @@ example (p : Prop) (a b c : Nat) : p → a = 0 → a = b → h a = h c → a = c set_option trace.grind.debug.proof true /-- error: `grind` failed +case grind α : Type a : α p q r : Prop diff --git a/tests/lean/run/grind_split.lean b/tests/lean/run/grind_split.lean index 917e62e6b7f6..7acf78263d59 100644 --- a/tests/lean/run/grind_split.lean +++ b/tests/lean/run/grind_split.lean @@ -5,10 +5,29 @@ example (p q : Prop) : p ∨ q → p ∨ ¬q → ¬p ∨ q → ¬p ∨ ¬q → F opaque R : Nat → Prop +/-- +info: [grind] working on goal `grind` +[grind.eqc] (if p then a else b) = c +[grind.eqc] R a = True +[grind.eqc] R b = True +[grind.eqc] R c = False +[grind.split] if p then a else b, generation: 0 +[grind] working on goal `grind.1` +[grind.eqc] p = True +[grind.eqc] (if p then a else b) = a +[grind.eqc] R a = R c +[grind] closed `grind.1` +[grind] working on goal `grind.2` +[grind.eqc] p = False +[grind.eqc] (if p then a else b) = b +[grind.eqc] R b = R c +[grind] closed `grind.2` +-/ +#guard_msgs (info) in +set_option trace.grind true in example (p : Prop) [Decidable p] (a b c : Nat) : (if p then a else b) = c → R a → R b → R c := by grind - namespace grind_test_induct_pred inductive Expr where @@ -29,6 +48,7 @@ inductive HasType : Expr → Ty → Prop | bool : HasType (.bool v) .bool | and : HasType a .bool → HasType b .bool → HasType (.and a b) .bool +set_option trace.grind true theorem HasType.det (h₁ : HasType e t₁) (h₂ : HasType e t₂) : t₁ = t₂ := by grind From 22a799524f07496407703f66af7e3f4a60d5980c Mon Sep 17 00:00:00 2001 From: Leonardo de Moura Date: Tue, 7 Jan 2025 16:21:13 -0800 Subject: [PATCH 045/100] feat: add support for `cast`, `Eq.rec`, `Eq.ndrec` to `grind` (#6568) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR adds basic support for cast-like operators to the grind tactic. Example: ```lean example (α : Type) (β : Type) (a₁ a₂ : α) (b₁ b₂ : β) (h₁ : α = β) (h₂ : h₁ ▸ a₁ = b₁) (h₃ : a₁ = a₂) (h₄ : b₁ = b₂) : HEq a₂ b₂ := by grind ``` --- src/Init/Grind/Lemmas.lean | 17 +++++++++ src/Lean/Meta/Tactic/Grind/EMatch.lean | 6 ++-- src/Lean/Meta/Tactic/Grind/Internalize.lean | 15 ++++++++ tests/lean/run/grind_t1.lean | 40 +++++++++++++++++++++ 4 files changed, 76 insertions(+), 2 deletions(-) diff --git a/src/Init/Grind/Lemmas.lean b/src/Init/Grind/Lemmas.lean index b19515b7e6c9..948033cc6bb8 100644 --- a/src/Init/Grind/Lemmas.lean +++ b/src/Init/Grind/Lemmas.lean @@ -78,4 +78,21 @@ theorem forall_propagator (p : Prop) (q : p → Prop) (q' : Prop) (h₁ : p = Tr theorem dite_cond_eq_true' {α : Sort u} {c : Prop} {_ : Decidable c} {a : c → α} {b : ¬ c → α} {r : α} (h₁ : c = True) (h₂ : a (of_eq_true h₁) = r) : (dite c a b) = r := by simp [h₁, h₂] theorem dite_cond_eq_false' {α : Sort u} {c : Prop} {_ : Decidable c} {a : c → α} {b : ¬ c → α} {r : α} (h₁ : c = False) (h₂ : b (of_eq_false h₁) = r) : (dite c a b) = r := by simp [h₁, h₂] +/-! Casts -/ + +theorem eqRec_heq.{u_1, u_2} {α : Sort u_2} {a : α} + {motive : (x : α) → a = x → Sort u_1} (v : motive a (Eq.refl a)) {b : α} (h : a = b) + : HEq (@Eq.rec α a motive v b h) v := by + subst h; rfl + +theorem eqRecOn_heq.{u_1, u_2} {α : Sort u_2} {a : α} + {motive : (x : α) → a = x → Sort u_1} {b : α} (h : a = b) (v : motive a (Eq.refl a)) + : HEq (@Eq.recOn α a motive b h v) v := by + subst h; rfl + +theorem eqNDRec_heq.{u_1, u_2} {α : Sort u_2} {a : α} + {motive : α → Sort u_1} (v : motive a) {b : α} (h : a = b) + : HEq (@Eq.ndrec α a motive v b h) v := by + subst h; rfl + end Lean.Grind diff --git a/src/Lean/Meta/Tactic/Grind/EMatch.lean b/src/Lean/Meta/Tactic/Grind/EMatch.lean index 67411cb8ed84..bce265b73818 100644 --- a/src/Lean/Meta/Tactic/Grind/EMatch.lean +++ b/src/Lean/Meta/Tactic/Grind/EMatch.lean @@ -245,8 +245,10 @@ private partial def instantiateTheorem (c : Choice) : M Unit := withDefault do w let v := c.assignment[numParams - i - 1]! unless isSameExpr v unassigned do let mvarId := mvars[i].mvarId! - unless (← isDefEq (← mvarId.getType) (← inferType v) <&&> mvarId.checkedAssign v) do - trace_goal[grind.issues] "type error constructing proof for {← thm.origin.pp}\nwhen assigning metavariable {mvars[i]} with {indentExpr v}" + let mvarIdType ← mvarId.getType + let vType ← inferType v + unless (← isDefEq mvarIdType vType <&&> mvarId.checkedAssign v) do + trace_goal[grind.issues] "type error constructing proof for {← thm.origin.pp}\nwhen assigning metavariable {mvars[i]} with {indentExpr v}\n{← mkHasTypeButIsExpectedMsg vType mvarIdType}" return () -- Synthesize instances for mvar in mvars, bi in bis do diff --git a/src/Lean/Meta/Tactic/Grind/Internalize.lean b/src/Lean/Meta/Tactic/Grind/Internalize.lean index 9ce47e2583c6..e6f76fe7d9db 100644 --- a/src/Lean/Meta/Tactic/Grind/Internalize.lean +++ b/src/Lean/Meta/Tactic/Grind/Internalize.lean @@ -5,6 +5,7 @@ Authors: Leonardo de Moura -/ prelude import Init.Grind.Util +import Init.Grind.Lemmas import Lean.Meta.LitValues import Lean.Meta.Match.MatcherInfo import Lean.Meta.Match.MatchEqsExt @@ -72,6 +73,19 @@ private def checkAndAddSplitCandidate (e : Expr) : GoalM Unit := do if (← isInductivePredicate declName) then addSplitCandidate e +/-- +If `e` is a `cast`-like term (e.g., `cast h a`), add `HEq e a` to the to-do list. +It could be an E-matching theorem, but we want to ensure it is always applied since +we want to rely on the fact that `cast h a` and `a` are in the same equivalence class. +-/ +private def pushCastHEqs (e : Expr) : GoalM Unit := do + match_expr e with + | f@cast α β h a => pushHEq e a (mkApp4 (mkConst ``cast_heq f.constLevels!) α β h a) + | f@Eq.rec α a motive v b h => pushHEq e v (mkApp6 (mkConst ``Grind.eqRec_heq f.constLevels!) α a motive v b h) + | f@Eq.ndrec α a motive v b h => pushHEq e v (mkApp6 (mkConst ``Grind.eqNDRec_heq f.constLevels!) α a motive v b h) + | f@Eq.recOn α a motive b h v => pushHEq e v (mkApp6 (mkConst ``Grind.eqRecOn_heq f.constLevels!) α a motive b h v) + | _ => return () + mutual /-- Internalizes the nested ground terms in the given pattern. -/ private partial def internalizePattern (pattern : Expr) (generation : Nat) : GoalM Expr := do @@ -150,6 +164,7 @@ partial def internalize (e : Expr) (generation : Nat) : GoalM Unit := do mkENode e generation else e.withApp fun f args => do checkAndAddSplitCandidate e + pushCastHEqs e addMatchEqns f generation if f.isConstOf ``Lean.Grind.nestedProof && args.size == 2 then -- We only internalize the proposition. We can skip the proof because of diff --git a/tests/lean/run/grind_t1.lean b/tests/lean/run/grind_t1.lean index 7686d4c0d832..9c75284afdf5 100644 --- a/tests/lean/run/grind_t1.lean +++ b/tests/lean/run/grind_t1.lean @@ -133,3 +133,43 @@ info: [grind.eqc] x = 2 * a set_option trace.grind.eqc true in example (a : Nat) : let_fun x := a + a; y = x → y = a + a := by grind + +example (α : Type) (β : Type) (a₁ a₂ : α) (b₁ b₂ : β) + (h₁ : α = β) + (h₂ : cast h₁ a₁ = b₁) + (h₃ : a₁ = a₂) + (h₄ : b₁ = b₂) + : HEq a₂ b₂ := by + grind + +example (α : Type) (β : Type) (a₁ a₂ : α) (b₁ b₂ : β) + (h₁ : α = β) + (h₂ : h₁ ▸ a₁ = b₁) + (h₃ : a₁ = a₂) + (h₄ : b₁ = b₂) + : HEq a₂ b₂ := by + grind + +example (α : Type) (β : Type) (a₁ a₂ : α) (b₁ b₂ : β) + (h₁ : α = β) + (h₂ : Eq.recOn h₁ a₁ = b₁) + (h₃ : a₁ = a₂) + (h₄ : b₁ = b₂) + : HEq a₂ b₂ := by + grind + +example (α : Type) (β : Type) (a₁ a₂ : α) (b₁ b₂ : β) + (h₁ : α = β) + (h₂ : Eq.ndrec (motive := id) a₁ h₁ = b₁) + (h₃ : a₁ = a₂) + (h₄ : b₁ = b₂) + : HEq a₂ b₂ := by + grind + +example (α : Type) (β : Type) (a₁ a₂ : α) (b₁ b₂ : β) + (h₁ : α = β) + (h₂ : Eq.rec (motive := fun x _ => x) a₁ h₁ = b₁) + (h₃ : a₁ = a₂) + (h₄ : b₁ = b₂) + : HEq a₂ b₂ := by + grind From 78ed072ab0191391e2f2673907e741e5a0a98327 Mon Sep 17 00:00:00 2001 From: Vlad Tsyrklevich Date: Wed, 8 Jan 2025 03:20:43 +0100 Subject: [PATCH 046/100] feat: add `Int.emod_sub_emod` and `Int.sub_emod_emod` (#6507) This PR adds the subtraction equivalents for `Int.emod_add_emod` (`(a % n + b) % n = (a + b) % n`) and `Int.add_emod_emod` (`(a + b % n) % n = (a + b) % n`). These are marked @[simp] like their addition equivalents. Discussed on Zulip in https://leanprover.zulipchat.com/#narrow/channel/270676-lean4/topic/Adding.20some.20sub_emod.20lemmas.20to.20DivModLemmas --- src/Init/Data/Int/DivModLemmas.lean | 7 +++++++ 1 file changed, 7 insertions(+) diff --git a/src/Init/Data/Int/DivModLemmas.lean b/src/Init/Data/Int/DivModLemmas.lean index d0c86c5c98d0..b8f121043a55 100644 --- a/src/Init/Data/Int/DivModLemmas.lean +++ b/src/Init/Data/Int/DivModLemmas.lean @@ -534,6 +534,13 @@ theorem mul_emod (a b n : Int) : (a * b) % n = (a % n) * (b % n) % n := by @[simp] theorem emod_emod (a b : Int) : (a % b) % b = a % b := by conv => rhs; rw [← emod_add_ediv a b, add_mul_emod_self_left] +@[simp] theorem emod_sub_emod (m n k : Int) : (m % n - k) % n = (m - k) % n := + Int.emod_add_emod m n (-k) + +@[simp] theorem sub_emod_emod (m n k : Int) : (m - n % k) % k = (m - n) % k := by + apply (emod_add_cancel_right (n % k)).mp + rw [Int.sub_add_cancel, Int.add_emod_emod, Int.sub_add_cancel] + theorem sub_emod (a b n : Int) : (a - b) % n = (a % n - b % n) % n := by apply (emod_add_cancel_right b).mp rw [Int.sub_add_cancel, ← Int.add_emod_emod, Int.sub_add_cancel, emod_emod] From 18b183f62bd744f8b61d5906df92914ca86578a4 Mon Sep 17 00:00:00 2001 From: Kyle Miller Date: Tue, 7 Jan 2025 18:25:21 -0800 Subject: [PATCH 047/100] feat: let `induction` take zero alteratives (#6486) This PR modifies the `induction`/`cases` syntax so that the `with` clause does not need to be followed by any alternatives. This improves friendliness of these tactics, since this lets them surface the names of the missing alternatives: ```lean example (n : Nat) : True := by induction n with /- ~~~~ alternative 'zero' has not been provided alternative 'succ' has not been provided -/ ``` Related to issue #3555 --- src/Init/Tactics.lean | 2 +- src/Lean/Elab/Tactic/Induction.lean | 46 ++++++++++++--------- tests/lean/inductionParse.lean.expected.out | 3 +- 3 files changed, 30 insertions(+), 21 deletions(-) diff --git a/src/Init/Tactics.lean b/src/Init/Tactics.lean index 05b0fe9aecad..e7f5a6eb28e8 100644 --- a/src/Init/Tactics.lean +++ b/src/Init/Tactics.lean @@ -818,7 +818,7 @@ syntax inductionAlt := ppDedent(ppLine) inductionAltLHS+ " => " (hole <|> synth After `with`, there is an optional tactic that runs on all branches, and then a list of alternatives. -/ -syntax inductionAlts := " with" (ppSpace colGt tactic)? withPosition((colGe inductionAlt)+) +syntax inductionAlts := " with" (ppSpace colGt tactic)? withPosition((colGe inductionAlt)*) /-- Assuming `x` is a variable in the local context with an inductive type, diff --git a/src/Lean/Elab/Tactic/Induction.lean b/src/Lean/Elab/Tactic/Induction.lean index bbc1830981c2..a9b4d650aa96 100644 --- a/src/Lean/Elab/Tactic/Induction.lean +++ b/src/Lean/Elab/Tactic/Induction.lean @@ -258,11 +258,11 @@ private def saveAltVarsInfo (altMVarId : MVarId) (altStx : Syntax) (fvarIds : Ar i := i + 1 open Language in -def evalAlts (elimInfo : ElimInfo) (alts : Array Alt) (optPreTac : Syntax) (altStxs : Array Syntax) +def evalAlts (elimInfo : ElimInfo) (alts : Array Alt) (optPreTac : Syntax) (altStxs? : Option (Array Syntax)) (initialInfo : Info) (numEqs : Nat := 0) (numGeneralized : Nat := 0) (toClear : Array FVarId := #[]) (toTag : Array (Ident × FVarId) := #[]) : TacticM Unit := do - let hasAlts := altStxs.size > 0 + let hasAlts := altStxs?.isSome if hasAlts then -- default to initial state outside of alts -- HACK: because this node has the same span as the original tactic, @@ -274,9 +274,7 @@ def evalAlts (elimInfo : ElimInfo) (alts : Array Alt) (optPreTac : Syntax) (altS where -- continuation in the correct info context goWithInfo := do - let hasAlts := altStxs.size > 0 - - if hasAlts then + if let some altStxs := altStxs? then if let some tacSnap := (← readThe Term.Context).tacSnap? then -- incrementality: create a new promise for each alternative, resolve current snapshot to -- them, eventually put each of them back in `Context.tacSnap?` in `applyAltStx` @@ -309,7 +307,8 @@ where -- continuation in the correct incrementality context goWithIncremental (tacSnaps : Array (SnapshotBundle TacticParsedSnapshot)) := do - let hasAlts := altStxs.size > 0 + let hasAlts := altStxs?.isSome + let altStxs := altStxs?.getD #[] let mut alts := alts -- initial sanity checks: named cases should be known, wildcards should be last @@ -343,12 +342,12 @@ where let altName := getAltName altStx if let some i := alts.findFinIdx? (·.1 == altName) then -- cover named alternative - applyAltStx tacSnaps altStxIdx altStx alts[i] + applyAltStx tacSnaps altStxs altStxIdx altStx alts[i] alts := alts.eraseIdx i else if !alts.isEmpty && isWildcard altStx then -- cover all alternatives for alt in alts do - applyAltStx tacSnaps altStxIdx altStx alt + applyAltStx tacSnaps altStxs altStxIdx altStx alt alts := #[] else throwErrorAt altStx "unused alternative '{altName}'" @@ -379,7 +378,7 @@ where altMVarIds.forM fun mvarId => admitGoal mvarId /-- Applies syntactic alternative to alternative goal. -/ - applyAltStx tacSnaps altStxIdx altStx alt := withRef altStx do + applyAltStx tacSnaps altStxs altStxIdx altStx alt := withRef altStx do let { name := altName, info, mvarId := altMVarId } := alt -- also checks for unknown alternatives let numFields ← getAltNumFields elimInfo altName @@ -476,7 +475,7 @@ private def generalizeVars (mvarId : MVarId) (stx : Syntax) (targets : Array Exp /-- Given `inductionAlts` of the form ``` -syntax inductionAlts := "with " (tactic)? withPosition( (colGe inductionAlt)+) +syntax inductionAlts := "with " (tactic)? withPosition( (colGe inductionAlt)*) ``` Return an array containing its alternatives. -/ @@ -486,21 +485,30 @@ private def getAltsOfInductionAlts (inductionAlts : Syntax) : Array Syntax := /-- Given `inductionAlts` of the form ``` -syntax inductionAlts := "with " (tactic)? withPosition( (colGe inductionAlt)+) +syntax inductionAlts := "with " (tactic)? withPosition( (colGe inductionAlt)*) ``` -runs `cont alts` where `alts` is an array containing all `inductionAlt`s while disabling incremental -reuse if any other syntax changed. +runs `cont (some alts)` where `alts` is an array containing all `inductionAlt`s while disabling incremental +reuse if any other syntax changed. If there's no `with` clause, then runs `cont none`. -/ private def withAltsOfOptInductionAlts (optInductionAlts : Syntax) - (cont : Array Syntax → TacticM α) : TacticM α := + (cont : Option (Array Syntax) → TacticM α) : TacticM α := Term.withNarrowedTacticReuse (stx := optInductionAlts) (fun optInductionAlts => if optInductionAlts.isNone then -- if there are no alternatives, what to compare is irrelevant as there will be no reuse (mkNullNode #[], mkNullNode #[]) else + -- if there are no alts, then use the `with` token for `inner` for a ref for messages + let altStxs := optInductionAlts[0].getArg 2 + let inner := if altStxs.getNumArgs > 0 then altStxs else optInductionAlts[0][0] -- `with` and tactic applied to all branches must be unchanged for reuse - (mkNullNode optInductionAlts[0].getArgs[:2], optInductionAlts[0].getArg 2)) - (fun alts => cont alts.getArgs) + (mkNullNode optInductionAlts[0].getArgs[:2], inner)) + (fun alts? => + if optInductionAlts.isNone then -- no `with` clause + cont none + else if alts?.isOfKind nullKind then -- has alts + cont (some alts?.getArgs) + else -- has `with` clause, but no alts + cont (some #[])) private def getOptPreTacOfOptInductionAlts (optInductionAlts : Syntax) : Syntax := if optInductionAlts.isNone then mkNullNode else optInductionAlts[0][1] @@ -518,7 +526,7 @@ private def expandMultiAlt? (alt : Syntax) : Option (Array Syntax) := Id.run do /-- Given `inductionAlts` of the form ``` -syntax inductionAlts := "with " (tactic)? withPosition( (colGe inductionAlt)+) +syntax inductionAlts := "with " (tactic)? withPosition( (colGe inductionAlt)*) ``` Return `some inductionAlts'` if one of the alternatives have multiple LHSs, in the new `inductionAlts'` all alternatives have a single LHS. @@ -700,10 +708,10 @@ def evalInduction : Tactic := fun stx => -- unchanged -- everything up to the alternatives must be unchanged for reuse Term.withNarrowedArgTacticReuse (stx := stx) (argIdx := 4) fun optInductionAlts => do - withAltsOfOptInductionAlts optInductionAlts fun alts => do + withAltsOfOptInductionAlts optInductionAlts fun alts? => do let optPreTac := getOptPreTacOfOptInductionAlts optInductionAlts mvarId.assign result.elimApp - ElimApp.evalAlts elimInfo result.alts optPreTac alts initInfo (numGeneralized := n) (toClear := targetFVarIds) + ElimApp.evalAlts elimInfo result.alts optPreTac alts? initInfo (numGeneralized := n) (toClear := targetFVarIds) appendGoals result.others.toList where checkTargets (targets : Array Expr) : MetaM Unit := do diff --git a/tests/lean/inductionParse.lean.expected.out b/tests/lean/inductionParse.lean.expected.out index da90e48b9105..39bafa0b0b87 100644 --- a/tests/lean/inductionParse.lean.expected.out +++ b/tests/lean/inductionParse.lean.expected.out @@ -1 +1,2 @@ -inductionParse.lean:4:18-5:6: error: unexpected identifier; expected '|' +inductionParse.lean:4:14-4:18: error: alternative 'zero' has not been provided +inductionParse.lean:4:14-4:18: error: alternative 'succ' has not been provided From 91cbd7c80e66fabbbc9cff20818129bbc27b9e00 Mon Sep 17 00:00:00 2001 From: Harun Khan <58151072+mhk119@users.noreply.github.com> Date: Tue, 7 Jan 2025 21:55:50 -0500 Subject: [PATCH 048/100] feat: `BitVec.toInt_shiftLeft` theorem (#6346) This PR completes the toNat/Int/Fin family for `shiftLeft`. --- src/Init/Data/BitVec/Lemmas.lean | 9 +++++++-- 1 file changed, 7 insertions(+), 2 deletions(-) diff --git a/src/Init/Data/BitVec/Lemmas.lean b/src/Init/Data/BitVec/Lemmas.lean index 0b6596bf4388..68375c1c2c4b 100644 --- a/src/Init/Data/BitVec/Lemmas.lean +++ b/src/Init/Data/BitVec/Lemmas.lean @@ -1142,11 +1142,16 @@ theorem getMsb_not {x : BitVec w} : /-! ### shiftLeft -/ @[simp, bv_toNat] theorem toNat_shiftLeft {x : BitVec v} : - BitVec.toNat (x <<< n) = BitVec.toNat x <<< n % 2^v := + (x <<< n).toNat = x.toNat <<< n % 2^v := BitVec.toNat_ofNat _ _ +@[simp] theorem toInt_shiftLeft {x : BitVec w} : + (x <<< n).toInt = (x.toNat <<< n : Int).bmod (2^w) := by + rw [toInt_eq_toNat_bmod, toNat_shiftLeft, Nat.shiftLeft_eq] + simp + @[simp] theorem toFin_shiftLeft {n : Nat} (x : BitVec w) : - BitVec.toFin (x <<< n) = Fin.ofNat' (2^w) (x.toNat <<< n) := rfl + (x <<< n).toFin = Fin.ofNat' (2^w) (x.toNat <<< n) := rfl @[simp] theorem shiftLeft_zero (x : BitVec w) : x <<< 0 = x := by From 9040108e2f069dcba488246424812649f37d9fca Mon Sep 17 00:00:00 2001 From: Tobias Grosser Date: Wed, 8 Jan 2025 02:57:53 +0000 Subject: [PATCH 049/100] feat: add BitVec.[toNat|toInt|toFin|getLsbD|getMsbD|getElem|msb]_fill (#6177) This PR implements `BitVec.*_fill`. We also add `toInt_allOnes` and `toFin_allOnes` as the former is needed here. This completes the allOnes API. --- src/Init/Data/BitVec/Lemmas.lean | 61 ++++++++++++++++++++++++++++++++ 1 file changed, 61 insertions(+) diff --git a/src/Init/Data/BitVec/Lemmas.lean b/src/Init/Data/BitVec/Lemmas.lean index 68375c1c2c4b..a4599aaddb49 100644 --- a/src/Init/Data/BitVec/Lemmas.lean +++ b/src/Init/Data/BitVec/Lemmas.lean @@ -785,6 +785,19 @@ theorem extractLsb'_eq_extractLsb {w : Nat} (x : BitVec w) (start len : Nat) (h unfold allOnes simp +@[simp] theorem toInt_allOnes : (allOnes w).toInt = if 0 < w then -1 else 0 := by + norm_cast + by_cases h : w = 0 + · subst h + simp + · have : 1 < 2 ^ w := by simp [h] + simp [BitVec.toInt] + omega + +@[simp] theorem toFin_allOnes : (allOnes w).toFin = Fin.ofNat' (2^w) (2^w - 1) := by + ext + simp + @[simp] theorem getLsbD_allOnes : (allOnes v).getLsbD i = decide (i < v) := by simp [allOnes] @@ -2382,6 +2395,54 @@ theorem not_neg (x : BitVec w) : ~~~(-x) = x + -1#w := by show (_ - x.toNat) % _ = _ by rw [Nat.mod_eq_of_lt (by omega)]] omega +/-! ### fill -/ + +@[simp] +theorem getLsbD_fill {w i : Nat} {v : Bool} : + (fill w v).getLsbD i = (v && decide (i < w)) := by + by_cases h : v + <;> simp [h, BitVec.fill, BitVec.negOne_eq_allOnes] + +@[simp] +theorem getMsbD_fill {w i : Nat} {v : Bool} : + (fill w v).getMsbD i = (v && decide (i < w)) := by + by_cases h : v + <;> simp [h, BitVec.fill, BitVec.negOne_eq_allOnes] + +@[simp] +theorem getElem_fill {w i : Nat} {v : Bool} (h : i < w) : + (fill w v)[i] = v := by + by_cases h : v + <;> simp [h, BitVec.fill, BitVec.negOne_eq_allOnes] + +@[simp] +theorem msb_fill {w : Nat} {v : Bool} : + (fill w v).msb = (v && decide (0 < w)) := by + simp [BitVec.msb] + +theorem fill_eq {w : Nat} {v : Bool} : fill w v = if v = true then allOnes w else 0#w := by + by_cases h : v <;> (simp only [h] ; ext ; simp) + +@[simp] +theorem fill_true {w : Nat} : fill w true = allOnes w := by + simp [fill_eq] + +@[simp] +theorem fill_false {w : Nat} : fill w false = 0#w := by + simp [fill_eq] + +@[simp] theorem fill_toNat {w : Nat} {v : Bool} : + (fill w v).toNat = if v = true then 2^w - 1 else 0 := by + by_cases h : v <;> simp [h] + +@[simp] theorem fill_toInt {w : Nat} {v : Bool} : + (fill w v).toInt = if v = true && 0 < w then -1 else 0 := by + by_cases h : v <;> simp [h] + +@[simp] theorem fill_toFin {w : Nat} {v : Bool} : + (fill w v).toFin = if v = true then (allOnes w).toFin else Fin.ofNat' (2 ^ w) 0 := by + by_cases h : v <;> simp [h] + /-! ### mul -/ theorem mul_def {n} {x y : BitVec n} : x * y = (ofFin <| x.toFin * y.toFin) := by rfl From 00ef231a6e03398c2ad3b577ab036f901ec88543 Mon Sep 17 00:00:00 2001 From: Leonardo de Moura Date: Tue, 7 Jan 2025 19:10:11 -0800 Subject: [PATCH 050/100] feat: split on `match`-expressions in the `grind` tactic (#6569) This PR adds support for case splitting on `match`-expressions in `grind`. We still need to add support for resolving the antecedents of `match`-conditional equations. --- src/Init/Grind/Util.lean | 7 +++ src/Lean/Meta/Tactic/Cases.lean | 2 +- src/Lean/Meta/Tactic/Grind.lean | 3 +- src/Lean/Meta/Tactic/Grind/CasesMatch.lean | 53 ++++++++++++++++++++++ src/Lean/Meta/Tactic/Grind/EMatch.lean | 16 ++++--- src/Lean/Meta/Tactic/Grind/Propagate.lean | 7 +++ src/Lean/Meta/Tactic/Grind/Split.lean | 13 ++++-- tests/lean/run/grind_match1.lean | 11 +++-- tests/lean/run/grind_match2.lean | 28 ++++++++++++ 9 files changed, 121 insertions(+), 19 deletions(-) create mode 100644 src/Lean/Meta/Tactic/Grind/CasesMatch.lean create mode 100644 tests/lean/run/grind_match2.lean diff --git a/src/Init/Grind/Util.lean b/src/Init/Grind/Util.lean index 9ef005388f82..f37dcc424898 100644 --- a/src/Init/Grind/Util.lean +++ b/src/Init/Grind/Util.lean @@ -21,6 +21,13 @@ def doNotSimp {α : Sort u} (a : α) : α := a /-- Gadget for representing offsets `t+k` in patterns. -/ def offset (a b : Nat) : Nat := a + b +/-- +Gadget for annotating the equalities in `match`-equations conclusions. +`_origin` is the term used to instantiate the `match`-equation using E-matching. +When `EqMatch a b origin` is `True`, we mark `origin` as a resolved case-split. +-/ +def EqMatch (a b : α) {_origin : α} : Prop := a = b + theorem nestedProof_congr (p q : Prop) (h : p = q) (hp : p) (hq : q) : HEq (nestedProof p hp) (nestedProof q hq) := by subst h; apply HEq.refl diff --git a/src/Lean/Meta/Tactic/Cases.lean b/src/Lean/Meta/Tactic/Cases.lean index 56e8b8514a27..7e4e7cf34f4c 100644 --- a/src/Lean/Meta/Tactic/Cases.lean +++ b/src/Lean/Meta/Tactic/Cases.lean @@ -33,7 +33,7 @@ private def mkEqAndProof (lhs rhs : Expr) : MetaM (Expr × Expr) := do else pure (mkApp4 (mkConst ``HEq [u]) lhsType lhs rhsType rhs, mkApp2 (mkConst ``HEq.refl [u]) lhsType lhs) -private partial def withNewEqs (targets targetsNew : Array Expr) (k : Array Expr → Array Expr → MetaM α) : MetaM α := +partial def withNewEqs (targets targetsNew : Array Expr) (k : Array Expr → Array Expr → MetaM α) : MetaM α := let rec loop (i : Nat) (newEqs : Array Expr) (newRefls : Array Expr) := do if i < targets.size then let (newEqType, newRefl) ← mkEqAndProof targets[i]! targetsNew[i]! diff --git a/src/Lean/Meta/Tactic/Grind.lean b/src/Lean/Meta/Tactic/Grind.lean index 593bb147c6ab..80edeeef6cac 100644 --- a/src/Lean/Meta/Tactic/Grind.lean +++ b/src/Lean/Meta/Tactic/Grind.lean @@ -23,7 +23,7 @@ import Lean.Meta.Tactic.Grind.Parser import Lean.Meta.Tactic.Grind.EMatchTheorem import Lean.Meta.Tactic.Grind.EMatch import Lean.Meta.Tactic.Grind.Main - +import Lean.Meta.Tactic.Grind.CasesMatch namespace Lean @@ -52,5 +52,6 @@ builtin_initialize registerTraceClass `grind.debug.proj builtin_initialize registerTraceClass `grind.debug.parent builtin_initialize registerTraceClass `grind.debug.final builtin_initialize registerTraceClass `grind.debug.forallPropagator +builtin_initialize registerTraceClass `grind.debug.split end Lean diff --git a/src/Lean/Meta/Tactic/Grind/CasesMatch.lean b/src/Lean/Meta/Tactic/Grind/CasesMatch.lean new file mode 100644 index 000000000000..1e5f07ed6a0d --- /dev/null +++ b/src/Lean/Meta/Tactic/Grind/CasesMatch.lean @@ -0,0 +1,53 @@ +/- +Copyright (c) 2025 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.Meta.Tactic.Util +import Lean.Meta.Tactic.Cases +import Lean.Meta.Match.MatcherApp + +namespace Lean.Meta.Grind + +def casesMatch (mvarId : MVarId) (e : Expr) : MetaM (List MVarId) := mvarId.withContext do + let some app ← matchMatcherApp? e + | throwTacticEx `grind.casesMatch mvarId m!"`match`-expression expected{indentExpr e}" + let (motive, eqRefls) ← mkMotiveAndRefls app + let target ← mvarId.getType + let mut us := app.matcherLevels + if let some i := app.uElimPos? then + us := us.set! i (← getLevel target) + let splitterName := (← Match.getEquationsFor app.matcherName).splitterName + let splitterApp := mkConst splitterName us.toList + let splitterApp := mkAppN splitterApp app.params + let splitterApp := mkApp splitterApp motive + let splitterApp := mkAppN splitterApp app.discrs + let (mvars, _, _) ← forallMetaBoundedTelescope (← inferType splitterApp) app.alts.size (kind := .syntheticOpaque) + let splitterApp := mkAppN splitterApp mvars + let val := mkAppN splitterApp eqRefls + mvarId.assign val + updateTags mvars + return mvars.toList.map (·.mvarId!) +where + mkMotiveAndRefls (app : MatcherApp) : MetaM (Expr × Array Expr) := do + let dummy := mkSort 0 + let aux := mkApp (mkAppN e.getAppFn app.params) dummy + forallBoundedTelescope (← inferType aux) app.discrs.size fun xs _ => do + withNewEqs app.discrs xs fun eqs eqRefls => do + let type ← mvarId.getType + let type ← mkForallFVars eqs type + let motive ← mkLambdaFVars xs type + return (motive, eqRefls) + + updateTags (mvars : Array Expr) : MetaM Unit := do + let tag ← mvarId.getTag + if mvars.size == 1 then + mvars[0]!.mvarId!.setTag tag + else + let mut idx := 1 + for mvar in mvars do + mvar.mvarId!.setTag (Name.num tag idx) + idx := idx + 1 + +end Lean.Meta.Grind diff --git a/src/Lean/Meta/Tactic/Grind/EMatch.lean b/src/Lean/Meta/Tactic/Grind/EMatch.lean index bce265b73818..fccb34b22f68 100644 --- a/src/Lean/Meta/Tactic/Grind/EMatch.lean +++ b/src/Lean/Meta/Tactic/Grind/EMatch.lean @@ -199,14 +199,18 @@ private def processContinue (c : Choice) (p : Expr) : M Unit := do let c := { c with gen := Nat.max gen c.gen } modify fun s => { s with choiceStack := c :: s.choiceStack } -/-- Helper function for marking parts of `match`-equation theorem as "do-not-simplify" -/ -private partial def annotateMatchEqnType (prop : Expr) : M Expr := do +/-- +Helper function for marking parts of `match`-equation theorem as "do-not-simplify" +`initApp` is the match-expression used to instantiate the `match`-equation. +-/ +private partial def annotateMatchEqnType (prop : Expr) (initApp : Expr) : M Expr := do if let .forallE n d b bi := prop then withLocalDecl n bi (← markAsDoNotSimp d) fun x => do - mkForallFVars #[x] (← annotateMatchEqnType (b.instantiate1 x)) + mkForallFVars #[x] (← annotateMatchEqnType (b.instantiate1 x) initApp) else let_expr f@Eq α lhs rhs := prop | return prop - return mkApp3 f α (← markAsDoNotSimp lhs) rhs + -- See comment at `Grind.EqMatch` + return mkApp4 (mkConst ``Grind.EqMatch f.constLevels!) α (← markAsDoNotSimp lhs) rhs initApp /-- Stores new theorem instance in the state. @@ -218,9 +222,7 @@ private def addNewInstance (origin : Origin) (proof : Expr) (generation : Nat) : check proof let mut prop ← inferType proof if Match.isMatchEqnTheorem (← getEnv) origin.key then - -- `initApp` is a match-application that we don't need to split at anymore. - markCaseSplitAsResolved (← read).initApp - prop ← annotateMatchEqnType prop + prop ← annotateMatchEqnType prop (← read).initApp trace_goal[grind.ematch.instance] "{← origin.pp}: {prop}" addTheoremInstance proof prop (generation+1) diff --git a/src/Lean/Meta/Tactic/Grind/Propagate.lean b/src/Lean/Meta/Tactic/Grind/Propagate.lean index 42d1135a14f0..bfe1a3d014a5 100644 --- a/src/Lean/Meta/Tactic/Grind/Propagate.lean +++ b/src/Lean/Meta/Tactic/Grind/Propagate.lean @@ -134,6 +134,13 @@ builtin_grind_propagator propagateEqDown ↓Eq := fun e => do let_expr Eq _ a b := e | return () pushEq a b <| mkApp2 (mkConst ``of_eq_true) e (← mkEqTrueProof e) +/-- Propagates `EqMatch` downwards -/ +builtin_grind_propagator propagateEqMatchDown ↓Grind.EqMatch := fun e => do + if (← isEqTrue e) then + let_expr Grind.EqMatch _ a b origin := e | return () + markCaseSplitAsResolved origin + pushEq a b <| mkApp2 (mkConst ``of_eq_true) e (← mkEqTrueProof e) + /-- Propagates `HEq` downwards -/ builtin_grind_propagator propagateHEqDown ↓HEq := fun e => do if (← isEqTrue e) then diff --git a/src/Lean/Meta/Tactic/Grind/Split.lean b/src/Lean/Meta/Tactic/Grind/Split.lean index 96ae7ce4160a..e57a5270e404 100644 --- a/src/Lean/Meta/Tactic/Grind/Split.lean +++ b/src/Lean/Meta/Tactic/Grind/Split.lean @@ -7,6 +7,7 @@ prelude import Lean.Meta.Tactic.Grind.Types import Lean.Meta.Tactic.Grind.Intro import Lean.Meta.Tactic.Grind.Cases +import Lean.Meta.Tactic.Grind.CasesMatch namespace Lean.Meta.Grind @@ -50,10 +51,10 @@ private def checkCaseSplitStatus (e : Expr) : GoalM CaseSplitStatus := do return .ready | _ => if (← isResolvedCaseSplit e) then + trace[grind.debug.split] "split resolved: {e}" return .resolved if (← isMatcherApp e) then - return .notReady -- TODO: implement splitters for `match` - -- return .ready + return .ready let .const declName .. := e.getAppFn | unreachable! if (← isInductivePredicate declName <&&> isEqTrue e) then return .ready @@ -111,9 +112,11 @@ def splitNext : GrindTactic := fun goal => do | return none let gen ← getGeneration c trace_goal[grind.split] "{c}, generation: {gen}" - -- TODO: `match` - let major ← mkCasesMajor c - let mvarIds ← cases (← get).mvarId major + let mvarIds ← if (← isMatcherApp c) then + casesMatch (← get).mvarId c + else + let major ← mkCasesMajor c + cases (← get).mvarId major let goal ← get let goals := mvarIds.map fun mvarId => { goal with mvarId } let goals ← introNewHyp goals [] (gen+1) diff --git a/tests/lean/run/grind_match1.lean b/tests/lean/run/grind_match1.lean index c50bfcf6e618..303b9f7172f0 100644 --- a/tests/lean/run/grind_match1.lean +++ b/tests/lean/run/grind_match1.lean @@ -24,15 +24,16 @@ info: [grind.assert] (match as, bs with [grind.assert] a₁ :: f 0 = as [grind.assert] f 0 = a₂ :: f 1 [grind.assert] ¬d = [] +[grind.assert] Lean.Grind.EqMatch + (match a₁ :: a₂ :: f 1, [] with + | [], x => bs + | head :: head_1 :: tail, [] => [] + | x :: xs, ys => x :: g xs ys) + [] [grind.split.resolved] match as, bs with | [], x => bs | head :: head_1 :: tail, [] => [] | x :: xs, ys => x :: g xs ys -[grind.assert] (match a₁ :: a₂ :: f 1, [] with - | [], x => bs - | head :: head_1 :: tail, [] => [] - | x :: xs, ys => x :: g xs ys) = - [] -/ #guard_msgs (info) in example (f : Nat → List Nat) : g as bs = d → bs = [] → a₁ :: f 0 = as → f 0 = a₂ :: f 1 → d = [] := by diff --git a/tests/lean/run/grind_match2.lean b/tests/lean/run/grind_match2.lean new file mode 100644 index 000000000000..24160dd19545 --- /dev/null +++ b/tests/lean/run/grind_match2.lean @@ -0,0 +1,28 @@ +def g (a : α) (as : List α) : List α := + match as with + | [] => [a] + | b::bs => a::a::b::bs + +set_option trace.grind true in +set_option trace.grind.assert true in +example : ¬ (g a as).isEmpty := by + unfold List.isEmpty + unfold g + grind + +def h (as : List Nat) := + match as with + | [] => 1 + | [_] => 2 + | _::_::_ => 3 + +/-- +info: [grind] closed `grind.1` +[grind] closed `grind.2` +[grind] closed `grind.3` +-/ +#guard_msgs (info) in +set_option trace.grind true in +example : h as ≠ 0 := by + unfold h + grind From f01471f6203168c8ab349e552214d552bebd229a Mon Sep 17 00:00:00 2001 From: Sebastian Graf Date: Wed, 8 Jan 2025 09:36:45 +0100 Subject: [PATCH 051/100] fix: proper "excess binders" error locations for `rintro` and `intro` (#6565) This PR fixes the location of the error emitted when the `rintro` and `intro` tactics cannot introduce the requested number of binders. This patch adds a few `withRef` wrappers to invocations of `MVarId.intro` to fix error locations. Perhaps `MVarId.intro` should take a syntax object to set the location itself in the future; however there are a couple other call sites which would need non-trivial fixup. Closes #5659. --- src/Init/Prelude.lean | 10 +++++++++ src/Lean/Elab/Tactic/BuiltinTactic.lean | 8 +++---- src/Lean/Elab/Tactic/RCases.lean | 2 +- tests/lean/interactive/5659.lean | 13 +++++++++++ tests/lean/interactive/5659.lean.expected.out | 22 +++++++++++++++++++ .../incrementalTactic.lean.expected.out | 5 +++-- 6 files changed, 53 insertions(+), 7 deletions(-) create mode 100644 tests/lean/interactive/5659.lean create mode 100644 tests/lean/interactive/5659.lean.expected.out diff --git a/src/Init/Prelude.lean b/src/Init/Prelude.lean index 0d9514ae77d3..0253c5df6fa0 100644 --- a/src/Init/Prelude.lean +++ b/src/Init/Prelude.lean @@ -4170,6 +4170,16 @@ def withRef [Monad m] [MonadRef m] {α} (ref : Syntax) (x : m α) : m α := let ref := replaceRef ref oldRef MonadRef.withRef ref x +/-- +If `ref? = some ref`, run `x : m α` with a modified value for the `ref` by calling `withRef`. +Otherwise, run `x` directly. +-/ +@[always_inline, inline] +def withRef? [Monad m] [MonadRef m] {α} (ref? : Option Syntax) (x : m α) : m α := + match ref? with + | some ref => withRef ref x + | _ => x + /-- A monad that supports syntax quotations. Syntax quotations (in term position) are monadic values that when executed retrieve the current "macro scope" from the monad and apply it to every identifier they introduce diff --git a/src/Lean/Elab/Tactic/BuiltinTactic.lean b/src/Lean/Elab/Tactic/BuiltinTactic.lean index aa02a0f84903..72c5d139c4e8 100644 --- a/src/Lean/Elab/Tactic/BuiltinTactic.lean +++ b/src/Lean/Elab/Tactic/BuiltinTactic.lean @@ -362,9 +362,9 @@ partial def evalChoiceAux (tactics : Array Syntax) (i : Nat) : TacticM Unit := | `(tactic| intro $h:term $hs:term*) => evalTactic (← `(tactic| intro $h:term; intro $hs:term*)) | _ => throwUnsupportedSyntax where - introStep (ref : Option Syntax) (n : Name) (typeStx? : Option Syntax := none) : TacticM Unit := do + introStep (ref? : Option Syntax) (n : Name) (typeStx? : Option Syntax := none) : TacticM Unit := do let fvarId ← liftMetaTacticAux fun mvarId => do - let (fvarId, mvarId) ← mvarId.intro n + let (fvarId, mvarId) ← withRef? ref? <| mvarId.intro n pure (fvarId, [mvarId]) if let some typeStx := typeStx? then withMainContext do @@ -374,9 +374,9 @@ where unless (← isDefEqGuarded type fvarType) do throwError "type mismatch at `intro {fvar}`{← mkHasTypeButIsExpectedMsg fvarType type}" liftMetaTactic fun mvarId => return [← mvarId.replaceLocalDeclDefEq fvarId type] - if let some stx := ref then + if let some ref := ref? then withMainContext do - Term.addLocalVarInfo stx (mkFVar fvarId) + Term.addLocalVarInfo ref (mkFVar fvarId) @[builtin_tactic Lean.Parser.Tactic.introMatch] def evalIntroMatch : Tactic := fun stx => do let matchAlts := stx[1] diff --git a/src/Lean/Elab/Tactic/RCases.lean b/src/Lean/Elab/Tactic/RCases.lean index d8ecac3b2161..31d3ecac7164 100644 --- a/src/Lean/Elab/Tactic/RCases.lean +++ b/src/Lean/Elab/Tactic/RCases.lean @@ -507,7 +507,7 @@ partial def rintroCore (g : MVarId) (fs : FVarSubst) (clears : Array FVarId) (a match pat with | `(rintroPat| $pat:rcasesPat) => let pat := (← RCasesPatt.parse pat).typed? ref ty? - let (v, g) ← g.intro (pat.name?.getD `_) + let (v, g) ← withRef pat.ref <| g.intro (pat.name?.getD `_) rcasesCore g fs clears (.fvar v) a pat cont | `(rintroPat| ($(pats)* $[: $ty?']?)) => let ref := if pats.size == 1 then pat.raw else .missing diff --git a/tests/lean/interactive/5659.lean b/tests/lean/interactive/5659.lean new file mode 100644 index 000000000000..92b03237bf60 --- /dev/null +++ b/tests/lean/interactive/5659.lean @@ -0,0 +1,13 @@ +/-! + # `rintro` and `intro` error message should point to excess arg + + Below, the errors should point to `h` because there is no lambda it binds. + The error should not point to `hn`; it would be OKish to underline the whole line. -/ + +example : (∃ n : Nat, n < 0) → False := by + rintro ⟨n, hn⟩ h + --^ collectDiagnostics + +example : (∃ n : Nat, n < 0) → False := by + intro ⟨n, hn⟩ h + --^ collectDiagnostics diff --git a/tests/lean/interactive/5659.lean.expected.out b/tests/lean/interactive/5659.lean.expected.out new file mode 100644 index 000000000000..95178ac23008 --- /dev/null +++ b/tests/lean/interactive/5659.lean.expected.out @@ -0,0 +1,22 @@ +{"version": 1, + "uri": "file:///5659.lean", + "diagnostics": + [{"source": "Lean 4", + "severity": 1, + "range": + {"start": {"line": 7, "character": 17}, "end": {"line": 7, "character": 18}}, + "message": + "tactic 'introN' failed, insufficient number of binders\ncase intro\nn : Nat\nhn : n < 0\n⊢ False", + "fullRange": + {"start": {"line": 7, "character": 17}, + "end": {"line": 7, "character": 18}}}, + {"source": "Lean 4", + "severity": 1, + "range": + {"start": {"line": 11, "character": 16}, + "end": {"line": 11, "character": 17}}, + "message": + "tactic 'introN' failed, insufficient number of binders\nn : Nat\nhn : n < 0\n⊢ False", + "fullRange": + {"start": {"line": 11, "character": 16}, + "end": {"line": 11, "character": 17}}}]} diff --git a/tests/lean/interactive/incrementalTactic.lean.expected.out b/tests/lean/interactive/incrementalTactic.lean.expected.out index 706ae32c688f..b429e40063b2 100644 --- a/tests/lean/interactive/incrementalTactic.lean.expected.out +++ b/tests/lean/interactive/incrementalTactic.lean.expected.out @@ -12,11 +12,12 @@ t 2 [{"source": "Lean 4", "severity": 1, "range": - {"start": {"line": 4, "character": 4}, "end": {"line": 4, "character": 13}}, + {"start": {"line": 4, "character": 12}, "end": {"line": 4, "character": 13}}, "message": "tactic 'introN' failed, insufficient number of binders\na n : Nat\n⊢ True", "fullRange": - {"start": {"line": 4, "character": 4}, "end": {"line": 4, "character": 13}}}, + {"start": {"line": 4, "character": 12}, + "end": {"line": 4, "character": 13}}}, {"source": "Lean 4", "severity": 1, "range": From 48eb3084a0e0e15d6c22287d1c17072cbe9f1541 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Henrik=20B=C3=B6ving?= Date: Wed, 8 Jan 2025 13:06:25 +0100 Subject: [PATCH 052/100] perf: speed up JSON serialisation (#6479) This PR speeds up JSON serialisation by using a lookup table to check whether a string needs to be escaped. The approach is based on https://byroot.github.io/ruby/json/2024/12/15/optimizing-ruby-json-part-1.html. --- src/Lean/Data/Json/Printer.lean | 37 ++++++++++++++++++++++++++++++++- 1 file changed, 36 insertions(+), 1 deletion(-) diff --git a/src/Lean/Data/Json/Printer.lean b/src/Lean/Data/Json/Printer.lean index 17bc93ee5638..d58f5fac298a 100644 --- a/src/Lean/Data/Json/Printer.lean +++ b/src/Lean/Data/Json/Printer.lean @@ -11,6 +11,22 @@ import Init.Data.List.Impl namespace Lean namespace Json +set_option maxRecDepth 1024 in +/-- +This table contains for each UTF-8 byte whether we need to escape a string that contains it. +-/ +private def escapeTable : { xs : ByteArray // xs.size = 256 } := + ⟨ByteArray.mk #[ + 1,1,1,1,1,1,1,1,1,1,1,1,1,1,1,1, 1,1,1,1,1,1,1,1,1,1,1,1,1,1,1,1, + 0,0,1,0,0,0,0,0,0,0,0,0,0,0,0,1, 0,0,0,0,0,0,0,0,0,0,0,0,0,0,0,0, + 0,0,0,0,0,0,0,0,0,0,0,0,0,0,0,0, 0,0,0,0,0,0,0,0,0,0,0,0,1,0,0,0, + 0,0,0,0,0,0,0,0,0,0,0,0,0,0,0,0, 0,0,0,0,0,0,0,0,0,0,0,0,0,0,0,0, + 1,1,1,1,1,1,1,1,1,1,1,1,1,1,1,1, 1,1,1,1,1,1,1,1,1,1,1,1,1,1,1,1, + 1,1,1,1,1,1,1,1,1,1,1,1,1,1,1,1, 1,1,1,1,1,1,1,1,1,1,1,1,1,1,1,1, + 1,1,1,1,1,1,1,1,1,1,1,1,1,1,1,1, 1,1,1,1,1,1,1,1,1,1,1,1,1,1,1,1, + 1,1,1,1,1,1,1,1,1,1,1,1,1,1,1,1, 1,1,1,1,1,1,1,1,1,1,1,1,1,1,1,1 + ], by rfl⟩ + private def escapeAux (acc : String) (c : Char) : String := -- escape ", \, \n and \r, keep all other characters ≥ 0x20 and render characters < 0x20 with \u if c = '"' then -- hack to prevent emacs from regarding the rest of the file as a string: " @@ -39,8 +55,27 @@ private def escapeAux (acc : String) (c : Char) : String := let d4 := Nat.digitChar (n % 16) acc ++ "\\u" |>.push d1 |>.push d2 |>.push d3 |>.push d4 +private def needEscape (s : String) : Bool := + go s 0 +where + go (s : String) (i : Nat) : Bool := + if h : i < s.utf8ByteSize then + let byte := s.getUtf8Byte i h + have h1 : byte.toNat < 256 := UInt8.toNat_lt_size byte + have h2 : escapeTable.val.size = 256 := escapeTable.property + if escapeTable.val.get byte.toNat (Nat.lt_of_lt_of_eq h1 h2.symm) == 0 then + go s (i + 1) + else + true + else + false + def escape (s : String) (acc : String := "") : String := - s.foldl escapeAux acc + -- If we don't have any characters that need to be escaped we can just append right away. + if needEscape s then + s.foldl escapeAux acc + else + acc ++ s def renderString (s : String) (acc : String := "") : String := let acc := acc ++ "\"" From 680ede7a89669d767368dc54953b7371324b5dde Mon Sep 17 00:00:00 2001 From: Sebastian Ullrich Date: Wed, 8 Jan 2025 12:56:27 +0000 Subject: [PATCH 053/100] fix: set LLVM sysroot consistently (#6574) This PR actually prevents Lake from accidentally picking up other toolchains installed on the machine. Fixes regression introduced in #6176 --- script/prepare-llvm-linux.sh | 4 ++-- script/prepare-llvm-macos.sh | 10 ++-------- script/prepare-llvm-mingw.sh | 2 +- 3 files changed, 5 insertions(+), 11 deletions(-) diff --git a/script/prepare-llvm-linux.sh b/script/prepare-llvm-linux.sh index 2ebb7a41b88e..ac634aee4ec8 100755 --- a/script/prepare-llvm-linux.sh +++ b/script/prepare-llvm-linux.sh @@ -63,8 +63,8 @@ else fi # use `-nostdinc` to make sure headers are not visible by default (in particular, not to `#include_next` in the clang headers), # but do not change sysroot so users can still link against system libs -echo -n " -DLEANC_INTERNAL_FLAGS='-nostdinc -isystem ROOT/include/clang' -DLEANC_CC=ROOT/bin/clang" -echo -n " -DLEANC_INTERNAL_LINKER_FLAGS='-L ROOT/lib -L ROOT/lib/glibc ROOT/lib/glibc/libc_nonshared.a ROOT/lib/glibc/libpthread_nonshared.a -Wl,--as-needed -Wl,-Bstatic -lgmp -lunwind -luv -Wl,-Bdynamic -Wl,--no-as-needed -fuse-ld=lld --ld-path=ROOT/bin/ld.lld'" +echo -n " -DLEANC_INTERNAL_FLAGS='--sysroot ROOT -nostdinc -isystem ROOT/include/clang' -DLEANC_CC=ROOT/bin/clang" +echo -n " -DLEANC_INTERNAL_LINKER_FLAGS='--sysroot ROOT -L ROOT/lib -L ROOT/lib/glibc ROOT/lib/glibc/libc_nonshared.a ROOT/lib/glibc/libpthread_nonshared.a -Wl,--as-needed -Wl,-Bstatic -lgmp -lunwind -luv -Wl,-Bdynamic -Wl,--no-as-needed -fuse-ld=lld'" # when not using the above flags, link GMP dynamically/as usual echo -n " -DLEAN_EXTRA_LINKER_FLAGS='-Wl,--as-needed -lgmp -luv -lpthread -ldl -lrt -Wl,--no-as-needed'" # do not set `LEAN_CC` for tests diff --git a/script/prepare-llvm-macos.sh b/script/prepare-llvm-macos.sh index 8c15fb115db4..715ba4e2f72a 100755 --- a/script/prepare-llvm-macos.sh +++ b/script/prepare-llvm-macos.sh @@ -52,13 +52,7 @@ if [[ -L llvm-host ]]; then else echo -n " -DCMAKE_C_COMPILER=$PWD/llvm-host/bin/clang -DLEANC_OPTS='--sysroot $PWD/stage1 -resource-dir $PWD/stage1/lib/clang/15.0.1 ${EXTRA_FLAGS:-}'" fi -echo -n " -DLEANC_INTERNAL_FLAGS='-nostdinc -isystem ROOT/include/clang' -DLEANC_CC=ROOT/bin/clang" -if [[ "$(uname -p)" == "i386" ]]; then - # `--ld-path` creates some platform_version troubles on macOS x64 but this is not a high-prio platform anymore and - # its users are not that likely to have conflicting `lld`s in their `PATH` - echo -n " -DLEANC_INTERNAL_LINKER_FLAGS='-L ROOT/lib -L ROOT/lib/libc -fuse-ld=lld'" -else - echo -n " -DLEANC_INTERNAL_LINKER_FLAGS='-L ROOT/lib -L ROOT/lib/libc -fuse-ld=lld --ld-path=ROOT/bin/ld64.lld'" -fi +echo -n " -DLEANC_INTERNAL_FLAGS='--sysroot ROOT -nostdinc -isystem ROOT/include/clang' -DLEANC_CC=ROOT/bin/clang" +echo -n " -DLEANC_INTERNAL_LINKER_FLAGS='--sysroot ROOT -L ROOT/lib -L ROOT/lib/libc -fuse-ld=lld'" # do not set `LEAN_CC` for tests echo -n " -DLEAN_TEST_VARS=''" diff --git a/script/prepare-llvm-mingw.sh b/script/prepare-llvm-mingw.sh index 0bd255f9edf3..1061630de0b5 100644 --- a/script/prepare-llvm-mingw.sh +++ b/script/prepare-llvm-mingw.sh @@ -43,7 +43,7 @@ echo -n " -DCMAKE_C_COMPILER=$PWD/stage1/bin/clang.exe -DCMAKE_C_COMPILER_WORKS= echo -n " -DSTAGE0_CMAKE_C_COMPILER=clang -DSTAGE0_CMAKE_CXX_COMPILER=clang++" echo -n " -DLEAN_EXTRA_CXX_FLAGS='--sysroot $PWD/llvm -idirafter /clang64/include/'" echo -n " -DLEANC_INTERNAL_FLAGS='--sysroot ROOT -nostdinc -isystem ROOT/include/clang' -DLEANC_CC=ROOT/bin/clang.exe" -echo -n " -DLEANC_INTERNAL_LINKER_FLAGS='-L ROOT/lib -Wl,-Bstatic -lgmp $(pkg-config --static --libs libuv) -lunwind -Wl,-Bdynamic -fuse-ld=lld --ld-path=ROOT/bin/ld.lld'" +echo -n " -DLEANC_INTERNAL_LINKER_FLAGS='--sysroot ROOT -L ROOT/lib -Wl,-Bstatic -lgmp $(pkg-config --static --libs libuv) -lunwind -Wl,-Bdynamic -fuse-ld=lld'" # when not using the above flags, link GMP dynamically/as usual. Always link ICU dynamically. echo -n " -DLEAN_EXTRA_LINKER_FLAGS='-lgmp $(pkg-config --libs libuv) -lucrtbase'" # do not set `LEAN_CC` for tests From 034bc2674029e53c282efd2cb220221480565c8c Mon Sep 17 00:00:00 2001 From: Sebastian Ullrich Date: Wed, 8 Jan 2025 13:04:31 +0000 Subject: [PATCH 054/100] feat: make `classical` tactic incremental (#6575) This PR ensures tactics are evaluated incrementally in the body of `classical`. --- src/Lean/Elab/Tactic/Classical.lean | 9 +++------ 1 file changed, 3 insertions(+), 6 deletions(-) diff --git a/src/Lean/Elab/Tactic/Classical.lean b/src/Lean/Elab/Tactic/Classical.lean index 787310edc0a5..1a1aeacf6c8e 100644 --- a/src/Lean/Elab/Tactic/Classical.lean +++ b/src/Lean/Elab/Tactic/Classical.lean @@ -24,11 +24,8 @@ def classical [Monad m] [MonadEnv m] [MonadFinally m] [MonadLiftT MetaM m] (t : finally modifyEnv Meta.instanceExtension.popScope -@[builtin_tactic Lean.Parser.Tactic.classical] -def evalClassical : Tactic := fun stx => do - match stx with - | `(tactic| classical $tacs:tacticSeq) => - classical <| Elab.Tactic.evalTactic tacs - | _ => throwUnsupportedSyntax +@[builtin_tactic Lean.Parser.Tactic.classical, builtin_incremental] +def evalClassical : Tactic := fun stx => + classical <| Term.withNarrowedArgTacticReuse (argIdx := 1) Elab.Tactic.evalTactic stx end Lean.Elab.Tactic From 5be241cba04e24491158e250926e06410cbbc398 Mon Sep 17 00:00:00 2001 From: Leonardo de Moura Date: Wed, 8 Jan 2025 10:03:31 -0800 Subject: [PATCH 055/100] fix: forall propagation in `grind` (#6578) This PR fixes and improves the propagator for forall-expressions in the `grind` tactic. --------- Co-authored-by: Kim Morrison --- src/Init/Grind/Lemmas.lean | 2 + src/Lean/Meta/Tactic/Grind/ForallProp.lean | 20 ++++++-- src/Lean/Meta/Tactic/Grind/Main.lean | 4 +- tests/lean/run/grind_cat.lean | 57 ++++++++++++++++++++++ 4 files changed, 76 insertions(+), 7 deletions(-) diff --git a/src/Init/Grind/Lemmas.lean b/src/Init/Grind/Lemmas.lean index 948033cc6bb8..6707d1b16146 100644 --- a/src/Init/Grind/Lemmas.lean +++ b/src/Init/Grind/Lemmas.lean @@ -73,6 +73,8 @@ theorem forall_propagator (p : Prop) (q : p → Prop) (q' : Prop) (h₁ : p = Tr · intro h'; exact Eq.mp h₂ (h' (of_eq_true h₁)) · intro h'; intros; exact Eq.mpr h₂ h' +theorem of_forall_eq_false (α : Sort u) (p : α → Prop) (h : (∀ x : α, p x) = False) : ∃ x : α, ¬ p x := by simp_all + /-! dite -/ theorem dite_cond_eq_true' {α : Sort u} {c : Prop} {_ : Decidable c} {a : c → α} {b : ¬ c → α} {r : α} (h₁ : c = True) (h₂ : a (of_eq_true h₁) = r) : (dite c a b) = r := by simp [h₁, h₂] diff --git a/src/Lean/Meta/Tactic/Grind/ForallProp.lean b/src/Lean/Meta/Tactic/Grind/ForallProp.lean index e89697fa803e..04d49e88a622 100644 --- a/src/Lean/Meta/Tactic/Grind/ForallProp.lean +++ b/src/Lean/Meta/Tactic/Grind/ForallProp.lean @@ -46,11 +46,21 @@ where -- b = True → (a → b) = True pushEqTrue e <| mkApp3 (mkConst ``Grind.imp_eq_of_eq_true_right) a b (← mkEqTrueProof b) -def propagateImpliesDown (e : Expr) : GoalM Unit := do +def propagateForallPropDown (e : Expr) : GoalM Unit := do + let .forallE n a b bi := e | return () if (← isEqFalse e) then - let .forallE _ a b _ := e | return () - let h ← mkEqFalseProof e - pushEqTrue a <| mkApp3 (mkConst ``Grind.eq_true_of_imp_eq_false) a b h - pushEqFalse b <| mkApp3 (mkConst ``Grind.eq_false_of_imp_eq_false) a b h + if b.hasLooseBVars then + let α := a + let p := b + -- `e` is of the form `∀ x : α, p x` + -- Add fact `∃ x : α, ¬ p x` + let u ← getLevel α + let prop := mkApp2 (mkConst ``Exists [u]) α (mkLambda n bi α (mkNot p)) + let proof := mkApp3 (mkConst ``Grind.of_forall_eq_false [u]) α (mkLambda n bi α p) (← mkEqFalseProof e) + addNewFact proof prop (← getGeneration e) + else + let h ← mkEqFalseProof e + pushEqTrue a <| mkApp3 (mkConst ``Grind.eq_true_of_imp_eq_false) a b h + pushEqFalse b <| mkApp3 (mkConst ``Grind.eq_false_of_imp_eq_false) a b h end Lean.Meta.Grind diff --git a/src/Lean/Meta/Tactic/Grind/Main.lean b/src/Lean/Meta/Tactic/Grind/Main.lean index 715191ed68f2..7272df1328a6 100644 --- a/src/Lean/Meta/Tactic/Grind/Main.lean +++ b/src/Lean/Meta/Tactic/Grind/Main.lean @@ -30,7 +30,7 @@ def mkMethods (fallback : Fallback) : CoreM Methods := do if let some prop := builtinPropagators.up[declName]? then prop e propagateDown := fun e => do - propagateImpliesDown e + propagateForallPropDown e let .const declName _ := e.getAppFn | return () if let some prop := builtinPropagators.down[declName]? then prop e @@ -70,7 +70,7 @@ def all (goals : List Goal) (f : Goal → GrindM (List Goal)) : GrindM (List Goa /-- A very simple strategy -/ private def simple (goals : List Goal) : GrindM (List Goal) := do - applyToAll (ematchStar >> (splitNext >> ematchStar).iterate) goals + applyToAll (assertAll >> ematchStar >> (splitNext >> assertAll >> ematchStar).iterate) goals def main (mvarId : MVarId) (config : Grind.Config) (mainDeclName : Name) (fallback : Fallback) : MetaM (List MVarId) := do let go : GrindM (List MVarId) := do diff --git a/tests/lean/run/grind_cat.lean b/tests/lean/run/grind_cat.lean index 9e501f97f97b..bc4101848b55 100644 --- a/tests/lean/run/grind_cat.lean +++ b/tests/lean/run/grind_cat.lean @@ -122,4 +122,61 @@ def hcomp {H I : Functor D E} (α : F ⟶ G) (β : H ⟶ I) : F.comp H ⟶ G.com -- rw [Functor.comp_map, Functor.comp_map, ← assoc, naturality, assoc, ← I.map_comp, naturality, -- map_comp, assoc] +structure Iso {C : Type u} [Category.{v} C] (X Y : C) where + hom : X ⟶ Y + inv : Y ⟶ X + hom_inv_id : hom ≫ inv = 𝟙 X := by cat_tac + inv_hom_id : inv ≫ hom = 𝟙 Y := by cat_tac + +attribute [grind =] Iso.hom_inv_id Iso.inv_hom_id + +/-- Notation for an isomorphism in a category. -/ +infixr:10 " ≅ " => Iso -- type as \cong or \iso + +variable {C : Type u} [Category.{v} C] {X Y Z : C} + +namespace Iso + +@[ext] +theorem ext ⦃α β : X ≅ Y⦄ (w : α.hom = β.hom) : α = β := + suffices α.inv = β.inv by + cases α + cases β + cases w + cases this + rfl + calc + α.inv = α.inv ≫ β.hom ≫ β.inv := by grind + _ = β.inv := by grind + + +/-- `LeftInverse g f` means that g is a left inverse to f. That is, `g ∘ f = id`. -/ +def Function.LeftInverse (g : β → α) (f : α → β) : Prop := + ∀ x, g (f x) = x + +/-- `RightInverse g f` means that g is a right inverse to f. That is, `f ∘ g = id`. -/ +def Function.RightInverse (g : β → α) (f : α → β) : Prop := + LeftInverse f g + +open Function + +/-- `α ≃ β` is the type of functions from `α → β` with a two-sided inverse. -/ +structure Equiv (α : Sort _) (β : Sort _) where + protected toFun : α → β + protected invFun : β → α + protected left_inv : LeftInverse invFun toFun + protected right_inv : RightInverse invFun toFun + +@[inherit_doc] +infixl:25 " ≃ " => Equiv + +attribute [local grind] Function.LeftInverse in +/-- The bijection `(Z ⟶ X) ≃ (Z ⟶ Y)` induced by `α : X ≅ Y`. -/ +def homToEquiv (α : X ≅ Y) {Z : C} : (Z ⟶ X) ≃ (Z ⟶ Y) where + toFun f := f ≫ α.hom + invFun g := g ≫ α.inv + left_inv := by cat_tac + right_inv := sorry + +end Iso end CategoryTheory From ddd454c9c13c2e9e4c320c2e7acdc3df3becb18d Mon Sep 17 00:00:00 2001 From: Leonardo de Moura Date: Wed, 8 Jan 2025 12:52:21 -0800 Subject: [PATCH 056/100] feat: add `grind` configuration options to control case-splitting (#6581) This PR adds the following configuration options to `Grind.Config`: `splitIte`, `splitMatch`, and `splitIndPred`. --- src/Init/Grind/Tactics.lean | 8 ++++++ src/Lean/Meta/Tactic/Grind/Internalize.lean | 27 ++++++++++++--------- tests/lean/run/grind_match2.lean | 5 ++++ tests/lean/run/grind_split.lean | 8 ++++++ 4 files changed, 36 insertions(+), 12 deletions(-) diff --git a/src/Init/Grind/Tactics.lean b/src/Init/Grind/Tactics.lean index f4678adfc688..951d3e2c0c77 100644 --- a/src/Init/Grind/Tactics.lean +++ b/src/Init/Grind/Tactics.lean @@ -37,6 +37,14 @@ structure Config where instances : Nat := 1000 /-- If `matchEqs` is `true`, `grind` uses `match`-equations as E-matching theorems. -/ matchEqs : Bool := true + /-- If `splitMatch` is `true`, `grind` performs case-splitting on `match`-expressions during the search. -/ + splitMatch : Bool := true + /-- If `splitIte` is `true`, `grind` performs case-splitting on `if-then-else` expressions during the search. -/ + splitIte : Bool := true + /-- + If `splitIndPred` is `true`, `grind` performs case-splitting on inductive predicates. + Otherwise, it performs case-splitting only on types marked with `[grind_split]` attribute. -/ + splitIndPred : Bool := true deriving Inhabited, BEq end Lean.Grind diff --git a/src/Lean/Meta/Tactic/Grind/Internalize.lean b/src/Lean/Meta/Tactic/Grind/Internalize.lean index e6f76fe7d9db..980aca39772f 100644 --- a/src/Lean/Meta/Tactic/Grind/Internalize.lean +++ b/src/Lean/Meta/Tactic/Grind/Internalize.lean @@ -56,22 +56,25 @@ private def forbiddenSplitTypes := [``Eq, ``HEq, ``True, ``False] /-- Inserts `e` into the list of case-split candidates if applicable. -/ private def checkAndAddSplitCandidate (e : Expr) : GoalM Unit := do unless e.isApp do return () - if e.isIte || e.isDIte then + if (← getConfig).splitIte && (e.isIte || e.isDIte) then addSplitCandidate e - else if (← isMatcherApp e) then - if let .reduced _ ← reduceMatcher? e then - -- When instantiating `match`-equations, we add `match`-applications that can be reduced, - -- and consequently don't need to be splitted. - return () - else - addSplitCandidate e - else - let .const declName _ := e.getAppFn | return () + return () + if (← getConfig).splitMatch then + if (← isMatcherApp e) then + if let .reduced _ ← reduceMatcher? e then + -- When instantiating `match`-equations, we add `match`-applications that can be reduced, + -- and consequently don't need to be splitted. + return () + else + addSplitCandidate e + return () + let .const declName _ := e.getAppFn | return () if forbiddenSplitTypes.contains declName then return () -- We should have a mechanism for letting users to select types to case-split. -- Right now, we just consider inductive predicates that are not in the forbidden list - if (← isInductivePredicate declName) then - addSplitCandidate e + if (← getConfig).splitIndPred then + if (← isInductivePredicate declName) then + addSplitCandidate e /-- If `e` is a `cast`-like term (e.g., `cast h a`), add `HEq e a` to the to-do list. diff --git a/tests/lean/run/grind_match2.lean b/tests/lean/run/grind_match2.lean index 24160dd19545..410e5c8819f5 100644 --- a/tests/lean/run/grind_match2.lean +++ b/tests/lean/run/grind_match2.lean @@ -26,3 +26,8 @@ set_option trace.grind true in example : h as ≠ 0 := by unfold h grind + +example : h as ≠ 0 := by + unfold h + fail_if_success grind (splitMatch := false) + sorry diff --git a/tests/lean/run/grind_split.lean b/tests/lean/run/grind_split.lean index 7acf78263d59..edf944cfae74 100644 --- a/tests/lean/run/grind_split.lean +++ b/tests/lean/run/grind_split.lean @@ -28,6 +28,10 @@ set_option trace.grind true in example (p : Prop) [Decidable p] (a b c : Nat) : (if p then a else b) = c → R a → R b → R c := by grind +example (p : Prop) [Decidable p] (a b c : Nat) : (if p then a else b) = c → R a → R b → R c := by + fail_if_success grind (splitIte := false) + sorry + namespace grind_test_induct_pred inductive Expr where @@ -52,4 +56,8 @@ set_option trace.grind true theorem HasType.det (h₁ : HasType e t₁) (h₂ : HasType e t₂) : t₁ = t₂ := by grind +theorem HasType.det' (h₁ : HasType e t₁) (h₂ : HasType e t₂) : t₁ = t₂ := by + fail_if_success grind (splitIndPred := false) + sorry + end grind_test_induct_pred From 0afa1d1e5dda3cfedd4a847421271fef64e9743c Mon Sep 17 00:00:00 2001 From: Leonardo de Moura Date: Wed, 8 Jan 2025 13:37:29 -0800 Subject: [PATCH 057/100] feat: apply E-matching for local lemmas in `grind` (#6582) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR adds support for creating local E-matching theorems for universal propositions known to be true. It allows `grind` to automatically solve examples such as: ```lean example (b : List α) (p : α → Prop) (h₁ : ∀ a ∈ b, p a) (h₂ : ∃ a ∈ b, ¬p a) : False := by grind ``` --- src/Lean/Meta/Tactic/Grind/EMatchTheorem.lean | 9 +++-- src/Lean/Meta/Tactic/Grind/ForallProp.lean | 29 +++++++++++++++ src/Lean/Meta/Tactic/Grind/Internalize.lean | 2 +- src/Lean/Meta/Tactic/Grind/Types.lean | 2 + tests/lean/run/grind_t1.lean | 37 +++++++++++++++++++ 5 files changed, 74 insertions(+), 5 deletions(-) diff --git a/src/Lean/Meta/Tactic/Grind/EMatchTheorem.lean b/src/Lean/Meta/Tactic/Grind/EMatchTheorem.lean index 96e6b0690186..1170a8dffa80 100644 --- a/src/Lean/Meta/Tactic/Grind/EMatchTheorem.lean +++ b/src/Lean/Meta/Tactic/Grind/EMatchTheorem.lean @@ -57,7 +57,8 @@ inductive Origin where is the provided grind argument. The `id` is a unique identifier for the call. -/ | stx (id : Name) (ref : Syntax) - | other (id : Name) + /-- It is local, but we don't have a local hypothesis for it. -/ + | local (id : Name) deriving Inhabited, Repr, BEq /-- A unique identifier corresponding to the origin. -/ @@ -65,14 +66,14 @@ def Origin.key : Origin → Name | .decl declName => declName | .fvar fvarId => fvarId.name | .stx id _ => id - | .other id => id + | .local id => id def Origin.pp [Monad m] [MonadEnv m] [MonadError m] (o : Origin) : m MessageData := do match o with | .decl declName => return MessageData.ofConst (← mkConstWithLevelParams declName) | .fvar fvarId => return mkFVar fvarId | .stx _ ref => return ref - | .other id => return id + | .local id => return id instance : BEq Origin where beq a b := a.key == b.key @@ -598,7 +599,7 @@ private def collectPatterns? (proof : Expr) (xs : Array Expr) (searchPlaces : Ar | return none return some (ps, s.symbols.toList) -private def mkEMatchTheoremWithKind? (origin : Origin) (levelParams : Array Name) (proof : Expr) (kind : TheoremKind) : MetaM (Option EMatchTheorem) := do +def mkEMatchTheoremWithKind? (origin : Origin) (levelParams : Array Name) (proof : Expr) (kind : TheoremKind) : MetaM (Option EMatchTheorem) := do if kind == .eqLhs then return (← mkEMatchEqTheoremCore origin levelParams proof (normalizePattern := false) (useLhs := true)) else if kind == .eqRhs then diff --git a/src/Lean/Meta/Tactic/Grind/ForallProp.lean b/src/Lean/Meta/Tactic/Grind/ForallProp.lean index 04d49e88a622..b8be9355c7f1 100644 --- a/src/Lean/Meta/Tactic/Grind/ForallProp.lean +++ b/src/Lean/Meta/Tactic/Grind/ForallProp.lean @@ -46,6 +46,32 @@ where -- b = True → (a → b) = True pushEqTrue e <| mkApp3 (mkConst ``Grind.imp_eq_of_eq_true_right) a b (← mkEqTrueProof b) +private def isEqTrueHyp? (proof : Expr) : Option FVarId := Id.run do + let_expr eq_true _ p := proof | return none + let .fvar fvarId := p | return none + return some fvarId + +private def addLocalEMatchTheorems (e : Expr) : GoalM Unit := do + let proof ← mkEqTrueProof e + let origin ← if let some fvarId := isEqTrueHyp? proof then + pure <| .fvar fvarId + else + let idx ← modifyGet fun s => (s.nextThmIdx, { s with nextThmIdx := s.nextThmIdx + 1 }) + pure <| .local ((`local).appendIndexAfter idx) + let proof := mkApp2 (mkConst ``of_eq_true) e proof + let size := (← get).newThms.size + let gen ← getGeneration e + -- TODO: we should have a flag for collecting all unary patterns in a local theorem + if let some thm ← mkEMatchTheoremWithKind? origin #[] proof .fwd then + activateTheorem thm gen + if let some thm ← mkEMatchTheoremWithKind? origin #[] proof .bwd then + activateTheorem thm gen + if (← get).newThms.size == size then + if let some thm ← mkEMatchTheoremWithKind? origin #[] proof .default then + activateTheorem thm gen + if (← get).newThms.size == size then + trace[grind.issues] "failed to create E-match local theorem for{indentExpr e}" + def propagateForallPropDown (e : Expr) : GoalM Unit := do let .forallE n a b bi := e | return () if (← isEqFalse e) then @@ -62,5 +88,8 @@ def propagateForallPropDown (e : Expr) : GoalM Unit := do let h ← mkEqFalseProof e pushEqTrue a <| mkApp3 (mkConst ``Grind.eq_true_of_imp_eq_false) a b h pushEqFalse b <| mkApp3 (mkConst ``Grind.eq_false_of_imp_eq_false) a b h + else if (← isEqTrue e) then + if b.hasLooseBVars then + addLocalEMatchTheorems e end Lean.Meta.Grind diff --git a/src/Lean/Meta/Tactic/Grind/Internalize.lean b/src/Lean/Meta/Tactic/Grind/Internalize.lean index 980aca39772f..eaaf18f107e0 100644 --- a/src/Lean/Meta/Tactic/Grind/Internalize.lean +++ b/src/Lean/Meta/Tactic/Grind/Internalize.lean @@ -101,7 +101,7 @@ private partial def internalizePattern (pattern : Expr) (generation : Nat) : Goa else pattern.withApp fun f args => do return mkAppN f (← args.mapM (internalizePattern · generation)) -private partial def activateTheorem (thm : EMatchTheorem) (generation : Nat) : GoalM Unit := do +partial def activateTheorem (thm : EMatchTheorem) (generation : Nat) : GoalM Unit := do -- Recall that we use the proof as part of the key for a set of instances found so far. -- We don't want to use structural equality when comparing keys. let proof ← shareCommon thm.proof diff --git a/src/Lean/Meta/Tactic/Grind/Types.lean b/src/Lean/Meta/Tactic/Grind/Types.lean index 3665bece75a7..50e072162a63 100644 --- a/src/Lean/Meta/Tactic/Grind/Types.lean +++ b/src/Lean/Meta/Tactic/Grind/Types.lean @@ -393,6 +393,8 @@ structure Goal where numSplits : Nat := 0 /-- Case-splits that do not have to be performed anymore. -/ resolvedSplits : PHashSet ENodeKey := {} + /-- Next local E-match theorem idx. -/ + nextThmIdx : Nat := 0 deriving Inhabited def Goal.admit (goal : Goal) : MetaM Unit := diff --git a/tests/lean/run/grind_t1.lean b/tests/lean/run/grind_t1.lean index 9c75284afdf5..3bf2a21185e0 100644 --- a/tests/lean/run/grind_t1.lean +++ b/tests/lean/run/grind_t1.lean @@ -173,3 +173,40 @@ example (α : Type) (β : Type) (a₁ a₂ : α) (b₁ b₂ : β) (h₄ : b₁ = b₂) : HEq a₂ b₂ := by grind + +/-- +info: [grind.assert] ∀ (a : α), a ∈ b → p a +[grind.ematch.pattern] h₁: [@Membership.mem `[α] `[List α] `[List.instMembership] `[b] #1] +[grind.ematch.pattern] h₁: [p #1] +[grind.assert] w ∈ b +[grind.assert] ¬p w +[grind.ematch.instance] h₁: w ∈ b → p w +[grind.assert] w ∈ b → p w +-/ +#guard_msgs (info) in +set_option trace.grind.ematch.pattern true in +set_option trace.grind.ematch.instance true in +set_option trace.grind.assert true in +example (b : List α) (p : α → Prop) (h₁ : ∀ a ∈ b, p a) (h₂ : ∃ a ∈ b, ¬p a) : False := by + grind + +/-- +info: [grind.assert] ∀ (x : α), Q x → P x +[grind.ematch.pattern] h₁: [Q #1] +[grind.ematch.pattern] h₁: [P #1] +[grind.assert] ∀ (x : α), R x → False = P x +[grind.ematch.pattern] h₂: [R #1] +[grind.ematch.pattern] h₂: [P #1] +[grind.assert] Q a +[grind.assert] R a +[grind.ematch.instance] h₁: Q a → P a +[grind.ematch.instance] h₂: R a → False = P a +[grind.assert] Q a → P a +[grind.assert] R a → False = P a +-/ +#guard_msgs (info) in +set_option trace.grind.ematch.pattern true in +set_option trace.grind.ematch.instance true in +set_option trace.grind.assert true in +example (P Q R : α → Prop) (h₁ : ∀ x, Q x → P x) (h₂ : ∀ x, R x → False = (P x)) : Q a → R a → False := by + grind From c5314da28ec6ffb5836075a05b739043a657be8e Mon Sep 17 00:00:00 2001 From: Leonardo de Moura Date: Wed, 8 Jan 2025 17:32:49 -0800 Subject: [PATCH 058/100] feat: add helper theorems for handling offsets in `grind` (#6584) This PR adds helper theorems to implement offset constraints in grind. --- src/Init/Grind.lean | 1 + src/Init/Grind/Offset.lean | 132 +++++++++++++++++++++++++++++++++++++ 2 files changed, 133 insertions(+) create mode 100644 src/Init/Grind/Offset.lean diff --git a/src/Init/Grind.lean b/src/Init/Grind.lean index 30c0d8381cda..80814ee4e30d 100644 --- a/src/Init/Grind.lean +++ b/src/Init/Grind.lean @@ -10,3 +10,4 @@ import Init.Grind.Lemmas import Init.Grind.Cases import Init.Grind.Propagator import Init.Grind.Util +import Init.Grind.Offset diff --git a/src/Init/Grind/Offset.lean b/src/Init/Grind/Offset.lean new file mode 100644 index 000000000000..56626662488f --- /dev/null +++ b/src/Init/Grind/Offset.lean @@ -0,0 +1,132 @@ +/- +Copyright (c) 2025 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 Init.Core +import Init.Omega + +namespace Lean.Grind + +abbrev Var := Nat +abbrev Context := Lean.RArray Nat + +def fixedVar := 100000000 -- Any big number should work here + +def Var.denote (ctx : Context) (v : Var) : Nat := + bif v == fixedVar then 1 else ctx.get v + +structure Cnstr where + x : Var + y : Var + k : Nat := 0 + l : Bool := true + deriving Repr, BEq, Inhabited + +def Cnstr.denote (c : Cnstr) (ctx : Context) : Prop := + if c.l then + c.x.denote ctx + c.k ≤ c.y.denote ctx + else + c.x.denote ctx ≤ c.y.denote ctx + c.k + +def trivialCnstr : Cnstr := { x := 0, y := 0, k := 0, l := true } + +@[simp] theorem denote_trivial (ctx : Context) : trivialCnstr.denote ctx := by + simp [Cnstr.denote, trivialCnstr] + +def Cnstr.trans (c₁ c₂ : Cnstr) : Cnstr := + if c₁.y = c₂.x then + let { x, k := k₁, l := l₁, .. } := c₁ + let { y, k := k₂, l := l₂, .. } := c₂ + match l₁, l₂ with + | false, false => + { x, y, k := k₁ + k₂, l := false } + | false, true => + if k₁ < k₂ then + { x, y, k := k₂ - k₁, l := true } + else + { x, y, k := k₁ - k₂, l := false } + | true, false => + if k₁ < k₂ then + { x, y, k := k₂ - k₁, l := false } + else + { x, y, k := k₁ - k₂, l := true } + | true, true => + { x, y, k := k₁ + k₂, l := true } + else + trivialCnstr + +@[simp] theorem Cnstr.denote_trans_easy (ctx : Context) (c₁ c₂ : Cnstr) (h : c₁.y ≠ c₂.x) : (c₁.trans c₂).denote ctx := by + simp [*, Cnstr.trans] + +@[simp] theorem Cnstr.denote_trans (ctx : Context) (c₁ c₂ : Cnstr) : c₁.denote ctx → c₂.denote ctx → (c₁.trans c₂).denote ctx := by + by_cases c₁.y = c₂.x + case neg => simp [*] + simp [trans, *] + let { x, k := k₁, l := l₁, .. } := c₁ + let { y, k := k₂, l := l₂, .. } := c₂ + simp_all; split + · simp [denote]; omega + · split <;> simp [denote] <;> omega + · split <;> simp [denote] <;> omega + · simp [denote]; omega + +def Cnstr.isTrivial (c : Cnstr) : Bool := c.x == c.y && c.k == 0 + +theorem Cnstr.of_isTrivial (ctx : Context) (c : Cnstr) : c.isTrivial = true → c.denote ctx := by + cases c; simp [isTrivial]; intros; simp [*, denote] + +def Cnstr.isFalse (c : Cnstr) : Bool := c.x == c.y && c.k != 0 && c.l == true + +theorem Cnstr.of_isFalse (ctx : Context) {c : Cnstr} : c.isFalse = true → ¬c.denote ctx := by + cases c; simp [isFalse]; intros; simp [*, denote]; omega + +def Certificate := List Cnstr + +def Certificate.denote' (ctx : Context) (c₁ : Cnstr) (c₂ : Certificate) : Prop := + match c₂ with + | [] => c₁.denote ctx + | c::cs => c₁.denote ctx ∧ Certificate.denote' ctx c cs + +theorem Certificate.denote'_trans (ctx : Context) (c₁ c : Cnstr) (cs : Certificate) : c₁.denote ctx → denote' ctx c cs → denote' ctx (c₁.trans c) cs := by + induction cs + next => simp [denote', *]; apply Cnstr.denote_trans + next c cs ih => simp [denote']; intros; simp [*] + +def Certificate.trans' (c₁ : Cnstr) (c₂ : Certificate) : Cnstr := + match c₂ with + | [] => c₁ + | c::c₂ => trans' (c₁.trans c) c₂ + +@[simp] theorem Certificate.denote'_trans' (ctx : Context) (c₁ : Cnstr) (c₂ : Certificate) : denote' ctx c₁ c₂ → (trans' c₁ c₂).denote ctx := by + induction c₂ generalizing c₁ + next => intros; simp_all [trans', denote'] + next c cs ih => simp [denote']; intros; simp [trans']; apply ih; apply denote'_trans <;> assumption + +def Certificate.denote (ctx : Context) (c : Certificate) : Prop := + match c with + | [] => True + | c::cs => denote' ctx c cs + +def Certificate.trans (c : Certificate) : Cnstr := + match c with + | [] => trivialCnstr + | c::cs => trans' c cs + +theorem Certificate.denote_trans {ctx : Context} {c : Certificate} : c.denote ctx → c.trans.denote ctx := by + cases c <;> simp [*, trans, Certificate.denote] <;> intros <;> simp [*] + +def Certificate.isFalse (c : Certificate) : Bool := + c.trans.isFalse + +theorem Certificate.unsat (ctx : Context) (c : Certificate) : c.isFalse = true → ¬ c.denote ctx := by + simp [isFalse]; intro h₁ h₂ + have := Certificate.denote_trans h₂ + have := Cnstr.of_isFalse ctx h₁ + contradiction + +theorem Certificate.imp (ctx : Context) (c : Certificate) : c.denote ctx → c.trans.denote ctx := by + apply denote_trans + +end Lean.Grind From cb9f198f01d38024ce92aec1c7eb7159299d6f78 Mon Sep 17 00:00:00 2001 From: Leonardo de Moura Date: Wed, 8 Jan 2025 18:23:46 -0800 Subject: [PATCH 059/100] fix: `grind` canonicalizer (#6585) This PR fixes a bug in the `grind` canonicalizer. --- src/Lean/Meta/Tactic/Grind.lean | 2 +- src/Lean/Meta/Tactic/Grind/Canon.lean | 17 ++++++++++++++--- tests/lean/run/grind_canon_bug.lean | 13 +++++++++++++ 3 files changed, 28 insertions(+), 4 deletions(-) create mode 100644 tests/lean/run/grind_canon_bug.lean diff --git a/src/Lean/Meta/Tactic/Grind.lean b/src/Lean/Meta/Tactic/Grind.lean index 80edeeef6cac..47f4b4911817 100644 --- a/src/Lean/Meta/Tactic/Grind.lean +++ b/src/Lean/Meta/Tactic/Grind.lean @@ -53,5 +53,5 @@ builtin_initialize registerTraceClass `grind.debug.parent builtin_initialize registerTraceClass `grind.debug.final builtin_initialize registerTraceClass `grind.debug.forallPropagator builtin_initialize registerTraceClass `grind.debug.split - +builtin_initialize registerTraceClass `grind.debug.canon end Lean diff --git a/src/Lean/Meta/Tactic/Grind/Canon.lean b/src/Lean/Meta/Tactic/Grind/Canon.lean index 6ac02c336ebd..17c3f43fc6de 100644 --- a/src/Lean/Meta/Tactic/Grind/Canon.lean +++ b/src/Lean/Meta/Tactic/Grind/Canon.lean @@ -66,11 +66,13 @@ def canonElemCore (f : Expr) (i : Nat) (e : Expr) (isInst : Bool) : StateT State -- in general safe to replace `e` with `c` if `c` has more free variables than `e`. -- However, we don't revert previously canonicalized elements in the `grind` tactic. modify fun s => { s with canon := s.canon.insert e c } + trace[grind.debug.canon] "found {e} ===> {c}" return c if isInst then - if (← isDiagnosticsEnabled <&&> pure (c.fvarsSubset e) <&&> (withDefault <| isDefEq e c)) then + if (← isDiagnosticsEnabled <&&> (withDefault <| isDefEq e c)) then -- TODO: consider storing this information in some structure that can be browsed later. trace[grind.issues] "the following `grind` static elements are definitionally equal with `default` transparency, but not with `instances` transparency{indentExpr e}\nand{indentExpr c}" + trace[grind.debug.canon] "({f}, {i}) ↦ {e}" modify fun s => { s with canon := s.canon.insert e e, argMap := s.argMap.insert key (e::cs) } return e @@ -92,6 +94,12 @@ private inductive ShouldCanonResult where visit deriving Inhabited +instance : Repr ShouldCanonResult where + reprPrec r _ := match r with + | .canonType => "canonType" + | .canonInst => "canonInst" + | .visit => "visit" + /-- See comments at `ShouldCanonResult`. -/ @@ -102,7 +110,9 @@ def shouldCanon (pinfos : Array ParamInfo) (i : Nat) (arg : Expr) : MetaM Should return .canonInst else if pinfo.isProp then return .visit - if (← isTypeFormer arg) then + if (← isProp arg) then + return .visit + else if (← isTypeFormer arg) then return .canonType else return .visit @@ -151,7 +161,8 @@ where return e' /-- Canonicalizes nested types, type formers, and instances in `e`. -/ -def canon (e : Expr) : StateT State MetaM Expr := +def canon (e : Expr) : StateT State MetaM Expr := do + trace[grind.debug.canon] "{e}" unsafe canonImpl e end Canon diff --git a/tests/lean/run/grind_canon_bug.lean b/tests/lean/run/grind_canon_bug.lean new file mode 100644 index 000000000000..1605f7a89213 --- /dev/null +++ b/tests/lean/run/grind_canon_bug.lean @@ -0,0 +1,13 @@ +open List + +attribute [grind =] getElem_cons_zero +attribute [grind =] getElem?_cons_zero + +example (h : (a :: t)[0]? = some b) : (a :: t)[0] = b := by + grind -- ✓ + +example [Inhabited α] : ∀ {l : List α} {i : Nat}, l[i]? = some b → l[i]! = b + | a::t, 0, _ => by + rw [getElem!_pos _ _ sorry] + grind + | _::l, _+1, e => sorry From 623dec1047d855a1c0659d0c44c016aeb823ab7d Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Thu, 9 Jan 2025 13:27:20 +1100 Subject: [PATCH 060/100] feat: aligning `List/Array/Vector` lemmas for `map` (#6586) This PR continues aligning `List/Array/Vector` lemmas, finishing up lemmas about `map`. --- src/Init/Data/Array/Lemmas.lean | 94 ++++++++++++++++------ src/Init/Data/List/Lemmas.lean | 7 +- src/Init/Data/Vector/Basic.lean | 2 +- src/Init/Data/Vector/Lemmas.lean | 133 +++++++++++++++++++++++++++++-- 4 files changed, 203 insertions(+), 33 deletions(-) diff --git a/src/Init/Data/Array/Lemmas.lean b/src/Init/Data/Array/Lemmas.lean index 8bb29aae6fae..a93201ab419b 100644 --- a/src/Init/Data/Array/Lemmas.lean +++ b/src/Init/Data/Array/Lemmas.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2022 Mario Carneiro. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Mario Carneiro +Authors: Mario Carneiro, Kim Morrison -/ prelude import Init.Data.Nat.Lemmas @@ -1001,8 +1001,6 @@ private theorem beq_of_beq_singleton [BEq α] {a b : α} : #[a] == #[b] → a == cases l₂ simp -/-! Content below this point has not yet been aligned with `List`. -/ - /-! ### map -/ theorem mapM_eq_foldlM [Monad m] [LawfulMonad m] (f : α → m β) (arr : Array α) : @@ -1036,11 +1034,6 @@ where simp only [← length_toList] simp -@[simp] theorem mapM_empty [Monad m] (f : α → m β) : mapM f #[] = pure #[] := by - rw [mapM, mapM.map]; rfl - -@[simp] theorem map_empty (f : α → β) : map f #[] = #[] := mapM_empty f - @[simp] theorem getElem_map (f : α → β) (a : Array α) (i : Nat) (hi : i < (a.map f).size) : (a.map f)[i] = f (a[i]'(by simpa using hi)) := by cases a @@ -1050,6 +1043,18 @@ where (as.map f)[i]? = as[i]?.map f := by simp [getElem?_def] +@[simp] theorem mapM_empty [Monad m] (f : α → m β) : mapM f #[] = pure #[] := by + rw [mapM, mapM.map]; rfl + +@[simp] theorem map_empty (f : α → β) : map f #[] = #[] := mapM_empty f + +@[simp] theorem map_push {f : α → β} {as : Array α} {x : α} : + (as.push x).map f = (as.map f).push (f x) := by + ext + · simp + · simp only [getElem_map, getElem_push, size_map] + split <;> rfl + @[simp] theorem map_id_fun : map (id : α → α) = id := by funext l induction l <;> simp_all @@ -1100,17 +1105,59 @@ theorem map_inj : map f = map g ↔ f = g := by theorem eq_empty_of_map_eq_empty {f : α → β} {l : Array α} (h : map f l = #[]) : l = #[] := map_eq_empty_iff.mp h +theorem map_eq_push_iff {f : α → β} {l : Array α} {l₂ : Array β} {b : β} : + map f l = l₂.push b ↔ ∃ l₁ a, l = l₁.push a ∧ map f l₁ = l₂ ∧ f a = b := by + rcases l with ⟨l⟩ + rcases l₂ with ⟨l₂⟩ + simp only [List.map_toArray, List.push_toArray, mk.injEq, List.map_eq_append_iff] + constructor + · rintro ⟨l₁, l₂, rfl, rfl, h⟩ + simp only [List.map_eq_singleton_iff] at h + obtain ⟨a, rfl, rfl⟩ := h + refine ⟨l₁.toArray, a, by simp⟩ + · rintro ⟨⟨l₁⟩, a, h₁, h₂, rfl⟩ + refine ⟨l₁, [a], by simp_all⟩ + +@[simp] theorem map_eq_singleton_iff {f : α → β} {l : Array α} {b : β} : + map f l = #[b] ↔ ∃ a, l = #[a] ∧ f a = b := by + cases l + simp + +theorem map_eq_map_iff {f g : α → β} {l : Array α} : + map f l = map g l ↔ ∀ a ∈ l, f a = g a := by + cases l <;> simp_all + +theorem map_eq_iff : map f l = l' ↔ ∀ i : Nat, l'[i]? = l[i]?.map f := by + cases l + cases l' + simp [List.map_eq_iff] + +theorem map_eq_foldl (f : α → β) (l : Array α) : + map f l = foldl (fun bs a => bs.push (f a)) #[] l := by + simpa using mapM_eq_foldlM (m := Id) f l + +@[simp] theorem map_set {f : α → β} {l : Array α} {i : Nat} {h : i < l.size} {a : α} : + (l.set i a).map f = (l.map f).set i (f a) (by simpa using h) := by + cases l + simp + +@[simp] theorem map_setIfInBounds {f : α → β} {l : Array α} {i : Nat} {a : α} : + (l.setIfInBounds i a).map f = (l.map f).setIfInBounds i (f a) := by + cases l + simp + +@[simp] theorem map_pop {f : α → β} {l : Array α} : l.pop.map f = (l.map f).pop := by + cases l + simp + +@[simp] theorem back?_map {f : α → β} {l : Array α} : (l.map f).back? = l.back?.map f := by + cases l + simp + @[simp] theorem map_map {f : α → β} {g : β → γ} {as : Array α} : (as.map f).map g = as.map (g ∘ f) := by cases as; simp -@[simp] theorem map_push {f : α → β} {as : Array α} {x : α} : - (as.push x).map f = (as.map f).push (f x) := by - ext - · simp - · simp only [getElem_map, getElem_push, size_map] - split <;> rfl - theorem mapM_eq_mapM_toList [Monad m] [LawfulMonad m] (f : α → m β) (arr : Array α) : arr.mapM f = List.toArray <$> (arr.toList.mapM f) := by rw [mapM_eq_foldlM, ← foldlM_toList, ← List.foldrM_reverse] @@ -1123,6 +1170,7 @@ theorem mapM_eq_mapM_toList [Monad m] [LawfulMonad m] (f : α → m β) (arr : A toList <$> arr.mapM f = arr.toList.mapM f := by simp [mapM_eq_mapM_toList] +@[deprecated "Use `mapM_eq_foldlM` instead" (since := "2025-01-08")] theorem mapM_map_eq_foldl (as : Array α) (f : α → β) (i) : mapM.map (m := Id) f as i b = as.foldl (start := i) (fun r a => r.push (f a)) b := by unfold mapM.map @@ -1139,9 +1187,7 @@ theorem mapM_map_eq_foldl (as : Array α) (f : α → β) (i) : rfl termination_by as.size - i -theorem map_eq_foldl (as : Array α) (f : α → β) : - as.map f = as.foldl (fun r a => r.push (f a)) #[] := - mapM_map_eq_foldl _ _ _ +/-! Content below this point has not yet been aligned with `List`. -/ theorem singleton_eq_toArray_singleton (a : α) : #[a] = [a].toArray := rfl @@ -1687,12 +1733,6 @@ theorem foldr_hom (f : β₁ → β₂) (g₁ : α → β₁ → β₁) (g₂ : /-! ### map -/ -@[simp] theorem map_pop {f : α → β} {as : Array α} : - as.pop.map f = (as.map f).pop := by - ext - · simp - · simp only [getElem_map, getElem_pop, size_map] - @[deprecated "Use `toList_map` or `List.map_toArray` to characterize `Array.map`." (since := "2025-01-06")] theorem map_induction (as : Array α) (f : α → β) (motive : Nat → Prop) (h0 : motive 0) (p : Fin as.size → β → Prop) (hs : ∀ i, motive i.1 → p i (f as[i]) ∧ motive (i+1)) : @@ -2368,6 +2408,12 @@ theorem foldr_map' (g : α → β) (f : α → α → α) (f' : β → β → β | nil => simp | cons xs xss ih => simp [ih] +/-! ### mkArray -/ + +@[simp] theorem mem_mkArray (a : α) (n : Nat) : b ∈ mkArray n a ↔ n ≠ 0 ∧ b = a := by + rw [mkArray, mem_toArray] + simp + /-! ### reverse -/ @[simp] theorem mem_reverse {x : α} {as : Array α} : x ∈ as.reverse ↔ x ∈ as := by diff --git a/src/Init/Data/List/Lemmas.lean b/src/Init/Data/List/Lemmas.lean index e29c161a31f0..49333f6b8b90 100644 --- a/src/Init/Data/List/Lemmas.lean +++ b/src/Init/Data/List/Lemmas.lean @@ -1,7 +1,8 @@ /- Copyright (c) 2014 Parikshit Khanna. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Parikshit Khanna, Jeremy Avigad, Leonardo de Moura, Floris van Doorn, Mario Carneiro +Authors: Parikshit Khanna, Jeremy Avigad, Leonardo de Moura, Floris van Doorn, Mario Carneiro, + Kim Morrison -/ prelude import Init.Data.Bool @@ -1114,6 +1115,10 @@ theorem map_eq_cons_iff' {f : α → β} {l : List α} : @[deprecated map_eq_cons' (since := "2024-09-05")] abbrev map_eq_cons' := @map_eq_cons_iff' +@[simp] theorem map_eq_singleton_iff {f : α → β} {l : List α} {b : β} : + map f l = [b] ↔ ∃ a, l = [a] ∧ f a = b := by + simp [map_eq_cons_iff] + theorem map_eq_map_iff : map f l = map g l ↔ ∀ a ∈ l, f a = g a := by induction l <;> simp diff --git a/src/Init/Data/Vector/Basic.lean b/src/Init/Data/Vector/Basic.lean index 31e3e430bc69..53b9327d212a 100644 --- a/src/Init/Data/Vector/Basic.lean +++ b/src/Init/Data/Vector/Basic.lean @@ -103,7 +103,7 @@ of bounds. @[inline] def head [NeZero n] (v : Vector α n) := v[0]'(Nat.pos_of_neZero n) /-- Push an element `x` to the end of a vector. -/ -@[inline] def push (x : α) (v : Vector α n) : Vector α (n + 1) := +@[inline] def push (v : Vector α n) (x : α) : Vector α (n + 1) := ⟨v.toArray.push x, by simp⟩ /-- Remove the last element of a vector. -/ diff --git a/src/Init/Data/Vector/Lemmas.lean b/src/Init/Data/Vector/Lemmas.lean index 88ccbc6dcb63..f8fb2eaf2e0c 100644 --- a/src/Init/Data/Vector/Lemmas.lean +++ b/src/Init/Data/Vector/Lemmas.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2024 Shreyas Srinivas. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Shreyas Srinivas, Francois Dorais +Authors: Shreyas Srinivas, Francois Dorais, Kim Morrison -/ prelude import Init.Data.Vector.Basic @@ -153,6 +153,14 @@ theorem toArray_mk (a : Array α) (h : a.size = n) : (Vector.mk a h).toArray = a @[simp] theorem all_mk (p : α → Bool) (a : Array α) (h : a.size = n) : (Vector.mk a h).all p = a.all p := rfl +@[simp] theorem eq_mk : v = Vector.mk a h ↔ v.toArray = a := by + cases v + simp + +@[simp] theorem mk_eq : Vector.mk a h = v ↔ a = v.toArray := by + cases v + simp + /-! ### toArray lemmas -/ @[simp] theorem getElem_toArray {α n} (xs : Vector α n) (i : Nat) (h : i < xs.toArray.size) : @@ -1035,8 +1043,6 @@ theorem mem_setIfInBounds (v : Vector α n) (i : Nat) (hi : i < n) (a : α) : cases l₂ simp -/-! Content below this point has not yet been aligned with `List` and `Array`. -/ - /-! ### map -/ @[simp] theorem getElem_map (f : α → β) (a : Vector α n) (i : Nat) (hi : i < n) : @@ -1044,16 +1050,129 @@ theorem mem_setIfInBounds (v : Vector α n) (i : Nat) (hi : i < n) (a : α) : cases a simp -@[simp] theorem getElem_ofFn {α n} (f : Fin n → α) (i : Nat) (h : i < n) : - (Vector.ofFn f)[i] = f ⟨i, by simpa using h⟩ := by - simp [ofFn] - /-- The empty vector maps to the empty vector. -/ @[simp] theorem map_empty (f : α → β) : map f #v[] = #v[] := by rw [map, mk.injEq] exact Array.map_empty f +@[simp] theorem map_push {f : α → β} {as : Vector α n} {x : α} : + (as.push x).map f = (as.map f).push (f x) := by + cases as + simp + +@[simp] theorem map_id_fun : map (n := n) (id : α → α) = id := by + funext l + induction l <;> simp_all + +/-- `map_id_fun'` differs from `map_id_fun` by representing the identity function as a lambda, rather than `id`. -/ +@[simp] theorem map_id_fun' : map (n := n) (fun (a : α) => a) = id := map_id_fun + +-- This is not a `@[simp]` lemma because `map_id_fun` will apply. +theorem map_id (l : Vector α n) : map (id : α → α) l = l := by + cases l <;> simp_all + +/-- `map_id'` differs from `map_id` by representing the identity function as a lambda, rather than `id`. -/ +-- This is not a `@[simp]` lemma because `map_id_fun'` will apply. +theorem map_id' (l : Vector α n) : map (fun (a : α) => a) l = l := map_id l + +/-- Variant of `map_id`, with a side condition that the function is pointwise the identity. -/ +theorem map_id'' {f : α → α} (h : ∀ x, f x = x) (l : Vector α n) : map f l = l := by + simp [show f = id from funext h] + +theorem map_singleton (f : α → β) (a : α) : map f #v[a] = #v[f a] := rfl + +@[simp] theorem mem_map {f : α → β} {l : Vector α n} : b ∈ l.map f ↔ ∃ a, a ∈ l ∧ f a = b := by + cases l + simp + +theorem exists_of_mem_map (h : b ∈ map f l) : ∃ a, a ∈ l ∧ f a = b := mem_map.1 h + +theorem mem_map_of_mem (f : α → β) (h : a ∈ l) : f a ∈ map f l := mem_map.2 ⟨_, h, rfl⟩ + +theorem forall_mem_map {f : α → β} {l : Vector α n} {P : β → Prop} : + (∀ (i) (_ : i ∈ l.map f), P i) ↔ ∀ (j) (_ : j ∈ l), P (f j) := by + simp + +@[simp] theorem map_inj_left {f g : α → β} : map f l = map g l ↔ ∀ a ∈ l, f a = g a := by + cases l <;> simp_all + +theorem map_congr_left (h : ∀ a ∈ l, f a = g a) : map f l = map g l := + map_inj_left.2 h + +theorem map_inj [NeZero n] : map (n := n) f = map g ↔ f = g := by + constructor + · intro h + ext a + replace h := congrFun h (mkVector n a) + simp only [mkVector, map_mk, mk.injEq, Array.map_inj_left, Array.mem_mkArray, and_imp, + forall_eq_apply_imp_iff] at h + exact h (NeZero.ne n) + · intro h; subst h; rfl + +theorem map_eq_push_iff {f : α → β} {l : Vector α (n + 1)} {l₂ : Vector β n} {b : β} : + map f l = l₂.push b ↔ ∃ l₁ a, l = l₁.push a ∧ map f l₁ = l₂ ∧ f a = b := by + rcases l with ⟨l, h⟩ + rcases l₂ with ⟨l₂, rfl⟩ + simp only [map_mk, push_mk, mk.injEq, Array.map_eq_push_iff] + constructor + · rintro ⟨l₁, a, rfl, rfl, rfl⟩ + refine ⟨⟨l₁, by simp⟩, a, by simp⟩ + · rintro ⟨l₁, a, h₁, h₂, rfl⟩ + refine ⟨l₁.toArray, a, by simp_all⟩ + +@[simp] theorem map_eq_singleton_iff {f : α → β} {l : Vector α 1} {b : β} : + map f l = #v[b] ↔ ∃ a, l = #v[a] ∧ f a = b := by + cases l + simp + +theorem map_eq_map_iff {f g : α → β} {l : Vector α n} : + map f l = map g l ↔ ∀ a ∈ l, f a = g a := by + cases l <;> simp_all + +theorem map_eq_iff {f : α → β} {l : Vector α n} {l' : Vector β n} : + map f l = l' ↔ ∀ i (h : i < n), l'[i] = f l[i] := by + rcases l with ⟨l, rfl⟩ + rcases l' with ⟨l', h'⟩ + simp only [map_mk, eq_mk, Array.map_eq_iff, getElem_mk] + constructor + · intro w i h + simpa [h, h'] using w i + · intro w i + if h : i < l.size then + simpa [h, h'] using w i h + else + rw [getElem?_neg, getElem?_neg, Option.map_none'] <;> omega + +@[simp] theorem map_set {f : α → β} {l : Vector α n} {i : Nat} {h : i < n} {a : α} : + (l.set i a).map f = (l.map f).set i (f a) (by simpa using h) := by + cases l + simp + +@[simp] theorem map_setIfInBounds {f : α → β} {l : Vector α n} {i : Nat} {a : α} : + (l.setIfInBounds i a).map f = (l.map f).setIfInBounds i (f a) := by + cases l + simp + +@[simp] theorem map_pop {f : α → β} {l : Vector α n} : l.pop.map f = (l.map f).pop := by + cases l + simp + +@[simp] theorem back?_map {f : α → β} {l : Vector α n} : (l.map f).back? = l.back?.map f := by + cases l + simp + +@[simp] theorem map_map {f : α → β} {g : β → γ} {as : Vector α n} : + (as.map f).map g = as.map (g ∘ f) := by + cases as + simp + +/-! Content below this point has not yet been aligned with `List` and `Array`. -/ + +@[simp] theorem getElem_ofFn {α n} (f : Fin n → α) (i : Nat) (h : i < n) : + (Vector.ofFn f)[i] = f ⟨i, by simpa using h⟩ := by + simp [ofFn] + @[simp] theorem getElem_push_last {v : Vector α n} {x : α} : (v.push x)[n] = x := by rcases v with ⟨data, rfl⟩ simp From 827c6676fd339b147258c05fda75b3cf6dc0fd31 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Thu, 9 Jan 2025 15:15:47 +1100 Subject: [PATCH 061/100] feat: align `List/Array` lemmas for `filter/filterMap` (#6589) This PR continues aligning `List/Array` lemmas, finishing `filter` and `filterMap`. --- src/Init/Data/Array/Lemmas.lean | 396 +++++++++++++++++++++++++------- src/Init/Data/List/Lemmas.lean | 2 +- 2 files changed, 318 insertions(+), 80 deletions(-) diff --git a/src/Init/Data/Array/Lemmas.lean b/src/Init/Data/Array/Lemmas.lean index a93201ab419b..7348c35f5338 100644 --- a/src/Init/Data/Array/Lemmas.lean +++ b/src/Init/Data/Array/Lemmas.lean @@ -124,6 +124,8 @@ theorem push_eq_push {a b : α} {xs ys : Array α} : xs.push a = ys.push b ↔ a · rintro ⟨rfl, rfl⟩ rfl +theorem push_eq_append_singleton (as : Array α) (x) : as.push x = as ++ #[x] := rfl + theorem exists_push_of_ne_empty {xs : Array α} (h : xs ≠ #[]) : ∃ (ys : Array α) (a : α), xs = ys.push a := by rcases xs with ⟨xs⟩ @@ -1187,6 +1189,321 @@ theorem mapM_map_eq_foldl (as : Array α) (f : α → β) (i) : rfl termination_by as.size - i +/-! ### filter -/ + +@[congr] +theorem filter_congr {as bs : Array α} (h : as = bs) + {f : α → Bool} {g : α → Bool} (h' : f = g) {start stop start' stop' : Nat} + (h₁ : start = start') (h₂ : stop = stop') : + filter f as start stop = filter g bs start' stop' := by + congr + +@[simp] theorem toList_filter' (p : α → Bool) (l : Array α) (h : stop = l.size) : + (l.filter p 0 stop).toList = l.toList.filter p := by + subst h + dsimp only [filter] + rw [← foldl_toList] + generalize l.toList = l + suffices ∀ a, (List.foldl (fun r a => if p a = true then push r a else r) a l).toList = + a.toList ++ List.filter p l by + simpa using this #[] + induction l with simp + | cons => split <;> simp [*] + +theorem toList_filter (p : α → Bool) (l : Array α) : + (l.filter p).toList = l.toList.filter p := by + simp + +@[simp] theorem _root_.List.filter_toArray' (p : α → Bool) (l : List α) (h : stop = l.length) : + l.toArray.filter p 0 stop = (l.filter p).toArray := by + apply ext' + simp [h] + +theorem _root_.List.filter_toArray (p : α → Bool) (l : List α) : + l.toArray.filter p = (l.filter p).toArray := by + simp + +@[simp] theorem filter_push_of_pos {p : α → Bool} {a : α} {l : Array α} + (h : p a) (w : stop = l.size + 1): + (l.push a).filter p 0 stop = (l.filter p).push a := by + subst w + rcases l with ⟨l⟩ + simp [h] + +@[simp] theorem filter_push_of_neg {p : α → Bool} {a : α} {l : Array α} + (h : ¬p a) (w : stop = l.size + 1) : + (l.push a).filter p 0 stop = l.filter p := by + subst w + rcases l with ⟨l⟩ + simp [h] + +theorem filter_push {p : α → Bool} {a : α} {l : Array α} : + (l.push a).filter p = if p a then (l.filter p).push a else l.filter p := by + split <;> simp [*] + +theorem size_filter_le (p : α → Bool) (l : Array α) : + (l.filter p).size ≤ l.size := by + rcases l with ⟨l⟩ + simpa using List.length_filter_le p l + +@[simp] theorem filter_eq_self {p : α → Bool} {l : Array α} : + filter p l = l ↔ ∀ a ∈ l, p a := by + rcases l with ⟨l⟩ + simp + +@[simp] theorem filter_size_eq_size {p : α → Bool} {l : Array α} : + (filter p l).size = l.size ↔ ∀ a ∈ l, p a := by + rcases l with ⟨l⟩ + simp + +@[simp] theorem mem_filter {p : α → Bool} {l : Array α} {a : α} : + a ∈ filter p l ↔ a ∈ l ∧ p a := by + rcases l with ⟨l⟩ + simp + +@[simp] theorem filter_eq_empty_iff {p : α → Bool} {l : Array α} : + filter p l = #[] ↔ ∀ a, a ∈ l → ¬p a := by + rcases l with ⟨l⟩ + simp + +theorem forall_mem_filter {p : α → Bool} {l : Array α} {P : α → Prop} : + (∀ (i) (_ : i ∈ l.filter p), P i) ↔ ∀ (j) (_ : j ∈ l), p j → P j := by + simp + +@[simp] theorem filter_filter (q) (l : Array α) : + filter p (filter q l) = filter (fun a => p a && q a) l := by + apply ext' + simp only [toList_filter, List.filter_filter] + +theorem foldl_filter (p : α → Bool) (f : β → α → β) (l : Array α) (init : β) : + (l.filter p).foldl f init = l.foldl (fun x y => if p y then f x y else x) init := by + rcases l with ⟨l⟩ + rw [List.filter_toArray] + simp [List.foldl_filter] + +theorem foldr_filter (p : α → Bool) (f : α → β → β) (l : Array α) (init : β) : + (l.filter p).foldr f init = l.foldr (fun x y => if p x then f x y else y) init := by + rcases l with ⟨l⟩ + rw [List.filter_toArray] + simp [List.foldr_filter] + +theorem filter_map (f : β → α) (l : Array β) : filter p (map f l) = map f (filter (p ∘ f) l) := by + rcases l with ⟨l⟩ + simp [List.filter_map] + +theorem map_filter_eq_foldl (f : α → β) (p : α → Bool) (l : Array α) : + map f (filter p l) = foldl (fun y x => bif p x then y.push (f x) else y) #[] l := by + rcases l with ⟨l⟩ + apply ext' + simp only [size_toArray, List.filter_toArray', List.map_toArray, List.foldl_toArray'] + rw [← List.reverse_reverse l] + generalize l.reverse = l + simp only [List.filter_reverse, List.map_reverse, List.foldl_reverse] + induction l with + | nil => rfl + | cons x l ih => + simp only [List.filter_cons, List.foldr_cons] + split <;> simp_all + +@[simp] theorem filter_append {p : α → Bool} (l₁ l₂ : Array α) : + filter p (l₁ ++ l₂) = filter p l₁ ++ filter p l₂ := by + rcases l₁ with ⟨l₁⟩ + rcases l₂ with ⟨l₂⟩ + simp [List.filter_append] + +theorem filter_eq_append_iff {p : α → Bool} : + filter p l = L₁ ++ L₂ ↔ ∃ l₁ l₂, l = l₁ ++ l₂ ∧ filter p l₁ = L₁ ∧ filter p l₂ = L₂ := by + rcases l with ⟨l⟩ + rcases L₁ with ⟨L₁⟩ + rcases L₂ with ⟨L₂⟩ + simp only [size_toArray, List.filter_toArray', List.append_toArray, mk.injEq, + List.filter_eq_append_iff] + constructor + · rintro ⟨l₁, l₂, rfl, rfl, rfl⟩ + refine ⟨l₁.toArray, l₂.toArray, by simp⟩ + · rintro ⟨⟨l₁⟩, ⟨l₂⟩, h₁, h₂, h₃⟩ + refine ⟨l₁, l₂, by simp_all⟩ + +theorem filter_eq_push_iff {p : α → Bool} {l l' : Array α} {a : α} : + filter p l = l'.push a ↔ + ∃ l₁ l₂, l = l₁.push a ++ l₂ ∧ filter p l₁ = l' ∧ p a ∧ (∀ x, x ∈ l₂ → ¬p x) := by + rcases l with ⟨l⟩ + rcases l' with ⟨l'⟩ + simp only [size_toArray, List.filter_toArray', List.push_toArray, mk.injEq, Bool.not_eq_true] + rw [← List.reverse_inj] + simp only [← List.filter_reverse] + simp only [List.reverse_append, List.reverse_cons, List.reverse_nil, List.nil_append, + List.singleton_append] + rw [List.filter_eq_cons_iff] + constructor + · rintro ⟨l₁, l₂, h₁, h₂, h₃, h₄⟩ + refine ⟨l₂.reverse.toArray, l₁.reverse.toArray, by simp_all⟩ + · rintro ⟨⟨l₁⟩, ⟨l₂⟩, h₁, h₂, h₃, h₄⟩ + refine ⟨l₂.reverse, l₁.reverse, by simp_all⟩ + +@[deprecated filter_map (since := "2024-06-15")] abbrev map_filter := @filter_map + +theorem mem_of_mem_filter {a : α} {l} (h : a ∈ filter p l) : a ∈ l := + (mem_filter.mp h).1 + +/-! ### filterMap -/ + +@[congr] +theorem filterMap_congr {as bs : Array α} (h : as = bs) + {f : α → Option β} {g : α → Option β} (h' : f = g) {start stop start' stop' : Nat} + (h₁ : start = start') (h₂ : stop = stop') : + filterMap f as start stop = filterMap g bs start' stop' := by + congr + +@[simp] theorem toList_filterMap' (f : α → Option β) (l : Array α) (w : stop = l.size) : + (l.filterMap f 0 stop).toList = l.toList.filterMap f := by + subst w + dsimp only [filterMap, filterMapM] + rw [← foldlM_toList] + generalize l.toList = l + have this : ∀ a : Array β, (Id.run (List.foldlM (m := Id) ?_ a l)).toList = + a.toList ++ List.filterMap f l := ?_ + exact this #[] + induction l + · simp_all [Id.run] + · simp_all [Id.run, List.filterMap_cons] + split <;> simp_all + +theorem toList_filterMap (f : α → Option β) (l : Array α) : + (l.filterMap f).toList = l.toList.filterMap f := by + simp [toList_filterMap'] + + +@[simp] theorem _root_.List.filterMap_toArray' (f : α → Option β) (l : List α) (h : stop = l.length) : + l.toArray.filterMap f 0 stop = (l.filterMap f).toArray := by + apply ext' + simp [h] + +theorem _root_.List.filterMap_toArray (f : α → Option β) (l : List α) : + l.toArray.filterMap f = (l.filterMap f).toArray := by + simp + +@[simp] theorem filterMap_push_none {f : α → Option β} {a : α} {l : Array α} + (h : f a = none) (w : stop = l.size + 1) : + filterMap f (l.push a) 0 stop = filterMap f l := by + subst w + rcases l with ⟨l⟩ + simp [h] + +@[simp] theorem filterMap_push_some {f : α → Option β} {a : α} {l : Array α} {b : β} + (h : f a = some b) (w : stop = l.size + 1) : + filterMap f (l.push a) 0 stop = (filterMap f l).push b := by + subst w + rcases l with ⟨l⟩ + simp [h] + +@[simp] theorem filterMap_eq_map (f : α → β) : filterMap (some ∘ f) = map f := by + funext l + cases l + simp + +theorem filterMap_some_fun : filterMap (some : α → Option α) = id := by + funext l + cases l + simp + +@[simp] theorem filterMap_some (l : Array α) : filterMap some l = l := by + cases l + simp + +theorem map_filterMap_some_eq_filterMap_isSome (f : α → Option β) (l : Array α) : + (l.filterMap f).map some = (l.map f).filter fun b => b.isSome := by + cases l + simp [List.map_filterMap_some_eq_filter_map_isSome] + +theorem size_filterMap_le (f : α → Option β) (l : Array α) : + (filterMap f l).size ≤ l.size := by + cases l + simp [List.length_filterMap_le] + +@[simp] theorem filterMap_size_eq_size {l} : + (filterMap f l).size = l.size ↔ ∀ a, a ∈ l → (f a).isSome := by + cases l + simp [List.filterMap_length_eq_length] + +@[simp] +theorem filterMap_eq_filter (p : α → Bool) : + filterMap (Option.guard (p ·)) = filter p := by + funext l + cases l + simp + +theorem filterMap_filterMap (f : α → Option β) (g : β → Option γ) (l : Array α) : + filterMap g (filterMap f l) = filterMap (fun x => (f x).bind g) l := by + cases l + simp [List.filterMap_filterMap] + +theorem map_filterMap (f : α → Option β) (g : β → γ) (l : Array α) : + map g (filterMap f l) = filterMap (fun x => (f x).map g) l := by + cases l + simp [List.map_filterMap] + +@[simp] theorem filterMap_map (f : α → β) (g : β → Option γ) (l : Array α) : + filterMap g (map f l) = filterMap (g ∘ f) l := by + cases l + simp [List.filterMap_map] + +theorem filter_filterMap (f : α → Option β) (p : β → Bool) (l : Array α) : + filter p (filterMap f l) = filterMap (fun x => (f x).filter p) l := by + cases l + simp [List.filter_filterMap] + +theorem filterMap_filter (p : α → Bool) (f : α → Option β) (l : Array α) : + filterMap f (filter p l) = filterMap (fun x => if p x then f x else none) l := by + cases l + simp [List.filterMap_filter] + +@[simp] theorem mem_filterMap {f : α → Option β} {l : Array α} {b : β} : + b ∈ filterMap f l ↔ ∃ a, a ∈ l ∧ f a = some b := by + simp only [mem_def, toList_filterMap, List.mem_filterMap] + +theorem forall_mem_filterMap {f : α → Option β} {l : Array α} {P : β → Prop} : + (∀ (i) (_ : i ∈ filterMap f l), P i) ↔ ∀ (j) (_ : j ∈ l) (b), f j = some b → P b := by + simp only [mem_filterMap, forall_exists_index, and_imp] + rw [forall_comm] + apply forall_congr' + intro a + rw [forall_comm] + +@[simp] theorem filterMap_append {α β : Type _} (l l' : Array α) (f : α → Option β) : + filterMap f (l ++ l') = filterMap f l ++ filterMap f l' := by + cases l + cases l' + simp + +theorem map_filterMap_of_inv (f : α → Option β) (g : β → α) (H : ∀ x : α, (f x).map g = some x) + (l : Array α) : map g (filterMap f l) = l := by simp only [map_filterMap, H, filterMap_some, id] + +theorem forall_none_of_filterMap_eq_empty (h : filterMap f xs = #[]) : ∀ x ∈ xs, f x = none := by + cases xs + simpa using h + +@[simp] theorem filterMap_eq_nil_iff {l} : filterMap f l = #[] ↔ ∀ a, a ∈ l → f a = none := by + cases l + simp + +theorem filterMap_eq_push_iff {f : α → Option β} {l : Array α} {l' : Array β} {b : β} : + filterMap f l = l'.push b ↔ ∃ l₁ a l₂, + l = l₁.push a ++ l₂ ∧ filterMap f l₁ = l' ∧ f a = some b ∧ (∀ x, x ∈ l₂ → f x = none) := by + rcases l with ⟨l⟩ + rcases l' with ⟨l'⟩ + simp only [size_toArray, List.filterMap_toArray', List.push_toArray, mk.injEq] + rw [← List.reverse_inj] + simp only [← List.filterMap_reverse] + simp only [List.reverse_append, List.reverse_cons, List.reverse_nil, List.nil_append, + List.singleton_append] + rw [List.filterMap_eq_cons_iff] + constructor + · rintro ⟨l₁, a, l₂, h₁, h₂, h₃, h₄⟩ + refine ⟨l₂.reverse.toArray, a, l₁.reverse.toArray, by simp_all⟩ + · rintro ⟨⟨l₁⟩, a, ⟨l₂⟩, h₁, h₂, h₃, h₄⟩ + refine ⟨l₂.reverse, a, l₁.reverse, by simp_all⟩ + /-! Content below this point has not yet been aligned with `List`. -/ theorem singleton_eq_toArray_singleton (a : α) : #[a] = [a].toArray := rfl @@ -1801,71 +2118,12 @@ theorem getElem?_modify {as : Array α} {i : Nat} {f : α → α} {j : Nat} : simp only [getElem?_def, size_modify, getElem_modify, Option.map_dif] split <;> split <;> rfl -/-! ### filter -/ - -@[simp] theorem toList_filter (p : α → Bool) (l : Array α) : - (l.filter p).toList = l.toList.filter p := by - dsimp only [filter] - rw [← foldl_toList] - generalize l.toList = l - suffices ∀ a, (List.foldl (fun r a => if p a = true then push r a else r) a l).toList = - a.toList ++ List.filter p l by - simpa using this #[] - induction l with simp - | cons => split <;> simp [*] - -@[simp] theorem filter_filter (q) (l : Array α) : - filter p (filter q l) = filter (fun a => p a && q a) l := by - apply ext' - simp only [toList_filter, List.filter_filter] - -@[simp] theorem mem_filter : x ∈ filter p as ↔ x ∈ as ∧ p x := by - simp only [mem_def, toList_filter, List.mem_filter] - -theorem mem_of_mem_filter {a : α} {l} (h : a ∈ filter p l) : a ∈ l := - (mem_filter.mp h).1 - -@[congr] -theorem filter_congr {as bs : Array α} (h : as = bs) - {f : α → Bool} {g : α → Bool} (h' : f = g) {start stop start' stop' : Nat} - (h₁ : start = start') (h₂ : stop = stop') : - filter f as start stop = filter g bs start' stop' := by - congr - -/-! ### filterMap -/ - -@[simp] theorem toList_filterMap (f : α → Option β) (l : Array α) : - (l.filterMap f).toList = l.toList.filterMap f := by - dsimp only [filterMap, filterMapM] - rw [← foldlM_toList] - generalize l.toList = l - have this : ∀ a : Array β, (Id.run (List.foldlM (m := Id) ?_ a l)).toList = - a.toList ++ List.filterMap f l := ?_ - exact this #[] - induction l - · simp_all [Id.run] - · simp_all [Id.run, List.filterMap_cons] - split <;> simp_all - -@[simp] theorem mem_filterMap {f : α → Option β} {l : Array α} {b : β} : - b ∈ filterMap f l ↔ ∃ a, a ∈ l ∧ f a = some b := by - simp only [mem_def, toList_filterMap, List.mem_filterMap] - -@[congr] -theorem filterMap_congr {as bs : Array α} (h : as = bs) - {f : α → Option β} {g : α → Option β} (h' : f = g) {start stop start' stop' : Nat} - (h₁ : start = start') (h₂ : stop = stop') : - filterMap f as start stop = filterMap g bs start' stop' := by - congr - /-! ### empty -/ theorem size_empty : (#[] : Array α).size = 0 := rfl /-! ### append -/ -theorem push_eq_append_singleton (as : Array α) (x) : as.push x = as ++ #[x] := rfl - @[simp] theorem mem_append {a : α} {s t : Array α} : a ∈ s ++ t ↔ a ∈ s ∨ a ∈ t := by simp only [mem_def, toList_append, List.mem_append] @@ -2267,26 +2525,6 @@ theorem uset_toArray (l : List α) (i : USize) (a : α) (h : i.toNat < l.toArray apply ext' simp -@[simp] theorem filter_toArray' (p : α → Bool) (l : List α) (h : stop = l.toArray.size) : - l.toArray.filter p 0 stop = (l.filter p).toArray := by - subst h - apply ext' - rw [toList_filter] - -@[simp] theorem filterMap_toArray' (f : α → Option β) (l : List α) (h : stop = l.toArray.size) : - l.toArray.filterMap f 0 stop = (l.filterMap f).toArray := by - subst h - apply ext' - rw [toList_filterMap] - -theorem filter_toArray (p : α → Bool) (l : List α) : - l.toArray.filter p = (l.filter p).toArray := by - simp - -theorem filterMap_toArray (f : α → Option β) (l : List α) : - l.toArray.filterMap f = (l.filterMap f).toArray := by - simp - @[simp] theorem flatten_toArray (l : List (List α)) : (l.toArray.map List.toArray).flatten = l.flatten.toArray := by apply ext' diff --git a/src/Init/Data/List/Lemmas.lean b/src/Init/Data/List/Lemmas.lean index 49333f6b8b90..934456403a80 100644 --- a/src/Init/Data/List/Lemmas.lean +++ b/src/Init/Data/List/Lemmas.lean @@ -1285,7 +1285,7 @@ theorem map_filter_eq_foldr (f : α → β) (p : α → Bool) (as : List α) : @[simp] theorem filter_append {p : α → Bool} : ∀ (l₁ l₂ : List α), filter p (l₁ ++ l₂) = filter p l₁ ++ filter p l₂ | [], _ => rfl - | a :: l₁, l₂ => by simp [filter]; split <;> simp [filter_append l₁] + | a :: l₁, l₂ => by simp only [cons_append, filter]; split <;> simp [filter_append l₁] theorem filter_eq_cons_iff {l} {a} {as} : filter p l = a :: as ↔ From dd6445515ddc71826c76b8020fe9e030579b432b Mon Sep 17 00:00:00 2001 From: Leonardo de Moura Date: Wed, 8 Jan 2025 22:21:42 -0800 Subject: [PATCH 062/100] feat: improve `grind` canonicalizer diagnostics (#6588) This PR improves the `grind` canonicalizer diagnostics. --------- Co-authored-by: Kim Morrison --- src/Lean/Meta/Tactic/Grind/Canon.lean | 49 +++- tests/lean/run/grind_cat2.lean | 311 ++++++++++++++++++++++++++ 2 files changed, 351 insertions(+), 9 deletions(-) create mode 100644 tests/lean/run/grind_cat2.lean diff --git a/src/Lean/Meta/Tactic/Grind/Canon.lean b/src/Lean/Meta/Tactic/Grind/Canon.lean index 17c3f43fc6de..2c9ac388a816 100644 --- a/src/Lean/Meta/Tactic/Grind/Canon.lean +++ b/src/Lean/Meta/Tactic/Grind/Canon.lean @@ -46,15 +46,35 @@ structure State where proofCanon : PHashMap Expr Expr := {} deriving Inhabited +inductive CanonElemKind where + | /-- + Type class instances are canonicalized using `TransparencyMode.instances`. + -/ + instance + | /-- + Types and Type formers are canonicalized using `TransparencyMode.default`. + Remark: propositions are just visited. We do not invoke `canonElemCore` for them. + -/ + type + | /-- + Implicit arguments that are not types, type formers, or instances, are canonicalized + using `TransparencyMode.reducible` + -/ + implicit + deriving BEq + +def CanonElemKind.explain : CanonElemKind → String + | .instance => "type class instances" + | .type => "types (or type formers)" + | .implicit => "implicit arguments (which are not type class instances or types)" + /-- Helper function for canonicalizing `e` occurring as the `i`th argument of an `f`-application. -`isInst` is true if `e` is an type class instance. -Recall that we use `TransparencyMode.instances` for checking whether two instances are definitionally equal or not. -Thus, if diagnostics are enabled, we also check them using `TransparencyMode.default`. If the result is different +Thus, if diagnostics are enabled, we also re-check them using `TransparencyMode.default`. If the result is different we report to the user. -/ -def canonElemCore (f : Expr) (i : Nat) (e : Expr) (isInst : Bool) : StateT State MetaM Expr := do +def canonElemCore (f : Expr) (i : Nat) (e : Expr) (kind : CanonElemKind) : StateT State MetaM Expr := do let s ← get if let some c := s.canon.find? e then return c @@ -68,16 +88,17 @@ def canonElemCore (f : Expr) (i : Nat) (e : Expr) (isInst : Bool) : StateT State modify fun s => { s with canon := s.canon.insert e c } trace[grind.debug.canon] "found {e} ===> {c}" return c - if isInst then - if (← isDiagnosticsEnabled <&&> (withDefault <| isDefEq e c)) then + if kind != .type then + if (← isTracingEnabledFor `grind.issues <&&> (withDefault <| isDefEq e c)) then -- TODO: consider storing this information in some structure that can be browsed later. - trace[grind.issues] "the following `grind` static elements are definitionally equal with `default` transparency, but not with `instances` transparency{indentExpr e}\nand{indentExpr c}" + trace[grind.issues] "the following {kind.explain} are definitionally equal with `default` transparency but not with a more restrictive transparency{indentExpr e}\nand{indentExpr c}" trace[grind.debug.canon] "({f}, {i}) ↦ {e}" modify fun s => { s with canon := s.canon.insert e e, argMap := s.argMap.insert key (e::cs) } return e -abbrev canonType (f : Expr) (i : Nat) (e : Expr) := withDefault <| canonElemCore f i e false -abbrev canonInst (f : Expr) (i : Nat) (e : Expr) := withReducibleAndInstances <| canonElemCore f i e true +abbrev canonType (f : Expr) (i : Nat) (e : Expr) := withDefault <| canonElemCore f i e .type +abbrev canonInst (f : Expr) (i : Nat) (e : Expr) := withReducibleAndInstances <| canonElemCore f i e .instance +abbrev canonImplicit (f : Expr) (i : Nat) (e : Expr) := withReducible <| canonElemCore f i e .implicit /-- Return type for the `shouldCanon` function. @@ -87,6 +108,8 @@ private inductive ShouldCanonResult where canonType | /- Nested instances are canonicalized. -/ canonInst + | /- Implicit argument that is not an instance nor a type. -/ + canonImplicit | /- Term is not a proof, type (former), nor an instance. Thus, it must be recursively visited by the canonizer. @@ -98,6 +121,7 @@ instance : Repr ShouldCanonResult where reprPrec r _ := match r with | .canonType => "canonType" | .canonInst => "canonInst" + | .canonImplicit => "canonImplicit" | .visit => "visit" /-- @@ -110,6 +134,11 @@ def shouldCanon (pinfos : Array ParamInfo) (i : Nat) (arg : Expr) : MetaM Should return .canonInst else if pinfo.isProp then return .visit + else if pinfo.isImplicit then + if (← isTypeFormer arg) then + return .canonType + else + return .canonImplicit if (← isProp arg) then return .visit else if (← isTypeFormer arg) then @@ -142,9 +171,11 @@ where let mut args := args.toVector for h : i in [:args.size] do let arg := args[i] + trace[grind.debug.canon] "[{repr (← shouldCanon pinfos i arg)}]: {arg} : {← inferType arg}" let arg' ← match (← shouldCanon pinfos i arg) with | .canonType => canonType f i arg | .canonInst => canonInst f i arg + | .canonImplicit => canonImplicit f i (← visit arg) | .visit => visit arg unless ptrEq arg arg' do args := args.set i arg' diff --git a/tests/lean/run/grind_cat2.lean b/tests/lean/run/grind_cat2.lean new file mode 100644 index 000000000000..b9b03ea222db --- /dev/null +++ b/tests/lean/run/grind_cat2.lean @@ -0,0 +1,311 @@ +-- import Lean.Meta.Tactic.Grind +universe v v₁ v₂ v₃ u u₁ u₂ u₃ + +namespace CategoryTheory + +macro "cat_tac" : tactic => `(tactic| (intros; (try ext); grind)) + +class Category (obj : Type u) : Type max u (v + 1) where + Hom : obj → obj → Type v + /-- The identity morphism on an object. -/ + id : ∀ X : obj, Hom X X + /-- Composition of morphisms in a category, written `f ≫ g`. -/ + comp : ∀ {X Y Z : obj}, (Hom X Y) → (Hom Y Z) → (Hom X Z) + /-- Identity morphisms are left identities for composition. -/ + id_comp : ∀ {X Y : obj} (f : Hom X Y), comp (id X) f = f := by cat_tac + /-- Identity morphisms are right identities for composition. -/ + comp_id : ∀ {X Y : obj} (f : Hom X Y), comp f (id Y) = f := by cat_tac + /-- Composition in a category is associative. -/ + assoc : ∀ {W X Y Z : obj} (f : Hom W X) (g : Hom X Y) (h : Hom Y Z), comp (comp f g) h = comp f (comp g h) := by cat_tac + +infixr:10 " ⟶ " => Category.Hom +scoped notation "𝟙" => Category.id -- type as \b1 +scoped infixr:80 " ≫ " => Category.comp + +attribute [simp] Category.id_comp Category.comp_id Category.assoc + +attribute [grind =] Category.id_comp Category.comp_id +attribute [grind _=_] Category.assoc + +structure Functor (C : Type u₁) [Category.{v₁} C] (D : Type u₂) [Category.{v₂} D] : Type max v₁ v₂ u₁ u₂ where + /-- The action of a functor on objects. -/ + obj : C → D + /-- The action of a functor on morphisms. -/ + map : ∀ {X Y : C}, (X ⟶ Y) → ((obj X) ⟶ (obj Y)) + /-- A functor preserves identity morphisms. -/ + map_id : ∀ X : C, map (𝟙 X) = 𝟙 (obj X) := by cat_tac + /-- A functor preserves composition. -/ + map_comp : ∀ {X Y Z : C} (f : X ⟶ Y) (g : Y ⟶ Z), map (f ≫ g) = (map f) ≫ (map g) := by cat_tac + +infixr:26 " ⥤ " => Functor + +attribute [simp] Functor.map_id Functor.map_comp + +attribute [grind =] Functor.map_id +attribute [grind _=_] Functor.map_comp + +variable {C : Type u₁} [Category.{v₁} C] {D : Type u₂} [Category.{v₂} D] {E : Type u₃} [Category.{v₃} E] +variable {F G H : Functor C D} + +namespace Functor + +def comp (F : Functor C D) (G : Functor D E) : Functor C E where + obj X := G.obj (F.obj X) + map f := G.map (F.map f) + -- Note `map_id` and `map_comp` are handled by `cat_tac`. + +infixr:80 " ⋙ " => Functor.comp + +variable {X Y : C} {G : Functor D E} + +@[simp, grind =] theorem comp_obj : (F.comp G).obj X = G.obj (F.obj X) := rfl +@[simp, grind =] theorem comp_map (f : X ⟶ Y) : (F.comp G).map f = G.map (F.map f) := rfl + +end Functor + +@[ext] +structure NatTrans [Category.{v₁, u₁} C] [Category.{v₂, u₂} D] (F G : Functor C D) : Type max u₁ v₂ where + /-- The component of a natural transformation. -/ + app : ∀ X : C, F.obj X ⟶ G.obj X + /-- The naturality square for a given morphism. -/ + naturality : ∀ ⦃X Y : C⦄ (f : X ⟶ Y), F.map f ≫ app Y = app X ≫ G.map f := by cat_tac + +attribute [simp, grind =] NatTrans.naturality + +namespace NatTrans + +variable {X : C} + +protected def id (F : Functor C D) : NatTrans F F where app X := 𝟙 (F.obj X) + +@[simp, grind =] theorem id_app : (NatTrans.id F).app X = 𝟙 (F.obj X) := rfl + +protected def vcomp (α : NatTrans F G) (β : NatTrans G H) : NatTrans F H where + app X := α.app X ≫ β.app X + -- `naturality` is now handled by `grind`; in Mathlib this relies on `@[reassoc]` attributes. + -- Manual proof: + -- rw [← Category.assoc] + -- rw [α.naturality f] + -- rw [Category.assoc] + -- rw [β.naturality f] + -- rw [← Category.assoc] + +@[simp, grind =] theorem vcomp_app (α : NatTrans F G) (β : NatTrans G H) (X : C) : + (α.vcomp β).app X = α.app X ≫ β.app X := rfl + +end NatTrans + +instance Functor.category : Category.{max u₁ v₂} (Functor C D) where + Hom F G := NatTrans F G + id F := NatTrans.id F + comp α β := NatTrans.vcomp α β + -- Here we're okay: all the proofs are handled by `cat_tac`. + +@[simp, grind =] +theorem id_app (F : Functor C D) (X : C) : (𝟙 F : F ⟶ F).app X = 𝟙 (F.obj X) := rfl + +@[simp, grind _=_] +theorem comp_app {F G H : Functor C D} (α : F ⟶ G) (β : G ⟶ H) (X : C) : + (α ≫ β).app X = α.app X ≫ β.app X := rfl + +theorem app_naturality {F G : Functor C (Functor D E)} (T : F ⟶ G) (X : C) {Y Z : D} (f : Y ⟶ Z) : + (F.obj X).map f ≫ (T.app X).app Z = (T.app X).app Y ≫ (G.obj X).map f := by + cat_tac + +theorem naturality_app {F G : Functor C (Functor D E)} (T : F ⟶ G) (Z : D) {X Y : C} (f : X ⟶ Y) : + (F.map f).app Z ≫ (T.app Y).app Z = (T.app X).app Z ≫ (G.map f).app Z := by + cat_tac -- this is done manually in Mathlib! + -- rw [← comp_app] + -- rw [T.naturality f] + -- rw [comp_app] + +open Category Functor NatTrans + +def hcomp {H I : Functor D E} (α : F ⟶ G) (β : H ⟶ I) : F.comp H ⟶ G.comp I where + app := fun X : C => β.app (F.obj X) ≫ I.map (α.app X) + -- `grind` can now handle `naturality`, while Mathlib does this manually: + -- rw [Functor.comp_map, Functor.comp_map, ← assoc, naturality, assoc, ← I.map_comp, naturality, + -- map_comp, assoc] + + + +structure Iso {C : Type u} [Category.{v} C] (X Y : C) where + hom : X ⟶ Y + inv : Y ⟶ X + hom_inv_id : hom ≫ inv = 𝟙 X := by cat_tac + inv_hom_id : inv ≫ hom = 𝟙 Y := by cat_tac + +attribute [grind =] Iso.hom_inv_id Iso.inv_hom_id + +/-- Notation for an isomorphism in a category. -/ +infixr:10 " ≅ " => Iso -- type as \cong or \iso + +variable {C : Type u} [Category.{v} C] {X Y Z : C} + +namespace Iso + +@[ext] +theorem ext ⦃α β : X ≅ Y⦄ (w : α.hom = β.hom) : α = β := + suffices α.inv = β.inv by + cases α + cases β + cases w + cases this + rfl + calc + α.inv = α.inv ≫ β.hom ≫ β.inv := by grind + _ = β.inv := by grind + + +/-- `LeftInverse g f` means that g is a left inverse to f. That is, `g ∘ f = id`. -/ +def Function.LeftInverse (g : β → α) (f : α → β) : Prop := + ∀ x, g (f x) = x + +/-- `RightInverse g f` means that g is a right inverse to f. That is, `f ∘ g = id`. -/ +def Function.RightInverse (g : β → α) (f : α → β) : Prop := + LeftInverse f g + +open Function + +/-- `α ≃ β` is the type of functions from `α → β` with a two-sided inverse. -/ +structure Equiv (α : Sort _) (β : Sort _) where + protected toFun : α → β + protected invFun : β → α + protected left_inv : LeftInverse invFun toFun + protected right_inv : RightInverse invFun toFun + +@[inherit_doc] +infixl:25 " ≃ " => Equiv + +attribute [local grind] Function.LeftInverse in +/-- The bijection `(Z ⟶ X) ≃ (Z ⟶ Y)` induced by `α : X ≅ Y`. -/ +def homToEquiv (α : X ≅ Y) {Z : C} : (Z ⟶ X) ≃ (Z ⟶ Y) where + toFun f := f ≫ α.hom + invFun g := g ≫ α.inv + left_inv := by cat_tac + right_inv := sorry + +end Iso + + +variable {C : Type u₁} [Category.{v₁} C] {D : Type u₂} [Category.{v₂} D] + +-- Perhaps in the future we could redefine `Functor` in terms of this, but that isn't the +-- immediate plan. +/-- An unbundled functor. -/ +class Functorial (F : C → D) : Type max v₁ v₂ u₁ u₂ where + /-- A functorial map extends to an action on morphisms. -/ + map' : ∀ {X Y : C}, (X ⟶ Y) → (F X ⟶ F Y) + /-- A functorial map preserves identities. -/ + map_id' : ∀ X : C, map' (𝟙 X) = 𝟙 (F X) := by cat_tac + /-- A functorial map preserves composition of morphisms. -/ + map_comp' : ∀ {X Y Z : C} (f : X ⟶ Y) (g : Y ⟶ Z), map' (f ≫ g) = map' f ≫ map' g := by + cat_tac + +def map (F : C → D) [Functorial.{v₁, v₂} F] {X Y : C} (f : X ⟶ Y) : F X ⟶ F Y := + Functorial.map'.{v₁, v₂} f + +@[simp, grind =] +theorem map'_as_map {F : C → D} [Functorial.{v₁, v₂} F] {X Y : C} {f : X ⟶ Y} : + Functorial.map'.{v₁, v₂} f = map F f := + rfl + +@[simp, grind =] +theorem Functorial.map_id {F : C → D} [Functorial.{v₁, v₂} F] {X : C} : map F (𝟙 X) = 𝟙 (F X) := + Functorial.map_id' X + +@[simp, grind =] +theorem Functorial.map_comp {F : C → D} [Functorial.{v₁, v₂} F] {X Y Z : C} {f : X ⟶ Y} + {g : Y ⟶ Z} : map F (f ≫ g) = map F f ≫ map F g := + Functorial.map_comp' f g + +namespace Functor + +/-- Bundle a functorial function as a functor. +-/ +def of (F : C → D) [I : Functorial.{v₁, v₂} F] : C ⥤ D := + { I with + obj := F + map := CategoryTheory.map F } + +end Functor + +instance (F : C ⥤ D) : Functorial.{v₁, v₂} F.obj := + { F with map' := F.map } + +@[simp, grind =] +theorem map_functorial_obj (F : C ⥤ D) {X Y : C} (f : X ⟶ Y) : map F.obj f = F.map f := + rfl + +attribute [grind] _root_.id + +instance functorial_id : Functorial.{v₁, v₁} (id : C → C) where map' f := f + +namespace Ex1 +variable {E : Type u₃} [Category.{v₃} E] + +/-- +info: [grind.issues] the following implicit arguments (which are not type class instances or types) are definitionally equal with `default` transparency but not with a more restrictive transparency + G (F X) + and + (G ∘ F) X +[grind.issues] the following implicit arguments (which are not type class instances or types) are definitionally equal with `default` transparency but not with a more restrictive transparency + G (F Y) + and + (G ∘ F) Y +[grind.issues] the following implicit arguments (which are not type class instances or types) are definitionally equal with `default` transparency but not with a more restrictive transparency + G (F Z) + and + (G ∘ F) Z +-/ +#guard_msgs (info) in +def functorial_comp (F : C → D) [Functorial.{v₁, v₂} F] (G : D → E) [Functorial.{v₂, v₃} G] : + Functorial.{v₁, v₃} (G ∘ F) := + { Functor.of F ⋙ Functor.of G with + map' := fun f => map G (map F f) + map_id' := sorry + map_comp' := by + set_option trace.grind.issues true in + fail_if_success grind + sorry + } +end Ex1 + +namespace Ex2 +variable {E : Type u₃} [Category.{v₃} E] + +-- Since in the trace above, `Function.comp` generated problems because it is not `reducible`, +-- we can fix the issue by instructing `grind` to unfold it. +attribute [local grind] Function.comp + + +def functorial_comp (F : C → D) [Functorial.{v₁, v₂} F] (G : D → E) [Functorial.{v₂, v₃} G] : + Functorial.{v₁, v₃} (G ∘ F) := + { Functor.of F ⋙ Functor.of G with + map' := fun f => map G (map F f) + map_id' := sorry + map_comp' := by grind + } + +end Ex2 + +namespace Ex3 +variable {E : Type u₃} [Category.{v₃} E] + +-- Since in the trace above, `Function.comp` generated problems because it is not `reducible`, +-- we set it to reducible. +set_option allowUnsafeReducibility true +attribute [reducible] Function.comp + +def functorial_comp (F : C → D) [Functorial.{v₁, v₂} F] (G : D → E) [Functorial.{v₂, v₃} G] : + Functorial.{v₁, v₃} (G ∘ F) := + { Functor.of F ⋙ Functor.of G with + map' := fun f => map G (map F f) + map_id' := sorry + map_comp' := by grind + } + +end Ex3 + + +end CategoryTheory From 1b4272821d034a2721893171f5c0d0294861753a Mon Sep 17 00:00:00 2001 From: David Thrane Christiansen Date: Thu, 9 Jan 2025 08:01:35 +0100 Subject: [PATCH 063/100] feat: add UInt32.{lt, le} (#6591) This PR adds less-than and less-than-or-equal-to relations to `UInt32`, consistent with the other `UIntN` types. --- src/Init/Data/UInt/Basic.lean | 4 ++++ 1 file changed, 4 insertions(+) diff --git a/src/Init/Data/UInt/Basic.lean b/src/Init/Data/UInt/Basic.lean index 8a41d8df3497..081979f23a1d 100644 --- a/src/Init/Data/UInt/Basic.lean +++ b/src/Init/Data/UInt/Basic.lean @@ -159,6 +159,8 @@ def UInt32.xor (a b : UInt32) : UInt32 := ⟨a.toBitVec ^^^ b.toBitVec⟩ def UInt32.shiftLeft (a b : UInt32) : UInt32 := ⟨a.toBitVec <<< (mod b 32).toBitVec⟩ @[extern "lean_uint32_shift_right"] def UInt32.shiftRight (a b : UInt32) : UInt32 := ⟨a.toBitVec >>> (mod b 32).toBitVec⟩ +def UInt32.lt (a b : UInt32) : Prop := a.toBitVec < b.toBitVec +def UInt32.le (a b : UInt32) : Prop := a.toBitVec ≤ b.toBitVec instance : Add UInt32 := ⟨UInt32.add⟩ instance : Sub UInt32 := ⟨UInt32.sub⟩ @@ -169,6 +171,8 @@ set_option linter.deprecated false in instance : HMod UInt32 Nat UInt32 := ⟨UInt32.modn⟩ instance : Div UInt32 := ⟨UInt32.div⟩ +instance : LT UInt32 := ⟨UInt32.lt⟩ +instance : LE UInt32 := ⟨UInt32.le⟩ @[extern "lean_uint32_complement"] def UInt32.complement (a : UInt32) : UInt32 := ⟨~~~a.toBitVec⟩ From a6789a73ffeeb80d6570312919e30d0e2e1caaa5 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Henrik=20B=C3=B6ving?= Date: Thu, 9 Jan 2025 10:33:03 +0100 Subject: [PATCH 064/100] feat: `Std.Net.Addr` (#6563) This PR implements `Std.Net.Addr` which contains structures around IP and socket addresses. While we could implement our own parser instead of going through the `addr_in`/`addr_in6` route we will need to implement these conversions to make proper system calls anyway. Hence this is likely the approach with the least amount of non trivial code overall. The only thing I am uncertain about is whether `ofString` should return `Option` or `Except`, unfortunately `libuv` doesn't hand out error messages on IP parsing. --- src/Std.lean | 1 + src/Std/Net.lean | 7 ++ src/Std/Net/Addr.lean | 197 +++++++++++++++++++++++++++++++++++ src/runtime/CMakeLists.txt | 2 +- src/runtime/uv/net_addr.cpp | 131 +++++++++++++++++++++++ src/runtime/uv/net_addr.h | 28 +++++ tests/lean/run/net_addr.lean | 39 +++++++ 7 files changed, 404 insertions(+), 1 deletion(-) create mode 100644 src/Std/Net.lean create mode 100644 src/Std/Net/Addr.lean create mode 100644 src/runtime/uv/net_addr.cpp create mode 100644 src/runtime/uv/net_addr.h create mode 100644 tests/lean/run/net_addr.lean diff --git a/src/Std.lean b/src/Std.lean index 0a2f53c641b6..f427d00b8a8f 100644 --- a/src/Std.lean +++ b/src/Std.lean @@ -10,3 +10,4 @@ import Std.Sync import Std.Time import Std.Tactic import Std.Internal +import Std.Net diff --git a/src/Std/Net.lean b/src/Std/Net.lean new file mode 100644 index 000000000000..3d96dac34807 --- /dev/null +++ b/src/Std/Net.lean @@ -0,0 +1,7 @@ +/- +Copyright (c) 2025 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Henrik Böving +-/ +prelude +import Std.Net.Addr diff --git a/src/Std/Net/Addr.lean b/src/Std/Net/Addr.lean new file mode 100644 index 000000000000..dffa107465fc --- /dev/null +++ b/src/Std/Net/Addr.lean @@ -0,0 +1,197 @@ +/- +Copyright (c) 2025 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Henrik Böving +-/ +prelude +import Init.Data.Vector.Basic + +/-! +This module contains Lean representations of IP and socket addresses: +- `IPv4Addr`: Representing IPv4 addresses. +- `SocketAddressV4`: Representing a pair of IPv4 address and port. +- `IPv6Addr`: Representing IPv6 addresses. +- `SocketAddressV6`: Representing a pair of IPv6 address and port. +- `IPAddr`: Can either be an `IPv4Addr` or an `IPv6Addr`. +- `SocketAddress`: Can either be a `SocketAddressV4` or `SocketAddressV6`. +-/ + +namespace Std +namespace Net + +/-- +Representation of an IPv4 address. +-/ +structure IPv4Addr where + /-- + This structure represents the address: `octets[0].octets[1].octets[2].octets[3]`. + -/ + octets : Vector UInt8 4 + deriving Inhabited, DecidableEq + +/-- +A pair of an `IPv4Addr` and a port. +-/ +structure SocketAddressV4 where + addr : IPv4Addr + port : UInt16 + deriving Inhabited, DecidableEq + +/-- +Representation of an IPv6 address. +-/ +structure IPv6Addr where + /-- + This structure represents the address: `segments[0]:segments[1]:...`. + -/ + segments : Vector UInt16 8 + deriving Inhabited, DecidableEq + +/-- +A pair of an `IPv6Addr` and a port. +-/ +structure SocketAddressV6 where + addr : IPv6Addr + port : UInt16 + deriving Inhabited, DecidableEq + +/-- +An IP address, either IPv4 or IPv6. +-/ +inductive IPAddr where + | v4 (addr : IPv4Addr) + | v6 (addr : IPv6Addr) + deriving Inhabited, DecidableEq + +/-- +Either a `SocketAddressV4` or `SocketAddressV6`. +-/ +inductive SocketAddress where + | v4 (addr : SocketAddressV4) + | v6 (addr : SocketAddressV6) + deriving Inhabited, DecidableEq + +/-- +The kinds of address families supported by Lean, currently only IP variants. +-/ +inductive AddressFamily where + | ipv4 + | ipv6 + deriving Inhabited, DecidableEq + +namespace IPv4Addr + +/-- +Build the IPv4 address `a.b.c.d`. +-/ +def ofParts (a b c d : UInt8) : IPv4Addr := + { octets := #v[a, b, c, d] } + +/-- +Try to parse `s` as an IPv4 address, returning `none` on failure. +-/ +@[extern "lean_uv_pton_v4"] +opaque ofString (s : @&String) : Option IPv4Addr + +/-- +Turn `addr` into a `String` in the usual IPv4 format. +-/ +@[extern "lean_uv_ntop_v4"] +opaque toString (addr : @&IPv4Addr) : String + +instance : ToString IPv4Addr where + toString := toString + +instance : Coe IPv4Addr IPAddr where + coe addr := .v4 addr + +end IPv4Addr + +namespace SocketAddressV4 + +instance : Coe SocketAddressV4 SocketAddress where + coe addr := .v4 addr + +end SocketAddressV4 + +namespace IPv6Addr + +/-- +Build the IPv6 address `a:b:c:d:e:f:g:h`. +-/ +def ofParts (a b c d e f g h : UInt16) : IPv6Addr := + { segments := #v[a, b, c, d, e, f, g, h] } + +/-- +Try to parse `s` as an IPv6 address according to +[RFC 2373](https://datatracker.ietf.org/doc/html/rfc2373), returning `none` on failure. +-/ +@[extern "lean_uv_pton_v6"] +opaque ofString (s : @&String) : Option IPv6Addr + +/-- +Turn `addr` into a `String` in the IPv6 format described in +[RFC 2373](https://datatracker.ietf.org/doc/html/rfc2373). +-/ +@[extern "lean_uv_ntop_v6"] +opaque toString (addr : @&IPv6Addr) : String + +instance : ToString IPv6Addr where + toString := toString + +instance : Coe IPv6Addr IPAddr where + coe addr := .v6 addr + +end IPv6Addr + +namespace SocketAddressV6 + +instance : Coe SocketAddressV6 SocketAddress where + coe addr := .v6 addr + +end SocketAddressV6 + +namespace IPAddr + +/-- +Obtain the `AddressFamily` associated with an `IPAddr`. +-/ +def family : IPAddr → AddressFamily + | .v4 .. => .ipv4 + | .v6 .. => .ipv6 + +def toString : IPAddr → String + | .v4 addr => addr.toString + | .v6 addr => addr.toString + +instance : ToString IPAddr where + toString := toString + +end IPAddr + +namespace SocketAddress + +/-- +Obtain the `AddressFamily` associated with a `SocketAddress`. +-/ +def family : SocketAddress → AddressFamily + | .v4 .. => .ipv4 + | .v6 .. => .ipv6 + +/-- +Obtain the `IPAddr` contained in a `SocketAddress`. +-/ +def ipAddr : SocketAddress → IPAddr + | .v4 sa => .v4 sa.addr + | .v6 sa => .v6 sa.addr + +/-- +Obtain the port contained in a `SocketAddress`. +-/ +def port : SocketAddress → UInt16 + | .v4 sa | .v6 sa => sa.port + +end SocketAddress + +end Net +end Std diff --git a/src/runtime/CMakeLists.txt b/src/runtime/CMakeLists.txt index 367c7c298038..d6d8b50b395d 100644 --- a/src/runtime/CMakeLists.txt +++ b/src/runtime/CMakeLists.txt @@ -2,7 +2,7 @@ set(RUNTIME_OBJS debug.cpp thread.cpp mpz.cpp utf8.cpp object.cpp apply.cpp exception.cpp interrupt.cpp memory.cpp stackinfo.cpp compact.cpp init_module.cpp load_dynlib.cpp io.cpp hash.cpp platform.cpp alloc.cpp allocprof.cpp sharecommon.cpp stack_overflow.cpp -process.cpp object_ref.cpp mpn.cpp mutex.cpp libuv.cpp) +process.cpp object_ref.cpp mpn.cpp mutex.cpp libuv.cpp uv/net_addr.cpp) add_library(leanrt_initial-exec STATIC ${RUNTIME_OBJS}) set_target_properties(leanrt_initial-exec PROPERTIES ARCHIVE_OUTPUT_DIRECTORY ${CMAKE_CURRENT_BINARY_DIR}) diff --git a/src/runtime/uv/net_addr.cpp b/src/runtime/uv/net_addr.cpp new file mode 100644 index 000000000000..f06a078f6874 --- /dev/null +++ b/src/runtime/uv/net_addr.cpp @@ -0,0 +1,131 @@ +/* +Copyright (c) 2025 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Author: Henrik Böving +*/ + +#include "runtime/uv/net_addr.h" +#include + +namespace lean { + +#ifndef LEAN_EMSCRIPTEN + +void lean_ipv4_addr_to_in_addr(b_obj_arg ipv4_addr, in_addr* out) { + out->s_addr = 0; + for (int i = 0; i < 4; i++) { + uint32_t octet = (uint32_t)(uint8_t)lean_unbox(array_uget(ipv4_addr, i)); + out->s_addr |= octet << (3 - i) * 8; + } + out->s_addr = htonl(out->s_addr); +} + +void lean_ipv6_addr_to_in6_addr(b_obj_arg ipv6_addr, in6_addr* out) { + for (int i = 0; i < 8; i++) { + uint16_t segment = htons((uint16_t)lean_unbox(array_uget(ipv6_addr, i))); + out->s6_addr[2 * i] = (uint8_t)segment; + out->s6_addr[2 * i + 1] = (uint8_t)(segment >> 8); + } +} + +lean_obj_res lean_in_addr_to_ipv4_addr(const in_addr* ipv4_addr) { + obj_res ret = alloc_array(0, 4); + uint32_t hostaddr = ntohl(ipv4_addr->s_addr); + + for (int i = 0; i < 4; i++) { + uint8_t octet = (uint8_t)(hostaddr >> (3 - i) * 8); + array_push(ret, lean_box(octet)); + } + + lean_assert(array_size(ret) == 4); + return ret; +} + +lean_obj_res lean_in6_addr_to_ipv6_addr(const in6_addr* ipv6_addr) { + obj_res ret = alloc_array(0, 8); + + for (int i = 0; i < 16; i += 2) { + uint16_t part1 = (uint16_t)ipv6_addr->s6_addr[i]; + uint16_t part2 = (uint16_t)ipv6_addr->s6_addr[i + 1]; + uint16_t segment = ntohs((part2 << 8) | part1); + array_push(ret, lean_box(segment)); + } + + lean_assert(array_size(ret) == 8); + return ret; +} + +/* Std.Net.IPV4Addr.ofString (s : @&String) : Option IPV4Addr */ +extern "C" LEAN_EXPORT lean_obj_res lean_uv_pton_v4(b_obj_arg str_obj) { + const char* str = string_cstr(str_obj); + in_addr internal; + if (uv_inet_pton(AF_INET, str, &internal) == 0) { + return mk_option_some(lean_in_addr_to_ipv4_addr(&internal)); + } else { + return mk_option_none(); + } +} + +/* Std.Net.IPV4Addr.toString (addr : @&IPV4Addr) : String */ +extern "C" LEAN_EXPORT lean_obj_res lean_uv_ntop_v4(b_obj_arg ipv4_addr) { + in_addr internal; + lean_ipv4_addr_to_in_addr(ipv4_addr, &internal); + char dst[INET_ADDRSTRLEN]; + int ret = uv_inet_ntop(AF_INET, &internal, dst, sizeof(dst)); + lean_always_assert(ret == 0); + return lean_mk_string(dst); +} + +/* Std.Net.IPV6Addr.ofString (s : @&String) : Option IPV6Addr */ +extern "C" LEAN_EXPORT lean_obj_res lean_uv_pton_v6(b_obj_arg str_obj) { + const char* str = string_cstr(str_obj); + in6_addr internal; + if (uv_inet_pton(AF_INET6, str, &internal) == 0) { + return mk_option_some(lean_in6_addr_to_ipv6_addr(&internal)); + } else { + return mk_option_none(); + } +} + +/* Std.Net.IPV6Addr.toString (addr : @&IPV6Addr) : String */ +extern "C" LEAN_EXPORT lean_obj_res lean_uv_ntop_v6(b_obj_arg ipv6_addr) { + in6_addr internal; + lean_ipv6_addr_to_in6_addr(ipv6_addr, &internal); + char dst[INET6_ADDRSTRLEN]; + int ret = uv_inet_ntop(AF_INET6, &internal, dst, sizeof(dst)); + lean_always_assert(ret == 0); + return lean_mk_string(dst); +} + +#else + +/* Std.Net.IPV4Addr.ofString (s : @&String) : Option IPV4Addr */ +extern "C" LEAN_EXPORT lean_obj_res lean_uv_pton_v4(b_obj_arg str_obj) { + lean_always_assert( + false && ("Please build a version of Lean4 with libuv to invoke this.") + ); +} + +/* Std.Net.IPV4Addr.toString (addr : @&IPV4Addr) : String */ +extern "C" LEAN_EXPORT lean_obj_res lean_uv_ntop_v4(b_obj_arg ipv4_addr) { + lean_always_assert( + false && ("Please build a version of Lean4 with libuv to invoke this.") + ); +} + +/* Std.Net.IPV6Addr.ofString (s : @&String) : Option IPV6Addr */ +extern "C" LEAN_EXPORT lean_obj_res lean_uv_pton_v6(b_obj_arg str_obj) { + lean_always_assert( + false && ("Please build a version of Lean4 with libuv to invoke this.") + ); +} + +/* Std.Net.IPV6Addr.toString (addr : @&IPV6Addr) : String */ +extern "C" LEAN_EXPORT lean_obj_res lean_uv_ntop_v6(b_obj_arg ipv6_addr) { + lean_always_assert( + false && ("Please build a version of Lean4 with libuv to invoke this.") + ); +} + +#endif +} diff --git a/src/runtime/uv/net_addr.h b/src/runtime/uv/net_addr.h new file mode 100644 index 000000000000..2d6783beb50b --- /dev/null +++ b/src/runtime/uv/net_addr.h @@ -0,0 +1,28 @@ +/* +Copyright (c) 2025 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Author: Henrik Böving +*/ +#pragma once +#include +#include "runtime/object.h" + + +namespace lean { + +#ifndef LEAN_EMSCRIPTEN +#include + +void lean_ipv4_addr_to_in_addr(b_obj_arg ipv4_addr, struct in_addr* out); +void lean_ipv6_addr_to_in6_addr(b_obj_arg ipv6_addr, struct in6_addr* out); +lean_obj_res lean_in_addr_to_ipv4_addr(const struct in_addr* ipv4_addr); +lean_obj_res lean_in6_addr_to_ipv6_addr(const struct in6_addr* ipv6_addr); + +#endif + +extern "C" LEAN_EXPORT lean_obj_res lean_uv_pton_v4(b_obj_arg str_obj); +extern "C" LEAN_EXPORT lean_obj_res lean_uv_ntop_v4(b_obj_arg ipv4_addr); +extern "C" LEAN_EXPORT lean_obj_res lean_uv_pton_v6(b_obj_arg str_obj); +extern "C" LEAN_EXPORT lean_obj_res lean_uv_ntop_v6(b_obj_arg ipv6_addr); + +} diff --git a/tests/lean/run/net_addr.lean b/tests/lean/run/net_addr.lean new file mode 100644 index 000000000000..5503e880310f --- /dev/null +++ b/tests/lean/run/net_addr.lean @@ -0,0 +1,39 @@ +import Std.Net.Addr + +open Std.Net + +/-- info: true -/ +#guard_msgs in +#eval (IPv4Addr.ofParts 192 168 178 120).toString == "192.168.178.120" + +/-- info: true -/ +#guard_msgs in +#eval (IPv4Addr.ofParts 1 2 3 4).toString == "1.2.3.4" + +/-- info: true -/ +#guard_msgs in +#eval (IPv6Addr.ofParts 0xdead 0xbeef 0 0 0 0 0 0).toString == "dead:beef::" + +/-- info: true -/ +#guard_msgs in +#eval (IPv6Addr.ofParts 0x1234 0x5678 0x9abc 0xdef1 0x4321 0x8765 0xcba9 0x1fed).toString == "1234:5678:9abc:def1:4321:8765:cba9:1fed" + +/-- info: true -/ +#guard_msgs in +#eval IPv4Addr.ofString "1.2.3.4" == some (IPv4Addr.ofParts 1 2 3 4) + +/-- info: true -/ +#guard_msgs in +#eval IPv4Addr.ofString "192.168.300.120" |>.isNone + +/-- info: true -/ +#guard_msgs in +#eval IPv6Addr.ofString "dead:beef::" == some (IPv6Addr.ofParts 0xdead 0xbeef 0 0 0 0 0 0) + +/-- info: true -/ +#guard_msgs in +#eval IPv6Addr.ofString "1234:5678:9abc:def1:4321:8765:cba9:1fed" == some (IPv6Addr.ofParts 0x1234 0x5678 0x9abc 0xdef1 0x4321 0x8765 0xcba9 0x1fed) + +/-- info: true -/ +#guard_msgs in +#eval IPv6Addr.ofString "dead:beef::badaddress" |>.isNone From d3699764749a456beecc250ddf17df57ebace155 Mon Sep 17 00:00:00 2001 From: Leonardo de Moura Date: Thu, 9 Jan 2025 12:43:30 -0800 Subject: [PATCH 065/100] feat: improve inequality offset support theorems for `grind` (#6595) This PR improves the theorems used to justify the steps performed by the inequality offset module. See new test for examples of how they are going to be used. --- src/Init/Grind/Offset.lean | 81 +++++++++++++++------- tests/lean/run/grind_offset_ineq_thms.lean | 31 +++++++++ 2 files changed, 88 insertions(+), 24 deletions(-) create mode 100644 tests/lean/run/grind_offset_ineq_thms.lean diff --git a/src/Init/Grind/Offset.lean b/src/Init/Grind/Offset.lean index 56626662488f..705ca31f3b73 100644 --- a/src/Init/Grind/Offset.lean +++ b/src/Init/Grind/Offset.lean @@ -7,7 +7,7 @@ prelude import Init.Core import Init.Omega -namespace Lean.Grind +namespace Lean.Grind.Offset abbrev Var := Nat abbrev Context := Lean.RArray Nat @@ -22,7 +22,7 @@ structure Cnstr where y : Var k : Nat := 0 l : Bool := true - deriving Repr, BEq, Inhabited + deriving Repr, DecidableEq, Inhabited def Cnstr.denote (c : Cnstr) (ctx : Context) : Prop := if c.l then @@ -82,51 +82,84 @@ def Cnstr.isFalse (c : Cnstr) : Bool := c.x == c.y && c.k != 0 && c.l == true theorem Cnstr.of_isFalse (ctx : Context) {c : Cnstr} : c.isFalse = true → ¬c.denote ctx := by cases c; simp [isFalse]; intros; simp [*, denote]; omega -def Certificate := List Cnstr +def Cnstrs := List Cnstr -def Certificate.denote' (ctx : Context) (c₁ : Cnstr) (c₂ : Certificate) : Prop := +def Cnstrs.denoteAnd' (ctx : Context) (c₁ : Cnstr) (c₂ : Cnstrs) : Prop := match c₂ with | [] => c₁.denote ctx - | c::cs => c₁.denote ctx ∧ Certificate.denote' ctx c cs + | c::cs => c₁.denote ctx ∧ Cnstrs.denoteAnd' ctx c cs -theorem Certificate.denote'_trans (ctx : Context) (c₁ c : Cnstr) (cs : Certificate) : c₁.denote ctx → denote' ctx c cs → denote' ctx (c₁.trans c) cs := by +theorem Cnstrs.denote'_trans (ctx : Context) (c₁ c : Cnstr) (cs : Cnstrs) : c₁.denote ctx → denoteAnd' ctx c cs → denoteAnd' ctx (c₁.trans c) cs := by induction cs - next => simp [denote', *]; apply Cnstr.denote_trans - next c cs ih => simp [denote']; intros; simp [*] + next => simp [denoteAnd', *]; apply Cnstr.denote_trans + next c cs ih => simp [denoteAnd']; intros; simp [*] -def Certificate.trans' (c₁ : Cnstr) (c₂ : Certificate) : Cnstr := +def Cnstrs.trans' (c₁ : Cnstr) (c₂ : Cnstrs) : Cnstr := match c₂ with | [] => c₁ | c::c₂ => trans' (c₁.trans c) c₂ -@[simp] theorem Certificate.denote'_trans' (ctx : Context) (c₁ : Cnstr) (c₂ : Certificate) : denote' ctx c₁ c₂ → (trans' c₁ c₂).denote ctx := by +@[simp] theorem Cnstrs.denote'_trans' (ctx : Context) (c₁ : Cnstr) (c₂ : Cnstrs) : denoteAnd' ctx c₁ c₂ → (trans' c₁ c₂).denote ctx := by induction c₂ generalizing c₁ - next => intros; simp_all [trans', denote'] - next c cs ih => simp [denote']; intros; simp [trans']; apply ih; apply denote'_trans <;> assumption + next => intros; simp_all [trans', denoteAnd'] + next c cs ih => simp [denoteAnd']; intros; simp [trans']; apply ih; apply denote'_trans <;> assumption -def Certificate.denote (ctx : Context) (c : Certificate) : Prop := +def Cnstrs.denoteAnd (ctx : Context) (c : Cnstrs) : Prop := match c with | [] => True - | c::cs => denote' ctx c cs + | c::cs => denoteAnd' ctx c cs -def Certificate.trans (c : Certificate) : Cnstr := +def Cnstrs.trans (c : Cnstrs) : Cnstr := match c with | [] => trivialCnstr | c::cs => trans' c cs -theorem Certificate.denote_trans {ctx : Context} {c : Certificate} : c.denote ctx → c.trans.denote ctx := by - cases c <;> simp [*, trans, Certificate.denote] <;> intros <;> simp [*] +theorem Cnstrs.of_denoteAnd_trans {ctx : Context} {c : Cnstrs} : c.denoteAnd ctx → c.trans.denote ctx := by + cases c <;> simp [*, trans, denoteAnd] <;> intros <;> simp [*] -def Certificate.isFalse (c : Certificate) : Bool := +def Cnstrs.isFalse (c : Cnstrs) : Bool := c.trans.isFalse -theorem Certificate.unsat (ctx : Context) (c : Certificate) : c.isFalse = true → ¬ c.denote ctx := by +theorem Cnstrs.unsat' (ctx : Context) (c : Cnstrs) : c.isFalse = true → ¬ c.denoteAnd ctx := by simp [isFalse]; intro h₁ h₂ - have := Certificate.denote_trans h₂ + have := of_denoteAnd_trans h₂ have := Cnstr.of_isFalse ctx h₁ contradiction -theorem Certificate.imp (ctx : Context) (c : Certificate) : c.denote ctx → c.trans.denote ctx := by - apply denote_trans - -end Lean.Grind +/-- `denote ctx [c_1, ..., c_n] C` is `c_1.denote ctx → ... → c_n.denote ctx → C` -/ +def Cnstrs.denote (ctx : Context) (cs : Cnstrs) (C : Prop) : Prop := + match cs with + | [] => C + | c::cs => c.denote ctx → denote ctx cs C + +theorem Cnstrs.not_denoteAnd'_eq (ctx : Context) (c : Cnstr) (cs : Cnstrs) (C : Prop) : (denoteAnd' ctx c cs → C) = denote ctx (c::cs) C := by + simp [denote] + induction cs generalizing c + next => simp [denoteAnd', denote] + next c' cs ih => + simp [denoteAnd', denote, *] + +theorem Cnstrs.not_denoteAnd_eq (ctx : Context) (cs : Cnstrs) (C : Prop) : (denoteAnd ctx cs → C) = denote ctx cs C := by + cases cs + next => simp [denoteAnd, denote] + next c cs => apply not_denoteAnd'_eq + +def Cnstr.isImpliedBy (cs : Cnstrs) (c : Cnstr) : Bool := + cs.trans == c + +/-! Main theorems used by `grind`. -/ + +/-- Auxiliary theorem used by `grind` to prove that a system of offset inequalities is unsatisfiable. -/ +theorem Cnstrs.unsat (ctx : Context) (cs : Cnstrs) : cs.isFalse = true → cs.denote ctx False := by + intro h + rw [← not_denoteAnd_eq] + apply unsat' + assumption + +/-- Auxiliary theorem used by `grind` to prove an implied offset inequality. -/ +theorem Cnstrs.imp (ctx : Context) (cs : Cnstrs) (c : Cnstr) (h : c.isImpliedBy cs = true) : cs.denote ctx (c.denote ctx) := by + rw [← eq_of_beq h] + rw [← not_denoteAnd_eq] + apply of_denoteAnd_trans + +end Lean.Grind.Offset diff --git a/tests/lean/run/grind_offset_ineq_thms.lean b/tests/lean/run/grind_offset_ineq_thms.lean new file mode 100644 index 000000000000..bb76dcfe205e --- /dev/null +++ b/tests/lean/run/grind_offset_ineq_thms.lean @@ -0,0 +1,31 @@ +import Lean + +elab tk:"#R[" ts:term,* "]" : term => do + let ts : Array Lean.Syntax := ts + let es ← ts.mapM fun stx => Lean.Elab.Term.elabTerm stx none + if h : 0 < es.size then + return (Lean.RArray.toExpr (← Lean.Meta.inferType es[0]!) id (Lean.RArray.ofArray es h)) + else + throwErrorAt tk "RArray cannot be empty" + +open Lean.Grind.Offset + +macro "C[" "#" x:term:max " ≤ " "#" y:term:max "]" : term => `({ x := $x, y := $y : Cnstr }) +macro "C[" "#" x:term:max " + " k:term:max " ≤ " "#" y:term:max "]" : term => `({ x := $x, y := $y, k := $k : Cnstr }) +macro "C[" "#" x:term:max " ≤ " "#"y:term:max " + " k:term:max "]" : term => `({ x := $x, y := $y, k := $k, l := false : Cnstr }) + +example (x y z : Nat) : x + 2 ≤ y → y ≤ z → z + 1 ≤ x → False := + Cnstrs.unsat #R[x, y, z] [ + C[ #0 + 2 ≤ #1 ], + C[ #1 ≤ #2 ], + C[ #2 + 1 ≤ #0 ] + ] rfl + +example (x y z w : Nat) : x + 2 ≤ y → y ≤ z → z ≤ w + 7 → x ≤ w + 5 := + Cnstrs.imp #R[x, y, z, w] [ + C[ #0 + 2 ≤ #1 ], + C[ #1 ≤ #2 ], + C[ #2 ≤ #3 + 7] + ] + C[ #0 ≤ #3 + 5 ] + rfl From c07948a168de1062f86f2c9f8b96da4aab049d6a Mon Sep 17 00:00:00 2001 From: jrr6 <7482866+jrr6@users.noreply.github.com> Date: Thu, 9 Jan 2025 20:42:17 -0500 Subject: [PATCH 066/100] feat: add `simp?` and `dsimp?` in conversion mode (#6593) This PR adds support for the `simp?` and `dsimp?` tactics in conversion mode. Closes #6164 --- src/Init/Conv.lean | 7 +++++++ src/Lean/Elab/Tactic/Conv/Simp.lean | 27 ++++++++++++++++++++++++++- tests/lean/run/6164.lean | 29 +++++++++++++++++++++++++++++ 3 files changed, 62 insertions(+), 1 deletion(-) create mode 100644 tests/lean/run/6164.lean diff --git a/src/Init/Conv.lean b/src/Init/Conv.lean index 323abdf177e5..4bd2f9c3e098 100644 --- a/src/Init/Conv.lean +++ b/src/Init/Conv.lean @@ -150,6 +150,10 @@ See the `simp` tactic for more information. -/ syntax (name := simp) "simp" optConfig (discharger)? (&" only")? (" [" withoutPosition((simpStar <|> simpErase <|> simpLemma),*) "]")? : conv +/-- `simp?` takes the same arguments as `simp`, but reports an equivalent call to `simp only` +that would be sufficient to close the goal. See the `simp?` tactic for more information. -/ +syntax (name := simpTrace) "simp?" optConfig (discharger)? (&" only")? (simpArgs)? : conv + /-- `dsimp` is the definitional simplifier in `conv`-mode. It differs from `simp` in that it only applies theorems that hold by reflexivity. @@ -167,6 +171,9 @@ example (a : Nat): (0 + 0) = a - a := by syntax (name := dsimp) "dsimp" optConfig (discharger)? (&" only")? (" [" withoutPosition((simpErase <|> simpLemma),*) "]")? : conv +@[inherit_doc simpTrace] +syntax (name := dsimpTrace) "dsimp?" optConfig (&" only")? (dsimpArgs)? : conv + /-- `simp_match` simplifies match expressions. For example, ``` match [a, b] with diff --git a/src/Lean/Elab/Tactic/Conv/Simp.lean b/src/Lean/Elab/Tactic/Conv/Simp.lean index 17c7f0648bfc..b764034d3275 100644 --- a/src/Lean/Elab/Tactic/Conv/Simp.lean +++ b/src/Lean/Elab/Tactic/Conv/Simp.lean @@ -7,9 +7,10 @@ prelude import Lean.Elab.Tactic.Simp import Lean.Elab.Tactic.Split import Lean.Elab.Tactic.Conv.Basic +import Lean.Elab.Tactic.SimpTrace namespace Lean.Elab.Tactic.Conv -open Meta +open Meta Tactic TryThis def applySimpResult (result : Simp.Result) : TacticM Unit := do if result.proof?.isNone then @@ -23,6 +24,19 @@ def applySimpResult (result : Simp.Result) : TacticM Unit := do let (result, _) ← dischargeWrapper.with fun d? => simp lhs ctx (simprocs := simprocs) (discharge? := d?) applySimpResult result +@[builtin_tactic Lean.Parser.Tactic.Conv.simpTrace] def evalSimpTrace : Tactic := fun stx => withMainContext do + match stx with + | `(conv| simp?%$tk $cfg:optConfig $(discharger)? $[only%$o]? $[[$args,*]]?) => do + let stx ← `(tactic| simp%$tk $cfg:optConfig $[$discharger]? $[only%$o]? $[[$args,*]]?) + let { ctx, simprocs, dischargeWrapper, .. } ← mkSimpContext stx (eraseLocal := false) + let lhs ← getLhs + let (result, stats) ← dischargeWrapper.with fun d? => + simp lhs ctx (simprocs := simprocs) (discharge? := d?) + applySimpResult result + let stx ← mkSimpCallStx stx stats.usedTheorems + addSuggestion tk stx (origSpan? := ← getRef) + | _ => throwUnsupportedSyntax + @[builtin_tactic Lean.Parser.Tactic.Conv.simpMatch] def evalSimpMatch : Tactic := fun _ => withMainContext do applySimpResult (← Split.simpMatch (← getLhs)) @@ -30,4 +44,15 @@ def applySimpResult (result : Simp.Result) : TacticM Unit := do let { ctx, .. } ← mkSimpContext stx (eraseLocal := false) (kind := .dsimp) changeLhs (← Lean.Meta.dsimp (← getLhs) ctx).1 +@[builtin_tactic Lean.Parser.Tactic.Conv.dsimpTrace] def evalDSimpTrace : Tactic := fun stx => withMainContext do + match stx with + | `(conv| dsimp?%$tk $cfg:optConfig $[only%$o]? $[[$args,*]]?) => + let stx ← `(tactic| dsimp%$tk $cfg:optConfig $[only%$o]? $[[$args,*]]?) + let { ctx, .. } ← mkSimpContext stx (eraseLocal := false) (kind := .dsimp) + let (result, stats) ← Lean.Meta.dsimp (← getLhs) ctx + changeLhs result + let stx ← mkSimpCallStx stx stats.usedTheorems + addSuggestion tk stx (origSpan? := ← getRef) + | _ => throwUnsupportedSyntax + end Lean.Elab.Tactic.Conv diff --git a/tests/lean/run/6164.lean b/tests/lean/run/6164.lean new file mode 100644 index 000000000000..62145603e68a --- /dev/null +++ b/tests/lean/run/6164.lean @@ -0,0 +1,29 @@ +/-! + # `simp?` in conversion mode + + Tests that `simp?` and `dsimp?` work properly in `conv` mode: namely, that each displays the + appropriate suggestion and applies the corresponding simplification to the focused expression. +-/ + +attribute [simp] Nat.two_mul + +/-- +info: Try this: simp only [Nat.two_mul] +-/ +#guard_msgs in +example (n : Nat) : 123 + 2 * n = 123 + (n + n) := by + conv => + enter [1, 2] + simp? + + +@[simp] def foo (n : Nat) := n + 1 + +/-- +info: Try this: dsimp only [foo] +-/ +#guard_msgs in +example (n : Nat) : foo n = n + 1 := by + conv => + lhs + dsimp? From d2c4471cfa4611977bf4927b5cd849df1a4272b7 Mon Sep 17 00:00:00 2001 From: Alex Keizer Date: Fri, 10 Jan 2025 02:31:16 +0000 Subject: [PATCH 067/100] feat: BitVec.{toInt, toFin, msb}_udiv (#6402) This PR adds a `toFin` and `msb` lemma for unsigned bitvector division. We *don't* have `toInt_udiv`, since the only truly general statement we can make does no better than unfolding the definition, and it's not uncontroversially clear how to unfold `toInt` (see `toInt_eq_msb_cond`/`toInt_eq_toNat_cond`/`toInt_eq_toNat_bmod` for a few options currently provided). Instead, we do have `toInt_udiv_of_msb` that's able to provide a more meaningful rewrite given an extra side-condition (that `x.msb = false`). This PR also upstreams a minor `Nat` theorem (`Nat.div_le_div_left`) needed for the above from Mathlib. --------- Co-authored-by: Kim Morrison --- src/Init/Data/BitVec/Lemmas.lean | 63 ++++++++++++++++++++++++++++--- src/Init/Data/Nat/Div/Lemmas.lean | 8 ++++ 2 files changed, 66 insertions(+), 5 deletions(-) diff --git a/src/Init/Data/BitVec/Lemmas.lean b/src/Init/Data/BitVec/Lemmas.lean index a4599aaddb49..56bab2c22666 100644 --- a/src/Init/Data/BitVec/Lemmas.lean +++ b/src/Init/Data/BitVec/Lemmas.lean @@ -9,6 +9,7 @@ import Init.Data.Bool import Init.Data.BitVec.Basic import Init.Data.Fin.Lemmas import Init.Data.Nat.Lemmas +import Init.Data.Nat.Div.Lemmas import Init.Data.Nat.Mod import Init.Data.Int.Bitwise.Lemmas import Init.Data.Int.Pow @@ -442,6 +443,10 @@ theorem toInt_eq_toNat_cond (x : BitVec n) : (x.toNat : Int) - (2^n : Nat) := rfl +theorem toInt_eq_toNat_of_lt {x : BitVec n} (h : 2 * x.toNat < 2^n) : + x.toInt = x.toNat := by + simp [toInt_eq_toNat_cond, h] + theorem msb_eq_false_iff_two_mul_lt {x : BitVec w} : x.msb = false ↔ 2 * x.toNat < 2^w := by cases w <;> simp [Nat.pow_succ, Nat.mul_comm _ 2, msb_eq_decide, toNat_of_zero_length] @@ -454,6 +459,9 @@ theorem toInt_eq_msb_cond (x : BitVec w) : simp only [BitVec.toInt, ← msb_eq_false_iff_two_mul_lt] cases x.msb <;> rfl +theorem toInt_eq_toNat_of_msb {x : BitVec w} (h : x.msb = false) : + x.toInt = x.toNat := by + simp [toInt_eq_msb_cond, h] theorem toInt_eq_toNat_bmod (x : BitVec n) : x.toInt = Int.bmod x.toNat (2^n) := by simp only [toInt_eq_toNat_cond] @@ -2300,6 +2308,12 @@ theorem ofNat_sub_ofNat {n} (x y : Nat) : BitVec.ofNat n x - BitVec.ofNat n y = @[simp, bv_toNat] theorem toNat_neg (x : BitVec n) : (- x).toNat = (2^n - x.toNat) % 2^n := by simp [Neg.neg, BitVec.neg] +theorem toNat_neg_of_pos {x : BitVec n} (h : 0#n < x) : + (- x).toNat = 2^n - x.toNat := by + change 0 < x.toNat at h + rw [toNat_neg, Nat.mod_eq_of_lt] + omega + theorem toInt_neg {x : BitVec w} : (-x).toInt = (-x.toInt).bmod (2 ^ w) := by rw [← BitVec.zero_sub, toInt_sub] @@ -2586,13 +2600,13 @@ theorem udiv_def {x y : BitVec n} : x / y = BitVec.ofNat n (x.toNat / y.toNat) : rw [← udiv_eq] simp [udiv, bv_toNat, h, Nat.mod_eq_of_lt] +@[simp] +theorem toFin_udiv {x y : BitVec n} : (x / y).toFin = x.toFin / y.toFin := by + rfl + @[simp, bv_toNat] theorem toNat_udiv {x y : BitVec n} : (x / y).toNat = x.toNat / y.toNat := by - rw [udiv_def] - by_cases h : y = 0 - · simp [h] - · rw [toNat_ofNat, Nat.mod_eq_of_lt] - exact Nat.lt_of_le_of_lt (Nat.div_le_self ..) (by omega) + rfl @[simp] theorem zero_udiv {x : BitVec w} : (0#w) / x = 0#w := by @@ -2628,6 +2642,45 @@ theorem udiv_self {x : BitVec w} : ↓reduceIte, toNat_udiv] rw [Nat.div_self (by omega), Nat.mod_eq_of_lt (by omega)] +theorem msb_udiv (x y : BitVec w) : + (x / y).msb = (x.msb && y == 1#w) := by + cases msb_x : x.msb + · suffices x.toNat / y.toNat < 2 ^ (w - 1) by simpa [msb_eq_decide] + calc + x.toNat / y.toNat ≤ x.toNat := by apply Nat.div_le_self + _ < 2 ^ (w - 1) := by simpa [msb_eq_decide] using msb_x + . rcases w with _|w + · contradiction + · have : (y == 1#_) = decide (y.toNat = 1) := by + simp [(· == ·), toNat_eq] + simp only [this, Bool.true_and] + match hy : y.toNat with + | 0 => + obtain rfl : y = 0#_ := eq_of_toNat_eq hy + simp + | 1 => + obtain rfl : y = 1#_ := eq_of_toNat_eq (by simp [hy]) + simpa using msb_x + | y + 2 => + suffices x.toNat / (y + 2) < 2 ^ w by + simp_all [msb_eq_decide, hy] + calc + x.toNat / (y + 2) + ≤ x.toNat / 2 := by apply Nat.div_add_le_right (by omega) + _ < 2 ^ w := by omega + +theorem msb_udiv_eq_false_of {x : BitVec w} (h : x.msb = false) (y : BitVec w) : + (x / y).msb = false := by + simp [msb_udiv, h] + +/-- +If `x` is nonnegative (i.e., does not have its msb set), +then `x / y` is nonnegative, thus `toInt` and `toNat` coincide. +-/ +theorem toInt_udiv_of_msb {x : BitVec w} (h : x.msb = false) (y : BitVec w) : + (x / y).toInt = x.toNat / y.toNat := by + simp [toInt_eq_msb_cond, msb_udiv_eq_false_of h] + /-! ### umod -/ theorem umod_def {x y : BitVec n} : diff --git a/src/Init/Data/Nat/Div/Lemmas.lean b/src/Init/Data/Nat/Div/Lemmas.lean index 8fb6b1fef73e..d5a26a0950ba 100644 --- a/src/Init/Data/Nat/Div/Lemmas.lean +++ b/src/Init/Data/Nat/Div/Lemmas.lean @@ -49,4 +49,12 @@ theorem lt_div_mul_self (h : 0 < k) (w : k ≤ x) : x - k < x / k * k := by have : x % k < k := mod_lt x h omega +theorem div_le_div_left (hcb : c ≤ b) (hc : 0 < c) : a / b ≤ a / c := + (Nat.le_div_iff_mul_le hc).2 <| + Nat.le_trans (Nat.mul_le_mul_left _ hcb) (Nat.div_mul_le_self a b) + +theorem div_add_le_right {z : Nat} (h : 0 < z) (x y : Nat) : + x / (y + z) ≤ x / z := + div_le_div_left (Nat.le_add_left z y) h + end Nat From ed309dc2a4cda70f376b98a001dd734a17924d87 Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Fri, 10 Jan 2025 04:34:46 -0300 Subject: [PATCH 068/100] feat: add decidable instances for comparison operation of time offset types (#6587) This PR adds decidable instances for the `LE` and `LT` instances for the `Offset` types defined in `Std.Time`. --- src/Std/Time/Date/Unit/Month.lean | 6 ++++++ src/Std/Time/Date/Unit/Week.lean | 6 ++++++ src/Std/Time/Time/Unit/Hour.lean | 8 +++++++- src/Std/Time/Time/Unit/Millisecond.lean | 6 ++++++ src/Std/Time/Time/Unit/Minute.lean | 8 +++++++- src/Std/Time/Time/Unit/Nanosecond.lean | 3 +++ src/Std/Time/Time/Unit/Second.lean | 6 ++++++ 7 files changed, 41 insertions(+), 2 deletions(-) diff --git a/src/Std/Time/Date/Unit/Month.lean b/src/Std/Time/Date/Unit/Month.lean index d26210080621..f700fbcd65ec 100644 --- a/src/Std/Time/Date/Unit/Month.lean +++ b/src/Std/Time/Date/Unit/Month.lean @@ -39,6 +39,12 @@ instance {x y : Ordinal} : Decidable (x < y) := def Offset : Type := Int deriving Repr, BEq, Inhabited, Add, Sub, Mul, Div, Neg, ToString, LT, LE, DecidableEq +instance {x y : Offset} : Decidable (x ≤ y) := + Int.decLe x y + +instance {x y : Offset} : Decidable (x < y) := + Int.decLt x y + instance : OfNat Offset n := ⟨Int.ofNat n⟩ diff --git a/src/Std/Time/Date/Unit/Week.lean b/src/Std/Time/Date/Unit/Week.lean index 84fe108728a2..fd23dc8432c9 100644 --- a/src/Std/Time/Date/Unit/Week.lean +++ b/src/Std/Time/Date/Unit/Week.lean @@ -39,6 +39,12 @@ instance : Inhabited Ordinal where def Offset : Type := UnitVal (86400 * 7) deriving Repr, BEq, Inhabited, Add, Sub, Neg, LE, LT, ToString +instance {x y : Offset} : Decidable (x ≤ y) := + inferInstanceAs (Decidable (x.val ≤ y.val)) + +instance {x y : Offset} : Decidable (x < y) := + inferInstanceAs (Decidable (x.val < y.val)) + instance : OfNat Offset n := ⟨UnitVal.ofNat n⟩ namespace Ordinal diff --git a/src/Std/Time/Time/Unit/Hour.lean b/src/Std/Time/Time/Unit/Hour.lean index 7eceb2c7648b..fa200bb8004c 100644 --- a/src/Std/Time/Time/Unit/Hour.lean +++ b/src/Std/Time/Time/Unit/Hour.lean @@ -40,7 +40,13 @@ instance {x y : Ordinal} : Decidable (x < y) := or differences in hours. -/ def Offset : Type := UnitVal 3600 - deriving Repr, BEq, Inhabited, Add, Sub, Neg, ToString + deriving Repr, BEq, Inhabited, Add, Sub, Neg, ToString, LT, LE + +instance { x y : Offset } : Decidable (x ≤ y) := + inferInstanceAs (Decidable (x.val ≤ y.val)) + +instance { x y : Offset } : Decidable (x < y) := + inferInstanceAs (Decidable (x.val < y.val)) instance : OfNat Offset n := ⟨UnitVal.ofNat n⟩ diff --git a/src/Std/Time/Time/Unit/Millisecond.lean b/src/Std/Time/Time/Unit/Millisecond.lean index 754ddf22c477..9452a51fac81 100644 --- a/src/Std/Time/Time/Unit/Millisecond.lean +++ b/src/Std/Time/Time/Unit/Millisecond.lean @@ -40,6 +40,12 @@ instance {x y : Ordinal} : Decidable (x < y) := def Offset : Type := UnitVal (1 / 1000) deriving Repr, BEq, Inhabited, Add, Sub, Neg, LE, LT, ToString +instance { x y : Offset } : Decidable (x ≤ y) := + inferInstanceAs (Decidable (x.val ≤ y.val)) + +instance { x y : Offset } : Decidable (x < y) := + inferInstanceAs (Decidable (x.val < y.val)) + instance : OfNat Offset n := ⟨UnitVal.ofNat n⟩ diff --git a/src/Std/Time/Time/Unit/Minute.lean b/src/Std/Time/Time/Unit/Minute.lean index 0accf7f45dab..14cddd68b9d9 100644 --- a/src/Std/Time/Time/Unit/Minute.lean +++ b/src/Std/Time/Time/Unit/Minute.lean @@ -38,7 +38,13 @@ instance {x y : Ordinal} : Decidable (x < y) := `Offset` represents a duration offset in minutes. -/ def Offset : Type := UnitVal 60 - deriving Repr, BEq, Inhabited, Add, Sub, Neg, ToString + deriving Repr, BEq, Inhabited, Add, Sub, Neg, ToString, LT, LE + +instance { x y : Offset } : Decidable (x ≤ y) := + inferInstanceAs (Decidable (x.val ≤ y.val)) + +instance { x y : Offset } : Decidable (x < y) := + inferInstanceAs (Decidable (x.val < y.val)) instance : OfNat Offset n := ⟨UnitVal.ofInt <| Int.ofNat n⟩ diff --git a/src/Std/Time/Time/Unit/Nanosecond.lean b/src/Std/Time/Time/Unit/Nanosecond.lean index ad32e8989565..8e81975f2c49 100644 --- a/src/Std/Time/Time/Unit/Nanosecond.lean +++ b/src/Std/Time/Time/Unit/Nanosecond.lean @@ -42,6 +42,9 @@ def Offset : Type := UnitVal (1 / 1000000000) instance { x y : Offset } : Decidable (x ≤ y) := inferInstanceAs (Decidable (x.val ≤ y.val)) +instance { x y : Offset } : Decidable (x < y) := + inferInstanceAs (Decidable (x.val < y.val)) + instance : OfNat Offset n := ⟨UnitVal.ofNat n⟩ diff --git a/src/Std/Time/Time/Unit/Second.lean b/src/Std/Time/Time/Unit/Second.lean index 39c409f56056..4c34081e57d6 100644 --- a/src/Std/Time/Time/Unit/Second.lean +++ b/src/Std/Time/Time/Unit/Second.lean @@ -51,6 +51,12 @@ instance {x y : Ordinal l} : Decidable (x < y) := def Offset : Type := UnitVal 1 deriving Repr, BEq, Inhabited, Add, Sub, Neg, LE, LT, ToString +instance { x y : Offset } : Decidable (x ≤ y) := + inferInstanceAs (Decidable (x.val ≤ y.val)) + +instance { x y : Offset } : Decidable (x < y) := + inferInstanceAs (Decidable (x.val < y.val)) + instance : OfNat Offset n := ⟨UnitVal.ofNat n⟩ From 0b5d97725cef10e064acb792faafc4b5cdd35b39 Mon Sep 17 00:00:00 2001 From: Harun Khan <58151072+mhk119@users.noreply.github.com> Date: Fri, 10 Jan 2025 06:03:58 -0500 Subject: [PATCH 069/100] feat: `BitVec.toNat` theorems for `rotateLeft` and `rotateRight` (#6347) This PR adds `BitVec.toNat_rotateLeft` and `BitVec.toNat_rotateLeft`. --------- Co-authored-by: Kim Morrison --- src/Init/Data/BitVec/Lemmas.lean | 22 +++++++++++++++++++++- 1 file changed, 21 insertions(+), 1 deletion(-) diff --git a/src/Init/Data/BitVec/Lemmas.lean b/src/Init/Data/BitVec/Lemmas.lean index 56bab2c22666..8b49a5d8beb0 100644 --- a/src/Init/Data/BitVec/Lemmas.lean +++ b/src/Init/Data/BitVec/Lemmas.lean @@ -2876,7 +2876,12 @@ theorem smod_zero {x : BitVec n} : x.smod 0#n = x := by /-! # Rotate Left -/ -/-- rotateLeft is invariant under `mod` by the bitwidth. -/ +/--`rotateLeft` is defined in terms of left and right shifts. -/ +theorem rotateLeft_def {x : BitVec w} {r : Nat} : + x.rotateLeft r = (x <<< (r % w)) ||| (x >>> (w - r % w)) := by + simp only [rotateLeft, rotateLeftAux] + +/-- `rotateLeft` is invariant under `mod` by the bitwidth. -/ @[simp] theorem rotateLeft_mod_eq_rotateLeft {x : BitVec w} {r : Nat} : x.rotateLeft (r % w) = x.rotateLeft r := by @@ -3020,8 +3025,18 @@ theorem msb_rotateLeft {m w : Nat} {x : BitVec w} : · simp omega +@[simp] +theorem toNat_rotateLeft {x : BitVec w} {r : Nat} : + (x.rotateLeft r).toNat = (x.toNat <<< (r % w)) % (2^w) ||| x.toNat >>> (w - r % w) := by + simp only [rotateLeft_def, toNat_shiftLeft, toNat_ushiftRight, toNat_or] + /-! ## Rotate Right -/ +/-- `rotateRight` is defined in terms of left and right shifts. -/ +theorem rotateRight_def {x : BitVec w} {r : Nat} : + x.rotateRight r = (x >>> (r % w)) ||| (x <<< (w - r % w)) := by + simp only [rotateRight, rotateRightAux] + /-- Accessing bits in `x.rotateRight r` the range `[0, w-r)` is equal to accessing bits `x` in the range `[r, w)`. @@ -3157,6 +3172,11 @@ theorem msb_rotateRight {r w : Nat} {x : BitVec w} : simp [h₁] · simp [show w = 0 by omega] +@[simp] +theorem toNat_rotateRight {x : BitVec w} {r : Nat} : + (x.rotateRight r).toNat = (x.toNat >>> (r % w)) ||| x.toNat <<< (w - r % w) % (2^w) := by + simp only [rotateRight_def, toNat_shiftLeft, toNat_ushiftRight, toNat_or] + /- ## twoPow -/ theorem twoPow_eq (w : Nat) (i : Nat) : twoPow w i = 1#w <<< i := by From 58cd01154bc6dfae7c0921d04d082efe00856fd1 Mon Sep 17 00:00:00 2001 From: Lean stage0 autoupdater <> Date: Fri, 10 Jan 2025 16:42:03 +0000 Subject: [PATCH 070/100] chore: update stage0 --- stage0/src/runtime/CMakeLists.txt | 2 +- stage0/src/runtime/uv/net_addr.cpp | 131 + stage0/src/runtime/uv/net_addr.h | 28 + stage0/stdlib/Init/Conv.c | 313 + stage0/stdlib/Init/Data/Array/Lemmas.c | 148 +- stage0/stdlib/Init/Data/BitVec/Lemmas.c | 6 +- stage0/stdlib/Init/Data/List/Lemmas.c | 414 +- stage0/stdlib/Init/Data/UInt/Basic.c | 22 + stage0/stdlib/Init/Data/Vector/Basic.c | 807 +- stage0/stdlib/Init/Grind.c | 6 +- stage0/stdlib/Init/Grind/Norm.c | 6 +- stage0/stdlib/Init/Grind/Offset.c | 1367 + stage0/stdlib/Init/Grind/Tactics.c | 1102 +- stage0/stdlib/Init/Grind/Util.c | 48 + stage0/stdlib/Init/Prelude.c | 39 + stage0/stdlib/Init/Tactics.c | 2 +- stage0/stdlib/Lake/Load/Lean/Elab.c | 4 +- stage0/stdlib/Lean/Data/Json/Printer.c | 4236 +- stage0/stdlib/Lean/Data/Lsp.c | 6 +- stage0/stdlib/Lean/Data/Lsp/Basic.c | 7449 ++- stage0/stdlib/Lean/Data/Lsp/CancelParams.c | 603 + stage0/stdlib/Lean/Data/Lsp/Client.c | 14 +- stage0/stdlib/Lean/Data/Lsp/CodeActions.c | 60 +- stage0/stdlib/Lean/Data/Lsp/Diagnostics.c | 112 +- stage0/stdlib/Lean/Data/Lsp/Extra.c | 82 +- stage0/stdlib/Lean/Data/Lsp/InitShutdown.c | 26 +- stage0/stdlib/Lean/Data/Lsp/Internal.c | 244 +- stage0/stdlib/Lean/Data/Lsp/Ipc.c | 6 +- .../stdlib/Lean/Data/Lsp/LanguageFeatures.c | 402 +- stage0/stdlib/Lean/Data/Lsp/TextSync.c | 98 +- stage0/stdlib/Lean/Data/Lsp/Utf16.c | 6 +- stage0/stdlib/Lean/Data/Lsp/Workspace.c | 42 +- stage0/stdlib/Lean/Elab/BuiltinCommand.c | 156 +- .../stdlib/Lean/Elab/Tactic/BuiltinTactic.c | 510 +- stage0/stdlib/Lean/Elab/Tactic/Classical.c | 1467 +- stage0/stdlib/Lean/Elab/Tactic/Conv/Simp.c | 2740 +- stage0/stdlib/Lean/Elab/Tactic/Grind.c | 1655 +- stage0/stdlib/Lean/Elab/Tactic/Induction.c | 2311 +- stage0/stdlib/Lean/Elab/Tactic/NormCast.c | 10 +- stage0/stdlib/Lean/Elab/Tactic/RCases.c | 209 +- stage0/stdlib/Lean/Linter/MissingDocs.c | 6 +- stage0/stdlib/Lean/Meta/Tactic/Cases.c | 1114 +- stage0/stdlib/Lean/Meta/Tactic/Grind.c | 641 +- stage0/stdlib/Lean/Meta/Tactic/Grind/Canon.c | 17994 ++++--- stage0/stdlib/Lean/Meta/Tactic/Grind/Cases.c | 12070 ++++- .../Lean/Meta/Tactic/Grind/CasesMatch.c | 1154 + .../Lean/Meta/Tactic/Grind/Combinators.c | 1388 + stage0/stdlib/Lean/Meta/Tactic/Grind/Core.c | 5978 ++- stage0/stdlib/Lean/Meta/Tactic/Grind/Ctor.c | 663 +- .../stdlib/Lean/Meta/Tactic/Grind/DoNotSimp.c | 608 + stage0/stdlib/Lean/Meta/Tactic/Grind/EMatch.c | 23017 +++++---- .../Lean/Meta/Tactic/Grind/EMatchTheorem.c | 40051 ++++++++++------ .../Lean/Meta/Tactic/Grind/ForallProp.c | 3704 +- .../Lean/Meta/Tactic/Grind/Internalize.c | 18366 ++++--- stage0/stdlib/Lean/Meta/Tactic/Grind/Intro.c | 7879 +-- stage0/stdlib/Lean/Meta/Tactic/Grind/Inv.c | 2539 +- stage0/stdlib/Lean/Meta/Tactic/Grind/Main.c | 3021 +- .../Lean/Meta/Tactic/Grind/MarkNestedProofs.c | 6484 +-- stage0/stdlib/Lean/Meta/Tactic/Grind/Proj.c | 154 +- stage0/stdlib/Lean/Meta/Tactic/Grind/Proof.c | 3545 +- .../stdlib/Lean/Meta/Tactic/Grind/Propagate.c | 3426 +- stage0/stdlib/Lean/Meta/Tactic/Grind/Simp.c | 1330 +- .../stdlib/Lean/Meta/Tactic/Grind/SimpUtil.c | 586 + stage0/stdlib/Lean/Meta/Tactic/Grind/Split.c | 5023 ++ stage0/stdlib/Lean/Meta/Tactic/Grind/Types.c | 26441 ++++++---- stage0/stdlib/Lean/Meta/Tactic/Grind/Util.c | 1072 +- stage0/stdlib/Lean/Meta/Tactic/TryThis.c | 16 +- .../Lean/PrettyPrinter/Delaborator/Builtins.c | 30 +- stage0/stdlib/Lean/Server/CodeActions/Basic.c | 8 +- stage0/stdlib/Lean/Server/Completion.c | 24 +- stage0/stdlib/Lean/Server/FileWorker.c | 18 +- .../Lean/Server/FileWorker/RequestHandling.c | 38 +- .../Lean/Server/FileWorker/WidgetRequests.c | 14 +- stage0/stdlib/Lean/Server/References.c | 40 +- stage0/stdlib/Lean/Server/Watchdog.c | 44 +- stage0/stdlib/Lean/Util/ShareCommon.c | 10 +- stage0/stdlib/Lean/Util/Trace.c | 1426 +- stage0/stdlib/Lean/Widget/InteractiveCode.c | 44 +- .../Lean/Widget/InteractiveDiagnostic.c | 154 +- stage0/stdlib/Lean/Widget/InteractiveGoal.c | 74 +- stage0/stdlib/Lean/Widget/UserWidget.c | 56 +- stage0/stdlib/Std.c | 6 +- stage0/stdlib/Std/Net.c | 29 + stage0/stdlib/Std/Net/Addr.c | 1256 + stage0/stdlib/Std/Sat/AIG/Basic.c | 6 +- .../LRAT/Internal/Formula/Implementation.c | 6 +- stage0/stdlib/Std/Time/Date/Unit/Month.c | 352 +- stage0/stdlib/Std/Time/Date/Unit/Week.c | 352 +- stage0/stdlib/Std/Time/DateTime/Timestamp.c | 6 +- stage0/stdlib/Std/Time/Internal/Bounded.c | 6 +- stage0/stdlib/Std/Time/Time/Unit/Hour.c | 64 + .../stdlib/Std/Time/Time/Unit/Millisecond.c | 42 + stage0/stdlib/Std/Time/Time/Unit/Minute.c | 64 + stage0/stdlib/Std/Time/Time/Unit/Nanosecond.c | 21 + stage0/stdlib/Std/Time/Time/Unit/Second.c | 42 + 95 files changed, 145579 insertions(+), 73812 deletions(-) create mode 100644 stage0/src/runtime/uv/net_addr.cpp create mode 100644 stage0/src/runtime/uv/net_addr.h create mode 100644 stage0/stdlib/Init/Grind/Offset.c create mode 100644 stage0/stdlib/Lean/Data/Lsp/CancelParams.c create mode 100644 stage0/stdlib/Lean/Meta/Tactic/Grind/CasesMatch.c create mode 100644 stage0/stdlib/Lean/Meta/Tactic/Grind/Combinators.c create mode 100644 stage0/stdlib/Lean/Meta/Tactic/Grind/DoNotSimp.c create mode 100644 stage0/stdlib/Lean/Meta/Tactic/Grind/SimpUtil.c create mode 100644 stage0/stdlib/Lean/Meta/Tactic/Grind/Split.c create mode 100644 stage0/stdlib/Std/Net.c create mode 100644 stage0/stdlib/Std/Net/Addr.c diff --git a/stage0/src/runtime/CMakeLists.txt b/stage0/src/runtime/CMakeLists.txt index 367c7c298038..d6d8b50b395d 100644 --- a/stage0/src/runtime/CMakeLists.txt +++ b/stage0/src/runtime/CMakeLists.txt @@ -2,7 +2,7 @@ set(RUNTIME_OBJS debug.cpp thread.cpp mpz.cpp utf8.cpp object.cpp apply.cpp exception.cpp interrupt.cpp memory.cpp stackinfo.cpp compact.cpp init_module.cpp load_dynlib.cpp io.cpp hash.cpp platform.cpp alloc.cpp allocprof.cpp sharecommon.cpp stack_overflow.cpp -process.cpp object_ref.cpp mpn.cpp mutex.cpp libuv.cpp) +process.cpp object_ref.cpp mpn.cpp mutex.cpp libuv.cpp uv/net_addr.cpp) add_library(leanrt_initial-exec STATIC ${RUNTIME_OBJS}) set_target_properties(leanrt_initial-exec PROPERTIES ARCHIVE_OUTPUT_DIRECTORY ${CMAKE_CURRENT_BINARY_DIR}) diff --git a/stage0/src/runtime/uv/net_addr.cpp b/stage0/src/runtime/uv/net_addr.cpp new file mode 100644 index 000000000000..f06a078f6874 --- /dev/null +++ b/stage0/src/runtime/uv/net_addr.cpp @@ -0,0 +1,131 @@ +/* +Copyright (c) 2025 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Author: Henrik Böving +*/ + +#include "runtime/uv/net_addr.h" +#include + +namespace lean { + +#ifndef LEAN_EMSCRIPTEN + +void lean_ipv4_addr_to_in_addr(b_obj_arg ipv4_addr, in_addr* out) { + out->s_addr = 0; + for (int i = 0; i < 4; i++) { + uint32_t octet = (uint32_t)(uint8_t)lean_unbox(array_uget(ipv4_addr, i)); + out->s_addr |= octet << (3 - i) * 8; + } + out->s_addr = htonl(out->s_addr); +} + +void lean_ipv6_addr_to_in6_addr(b_obj_arg ipv6_addr, in6_addr* out) { + for (int i = 0; i < 8; i++) { + uint16_t segment = htons((uint16_t)lean_unbox(array_uget(ipv6_addr, i))); + out->s6_addr[2 * i] = (uint8_t)segment; + out->s6_addr[2 * i + 1] = (uint8_t)(segment >> 8); + } +} + +lean_obj_res lean_in_addr_to_ipv4_addr(const in_addr* ipv4_addr) { + obj_res ret = alloc_array(0, 4); + uint32_t hostaddr = ntohl(ipv4_addr->s_addr); + + for (int i = 0; i < 4; i++) { + uint8_t octet = (uint8_t)(hostaddr >> (3 - i) * 8); + array_push(ret, lean_box(octet)); + } + + lean_assert(array_size(ret) == 4); + return ret; +} + +lean_obj_res lean_in6_addr_to_ipv6_addr(const in6_addr* ipv6_addr) { + obj_res ret = alloc_array(0, 8); + + for (int i = 0; i < 16; i += 2) { + uint16_t part1 = (uint16_t)ipv6_addr->s6_addr[i]; + uint16_t part2 = (uint16_t)ipv6_addr->s6_addr[i + 1]; + uint16_t segment = ntohs((part2 << 8) | part1); + array_push(ret, lean_box(segment)); + } + + lean_assert(array_size(ret) == 8); + return ret; +} + +/* Std.Net.IPV4Addr.ofString (s : @&String) : Option IPV4Addr */ +extern "C" LEAN_EXPORT lean_obj_res lean_uv_pton_v4(b_obj_arg str_obj) { + const char* str = string_cstr(str_obj); + in_addr internal; + if (uv_inet_pton(AF_INET, str, &internal) == 0) { + return mk_option_some(lean_in_addr_to_ipv4_addr(&internal)); + } else { + return mk_option_none(); + } +} + +/* Std.Net.IPV4Addr.toString (addr : @&IPV4Addr) : String */ +extern "C" LEAN_EXPORT lean_obj_res lean_uv_ntop_v4(b_obj_arg ipv4_addr) { + in_addr internal; + lean_ipv4_addr_to_in_addr(ipv4_addr, &internal); + char dst[INET_ADDRSTRLEN]; + int ret = uv_inet_ntop(AF_INET, &internal, dst, sizeof(dst)); + lean_always_assert(ret == 0); + return lean_mk_string(dst); +} + +/* Std.Net.IPV6Addr.ofString (s : @&String) : Option IPV6Addr */ +extern "C" LEAN_EXPORT lean_obj_res lean_uv_pton_v6(b_obj_arg str_obj) { + const char* str = string_cstr(str_obj); + in6_addr internal; + if (uv_inet_pton(AF_INET6, str, &internal) == 0) { + return mk_option_some(lean_in6_addr_to_ipv6_addr(&internal)); + } else { + return mk_option_none(); + } +} + +/* Std.Net.IPV6Addr.toString (addr : @&IPV6Addr) : String */ +extern "C" LEAN_EXPORT lean_obj_res lean_uv_ntop_v6(b_obj_arg ipv6_addr) { + in6_addr internal; + lean_ipv6_addr_to_in6_addr(ipv6_addr, &internal); + char dst[INET6_ADDRSTRLEN]; + int ret = uv_inet_ntop(AF_INET6, &internal, dst, sizeof(dst)); + lean_always_assert(ret == 0); + return lean_mk_string(dst); +} + +#else + +/* Std.Net.IPV4Addr.ofString (s : @&String) : Option IPV4Addr */ +extern "C" LEAN_EXPORT lean_obj_res lean_uv_pton_v4(b_obj_arg str_obj) { + lean_always_assert( + false && ("Please build a version of Lean4 with libuv to invoke this.") + ); +} + +/* Std.Net.IPV4Addr.toString (addr : @&IPV4Addr) : String */ +extern "C" LEAN_EXPORT lean_obj_res lean_uv_ntop_v4(b_obj_arg ipv4_addr) { + lean_always_assert( + false && ("Please build a version of Lean4 with libuv to invoke this.") + ); +} + +/* Std.Net.IPV6Addr.ofString (s : @&String) : Option IPV6Addr */ +extern "C" LEAN_EXPORT lean_obj_res lean_uv_pton_v6(b_obj_arg str_obj) { + lean_always_assert( + false && ("Please build a version of Lean4 with libuv to invoke this.") + ); +} + +/* Std.Net.IPV6Addr.toString (addr : @&IPV6Addr) : String */ +extern "C" LEAN_EXPORT lean_obj_res lean_uv_ntop_v6(b_obj_arg ipv6_addr) { + lean_always_assert( + false && ("Please build a version of Lean4 with libuv to invoke this.") + ); +} + +#endif +} diff --git a/stage0/src/runtime/uv/net_addr.h b/stage0/src/runtime/uv/net_addr.h new file mode 100644 index 000000000000..2d6783beb50b --- /dev/null +++ b/stage0/src/runtime/uv/net_addr.h @@ -0,0 +1,28 @@ +/* +Copyright (c) 2025 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Author: Henrik Böving +*/ +#pragma once +#include +#include "runtime/object.h" + + +namespace lean { + +#ifndef LEAN_EMSCRIPTEN +#include + +void lean_ipv4_addr_to_in_addr(b_obj_arg ipv4_addr, struct in_addr* out); +void lean_ipv6_addr_to_in6_addr(b_obj_arg ipv6_addr, struct in6_addr* out); +lean_obj_res lean_in_addr_to_ipv4_addr(const struct in_addr* ipv4_addr); +lean_obj_res lean_in6_addr_to_ipv6_addr(const struct in6_addr* ipv6_addr); + +#endif + +extern "C" LEAN_EXPORT lean_obj_res lean_uv_pton_v4(b_obj_arg str_obj); +extern "C" LEAN_EXPORT lean_obj_res lean_uv_ntop_v4(b_obj_arg ipv4_addr); +extern "C" LEAN_EXPORT lean_obj_res lean_uv_pton_v6(b_obj_arg str_obj); +extern "C" LEAN_EXPORT lean_obj_res lean_uv_ntop_v6(b_obj_arg ipv6_addr); + +} diff --git a/stage0/stdlib/Init/Conv.c b/stage0/stdlib/Init/Conv.c index 4038d359657a..bba365609146 100644 --- a/stage0/stdlib/Init/Conv.c +++ b/stage0/stdlib/Init/Conv.c @@ -21,6 +21,7 @@ static lean_object* l_Lean_Parser_Tactic_Conv_convTry_____closed__6; static lean_object* l_Lean_Parser_Tactic_Conv_convLeft___closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convRfl__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_Conv_first___closed__9; +static lean_object* l_Lean_Parser_Tactic_Conv_simpTrace___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convTrace__state__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_Conv_convRepeat_____closed__3; static lean_object* l_Lean_Parser_Tactic_Conv_occsIndexed___closed__5; @@ -30,6 +31,7 @@ static lean_object* l_Lean_Parser_Tactic_Conv_nestedTactic___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv_dsimp; static lean_object* l_Lean_Parser_Tactic_Conv_simpMatch___closed__2; static lean_object* l_Lean_Parser_Tactic_Conv_conv_quot___closed__10; +static lean_object* l_Lean_Parser_Tactic_Conv_dsimpTrace___closed__5; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__failIfSuccess__1___boxed(lean_object*, lean_object*, lean_object*); lean_object* lean_mk_empty_array_with_capacity(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv_rewrite; @@ -50,6 +52,7 @@ static lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules_ LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv_change; static lean_object* l_Lean_Parser_Tactic_Conv_first___closed__20; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convRepeat____1(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv_dsimpTrace; static lean_object* l_Lean_Parser_Tactic_Conv_case___closed__2; static lean_object* l_Lean_Parser_Tactic_Conv_case_x27___closed__3; static lean_object* l_Lean_Parser_Tactic_Conv_conv_xb7_x2e_____closed__9; @@ -69,6 +72,7 @@ static lean_object* l_Lean_Parser_Tactic_Conv_convSeq___closed__5; static lean_object* l_Lean_Parser_Tactic_Conv_simpMatch___closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convDone__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv_zeta; +static lean_object* l_Lean_Parser_Tactic_Conv_dsimpTrace___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv_convSeq1Indented; static lean_object* l_Lean_Parser_Tactic_Conv_enter___closed__8; static lean_object* l_Lean_Parser_Tactic_Conv_ext___closed__7; @@ -150,6 +154,7 @@ uint8_t l_Lean_Syntax_isOfKind(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convErw______1___closed__1; lean_object* l_Lean_Name_mkStr5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_Conv_conv___x3c_x3b_x3e_____closed__2; +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv_simpTrace; static lean_object* l_Lean_Parser_Tactic_Conv_conv_xb7_x2e_____closed__12; static lean_object* l_Lean_Parser_Tactic_Conv_conv_quot___closed__18; static lean_object* l_Lean_Parser_Tactic_Conv_pattern___closed__2; @@ -161,6 +166,7 @@ static lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules_ static lean_object* l_Lean_Parser_Tactic_Conv_conv___closed__12; static lean_object* l_Lean_Parser_Tactic_Conv_convRfl___closed__2; static lean_object* l_Lean_Parser_Tactic_Conv_fun___closed__4; +static lean_object* l_Lean_Parser_Tactic_Conv_dsimpTrace___closed__7; static lean_object* l_Lean_Parser_Tactic_Conv_enter___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convDone__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_Conv_occsIndexed___closed__6; @@ -234,6 +240,7 @@ static lean_object* l_Lean_Parser_Tactic_Conv_argArg___closed__11; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__case__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_Conv_convConvSeq___closed__2; static lean_object* l_Lean_Parser_Tactic_Conv_nestedTacticCore___closed__5; +static lean_object* l_Lean_Parser_Tactic_Conv_simpTrace___closed__4; static lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convTrace__state__1___closed__1; static lean_object* l_Lean_Parser_Tactic_Conv_convTactic___closed__6; static lean_object* l_Lean_Parser_Tactic_Conv_case___closed__10; @@ -242,8 +249,10 @@ static lean_object* l_Lean_Parser_Tactic_Conv_convRepeat_____closed__6; static lean_object* l_Lean_Parser_Tactic_Conv_change___closed__1; static lean_object* l_Lean_Parser_Tactic_Conv_simpMatch___closed__3; static lean_object* l_Lean_Parser_Tactic_Conv_convNext_____x3d_x3e_____closed__7; +static lean_object* l_Lean_Parser_Tactic_Conv_simpTrace___closed__7; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convArgs__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_Conv_nestedTactic___closed__3; +static lean_object* l_Lean_Parser_Tactic_Conv_simpTrace___closed__6; static lean_object* l_Lean_Parser_Tactic_Conv_conv___x3c_x3b_x3e_____closed__6; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv_conv_quot; static lean_object* l_Lean_Parser_Tactic_Conv_occsIndexed___closed__9; @@ -266,9 +275,11 @@ static lean_object* l_Lean_Parser_Tactic_Conv_convApply_____closed__2; static lean_object* l_Lean_Parser_Tactic_Conv_case_x27___closed__7; static lean_object* l_Lean_Parser_Tactic_Conv_convApply_____closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__focus__1___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Tactic_Conv_simpTrace___closed__9; static lean_object* l_Lean_Parser_Tactic_Conv_arg___closed__1; static lean_object* l_Lean_Parser_Tactic_Conv_first___closed__14; static lean_object* l_Lean_Parser_Tactic_Conv_change___closed__9; +static lean_object* l_Lean_Parser_Tactic_Conv_dsimpTrace___closed__2; static lean_object* l_Lean_Parser_Tactic_Conv_convConvSeq___closed__6; static lean_object* l_Lean_Parser_Tactic_Conv_enter___closed__5; static lean_object* l_Lean_Parser_Tactic_Conv_simp___closed__2; @@ -278,6 +289,7 @@ static lean_object* l_Lean_Parser_Tactic_Conv_case___closed__4; static lean_object* l_Lean_Parser_Tactic_Conv_lhs___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__conv_xb7_x2e____1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_Conv_convNext_____x3d_x3e_____closed__9; +static lean_object* l_Lean_Parser_Tactic_Conv_simpTrace___closed__8; static lean_object* l_Lean_Parser_Tactic_Conv_convSeq___closed__4; static lean_object* l_Lean_Parser_Tactic_Conv_allGoals___closed__1; extern lean_object* l_Lean_Parser_Tactic_simpLemma; @@ -305,6 +317,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv_enterArg; static lean_object* l_Lean_Parser_Tactic_Conv_skip___closed__3; static lean_object* l_Lean_Parser_Tactic_Conv_focus___closed__3; static lean_object* l_Lean_Parser_Tactic_Conv_conv_xb7_x2e_____closed__15; +static lean_object* l_Lean_Parser_Tactic_Conv_simpTrace___closed__1; static lean_object* l_Lean_Parser_Tactic_Conv_failIfSuccess___closed__4; static lean_object* l_Lean_Parser_Tactic_Conv_convSeqBracketed___closed__8; static lean_object* l_Lean_Parser_Tactic_Conv_convNext_____x3d_x3e_____closed__2; @@ -481,6 +494,7 @@ static lean_object* l_Lean_Parser_Tactic_Conv_first___closed__3; static lean_object* l_Lean_Parser_Tactic_Conv_nestedConv___closed__2; static lean_object* l_Lean_Parser_Tactic_Conv_change___closed__6; static lean_object* l_Lean_Parser_Tactic_Conv_simp___closed__7; +static lean_object* l_Lean_Parser_Tactic_Conv_dsimpTrace___closed__8; static lean_object* l_Lean_Parser_Tactic_Conv_dsimp___closed__5; static lean_object* l_Lean_Parser_Tactic_Conv_rewrite___closed__1; static lean_object* l_Lean_Parser_Tactic_Conv_ext___closed__1; @@ -490,6 +504,7 @@ static lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules_ LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convRfl__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_Conv_conv_quot___closed__2; static lean_object* l_Lean_Parser_Tactic_Conv_argArg___closed__1; +static lean_object* l_Lean_Parser_Tactic_Conv_dsimpTrace___closed__9; lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convErw______1___closed__18; static lean_object* l_Lean_Parser_Tactic_Conv_convConvSeq___closed__5; @@ -540,6 +555,7 @@ static lean_object* l_Lean_Parser_Tactic_Conv_normCast___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv_unfold; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convTry____1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_Conv_change___closed__4; +static lean_object* l_Lean_Parser_Tactic_Conv_dsimpTrace___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv_congr; static lean_object* l_Lean_Parser_Tactic_Conv_allGoals___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv_rhs; @@ -568,9 +584,11 @@ static lean_object* l_Lean_Parser_Tactic_Conv_convApply_____closed__3; static lean_object* l_Lean_Parser_Tactic_Conv_conv___closed__9; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv_focus; static lean_object* l_Lean_Parser_Tactic_Conv_conv_xb7_x2e_____closed__5; +extern lean_object* l_Lean_Parser_Tactic_dsimpArgs; static lean_object* l_Lean_Parser_Tactic_Conv_ext___closed__10; LEAN_EXPORT lean_object* l_Lean_Parser_Category_conv; static lean_object* l_Lean_Parser_Tactic_Conv_dsimp___closed__1; +extern lean_object* l_Lean_Parser_Tactic_simpArgs; static lean_object* l_Lean_Parser_Tactic_Conv_fun___closed__2; static lean_object* l_Lean_Parser_Tactic_Conv_first___closed__27; static lean_object* l_Lean_Parser_Tactic_Conv_convConvSeq___closed__1; @@ -582,6 +600,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv_convTry__; static lean_object* l_Lean_Parser_Tactic_Conv_ext___closed__4; static lean_object* l_Lean_Parser_Tactic_Conv_convArgs___closed__4; static lean_object* l_Lean_Parser_Tactic_Conv_skip___closed__1; +static lean_object* l_Lean_Parser_Tactic_Conv_dsimpTrace___closed__6; static lean_object* l_Lean_Parser_Tactic_Conv_failIfSuccess___closed__3; static lean_object* l_Lean_Parser_Tactic_Conv_argArg___closed__8; size_t lean_array_size(lean_object*); @@ -595,6 +614,7 @@ static lean_object* l_Lean_Parser_Tactic_Conv_first___closed__19; static lean_object* l_Lean_Parser_Tactic_Conv_anyGoals___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__conv_xb7_x2e____1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_Conv_convRfl___closed__1; +static lean_object* l_Lean_Parser_Tactic_Conv_dsimpTrace___closed__4; lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_Conv_pattern___closed__3; static lean_object* l_Lean_Parser_Tactic_Conv_nestedTacticCore___closed__3; @@ -606,6 +626,8 @@ static lean_object* l_Lean_Parser_Tactic_Conv_first___closed__2; static lean_object* l_Lean_Parser_Tactic_Conv_focus___closed__4; static lean_object* l_Lean_Parser_Tactic_Conv_paren___closed__1; static lean_object* l_Lean_Parser_Tactic_Conv_conv_quot___closed__19; +static lean_object* l_Lean_Parser_Tactic_Conv_simpTrace___closed__5; +static lean_object* l_Lean_Parser_Tactic_Conv_simpTrace___closed__3; static lean_object* l_Lean_Parser_Tactic_Conv_convRw_______closed__2; static lean_object* l_Lean_Parser_Tactic_Conv_simp___closed__24; static lean_object* l_Lean_Parser_Tactic_Conv_occsWildcard___closed__4; @@ -613,6 +635,7 @@ static lean_object* l_Lean_Parser_Tactic_Conv_case_x27___closed__8; static lean_object* l_Lean_Parser_Tactic_Conv_simp___closed__22; static lean_object* l_Lean_Parser_Tactic_Conv_simp___closed__12; static lean_object* l_Lean_Parser_Tactic_Conv_occsWildcard___closed__2; +static lean_object* l_Lean_Parser_Tactic_Conv_simpTrace___closed__10; static lean_object* l_Lean_Parser_Tactic_Conv_convErw_______closed__2; static lean_object* l_Lean_Parser_Tactic_Conv_first___closed__15; static lean_object* l_Lean_Parser_Tactic_Conv_convRw_______closed__4; @@ -3365,6 +3388,137 @@ x_1 = l_Lean_Parser_Tactic_Conv_simp___closed__26; return x_1; } } +static lean_object* _init_l_Lean_Parser_Tactic_Conv_simpTrace___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("simpTrace", 9, 9); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Tactic_Conv_simpTrace___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_Parser_Tactic_Conv_conv_quot___closed__1; +x_2 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__2; +x_3 = l_Lean_Parser_Tactic_Conv_convSeq1Indented___closed__1; +x_4 = l_Lean_Parser_Tactic_Conv_convSeq1Indented___closed__2; +x_5 = l_Lean_Parser_Tactic_Conv_simpTrace___closed__1; +x_6 = l_Lean_Name_mkStr5(x_1, x_2, x_3, x_4, x_5); +return x_6; +} +} +static lean_object* _init_l_Lean_Parser_Tactic_Conv_simpTrace___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("simp\?", 5, 5); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Tactic_Conv_simpTrace___closed__4() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Tactic_Conv_simpTrace___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_Lean_Parser_Tactic_Conv_simpTrace___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__9; +x_2 = l_Lean_Parser_Tactic_Conv_simpTrace___closed__4; +x_3 = l_Lean_Parser_Tactic_optConfig; +x_4 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Parser_Tactic_Conv_simpTrace___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__9; +x_2 = l_Lean_Parser_Tactic_Conv_simpTrace___closed__5; +x_3 = l_Lean_Parser_Tactic_Conv_simp___closed__5; +x_4 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Parser_Tactic_Conv_simpTrace___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__9; +x_2 = l_Lean_Parser_Tactic_Conv_simpTrace___closed__6; +x_3 = l_Lean_Parser_Tactic_Conv_simp___closed__9; +x_4 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Parser_Tactic_Conv_simpTrace___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Tactic_Conv_argArg___closed__4; +x_2 = l_Lean_Parser_Tactic_simpArgs; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Tactic_Conv_simpTrace___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__9; +x_2 = l_Lean_Parser_Tactic_Conv_simpTrace___closed__7; +x_3 = l_Lean_Parser_Tactic_Conv_simpTrace___closed__8; +x_4 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Parser_Tactic_Conv_simpTrace___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Parser_Tactic_Conv_simpTrace___closed__2; +x_2 = lean_unsigned_to_nat(1022u); +x_3 = l_Lean_Parser_Tactic_Conv_simpTrace___closed__9; +x_4 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Parser_Tactic_Conv_simpTrace() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Parser_Tactic_Conv_simpTrace___closed__10; +return x_1; +} +} static lean_object* _init_l_Lean_Parser_Tactic_Conv_dsimp___closed__1() { _start: { @@ -3544,6 +3698,123 @@ x_1 = l_Lean_Parser_Tactic_Conv_dsimp___closed__13; return x_1; } } +static lean_object* _init_l_Lean_Parser_Tactic_Conv_dsimpTrace___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("dsimpTrace", 10, 10); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Tactic_Conv_dsimpTrace___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_Parser_Tactic_Conv_conv_quot___closed__1; +x_2 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__2; +x_3 = l_Lean_Parser_Tactic_Conv_convSeq1Indented___closed__1; +x_4 = l_Lean_Parser_Tactic_Conv_convSeq1Indented___closed__2; +x_5 = l_Lean_Parser_Tactic_Conv_dsimpTrace___closed__1; +x_6 = l_Lean_Name_mkStr5(x_1, x_2, x_3, x_4, x_5); +return x_6; +} +} +static lean_object* _init_l_Lean_Parser_Tactic_Conv_dsimpTrace___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("dsimp\?", 6, 6); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Tactic_Conv_dsimpTrace___closed__4() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Tactic_Conv_dsimpTrace___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_Lean_Parser_Tactic_Conv_dsimpTrace___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__9; +x_2 = l_Lean_Parser_Tactic_Conv_dsimpTrace___closed__4; +x_3 = l_Lean_Parser_Tactic_optConfig; +x_4 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Parser_Tactic_Conv_dsimpTrace___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__9; +x_2 = l_Lean_Parser_Tactic_Conv_dsimpTrace___closed__5; +x_3 = l_Lean_Parser_Tactic_Conv_simp___closed__9; +x_4 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Parser_Tactic_Conv_dsimpTrace___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Tactic_Conv_argArg___closed__4; +x_2 = l_Lean_Parser_Tactic_dsimpArgs; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Tactic_Conv_dsimpTrace___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__9; +x_2 = l_Lean_Parser_Tactic_Conv_dsimpTrace___closed__6; +x_3 = l_Lean_Parser_Tactic_Conv_dsimpTrace___closed__7; +x_4 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Parser_Tactic_Conv_dsimpTrace___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Parser_Tactic_Conv_dsimpTrace___closed__2; +x_2 = lean_unsigned_to_nat(1022u); +x_3 = l_Lean_Parser_Tactic_Conv_dsimpTrace___closed__8; +x_4 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Parser_Tactic_Conv_dsimpTrace() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Parser_Tactic_Conv_dsimpTrace___closed__9; +return x_1; +} +} static lean_object* _init_l_Lean_Parser_Tactic_Conv_simpMatch___closed__1() { _start: { @@ -9510,6 +9781,28 @@ l_Lean_Parser_Tactic_Conv_simp___closed__26 = _init_l_Lean_Parser_Tactic_Conv_si lean_mark_persistent(l_Lean_Parser_Tactic_Conv_simp___closed__26); l_Lean_Parser_Tactic_Conv_simp = _init_l_Lean_Parser_Tactic_Conv_simp(); lean_mark_persistent(l_Lean_Parser_Tactic_Conv_simp); +l_Lean_Parser_Tactic_Conv_simpTrace___closed__1 = _init_l_Lean_Parser_Tactic_Conv_simpTrace___closed__1(); +lean_mark_persistent(l_Lean_Parser_Tactic_Conv_simpTrace___closed__1); +l_Lean_Parser_Tactic_Conv_simpTrace___closed__2 = _init_l_Lean_Parser_Tactic_Conv_simpTrace___closed__2(); +lean_mark_persistent(l_Lean_Parser_Tactic_Conv_simpTrace___closed__2); +l_Lean_Parser_Tactic_Conv_simpTrace___closed__3 = _init_l_Lean_Parser_Tactic_Conv_simpTrace___closed__3(); +lean_mark_persistent(l_Lean_Parser_Tactic_Conv_simpTrace___closed__3); +l_Lean_Parser_Tactic_Conv_simpTrace___closed__4 = _init_l_Lean_Parser_Tactic_Conv_simpTrace___closed__4(); +lean_mark_persistent(l_Lean_Parser_Tactic_Conv_simpTrace___closed__4); +l_Lean_Parser_Tactic_Conv_simpTrace___closed__5 = _init_l_Lean_Parser_Tactic_Conv_simpTrace___closed__5(); +lean_mark_persistent(l_Lean_Parser_Tactic_Conv_simpTrace___closed__5); +l_Lean_Parser_Tactic_Conv_simpTrace___closed__6 = _init_l_Lean_Parser_Tactic_Conv_simpTrace___closed__6(); +lean_mark_persistent(l_Lean_Parser_Tactic_Conv_simpTrace___closed__6); +l_Lean_Parser_Tactic_Conv_simpTrace___closed__7 = _init_l_Lean_Parser_Tactic_Conv_simpTrace___closed__7(); +lean_mark_persistent(l_Lean_Parser_Tactic_Conv_simpTrace___closed__7); +l_Lean_Parser_Tactic_Conv_simpTrace___closed__8 = _init_l_Lean_Parser_Tactic_Conv_simpTrace___closed__8(); +lean_mark_persistent(l_Lean_Parser_Tactic_Conv_simpTrace___closed__8); +l_Lean_Parser_Tactic_Conv_simpTrace___closed__9 = _init_l_Lean_Parser_Tactic_Conv_simpTrace___closed__9(); +lean_mark_persistent(l_Lean_Parser_Tactic_Conv_simpTrace___closed__9); +l_Lean_Parser_Tactic_Conv_simpTrace___closed__10 = _init_l_Lean_Parser_Tactic_Conv_simpTrace___closed__10(); +lean_mark_persistent(l_Lean_Parser_Tactic_Conv_simpTrace___closed__10); +l_Lean_Parser_Tactic_Conv_simpTrace = _init_l_Lean_Parser_Tactic_Conv_simpTrace(); +lean_mark_persistent(l_Lean_Parser_Tactic_Conv_simpTrace); l_Lean_Parser_Tactic_Conv_dsimp___closed__1 = _init_l_Lean_Parser_Tactic_Conv_dsimp___closed__1(); lean_mark_persistent(l_Lean_Parser_Tactic_Conv_dsimp___closed__1); l_Lean_Parser_Tactic_Conv_dsimp___closed__2 = _init_l_Lean_Parser_Tactic_Conv_dsimp___closed__2(); @@ -9538,6 +9831,26 @@ l_Lean_Parser_Tactic_Conv_dsimp___closed__13 = _init_l_Lean_Parser_Tactic_Conv_d lean_mark_persistent(l_Lean_Parser_Tactic_Conv_dsimp___closed__13); l_Lean_Parser_Tactic_Conv_dsimp = _init_l_Lean_Parser_Tactic_Conv_dsimp(); lean_mark_persistent(l_Lean_Parser_Tactic_Conv_dsimp); +l_Lean_Parser_Tactic_Conv_dsimpTrace___closed__1 = _init_l_Lean_Parser_Tactic_Conv_dsimpTrace___closed__1(); +lean_mark_persistent(l_Lean_Parser_Tactic_Conv_dsimpTrace___closed__1); +l_Lean_Parser_Tactic_Conv_dsimpTrace___closed__2 = _init_l_Lean_Parser_Tactic_Conv_dsimpTrace___closed__2(); +lean_mark_persistent(l_Lean_Parser_Tactic_Conv_dsimpTrace___closed__2); +l_Lean_Parser_Tactic_Conv_dsimpTrace___closed__3 = _init_l_Lean_Parser_Tactic_Conv_dsimpTrace___closed__3(); +lean_mark_persistent(l_Lean_Parser_Tactic_Conv_dsimpTrace___closed__3); +l_Lean_Parser_Tactic_Conv_dsimpTrace___closed__4 = _init_l_Lean_Parser_Tactic_Conv_dsimpTrace___closed__4(); +lean_mark_persistent(l_Lean_Parser_Tactic_Conv_dsimpTrace___closed__4); +l_Lean_Parser_Tactic_Conv_dsimpTrace___closed__5 = _init_l_Lean_Parser_Tactic_Conv_dsimpTrace___closed__5(); +lean_mark_persistent(l_Lean_Parser_Tactic_Conv_dsimpTrace___closed__5); +l_Lean_Parser_Tactic_Conv_dsimpTrace___closed__6 = _init_l_Lean_Parser_Tactic_Conv_dsimpTrace___closed__6(); +lean_mark_persistent(l_Lean_Parser_Tactic_Conv_dsimpTrace___closed__6); +l_Lean_Parser_Tactic_Conv_dsimpTrace___closed__7 = _init_l_Lean_Parser_Tactic_Conv_dsimpTrace___closed__7(); +lean_mark_persistent(l_Lean_Parser_Tactic_Conv_dsimpTrace___closed__7); +l_Lean_Parser_Tactic_Conv_dsimpTrace___closed__8 = _init_l_Lean_Parser_Tactic_Conv_dsimpTrace___closed__8(); +lean_mark_persistent(l_Lean_Parser_Tactic_Conv_dsimpTrace___closed__8); +l_Lean_Parser_Tactic_Conv_dsimpTrace___closed__9 = _init_l_Lean_Parser_Tactic_Conv_dsimpTrace___closed__9(); +lean_mark_persistent(l_Lean_Parser_Tactic_Conv_dsimpTrace___closed__9); +l_Lean_Parser_Tactic_Conv_dsimpTrace = _init_l_Lean_Parser_Tactic_Conv_dsimpTrace(); +lean_mark_persistent(l_Lean_Parser_Tactic_Conv_dsimpTrace); l_Lean_Parser_Tactic_Conv_simpMatch___closed__1 = _init_l_Lean_Parser_Tactic_Conv_simpMatch___closed__1(); lean_mark_persistent(l_Lean_Parser_Tactic_Conv_simpMatch___closed__1); l_Lean_Parser_Tactic_Conv_simpMatch___closed__2 = _init_l_Lean_Parser_Tactic_Conv_simpMatch___closed__2(); diff --git a/stage0/stdlib/Init/Data/Array/Lemmas.c b/stage0/stdlib/Init/Data/Array/Lemmas.c index 7ee06a23c847..19ba875cf9dc 100644 --- a/stage0/stdlib/Init/Data/Array/Lemmas.c +++ b/stage0/stdlib/Init/Data/Array/Lemmas.c @@ -206,6 +206,80 @@ x_6 = lean_box(x_5); return x_6; } } +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_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_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_match__1_splitter___rarg___boxed), 3, 0); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Lemmas_0__List_filterMap_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_match__1_splitter___rarg(x_1, x_2, x_3); +lean_dec(x_2); +return x_4; +} +} LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Array_toListRev___spec__1___rarg(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { _start: { @@ -507,80 +581,6 @@ lean_dec(x_2); return x_4; } } -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_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_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_match__1_splitter___rarg___boxed), 3, 0); -return x_3; -} -} -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Lemmas_0__List_filterMap_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_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: { diff --git a/stage0/stdlib/Init/Data/BitVec/Lemmas.c b/stage0/stdlib/Init/Data/BitVec/Lemmas.c index bbc93a815a7e..325daedf3fa2 100644 --- a/stage0/stdlib/Init/Data/BitVec/Lemmas.c +++ b/stage0/stdlib/Init/Data/BitVec/Lemmas.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Init.Data.BitVec.Lemmas -// Imports: Init.Data.Bool Init.Data.BitVec.Basic Init.Data.Fin.Lemmas Init.Data.Nat.Lemmas Init.Data.Nat.Mod Init.Data.Int.Bitwise.Lemmas Init.Data.Int.Pow +// Imports: Init.Data.Bool Init.Data.BitVec.Basic Init.Data.Fin.Lemmas Init.Data.Nat.Lemmas Init.Data.Nat.Div.Lemmas Init.Data.Nat.Mod Init.Data.Int.Bitwise.Lemmas Init.Data.Int.Pow #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -621,6 +621,7 @@ lean_object* initialize_Init_Data_Bool(uint8_t builtin, lean_object*); lean_object* initialize_Init_Data_BitVec_Basic(uint8_t builtin, lean_object*); lean_object* initialize_Init_Data_Fin_Lemmas(uint8_t builtin, lean_object*); lean_object* initialize_Init_Data_Nat_Lemmas(uint8_t builtin, lean_object*); +lean_object* initialize_Init_Data_Nat_Div_Lemmas(uint8_t builtin, lean_object*); lean_object* initialize_Init_Data_Nat_Mod(uint8_t builtin, lean_object*); lean_object* initialize_Init_Data_Int_Bitwise_Lemmas(uint8_t builtin, lean_object*); lean_object* initialize_Init_Data_Int_Pow(uint8_t builtin, lean_object*); @@ -641,6 +642,9 @@ 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); +res = initialize_Init_Data_Nat_Div_Lemmas(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); res = initialize_Init_Data_Nat_Mod(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); diff --git a/stage0/stdlib/Init/Data/List/Lemmas.c b/stage0/stdlib/Init/Data/List/Lemmas.c index 82a167b40dcc..65fc6dc69c83 100644 --- a/stage0/stdlib/Init/Data/List/Lemmas.c +++ b/stage0/stdlib/Init/Data/List/Lemmas.c @@ -192,39 +192,60 @@ x_3 = lean_alloc_closure((void*)(l___private_Init_Data_List_Lemmas_0__List_isEqv return x_3; } } -LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_foldl_match__1_splitter___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_getLast_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: { -if (lean_obj_tag(x_2) == 0) +if (lean_obj_tag(x_1) == 0) { -lean_object* x_5; +lean_object* x_6; +lean_dec(x_5); lean_dec(x_4); -x_5 = lean_apply_1(x_3, x_1); -return x_5; +x_6 = lean_apply_1(x_3, lean_box(0)); +return x_6; } else { -lean_object* x_6; lean_object* x_7; lean_object* x_8; +lean_object* x_7; lean_dec(x_3); -x_6 = lean_ctor_get(x_2, 0); -lean_inc(x_6); -x_7 = lean_ctor_get(x_2, 1); +x_7 = lean_ctor_get(x_1, 1); lean_inc(x_7); -lean_dec(x_2); -x_8 = lean_apply_3(x_4, x_1, x_6, x_7); -return x_8; +if (lean_obj_tag(x_7) == 0) +{ +lean_object* x_8; lean_object* x_9; +lean_dec(x_5); +x_8 = lean_ctor_get(x_1, 0); +lean_inc(x_8); +lean_dec(x_1); +x_9 = lean_apply_2(x_4, x_8, lean_box(0)); +return x_9; +} +else +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +lean_dec(x_4); +x_10 = lean_ctor_get(x_1, 0); +lean_inc(x_10); +lean_dec(x_1); +x_11 = lean_ctor_get(x_7, 0); +lean_inc(x_11); +x_12 = lean_ctor_get(x_7, 1); +lean_inc(x_12); +lean_dec(x_7); +x_13 = lean_apply_4(x_5, x_10, x_11, x_12, lean_box(0)); +return x_13; } } } -LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_foldl_match__1_splitter(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +} +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_getLast_match__1_splitter(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_4; -x_4 = lean_alloc_closure((void*)(l___private_Init_Data_List_Lemmas_0__List_foldl_match__1_splitter___rarg), 4, 0); -return x_4; +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l___private_Init_Data_List_Lemmas_0__List_getLast_match__1_splitter___rarg), 5, 0); +return x_3; } } -LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_getLast_x3f_match__1_splitter___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_getLast_x21_match__1_splitter___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_1) == 0) @@ -246,24 +267,24 @@ return x_6; } } } -LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_getLast_x3f_match__1_splitter(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_getLast_x21_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_getLast_x3f_match__1_splitter___rarg___boxed), 3, 0); +x_3 = lean_alloc_closure((void*)(l___private_Init_Data_List_Lemmas_0__List_getLast_x21_match__1_splitter___rarg___boxed), 3, 0); return x_3; } } -LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_getLast_x3f_match__1_splitter___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_getLast_x21_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_getLast_x3f_match__1_splitter___rarg(x_1, x_2, x_3); +x_4 = l___private_Init_Data_List_Lemmas_0__List_getLast_x21_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_filterMap_match__1_splitter___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_getLast_x3f_match__1_splitter___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_1) == 0) @@ -274,221 +295,178 @@ return x_2; } else { -lean_object* x_4; lean_object* x_5; +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_5 = lean_apply_1(x_3, x_4); -return x_5; +x_6 = lean_apply_2(x_3, x_4, x_5); +return x_6; } } } -LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_filterMap_match__1_splitter(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_getLast_x3f_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_match__1_splitter___rarg___boxed), 3, 0); +x_3 = lean_alloc_closure((void*)(l___private_Init_Data_List_Lemmas_0__List_getLast_x3f_match__1_splitter___rarg___boxed), 3, 0); return x_3; } } -LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_filterMap_match__1_splitter___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_getLast_x3f_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_match__1_splitter___rarg(x_1, x_2, x_3); +x_4 = l___private_Init_Data_List_Lemmas_0__List_getLast_x3f_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_foldl__filterMap_match__1_splitter___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_getLastD_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_dec(x_2); -lean_inc(x_3); -return x_3; +lean_object* x_5; +lean_dec(x_4); +x_5 = lean_apply_1(x_3, x_2); +return x_5; } else { -lean_object* x_4; lean_object* x_5; -x_4 = lean_ctor_get(x_1, 0); -lean_inc(x_4); +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_5 = lean_apply_1(x_2, x_4); -return x_5; +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_foldl__filterMap_match__1_splitter(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_getLastD_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_foldl__filterMap_match__1_splitter___rarg___boxed), 3, 0); +x_3 = lean_alloc_closure((void*)(l___private_Init_Data_List_Lemmas_0__List_getLastD_match__1_splitter___rarg), 4, 0); return x_3; } } -LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_foldl__filterMap_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_foldl__filterMap_match__1_splitter___rarg(x_1, x_2, x_3); -lean_dec(x_3); -return x_4; -} -} -LEAN_EXPORT lean_object* l_List_foldlRecOn___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; -x_6 = lean_apply_4(x_1, x_2, x_3, x_4, lean_box(0)); -return x_6; -} -} -LEAN_EXPORT lean_object* l_List_foldlRecOn___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___private_Init_Data_List_Lemmas_0__List_filterMap_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_5); lean_dec(x_3); -lean_dec(x_2); -return x_4; +lean_inc(x_2); +return x_2; } else { -lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; -x_6 = lean_ctor_get(x_1, 0); -lean_inc(x_6); -x_7 = lean_ctor_get(x_1, 1); -lean_inc(x_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); -lean_inc(x_2); -lean_inc(x_6); -lean_inc(x_3); -x_8 = lean_apply_2(x_2, x_3, x_6); -lean_inc(x_5); -x_9 = lean_apply_4(x_5, x_3, x_4, x_6, lean_box(0)); -x_10 = lean_alloc_closure((void*)(l_List_foldlRecOn___rarg___lambda__1), 5, 1); -lean_closure_set(x_10, 0, x_5); -x_1 = x_7; -x_3 = x_8; -x_4 = x_9; -x_5 = x_10; -goto _start; +x_5 = lean_apply_1(x_3, x_4); +return x_5; } } } -LEAN_EXPORT lean_object* l_List_foldlRecOn(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_filterMap_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_match__1_splitter___rarg___boxed), 3, 0); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_filterMap_match__1_splitter___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = lean_alloc_closure((void*)(l_List_foldlRecOn___rarg), 5, 0); +x_4 = l___private_Init_Data_List_Lemmas_0__List_filterMap_match__1_splitter___rarg(x_1, x_2, x_3); +lean_dec(x_2); return x_4; } } -LEAN_EXPORT lean_object* l_List_foldrRecOn___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___private_Init_Data_List_Lemmas_0__Option_isSome_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_5); lean_dec(x_2); -lean_inc(x_4); -return x_4; +lean_inc(x_3); +return x_3; } else { -lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; -x_6 = lean_ctor_get(x_1, 0); -lean_inc(x_6); -x_7 = lean_ctor_get(x_1, 1); -lean_inc(x_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); -lean_inc(x_7); -lean_inc(x_2); -x_8 = l_List_foldr___rarg(x_2, x_3, x_7); -lean_inc(x_5); -x_9 = lean_alloc_closure((void*)(l_List_foldlRecOn___rarg___lambda__1), 5, 1); -lean_closure_set(x_9, 0, x_5); -x_10 = l_List_foldrRecOn___rarg(x_7, x_2, x_3, x_4, x_9); -x_11 = lean_apply_4(x_5, x_8, x_10, x_6, lean_box(0)); -return x_11; +x_5 = lean_apply_1(x_2, x_4); +return x_5; } } } -LEAN_EXPORT lean_object* l_List_foldrRecOn(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__Option_isSome_match__1_splitter(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_4; -x_4 = lean_alloc_closure((void*)(l_List_foldrRecOn___rarg___boxed), 5, 0); -return x_4; +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l___private_Init_Data_List_Lemmas_0__Option_isSome_match__1_splitter___rarg___boxed), 3, 0); +return x_3; } } -LEAN_EXPORT lean_object* l_List_foldrRecOn___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___private_Init_Data_List_Lemmas_0__Option_isSome_match__1_splitter___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_6; -x_6 = l_List_foldrRecOn___rarg(x_1, x_2, x_3, x_4, x_5); -lean_dec(x_4); +lean_object* x_4; +x_4 = l___private_Init_Data_List_Lemmas_0__Option_isSome_match__1_splitter___rarg(x_1, x_2, x_3); lean_dec(x_3); -return x_6; +return x_4; } } -LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_getLast_match__1_splitter___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___private_Init_Data_List_Lemmas_0__List_findSome_x3f_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_object* x_6; -lean_dec(x_5); -lean_dec(x_4); -x_6 = lean_apply_1(x_3, lean_box(0)); -return x_6; -} -else -{ -lean_object* x_7; -lean_dec(x_3); -x_7 = lean_ctor_get(x_1, 1); -lean_inc(x_7); -if (lean_obj_tag(x_7) == 0) -{ -lean_object* x_8; lean_object* x_9; -lean_dec(x_5); -x_8 = lean_ctor_get(x_1, 0); -lean_inc(x_8); -lean_dec(x_1); -x_9 = lean_apply_2(x_4, x_8, lean_box(0)); -return x_9; +lean_dec(x_2); +lean_inc(x_3); +return x_3; } else { -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; -lean_dec(x_4); -x_10 = lean_ctor_get(x_1, 0); -lean_inc(x_10); +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_11 = lean_ctor_get(x_7, 0); -lean_inc(x_11); -x_12 = lean_ctor_get(x_7, 1); -lean_inc(x_12); -lean_dec(x_7); -x_13 = lean_apply_4(x_5, x_10, x_11, x_12, lean_box(0)); -return x_13; -} +x_5 = lean_apply_1(x_2, x_4); +return x_5; } } } -LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_getLast_match__1_splitter(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_findSome_x3f_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_getLast_match__1_splitter___rarg), 5, 0); +x_3 = lean_alloc_closure((void*)(l___private_Init_Data_List_Lemmas_0__List_findSome_x3f_match__1_splitter___rarg___boxed), 3, 0); return x_3; } } -LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_getLast_x21_match__1_splitter___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_findSome_x3f_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_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_List_Lemmas_0__List_filterMap__replicate_match__1_splitter___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_1) == 0) @@ -499,67 +477,65 @@ return x_2; } else { -lean_object* x_4; lean_object* x_5; lean_object* x_6; +lean_object* x_4; lean_object* x_5; 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; +x_5 = lean_apply_1(x_3, x_4); +return x_5; } } } -LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_getLast_x21_match__1_splitter(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_filterMap__replicate_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_getLast_x21_match__1_splitter___rarg___boxed), 3, 0); +x_3 = lean_alloc_closure((void*)(l___private_Init_Data_List_Lemmas_0__List_filterMap__replicate_match__1_splitter___rarg___boxed), 3, 0); return x_3; } } -LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_getLast_x21_match__1_splitter___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_filterMap__replicate_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_getLast_x21_match__1_splitter___rarg(x_1, x_2, x_3); +x_4 = l___private_Init_Data_List_Lemmas_0__List_filterMap__replicate_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_getLastD_match__1_splitter___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_foldl_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) +if (lean_obj_tag(x_2) == 0) { lean_object* x_5; lean_dec(x_4); -x_5 = lean_apply_1(x_3, x_2); +x_5 = lean_apply_1(x_3, x_1); 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); +x_6 = lean_ctor_get(x_2, 0); lean_inc(x_6); -x_7 = lean_ctor_get(x_1, 1); +x_7 = lean_ctor_get(x_2, 1); lean_inc(x_7); -lean_dec(x_1); -x_8 = lean_apply_3(x_4, x_6, x_7, x_2); +lean_dec(x_2); +x_8 = lean_apply_3(x_4, x_1, x_6, x_7); return x_8; } } } -LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_getLastD_match__1_splitter(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_foldl_match__1_splitter(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l___private_Init_Data_List_Lemmas_0__List_getLastD_match__1_splitter___rarg), 4, 0); -return x_3; +lean_object* x_4; +x_4 = lean_alloc_closure((void*)(l___private_Init_Data_List_Lemmas_0__List_foldl_match__1_splitter___rarg), 4, 0); +return x_4; } } -LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__Option_isSome_match__1_splitter___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_foldl__filterMap_match__1_splitter___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_1) == 0) @@ -579,95 +555,119 @@ return x_5; } } } -LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__Option_isSome_match__1_splitter(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_foldl__filterMap_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__Option_isSome_match__1_splitter___rarg___boxed), 3, 0); +x_3 = lean_alloc_closure((void*)(l___private_Init_Data_List_Lemmas_0__List_foldl__filterMap_match__1_splitter___rarg___boxed), 3, 0); return x_3; } } -LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__Option_isSome_match__1_splitter___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_foldl__filterMap_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__Option_isSome_match__1_splitter___rarg(x_1, x_2, x_3); +x_4 = l___private_Init_Data_List_Lemmas_0__List_foldl__filterMap_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_List_Lemmas_0__List_findSome_x3f_match__1_splitter___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_List_foldlRecOn___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; +x_6 = lean_apply_4(x_1, x_2, x_3, x_4, lean_box(0)); +return x_6; +} +} +LEAN_EXPORT lean_object* l_List_foldlRecOn___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_1) == 0) { +lean_dec(x_5); +lean_dec(x_3); lean_dec(x_2); -lean_inc(x_3); -return x_3; +return x_4; } else { -lean_object* x_4; lean_object* x_5; -x_4 = lean_ctor_get(x_1, 0); -lean_inc(x_4); +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_6 = lean_ctor_get(x_1, 0); +lean_inc(x_6); +x_7 = lean_ctor_get(x_1, 1); +lean_inc(x_7); lean_dec(x_1); -x_5 = lean_apply_1(x_2, x_4); -return x_5; -} -} +lean_inc(x_2); +lean_inc(x_6); +lean_inc(x_3); +x_8 = lean_apply_2(x_2, x_3, x_6); +lean_inc(x_5); +x_9 = lean_apply_4(x_5, x_3, x_4, x_6, lean_box(0)); +x_10 = lean_alloc_closure((void*)(l_List_foldlRecOn___rarg___lambda__1), 5, 1); +lean_closure_set(x_10, 0, x_5); +x_1 = x_7; +x_3 = x_8; +x_4 = x_9; +x_5 = x_10; +goto _start; } -LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_findSome_x3f_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_match__1_splitter___rarg___boxed), 3, 0); -return x_3; } } -LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_findSome_x3f_match__1_splitter___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_List_foldlRecOn(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_match__1_splitter___rarg(x_1, x_2, x_3); -lean_dec(x_3); +x_4 = lean_alloc_closure((void*)(l_List_foldlRecOn___rarg), 5, 0); return x_4; } } -LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_filterMap__replicate_match__1_splitter___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_List_foldrRecOn___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_1) == 0) { -lean_dec(x_3); -lean_inc(x_2); -return x_2; +lean_dec(x_5); +lean_dec(x_2); +lean_inc(x_4); +return x_4; } else { -lean_object* x_4; lean_object* x_5; -x_4 = lean_ctor_get(x_1, 0); -lean_inc(x_4); +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_6 = lean_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_5 = lean_apply_1(x_3, x_4); -return x_5; +lean_inc(x_7); +lean_inc(x_2); +x_8 = l_List_foldr___rarg(x_2, x_3, x_7); +lean_inc(x_5); +x_9 = lean_alloc_closure((void*)(l_List_foldlRecOn___rarg___lambda__1), 5, 1); +lean_closure_set(x_9, 0, x_5); +x_10 = l_List_foldrRecOn___rarg(x_7, x_2, x_3, x_4, x_9); +x_11 = lean_apply_4(x_5, x_8, x_10, x_6, lean_box(0)); +return x_11; } } } -LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_filterMap__replicate_match__1_splitter(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_List_foldrRecOn(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l___private_Init_Data_List_Lemmas_0__List_filterMap__replicate_match__1_splitter___rarg___boxed), 3, 0); -return x_3; +lean_object* x_4; +x_4 = lean_alloc_closure((void*)(l_List_foldrRecOn___rarg___boxed), 5, 0); +return x_4; } } -LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_filterMap__replicate_match__1_splitter___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_List_foldrRecOn___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_4; -x_4 = l___private_Init_Data_List_Lemmas_0__List_filterMap__replicate_match__1_splitter___rarg(x_1, x_2, x_3); -lean_dec(x_2); -return x_4; +lean_object* x_6; +x_6 = l_List_foldrRecOn___rarg(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_6; } } LEAN_EXPORT lean_object* l___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) { diff --git a/stage0/stdlib/Init/Data/UInt/Basic.c b/stage0/stdlib/Init/Data/UInt/Basic.c index dcb9d3a99120..7be2b8c067cc 100644 --- a/stage0/stdlib/Init/Data/UInt/Basic.c +++ b/stage0/stdlib/Init/Data/UInt/Basic.c @@ -15,6 +15,7 @@ extern "C" { #endif LEAN_EXPORT lean_object* l_UInt64_modn___boxed(lean_object*, lean_object*); uint8_t lean_uint8_sub(uint8_t, uint8_t); +LEAN_EXPORT lean_object* l_instLEUInt32__1; LEAN_EXPORT uint8_t l_instDecidableLtUInt16(uint16_t, uint16_t); LEAN_EXPORT lean_object* l_instComplementUInt32; LEAN_EXPORT lean_object* l_USize_toUInt16___boxed(lean_object*); @@ -122,6 +123,7 @@ LEAN_EXPORT lean_object* l_instDecidableLtUInt8___boxed(lean_object*, lean_objec static lean_object* l_instMulUSize___closed__1; LEAN_EXPORT lean_object* l_Nat_cast___at_UInt8_modn___spec__1___boxed(lean_object*); LEAN_EXPORT uint8_t l_instDecidableLtUInt8(uint8_t, uint8_t); +LEAN_EXPORT lean_object* l_instLTUInt32__1; uint32_t lean_uint32_mul(uint32_t, uint32_t); LEAN_EXPORT lean_object* l_USize_modn(size_t, lean_object*); size_t lean_usize_lor(size_t, size_t); @@ -1669,6 +1671,22 @@ x_1 = l_instDivUInt32___closed__1; return x_1; } } +static lean_object* _init_l_instLTUInt32__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_box(0); +return x_1; +} +} +static lean_object* _init_l_instLEUInt32__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_box(0); +return x_1; +} +} LEAN_EXPORT lean_object* l_UInt32_complement___boxed(lean_object* x_1) { _start: { @@ -2940,6 +2958,10 @@ l_instDivUInt32___closed__1 = _init_l_instDivUInt32___closed__1(); lean_mark_persistent(l_instDivUInt32___closed__1); l_instDivUInt32 = _init_l_instDivUInt32(); lean_mark_persistent(l_instDivUInt32); +l_instLTUInt32__1 = _init_l_instLTUInt32__1(); +lean_mark_persistent(l_instLTUInt32__1); +l_instLEUInt32__1 = _init_l_instLEUInt32__1(); +lean_mark_persistent(l_instLEUInt32__1); l_instComplementUInt32___closed__1 = _init_l_instComplementUInt32___closed__1(); lean_mark_persistent(l_instComplementUInt32___closed__1); l_instComplementUInt32 = _init_l_instComplementUInt32(); diff --git a/stage0/stdlib/Init/Data/Vector/Basic.c b/stage0/stdlib/Init/Data/Vector/Basic.c index 77c6325f56a5..d9cb8ca75d85 100644 --- a/stage0/stdlib/Init/Data/Vector/Basic.c +++ b/stage0/stdlib/Init/Data/Vector/Basic.c @@ -22,7 +22,6 @@ static lean_object* l_Vector_lex___rarg___closed__3; LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Vector_lex___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_uget___rarg(lean_object*, size_t, lean_object*); static lean_object* l_Std_Range_forIn_x27_loop___at_Vector_lex___spec__1___rarg___closed__2; -static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__9; LEAN_EXPORT lean_object* l_Vector_drop___boxed(lean_object*, lean_object*); static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_1091____closed__2; static lean_object* l___private_Init_Data_Vector_Basic_0__reprVector____x40_Init_Data_Vector_Basic___hyg_32____rarg___closed__28; @@ -32,20 +31,20 @@ LEAN_EXPORT uint8_t l_Vector_instBEq___rarg(lean_object*, lean_object*, lean_obj static lean_object* l___private_Init_Data_Vector_Basic_0__reprVector____x40_Init_Data_Vector_Basic___hyg_32____rarg___closed__36; uint8_t l_Array_isEqvAux___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Init_Data_Vector_Basic_0__reprVector____x40_Init_Data_Vector_Basic___hyg_32____rarg___closed__20; +static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__33; static lean_object* l_Vector_term_x23v_x5b___x2c_x5d___closed__19; lean_object* lean_mk_empty_array_with_capacity(lean_object*); -static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__19; +static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__25; static lean_object* l___private_Init_Data_Vector_Basic_0__reprVector____x40_Init_Data_Vector_Basic___hyg_32____rarg___closed__31; LEAN_EXPORT lean_object* l_Vector_map___rarg(lean_object*, lean_object*); static lean_object* l_Vector_term_x23v_x5b___x2c_x5d___closed__9; LEAN_EXPORT lean_object* l_Vector_isEqv(lean_object*); LEAN_EXPORT lean_object* l_Vector_back_x21___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_mkVector___rarg(lean_object*, lean_object*); -static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__5; +static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__8; LEAN_EXPORT lean_object* l_Vector_back_x3f___rarg___boxed(lean_object*); static lean_object* l___private_Init_Data_Vector_Basic_0__reprVector____x40_Init_Data_Vector_Basic___hyg_32____rarg___closed__17; static lean_object* l_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__22; -static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__14; LEAN_EXPORT lean_object* l_Vector_push(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Vector_any___spec__1(lean_object*); static lean_object* l___private_Init_Data_Vector_Basic_0__reprVector____x40_Init_Data_Vector_Basic___hyg_32____rarg___closed__30; @@ -56,17 +55,18 @@ static lean_object* l_Vector___aux__Init__Data__Vector__Basic______macroRules__V LEAN_EXPORT lean_object* l_Vector_eraseIdx_x21___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Vector_lex___spec__1(lean_object*); static lean_object* l___private_Init_Data_Vector_Basic_0__reprVector____x40_Init_Data_Vector_Basic___hyg_32____rarg___closed__37; +LEAN_EXPORT lean_object* l_Vector_foldl(lean_object*, lean_object*, lean_object*); static lean_object* l_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__10; LEAN_EXPORT lean_object* l_Vector_set_x21___rarg___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__6; +static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__24; LEAN_EXPORT lean_object* l_Array_toVector___rarg(lean_object*); +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Vector_foldr___spec__2___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Vector_lex___rarg___closed__2; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Vector_map___spec__1(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Vector_isEqv___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_pop___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_eraseIdx___boxed(lean_object*, lean_object*); static lean_object* l_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__14; -static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__18; LEAN_EXPORT lean_object* l_instReprVector___rarg(lean_object*); lean_object* lean_array_fswap(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_instReprVector___boxed(lean_object*, lean_object*); @@ -79,6 +79,7 @@ LEAN_EXPORT lean_object* l_List_foldl___at___private_Init_Data_Vector_Basic_0__r LEAN_EXPORT lean_object* l_Vector_back___boxed(lean_object*, lean_object*, lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_back_x21(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Vector_foldr___spec__1___rarg(lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Array_toVector(lean_object*); LEAN_EXPORT lean_object* l_Vector_get(lean_object*, lean_object*); lean_object* lean_mk_array(lean_object*, lean_object*); @@ -86,17 +87,22 @@ static lean_object* l___private_Init_Data_Vector_Basic_0__reprVector____x40_Init static lean_object* l___private_Init_Data_Vector_Basic_0__reprVector____x40_Init_Data_Vector_Basic___hyg_32____rarg___closed__25; LEAN_EXPORT lean_object* l_Vector_swap___boxed(lean_object*, lean_object*); uint8_t lean_usize_dec_eq(size_t, size_t); +static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__21; LEAN_EXPORT lean_object* l_Vector_swap___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__19; static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_1091____closed__8; -static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__34; LEAN_EXPORT lean_object* l_Vector_reverse___boxed(lean_object*, lean_object*); lean_object* l_Lean_Syntax_getArgs(lean_object*); +LEAN_EXPORT lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_1599_; static lean_object* l_Vector_term_x23v_x5b___x2c_x5d___closed__13; +static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__19; +LEAN_EXPORT lean_object* l_Vector_foldrM___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_swap(lean_object*, lean_object*, lean_object*); static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_1091____closed__7; static lean_object* l_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__1; +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Vector_foldl___spec__1(lean_object*, lean_object*); lean_object* lean_array_fget(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Vector_foldr___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Vector_term_x23v_x5b___x2c_x5d___closed__14; lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_set_x21___rarg(lean_object*, lean_object*, lean_object*); @@ -105,28 +111,32 @@ LEAN_EXPORT uint8_t l_Array_isEqvAux___at_Vector_instBEq___spec__1___rarg(lean_o LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Vector_any___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Init_Data_Vector_Basic_0__reprVector____x40_Init_Data_Vector_Basic___hyg_32____rarg___closed__4; lean_object* l_Lean_Syntax_TSepArray_getElems___rarg(lean_object*); +static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__4; static lean_object* l_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__25; LEAN_EXPORT uint8_t l_Vector_contains___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_instGetElemNatLt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_all(lean_object*, lean_object*); +static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__31; LEAN_EXPORT lean_object* l___private_Init_Data_Vector_Basic_0__decEqVector____x40_Init_Data_Vector_Basic___hyg_98____boxed(lean_object*, lean_object*); -static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__22; LEAN_EXPORT lean_object* l_Vector_eraseIdx___rarg(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_1443_; -static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__28; static lean_object* l___private_Init_Data_Vector_Basic_0__reprVector____x40_Init_Data_Vector_Basic___hyg_32____rarg___closed__40; LEAN_EXPORT lean_object* l___private_Init_Data_Vector_Basic_0__reprVector____x40_Init_Data_Vector_Basic___hyg_32_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_instLE___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Vector_foldrM(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Init_Data_Vector_Basic_0__reprVector____x40_Init_Data_Vector_Basic___hyg_32____rarg___closed__16; static lean_object* l_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__26; LEAN_EXPORT lean_object* l_Vector_eraseIdx(lean_object*, lean_object*); -static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__21; lean_object* l_Lean_Syntax_node5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Syntax_isOfKind(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_uget___rarg___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__20; LEAN_EXPORT lean_object* l_Vector_singleton___rarg(lean_object*); LEAN_EXPORT lean_object* l_Vector_contains(lean_object*, lean_object*); +static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__27; static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_1091____closed__13; +static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__14; +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Vector_foldr___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__3; LEAN_EXPORT lean_object* l_Vector_all___rarg___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_isEqvAux___at_Vector_instBEq___spec__1(lean_object*); LEAN_EXPORT lean_object* l_Vector_instLT___boxed(lean_object*, lean_object*, lean_object*); @@ -134,11 +144,12 @@ LEAN_EXPORT lean_object* l_Vector_instGetElemNatLt___boxed(lean_object*, lean_ob static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_1091____closed__14; LEAN_EXPORT lean_object* l_Vector_any___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_indexOf_x3f(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_1929_; LEAN_EXPORT lean_object* l___private_Init_Data_Vector_Basic_0__reprVector____x40_Init_Data_Vector_Basic___hyg_32____rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_anyM___rarg(lean_object*, lean_object*, lean_object*); +static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__5; LEAN_EXPORT lean_object* l_Vector_extract(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_tail___rarg(lean_object*, lean_object*); -static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__4; LEAN_EXPORT lean_object* l_Vector_setIfInBounds(lean_object*, lean_object*); static lean_object* l___private_Init_Data_Vector_Basic_0__reprVector____x40_Init_Data_Vector_Basic___hyg_32____rarg___closed__34; LEAN_EXPORT lean_object* l_Vector_get___rarg(lean_object*, lean_object*); @@ -148,12 +159,12 @@ static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_1091____closed_ LEAN_EXPORT lean_object* l_Vector_cast___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Vector_all___spec__1(lean_object*); LEAN_EXPORT uint8_t l_Vector_isPrefixOf___rarg(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Vector_foldl___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_mkEmpty___rarg(lean_object*); LEAN_EXPORT lean_object* l_Vector_elimAsArray___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Vector_Basic_0__decEqVector____x40_Init_Data_Vector_Basic___hyg_98_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_instHAppendHAddNat___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_head___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__31; lean_object* l_Array_range___lambda__1___boxed(lean_object*); size_t lean_usize_of_nat(lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Vector_allM___spec__1___rarg___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -164,36 +175,37 @@ LEAN_EXPORT lean_object* l_Vector_swapIfInBounds___rarg___boxed(lean_object*, le static lean_object* l_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__8; static lean_object* l_Vector_swapAt_x21___rarg___closed__3; LEAN_EXPORT lean_object* l_instReprVector(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Vector_foldr(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_back(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Init_Data_Vector_Basic_0__reprVector____x40_Init_Data_Vector_Basic___hyg_32____rarg___closed__8; lean_object* l_Array_reverse___rarg(lean_object*); -static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__1; -static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__15; LEAN_EXPORT lean_object* l_Vector_getD___rarg___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Init_Data_Vector_Basic_0__reprVector____x40_Init_Data_Vector_Basic___hyg_32____rarg___closed__7; +static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__9; LEAN_EXPORT lean_object* l_Vector_pop(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_uget___boxed(lean_object*, lean_object*); +static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__29; LEAN_EXPORT lean_object* l_Vector_extract___rarg(lean_object*, lean_object*, lean_object*); +static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__13; LEAN_EXPORT lean_object* l_Vector_allM(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Vector_Basic_0__reprVector____x40_Init_Data_Vector_Basic___hyg_32____rarg___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Vector_foldl___rarg(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_SourceInfo_fromRef(lean_object*, uint8_t); -static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__20; LEAN_EXPORT lean_object* l_Vector_cast___rarg(lean_object*); lean_object* l_Array_zipWithAux___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Vector_foldr___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_get___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_back_x3f___boxed(lean_object*, lean_object*); lean_object* lean_nat_to_int(lean_object*); LEAN_EXPORT uint8_t l_Vector_all___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_isPrefixOf___boxed(lean_object*, lean_object*, lean_object*); uint8_t l_Array_isPrefixOf___rarg(lean_object*, lean_object*, lean_object*); -static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__32; static lean_object* l_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__16; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Vector_map___spec__1___rarg(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Vector_contains___boxed(lean_object*, lean_object*); lean_object* l_Array_ofFn___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_all___boxed(lean_object*, lean_object*); static lean_object* l_Vector_term_x23v_x5b___x2c_x5d___closed__4; -static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__23; LEAN_EXPORT lean_object* l_Vector_cast___rarg___boxed(lean_object*); LEAN_EXPORT lean_object* l_Vector_extract___boxed(lean_object*, lean_object*); static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_1091____closed__15; @@ -206,10 +218,14 @@ static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_1091____closed_ LEAN_EXPORT lean_object* l_Vector_head___rarg___boxed(lean_object*); static lean_object* l_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__9; LEAN_EXPORT lean_object* l_Vector_isPrefixOf___rarg___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Vector_foldr___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_swap___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_1091____closed__10; +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Vector_foldl___spec__1___rarg(lean_object*, lean_object*, size_t, size_t, lean_object*); lean_object* l_Array_back_x3f___rarg(lean_object*); -static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__27; +static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__34; +static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__18; +static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__30; LEAN_EXPORT lean_object* l_panic___at_Vector_swapAt_x21___spec__1___boxed(lean_object*, lean_object*); lean_object* lean_array_pop(lean_object*); LEAN_EXPORT lean_object* l_Vector_reverse(lean_object*, lean_object*); @@ -217,17 +233,16 @@ static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_1091____closed_ static lean_object* l_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__27; LEAN_EXPORT lean_object* l_Vector_zipWith___rarg(lean_object*, lean_object*, lean_object*); lean_object* lean_array_to_list(lean_object*); +static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__6; static lean_object* l___private_Init_Data_Vector_Basic_0__reprVector____x40_Init_Data_Vector_Basic___hyg_32____rarg___closed__9; -static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__17; static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_1091____closed__16; -static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__25; +static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__7; lean_object* l_Lean_Syntax_node3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_map(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Vector_allM___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_take___rarg___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_getD(lean_object*, lean_object*); static lean_object* l_Vector_term_x23v_x5b___x2c_x5d___closed__8; -static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__24; LEAN_EXPORT lean_object* l_Vector_anyM___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_set_x21(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_setIfInBounds___rarg___boxed(lean_object*, lean_object*, lean_object*); @@ -235,7 +250,6 @@ LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Vector_all___spec__1___rarg static lean_object* l___private_Init_Data_Vector_Basic_0__reprVector____x40_Init_Data_Vector_Basic___hyg_32____rarg___closed__18; static lean_object* l_Vector_lex___rarg___closed__1; LEAN_EXPORT lean_object* l_Vector_back_x3f(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_1517_; LEAN_EXPORT lean_object* l_Vector_append(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at_Vector_swapAt_x21___spec__1___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_lex(lean_object*); @@ -243,6 +257,7 @@ LEAN_EXPORT lean_object* l_Vector_swapAt_x21(lean_object*); static lean_object* l_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__18; static lean_object* l_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__20; lean_object* l_Array_anyMUnsafe_any___rarg(lean_object*, lean_object*, lean_object*, size_t, size_t); +LEAN_EXPORT lean_object* l_Vector_foldrM___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Init_Data_Vector_Basic_0__reprVector____x40_Init_Data_Vector_Basic___hyg_32____rarg___closed__19; LEAN_EXPORT uint8_t l_Vector_lex___rarg___lambda__1(lean_object*); LEAN_EXPORT lean_object* l_Vector_ofFn___boxed(lean_object*, lean_object*, lean_object*); @@ -253,10 +268,14 @@ static lean_object* l_Vector_eraseIdx_x21___rarg___closed__4; LEAN_EXPORT lean_object* l_Vector_elimAsList___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_instInhabited___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_get___rarg___boxed(lean_object*, lean_object*); +static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__26; lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Vector_foldl___rarg___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Vector_foldr___rarg(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_node2(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Vector_swapAt_x21___rarg___closed__1; LEAN_EXPORT lean_object* l_Vector_drop___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2352_; static lean_object* l_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__13; LEAN_EXPORT lean_object* l_Vector_elimAsArray___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_lex___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -265,25 +284,26 @@ lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_eraseIdx_x21(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_Vector_lex___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__12; static lean_object* l_Vector_swapAt_x21___rarg___closed__4; LEAN_EXPORT lean_object* l_Vector_back_x3f___rarg(lean_object*); static lean_object* l___private_Init_Data_Vector_Basic_0__reprVector____x40_Init_Data_Vector_Basic___hyg_32____rarg___closed__26; static lean_object* l___private_Init_Data_Vector_Basic_0__reprVector____x40_Init_Data_Vector_Basic___hyg_32____rarg___closed__5; LEAN_EXPORT lean_object* l_Vector_swapAt_x21___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Vector_eraseIdx_x21___rarg___closed__2; -static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__30; LEAN_EXPORT lean_object* l_Vector_set___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__17; lean_object* l_Array_indexOfAux___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__29; -static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__16; +LEAN_EXPORT lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_1591_; static lean_object* l_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__7; LEAN_EXPORT lean_object* l_Vector_back_x21___rarg___boxed(lean_object*, lean_object*); lean_object* l_Array_eraseIdx___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_take___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_any(lean_object*, lean_object*); static lean_object* l_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__3; +static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__16; LEAN_EXPORT lean_object* l_Vector_indexOf_x3f___rarg___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Vector_foldlM___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_head___rarg(lean_object*); LEAN_EXPORT lean_object* l_Vector_zipWith___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_swapAt___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -296,24 +316,25 @@ LEAN_EXPORT lean_object* l_Vector_uget(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_swapAt___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_swapAt_x21___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Init_Data_Vector_Basic_0__reprVector____x40_Init_Data_Vector_Basic___hyg_32____rarg___closed__6; -LEAN_EXPORT lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_1435_; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Vector_allM___spec__1___rarg___lambda__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at_Vector_eraseIdx_x21___spec__1___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Vector_foldr___spec__2___rarg(lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Vector_instLE(lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_x27_loop___at_Vector_lex___spec__1___rarg___closed__3; LEAN_EXPORT lean_object* l___private_Init_Data_Vector_Basic_0__decEqVector____x40_Init_Data_Vector_Basic___hyg_98____rarg___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Init_Data_Vector_Basic_0__reprVector____x40_Init_Data_Vector_Basic___hyg_32____rarg___closed__3; -static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__8; lean_object* l_Array_extract___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Vector_allM___spec__1___rarg___lambda__1(lean_object*, uint8_t); static lean_object* l___private_Init_Data_Vector_Basic_0__reprVector____x40_Init_Data_Vector_Basic___hyg_32____rarg___closed__39; +static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__22; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Vector_allM___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l_Vector_extract___rarg___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Vector_foldl___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_lex___rarg___lambda__1___boxed(lean_object*); static lean_object* l_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__2; LEAN_EXPORT lean_object* l_Vector_getD___boxed(lean_object*, lean_object*); -static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__13; LEAN_EXPORT lean_object* l_Vector_instGetElemNatLt___rarg___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__11; LEAN_EXPORT lean_object* l_Vector_take(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_instBEq___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_push___boxed(lean_object*, lean_object*); @@ -323,26 +344,28 @@ static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_1091____closed_ static lean_object* l_Vector_eraseIdx_x21___rarg___closed__3; static lean_object* l_Vector_swapAt_x21___rarg___closed__2; static lean_object* l_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__31; +static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__2; LEAN_EXPORT lean_object* l_Array_isEqvAux___at_Vector_instBEq___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_Vector_any___rarg___boxed(lean_object*, lean_object*); static lean_object* l___private_Init_Data_Vector_Basic_0__reprVector____x40_Init_Data_Vector_Basic___hyg_32____rarg___closed__22; +LEAN_EXPORT lean_object* l_Vector_foldlM___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__29; LEAN_EXPORT lean_object* l_panic___at_Vector_swapAt_x21___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Format_joinSep___at___private_Init_Data_Vector_Basic_0__reprVector____x40_Init_Data_Vector_Basic___hyg_32____spec__1(lean_object*); lean_object* lean_string_length(lean_object*); static lean_object* l_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__21; LEAN_EXPORT lean_object* l_Vector_allM___rarg(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2196_; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); -static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__7; uint8_t l_Array_contains___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Init_Data_Vector_Basic_0__reprVector____x40_Init_Data_Vector_Basic___hyg_32____rarg___closed__23; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); static lean_object* l_Vector_instHAppendHAddNat___closed__1; static lean_object* l___private_Init_Data_Vector_Basic_0__reprVector____x40_Init_Data_Vector_Basic___hyg_32____rarg___closed__33; +LEAN_EXPORT lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_1673_; LEAN_EXPORT lean_object* l_Vector_isPrefixOf(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_swapIfInBounds(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Vector_foldlM(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Vector_term_x23v_x5b___x2c_x5d___closed__10; static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_1091____closed__6; static lean_object* l_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__23; @@ -351,10 +374,10 @@ static lean_object* l_Vector___aux__Init__Data__Vector__Basic______macroRules__V LEAN_EXPORT lean_object* l_Array_toVector___rarg___boxed(lean_object*); LEAN_EXPORT lean_object* l_Vector_cast(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_1091____closed__17; +static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__1; lean_object* lean_array_set(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_swapAt___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Init_Data_Vector_Basic_0__reprVector____x40_Init_Data_Vector_Basic___hyg_32____rarg___closed__21; -static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__3; lean_object* lean_panic_fn(lean_object*, lean_object*); uint8_t l_Array_instDecidableEq___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_1091____closed__4; @@ -365,6 +388,7 @@ lean_object* l_Array_back_x21___rarg(lean_object*, lean_object*); static lean_object* l___private_Init_Data_Vector_Basic_0__reprVector____x40_Init_Data_Vector_Basic___hyg_32____rarg___closed__38; static lean_object* l_Vector_range___closed__1; LEAN_EXPORT lean_object* l_Vector_append___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__10; static lean_object* l_Vector_term_x23v_x5b___x2c_x5d___closed__7; LEAN_EXPORT lean_object* l_Vector_tail___rarg___boxed(lean_object*, lean_object*); static lean_object* l_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__5; @@ -372,7 +396,7 @@ LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Vector_all___spec__1___rarg(lean static lean_object* l___private_Init_Data_Vector_Basic_0__reprVector____x40_Init_Data_Vector_Basic___hyg_32____rarg___closed__12; LEAN_EXPORT lean_object* l_Vector_mkEmpty___rarg___boxed(lean_object*); LEAN_EXPORT lean_object* l_Vector_set___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__12; +size_t lean_usize_sub(size_t, size_t); lean_object* lean_array_mk(lean_object*); static lean_object* l_Vector_term_x23v_x5b___x2c_x5d___closed__15; LEAN_EXPORT lean_object* l_panic___at_Vector_eraseIdx_x21___spec__1(lean_object*, lean_object*); @@ -385,22 +409,23 @@ LEAN_EXPORT lean_object* l_Vector_head(lean_object*, lean_object*, lean_object*) LEAN_EXPORT lean_object* l_Vector_allM___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Init_Data_Vector_Basic_0__reprVector____x40_Init_Data_Vector_Basic___hyg_32____rarg___closed__35; LEAN_EXPORT lean_object* l_instDecidableEqVector(lean_object*, lean_object*); +lean_object* l_Array_foldrMUnsafe_fold___rarg(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); +static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__23; LEAN_EXPORT lean_object* l_Vector_back___rarg___boxed(lean_object*); lean_object* lean_array_uget(lean_object*, size_t); static lean_object* l___private_Init_Data_Vector_Basic_0__reprVector____x40_Init_Data_Vector_Basic___hyg_32____rarg___closed__27; LEAN_EXPORT lean_object* l_Vector_reverse___rarg(lean_object*); size_t lean_array_size(lean_object*); +lean_object* l_Array_foldlMUnsafe_fold___rarg(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_instDecidableEqVector___boxed(lean_object*, lean_object*); static lean_object* l___private_Init_Data_Vector_Basic_0__reprVector____x40_Init_Data_Vector_Basic___hyg_32____rarg___closed__11; -static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__26; LEAN_EXPORT lean_object* l_Vector_instInhabited(lean_object*); static lean_object* l_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__4; LEAN_EXPORT lean_object* l_Vector_swapIfInBounds___boxed(lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_1091____closed__9; LEAN_EXPORT lean_object* l_Vector_swapIfInBounds___rarg(lean_object*, lean_object*, lean_object*); -static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__11; -static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__10; +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Vector_foldr___spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_set_x21___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_drop(lean_object*, lean_object*); lean_object* lean_string_append(lean_object*, lean_object*); @@ -410,28 +435,31 @@ LEAN_EXPORT lean_object* l_Vector_zipWith___rarg___boxed(lean_object*, lean_obje static lean_object* l_Vector_eraseIdx_x21___rarg___closed__1; LEAN_EXPORT lean_object* l_Vector_instMembership___boxed(lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); -LEAN_EXPORT lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_1773_; LEAN_EXPORT lean_object* l_Vector_indexOf_x3f___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Vector_Basic_0__reprVector____x40_Init_Data_Vector_Basic___hyg_32____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_append___rarg___boxed(lean_object*, lean_object*); static lean_object* l_Vector_term_x23v_x5b___x2c_x5d___closed__5; LEAN_EXPORT lean_object* l_Vector_set(lean_object*, lean_object*); -static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__33; +uint8_t lean_nat_dec_le(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_swap(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_instGetElemNatLt___rarg(lean_object*, lean_object*, lean_object*); +static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__15; uint8_t lean_usize_dec_lt(size_t, size_t); LEAN_EXPORT lean_object* l_Vector_ofFn(lean_object*, lean_object*, lean_object*); static lean_object* l_Vector_term_x23v_x5b___x2c_x5d___closed__1; static lean_object* l_Vector_term_x23v_x5b___x2c_x5d___closed__3; static lean_object* l___private_Init_Data_Vector_Basic_0__reprVector____x40_Init_Data_Vector_Basic___hyg_32____rarg___closed__15; static lean_object* l_Std_Range_forIn_x27_loop___at_Vector_lex___spec__1___rarg___closed__4; +static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__28; lean_object* lean_nat_add(lean_object*, lean_object*); +static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__32; LEAN_EXPORT lean_object* l_Vector_setIfInBounds___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_Vector_term_x23v_x5b___x2c_x5d___closed__18; LEAN_EXPORT lean_object* l_Std_Format_joinSep___at___private_Init_Data_Vector_Basic_0__reprVector____x40_Init_Data_Vector_Basic___hyg_32____spec__1___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_zipWith(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__12; LEAN_EXPORT lean_object* l_Vector_range(lean_object*); +static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__17; static lean_object* l_Vector_term_x23v_x5b___x2c_x5d___closed__17; LEAN_EXPORT lean_object* l_Vector_mkEmpty(lean_object*); LEAN_EXPORT lean_object* l_Vector_eraseIdx_x21___rarg___boxed(lean_object*, lean_object*, lean_object*); @@ -443,7 +471,6 @@ LEAN_EXPORT lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_1091_; lean_object* lean_array_uset(lean_object*, size_t, lean_object*); lean_object* l___private_Init_Data_Repr_0__Nat_reprFast(lean_object*); LEAN_EXPORT lean_object* l_Vector_anyM(lean_object*, lean_object*, lean_object*); -static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__2; static lean_object* l___private_Init_Data_Vector_Basic_0__reprVector____x40_Init_Data_Vector_Basic___hyg_32____rarg___closed__13; LEAN_EXPORT lean_object* l_Vector_elimAsList(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_mkVector(lean_object*); @@ -2327,7 +2354,7 @@ LEAN_EXPORT lean_object* l_Vector_push___rarg(lean_object* x_1, lean_object* x_2 _start: { lean_object* x_3; -x_3 = lean_array_push(x_2, x_1); +x_3 = lean_array_push(x_1, x_2); return x_3; } } @@ -2677,6 +2704,398 @@ lean_dec(x_2); return x_3; } } +LEAN_EXPORT lean_object* l_Vector_foldlM___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; uint8_t x_7; +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_object* x_9; lean_object* x_10; +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +x_8 = lean_ctor_get(x_1, 0); +lean_inc(x_8); +lean_dec(x_1); +x_9 = lean_ctor_get(x_8, 1); +lean_inc(x_9); +lean_dec(x_8); +x_10 = lean_apply_2(x_9, lean_box(0), x_3); +return x_10; +} +else +{ +uint8_t x_11; +x_11 = lean_nat_dec_le(x_5, x_5); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +x_12 = lean_ctor_get(x_1, 0); +lean_inc(x_12); +lean_dec(x_1); +x_13 = lean_ctor_get(x_12, 1); +lean_inc(x_13); +lean_dec(x_12); +x_14 = lean_apply_2(x_13, lean_box(0), x_3); +return x_14; +} +else +{ +size_t x_15; size_t x_16; lean_object* x_17; +x_15 = 0; +x_16 = lean_usize_of_nat(x_5); +lean_dec(x_5); +x_17 = l_Array_foldlMUnsafe_fold___rarg(x_1, x_2, x_4, x_15, x_16, x_3); +return x_17; +} +} +} +} +LEAN_EXPORT lean_object* l_Vector_foldlM(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_Vector_foldlM___rarg), 4, 0); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Vector_foldlM___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_Vector_foldlM(x_1, x_2, x_3, x_4); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Vector_foldrM___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; uint8_t x_7; +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_object* x_9; lean_object* x_10; +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +x_8 = lean_ctor_get(x_1, 0); +lean_inc(x_8); +lean_dec(x_1); +x_9 = lean_ctor_get(x_8, 1); +lean_inc(x_9); +lean_dec(x_8); +x_10 = lean_apply_2(x_9, lean_box(0), x_3); +return x_10; +} +else +{ +size_t x_11; size_t x_12; lean_object* x_13; +x_11 = lean_usize_of_nat(x_5); +lean_dec(x_5); +x_12 = 0; +x_13 = l_Array_foldrMUnsafe_fold___rarg(x_1, x_2, x_4, x_11, x_12, x_3); +return x_13; +} +} +} +LEAN_EXPORT lean_object* l_Vector_foldrM(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_Vector_foldrM___rarg), 4, 0); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Vector_foldrM___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_Vector_foldrM(x_1, x_2, x_3, x_4); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Vector_foldl___spec__1___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; lean_object* x_8; size_t x_9; size_t x_10; +x_7 = lean_array_uget(x_2, x_3); +lean_inc(x_1); +x_8 = lean_apply_2(x_1, x_5, x_7); +x_9 = 1; +x_10 = lean_usize_add(x_3, x_9); +x_3 = x_10; +x_5 = x_8; +goto _start; +} +else +{ +lean_dec(x_1); +return x_5; +} +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Vector_foldl___spec__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Array_foldlMUnsafe_fold___at_Vector_foldl___spec__1___rarg___boxed), 5, 0); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Vector_foldl___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_3); +x_5 = lean_unsigned_to_nat(0u); +x_6 = lean_nat_dec_lt(x_5, x_4); +if (x_6 == 0) +{ +lean_dec(x_4); +lean_dec(x_1); +return x_2; +} +else +{ +uint8_t x_7; +x_7 = lean_nat_dec_le(x_4, x_4); +if (x_7 == 0) +{ +lean_dec(x_4); +lean_dec(x_1); +return x_2; +} +else +{ +size_t x_8; size_t x_9; lean_object* x_10; +x_8 = 0; +x_9 = lean_usize_of_nat(x_4); +lean_dec(x_4); +x_10 = l_Array_foldlMUnsafe_fold___at_Vector_foldl___spec__1___rarg(x_1, x_3, x_8, x_9, x_2); +return x_10; +} +} +} +} +LEAN_EXPORT lean_object* l_Vector_foldl(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = lean_alloc_closure((void*)(l_Vector_foldl___rarg___boxed), 3, 0); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Vector_foldl___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: +{ +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_Vector_foldl___spec__1___rarg(x_1, x_2, x_6, x_7, x_5); +lean_dec(x_2); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Vector_foldl___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Vector_foldl___rarg(x_1, x_2, x_3); +lean_dec(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Vector_foldl___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Vector_foldl(x_1, x_2, x_3); +lean_dec(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Vector_foldr___spec__1___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) +{ +size_t x_7; size_t x_8; lean_object* x_9; lean_object* x_10; +x_7 = 1; +x_8 = lean_usize_sub(x_3, x_7); +x_9 = lean_array_uget(x_2, x_8); +lean_inc(x_1); +x_10 = lean_apply_2(x_1, x_9, x_5); +x_3 = x_8; +x_5 = x_10; +goto _start; +} +else +{ +lean_dec(x_1); +return x_5; +} +} +} +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Vector_foldr___spec__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Array_foldrMUnsafe_fold___at_Vector_foldr___spec__1___rarg___boxed), 5, 0); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Vector_foldr___spec__2___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) +{ +size_t x_7; size_t x_8; lean_object* x_9; lean_object* x_10; +x_7 = 1; +x_8 = lean_usize_sub(x_3, x_7); +x_9 = lean_array_uget(x_2, x_8); +lean_inc(x_1); +x_10 = lean_apply_2(x_1, x_9, x_5); +x_3 = x_8; +x_5 = x_10; +goto _start; +} +else +{ +lean_dec(x_1); +return x_5; +} +} +} +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Vector_foldr___spec__2(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Array_foldrMUnsafe_fold___at_Vector_foldr___spec__2___rarg___boxed), 5, 0); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Vector_foldr___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_array_get_size(x_3); +x_5 = lean_nat_dec_le(x_4, x_4); +if (x_5 == 0) +{ +lean_object* x_6; uint8_t x_7; +x_6 = lean_unsigned_to_nat(0u); +x_7 = lean_nat_dec_lt(x_6, x_4); +if (x_7 == 0) +{ +lean_dec(x_4); +lean_dec(x_1); +return x_2; +} +else +{ +size_t x_8; size_t x_9; lean_object* x_10; +x_8 = lean_usize_of_nat(x_4); +lean_dec(x_4); +x_9 = 0; +x_10 = l_Array_foldrMUnsafe_fold___at_Vector_foldr___spec__1___rarg(x_1, x_3, x_8, x_9, x_2); +return x_10; +} +} +else +{ +lean_object* x_11; uint8_t x_12; +x_11 = lean_unsigned_to_nat(0u); +x_12 = lean_nat_dec_lt(x_11, x_4); +if (x_12 == 0) +{ +lean_dec(x_4); +lean_dec(x_1); +return x_2; +} +else +{ +size_t x_13; size_t x_14; lean_object* x_15; +x_13 = lean_usize_of_nat(x_4); +lean_dec(x_4); +x_14 = 0; +x_15 = l_Array_foldrMUnsafe_fold___at_Vector_foldr___spec__2___rarg(x_1, x_3, x_13, x_14, x_2); +return x_15; +} +} +} +} +LEAN_EXPORT lean_object* l_Vector_foldr(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = lean_alloc_closure((void*)(l_Vector_foldr___rarg___boxed), 3, 0); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Vector_foldr___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: +{ +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_foldrMUnsafe_fold___at_Vector_foldr___spec__1___rarg(x_1, x_2, x_6, x_7, x_5); +lean_dec(x_2); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Vector_foldr___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; 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_foldrMUnsafe_fold___at_Vector_foldr___spec__2___rarg(x_1, x_2, x_6, x_7, x_5); +lean_dec(x_2); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Vector_foldr___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Vector_foldr___rarg(x_1, x_2, x_3); +lean_dec(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Vector_foldr___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Vector_foldr(x_1, x_2, x_3); +lean_dec(x_3); +return x_4; +} +} LEAN_EXPORT lean_object* l_Vector_append___rarg(lean_object* x_1, lean_object* x_2) { _start: { @@ -2944,7 +3363,7 @@ lean_dec(x_1); return x_4; } } -static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_1435_() { +static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_1591_() { _start: { lean_object* x_1; @@ -2952,7 +3371,7 @@ x_1 = l___auto____x40_Init_Data_Vector_Basic___hyg_1091____closed__17; return x_1; } } -static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_1443_() { +static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_1599_() { _start: { lean_object* x_1; @@ -3030,7 +3449,7 @@ lean_dec(x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_1517_() { +static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_1673_() { _start: { lean_object* x_1; @@ -3450,7 +3869,7 @@ lean_dec(x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_1773_() { +static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_1929_() { _start: { lean_object* x_1; @@ -3530,7 +3949,7 @@ static lean_object* _init_l_Vector_eraseIdx_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_Vector_eraseIdx_x21___rarg___closed__1; x_2 = l_Vector_eraseIdx_x21___rarg___closed__2; -x_3 = lean_unsigned_to_nat(249u); +x_3 = lean_unsigned_to_nat(261u); x_4 = lean_unsigned_to_nat(4u); x_5 = l_Vector_eraseIdx_x21___rarg___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -4256,7 +4675,7 @@ lean_dec(x_2); return x_4; } } -static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__1() { +static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__1() { _start: { lean_object* x_1; @@ -4264,41 +4683,41 @@ x_1 = lean_mk_string_unchecked("exact", 5, 5); return x_1; } } -static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__2() { +static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____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_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__1; x_2 = l_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__2; x_3 = l___auto____x40_Init_Data_Vector_Basic___hyg_1091____closed__1; -x_4 = l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__1; +x_4 = l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__1; 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_Data_Vector_Basic___hyg_2196____closed__3() { +static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(2); -x_2 = l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__1; +x_2 = l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__1; 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_Data_Vector_Basic___hyg_2196____closed__4() { +static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__26; -x_2 = l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__3; +x_2 = l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__3; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__5() { +static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__5() { _start: { lean_object* x_1; @@ -4306,19 +4725,19 @@ x_1 = lean_mk_string_unchecked("paren", 5, 5); return x_1; } } -static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__6() { +static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____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_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__1; x_2 = l_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__2; x_3 = l_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__3; -x_4 = l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__5; +x_4 = l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__5; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__7() { +static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -4330,17 +4749,17 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__8() { +static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__26; -x_2 = l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__7; +x_2 = l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__7; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__9() { +static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__9() { _start: { lean_object* x_1; @@ -4348,17 +4767,17 @@ x_1 = lean_mk_string_unchecked("term_<_", 7, 7); return x_1; } } -static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__10() { +static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____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_Data_Vector_Basic___hyg_2196____closed__9; +x_2 = l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__9; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__11() { +static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__11() { _start: { lean_object* x_1; @@ -4366,19 +4785,19 @@ x_1 = lean_mk_string_unchecked("cdot", 4, 4); return x_1; } } -static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__12() { +static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____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_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__1; x_2 = l_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__2; x_3 = l_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__3; -x_4 = l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__11; +x_4 = l___auto____x40_Init_Data_Vector_Basic___hyg_2352____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_Data_Vector_Basic___hyg_2196____closed__13() { +static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__13() { _start: { lean_object* x_1; @@ -4386,35 +4805,35 @@ x_1 = lean_mk_string_unchecked("·", 2, 1); return x_1; } } -static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__14() { +static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(2); -x_2 = l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__13; +x_2 = l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__13; x_3 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__15() { +static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__26; -x_2 = l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__14; +x_2 = l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__14; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__16() { +static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__16() { _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_Data_Vector_Basic___hyg_2196____closed__12; -x_3 = l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__15; +x_2 = l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__12; +x_3 = l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__15; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -4422,17 +4841,17 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__17() { +static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__26; -x_2 = l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__16; +x_2 = l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__16; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__18() { +static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__18() { _start: { lean_object* x_1; @@ -4440,45 +4859,45 @@ x_1 = lean_mk_string_unchecked("<", 1, 1); return x_1; } } -static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__19() { +static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__19() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(2); -x_2 = l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__18; +x_2 = l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__18; x_3 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__20() { +static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__20() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__17; -x_2 = l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__19; +x_1 = l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__17; +x_2 = l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__19; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__21() { +static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__21() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__20; -x_2 = l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__16; +x_1 = l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__20; +x_2 = l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__16; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__22() { +static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____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_Init_Data_Vector_Basic___hyg_2196____closed__10; -x_3 = l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__21; +x_2 = l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__10; +x_3 = l___auto____x40_Init_Data_Vector_Basic___hyg_2352____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); @@ -4486,17 +4905,17 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__23() { +static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__23() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__8; -x_2 = l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__22; +x_1 = l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__8; +x_2 = l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__22; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__24() { +static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__24() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -4508,23 +4927,23 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__25() { +static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__25() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__23; -x_2 = l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__24; +x_1 = l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__23; +x_2 = l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__24; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__26() { +static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____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_Init_Data_Vector_Basic___hyg_2196____closed__6; -x_3 = l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__25; +x_2 = l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__6; +x_3 = l___auto____x40_Init_Data_Vector_Basic___hyg_2352____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); @@ -4532,23 +4951,23 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__27() { +static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__27() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__4; -x_2 = l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__26; +x_1 = l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__4; +x_2 = l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__26; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__28() { +static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____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_Init_Data_Vector_Basic___hyg_2196____closed__2; -x_3 = l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__27; +x_2 = l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__2; +x_3 = l___auto____x40_Init_Data_Vector_Basic___hyg_2352____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); @@ -4556,23 +4975,23 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__29() { +static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__29() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__26; -x_2 = l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__28; +x_2 = l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__28; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__30() { +static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__30() { _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_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__15; -x_3 = l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__29; +x_3 = l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__29; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -4580,23 +4999,23 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__31() { +static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__31() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__26; -x_2 = l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__30; +x_2 = l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__30; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__32() { +static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__32() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); x_2 = l___auto____x40_Init_Data_Vector_Basic___hyg_1091____closed__5; -x_3 = l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__31; +x_3 = l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__31; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -4604,23 +5023,23 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__33() { +static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__33() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__26; -x_2 = l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__32; +x_2 = l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__32; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__34() { +static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__34() { _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_Data_Vector_Basic___hyg_1091____closed__3; -x_3 = l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__33; +x_3 = l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__33; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -4628,11 +5047,11 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196_() { +static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352_() { _start: { lean_object* x_1; -x_1 = l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__34; +x_1 = l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__34; return x_1; } } @@ -5116,12 +5535,12 @@ l_Vector_instHAppendHAddNat___closed__1 = _init_l_Vector_instHAppendHAddNat___cl lean_mark_persistent(l_Vector_instHAppendHAddNat___closed__1); l_Vector_zipWith___rarg___closed__1 = _init_l_Vector_zipWith___rarg___closed__1(); lean_mark_persistent(l_Vector_zipWith___rarg___closed__1); -l___auto____x40_Init_Data_Vector_Basic___hyg_1435_ = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_1435_(); -lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_1435_); -l___auto____x40_Init_Data_Vector_Basic___hyg_1443_ = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_1443_(); -lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_1443_); -l___auto____x40_Init_Data_Vector_Basic___hyg_1517_ = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_1517_(); -lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_1517_); +l___auto____x40_Init_Data_Vector_Basic___hyg_1591_ = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_1591_(); +lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_1591_); +l___auto____x40_Init_Data_Vector_Basic___hyg_1599_ = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_1599_(); +lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_1599_); +l___auto____x40_Init_Data_Vector_Basic___hyg_1673_ = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_1673_(); +lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_1673_); l_Vector_swapAt_x21___rarg___closed__1 = _init_l_Vector_swapAt_x21___rarg___closed__1(); lean_mark_persistent(l_Vector_swapAt_x21___rarg___closed__1); l_Vector_swapAt_x21___rarg___closed__2 = _init_l_Vector_swapAt_x21___rarg___closed__2(); @@ -5132,8 +5551,8 @@ l_Vector_swapAt_x21___rarg___closed__4 = _init_l_Vector_swapAt_x21___rarg___clos lean_mark_persistent(l_Vector_swapAt_x21___rarg___closed__4); l_Vector_range___closed__1 = _init_l_Vector_range___closed__1(); lean_mark_persistent(l_Vector_range___closed__1); -l___auto____x40_Init_Data_Vector_Basic___hyg_1773_ = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_1773_(); -lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_1773_); +l___auto____x40_Init_Data_Vector_Basic___hyg_1929_ = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_1929_(); +lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_1929_); l_Vector_eraseIdx_x21___rarg___closed__1 = _init_l_Vector_eraseIdx_x21___rarg___closed__1(); lean_mark_persistent(l_Vector_eraseIdx_x21___rarg___closed__1); l_Vector_eraseIdx_x21___rarg___closed__2 = _init_l_Vector_eraseIdx_x21___rarg___closed__2(); @@ -5142,76 +5561,76 @@ l_Vector_eraseIdx_x21___rarg___closed__3 = _init_l_Vector_eraseIdx_x21___rarg___ lean_mark_persistent(l_Vector_eraseIdx_x21___rarg___closed__3); l_Vector_eraseIdx_x21___rarg___closed__4 = _init_l_Vector_eraseIdx_x21___rarg___closed__4(); lean_mark_persistent(l_Vector_eraseIdx_x21___rarg___closed__4); -l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__1 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__1(); -lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__1); -l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__2 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__2(); -lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__2); -l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__3 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__3(); -lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__3); -l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__4 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__4(); -lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__4); -l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__5 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__5(); -lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__5); -l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__6 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__6(); -lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__6); -l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__7 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__7(); -lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__7); -l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__8 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__8(); -lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__8); -l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__9 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__9(); -lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__9); -l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__10 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__10(); -lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__10); -l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__11 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__11(); -lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__11); -l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__12 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__12(); -lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__12); -l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__13 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__13(); -lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__13); -l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__14 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__14(); -lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__14); -l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__15 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__15(); -lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__15); -l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__16 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__16(); -lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__16); -l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__17 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__17(); -lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__17); -l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__18 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__18(); -lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__18); -l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__19 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__19(); -lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__19); -l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__20 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__20(); -lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__20); -l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__21 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__21(); -lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__21); -l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__22 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__22(); -lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__22); -l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__23 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__23(); -lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__23); -l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__24 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__24(); -lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__24); -l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__25 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__25(); -lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__25); -l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__26 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__26(); -lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__26); -l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__27 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__27(); -lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__27); -l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__28 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__28(); -lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__28); -l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__29 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__29(); -lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__29); -l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__30 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__30(); -lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__30); -l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__31 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__31(); -lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__31); -l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__32 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__32(); -lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__32); -l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__33 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__33(); -lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__33); -l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__34 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__34(); -lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__34); -l___auto____x40_Init_Data_Vector_Basic___hyg_2196_ = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196_(); -lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2196_); +l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__1 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__1(); +lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__1); +l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__2 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__2(); +lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__2); +l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__3 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__3(); +lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__3); +l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__4 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__4(); +lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__4); +l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__5 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__5(); +lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__5); +l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__6 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__6(); +lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__6); +l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__7 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__7(); +lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__7); +l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__8 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__8(); +lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__8); +l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__9 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__9(); +lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__9); +l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__10 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__10(); +lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__10); +l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__11 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__11(); +lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__11); +l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__12 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__12(); +lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__12); +l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__13 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__13(); +lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__13); +l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__14 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__14(); +lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__14); +l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__15 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__15(); +lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__15); +l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__16 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__16(); +lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__16); +l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__17 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__17(); +lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__17); +l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__18 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__18(); +lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__18); +l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__19 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__19(); +lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__19); +l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__20 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__20(); +lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__20); +l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__21 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__21(); +lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__21); +l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__22 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__22(); +lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__22); +l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__23 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__23(); +lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__23); +l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__24 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__24(); +lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__24); +l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__25 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__25(); +lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__25); +l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__26 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__26(); +lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__26); +l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__27 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__27(); +lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__27); +l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__28 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__28(); +lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__28); +l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__29 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__29(); +lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__29); +l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__30 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__30(); +lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__30); +l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__31 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__31(); +lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__31); +l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__32 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__32(); +lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__32); +l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__33 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__33(); +lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__33); +l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__34 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__34(); +lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__34); +l___auto____x40_Init_Data_Vector_Basic___hyg_2352_ = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352_(); +lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2352_); l_Std_Range_forIn_x27_loop___at_Vector_lex___spec__1___rarg___closed__1 = _init_l_Std_Range_forIn_x27_loop___at_Vector_lex___spec__1___rarg___closed__1(); lean_mark_persistent(l_Std_Range_forIn_x27_loop___at_Vector_lex___spec__1___rarg___closed__1); l_Std_Range_forIn_x27_loop___at_Vector_lex___spec__1___rarg___closed__2 = _init_l_Std_Range_forIn_x27_loop___at_Vector_lex___spec__1___rarg___closed__2(); diff --git a/stage0/stdlib/Init/Grind.c b/stage0/stdlib/Init/Grind.c index 66af4555bd5a..883821f27a08 100644 --- a/stage0/stdlib/Init/Grind.c +++ b/stage0/stdlib/Init/Grind.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Init.Grind -// Imports: Init.Grind.Norm Init.Grind.Tactics Init.Grind.Lemmas Init.Grind.Cases Init.Grind.Propagator Init.Grind.Util +// Imports: Init.Grind.Norm Init.Grind.Tactics Init.Grind.Lemmas Init.Grind.Cases Init.Grind.Propagator Init.Grind.Util Init.Grind.Offset #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -19,6 +19,7 @@ lean_object* initialize_Init_Grind_Lemmas(uint8_t builtin, lean_object*); lean_object* initialize_Init_Grind_Cases(uint8_t builtin, lean_object*); lean_object* initialize_Init_Grind_Propagator(uint8_t builtin, lean_object*); lean_object* initialize_Init_Grind_Util(uint8_t builtin, lean_object*); +lean_object* initialize_Init_Grind_Offset(uint8_t builtin, lean_object*); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Init_Grind(uint8_t builtin, lean_object* w) { lean_object * res; @@ -42,6 +43,9 @@ lean_dec_ref(res); res = initialize_Init_Grind_Util(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +res = initialize_Init_Grind_Offset(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/Grind/Norm.c b/stage0/stdlib/Init/Grind/Norm.c index 3d8f931a3d9a..dce02ea78b0d 100644 --- a/stage0/stdlib/Init/Grind/Norm.c +++ b/stage0/stdlib/Init/Grind/Norm.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Init.Grind.Norm -// Imports: Init.SimpLemmas Init.Classical Init.ByCases +// Imports: Init.SimpLemmas Init.PropLemmas Init.Classical Init.ByCases #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -14,6 +14,7 @@ extern "C" { #endif lean_object* initialize_Init_SimpLemmas(uint8_t builtin, lean_object*); +lean_object* initialize_Init_PropLemmas(uint8_t builtin, lean_object*); lean_object* initialize_Init_Classical(uint8_t builtin, lean_object*); lean_object* initialize_Init_ByCases(uint8_t builtin, lean_object*); static bool _G_initialized = false; @@ -24,6 +25,9 @@ _G_initialized = true; res = initialize_Init_SimpLemmas(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +res = initialize_Init_PropLemmas(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); res = initialize_Init_Classical(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); diff --git a/stage0/stdlib/Init/Grind/Offset.c b/stage0/stdlib/Init/Grind/Offset.c new file mode 100644 index 000000000000..755f50a70b76 --- /dev/null +++ b/stage0/stdlib/Init/Grind/Offset.c @@ -0,0 +1,1367 @@ +// Lean compiler output +// Module: Init.Grind.Offset +// Imports: Init.Core Init.Omega +#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___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__22; +static lean_object* l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__16; +LEAN_EXPORT uint8_t l_Lean_Grind_Offset_Cnstr_isImpliedBy(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Lean_Grind_Offset_instDecidableEqCnstr(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Grind_Offset_instInhabitedCnstr; +LEAN_EXPORT lean_object* l_Lean_Grind_Offset_Var_denote___boxed(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Lean_Grind_Offset_Cnstrs_isFalse(lean_object*); +lean_object* l_Lean_RArray_getImpl___rarg(lean_object*, lean_object*); +static lean_object* l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__3; +LEAN_EXPORT lean_object* l_Lean_Grind_Offset_Cnstr_isImpliedBy___boxed(lean_object*, lean_object*); +static lean_object* l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__4; +static lean_object* l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__21; +LEAN_EXPORT lean_object* l_Lean_Grind_Offset_Cnstrs_trans(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Grind_Offset_instReprCnstr; +LEAN_EXPORT lean_object* l___private_Init_Grind_Offset_0__Lean_Grind_Offset_Cnstr_trans_match__1_splitter___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__6; +static lean_object* l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__5; +static lean_object* l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__1; +static lean_object* l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__7; +static lean_object* l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__18; +LEAN_EXPORT lean_object* l_Lean_Grind_Offset_Cnstr_isFalse___boxed(lean_object*); +lean_object* lean_nat_to_int(lean_object*); +static lean_object* l_Lean_Grind_Offset_instInhabitedCnstr___closed__1; +LEAN_EXPORT uint8_t l_Lean_Grind_Offset_Cnstr_isTrivial(lean_object*); +static lean_object* l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__20; +static lean_object* l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__28; +LEAN_EXPORT lean_object* l_Lean_Grind_Offset_instDecidableEqCnstr___boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Grind_Offset_instReprCnstr___closed__1; +static lean_object* l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__14; +LEAN_EXPORT lean_object* l_Lean_Grind_Offset_Cnstr_isTrivial___boxed(lean_object*); +static lean_object* l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__17; +static lean_object* l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__29; +LEAN_EXPORT lean_object* l___private_Init_Grind_Offset_0__Lean_Grind_Offset_Cnstr_trans_match__2_splitter(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Grind_Offset_Var_denote(lean_object*, lean_object*); +static lean_object* l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__23; +LEAN_EXPORT lean_object* l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75_(lean_object*, lean_object*); +static lean_object* l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__2; +LEAN_EXPORT lean_object* l_Lean_Grind_Offset_Cnstr_trans(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Lean_Grind_Offset_Cnstr_isFalse(lean_object*); +static lean_object* l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__11; +static lean_object* l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__15; +static lean_object* l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__8; +LEAN_EXPORT lean_object* l___private_Init_Grind_Offset_0__Lean_Grind_Offset_decEqCnstr____x40_Init_Grind_Offset___hyg_173____boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Grind_Offset_trivialCnstr___closed__1; +static lean_object* l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__25; +LEAN_EXPORT lean_object* l_Lean_Grind_Offset_trivialCnstr; +static lean_object* l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__26; +LEAN_EXPORT lean_object* l_Lean_Grind_Offset_Cnstr_trans___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Grind_Offset_0__Lean_Grind_Offset_Cnstrs_denoteAnd_x27_match__1_splitter___rarg___boxed(lean_object*, lean_object*, lean_object*); +lean_object* lean_string_length(lean_object*); +uint8_t lean_nat_dec_eq(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Init_Grind_Offset_0__Lean_Grind_Offset_decEqCnstr____x40_Init_Grind_Offset___hyg_173_(lean_object*, lean_object*); +uint8_t lean_nat_dec_lt(lean_object*, lean_object*); +static lean_object* l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__13; +static lean_object* l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__19; +static lean_object* l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__12; +lean_object* lean_nat_sub(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Grind_Offset_fixedVar; +LEAN_EXPORT lean_object* l___private_Init_Grind_Offset_0__Lean_Grind_Offset_Cnstr_trans_match__1_splitter(lean_object*); +static lean_object* l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__10; +LEAN_EXPORT lean_object* l___private_Init_Grind_Offset_0__Lean_Grind_Offset_Cnstr_trans_match__2_splitter___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Grind_Offset_Cnstrs_trans_x27(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Grind_Offset_0__Lean_Grind_Offset_Cnstrs_denoteAnd_x27_match__1_splitter___rarg(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Grind_Offset_Cnstrs_isFalse___boxed(lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Grind_Offset_0__Lean_Grind_Offset_Cnstrs_denoteAnd_x27_match__1_splitter(lean_object*); +static lean_object* l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__27; +LEAN_EXPORT lean_object* l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____boxed(lean_object*, lean_object*); +static lean_object* l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__9; +lean_object* lean_nat_add(lean_object*, lean_object*); +lean_object* l___private_Init_Data_Repr_0__Nat_reprFast(lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Grind_Offset_0__Lean_Grind_Offset_Cnstr_trans_match__1_splitter___rarg(uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__24; +static lean_object* _init_l_Lean_Grind_Offset_fixedVar() { +_start: +{ +lean_object* x_1; +x_1 = lean_unsigned_to_nat(100000000u); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Grind_Offset_Var_denote(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; uint8_t x_4; +x_3 = l_Lean_Grind_Offset_fixedVar; +x_4 = lean_nat_dec_eq(x_2, x_3); +if (x_4 == 0) +{ +lean_object* x_5; +x_5 = l_Lean_RArray_getImpl___rarg(x_1, x_2); +return x_5; +} +else +{ +lean_object* x_6; +x_6 = lean_unsigned_to_nat(1u); +return x_6; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Grind_Offset_Var_denote___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_Grind_Offset_Var_denote(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +static lean_object* _init_l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("x", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____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_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____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_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____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_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(" := ", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____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_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__3; +x_2 = l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____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_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(5u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__8() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(",", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____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_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__10() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("y", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____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_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__12() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("k", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____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_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__14() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("l", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__15() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____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_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__16() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("{ ", 2, 2); +return x_1; +} +} +static lean_object* _init_l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__17() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__16; +x_2 = lean_string_length(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__18() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__17; +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__19() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____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_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__20() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(" }", 2, 2); +return x_1; +} +} +static lean_object* _init_l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__21() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____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_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__22() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("false", 5, 5); +return x_1; +} +} +static lean_object* _init_l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__23() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__22; +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_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__24() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__7; +x_2 = l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__23; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__25() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__24; +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_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__26() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("true", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__27() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____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_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__28() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__7; +x_2 = l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__27; +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_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__29() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__28; +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_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75_(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; lean_object* x_41; uint8_t x_42; +x_3 = lean_ctor_get(x_1, 0); +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_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____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); +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_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____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_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____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); +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_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____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_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____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___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 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_23, 0, x_6); +lean_ctor_set(x_23, 1, x_22); +x_24 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set_uint8(x_24, sizeof(void*)*1, x_8); +x_25 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_25, 0, x_19); +lean_ctor_set(x_25, 1, x_24); +x_26 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_26, 0, x_25); +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_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__13; +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 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_18); +x_31 = lean_ctor_get(x_1, 2); +lean_inc(x_31); +x_32 = l___private_Init_Data_Repr_0__Nat_reprFast(x_31); +x_33 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_33, 0, x_32); +x_34 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_34, 0, x_6); +lean_ctor_set(x_34, 1, x_33); +x_35 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_35, 0, x_34); +lean_ctor_set_uint8(x_35, sizeof(void*)*1, x_8); +x_36 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_36, 0, x_30); +lean_ctor_set(x_36, 1, x_35); +x_37 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_37, 0, x_36); +lean_ctor_set(x_37, 1, x_12); +x_38 = 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_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__15; +x_40 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_40, 0, x_38); +lean_ctor_set(x_40, 1, x_39); +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*)*3); +lean_dec(x_1); +if (x_42 == 0) +{ +lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_43 = l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__25; +x_44 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_44, 0, x_41); +lean_ctor_set(x_44, 1, x_43); +x_45 = l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__19; +x_46 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_44); +x_47 = l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__21; +x_48 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_48, 0, x_46); +lean_ctor_set(x_48, 1, x_47); +x_49 = l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__18; +x_50 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_50, 0, x_49); +lean_ctor_set(x_50, 1, x_48); +x_51 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_51, 0, x_50); +lean_ctor_set_uint8(x_51, sizeof(void*)*1, x_8); +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; +x_52 = l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__29; +x_53 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_53, 0, x_41); +lean_ctor_set(x_53, 1, x_52); +x_54 = l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__19; +x_55 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_55, 0, x_54); +lean_ctor_set(x_55, 1, x_53); +x_56 = l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__21; +x_57 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_57, 0, x_55); +lean_ctor_set(x_57, 1, x_56); +x_58 = l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__18; +x_59 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_59, 0, x_58); +lean_ctor_set(x_59, 1, x_57); +x_60 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_60, 0, x_59); +lean_ctor_set_uint8(x_60, sizeof(void*)*1, x_8); +return x_60; +} +} +} +LEAN_EXPORT lean_object* l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75_(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Grind_Offset_instReprCnstr___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Grind_Offset_instReprCnstr() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Grind_Offset_instReprCnstr___closed__1; +return x_1; +} +} +LEAN_EXPORT uint8_t l___private_Init_Grind_Offset_0__Lean_Grind_Offset_decEqCnstr____x40_Init_Grind_Offset___hyg_173_(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; uint8_t x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; uint8_t x_11; +x_3 = lean_ctor_get(x_1, 0); +x_4 = lean_ctor_get(x_1, 1); +x_5 = lean_ctor_get(x_1, 2); +x_6 = lean_ctor_get_uint8(x_1, sizeof(void*)*3); +x_7 = lean_ctor_get(x_2, 0); +x_8 = lean_ctor_get(x_2, 1); +x_9 = lean_ctor_get(x_2, 2); +x_10 = lean_ctor_get_uint8(x_2, sizeof(void*)*3); +x_11 = lean_nat_dec_eq(x_3, x_7); +if (x_11 == 0) +{ +uint8_t x_12; +x_12 = 0; +return x_12; +} +else +{ +uint8_t x_13; +x_13 = lean_nat_dec_eq(x_4, x_8); +if (x_13 == 0) +{ +uint8_t x_14; +x_14 = 0; +return x_14; +} +else +{ +uint8_t x_15; +x_15 = lean_nat_dec_eq(x_5, x_9); +if (x_15 == 0) +{ +uint8_t x_16; +x_16 = 0; +return x_16; +} +else +{ +if (x_6 == 0) +{ +if (x_10 == 0) +{ +uint8_t x_17; +x_17 = 1; +return x_17; +} +else +{ +uint8_t x_18; +x_18 = 0; +return x_18; +} +} +else +{ +return x_10; +} +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Init_Grind_Offset_0__Lean_Grind_Offset_decEqCnstr____x40_Init_Grind_Offset___hyg_173____boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l___private_Init_Grind_Offset_0__Lean_Grind_Offset_decEqCnstr____x40_Init_Grind_Offset___hyg_173_(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT uint8_t l_Lean_Grind_Offset_instDecidableEqCnstr(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = l___private_Init_Grind_Offset_0__Lean_Grind_Offset_decEqCnstr____x40_Init_Grind_Offset___hyg_173_(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Grind_Offset_instDecidableEqCnstr___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Lean_Grind_Offset_instDecidableEqCnstr(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Grind_Offset_instInhabitedCnstr___closed__1() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = lean_unsigned_to_nat(0u); +x_2 = 0; +x_3 = lean_alloc_ctor(0, 3, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_1); +lean_ctor_set(x_3, 2, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*3, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Grind_Offset_instInhabitedCnstr() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Grind_Offset_instInhabitedCnstr___closed__1; +return x_1; +} +} +static lean_object* _init_l_Lean_Grind_Offset_trivialCnstr___closed__1() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = lean_unsigned_to_nat(0u); +x_2 = 1; +x_3 = lean_alloc_ctor(0, 3, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_1); +lean_ctor_set(x_3, 2, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*3, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Grind_Offset_trivialCnstr() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Grind_Offset_trivialCnstr___closed__1; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Grind_Offset_Cnstr_trans(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; uint8_t x_5; +x_3 = lean_ctor_get(x_1, 1); +x_4 = lean_ctor_get(x_2, 0); +lean_inc(x_4); +x_5 = lean_nat_dec_eq(x_3, x_4); +lean_dec(x_4); +if (x_5 == 0) +{ +lean_object* x_6; +lean_dec(x_2); +x_6 = l_Lean_Grind_Offset_trivialCnstr; +return x_6; +} +else +{ +uint8_t x_7; +x_7 = lean_ctor_get_uint8(x_1, sizeof(void*)*3); +if (x_7 == 0) +{ +uint8_t x_8; +x_8 = lean_ctor_get_uint8(x_2, sizeof(void*)*3); +if (x_8 == 0) +{ +uint8_t x_9; +x_9 = !lean_is_exclusive(x_2); +if (x_9 == 0) +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_10 = lean_ctor_get(x_1, 0); +x_11 = lean_ctor_get(x_1, 2); +x_12 = lean_ctor_get(x_2, 2); +x_13 = lean_ctor_get(x_2, 0); +lean_dec(x_13); +x_14 = lean_nat_add(x_11, x_12); +lean_dec(x_12); +x_15 = 0; +lean_inc(x_10); +lean_ctor_set(x_2, 2, x_14); +lean_ctor_set(x_2, 0, x_10); +lean_ctor_set_uint8(x_2, sizeof(void*)*3, x_15); +return x_2; +} +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; lean_object* x_22; +x_16 = lean_ctor_get(x_1, 0); +x_17 = lean_ctor_get(x_1, 2); +x_18 = lean_ctor_get(x_2, 1); +x_19 = lean_ctor_get(x_2, 2); +lean_inc(x_19); +lean_inc(x_18); +lean_dec(x_2); +x_20 = lean_nat_add(x_17, x_19); +lean_dec(x_19); +x_21 = 0; +lean_inc(x_16); +x_22 = lean_alloc_ctor(0, 3, 1); +lean_ctor_set(x_22, 0, x_16); +lean_ctor_set(x_22, 1, x_18); +lean_ctor_set(x_22, 2, x_20); +lean_ctor_set_uint8(x_22, sizeof(void*)*3, x_21); +return x_22; +} +} +else +{ +uint8_t x_23; +x_23 = !lean_is_exclusive(x_2); +if (x_23 == 0) +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; +x_24 = lean_ctor_get(x_1, 0); +x_25 = lean_ctor_get(x_1, 2); +x_26 = lean_ctor_get(x_2, 2); +x_27 = lean_ctor_get(x_2, 0); +lean_dec(x_27); +x_28 = lean_nat_dec_lt(x_25, x_26); +if (x_28 == 0) +{ +lean_object* x_29; uint8_t x_30; +x_29 = lean_nat_sub(x_25, x_26); +lean_dec(x_26); +x_30 = 0; +lean_inc(x_24); +lean_ctor_set(x_2, 2, x_29); +lean_ctor_set(x_2, 0, x_24); +lean_ctor_set_uint8(x_2, sizeof(void*)*3, x_30); +return x_2; +} +else +{ +lean_object* x_31; uint8_t x_32; +x_31 = lean_nat_sub(x_26, x_25); +lean_dec(x_26); +x_32 = 1; +lean_inc(x_24); +lean_ctor_set(x_2, 2, x_31); +lean_ctor_set(x_2, 0, x_24); +lean_ctor_set_uint8(x_2, sizeof(void*)*3, x_32); +return x_2; +} +} +else +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; +x_33 = lean_ctor_get(x_1, 0); +x_34 = lean_ctor_get(x_1, 2); +x_35 = lean_ctor_get(x_2, 1); +x_36 = lean_ctor_get(x_2, 2); +lean_inc(x_36); +lean_inc(x_35); +lean_dec(x_2); +x_37 = lean_nat_dec_lt(x_34, x_36); +if (x_37 == 0) +{ +lean_object* x_38; uint8_t x_39; lean_object* x_40; +x_38 = lean_nat_sub(x_34, x_36); +lean_dec(x_36); +x_39 = 0; +lean_inc(x_33); +x_40 = lean_alloc_ctor(0, 3, 1); +lean_ctor_set(x_40, 0, x_33); +lean_ctor_set(x_40, 1, x_35); +lean_ctor_set(x_40, 2, x_38); +lean_ctor_set_uint8(x_40, sizeof(void*)*3, x_39); +return x_40; +} +else +{ +lean_object* x_41; uint8_t x_42; lean_object* x_43; +x_41 = lean_nat_sub(x_36, x_34); +lean_dec(x_36); +x_42 = 1; +lean_inc(x_33); +x_43 = lean_alloc_ctor(0, 3, 1); +lean_ctor_set(x_43, 0, x_33); +lean_ctor_set(x_43, 1, x_35); +lean_ctor_set(x_43, 2, x_41); +lean_ctor_set_uint8(x_43, sizeof(void*)*3, x_42); +return x_43; +} +} +} +} +else +{ +uint8_t x_44; +x_44 = lean_ctor_get_uint8(x_2, sizeof(void*)*3); +if (x_44 == 0) +{ +uint8_t x_45; +x_45 = !lean_is_exclusive(x_2); +if (x_45 == 0) +{ +lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; uint8_t x_50; +x_46 = lean_ctor_get(x_1, 0); +x_47 = lean_ctor_get(x_1, 2); +x_48 = lean_ctor_get(x_2, 2); +x_49 = lean_ctor_get(x_2, 0); +lean_dec(x_49); +x_50 = lean_nat_dec_lt(x_47, x_48); +if (x_50 == 0) +{ +lean_object* x_51; uint8_t x_52; +x_51 = lean_nat_sub(x_47, x_48); +lean_dec(x_48); +x_52 = 1; +lean_inc(x_46); +lean_ctor_set(x_2, 2, x_51); +lean_ctor_set(x_2, 0, x_46); +lean_ctor_set_uint8(x_2, sizeof(void*)*3, x_52); +return x_2; +} +else +{ +lean_object* x_53; uint8_t x_54; +x_53 = lean_nat_sub(x_48, x_47); +lean_dec(x_48); +x_54 = 0; +lean_inc(x_46); +lean_ctor_set(x_2, 2, x_53); +lean_ctor_set(x_2, 0, x_46); +lean_ctor_set_uint8(x_2, sizeof(void*)*3, x_54); +return x_2; +} +} +else +{ +lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; uint8_t x_59; +x_55 = lean_ctor_get(x_1, 0); +x_56 = lean_ctor_get(x_1, 2); +x_57 = lean_ctor_get(x_2, 1); +x_58 = lean_ctor_get(x_2, 2); +lean_inc(x_58); +lean_inc(x_57); +lean_dec(x_2); +x_59 = lean_nat_dec_lt(x_56, x_58); +if (x_59 == 0) +{ +lean_object* x_60; uint8_t x_61; lean_object* x_62; +x_60 = lean_nat_sub(x_56, x_58); +lean_dec(x_58); +x_61 = 1; +lean_inc(x_55); +x_62 = lean_alloc_ctor(0, 3, 1); +lean_ctor_set(x_62, 0, x_55); +lean_ctor_set(x_62, 1, x_57); +lean_ctor_set(x_62, 2, x_60); +lean_ctor_set_uint8(x_62, sizeof(void*)*3, x_61); +return x_62; +} +else +{ +lean_object* x_63; uint8_t x_64; lean_object* x_65; +x_63 = lean_nat_sub(x_58, x_56); +lean_dec(x_58); +x_64 = 0; +lean_inc(x_55); +x_65 = lean_alloc_ctor(0, 3, 1); +lean_ctor_set(x_65, 0, x_55); +lean_ctor_set(x_65, 1, x_57); +lean_ctor_set(x_65, 2, x_63); +lean_ctor_set_uint8(x_65, sizeof(void*)*3, x_64); +return x_65; +} +} +} +else +{ +uint8_t x_66; +x_66 = !lean_is_exclusive(x_2); +if (x_66 == 0) +{ +lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; uint8_t x_72; +x_67 = lean_ctor_get(x_1, 0); +x_68 = lean_ctor_get(x_1, 2); +x_69 = lean_ctor_get(x_2, 2); +x_70 = lean_ctor_get(x_2, 0); +lean_dec(x_70); +x_71 = lean_nat_add(x_68, x_69); +lean_dec(x_69); +x_72 = 1; +lean_inc(x_67); +lean_ctor_set(x_2, 2, x_71); +lean_ctor_set(x_2, 0, x_67); +lean_ctor_set_uint8(x_2, sizeof(void*)*3, x_72); +return x_2; +} +else +{ +lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; uint8_t x_78; lean_object* x_79; +x_73 = lean_ctor_get(x_1, 0); +x_74 = lean_ctor_get(x_1, 2); +x_75 = lean_ctor_get(x_2, 1); +x_76 = lean_ctor_get(x_2, 2); +lean_inc(x_76); +lean_inc(x_75); +lean_dec(x_2); +x_77 = lean_nat_add(x_74, x_76); +lean_dec(x_76); +x_78 = 1; +lean_inc(x_73); +x_79 = lean_alloc_ctor(0, 3, 1); +lean_ctor_set(x_79, 0, x_73); +lean_ctor_set(x_79, 1, x_75); +lean_ctor_set(x_79, 2, x_77); +lean_ctor_set_uint8(x_79, sizeof(void*)*3, x_78); +return x_79; +} +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Grind_Offset_Cnstr_trans___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_Grind_Offset_Cnstr_trans(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Init_Grind_Offset_0__Lean_Grind_Offset_Cnstr_trans_match__2_splitter___rarg(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; uint8_t x_6; lean_object* x_7; lean_object* x_8; +x_3 = lean_ctor_get(x_1, 0); +lean_inc(x_3); +x_4 = lean_ctor_get(x_1, 1); +lean_inc(x_4); +x_5 = lean_ctor_get(x_1, 2); +lean_inc(x_5); +x_6 = lean_ctor_get_uint8(x_1, sizeof(void*)*3); +lean_dec(x_1); +x_7 = lean_box(x_6); +x_8 = lean_apply_4(x_2, x_3, x_4, x_5, x_7); +return x_8; +} +} +LEAN_EXPORT lean_object* l___private_Init_Grind_Offset_0__Lean_Grind_Offset_Cnstr_trans_match__2_splitter(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l___private_Init_Grind_Offset_0__Lean_Grind_Offset_Cnstr_trans_match__2_splitter___rarg), 2, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Init_Grind_Offset_0__Lean_Grind_Offset_Cnstr_trans_match__1_splitter___rarg(uint8_t x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +if (x_1 == 0) +{ +if (x_2 == 0) +{ +lean_inc(x_3); +return x_3; +} +else +{ +lean_inc(x_4); +return x_4; +} +} +else +{ +if (x_2 == 0) +{ +lean_inc(x_5); +return x_5; +} +else +{ +lean_inc(x_6); +return x_6; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Init_Grind_Offset_0__Lean_Grind_Offset_Cnstr_trans_match__1_splitter(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l___private_Init_Grind_Offset_0__Lean_Grind_Offset_Cnstr_trans_match__1_splitter___rarg___boxed), 6, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Init_Grind_Offset_0__Lean_Grind_Offset_Cnstr_trans_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, lean_object* x_6) { +_start: +{ +uint8_t x_7; uint8_t x_8; lean_object* x_9; +x_7 = lean_unbox(x_1); +lean_dec(x_1); +x_8 = lean_unbox(x_2); +lean_dec(x_2); +x_9 = l___private_Init_Grind_Offset_0__Lean_Grind_Offset_Cnstr_trans_match__1_splitter___rarg(x_7, x_8, x_3, x_4, x_5, x_6); +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_Grind_Offset_Cnstr_isTrivial(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; uint8_t x_4; +x_2 = lean_ctor_get(x_1, 0); +x_3 = lean_ctor_get(x_1, 1); +x_4 = lean_nat_dec_eq(x_2, x_3); +if (x_4 == 0) +{ +uint8_t x_5; +x_5 = 0; +return x_5; +} +else +{ +lean_object* x_6; lean_object* x_7; uint8_t x_8; +x_6 = lean_ctor_get(x_1, 2); +x_7 = lean_unsigned_to_nat(0u); +x_8 = lean_nat_dec_eq(x_6, x_7); +return x_8; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Grind_Offset_Cnstr_isTrivial___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = l_Lean_Grind_Offset_Cnstr_isTrivial(x_1); +lean_dec(x_1); +x_3 = lean_box(x_2); +return x_3; +} +} +LEAN_EXPORT uint8_t l_Lean_Grind_Offset_Cnstr_isFalse(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; uint8_t x_4; +x_2 = lean_ctor_get(x_1, 0); +x_3 = lean_ctor_get(x_1, 1); +x_4 = lean_nat_dec_eq(x_2, x_3); +if (x_4 == 0) +{ +uint8_t x_5; +x_5 = 0; +return x_5; +} +else +{ +lean_object* x_6; lean_object* x_7; uint8_t x_8; +x_6 = lean_ctor_get(x_1, 2); +x_7 = lean_unsigned_to_nat(0u); +x_8 = lean_nat_dec_eq(x_6, x_7); +if (x_8 == 0) +{ +uint8_t x_9; +x_9 = lean_ctor_get_uint8(x_1, sizeof(void*)*3); +return x_9; +} +else +{ +uint8_t x_10; +x_10 = 0; +return x_10; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Grind_Offset_Cnstr_isFalse___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = l_Lean_Grind_Offset_Cnstr_isFalse(x_1); +lean_dec(x_1); +x_3 = lean_box(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Init_Grind_Offset_0__Lean_Grind_Offset_Cnstrs_denoteAnd_x27_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_Grind_Offset_0__Lean_Grind_Offset_Cnstrs_denoteAnd_x27_match__1_splitter(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l___private_Init_Grind_Offset_0__Lean_Grind_Offset_Cnstrs_denoteAnd_x27_match__1_splitter___rarg___boxed), 3, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Init_Grind_Offset_0__Lean_Grind_Offset_Cnstrs_denoteAnd_x27_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_Grind_Offset_0__Lean_Grind_Offset_Cnstrs_denoteAnd_x27_match__1_splitter___rarg(x_1, x_2, x_3); +lean_dec(x_2); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_Grind_Offset_Cnstrs_trans_x27(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +return x_1; +} +else +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = lean_ctor_get(x_2, 0); +lean_inc(x_3); +x_4 = lean_ctor_get(x_2, 1); +lean_inc(x_4); +lean_dec(x_2); +x_5 = l_Lean_Grind_Offset_Cnstr_trans(x_1, x_3); +lean_dec(x_1); +x_1 = x_5; +x_2 = x_4; +goto _start; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Grind_Offset_Cnstrs_trans(lean_object* x_1) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_2; +x_2 = l_Lean_Grind_Offset_trivialCnstr; +return x_2; +} +else +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = lean_ctor_get(x_1, 0); +lean_inc(x_3); +x_4 = lean_ctor_get(x_1, 1); +lean_inc(x_4); +lean_dec(x_1); +x_5 = l_Lean_Grind_Offset_Cnstrs_trans_x27(x_3, x_4); +return x_5; +} +} +} +LEAN_EXPORT uint8_t l_Lean_Grind_Offset_Cnstrs_isFalse(lean_object* x_1) { +_start: +{ +lean_object* x_2; uint8_t x_3; +x_2 = l_Lean_Grind_Offset_Cnstrs_trans(x_1); +x_3 = l_Lean_Grind_Offset_Cnstr_isFalse(x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Grind_Offset_Cnstrs_isFalse___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = l_Lean_Grind_Offset_Cnstrs_isFalse(x_1); +x_3 = lean_box(x_2); +return x_3; +} +} +LEAN_EXPORT uint8_t l_Lean_Grind_Offset_Cnstr_isImpliedBy(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; uint8_t x_4; +x_3 = l_Lean_Grind_Offset_Cnstrs_trans(x_1); +x_4 = l___private_Init_Grind_Offset_0__Lean_Grind_Offset_decEqCnstr____x40_Init_Grind_Offset___hyg_173_(x_3, x_2); +lean_dec(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_Grind_Offset_Cnstr_isImpliedBy___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Lean_Grind_Offset_Cnstr_isImpliedBy(x_1, x_2); +lean_dec(x_2); +x_4 = lean_box(x_3); +return x_4; +} +} +lean_object* initialize_Init_Core(uint8_t builtin, lean_object*); +lean_object* initialize_Init_Omega(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Init_Grind_Offset(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_Core(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Init_Omega(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l_Lean_Grind_Offset_fixedVar = _init_l_Lean_Grind_Offset_fixedVar(); +lean_mark_persistent(l_Lean_Grind_Offset_fixedVar); +l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__1 = _init_l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__1(); +lean_mark_persistent(l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__1); +l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__2 = _init_l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__2(); +lean_mark_persistent(l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__2); +l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__3 = _init_l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__3(); +lean_mark_persistent(l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__3); +l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__4 = _init_l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__4(); +lean_mark_persistent(l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__4); +l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__5 = _init_l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__5(); +lean_mark_persistent(l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__5); +l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__6 = _init_l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__6(); +lean_mark_persistent(l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__6); +l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__7 = _init_l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__7(); +lean_mark_persistent(l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__7); +l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__8 = _init_l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__8(); +lean_mark_persistent(l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__8); +l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__9 = _init_l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__9(); +lean_mark_persistent(l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__9); +l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__10 = _init_l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__10(); +lean_mark_persistent(l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__10); +l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__11 = _init_l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__11(); +lean_mark_persistent(l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__11); +l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__12 = _init_l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__12(); +lean_mark_persistent(l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__12); +l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__13 = _init_l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__13(); +lean_mark_persistent(l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__13); +l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__14 = _init_l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__14(); +lean_mark_persistent(l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__14); +l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__15 = _init_l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__15(); +lean_mark_persistent(l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__15); +l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__16 = _init_l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__16(); +lean_mark_persistent(l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__16); +l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__17 = _init_l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__17(); +lean_mark_persistent(l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__17); +l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__18 = _init_l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__18(); +lean_mark_persistent(l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__18); +l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__19 = _init_l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__19(); +lean_mark_persistent(l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__19); +l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__20 = _init_l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__20(); +lean_mark_persistent(l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__20); +l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__21 = _init_l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__21(); +lean_mark_persistent(l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__21); +l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__22 = _init_l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__22(); +lean_mark_persistent(l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__22); +l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__23 = _init_l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__23(); +lean_mark_persistent(l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__23); +l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__24 = _init_l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__24(); +lean_mark_persistent(l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__24); +l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__25 = _init_l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__25(); +lean_mark_persistent(l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__25); +l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__26 = _init_l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__26(); +lean_mark_persistent(l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__26); +l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__27 = _init_l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__27(); +lean_mark_persistent(l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__27); +l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__28 = _init_l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__28(); +lean_mark_persistent(l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__28); +l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__29 = _init_l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__29(); +lean_mark_persistent(l___private_Init_Grind_Offset_0__Lean_Grind_Offset_reprCnstr____x40_Init_Grind_Offset___hyg_75____closed__29); +l_Lean_Grind_Offset_instReprCnstr___closed__1 = _init_l_Lean_Grind_Offset_instReprCnstr___closed__1(); +lean_mark_persistent(l_Lean_Grind_Offset_instReprCnstr___closed__1); +l_Lean_Grind_Offset_instReprCnstr = _init_l_Lean_Grind_Offset_instReprCnstr(); +lean_mark_persistent(l_Lean_Grind_Offset_instReprCnstr); +l_Lean_Grind_Offset_instInhabitedCnstr___closed__1 = _init_l_Lean_Grind_Offset_instInhabitedCnstr___closed__1(); +lean_mark_persistent(l_Lean_Grind_Offset_instInhabitedCnstr___closed__1); +l_Lean_Grind_Offset_instInhabitedCnstr = _init_l_Lean_Grind_Offset_instInhabitedCnstr(); +lean_mark_persistent(l_Lean_Grind_Offset_instInhabitedCnstr); +l_Lean_Grind_Offset_trivialCnstr___closed__1 = _init_l_Lean_Grind_Offset_trivialCnstr___closed__1(); +lean_mark_persistent(l_Lean_Grind_Offset_trivialCnstr___closed__1); +l_Lean_Grind_Offset_trivialCnstr = _init_l_Lean_Grind_Offset_trivialCnstr(); +lean_mark_persistent(l_Lean_Grind_Offset_trivialCnstr); +return lean_io_result_mk_ok(lean_box(0)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Init/Grind/Tactics.c b/stage0/stdlib/Init/Grind/Tactics.c index 4e46c7564577..22cb3c2f33b5 100644 --- a/stage0/stdlib/Init/Grind/Tactics.c +++ b/stage0/stdlib/Init/Grind/Tactics.c @@ -13,39 +13,679 @@ #ifdef __cplusplus extern "C" { #endif +static lean_object* l_Lean_Parser_Attr_grindEqBoth___closed__12; +static lean_object* l_Lean_Parser_Tactic_grind___closed__11; +static lean_object* l_Lean_Parser_Attr_grindEqBoth___closed__4; static lean_object* l_Lean_Parser_Tactic_grind___closed__7; +LEAN_EXPORT lean_object* l_Lean_Parser_Attr_grindBwd; +static lean_object* l_Lean_Parser_Attr_grind___closed__13; +static lean_object* l_Lean_Parser_Attr_grindBwd___closed__1; +LEAN_EXPORT lean_object* l_Lean_Parser_Attr_grindFwd; extern lean_object* l_Lean_Parser_Tactic_optConfig; +static lean_object* l_Lean_Parser_Attr_grindEqBoth___closed__8; +static lean_object* l_Lean_Parser_Attr_grindFwd___closed__4; +static lean_object* l_Lean_Parser_Attr_grindEq___closed__3; +static lean_object* l_Lean_Parser_Attr_grind___closed__4; static lean_object* l_Lean_Grind_instInhabitedConfig___closed__1; +static lean_object* l_Lean_Parser_Attr_grindEq___closed__4; +static lean_object* l_Lean_Parser_Attr_grindEqRhs___closed__4; static lean_object* l_Lean_Parser_Tactic_grind___closed__10; +LEAN_EXPORT lean_object* l_Lean_Parser_Attr_grindEqRhs; +static lean_object* l_Lean_Parser_Attr_grindEqBoth___closed__6; +static lean_object* l_Lean_Parser_Attr_grindEq___closed__8; +static lean_object* l_Lean_Parser_Attr_grindEqBoth___closed__9; +LEAN_EXPORT lean_object* l_Lean_Parser_Attr_grind; +static lean_object* l_Lean_Parser_Attr_grindEq___closed__5; +static lean_object* l_Lean_Parser_Attr_grindBwd___closed__2; +static lean_object* l_Lean_Parser_Attr_grindEqRhs___closed__1; static lean_object* l_Lean_Parser_Tactic_grind___closed__6; +static lean_object* l_Lean_Parser_Attr_grindEq___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_grind; static lean_object* l_Lean_Grind_instBEqConfig___closed__1; static lean_object* l_Lean_Parser_Tactic_grind___closed__2; LEAN_EXPORT lean_object* l_Lean_Grind_instBEqConfig; +static lean_object* l_Lean_Parser_Attr_grindEqBoth___closed__7; static lean_object* l_Lean_Parser_Tactic_grind___closed__9; +static lean_object* l_Lean_Parser_Attr_grindFwd___closed__3; +static lean_object* l_Lean_Parser_Attr_grindEqBoth___closed__5; +static lean_object* l_Lean_Parser_Attr_grindEqBoth___closed__3; +LEAN_EXPORT lean_object* l___private_Init_Grind_Tactics_0__Lean_Grind_beqConfig____x40_Init_Grind_Tactics___hyg_134____boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Attr_grindEqBoth___closed__1; static lean_object* l_Lean_Parser_Tactic_grind___closed__1; +LEAN_EXPORT lean_object* l_Lean_Parser_Attr_grindEqBoth; +static lean_object* l_Lean_Parser_Attr_grind___closed__10; lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Grind_instInhabitedConfig; +static lean_object* l_Lean_Parser_Attr_grind___closed__3; static lean_object* l_Lean_Parser_Tactic_grind___closed__5; +static lean_object* l_Lean_Parser_Attr_grindFwd___closed__1; +static lean_object* l_Lean_Parser_Attr_grindFwd___closed__5; +static lean_object* l_Lean_Parser_Attr_grind___closed__2; +static lean_object* l_Lean_Parser_Attr_grindEq___closed__7; +LEAN_EXPORT uint8_t l___private_Init_Grind_Tactics_0__Lean_Grind_beqConfig____x40_Init_Grind_Tactics___hyg_134_(lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Attr_grindEqRhs___closed__5; static lean_object* l_Lean_Parser_Tactic_grind___closed__3; +static lean_object* l_Lean_Parser_Attr_grind___closed__6; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Attr_grind___closed__11; +static lean_object* l_Lean_Parser_Attr_grindEqBoth___closed__2; +static lean_object* l_Lean_Parser_Attr_grindBwd___closed__3; +static lean_object* l_Lean_Parser_Attr_grindEqRhs___closed__3; +static lean_object* l_Lean_Parser_Attr_grindEq___closed__2; +static lean_object* l_Lean_Parser_Attr_grindEq___closed__6; +LEAN_EXPORT lean_object* l_Lean_Parser_Attr_grindEq; static lean_object* l_Lean_Parser_Tactic_grind___closed__4; -LEAN_EXPORT lean_object* l___private_Init_Grind_Tactics_0__Lean_Grind_beqConfig____x40_Init_Grind_Tactics___hyg_30____boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Attr_grindEqBoth___closed__11; +static lean_object* l_Lean_Parser_Attr_grind___closed__5; static lean_object* l_Lean_Parser_Tactic_grind___closed__8; +static lean_object* l_Lean_Parser_Attr_grindEqBoth___closed__10; lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT uint8_t l___private_Init_Grind_Tactics_0__Lean_Grind_beqConfig____x40_Init_Grind_Tactics___hyg_30_(lean_object*, lean_object*); -static lean_object* _init_l_Lean_Grind_instInhabitedConfig___closed__1() { +static lean_object* l_Lean_Parser_Attr_grind___closed__12; +static lean_object* l_Lean_Parser_Attr_grind___closed__8; +static lean_object* l_Lean_Parser_Attr_grindBwd___closed__4; +static lean_object* l_Lean_Parser_Attr_grind___closed__7; +static lean_object* l_Lean_Parser_Attr_grind___closed__14; +static lean_object* l_Lean_Parser_Attr_grindBwd___closed__5; +static lean_object* l_Lean_Parser_Attr_grindFwd___closed__2; +static lean_object* l_Lean_Parser_Attr_grindEqRhs___closed__2; +static lean_object* l_Lean_Parser_Attr_grind___closed__1; +static lean_object* l_Lean_Parser_Attr_grind___closed__9; +static lean_object* l_Lean_Parser_Tactic_grind___closed__12; +static lean_object* _init_l_Lean_Parser_Attr_grindEq___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean", 4, 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grindEq___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Parser", 6, 6); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grindEq___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Attr", 4, 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grindEq___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("grindEq", 7, 7); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grindEq___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_Parser_Attr_grindEq___closed__1; +x_2 = l_Lean_Parser_Attr_grindEq___closed__2; +x_3 = l_Lean_Parser_Attr_grindEq___closed__3; +x_4 = l_Lean_Parser_Attr_grindEq___closed__4; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grindEq___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("=", 1, 1); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grindEq___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Attr_grindEq___closed__6; +x_2 = lean_alloc_ctor(5, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grindEq___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Parser_Attr_grindEq___closed__4; +x_2 = l_Lean_Parser_Attr_grindEq___closed__5; +x_3 = l_Lean_Parser_Attr_grindEq___closed__7; +x_4 = lean_alloc_ctor(9, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grindEq() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Parser_Attr_grindEq___closed__8; +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grindEqBoth___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("grindEqBoth", 11, 11); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grindEqBoth___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Lean_Parser_Attr_grindEq___closed__1; +x_2 = l_Lean_Parser_Attr_grindEq___closed__2; +x_3 = l_Lean_Parser_Attr_grindEq___closed__3; +x_4 = l_Lean_Parser_Attr_grindEqBoth___closed__1; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grindEqBoth___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("atomic", 6, 6); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grindEqBoth___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Parser_Attr_grindEqBoth___closed__3; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grindEqBoth___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("andthen", 7, 7); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grindEqBoth___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Parser_Attr_grindEqBoth___closed__5; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grindEqBoth___closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("_", 1, 1); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grindEqBoth___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Attr_grindEqBoth___closed__7; +x_2 = lean_alloc_ctor(5, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grindEqBoth___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Parser_Attr_grindEqBoth___closed__6; +x_2 = l_Lean_Parser_Attr_grindEqBoth___closed__8; +x_3 = l_Lean_Parser_Attr_grindEq___closed__7; +x_4 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grindEqBoth___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Parser_Attr_grindEqBoth___closed__6; +x_2 = l_Lean_Parser_Attr_grindEqBoth___closed__9; +x_3 = l_Lean_Parser_Attr_grindEqBoth___closed__8; +x_4 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grindEqBoth___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Attr_grindEqBoth___closed__4; +x_2 = l_Lean_Parser_Attr_grindEqBoth___closed__10; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grindEqBoth___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Parser_Attr_grindEqBoth___closed__1; +x_2 = l_Lean_Parser_Attr_grindEqBoth___closed__2; +x_3 = l_Lean_Parser_Attr_grindEqBoth___closed__11; +x_4 = lean_alloc_ctor(9, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grindEqBoth() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Parser_Attr_grindEqBoth___closed__12; +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grindEqRhs___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("grindEqRhs", 10, 10); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grindEqRhs___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Lean_Parser_Attr_grindEq___closed__1; +x_2 = l_Lean_Parser_Attr_grindEq___closed__2; +x_3 = l_Lean_Parser_Attr_grindEq___closed__3; +x_4 = l_Lean_Parser_Attr_grindEqRhs___closed__1; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grindEqRhs___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Parser_Attr_grindEqBoth___closed__6; +x_2 = l_Lean_Parser_Attr_grindEq___closed__7; +x_3 = l_Lean_Parser_Attr_grindEqBoth___closed__8; +x_4 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grindEqRhs___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Attr_grindEqBoth___closed__4; +x_2 = l_Lean_Parser_Attr_grindEqRhs___closed__3; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grindEqRhs___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Parser_Attr_grindEqRhs___closed__1; +x_2 = l_Lean_Parser_Attr_grindEqRhs___closed__2; +x_3 = l_Lean_Parser_Attr_grindEqRhs___closed__4; +x_4 = lean_alloc_ctor(9, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grindEqRhs() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Parser_Attr_grindEqRhs___closed__5; +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grindBwd___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("grindBwd", 8, 8); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grindBwd___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Lean_Parser_Attr_grindEq___closed__1; +x_2 = l_Lean_Parser_Attr_grindEq___closed__2; +x_3 = l_Lean_Parser_Attr_grindEq___closed__3; +x_4 = l_Lean_Parser_Attr_grindBwd___closed__1; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grindBwd___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("←", 3, 1); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grindBwd___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Attr_grindBwd___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_Attr_grindBwd___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Parser_Attr_grindBwd___closed__1; +x_2 = l_Lean_Parser_Attr_grindBwd___closed__2; +x_3 = l_Lean_Parser_Attr_grindBwd___closed__4; +x_4 = lean_alloc_ctor(9, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grindBwd() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Parser_Attr_grindBwd___closed__5; +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grindFwd___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("grindFwd", 8, 8); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grindFwd___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Lean_Parser_Attr_grindEq___closed__1; +x_2 = l_Lean_Parser_Attr_grindEq___closed__2; +x_3 = l_Lean_Parser_Attr_grindEq___closed__3; +x_4 = l_Lean_Parser_Attr_grindFwd___closed__1; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grindFwd___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("→", 3, 1); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grindFwd___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Attr_grindFwd___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_Attr_grindFwd___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Parser_Attr_grindFwd___closed__1; +x_2 = l_Lean_Parser_Attr_grindFwd___closed__2; +x_3 = l_Lean_Parser_Attr_grindFwd___closed__4; +x_4 = lean_alloc_ctor(9, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grindFwd() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Parser_Attr_grindFwd___closed__5; +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grind___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("grind", 5, 5); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grind___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Lean_Parser_Attr_grindEq___closed__1; +x_2 = l_Lean_Parser_Attr_grindEq___closed__2; +x_3 = l_Lean_Parser_Attr_grindEq___closed__3; +x_4 = l_Lean_Parser_Attr_grind___closed__1; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grind___closed__3() { _start: { -uint8_t x_1; lean_object* x_2; lean_object* x_3; -x_1 = 0; -x_2 = lean_unsigned_to_nat(0u); -x_3 = lean_alloc_ctor(0, 4, 1); -lean_ctor_set(x_3, 0, x_2); +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Attr_grind___closed__1; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grind___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("optional", 8, 8); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grind___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_Attr_grind___closed__4; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grind___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("orelse", 6, 6); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grind___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_Attr_grind___closed__6; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grind___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Parser_Attr_grind___closed__7; +x_2 = l_Lean_Parser_Attr_grindBwd; +x_3 = l_Lean_Parser_Attr_grindFwd; +x_4 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grind___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Parser_Attr_grind___closed__7; +x_2 = l_Lean_Parser_Attr_grindEq; +x_3 = l_Lean_Parser_Attr_grind___closed__8; +x_4 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grind___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Parser_Attr_grind___closed__7; +x_2 = l_Lean_Parser_Attr_grindEqRhs; +x_3 = l_Lean_Parser_Attr_grind___closed__9; +x_4 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grind___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Parser_Attr_grind___closed__7; +x_2 = l_Lean_Parser_Attr_grindEqBoth; +x_3 = l_Lean_Parser_Attr_grind___closed__10; +x_4 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grind___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Attr_grind___closed__5; +x_2 = l_Lean_Parser_Attr_grind___closed__11; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); -lean_ctor_set(x_3, 2, x_2); -lean_ctor_set(x_3, 3, x_2); -lean_ctor_set_uint8(x_3, sizeof(void*)*4, x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grind___closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Parser_Attr_grindEqBoth___closed__6; +x_2 = l_Lean_Parser_Attr_grind___closed__3; +x_3 = l_Lean_Parser_Attr_grind___closed__12; +x_4 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grind___closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Parser_Attr_grind___closed__2; +x_2 = lean_unsigned_to_nat(1022u); +x_3 = l_Lean_Parser_Attr_grind___closed__13; +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_Attr_grind() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Parser_Attr_grind___closed__14; +return x_1; +} +} +static lean_object* _init_l_Lean_Grind_instInhabitedConfig___closed__1() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = lean_unsigned_to_nat(0u); +x_2 = 0; +x_3 = lean_alloc_ctor(0, 4, 4); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_1); +lean_ctor_set(x_3, 2, x_1); +lean_ctor_set(x_3, 3, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*4, x_2); +lean_ctor_set_uint8(x_3, sizeof(void*)*4 + 1, x_2); +lean_ctor_set_uint8(x_3, sizeof(void*)*4 + 2, x_2); +lean_ctor_set_uint8(x_3, sizeof(void*)*4 + 3, x_2); return x_3; } } @@ -57,109 +697,233 @@ x_1 = l_Lean_Grind_instInhabitedConfig___closed__1; return x_1; } } -LEAN_EXPORT uint8_t l___private_Init_Grind_Tactics_0__Lean_Grind_beqConfig____x40_Init_Grind_Tactics___hyg_30_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Init_Grind_Tactics_0__Lean_Grind_beqConfig____x40_Init_Grind_Tactics___hyg_134_(lean_object* x_1, lean_object* x_2) { _start: { -uint8_t x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; -x_3 = lean_ctor_get_uint8(x_1, sizeof(void*)*4); -x_4 = lean_ctor_get(x_1, 0); -x_5 = lean_ctor_get(x_1, 1); -x_6 = lean_ctor_get(x_1, 2); -x_7 = lean_ctor_get(x_1, 3); -x_8 = lean_ctor_get_uint8(x_2, sizeof(void*)*4); -x_9 = lean_ctor_get(x_2, 0); -x_10 = lean_ctor_get(x_2, 1); -x_11 = lean_ctor_get(x_2, 2); -x_12 = lean_ctor_get(x_2, 3); -if (x_3 == 0) +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; uint8_t x_8; uint8_t 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; uint8_t x_16; uint8_t x_17; uint8_t x_18; uint8_t x_19; uint8_t x_24; uint8_t x_31; uint8_t x_38; +x_3 = lean_ctor_get(x_1, 0); +x_4 = lean_ctor_get(x_1, 1); +x_5 = lean_ctor_get(x_1, 2); +x_6 = lean_ctor_get(x_1, 3); +x_7 = lean_ctor_get_uint8(x_1, sizeof(void*)*4); +x_8 = lean_ctor_get_uint8(x_1, sizeof(void*)*4 + 1); +x_9 = lean_ctor_get_uint8(x_1, sizeof(void*)*4 + 2); +x_10 = lean_ctor_get_uint8(x_1, sizeof(void*)*4 + 3); +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_uint8(x_2, sizeof(void*)*4); +x_16 = lean_ctor_get_uint8(x_2, sizeof(void*)*4 + 1); +x_17 = lean_ctor_get_uint8(x_2, sizeof(void*)*4 + 2); +x_18 = lean_ctor_get_uint8(x_2, sizeof(void*)*4 + 3); +x_38 = lean_nat_dec_eq(x_3, x_11); +if (x_38 == 0) { -if (x_8 == 0) +uint8_t x_39; +x_39 = 0; +return x_39; +} +else { -uint8_t x_23; -x_23 = 1; -x_13 = x_23; -goto block_22; +uint8_t x_40; +x_40 = lean_nat_dec_eq(x_4, x_12); +if (x_40 == 0) +{ +uint8_t x_41; +x_41 = 0; +return x_41; } else { -uint8_t x_24; -x_24 = 0; -x_13 = x_24; -goto block_22; +uint8_t x_42; +x_42 = lean_nat_dec_eq(x_5, x_13); +if (x_42 == 0) +{ +uint8_t x_43; +x_43 = 0; +return x_43; } +else +{ +uint8_t x_44; +x_44 = lean_nat_dec_eq(x_6, x_14); +if (x_44 == 0) +{ +uint8_t x_45; +x_45 = 0; +return x_45; } else { -if (x_8 == 0) +if (x_7 == 0) +{ +if (x_15 == 0) +{ +uint8_t x_46; +x_46 = 1; +x_31 = x_46; +goto block_37; +} +else +{ +uint8_t x_47; +x_47 = 0; +x_31 = x_47; +goto block_37; +} +} +else +{ +if (x_15 == 0) +{ +uint8_t x_48; +x_48 = 0; +x_31 = x_48; +goto block_37; +} +else +{ +uint8_t x_49; +x_49 = 1; +x_31 = x_49; +goto block_37; +} +} +} +} +} +} +block_23: +{ +if (x_19 == 0) +{ +uint8_t x_20; +x_20 = 0; +return x_20; +} +else +{ +if (x_10 == 0) +{ +if (x_18 == 0) +{ +uint8_t x_21; +x_21 = 1; +return x_21; +} +else +{ +uint8_t x_22; +x_22 = 0; +return x_22; +} +} +else +{ +return x_18; +} +} +} +block_30: +{ +if (x_24 == 0) { uint8_t x_25; x_25 = 0; -x_13 = x_25; -goto block_22; +return x_25; } else { +if (x_9 == 0) +{ +if (x_17 == 0) +{ uint8_t x_26; x_26 = 1; -x_13 = x_26; -goto block_22; +x_19 = x_26; +goto block_23; +} +else +{ +uint8_t x_27; +x_27 = 0; +x_19 = x_27; +goto block_23; } } -block_22: +else { -if (x_13 == 0) +if (x_17 == 0) { -uint8_t x_14; -x_14 = 0; -return x_14; +uint8_t x_28; +x_28 = 0; +x_19 = x_28; +goto block_23; } else { -uint8_t x_15; -x_15 = lean_nat_dec_eq(x_4, x_9); -if (x_15 == 0) +uint8_t x_29; +x_29 = 1; +x_19 = x_29; +goto block_23; +} +} +} +} +block_37: +{ +if (x_31 == 0) { -uint8_t x_16; -x_16 = 0; -return x_16; +uint8_t x_32; +x_32 = 0; +return x_32; } else { -uint8_t x_17; -x_17 = lean_nat_dec_eq(x_5, x_10); -if (x_17 == 0) +if (x_8 == 0) { -uint8_t x_18; -x_18 = 0; -return x_18; +if (x_16 == 0) +{ +uint8_t x_33; +x_33 = 1; +x_24 = x_33; +goto block_30; } else { -uint8_t x_19; -x_19 = lean_nat_dec_eq(x_6, x_11); -if (x_19 == 0) -{ -uint8_t x_20; -x_20 = 0; -return x_20; +uint8_t x_34; +x_34 = 0; +x_24 = x_34; +goto block_30; +} } else { -uint8_t x_21; -x_21 = lean_nat_dec_eq(x_7, x_12); -return x_21; +if (x_16 == 0) +{ +uint8_t x_35; +x_35 = 0; +x_24 = x_35; +goto block_30; } +else +{ +uint8_t x_36; +x_36 = 1; +x_24 = x_36; +goto block_30; } } } } } } -LEAN_EXPORT lean_object* l___private_Init_Grind_Tactics_0__Lean_Grind_beqConfig____x40_Init_Grind_Tactics___hyg_30____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_Grind_Tactics_0__Lean_Grind_beqConfig____x40_Init_Grind_Tactics___hyg_134____boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l___private_Init_Grind_Tactics_0__Lean_Grind_beqConfig____x40_Init_Grind_Tactics___hyg_30_(x_1, x_2); +x_3 = l___private_Init_Grind_Tactics_0__Lean_Grind_beqConfig____x40_Init_Grind_Tactics___hyg_134_(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); @@ -170,7 +934,7 @@ static lean_object* _init_l_Lean_Grind_instBEqConfig___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Init_Grind_Tactics_0__Lean_Grind_beqConfig____x40_Init_Grind_Tactics___hyg_30____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Init_Grind_Tactics_0__Lean_Grind_beqConfig____x40_Init_Grind_Tactics___hyg_134____boxed), 2, 0); return x_1; } } @@ -186,51 +950,59 @@ static lean_object* _init_l_Lean_Parser_Tactic_grind___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean", 4, 4); +x_1 = lean_mk_string_unchecked("Tactic", 6, 6); return x_1; } } static lean_object* _init_l_Lean_Parser_Tactic_grind___closed__2() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Parser", 6, 6); -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_Parser_Attr_grindEq___closed__1; +x_2 = l_Lean_Parser_Attr_grindEq___closed__2; +x_3 = l_Lean_Parser_Tactic_grind___closed__1; +x_4 = l_Lean_Parser_Attr_grind___closed__1; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; } } static lean_object* _init_l_Lean_Parser_Tactic_grind___closed__3() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Tactic", 6, 6); -return x_1; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Parser_Attr_grindEqBoth___closed__6; +x_2 = l_Lean_Parser_Attr_grind___closed__3; +x_3 = l_Lean_Parser_Tactic_optConfig; +x_4 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; } } static lean_object* _init_l_Lean_Parser_Tactic_grind___closed__4() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("grind", 5, 5); +x_1 = lean_mk_string_unchecked("on_failure ", 11, 11); return x_1; } } static lean_object* _init_l_Lean_Parser_Tactic_grind___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_Parser_Tactic_grind___closed__1; -x_2 = l_Lean_Parser_Tactic_grind___closed__2; -x_3 = l_Lean_Parser_Tactic_grind___closed__3; -x_4 = l_Lean_Parser_Tactic_grind___closed__4; -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; +x_1 = l_Lean_Parser_Tactic_grind___closed__4; +x_2 = lean_alloc_ctor(5, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; } } static lean_object* _init_l_Lean_Parser_Tactic_grind___closed__6() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("andthen", 7, 7); +x_1 = lean_mk_string_unchecked("term", 4, 4); return x_1; } } @@ -247,12 +1019,12 @@ return x_3; static lean_object* _init_l_Lean_Parser_Tactic_grind___closed__8() { _start: { -lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_grind___closed__4; -x_2 = 0; -x_3 = lean_alloc_ctor(6, 1, 1); +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Tactic_grind___closed__7; +x_2 = lean_unsigned_to_nat(0u); +x_3 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_3, 0, x_1); -lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +lean_ctor_set(x_3, 1, x_2); return x_3; } } @@ -260,9 +1032,9 @@ static lean_object* _init_l_Lean_Parser_Tactic_grind___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_grind___closed__7; -x_2 = l_Lean_Parser_Tactic_grind___closed__8; -x_3 = l_Lean_Parser_Tactic_optConfig; +x_1 = l_Lean_Parser_Attr_grindEqBoth___closed__6; +x_2 = l_Lean_Parser_Tactic_grind___closed__5; +x_3 = l_Lean_Parser_Tactic_grind___closed__8; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -273,10 +1045,36 @@ return x_4; static lean_object* _init_l_Lean_Parser_Tactic_grind___closed__10() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Attr_grind___closed__5; +x_2 = l_Lean_Parser_Tactic_grind___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_grind___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Parser_Attr_grindEqBoth___closed__6; +x_2 = l_Lean_Parser_Tactic_grind___closed__3; +x_3 = l_Lean_Parser_Tactic_grind___closed__10; +x_4 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Parser_Tactic_grind___closed__12() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_grind___closed__5; +x_1 = l_Lean_Parser_Tactic_grind___closed__2; x_2 = lean_unsigned_to_nat(1022u); -x_3 = l_Lean_Parser_Tactic_grind___closed__9; +x_3 = l_Lean_Parser_Tactic_grind___closed__11; x_4 = lean_alloc_ctor(3, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -288,7 +1086,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_grind() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Tactic_grind___closed__10; +x_1 = l_Lean_Parser_Tactic_grind___closed__12; return x_1; } } @@ -301,6 +1099,116 @@ _G_initialized = true; res = initialize_Init_Tactics(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +l_Lean_Parser_Attr_grindEq___closed__1 = _init_l_Lean_Parser_Attr_grindEq___closed__1(); +lean_mark_persistent(l_Lean_Parser_Attr_grindEq___closed__1); +l_Lean_Parser_Attr_grindEq___closed__2 = _init_l_Lean_Parser_Attr_grindEq___closed__2(); +lean_mark_persistent(l_Lean_Parser_Attr_grindEq___closed__2); +l_Lean_Parser_Attr_grindEq___closed__3 = _init_l_Lean_Parser_Attr_grindEq___closed__3(); +lean_mark_persistent(l_Lean_Parser_Attr_grindEq___closed__3); +l_Lean_Parser_Attr_grindEq___closed__4 = _init_l_Lean_Parser_Attr_grindEq___closed__4(); +lean_mark_persistent(l_Lean_Parser_Attr_grindEq___closed__4); +l_Lean_Parser_Attr_grindEq___closed__5 = _init_l_Lean_Parser_Attr_grindEq___closed__5(); +lean_mark_persistent(l_Lean_Parser_Attr_grindEq___closed__5); +l_Lean_Parser_Attr_grindEq___closed__6 = _init_l_Lean_Parser_Attr_grindEq___closed__6(); +lean_mark_persistent(l_Lean_Parser_Attr_grindEq___closed__6); +l_Lean_Parser_Attr_grindEq___closed__7 = _init_l_Lean_Parser_Attr_grindEq___closed__7(); +lean_mark_persistent(l_Lean_Parser_Attr_grindEq___closed__7); +l_Lean_Parser_Attr_grindEq___closed__8 = _init_l_Lean_Parser_Attr_grindEq___closed__8(); +lean_mark_persistent(l_Lean_Parser_Attr_grindEq___closed__8); +l_Lean_Parser_Attr_grindEq = _init_l_Lean_Parser_Attr_grindEq(); +lean_mark_persistent(l_Lean_Parser_Attr_grindEq); +l_Lean_Parser_Attr_grindEqBoth___closed__1 = _init_l_Lean_Parser_Attr_grindEqBoth___closed__1(); +lean_mark_persistent(l_Lean_Parser_Attr_grindEqBoth___closed__1); +l_Lean_Parser_Attr_grindEqBoth___closed__2 = _init_l_Lean_Parser_Attr_grindEqBoth___closed__2(); +lean_mark_persistent(l_Lean_Parser_Attr_grindEqBoth___closed__2); +l_Lean_Parser_Attr_grindEqBoth___closed__3 = _init_l_Lean_Parser_Attr_grindEqBoth___closed__3(); +lean_mark_persistent(l_Lean_Parser_Attr_grindEqBoth___closed__3); +l_Lean_Parser_Attr_grindEqBoth___closed__4 = _init_l_Lean_Parser_Attr_grindEqBoth___closed__4(); +lean_mark_persistent(l_Lean_Parser_Attr_grindEqBoth___closed__4); +l_Lean_Parser_Attr_grindEqBoth___closed__5 = _init_l_Lean_Parser_Attr_grindEqBoth___closed__5(); +lean_mark_persistent(l_Lean_Parser_Attr_grindEqBoth___closed__5); +l_Lean_Parser_Attr_grindEqBoth___closed__6 = _init_l_Lean_Parser_Attr_grindEqBoth___closed__6(); +lean_mark_persistent(l_Lean_Parser_Attr_grindEqBoth___closed__6); +l_Lean_Parser_Attr_grindEqBoth___closed__7 = _init_l_Lean_Parser_Attr_grindEqBoth___closed__7(); +lean_mark_persistent(l_Lean_Parser_Attr_grindEqBoth___closed__7); +l_Lean_Parser_Attr_grindEqBoth___closed__8 = _init_l_Lean_Parser_Attr_grindEqBoth___closed__8(); +lean_mark_persistent(l_Lean_Parser_Attr_grindEqBoth___closed__8); +l_Lean_Parser_Attr_grindEqBoth___closed__9 = _init_l_Lean_Parser_Attr_grindEqBoth___closed__9(); +lean_mark_persistent(l_Lean_Parser_Attr_grindEqBoth___closed__9); +l_Lean_Parser_Attr_grindEqBoth___closed__10 = _init_l_Lean_Parser_Attr_grindEqBoth___closed__10(); +lean_mark_persistent(l_Lean_Parser_Attr_grindEqBoth___closed__10); +l_Lean_Parser_Attr_grindEqBoth___closed__11 = _init_l_Lean_Parser_Attr_grindEqBoth___closed__11(); +lean_mark_persistent(l_Lean_Parser_Attr_grindEqBoth___closed__11); +l_Lean_Parser_Attr_grindEqBoth___closed__12 = _init_l_Lean_Parser_Attr_grindEqBoth___closed__12(); +lean_mark_persistent(l_Lean_Parser_Attr_grindEqBoth___closed__12); +l_Lean_Parser_Attr_grindEqBoth = _init_l_Lean_Parser_Attr_grindEqBoth(); +lean_mark_persistent(l_Lean_Parser_Attr_grindEqBoth); +l_Lean_Parser_Attr_grindEqRhs___closed__1 = _init_l_Lean_Parser_Attr_grindEqRhs___closed__1(); +lean_mark_persistent(l_Lean_Parser_Attr_grindEqRhs___closed__1); +l_Lean_Parser_Attr_grindEqRhs___closed__2 = _init_l_Lean_Parser_Attr_grindEqRhs___closed__2(); +lean_mark_persistent(l_Lean_Parser_Attr_grindEqRhs___closed__2); +l_Lean_Parser_Attr_grindEqRhs___closed__3 = _init_l_Lean_Parser_Attr_grindEqRhs___closed__3(); +lean_mark_persistent(l_Lean_Parser_Attr_grindEqRhs___closed__3); +l_Lean_Parser_Attr_grindEqRhs___closed__4 = _init_l_Lean_Parser_Attr_grindEqRhs___closed__4(); +lean_mark_persistent(l_Lean_Parser_Attr_grindEqRhs___closed__4); +l_Lean_Parser_Attr_grindEqRhs___closed__5 = _init_l_Lean_Parser_Attr_grindEqRhs___closed__5(); +lean_mark_persistent(l_Lean_Parser_Attr_grindEqRhs___closed__5); +l_Lean_Parser_Attr_grindEqRhs = _init_l_Lean_Parser_Attr_grindEqRhs(); +lean_mark_persistent(l_Lean_Parser_Attr_grindEqRhs); +l_Lean_Parser_Attr_grindBwd___closed__1 = _init_l_Lean_Parser_Attr_grindBwd___closed__1(); +lean_mark_persistent(l_Lean_Parser_Attr_grindBwd___closed__1); +l_Lean_Parser_Attr_grindBwd___closed__2 = _init_l_Lean_Parser_Attr_grindBwd___closed__2(); +lean_mark_persistent(l_Lean_Parser_Attr_grindBwd___closed__2); +l_Lean_Parser_Attr_grindBwd___closed__3 = _init_l_Lean_Parser_Attr_grindBwd___closed__3(); +lean_mark_persistent(l_Lean_Parser_Attr_grindBwd___closed__3); +l_Lean_Parser_Attr_grindBwd___closed__4 = _init_l_Lean_Parser_Attr_grindBwd___closed__4(); +lean_mark_persistent(l_Lean_Parser_Attr_grindBwd___closed__4); +l_Lean_Parser_Attr_grindBwd___closed__5 = _init_l_Lean_Parser_Attr_grindBwd___closed__5(); +lean_mark_persistent(l_Lean_Parser_Attr_grindBwd___closed__5); +l_Lean_Parser_Attr_grindBwd = _init_l_Lean_Parser_Attr_grindBwd(); +lean_mark_persistent(l_Lean_Parser_Attr_grindBwd); +l_Lean_Parser_Attr_grindFwd___closed__1 = _init_l_Lean_Parser_Attr_grindFwd___closed__1(); +lean_mark_persistent(l_Lean_Parser_Attr_grindFwd___closed__1); +l_Lean_Parser_Attr_grindFwd___closed__2 = _init_l_Lean_Parser_Attr_grindFwd___closed__2(); +lean_mark_persistent(l_Lean_Parser_Attr_grindFwd___closed__2); +l_Lean_Parser_Attr_grindFwd___closed__3 = _init_l_Lean_Parser_Attr_grindFwd___closed__3(); +lean_mark_persistent(l_Lean_Parser_Attr_grindFwd___closed__3); +l_Lean_Parser_Attr_grindFwd___closed__4 = _init_l_Lean_Parser_Attr_grindFwd___closed__4(); +lean_mark_persistent(l_Lean_Parser_Attr_grindFwd___closed__4); +l_Lean_Parser_Attr_grindFwd___closed__5 = _init_l_Lean_Parser_Attr_grindFwd___closed__5(); +lean_mark_persistent(l_Lean_Parser_Attr_grindFwd___closed__5); +l_Lean_Parser_Attr_grindFwd = _init_l_Lean_Parser_Attr_grindFwd(); +lean_mark_persistent(l_Lean_Parser_Attr_grindFwd); +l_Lean_Parser_Attr_grind___closed__1 = _init_l_Lean_Parser_Attr_grind___closed__1(); +lean_mark_persistent(l_Lean_Parser_Attr_grind___closed__1); +l_Lean_Parser_Attr_grind___closed__2 = _init_l_Lean_Parser_Attr_grind___closed__2(); +lean_mark_persistent(l_Lean_Parser_Attr_grind___closed__2); +l_Lean_Parser_Attr_grind___closed__3 = _init_l_Lean_Parser_Attr_grind___closed__3(); +lean_mark_persistent(l_Lean_Parser_Attr_grind___closed__3); +l_Lean_Parser_Attr_grind___closed__4 = _init_l_Lean_Parser_Attr_grind___closed__4(); +lean_mark_persistent(l_Lean_Parser_Attr_grind___closed__4); +l_Lean_Parser_Attr_grind___closed__5 = _init_l_Lean_Parser_Attr_grind___closed__5(); +lean_mark_persistent(l_Lean_Parser_Attr_grind___closed__5); +l_Lean_Parser_Attr_grind___closed__6 = _init_l_Lean_Parser_Attr_grind___closed__6(); +lean_mark_persistent(l_Lean_Parser_Attr_grind___closed__6); +l_Lean_Parser_Attr_grind___closed__7 = _init_l_Lean_Parser_Attr_grind___closed__7(); +lean_mark_persistent(l_Lean_Parser_Attr_grind___closed__7); +l_Lean_Parser_Attr_grind___closed__8 = _init_l_Lean_Parser_Attr_grind___closed__8(); +lean_mark_persistent(l_Lean_Parser_Attr_grind___closed__8); +l_Lean_Parser_Attr_grind___closed__9 = _init_l_Lean_Parser_Attr_grind___closed__9(); +lean_mark_persistent(l_Lean_Parser_Attr_grind___closed__9); +l_Lean_Parser_Attr_grind___closed__10 = _init_l_Lean_Parser_Attr_grind___closed__10(); +lean_mark_persistent(l_Lean_Parser_Attr_grind___closed__10); +l_Lean_Parser_Attr_grind___closed__11 = _init_l_Lean_Parser_Attr_grind___closed__11(); +lean_mark_persistent(l_Lean_Parser_Attr_grind___closed__11); +l_Lean_Parser_Attr_grind___closed__12 = _init_l_Lean_Parser_Attr_grind___closed__12(); +lean_mark_persistent(l_Lean_Parser_Attr_grind___closed__12); +l_Lean_Parser_Attr_grind___closed__13 = _init_l_Lean_Parser_Attr_grind___closed__13(); +lean_mark_persistent(l_Lean_Parser_Attr_grind___closed__13); +l_Lean_Parser_Attr_grind___closed__14 = _init_l_Lean_Parser_Attr_grind___closed__14(); +lean_mark_persistent(l_Lean_Parser_Attr_grind___closed__14); +l_Lean_Parser_Attr_grind = _init_l_Lean_Parser_Attr_grind(); +lean_mark_persistent(l_Lean_Parser_Attr_grind); l_Lean_Grind_instInhabitedConfig___closed__1 = _init_l_Lean_Grind_instInhabitedConfig___closed__1(); lean_mark_persistent(l_Lean_Grind_instInhabitedConfig___closed__1); l_Lean_Grind_instInhabitedConfig = _init_l_Lean_Grind_instInhabitedConfig(); @@ -329,6 +1237,10 @@ l_Lean_Parser_Tactic_grind___closed__9 = _init_l_Lean_Parser_Tactic_grind___clos lean_mark_persistent(l_Lean_Parser_Tactic_grind___closed__9); l_Lean_Parser_Tactic_grind___closed__10 = _init_l_Lean_Parser_Tactic_grind___closed__10(); lean_mark_persistent(l_Lean_Parser_Tactic_grind___closed__10); +l_Lean_Parser_Tactic_grind___closed__11 = _init_l_Lean_Parser_Tactic_grind___closed__11(); +lean_mark_persistent(l_Lean_Parser_Tactic_grind___closed__11); +l_Lean_Parser_Tactic_grind___closed__12 = _init_l_Lean_Parser_Tactic_grind___closed__12(); +lean_mark_persistent(l_Lean_Parser_Tactic_grind___closed__12); l_Lean_Parser_Tactic_grind = _init_l_Lean_Parser_Tactic_grind(); lean_mark_persistent(l_Lean_Parser_Tactic_grind); return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Init/Grind/Util.c b/stage0/stdlib/Init/Grind/Util.c index 02d8bb9beb3d..9a31181dd992 100644 --- a/stage0/stdlib/Init/Grind/Util.c +++ b/stage0/stdlib/Init/Grind/Util.c @@ -13,6 +13,54 @@ #ifdef __cplusplus extern "C" { #endif +LEAN_EXPORT lean_object* l_Lean_Grind_doNotSimp(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Grind_offset(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Grind_doNotSimp___rarg___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Grind_offset___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Grind_doNotSimp___rarg(lean_object*); +lean_object* lean_nat_add(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Grind_doNotSimp___rarg(lean_object* x_1) { +_start: +{ +lean_inc(x_1); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Grind_doNotSimp(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_Grind_doNotSimp___rarg___boxed), 1, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Grind_doNotSimp___rarg___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Lean_Grind_doNotSimp___rarg(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Grind_offset(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_nat_add(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Grind_offset___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_Grind_offset(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} lean_object* initialize_Init_Core(uint8_t builtin, lean_object*); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Init_Grind_Util(uint8_t builtin, lean_object* w) { diff --git a/stage0/stdlib/Init/Prelude.c b/stage0/stdlib/Init/Prelude.c index 350b5f3c331a..3dd6be75eba9 100644 --- a/stage0/stdlib/Init/Prelude.c +++ b/stage0/stdlib/Init/Prelude.c @@ -13,6 +13,7 @@ #ifdef __cplusplus extern "C" { #endif +LEAN_EXPORT lean_object* l_Lean_withRef_x3f___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_instMulNat___closed__1; LEAN_EXPORT lean_object* l_instInhabitedReaderT___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_ReaderT_read(lean_object*, lean_object*); @@ -729,6 +730,7 @@ LEAN_EXPORT lean_object* l_decEq(lean_object*); LEAN_EXPORT lean_object* l_EStateM_instOrElseOfBacktrackable(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instMonadRefOfMonadLiftOfMonadFunctor___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_instMonadLiftTOfMonadLift(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withRef_x3f(lean_object*); LEAN_EXPORT lean_object* l_instLTNat; LEAN_EXPORT lean_object* l_instDecidableAnd(lean_object*, lean_object*); static lean_object* l_EStateM_instMonad___closed__6; @@ -9603,6 +9605,43 @@ lean_dec(x_1); return x_5; } } +LEAN_EXPORT lean_object* l_Lean_withRef_x3f___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_4) == 0) +{ +lean_dec(x_2); +lean_dec(x_1); +return x_5; +} +else +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_6 = lean_ctor_get(x_4, 0); +lean_inc(x_6); +lean_dec(x_4); +x_7 = lean_ctor_get(x_1, 1); +lean_inc(x_7); +lean_dec(x_1); +x_8 = lean_ctor_get(x_2, 0); +lean_inc(x_8); +x_9 = lean_alloc_closure((void*)(l_Lean_withRef___rarg___lambda__1___boxed), 4, 3); +lean_closure_set(x_9, 0, x_6); +lean_closure_set(x_9, 1, x_2); +lean_closure_set(x_9, 2, x_5); +x_10 = lean_apply_4(x_7, lean_box(0), lean_box(0), x_8, x_9); +return x_10; +} +} +} +LEAN_EXPORT lean_object* l_Lean_withRef_x3f(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_withRef_x3f___rarg), 5, 0); +return x_2; +} +} LEAN_EXPORT lean_object* l_Lean_MonadRef_mkInfoFromRefPos___rarg___lambda__1(lean_object* x_1, lean_object* x_2) { _start: { diff --git a/stage0/stdlib/Init/Tactics.c b/stage0/stdlib/Init/Tactics.c index 1aae711684f2..1c0cc09ecdfc 100644 --- a/stage0/stdlib/Init/Tactics.c +++ b/stage0/stdlib/Init/Tactics.c @@ -16374,7 +16374,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_inductionAlts___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_revert___closed__5; +x_1 = l_Lean_Parser_Tactic_intro___closed__11; x_2 = l_Lean_Parser_Tactic_inductionAlts___closed__6; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); diff --git a/stage0/stdlib/Lake/Load/Lean/Elab.c b/stage0/stdlib/Lake/Load/Lean/Elab.c index dc13c42845c8..6426e614fb52 100644 --- a/stage0/stdlib/Lake/Load/Lean/Elab.c +++ b/stage0/stdlib/Lake/Load/Lean/Elab.c @@ -134,7 +134,6 @@ lean_object* lean_st_mk_ref(lean_object*, lean_object*); lean_object* l_System_FilePath_fileName(lean_object*); static lean_object* l_Lake_importConfigFileCore_lakeExts___closed__34; static lean_object* l_Lake_importConfigFileCore_lakeExts___closed__23; -lean_object* l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(lean_object*, lean_object*); static uint64_t l_Lake_importModulesUsingCache___closed__3; static lean_object* l_Lake_importConfigFileCore_lakeExts___closed__18; static lean_object* l_Lake_importConfigFileCore_lakeExts___closed__5; @@ -285,6 +284,7 @@ static lean_object* l___private_Lake_Load_Lean_Elab_0__Lake_fromJsonConfigTrace_ lean_object* lean_environment_add(lean_object*, lean_object*); lean_object* lean_array_uset(lean_object*, size_t, lean_object*); static lean_object* l_Lake_importConfigFileCore_lakeExts___closed__28; +lean_object* l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lake_Load_Lean_Elab_0__Lake_addToEnv___boxed(lean_object*, lean_object*); lean_object* l_Lean_MessageLog_add(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lake_importModulesUsingCache___spec__3(lean_object*); @@ -4643,7 +4643,7 @@ x_30 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_30, 0, x_10); lean_ctor_set(x_30, 1, x_29); x_31 = l___private_Lake_Load_Lean_Elab_0__Lake_toJsonConfigTrace____x40_Lake_Load_Lean_Elab___hyg_1030____closed__4; -x_32 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_30, x_31); +x_32 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_30, x_31); x_33 = l_Lean_Json_mkObj(x_32); return x_33; } diff --git a/stage0/stdlib/Lean/Data/Json/Printer.c b/stage0/stdlib/Lean/Data/Json/Printer.c index 39b0a4548a8a..8131e3c14116 100644 --- a/stage0/stdlib/Lean/Data/Json/Printer.c +++ b/stage0/stdlib/Lean/Data/Json/Printer.c @@ -13,89 +13,3713 @@ #ifdef __cplusplus extern "C" { #endif +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__44; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__254; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__245; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__151; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__111; lean_object* lean_format_pretty(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__99; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__221; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__141; LEAN_EXPORT lean_object* l_Lean_Json_compress(lean_object*); +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__94; uint32_t lean_string_utf8_get(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__181; +lean_object* lean_byte_array_mk(lean_object*); +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__114; static lean_object* l_Lean_Json_compress_go___closed__1; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__13; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__137; lean_object* l_Lean_JsonNumber_toString(lean_object*); +uint8_t lean_byte_array_fget(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__219; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__242; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__183; lean_object* lean_uint32_to_nat(uint32_t); +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__46; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__159; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__7; LEAN_EXPORT lean_object* l_Lean_RBNode_fold___at_Lean_Json_compress_go___spec__3(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__47; LEAN_EXPORT lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeAux___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__1; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__85; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__83; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__119; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__77; static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeAux___closed__3; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__30; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__26; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__154; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__62; static lean_object* l_Lean_Json_render___closed__9; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__168; static lean_object* l_Lean_Json_render___closed__19; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__96; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__147; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__63; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__189; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__21; uint8_t lean_usize_dec_eq(size_t, size_t); +uint8_t lean_string_get_byte_fast(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__173; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__11; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__27; static lean_object* l_Lean_Json_renderString___closed__1; static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeAux___closed__2; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__110; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__18; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__176; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__106; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__53; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__121; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__201; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__184; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__157; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__41; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__207; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__233; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__134; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__203; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__214; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__19; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__138; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__58; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__103; static lean_object* l_Lean_Json_render___closed__12; lean_object* lean_string_utf8_byte_size(lean_object*); lean_object* lean_string_push(lean_object*, uint32_t); static lean_object* l_Lean_Json_render___closed__20; static lean_object* l_Lean_RBNode_fold___at_Lean_Json_render___spec__2___closed__1; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__224; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__186; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__24; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__101; static lean_object* l_Lean_Json_render___closed__6; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__171; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__229; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__78; +LEAN_EXPORT uint8_t l___private_Lean_Data_Json_Printer_0__Lean_Json_needEscape(lean_object*); +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__158; lean_object* l_List_appendTR___rarg(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__23; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__59; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__105; size_t lean_usize_of_nat(lean_object*); +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__206; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__116; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__175; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__196; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__236; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__35; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__38; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__9; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__167; static lean_object* l_Lean_Json_render___closed__14; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__162; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__10; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__247; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__240; lean_object* lean_string_utf8_next(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__204; LEAN_EXPORT lean_object* l_Lean_Json_escape(lean_object*, lean_object*); uint32_t l_Nat_digitChar(lean_object*); +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__97; LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Lean_Json_compress_go___spec__2(lean_object*, size_t, size_t, lean_object*); uint8_t lean_uint32_dec_le(uint32_t, uint32_t); lean_object* lean_nat_to_int(lean_object*); +LEAN_EXPORT uint8_t l___private_Lean_Data_Json_Printer_0__Lean_Json_needEscape_go(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__239; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__163; lean_object* lean_nat_div(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__256; static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeAux___closed__1; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__28; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__33; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__234; static lean_object* l_Lean_Json_render___closed__17; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__74; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__61; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__252; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__37; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__197; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__250; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__87; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__39; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__238; static lean_object* l_Lean_Json_render___closed__4; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__31; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__212; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__113; LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Lean_Json_compress_go___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__216; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__257; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__195; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__100; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__42; static lean_object* l_Lean_Json_render___closed__11; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__79; lean_object* lean_array_to_list(lean_object*); +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__6; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__249; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Json_compress_go___spec__1(size_t, size_t, lean_object*); static lean_object* l_Lean_Json_render___closed__16; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__205; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__73; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__227; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__161; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__228; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__142; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__199; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__172; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__57; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__188; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__215; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__15; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__133; static lean_object* l_Lean_Json_render___closed__13; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__155; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__65; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__131; LEAN_EXPORT lean_object* l_Lean_Json_renderString___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__169; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__120; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__25; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__8; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__153; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__235; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__144; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__48; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__71; static lean_object* l_Lean_Json_instToFormat___closed__1; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__190; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__36; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__66; static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeAux___closed__5; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__108; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__180; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__2; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__200; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__156; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__243; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__67; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__135; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__64; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__202; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__217; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__84; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__223; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__160; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__68; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__55; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__177; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__45; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Json_render___spec__1___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__81; static lean_object* l_Lean_Json_render___closed__5; static lean_object* l_Lean_Json_render___closed__7; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__248; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__112; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__148; static lean_object* l_Lean_Json_render___closed__15; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__107; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__92; static lean_object* l_Lean_Json_render___closed__3; static lean_object* l_Lean_Json_render___closed__10; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__70; LEAN_EXPORT lean_object* l_String_foldlAux___at_Lean_Json_escape___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__86; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__98; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__222; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__130; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__3; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__237; static lean_object* l_Lean_Json_render___closed__1; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__211; +LEAN_EXPORT lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_needEscape___boxed(lean_object*); +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__17; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__76; lean_object* lean_string_length(lean_object*); +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__244; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__191; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__139; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Json_compress_go___spec__1___boxed(lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__122; lean_object* lean_nat_mod(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__4; LEAN_EXPORT lean_object* l_Lean_Json_instToString(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Json_render___spec__1(size_t, size_t, lean_object*); uint8_t lean_uint32_dec_eq(uint32_t, uint32_t); LEAN_EXPORT lean_object* l_Lean_RBNode_fold___at_Lean_Json_compress_go___spec__3___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__90; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__170; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__251; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__124; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__72; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__20; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__166; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__150; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__128; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__43; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__54; lean_object* l_Std_Format_joinSep___at_Prod_repr___spec__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__230; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__56; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__93; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__209; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__198; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__91; LEAN_EXPORT lean_object* l_String_foldlAux___at_Lean_Json_escape___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_renderString(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_compress_go(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__89; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__208; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__34; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__225; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__226; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__52; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__49; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__115; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__146; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__149; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__80; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__145; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__213; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__102; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__232; static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeAux___closed__4; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__187; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__193; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__185; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__136; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__246; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__129; size_t lean_usize_sub(size_t, size_t); +lean_object* lean_array_mk(lean_object*); +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__164; +lean_object* lean_uint8_to_nat(uint8_t); size_t lean_usize_add(size_t, size_t); +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__140; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__104; static lean_object* l_Lean_Json_render___closed__18; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__220; lean_object* lean_array_uget(lean_object*, size_t); size_t lean_array_size(lean_object*); +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__32; LEAN_EXPORT lean_object* l_Lean_Json_escape___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__132; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__174; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__50; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__253; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__165; static lean_object* l_Lean_Json_render___closed__8; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__5; lean_object* lean_string_append(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__125; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__194; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__117; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__241; lean_object* lean_array_get_size(lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_render(lean_object*); +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__51; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__88; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__82; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__22; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__40; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__60; static lean_object* l_Lean_Json_render___closed__2; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__109; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__29; uint8_t lean_usize_dec_lt(size_t, size_t); static lean_object* l_Lean_RBNode_fold___at_Lean_Json_render___spec__2___closed__2; +LEAN_EXPORT lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_needEscape_go___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__14; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__143; +lean_object* lean_nat_add(lean_object*, lean_object*); static lean_object* l_Lean_Json_render___closed__21; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__95; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__75; LEAN_EXPORT lean_object* l_Lean_RBNode_fold___at_Lean_Json_render___spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_instToFormat; LEAN_EXPORT lean_object* l_Lean_Json_pretty(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__69; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__126; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__218; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__118; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__258; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__179; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__16; +uint8_t lean_uint8_dec_eq(uint8_t, uint8_t); +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__178; lean_object* lean_array_uset(lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeAux(lean_object*, uint32_t); static lean_object* l_Lean_RBNode_fold___at_Lean_Json_render___spec__2___closed__3; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__12; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__182; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__255; +LEAN_EXPORT lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__152; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__192; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__231; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__123; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__210; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__127; +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__1() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__2() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__1; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__3() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__2; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__4() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__3; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__4; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__6() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__5; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__7() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__6; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__8() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__7; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__9() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__8; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__10() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__9; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___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___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__10; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__12() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__11; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__13() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__12; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__14() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__13; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__15() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__14; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__16() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__15; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__17() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__16; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__18() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__17; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__19() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__18; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__20() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__19; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__21() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__20; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__22() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__21; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__23() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__22; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__24() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__23; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__25() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__24; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__26() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__25; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__27() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__26; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__28() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__27; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__29() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__28; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__30() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__29; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__31() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__30; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__32() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__31; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__33() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__32; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__34() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__33; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__35() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__34; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__36() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__35; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__37() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__36; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__38() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__37; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__39() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__38; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__40() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__39; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__41() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__40; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__42() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__41; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__43() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__42; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__44() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__43; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__45() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__44; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__46() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__45; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__47() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__46; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__48() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__47; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__49() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__48; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__50() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__49; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__51() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__50; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__52() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__51; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__53() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__52; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__54() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__53; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__55() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__54; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__56() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__55; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__57() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__56; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__58() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__57; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__59() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__58; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__60() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__59; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__61() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__60; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__62() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__61; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__63() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__62; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__64() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__63; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__65() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__64; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__66() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__65; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__67() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__66; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__68() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__67; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__69() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__68; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__70() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__69; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__71() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__70; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__72() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__71; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__73() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__72; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__74() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__73; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__75() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__74; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__76() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__75; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__77() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__76; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__78() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__77; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__79() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__78; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__80() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__79; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__81() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__80; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__82() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__81; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__83() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__82; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__84() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__83; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__85() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__84; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__86() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__85; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__87() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__86; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__88() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__87; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__89() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__88; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__90() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__89; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__91() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__90; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__92() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__91; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__93() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__92; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__94() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__93; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__95() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__94; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__96() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__95; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__97() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__96; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__98() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__97; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__99() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__98; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__100() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__99; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__101() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__100; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__102() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__101; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__103() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__102; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__104() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__103; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__105() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__104; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__106() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__105; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__107() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__106; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__108() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__107; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__109() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__108; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__110() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__109; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__111() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__110; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__112() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__111; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__113() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__112; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__114() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__113; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__115() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__114; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__116() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__115; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__117() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__116; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__118() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__117; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__119() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__118; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__120() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__119; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__121() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__120; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__122() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__121; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__123() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__122; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__124() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__123; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__125() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__124; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__126() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__125; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__127() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__126; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__128() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__127; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__129() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; +x_1 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__128; +x_2 = 0; +x_3 = lean_box(x_2); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_1); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__130() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__129; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__131() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__130; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__132() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__131; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__133() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__132; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__134() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__133; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__135() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__134; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__136() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__135; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__137() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__136; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__138() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__137; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__139() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__138; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__140() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__139; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__141() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__140; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__142() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__141; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__143() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__142; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__144() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__143; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__145() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__144; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__146() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__145; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__147() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__146; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__148() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__147; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__149() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__148; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__150() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__149; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__151() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__150; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__152() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__151; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__153() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__152; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__154() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__153; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__155() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__154; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__156() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__155; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__157() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__156; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__158() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__157; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__159() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__158; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__160() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__159; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__161() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__160; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__162() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__161; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__163() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__162; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__164() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__163; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__165() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__164; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__166() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__165; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__167() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__166; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__168() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__167; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__169() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__168; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__170() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__169; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__171() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__170; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__172() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__171; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__173() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__172; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__174() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__173; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__175() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__174; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__176() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__175; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__177() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__176; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__178() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__177; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__179() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__178; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__180() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__179; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__181() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__180; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__182() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__181; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__183() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__182; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__184() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__183; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__185() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__184; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__186() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__185; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__187() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__186; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__188() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__187; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__189() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__188; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__190() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__189; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__191() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__190; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__192() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__191; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__193() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__192; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__194() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__193; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__195() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__194; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__196() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__195; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__197() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__196; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__198() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__197; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__199() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__198; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__200() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__199; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__201() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__200; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__202() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__201; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__203() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__202; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__204() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__203; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__205() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__204; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__206() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__205; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__207() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__206; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__208() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__207; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__209() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__208; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__210() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__209; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__211() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__210; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__212() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__211; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__213() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__212; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__214() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__213; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__215() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__214; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__216() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__215; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__217() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__216; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__218() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__217; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__219() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__218; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__220() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__219; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__221() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__220; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__222() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__221; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__223() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__222; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__224() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__223; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__225() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__224; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__226() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__225; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__227() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__226; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__228() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__227; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__229() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__228; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__230() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__229; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__231() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__230; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__232() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__231; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__233() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__232; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__234() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__233; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__235() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__234; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__236() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__235; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__237() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__236; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__238() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__237; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__239() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__238; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__240() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__239; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__241() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__240; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__242() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__241; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__243() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__242; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__244() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__243; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__245() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__244; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__246() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__245; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__247() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__246; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__248() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__247; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__249() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__248; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__250() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__249; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__251() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__250; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__252() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__251; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__253() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__252; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__254() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__253; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__255() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__254; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__256() { +_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_Data_Json_Printer_0__Lean_Json_escapeTable___closed__255; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__257() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__256; +x_2 = lean_array_mk(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__258() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__257; +x_2 = lean_byte_array_mk(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable() { +_start: +{ +lean_object* x_1; +x_1 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__258; +return x_1; +} +} static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeAux___closed__1() { _start: { @@ -265,6 +3889,79 @@ x_4 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeAux(x_1, x_3); return x_4; } } +LEAN_EXPORT uint8_t l___private_Lean_Data_Json_Printer_0__Lean_Json_needEscape_go(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; uint8_t x_4; +x_3 = lean_string_utf8_byte_size(x_1); +x_4 = lean_nat_dec_lt(x_2, x_3); +lean_dec(x_3); +if (x_4 == 0) +{ +uint8_t x_5; +lean_dec(x_2); +x_5 = 0; +return x_5; +} +else +{ +uint8_t x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; uint8_t x_10; uint8_t x_11; +lean_inc(x_2); +x_6 = lean_string_get_byte_fast(x_1, x_2); +x_7 = lean_uint8_to_nat(x_6); +x_8 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable; +x_9 = lean_byte_array_fget(x_8, x_7); +lean_dec(x_7); +x_10 = 0; +x_11 = lean_uint8_dec_eq(x_9, x_10); +if (x_11 == 0) +{ +uint8_t x_12; +lean_dec(x_2); +x_12 = 1; +return x_12; +} +else +{ +lean_object* x_13; lean_object* x_14; +x_13 = lean_unsigned_to_nat(1u); +x_14 = lean_nat_add(x_2, x_13); +lean_dec(x_2); +x_2 = x_14; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_needEscape_go___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l___private_Lean_Data_Json_Printer_0__Lean_Json_needEscape_go(x_1, x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT uint8_t l___private_Lean_Data_Json_Printer_0__Lean_Json_needEscape(lean_object* x_1) { +_start: +{ +lean_object* x_2; uint8_t x_3; +x_2 = lean_unsigned_to_nat(0u); +x_3 = l___private_Lean_Data_Json_Printer_0__Lean_Json_needEscape_go(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_needEscape___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_needEscape(x_1); +lean_dec(x_1); +x_3 = lean_box(x_2); +return x_3; +} +} LEAN_EXPORT lean_object* l_String_foldlAux___at_Lean_Json_escape___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { @@ -291,13 +3988,24 @@ goto _start; LEAN_EXPORT lean_object* l_Lean_Json_escape(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_3 = lean_string_utf8_byte_size(x_1); -x_4 = lean_unsigned_to_nat(0u); -x_5 = l_String_foldlAux___at_Lean_Json_escape___spec__1(x_1, x_3, x_4, x_2); -lean_dec(x_3); +lean_object* x_3; uint8_t x_4; +x_3 = lean_unsigned_to_nat(0u); +x_4 = l___private_Lean_Data_Json_Printer_0__Lean_Json_needEscape_go(x_1, x_3); +if (x_4 == 0) +{ +lean_object* x_5; +x_5 = lean_string_append(x_2, x_1); return x_5; } +else +{ +lean_object* x_6; lean_object* x_7; +x_6 = lean_string_utf8_byte_size(x_1); +x_7 = l_String_foldlAux___at_Lean_Json_escape___spec__1(x_1, x_6, x_3, x_2); +lean_dec(x_6); +return x_7; +} +} } LEAN_EXPORT lean_object* l_String_foldlAux___at_Lean_Json_escape___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: @@ -2030,6 +5738,524 @@ 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); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__1 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__1(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__1); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__2 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__2(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__2); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__3 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__3(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__3); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__4 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__4(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__4); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__5 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__5(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__5); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__6 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__6(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__6); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__7 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__7(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__7); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__8 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__8(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__8); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__9 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__9(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__9); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__10 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__10(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__10); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__11 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__11(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__11); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__12 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__12(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__12); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__13 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__13(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__13); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__14 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__14(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__14); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__15 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__15(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__15); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__16 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__16(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__16); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__17 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__17(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__17); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__18 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__18(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__18); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__19 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__19(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__19); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__20 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__20(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__20); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__21 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__21(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__21); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__22 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__22(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__22); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__23 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__23(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__23); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__24 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__24(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__24); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__25 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__25(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__25); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__26 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__26(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__26); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__27 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__27(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__27); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__28 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__28(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__28); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__29 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__29(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__29); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__30 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__30(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__30); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__31 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__31(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__31); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__32 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__32(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__32); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__33 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__33(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__33); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__34 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__34(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__34); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__35 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__35(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__35); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__36 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__36(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__36); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__37 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__37(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__37); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__38 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__38(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__38); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__39 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__39(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__39); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__40 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__40(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__40); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__41 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__41(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__41); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__42 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__42(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__42); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__43 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__43(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__43); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__44 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__44(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__44); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__45 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__45(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__45); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__46 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__46(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__46); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__47 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__47(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__47); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__48 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__48(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__48); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__49 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__49(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__49); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__50 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__50(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__50); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__51 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__51(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__51); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__52 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__52(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__52); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__53 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__53(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__53); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__54 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__54(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__54); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__55 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__55(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__55); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__56 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__56(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__56); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__57 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__57(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__57); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__58 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__58(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__58); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__59 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__59(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__59); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__60 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__60(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__60); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__61 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__61(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__61); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__62 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__62(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__62); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__63 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__63(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__63); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__64 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__64(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__64); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__65 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__65(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__65); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__66 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__66(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__66); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__67 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__67(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__67); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__68 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__68(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__68); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__69 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__69(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__69); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__70 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__70(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__70); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__71 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__71(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__71); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__72 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__72(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__72); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__73 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__73(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__73); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__74 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__74(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__74); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__75 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__75(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__75); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__76 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__76(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__76); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__77 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__77(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__77); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__78 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__78(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__78); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__79 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__79(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__79); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__80 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__80(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__80); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__81 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__81(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__81); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__82 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__82(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__82); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__83 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__83(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__83); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__84 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__84(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__84); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__85 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__85(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__85); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__86 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__86(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__86); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__87 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__87(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__87); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__88 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__88(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__88); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__89 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__89(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__89); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__90 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__90(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__90); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__91 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__91(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__91); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__92 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__92(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__92); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__93 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__93(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__93); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__94 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__94(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__94); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__95 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__95(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__95); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__96 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__96(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__96); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__97 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__97(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__97); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__98 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__98(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__98); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__99 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__99(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__99); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__100 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__100(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__100); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__101 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__101(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__101); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__102 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__102(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__102); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__103 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__103(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__103); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__104 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__104(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__104); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__105 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__105(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__105); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__106 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__106(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__106); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__107 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__107(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__107); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__108 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__108(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__108); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__109 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__109(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__109); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__110 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__110(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__110); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__111 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__111(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__111); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__112 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__112(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__112); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__113 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__113(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__113); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__114 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__114(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__114); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__115 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__115(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__115); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__116 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__116(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__116); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__117 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__117(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__117); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__118 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__118(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__118); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__119 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__119(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__119); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__120 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__120(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__120); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__121 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__121(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__121); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__122 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__122(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__122); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__123 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__123(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__123); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__124 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__124(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__124); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__125 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__125(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__125); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__126 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__126(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__126); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__127 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__127(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__127); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__128 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__128(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__128); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__129 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__129(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__129); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__130 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__130(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__130); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__131 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__131(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__131); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__132 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__132(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__132); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__133 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__133(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__133); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__134 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__134(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__134); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__135 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__135(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__135); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__136 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__136(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__136); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__137 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__137(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__137); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__138 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__138(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__138); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__139 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__139(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__139); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__140 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__140(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__140); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__141 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__141(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__141); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__142 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__142(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__142); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__143 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__143(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__143); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__144 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__144(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__144); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__145 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__145(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__145); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__146 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__146(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__146); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__147 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__147(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__147); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__148 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__148(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__148); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__149 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__149(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__149); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__150 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__150(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__150); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__151 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__151(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__151); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__152 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__152(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__152); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__153 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__153(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__153); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__154 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__154(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__154); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__155 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__155(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__155); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__156 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__156(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__156); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__157 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__157(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__157); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__158 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__158(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__158); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__159 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__159(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__159); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__160 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__160(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__160); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__161 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__161(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__161); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__162 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__162(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__162); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__163 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__163(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__163); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__164 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__164(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__164); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__165 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__165(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__165); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__166 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__166(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__166); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__167 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__167(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__167); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__168 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__168(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__168); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__169 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__169(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__169); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__170 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__170(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__170); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__171 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__171(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__171); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__172 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__172(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__172); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__173 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__173(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__173); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__174 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__174(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__174); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__175 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__175(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__175); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__176 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__176(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__176); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__177 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__177(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__177); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__178 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__178(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__178); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__179 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__179(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__179); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__180 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__180(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__180); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__181 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__181(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__181); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__182 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__182(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__182); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__183 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__183(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__183); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__184 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__184(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__184); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__185 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__185(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__185); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__186 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__186(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__186); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__187 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__187(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__187); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__188 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__188(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__188); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__189 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__189(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__189); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__190 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__190(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__190); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__191 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__191(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__191); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__192 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__192(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__192); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__193 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__193(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__193); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__194 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__194(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__194); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__195 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__195(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__195); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__196 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__196(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__196); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__197 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__197(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__197); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__198 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__198(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__198); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__199 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__199(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__199); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__200 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__200(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__200); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__201 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__201(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__201); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__202 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__202(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__202); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__203 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__203(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__203); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__204 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__204(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__204); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__205 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__205(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__205); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__206 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__206(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__206); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__207 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__207(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__207); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__208 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__208(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__208); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__209 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__209(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__209); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__210 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__210(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__210); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__211 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__211(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__211); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__212 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__212(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__212); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__213 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__213(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__213); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__214 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__214(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__214); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__215 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__215(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__215); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__216 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__216(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__216); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__217 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__217(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__217); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__218 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__218(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__218); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__219 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__219(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__219); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__220 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__220(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__220); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__221 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__221(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__221); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__222 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__222(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__222); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__223 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__223(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__223); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__224 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__224(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__224); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__225 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__225(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__225); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__226 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__226(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__226); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__227 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__227(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__227); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__228 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__228(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__228); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__229 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__229(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__229); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__230 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__230(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__230); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__231 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__231(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__231); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__232 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__232(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__232); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__233 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__233(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__233); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__234 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__234(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__234); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__235 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__235(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__235); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__236 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__236(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__236); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__237 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__237(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__237); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__238 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__238(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__238); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__239 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__239(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__239); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__240 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__240(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__240); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__241 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__241(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__241); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__242 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__242(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__242); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__243 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__243(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__243); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__244 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__244(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__244); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__245 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__245(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__245); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__246 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__246(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__246); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__247 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__247(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__247); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__248 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__248(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__248); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__249 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__249(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__249); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__250 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__250(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__250); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__251 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__251(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__251); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__252 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__252(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__252); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__253 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__253(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__253); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__254 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__254(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__254); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__255 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__255(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__255); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__256 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__256(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__256); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__257 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__257(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__257); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__258 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__258(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__258); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable); l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeAux___closed__1 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeAux___closed__1(); lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeAux___closed__1); l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeAux___closed__2 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeAux___closed__2(); diff --git a/stage0/stdlib/Lean/Data/Lsp.c b/stage0/stdlib/Lean/Data/Lsp.c index 0055ee3f3b15..bd31f38bcb5e 100644 --- a/stage0/stdlib/Lean/Data/Lsp.c +++ b/stage0/stdlib/Lean/Data/Lsp.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Lean.Data.Lsp -// Imports: Lean.Data.Lsp.Basic Lean.Data.Lsp.Capabilities Lean.Data.Lsp.Client Lean.Data.Lsp.Communication Lean.Data.Lsp.Diagnostics Lean.Data.Lsp.Extra Lean.Data.Lsp.InitShutdown Lean.Data.Lsp.Internal Lean.Data.Lsp.LanguageFeatures Lean.Data.Lsp.TextSync Lean.Data.Lsp.Utf16 Lean.Data.Lsp.Workspace Lean.Data.Lsp.Ipc Lean.Data.Lsp.CodeActions +// Imports: Lean.Data.Lsp.Basic Lean.Data.Lsp.CancelParams Lean.Data.Lsp.Capabilities Lean.Data.Lsp.Client Lean.Data.Lsp.Communication Lean.Data.Lsp.Diagnostics Lean.Data.Lsp.Extra Lean.Data.Lsp.InitShutdown Lean.Data.Lsp.Internal Lean.Data.Lsp.LanguageFeatures Lean.Data.Lsp.TextSync Lean.Data.Lsp.Utf16 Lean.Data.Lsp.Workspace Lean.Data.Lsp.Ipc Lean.Data.Lsp.CodeActions #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -14,6 +14,7 @@ extern "C" { #endif lean_object* initialize_Lean_Data_Lsp_Basic(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Data_Lsp_CancelParams(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Data_Lsp_Capabilities(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Data_Lsp_Client(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Data_Lsp_Communication(uint8_t builtin, lean_object*); @@ -35,6 +36,9 @@ _G_initialized = true; res = initialize_Lean_Data_Lsp_Basic(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +res = initialize_Lean_Data_Lsp_CancelParams(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); res = initialize_Lean_Data_Lsp_Capabilities(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); diff --git a/stage0/stdlib/Lean/Data/Lsp/Basic.c b/stage0/stdlib/Lean/Data/Lsp/Basic.c index 0b3c563a5b67..fbfb4b502ed4 100644 --- a/stage0/stdlib/Lean/Data/Lsp/Basic.c +++ b/stage0/stdlib/Lean/Data/Lsp/Basic.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Lean.Data.Lsp.Basic -// Imports: Lean.Data.Json Lean.Data.JsonRpc +// Imports: Lean.Data.Json #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -13,1358 +13,858 @@ #ifdef __cplusplus extern "C" { #endif +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__7; lean_object* l_Lean_JsonNumber_fromNat(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__9; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__13; static lean_object* l_Lean_Lsp_instToJsonMarkupContent___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__7; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693_(lean_object*); static lean_object* l_Lean_Lsp_instBEqRange___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__7; LEAN_EXPORT lean_object* l_Lean_Lsp_MarkupKind_noConfusion___rarg___lambda__1___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_CreateFile_instFromJsonOptions; static lean_object* l_Lean_Lsp_instToStringPosition___closed__2; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3906_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__5; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonChangeAnnotation; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__8; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____closed__1; -LEAN_EXPORT lean_object* l_Lean_RBNode_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____spec__2(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__9; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__15; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__12; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6641____closed__1; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____spec__2(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__5; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__5; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__15; lean_object* l_Lean_Json_getObj_x3f(lean_object*); static lean_object* l_Lean_Lsp_instToJsonStaticRegistrationOptions___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__8; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__9; LEAN_EXPORT lean_object* l_Lean_Lsp_instInhabitedLocation; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__14___boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__9; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__10; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__11; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonTextDocumentPositionParams; static lean_object* l_Lean_Lsp_instToJsonVersionedTextDocumentIdentifier___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__8; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____spec__1(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__11; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____closed__2; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__13; -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__25___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__26(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonCancelParams; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____closed__2; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6531____rarg___closed__1; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____spec__1___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__5; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__6; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____spec__1___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__7; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__1; +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__19(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__1; static lean_object* l_Lean_Lsp_instInhabitedRange___closed__1; static lean_object* l_Lean_Lsp_instFromJsonPartialResultParams___closed__1; -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3399____spec__1___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__21___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_5974____boxed(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3721____spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonProgressParams(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__1; -LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_627_(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__3; -static lean_object* l_Lean_Lsp_instToJsonCancelParams___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordLocation____x40_Lean_Data_Lsp_Basic___hyg_1238____boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__13; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__7; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5505____closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupKind____x40_Lean_Data_Lsp_Basic___hyg_6036____boxed(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__9; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__1; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__4; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__4; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__9; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__6; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__3; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__7; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____lambda__1___boxed(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6456_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__18; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__1; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__1; static lean_object* l_Lean_Lsp_instHashablePosition___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__5; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__6; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonDocumentFilter; LEAN_EXPORT lean_object* l_Lean_Lsp_instToStringTextDocumentPositionParams(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__1; static lean_object* l_Lean_Lsp_instFromJsonCommand___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3024____boxed(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__6; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__6; -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__21___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; -LEAN_EXPORT lean_object* l_Lean_Lsp_instBEqCancelParams; -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3906____spec__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__15; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3214_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__3; +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__23___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__24(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1655_(lean_object*); lean_object* l_Lean_Json_mkObj(lean_object*); static lean_object* l_Lean_Lsp_DeleteFile_instToJsonOptions___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__12; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3399_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__5; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__8; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__17; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__2; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__7; +LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqLocation____x40_Lean_Data_Lsp_Basic___hyg_821_(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__4___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__5(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__4; LEAN_EXPORT lean_object* l_Lean_Lsp_instHashableMarkupContent; LEAN_EXPORT lean_object* l_Lean_Lsp_instHashableMarkupKind; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__3; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__9; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__6; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__3; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__4; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__11; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__16; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__2; +LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashPosition____x40_Lean_Data_Lsp_Basic___hyg_180_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonDocumentSelector(lean_object*); -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__10___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__11(lean_object*, lean_object*); -LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordRange____x40_Lean_Data_Lsp_Basic___hyg_900_(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__5; +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__27___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__28(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instToJsonDocumentChange___closed__4; -static lean_object* l_Lean_Lsp_instInhabitedCancelParams___closed__1; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__5; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__6; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6346____rarg___closed__1; static lean_object* l_Lean_Lsp_instOrdRange___closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_MarkupKind_toCtorIdx(uint8_t); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__9; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__13; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__8; uint64_t lean_uint64_of_nat(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonPosition; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__4; uint64_t lean_uint64_mix_hash(uint64_t, uint64_t); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__3; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__6; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__12; static lean_object* l_Lean_Lsp_instToJsonTextDocumentItem___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__12___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__13(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__12; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__5; static lean_object* l_Lean_Lsp_instToJsonMarkupKind___closed__1; static lean_object* l_Lean_Lsp_instFromJsonApplyWorkspaceEditParams___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__4; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__16; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__9; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__11; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__5; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3208____closed__2; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__8; -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__27(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__5; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_627____boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__2; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895____closed__2; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__7; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3445_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__27___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__8; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__19; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__1; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__7; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__1; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____spec__2___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__19; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonMarkupContent; static lean_object* l_Lean_Lsp_WorkspaceEdit_instAppend___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__8; -LEAN_EXPORT lean_object* l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__18(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____spec__1___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155____closed__3; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_442____boxed(lean_object*, lean_object*); lean_object* l_Lean_Name_toString(lean_object*, uint8_t, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3023____closed__2; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__10; static lean_object* l_Lean_Lsp_instFromJsonCreateFile___closed__1; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__10; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__9; +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__23___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonWorkDoneProgressOptions; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__6; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__1; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__9; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3721_(lean_object*); uint8_t l_Lean_RBNode_isRed___rarg(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__11; LEAN_EXPORT uint8_t l_Lean_Lsp_MarkupKind_ofNat(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__3; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonWorkDoneProgressReport; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__6; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__3; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Lsp_instFromJsonTextEditBatch___spec__1(size_t, size_t, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__2; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__11; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonTextDocumentEdit; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressBegin____x40_Lean_Data_Lsp_Basic___hyg_6722_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(lean_object*, lean_object*); -LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6462_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____spec__1___boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__1; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_5974____closed__3; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273_(lean_object*); static lean_object* l_Lean_Lsp_instToJsonPartialResultParams___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6641____closed__3; -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__25___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__3; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__2___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__20; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2839____closed__1; static lean_object* l_Lean_Lsp_WorkspaceEdit_instAppend___closed__3; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__3; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____spec__1___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__8___boxed(lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__11; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__3; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6159____closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__12; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080____closed__2; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__16; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5651____closed__1; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2612____closed__3; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__10___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__8; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6346_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__6___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__20; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6905____boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instOrdPosition; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__4; +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6689____spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2128_(lean_object*); static lean_object* l_Lean_Lsp_instToJsonLocation___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1840_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____spec__3(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_WorkspaceEdit_ofTextEdit(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____spec__3(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__2; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5727_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__5; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__8; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__12; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7090_(uint8_t); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____spec__4___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__5; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__14; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__12; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__3; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__5; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____closed__2; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__11; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__9; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__4; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__11; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5542_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__2; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166____closed__3; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__14; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__8; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6905_(uint8_t); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__9; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____spec__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__4; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__10; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonTextDocumentEdit; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____closed__3; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonWorkDoneProgressParams; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____spec__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__11; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonRenameFile; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__21; -LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashPosition____x40_Lean_Data_Lsp_Basic___hyg_365_(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660_(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6531____rarg(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____spec__4(size_t, size_t, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__5; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__6; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__1; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____spec__2(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6689_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__15; static lean_object* l_Lean_Lsp_instBEqLocation___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__6; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__18; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____spec__1___boxed(lean_object*, lean_object*); uint64_t lean_string_hash(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__16; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__4; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5320____closed__2; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonRange; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__2; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351____closed__3; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__4; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__8; static lean_object* l_Lean_Lsp_instOrdPosition___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__3; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__5; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__12; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__8; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__10; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__3; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____spec__1___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__9; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonLocation; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2440_(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6641____closed__2; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742____closed__2; +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__25___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__15; LEAN_EXPORT lean_object* l_Lean_Lsp_instEmptyCollectionTextEditBatch; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__17; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__4; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__3; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__2; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__12; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__13; +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__12___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__13(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instToJsonCreateFile___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__4; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____lambda__1___boxed(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__12; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3630____closed__2; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__11; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6456____closed__1; +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5651____spec__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__1; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__18; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__14; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__9; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__2; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__11; uint8_t lean_string_dec_eq(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3214____spec__1___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__3; static lean_object* l_Lean_Lsp_instToStringTextDocumentPositionParams___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__15; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonApplyWorkspaceEditParams; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__10; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__7; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_instBEqLocation; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__4; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonLocation; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__9; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__13; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6531_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____closed__3; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__8; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__1; +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__6___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__7(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__6; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__3; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__8; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__1; +LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordLocation____x40_Lean_Data_Lsp_Basic___hyg_1053_(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943_(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__9; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__4; +static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4805____closed__1; lean_object* l_Lean_Json_getObjVal_x3f(lean_object*, lean_object*); lean_object* l_Lean_Json_getBool_x3f(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____spec__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4805____closed__2; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__14; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__2; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____spec__2___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instToJsonSnippetString___closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_DeleteFile_instToJsonOptions; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__4; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__2; -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5836____spec__1(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__5; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__14; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351____closed__5; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____spec__2(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__15; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5836_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____spec__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6641_(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853_(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211_(lean_object*); +LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqPosition____x40_Lean_Data_Lsp_Basic___hyg_41_(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6905____closed__1; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__3; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__9; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__8; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5089____closed__1; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3214____closed__1; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__6; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6456____closed__2; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____spec__1___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__10; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__10; static lean_object* l_Lean_Lsp_instFromJsonRange___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6874_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__13; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonTextDocumentIdentifier; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__8; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2797____closed__2; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__6; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__10; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonDocumentSelector(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__9; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__1; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__7; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonTextEditBatch(lean_object*); lean_object* l_Lean_Json_getStr_x3f(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__6; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2608_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonDocumentChange(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonPosition; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__1; -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__6(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__8(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__3; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6798____closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonCommand; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__2; -LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordLocation____x40_Lean_Data_Lsp_Basic___hyg_1238_(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__7; lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(size_t, size_t, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonDocumentChange(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____spec__1(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__20; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__27___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__4___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__5(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__9; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordRange____x40_Lean_Data_Lsp_Basic___hyg_715____boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__19___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__20(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__5; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__25___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__26(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__9; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonTextEdit; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__11; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____spec__1___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__15; static lean_object* l_Lean_Lsp_instFromJsonSnippetString___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6983_(lean_object*); uint8_t lean_string_dec_lt(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__3; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__13; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instToJsonCommand___closed__1; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__16; static lean_object* l_Lean_Lsp_instFromJsonMarkupKind___closed__6; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__10; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__6; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__1; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__6; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166____closed__1; static lean_object* l_Lean_Lsp_instFromJsonTextEdit___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3024_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__3; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__2; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__10; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__4; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__15___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__1; static lean_object* l_Lean_Lsp_instFromJsonTextDocumentPositionParams___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__17; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3024____closed__2; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__1; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__5; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2255____closed__1; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__16; LEAN_EXPORT lean_object* l_Lean_Lsp_instLTRange; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__12; +LEAN_EXPORT lean_object* l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__17(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__8; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__12; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__7; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3399____closed__1; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__15___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__10; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155____spec__1(lean_object*, lean_object*); +static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; static lean_object* l_Lean_Lsp_instToJsonTextDocumentPositionParams___closed__1; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____spec__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_WorkspaceEdit_instAppend___closed__2; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__8; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__14; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__5; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5651_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__4; static lean_object* l_Lean_Lsp_instHashableMarkupKind___closed__1; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____spec__1___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_5974____closed__2; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__8; LEAN_EXPORT uint8_t l_Lean_Lsp_instDecidableEqMarkupKind(uint8_t, uint8_t); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__1; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__1(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__10; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__3; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885_(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Lsp_instFromJsonTextEditBatch___spec__1___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__12; lean_object* l_Lean_RBNode_setBlack___rarg(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__5; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonChangeAnnotation; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_WorkspaceEdit_ofTextDocumentEdit(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2423____closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonWorkDoneProgressOptions; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__5; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__3; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonWorkDoneProgressParams; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__15; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__2___boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__2; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__10; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__11; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__14; +LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordRange____x40_Lean_Data_Lsp_Basic___hyg_715_(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__9; +LEAN_EXPORT lean_object* l_Lean_RBNode_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____spec__2(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__8; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2255_(lean_object*); static lean_object* l_Lean_Lsp_instFromJsonDocumentFilter___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225_(lean_object*); lean_object* l_Lean_Json_setObjVal_x21(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__7; static lean_object* l_Lean_Lsp_instToJsonTextDocumentIdentifier___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__16___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5089_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__16; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_WorkspaceEdit_instAppend___lambda__3(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6159_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__8; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__7; lean_object* l_Lean_Json_getObjValD(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_WorkspaceEdit_instEmptyCollection; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__11; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__7; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__3; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__18; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__3; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__2; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__3; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__3; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__7; +LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____lambda__1(lean_object*); +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__25(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5320_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__11; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____closed__3; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166____closed__2; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__13; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_WorkspaceEdit_instAppend(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__11; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__9; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351____closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4806_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__12(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instHashablePosition; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4990_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__7; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__9; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__2; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__19; LEAN_EXPORT lean_object* l_Lean_Lsp_instOrdRange; -LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqPosition____x40_Lean_Data_Lsp_Basic___hyg_226_(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__8; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____spec__1___boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__19; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__7; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__15; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqPosition____x40_Lean_Data_Lsp_Basic___hyg_226____boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__2; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__9; -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__4(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonCancelParams; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__8; +LEAN_EXPORT lean_object* l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__3(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__8; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6456____closed__3; +static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____spec__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonWorkspaceEdit; -LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_701_(lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__15(size_t, size_t, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__10; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__6; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__5; static lean_object* l_Lean_Lsp_instToJsonDocumentChange___closed__1; static lean_object* l_Lean_Lsp_instFromJsonWorkDoneProgressParams___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__1; static lean_object* l_Lean_Lsp_instToJsonWorkspaceEdit___closed__1; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____spec__2___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonSnippetString; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonTextDocumentPositionParams; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__1; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__3; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonWorkDoneProgressBegin; static lean_object* l_Lean_Lsp_instFromJsonTextDocumentItem___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__13; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__8; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4806____closed__1; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__4; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__17; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instCoeTextEditTextEditBatch(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__8; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____spec__1___boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__2; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__6; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____spec__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2313_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__17; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__7; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__13; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__5; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_MarkupKind_toCtorIdx___boxed(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_516____boxed(lean_object*); static lean_object* l_Lean_Lsp_instToJsonRange___closed__1; lean_object* lean_array_to_list(lean_object*); -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__12___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__9; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__14; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__10; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__5; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__6; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__11; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__10; +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__6(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__8; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____closed__2; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__12; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Lsp_instToJsonDocumentSelector___spec__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonMarkupKind(lean_object*); static lean_object* l_Lean_Lsp_instToJsonDocumentChange___closed__3; +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instToJsonMarkupKind___closed__2; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__4; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__7; LEAN_EXPORT lean_object* l_Lean_Lsp_instToStringPosition(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__15; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__8; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6346____rarg(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__11; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonCreateFile; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__19; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__12; -LEAN_EXPORT lean_object* l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__10; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__2; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonLocationLink; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonWorkspaceEdit; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__5; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__1; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____closed__2; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166_(lean_object*); static lean_object* l_Lean_Lsp_WorkspaceEdit_instEmptyCollection___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instAppendTextEditBatch; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__8; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__4; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5505____closed__3; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__6; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__7; +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__19___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__5; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_DeleteFile_instFromJsonOptions___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__8; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4990____closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__11; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6159____closed__4; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__6; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonDeleteFile; static lean_object* l_Lean_Lsp_instToJsonWorkDoneProgressParams___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__2; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__5; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____closed__1; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__11; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__6; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__11; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____spec__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__10; static lean_object* l_Lean_Lsp_instToJsonTextDocumentEdit___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__3; -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__25(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2440____closed__1; -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__6___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__1___boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__9; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6798_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__10___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__11(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressEnd____x40_Lean_Data_Lsp_Basic___hyg_6633_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__18; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826_(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__3; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__3; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__2; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__10; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Lsp_instFromJsonDocumentSelector___spec__1___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__2; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848_(lean_object*); static lean_object* l_Lean_Lsp_instHashableMarkupContent___closed__1; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__2(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__16(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3023____closed__1; lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1840____closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__8; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__10; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__14(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__15; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__16; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__3; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__8; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__10; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__8; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__7; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonPartialResultParams; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__9; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__18; -LEAN_EXPORT lean_object* l_Lean_Lsp_instInhabitedCancelParams; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__8; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__5; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_MarkupKind_noConfusion(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__4; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__3; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__5; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__5; LEAN_EXPORT uint8_t l_Lean_Lsp_instDecidableEqMarkupContent(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__10; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__17; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5505____closed__2; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__9; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__7; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__17; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__9; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashPosition____x40_Lean_Data_Lsp_Basic___hyg_180____boxed(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__5; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__2; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupKind____x40_Lean_Data_Lsp_Basic___hyg_5851____boxed(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordLocation____x40_Lean_Data_Lsp_Basic___hyg_1053____boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__8; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__11; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__2; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__2; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonDocumentFilter; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__7; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__6; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__6; +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2255____spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__15(size_t, size_t, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_MarkupKind_noConfusion___rarg___lambda__1(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____spec__2___boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__5; +LEAN_EXPORT lean_object* l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__2(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonPartialResultParams; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__2; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3208____closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__14(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__3; static lean_object* l_Lean_Lsp_CreateFile_instFromJsonOptions___closed__1; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonMarkupContent; -LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupKind____x40_Lean_Data_Lsp_Basic___hyg_6036_(uint8_t); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonVersionedTextDocumentIdentifier; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__6; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__14; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__18; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__4; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__8; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____closed__3; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__4; +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__4(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__4; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__4; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__7; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__16___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__3; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____spec__1___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2839____closed__2; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__2; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__16; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__2; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__7; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6277____boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonVersionedTextDocumentIdentifier; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__19; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__16; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__11; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__12; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6159____closed__2; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__7; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__2(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____spec__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__12; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonMarkupKind___boxed(lean_object*); -LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordPosition____x40_Lean_Data_Lsp_Basic___hyg_300_(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__9; static lean_object* l_Lean_Lsp_instOrdLocation___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__7; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__6; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__7; +static lean_object* l_Lean_Lsp_instInhabitedLocation___closed__2; +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____spec__2(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__9; LEAN_EXPORT lean_object* l_Lean_Lsp_instLEPosition; lean_object* l_Array_append___rarg(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166____closed__4; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instBEqPosition; +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_CreateFile_instToJsonOptions; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__5; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6159____closed__3; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690_(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__4___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__11; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__11; static lean_object* l_Lean_Lsp_instToJsonApplyWorkspaceEditParams___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4990____closed__2; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Lsp_instToJsonDocumentSelector___spec__1(size_t, size_t, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____spec__1___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__7; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2612____closed__2; static lean_object* l_Lean_Lsp_instToJsonDeleteFile___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__7; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__12; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____spec__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557____closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonTextEdit; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonTextDocumentRegistrationOptions; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__19; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____spec__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instFromJsonMarkupKind___closed__5; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__8; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__17; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__5; +LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6277_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instOrdLocation; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__10; -LEAN_EXPORT lean_object* l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__17(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__15; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5320____closed__1; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__4; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonSnippetString; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__5; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__1; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__6; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__2; static lean_object* l_Lean_Lsp_instFromJsonStaticRegistrationOptions___closed__1; -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__19___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__16; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____spec__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__21___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__22(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6689____closed__1; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__13; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Lsp_instToJsonTextEditBatch___spec__1(size_t, size_t, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2797____closed__3; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonMarkupKind(uint8_t); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__2; -static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____spec__1___closed__1; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__1___boxed(lean_object*, lean_object*); +LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_516_(lean_object*); LEAN_EXPORT uint8_t l_Lean_Lsp_WorkspaceEdit_instAppend___lambda__1(lean_object*, lean_object*); -LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_decEqMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6317_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonDeleteFile; LEAN_EXPORT lean_object* l_Lean_Lsp_instBEqRange; static lean_object* l_Lean_Lsp_instFromJsonTextDocumentIdentifier___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__20; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__10; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_decEqMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6317____boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__2; lean_object* l_Array_append___rarg___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_CreateFile_instToJsonOptions___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__7; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__11; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__2; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Lsp_instToJsonTextEditBatch___spec__1___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__19___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__20(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__16; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891_(lean_object*); static lean_object* l_Lean_Lsp_instToJsonDocumentChange___closed__6; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080____closed__1; -LEAN_EXPORT lean_object* l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__3(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__2(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__2; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__8; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__7; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__17; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__4; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonProgressParams___rarg(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406____closed__2; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__1; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(size_t, size_t, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__18; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonStaticRegistrationOptions; -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__27___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__28(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6874____spec__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__11; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406____closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3630_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__21(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__8; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__2; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__8; LEAN_EXPORT lean_object* l_Lean_Lsp_instLTPosition; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordPosition____x40_Lean_Data_Lsp_Basic___hyg_300____boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2612____closed__1; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____spec__1___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__9; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7090____boxed(lean_object*); -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__23___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__24(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__15; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__1; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__10; +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3214____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonStaticRegistrationOptions; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__4; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__10; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351____closed__6; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__7; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2608____closed__2; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__12; static lean_object* l_Lean_Lsp_instFromJsonLocation___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__5; -static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__6; static lean_object* l_Lean_Lsp_instFromJsonLocationLink___closed__1; static lean_object* l_Lean_Lsp_instInhabitedPosition___closed__1; -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__4___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__3; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__13; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__3; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____closed__2; static lean_object* l_Lean_Lsp_instFromJsonPosition___closed__1; -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__10___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__2; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__2; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____lambda__1(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__4; static lean_object* l_Lean_Lsp_instFromJsonTextDocumentRegistrationOptions___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__3; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__15; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__14; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__2; LEAN_EXPORT lean_object* l_Lean_Lsp_WorkspaceEdit_instAppend___lambda__2(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__9; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__1; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__21; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__7; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__13; static lean_object* l_Lean_Lsp_instInhabitedLocation___closed__1; static lean_object* l_Lean_Lsp_instFromJsonMarkupContent___closed__1; -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__19(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__12; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__7; static lean_object* l_Lean_Lsp_instToJsonTextDocumentRegistrationOptions___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__8; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__4; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3208_(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqLocation____x40_Lean_Data_Lsp_Basic___hyg_821____boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__4; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__12; static lean_object* l_Lean_Lsp_instFromJsonWorkDoneProgressOptions___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____closed__2; +LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_442_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_DeleteFile_instFromJsonOptions; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__4; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__13; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__6; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__6; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__3; LEAN_EXPORT lean_object* l_Lean_Lsp_instDecidableEqMarkupKind___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____spec__1___boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__6; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__12; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__4; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3630____closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__7; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2797_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__17; -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__23(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____spec__5(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqPosition____x40_Lean_Data_Lsp_Basic___hyg_41____boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_5974____closed__4; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressBegin____x40_Lean_Data_Lsp_Basic___hyg_6537_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__5; LEAN_EXPORT lean_object* l_Lean_Lsp_instInhabitedRange; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5274_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__5; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__9; -LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqLocation____x40_Lean_Data_Lsp_Basic___hyg_1006_(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__3; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6874____closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__14; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____spec__2(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__16; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__9; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__6; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__10; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__8; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__3; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__5; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4621____closed__1; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__5; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____spec__1___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__6; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__8; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__5; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____closed__1; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____spec__4___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____spec__1___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__8; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__1___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__6; static lean_object* l_Lean_Lsp_instFromJsonMarkupKind___closed__4; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__2; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonTextDocumentRegistrationOptions; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__17; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297_(lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instLERange; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__8; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6462____boxed(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__9; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__4; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____spec__1___boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__7; -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3399____spec__1(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__11; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__12; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__7; -static lean_object* l_Lean_Lsp_instBEqCancelParams___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__8; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__4; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__15; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__9; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2612_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__21___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__22(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__16; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__8; lean_object* l_Lean_Json_getNat_x3f(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6983____closed__1; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__2; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__3; static lean_object* l_Lean_Lsp_instToJsonDocumentChange___closed__7; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__9; static lean_object* l_Lean_Lsp_instFromJsonMarkupKind___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__6; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__4; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__1; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____spec__1(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__6; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__5; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__6; LEAN_EXPORT lean_object* l_Lean_Lsp_WorkspaceEdit_instAppend___lambda__1___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__6; lean_object* l_List_foldl___at_Array_appendList___spec__1___rarg(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3208____boxed(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895____closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_instHashableRange; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166____closed__5; static lean_object* l_Lean_Lsp_instAppendTextEditBatch___closed__1; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____spec__1(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__8; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351____closed__2; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__2; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5505_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__8___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__21(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__23(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instToJsonWorkDoneProgressEnd___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__6; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__3; -LEAN_EXPORT lean_object* l_Lean_RBNode_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____spec__6(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__17; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__12___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__4; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instBEqPosition___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2797____closed__1; -static lean_object* l_Lean_Lsp_instInhabitedCancelParams___closed__2; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__7; -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__3; static lean_object* l_Lean_Lsp_instToJsonRenameFile___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__4; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____spec__1___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__10(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__13; +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__8___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__9(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__6; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__4; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__7; +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____spec__5(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instFromJsonVersionedTextDocumentIdentifier___closed__1; lean_object* lean_array_mk(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_WorkspaceEdit_instAppend___lambda__2___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__6; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__13; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__8; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____closed__3; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__5; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__2; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__2; static lean_object* l_Lean_Lsp_instFromJsonWorkspaceEdit___closed__1; -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__8___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__9(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__10; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__10; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__7; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__11; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__17; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____spec__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_5974____closed__1; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__10; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__14; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4805_(lean_object*); size_t lean_usize_add(size_t, size_t); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__13; -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2440____spec__1(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__9; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__4; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____closed__3; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__14; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__4; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__3; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557____closed__2; LEAN_EXPORT lean_object* l_Lean_Lsp_WorkspaceEdit_instAppend___lambda__3___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonWorkDoneProgressEnd; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__7; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__19; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2839____boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_MarkupKind_ofNat___boxed(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__6; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonTextDocumentItem; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__9; static lean_object* l_Lean_Lsp_instFromJsonChangeAnnotation___closed__1; -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__23___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__5; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__7; static lean_object* l_Lean_Lsp_instToJsonTextEdit___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__12; LEAN_EXPORT lean_object* l_Lean_Lsp_instInhabitedPosition; lean_object* lean_array_uget(lean_object*, size_t); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__12; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1655____closed__1; size_t lean_array_size(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__16; LEAN_EXPORT lean_object* l_Lean_Lsp_MarkupKind_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(size_t, size_t, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__2; static lean_object* l_Lean_Lsp_instToJsonDocumentFilter___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__11; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__18; -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__6___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__7(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____spec__1(lean_object*, lean_object*); -uint8_t l___private_Lean_Data_JsonRpc_0__Lean_JsonRpc_beqRequestID____x40_Lean_Data_JsonRpc___hyg_36_(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2839_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__2; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155____closed__4; static lean_object* l_Lean_Lsp_instToStringPosition___closed__3; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__18; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____spec__1(lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__8; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__3; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__10; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__2; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__12; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_5974_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instDecidableEqMarkupContent___boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__11; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7090____closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__9; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__4; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__9; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2423_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__1; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__7; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__7; +LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordPosition____x40_Lean_Data_Lsp_Basic___hyg_115_(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__6; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__1; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__4; static lean_object* l_Lean_Lsp_instToJsonChangeAnnotation___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__10; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__5; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____spec__1___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__11; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__2; +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3721____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonCommand; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonTextDocumentItem; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__2; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155____closed__2; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155____closed__1; lean_object* lean_string_append(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__6; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__9; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__7; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__4; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__17; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__3; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2423____closed__2; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonLocationLink; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__16(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__5; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__9; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__9; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__3; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__12; static lean_object* l_Lean_Lsp_instHashableRange___closed__1; -lean_object* l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instToJsonWorkDoneProgressBegin___closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonRenameFile; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__2; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__2; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__5; -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3906____spec__1___boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__2; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__10; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__5; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__8; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__15; +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__10(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__14___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instFromJsonDeleteFile___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__6; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____closed__4; +LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_decEqMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6132_(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordPosition____x40_Lean_Data_Lsp_Basic___hyg_115____boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__10; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__10; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__11; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__16; lean_object* l_Lean_RBNode_fold___at_Lean_RBMap_mergeBy___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__8(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__4; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__5; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__8; static lean_object* l_Lean_Lsp_instFromJsonMarkupKind___closed__2; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__1___boxed(lean_object*, lean_object*); uint8_t lean_usize_dec_lt(size_t, size_t); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__6; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5836____closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonCreateFile; static lean_object* l_Lean_Lsp_instFromJsonTextDocumentEdit___closed__1; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____spec__1___boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742____closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6159____boxed(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__7; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__5; LEAN_EXPORT lean_object* l_Lean_Lsp_MarkupKind_noConfusion___rarg(uint8_t, uint8_t, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__2; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__10; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__6; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__3; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__1; static lean_object* l_Lean_Lsp_instToJsonDocumentChange___closed__2; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__2; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5274____closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__3; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__2; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__5; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_decEqMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6132____boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_RBNode_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____spec__6(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__14; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__9; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonTextDocumentIdentifier; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__10; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__1___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__4; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__9; lean_object* l_Lean_Json_pretty(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instFromJsonMarkupKind___closed__3; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__4; -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__1(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__13; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashPosition____x40_Lean_Data_Lsp_Basic___hyg_365____boxed(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__12; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__3; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____spec__1(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__7; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__2; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__4; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166____closed__6; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5320____closed__3; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__12; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__14; static lean_object* l_Lean_Lsp_instToJsonWorkDoneProgressReport___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__2; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__4; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__6; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__6; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3023____boxed(lean_object*); +LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupKind____x40_Lean_Data_Lsp_Basic___hyg_5851_(uint8_t); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__7; static lean_object* l_Lean_Lsp_instToJsonDocumentChange___closed__5; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordRange____x40_Lean_Data_Lsp_Basic___hyg_900____boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__12; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonApplyWorkspaceEditParams; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__14; -static lean_object* l_Lean_Lsp_instFromJsonCancelParams___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__4; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__10; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__1; static lean_object* l_Lean_Lsp_instToJsonWorkDoneProgressOptions___closed__1; lean_object* lean_array_uset(lean_object*, size_t, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqLocation____x40_Lean_Data_Lsp_Basic___hyg_1006____boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__3; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__14; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonMarkupKind___boxed(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____closed__3; lean_object* l___private_Init_Data_Repr_0__Nat_reprFast(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__12; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonRange; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__13; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__11; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__2; static lean_object* l_Lean_Lsp_instToStringPosition___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3024____closed__1; +LEAN_EXPORT lean_object* l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947_(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767_(lean_object*); static lean_object* l_Lean_Lsp_instToJsonPosition___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351____closed__4; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__7; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__13; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__16; static lean_object* l_Lean_Lsp_MarkupKind_noConfusion___rarg___closed__1; static lean_object* l_Lean_Lsp_instFromJsonRenameFile___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____spec__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__1; +LEAN_EXPORT lean_object* l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__18(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__13; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Lsp_instFromJsonDocumentSelector___spec__1(size_t, size_t, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__4; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__17; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4621_(lean_object*); static lean_object* l_Lean_Lsp_instToJsonLocationLink___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__6; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonTextEditBatch(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__4; -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__12(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2608____closed__1; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____spec__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqCancelParams____x40_Lean_Data_Lsp_Basic___hyg_28____boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__3; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_701____boxed(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__4; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__7; -LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqCancelParams____x40_Lean_Data_Lsp_Basic___hyg_28_(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressEnd____x40_Lean_Data_Lsp_Basic___hyg_6818_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__3; -static lean_object* _init_l_Lean_Lsp_instInhabitedCancelParams___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("", 0, 0); -return x_1; -} -} -static lean_object* _init_l_Lean_Lsp_instInhabitedCancelParams___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Lsp_instInhabitedCancelParams___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_Lsp_instInhabitedCancelParams() { -_start: -{ -lean_object* x_1; -x_1 = l_Lean_Lsp_instInhabitedCancelParams___closed__2; -return x_1; -} -} -LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqCancelParams____x40_Lean_Data_Lsp_Basic___hyg_28_(lean_object* x_1, lean_object* x_2) { -_start: -{ -uint8_t x_3; -x_3 = l___private_Lean_Data_JsonRpc_0__Lean_JsonRpc_beqRequestID____x40_Lean_Data_JsonRpc___hyg_36_(x_1, x_2); -return x_3; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqCancelParams____x40_Lean_Data_Lsp_Basic___hyg_28____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_Basic_0__Lean_Lsp_beqCancelParams____x40_Lean_Data_Lsp_Basic___hyg_28_(x_1, x_2); -lean_dec(x_2); -lean_dec(x_1); -x_4 = lean_box(x_3); -return x_4; -} -} -static lean_object* _init_l_Lean_Lsp_instBEqCancelParams___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqCancelParams____x40_Lean_Data_Lsp_Basic___hyg_28____boxed), 2, 0); -return x_1; -} -} -static lean_object* _init_l_Lean_Lsp_instBEqCancelParams() { -_start: -{ -lean_object* x_1; -x_1 = l_Lean_Lsp_instBEqCancelParams___closed__1; -return x_1; -} -} -LEAN_EXPORT lean_object* l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(lean_object* x_1, lean_object* x_2) { -_start: -{ -if (lean_obj_tag(x_1) == 0) -{ -lean_object* x_3; -x_3 = lean_array_to_list(x_2); -return x_3; -} -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 = l_List_foldl___at_Array_appendList___spec__1___rarg(x_2, x_4); -x_1 = x_5; -x_2 = x_6; -goto _start; -} -} -} -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = lean_box(0); -x_2 = lean_array_mk(x_1); -return x_2; -} -} -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("id", 2, 2); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__3() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__2; -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; -} -} -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____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_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__3; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set(x_3, 1, x_1); -return x_3; -} -} -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____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_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____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; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86_(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = lean_box(0); -switch (lean_obj_tag(x_1)) { -case 0: -{ -uint8_t x_3; -x_3 = !lean_is_exclusive(x_1); -if (x_3 == 0) -{ -lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; -lean_ctor_set_tag(x_1, 3); -x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____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_1); -x_6 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_6, 0, x_5); -lean_ctor_set(x_6, 1, x_2); -x_7 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_7, 0, x_6); -lean_ctor_set(x_7, 1, x_2); -x_8 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_9 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_7, x_8); -x_10 = l_Lean_Json_mkObj(x_9); -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; -x_11 = lean_ctor_get(x_1, 0); -lean_inc(x_11); -lean_dec(x_1); -x_12 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_12, 0, x_11); -x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__2; -x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_13); -lean_ctor_set(x_14, 1, x_12); -x_15 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_15, 0, x_14); -lean_ctor_set(x_15, 1, x_2); -x_16 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_16, 0, x_15); -lean_ctor_set(x_16, 1, x_2); -x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_18 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_16, x_17); -x_19 = l_Lean_Json_mkObj(x_18); -return x_19; -} -} -case 1: -{ -uint8_t x_20; -x_20 = !lean_is_exclusive(x_1); -if (x_20 == 0) -{ -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; -lean_ctor_set_tag(x_1, 2); -x_21 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__2; -x_22 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_22, 0, x_21); -lean_ctor_set(x_22, 1, x_1); -x_23 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_23, 0, x_22); -lean_ctor_set(x_23, 1, x_2); -x_24 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_24, 0, x_23); -lean_ctor_set(x_24, 1, x_2); -x_25 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_26 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_24, x_25); -x_27 = l_Lean_Json_mkObj(x_26); -return x_27; -} -else -{ -lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_28 = lean_ctor_get(x_1, 0); -lean_inc(x_28); -lean_dec(x_1); -x_29 = lean_alloc_ctor(2, 1, 0); -lean_ctor_set(x_29, 0, x_28); -x_30 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__2; -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 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_32, 0, x_31); -lean_ctor_set(x_32, 1, x_2); -x_33 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_33, 0, x_32); -lean_ctor_set(x_33, 1, x_2); -x_34 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_35 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_33, x_34); -x_36 = l_Lean_Json_mkObj(x_35); -return x_36; -} -} -default: -{ -lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_37 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__5; -x_38 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_39 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_37, x_38); -x_40 = l_Lean_Json_mkObj(x_39); -return x_40; -} -} -} -} -static lean_object* _init_l_Lean_Lsp_instToJsonCancelParams___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86_), 1, 0); -return x_1; -} -} -static lean_object* _init_l_Lean_Lsp_instToJsonCancelParams() { -_start: -{ -lean_object* x_1; -x_1 = l_Lean_Lsp_instToJsonCancelParams___closed__1; -return x_1; -} -} -LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____lambda__1(lean_object* x_1) { -_start: -{ -uint8_t x_2; -x_2 = 0; -return x_2; -} -} -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean", 4, 4); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lsp", 3, 3); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__3() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("CancelParams", 12, 12); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__4() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__1; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__2; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__3; -x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); -return x_4; -} -} -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____lambda__1___boxed), 1, 0); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__6() { -_start: -{ -lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__4; -x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; -x_4 = l_Lean_Name_toString(x_1, x_2, x_3); -return x_4; -} -} -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__7() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked(".", 1, 1); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__8() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__6; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__7; -x_3 = lean_string_append(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__9() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____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_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__10() { -_start: -{ -lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__9; -x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; -x_4 = l_Lean_Name_toString(x_1, x_2, x_3); -return x_4; -} -} -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__11() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__8; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__10; -x_3 = lean_string_append(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked(": ", 2, 2); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__13() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__11; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; -x_3 = lean_string_append(x_1, x_2); -return x_3; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124_(lean_object* x_1) { -_start: -{ -lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__2; -x_3 = l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__1(x_1, x_2); -if (lean_obj_tag(x_3) == 0) -{ -uint8_t x_4; -x_4 = !lean_is_exclusive(x_3); -if (x_4 == 0) -{ -lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__13; -x_7 = lean_string_append(x_6, x_5); -lean_dec(x_5); -lean_ctor_set(x_3, 0, x_7); -return x_3; -} -else -{ -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_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__13; -x_10 = lean_string_append(x_9, x_8); -lean_dec(x_8); -x_11 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_11, 0, x_10); -return x_11; -} -} -else -{ -uint8_t x_12; -x_12 = !lean_is_exclusive(x_3); -if (x_12 == 0) -{ -return x_3; -} -else -{ -lean_object* x_13; lean_object* x_14; -x_13 = lean_ctor_get(x_3, 0); -lean_inc(x_13); -lean_dec(x_3); -x_14 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_14, 0, x_13); -return x_14; -} -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____lambda__1___boxed(lean_object* x_1) { -_start: -{ -uint8_t x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____lambda__1(x_1); -lean_dec(x_1); -x_3 = lean_box(x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Lsp_instFromJsonCancelParams___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124_), 1, 0); -return x_1; -} -} -static lean_object* _init_l_Lean_Lsp_instFromJsonCancelParams() { -_start: -{ -lean_object* x_1; -x_1 = l_Lean_Lsp_instFromJsonCancelParams___closed__1; -return x_1; -} -} +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__9; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5542____closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3023_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__6; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____spec__4(size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__27(lean_object*, lean_object*, lean_object*); static lean_object* _init_l_Lean_Lsp_instInhabitedPosition___closed__1() { _start: { @@ -1384,7 +884,7 @@ x_1 = l_Lean_Lsp_instInhabitedPosition___closed__1; return x_1; } } -LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqPosition____x40_Lean_Data_Lsp_Basic___hyg_226_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqPosition____x40_Lean_Data_Lsp_Basic___hyg_41_(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; @@ -1407,11 +907,11 @@ return x_9; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqPosition____x40_Lean_Data_Lsp_Basic___hyg_226____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqPosition____x40_Lean_Data_Lsp_Basic___hyg_41____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_Basic_0__Lean_Lsp_beqPosition____x40_Lean_Data_Lsp_Basic___hyg_226_(x_1, x_2); +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqPosition____x40_Lean_Data_Lsp_Basic___hyg_41_(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); @@ -1422,7 +922,7 @@ static lean_object* _init_l_Lean_Lsp_instBEqPosition___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqPosition____x40_Lean_Data_Lsp_Basic___hyg_226____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqPosition____x40_Lean_Data_Lsp_Basic___hyg_41____boxed), 2, 0); return x_1; } } @@ -1434,7 +934,7 @@ x_1 = l_Lean_Lsp_instBEqPosition___closed__1; return x_1; } } -LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordPosition____x40_Lean_Data_Lsp_Basic___hyg_300_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordPosition____x40_Lean_Data_Lsp_Basic___hyg_115_(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; @@ -1489,11 +989,11 @@ return x_15; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordPosition____x40_Lean_Data_Lsp_Basic___hyg_300____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordPosition____x40_Lean_Data_Lsp_Basic___hyg_115____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_Basic_0__Lean_Lsp_ordPosition____x40_Lean_Data_Lsp_Basic___hyg_300_(x_1, x_2); +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordPosition____x40_Lean_Data_Lsp_Basic___hyg_115_(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); @@ -1504,7 +1004,7 @@ static lean_object* _init_l_Lean_Lsp_instOrdPosition___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordPosition____x40_Lean_Data_Lsp_Basic___hyg_300____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordPosition____x40_Lean_Data_Lsp_Basic___hyg_115____boxed), 2, 0); return x_1; } } @@ -1516,7 +1016,7 @@ x_1 = l_Lean_Lsp_instOrdPosition___closed__1; return x_1; } } -LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashPosition____x40_Lean_Data_Lsp_Basic___hyg_365_(lean_object* x_1) { +LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashPosition____x40_Lean_Data_Lsp_Basic___hyg_180_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; uint64_t x_4; uint64_t x_5; uint64_t x_6; uint64_t x_7; uint64_t x_8; @@ -1530,11 +1030,11 @@ x_8 = lean_uint64_mix_hash(x_6, x_7); return x_8; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashPosition____x40_Lean_Data_Lsp_Basic___hyg_365____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashPosition____x40_Lean_Data_Lsp_Basic___hyg_180____boxed(lean_object* x_1) { _start: { uint64_t x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashPosition____x40_Lean_Data_Lsp_Basic___hyg_365_(x_1); +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashPosition____x40_Lean_Data_Lsp_Basic___hyg_180_(x_1); lean_dec(x_1); x_3 = lean_box_uint64(x_2); return x_3; @@ -1544,7 +1044,7 @@ static lean_object* _init_l_Lean_Lsp_instHashablePosition___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashPosition____x40_Lean_Data_Lsp_Basic___hyg_365____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashPosition____x40_Lean_Data_Lsp_Basic___hyg_180____boxed), 1, 0); return x_1; } } @@ -1556,7 +1056,31 @@ x_1 = l_Lean_Lsp_instHashablePosition___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406____closed__1() { +LEAN_EXPORT lean_object* l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_3; +x_3 = lean_array_to_list(x_2); +return x_3; +} +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 = l_List_foldl___at_Array_appendList___spec__1___rarg(x_2, x_4); +x_1 = x_5; +x_2 = x_6; +goto _start; +} +} +} +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__1() { _start: { lean_object* x_1; @@ -1564,7 +1088,7 @@ x_1 = lean_mk_string_unchecked("line", 4, 4); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__2() { _start: { lean_object* x_1; @@ -1572,7 +1096,16 @@ x_1 = lean_mk_string_unchecked("character", 9, 9); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406_(lean_object* x_1) { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(0); +x_2 = lean_array_mk(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221_(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; @@ -1581,7 +1114,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_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406____closed__1; +x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____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); @@ -1595,7 +1128,7 @@ lean_dec(x_1); 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_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406____closed__2; +x_12 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____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); @@ -1608,8 +1141,8 @@ lean_ctor_set(x_15, 1, x_7); x_16 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_16, 0, x_8); lean_ctor_set(x_16, 1, x_15); -x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_18 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_16, x_17); +x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3; +x_18 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_16, x_17); x_19 = l_Lean_Json_mkObj(x_18); return x_19; } @@ -1618,7 +1151,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonPosition___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221_), 1, 0); return x_1; } } @@ -1630,7 +1163,7 @@ x_1 = l_Lean_Lsp_instToJsonPosition___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____spec__1(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; @@ -1639,7 +1172,31 @@ x_4 = l_Lean_Json_getNat_x3f(x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__1() { +LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____lambda__1(lean_object* x_1) { +_start: +{ +uint8_t x_2; +x_2 = 0; +return x_2; +} +} +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lsp", 3, 3); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__3() { _start: { lean_object* x_1; @@ -1647,127 +1204,151 @@ x_1 = lean_mk_string_unchecked("Position", 8, 8); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__1; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__2; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__1; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__2; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__3; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____lambda__1___boxed), 1, 0); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__6() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__2; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__4; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(".", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__6; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____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_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__10() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__5; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__9; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__6; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__8; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__10; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(": ", 2, 2); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__7; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__11; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____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_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406____closed__2; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____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_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__10() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__15() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__9; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__14; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__11() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__10; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__8; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__15; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__12() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__11; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__16; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -1777,7 +1358,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_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__8; +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__13; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -1789,7 +1370,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_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__8; +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__13; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -1803,8 +1384,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_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406____closed__2; -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____spec__1(x_1, x_13); +x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__2; +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -1814,7 +1395,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_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__12; +x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__17; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -1826,7 +1407,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_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__12; +x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__17; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -1865,20 +1446,30 @@ return x_28; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____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_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____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_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____spec__1(x_1, x_2); lean_dec(x_2); return x_3; } } +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____lambda__1___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____lambda__1(x_1); +lean_dec(x_1); +x_3 = lean_box(x_2); +return x_3; +} +} static lean_object* _init_l_Lean_Lsp_instFromJsonPosition___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273_), 1, 0); return x_1; } } @@ -1972,7 +1563,7 @@ x_1 = l_Lean_Lsp_instInhabitedRange___closed__1; return x_1; } } -LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_627_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_442_(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; @@ -1980,7 +1571,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 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqPosition____x40_Lean_Data_Lsp_Basic___hyg_226_(x_3, x_5); +x_7 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqPosition____x40_Lean_Data_Lsp_Basic___hyg_41_(x_3, x_5); if (x_7 == 0) { uint8_t x_8; @@ -1990,16 +1581,16 @@ return x_8; else { uint8_t x_9; -x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqPosition____x40_Lean_Data_Lsp_Basic___hyg_226_(x_4, x_6); +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqPosition____x40_Lean_Data_Lsp_Basic___hyg_41_(x_4, x_6); return x_9; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_627____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_442____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_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_627_(x_1, x_2); +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_442_(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); @@ -2010,7 +1601,7 @@ static lean_object* _init_l_Lean_Lsp_instBEqRange___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_627____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_442____boxed), 2, 0); return x_1; } } @@ -2022,25 +1613,25 @@ x_1 = l_Lean_Lsp_instBEqRange___closed__1; return x_1; } } -LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_701_(lean_object* x_1) { +LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_516_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; uint64_t x_4; uint64_t x_5; uint64_t x_6; uint64_t x_7; uint64_t x_8; x_2 = lean_ctor_get(x_1, 0); x_3 = lean_ctor_get(x_1, 1); x_4 = 0; -x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashPosition____x40_Lean_Data_Lsp_Basic___hyg_365_(x_2); +x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashPosition____x40_Lean_Data_Lsp_Basic___hyg_180_(x_2); x_6 = lean_uint64_mix_hash(x_4, x_5); -x_7 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashPosition____x40_Lean_Data_Lsp_Basic___hyg_365_(x_3); +x_7 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashPosition____x40_Lean_Data_Lsp_Basic___hyg_180_(x_3); x_8 = lean_uint64_mix_hash(x_6, x_7); return x_8; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_701____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_516____boxed(lean_object* x_1) { _start: { uint64_t x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_701_(x_1); +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_516_(x_1); lean_dec(x_1); x_3 = lean_box_uint64(x_2); return x_3; @@ -2050,7 +1641,7 @@ static lean_object* _init_l_Lean_Lsp_instHashableRange___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_701____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_516____boxed), 1, 0); return x_1; } } @@ -2062,7 +1653,7 @@ x_1 = l_Lean_Lsp_instHashableRange___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557____closed__1() { _start: { lean_object* x_1; @@ -2070,7 +1661,7 @@ x_1 = lean_mk_string_unchecked("start", 5, 5); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557____closed__2() { _start: { lean_object* x_1; @@ -2078,14 +1669,14 @@ x_1 = lean_mk_string_unchecked("end", 3, 3); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(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; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406_(x_2); -x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742____closed__1; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221_(x_2); +x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557____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); @@ -2096,8 +1687,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_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406_(x_8); -x_10 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742____closed__2; +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221_(x_8); +x_10 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557____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); @@ -2110,8 +1701,8 @@ lean_ctor_set(x_13, 1, x_6); x_14 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_14, 0, x_7); lean_ctor_set(x_14, 1, x_13); -x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_16 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_14, x_15); +x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3; +x_16 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_14, x_15); x_17 = l_Lean_Json_mkObj(x_16); return x_17; } @@ -2120,7 +1711,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonRange___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_), 1, 0); return x_1; } } @@ -2132,16 +1723,16 @@ x_1 = l_Lean_Lsp_instToJsonRange___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____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_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458_(x_3); +x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273_(x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__1() { _start: { lean_object* x_1; @@ -2149,127 +1740,127 @@ x_1 = lean_mk_string_unchecked("Range", 5, 5); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____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_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__1; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__2; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__1; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__2; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____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_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__2; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__2; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__3; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____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_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557____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_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__6() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__5; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__5; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__6; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__6; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__7; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__7; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742____closed__2; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557____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_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__10() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__10() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__9; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__9; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__11() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__10; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__10; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__12() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__11; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__11; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -2279,7 +1870,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_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__8; +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__8; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -2291,7 +1882,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_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__8; +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__8; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -2305,8 +1896,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_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742____closed__2; -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____spec__1(x_1, x_13); +x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557____closed__2; +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -2316,7 +1907,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_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__12; +x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__12; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -2328,7 +1919,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_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__12; +x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__12; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -2367,11 +1958,11 @@ return x_28; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____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_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____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_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____spec__1(x_1, x_2); lean_dec(x_2); return x_3; } @@ -2380,7 +1971,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonRange___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609_), 1, 0); return x_1; } } @@ -2392,7 +1983,7 @@ x_1 = l_Lean_Lsp_instFromJsonRange___closed__1; return x_1; } } -LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordRange____x40_Lean_Data_Lsp_Basic___hyg_900_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordRange____x40_Lean_Data_Lsp_Basic___hyg_715_(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; uint8_t x_8; lean_object* x_9; @@ -2400,8 +1991,8 @@ 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 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordPosition____x40_Lean_Data_Lsp_Basic___hyg_300_(x_3, x_5); -x_8 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordPosition____x40_Lean_Data_Lsp_Basic___hyg_300_(x_4, x_6); +x_7 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordPosition____x40_Lean_Data_Lsp_Basic___hyg_115_(x_3, x_5); +x_8 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordPosition____x40_Lean_Data_Lsp_Basic___hyg_115_(x_4, x_6); x_9 = lean_box(x_8); if (lean_obj_tag(x_9) == 1) { @@ -2436,11 +2027,11 @@ return x_7; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordRange____x40_Lean_Data_Lsp_Basic___hyg_900____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordRange____x40_Lean_Data_Lsp_Basic___hyg_715____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_Basic_0__Lean_Lsp_ordRange____x40_Lean_Data_Lsp_Basic___hyg_900_(x_1, x_2); +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordRange____x40_Lean_Data_Lsp_Basic___hyg_715_(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); @@ -2451,7 +2042,7 @@ static lean_object* _init_l_Lean_Lsp_instOrdRange___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordRange____x40_Lean_Data_Lsp_Basic___hyg_900____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordRange____x40_Lean_Data_Lsp_Basic___hyg_715____boxed), 2, 0); return x_1; } } @@ -2482,8 +2073,16 @@ return x_1; static lean_object* _init_l_Lean_Lsp_instInhabitedLocation___closed__1() { _start: { +lean_object* x_1; +x_1 = lean_mk_string_unchecked("", 0, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Lsp_instInhabitedLocation___closed__2() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Lsp_instInhabitedCancelParams___closed__1; +x_1 = l_Lean_Lsp_instInhabitedLocation___closed__1; x_2 = l_Lean_Lsp_instInhabitedRange___closed__1; x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -2495,11 +2094,11 @@ static lean_object* _init_l_Lean_Lsp_instInhabitedLocation() { _start: { lean_object* x_1; -x_1 = l_Lean_Lsp_instInhabitedLocation___closed__1; +x_1 = l_Lean_Lsp_instInhabitedLocation___closed__2; return x_1; } } -LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqLocation____x40_Lean_Data_Lsp_Basic___hyg_1006_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqLocation____x40_Lean_Data_Lsp_Basic___hyg_821_(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; @@ -2517,16 +2116,16 @@ return x_8; else { uint8_t x_9; -x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_627_(x_4, x_6); +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_442_(x_4, x_6); return x_9; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqLocation____x40_Lean_Data_Lsp_Basic___hyg_1006____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqLocation____x40_Lean_Data_Lsp_Basic___hyg_821____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_Basic_0__Lean_Lsp_beqLocation____x40_Lean_Data_Lsp_Basic___hyg_1006_(x_1, x_2); +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqLocation____x40_Lean_Data_Lsp_Basic___hyg_821_(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); @@ -2537,7 +2136,7 @@ static lean_object* _init_l_Lean_Lsp_instBEqLocation___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqLocation____x40_Lean_Data_Lsp_Basic___hyg_1006____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqLocation____x40_Lean_Data_Lsp_Basic___hyg_821____boxed), 2, 0); return x_1; } } @@ -2549,7 +2148,7 @@ x_1 = l_Lean_Lsp_instBEqLocation___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895____closed__1() { _start: { lean_object* x_1; @@ -2557,7 +2156,7 @@ x_1 = lean_mk_string_unchecked("uri", 3, 3); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895____closed__2() { _start: { lean_object* x_1; @@ -2565,7 +2164,7 @@ x_1 = lean_mk_string_unchecked("range", 5, 5); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895_(lean_object* x_1) { _start: { uint8_t x_2; @@ -2577,15 +2176,15 @@ x_3 = lean_ctor_get(x_1, 0); x_4 = lean_ctor_get(x_1, 1); x_5 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_5, 0, x_3); -x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080____closed__1; +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895____closed__1; lean_ctor_set(x_1, 1, x_5); lean_ctor_set(x_1, 0, x_6); x_7 = lean_box(0); x_8 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_8, 0, x_1); lean_ctor_set(x_8, 1, x_7); -x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_4); -x_10 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080____closed__2; +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_4); +x_10 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895____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); @@ -2598,8 +2197,8 @@ lean_ctor_set(x_13, 1, x_7); x_14 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_14, 0, x_8); lean_ctor_set(x_14, 1, x_13); -x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_16 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_14, x_15); +x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3; +x_16 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_14, x_15); x_17 = l_Lean_Json_mkObj(x_16); return x_17; } @@ -2613,7 +2212,7 @@ lean_inc(x_18); lean_dec(x_1); x_20 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_20, 0, x_18); -x_21 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080____closed__1; +x_21 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895____closed__1; x_22 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_22, 0, x_21); lean_ctor_set(x_22, 1, x_20); @@ -2621,8 +2220,8 @@ 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 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_19); -x_26 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080____closed__2; +x_25 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_19); +x_26 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895____closed__2; x_27 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_27, 0, x_26); lean_ctor_set(x_27, 1, x_25); @@ -2635,8 +2234,8 @@ lean_ctor_set(x_29, 1, x_23); x_30 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_30, 0, x_24); lean_ctor_set(x_30, 1, x_29); -x_31 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_32 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_30, x_31); +x_31 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3; +x_32 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_30, x_31); x_33 = l_Lean_Json_mkObj(x_32); return x_33; } @@ -2646,7 +2245,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonLocation___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895_), 1, 0); return x_1; } } @@ -2658,7 +2257,7 @@ x_1 = l_Lean_Lsp_instToJsonLocation___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__1(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; @@ -2667,16 +2266,16 @@ x_4 = l_Lean_Json_getStr_x3f(x_3); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__2(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__2(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_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794_(x_3); +x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609_(x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__1() { _start: { lean_object* x_1; @@ -2684,127 +2283,127 @@ x_1 = lean_mk_string_unchecked("Location", 8, 8); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____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_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__1; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__2; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__1; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__2; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____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_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__2; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__2; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__3; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____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_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895____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_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__6() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__5; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__5; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__6; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__6; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__7; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__7; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080____closed__2; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895____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_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__10() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__10() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__9; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__9; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__11() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__10; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__10; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__12() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__11; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__11; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -2814,7 +2413,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_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__8; +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__8; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -2826,7 +2425,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_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__8; +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__8; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -2840,8 +2439,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_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080____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_1132____spec__2(x_1, x_13); +x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895____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_947____spec__2(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -2851,7 +2450,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_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__12; +x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__12; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -2863,7 +2462,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_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__12; +x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__12; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -2902,20 +2501,20 @@ return x_28; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____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_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____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_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__1(x_1, x_2); lean_dec(x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____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_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____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_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__2(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__2(x_1, x_2); lean_dec(x_2); return x_3; } @@ -2924,7 +2523,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonLocation___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947_), 1, 0); return x_1; } } @@ -2936,7 +2535,7 @@ x_1 = l_Lean_Lsp_instFromJsonLocation___closed__1; return x_1; } } -LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordLocation____x40_Lean_Data_Lsp_Basic___hyg_1238_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordLocation____x40_Lean_Data_Lsp_Basic___hyg_1053_(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; @@ -2948,7 +2547,7 @@ x_7 = lean_string_dec_lt(x_3, x_5); if (x_7 == 0) { uint8_t x_8; uint8_t x_9; -x_8 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordRange____x40_Lean_Data_Lsp_Basic___hyg_900_(x_4, x_6); +x_8 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordRange____x40_Lean_Data_Lsp_Basic___hyg_715_(x_4, x_6); x_9 = lean_string_dec_eq(x_3, x_5); if (x_9 == 0) { @@ -2981,11 +2580,11 @@ return x_13; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordLocation____x40_Lean_Data_Lsp_Basic___hyg_1238____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordLocation____x40_Lean_Data_Lsp_Basic___hyg_1053____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_Basic_0__Lean_Lsp_ordLocation____x40_Lean_Data_Lsp_Basic___hyg_1238_(x_1, x_2); +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordLocation____x40_Lean_Data_Lsp_Basic___hyg_1053_(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); @@ -2996,7 +2595,7 @@ static lean_object* _init_l_Lean_Lsp_instOrdLocation___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordLocation____x40_Lean_Data_Lsp_Basic___hyg_1238____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordLocation____x40_Lean_Data_Lsp_Basic___hyg_1053____boxed), 2, 0); return x_1; } } @@ -3008,7 +2607,7 @@ x_1 = l_Lean_Lsp_instOrdLocation___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155____spec__1(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) @@ -3024,7 +2623,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_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_4); +x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(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); @@ -3036,7 +2635,7 @@ return x_8; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155____closed__1() { _start: { lean_object* x_1; @@ -3044,7 +2643,7 @@ x_1 = lean_mk_string_unchecked("originSelectionRange", 20, 20); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155____closed__2() { _start: { lean_object* x_1; @@ -3052,7 +2651,7 @@ x_1 = lean_mk_string_unchecked("targetUri", 9, 9); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155____closed__3() { _start: { lean_object* x_1; @@ -3060,7 +2659,7 @@ x_1 = lean_mk_string_unchecked("targetRange", 11, 11); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155____closed__4() { _start: { lean_object* x_1; @@ -3068,19 +2667,19 @@ x_1 = lean_mk_string_unchecked("targetSelectionRange", 20, 20); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155_(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; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____closed__1; -x_4 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____spec__1(x_3, x_2); +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155____closed__1; +x_4 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155____spec__1(x_3, x_2); x_5 = lean_ctor_get(x_1, 1); lean_inc(x_5); x_6 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_6, 0, x_5); -x_7 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____closed__2; +x_7 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155____closed__2; x_8 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_8, 0, x_7); lean_ctor_set(x_8, 1, x_6); @@ -3090,8 +2689,8 @@ lean_ctor_set(x_10, 0, x_8); lean_ctor_set(x_10, 1, x_9); x_11 = lean_ctor_get(x_1, 2); lean_inc(x_11); -x_12 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_11); -x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____closed__3; +x_12 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_11); +x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155____closed__3; x_14 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_14, 0, x_13); lean_ctor_set(x_14, 1, x_12); @@ -3101,8 +2700,8 @@ lean_ctor_set(x_15, 1, x_9); x_16 = lean_ctor_get(x_1, 3); lean_inc(x_16); lean_dec(x_1); -x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_16); -x_18 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____closed__4; +x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_16); +x_18 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155____closed__4; x_19 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_19, 0, x_18); lean_ctor_set(x_19, 1, x_17); @@ -3121,8 +2720,8 @@ lean_ctor_set(x_23, 1, x_22); x_24 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_24, 0, x_4); lean_ctor_set(x_24, 1, x_23); -x_25 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_26 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_24, x_25); +x_25 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3; +x_26 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_24, x_25); x_27 = l_Lean_Json_mkObj(x_26); return x_27; } @@ -3131,7 +2730,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonLocationLink___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155_), 1, 0); return x_1; } } @@ -3143,7 +2742,7 @@ x_1 = l_Lean_Lsp_instToJsonLocationLink___closed__1; return x_1; } } -static lean_object* _init_l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____spec__1___closed__1() { +static lean_object* _init_l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____spec__1___closed__1() { _start: { lean_object* x_1; lean_object* x_2; @@ -3153,7 +2752,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____spec__1(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; @@ -3162,13 +2761,13 @@ switch (lean_obj_tag(x_3)) { case 0: { lean_object* x_4; -x_4 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____spec__1___closed__1; +x_4 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____spec__1___closed__1; return x_4; } case 1: { lean_object* x_5; -x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794_(x_3); +x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609_(x_3); if (lean_obj_tag(x_5) == 0) { uint8_t x_6; @@ -3219,7 +2818,7 @@ return x_14; { lean_object* x_15; uint8_t x_16; lean_inc(x_3); -x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794_(x_3); +x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609_(x_3); x_16 = !lean_is_exclusive(x_3); if (x_16 == 0) { @@ -3323,7 +2922,7 @@ return x_31; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__1() { _start: { lean_object* x_1; @@ -3331,39 +2930,39 @@ x_1 = lean_mk_string_unchecked("LocationLink", 12, 12); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____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_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__1; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__2; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__1; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__2; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____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_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__2; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__2; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__3; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__5() { _start: { lean_object* x_1; @@ -3371,177 +2970,177 @@ x_1 = lean_mk_string_unchecked("originSelectionRange\?", 21, 21); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____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_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__5; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__5; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__7() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__6; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__6; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__8; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__8; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__10() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____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_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____closed__2; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155____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_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__11() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__11() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__10; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__10; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__12() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__11; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__11; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__13() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__12; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__12; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__14() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____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_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____closed__3; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155____closed__3; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__15() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__15() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__14; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__14; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__16() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__15; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__15; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__17() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__16; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__16; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__18() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____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_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155____closed__4; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__19() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__19() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__18; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__18; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__20() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__20() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__19; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__19; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__21() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__21() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__20; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__20; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -3551,7 +3150,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_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__9; +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__9; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -3563,7 +3162,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_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__9; +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__9; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -3577,9 +3176,9 @@ 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_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____closed__2; +x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155____closed__2; lean_inc(x_1); -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__1(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -3590,7 +3189,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_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__13; +x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__13; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -3602,7 +3201,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_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__13; +x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__13; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -3616,9 +3215,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_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____closed__3; +x_24 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155____closed__3; lean_inc(x_1); -x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__2(x_1, x_24); +x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__2(x_1, x_24); if (lean_obj_tag(x_25) == 0) { uint8_t x_26; @@ -3630,7 +3229,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_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__17; +x_28 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__17; x_29 = lean_string_append(x_28, x_27); lean_dec(x_27); lean_ctor_set(x_25, 0, x_29); @@ -3642,7 +3241,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_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__17; +x_31 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__17; x_32 = lean_string_append(x_31, x_30); lean_dec(x_30); x_33 = lean_alloc_ctor(0, 1, 0); @@ -3656,8 +3255,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_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____closed__4; -x_36 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__2(x_1, x_35); +x_35 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155____closed__4; +x_36 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__2(x_1, x_35); if (lean_obj_tag(x_36) == 0) { uint8_t x_37; @@ -3669,7 +3268,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_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__21; +x_39 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__21; x_40 = lean_string_append(x_39, x_38); lean_dec(x_38); lean_ctor_set(x_36, 0, x_40); @@ -3681,7 +3280,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_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__21; +x_42 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__21; x_43 = lean_string_append(x_42, x_41); lean_dec(x_41); x_44 = lean_alloc_ctor(0, 1, 0); @@ -3726,11 +3325,11 @@ return x_50; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____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_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____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_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____spec__1(x_1, x_2); lean_dec(x_2); return x_3; } @@ -3739,7 +3338,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonLocationLink___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225_), 1, 0); return x_1; } } @@ -3751,7 +3350,7 @@ x_1 = l_Lean_Lsp_instFromJsonLocationLink___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____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_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(size_t x_1, size_t x_2, lean_object* x_3) { _start: { uint8_t x_4; @@ -3775,7 +3374,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__1(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) @@ -3795,7 +3394,7 @@ lean_object* x_5; size_t x_6; size_t x_7; lean_object* x_8; lean_object* x_9; le x_5 = lean_ctor_get(x_2, 0); x_6 = lean_array_size(x_5); x_7 = 0; -x_8 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_6, x_7, x_5); +x_8 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_6, x_7, x_5); lean_ctor_set_tag(x_2, 4); lean_ctor_set(x_2, 0, x_8); x_9 = lean_alloc_ctor(0, 2, 0); @@ -3815,7 +3414,7 @@ lean_inc(x_12); lean_dec(x_2); x_13 = lean_array_size(x_12); x_14 = 0; -x_15 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_13, x_14, x_12); +x_15 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_13, x_14, x_12); x_16 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_16, 0, x_15); x_17 = lean_alloc_ctor(0, 2, 0); @@ -3830,7 +3429,7 @@ return x_19; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____closed__1() { _start: { lean_object* x_1; @@ -3838,7 +3437,7 @@ x_1 = lean_mk_string_unchecked("title", 5, 5); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____closed__2() { _start: { lean_object* x_1; @@ -3846,7 +3445,7 @@ x_1 = lean_mk_string_unchecked("command", 7, 7); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____closed__3() { _start: { lean_object* x_1; @@ -3854,7 +3453,7 @@ x_1 = lean_mk_string_unchecked("arguments", 9, 9); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442_(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; @@ -3867,7 +3466,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_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____closed__1; +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____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); @@ -3877,15 +3476,15 @@ lean_ctor_set(x_9, 0, x_7); lean_ctor_set(x_9, 1, x_8); x_10 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_10, 0, x_3); -x_11 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____closed__2; +x_11 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____closed__2; x_12 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_12, 0, x_11); lean_ctor_set(x_12, 1, x_10); 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_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____closed__3; -x_15 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__1(x_14, x_4); +x_14 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____closed__3; +x_15 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__1(x_14, x_4); x_16 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_16, 0, x_15); lean_ctor_set(x_16, 1, x_8); @@ -3895,13 +3494,13 @@ lean_ctor_set(x_17, 1, x_16); x_18 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_18, 0, x_9); lean_ctor_set(x_18, 1, x_17); -x_19 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_20 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_18, x_19); +x_19 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3; +x_20 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_18, x_19); x_21 = l_Lean_Json_mkObj(x_20); return x_21; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____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_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____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; @@ -3909,7 +3508,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_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_4, x_5, x_3); +x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_4, x_5, x_3); return x_6; } } @@ -3917,7 +3516,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonCommand___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442_), 1, 0); return x_1; } } @@ -3929,7 +3528,16 @@ x_1 = l_Lean_Lsp_instToJsonCommand___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(size_t x_1, size_t x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____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_Lean_Json_getStr_x3f(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(size_t x_1, size_t x_2, lean_object* x_3) { _start: { uint8_t x_4; @@ -3956,7 +3564,7 @@ goto _start; } } } -static lean_object* _init_l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1() { +static lean_object* _init_l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1() { _start: { lean_object* x_1; @@ -3964,7 +3572,7 @@ x_1 = lean_mk_string_unchecked("expected JSON array, got '", 26, 26); return x_1; } } -static lean_object* _init_l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2() { +static lean_object* _init_l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2() { _start: { lean_object* x_1; @@ -3972,7 +3580,7 @@ x_1 = lean_mk_string_unchecked("'", 1, 1); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; @@ -3981,7 +3589,7 @@ switch (lean_obj_tag(x_3)) { case 0: { lean_object* x_4; -x_4 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____spec__1___closed__1; +x_4 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____spec__1___closed__1; return x_4; } case 1: @@ -3989,10 +3597,10 @@ case 1: lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_5 = lean_unsigned_to_nat(80u); x_6 = l_Lean_Json_pretty(x_3, x_5); -x_7 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_7 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_8 = lean_string_append(x_7, x_6); lean_dec(x_6); -x_9 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_9 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_10 = lean_string_append(x_8, x_9); x_11 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_11, 0, x_10); @@ -4008,7 +3616,7 @@ lean_object* x_13; size_t x_14; size_t x_15; lean_object* x_16; uint8_t x_17; x_13 = lean_ctor_get(x_3, 0); x_14 = lean_array_size(x_13); x_15 = 0; -x_16 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_14, x_15, x_13); +x_16 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_14, x_15, x_13); x_17 = !lean_is_exclusive(x_16); if (x_17 == 0) { @@ -4040,7 +3648,7 @@ lean_inc(x_21); lean_dec(x_3); x_22 = lean_array_size(x_21); x_23 = 0; -x_24 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_22, x_23, x_21); +x_24 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_22, x_23, x_21); x_25 = lean_ctor_get(x_24, 0); lean_inc(x_25); if (lean_is_exclusive(x_24)) { @@ -4073,10 +3681,10 @@ if (x_31 == 0) lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; x_32 = lean_ctor_get(x_3, 0); lean_dec(x_32); -x_33 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_33 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_34 = lean_string_append(x_33, x_30); lean_dec(x_30); -x_35 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_35 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_36 = lean_string_append(x_34, x_35); lean_ctor_set_tag(x_3, 0); lean_ctor_set(x_3, 0, x_36); @@ -4086,10 +3694,10 @@ else { lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_dec(x_3); -x_37 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_37 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_38 = lean_string_append(x_37, x_30); lean_dec(x_30); -x_39 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_39 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_40 = lean_string_append(x_38, x_39); x_41 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_41, 0, x_40); @@ -4099,7 +3707,7 @@ return x_41; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__1() { _start: { lean_object* x_1; @@ -4107,121 +3715,121 @@ x_1 = lean_mk_string_unchecked("Command", 7, 7); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____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_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__1; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__2; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__1; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__2; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____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_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__2; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__2; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__3; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____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_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____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_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__6() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__5; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__5; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__6; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__6; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__7; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__7; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____closed__2; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____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_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__10() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__10() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__9; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__9; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__11() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__10; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__10; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__12() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__11; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__11; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__13() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__13() { _start: { lean_object* x_1; @@ -4229,54 +3837,54 @@ x_1 = lean_mk_string_unchecked("arguments\?", 10, 10); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__14() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____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_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__13; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__13; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__15() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__15() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__14; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__14; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__16() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__15; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__15; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__17() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__16; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__16; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -4286,7 +3894,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_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__8; +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__8; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -4298,7 +3906,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_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__8; +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__8; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -4312,9 +3920,9 @@ 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_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____closed__2; +x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____closed__2; lean_inc(x_1); -x_14 = l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -4325,7 +3933,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_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__12; +x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__12; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -4337,7 +3945,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_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__12; +x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__12; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -4351,8 +3959,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_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____closed__3; -x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1(x_1, x_24); +x_24 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____closed__3; +x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2(x_1, x_24); if (lean_obj_tag(x_25) == 0) { uint8_t x_26; @@ -4363,7 +3971,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_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__17; +x_28 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__17; x_29 = lean_string_append(x_28, x_27); lean_dec(x_27); lean_ctor_set(x_25, 0, x_29); @@ -4375,7 +3983,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_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__17; +x_31 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__17; x_32 = lean_string_append(x_31, x_30); lean_dec(x_30); x_33 = lean_alloc_ctor(0, 1, 0); @@ -4417,7 +4025,16 @@ return x_39; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2___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_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____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_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__1(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { size_t x_4; size_t x_5; lean_object* x_6; @@ -4425,15 +4042,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_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_4, x_5, x_3); +x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_4, x_5, x_3); return x_6; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____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_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____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_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2(x_1, x_2); lean_dec(x_2); return x_3; } @@ -4442,7 +4059,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonCommand___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498_), 1, 0); return x_1; } } @@ -4454,7 +4071,7 @@ x_1 = l_Lean_Lsp_instFromJsonCommand___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1840____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1655____closed__1() { _start: { lean_object* x_1; @@ -4462,13 +4079,13 @@ x_1 = lean_mk_string_unchecked("value", 5, 5); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1840_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1655_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1840____closed__1; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1655____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); @@ -4479,8 +4096,8 @@ 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_5); -x_8 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_9 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_7, x_8); +x_8 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3; +x_9 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_7, x_8); x_10 = l_Lean_Json_mkObj(x_9); return x_10; } @@ -4489,7 +4106,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonSnippetString___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1840_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1655_), 1, 0); return x_1; } } @@ -4501,7 +4118,7 @@ x_1 = l_Lean_Lsp_instToJsonSnippetString___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__1() { _start: { lean_object* x_1; @@ -4509,85 +4126,85 @@ x_1 = lean_mk_string_unchecked("SnippetString", 13, 13); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____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_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__1; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__2; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__1; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__2; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____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_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__2; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__2; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__3; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____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_Basic_0__Lean_Lsp_toJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1840____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1655____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_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__6() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__5; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__5; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__6; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__6; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__7; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__7; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1840____closed__1; -x_3 = l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(x_1, x_2); +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1655____closed__1; +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -4596,7 +4213,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_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__8; +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__8; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -4608,7 +4225,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_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__8; +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__8; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -4641,7 +4258,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonSnippetString___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693_), 1, 0); return x_1; } } @@ -4653,7 +4270,7 @@ x_1 = l_Lean_Lsp_instFromJsonSnippetString___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____spec__1(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) @@ -4669,7 +4286,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_Basic_0__Lean_Lsp_toJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1840_(x_4); +x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1655_(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); @@ -4681,7 +4298,7 @@ return x_8; } } } -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____spec__2(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) @@ -4728,7 +4345,7 @@ return x_12; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____closed__1() { _start: { lean_object* x_1; @@ -4736,7 +4353,7 @@ x_1 = lean_mk_string_unchecked("newText", 7, 7); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____closed__2() { _start: { lean_object* x_1; @@ -4744,7 +4361,7 @@ x_1 = lean_mk_string_unchecked("leanExtSnippet", 14, 14); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____closed__3() { _start: { lean_object* x_1; @@ -4752,14 +4369,14 @@ x_1 = lean_mk_string_unchecked("annotationId", 12, 12); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800_(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; 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_742_(x_2); -x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080____closed__2; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_2); +x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895____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); @@ -4771,7 +4388,7 @@ x_8 = lean_ctor_get(x_1, 1); lean_inc(x_8); x_9 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_9, 0, x_8); -x_10 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____closed__1; +x_10 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____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); @@ -4780,13 +4397,13 @@ lean_ctor_set(x_12, 0, x_11); lean_ctor_set(x_12, 1, x_6); x_13 = lean_ctor_get(x_1, 2); lean_inc(x_13); -x_14 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____closed__2; -x_15 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__1(x_14, x_13); +x_14 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____closed__2; +x_15 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____spec__1(x_14, x_13); x_16 = lean_ctor_get(x_1, 3); lean_inc(x_16); lean_dec(x_1); -x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____closed__3; -x_18 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(x_17, x_16); +x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____closed__3; +x_18 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____spec__2(x_17, x_16); x_19 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_19, 0, x_18); lean_ctor_set(x_19, 1, x_6); @@ -4799,8 +4416,8 @@ lean_ctor_set(x_21, 1, x_20); x_22 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_22, 0, x_7); lean_ctor_set(x_22, 1, x_21); -x_23 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_24 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_22, x_23); +x_23 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3; +x_24 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_22, x_23); x_25 = l_Lean_Json_mkObj(x_24); return x_25; } @@ -4809,7 +4426,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonTextEdit___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800_), 1, 0); return x_1; } } @@ -4821,7 +4438,7 @@ x_1 = l_Lean_Lsp_instToJsonTextEdit___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____spec__1(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; @@ -4830,13 +4447,13 @@ switch (lean_obj_tag(x_3)) { case 0: { lean_object* x_4; -x_4 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____spec__1___closed__1; +x_4 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____spec__1___closed__1; return x_4; } case 1: { lean_object* x_5; -x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878_(x_3); +x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693_(x_3); if (lean_obj_tag(x_5) == 0) { uint8_t x_6; @@ -4887,7 +4504,7 @@ return x_14; { lean_object* x_15; uint8_t x_16; lean_inc(x_3); -x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878_(x_3); +x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693_(x_3); x_16 = !lean_is_exclusive(x_3); if (x_16 == 0) { @@ -4991,7 +4608,7 @@ return x_31; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____spec__2(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____spec__2(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; @@ -5000,7 +4617,7 @@ switch (lean_obj_tag(x_3)) { case 0: { lean_object* x_4; -x_4 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____spec__1___closed__1; +x_4 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____spec__1___closed__1; return x_4; } case 1: @@ -5161,7 +4778,7 @@ return x_31; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__1() { _start: { lean_object* x_1; @@ -5169,100 +4786,100 @@ x_1 = lean_mk_string_unchecked("TextEdit", 8, 8); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____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_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__1; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__2; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__1; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__2; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____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_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__2; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__2; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__3; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__10; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__10; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__5; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__5; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____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_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____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_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__8() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__7; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__8; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__8; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__10() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__9; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__9; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__11() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__11() { _start: { lean_object* x_1; @@ -5270,48 +4887,48 @@ x_1 = lean_mk_string_unchecked("leanExtSnippet\?", 15, 15); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__12() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____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_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__11; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____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_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__13() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__13() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__12; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__14() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__13; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__15() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__14; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__14; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__16() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__16() { _start: { lean_object* x_1; @@ -5319,54 +4936,54 @@ x_1 = lean_mk_string_unchecked("annotationId\?", 13, 13); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__17() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____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_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__16; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____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_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__18() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__18() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__17; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__17; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__19() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__19() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__18; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__18; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__20() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__20() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__19; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__19; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080____closed__2; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895____closed__2; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__2(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__2(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -5376,7 +4993,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_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__6; +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__6; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -5388,7 +5005,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_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__6; +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__6; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -5402,9 +5019,9 @@ 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_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____closed__1; +x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____closed__1; lean_inc(x_1); -x_14 = l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -5415,7 +5032,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_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__10; +x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__10; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -5427,7 +5044,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_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__10; +x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__10; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -5441,9 +5058,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_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____closed__2; +x_24 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____closed__2; lean_inc(x_1); -x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____spec__1(x_1, x_24); +x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____spec__1(x_1, x_24); if (lean_obj_tag(x_25) == 0) { uint8_t x_26; @@ -5455,7 +5072,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_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__15; +x_28 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__15; x_29 = lean_string_append(x_28, x_27); lean_dec(x_27); lean_ctor_set(x_25, 0, x_29); @@ -5467,7 +5084,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_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__15; +x_31 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__15; x_32 = lean_string_append(x_31, x_30); lean_dec(x_30); x_33 = lean_alloc_ctor(0, 1, 0); @@ -5481,8 +5098,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_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____closed__3; -x_36 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____spec__2(x_1, x_35); +x_35 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____closed__3; +x_36 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____spec__2(x_1, x_35); if (lean_obj_tag(x_36) == 0) { uint8_t x_37; @@ -5494,7 +5111,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_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__20; +x_39 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__20; x_40 = lean_string_append(x_39, x_38); lean_dec(x_38); lean_ctor_set(x_36, 0, x_40); @@ -5506,7 +5123,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_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__20; +x_42 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__20; x_43 = lean_string_append(x_42, x_41); lean_dec(x_41); x_44 = lean_alloc_ctor(0, 1, 0); @@ -5551,20 +5168,20 @@ return x_50; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____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_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____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_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____spec__1(x_1, x_2); lean_dec(x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____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_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____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_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____spec__2(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____spec__2(x_1, x_2); lean_dec(x_2); return x_3; } @@ -5573,7 +5190,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonTextEdit___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860_), 1, 0); return x_1; } } @@ -5603,7 +5220,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_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045_(x_6); +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860_(x_6); if (lean_obj_tag(x_9) == 0) { uint8_t x_10; @@ -5649,10 +5266,10 @@ case 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; x_2 = lean_unsigned_to_nat(80u); x_3 = l_Lean_Json_pretty(x_1, x_2); -x_4 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_4 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_5 = lean_string_append(x_4, x_3); lean_dec(x_3); -x_6 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_6 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_7 = lean_string_append(x_5, x_6); x_8 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_8, 0, x_7); @@ -5663,10 +5280,10 @@ case 1: lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; x_9 = lean_unsigned_to_nat(80u); x_10 = l_Lean_Json_pretty(x_1, x_9); -x_11 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_11 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_12 = lean_string_append(x_11, x_10); lean_dec(x_10); -x_13 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_13 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_14 = lean_string_append(x_12, x_13); x_15 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_15, 0, x_14); @@ -5695,10 +5312,10 @@ if (x_22 == 0) lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; x_23 = lean_ctor_get(x_1, 0); lean_dec(x_23); -x_24 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_24 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_25 = lean_string_append(x_24, x_21); lean_dec(x_21); -x_26 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_26 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_27 = lean_string_append(x_25, x_26); lean_ctor_set_tag(x_1, 0); lean_ctor_set(x_1, 0, x_27); @@ -5708,10 +5325,10 @@ else { lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_dec(x_1); -x_28 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_28 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_29 = lean_string_append(x_28, x_21); lean_dec(x_21); -x_30 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_30 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_31 = lean_string_append(x_29, x_30); x_32 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_32, 0, x_31); @@ -5748,7 +5365,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_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985_(x_5); +x_8 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800_(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); @@ -5786,7 +5403,7 @@ static lean_object* _init_l_Lean_Lsp_instEmptyCollectionTextEditBatch() { _start: { lean_object* x_1; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3; return x_1; } } @@ -5818,13 +5435,13 @@ x_4 = lean_array_mk(x_3); return x_4; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2313_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2128_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080____closed__1; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895____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); @@ -5835,8 +5452,8 @@ 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_5); -x_8 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_9 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_7, x_8); +x_8 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3; +x_9 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_7, x_8); x_10 = l_Lean_Json_mkObj(x_9); return x_10; } @@ -5845,7 +5462,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonTextDocumentIdentifier___closed__ _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2313_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2128_), 1, 0); return x_1; } } @@ -5857,7 +5474,7 @@ x_1 = l_Lean_Lsp_instToJsonTextDocumentIdentifier___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166____closed__1() { _start: { lean_object* x_1; @@ -5865,64 +5482,64 @@ x_1 = lean_mk_string_unchecked("TextDocumentIdentifier", 22, 22); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166____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_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__1; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__2; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351____closed__1; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__2; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166____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_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351____closed__2; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166____closed__2; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351____closed__3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166____closed__3; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__6; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__6; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351____closed__5; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166____closed__5; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080____closed__1; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__1(x_1, x_2); +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895____closed__1; +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -5931,7 +5548,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_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351____closed__6; +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166____closed__6; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -5943,7 +5560,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_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351____closed__6; +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166____closed__6; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -5976,7 +5593,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonTextDocumentIdentifier___closed _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166_), 1, 0); return x_1; } } @@ -5988,7 +5605,7 @@ x_1 = l_Lean_Lsp_instFromJsonTextDocumentIdentifier___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2440____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2255____spec__1(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) @@ -6039,7 +5656,7 @@ return x_15; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2440____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2255____closed__1() { _start: { lean_object* x_1; @@ -6047,7 +5664,7 @@ x_1 = lean_mk_string_unchecked("version", 7, 7); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2440_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2255_(lean_object* x_1) { _start: { uint8_t x_2; @@ -6059,23 +5676,23 @@ x_3 = lean_ctor_get(x_1, 0); x_4 = lean_ctor_get(x_1, 1); x_5 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_5, 0, x_3); -x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080____closed__1; +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895____closed__1; lean_ctor_set(x_1, 1, x_5); lean_ctor_set(x_1, 0, x_6); x_7 = lean_box(0); x_8 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_8, 0, x_1); lean_ctor_set(x_8, 1, x_7); -x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2440____closed__1; -x_10 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2440____spec__1(x_9, x_4); +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2255____closed__1; +x_10 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2255____spec__1(x_9, x_4); x_11 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_11, 0, x_10); lean_ctor_set(x_11, 1, x_7); x_12 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_12, 0, x_8); lean_ctor_set(x_12, 1, x_11); -x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_14 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_12, x_13); +x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3; +x_14 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_12, x_13); x_15 = l_Lean_Json_mkObj(x_14); return x_15; } @@ -6089,7 +5706,7 @@ lean_inc(x_16); lean_dec(x_1); x_18 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_18, 0, x_16); -x_19 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080____closed__1; +x_19 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895____closed__1; x_20 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_20, 0, x_19); lean_ctor_set(x_20, 1, x_18); @@ -6097,16 +5714,16 @@ x_21 = lean_box(0); x_22 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_22, 0, x_20); lean_ctor_set(x_22, 1, x_21); -x_23 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2440____closed__1; -x_24 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2440____spec__1(x_23, x_17); +x_23 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2255____closed__1; +x_24 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2255____spec__1(x_23, x_17); x_25 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_25, 0, x_24); lean_ctor_set(x_25, 1, x_21); x_26 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_26, 0, x_22); lean_ctor_set(x_26, 1, x_25); -x_27 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_28 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_26, x_27); +x_27 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3; +x_28 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_26, x_27); x_29 = l_Lean_Json_mkObj(x_28); return x_29; } @@ -6116,7 +5733,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonVersionedTextDocumentIdentifier__ _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2440_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2255_), 1, 0); return x_1; } } @@ -6128,7 +5745,7 @@ x_1 = l_Lean_Lsp_instToJsonVersionedTextDocumentIdentifier___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____spec__1(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; @@ -6137,7 +5754,7 @@ switch (lean_obj_tag(x_3)) { case 0: { lean_object* x_4; -x_4 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____spec__1___closed__1; +x_4 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____spec__1___closed__1; return x_4; } case 1: @@ -6298,7 +5915,7 @@ return x_31; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__1() { _start: { lean_object* x_1; @@ -6306,59 +5923,59 @@ x_1 = lean_mk_string_unchecked("VersionedTextDocumentIdentifier", 31, 31); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____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_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__1; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__2; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__1; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__2; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____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_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__2; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__2; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__3; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__6; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__6; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__5; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__5; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__7() { _start: { lean_object* x_1; @@ -6366,54 +5983,54 @@ x_1 = lean_mk_string_unchecked("version\?", 8, 8); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____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_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__7; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__7; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__9() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__8; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__8; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__10() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__9; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__9; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__11() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__10; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__10; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -6423,7 +6040,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_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__6; +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__6; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -6435,7 +6052,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_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__6; +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__6; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -6449,8 +6066,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_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2440____closed__1; -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____spec__1(x_1, x_13); +x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2255____closed__1; +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -6460,7 +6077,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_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__11; +x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__11; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -6472,7 +6089,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_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__11; +x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__11; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -6511,11 +6128,11 @@ return x_28; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____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_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____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_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____spec__1(x_1, x_2); lean_dec(x_2); return x_3; } @@ -6524,7 +6141,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonVersionedTextDocumentIdentifier _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297_), 1, 0); return x_1; } } @@ -6536,7 +6153,7 @@ x_1 = l_Lean_Lsp_instFromJsonVersionedTextDocumentIdentifier___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2608____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2423____closed__1() { _start: { lean_object* x_1; @@ -6544,7 +6161,7 @@ x_1 = lean_mk_string_unchecked("textDocument", 12, 12); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2608____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2423____closed__2() { _start: { lean_object* x_1; @@ -6552,14 +6169,14 @@ x_1 = lean_mk_string_unchecked("edits", 5, 5); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2608_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2423_(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; size_t x_9; size_t x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2440_(x_2); -x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2608____closed__1; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2255_(x_2); +x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2423____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); @@ -6575,7 +6192,7 @@ x_10 = 0; x_11 = l_Array_mapMUnsafe_map___at_Lean_Lsp_instToJsonTextEditBatch___spec__1(x_9, x_10, x_8); x_12 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_12, 0, x_11); -x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2608____closed__2; +x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2423____closed__2; x_14 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_14, 0, x_13); lean_ctor_set(x_14, 1, x_12); @@ -6588,8 +6205,8 @@ lean_ctor_set(x_16, 1, x_6); x_17 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_17, 0, x_7); lean_ctor_set(x_17, 1, x_16); -x_18 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_19 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_17, x_18); +x_18 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3; +x_19 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_17, x_18); x_20 = l_Lean_Json_mkObj(x_19); return x_20; } @@ -6598,7 +6215,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonTextDocumentEdit___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2608_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2423_), 1, 0); return x_1; } } @@ -6610,16 +6227,16 @@ x_1 = l_Lean_Lsp_instToJsonTextDocumentEdit___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____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_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482_(x_3); +x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297_(x_3); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____spec__2(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____spec__2(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; @@ -6630,10 +6247,10 @@ case 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_unsigned_to_nat(80u); x_5 = l_Lean_Json_pretty(x_3, x_4); -x_6 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_6 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); -x_8 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_8 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_9 = lean_string_append(x_7, x_8); x_10 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_10, 0, x_9); @@ -6644,10 +6261,10 @@ case 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; x_11 = lean_unsigned_to_nat(80u); x_12 = l_Lean_Json_pretty(x_3, x_11); -x_13 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_13 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_14 = lean_string_append(x_13, x_12); lean_dec(x_12); -x_15 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_15 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_16 = lean_string_append(x_14, x_15); x_17 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_17, 0, x_16); @@ -6676,10 +6293,10 @@ if (x_24 == 0) lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; x_25 = lean_ctor_get(x_3, 0); lean_dec(x_25); -x_26 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_26 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_27 = lean_string_append(x_26, x_23); lean_dec(x_23); -x_28 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_28 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_29 = lean_string_append(x_27, x_28); lean_ctor_set_tag(x_3, 0); lean_ctor_set(x_3, 0, x_29); @@ -6689,10 +6306,10 @@ else { lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_dec(x_3); -x_30 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_30 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_31 = lean_string_append(x_30, x_23); lean_dec(x_23); -x_32 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_32 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_33 = lean_string_append(x_31, x_32); x_34 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_34, 0, x_33); @@ -6702,7 +6319,7 @@ return x_34; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__1() { _start: { lean_object* x_1; @@ -6710,127 +6327,127 @@ x_1 = lean_mk_string_unchecked("TextDocumentEdit", 16, 16); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____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_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__1; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__2; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__1; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__2; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____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_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__2; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__2; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__3; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____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_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2608____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2423____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_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__6() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__5; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__5; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__6; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__6; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__7; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__7; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2608____closed__2; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2423____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_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__10() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__10() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__9; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__9; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__11() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__10; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__10; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__12() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__11; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__11; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2608____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2423____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -6840,7 +6457,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_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__8; +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__8; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -6852,7 +6469,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_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__8; +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__8; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -6866,8 +6483,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_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2608____closed__2; -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____spec__2(x_1, x_13); +x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2423____closed__2; +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____spec__2(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -6877,7 +6494,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_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__12; +x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__12; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -6889,7 +6506,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_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__12; +x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__12; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -6928,20 +6545,20 @@ return x_28; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____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_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____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_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____spec__1(x_1, x_2); lean_dec(x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____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_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____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_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____spec__2(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____spec__2(x_1, x_2); lean_dec(x_2); return x_3; } @@ -6950,7 +6567,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonTextDocumentEdit___closed__1() _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475_), 1, 0); return x_1; } } @@ -6962,7 +6579,7 @@ x_1 = l_Lean_Lsp_instFromJsonTextDocumentEdit___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2797____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2612____closed__1() { _start: { lean_object* x_1; @@ -6970,7 +6587,7 @@ x_1 = lean_mk_string_unchecked("label", 5, 5); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2797____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2612____closed__2() { _start: { lean_object* x_1; @@ -6978,7 +6595,7 @@ x_1 = lean_mk_string_unchecked("needsConfirmation", 17, 17); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2797____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2612____closed__3() { _start: { lean_object* x_1; @@ -6986,7 +6603,7 @@ x_1 = lean_mk_string_unchecked("description", 11, 11); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2797_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2612_(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; @@ -6998,7 +6615,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_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2797____closed__1; +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2612____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); @@ -7008,15 +6625,15 @@ lean_ctor_set(x_9, 0, x_7); lean_ctor_set(x_9, 1, x_8); x_10 = lean_alloc_ctor(1, 0, 1); lean_ctor_set_uint8(x_10, 0, x_3); -x_11 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2797____closed__2; +x_11 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2612____closed__2; x_12 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_12, 0, x_11); lean_ctor_set(x_12, 1, x_10); 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_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2797____closed__3; -x_15 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(x_14, x_4); +x_14 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2612____closed__3; +x_15 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____spec__2(x_14, x_4); x_16 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_16, 0, x_15); lean_ctor_set(x_16, 1, x_8); @@ -7026,8 +6643,8 @@ lean_ctor_set(x_17, 1, x_16); x_18 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_18, 0, x_9); lean_ctor_set(x_18, 1, x_17); -x_19 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_20 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_18, x_19); +x_19 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3; +x_20 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_18, x_19); x_21 = l_Lean_Json_mkObj(x_20); return x_21; } @@ -7036,7 +6653,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonChangeAnnotation___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2797_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2612_), 1, 0); return x_1; } } @@ -7048,7 +6665,7 @@ x_1 = l_Lean_Lsp_instToJsonChangeAnnotation___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____spec__1(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; @@ -7058,7 +6675,7 @@ lean_dec(x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__1() { _start: { lean_object* x_1; @@ -7066,121 +6683,121 @@ x_1 = lean_mk_string_unchecked("ChangeAnnotation", 16, 16); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____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_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__1; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__2; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__1; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__2; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____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_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__2; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__2; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__3; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____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_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2797____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2612____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_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__6() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__5; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__5; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__6; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__6; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__7; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__7; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2797____closed__2; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2612____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_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__10() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__10() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__9; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__9; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__11() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__10; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__10; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__12() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__11; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__11; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__13() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__13() { _start: { lean_object* x_1; @@ -7188,54 +6805,54 @@ x_1 = lean_mk_string_unchecked("description\?", 12, 12); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__14() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____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_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__13; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__13; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__15() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__15() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__14; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__14; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__16() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__15; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__15; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__17() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__16; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__16; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2797____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2612____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -7245,7 +6862,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_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__8; +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__8; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -7257,7 +6874,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_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__8; +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__8; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -7271,9 +6888,9 @@ 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_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2797____closed__2; +x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2612____closed__2; lean_inc(x_1); -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____spec__1(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -7284,7 +6901,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_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__12; +x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__12; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -7296,7 +6913,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_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__12; +x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__12; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -7310,8 +6927,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_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2797____closed__3; -x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____spec__2(x_1, x_24); +x_24 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2612____closed__3; +x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____spec__2(x_1, x_24); if (lean_obj_tag(x_25) == 0) { uint8_t x_26; @@ -7322,7 +6939,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_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__17; +x_28 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__17; x_29 = lean_string_append(x_28, x_27); lean_dec(x_27); lean_ctor_set(x_25, 0, x_29); @@ -7334,7 +6951,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_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__17; +x_31 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__17; x_32 = lean_string_append(x_31, x_30); lean_dec(x_30); x_33 = lean_alloc_ctor(0, 1, 0); @@ -7380,11 +6997,11 @@ return x_41; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____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_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____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_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____spec__1(x_1, x_2); lean_dec(x_2); return x_3; } @@ -7393,7 +7010,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonChangeAnnotation___closed__1() _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668_), 1, 0); return x_1; } } @@ -7405,7 +7022,7 @@ x_1 = l_Lean_Lsp_instFromJsonChangeAnnotation___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3024____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2839____closed__1() { _start: { lean_object* x_1; @@ -7413,7 +7030,7 @@ x_1 = lean_mk_string_unchecked("overwrite", 9, 9); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3024____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2839____closed__2() { _start: { lean_object* x_1; @@ -7421,7 +7038,7 @@ x_1 = lean_mk_string_unchecked("ignoreIfExists", 14, 14); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3024_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2839_(lean_object* x_1) { _start: { 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; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; @@ -7429,7 +7046,7 @@ x_2 = lean_ctor_get_uint8(x_1, 0); x_3 = lean_ctor_get_uint8(x_1, 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_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3024____closed__1; +x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2839____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); @@ -7439,7 +7056,7 @@ lean_ctor_set(x_8, 0, x_6); lean_ctor_set(x_8, 1, x_7); x_9 = lean_alloc_ctor(1, 0, 1); lean_ctor_set_uint8(x_9, 0, x_3); -x_10 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3024____closed__2; +x_10 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2839____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); @@ -7452,17 +7069,17 @@ lean_ctor_set(x_13, 1, x_7); x_14 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_14, 0, x_8); lean_ctor_set(x_14, 1, x_13); -x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_16 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_14, x_15); +x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3; +x_16 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_14, x_15); x_17 = l_Lean_Json_mkObj(x_16); return x_17; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3024____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2839____boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3024_(x_1); +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2839_(x_1); lean_dec(x_1); return x_2; } @@ -7471,7 +7088,7 @@ static lean_object* _init_l_Lean_Lsp_CreateFile_instToJsonOptions___closed__1() _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3024____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2839____boxed), 1, 0); return x_1; } } @@ -7483,7 +7100,7 @@ x_1 = l_Lean_Lsp_CreateFile_instToJsonOptions___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__1() { _start: { lean_object* x_1; @@ -7491,7 +7108,7 @@ x_1 = lean_mk_string_unchecked("CreateFile", 10, 10); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__2() { _start: { lean_object* x_1; @@ -7499,128 +7116,128 @@ x_1 = lean_mk_string_unchecked("Options", 7, 7); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____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_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__1; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__2; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__1; -x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__2; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__2; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__1; +x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__2; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__4() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__3; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__3; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____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_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3024____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2839____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_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__7() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__6; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__6; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__5; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__5; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__8; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__8; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__10() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____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_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3024____closed__2; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2839____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_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__11() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__11() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__10; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__10; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__12() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__5; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__11; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__5; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__11; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__13() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__12; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__12; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3024____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2839____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -7630,7 +7247,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_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__9; +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__9; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -7642,7 +7259,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_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__9; +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__9; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -7656,8 +7273,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_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3024____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_2853____spec__1(x_1, x_13); +x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2839____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_2668____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -7667,7 +7284,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_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__13; +x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__13; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -7679,7 +7296,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_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__13; +x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__13; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -7730,7 +7347,7 @@ static lean_object* _init_l_Lean_Lsp_CreateFile_instFromJsonOptions___closed__1( _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891_), 1, 0); return x_1; } } @@ -7742,7 +7359,7 @@ x_1 = l_Lean_Lsp_CreateFile_instFromJsonOptions___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3208____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3023____closed__1() { _start: { lean_object* x_1; @@ -7750,7 +7367,7 @@ x_1 = lean_mk_string_unchecked("recursive", 9, 9); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3208____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3023____closed__2() { _start: { lean_object* x_1; @@ -7758,7 +7375,7 @@ x_1 = lean_mk_string_unchecked("ignoreIfNotExists", 17, 17); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3208_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3023_(lean_object* x_1) { _start: { 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; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; @@ -7766,7 +7383,7 @@ x_2 = lean_ctor_get_uint8(x_1, 0); x_3 = lean_ctor_get_uint8(x_1, 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_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3208____closed__1; +x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3023____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); @@ -7776,7 +7393,7 @@ lean_ctor_set(x_8, 0, x_6); lean_ctor_set(x_8, 1, x_7); x_9 = lean_alloc_ctor(1, 0, 1); lean_ctor_set_uint8(x_9, 0, x_3); -x_10 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3208____closed__2; +x_10 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3023____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); @@ -7789,17 +7406,17 @@ lean_ctor_set(x_13, 1, x_7); x_14 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_14, 0, x_8); lean_ctor_set(x_14, 1, x_13); -x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_16 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_14, x_15); +x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3; +x_16 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_14, x_15); x_17 = l_Lean_Json_mkObj(x_16); return x_17; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3208____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3023____boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3208_(x_1); +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3023_(x_1); lean_dec(x_1); return x_2; } @@ -7808,7 +7425,7 @@ static lean_object* _init_l_Lean_Lsp_DeleteFile_instToJsonOptions___closed__1() _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3208____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3023____boxed), 1, 0); return x_1; } } @@ -7820,7 +7437,7 @@ x_1 = l_Lean_Lsp_DeleteFile_instToJsonOptions___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__1() { _start: { lean_object* x_1; @@ -7828,128 +7445,128 @@ x_1 = lean_mk_string_unchecked("DeleteFile", 10, 10); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____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_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__1; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__2; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__1; -x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__2; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__2; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__1; +x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__2; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__2; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__2; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__3; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____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_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3208____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3023____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_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__6() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__5; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__5; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__6; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__6; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__7; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__7; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3208____closed__2; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3023____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_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__10() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__10() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__9; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__9; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__11() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__10; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__10; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__12() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__11; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__11; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3208____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3023____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -7959,7 +7576,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_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__8; +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__8; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -7971,7 +7588,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_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__8; +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__8; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -7985,8 +7602,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_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3208____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_2853____spec__1(x_1, x_13); +x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3023____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_2668____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -7996,7 +7613,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_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__12; +x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__12; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -8008,7 +7625,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_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__12; +x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__12; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -8059,7 +7676,7 @@ static lean_object* _init_l_Lean_Lsp_DeleteFile_instFromJsonOptions___closed__1( _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075_), 1, 0); return x_1; } } @@ -8071,7 +7688,7 @@ x_1 = l_Lean_Lsp_DeleteFile_instFromJsonOptions___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3399____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3214____spec__1(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) @@ -8085,7 +7702,7 @@ else { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; x_4 = lean_ctor_get(x_2, 0); -x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3024_(x_4); +x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2839_(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); @@ -8097,7 +7714,7 @@ return x_8; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3399____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3214____closed__1() { _start: { lean_object* x_1; @@ -8105,7 +7722,7 @@ x_1 = lean_mk_string_unchecked("options", 7, 7); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3399_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3214_(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; @@ -8118,7 +7735,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_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080____closed__1; +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895____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); @@ -8126,11 +7743,11 @@ x_8 = lean_box(0); 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_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3399____closed__1; -x_11 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3399____spec__1(x_10, x_3); +x_10 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3214____closed__1; +x_11 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3214____spec__1(x_10, x_3); lean_dec(x_3); -x_12 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____closed__3; -x_13 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(x_12, x_4); +x_12 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____closed__3; +x_13 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____spec__2(x_12, x_4); x_14 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_14, 0, x_13); lean_ctor_set(x_14, 1, x_8); @@ -8140,17 +7757,17 @@ lean_ctor_set(x_15, 1, x_14); x_16 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_16, 0, x_9); lean_ctor_set(x_16, 1, x_15); -x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_18 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_16, x_17); +x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3; +x_18 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_16, x_17); x_19 = l_Lean_Json_mkObj(x_18); return x_19; } } -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3399____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3214____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_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3399____spec__1(x_1, x_2); +x_3 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3214____spec__1(x_1, x_2); lean_dec(x_2); return x_3; } @@ -8159,7 +7776,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonCreateFile___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3399_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3214_), 1, 0); return x_1; } } @@ -8171,7 +7788,7 @@ x_1 = l_Lean_Lsp_instToJsonCreateFile___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____spec__1(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; @@ -8180,13 +7797,13 @@ switch (lean_obj_tag(x_3)) { case 0: { lean_object* x_4; -x_4 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____spec__1___closed__1; +x_4 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____spec__1___closed__1; return x_4; } case 1: { lean_object* x_5; -x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076_(x_3); +x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891_(x_3); if (lean_obj_tag(x_5) == 0) { uint8_t x_6; @@ -8237,7 +7854,7 @@ return x_14; { lean_object* x_15; uint8_t x_16; lean_inc(x_3); -x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076_(x_3); +x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891_(x_3); x_16 = !lean_is_exclusive(x_3); if (x_16 == 0) { @@ -8341,59 +7958,59 @@ return x_31; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__1; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__2; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__1; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__2; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____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_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__2() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__1; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__1; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__2; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__2; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__6; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__3; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__6; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__6() { _start: { lean_object* x_1; @@ -8401,74 +8018,74 @@ x_1 = lean_mk_string_unchecked("options\?", 8, 8); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____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_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__6; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____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_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__8() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__7; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__8; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__3; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__8; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__10() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__9; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__9; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__11() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__18; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__3; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__18; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__12() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__11; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__11; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -8478,7 +8095,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_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__5; +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__5; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -8490,7 +8107,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_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__5; +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__5; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -8504,9 +8121,9 @@ 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_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3399____closed__1; +x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3214____closed__1; lean_inc(x_1); -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____spec__1(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -8517,7 +8134,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_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__10; +x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__10; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -8529,7 +8146,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_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__10; +x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__10; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -8543,8 +8160,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_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____closed__3; -x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____spec__2(x_1, x_24); +x_24 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____closed__3; +x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____spec__2(x_1, x_24); if (lean_obj_tag(x_25) == 0) { uint8_t x_26; @@ -8555,7 +8172,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_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__12; +x_28 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__12; x_29 = lean_string_append(x_28, x_27); lean_dec(x_27); lean_ctor_set(x_25, 0, x_29); @@ -8567,7 +8184,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_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__12; +x_31 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__12; x_32 = lean_string_append(x_31, x_30); lean_dec(x_30); x_33 = lean_alloc_ctor(0, 1, 0); @@ -8609,11 +8226,11 @@ return x_39; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____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_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____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_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____spec__1(x_1, x_2); lean_dec(x_2); return x_3; } @@ -8622,7 +8239,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonCreateFile___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260_), 1, 0); return x_1; } } @@ -8634,7 +8251,7 @@ x_1 = l_Lean_Lsp_instFromJsonCreateFile___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3630____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__1() { _start: { lean_object* x_1; @@ -8642,7 +8259,7 @@ x_1 = lean_mk_string_unchecked("oldUri", 6, 6); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3630____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__2() { _start: { lean_object* x_1; @@ -8650,7 +8267,7 @@ x_1 = lean_mk_string_unchecked("newUri", 6, 6); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3630_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3445_(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; @@ -8665,7 +8282,7 @@ lean_inc(x_5); lean_dec(x_1); x_6 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_6, 0, x_2); -x_7 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3630____closed__1; +x_7 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__1; x_8 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_8, 0, x_7); lean_ctor_set(x_8, 1, x_6); @@ -8675,18 +8292,18 @@ lean_ctor_set(x_10, 0, x_8); lean_ctor_set(x_10, 1, x_9); x_11 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_11, 0, x_3); -x_12 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3630____closed__2; +x_12 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3445____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); x_14 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_14, 0, x_13); lean_ctor_set(x_14, 1, x_9); -x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3399____closed__1; -x_16 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3399____spec__1(x_15, x_4); +x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3214____closed__1; +x_16 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3214____spec__1(x_15, x_4); lean_dec(x_4); -x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____closed__3; -x_18 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(x_17, x_5); +x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____closed__3; +x_18 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____spec__2(x_17, x_5); x_19 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_19, 0, x_18); lean_ctor_set(x_19, 1, x_9); @@ -8699,8 +8316,8 @@ lean_ctor_set(x_21, 1, x_20); x_22 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_22, 0, x_10); lean_ctor_set(x_22, 1, x_21); -x_23 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_24 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_22, x_23); +x_23 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3; +x_24 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_22, x_23); x_25 = l_Lean_Json_mkObj(x_24); return x_25; } @@ -8709,7 +8326,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonRenameFile___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3630_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3445_), 1, 0); return x_1; } } @@ -8721,7 +8338,7 @@ x_1 = l_Lean_Lsp_instToJsonRenameFile___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__1() { _start: { lean_object* x_1; @@ -8729,167 +8346,167 @@ x_1 = lean_mk_string_unchecked("RenameFile", 10, 10); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____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_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__1; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__2; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__1; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__2; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____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_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__2; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__2; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__3; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____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_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3630____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3445____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_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__6() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__5; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__5; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__6; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__6; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__7; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__7; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3630____closed__2; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3445____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_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__10() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__10() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__9; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__9; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__11() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__10; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__10; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__12() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__11; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__11; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__13() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__8; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__8; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__14() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__13; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__13; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__15() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__18; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__18; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__16() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__15; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__15; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3630____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -8899,7 +8516,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_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__8; +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__8; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -8911,7 +8528,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_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__8; +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__8; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -8925,9 +8542,9 @@ 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_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3630____closed__2; +x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__2; lean_inc(x_1); -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__1(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -8938,7 +8555,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_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__12; +x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__12; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -8950,7 +8567,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_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__12; +x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__12; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -8964,9 +8581,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_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3399____closed__1; +x_24 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3214____closed__1; lean_inc(x_1); -x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____spec__1(x_1, x_24); +x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____spec__1(x_1, x_24); if (lean_obj_tag(x_25) == 0) { uint8_t x_26; @@ -8978,7 +8595,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_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__14; +x_28 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__14; x_29 = lean_string_append(x_28, x_27); lean_dec(x_27); lean_ctor_set(x_25, 0, x_29); @@ -8990,7 +8607,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_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__14; +x_31 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__14; x_32 = lean_string_append(x_31, x_30); lean_dec(x_30); x_33 = lean_alloc_ctor(0, 1, 0); @@ -9004,8 +8621,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_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____closed__3; -x_36 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____spec__2(x_1, x_35); +x_35 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____closed__3; +x_36 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____spec__2(x_1, x_35); if (lean_obj_tag(x_36) == 0) { uint8_t x_37; @@ -9017,7 +8634,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_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__16; +x_39 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__16; x_40 = lean_string_append(x_39, x_38); lean_dec(x_38); lean_ctor_set(x_36, 0, x_40); @@ -9029,7 +8646,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_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__16; +x_42 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__16; x_43 = lean_string_append(x_42, x_41); lean_dec(x_41); x_44 = lean_alloc_ctor(0, 1, 0); @@ -9078,7 +8695,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonRenameFile___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505_), 1, 0); return x_1; } } @@ -9090,7 +8707,7 @@ x_1 = l_Lean_Lsp_instFromJsonRenameFile___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3906____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3721____spec__1(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) @@ -9104,7 +8721,7 @@ else { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; x_4 = lean_ctor_get(x_2, 0); -x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3208_(x_4); +x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3023_(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); @@ -9116,7 +8733,7 @@ return x_8; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3906_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3721_(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; @@ -9129,7 +8746,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_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080____closed__1; +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895____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); @@ -9137,11 +8754,11 @@ x_8 = lean_box(0); 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_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3399____closed__1; -x_11 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3906____spec__1(x_10, x_3); +x_10 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3214____closed__1; +x_11 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3721____spec__1(x_10, x_3); lean_dec(x_3); -x_12 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____closed__3; -x_13 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(x_12, x_4); +x_12 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____closed__3; +x_13 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____spec__2(x_12, x_4); x_14 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_14, 0, x_13); lean_ctor_set(x_14, 1, x_8); @@ -9151,17 +8768,17 @@ lean_ctor_set(x_15, 1, x_14); x_16 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_16, 0, x_9); lean_ctor_set(x_16, 1, x_15); -x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_18 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_16, x_17); +x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3; +x_18 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_16, x_17); x_19 = l_Lean_Json_mkObj(x_18); return x_19; } } -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3906____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3721____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_Basic_0__Lean_Lsp_toJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3906____spec__1(x_1, x_2); +x_3 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3721____spec__1(x_1, x_2); lean_dec(x_2); return x_3; } @@ -9170,7 +8787,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonDeleteFile___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3906_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3721_), 1, 0); return x_1; } } @@ -9182,7 +8799,7 @@ x_1 = l_Lean_Lsp_instToJsonDeleteFile___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____spec__1(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; @@ -9191,13 +8808,13 @@ switch (lean_obj_tag(x_3)) { case 0: { lean_object* x_4; -x_4 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____spec__1___closed__1; +x_4 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____spec__1___closed__1; return x_4; } case 1: { lean_object* x_5; -x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260_(x_3); +x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075_(x_3); if (lean_obj_tag(x_5) == 0) { uint8_t x_6; @@ -9248,7 +8865,7 @@ return x_14; { lean_object* x_15; uint8_t x_16; lean_inc(x_3); -x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260_(x_3); +x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075_(x_3); x_16 = !lean_is_exclusive(x_3); if (x_16 == 0) { @@ -9352,105 +8969,105 @@ return x_31; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__1; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__2; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__1; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__2; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____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_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__2() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__1; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__1; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__2; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__2; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__6; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__3; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__6; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__8; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__3; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__8; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__6; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__6; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__18; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__3; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__18; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__8; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__8; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -9460,7 +9077,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_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__5; +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__5; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -9472,7 +9089,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_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__5; +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__5; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -9486,9 +9103,9 @@ 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_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3399____closed__1; +x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3214____closed__1; lean_inc(x_1); -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____spec__1(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -9499,7 +9116,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_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__7; +x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__7; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -9511,7 +9128,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_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__7; +x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__7; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -9525,8 +9142,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_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____closed__3; -x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____spec__2(x_1, x_24); +x_24 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____closed__3; +x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____spec__2(x_1, x_24); if (lean_obj_tag(x_25) == 0) { uint8_t x_26; @@ -9537,7 +9154,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_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__9; +x_28 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__9; x_29 = lean_string_append(x_28, x_27); lean_dec(x_27); lean_ctor_set(x_25, 0, x_29); @@ -9549,7 +9166,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_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__9; +x_31 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__9; x_32 = lean_string_append(x_31, x_30); lean_dec(x_30); x_33 = lean_alloc_ctor(0, 1, 0); @@ -9591,11 +9208,11 @@ return x_39; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____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_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____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_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____spec__1(x_1, x_2); lean_dec(x_2); return x_3; } @@ -9604,7 +9221,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonDeleteFile___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767_), 1, 0); return x_1; } } @@ -9688,7 +9305,7 @@ lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_obj x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); lean_dec(x_1); -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3399_(x_2); +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3214_(x_2); x_4 = l_Lean_Lsp_instToJsonDocumentChange___closed__3; x_5 = l_Lean_Lsp_instToJsonDocumentChange___closed__2; x_6 = l_Lean_Json_setObjVal_x21(x_3, x_4, x_5); @@ -9700,7 +9317,7 @@ lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_ob x_7 = lean_ctor_get(x_1, 0); lean_inc(x_7); lean_dec(x_1); -x_8 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3630_(x_7); +x_8 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3445_(x_7); x_9 = l_Lean_Lsp_instToJsonDocumentChange___closed__3; x_10 = l_Lean_Lsp_instToJsonDocumentChange___closed__5; x_11 = l_Lean_Json_setObjVal_x21(x_8, x_9, x_10); @@ -9712,7 +9329,7 @@ lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean x_12 = lean_ctor_get(x_1, 0); lean_inc(x_12); lean_dec(x_1); -x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3906_(x_12); +x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3721_(x_12); x_14 = l_Lean_Lsp_instToJsonDocumentChange___closed__3; x_15 = l_Lean_Lsp_instToJsonDocumentChange___closed__7; x_16 = l_Lean_Json_setObjVal_x21(x_13, x_14, x_15); @@ -9724,7 +9341,7 @@ lean_object* x_17; lean_object* x_18; x_17 = lean_ctor_get(x_1, 0); lean_inc(x_17); lean_dec(x_1); -x_18 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2608_(x_17); +x_18 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2423_(x_17); return x_18; } } @@ -9784,7 +9401,7 @@ else { lean_object* x_27; lean_inc(x_1); -x_27 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952_(x_1); +x_27 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767_(x_1); if (lean_obj_tag(x_27) == 0) { lean_object* x_28; @@ -9828,7 +9445,7 @@ else lean_object* x_33; lean_dec(x_19); lean_inc(x_1); -x_33 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690_(x_1); +x_33 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505_(x_1); if (lean_obj_tag(x_33) == 0) { lean_object* x_34; @@ -9872,7 +9489,7 @@ else lean_object* x_39; lean_dec(x_19); lean_inc(x_1); -x_39 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445_(x_1); +x_39 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260_(x_1); if (lean_obj_tag(x_39) == 0) { lean_object* x_40; @@ -9941,7 +9558,7 @@ else { lean_object* x_53; lean_inc(x_1); -x_53 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952_(x_1); +x_53 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767_(x_1); if (lean_obj_tag(x_53) == 0) { lean_object* x_54; @@ -9980,7 +9597,7 @@ else lean_object* x_59; lean_dec(x_45); lean_inc(x_1); -x_59 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690_(x_1); +x_59 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505_(x_1); if (lean_obj_tag(x_59) == 0) { lean_object* x_60; @@ -10019,7 +9636,7 @@ else lean_object* x_65; lean_dec(x_45); lean_inc(x_1); -x_65 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445_(x_1); +x_65 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260_(x_1); if (lean_obj_tag(x_65) == 0) { lean_object* x_66; @@ -10067,7 +9684,7 @@ goto block_13; { lean_object* x_3; lean_dec(x_2); -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660_(x_1); +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475_(x_1); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -10116,7 +9733,7 @@ return x_12; } } } -LEAN_EXPORT lean_object* l_Lean_RBNode_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____spec__2(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_RBNode_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____spec__2(lean_object* x_1) { _start: { if (lean_obj_tag(x_1) == 0) @@ -10135,13 +9752,13 @@ lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; size_t x x_4 = lean_ctor_get(x_1, 0); x_5 = lean_ctor_get(x_1, 2); x_6 = lean_ctor_get(x_1, 3); -x_7 = l_Lean_RBNode_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____spec__2(x_4); +x_7 = l_Lean_RBNode_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____spec__2(x_4); x_8 = lean_array_size(x_5); x_9 = 0; x_10 = l_Array_mapMUnsafe_map___at_Lean_Lsp_instToJsonTextEditBatch___spec__1(x_8, x_9, x_5); x_11 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_11, 0, x_10); -x_12 = l_Lean_RBNode_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____spec__2(x_6); +x_12 = l_Lean_RBNode_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____spec__2(x_6); lean_ctor_set(x_1, 3, x_12); lean_ctor_set(x_1, 2, x_11); lean_ctor_set(x_1, 0, x_7); @@ -10160,13 +9777,13 @@ lean_inc(x_16); lean_inc(x_15); lean_inc(x_14); lean_dec(x_1); -x_18 = l_Lean_RBNode_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____spec__2(x_14); +x_18 = l_Lean_RBNode_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____spec__2(x_14); x_19 = lean_array_size(x_16); x_20 = 0; x_21 = l_Array_mapMUnsafe_map___at_Lean_Lsp_instToJsonTextEditBatch___spec__1(x_19, x_20, x_16); x_22 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_22, 0, x_21); -x_23 = l_Lean_RBNode_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____spec__2(x_17); +x_23 = l_Lean_RBNode_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____spec__2(x_17); x_24 = lean_alloc_ctor(1, 4, 1); lean_ctor_set(x_24, 0, x_18); lean_ctor_set(x_24, 1, x_15); @@ -10178,7 +9795,7 @@ return x_24; } } } -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____spec__1(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) @@ -10196,7 +9813,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; x_5 = lean_ctor_get(x_2, 0); -x_6 = l_Lean_RBNode_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____spec__2(x_5); +x_6 = l_Lean_RBNode_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____spec__2(x_5); lean_ctor_set_tag(x_2, 5); lean_ctor_set(x_2, 0, x_6); x_7 = lean_alloc_ctor(0, 2, 0); @@ -10214,7 +9831,7 @@ lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean x_10 = lean_ctor_get(x_2, 0); lean_inc(x_10); lean_dec(x_2); -x_11 = l_Lean_RBNode_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____spec__2(x_10); +x_11 = l_Lean_RBNode_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____spec__2(x_10); x_12 = lean_alloc_ctor(5, 1, 0); lean_ctor_set(x_12, 0, x_11); x_13 = lean_alloc_ctor(0, 2, 0); @@ -10229,7 +9846,7 @@ return x_15; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____spec__4(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_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____spec__4(size_t x_1, size_t x_2, lean_object* x_3) { _start: { uint8_t x_4; @@ -10253,7 +9870,7 @@ lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean x_10 = lean_ctor_get(x_5, 0); lean_inc(x_10); lean_dec(x_5); -x_11 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3399_(x_10); +x_11 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3214_(x_10); x_12 = l_Lean_Lsp_instToJsonDocumentChange___closed__3; x_13 = l_Lean_Lsp_instToJsonDocumentChange___closed__2; x_14 = l_Lean_Json_setObjVal_x21(x_11, x_12, x_13); @@ -10268,7 +9885,7 @@ lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean x_17 = lean_ctor_get(x_5, 0); lean_inc(x_17); lean_dec(x_5); -x_18 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3630_(x_17); +x_18 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3445_(x_17); x_19 = l_Lean_Lsp_instToJsonDocumentChange___closed__3; x_20 = l_Lean_Lsp_instToJsonDocumentChange___closed__5; x_21 = l_Lean_Json_setObjVal_x21(x_18, x_19, x_20); @@ -10283,7 +9900,7 @@ lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean x_24 = lean_ctor_get(x_5, 0); lean_inc(x_24); lean_dec(x_5); -x_25 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3906_(x_24); +x_25 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3721_(x_24); x_26 = l_Lean_Lsp_instToJsonDocumentChange___closed__3; x_27 = l_Lean_Lsp_instToJsonDocumentChange___closed__7; x_28 = l_Lean_Json_setObjVal_x21(x_25, x_26, x_27); @@ -10298,7 +9915,7 @@ lean_object* x_31; lean_object* x_32; lean_object* x_33; x_31 = lean_ctor_get(x_5, 0); lean_inc(x_31); lean_dec(x_5); -x_32 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2608_(x_31); +x_32 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2423_(x_31); x_33 = lean_array_uset(x_7, x_2, x_32); x_2 = x_9; x_3 = x_33; @@ -10308,7 +9925,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____spec__3(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____spec__3(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) @@ -10328,7 +9945,7 @@ lean_object* x_5; size_t x_6; size_t x_7; lean_object* x_8; lean_object* x_9; le x_5 = lean_ctor_get(x_2, 0); x_6 = lean_array_size(x_5); x_7 = 0; -x_8 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____spec__4(x_6, x_7, x_5); +x_8 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____spec__4(x_6, x_7, x_5); lean_ctor_set_tag(x_2, 4); lean_ctor_set(x_2, 0, x_8); x_9 = lean_alloc_ctor(0, 2, 0); @@ -10348,7 +9965,7 @@ lean_inc(x_12); lean_dec(x_2); x_13 = lean_array_size(x_12); x_14 = 0; -x_15 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____spec__4(x_13, x_14, x_12); +x_15 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____spec__4(x_13, x_14, x_12); x_16 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_16, 0, x_15); x_17 = lean_alloc_ctor(0, 2, 0); @@ -10363,7 +9980,7 @@ return x_19; } } } -LEAN_EXPORT lean_object* l_Lean_RBNode_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____spec__6(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_RBNode_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____spec__6(lean_object* x_1) { _start: { if (lean_obj_tag(x_1) == 0) @@ -10382,9 +9999,9 @@ lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_obj x_4 = lean_ctor_get(x_1, 0); x_5 = lean_ctor_get(x_1, 2); x_6 = lean_ctor_get(x_1, 3); -x_7 = l_Lean_RBNode_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____spec__6(x_4); -x_8 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2797_(x_5); -x_9 = l_Lean_RBNode_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____spec__6(x_6); +x_7 = l_Lean_RBNode_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____spec__6(x_4); +x_8 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2612_(x_5); +x_9 = l_Lean_RBNode_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____spec__6(x_6); lean_ctor_set(x_1, 3, x_9); lean_ctor_set(x_1, 2, x_8); lean_ctor_set(x_1, 0, x_7); @@ -10403,9 +10020,9 @@ lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); lean_dec(x_1); -x_15 = l_Lean_RBNode_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____spec__6(x_11); -x_16 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2797_(x_13); -x_17 = l_Lean_RBNode_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____spec__6(x_14); +x_15 = l_Lean_RBNode_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____spec__6(x_11); +x_16 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2612_(x_13); +x_17 = l_Lean_RBNode_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____spec__6(x_14); x_18 = lean_alloc_ctor(1, 4, 1); lean_ctor_set(x_18, 0, x_15); lean_ctor_set(x_18, 1, x_12); @@ -10417,7 +10034,7 @@ return x_18; } } } -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____spec__5(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____spec__5(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) @@ -10435,7 +10052,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; x_5 = lean_ctor_get(x_2, 0); -x_6 = l_Lean_RBNode_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____spec__6(x_5); +x_6 = l_Lean_RBNode_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____spec__6(x_5); lean_ctor_set_tag(x_2, 5); lean_ctor_set(x_2, 0, x_6); x_7 = lean_alloc_ctor(0, 2, 0); @@ -10453,7 +10070,7 @@ lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean x_10 = lean_ctor_get(x_2, 0); lean_inc(x_10); lean_dec(x_2); -x_11 = l_Lean_RBNode_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____spec__6(x_10); +x_11 = l_Lean_RBNode_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____spec__6(x_10); x_12 = lean_alloc_ctor(5, 1, 0); lean_ctor_set(x_12, 0, x_11); x_13 = lean_alloc_ctor(0, 2, 0); @@ -10468,7 +10085,7 @@ return x_15; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____closed__1() { _start: { lean_object* x_1; @@ -10476,7 +10093,7 @@ x_1 = lean_mk_string_unchecked("changes", 7, 7); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____closed__2() { _start: { lean_object* x_1; @@ -10484,7 +10101,7 @@ x_1 = lean_mk_string_unchecked("documentChanges", 15, 15); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____closed__3() { _start: { lean_object* x_1; @@ -10492,23 +10109,23 @@ x_1 = lean_mk_string_unchecked("changeAnnotations", 17, 17); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223_(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; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____closed__1; -x_4 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____spec__1(x_3, x_2); +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____closed__1; +x_4 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____spec__1(x_3, x_2); x_5 = lean_ctor_get(x_1, 1); lean_inc(x_5); -x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____closed__2; -x_7 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____spec__3(x_6, x_5); +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____closed__2; +x_7 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____spec__3(x_6, x_5); x_8 = lean_ctor_get(x_1, 2); lean_inc(x_8); lean_dec(x_1); -x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____closed__3; -x_10 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____spec__5(x_9, x_8); +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____closed__3; +x_10 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____spec__5(x_9, x_8); x_11 = lean_box(0); x_12 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_12, 0, x_10); @@ -10519,13 +10136,13 @@ lean_ctor_set(x_13, 1, x_12); x_14 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_14, 0, x_4); lean_ctor_set(x_14, 1, x_13); -x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_16 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_14, x_15); +x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3; +x_16 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_14, x_15); x_17 = l_Lean_Json_mkObj(x_16); return x_17; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____spec__4___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_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { size_t x_4; size_t x_5; lean_object* x_6; @@ -10533,7 +10150,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_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____spec__4(x_4, x_5, x_3); +x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____spec__4(x_4, x_5, x_3); return x_6; } } @@ -10541,7 +10158,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonWorkspaceEdit___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223_), 1, 0); return x_1; } } @@ -10553,7 +10170,7 @@ x_1 = l_Lean_Lsp_instToJsonWorkspaceEdit___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_1) == 0) @@ -10592,7 +10209,7 @@ x_14 = lean_string_dec_eq(x_2, x_10); if (x_14 == 0) { lean_object* x_15; uint8_t x_16; -x_15 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__3(x_12, x_2, x_3); +x_15 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__3(x_12, x_2, x_3); x_16 = 0; lean_ctor_set(x_1, 3, x_15); lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_16); @@ -10613,7 +10230,7 @@ return x_1; else { lean_object* x_18; uint8_t x_19; -x_18 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__3(x_9, x_2, x_3); +x_18 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__3(x_9, x_2, x_3); x_19 = 0; lean_ctor_set(x_1, 0, x_18); lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_19); @@ -10640,7 +10257,7 @@ x_25 = lean_string_dec_eq(x_2, x_21); if (x_25 == 0) { lean_object* x_26; uint8_t x_27; lean_object* x_28; -x_26 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__3(x_23, x_2, x_3); +x_26 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__3(x_23, x_2, x_3); x_27 = 0; x_28 = lean_alloc_ctor(1, 4, 1); lean_ctor_set(x_28, 0, x_20); @@ -10668,7 +10285,7 @@ return x_30; else { lean_object* x_31; uint8_t x_32; lean_object* x_33; -x_31 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__3(x_20, x_2, x_3); +x_31 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__3(x_20, x_2, x_3); x_32 = 0; x_33 = lean_alloc_ctor(1, 4, 1); lean_ctor_set(x_33, 0, x_31); @@ -10699,7 +10316,7 @@ x_40 = lean_string_dec_eq(x_2, x_36); if (x_40 == 0) { lean_object* x_41; uint8_t x_42; -x_41 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__3(x_38, x_2, x_3); +x_41 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__3(x_38, x_2, x_3); x_42 = lean_ctor_get_uint8(x_41, sizeof(void*)*4); if (x_42 == 0) { @@ -11390,7 +11007,7 @@ return x_1; else { lean_object* x_193; uint8_t x_194; -x_193 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__3(x_35, x_2, x_3); +x_193 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__3(x_35, x_2, x_3); x_194 = lean_ctor_get_uint8(x_193, sizeof(void*)*4); if (x_194 == 0) { @@ -12096,7 +11713,7 @@ x_350 = lean_string_dec_eq(x_2, x_346); if (x_350 == 0) { lean_object* x_351; uint8_t x_352; -x_351 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__3(x_348, x_2, x_3); +x_351 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__3(x_348, x_2, x_3); x_352 = lean_ctor_get_uint8(x_351, sizeof(void*)*4); if (x_352 == 0) { @@ -12526,7 +12143,7 @@ return x_427; else { lean_object* x_428; uint8_t x_429; -x_428 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__3(x_345, x_2, x_3); +x_428 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__3(x_345, x_2, x_3); x_429 = lean_ctor_get_uint8(x_428, sizeof(void*)*4); if (x_429 == 0) { @@ -12943,7 +12560,7 @@ return x_502; } } } -LEAN_EXPORT lean_object* l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { uint8_t x_4; @@ -12951,19 +12568,19 @@ x_4 = l_Lean_RBNode_isRed___rarg(x_1); if (x_4 == 0) { lean_object* x_5; -x_5 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__3(x_1, x_2, x_3); +x_5 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__3(x_1, x_2, x_3); return x_5; } else { lean_object* x_6; lean_object* x_7; -x_6 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__3(x_1, x_2, x_3); +x_6 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__3(x_1, x_2, x_3); x_7 = l_Lean_RBNode_setBlack___rarg(x_6); return x_7; } } } -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_3) == 0) @@ -12985,7 +12602,7 @@ lean_inc(x_7); x_8 = lean_ctor_get(x_3, 3); lean_inc(x_8); lean_dec(x_3); -x_9 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__4(x_1, x_2, x_5); +x_9 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__4(x_1, x_2, x_5); if (lean_obj_tag(x_9) == 0) { uint8_t x_10; @@ -13050,7 +12667,7 @@ 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 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__2(x_13, x_6, x_21); +x_22 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__2(x_13, x_6, x_21); x_2 = x_22; x_3 = x_8; goto _start; @@ -13069,10 +12686,10 @@ x_25 = lean_ctor_get(x_9, 0); lean_dec(x_25); x_26 = lean_unsigned_to_nat(80u); x_27 = l_Lean_Json_pretty(x_7, x_26); -x_28 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_28 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_29 = lean_string_append(x_28, x_27); lean_dec(x_27); -x_30 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_30 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_31 = lean_string_append(x_29, x_30); lean_ctor_set_tag(x_9, 0); lean_ctor_set(x_9, 0, x_31); @@ -13084,10 +12701,10 @@ lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean lean_dec(x_9); x_32 = lean_unsigned_to_nat(80u); x_33 = l_Lean_Json_pretty(x_7, x_32); -x_34 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_34 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_35 = lean_string_append(x_34, x_33); lean_dec(x_33); -x_36 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_36 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_37 = lean_string_append(x_35, x_36); x_38 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_38, 0, x_37); @@ -13098,7 +12715,7 @@ return x_38; } } } -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__4___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__5(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__4___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__5(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) @@ -13120,7 +12737,7 @@ lean_inc(x_6); x_7 = lean_ctor_get(x_2, 3); lean_inc(x_7); lean_dec(x_2); -x_8 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__4___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__5(x_1, x_4); +x_8 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__4___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__5(x_1, x_4); if (lean_obj_tag(x_8) == 0) { uint8_t x_9; @@ -13185,7 +12802,7 @@ 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 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__2(x_12, x_5, x_20); +x_21 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__2(x_12, x_5, x_20); x_1 = x_21; x_2 = x_7; goto _start; @@ -13204,10 +12821,10 @@ x_24 = lean_ctor_get(x_8, 0); lean_dec(x_24); x_25 = lean_unsigned_to_nat(80u); x_26 = l_Lean_Json_pretty(x_6, x_25); -x_27 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_27 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_28 = lean_string_append(x_27, x_26); lean_dec(x_26); -x_29 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_29 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_30 = lean_string_append(x_28, x_29); lean_ctor_set_tag(x_8, 0); lean_ctor_set(x_8, 0, x_30); @@ -13219,10 +12836,10 @@ lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean lean_dec(x_8); x_31 = lean_unsigned_to_nat(80u); x_32 = l_Lean_Json_pretty(x_6, x_31); -x_33 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_33 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_34 = lean_string_append(x_33, x_32); lean_dec(x_32); -x_35 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_35 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_36 = lean_string_append(x_34, x_35); x_37 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_37, 0, x_36); @@ -13233,7 +12850,7 @@ return x_37; } } } -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_3) == 0) @@ -13255,7 +12872,7 @@ lean_inc(x_7); x_8 = lean_ctor_get(x_3, 3); lean_inc(x_8); lean_dec(x_3); -x_9 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__6(x_1, x_2, x_5); +x_9 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__6(x_1, x_2, x_5); if (lean_obj_tag(x_9) == 0) { uint8_t x_10; @@ -13320,7 +12937,7 @@ 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 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__2(x_13, x_6, x_21); +x_22 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__2(x_13, x_6, x_21); x_2 = x_22; x_3 = x_8; goto _start; @@ -13339,10 +12956,10 @@ x_25 = lean_ctor_get(x_9, 0); lean_dec(x_25); x_26 = lean_unsigned_to_nat(80u); x_27 = l_Lean_Json_pretty(x_7, x_26); -x_28 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_28 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_29 = lean_string_append(x_28, x_27); lean_dec(x_27); -x_30 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_30 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_31 = lean_string_append(x_29, x_30); lean_ctor_set_tag(x_9, 0); lean_ctor_set(x_9, 0, x_31); @@ -13354,10 +12971,10 @@ lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean lean_dec(x_9); x_32 = lean_unsigned_to_nat(80u); x_33 = l_Lean_Json_pretty(x_7, x_32); -x_34 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_34 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_35 = lean_string_append(x_34, x_33); lean_dec(x_33); -x_36 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_36 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_37 = lean_string_append(x_35, x_36); x_38 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_38, 0, x_37); @@ -13368,7 +12985,7 @@ return x_38; } } } -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__6___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__7(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__6___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__7(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) @@ -13390,7 +13007,7 @@ lean_inc(x_6); x_7 = lean_ctor_get(x_2, 3); lean_inc(x_7); lean_dec(x_2); -x_8 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__6___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__7(x_1, x_4); +x_8 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__6___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__7(x_1, x_4); if (lean_obj_tag(x_8) == 0) { uint8_t x_9; @@ -13455,7 +13072,7 @@ 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 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__2(x_12, x_5, x_20); +x_21 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__2(x_12, x_5, x_20); x_1 = x_21; x_2 = x_7; goto _start; @@ -13474,10 +13091,10 @@ x_24 = lean_ctor_get(x_8, 0); lean_dec(x_24); x_25 = lean_unsigned_to_nat(80u); x_26 = l_Lean_Json_pretty(x_6, x_25); -x_27 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_27 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_28 = lean_string_append(x_27, x_26); lean_dec(x_26); -x_29 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_29 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_30 = lean_string_append(x_28, x_29); lean_ctor_set_tag(x_8, 0); lean_ctor_set(x_8, 0, x_30); @@ -13489,10 +13106,10 @@ lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean lean_dec(x_8); x_31 = lean_unsigned_to_nat(80u); x_32 = l_Lean_Json_pretty(x_6, x_31); -x_33 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_33 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_34 = lean_string_append(x_33, x_32); lean_dec(x_32); -x_35 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_35 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_36 = lean_string_append(x_34, x_35); x_37 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_37, 0, x_36); @@ -13503,7 +13120,7 @@ return x_37; } } } -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__8(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__8(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_3) == 0) @@ -13525,7 +13142,7 @@ lean_inc(x_7); x_8 = lean_ctor_get(x_3, 3); lean_inc(x_8); lean_dec(x_3); -x_9 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__8(x_1, x_2, x_5); +x_9 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__8(x_1, x_2, x_5); if (lean_obj_tag(x_9) == 0) { uint8_t x_10; @@ -13590,7 +13207,7 @@ 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 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__2(x_13, x_6, x_21); +x_22 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__2(x_13, x_6, x_21); x_2 = x_22; x_3 = x_8; goto _start; @@ -13609,10 +13226,10 @@ x_25 = lean_ctor_get(x_9, 0); lean_dec(x_25); x_26 = lean_unsigned_to_nat(80u); x_27 = l_Lean_Json_pretty(x_7, x_26); -x_28 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_28 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_29 = lean_string_append(x_28, x_27); lean_dec(x_27); -x_30 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_30 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_31 = lean_string_append(x_29, x_30); lean_ctor_set_tag(x_9, 0); lean_ctor_set(x_9, 0, x_31); @@ -13624,10 +13241,10 @@ lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean lean_dec(x_9); x_32 = lean_unsigned_to_nat(80u); x_33 = l_Lean_Json_pretty(x_7, x_32); -x_34 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_34 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_35 = lean_string_append(x_34, x_33); lean_dec(x_33); -x_36 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_36 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_37 = lean_string_append(x_35, x_36); x_38 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_38, 0, x_37); @@ -13638,7 +13255,7 @@ return x_38; } } } -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__8___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__9(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__8___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__9(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) @@ -13660,7 +13277,7 @@ lean_inc(x_6); x_7 = lean_ctor_get(x_2, 3); lean_inc(x_7); lean_dec(x_2); -x_8 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__8___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__9(x_1, x_4); +x_8 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__8___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__9(x_1, x_4); if (lean_obj_tag(x_8) == 0) { uint8_t x_9; @@ -13725,7 +13342,7 @@ 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 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__2(x_12, x_5, x_20); +x_21 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__2(x_12, x_5, x_20); x_1 = x_21; x_2 = x_7; goto _start; @@ -13744,10 +13361,10 @@ x_24 = lean_ctor_get(x_8, 0); lean_dec(x_24); x_25 = lean_unsigned_to_nat(80u); x_26 = l_Lean_Json_pretty(x_6, x_25); -x_27 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_27 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_28 = lean_string_append(x_27, x_26); lean_dec(x_26); -x_29 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_29 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_30 = lean_string_append(x_28, x_29); lean_ctor_set_tag(x_8, 0); lean_ctor_set(x_8, 0, x_30); @@ -13759,10 +13376,10 @@ lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean lean_dec(x_8); x_31 = lean_unsigned_to_nat(80u); x_32 = l_Lean_Json_pretty(x_6, x_31); -x_33 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_33 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_34 = lean_string_append(x_33, x_32); lean_dec(x_32); -x_35 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_35 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_36 = lean_string_append(x_34, x_35); x_37 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_37, 0, x_36); @@ -13773,7 +13390,7 @@ return x_37; } } } -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__10(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__10(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_3) == 0) @@ -13795,7 +13412,7 @@ lean_inc(x_7); x_8 = lean_ctor_get(x_3, 3); lean_inc(x_8); lean_dec(x_3); -x_9 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__10(x_1, x_2, x_5); +x_9 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__10(x_1, x_2, x_5); if (lean_obj_tag(x_9) == 0) { uint8_t x_10; @@ -13860,7 +13477,7 @@ 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 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__2(x_13, x_6, x_21); +x_22 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__2(x_13, x_6, x_21); x_2 = x_22; x_3 = x_8; goto _start; @@ -13879,10 +13496,10 @@ x_25 = lean_ctor_get(x_9, 0); lean_dec(x_25); x_26 = lean_unsigned_to_nat(80u); x_27 = l_Lean_Json_pretty(x_7, x_26); -x_28 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_28 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_29 = lean_string_append(x_28, x_27); lean_dec(x_27); -x_30 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_30 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_31 = lean_string_append(x_29, x_30); lean_ctor_set_tag(x_9, 0); lean_ctor_set(x_9, 0, x_31); @@ -13894,10 +13511,10 @@ lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean lean_dec(x_9); x_32 = lean_unsigned_to_nat(80u); x_33 = l_Lean_Json_pretty(x_7, x_32); -x_34 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_34 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_35 = lean_string_append(x_34, x_33); lean_dec(x_33); -x_36 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_36 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_37 = lean_string_append(x_35, x_36); x_38 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_38, 0, x_37); @@ -13908,7 +13525,7 @@ return x_38; } } } -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__10___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__11(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__10___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__11(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) @@ -13930,7 +13547,7 @@ lean_inc(x_6); x_7 = lean_ctor_get(x_2, 3); lean_inc(x_7); lean_dec(x_2); -x_8 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__10___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__11(x_1, x_4); +x_8 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__10___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__11(x_1, x_4); if (lean_obj_tag(x_8) == 0) { uint8_t x_9; @@ -13995,7 +13612,7 @@ 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 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__2(x_12, x_5, x_20); +x_21 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__2(x_12, x_5, x_20); x_1 = x_21; x_2 = x_7; goto _start; @@ -14014,10 +13631,10 @@ x_24 = lean_ctor_get(x_8, 0); lean_dec(x_24); x_25 = lean_unsigned_to_nat(80u); x_26 = l_Lean_Json_pretty(x_6, x_25); -x_27 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_27 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_28 = lean_string_append(x_27, x_26); lean_dec(x_26); -x_29 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_29 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_30 = lean_string_append(x_28, x_29); lean_ctor_set_tag(x_8, 0); lean_ctor_set(x_8, 0, x_30); @@ -14029,10 +13646,10 @@ lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean lean_dec(x_8); x_31 = lean_unsigned_to_nat(80u); x_32 = l_Lean_Json_pretty(x_6, x_31); -x_33 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_33 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_34 = lean_string_append(x_33, x_32); lean_dec(x_32); -x_35 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_35 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_36 = lean_string_append(x_34, x_35); x_37 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_37, 0, x_36); @@ -14043,7 +13660,7 @@ return x_37; } } } -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__12(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__12(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_3) == 0) @@ -14065,7 +13682,7 @@ lean_inc(x_7); x_8 = lean_ctor_get(x_3, 3); lean_inc(x_8); lean_dec(x_3); -x_9 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__12(x_1, x_2, x_5); +x_9 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__12(x_1, x_2, x_5); if (lean_obj_tag(x_9) == 0) { uint8_t x_10; @@ -14130,7 +13747,7 @@ 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 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__2(x_13, x_6, x_21); +x_22 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__2(x_13, x_6, x_21); x_2 = x_22; x_3 = x_8; goto _start; @@ -14149,10 +13766,10 @@ x_25 = lean_ctor_get(x_9, 0); lean_dec(x_25); x_26 = lean_unsigned_to_nat(80u); x_27 = l_Lean_Json_pretty(x_7, x_26); -x_28 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_28 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_29 = lean_string_append(x_28, x_27); lean_dec(x_27); -x_30 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_30 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_31 = lean_string_append(x_29, x_30); lean_ctor_set_tag(x_9, 0); lean_ctor_set(x_9, 0, x_31); @@ -14164,10 +13781,10 @@ lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean lean_dec(x_9); x_32 = lean_unsigned_to_nat(80u); x_33 = l_Lean_Json_pretty(x_7, x_32); -x_34 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_34 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_35 = lean_string_append(x_34, x_33); lean_dec(x_33); -x_36 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_36 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_37 = lean_string_append(x_35, x_36); x_38 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_38, 0, x_37); @@ -14178,7 +13795,7 @@ return x_38; } } } -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__12___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__13(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__12___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__13(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) @@ -14200,7 +13817,7 @@ lean_inc(x_6); x_7 = lean_ctor_get(x_2, 3); lean_inc(x_7); lean_dec(x_2); -x_8 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__12___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__13(x_1, x_4); +x_8 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__12___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__13(x_1, x_4); if (lean_obj_tag(x_8) == 0) { uint8_t x_9; @@ -14265,7 +13882,7 @@ 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 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__2(x_12, x_5, x_20); +x_21 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__2(x_12, x_5, x_20); x_1 = x_21; x_2 = x_7; goto _start; @@ -14284,10 +13901,10 @@ x_24 = lean_ctor_get(x_8, 0); lean_dec(x_24); x_25 = lean_unsigned_to_nat(80u); x_26 = l_Lean_Json_pretty(x_6, x_25); -x_27 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_27 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_28 = lean_string_append(x_27, x_26); lean_dec(x_26); -x_29 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_29 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_30 = lean_string_append(x_28, x_29); lean_ctor_set_tag(x_8, 0); lean_ctor_set(x_8, 0, x_30); @@ -14299,10 +13916,10 @@ lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean lean_dec(x_8); x_31 = lean_unsigned_to_nat(80u); x_32 = l_Lean_Json_pretty(x_6, x_31); -x_33 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_33 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_34 = lean_string_append(x_33, x_32); lean_dec(x_32); -x_35 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_35 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_36 = lean_string_append(x_34, x_35); x_37 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_37, 0, x_36); @@ -14313,7 +13930,7 @@ return x_37; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__1(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; @@ -14322,7 +13939,7 @@ switch (lean_obj_tag(x_3)) { case 0: { lean_object* x_4; -x_4 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____spec__1___closed__1; +x_4 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____spec__1___closed__1; return x_4; } case 1: @@ -14355,7 +13972,7 @@ x_9 = lean_ctor_get(x_5, 0); lean_inc(x_9); lean_dec(x_5); x_10 = lean_box(0); -x_11 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__4___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__5(x_10, x_9); +x_11 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__4___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__5(x_10, x_9); if (lean_obj_tag(x_11) == 0) { uint8_t x_12; @@ -14441,7 +14058,7 @@ x_27 = lean_ctor_get(x_21, 0); lean_inc(x_27); lean_dec(x_21); x_28 = lean_box(0); -x_29 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__6___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__7(x_28, x_27); +x_29 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__6___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__7(x_28, x_27); if (lean_obj_tag(x_29) == 0) { uint8_t x_30; @@ -14520,7 +14137,7 @@ x_40 = lean_ctor_get(x_21, 0); lean_inc(x_40); lean_dec(x_21); x_41 = lean_box(0); -x_42 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__6___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__7(x_41, x_40); +x_42 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__6___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__7(x_41, x_40); if (lean_obj_tag(x_42) == 0) { lean_object* x_43; lean_object* x_44; lean_object* x_45; @@ -14604,7 +14221,7 @@ x_56 = lean_ctor_get(x_50, 0); lean_inc(x_56); lean_dec(x_50); x_57 = lean_box(0); -x_58 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__8___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__9(x_57, x_56); +x_58 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__8___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__9(x_57, x_56); if (lean_obj_tag(x_58) == 0) { uint8_t x_59; @@ -14683,7 +14300,7 @@ x_69 = lean_ctor_get(x_50, 0); lean_inc(x_69); lean_dec(x_50); x_70 = lean_box(0); -x_71 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__8___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__9(x_70, x_69); +x_71 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__8___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__9(x_70, x_69); if (lean_obj_tag(x_71) == 0) { lean_object* x_72; lean_object* x_73; lean_object* x_74; @@ -14767,7 +14384,7 @@ x_85 = lean_ctor_get(x_79, 0); lean_inc(x_85); lean_dec(x_79); x_86 = lean_box(0); -x_87 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__10___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__11(x_86, x_85); +x_87 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__10___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__11(x_86, x_85); if (lean_obj_tag(x_87) == 0) { uint8_t x_88; @@ -14846,7 +14463,7 @@ x_98 = lean_ctor_get(x_79, 0); lean_inc(x_98); lean_dec(x_79); x_99 = lean_box(0); -x_100 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__10___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__11(x_99, x_98); +x_100 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__10___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__11(x_99, x_98); if (lean_obj_tag(x_100) == 0) { lean_object* x_101; lean_object* x_102; lean_object* x_103; @@ -14930,7 +14547,7 @@ x_114 = lean_ctor_get(x_108, 0); lean_inc(x_114); lean_dec(x_108); x_115 = lean_box(0); -x_116 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__12___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__13(x_115, x_114); +x_116 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__12___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__13(x_115, x_114); if (lean_obj_tag(x_116) == 0) { uint8_t x_117; @@ -15009,7 +14626,7 @@ x_127 = lean_ctor_get(x_108, 0); lean_inc(x_127); lean_dec(x_108); x_128 = lean_box(0); -x_129 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__12___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__13(x_128, x_127); +x_129 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__12___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__13(x_128, x_127); if (lean_obj_tag(x_129) == 0) { lean_object* x_130; lean_object* x_131; lean_object* x_132; @@ -15058,7 +14675,7 @@ return x_136; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__15(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_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__15(size_t x_1, size_t x_2, lean_object* x_3) { _start: { uint8_t x_4; @@ -15123,7 +14740,7 @@ else { lean_object* x_33; lean_inc(x_6); -x_33 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952_(x_6); +x_33 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767_(x_6); if (lean_obj_tag(x_33) == 0) { lean_object* x_34; @@ -15171,7 +14788,7 @@ else lean_object* x_46; lean_dec(x_25); lean_inc(x_6); -x_46 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690_(x_6); +x_46 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505_(x_6); if (lean_obj_tag(x_46) == 0) { lean_object* x_47; @@ -15218,7 +14835,7 @@ else lean_object* x_59; lean_dec(x_25); lean_inc(x_6); -x_59 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445_(x_6); +x_59 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260_(x_6); if (lean_obj_tag(x_59) == 0) { lean_object* x_60; @@ -15274,7 +14891,7 @@ goto block_20; { lean_object* x_10; lean_dec(x_9); -x_10 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660_(x_6); +x_10 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475_(x_6); if (lean_obj_tag(x_10) == 0) { uint8_t x_11; @@ -15314,7 +14931,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__14(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__14(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; @@ -15323,7 +14940,7 @@ switch (lean_obj_tag(x_3)) { case 0: { lean_object* x_4; -x_4 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____spec__1___closed__1; +x_4 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____spec__1___closed__1; return x_4; } case 1: @@ -15331,10 +14948,10 @@ case 1: lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_5 = lean_unsigned_to_nat(80u); x_6 = l_Lean_Json_pretty(x_3, x_5); -x_7 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_7 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_8 = lean_string_append(x_7, x_6); lean_dec(x_6); -x_9 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_9 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_10 = lean_string_append(x_8, x_9); x_11 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_11, 0, x_10); @@ -15350,7 +14967,7 @@ lean_object* x_13; size_t x_14; size_t x_15; lean_object* x_16; x_13 = lean_ctor_get(x_3, 0); x_14 = lean_array_size(x_13); x_15 = 0; -x_16 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__15(x_14, x_15, x_13); +x_16 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__15(x_14, x_15, x_13); if (lean_obj_tag(x_16) == 0) { uint8_t x_17; @@ -15406,7 +15023,7 @@ lean_inc(x_24); lean_dec(x_3); x_25 = lean_array_size(x_24); x_26 = 0; -x_27 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__15(x_25, x_26, x_24); +x_27 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__15(x_25, x_26, x_24); if (lean_obj_tag(x_27) == 0) { lean_object* x_28; lean_object* x_29; lean_object* x_30; @@ -15463,10 +15080,10 @@ if (x_37 == 0) lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; x_38 = lean_ctor_get(x_3, 0); lean_dec(x_38); -x_39 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_39 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_40 = lean_string_append(x_39, x_36); lean_dec(x_36); -x_41 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_41 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_42 = lean_string_append(x_40, x_41); lean_ctor_set_tag(x_3, 0); lean_ctor_set(x_3, 0, x_42); @@ -15476,10 +15093,10 @@ else { lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_dec(x_3); -x_43 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_43 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_44 = lean_string_append(x_43, x_36); lean_dec(x_36); -x_45 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_45 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_46 = lean_string_append(x_44, x_45); x_47 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_47, 0, x_46); @@ -15489,7 +15106,7 @@ return x_47; } } } -LEAN_EXPORT lean_object* l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__18(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__18(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_1) == 0) @@ -15528,7 +15145,7 @@ x_14 = lean_string_dec_eq(x_2, x_10); if (x_14 == 0) { lean_object* x_15; uint8_t x_16; -x_15 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__18(x_12, x_2, x_3); +x_15 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__18(x_12, x_2, x_3); x_16 = 0; lean_ctor_set(x_1, 3, x_15); lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_16); @@ -15549,7 +15166,7 @@ return x_1; else { lean_object* x_18; uint8_t x_19; -x_18 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__18(x_9, x_2, x_3); +x_18 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__18(x_9, x_2, x_3); x_19 = 0; lean_ctor_set(x_1, 0, x_18); lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_19); @@ -15576,7 +15193,7 @@ x_25 = lean_string_dec_eq(x_2, x_21); if (x_25 == 0) { lean_object* x_26; uint8_t x_27; lean_object* x_28; -x_26 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__18(x_23, x_2, x_3); +x_26 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__18(x_23, x_2, x_3); x_27 = 0; x_28 = lean_alloc_ctor(1, 4, 1); lean_ctor_set(x_28, 0, x_20); @@ -15604,7 +15221,7 @@ return x_30; else { lean_object* x_31; uint8_t x_32; lean_object* x_33; -x_31 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__18(x_20, x_2, x_3); +x_31 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__18(x_20, x_2, x_3); x_32 = 0; x_33 = lean_alloc_ctor(1, 4, 1); lean_ctor_set(x_33, 0, x_31); @@ -15635,7 +15252,7 @@ x_40 = lean_string_dec_eq(x_2, x_36); if (x_40 == 0) { lean_object* x_41; uint8_t x_42; -x_41 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__18(x_38, x_2, x_3); +x_41 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__18(x_38, x_2, x_3); x_42 = lean_ctor_get_uint8(x_41, sizeof(void*)*4); if (x_42 == 0) { @@ -16326,7 +15943,7 @@ return x_1; else { lean_object* x_193; uint8_t x_194; -x_193 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__18(x_35, x_2, x_3); +x_193 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__18(x_35, x_2, x_3); x_194 = lean_ctor_get_uint8(x_193, sizeof(void*)*4); if (x_194 == 0) { @@ -17032,7 +16649,7 @@ x_350 = lean_string_dec_eq(x_2, x_346); if (x_350 == 0) { lean_object* x_351; uint8_t x_352; -x_351 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__18(x_348, x_2, x_3); +x_351 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__18(x_348, x_2, x_3); x_352 = lean_ctor_get_uint8(x_351, sizeof(void*)*4); if (x_352 == 0) { @@ -17462,7 +17079,7 @@ return x_427; else { lean_object* x_428; uint8_t x_429; -x_428 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__18(x_345, x_2, x_3); +x_428 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__18(x_345, x_2, x_3); x_429 = lean_ctor_get_uint8(x_428, sizeof(void*)*4); if (x_429 == 0) { @@ -17879,7 +17496,7 @@ return x_502; } } } -LEAN_EXPORT lean_object* l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__17(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__17(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { uint8_t x_4; @@ -17887,19 +17504,19 @@ x_4 = l_Lean_RBNode_isRed___rarg(x_1); if (x_4 == 0) { lean_object* x_5; -x_5 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__18(x_1, x_2, x_3); +x_5 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__18(x_1, x_2, x_3); return x_5; } else { lean_object* x_6; lean_object* x_7; -x_6 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__18(x_1, x_2, x_3); +x_6 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__18(x_1, x_2, x_3); x_7 = l_Lean_RBNode_setBlack___rarg(x_6); return x_7; } } } -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__19(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__19(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_3) == 0) @@ -17921,7 +17538,7 @@ lean_inc(x_7); x_8 = lean_ctor_get(x_3, 3); lean_inc(x_8); lean_dec(x_3); -x_9 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__19(x_1, x_2, x_5); +x_9 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__19(x_1, x_2, x_5); if (lean_obj_tag(x_9) == 0) { uint8_t x_10; @@ -17950,7 +17567,7 @@ lean_object* x_13; lean_object* x_14; x_13 = lean_ctor_get(x_9, 0); lean_inc(x_13); lean_dec(x_9); -x_14 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853_(x_7); +x_14 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668_(x_7); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -17979,7 +17596,7 @@ lean_object* x_18; lean_object* x_19; x_18 = lean_ctor_get(x_14, 0); lean_inc(x_18); lean_dec(x_14); -x_19 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__17(x_13, x_6, x_18); +x_19 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__17(x_13, x_6, x_18); x_2 = x_19; x_3 = x_8; goto _start; @@ -17988,7 +17605,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__19___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__20(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__19___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__20(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) @@ -18010,7 +17627,7 @@ lean_inc(x_6); x_7 = lean_ctor_get(x_2, 3); lean_inc(x_7); lean_dec(x_2); -x_8 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__19___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__20(x_1, x_4); +x_8 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__19___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__20(x_1, x_4); if (lean_obj_tag(x_8) == 0) { uint8_t x_9; @@ -18039,7 +17656,7 @@ 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 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853_(x_6); +x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668_(x_6); if (lean_obj_tag(x_13) == 0) { uint8_t x_14; @@ -18068,7 +17685,7 @@ lean_object* x_17; lean_object* x_18; x_17 = lean_ctor_get(x_13, 0); lean_inc(x_17); lean_dec(x_13); -x_18 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__17(x_12, x_5, x_17); +x_18 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__17(x_12, x_5, x_17); x_1 = x_18; x_2 = x_7; goto _start; @@ -18077,7 +17694,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__21(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__21(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_3) == 0) @@ -18099,7 +17716,7 @@ lean_inc(x_7); x_8 = lean_ctor_get(x_3, 3); lean_inc(x_8); lean_dec(x_3); -x_9 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__21(x_1, x_2, x_5); +x_9 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__21(x_1, x_2, x_5); if (lean_obj_tag(x_9) == 0) { uint8_t x_10; @@ -18128,7 +17745,7 @@ lean_object* x_13; lean_object* x_14; x_13 = lean_ctor_get(x_9, 0); lean_inc(x_13); lean_dec(x_9); -x_14 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853_(x_7); +x_14 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668_(x_7); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -18157,7 +17774,7 @@ lean_object* x_18; lean_object* x_19; x_18 = lean_ctor_get(x_14, 0); lean_inc(x_18); lean_dec(x_14); -x_19 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__17(x_13, x_6, x_18); +x_19 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__17(x_13, x_6, x_18); x_2 = x_19; x_3 = x_8; goto _start; @@ -18166,7 +17783,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__21___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__22(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__21___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__22(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) @@ -18188,7 +17805,7 @@ lean_inc(x_6); x_7 = lean_ctor_get(x_2, 3); lean_inc(x_7); lean_dec(x_2); -x_8 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__21___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__22(x_1, x_4); +x_8 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__21___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__22(x_1, x_4); if (lean_obj_tag(x_8) == 0) { uint8_t x_9; @@ -18217,7 +17834,7 @@ 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 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853_(x_6); +x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668_(x_6); if (lean_obj_tag(x_13) == 0) { uint8_t x_14; @@ -18246,7 +17863,7 @@ lean_object* x_17; lean_object* x_18; x_17 = lean_ctor_get(x_13, 0); lean_inc(x_17); lean_dec(x_13); -x_18 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__17(x_12, x_5, x_17); +x_18 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__17(x_12, x_5, x_17); x_1 = x_18; x_2 = x_7; goto _start; @@ -18255,7 +17872,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__23(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__23(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_3) == 0) @@ -18277,7 +17894,7 @@ lean_inc(x_7); x_8 = lean_ctor_get(x_3, 3); lean_inc(x_8); lean_dec(x_3); -x_9 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__23(x_1, x_2, x_5); +x_9 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__23(x_1, x_2, x_5); if (lean_obj_tag(x_9) == 0) { uint8_t x_10; @@ -18306,7 +17923,7 @@ lean_object* x_13; lean_object* x_14; x_13 = lean_ctor_get(x_9, 0); lean_inc(x_13); lean_dec(x_9); -x_14 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853_(x_7); +x_14 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668_(x_7); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -18335,7 +17952,7 @@ lean_object* x_18; lean_object* x_19; x_18 = lean_ctor_get(x_14, 0); lean_inc(x_18); lean_dec(x_14); -x_19 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__17(x_13, x_6, x_18); +x_19 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__17(x_13, x_6, x_18); x_2 = x_19; x_3 = x_8; goto _start; @@ -18344,7 +17961,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__23___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__24(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__23___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__24(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) @@ -18366,7 +17983,7 @@ lean_inc(x_6); x_7 = lean_ctor_get(x_2, 3); lean_inc(x_7); lean_dec(x_2); -x_8 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__23___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__24(x_1, x_4); +x_8 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__23___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__24(x_1, x_4); if (lean_obj_tag(x_8) == 0) { uint8_t x_9; @@ -18395,7 +18012,7 @@ 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 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853_(x_6); +x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668_(x_6); if (lean_obj_tag(x_13) == 0) { uint8_t x_14; @@ -18424,7 +18041,7 @@ lean_object* x_17; lean_object* x_18; x_17 = lean_ctor_get(x_13, 0); lean_inc(x_17); lean_dec(x_13); -x_18 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__17(x_12, x_5, x_17); +x_18 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__17(x_12, x_5, x_17); x_1 = x_18; x_2 = x_7; goto _start; @@ -18433,7 +18050,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__25(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__25(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_3) == 0) @@ -18455,7 +18072,7 @@ lean_inc(x_7); x_8 = lean_ctor_get(x_3, 3); lean_inc(x_8); lean_dec(x_3); -x_9 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__25(x_1, x_2, x_5); +x_9 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__25(x_1, x_2, x_5); if (lean_obj_tag(x_9) == 0) { uint8_t x_10; @@ -18484,7 +18101,7 @@ lean_object* x_13; lean_object* x_14; x_13 = lean_ctor_get(x_9, 0); lean_inc(x_13); lean_dec(x_9); -x_14 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853_(x_7); +x_14 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668_(x_7); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -18513,7 +18130,7 @@ lean_object* x_18; lean_object* x_19; x_18 = lean_ctor_get(x_14, 0); lean_inc(x_18); lean_dec(x_14); -x_19 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__17(x_13, x_6, x_18); +x_19 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__17(x_13, x_6, x_18); x_2 = x_19; x_3 = x_8; goto _start; @@ -18522,7 +18139,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__25___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__26(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__25___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__26(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) @@ -18544,7 +18161,7 @@ lean_inc(x_6); x_7 = lean_ctor_get(x_2, 3); lean_inc(x_7); lean_dec(x_2); -x_8 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__25___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__26(x_1, x_4); +x_8 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__25___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__26(x_1, x_4); if (lean_obj_tag(x_8) == 0) { uint8_t x_9; @@ -18573,7 +18190,7 @@ 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 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853_(x_6); +x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668_(x_6); if (lean_obj_tag(x_13) == 0) { uint8_t x_14; @@ -18602,7 +18219,7 @@ lean_object* x_17; lean_object* x_18; x_17 = lean_ctor_get(x_13, 0); lean_inc(x_17); lean_dec(x_13); -x_18 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__17(x_12, x_5, x_17); +x_18 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__17(x_12, x_5, x_17); x_1 = x_18; x_2 = x_7; goto _start; @@ -18611,7 +18228,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__27(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__27(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_3) == 0) @@ -18633,7 +18250,7 @@ lean_inc(x_7); x_8 = lean_ctor_get(x_3, 3); lean_inc(x_8); lean_dec(x_3); -x_9 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__27(x_1, x_2, x_5); +x_9 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__27(x_1, x_2, x_5); if (lean_obj_tag(x_9) == 0) { uint8_t x_10; @@ -18662,7 +18279,7 @@ lean_object* x_13; lean_object* x_14; x_13 = lean_ctor_get(x_9, 0); lean_inc(x_13); lean_dec(x_9); -x_14 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853_(x_7); +x_14 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668_(x_7); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -18691,7 +18308,7 @@ lean_object* x_18; lean_object* x_19; x_18 = lean_ctor_get(x_14, 0); lean_inc(x_18); lean_dec(x_14); -x_19 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__17(x_13, x_6, x_18); +x_19 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__17(x_13, x_6, x_18); x_2 = x_19; x_3 = x_8; goto _start; @@ -18700,7 +18317,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__27___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__28(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__27___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__28(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) @@ -18722,7 +18339,7 @@ lean_inc(x_6); x_7 = lean_ctor_get(x_2, 3); lean_inc(x_7); lean_dec(x_2); -x_8 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__27___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__28(x_1, x_4); +x_8 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__27___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__28(x_1, x_4); if (lean_obj_tag(x_8) == 0) { uint8_t x_9; @@ -18751,7 +18368,7 @@ 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 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853_(x_6); +x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668_(x_6); if (lean_obj_tag(x_13) == 0) { uint8_t x_14; @@ -18780,7 +18397,7 @@ lean_object* x_17; lean_object* x_18; x_17 = lean_ctor_get(x_13, 0); lean_inc(x_17); lean_dec(x_13); -x_18 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__17(x_12, x_5, x_17); +x_18 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__17(x_12, x_5, x_17); x_1 = x_18; x_2 = x_7; goto _start; @@ -18789,7 +18406,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__16(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__16(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; @@ -18798,7 +18415,7 @@ switch (lean_obj_tag(x_3)) { case 0: { lean_object* x_4; -x_4 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____spec__1___closed__1; +x_4 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____spec__1___closed__1; return x_4; } case 1: @@ -18831,7 +18448,7 @@ x_9 = lean_ctor_get(x_5, 0); lean_inc(x_9); lean_dec(x_5); x_10 = lean_box(0); -x_11 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__19___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__20(x_10, x_9); +x_11 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__19___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__20(x_10, x_9); if (lean_obj_tag(x_11) == 0) { uint8_t x_12; @@ -18917,7 +18534,7 @@ x_27 = lean_ctor_get(x_21, 0); lean_inc(x_27); lean_dec(x_21); x_28 = lean_box(0); -x_29 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__21___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__22(x_28, x_27); +x_29 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__21___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__22(x_28, x_27); if (lean_obj_tag(x_29) == 0) { uint8_t x_30; @@ -18996,7 +18613,7 @@ x_40 = lean_ctor_get(x_21, 0); lean_inc(x_40); lean_dec(x_21); x_41 = lean_box(0); -x_42 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__21___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__22(x_41, x_40); +x_42 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__21___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__22(x_41, x_40); if (lean_obj_tag(x_42) == 0) { lean_object* x_43; lean_object* x_44; lean_object* x_45; @@ -19080,7 +18697,7 @@ x_56 = lean_ctor_get(x_50, 0); lean_inc(x_56); lean_dec(x_50); x_57 = lean_box(0); -x_58 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__23___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__24(x_57, x_56); +x_58 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__23___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__24(x_57, x_56); if (lean_obj_tag(x_58) == 0) { uint8_t x_59; @@ -19159,7 +18776,7 @@ x_69 = lean_ctor_get(x_50, 0); lean_inc(x_69); lean_dec(x_50); x_70 = lean_box(0); -x_71 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__23___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__24(x_70, x_69); +x_71 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__23___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__24(x_70, x_69); if (lean_obj_tag(x_71) == 0) { lean_object* x_72; lean_object* x_73; lean_object* x_74; @@ -19243,7 +18860,7 @@ x_85 = lean_ctor_get(x_79, 0); lean_inc(x_85); lean_dec(x_79); x_86 = lean_box(0); -x_87 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__25___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__26(x_86, x_85); +x_87 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__25___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__26(x_86, x_85); if (lean_obj_tag(x_87) == 0) { uint8_t x_88; @@ -19322,7 +18939,7 @@ x_98 = lean_ctor_get(x_79, 0); lean_inc(x_98); lean_dec(x_79); x_99 = lean_box(0); -x_100 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__25___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__26(x_99, x_98); +x_100 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__25___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__26(x_99, x_98); if (lean_obj_tag(x_100) == 0) { lean_object* x_101; lean_object* x_102; lean_object* x_103; @@ -19406,7 +19023,7 @@ x_114 = lean_ctor_get(x_108, 0); lean_inc(x_114); lean_dec(x_108); x_115 = lean_box(0); -x_116 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__27___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__28(x_115, x_114); +x_116 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__27___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__28(x_115, x_114); if (lean_obj_tag(x_116) == 0) { uint8_t x_117; @@ -19485,7 +19102,7 @@ x_127 = lean_ctor_get(x_108, 0); lean_inc(x_127); lean_dec(x_108); x_128 = lean_box(0); -x_129 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__27___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__28(x_128, x_127); +x_129 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__27___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__28(x_128, x_127); if (lean_obj_tag(x_129) == 0) { lean_object* x_130; lean_object* x_131; lean_object* x_132; @@ -19534,7 +19151,7 @@ return x_136; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__1() { _start: { lean_object* x_1; @@ -19542,39 +19159,39 @@ x_1 = lean_mk_string_unchecked("WorkspaceEdit", 13, 13); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____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_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__1; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__2; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__1; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__2; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____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_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__2; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__2; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__3; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__5() { _start: { lean_object* x_1; @@ -19582,48 +19199,48 @@ x_1 = lean_mk_string_unchecked("changes\?", 8, 8); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____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_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__5; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__5; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__7() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__6; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__6; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__8; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__8; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__10() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__10() { _start: { lean_object* x_1; @@ -19631,48 +19248,48 @@ x_1 = lean_mk_string_unchecked("documentChanges\?", 16, 16); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__11() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____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_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__10; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____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_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__12() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__12() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__11; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__11; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__13() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__14() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__13; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__13; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__15() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__15() { _start: { lean_object* x_1; @@ -19680,54 +19297,54 @@ x_1 = lean_mk_string_unchecked("changeAnnotations\?", 18, 18); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__16() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____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_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__15; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____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_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__17() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__17() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__16; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__16; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__18() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__17; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__17; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__19() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__19() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__18; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__18; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -19737,7 +19354,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_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__9; +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__9; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -19749,7 +19366,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_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__9; +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__9; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -19763,9 +19380,9 @@ 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_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____closed__2; +x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____closed__2; lean_inc(x_1); -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__14(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__14(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -19776,7 +19393,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_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__14; +x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__14; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -19788,7 +19405,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_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__14; +x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__14; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -19802,8 +19419,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_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____closed__3; -x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__16(x_1, x_24); +x_24 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____closed__3; +x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__16(x_1, x_24); if (lean_obj_tag(x_25) == 0) { uint8_t x_26; @@ -19814,7 +19431,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_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__19; +x_28 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__19; x_29 = lean_string_append(x_28, x_27); lean_dec(x_27); lean_ctor_set(x_25, 0, x_29); @@ -19826,7 +19443,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_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__19; +x_31 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__19; x_32 = lean_string_append(x_31, x_30); lean_dec(x_30); x_33 = lean_alloc_ctor(0, 1, 0); @@ -19868,61 +19485,61 @@ return x_39; } } } -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__4(x_1, x_2, x_3); +x_4 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__4(x_1, x_2, x_3); lean_dec(x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__6(x_1, x_2, x_3); +x_4 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__6(x_1, x_2, x_3); lean_dec(x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__8(x_1, x_2, x_3); +x_4 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__8(x_1, x_2, x_3); lean_dec(x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__10___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__10___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__10(x_1, x_2, x_3); +x_4 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__10(x_1, x_2, x_3); lean_dec(x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__12___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__12___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__12(x_1, x_2, x_3); +x_4 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__12(x_1, x_2, x_3); lean_dec(x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____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_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____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_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__1(x_1, x_2); lean_dec(x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__15___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_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__15___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; @@ -19930,69 +19547,69 @@ 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_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__15(x_4, x_5, x_3); +x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__15(x_4, x_5, x_3); return x_6; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__14___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__14___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_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__14(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__14(x_1, x_2); lean_dec(x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__19___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__19___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__19(x_1, x_2, x_3); +x_4 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__19(x_1, x_2, x_3); lean_dec(x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__21___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__21___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__21(x_1, x_2, x_3); +x_4 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__21(x_1, x_2, x_3); lean_dec(x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__23___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__23___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__23(x_1, x_2, x_3); +x_4 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__23(x_1, x_2, x_3); lean_dec(x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__25___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__25___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__25(x_1, x_2, x_3); +x_4 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__25(x_1, x_2, x_3); lean_dec(x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__27___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__27___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__27(x_1, x_2, x_3); +x_4 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__27(x_1, x_2, x_3); lean_dec(x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__16___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__16___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_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__16(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__16(x_1, x_2); lean_dec(x_2); return x_3; } @@ -20001,7 +19618,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonWorkspaceEdit___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259_), 1, 0); return x_1; } } @@ -20451,7 +20068,7 @@ x_7 = l_Lean_Lsp_WorkspaceEdit_ofTextDocumentEdit(x_6); return x_7; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4806____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4621____closed__1() { _start: { lean_object* x_1; @@ -20459,19 +20076,19 @@ x_1 = lean_mk_string_unchecked("edit", 4, 4); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4806_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4621_(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; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2797____closed__1; -x_4 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(x_3, x_2); +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2612____closed__1; +x_4 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____spec__2(x_3, x_2); x_5 = lean_ctor_get(x_1, 1); lean_inc(x_5); lean_dec(x_1); -x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408_(x_5); -x_7 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4806____closed__1; +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223_(x_5); +x_7 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4621____closed__1; x_8 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_8, 0, x_7); lean_ctor_set(x_8, 1, x_6); @@ -20485,8 +20102,8 @@ lean_ctor_set(x_11, 1, x_9); x_12 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_12, 0, x_4); lean_ctor_set(x_12, 1, x_11); -x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_14 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_12, x_13); +x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3; +x_14 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_12, x_13); x_15 = l_Lean_Json_mkObj(x_14); return x_15; } @@ -20495,7 +20112,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonApplyWorkspaceEditParams___closed _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4806_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4621_), 1, 0); return x_1; } } @@ -20507,16 +20124,16 @@ x_1 = l_Lean_Lsp_instToJsonApplyWorkspaceEditParams___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____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_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444_(x_3); +x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259_(x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__1() { _start: { lean_object* x_1; @@ -20524,39 +20141,39 @@ x_1 = lean_mk_string_unchecked("ApplyWorkspaceEditParams", 24, 24); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____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_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__1; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__2; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__1; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__2; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____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_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__2; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__2; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__3; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__5() { _start: { lean_object* x_1; @@ -20564,95 +20181,95 @@ x_1 = lean_mk_string_unchecked("label\?", 6, 6); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____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_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__5; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__5; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__7() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__6; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__6; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__8; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__8; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__10() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____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_Basic_0__Lean_Lsp_toJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4806____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4621____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_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__11() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__11() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__10; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__10; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__12() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__11; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__11; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__13() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__12; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__12; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2797____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2612____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____spec__2(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____spec__2(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -20662,7 +20279,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_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__9; +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__9; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -20674,7 +20291,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_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__9; +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__9; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -20688,8 +20305,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_Basic_0__Lean_Lsp_toJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4806____closed__1; -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____spec__1(x_1, x_13); +x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4621____closed__1; +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -20699,7 +20316,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_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__13; +x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__13; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -20711,7 +20328,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_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__13; +x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__13; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -20750,11 +20367,11 @@ return x_28; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____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_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____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_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____spec__1(x_1, x_2); lean_dec(x_2); return x_3; } @@ -20763,7 +20380,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonApplyWorkspaceEditParams___clos _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663_), 1, 0); return x_1; } } @@ -20775,7 +20392,7 @@ x_1 = l_Lean_Lsp_instFromJsonApplyWorkspaceEditParams___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4990____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4805____closed__1() { _start: { lean_object* x_1; @@ -20783,7 +20400,7 @@ x_1 = lean_mk_string_unchecked("languageId", 10, 10); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4990____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4805____closed__2() { _start: { lean_object* x_1; @@ -20791,7 +20408,7 @@ x_1 = lean_mk_string_unchecked("text", 4, 4); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4990_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4805_(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; @@ -20806,7 +20423,7 @@ lean_inc(x_5); lean_dec(x_1); x_6 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_6, 0, x_2); -x_7 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080____closed__1; +x_7 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895____closed__1; x_8 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_8, 0, x_7); lean_ctor_set(x_8, 1, x_6); @@ -20816,7 +20433,7 @@ lean_ctor_set(x_10, 0, x_8); lean_ctor_set(x_10, 1, x_9); x_11 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_11, 0, x_3); -x_12 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4990____closed__1; +x_12 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4805____closed__1; x_13 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_13, 0, x_12); lean_ctor_set(x_13, 1, x_11); @@ -20826,7 +20443,7 @@ lean_ctor_set(x_14, 1, x_9); x_15 = l_Lean_JsonNumber_fromNat(x_4); x_16 = lean_alloc_ctor(2, 1, 0); lean_ctor_set(x_16, 0, x_15); -x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2440____closed__1; +x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2255____closed__1; x_18 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_18, 0, x_17); lean_ctor_set(x_18, 1, x_16); @@ -20835,7 +20452,7 @@ lean_ctor_set(x_19, 0, x_18); lean_ctor_set(x_19, 1, x_9); x_20 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_20, 0, x_5); -x_21 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4990____closed__2; +x_21 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4805____closed__2; x_22 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_22, 0, x_21); lean_ctor_set(x_22, 1, x_20); @@ -20854,8 +20471,8 @@ lean_ctor_set(x_26, 1, x_25); x_27 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_27, 0, x_10); lean_ctor_set(x_27, 1, x_26); -x_28 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_29 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_27, x_28); +x_28 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3; +x_29 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_27, x_28); x_30 = l_Lean_Json_mkObj(x_29); return x_30; } @@ -20864,7 +20481,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonTextDocumentItem___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4990_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4805_), 1, 0); return x_1; } } @@ -20876,7 +20493,7 @@ x_1 = l_Lean_Lsp_instToJsonTextDocumentItem___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__1() { _start: { lean_object* x_1; @@ -20884,188 +20501,188 @@ x_1 = lean_mk_string_unchecked("TextDocumentItem", 16, 16); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____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_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__1; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__2; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__1; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__2; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____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_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__2; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__2; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__3; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__6; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__6; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__5; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__5; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____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_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4990____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4805____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_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__8() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__7; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__8; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__8; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__10() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__9; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__9; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__11() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____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_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2440____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2255____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_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__12() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__12() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__11; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__11; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__13() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__14() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__13; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__13; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__15() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____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_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4990____closed__2; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4805____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_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__16() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__16() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__15; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__15; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__17() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__16; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__16; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__18() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__17; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__17; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -21075,7 +20692,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_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__6; +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__6; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -21087,7 +20704,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_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__6; +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__6; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -21101,9 +20718,9 @@ 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_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4990____closed__1; +x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4805____closed__1; lean_inc(x_1); -x_14 = l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -21114,7 +20731,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_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__10; +x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__10; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -21126,7 +20743,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_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__10; +x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__10; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -21140,9 +20757,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_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2440____closed__1; +x_24 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2255____closed__1; lean_inc(x_1); -x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____spec__1(x_1, x_24); +x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____spec__1(x_1, x_24); if (lean_obj_tag(x_25) == 0) { uint8_t x_26; @@ -21154,7 +20771,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_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__14; +x_28 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__14; x_29 = lean_string_append(x_28, x_27); lean_dec(x_27); lean_ctor_set(x_25, 0, x_29); @@ -21166,7 +20783,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_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__14; +x_31 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__14; x_32 = lean_string_append(x_31, x_30); lean_dec(x_30); x_33 = lean_alloc_ctor(0, 1, 0); @@ -21180,8 +20797,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_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4990____closed__2; -x_36 = l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(x_1, x_35); +x_35 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4805____closed__2; +x_36 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__1(x_1, x_35); if (lean_obj_tag(x_36) == 0) { uint8_t x_37; @@ -21193,7 +20810,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_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__18; +x_39 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__18; x_40 = lean_string_append(x_39, x_38); lean_dec(x_38); lean_ctor_set(x_36, 0, x_40); @@ -21205,7 +20822,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_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__18; +x_42 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__18; x_43 = lean_string_append(x_42, x_41); lean_dec(x_41); x_44 = lean_alloc_ctor(0, 1, 0); @@ -21254,7 +20871,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonTextDocumentItem___closed__1() _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885_), 1, 0); return x_1; } } @@ -21266,7 +20883,7 @@ x_1 = l_Lean_Lsp_instFromJsonTextDocumentItem___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5274____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5089____closed__1() { _start: { lean_object* x_1; @@ -21274,14 +20891,14 @@ x_1 = lean_mk_string_unchecked("position", 8, 8); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5274_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5089_(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; 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_2313_(x_2); -x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2608____closed__1; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2128_(x_2); +x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2423____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); @@ -21292,8 +20909,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_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406_(x_8); -x_10 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5274____closed__1; +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221_(x_8); +x_10 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5089____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); @@ -21306,8 +20923,8 @@ lean_ctor_set(x_13, 1, x_6); x_14 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_14, 0, x_7); lean_ctor_set(x_14, 1, x_13); -x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_16 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_14, x_15); +x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3; +x_16 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_14, x_15); x_17 = l_Lean_Json_mkObj(x_16); return x_17; } @@ -21316,7 +20933,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonTextDocumentPositionParams___clos _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5274_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5089_), 1, 0); return x_1; } } @@ -21328,16 +20945,16 @@ x_1 = l_Lean_Lsp_instToJsonTextDocumentPositionParams___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____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_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351_(x_3); +x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166_(x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__1() { _start: { lean_object* x_1; @@ -21345,106 +20962,106 @@ x_1 = lean_mk_string_unchecked("TextDocumentPositionParams", 26, 26); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____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_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__1; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__2; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__1; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__2; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____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_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__2; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__2; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__3; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__6; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__6; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__5; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__5; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____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_Basic_0__Lean_Lsp_toJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5274____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5089____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_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__8() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__7; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__8; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__8; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__10() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__9; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__9; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2608____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2423____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -21454,7 +21071,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_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__6; +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__6; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -21466,7 +21083,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_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__6; +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__6; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -21480,8 +21097,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_Basic_0__Lean_Lsp_toJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5274____closed__1; -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____spec__1(x_1, x_13); +x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5089____closed__1; +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -21491,7 +21108,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_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__10; +x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__10; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -21503,7 +21120,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_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__10; +x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__10; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -21542,11 +21159,11 @@ return x_28; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____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_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____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_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____spec__1(x_1, x_2); lean_dec(x_2); return x_3; } @@ -21555,7 +21172,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonTextDocumentPositionParams___cl _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141_), 1, 0); return x_1; } } @@ -21581,7 +21198,7 @@ LEAN_EXPORT lean_object* l_Lean_Lsp_instToStringTextDocumentPositionParams(lean_ lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; 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_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); -x_3 = l_Lean_Lsp_instInhabitedCancelParams___closed__1; +x_3 = l_Lean_Lsp_instInhabitedLocation___closed__1; x_4 = lean_string_append(x_3, x_2); lean_dec(x_2); x_5 = l_Lean_Lsp_instToStringTextDocumentPositionParams___closed__1; @@ -21605,7 +21222,7 @@ x_15 = lean_string_append(x_14, x_3); return x_15; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5505____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5320____closed__1() { _start: { lean_object* x_1; @@ -21613,7 +21230,7 @@ x_1 = lean_mk_string_unchecked("language", 8, 8); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5505____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5320____closed__2() { _start: { lean_object* x_1; @@ -21621,7 +21238,7 @@ x_1 = lean_mk_string_unchecked("scheme", 6, 6); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5505____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5320____closed__3() { _start: { lean_object* x_1; @@ -21629,23 +21246,23 @@ x_1 = lean_mk_string_unchecked("pattern", 7, 7); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5505_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5320_(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; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5505____closed__1; -x_4 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(x_3, x_2); +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5320____closed__1; +x_4 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____spec__2(x_3, x_2); x_5 = lean_ctor_get(x_1, 1); lean_inc(x_5); -x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5505____closed__2; -x_7 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(x_6, x_5); +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5320____closed__2; +x_7 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____spec__2(x_6, x_5); x_8 = lean_ctor_get(x_1, 2); lean_inc(x_8); lean_dec(x_1); -x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5505____closed__3; -x_10 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(x_9, x_8); +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5320____closed__3; +x_10 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____spec__2(x_9, x_8); x_11 = lean_box(0); x_12 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_12, 0, x_10); @@ -21656,8 +21273,8 @@ lean_ctor_set(x_13, 1, x_12); x_14 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_14, 0, x_4); lean_ctor_set(x_14, 1, x_13); -x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_16 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_14, x_15); +x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3; +x_16 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_14, x_15); x_17 = l_Lean_Json_mkObj(x_16); return x_17; } @@ -21666,7 +21283,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonDocumentFilter___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5505_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5320_), 1, 0); return x_1; } } @@ -21678,7 +21295,7 @@ x_1 = l_Lean_Lsp_instToJsonDocumentFilter___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__1() { _start: { lean_object* x_1; @@ -21686,39 +21303,39 @@ x_1 = lean_mk_string_unchecked("DocumentFilter", 14, 14); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____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_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__1; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__2; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__1; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__2; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____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_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__2; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__2; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__3; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__5() { _start: { lean_object* x_1; @@ -21726,48 +21343,48 @@ x_1 = lean_mk_string_unchecked("language\?", 9, 9); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____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_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__5; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__5; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__7() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__6; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__6; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__8; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__8; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__10() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__10() { _start: { lean_object* x_1; @@ -21775,48 +21392,48 @@ x_1 = lean_mk_string_unchecked("scheme\?", 7, 7); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__11() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____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_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__10; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____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_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__12() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__12() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__11; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__11; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__13() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__14() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__13; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__13; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__15() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__15() { _start: { lean_object* x_1; @@ -21824,54 +21441,54 @@ x_1 = lean_mk_string_unchecked("pattern\?", 8, 8); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__16() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____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_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__15; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____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_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__17() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__17() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__16; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__16; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__18() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__17; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__17; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__19() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__19() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__18; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__18; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5505____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5320____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____spec__2(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____spec__2(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -21881,7 +21498,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_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__9; +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__9; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -21893,7 +21510,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_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__9; +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__9; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -21907,9 +21524,9 @@ 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_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5505____closed__2; +x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5320____closed__2; lean_inc(x_1); -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____spec__2(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____spec__2(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -21920,7 +21537,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_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__14; +x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__14; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -21932,7 +21549,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_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__14; +x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__14; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -21946,8 +21563,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_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5505____closed__3; -x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____spec__2(x_1, x_24); +x_24 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5320____closed__3; +x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____spec__2(x_1, x_24); if (lean_obj_tag(x_25) == 0) { uint8_t x_26; @@ -21958,7 +21575,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_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__19; +x_28 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__19; x_29 = lean_string_append(x_28, x_27); lean_dec(x_27); lean_ctor_set(x_25, 0, x_29); @@ -21970,7 +21587,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_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__19; +x_31 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__19; x_32 = lean_string_append(x_31, x_30); lean_dec(x_30); x_33 = lean_alloc_ctor(0, 1, 0); @@ -22016,7 +21633,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonDocumentFilter___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356_), 1, 0); return x_1; } } @@ -22046,7 +21663,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_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541_(x_6); +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356_(x_6); if (lean_obj_tag(x_9) == 0) { uint8_t x_10; @@ -22092,10 +21709,10 @@ case 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; x_2 = lean_unsigned_to_nat(80u); x_3 = l_Lean_Json_pretty(x_1, x_2); -x_4 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_4 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_5 = lean_string_append(x_4, x_3); lean_dec(x_3); -x_6 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_6 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_7 = lean_string_append(x_5, x_6); x_8 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_8, 0, x_7); @@ -22106,10 +21723,10 @@ case 1: lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; x_9 = lean_unsigned_to_nat(80u); x_10 = l_Lean_Json_pretty(x_1, x_9); -x_11 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_11 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_12 = lean_string_append(x_11, x_10); lean_dec(x_10); -x_13 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_13 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_14 = lean_string_append(x_12, x_13); x_15 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_15, 0, x_14); @@ -22138,10 +21755,10 @@ if (x_22 == 0) lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; x_23 = lean_ctor_get(x_1, 0); lean_dec(x_23); -x_24 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_24 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_25 = lean_string_append(x_24, x_21); lean_dec(x_21); -x_26 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_26 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_27 = lean_string_append(x_25, x_26); lean_ctor_set_tag(x_1, 0); lean_ctor_set(x_1, 0, x_27); @@ -22151,10 +21768,10 @@ else { lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_dec(x_1); -x_28 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_28 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_29 = lean_string_append(x_28, x_21); lean_dec(x_21); -x_30 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_30 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_31 = lean_string_append(x_29, x_30); x_32 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_32, 0, x_31); @@ -22191,7 +21808,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_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5505_(x_5); +x_8 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5320_(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); @@ -22225,18 +21842,26 @@ x_6 = l_Array_mapMUnsafe_map___at_Lean_Lsp_instToJsonDocumentSelector___spec__1( return x_6; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5727_(lean_object* x_1) { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5542____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("id", 2, 2); +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5542_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__2; -x_3 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(x_2, x_1); +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5542____closed__1; +x_3 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____spec__2(x_2, x_1); x_4 = lean_box(0); x_5 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_5, 0, x_3); lean_ctor_set(x_5, 1, x_4); -x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_7 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_5, x_6); +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3; +x_7 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_5, x_6); x_8 = l_Lean_Json_mkObj(x_7); return x_8; } @@ -22245,7 +21870,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonStaticRegistrationOptions___close _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5727_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5542_), 1, 0); return x_1; } } @@ -22257,7 +21882,7 @@ x_1 = l_Lean_Lsp_instToJsonStaticRegistrationOptions___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__1() { _start: { lean_object* x_1; @@ -22265,39 +21890,39 @@ x_1 = lean_mk_string_unchecked("StaticRegistrationOptions", 25, 25); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____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_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__1; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__2; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__1; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__2; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____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_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__2; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__2; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__3; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__5() { _start: { lean_object* x_1; @@ -22305,53 +21930,53 @@ x_1 = lean_mk_string_unchecked("id\?", 3, 3); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____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_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__5; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__5; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__7() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__6; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__6; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__8; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__8; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__2; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____spec__2(x_1, x_2); +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5542____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_1860____spec__2(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -22360,7 +21985,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_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__9; +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__9; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -22372,7 +21997,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_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__9; +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__9; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -22405,7 +22030,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonStaticRegistrationOptions___clo _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570_), 1, 0); return x_1; } } @@ -22417,7 +22042,7 @@ x_1 = l_Lean_Lsp_instFromJsonStaticRegistrationOptions___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5836____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5651____spec__1(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) @@ -22472,7 +22097,7 @@ return x_19; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5836____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5651____closed__1() { _start: { lean_object* x_1; @@ -22480,18 +22105,18 @@ x_1 = lean_mk_string_unchecked("documentSelector", 16, 16); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5836_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5651_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5836____closed__1; -x_3 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5836____spec__1(x_2, x_1); +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5651____closed__1; +x_3 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5651____spec__1(x_2, x_1); x_4 = lean_box(0); x_5 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_5, 0, x_3); lean_ctor_set(x_5, 1, x_4); -x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_7 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_5, x_6); +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3; +x_7 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_5, x_6); x_8 = l_Lean_Json_mkObj(x_7); return x_8; } @@ -22500,7 +22125,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonTextDocumentRegistrationOptions__ _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5836_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5651_), 1, 0); return x_1; } } @@ -22512,7 +22137,7 @@ x_1 = l_Lean_Lsp_instToJsonTextDocumentRegistrationOptions___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____spec__1(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; @@ -22521,7 +22146,7 @@ switch (lean_obj_tag(x_3)) { case 0: { lean_object* x_4; -x_4 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____spec__1___closed__1; +x_4 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____spec__1___closed__1; return x_4; } case 1: @@ -22529,10 +22154,10 @@ case 1: lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_5 = lean_unsigned_to_nat(80u); x_6 = l_Lean_Json_pretty(x_3, x_5); -x_7 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_7 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_8 = lean_string_append(x_7, x_6); lean_dec(x_6); -x_9 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_9 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_10 = lean_string_append(x_8, x_9); x_11 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_11, 0, x_10); @@ -22661,10 +22286,10 @@ if (x_37 == 0) lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; x_38 = lean_ctor_get(x_3, 0); lean_dec(x_38); -x_39 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_39 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_40 = lean_string_append(x_39, x_36); lean_dec(x_36); -x_41 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_41 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_42 = lean_string_append(x_40, x_41); lean_ctor_set_tag(x_3, 0); lean_ctor_set(x_3, 0, x_42); @@ -22674,10 +22299,10 @@ else { lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_dec(x_3); -x_43 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_43 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_44 = lean_string_append(x_43, x_36); lean_dec(x_36); -x_45 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_45 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_46 = lean_string_append(x_44, x_45); x_47 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_47, 0, x_46); @@ -22687,7 +22312,7 @@ return x_47; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__1() { _start: { lean_object* x_1; @@ -22695,39 +22320,39 @@ x_1 = lean_mk_string_unchecked("TextDocumentRegistrationOptions", 31, 31); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____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_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__1; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__2; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__1; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__2; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____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_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__2; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__2; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__3; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__5() { _start: { lean_object* x_1; @@ -22735,53 +22360,53 @@ x_1 = lean_mk_string_unchecked("documentSelector\?", 17, 17); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____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_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__5; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__5; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__7() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__6; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__6; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__8; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__8; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5836____closed__1; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____spec__1(x_1, x_2); +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5651____closed__1; +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -22790,7 +22415,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_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__9; +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__9; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -22802,7 +22427,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_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__9; +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__9; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -22831,11 +22456,11 @@ return x_14; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____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_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____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_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____spec__1(x_1, x_2); lean_dec(x_2); return x_3; } @@ -22844,7 +22469,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonTextDocumentRegistrationOptions _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679_), 1, 0); return x_1; } } @@ -22990,7 +22615,7 @@ x_6 = lean_box(x_5); return x_6; } } -LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupKind____x40_Lean_Data_Lsp_Basic___hyg_6036_(uint8_t x_1) { +LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupKind____x40_Lean_Data_Lsp_Basic___hyg_5851_(uint8_t x_1) { _start: { if (x_1 == 0) @@ -23007,13 +22632,13 @@ return x_3; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupKind____x40_Lean_Data_Lsp_Basic___hyg_6036____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupKind____x40_Lean_Data_Lsp_Basic___hyg_5851____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_Basic_0__Lean_Lsp_hashMarkupKind____x40_Lean_Data_Lsp_Basic___hyg_6036_(x_2); +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupKind____x40_Lean_Data_Lsp_Basic___hyg_5851_(x_2); x_4 = lean_box_uint64(x_3); return x_4; } @@ -23022,7 +22647,7 @@ static lean_object* _init_l_Lean_Lsp_instHashableMarkupKind___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupKind____x40_Lean_Data_Lsp_Basic___hyg_6036____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupKind____x40_Lean_Data_Lsp_Basic___hyg_5851____boxed), 1, 0); return x_1; } } @@ -23188,7 +22813,7 @@ x_3 = l_Lean_Lsp_instToJsonMarkupKind(x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6159____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_5974____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -23200,19 +22825,19 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6159____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_5974____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6159____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_5974____closed__1; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6159____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_5974____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -23224,19 +22849,19 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6159____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_5974____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_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6159____closed__3; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_5974____closed__3; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6159_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_5974_(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; @@ -23246,7 +22871,7 @@ x_4 = lean_ctor_get(x_1, 0); lean_inc(x_4); x_5 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_5, 0, x_4); -x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1840____closed__1; +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1655____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); @@ -23259,34 +22884,34 @@ lean_ctor_set(x_9, 1, x_3); if (x_2 == 0) { lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_10 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6159____closed__2; +x_10 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_5974____closed__2; x_11 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_11, 0, x_10); lean_ctor_set(x_11, 1, x_9); -x_12 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_13 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_11, x_12); +x_12 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3; +x_13 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_11, x_12); x_14 = l_Lean_Json_mkObj(x_13); return x_14; } else { lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6159____closed__4; +x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_5974____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_9); -x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_18 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_16, x_17); +x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3; +x_18 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_16, x_17); x_19 = l_Lean_Json_mkObj(x_18); return x_19; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6159____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_5974____boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6159_(x_1); +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_5974_(x_1); lean_dec(x_1); return x_2; } @@ -23295,7 +22920,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonMarkupContent___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6159____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_5974____boxed), 1, 0); return x_1; } } @@ -23307,7 +22932,7 @@ x_1 = l_Lean_Lsp_instToJsonMarkupContent___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____spec__1(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; @@ -23356,7 +22981,7 @@ return x_12; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__1() { _start: { lean_object* x_1; @@ -23364,39 +22989,39 @@ x_1 = lean_mk_string_unchecked("MarkupContent", 13, 13); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____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_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__1; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__2; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__1; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__2; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____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_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__2; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__2; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__3; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -23406,64 +23031,64 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__6() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__5; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__5; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__6; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__6; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__7; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__7; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__6; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__6; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__10() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__9; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__9; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; x_2 = l_Lean_Lsp_instToJsonDocumentChange___closed__3; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -23473,7 +23098,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_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__8; +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__8; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -23485,7 +23110,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_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__8; +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__8; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -23499,8 +23124,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_Basic_0__Lean_Lsp_toJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1840____closed__1; -x_14 = l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(x_1, x_13); +x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1655____closed__1; +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -23510,7 +23135,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_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__10; +x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__10; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -23522,7 +23147,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_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__10; +x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__10; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -23565,11 +23190,11 @@ return x_30; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____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_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____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_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____spec__1(x_1, x_2); lean_dec(x_2); return x_3; } @@ -23578,7 +23203,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonMarkupContent___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026_), 1, 0); return x_1; } } @@ -23590,7 +23215,7 @@ x_1 = l_Lean_Lsp_instFromJsonMarkupContent___closed__1; return x_1; } } -LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_decEqMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6317_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_decEqMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6132_(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; uint8_t x_5; lean_object* x_6; uint8_t x_7; @@ -23613,11 +23238,11 @@ return x_9; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_decEqMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6317____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_decEqMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6132____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_Basic_0__Lean_Lsp_decEqMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6317_(x_1, x_2); +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_decEqMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6132_(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); @@ -23628,7 +23253,7 @@ LEAN_EXPORT uint8_t l_Lean_Lsp_instDecidableEqMarkupContent(lean_object* x_1, le _start: { uint8_t x_3; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_decEqMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6317_(x_1, x_2); +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_decEqMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6132_(x_1, x_2); return x_3; } } @@ -23643,25 +23268,25 @@ x_4 = lean_box(x_3); return x_4; } } -LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6462_(lean_object* x_1) { +LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6277_(lean_object* x_1) { _start: { uint8_t x_2; lean_object* x_3; uint64_t x_4; uint64_t x_5; uint64_t x_6; uint64_t x_7; uint64_t x_8; x_2 = lean_ctor_get_uint8(x_1, sizeof(void*)*1); x_3 = lean_ctor_get(x_1, 0); x_4 = 0; -x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupKind____x40_Lean_Data_Lsp_Basic___hyg_6036_(x_2); +x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupKind____x40_Lean_Data_Lsp_Basic___hyg_5851_(x_2); x_6 = lean_uint64_mix_hash(x_4, x_5); x_7 = lean_string_hash(x_3); x_8 = lean_uint64_mix_hash(x_6, x_7); return x_8; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6462____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6277____boxed(lean_object* x_1) { _start: { uint64_t x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6462_(x_1); +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6277_(x_1); lean_dec(x_1); x_3 = lean_box_uint64(x_2); return x_3; @@ -23671,7 +23296,7 @@ static lean_object* _init_l_Lean_Lsp_instHashableMarkupContent___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6462____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6277____boxed), 1, 0); return x_1; } } @@ -23683,7 +23308,7 @@ x_1 = l_Lean_Lsp_instHashableMarkupContent___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6531____rarg___closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6346____rarg___closed__1() { _start: { lean_object* x_1; @@ -23691,7 +23316,7 @@ x_1 = lean_mk_string_unchecked("token", 5, 5); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6531____rarg(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6346____rarg(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; @@ -23703,7 +23328,7 @@ x_4 = lean_ctor_get(x_2, 0); x_5 = lean_ctor_get(x_2, 1); x_6 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_6, 0, x_4); -x_7 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6531____rarg___closed__1; +x_7 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6346____rarg___closed__1; lean_ctor_set(x_2, 1, x_6); lean_ctor_set(x_2, 0, x_7); x_8 = lean_box(0); @@ -23711,7 +23336,7 @@ x_9 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_9, 0, x_2); lean_ctor_set(x_9, 1, x_8); x_10 = lean_apply_1(x_1, x_5); -x_11 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1840____closed__1; +x_11 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1655____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); @@ -23724,8 +23349,8 @@ lean_ctor_set(x_14, 1, x_8); x_15 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_15, 0, x_9); lean_ctor_set(x_15, 1, x_14); -x_16 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_17 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_15, x_16); +x_16 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3; +x_17 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_15, x_16); x_18 = l_Lean_Json_mkObj(x_17); return x_18; } @@ -23739,7 +23364,7 @@ lean_inc(x_19); lean_dec(x_2); x_21 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_21, 0, x_19); -x_22 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6531____rarg___closed__1; +x_22 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6346____rarg___closed__1; x_23 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_23, 0, x_22); lean_ctor_set(x_23, 1, x_21); @@ -23748,7 +23373,7 @@ x_25 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_25, 0, x_23); lean_ctor_set(x_25, 1, x_24); x_26 = lean_apply_1(x_1, x_20); -x_27 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1840____closed__1; +x_27 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1655____closed__1; x_28 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_28, 0, x_27); lean_ctor_set(x_28, 1, x_26); @@ -23761,18 +23386,18 @@ lean_ctor_set(x_30, 1, x_24); x_31 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_31, 0, x_25); lean_ctor_set(x_31, 1, x_30); -x_32 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_33 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_31, x_32); +x_32 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3; +x_33 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_31, x_32); x_34 = l_Lean_Json_mkObj(x_33); return x_34; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6531_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6346_(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6531____rarg), 2, 0); +x_2 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6346____rarg), 2, 0); return x_2; } } @@ -23780,7 +23405,7 @@ LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonProgressParams___rarg(lean_object* _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6531____rarg), 2, 1); +x_2 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6346____rarg), 2, 1); lean_closure_set(x_2, 0, x_1); return x_2; } @@ -23793,7 +23418,7 @@ x_2 = lean_alloc_closure((void*)(l_Lean_Lsp_instToJsonProgressParams___rarg), 1, return x_2; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6641____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6456____closed__1() { _start: { lean_object* x_1; @@ -23801,7 +23426,7 @@ x_1 = lean_mk_string_unchecked("message", 7, 7); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6641____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6456____closed__2() { _start: { lean_object* x_1; @@ -23809,7 +23434,7 @@ x_1 = lean_mk_string_unchecked("cancellable", 11, 11); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6641____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6456____closed__3() { _start: { lean_object* x_1; @@ -23817,7 +23442,7 @@ x_1 = lean_mk_string_unchecked("percentage", 10, 10); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6641_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6456_(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; @@ -23839,19 +23464,19 @@ x_9 = lean_box(0); x_10 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_10, 0, x_8); lean_ctor_set(x_10, 1, x_9); -x_11 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6641____closed__1; -x_12 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(x_11, x_3); +x_11 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6456____closed__1; +x_12 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____spec__2(x_11, x_3); x_13 = lean_alloc_ctor(1, 0, 1); lean_ctor_set_uint8(x_13, 0, x_4); -x_14 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6641____closed__2; +x_14 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6456____closed__2; x_15 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_15, 0, x_14); lean_ctor_set(x_15, 1, x_13); x_16 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_16, 0, x_15); lean_ctor_set(x_16, 1, x_9); -x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6641____closed__3; -x_18 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2440____spec__1(x_17, x_5); +x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6456____closed__3; +x_18 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2255____spec__1(x_17, x_5); x_19 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_19, 0, x_18); lean_ctor_set(x_19, 1, x_9); @@ -23864,8 +23489,8 @@ lean_ctor_set(x_21, 1, x_20); x_22 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_22, 0, x_10); lean_ctor_set(x_22, 1, x_21); -x_23 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_24 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_22, x_23); +x_23 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3; +x_24 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_22, x_23); x_25 = l_Lean_Json_mkObj(x_24); return x_25; } @@ -23874,7 +23499,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonWorkDoneProgressReport___closed__ _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6641_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6456_), 1, 0); return x_1; } } @@ -23886,7 +23511,7 @@ x_1 = l_Lean_Lsp_instToJsonWorkDoneProgressReport___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressBegin____x40_Lean_Data_Lsp_Basic___hyg_6722_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressBegin____x40_Lean_Data_Lsp_Basic___hyg_6537_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; uint8_t x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; @@ -23910,25 +23535,25 @@ x_10 = lean_box(0); x_11 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_11, 0, x_9); lean_ctor_set(x_11, 1, x_10); -x_12 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6641____closed__1; -x_13 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(x_12, x_4); +x_12 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6456____closed__1; +x_13 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____spec__2(x_12, x_4); x_14 = lean_alloc_ctor(1, 0, 1); lean_ctor_set_uint8(x_14, 0, x_5); -x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6641____closed__2; +x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6456____closed__2; x_16 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_16, 0, x_15); lean_ctor_set(x_16, 1, x_14); x_17 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_17, 0, x_16); lean_ctor_set(x_17, 1, x_10); -x_18 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6641____closed__3; -x_19 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2440____spec__1(x_18, x_6); +x_18 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6456____closed__3; +x_19 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2255____spec__1(x_18, x_6); x_20 = lean_ctor_get(x_1, 1); lean_inc(x_20); lean_dec(x_1); x_21 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_21, 0, x_20); -x_22 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____closed__1; +x_22 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____closed__1; x_23 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_23, 0, x_22); lean_ctor_set(x_23, 1, x_21); @@ -23950,8 +23575,8 @@ lean_ctor_set(x_28, 1, x_27); x_29 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_29, 0, x_11); lean_ctor_set(x_29, 1, x_28); -x_30 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_31 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_29, x_30); +x_30 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3; +x_31 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_29, x_30); x_32 = l_Lean_Json_mkObj(x_31); return x_32; } @@ -23960,7 +23585,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonWorkDoneProgressBegin___closed__1 _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressBegin____x40_Lean_Data_Lsp_Basic___hyg_6722_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressBegin____x40_Lean_Data_Lsp_Basic___hyg_6537_), 1, 0); return x_1; } } @@ -23972,7 +23597,7 @@ x_1 = l_Lean_Lsp_instToJsonWorkDoneProgressBegin___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressEnd____x40_Lean_Data_Lsp_Basic___hyg_6818_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressEnd____x40_Lean_Data_Lsp_Basic___hyg_6633_(lean_object* x_1) { _start: { uint8_t x_2; @@ -23991,16 +23616,16 @@ x_7 = lean_box(0); x_8 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_8, 0, x_1); lean_ctor_set(x_8, 1, x_7); -x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6641____closed__1; -x_10 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(x_9, x_4); +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6456____closed__1; +x_10 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____spec__2(x_9, x_4); x_11 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_11, 0, x_10); lean_ctor_set(x_11, 1, x_7); x_12 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_12, 0, x_8); lean_ctor_set(x_12, 1, x_11); -x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_14 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_12, x_13); +x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3; +x_14 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_12, x_13); x_15 = l_Lean_Json_mkObj(x_14); return x_15; } @@ -24022,16 +23647,16 @@ x_21 = lean_box(0); x_22 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_22, 0, x_20); lean_ctor_set(x_22, 1, x_21); -x_23 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6641____closed__1; -x_24 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(x_23, x_17); +x_23 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6456____closed__1; +x_24 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____spec__2(x_23, x_17); x_25 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_25, 0, x_24); lean_ctor_set(x_25, 1, x_21); x_26 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_26, 0, x_22); lean_ctor_set(x_26, 1, x_25); -x_27 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_28 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_26, x_27); +x_27 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3; +x_28 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_26, x_27); x_29 = l_Lean_Json_mkObj(x_28); return x_29; } @@ -24041,7 +23666,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonWorkDoneProgressEnd___closed__1() _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressEnd____x40_Lean_Data_Lsp_Basic___hyg_6818_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressEnd____x40_Lean_Data_Lsp_Basic___hyg_6633_), 1, 0); return x_1; } } @@ -24053,7 +23678,7 @@ x_1 = l_Lean_Lsp_instToJsonWorkDoneProgressEnd___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6874____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6689____spec__1(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) @@ -24100,7 +23725,7 @@ return x_12; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6874____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6689____closed__1() { _start: { lean_object* x_1; @@ -24108,18 +23733,18 @@ x_1 = lean_mk_string_unchecked("workDoneToken", 13, 13); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6874_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6689_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6874____closed__1; -x_3 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6874____spec__1(x_2, x_1); +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6689____closed__1; +x_3 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6689____spec__1(x_2, x_1); x_4 = lean_box(0); x_5 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_5, 0, x_3); lean_ctor_set(x_5, 1, x_4); -x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_7 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_5, x_6); +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3; +x_7 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_5, x_6); x_8 = l_Lean_Json_mkObj(x_7); return x_8; } @@ -24128,7 +23753,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonWorkDoneProgressParams___closed__ _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6874_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6689_), 1, 0); return x_1; } } @@ -24140,7 +23765,7 @@ x_1 = l_Lean_Lsp_instToJsonWorkDoneProgressParams___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____spec__1(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; @@ -24149,7 +23774,7 @@ switch (lean_obj_tag(x_3)) { case 0: { lean_object* x_4; -x_4 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____spec__1___closed__1; +x_4 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____spec__1___closed__1; return x_4; } case 1: @@ -24310,7 +23935,7 @@ return x_31; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__1() { _start: { lean_object* x_1; @@ -24318,39 +23943,39 @@ x_1 = lean_mk_string_unchecked("WorkDoneProgressParams", 22, 22); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____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_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__1; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__2; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__1; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__2; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____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_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__2; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__2; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__3; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__5() { _start: { lean_object* x_1; @@ -24358,53 +23983,53 @@ x_1 = lean_mk_string_unchecked("workDoneToken\?", 14, 14); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____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_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__5; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__5; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__7() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__6; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__6; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__8; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__8; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6874____closed__1; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____spec__1(x_1, x_2); +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6689____closed__1; +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -24413,7 +24038,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_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__9; +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__9; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -24425,7 +24050,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_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__9; +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__9; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -24454,11 +24079,11 @@ return x_14; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____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_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____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_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____spec__1(x_1, x_2); lean_dec(x_2); return x_3; } @@ -24467,7 +24092,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonWorkDoneProgressParams___closed _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717_), 1, 0); return x_1; } } @@ -24479,7 +24104,7 @@ x_1 = l_Lean_Lsp_instFromJsonWorkDoneProgressParams___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6983____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6798____closed__1() { _start: { lean_object* x_1; @@ -24487,18 +24112,18 @@ x_1 = lean_mk_string_unchecked("partialResultToken", 18, 18); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6983_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6798_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6983____closed__1; -x_3 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6874____spec__1(x_2, x_1); +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6798____closed__1; +x_3 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6689____spec__1(x_2, x_1); x_4 = lean_box(0); x_5 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_5, 0, x_3); lean_ctor_set(x_5, 1, x_4); -x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_7 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_5, x_6); +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3; +x_7 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_5, x_6); x_8 = l_Lean_Json_mkObj(x_7); return x_8; } @@ -24507,7 +24132,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonPartialResultParams___closed__1() _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6983_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6798_), 1, 0); return x_1; } } @@ -24519,7 +24144,7 @@ x_1 = l_Lean_Lsp_instToJsonPartialResultParams___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__1() { _start: { lean_object* x_1; @@ -24527,39 +24152,39 @@ x_1 = lean_mk_string_unchecked("PartialResultParams", 19, 19); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____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_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__1; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__2; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__1; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__2; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____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_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__2; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__2; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__3; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__5() { _start: { lean_object* x_1; @@ -24567,53 +24192,53 @@ x_1 = lean_mk_string_unchecked("partialResultToken\?", 19, 19); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____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_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__5; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__5; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__7() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__6; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__6; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__8; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__8; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6983____closed__1; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____spec__1(x_1, x_2); +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6798____closed__1; +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -24622,7 +24247,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_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__9; +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__9; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -24634,7 +24259,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_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__9; +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__9; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -24667,7 +24292,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonPartialResultParams___closed__1 _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826_), 1, 0); return x_1; } } @@ -24679,7 +24304,7 @@ x_1 = l_Lean_Lsp_instFromJsonPartialResultParams___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7090____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6905____closed__1() { _start: { lean_object* x_1; @@ -24687,13 +24312,13 @@ x_1 = lean_mk_string_unchecked("workDoneProgress", 16, 16); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7090_(uint8_t x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6905_(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; lean_object* x_10; x_2 = lean_alloc_ctor(1, 0, 1); lean_ctor_set_uint8(x_2, 0, x_1); -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7090____closed__1; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6905____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); @@ -24704,19 +24329,19 @@ 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_5); -x_8 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_9 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_7, x_8); +x_8 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3; +x_9 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_7, x_8); x_10 = l_Lean_Json_mkObj(x_9); return x_10; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7090____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6905____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_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7090_(x_2); +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6905_(x_2); return x_3; } } @@ -24724,7 +24349,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonWorkDoneProgressOptions___closed_ _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7090____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6905____boxed), 1, 0); return x_1; } } @@ -24736,7 +24361,7 @@ x_1 = l_Lean_Lsp_instToJsonWorkDoneProgressOptions___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__1() { _start: { lean_object* x_1; @@ -24744,85 +24369,85 @@ x_1 = lean_mk_string_unchecked("WorkDoneProgressOptions", 23, 23); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____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_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__1; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__2; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__1; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__2; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____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_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__2; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__2; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__3; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____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_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7090____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6905____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_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__6() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__5; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__5; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__6; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__6; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__7; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__7; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7090____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_2853____spec__1(x_1, x_2); +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6905____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_2668____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -24831,7 +24456,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_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__8; +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__8; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -24843,7 +24468,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_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__8; +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__8; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -24876,7 +24501,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonWorkDoneProgressOptions___close _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943_), 1, 0); return x_1; } } @@ -24889,7 +24514,6 @@ return x_1; } } lean_object* initialize_Lean_Data_Json(uint8_t builtin, lean_object*); -lean_object* initialize_Lean_Data_JsonRpc(uint8_t builtin, lean_object*); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Lean_Data_Lsp_Basic(uint8_t builtin, lean_object* w) { lean_object * res; @@ -24898,63 +24522,6 @@ _G_initialized = true; res = initialize_Lean_Data_Json(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -res = initialize_Lean_Data_JsonRpc(builtin, lean_io_mk_world()); -if (lean_io_result_is_error(res)) return res; -lean_dec_ref(res); -l_Lean_Lsp_instInhabitedCancelParams___closed__1 = _init_l_Lean_Lsp_instInhabitedCancelParams___closed__1(); -lean_mark_persistent(l_Lean_Lsp_instInhabitedCancelParams___closed__1); -l_Lean_Lsp_instInhabitedCancelParams___closed__2 = _init_l_Lean_Lsp_instInhabitedCancelParams___closed__2(); -lean_mark_persistent(l_Lean_Lsp_instInhabitedCancelParams___closed__2); -l_Lean_Lsp_instInhabitedCancelParams = _init_l_Lean_Lsp_instInhabitedCancelParams(); -lean_mark_persistent(l_Lean_Lsp_instInhabitedCancelParams); -l_Lean_Lsp_instBEqCancelParams___closed__1 = _init_l_Lean_Lsp_instBEqCancelParams___closed__1(); -lean_mark_persistent(l_Lean_Lsp_instBEqCancelParams___closed__1); -l_Lean_Lsp_instBEqCancelParams = _init_l_Lean_Lsp_instBEqCancelParams(); -lean_mark_persistent(l_Lean_Lsp_instBEqCancelParams); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__3); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__4); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__5); -l_Lean_Lsp_instToJsonCancelParams___closed__1 = _init_l_Lean_Lsp_instToJsonCancelParams___closed__1(); -lean_mark_persistent(l_Lean_Lsp_instToJsonCancelParams___closed__1); -l_Lean_Lsp_instToJsonCancelParams = _init_l_Lean_Lsp_instToJsonCancelParams(); -lean_mark_persistent(l_Lean_Lsp_instToJsonCancelParams); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__3); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__4); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__6); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__7); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__8); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__9); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__10(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__10); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__11 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__11(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__11); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__13 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__13(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__13); -l_Lean_Lsp_instFromJsonCancelParams___closed__1 = _init_l_Lean_Lsp_instFromJsonCancelParams___closed__1(); -lean_mark_persistent(l_Lean_Lsp_instFromJsonCancelParams___closed__1); -l_Lean_Lsp_instFromJsonCancelParams = _init_l_Lean_Lsp_instFromJsonCancelParams(); -lean_mark_persistent(l_Lean_Lsp_instFromJsonCancelParams); l_Lean_Lsp_instInhabitedPosition___closed__1 = _init_l_Lean_Lsp_instInhabitedPosition___closed__1(); lean_mark_persistent(l_Lean_Lsp_instInhabitedPosition___closed__1); l_Lean_Lsp_instInhabitedPosition = _init_l_Lean_Lsp_instInhabitedPosition(); @@ -24971,38 +24538,50 @@ l_Lean_Lsp_instHashablePosition___closed__1 = _init_l_Lean_Lsp_instHashablePosit lean_mark_persistent(l_Lean_Lsp_instHashablePosition___closed__1); l_Lean_Lsp_instHashablePosition = _init_l_Lean_Lsp_instHashablePosition(); lean_mark_persistent(l_Lean_Lsp_instHashablePosition); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3); l_Lean_Lsp_instToJsonPosition___closed__1 = _init_l_Lean_Lsp_instToJsonPosition___closed__1(); lean_mark_persistent(l_Lean_Lsp_instToJsonPosition___closed__1); l_Lean_Lsp_instToJsonPosition = _init_l_Lean_Lsp_instToJsonPosition(); lean_mark_persistent(l_Lean_Lsp_instToJsonPosition); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__3); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__4); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__5); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__6); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__7); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__8); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__9); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__10(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__10); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__11 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__11(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__11); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__12 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__12(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__12); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__3); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__4); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__6); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__7); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__8); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__9); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__10(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__10); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__11 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__11(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__11); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__13 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__13(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__13); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__14 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__14(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__14); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__15 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__15(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__15); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__16 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__16(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__16); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__17 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__17(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__17); l_Lean_Lsp_instFromJsonPosition___closed__1 = _init_l_Lean_Lsp_instFromJsonPosition___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonPosition___closed__1); l_Lean_Lsp_instFromJsonPosition = _init_l_Lean_Lsp_instFromJsonPosition(); @@ -25029,38 +24608,38 @@ l_Lean_Lsp_instHashableRange___closed__1 = _init_l_Lean_Lsp_instHashableRange___ lean_mark_persistent(l_Lean_Lsp_instHashableRange___closed__1); l_Lean_Lsp_instHashableRange = _init_l_Lean_Lsp_instHashableRange(); lean_mark_persistent(l_Lean_Lsp_instHashableRange); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557____closed__2); l_Lean_Lsp_instToJsonRange___closed__1 = _init_l_Lean_Lsp_instToJsonRange___closed__1(); lean_mark_persistent(l_Lean_Lsp_instToJsonRange___closed__1); l_Lean_Lsp_instToJsonRange = _init_l_Lean_Lsp_instToJsonRange(); lean_mark_persistent(l_Lean_Lsp_instToJsonRange); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__3); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__4); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__5); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__6); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__7); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__8); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__9); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__10(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__10); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__11 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__11(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__11); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__12 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__12(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__12); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__3); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__4); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__5); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__6); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__7); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__8); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__9); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__10(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__10); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__11 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__11(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__11); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__12 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__12(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__12); l_Lean_Lsp_instFromJsonRange___closed__1 = _init_l_Lean_Lsp_instFromJsonRange___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonRange___closed__1); l_Lean_Lsp_instFromJsonRange = _init_l_Lean_Lsp_instFromJsonRange(); @@ -25075,44 +24654,46 @@ l_Lean_Lsp_instLERange = _init_l_Lean_Lsp_instLERange(); lean_mark_persistent(l_Lean_Lsp_instLERange); l_Lean_Lsp_instInhabitedLocation___closed__1 = _init_l_Lean_Lsp_instInhabitedLocation___closed__1(); lean_mark_persistent(l_Lean_Lsp_instInhabitedLocation___closed__1); +l_Lean_Lsp_instInhabitedLocation___closed__2 = _init_l_Lean_Lsp_instInhabitedLocation___closed__2(); +lean_mark_persistent(l_Lean_Lsp_instInhabitedLocation___closed__2); l_Lean_Lsp_instInhabitedLocation = _init_l_Lean_Lsp_instInhabitedLocation(); lean_mark_persistent(l_Lean_Lsp_instInhabitedLocation); l_Lean_Lsp_instBEqLocation___closed__1 = _init_l_Lean_Lsp_instBEqLocation___closed__1(); lean_mark_persistent(l_Lean_Lsp_instBEqLocation___closed__1); l_Lean_Lsp_instBEqLocation = _init_l_Lean_Lsp_instBEqLocation(); lean_mark_persistent(l_Lean_Lsp_instBEqLocation); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895____closed__2); l_Lean_Lsp_instToJsonLocation___closed__1 = _init_l_Lean_Lsp_instToJsonLocation___closed__1(); lean_mark_persistent(l_Lean_Lsp_instToJsonLocation___closed__1); l_Lean_Lsp_instToJsonLocation = _init_l_Lean_Lsp_instToJsonLocation(); lean_mark_persistent(l_Lean_Lsp_instToJsonLocation); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__3); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__4); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__5); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__6); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__7); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__8); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__9); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__10(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__10); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__11 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__11(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__11); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__12 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__12(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__12); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__3); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__4); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__5); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__6); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__7); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__8); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__9); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__10(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__10); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__11 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__11(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__11); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__12 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__12(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__12); l_Lean_Lsp_instFromJsonLocation___closed__1 = _init_l_Lean_Lsp_instFromJsonLocation___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonLocation___closed__1); l_Lean_Lsp_instFromJsonLocation = _init_l_Lean_Lsp_instFromJsonLocation(); @@ -25121,194 +24702,194 @@ l_Lean_Lsp_instOrdLocation___closed__1 = _init_l_Lean_Lsp_instOrdLocation___clos lean_mark_persistent(l_Lean_Lsp_instOrdLocation___closed__1); l_Lean_Lsp_instOrdLocation = _init_l_Lean_Lsp_instOrdLocation(); lean_mark_persistent(l_Lean_Lsp_instOrdLocation); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____closed__3); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____closed__4); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155____closed__3); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155____closed__4); l_Lean_Lsp_instToJsonLocationLink___closed__1 = _init_l_Lean_Lsp_instToJsonLocationLink___closed__1(); lean_mark_persistent(l_Lean_Lsp_instToJsonLocationLink___closed__1); l_Lean_Lsp_instToJsonLocationLink = _init_l_Lean_Lsp_instToJsonLocationLink(); lean_mark_persistent(l_Lean_Lsp_instToJsonLocationLink); -l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____spec__1___closed__1 = _init_l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____spec__1___closed__1(); -lean_mark_persistent(l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____spec__1___closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__3); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__4); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__5); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__6); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__7); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__8); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__9); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__10(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__10); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__11 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__11(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__11); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__12 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__12(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__12); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__13 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__13(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__13); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__14 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__14(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__14); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__15 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__15(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__15); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__16 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__16(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__16); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__17 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__17(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__17); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__18 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__18(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__18); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__19 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__19(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__19); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__20 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__20(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__20); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__21 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__21(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__21); +l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____spec__1___closed__1 = _init_l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____spec__1___closed__1(); +lean_mark_persistent(l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____spec__1___closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__3); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__4); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__5); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__6); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__7); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__8); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__9); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__10(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__10); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__11 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__11(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__11); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__12 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__12(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__12); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__13 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__13(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__13); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__14 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__14(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__14); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__15 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__15(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__15); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__16 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__16(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__16); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__17 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__17(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__17); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__18 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__18(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__18); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__19 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__19(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__19); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__20 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__20(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__20); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__21 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__21(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__21); l_Lean_Lsp_instFromJsonLocationLink___closed__1 = _init_l_Lean_Lsp_instFromJsonLocationLink___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonLocationLink___closed__1); l_Lean_Lsp_instFromJsonLocationLink = _init_l_Lean_Lsp_instFromJsonLocationLink(); lean_mark_persistent(l_Lean_Lsp_instFromJsonLocationLink); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____closed__3); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____closed__3); l_Lean_Lsp_instToJsonCommand___closed__1 = _init_l_Lean_Lsp_instToJsonCommand___closed__1(); lean_mark_persistent(l_Lean_Lsp_instToJsonCommand___closed__1); l_Lean_Lsp_instToJsonCommand = _init_l_Lean_Lsp_instToJsonCommand(); lean_mark_persistent(l_Lean_Lsp_instToJsonCommand); -l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1 = _init_l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1(); -lean_mark_persistent(l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1); -l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2 = _init_l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2(); -lean_mark_persistent(l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__3); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__4); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__5); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__6); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__7); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__8); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__9); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__10(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__10); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__11 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__11(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__11); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__12 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__12(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__12); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__13 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__13(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__13); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__14 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__14(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__14); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__15 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__15(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__15); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__16 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__16(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__16); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__17 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__17(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__17); +l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1 = _init_l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1(); +lean_mark_persistent(l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1); +l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2 = _init_l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2(); +lean_mark_persistent(l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__3); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__4); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__5); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__6); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__7); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__8); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__9); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__10(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__10); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__11 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__11(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__11); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__12 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__12(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__12); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__13 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__13(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__13); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__14 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__14(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__14); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__15 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__15(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__15); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__16 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__16(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__16); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__17 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__17(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__17); l_Lean_Lsp_instFromJsonCommand___closed__1 = _init_l_Lean_Lsp_instFromJsonCommand___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonCommand___closed__1); l_Lean_Lsp_instFromJsonCommand = _init_l_Lean_Lsp_instFromJsonCommand(); lean_mark_persistent(l_Lean_Lsp_instFromJsonCommand); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1840____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1840____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1840____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1655____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1655____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1655____closed__1); l_Lean_Lsp_instToJsonSnippetString___closed__1 = _init_l_Lean_Lsp_instToJsonSnippetString___closed__1(); lean_mark_persistent(l_Lean_Lsp_instToJsonSnippetString___closed__1); l_Lean_Lsp_instToJsonSnippetString = _init_l_Lean_Lsp_instToJsonSnippetString(); lean_mark_persistent(l_Lean_Lsp_instToJsonSnippetString); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__3); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__4); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__5); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__6); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__7); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__8); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__3); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__4); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__5); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__6); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__7); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__8); l_Lean_Lsp_instFromJsonSnippetString___closed__1 = _init_l_Lean_Lsp_instFromJsonSnippetString___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonSnippetString___closed__1); l_Lean_Lsp_instFromJsonSnippetString = _init_l_Lean_Lsp_instFromJsonSnippetString(); lean_mark_persistent(l_Lean_Lsp_instFromJsonSnippetString); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____closed__3); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____closed__3); l_Lean_Lsp_instToJsonTextEdit___closed__1 = _init_l_Lean_Lsp_instToJsonTextEdit___closed__1(); lean_mark_persistent(l_Lean_Lsp_instToJsonTextEdit___closed__1); l_Lean_Lsp_instToJsonTextEdit = _init_l_Lean_Lsp_instToJsonTextEdit(); lean_mark_persistent(l_Lean_Lsp_instToJsonTextEdit); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__3); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__4); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__5); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__6); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__7); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__8); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__9); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__10(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__10); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__11 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__11(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__11); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__12 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__12(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__12); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__13 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__13(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__13); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__14 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__14(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__14); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__15 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__15(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__15); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__16 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__16(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__16); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__17 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__17(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__17); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__18 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__18(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__18); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__19 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__19(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__19); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__20 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__20(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__20); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__3); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__4); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__5); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__6); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__7); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__8); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__9); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__10(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__10); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__11 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__11(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__11); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__12 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__12(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__12); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__13 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__13(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__13); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__14 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__14(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__14); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__15 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__15(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__15); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__16 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__16(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__16); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__17 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__17(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__17); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__18 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__18(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__18); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__19 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__19(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__19); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__20 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__20(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__20); l_Lean_Lsp_instFromJsonTextEdit___closed__1 = _init_l_Lean_Lsp_instFromJsonTextEdit___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonTextEdit___closed__1); l_Lean_Lsp_instFromJsonTextEdit = _init_l_Lean_Lsp_instFromJsonTextEdit(); @@ -25323,286 +24904,286 @@ l_Lean_Lsp_instToJsonTextDocumentIdentifier___closed__1 = _init_l_Lean_Lsp_instT lean_mark_persistent(l_Lean_Lsp_instToJsonTextDocumentIdentifier___closed__1); l_Lean_Lsp_instToJsonTextDocumentIdentifier = _init_l_Lean_Lsp_instToJsonTextDocumentIdentifier(); lean_mark_persistent(l_Lean_Lsp_instToJsonTextDocumentIdentifier); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351____closed__3); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351____closed__4); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351____closed__5); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351____closed__6); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166____closed__3); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166____closed__4); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166____closed__5); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166____closed__6); l_Lean_Lsp_instFromJsonTextDocumentIdentifier___closed__1 = _init_l_Lean_Lsp_instFromJsonTextDocumentIdentifier___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonTextDocumentIdentifier___closed__1); l_Lean_Lsp_instFromJsonTextDocumentIdentifier = _init_l_Lean_Lsp_instFromJsonTextDocumentIdentifier(); lean_mark_persistent(l_Lean_Lsp_instFromJsonTextDocumentIdentifier); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2440____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2440____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2440____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2255____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2255____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2255____closed__1); l_Lean_Lsp_instToJsonVersionedTextDocumentIdentifier___closed__1 = _init_l_Lean_Lsp_instToJsonVersionedTextDocumentIdentifier___closed__1(); lean_mark_persistent(l_Lean_Lsp_instToJsonVersionedTextDocumentIdentifier___closed__1); l_Lean_Lsp_instToJsonVersionedTextDocumentIdentifier = _init_l_Lean_Lsp_instToJsonVersionedTextDocumentIdentifier(); lean_mark_persistent(l_Lean_Lsp_instToJsonVersionedTextDocumentIdentifier); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__3); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__4); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__5); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__6); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__7); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__8); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__9); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__10(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__10); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__11 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__11(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__11); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__3); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__4); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__5); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__6); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__7); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__8); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__9); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__10(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__10); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__11 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__11(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__11); l_Lean_Lsp_instFromJsonVersionedTextDocumentIdentifier___closed__1 = _init_l_Lean_Lsp_instFromJsonVersionedTextDocumentIdentifier___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonVersionedTextDocumentIdentifier___closed__1); l_Lean_Lsp_instFromJsonVersionedTextDocumentIdentifier = _init_l_Lean_Lsp_instFromJsonVersionedTextDocumentIdentifier(); lean_mark_persistent(l_Lean_Lsp_instFromJsonVersionedTextDocumentIdentifier); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2608____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2608____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2608____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2608____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2608____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2608____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2423____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2423____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2423____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2423____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2423____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2423____closed__2); l_Lean_Lsp_instToJsonTextDocumentEdit___closed__1 = _init_l_Lean_Lsp_instToJsonTextDocumentEdit___closed__1(); lean_mark_persistent(l_Lean_Lsp_instToJsonTextDocumentEdit___closed__1); l_Lean_Lsp_instToJsonTextDocumentEdit = _init_l_Lean_Lsp_instToJsonTextDocumentEdit(); lean_mark_persistent(l_Lean_Lsp_instToJsonTextDocumentEdit); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__3); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__4); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__5); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__6); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__7); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__8); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__9); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__10(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__10); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__11 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__11(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__11); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__12 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__12(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__12); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__3); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__4); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__5); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__6); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__7); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__8); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__9); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__10(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__10); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__11 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__11(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__11); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__12 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__12(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__12); l_Lean_Lsp_instFromJsonTextDocumentEdit___closed__1 = _init_l_Lean_Lsp_instFromJsonTextDocumentEdit___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonTextDocumentEdit___closed__1); l_Lean_Lsp_instFromJsonTextDocumentEdit = _init_l_Lean_Lsp_instFromJsonTextDocumentEdit(); lean_mark_persistent(l_Lean_Lsp_instFromJsonTextDocumentEdit); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2797____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2797____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2797____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2797____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2797____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2797____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2797____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2797____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2797____closed__3); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2612____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2612____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2612____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2612____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2612____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2612____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2612____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2612____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2612____closed__3); l_Lean_Lsp_instToJsonChangeAnnotation___closed__1 = _init_l_Lean_Lsp_instToJsonChangeAnnotation___closed__1(); lean_mark_persistent(l_Lean_Lsp_instToJsonChangeAnnotation___closed__1); l_Lean_Lsp_instToJsonChangeAnnotation = _init_l_Lean_Lsp_instToJsonChangeAnnotation(); lean_mark_persistent(l_Lean_Lsp_instToJsonChangeAnnotation); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__3); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__4); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__5); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__6); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__7); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__8); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__9); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__10(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__10); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__11 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__11(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__11); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__12 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__12(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__12); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__13 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__13(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__13); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__14 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__14(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__14); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__15 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__15(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__15); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__16 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__16(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__16); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__17 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__17(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__17); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__3); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__4); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__5); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__6); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__7); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__8); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__9); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__10(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__10); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__11 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__11(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__11); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__12 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__12(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__12); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__13 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__13(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__13); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__14 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__14(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__14); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__15 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__15(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__15); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__16 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__16(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__16); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__17 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__17(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__17); l_Lean_Lsp_instFromJsonChangeAnnotation___closed__1 = _init_l_Lean_Lsp_instFromJsonChangeAnnotation___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonChangeAnnotation___closed__1); l_Lean_Lsp_instFromJsonChangeAnnotation = _init_l_Lean_Lsp_instFromJsonChangeAnnotation(); lean_mark_persistent(l_Lean_Lsp_instFromJsonChangeAnnotation); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3024____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3024____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3024____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3024____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3024____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3024____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2839____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2839____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2839____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2839____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2839____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2839____closed__2); l_Lean_Lsp_CreateFile_instToJsonOptions___closed__1 = _init_l_Lean_Lsp_CreateFile_instToJsonOptions___closed__1(); lean_mark_persistent(l_Lean_Lsp_CreateFile_instToJsonOptions___closed__1); l_Lean_Lsp_CreateFile_instToJsonOptions = _init_l_Lean_Lsp_CreateFile_instToJsonOptions(); lean_mark_persistent(l_Lean_Lsp_CreateFile_instToJsonOptions); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__3); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__4); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__5); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__6); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__7); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__8); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__9); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__10(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__10); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__11 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__11(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__11); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__12 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__12(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__12); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__13 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__13(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__13); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__3); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__4); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__5); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__6); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__7); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__8); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__9); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__10(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__10); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__11 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__11(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__11); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__12 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__12(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__12); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__13 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__13(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__13); l_Lean_Lsp_CreateFile_instFromJsonOptions___closed__1 = _init_l_Lean_Lsp_CreateFile_instFromJsonOptions___closed__1(); lean_mark_persistent(l_Lean_Lsp_CreateFile_instFromJsonOptions___closed__1); l_Lean_Lsp_CreateFile_instFromJsonOptions = _init_l_Lean_Lsp_CreateFile_instFromJsonOptions(); lean_mark_persistent(l_Lean_Lsp_CreateFile_instFromJsonOptions); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3208____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3208____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3208____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3208____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3208____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3208____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3023____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3023____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3023____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3023____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3023____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3023____closed__2); l_Lean_Lsp_DeleteFile_instToJsonOptions___closed__1 = _init_l_Lean_Lsp_DeleteFile_instToJsonOptions___closed__1(); lean_mark_persistent(l_Lean_Lsp_DeleteFile_instToJsonOptions___closed__1); l_Lean_Lsp_DeleteFile_instToJsonOptions = _init_l_Lean_Lsp_DeleteFile_instToJsonOptions(); lean_mark_persistent(l_Lean_Lsp_DeleteFile_instToJsonOptions); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__3); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__4); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__5); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__6); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__7); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__8); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__9); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__10(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__10); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__11 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__11(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__11); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__12 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__12(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__12); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__3); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__4); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__5); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__6); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__7); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__8); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__9); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__10(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__10); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__11 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__11(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__11); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__12 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__12(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__12); l_Lean_Lsp_DeleteFile_instFromJsonOptions___closed__1 = _init_l_Lean_Lsp_DeleteFile_instFromJsonOptions___closed__1(); lean_mark_persistent(l_Lean_Lsp_DeleteFile_instFromJsonOptions___closed__1); l_Lean_Lsp_DeleteFile_instFromJsonOptions = _init_l_Lean_Lsp_DeleteFile_instFromJsonOptions(); lean_mark_persistent(l_Lean_Lsp_DeleteFile_instFromJsonOptions); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3399____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3399____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3399____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3214____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3214____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3214____closed__1); l_Lean_Lsp_instToJsonCreateFile___closed__1 = _init_l_Lean_Lsp_instToJsonCreateFile___closed__1(); lean_mark_persistent(l_Lean_Lsp_instToJsonCreateFile___closed__1); l_Lean_Lsp_instToJsonCreateFile = _init_l_Lean_Lsp_instToJsonCreateFile(); lean_mark_persistent(l_Lean_Lsp_instToJsonCreateFile); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__3); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__4); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__5); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__6); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__7); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__8); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__9); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__10(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__10); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__11 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__11(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__11); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__12 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__12(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__12); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__3); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__4); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__5); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__6); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__7); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__8); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__9); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__10(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__10); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__11 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__11(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__11); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__12 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__12(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__12); l_Lean_Lsp_instFromJsonCreateFile___closed__1 = _init_l_Lean_Lsp_instFromJsonCreateFile___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonCreateFile___closed__1); l_Lean_Lsp_instFromJsonCreateFile = _init_l_Lean_Lsp_instFromJsonCreateFile(); lean_mark_persistent(l_Lean_Lsp_instFromJsonCreateFile); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3630____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3630____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3630____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3630____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3630____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3630____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__2); l_Lean_Lsp_instToJsonRenameFile___closed__1 = _init_l_Lean_Lsp_instToJsonRenameFile___closed__1(); lean_mark_persistent(l_Lean_Lsp_instToJsonRenameFile___closed__1); l_Lean_Lsp_instToJsonRenameFile = _init_l_Lean_Lsp_instToJsonRenameFile(); lean_mark_persistent(l_Lean_Lsp_instToJsonRenameFile); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__3); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__4); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__5); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__6); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__7); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__8); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__9); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__10(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__10); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__11 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__11(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__11); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__12 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__12(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__12); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__13 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__13(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__13); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__14 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__14(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__14); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__15 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__15(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__15); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__16 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__16(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__16); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__3); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__4); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__5); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__6); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__7); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__8); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__9); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__10(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__10); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__11 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__11(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__11); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__12 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__12(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__12); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__13 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__13(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__13); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__14 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__14(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__14); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__15 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__15(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__15); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__16 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__16(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__16); l_Lean_Lsp_instFromJsonRenameFile___closed__1 = _init_l_Lean_Lsp_instFromJsonRenameFile___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonRenameFile___closed__1); l_Lean_Lsp_instFromJsonRenameFile = _init_l_Lean_Lsp_instFromJsonRenameFile(); @@ -25611,24 +25192,24 @@ l_Lean_Lsp_instToJsonDeleteFile___closed__1 = _init_l_Lean_Lsp_instToJsonDeleteF lean_mark_persistent(l_Lean_Lsp_instToJsonDeleteFile___closed__1); l_Lean_Lsp_instToJsonDeleteFile = _init_l_Lean_Lsp_instToJsonDeleteFile(); lean_mark_persistent(l_Lean_Lsp_instToJsonDeleteFile); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__3); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__4); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__5); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__6); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__7); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__8); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__9); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__3); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__4); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__5); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__6); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__7); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__8); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__9); l_Lean_Lsp_instFromJsonDeleteFile___closed__1 = _init_l_Lean_Lsp_instFromJsonDeleteFile___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonDeleteFile___closed__1); l_Lean_Lsp_instFromJsonDeleteFile = _init_l_Lean_Lsp_instFromJsonDeleteFile(); @@ -25647,54 +25228,54 @@ l_Lean_Lsp_instToJsonDocumentChange___closed__6 = _init_l_Lean_Lsp_instToJsonDoc lean_mark_persistent(l_Lean_Lsp_instToJsonDocumentChange___closed__6); l_Lean_Lsp_instToJsonDocumentChange___closed__7 = _init_l_Lean_Lsp_instToJsonDocumentChange___closed__7(); lean_mark_persistent(l_Lean_Lsp_instToJsonDocumentChange___closed__7); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____closed__3); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____closed__3); l_Lean_Lsp_instToJsonWorkspaceEdit___closed__1 = _init_l_Lean_Lsp_instToJsonWorkspaceEdit___closed__1(); lean_mark_persistent(l_Lean_Lsp_instToJsonWorkspaceEdit___closed__1); l_Lean_Lsp_instToJsonWorkspaceEdit = _init_l_Lean_Lsp_instToJsonWorkspaceEdit(); lean_mark_persistent(l_Lean_Lsp_instToJsonWorkspaceEdit); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__3); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__4); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__5); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__6); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__7); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__8); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__9); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__10(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__10); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__11 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__11(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__11); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__12 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__12(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__12); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__13 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__13(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__13); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__14 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__14(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__14); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__15 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__15(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__15); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__16 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__16(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__16); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__17 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__17(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__17); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__18 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__18(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__18); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__19 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__19(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__19); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__3); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__4); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__5); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__6); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__7); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__8); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__9); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__10(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__10); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__11 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__11(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__11); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__12 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__12(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__12); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__13 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__13(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__13); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__14 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__14(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__14); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__15 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__15(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__15); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__16 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__16(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__16); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__17 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__17(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__17); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__18 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__18(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__18); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__19 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__19(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__19); l_Lean_Lsp_instFromJsonWorkspaceEdit___closed__1 = _init_l_Lean_Lsp_instFromJsonWorkspaceEdit___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonWorkspaceEdit___closed__1); l_Lean_Lsp_instFromJsonWorkspaceEdit = _init_l_Lean_Lsp_instFromJsonWorkspaceEdit(); @@ -25709,224 +25290,226 @@ l_Lean_Lsp_WorkspaceEdit_instAppend___closed__2 = _init_l_Lean_Lsp_WorkspaceEdit lean_mark_persistent(l_Lean_Lsp_WorkspaceEdit_instAppend___closed__2); l_Lean_Lsp_WorkspaceEdit_instAppend___closed__3 = _init_l_Lean_Lsp_WorkspaceEdit_instAppend___closed__3(); lean_mark_persistent(l_Lean_Lsp_WorkspaceEdit_instAppend___closed__3); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4806____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4806____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4806____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4621____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4621____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4621____closed__1); l_Lean_Lsp_instToJsonApplyWorkspaceEditParams___closed__1 = _init_l_Lean_Lsp_instToJsonApplyWorkspaceEditParams___closed__1(); lean_mark_persistent(l_Lean_Lsp_instToJsonApplyWorkspaceEditParams___closed__1); l_Lean_Lsp_instToJsonApplyWorkspaceEditParams = _init_l_Lean_Lsp_instToJsonApplyWorkspaceEditParams(); lean_mark_persistent(l_Lean_Lsp_instToJsonApplyWorkspaceEditParams); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__3); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__4); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__5); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__6); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__7); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__8); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__9); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__10(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__10); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__11 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__11(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__11); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__12 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__12(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__12); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__13 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__13(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__13); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__3); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__4); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__5); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__6); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__7); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__8); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__9); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__10(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__10); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__11 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__11(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__11); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__12 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__12(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__12); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__13 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__13(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__13); l_Lean_Lsp_instFromJsonApplyWorkspaceEditParams___closed__1 = _init_l_Lean_Lsp_instFromJsonApplyWorkspaceEditParams___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonApplyWorkspaceEditParams___closed__1); l_Lean_Lsp_instFromJsonApplyWorkspaceEditParams = _init_l_Lean_Lsp_instFromJsonApplyWorkspaceEditParams(); lean_mark_persistent(l_Lean_Lsp_instFromJsonApplyWorkspaceEditParams); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4990____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4990____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4990____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4990____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4990____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4990____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4805____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4805____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4805____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4805____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4805____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4805____closed__2); l_Lean_Lsp_instToJsonTextDocumentItem___closed__1 = _init_l_Lean_Lsp_instToJsonTextDocumentItem___closed__1(); lean_mark_persistent(l_Lean_Lsp_instToJsonTextDocumentItem___closed__1); l_Lean_Lsp_instToJsonTextDocumentItem = _init_l_Lean_Lsp_instToJsonTextDocumentItem(); lean_mark_persistent(l_Lean_Lsp_instToJsonTextDocumentItem); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__3); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__4); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__5); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__6); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__7); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__8); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__9); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__10(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__10); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__11 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__11(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__11); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__12 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__12(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__12); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__13 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__13(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__13); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__14 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__14(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__14); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__15 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__15(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__15); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__16 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__16(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__16); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__17 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__17(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__17); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__18 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__18(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__18); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__3); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__4); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__5); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__6); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__7); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__8); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__9); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__10(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__10); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__11 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__11(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__11); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__12 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__12(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__12); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__13 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__13(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__13); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__14 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__14(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__14); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__15 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__15(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__15); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__16 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__16(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__16); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__17 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__17(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__17); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__18 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__18(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__18); l_Lean_Lsp_instFromJsonTextDocumentItem___closed__1 = _init_l_Lean_Lsp_instFromJsonTextDocumentItem___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonTextDocumentItem___closed__1); l_Lean_Lsp_instFromJsonTextDocumentItem = _init_l_Lean_Lsp_instFromJsonTextDocumentItem(); lean_mark_persistent(l_Lean_Lsp_instFromJsonTextDocumentItem); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5274____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5274____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5274____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5089____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5089____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5089____closed__1); l_Lean_Lsp_instToJsonTextDocumentPositionParams___closed__1 = _init_l_Lean_Lsp_instToJsonTextDocumentPositionParams___closed__1(); lean_mark_persistent(l_Lean_Lsp_instToJsonTextDocumentPositionParams___closed__1); l_Lean_Lsp_instToJsonTextDocumentPositionParams = _init_l_Lean_Lsp_instToJsonTextDocumentPositionParams(); lean_mark_persistent(l_Lean_Lsp_instToJsonTextDocumentPositionParams); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__3); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__4); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__5); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__6); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__7); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__8); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__9); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__10(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__10); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__3); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__4); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__5); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__6); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__7); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__8); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__9); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__10(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__10); l_Lean_Lsp_instFromJsonTextDocumentPositionParams___closed__1 = _init_l_Lean_Lsp_instFromJsonTextDocumentPositionParams___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonTextDocumentPositionParams___closed__1); l_Lean_Lsp_instFromJsonTextDocumentPositionParams = _init_l_Lean_Lsp_instFromJsonTextDocumentPositionParams(); lean_mark_persistent(l_Lean_Lsp_instFromJsonTextDocumentPositionParams); l_Lean_Lsp_instToStringTextDocumentPositionParams___closed__1 = _init_l_Lean_Lsp_instToStringTextDocumentPositionParams___closed__1(); lean_mark_persistent(l_Lean_Lsp_instToStringTextDocumentPositionParams___closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5505____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5505____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5505____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5505____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5505____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5505____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5505____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5505____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5505____closed__3); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5320____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5320____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5320____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5320____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5320____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5320____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5320____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5320____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5320____closed__3); l_Lean_Lsp_instToJsonDocumentFilter___closed__1 = _init_l_Lean_Lsp_instToJsonDocumentFilter___closed__1(); lean_mark_persistent(l_Lean_Lsp_instToJsonDocumentFilter___closed__1); l_Lean_Lsp_instToJsonDocumentFilter = _init_l_Lean_Lsp_instToJsonDocumentFilter(); lean_mark_persistent(l_Lean_Lsp_instToJsonDocumentFilter); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__3); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__4); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__5); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__6); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__7); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__8); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__9); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__10(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__10); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__11 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__11(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__11); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__12 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__12(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__12); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__13 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__13(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__13); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__14 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__14(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__14); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__15 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__15(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__15); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__16 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__16(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__16); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__17 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__17(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__17); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__18 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__18(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__18); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__19 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__19(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__19); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__3); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__4); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__5); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__6); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__7); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__8); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__9); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__10(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__10); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__11 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__11(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__11); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__12 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__12(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__12); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__13 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__13(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__13); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__14 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__14(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__14); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__15 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__15(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__15); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__16 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__16(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__16); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__17 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__17(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__17); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__18 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__18(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__18); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__19 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__19(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__19); l_Lean_Lsp_instFromJsonDocumentFilter___closed__1 = _init_l_Lean_Lsp_instFromJsonDocumentFilter___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonDocumentFilter___closed__1); l_Lean_Lsp_instFromJsonDocumentFilter = _init_l_Lean_Lsp_instFromJsonDocumentFilter(); lean_mark_persistent(l_Lean_Lsp_instFromJsonDocumentFilter); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5542____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5542____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5542____closed__1); l_Lean_Lsp_instToJsonStaticRegistrationOptions___closed__1 = _init_l_Lean_Lsp_instToJsonStaticRegistrationOptions___closed__1(); lean_mark_persistent(l_Lean_Lsp_instToJsonStaticRegistrationOptions___closed__1); l_Lean_Lsp_instToJsonStaticRegistrationOptions = _init_l_Lean_Lsp_instToJsonStaticRegistrationOptions(); lean_mark_persistent(l_Lean_Lsp_instToJsonStaticRegistrationOptions); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__3); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__4); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__5); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__6); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__7); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__8); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__9); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__3); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__4); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__5); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__6); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__7); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__8); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__9); l_Lean_Lsp_instFromJsonStaticRegistrationOptions___closed__1 = _init_l_Lean_Lsp_instFromJsonStaticRegistrationOptions___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonStaticRegistrationOptions___closed__1); l_Lean_Lsp_instFromJsonStaticRegistrationOptions = _init_l_Lean_Lsp_instFromJsonStaticRegistrationOptions(); lean_mark_persistent(l_Lean_Lsp_instFromJsonStaticRegistrationOptions); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5836____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5836____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5836____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5651____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5651____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5651____closed__1); l_Lean_Lsp_instToJsonTextDocumentRegistrationOptions___closed__1 = _init_l_Lean_Lsp_instToJsonTextDocumentRegistrationOptions___closed__1(); lean_mark_persistent(l_Lean_Lsp_instToJsonTextDocumentRegistrationOptions___closed__1); l_Lean_Lsp_instToJsonTextDocumentRegistrationOptions = _init_l_Lean_Lsp_instToJsonTextDocumentRegistrationOptions(); lean_mark_persistent(l_Lean_Lsp_instToJsonTextDocumentRegistrationOptions); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__3); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__4); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__5); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__6); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__7); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__8); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__9); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__3); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__4); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__5); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__6); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__7); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__8); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__9); l_Lean_Lsp_instFromJsonTextDocumentRegistrationOptions___closed__1 = _init_l_Lean_Lsp_instFromJsonTextDocumentRegistrationOptions___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonTextDocumentRegistrationOptions___closed__1); l_Lean_Lsp_instFromJsonTextDocumentRegistrationOptions = _init_l_Lean_Lsp_instFromJsonTextDocumentRegistrationOptions(); @@ -25953,38 +25536,38 @@ l_Lean_Lsp_instToJsonMarkupKind___closed__1 = _init_l_Lean_Lsp_instToJsonMarkupK lean_mark_persistent(l_Lean_Lsp_instToJsonMarkupKind___closed__1); l_Lean_Lsp_instToJsonMarkupKind___closed__2 = _init_l_Lean_Lsp_instToJsonMarkupKind___closed__2(); lean_mark_persistent(l_Lean_Lsp_instToJsonMarkupKind___closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6159____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6159____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6159____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6159____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6159____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6159____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6159____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6159____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6159____closed__3); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6159____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6159____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6159____closed__4); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_5974____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_5974____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_5974____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_5974____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_5974____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_5974____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_5974____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_5974____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_5974____closed__3); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_5974____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_5974____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_5974____closed__4); l_Lean_Lsp_instToJsonMarkupContent___closed__1 = _init_l_Lean_Lsp_instToJsonMarkupContent___closed__1(); lean_mark_persistent(l_Lean_Lsp_instToJsonMarkupContent___closed__1); l_Lean_Lsp_instToJsonMarkupContent = _init_l_Lean_Lsp_instToJsonMarkupContent(); lean_mark_persistent(l_Lean_Lsp_instToJsonMarkupContent); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__3); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__4); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__5); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__6); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__7); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__8); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__9); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__10(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__10); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__3); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__4); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__5); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__6); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__7); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__8); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__9); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__10(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__10); l_Lean_Lsp_instFromJsonMarkupContent___closed__1 = _init_l_Lean_Lsp_instFromJsonMarkupContent___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonMarkupContent___closed__1); l_Lean_Lsp_instFromJsonMarkupContent = _init_l_Lean_Lsp_instFromJsonMarkupContent(); @@ -25993,14 +25576,14 @@ l_Lean_Lsp_instHashableMarkupContent___closed__1 = _init_l_Lean_Lsp_instHashable lean_mark_persistent(l_Lean_Lsp_instHashableMarkupContent___closed__1); l_Lean_Lsp_instHashableMarkupContent = _init_l_Lean_Lsp_instHashableMarkupContent(); lean_mark_persistent(l_Lean_Lsp_instHashableMarkupContent); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6531____rarg___closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6531____rarg___closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6531____rarg___closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6641____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6641____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6641____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6641____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6641____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6641____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6641____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6641____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6641____closed__3); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6346____rarg___closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6346____rarg___closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6346____rarg___closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6456____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6456____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6456____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6456____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6456____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6456____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6456____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6456____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6456____closed__3); l_Lean_Lsp_instToJsonWorkDoneProgressReport___closed__1 = _init_l_Lean_Lsp_instToJsonWorkDoneProgressReport___closed__1(); lean_mark_persistent(l_Lean_Lsp_instToJsonWorkDoneProgressReport___closed__1); l_Lean_Lsp_instToJsonWorkDoneProgressReport = _init_l_Lean_Lsp_instToJsonWorkDoneProgressReport(); @@ -26013,84 +25596,84 @@ l_Lean_Lsp_instToJsonWorkDoneProgressEnd___closed__1 = _init_l_Lean_Lsp_instToJs lean_mark_persistent(l_Lean_Lsp_instToJsonWorkDoneProgressEnd___closed__1); l_Lean_Lsp_instToJsonWorkDoneProgressEnd = _init_l_Lean_Lsp_instToJsonWorkDoneProgressEnd(); lean_mark_persistent(l_Lean_Lsp_instToJsonWorkDoneProgressEnd); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6874____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6874____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6874____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6689____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6689____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6689____closed__1); l_Lean_Lsp_instToJsonWorkDoneProgressParams___closed__1 = _init_l_Lean_Lsp_instToJsonWorkDoneProgressParams___closed__1(); lean_mark_persistent(l_Lean_Lsp_instToJsonWorkDoneProgressParams___closed__1); l_Lean_Lsp_instToJsonWorkDoneProgressParams = _init_l_Lean_Lsp_instToJsonWorkDoneProgressParams(); lean_mark_persistent(l_Lean_Lsp_instToJsonWorkDoneProgressParams); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__3); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__4); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__5); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__6); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__7); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__8); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__9); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__3); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__4); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__5); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__6); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__7); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__8); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__9); l_Lean_Lsp_instFromJsonWorkDoneProgressParams___closed__1 = _init_l_Lean_Lsp_instFromJsonWorkDoneProgressParams___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonWorkDoneProgressParams___closed__1); l_Lean_Lsp_instFromJsonWorkDoneProgressParams = _init_l_Lean_Lsp_instFromJsonWorkDoneProgressParams(); lean_mark_persistent(l_Lean_Lsp_instFromJsonWorkDoneProgressParams); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6983____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6983____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6983____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6798____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6798____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6798____closed__1); l_Lean_Lsp_instToJsonPartialResultParams___closed__1 = _init_l_Lean_Lsp_instToJsonPartialResultParams___closed__1(); lean_mark_persistent(l_Lean_Lsp_instToJsonPartialResultParams___closed__1); l_Lean_Lsp_instToJsonPartialResultParams = _init_l_Lean_Lsp_instToJsonPartialResultParams(); lean_mark_persistent(l_Lean_Lsp_instToJsonPartialResultParams); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__3); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__4); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__5); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__6); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__7); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__8); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__9); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__3); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__4); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__5); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__6); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__7); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__8); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__9); l_Lean_Lsp_instFromJsonPartialResultParams___closed__1 = _init_l_Lean_Lsp_instFromJsonPartialResultParams___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonPartialResultParams___closed__1); l_Lean_Lsp_instFromJsonPartialResultParams = _init_l_Lean_Lsp_instFromJsonPartialResultParams(); lean_mark_persistent(l_Lean_Lsp_instFromJsonPartialResultParams); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7090____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7090____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7090____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6905____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6905____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6905____closed__1); l_Lean_Lsp_instToJsonWorkDoneProgressOptions___closed__1 = _init_l_Lean_Lsp_instToJsonWorkDoneProgressOptions___closed__1(); lean_mark_persistent(l_Lean_Lsp_instToJsonWorkDoneProgressOptions___closed__1); l_Lean_Lsp_instToJsonWorkDoneProgressOptions = _init_l_Lean_Lsp_instToJsonWorkDoneProgressOptions(); lean_mark_persistent(l_Lean_Lsp_instToJsonWorkDoneProgressOptions); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__3); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__4); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__5); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__6); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__7); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__8); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__3); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__4); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__5); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__6); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__7); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__8); l_Lean_Lsp_instFromJsonWorkDoneProgressOptions___closed__1 = _init_l_Lean_Lsp_instFromJsonWorkDoneProgressOptions___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonWorkDoneProgressOptions___closed__1); l_Lean_Lsp_instFromJsonWorkDoneProgressOptions = _init_l_Lean_Lsp_instFromJsonWorkDoneProgressOptions(); diff --git a/stage0/stdlib/Lean/Data/Lsp/CancelParams.c b/stage0/stdlib/Lean/Data/Lsp/CancelParams.c new file mode 100644 index 000000000000..27bade501d73 --- /dev/null +++ b/stage0/stdlib/Lean/Data/Lsp/CancelParams.c @@ -0,0 +1,603 @@ +// Lean compiler output +// Module: Lean.Data.Lsp.CancelParams +// Imports: Lean.Data.JsonRpc +#include +#if defined(__clang__) +#pragma clang diagnostic ignored "-Wunused-parameter" +#pragma clang diagnostic ignored "-Wunused-label" +#elif defined(__GNUC__) && !defined(__CLANG__) +#pragma GCC diagnostic ignored "-Wunused-parameter" +#pragma GCC diagnostic ignored "-Wunused-label" +#pragma GCC diagnostic ignored "-Wunused-but-set-variable" +#endif +#ifdef __cplusplus +extern "C" { +#endif +LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonCancelParams; +static lean_object* l_Lean_Lsp_instToJsonCancelParams___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__7; +LEAN_EXPORT lean_object* l_Lean_Lsp_instBEqCancelParams; +lean_object* l_Lean_Json_mkObj(lean_object*); +LEAN_EXPORT lean_object* l_List_flatMapTR_go___at___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____spec__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__6; +static lean_object* l_Lean_Lsp_instInhabitedCancelParams___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_beqCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_28____boxed(lean_object*, lean_object*); +lean_object* l_Lean_Name_toString(lean_object*, uint8_t, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__12; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____lambda__1___boxed(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__13; +lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_beqCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_28_(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____lambda__1(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonCancelParams; +lean_object* lean_array_to_list(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____closed__4; +static lean_object* l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____closed__2; +static lean_object* l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____closed__1; +lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Lsp_instInhabitedCancelParams; +static lean_object* l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__8; +static lean_object* l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____closed__5; +static lean_object* l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__2; +static lean_object* l_Lean_Lsp_instBEqCancelParams___closed__1; +static lean_object* l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__9; +static lean_object* l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____closed__3; +lean_object* l_List_foldl___at_Array_appendList___spec__1___rarg(lean_object*, lean_object*); +static lean_object* l_Lean_Lsp_instInhabitedCancelParams___closed__2; +lean_object* lean_array_mk(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__11; +static lean_object* l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__4; +uint8_t l___private_Lean_Data_JsonRpc_0__Lean_JsonRpc_beqRequestID____x40_Lean_Data_JsonRpc___hyg_36_(lean_object*, lean_object*); +lean_object* lean_string_append(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__1; +lean_object* l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__5; +static lean_object* l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__3; +static lean_object* l_Lean_Lsp_instFromJsonCancelParams___closed__1; +static lean_object* l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__10; +static lean_object* _init_l_Lean_Lsp_instInhabitedCancelParams___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("", 0, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Lsp_instInhabitedCancelParams___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Lsp_instInhabitedCancelParams___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_Lsp_instInhabitedCancelParams() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Lsp_instInhabitedCancelParams___closed__2; +return x_1; +} +} +LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_beqCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_28_(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = l___private_Lean_Data_JsonRpc_0__Lean_JsonRpc_beqRequestID____x40_Lean_Data_JsonRpc___hyg_36_(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_beqCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_28____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_CancelParams_0__Lean_Lsp_beqCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_28_(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Lsp_instBEqCancelParams___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_beqCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_28____boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Lsp_instBEqCancelParams() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Lsp_instBEqCancelParams___closed__1; +return x_1; +} +} +LEAN_EXPORT lean_object* l_List_flatMapTR_go___at___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____spec__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_3; +x_3 = lean_array_to_list(x_2); +return x_3; +} +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 = l_List_foldl___at_Array_appendList___spec__1___rarg(x_2, x_4); +x_1 = x_5; +x_2 = x_6; +goto _start; +} +} +} +static lean_object* _init_l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(0); +x_2 = lean_array_mk(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("id", 2, 2); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____closed__2; +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; +} +} +static lean_object* _init_l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____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_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____closed__3; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____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_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____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; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86_(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_box(0); +switch (lean_obj_tag(x_1)) { +case 0: +{ +uint8_t x_3; +x_3 = !lean_is_exclusive(x_1); +if (x_3 == 0) +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +lean_ctor_set_tag(x_1, 3); +x_4 = l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____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_1); +x_6 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_6, 0, x_5); +lean_ctor_set(x_6, 1, x_2); +x_7 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_7, 0, x_6); +lean_ctor_set(x_7, 1, x_2); +x_8 = l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____closed__1; +x_9 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____spec__1(x_7, x_8); +x_10 = l_Lean_Json_mkObj(x_9); +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; +x_11 = lean_ctor_get(x_1, 0); +lean_inc(x_11); +lean_dec(x_1); +x_12 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_12, 0, x_11); +x_13 = l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____closed__2; +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_12); +x_15 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_15, 0, x_14); +lean_ctor_set(x_15, 1, x_2); +x_16 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_16, 0, x_15); +lean_ctor_set(x_16, 1, x_2); +x_17 = l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____closed__1; +x_18 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____spec__1(x_16, x_17); +x_19 = l_Lean_Json_mkObj(x_18); +return x_19; +} +} +case 1: +{ +uint8_t x_20; +x_20 = !lean_is_exclusive(x_1); +if (x_20 == 0) +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +lean_ctor_set_tag(x_1, 2); +x_21 = l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____closed__2; +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_1); +x_23 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_2); +x_24 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set(x_24, 1, x_2); +x_25 = l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____closed__1; +x_26 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____spec__1(x_24, x_25); +x_27 = l_Lean_Json_mkObj(x_26); +return x_27; +} +else +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_28 = lean_ctor_get(x_1, 0); +lean_inc(x_28); +lean_dec(x_1); +x_29 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_29, 0, x_28); +x_30 = l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____closed__2; +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 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_2); +x_33 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_2); +x_34 = l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____closed__1; +x_35 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____spec__1(x_33, x_34); +x_36 = l_Lean_Json_mkObj(x_35); +return x_36; +} +} +default: +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_37 = l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____closed__5; +x_38 = l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____closed__1; +x_39 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____spec__1(x_37, x_38); +x_40 = l_Lean_Json_mkObj(x_39); +return x_40; +} +} +} +} +static lean_object* _init_l_Lean_Lsp_instToJsonCancelParams___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86_), 1, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Lsp_instToJsonCancelParams() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Lsp_instToJsonCancelParams___closed__1; +return x_1; +} +} +LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____lambda__1(lean_object* x_1) { +_start: +{ +uint8_t x_2; +x_2 = 0; +return x_2; +} +} +static lean_object* _init_l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lsp", 3, 3); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("CancelParams", 12, 12); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__1; +x_2 = l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__2; +x_3 = l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__3; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____lambda__1___boxed), 1, 0); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__6() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; +x_1 = l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__4; +x_2 = 1; +x_3 = l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__5; +x_4 = l_Lean_Name_toString(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(".", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__6; +x_2 = l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__7; +x_3 = lean_string_append(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____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_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__10() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; +x_1 = l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__9; +x_2 = 1; +x_3 = l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__5; +x_4 = l_Lean_Name_toString(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__8; +x_2 = l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__10; +x_3 = lean_string_append(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__12() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(": ", 2, 2); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__11; +x_2 = l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__12; +x_3 = lean_string_append(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124_(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____closed__2; +x_3 = l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__1(x_1, x_2); +if (lean_obj_tag(x_3) == 0) +{ +uint8_t x_4; +x_4 = !lean_is_exclusive(x_3); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_5 = lean_ctor_get(x_3, 0); +x_6 = l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__13; +x_7 = lean_string_append(x_6, x_5); +lean_dec(x_5); +lean_ctor_set(x_3, 0, x_7); +return x_3; +} +else +{ +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_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__13; +x_10 = lean_string_append(x_9, x_8); +lean_dec(x_8); +x_11 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_11, 0, x_10); +return x_11; +} +} +else +{ +uint8_t x_12; +x_12 = !lean_is_exclusive(x_3); +if (x_12 == 0) +{ +return x_3; +} +else +{ +lean_object* x_13; lean_object* x_14; +x_13 = lean_ctor_get(x_3, 0); +lean_inc(x_13); +lean_dec(x_3); +x_14 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_14, 0, x_13); +return x_14; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____lambda__1___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____lambda__1(x_1); +lean_dec(x_1); +x_3 = lean_box(x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Lsp_instFromJsonCancelParams___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124_), 1, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Lsp_instFromJsonCancelParams() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Lsp_instFromJsonCancelParams___closed__1; +return x_1; +} +} +lean_object* initialize_Lean_Data_JsonRpc(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Lean_Data_Lsp_CancelParams(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_Data_JsonRpc(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l_Lean_Lsp_instInhabitedCancelParams___closed__1 = _init_l_Lean_Lsp_instInhabitedCancelParams___closed__1(); +lean_mark_persistent(l_Lean_Lsp_instInhabitedCancelParams___closed__1); +l_Lean_Lsp_instInhabitedCancelParams___closed__2 = _init_l_Lean_Lsp_instInhabitedCancelParams___closed__2(); +lean_mark_persistent(l_Lean_Lsp_instInhabitedCancelParams___closed__2); +l_Lean_Lsp_instInhabitedCancelParams = _init_l_Lean_Lsp_instInhabitedCancelParams(); +lean_mark_persistent(l_Lean_Lsp_instInhabitedCancelParams); +l_Lean_Lsp_instBEqCancelParams___closed__1 = _init_l_Lean_Lsp_instBEqCancelParams___closed__1(); +lean_mark_persistent(l_Lean_Lsp_instBEqCancelParams___closed__1); +l_Lean_Lsp_instBEqCancelParams = _init_l_Lean_Lsp_instBEqCancelParams(); +lean_mark_persistent(l_Lean_Lsp_instBEqCancelParams); +l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____closed__1 = _init_l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____closed__1); +l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____closed__2 = _init_l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____closed__2); +l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____closed__3 = _init_l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____closed__3); +l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____closed__4 = _init_l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____closed__4); +l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____closed__5 = _init_l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____closed__5); +l_Lean_Lsp_instToJsonCancelParams___closed__1 = _init_l_Lean_Lsp_instToJsonCancelParams___closed__1(); +lean_mark_persistent(l_Lean_Lsp_instToJsonCancelParams___closed__1); +l_Lean_Lsp_instToJsonCancelParams = _init_l_Lean_Lsp_instToJsonCancelParams(); +lean_mark_persistent(l_Lean_Lsp_instToJsonCancelParams); +l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__1 = _init_l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__1); +l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__2 = _init_l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__2); +l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__3 = _init_l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__3); +l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__4 = _init_l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__4); +l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__5 = _init_l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__5); +l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__6 = _init_l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__6); +l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__7 = _init_l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__7); +l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__8 = _init_l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__8); +l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__9 = _init_l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__9); +l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__10 = _init_l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__10(); +lean_mark_persistent(l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__10); +l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__11 = _init_l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__11(); +lean_mark_persistent(l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__11); +l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__12 = _init_l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__12(); +lean_mark_persistent(l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__12); +l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__13 = _init_l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__13(); +lean_mark_persistent(l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__13); +l_Lean_Lsp_instFromJsonCancelParams___closed__1 = _init_l_Lean_Lsp_instFromJsonCancelParams___closed__1(); +lean_mark_persistent(l_Lean_Lsp_instFromJsonCancelParams___closed__1); +l_Lean_Lsp_instFromJsonCancelParams = _init_l_Lean_Lsp_instFromJsonCancelParams(); +lean_mark_persistent(l_Lean_Lsp_instFromJsonCancelParams); +return lean_io_result_mk_ok(lean_box(0)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Lean/Data/Lsp/Client.c b/stage0/stdlib/Lean/Data/Lsp/Client.c index 51a4f39710a6..d665aa401a6a 100644 --- a/stage0/stdlib/Lean/Data/Lsp/Client.c +++ b/stage0/stdlib/Lean/Data/Lsp/Client.c @@ -27,7 +27,6 @@ static lean_object* l___private_Lean_Data_Lsp_Client_0__Lean_Lsp_fromJsonRegistr LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Client_0__Lean_Lsp_fromJsonRegistrationParams____x40_Lean_Data_Lsp_Client___hyg_296____spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instFromJsonRegistrationParams___closed__1; static lean_object* l___private_Lean_Data_Lsp_Client_0__Lean_Lsp_fromJsonRegistration____x40_Lean_Data_Lsp_Client___hyg_100____closed__18; -lean_object* l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Client_0__Lean_Lsp_fromJsonRegistration____x40_Lean_Data_Lsp_Client___hyg_100____closed__19; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Client_0__Lean_Lsp_fromJsonRegistration____x40_Lean_Data_Lsp_Client___hyg_100____lambda__1___boxed(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Client_0__Lean_Lsp_toJsonRegistrationParams____x40_Lean_Data_Lsp_Client___hyg_258____spec__1___boxed(lean_object*, lean_object*, lean_object*); @@ -44,7 +43,6 @@ static lean_object* l___private_Lean_Data_Lsp_Client_0__Lean_Lsp_fromJsonRegistr static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Client_0__Lean_Lsp_fromJsonRegistrationParams____x40_Lean_Data_Lsp_Client___hyg_296____spec__1___closed__1; lean_object* l_Lean_Json_getObjValD(lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Client_0__Lean_Lsp_fromJsonRegistration____x40_Lean_Data_Lsp_Client___hyg_100____lambda__1(lean_object*); -lean_object* l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Client_0__Lean_Lsp_toJsonRegistration____x40_Lean_Data_Lsp_Client___hyg_34____closed__5; static lean_object* l___private_Lean_Data_Lsp_Client_0__Lean_Lsp_fromJsonRegistrationParams____x40_Lean_Data_Lsp_Client___hyg_296____closed__2; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonRegistration; @@ -68,6 +66,7 @@ static lean_object* l_Lean_Lsp_instToJsonRegistrationParams___closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonRegistration; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Client_0__Lean_Lsp_fromJsonRegistrationParams____x40_Lean_Data_Lsp_Client___hyg_296_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Client_0__Lean_Lsp_toJsonRegistration____x40_Lean_Data_Lsp_Client___hyg_34____closed__1; +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__1(lean_object*, lean_object*); lean_object* lean_array_mk(lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Client_0__Lean_Lsp_fromJsonRegistration____x40_Lean_Data_Lsp_Client___hyg_100____spec__1(lean_object*, lean_object*); size_t lean_usize_add(size_t, size_t); @@ -89,6 +88,7 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Cli static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Client_0__Lean_Lsp_fromJsonRegistration____x40_Lean_Data_Lsp_Client___hyg_100____spec__1___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Client_0__Lean_Lsp_toJsonRegistration____x40_Lean_Data_Lsp_Client___hyg_34____boxed(lean_object*); lean_object* lean_array_uset(lean_object*, size_t, lean_object*); +lean_object* l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Client_0__Lean_Lsp_fromJsonRegistration____x40_Lean_Data_Lsp_Client___hyg_100____closed__6; static lean_object* l___private_Lean_Data_Lsp_Client_0__Lean_Lsp_toJsonRegistration____x40_Lean_Data_Lsp_Client___hyg_34____closed__4; static lean_object* _init_l___private_Lean_Data_Lsp_Client_0__Lean_Lsp_toJsonRegistration____x40_Lean_Data_Lsp_Client___hyg_34____closed__1() { @@ -199,7 +199,7 @@ x_16 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_16, 0, x_9); lean_ctor_set(x_16, 1, x_15); x_17 = l___private_Lean_Data_Lsp_Client_0__Lean_Lsp_toJsonRegistration____x40_Lean_Data_Lsp_Client___hyg_34____closed__3; -x_18 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_16, x_17); +x_18 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_16, x_17); x_19 = l_Lean_Json_mkObj(x_18); return x_19; } @@ -225,7 +225,7 @@ x_26 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_26, 0, x_9); lean_ctor_set(x_26, 1, x_25); x_27 = l___private_Lean_Data_Lsp_Client_0__Lean_Lsp_toJsonRegistration____x40_Lean_Data_Lsp_Client___hyg_34____closed__3; -x_28 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_26, x_27); +x_28 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_26, x_27); x_29 = l_Lean_Json_mkObj(x_28); return x_29; } @@ -505,7 +505,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Client_0__Lean_Lsp_fromJsonRe lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_Client_0__Lean_Lsp_toJsonRegistration____x40_Lean_Data_Lsp_Client___hyg_34____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -543,7 +543,7 @@ lean_inc(x_12); lean_dec(x_3); x_13 = l___private_Lean_Data_Lsp_Client_0__Lean_Lsp_toJsonRegistration____x40_Lean_Data_Lsp_Client___hyg_34____closed__2; lean_inc(x_1); -x_14 = l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -702,7 +702,7 @@ x_10 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_10, 0, x_9); lean_ctor_set(x_10, 1, x_8); x_11 = l___private_Lean_Data_Lsp_Client_0__Lean_Lsp_toJsonRegistration____x40_Lean_Data_Lsp_Client___hyg_34____closed__3; -x_12 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_10, x_11); +x_12 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_10, x_11); x_13 = l_Lean_Json_mkObj(x_12); return x_13; } diff --git a/stage0/stdlib/Lean/Data/Lsp/CodeActions.c b/stage0/stdlib/Lean/Data/Lsp/CodeActions.c index e020972f5b90..209fa40073ed 100644 --- a/stage0/stdlib/Lean/Data/Lsp/CodeActions.c +++ b/stage0/stdlib/Lean/Data/Lsp/CodeActions.c @@ -36,6 +36,7 @@ LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonCodeActionLiteralSupportValueSet; static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionClientCapabilities____x40_Lean_Data_Lsp_CodeActions___hyg_2050____closed__45; static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionParams____x40_Lean_Data_Lsp_CodeActions___hyg_390____closed__4; static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionClientCapabilities____x40_Lean_Data_Lsp_CodeActions___hyg_2050____closed__41; +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_CodeActionTriggerKind_toCtorIdx___boxed(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_toJsonCodeActionClientCapabilities____x40_Lean_Data_Lsp_CodeActions___hyg_2351_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonCodeActionTriggerKind(uint8_t); @@ -63,8 +64,10 @@ static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCo LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeAction____x40_Lean_Data_Lsp_CodeActions___hyg_1205____spec__3___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionContext____x40_Lean_Data_Lsp_CodeActions___hyg_148____closed__5; static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeAction____x40_Lean_Data_Lsp_CodeActions___hyg_1205____closed__31; +lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6689____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_toJsonCodeActionOptions____x40_Lean_Data_Lsp_CodeActions___hyg_839____closed__3; LEAN_EXPORT lean_object* l_Lean_Lsp_CodeActionTriggerKind_noConfusion___rarg(uint8_t, uint8_t, lean_object*); +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2128_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionParams____x40_Lean_Data_Lsp_CodeActions___hyg_390____closed__24; static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionParams____x40_Lean_Data_Lsp_CodeActions___hyg_390____closed__3; static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionContext____x40_Lean_Data_Lsp_CodeActions___hyg_148____closed__1; @@ -74,12 +77,12 @@ static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCo static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionClientCapabilities____x40_Lean_Data_Lsp_CodeActions___hyg_2050____closed__25; static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeAction____x40_Lean_Data_Lsp_CodeActions___hyg_1205____closed__4; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_toJsonCodeActionOptions____x40_Lean_Data_Lsp_CodeActions___hyg_839_(lean_object*); -lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_toJsonCodeAction____x40_Lean_Data_Lsp_CodeActions___hyg_1131____closed__5; LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeAction____x40_Lean_Data_Lsp_CodeActions___hyg_1205____spec__2(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeAction____x40_Lean_Data_Lsp_CodeActions___hyg_1205____closed__22; static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionClientCapabilities____x40_Lean_Data_Lsp_CodeActions___hyg_2050____closed__29; static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionClientCapabilities____x40_Lean_Data_Lsp_CodeActions___hyg_2050____closed__26; +lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1455____spec__7(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionLiteralSupportValueSet____x40_Lean_Data_Lsp_CodeActions___hyg_1754____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionOptions____x40_Lean_Data_Lsp_CodeActions___hyg_885____closed__17; static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionClientCapabilities____x40_Lean_Data_Lsp_CodeActions___hyg_2050____closed__24; @@ -98,7 +101,6 @@ static lean_object* l_Lean_Lsp_instFromJsonCodeActionDisabled___closed__1; LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_toJsonCodeActionContext____x40_Lean_Data_Lsp_CodeActions___hyg_293____spec__3___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionOptions____x40_Lean_Data_Lsp_CodeActions___hyg_885____closed__2; lean_object* l_Lean_Json_getBool_x3f(lean_object*); -lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionOptions____x40_Lean_Data_Lsp_CodeActions___hyg_885____closed__9; static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionContext____x40_Lean_Data_Lsp_CodeActions___hyg_148____closed__10; static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionDisabled____x40_Lean_Data_Lsp_CodeActions___hyg_699____closed__1; @@ -134,25 +136,24 @@ static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCo static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionClientCapabilities____x40_Lean_Data_Lsp_CodeActions___hyg_2050____closed__9; static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionClientCapabilities____x40_Lean_Data_Lsp_CodeActions___hyg_2050____closed__19; LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionOptions____x40_Lean_Data_Lsp_CodeActions___hyg_885____spec__1___boxed(lean_object*, lean_object*); +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeAction____x40_Lean_Data_Lsp_CodeActions___hyg_1205____spec__4___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instFromJsonCodeActionTriggerKind___closed__4; static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonResolveSupport____x40_Lean_Data_Lsp_CodeActions___hyg_1636____closed__3; +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionParams____x40_Lean_Data_Lsp_CodeActions___hyg_390____closed__27; LEAN_EXPORT lean_object* l_Lean_Lsp_CodeActionTriggerKind_noConfusion___rarg___lambda__1(lean_object*); static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionContext____x40_Lean_Data_Lsp_CodeActions___hyg_148____closed__4; static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionOptions____x40_Lean_Data_Lsp_CodeActions___hyg_885____closed__12; -lean_object* l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionParams____x40_Lean_Data_Lsp_CodeActions___hyg_390____closed__30; static lean_object* l_Lean_Lsp_instFromJsonCodeActionOptions___closed__1; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_toJsonResolveSupport____x40_Lean_Data_Lsp_CodeActions___hyg_1703____spec__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionDisabled____x40_Lean_Data_Lsp_CodeActions___hyg_699____closed__8; -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionDisabled____x40_Lean_Data_Lsp_CodeActions___hyg_699____closed__7; static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionDisabled____x40_Lean_Data_Lsp_CodeActions___hyg_699____closed__5; static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_toJsonCodeActionContext____x40_Lean_Data_Lsp_CodeActions___hyg_293____closed__1; static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionContext____x40_Lean_Data_Lsp_CodeActions___hyg_148____closed__3; lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonPublishDiagnosticsParams____x40_Lean_Data_Lsp_Diagnostics___hyg_2470____spec__2(lean_object*, lean_object*); -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683_(lean_object*); static lean_object* l_Lean_Lsp_instToJsonResolveSupport___closed__1; static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_toJsonCodeAction____x40_Lean_Data_Lsp_CodeActions___hyg_1131____closed__4; lean_object* l_Lean_Json_getObjValD(lean_object*, lean_object*); @@ -189,7 +190,7 @@ static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCo static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionContext____x40_Lean_Data_Lsp_CodeActions___hyg_148____closed__6; static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionLiteralSupportValueSet____x40_Lean_Data_Lsp_CodeActions___hyg_1754____closed__8; LEAN_EXPORT lean_object* l_Lean_Lsp_CodeActionTriggerKind_noConfusion___rarg___lambda__1___boxed(lean_object*); -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2313_(lean_object*); +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionClientCapabilities____x40_Lean_Data_Lsp_CodeActions___hyg_2050____closed__12; lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Message_0__Lean_fromJsonBaseMessage____x40_Lean_Message___hyg_2877____spec__4(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeAction____x40_Lean_Data_Lsp_CodeActions___hyg_1205____closed__9; @@ -209,13 +210,13 @@ static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCo static lean_object* l_Lean_Lsp_instToJsonCodeActionTriggerKind___closed__4; static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionClientCapabilities____x40_Lean_Data_Lsp_CodeActions___hyg_2050____closed__39; static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionClientCapabilities____x40_Lean_Data_Lsp_CodeActions___hyg_2050____closed__16; +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionClientCapabilities____x40_Lean_Data_Lsp_CodeActions___hyg_2050____closed__21; static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionOptions____x40_Lean_Data_Lsp_CodeActions___hyg_885____closed__15; static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionContext____x40_Lean_Data_Lsp_CodeActions___hyg_148____closed__2; LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_toJsonCodeAction____x40_Lean_Data_Lsp_CodeActions___hyg_1131____spec__5(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeAction____x40_Lean_Data_Lsp_CodeActions___hyg_1205____closed__41; static lean_object* l_Lean_Lsp_instToJsonCodeActionTriggerKind___closed__1; -lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__2(lean_object*, lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionLiteralSupportValueSet____x40_Lean_Data_Lsp_CodeActions___hyg_1754____closed__9; static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionParams____x40_Lean_Data_Lsp_CodeActions___hyg_390____closed__14; @@ -235,14 +236,16 @@ static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCo static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionClientCapabilities____x40_Lean_Data_Lsp_CodeActions___hyg_2050____closed__32; static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_toJsonCodeAction____x40_Lean_Data_Lsp_CodeActions___hyg_1131____closed__6; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionDisabled____x40_Lean_Data_Lsp_CodeActions___hyg_699_(lean_object*); -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionClientCapabilities____x40_Lean_Data_Lsp_CodeActions___hyg_2050____closed__7; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonCodeActionLiteralSupportValueSet; +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__2(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionContext____x40_Lean_Data_Lsp_CodeActions___hyg_148____closed__16; static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeAction____x40_Lean_Data_Lsp_CodeActions___hyg_1205____closed__42; static lean_object* l_Lean_Lsp_instFromJsonCodeActionTriggerKind___closed__1; +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeAction____x40_Lean_Data_Lsp_CodeActions___hyg_1205____closed__13; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_toJsonCodeActionContext____x40_Lean_Data_Lsp_CodeActions___hyg_293____spec__2___boxed(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_CodeActionTriggerKind_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionOptions____x40_Lean_Data_Lsp_CodeActions___hyg_885____closed__13; LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_toJsonCodeActionOptions____x40_Lean_Data_Lsp_CodeActions___hyg_839____spec__1___boxed(lean_object*, lean_object*); @@ -262,7 +265,6 @@ static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCo static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionDisabled____x40_Lean_Data_Lsp_CodeActions___hyg_699____closed__9; static lean_object* l_Lean_Lsp_instToJsonCodeActionTriggerKind___closed__3; LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeAction____x40_Lean_Data_Lsp_CodeActions___hyg_1205____spec__2___boxed(lean_object*, lean_object*); -lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6874____spec__1(lean_object*, lean_object*); lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonPublishDiagnosticsParams____x40_Lean_Data_Lsp_Diagnostics___hyg_2414____spec__3(size_t, size_t, lean_object*); static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeAction____x40_Lean_Data_Lsp_CodeActions___hyg_1205____closed__7; static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionClientCapabilities____x40_Lean_Data_Lsp_CodeActions___hyg_2050____closed__33; @@ -296,7 +298,6 @@ LEAN_EXPORT lean_object* l_Lean_Lsp_CodeActionTriggerKind_noConfusion(lean_objec static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionClientCapabilities____x40_Lean_Data_Lsp_CodeActions___hyg_2050____closed__1; static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionOptions____x40_Lean_Data_Lsp_CodeActions___hyg_885____closed__7; static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeAction____x40_Lean_Data_Lsp_CodeActions___hyg_1205____closed__27; -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionClientCapabilities____x40_Lean_Data_Lsp_CodeActions___hyg_2050____closed__6; static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionClientCapabilities____x40_Lean_Data_Lsp_CodeActions___hyg_2050____closed__27; static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionParams____x40_Lean_Data_Lsp_CodeActions___hyg_390____closed__9; @@ -344,7 +345,6 @@ LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Ls static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeAction____x40_Lean_Data_Lsp_CodeActions___hyg_1205____closed__30; LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionClientCapabilities____x40_Lean_Data_Lsp_CodeActions___hyg_2050____spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionContext____x40_Lean_Data_Lsp_CodeActions___hyg_148____spec__2___boxed(lean_object*, lean_object*, lean_object*); -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonResolveSupport____x40_Lean_Data_Lsp_CodeActions___hyg_1636_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_toJsonCodeActionOptions____x40_Lean_Data_Lsp_CodeActions___hyg_839____spec__1(lean_object*, lean_object*); lean_object* lean_string_append(lean_object*, lean_object*); @@ -2100,7 +2100,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJ lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionParams____x40_Lean_Data_Lsp_CodeActions___hyg_390____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -2138,7 +2138,7 @@ lean_inc(x_12); lean_dec(x_3); x_13 = l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionParams____x40_Lean_Data_Lsp_CodeActions___hyg_390____closed__11; lean_inc(x_1); -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____spec__1(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -2177,7 +2177,7 @@ lean_inc(x_23); lean_dec(x_14); x_24 = l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionParams____x40_Lean_Data_Lsp_CodeActions___hyg_390____closed__17; lean_inc(x_1); -x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____spec__1(x_1, x_24); +x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____spec__1(x_1, x_24); if (lean_obj_tag(x_25) == 0) { uint8_t x_26; @@ -2217,7 +2217,7 @@ lean_inc(x_34); lean_dec(x_25); x_35 = l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionParams____x40_Lean_Data_Lsp_CodeActions___hyg_390____closed__22; lean_inc(x_1); -x_36 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__2(x_1, x_35); +x_36 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__2(x_1, x_35); if (lean_obj_tag(x_36) == 0) { uint8_t x_37; @@ -2362,14 +2362,14 @@ lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_obj x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); x_3 = l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionParams____x40_Lean_Data_Lsp_CodeActions___hyg_390____closed__1; -x_4 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6874____spec__1(x_3, x_2); +x_4 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6689____spec__1(x_3, x_2); x_5 = lean_ctor_get(x_1, 1); lean_inc(x_5); x_6 = l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionParams____x40_Lean_Data_Lsp_CodeActions___hyg_390____closed__11; -x_7 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6874____spec__1(x_6, x_5); +x_7 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6689____spec__1(x_6, x_5); x_8 = lean_ctor_get(x_1, 2); lean_inc(x_8); -x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2313_(x_8); +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2128_(x_8); x_10 = l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionParams____x40_Lean_Data_Lsp_CodeActions___hyg_390____closed__17; x_11 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_11, 0, x_10); @@ -2380,7 +2380,7 @@ lean_ctor_set(x_13, 0, x_11); lean_ctor_set(x_13, 1, x_12); x_14 = lean_ctor_get(x_1, 3); lean_inc(x_14); -x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_14); +x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_14); x_16 = l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionParams____x40_Lean_Data_Lsp_CodeActions___hyg_390____closed__22; x_17 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_17, 0, x_16); @@ -3422,7 +3422,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_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408_(x_4); +x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223_(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); @@ -3450,7 +3450,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_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627_(x_4); +x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442_(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); @@ -3525,11 +3525,11 @@ lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_obj x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); x_3 = l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionParams____x40_Lean_Data_Lsp_CodeActions___hyg_390____closed__1; -x_4 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6874____spec__1(x_3, x_2); +x_4 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6689____spec__1(x_3, x_2); x_5 = lean_ctor_get(x_1, 1); lean_inc(x_5); x_6 = l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionParams____x40_Lean_Data_Lsp_CodeActions___hyg_390____closed__11; -x_7 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6874____spec__1(x_6, x_5); +x_7 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6689____spec__1(x_6, x_5); x_8 = lean_ctor_get(x_1, 2); lean_inc(x_8); x_9 = lean_alloc_ctor(3, 1, 0); @@ -3571,7 +3571,7 @@ x_32 = lean_ctor_get(x_1, 9); lean_inc(x_32); lean_dec(x_1); x_33 = l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_toJsonCodeAction____x40_Lean_Data_Lsp_CodeActions___hyg_1131____closed__7; -x_34 = l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(x_33, x_32); +x_34 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1455____spec__7(x_33, x_32); lean_dec(x_32); x_35 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_35, 0, x_34); @@ -4155,7 +4155,7 @@ return x_4; case 1: { lean_object* x_5; -x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444_(x_3); +x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259_(x_3); if (lean_obj_tag(x_5) == 0) { uint8_t x_6; @@ -4206,7 +4206,7 @@ return x_14; { lean_object* x_15; uint8_t x_16; lean_inc(x_3); -x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444_(x_3); +x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259_(x_3); x_16 = !lean_is_exclusive(x_3); if (x_16 == 0) { @@ -4325,7 +4325,7 @@ return x_4; case 1: { lean_object* x_5; -x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683_(x_3); +x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498_(x_3); if (lean_obj_tag(x_5) == 0) { uint8_t x_6; @@ -4376,7 +4376,7 @@ return x_14; { lean_object* x_15; uint8_t x_16; lean_inc(x_3); -x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683_(x_3); +x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498_(x_3); x_16 = !lean_is_exclusive(x_3); if (x_16 == 0) { @@ -4950,7 +4950,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJ lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionParams____x40_Lean_Data_Lsp_CodeActions___hyg_390____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -4988,7 +4988,7 @@ lean_inc(x_12); lean_dec(x_3); x_13 = l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionParams____x40_Lean_Data_Lsp_CodeActions___hyg_390____closed__11; lean_inc(x_1); -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____spec__1(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; diff --git a/stage0/stdlib/Lean/Data/Lsp/Diagnostics.c b/stage0/stdlib/Lean/Data/Lsp/Diagnostics.c index 5c5aa0e4d93a..49e76388f362 100644 --- a/stage0/stdlib/Lean/Data/Lsp/Diagnostics.c +++ b/stage0/stdlib/Lean/Data/Lsp/Diagnostics.c @@ -27,20 +27,19 @@ LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Ls static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1553____rarg___closed__12; lean_object* lean_mk_empty_array_with_capacity(lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1553____spec__3(lean_object*, lean_object*); -uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_627_(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticRelatedInformation____x40_Lean_Data_Lsp_Diagnostics___hyg_956____closed__16; static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1553____rarg___closed__19; static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1553____rarg___closed__9; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1553____spec__6___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1553____spec__1___closed__9; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonDiagnosticWith(lean_object*); -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonDiagnosticSeverity___boxed(lean_object*); static lean_object* l_Lean_Lsp_instFromJsonDiagnosticTag___closed__4; lean_object* l_Lean_Json_mkObj(lean_object*); static lean_object* l_Lean_Lsp_DiagnosticSeverity_noConfusion___rarg___closed__1; +uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqLocation____x40_Lean_Data_Lsp_Basic___hyg_821_(lean_object*, lean_object*); +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1553_(lean_object*); -uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordRange____x40_Lean_Data_Lsp_Basic___hyg_900_(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticRelatedInformation____x40_Lean_Data_Lsp_Diagnostics___hyg_956____closed__4; static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1553____rarg___closed__21; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonDiagnosticTag(lean_object*); @@ -56,7 +55,6 @@ LEAN_EXPORT uint8_t l_Array_isEqvAux___at___private_Lean_Data_Lsp_Diagnostics_0_ static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1553____rarg___closed__40; LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticRelatedInformation____x40_Lean_Data_Lsp_Diagnostics___hyg_956____spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instOrdDiagnosticCode; -lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1553____rarg___closed__28; static lean_object* l_Lean_Lsp_instFromJsonDiagnosticSeverity___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1455____at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonPublishDiagnosticsParams____x40_Lean_Data_Lsp_Diagnostics___hyg_2414____spec__2(lean_object*); @@ -69,14 +67,14 @@ static lean_object* l_Lean_Lsp_instFromJsonDiagnosticTag___closed__2; static lean_object* l_Lean_Lsp_instInhabitedDiagnosticRelatedInformation___closed__2; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1455____spec__4___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1254____rarg(lean_object*, lean_object*, lean_object*); -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080_(lean_object*); static lean_object* l_Lean_Lsp_instFromJsonDiagnosticTag___closed__3; static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1553____spec__1___closed__5; +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____spec__2(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1553____rarg___closed__22; static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1553____rarg___closed__43; static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1553____rarg___closed__5; static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticRelatedInformation____x40_Lean_Data_Lsp_Diagnostics___hyg_956____closed__6; -lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1455____spec__7(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonPublishDiagnosticsParams____x40_Lean_Data_Lsp_Diagnostics___hyg_2470____closed__7; static lean_object* l_Lean_Lsp_instToJsonDiagnosticSeverity___closed__5; static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1553____spec__3___closed__1; @@ -93,6 +91,7 @@ static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDi static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1553____rarg___closed__35; LEAN_EXPORT lean_object* l_Lean_Lsp_DiagnosticTag_noConfusion___rarg(uint8_t, uint8_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1455____spec__1(lean_object*, lean_object*); +uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordLocation____x40_Lean_Data_Lsp_Basic___hyg_1053_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_DiagnosticSeverity_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); 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_Diagnostics_0__Lean_Lsp_fromJsonPublishDiagnosticsParams____x40_Lean_Data_Lsp_Diagnostics___hyg_2470____closed__14; @@ -110,7 +109,6 @@ static lean_object* l_Lean_Lsp_instBEqPublishDiagnosticsParams___closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_instBEqDiagnosticCode; LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticTag____x40_Lean_Data_Lsp_Diagnostics___hyg_628_(uint8_t, uint8_t); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1455_(lean_object*); -uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordLocation____x40_Lean_Data_Lsp_Basic___hyg_1238_(lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_DiagnosticWith_fullRange(lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1553____spec__7___boxed(lean_object*, lean_object*); @@ -127,15 +125,16 @@ static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDi LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1254____at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqPublishDiagnosticsParams____x40_Lean_Data_Lsp_Diagnostics___hyg_2324____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonDiagnosticSeverity(uint8_t); LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1254____spec__2___boxed(lean_object*, lean_object*); +lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonPublishDiagnosticsParams; LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_ordDiagnosticTag____x40_Lean_Data_Lsp_Diagnostics___hyg_646_(uint8_t, uint8_t); LEAN_EXPORT lean_object* l_Lean_Lsp_instInhabitedDiagnosticWith(lean_object*); -lean_object* l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonPublishDiagnosticsParams____x40_Lean_Data_Lsp_Diagnostics___hyg_2414____spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instBEqDiagnosticCode___closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_DiagnosticSeverity_noConfusion___rarg(uint8_t, uint8_t, lean_object*); lean_object* lean_nat_to_int(lean_object*); static lean_object* l_Lean_Lsp_instFromJsonDiagnosticSeverity___closed__2; +uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordRange____x40_Lean_Data_Lsp_Basic___hyg_715_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1254____spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instBEqDiagnosticSeverity; LEAN_EXPORT uint8_t l_Lean_Lsp_instInhabitedDiagnosticTag; @@ -186,7 +185,6 @@ static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Dia LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_DiagnosticWith_ordUserVisible____x40_Lean_Data_Lsp_Diagnostics___hyg_2056_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonDiagnosticWith___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instBEqPublishDiagnosticsParams; -lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__2(lean_object*, lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1553____rarg___closed__34; LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_ordDiagnosticRelatedInformation____x40_Lean_Data_Lsp_Diagnostics___hyg_1062_(lean_object*, lean_object*); @@ -202,13 +200,16 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDi static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1553____rarg___closed__31; static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticRelatedInformation____x40_Lean_Data_Lsp_Diagnostics___hyg_956____closed__8; static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonPublishDiagnosticsParams____x40_Lean_Data_Lsp_Diagnostics___hyg_2470____closed__15; +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1553____rarg___closed__39; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonPublishDiagnosticsParams____x40_Lean_Data_Lsp_Diagnostics___hyg_2414____spec__3___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticRelatedInformation____x40_Lean_Data_Lsp_Diagnostics___hyg_956____closed__17; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1455____spec__6___boxed(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__2(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instToJsonDiagnosticSeverity___closed__4; static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1553____rarg___closed__14; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_ordDiagnosticSeverity____x40_Lean_Data_Lsp_Diagnostics___hyg_39____boxed(lean_object*, lean_object*); +lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____spec__2(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1553____rarg___closed__16; static lean_object* l_Lean_Lsp_instToJsonDiagnosticSeverity___closed__7; LEAN_EXPORT lean_object* l_Lean_Lsp_instInhabitedDiagnosticRelatedInformation; @@ -252,13 +253,12 @@ static lean_object* l_Lean_Lsp_instToJsonDiagnosticSeverity___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1254____rarg___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticRelatedInformation____x40_Lean_Data_Lsp_Diagnostics___hyg_956____closed__10; static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonPublishDiagnosticsParams____x40_Lean_Data_Lsp_Diagnostics___hyg_2414____closed__2; +uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_442_(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1553____rarg___closed__38; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonDiagnosticTag(uint8_t); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonDiagnosticCode(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_ordDiagnosticRelatedInformation____x40_Lean_Data_Lsp_Diagnostics___hyg_1062____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonDiagnosticRelatedInformation; -uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqLocation____x40_Lean_Data_Lsp_Basic___hyg_1006_(lean_object*, lean_object*); -lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_DiagnosticWith_UserVisible_ofDiagnostic___rarg___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_DiagnosticSeverity_noConfusion___rarg___lambda__1(lean_object*); LEAN_EXPORT uint8_t l_Lean_Lsp_instInhabitedDiagnosticSeverity; @@ -271,7 +271,6 @@ lean_object* l_Lean_Json_getNat_x3f(lean_object*); static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1553____spec__1___closed__2; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonPublishDiagnosticsParams____x40_Lean_Data_Lsp_Diagnostics___hyg_2470____spec__4(size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticCode____x40_Lean_Data_Lsp_Diagnostics___hyg_327____boxed(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_1410____spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1553____spec__1___closed__7; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonDiagnosticSeverity(lean_object*); static lean_object* l_Lean_Lsp_instInhabitedPublishDiagnosticsParams___closed__2; @@ -296,11 +295,10 @@ lean_object* lean_array_uget(lean_object*, size_t); size_t lean_array_size(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1553____rarg___closed__26; LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_ordDiagnosticSeverity____x40_Lean_Data_Lsp_Diagnostics___hyg_39_(uint8_t, uint8_t); -lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____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_1225____spec__1(lean_object*, lean_object*); uint8_t lean_int_dec_eq(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1553____rarg___closed__23; static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1553____rarg___closed__17; -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(lean_object*); lean_object* lean_string_append(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticRelatedInformation____x40_Lean_Data_Lsp_Diagnostics___hyg_904____closed__2; lean_object* lean_array_get_size(lean_object*); @@ -319,6 +317,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_Diagn static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1553____rarg___closed__25; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticRelatedInformation____x40_Lean_Data_Lsp_Diagnostics___hyg_904_(lean_object*); lean_object* l_Lean_Json_pretty(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1455____spec__7___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1455____rarg___closed__8; static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1553____rarg___closed__30; LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticRelatedInformation____x40_Lean_Data_Lsp_Diagnostics___hyg_830_(lean_object*, lean_object*); @@ -335,6 +334,8 @@ static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonPu static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonPublishDiagnosticsParams____x40_Lean_Data_Lsp_Diagnostics___hyg_2414____closed__3; LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1455____spec__3(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1455____rarg(lean_object*, lean_object*); +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947_(lean_object*); +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895_(lean_object*); LEAN_EXPORT uint8_t l_Array_isEqvAux___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqPublishDiagnosticsParams____x40_Lean_Data_Lsp_Diagnostics___hyg_2324____spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticSeverity____x40_Lean_Data_Lsp_Diagnostics___hyg_21_(uint8_t, uint8_t); static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticRelatedInformation____x40_Lean_Data_Lsp_Diagnostics___hyg_956____closed__3; @@ -1804,7 +1805,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 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqLocation____x40_Lean_Data_Lsp_Basic___hyg_1006_(x_3, x_5); +x_7 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqLocation____x40_Lean_Data_Lsp_Basic___hyg_821_(x_3, x_5); if (x_7 == 0) { uint8_t x_8; @@ -1877,7 +1878,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJso lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; 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_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080_(x_2); +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895_(x_2); x_4 = l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticRelatedInformation____x40_Lean_Data_Lsp_Diagnostics___hyg_904____closed__1; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); @@ -1931,7 +1932,7 @@ LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Ls { 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_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132_(x_3); +x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947_(x_3); return x_4; } } @@ -2252,7 +2253,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 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordLocation____x40_Lean_Data_Lsp_Basic___hyg_1238_(x_3, x_5); +x_7 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordLocation____x40_Lean_Data_Lsp_Basic___hyg_1053_(x_3, x_5); x_8 = lean_string_dec_lt(x_4, x_6); if (x_8 == 0) { @@ -2394,7 +2395,7 @@ else lean_object* x_6; lean_object* x_7; uint8_t x_8; x_6 = lean_ctor_get(x_1, 0); x_7 = lean_ctor_get(x_2, 0); -x_8 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_627_(x_6, x_7); +x_8 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_442_(x_6, x_7); return x_8; } } @@ -2751,7 +2752,7 @@ lean_inc(x_20); x_21 = lean_ctor_get(x_3, 8); lean_inc(x_21); lean_dec(x_3); -x_22 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_627_(x_4, x_13); +x_22 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_442_(x_4, x_13); lean_dec(x_13); lean_dec(x_4); if (x_22 == 0) @@ -3400,6 +3401,32 @@ return x_19; } } } +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1455____spec__7(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_3; +lean_dec(x_1); +x_3 = lean_box(0); +return x_3; +} +else +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_4 = lean_ctor_get(x_2, 0); +lean_inc(x_4); +x_5 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_5, 0, x_1); +lean_ctor_set(x_5, 1, x_4); +x_6 = lean_box(0); +x_7 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_7, 0, x_5); +lean_ctor_set(x_7, 1, x_6); +return x_7; +} +} +} static lean_object* _init_l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1455____rarg___closed__1() { _start: { @@ -3470,7 +3497,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJso lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; 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_3 = lean_ctor_get(x_2, 0); lean_inc(x_3); -x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_3); +x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_3); x_5 = l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1455____rarg___closed__1; x_6 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_6, 0, x_5); @@ -3482,7 +3509,7 @@ lean_ctor_set(x_8, 1, x_7); x_9 = lean_ctor_get(x_2, 1); lean_inc(x_9); x_10 = l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1455____rarg___closed__2; -x_11 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____spec__1(x_10, x_9); +x_11 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155____spec__1(x_10, x_9); x_12 = lean_ctor_get(x_2, 2); lean_inc(x_12); x_13 = l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1455____rarg___closed__3; @@ -3495,7 +3522,7 @@ x_17 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJs x_18 = lean_ctor_get(x_2, 4); lean_inc(x_18); x_19 = l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1455____rarg___closed__5; -x_20 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(x_19, x_18); +x_20 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____spec__2(x_19, x_18); x_21 = lean_ctor_get(x_2, 5); lean_inc(x_21); x_22 = lean_apply_1(x_1, x_21); @@ -3518,7 +3545,7 @@ x_32 = lean_ctor_get(x_2, 8); lean_inc(x_32); lean_dec(x_2); x_33 = l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1455____rarg___closed__8; -x_34 = l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(x_33, x_32); +x_34 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1455____spec__7(x_33, x_32); lean_dec(x_32); x_35 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_35, 0, x_34); @@ -3594,6 +3621,15 @@ x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Ls return x_6; } } +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1455____spec__7___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_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1455____spec__7(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonDiagnosticWith___rarg(lean_object* x_1) { _start: { @@ -5048,7 +5084,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJ lean_object* x_3; lean_object* x_4; x_3 = l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1455____rarg___closed__1; lean_inc(x_2); -x_4 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__2(x_2, x_3); +x_4 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__2(x_2, x_3); if (lean_obj_tag(x_4) == 0) { uint8_t x_5; @@ -5087,7 +5123,7 @@ lean_inc(x_13); lean_dec(x_4); x_14 = l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1455____rarg___closed__2; lean_inc(x_2); -x_15 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____spec__1(x_2, x_14); +x_15 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____spec__1(x_2, x_14); if (lean_obj_tag(x_15) == 0) { uint8_t x_16; @@ -5210,7 +5246,7 @@ lean_inc(x_46); lean_dec(x_37); x_47 = l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1455____rarg___closed__5; lean_inc(x_2); -x_48 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____spec__2(x_2, x_47); +x_48 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____spec__2(x_2, x_47); if (lean_obj_tag(x_48) == 0) { uint8_t x_49; @@ -5603,7 +5639,7 @@ lean_inc(x_18); x_19 = lean_ctor_get(x_3, 7); lean_inc(x_19); lean_dec(x_3); -x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordRange____x40_Lean_Data_Lsp_Basic___hyg_900_(x_4, x_12); +x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordRange____x40_Lean_Data_Lsp_Basic___hyg_715_(x_4, x_12); lean_dec(x_12); lean_dec(x_4); x_21 = lean_apply_2(x_1, x_9, x_17); @@ -5646,7 +5682,7 @@ lean_dec(x_5); x_125 = lean_ctor_get(x_13, 0); lean_inc(x_125); lean_dec(x_13); -x_126 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordRange____x40_Lean_Data_Lsp_Basic___hyg_900_(x_124, x_125); +x_126 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordRange____x40_Lean_Data_Lsp_Basic___hyg_715_(x_124, x_125); lean_dec(x_125); lean_dec(x_124); x_23 = x_126; @@ -6422,7 +6458,7 @@ lean_inc(x_19); x_20 = lean_ctor_get(x_2, 8); lean_inc(x_20); lean_dec(x_2); -x_21 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_627_(x_3, x_12); +x_21 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_442_(x_3, x_12); lean_dec(x_12); lean_dec(x_3); if (x_21 == 0) @@ -6804,7 +6840,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJso lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; 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_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_742_(x_2); +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_2); x_4 = l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1455____rarg___closed__1; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); @@ -6816,7 +6852,7 @@ lean_ctor_set(x_7, 1, x_6); x_8 = lean_ctor_get(x_1, 1); lean_inc(x_8); x_9 = l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1455____rarg___closed__2; -x_10 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____spec__1(x_9, x_8); +x_10 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155____spec__1(x_9, x_8); x_11 = lean_ctor_get(x_1, 2); lean_inc(x_11); x_12 = l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1455____rarg___closed__3; @@ -6829,7 +6865,7 @@ x_16 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJs x_17 = lean_ctor_get(x_1, 4); lean_inc(x_17); x_18 = l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1455____rarg___closed__5; -x_19 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(x_18, x_17); +x_19 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____spec__2(x_18, x_17); x_20 = lean_ctor_get(x_1, 5); lean_inc(x_20); x_21 = lean_alloc_ctor(3, 1, 0); @@ -6853,7 +6889,7 @@ x_31 = lean_ctor_get(x_1, 8); lean_inc(x_31); lean_dec(x_1); x_32 = l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1455____rarg___closed__8; -x_33 = l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(x_32, x_31); +x_33 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1455____spec__7(x_32, x_31); lean_dec(x_31); x_34 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_34, 0, x_33); @@ -7191,7 +7227,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJ lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1455____rarg___closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__2(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__2(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -7229,7 +7265,7 @@ lean_inc(x_12); lean_dec(x_3); x_13 = l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1455____rarg___closed__2; lean_inc(x_1); -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____spec__1(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -7349,7 +7385,7 @@ lean_inc(x_45); lean_dec(x_36); x_46 = l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1455____rarg___closed__5; lean_inc(x_1); -x_47 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____spec__2(x_1, x_46); +x_47 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____spec__2(x_1, x_46); if (lean_obj_tag(x_47) == 0) { uint8_t x_48; @@ -7886,7 +7922,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJ lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonPublishDiagnosticsParams____x40_Lean_Data_Lsp_Diagnostics___hyg_2414____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; diff --git a/stage0/stdlib/Lean/Data/Lsp/Extra.c b/stage0/stdlib/Lean/Data/Lsp/Extra.c index 0b4afba07b51..938bdfbddaf7 100644 --- a/stage0/stdlib/Lean/Data/Lsp/Extra.c +++ b/stage0/stdlib/Lean/Data/Lsp/Extra.c @@ -37,6 +37,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_reprLineRan static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_reprLineRange____x40_Lean_Data_Lsp_Extra___hyg_2780____closed__12; lean_object* l_Lean_Json_mkObj(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLeanDidOpenTextDocumentParams____x40_Lean_Data_Lsp_Extra___hyg_203____closed__8; +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLeanDidOpenTextDocumentParams____x40_Lean_Data_Lsp_Extra___hyg_203____closed__15; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_toJsonLineRange____x40_Lean_Data_Lsp_Extra___hyg_2944_(lean_object*); static lean_object* l_Lean_Lsp_instToJsonPlainGoalParams___closed__1; @@ -76,11 +77,11 @@ static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcCallP static lean_object* l_Lean_Lsp_instFromJsonLeanDidOpenTextDocumentParams___closed__1; uint8_t l_Lean_Name_isAnonymous(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPlainGoal____x40_Lean_Data_Lsp_Extra___hyg_1238____spec__2(size_t, size_t, lean_object*); -lean_object* l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_reprLineRange____x40_Lean_Data_Lsp_Extra___hyg_2780____closed__17; static lean_object* l_Lean_Lsp_instToJsonLeanFileProgressKind___closed__4; static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLineRange____x40_Lean_Data_Lsp_Extra___hyg_2838____closed__1; static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_reprLineRange____x40_Lean_Data_Lsp_Extra___hyg_2780____closed__4; +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2128_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcConnectParams____x40_Lean_Data_Lsp_Extra___hyg_1756____closed__3; static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLeanDidOpenTextDocumentParams____x40_Lean_Data_Lsp_Extra___hyg_203____spec__1___closed__1; static lean_object* l_Lean_Lsp_instFromJsonWaitForDiagnostics___closed__1; @@ -99,14 +100,12 @@ static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLeanFile static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcCallParams____x40_Lean_Data_Lsp_Extra___hyg_2014____closed__10; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcConnected____x40_Lean_Data_Lsp_Extra___hyg_1873_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLeanDidOpenTextDocumentParams____x40_Lean_Data_Lsp_Extra___hyg_203____spec__1(lean_object*, lean_object*); -lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLineRange____x40_Lean_Data_Lsp_Extra___hyg_2838____closed__7; static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPlainTermGoal____x40_Lean_Data_Lsp_Extra___hyg_1586____closed__1; static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPlainTermGoalParams____x40_Lean_Data_Lsp_Extra___hyg_1408____closed__6; static lean_object* l_Lean_Lsp_instToJsonPlainTermGoalParams___closed__1; static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLeanDidOpenTextDocumentParams____x40_Lean_Data_Lsp_Extra___hyg_203____closed__4; static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_reprLineRange____x40_Lean_Data_Lsp_Extra___hyg_2780____closed__3; -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2440_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLeanDidOpenTextDocumentParams____x40_Lean_Data_Lsp_Extra___hyg_203____closed__20; lean_object* l_List_flatMapTR_go___at___private_Lean_Server_Rpc_Basic_0__Lean_Lsp_toJsonRpcRef____x40_Lean_Server_Rpc_Basic___hyg_173____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLeanFileProgressProcessingInfo____x40_Lean_Data_Lsp_Extra___hyg_710____closed__6; @@ -119,12 +118,12 @@ static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLeanFile static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_reprLineRange____x40_Lean_Data_Lsp_Extra___hyg_2780____closed__1; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_toJsonLeanFileProgressParams____x40_Lean_Data_Lsp_Extra___hyg_995____spec__1___boxed(lean_object*, lean_object*, lean_object*); uint8_t lean_string_dec_eq(lean_object*, lean_object*); +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____spec__1(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Lsp_instInhabitedDependencyBuildMode; static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLeanDidOpenTextDocumentParams____x40_Lean_Data_Lsp_Extra___hyg_203____closed__19; static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcCallParams____x40_Lean_Data_Lsp_Extra___hyg_2014____closed__6; static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPlainTermGoalParams____x40_Lean_Data_Lsp_Extra___hyg_1408____closed__4; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonPlainTermGoalParams; -lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPlainTermGoalParams____x40_Lean_Data_Lsp_Extra___hyg_1408____closed__7; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonRpcCallParams; static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPlainGoal____x40_Lean_Data_Lsp_Extra___hyg_1238____closed__14; @@ -148,6 +147,7 @@ static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcCallP LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_toJsonDependencyBuildMode____x40_Lean_Data_Lsp_Extra___hyg_133_(uint8_t); static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcConnectParams____x40_Lean_Data_Lsp_Extra___hyg_1756____closed__2; static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcCallParams____x40_Lean_Data_Lsp_Extra___hyg_2014____spec__1___closed__1; +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLeanFileProgressProcessingInfo____x40_Lean_Data_Lsp_Extra___hyg_710____closed__1; static lean_object* l_Lean_Lsp_instToJsonLeanDidOpenTextDocumentParams___closed__1; static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_reprLineRange____x40_Lean_Data_Lsp_Extra___hyg_2780____closed__15; @@ -167,6 +167,7 @@ static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLineRang lean_object* lean_nat_to_int(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcReleaseParams____x40_Lean_Data_Lsp_Extra___hyg_2360____closed__1; static lean_object* l_Lean_Lsp_instFromJsonLeanFileProgressKind___closed__1; +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2255_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_toJsonDependencyBuildMode____x40_Lean_Data_Lsp_Extra___hyg_133____closed__3; static lean_object* l_Lean_Lsp_instFromJsonRpcConnectParams___closed__1; static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLeanFileProgressParams____x40_Lean_Data_Lsp_Extra___hyg_889____closed__8; @@ -177,7 +178,6 @@ static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcCallP static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPlainTermGoal____x40_Lean_Data_Lsp_Extra___hyg_1586____closed__10; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_toJsonLeanFileProgressParams____x40_Lean_Data_Lsp_Extra___hyg_995_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPlainGoal____x40_Lean_Data_Lsp_Extra___hyg_1238____closed__7; -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4990_(lean_object*); static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLeanFileProgressParams____x40_Lean_Data_Lsp_Extra___hyg_889____spec__1___closed__1; lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidOpenTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_164____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLeanDidOpenTextDocumentParams____x40_Lean_Data_Lsp_Extra___hyg_203____spec__1___boxed(lean_object*, lean_object*); @@ -199,10 +199,8 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPla static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPlainGoalParams____x40_Lean_Data_Lsp_Extra___hyg_1059____closed__5; LEAN_EXPORT lean_object* l_Lean_Lsp_LeanFileProgressKind_toCtorIdx___boxed(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcCallParams____x40_Lean_Data_Lsp_Extra___hyg_2014____closed__11; -lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPlainGoal____x40_Lean_Data_Lsp_Extra___hyg_1238____spec__2___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPlainGoal____x40_Lean_Data_Lsp_Extra___hyg_1238_(lean_object*); -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2313_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_LeanFileProgressKind_noConfusion(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPlainGoal____x40_Lean_Data_Lsp_Extra___hyg_1238____closed__4; static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonWaitForDiagnosticsParams____x40_Lean_Data_Lsp_Extra___hyg_371____closed__5; @@ -226,6 +224,7 @@ LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonLineRange; static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcReleaseParams____x40_Lean_Data_Lsp_Extra___hyg_2360____closed__12; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonRpcConnected; static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcConnected____x40_Lean_Data_Lsp_Extra___hyg_1873____spec__1___closed__2; +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonDependencyBuildMode____x40_Lean_Data_Lsp_Extra___hyg_11____lambda__2___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPlainGoalParams____x40_Lean_Data_Lsp_Extra___hyg_1059____closed__8; static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPlainGoalParams____x40_Lean_Data_Lsp_Extra___hyg_1059____closed__10; @@ -236,7 +235,6 @@ LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Ls static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLeanFileProgressProcessingInfo____x40_Lean_Data_Lsp_Extra___hyg_710____closed__9; LEAN_EXPORT lean_object* l_Lean_Lsp_DependencyBuildMode_toCtorIdx(uint8_t); static lean_object* l_Lean_Lsp_instToJsonLeanFileProgressKind___closed__2; -lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__2(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_toJsonLeanFileProgressProcessingInfo____x40_Lean_Data_Lsp_Extra___hyg_816____closed__3; lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonWaitForDiagnosticsParams; @@ -261,6 +259,7 @@ LEAN_EXPORT lean_object* l_Lean_Lsp_LeanFileProgressKind_toCtorIdx(uint8_t); static lean_object* l_Lean_Lsp_instInhabitedLineRange___closed__1; static lean_object* l_Lean_Lsp_instToJsonLeanFileProgressKind___closed__1; static lean_object* l_Lean_Lsp_instToJsonPlainTermGoal___closed__1; +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonWaitForDiagnostics(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPlainGoal____x40_Lean_Data_Lsp_Extra___hyg_1238____closed__1; static lean_object* l_Lean_Lsp_instFromJsonDependencyBuildMode___closed__1; @@ -269,11 +268,13 @@ LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonPlainTermGoalParams; static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcKeepAliveParams____x40_Lean_Data_Lsp_Extra___hyg_2591____closed__6; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonDependencyBuildMode____x40_Lean_Data_Lsp_Extra___hyg_11____lambda__1(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPlainTermGoal____x40_Lean_Data_Lsp_Extra___hyg_1586_(lean_object*); +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__2(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_toJsonDependencyBuildMode____x40_Lean_Data_Lsp_Extra___hyg_133____closed__1; static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPlainTermGoalParams____x40_Lean_Data_Lsp_Extra___hyg_1408____closed__8; static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPlainGoalParams____x40_Lean_Data_Lsp_Extra___hyg_1059____closed__3; static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLeanDidOpenTextDocumentParams____x40_Lean_Data_Lsp_Extra___hyg_203____closed__14; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLeanFileProgressProcessingInfo____x40_Lean_Data_Lsp_Extra___hyg_710_(lean_object*); +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonRpcConnected; static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonWaitForDiagnosticsParams____x40_Lean_Data_Lsp_Extra___hyg_371____closed__10; static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcCallParams____x40_Lean_Data_Lsp_Extra___hyg_2014____closed__16; @@ -306,7 +307,6 @@ static lean_object* l_Lean_Lsp_instFromJsonPlainGoalParams___closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonPlainGoalParams; static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLeanFileProgressParams____x40_Lean_Data_Lsp_Extra___hyg_889____closed__4; static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonDependencyBuildMode____x40_Lean_Data_Lsp_Extra___hyg_11____lambda__3___closed__2; -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_reprLineRange____x40_Lean_Data_Lsp_Extra___hyg_2780____closed__16; lean_object* lean_string_length(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_toJsonRpcConnected____x40_Lean_Data_Lsp_Extra___hyg_1940_(uint64_t); @@ -373,9 +373,9 @@ LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Ls LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_toJsonRpcCallParams____x40_Lean_Data_Lsp_Extra___hyg_2237_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcConnected____x40_Lean_Data_Lsp_Extra___hyg_1873____spec__1___lambda__1(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instBEqLeanFileProgressKind___closed__1; -lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPlainGoal____x40_Lean_Data_Lsp_Extra___hyg_1238____closed__11; lean_object* l_Lean_bignumToJson(lean_object*); +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonPlainTermGoal; static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcConnected____x40_Lean_Data_Lsp_Extra___hyg_1873____closed__7; LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLeanFileProgressParams____x40_Lean_Data_Lsp_Extra___hyg_889____spec__1___boxed(lean_object*, lean_object*); @@ -387,6 +387,7 @@ lean_object* lean_array_mk(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPlainTermGoalParams____x40_Lean_Data_Lsp_Extra___hyg_1408____closed__2; static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_reprLineRange____x40_Lean_Data_Lsp_Extra___hyg_2780____closed__10; static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcReleaseParams____x40_Lean_Data_Lsp_Extra___hyg_2360____closed__10; +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4805_(lean_object*); size_t lean_usize_add(size_t, size_t); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonDependencyBuildMode____x40_Lean_Data_Lsp_Extra___hyg_11____lambda__1___boxed(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_toJsonLeanFileProgressParams____x40_Lean_Data_Lsp_Extra___hyg_995____spec__1(size_t, size_t, lean_object*); @@ -403,7 +404,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPla static lean_object* l_Lean_Lsp_instToJsonWaitForDiagnostics___closed__1; static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcConnectParams____x40_Lean_Data_Lsp_Extra___hyg_1756____closed__5; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLeanFileProgressParams____x40_Lean_Data_Lsp_Extra___hyg_889_(lean_object*); -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonWaitForDiagnosticsParams; static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonDependencyBuildMode____x40_Lean_Data_Lsp_Extra___hyg_11____closed__1; lean_object* lean_string_append(lean_object*, lean_object*); @@ -426,12 +426,12 @@ static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonWaitForD LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonLeanFileProgressParams; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_toJsonRpcKeepAliveParams____x40_Lean_Data_Lsp_Extra___hyg_2697_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPlainGoal____x40_Lean_Data_Lsp_Extra___hyg_1238____closed__6; +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221_(lean_object*); lean_object* l_Lean_Json_pretty(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcKeepAliveParams____x40_Lean_Data_Lsp_Extra___hyg_2591____closed__5; LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcCallParams____x40_Lean_Data_Lsp_Extra___hyg_2014____spec__1___lambda__1___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcReleaseParams____x40_Lean_Data_Lsp_Extra___hyg_2360____closed__13; static lean_object* l_Lean_Lsp_instToJsonLineRange___closed__1; -lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPlainGoalParams____x40_Lean_Data_Lsp_Extra___hyg_1059____closed__4; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_toJsonLeanFileProgressProcessingInfo____x40_Lean_Data_Lsp_Extra___hyg_816_(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_toJsonPlainGoal____x40_Lean_Data_Lsp_Extra___hyg_1344____spec__1(size_t, size_t, lean_object*); @@ -1475,7 +1475,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_toJsonLeanD lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; 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_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4990_(x_2); +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4805_(x_2); x_4 = l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLeanDidOpenTextDocumentParams____x40_Lean_Data_Lsp_Extra___hyg_203____closed__1; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); @@ -1661,7 +1661,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonWai lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonWaitForDiagnosticsParams____x40_Lean_Data_Lsp_Extra___hyg_371____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -1698,7 +1698,7 @@ x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); x_13 = l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonWaitForDiagnosticsParams____x40_Lean_Data_Lsp_Extra___hyg_371____closed__10; -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____spec__1(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -2604,7 +2604,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLea lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLeanFileProgressProcessingInfo____x40_Lean_Data_Lsp_Extra___hyg_710____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__2(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__2(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -2809,7 +2809,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_toJsonLeanF 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; 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_742_(x_2); +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_2); x_4 = l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLeanFileProgressProcessingInfo____x40_Lean_Data_Lsp_Extra___hyg_710____closed__1; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); @@ -3123,7 +3123,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLea lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLeanDidOpenTextDocumentParams____x40_Lean_Data_Lsp_Extra___hyg_203____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -3289,7 +3289,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_toJsonLeanF lean_object* x_2; 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_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_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2440_(x_2); +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2255_(x_2); x_4 = l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLeanDidOpenTextDocumentParams____x40_Lean_Data_Lsp_Extra___hyg_203____closed__1; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); @@ -3468,7 +3468,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPla lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLeanDidOpenTextDocumentParams____x40_Lean_Data_Lsp_Extra___hyg_203____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -3505,7 +3505,7 @@ x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); x_13 = l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPlainGoalParams____x40_Lean_Data_Lsp_Extra___hyg_1059____closed__7; -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____spec__1(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -3588,7 +3588,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_toJsonPlain lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; 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_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_2313_(x_2); +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2128_(x_2); x_4 = l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLeanDidOpenTextDocumentParams____x40_Lean_Data_Lsp_Extra___hyg_203____closed__1; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); @@ -3600,7 +3600,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_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406_(x_8); +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221_(x_8); x_10 = l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPlainGoalParams____x40_Lean_Data_Lsp_Extra___hyg_1059____closed__7; x_11 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_11, 0, x_10); @@ -3918,7 +3918,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPla lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPlainGoal____x40_Lean_Data_Lsp_Extra___hyg_1238____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -4278,7 +4278,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPla lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLeanDidOpenTextDocumentParams____x40_Lean_Data_Lsp_Extra___hyg_203____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -4315,7 +4315,7 @@ x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); x_13 = l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPlainGoalParams____x40_Lean_Data_Lsp_Extra___hyg_1059____closed__7; -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____spec__1(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -4398,7 +4398,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_toJsonPlain lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; 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_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_2313_(x_2); +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2128_(x_2); x_4 = l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLeanDidOpenTextDocumentParams____x40_Lean_Data_Lsp_Extra___hyg_203____closed__1; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); @@ -4410,7 +4410,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_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406_(x_8); +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221_(x_8); x_10 = l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPlainGoalParams____x40_Lean_Data_Lsp_Extra___hyg_1059____closed__7; x_11 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_11, 0, x_10); @@ -4561,7 +4561,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPla lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPlainTermGoal____x40_Lean_Data_Lsp_Extra___hyg_1586____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -4598,7 +4598,7 @@ x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); x_13 = l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLeanFileProgressProcessingInfo____x40_Lean_Data_Lsp_Extra___hyg_710____closed__1; -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__2(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__2(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -4694,7 +4694,7 @@ x_7 = lean_box(0); x_8 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_8, 0, x_1); lean_ctor_set(x_8, 1, x_7); -x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_4); +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_4); x_10 = l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLeanFileProgressProcessingInfo____x40_Lean_Data_Lsp_Extra___hyg_710____closed__1; x_11 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_11, 0, x_10); @@ -4731,7 +4731,7 @@ 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 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_19); +x_25 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_19); x_26 = l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLeanFileProgressProcessingInfo____x40_Lean_Data_Lsp_Extra___hyg_710____closed__1; x_27 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_27, 0, x_26); @@ -4833,7 +4833,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpc { lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonWaitForDiagnosticsParams____x40_Lean_Data_Lsp_Extra___hyg_371____closed__1; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -5626,7 +5626,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpc lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLeanDidOpenTextDocumentParams____x40_Lean_Data_Lsp_Extra___hyg_203____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -5664,7 +5664,7 @@ lean_inc(x_12); lean_dec(x_3); x_13 = l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPlainGoalParams____x40_Lean_Data_Lsp_Extra___hyg_1059____closed__7; lean_inc(x_1); -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____spec__1(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -5879,7 +5879,7 @@ x_2 = lean_ctor_get(x_1, 0); 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_2313_(x_3); +x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2128_(x_3); x_5 = l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLeanDidOpenTextDocumentParams____x40_Lean_Data_Lsp_Extra___hyg_203____closed__1; x_6 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_6, 0, x_5); @@ -5891,7 +5891,7 @@ lean_ctor_set(x_8, 1, x_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_406_(x_9); +x_10 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221_(x_9); x_11 = l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPlainGoalParams____x40_Lean_Data_Lsp_Extra___hyg_1059____closed__7; x_12 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_12, 0, x_11); @@ -6243,7 +6243,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpc lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonWaitForDiagnosticsParams____x40_Lean_Data_Lsp_Extra___hyg_371____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -6622,7 +6622,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpc lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonWaitForDiagnosticsParams____x40_Lean_Data_Lsp_Extra___hyg_371____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -7214,7 +7214,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLin lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_reprLineRange____x40_Lean_Data_Lsp_Extra___hyg_2780____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -7251,7 +7251,7 @@ x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); x_13 = l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_reprLineRange____x40_Lean_Data_Lsp_Extra___hyg_2780____closed__10; -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____spec__1(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; diff --git a/stage0/stdlib/Lean/Data/Lsp/InitShutdown.c b/stage0/stdlib/Lean/Data/Lsp/InitShutdown.c index f9023328c474..89e4465c0a6c 100644 --- a/stage0/stdlib/Lean/Data/Lsp/InitShutdown.c +++ b/stage0/stdlib/Lean/Data/Lsp/InitShutdown.c @@ -44,7 +44,6 @@ LEAN_EXPORT lean_object* l_Lean_Lsp_InitializedParams_toCtorIdx___boxed(lean_obj static lean_object* l___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_fromJsonInitializeResult____x40_Lean_Data_Lsp_InitShutdown___hyg_970____closed__10; static lean_object* l___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_fromJsonClientInfo____x40_Lean_Data_Lsp_InitShutdown___hyg_70____closed__9; LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at_Lean_Lsp_instFromJsonInitializeParams___spec__6(lean_object*, lean_object*); -lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_toJsonInitializeParams____x40_Lean_Data_Lsp_InitShutdown___hyg_523____closed__6; static lean_object* l___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_toJsonInitializeParams____x40_Lean_Data_Lsp_InitShutdown___hyg_523____closed__10; static lean_object* l___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_fromJsonClientInfo____x40_Lean_Data_Lsp_InitShutdown___hyg_70____closed__10; @@ -56,6 +55,7 @@ static lean_object* l___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_fromJsonC LEAN_EXPORT lean_object* l_Lean_Lsp_Trace_hasToJson(uint8_t); static lean_object* l___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_fromJsonInitializationOptions____x40_Lean_Data_Lsp_InitShutdown___hyg_343____closed__3; static lean_object* l___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_toJsonInitializeParams____x40_Lean_Data_Lsp_InitShutdown___hyg_523____closed__12; +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonInitializeResult; static lean_object* l_Lean_Lsp_instToJsonClientInfo___closed__1; static lean_object* l___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_toJsonInitializeParams____x40_Lean_Data_Lsp_InitShutdown___hyg_523____closed__3; @@ -76,7 +76,6 @@ static lean_object* l___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_fromJsonI lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_fromJsonInitializeResult____x40_Lean_Data_Lsp_InitShutdown___hyg_970____closed__4; LEAN_EXPORT lean_object* l_Lean_Lsp_InitializedParams_toCtorIdx(lean_object*); -lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_Trace_toCtorIdx___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at_Lean_Lsp_instFromJsonInitializeParams___spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_fromJsonServerInfo____x40_Lean_Data_Lsp_InitShutdown___hyg_800_(lean_object*); @@ -119,12 +118,14 @@ static lean_object* l___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_toJsonIni static lean_object* l___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_fromJsonClientInfo____x40_Lean_Data_Lsp_InitShutdown___hyg_70____closed__1; static lean_object* l_Lean_Lsp_instFromJsonServerInfo___closed__1; static lean_object* l_Lean_Lsp_Trace_hasToJson___closed__1; +lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2255____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_InitializeParams_editDelay(lean_object*); static lean_object* l___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_fromJsonInitializeResult____x40_Lean_Data_Lsp_InitShutdown___hyg_970____closed__5; lean_object* l___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_fromJsonClientCapabilities____x40_Lean_Data_Lsp_Capabilities___hyg_1171_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at_Lean_Lsp_instFromJsonInitializeParams___spec__5___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at_Lean_Lsp_instFromJsonInitializeParams___spec__3___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_fromJsonInitializeResult____x40_Lean_Data_Lsp_InitShutdown___hyg_970____closed__9; +lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____spec__2(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_toJsonInitializeParams____x40_Lean_Data_Lsp_InitShutdown___hyg_523____closed__9; static lean_object* l___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_toJsonInitializeResult____x40_Lean_Data_Lsp_InitShutdown___hyg_928____closed__1; static lean_object* l___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_toJsonClientInfo____x40_Lean_Data_Lsp_InitShutdown___hyg_28____closed__3; @@ -149,7 +150,6 @@ static lean_object* l___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_fromJsonC LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_toJsonInitializeParams____x40_Lean_Data_Lsp_InitShutdown___hyg_523____spec__4___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_toJsonInitializeParams____x40_Lean_Data_Lsp_InitShutdown___hyg_523____closed__2; static lean_object* l___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_toJsonClientInfo____x40_Lean_Data_Lsp_InitShutdown___hyg_28____closed__2; -lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____spec__2(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_fromJsonInitializeResult____x40_Lean_Data_Lsp_InitShutdown___hyg_970____closed__3; LEAN_EXPORT lean_object* l_Lean_Lsp_InitializedParams_noConfusion___rarg___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_fromJsonInitializeResult____x40_Lean_Data_Lsp_InitShutdown___hyg_970____spec__2(lean_object*, lean_object*); @@ -165,8 +165,8 @@ static lean_object* l___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_fromJsonC LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_toJsonServerInfo____x40_Lean_Data_Lsp_InitShutdown___hyg_758_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_InitializedParams_noConfusion___rarg(lean_object*); static lean_object* l_Lean_Json_getObjValAs_x3f___at_Lean_Lsp_instFromJsonInitializeParams___spec__6___closed__2; +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____spec__1(lean_object*, lean_object*); size_t lean_usize_add(size_t, size_t); -lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2440____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_fromJsonInitializeResult____x40_Lean_Data_Lsp_InitShutdown___hyg_970_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_Trace_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_toJsonInitializeResult____x40_Lean_Data_Lsp_InitShutdown___hyg_928_(lean_object*); @@ -253,7 +253,7 @@ x_8 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_8, 0, x_1); lean_ctor_set(x_8, 1, x_7); x_9 = l___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_toJsonClientInfo____x40_Lean_Data_Lsp_InitShutdown___hyg_28____closed__2; -x_10 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(x_9, x_4); +x_10 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____spec__2(x_9, x_4); x_11 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_11, 0, x_10); lean_ctor_set(x_11, 1, x_7); @@ -284,7 +284,7 @@ x_22 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_22, 0, x_20); lean_ctor_set(x_22, 1, x_21); x_23 = l___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_toJsonClientInfo____x40_Lean_Data_Lsp_InitShutdown___hyg_28____closed__2; -x_24 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(x_23, x_17); +x_24 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____spec__2(x_23, x_17); x_25 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_25, 0, x_24); lean_ctor_set(x_25, 1, x_21); @@ -535,7 +535,7 @@ x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); x_13 = l___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_toJsonClientInfo____x40_Lean_Data_Lsp_InitShutdown___hyg_28____closed__2; -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____spec__2(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____spec__2(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -934,7 +934,7 @@ lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_obj x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); x_3 = l___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_toJsonInitializationOptions____x40_Lean_Data_Lsp_InitShutdown___hyg_311____closed__1; -x_4 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2440____spec__1(x_3, x_2); +x_4 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2255____spec__1(x_3, x_2); x_5 = lean_ctor_get(x_1, 1); lean_inc(x_5); lean_dec(x_1); @@ -1114,7 +1114,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_from lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_toJsonInitializationOptions____x40_Lean_Data_Lsp_InitShutdown___hyg_311____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -1507,7 +1507,7 @@ x_7 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_toJs x_8 = lean_ctor_get(x_1, 2); lean_inc(x_8); x_9 = l___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_toJsonInitializeParams____x40_Lean_Data_Lsp_InitShutdown___hyg_523____closed__3; -x_10 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(x_9, x_8); +x_10 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____spec__2(x_9, x_8); x_11 = lean_ctor_get(x_1, 3); lean_inc(x_11); x_12 = l___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_toJsonInitializeParams____x40_Lean_Data_Lsp_InitShutdown___hyg_523____closed__4; @@ -4191,7 +4191,7 @@ x_8 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_8, 0, x_1); lean_ctor_set(x_8, 1, x_7); x_9 = l___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_toJsonClientInfo____x40_Lean_Data_Lsp_InitShutdown___hyg_28____closed__2; -x_10 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(x_9, x_4); +x_10 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____spec__2(x_9, x_4); x_11 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_11, 0, x_10); lean_ctor_set(x_11, 1, x_7); @@ -4222,7 +4222,7 @@ x_22 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_22, 0, x_20); lean_ctor_set(x_22, 1, x_21); x_23 = l___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_toJsonClientInfo____x40_Lean_Data_Lsp_InitShutdown___hyg_28____closed__2; -x_24 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(x_23, x_17); +x_24 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____spec__2(x_23, x_17); x_25 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_25, 0, x_24); lean_ctor_set(x_25, 1, x_21); @@ -4375,7 +4375,7 @@ x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); x_13 = l___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_toJsonClientInfo____x40_Lean_Data_Lsp_InitShutdown___hyg_28____closed__2; -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____spec__2(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____spec__2(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; diff --git a/stage0/stdlib/Lean/Data/Lsp/Internal.c b/stage0/stdlib/Lean/Data/Lsp/Internal.c index aa45fbde2173..32d55b6945f8 100644 --- a/stage0/stdlib/Lean/Data/Lsp/Internal.c +++ b/stage0/stdlib/Lean/Data/Lsp/Internal.c @@ -30,6 +30,7 @@ static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanI static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1923____closed__4; lean_object* l_Lean_Json_mkObj(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_RefIdent_fromJsonRefIdentJsonRepr____x40_Lean_Data_Lsp_Internal___hyg_270____lambda__1___closed__2; +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Lsp_instToJsonModuleRefs___spec__1(size_t, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_RefIdent_fromJsonRefIdentJsonRepr____x40_Lean_Data_Lsp_Internal___hyg_270____lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Lsp_RefInfo_instInhabitedLocation___closed__2; @@ -58,12 +59,12 @@ LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonLeanIleanInfoParams; static lean_object* l_Lean_Lsp_RefIdent_instFromJsonRefIdentJsonRepr___closed__1; uint64_t lean_string_hash(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonModuleRefs(lean_object*); -lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instToJsonLeanImportClosureParams___closed__1; extern lean_object* l_Lean_instInhabitedJson; lean_object* l_Nat_nextPowerOfTwo_go(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanStaleDependencyParams____x40_Lean_Data_Lsp_Internal___hyg_2211____closed__9; uint8_t lean_string_dec_eq(lean_object*, lean_object*); +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonLeanStaleDependencyParams; static lean_object* l_Lean_Lsp_instInhabitedRefIdent___closed__1; static lean_object* l_Lean_Lsp_RefIdent_instToJsonRefIdentJsonRepr___closed__1; @@ -75,7 +76,6 @@ static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanI lean_object* l_Lean_Json_getStr_x3f(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_RefInfo_toJsonParentDecl____x40_Lean_Data_Lsp_Internal___hyg_723____closed__2; lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); -lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(size_t, size_t, lean_object*); lean_object* l_List_appendTR___rarg(lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Lsp_instFromJsonRefInfo___spec__2___closed__2; static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanStaleDependencyParams____x40_Lean_Data_Lsp_Internal___hyg_2211____closed__7; @@ -104,14 +104,12 @@ lean_object* l_outOfBounds___rarg(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_toJsonLeanImportClosureParams____x40_Lean_Data_Lsp_Internal___hyg_2161____spec__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Lsp_instToJsonRefInfo___spec__2(size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Lsp_instToJsonRefInfo___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonRefInfo(lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1923____spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instInhabitedRefIdent___closed__2; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Lsp_instToJsonModuleRefs___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_RBNode_foldM___at_Lean_Lsp_instFromJsonModuleRefs___spec__6___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_RefIdent_fromJsonRefIdentJsonRepr____x40_Lean_Data_Lsp_Internal___hyg_270____lambda__1___boxed(lean_object*); -lean_object* l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at_Lean_Lsp_instFromJsonRefInfo___spec__1___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanStaleDependencyParams____x40_Lean_Data_Lsp_Internal___hyg_2211____closed__3; LEAN_EXPORT lean_object* l_List_mapTR_loop___at___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_toJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_2029____spec__5(lean_object*, lean_object*, lean_object*); @@ -125,6 +123,7 @@ LEAN_EXPORT lean_object* l_Lean_Lsp_RefIdent_fromJson_x3f(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1923____closed__16; static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_RefIdent_fromJsonRefIdentJsonRepr____x40_Lean_Data_Lsp_Internal___hyg_270____closed__5; LEAN_EXPORT lean_object* l_List_mapTR_loop___at___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_toJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_2029____spec__2(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__1(lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_beqRefIdent____x40_Lean_Data_Lsp_Internal___hyg_45_(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1923____closed__9; static lean_object* l_Array_mapMUnsafe_map___at_Lean_Lsp_instFromJsonRefInfo___spec__2___closed__1; @@ -153,6 +152,7 @@ LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Lsp_instToJsonModuleRefs___ lean_object* l_Lean_Json_parseTagged(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanStaleDependencyParams____x40_Lean_Data_Lsp_Internal___hyg_2211_(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Lsp_instFromJsonRefInfo___spec__2(size_t, size_t, lean_object*); +lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(size_t, size_t, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanStaleDependencyParams____x40_Lean_Data_Lsp_Internal___hyg_2211____closed__2; static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_RefIdent_fromJsonRefIdentJsonRepr____x40_Lean_Data_Lsp_Internal___hyg_270____lambda__2___closed__1; static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanStaleDependencyParams____x40_Lean_Data_Lsp_Internal___hyg_2211____closed__1; @@ -196,6 +196,7 @@ static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanI static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanStaleDependencyParams____x40_Lean_Data_Lsp_Internal___hyg_2211____closed__8; LEAN_EXPORT uint8_t l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Lsp_instFromJsonModuleRefs___spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1923____closed__19; +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2(lean_object*, lean_object*); size_t lean_usize_add(size_t, size_t); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Lsp_instFromJsonRefInfo___spec__3___lambda__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instToJsonRefInfo___closed__1; @@ -203,13 +204,12 @@ LEAN_EXPORT lean_object* l_Lean_Lsp_RefIdent_instFromJson; lean_object* lean_array_uget(lean_object*, size_t); size_t lean_array_size(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Lsp_instFromJsonRefInfo___spec__2___boxed(lean_object*, lean_object*, lean_object*); +lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(size_t, size_t, lean_object*); static lean_object* l_Lean_Lsp_RefInfo_instInhabitedLocation___closed__1; static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_RefIdent_fromJsonRefIdentJsonRepr____x40_Lean_Data_Lsp_Internal___hyg_270____lambda__2___closed__3; static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1923____closed__2; static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_RefIdent_fromJsonRefIdentJsonRepr____x40_Lean_Data_Lsp_Internal___hyg_270____closed__3; -lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(size_t, size_t, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanImportClosureParams____x40_Lean_Data_Lsp_Internal___hyg_2094____closed__9; -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(lean_object*); static lean_object* l_Lean_Lsp_RefInfo_instInhabitedLocation___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanImportClosureParams____x40_Lean_Data_Lsp_Internal___hyg_2094_(lean_object*); lean_object* lean_string_append(lean_object*, lean_object*); @@ -226,13 +226,13 @@ lean_object* lean_nat_add(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Lsp_instFromJsonModuleRefs___spec__5(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_toJsonLeanImportClosureParams____x40_Lean_Data_Lsp_Internal___hyg_2161_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1923____closed__10; -lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1(lean_object*, lean_object*); lean_object* l_Lean_Json_pretty(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_RefIdent_instToJson; LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanImportClosureParams____x40_Lean_Data_Lsp_Internal___hyg_2094____spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanImportClosureParams____x40_Lean_Data_Lsp_Internal___hyg_2094____spec__2___boxed(lean_object*, lean_object*, lean_object*); lean_object* lean_array_uset(lean_object*, size_t, lean_object*); lean_object* l___private_Init_Data_Repr_0__Nat_reprFast(lean_object*); +lean_object* l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instToJsonLeanIleanInfoParams___closed__1; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Lsp_instToJsonRefInfo___spec__3(size_t, size_t, size_t, lean_object*); size_t lean_usize_land(size_t, size_t); @@ -1780,7 +1780,7 @@ x_8 = lean_box(0); 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_742_(x_3); +x_10 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_3); x_11 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_RefInfo_toJsonParentDecl____x40_Lean_Data_Lsp_Internal___hyg_723____closed__2; x_12 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_12, 0, x_11); @@ -1788,7 +1788,7 @@ 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_8); -x_14 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_4); +x_14 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_4); x_15 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_RefInfo_toJsonParentDecl____x40_Lean_Data_Lsp_Internal___hyg_723____closed__3; x_16 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_16, 0, x_15); @@ -1806,7 +1806,7 @@ x_20 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_20, 0, x_9); lean_ctor_set(x_20, 1, x_19); x_21 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_RefInfo_toJsonParentDecl____x40_Lean_Data_Lsp_Internal___hyg_723____closed__4; -x_22 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_20, x_21); +x_22 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_20, x_21); x_23 = l_Lean_Json_mkObj(x_22); return x_23; } @@ -2873,7 +2873,7 @@ x_7 = lean_unsigned_to_nat(0u); x_8 = lean_array_uset(x_4, x_3, x_7); x_9 = lean_array_mk(x_6); x_10 = lean_array_size(x_9); -x_11 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_10, x_1, x_9); +x_11 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_10, x_1, x_9); x_12 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_12, 0, x_11); x_13 = 1; @@ -2996,7 +2996,7 @@ lean_object* x_32; lean_object* x_33; size_t x_34; lean_object* x_35; lean_objec x_32 = l_List_appendTR___rarg(x_30, x_12); x_33 = lean_array_mk(x_32); x_34 = lean_array_size(x_33); -x_35 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_34, x_5, x_33); +x_35 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_34, x_5, x_33); lean_ctor_set_tag(x_2, 4); lean_ctor_set(x_2, 0, x_35); x_36 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -3094,7 +3094,7 @@ x_69 = l_List_appendTR___rarg(x_68, x_66); x_70 = l_List_appendTR___rarg(x_30, x_69); x_71 = lean_array_mk(x_70); x_72 = lean_array_size(x_71); -x_73 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_72, x_5, x_71); +x_73 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_72, x_5, x_71); lean_ctor_set_tag(x_2, 4); lean_ctor_set(x_2, 0, x_73); x_74 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -3138,7 +3138,7 @@ x_88 = l_List_appendTR___rarg(x_87, x_85); x_89 = l_List_appendTR___rarg(x_30, x_88); x_90 = lean_array_mk(x_89); x_91 = lean_array_size(x_90); -x_92 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_91, x_5, x_90); +x_92 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_91, x_5, x_90); lean_ctor_set_tag(x_2, 4); lean_ctor_set(x_2, 0, x_92); x_93 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -3198,7 +3198,7 @@ x_109 = l_List_appendTR___rarg(x_108, x_106); x_110 = l_List_appendTR___rarg(x_30, x_109); x_111 = lean_array_mk(x_110); x_112 = lean_array_size(x_111); -x_113 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_112, x_5, x_111); +x_113 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_112, x_5, x_111); lean_ctor_set_tag(x_2, 4); lean_ctor_set(x_2, 0, x_113); x_114 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -3295,7 +3295,7 @@ x_141 = l_List_appendTR___rarg(x_140, x_138); x_142 = l_List_appendTR___rarg(x_30, x_141); x_143 = lean_array_mk(x_142); x_144 = lean_array_size(x_143); -x_145 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_144, x_5, x_143); +x_145 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_144, x_5, x_143); lean_ctor_set_tag(x_2, 4); lean_ctor_set(x_2, 0, x_145); x_146 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -3407,7 +3407,7 @@ x_175 = l_List_appendTR___rarg(x_174, x_172); x_176 = l_List_appendTR___rarg(x_30, x_175); x_177 = lean_array_mk(x_176); x_178 = lean_array_size(x_177); -x_179 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_178, x_5, x_177); +x_179 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_178, x_5, x_177); lean_ctor_set_tag(x_2, 4); lean_ctor_set(x_2, 0, x_179); x_180 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -3548,7 +3548,7 @@ x_217 = l_List_appendTR___rarg(x_216, x_214); x_218 = l_List_appendTR___rarg(x_30, x_217); x_219 = lean_array_mk(x_218); x_220 = lean_array_size(x_219); -x_221 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_220, x_5, x_219); +x_221 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_220, x_5, x_219); lean_ctor_set_tag(x_2, 4); lean_ctor_set(x_2, 0, x_221); x_222 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -3595,7 +3595,7 @@ lean_object* x_235; lean_object* x_236; size_t x_237; lean_object* x_238; lean_o x_235 = l_List_appendTR___rarg(x_233, x_12); x_236 = lean_array_mk(x_235); x_237 = lean_array_size(x_236); -x_238 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_237, x_5, x_236); +x_238 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_237, x_5, x_236); lean_ctor_set_tag(x_2, 4); lean_ctor_set(x_2, 0, x_238); x_239 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -3746,7 +3746,7 @@ x_277 = l_List_appendTR___rarg(x_276, x_274); x_278 = l_List_appendTR___rarg(x_233, x_277); x_279 = lean_array_mk(x_278); x_280 = lean_array_size(x_279); -x_281 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_280, x_5, x_279); +x_281 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_280, x_5, x_279); lean_ctor_set_tag(x_2, 4); lean_ctor_set(x_2, 0, x_281); x_282 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -3808,7 +3808,7 @@ lean_object* x_297; lean_object* x_298; size_t x_299; lean_object* x_300; lean_o x_297 = l_List_appendTR___rarg(x_295, x_12); x_298 = lean_array_mk(x_297); x_299 = lean_array_size(x_298); -x_300 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_299, x_5, x_298); +x_300 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_299, x_5, x_298); lean_ctor_set_tag(x_2, 4); lean_ctor_set(x_2, 0, x_300); x_301 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -3959,7 +3959,7 @@ x_339 = l_List_appendTR___rarg(x_338, x_336); x_340 = l_List_appendTR___rarg(x_295, x_339); x_341 = lean_array_mk(x_340); x_342 = lean_array_size(x_341); -x_343 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_342, x_5, x_341); +x_343 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_342, x_5, x_341); lean_ctor_set_tag(x_2, 4); lean_ctor_set(x_2, 0, x_343); x_344 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -4043,7 +4043,7 @@ lean_object* x_364; lean_object* x_365; size_t x_366; lean_object* x_367; lean_o x_364 = l_List_appendTR___rarg(x_362, x_12); x_365 = lean_array_mk(x_364); x_366 = lean_array_size(x_365); -x_367 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_366, x_5, x_365); +x_367 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_366, x_5, x_365); x_368 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_368, 0, x_367); x_369 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -4194,7 +4194,7 @@ x_407 = l_List_appendTR___rarg(x_406, x_404); x_408 = l_List_appendTR___rarg(x_362, x_407); x_409 = lean_array_mk(x_408); x_410 = lean_array_size(x_409); -x_411 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_410, x_5, x_409); +x_411 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_410, x_5, x_409); x_412 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_412, 0, x_411); x_413 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -4310,7 +4310,7 @@ x_22 = lean_unsigned_to_nat(0u); x_23 = lean_array_uset(x_3, x_2, x_22); x_24 = lean_array_size(x_21); x_25 = 0; -x_26 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_24, x_25, x_21); +x_26 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_24, x_25, x_21); x_27 = lean_ctor_get(x_26, 0); lean_inc(x_27); lean_dec(x_26); @@ -6344,7 +6344,7 @@ LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonRefInfo(lean_object* x_1) { lean_object* x_2; lean_object* x_3; x_2 = l_Lean_Lsp_instToJsonRefInfo___closed__2; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -6651,7 +6651,7 @@ x_7 = lean_unsigned_to_nat(0u); x_8 = lean_array_uset(x_4, x_3, x_7); x_9 = lean_array_mk(x_6); x_10 = lean_array_size(x_9); -x_11 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_10, x_1, x_9); +x_11 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_10, x_1, x_9); x_12 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_12, 0, x_11); x_13 = 1; @@ -6800,7 +6800,7 @@ lean_object* x_43; lean_object* x_44; size_t x_45; lean_object* x_46; lean_objec x_43 = l_List_appendTR___rarg(x_41, x_21); x_44 = lean_array_mk(x_43); x_45 = lean_array_size(x_44); -x_46 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_45, x_15, x_44); +x_46 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_45, x_15, x_44); lean_ctor_set_tag(x_12, 4); lean_ctor_set(x_12, 0, x_46); x_47 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -6906,7 +6906,7 @@ x_83 = l_List_appendTR___rarg(x_82, x_80); x_84 = l_List_appendTR___rarg(x_41, x_83); x_85 = lean_array_mk(x_84); x_86 = lean_array_size(x_85); -x_87 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_86, x_15, x_85); +x_87 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_86, x_15, x_85); lean_ctor_set_tag(x_12, 4); lean_ctor_set(x_12, 0, x_87); x_88 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -6958,7 +6958,7 @@ x_105 = l_List_appendTR___rarg(x_104, x_102); x_106 = l_List_appendTR___rarg(x_41, x_105); x_107 = lean_array_mk(x_106); x_108 = lean_array_size(x_107); -x_109 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_108, x_15, x_107); +x_109 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_108, x_15, x_107); lean_ctor_set_tag(x_12, 4); lean_ctor_set(x_12, 0, x_109); x_110 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -7026,7 +7026,7 @@ x_129 = l_List_appendTR___rarg(x_128, x_126); x_130 = l_List_appendTR___rarg(x_41, x_129); x_131 = lean_array_mk(x_130); x_132 = lean_array_size(x_131); -x_133 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_132, x_15, x_131); +x_133 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_132, x_15, x_131); lean_ctor_set_tag(x_12, 4); lean_ctor_set(x_12, 0, x_133); x_134 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -7131,7 +7131,7 @@ x_164 = l_List_appendTR___rarg(x_163, x_161); x_165 = l_List_appendTR___rarg(x_41, x_164); x_166 = lean_array_mk(x_165); x_167 = lean_array_size(x_166); -x_168 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_167, x_15, x_166); +x_168 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_167, x_15, x_166); lean_ctor_set_tag(x_12, 4); lean_ctor_set(x_12, 0, x_168); x_169 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -7251,7 +7251,7 @@ x_201 = l_List_appendTR___rarg(x_200, x_198); x_202 = l_List_appendTR___rarg(x_41, x_201); x_203 = lean_array_mk(x_202); x_204 = lean_array_size(x_203); -x_205 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_204, x_15, x_203); +x_205 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_204, x_15, x_203); lean_ctor_set_tag(x_12, 4); lean_ctor_set(x_12, 0, x_205); x_206 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -7400,7 +7400,7 @@ x_246 = l_List_appendTR___rarg(x_245, x_243); x_247 = l_List_appendTR___rarg(x_41, x_246); x_248 = lean_array_mk(x_247); x_249 = lean_array_size(x_248); -x_250 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_249, x_15, x_248); +x_250 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_249, x_15, x_248); lean_ctor_set_tag(x_12, 4); lean_ctor_set(x_12, 0, x_250); x_251 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -7455,7 +7455,7 @@ lean_object* x_267; lean_object* x_268; size_t x_269; lean_object* x_270; lean_o x_267 = l_List_appendTR___rarg(x_265, x_21); x_268 = lean_array_mk(x_267); x_269 = lean_array_size(x_268); -x_270 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_269, x_15, x_268); +x_270 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_269, x_15, x_268); lean_ctor_set_tag(x_12, 4); lean_ctor_set(x_12, 0, x_270); x_271 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -7614,7 +7614,7 @@ x_312 = l_List_appendTR___rarg(x_311, x_309); x_313 = l_List_appendTR___rarg(x_265, x_312); x_314 = lean_array_mk(x_313); x_315 = lean_array_size(x_314); -x_316 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_315, x_15, x_314); +x_316 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_315, x_15, x_314); lean_ctor_set_tag(x_12, 4); lean_ctor_set(x_12, 0, x_316); x_317 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -7684,7 +7684,7 @@ lean_object* x_335; lean_object* x_336; size_t x_337; lean_object* x_338; lean_o x_335 = l_List_appendTR___rarg(x_333, x_21); x_336 = lean_array_mk(x_335); x_337 = lean_array_size(x_336); -x_338 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_337, x_15, x_336); +x_338 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_337, x_15, x_336); lean_ctor_set_tag(x_12, 4); lean_ctor_set(x_12, 0, x_338); x_339 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -7843,7 +7843,7 @@ x_380 = l_List_appendTR___rarg(x_379, x_377); x_381 = l_List_appendTR___rarg(x_333, x_380); x_382 = lean_array_mk(x_381); x_383 = lean_array_size(x_382); -x_384 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_383, x_15, x_382); +x_384 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_383, x_15, x_382); lean_ctor_set_tag(x_12, 4); lean_ctor_set(x_12, 0, x_384); x_385 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -7935,7 +7935,7 @@ lean_object* x_408; lean_object* x_409; size_t x_410; lean_object* x_411; lean_o x_408 = l_List_appendTR___rarg(x_406, x_21); x_409 = lean_array_mk(x_408); x_410 = lean_array_size(x_409); -x_411 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_410, x_15, x_409); +x_411 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_410, x_15, x_409); x_412 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_412, 0, x_411); x_413 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -8094,7 +8094,7 @@ x_454 = l_List_appendTR___rarg(x_453, x_451); x_455 = l_List_appendTR___rarg(x_406, x_454); x_456 = lean_array_mk(x_455); x_457 = lean_array_size(x_456); -x_458 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_457, x_15, x_456); +x_458 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_457, x_15, x_456); x_459 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_459, 0, x_458); x_460 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -8241,7 +8241,7 @@ lean_object* x_506; lean_object* x_507; size_t x_508; lean_object* x_509; lean_o x_506 = l_List_appendTR___rarg(x_504, x_482); x_507 = lean_array_mk(x_506); x_508 = lean_array_size(x_507); -x_509 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_508, x_475, x_507); +x_509 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_508, x_475, x_507); if (lean_is_scalar(x_490)) { x_510 = lean_alloc_ctor(4, 1, 0); } else { @@ -8405,7 +8405,7 @@ x_552 = l_List_appendTR___rarg(x_551, x_549); x_553 = l_List_appendTR___rarg(x_504, x_552); x_554 = lean_array_mk(x_553); x_555 = lean_array_size(x_554); -x_556 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_555, x_475, x_554); +x_556 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_555, x_475, x_554); if (lean_is_scalar(x_490)) { x_557 = lean_alloc_ctor(4, 1, 0); } else { @@ -8573,7 +8573,7 @@ lean_object* x_607; lean_object* x_608; size_t x_609; lean_object* x_610; lean_o x_607 = l_List_appendTR___rarg(x_605, x_582); x_608 = lean_array_mk(x_607); x_609 = lean_array_size(x_608); -x_610 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_609, x_575, x_608); +x_610 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_609, x_575, x_608); if (lean_is_scalar(x_591)) { x_611 = lean_alloc_ctor(4, 1, 0); } else { @@ -8737,7 +8737,7 @@ x_653 = l_List_appendTR___rarg(x_652, x_650); x_654 = l_List_appendTR___rarg(x_605, x_653); x_655 = lean_array_mk(x_654); x_656 = lean_array_size(x_655); -x_657 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_656, x_575, x_655); +x_657 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_656, x_575, x_655); if (lean_is_scalar(x_591)) { x_658 = lean_alloc_ctor(4, 1, 0); } else { @@ -9210,7 +9210,7 @@ lean_inc(x_23); lean_dec(x_19); x_100 = l_Lean_Lsp_instToJsonRefInfo___closed__2; lean_inc(x_6); -x_101 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1(x_6, x_100); +x_101 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2(x_6, x_100); if (lean_obj_tag(x_101) == 0) { uint8_t x_102; @@ -10101,7 +10101,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJson lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1923____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -10251,7 +10251,7 @@ x_7 = lean_unsigned_to_nat(0u); x_8 = lean_array_uset(x_4, x_3, x_7); x_9 = lean_array_mk(x_6); x_10 = lean_array_size(x_9); -x_11 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_10, x_1, x_9); +x_11 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_10, x_1, x_9); x_12 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_12, 0, x_11); x_13 = 1; @@ -10375,7 +10375,7 @@ lean_object* x_44; lean_object* x_45; size_t x_46; lean_object* x_47; lean_objec x_44 = l_List_appendTR___rarg(x_42, x_39); x_45 = lean_array_mk(x_44); x_46 = lean_array_size(x_45); -x_47 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_46, x_16, x_45); +x_47 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_46, x_16, x_45); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_47); x_48 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -10481,7 +10481,7 @@ x_84 = l_List_appendTR___rarg(x_83, x_81); x_85 = l_List_appendTR___rarg(x_42, x_84); x_86 = lean_array_mk(x_85); x_87 = lean_array_size(x_86); -x_88 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_87, x_16, x_86); +x_88 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_87, x_16, x_86); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_88); x_89 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -10533,7 +10533,7 @@ x_106 = l_List_appendTR___rarg(x_105, x_103); x_107 = l_List_appendTR___rarg(x_42, x_106); x_108 = lean_array_mk(x_107); x_109 = lean_array_size(x_108); -x_110 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_109, x_16, x_108); +x_110 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_109, x_16, x_108); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_110); x_111 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -10601,7 +10601,7 @@ x_130 = l_List_appendTR___rarg(x_129, x_127); x_131 = l_List_appendTR___rarg(x_42, x_130); x_132 = lean_array_mk(x_131); x_133 = lean_array_size(x_132); -x_134 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_133, x_16, x_132); +x_134 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_133, x_16, x_132); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_134); x_135 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -10706,7 +10706,7 @@ x_165 = l_List_appendTR___rarg(x_164, x_162); x_166 = l_List_appendTR___rarg(x_42, x_165); x_167 = lean_array_mk(x_166); x_168 = lean_array_size(x_167); -x_169 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_168, x_16, x_167); +x_169 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_168, x_16, x_167); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_169); x_170 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -10826,7 +10826,7 @@ x_202 = l_List_appendTR___rarg(x_201, x_199); x_203 = l_List_appendTR___rarg(x_42, x_202); x_204 = lean_array_mk(x_203); x_205 = lean_array_size(x_204); -x_206 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_205, x_16, x_204); +x_206 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_205, x_16, x_204); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_206); x_207 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -10975,7 +10975,7 @@ x_247 = l_List_appendTR___rarg(x_246, x_244); x_248 = l_List_appendTR___rarg(x_42, x_247); x_249 = lean_array_mk(x_248); x_250 = lean_array_size(x_249); -x_251 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_250, x_16, x_249); +x_251 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_250, x_16, x_249); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_251); x_252 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -11031,7 +11031,7 @@ lean_object* x_269; lean_object* x_270; size_t x_271; lean_object* x_272; lean_o x_269 = l_List_appendTR___rarg(x_267, x_263); x_270 = lean_array_mk(x_269); x_271 = lean_array_size(x_270); -x_272 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_271, x_16, x_270); +x_272 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_271, x_16, x_270); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_272); x_273 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -11190,7 +11190,7 @@ x_314 = l_List_appendTR___rarg(x_313, x_311); x_315 = l_List_appendTR___rarg(x_267, x_314); x_316 = lean_array_mk(x_315); x_317 = lean_array_size(x_316); -x_318 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_317, x_16, x_316); +x_318 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_317, x_16, x_316); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_318); x_319 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -11261,7 +11261,7 @@ lean_object* x_338; lean_object* x_339; size_t x_340; lean_object* x_341; lean_o x_338 = l_List_appendTR___rarg(x_336, x_331); x_339 = lean_array_mk(x_338); x_340 = lean_array_size(x_339); -x_341 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_340, x_16, x_339); +x_341 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_340, x_16, x_339); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_341); x_342 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -11420,7 +11420,7 @@ x_383 = l_List_appendTR___rarg(x_382, x_380); x_384 = l_List_appendTR___rarg(x_336, x_383); x_385 = lean_array_mk(x_384); x_386 = lean_array_size(x_385); -x_387 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_386, x_16, x_385); +x_387 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_386, x_16, x_385); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_387); x_388 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -11513,7 +11513,7 @@ lean_object* x_412; lean_object* x_413; size_t x_414; lean_object* x_415; lean_o x_412 = l_List_appendTR___rarg(x_410, x_405); x_413 = lean_array_mk(x_412); x_414 = lean_array_size(x_413); -x_415 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_414, x_16, x_413); +x_415 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_414, x_16, x_413); x_416 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_416, 0, x_415); x_417 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -11672,7 +11672,7 @@ x_458 = l_List_appendTR___rarg(x_457, x_455); x_459 = l_List_appendTR___rarg(x_410, x_458); x_460 = lean_array_mk(x_459); x_461 = lean_array_size(x_460); -x_462 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_461, x_16, x_460); +x_462 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_461, x_16, x_460); x_463 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_463, 0, x_462); x_464 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -11820,7 +11820,7 @@ lean_object* x_510; lean_object* x_511; size_t x_512; lean_object* x_513; lean_o x_510 = l_List_appendTR___rarg(x_508, x_503); x_511 = lean_array_mk(x_510); x_512 = lean_array_size(x_511); -x_513 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_512, x_479, x_511); +x_513 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_512, x_479, x_511); if (lean_is_scalar(x_493)) { x_514 = lean_alloc_ctor(4, 1, 0); } else { @@ -11984,7 +11984,7 @@ x_556 = l_List_appendTR___rarg(x_555, x_553); x_557 = l_List_appendTR___rarg(x_508, x_556); x_558 = lean_array_mk(x_557); x_559 = lean_array_size(x_558); -x_560 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_559, x_479, x_558); +x_560 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_559, x_479, x_558); if (lean_is_scalar(x_493)) { x_561 = lean_alloc_ctor(4, 1, 0); } else { @@ -12153,7 +12153,7 @@ lean_object* x_611; lean_object* x_612; size_t x_613; lean_object* x_614; lean_o x_611 = l_List_appendTR___rarg(x_609, x_604); x_612 = lean_array_mk(x_611); x_613 = lean_array_size(x_612); -x_614 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_613, x_579, x_612); +x_614 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_613, x_579, x_612); if (lean_is_scalar(x_594)) { x_615 = lean_alloc_ctor(4, 1, 0); } else { @@ -12317,7 +12317,7 @@ x_657 = l_List_appendTR___rarg(x_656, x_654); x_658 = l_List_appendTR___rarg(x_609, x_657); x_659 = lean_array_mk(x_658); x_660 = lean_array_size(x_659); -x_661 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_660, x_579, x_659); +x_661 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_660, x_579, x_659); if (lean_is_scalar(x_594)) { x_662 = lean_alloc_ctor(4, 1, 0); } else { @@ -12460,7 +12460,7 @@ lean_object* x_44; lean_object* x_45; size_t x_46; lean_object* x_47; lean_objec x_44 = l_List_appendTR___rarg(x_42, x_39); x_45 = lean_array_mk(x_44); x_46 = lean_array_size(x_45); -x_47 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_46, x_16, x_45); +x_47 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_46, x_16, x_45); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_47); x_48 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -12566,7 +12566,7 @@ x_84 = l_List_appendTR___rarg(x_83, x_81); x_85 = l_List_appendTR___rarg(x_42, x_84); x_86 = lean_array_mk(x_85); x_87 = lean_array_size(x_86); -x_88 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_87, x_16, x_86); +x_88 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_87, x_16, x_86); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_88); x_89 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -12618,7 +12618,7 @@ x_106 = l_List_appendTR___rarg(x_105, x_103); x_107 = l_List_appendTR___rarg(x_42, x_106); x_108 = lean_array_mk(x_107); x_109 = lean_array_size(x_108); -x_110 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_109, x_16, x_108); +x_110 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_109, x_16, x_108); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_110); x_111 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -12686,7 +12686,7 @@ x_130 = l_List_appendTR___rarg(x_129, x_127); x_131 = l_List_appendTR___rarg(x_42, x_130); x_132 = lean_array_mk(x_131); x_133 = lean_array_size(x_132); -x_134 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_133, x_16, x_132); +x_134 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_133, x_16, x_132); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_134); x_135 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -12791,7 +12791,7 @@ x_165 = l_List_appendTR___rarg(x_164, x_162); x_166 = l_List_appendTR___rarg(x_42, x_165); x_167 = lean_array_mk(x_166); x_168 = lean_array_size(x_167); -x_169 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_168, x_16, x_167); +x_169 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_168, x_16, x_167); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_169); x_170 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -12911,7 +12911,7 @@ x_202 = l_List_appendTR___rarg(x_201, x_199); x_203 = l_List_appendTR___rarg(x_42, x_202); x_204 = lean_array_mk(x_203); x_205 = lean_array_size(x_204); -x_206 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_205, x_16, x_204); +x_206 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_205, x_16, x_204); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_206); x_207 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -13060,7 +13060,7 @@ x_247 = l_List_appendTR___rarg(x_246, x_244); x_248 = l_List_appendTR___rarg(x_42, x_247); x_249 = lean_array_mk(x_248); x_250 = lean_array_size(x_249); -x_251 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_250, x_16, x_249); +x_251 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_250, x_16, x_249); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_251); x_252 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -13116,7 +13116,7 @@ lean_object* x_269; lean_object* x_270; size_t x_271; lean_object* x_272; lean_o x_269 = l_List_appendTR___rarg(x_267, x_263); x_270 = lean_array_mk(x_269); x_271 = lean_array_size(x_270); -x_272 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_271, x_16, x_270); +x_272 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_271, x_16, x_270); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_272); x_273 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -13275,7 +13275,7 @@ x_314 = l_List_appendTR___rarg(x_313, x_311); x_315 = l_List_appendTR___rarg(x_267, x_314); x_316 = lean_array_mk(x_315); x_317 = lean_array_size(x_316); -x_318 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_317, x_16, x_316); +x_318 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_317, x_16, x_316); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_318); x_319 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -13346,7 +13346,7 @@ lean_object* x_338; lean_object* x_339; size_t x_340; lean_object* x_341; lean_o x_338 = l_List_appendTR___rarg(x_336, x_331); x_339 = lean_array_mk(x_338); x_340 = lean_array_size(x_339); -x_341 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_340, x_16, x_339); +x_341 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_340, x_16, x_339); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_341); x_342 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -13505,7 +13505,7 @@ x_383 = l_List_appendTR___rarg(x_382, x_380); x_384 = l_List_appendTR___rarg(x_336, x_383); x_385 = lean_array_mk(x_384); x_386 = lean_array_size(x_385); -x_387 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_386, x_16, x_385); +x_387 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_386, x_16, x_385); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_387); x_388 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -13598,7 +13598,7 @@ lean_object* x_412; lean_object* x_413; size_t x_414; lean_object* x_415; lean_o x_412 = l_List_appendTR___rarg(x_410, x_405); x_413 = lean_array_mk(x_412); x_414 = lean_array_size(x_413); -x_415 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_414, x_16, x_413); +x_415 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_414, x_16, x_413); x_416 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_416, 0, x_415); x_417 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -13757,7 +13757,7 @@ x_458 = l_List_appendTR___rarg(x_457, x_455); x_459 = l_List_appendTR___rarg(x_410, x_458); x_460 = lean_array_mk(x_459); x_461 = lean_array_size(x_460); -x_462 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_461, x_16, x_460); +x_462 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_461, x_16, x_460); x_463 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_463, 0, x_462); x_464 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -13905,7 +13905,7 @@ lean_object* x_510; lean_object* x_511; size_t x_512; lean_object* x_513; lean_o x_510 = l_List_appendTR___rarg(x_508, x_503); x_511 = lean_array_mk(x_510); x_512 = lean_array_size(x_511); -x_513 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_512, x_479, x_511); +x_513 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_512, x_479, x_511); if (lean_is_scalar(x_493)) { x_514 = lean_alloc_ctor(4, 1, 0); } else { @@ -14069,7 +14069,7 @@ x_556 = l_List_appendTR___rarg(x_555, x_553); x_557 = l_List_appendTR___rarg(x_508, x_556); x_558 = lean_array_mk(x_557); x_559 = lean_array_size(x_558); -x_560 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_559, x_479, x_558); +x_560 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_559, x_479, x_558); if (lean_is_scalar(x_493)) { x_561 = lean_alloc_ctor(4, 1, 0); } else { @@ -14238,7 +14238,7 @@ lean_object* x_611; lean_object* x_612; size_t x_613; lean_object* x_614; lean_o x_611 = l_List_appendTR___rarg(x_609, x_604); x_612 = lean_array_mk(x_611); x_613 = lean_array_size(x_612); -x_614 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_613, x_579, x_612); +x_614 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_613, x_579, x_612); if (lean_is_scalar(x_594)) { x_615 = lean_alloc_ctor(4, 1, 0); } else { @@ -14402,7 +14402,7 @@ x_657 = l_List_appendTR___rarg(x_656, x_654); x_658 = l_List_appendTR___rarg(x_609, x_657); x_659 = lean_array_mk(x_658); x_660 = lean_array_size(x_659); -x_661 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_660, x_579, x_659); +x_661 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_660, x_579, x_659); if (lean_is_scalar(x_594)) { x_662 = lean_alloc_ctor(4, 1, 0); } else { @@ -14545,7 +14545,7 @@ lean_object* x_44; lean_object* x_45; size_t x_46; lean_object* x_47; lean_objec x_44 = l_List_appendTR___rarg(x_42, x_39); x_45 = lean_array_mk(x_44); x_46 = lean_array_size(x_45); -x_47 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_46, x_16, x_45); +x_47 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_46, x_16, x_45); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_47); x_48 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -14651,7 +14651,7 @@ x_84 = l_List_appendTR___rarg(x_83, x_81); x_85 = l_List_appendTR___rarg(x_42, x_84); x_86 = lean_array_mk(x_85); x_87 = lean_array_size(x_86); -x_88 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_87, x_16, x_86); +x_88 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_87, x_16, x_86); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_88); x_89 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -14703,7 +14703,7 @@ x_106 = l_List_appendTR___rarg(x_105, x_103); x_107 = l_List_appendTR___rarg(x_42, x_106); x_108 = lean_array_mk(x_107); x_109 = lean_array_size(x_108); -x_110 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_109, x_16, x_108); +x_110 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_109, x_16, x_108); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_110); x_111 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -14771,7 +14771,7 @@ x_130 = l_List_appendTR___rarg(x_129, x_127); x_131 = l_List_appendTR___rarg(x_42, x_130); x_132 = lean_array_mk(x_131); x_133 = lean_array_size(x_132); -x_134 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_133, x_16, x_132); +x_134 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_133, x_16, x_132); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_134); x_135 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -14876,7 +14876,7 @@ x_165 = l_List_appendTR___rarg(x_164, x_162); x_166 = l_List_appendTR___rarg(x_42, x_165); x_167 = lean_array_mk(x_166); x_168 = lean_array_size(x_167); -x_169 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_168, x_16, x_167); +x_169 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_168, x_16, x_167); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_169); x_170 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -14996,7 +14996,7 @@ x_202 = l_List_appendTR___rarg(x_201, x_199); x_203 = l_List_appendTR___rarg(x_42, x_202); x_204 = lean_array_mk(x_203); x_205 = lean_array_size(x_204); -x_206 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_205, x_16, x_204); +x_206 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_205, x_16, x_204); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_206); x_207 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -15145,7 +15145,7 @@ x_247 = l_List_appendTR___rarg(x_246, x_244); x_248 = l_List_appendTR___rarg(x_42, x_247); x_249 = lean_array_mk(x_248); x_250 = lean_array_size(x_249); -x_251 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_250, x_16, x_249); +x_251 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_250, x_16, x_249); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_251); x_252 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -15201,7 +15201,7 @@ lean_object* x_269; lean_object* x_270; size_t x_271; lean_object* x_272; lean_o x_269 = l_List_appendTR___rarg(x_267, x_263); x_270 = lean_array_mk(x_269); x_271 = lean_array_size(x_270); -x_272 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_271, x_16, x_270); +x_272 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_271, x_16, x_270); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_272); x_273 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -15360,7 +15360,7 @@ x_314 = l_List_appendTR___rarg(x_313, x_311); x_315 = l_List_appendTR___rarg(x_267, x_314); x_316 = lean_array_mk(x_315); x_317 = lean_array_size(x_316); -x_318 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_317, x_16, x_316); +x_318 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_317, x_16, x_316); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_318); x_319 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -15431,7 +15431,7 @@ lean_object* x_338; lean_object* x_339; size_t x_340; lean_object* x_341; lean_o x_338 = l_List_appendTR___rarg(x_336, x_331); x_339 = lean_array_mk(x_338); x_340 = lean_array_size(x_339); -x_341 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_340, x_16, x_339); +x_341 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_340, x_16, x_339); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_341); x_342 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -15590,7 +15590,7 @@ x_383 = l_List_appendTR___rarg(x_382, x_380); x_384 = l_List_appendTR___rarg(x_336, x_383); x_385 = lean_array_mk(x_384); x_386 = lean_array_size(x_385); -x_387 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_386, x_16, x_385); +x_387 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_386, x_16, x_385); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_387); x_388 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -15683,7 +15683,7 @@ lean_object* x_412; lean_object* x_413; size_t x_414; lean_object* x_415; lean_o x_412 = l_List_appendTR___rarg(x_410, x_405); x_413 = lean_array_mk(x_412); x_414 = lean_array_size(x_413); -x_415 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_414, x_16, x_413); +x_415 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_414, x_16, x_413); x_416 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_416, 0, x_415); x_417 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -15842,7 +15842,7 @@ x_458 = l_List_appendTR___rarg(x_457, x_455); x_459 = l_List_appendTR___rarg(x_410, x_458); x_460 = lean_array_mk(x_459); x_461 = lean_array_size(x_460); -x_462 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_461, x_16, x_460); +x_462 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_461, x_16, x_460); x_463 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_463, 0, x_462); x_464 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -15990,7 +15990,7 @@ lean_object* x_510; lean_object* x_511; size_t x_512; lean_object* x_513; lean_o x_510 = l_List_appendTR___rarg(x_508, x_503); x_511 = lean_array_mk(x_510); x_512 = lean_array_size(x_511); -x_513 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_512, x_479, x_511); +x_513 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_512, x_479, x_511); if (lean_is_scalar(x_493)) { x_514 = lean_alloc_ctor(4, 1, 0); } else { @@ -16154,7 +16154,7 @@ x_556 = l_List_appendTR___rarg(x_555, x_553); x_557 = l_List_appendTR___rarg(x_508, x_556); x_558 = lean_array_mk(x_557); x_559 = lean_array_size(x_558); -x_560 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_559, x_479, x_558); +x_560 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_559, x_479, x_558); if (lean_is_scalar(x_493)) { x_561 = lean_alloc_ctor(4, 1, 0); } else { @@ -16323,7 +16323,7 @@ lean_object* x_611; lean_object* x_612; size_t x_613; lean_object* x_614; lean_o x_611 = l_List_appendTR___rarg(x_609, x_604); x_612 = lean_array_mk(x_611); x_613 = lean_array_size(x_612); -x_614 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_613, x_579, x_612); +x_614 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_613, x_579, x_612); if (lean_is_scalar(x_594)) { x_615 = lean_alloc_ctor(4, 1, 0); } else { @@ -16487,7 +16487,7 @@ x_657 = l_List_appendTR___rarg(x_656, x_654); x_658 = l_List_appendTR___rarg(x_609, x_657); x_659 = lean_array_mk(x_658); x_660 = lean_array_size(x_659); -x_661 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_660, x_579, x_659); +x_661 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_660, x_579, x_659); if (lean_is_scalar(x_594)) { x_662 = lean_alloc_ctor(4, 1, 0); } else { @@ -16630,7 +16630,7 @@ lean_object* x_44; lean_object* x_45; size_t x_46; lean_object* x_47; lean_objec x_44 = l_List_appendTR___rarg(x_42, x_39); x_45 = lean_array_mk(x_44); x_46 = lean_array_size(x_45); -x_47 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_46, x_16, x_45); +x_47 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_46, x_16, x_45); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_47); x_48 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -16736,7 +16736,7 @@ x_84 = l_List_appendTR___rarg(x_83, x_81); x_85 = l_List_appendTR___rarg(x_42, x_84); x_86 = lean_array_mk(x_85); x_87 = lean_array_size(x_86); -x_88 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_87, x_16, x_86); +x_88 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_87, x_16, x_86); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_88); x_89 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -16788,7 +16788,7 @@ x_106 = l_List_appendTR___rarg(x_105, x_103); x_107 = l_List_appendTR___rarg(x_42, x_106); x_108 = lean_array_mk(x_107); x_109 = lean_array_size(x_108); -x_110 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_109, x_16, x_108); +x_110 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_109, x_16, x_108); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_110); x_111 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -16856,7 +16856,7 @@ x_130 = l_List_appendTR___rarg(x_129, x_127); x_131 = l_List_appendTR___rarg(x_42, x_130); x_132 = lean_array_mk(x_131); x_133 = lean_array_size(x_132); -x_134 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_133, x_16, x_132); +x_134 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_133, x_16, x_132); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_134); x_135 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -16961,7 +16961,7 @@ x_165 = l_List_appendTR___rarg(x_164, x_162); x_166 = l_List_appendTR___rarg(x_42, x_165); x_167 = lean_array_mk(x_166); x_168 = lean_array_size(x_167); -x_169 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_168, x_16, x_167); +x_169 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_168, x_16, x_167); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_169); x_170 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -17081,7 +17081,7 @@ x_202 = l_List_appendTR___rarg(x_201, x_199); x_203 = l_List_appendTR___rarg(x_42, x_202); x_204 = lean_array_mk(x_203); x_205 = lean_array_size(x_204); -x_206 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_205, x_16, x_204); +x_206 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_205, x_16, x_204); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_206); x_207 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -17230,7 +17230,7 @@ x_247 = l_List_appendTR___rarg(x_246, x_244); x_248 = l_List_appendTR___rarg(x_42, x_247); x_249 = lean_array_mk(x_248); x_250 = lean_array_size(x_249); -x_251 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_250, x_16, x_249); +x_251 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_250, x_16, x_249); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_251); x_252 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -17286,7 +17286,7 @@ lean_object* x_269; lean_object* x_270; size_t x_271; lean_object* x_272; lean_o x_269 = l_List_appendTR___rarg(x_267, x_263); x_270 = lean_array_mk(x_269); x_271 = lean_array_size(x_270); -x_272 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_271, x_16, x_270); +x_272 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_271, x_16, x_270); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_272); x_273 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -17445,7 +17445,7 @@ x_314 = l_List_appendTR___rarg(x_313, x_311); x_315 = l_List_appendTR___rarg(x_267, x_314); x_316 = lean_array_mk(x_315); x_317 = lean_array_size(x_316); -x_318 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_317, x_16, x_316); +x_318 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_317, x_16, x_316); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_318); x_319 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -17516,7 +17516,7 @@ lean_object* x_338; lean_object* x_339; size_t x_340; lean_object* x_341; lean_o x_338 = l_List_appendTR___rarg(x_336, x_331); x_339 = lean_array_mk(x_338); x_340 = lean_array_size(x_339); -x_341 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_340, x_16, x_339); +x_341 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_340, x_16, x_339); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_341); x_342 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -17675,7 +17675,7 @@ x_383 = l_List_appendTR___rarg(x_382, x_380); x_384 = l_List_appendTR___rarg(x_336, x_383); x_385 = lean_array_mk(x_384); x_386 = lean_array_size(x_385); -x_387 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_386, x_16, x_385); +x_387 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_386, x_16, x_385); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_387); x_388 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -17768,7 +17768,7 @@ lean_object* x_412; lean_object* x_413; size_t x_414; lean_object* x_415; lean_o x_412 = l_List_appendTR___rarg(x_410, x_405); x_413 = lean_array_mk(x_412); x_414 = lean_array_size(x_413); -x_415 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_414, x_16, x_413); +x_415 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_414, x_16, x_413); x_416 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_416, 0, x_415); x_417 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -17927,7 +17927,7 @@ x_458 = l_List_appendTR___rarg(x_457, x_455); x_459 = l_List_appendTR___rarg(x_410, x_458); x_460 = lean_array_mk(x_459); x_461 = lean_array_size(x_460); -x_462 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_461, x_16, x_460); +x_462 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_461, x_16, x_460); x_463 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_463, 0, x_462); x_464 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -18075,7 +18075,7 @@ lean_object* x_510; lean_object* x_511; size_t x_512; lean_object* x_513; lean_o x_510 = l_List_appendTR___rarg(x_508, x_503); x_511 = lean_array_mk(x_510); x_512 = lean_array_size(x_511); -x_513 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_512, x_479, x_511); +x_513 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_512, x_479, x_511); if (lean_is_scalar(x_493)) { x_514 = lean_alloc_ctor(4, 1, 0); } else { @@ -18239,7 +18239,7 @@ x_556 = l_List_appendTR___rarg(x_555, x_553); x_557 = l_List_appendTR___rarg(x_508, x_556); x_558 = lean_array_mk(x_557); x_559 = lean_array_size(x_558); -x_560 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_559, x_479, x_558); +x_560 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_559, x_479, x_558); if (lean_is_scalar(x_493)) { x_561 = lean_alloc_ctor(4, 1, 0); } else { @@ -18408,7 +18408,7 @@ lean_object* x_611; lean_object* x_612; size_t x_613; lean_object* x_614; lean_o x_611 = l_List_appendTR___rarg(x_609, x_604); x_612 = lean_array_mk(x_611); x_613 = lean_array_size(x_612); -x_614 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_613, x_579, x_612); +x_614 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_613, x_579, x_612); if (lean_is_scalar(x_594)) { x_615 = lean_alloc_ctor(4, 1, 0); } else { @@ -18572,7 +18572,7 @@ x_657 = l_List_appendTR___rarg(x_656, x_654); x_658 = l_List_appendTR___rarg(x_609, x_657); x_659 = lean_array_mk(x_658); x_660 = lean_array_size(x_659); -x_661 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_660, x_579, x_659); +x_661 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_660, x_579, x_659); if (lean_is_scalar(x_594)) { x_662 = lean_alloc_ctor(4, 1, 0); } else { @@ -18654,7 +18654,7 @@ x_21 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_21, 0, x_8); lean_ctor_set(x_21, 1, x_20); x_22 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_RefInfo_toJsonParentDecl____x40_Lean_Data_Lsp_Internal___hyg_723____closed__4; -x_23 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_21, x_22); +x_23 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_21, x_22); x_24 = l_Lean_Json_mkObj(x_23); return x_24; } @@ -18682,7 +18682,7 @@ x_34 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_34, 0, x_8); lean_ctor_set(x_34, 1, x_33); x_35 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_RefInfo_toJsonParentDecl____x40_Lean_Data_Lsp_Internal___hyg_723____closed__4; -x_36 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_34, x_35); +x_36 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_34, x_35); x_37 = l_Lean_Json_mkObj(x_36); return x_37; } @@ -18713,7 +18713,7 @@ x_46 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_46, 0, x_8); lean_ctor_set(x_46, 1, x_45); x_47 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_RefInfo_toJsonParentDecl____x40_Lean_Data_Lsp_Internal___hyg_723____closed__4; -x_48 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_46, x_47); +x_48 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_46, x_47); x_49 = l_Lean_Json_mkObj(x_48); return x_49; } @@ -18741,7 +18741,7 @@ x_59 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_59, 0, x_8); lean_ctor_set(x_59, 1, x_58); x_60 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_RefInfo_toJsonParentDecl____x40_Lean_Data_Lsp_Internal___hyg_723____closed__4; -x_61 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_59, x_60); +x_61 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_59, x_60); x_62 = l_Lean_Json_mkObj(x_61); return x_62; } @@ -19144,7 +19144,7 @@ x_10 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_10, 0, x_9); lean_ctor_set(x_10, 1, x_8); x_11 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_RefInfo_toJsonParentDecl____x40_Lean_Data_Lsp_Internal___hyg_723____closed__4; -x_12 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_10, x_11); +x_12 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_10, x_11); x_13 = l_Lean_Json_mkObj(x_12); return x_13; } @@ -19271,7 +19271,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJson { lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanStaleDependencyParams____x40_Lean_Data_Lsp_Internal___hyg_2211____closed__1; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -19355,7 +19355,7 @@ x_7 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_7, 0, x_6); lean_ctor_set(x_7, 1, x_5); x_8 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_RefInfo_toJsonParentDecl____x40_Lean_Data_Lsp_Internal___hyg_723____closed__4; -x_9 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_7, x_8); +x_9 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_7, x_8); x_10 = l_Lean_Json_mkObj(x_9); return x_10; } diff --git a/stage0/stdlib/Lean/Data/Lsp/Ipc.c b/stage0/stdlib/Lean/Data/Lsp/Ipc.c index 6cffa9fe10ec..77aeb0e48fea 100644 --- a/stage0/stdlib/Lean/Data/Lsp/Ipc.c +++ b/stage0/stdlib/Lean/Data/Lsp/Ipc.c @@ -48,6 +48,7 @@ LEAN_EXPORT lean_object* l_Lean_Lsp_Ipc_collectDiagnostics_loop___boxed(lean_obj static lean_object* l_Lean_Lsp_Ipc_readResponseAs___closed__25; static lean_object* l_Lean_Lsp_Ipc_readResponseAs___closed__52; LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at_Lean_Lsp_Ipc_shutdown___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1455____spec__7(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_Ipc_readResponseAs___closed__20; static lean_object* l_Lean_Loop_forIn_loop___at_Lean_Lsp_Ipc_shutdown___spec__5___closed__7; LEAN_EXPORT lean_object* l_Lean_Lsp_Ipc_shutdown(lean_object*, lean_object*, lean_object*); @@ -68,7 +69,6 @@ static lean_object* l_Lean_Lsp_Ipc_readResponseAs___closed__2; LEAN_EXPORT lean_object* l_Lean_Json_toStructured_x3f___at_Lean_Lsp_Ipc_collectDiagnostics___spec__3(lean_object*); LEAN_EXPORT lean_object* l_IO_FS_Stream_writeLspRequest___at_Lean_Lsp_Ipc_collectDiagnostics___spec__2(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_toStructured_x3f___at_Lean_Lsp_Ipc_shutdown___spec__2(lean_object*); -lean_object* l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_Ipc_readResponseAs___closed__31; lean_object* lean_nat_to_int(lean_object*); lean_object* l_IO_FS_Stream_readLspMessage(lean_object*, lean_object*); @@ -2613,7 +2613,7 @@ x_155 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_155, 0, x_153); lean_ctor_set(x_155, 1, x_154); x_156 = l_Lean_Lsp_Ipc_readResponseAs___closed__8; -x_157 = l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(x_156, x_150); +x_157 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1455____spec__7(x_156, x_150); lean_dec(x_150); switch (lean_obj_tag(x_147)) { case 0: @@ -3130,7 +3130,7 @@ x_281 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_281, 0, x_279); lean_ctor_set(x_281, 1, x_280); x_282 = l_Lean_Lsp_Ipc_readResponseAs___closed__8; -x_283 = l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(x_282, x_276); +x_283 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1455____spec__7(x_282, x_276); lean_dec(x_276); switch (lean_obj_tag(x_273)) { case 0: diff --git a/stage0/stdlib/Lean/Data/Lsp/LanguageFeatures.c b/stage0/stdlib/Lean/Data/Lsp/LanguageFeatures.c index 44b2494a499b..5ab65ad2fb1c 100644 --- a/stage0/stdlib/Lean/Data/Lsp/LanguageFeatures.c +++ b/stage0/stdlib/Lean/Data/Lsp/LanguageFeatures.c @@ -81,7 +81,6 @@ static lean_object* l_Lean_Lsp_instToJsonCallHierarchyIncomingCall___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_5367____spec__3___rarg(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_CompletionItemKind_toCtorIdx(uint8_t); LEAN_EXPORT lean_object* l_Lean_Lsp_SemanticTokenModifier_noConfusion___rarg(uint8_t, uint8_t, lean_object*); -uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_627_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonCallHierarchyOutgoingCall; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5367____rarg___closed__5; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1403____closed__7; @@ -118,7 +117,6 @@ LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__Option_beqOption_ static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9157____closed__8; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9750____closed__3; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4438____closed__5; -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132_(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8921____boxed(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instToJsonDocumentHighlight___closed__1; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__28; @@ -136,6 +134,7 @@ LEAN_EXPORT lean_object* l_Lean_Lsp_FoldingRangeKind_noConfusion___rarg(uint8_t, LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5646____spec__2(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9238____lambda__6___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_349____closed__110; +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8085____lambda__15(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonCallHierarchyPrepareParams; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7438____closed__5; @@ -212,8 +211,6 @@ LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Ls static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_349____closed__136; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_349____closed__147; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9238_(lean_object*); -lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(lean_object*, lean_object*); -uint64_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6462_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7438____closed__14; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9856_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9157____closed__12; @@ -222,9 +219,9 @@ static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJ uint8_t lean_usize_dec_eq(size_t, size_t); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9157____closed__5; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9238____lambda__3___closed__1; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2414____spec__6___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_11217____closed__2; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5367____rarg(lean_object*, lean_object*); -lean_object* l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__34; static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__16; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5367____rarg___closed__22; @@ -244,6 +241,7 @@ LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Ls static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_349____closed__122; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__15; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8085____lambda__23(lean_object*, lean_object*, lean_object*); +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2128_(lean_object*); 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_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2953____closed__8; 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_44____spec__1(lean_object*, lean_object*); @@ -276,6 +274,7 @@ LEAN_EXPORT lean_object* l_Lean_Lsp_SymbolTag_noConfusion___rarg___boxed(lean_ob static lean_object* l_Lean_Lsp_instToJsonDocumentSymbol___closed__1; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1403____closed__15; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8085____lambda__9___boxed(lean_object*, lean_object*, 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_2668____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2783____closed__1; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10869____closed__9; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6572____closed__5; @@ -283,7 +282,6 @@ static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJ static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_349____closed__15; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__43; static lean_object* l_Lean_Lsp_instToJsonCallHierarchyPrepareParams___closed__1; -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_349____closed__78; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3973____closed__7; static lean_object* l_Lean_Lsp_instFromJsonReferenceContext___closed__1; @@ -305,6 +303,7 @@ LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonSymbolTag___boxed(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4040_(uint8_t); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9157____closed__16; static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__3; +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____spec__2(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_349____closed__13; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_349____closed__152; static lean_object* l_Lean_Lsp_instFromJsonHover___closed__1; @@ -326,7 +325,6 @@ static lean_object* l_Lean_Lsp_instToJsonCompletionItemTag___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8085____lambda__22(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_349____closed__8; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9750____closed__13; -lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__1(lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8921_(uint8_t, uint8_t); static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__11; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonHover; @@ -391,7 +389,6 @@ 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_349____closed__34; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6337_(lean_object*); -lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____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_349____closed__46; static lean_object* l_Lean_Lsp_instFromJsonFoldingRangeParams___closed__1; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__13; @@ -409,7 +406,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6572_(lean_object*); static lean_object* l_Lean_Lsp_instFromJsonSemanticTokensOptions___closed__1; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__8; -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instBEqSemanticTokenType; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7312____closed__6; LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____lambda__1(lean_object*); @@ -449,6 +445,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____lambda__1___boxed(lean_object*); size_t lean_usize_of_nat(lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6572____closed__3; +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4321____closed__2; static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__18; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4321____closed__3; @@ -505,6 +502,7 @@ static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__16; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6572____closed__15; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9157____closed__6; +lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155____spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_CompletionItemKind_noConfusion___rarg___closed__1; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_349____closed__84; static lean_object* l_Lean_Lsp_instFromJsonCompletionItem___closed__1; @@ -520,7 +518,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8085____lambda__15___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_9238____lambda__1___closed__2; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1403____closed__5; -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_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__2; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5367____rarg___closed__11; @@ -566,7 +563,6 @@ static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJ 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_4098____spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5646_(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8085____lambda__16(lean_object*, lean_object*, lean_object*); -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6159_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_349____closed__32; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__45; lean_object* l_Lean_Json_getObjValD(lean_object*, lean_object*); @@ -617,7 +613,6 @@ static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__23; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__39; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4098____closed__7; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7920____closed__31; -uint64_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_701_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10159____closed__2; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6018____spec__2(size_t, size_t, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_349____closed__86; @@ -657,7 +652,6 @@ static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJ static lean_object* l_Lean_Lsp_instBEqCallHierarchyItem___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8085____lambda__23___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_6572____closed__9; -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2313_(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*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10284____closed__1; @@ -683,8 +677,8 @@ static lean_object* l_Lean_Lsp_instFromJsonTypeDefinitionParams___closed__1; static lean_object* l_Lean_Lsp_instHashableSymbolTag___closed__1; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2074____closed__17; static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__8; -lean_object* l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____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_7920____closed__24; +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2414____spec__4___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7920____closed__44; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4098____closed__8; LEAN_EXPORT lean_object* l_Lean_Lsp_DocumentHighlightKind_noConfusion___rarg(uint8_t, uint8_t, lean_object*); @@ -756,7 +750,6 @@ static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJso LEAN_EXPORT lean_object* l_Lean_Lsp_SemanticTokenType_toCtorIdx(uint8_t); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____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_2074____closed__21; -lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__2(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10994____closed__2; 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*); @@ -834,6 +827,7 @@ static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJ static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1851____rarg___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_5367____spec__2___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2667____spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonSemanticTokensLegend; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__4; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_349____closed__101; @@ -858,11 +852,13 @@ static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJ static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_349____closed__70; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_349____closed__130; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8939____boxed(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2414____spec__5(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instBEqCompletionItem; LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2497____spec__5(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3175____closed__13; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10081_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3803____closed__8; +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__2(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9936____closed__15; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__16; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9936____closed__11; @@ -882,12 +878,14 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2074____closed__1; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7438____closed__10; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__18; +lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____spec__2(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instFromJsonInsertReplaceEdit___closed__1; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4321____closed__1; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__46; LEAN_EXPORT lean_object* l_Lean_Lsp_instInhabitedCallHierarchyItem; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4438____closed__4; static lean_object* l_Lean_Lsp_instInhabitedCompletionItem___closed__2; +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instToJsonReferenceContext___closed__1; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__31; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonFoldingRange; @@ -911,6 +909,7 @@ static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__9; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_349____closed__37; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7438____closed__3; LEAN_EXPORT lean_object* l_Lean_Lsp_DocumentHighlightKind_noConfusion(lean_object*); +uint64_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6277_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_11217____closed__3; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4321____closed__9; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9936____closed__3; @@ -932,12 +931,12 @@ static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprC 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_5367____spec__2___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instToJsonSemanticTokensLegend___closed__1; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_349____closed__125; +uint64_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_516_(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_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7438____closed__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_10800____spec__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instFromJsonRenameOptions___closed__1; LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2497____spec__2(lean_object*, lean_object*); -uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_decEqMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6317_(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__21; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7920____closed__6; static lean_object* l_Lean_Lsp_instFromJsonSemanticTokenType___closed__1; @@ -996,7 +995,6 @@ LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hash static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__26; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7920____closed__27; LEAN_EXPORT lean_object* l_Lean_Lsp_DocumentHighlightKind_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406_(lean_object*); static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__17; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10625_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonCompletionParams; @@ -1069,6 +1067,7 @@ static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJ static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7745____closed__3; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__18; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_349____closed__131; +uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_442_(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10159____closed__3; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3133____closed__2; LEAN_EXPORT lean_object* l_Lean_Lsp_SemanticTokenModifier_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); @@ -1114,7 +1113,6 @@ static lean_object* l_Lean_Lsp_instInhabitedCallHierarchyIncomingCall___closed__ LEAN_EXPORT lean_object* l_Lean_Lsp_SymbolKind_toCtorIdx___boxed(lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7920____closed__23; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8085____lambda__12___closed__1; -lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8085____lambda__21___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instHashableSymbolKind; static lean_object* l_Lean_Lsp_instToJsonInsertReplaceEdit___closed__1; @@ -1167,10 +1165,10 @@ static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJ LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonSemanticTokensOptions; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3739_(lean_object*); static lean_object* l_Lean_Lsp_instBEqCompletionItem___closed__1; +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9157____closed__14; LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2497____spec__3___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_FoldingRangeKind_toCtorIdx___boxed(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_1410____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_6018____closed__17; 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_5367____spec__2(lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_11217____closed__8; @@ -1184,12 +1182,13 @@ static lean_object* l_Lean_Lsp_SemanticTokenModifier_names___closed__8; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2783____closed__12; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2074____closed__45; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9238____lambda__8___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2414____spec__6(size_t, size_t, lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2497_(lean_object*, lean_object*); -lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9750____closed__7; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10465____closed__1; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7920____closed__3; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7188____spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5367____rarg___closed__8; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9157____closed__9; LEAN_EXPORT lean_object* l_Lean_Lsp_CompletionItemKind_noConfusion___rarg___lambda__1(lean_object*); @@ -1232,7 +1231,6 @@ static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJ static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7312____closed__3; LEAN_EXPORT uint8_t l_Lean_Lsp_instInhabitedCompletionItemKind; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7920____closed__14; -lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____spec__1(lean_object*, lean_object*); size_t lean_usize_add(size_t, size_t); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1851____rarg___closed__5; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlight____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4658_(lean_object*); @@ -1286,15 +1284,15 @@ LEAN_EXPORT lean_object* l_Lean_Lsp_SymbolTag_noConfusion(lean_object*, lean_obj static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_349____closed__148; static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__15; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7438____closed__9; -lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____spec__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2414____spec__5___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6337____closed__6; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7920____closed__16; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_349____closed__92; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_349____closed__104; static lean_object* l_Lean_Lsp_instFromJsonRenameParams___closed__1; +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____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_349____closed__107; static lean_object* l_Lean_Lsp_instToJsonRenameParams___closed__1; +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_5974_(lean_object*); 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_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8085____lambda__15___closed__1; @@ -1309,7 +1307,6 @@ static lean_object* l_Lean_Lsp_SemanticTokenModifier_names___closed__9; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6572____closed__10; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8085____lambda__13___closed__1; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_349____closed__60; -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10994____closed__12; LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlight____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4658____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_189____spec__2___boxed(lean_object*, lean_object*, lean_object*); @@ -1345,6 +1342,7 @@ static lean_object* l_Lean_Lsp_instFromJsonCallHierarchyItem___closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonDocumentHighlightKind(uint8_t); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2953____closed__12; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1403____closed__8; +uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_decEqMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6132_(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_9750____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9750____closed__6; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1403____closed__17; @@ -1362,6 +1360,7 @@ static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJ static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10284____closed__4; 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_2074____spec__3(lean_object*, lean_object*); uint8_t lean_usize_dec_lt(size_t, size_t); +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8085____lambda__11___closed__1; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_349____closed__132; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9750____closed__4; @@ -1394,6 +1393,7 @@ static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__44; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9936____closed__7; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_349____closed__66; LEAN_EXPORT lean_object* l_Lean_Lsp_instBEqInsertReplaceEdit; +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221_(lean_object*); static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___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_5646____spec__2___rarg(lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____spec__1___closed__1; @@ -1439,8 +1439,11 @@ 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_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10284____closed__5; static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__10; +lean_object* l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(lean_object*, lean_object*); +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonSemanticTokenType; LEAN_EXPORT lean_object* l_Lean_Lsp_CompletionItemKind_noConfusion___rarg___lambda__1___boxed(lean_object*); +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9157____closed__2; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_349____closed__102; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_349____closed__140; @@ -1450,9 +1453,7 @@ static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__7; static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____spec__1___closed__3; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonCallHierarchyIncomingCallsParams; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3463____closed__3; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2414____spec__5(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_4438____closed__2; -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794_(lean_object*); static lean_object* l_Lean_Lsp_SemanticTokenModifier_names___closed__5; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_349____closed__135; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6018____closed__15; @@ -2070,7 +2071,7 @@ 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_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__22; -x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____spec__1(x_1, x_24); +x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____spec__1(x_1, x_24); if (lean_obj_tag(x_25) == 0) { uint8_t x_26; @@ -2310,7 +2311,7 @@ x_16 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_16, 0, x_4); lean_ctor_set(x_16, 1, x_15); x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____closed__1; -x_18 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_16, x_17); +x_18 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_16, x_17); x_19 = l_Lean_Json_mkObj(x_18); return x_19; } @@ -5555,7 +5556,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ 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_1403____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -5593,7 +5594,7 @@ 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_1403____closed__10; lean_inc(x_1); -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__2(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__2(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -5631,7 +5632,7 @@ 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_1403____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_1132____spec__2(x_1, x_24); +x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__2(x_1, x_24); if (lean_obj_tag(x_25) == 0) { uint8_t x_26; @@ -5733,7 +5734,7 @@ x_8 = lean_box(0); 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_742_(x_3); +x_10 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_3); x_11 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1403____closed__10; x_12 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_12, 0, x_11); @@ -5741,7 +5742,7 @@ 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_8); -x_14 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_4); +x_14 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_4); x_15 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1403____closed__15; x_16 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_16, 0, x_15); @@ -5759,7 +5760,7 @@ x_20 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_20, 0, x_9); lean_ctor_set(x_20, 1, x_19); x_21 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____closed__1; -x_22 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_20, x_21); +x_22 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_20, x_21); x_23 = l_Lean_Json_mkObj(x_22); return x_23; } @@ -5800,7 +5801,7 @@ return x_10; else { uint8_t x_11; -x_11 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_627_(x_4, x_7); +x_11 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_442_(x_4, x_7); if (x_11 == 0) { uint8_t x_12; @@ -5810,7 +5811,7 @@ return x_12; else { uint8_t x_13; -x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_627_(x_5, x_8); +x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_442_(x_5, x_8); return x_13; } } @@ -5853,9 +5854,9 @@ x_4 = lean_ctor_get(x_1, 2); x_5 = 0; x_6 = lean_string_hash(x_2); x_7 = lean_uint64_mix_hash(x_5, x_6); -x_8 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_701_(x_3); +x_8 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_516_(x_3); x_9 = lean_uint64_mix_hash(x_7, x_8); -x_10 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_701_(x_4); +x_10 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_516_(x_4); x_11 = lean_uint64_mix_hash(x_9, x_10); return x_11; } @@ -6239,7 +6240,7 @@ return x_4; case 1: { lean_object* x_5; -x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211_(x_3); +x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026_(x_3); if (lean_obj_tag(x_5) == 0) { uint8_t x_6; @@ -6290,7 +6291,7 @@ return x_14; { lean_object* x_15; uint8_t x_16; lean_inc(x_3); -x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211_(x_3); +x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026_(x_3); x_16 = !lean_is_exclusive(x_3); if (x_16 == 0) { @@ -7509,7 +7510,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ 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_2074____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -7547,7 +7548,7 @@ 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_2074____closed__10; lean_inc(x_1); -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____spec__2(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____spec__2(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -7709,7 +7710,7 @@ 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_2074____closed__34; lean_inc(x_1); -x_58 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____spec__2(x_1, x_57); +x_58 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____spec__2(x_1, x_57); if (lean_obj_tag(x_58) == 0) { uint8_t x_59; @@ -7928,7 +7929,7 @@ else { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; x_4 = lean_ctor_get(x_2, 0); -x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6159_(x_4); +x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_5974_(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); @@ -8031,7 +8032,33 @@ return x_8; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2414____spec__5(size_t x_1, size_t x_2, lean_object* x_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_2414____spec__4(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_3; +lean_dec(x_1); +x_3 = lean_box(0); +return x_3; +} +else +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_4 = lean_ctor_get(x_2, 0); +lean_inc(x_4); +x_5 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_5, 0, x_1); +lean_ctor_set(x_5, 1, x_4); +x_6 = lean_box(0); +x_7 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_7, 0, x_5); +lean_ctor_set(x_7, 1, x_6); +return x_7; +} +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2414____spec__6(size_t x_1, size_t x_2, lean_object* x_3) { _start: { uint8_t x_4; @@ -8055,7 +8082,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2414____spec__4(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_2414____spec__5(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) @@ -8075,7 +8102,7 @@ lean_object* x_5; size_t x_6; size_t x_7; lean_object* x_8; lean_object* x_9; le x_5 = lean_ctor_get(x_2, 0); x_6 = lean_array_size(x_5); x_7 = 0; -x_8 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2414____spec__5(x_6, x_7, x_5); +x_8 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2414____spec__6(x_6, x_7, x_5); lean_ctor_set_tag(x_2, 4); lean_ctor_set(x_2, 0, x_8); x_9 = lean_alloc_ctor(0, 2, 0); @@ -8095,7 +8122,7 @@ lean_inc(x_12); lean_dec(x_2); x_13 = lean_array_size(x_12); x_14 = 0; -x_15 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2414____spec__5(x_13, x_14, x_12); +x_15 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2414____spec__6(x_13, x_14, x_12); x_16 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_16, 0, x_15); x_17 = lean_alloc_ctor(0, 2, 0); @@ -8142,7 +8169,7 @@ 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_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2074____closed__10; -x_16 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(x_15, x_3); +x_16 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____spec__2(x_15, x_3); x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2074____closed__16; x_18 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2414____spec__1(x_17, x_4); lean_dec(x_4); @@ -8151,12 +8178,12 @@ x_20 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp x_21 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2074____closed__28; x_22 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2414____spec__3(x_21, x_6); x_23 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2074____closed__34; -x_24 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(x_23, x_7); +x_24 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____spec__2(x_23, x_7); x_25 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2074____closed__40; -x_26 = l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(x_25, x_8); +x_26 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2414____spec__4(x_25, x_8); lean_dec(x_8); x_27 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2074____closed__46; -x_28 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2414____spec__4(x_27, x_9); +x_28 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2414____spec__5(x_27, x_9); x_29 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_29, 0, x_28); lean_ctor_set(x_29, 1, x_13); @@ -8182,7 +8209,7 @@ x_36 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_36, 0, x_14); lean_ctor_set(x_36, 1, x_35); x_37 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____closed__1; -x_38 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_36, x_37); +x_38 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_36, x_37); x_39 = l_Lean_Json_mkObj(x_38); return x_39; } @@ -8196,7 +8223,16 @@ lean_dec(x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2414____spec__5___boxed(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_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2414____spec__4___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_2414____spec__4(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2414____spec__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { size_t x_4; size_t x_5; lean_object* x_6; @@ -8204,7 +8240,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_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2414____spec__5(x_4, x_5, x_3); +x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2414____spec__6(x_4, x_5, x_3); return x_6; } } @@ -8326,7 +8362,7 @@ else lean_object* x_6; lean_object* x_7; uint8_t x_8; x_6 = lean_ctor_get(x_1, 0); x_7 = lean_ctor_get(x_2, 0); -x_8 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_decEqMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6317_(x_6, x_7); +x_8 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_decEqMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6132_(x_6, x_7); return x_8; } } @@ -8915,7 +8951,7 @@ lean_object* x_70; uint64_t x_71; uint64_t x_72; uint64_t x_73; x_70 = lean_ctor_get(x_4, 0); lean_inc(x_70); lean_dec(x_4); -x_71 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6462_(x_70); +x_71 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6277_(x_70); lean_dec(x_70); x_72 = 13; x_73 = lean_uint64_mix_hash(x_71, x_72); @@ -9410,7 +9446,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ 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_2783____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -9611,7 +9647,7 @@ x_17 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_17, 0, x_8); lean_ctor_set(x_17, 1, x_16); x_18 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____closed__1; -x_19 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_17, x_18); +x_19 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_17, x_18); x_20 = l_Lean_Json_mkObj(x_19); return x_20; } @@ -9788,7 +9824,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ 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_2953____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -9825,7 +9861,7 @@ 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_2953____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_794____spec__1(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -9908,7 +9944,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; 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_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_2313_(x_2); +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2128_(x_2); x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2953____closed__1; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); @@ -9920,7 +9956,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_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406_(x_8); +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221_(x_8); x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2953____closed__10; x_11 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_11, 0, x_10); @@ -9935,7 +9971,7 @@ x_14 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_14, 0, x_7); lean_ctor_set(x_14, 1, x_13); x_15 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____closed__1; -x_16 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_14, x_15); +x_16 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_14, x_15); x_17 = l_Lean_Json_mkObj(x_16); return x_17; } @@ -9978,7 +10014,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; 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_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_6159_(x_2); +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_5974_(x_2); lean_dec(x_2); x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3133____closed__1; x_5 = lean_alloc_ctor(0, 2, 0); @@ -9992,7 +10028,7 @@ 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_3133____closed__2; -x_10 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____spec__1(x_9, x_8); +x_10 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155____spec__1(x_9, x_8); x_11 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_11, 0, x_10); lean_ctor_set(x_11, 1, x_6); @@ -10000,7 +10036,7 @@ x_12 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_12, 0, x_7); lean_ctor_set(x_12, 1, x_11); x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____closed__1; -x_14 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_12, x_13); +x_14 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_12, x_13); x_15 = l_Lean_Json_mkObj(x_14); return x_15; } @@ -10026,7 +10062,7 @@ LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Ls { 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_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211_(x_3); +x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026_(x_3); return x_4; } } @@ -10203,7 +10239,7 @@ 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_3133____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_1410____spec__1(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -10375,7 +10411,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ 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_2953____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -10412,7 +10448,7 @@ 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_2953____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_794____spec__1(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -10495,7 +10531,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; 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_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_2313_(x_2); +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2128_(x_2); x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2953____closed__1; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); @@ -10507,7 +10543,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_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406_(x_8); +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221_(x_8); x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2953____closed__10; x_11 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_11, 0, x_10); @@ -10522,7 +10558,7 @@ x_14 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_14, 0, x_7); lean_ctor_set(x_14, 1, x_13); x_15 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____closed__1; -x_16 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_14, x_15); +x_16 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_14, x_15); x_17 = l_Lean_Json_mkObj(x_16); return x_17; } @@ -10629,7 +10665,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ 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_2953____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -10666,7 +10702,7 @@ 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_2953____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_794____spec__1(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -10749,7 +10785,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; 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_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_2313_(x_2); +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2128_(x_2); x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2953____closed__1; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); @@ -10761,7 +10797,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_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406_(x_8); +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221_(x_8); x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2953____closed__10; x_11 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_11, 0, x_10); @@ -10776,7 +10812,7 @@ x_14 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_14, 0, x_7); lean_ctor_set(x_14, 1, x_13); x_15 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____closed__1; -x_16 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_14, x_15); +x_16 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_14, x_15); x_17 = l_Lean_Json_mkObj(x_16); return x_17; } @@ -10883,7 +10919,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ 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_2953____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -10920,7 +10956,7 @@ 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_2953____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_794____spec__1(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -11003,7 +11039,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; 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_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_2313_(x_2); +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2128_(x_2); x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2953____closed__1; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); @@ -11015,7 +11051,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_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406_(x_8); +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221_(x_8); x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2953____closed__10; x_11 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_11, 0, x_10); @@ -11030,7 +11066,7 @@ x_14 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_14, 0, x_7); lean_ctor_set(x_14, 1, x_13); x_15 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____closed__1; -x_16 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_14, x_15); +x_16 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_14, x_15); x_17 = l_Lean_Json_mkObj(x_16); return x_17; } @@ -11137,7 +11173,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ 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_2953____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -11174,7 +11210,7 @@ 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_2953____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_794____spec__1(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -11257,7 +11293,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; 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_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_2313_(x_2); +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2128_(x_2); x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2953____closed__1; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); @@ -11269,7 +11305,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_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406_(x_8); +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221_(x_8); x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2953____closed__10; x_11 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_11, 0, x_10); @@ -11284,7 +11320,7 @@ x_14 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_14, 0, x_7); lean_ctor_set(x_14, 1, x_13); x_15 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____closed__1; -x_16 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_14, x_15); +x_16 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_14, x_15); x_17 = l_Lean_Json_mkObj(x_16); return x_17; } @@ -11399,7 +11435,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ { 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_3973____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_2853____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -11483,7 +11519,7 @@ x_7 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_7, 0, x_6); lean_ctor_set(x_7, 1, x_5); x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____closed__1; -x_9 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_7, x_8); +x_9 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_7, x_8); x_10 = l_Lean_Json_mkObj(x_9); return x_10; } @@ -11658,7 +11694,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ 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_2953____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -11696,7 +11732,7 @@ 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_2953____closed__10; lean_inc(x_1); -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____spec__1(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -11836,7 +11872,7 @@ x_2 = lean_ctor_get(x_1, 0); 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_2313_(x_3); +x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2128_(x_3); x_5 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2953____closed__1; x_6 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_6, 0, x_5); @@ -11848,7 +11884,7 @@ lean_ctor_set(x_8, 1, x_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_406_(x_9); +x_10 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221_(x_9); x_11 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2953____closed__10; x_12 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_12, 0, x_11); @@ -11879,7 +11915,7 @@ x_22 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_22, 0, x_8); lean_ctor_set(x_22, 1, x_21); x_23 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____closed__1; -x_24 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_22, x_23); +x_24 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_22, x_23); x_25 = l_Lean_Json_mkObj(x_24); return x_25; } @@ -11994,7 +12030,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ { 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_4321____closed__1; -x_3 = l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -12078,7 +12114,7 @@ x_7 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_7, 0, x_6); lean_ctor_set(x_7, 1, x_5); x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____closed__1; -x_9 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_7, x_8); +x_9 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_7, x_8); x_10 = l_Lean_Json_mkObj(x_9); return x_10; } @@ -12185,7 +12221,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ 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_2953____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -12222,7 +12258,7 @@ 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_2953____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_794____spec__1(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -12305,7 +12341,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; 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_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_2313_(x_2); +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2128_(x_2); x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2953____closed__1; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); @@ -12317,7 +12353,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_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406_(x_8); +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221_(x_8); x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2953____closed__10; x_11 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_11, 0, x_10); @@ -12332,7 +12368,7 @@ x_14 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_14, 0, x_7); lean_ctor_set(x_14, 1, x_13); x_15 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____closed__1; -x_16 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_14, x_15); +x_16 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_14, x_15); x_17 = l_Lean_Json_mkObj(x_16); return x_17; } @@ -12552,7 +12588,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; 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_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_742_(x_2); +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_2); x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3133____closed__2; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); @@ -12574,7 +12610,7 @@ x_12 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_12, 0, x_7); lean_ctor_set(x_12, 1, x_11); x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____closed__1; -x_14 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_12, x_13); +x_14 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_12, x_13); x_15 = l_Lean_Json_mkObj(x_14); return x_15; } @@ -12669,7 +12705,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ { 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_2953____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_5326____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -12739,7 +12775,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2313_(x_1); +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2128_(x_1); x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2953____closed__1; x_4 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_4, 0, x_3); @@ -12752,7 +12788,7 @@ x_7 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_7, 0, x_6); lean_ctor_set(x_7, 1, x_5); x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____closed__1; -x_9 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_7, x_8); +x_9 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_7, x_8); x_10 = l_Lean_Json_mkObj(x_9); return x_10; } @@ -21885,7 +21921,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ 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_5367____rarg___closed__1; lean_inc(x_2); -x_4 = l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(x_2, x_3); +x_4 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__1(x_2, x_3); if (lean_obj_tag(x_4) == 0) { uint8_t x_5; @@ -21924,7 +21960,7 @@ 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_2074____closed__10; lean_inc(x_2); -x_15 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____spec__2(x_2, x_14); +x_15 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____spec__2(x_2, x_14); if (lean_obj_tag(x_15) == 0) { uint8_t x_16; @@ -22005,7 +22041,7 @@ 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_3133____closed__2; lean_inc(x_2); -x_37 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__2(x_2, x_36); +x_37 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__2(x_2, x_36); if (lean_obj_tag(x_37) == 0) { uint8_t x_38; @@ -22047,7 +22083,7 @@ 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_5367____rarg___closed__20; lean_inc(x_2); -x_48 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__2(x_2, x_47); +x_48 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__2(x_2, x_47); if (lean_obj_tag(x_48) == 0) { uint8_t x_49; @@ -22351,8 +22387,8 @@ 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_2074____closed__10; -x_15 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(x_14, x_4); -x_16 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_6); +x_15 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____spec__2(x_14, x_4); +x_16 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_6); x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3133____closed__2; x_18 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_18, 0, x_17); @@ -22360,7 +22396,7 @@ lean_ctor_set(x_18, 1, x_16); x_19 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_19, 0, x_18); lean_ctor_set(x_19, 1, x_12); -x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_7); +x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_7); x_21 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5367____rarg___closed__20; x_22 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_22, 0, x_21); @@ -22583,7 +22619,7 @@ x_35 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_35, 0, x_13); lean_ctor_set(x_35, 1, x_34); x_36 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____closed__1; -x_37 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_35, x_36); +x_37 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_35, x_36); x_38 = l_Lean_Json_mkObj(x_37); return x_38; } @@ -22733,8 +22769,8 @@ 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_2074____closed__10; -x_14 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(x_13, x_3); -x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_5); +x_14 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____spec__2(x_13, x_3); +x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_5); x_16 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3133____closed__2; x_17 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_17, 0, x_16); @@ -22742,7 +22778,7 @@ lean_ctor_set(x_17, 1, x_15); x_18 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_18, 0, x_17); lean_ctor_set(x_18, 1, x_11); -x_19 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_6); +x_19 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_6); x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5367____rarg___closed__20; x_21 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_21, 0, x_20); @@ -22965,7 +23001,7 @@ x_34 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_34, 0, x_12); lean_ctor_set(x_34, 1, x_33); x_35 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____closed__1; -x_36 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_34, x_35); +x_36 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_34, x_35); x_37 = l_Lean_Json_mkObj(x_36); return x_37; } @@ -23795,7 +23831,7 @@ LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Ls { 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_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132_(x_3); +x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947_(x_3); return x_4; } } @@ -24032,7 +24068,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ 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_5367____rarg___closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -24189,7 +24225,7 @@ 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_6018____closed__18; -x_47 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____spec__2(x_1, x_46); +x_47 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____spec__2(x_1, x_46); if (lean_obj_tag(x_47) == 0) { uint8_t x_48; @@ -24372,7 +24408,7 @@ lean_ctor_set(x_17, 1, x_15); x_18 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_18, 0, x_17); lean_ctor_set(x_18, 1, x_10); -x_19 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080_(x_5); +x_19 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895_(x_5); x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6018____closed__13; x_21 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_21, 0, x_20); @@ -24381,7 +24417,7 @@ x_22 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_22, 0, x_21); lean_ctor_set(x_22, 1, x_10); x_23 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6018____closed__18; -x_24 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(x_23, x_6); +x_24 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____spec__2(x_23, x_6); x_25 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_25, 0, x_24); lean_ctor_set(x_25, 1, x_10); @@ -24592,7 +24628,7 @@ x_33 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_33, 0, x_11); lean_ctor_set(x_33, 1, x_32); x_34 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____closed__1; -x_35 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_33, x_34); +x_35 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_33, x_34); x_36 = l_Lean_Json_mkObj(x_35); return x_36; } @@ -24712,7 +24748,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ 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_2953____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -24749,7 +24785,7 @@ 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_2953____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_794____spec__1(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -24832,7 +24868,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; 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_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_2313_(x_2); +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2128_(x_2); x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2953____closed__1; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); @@ -24844,7 +24880,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_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406_(x_8); +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221_(x_8); x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2953____closed__10; x_11 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_11, 0, x_10); @@ -24859,7 +24895,7 @@ x_14 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_14, 0, x_7); lean_ctor_set(x_14, 1, x_13); x_15 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____closed__1; -x_16 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_14, x_15); +x_16 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_14, x_15); x_17 = l_Lean_Json_mkObj(x_16); return x_17; } @@ -25290,7 +25326,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ 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_5367____rarg___closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -25407,7 +25443,7 @@ 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_2074____closed__10; lean_inc(x_1); -x_36 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____spec__2(x_1, x_35); +x_36 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____spec__2(x_1, x_35); if (lean_obj_tag(x_36) == 0) { uint8_t x_37; @@ -25448,7 +25484,7 @@ 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_6572____closed__13; lean_inc(x_1); -x_47 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__1(x_1, x_46); +x_47 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__1(x_1, x_46); if (lean_obj_tag(x_47) == 0) { uint8_t x_48; @@ -25490,7 +25526,7 @@ 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_3133____closed__2; lean_inc(x_1); -x_58 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__2(x_1, x_57); +x_58 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__2(x_1, x_57); if (lean_obj_tag(x_58) == 0) { uint8_t x_59; @@ -25533,7 +25569,7 @@ 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_5367____rarg___closed__20; lean_inc(x_1); -x_69 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__2(x_1, x_68); +x_69 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__2(x_1, x_68); if (lean_obj_tag(x_69) == 0) { uint8_t x_70; @@ -25739,7 +25775,7 @@ lean_ctor_set(x_14, 1, x_13); x_15 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2074____closed__46; x_16 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6912____spec__1(x_15, x_4); x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2074____closed__10; -x_18 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(x_17, x_5); +x_18 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____spec__2(x_17, 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_6572____closed__13; @@ -25749,7 +25785,7 @@ 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_13); -x_23 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_7); +x_23 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_7); x_24 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3133____closed__2; x_25 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_25, 0, x_24); @@ -25757,7 +25793,7 @@ lean_ctor_set(x_25, 1, x_23); 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_742_(x_8); +x_27 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_8); x_28 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5367____rarg___closed__20; x_29 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_29, 0, x_28); @@ -25766,7 +25802,7 @@ 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_2074____closed__40; -x_32 = l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(x_31, x_9); +x_32 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2414____spec__4(x_31, x_9); lean_dec(x_9); x_33 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_33, 0, x_32); @@ -25987,7 +26023,7 @@ x_44 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_44, 0, x_14); lean_ctor_set(x_44, 1, x_43); x_45 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____closed__1; -x_46 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_44, x_45); +x_46 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_44, x_45); x_47 = l_Lean_Json_mkObj(x_46); return x_47; } @@ -26157,7 +26193,7 @@ return x_28; else { uint8_t x_29; -x_29 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_627_(x_8, x_16); +x_29 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_442_(x_8, x_16); if (x_29 == 0) { uint8_t x_30; @@ -26167,7 +26203,7 @@ return x_30; else { uint8_t x_31; -x_31 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_627_(x_9, x_17); +x_31 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_442_(x_9, x_17); if (x_31 == 0) { uint8_t x_32; @@ -26278,8 +26314,8 @@ 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_4846_(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_701_(x_7); -x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_701_(x_8); +x_16 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_516_(x_7); +x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_516_(x_8); if (lean_obj_tag(x_4) == 0) { uint64_t x_46; @@ -26678,7 +26714,7 @@ x_7 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_7, 0, x_6); lean_ctor_set(x_7, 1, x_5); x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____closed__1; -x_9 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_7, x_8); +x_9 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_7, x_8); x_10 = l_Lean_Json_mkObj(x_9); return x_10; } @@ -26717,7 +26753,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_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794_(x_6); +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609_(x_6); if (lean_obj_tag(x_9) == 0) { uint8_t x_10; @@ -27131,7 +27167,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_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_5); +x_8 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(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); @@ -27178,7 +27214,7 @@ x_17 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_17, 0, x_7); lean_ctor_set(x_17, 1, x_16); x_18 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____closed__1; -x_19 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_17, x_18); +x_19 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_17, x_18); x_20 = l_Lean_Json_mkObj(x_19); return x_20; } @@ -27388,7 +27424,7 @@ x_7 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_7, 0, x_6); lean_ctor_set(x_7, 1, x_5); x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____closed__1; -x_9 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_7, x_8); +x_9 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_7, x_8); x_10 = l_Lean_Json_mkObj(x_9); return x_10; } @@ -27675,7 +27711,7 @@ x_17 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_17, 0, x_7); lean_ctor_set(x_17, 1, x_16); x_18 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____closed__1; -x_19 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_17, x_18); +x_19 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_17, x_18); x_20 = l_Lean_Json_mkObj(x_19); return x_20; } @@ -32198,7 +32234,7 @@ x_19 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_19, 0, x_10); lean_ctor_set(x_19, 1, x_18); x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____closed__1; -x_21 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_19, x_20); +x_21 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_19, x_20); x_22 = l_Lean_Json_mkObj(x_21); return x_22; } @@ -32430,7 +32466,7 @@ 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_3133____closed__2; lean_inc(x_1); -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____spec__1(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -32468,7 +32504,7 @@ 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_9936____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_2853____spec__1(x_1, x_24); +x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____spec__1(x_1, x_24); if (lean_obj_tag(x_25) == 0) { uint8_t x_26; @@ -32612,7 +32648,7 @@ x_20 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_20, 0, x_7); lean_ctor_set(x_20, 1, x_19); x_21 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____closed__1; -x_22 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_20, x_21); +x_22 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_20, x_21); x_23 = l_Lean_Json_mkObj(x_22); return x_23; } @@ -32698,7 +32734,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ { 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_2953____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_5326____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -32768,7 +32804,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2313_(x_1); +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2128_(x_1); x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2953____closed__1; x_4 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_4, 0, x_3); @@ -32781,7 +32817,7 @@ x_7 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_7, 0, x_6); lean_ctor_set(x_7, 1, x_5); x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____closed__1; -x_9 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_7, x_8); +x_9 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_7, x_8); x_10 = l_Lean_Json_mkObj(x_9); return x_10; } @@ -32888,7 +32924,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ 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_2953____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -32925,7 +32961,7 @@ 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_3133____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_1132____spec__2(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__2(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -33008,7 +33044,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; 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_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_2313_(x_2); +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2128_(x_2); x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2953____closed__1; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); @@ -33020,7 +33056,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_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_8); +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_8); x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3133____closed__2; x_11 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_11, 0, x_10); @@ -33035,7 +33071,7 @@ x_14 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_14, 0, x_7); lean_ctor_set(x_14, 1, x_13); x_15 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____closed__1; -x_16 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_14, x_15); +x_16 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_14, x_15); x_17 = l_Lean_Json_mkObj(x_16); return x_17; } @@ -33338,7 +33374,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ 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_10465____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____spec__2(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____spec__2(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -33507,7 +33543,7 @@ lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; size_t x 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_10465____closed__1; -x_4 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(x_3, x_2); +x_4 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____spec__2(x_3, x_2); x_5 = lean_ctor_get(x_1, 1); lean_inc(x_5); lean_dec(x_1); @@ -33531,7 +33567,7 @@ x_15 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_15, 0, x_4); lean_ctor_set(x_15, 1, x_14); x_16 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____closed__1; -x_17 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_15, x_16); +x_17 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_15, x_16); x_18 = l_Lean_Json_mkObj(x_17); return x_18; } @@ -33629,7 +33665,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ { 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_2953____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_5326____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -33699,7 +33735,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2313_(x_1); +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2128_(x_1); x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2953____closed__1; x_4 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_4, 0, x_3); @@ -33712,7 +33748,7 @@ x_7 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_7, 0, x_6); lean_ctor_set(x_7, 1, x_5); x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____closed__1; -x_9 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_7, x_8); +x_9 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_7, x_8); x_10 = l_Lean_Json_mkObj(x_9); return x_10; } @@ -33985,7 +34021,7 @@ x_20 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_20, 0, x_8); lean_ctor_set(x_20, 1, x_19); x_21 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____closed__1; -x_22 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_20, x_21); +x_22 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_20, x_21); x_23 = l_Lean_Json_mkObj(x_22); return x_23; } @@ -34109,7 +34145,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ { 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_10869____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_2853____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -34193,7 +34229,7 @@ x_7 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_7, 0, x_6); lean_ctor_set(x_7, 1, x_5); x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____closed__1; -x_9 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_7, x_8); +x_9 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_7, x_8); x_10 = l_Lean_Json_mkObj(x_9); return x_10; } @@ -34359,7 +34395,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ 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_2953____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -34397,7 +34433,7 @@ 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_2953____closed__10; lean_inc(x_1); -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____spec__1(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -34435,7 +34471,7 @@ 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_10994____closed__9; -x_25 = l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(x_1, x_24); +x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__1(x_1, x_24); if (lean_obj_tag(x_25) == 0) { uint8_t x_26; @@ -34528,7 +34564,7 @@ x_2 = lean_ctor_get(x_1, 0); 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_2313_(x_3); +x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2128_(x_3); x_5 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2953____closed__1; x_6 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_6, 0, x_5); @@ -34540,7 +34576,7 @@ lean_ctor_set(x_8, 1, x_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_406_(x_9); +x_10 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221_(x_9); x_11 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2953____closed__10; x_12 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_12, 0, x_11); @@ -34570,7 +34606,7 @@ x_21 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_21, 0, x_8); lean_ctor_set(x_21, 1, x_20); x_22 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____closed__1; -x_23 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_21, x_22); +x_23 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_21, x_22); x_24 = l_Lean_Json_mkObj(x_23); return x_24; } @@ -34677,7 +34713,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ 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_2953____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -34714,7 +34750,7 @@ 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_2953____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_794____spec__1(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -34797,7 +34833,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; 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_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_2313_(x_2); +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2128_(x_2); x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2953____closed__1; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); @@ -34809,7 +34845,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_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406_(x_8); +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221_(x_8); x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2953____closed__10; x_11 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_11, 0, x_10); @@ -34824,7 +34860,7 @@ x_14 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_14, 0, x_7); lean_ctor_set(x_14, 1, x_13); x_15 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____closed__1; -x_16 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_14, x_15); +x_16 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_14, x_15); x_17 = l_Lean_Json_mkObj(x_16); return x_17; } diff --git a/stage0/stdlib/Lean/Data/Lsp/TextSync.c b/stage0/stdlib/Lean/Data/Lsp/TextSync.c index ee876cde5ce2..6e0902ce108b 100644 --- a/stage0/stdlib/Lean/Data/Lsp/TextSync.c +++ b/stage0/stdlib/Lean/Data/Lsp/TextSync.c @@ -24,6 +24,7 @@ static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidCl static lean_object* l_Lean_Lsp_instToJsonSaveOptions___closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_TextDocumentContentChangeEvent_hasToJson(lean_object*); lean_object* l_Lean_Json_mkObj(lean_object*); +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonDidOpenTextDocumentParams; static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonTextDocumentSyncOptions____x40_Lean_Data_Lsp_TextSync___hyg_1223____closed__5; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonTextDocumentChangeRegistrationOptions____x40_Lean_Data_Lsp_TextSync___hyg_253_(lean_object*); @@ -36,10 +37,8 @@ static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonDidOpen LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonSaveOptions____x40_Lean_Data_Lsp_TextSync___hyg_872____boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_TextDocumentSyncKind_toCtorIdx(uint8_t); static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidChangeTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_585____closed__9; -lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidChangeTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_585____closed__7; static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidChangeTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_585____spec__1___closed__1; -lean_object* l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonTextDocumentChangeRegistrationOptions____x40_Lean_Data_Lsp_TextSync___hyg_253____closed__7; static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonSaveOptions____x40_Lean_Data_Lsp_TextSync___hyg_910____closed__5; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonTextDocumentContentChangeEvent(lean_object*); @@ -47,27 +46,27 @@ static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonSaveO static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidChangeTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_585____closed__3; static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidOpenTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_164____closed__3; static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidChangeTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_585____closed__4; +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2128_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonSaveOptions; static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonSaveOptions____x40_Lean_Data_Lsp_TextSync___hyg_910____closed__7; static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonSaveOptions____x40_Lean_Data_Lsp_TextSync___hyg_910____closed__3; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidOpenTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_164_(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_2668____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidOpenTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_164____closed__10; static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonTextDocumentSyncOptions____x40_Lean_Data_Lsp_TextSync___hyg_1223____spec__1___closed__1; static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidOpenTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_164____closed__5; static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidSaveTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_754____closed__6; static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonTextDocumentChangeRegistrationOptions____x40_Lean_Data_Lsp_TextSync___hyg_253____closed__9; static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonTextDocumentSyncOptions____x40_Lean_Data_Lsp_TextSync___hyg_1223____closed__22; +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____spec__2(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidChangeTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_585____closed__1; static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidSaveTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_754____closed__7; -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2440_(lean_object*); -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonTextDocumentSyncOptions____x40_Lean_Data_Lsp_TextSync___hyg_1139____closed__11; static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonTextDocumentSyncOptions____x40_Lean_Data_Lsp_TextSync___hyg_1139____closed__10; static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonTextDocumentSyncOptions____x40_Lean_Data_Lsp_TextSync___hyg_1223____closed__18; LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidOpenTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_164____spec__1___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonTextDocumentSyncOptions____x40_Lean_Data_Lsp_TextSync___hyg_1223____closed__6; static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonSaveOptions____x40_Lean_Data_Lsp_TextSync___hyg_910____closed__6; -lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonTextDocumentSyncOptions____x40_Lean_Data_Lsp_TextSync___hyg_1223____spec__1___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonDidChangeTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_533____closed__1; static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidChangeTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_585____closed__6; @@ -81,6 +80,7 @@ static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidSa static lean_object* l_Lean_Lsp_instFromJsonTextDocumentSyncKind___closed__1; static lean_object* l_Lean_Lsp_instToJsonTextDocumentSyncKind___closed__1; static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonTextDocumentSyncOptions____x40_Lean_Data_Lsp_TextSync___hyg_1223____closed__11; +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonTextDocumentChangeRegistrationOptions____x40_Lean_Data_Lsp_TextSync___hyg_253____closed__13; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonDidOpenTextDocumentParams; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidChangeTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_585____spec__2___boxed(lean_object*, lean_object*, lean_object*); @@ -89,13 +89,14 @@ LEAN_EXPORT lean_object* l_Lean_Lsp_TextDocumentSyncKind_noConfusion(lean_object static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonSaveOptions____x40_Lean_Data_Lsp_TextSync___hyg_910____closed__4; static lean_object* l_Lean_Lsp_instFromJsonTextDocumentSyncKind___closed__2; static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidSaveTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_754____closed__1; +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidOpenTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_164____closed__2; +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2255_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonDidChangeTextDocumentParams; static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidSaveTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_754____closed__11; lean_object* l_Lean_Json_getObjValD(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_TextDocumentSyncKind_noConfusion___rarg(uint8_t, uint8_t, lean_object*); static lean_object* l_Lean_Lsp_instToJsonTextDocumentSyncKind___closed__4; -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4990_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidOpenTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_164____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonTextDocumentSyncKind(lean_object*); static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidOpenTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_164____closed__11; @@ -103,7 +104,6 @@ static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidOp LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidChangeTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_585_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_TextDocumentSyncKind_toCtorIdx___boxed(lean_object*); static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonTextDocumentSyncOptions____x40_Lean_Data_Lsp_TextSync___hyg_1223____closed__14; -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2313_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidSaveTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_754____closed__9; static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonTextDocumentSyncOptions____x40_Lean_Data_Lsp_TextSync___hyg_1139____closed__4; static lean_object* l_Lean_Lsp_instToJsonDidSaveTextDocumentParams___closed__1; @@ -111,13 +111,12 @@ static lean_object* l_Lean_Lsp_instFromJsonTextDocumentSyncKind___closed__3; static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidCloseTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_1027____closed__6; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonSaveOptions____x40_Lean_Data_Lsp_TextSync___hyg_872_(uint8_t); static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonTextDocumentSyncOptions____x40_Lean_Data_Lsp_TextSync___hyg_1223____closed__20; -lean_object* l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonDidCloseTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_989_(lean_object*); +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidChangeTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_585____spec__1___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonDidChangeTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_533_(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonDidSaveTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_712_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonTextDocumentSyncOptions____x40_Lean_Data_Lsp_TextSync___hyg_1223____closed__10; -lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__2(lean_object*, lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonTextDocumentSyncOptions____x40_Lean_Data_Lsp_TextSync___hyg_1223_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidChangeTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_585____closed__5; @@ -132,9 +131,12 @@ static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonTextD static lean_object* l_Lean_Lsp_instFromJsonTextDocumentSyncKind___closed__4; static lean_object* l_Lean_Lsp_TextDocumentSyncKind_noConfusion___rarg___closed__1; static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidOpenTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_164____closed__12; +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__2(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonTextDocumentSyncOptions____x40_Lean_Data_Lsp_TextSync___hyg_1223____closed__9; static lean_object* l_Lean_Lsp_instFromJsonDidSaveTextDocumentParams___closed__1; static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidCloseTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_1027____closed__1; +lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____spec__2(lean_object*, lean_object*); +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonTextDocumentChangeRegistrationOptions____x40_Lean_Data_Lsp_TextSync___hyg_253____closed__10; static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonTextDocumentSyncOptions____x40_Lean_Data_Lsp_TextSync___hyg_1223____closed__13; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonDidCloseTextDocumentParams; @@ -144,7 +146,6 @@ static lean_object* l_Lean_Lsp_instToJsonTextDocumentSyncKind___closed__2; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonTextDocumentSyncOptions; LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonTextDocumentChangeRegistrationOptions____x40_Lean_Data_Lsp_TextSync___hyg_253____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonDidChangeTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_533____spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonSaveOptions____x40_Lean_Data_Lsp_TextSync___hyg_910____closed__2; static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonTextDocumentChangeRegistrationOptions____x40_Lean_Data_Lsp_TextSync___hyg_253____closed__15; static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonTextDocumentChangeRegistrationOptions____x40_Lean_Data_Lsp_TextSync___hyg_253____closed__14; @@ -169,7 +170,6 @@ static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonTextD static lean_object* l_Lean_Lsp_instFromJsonDidCloseTextDocumentParams___closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonSaveOptions; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonTextDocumentChangeRegistrationOptions; -lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____spec__2(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonTextDocumentSyncOptions____x40_Lean_Data_Lsp_TextSync___hyg_1223____closed__2; static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidSaveTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_754____closed__8; static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonTextDocumentSyncOptions____x40_Lean_Data_Lsp_TextSync___hyg_1139____closed__6; @@ -182,12 +182,13 @@ static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonTextD static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonTextDocumentSyncOptions____x40_Lean_Data_Lsp_TextSync___hyg_1223____closed__25; static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidCloseTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_1027____closed__4; static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidCloseTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_1027____closed__2; +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonTextDocumentSyncOptions____x40_Lean_Data_Lsp_TextSync___hyg_1139____closed__3; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidSaveTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_754_(lean_object*); lean_object* lean_array_mk(lean_object*); static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidSaveTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_754____closed__10; static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonTextDocumentChangeRegistrationOptions____x40_Lean_Data_Lsp_TextSync___hyg_253____closed__8; -lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____spec__1(lean_object*, lean_object*); +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4805_(lean_object*); size_t lean_usize_add(size_t, size_t); static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonTextDocumentChangeRegistrationOptions____x40_Lean_Data_Lsp_TextSync___hyg_253____closed__4; static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidOpenTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_164____closed__4; @@ -204,7 +205,6 @@ LEAN_EXPORT lean_object* l_Lean_Lsp_TextDocumentSyncKind_noConfusion___rarg___bo static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidCloseTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_1027____closed__3; static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonTextDocumentChangeRegistrationOptions____x40_Lean_Data_Lsp_TextSync___hyg_253____closed__6; static lean_object* l_Lean_Lsp_instFromJsonTextDocumentContentChangeEvent___closed__1; -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(lean_object*); lean_object* lean_string_append(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instToJsonDidChangeTextDocumentParams___closed__1; static lean_object* l_Lean_Lsp_instToJsonTextDocumentSyncKind___closed__5; @@ -221,11 +221,11 @@ LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Ls lean_object* l_Lean_Json_pretty(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonDidOpenTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_126____closed__2; static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidSaveTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_754____closed__5; -lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidOpenTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_164____closed__8; LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonTextDocumentSyncOptions____x40_Lean_Data_Lsp_TextSync___hyg_1139____spec__1(lean_object*, lean_object*); lean_object* lean_array_uset(lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonDidSaveTextDocumentParams; +lean_object* l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidChangeTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_585____spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidOpenTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_164____lambda__1___boxed(lean_object*); static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonTextDocumentSyncOptions____x40_Lean_Data_Lsp_TextSync___hyg_1139____closed__5; @@ -543,7 +543,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonDi _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4990_(x_1); +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4805_(x_1); x_3 = l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonDidOpenTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_126____closed__1; x_4 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_4, 0, x_3); @@ -556,7 +556,7 @@ x_7 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_7, 0, x_6); lean_ctor_set(x_7, 1, x_5); x_8 = l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonDidOpenTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_126____closed__2; -x_9 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_7, x_8); +x_9 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_7, x_8); x_10 = l_Lean_Json_mkObj(x_9); return x_10; } @@ -582,7 +582,7 @@ LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Ls { 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_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070_(x_3); +x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885_(x_3); return x_4; } } @@ -1020,7 +1020,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJson lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonTextDocumentChangeRegistrationOptions____x40_Lean_Data_Lsp_TextSync___hyg_253____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -1169,7 +1169,7 @@ LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonTextDocumentContentChangeEvent(l lean_object* x_2; lean_object* x_3; x_2 = l_Lean_Lsp_instFromJsonTextDocumentContentChangeEvent___closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__2(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__2(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -1180,7 +1180,7 @@ lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); lean_dec(x_5); x_6 = l_Lean_Lsp_instFromJsonTextDocumentContentChangeEvent___closed__2; -x_7 = l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(x_1, x_6); +x_7 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__1(x_1, x_6); if (lean_obj_tag(x_7) == 0) { uint8_t x_8; @@ -1233,7 +1233,7 @@ else lean_object* x_15; lean_object* x_16; lean_dec(x_3); x_15 = l_Lean_Lsp_instFromJsonTextDocumentContentChangeEvent___closed__2; -x_16 = l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(x_1, x_15); +x_16 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__1(x_1, x_15); if (lean_obj_tag(x_16) == 0) { lean_object* x_17; lean_object* x_18; lean_object* x_19; @@ -1285,7 +1285,7 @@ x_24 = lean_ctor_get(x_3, 0); lean_inc(x_24); lean_dec(x_3); x_25 = l_Lean_Lsp_instFromJsonTextDocumentContentChangeEvent___closed__2; -x_26 = l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(x_1, x_25); +x_26 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__1(x_1, x_25); if (lean_obj_tag(x_26) == 0) { uint8_t x_27; @@ -1349,7 +1349,7 @@ if (x_2 == 0) lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; x_3 = lean_ctor_get(x_1, 0); x_4 = lean_ctor_get(x_1, 1); -x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_3); +x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_3); x_6 = l_Lean_Lsp_instFromJsonTextDocumentContentChangeEvent___closed__1; lean_ctor_set(x_1, 1, x_5); lean_ctor_set(x_1, 0, x_6); @@ -1377,7 +1377,7 @@ x_15 = lean_ctor_get(x_1, 1); lean_inc(x_15); lean_inc(x_14); lean_dec(x_1); -x_16 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_14); +x_16 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_14); x_17 = l_Lean_Lsp_instFromJsonTextDocumentContentChangeEvent___closed__1; x_18 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_18, 0, x_17); @@ -1467,7 +1467,7 @@ if (x_11 == 0) lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; x_12 = lean_ctor_get(x_6, 0); x_13 = lean_ctor_get(x_6, 1); -x_14 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_12); +x_14 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_12); x_15 = l_Lean_Lsp_instFromJsonTextDocumentContentChangeEvent___closed__1; lean_ctor_set(x_6, 1, x_14); lean_ctor_set(x_6, 0, x_15); @@ -1498,7 +1498,7 @@ x_25 = lean_ctor_get(x_6, 1); lean_inc(x_25); lean_inc(x_24); lean_dec(x_6); -x_26 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_24); +x_26 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_24); x_27 = l_Lean_Lsp_instFromJsonTextDocumentContentChangeEvent___closed__1; x_28 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_28, 0, x_27); @@ -1585,7 +1585,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonDi lean_object* x_2; 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_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_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2440_(x_2); +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2255_(x_2); x_4 = l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonDidOpenTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_126____closed__1; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); @@ -1616,7 +1616,7 @@ x_17 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_17, 0, x_7); lean_ctor_set(x_17, 1, x_16); x_18 = l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonDidOpenTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_126____closed__2; -x_19 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_17, x_18); +x_19 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_17, x_18); x_20 = l_Lean_Json_mkObj(x_19); return x_20; } @@ -1669,7 +1669,7 @@ x_7 = lean_unsigned_to_nat(0u); x_8 = lean_array_uset(x_3, x_2, x_7); x_9 = l_Lean_Lsp_instFromJsonTextDocumentContentChangeEvent___closed__1; lean_inc(x_6); -x_10 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__2(x_6, x_9); +x_10 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__2(x_6, x_9); if (lean_obj_tag(x_10) == 0) { uint8_t x_11; @@ -1680,7 +1680,7 @@ lean_object* x_12; lean_object* x_13; lean_object* x_14; x_12 = lean_ctor_get(x_10, 0); lean_dec(x_12); x_13 = l_Lean_Lsp_instFromJsonTextDocumentContentChangeEvent___closed__2; -x_14 = l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(x_6, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__1(x_6, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -1723,7 +1723,7 @@ else lean_object* x_23; lean_object* x_24; lean_dec(x_10); x_23 = l_Lean_Lsp_instFromJsonTextDocumentContentChangeEvent___closed__2; -x_24 = l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(x_6, x_23); +x_24 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__1(x_6, x_23); if (lean_obj_tag(x_24) == 0) { lean_object* x_25; lean_object* x_26; lean_object* x_27; @@ -1769,7 +1769,7 @@ x_34 = lean_ctor_get(x_10, 0); lean_inc(x_34); lean_dec(x_10); x_35 = l_Lean_Lsp_instFromJsonTextDocumentContentChangeEvent___closed__2; -x_36 = l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(x_6, x_35); +x_36 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__1(x_6, x_35); if (lean_obj_tag(x_36) == 0) { uint8_t x_37; @@ -2017,7 +2017,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJson lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonDidOpenTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_126____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -2158,7 +2158,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonDi lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; 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_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_2313_(x_2); +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2128_(x_2); x_4 = l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonDidOpenTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_126____closed__1; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); @@ -2171,7 +2171,7 @@ x_8 = lean_ctor_get(x_1, 1); lean_inc(x_8); lean_dec(x_1); x_9 = l_Lean_Lsp_instFromJsonTextDocumentContentChangeEvent___closed__2; -x_10 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(x_9, x_8); +x_10 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____spec__2(x_9, x_8); x_11 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_11, 0, x_10); lean_ctor_set(x_11, 1, x_6); @@ -2179,7 +2179,7 @@ x_12 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_12, 0, x_7); lean_ctor_set(x_12, 1, x_11); x_13 = l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonDidOpenTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_126____closed__2; -x_14 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_12, x_13); +x_14 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_12, x_13); x_15 = l_Lean_Json_mkObj(x_14); return x_15; } @@ -2315,7 +2315,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJson lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonDidOpenTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_126____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -2352,7 +2352,7 @@ x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); x_13 = l_Lean_Lsp_instFromJsonTextDocumentContentChangeEvent___closed__2; -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____spec__2(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____spec__2(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -2455,7 +2455,7 @@ x_7 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_7, 0, x_6); lean_ctor_set(x_7, 1, x_5); x_8 = l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonDidOpenTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_126____closed__2; -x_9 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_7, x_8); +x_9 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_7, x_8); x_10 = l_Lean_Json_mkObj(x_9); return x_10; } @@ -2572,7 +2572,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJson { lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonSaveOptions____x40_Lean_Data_Lsp_TextSync___hyg_872____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_2853____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -2642,7 +2642,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonDi _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2313_(x_1); +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2128_(x_1); x_3 = l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonDidOpenTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_126____closed__1; x_4 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_4, 0, x_3); @@ -2655,7 +2655,7 @@ x_7 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_7, 0, x_6); lean_ctor_set(x_7, 1, x_5); x_8 = l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonDidOpenTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_126____closed__2; -x_9 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_7, x_8); +x_9 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_7, x_8); x_10 = l_Lean_Json_mkObj(x_9); return x_10; } @@ -2741,7 +2741,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJson { lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonDidOpenTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_126____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_5326____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -3011,7 +3011,7 @@ x_27 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_27, 0, x_11); lean_ctor_set(x_27, 1, x_26); x_28 = l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonDidOpenTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_126____closed__2; -x_29 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_27, x_28); +x_29 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_27, x_28); x_30 = l_Lean_Json_mkObj(x_29); return x_30; } @@ -3026,7 +3026,7 @@ x_33 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_33, 0, x_11); lean_ctor_set(x_33, 1, x_32); x_34 = l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonDidOpenTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_126____closed__2; -x_35 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_33, x_34); +x_35 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_33, x_34); x_36 = l_Lean_Json_mkObj(x_35); return x_36; } @@ -3041,7 +3041,7 @@ x_39 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_39, 0, x_11); lean_ctor_set(x_39, 1, x_38); x_40 = l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonDidOpenTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_126____closed__2; -x_41 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_39, x_40); +x_41 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_39, x_40); x_42 = l_Lean_Json_mkObj(x_41); return x_42; } @@ -3503,7 +3503,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJson lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonTextDocumentSyncOptions____x40_Lean_Data_Lsp_TextSync___hyg_1139____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -3580,7 +3580,7 @@ lean_inc(x_23); lean_dec(x_14); x_24 = l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonTextDocumentSyncOptions____x40_Lean_Data_Lsp_TextSync___hyg_1139____closed__2; lean_inc(x_1); -x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____spec__1(x_1, x_24); +x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____spec__1(x_1, x_24); if (lean_obj_tag(x_25) == 0) { uint8_t x_26; @@ -3620,7 +3620,7 @@ lean_inc(x_34); lean_dec(x_25); x_35 = l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonTextDocumentSyncOptions____x40_Lean_Data_Lsp_TextSync___hyg_1139____closed__3; lean_inc(x_1); -x_36 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____spec__1(x_1, x_35); +x_36 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____spec__1(x_1, x_35); if (lean_obj_tag(x_36) == 0) { uint8_t x_37; diff --git a/stage0/stdlib/Lean/Data/Lsp/Utf16.c b/stage0/stdlib/Lean/Data/Lsp/Utf16.c index b1128d2ef78d..c43a16aa59cb 100644 --- a/stage0/stdlib/Lean/Data/Lsp/Utf16.c +++ b/stage0/stdlib/Lean/Data/Lsp/Utf16.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Lean.Data.Lsp.Utf16 -// Imports: Init.Data.String Init.Data.Array Lean.Data.Lsp.Basic Lean.Data.Position Lean.DeclarationRange +// Imports: Init.Data.String Lean.Data.Lsp.Basic Lean.Data.Position Lean.DeclarationRange #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -561,7 +561,6 @@ return x_2; } } lean_object* initialize_Init_Data_String(uint8_t builtin, lean_object*); -lean_object* initialize_Init_Data_Array(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Data_Lsp_Basic(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Data_Position(uint8_t builtin, lean_object*); lean_object* initialize_Lean_DeclarationRange(uint8_t builtin, lean_object*); @@ -573,9 +572,6 @@ _G_initialized = true; res = initialize_Init_Data_String(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -res = initialize_Init_Data_Array(builtin, lean_io_mk_world()); -if (lean_io_result_is_error(res)) return res; -lean_dec_ref(res); res = initialize_Lean_Data_Lsp_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/Data/Lsp/Workspace.c b/stage0/stdlib/Lean/Data/Lsp/Workspace.c index 0ec53b37a480..efcad0521383 100644 --- a/stage0/stdlib/Lean/Data/Lsp/Workspace.c +++ b/stage0/stdlib/Lean/Data/Lsp/Workspace.c @@ -40,7 +40,6 @@ LEAN_EXPORT lean_object* l_Lean_Lsp_FileSystemWatcher_delete; static lean_object* l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_fromJsonWorkspaceFolder____x40_Lean_Data_Lsp_Workspace___hyg_77____closed__15; static lean_object* l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_fromJsonWorkspaceFolder____x40_Lean_Data_Lsp_Workspace___hyg_77____closed__16; static lean_object* l_Lean_Lsp_instFromJsonFileChangeType___closed__1; -lean_object* l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_fromJsonFileSystemWatcher____x40_Lean_Data_Lsp_Workspace___hyg_205____closed__1; static lean_object* l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_fromJsonWorkspaceFolder____x40_Lean_Data_Lsp_Workspace___hyg_77____closed__6; static lean_object* l_Lean_Lsp_instFromJsonWorkspaceFolder___closed__1; @@ -52,7 +51,6 @@ static lean_object* l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_fromJsonWork static lean_object* l_Lean_Lsp_instToJsonDidChangeWatchedFilesRegistrationOptions___closed__1; static lean_object* l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_fromJsonDidChangeWatchedFilesParams____x40_Lean_Data_Lsp_Workspace___hyg_816____closed__2; LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_fromJsonWorkspaceFolder____x40_Lean_Data_Lsp_Workspace___hyg_77____lambda__1(lean_object*); -lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_toJsonDidChangeWatchedFilesParams____x40_Lean_Data_Lsp_Workspace___hyg_883_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_fromJsonFileSystemWatcher____x40_Lean_Data_Lsp_Workspace___hyg_205____closed__9; static lean_object* l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_fromJsonFileEvent____x40_Lean_Data_Lsp_Workspace___hyg_645____closed__8; @@ -70,7 +68,6 @@ LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonFileChangeType___boxed(lean_object LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonFileChangeType(lean_object*); static lean_object* l_Lean_Lsp_instToJsonDidChangeWatchedFilesParams___closed__1; static lean_object* l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_fromJsonFileSystemWatcher____x40_Lean_Data_Lsp_Workspace___hyg_205____closed__8; -lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_fromJsonDidChangeWatchedFilesParams____x40_Lean_Data_Lsp_Workspace___hyg_816____spec__2___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_toJsonWorkspaceFolder____x40_Lean_Data_Lsp_Workspace___hyg_25____closed__1; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_toJsonWorkspaceFolder____x40_Lean_Data_Lsp_Workspace___hyg_25_(lean_object*); @@ -95,7 +92,6 @@ static lean_object* l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_toJsonFileEv static lean_object* l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_fromJsonDidChangeWatchedFilesRegistrationOptions____x40_Lean_Data_Lsp_Workspace___hyg_390____closed__5; static lean_object* l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_fromJsonDidChangeWatchedFilesRegistrationOptions____x40_Lean_Data_Lsp_Workspace___hyg_390____closed__2; static lean_object* l_Lean_Lsp_instFromJsonFileChangeType___closed__5; -lean_object* l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_fromJsonFileEvent____x40_Lean_Data_Lsp_Workspace___hyg_645____closed__11; static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_fromJsonDidChangeWatchedFilesRegistrationOptions____x40_Lean_Data_Lsp_Workspace___hyg_390____spec__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_fromJsonDidChangeWatchedFilesRegistrationOptions____x40_Lean_Data_Lsp_Workspace___hyg_390____spec__1___boxed(lean_object*, lean_object*); @@ -117,6 +113,7 @@ static lean_object* l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_fromJsonFile LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_toJsonDidChangeWatchedFilesRegistrationOptions____x40_Lean_Data_Lsp_Workspace___hyg_457____spec__1(size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonFileEvent; static lean_object* l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_fromJsonDidChangeWatchedFilesRegistrationOptions____x40_Lean_Data_Lsp_Workspace___hyg_390____closed__1; +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instToJsonFileSystemWatcher___closed__1; static lean_object* l_Lean_Lsp_instToJsonFileChangeType___closed__5; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_toJsonFileSystemWatcher____x40_Lean_Data_Lsp_Workspace___hyg_311_(lean_object*); @@ -155,8 +152,10 @@ static lean_object* l_Lean_Lsp_instFromJsonDidChangeWatchedFilesParams___closed_ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_fromJsonDidChangeWatchedFilesParams____x40_Lean_Data_Lsp_Workspace___hyg_816_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_toJsonFileEvent____x40_Lean_Data_Lsp_Workspace___hyg_751____closed__9; static lean_object* l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_fromJsonWorkspaceFolder____x40_Lean_Data_Lsp_Workspace___hyg_77____closed__12; +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__1(lean_object*, lean_object*); lean_object* lean_array_mk(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_FileChangeType_noConfusion(lean_object*); +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_FileChangeType_noConfusion___rarg(uint8_t, uint8_t, lean_object*); size_t lean_usize_add(size_t, size_t); static lean_object* l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_fromJsonFileEvent____x40_Lean_Data_Lsp_Workspace___hyg_645____closed__1; @@ -176,6 +175,7 @@ static lean_object* l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_toJsonFileEv static lean_object* l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_fromJsonDidChangeWatchedFilesParams____x40_Lean_Data_Lsp_Workspace___hyg_816____closed__5; LEAN_EXPORT lean_object* l_Lean_Lsp_FileSystemWatcher_create; lean_object* lean_array_uset(lean_object*, size_t, lean_object*); +lean_object* l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonDidChangeWatchedFilesRegistrationOptions; static lean_object* l_Lean_Lsp_instToJsonFileChangeType___closed__4; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_toJsonDidChangeWatchedFilesRegistrationOptions____x40_Lean_Data_Lsp_Workspace___hyg_457____spec__1___boxed(lean_object*, lean_object*, lean_object*); @@ -242,7 +242,7 @@ x_14 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_14, 0, x_8); lean_ctor_set(x_14, 1, x_13); x_15 = l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_toJsonWorkspaceFolder____x40_Lean_Data_Lsp_Workspace___hyg_25____closed__3; -x_16 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_14, x_15); +x_16 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_14, x_15); x_17 = l_Lean_Json_mkObj(x_16); return x_17; } @@ -280,7 +280,7 @@ x_30 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_30, 0, x_24); lean_ctor_set(x_30, 1, x_29); x_31 = l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_toJsonWorkspaceFolder____x40_Lean_Data_Lsp_Workspace___hyg_25____closed__3; -x_32 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_30, x_31); +x_32 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_30, x_31); x_33 = l_Lean_Json_mkObj(x_32); return x_33; } @@ -478,7 +478,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_fromJso lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_toJsonWorkspaceFolder____x40_Lean_Data_Lsp_Workspace___hyg_25____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -515,7 +515,7 @@ x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); x_13 = l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_toJsonWorkspaceFolder____x40_Lean_Data_Lsp_Workspace___hyg_25____closed__2; -x_14 = l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -746,7 +746,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_fromJso lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_fromJsonFileSystemWatcher____x40_Lean_Data_Lsp_Workspace___hyg_205____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -783,7 +783,7 @@ x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); x_13 = l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_fromJsonFileSystemWatcher____x40_Lean_Data_Lsp_Workspace___hyg_205____closed__10; -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____spec__1(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -923,7 +923,7 @@ x_10 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_10, 0, x_8); lean_ctor_set(x_10, 1, x_9); x_11 = l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_toJsonWorkspaceFolder____x40_Lean_Data_Lsp_Workspace___hyg_25____closed__3; -x_12 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_10, x_11); +x_12 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_10, x_11); x_13 = l_Lean_Json_mkObj(x_12); return x_13; } @@ -952,7 +952,7 @@ x_21 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_21, 0, x_8); lean_ctor_set(x_21, 1, x_20); x_22 = l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_toJsonWorkspaceFolder____x40_Lean_Data_Lsp_Workspace___hyg_25____closed__3; -x_23 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_21, x_22); +x_23 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_21, x_22); x_24 = l_Lean_Json_mkObj(x_23); return x_24; } @@ -979,7 +979,7 @@ x_32 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_32, 0, x_8); lean_ctor_set(x_32, 1, x_31); x_33 = l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_toJsonWorkspaceFolder____x40_Lean_Data_Lsp_Workspace___hyg_25____closed__3; -x_34 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_32, x_33); +x_34 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_32, x_33); x_35 = l_Lean_Json_mkObj(x_34); return x_35; } @@ -1011,7 +1011,7 @@ x_44 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_44, 0, x_42); lean_ctor_set(x_44, 1, x_43); x_45 = l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_toJsonWorkspaceFolder____x40_Lean_Data_Lsp_Workspace___hyg_25____closed__3; -x_46 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_44, x_45); +x_46 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_44, x_45); x_47 = l_Lean_Json_mkObj(x_46); return x_47; } @@ -1049,7 +1049,7 @@ x_56 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_56, 0, x_42); lean_ctor_set(x_56, 1, x_55); x_57 = l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_toJsonWorkspaceFolder____x40_Lean_Data_Lsp_Workspace___hyg_25____closed__3; -x_58 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_56, x_57); +x_58 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_56, x_57); x_59 = l_Lean_Json_mkObj(x_58); return x_59; } @@ -1477,7 +1477,7 @@ x_10 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_10, 0, x_9); lean_ctor_set(x_10, 1, x_8); x_11 = l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_toJsonWorkspaceFolder____x40_Lean_Data_Lsp_Workspace___hyg_25____closed__3; -x_12 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_10, x_11); +x_12 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_10, x_11); x_13 = l_Lean_Json_mkObj(x_12); return x_13; } @@ -2158,7 +2158,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_fromJso lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_toJsonWorkspaceFolder____x40_Lean_Data_Lsp_Workspace___hyg_25____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -2419,7 +2419,7 @@ x_10 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_10, 0, x_8); lean_ctor_set(x_10, 1, x_9); x_11 = l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_toJsonWorkspaceFolder____x40_Lean_Data_Lsp_Workspace___hyg_25____closed__3; -x_12 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_10, x_11); +x_12 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_10, x_11); x_13 = l_Lean_Json_mkObj(x_12); return x_13; } @@ -2431,7 +2431,7 @@ x_15 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_15, 0, x_8); lean_ctor_set(x_15, 1, x_14); x_16 = l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_toJsonWorkspaceFolder____x40_Lean_Data_Lsp_Workspace___hyg_25____closed__3; -x_17 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_15, x_16); +x_17 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_15, x_16); x_18 = l_Lean_Json_mkObj(x_17); return x_18; } @@ -2443,7 +2443,7 @@ x_20 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_20, 0, x_8); lean_ctor_set(x_20, 1, x_19); x_21 = l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_toJsonWorkspaceFolder____x40_Lean_Data_Lsp_Workspace___hyg_25____closed__3; -x_22 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_20, x_21); +x_22 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_20, x_21); x_23 = l_Lean_Json_mkObj(x_22); return x_23; } @@ -2841,7 +2841,7 @@ x_10 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_10, 0, x_9); lean_ctor_set(x_10, 1, x_8); x_11 = l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_toJsonWorkspaceFolder____x40_Lean_Data_Lsp_Workspace___hyg_25____closed__3; -x_12 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_10, x_11); +x_12 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_10, x_11); x_13 = l_Lean_Json_mkObj(x_12); return x_13; } diff --git a/stage0/stdlib/Lean/Elab/BuiltinCommand.c b/stage0/stdlib/Lean/Elab/BuiltinCommand.c index f3709a66ad21..0ced9b85a6d1 100644 --- a/stage0/stdlib/Lean/Elab/BuiltinCommand.c +++ b/stage0/stdlib/Lean/Elab/BuiltinCommand.c @@ -87,7 +87,6 @@ static lean_object* l_Lean_Elab_Command_elabOpen___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Command_elabEnd_declRange__1___closed__2; static lean_object* l_Lean_resolveNamespaceCore___at_Lean_Elab_Command_elabExport___spec__4___closed__1; uint8_t l___private_Lean_Data_KVMap_0__Lean_beqDataValue____x40_Lean_Data_KVMap___hyg_70_(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__5___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Command_hasNoErrorMessages(lean_object*); LEAN_EXPORT lean_object* l_Lean_popScope___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_popScopes___spec__1___boxed(lean_object*, lean_object*, lean_object*); @@ -145,6 +144,7 @@ lean_object* l_Lean_FileMap_toPosition(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabImport___rarg___closed__1; static lean_object* l_Lean_Elab_Command_elabOmit___lambda__1___closed__1; static lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOptions___spec__1___lambda__2___closed__5; +static lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__3; lean_object* l_Lean_MessageData_joinSep(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_forArgsM___at_Lean_Elab_Command_elabUniverse___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_uint64_to_usize(uint64_t); @@ -158,6 +158,7 @@ LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_elabOpen___spe lean_object* l_Lean_Elab_addCompletionInfo___at_Lean_Elab_Term_addDotCompletionInfo___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Command_elabInclude___spec__1___closed__2; lean_object* l_Lean_Name_toString(lean_object*, uint8_t, lean_object*); +static lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__5; LEAN_EXPORT lean_object* l___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_containsId___boxed(lean_object*, lean_object*); lean_object* l_Lean_Syntax_getId(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabInitQuot___boxed(lean_object*); @@ -481,6 +482,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabVersion___rarg(lean_object*, le static lean_object* l___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_typelessBinder_x3f___closed__5; static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Command_elabOmit___spec__13___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Command_elabVersion__1___closed__3; +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_OpenDecl_elabOpenDecl___at_Lean_Elab_Command_elabOpen___spec__1___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Command_elabReduce__1___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Command_elabNonComputableSection__1___closed__2; @@ -562,7 +564,6 @@ LEAN_EXPORT lean_object* l_Lean_ensureNoOverload___at_Lean_Elab_Command_elabExpo lean_object* lean_array_to_list(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_addScopes___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabVersion___rarg___lambda__1___closed__33; -static lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabInclude___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabVersion___rarg___lambda__1___closed__27; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Command_elabUniverse_declRange__1(lean_object*); @@ -628,6 +629,7 @@ static lean_object* l___regBuiltin_Lean_Elab_Command_elabImport_declRange__1___c LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___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*); static lean_object* l_Lean_Elab_Command_elabExport___closed__4; static lean_object* l_Lean_Elab_Command_elabEnd___lambda__1___closed__2; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_addDocString___at_Lean_Elab_Command_elabAddDeclDoc___spec__5___lambda__1___closed__1; uint8_t lean_name_eq(lean_object*, lean_object*); lean_object* l_List_filterMapTR_go___at_Lean_resolveNamespace___spec__1(lean_object*, lean_object*); @@ -677,7 +679,6 @@ LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_elabExport___s static lean_object* l___regBuiltin_Lean_Elab_Command_elabChoice_declRange__1___closed__7; uint8_t l_Lean_Syntax_matchesNull(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_pushScope___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_addScope___spec__1___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___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___regBuiltin_Lean_Elab_Command_elabSection_declRange__1___closed__7; static lean_object* l___regBuiltin_Lean_Elab_Command_elabEoi_declRange__1___closed__5; extern uint8_t l_System_Platform_isEmscripten; @@ -730,8 +731,8 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabReduce(lean_object*, lean_objec static lean_object* l_Lean_Elab_Command_elabVersion___rarg___lambda__1___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Command_elabEnd_declRange__1___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Command_elabVariable_declRange__1___closed__1; +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_nestedExceptionToMessageData___at_Lean_Elab_Command_elabExport___spec__20___closed__2; -static lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabImport___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_expandInCmd___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabOpen(lean_object*, lean_object*, lean_object*, lean_object*); @@ -754,7 +755,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Comma lean_object* l_Lean_MessageData_ofExpr(lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Command_expandInCmd_declRange__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_pushScope___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_addScope___spec__1(lean_object*, lean_object*, lean_object*); -static lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__3; static lean_object* l_Lean_Elab_Command_elabCheckCore___lambda__4___closed__1; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Command_elabOpen___spec__25___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabExit___rarg(lean_object*, lean_object*, lean_object*); @@ -801,6 +801,7 @@ lean_object* l_Lean_Elab_Command_withFreshMacroScope___rarg(lean_object*, lean_o LEAN_EXPORT lean_object* l_Lean_Elab_OpenDecl_resolveNameUsingNamespaces___at_Lean_Elab_Command_elabExport___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_withNamespace___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand_go___at_Lean_Elab_Command_elabOmit___spec__7(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3(size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabInitQuot___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Command_elabSection_declRange__1(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__2(size_t, size_t, lean_object*); @@ -989,7 +990,6 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabUn static lean_object* l_Lean_Elab_nestedExceptionToMessageData___at_Lean_Elab_Command_elabExport___spec__20___closed__1; lean_object* l_List_toString___at_Lean_ensureNoOverload___spec__2(lean_object*); LEAN_EXPORT lean_object* l_Lean_activateScoped___at_Lean_Elab_Command_elabOpen___spec__18(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__4; lean_object* l_List_reverse___rarg(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabCheck_declRange__1___closed__2; static lean_object* l_Lean_Elab_Command_elabOpen___closed__1; @@ -1049,6 +1049,7 @@ lean_object* lean_io_error_to_string(lean_object*); static lean_object* l_Lean_Elab_Command_elabSynth___lambda__2___closed__2; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabUniverse___spec__3___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); +static lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Command_elabImport_declRange__1___closed__2; static lean_object* l___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___closed__1; static lean_object* l_Lean_Elab_Command_elabCheckCore___lambda__4___closed__2; @@ -1064,6 +1065,7 @@ lean_object* l_Lean_Syntax_unsetTrailing(lean_object*); static lean_object* l_Lean_resolveNamespaceCore___at_Lean_Elab_Command_elabExport___spec__4___closed__3; static lean_object* l_Lean_Elab_Command_elabVersion___rarg___lambda__1___closed__25; LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabWhere_describeOptions___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Command_elabOpen__1___closed__1; static uint8_t l_Lean_Elab_Command_elabVersion___rarg___closed__3; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Command_elabInclude___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1074,7 +1076,6 @@ static lean_object* l_Lean_Elab_OpenDecl_elabOpenDecl___at_Lean_Elab_Command_ela LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Command_elabOpen__1(lean_object*); lean_object* lean_string_append(lean_object*, lean_object*); uint8_t l_List_beq___at___private_Lean_Data_OpenDecl_0__Lean_beqOpenDecl____x40_Lean_Data_OpenDecl___hyg_41____spec__1(lean_object*, lean_object*); -static lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Command_elabReduce_declRange__1(lean_object*); static lean_object* l_Lean_Elab_Command_elabVersion___rarg___lambda__1___closed__37; static lean_object* l___regBuiltin_Lean_Elab_Command_elabEoi_declRange__1___closed__6; @@ -1104,6 +1105,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___regBuiltin_Lean_Elab_Command_elabEoi_declRange__1___closed__2; extern lean_object* l_Lean_Elab_unsupportedSyntaxExceptionId; +static lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__4; static lean_object* l_Lean_Elab_OpenDecl_elabOpenDecl___at_Lean_Elab_Command_elabOpen___spec__1___closed__4; LEAN_EXPORT lean_object* l___private_Lean_ResolveName_0__Lean_resolveGlobalConstCore___at_Lean_Elab_Command_elabExport___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_usize_dec_lt(size_t, size_t); @@ -1183,7 +1185,6 @@ static lean_object* l_Lean_resolveNamespaceCore___at_Lean_Elab_Command_elabExpor LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Command_elabOmit___spec__11___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Command_elabOmit___spec__12(lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__2___boxed(lean_object*, lean_object*, lean_object*); -lean_object* l_Array_mapMUnsafe_map___at_Std_DHashMap_Internal_Raw_u2080___aux__Std__Data__DHashMap__Internal__RawLemmas______macroRules__Std__DHashMap__Internal__Raw_u2080__tacticSimp__to__modelUsing____1___spec__3(size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabAddDeclDoc___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_throwErrorWithNestedErrors___at_Lean_Elab_Command_elabExport___spec__19___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabCheckCore___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -31828,6 +31829,31 @@ return x_3; } else { +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x_9; size_t x_10; lean_object* x_11; +x_5 = lean_array_uget(x_3, x_2); +x_6 = lean_unsigned_to_nat(0u); +x_7 = lean_array_uset(x_3, x_2, x_6); +x_8 = lean_mk_syntax_ident(x_5); +x_9 = 1; +x_10 = lean_usize_add(x_2, x_9); +x_11 = lean_array_uset(x_7, x_2, x_8); +x_2 = x_10; +x_3 = x_11; +goto _start; +} +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__2(size_t x_1, size_t x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; +x_4 = lean_usize_dec_lt(x_2, x_1); +if (x_4 == 0) +{ +return x_3; +} +else +{ lean_object* x_5; lean_object* x_6; lean_object* x_7; size_t x_8; size_t x_9; lean_object* x_10; x_5 = lean_array_uget(x_3, x_2); x_6 = lean_unsigned_to_nat(0u); @@ -31841,7 +31867,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__2(size_t x_1, size_t x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3(size_t x_1, size_t x_2, lean_object* x_3) { _start: { uint8_t x_4; @@ -31866,7 +31892,7 @@ goto _start; } } } -static lean_object* _init_l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__1() { +static lean_object* _init_l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__1() { _start: { lean_object* x_1; @@ -31874,7 +31900,7 @@ x_1 = lean_mk_string_unchecked("hiding", 6, 6); return x_1; } } -static lean_object* _init_l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__2() { +static lean_object* _init_l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__2() { _start: { lean_object* x_1; @@ -31882,7 +31908,7 @@ x_1 = lean_mk_string_unchecked("renaming", 8, 8); return x_1; } } -static lean_object* _init_l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__3() { +static lean_object* _init_l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__3() { _start: { lean_object* x_1; @@ -31890,19 +31916,19 @@ x_1 = lean_mk_string_unchecked("openRenamingItem", 16, 16); return x_1; } } -static lean_object* _init_l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__4() { +static lean_object* _init_l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___regBuiltin_Lean_Elab_Command_elabModuleDoc__1___closed__1; x_2 = l___regBuiltin_Lean_Elab_Command_elabModuleDoc__1___closed__2; x_3 = l___regBuiltin_Lean_Elab_Command_elabModuleDoc__1___closed__3; -x_4 = l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__3; +x_4 = l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__3; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__5() { +static lean_object* _init_l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__5() { _start: { lean_object* x_1; @@ -31910,7 +31936,7 @@ x_1 = lean_mk_string_unchecked("→", 3, 1); return x_1; } } -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { if (lean_obj_tag(x_6) == 0) @@ -32014,7 +32040,7 @@ lean_ctor_set(x_122, 1, x_126); lean_ctor_set(x_122, 0, x_117); x_127 = lean_array_size(x_16); x_128 = 0; -x_129 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__2(x_127, x_128, x_16); +x_129 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3(x_127, x_128, x_16); x_130 = l_Lean_Elab_OpenDecl_elabOpenDecl___at_Lean_Elab_Command_elabOpen___spec__1___closed__12; x_131 = l_Array_append___rarg(x_130, x_129); lean_dec(x_129); @@ -32051,7 +32077,7 @@ lean_ctor_set(x_142, 0, x_117); lean_ctor_set(x_142, 1, x_141); x_143 = lean_array_size(x_16); x_144 = 0; -x_145 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__2(x_143, x_144, x_16); +x_145 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3(x_143, x_144, x_16); x_146 = l_Lean_Elab_OpenDecl_elabOpenDecl___at_Lean_Elab_Command_elabOpen___spec__1___closed__12; x_147 = l_Array_append___rarg(x_146, x_145); lean_dec(x_145); @@ -32105,7 +32131,7 @@ lean_ctor_set(x_161, 0, x_117); lean_ctor_set(x_161, 1, x_160); x_162 = lean_array_size(x_16); x_163 = 0; -x_164 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__2(x_162, x_163, x_16); +x_164 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3(x_162, x_163, x_16); x_165 = l_Lean_Elab_OpenDecl_elabOpenDecl___at_Lean_Elab_Command_elabOpen___spec__1___closed__12; x_166 = l_Array_append___rarg(x_165, x_164); lean_dec(x_164); @@ -32204,7 +32230,7 @@ lean_ctor_set_tag(x_34, 2); lean_ctor_set(x_34, 1, x_38); lean_ctor_set(x_34, 0, x_29); x_39 = lean_mk_syntax_ident(x_18); -x_40 = l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__1; +x_40 = l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__1; lean_inc(x_29); lean_ctor_set_tag(x_30, 2); lean_ctor_set(x_30, 1, x_40); @@ -32212,9 +32238,9 @@ lean_ctor_set(x_30, 0, x_29); x_41 = lean_array_mk(x_19); x_42 = lean_array_size(x_41); x_43 = 0; -x_44 = l_Array_mapMUnsafe_map___at_Std_DHashMap_Internal_Raw_u2080___aux__Std__Data__DHashMap__Internal__RawLemmas______macroRules__Std__DHashMap__Internal__Raw_u2080__tacticSimp__to__modelUsing____1___spec__3(x_42, x_43, x_41); +x_44 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__1(x_42, x_43, x_41); x_45 = lean_array_size(x_44); -x_46 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__1(x_45, x_43, x_44); +x_46 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__2(x_45, x_43, x_44); x_47 = l_Lean_Elab_OpenDecl_elabOpenDecl___at_Lean_Elab_Command_elabOpen___spec__1___closed__12; x_48 = l_Array_append___rarg(x_47, x_46); lean_dec(x_46); @@ -32256,7 +32282,7 @@ x_61 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_61, 0, x_29); lean_ctor_set(x_61, 1, x_60); x_62 = lean_mk_syntax_ident(x_18); -x_63 = l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__1; +x_63 = l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__1; lean_inc(x_29); lean_ctor_set_tag(x_30, 2); lean_ctor_set(x_30, 1, x_63); @@ -32264,9 +32290,9 @@ lean_ctor_set(x_30, 0, x_29); x_64 = lean_array_mk(x_19); x_65 = lean_array_size(x_64); x_66 = 0; -x_67 = l_Array_mapMUnsafe_map___at_Std_DHashMap_Internal_Raw_u2080___aux__Std__Data__DHashMap__Internal__RawLemmas______macroRules__Std__DHashMap__Internal__Raw_u2080__tacticSimp__to__modelUsing____1___spec__3(x_65, x_66, x_64); +x_67 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__1(x_65, x_66, x_64); x_68 = lean_array_size(x_67); -x_69 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__1(x_68, x_66, x_67); +x_69 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__2(x_68, x_66, x_67); x_70 = l_Lean_Elab_OpenDecl_elabOpenDecl___at_Lean_Elab_Command_elabOpen___spec__1___closed__12; x_71 = l_Array_append___rarg(x_70, x_69); lean_dec(x_69); @@ -32325,7 +32351,7 @@ if (lean_is_scalar(x_85)) { lean_ctor_set(x_87, 0, x_29); lean_ctor_set(x_87, 1, x_86); x_88 = lean_mk_syntax_ident(x_18); -x_89 = l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__1; +x_89 = l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__1; lean_inc(x_29); x_90 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_90, 0, x_29); @@ -32333,9 +32359,9 @@ lean_ctor_set(x_90, 1, x_89); x_91 = lean_array_mk(x_19); x_92 = lean_array_size(x_91); x_93 = 0; -x_94 = l_Array_mapMUnsafe_map___at_Std_DHashMap_Internal_Raw_u2080___aux__Std__Data__DHashMap__Internal__RawLemmas______macroRules__Std__DHashMap__Internal__Raw_u2080__tacticSimp__to__modelUsing____1___spec__3(x_92, x_93, x_91); +x_94 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__1(x_92, x_93, x_91); x_95 = lean_array_size(x_94); -x_96 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__1(x_95, x_93, x_94); +x_96 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__2(x_95, x_93, x_94); x_97 = l_Lean_Elab_OpenDecl_elabOpenDecl___at_Lean_Elab_Command_elabOpen___spec__1___closed__12; x_98 = l_Array_append___rarg(x_97, x_96); lean_dec(x_96); @@ -32457,19 +32483,19 @@ lean_ctor_set_tag(x_208, 2); lean_ctor_set(x_208, 1, x_212); lean_ctor_set(x_208, 0, x_203); x_213 = lean_mk_syntax_ident(x_194); -x_214 = l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__2; +x_214 = l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__2; lean_inc(x_203); lean_ctor_set_tag(x_204, 2); lean_ctor_set(x_204, 1, x_214); lean_ctor_set(x_204, 0, x_203); x_215 = lean_mk_syntax_ident(x_197); -x_216 = l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__5; +x_216 = l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__5; lean_inc(x_203); lean_ctor_set_tag(x_198, 2); lean_ctor_set(x_198, 1, x_216); lean_ctor_set(x_198, 0, x_203); x_217 = lean_mk_syntax_ident(x_187); -x_218 = l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__4; +x_218 = l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__4; lean_inc(x_203); x_219 = l_Lean_Syntax_node3(x_203, x_218, x_215, x_198, x_217); x_220 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__5___closed__3; @@ -32507,19 +32533,19 @@ x_232 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_232, 0, x_203); lean_ctor_set(x_232, 1, x_231); x_233 = lean_mk_syntax_ident(x_194); -x_234 = l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__2; +x_234 = l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__2; lean_inc(x_203); lean_ctor_set_tag(x_204, 2); lean_ctor_set(x_204, 1, x_234); lean_ctor_set(x_204, 0, x_203); x_235 = lean_mk_syntax_ident(x_197); -x_236 = l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__5; +x_236 = l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__5; lean_inc(x_203); lean_ctor_set_tag(x_198, 2); lean_ctor_set(x_198, 1, x_236); lean_ctor_set(x_198, 0, x_203); x_237 = lean_mk_syntax_ident(x_187); -x_238 = l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__4; +x_238 = l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__4; lean_inc(x_203); x_239 = l_Lean_Syntax_node3(x_203, x_238, x_235, x_198, x_237); x_240 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__5___closed__3; @@ -32574,19 +32600,19 @@ if (lean_is_scalar(x_253)) { lean_ctor_set(x_255, 0, x_203); lean_ctor_set(x_255, 1, x_254); x_256 = lean_mk_syntax_ident(x_194); -x_257 = l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__2; +x_257 = l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__2; lean_inc(x_203); x_258 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_258, 0, x_203); lean_ctor_set(x_258, 1, x_257); x_259 = lean_mk_syntax_ident(x_197); -x_260 = l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__5; +x_260 = l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__5; lean_inc(x_203); lean_ctor_set_tag(x_198, 2); lean_ctor_set(x_198, 1, x_260); lean_ctor_set(x_198, 0, x_203); x_261 = lean_mk_syntax_ident(x_187); -x_262 = l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__4; +x_262 = l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__4; lean_inc(x_203); x_263 = l_Lean_Syntax_node3(x_203, x_262, x_259, x_198, x_261); x_264 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__5___closed__3; @@ -32657,7 +32683,7 @@ if (lean_is_scalar(x_283)) { lean_ctor_set(x_285, 0, x_277); lean_ctor_set(x_285, 1, x_284); x_286 = lean_mk_syntax_ident(x_194); -x_287 = l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__2; +x_287 = l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__2; lean_inc(x_277); if (lean_is_scalar(x_280)) { x_288 = lean_alloc_ctor(2, 2, 0); @@ -32668,13 +32694,13 @@ if (lean_is_scalar(x_280)) { lean_ctor_set(x_288, 0, x_277); lean_ctor_set(x_288, 1, x_287); x_289 = lean_mk_syntax_ident(x_197); -x_290 = l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__5; +x_290 = l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__5; lean_inc(x_277); x_291 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_291, 0, x_277); lean_ctor_set(x_291, 1, x_290); x_292 = lean_mk_syntax_ident(x_187); -x_293 = l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__4; +x_293 = l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__4; lean_inc(x_277); x_294 = l_Lean_Syntax_node3(x_277, x_293, x_289, x_291, x_292); x_295 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__5___closed__3; @@ -32739,7 +32765,7 @@ lean_ctor_set(x_317, 1, x_321); lean_ctor_set(x_317, 0, x_312); x_322 = lean_array_size(x_185); x_323 = 0; -x_324 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__2(x_322, x_323, x_185); +x_324 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3(x_322, x_323, x_185); x_325 = l_Lean_Elab_OpenDecl_elabOpenDecl___at_Lean_Elab_Command_elabOpen___spec__1___closed__12; x_326 = l_Array_append___rarg(x_325, x_324); lean_dec(x_324); @@ -32776,7 +32802,7 @@ lean_ctor_set(x_337, 0, x_312); lean_ctor_set(x_337, 1, x_336); x_338 = lean_array_size(x_185); x_339 = 0; -x_340 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__2(x_338, x_339, x_185); +x_340 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3(x_338, x_339, x_185); x_341 = l_Lean_Elab_OpenDecl_elabOpenDecl___at_Lean_Elab_Command_elabOpen___spec__1___closed__12; x_342 = l_Array_append___rarg(x_341, x_340); lean_dec(x_340); @@ -32830,7 +32856,7 @@ lean_ctor_set(x_356, 0, x_312); lean_ctor_set(x_356, 1, x_355); x_357 = lean_array_size(x_185); x_358 = 0; -x_359 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__2(x_357, x_358, x_185); +x_359 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3(x_357, x_358, x_185); x_360 = l_Lean_Elab_OpenDecl_elabOpenDecl___at_Lean_Elab_Command_elabOpen___spec__1___closed__12; x_361 = l_Array_append___rarg(x_360, x_359); lean_dec(x_359); @@ -32912,7 +32938,7 @@ x_27 = lean_box(0); x_28 = l_Lean_resolveNamespace___at_Lean_Elab_Command_elabExport___spec__1___closed__4; x_29 = l___private_Lean_Elab_Open_0__Lean_Elab_OpenDecl_resolveNameUsingNamespacesCore___at_Lean_Elab_Command_elabExport___spec__6___closed__1; lean_inc(x_1); -x_30 = l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3(x_1, x_26, x_28, x_27, x_1, x_1, x_29, lean_box(0), x_2, x_3, x_4); +x_30 = l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4(x_1, x_26, x_28, x_27, x_1, x_1, x_29, lean_box(0), x_2, x_3, x_4); lean_dec(x_1); x_31 = lean_ctor_get(x_30, 0); lean_inc(x_31); @@ -33049,7 +33075,7 @@ lean_ctor_set(x_47, 1, x_51); lean_ctor_set(x_47, 0, x_42); x_52 = lean_array_size(x_34); x_53 = 0; -x_54 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__2(x_52, x_53, x_34); +x_54 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3(x_52, x_53, x_34); x_55 = l_Lean_Elab_OpenDecl_elabOpenDecl___at_Lean_Elab_Command_elabOpen___spec__1___closed__12; x_56 = l_Array_append___rarg(x_55, x_54); lean_dec(x_54); @@ -33085,7 +33111,7 @@ lean_ctor_set(x_67, 0, x_42); lean_ctor_set(x_67, 1, x_66); x_68 = lean_array_size(x_34); x_69 = 0; -x_70 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__2(x_68, x_69, x_34); +x_70 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3(x_68, x_69, x_34); x_71 = l_Lean_Elab_OpenDecl_elabOpenDecl___at_Lean_Elab_Command_elabOpen___spec__1___closed__12; x_72 = l_Array_append___rarg(x_71, x_70); lean_dec(x_70); @@ -33138,7 +33164,7 @@ lean_ctor_set(x_86, 0, x_42); lean_ctor_set(x_86, 1, x_85); x_87 = lean_array_size(x_34); x_88 = 0; -x_89 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__2(x_87, x_88, x_34); +x_89 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3(x_87, x_88, x_34); x_90 = l_Lean_Elab_OpenDecl_elabOpenDecl___at_Lean_Elab_Command_elabOpen___spec__1___closed__12; x_91 = l_Array_append___rarg(x_90, x_89); lean_dec(x_89); @@ -33204,11 +33230,23 @@ x_6 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabWhere_describeOpenDecls_ return x_6; } } -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___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_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +size_t x_4; size_t x_5; lean_object* x_6; +x_4 = lean_unbox_usize(x_1); +lean_dec(x_1); +x_5 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_6 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3(x_4, x_5, x_3); +return x_6; +} +} +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { lean_object* x_12; -x_12 = l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_12 = l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_5); @@ -36201,16 +36239,16 @@ lean_mark_persistent(l___regBuiltin_Lean_Elab_Command_elabEoi_declRange__1___clo if (builtin) {res = l___regBuiltin_Lean_Elab_Command_elabEoi_declRange__1(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -}l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__1 = _init_l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__1(); -lean_mark_persistent(l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__1); -l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__2 = _init_l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__2(); -lean_mark_persistent(l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__2); -l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__3 = _init_l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__3(); -lean_mark_persistent(l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__3); -l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__4 = _init_l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__4(); -lean_mark_persistent(l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__4); -l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__5 = _init_l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__5(); -lean_mark_persistent(l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__5); +}l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__1 = _init_l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__1(); +lean_mark_persistent(l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__1); +l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__2 = _init_l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__2(); +lean_mark_persistent(l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__2); +l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__3 = _init_l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__3(); +lean_mark_persistent(l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__3); +l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__4 = _init_l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__4(); +lean_mark_persistent(l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__4); +l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__5 = _init_l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__5(); +lean_mark_persistent(l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__5); l_Lean_Elab_Command_elabWhere_describeOpenDecls___closed__1 = _init_l_Lean_Elab_Command_elabWhere_describeOpenDecls___closed__1(); lean_mark_persistent(l_Lean_Elab_Command_elabWhere_describeOpenDecls___closed__1); l_Lean_Elab_Command_elabWhere_describeOpenDecls___closed__2 = _init_l_Lean_Elab_Command_elabWhere_describeOpenDecls___closed__2(); diff --git a/stage0/stdlib/Lean/Elab/Tactic/BuiltinTactic.c b/stage0/stdlib/Lean/Elab/Tactic/BuiltinTactic.c index f8a296c97bb7..39fdeb78d0ba 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/BuiltinTactic.c +++ b/stage0/stdlib/Lean/Elab/Tactic/BuiltinTactic.c @@ -1289,7 +1289,7 @@ static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalRenameInaccessibles_decl lean_object* lean_array_mk(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwUnknownConstant___at_Lean_Elab_Tactic_evalOpen___spec__17___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_elabSetOption___at_Lean_Elab_Tactic_elabSetOption___spec__1___closed__3; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalIntro_introStep___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalIntro_introStep___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalSeq1_declRange__1___closed__7; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_renameInaccessibles___spec__15(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_evalIntro_introStep___lambda__6___closed__3; @@ -1308,7 +1308,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_renameInaccessibles___lambda__3(lean LEAN_EXPORT lean_object* l___private_Lean_Elab_Open_0__Lean_Elab_OpenDecl_resolveNameUsingNamespacesCore___at_Lean_Elab_Tactic_evalOpen___spec__34___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalContradiction___boxed(lean_object*); 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*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalIntro_introStep___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_evalIntro_introStep___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_evalSeq1__1___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalSkip_declRange__1___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalRunTac__1___closed__3; @@ -28943,213 +28943,436 @@ x_4 = l_Lean_addBuiltinDeclarationRanges(x_2, x_3, x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalIntro_introStep___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalIntro_introStep___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -lean_object* x_11; -x_11 = l_Lean_Elab_Tactic_getMainGoal(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_11) == 0) +lean_object* x_12; +x_12 = l_Lean_Elab_Tactic_getMainGoal(x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_12) == 0) { -lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); -x_13 = lean_ctor_get(x_11, 1); +if (lean_obj_tag(x_1) == 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); -lean_dec(x_11); +x_14 = lean_ctor_get(x_12, 1); +lean_inc(x_14); +lean_dec(x_12); +lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); -lean_inc(x_6); -x_14 = l_Lean_MVarId_intro(x_12, x_1, x_6, x_7, x_8, x_9, x_13); -if (lean_obj_tag(x_14) == 0) +x_15 = l_Lean_MVarId_intro(x_13, x_2, x_7, x_8, x_9, x_10, x_14); +if (lean_obj_tag(x_15) == 0) { -lean_object* x_15; lean_object* x_16; uint8_t x_17; -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; +x_16 = lean_ctor_get(x_15, 0); lean_inc(x_16); -lean_dec(x_14); -x_17 = !lean_is_exclusive(x_15); -if (x_17 == 0) +x_17 = lean_ctor_get(x_15, 1); +lean_inc(x_17); +lean_dec(x_15); +x_18 = !lean_is_exclusive(x_16); +if (x_18 == 0) { -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_18 = lean_ctor_get(x_15, 0); -x_19 = lean_ctor_get(x_15, 1); -x_20 = lean_box(0); -lean_ctor_set_tag(x_15, 1); -lean_ctor_set(x_15, 1, x_20); -lean_ctor_set(x_15, 0, x_19); -x_21 = l_Lean_Elab_Tactic_replaceMainGoal(x_15, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_16); +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_19 = lean_ctor_get(x_16, 0); +x_20 = lean_ctor_get(x_16, 1); +x_21 = lean_box(0); +lean_ctor_set_tag(x_16, 1); +lean_ctor_set(x_16, 1, x_21); +lean_ctor_set(x_16, 0, x_20); +x_22 = l_Lean_Elab_Tactic_replaceMainGoal(x_16, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_17); +lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_6); -if (lean_obj_tag(x_21) == 0) +if (lean_obj_tag(x_22) == 0) { -uint8_t x_22; -x_22 = !lean_is_exclusive(x_21); -if (x_22 == 0) +uint8_t x_23; +x_23 = !lean_is_exclusive(x_22); +if (x_23 == 0) { -lean_object* x_23; -x_23 = lean_ctor_get(x_21, 0); -lean_dec(x_23); -lean_ctor_set(x_21, 0, x_18); -return x_21; +lean_object* x_24; +x_24 = lean_ctor_get(x_22, 0); +lean_dec(x_24); +lean_ctor_set(x_22, 0, x_19); +return x_22; } else { -lean_object* x_24; lean_object* x_25; -x_24 = lean_ctor_get(x_21, 1); -lean_inc(x_24); -lean_dec(x_21); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_18); -lean_ctor_set(x_25, 1, x_24); -return x_25; +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_19); +lean_ctor_set(x_26, 1, x_25); +return x_26; } } else { -uint8_t x_26; -lean_dec(x_18); -x_26 = !lean_is_exclusive(x_21); -if (x_26 == 0) +uint8_t x_27; +lean_dec(x_19); +x_27 = !lean_is_exclusive(x_22); +if (x_27 == 0) { -return x_21; +return x_22; } else { -lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_27 = lean_ctor_get(x_21, 0); -x_28 = lean_ctor_get(x_21, 1); +lean_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_inc(x_27); -lean_dec(x_21); -x_29 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_29, 0, x_27); -lean_ctor_set(x_29, 1, x_28); -return x_29; +lean_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; } } } else { -lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_30 = lean_ctor_get(x_15, 0); -x_31 = lean_ctor_get(x_15, 1); +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_31 = lean_ctor_get(x_16, 0); +x_32 = lean_ctor_get(x_16, 1); +lean_inc(x_32); lean_inc(x_31); -lean_inc(x_30); -lean_dec(x_15); -x_32 = lean_box(0); -x_33 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_33, 0, x_31); -lean_ctor_set(x_33, 1, x_32); -x_34 = l_Lean_Elab_Tactic_replaceMainGoal(x_33, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_16); +lean_dec(x_16); +x_33 = lean_box(0); +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_35 = l_Lean_Elab_Tactic_replaceMainGoal(x_34, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_17); +lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_6); -if (lean_obj_tag(x_34) == 0) +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, 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; +lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_35, 1); +lean_inc(x_36); +if (lean_is_exclusive(x_35)) { + lean_ctor_release(x_35, 0); + lean_ctor_release(x_35, 1); + x_37 = x_35; } else { - lean_dec_ref(x_34); - x_36 = lean_box(0); + lean_dec_ref(x_35); + x_37 = lean_box(0); } -if (lean_is_scalar(x_36)) { - x_37 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_37)) { + x_38 = lean_alloc_ctor(0, 2, 0); } else { - x_37 = x_36; + x_38 = x_37; } -lean_ctor_set(x_37, 0, x_30); -lean_ctor_set(x_37, 1, x_35); -return x_37; +lean_ctor_set(x_38, 0, x_31); +lean_ctor_set(x_38, 1, x_36); +return x_38; } else { -lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; -lean_dec(x_30); -x_38 = lean_ctor_get(x_34, 0); -lean_inc(x_38); -x_39 = lean_ctor_get(x_34, 1); +lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; +lean_dec(x_31); +x_39 = lean_ctor_get(x_35, 0); lean_inc(x_39); -if (lean_is_exclusive(x_34)) { - lean_ctor_release(x_34, 0); - lean_ctor_release(x_34, 1); - x_40 = x_34; +x_40 = lean_ctor_get(x_35, 1); +lean_inc(x_40); +if (lean_is_exclusive(x_35)) { + lean_ctor_release(x_35, 0); + lean_ctor_release(x_35, 1); + x_41 = x_35; } else { - lean_dec_ref(x_34); - x_40 = lean_box(0); + lean_dec_ref(x_35); + x_41 = lean_box(0); } -if (lean_is_scalar(x_40)) { - x_41 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_41)) { + x_42 = lean_alloc_ctor(1, 2, 0); } else { - x_41 = x_40; + x_42 = x_41; } -lean_ctor_set(x_41, 0, x_38); -lean_ctor_set(x_41, 1, x_39); -return x_41; +lean_ctor_set(x_42, 0, x_39); +lean_ctor_set(x_42, 1, x_40); +return x_42; } } } else { -uint8_t x_42; +uint8_t x_43; +lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_6); -x_42 = !lean_is_exclusive(x_14); -if (x_42 == 0) +x_43 = !lean_is_exclusive(x_15); +if (x_43 == 0) { -return x_14; +return x_15; } else { -lean_object* x_43; lean_object* x_44; lean_object* x_45; -x_43 = lean_ctor_get(x_14, 0); -x_44 = lean_ctor_get(x_14, 1); +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_15, 0); +x_45 = lean_ctor_get(x_15, 1); +lean_inc(x_45); lean_inc(x_44); -lean_inc(x_43); -lean_dec(x_14); -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_15); +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; +lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; 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; uint8_t x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_47 = lean_ctor_get(x_12, 0); +lean_inc(x_47); +x_48 = lean_ctor_get(x_12, 1); +lean_inc(x_48); +lean_dec(x_12); +x_49 = lean_ctor_get(x_1, 0); +x_50 = lean_ctor_get(x_9, 0); +lean_inc(x_50); +x_51 = lean_ctor_get(x_9, 1); +lean_inc(x_51); +x_52 = lean_ctor_get(x_9, 2); +lean_inc(x_52); +x_53 = lean_ctor_get(x_9, 3); +lean_inc(x_53); +x_54 = lean_ctor_get(x_9, 4); +lean_inc(x_54); +x_55 = lean_ctor_get(x_9, 5); +lean_inc(x_55); +x_56 = lean_ctor_get(x_9, 6); +lean_inc(x_56); +x_57 = lean_ctor_get(x_9, 7); +lean_inc(x_57); +x_58 = lean_ctor_get(x_9, 8); +lean_inc(x_58); +x_59 = lean_ctor_get(x_9, 9); +lean_inc(x_59); +x_60 = lean_ctor_get(x_9, 10); +lean_inc(x_60); +x_61 = lean_ctor_get_uint8(x_9, sizeof(void*)*12); +x_62 = lean_ctor_get(x_9, 11); +lean_inc(x_62); +x_63 = lean_ctor_get_uint8(x_9, sizeof(void*)*12 + 1); +x_64 = l_Lean_replaceRef(x_49, x_55); +lean_dec(x_55); +x_65 = lean_alloc_ctor(0, 12, 2); +lean_ctor_set(x_65, 0, x_50); +lean_ctor_set(x_65, 1, x_51); +lean_ctor_set(x_65, 2, x_52); +lean_ctor_set(x_65, 3, x_53); +lean_ctor_set(x_65, 4, x_54); +lean_ctor_set(x_65, 5, x_64); +lean_ctor_set(x_65, 6, x_56); +lean_ctor_set(x_65, 7, x_57); +lean_ctor_set(x_65, 8, x_58); +lean_ctor_set(x_65, 9, x_59); +lean_ctor_set(x_65, 10, x_60); +lean_ctor_set(x_65, 11, x_62); +lean_ctor_set_uint8(x_65, sizeof(void*)*12, x_61); +lean_ctor_set_uint8(x_65, sizeof(void*)*12 + 1, x_63); +lean_inc(x_10); +lean_inc(x_8); +lean_inc(x_7); +x_66 = l_Lean_MVarId_intro(x_47, x_2, x_7, x_8, x_65, x_10, x_48); +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_is_exclusive(x_67); +if (x_69 == 0) +{ +lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; +x_70 = lean_ctor_get(x_67, 0); +x_71 = lean_ctor_get(x_67, 1); +x_72 = lean_box(0); +lean_ctor_set_tag(x_67, 1); +lean_ctor_set(x_67, 1, x_72); +lean_ctor_set(x_67, 0, x_71); +x_73 = l_Lean_Elab_Tactic_replaceMainGoal(x_67, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_68); +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_46 = !lean_is_exclusive(x_11); -if (x_46 == 0) +if (lean_obj_tag(x_73) == 0) { -return x_11; +uint8_t x_74; +x_74 = !lean_is_exclusive(x_73); +if (x_74 == 0) +{ +lean_object* x_75; +x_75 = lean_ctor_get(x_73, 0); +lean_dec(x_75); +lean_ctor_set(x_73, 0, x_70); +return x_73; } else { -lean_object* x_47; lean_object* x_48; lean_object* x_49; -x_47 = lean_ctor_get(x_11, 0); -x_48 = lean_ctor_get(x_11, 1); -lean_inc(x_48); -lean_inc(x_47); -lean_dec(x_11); -x_49 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_49, 0, x_47); -lean_ctor_set(x_49, 1, x_48); -return x_49; +lean_object* x_76; lean_object* x_77; +x_76 = lean_ctor_get(x_73, 1); +lean_inc(x_76); +lean_dec(x_73); +x_77 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_77, 0, x_70); +lean_ctor_set(x_77, 1, x_76); +return x_77; +} +} +else +{ +uint8_t x_78; +lean_dec(x_70); +x_78 = !lean_is_exclusive(x_73); +if (x_78 == 0) +{ +return x_73; +} +else +{ +lean_object* x_79; lean_object* x_80; lean_object* x_81; +x_79 = lean_ctor_get(x_73, 0); +x_80 = lean_ctor_get(x_73, 1); +lean_inc(x_80); +lean_inc(x_79); +lean_dec(x_73); +x_81 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_81, 0, x_79); +lean_ctor_set(x_81, 1, x_80); +return x_81; +} +} +} +else +{ +lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; +x_82 = lean_ctor_get(x_67, 0); +x_83 = lean_ctor_get(x_67, 1); +lean_inc(x_83); +lean_inc(x_82); +lean_dec(x_67); +x_84 = lean_box(0); +x_85 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_85, 0, x_83); +lean_ctor_set(x_85, 1, x_84); +x_86 = l_Lean_Elab_Tactic_replaceMainGoal(x_85, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_68); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +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(0, 2, 0); +} else { + x_89 = x_88; +} +lean_ctor_set(x_89, 0, x_82); +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_82); +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; +} +} +} +else +{ +uint8_t x_94; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +x_94 = !lean_is_exclusive(x_66); +if (x_94 == 0) +{ +return x_66; +} +else +{ +lean_object* x_95; lean_object* x_96; lean_object* x_97; +x_95 = lean_ctor_get(x_66, 0); +x_96 = lean_ctor_get(x_66, 1); +lean_inc(x_96); +lean_inc(x_95); +lean_dec(x_66); +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; +} +} +} +} +else +{ +uint8_t x_98; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_2); +x_98 = !lean_is_exclusive(x_12); +if (x_98 == 0) +{ +return x_12; +} +else +{ +lean_object* x_99; lean_object* x_100; lean_object* x_101; +x_99 = lean_ctor_get(x_12, 0); +x_100 = lean_ctor_get(x_12, 1); +lean_inc(x_100); +lean_inc(x_99); +lean_dec(x_12); +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; } } } @@ -29651,8 +29874,10 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalIntro_introStep(lean_object* x_1 _start: { lean_object* x_13; lean_object* x_14; -x_13 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_evalIntro_introStep___lambda__1___boxed), 10, 1); -lean_closure_set(x_13, 0, x_2); +lean_inc(x_1); +x_13 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_evalIntro_introStep___lambda__1___boxed), 11, 2); +lean_closure_set(x_13, 0, x_1); +lean_closure_set(x_13, 1, x_2); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); @@ -29780,16 +30005,17 @@ return x_34; } } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalIntro_introStep___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalIntro_introStep___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -lean_object* x_11; -x_11 = l_Lean_Elab_Tactic_evalIntro_introStep___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_object* x_12; +x_12 = l_Lean_Elab_Tactic_evalIntro_introStep___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_11; +lean_dec(x_1); +return x_12; } } LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalIntro_introStep___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) { diff --git a/stage0/stdlib/Lean/Elab/Tactic/Classical.c b/stage0/stdlib/Lean/Elab/Tactic/Classical.c index 91417d16259c..faf7dfb106e5 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/Classical.c +++ b/stage0/stdlib/Lean/Elab/Tactic/Classical.c @@ -13,66 +13,73 @@ #ifdef __cplusplus extern "C" { #endif -static lean_object* l_Lean_Elab_Tactic_classical___at_Lean_Elab_Tactic_evalClassical___spec__2___closed__4; lean_object* l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_classical___rarg___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_evalClassical___closed__3; lean_object* l_Lean_ScopedEnvExtension_pushScope___rarg___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_classical___rarg___lambda__3___closed__1; extern lean_object* l_Lean_Meta_instanceExtension; static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalClassical__1___closed__1; -static lean_object* l_Lean_Elab_Tactic_evalClassical___closed__7; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_classical___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Elab_Tactic_tacticElabAttribute; +lean_object* l_Array_toSubarray___rarg(lean_object*, lean_object*, lean_object*); +static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalClassical__1___closed__10; +static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalClassical__1___closed__8; +lean_object* l_Lean_Syntax_getArgs(lean_object*); static lean_object* l_Lean_Elab_Tactic_evalClassical___closed__1; +static lean_object* l_Lean_Elab_Tactic_classical___at_Lean_Elab_Tactic_evalClassical___spec__4___closed__3; lean_object* l_Lean_ScopedEnvExtension_pushScope___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalClassical(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_evalClassical__1___closed__4; -static lean_object* l_Lean_Elab_Tactic_classical___at_Lean_Elab_Tactic_evalClassical___spec__2___closed__2; -uint8_t l_Lean_Syntax_isOfKind(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_classical___at_Lean_Elab_Tactic_evalClassical___spec__4___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_classical___rarg___lambda__2___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_classical___at_Lean_Elab_Tactic_evalClassical___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_classical___rarg___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_classical___rarg___lambda__4___closed__4; -static lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalClassical___spec__1___rarg___closed__2; -LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalClassical___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_withNarrowedArgTacticReuse___at_Lean_Elab_Tactic_evalClassical___spec__1___lambda__1(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalClassical__1___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalClassical__1___closed__5; -static lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalClassical___spec__1___rarg___closed__1; -static lean_object* l_Lean_Elab_Tactic_evalClassical___closed__2; +static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalClassical__1___closed__7; +lean_object* l_Lean_Syntax_eqWithInfoAndTraceReuse(lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_take(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_withReuseContext___at_Lean_Elab_Tactic_evalClassical___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___regBuiltin_Lean_Elab_Tactic_evalClassical__1___closed__6; +LEAN_EXPORT lean_object* l_Lean_Elab_Term_withNarrowedTacticReuse___at_Lean_Elab_Tactic_evalClassical___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_Lean_Elab_Tactic_classical___rarg___lambda__1(lean_object*, lean_object*); lean_object* l_Lean_ScopedEnvExtension_popScope___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_classical___rarg___lambda__3___closed__2; +LEAN_EXPORT lean_object* l_Lean_Elab_Term_withNarrowedArgTacticReuse___at_Lean_Elab_Tactic_evalClassical___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_classical___rarg___lambda__4___closed__3; -static lean_object* l_Lean_Elab_Tactic_classical___at_Lean_Elab_Tactic_evalClassical___spec__2___closed__1; +lean_object* l_Lean_Elab_addBuiltinIncrementalElab(lean_object*, lean_object*); +static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalClassical__1___closed__9; static lean_object* l_Lean_Elab_Tactic_classical___rarg___lambda__3___closed__3; +static lean_object* l_Lean_Elab_Tactic_classical___at_Lean_Elab_Tactic_evalClassical___spec__4___closed__2; +lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_withNarrowedArgTacticReuse___at_Lean_Elab_Tactic_evalClassical___spec__1___lambda__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_classical___rarg___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_classical___rarg___closed__1; lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_evalClassical___closed__5; -LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalClassical___spec__1___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_classical___rarg___lambda__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_classical___rarg___lambda__4___closed__1; lean_object* l_Lean_ScopedEnvExtension_popScope___rarg___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_classical___rarg___lambda__2(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalClassical___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_Tactic_classical___at_Lean_Elab_Tactic_evalClassical___spec__2___closed__3; -static lean_object* l_Lean_Elab_Tactic_evalClassical___closed__4; -static lean_object* l_Lean_Elab_Tactic_evalClassical___closed__6; lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_evalClassical__2(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_classical___rarg___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_addInstance(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_classical___at_Lean_Elab_Tactic_evalClassical___spec__4___closed__4; lean_object* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); +lean_object* l_Array_ofSubarray___rarg(lean_object*); +uint8_t l_Lean_Syntax_hasMissing(lean_object*); +static lean_object* l_Lean_Elab_Term_withNarrowedArgTacticReuse___at_Lean_Elab_Tactic_evalClassical___spec__1___lambda__1___closed__2; lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_withNarrowedArgTacticReuse___at_Lean_Elab_Tactic_evalClassical___spec__1___lambda__1___closed__1; lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Tactic_evalTactic(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -extern lean_object* l_Lean_Elab_unsupportedSyntaxExceptionId; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_classical(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_evalClassical__1(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalClassical__1___closed__2; lean_object* l_Lean_Meta_addInstance___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_classical___rarg___lambda__4___closed__2; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_classical___at_Lean_Elab_Tactic_evalClassical___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_classical___rarg___lambda__1(lean_object* x_1, lean_object* x_2) { _start: { @@ -273,46 +280,1192 @@ lean_dec(x_7); return x_8; } } -static lean_object* _init_l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalClassical___spec__1___rarg___closed__1() { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_withReuseContext___at_Lean_Elab_Tactic_evalClassical___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +uint8_t x_12; +x_12 = lean_ctor_get_uint8(x_9, sizeof(void*)*12 + 1); +if (x_12 == 0) +{ +uint8_t x_13; +x_13 = !lean_is_exclusive(x_9); +if (x_13 == 0) +{ +lean_object* x_14; uint8_t x_15; lean_object* x_16; +x_14 = lean_ctor_get(x_9, 5); +lean_dec(x_14); +x_15 = 0; +lean_ctor_set(x_9, 5, x_1); +lean_ctor_set_uint8(x_9, sizeof(void*)*12 + 1, x_15); +x_16 = lean_apply_9(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +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; uint8_t x_27; lean_object* x_28; uint8_t x_29; lean_object* x_30; lean_object* x_31; +x_17 = lean_ctor_get(x_9, 0); +x_18 = lean_ctor_get(x_9, 1); +x_19 = lean_ctor_get(x_9, 2); +x_20 = lean_ctor_get(x_9, 3); +x_21 = lean_ctor_get(x_9, 4); +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*)*12); +x_28 = lean_ctor_get(x_9, 11); +lean_inc(x_28); +lean_inc(x_26); +lean_inc(x_25); +lean_inc(x_24); +lean_inc(x_23); +lean_inc(x_22); +lean_inc(x_21); +lean_inc(x_20); +lean_inc(x_19); +lean_inc(x_18); +lean_inc(x_17); +lean_dec(x_9); +x_29 = 0; +x_30 = lean_alloc_ctor(0, 12, 2); +lean_ctor_set(x_30, 0, x_17); +lean_ctor_set(x_30, 1, x_18); +lean_ctor_set(x_30, 2, x_19); +lean_ctor_set(x_30, 3, x_20); +lean_ctor_set(x_30, 4, x_21); +lean_ctor_set(x_30, 5, x_1); +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(x_30, 11, x_28); +lean_ctor_set_uint8(x_30, sizeof(void*)*12, x_27); +lean_ctor_set_uint8(x_30, sizeof(void*)*12 + 1, x_29); +x_31 = lean_apply_9(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_30, x_10, x_11); +return x_31; +} +} +else +{ +uint8_t x_32; +x_32 = !lean_is_exclusive(x_9); +if (x_32 == 0) +{ +lean_object* x_33; uint8_t x_34; lean_object* x_35; +x_33 = lean_ctor_get(x_9, 5); +lean_dec(x_33); +lean_inc(x_1); +x_34 = l_Lean_Syntax_hasMissing(x_1); +lean_ctor_set(x_9, 5, x_1); +lean_ctor_set_uint8(x_9, sizeof(void*)*12 + 1, x_34); +x_35 = lean_apply_9(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +return x_35; +} +else +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; uint8_t x_46; lean_object* x_47; uint8_t x_48; lean_object* x_49; lean_object* x_50; +x_36 = lean_ctor_get(x_9, 0); +x_37 = lean_ctor_get(x_9, 1); +x_38 = lean_ctor_get(x_9, 2); +x_39 = lean_ctor_get(x_9, 3); +x_40 = lean_ctor_get(x_9, 4); +x_41 = lean_ctor_get(x_9, 6); +x_42 = lean_ctor_get(x_9, 7); +x_43 = lean_ctor_get(x_9, 8); +x_44 = lean_ctor_get(x_9, 9); +x_45 = lean_ctor_get(x_9, 10); +x_46 = lean_ctor_get_uint8(x_9, sizeof(void*)*12); +x_47 = lean_ctor_get(x_9, 11); +lean_inc(x_47); +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_dec(x_9); +lean_inc(x_1); +x_48 = l_Lean_Syntax_hasMissing(x_1); +x_49 = lean_alloc_ctor(0, 12, 2); +lean_ctor_set(x_49, 0, x_36); +lean_ctor_set(x_49, 1, x_37); +lean_ctor_set(x_49, 2, x_38); +lean_ctor_set(x_49, 3, x_39); +lean_ctor_set(x_49, 4, x_40); +lean_ctor_set(x_49, 5, x_1); +lean_ctor_set(x_49, 6, x_41); +lean_ctor_set(x_49, 7, x_42); +lean_ctor_set(x_49, 8, x_43); +lean_ctor_set(x_49, 9, x_44); +lean_ctor_set(x_49, 10, x_45); +lean_ctor_set(x_49, 11, x_47); +lean_ctor_set_uint8(x_49, sizeof(void*)*12, x_46); +lean_ctor_set_uint8(x_49, sizeof(void*)*12 + 1, x_48); +x_50 = lean_apply_9(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_49, x_10, x_11); +return x_50; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Term_withNarrowedTacticReuse___at_Lean_Elab_Tactic_evalClassical___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +lean_inc(x_1); +x_13 = lean_apply_1(x_1, x_3); +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +x_16 = lean_ctor_get(x_10, 2); +lean_inc(x_16); +lean_inc(x_15); +x_17 = lean_apply_1(x_2, x_15); +x_18 = lean_ctor_get(x_6, 8); +lean_inc(x_18); +if (lean_obj_tag(x_18) == 0) +{ +uint8_t x_19; +lean_dec(x_16); +lean_dec(x_14); +lean_dec(x_1); +x_19 = !lean_is_exclusive(x_6); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_6, 8); +lean_dec(x_20); +x_21 = lean_box(0); +lean_ctor_set(x_6, 8, x_21); +x_22 = l_Lean_Elab_Term_withReuseContext___at_Lean_Elab_Tactic_evalClassical___spec__3(x_15, x_17, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_22; +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; uint8_t x_27; uint8_t x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; uint8_t x_34; uint8_t x_35; uint8_t x_36; uint8_t x_37; lean_object* x_38; uint8_t x_39; uint8_t x_40; uint8_t x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_23 = lean_ctor_get(x_6, 0); +x_24 = lean_ctor_get(x_6, 1); +x_25 = lean_ctor_get(x_6, 2); +x_26 = lean_ctor_get_uint8(x_6, sizeof(void*)*9); +x_27 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 1); +x_28 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 2); +x_29 = lean_ctor_get(x_6, 3); +x_30 = lean_ctor_get(x_6, 4); +x_31 = lean_ctor_get(x_6, 5); +x_32 = lean_ctor_get(x_6, 6); +x_33 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 3); +x_34 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 4); +x_35 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 5); +x_36 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 6); +x_37 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 7); +x_38 = lean_ctor_get(x_6, 7); +x_39 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 8); +x_40 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 9); +x_41 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 10); +lean_inc(x_38); +lean_inc(x_32); +lean_inc(x_31); +lean_inc(x_30); +lean_inc(x_29); +lean_inc(x_25); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_6); +x_42 = lean_box(0); +x_43 = lean_alloc_ctor(0, 9, 11); +lean_ctor_set(x_43, 0, x_23); +lean_ctor_set(x_43, 1, x_24); +lean_ctor_set(x_43, 2, x_25); +lean_ctor_set(x_43, 3, x_29); +lean_ctor_set(x_43, 4, x_30); +lean_ctor_set(x_43, 5, x_31); +lean_ctor_set(x_43, 6, x_32); +lean_ctor_set(x_43, 7, x_38); +lean_ctor_set(x_43, 8, x_42); +lean_ctor_set_uint8(x_43, sizeof(void*)*9, x_26); +lean_ctor_set_uint8(x_43, sizeof(void*)*9 + 1, x_27); +lean_ctor_set_uint8(x_43, sizeof(void*)*9 + 2, x_28); +lean_ctor_set_uint8(x_43, sizeof(void*)*9 + 3, x_33); +lean_ctor_set_uint8(x_43, sizeof(void*)*9 + 4, x_34); +lean_ctor_set_uint8(x_43, sizeof(void*)*9 + 5, x_35); +lean_ctor_set_uint8(x_43, sizeof(void*)*9 + 6, x_36); +lean_ctor_set_uint8(x_43, sizeof(void*)*9 + 7, x_37); +lean_ctor_set_uint8(x_43, sizeof(void*)*9 + 8, x_39); +lean_ctor_set_uint8(x_43, sizeof(void*)*9 + 9, x_40); +lean_ctor_set_uint8(x_43, sizeof(void*)*9 + 10, x_41); +x_44 = l_Lean_Elab_Term_withReuseContext___at_Lean_Elab_Tactic_evalClassical___spec__3(x_15, x_17, x_4, x_5, x_43, x_7, x_8, x_9, x_10, x_11, x_12); +return x_44; +} +} +else +{ +uint8_t x_45; +x_45 = !lean_is_exclusive(x_18); +if (x_45 == 0) +{ +lean_object* x_46; lean_object* x_47; +x_46 = lean_ctor_get(x_18, 0); +x_47 = lean_ctor_get(x_46, 0); +lean_inc(x_47); +if (lean_obj_tag(x_47) == 0) +{ +uint8_t x_48; +lean_dec(x_16); +lean_dec(x_14); +lean_dec(x_1); +x_48 = !lean_is_exclusive(x_6); +if (x_48 == 0) +{ +lean_object* x_49; uint8_t x_50; +x_49 = lean_ctor_get(x_6, 8); +lean_dec(x_49); +x_50 = !lean_is_exclusive(x_46); +if (x_50 == 0) +{ +lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_51 = lean_ctor_get(x_46, 0); +lean_dec(x_51); +x_52 = lean_box(0); +lean_ctor_set(x_46, 0, x_52); +x_53 = l_Lean_Elab_Term_withReuseContext___at_Lean_Elab_Tactic_evalClassical___spec__3(x_15, x_17, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_53; +} +else +{ +lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_54 = lean_ctor_get(x_46, 1); +lean_inc(x_54); +lean_dec(x_46); +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_54); +lean_ctor_set(x_18, 0, x_56); +x_57 = l_Lean_Elab_Term_withReuseContext___at_Lean_Elab_Tactic_evalClassical___spec__3(x_15, x_17, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_57; +} +} +else +{ +lean_object* x_58; lean_object* x_59; lean_object* x_60; uint8_t x_61; uint8_t x_62; uint8_t 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; uint8_t x_70; uint8_t x_71; uint8_t x_72; lean_object* x_73; uint8_t x_74; uint8_t 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; lean_object* x_82; +x_58 = lean_ctor_get(x_6, 0); +x_59 = lean_ctor_get(x_6, 1); +x_60 = lean_ctor_get(x_6, 2); +x_61 = lean_ctor_get_uint8(x_6, sizeof(void*)*9); +x_62 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 1); +x_63 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 2); +x_64 = lean_ctor_get(x_6, 3); +x_65 = lean_ctor_get(x_6, 4); +x_66 = lean_ctor_get(x_6, 5); +x_67 = lean_ctor_get(x_6, 6); +x_68 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 3); +x_69 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 4); +x_70 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 5); +x_71 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 6); +x_72 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 7); +x_73 = lean_ctor_get(x_6, 7); +x_74 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 8); +x_75 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 9); +x_76 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 10); +lean_inc(x_73); +lean_inc(x_67); +lean_inc(x_66); +lean_inc(x_65); +lean_inc(x_64); +lean_inc(x_60); +lean_inc(x_59); +lean_inc(x_58); +lean_dec(x_6); +x_77 = lean_ctor_get(x_46, 1); +lean_inc(x_77); +if (lean_is_exclusive(x_46)) { + lean_ctor_release(x_46, 0); + lean_ctor_release(x_46, 1); + x_78 = x_46; +} else { + lean_dec_ref(x_46); + x_78 = lean_box(0); +} +x_79 = lean_box(0); +if (lean_is_scalar(x_78)) { + x_80 = lean_alloc_ctor(0, 2, 0); +} else { + x_80 = x_78; +} +lean_ctor_set(x_80, 0, x_79); +lean_ctor_set(x_80, 1, x_77); +lean_ctor_set(x_18, 0, x_80); +x_81 = lean_alloc_ctor(0, 9, 11); +lean_ctor_set(x_81, 0, x_58); +lean_ctor_set(x_81, 1, x_59); +lean_ctor_set(x_81, 2, x_60); +lean_ctor_set(x_81, 3, x_64); +lean_ctor_set(x_81, 4, x_65); +lean_ctor_set(x_81, 5, x_66); +lean_ctor_set(x_81, 6, x_67); +lean_ctor_set(x_81, 7, x_73); +lean_ctor_set(x_81, 8, x_18); +lean_ctor_set_uint8(x_81, sizeof(void*)*9, x_61); +lean_ctor_set_uint8(x_81, sizeof(void*)*9 + 1, x_62); +lean_ctor_set_uint8(x_81, sizeof(void*)*9 + 2, x_63); +lean_ctor_set_uint8(x_81, sizeof(void*)*9 + 3, x_68); +lean_ctor_set_uint8(x_81, sizeof(void*)*9 + 4, x_69); +lean_ctor_set_uint8(x_81, sizeof(void*)*9 + 5, x_70); +lean_ctor_set_uint8(x_81, sizeof(void*)*9 + 6, x_71); +lean_ctor_set_uint8(x_81, sizeof(void*)*9 + 7, x_72); +lean_ctor_set_uint8(x_81, sizeof(void*)*9 + 8, x_74); +lean_ctor_set_uint8(x_81, sizeof(void*)*9 + 9, x_75); +lean_ctor_set_uint8(x_81, sizeof(void*)*9 + 10, x_76); +x_82 = l_Lean_Elab_Term_withReuseContext___at_Lean_Elab_Tactic_evalClassical___spec__3(x_15, x_17, x_4, x_5, x_81, x_7, x_8, x_9, x_10, x_11, x_12); +return x_82; +} +} +else +{ +uint8_t x_83; +x_83 = !lean_is_exclusive(x_47); +if (x_83 == 0) +{ +uint8_t x_84; +x_84 = !lean_is_exclusive(x_6); +if (x_84 == 0) +{ +lean_object* x_85; lean_object* x_86; uint8_t x_87; +x_85 = lean_ctor_get(x_47, 0); +x_86 = lean_ctor_get(x_6, 8); +lean_dec(x_86); +x_87 = !lean_is_exclusive(x_46); +if (x_87 == 0) +{ +lean_object* x_88; uint8_t x_89; +x_88 = lean_ctor_get(x_46, 0); +lean_dec(x_88); +x_89 = !lean_is_exclusive(x_85); +if (x_89 == 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; uint8_t x_96; +x_90 = lean_ctor_get(x_85, 0); +x_91 = lean_ctor_get(x_85, 1); +x_92 = lean_apply_1(x_1, 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 = l_Lean_Syntax_eqWithInfoAndTraceReuse(x_16, x_14, x_93); +lean_dec(x_16); +x_96 = lean_unbox(x_95); +lean_dec(x_95); +if (x_96 == 0) +{ +lean_object* x_97; lean_object* x_98; +lean_dec(x_94); +lean_free_object(x_85); +lean_dec(x_91); +lean_free_object(x_18); +x_97 = lean_box(0); +lean_ctor_set(x_46, 0, x_97); +lean_ctor_set(x_47, 0, x_46); +lean_ctor_set(x_6, 8, x_47); +x_98 = l_Lean_Elab_Term_withReuseContext___at_Lean_Elab_Tactic_evalClassical___spec__3(x_15, x_17, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_98; +} +else +{ +lean_object* x_99; +lean_ctor_set(x_85, 0, x_94); +x_99 = l_Lean_Elab_Term_withReuseContext___at_Lean_Elab_Tactic_evalClassical___spec__3(x_15, x_17, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_99; +} +} +else +{ +lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; uint8_t x_106; +x_100 = lean_ctor_get(x_85, 0); +x_101 = lean_ctor_get(x_85, 1); +lean_inc(x_101); +lean_inc(x_100); +lean_dec(x_85); +x_102 = lean_apply_1(x_1, x_100); +x_103 = lean_ctor_get(x_102, 0); +lean_inc(x_103); +x_104 = lean_ctor_get(x_102, 1); +lean_inc(x_104); +lean_dec(x_102); +x_105 = l_Lean_Syntax_eqWithInfoAndTraceReuse(x_16, x_14, x_103); +lean_dec(x_16); +x_106 = lean_unbox(x_105); +lean_dec(x_105); +if (x_106 == 0) +{ +lean_object* x_107; lean_object* x_108; +lean_dec(x_104); +lean_dec(x_101); +lean_free_object(x_18); +x_107 = lean_box(0); +lean_ctor_set(x_46, 0, x_107); +lean_ctor_set(x_47, 0, x_46); +lean_ctor_set(x_6, 8, x_47); +x_108 = l_Lean_Elab_Term_withReuseContext___at_Lean_Elab_Tactic_evalClassical___spec__3(x_15, x_17, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_108; +} +else +{ +lean_object* x_109; lean_object* x_110; +x_109 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_109, 0, x_104); +lean_ctor_set(x_109, 1, x_101); +lean_ctor_set(x_47, 0, x_109); +x_110 = l_Lean_Elab_Term_withReuseContext___at_Lean_Elab_Tactic_evalClassical___spec__3(x_15, x_17, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_110; +} +} +} +else +{ +lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; uint8_t x_119; +x_111 = lean_ctor_get(x_46, 1); +lean_inc(x_111); +lean_dec(x_46); +x_112 = lean_ctor_get(x_85, 0); +lean_inc(x_112); +x_113 = lean_ctor_get(x_85, 1); +lean_inc(x_113); +if (lean_is_exclusive(x_85)) { + lean_ctor_release(x_85, 0); + lean_ctor_release(x_85, 1); + x_114 = x_85; +} else { + lean_dec_ref(x_85); + x_114 = lean_box(0); +} +x_115 = lean_apply_1(x_1, x_112); +x_116 = lean_ctor_get(x_115, 0); +lean_inc(x_116); +x_117 = lean_ctor_get(x_115, 1); +lean_inc(x_117); +lean_dec(x_115); +x_118 = l_Lean_Syntax_eqWithInfoAndTraceReuse(x_16, x_14, x_116); +lean_dec(x_16); +x_119 = lean_unbox(x_118); +lean_dec(x_118); +if (x_119 == 0) +{ +lean_object* x_120; lean_object* x_121; lean_object* x_122; +lean_dec(x_117); +lean_dec(x_114); +lean_dec(x_113); +lean_free_object(x_18); +x_120 = lean_box(0); +x_121 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_121, 0, x_120); +lean_ctor_set(x_121, 1, x_111); +lean_ctor_set(x_47, 0, x_121); +lean_ctor_set(x_6, 8, x_47); +x_122 = l_Lean_Elab_Term_withReuseContext___at_Lean_Elab_Tactic_evalClassical___spec__3(x_15, x_17, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_122; +} +else +{ +lean_object* x_123; lean_object* x_124; lean_object* x_125; +if (lean_is_scalar(x_114)) { + x_123 = lean_alloc_ctor(0, 2, 0); +} else { + x_123 = x_114; +} +lean_ctor_set(x_123, 0, x_117); +lean_ctor_set(x_123, 1, x_113); +lean_ctor_set(x_47, 0, x_123); +x_124 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_124, 0, x_47); +lean_ctor_set(x_124, 1, x_111); +lean_ctor_set(x_18, 0, x_124); +x_125 = l_Lean_Elab_Term_withReuseContext___at_Lean_Elab_Tactic_evalClassical___spec__3(x_15, x_17, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_125; +} +} +} +else +{ +lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; uint8_t x_130; uint8_t x_131; uint8_t x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; uint8_t x_137; uint8_t x_138; uint8_t x_139; uint8_t x_140; uint8_t x_141; lean_object* x_142; uint8_t x_143; uint8_t 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; uint8_t x_155; +x_126 = lean_ctor_get(x_47, 0); +x_127 = lean_ctor_get(x_6, 0); +x_128 = lean_ctor_get(x_6, 1); +x_129 = lean_ctor_get(x_6, 2); +x_130 = lean_ctor_get_uint8(x_6, sizeof(void*)*9); +x_131 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 1); +x_132 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 2); +x_133 = lean_ctor_get(x_6, 3); +x_134 = lean_ctor_get(x_6, 4); +x_135 = lean_ctor_get(x_6, 5); +x_136 = lean_ctor_get(x_6, 6); +x_137 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 3); +x_138 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 4); +x_139 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 5); +x_140 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 6); +x_141 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 7); +x_142 = lean_ctor_get(x_6, 7); +x_143 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 8); +x_144 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 9); +x_145 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 10); +lean_inc(x_142); +lean_inc(x_136); +lean_inc(x_135); +lean_inc(x_134); +lean_inc(x_133); +lean_inc(x_129); +lean_inc(x_128); +lean_inc(x_127); +lean_dec(x_6); +x_146 = lean_ctor_get(x_46, 1); +lean_inc(x_146); +if (lean_is_exclusive(x_46)) { + lean_ctor_release(x_46, 0); + lean_ctor_release(x_46, 1); + x_147 = x_46; +} else { + lean_dec_ref(x_46); + x_147 = lean_box(0); +} +x_148 = lean_ctor_get(x_126, 0); +lean_inc(x_148); +x_149 = lean_ctor_get(x_126, 1); +lean_inc(x_149); +if (lean_is_exclusive(x_126)) { + lean_ctor_release(x_126, 0); + lean_ctor_release(x_126, 1); + x_150 = x_126; +} else { + lean_dec_ref(x_126); + x_150 = lean_box(0); +} +x_151 = lean_apply_1(x_1, x_148); +x_152 = lean_ctor_get(x_151, 0); +lean_inc(x_152); +x_153 = lean_ctor_get(x_151, 1); +lean_inc(x_153); +lean_dec(x_151); +x_154 = l_Lean_Syntax_eqWithInfoAndTraceReuse(x_16, x_14, x_152); +lean_dec(x_16); +x_155 = lean_unbox(x_154); +lean_dec(x_154); +if (x_155 == 0) +{ +lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; +lean_dec(x_153); +lean_dec(x_150); +lean_dec(x_149); +lean_free_object(x_18); +x_156 = lean_box(0); +if (lean_is_scalar(x_147)) { + x_157 = lean_alloc_ctor(0, 2, 0); +} else { + x_157 = x_147; +} +lean_ctor_set(x_157, 0, x_156); +lean_ctor_set(x_157, 1, x_146); +lean_ctor_set(x_47, 0, x_157); +x_158 = lean_alloc_ctor(0, 9, 11); +lean_ctor_set(x_158, 0, x_127); +lean_ctor_set(x_158, 1, x_128); +lean_ctor_set(x_158, 2, x_129); +lean_ctor_set(x_158, 3, x_133); +lean_ctor_set(x_158, 4, x_134); +lean_ctor_set(x_158, 5, x_135); +lean_ctor_set(x_158, 6, x_136); +lean_ctor_set(x_158, 7, x_142); +lean_ctor_set(x_158, 8, x_47); +lean_ctor_set_uint8(x_158, sizeof(void*)*9, x_130); +lean_ctor_set_uint8(x_158, sizeof(void*)*9 + 1, x_131); +lean_ctor_set_uint8(x_158, sizeof(void*)*9 + 2, x_132); +lean_ctor_set_uint8(x_158, sizeof(void*)*9 + 3, x_137); +lean_ctor_set_uint8(x_158, sizeof(void*)*9 + 4, x_138); +lean_ctor_set_uint8(x_158, sizeof(void*)*9 + 5, x_139); +lean_ctor_set_uint8(x_158, sizeof(void*)*9 + 6, x_140); +lean_ctor_set_uint8(x_158, sizeof(void*)*9 + 7, x_141); +lean_ctor_set_uint8(x_158, sizeof(void*)*9 + 8, x_143); +lean_ctor_set_uint8(x_158, sizeof(void*)*9 + 9, x_144); +lean_ctor_set_uint8(x_158, sizeof(void*)*9 + 10, x_145); +x_159 = l_Lean_Elab_Term_withReuseContext___at_Lean_Elab_Tactic_evalClassical___spec__3(x_15, x_17, x_4, x_5, x_158, x_7, x_8, x_9, x_10, x_11, x_12); +return x_159; +} +else +{ +lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; +if (lean_is_scalar(x_150)) { + x_160 = lean_alloc_ctor(0, 2, 0); +} else { + x_160 = x_150; +} +lean_ctor_set(x_160, 0, x_153); +lean_ctor_set(x_160, 1, x_149); +lean_ctor_set(x_47, 0, x_160); +if (lean_is_scalar(x_147)) { + x_161 = lean_alloc_ctor(0, 2, 0); +} else { + x_161 = x_147; +} +lean_ctor_set(x_161, 0, x_47); +lean_ctor_set(x_161, 1, x_146); +lean_ctor_set(x_18, 0, x_161); +x_162 = lean_alloc_ctor(0, 9, 11); +lean_ctor_set(x_162, 0, x_127); +lean_ctor_set(x_162, 1, x_128); +lean_ctor_set(x_162, 2, x_129); +lean_ctor_set(x_162, 3, x_133); +lean_ctor_set(x_162, 4, x_134); +lean_ctor_set(x_162, 5, x_135); +lean_ctor_set(x_162, 6, x_136); +lean_ctor_set(x_162, 7, x_142); +lean_ctor_set(x_162, 8, x_18); +lean_ctor_set_uint8(x_162, sizeof(void*)*9, x_130); +lean_ctor_set_uint8(x_162, sizeof(void*)*9 + 1, x_131); +lean_ctor_set_uint8(x_162, sizeof(void*)*9 + 2, x_132); +lean_ctor_set_uint8(x_162, sizeof(void*)*9 + 3, x_137); +lean_ctor_set_uint8(x_162, sizeof(void*)*9 + 4, x_138); +lean_ctor_set_uint8(x_162, sizeof(void*)*9 + 5, x_139); +lean_ctor_set_uint8(x_162, sizeof(void*)*9 + 6, x_140); +lean_ctor_set_uint8(x_162, sizeof(void*)*9 + 7, x_141); +lean_ctor_set_uint8(x_162, sizeof(void*)*9 + 8, x_143); +lean_ctor_set_uint8(x_162, sizeof(void*)*9 + 9, x_144); +lean_ctor_set_uint8(x_162, sizeof(void*)*9 + 10, x_145); +x_163 = l_Lean_Elab_Term_withReuseContext___at_Lean_Elab_Tactic_evalClassical___spec__3(x_15, x_17, x_4, x_5, x_162, x_7, x_8, x_9, x_10, x_11, x_12); +return x_163; +} +} +} +else +{ +lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; uint8_t x_168; uint8_t x_169; uint8_t x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; uint8_t x_175; uint8_t x_176; uint8_t x_177; uint8_t x_178; uint8_t x_179; lean_object* x_180; uint8_t x_181; uint8_t x_182; uint8_t 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; uint8_t x_194; +x_164 = lean_ctor_get(x_47, 0); +lean_inc(x_164); +lean_dec(x_47); +x_165 = lean_ctor_get(x_6, 0); +lean_inc(x_165); +x_166 = lean_ctor_get(x_6, 1); +lean_inc(x_166); +x_167 = lean_ctor_get(x_6, 2); +lean_inc(x_167); +x_168 = lean_ctor_get_uint8(x_6, sizeof(void*)*9); +x_169 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 1); +x_170 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 2); +x_171 = lean_ctor_get(x_6, 3); +lean_inc(x_171); +x_172 = lean_ctor_get(x_6, 4); +lean_inc(x_172); +x_173 = lean_ctor_get(x_6, 5); +lean_inc(x_173); +x_174 = lean_ctor_get(x_6, 6); +lean_inc(x_174); +x_175 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 3); +x_176 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 4); +x_177 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 5); +x_178 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 6); +x_179 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 7); +x_180 = lean_ctor_get(x_6, 7); +lean_inc(x_180); +x_181 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 8); +x_182 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 9); +x_183 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 10); +if (lean_is_exclusive(x_6)) { + lean_ctor_release(x_6, 0); + lean_ctor_release(x_6, 1); + lean_ctor_release(x_6, 2); + lean_ctor_release(x_6, 3); + lean_ctor_release(x_6, 4); + lean_ctor_release(x_6, 5); + lean_ctor_release(x_6, 6); + lean_ctor_release(x_6, 7); + lean_ctor_release(x_6, 8); + x_184 = x_6; +} else { + lean_dec_ref(x_6); + x_184 = lean_box(0); +} +x_185 = lean_ctor_get(x_46, 1); +lean_inc(x_185); +if (lean_is_exclusive(x_46)) { + lean_ctor_release(x_46, 0); + lean_ctor_release(x_46, 1); + x_186 = x_46; +} else { + lean_dec_ref(x_46); + x_186 = lean_box(0); +} +x_187 = lean_ctor_get(x_164, 0); +lean_inc(x_187); +x_188 = lean_ctor_get(x_164, 1); +lean_inc(x_188); +if (lean_is_exclusive(x_164)) { + lean_ctor_release(x_164, 0); + lean_ctor_release(x_164, 1); + x_189 = x_164; +} else { + lean_dec_ref(x_164); + x_189 = lean_box(0); +} +x_190 = lean_apply_1(x_1, x_187); +x_191 = lean_ctor_get(x_190, 0); +lean_inc(x_191); +x_192 = lean_ctor_get(x_190, 1); +lean_inc(x_192); +lean_dec(x_190); +x_193 = l_Lean_Syntax_eqWithInfoAndTraceReuse(x_16, x_14, x_191); +lean_dec(x_16); +x_194 = lean_unbox(x_193); +lean_dec(x_193); +if (x_194 == 0) +{ +lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; +lean_dec(x_192); +lean_dec(x_189); +lean_dec(x_188); +lean_free_object(x_18); +x_195 = lean_box(0); +if (lean_is_scalar(x_186)) { + x_196 = lean_alloc_ctor(0, 2, 0); +} else { + x_196 = x_186; +} +lean_ctor_set(x_196, 0, x_195); +lean_ctor_set(x_196, 1, x_185); +x_197 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_197, 0, x_196); +if (lean_is_scalar(x_184)) { + x_198 = lean_alloc_ctor(0, 9, 11); +} else { + x_198 = x_184; +} +lean_ctor_set(x_198, 0, x_165); +lean_ctor_set(x_198, 1, x_166); +lean_ctor_set(x_198, 2, x_167); +lean_ctor_set(x_198, 3, x_171); +lean_ctor_set(x_198, 4, x_172); +lean_ctor_set(x_198, 5, x_173); +lean_ctor_set(x_198, 6, x_174); +lean_ctor_set(x_198, 7, x_180); +lean_ctor_set(x_198, 8, x_197); +lean_ctor_set_uint8(x_198, sizeof(void*)*9, x_168); +lean_ctor_set_uint8(x_198, sizeof(void*)*9 + 1, x_169); +lean_ctor_set_uint8(x_198, sizeof(void*)*9 + 2, x_170); +lean_ctor_set_uint8(x_198, sizeof(void*)*9 + 3, x_175); +lean_ctor_set_uint8(x_198, sizeof(void*)*9 + 4, x_176); +lean_ctor_set_uint8(x_198, sizeof(void*)*9 + 5, x_177); +lean_ctor_set_uint8(x_198, sizeof(void*)*9 + 6, x_178); +lean_ctor_set_uint8(x_198, sizeof(void*)*9 + 7, x_179); +lean_ctor_set_uint8(x_198, sizeof(void*)*9 + 8, x_181); +lean_ctor_set_uint8(x_198, sizeof(void*)*9 + 9, x_182); +lean_ctor_set_uint8(x_198, sizeof(void*)*9 + 10, x_183); +x_199 = l_Lean_Elab_Term_withReuseContext___at_Lean_Elab_Tactic_evalClassical___spec__3(x_15, x_17, x_4, x_5, x_198, x_7, x_8, x_9, x_10, x_11, x_12); +return x_199; +} +else +{ +lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; +if (lean_is_scalar(x_189)) { + x_200 = lean_alloc_ctor(0, 2, 0); +} else { + x_200 = x_189; +} +lean_ctor_set(x_200, 0, x_192); +lean_ctor_set(x_200, 1, x_188); +x_201 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_201, 0, x_200); +if (lean_is_scalar(x_186)) { + x_202 = lean_alloc_ctor(0, 2, 0); +} else { + x_202 = x_186; +} +lean_ctor_set(x_202, 0, x_201); +lean_ctor_set(x_202, 1, x_185); +lean_ctor_set(x_18, 0, x_202); +if (lean_is_scalar(x_184)) { + x_203 = lean_alloc_ctor(0, 9, 11); +} else { + x_203 = x_184; +} +lean_ctor_set(x_203, 0, x_165); +lean_ctor_set(x_203, 1, x_166); +lean_ctor_set(x_203, 2, x_167); +lean_ctor_set(x_203, 3, x_171); +lean_ctor_set(x_203, 4, x_172); +lean_ctor_set(x_203, 5, x_173); +lean_ctor_set(x_203, 6, x_174); +lean_ctor_set(x_203, 7, x_180); +lean_ctor_set(x_203, 8, x_18); +lean_ctor_set_uint8(x_203, sizeof(void*)*9, x_168); +lean_ctor_set_uint8(x_203, sizeof(void*)*9 + 1, x_169); +lean_ctor_set_uint8(x_203, sizeof(void*)*9 + 2, x_170); +lean_ctor_set_uint8(x_203, sizeof(void*)*9 + 3, x_175); +lean_ctor_set_uint8(x_203, sizeof(void*)*9 + 4, x_176); +lean_ctor_set_uint8(x_203, sizeof(void*)*9 + 5, x_177); +lean_ctor_set_uint8(x_203, sizeof(void*)*9 + 6, x_178); +lean_ctor_set_uint8(x_203, sizeof(void*)*9 + 7, x_179); +lean_ctor_set_uint8(x_203, sizeof(void*)*9 + 8, x_181); +lean_ctor_set_uint8(x_203, sizeof(void*)*9 + 9, x_182); +lean_ctor_set_uint8(x_203, sizeof(void*)*9 + 10, x_183); +x_204 = l_Lean_Elab_Term_withReuseContext___at_Lean_Elab_Tactic_evalClassical___spec__3(x_15, x_17, x_4, x_5, x_203, x_7, x_8, x_9, x_10, x_11, x_12); +return x_204; +} +} +} +} +else +{ +lean_object* x_205; lean_object* x_206; +x_205 = lean_ctor_get(x_18, 0); +lean_inc(x_205); +lean_dec(x_18); +x_206 = lean_ctor_get(x_205, 0); +lean_inc(x_206); +if (lean_obj_tag(x_206) == 0) +{ +lean_object* x_207; lean_object* x_208; lean_object* x_209; uint8_t x_210; uint8_t x_211; uint8_t x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; uint8_t x_217; uint8_t x_218; uint8_t x_219; uint8_t x_220; uint8_t x_221; lean_object* x_222; uint8_t x_223; uint8_t x_224; uint8_t 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_dec(x_16); +lean_dec(x_14); +lean_dec(x_1); +x_207 = lean_ctor_get(x_6, 0); +lean_inc(x_207); +x_208 = lean_ctor_get(x_6, 1); +lean_inc(x_208); +x_209 = lean_ctor_get(x_6, 2); +lean_inc(x_209); +x_210 = lean_ctor_get_uint8(x_6, sizeof(void*)*9); +x_211 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 1); +x_212 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 2); +x_213 = lean_ctor_get(x_6, 3); +lean_inc(x_213); +x_214 = lean_ctor_get(x_6, 4); +lean_inc(x_214); +x_215 = lean_ctor_get(x_6, 5); +lean_inc(x_215); +x_216 = lean_ctor_get(x_6, 6); +lean_inc(x_216); +x_217 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 3); +x_218 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 4); +x_219 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 5); +x_220 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 6); +x_221 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 7); +x_222 = lean_ctor_get(x_6, 7); +lean_inc(x_222); +x_223 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 8); +x_224 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 9); +x_225 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 10); +if (lean_is_exclusive(x_6)) { + lean_ctor_release(x_6, 0); + lean_ctor_release(x_6, 1); + lean_ctor_release(x_6, 2); + lean_ctor_release(x_6, 3); + lean_ctor_release(x_6, 4); + lean_ctor_release(x_6, 5); + lean_ctor_release(x_6, 6); + lean_ctor_release(x_6, 7); + lean_ctor_release(x_6, 8); + x_226 = x_6; +} else { + lean_dec_ref(x_6); + x_226 = lean_box(0); +} +x_227 = lean_ctor_get(x_205, 1); +lean_inc(x_227); +if (lean_is_exclusive(x_205)) { + lean_ctor_release(x_205, 0); + lean_ctor_release(x_205, 1); + x_228 = x_205; +} else { + lean_dec_ref(x_205); + x_228 = lean_box(0); +} +x_229 = lean_box(0); +if (lean_is_scalar(x_228)) { + x_230 = lean_alloc_ctor(0, 2, 0); +} else { + x_230 = x_228; +} +lean_ctor_set(x_230, 0, x_229); +lean_ctor_set(x_230, 1, x_227); +x_231 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_231, 0, x_230); +if (lean_is_scalar(x_226)) { + x_232 = lean_alloc_ctor(0, 9, 11); +} else { + x_232 = x_226; +} +lean_ctor_set(x_232, 0, x_207); +lean_ctor_set(x_232, 1, x_208); +lean_ctor_set(x_232, 2, x_209); +lean_ctor_set(x_232, 3, x_213); +lean_ctor_set(x_232, 4, x_214); +lean_ctor_set(x_232, 5, x_215); +lean_ctor_set(x_232, 6, x_216); +lean_ctor_set(x_232, 7, x_222); +lean_ctor_set(x_232, 8, x_231); +lean_ctor_set_uint8(x_232, sizeof(void*)*9, x_210); +lean_ctor_set_uint8(x_232, sizeof(void*)*9 + 1, x_211); +lean_ctor_set_uint8(x_232, sizeof(void*)*9 + 2, x_212); +lean_ctor_set_uint8(x_232, sizeof(void*)*9 + 3, x_217); +lean_ctor_set_uint8(x_232, sizeof(void*)*9 + 4, x_218); +lean_ctor_set_uint8(x_232, sizeof(void*)*9 + 5, x_219); +lean_ctor_set_uint8(x_232, sizeof(void*)*9 + 6, x_220); +lean_ctor_set_uint8(x_232, sizeof(void*)*9 + 7, x_221); +lean_ctor_set_uint8(x_232, sizeof(void*)*9 + 8, x_223); +lean_ctor_set_uint8(x_232, sizeof(void*)*9 + 9, x_224); +lean_ctor_set_uint8(x_232, sizeof(void*)*9 + 10, x_225); +x_233 = l_Lean_Elab_Term_withReuseContext___at_Lean_Elab_Tactic_evalClassical___spec__3(x_15, x_17, x_4, x_5, x_232, x_7, x_8, x_9, x_10, x_11, x_12); +return x_233; +} +else +{ +lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; uint8_t x_239; uint8_t x_240; uint8_t x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; uint8_t x_246; uint8_t x_247; uint8_t x_248; uint8_t x_249; uint8_t x_250; lean_object* x_251; uint8_t x_252; uint8_t x_253; uint8_t 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; uint8_t x_265; +x_234 = lean_ctor_get(x_206, 0); +lean_inc(x_234); +if (lean_is_exclusive(x_206)) { + lean_ctor_release(x_206, 0); + x_235 = x_206; +} else { + lean_dec_ref(x_206); + x_235 = lean_box(0); +} +x_236 = lean_ctor_get(x_6, 0); +lean_inc(x_236); +x_237 = lean_ctor_get(x_6, 1); +lean_inc(x_237); +x_238 = lean_ctor_get(x_6, 2); +lean_inc(x_238); +x_239 = lean_ctor_get_uint8(x_6, sizeof(void*)*9); +x_240 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 1); +x_241 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 2); +x_242 = lean_ctor_get(x_6, 3); +lean_inc(x_242); +x_243 = lean_ctor_get(x_6, 4); +lean_inc(x_243); +x_244 = lean_ctor_get(x_6, 5); +lean_inc(x_244); +x_245 = lean_ctor_get(x_6, 6); +lean_inc(x_245); +x_246 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 3); +x_247 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 4); +x_248 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 5); +x_249 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 6); +x_250 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 7); +x_251 = lean_ctor_get(x_6, 7); +lean_inc(x_251); +x_252 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 8); +x_253 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 9); +x_254 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 10); +if (lean_is_exclusive(x_6)) { + lean_ctor_release(x_6, 0); + lean_ctor_release(x_6, 1); + lean_ctor_release(x_6, 2); + lean_ctor_release(x_6, 3); + lean_ctor_release(x_6, 4); + lean_ctor_release(x_6, 5); + lean_ctor_release(x_6, 6); + lean_ctor_release(x_6, 7); + lean_ctor_release(x_6, 8); + x_255 = x_6; +} else { + lean_dec_ref(x_6); + x_255 = lean_box(0); +} +x_256 = lean_ctor_get(x_205, 1); +lean_inc(x_256); +if (lean_is_exclusive(x_205)) { + lean_ctor_release(x_205, 0); + lean_ctor_release(x_205, 1); + x_257 = x_205; +} else { + lean_dec_ref(x_205); + x_257 = lean_box(0); +} +x_258 = lean_ctor_get(x_234, 0); +lean_inc(x_258); +x_259 = lean_ctor_get(x_234, 1); +lean_inc(x_259); +if (lean_is_exclusive(x_234)) { + lean_ctor_release(x_234, 0); + lean_ctor_release(x_234, 1); + x_260 = x_234; +} else { + lean_dec_ref(x_234); + x_260 = lean_box(0); +} +x_261 = lean_apply_1(x_1, x_258); +x_262 = lean_ctor_get(x_261, 0); +lean_inc(x_262); +x_263 = lean_ctor_get(x_261, 1); +lean_inc(x_263); +lean_dec(x_261); +x_264 = l_Lean_Syntax_eqWithInfoAndTraceReuse(x_16, x_14, x_262); +lean_dec(x_16); +x_265 = lean_unbox(x_264); +lean_dec(x_264); +if (x_265 == 0) +{ +lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; +lean_dec(x_263); +lean_dec(x_260); +lean_dec(x_259); +x_266 = lean_box(0); +if (lean_is_scalar(x_257)) { + x_267 = lean_alloc_ctor(0, 2, 0); +} else { + x_267 = x_257; +} +lean_ctor_set(x_267, 0, x_266); +lean_ctor_set(x_267, 1, x_256); +if (lean_is_scalar(x_235)) { + x_268 = lean_alloc_ctor(1, 1, 0); +} else { + x_268 = x_235; +} +lean_ctor_set(x_268, 0, x_267); +if (lean_is_scalar(x_255)) { + x_269 = lean_alloc_ctor(0, 9, 11); +} else { + x_269 = x_255; +} +lean_ctor_set(x_269, 0, x_236); +lean_ctor_set(x_269, 1, x_237); +lean_ctor_set(x_269, 2, x_238); +lean_ctor_set(x_269, 3, x_242); +lean_ctor_set(x_269, 4, x_243); +lean_ctor_set(x_269, 5, x_244); +lean_ctor_set(x_269, 6, x_245); +lean_ctor_set(x_269, 7, x_251); +lean_ctor_set(x_269, 8, x_268); +lean_ctor_set_uint8(x_269, sizeof(void*)*9, x_239); +lean_ctor_set_uint8(x_269, sizeof(void*)*9 + 1, x_240); +lean_ctor_set_uint8(x_269, sizeof(void*)*9 + 2, x_241); +lean_ctor_set_uint8(x_269, sizeof(void*)*9 + 3, x_246); +lean_ctor_set_uint8(x_269, sizeof(void*)*9 + 4, x_247); +lean_ctor_set_uint8(x_269, sizeof(void*)*9 + 5, x_248); +lean_ctor_set_uint8(x_269, sizeof(void*)*9 + 6, x_249); +lean_ctor_set_uint8(x_269, sizeof(void*)*9 + 7, x_250); +lean_ctor_set_uint8(x_269, sizeof(void*)*9 + 8, x_252); +lean_ctor_set_uint8(x_269, sizeof(void*)*9 + 9, x_253); +lean_ctor_set_uint8(x_269, sizeof(void*)*9 + 10, x_254); +x_270 = l_Lean_Elab_Term_withReuseContext___at_Lean_Elab_Tactic_evalClassical___spec__3(x_15, x_17, x_4, x_5, x_269, x_7, x_8, x_9, x_10, x_11, x_12); +return x_270; +} +else +{ +lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; +if (lean_is_scalar(x_260)) { + x_271 = lean_alloc_ctor(0, 2, 0); +} else { + x_271 = x_260; +} +lean_ctor_set(x_271, 0, x_263); +lean_ctor_set(x_271, 1, x_259); +if (lean_is_scalar(x_235)) { + x_272 = lean_alloc_ctor(1, 1, 0); +} else { + x_272 = x_235; +} +lean_ctor_set(x_272, 0, x_271); +if (lean_is_scalar(x_257)) { + x_273 = lean_alloc_ctor(0, 2, 0); +} else { + x_273 = x_257; +} +lean_ctor_set(x_273, 0, x_272); +lean_ctor_set(x_273, 1, x_256); +x_274 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_274, 0, x_273); +if (lean_is_scalar(x_255)) { + x_275 = lean_alloc_ctor(0, 9, 11); +} else { + x_275 = x_255; +} +lean_ctor_set(x_275, 0, x_236); +lean_ctor_set(x_275, 1, x_237); +lean_ctor_set(x_275, 2, x_238); +lean_ctor_set(x_275, 3, x_242); +lean_ctor_set(x_275, 4, x_243); +lean_ctor_set(x_275, 5, x_244); +lean_ctor_set(x_275, 6, x_245); +lean_ctor_set(x_275, 7, x_251); +lean_ctor_set(x_275, 8, x_274); +lean_ctor_set_uint8(x_275, sizeof(void*)*9, x_239); +lean_ctor_set_uint8(x_275, sizeof(void*)*9 + 1, x_240); +lean_ctor_set_uint8(x_275, sizeof(void*)*9 + 2, x_241); +lean_ctor_set_uint8(x_275, sizeof(void*)*9 + 3, x_246); +lean_ctor_set_uint8(x_275, sizeof(void*)*9 + 4, x_247); +lean_ctor_set_uint8(x_275, sizeof(void*)*9 + 5, x_248); +lean_ctor_set_uint8(x_275, sizeof(void*)*9 + 6, x_249); +lean_ctor_set_uint8(x_275, sizeof(void*)*9 + 7, x_250); +lean_ctor_set_uint8(x_275, sizeof(void*)*9 + 8, x_252); +lean_ctor_set_uint8(x_275, sizeof(void*)*9 + 9, x_253); +lean_ctor_set_uint8(x_275, sizeof(void*)*9 + 10, x_254); +x_276 = l_Lean_Elab_Term_withReuseContext___at_Lean_Elab_Tactic_evalClassical___spec__3(x_15, x_17, x_4, x_5, x_275, x_7, x_8, x_9, x_10, x_11, x_12); +return x_276; +} +} +} +} +} +} +static lean_object* _init_l_Lean_Elab_Term_withNarrowedArgTacticReuse___at_Lean_Elab_Tactic_evalClassical___spec__1___lambda__1___closed__1() { _start: { lean_object* x_1; -x_1 = l_Lean_Elab_unsupportedSyntaxExceptionId; +x_1 = lean_mk_string_unchecked("null", 4, 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalClassical___spec__1___rarg___closed__2() { +static lean_object* _init_l_Lean_Elab_Term_withNarrowedArgTacticReuse___at_Lean_Elab_Tactic_evalClassical___spec__1___lambda__1___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalClassical___spec__1___rarg___closed__1; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set(x_3, 1, x_1); +x_2 = l_Lean_Elab_Term_withNarrowedArgTacticReuse___at_Lean_Elab_Tactic_evalClassical___spec__1___lambda__1___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalClassical___spec__1___rarg(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_withNarrowedArgTacticReuse___at_Lean_Elab_Tactic_evalClassical___spec__1___lambda__1(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalClassical___spec__1___rarg___closed__2; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set(x_3, 1, x_1); -return x_3; +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_3 = l_Lean_Syntax_getArgs(x_2); +x_4 = lean_unsigned_to_nat(0u); +lean_inc(x_1); +x_5 = l_Array_toSubarray___rarg(x_3, x_4, x_1); +x_6 = l_Array_ofSubarray___rarg(x_5); +lean_dec(x_5); +x_7 = lean_box(2); +x_8 = l_Lean_Elab_Term_withNarrowedArgTacticReuse___at_Lean_Elab_Tactic_evalClassical___spec__1___lambda__1___closed__2; +x_9 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_9, 0, x_7); +lean_ctor_set(x_9, 1, x_8); +lean_ctor_set(x_9, 2, x_6); +x_10 = l_Lean_Syntax_getArg(x_2, x_1); +lean_dec(x_1); +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_9); +lean_ctor_set(x_11, 1, x_10); +return x_11; } } -LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalClassical___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_Elab_Term_withNarrowedArgTacticReuse___at_Lean_Elab_Tactic_evalClassical___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: { -lean_object* x_9; -x_9 = lean_alloc_closure((void*)(l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalClassical___spec__1___rarg), 1, 0); -return x_9; +lean_object* x_13; lean_object* x_14; +x_13 = lean_alloc_closure((void*)(l_Lean_Elab_Term_withNarrowedArgTacticReuse___at_Lean_Elab_Tactic_evalClassical___spec__1___lambda__1___boxed), 2, 1); +lean_closure_set(x_13, 0, x_1); +x_14 = l_Lean_Elab_Term_withNarrowedTacticReuse___at_Lean_Elab_Tactic_evalClassical___spec__2(x_13, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_14; } } -static lean_object* _init_l_Lean_Elab_Tactic_classical___at_Lean_Elab_Tactic_evalClassical___spec__2___closed__1() { +static lean_object* _init_l_Lean_Elab_Tactic_classical___at_Lean_Elab_Tactic_evalClassical___spec__4___closed__1() { _start: { lean_object* x_1; @@ -320,32 +1473,32 @@ x_1 = l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_box(0), lean_box(0)); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_classical___at_Lean_Elab_Tactic_evalClassical___spec__2___closed__2() { +static lean_object* _init_l_Lean_Elab_Tactic_classical___at_Lean_Elab_Tactic_evalClassical___spec__4___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic_classical___at_Lean_Elab_Tactic_evalClassical___spec__2___closed__1; +x_1 = l_Lean_Elab_Tactic_classical___at_Lean_Elab_Tactic_evalClassical___spec__4___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_Tactic_classical___at_Lean_Elab_Tactic_evalClassical___spec__2___closed__3() { +static lean_object* _init_l_Lean_Elab_Tactic_classical___at_Lean_Elab_Tactic_evalClassical___spec__4___closed__3() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic_classical___at_Lean_Elab_Tactic_evalClassical___spec__2___closed__2; +x_1 = l_Lean_Elab_Tactic_classical___at_Lean_Elab_Tactic_evalClassical___spec__4___closed__2; x_2 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_2, 0, x_1); lean_ctor_set(x_2, 1, x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic_classical___at_Lean_Elab_Tactic_evalClassical___spec__2___closed__4() { +static lean_object* _init_l_Lean_Elab_Tactic_classical___at_Lean_Elab_Tactic_evalClassical___spec__4___closed__4() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic_classical___at_Lean_Elab_Tactic_evalClassical___spec__2___closed__2; +x_1 = l_Lean_Elab_Tactic_classical___at_Lean_Elab_Tactic_evalClassical___spec__4___closed__2; x_2 = lean_alloc_ctor(0, 6, 0); lean_ctor_set(x_2, 0, x_1); lean_ctor_set(x_2, 1, x_1); @@ -356,7 +1509,7 @@ lean_ctor_set(x_2, 5, x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_classical___at_Lean_Elab_Tactic_evalClassical___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_EXPORT lean_object* l_Lean_Elab_Tactic_classical___at_Lean_Elab_Tactic_evalClassical___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { lean_object* x_11; lean_object* x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; @@ -375,7 +1528,7 @@ x_29 = lean_ctor_get(x_25, 4); lean_dec(x_29); x_30 = l_Lean_Elab_Tactic_classical___rarg___lambda__3___closed__1; x_31 = l_Lean_ScopedEnvExtension_pushScope___rarg(x_30, x_28); -x_32 = l_Lean_Elab_Tactic_classical___at_Lean_Elab_Tactic_evalClassical___spec__2___closed__3; +x_32 = l_Lean_Elab_Tactic_classical___at_Lean_Elab_Tactic_evalClassical___spec__4___closed__3; lean_ctor_set(x_25, 4, x_32); lean_ctor_set(x_25, 0, x_31); x_33 = lean_st_ref_set(x_9, x_25, x_26); @@ -394,7 +1547,7 @@ if (x_38 == 0) lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; uint8_t x_44; lean_object* x_45; lean_object* x_46; x_39 = lean_ctor_get(x_36, 1); lean_dec(x_39); -x_40 = l_Lean_Elab_Tactic_classical___at_Lean_Elab_Tactic_evalClassical___spec__2___closed__4; +x_40 = l_Lean_Elab_Tactic_classical___at_Lean_Elab_Tactic_evalClassical___spec__4___closed__4; lean_ctor_set(x_36, 1, x_40); x_41 = lean_st_ref_set(x_7, x_36, x_37); x_42 = lean_ctor_get(x_41, 1); @@ -953,7 +2106,7 @@ lean_inc(x_185); lean_inc(x_184); lean_inc(x_183); lean_dec(x_36); -x_187 = l_Lean_Elab_Tactic_classical___at_Lean_Elab_Tactic_evalClassical___spec__2___closed__4; +x_187 = l_Lean_Elab_Tactic_classical___at_Lean_Elab_Tactic_evalClassical___spec__4___closed__4; x_188 = lean_alloc_ctor(0, 5, 0); lean_ctor_set(x_188, 0, x_183); lean_ctor_set(x_188, 1, x_187); @@ -1289,7 +2442,7 @@ lean_inc(x_264); lean_dec(x_25); x_271 = l_Lean_Elab_Tactic_classical___rarg___lambda__3___closed__1; x_272 = l_Lean_ScopedEnvExtension_pushScope___rarg(x_271, x_264); -x_273 = l_Lean_Elab_Tactic_classical___at_Lean_Elab_Tactic_evalClassical___spec__2___closed__3; +x_273 = l_Lean_Elab_Tactic_classical___at_Lean_Elab_Tactic_evalClassical___spec__4___closed__3; x_274 = lean_alloc_ctor(0, 8, 0); lean_ctor_set(x_274, 0, x_272); lean_ctor_set(x_274, 1, x_265); @@ -1328,7 +2481,7 @@ if (lean_is_exclusive(x_278)) { lean_dec_ref(x_278); x_284 = lean_box(0); } -x_285 = l_Lean_Elab_Tactic_classical___at_Lean_Elab_Tactic_evalClassical___spec__2___closed__4; +x_285 = l_Lean_Elab_Tactic_classical___at_Lean_Elab_Tactic_evalClassical___spec__4___closed__4; if (lean_is_scalar(x_284)) { x_286 = lean_alloc_ctor(0, 5, 0); } else { @@ -1709,140 +2862,78 @@ static lean_object* _init_l_Lean_Elab_Tactic_evalClassical___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean", 4, 4); +x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_evalTactic), 10, 0); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalClassical___closed__2() { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalClassical(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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_unsigned_to_nat(1u); +x_12 = l_Lean_Elab_Tactic_evalClassical___closed__1; +x_13 = lean_alloc_closure((void*)(l_Lean_Elab_Term_withNarrowedArgTacticReuse___at_Lean_Elab_Tactic_evalClassical___spec__1), 12, 3); +lean_closure_set(x_13, 0, x_11); +lean_closure_set(x_13, 1, x_12); +lean_closure_set(x_13, 2, x_1); +x_14 = l_Lean_Elab_Tactic_classical___at_Lean_Elab_Tactic_evalClassical___spec__4(x_13, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Term_withNarrowedArgTacticReuse___at_Lean_Elab_Tactic_evalClassical___spec__1___lambda__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Parser", 6, 6); -return x_1; +lean_object* x_3; +x_3 = l_Lean_Elab_Term_withNarrowedArgTacticReuse___at_Lean_Elab_Tactic_evalClassical___spec__1___lambda__1(x_1, x_2); +lean_dec(x_2); +return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalClassical___closed__3() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalClassical__1___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("Tactic", 6, 6); +x_1 = lean_mk_string_unchecked("Lean", 4, 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalClassical___closed__4() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalClassical__1___closed__2() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("classical", 9, 9); +x_1 = lean_mk_string_unchecked("Parser", 6, 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalClassical___closed__5() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalClassical__1___closed__3() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l_Lean_Elab_Tactic_evalClassical___closed__1; -x_2 = l_Lean_Elab_Tactic_evalClassical___closed__2; -x_3 = l_Lean_Elab_Tactic_evalClassical___closed__3; -x_4 = l_Lean_Elab_Tactic_evalClassical___closed__4; -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_unchecked("Tactic", 6, 6); +return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalClassical___closed__6() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalClassical__1___closed__4() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("tacticSeq", 9, 9); +x_1 = lean_mk_string_unchecked("classical", 9, 9); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalClassical___closed__7() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalClassical__1___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l_Lean_Elab_Tactic_evalClassical___closed__1; -x_2 = l_Lean_Elab_Tactic_evalClassical___closed__2; -x_3 = l_Lean_Elab_Tactic_evalClassical___closed__3; -x_4 = l_Lean_Elab_Tactic_evalClassical___closed__6; +x_1 = l___regBuiltin_Lean_Elab_Tactic_evalClassical__1___closed__1; +x_2 = l___regBuiltin_Lean_Elab_Tactic_evalClassical__1___closed__2; +x_3 = l___regBuiltin_Lean_Elab_Tactic_evalClassical__1___closed__3; +x_4 = l___regBuiltin_Lean_Elab_Tactic_evalClassical__1___closed__4; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalClassical(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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 = l_Lean_Elab_Tactic_evalClassical___closed__5; -lean_inc(x_1); -x_12 = l_Lean_Syntax_isOfKind(x_1, x_11); -if (x_12 == 0) -{ -lean_object* x_13; -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_13 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalClassical___spec__1___rarg(x_10); -return x_13; -} -else -{ -lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; -x_14 = lean_unsigned_to_nat(1u); -x_15 = l_Lean_Syntax_getArg(x_1, x_14); -lean_dec(x_1); -x_16 = l_Lean_Elab_Tactic_evalClassical___closed__7; -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_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_18 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalClassical___spec__1___rarg(x_10); -return x_18; -} -else -{ -lean_object* x_19; lean_object* x_20; -x_19 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_evalTactic), 10, 1); -lean_closure_set(x_19, 0, x_15); -x_20 = l_Lean_Elab_Tactic_classical___at_Lean_Elab_Tactic_evalClassical___spec__2(x_19, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_20; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalClassical___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_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalClassical___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_9; -} -} -static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalClassical__1___closed__1() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalClassical__1___closed__6() { _start: { lean_object* x_1; @@ -1850,7 +2941,7 @@ x_1 = lean_mk_string_unchecked("Elab", 4, 4); return x_1; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalClassical__1___closed__2() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalClassical__1___closed__7() { _start: { lean_object* x_1; @@ -1858,19 +2949,19 @@ x_1 = lean_mk_string_unchecked("evalClassical", 13, 13); return x_1; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalClassical__1___closed__3() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalClassical__1___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l_Lean_Elab_Tactic_evalClassical___closed__1; -x_2 = l___regBuiltin_Lean_Elab_Tactic_evalClassical__1___closed__1; -x_3 = l_Lean_Elab_Tactic_evalClassical___closed__3; -x_4 = l___regBuiltin_Lean_Elab_Tactic_evalClassical__1___closed__2; +x_1 = l___regBuiltin_Lean_Elab_Tactic_evalClassical__1___closed__1; +x_2 = l___regBuiltin_Lean_Elab_Tactic_evalClassical__1___closed__6; +x_3 = l___regBuiltin_Lean_Elab_Tactic_evalClassical__1___closed__3; +x_4 = l___regBuiltin_Lean_Elab_Tactic_evalClassical__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___regBuiltin_Lean_Elab_Tactic_evalClassical__1___closed__4() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalClassical__1___closed__9() { _start: { lean_object* x_1; @@ -1878,7 +2969,7 @@ x_1 = l_Lean_Elab_Tactic_tacticElabAttribute; return x_1; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalClassical__1___closed__5() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalClassical__1___closed__10() { _start: { lean_object* x_1; @@ -1890,14 +2981,23 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_evalClassical__1(lean_o _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_2 = l___regBuiltin_Lean_Elab_Tactic_evalClassical__1___closed__4; -x_3 = l_Lean_Elab_Tactic_evalClassical___closed__5; -x_4 = l___regBuiltin_Lean_Elab_Tactic_evalClassical__1___closed__3; -x_5 = l___regBuiltin_Lean_Elab_Tactic_evalClassical__1___closed__5; +x_2 = l___regBuiltin_Lean_Elab_Tactic_evalClassical__1___closed__9; +x_3 = l___regBuiltin_Lean_Elab_Tactic_evalClassical__1___closed__5; +x_4 = l___regBuiltin_Lean_Elab_Tactic_evalClassical__1___closed__8; +x_5 = l___regBuiltin_Lean_Elab_Tactic_evalClassical__1___closed__10; 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___regBuiltin_Lean_Elab_Tactic_evalClassical__2(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l___regBuiltin_Lean_Elab_Tactic_evalClassical__1___closed__8; +x_3 = l_Lean_Elab_addBuiltinIncrementalElab(x_2, x_1); +return x_3; +} +} lean_object* initialize_Lean_Elab_Tactic_Basic(uint8_t builtin, lean_object*); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Lean_Elab_Tactic_Classical(uint8_t builtin, lean_object* w) { @@ -1923,32 +3023,20 @@ l_Lean_Elab_Tactic_classical___rarg___lambda__4___closed__4 = _init_l_Lean_Elab_ lean_mark_persistent(l_Lean_Elab_Tactic_classical___rarg___lambda__4___closed__4); l_Lean_Elab_Tactic_classical___rarg___closed__1 = _init_l_Lean_Elab_Tactic_classical___rarg___closed__1(); lean_mark_persistent(l_Lean_Elab_Tactic_classical___rarg___closed__1); -l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalClassical___spec__1___rarg___closed__1 = _init_l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalClassical___spec__1___rarg___closed__1(); -lean_mark_persistent(l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalClassical___spec__1___rarg___closed__1); -l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalClassical___spec__1___rarg___closed__2 = _init_l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalClassical___spec__1___rarg___closed__2(); -lean_mark_persistent(l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalClassical___spec__1___rarg___closed__2); -l_Lean_Elab_Tactic_classical___at_Lean_Elab_Tactic_evalClassical___spec__2___closed__1 = _init_l_Lean_Elab_Tactic_classical___at_Lean_Elab_Tactic_evalClassical___spec__2___closed__1(); -lean_mark_persistent(l_Lean_Elab_Tactic_classical___at_Lean_Elab_Tactic_evalClassical___spec__2___closed__1); -l_Lean_Elab_Tactic_classical___at_Lean_Elab_Tactic_evalClassical___spec__2___closed__2 = _init_l_Lean_Elab_Tactic_classical___at_Lean_Elab_Tactic_evalClassical___spec__2___closed__2(); -lean_mark_persistent(l_Lean_Elab_Tactic_classical___at_Lean_Elab_Tactic_evalClassical___spec__2___closed__2); -l_Lean_Elab_Tactic_classical___at_Lean_Elab_Tactic_evalClassical___spec__2___closed__3 = _init_l_Lean_Elab_Tactic_classical___at_Lean_Elab_Tactic_evalClassical___spec__2___closed__3(); -lean_mark_persistent(l_Lean_Elab_Tactic_classical___at_Lean_Elab_Tactic_evalClassical___spec__2___closed__3); -l_Lean_Elab_Tactic_classical___at_Lean_Elab_Tactic_evalClassical___spec__2___closed__4 = _init_l_Lean_Elab_Tactic_classical___at_Lean_Elab_Tactic_evalClassical___spec__2___closed__4(); -lean_mark_persistent(l_Lean_Elab_Tactic_classical___at_Lean_Elab_Tactic_evalClassical___spec__2___closed__4); +l_Lean_Elab_Term_withNarrowedArgTacticReuse___at_Lean_Elab_Tactic_evalClassical___spec__1___lambda__1___closed__1 = _init_l_Lean_Elab_Term_withNarrowedArgTacticReuse___at_Lean_Elab_Tactic_evalClassical___spec__1___lambda__1___closed__1(); +lean_mark_persistent(l_Lean_Elab_Term_withNarrowedArgTacticReuse___at_Lean_Elab_Tactic_evalClassical___spec__1___lambda__1___closed__1); +l_Lean_Elab_Term_withNarrowedArgTacticReuse___at_Lean_Elab_Tactic_evalClassical___spec__1___lambda__1___closed__2 = _init_l_Lean_Elab_Term_withNarrowedArgTacticReuse___at_Lean_Elab_Tactic_evalClassical___spec__1___lambda__1___closed__2(); +lean_mark_persistent(l_Lean_Elab_Term_withNarrowedArgTacticReuse___at_Lean_Elab_Tactic_evalClassical___spec__1___lambda__1___closed__2); +l_Lean_Elab_Tactic_classical___at_Lean_Elab_Tactic_evalClassical___spec__4___closed__1 = _init_l_Lean_Elab_Tactic_classical___at_Lean_Elab_Tactic_evalClassical___spec__4___closed__1(); +lean_mark_persistent(l_Lean_Elab_Tactic_classical___at_Lean_Elab_Tactic_evalClassical___spec__4___closed__1); +l_Lean_Elab_Tactic_classical___at_Lean_Elab_Tactic_evalClassical___spec__4___closed__2 = _init_l_Lean_Elab_Tactic_classical___at_Lean_Elab_Tactic_evalClassical___spec__4___closed__2(); +lean_mark_persistent(l_Lean_Elab_Tactic_classical___at_Lean_Elab_Tactic_evalClassical___spec__4___closed__2); +l_Lean_Elab_Tactic_classical___at_Lean_Elab_Tactic_evalClassical___spec__4___closed__3 = _init_l_Lean_Elab_Tactic_classical___at_Lean_Elab_Tactic_evalClassical___spec__4___closed__3(); +lean_mark_persistent(l_Lean_Elab_Tactic_classical___at_Lean_Elab_Tactic_evalClassical___spec__4___closed__3); +l_Lean_Elab_Tactic_classical___at_Lean_Elab_Tactic_evalClassical___spec__4___closed__4 = _init_l_Lean_Elab_Tactic_classical___at_Lean_Elab_Tactic_evalClassical___spec__4___closed__4(); +lean_mark_persistent(l_Lean_Elab_Tactic_classical___at_Lean_Elab_Tactic_evalClassical___spec__4___closed__4); l_Lean_Elab_Tactic_evalClassical___closed__1 = _init_l_Lean_Elab_Tactic_evalClassical___closed__1(); lean_mark_persistent(l_Lean_Elab_Tactic_evalClassical___closed__1); -l_Lean_Elab_Tactic_evalClassical___closed__2 = _init_l_Lean_Elab_Tactic_evalClassical___closed__2(); -lean_mark_persistent(l_Lean_Elab_Tactic_evalClassical___closed__2); -l_Lean_Elab_Tactic_evalClassical___closed__3 = _init_l_Lean_Elab_Tactic_evalClassical___closed__3(); -lean_mark_persistent(l_Lean_Elab_Tactic_evalClassical___closed__3); -l_Lean_Elab_Tactic_evalClassical___closed__4 = _init_l_Lean_Elab_Tactic_evalClassical___closed__4(); -lean_mark_persistent(l_Lean_Elab_Tactic_evalClassical___closed__4); -l_Lean_Elab_Tactic_evalClassical___closed__5 = _init_l_Lean_Elab_Tactic_evalClassical___closed__5(); -lean_mark_persistent(l_Lean_Elab_Tactic_evalClassical___closed__5); -l_Lean_Elab_Tactic_evalClassical___closed__6 = _init_l_Lean_Elab_Tactic_evalClassical___closed__6(); -lean_mark_persistent(l_Lean_Elab_Tactic_evalClassical___closed__6); -l_Lean_Elab_Tactic_evalClassical___closed__7 = _init_l_Lean_Elab_Tactic_evalClassical___closed__7(); -lean_mark_persistent(l_Lean_Elab_Tactic_evalClassical___closed__7); l___regBuiltin_Lean_Elab_Tactic_evalClassical__1___closed__1 = _init_l___regBuiltin_Lean_Elab_Tactic_evalClassical__1___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Elab_Tactic_evalClassical__1___closed__1); l___regBuiltin_Lean_Elab_Tactic_evalClassical__1___closed__2 = _init_l___regBuiltin_Lean_Elab_Tactic_evalClassical__1___closed__2(); @@ -1959,9 +3047,22 @@ l___regBuiltin_Lean_Elab_Tactic_evalClassical__1___closed__4 = _init_l___regBuil lean_mark_persistent(l___regBuiltin_Lean_Elab_Tactic_evalClassical__1___closed__4); l___regBuiltin_Lean_Elab_Tactic_evalClassical__1___closed__5 = _init_l___regBuiltin_Lean_Elab_Tactic_evalClassical__1___closed__5(); lean_mark_persistent(l___regBuiltin_Lean_Elab_Tactic_evalClassical__1___closed__5); +l___regBuiltin_Lean_Elab_Tactic_evalClassical__1___closed__6 = _init_l___regBuiltin_Lean_Elab_Tactic_evalClassical__1___closed__6(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Tactic_evalClassical__1___closed__6); +l___regBuiltin_Lean_Elab_Tactic_evalClassical__1___closed__7 = _init_l___regBuiltin_Lean_Elab_Tactic_evalClassical__1___closed__7(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Tactic_evalClassical__1___closed__7); +l___regBuiltin_Lean_Elab_Tactic_evalClassical__1___closed__8 = _init_l___regBuiltin_Lean_Elab_Tactic_evalClassical__1___closed__8(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Tactic_evalClassical__1___closed__8); +l___regBuiltin_Lean_Elab_Tactic_evalClassical__1___closed__9 = _init_l___regBuiltin_Lean_Elab_Tactic_evalClassical__1___closed__9(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Tactic_evalClassical__1___closed__9); +l___regBuiltin_Lean_Elab_Tactic_evalClassical__1___closed__10 = _init_l___regBuiltin_Lean_Elab_Tactic_evalClassical__1___closed__10(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Tactic_evalClassical__1___closed__10); if (builtin) {res = l___regBuiltin_Lean_Elab_Tactic_evalClassical__1(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +}if (builtin) {res = l___regBuiltin_Lean_Elab_Tactic_evalClassical__2(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/Elab/Tactic/Conv/Simp.c b/stage0/stdlib/Lean/Elab/Tactic/Conv/Simp.c index a5c73f0753a9..a59a191776c0 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/Conv/Simp.c +++ b/stage0/stdlib/Lean/Elab/Tactic/Conv/Simp.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Lean.Elab.Tactic.Conv.Simp -// Imports: Lean.Elab.Tactic.Simp Lean.Elab.Tactic.Split Lean.Elab.Tactic.Conv.Basic +// Imports: Lean.Elab.Tactic.Simp Lean.Elab.Tactic.Split Lean.Elab.Tactic.Conv.Basic Lean.Elab.Tactic.SimpTrace #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -13,10 +13,14 @@ #ifdef __cplusplus extern "C" { #endif +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Conv_evalDSimpTrace___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_Tactic_Conv_evalSimpTrace___lambda__1___closed__9; lean_object* l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Simp_Result_getProof(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp__1___closed__5; +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Conv_evalSimpTrace___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* l_Lean_Meta_getSimpTheorems___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__2___closed__2; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp_declRange__1(lean_object*); static lean_object* l_Lean_Elab_Tactic_Conv_evalDSimp___lambda__1___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpMatch__1___closed__2; @@ -24,19 +28,27 @@ 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__1___closed__4; +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Conv_evalDSimpTrace___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Conv_evalSimp___lambda__1___closed__5; static lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp_declRange__1___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpMatch__1___closed__5; extern lean_object* l_Lean_Elab_Tactic_tacticElabAttribute; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp__1(lean_object*); +lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalExact___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpMatch_declRange__1___closed__1; +static lean_object* l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__10; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Conv_evalSimpMatch___boxed(lean_object*); lean_object* l_Lean_Meta_Split_simpMatch(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Syntax_getArgs(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimp_declRange__1___closed__3; +static lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimpTrace__1___closed__1; static lean_object* l_Lean_Elab_Tactic_Conv_evalSimp___lambda__1___closed__4; lean_object* l_Lean_addBuiltinDeclarationRanges(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_Tactic_TryThis_addSuggestion(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimp_declRange__1___closed__2; +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___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimpTrace__1___closed__2; 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*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp_declRange__1___closed__6; static lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp__1___closed__9; @@ -46,40 +58,80 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Conv_evalDSimp(lean_object*, lean_ob static lean_object* l_Lean_Elab_Tactic_Conv_evalSimp___lambda__2___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimp__1___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpMatch_declRange__1___closed__5; +static lean_object* l_Lean_Elab_Tactic_Conv_evalSimpTrace___closed__3; +static lean_object* l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__3; +static lean_object* l_Lean_Elab_Tactic_Conv_evalDSimpTrace___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp__1___closed__11; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Conv_evalSimp___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_Tactic_Conv_evalSimpTrace___lambda__1___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp__1___closed__10; +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_SourceInfo_fromRef(lean_object*, uint8_t); static lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpMatch_declRange__1___closed__4; static lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimp__1___closed__3; +lean_object* l_Lean_Syntax_node6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__8; lean_object* l_Lean_Elab_Tactic_Conv_changeLhs(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___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpMatch__1___closed__4; +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Conv_evalDSimpTrace(lean_object*, lean_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_evalDSimpTrace___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* 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___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpTrace__1___closed__1; +static lean_object* l_Lean_Elab_Tactic_Conv_evalSimpTrace___closed__4; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimp_declRange__1(lean_object*); lean_object* l_Lean_Elab_Tactic_withMainContext___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimp__1(lean_object*); +lean_object* lean_st_ref_get(lean_object*, lean_object*); +static lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpTrace__1___closed__3; +lean_object* l_Lean_Syntax_getOptional_x3f(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimp_declRange__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Conv_evalSimp___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___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpMatch_declRange__1___closed__6; static lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp__1___closed__6; static lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp_declRange__1___closed__7; +static lean_object* l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__11; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpMatch_declRange__1(lean_object*); static lean_object* l_Lean_Elab_Tactic_Conv_evalSimp___lambda__1___closed__7; +lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp__1___closed__4; +lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); +uint8_t l_Lean_Syntax_matchesNull(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Conv_evalSimp___lambda__1___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimp_declRange__1___closed__7; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Conv_evalSimp___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__6; static lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimp__1___closed__5; static lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp__1___closed__1; +lean_object* l_Array_append___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Conv_evalSimp___lambda__1___closed__1; +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimpTrace__1(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Conv_evalSimpTrace___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_Elab_Tactic_Conv_evalSimpTrace___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp__1___closed__8; +static lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpTrace__1___closed__2; +static lean_object* l_Lean_Elab_Tactic_Conv_evalDSimpTrace___lambda__2___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp_declRange__1___closed__3; +static lean_object* l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__4; +lean_object* l_Array_mkArray3___rarg(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimp__1___closed__1; +static lean_object* l_Lean_Elab_Tactic_Conv_evalDSimpTrace___closed__2; +static lean_object* l_Lean_Elab_Tactic_Conv_evalSimpTrace___closed__2; lean_object* l_Lean_Elab_Tactic_mkSimpContext(lean_object*, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Conv_evalSimpMatch___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Conv_evalSimp(lean_object*, 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_evalDSimpTrace___lambda__1___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimp_declRange__1___closed__5; +lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalExact___spec__1___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Conv_evalSimpMatch(lean_object*); +static lean_object* l_Lean_Elab_Tactic_Conv_evalDSimpTrace___lambda__2___closed__1; +static lean_object* l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__2___closed__1; +static lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimpTrace__1___closed__3; +uint8_t l_Lean_Syntax_isNone(lean_object*); +lean_object* l_Array_mkArray1___rarg(lean_object*); +static lean_object* l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__5; static lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp_declRange__1___closed__5; static lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp__1___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpMatch__1___closed__3; @@ -90,17 +142,24 @@ lean_object* lean_array_mk(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimp__1___closed__4; static lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp__1___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpMatch_declRange__1___closed__7; +lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Conv_evalSimpTrace(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_evalSimpMatch__1___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp_declRange__1___closed__2; +static lean_object* l_Lean_Elab_Tactic_Conv_evalSimpTrace___closed__5; static lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimp_declRange__1___closed__4; LEAN_EXPORT 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_Elab_Tactic_Conv_evalDSimp___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_mkSimpCallStx(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpTrace__1(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Conv_evalDSimpTrace___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___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpMatch_declRange__1___closed__3; 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_declRange__1___closed__6; static lean_object* l_Lean_Elab_Tactic_Conv_evalSimp___lambda__1___closed__6; static lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpMatch_declRange__1___closed__2; 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*); +static lean_object* l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__7; 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: { @@ -689,42 +748,1445 @@ x_4 = l_Lean_addBuiltinDeclarationRanges(x_2, x_3, x_1); return x_4; } } +static lean_object* _init_l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("tactic", 6, 6); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_Tactic_Conv_evalSimpTrace___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_Tactic_Conv_evalSimpTrace___lambda__1___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Try this: ", 10, 10); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp__1___closed__1; +x_2 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp__1___closed__2; +x_3 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp__1___closed__3; +x_4 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp__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_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("null", 4, 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__5; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = lean_mk_empty_array_with_capacity(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__7; +x_2 = l_Array_append___rarg(x_1, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__9() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("[", 1, 1); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__10() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("]", 1, 1); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__11() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("only", 4, 4); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Conv_evalSimpTrace___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_object* x_78; uint8_t x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; +x_16 = l_Lean_Syntax_getOptional_x3f(x_1); +x_78 = lean_ctor_get(x_13, 5); +lean_inc(x_78); +x_79 = 0; +x_80 = l_Lean_SourceInfo_fromRef(x_78, x_79); +lean_dec(x_78); +x_81 = lean_st_ref_get(x_14, x_15); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_191; +x_191 = lean_box(0); +x_82 = x_191; +goto block_190; +} +else +{ +uint8_t x_192; +x_192 = !lean_is_exclusive(x_16); +if (x_192 == 0) +{ +x_82 = x_16; +goto block_190; +} +else +{ +lean_object* x_193; lean_object* x_194; +x_193 = lean_ctor_get(x_16, 0); +lean_inc(x_193); +lean_dec(x_16); +x_194 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_194, 0, x_193); +x_82 = x_194; +goto block_190; +} +} +block_77: +{ +uint8_t x_19; uint8_t x_20; lean_object* x_21; lean_object* x_22; +x_19 = 0; +x_20 = 0; +x_21 = l_Lean_Elab_Tactic_Conv_evalSimp___lambda__2___closed__1; +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_17); +x_22 = l_Lean_Elab_Tactic_mkSimpContext(x_17, x_19, x_20, x_19, x_21, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_18); +if (lean_obj_tag(x_22) == 0) +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +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); +x_26 = lean_ctor_get(x_23, 1); +lean_inc(x_26); +x_27 = lean_ctor_get(x_23, 2); +lean_inc(x_27); +lean_dec(x_23); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +x_28 = l_Lean_Elab_Tactic_Conv_getLhs(x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_24); +if (lean_obj_tag(x_28) == 0) +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_29 = lean_ctor_get(x_28, 0); +lean_inc(x_29); +x_30 = lean_ctor_get(x_28, 1); +lean_inc(x_30); +lean_dec(x_28); +x_31 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Conv_evalSimp___lambda__1___boxed), 13, 3); +lean_closure_set(x_31, 0, x_29); +lean_closure_set(x_31, 1, x_25); +lean_closure_set(x_31, 2, x_26); +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); +x_32 = l_Lean_Elab_Tactic_Simp_DischargeWrapper_with___rarg(x_27, x_31, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_30); +lean_dec(x_27); +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; +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, 0); +lean_inc(x_35); +x_36 = lean_ctor_get(x_33, 1); +lean_inc(x_36); +lean_dec(x_33); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +x_37 = l_Lean_Elab_Tactic_Conv_applySimpResult(x_35, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_34); +if (lean_obj_tag(x_37) == 0) +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; uint8_t x_41; +x_38 = lean_ctor_get(x_37, 1); +lean_inc(x_38); +lean_dec(x_37); +x_39 = lean_ctor_get(x_36, 0); +lean_inc(x_39); +lean_dec(x_36); +lean_inc(x_13); +lean_inc(x_11); +x_40 = l_Lean_Elab_Tactic_mkSimpCallStx(x_17, x_39, x_11, x_12, x_13, x_14, x_38); +lean_dec(x_39); +x_41 = !lean_is_exclusive(x_40); +if (x_41 == 0) +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_42 = lean_ctor_get(x_40, 0); +x_43 = lean_ctor_get(x_40, 1); +x_44 = lean_ctor_get(x_13, 5); +lean_inc(x_44); +x_45 = l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__2; +lean_ctor_set(x_40, 1, x_42); +lean_ctor_set(x_40, 0, x_45); +x_46 = lean_box(0); +x_47 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_47, 0, x_40); +lean_ctor_set(x_47, 1, x_46); +lean_ctor_set(x_47, 2, x_46); +lean_ctor_set(x_47, 3, x_46); +lean_ctor_set(x_47, 4, x_46); +lean_ctor_set(x_47, 5, x_46); +x_48 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_48, 0, x_44); +x_49 = l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__3; +x_50 = l_Lean_Meta_Tactic_TryThis_addSuggestion(x_2, x_47, x_48, x_49, x_46, x_11, x_12, x_13, x_14, x_43); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(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; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_51 = lean_ctor_get(x_40, 0); +x_52 = lean_ctor_get(x_40, 1); +lean_inc(x_52); +lean_inc(x_51); +lean_dec(x_40); +x_53 = lean_ctor_get(x_13, 5); +lean_inc(x_53); +x_54 = l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__2; +x_55 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_55, 0, x_54); +lean_ctor_set(x_55, 1, x_51); +x_56 = lean_box(0); +x_57 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_57, 0, x_55); +lean_ctor_set(x_57, 1, x_56); +lean_ctor_set(x_57, 2, x_56); +lean_ctor_set(x_57, 3, x_56); +lean_ctor_set(x_57, 4, x_56); +lean_ctor_set(x_57, 5, x_56); +x_58 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_58, 0, x_53); +x_59 = l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__3; +x_60 = l_Lean_Meta_Tactic_TryThis_addSuggestion(x_2, x_57, x_58, x_59, x_56, x_11, x_12, x_13, x_14, x_52); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_58); +return x_60; +} +} +else +{ +uint8_t x_61; +lean_dec(x_36); +lean_dec(x_17); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +x_61 = !lean_is_exclusive(x_37); +if (x_61 == 0) +{ +return x_37; +} +else +{ +lean_object* x_62; lean_object* x_63; lean_object* x_64; +x_62 = lean_ctor_get(x_37, 0); +x_63 = lean_ctor_get(x_37, 1); +lean_inc(x_63); +lean_inc(x_62); +lean_dec(x_37); +x_64 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_64, 0, x_62); +lean_ctor_set(x_64, 1, x_63); +return x_64; +} +} +} +else +{ +uint8_t x_65; +lean_dec(x_17); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +x_65 = !lean_is_exclusive(x_32); +if (x_65 == 0) +{ +return x_32; +} +else +{ +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 +{ +uint8_t x_69; +lean_dec(x_27); +lean_dec(x_26); +lean_dec(x_25); +lean_dec(x_17); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +x_69 = !lean_is_exclusive(x_28); +if (x_69 == 0) +{ +return x_28; +} +else +{ +lean_object* x_70; lean_object* x_71; lean_object* x_72; +x_70 = lean_ctor_get(x_28, 0); +x_71 = lean_ctor_get(x_28, 1); +lean_inc(x_71); +lean_inc(x_70); +lean_dec(x_28); +x_72 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_72, 0, x_70); +lean_ctor_set(x_72, 1, x_71); +return x_72; +} +} +} +else +{ +uint8_t x_73; +lean_dec(x_17); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +x_73 = !lean_is_exclusive(x_22); +if (x_73 == 0) +{ +return x_22; +} +else +{ +lean_object* x_74; lean_object* x_75; lean_object* x_76; +x_74 = lean_ctor_get(x_22, 0); +x_75 = lean_ctor_get(x_22, 1); +lean_inc(x_75); +lean_inc(x_74); +lean_dec(x_22); +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_190: +{ +uint8_t x_83; +x_83 = !lean_is_exclusive(x_81); +if (x_83 == 0) +{ +lean_object* 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; lean_object* x_92; +x_84 = lean_ctor_get(x_81, 1); +x_85 = lean_ctor_get(x_81, 0); +lean_dec(x_85); +x_86 = 1; +x_87 = l_Lean_SourceInfo_fromRef(x_2, x_86); +x_88 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp__1___closed__5; +lean_ctor_set_tag(x_81, 2); +lean_ctor_set(x_81, 1, x_88); +lean_ctor_set(x_81, 0, x_87); +x_89 = l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__6; +x_90 = l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__7; +lean_inc(x_80); +x_91 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_91, 0, x_80); +lean_ctor_set(x_91, 1, x_89); +lean_ctor_set(x_91, 2, x_90); +if (lean_obj_tag(x_82) == 0) +{ +x_92 = x_90; +goto block_134; +} +else +{ +lean_object* x_135; lean_object* x_136; +x_135 = lean_ctor_get(x_82, 0); +lean_inc(x_135); +lean_dec(x_82); +x_136 = l_Array_mkArray1___rarg(x_135); +x_92 = x_136; +goto block_134; +} +block_134: +{ +lean_object* x_93; lean_object* x_94; +x_93 = l_Array_append___rarg(x_90, x_92); +lean_dec(x_92); +lean_inc(x_80); +x_94 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_94, 0, x_80); +lean_ctor_set(x_94, 1, x_89); +lean_ctor_set(x_94, 2, x_93); +if (lean_obj_tag(x_3) == 0) +{ +lean_object* x_95; lean_object* x_96; +x_95 = l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__8; +lean_inc(x_80); +x_96 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_96, 0, x_80); +lean_ctor_set(x_96, 1, x_89); +lean_ctor_set(x_96, 2, x_95); +if (lean_obj_tag(x_6) == 0) +{ +lean_object* x_97; lean_object* x_98; +x_97 = l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__4; +lean_inc(x_96); +x_98 = l_Lean_Syntax_node6(x_80, x_97, x_81, x_4, x_94, x_96, x_96, x_91); +x_17 = x_98; +x_18 = x_84; +goto block_77; +} +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; +x_99 = lean_ctor_get(x_6, 0); +x_100 = l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__9; +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_Array_append___rarg(x_90, x_99); +lean_inc(x_80); +x_103 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_103, 0, x_80); +lean_ctor_set(x_103, 1, x_89); +lean_ctor_set(x_103, 2, x_102); +x_104 = l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__10; +lean_inc(x_80); +x_105 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_105, 0, x_80); +lean_ctor_set(x_105, 1, x_104); +x_106 = l_Array_mkArray3___rarg(x_101, x_103, x_105); +x_107 = l_Array_append___rarg(x_90, x_106); +lean_dec(x_106); +lean_inc(x_80); +x_108 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_108, 0, x_80); +lean_ctor_set(x_108, 1, x_89); +lean_ctor_set(x_108, 2, x_107); +x_109 = l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__4; +x_110 = l_Lean_Syntax_node6(x_80, x_109, x_81, x_4, x_94, x_96, x_108, x_91); +x_17 = x_110; +x_18 = x_84; +goto block_77; +} +} +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; +x_111 = lean_ctor_get(x_3, 0); +x_112 = l_Lean_SourceInfo_fromRef(x_111, x_86); +x_113 = l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__11; +x_114 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_114, 0, x_112); +lean_ctor_set(x_114, 1, x_113); +x_115 = l_Array_mkArray1___rarg(x_114); +x_116 = l_Array_append___rarg(x_90, x_115); +lean_dec(x_115); +lean_inc(x_80); +x_117 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_117, 0, x_80); +lean_ctor_set(x_117, 1, x_89); +lean_ctor_set(x_117, 2, x_116); +if (lean_obj_tag(x_6) == 0) +{ +lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; +x_118 = l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__8; +lean_inc(x_80); +x_119 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_119, 0, x_80); +lean_ctor_set(x_119, 1, x_89); +lean_ctor_set(x_119, 2, x_118); +x_120 = l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__4; +x_121 = l_Lean_Syntax_node6(x_80, x_120, x_81, x_4, x_94, x_117, x_119, x_91); +x_17 = x_121; +x_18 = x_84; +goto block_77; +} +else +{ +lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_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_6, 0); +x_123 = l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__9; +lean_inc(x_80); +x_124 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_124, 0, x_80); +lean_ctor_set(x_124, 1, x_123); +x_125 = l_Array_append___rarg(x_90, x_122); +lean_inc(x_80); +x_126 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_126, 0, x_80); +lean_ctor_set(x_126, 1, x_89); +lean_ctor_set(x_126, 2, x_125); +x_127 = l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__10; +lean_inc(x_80); +x_128 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_128, 0, x_80); +lean_ctor_set(x_128, 1, x_127); +x_129 = l_Array_mkArray3___rarg(x_124, x_126, x_128); +x_130 = l_Array_append___rarg(x_90, x_129); +lean_dec(x_129); +lean_inc(x_80); +x_131 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_131, 0, x_80); +lean_ctor_set(x_131, 1, x_89); +lean_ctor_set(x_131, 2, x_130); +x_132 = l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__4; +x_133 = l_Lean_Syntax_node6(x_80, x_132, x_81, x_4, x_94, x_117, x_131, x_91); +x_17 = x_133; +x_18 = x_84; +goto block_77; +} +} +} +} +else +{ +lean_object* x_137; uint8_t 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_137 = lean_ctor_get(x_81, 1); +lean_inc(x_137); +lean_dec(x_81); +x_138 = 1; +x_139 = l_Lean_SourceInfo_fromRef(x_2, x_138); +x_140 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp__1___closed__5; +x_141 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_141, 0, x_139); +lean_ctor_set(x_141, 1, x_140); +x_142 = l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__6; +x_143 = l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__7; +lean_inc(x_80); +x_144 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_144, 0, x_80); +lean_ctor_set(x_144, 1, x_142); +lean_ctor_set(x_144, 2, x_143); +if (lean_obj_tag(x_82) == 0) +{ +x_145 = x_143; +goto block_187; +} +else +{ +lean_object* x_188; lean_object* x_189; +x_188 = lean_ctor_get(x_82, 0); +lean_inc(x_188); +lean_dec(x_82); +x_189 = l_Array_mkArray1___rarg(x_188); +x_145 = x_189; +goto block_187; +} +block_187: +{ +lean_object* x_146; lean_object* x_147; +x_146 = l_Array_append___rarg(x_143, x_145); +lean_dec(x_145); +lean_inc(x_80); +x_147 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_147, 0, x_80); +lean_ctor_set(x_147, 1, x_142); +lean_ctor_set(x_147, 2, x_146); +if (lean_obj_tag(x_3) == 0) +{ +lean_object* x_148; lean_object* x_149; +x_148 = l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__8; +lean_inc(x_80); +x_149 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_149, 0, x_80); +lean_ctor_set(x_149, 1, x_142); +lean_ctor_set(x_149, 2, x_148); +if (lean_obj_tag(x_6) == 0) +{ +lean_object* x_150; lean_object* x_151; +x_150 = l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__4; +lean_inc(x_149); +x_151 = l_Lean_Syntax_node6(x_80, x_150, x_141, x_4, x_147, x_149, x_149, x_144); +x_17 = x_151; +x_18 = x_137; +goto block_77; +} +else +{ +lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; +x_152 = lean_ctor_get(x_6, 0); +x_153 = l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__9; +lean_inc(x_80); +x_154 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_154, 0, x_80); +lean_ctor_set(x_154, 1, x_153); +x_155 = l_Array_append___rarg(x_143, x_152); +lean_inc(x_80); +x_156 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_156, 0, x_80); +lean_ctor_set(x_156, 1, x_142); +lean_ctor_set(x_156, 2, x_155); +x_157 = l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__10; +lean_inc(x_80); +x_158 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_158, 0, x_80); +lean_ctor_set(x_158, 1, x_157); +x_159 = l_Array_mkArray3___rarg(x_154, x_156, x_158); +x_160 = l_Array_append___rarg(x_143, x_159); +lean_dec(x_159); +lean_inc(x_80); +x_161 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_161, 0, x_80); +lean_ctor_set(x_161, 1, x_142); +lean_ctor_set(x_161, 2, x_160); +x_162 = l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__4; +x_163 = l_Lean_Syntax_node6(x_80, x_162, x_141, x_4, x_147, x_149, x_161, x_144); +x_17 = x_163; +x_18 = x_137; +goto block_77; +} +} +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; +x_164 = lean_ctor_get(x_3, 0); +x_165 = l_Lean_SourceInfo_fromRef(x_164, x_138); +x_166 = l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__11; +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_Array_mkArray1___rarg(x_167); +x_169 = l_Array_append___rarg(x_143, x_168); +lean_dec(x_168); +lean_inc(x_80); +x_170 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_170, 0, x_80); +lean_ctor_set(x_170, 1, x_142); +lean_ctor_set(x_170, 2, x_169); +if (lean_obj_tag(x_6) == 0) +{ +lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; +x_171 = l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__8; +lean_inc(x_80); +x_172 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_172, 0, x_80); +lean_ctor_set(x_172, 1, x_142); +lean_ctor_set(x_172, 2, x_171); +x_173 = l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__4; +x_174 = l_Lean_Syntax_node6(x_80, x_173, x_141, x_4, x_147, x_170, x_172, x_144); +x_17 = x_174; +x_18 = x_137; +goto block_77; +} +else +{ +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; +x_175 = lean_ctor_get(x_6, 0); +x_176 = l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__9; +lean_inc(x_80); +x_177 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_177, 0, x_80); +lean_ctor_set(x_177, 1, x_176); +x_178 = l_Array_append___rarg(x_143, x_175); +lean_inc(x_80); +x_179 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_179, 0, x_80); +lean_ctor_set(x_179, 1, x_142); +lean_ctor_set(x_179, 2, x_178); +x_180 = l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__10; +lean_inc(x_80); +x_181 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_181, 0, x_80); +lean_ctor_set(x_181, 1, x_180); +x_182 = l_Array_mkArray3___rarg(x_177, x_179, x_181); +x_183 = l_Array_append___rarg(x_143, x_182); +lean_dec(x_182); +lean_inc(x_80); +x_184 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_184, 0, x_80); +lean_ctor_set(x_184, 1, x_142); +lean_ctor_set(x_184, 2, x_183); +x_185 = l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__4; +x_186 = l_Lean_Syntax_node6(x_80, x_185, x_141, x_4, x_147, x_170, x_184, x_144); +x_17 = x_186; +x_18 = x_137; +goto block_77; +} +} +} +} +} +} +} +static lean_object* _init_l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__2___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("simpArgs", 8, 8); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__2___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp__1___closed__1; +x_2 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp__1___closed__2; +x_3 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp__1___closed__3; +x_4 = l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__2___closed__1; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +_start: +{ +lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_16 = lean_unsigned_to_nat(4u); +x_17 = l_Lean_Syntax_getArg(x_1, x_16); +x_18 = l_Lean_Syntax_isNone(x_17); +if (x_18 == 0) +{ +lean_object* x_19; uint8_t x_20; +x_19 = lean_unsigned_to_nat(1u); +lean_inc(x_17); +x_20 = l_Lean_Syntax_matchesNull(x_17, x_19); +if (x_20 == 0) +{ +lean_object* x_21; +lean_dec(x_17); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_4); +x_21 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalExact___spec__1___rarg(x_15); +return x_21; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; +x_22 = lean_unsigned_to_nat(0u); +x_23 = l_Lean_Syntax_getArg(x_17, x_22); +lean_dec(x_17); +x_24 = l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__2___closed__2; +lean_inc(x_23); +x_25 = l_Lean_Syntax_isOfKind(x_23, x_24); +if (x_25 == 0) +{ +lean_object* x_26; +lean_dec(x_23); +lean_dec(x_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_4); +x_26 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalExact___spec__1___rarg(x_15); +return x_26; +} +else +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_27 = l_Lean_Syntax_getArg(x_23, x_19); +lean_dec(x_23); +x_28 = l_Lean_Syntax_getArgs(x_27); +lean_dec(x_27); +x_29 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_29, 0, x_28); +x_30 = lean_box(0); +x_31 = l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1(x_2, x_3, x_6, x_4, x_30, x_29, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +lean_dec(x_29); +return x_31; +} +} +} +else +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; +lean_dec(x_17); +x_32 = lean_box(0); +x_33 = lean_box(0); +x_34 = l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1(x_2, x_3, x_6, x_4, x_33, x_32, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +return x_34; +} +} +} +static lean_object* _init_l_Lean_Elab_Tactic_Conv_evalSimpTrace___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("simpTrace", 9, 9); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Tactic_Conv_evalSimpTrace___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___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp__1___closed__1; +x_2 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp__1___closed__2; +x_3 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp__1___closed__3; +x_4 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp__1___closed__4; +x_5 = l_Lean_Elab_Tactic_Conv_evalSimpTrace___closed__1; +x_6 = l_Lean_Name_mkStr5(x_1, x_2, x_3, x_4, x_5); +return x_6; +} +} +static lean_object* _init_l_Lean_Elab_Tactic_Conv_evalSimpTrace___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalExact___spec__1___boxed), 8, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Tactic_Conv_evalSimpTrace___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("optConfig", 9, 9); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Tactic_Conv_evalSimpTrace___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___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp__1___closed__1; +x_2 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp__1___closed__2; +x_3 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp__1___closed__3; +x_4 = l_Lean_Elab_Tactic_Conv_evalSimpTrace___closed__4; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Conv_evalSimpTrace(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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 = l_Lean_Elab_Tactic_Conv_evalSimpTrace___closed__2; +lean_inc(x_1); +x_12 = l_Lean_Syntax_isOfKind(x_1, x_11); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; +lean_dec(x_1); +x_13 = l_Lean_Elab_Tactic_Conv_evalSimpTrace___closed__3; +x_14 = l_Lean_Elab_Tactic_withMainContext___rarg(x_13, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_14; +} +else +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_15 = lean_unsigned_to_nat(0u); +x_16 = l_Lean_Syntax_getArg(x_1, x_15); +x_17 = lean_unsigned_to_nat(1u); +x_18 = l_Lean_Syntax_getArg(x_1, x_17); +x_19 = l_Lean_Elab_Tactic_Conv_evalSimpTrace___closed__5; +lean_inc(x_18); +x_20 = l_Lean_Syntax_isOfKind(x_18, x_19); +if (x_20 == 0) +{ +lean_object* x_21; lean_object* x_22; +lean_dec(x_18); +lean_dec(x_16); +lean_dec(x_1); +x_21 = l_Lean_Elab_Tactic_Conv_evalSimpTrace___closed__3; +x_22 = l_Lean_Elab_Tactic_withMainContext___rarg(x_21, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_22; +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; +x_23 = lean_unsigned_to_nat(2u); +x_24 = l_Lean_Syntax_getArg(x_1, x_23); +x_25 = lean_unsigned_to_nat(3u); +x_26 = l_Lean_Syntax_getArg(x_1, x_25); +x_27 = l_Lean_Syntax_isNone(x_26); +if (x_27 == 0) +{ +uint8_t x_28; +lean_inc(x_26); +x_28 = l_Lean_Syntax_matchesNull(x_26, x_17); +if (x_28 == 0) +{ +lean_object* x_29; lean_object* x_30; +lean_dec(x_26); +lean_dec(x_24); +lean_dec(x_18); +lean_dec(x_16); +lean_dec(x_1); +x_29 = l_Lean_Elab_Tactic_Conv_evalSimpTrace___closed__3; +x_30 = l_Lean_Elab_Tactic_withMainContext___rarg(x_29, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_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; +x_31 = l_Lean_Syntax_getArg(x_26, x_15); +lean_dec(x_26); +x_32 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_32, 0, x_31); +x_33 = lean_box(0); +x_34 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__2___boxed), 15, 6); +lean_closure_set(x_34, 0, x_1); +lean_closure_set(x_34, 1, x_24); +lean_closure_set(x_34, 2, x_16); +lean_closure_set(x_34, 3, x_18); +lean_closure_set(x_34, 4, x_33); +lean_closure_set(x_34, 5, x_32); +x_35 = l_Lean_Elab_Tactic_withMainContext___rarg(x_34, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_35; +} +} +else +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; +lean_dec(x_26); +x_36 = lean_box(0); +x_37 = lean_box(0); +x_38 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__2___boxed), 15, 6); +lean_closure_set(x_38, 0, x_1); +lean_closure_set(x_38, 1, x_24); +lean_closure_set(x_38, 2, x_16); +lean_closure_set(x_38, 3, x_18); +lean_closure_set(x_38, 4, x_37); +lean_closure_set(x_38, 5, x_36); +x_39 = l_Lean_Elab_Tactic_withMainContext___rarg(x_38, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_39; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +_start: +{ +lean_object* x_16; +x_16 = l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_16; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +_start: +{ +lean_object* x_16; +x_16 = l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_16; +} +} +static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpTrace__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("evalSimpTrace", 13, 13); +return x_1; +} +} +static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpTrace__1___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___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp__1___closed__1; +x_2 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp__1___closed__7; +x_3 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp__1___closed__3; +x_4 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp__1___closed__4; +x_5 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpTrace__1___closed__1; +x_6 = l_Lean_Name_mkStr5(x_1, x_2, x_3, x_4, x_5); +return x_6; +} +} +static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpTrace__1___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Conv_evalSimpTrace), 10, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpTrace__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; +x_2 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp__1___closed__10; +x_3 = l_Lean_Elab_Tactic_Conv_evalSimpTrace___closed__2; +x_4 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpTrace__1___closed__2; +x_5 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpTrace__1___closed__3; +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_Elab_Tactic_Conv_evalSimpMatch___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_10; +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +x_10 = l_Lean_Elab_Tactic_Conv_getLhs(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_10) == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; +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_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +x_13 = l_Lean_Meta_Split_simpMatch(x_11, 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; +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +x_16 = l_Lean_Elab_Tactic_Conv_applySimpResult(x_14, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_15); +return x_16; +} +else +{ +uint8_t x_17; +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_17 = !lean_is_exclusive(x_13); +if (x_17 == 0) +{ +return x_13; +} +else +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_18 = lean_ctor_get(x_13, 0); +x_19 = lean_ctor_get(x_13, 1); +lean_inc(x_19); +lean_inc(x_18); +lean_dec(x_13); +x_20 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_20, 0, x_18); +lean_ctor_set(x_20, 1, x_19); +return x_20; +} +} +} +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); +lean_dec(x_1); +x_21 = !lean_is_exclusive(x_10); +if (x_21 == 0) +{ +return x_10; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_22 = lean_ctor_get(x_10, 0); +x_23 = lean_ctor_get(x_10, 1); +lean_inc(x_23); +lean_inc(x_22); +lean_dec(x_10); +x_24 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_24, 0, x_22); +lean_ctor_set(x_24, 1, x_23); +return x_24; +} +} +} +} +static lean_object* _init_l_Lean_Elab_Tactic_Conv_evalSimpMatch___rarg___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Conv_evalSimpMatch___rarg___lambda__1), 9, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Conv_evalSimpMatch___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_Elab_Tactic_Conv_evalSimpMatch___rarg___closed__1; +x_11 = l_Lean_Elab_Tactic_withMainContext___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_EXPORT lean_object* l_Lean_Elab_Tactic_Conv_evalSimpMatch(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Conv_evalSimpMatch___rarg), 9, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Conv_evalSimpMatch___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Lean_Elab_Tactic_Conv_evalSimpMatch(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpMatch__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("simpMatch", 9, 9); +return x_1; +} +} +static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpMatch__1___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___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp__1___closed__1; +x_2 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp__1___closed__2; +x_3 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp__1___closed__3; +x_4 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp__1___closed__4; +x_5 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpMatch__1___closed__1; +x_6 = l_Lean_Name_mkStr5(x_1, x_2, x_3, x_4, x_5); +return x_6; +} +} +static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpMatch__1___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("evalSimpMatch", 13, 13); +return x_1; +} +} +static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpMatch__1___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp__1___closed__1; +x_2 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp__1___closed__7; +x_3 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp__1___closed__3; +x_4 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp__1___closed__4; +x_5 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpMatch__1___closed__3; +x_6 = l_Lean_Name_mkStr5(x_1, x_2, x_3, x_4, x_5); +return x_6; +} +} +static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpMatch__1___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Conv_evalSimpMatch___boxed), 1, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpMatch__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; +x_2 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp__1___closed__10; +x_3 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpMatch__1___closed__2; +x_4 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpMatch__1___closed__4; +x_5 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpMatch__1___closed__5; +x_6 = l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(x_2, x_3, x_4, x_5, x_1); +return x_6; +} +} +static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpMatch_declRange__1___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_unsigned_to_nat(26u); +x_2 = lean_unsigned_to_nat(52u); +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpMatch_declRange__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_unsigned_to_nat(27u); +x_2 = lean_unsigned_to_nat(48u); +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_Conv_evalSimpMatch_declRange__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___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpMatch_declRange__1___closed__1; +x_2 = lean_unsigned_to_nat(52u); +x_3 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpMatch_declRange__1___closed__2; +x_4 = lean_unsigned_to_nat(48u); +x_5 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_5, 0, x_1); +lean_ctor_set(x_5, 1, x_2); +lean_ctor_set(x_5, 2, x_3); +lean_ctor_set(x_5, 3, x_4); +return x_5; +} +} +static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpMatch_declRange__1___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_unsigned_to_nat(26u); +x_2 = lean_unsigned_to_nat(56u); +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_Conv_evalSimpMatch_declRange__1___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_unsigned_to_nat(26u); +x_2 = lean_unsigned_to_nat(69u); +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpMatch_declRange__1___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpMatch_declRange__1___closed__4; +x_2 = lean_unsigned_to_nat(56u); +x_3 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpMatch_declRange__1___closed__5; +x_4 = lean_unsigned_to_nat(69u); +x_5 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_5, 0, x_1); +lean_ctor_set(x_5, 1, x_2); +lean_ctor_set(x_5, 2, x_3); +lean_ctor_set(x_5, 3, x_4); +return x_5; +} +} +static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpMatch_declRange__1___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpMatch_declRange__1___closed__3; +x_2 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpMatch_declRange__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; +} +} +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpMatch_declRange__1(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_2 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpMatch__1___closed__4; +x_3 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpMatch_declRange__1___closed__7; +x_4 = l_Lean_addBuiltinDeclarationRanges(x_2, x_3, x_1); +return x_4; +} +} +static lean_object* _init_l_Lean_Elab_Tactic_Conv_evalDSimp___lambda__1___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(0); +x_2 = lean_array_mk(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Conv_evalDSimp___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: +{ +uint8_t x_11; uint8_t x_12; lean_object* x_13; lean_object* x_14; +x_11 = 0; +x_12 = 2; +x_13 = l_Lean_Elab_Tactic_Conv_evalSimp___lambda__2___closed__1; +lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -x_10 = l_Lean_Elab_Tactic_Conv_getLhs(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_10) == 0) +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +x_14 = l_Lean_Elab_Tactic_mkSimpContext(x_1, x_11, x_12, x_11, x_13, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_14) == 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); +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_ctor_get(x_15, 0); +lean_inc(x_17); +lean_dec(x_15); +lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -lean_inc(x_5); -x_13 = l_Lean_Meta_Split_simpMatch(x_11, x_5, x_6, x_7, x_8, x_12); -if (lean_obj_tag(x_13) == 0) +x_18 = l_Lean_Elab_Tactic_Conv_getLhs(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_16); +if (lean_obj_tag(x_18) == 0) { -lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_14 = lean_ctor_get(x_13, 0); -lean_inc(x_14); -x_15 = lean_ctor_get(x_13, 1); -lean_inc(x_15); -lean_dec(x_13); -x_16 = l_Lean_Elab_Tactic_Conv_applySimpResult(x_14, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_15); -return x_16; +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_inc(x_19); +x_20 = lean_ctor_get(x_18, 1); +lean_inc(x_20); +lean_dec(x_18); +x_21 = l_Lean_Elab_Tactic_Conv_evalDSimp___lambda__1___closed__1; +x_22 = l_Lean_Elab_Tactic_Conv_evalSimp___lambda__1___closed__8; +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +x_23 = l_Lean_Meta_dsimp(x_19, x_17, x_21, x_22, x_6, x_7, x_8, x_9, 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; +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +x_25 = lean_ctor_get(x_23, 1); +lean_inc(x_25); +lean_dec(x_23); +x_26 = lean_ctor_get(x_24, 0); +lean_inc(x_26); +lean_dec(x_24); +x_27 = l_Lean_Elab_Tactic_Conv_changeLhs(x_26, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_25); +return x_27; } else { -uint8_t x_17; +uint8_t x_28; +lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -732,30 +2194,31 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -x_17 = !lean_is_exclusive(x_13); -if (x_17 == 0) +x_28 = !lean_is_exclusive(x_23); +if (x_28 == 0) { -return x_13; +return x_23; } else { -lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_18 = lean_ctor_get(x_13, 0); -x_19 = lean_ctor_get(x_13, 1); -lean_inc(x_19); -lean_inc(x_18); -lean_dec(x_13); -x_20 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_20, 0, x_18); -lean_ctor_set(x_20, 1, x_19); -return x_20; +lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_29 = lean_ctor_get(x_23, 0); +x_30 = lean_ctor_get(x_23, 1); +lean_inc(x_30); +lean_inc(x_29); +lean_dec(x_23); +x_31 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_31, 0, x_29); +lean_ctor_set(x_31, 1, x_30); +return x_31; } } } else { -uint8_t x_21; +uint8_t x_32; +lean_dec(x_17); +lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -763,71 +2226,77 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -x_21 = !lean_is_exclusive(x_10); -if (x_21 == 0) +x_32 = !lean_is_exclusive(x_18); +if (x_32 == 0) { -return x_10; +return x_18; } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_10, 0); -x_23 = lean_ctor_get(x_10, 1); -lean_inc(x_23); -lean_inc(x_22); -lean_dec(x_10); -x_24 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_24, 0, x_22); -lean_ctor_set(x_24, 1, x_23); -return x_24; -} +lean_object* x_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_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; } } } -static lean_object* _init_l_Lean_Elab_Tactic_Conv_evalSimpMatch___rarg___closed__1() { -_start: +else { -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Conv_evalSimpMatch___rarg___lambda__1), 9, 0); -return x_1; -} +uint8_t x_36; +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_36 = !lean_is_exclusive(x_14); +if (x_36 == 0) +{ +return x_14; } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Conv_evalSimpMatch___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 = l_Lean_Elab_Tactic_Conv_evalSimpMatch___rarg___closed__1; -x_11 = l_Lean_Elab_Tactic_withMainContext___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_37; lean_object* x_38; lean_object* x_39; +x_37 = lean_ctor_get(x_14, 0); +x_38 = lean_ctor_get(x_14, 1); +lean_inc(x_38); +lean_inc(x_37); +lean_dec(x_14); +x_39 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_39, 0, x_37); +lean_ctor_set(x_39, 1, x_38); +return x_39; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Conv_evalSimpMatch(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Conv_evalSimpMatch___rarg), 9, 0); -return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Conv_evalSimpMatch___boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Conv_evalDSimp(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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 = l_Lean_Elab_Tactic_Conv_evalSimpMatch(x_1); -lean_dec(x_1); -return x_2; +{ +lean_object* x_11; lean_object* x_12; +x_11 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Conv_evalDSimp___lambda__1), 10, 1); +lean_closure_set(x_11, 0, x_1); +x_12 = l_Lean_Elab_Tactic_withMainContext___rarg(x_11, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_12; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpMatch__1___closed__1() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimp__1___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("simpMatch", 9, 9); +x_1 = lean_mk_string_unchecked("dsimp", 5, 5); return x_1; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpMatch__1___closed__2() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimp__1___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; @@ -835,20 +2304,20 @@ x_1 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp__1___closed__1; x_2 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp__1___closed__2; x_3 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp__1___closed__3; x_4 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp__1___closed__4; -x_5 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpMatch__1___closed__1; +x_5 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimp__1___closed__1; x_6 = l_Lean_Name_mkStr5(x_1, x_2, x_3, x_4, x_5); return x_6; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpMatch__1___closed__3() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimp__1___closed__3() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("evalSimpMatch", 13, 13); +x_1 = lean_mk_string_unchecked("evalDSimp", 9, 9); return x_1; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpMatch__1___closed__4() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimp__1___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; @@ -856,48 +2325,48 @@ x_1 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp__1___closed__1; x_2 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp__1___closed__7; x_3 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp__1___closed__3; x_4 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp__1___closed__4; -x_5 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpMatch__1___closed__3; +x_5 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimp__1___closed__3; x_6 = l_Lean_Name_mkStr5(x_1, x_2, x_3, x_4, x_5); return x_6; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpMatch__1___closed__5() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimp__1___closed__5() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Conv_evalSimpMatch___boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Conv_evalDSimp), 10, 0); return x_1; } } -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpMatch__1(lean_object* x_1) { +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimp__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; x_2 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp__1___closed__10; -x_3 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpMatch__1___closed__2; -x_4 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpMatch__1___closed__4; -x_5 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpMatch__1___closed__5; +x_3 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimp__1___closed__2; +x_4 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimp__1___closed__4; +x_5 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimp__1___closed__5; x_6 = l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(x_2, x_3, x_4, x_5, x_1); return x_6; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpMatch_declRange__1___closed__1() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimp_declRange__1___closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(26u); -x_2 = lean_unsigned_to_nat(52u); +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); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpMatch_declRange__1___closed__2() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimp_declRange__1___closed__2() { _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(31u); x_2 = lean_unsigned_to_nat(48u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -905,54 +2374,53 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpMatch_declRange__1___closed__3() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimp_declRange__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___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpMatch_declRange__1___closed__1; -x_2 = lean_unsigned_to_nat(52u); -x_3 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpMatch_declRange__1___closed__2; -x_4 = lean_unsigned_to_nat(48u); -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; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimp_declRange__1___closed__1; +x_2 = lean_unsigned_to_nat(48u); +x_3 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimp_declRange__1___closed__2; +x_4 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +lean_ctor_set(x_4, 3, x_2); +return x_4; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpMatch_declRange__1___closed__4() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimp_declRange__1___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(26u); -x_2 = lean_unsigned_to_nat(56u); +x_1 = lean_unsigned_to_nat(29u); +x_2 = lean_unsigned_to_nat(52u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpMatch_declRange__1___closed__5() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimp_declRange__1___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(26u); -x_2 = lean_unsigned_to_nat(69u); +x_1 = lean_unsigned_to_nat(29u); +x_2 = lean_unsigned_to_nat(61u); 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_Conv_evalSimpMatch_declRange__1___closed__6() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimp_declRange__1___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpMatch_declRange__1___closed__4; -x_2 = lean_unsigned_to_nat(56u); -x_3 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpMatch_declRange__1___closed__5; -x_4 = lean_unsigned_to_nat(69u); +x_1 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimp_declRange__1___closed__4; +x_2 = lean_unsigned_to_nat(52u); +x_3 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimp_declRange__1___closed__5; +x_4 = lean_unsigned_to_nat(61u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); lean_ctor_set(x_5, 1, x_2); @@ -961,369 +2429,892 @@ lean_ctor_set(x_5, 3, x_4); return x_5; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpMatch_declRange__1___closed__7() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimp_declRange__1___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpMatch_declRange__1___closed__3; -x_2 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpMatch_declRange__1___closed__6; +x_1 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimp_declRange__1___closed__3; +x_2 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimp_declRange__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; } } -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpMatch_declRange__1(lean_object* x_1) { +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimp_declRange__1(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_2 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpMatch__1___closed__4; -x_3 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpMatch_declRange__1___closed__7; +x_2 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimp__1___closed__4; +x_3 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimp_declRange__1___closed__7; x_4 = l_Lean_addBuiltinDeclarationRanges(x_2, x_3, x_1); return x_4; } } -static lean_object* _init_l_Lean_Elab_Tactic_Conv_evalDSimp___lambda__1___closed__1() { +static lean_object* _init_l_Lean_Elab_Tactic_Conv_evalDSimpTrace___lambda__1___closed__1() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = lean_box(0); -x_2 = lean_array_mk(x_1); -return x_2; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp__1___closed__1; +x_2 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp__1___closed__2; +x_3 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp__1___closed__3; +x_4 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimp__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_Elab_Tactic_Conv_evalDSimp___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_EXPORT lean_object* l_Lean_Elab_Tactic_Conv_evalDSimpTrace___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +lean_object* x_15; lean_object* x_16; lean_object* x_75; uint8_t x_76; lean_object* x_77; lean_object* x_78; uint8_t x_79; +x_75 = lean_ctor_get(x_12, 5); +lean_inc(x_75); +x_76 = 0; +x_77 = l_Lean_SourceInfo_fromRef(x_75, x_76); +lean_dec(x_75); +x_78 = lean_st_ref_get(x_13, x_14); +x_79 = !lean_is_exclusive(x_78); +if (x_79 == 0) +{ +lean_object* x_80; lean_object* x_81; uint8_t x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; +x_80 = lean_ctor_get(x_78, 1); +x_81 = lean_ctor_get(x_78, 0); +lean_dec(x_81); +x_82 = 1; +x_83 = l_Lean_SourceInfo_fromRef(x_1, x_82); +x_84 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimp__1___closed__1; +lean_ctor_set_tag(x_78, 2); +lean_ctor_set(x_78, 1, x_84); +lean_ctor_set(x_78, 0, x_83); +x_85 = l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__6; +x_86 = l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__7; +lean_inc(x_77); +x_87 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_87, 0, x_77); +lean_ctor_set(x_87, 1, x_85); +lean_ctor_set(x_87, 2, x_86); +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_88; lean_object* x_89; +x_88 = l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__8; +lean_inc(x_77); +x_89 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_89, 0, x_77); +lean_ctor_set(x_89, 1, x_85); +lean_ctor_set(x_89, 2, x_88); +if (lean_obj_tag(x_5) == 0) +{ +lean_object* x_90; lean_object* x_91; +x_90 = l_Lean_Elab_Tactic_Conv_evalDSimpTrace___lambda__1___closed__1; +lean_inc(x_89); +lean_inc(x_87); +x_91 = l_Lean_Syntax_node6(x_77, x_90, x_78, x_3, x_87, x_89, x_89, x_87); +x_15 = x_91; +x_16 = x_80; +goto block_74; +} +else { -uint8_t x_11; uint8_t x_12; lean_object* x_13; lean_object* x_14; -x_11 = 0; -x_12 = 2; -x_13 = l_Lean_Elab_Tactic_Conv_evalSimp___lambda__2___closed__1; -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -lean_inc(x_2); -x_14 = l_Lean_Elab_Tactic_mkSimpContext(x_1, x_11, x_12, x_11, x_13, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_14) == 0) +lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; +x_92 = lean_ctor_get(x_5, 0); +x_93 = l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__9; +lean_inc(x_77); +x_94 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_94, 0, x_77); +lean_ctor_set(x_94, 1, x_93); +x_95 = l_Array_append___rarg(x_86, x_92); +lean_inc(x_77); +x_96 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_96, 0, x_77); +lean_ctor_set(x_96, 1, x_85); +lean_ctor_set(x_96, 2, x_95); +x_97 = l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__10; +lean_inc(x_77); +x_98 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_98, 0, x_77); +lean_ctor_set(x_98, 1, x_97); +x_99 = l_Array_mkArray3___rarg(x_94, x_96, x_98); +x_100 = l_Array_append___rarg(x_86, x_99); +lean_dec(x_99); +lean_inc(x_77); +x_101 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_101, 0, x_77); +lean_ctor_set(x_101, 1, x_85); +lean_ctor_set(x_101, 2, x_100); +x_102 = l_Lean_Elab_Tactic_Conv_evalDSimpTrace___lambda__1___closed__1; +lean_inc(x_87); +x_103 = l_Lean_Syntax_node6(x_77, x_102, x_78, x_3, x_87, x_89, x_101, x_87); +x_15 = x_103; +x_16 = x_80; +goto block_74; +} +} +else { -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_ctor_get(x_15, 0); -lean_inc(x_17); -lean_dec(x_15); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -x_18 = l_Lean_Elab_Tactic_Conv_getLhs(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_16); -if (lean_obj_tag(x_18) == 0) +lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; +x_104 = lean_ctor_get(x_2, 0); +x_105 = l_Lean_SourceInfo_fromRef(x_104, x_82); +x_106 = l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__11; +x_107 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_107, 0, x_105); +lean_ctor_set(x_107, 1, x_106); +x_108 = l_Array_mkArray1___rarg(x_107); +x_109 = l_Array_append___rarg(x_86, x_108); +lean_dec(x_108); +lean_inc(x_77); +x_110 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_110, 0, x_77); +lean_ctor_set(x_110, 1, x_85); +lean_ctor_set(x_110, 2, x_109); +if (lean_obj_tag(x_5) == 0) +{ +lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; +x_111 = l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__8; +lean_inc(x_77); +x_112 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_112, 0, x_77); +lean_ctor_set(x_112, 1, x_85); +lean_ctor_set(x_112, 2, x_111); +x_113 = l_Lean_Elab_Tactic_Conv_evalDSimpTrace___lambda__1___closed__1; +lean_inc(x_87); +x_114 = l_Lean_Syntax_node6(x_77, x_113, x_78, x_3, x_87, x_110, x_112, x_87); +x_15 = x_114; +x_16 = x_80; +goto block_74; +} +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_18, 0); -lean_inc(x_19); -x_20 = lean_ctor_get(x_18, 1); -lean_inc(x_20); -lean_dec(x_18); -x_21 = l_Lean_Elab_Tactic_Conv_evalDSimp___lambda__1___closed__1; -x_22 = l_Lean_Elab_Tactic_Conv_evalSimp___lambda__1___closed__8; +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_115 = lean_ctor_get(x_5, 0); +x_116 = l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__9; +lean_inc(x_77); +x_117 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_117, 0, x_77); +lean_ctor_set(x_117, 1, x_116); +x_118 = l_Array_append___rarg(x_86, x_115); +lean_inc(x_77); +x_119 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_119, 0, x_77); +lean_ctor_set(x_119, 1, x_85); +lean_ctor_set(x_119, 2, x_118); +x_120 = l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__10; +lean_inc(x_77); +x_121 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_121, 0, x_77); +lean_ctor_set(x_121, 1, x_120); +x_122 = l_Array_mkArray3___rarg(x_117, x_119, x_121); +x_123 = l_Array_append___rarg(x_86, x_122); +lean_dec(x_122); +lean_inc(x_77); +x_124 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_124, 0, x_77); +lean_ctor_set(x_124, 1, x_85); +lean_ctor_set(x_124, 2, x_123); +x_125 = l_Lean_Elab_Tactic_Conv_evalDSimpTrace___lambda__1___closed__1; +lean_inc(x_87); +x_126 = l_Lean_Syntax_node6(x_77, x_125, x_78, x_3, x_87, x_110, x_124, x_87); +x_15 = x_126; +x_16 = x_80; +goto block_74; +} +} +} +else +{ +lean_object* x_127; uint8_t x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; +x_127 = lean_ctor_get(x_78, 1); +lean_inc(x_127); +lean_dec(x_78); +x_128 = 1; +x_129 = l_Lean_SourceInfo_fromRef(x_1, x_128); +x_130 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimp__1___closed__1; +x_131 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_131, 0, x_129); +lean_ctor_set(x_131, 1, x_130); +x_132 = l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__6; +x_133 = l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__7; +lean_inc(x_77); +x_134 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_134, 0, x_77); +lean_ctor_set(x_134, 1, x_132); +lean_ctor_set(x_134, 2, x_133); +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_135; lean_object* x_136; +x_135 = l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__8; +lean_inc(x_77); +x_136 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_136, 0, x_77); +lean_ctor_set(x_136, 1, x_132); +lean_ctor_set(x_136, 2, x_135); +if (lean_obj_tag(x_5) == 0) +{ +lean_object* x_137; lean_object* x_138; +x_137 = l_Lean_Elab_Tactic_Conv_evalDSimpTrace___lambda__1___closed__1; +lean_inc(x_136); +lean_inc(x_134); +x_138 = l_Lean_Syntax_node6(x_77, x_137, x_131, x_3, x_134, x_136, x_136, x_134); +x_15 = x_138; +x_16 = x_127; +goto block_74; +} +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; +x_139 = lean_ctor_get(x_5, 0); +x_140 = l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__9; +lean_inc(x_77); +x_141 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_141, 0, x_77); +lean_ctor_set(x_141, 1, x_140); +x_142 = l_Array_append___rarg(x_133, x_139); +lean_inc(x_77); +x_143 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_143, 0, x_77); +lean_ctor_set(x_143, 1, x_132); +lean_ctor_set(x_143, 2, x_142); +x_144 = l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__10; +lean_inc(x_77); +x_145 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_145, 0, x_77); +lean_ctor_set(x_145, 1, x_144); +x_146 = l_Array_mkArray3___rarg(x_141, x_143, x_145); +x_147 = l_Array_append___rarg(x_133, x_146); +lean_dec(x_146); +lean_inc(x_77); +x_148 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_148, 0, x_77); +lean_ctor_set(x_148, 1, x_132); +lean_ctor_set(x_148, 2, x_147); +x_149 = l_Lean_Elab_Tactic_Conv_evalDSimpTrace___lambda__1___closed__1; +lean_inc(x_134); +x_150 = l_Lean_Syntax_node6(x_77, x_149, x_131, x_3, x_134, x_136, x_148, x_134); +x_15 = x_150; +x_16 = x_127; +goto block_74; +} +} +else +{ +lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; +x_151 = lean_ctor_get(x_2, 0); +x_152 = l_Lean_SourceInfo_fromRef(x_151, x_128); +x_153 = l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__11; +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_Array_mkArray1___rarg(x_154); +x_156 = l_Array_append___rarg(x_133, x_155); +lean_dec(x_155); +lean_inc(x_77); +x_157 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_157, 0, x_77); +lean_ctor_set(x_157, 1, x_132); +lean_ctor_set(x_157, 2, x_156); +if (lean_obj_tag(x_5) == 0) +{ +lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; +x_158 = l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__8; +lean_inc(x_77); +x_159 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_159, 0, x_77); +lean_ctor_set(x_159, 1, x_132); +lean_ctor_set(x_159, 2, x_158); +x_160 = l_Lean_Elab_Tactic_Conv_evalDSimpTrace___lambda__1___closed__1; +lean_inc(x_134); +x_161 = l_Lean_Syntax_node6(x_77, x_160, x_131, x_3, x_134, x_157, x_159, x_134); +x_15 = x_161; +x_16 = x_127; +goto block_74; +} +else +{ +lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; +x_162 = lean_ctor_get(x_5, 0); +x_163 = l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__9; +lean_inc(x_77); +x_164 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_164, 0, x_77); +lean_ctor_set(x_164, 1, x_163); +x_165 = l_Array_append___rarg(x_133, x_162); +lean_inc(x_77); +x_166 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_166, 0, x_77); +lean_ctor_set(x_166, 1, x_132); +lean_ctor_set(x_166, 2, x_165); +x_167 = l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__10; +lean_inc(x_77); +x_168 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_168, 0, x_77); +lean_ctor_set(x_168, 1, x_167); +x_169 = l_Array_mkArray3___rarg(x_164, x_166, x_168); +x_170 = l_Array_append___rarg(x_133, x_169); +lean_dec(x_169); +lean_inc(x_77); +x_171 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_171, 0, x_77); +lean_ctor_set(x_171, 1, x_132); +lean_ctor_set(x_171, 2, x_170); +x_172 = l_Lean_Elab_Tactic_Conv_evalDSimpTrace___lambda__1___closed__1; +lean_inc(x_134); +x_173 = l_Lean_Syntax_node6(x_77, x_172, x_131, x_3, x_134, x_157, x_171, x_134); +x_15 = x_173; +x_16 = x_127; +goto block_74; +} +} +} +block_74: +{ +uint8_t x_17; uint8_t x_18; lean_object* x_19; lean_object* x_20; +x_17 = 0; +x_18 = 2; +x_19 = l_Lean_Elab_Tactic_Conv_evalSimp___lambda__2___closed__1; +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -x_23 = l_Lean_Meta_dsimp(x_19, x_17, x_21, x_22, x_6, x_7, x_8, x_9, x_20); -if (lean_obj_tag(x_23) == 0) +lean_inc(x_15); +x_20 = l_Lean_Elab_Tactic_mkSimpContext(x_15, x_17, x_18, x_17, x_19, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_16); +if (lean_obj_tag(x_20) == 0) { -lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -x_25 = lean_ctor_get(x_23, 1); +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_21 = lean_ctor_get(x_20, 0); +lean_inc(x_21); +x_22 = lean_ctor_get(x_20, 1); +lean_inc(x_22); +lean_dec(x_20); +x_23 = lean_ctor_get(x_21, 0); +lean_inc(x_23); +lean_dec(x_21); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +x_24 = l_Lean_Elab_Tactic_Conv_getLhs(x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_22); +if (lean_obj_tag(x_24) == 0) +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_25 = lean_ctor_get(x_24, 0); lean_inc(x_25); -lean_dec(x_23); -x_26 = lean_ctor_get(x_24, 0); +x_26 = lean_ctor_get(x_24, 1); lean_inc(x_26); lean_dec(x_24); -x_27 = l_Lean_Elab_Tactic_Conv_changeLhs(x_26, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_25); -return x_27; +x_27 = l_Lean_Elab_Tactic_Conv_evalDSimp___lambda__1___closed__1; +x_28 = l_Lean_Elab_Tactic_Conv_evalSimp___lambda__1___closed__8; +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +x_29 = l_Lean_Meta_dsimp(x_25, x_23, x_27, x_28, x_10, x_11, x_12, x_13, x_26); +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; +x_30 = lean_ctor_get(x_29, 0); +lean_inc(x_30); +x_31 = lean_ctor_get(x_29, 1); +lean_inc(x_31); +lean_dec(x_29); +x_32 = lean_ctor_get(x_30, 0); +lean_inc(x_32); +x_33 = lean_ctor_get(x_30, 1); +lean_inc(x_33); +lean_dec(x_30); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +x_34 = l_Lean_Elab_Tactic_Conv_changeLhs(x_32, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_31); +if (lean_obj_tag(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_34, 1); +lean_inc(x_35); +lean_dec(x_34); +x_36 = lean_ctor_get(x_33, 0); +lean_inc(x_36); +lean_dec(x_33); +lean_inc(x_12); +lean_inc(x_10); +x_37 = l_Lean_Elab_Tactic_mkSimpCallStx(x_15, x_36, x_10, x_11, x_12, x_13, x_35); +lean_dec(x_36); +x_38 = !lean_is_exclusive(x_37); +if (x_38 == 0) +{ +lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_39 = lean_ctor_get(x_37, 0); +x_40 = lean_ctor_get(x_37, 1); +x_41 = lean_ctor_get(x_12, 5); +lean_inc(x_41); +x_42 = l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__2; +lean_ctor_set(x_37, 1, x_39); +lean_ctor_set(x_37, 0, x_42); +x_43 = lean_box(0); +x_44 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_44, 0, x_37); +lean_ctor_set(x_44, 1, x_43); +lean_ctor_set(x_44, 2, x_43); +lean_ctor_set(x_44, 3, x_43); +lean_ctor_set(x_44, 4, x_43); +lean_ctor_set(x_44, 5, x_43); +x_45 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_45, 0, x_41); +x_46 = l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__3; +x_47 = l_Lean_Meta_Tactic_TryThis_addSuggestion(x_1, x_44, x_45, x_46, x_43, x_10, x_11, x_12, x_13, x_40); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_45); +return x_47; } else { -uint8_t x_28; +lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; 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_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_ctor_get(x_12, 5); +lean_inc(x_50); +x_51 = l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__2; +x_52 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_52, 0, x_51); +lean_ctor_set(x_52, 1, x_48); +x_53 = lean_box(0); +x_54 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_54, 0, x_52); +lean_ctor_set(x_54, 1, x_53); +lean_ctor_set(x_54, 2, x_53); +lean_ctor_set(x_54, 3, x_53); +lean_ctor_set(x_54, 4, x_53); +lean_ctor_set(x_54, 5, x_53); +x_55 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_55, 0, x_50); +x_56 = l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__3; +x_57 = l_Lean_Meta_Tactic_TryThis_addSuggestion(x_1, x_54, x_55, x_56, x_53, x_10, x_11, x_12, x_13, x_49); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_55); +return x_57; +} +} +else +{ +uint8_t x_58; +lean_dec(x_33); +lean_dec(x_15); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +x_58 = !lean_is_exclusive(x_34); +if (x_58 == 0) +{ +return x_34; +} +else +{ +lean_object* x_59; lean_object* x_60; lean_object* x_61; +x_59 = lean_ctor_get(x_34, 0); +x_60 = lean_ctor_get(x_34, 1); +lean_inc(x_60); +lean_inc(x_59); +lean_dec(x_34); +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_62; +lean_dec(x_15); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -x_28 = !lean_is_exclusive(x_23); -if (x_28 == 0) +x_62 = !lean_is_exclusive(x_29); +if (x_62 == 0) { -return x_23; +return x_29; } else { -lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_29 = lean_ctor_get(x_23, 0); -x_30 = lean_ctor_get(x_23, 1); -lean_inc(x_30); -lean_inc(x_29); -lean_dec(x_23); -x_31 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_31, 0, x_29); -lean_ctor_set(x_31, 1, x_30); -return x_31; +lean_object* x_63; lean_object* x_64; lean_object* x_65; +x_63 = lean_ctor_get(x_29, 0); +x_64 = lean_ctor_get(x_29, 1); +lean_inc(x_64); +lean_inc(x_63); +lean_dec(x_29); +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_32; -lean_dec(x_17); +{ +uint8_t x_66; +lean_dec(x_23); +lean_dec(x_15); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -x_32 = !lean_is_exclusive(x_18); -if (x_32 == 0) +x_66 = !lean_is_exclusive(x_24); +if (x_66 == 0) { -return x_18; +return x_24; } else { -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_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; +lean_object* x_67; lean_object* x_68; lean_object* x_69; +x_67 = lean_ctor_get(x_24, 0); +x_68 = lean_ctor_get(x_24, 1); +lean_inc(x_68); +lean_inc(x_67); +lean_dec(x_24); +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_36; +uint8_t x_70; +lean_dec(x_15); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -x_36 = !lean_is_exclusive(x_14); -if (x_36 == 0) +x_70 = !lean_is_exclusive(x_20); +if (x_70 == 0) { -return x_14; +return x_20; } else { -lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_37 = lean_ctor_get(x_14, 0); -x_38 = lean_ctor_get(x_14, 1); -lean_inc(x_38); -lean_inc(x_37); -lean_dec(x_14); -x_39 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_39, 0, x_37); -lean_ctor_set(x_39, 1, x_38); -return x_39; -} +lean_object* x_71; lean_object* x_72; lean_object* x_73; +x_71 = lean_ctor_get(x_20, 0); +x_72 = lean_ctor_get(x_20, 1); +lean_inc(x_72); +lean_inc(x_71); +lean_dec(x_20); +x_73 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_73, 0, x_71); +lean_ctor_set(x_73, 1, x_72); +return x_73; } } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Conv_evalDSimp(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_11; lean_object* x_12; -x_11 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Conv_evalDSimp___lambda__1), 10, 1); -lean_closure_set(x_11, 0, x_1); -x_12 = l_Lean_Elab_Tactic_withMainContext___rarg(x_11, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_12; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimp__1___closed__1() { +static lean_object* _init_l_Lean_Elab_Tactic_Conv_evalDSimpTrace___lambda__2___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("dsimp", 5, 5); +x_1 = lean_mk_string_unchecked("dsimpArgs", 9, 9); return x_1; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimp__1___closed__2() { +static lean_object* _init_l_Lean_Elab_Tactic_Conv_evalDSimpTrace___lambda__2___closed__2() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp__1___closed__1; x_2 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp__1___closed__2; x_3 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp__1___closed__3; -x_4 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp__1___closed__4; -x_5 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimp__1___closed__1; -x_6 = l_Lean_Name_mkStr5(x_1, x_2, x_3, x_4, x_5); -return x_6; +x_4 = l_Lean_Elab_Tactic_Conv_evalDSimpTrace___lambda__2___closed__1; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimp__1___closed__3() { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Conv_evalDSimpTrace___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_15 = lean_unsigned_to_nat(3u); +x_16 = l_Lean_Syntax_getArg(x_1, x_15); +x_17 = l_Lean_Syntax_isNone(x_16); +if (x_17 == 0) +{ +lean_object* x_18; uint8_t x_19; +x_18 = lean_unsigned_to_nat(1u); +lean_inc(x_16); +x_19 = l_Lean_Syntax_matchesNull(x_16, x_18); +if (x_19 == 0) +{ +lean_object* x_20; +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_3); +x_20 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalExact___spec__1___rarg(x_14); +return x_20; +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_21 = lean_unsigned_to_nat(0u); +x_22 = l_Lean_Syntax_getArg(x_16, x_21); +lean_dec(x_16); +x_23 = l_Lean_Elab_Tactic_Conv_evalDSimpTrace___lambda__2___closed__2; +lean_inc(x_22); +x_24 = l_Lean_Syntax_isOfKind(x_22, x_23); +if (x_24 == 0) +{ +lean_object* x_25; +lean_dec(x_22); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_3); +x_25 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalExact___spec__1___rarg(x_14); +return x_25; +} +else +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_26 = l_Lean_Syntax_getArg(x_22, x_18); +lean_dec(x_22); +x_27 = l_Lean_Syntax_getArgs(x_26); +lean_dec(x_26); +x_28 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_28, 0, x_27); +x_29 = lean_box(0); +x_30 = l_Lean_Elab_Tactic_Conv_evalDSimpTrace___lambda__1(x_2, x_5, x_3, x_29, x_28, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +lean_dec(x_28); +return x_30; +} +} +} +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; +lean_dec(x_16); +x_31 = lean_box(0); +x_32 = lean_box(0); +x_33 = l_Lean_Elab_Tactic_Conv_evalDSimpTrace___lambda__1(x_2, x_5, x_3, x_32, x_31, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +return x_33; +} +} +} +static lean_object* _init_l_Lean_Elab_Tactic_Conv_evalDSimpTrace___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("evalDSimp", 9, 9); +x_1 = lean_mk_string_unchecked("dsimpTrace", 10, 10); return x_1; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimp__1___closed__4() { +static lean_object* _init_l_Lean_Elab_Tactic_Conv_evalDSimpTrace___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___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp__1___closed__1; -x_2 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp__1___closed__7; +x_2 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp__1___closed__2; x_3 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp__1___closed__3; x_4 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp__1___closed__4; -x_5 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimp__1___closed__3; +x_5 = l_Lean_Elab_Tactic_Conv_evalDSimpTrace___closed__1; x_6 = l_Lean_Name_mkStr5(x_1, x_2, x_3, x_4, x_5); return x_6; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimp__1___closed__5() { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Conv_evalDSimpTrace(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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_alloc_closure((void*)(l_Lean_Elab_Tactic_Conv_evalDSimp), 10, 0); -return x_1; -} +lean_object* x_11; uint8_t x_12; +x_11 = l_Lean_Elab_Tactic_Conv_evalDSimpTrace___closed__2; +lean_inc(x_1); +x_12 = l_Lean_Syntax_isOfKind(x_1, x_11); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; +lean_dec(x_1); +x_13 = l_Lean_Elab_Tactic_Conv_evalSimpTrace___closed__3; +x_14 = l_Lean_Elab_Tactic_withMainContext___rarg(x_13, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_14; } -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimp__1(lean_object* x_1) { -_start: +else { -lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_2 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp__1___closed__10; -x_3 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimp__1___closed__2; -x_4 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimp__1___closed__4; -x_5 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimp__1___closed__5; -x_6 = l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(x_2, x_3, x_4, x_5, x_1); -return x_6; +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_15 = lean_unsigned_to_nat(0u); +x_16 = l_Lean_Syntax_getArg(x_1, x_15); +x_17 = lean_unsigned_to_nat(1u); +x_18 = l_Lean_Syntax_getArg(x_1, x_17); +x_19 = l_Lean_Elab_Tactic_Conv_evalSimpTrace___closed__5; +lean_inc(x_18); +x_20 = l_Lean_Syntax_isOfKind(x_18, x_19); +if (x_20 == 0) +{ +lean_object* x_21; lean_object* x_22; +lean_dec(x_18); +lean_dec(x_16); +lean_dec(x_1); +x_21 = l_Lean_Elab_Tactic_Conv_evalSimpTrace___closed__3; +x_22 = l_Lean_Elab_Tactic_withMainContext___rarg(x_21, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_22; } +else +{ +lean_object* x_23; lean_object* x_24; uint8_t x_25; +x_23 = lean_unsigned_to_nat(2u); +x_24 = l_Lean_Syntax_getArg(x_1, x_23); +x_25 = l_Lean_Syntax_isNone(x_24); +if (x_25 == 0) +{ +uint8_t x_26; +lean_inc(x_24); +x_26 = l_Lean_Syntax_matchesNull(x_24, x_17); +if (x_26 == 0) +{ +lean_object* x_27; lean_object* x_28; +lean_dec(x_24); +lean_dec(x_18); +lean_dec(x_16); +lean_dec(x_1); +x_27 = l_Lean_Elab_Tactic_Conv_evalSimpTrace___closed__3; +x_28 = l_Lean_Elab_Tactic_withMainContext___rarg(x_27, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_28; } -static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimp_declRange__1___closed__1() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -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); -lean_ctor_set(x_3, 1, x_2); -return x_3; +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_29 = l_Lean_Syntax_getArg(x_24, x_15); +lean_dec(x_24); +x_30 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_30, 0, x_29); +x_31 = lean_box(0); +x_32 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Conv_evalDSimpTrace___lambda__2___boxed), 14, 5); +lean_closure_set(x_32, 0, x_1); +lean_closure_set(x_32, 1, x_16); +lean_closure_set(x_32, 2, x_18); +lean_closure_set(x_32, 3, x_31); +lean_closure_set(x_32, 4, x_30); +x_33 = l_Lean_Elab_Tactic_withMainContext___rarg(x_32, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_33; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimp_declRange__1___closed__2() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(31u); -x_2 = lean_unsigned_to_nat(48u); -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_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; +lean_dec(x_24); +x_34 = lean_box(0); +x_35 = lean_box(0); +x_36 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Conv_evalDSimpTrace___lambda__2___boxed), 14, 5); +lean_closure_set(x_36, 0, x_1); +lean_closure_set(x_36, 1, x_16); +lean_closure_set(x_36, 2, x_18); +lean_closure_set(x_36, 3, x_35); +lean_closure_set(x_36, 4, x_34); +x_37 = l_Lean_Elab_Tactic_withMainContext___rarg(x_36, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_37; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimp_declRange__1___closed__3() { +} +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Conv_evalDSimpTrace___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_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimp_declRange__1___closed__1; -x_2 = lean_unsigned_to_nat(48u); -x_3 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimp_declRange__1___closed__2; -x_4 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_4, 0, x_1); -lean_ctor_set(x_4, 1, x_2); -lean_ctor_set(x_4, 2, x_3); -lean_ctor_set(x_4, 3, x_2); -return x_4; +lean_object* x_15; +x_15 = l_Lean_Elab_Tactic_Conv_evalDSimpTrace___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +return x_15; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimp_declRange__1___closed__4() { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Conv_evalDSimpTrace___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(29u); -x_2 = lean_unsigned_to_nat(52u); -x_3 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set(x_3, 1, x_2); -return x_3; +lean_object* x_15; +x_15 = l_Lean_Elab_Tactic_Conv_evalDSimpTrace___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +return x_15; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimp_declRange__1___closed__5() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimpTrace__1___closed__1() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(29u); -x_2 = lean_unsigned_to_nat(61u); -x_3 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set(x_3, 1, x_2); -return x_3; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("evalDSimpTrace", 14, 14); +return x_1; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimp_declRange__1___closed__6() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimpTrace__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___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimp_declRange__1___closed__4; -x_2 = lean_unsigned_to_nat(52u); -x_3 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimp_declRange__1___closed__5; -x_4 = lean_unsigned_to_nat(61u); -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; +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___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp__1___closed__1; +x_2 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp__1___closed__7; +x_3 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp__1___closed__3; +x_4 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp__1___closed__4; +x_5 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimpTrace__1___closed__1; +x_6 = l_Lean_Name_mkStr5(x_1, x_2, x_3, x_4, x_5); +return x_6; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimp_declRange__1___closed__7() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimpTrace__1___closed__3() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimp_declRange__1___closed__3; -x_2 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimp_declRange__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; +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Conv_evalDSimpTrace), 10, 0); +return x_1; } } -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimp_declRange__1(lean_object* x_1) { +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimpTrace__1(lean_object* x_1) { _start: { -lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_2 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimp__1___closed__4; -x_3 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimp_declRange__1___closed__7; -x_4 = l_Lean_addBuiltinDeclarationRanges(x_2, x_3, x_1); -return x_4; +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_2 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp__1___closed__10; +x_3 = l_Lean_Elab_Tactic_Conv_evalDSimpTrace___closed__2; +x_4 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimpTrace__1___closed__2; +x_5 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimpTrace__1___closed__3; +x_6 = l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(x_2, x_3, x_4, x_5, x_1); +return x_6; } } lean_object* initialize_Lean_Elab_Tactic_Simp(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Elab_Tactic_Split(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Elab_Tactic_Conv_Basic(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Elab_Tactic_SimpTrace(uint8_t builtin, lean_object*); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Lean_Elab_Tactic_Conv_Simp(uint8_t builtin, lean_object* w) { lean_object * res; @@ -1338,6 +3329,9 @@ lean_dec_ref(res); res = initialize_Lean_Elab_Tactic_Conv_Basic(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +res = initialize_Lean_Elab_Tactic_SimpTrace(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); l_Lean_Elab_Tactic_Conv_evalSimp___lambda__1___closed__1 = _init_l_Lean_Elab_Tactic_Conv_evalSimp___lambda__1___closed__1(); lean_mark_persistent(l_Lean_Elab_Tactic_Conv_evalSimp___lambda__1___closed__1); l_Lean_Elab_Tactic_Conv_evalSimp___lambda__1___closed__2 = _init_l_Lean_Elab_Tactic_Conv_evalSimp___lambda__1___closed__2(); @@ -1398,6 +3392,51 @@ lean_mark_persistent(l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp_declRange__1_ if (builtin) {res = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp_declRange__1(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +}l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__1 = _init_l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__1(); +lean_mark_persistent(l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__1); +l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__2 = _init_l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__2(); +lean_mark_persistent(l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__2); +l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__3 = _init_l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__3(); +lean_mark_persistent(l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__3); +l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__4 = _init_l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__4(); +lean_mark_persistent(l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__4); +l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__5 = _init_l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__5(); +lean_mark_persistent(l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__5); +l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__6 = _init_l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__6(); +lean_mark_persistent(l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__6); +l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__7 = _init_l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__7(); +lean_mark_persistent(l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__7); +l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__8 = _init_l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__8(); +lean_mark_persistent(l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__8); +l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__9 = _init_l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__9(); +lean_mark_persistent(l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__9); +l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__10 = _init_l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__10(); +lean_mark_persistent(l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__10); +l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__11 = _init_l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__11(); +lean_mark_persistent(l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__1___closed__11); +l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__2___closed__1 = _init_l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__2___closed__1(); +lean_mark_persistent(l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__2___closed__1); +l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__2___closed__2 = _init_l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__2___closed__2(); +lean_mark_persistent(l_Lean_Elab_Tactic_Conv_evalSimpTrace___lambda__2___closed__2); +l_Lean_Elab_Tactic_Conv_evalSimpTrace___closed__1 = _init_l_Lean_Elab_Tactic_Conv_evalSimpTrace___closed__1(); +lean_mark_persistent(l_Lean_Elab_Tactic_Conv_evalSimpTrace___closed__1); +l_Lean_Elab_Tactic_Conv_evalSimpTrace___closed__2 = _init_l_Lean_Elab_Tactic_Conv_evalSimpTrace___closed__2(); +lean_mark_persistent(l_Lean_Elab_Tactic_Conv_evalSimpTrace___closed__2); +l_Lean_Elab_Tactic_Conv_evalSimpTrace___closed__3 = _init_l_Lean_Elab_Tactic_Conv_evalSimpTrace___closed__3(); +lean_mark_persistent(l_Lean_Elab_Tactic_Conv_evalSimpTrace___closed__3); +l_Lean_Elab_Tactic_Conv_evalSimpTrace___closed__4 = _init_l_Lean_Elab_Tactic_Conv_evalSimpTrace___closed__4(); +lean_mark_persistent(l_Lean_Elab_Tactic_Conv_evalSimpTrace___closed__4); +l_Lean_Elab_Tactic_Conv_evalSimpTrace___closed__5 = _init_l_Lean_Elab_Tactic_Conv_evalSimpTrace___closed__5(); +lean_mark_persistent(l_Lean_Elab_Tactic_Conv_evalSimpTrace___closed__5); +l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpTrace__1___closed__1 = _init_l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpTrace__1___closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpTrace__1___closed__1); +l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpTrace__1___closed__2 = _init_l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpTrace__1___closed__2(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpTrace__1___closed__2); +l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpTrace__1___closed__3 = _init_l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpTrace__1___closed__3(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpTrace__1___closed__3); +if (builtin) {res = l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpTrace__1(lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); }l_Lean_Elab_Tactic_Conv_evalSimpMatch___rarg___closed__1 = _init_l_Lean_Elab_Tactic_Conv_evalSimpMatch___rarg___closed__1(); lean_mark_persistent(l_Lean_Elab_Tactic_Conv_evalSimpMatch___rarg___closed__1); l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpMatch__1___closed__1 = _init_l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpMatch__1___closed__1(); @@ -1462,6 +3501,25 @@ lean_mark_persistent(l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimp_declRange__1 if (builtin) {res = l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimp_declRange__1(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +}l_Lean_Elab_Tactic_Conv_evalDSimpTrace___lambda__1___closed__1 = _init_l_Lean_Elab_Tactic_Conv_evalDSimpTrace___lambda__1___closed__1(); +lean_mark_persistent(l_Lean_Elab_Tactic_Conv_evalDSimpTrace___lambda__1___closed__1); +l_Lean_Elab_Tactic_Conv_evalDSimpTrace___lambda__2___closed__1 = _init_l_Lean_Elab_Tactic_Conv_evalDSimpTrace___lambda__2___closed__1(); +lean_mark_persistent(l_Lean_Elab_Tactic_Conv_evalDSimpTrace___lambda__2___closed__1); +l_Lean_Elab_Tactic_Conv_evalDSimpTrace___lambda__2___closed__2 = _init_l_Lean_Elab_Tactic_Conv_evalDSimpTrace___lambda__2___closed__2(); +lean_mark_persistent(l_Lean_Elab_Tactic_Conv_evalDSimpTrace___lambda__2___closed__2); +l_Lean_Elab_Tactic_Conv_evalDSimpTrace___closed__1 = _init_l_Lean_Elab_Tactic_Conv_evalDSimpTrace___closed__1(); +lean_mark_persistent(l_Lean_Elab_Tactic_Conv_evalDSimpTrace___closed__1); +l_Lean_Elab_Tactic_Conv_evalDSimpTrace___closed__2 = _init_l_Lean_Elab_Tactic_Conv_evalDSimpTrace___closed__2(); +lean_mark_persistent(l_Lean_Elab_Tactic_Conv_evalDSimpTrace___closed__2); +l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimpTrace__1___closed__1 = _init_l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimpTrace__1___closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimpTrace__1___closed__1); +l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimpTrace__1___closed__2 = _init_l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimpTrace__1___closed__2(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimpTrace__1___closed__2); +l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimpTrace__1___closed__3 = _init_l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimpTrace__1___closed__3(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimpTrace__1___closed__3); +if (builtin) {res = l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimpTrace__1(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/Elab/Tactic/Grind.c b/stage0/stdlib/Lean/Elab/Tactic/Grind.c index 3fc44a87b94b..f011d032d6b8 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/Grind.c +++ b/stage0/stdlib/Lean/Elab/Tactic/Grind.c @@ -13,14 +13,19 @@ #ifdef __cplusplus extern "C" { #endif +lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); lean_object* l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_format_pretty(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___elambda__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_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__4; static lean_object* l_Lean_ensureNonAmbiguous___at_Lean_Elab_Tactic_elabGrindPattern___spec__9___closed__5; -static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_elabGrindPattern___spec__12___closed__2; LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Tactic_elabGrindPattern___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_grind___closed__2; +LEAN_EXPORT lean_object* l_Lean_ofExcept___at___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback_unsafe__1___spec__2(lean_object*); +lean_object* lean_mk_empty_array_with_capacity(lean_object*); static lean_object* l_Lean_throwUnknownConstant___at_Lean_Elab_Tactic_elabGrindPattern___spec__6___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescope___at_Lean_Elab_Tactic_elabGrindPattern___spec__13___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___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalApplyRfl__1___closed__1; static lean_object* l_Lean_throwUnknownConstant___at_Lean_Elab_Tactic_elabGrindPattern___spec__6___closed__4; lean_object* l___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -28,6 +33,7 @@ static lean_object* l_Lean_Elab_Tactic_elabGrindConfig___lambda__3___closed__1; LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Tactic_elabGrindPattern___spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_throwError___at_Lean_Elab_Term_mkAuxName___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_ConstantInfo_type(lean_object*); +static lean_object* l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__19; LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_elabGrindPattern___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_elabGrindConfig___lambda__1___closed__6; lean_object* l_Lean_Meta_Grind_addEMatchTheorem(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -37,12 +43,14 @@ static lean_object* l_Lean_Elab_Tactic_elabGrindConfig___lambda__1___closed__2; lean_object* l_Lean_Syntax_formatStxAux(lean_object*, uint8_t, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalApplyRfl__1___closed__4; lean_object* l_Lean_Elab_Term_elabTerm(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Meta_Grind_unfoldReducible___lambda__3___boxed(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*); LEAN_EXPORT lean_object* l_Lean_ensureNonAmbiguous___at_Lean_Elab_Tactic_elabGrindPattern___spec__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); 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*); extern lean_object* l_Lean_Elab_Tactic_tacticElabAttribute; static lean_object* l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Grind___hyg_5____closed__3; static lean_object* l_Lean_Elab_Tactic_grind___closed__1; +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalApplyRfl___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback_unsafe__1___spec__3___rarg(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_object* l_Lean_replaceRef(lean_object*, lean_object*); @@ -50,7 +58,7 @@ static lean_object* l_Lean_preprocessSyntaxAndResolve___at_Lean_Elab_Tactic_elab lean_object* l_Lean_instantiateMVars___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescope___at_Lean_Elab_Tactic_elabGrindPattern___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* l_Lean_Syntax_TSepArray_getElems___rarg(lean_object*); -static lean_object* l_Lean_Elab_Tactic_evalApplyRfl___closed__6; +static lean_object* l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__17; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_elabGrindConfig___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_ensureNonAmbiguous___at_Lean_Elab_Tactic_elabGrindPattern___spec__9___closed__4; static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalApplyRfl__1___closed__3; @@ -74,36 +82,40 @@ uint8_t l_Lean_Expr_hasSyntheticSorry(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescope___at_Lean_Elab_Tactic_elabGrindPattern___spec__13(lean_object*); lean_object* l_Lean_Parser_Tactic_getConfigItems(lean_object*); lean_object* l_List_filterMapTR_go___at_Lean_preprocessSyntaxAndResolve___spec__1(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_evalApplyRfl___closed__9; lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_elabGrindPattern___spec__1___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Tactic_elabGrindConfig___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_Tactic_elabGrindPattern___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_expr_abstract(lean_object*, lean_object*); +lean_object* l_Lean_Meta_withLCtx___at_Lean_Elab_Term_MVarErrorInfo_logError___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_List_mapTR_loop___at_Lean_filterFieldList___spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_elabGrindConfig___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_Grind_preprocessPattern(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_throwUnknownConstant___at_Lean_Elab_Tactic_elabGrindPattern___spec__6___closed__1; -lean_object* l_Lean_Core_transform___at_Lean_Meta_Grind_unfoldReducible___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_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__13; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_elabGrindConfig___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_elabGrindConfig___lambda__1___closed__5; LEAN_EXPORT lean_object* l_Lean_filterFieldList___at_Lean_Elab_Tactic_elabGrindPattern___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Elab_Command_commandElabAttribute; static lean_object* l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Grind___hyg_5____closed__4; static lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_elabGrindPattern___spec__1___rarg___closed__2; -static lean_object* l_Lean_Elab_Tactic_evalApplyRfl___closed__8; +static lean_object* l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__11; static lean_object* l_Lean_Elab_Tactic_elabGrindConfig___lambda__1___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*); -static lean_object* l_Lean_Elab_Tactic_evalApplyRfl___closed__5; +lean_object* lean_eval_const(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__5; lean_object* l_Lean_Elab_goalsToMessageData(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_elabGrindPattern___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofFormat(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwUnknownConstant___at_Lean_Elab_Tactic_elabGrindPattern___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback(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*); LEAN_EXPORT lean_object* l_Lean_throwUnknownConstant___at_Lean_Elab_Tactic_elabGrindPattern___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_evalApplyRfl___closed__4; static lean_object* l_Lean_Elab_Tactic_elabGrindPattern___closed__5; static lean_object* l___regBuiltin_Lean_Elab_Tactic_elabGrindPattern__1___closed__2; lean_object* l_Lean_Elab_Tactic_withMainContext___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___elambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Command_liftTermElabM___rarg(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_get(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*); @@ -114,26 +126,39 @@ lean_object* lean_array_to_list(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Tactic_elabGrindPattern___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Term_getDeclName_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_elabGrindPattern___closed__1; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalApplyRfl___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalApplyRfl___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_evalApplyRfl___lambda__2___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_elabGrindConfig___lambda__3(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_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback_unsafe__1___spec__3(lean_object*); lean_object* l___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_mkConfigItemViews(lean_object*); static lean_object* l_Lean_preprocessSyntaxAndResolve___at_Lean_Elab_Tactic_elabGrindPattern___spec__7___closed__3; static lean_object* l_Lean_Elab_Tactic_elabGrindConfig___lambda__4___closed__5; static lean_object* l_Lean_Elab_Tactic_elabGrindConfig___lambda__4___closed__2; +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalApplyRfl___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* l_Lean_Name_str___override(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_elabGrindPattern___closed__6; +LEAN_EXPORT lean_object* l_Lean_evalConst___at___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback_unsafe__1___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Elab_Term_synthesizeSyntheticMVarsUsingDefault(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_evalApplyRfl___lambda__2___closed__5; 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*); LEAN_EXPORT lean_object* l_Lean_filterFieldList___at_Lean_Elab_Tactic_elabGrindPattern___spec__5___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Syntax_matchesNull(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_elabGrindConfig___lambda__4___closed__3; +static lean_object* l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__20; +static lean_object* l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__12; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_elabGrindConfig___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_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__15; lean_object* l_Lean_resolveGlobalName___at_Lean_Elab_Term_resolveLocalName_loop___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_elabGrindConfig___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_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__7; static lean_object* l_Lean_Elab_Tactic_elabGrindConfig___lambda__4___closed__1; extern lean_object* l_Std_Format_defWidth; +static lean_object* l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__14; static lean_object* l_Lean_Elab_Tactic_elabGrindPattern___closed__3; static lean_object* l_Lean_preprocessSyntaxAndResolve___at_Lean_Elab_Tactic_elabGrindPattern___spec__7___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_elabGrindConfig(lean_object*, 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_evalApplyRfl___lambda__2___closed__3; lean_object* l_Lean_MessageData_ofConstName(lean_object*, uint8_t); uint8_t l_Lean_Environment_contains(lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofExpr(lean_object*); @@ -141,28 +166,39 @@ static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalApplyRfl__1___closed__2; lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_forallTelescopeReducingAuxAux___rarg(uint8_t, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_elabGrindPattern___spec__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_elabGrindConfig___lambda__4___closed__4; -lean_object* l_Lean_Meta_Grind_main(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__16; +lean_object* l_Lean_Meta_Grind_main(lean_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_elabGrindConfig___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_Tactic_evalApplyRfl___lambda__2___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Grind___hyg_5_(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_resolveGlobalConst___at_Lean_Elab_Tactic_elabGrindPattern___spec__3___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_elabGrindPattern__1(lean_object*); +static lean_object* l_Lean_Elab_Tactic_evalApplyRfl___lambda__2___closed__1; static lean_object* l_panic___at_Lean_Elab_Tactic_elabGrindPattern___spec__10___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_elabGrindPattern___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_Tactic_evalApplyRfl___closed__3; +lean_object* l_Lean_Expr_app___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_filterFieldList___at_Lean_Elab_Tactic_elabGrindPattern___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___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback_unsafe__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_preprocessSyntaxAndResolve___at_Lean_Elab_Tactic_elabGrindPattern___spec__7___closed__1; static lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_elabGrindPattern___spec__1___rarg___closed__1; -static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_elabGrindPattern___spec__12___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*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_elabGrindPattern__1___closed__5; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_grind(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_grind(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_evalApplyRfl___closed__2; static lean_object* l_Lean_Elab_Tactic_elabGrindPattern___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_Lean_throwError___at___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback_unsafe__1___spec__3___rarg___boxed(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_Lean_ofExcept___at___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback_unsafe__1___spec__2___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_panic_fn(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Grind___hyg_5____boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +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*); +LEAN_EXPORT lean_object* l_Lean_evalConst___at___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback_unsafe__1___spec__1(lean_object*); static lean_object* l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Grind___hyg_5____closed__1; LEAN_EXPORT lean_object* l___private_Lean_ResolveName_0__Lean_resolveGlobalConstCore___at_Lean_Elab_Tactic_elabGrindPattern___spec__4___boxed(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___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__6; LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalApplyRfl___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_throwUnsupportedSyntax___at_Lean_Elab_Tactic_elabGrindPattern___spec__1___boxed(lean_object*, lean_object*); lean_object* l_List_toString___at_Lean_ensureNoOverload___spec__2(lean_object*); @@ -182,29 +218,38 @@ size_t lean_array_size(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_elabGrindPattern__1___closed__6; lean_object* l_Lean_Elab_Term_addTermInfo(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_elabGrindConfig___lambda__2___closed__2; +LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback_unsafe__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_ResolveName_0__Lean_resolveGlobalConstCore___at_Lean_Elab_Tactic_elabGrindPattern___spec__4(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___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___elambda__1___rarg(lean_object*); +static lean_object* l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__10; lean_object* l_Lean_Elab_Tactic_liftMetaFinishingTactic(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_preprocessSyntaxAndResolve___at_Lean_Elab_Tactic_elabGrindPattern___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__1; lean_object* lean_string_append(lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__8; lean_object* l_Lean_Meta_evalExpr_x27___rarg(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_elabGrindPattern__1___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_elabGrindPattern(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_ensureNonAmbiguous___at_Lean_Elab_Tactic_elabGrindPattern___spec__9___closed__1; +static lean_object* l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__9; extern lean_object* l_Lean_Elab_unsupportedSyntaxExceptionId; uint8_t lean_usize_dec_lt(size_t, size_t); +LEAN_EXPORT lean_object* l_Lean_ofExcept___at___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback_unsafe__1___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_evalConst___at___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback_unsafe__1___spec__1___rarg(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*); uint8_t l_Lean_Exception_isRuntime(lean_object*); 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_object* l_Lean_Meta_Grind_unfoldReducible___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalApplyRfl___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_uset(lean_object*, size_t, lean_object*); static lean_object* l_Lean_throwUnknownConstant___at_Lean_Elab_Tactic_elabGrindPattern___spec__6___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Tactic_elabGrindPattern__1___closed__3; lean_object* l_Lean_MessageData_ofName(lean_object*); -static lean_object* l_Lean_Elab_Tactic_evalApplyRfl___closed__7; +static lean_object* l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__3; lean_object* l_List_filterTR_loop___at_Lean_filterFieldList___spec__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__18; static lean_object* l_Lean_Elab_Tactic_elabGrindConfig___lambda__1___closed__1; static lean_object* l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Grind___hyg_5____closed__2; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_elabGrindPattern___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*); @@ -675,18 +720,20 @@ return x_18; static lean_object* _init_l_Lean_Elab_Tactic_elabGrindConfig___lambda__3___closed__1() { _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 = lean_unsigned_to_nat(100u); -x_3 = lean_unsigned_to_nat(5u); -x_4 = lean_unsigned_to_nat(1000u); -x_5 = lean_alloc_ctor(0, 4, 1); -lean_ctor_set(x_5, 0, x_2); -lean_ctor_set(x_5, 1, x_3); -lean_ctor_set(x_5, 2, x_3); -lean_ctor_set(x_5, 3, x_4); -lean_ctor_set_uint8(x_5, sizeof(void*)*4, x_1); -return x_5; +lean_object* x_1; lean_object* x_2; uint8_t x_3; lean_object* x_4; +x_1 = lean_unsigned_to_nat(5u); +x_2 = lean_unsigned_to_nat(1000u); +x_3 = 1; +x_4 = lean_alloc_ctor(0, 4, 4); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_1); +lean_ctor_set(x_4, 2, x_1); +lean_ctor_set(x_4, 3, x_2); +lean_ctor_set_uint8(x_4, sizeof(void*)*4, x_3); +lean_ctor_set_uint8(x_4, sizeof(void*)*4 + 1, x_3); +lean_ctor_set_uint8(x_4, sizeof(void*)*4 + 2, x_3); +lean_ctor_set_uint8(x_4, sizeof(void*)*4 + 3, x_3); +return x_4; } } LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_elabGrindConfig___lambda__3(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) { @@ -1746,22 +1793,6 @@ return x_17; } } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_elabGrindPattern___spec__12___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_unfoldReducible___lambda__2), 6, 0); -return x_1; -} -} -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_elabGrindPattern___spec__12___closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_unfoldReducible___lambda__3___boxed), 6, 0); -return x_1; -} -} LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_elabGrindPattern___spec__12(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { @@ -1799,25 +1830,36 @@ lean_inc(x_1); x_19 = l_Lean_Elab_Term_elabTerm(x_15, x_1, x_18, x_18, x_6, x_7, x_8, x_9, x_10, x_11, x_12); if (lean_obj_tag(x_19) == 0) { -lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +lean_object* x_20; lean_object* x_21; lean_object* x_22; x_20 = lean_ctor_get(x_19, 0); lean_inc(x_20); x_21 = lean_ctor_get(x_19, 1); lean_inc(x_21); lean_dec(x_19); -x_22 = l_Lean_instantiateMVars___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__1(x_20, x_6, x_7, x_8, x_9, x_10, x_11, x_21); -x_23 = lean_ctor_get(x_22, 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_22 = l_Lean_Elab_Term_synthesizeSyntheticMVarsUsingDefault(x_6, x_7, x_8, x_9, x_10, x_11, x_21); +if (lean_obj_tag(x_22) == 0) +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +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_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_elabGrindPattern___spec__12___closed__1; -x_26 = l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_elabGrindPattern___spec__12___closed__2; +x_24 = l_Lean_instantiateMVars___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__1(x_20, x_6, x_7, x_8, x_9, x_10, x_11, x_23); +x_25 = lean_ctor_get(x_24, 0); +lean_inc(x_25); +x_26 = lean_ctor_get(x_24, 1); +lean_inc(x_26); +lean_dec(x_24); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); -x_27 = l_Lean_Core_transform___at_Lean_Meta_Grind_unfoldReducible___spec__1(x_23, x_25, x_26, x_8, x_9, x_10, x_11, x_24); +x_27 = l_Lean_Meta_Grind_preprocessPattern(x_25, x_18, x_8, x_9, x_10, x_11, x_26); if (lean_obj_tag(x_27) == 0) { lean_object* x_28; lean_object* x_29; lean_object* x_30; size_t x_31; size_t x_32; lean_object* x_33; @@ -1870,6 +1912,7 @@ return x_38; else { uint8_t x_39; +lean_dec(x_20); lean_dec(x_17); lean_dec(x_11); lean_dec(x_10); @@ -1878,19 +1921,19 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_1); -x_39 = !lean_is_exclusive(x_19); +x_39 = !lean_is_exclusive(x_22); if (x_39 == 0) { -return x_19; +return x_22; } else { lean_object* x_40; lean_object* x_41; lean_object* x_42; -x_40 = lean_ctor_get(x_19, 0); -x_41 = lean_ctor_get(x_19, 1); +x_40 = lean_ctor_get(x_22, 0); +x_41 = lean_ctor_get(x_22, 1); lean_inc(x_41); lean_inc(x_40); -lean_dec(x_19); +lean_dec(x_22); x_42 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_42, 0, x_40); lean_ctor_set(x_42, 1, x_41); @@ -1898,6 +1941,37 @@ return x_42; } } } +else +{ +uint8_t x_43; +lean_dec(x_17); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_1); +x_43 = !lean_is_exclusive(x_19); +if (x_43 == 0) +{ +return x_19; +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_19, 0); +x_45 = lean_ctor_get(x_19, 1); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_19); +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_EXPORT lean_object* l_Lean_Meta_forallTelescope___at_Lean_Elab_Tactic_elabGrindPattern___spec__13___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) { @@ -2535,225 +2609,1008 @@ x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_grind(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_grind(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { -lean_object* x_9; +lean_object* x_10; +lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -lean_inc(x_4); -x_9 = l_Lean_Meta_Grind_main(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_9) == 0) +x_10 = l_Lean_Meta_Grind_main(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_10; -x_10 = !lean_is_exclusive(x_9); -if (x_10 == 0) +uint8_t x_11; +x_11 = !lean_is_exclusive(x_10); +if (x_11 == 0) { -lean_object* x_11; lean_object* x_12; uint8_t x_13; -x_11 = lean_ctor_get(x_9, 0); -x_12 = lean_ctor_get(x_9, 1); -x_13 = l_List_isEmpty___rarg(x_11); -if (x_13 == 0) +lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_12 = lean_ctor_get(x_10, 0); +x_13 = lean_ctor_get(x_10, 1); +x_14 = l_List_isEmpty___rarg(x_12); +if (x_14 == 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_free_object(x_9); -x_14 = l_Lean_Elab_goalsToMessageData(x_11); -x_15 = l_Lean_Elab_Tactic_grind___closed__2; -x_16 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_16, 0, x_15); -lean_ctor_set(x_16, 1, x_14); -x_17 = l_Lean_Elab_Tactic_elabGrindConfig___lambda__1___closed__6; -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___at_Lean_Meta_setInlineAttribute___spec__1(x_18, x_4, x_5, x_6, x_7, x_12); +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +lean_free_object(x_10); +x_15 = l_Lean_Elab_goalsToMessageData(x_12); +x_16 = l_Lean_Elab_Tactic_grind___closed__2; +x_17 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_17, 0, x_16); +lean_ctor_set(x_17, 1, x_15); +x_18 = l_Lean_Elab_Tactic_elabGrindConfig___lambda__1___closed__6; +x_19 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_19, 0, x_17); +lean_ctor_set(x_19, 1, x_18); +x_20 = l_Lean_throwError___at_Lean_Meta_setInlineAttribute___spec__1(x_19, x_5, x_6, x_7, x_8, x_13); +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); -return x_19; +return x_20; } else { -lean_object* x_20; -lean_dec(x_11); +lean_object* x_21; +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_20 = lean_box(0); -lean_ctor_set(x_9, 0, x_20); -return x_9; +x_21 = lean_box(0); +lean_ctor_set(x_10, 0, x_21); +return x_10; } } else { -lean_object* x_21; lean_object* x_22; uint8_t x_23; -x_21 = lean_ctor_get(x_9, 0); -x_22 = lean_ctor_get(x_9, 1); +lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_22 = lean_ctor_get(x_10, 0); +x_23 = lean_ctor_get(x_10, 1); +lean_inc(x_23); lean_inc(x_22); -lean_inc(x_21); -lean_dec(x_9); -x_23 = l_List_isEmpty___rarg(x_21); -if (x_23 == 0) +lean_dec(x_10); +x_24 = l_List_isEmpty___rarg(x_22); +if (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 = l_Lean_Elab_goalsToMessageData(x_21); -x_25 = l_Lean_Elab_Tactic_grind___closed__2; -x_26 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_26, 0, x_25); -lean_ctor_set(x_26, 1, x_24); -x_27 = l_Lean_Elab_Tactic_elabGrindConfig___lambda__1___closed__6; -x_28 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set(x_28, 1, x_27); -x_29 = l_Lean_throwError___at_Lean_Meta_setInlineAttribute___spec__1(x_28, x_4, x_5, x_6, x_7, x_22); +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 = l_Lean_Elab_goalsToMessageData(x_22); +x_26 = l_Lean_Elab_Tactic_grind___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_Lean_Elab_Tactic_elabGrindConfig___lambda__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_throwError___at_Lean_Meta_setInlineAttribute___spec__1(x_29, 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); -return x_29; +return x_30; } else { -lean_object* x_30; lean_object* x_31; -lean_dec(x_21); +lean_object* x_31; lean_object* x_32; +lean_dec(x_22); +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); -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_22); -return x_31; +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_23); +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); -x_32 = !lean_is_exclusive(x_9); -if (x_32 == 0) +x_33 = !lean_is_exclusive(x_10); +if (x_33 == 0) { -return x_9; +return x_10; } else { -lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_33 = lean_ctor_get(x_9, 0); -x_34 = lean_ctor_get(x_9, 1); +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_10, 0); +x_35 = lean_ctor_get(x_10, 1); +lean_inc(x_35); lean_inc(x_34); -lean_inc(x_33); -lean_dec(x_9); -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_10); +x_36 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_36, 0, x_34); +lean_ctor_set(x_36, 1, x_35); +return x_36; } } } } -LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalApplyRfl___spec__1___rarg(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback_unsafe__1___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: { -lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_elabGrindPattern___spec__1___rarg___closed__2; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set(x_3, 1, x_1); -return x_3; -} +lean_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, 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_is_exclusive(x_12); +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_12, 0); +x_15 = lean_ctor_get(x_12, 1); +x_16 = l_Lean_Elab_addMacroStack___at_Lean_Elab_Term_instAddErrorMessageContextTermElabM___spec__1(x_14, x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_15); +lean_dec(x_2); +x_17 = !lean_is_exclusive(x_16); +if (x_17 == 0) +{ +lean_object* x_18; +x_18 = lean_ctor_get(x_16, 0); +lean_ctor_set(x_12, 1, x_18); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set_tag(x_16, 1); +lean_ctor_set(x_16, 0, x_12); +return x_16; } -LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalApplyRfl___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; -x_9 = lean_alloc_closure((void*)(l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalApplyRfl___spec__1___rarg), 1, 0); -return x_9; +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_16, 0); +x_20 = lean_ctor_get(x_16, 1); +lean_inc(x_20); +lean_inc(x_19); +lean_dec(x_16); +lean_ctor_set(x_12, 1, x_19); +lean_ctor_set(x_12, 0, x_11); +x_21 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_21, 0, x_12); +lean_ctor_set(x_21, 1, x_20); +return x_21; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalApplyRfl___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: +else { -lean_object* x_9; -x_9 = l_Lean_Elab_Tactic_grind(x_3, x_1, x_2, x_4, x_5, x_6, x_7, x_8); -return x_9; +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_22 = lean_ctor_get(x_12, 0); +x_23 = lean_ctor_get(x_12, 1); +lean_inc(x_23); +lean_inc(x_22); +lean_dec(x_12); +x_24 = l_Lean_Elab_addMacroStack___at_Lean_Elab_Term_instAddErrorMessageContextTermElabM___spec__1(x_22, x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_23); +lean_dec(x_2); +x_25 = lean_ctor_get(x_24, 0); +lean_inc(x_25); +x_26 = lean_ctor_get(x_24, 1); +lean_inc(x_26); +if (lean_is_exclusive(x_24)) { + lean_ctor_release(x_24, 0); + lean_ctor_release(x_24, 1); + x_27 = x_24; +} else { + lean_dec_ref(x_24); + x_27 = lean_box(0); +} +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_11); +lean_ctor_set(x_28, 1, x_25); +if (lean_is_scalar(x_27)) { + x_29 = lean_alloc_ctor(1, 2, 0); +} else { + x_29 = x_27; + lean_ctor_set_tag(x_29, 1); } +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_26); +return x_29; } -static lean_object* _init_l_Lean_Elab_Tactic_evalApplyRfl___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("grind", 5, 5); -return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalApplyRfl___closed__2() { +LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback_unsafe__1___spec__3(lean_object* x_1) { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Grind___hyg_5____closed__1; -x_2 = l_Lean_Elab_Tactic_elabGrindPattern___closed__1; -x_3 = l___regBuiltin_Lean_Elab_Tactic_elabGrindPattern__1___closed__2; -x_4 = l_Lean_Elab_Tactic_evalApplyRfl___closed__1; -x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); -return x_5; +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_throwError___at___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback_unsafe__1___spec__3___rarg___boxed), 8, 0); +return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalApplyRfl___closed__3() { +LEAN_EXPORT lean_object* l_Lean_ofExcept___at___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback_unsafe__1___spec__2___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("optConfig", 9, 9); -return x_1; +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_9 = lean_ctor_get(x_1, 0); +x_10 = l_Lean_stringToMessageData(x_9); +x_11 = l_Lean_throwError___at___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback_unsafe__1___spec__3___rarg(x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +return x_11; +} +else +{ +lean_object* x_12; lean_object* x_13; +lean_dec(x_2); +x_12 = lean_ctor_get(x_1, 0); +lean_inc(x_12); +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; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalApplyRfl___closed__4() { +} +LEAN_EXPORT lean_object* l_Lean_ofExcept___at___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback_unsafe__1___spec__2(lean_object* x_1) { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Grind___hyg_5____closed__1; -x_2 = l_Lean_Elab_Tactic_elabGrindPattern___closed__1; -x_3 = l___regBuiltin_Lean_Elab_Tactic_elabGrindPattern__1___closed__2; -x_4 = l_Lean_Elab_Tactic_evalApplyRfl___closed__3; -x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); -return x_5; +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_ofExcept___at___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback_unsafe__1___spec__2___rarg___boxed), 8, 0); +return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalApplyRfl___closed__5() { +LEAN_EXPORT lean_object* l_Lean_evalConst___at___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback_unsafe__1___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) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("The `grind` tactic is experimental and still under development. Avoid using it in production projects", 101, 101); +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_9 = lean_st_ref_get(x_7, x_8); +x_10 = lean_ctor_get(x_9, 0); +lean_inc(x_10); +x_11 = lean_ctor_get(x_9, 1); +lean_inc(x_11); +lean_dec(x_9); +x_12 = lean_ctor_get(x_10, 0); +lean_inc(x_12); +lean_dec(x_10); +x_13 = lean_ctor_get(x_6, 2); +x_14 = lean_eval_const(x_12, x_13, x_1); +lean_dec(x_12); +x_15 = l_Lean_ofExcept___at___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback_unsafe__1___spec__2___rarg(x_14, x_2, x_3, x_4, x_5, x_6, x_7, x_11); +lean_dec(x_14); +return x_15; +} +} +LEAN_EXPORT lean_object* l_Lean_evalConst___at___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback_unsafe__1___spec__1(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_evalConst___at___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback_unsafe__1___spec__1___rarg___boxed), 8, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback_unsafe__1(lean_object* x_1, lean_object* x_2, lean_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_evalConst___at___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback_unsafe__1___spec__1___rarg(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback_unsafe__1___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_throwError___at___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback_unsafe__1___spec__3___rarg(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_ofExcept___at___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback_unsafe__1___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) { +_start: +{ +lean_object* x_9; +x_9 = l_Lean_ofExcept___at___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback_unsafe__1___spec__2___rarg(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_1); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_evalConst___at___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback_unsafe__1___spec__1___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; +x_9 = l_Lean_evalConst___at___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback_unsafe__1___spec__1___rarg(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_1); +return x_9; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback_unsafe__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; +x_9 = l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback_unsafe__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_1); +return x_9; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___elambda__1___rarg(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = lean_box(0); +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___elambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; +x_9 = lean_alloc_closure((void*)(l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___elambda__1___rarg), 1, 0); +return x_9; +} +} +static lean_object* _init_l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___elambda__1___boxed), 8, 0); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Meta", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("GoalM", 5, 5); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___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_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Grind___hyg_5____closed__1; +x_2 = l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__2; +x_3 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Grind___hyg_5____closed__2; +x_4 = l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__3; +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_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__4; +x_3 = l_Lean_Expr_const___override(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Unit", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__6; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__7; +x_3 = l_Lean_Expr_const___override(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__5; +x_2 = l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__8; +x_3 = l_Lean_Expr_app___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__10() { +_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___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__10; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(32u); +x_2 = lean_mk_empty_array_with_capacity(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__12; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__14() { +_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___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__13; +x_3 = l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__12; +x_4 = lean_unsigned_to_nat(0u); +x_5 = lean_alloc_ctor(0, 4, sizeof(size_t)*1); +lean_ctor_set(x_5, 0, x_2); +lean_ctor_set(x_5, 1, x_3); +lean_ctor_set(x_5, 2, x_4); +lean_ctor_set(x_5, 3, x_4); +lean_ctor_set_usize(x_5, 4, x_1); +return x_5; +} +} +static lean_object* _init_l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__15() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__11; +x_2 = l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__14; +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_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__16() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = lean_mk_empty_array_with_capacity(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__17() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__9; +x_2 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__18() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_evalConst___at___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback_unsafe__1___spec__1___rarg___boxed), 8, 0); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__19() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("_grind_fallback", 15, 15); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__20() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__19; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_9; lean_object* x_10; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_9 = l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___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); +return x_10; +} +else +{ +uint8_t x_11; +x_11 = !lean_is_exclusive(x_1); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_12 = lean_ctor_get(x_1, 0); +x_13 = lean_box(0); +x_14 = l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__17; +x_15 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabTermAndSynthesize), 9, 2); +lean_closure_set(x_15, 0, x_12); +lean_closure_set(x_15, 1, x_14); +x_16 = l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__15; +x_17 = l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__16; +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_18 = l_Lean_Meta_withLCtx___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__3___rarg(x_16, x_17, x_15, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_18) == 0) +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_18, 1); +lean_inc(x_20); +lean_dec(x_18); +x_21 = l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__18; +if (lean_obj_tag(x_19) == 4) +{ +lean_object* x_22; lean_object* x_23; +lean_free_object(x_1); +x_22 = lean_ctor_get(x_19, 0); +lean_inc(x_22); +lean_dec(x_19); +x_23 = lean_apply_8(x_21, x_22, x_2, x_3, x_4, x_5, x_6, x_7, x_20); +return x_23; +} +else +{ +lean_object* x_24; lean_object* x_25; +x_24 = l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__20; +lean_inc(x_2); +x_25 = l_Lean_Elab_Term_mkAuxName(x_24, x_2, x_3, x_4, x_5, x_6, x_7, x_20); +if (lean_obj_tag(x_25) == 0) +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; uint8_t x_32; lean_object* x_33; lean_object* x_34; +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___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__9; +lean_inc(x_26); +x_29 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_29, 0, x_26); +lean_ctor_set(x_29, 1, x_13); +lean_ctor_set(x_29, 2, x_28); +lean_inc(x_26); +x_30 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_30, 0, x_26); +lean_ctor_set(x_30, 1, x_13); +x_31 = lean_box(0); +x_32 = 1; +x_33 = lean_alloc_ctor(0, 4, 1); +lean_ctor_set(x_33, 0, x_29); +lean_ctor_set(x_33, 1, x_19); +lean_ctor_set(x_33, 2, x_31); +lean_ctor_set(x_33, 3, x_30); +lean_ctor_set_uint8(x_33, sizeof(void*)*4, x_32); +lean_ctor_set(x_1, 0, x_33); +lean_inc(x_7); +lean_inc(x_6); +x_34 = l_Lean_addAndCompile(x_1, x_6, x_7, x_27); +if (lean_obj_tag(x_34) == 0) +{ +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 = lean_apply_8(x_21, x_26, x_2, x_3, x_4, x_5, x_6, x_7, x_35); +return x_36; +} +else +{ +uint8_t x_37; +lean_dec(x_26); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_37 = !lean_is_exclusive(x_34); +if (x_37 == 0) +{ +return x_34; +} +else +{ +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_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; +} +} +} +else +{ +uint8_t x_41; +lean_dec(x_19); +lean_free_object(x_1); +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_41 = !lean_is_exclusive(x_25); +if (x_41 == 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_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; +} +} +} +} +else +{ +uint8_t x_45; +lean_free_object(x_1); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_45 = !lean_is_exclusive(x_18); +if (x_45 == 0) +{ +return x_18; +} +else +{ +lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_46 = lean_ctor_get(x_18, 0); +x_47 = lean_ctor_get(x_18, 1); +lean_inc(x_47); +lean_inc(x_46); +lean_dec(x_18); +x_48 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_48, 0, x_46); +lean_ctor_set(x_48, 1, x_47); +return x_48; +} +} +} +else +{ +lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_49 = lean_ctor_get(x_1, 0); +lean_inc(x_49); +lean_dec(x_1); +x_50 = lean_box(0); +x_51 = l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__17; +x_52 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabTermAndSynthesize), 9, 2); +lean_closure_set(x_52, 0, x_49); +lean_closure_set(x_52, 1, x_51); +x_53 = l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__15; +x_54 = l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__16; +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_55 = l_Lean_Meta_withLCtx___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__3___rarg(x_53, x_54, x_52, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +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); +x_58 = l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__18; +if (lean_obj_tag(x_56) == 4) +{ +lean_object* x_59; lean_object* x_60; +x_59 = lean_ctor_get(x_56, 0); +lean_inc(x_59); +lean_dec(x_56); +x_60 = lean_apply_8(x_58, x_59, x_2, x_3, x_4, x_5, x_6, x_7, x_57); +return x_60; +} +else +{ +lean_object* x_61; lean_object* x_62; +x_61 = l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__20; +lean_inc(x_2); +x_62 = l_Lean_Elab_Term_mkAuxName(x_61, x_2, x_3, x_4, x_5, x_6, x_7, x_57); +if (lean_obj_tag(x_62) == 0) +{ +lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; uint8_t x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; +x_63 = lean_ctor_get(x_62, 0); +lean_inc(x_63); +x_64 = lean_ctor_get(x_62, 1); +lean_inc(x_64); +lean_dec(x_62); +x_65 = l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__9; +lean_inc(x_63); +x_66 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_66, 0, x_63); +lean_ctor_set(x_66, 1, x_50); +lean_ctor_set(x_66, 2, x_65); +lean_inc(x_63); +x_67 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_67, 0, x_63); +lean_ctor_set(x_67, 1, x_50); +x_68 = lean_box(0); +x_69 = 1; +x_70 = lean_alloc_ctor(0, 4, 1); +lean_ctor_set(x_70, 0, x_66); +lean_ctor_set(x_70, 1, x_56); +lean_ctor_set(x_70, 2, x_68); +lean_ctor_set(x_70, 3, x_67); +lean_ctor_set_uint8(x_70, sizeof(void*)*4, x_69); +x_71 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_71, 0, x_70); +lean_inc(x_7); +lean_inc(x_6); +x_72 = l_Lean_addAndCompile(x_71, x_6, x_7, x_64); +if (lean_obj_tag(x_72) == 0) +{ +lean_object* x_73; lean_object* x_74; +x_73 = lean_ctor_get(x_72, 1); +lean_inc(x_73); +lean_dec(x_72); +x_74 = lean_apply_8(x_58, x_63, x_2, x_3, x_4, x_5, x_6, x_7, x_73); +return x_74; +} +else +{ +lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; +lean_dec(x_63); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_75 = lean_ctor_get(x_72, 0); +lean_inc(x_75); +x_76 = lean_ctor_get(x_72, 1); +lean_inc(x_76); +if (lean_is_exclusive(x_72)) { + lean_ctor_release(x_72, 0); + lean_ctor_release(x_72, 1); + x_77 = x_72; +} else { + lean_dec_ref(x_72); + x_77 = lean_box(0); +} +if (lean_is_scalar(x_77)) { + x_78 = lean_alloc_ctor(1, 2, 0); +} else { + x_78 = x_77; +} +lean_ctor_set(x_78, 0, x_75); +lean_ctor_set(x_78, 1, x_76); +return x_78; +} +} +else +{ +lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; +lean_dec(x_56); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_79 = lean_ctor_get(x_62, 0); +lean_inc(x_79); +x_80 = lean_ctor_get(x_62, 1); +lean_inc(x_80); +if (lean_is_exclusive(x_62)) { + lean_ctor_release(x_62, 0); + lean_ctor_release(x_62, 1); + x_81 = x_62; +} else { + lean_dec_ref(x_62); + 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; lean_object* x_86; +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_83 = lean_ctor_get(x_55, 0); +lean_inc(x_83); +x_84 = lean_ctor_get(x_55, 1); +lean_inc(x_84); +if (lean_is_exclusive(x_55)) { + lean_ctor_release(x_55, 0); + lean_ctor_release(x_55, 1); + x_85 = x_55; +} else { + lean_dec_ref(x_55); + 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_EXPORT lean_object* l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___elambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; +x_9 = l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___elambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_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_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalApplyRfl___spec__1___rarg(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_elabGrindPattern___spec__1___rarg___closed__2; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalApplyRfl___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 = lean_alloc_closure((void*)(l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalApplyRfl___spec__1___rarg), 1, 0); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalApplyRfl___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l_Lean_Elab_Tactic_grind(x_4, x_1, x_2, x_3, x_5, x_6, x_7, x_8, x_9); +return x_10; +} +} +static lean_object* _init_l_Lean_Elab_Tactic_evalApplyRfl___lambda__2___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("The `grind` tactic is experimental and still under development. Avoid using it in production projects", 101, 101); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalApplyRfl___closed__6() { +static lean_object* _init_l_Lean_Elab_Tactic_evalApplyRfl___lambda__2___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic_evalApplyRfl___closed__5; +x_1 = l_Lean_Elab_Tactic_evalApplyRfl___lambda__2___closed__1; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalApplyRfl___closed__7() { +static lean_object* _init_l_Lean_Elab_Tactic_evalApplyRfl___lambda__2___closed__3() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic_evalApplyRfl___closed__6; +x_1 = l_Lean_Elab_Tactic_evalApplyRfl___lambda__2___closed__2; x_2 = l_Lean_MessageData_ofFormat(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalApplyRfl___closed__8() { +static lean_object* _init_l_Lean_Elab_Tactic_evalApplyRfl___lambda__2___closed__4() { _start: { lean_object* x_1; @@ -2761,16 +3618,236 @@ x_1 = lean_mk_string_unchecked("_grind", 6, 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalApplyRfl___closed__9() { +static lean_object* _init_l_Lean_Elab_Tactic_evalApplyRfl___lambda__2___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic_evalApplyRfl___closed__8; +x_2 = l_Lean_Elab_Tactic_evalApplyRfl___lambda__2___closed__4; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalApplyRfl___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +lean_object* x_14; +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +x_14 = l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback(x_4, x_7, x_8, 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; 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; +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +x_16 = lean_ctor_get(x_14, 1); +lean_inc(x_16); +lean_dec(x_14); +x_17 = l_Lean_Elab_Tactic_evalApplyRfl___lambda__2___closed__3; +x_18 = 1; +lean_inc(x_11); +x_19 = l_Lean_logAt___at_Lean_Elab_Tactic_closeUsingOrAdmit___spec__2(x_1, x_17, x_18, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_16); +x_20 = lean_ctor_get(x_19, 1); +lean_inc(x_20); +lean_dec(x_19); +x_21 = l_Lean_Elab_Term_getDeclName_x3f(x_7, x_8, x_9, x_10, x_11, x_12, x_20); +x_22 = lean_ctor_get(x_21, 0); +lean_inc(x_22); +x_23 = lean_ctor_get(x_21, 1); +lean_inc(x_23); +lean_dec(x_21); +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_24 = l_Lean_Elab_Tactic_elabGrindConfig(x_2, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_23); +if (lean_obj_tag(x_22) == 0) +{ +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; +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_evalApplyRfl___lambda__2___closed__5; +x_28 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_evalApplyRfl___lambda__1), 9, 3); +lean_closure_set(x_28, 0, x_25); +lean_closure_set(x_28, 1, x_27); +lean_closure_set(x_28, 2, x_15); +x_29 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_liftMetaFinishingTactic), 10, 1); +lean_closure_set(x_29, 0, x_28); +x_30 = l_Lean_Elab_Tactic_withMainContext___rarg(x_29, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_26); +return x_30; +} +else +{ +uint8_t x_31; +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); +x_31 = !lean_is_exclusive(x_24); +if (x_31 == 0) +{ +return x_24; +} +else +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_24, 0); +x_33 = lean_ctor_get(x_24, 1); +lean_inc(x_33); +lean_inc(x_32); +lean_dec(x_24); +x_34 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +return x_34; +} +} +} +else +{ +if (lean_obj_tag(x_24) == 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; +x_35 = lean_ctor_get(x_22, 0); +lean_inc(x_35); +lean_dec(x_22); +x_36 = lean_ctor_get(x_24, 0); +lean_inc(x_36); +x_37 = lean_ctor_get(x_24, 1); +lean_inc(x_37); +lean_dec(x_24); +x_38 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_evalApplyRfl___lambda__1), 9, 3); +lean_closure_set(x_38, 0, x_36); +lean_closure_set(x_38, 1, x_35); +lean_closure_set(x_38, 2, x_15); +x_39 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_liftMetaFinishingTactic), 10, 1); +lean_closure_set(x_39, 0, x_38); +x_40 = l_Lean_Elab_Tactic_withMainContext___rarg(x_39, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_37); +return x_40; +} +else +{ +uint8_t x_41; +lean_dec(x_22); +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); +x_41 = !lean_is_exclusive(x_24); +if (x_41 == 0) +{ +return x_24; +} +else +{ +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_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_45; +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_2); +x_45 = !lean_is_exclusive(x_14); +if (x_45 == 0) +{ +return x_14; +} +else +{ +lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_46 = lean_ctor_get(x_14, 0); +x_47 = lean_ctor_get(x_14, 1); +lean_inc(x_47); +lean_inc(x_46); +lean_dec(x_14); +x_48 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_48, 0, x_46); +lean_ctor_set(x_48, 1, x_47); +return x_48; +} +} +} +} +static lean_object* _init_l_Lean_Elab_Tactic_evalApplyRfl___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("grind", 5, 5); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Tactic_evalApplyRfl___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Grind___hyg_5____closed__1; +x_2 = l_Lean_Elab_Tactic_elabGrindPattern___closed__1; +x_3 = l___regBuiltin_Lean_Elab_Tactic_elabGrindPattern__1___closed__2; +x_4 = l_Lean_Elab_Tactic_evalApplyRfl___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_Tactic_evalApplyRfl___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("optConfig", 9, 9); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Tactic_evalApplyRfl___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_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Grind___hyg_5____closed__1; +x_2 = l_Lean_Elab_Tactic_elabGrindPattern___closed__1; +x_3 = l___regBuiltin_Lean_Elab_Tactic_elabGrindPattern__1___closed__2; +x_4 = l_Lean_Elab_Tactic_evalApplyRfl___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_Tactic_evalApplyRfl(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { @@ -2819,50 +3896,20 @@ return x_18; } else { -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; -x_19 = l_Lean_Elab_Tactic_evalApplyRfl___closed__7; -x_20 = 1; -lean_inc(x_8); -x_21 = l_Lean_logAt___at_Lean_Elab_Tactic_closeUsingOrAdmit___spec__2(x_1, x_19, x_20, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_1); -x_22 = lean_ctor_get(x_21, 1); -lean_inc(x_22); -lean_dec(x_21); -x_23 = l_Lean_Elab_Term_getDeclName_x3f(x_4, x_5, x_6, x_7, x_8, x_9, x_22); -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -x_25 = lean_ctor_get(x_23, 1); -lean_inc(x_25); -lean_dec(x_23); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -x_26 = l_Lean_Elab_Tactic_elabGrindConfig(x_15, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_25); -if (lean_obj_tag(x_24) == 0) -{ -if (lean_obj_tag(x_26) == 0) +lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_19 = lean_unsigned_to_nat(2u); +x_20 = l_Lean_Syntax_getArg(x_1, x_19); +x_21 = l_Lean_Syntax_isNone(x_20); +if (x_21 == 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; -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_evalApplyRfl___closed__9; -x_30 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_evalApplyRfl___lambda__1), 8, 2); -lean_closure_set(x_30, 0, x_27); -lean_closure_set(x_30, 1, x_29); -x_31 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_liftMetaFinishingTactic), 10, 1); -lean_closure_set(x_31, 0, x_30); -x_32 = l_Lean_Elab_Tactic_withMainContext___rarg(x_31, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_28); -return x_32; -} -else +uint8_t x_22; +lean_inc(x_20); +x_22 = l_Lean_Syntax_matchesNull(x_20, x_19); +if (x_22 == 0) { -uint8_t x_33; +lean_object* x_23; +lean_dec(x_20); +lean_dec(x_15); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -2871,78 +3918,32 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_33 = !lean_is_exclusive(x_26); -if (x_33 == 0) -{ -return x_26; -} -else -{ -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_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; -} -} +lean_dec(x_1); +x_23 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalApplyRfl___spec__1___rarg(x_10); +return x_23; } else { -if (lean_obj_tag(x_26) == 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_24, 0); -lean_inc(x_37); -lean_dec(x_24); -x_38 = lean_ctor_get(x_26, 0); -lean_inc(x_38); -x_39 = lean_ctor_get(x_26, 1); -lean_inc(x_39); -lean_dec(x_26); -x_40 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_evalApplyRfl___lambda__1), 8, 2); -lean_closure_set(x_40, 0, x_38); -lean_closure_set(x_40, 1, x_37); -x_41 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_liftMetaFinishingTactic), 10, 1); -lean_closure_set(x_41, 0, x_40); -x_42 = l_Lean_Elab_Tactic_withMainContext___rarg(x_41, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_39); -return x_42; +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_24 = l_Lean_Syntax_getArg(x_20, x_14); +lean_dec(x_20); +x_25 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_25, 0, x_24); +x_26 = lean_box(0); +x_27 = l_Lean_Elab_Tactic_evalApplyRfl___lambda__2(x_1, x_15, x_26, x_25, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_1); +return x_27; } -else -{ -uint8_t x_43; -lean_dec(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_2); -x_43 = !lean_is_exclusive(x_26); -if (x_43 == 0) -{ -return x_26; } else { -lean_object* x_44; lean_object* x_45; lean_object* x_46; -x_44 = lean_ctor_get(x_26, 0); -x_45 = lean_ctor_get(x_26, 1); -lean_inc(x_45); -lean_inc(x_44); -lean_dec(x_26); -x_46 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_45); -return x_46; -} -} +lean_object* x_28; lean_object* x_29; lean_object* x_30; +lean_dec(x_20); +x_28 = lean_box(0); +x_29 = lean_box(0); +x_30 = l_Lean_Elab_Tactic_evalApplyRfl___lambda__2(x_1, x_15, x_29, x_28, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_1); +return x_30; } } } @@ -2964,6 +3965,16 @@ lean_dec(x_1); return x_9; } } +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalApplyRfl___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +lean_object* x_14; +x_14 = l_Lean_Elab_Tactic_evalApplyRfl___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_3); +lean_dec(x_1); +return x_14; +} +} static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalApplyRfl__1___closed__1() { _start: { @@ -3109,10 +4120,6 @@ l_Lean_ensureNonAmbiguous___at_Lean_Elab_Tactic_elabGrindPattern___spec__9___clo lean_mark_persistent(l_Lean_ensureNonAmbiguous___at_Lean_Elab_Tactic_elabGrindPattern___spec__9___closed__5); l_Lean_ensureNonAmbiguous___at_Lean_Elab_Tactic_elabGrindPattern___spec__9___closed__6 = _init_l_Lean_ensureNonAmbiguous___at_Lean_Elab_Tactic_elabGrindPattern___spec__9___closed__6(); lean_mark_persistent(l_Lean_ensureNonAmbiguous___at_Lean_Elab_Tactic_elabGrindPattern___spec__9___closed__6); -l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_elabGrindPattern___spec__12___closed__1 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_elabGrindPattern___spec__12___closed__1(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_elabGrindPattern___spec__12___closed__1); -l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_elabGrindPattern___spec__12___closed__2 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_elabGrindPattern___spec__12___closed__2(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_elabGrindPattern___spec__12___closed__2); l_Lean_Elab_Tactic_elabGrindPattern___closed__1 = _init_l_Lean_Elab_Tactic_elabGrindPattern___closed__1(); lean_mark_persistent(l_Lean_Elab_Tactic_elabGrindPattern___closed__1); l_Lean_Elab_Tactic_elabGrindPattern___closed__2 = _init_l_Lean_Elab_Tactic_elabGrindPattern___closed__2(); @@ -3144,6 +4151,56 @@ lean_dec_ref(res); lean_mark_persistent(l_Lean_Elab_Tactic_grind___closed__1); l_Lean_Elab_Tactic_grind___closed__2 = _init_l_Lean_Elab_Tactic_grind___closed__2(); lean_mark_persistent(l_Lean_Elab_Tactic_grind___closed__2); +l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__1 = _init_l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__1(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__1); +l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__2 = _init_l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__2(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__2); +l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__3 = _init_l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__3(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__3); +l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__4 = _init_l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__4(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__4); +l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__5 = _init_l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__5(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__5); +l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__6 = _init_l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__6(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__6); +l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__7 = _init_l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__7(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__7); +l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__8 = _init_l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__8(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__8); +l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__9 = _init_l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__9(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__9); +l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__10 = _init_l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__10(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__10); +l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__11 = _init_l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__11(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__11); +l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__12 = _init_l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__12(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__12); +l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__13 = _init_l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__13(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__13); +l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__14 = _init_l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__14(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__14); +l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__15 = _init_l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__15(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__15); +l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__16 = _init_l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__16(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__16); +l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__17 = _init_l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__17(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__17); +l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__18 = _init_l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__18(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__18); +l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__19 = _init_l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__19(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__19); +l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__20 = _init_l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__20(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__20); +l_Lean_Elab_Tactic_evalApplyRfl___lambda__2___closed__1 = _init_l_Lean_Elab_Tactic_evalApplyRfl___lambda__2___closed__1(); +lean_mark_persistent(l_Lean_Elab_Tactic_evalApplyRfl___lambda__2___closed__1); +l_Lean_Elab_Tactic_evalApplyRfl___lambda__2___closed__2 = _init_l_Lean_Elab_Tactic_evalApplyRfl___lambda__2___closed__2(); +lean_mark_persistent(l_Lean_Elab_Tactic_evalApplyRfl___lambda__2___closed__2); +l_Lean_Elab_Tactic_evalApplyRfl___lambda__2___closed__3 = _init_l_Lean_Elab_Tactic_evalApplyRfl___lambda__2___closed__3(); +lean_mark_persistent(l_Lean_Elab_Tactic_evalApplyRfl___lambda__2___closed__3); +l_Lean_Elab_Tactic_evalApplyRfl___lambda__2___closed__4 = _init_l_Lean_Elab_Tactic_evalApplyRfl___lambda__2___closed__4(); +lean_mark_persistent(l_Lean_Elab_Tactic_evalApplyRfl___lambda__2___closed__4); +l_Lean_Elab_Tactic_evalApplyRfl___lambda__2___closed__5 = _init_l_Lean_Elab_Tactic_evalApplyRfl___lambda__2___closed__5(); +lean_mark_persistent(l_Lean_Elab_Tactic_evalApplyRfl___lambda__2___closed__5); l_Lean_Elab_Tactic_evalApplyRfl___closed__1 = _init_l_Lean_Elab_Tactic_evalApplyRfl___closed__1(); lean_mark_persistent(l_Lean_Elab_Tactic_evalApplyRfl___closed__1); l_Lean_Elab_Tactic_evalApplyRfl___closed__2 = _init_l_Lean_Elab_Tactic_evalApplyRfl___closed__2(); @@ -3152,16 +4209,6 @@ l_Lean_Elab_Tactic_evalApplyRfl___closed__3 = _init_l_Lean_Elab_Tactic_evalApply lean_mark_persistent(l_Lean_Elab_Tactic_evalApplyRfl___closed__3); l_Lean_Elab_Tactic_evalApplyRfl___closed__4 = _init_l_Lean_Elab_Tactic_evalApplyRfl___closed__4(); lean_mark_persistent(l_Lean_Elab_Tactic_evalApplyRfl___closed__4); -l_Lean_Elab_Tactic_evalApplyRfl___closed__5 = _init_l_Lean_Elab_Tactic_evalApplyRfl___closed__5(); -lean_mark_persistent(l_Lean_Elab_Tactic_evalApplyRfl___closed__5); -l_Lean_Elab_Tactic_evalApplyRfl___closed__6 = _init_l_Lean_Elab_Tactic_evalApplyRfl___closed__6(); -lean_mark_persistent(l_Lean_Elab_Tactic_evalApplyRfl___closed__6); -l_Lean_Elab_Tactic_evalApplyRfl___closed__7 = _init_l_Lean_Elab_Tactic_evalApplyRfl___closed__7(); -lean_mark_persistent(l_Lean_Elab_Tactic_evalApplyRfl___closed__7); -l_Lean_Elab_Tactic_evalApplyRfl___closed__8 = _init_l_Lean_Elab_Tactic_evalApplyRfl___closed__8(); -lean_mark_persistent(l_Lean_Elab_Tactic_evalApplyRfl___closed__8); -l_Lean_Elab_Tactic_evalApplyRfl___closed__9 = _init_l_Lean_Elab_Tactic_evalApplyRfl___closed__9(); -lean_mark_persistent(l_Lean_Elab_Tactic_evalApplyRfl___closed__9); l___regBuiltin_Lean_Elab_Tactic_evalApplyRfl__1___closed__1 = _init_l___regBuiltin_Lean_Elab_Tactic_evalApplyRfl__1___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Elab_Tactic_evalApplyRfl__1___closed__1); l___regBuiltin_Lean_Elab_Tactic_evalApplyRfl__1___closed__2 = _init_l___regBuiltin_Lean_Elab_Tactic_evalApplyRfl__1___closed__2(); diff --git a/stage0/stdlib/Lean/Elab/Tactic/Induction.c b/stage0/stdlib/Lean/Elab/Tactic/Induction.c index 8051c7745a75..2646b3798abd 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/Induction.c +++ b/stage0/stdlib/Lean/Elab/Tactic/Induction.c @@ -13,7 +13,7 @@ #ifdef __cplusplus extern "C" { #endif -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_applyAltStx___spec__3(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forM___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_bindingName_x21(lean_object*); @@ -85,9 +85,11 @@ LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Tactic_ElimApp_eva LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_expandInductionAlts_x3f(lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_getElimNameInfo___lambda__5___closed__1; static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_evalInduction_checkTargets___spec__1___lambda__2___closed__1; +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__14; LEAN_EXPORT uint8_t l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_isMultiAlt(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_ElimApp_mkElimApp_loop___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_generalizeVars(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__16; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_io_promise_new(lean_object*); uint8_t l_Lean_Exception_isInterrupt(lean_object*); @@ -96,7 +98,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tac LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo___boxed(lean_object**); static lean_object* l_Lean_Elab_Tactic_ElimApp_evalAlts_applyAltStx___lambda__6___closed__2; lean_object* l_Lean_Syntax_formatStxAux(lean_object*, uint8_t, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__9; 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_Array_mapMUnsafe_map___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_expandMultiAlt_x3f___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalInduction_declRange__1___closed__2; @@ -107,7 +108,6 @@ LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_Elim LEAN_EXPORT lean_object* l_Lean_Language_withAlwaysResolvedPromises___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo___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_Lean_Elab_Tactic_evalCases___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_elabCasesTargets___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__10; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_ElimApp_mkElimApp_loop(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_toString(lean_object*, uint8_t, lean_object*); lean_object* l_Lean_Meta_mkGeneralizationForbiddenSet(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -124,7 +124,6 @@ LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_eval static lean_object* l_Lean_Elab_Tactic_elabCasesTargets___lambda__1___closed__2; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_generalizeTargets___spec__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_getAltVars___boxed(lean_object*); -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__13; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_evalInduction___spec__1(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_mk_array(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalInduction_declRange__1___closed__6; @@ -171,9 +170,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tac lean_object* l_Lean_addBuiltinDeclarationRanges(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalInduction___lambda__5___boxed(lean_object**); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_getElimNameInfo___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572____closed__5; LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapMAux___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_saveAltVarsInfo___spec__10(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__15; LEAN_EXPORT lean_object* l_Lean_Elab_Term_withReuseContext___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_withAltsOfOptInductionAlts___spec__3(lean_object*); LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_applyAltStx___spec__5___at_Lean_Elab_Tactic_ElimApp_evalAlts_applyAltStx___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_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_getNumExplicitFields___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -228,12 +225,10 @@ LEAN_EXPORT uint8_t l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_a LEAN_EXPORT uint8_t l_Lean_Elab_Tactic_isHoleRHS(lean_object*); static lean_object* l_Lean_Elab_Tactic_getInductiveValFromMajor___lambda__2___closed__11; LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_checkAltNames___spec__4___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_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__4; uint8_t l_Lean_Expr_isType(lean_object*); lean_object* l_Lean_Syntax_getNumArgs(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_elabCasesTargets___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_generalizeTargets___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__11; static lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_getAltName___closed__5; LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_getElimNameInfo___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_ElimApp_evalAlts_applyAltStx___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -265,9 +260,9 @@ lean_object* l_Lean_Elab_Tactic_getNameOfIdent_x27(lean_object*); lean_object* l_Lean_Elab_Tactic_getGoals___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Tactic_getUnsolvedGoals(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_ElimApp_setMotiveArg___closed__2; +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__3; static lean_object* l_Lean_Elab_Tactic_ElimApp_evalAlts_applyAltStx___lambda__8___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_ElimApp_evalAlts_applyAltStx___lambda__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__8; LEAN_EXPORT lean_object* l_Lean_Elab_Term_withReuseContext___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_withAltsOfOptInductionAlts___spec__4___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Tactic_elabTerm(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_ElimApp_instInhabitedAlt___closed__1; @@ -287,6 +282,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_getInductiveValFromMajor(uint8_t, le static lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_getElimNameInfo___lambda__4___closed__2; static lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_getElimNameInfo___closed__1; static lean_object* l_Lean_Elab_withSaveInfoContext___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_saveAltVarsInfo___spec__2___closed__1; +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704____closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_ElimApp_evalAlts___boxed(lean_object**); static lean_object* l_Lean_Elab_Tactic_getInductiveValFromMajor___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_checkAltsOfOptInductionAlts___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*); @@ -310,6 +306,7 @@ LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Tactic_ElimApp_eva LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_generalizeVars___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MVarId_getType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalInduction___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*); +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704____closed__5; 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_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_evalInduction___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_getInductiveValFromMajor___lambda__2___closed__4; @@ -325,7 +322,7 @@ lean_object* l_Lean_Name_getPrefix(lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_elabCasesTargets___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*); static lean_object* l_Lean_Elab_Tactic_isHoleRHS___closed__2; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_generalizeVars___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___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_withAltsOfOptInductionAlts___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_withAltsOfOptInductionAlts___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_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_elabCasesTargets___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_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_getElimNameInfo___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*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalInduction__1___closed__4; @@ -346,6 +343,7 @@ lean_object* l_Lean_MVarId_withContext___at___private_Lean_Meta_SynthInstance_0_ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_ElimApp_evalAlts_applyAltStx___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_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_checkAltsOfOptInductionAlts___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___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_getAltVars(lean_object*); +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_isWildcard___boxed(lean_object*); lean_object* l_Lean_Meta_getCustomEliminator_x3f(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_FVarId_getDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -372,14 +370,12 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_withNarrowedArgTacticReuse___at_Lean_E lean_object* lean_st_mk_ref(lean_object*, lean_object*); lean_object* l_Lean_Elab_Term_synthesizeInstMVarCore(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_to_list(lean_object*); -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__14; LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___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_ElimApp_evalAlts_applyAltStx___lambda__1(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572____closed__2; static lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_getNumExplicitFields___lambda__1___closed__1; lean_object* l_Lean_Meta_abstractMVars(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_checkAltsOfOptInductionAlts___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___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*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__8___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_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_getInductiveValFromMajor___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MVarId_withContext___at_Lean_Elab_Tactic_withMainContext___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); @@ -396,7 +392,8 @@ extern lean_object* l_Lean_levelZero; LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_withAltsOfOptInductionAlts___rarg___lambda__1(lean_object*); lean_object* l_Lean_MVarId_revert(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_checkAltsOfOptInductionAlts___spec__1___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_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763_(lean_object*); static lean_object* l_Lean_Elab_Tactic_getInductiveValFromMajor___lambda__2___closed__9; extern lean_object* l_Lean_instInhabitedExpr; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalAlt___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -407,6 +404,7 @@ lean_object* l_Lean_Expr_constName_x3f(lean_object*); lean_object* l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_Term_addDotCompletionInfo___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_getAltNumFields___closed__1; static lean_object* l_Lean_Elab_Tactic_getInductiveValFromMajor___lambda__2___closed__5; +static lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_withAltsOfOptInductionAlts___rarg___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalCases___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_ElimApp_mkElimApp___closed__2; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -434,6 +432,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_elabCasesTargets___lambda__2___boxed LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_getAltsOfInductionAlts___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_withReuseContext___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_withAltsOfOptInductionAlts___spec__4(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_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__9; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_saveAltVarsInfo___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_Tactic_evalCases___lambda__4___boxed(lean_object**); lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); @@ -445,6 +444,7 @@ LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Elab_Tac LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_checkAltNames___spec__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_withoutTacticIncrementality___at_Lean_Elab_Tactic_ElimApp_evalAlts_applyAltStx___spec__7(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_generalizeVars___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_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__15; 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_object* l_Lean_Meta_saveState___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_evalInduction_checkTargets___spec__1___at_Lean_Elab_Tactic_evalInduction_checkTargets___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -453,10 +453,8 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalInduction___lambda__6___boxed(le static lean_object* l_Lean_Elab_Term_withoutTacticIncrementality___at_Lean_Elab_Tactic_ElimApp_evalAlts_applyPreTac___spec__1___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_addConstInfo___at_Lean_Elab_Tactic_ElimApp_evalAlts_applyAltStx___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_Elab_withInfoTreeContext___at_Lean_Elab_Tactic_evalTactic_eval___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_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__12; extern lean_object* l_Lean_Language_Snapshot_Diagnostics_empty; static lean_object* l_Lean_Elab_Tactic_getInductiveValFromMajor___lambda__2___closed__7; -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__6; static lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_getElimNameInfo___lambda__4___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_ElimApp_evalAlts_applyAltStx___lambda__7___boxed(lean_object**); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalInduction_checkTargets___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -481,7 +479,6 @@ static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic static lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_generalizeTargets___closed__1; static lean_object* l_Lean_Elab_Tactic_ElimApp_evalAlts_applyAltStx___lambda__8___closed__8; lean_object* l_Array_append___rarg(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__16; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_ElimApp_evalAlts(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofExpr(lean_object*); lean_object* l_Lean_Option_register___at_Lean_Elab_initFn____x40_Lean_Elab_AutoBound___hyg_6____spec__1(lean_object*, lean_object*, lean_object*, lean_object*); @@ -535,12 +532,10 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Tact LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_getAltDArrow___boxed(lean_object*); lean_object* l_Lean_Elab_Term_addLocalVarInfo(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_bindingDomain_x21(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572_(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_saveAltVarsInfo___spec__13___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_ElimApp_evalAlts_applyAltStx___lambda__8___closed__6; static lean_object* l_Lean_Elab_Tactic_ElimApp_evalAlts_applyAltStx___lambda__8___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_withReuseContext___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_withAltsOfOptInductionAlts___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*); -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_ElimApp_evalAlts_applyPreTac___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_getFType___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_elabTermForElim___closed__1; @@ -557,14 +552,18 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_evalInduction__2(lean_o 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_Tactic_Induction_0__Lean_Elab_Tactic_getAltDArrow(lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704____closed__4; LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_getUserGeneralizingFVarIds___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_x27_loop___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_checkAltNames___spec__4___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__11; LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_getOptPreTacOfOptInductionAlts(lean_object*); +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__10; LEAN_EXPORT uint8_t l_Std_Range_forIn_x27_loop___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_checkAltNames___spec__4___lambda__2(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_saveAltVarsInfo___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*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapM___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_saveAltVarsInfo___spec__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_evalInduction_checkTargets___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* l_Lean_Name_mkStr2(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__7; lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Generalize_0__Lean_Meta_generalizeCore___spec__1(size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalInduction___boxed__const__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_getUserGeneralizingFVarIds___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -573,22 +572,20 @@ LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_Elim lean_object* l_Lean_indentExpr(lean_object*); uint8_t l_Lean_BinderInfo_isExplicit(uint8_t); static lean_object* l_Lean_Elab_Tactic_evalInduction___lambda__4___closed__1; -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__7; lean_object* l_Lean_Meta_getMVarsNoDelayed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704____closed__6; extern lean_object* l_Lean_Elab_Tactic_instInhabitedTacticFinishedSnapshot; lean_object* l_Lean_Meta_isExprDefEqGuarded(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___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_EXPORT lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631_(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_getNumExplicitFields___spec__1(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___boxed(lean_object**); lean_object* l_Lean_Meta_getFVarSetToGeneralize(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_ElimApp_setMotiveArg___closed__6; -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__5; -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572____closed__6; static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__8___lambda__2___closed__1; lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Tactic_evalTactic___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_getUserGeneralizingFVarIds___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_getInductiveValFromMajor___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__13; lean_object* l_Lean_Meta_mkHasTypeButIsExpectedMsg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_getInductiveValFromMajor___lambda__2___closed__13; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_getAltNumFields___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*); @@ -624,16 +621,16 @@ LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Elab_Tac lean_object* l_Lean_Syntax_getRange_x3f(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_elabCasesTargets___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_Lean_Elab_Tactic_ElimApp_evalAlts___lambda__1___boxed(lean_object**); -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572____closed__3; LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_checkAltsOfOptInductionAlts___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_ElimApp_evalAlts___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_withoutTacticIncrementality___at_Lean_Elab_Tactic_ElimApp_evalAlts_applyPreTac___spec__1___closed__4; lean_object* l_List_iotaTR(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalInduction_declRange__1___closed__4; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__8___lambda__1(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704____closed__1; lean_object* l_Array_ofSubarray___rarg(lean_object*); uint8_t l_Lean_LocalDecl_binderInfo(lean_object*); -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___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*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___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*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_expandInductionAlts_x3f___spec__1(lean_object*, size_t, size_t); lean_object* l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalCases_declRange__1___closed__5; @@ -641,6 +638,7 @@ static lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_Tactic_ElimApp_eva LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_Tactic_elabCasesTargets___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*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_applyAltStx___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*, lean_object*, lean_object*); lean_object* lean_erase_macro_scopes(lean_object*); +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__5; LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_getBindingName___boxed(lean_object*); lean_object* l_List_reverse___rarg(lean_object*); static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_checkAltNames___spec__4___lambda__3___closed__2; @@ -648,7 +646,6 @@ static lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_g LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_elabCasesTargets___spec__4(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_applyAltStx___spec__5___at_Lean_Elab_Tactic_ElimApp_evalAlts_applyAltStx___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_array_mk(lean_object*); -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572____closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_withReuseContext___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_withAltsOfOptInductionAlts___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*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_getFType___boxed(lean_object*); static lean_object* l_Lean_Elab_Term_withoutTacticIncrementality___at_Lean_Elab_Tactic_ElimApp_evalAlts_applyPreTac___spec__1___closed__5; @@ -662,9 +659,11 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalInduction___lambda__3___boxed(le static lean_object* l_Lean_Elab_Tactic_ElimApp_mkElimApp_loop___closed__4; static lean_object* l_Lean_Elab_Tactic_ElimApp_evalAlts_applyAltStx___lambda__8___closed__7; static lean_object* l_Lean_Elab_Tactic_ElimApp_evalAlts_applyAltStx___lambda__6___closed__4; +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__1; lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_evalGeneralize___spec__9(size_t, size_t, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_withAltsOfOptInductionAlts___rarg___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_withAltsOfOptInductionAlts___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_array_uget(lean_object*, size_t); +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__6; LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__4___boxed(lean_object**); lean_object* l_Lean_Expr_fvar___override(lean_object*); size_t lean_array_size(lean_object*); @@ -721,6 +720,7 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Tact static lean_object* l_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo___lambda__2___closed__10; uint8_t lean_nat_dec_le(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_getInductiveValFromMajor___lambda__2___closed__14; +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704_(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*); static lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_withAltsOfOptInductionAlts___rarg___lambda__1___closed__2; @@ -736,8 +736,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalCases(lean_object*, lean_object* lean_object* lean_nat_add(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_withoutTacticIncrementality___at_Lean_Elab_Tactic_ElimApp_evalAlts_applyPreTac___spec__1___lambda__2___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_generalizeVars___spec__1___boxed(lean_object**); -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572____closed__4; -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__2; uint8_t l_Lean_Exception_isRuntime(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Tactic_ElimApp_mkElimApp___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); 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*); @@ -745,6 +743,7 @@ LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_Elim LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_getInductiveValFromMajor___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_ElimApp_evalAlts_goWithInfo___lambda__1___boxed(lean_object**); static lean_object* l_Lean_Elab_Tactic_getInductiveValFromMajor___lambda__2___closed__8; +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__4; LEAN_EXPORT lean_object* l_Array_mapFinIdxM_map___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isFVar(lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_evalInduction_checkTargets___spec__1___closed__1; @@ -759,7 +758,6 @@ uint64_t l___private_Lean_Meta_Basic_0__Lean_Meta_Config_toKey(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalInduction_declRange__1___closed__5; static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_checkAltNames___spec__4___closed__2; static lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_getAltNumFields___lambda__1___closed__2; -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__3; lean_object* lean_array_uset(lean_object*, size_t, lean_object*); lean_object* lean_dbg_trace(lean_object*, lean_object*); 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*); @@ -776,15 +774,18 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_evalCases_ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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_x27_loop___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_checkAltNames___spec__4___lambda__1___closed__5; LEAN_EXPORT uint8_t l_Lean_Elab_Term_withoutTacticIncrementality___at_Lean_Elab_Tactic_ElimApp_evalAlts_applyPreTac___spec__1___lambda__1(uint8_t, lean_object*); +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704____closed__3; static lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_withAltsOfOptInductionAlts___rarg___closed__1; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_elabCasesTargets___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__8; LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_getElimNameInfo___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___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_getUserGeneralizingFVarIds(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo___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___regBuiltin_Lean_Elab_Tactic_evalCases_declRange__1___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalAlt___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_applyAltStx___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__12; static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_evalInduction_checkTargets___spec__1___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_getFType___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo___lambda__3___boxed(lean_object**); @@ -13864,7 +13865,7 @@ lean_inc(x_33); x_34 = lean_ctor_get(x_32, 1); lean_inc(x_34); lean_dec(x_32); -x_35 = l_Lean_Elab_Tactic_ElimApp_evalAlts_applyAltStx___lambda__8(x_30, x_26, x_28, x_10, x_24, x_4, x_25, x_6, x_7, x_2, x_3, x_8, x_9, x_5, x_33, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_34); +x_35 = l_Lean_Elab_Tactic_ElimApp_evalAlts_applyAltStx___lambda__8(x_30, x_26, x_28, x_10, x_24, x_3, x_25, x_5, x_6, x_2, x_8, x_7, x_9, x_4, x_33, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_34); return x_35; } else @@ -13885,8 +13886,8 @@ lean_dec(x_13); lean_dec(x_12); lean_dec(x_10); lean_dec(x_9); -lean_dec(x_7); -lean_dec(x_5); +lean_dec(x_8); +lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); x_36 = !lean_is_exclusive(x_32); @@ -13913,7 +13914,7 @@ else { lean_object* x_40; lean_inc(x_28); -x_40 = l_Lean_Elab_Tactic_ElimApp_evalAlts_applyAltStx___lambda__8(x_30, x_26, x_28, x_10, x_24, x_4, x_25, x_6, x_7, x_2, x_3, x_8, x_9, x_5, x_28, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_29); +x_40 = l_Lean_Elab_Tactic_ElimApp_evalAlts_applyAltStx___lambda__8(x_30, x_26, x_28, x_10, x_24, x_3, x_25, x_5, x_6, x_2, x_8, x_7, x_9, x_4, x_28, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_29); return x_40; } } @@ -13933,8 +13934,8 @@ lean_dec(x_13); lean_dec(x_12); lean_dec(x_10); lean_dec(x_9); -lean_dec(x_7); -lean_dec(x_5); +lean_dec(x_8); +lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); x_41 = !lean_is_exclusive(x_27); @@ -14042,7 +14043,7 @@ lean_inc(x_70); x_71 = lean_ctor_get(x_69, 1); lean_inc(x_71); lean_dec(x_69); -x_72 = l_Lean_Elab_Tactic_ElimApp_evalAlts_applyAltStx___lambda__8(x_67, x_63, x_65, x_10, x_61, x_4, x_62, x_6, x_7, x_2, x_3, x_8, x_9, x_5, x_70, x_12, x_13, x_14, x_15, x_16, x_17, x_60, x_19, x_71); +x_72 = l_Lean_Elab_Tactic_ElimApp_evalAlts_applyAltStx___lambda__8(x_67, x_63, x_65, x_10, x_61, x_3, x_62, x_5, x_6, x_2, x_8, x_7, x_9, x_4, x_70, x_12, x_13, x_14, x_15, x_16, x_17, x_60, x_19, x_71); return x_72; } else @@ -14063,8 +14064,8 @@ lean_dec(x_13); lean_dec(x_12); lean_dec(x_10); lean_dec(x_9); -lean_dec(x_7); -lean_dec(x_5); +lean_dec(x_8); +lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); x_73 = lean_ctor_get(x_69, 0); @@ -14093,7 +14094,7 @@ else { lean_object* x_77; lean_inc(x_65); -x_77 = l_Lean_Elab_Tactic_ElimApp_evalAlts_applyAltStx___lambda__8(x_67, x_63, x_65, x_10, x_61, x_4, x_62, x_6, x_7, x_2, x_3, x_8, x_9, x_5, x_65, x_12, x_13, x_14, x_15, x_16, x_17, x_60, x_19, x_66); +x_77 = l_Lean_Elab_Tactic_ElimApp_evalAlts_applyAltStx___lambda__8(x_67, x_63, x_65, x_10, x_61, x_3, x_62, x_5, x_6, x_2, x_8, x_7, x_9, x_4, x_65, x_12, x_13, x_14, x_15, x_16, x_17, x_60, x_19, x_66); return x_77; } } @@ -14113,8 +14114,8 @@ lean_dec(x_13); lean_dec(x_12); lean_dec(x_10); lean_dec(x_9); -lean_dec(x_7); -lean_dec(x_5); +lean_dec(x_8); +lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); x_78 = lean_ctor_get(x_64, 0); @@ -14493,8 +14494,8 @@ lean_object* x_20 = _args[19]; { lean_object* x_21; x_21 = l_Lean_Elab_Tactic_ElimApp_evalAlts_applyAltStx(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20); -lean_dec(x_8); -lean_dec(x_6); +lean_dec(x_7); +lean_dec(x_5); lean_dec(x_2); lean_dec(x_1); return x_21; @@ -14625,8 +14626,8 @@ lean_dec(x_18); lean_dec(x_17); lean_dec(x_11); lean_dec(x_9); -lean_dec(x_7); -lean_dec(x_5); +lean_dec(x_8); +lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); x_27 = lean_alloc_ctor(0, 2, 0); @@ -14649,8 +14650,8 @@ lean_inc(x_18); lean_inc(x_17); lean_inc(x_11); lean_inc(x_9); -lean_inc(x_7); -lean_inc(x_5); +lean_inc(x_8); +lean_inc(x_6); lean_inc(x_4); lean_inc(x_3); x_29 = l_Lean_Elab_Tactic_ElimApp_evalAlts_applyAltStx(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_11, x_28, x_17, x_18, x_19, x_20, x_21, x_22, x_23, x_24, x_25); @@ -14681,8 +14682,8 @@ lean_dec(x_18); lean_dec(x_17); lean_dec(x_11); lean_dec(x_9); -lean_dec(x_7); -lean_dec(x_5); +lean_dec(x_8); +lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); x_35 = !lean_is_exclusive(x_29); @@ -14743,15 +14744,16 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19, lean_object* x_20, lean_object* x_21, lean_object* x_22, lean_object* x_23) { +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19, lean_object* x_20, lean_object* x_21, lean_object* x_22, lean_object* x_23, lean_object* x_24) { _start: { -lean_object* x_24; uint8_t x_25; -x_24 = lean_ctor_get(x_10, 1); -x_25 = lean_nat_dec_lt(x_12, x_24); -if (x_25 == 0) +lean_object* x_25; uint8_t x_26; +x_25 = lean_ctor_get(x_11, 1); +x_26 = lean_nat_dec_lt(x_13, x_25); +if (x_26 == 0) { -lean_object* x_26; +lean_object* x_27; +lean_dec(x_23); lean_dec(x_22); lean_dec(x_21); lean_dec(x_20); @@ -14759,56 +14761,56 @@ lean_dec(x_19); lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_12); +lean_dec(x_13); +lean_dec(x_9); lean_dec(x_7); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); -x_26 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_26, 0, x_11); -lean_ctor_set(x_26, 1, x_23); -return x_26; +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_12); +lean_ctor_set(x_27, 1, x_24); +return x_27; } else { -lean_object* x_27; lean_object* x_28; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_48; lean_object* x_49; lean_object* x_50; -x_34 = lean_array_fget(x_3, x_12); -x_35 = l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_getAltName(x_34); -lean_inc(x_35); -x_48 = lean_alloc_closure((void*)(l_Std_Range_forIn_x27_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__4___lambda__1___boxed), 2, 1); -lean_closure_set(x_48, 0, x_35); -x_49 = lean_unsigned_to_nat(0u); -x_50 = l_Array_findFinIdx_x3f_loop___rarg(x_48, x_11, x_49); -if (lean_obj_tag(x_50) == 0) -{ -uint8_t x_51; -x_51 = l_Array_isEmpty___rarg(x_11); -if (x_51 == 0) +lean_object* x_28; lean_object* x_29; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_35 = lean_array_fget(x_9, x_13); +x_36 = l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_getAltName(x_35); +lean_inc(x_36); +x_49 = lean_alloc_closure((void*)(l_Std_Range_forIn_x27_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__4___lambda__1___boxed), 2, 1); +lean_closure_set(x_49, 0, x_36); +x_50 = lean_unsigned_to_nat(0u); +x_51 = l_Array_findFinIdx_x3f_loop___rarg(x_49, x_12, x_50); +if (lean_obj_tag(x_51) == 0) { uint8_t x_52; -x_52 = l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_isWildcard(x_34); +x_52 = l_Array_isEmpty___rarg(x_12); if (x_52 == 0) { -lean_object* x_53; +uint8_t x_53; +x_53 = l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_isWildcard(x_35); +if (x_53 == 0) +{ +lean_object* x_54; +lean_dec(x_13); lean_dec(x_12); -lean_dec(x_11); +lean_dec(x_9); lean_dec(x_7); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); -x_53 = lean_box(0); -x_36 = x_53; -goto block_47; +x_54 = lean_box(0); +x_37 = x_54; +goto block_48; } else { -lean_object* x_54; size_t x_55; size_t x_56; lean_object* x_57; lean_object* x_58; -lean_dec(x_35); -x_54 = lean_box(0); -x_55 = lean_array_size(x_11); -x_56 = 0; -x_57 = lean_box(0); +lean_object* x_55; size_t x_56; size_t x_57; lean_object* x_58; lean_object* x_59; +lean_dec(x_36); +x_55 = lean_box(0); +x_56 = lean_array_size(x_12); +x_57 = 0; +x_58 = lean_box(0); +lean_inc(x_23); lean_inc(x_22); lean_inc(x_21); lean_inc(x_20); @@ -14816,28 +14818,28 @@ 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_13); +lean_inc(x_9); lean_inc(x_7); lean_inc(x_5); lean_inc(x_4); -lean_inc(x_3); -x_58 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_12, x_11, x_34, x_54, x_11, x_55, x_56, x_57, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_23); -lean_dec(x_11); -if (lean_obj_tag(x_58) == 0) +x_59 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__3(x_1, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_13, x_12, x_35, x_55, x_12, x_56, x_57, x_58, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_23, x_24); +lean_dec(x_12); +if (lean_obj_tag(x_59) == 0) { -lean_object* x_59; lean_object* x_60; -x_59 = lean_ctor_get(x_58, 1); -lean_inc(x_59); -lean_dec(x_58); -x_60 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__4___closed__3; -x_27 = x_60; -x_28 = x_59; -goto block_33; +lean_object* x_60; lean_object* x_61; +x_60 = lean_ctor_get(x_59, 1); +lean_inc(x_60); +lean_dec(x_59); +x_61 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__4___closed__3; +x_28 = x_61; +x_29 = x_60; +goto block_34; } else { -uint8_t x_61; +uint8_t x_62; +lean_dec(x_23); lean_dec(x_22); lean_dec(x_21); lean_dec(x_20); @@ -14845,57 +14847,57 @@ lean_dec(x_19); lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_12); +lean_dec(x_13); +lean_dec(x_9); lean_dec(x_7); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); -x_61 = !lean_is_exclusive(x_58); -if (x_61 == 0) +x_62 = !lean_is_exclusive(x_59); +if (x_62 == 0) { -return x_58; +return x_59; } else { -lean_object* x_62; lean_object* x_63; lean_object* x_64; -x_62 = lean_ctor_get(x_58, 0); -x_63 = lean_ctor_get(x_58, 1); +lean_object* x_63; lean_object* x_64; lean_object* x_65; +x_63 = lean_ctor_get(x_59, 0); +x_64 = lean_ctor_get(x_59, 1); +lean_inc(x_64); lean_inc(x_63); -lean_inc(x_62); -lean_dec(x_58); -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_59); +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_65; +lean_object* x_66; +lean_dec(x_13); lean_dec(x_12); -lean_dec(x_11); +lean_dec(x_9); lean_dec(x_7); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); -x_65 = lean_box(0); -x_36 = x_65; -goto block_47; +x_66 = lean_box(0); +x_37 = x_66; +goto block_48; } } else { -uint8_t x_66; -lean_dec(x_35); -x_66 = !lean_is_exclusive(x_50); -if (x_66 == 0) +uint8_t x_67; +lean_dec(x_36); +x_67 = !lean_is_exclusive(x_51); +if (x_67 == 0) { -lean_object* x_67; lean_object* x_68; lean_object* x_69; -x_67 = lean_ctor_get(x_50, 0); -x_68 = lean_array_fget(x_11, x_67); +lean_object* x_68; lean_object* x_69; lean_object* x_70; +x_68 = lean_ctor_get(x_51, 0); +x_69 = lean_array_fget(x_12, x_68); +lean_inc(x_23); lean_inc(x_22); lean_inc(x_21); lean_inc(x_20); @@ -14903,30 +14905,30 @@ 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_13); +lean_inc(x_9); lean_inc(x_7); lean_inc(x_5); lean_inc(x_4); -lean_inc(x_3); -x_69 = l_Lean_Elab_Tactic_ElimApp_evalAlts_applyAltStx(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_12, x_34, x_68, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_23); -if (lean_obj_tag(x_69) == 0) +x_70 = l_Lean_Elab_Tactic_ElimApp_evalAlts_applyAltStx(x_1, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_13, x_35, x_69, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_23, x_24); +if (lean_obj_tag(x_70) == 0) { -lean_object* x_70; lean_object* x_71; -x_70 = lean_ctor_get(x_69, 1); -lean_inc(x_70); -lean_dec(x_69); -x_71 = l_Array_eraseIdx___rarg(x_11, x_67, lean_box(0)); -lean_ctor_set(x_50, 0, x_71); -x_27 = x_50; -x_28 = x_70; -goto block_33; +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_Array_eraseIdx___rarg(x_12, x_68, lean_box(0)); +lean_ctor_set(x_51, 0, x_72); +x_28 = x_51; +x_29 = x_71; +goto block_34; } else { -uint8_t x_72; -lean_free_object(x_50); -lean_dec(x_67); +uint8_t x_73; +lean_free_object(x_51); +lean_dec(x_68); +lean_dec(x_23); lean_dec(x_22); lean_dec(x_21); lean_dec(x_20); @@ -14934,40 +14936,40 @@ lean_dec(x_19); lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); -lean_dec(x_15); +lean_dec(x_13); lean_dec(x_12); -lean_dec(x_11); +lean_dec(x_9); lean_dec(x_7); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); -x_72 = !lean_is_exclusive(x_69); -if (x_72 == 0) +x_73 = !lean_is_exclusive(x_70); +if (x_73 == 0) { -return x_69; +return x_70; } else { -lean_object* x_73; lean_object* x_74; lean_object* x_75; -x_73 = lean_ctor_get(x_69, 0); -x_74 = lean_ctor_get(x_69, 1); +lean_object* x_74; lean_object* x_75; lean_object* x_76; +x_74 = lean_ctor_get(x_70, 0); +x_75 = lean_ctor_get(x_70, 1); +lean_inc(x_75); lean_inc(x_74); -lean_inc(x_73); -lean_dec(x_69); -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_70); +x_76 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_76, 0, x_74); +lean_ctor_set(x_76, 1, x_75); +return x_76; } } } else { -lean_object* x_76; lean_object* x_77; lean_object* x_78; -x_76 = lean_ctor_get(x_50, 0); -lean_inc(x_76); -lean_dec(x_50); -x_77 = lean_array_fget(x_11, x_76); +lean_object* x_77; lean_object* x_78; lean_object* x_79; +x_77 = lean_ctor_get(x_51, 0); +lean_inc(x_77); +lean_dec(x_51); +x_78 = lean_array_fget(x_12, x_77); +lean_inc(x_23); lean_inc(x_22); lean_inc(x_21); lean_inc(x_20); @@ -14975,30 +14977,30 @@ 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_13); +lean_inc(x_9); lean_inc(x_7); lean_inc(x_5); lean_inc(x_4); -lean_inc(x_3); -x_78 = l_Lean_Elab_Tactic_ElimApp_evalAlts_applyAltStx(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_12, x_34, x_77, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_23); -if (lean_obj_tag(x_78) == 0) +x_79 = l_Lean_Elab_Tactic_ElimApp_evalAlts_applyAltStx(x_1, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_13, x_35, x_78, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_23, x_24); +if (lean_obj_tag(x_79) == 0) { -lean_object* x_79; lean_object* x_80; lean_object* x_81; -x_79 = lean_ctor_get(x_78, 1); -lean_inc(x_79); -lean_dec(x_78); -x_80 = l_Array_eraseIdx___rarg(x_11, x_76, lean_box(0)); -x_81 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_81, 0, x_80); -x_27 = x_81; -x_28 = x_79; -goto block_33; +lean_object* x_80; lean_object* x_81; lean_object* x_82; +x_80 = lean_ctor_get(x_79, 1); +lean_inc(x_80); +lean_dec(x_79); +x_81 = l_Array_eraseIdx___rarg(x_12, x_77, lean_box(0)); +x_82 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_82, 0, x_81); +x_28 = x_82; +x_29 = x_80; +goto block_34; } else { -lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; -lean_dec(x_76); +lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; +lean_dec(x_77); +lean_dec(x_23); lean_dec(x_22); lean_dec(x_21); lean_dec(x_20); @@ -15006,91 +15008,90 @@ lean_dec(x_19); lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); -lean_dec(x_15); +lean_dec(x_13); lean_dec(x_12); -lean_dec(x_11); +lean_dec(x_9); lean_dec(x_7); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); -x_82 = lean_ctor_get(x_78, 0); -lean_inc(x_82); -x_83 = lean_ctor_get(x_78, 1); +x_83 = lean_ctor_get(x_79, 0); lean_inc(x_83); -if (lean_is_exclusive(x_78)) { - lean_ctor_release(x_78, 0); - lean_ctor_release(x_78, 1); - x_84 = x_78; +x_84 = lean_ctor_get(x_79, 1); +lean_inc(x_84); +if (lean_is_exclusive(x_79)) { + lean_ctor_release(x_79, 0); + lean_ctor_release(x_79, 1); + x_85 = x_79; } else { - lean_dec_ref(x_78); - x_84 = lean_box(0); + lean_dec_ref(x_79); + x_85 = lean_box(0); } -if (lean_is_scalar(x_84)) { - x_85 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_85)) { + x_86 = lean_alloc_ctor(1, 2, 0); } else { - x_85 = x_84; + x_86 = x_85; } -lean_ctor_set(x_85, 0, x_82); -lean_ctor_set(x_85, 1, x_83); -return x_85; +lean_ctor_set(x_86, 0, x_83); +lean_ctor_set(x_86, 1, x_84); +return x_86; } } } -block_33: +block_34: { -lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_29 = lean_ctor_get(x_27, 0); -lean_inc(x_29); -lean_dec(x_27); -x_30 = lean_ctor_get(x_10, 2); -x_31 = lean_nat_add(x_12, x_30); -lean_dec(x_12); -x_11 = x_29; -x_12 = x_31; -x_13 = lean_box(0); +lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_30 = lean_ctor_get(x_28, 0); +lean_inc(x_30); +lean_dec(x_28); +x_31 = lean_ctor_get(x_11, 2); +x_32 = lean_nat_add(x_13, x_31); +lean_dec(x_13); +x_12 = x_30; +x_13 = x_32; x_14 = lean_box(0); -x_23 = x_28; +x_15 = lean_box(0); +x_24 = x_29; goto _start; } -block_47: +block_48: { -lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; uint8_t x_43; -lean_dec(x_36); -x_37 = l_Lean_MessageData_ofName(x_35); -x_38 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__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_Lean_Elab_Tactic_ElimApp_mkElimApp_loop___closed__4; -x_41 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_41, 0, x_39); -lean_ctor_set(x_41, 1, x_40); -x_42 = l_Lean_throwErrorAt___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__1(x_34, x_41, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_23); -lean_dec(x_22); +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_dec(x_37); +x_38 = l_Lean_MessageData_ofName(x_36); +x_39 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__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_Lean_Elab_Tactic_ElimApp_mkElimApp_loop___closed__4; +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_throwErrorAt___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__1(x_35, x_42, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_23, x_24); +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_16); -lean_dec(x_15); -lean_dec(x_34); -x_43 = !lean_is_exclusive(x_42); -if (x_43 == 0) +lean_dec(x_35); +x_44 = !lean_is_exclusive(x_43); +if (x_44 == 0) { -return x_42; +return x_43; } else { -lean_object* x_44; lean_object* x_45; lean_object* x_46; -x_44 = lean_ctor_get(x_42, 0); -x_45 = lean_ctor_get(x_42, 1); +lean_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_inc(x_44); -lean_dec(x_42); -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_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; } } } @@ -15407,7 +15408,7 @@ x_2 = lean_box_usize(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___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, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17) { +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__8___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_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; size_t x_19; size_t x_20; lean_object* x_21; @@ -15463,28 +15464,23 @@ lean_inc(x_9); x_30 = l_Lean_Elab_Tactic_ElimApp_evalAlts_applyPreTac(x_4, x_22, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_29); if (lean_obj_tag(x_30) == 0) { -uint8_t x_31; -x_31 = !lean_is_exclusive(x_30); -if (x_31 == 0) -{ -lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; -x_32 = lean_ctor_get(x_30, 0); -x_33 = lean_ctor_get(x_30, 1); -x_34 = lean_unsigned_to_nat(0u); -x_35 = lean_nat_dec_lt(x_34, x_5); -if (x_35 == 0) +if (x_5 == 0) { -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_free_object(x_30); +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; uint8_t x_38; lean_dec(x_6); -x_36 = l_Lean_Elab_Tactic_getGoals___rarg(x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_33); -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_List_appendTR___rarg(x_37, x_32); -x_40 = l_Lean_Elab_Tactic_setGoals(x_39, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_38); +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_Tactic_getGoals___rarg(x_10, x_11, x_12, x_13, x_14, x_15, 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 = l_List_appendTR___rarg(x_34, x_31); +x_37 = l_Lean_Elab_Tactic_setGoals(x_36, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_35); lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); @@ -15493,33 +15489,39 @@ lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); -x_41 = !lean_is_exclusive(x_40); -if (x_41 == 0) +x_38 = !lean_is_exclusive(x_37); +if (x_38 == 0) { -lean_object* x_42; lean_object* x_43; -x_42 = lean_ctor_get(x_40, 0); -lean_dec(x_42); -x_43 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__8___lambda__2___closed__1; -lean_ctor_set(x_40, 0, x_43); -return x_40; +lean_object* x_39; lean_object* x_40; +x_39 = lean_ctor_get(x_37, 0); +lean_dec(x_39); +x_40 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__8___lambda__2___closed__1; +lean_ctor_set(x_37, 0, x_40); +return x_37; } 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 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__8___lambda__2___closed__1; -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_41; lean_object* x_42; lean_object* x_43; +x_41 = lean_ctor_get(x_37, 1); +lean_inc(x_41); +lean_dec(x_37); +x_42 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__8___lambda__2___closed__1; +x_43 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_43, 0, x_42); +lean_ctor_set(x_43, 1, x_41); +return x_43; } } else { -uint8_t x_47; -x_47 = l_List_isEmpty___rarg(x_32); +uint8_t x_44; +x_44 = !lean_is_exclusive(x_30); +if (x_44 == 0) +{ +lean_object* x_45; lean_object* x_46; uint8_t x_47; +x_45 = lean_ctor_get(x_30, 0); +x_46 = lean_ctor_get(x_30, 1); +x_47 = l_List_isEmpty___rarg(x_45); if (x_47 == 0) { lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; uint8_t x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; @@ -15535,11 +15537,11 @@ lean_ctor_set(x_52, 0, x_50); lean_ctor_set(x_52, 1, x_51); x_53 = 2; lean_inc(x_15); -x_54 = l_Lean_log___at_Lean_Elab_Tactic_closeUsingOrAdmit___spec__3(x_52, x_53, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_33); +x_54 = l_Lean_log___at_Lean_Elab_Tactic_closeUsingOrAdmit___spec__3(x_52, x_53, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_46); x_55 = lean_ctor_get(x_54, 1); lean_inc(x_55); lean_dec(x_54); -x_56 = l_List_forM___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__7(x_32, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_55); +x_56 = l_List_forM___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__7(x_45, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_55); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); @@ -15596,7 +15598,7 @@ return x_66; else { lean_object* x_67; -lean_dec(x_32); +lean_dec(x_45); lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); @@ -15611,135 +15613,89 @@ lean_ctor_set(x_30, 0, x_67); return x_30; } } -} else { -lean_object* x_68; lean_object* x_69; lean_object* x_70; uint8_t x_71; +lean_object* x_68; lean_object* x_69; uint8_t x_70; x_68 = lean_ctor_get(x_30, 0); x_69 = lean_ctor_get(x_30, 1); lean_inc(x_69); lean_inc(x_68); lean_dec(x_30); -x_70 = lean_unsigned_to_nat(0u); -x_71 = lean_nat_dec_lt(x_70, x_5); -if (x_71 == 0) -{ -lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; -lean_dec(x_6); -x_72 = l_Lean_Elab_Tactic_getGoals___rarg(x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_69); -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_List_appendTR___rarg(x_73, x_68); -x_76 = l_Lean_Elab_Tactic_setGoals(x_75, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_74); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -x_77 = lean_ctor_get(x_76, 1); -lean_inc(x_77); -if (lean_is_exclusive(x_76)) { - lean_ctor_release(x_76, 0); - lean_ctor_release(x_76, 1); - x_78 = x_76; -} else { - lean_dec_ref(x_76); - x_78 = lean_box(0); -} -x_79 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__8___lambda__2___closed__1; -if (lean_is_scalar(x_78)) { - x_80 = lean_alloc_ctor(0, 2, 0); -} else { - x_80 = x_78; -} -lean_ctor_set(x_80, 0, x_79); -lean_ctor_set(x_80, 1, x_77); -return x_80; -} -else -{ -uint8_t x_81; -x_81 = l_List_isEmpty___rarg(x_68); -if (x_81 == 0) +x_70 = l_List_isEmpty___rarg(x_68); +if (x_70 == 0) { -lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; uint8_t x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; -x_82 = l_Lean_MessageData_ofName(x_6); -x_83 = l_Lean_Elab_Tactic_ElimApp_evalAlts_applyAltStx___lambda__6___closed__2; -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_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__8___lambda__2___closed__3; -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 = 2; +lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; uint8_t x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; +x_71 = l_Lean_MessageData_ofName(x_6); +x_72 = l_Lean_Elab_Tactic_ElimApp_evalAlts_applyAltStx___lambda__6___closed__2; +x_73 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_73, 0, x_72); +lean_ctor_set(x_73, 1, x_71); +x_74 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__8___lambda__2___closed__3; +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 = 2; lean_inc(x_15); -x_88 = l_Lean_log___at_Lean_Elab_Tactic_closeUsingOrAdmit___spec__3(x_86, x_87, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_69); -x_89 = lean_ctor_get(x_88, 1); -lean_inc(x_89); -lean_dec(x_88); -x_90 = l_List_forM___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__7(x_68, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_89); +x_77 = l_Lean_log___at_Lean_Elab_Tactic_closeUsingOrAdmit___spec__3(x_75, x_76, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_69); +x_78 = lean_ctor_get(x_77, 1); +lean_inc(x_78); +lean_dec(x_77); +x_79 = l_List_forM___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__7(x_68, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_78); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); -if (lean_obj_tag(x_90) == 0) +if (lean_obj_tag(x_79) == 0) { -lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; -x_91 = lean_ctor_get(x_90, 1); -lean_inc(x_91); -if (lean_is_exclusive(x_90)) { - lean_ctor_release(x_90, 0); - lean_ctor_release(x_90, 1); - x_92 = x_90; +lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; +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_90); - x_92 = lean_box(0); + lean_dec_ref(x_79); + x_81 = lean_box(0); } -x_93 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__8___lambda__2___closed__1; -if (lean_is_scalar(x_92)) { - x_94 = lean_alloc_ctor(0, 2, 0); +x_82 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__8___lambda__2___closed__1; +if (lean_is_scalar(x_81)) { + x_83 = lean_alloc_ctor(0, 2, 0); } else { - x_94 = x_92; + x_83 = x_81; } -lean_ctor_set(x_94, 0, x_93); -lean_ctor_set(x_94, 1, x_91); -return x_94; +lean_ctor_set(x_83, 0, x_82); +lean_ctor_set(x_83, 1, x_80); +return x_83; } else { -lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; -x_95 = lean_ctor_get(x_90, 0); -lean_inc(x_95); -x_96 = lean_ctor_get(x_90, 1); -lean_inc(x_96); -if (lean_is_exclusive(x_90)) { - lean_ctor_release(x_90, 0); - lean_ctor_release(x_90, 1); - x_97 = x_90; +lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; +x_84 = lean_ctor_get(x_79, 0); +lean_inc(x_84); +x_85 = lean_ctor_get(x_79, 1); +lean_inc(x_85); +if (lean_is_exclusive(x_79)) { + lean_ctor_release(x_79, 0); + lean_ctor_release(x_79, 1); + x_86 = x_79; } else { - lean_dec_ref(x_90); - x_97 = lean_box(0); + lean_dec_ref(x_79); + x_86 = lean_box(0); } -if (lean_is_scalar(x_97)) { - x_98 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_86)) { + x_87 = lean_alloc_ctor(1, 2, 0); } else { - x_98 = x_97; + x_87 = x_86; } -lean_ctor_set(x_98, 0, x_95); -lean_ctor_set(x_98, 1, x_96); -return x_98; +lean_ctor_set(x_87, 0, x_84); +lean_ctor_set(x_87, 1, x_85); +return x_87; } } else { -lean_object* x_99; lean_object* x_100; +lean_object* x_88; lean_object* x_89; lean_dec(x_68); lean_dec(x_16); lean_dec(x_15); @@ -15750,18 +15706,18 @@ lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_6); -x_99 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__8___lambda__2___closed__1; -x_100 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_100, 0, x_99); -lean_ctor_set(x_100, 1, x_69); -return x_100; +x_88 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__8___lambda__2___closed__1; +x_89 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_89, 0, x_88); +lean_ctor_set(x_89, 1, x_69); +return x_89; } } } } else { -uint8_t x_101; +uint8_t x_90; lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); @@ -15771,29 +15727,29 @@ lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_6); -x_101 = !lean_is_exclusive(x_30); -if (x_101 == 0) +x_90 = !lean_is_exclusive(x_30); +if (x_90 == 0) { return x_30; } else { -lean_object* x_102; lean_object* x_103; lean_object* x_104; -x_102 = lean_ctor_get(x_30, 0); -x_103 = lean_ctor_get(x_30, 1); -lean_inc(x_103); -lean_inc(x_102); +lean_object* x_91; lean_object* x_92; lean_object* x_93; +x_91 = lean_ctor_get(x_30, 0); +x_92 = lean_ctor_get(x_30, 1); +lean_inc(x_92); +lean_inc(x_91); lean_dec(x_30); -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; +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 { -uint8_t x_105; +uint8_t x_94; lean_dec(x_22); lean_dec(x_16); lean_dec(x_15); @@ -15804,29 +15760,29 @@ lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_6); -x_105 = !lean_is_exclusive(x_28); -if (x_105 == 0) +x_94 = !lean_is_exclusive(x_28); +if (x_94 == 0) { return x_28; } else { -lean_object* x_106; lean_object* x_107; lean_object* x_108; -x_106 = lean_ctor_get(x_28, 0); -x_107 = lean_ctor_get(x_28, 1); -lean_inc(x_107); -lean_inc(x_106); +lean_object* x_95; lean_object* x_96; lean_object* x_97; +x_95 = lean_ctor_get(x_28, 0); +x_96 = lean_ctor_get(x_28, 1); +lean_inc(x_96); +lean_inc(x_95); lean_dec(x_28); -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_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; } } } else { -uint8_t x_109; +uint8_t x_98; lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); @@ -15838,36 +15794,35 @@ lean_dec(x_9); lean_dec(x_6); lean_dec(x_3); lean_dec(x_2); -x_109 = !lean_is_exclusive(x_21); -if (x_109 == 0) +x_98 = !lean_is_exclusive(x_21); +if (x_98 == 0) { return x_21; } else { -lean_object* x_110; lean_object* x_111; lean_object* x_112; -x_110 = lean_ctor_get(x_21, 0); -x_111 = lean_ctor_get(x_21, 1); -lean_inc(x_111); -lean_inc(x_110); +lean_object* x_99; lean_object* x_100; lean_object* x_101; +x_99 = lean_ctor_get(x_21, 0); +x_100 = lean_ctor_get(x_21, 1); +lean_inc(x_100); +lean_inc(x_99); lean_dec(x_21); -x_112 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_112, 0, x_110); -lean_ctor_set(x_112, 1, x_111); -return x_112; +x_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_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___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, 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_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___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, uint8_t x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, size_t x_11, size_t x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19, lean_object* x_20, lean_object* x_21, lean_object* x_22) { _start: { -uint8_t x_24; -x_24 = lean_usize_dec_lt(x_13, x_12); -if (x_24 == 0) +uint8_t x_23; +x_23 = lean_usize_dec_lt(x_12, x_11); +if (x_23 == 0) { -lean_object* x_25; -lean_dec(x_22); +lean_object* x_24; lean_dec(x_21); lean_dec(x_20); lean_dec(x_19); @@ -15875,103 +15830,103 @@ lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); -lean_dec(x_7); -lean_dec(x_5); +lean_dec(x_14); +lean_dec(x_6); lean_dec(x_4); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_14); -lean_ctor_set(x_25, 1, x_23); -return x_25; +lean_dec(x_3); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_13); +lean_ctor_set(x_24, 1, x_22); +return x_24; } else { -lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; -lean_dec(x_14); -x_26 = lean_array_uget(x_11, x_13); -x_36 = lean_ctor_get(x_26, 0); +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; +lean_dec(x_13); +x_25 = lean_array_uget(x_10, x_12); +x_35 = lean_ctor_get(x_25, 0); +lean_inc(x_35); +x_36 = lean_ctor_get(x_25, 1); lean_inc(x_36); -x_37 = lean_ctor_get(x_26, 1); +x_37 = lean_ctor_get(x_25, 2); lean_inc(x_37); -x_38 = lean_ctor_get(x_26, 2); -lean_inc(x_38); -lean_dec(x_26); -lean_inc(x_17); -lean_inc(x_36); -x_39 = l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_getAltNumFields(x_1, x_36, x_17, x_18, x_19, x_20, x_21, x_22, x_23); -if (lean_obj_tag(x_39) == 0) +lean_dec(x_25); +lean_inc(x_16); +lean_inc(x_35); +x_38 = l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_getAltNumFields(x_1, x_35, x_16, x_17, x_18, x_19, x_20, x_21, x_22); +if (lean_obj_tag(x_38) == 0) { -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_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_38, 0); +lean_inc(x_39); +x_40 = lean_ctor_get(x_38, 1); lean_inc(x_40); -x_41 = lean_ctor_get(x_39, 1); -lean_inc(x_41); -lean_dec(x_39); -x_42 = lean_box(0); -x_43 = 0; -lean_inc(x_22); +lean_dec(x_38); +x_41 = lean_box(0); +x_42 = 0; lean_inc(x_21); lean_inc(x_20); lean_inc(x_19); -x_44 = l_Lean_Meta_introNCore(x_38, x_40, x_42, x_43, x_43, x_19, x_20, x_21, x_22, x_41); -if (lean_obj_tag(x_44) == 0) +lean_inc(x_18); +x_43 = l_Lean_Meta_introNCore(x_37, x_39, x_41, x_42, x_42, x_18, x_19, x_20, x_21, x_40); +if (lean_obj_tag(x_43) == 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; -x_45 = lean_ctor_get(x_44, 0); +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_44 = lean_ctor_get(x_43, 0); +lean_inc(x_44); +x_45 = lean_ctor_get(x_43, 1); lean_inc(x_45); +lean_dec(x_43); x_46 = lean_ctor_get(x_44, 1); lean_inc(x_46); lean_dec(x_44); -x_47 = lean_ctor_get(x_45, 1); -lean_inc(x_47); -lean_dec(x_45); +x_47 = lean_box(0); x_48 = lean_box(0); -x_49 = lean_box(0); -lean_inc(x_22); lean_inc(x_21); lean_inc(x_20); lean_inc(x_19); -lean_inc(x_47); -lean_inc(x_4); -x_50 = l_Lean_Meta_Cases_unifyEqs_x3f(x_4, x_47, x_48, x_49, x_19, x_20, x_21, x_22, x_46); -if (lean_obj_tag(x_50) == 0) +lean_inc(x_18); +lean_inc(x_46); +lean_inc(x_3); +x_49 = l_Lean_Meta_Cases_unifyEqs_x3f(x_3, x_46, x_47, x_48, x_18, x_19, x_20, x_21, x_45); +if (lean_obj_tag(x_49) == 0) { -lean_object* x_51; -x_51 = lean_ctor_get(x_50, 0); -lean_inc(x_51); -if (lean_obj_tag(x_51) == 0) +lean_object* x_50; +x_50 = lean_ctor_get(x_49, 0); +lean_inc(x_50); +if (lean_obj_tag(x_50) == 0) { -lean_object* x_52; lean_object* x_53; -lean_dec(x_47); -lean_dec(x_37); +lean_object* x_51; lean_object* x_52; +lean_dec(x_46); lean_dec(x_36); -x_52 = lean_ctor_get(x_50, 1); -lean_inc(x_52); -lean_dec(x_50); -x_53 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__8___lambda__2___closed__1; -x_27 = x_53; -x_28 = x_52; -goto block_35; +lean_dec(x_35); +x_51 = lean_ctor_get(x_49, 1); +lean_inc(x_51); +lean_dec(x_49); +x_52 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__8___lambda__2___closed__1; +x_26 = x_52; +x_27 = x_51; +goto block_34; } else { -lean_object* x_54; uint8_t x_55; -x_54 = lean_ctor_get(x_51, 0); -lean_inc(x_54); -lean_dec(x_51); -x_55 = lean_ctor_get_uint8(x_37, sizeof(void*)*3); -lean_dec(x_37); -if (x_55 == 0) +lean_object* x_53; uint8_t x_54; +x_53 = lean_ctor_get(x_50, 0); +lean_inc(x_53); +lean_dec(x_50); +x_54 = lean_ctor_get_uint8(x_36, sizeof(void*)*3); +lean_dec(x_36); +if (x_54 == 0) { -lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; -x_56 = lean_ctor_get(x_50, 1); +lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; +x_55 = lean_ctor_get(x_49, 1); +lean_inc(x_55); +lean_dec(x_49); +x_56 = lean_ctor_get(x_53, 0); lean_inc(x_56); -lean_dec(x_50); -x_57 = lean_ctor_get(x_54, 0); +x_57 = lean_ctor_get(x_53, 1); lean_inc(x_57); -x_58 = lean_ctor_get(x_54, 1); -lean_inc(x_58); -lean_dec(x_54); -lean_inc(x_22); +lean_dec(x_53); lean_inc(x_21); lean_inc(x_20); lean_inc(x_19); @@ -15979,25 +15934,25 @@ lean_inc(x_18); lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); -lean_inc(x_7); -x_59 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__8___lambda__2(x_6, x_7, x_58, x_2, x_8, x_36, x_47, x_57, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_56); -lean_dec(x_47); -if (lean_obj_tag(x_59) == 0) +lean_inc(x_14); +lean_inc(x_6); +x_58 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__8___lambda__2(x_5, x_6, x_57, x_2, x_7, x_35, x_46, x_56, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_55); +lean_dec(x_46); +if (lean_obj_tag(x_58) == 0) { -lean_object* x_60; lean_object* x_61; -x_60 = lean_ctor_get(x_59, 0); +lean_object* x_59; lean_object* x_60; +x_59 = lean_ctor_get(x_58, 0); +lean_inc(x_59); +x_60 = lean_ctor_get(x_58, 1); lean_inc(x_60); -x_61 = lean_ctor_get(x_59, 1); -lean_inc(x_61); -lean_dec(x_59); +lean_dec(x_58); +x_26 = x_59; x_27 = x_60; -x_28 = x_61; -goto block_35; +goto block_34; } else { -uint8_t x_62; -lean_dec(x_22); +uint8_t x_61; lean_dec(x_21); lean_dec(x_20); lean_dec(x_19); @@ -16005,60 +15960,60 @@ lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); -lean_dec(x_7); -lean_dec(x_5); +lean_dec(x_14); +lean_dec(x_6); lean_dec(x_4); -x_62 = !lean_is_exclusive(x_59); -if (x_62 == 0) +lean_dec(x_3); +x_61 = !lean_is_exclusive(x_58); +if (x_61 == 0) { -return x_59; +return x_58; } else { -lean_object* x_63; lean_object* x_64; lean_object* x_65; -x_63 = lean_ctor_get(x_59, 0); -x_64 = lean_ctor_get(x_59, 1); -lean_inc(x_64); +lean_object* x_62; lean_object* x_63; lean_object* x_64; +x_62 = lean_ctor_get(x_58, 0); +x_63 = lean_ctor_get(x_58, 1); lean_inc(x_63); -lean_dec(x_59); -x_65 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_65, 0, x_63); -lean_ctor_set(x_65, 1, x_64); -return x_65; +lean_inc(x_62); +lean_dec(x_58); +x_64 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_64, 0, x_62); +lean_ctor_set(x_64, 1, x_63); +return x_64; } } } else { -lean_object* x_66; lean_object* x_67; lean_object* x_68; uint8_t x_69; lean_object* x_70; -lean_dec(x_47); -x_66 = lean_ctor_get(x_50, 1); +lean_object* x_65; lean_object* x_66; lean_object* x_67; uint8_t x_68; lean_object* x_69; +lean_dec(x_46); +x_65 = lean_ctor_get(x_49, 1); +lean_inc(x_65); +lean_dec(x_49); +x_66 = lean_ctor_get(x_53, 0); lean_inc(x_66); -lean_dec(x_50); -x_67 = lean_ctor_get(x_54, 0); +x_67 = lean_ctor_get(x_53, 1); lean_inc(x_67); -x_68 = lean_ctor_get(x_54, 1); -lean_inc(x_68); -lean_dec(x_54); -x_69 = 1; -lean_inc(x_22); +lean_dec(x_53); +x_68 = 1; lean_inc(x_21); lean_inc(x_20); lean_inc(x_19); -lean_inc(x_5); -x_70 = l_Lean_Meta_introNCore(x_67, x_5, x_42, x_43, x_69, x_19, x_20, x_21, x_22, x_66); -if (lean_obj_tag(x_70) == 0) +lean_inc(x_18); +lean_inc(x_4); +x_69 = l_Lean_Meta_introNCore(x_66, x_4, x_41, x_42, x_68, x_18, x_19, x_20, x_21, x_65); +if (lean_obj_tag(x_69) == 0) { -lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; -x_71 = lean_ctor_get(x_70, 0); +lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; +x_70 = lean_ctor_get(x_69, 0); +lean_inc(x_70); +x_71 = lean_ctor_get(x_69, 1); lean_inc(x_71); +lean_dec(x_69); x_72 = lean_ctor_get(x_70, 1); lean_inc(x_72); lean_dec(x_70); -x_73 = lean_ctor_get(x_71, 1); -lean_inc(x_73); -lean_dec(x_71); -lean_inc(x_22); lean_inc(x_21); lean_inc(x_20); lean_inc(x_19); @@ -16066,26 +16021,26 @@ lean_inc(x_18); lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); -lean_inc(x_73); -lean_inc(x_7); -x_74 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__8___lambda__2(x_6, x_7, x_68, x_2, x_8, x_36, x_73, x_73, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_72); -lean_dec(x_73); -if (lean_obj_tag(x_74) == 0) +lean_inc(x_14); +lean_inc(x_72); +lean_inc(x_6); +x_73 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__8___lambda__2(x_5, x_6, x_67, x_2, x_7, x_35, x_72, x_72, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_71); +lean_dec(x_72); +if (lean_obj_tag(x_73) == 0) { -lean_object* x_75; lean_object* x_76; -x_75 = lean_ctor_get(x_74, 0); +lean_object* x_74; lean_object* x_75; +x_74 = lean_ctor_get(x_73, 0); +lean_inc(x_74); +x_75 = lean_ctor_get(x_73, 1); lean_inc(x_75); -x_76 = lean_ctor_get(x_74, 1); -lean_inc(x_76); -lean_dec(x_74); +lean_dec(x_73); +x_26 = x_74; x_27 = x_75; -x_28 = x_76; -goto block_35; +goto block_34; } else { -uint8_t x_77; -lean_dec(x_22); +uint8_t x_76; lean_dec(x_21); lean_dec(x_20); lean_dec(x_19); @@ -16093,35 +16048,35 @@ lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); -lean_dec(x_7); -lean_dec(x_5); +lean_dec(x_14); +lean_dec(x_6); lean_dec(x_4); -x_77 = !lean_is_exclusive(x_74); -if (x_77 == 0) +lean_dec(x_3); +x_76 = !lean_is_exclusive(x_73); +if (x_76 == 0) { -return x_74; +return x_73; } else { -lean_object* x_78; lean_object* x_79; lean_object* x_80; -x_78 = lean_ctor_get(x_74, 0); -x_79 = lean_ctor_get(x_74, 1); -lean_inc(x_79); +lean_object* x_77; lean_object* x_78; lean_object* x_79; +x_77 = lean_ctor_get(x_73, 0); +x_78 = lean_ctor_get(x_73, 1); lean_inc(x_78); -lean_dec(x_74); -x_80 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_80, 0, x_78); -lean_ctor_set(x_80, 1, x_79); -return x_80; +lean_inc(x_77); +lean_dec(x_73); +x_79 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_79, 0, x_77); +lean_ctor_set(x_79, 1, x_78); +return x_79; } } } else { -uint8_t x_81; -lean_dec(x_68); -lean_dec(x_36); -lean_dec(x_22); +uint8_t x_80; +lean_dec(x_67); +lean_dec(x_35); lean_dec(x_21); lean_dec(x_20); lean_dec(x_19); @@ -16129,26 +16084,27 @@ lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); -lean_dec(x_7); -lean_dec(x_5); +lean_dec(x_14); +lean_dec(x_6); lean_dec(x_4); -x_81 = !lean_is_exclusive(x_70); -if (x_81 == 0) +lean_dec(x_3); +x_80 = !lean_is_exclusive(x_69); +if (x_80 == 0) { -return x_70; +return x_69; } else { -lean_object* x_82; lean_object* x_83; lean_object* x_84; -x_82 = lean_ctor_get(x_70, 0); -x_83 = lean_ctor_get(x_70, 1); -lean_inc(x_83); +lean_object* x_81; lean_object* x_82; lean_object* x_83; +x_81 = lean_ctor_get(x_69, 0); +x_82 = lean_ctor_get(x_69, 1); lean_inc(x_82); -lean_dec(x_70); -x_84 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_84, 0, x_82); -lean_ctor_set(x_84, 1, x_83); -return x_84; +lean_inc(x_81); +lean_dec(x_69); +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; } } } @@ -16156,11 +16112,10 @@ return x_84; } else { -uint8_t x_85; -lean_dec(x_47); -lean_dec(x_37); +uint8_t x_84; +lean_dec(x_46); lean_dec(x_36); -lean_dec(x_22); +lean_dec(x_35); lean_dec(x_21); lean_dec(x_20); lean_dec(x_19); @@ -16168,35 +16123,35 @@ lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); -lean_dec(x_7); -lean_dec(x_5); +lean_dec(x_14); +lean_dec(x_6); lean_dec(x_4); -x_85 = !lean_is_exclusive(x_50); -if (x_85 == 0) +lean_dec(x_3); +x_84 = !lean_is_exclusive(x_49); +if (x_84 == 0) { -return x_50; +return x_49; } else { -lean_object* x_86; lean_object* x_87; lean_object* x_88; -x_86 = lean_ctor_get(x_50, 0); -x_87 = lean_ctor_get(x_50, 1); -lean_inc(x_87); +lean_object* x_85; lean_object* x_86; lean_object* x_87; +x_85 = lean_ctor_get(x_49, 0); +x_86 = lean_ctor_get(x_49, 1); lean_inc(x_86); -lean_dec(x_50); -x_88 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_88, 0, x_86); -lean_ctor_set(x_88, 1, x_87); -return x_88; +lean_inc(x_85); +lean_dec(x_49); +x_87 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_87, 0, x_85); +lean_ctor_set(x_87, 1, x_86); +return x_87; } } } else { -uint8_t x_89; -lean_dec(x_37); +uint8_t x_88; lean_dec(x_36); -lean_dec(x_22); +lean_dec(x_35); lean_dec(x_21); lean_dec(x_20); lean_dec(x_19); @@ -16204,36 +16159,36 @@ lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); -lean_dec(x_7); -lean_dec(x_5); +lean_dec(x_14); +lean_dec(x_6); lean_dec(x_4); -x_89 = !lean_is_exclusive(x_44); -if (x_89 == 0) +lean_dec(x_3); +x_88 = !lean_is_exclusive(x_43); +if (x_88 == 0) { -return x_44; +return x_43; } else { -lean_object* x_90; lean_object* x_91; lean_object* x_92; -x_90 = lean_ctor_get(x_44, 0); -x_91 = lean_ctor_get(x_44, 1); -lean_inc(x_91); +lean_object* x_89; lean_object* x_90; lean_object* x_91; +x_89 = lean_ctor_get(x_43, 0); +x_90 = lean_ctor_get(x_43, 1); lean_inc(x_90); -lean_dec(x_44); -x_92 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_92, 0, x_90); -lean_ctor_set(x_92, 1, x_91); -return x_92; +lean_inc(x_89); +lean_dec(x_43); +x_91 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_91, 0, x_89); +lean_ctor_set(x_91, 1, x_90); +return x_91; } } } else { -uint8_t x_93; -lean_dec(x_38); +uint8_t x_92; lean_dec(x_37); lean_dec(x_36); -lean_dec(x_22); +lean_dec(x_35); lean_dec(x_21); lean_dec(x_20); lean_dec(x_19); @@ -16241,34 +16196,34 @@ lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); -lean_dec(x_7); -lean_dec(x_5); +lean_dec(x_14); +lean_dec(x_6); lean_dec(x_4); -x_93 = !lean_is_exclusive(x_39); -if (x_93 == 0) +lean_dec(x_3); +x_92 = !lean_is_exclusive(x_38); +if (x_92 == 0) { -return x_39; +return x_38; } else { -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_object* x_93; lean_object* x_94; lean_object* x_95; +x_93 = lean_ctor_get(x_38, 0); +x_94 = lean_ctor_get(x_38, 1); lean_inc(x_94); -lean_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; +lean_inc(x_93); +lean_dec(x_38); +x_95 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_95, 0, x_93); +lean_ctor_set(x_95, 1, x_94); +return x_95; } } -block_35: +block_34: { -if (lean_obj_tag(x_27) == 0) +if (lean_obj_tag(x_26) == 0) { -lean_object* x_29; lean_object* x_30; -lean_dec(x_22); +lean_object* x_28; lean_object* x_29; lean_dec(x_21); lean_dec(x_20); lean_dec(x_19); @@ -16276,28 +16231,29 @@ lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); -lean_dec(x_7); -lean_dec(x_5); +lean_dec(x_14); +lean_dec(x_6); lean_dec(x_4); -x_29 = lean_ctor_get(x_27, 0); -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_dec(x_3); +x_28 = lean_ctor_get(x_26, 0); +lean_inc(x_28); +lean_dec(x_26); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_27); +return x_29; } else { -lean_object* x_31; size_t x_32; size_t x_33; -x_31 = lean_ctor_get(x_27, 0); -lean_inc(x_31); -lean_dec(x_27); -x_32 = 1; -x_33 = lean_usize_add(x_13, x_32); -x_13 = x_33; -x_14 = x_31; -x_23 = x_28; +lean_object* x_30; size_t x_31; size_t x_32; +x_30 = lean_ctor_get(x_26, 0); +lean_inc(x_30); +lean_dec(x_26); +x_31 = 1; +x_32 = lean_usize_add(x_12, x_31); +x_12 = x_32; +x_13 = x_30; +x_22 = x_27; goto _start; } } @@ -16307,23 +16263,57 @@ goto _start; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, 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; +uint8_t x_19; +if (lean_obj_tag(x_4) == 0) +{ +uint8_t x_55; +x_55 = 0; +x_19 = x_55; +goto block_54; +} +else +{ +uint8_t x_56; +x_56 = 1; +x_19 = x_56; +goto block_54; +} +block_54: +{ +lean_object* x_20; +if (lean_obj_tag(x_4) == 0) +{ +lean_object* x_52; +x_52 = l_Lean_Elab_Tactic_ElimApp_mkElimApp___closed__1; +x_20 = x_52; +goto block_51; +} +else +{ +lean_object* x_53; +x_53 = lean_ctor_get(x_4, 0); +lean_inc(x_53); +x_20 = x_53; +goto block_51; +} +block_51: +{ +lean_object* x_21; lean_inc(x_16); -x_19 = l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_checkAltNames(x_2, x_4, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); -if (lean_obj_tag(x_19) == 0) +x_21 = l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_checkAltNames(x_2, x_20, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +if (lean_obj_tag(x_21) == 0) { -lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; -x_20 = lean_ctor_get(x_19, 1); -lean_inc(x_20); -lean_dec(x_19); -x_21 = lean_array_get_size(x_4); -x_22 = lean_unsigned_to_nat(0u); -x_23 = lean_unsigned_to_nat(1u); -lean_inc(x_21); -x_24 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_24, 0, x_22); -lean_ctor_set(x_24, 1, x_21); -lean_ctor_set(x_24, 2, x_23); +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, 1); +lean_inc(x_22); +lean_dec(x_21); +x_23 = lean_array_get_size(x_20); +x_24 = lean_unsigned_to_nat(0u); +x_25 = lean_unsigned_to_nat(1u); +x_26 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_26, 0, x_24); +lean_ctor_set(x_26, 1, x_23); +lean_ctor_set(x_26, 2, x_25); lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); @@ -16335,76 +16325,73 @@ lean_inc(x_10); lean_inc(x_8); lean_inc(x_6); lean_inc(x_5); -lean_inc(x_4); -x_25 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__4(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_24, x_24, x_2, x_22, lean_box(0), lean_box(0), x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_20); -lean_dec(x_24); -if (lean_obj_tag(x_25) == 0) -{ -lean_object* x_26; lean_object* x_27; lean_object* x_28; size_t x_29; size_t x_30; lean_object* x_31; lean_object* x_32; -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_box(0); -x_29 = lean_array_size(x_26); -x_30 = 0; -x_31 = lean_box(0); -x_32 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__8(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_21, x_26, x_28, x_26, x_29, x_30, x_31, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_27); +x_27 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__4(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_20, x_26, x_26, x_2, x_24, lean_box(0), lean_box(0), x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_22); lean_dec(x_26); -lean_dec(x_21); lean_dec(x_4); -if (lean_obj_tag(x_32) == 0) +if (lean_obj_tag(x_27) == 0) { -uint8_t x_33; -x_33 = !lean_is_exclusive(x_32); -if (x_33 == 0) +lean_object* x_28; lean_object* x_29; lean_object* x_30; size_t x_31; size_t x_32; lean_object* x_33; lean_object* x_34; +x_28 = lean_ctor_get(x_27, 0); +lean_inc(x_28); +x_29 = lean_ctor_get(x_27, 1); +lean_inc(x_29); +lean_dec(x_27); +x_30 = lean_box(0); +x_31 = lean_array_size(x_28); +x_32 = 0; +x_33 = lean_box(0); +x_34 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__8(x_1, x_3, x_5, x_6, x_7, x_8, x_19, x_28, x_30, x_28, x_31, x_32, x_33, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_29); +lean_dec(x_28); +if (lean_obj_tag(x_34) == 0) { -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; +uint8_t x_35; +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_33); +return x_34; } 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_31); -lean_ctor_set(x_36, 1, x_35); -return x_36; +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_33); +lean_ctor_set(x_38, 1, x_37); +return x_38; } } else { -uint8_t x_37; -x_37 = !lean_is_exclusive(x_32); -if (x_37 == 0) +uint8_t x_39; +x_39 = !lean_is_exclusive(x_34); +if (x_39 == 0) { -return x_32; +return x_34; } 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; +lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_40 = lean_ctor_get(x_34, 0); +x_41 = lean_ctor_get(x_34, 1); +lean_inc(x_41); +lean_inc(x_40); +lean_dec(x_34); +x_42 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_42, 0, x_40); +lean_ctor_set(x_42, 1, x_41); +return x_42; } } } else { -uint8_t x_41; -lean_dec(x_21); +uint8_t x_43; lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); @@ -16416,30 +16403,30 @@ lean_dec(x_10); lean_dec(x_8); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); -x_41 = !lean_is_exclusive(x_25); -if (x_41 == 0) +x_43 = !lean_is_exclusive(x_27); +if (x_43 == 0) { -return x_25; +return x_27; } 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_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; +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_27, 0); +x_45 = lean_ctor_get(x_27, 1); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_27); +x_46 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_46, 0, x_44); +lean_ctor_set(x_46, 1, x_45); +return x_46; } } } else { -uint8_t x_45; +uint8_t x_47; +lean_dec(x_20); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); @@ -16453,23 +16440,25 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); -x_45 = !lean_is_exclusive(x_19); -if (x_45 == 0) +x_47 = !lean_is_exclusive(x_21); +if (x_47 == 0) { -return x_19; +return x_21; } 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_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; +lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_48 = lean_ctor_get(x_21, 0); +x_49 = lean_ctor_get(x_21, 1); +lean_inc(x_49); +lean_inc(x_48); +lean_dec(x_21); +x_50 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_50, 0, x_48); +lean_ctor_set(x_50, 1, x_49); +return x_50; +} +} } } } @@ -16543,8 +16532,8 @@ x_28 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWit lean_dec(x_13); lean_dec(x_12); lean_dec(x_10); -lean_dec(x_8); -lean_dec(x_6); +lean_dec(x_7); +lean_dec(x_5); lean_dec(x_2); lean_dec(x_1); return x_28; @@ -16585,17 +16574,19 @@ lean_object* x_20 = _args[19]; lean_object* x_21 = _args[20]; lean_object* x_22 = _args[21]; lean_object* x_23 = _args[22]; +lean_object* x_24 = _args[23]; _start: { -lean_object* x_24; -x_24 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_23); +lean_object* x_25; +x_25 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_23, x_24); +lean_dec(x_11); lean_dec(x_10); -lean_dec(x_9); lean_dec(x_8); lean_dec(x_6); +lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -return x_24; +return x_25; } } LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { @@ -16684,13 +16675,14 @@ lean_object* x_16 = _args[15]; lean_object* x_17 = _args[16]; _start: { -lean_object* x_18; -x_18 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___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, x_12, x_13, x_14, x_15, x_16, x_17); -lean_dec(x_7); +uint8_t x_18; lean_object* x_19; +x_18 = lean_unbox(x_5); lean_dec(x_5); +x_19 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__8___lambda__2(x_1, x_2, x_3, x_4, x_18, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +lean_dec(x_7); lean_dec(x_4); lean_dec(x_1); -return x_18; +return x_19; } } LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__8___boxed(lean_object** _args) { @@ -16716,21 +16708,20 @@ lean_object* x_19 = _args[18]; lean_object* x_20 = _args[19]; lean_object* x_21 = _args[20]; lean_object* x_22 = _args[21]; -lean_object* x_23 = _args[22]; _start: { -size_t x_24; size_t x_25; lean_object* x_26; -x_24 = lean_unbox_usize(x_12); -lean_dec(x_12); -x_25 = lean_unbox_usize(x_13); -lean_dec(x_13); -x_26 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___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_24, x_25, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_23); +uint8_t x_23; size_t x_24; size_t x_25; lean_object* x_26; +x_23 = lean_unbox(x_7); +lean_dec(x_7); +x_24 = lean_unbox_usize(x_11); lean_dec(x_11); +x_25 = lean_unbox_usize(x_12); +lean_dec(x_12); +x_26 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__8(x_1, x_2, x_3, x_4, x_5, x_6, x_23, x_8, x_9, x_10, x_24, x_25, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); -lean_dec(x_6); -lean_dec(x_3); +lean_dec(x_5); lean_dec(x_2); lean_dec(x_1); return x_26; @@ -17393,162 +17384,163 @@ lean_ctor_set_uint8(x_6, sizeof(void*)*4, x_5); return x_6; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19, lean_object* x_20) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19, lean_object* x_20, lean_object* x_21) { _start: { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; -x_21 = lean_box(0); -x_22 = lean_box(2); -x_23 = l_Lean_Elab_Tactic_withCaseRef___at_Lean_Elab_Tactic_evalAlt___spec__1___closed__2; +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_22 = lean_box(0); +x_23 = lean_box(2); +x_24 = l_Lean_Elab_Tactic_withCaseRef___at_Lean_Elab_Tactic_evalAlt___spec__1___closed__2; lean_inc(x_1); -x_24 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_24, 0, x_22); -lean_ctor_set(x_24, 1, x_23); -lean_ctor_set(x_24, 2, x_1); +x_25 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_24); +lean_ctor_set(x_25, 2, x_1); lean_inc(x_2); -x_25 = lean_io_promise_result(x_2); -x_26 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_26, 0, x_21); -lean_ctor_set(x_26, 1, x_25); -x_27 = lean_unsigned_to_nat(0u); -x_28 = l_Lean_Elab_Tactic_ElimApp_mkElimApp___closed__1; -x_29 = l_Array_zipWithAux___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo___spec__1(x_1, x_11, x_27, x_28); -x_30 = l_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo___lambda__2___closed__11; -x_31 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_31, 0, x_30); -lean_ctor_set(x_31, 1, x_24); -lean_ctor_set(x_31, 2, x_21); -lean_ctor_set(x_31, 3, x_26); -lean_ctor_set(x_31, 4, x_29); -x_32 = lean_ctor_get(x_3, 1); -lean_inc(x_32); -x_33 = lean_io_promise_resolve(x_31, x_32, x_20); -lean_dec(x_32); -x_34 = lean_ctor_get(x_33, 1); -lean_inc(x_34); +x_26 = lean_io_promise_result(x_2); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_22); +lean_ctor_set(x_27, 1, x_26); +x_28 = lean_unsigned_to_nat(0u); +x_29 = l_Lean_Elab_Tactic_ElimApp_mkElimApp___closed__1; +x_30 = l_Array_zipWithAux___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo___spec__1(x_1, x_12, x_28, x_29); +lean_dec(x_1); +x_31 = l_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo___lambda__2___closed__11; +x_32 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_25); +lean_ctor_set(x_32, 2, x_22); +lean_ctor_set(x_32, 3, x_27); +lean_ctor_set(x_32, 4, x_30); +x_33 = lean_ctor_get(x_3, 1); +lean_inc(x_33); +x_34 = lean_io_promise_resolve(x_32, x_33, x_21); lean_dec(x_33); -x_35 = lean_array_get_size(x_11); -x_36 = lean_mk_empty_array_with_capacity(x_35); -x_37 = l_Array_mapFinIdxM_map___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo___spec__2(x_3, x_11, x_11, x_35, x_27, lean_box(0), x_36); +x_35 = lean_ctor_get(x_34, 1); +lean_inc(x_35); +lean_dec(x_34); +x_36 = lean_array_get_size(x_12); +x_37 = lean_mk_empty_array_with_capacity(x_36); +x_38 = l_Array_mapFinIdxM_map___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo___spec__2(x_3, x_12, x_12, x_36, x_28, lean_box(0), x_37); +lean_inc(x_20); lean_inc(x_19); lean_inc(x_18); lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); lean_inc(x_14); -lean_inc(x_13); -x_38 = l_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental(x_4, x_5, x_6, x_1, x_7, x_8, x_9, x_10, x_37, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_34); -lean_dec(x_37); -if (lean_obj_tag(x_38) == 0) -{ -lean_object* x_39; lean_object* x_40; uint8_t x_41; -x_39 = lean_ctor_get(x_38, 1); -lean_inc(x_39); +x_39 = l_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental(x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_38, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_35); lean_dec(x_38); -x_40 = l_Lean_Elab_Tactic_saveState___rarg(x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_39); +if (lean_obj_tag(x_39) == 0) +{ +lean_object* x_40; lean_object* x_41; uint8_t x_42; +x_40 = lean_ctor_get(x_39, 1); +lean_inc(x_40); +lean_dec(x_39); +x_41 = l_Lean_Elab_Tactic_saveState___rarg(x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_40); +lean_dec(x_20); lean_dec(x_19); lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); -lean_dec(x_13); -x_41 = !lean_is_exclusive(x_40); -if (x_41 == 0) +x_42 = !lean_is_exclusive(x_41); +if (x_42 == 0) { -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_40, 0); -x_43 = lean_ctor_get(x_40, 1); -x_44 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_44, 0, x_42); -lean_ctor_set(x_40, 1, x_44); -lean_ctor_set(x_40, 0, x_30); -x_45 = lean_io_promise_resolve(x_40, x_2, x_43); +lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; uint8_t x_47; +x_43 = lean_ctor_get(x_41, 0); +x_44 = lean_ctor_get(x_41, 1); +x_45 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_41, 1, x_45); +lean_ctor_set(x_41, 0, x_31); +x_46 = lean_io_promise_resolve(x_41, x_2, x_44); lean_dec(x_2); -x_46 = !lean_is_exclusive(x_45); -if (x_46 == 0) +x_47 = !lean_is_exclusive(x_46); +if (x_47 == 0) { -return x_45; +return x_46; } else { -lean_object* x_47; lean_object* x_48; lean_object* x_49; -x_47 = lean_ctor_get(x_45, 0); -x_48 = lean_ctor_get(x_45, 1); +lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_48 = lean_ctor_get(x_46, 0); +x_49 = lean_ctor_get(x_46, 1); +lean_inc(x_49); lean_inc(x_48); -lean_inc(x_47); -lean_dec(x_45); -x_49 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_49, 0, x_47); -lean_ctor_set(x_49, 1, x_48); -return x_49; +lean_dec(x_46); +x_50 = lean_alloc_ctor(0, 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_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; -x_50 = lean_ctor_get(x_40, 0); -x_51 = lean_ctor_get(x_40, 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; +x_51 = lean_ctor_get(x_41, 0); +x_52 = lean_ctor_get(x_41, 1); +lean_inc(x_52); lean_inc(x_51); -lean_inc(x_50); -lean_dec(x_40); -x_52 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_52, 0, x_50); -x_53 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_53, 0, x_30); -lean_ctor_set(x_53, 1, x_52); -x_54 = lean_io_promise_resolve(x_53, x_2, x_51); +lean_dec(x_41); +x_53 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_53, 0, x_51); +x_54 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_54, 0, x_31); +lean_ctor_set(x_54, 1, x_53); +x_55 = lean_io_promise_resolve(x_54, x_2, x_52); lean_dec(x_2); -x_55 = lean_ctor_get(x_54, 0); -lean_inc(x_55); -x_56 = lean_ctor_get(x_54, 1); +x_56 = lean_ctor_get(x_55, 0); 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; +x_57 = lean_ctor_get(x_55, 1); +lean_inc(x_57); +if (lean_is_exclusive(x_55)) { + lean_ctor_release(x_55, 0); + lean_ctor_release(x_55, 1); + x_58 = x_55; } else { - lean_dec_ref(x_54); - x_57 = lean_box(0); + lean_dec_ref(x_55); + x_58 = lean_box(0); } -if (lean_is_scalar(x_57)) { - x_58 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_58)) { + x_59 = lean_alloc_ctor(0, 2, 0); } else { - x_58 = x_57; + x_59 = x_58; } -lean_ctor_set(x_58, 0, x_55); -lean_ctor_set(x_58, 1, x_56); -return x_58; +lean_ctor_set(x_59, 0, x_56); +lean_ctor_set(x_59, 1, x_57); +return x_59; } } else { -uint8_t x_59; +uint8_t x_60; +lean_dec(x_20); lean_dec(x_19); lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); -lean_dec(x_13); lean_dec(x_2); -x_59 = !lean_is_exclusive(x_38); -if (x_59 == 0) +x_60 = !lean_is_exclusive(x_39); +if (x_60 == 0) { -return x_38; +return x_39; } else { -lean_object* x_60; lean_object* x_61; lean_object* x_62; -x_60 = lean_ctor_get(x_38, 0); -x_61 = lean_ctor_get(x_38, 1); +lean_object* x_61; lean_object* x_62; lean_object* x_63; +x_61 = lean_ctor_get(x_39, 0); +x_62 = lean_ctor_get(x_39, 1); +lean_inc(x_62); lean_inc(x_61); -lean_inc(x_60); -lean_dec(x_38); -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_39); +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; } } } @@ -17556,122 +17548,120 @@ return x_62; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo___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, lean_object* x_19, lean_object* x_20) { _start: { -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo___lambda__2___boxed), 20, 10); -lean_closure_set(x_21, 0, x_1); -lean_closure_set(x_21, 1, x_11); -lean_closure_set(x_21, 2, x_2); -lean_closure_set(x_21, 3, x_3); -lean_closure_set(x_21, 4, x_4); -lean_closure_set(x_21, 5, x_5); -lean_closure_set(x_21, 6, x_6); -lean_closure_set(x_21, 7, x_7); -lean_closure_set(x_21, 8, x_8); -lean_closure_set(x_21, 9, x_9); -x_22 = l_Lean_Elab_Tactic_instInhabitedTacticParsedSnapshot; -x_23 = l_Lean_Language_withAlwaysResolvedPromises___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo___spec__3(x_22, x_10, x_21, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20); -return x_23; +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_21 = lean_array_get_size(x_1); +x_22 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo___lambda__2___boxed), 21, 11); +lean_closure_set(x_22, 0, x_1); +lean_closure_set(x_22, 1, x_11); +lean_closure_set(x_22, 2, x_2); +lean_closure_set(x_22, 3, x_3); +lean_closure_set(x_22, 4, x_4); +lean_closure_set(x_22, 5, x_5); +lean_closure_set(x_22, 6, x_6); +lean_closure_set(x_22, 7, x_7); +lean_closure_set(x_22, 8, x_8); +lean_closure_set(x_22, 9, x_9); +lean_closure_set(x_22, 10, x_10); +x_23 = l_Lean_Elab_Tactic_instInhabitedTacticParsedSnapshot; +x_24 = l_Lean_Language_withAlwaysResolvedPromises___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo___spec__3(x_23, x_21, x_22, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20); +return x_24; } } LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, 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; uint8_t x_20; -x_18 = lean_array_get_size(x_4); -x_19 = lean_unsigned_to_nat(0u); -x_20 = lean_nat_dec_lt(x_19, x_18); -if (x_20 == 0) +if (lean_obj_tag(x_4) == 0) { -lean_object* x_21; lean_object* x_22; -lean_dec(x_18); -x_21 = lean_box(0); -x_22 = l_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_21, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +lean_object* x_18; lean_object* x_19; +x_18 = lean_box(0); +x_19 = l_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_18, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); lean_dec(x_7); lean_dec(x_3); lean_dec(x_1); -return x_22; +return x_19; } else { -lean_object* x_23; -x_23 = lean_ctor_get(x_11, 8); -lean_inc(x_23); -if (lean_obj_tag(x_23) == 0) +lean_object* x_20; +x_20 = lean_ctor_get(x_11, 8); +lean_inc(x_20); +if (lean_obj_tag(x_20) == 0) { -lean_object* x_24; lean_object* x_25; -lean_dec(x_18); -x_24 = lean_box(0); -x_25 = l_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_24, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +lean_object* x_21; lean_object* x_22; +x_21 = lean_box(0); +x_22 = l_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_21, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); lean_dec(x_7); lean_dec(x_3); lean_dec(x_1); -return x_25; +return x_22; } else { -lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_26 = lean_ctor_get(x_23, 0); -lean_inc(x_26); -lean_dec(x_23); -x_27 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo___lambda__3___boxed), 20, 10); -lean_closure_set(x_27, 0, x_4); -lean_closure_set(x_27, 1, x_26); -lean_closure_set(x_27, 2, x_1); -lean_closure_set(x_27, 3, x_2); -lean_closure_set(x_27, 4, x_3); -lean_closure_set(x_27, 5, x_5); -lean_closure_set(x_27, 6, x_6); -lean_closure_set(x_27, 7, x_7); -lean_closure_set(x_27, 8, x_8); -lean_closure_set(x_27, 9, x_18); -x_28 = l_Lean_Elab_Tactic_instInhabitedTacticFinishedSnapshot; -x_29 = l_Lean_Language_withAlwaysResolvedPromise___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo___spec__7(x_28, x_27, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); -if (lean_obj_tag(x_29) == 0) +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_23 = lean_ctor_get(x_4, 0); +lean_inc(x_23); +x_24 = lean_ctor_get(x_20, 0); +lean_inc(x_24); +lean_dec(x_20); +x_25 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo___lambda__3___boxed), 20, 10); +lean_closure_set(x_25, 0, x_23); +lean_closure_set(x_25, 1, x_24); +lean_closure_set(x_25, 2, x_1); +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_6); +lean_closure_set(x_25, 8, x_7); +lean_closure_set(x_25, 9, x_8); +x_26 = l_Lean_Elab_Tactic_instInhabitedTacticFinishedSnapshot; +x_27 = l_Lean_Language_withAlwaysResolvedPromise___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo___spec__7(x_26, x_25, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +if (lean_obj_tag(x_27) == 0) { -uint8_t x_30; -x_30 = !lean_is_exclusive(x_29); -if (x_30 == 0) +uint8_t x_28; +x_28 = !lean_is_exclusive(x_27); +if (x_28 == 0) { -lean_object* x_31; lean_object* x_32; -x_31 = lean_ctor_get(x_29, 0); -lean_dec(x_31); -x_32 = lean_box(0); -lean_ctor_set(x_29, 0, x_32); -return x_29; +lean_object* x_29; lean_object* x_30; +x_29 = lean_ctor_get(x_27, 0); +lean_dec(x_29); +x_30 = lean_box(0); +lean_ctor_set(x_27, 0, x_30); +return x_27; } else { -lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_33 = lean_ctor_get(x_29, 1); -lean_inc(x_33); -lean_dec(x_29); -x_34 = lean_box(0); -x_35 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_35, 0, x_34); -lean_ctor_set(x_35, 1, x_33); -return x_35; +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_27, 1); +lean_inc(x_31); +lean_dec(x_27); +x_32 = lean_box(0); +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_31); +return x_33; } } else { -uint8_t x_36; -x_36 = !lean_is_exclusive(x_29); -if (x_36 == 0) +uint8_t x_34; +x_34 = !lean_is_exclusive(x_27); +if (x_34 == 0) { -return x_29; +return x_27; } else { -lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_37 = lean_ctor_get(x_29, 0); -x_38 = lean_ctor_get(x_29, 1); -lean_inc(x_38); -lean_inc(x_37); -lean_dec(x_29); -x_39 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_39, 0, x_37); -lean_ctor_set(x_39, 1, x_38); -return x_39; +lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_35 = lean_ctor_get(x_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; } } } @@ -17815,15 +17805,16 @@ lean_object* x_17 = _args[16]; lean_object* x_18 = _args[17]; lean_object* x_19 = _args[18]; lean_object* x_20 = _args[19]; +lean_object* x_21 = _args[20]; _start: { -lean_object* x_21; -x_21 = l_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20); -lean_dec(x_11); -lean_dec(x_9); +lean_object* x_22; +x_22 = l_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21); +lean_dec(x_12); +lean_dec(x_10); lean_dec(x_6); lean_dec(x_4); -return x_21; +return x_22; } } LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo___lambda__3___boxed(lean_object** _args) { @@ -18002,41 +17993,36 @@ return x_13; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_ElimApp_evalAlts(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, 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; uint8_t x_21; -x_19 = lean_array_get_size(x_4); -x_20 = lean_unsigned_to_nat(0u); -x_21 = lean_nat_dec_lt(x_20, x_19); -lean_dec(x_19); -if (x_21 == 0) +if (lean_obj_tag(x_4) == 0) { -lean_object* x_22; +lean_object* x_19; lean_dec(x_5); -x_22 = l_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo(x_1, x_2, x_3, x_4, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); -return x_22; +x_19 = l_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo(x_1, x_2, x_3, x_4, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +return x_19; } else { -lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_23 = l_Lean_Elab_getResetInfoTrees___at_Lean_Elab_Tactic_withTacticInfoContext___spec__2___rarg(x_17, x_18); -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -x_25 = lean_ctor_get(x_23, 1); -lean_inc(x_25); -lean_dec(x_23); -x_26 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_ElimApp_evalAlts___lambda__1___boxed), 18, 9); -lean_closure_set(x_26, 0, x_24); -lean_closure_set(x_26, 1, x_1); -lean_closure_set(x_26, 2, x_2); -lean_closure_set(x_26, 3, x_3); -lean_closure_set(x_26, 4, x_4); -lean_closure_set(x_26, 5, x_6); -lean_closure_set(x_26, 6, x_7); -lean_closure_set(x_26, 7, x_8); -lean_closure_set(x_26, 8, x_9); -x_27 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_ElimApp_evalAlts___lambda__2___boxed), 11, 1); -lean_closure_set(x_27, 0, x_5); -x_28 = l_Lean_Elab_withInfoTreeContext___at_Lean_Elab_Tactic_evalTactic_eval___spec__3(x_26, x_27, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_25); -return x_28; +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_20 = l_Lean_Elab_getResetInfoTrees___at_Lean_Elab_Tactic_withTacticInfoContext___spec__2___rarg(x_17, 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_closure((void*)(l_Lean_Elab_Tactic_ElimApp_evalAlts___lambda__1___boxed), 18, 9); +lean_closure_set(x_23, 0, x_21); +lean_closure_set(x_23, 1, x_1); +lean_closure_set(x_23, 2, x_2); +lean_closure_set(x_23, 3, x_3); +lean_closure_set(x_23, 4, x_4); +lean_closure_set(x_23, 5, x_6); +lean_closure_set(x_23, 6, x_7); +lean_closure_set(x_23, 7, x_8); +lean_closure_set(x_23, 8, x_9); +x_24 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_ElimApp_evalAlts___lambda__2___boxed), 11, 1); +lean_closure_set(x_24, 0, x_5); +x_25 = l_Lean_Elab_withInfoTreeContext___at_Lean_Elab_Tactic_evalTactic_eval___spec__3(x_23, x_24, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_22); +return x_25; } } } @@ -20687,42 +20673,101 @@ uint8_t x_2; x_2 = l_Lean_Syntax_isNone(x_1); if (x_2 == 0) { -lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +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; x_3 = lean_unsigned_to_nat(0u); x_4 = l_Lean_Syntax_getArg(x_1, x_3); -x_5 = l_Lean_Syntax_getArgs(x_4); -x_6 = lean_unsigned_to_nat(2u); -x_7 = l_Array_toSubarray___rarg(x_5, x_3, x_6); -x_8 = l_Array_ofSubarray___rarg(x_7); +x_5 = lean_unsigned_to_nat(2u); +x_6 = l_Lean_Syntax_getArg(x_4, x_5); +x_7 = l_Lean_Syntax_getNumArgs(x_6); +x_8 = lean_nat_dec_lt(x_3, x_7); lean_dec(x_7); -x_9 = lean_box(2); -x_10 = l_Lean_Elab_Tactic_withCaseRef___at_Lean_Elab_Tactic_evalAlt___spec__1___closed__2; -x_11 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_11, 0, x_9); -lean_ctor_set(x_11, 1, x_10); -lean_ctor_set(x_11, 2, x_8); -x_12 = l_Lean_Syntax_getArg(x_4, x_6); +x_9 = l_Lean_Syntax_getArgs(x_4); +x_10 = l_Array_toSubarray___rarg(x_9, x_3, x_5); +x_11 = l_Array_ofSubarray___rarg(x_10); +lean_dec(x_10); +x_12 = lean_box(2); +x_13 = l_Lean_Elab_Tactic_withCaseRef___at_Lean_Elab_Tactic_evalAlt___spec__1___closed__2; +x_14 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_14, 0, x_12); +lean_ctor_set(x_14, 1, x_13); +lean_ctor_set(x_14, 2, x_11); +if (x_8 == 0) +{ +lean_object* x_15; lean_object* x_16; +lean_dec(x_6); +x_15 = l_Lean_Syntax_getArg(x_4, x_3); lean_dec(x_4); -x_13 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_13, 0, x_11); -lean_ctor_set(x_13, 1, x_12); -return x_13; +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_14); +lean_ctor_set(x_16, 1, x_15); +return x_16; } else { -lean_object* x_14; -x_14 = l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_withAltsOfOptInductionAlts___rarg___lambda__1___closed__2; -return x_14; +lean_object* x_17; +lean_dec(x_4); +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_14); +lean_ctor_set(x_17, 1, x_6); +return x_17; +} +} +else +{ +lean_object* x_18; +x_18 = l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_withAltsOfOptInductionAlts___rarg___lambda__1___closed__2; +return x_18; } } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_withAltsOfOptInductionAlts___rarg___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +static lean_object* _init_l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_withAltsOfOptInductionAlts___rarg___lambda__2___closed__1() { _start: { -lean_object* x_12; lean_object* x_13; -x_12 = l_Lean_Syntax_getArgs(x_2); -x_13 = lean_apply_10(x_1, x_12, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_13; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Elab_Tactic_ElimApp_mkElimApp___closed__1; +x_2 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_withAltsOfOptInductionAlts___rarg___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +uint8_t x_13; +x_13 = l_Lean_Syntax_isNone(x_1); +if (x_13 == 0) +{ +lean_object* x_14; uint8_t x_15; +x_14 = l_Lean_Elab_Tactic_withCaseRef___at_Lean_Elab_Tactic_evalAlt___spec__1___closed__2; +lean_inc(x_3); +x_15 = l_Lean_Syntax_isOfKind(x_3, x_14); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; +lean_dec(x_3); +x_16 = l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_withAltsOfOptInductionAlts___rarg___lambda__2___closed__1; +x_17 = lean_apply_10(x_2, x_16, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_17; +} +else +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_18 = l_Lean_Syntax_getArgs(x_3); +lean_dec(x_3); +x_19 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_19, 0, x_18); +x_20 = lean_apply_10(x_2, x_19, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_20; +} +} +else +{ +lean_object* x_21; lean_object* x_22; +lean_dec(x_3); +x_21 = lean_box(0); +x_22 = lean_apply_10(x_2, x_21, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_22; +} } } static lean_object* _init_l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_withAltsOfOptInductionAlts___rarg___closed__1() { @@ -20737,8 +20782,10 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tac _start: { lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_12 = lean_alloc_closure((void*)(l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_withAltsOfOptInductionAlts___rarg___lambda__2___boxed), 11, 1); -lean_closure_set(x_12, 0, x_2); +lean_inc(x_1); +x_12 = lean_alloc_closure((void*)(l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_withAltsOfOptInductionAlts___rarg___lambda__2___boxed), 12, 2); +lean_closure_set(x_12, 0, x_1); +lean_closure_set(x_12, 1, x_2); x_13 = l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_withAltsOfOptInductionAlts___rarg___closed__1; x_14 = l_Lean_Elab_Term_withNarrowedTacticReuse___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_withAltsOfOptInductionAlts___spec__1___rarg(x_13, x_12, x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); return x_14; @@ -20761,13 +20808,13 @@ lean_dec(x_1); return x_2; } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_withAltsOfOptInductionAlts___rarg___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_withAltsOfOptInductionAlts___rarg___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -lean_object* x_12; -x_12 = l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_withAltsOfOptInductionAlts___rarg___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_2); -return x_12; +lean_object* x_13; +x_13 = l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_withAltsOfOptInductionAlts___rarg___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_1); +return x_13; } } static lean_object* _init_l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_getOptPreTacOfOptInductionAlts___closed__1() { @@ -25286,7 +25333,7 @@ lean_dec(x_2); return x_10; } } -static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572____closed__1() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704____closed__1() { _start: { lean_object* x_1; @@ -25294,7 +25341,7 @@ x_1 = lean_mk_string_unchecked("tactic", 6, 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572____closed__2() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704____closed__2() { _start: { lean_object* x_1; @@ -25302,17 +25349,17 @@ x_1 = lean_mk_string_unchecked("customEliminators", 17, 17); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572____closed__3() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572____closed__1; -x_2 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572____closed__2; +x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704____closed__1; +x_2 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704____closed__2; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572____closed__4() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704____closed__4() { _start: { lean_object* x_1; @@ -25320,13 +25367,13 @@ x_1 = lean_mk_string_unchecked("enable using custom eliminators in the 'inductio return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572____closed__5() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704____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_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572____closed__1; -x_3 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572____closed__4; +x_2 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704____closed__1; +x_3 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704____closed__4; x_4 = lean_box(x_1); x_5 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_5, 0, x_4); @@ -25335,26 +25382,26 @@ lean_ctor_set(x_5, 2, x_3); return x_5; } } -static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572____closed__6() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_getAltName___closed__1; x_2 = l_Lean_Elab_Term_withoutTacticIncrementality___at_Lean_Elab_Tactic_ElimApp_evalAlts_applyPreTac___spec__1___closed__2; x_3 = l_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo___lambda__2___closed__1; -x_4 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572____closed__1; -x_5 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572____closed__2; +x_4 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704____closed__1; +x_5 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704____closed__2; x_6 = l_Lean_Name_mkStr5(x_1, x_2, x_3, x_4, x_5); return x_6; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704_(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_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572____closed__3; -x_3 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572____closed__5; -x_4 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572____closed__6; +x_2 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704____closed__3; +x_3 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704____closed__5; +x_4 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704____closed__6; x_5 = l_Lean_Option_register___at_Lean_Elab_initFn____x40_Lean_Elab_AutoBound___hyg_6____spec__1(x_2, x_3, x_4, x_1); return x_5; } @@ -30801,7 +30848,7 @@ static lean_object* _init_l_Lean_Elab_Tactic_elabCasesTargets___lambda__1___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_Tactic_elabCasesTargets___lambda__1___closed__4; x_2 = l_Lean_Elab_Tactic_elabCasesTargets___lambda__1___closed__5; -x_3 = lean_unsigned_to_nat(744u); +x_3 = lean_unsigned_to_nat(752u); x_4 = lean_unsigned_to_nat(8u); x_5 = l_Lean_Elab_Tactic_elabCasesTargets___lambda__1___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -33151,7 +33198,7 @@ x_3 = l_Lean_Elab_addBuiltinIncrementalElab(x_2, x_1); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__1() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -33161,7 +33208,7 @@ x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__2() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -33171,27 +33218,27 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__3() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__2; +x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__2; x_2 = l_Lean_Elab_Term_withoutTacticIncrementality___at_Lean_Elab_Tactic_ElimApp_evalAlts_applyPreTac___spec__1___closed__2; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__4() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__3; +x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__3; x_2 = l_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo___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_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__5() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__5() { _start: { lean_object* x_1; @@ -33199,17 +33246,17 @@ x_1 = lean_mk_string_unchecked("initFn", 6, 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__6() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__4; -x_2 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__5; +x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__4; +x_2 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__5; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__7() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__7() { _start: { lean_object* x_1; @@ -33217,47 +33264,47 @@ x_1 = lean_mk_string_unchecked("_@", 2, 2); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__8() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__6; -x_2 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__7; +x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__6; +x_2 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__7; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__9() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__8; +x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__8; x_2 = l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_getAltName___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__10() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__9; +x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__9; x_2 = l_Lean_Elab_Term_withoutTacticIncrementality___at_Lean_Elab_Tactic_ElimApp_evalAlts_applyPreTac___spec__1___closed__2; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__11() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__10; +x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__10; x_2 = l_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo___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_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__12() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__12() { _start: { lean_object* x_1; @@ -33265,17 +33312,17 @@ x_1 = lean_mk_string_unchecked("Induction", 9, 9); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__13() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__11; -x_2 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__12; +x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__11; +x_2 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__12; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__14() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__14() { _start: { lean_object* x_1; @@ -33283,33 +33330,33 @@ x_1 = lean_mk_string_unchecked("_hyg", 4, 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__15() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__13; -x_2 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__14; +x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__13; +x_2 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__14; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__16() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__15; -x_2 = lean_unsigned_to_nat(11631u); +x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__15; +x_2 = lean_unsigned_to_nat(11763u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763_(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; -x_2 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__1; +x_2 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__1; x_3 = 0; -x_4 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__16; +x_4 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__16; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); if (lean_obj_tag(x_5) == 0) { @@ -33578,6 +33625,8 @@ l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_withAltsOfOptInductio lean_mark_persistent(l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_withAltsOfOptInductionAlts___rarg___lambda__1___closed__1); l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_withAltsOfOptInductionAlts___rarg___lambda__1___closed__2 = _init_l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_withAltsOfOptInductionAlts___rarg___lambda__1___closed__2(); lean_mark_persistent(l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_withAltsOfOptInductionAlts___rarg___lambda__1___closed__2); +l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_withAltsOfOptInductionAlts___rarg___lambda__2___closed__1 = _init_l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_withAltsOfOptInductionAlts___rarg___lambda__2___closed__1(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_withAltsOfOptInductionAlts___rarg___lambda__2___closed__1); l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_withAltsOfOptInductionAlts___rarg___closed__1 = _init_l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_withAltsOfOptInductionAlts___rarg___closed__1(); lean_mark_persistent(l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_withAltsOfOptInductionAlts___rarg___closed__1); l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_getOptPreTacOfOptInductionAlts___closed__1 = _init_l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_getOptPreTacOfOptInductionAlts___closed__1(); @@ -33630,19 +33679,19 @@ l_Lean_Elab_Tactic_getInductiveValFromMajor___lambda__2___closed__14 = _init_l_L lean_mark_persistent(l_Lean_Elab_Tactic_getInductiveValFromMajor___lambda__2___closed__14); l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_elabTermForElim___closed__1 = _init_l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_elabTermForElim___closed__1(); lean_mark_persistent(l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_elabTermForElim___closed__1); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572____closed__1 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572____closed__1(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572____closed__1); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572____closed__2 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572____closed__2(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572____closed__2); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572____closed__3 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572____closed__3(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572____closed__3); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572____closed__4 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572____closed__4(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572____closed__4); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572____closed__5 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572____closed__5(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572____closed__5); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572____closed__6 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572____closed__6(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572____closed__6); -if (builtin) {res = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572_(lean_io_mk_world()); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704____closed__1 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704____closed__1(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704____closed__1); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704____closed__2 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704____closed__2(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704____closed__2); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704____closed__3 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704____closed__3(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704____closed__3); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704____closed__4 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704____closed__4(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704____closed__4); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704____closed__5 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704____closed__5(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704____closed__5); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704____closed__6 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704____closed__6(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704____closed__6); +if (builtin) {res = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l_Lean_Elab_Tactic_tactic_customEliminators = lean_io_result_get_value(res); lean_mark_persistent(l_Lean_Elab_Tactic_tactic_customEliminators); @@ -33769,39 +33818,39 @@ lean_dec_ref(res); }if (builtin) {res = l___regBuiltin_Lean_Elab_Tactic_evalCases__2(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -}l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__1 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__1(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__1); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__2 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__2(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__2); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__3 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__3(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__3); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__4 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__4(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__4); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__5 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__5(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__5); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__6 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__6(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__6); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__7 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__7(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__7); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__8 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__8(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__8); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__9 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__9(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__9); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__10 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__10(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__10); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__11 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__11(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__11); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__12 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__12(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__12); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__13 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__13(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__13); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__14 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__14(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__14); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__15 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__15(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__15); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__16 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__16(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__16); -if (builtin) {res = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631_(lean_io_mk_world()); +}l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__1 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__1(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__1); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__2 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__2(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__2); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__3 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__3(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__3); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__4 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__4(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__4); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__5 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__5(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__5); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__6 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__6(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__6); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__7 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__7(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__7); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__8 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__8(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__8); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__9 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__9(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__9); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__10 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__10(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__10); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__11 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__11(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__11); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__12 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__12(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__12); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__13 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__13(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__13); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__14 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__14(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__14); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__15 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__15(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__15); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__16 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__16(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__16); +if (builtin) {res = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763_(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/NormCast.c b/stage0/stdlib/Lean/Elab/Tactic/NormCast.c index f4b6f22866bc..e1bdeb14a200 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/NormCast.c +++ b/stage0/stdlib/Lean/Elab/Tactic/NormCast.c @@ -202,8 +202,6 @@ static lean_object* l___regBuiltin_Lean_Elab_Tactic_NormCast_elabModCast_declRan lean_object* l_Lean_Meta_Simp_Result_mkEqSymm(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_NormCast_mkCoe(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_NormCast_elabModCast__1___closed__1; -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_Lean_Elab_Tactic_NormCast_elabNormCastConfig___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_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*); @@ -391,6 +389,7 @@ 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 lean_object* l_Lean_Elab_Tactic_NormCast_elabModCast___lambda__2___closed__3; static double l_Lean_withTraceNode___at_Lean_Elab_Tactic_NormCast_proveEqUsingDown___spec__1___lambda__4___closed__5; +lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_elabLetMVar___spec__1___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_NormCast_elabAddElim___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalExact___spec__1___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_NormCast_derive___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -465,6 +464,7 @@ lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); 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*); static lean_object* l_Lean_Elab_Tactic_NormCast_elabModCast___lambda__2___closed__4; +lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Widget_elabShowPanelWidgetsCmd___spec__1___rarg(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_NormCast_evalPushCast_declRange__1___closed__4; static lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_NormCast_proveEqUsingDown___spec__1___lambda__4___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_NormCast_splittingProcedure___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*); @@ -14562,7 +14562,7 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_12 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_elabForall___spec__1___rarg(x_9); +x_12 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_elabLetMVar___spec__1___rarg(x_9); return x_12; } else @@ -17376,7 +17376,7 @@ if (x_6 == 0) { lean_object* x_7; lean_dec(x_1); -x_7 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabAuxDef___spec__1___rarg(x_4); +x_7 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Widget_elabShowPanelWidgetsCmd___spec__1___rarg(x_4); return x_7; } else @@ -17392,7 +17392,7 @@ if (x_11 == 0) { lean_object* x_12; lean_dec(x_9); -x_12 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabAuxDef___spec__1___rarg(x_4); +x_12 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Widget_elabShowPanelWidgetsCmd___spec__1___rarg(x_4); return x_12; } else diff --git a/stage0/stdlib/Lean/Elab/Tactic/RCases.c b/stage0/stdlib/Lean/Elab/Tactic/RCases.c index 971d3a6bf7fb..1f39e01f9e7e 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/RCases.c +++ b/stage0/stdlib/Lean/Elab/Tactic/RCases.c @@ -32374,7 +32374,7 @@ lean_dec(x_6); x_37 = l_Lean_Elab_Tactic_RCases_RCasesPatt_parse(x_36, x_11, x_12, x_13, x_14, x_15); 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_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; 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; lean_object* x_55; uint8_t x_56; lean_object* x_57; lean_object* x_58; x_38 = lean_ctor_get(x_37, 0); lean_inc(x_38); x_39 = lean_ctor_get(x_37, 1); @@ -32382,36 +32382,80 @@ lean_inc(x_39); lean_dec(x_37); x_40 = l_Lean_Elab_Tactic_RCases_RCasesPatt_typed_x3f(x_5, x_38, x_7); lean_dec(x_7); -x_41 = l_Lean_Elab_Tactic_RCases_RCasesPatt_name_x3f(x_40); -if (lean_obj_tag(x_41) == 0) +x_41 = l_Lean_Elab_Tactic_RCases_RCasesPatt_ref(x_40); +x_42 = l_Lean_Elab_Tactic_RCases_RCasesPatt_name_x3f(x_40); +x_43 = lean_ctor_get(x_13, 0); +lean_inc(x_43); +x_44 = lean_ctor_get(x_13, 1); +lean_inc(x_44); +x_45 = lean_ctor_get(x_13, 2); +lean_inc(x_45); +x_46 = lean_ctor_get(x_13, 3); +lean_inc(x_46); +x_47 = lean_ctor_get(x_13, 4); +lean_inc(x_47); +x_48 = lean_ctor_get(x_13, 5); +lean_inc(x_48); +x_49 = lean_ctor_get(x_13, 6); +lean_inc(x_49); +x_50 = lean_ctor_get(x_13, 7); +lean_inc(x_50); +x_51 = lean_ctor_get(x_13, 8); +lean_inc(x_51); +x_52 = lean_ctor_get(x_13, 9); +lean_inc(x_52); +x_53 = lean_ctor_get(x_13, 10); +lean_inc(x_53); +x_54 = lean_ctor_get_uint8(x_13, sizeof(void*)*12); +x_55 = lean_ctor_get(x_13, 11); +lean_inc(x_55); +x_56 = lean_ctor_get_uint8(x_13, sizeof(void*)*12 + 1); +x_57 = l_Lean_replaceRef(x_41, x_48); +lean_dec(x_48); +lean_dec(x_41); +x_58 = lean_alloc_ctor(0, 12, 2); +lean_ctor_set(x_58, 0, x_43); +lean_ctor_set(x_58, 1, x_44); +lean_ctor_set(x_58, 2, x_45); +lean_ctor_set(x_58, 3, x_46); +lean_ctor_set(x_58, 4, x_47); +lean_ctor_set(x_58, 5, x_57); +lean_ctor_set(x_58, 6, x_49); +lean_ctor_set(x_58, 7, x_50); +lean_ctor_set(x_58, 8, x_51); +lean_ctor_set(x_58, 9, x_52); +lean_ctor_set(x_58, 10, x_53); +lean_ctor_set(x_58, 11, x_55); +lean_ctor_set_uint8(x_58, sizeof(void*)*12, x_54); +lean_ctor_set_uint8(x_58, sizeof(void*)*12 + 1, x_56); +if (lean_obj_tag(x_42) == 0) { -lean_object* x_42; lean_object* x_43; -x_42 = l_Lean_Elab_Tactic_RCases_RCasesPatt_instInhabited___closed__2; +lean_object* x_59; lean_object* x_60; +x_59 = l_Lean_Elab_Tactic_RCases_RCasesPatt_instInhabited___closed__2; lean_inc(x_14); -lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); -x_43 = l_Lean_MVarId_intro(x_1, x_42, x_11, x_12, x_13, x_14, x_39); -if (lean_obj_tag(x_43) == 0) +x_60 = l_Lean_MVarId_intro(x_1, x_59, x_11, x_12, x_58, x_14, x_39); +if (lean_obj_tag(x_60) == 0) { -lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; -x_44 = lean_ctor_get(x_43, 0); -lean_inc(x_44); -x_45 = lean_ctor_get(x_43, 1); -lean_inc(x_45); -lean_dec(x_43); -x_46 = lean_ctor_get(x_44, 0); -lean_inc(x_46); -x_47 = lean_ctor_get(x_44, 1); -lean_inc(x_47); -lean_dec(x_44); -x_48 = l_Lean_Expr_fvar___override(x_46); -x_49 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg(x_47, x_2, x_3, x_48, x_4, x_40, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_45); -return x_49; +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_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); +x_64 = lean_ctor_get(x_61, 1); +lean_inc(x_64); +lean_dec(x_61); +x_65 = l_Lean_Expr_fvar___override(x_63); +x_66 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg(x_64, x_2, x_3, x_65, x_4, x_40, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_62); +return x_66; } else { -uint8_t x_50; +uint8_t x_67; lean_dec(x_40); lean_dec(x_14); lean_dec(x_13); @@ -32423,57 +32467,56 @@ lean_dec(x_8); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_50 = !lean_is_exclusive(x_43); -if (x_50 == 0) +x_67 = !lean_is_exclusive(x_60); +if (x_67 == 0) { -return x_43; +return x_60; } else { -lean_object* x_51; lean_object* x_52; lean_object* x_53; -x_51 = lean_ctor_get(x_43, 0); -x_52 = lean_ctor_get(x_43, 1); -lean_inc(x_52); -lean_inc(x_51); -lean_dec(x_43); -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_68; lean_object* x_69; lean_object* x_70; +x_68 = lean_ctor_get(x_60, 0); +x_69 = lean_ctor_get(x_60, 1); +lean_inc(x_69); +lean_inc(x_68); +lean_dec(x_60); +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_54; lean_object* x_55; -x_54 = lean_ctor_get(x_41, 0); -lean_inc(x_54); -lean_dec(x_41); +lean_object* x_71; lean_object* x_72; +x_71 = lean_ctor_get(x_42, 0); +lean_inc(x_71); +lean_dec(x_42); lean_inc(x_14); -lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); -x_55 = l_Lean_MVarId_intro(x_1, x_54, x_11, x_12, x_13, x_14, x_39); -if (lean_obj_tag(x_55) == 0) +x_72 = l_Lean_MVarId_intro(x_1, x_71, x_11, x_12, x_58, x_14, x_39); +if (lean_obj_tag(x_72) == 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_55, 0); -lean_inc(x_56); -x_57 = lean_ctor_get(x_55, 1); -lean_inc(x_57); -lean_dec(x_55); -x_58 = lean_ctor_get(x_56, 0); -lean_inc(x_58); -x_59 = lean_ctor_get(x_56, 1); -lean_inc(x_59); -lean_dec(x_56); -x_60 = l_Lean_Expr_fvar___override(x_58); -x_61 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg(x_59, x_2, x_3, x_60, x_4, x_40, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_57); -return x_61; +lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; +x_73 = lean_ctor_get(x_72, 0); +lean_inc(x_73); +x_74 = lean_ctor_get(x_72, 1); +lean_inc(x_74); +lean_dec(x_72); +x_75 = lean_ctor_get(x_73, 0); +lean_inc(x_75); +x_76 = lean_ctor_get(x_73, 1); +lean_inc(x_76); +lean_dec(x_73); +x_77 = l_Lean_Expr_fvar___override(x_75); +x_78 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg(x_76, x_2, x_3, x_77, x_4, x_40, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_74); +return x_78; } else { -uint8_t x_62; +uint8_t x_79; lean_dec(x_40); lean_dec(x_14); lean_dec(x_13); @@ -32485,30 +32528,30 @@ lean_dec(x_8); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_62 = !lean_is_exclusive(x_55); -if (x_62 == 0) +x_79 = !lean_is_exclusive(x_72); +if (x_79 == 0) { -return x_55; +return x_72; } else { -lean_object* x_63; lean_object* x_64; lean_object* x_65; -x_63 = lean_ctor_get(x_55, 0); -x_64 = lean_ctor_get(x_55, 1); -lean_inc(x_64); -lean_inc(x_63); -lean_dec(x_55); -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_80; lean_object* x_81; lean_object* x_82; +x_80 = lean_ctor_get(x_72, 0); +x_81 = lean_ctor_get(x_72, 1); +lean_inc(x_81); +lean_inc(x_80); +lean_dec(x_72); +x_82 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_82, 0, x_80); +lean_ctor_set(x_82, 1, x_81); +return x_82; } } } } else { -uint8_t x_66; +uint8_t x_83; lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); @@ -32522,23 +32565,23 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_66 = !lean_is_exclusive(x_37); -if (x_66 == 0) +x_83 = !lean_is_exclusive(x_37); +if (x_83 == 0) { return x_37; } else { -lean_object* x_67; lean_object* x_68; lean_object* x_69; -x_67 = lean_ctor_get(x_37, 0); -x_68 = lean_ctor_get(x_37, 1); -lean_inc(x_68); -lean_inc(x_67); +lean_object* x_84; lean_object* x_85; lean_object* x_86; +x_84 = lean_ctor_get(x_37, 0); +x_85 = lean_ctor_get(x_37, 1); +lean_inc(x_85); +lean_inc(x_84); lean_dec(x_37); -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; +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; } } } diff --git a/stage0/stdlib/Lean/Linter/MissingDocs.c b/stage0/stdlib/Lean/Linter/MissingDocs.c index 83759bb8497e..b5c70fba95ba 100644 --- a/stage0/stdlib/Lean/Linter/MissingDocs.c +++ b/stage0/stdlib/Lean/Linter/MissingDocs.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Lean.Linter.MissingDocs -// Imports: Lean.Meta.Tactic.Simp.RegisterCommand Lean.Elab.Command Lean.Elab.SetOption Lean.Linter.Util +// Imports: Lean.Parser.Syntax Lean.Meta.Tactic.Simp.RegisterCommand Lean.Elab.Command Lean.Elab.SetOption Lean.Linter.Util #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -8777,6 +8777,7 @@ x_4 = l_Lean_Linter_MissingDocs_addBuiltinHandler(x_2, x_3, x_1); return x_4; } } +lean_object* initialize_Lean_Parser_Syntax(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_Tactic_Simp_RegisterCommand(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Elab_Command(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Elab_SetOption(uint8_t builtin, lean_object*); @@ -8786,6 +8787,9 @@ LEAN_EXPORT lean_object* initialize_Lean_Linter_MissingDocs(uint8_t builtin, lea lean_object * res; if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); _G_initialized = true; +res = initialize_Lean_Parser_Syntax(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); res = initialize_Lean_Meta_Tactic_Simp_RegisterCommand(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/Cases.c b/stage0/stdlib/Lean/Meta/Tactic/Cases.c index 662aaea31e43..b5753efe8309 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Cases.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Cases.c @@ -14,28 +14,27 @@ extern "C" { #endif static lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_toByCasesSubgoal___closed__1; -static lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__6___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_generalizeIndices___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); lean_object* l_Lean_Meta_FVarSubst_apply(lean_object*, lean_object*); static lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__56___closed__3; static lean_object* l_Lean_MVarId_byCasesDec___closed__3; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__14; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_mkEqAndProof(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_generalizeIndices___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_EXPORT lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_inductionCasesOn___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__44___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_MVarId_casesAnd___closed__3; lean_object* l_Lean_Meta_throwTacticEx___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___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_Tactic_Cases_0__Lean_Meta_withNewEqs_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_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_mk_empty_array_with_capacity(lean_object*); lean_object* l_Lean_mkAppN(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900_(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_elimAuxIndices___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__25(lean_object*, lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__45___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyMAux___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__40___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Meta_Cases_unifyEqs_x3f___spec__1___closed__4; +LEAN_EXPORT lean_object* l_Lean_Meta_withNewEqs_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* l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__37(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -52,6 +51,7 @@ extern lean_object* l_Lean_maxRecDepthErrorMessage; lean_object* l_Lean_MetavarContext_getDecl(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyMAux___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__24___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__4(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__3; lean_object* l_Lean_MVarId_checkNotAssigned(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Exception_isInterrupt(lean_object*); lean_object* l_Lean_mkAppB(lean_object*, lean_object*, lean_object*); @@ -62,6 +62,7 @@ static lean_object* l_Lean_Meta_Cases_unifyEqs_x3f___closed__1; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__41___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__30___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__1; uint8_t l_Lean_Expr_isAppOfArity(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Cases_cases___lambda__2___closed__2; static lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_mkEqAndProof___closed__2; @@ -76,8 +77,8 @@ lean_object* lean_array_push(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyM___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__31___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyM___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__53(lean_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_Cases_0__Lean_Meta_Cases_inductionCasesOn___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Nat_anyTR_loop___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__2; static lean_object* l_Lean_Meta_generalizeTargetsEq___lambda__3___closed__2; lean_object* lean_mk_array(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__56(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -90,29 +91,29 @@ lean_object* lean_array_fget(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__35(lean_object*, lean_object*, lean_object*, size_t, size_t); extern lean_object* l_Lean_casesOnSuffix; LEAN_EXPORT lean_object* l_Lean_MVarId_byCasesDec(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_Tactic_Cases___hyg_3794____closed__12; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__41(lean_object*, lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_MVarId_casesRec___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*); static lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_mkEqAndProof___closed__6; lean_object* l_Lean_Meta_FVarSubst_get(lean_object*, lean_object*); lean_object* l_Lean_MVarId_getTag(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_withNewEqs___rarg___closed__1; lean_object* l_Lean_Expr_fvarId_x21(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyM___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__47___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_environment_find(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__10; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_MVarId_casesRec___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*); +static lean_object* l_Lean_Meta_generalizeIndices_x27___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_throwInductiveTypeExpected___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_throwNestedTacticEx___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__5; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__8; LEAN_EXPORT lean_object* l_Lean_Meta_generalizeTargetsEq___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_Expr_withAppAux___at_Lean_Meta_generalizeIndices___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_Lean_PersistentArray_anyMAux___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__8___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_forallTelescopeReducing___at_Lean_Meta_getParamNames___spec__2___rarg(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_withNewEqs_loop___rarg___closed__2; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_elimAuxIndices___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_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__25___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_throwInductiveTypeExpected___rarg___closed__1; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__3(lean_object*, lean_object*, size_t, size_t); +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___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*); static lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__56___closed__1; lean_object* l_Nat_nextPowerOfTwo_go(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_generalizeTargetsEq___closed__1; @@ -129,31 +130,36 @@ LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases LEAN_EXPORT lean_object* l_Lean_Meta_Cases_cases___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_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__57(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_MVarId_casesRec(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__5___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_withNewEqs_loop___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_elimAuxIndices___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Meta_Cases_unifyEqs_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__18___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_throwInductiveTypeExpected___rarg___closed__4; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__27___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___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*); LEAN_EXPORT lean_object* l_Lean_Meta_Cases_cases___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_MVarId_casesAnd___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Meta_Cases_unifyEqs_x3f___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__37___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Cases_cases___lambda__2___closed__3; static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Meta_Cases_unifyEqs_x3f___spec__1___closed__6; lean_object* l_Lean_AssocList_erase___at_Lean_Meta_FVarSubst_erase___spec__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs_loop___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_toCasesSubgoals___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__52___boxed(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_hasMVar(lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__55___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs_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* l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__9; +static lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__7___closed__2; lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyMAux___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__40(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___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*); static lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_mkEqAndProof___closed__1; -static lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__5; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_of_nat(lean_object*); +static lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__2; static lean_object* l_Lean_Meta_generalizeTargetsEq___closed__2; LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyM___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__39(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_throwInductiveTypeExpected___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -161,18 +167,18 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_h LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__36___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyM___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__23(lean_object*, lean_object*, lean_object*); static lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__56___closed__2; -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___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_PersistentArray_anyM___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__7___boxed(lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_take(lean_object*, lean_object*); static lean_object* l_Lean_MVarId_substEqs___closed__1; static lean_object* l_Lean_Meta_generalizeTargetsEq___lambda__3___closed__1; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__10; static lean_object* l_Lean_MVarId_byCases___closed__1; static lean_object* l_Lean_Meta_Cases_cases___lambda__2___closed__10; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__9(lean_object*, lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyMAux___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__48___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs_loop___rarg___closed__1; +static lean_object* l_Lean_Meta_generalizeIndices_x27___lambda__1___closed__2; uint8_t lean_expr_eqv(lean_object*, lean_object*); -static lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__6___closed__2; +LEAN_EXPORT lean_object* l_Lean_Meta_generalizeIndices_x27___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__36(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_instantiateMVars___at___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___spec__7(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -182,17 +188,18 @@ static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Meta_Cases_unifyEqs_x3f_ lean_object* l_Lean_MVarId_getType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isEq(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyMAux___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__54___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_Tactic_Cases___hyg_3794____closed__11; LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyM___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__53___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_Cases_cases___lambda__2___closed__4; +static lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__6; LEAN_EXPORT lean_object* l_Lean_MVarId_byCases(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___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__51(lean_object*, lean_object*, lean_object*, size_t, size_t); lean_object* l_Lean_registerTraceClass(lean_object*, uint8_t, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___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_generalizeIndices_x27(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_SavedState_restore(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_generalizeTargetsEq___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_PersistentArray_forIn___at_Lean_MVarId_casesRec___spec__3___lambda__1(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___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__30(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_throwError___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_toByCasesSubgoal___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MVarId_acyclic(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_toByCasesSubgoal___closed__2; @@ -200,17 +207,16 @@ LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_v LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at_Lean_MVarId_casesRec___spec__3___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofFormat(lean_object*); uint8_t l_Lean_Expr_isMVar(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794_(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*); static lean_object* l_Lean_Meta_Cases_cases___lambda__2___closed__9; lean_object* l_Lean_Meta_getLevel(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_FVarId_getDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_outOfBounds___rarg(lean_object*); -static lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__6___closed__4; static lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_mkEqAndProof___closed__7; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_MVarId_casesRec___spec__5(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_st_ref_get(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_inductionCasesOn___lambda__2___closed__1; +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___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_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_mkCasesContext_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___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__38___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Nat_anyTR_loop___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -220,11 +226,10 @@ LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyM___at___private_Lean_Meta_Ta LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_MVarId_casesRec___spec__7(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_array_to_list(lean_object*); static lean_object* l_Lean_Meta_Cases_cases___lambda__2___closed__8; -static lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__1; lean_object* l_Lean_MVarId_induction(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__3; lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__44(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__4; LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__14(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__10(lean_object*, lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l_Array_mapFinIdxM_map___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_toCasesSubgoals___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -235,22 +240,24 @@ extern lean_object* l_Lean_levelZero; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_unifyCasesEqs___spec__2(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_MVarId_byCases___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_instInhabitedExpr; +LEAN_EXPORT lean_object* l_Lean_Meta_withNewEqs_loop(lean_object*); lean_object* l_Lean_Meta_whnfD(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_casesAnd(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyMAux___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__16___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___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* l_List_lengthTRAux___rarg(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyMAux___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__8(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_elimAuxIndices___spec__1(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__13; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyM___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__31(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__22(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__2; uint8_t lean_name_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__58___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); lean_object* l_Lean_MVarId_clear(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__58(lean_object*, size_t, size_t); -static lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__4; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__6; LEAN_EXPORT lean_object* l_Lean_observing_x3f___at_Lean_MVarId_casesRec___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_MVarId_byCasesDec___closed__4; LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_mkCasesContext_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -269,18 +276,22 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_i LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_toByCasesSubgoal___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_MVarId_casesAnd___closed__4; LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyMAux___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__16(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at_Lean_MVarId_casesRec___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__3; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__15; +static lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__7___closed__3; lean_object* l_Lean_MVarId_assert(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_getInductiveUniverseAndParams___closed__1; +static lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__7; static lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_mkEqAndProof___closed__5; static lean_object* l_Lean_Meta_generalizeTargetsEq___lambda__3___closed__4; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_inductionCasesOn___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_generalizeTargetsEq___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_Environment_contains(lean_object*, lean_object*); +static lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__8; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__33(lean_object*, lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l_Lean_MVarId_substEqs(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_withNewEqs_loop___rarg___closed__1; +static lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__5; LEAN_EXPORT lean_object* l_Lean_MVarId_cases(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyM___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__15(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_substEqs___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -293,7 +304,7 @@ LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_ LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__43(lean_object*, lean_object*, lean_object*, size_t, size_t); lean_object* l_Lean_LocalDecl_userName(lean_object*); lean_object* l_Array_extract___rarg(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__11; lean_object* l_Lean_Meta_exactlyOne(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_casesRec___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Cases_cases___closed__1; @@ -301,13 +312,16 @@ lean_object* l_Lean_Meta_introNCore(lean_object*, lean_object*, lean_object*, ui LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_inductionCasesOn___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__34___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyMAux___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__32___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__7; +static lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__5___closed__2; static lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_throwInductiveTypeExpected___rarg___closed__2; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__34(lean_object*, lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__11(lean_object*, lean_object*, lean_object*, size_t, size_t); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__12; lean_object* l_Lean_Meta_ensureAtMostOne(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_generalizeIndices_x27___spec__1___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyMAux___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__32(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at_Lean_MVarId_casesRec___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_Meta_withNewEqs(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyM___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__39___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Meta_Cases_unifyEqs_x3f___spec__1___closed__3; lean_object* l_Lean_Expr_app___override(lean_object*, lean_object*); @@ -324,7 +338,6 @@ static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Meta_Cases_unifyEqs_x3f_ lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_getLocalInstances(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyM___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__23___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__8; LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__20(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); lean_object* l_Lean_indentExpr(lean_object*); @@ -338,9 +351,9 @@ lean_object* l_Lean_Meta_mkForallFVars(lean_object*, lean_object*, uint8_t, uint lean_object* lean_array_set(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* l_Lean_LocalDecl_type(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_withNewEqs___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_generalizeTargetsEq___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_MVarId_byCasesDec___closed__1; -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__15; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_MVarId_casesRec___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___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__12(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_nat_sub(lean_object*, lean_object*); @@ -352,13 +365,9 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Cases_cases(lean_object*, lean_object*, lea static lean_object* l_Lean_MVarId_byCases___closed__2; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__46___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs_loop___rarg___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__28___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__8; lean_object* l_List_reverse___rarg(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___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_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__6; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__18(lean_object*, lean_object*, lean_object*, size_t, size_t); lean_object* lean_array_mk(lean_object*); lean_object* l_Lean_Meta_saturate(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -367,14 +376,15 @@ lean_object* l_Lean_mkEM(lean_object*); static lean_object* l_Lean_Meta_Cases_cases___lambda__2___closed__6; LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyMAux___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__48(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__55(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_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__14; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__13; LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at_Lean_MVarId_casesRec___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__4; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__5; LEAN_EXPORT lean_object* l_Lean_Meta_Cases_cases___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___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_Array_mapFinIdxM_map___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_toCasesSubgoals___spec__1___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_add(size_t, size_t); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__1; LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_MVarId_casesRec___spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___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*); uint8_t l_Lean_Expr_hasFVar(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at_Lean_MVarId_casesRec___spec__4___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_unifyCasesEqs(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -385,23 +395,21 @@ lean_object* lean_array_uget(lean_object*, size_t); lean_object* l_Lean_Expr_fvar___override(lean_object*); size_t lean_array_size(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__14___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__9; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_throwInductiveTypeExpected(lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__7; lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__7___closed__4; lean_object* l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___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_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_throwInductiveTypeExpected___spec__1(lean_object*); +static lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__7___closed__1; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__49(lean_object*, lean_object*, lean_object*, size_t, size_t); lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Meta_addPPExplicitToExposeDiff_visit___spec__4___rarg(lean_object*, uint8_t, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Cases_cases___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_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__57___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_generalizeIndices___lambda__1___closed__2; LEAN_EXPORT uint8_t l_Nat_anyTR_loop___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__2(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_EXPORT lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs(lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_casesAnd___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs_loop(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_throwError___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_toByCasesSubgoal___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_MVarId_casesAnd___closed__2; @@ -416,10 +424,11 @@ static lean_object* l_Lean_MVarId_byCases___closed__3; 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_Meta_Cases_unifyEqs_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_getExprMVarAssignment_x3f___at___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_throwInductiveTypeExpected___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_generalizeIndices___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at_Lean_MVarId_casesRec___spec__4___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_withNewEqs_loop___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_nat_add(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__13(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_inductionCasesOn___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*); @@ -434,25 +443,21 @@ lean_object* l_Lean_Expr_mvarId_x21(lean_object*); static lean_object* l_Lean_MVarId_byCases___closed__4; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__17(lean_object*, lean_object*, lean_object*, size_t, size_t); lean_object* lean_array_uset(lean_object*, size_t, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___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*); lean_object* l___private_Init_Data_Repr_0__Nat_reprFast(lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__19(lean_object*, lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__51___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__3; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__4; static lean_object* l_Lean_MVarId_casesAnd___closed__1; -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___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_EXPORT uint8_t l_Lean_PersistentArray_anyM___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__47(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__6___closed__3; -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__7; LEAN_EXPORT lean_object* l_Nat_anyTR_loop___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_generalizeTargetsEq___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_throwInductiveTypeExpected___rarg___closed__3; -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__6; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__19___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Array_isEmpty___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_generalizeTargetsEq___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_Meta_Tactic_Cases_0__Lean_Meta_Cases_unifyCasesEqs___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_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__2; +static lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__1; LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_throwInductiveTypeExpected___spec__1___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { @@ -1162,7 +1167,7 @@ return x_79; } } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs_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_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l_Lean_Meta_withNewEqs_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_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; @@ -1170,11 +1175,11 @@ x_14 = lean_unsigned_to_nat(1u); x_15 = lean_nat_add(x_1, x_14); x_16 = lean_array_push(x_2, x_8); x_17 = lean_array_push(x_3, x_4); -x_18 = l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs_loop___rarg(x_5, x_6, x_7, x_15, x_16, x_17, x_9, x_10, x_11, x_12, x_13); +x_18 = l_Lean_Meta_withNewEqs_loop___rarg(x_5, x_6, x_7, x_15, x_16, x_17, x_9, x_10, x_11, x_12, x_13); return x_18; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs_loop___rarg___closed__1() { +static lean_object* _init_l_Lean_Meta_withNewEqs_loop___rarg___closed__1() { _start: { lean_object* x_1; @@ -1182,17 +1187,17 @@ x_1 = lean_mk_string_unchecked("h", 1, 1); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs_loop___rarg___closed__2() { +static lean_object* _init_l_Lean_Meta_withNewEqs_loop___rarg___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs_loop___rarg___closed__1; +x_2 = l_Lean_Meta_withNewEqs_loop___rarg___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs_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_EXPORT lean_object* l_Lean_Meta_withNewEqs_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) { _start: { lean_object* x_12; uint8_t x_13; @@ -1238,7 +1243,7 @@ lean_inc(x_23); x_24 = lean_ctor_get(x_21, 1); lean_inc(x_24); lean_dec(x_21); -x_25 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs_loop___rarg___lambda__1___boxed), 13, 7); +x_25 = lean_alloc_closure((void*)(l_Lean_Meta_withNewEqs_loop___rarg___lambda__1___boxed), 13, 7); lean_closure_set(x_25, 0, x_4); lean_closure_set(x_25, 1, x_5); lean_closure_set(x_25, 2, x_6); @@ -1246,7 +1251,7 @@ lean_closure_set(x_25, 3, x_24); lean_closure_set(x_25, 4, x_1); lean_closure_set(x_25, 5, x_2); lean_closure_set(x_25, 6, x_3); -x_26 = l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs_loop___rarg___closed__2; +x_26 = l_Lean_Meta_withNewEqs_loop___rarg___closed__2; x_27 = 0; x_28 = 0; x_29 = l_Lean_Meta_withLocalDecl___at_Lean_Meta_addPPExplicitToExposeDiff_visit___spec__4___rarg(x_26, x_27, x_23, x_25, x_28, x_7, x_8, x_9, x_10, x_22); @@ -1307,7 +1312,7 @@ lean_inc(x_38); x_39 = lean_ctor_get(x_36, 1); lean_inc(x_39); lean_dec(x_36); -x_40 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs_loop___rarg___lambda__1___boxed), 13, 7); +x_40 = lean_alloc_closure((void*)(l_Lean_Meta_withNewEqs_loop___rarg___lambda__1___boxed), 13, 7); lean_closure_set(x_40, 0, x_4); lean_closure_set(x_40, 1, x_5); lean_closure_set(x_40, 2, x_6); @@ -1315,7 +1320,7 @@ lean_closure_set(x_40, 3, x_39); lean_closure_set(x_40, 4, x_1); lean_closure_set(x_40, 5, x_2); lean_closure_set(x_40, 6, x_3); -x_41 = l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs_loop___rarg___closed__2; +x_41 = l_Lean_Meta_withNewEqs_loop___rarg___closed__2; x_42 = 0; x_43 = 0; x_44 = l_Lean_Meta_withLocalDecl___at_Lean_Meta_addPPExplicitToExposeDiff_visit___spec__4___rarg(x_41, x_42, x_38, x_40, x_43, x_7, x_8, x_9, x_10, x_37); @@ -1357,24 +1362,24 @@ return x_48; } } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs_loop(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Meta_withNewEqs_loop(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs_loop___rarg), 11, 0); +x_2 = lean_alloc_closure((void*)(l_Lean_Meta_withNewEqs_loop___rarg), 11, 0); return x_2; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs_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, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l_Lean_Meta_withNewEqs_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, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { lean_object* x_14; -x_14 = l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs_loop___rarg___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +x_14 = l_Lean_Meta_withNewEqs_loop___rarg___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); lean_dec(x_1); return x_14; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs___rarg___closed__1() { +static lean_object* _init_l_Lean_Meta_withNewEqs___rarg___closed__1() { _start: { lean_object* x_1; lean_object* x_2; @@ -1383,21 +1388,21 @@ x_2 = lean_array_mk(x_1); return x_2; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Lean_Meta_withNewEqs___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; x_9 = lean_unsigned_to_nat(0u); -x_10 = l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs___rarg___closed__1; -x_11 = l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs_loop___rarg(x_1, x_2, x_3, x_9, x_10, x_10, x_4, x_5, x_6, x_7, x_8); +x_10 = l_Lean_Meta_withNewEqs___rarg___closed__1; +x_11 = l_Lean_Meta_withNewEqs_loop___rarg(x_1, x_2, x_3, x_9, x_10, x_10, x_4, x_5, x_6, x_7, x_8); return x_11; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Meta_withNewEqs(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs___rarg), 8, 0); +x_2 = lean_alloc_closure((void*)(l_Lean_Meta_withNewEqs___rarg), 8, 0); return x_2; } } @@ -1539,7 +1544,7 @@ lean_inc(x_2); x_10 = lean_alloc_closure((void*)(l_Lean_Meta_generalizeTargetsEq___lambda__1___boxed), 9, 2); lean_closure_set(x_10, 0, x_1); lean_closure_set(x_10, 1, x_2); -x_11 = l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs___rarg(x_3, x_2, x_10, x_5, x_6, x_7, x_8, x_9); +x_11 = l_Lean_Meta_withNewEqs___rarg(x_3, x_2, x_10, x_5, x_6, x_7, x_8, x_9); return x_11; } } @@ -1892,7 +1897,7 @@ lean_dec(x_4); return x_10; } } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { _start: { lean_object* x_16; lean_object* x_17; @@ -2261,52 +2266,50 @@ return x_96; } } } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___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_object* x_12, lean_object* x_13, lean_object* x_14) { +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___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_object* x_12, lean_object* x_13, lean_object* x_14) { _start: { -lean_object* x_15; lean_object* x_16; -x_15 = l_Lean_LocalDecl_toExpr(x_1); +lean_object* x_15; lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_2); -lean_inc(x_15); -x_16 = l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_mkEqAndProof(x_15, x_2, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_16) == 0) +lean_inc(x_1); +x_15 = l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_mkEqAndProof(x_1, x_2, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_15) == 0) { -lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; uint8_t x_25; lean_object* x_26; -x_17 = lean_ctor_get(x_16, 0); +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; uint8_t x_24; lean_object* x_25; +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_15, 1); lean_inc(x_17); -x_18 = lean_ctor_get(x_16, 1); +lean_dec(x_15); +x_18 = lean_ctor_get(x_16, 0); lean_inc(x_18); -lean_dec(x_16); -x_19 = lean_ctor_get(x_17, 0); +x_19 = lean_ctor_get(x_16, 1); lean_inc(x_19); -x_20 = lean_ctor_get(x_17, 1); -lean_inc(x_20); -lean_dec(x_17); -x_21 = lean_array_push(x_9, x_20); -x_22 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__1___boxed), 15, 9); -lean_closure_set(x_22, 0, x_8); -lean_closure_set(x_22, 1, x_3); -lean_closure_set(x_22, 2, x_2); -lean_closure_set(x_22, 3, x_4); -lean_closure_set(x_22, 4, x_5); -lean_closure_set(x_22, 5, x_6); -lean_closure_set(x_22, 6, x_7); -lean_closure_set(x_22, 7, x_15); -lean_closure_set(x_22, 8, x_21); -x_23 = l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs_loop___rarg___closed__2; +lean_dec(x_16); +x_20 = lean_array_push(x_9, x_19); +x_21 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__1___boxed), 15, 9); +lean_closure_set(x_21, 0, x_8); +lean_closure_set(x_21, 1, x_3); +lean_closure_set(x_21, 2, x_2); +lean_closure_set(x_21, 3, x_4); +lean_closure_set(x_21, 4, x_5); +lean_closure_set(x_21, 5, x_6); +lean_closure_set(x_21, 6, x_7); +lean_closure_set(x_21, 7, x_1); +lean_closure_set(x_21, 8, x_20); +x_22 = l_Lean_Meta_withNewEqs_loop___rarg___closed__2; +x_23 = 0; x_24 = 0; -x_25 = 0; -x_26 = l_Lean_Meta_withLocalDecl___at_Lean_Meta_addPPExplicitToExposeDiff_visit___spec__4___rarg(x_23, x_24, x_19, x_22, x_25, x_10, x_11, x_12, x_13, x_18); -return x_26; +x_25 = l_Lean_Meta_withLocalDecl___at_Lean_Meta_addPPExplicitToExposeDiff_visit___spec__4___rarg(x_22, x_23, x_18, x_21, x_24, x_10, x_11, x_12, x_13, x_17); +return x_25; } else { -uint8_t x_27; -lean_dec(x_15); +uint8_t x_26; lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); @@ -2319,34 +2322,35 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_27 = !lean_is_exclusive(x_16); -if (x_27 == 0) +lean_dec(x_1); +x_26 = !lean_is_exclusive(x_15); +if (x_26 == 0) { -return x_16; +return x_15; } 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_inc(x_29); +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_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; +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; } } } } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { lean_object* x_13; lean_object* x_14; lean_inc(x_6); lean_inc(x_3); -x_13 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__2___boxed), 14, 7); +x_13 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__2), 14, 7); lean_closure_set(x_13, 0, x_1); lean_closure_set(x_13, 1, x_7); lean_closure_set(x_13, 2, x_2); @@ -2354,103 +2358,150 @@ lean_closure_set(x_13, 3, x_3); lean_closure_set(x_13, 4, x_4); lean_closure_set(x_13, 5, x_5); lean_closure_set(x_13, 6, x_6); -x_14 = l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs___rarg(x_6, x_3, x_13, x_8, x_9, x_10, x_11, x_12); +x_14 = l_Lean_Meta_withNewEqs___rarg(x_6, x_3, x_13, x_8, x_9, x_10, x_11, x_12); return x_14; } } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_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_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { -lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; uint8_t x_18; lean_object* x_19; -x_14 = l_Lean_mkAppN(x_1, x_7); -x_15 = l_Lean_LocalDecl_userName(x_2); -x_16 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__3), 12, 6); -lean_closure_set(x_16, 0, x_2); -lean_closure_set(x_16, 1, x_3); -lean_closure_set(x_16, 2, x_7); -lean_closure_set(x_16, 3, x_4); -lean_closure_set(x_16, 4, x_5); -lean_closure_set(x_16, 5, x_6); -x_17 = 0; -x_18 = 0; -x_19 = l_Lean_Meta_withLocalDecl___at_Lean_Meta_addPPExplicitToExposeDiff_visit___spec__4___rarg(x_15, x_17, x_14, x_16, x_18, x_9, x_10, x_11, x_12, x_13); -return x_19; +lean_object* x_14; uint8_t x_15; uint8_t x_16; lean_object* x_17; +x_14 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__3), 12, 6); +lean_closure_set(x_14, 0, x_1); +lean_closure_set(x_14, 1, x_2); +lean_closure_set(x_14, 2, x_3); +lean_closure_set(x_14, 3, x_4); +lean_closure_set(x_14, 4, x_5); +lean_closure_set(x_14, 5, x_6); +x_15 = 0; +x_16 = 0; +x_17 = l_Lean_Meta_withLocalDecl___at_Lean_Meta_addPPExplicitToExposeDiff_visit___spec__4___rarg(x_8, x_15, x_7, x_14, x_16, x_9, x_10, x_11, x_12, x_13); +return x_17; } } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, 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_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__5___closed__1() { _start: { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_14 = lean_array_get_size(x_1); -x_15 = lean_ctor_get(x_2, 2); -x_16 = lean_nat_sub(x_14, x_15); -x_17 = l_Array_extract___rarg(x_1, x_16, x_14); -lean_dec(x_14); -x_18 = lean_ctor_get(x_2, 1); -x_19 = lean_unsigned_to_nat(0u); -x_20 = l_Array_extract___rarg(x_1, x_19, x_18); -x_21 = l_Lean_mkAppN(x_3, x_20); -lean_dec(x_20); +lean_object* x_1; +x_1 = lean_mk_string_unchecked("x", 1, 1); +return x_1; +} +} +static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__5___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__5___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, 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_mkAppN(x_1, x_8); +if (lean_obj_tag(x_7) == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_16 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__5___closed__2; +x_17 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_16, x_12, x_13, x_14); +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_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__4(x_2, x_3, x_8, x_4, x_5, x_6, x_15, x_18, x_10, x_11, x_12, x_13, x_19); +return x_20; +} +else +{ +lean_object* x_21; lean_object* x_22; +x_21 = lean_ctor_get(x_7, 0); +lean_inc(x_21); +lean_dec(x_7); +x_22 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__4(x_2, x_3, x_8, x_4, x_5, x_6, x_15, x_21, x_10, x_11, x_12, x_13, x_14); +return x_22; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_15 = lean_array_get_size(x_1); +x_16 = lean_ctor_get(x_2, 2); +x_17 = lean_nat_sub(x_15, x_16); +x_18 = l_Array_extract___rarg(x_1, x_17, x_15); +lean_dec(x_15); +x_19 = lean_ctor_get(x_2, 1); +x_20 = lean_unsigned_to_nat(0u); +x_21 = l_Array_extract___rarg(x_1, x_20, x_19); +x_22 = l_Lean_mkAppN(x_3, 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); -lean_inc(x_21); -x_22 = lean_infer_type(x_21, x_9, x_10, x_11, x_12, x_13); -if (lean_obj_tag(x_22) == 0) +lean_inc(x_22); +x_23 = lean_infer_type(x_22, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_23) == 0) { -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, 0); -lean_inc(x_23); -x_24 = lean_ctor_get(x_22, 1); +lean_object* x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; lean_object* x_28; +x_24 = lean_ctor_get(x_23, 0); lean_inc(x_24); -lean_dec(x_22); -x_25 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__4___boxed), 13, 6); -lean_closure_set(x_25, 0, x_21); -lean_closure_set(x_25, 1, x_4); -lean_closure_set(x_25, 2, x_5); -lean_closure_set(x_25, 3, x_6); -lean_closure_set(x_25, 4, x_7); -lean_closure_set(x_25, 5, x_17); -x_26 = 0; -x_27 = l_Lean_Meta_forallTelescopeReducing___at_Lean_Meta_getParamNames___spec__2___rarg(x_23, x_25, x_26, x_9, x_10, x_11, x_12, x_24); -return x_27; +x_25 = lean_ctor_get(x_23, 1); +lean_inc(x_25); +lean_dec(x_23); +x_26 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__5___boxed), 14, 7); +lean_closure_set(x_26, 0, x_22); +lean_closure_set(x_26, 1, x_4); +lean_closure_set(x_26, 2, x_5); +lean_closure_set(x_26, 3, x_6); +lean_closure_set(x_26, 4, x_7); +lean_closure_set(x_26, 5, x_18); +lean_closure_set(x_26, 6, x_8); +x_27 = 0; +x_28 = l_Lean_Meta_forallTelescopeReducing___at_Lean_Meta_getParamNames___spec__2___rarg(x_24, x_26, x_27, x_10, x_11, x_12, x_13, x_25); +return x_28; } else { -uint8_t x_28; -lean_dec(x_21); -lean_dec(x_17); +uint8_t x_29; +lean_dec(x_22); +lean_dec(x_18); +lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); -lean_dec(x_9); +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_28 = !lean_is_exclusive(x_22); -if (x_28 == 0) +x_29 = !lean_is_exclusive(x_23); +if (x_29 == 0) { -return x_22; +return x_23; } else { -lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_29 = lean_ctor_get(x_22, 0); -x_30 = lean_ctor_get(x_22, 1); +lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_30 = lean_ctor_get(x_23, 0); +x_31 = lean_ctor_get(x_23, 1); +lean_inc(x_31); lean_inc(x_30); -lean_inc(x_29); -lean_dec(x_22); -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_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; } } } } -static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__6___closed__1() { +static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__7___closed__1() { _start: { lean_object* x_1; @@ -2458,89 +2509,90 @@ x_1 = lean_mk_string_unchecked("ill-formed inductive datatype", 29, 29); return x_1; } } -static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__6___closed__2() { +static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__7___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__6___closed__1; +x_1 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__7___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_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__6___closed__3() { +static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__7___closed__3() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__6___closed__2; +x_1 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__7___closed__2; x_2 = l_Lean_MessageData_ofFormat(x_1); return x_2; } } -static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__6___closed__4() { +static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__7___closed__4() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__6___closed__3; +x_1 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__7___closed__3; 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_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___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_object* x_15) { _start: { -lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; -x_15 = lean_array_get_size(x_1); -x_16 = lean_ctor_get(x_2, 2); -x_17 = lean_ctor_get(x_2, 1); -x_18 = lean_nat_add(x_16, x_17); -x_19 = lean_nat_dec_eq(x_15, x_18); -lean_dec(x_18); -lean_dec(x_15); -if (x_19 == 0) +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_16 = lean_array_get_size(x_1); +x_17 = lean_ctor_get(x_2, 2); +x_18 = lean_ctor_get(x_2, 1); +x_19 = lean_nat_add(x_17, x_18); +x_20 = lean_nat_dec_eq(x_16, x_19); +lean_dec(x_19); +lean_dec(x_16); +if (x_20 == 0) { -lean_object* x_20; lean_object* x_21; uint8_t x_22; +lean_object* x_21; lean_object* x_22; uint8_t x_23; +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); -x_20 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__6___closed__4; -x_21 = l_Lean_Meta_throwTacticEx___rarg(x_8, x_5, x_20, x_10, x_11, x_12, x_13, x_14); +x_21 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__7___closed__4; +x_22 = l_Lean_Meta_throwTacticEx___rarg(x_9, x_5, x_21, x_11, x_12, x_13, x_14, x_15); +lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); -lean_dec(x_10); -x_22 = !lean_is_exclusive(x_21); -if (x_22 == 0) +x_23 = !lean_is_exclusive(x_22); +if (x_23 == 0) { -return x_21; +return x_22; } 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_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_22, 0); +x_25 = lean_ctor_get(x_22, 1); +lean_inc(x_25); lean_inc(x_24); -lean_inc(x_23); -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; +lean_dec(x_22); +x_26 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_26, 0, x_24); +lean_ctor_set(x_26, 1, x_25); +return x_26; } } else { -lean_object* x_26; lean_object* x_27; -lean_dec(x_8); -x_26 = lean_box(0); -x_27 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__5(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_26, x_10, x_11, x_12, x_13, x_14); -return x_27; +lean_object* x_27; lean_object* x_28; +lean_dec(x_9); +x_27 = lean_box(0); +x_28 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__6(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_27, x_11, x_12, x_13, x_14, x_15); +return x_28; } } } -static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__1() { +static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__1() { _start: { lean_object* x_1; @@ -2548,36 +2600,36 @@ x_1 = lean_mk_string_unchecked("inductive type expected", 23, 23); return x_1; } } -static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__2() { +static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__1; +x_1 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__1; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__3() { +static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__3() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__2; +x_1 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__2; x_2 = l_Lean_MessageData_ofFormat(x_1); return x_2; } } -static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__4() { +static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__4() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__3; +x_1 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__3; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__5() { +static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__5() { _start: { lean_object* x_1; @@ -2585,189 +2637,193 @@ x_1 = lean_mk_string_unchecked("indexed inductive type expected", 31, 31); return x_1; } } -static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__6() { +static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__6() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__5; +x_1 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__5; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__7() { +static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__7() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__6; +x_1 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__6; x_2 = l_Lean_MessageData_ofFormat(x_1); return x_2; } } -static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__8() { +static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__8() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__7; +x_1 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__7; 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_withAppAux___at_Lean_Meta_generalizeIndices___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { _start: { -switch (lean_obj_tag(x_6)) { +switch (lean_obj_tag(x_7)) { case 4: { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; -lean_dec(x_8); -x_14 = lean_ctor_get(x_6, 0); -lean_inc(x_14); -x_15 = lean_st_ref_get(x_12, x_13); -x_16 = lean_ctor_get(x_15, 0); -lean_inc(x_16); -x_17 = lean_ctor_get(x_15, 1); +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +lean_dec(x_9); +x_15 = lean_ctor_get(x_7, 0); +lean_inc(x_15); +x_16 = lean_st_ref_get(x_13, x_14); +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); lean_dec(x_16); -x_19 = lean_environment_find(x_18, x_14); -if (lean_obj_tag(x_19) == 0) +x_19 = lean_ctor_get(x_17, 0); +lean_inc(x_19); +lean_dec(x_17); +x_20 = lean_environment_find(x_19, x_15); +if (lean_obj_tag(x_20) == 0) { -lean_object* x_20; lean_object* x_21; +lean_object* x_21; lean_object* x_22; +lean_dec(x_8); lean_dec(x_7); -lean_dec(x_6); lean_dec(x_5); +lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_20 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__4; -x_21 = l_Lean_Meta_throwTacticEx___rarg(x_4, x_1, x_20, x_9, x_10, x_11, x_12, x_17); +x_21 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__4; +x_22 = l_Lean_Meta_throwTacticEx___rarg(x_6, x_1, x_21, x_10, x_11, x_12, x_13, x_18); +lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); -lean_dec(x_9); -return x_21; +return x_22; } else { -lean_object* x_22; -x_22 = lean_ctor_get(x_19, 0); -lean_inc(x_22); -lean_dec(x_19); -if (lean_obj_tag(x_22) == 5) -{ -lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; -x_23 = lean_ctor_get(x_22, 0); +lean_object* x_23; +x_23 = lean_ctor_get(x_20, 0); lean_inc(x_23); -lean_dec(x_22); -x_24 = lean_ctor_get(x_23, 2); -lean_inc(x_24); -x_25 = lean_unsigned_to_nat(0u); -x_26 = lean_nat_dec_lt(x_25, x_24); -lean_dec(x_24); -if (x_26 == 0) +lean_dec(x_20); +if (lean_obj_tag(x_23) == 5) { -lean_object* x_27; lean_object* x_28; uint8_t x_29; +lean_object* x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); lean_dec(x_23); +x_25 = lean_ctor_get(x_24, 2); +lean_inc(x_25); +x_26 = lean_unsigned_to_nat(0u); +x_27 = lean_nat_dec_lt(x_26, x_25); +lean_dec(x_25); +if (x_27 == 0) +{ +lean_object* x_28; lean_object* x_29; uint8_t x_30; +lean_dec(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); lean_dec(x_2); -x_27 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__8; -x_28 = l_Lean_Meta_throwTacticEx___rarg(x_4, x_1, x_27, x_9, x_10, x_11, x_12, x_17); +x_28 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__8; +x_29 = l_Lean_Meta_throwTacticEx___rarg(x_6, x_1, x_28, x_10, x_11, x_12, x_13, x_18); +lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); -lean_dec(x_9); -x_29 = !lean_is_exclusive(x_28); -if (x_29 == 0) +x_30 = !lean_is_exclusive(x_29); +if (x_30 == 0) { -return x_28; +return x_29; } else { -lean_object* x_30; lean_object* x_31; lean_object* x_32; -x_30 = lean_ctor_get(x_28, 0); -x_31 = lean_ctor_get(x_28, 1); +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_29, 0); +x_32 = lean_ctor_get(x_29, 1); +lean_inc(x_32); lean_inc(x_31); -lean_inc(x_30); -lean_dec(x_28); -x_32 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_32, 0, x_30); -lean_ctor_set(x_32, 1, x_31); -return x_32; +lean_dec(x_29); +x_33 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_33, 0, x_31); +lean_ctor_set(x_33, 1, x_32); +return x_33; } } else { -lean_object* x_33; lean_object* x_34; -x_33 = lean_box(0); -x_34 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__6(x_7, x_23, x_6, x_5, x_1, x_2, x_3, x_4, x_33, x_9, x_10, x_11, x_12, x_17); -lean_dec(x_23); -lean_dec(x_7); -return x_34; +lean_object* x_34; lean_object* x_35; +x_34 = lean_box(0); +x_35 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__7(x_8, x_24, x_7, x_2, x_1, x_4, x_5, x_3, x_6, x_34, x_10, x_11, x_12, x_13, x_18); +lean_dec(x_24); +lean_dec(x_8); +return x_35; } } else { -lean_object* x_35; lean_object* x_36; -lean_dec(x_22); +lean_object* x_36; lean_object* x_37; +lean_dec(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); lean_dec(x_2); -x_35 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__4; -x_36 = l_Lean_Meta_throwTacticEx___rarg(x_4, x_1, x_35, x_9, x_10, x_11, x_12, x_17); +x_36 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__4; +x_37 = l_Lean_Meta_throwTacticEx___rarg(x_6, x_1, x_36, x_10, x_11, x_12, x_13, x_18); +lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); -lean_dec(x_9); -return x_36; +return x_37; } } } case 5: { -lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; -x_37 = lean_ctor_get(x_6, 0); -lean_inc(x_37); -x_38 = lean_ctor_get(x_6, 1); +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_38 = lean_ctor_get(x_7, 0); lean_inc(x_38); -lean_dec(x_6); -x_39 = lean_array_set(x_7, x_8, x_38); -x_40 = lean_unsigned_to_nat(1u); -x_41 = lean_nat_sub(x_8, x_40); -lean_dec(x_8); -x_6 = x_37; -x_7 = x_39; -x_8 = x_41; +x_39 = lean_ctor_get(x_7, 1); +lean_inc(x_39); +lean_dec(x_7); +x_40 = lean_array_set(x_8, x_9, x_39); +x_41 = lean_unsigned_to_nat(1u); +x_42 = lean_nat_sub(x_9, x_41); +lean_dec(x_9); +x_7 = x_38; +x_8 = x_40; +x_9 = x_42; goto _start; } default: { -lean_object* x_43; lean_object* x_44; +lean_object* x_44; lean_object* x_45; +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_43 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__4; -x_44 = l_Lean_Meta_throwTacticEx___rarg(x_4, x_1, x_43, x_9, x_10, x_11, x_12, x_13); +x_44 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__4; +x_45 = l_Lean_Meta_throwTacticEx___rarg(x_6, x_1, x_44, x_10, x_11, x_12, x_13, x_14); +lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); -lean_dec(x_9); -return x_44; +return x_45; } } } } -static lean_object* _init_l_Lean_Meta_generalizeIndices___lambda__1___closed__1() { +static lean_object* _init_l_Lean_Meta_generalizeIndices_x27___lambda__1___closed__1() { _start: { lean_object* x_1; @@ -2775,53 +2831,56 @@ x_1 = lean_mk_string_unchecked("generalizeIndices", 17, 17); return x_1; } } -static lean_object* _init_l_Lean_Meta_generalizeIndices___lambda__1___closed__2() { +static lean_object* _init_l_Lean_Meta_generalizeIndices_x27___lambda__1___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Meta_generalizeIndices___lambda__1___closed__1; +x_2 = l_Lean_Meta_generalizeIndices_x27___lambda__1___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Meta_generalizeIndices___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Lean_Meta_generalizeIndices_x27___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { -lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_8 = lean_ctor_get(x_3, 2); -lean_inc(x_8); -x_9 = l_Lean_Meta_getLocalInstances(x_3, x_4, x_5, x_6, x_7); -x_10 = lean_ctor_get(x_9, 0); -lean_inc(x_10); -x_11 = lean_ctor_get(x_9, 1); +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_9 = lean_ctor_get(x_4, 2); +lean_inc(x_9); +x_10 = l_Lean_Meta_getLocalInstances(x_4, x_5, x_6, x_7, x_8); +x_11 = lean_ctor_get(x_10, 0); lean_inc(x_11); -lean_dec(x_9); -x_12 = l_Lean_Meta_generalizeIndices___lambda__1___closed__2; +x_12 = lean_ctor_get(x_10, 1); +lean_inc(x_12); +lean_dec(x_10); +x_13 = l_Lean_Meta_generalizeIndices_x27___lambda__1___closed__2; lean_inc(x_1); -x_13 = l_Lean_MVarId_checkNotAssigned(x_1, x_12, x_3, x_4, x_5, x_6, x_11); -if (lean_obj_tag(x_13) == 0) +x_14 = l_Lean_MVarId_checkNotAssigned(x_1, x_13, x_4, x_5, x_6, x_7, x_12); +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_inc(x_3); -x_15 = l_Lean_FVarId_getDecl(x_2, x_3, x_4, x_5, x_6, x_14); -if (lean_obj_tag(x_15) == 0) +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_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_2); +x_16 = lean_infer_type(x_2, x_4, x_5, x_6, x_7, 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; -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; +x_17 = lean_ctor_get(x_16, 0); lean_inc(x_17); -lean_dec(x_15); -x_18 = l_Lean_LocalDecl_type(x_16); +x_18 = lean_ctor_get(x_16, 1); +lean_inc(x_18); +lean_dec(x_16); +lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -lean_inc(x_3); -x_19 = lean_whnf(x_18, x_3, x_4, x_5, x_6, x_17); +x_19 = l_Lean_Meta_whnfD(x_17, x_4, x_5, x_6, x_7, x_18); 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; @@ -2838,19 +2897,20 @@ x_25 = lean_mk_array(x_23, x_24); x_26 = lean_unsigned_to_nat(1u); x_27 = lean_nat_sub(x_23, x_26); lean_dec(x_23); -x_28 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1(x_1, x_8, x_10, x_12, x_16, x_20, x_25, x_27, x_3, x_4, x_5, x_6, x_21); +x_28 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1(x_1, x_2, x_3, x_9, x_11, x_13, x_20, x_25, x_27, x_4, x_5, x_6, x_7, x_21); return x_28; } else { uint8_t x_29; -lean_dec(x_16); -lean_dec(x_10); -lean_dec(x_8); +lean_dec(x_11); +lean_dec(x_9); +lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); +lean_dec(x_2); lean_dec(x_1); x_29 = !lean_is_exclusive(x_19); if (x_29 == 0) @@ -2875,26 +2935,28 @@ return x_32; else { uint8_t x_33; -lean_dec(x_10); -lean_dec(x_8); +lean_dec(x_11); +lean_dec(x_9); +lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); +lean_dec(x_2); lean_dec(x_1); -x_33 = !lean_is_exclusive(x_15); +x_33 = !lean_is_exclusive(x_16); if (x_33 == 0) { -return x_15; +return x_16; } else { lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_34 = lean_ctor_get(x_15, 0); -x_35 = lean_ctor_get(x_15, 1); +x_34 = lean_ctor_get(x_16, 0); +x_35 = lean_ctor_get(x_16, 1); lean_inc(x_35); lean_inc(x_34); -lean_dec(x_15); +lean_dec(x_16); x_36 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_36, 0, x_34); lean_ctor_set(x_36, 1, x_35); @@ -2905,27 +2967,28 @@ return x_36; else { uint8_t x_37; -lean_dec(x_10); -lean_dec(x_8); +lean_dec(x_11); +lean_dec(x_9); +lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_37 = !lean_is_exclusive(x_13); +x_37 = !lean_is_exclusive(x_14); if (x_37 == 0) { -return x_13; +return x_14; } else { lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_38 = lean_ctor_get(x_13, 0); -x_39 = lean_ctor_get(x_13, 1); +x_38 = lean_ctor_get(x_14, 0); +x_39 = lean_ctor_get(x_14, 1); lean_inc(x_39); lean_inc(x_38); -lean_dec(x_13); +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); @@ -2934,67 +2997,122 @@ return x_40; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_generalizeIndices(lean_object* x_1, lean_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_generalizeIndices_x27(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { -lean_object* x_8; lean_object* x_9; +lean_object* x_9; lean_object* x_10; lean_inc(x_1); -x_8 = lean_alloc_closure((void*)(l_Lean_Meta_generalizeIndices___lambda__1), 7, 2); -lean_closure_set(x_8, 0, x_1); -lean_closure_set(x_8, 1, x_2); -x_9 = l_Lean_MVarId_withContext___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___spec__2___rarg(x_1, x_8, x_3, x_4, x_5, x_6, x_7); -return x_9; +x_9 = lean_alloc_closure((void*)(l_Lean_Meta_generalizeIndices_x27___lambda__1), 8, 3); +lean_closure_set(x_9, 0, x_1); +lean_closure_set(x_9, 1, x_2); +lean_closure_set(x_9, 2, x_3); +x_10 = l_Lean_MVarId_withContext___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___spec__2___rarg(x_1, x_9, x_4, x_5, x_6, x_7, x_8); +return x_10; } } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { _start: { lean_object* x_16; -x_16 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +x_16 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); lean_dec(x_9); lean_dec(x_7); lean_dec(x_4); return x_16; } } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___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: { lean_object* x_15; -x_15 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___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, x_13, x_14); -lean_dec(x_1); +x_15 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__5(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +lean_dec(x_9); return x_15; } } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___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_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { _start: { -lean_object* x_14; -x_14 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___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); -lean_dec(x_8); -return x_14; +lean_object* x_15; +x_15 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__6(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +lean_dec(x_9); +lean_dec(x_2); +lean_dec(x_1); +return x_15; } } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___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_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { _start: { -lean_object* x_14; -x_14 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___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_8); +lean_object* x_16; +x_16 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___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); +lean_dec(x_10); lean_dec(x_2); lean_dec(x_1); -return x_14; +return x_16; } } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +LEAN_EXPORT lean_object* l_Lean_Meta_generalizeIndices___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { -lean_object* x_15; -x_15 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__6(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +lean_object* x_8; +lean_inc(x_3); +x_8 = l_Lean_FVarId_getDecl(x_1, x_3, x_4, x_5, x_6, x_7); +if (lean_obj_tag(x_8) == 0) +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_9 = lean_ctor_get(x_8, 0); +lean_inc(x_9); +x_10 = lean_ctor_get(x_8, 1); +lean_inc(x_10); +lean_dec(x_8); +x_11 = l_Lean_LocalDecl_toExpr(x_9); +x_12 = l_Lean_LocalDecl_userName(x_9); lean_dec(x_9); +x_13 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_13, 0, x_12); +x_14 = l_Lean_Meta_generalizeIndices_x27(x_2, x_11, x_13, x_3, x_4, x_5, x_6, x_10); +return x_14; +} +else +{ +uint8_t x_15; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -return x_15; +x_15 = !lean_is_exclusive(x_8); +if (x_15 == 0) +{ +return x_8; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_8, 0); +x_17 = lean_ctor_get(x_8, 1); +lean_inc(x_17); +lean_inc(x_16); +lean_dec(x_8); +x_18 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_18, 0, x_16); +lean_ctor_set(x_18, 1, x_17); +return x_18; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_generalizeIndices(lean_object* x_1, 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_inc(x_1); +x_8 = lean_alloc_closure((void*)(l_Lean_Meta_generalizeIndices___lambda__1), 7, 2); +lean_closure_set(x_8, 0, x_2); +lean_closure_set(x_8, 1, x_1); +x_9 = l_Lean_MVarId_withContext___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___spec__2___rarg(x_1, x_8, x_3, x_4, x_5, x_6, x_7); +return x_9; } } static lean_object* _init_l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_mkCasesContext_x3f___spec__1___closed__1() { @@ -13520,7 +13638,7 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_11 = l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs___rarg___closed__1; +x_11 = l_Lean_Meta_withNewEqs___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_7); @@ -13539,7 +13657,7 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_14 = l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs___rarg___closed__1; +x_14 = l_Lean_Meta_withNewEqs___rarg___closed__1; x_15 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_15, 0, x_14); lean_ctor_set(x_15, 1, x_7); @@ -13551,7 +13669,7 @@ size_t x_16; size_t x_17; lean_object* x_18; lean_object* x_19; x_16 = 0; x_17 = lean_usize_of_nat(x_8); lean_dec(x_8); -x_18 = l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs___rarg___closed__1; +x_18 = l_Lean_Meta_withNewEqs___rarg___closed__1; x_19 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_unifyCasesEqs___spec__2(x_1, x_2, x_16, x_17, x_18, x_3, x_4, x_5, x_6, x_7); return x_19; } @@ -15313,7 +15431,7 @@ lean_dec(x_35); x_41 = l_Lean_LocalDecl_fvarId(x_34); lean_dec(x_34); x_42 = lean_box(0); -x_43 = l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs___rarg___closed__1; +x_43 = l_Lean_Meta_withNewEqs___rarg___closed__1; lean_inc(x_2); x_44 = lean_alloc_closure((void*)(l_Array_forIn_x27Unsafe_loop___at_Lean_MVarId_casesRec___spec__6___lambda__1), 9, 4); lean_closure_set(x_44, 0, x_2); @@ -15497,7 +15615,7 @@ lean_dec(x_66); x_72 = l_Lean_LocalDecl_fvarId(x_65); lean_dec(x_65); x_73 = lean_box(0); -x_74 = l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs___rarg___closed__1; +x_74 = l_Lean_Meta_withNewEqs___rarg___closed__1; lean_inc(x_2); x_75 = lean_alloc_closure((void*)(l_Array_forIn_x27Unsafe_loop___at_Lean_MVarId_casesRec___spec__6___lambda__1), 9, 4); lean_closure_set(x_75, 0, x_2); @@ -15999,7 +16117,7 @@ lean_dec(x_39); x_45 = l_Lean_LocalDecl_fvarId(x_38); lean_dec(x_38); x_46 = lean_box(0); -x_47 = l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs___rarg___closed__1; +x_47 = l_Lean_Meta_withNewEqs___rarg___closed__1; lean_inc(x_2); x_48 = lean_alloc_closure((void*)(l_Array_forIn_x27Unsafe_loop___at_Lean_MVarId_casesRec___spec__6___lambda__1), 9, 4); lean_closure_set(x_48, 0, x_2); @@ -16183,7 +16301,7 @@ lean_dec(x_70); x_76 = l_Lean_LocalDecl_fvarId(x_69); lean_dec(x_69); x_77 = lean_box(0); -x_78 = l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs___rarg___closed__1; +x_78 = l_Lean_Meta_withNewEqs___rarg___closed__1; lean_inc(x_2); x_79 = lean_alloc_closure((void*)(l_Array_forIn_x27Unsafe_loop___at_Lean_MVarId_casesRec___spec__6___lambda__1), 9, 4); lean_closure_set(x_79, 0, x_2); @@ -18325,7 +18443,7 @@ return x_109; } } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__1() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__1() { _start: { lean_object* x_1; @@ -18333,27 +18451,27 @@ x_1 = lean_mk_string_unchecked("Lean", 4, 4); return x_1; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__2() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____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_Cases___hyg_3794____closed__1; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____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_Cases___hyg_3794____closed__3() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__2; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__2; x_2 = l_Lean_Meta_Cases_cases___lambda__2___closed__5; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__4() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__4() { _start: { lean_object* x_1; @@ -18361,17 +18479,17 @@ x_1 = lean_mk_string_unchecked("initFn", 6, 6); return x_1; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__5() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__3; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__4; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__3; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__4; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__6() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__6() { _start: { lean_object* x_1; @@ -18379,47 +18497,47 @@ x_1 = lean_mk_string_unchecked("_@", 2, 2); return x_1; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__7() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__5; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__6; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__5; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__6; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__8() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__7; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__1; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__7; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____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_Cases___hyg_3794____closed__9() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__8; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__8; x_2 = l_Lean_Meta_Cases_cases___lambda__2___closed__5; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__10() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__9; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__9; x_2 = l_Lean_Meta_Cases_cases___lambda__2___closed__6; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__11() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__11() { _start: { lean_object* x_1; @@ -18427,17 +18545,17 @@ x_1 = lean_mk_string_unchecked("Cases", 5, 5); return x_1; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__12() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__10; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__11; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__10; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__11; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__13() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__13() { _start: { lean_object* x_1; @@ -18445,33 +18563,33 @@ x_1 = lean_mk_string_unchecked("_hyg", 4, 4); return x_1; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__14() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__12; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__13; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__12; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__13; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__15() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__14; -x_2 = lean_unsigned_to_nat(3794u); +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__14; +x_2 = lean_unsigned_to_nat(3900u); 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_Tactic_Cases___hyg_3794_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900_(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_Cases_cases___lambda__2___closed__8; x_3 = 0; -x_4 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__15; +x_4 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__15; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); return x_5; } @@ -18533,12 +18651,12 @@ l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_mkEqAndProof___closed__6 = _init lean_mark_persistent(l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_mkEqAndProof___closed__6); l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_mkEqAndProof___closed__7 = _init_l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_mkEqAndProof___closed__7(); lean_mark_persistent(l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_mkEqAndProof___closed__7); -l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs_loop___rarg___closed__1 = _init_l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs_loop___rarg___closed__1(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs_loop___rarg___closed__1); -l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs_loop___rarg___closed__2 = _init_l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs_loop___rarg___closed__2(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs_loop___rarg___closed__2); -l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs___rarg___closed__1 = _init_l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs___rarg___closed__1(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs___rarg___closed__1); +l_Lean_Meta_withNewEqs_loop___rarg___closed__1 = _init_l_Lean_Meta_withNewEqs_loop___rarg___closed__1(); +lean_mark_persistent(l_Lean_Meta_withNewEqs_loop___rarg___closed__1); +l_Lean_Meta_withNewEqs_loop___rarg___closed__2 = _init_l_Lean_Meta_withNewEqs_loop___rarg___closed__2(); +lean_mark_persistent(l_Lean_Meta_withNewEqs_loop___rarg___closed__2); +l_Lean_Meta_withNewEqs___rarg___closed__1 = _init_l_Lean_Meta_withNewEqs___rarg___closed__1(); +lean_mark_persistent(l_Lean_Meta_withNewEqs___rarg___closed__1); l_Lean_Meta_generalizeTargetsEq___lambda__3___closed__1 = _init_l_Lean_Meta_generalizeTargetsEq___lambda__3___closed__1(); lean_mark_persistent(l_Lean_Meta_generalizeTargetsEq___lambda__3___closed__1); l_Lean_Meta_generalizeTargetsEq___lambda__3___closed__2 = _init_l_Lean_Meta_generalizeTargetsEq___lambda__3___closed__2(); @@ -18551,34 +18669,38 @@ l_Lean_Meta_generalizeTargetsEq___closed__1 = _init_l_Lean_Meta_generalizeTarget lean_mark_persistent(l_Lean_Meta_generalizeTargetsEq___closed__1); l_Lean_Meta_generalizeTargetsEq___closed__2 = _init_l_Lean_Meta_generalizeTargetsEq___closed__2(); lean_mark_persistent(l_Lean_Meta_generalizeTargetsEq___closed__2); -l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__6___closed__1 = _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__6___closed__1(); -lean_mark_persistent(l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__6___closed__1); -l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__6___closed__2 = _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__6___closed__2(); -lean_mark_persistent(l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__6___closed__2); -l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__6___closed__3 = _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__6___closed__3(); -lean_mark_persistent(l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__6___closed__3); -l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__6___closed__4 = _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__6___closed__4(); -lean_mark_persistent(l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__6___closed__4); -l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__1 = _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__1(); -lean_mark_persistent(l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__1); -l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__2 = _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__2(); -lean_mark_persistent(l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__2); -l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__3 = _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__3(); -lean_mark_persistent(l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__3); -l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__4 = _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__4(); -lean_mark_persistent(l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__4); -l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__5 = _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__5(); -lean_mark_persistent(l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__5); -l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__6 = _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__6(); -lean_mark_persistent(l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__6); -l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__7 = _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__7(); -lean_mark_persistent(l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__7); -l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__8 = _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__8(); -lean_mark_persistent(l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__8); -l_Lean_Meta_generalizeIndices___lambda__1___closed__1 = _init_l_Lean_Meta_generalizeIndices___lambda__1___closed__1(); -lean_mark_persistent(l_Lean_Meta_generalizeIndices___lambda__1___closed__1); -l_Lean_Meta_generalizeIndices___lambda__1___closed__2 = _init_l_Lean_Meta_generalizeIndices___lambda__1___closed__2(); -lean_mark_persistent(l_Lean_Meta_generalizeIndices___lambda__1___closed__2); +l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__5___closed__1 = _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__5___closed__1(); +lean_mark_persistent(l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__5___closed__1); +l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__5___closed__2 = _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__5___closed__2(); +lean_mark_persistent(l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__5___closed__2); +l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__7___closed__1 = _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__7___closed__1(); +lean_mark_persistent(l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__7___closed__1); +l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__7___closed__2 = _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__7___closed__2(); +lean_mark_persistent(l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__7___closed__2); +l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__7___closed__3 = _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__7___closed__3(); +lean_mark_persistent(l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__7___closed__3); +l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__7___closed__4 = _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__7___closed__4(); +lean_mark_persistent(l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__7___closed__4); +l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__1 = _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__1(); +lean_mark_persistent(l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__1); +l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__2 = _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__2(); +lean_mark_persistent(l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__2); +l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__3 = _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__3(); +lean_mark_persistent(l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__3); +l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__4 = _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__4(); +lean_mark_persistent(l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__4); +l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__5 = _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__5(); +lean_mark_persistent(l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__5); +l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__6 = _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__6(); +lean_mark_persistent(l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__6); +l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__7 = _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__7(); +lean_mark_persistent(l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__7); +l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__8 = _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__8(); +lean_mark_persistent(l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__8); +l_Lean_Meta_generalizeIndices_x27___lambda__1___closed__1 = _init_l_Lean_Meta_generalizeIndices_x27___lambda__1___closed__1(); +lean_mark_persistent(l_Lean_Meta_generalizeIndices_x27___lambda__1___closed__1); +l_Lean_Meta_generalizeIndices_x27___lambda__1___closed__2 = _init_l_Lean_Meta_generalizeIndices_x27___lambda__1___closed__2(); +lean_mark_persistent(l_Lean_Meta_generalizeIndices_x27___lambda__1___closed__2); l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_mkCasesContext_x3f___spec__1___closed__1 = _init_l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_mkCasesContext_x3f___spec__1___closed__1(); lean_mark_persistent(l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_mkCasesContext_x3f___spec__1___closed__1); l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__56___closed__1 = _init_l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__56___closed__1(); @@ -18669,37 +18791,37 @@ l_Lean_MVarId_byCasesDec___closed__4 = _init_l_Lean_MVarId_byCasesDec___closed__ lean_mark_persistent(l_Lean_MVarId_byCasesDec___closed__4); l_Lean_MVarId_byCasesDec___closed__5 = _init_l_Lean_MVarId_byCasesDec___closed__5(); lean_mark_persistent(l_Lean_MVarId_byCasesDec___closed__5); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__1 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__1(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__1); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__2 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__2(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__2); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__3 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__3(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__3); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__4 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__4(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__4); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__5 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__5(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__5); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__6 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__6(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__6); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__7 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__7(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__7); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__8 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__8(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__8); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__9 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__9(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__9); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__10 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__10(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__10); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__11 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__11(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__11); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__12 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__12(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__12); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__13 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__13(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__13); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__14 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__14(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__14); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__15 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__15(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__15); -if (builtin) {res = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794_(lean_io_mk_world()); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__1 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__1(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__1); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__2 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__2(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__2); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__3 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__3(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__3); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__4 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__4(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__4); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__5 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__5(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__5); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__6 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__6(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__6); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__7 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__7(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__7); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__8 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__8(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__8); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__9 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__9(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__9); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__10 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__10(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__10); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__11 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__11(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__11); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__12 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__12(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__12); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__13 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__13(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__13); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__14 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__14(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__14); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__15 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__15(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__15); +if (builtin) {res = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900_(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/Grind.c b/stage0/stdlib/Lean/Meta/Tactic/Grind.c index 08ec47f8a228..1d3922f2ccad 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Grind.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Grind.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Lean.Meta.Tactic.Grind -// Imports: Lean.Meta.Tactic.Grind.Attr Lean.Meta.Tactic.Grind.RevertAll Lean.Meta.Tactic.Grind.Types Lean.Meta.Tactic.Grind.Util Lean.Meta.Tactic.Grind.Cases Lean.Meta.Tactic.Grind.Injection Lean.Meta.Tactic.Grind.Core Lean.Meta.Tactic.Grind.Canon Lean.Meta.Tactic.Grind.MarkNestedProofs Lean.Meta.Tactic.Grind.Inv Lean.Meta.Tactic.Grind.Proof Lean.Meta.Tactic.Grind.Propagate Lean.Meta.Tactic.Grind.PP Lean.Meta.Tactic.Grind.Simp Lean.Meta.Tactic.Grind.Ctor Lean.Meta.Tactic.Grind.Parser Lean.Meta.Tactic.Grind.EMatchTheorem Lean.Meta.Tactic.Grind.EMatch Lean.Meta.Tactic.Grind.Main +// Imports: Lean.Meta.Tactic.Grind.Attr Lean.Meta.Tactic.Grind.RevertAll Lean.Meta.Tactic.Grind.Types Lean.Meta.Tactic.Grind.Util Lean.Meta.Tactic.Grind.Cases Lean.Meta.Tactic.Grind.Injection Lean.Meta.Tactic.Grind.Core Lean.Meta.Tactic.Grind.Canon Lean.Meta.Tactic.Grind.MarkNestedProofs Lean.Meta.Tactic.Grind.Inv Lean.Meta.Tactic.Grind.Proof Lean.Meta.Tactic.Grind.Propagate Lean.Meta.Tactic.Grind.PP Lean.Meta.Tactic.Grind.Simp Lean.Meta.Tactic.Grind.Ctor Lean.Meta.Tactic.Grind.Parser Lean.Meta.Tactic.Grind.EMatchTheorem Lean.Meta.Tactic.Grind.EMatch Lean.Meta.Tactic.Grind.Main Lean.Meta.Tactic.Grind.CasesMatch #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -13,86 +13,117 @@ #ifdef __cplusplus extern "C" { #endif +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_745____closed__2; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_78____closed__2; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_634____closed__3; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_856____closed__3; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__8; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__9; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_671____closed__3; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_300____closed__3; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__1; +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_597_(lean_object*); LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_226_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_412_(lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_708____closed__1; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__14; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__10; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_226____closed__3; +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_671_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_819_(lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_448____closed__2; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__2; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_485____closed__1; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_708____closed__2; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_337____closed__3; -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_375_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_745_(lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_448____closed__1; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__13; -static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_375____closed__1; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_411____closed__1; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_523____closed__2; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_152____closed__1; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_189____closed__2; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__11; -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_449_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_448_(lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_819____closed__1; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_411____closed__2; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_745____closed__1; lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_115_(lean_object*); static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_263____closed__3; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_115____closed__3; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_523____closed__1; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__3; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_671____closed__2; +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_634_(lean_object*); static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_41____closed__3; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_337____closed__1; lean_object* l_Lean_registerTraceClass(lean_object*, uint8_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_152_(lean_object*); LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_300_(lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_486____closed__3; -static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_375____closed__3; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_485____closed__3; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_374____closed__2; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_226____closed__1; -static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_375____closed__2; +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_374_(lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_597____closed__3; lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_449____closed__3; -static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_412____closed__3; +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_708_(lean_object*); static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__6; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_300____closed__1; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_634____closed__1; LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_523_(lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_412____closed__2; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_671____closed__1; LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_78_(lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__5; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_189____closed__3; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_597____closed__2; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_263____closed__2; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_523____closed__3; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_41____closed__1; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__12; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_560____closed__2; -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_486_(lean_object*); static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_78____closed__3; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_485____closed__2; LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_189_(lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_486____closed__2; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_189____closed__1; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_226____closed__2; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__16; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_448____closed__3; lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_152____closed__3; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_560____closed__1; -static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_449____closed__2; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_374____closed__1; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_856____closed__1; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_856____closed__2; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_819____closed__2; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_745____closed__3; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_374____closed__3; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_782____closed__1; LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_41_(lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_486____closed__1; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_634____closed__2; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_782____closed__2; +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_782_(lean_object*); static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_337____closed__2; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__15; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_263____closed__1; +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_411_(lean_object*); static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_300____closed__2; LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_560_(lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_449____closed__1; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__17; LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_337_(lean_object*); lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_411____closed__3; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_115____closed__1; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__18; -static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_412____closed__1; +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_856_(lean_object*); LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4_(lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_597____closed__1; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_708____closed__3; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__7; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__4; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_115____closed__2; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_782____closed__3; +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_485_(lean_object*); LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_263_(lean_object*); static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_152____closed__2; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_78____closed__1; @@ -473,19 +504,20 @@ static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_226_ _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("instance", 8, 8); +x_1 = lean_mk_string_unchecked("search", 6, 6); return x_1; } } static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_226____closed__2() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__1; x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_152____closed__1; -x_3 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_226____closed__1; -x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); -return x_4; +x_3 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_189____closed__1; +x_4 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_226____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_initFn____x40_Lean_Meta_Tactic_Grind___hyg_226____closed__3() { @@ -513,20 +545,19 @@ static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_263_ _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("assignment", 10, 10); +x_1 = lean_mk_string_unchecked("instance", 8, 8); return x_1; } } static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_263____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_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__1; x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_152____closed__1; -x_3 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_226____closed__1; -x_4 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_263____closed__1; -x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); -return x_5; +x_3 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_263____closed__1; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; } } static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_263____closed__3() { @@ -554,18 +585,20 @@ static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_300_ _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("issues", 6, 6); +x_1 = lean_mk_string_unchecked("assignment", 10, 10); return x_1; } } static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_300____closed__2() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* 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_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__1; -x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_300____closed__1; -x_3 = l_Lean_Name_mkStr2(x_1, x_2); -return x_3; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_152____closed__1; +x_3 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_263____closed__1; +x_4 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_300____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_initFn____x40_Lean_Meta_Tactic_Grind___hyg_300____closed__3() { @@ -593,7 +626,7 @@ static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_337_ _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("simp", 4, 4); +x_1 = lean_mk_string_unchecked("issues", 6, 6); return x_1; } } @@ -628,161 +661,160 @@ x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); return x_5; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_375____closed__1() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_374____closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("debug", 5, 5); +x_1 = lean_mk_string_unchecked("simp", 4, 4); return x_1; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_375____closed__2() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_374____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__1; -x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_375____closed__1; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_374____closed__1; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_375____closed__3() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_374____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__17; -x_2 = lean_unsigned_to_nat(375u); +x_2 = lean_unsigned_to_nat(374u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_375_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_374_(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_Meta_Tactic_Grind___hyg_375____closed__2; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_374____closed__2; x_3 = 0; -x_4 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_375____closed__3; +x_4 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_374____closed__3; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); return x_5; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_412____closed__1() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_411____closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("proofs", 6, 6); +x_1 = lean_mk_string_unchecked("split", 5, 5); return x_1; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_412____closed__2() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_411____closed__2() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__1; -x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_375____closed__1; -x_3 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_412____closed__1; -x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); -return x_4; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_411____closed__1; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_412____closed__3() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_411____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__17; -x_2 = lean_unsigned_to_nat(412u); +x_2 = lean_unsigned_to_nat(411u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_412_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_411_(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_Meta_Tactic_Grind___hyg_412____closed__2; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_411____closed__2; x_3 = 0; -x_4 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_412____closed__3; +x_4 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_411____closed__3; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); return x_5; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_449____closed__1() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_448____closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("congr", 5, 5); +x_1 = lean_mk_string_unchecked("candidate", 9, 9); return x_1; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_449____closed__2() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_448____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__1; -x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_375____closed__1; -x_3 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_449____closed__1; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_411____closed__1; +x_3 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_448____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_449____closed__3() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_448____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__17; -x_2 = lean_unsigned_to_nat(449u); +x_2 = lean_unsigned_to_nat(448u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_449_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_448_(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_Meta_Tactic_Grind___hyg_449____closed__2; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_448____closed__2; x_3 = 0; -x_4 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_449____closed__3; +x_4 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_448____closed__3; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); return x_5; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_486____closed__1() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_485____closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("proof", 5, 5); +x_1 = lean_mk_string_unchecked("resolved", 8, 8); return x_1; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_486____closed__2() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_485____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__1; -x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_375____closed__1; -x_3 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_486____closed__1; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_411____closed__1; +x_3 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_485____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_486____closed__3() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_485____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__17; -x_2 = lean_unsigned_to_nat(486u); +x_2 = lean_unsigned_to_nat(485u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_486_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_485_(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_Meta_Tactic_Grind___hyg_486____closed__2; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_485____closed__2; x_3 = 0; -x_4 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_486____closed__3; +x_4 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_485____closed__3; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); return x_5; } @@ -791,19 +823,18 @@ static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_523_ _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("proj", 4, 4); +x_1 = lean_mk_string_unchecked("debug", 5, 5); return x_1; } } static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_523____closed__2() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__1; -x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_375____closed__1; -x_3 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_523____closed__1; -x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); -return x_4; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_523____closed__1; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; } } static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_523____closed__3() { @@ -831,7 +862,7 @@ static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_560_ _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("parent", 6, 6); +x_1 = lean_mk_string_unchecked("proofs", 6, 6); return x_1; } } @@ -840,7 +871,7 @@ static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_560_ { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__1; -x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_375____closed__1; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_523____closed__1; x_3 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_560____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; @@ -867,6 +898,318 @@ x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); return x_5; } } +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_597____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("congr", 5, 5); +return x_1; +} +} +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_597____closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__1; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_523____closed__1; +x_3 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_597____closed__1; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_597____closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__17; +x_2 = lean_unsigned_to_nat(597u); +x_3 = l_Lean_Name_num___override(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_597_(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_Meta_Tactic_Grind___hyg_597____closed__2; +x_3 = 0; +x_4 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_597____closed__3; +x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); +return x_5; +} +} +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_634____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("proof", 5, 5); +return x_1; +} +} +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_634____closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__1; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_523____closed__1; +x_3 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_634____closed__1; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_634____closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__17; +x_2 = lean_unsigned_to_nat(634u); +x_3 = l_Lean_Name_num___override(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_634_(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_Meta_Tactic_Grind___hyg_634____closed__2; +x_3 = 0; +x_4 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_634____closed__3; +x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); +return x_5; +} +} +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_671____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("proj", 4, 4); +return x_1; +} +} +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_671____closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__1; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_523____closed__1; +x_3 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_671____closed__1; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_671____closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__17; +x_2 = lean_unsigned_to_nat(671u); +x_3 = l_Lean_Name_num___override(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_671_(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_Meta_Tactic_Grind___hyg_671____closed__2; +x_3 = 0; +x_4 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_671____closed__3; +x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); +return x_5; +} +} +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_708____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("parent", 6, 6); +return x_1; +} +} +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_708____closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__1; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_523____closed__1; +x_3 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_708____closed__1; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_708____closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__17; +x_2 = lean_unsigned_to_nat(708u); +x_3 = l_Lean_Name_num___override(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_708_(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_Meta_Tactic_Grind___hyg_708____closed__2; +x_3 = 0; +x_4 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_708____closed__3; +x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); +return x_5; +} +} +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_745____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("final", 5, 5); +return x_1; +} +} +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_745____closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__1; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_523____closed__1; +x_3 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_745____closed__1; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_745____closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__17; +x_2 = lean_unsigned_to_nat(745u); +x_3 = l_Lean_Name_num___override(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_745_(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_Meta_Tactic_Grind___hyg_745____closed__2; +x_3 = 0; +x_4 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_745____closed__3; +x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); +return x_5; +} +} +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_782____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("forallPropagator", 16, 16); +return x_1; +} +} +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_782____closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__1; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_523____closed__1; +x_3 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_782____closed__1; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_782____closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__17; +x_2 = lean_unsigned_to_nat(782u); +x_3 = l_Lean_Name_num___override(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_782_(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_Meta_Tactic_Grind___hyg_782____closed__2; +x_3 = 0; +x_4 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_782____closed__3; +x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); +return x_5; +} +} +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_819____closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__1; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_523____closed__1; +x_3 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_411____closed__1; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_819____closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__17; +x_2 = lean_unsigned_to_nat(819u); +x_3 = l_Lean_Name_num___override(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_819_(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_Meta_Tactic_Grind___hyg_819____closed__1; +x_3 = 0; +x_4 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_819____closed__2; +x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); +return x_5; +} +} +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_856____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("canon", 5, 5); +return x_1; +} +} +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_856____closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__1; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_523____closed__1; +x_3 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_856____closed__1; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_856____closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__17; +x_2 = lean_unsigned_to_nat(856u); +x_3 = l_Lean_Name_num___override(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_856_(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_Meta_Tactic_Grind___hyg_856____closed__2; +x_3 = 0; +x_4 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_856____closed__3; +x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); +return x_5; +} +} lean_object* initialize_Lean_Meta_Tactic_Grind_Attr(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_Tactic_Grind_RevertAll(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_Tactic_Grind_Types(uint8_t builtin, lean_object*); @@ -886,6 +1229,7 @@ lean_object* initialize_Lean_Meta_Tactic_Grind_Parser(uint8_t builtin, lean_obje lean_object* initialize_Lean_Meta_Tactic_Grind_EMatchTheorem(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_Tactic_Grind_EMatch(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_Tactic_Grind_Main(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Meta_Tactic_Grind_CasesMatch(uint8_t builtin, lean_object*); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Lean_Meta_Tactic_Grind(uint8_t builtin, lean_object* w) { lean_object * res; @@ -948,6 +1292,9 @@ lean_dec_ref(res); res = initialize_Lean_Meta_Tactic_Grind_Main(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +res = initialize_Lean_Meta_Tactic_Grind_CasesMatch(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_Grind___hyg_4____closed__1 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__1(); lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__1); l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__2 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__2(); @@ -1068,40 +1415,40 @@ lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_337____cl if (builtin) {res = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_337_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -}l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_375____closed__1 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_375____closed__1(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_375____closed__1); -l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_375____closed__2 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_375____closed__2(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_375____closed__2); -l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_375____closed__3 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_375____closed__3(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_375____closed__3); -if (builtin) {res = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_375_(lean_io_mk_world()); +}l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_374____closed__1 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_374____closed__1(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_374____closed__1); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_374____closed__2 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_374____closed__2(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_374____closed__2); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_374____closed__3 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_374____closed__3(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_374____closed__3); +if (builtin) {res = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_374_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -}l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_412____closed__1 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_412____closed__1(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_412____closed__1); -l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_412____closed__2 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_412____closed__2(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_412____closed__2); -l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_412____closed__3 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_412____closed__3(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_412____closed__3); -if (builtin) {res = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_412_(lean_io_mk_world()); +}l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_411____closed__1 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_411____closed__1(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_411____closed__1); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_411____closed__2 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_411____closed__2(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_411____closed__2); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_411____closed__3 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_411____closed__3(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_411____closed__3); +if (builtin) {res = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_411_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -}l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_449____closed__1 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_449____closed__1(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_449____closed__1); -l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_449____closed__2 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_449____closed__2(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_449____closed__2); -l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_449____closed__3 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_449____closed__3(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_449____closed__3); -if (builtin) {res = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_449_(lean_io_mk_world()); +}l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_448____closed__1 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_448____closed__1(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_448____closed__1); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_448____closed__2 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_448____closed__2(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_448____closed__2); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_448____closed__3 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_448____closed__3(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_448____closed__3); +if (builtin) {res = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_448_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -}l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_486____closed__1 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_486____closed__1(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_486____closed__1); -l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_486____closed__2 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_486____closed__2(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_486____closed__2); -l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_486____closed__3 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_486____closed__3(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_486____closed__3); -if (builtin) {res = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_486_(lean_io_mk_world()); +}l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_485____closed__1 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_485____closed__1(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_485____closed__1); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_485____closed__2 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_485____closed__2(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_485____closed__2); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_485____closed__3 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_485____closed__3(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_485____closed__3); +if (builtin) {res = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_485_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); }l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_523____closed__1 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_523____closed__1(); @@ -1122,6 +1469,76 @@ lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_560____cl if (builtin) {res = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_560_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +}l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_597____closed__1 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_597____closed__1(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_597____closed__1); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_597____closed__2 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_597____closed__2(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_597____closed__2); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_597____closed__3 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_597____closed__3(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_597____closed__3); +if (builtin) {res = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_597_(lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +}l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_634____closed__1 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_634____closed__1(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_634____closed__1); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_634____closed__2 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_634____closed__2(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_634____closed__2); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_634____closed__3 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_634____closed__3(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_634____closed__3); +if (builtin) {res = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_634_(lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +}l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_671____closed__1 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_671____closed__1(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_671____closed__1); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_671____closed__2 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_671____closed__2(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_671____closed__2); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_671____closed__3 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_671____closed__3(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_671____closed__3); +if (builtin) {res = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_671_(lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +}l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_708____closed__1 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_708____closed__1(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_708____closed__1); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_708____closed__2 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_708____closed__2(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_708____closed__2); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_708____closed__3 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_708____closed__3(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_708____closed__3); +if (builtin) {res = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_708_(lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +}l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_745____closed__1 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_745____closed__1(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_745____closed__1); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_745____closed__2 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_745____closed__2(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_745____closed__2); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_745____closed__3 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_745____closed__3(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_745____closed__3); +if (builtin) {res = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_745_(lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +}l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_782____closed__1 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_782____closed__1(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_782____closed__1); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_782____closed__2 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_782____closed__2(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_782____closed__2); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_782____closed__3 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_782____closed__3(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_782____closed__3); +if (builtin) {res = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_782_(lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +}l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_819____closed__1 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_819____closed__1(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_819____closed__1); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_819____closed__2 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_819____closed__2(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_819____closed__2); +if (builtin) {res = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_819_(lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +}l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_856____closed__1 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_856____closed__1(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_856____closed__1); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_856____closed__2 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_856____closed__2(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_856____closed__2); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_856____closed__3 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_856____closed__3(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_856____closed__3); +if (builtin) {res = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_856_(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/Tactic/Grind/Canon.c b/stage0/stdlib/Lean/Meta/Tactic/Grind/Canon.c index 29d5704f10e6..b1a5fae1b06b 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Grind/Canon.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Grind/Canon.c @@ -13,191 +13,273 @@ #ifdef __cplusplus extern "C" { #endif +LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Meta_Grind_Canon_canonImpl_visit___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_PersistentHashMap_findAux___at_Lean_Meta_Grind_Canon_canonElemCore___spec__16___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_canonElemCore___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10___lambda__3___closed__5; +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__6(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__12___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_CanonElemKind_explain(uint8_t); lean_object* l_Lean_mkAppN(lean_object*, lean_object*); size_t lean_usize_shift_right(size_t, size_t); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Canon_0__Lean_Meta_Grind_Canon_ShouldCanonResult_noConfusion___rarg(uint8_t, uint8_t, lean_object*); -static lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__3; +static lean_object* l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__3___closed__5; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_CanonElemKind_toCtorIdx___boxed(lean_object*); lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_Canon_canonElemCore___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_Canon_canonElemCore___lambda__2___closed__6; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Canon_0__Lean_Meta_Grind_Canon_ShouldCanonResult_noConfusion(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_Canon_canonElemCore___spec__8(size_t, 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_Meta_Grind_Canon_shouldCanon(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_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__1(lean_object*, lean_object*); uint64_t lean_uint64_of_nat(lean_object*); static lean_object* l_Lean_Meta_Grind_Canon_canonImpl___closed__1; uint64_t lean_uint64_mix_hash(uint64_t, uint64_t); +static lean_object* l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Canon_0__Lean_Meta_Grind_Canon_beqCanonElemKind____x40_Lean_Meta_Tactic_Grind_Canon___hyg_87____boxed(lean_object*, lean_object*); lean_object* l_Lean_Meta_isExprDefEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__4___closed__1; size_t lean_uint64_to_usize(uint64_t); uint64_t lean_uint64_lor(uint64_t, uint64_t); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10___lambda__4___boxed(lean_object**); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_canonElemCore(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Expr_isApp(lean_object*); +static lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10___lambda__3___closed__4; +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__11(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___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*); +LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___boxed(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*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_Canon_canonElemCore___spec__3___boxed(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*); size_t lean_usize_mul(size_t, size_t); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Canon_0__Lean_Meta_Grind_Canon_ShouldCanonResult_noConfusion___rarg___lambda__1___boxed(lean_object*); lean_object* lean_mk_array(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__4___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__5(lean_object*, lean_object*); uint8_t lean_usize_dec_eq(size_t, size_t); +static lean_object* l_Lean_Meta_Grind_Canon_instReprShouldCanonResult___closed__2; LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_Canon_canonElemCore___lambda__2___closed__5; lean_object* lean_array_fget(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__7(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__3(lean_object*); -static lean_object* l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__1___closed__1; lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__11(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_Canon_canonImpl_visit___closed__2; +static lean_object* l_Lean_Meta_Grind_Canon_CanonElemKind_explain___closed__1; +static lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10___lambda__3___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_Canon_canonElemCore___spec__16(lean_object*, size_t, lean_object*); -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand_go___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__4(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_Canon_canonElemCore___spec__9(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_Canon_instBEqCanonElemKind___closed__1; static lean_object* l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1___closed__5; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_CanonElemKind_noConfusion(lean_object*); +static lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__11___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_canonElemCore___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static double l_Lean_addTrace___at_Lean_Meta_Grind_Canon_canonElemCore___spec__5___closed__1; +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10___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_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__7___closed__2; lean_object* l_Lean_stringToMessageData(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_canonElemCore___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_Grind_Canon_instReprShouldCanonResult(uint8_t, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Canon_0__Lean_Meta_Grind_Canon_ShouldCanonResult_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__3___closed__2; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_instReprShouldCanonResult___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__2(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___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_EXPORT lean_object* l_Lean_Meta_Grind_Canon_CanonElemKind_noConfusion___rarg___lambda__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_canonImpl_visit___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_Grind_Canon_canonImpl_visit___lambda__4___closed__2; +static lean_object* l_Lean_Meta_Grind_Canon_instReprShouldCanonResult___closed__6; +static lean_object* l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___closed__4; size_t lean_ptr_addr(lean_object*); lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkPtrMap___rarg(lean_object*); -LEAN_EXPORT uint8_t l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__2(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_Canon_canonElemCore___lambda__2___closed__3; +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__4(lean_object*, lean_object*, lean_object*); size_t lean_usize_of_nat(lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10___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*, lean_object*); static lean_object* l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_Canon_canonElemCore___spec__7(lean_object*, size_t, size_t, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_Canon_canonImpl_visit___closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_canonElemCore___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_PersistentHashMap_insertAux___at_Lean_Meta_Grind_Canon_canonElemCore___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_take(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_Canon_canonElemCore___lambda__2___closed__2; static lean_object* l_Lean_Meta_Grind_Canon_instInhabitedState___closed__1; uint8_t lean_expr_eqv(lean_object*, lean_object*); static lean_object* l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1___closed__8; uint64_t lean_uint64_shift_right(uint64_t, uint64_t); static lean_object* l_Lean_addTrace___at_Lean_Meta_Grind_Canon_canonElemCore___spec__5___closed__2; -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__2___boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__3___closed__3; lean_object* l_Lean_Meta_getFunInfo(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint64_t lean_usize_to_uint64(size_t); +static lean_object* l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___closed__3; lean_object* lean_nat_div(lean_object*, lean_object*); +uint8_t l_Lean_Meta_ParamInfo_isImplicit(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_Canon_canonElemCore___spec__13___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Canon_0__Lean_Meta_Grind_Canon_ShouldCanonResult_noConfusion___rarg___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Canon_0__Lean_Meta_Grind_Canon_ShouldCanonResult_noConfusion___rarg___lambda__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_Canon_canonElemCore___spec__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_MessageData_ofFormat(lean_object*); +lean_object* l_ReaderT_instMonadLift(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_canonImpl_visit___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_Grind_Canon_canon___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_Grind_Canon_instReprShouldCanonResult___closed__1; +lean_object* l_Lean_Expr_forallE___override(lean_object*, lean_object*, lean_object*, uint8_t); lean_object* l_outOfBounds___rarg(lean_object*); +LEAN_EXPORT lean_object* l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__2(lean_object*); +static lean_object* l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__3___closed__4; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_get(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__11___closed__4; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_CanonElemKind_toCtorIdx(uint8_t); lean_object* lean_st_mk_ref(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_Canon_canonImpl_visit___closed__1; +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Expr_hasLooseBVars(lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__1___boxed(lean_object*, lean_object*); extern lean_object* l_Lean_levelZero; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_Canon_canonElemCore___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Meta_Grind_Canon_instInhabitedShouldCanonResult; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_Canon_canonElemCore___spec__12(lean_object*, size_t, size_t, lean_object*, lean_object*); extern lean_object* l_Lean_instInhabitedExpr; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_Canon_canonElemCore___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_canonElemCore___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*); +static lean_object* l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___closed__7; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_canonElemCore___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_PersistentHashMap_insertAux___at_Lean_Meta_Grind_Canon_canonElemCore___spec__7___closed__1; uint64_t l_Lean_Expr_hash(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_Canon_canonElemCore___spec__2___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__3___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_canon___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Canon_0__Lean_Meta_Grind_Canon_ShouldCanonResult_toCtorIdx(uint8_t); lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Meta_ParamInfo_isInstImplicit(lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10___boxed(lean_object**); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_shouldCanon___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_Meta_Grind_Canon_canonImpl_visit___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_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Meta_instMonadMetaM; +static lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__11___closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_canonImpl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__5___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__6(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_Canon_canonElemCore___lambda__3___closed__1; lean_object* lean_usize_to_nat(size_t); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_canon(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_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*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_MessageData_ofExpr(lean_object*); +static lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__11___closed__5; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_canonElemCore___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*); double l_Float_ofScientific(lean_object*, uint8_t, lean_object*); -uint8_t l_Lean_Expr_fvarsSubset(lean_object*, lean_object*); +static lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10___lambda__3___closed__6; +static lean_object* l_Lean_Meta_Grind_Canon_instReprShouldCanonResult___closed__4; +static lean_object* l_Lean_Meta_Grind_Canon_CanonElemKind_noConfusion___rarg___closed__1; +lean_object* l_StateT_lift___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_shouldCanon___lambda__1___boxed(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_Grind_Canon_canonElemCore___spec__13(size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__1___boxed(lean_object*, lean_object*); lean_object* l_StateT_instMonad___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Meta_Grind_Canon_canonElemCore___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_shouldCanon___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_canonImplicit(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_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1(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_addTrace___at_Lean_Meta_Grind_Canon_canonElemCore___spec__5___closed__3; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_Canon_canonElemCore___spec__2(lean_object*, size_t, lean_object*); +static lean_object* l_Lean_Meta_Grind_Canon_CanonElemKind_explain___closed__3; +static lean_object* l_Lean_Meta_Grind_Canon_instReprShouldCanonResult___closed__7; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_CanonElemKind_noConfusion___rarg___lambda__1___boxed(lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_Canon_instReprShouldCanonResult___closed__3; +static lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10___lambda__3___closed__2; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_canon_unsafe__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_Canon_canonElemCore___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); +static lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10___lambda__3___closed__3; +static uint64_t l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1___closed__11; lean_object* l_Lean_indentExpr(lean_object*); lean_object* l_Lean_PersistentHashMap_mkEmptyEntries(lean_object*, lean_object*); -static lean_object* l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__1___closed__3; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_CanonElemKind_explain___boxed(lean_object*); +static lean_object* l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__7___closed__3; static lean_object* l_Lean_Meta_Grind_Canon_canonElemCore___lambda__2___closed__1; static size_t l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_Canon_canonElemCore___spec__2___closed__1; static lean_object* l_Lean_Meta_Grind_Canon_instInhabitedState___closed__2; +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10___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* l_StateRefT_x27_lift(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isConstOf(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___closed__6; lean_object* lean_array_set(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*); uint64_t lean_uint64_xor(uint64_t, uint64_t); lean_object* lean_panic_fn(lean_object*, lean_object*); +static uint64_t l_Lean_Meta_Grind_Canon_canonImplicit___closed__1; LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Meta_Grind_Canon_canonElemCore___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__5(lean_object*, lean_object*, lean_object*); lean_object* lean_nat_sub(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__1___closed__1; -static lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__1; lean_object* lean_nat_mul(lean_object*, lean_object*); lean_object* l_Lean_Meta_isTypeFormer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_Canon_canonElemCore___lambda__2___closed__4; uint64_t lean_uint64_shift_left(uint64_t, uint64_t); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_canonImpl_visit(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1___closed__2; -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8___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* l_Lean_PersistentHashMap_mkCollisionNode___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1___closed__7; lean_object* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_instBEqCanonElemKind; size_t lean_usize_sub(size_t, size_t); lean_object* lean_array_mk(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_canonElemCore___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_Grind_Canon_CanonElemKind_explain___closed__2; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_CanonElemKind_noConfusion___rarg(uint8_t, uint8_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_canonType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1___closed__10; static uint64_t l_Lean_Meta_Grind_Canon_canonInst___closed__1; lean_object* l_Lean_isTracingEnabledForCore(lean_object*, lean_object*, lean_object*); size_t lean_usize_add(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_Canon_canonElemCore___spec__17___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_Canon_canonElemCore___spec__17(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_uget(lean_object*, size_t); -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___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 uint8_t l___private_Lean_Meta_Tactic_Grind_Canon_0__Lean_Meta_Grind_Canon_beqCanonElemKind____x40_Lean_Meta_Tactic_Grind_Canon___hyg_87_(uint8_t, uint8_t); lean_object* l_instInhabitedOfMonad___rarg(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_Canon_canonImpl_visit___closed__4; lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); size_t lean_usize_shift_left(size_t, size_t); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__12(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_instInhabitedState; +static lean_object* l_Lean_Meta_Grind_Canon_instReprShouldCanonResult___closed__8; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_canonInst(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___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_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); +static lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__11___closed__1; +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__2(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*); uint64_t l_Lean_Meta_TransparencyMode_toUInt64(uint8_t); -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___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*, lean_object*, lean_object*); lean_object* lean_nat_add(lean_object*, lean_object*); static lean_object* l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1___closed__6; -static uint64_t l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1___closed__9; +static lean_object* l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1___closed__9; lean_object* l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(lean_object*); static lean_object* l_Lean_Meta_Grind_Canon_instInhabitedState___closed__3; +static lean_object* l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__7___closed__1; +static lean_object* l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___closed__5; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(lean_object*, lean_object*); static lean_object* l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1___closed__3; -lean_object* l_Lean_isDiagnosticsEnabled(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__2; +static lean_object* l_Lean_Meta_Grind_Canon_instReprShouldCanonResult___closed__5; +uint8_t l_Lean_Expr_isForall(lean_object*); +static lean_object* l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___closed__2; static lean_object* l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1___closed__4; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Canon_0__Lean_Meta_Grind_Canon_ShouldCanonResult_toCtorIdx___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_uset(lean_object*, size_t, lean_object*); lean_object* lean_array_get(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10(lean_object*, lean_object*); -static lean_object* l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__1___closed__2; +lean_object* l___private_Init_Data_Repr_0__Nat_reprFast(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_Canon_canonElemCore___spec__14(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__4; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_canonImpl_visit___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* l_ReaderT_instMonad___rarg(lean_object*); size_t lean_usize_land(size_t, size_t); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand_go___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__3(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_CanonElemKind_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); static size_t l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_Canon_canonElemCore___spec__2___closed__2; +uint8_t l___private_Lean_Expr_0__Lean_beqBinderInfo____x40_Lean_Expr___hyg_406_(uint8_t, uint8_t); static lean_object* _init_l_Lean_Meta_Grind_Canon_instInhabitedState___closed__1() { _start: { @@ -236,6 +318,193 @@ x_1 = l_Lean_Meta_Grind_Canon_instInhabitedState___closed__3; return x_1; } } +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_CanonElemKind_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_Meta_Grind_Canon_CanonElemKind_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_Meta_Grind_Canon_CanonElemKind_toCtorIdx(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_CanonElemKind_noConfusion___rarg___lambda__1(lean_object* x_1) { +_start: +{ +lean_inc(x_1); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_Canon_CanonElemKind_noConfusion___rarg___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_Canon_CanonElemKind_noConfusion___rarg___lambda__1___boxed), 1, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_CanonElemKind_noConfusion___rarg(uint8_t x_1, uint8_t x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Lean_Meta_Grind_Canon_CanonElemKind_noConfusion___rarg___closed__1; +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_CanonElemKind_noConfusion(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_Canon_CanonElemKind_noConfusion___rarg___boxed), 3, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_CanonElemKind_noConfusion___rarg___lambda__1___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Lean_Meta_Grind_Canon_CanonElemKind_noConfusion___rarg___lambda__1(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_CanonElemKind_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_Meta_Grind_Canon_CanonElemKind_noConfusion___rarg(x_4, x_5, x_3); +return x_6; +} +} +LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_Grind_Canon_0__Lean_Meta_Grind_Canon_beqCanonElemKind____x40_Lean_Meta_Tactic_Grind_Canon___hyg_87_(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_Meta_Grind_Canon_CanonElemKind_toCtorIdx(x_1); +x_4 = l_Lean_Meta_Grind_Canon_CanonElemKind_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_Meta_Tactic_Grind_Canon_0__Lean_Meta_Grind_Canon_beqCanonElemKind____x40_Lean_Meta_Tactic_Grind_Canon___hyg_87____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_Meta_Tactic_Grind_Canon_0__Lean_Meta_Grind_Canon_beqCanonElemKind____x40_Lean_Meta_Tactic_Grind_Canon___hyg_87_(x_3, x_4); +x_6 = lean_box(x_5); +return x_6; +} +} +static lean_object* _init_l_Lean_Meta_Grind_Canon_instBEqCanonElemKind___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Canon_0__Lean_Meta_Grind_Canon_beqCanonElemKind____x40_Lean_Meta_Tactic_Grind_Canon___hyg_87____boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_Canon_instBEqCanonElemKind() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Meta_Grind_Canon_instBEqCanonElemKind___closed__1; +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_Canon_CanonElemKind_explain___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("type class instances", 20, 20); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_Canon_CanonElemKind_explain___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("types (or type formers)", 23, 23); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_Canon_CanonElemKind_explain___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("implicit arguments (which are not type class instances or types)", 64, 64); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_CanonElemKind_explain(uint8_t x_1) { +_start: +{ +switch (x_1) { +case 0: +{ +lean_object* x_2; +x_2 = l_Lean_Meta_Grind_Canon_CanonElemKind_explain___closed__1; +return x_2; +} +case 1: +{ +lean_object* x_3; +x_3 = l_Lean_Meta_Grind_Canon_CanonElemKind_explain___closed__2; +return x_3; +} +default: +{ +lean_object* x_4; +x_4 = l_Lean_Meta_Grind_Canon_CanonElemKind_explain___closed__3; +return x_4; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_CanonElemKind_explain___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_Meta_Grind_Canon_CanonElemKind_explain(x_2); +return x_3; +} +} LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_Canon_canonElemCore___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { @@ -1581,7 +1850,7 @@ static lean_object* _init_l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canon _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("the following `grind` static elements are definitionally equal with `default` transparency, but not with `instances` transparency", 129, 129); +x_1 = lean_mk_string_unchecked("the following ", 14, 14); return x_1; } } @@ -1598,7 +1867,7 @@ static lean_object* _init_l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canon _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("\nand", 4, 4); +x_1 = lean_mk_string_unchecked(" are definitionally equal with `default` transparency but not with a more restrictive transparency", 98, 98); return x_1; } } @@ -1614,13 +1883,30 @@ return x_2; static lean_object* _init_l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1___closed__8() { _start: { +lean_object* x_1; +x_1 = lean_mk_string_unchecked("\nand", 4, 4); +return x_1; +} +} +static lean_object* _init_l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1___closed__8; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1___closed__10() { +_start: +{ lean_object* x_1; lean_object* x_2; x_1 = l_Lean_addTrace___at_Lean_Meta_Grind_Canon_canonElemCore___spec__5___closed__2; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static uint64_t _init_l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1___closed__9() { +static uint64_t _init_l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1___closed__11() { _start: { uint8_t x_1; uint64_t x_2; @@ -1632,595 +1918,551 @@ return x_2; LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___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, lean_object* x_11) { _start: { -if (x_1 == 0) -{ -lean_object* x_12; lean_object* x_13; lean_object* x_14; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_4); -lean_dec(x_3); -x_12 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_12, 0, x_2); -x_13 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_13, 0, x_12); -lean_ctor_set(x_13, 1, x_6); -x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_13); -lean_ctor_set(x_14, 1, x_11); -return x_14; -} -else -{ -lean_object* x_15; lean_object* x_16; lean_object* x_75; lean_object* x_76; uint8_t x_77; -x_75 = l_Lean_isDiagnosticsEnabled(x_9, x_10, x_11); -x_76 = lean_ctor_get(x_75, 0); -lean_inc(x_76); -x_77 = lean_unbox(x_76); -if (x_77 == 0) +uint8_t x_12; uint8_t x_13; +x_12 = 1; +x_13 = l___private_Lean_Meta_Tactic_Grind_Canon_0__Lean_Meta_Grind_Canon_beqCanonElemKind____x40_Lean_Meta_Tactic_Grind_Canon___hyg_87_(x_1, x_12); +if (x_13 == 0) { -uint8_t x_78; -x_78 = !lean_is_exclusive(x_75); -if (x_78 == 0) +lean_object* x_14; lean_object* x_15; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; uint8_t x_83; +x_79 = l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1___closed__3; +x_80 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_Canon_canonElemCore___spec__4(x_79, x_6, x_7, x_8, x_9, x_10, x_11); +x_81 = lean_ctor_get(x_80, 0); +lean_inc(x_81); +x_82 = lean_ctor_get(x_81, 0); +lean_inc(x_82); +x_83 = lean_unbox(x_82); +if (x_83 == 0) { -lean_object* x_79; lean_object* x_80; -x_79 = lean_ctor_get(x_75, 1); -x_80 = lean_ctor_get(x_75, 0); +lean_object* x_84; uint8_t x_85; +x_84 = lean_ctor_get(x_80, 1); +lean_inc(x_84); lean_dec(x_80); -lean_ctor_set(x_75, 1, x_6); -x_15 = x_75; -x_16 = x_79; -goto block_74; +x_85 = !lean_is_exclusive(x_81); +if (x_85 == 0) +{ +lean_object* x_86; +x_86 = lean_ctor_get(x_81, 0); +lean_dec(x_86); +x_14 = x_81; +x_15 = x_84; +goto block_78; } else { -lean_object* x_81; lean_object* x_82; -x_81 = lean_ctor_get(x_75, 1); -lean_inc(x_81); -lean_dec(x_75); -x_82 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_82, 0, x_76); -lean_ctor_set(x_82, 1, x_6); -x_15 = x_82; -x_16 = x_81; -goto block_74; +lean_object* x_87; lean_object* x_88; +x_87 = lean_ctor_get(x_81, 1); +lean_inc(x_87); +lean_dec(x_81); +x_88 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_88, 0, x_82); +lean_ctor_set(x_88, 1, x_87); +x_14 = x_88; +x_15 = x_84; +goto block_78; } } else { -uint8_t x_83; -lean_dec(x_76); -x_83 = !lean_is_exclusive(x_75); -if (x_83 == 0) -{ -lean_object* x_84; lean_object* x_85; uint8_t x_86; -x_84 = lean_ctor_get(x_75, 1); -x_85 = lean_ctor_get(x_75, 0); -lean_dec(x_85); -lean_inc(x_3); -x_86 = l_Lean_Expr_fvarsSubset(x_4, x_3); -if (x_86 == 0) -{ -lean_object* x_87; -x_87 = lean_box(x_86); -lean_ctor_set(x_75, 1, x_6); -lean_ctor_set(x_75, 0, x_87); -x_15 = x_75; -x_16 = x_84; -goto block_74; -} -else +lean_object* x_89; lean_object* x_90; uint8_t x_91; +lean_dec(x_82); +x_89 = lean_ctor_get(x_7, 0); +lean_inc(x_89); +x_90 = lean_ctor_get(x_80, 1); +lean_inc(x_90); +lean_dec(x_80); +x_91 = !lean_is_exclusive(x_81); +if (x_91 == 0) { -lean_object* x_88; uint64_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; uint8_t x_97; -x_88 = lean_ctor_get(x_7, 0); -lean_inc(x_88); -x_89 = lean_ctor_get_uint64(x_7, sizeof(void*)*7); -x_90 = lean_ctor_get_uint8(x_7, sizeof(void*)*7 + 8); -x_91 = lean_ctor_get(x_7, 1); -lean_inc(x_91); -x_92 = lean_ctor_get(x_7, 2); -lean_inc(x_92); -x_93 = lean_ctor_get(x_7, 3); -lean_inc(x_93); -x_94 = lean_ctor_get(x_7, 4); -lean_inc(x_94); -x_95 = lean_ctor_get(x_7, 5); -lean_inc(x_95); -x_96 = lean_ctor_get(x_7, 6); +lean_object* x_92; lean_object* x_93; uint64_t 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; uint8_t x_102; +x_92 = lean_ctor_get(x_81, 1); +x_93 = lean_ctor_get(x_81, 0); +lean_dec(x_93); +x_94 = lean_ctor_get_uint64(x_7, sizeof(void*)*7); +x_95 = lean_ctor_get_uint8(x_7, sizeof(void*)*7 + 8); +x_96 = lean_ctor_get(x_7, 1); lean_inc(x_96); -x_97 = !lean_is_exclusive(x_88); -if (x_97 == 0) -{ -uint8_t x_98; uint8_t x_99; uint8_t x_100; uint64_t x_101; uint64_t x_102; uint64_t x_103; uint64_t x_104; uint64_t x_105; lean_object* x_106; lean_object* x_107; -x_98 = lean_ctor_get_uint8(x_7, sizeof(void*)*7 + 9); -x_99 = lean_ctor_get_uint8(x_7, sizeof(void*)*7 + 10); -x_100 = 1; -lean_ctor_set_uint8(x_88, 9, x_100); -x_101 = 2; -x_102 = lean_uint64_shift_right(x_89, x_101); -x_103 = lean_uint64_shift_left(x_102, x_101); -x_104 = l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1___closed__9; -x_105 = lean_uint64_lor(x_103, x_104); -x_106 = lean_alloc_ctor(0, 7, 11); -lean_ctor_set(x_106, 0, x_88); -lean_ctor_set(x_106, 1, x_91); -lean_ctor_set(x_106, 2, x_92); -lean_ctor_set(x_106, 3, x_93); -lean_ctor_set(x_106, 4, x_94); -lean_ctor_set(x_106, 5, x_95); -lean_ctor_set(x_106, 6, x_96); -lean_ctor_set_uint64(x_106, sizeof(void*)*7, x_105); -lean_ctor_set_uint8(x_106, sizeof(void*)*7 + 8, x_90); -lean_ctor_set_uint8(x_106, sizeof(void*)*7 + 9, x_98); -lean_ctor_set_uint8(x_106, sizeof(void*)*7 + 10, x_99); +x_97 = lean_ctor_get(x_7, 2); +lean_inc(x_97); +x_98 = lean_ctor_get(x_7, 3); +lean_inc(x_98); +x_99 = lean_ctor_get(x_7, 4); +lean_inc(x_99); +x_100 = lean_ctor_get(x_7, 5); +lean_inc(x_100); +x_101 = lean_ctor_get(x_7, 6); +lean_inc(x_101); +x_102 = !lean_is_exclusive(x_89); +if (x_102 == 0) +{ +uint8_t x_103; uint8_t x_104; uint8_t x_105; uint64_t x_106; uint64_t x_107; uint64_t x_108; uint64_t x_109; uint64_t x_110; lean_object* x_111; lean_object* x_112; +x_103 = lean_ctor_get_uint8(x_7, sizeof(void*)*7 + 9); +x_104 = lean_ctor_get_uint8(x_7, sizeof(void*)*7 + 10); +x_105 = 1; +lean_ctor_set_uint8(x_89, 9, x_105); +x_106 = 2; +x_107 = lean_uint64_shift_right(x_94, x_106); +x_108 = lean_uint64_shift_left(x_107, x_106); +x_109 = l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1___closed__11; +x_110 = lean_uint64_lor(x_108, x_109); +x_111 = lean_alloc_ctor(0, 7, 11); +lean_ctor_set(x_111, 0, x_89); +lean_ctor_set(x_111, 1, x_96); +lean_ctor_set(x_111, 2, x_97); +lean_ctor_set(x_111, 3, x_98); +lean_ctor_set(x_111, 4, x_99); +lean_ctor_set(x_111, 5, x_100); +lean_ctor_set(x_111, 6, x_101); +lean_ctor_set_uint64(x_111, sizeof(void*)*7, x_110); +lean_ctor_set_uint8(x_111, sizeof(void*)*7 + 8, x_95); +lean_ctor_set_uint8(x_111, sizeof(void*)*7 + 9, x_103); +lean_ctor_set_uint8(x_111, sizeof(void*)*7 + 10, x_104); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_4); lean_inc(x_3); -x_107 = l_Lean_Meta_isExprDefEq(x_3, x_4, x_106, x_8, x_9, x_10, x_84); -if (lean_obj_tag(x_107) == 0) +x_112 = l_Lean_Meta_isExprDefEq(x_3, x_4, x_111, x_8, x_9, x_10, x_90); +if (lean_obj_tag(x_112) == 0) { -lean_object* x_108; lean_object* x_109; -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); -lean_ctor_set(x_75, 1, x_6); -lean_ctor_set(x_75, 0, x_108); -x_15 = x_75; -x_16 = x_109; -goto block_74; +lean_object* x_113; lean_object* x_114; +x_113 = lean_ctor_get(x_112, 0); +lean_inc(x_113); +x_114 = lean_ctor_get(x_112, 1); +lean_inc(x_114); +lean_dec(x_112); +lean_ctor_set(x_81, 0, x_113); +x_14 = x_81; +x_15 = x_114; +goto block_78; } else { -uint8_t x_110; -lean_free_object(x_75); +uint8_t x_115; +lean_free_object(x_81); +lean_dec(x_92); 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_110 = !lean_is_exclusive(x_107); -if (x_110 == 0) +x_115 = !lean_is_exclusive(x_112); +if (x_115 == 0) { -return x_107; +return x_112; } else { -lean_object* x_111; lean_object* x_112; lean_object* x_113; -x_111 = lean_ctor_get(x_107, 0); -x_112 = lean_ctor_get(x_107, 1); -lean_inc(x_112); -lean_inc(x_111); -lean_dec(x_107); -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 -{ -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; uint8_t x_123; uint8_t x_124; uint8_t x_125; uint8_t x_126; uint8_t x_127; uint8_t x_128; uint8_t x_129; uint8_t x_130; uint8_t x_131; uint8_t x_132; lean_object* x_133; uint64_t x_134; uint64_t x_135; uint64_t x_136; uint64_t x_137; uint64_t x_138; lean_object* x_139; lean_object* x_140; -x_114 = lean_ctor_get_uint8(x_7, sizeof(void*)*7 + 9); -x_115 = lean_ctor_get_uint8(x_7, sizeof(void*)*7 + 10); -x_116 = lean_ctor_get_uint8(x_88, 0); -x_117 = lean_ctor_get_uint8(x_88, 1); -x_118 = lean_ctor_get_uint8(x_88, 2); -x_119 = lean_ctor_get_uint8(x_88, 3); -x_120 = lean_ctor_get_uint8(x_88, 4); -x_121 = lean_ctor_get_uint8(x_88, 5); -x_122 = lean_ctor_get_uint8(x_88, 6); -x_123 = lean_ctor_get_uint8(x_88, 7); -x_124 = lean_ctor_get_uint8(x_88, 8); -x_125 = lean_ctor_get_uint8(x_88, 10); -x_126 = lean_ctor_get_uint8(x_88, 11); -x_127 = lean_ctor_get_uint8(x_88, 12); -x_128 = lean_ctor_get_uint8(x_88, 13); -x_129 = lean_ctor_get_uint8(x_88, 14); -x_130 = lean_ctor_get_uint8(x_88, 15); -x_131 = lean_ctor_get_uint8(x_88, 16); -lean_dec(x_88); -x_132 = 1; -x_133 = lean_alloc_ctor(0, 0, 17); -lean_ctor_set_uint8(x_133, 0, x_116); -lean_ctor_set_uint8(x_133, 1, x_117); -lean_ctor_set_uint8(x_133, 2, x_118); -lean_ctor_set_uint8(x_133, 3, x_119); -lean_ctor_set_uint8(x_133, 4, x_120); -lean_ctor_set_uint8(x_133, 5, x_121); -lean_ctor_set_uint8(x_133, 6, x_122); -lean_ctor_set_uint8(x_133, 7, x_123); -lean_ctor_set_uint8(x_133, 8, x_124); -lean_ctor_set_uint8(x_133, 9, x_132); -lean_ctor_set_uint8(x_133, 10, x_125); -lean_ctor_set_uint8(x_133, 11, x_126); -lean_ctor_set_uint8(x_133, 12, x_127); -lean_ctor_set_uint8(x_133, 13, x_128); -lean_ctor_set_uint8(x_133, 14, x_129); -lean_ctor_set_uint8(x_133, 15, x_130); -lean_ctor_set_uint8(x_133, 16, x_131); -x_134 = 2; -x_135 = lean_uint64_shift_right(x_89, x_134); -x_136 = lean_uint64_shift_left(x_135, x_134); -x_137 = l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1___closed__9; -x_138 = lean_uint64_lor(x_136, x_137); -x_139 = lean_alloc_ctor(0, 7, 11); -lean_ctor_set(x_139, 0, x_133); -lean_ctor_set(x_139, 1, x_91); -lean_ctor_set(x_139, 2, x_92); -lean_ctor_set(x_139, 3, x_93); -lean_ctor_set(x_139, 4, x_94); -lean_ctor_set(x_139, 5, x_95); -lean_ctor_set(x_139, 6, x_96); -lean_ctor_set_uint64(x_139, sizeof(void*)*7, x_138); -lean_ctor_set_uint8(x_139, sizeof(void*)*7 + 8, x_90); -lean_ctor_set_uint8(x_139, sizeof(void*)*7 + 9, x_114); -lean_ctor_set_uint8(x_139, sizeof(void*)*7 + 10, x_115); +lean_object* x_116; lean_object* x_117; lean_object* x_118; +x_116 = lean_ctor_get(x_112, 0); +x_117 = lean_ctor_get(x_112, 1); +lean_inc(x_117); +lean_inc(x_116); +lean_dec(x_112); +x_118 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_118, 0, x_116); +lean_ctor_set(x_118, 1, x_117); +return x_118; +} +} +} +else +{ +uint8_t x_119; uint8_t x_120; uint8_t x_121; uint8_t x_122; uint8_t x_123; uint8_t x_124; uint8_t x_125; uint8_t x_126; uint8_t x_127; uint8_t x_128; uint8_t x_129; uint8_t x_130; uint8_t x_131; uint8_t x_132; uint8_t x_133; uint8_t x_134; uint8_t x_135; uint8_t x_136; uint8_t x_137; lean_object* x_138; uint64_t x_139; uint64_t x_140; uint64_t x_141; uint64_t x_142; uint64_t x_143; lean_object* x_144; lean_object* x_145; +x_119 = lean_ctor_get_uint8(x_7, sizeof(void*)*7 + 9); +x_120 = lean_ctor_get_uint8(x_7, sizeof(void*)*7 + 10); +x_121 = lean_ctor_get_uint8(x_89, 0); +x_122 = lean_ctor_get_uint8(x_89, 1); +x_123 = lean_ctor_get_uint8(x_89, 2); +x_124 = lean_ctor_get_uint8(x_89, 3); +x_125 = lean_ctor_get_uint8(x_89, 4); +x_126 = lean_ctor_get_uint8(x_89, 5); +x_127 = lean_ctor_get_uint8(x_89, 6); +x_128 = lean_ctor_get_uint8(x_89, 7); +x_129 = lean_ctor_get_uint8(x_89, 8); +x_130 = lean_ctor_get_uint8(x_89, 10); +x_131 = lean_ctor_get_uint8(x_89, 11); +x_132 = lean_ctor_get_uint8(x_89, 12); +x_133 = lean_ctor_get_uint8(x_89, 13); +x_134 = lean_ctor_get_uint8(x_89, 14); +x_135 = lean_ctor_get_uint8(x_89, 15); +x_136 = lean_ctor_get_uint8(x_89, 16); +lean_dec(x_89); +x_137 = 1; +x_138 = lean_alloc_ctor(0, 0, 17); +lean_ctor_set_uint8(x_138, 0, x_121); +lean_ctor_set_uint8(x_138, 1, x_122); +lean_ctor_set_uint8(x_138, 2, x_123); +lean_ctor_set_uint8(x_138, 3, x_124); +lean_ctor_set_uint8(x_138, 4, x_125); +lean_ctor_set_uint8(x_138, 5, x_126); +lean_ctor_set_uint8(x_138, 6, x_127); +lean_ctor_set_uint8(x_138, 7, x_128); +lean_ctor_set_uint8(x_138, 8, x_129); +lean_ctor_set_uint8(x_138, 9, x_137); +lean_ctor_set_uint8(x_138, 10, x_130); +lean_ctor_set_uint8(x_138, 11, x_131); +lean_ctor_set_uint8(x_138, 12, x_132); +lean_ctor_set_uint8(x_138, 13, x_133); +lean_ctor_set_uint8(x_138, 14, x_134); +lean_ctor_set_uint8(x_138, 15, x_135); +lean_ctor_set_uint8(x_138, 16, x_136); +x_139 = 2; +x_140 = lean_uint64_shift_right(x_94, x_139); +x_141 = lean_uint64_shift_left(x_140, x_139); +x_142 = l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1___closed__11; +x_143 = lean_uint64_lor(x_141, x_142); +x_144 = lean_alloc_ctor(0, 7, 11); +lean_ctor_set(x_144, 0, x_138); +lean_ctor_set(x_144, 1, x_96); +lean_ctor_set(x_144, 2, x_97); +lean_ctor_set(x_144, 3, x_98); +lean_ctor_set(x_144, 4, x_99); +lean_ctor_set(x_144, 5, x_100); +lean_ctor_set(x_144, 6, x_101); +lean_ctor_set_uint64(x_144, sizeof(void*)*7, x_143); +lean_ctor_set_uint8(x_144, sizeof(void*)*7 + 8, x_95); +lean_ctor_set_uint8(x_144, sizeof(void*)*7 + 9, x_119); +lean_ctor_set_uint8(x_144, sizeof(void*)*7 + 10, x_120); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_4); lean_inc(x_3); -x_140 = l_Lean_Meta_isExprDefEq(x_3, x_4, x_139, x_8, x_9, x_10, x_84); -if (lean_obj_tag(x_140) == 0) +x_145 = l_Lean_Meta_isExprDefEq(x_3, x_4, x_144, x_8, x_9, x_10, x_90); +if (lean_obj_tag(x_145) == 0) { -lean_object* x_141; lean_object* x_142; -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); -lean_ctor_set(x_75, 1, x_6); -lean_ctor_set(x_75, 0, x_141); -x_15 = x_75; -x_16 = x_142; -goto block_74; +lean_object* x_146; lean_object* x_147; +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); +lean_ctor_set(x_81, 0, x_146); +x_14 = x_81; +x_15 = x_147; +goto block_78; } else { -lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; -lean_free_object(x_75); +lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; +lean_free_object(x_81); +lean_dec(x_92); 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_143 = lean_ctor_get(x_140, 0); -lean_inc(x_143); -x_144 = lean_ctor_get(x_140, 1); -lean_inc(x_144); -if (lean_is_exclusive(x_140)) { - lean_ctor_release(x_140, 0); - lean_ctor_release(x_140, 1); - x_145 = x_140; +x_148 = lean_ctor_get(x_145, 0); +lean_inc(x_148); +x_149 = lean_ctor_get(x_145, 1); +lean_inc(x_149); +if (lean_is_exclusive(x_145)) { + lean_ctor_release(x_145, 0); + lean_ctor_release(x_145, 1); + x_150 = x_145; } else { - lean_dec_ref(x_140); - x_145 = lean_box(0); + lean_dec_ref(x_145); + x_150 = lean_box(0); } -if (lean_is_scalar(x_145)) { - x_146 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_150)) { + x_151 = lean_alloc_ctor(1, 2, 0); } else { - x_146 = x_145; -} -lean_ctor_set(x_146, 0, x_143); -lean_ctor_set(x_146, 1, x_144); -return x_146; + x_151 = x_150; } +lean_ctor_set(x_151, 0, x_148); +lean_ctor_set(x_151, 1, x_149); +return x_151; } } } else { -lean_object* x_147; uint8_t x_148; -x_147 = lean_ctor_get(x_75, 1); -lean_inc(x_147); -lean_dec(x_75); -lean_inc(x_3); -x_148 = l_Lean_Expr_fvarsSubset(x_4, x_3); -if (x_148 == 0) -{ -lean_object* x_149; lean_object* x_150; -x_149 = lean_box(x_148); -x_150 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_150, 0, x_149); -lean_ctor_set(x_150, 1, x_6); -x_15 = x_150; -x_16 = x_147; -goto block_74; -} -else -{ -lean_object* x_151; uint64_t x_152; uint8_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; uint8_t x_160; uint8_t x_161; uint8_t x_162; uint8_t x_163; uint8_t x_164; uint8_t x_165; uint8_t x_166; uint8_t x_167; uint8_t x_168; 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; lean_object* x_178; uint8_t x_179; lean_object* x_180; uint64_t x_181; uint64_t x_182; uint64_t x_183; uint64_t x_184; uint64_t x_185; lean_object* x_186; lean_object* x_187; -x_151 = lean_ctor_get(x_7, 0); -lean_inc(x_151); -x_152 = lean_ctor_get_uint64(x_7, sizeof(void*)*7); -x_153 = lean_ctor_get_uint8(x_7, sizeof(void*)*7 + 8); -x_154 = lean_ctor_get(x_7, 1); -lean_inc(x_154); -x_155 = lean_ctor_get(x_7, 2); +lean_object* x_152; uint64_t 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; uint8_t x_161; uint8_t x_162; uint8_t x_163; uint8_t x_164; uint8_t x_165; uint8_t x_166; uint8_t x_167; uint8_t x_168; 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; lean_object* x_179; uint8_t x_180; lean_object* x_181; uint64_t x_182; uint64_t x_183; uint64_t x_184; uint64_t x_185; uint64_t x_186; lean_object* x_187; lean_object* x_188; +x_152 = lean_ctor_get(x_81, 1); +lean_inc(x_152); +lean_dec(x_81); +x_153 = lean_ctor_get_uint64(x_7, sizeof(void*)*7); +x_154 = lean_ctor_get_uint8(x_7, sizeof(void*)*7 + 8); +x_155 = lean_ctor_get(x_7, 1); lean_inc(x_155); -x_156 = lean_ctor_get(x_7, 3); +x_156 = lean_ctor_get(x_7, 2); lean_inc(x_156); -x_157 = lean_ctor_get(x_7, 4); +x_157 = lean_ctor_get(x_7, 3); lean_inc(x_157); -x_158 = lean_ctor_get(x_7, 5); +x_158 = lean_ctor_get(x_7, 4); lean_inc(x_158); -x_159 = lean_ctor_get(x_7, 6); +x_159 = lean_ctor_get(x_7, 5); lean_inc(x_159); -x_160 = lean_ctor_get_uint8(x_7, sizeof(void*)*7 + 9); -x_161 = lean_ctor_get_uint8(x_7, sizeof(void*)*7 + 10); -x_162 = lean_ctor_get_uint8(x_151, 0); -x_163 = lean_ctor_get_uint8(x_151, 1); -x_164 = lean_ctor_get_uint8(x_151, 2); -x_165 = lean_ctor_get_uint8(x_151, 3); -x_166 = lean_ctor_get_uint8(x_151, 4); -x_167 = lean_ctor_get_uint8(x_151, 5); -x_168 = lean_ctor_get_uint8(x_151, 6); -x_169 = lean_ctor_get_uint8(x_151, 7); -x_170 = lean_ctor_get_uint8(x_151, 8); -x_171 = lean_ctor_get_uint8(x_151, 10); -x_172 = lean_ctor_get_uint8(x_151, 11); -x_173 = lean_ctor_get_uint8(x_151, 12); -x_174 = lean_ctor_get_uint8(x_151, 13); -x_175 = lean_ctor_get_uint8(x_151, 14); -x_176 = lean_ctor_get_uint8(x_151, 15); -x_177 = lean_ctor_get_uint8(x_151, 16); -if (lean_is_exclusive(x_151)) { - x_178 = x_151; +x_160 = lean_ctor_get(x_7, 6); +lean_inc(x_160); +x_161 = lean_ctor_get_uint8(x_7, sizeof(void*)*7 + 9); +x_162 = lean_ctor_get_uint8(x_7, sizeof(void*)*7 + 10); +x_163 = lean_ctor_get_uint8(x_89, 0); +x_164 = lean_ctor_get_uint8(x_89, 1); +x_165 = lean_ctor_get_uint8(x_89, 2); +x_166 = lean_ctor_get_uint8(x_89, 3); +x_167 = lean_ctor_get_uint8(x_89, 4); +x_168 = lean_ctor_get_uint8(x_89, 5); +x_169 = lean_ctor_get_uint8(x_89, 6); +x_170 = lean_ctor_get_uint8(x_89, 7); +x_171 = lean_ctor_get_uint8(x_89, 8); +x_172 = lean_ctor_get_uint8(x_89, 10); +x_173 = lean_ctor_get_uint8(x_89, 11); +x_174 = lean_ctor_get_uint8(x_89, 12); +x_175 = lean_ctor_get_uint8(x_89, 13); +x_176 = lean_ctor_get_uint8(x_89, 14); +x_177 = lean_ctor_get_uint8(x_89, 15); +x_178 = lean_ctor_get_uint8(x_89, 16); +if (lean_is_exclusive(x_89)) { + x_179 = x_89; } else { - lean_dec_ref(x_151); - x_178 = lean_box(0); -} -x_179 = 1; -if (lean_is_scalar(x_178)) { - x_180 = lean_alloc_ctor(0, 0, 17); -} else { - x_180 = x_178; -} -lean_ctor_set_uint8(x_180, 0, x_162); -lean_ctor_set_uint8(x_180, 1, x_163); -lean_ctor_set_uint8(x_180, 2, x_164); -lean_ctor_set_uint8(x_180, 3, x_165); -lean_ctor_set_uint8(x_180, 4, x_166); -lean_ctor_set_uint8(x_180, 5, x_167); -lean_ctor_set_uint8(x_180, 6, x_168); -lean_ctor_set_uint8(x_180, 7, x_169); -lean_ctor_set_uint8(x_180, 8, x_170); -lean_ctor_set_uint8(x_180, 9, x_179); -lean_ctor_set_uint8(x_180, 10, x_171); -lean_ctor_set_uint8(x_180, 11, x_172); -lean_ctor_set_uint8(x_180, 12, x_173); -lean_ctor_set_uint8(x_180, 13, x_174); -lean_ctor_set_uint8(x_180, 14, x_175); -lean_ctor_set_uint8(x_180, 15, x_176); -lean_ctor_set_uint8(x_180, 16, x_177); -x_181 = 2; -x_182 = lean_uint64_shift_right(x_152, x_181); -x_183 = lean_uint64_shift_left(x_182, x_181); -x_184 = l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1___closed__9; -x_185 = lean_uint64_lor(x_183, x_184); -x_186 = lean_alloc_ctor(0, 7, 11); -lean_ctor_set(x_186, 0, x_180); -lean_ctor_set(x_186, 1, x_154); -lean_ctor_set(x_186, 2, x_155); -lean_ctor_set(x_186, 3, x_156); -lean_ctor_set(x_186, 4, x_157); -lean_ctor_set(x_186, 5, x_158); -lean_ctor_set(x_186, 6, x_159); -lean_ctor_set_uint64(x_186, sizeof(void*)*7, x_185); -lean_ctor_set_uint8(x_186, sizeof(void*)*7 + 8, x_153); -lean_ctor_set_uint8(x_186, sizeof(void*)*7 + 9, x_160); -lean_ctor_set_uint8(x_186, sizeof(void*)*7 + 10, x_161); + lean_dec_ref(x_89); + x_179 = lean_box(0); +} +x_180 = 1; +if (lean_is_scalar(x_179)) { + x_181 = lean_alloc_ctor(0, 0, 17); +} else { + x_181 = x_179; +} +lean_ctor_set_uint8(x_181, 0, x_163); +lean_ctor_set_uint8(x_181, 1, x_164); +lean_ctor_set_uint8(x_181, 2, x_165); +lean_ctor_set_uint8(x_181, 3, x_166); +lean_ctor_set_uint8(x_181, 4, x_167); +lean_ctor_set_uint8(x_181, 5, x_168); +lean_ctor_set_uint8(x_181, 6, x_169); +lean_ctor_set_uint8(x_181, 7, x_170); +lean_ctor_set_uint8(x_181, 8, x_171); +lean_ctor_set_uint8(x_181, 9, x_180); +lean_ctor_set_uint8(x_181, 10, x_172); +lean_ctor_set_uint8(x_181, 11, x_173); +lean_ctor_set_uint8(x_181, 12, x_174); +lean_ctor_set_uint8(x_181, 13, x_175); +lean_ctor_set_uint8(x_181, 14, x_176); +lean_ctor_set_uint8(x_181, 15, x_177); +lean_ctor_set_uint8(x_181, 16, x_178); +x_182 = 2; +x_183 = lean_uint64_shift_right(x_153, x_182); +x_184 = lean_uint64_shift_left(x_183, x_182); +x_185 = l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1___closed__11; +x_186 = lean_uint64_lor(x_184, x_185); +x_187 = lean_alloc_ctor(0, 7, 11); +lean_ctor_set(x_187, 0, x_181); +lean_ctor_set(x_187, 1, x_155); +lean_ctor_set(x_187, 2, x_156); +lean_ctor_set(x_187, 3, x_157); +lean_ctor_set(x_187, 4, x_158); +lean_ctor_set(x_187, 5, x_159); +lean_ctor_set(x_187, 6, x_160); +lean_ctor_set_uint64(x_187, sizeof(void*)*7, x_186); +lean_ctor_set_uint8(x_187, sizeof(void*)*7 + 8, x_154); +lean_ctor_set_uint8(x_187, sizeof(void*)*7 + 9, x_161); +lean_ctor_set_uint8(x_187, sizeof(void*)*7 + 10, x_162); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_4); lean_inc(x_3); -x_187 = l_Lean_Meta_isExprDefEq(x_3, x_4, x_186, x_8, x_9, x_10, x_147); -if (lean_obj_tag(x_187) == 0) +x_188 = l_Lean_Meta_isExprDefEq(x_3, x_4, x_187, x_8, x_9, x_10, x_90); +if (lean_obj_tag(x_188) == 0) { -lean_object* x_188; lean_object* x_189; lean_object* x_190; -x_188 = lean_ctor_get(x_187, 0); -lean_inc(x_188); -x_189 = lean_ctor_get(x_187, 1); +lean_object* x_189; lean_object* x_190; lean_object* x_191; +x_189 = lean_ctor_get(x_188, 0); lean_inc(x_189); -lean_dec(x_187); -x_190 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_190, 0, x_188); -lean_ctor_set(x_190, 1, x_6); +x_190 = lean_ctor_get(x_188, 1); +lean_inc(x_190); +lean_dec(x_188); +x_191 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_191, 0, x_189); +lean_ctor_set(x_191, 1, x_152); +x_14 = x_191; x_15 = x_190; -x_16 = x_189; -goto block_74; +goto block_78; } 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_152); 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_191 = lean_ctor_get(x_187, 0); -lean_inc(x_191); -x_192 = lean_ctor_get(x_187, 1); +x_192 = lean_ctor_get(x_188, 0); lean_inc(x_192); -if (lean_is_exclusive(x_187)) { - lean_ctor_release(x_187, 0); - lean_ctor_release(x_187, 1); - x_193 = x_187; +x_193 = lean_ctor_get(x_188, 1); +lean_inc(x_193); +if (lean_is_exclusive(x_188)) { + lean_ctor_release(x_188, 0); + lean_ctor_release(x_188, 1); + x_194 = x_188; } else { - lean_dec_ref(x_187); - x_193 = lean_box(0); + lean_dec_ref(x_188); + x_194 = lean_box(0); } -if (lean_is_scalar(x_193)) { - x_194 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_194)) { + x_195 = lean_alloc_ctor(1, 2, 0); } else { - x_194 = x_193; -} -lean_ctor_set(x_194, 0, x_191); -lean_ctor_set(x_194, 1, x_192); -return x_194; + x_195 = x_194; } +lean_ctor_set(x_195, 0, x_192); +lean_ctor_set(x_195, 1, x_193); +return x_195; } } } -block_74: +block_78: { -lean_object* x_17; uint8_t x_18; -x_17 = lean_ctor_get(x_15, 0); -lean_inc(x_17); -x_18 = lean_unbox(x_17); -lean_dec(x_17); -if (x_18 == 0) +lean_object* x_16; uint8_t x_17; +x_16 = lean_ctor_get(x_14, 0); +lean_inc(x_16); +x_17 = lean_unbox(x_16); +lean_dec(x_16); +if (x_17 == 0) { -uint8_t x_19; +uint8_t x_18; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_4); lean_dec(x_3); -x_19 = !lean_is_exclusive(x_15); -if (x_19 == 0) +x_18 = !lean_is_exclusive(x_14); +if (x_18 == 0) { -lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_20 = lean_ctor_get(x_15, 0); -lean_dec(x_20); -x_21 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_21, 0, x_2); -lean_ctor_set(x_15, 0, x_21); -x_22 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_22, 0, x_15); -lean_ctor_set(x_22, 1, x_16); -return x_22; +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_14, 0); +lean_dec(x_19); +x_20 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_20, 0, x_2); +lean_ctor_set(x_14, 0, x_20); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_14); +lean_ctor_set(x_21, 1, x_15); +return x_21; } else { -lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_23 = lean_ctor_get(x_15, 1); -lean_inc(x_23); -lean_dec(x_15); -x_24 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_24, 0, x_2); +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_22 = lean_ctor_get(x_14, 1); +lean_inc(x_22); +lean_dec(x_14); +x_23 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_23, 0, x_2); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_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_23); -x_26 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_26, 0, x_25); -lean_ctor_set(x_26, 1, x_16); -return x_26; +lean_ctor_set(x_25, 1, x_15); +return x_25; } } 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_15, 1); -lean_inc(x_27); -lean_dec(x_15); -x_28 = l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1___closed__3; -x_29 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_Canon_canonElemCore___spec__4(x_28, x_27, x_7, x_8, x_9, x_10, x_16); +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; +x_26 = lean_ctor_get(x_14, 1); +lean_inc(x_26); +lean_dec(x_14); +x_27 = l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1___closed__3; +x_28 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_Canon_canonElemCore___spec__4(x_27, x_26, x_7, x_8, x_9, x_10, x_15); +x_29 = lean_ctor_get(x_28, 0); +lean_inc(x_29); x_30 = lean_ctor_get(x_29, 0); lean_inc(x_30); -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) +x_31 = lean_unbox(x_30); +lean_dec(x_30); +if (x_31 == 0) { -uint8_t x_33; +uint8_t x_32; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_4); lean_dec(x_3); -x_33 = !lean_is_exclusive(x_29); -if (x_33 == 0) +x_32 = !lean_is_exclusive(x_28); +if (x_32 == 0) { -lean_object* x_34; uint8_t x_35; -x_34 = lean_ctor_get(x_29, 0); -lean_dec(x_34); -x_35 = !lean_is_exclusive(x_30); -if (x_35 == 0) +lean_object* x_33; uint8_t x_34; +x_33 = lean_ctor_get(x_28, 0); +lean_dec(x_33); +x_34 = !lean_is_exclusive(x_29); +if (x_34 == 0) { -lean_object* x_36; lean_object* x_37; -x_36 = lean_ctor_get(x_30, 0); -lean_dec(x_36); -x_37 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_37, 0, x_2); -lean_ctor_set(x_30, 0, x_37); -return x_29; +lean_object* x_35; lean_object* x_36; +x_35 = lean_ctor_get(x_29, 0); +lean_dec(x_35); +x_36 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_36, 0, x_2); +lean_ctor_set(x_29, 0, x_36); +return x_28; } else { -lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_38 = lean_ctor_get(x_30, 1); -lean_inc(x_38); -lean_dec(x_30); -x_39 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_39, 0, x_2); -x_40 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_40, 0, x_39); -lean_ctor_set(x_40, 1, x_38); -lean_ctor_set(x_29, 0, x_40); -return x_29; +lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_37 = lean_ctor_get(x_29, 1); +lean_inc(x_37); +lean_dec(x_29); +x_38 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_38, 0, x_2); +x_39 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_37); +lean_ctor_set(x_28, 0, x_39); +return x_28; } } 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_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_40 = lean_ctor_get(x_28, 1); +lean_inc(x_40); +lean_dec(x_28); x_41 = lean_ctor_get(x_29, 1); lean_inc(x_41); -lean_dec(x_29); -x_42 = lean_ctor_get(x_30, 1); -lean_inc(x_42); -if (lean_is_exclusive(x_30)) { - lean_ctor_release(x_30, 0); - lean_ctor_release(x_30, 1); - x_43 = x_30; +if (lean_is_exclusive(x_29)) { + lean_ctor_release(x_29, 0); + lean_ctor_release(x_29, 1); + x_42 = x_29; } else { - lean_dec_ref(x_30); - x_43 = lean_box(0); + lean_dec_ref(x_29); + x_42 = lean_box(0); } -x_44 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_44, 0, x_2); -if (lean_is_scalar(x_43)) { - x_45 = lean_alloc_ctor(0, 2, 0); +x_43 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_43, 0, x_2); +if (lean_is_scalar(x_42)) { + x_44 = lean_alloc_ctor(0, 2, 0); } else { - x_45 = x_43; + x_44 = x_42; } +lean_ctor_set(x_44, 0, x_43); +lean_ctor_set(x_44, 1, x_41); +x_45 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_45, 0, x_44); -lean_ctor_set(x_45, 1, x_42); -x_46 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_46, 0, x_45); -lean_ctor_set(x_46, 1, x_41); -return x_46; +lean_ctor_set(x_45, 1, x_40); +return x_45; } } 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_object* x_56; lean_object* x_57; lean_object* x_58; uint8_t x_59; +lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; uint8_t x_63; +x_46 = lean_ctor_get(x_28, 1); +lean_inc(x_46); +lean_dec(x_28); x_47 = lean_ctor_get(x_29, 1); lean_inc(x_47); lean_dec(x_29); -x_48 = lean_ctor_get(x_30, 1); -lean_inc(x_48); -lean_dec(x_30); -x_49 = l_Lean_indentExpr(x_3); +x_48 = l_Lean_Meta_Grind_Canon_CanonElemKind_explain(x_1); +x_49 = l_Lean_stringToMessageData(x_48); +lean_dec(x_48); x_50 = l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1___closed__5; x_51 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_51, 0, x_50); @@ -2229,86 +2471,195 @@ x_52 = l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10 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_Lean_indentExpr(x_4); +x_54 = l_Lean_indentExpr(x_3); x_55 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_55, 0, x_53); lean_ctor_set(x_55, 1, x_54); -x_56 = l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1___closed__8; +x_56 = l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1___closed__9; x_57 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_57, 0, x_55); lean_ctor_set(x_57, 1, x_56); -x_58 = l_Lean_addTrace___at_Lean_Meta_Grind_Canon_canonElemCore___spec__5(x_28, x_57, x_48, x_7, x_8, x_9, x_10, x_47); +x_58 = l_Lean_indentExpr(x_4); +x_59 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_59, 0, x_57); +lean_ctor_set(x_59, 1, x_58); +x_60 = l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1___closed__10; +x_61 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_61, 0, x_59); +lean_ctor_set(x_61, 1, x_60); +x_62 = l_Lean_addTrace___at_Lean_Meta_Grind_Canon_canonElemCore___spec__5(x_27, x_61, x_47, x_7, x_8, x_9, x_10, x_46); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -x_59 = !lean_is_exclusive(x_58); -if (x_59 == 0) +x_63 = !lean_is_exclusive(x_62); +if (x_63 == 0) { -lean_object* x_60; uint8_t x_61; -x_60 = lean_ctor_get(x_58, 0); -x_61 = !lean_is_exclusive(x_60); -if (x_61 == 0) +lean_object* x_64; uint8_t x_65; +x_64 = lean_ctor_get(x_62, 0); +x_65 = !lean_is_exclusive(x_64); +if (x_65 == 0) { -lean_object* x_62; lean_object* x_63; -x_62 = lean_ctor_get(x_60, 0); -lean_dec(x_62); -x_63 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_63, 0, x_2); -lean_ctor_set(x_60, 0, x_63); -return x_58; +lean_object* x_66; lean_object* x_67; +x_66 = lean_ctor_get(x_64, 0); +lean_dec(x_66); +x_67 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_67, 0, x_2); +lean_ctor_set(x_64, 0, x_67); +return x_62; } else { -lean_object* x_64; lean_object* x_65; lean_object* x_66; -x_64 = lean_ctor_get(x_60, 1); -lean_inc(x_64); -lean_dec(x_60); -x_65 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_65, 0, x_2); -x_66 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_66, 0, x_65); -lean_ctor_set(x_66, 1, x_64); -lean_ctor_set(x_58, 0, x_66); -return x_58; +lean_object* x_68; lean_object* x_69; lean_object* x_70; +x_68 = lean_ctor_get(x_64, 1); +lean_inc(x_68); +lean_dec(x_64); +x_69 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_69, 0, x_2); +x_70 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_70, 0, x_69); +lean_ctor_set(x_70, 1, x_68); +lean_ctor_set(x_62, 0, x_70); +return x_62; } } else { -lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; -x_67 = lean_ctor_get(x_58, 0); -x_68 = lean_ctor_get(x_58, 1); -lean_inc(x_68); -lean_inc(x_67); -lean_dec(x_58); -x_69 = lean_ctor_get(x_67, 1); -lean_inc(x_69); -if (lean_is_exclusive(x_67)) { - lean_ctor_release(x_67, 0); - lean_ctor_release(x_67, 1); - x_70 = x_67; +lean_object* x_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_71 = lean_ctor_get(x_62, 0); +x_72 = lean_ctor_get(x_62, 1); +lean_inc(x_72); +lean_inc(x_71); +lean_dec(x_62); +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; } else { - lean_dec_ref(x_67); - x_70 = lean_box(0); + lean_dec_ref(x_71); + x_74 = lean_box(0); } -x_71 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_71, 0, x_2); -if (lean_is_scalar(x_70)) { - x_72 = lean_alloc_ctor(0, 2, 0); +x_75 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_75, 0, x_2); +if (lean_is_scalar(x_74)) { + x_76 = lean_alloc_ctor(0, 2, 0); } else { - x_72 = x_70; + x_76 = x_74; +} +lean_ctor_set(x_76, 0, x_75); +lean_ctor_set(x_76, 1, x_73); +x_77 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_77, 0, x_76); +lean_ctor_set(x_77, 1, x_72); +return x_77; +} +} +} +} +} +else +{ +lean_object* x_196; lean_object* x_197; lean_object* x_198; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_4); +lean_dec(x_3); +x_196 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_196, 0, x_2); +x_197 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_197, 0, x_196); +lean_ctor_set(x_197, 1, x_6); +x_198 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_198, 0, x_197); +lean_ctor_set(x_198, 1, x_11); +return x_198; +} +} +} +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_9 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_9, 0, x_1); +x_10 = lean_box(0); +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, 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_3); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_8); +return x_14; +} +} +static lean_object* _init_l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("debug", 5, 5); +return x_1; +} +} +static lean_object* _init_l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("canon", 5, 5); +return x_1; +} +} +static lean_object* _init_l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1___closed__1; +x_2 = l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___closed__1; +x_3 = l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___closed__2; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("found ", 6, 6); +return x_1; } -lean_ctor_set(x_72, 0, x_71); -lean_ctor_set(x_72, 1, x_69); -x_73 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_73, 0, x_72); -lean_ctor_set(x_73, 1, x_68); -return x_73; } +static lean_object* _init_l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___closed__4; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } } +static lean_object* _init_l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(" ===> ", 6, 6); +return x_1; } } +static lean_object* _init_l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___closed__6; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } } LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { @@ -2333,293 +2684,431 @@ return x_17; } else { -uint8_t x_18; +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_36; lean_dec(x_8); -x_18 = !lean_is_exclusive(x_7); -if (x_18 == 0) -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_19 = lean_ctor_get(x_7, 0); -x_20 = lean_ctor_get(x_7, 1); +x_18 = lean_ctor_get(x_7, 0); +lean_inc(x_18); +x_19 = lean_ctor_get(x_7, 1); +lean_inc(x_19); +if (lean_is_exclusive(x_7)) { + lean_ctor_release(x_7, 0); + lean_ctor_release(x_7, 1); + x_20 = x_7; +} else { + lean_dec_ref(x_7); + x_20 = lean_box(0); +} lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); -lean_inc(x_19); +lean_inc(x_18); lean_inc(x_1); -x_21 = l_Lean_Meta_isExprDefEq(x_1, x_19, x_11, x_12, x_13, x_14, x_15); -if (lean_obj_tag(x_21) == 0) +x_36 = l_Lean_Meta_isExprDefEq(x_1, x_18, x_11, x_12, x_13, x_14, x_15); +if (lean_obj_tag(x_36) == 0) { -lean_object* x_22; uint8_t x_23; -x_22 = lean_ctor_get(x_21, 0); -lean_inc(x_22); -x_23 = lean_unbox(x_22); -lean_dec(x_22); -if (x_23 == 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_24; lean_object* x_25; lean_object* x_26; -lean_free_object(x_7); -x_24 = lean_ctor_get(x_21, 1); -lean_inc(x_24); -lean_dec(x_21); -x_25 = lean_box(0); +lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_39 = lean_ctor_get(x_36, 1); +lean_inc(x_39); +lean_dec(x_36); +x_40 = lean_box(0); lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); lean_inc(x_1); lean_inc(x_4); -x_26 = l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1(x_2, x_4, x_1, x_19, x_25, x_10, x_11, x_12, x_13, x_14, x_24); -if (lean_obj_tag(x_26) == 0) +x_41 = l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1(x_2, x_4, x_1, x_18, x_40, x_10, x_11, x_12, x_13, x_14, x_39); +if (lean_obj_tag(x_41) == 0) { -lean_object* x_27; lean_object* x_28; -x_27 = lean_ctor_get(x_26, 0); -lean_inc(x_27); -x_28 = lean_ctor_get(x_27, 0); -lean_inc(x_28); -if (lean_obj_tag(x_28) == 0) +lean_object* x_42; lean_object* x_43; +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_21 = x_42; +x_22 = x_43; +goto block_35; +} +else { -uint8_t x_29; +uint8_t x_44; lean_dec(x_20); +lean_dec(x_19); lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_4); lean_dec(x_1); -x_29 = !lean_is_exclusive(x_26); -if (x_29 == 0) +x_44 = !lean_is_exclusive(x_41); +if (x_44 == 0) { -lean_object* x_30; uint8_t x_31; -x_30 = lean_ctor_get(x_26, 0); -lean_dec(x_30); -x_31 = !lean_is_exclusive(x_27); -if (x_31 == 0) +return x_41; +} +else { -lean_object* x_32; lean_object* x_33; -x_32 = lean_ctor_get(x_27, 0); -lean_dec(x_32); -x_33 = lean_ctor_get(x_28, 0); -lean_inc(x_33); -lean_dec(x_28); -lean_ctor_set(x_27, 0, x_33); -return x_26; +lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_45 = lean_ctor_get(x_41, 0); +x_46 = lean_ctor_get(x_41, 1); +lean_inc(x_46); +lean_inc(x_45); +lean_dec(x_41); +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_34; lean_object* x_35; lean_object* x_36; -x_34 = lean_ctor_get(x_27, 1); -lean_inc(x_34); -lean_dec(x_27); -x_35 = lean_ctor_get(x_28, 0); -lean_inc(x_35); -lean_dec(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); -lean_ctor_set(x_26, 0, x_36); -return x_26; } } else { -lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; -x_37 = lean_ctor_get(x_26, 1); -lean_inc(x_37); -lean_dec(x_26); -x_38 = lean_ctor_get(x_27, 1); -lean_inc(x_38); -if (lean_is_exclusive(x_27)) { - lean_ctor_release(x_27, 0); - lean_ctor_release(x_27, 1); - x_39 = x_27; -} else { - lean_dec_ref(x_27); - x_39 = lean_box(0); -} -x_40 = lean_ctor_get(x_28, 0); -lean_inc(x_40); -lean_dec(x_28); -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); -x_42 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_42, 0, x_41); -lean_ctor_set(x_42, 1, x_37); -return x_42; -} -} -else +lean_object* x_48; uint8_t x_49; +x_48 = lean_ctor_get(x_36, 1); +lean_inc(x_48); +lean_dec(x_36); +x_49 = !lean_is_exclusive(x_10); +if (x_49 == 0) { -lean_object* x_43; lean_object* x_44; lean_object* x_45; -x_43 = lean_ctor_get(x_26, 1); -lean_inc(x_43); -lean_dec(x_26); -x_44 = lean_ctor_get(x_27, 1); -lean_inc(x_44); -lean_dec(x_27); -x_45 = lean_ctor_get(x_28, 0); -lean_inc(x_45); -lean_dec(x_28); -x_7 = x_20; -x_8 = x_45; -x_9 = lean_box(0); -x_10 = x_44; -x_15 = x_43; -goto _start; -} +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_50 = lean_ctor_get(x_10, 1); +lean_inc(x_18); +lean_inc(x_1); +x_51 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_50, x_1, x_18); +lean_ctor_set(x_10, 1, x_51); +x_52 = l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___closed__3; +x_53 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_Canon_canonElemCore___spec__4(x_52, x_10, x_11, x_12, x_13, x_14, x_48); +x_54 = lean_ctor_get(x_53, 0); +lean_inc(x_54); +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_object* x_61; lean_object* x_62; +x_57 = lean_ctor_get(x_53, 1); +lean_inc(x_57); +lean_dec(x_53); +x_58 = lean_ctor_get(x_54, 1); +lean_inc(x_58); +lean_dec(x_54); +x_59 = lean_box(0); +x_60 = l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__2(x_18, x_59, x_58, x_11, x_12, x_13, x_14, x_57); +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_21 = x_61; +x_22 = x_62; +goto block_35; } else { -uint8_t x_47; -lean_dec(x_20); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_4); -lean_dec(x_1); -x_47 = !lean_is_exclusive(x_26); -if (x_47 == 0) +uint8_t x_63; +x_63 = !lean_is_exclusive(x_53); +if (x_63 == 0) { -return x_26; +lean_object* x_64; lean_object* x_65; uint8_t x_66; +x_64 = lean_ctor_get(x_53, 1); +x_65 = lean_ctor_get(x_53, 0); +lean_dec(x_65); +x_66 = !lean_is_exclusive(x_54); +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; 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_54, 1); +x_68 = lean_ctor_get(x_54, 0); +lean_dec(x_68); +lean_inc(x_1); +x_69 = l_Lean_MessageData_ofExpr(x_1); +x_70 = l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___closed__5; +lean_ctor_set_tag(x_54, 7); +lean_ctor_set(x_54, 1, x_69); +lean_ctor_set(x_54, 0, x_70); +x_71 = l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___closed__7; +lean_ctor_set_tag(x_53, 7); +lean_ctor_set(x_53, 1, x_71); +lean_inc(x_18); +x_72 = l_Lean_MessageData_ofExpr(x_18); +x_73 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_73, 0, x_53); +lean_ctor_set(x_73, 1, x_72); +x_74 = l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1___closed__10; +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_addTrace___at_Lean_Meta_Grind_Canon_canonElemCore___spec__5(x_52, x_75, x_67, x_11, x_12, x_13, x_14, x_64); +x_77 = lean_ctor_get(x_76, 0); +lean_inc(x_77); +x_78 = lean_ctor_get(x_76, 1); +lean_inc(x_78); +lean_dec(x_76); +x_79 = lean_ctor_get(x_77, 0); +lean_inc(x_79); +x_80 = lean_ctor_get(x_77, 1); +lean_inc(x_80); +lean_dec(x_77); +x_81 = l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__2(x_18, x_79, x_80, x_11, x_12, x_13, x_14, x_78); +lean_dec(x_79); +x_82 = lean_ctor_get(x_81, 0); +lean_inc(x_82); +x_83 = lean_ctor_get(x_81, 1); +lean_inc(x_83); +lean_dec(x_81); +x_21 = x_82; +x_22 = x_83; +goto block_35; } else { -lean_object* x_48; lean_object* x_49; lean_object* x_50; -x_48 = lean_ctor_get(x_26, 0); -x_49 = lean_ctor_get(x_26, 1); -lean_inc(x_49); -lean_inc(x_48); -lean_dec(x_26); -x_50 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_50, 0, x_48); -lean_ctor_set(x_50, 1, x_49); -return x_50; -} +lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; +x_84 = lean_ctor_get(x_54, 1); +lean_inc(x_84); +lean_dec(x_54); +lean_inc(x_1); +x_85 = l_Lean_MessageData_ofExpr(x_1); +x_86 = l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___closed__5; +x_87 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_87, 0, x_86); +lean_ctor_set(x_87, 1, x_85); +x_88 = l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___closed__7; +lean_ctor_set_tag(x_53, 7); +lean_ctor_set(x_53, 1, x_88); +lean_ctor_set(x_53, 0, x_87); +lean_inc(x_18); +x_89 = l_Lean_MessageData_ofExpr(x_18); +x_90 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_90, 0, x_53); +lean_ctor_set(x_90, 1, x_89); +x_91 = l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1___closed__10; +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_Lean_addTrace___at_Lean_Meta_Grind_Canon_canonElemCore___spec__5(x_52, x_92, x_84, x_11, x_12, x_13, x_14, x_64); +x_94 = lean_ctor_get(x_93, 0); +lean_inc(x_94); +x_95 = lean_ctor_get(x_93, 1); +lean_inc(x_95); +lean_dec(x_93); +x_96 = lean_ctor_get(x_94, 0); +lean_inc(x_96); +x_97 = lean_ctor_get(x_94, 1); +lean_inc(x_97); +lean_dec(x_94); +x_98 = l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__2(x_18, x_96, x_97, x_11, x_12, x_13, x_14, x_95); +lean_dec(x_96); +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_21 = x_99; +x_22 = x_100; +goto block_35; } } else { -uint8_t x_51; -lean_dec(x_20); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_4); -x_51 = !lean_is_exclusive(x_21); -if (x_51 == 0) -{ -lean_object* x_52; uint8_t x_53; -x_52 = lean_ctor_get(x_21, 0); -lean_dec(x_52); -x_53 = !lean_is_exclusive(x_10); -if (x_53 == 0) -{ -lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; -x_54 = lean_ctor_get(x_10, 1); -lean_inc(x_19); -x_55 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_54, x_1, x_19); -lean_ctor_set(x_10, 1, x_55); -x_56 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_56, 0, x_19); -x_57 = lean_box(0); -lean_ctor_set_tag(x_7, 0); -lean_ctor_set(x_7, 1, x_57); -lean_ctor_set(x_7, 0, 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_10); -lean_ctor_set(x_21, 0, x_58); -return x_21; +lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; +x_101 = lean_ctor_get(x_53, 1); +lean_inc(x_101); +lean_dec(x_53); +x_102 = lean_ctor_get(x_54, 1); +lean_inc(x_102); +if (lean_is_exclusive(x_54)) { + lean_ctor_release(x_54, 0); + lean_ctor_release(x_54, 1); + x_103 = x_54; +} else { + lean_dec_ref(x_54); + x_103 = lean_box(0); +} +lean_inc(x_1); +x_104 = l_Lean_MessageData_ofExpr(x_1); +x_105 = l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___closed__5; +if (lean_is_scalar(x_103)) { + x_106 = lean_alloc_ctor(7, 2, 0); +} else { + x_106 = x_103; + lean_ctor_set_tag(x_106, 7); +} +lean_ctor_set(x_106, 0, x_105); +lean_ctor_set(x_106, 1, x_104); +x_107 = l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___closed__7; +x_108 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_108, 0, x_106); +lean_ctor_set(x_108, 1, x_107); +lean_inc(x_18); +x_109 = l_Lean_MessageData_ofExpr(x_18); +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_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1___closed__10; +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_Meta_Grind_Canon_canonElemCore___spec__5(x_52, x_112, x_102, x_11, x_12, x_13, x_14, x_101); +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); +x_117 = lean_ctor_get(x_114, 1); +lean_inc(x_117); +lean_dec(x_114); +x_118 = l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__2(x_18, x_116, x_117, x_11, x_12, x_13, x_14, x_115); +lean_dec(x_116); +x_119 = lean_ctor_get(x_118, 0); +lean_inc(x_119); +x_120 = lean_ctor_get(x_118, 1); +lean_inc(x_120); +lean_dec(x_118); +x_21 = x_119; +x_22 = x_120; +goto block_35; +} +} } else { -lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; -x_59 = lean_ctor_get(x_10, 0); -x_60 = lean_ctor_get(x_10, 1); -x_61 = lean_ctor_get(x_10, 2); -lean_inc(x_61); -lean_inc(x_60); -lean_inc(x_59); +lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; uint8_t x_130; +x_121 = lean_ctor_get(x_10, 0); +x_122 = lean_ctor_get(x_10, 1); +x_123 = lean_ctor_get(x_10, 2); +lean_inc(x_123); +lean_inc(x_122); +lean_inc(x_121); lean_dec(x_10); -lean_inc(x_19); -x_62 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_60, x_1, x_19); -x_63 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_63, 0, x_59); -lean_ctor_set(x_63, 1, x_62); -lean_ctor_set(x_63, 2, x_61); -x_64 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_64, 0, x_19); -x_65 = lean_box(0); -lean_ctor_set_tag(x_7, 0); -lean_ctor_set(x_7, 1, x_65); -lean_ctor_set(x_7, 0, x_64); -x_66 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_66, 0, x_7); -lean_ctor_set(x_66, 1, x_63); -lean_ctor_set(x_21, 0, x_66); -return x_21; -} +lean_inc(x_18); +lean_inc(x_1); +x_124 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_122, x_1, x_18); +x_125 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_125, 0, x_121); +lean_ctor_set(x_125, 1, x_124); +lean_ctor_set(x_125, 2, x_123); +x_126 = l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___closed__3; +x_127 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_Canon_canonElemCore___spec__4(x_126, x_125, x_11, x_12, x_13, x_14, x_48); +x_128 = lean_ctor_get(x_127, 0); +lean_inc(x_128); +x_129 = lean_ctor_get(x_128, 0); +lean_inc(x_129); +x_130 = lean_unbox(x_129); +lean_dec(x_129); +if (x_130 == 0) +{ +lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; +x_131 = lean_ctor_get(x_127, 1); +lean_inc(x_131); +lean_dec(x_127); +x_132 = lean_ctor_get(x_128, 1); +lean_inc(x_132); +lean_dec(x_128); +x_133 = lean_box(0); +x_134 = l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__2(x_18, x_133, x_132, x_11, x_12, x_13, x_14, x_131); +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_21 = x_135; +x_22 = x_136; +goto block_35; } else { -lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; -x_67 = lean_ctor_get(x_21, 1); -lean_inc(x_67); -lean_dec(x_21); -x_68 = lean_ctor_get(x_10, 0); -lean_inc(x_68); -x_69 = lean_ctor_get(x_10, 1); -lean_inc(x_69); -x_70 = lean_ctor_get(x_10, 2); -lean_inc(x_70); -if (lean_is_exclusive(x_10)) { - lean_ctor_release(x_10, 0); - lean_ctor_release(x_10, 1); - lean_ctor_release(x_10, 2); - x_71 = x_10; +lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; +x_137 = lean_ctor_get(x_127, 1); +lean_inc(x_137); +if (lean_is_exclusive(x_127)) { + lean_ctor_release(x_127, 0); + lean_ctor_release(x_127, 1); + x_138 = x_127; } else { - lean_dec_ref(x_10); - x_71 = lean_box(0); + lean_dec_ref(x_127); + x_138 = lean_box(0); +} +x_139 = lean_ctor_get(x_128, 1); +lean_inc(x_139); +if (lean_is_exclusive(x_128)) { + lean_ctor_release(x_128, 0); + lean_ctor_release(x_128, 1); + x_140 = x_128; +} else { + lean_dec_ref(x_128); + x_140 = lean_box(0); } -lean_inc(x_19); -x_72 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_69, x_1, x_19); -if (lean_is_scalar(x_71)) { - x_73 = lean_alloc_ctor(0, 3, 0); +lean_inc(x_1); +x_141 = l_Lean_MessageData_ofExpr(x_1); +x_142 = l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___closed__5; +if (lean_is_scalar(x_140)) { + x_143 = lean_alloc_ctor(7, 2, 0); } else { - x_73 = x_71; + x_143 = x_140; + lean_ctor_set_tag(x_143, 7); +} +lean_ctor_set(x_143, 0, x_142); +lean_ctor_set(x_143, 1, x_141); +x_144 = l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___closed__7; +if (lean_is_scalar(x_138)) { + x_145 = lean_alloc_ctor(7, 2, 0); +} else { + x_145 = x_138; + lean_ctor_set_tag(x_145, 7); +} +lean_ctor_set(x_145, 0, x_143); +lean_ctor_set(x_145, 1, x_144); +lean_inc(x_18); +x_146 = l_Lean_MessageData_ofExpr(x_18); +x_147 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_147, 0, x_145); +lean_ctor_set(x_147, 1, x_146); +x_148 = l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1___closed__10; +x_149 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_149, 0, x_147); +lean_ctor_set(x_149, 1, x_148); +x_150 = l_Lean_addTrace___at_Lean_Meta_Grind_Canon_canonElemCore___spec__5(x_126, x_149, x_139, x_11, x_12, x_13, x_14, x_137); +x_151 = lean_ctor_get(x_150, 0); +lean_inc(x_151); +x_152 = lean_ctor_get(x_150, 1); +lean_inc(x_152); +lean_dec(x_150); +x_153 = lean_ctor_get(x_151, 0); +lean_inc(x_153); +x_154 = lean_ctor_get(x_151, 1); +lean_inc(x_154); +lean_dec(x_151); +x_155 = l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__2(x_18, x_153, x_154, x_11, x_12, x_13, x_14, x_152); +lean_dec(x_153); +x_156 = lean_ctor_get(x_155, 0); +lean_inc(x_156); +x_157 = lean_ctor_get(x_155, 1); +lean_inc(x_157); +lean_dec(x_155); +x_21 = x_156; +x_22 = x_157; +goto block_35; } -lean_ctor_set(x_73, 0, x_68); -lean_ctor_set(x_73, 1, x_72); -lean_ctor_set(x_73, 2, x_70); -x_74 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_74, 0, x_19); -x_75 = lean_box(0); -lean_ctor_set_tag(x_7, 0); -lean_ctor_set(x_7, 1, x_75); -lean_ctor_set(x_7, 0, x_74); -x_76 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_76, 0, x_7); -lean_ctor_set(x_76, 1, x_73); -x_77 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_77, 0, x_76); -lean_ctor_set(x_77, 1, x_67); -return x_77; } } } else { -uint8_t x_78; -lean_free_object(x_7); +uint8_t x_158; lean_dec(x_20); lean_dec(x_19); +lean_dec(x_18); lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); @@ -2627,267 +3116,101 @@ lean_dec(x_11); lean_dec(x_10); lean_dec(x_4); lean_dec(x_1); -x_78 = !lean_is_exclusive(x_21); -if (x_78 == 0) +x_158 = !lean_is_exclusive(x_36); +if (x_158 == 0) { -return x_21; +return x_36; } else { -lean_object* x_79; lean_object* x_80; lean_object* x_81; -x_79 = lean_ctor_get(x_21, 0); -x_80 = lean_ctor_get(x_21, 1); -lean_inc(x_80); -lean_inc(x_79); -lean_dec(x_21); -x_81 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_81, 0, x_79); -lean_ctor_set(x_81, 1, x_80); -return x_81; -} +lean_object* x_159; lean_object* x_160; lean_object* x_161; +x_159 = lean_ctor_get(x_36, 0); +x_160 = lean_ctor_get(x_36, 1); +lean_inc(x_160); +lean_inc(x_159); +lean_dec(x_36); +x_161 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_161, 0, x_159); +lean_ctor_set(x_161, 1, x_160); +return x_161; } } -else -{ -lean_object* x_82; lean_object* x_83; lean_object* x_84; -x_82 = lean_ctor_get(x_7, 0); -x_83 = lean_ctor_get(x_7, 1); -lean_inc(x_83); -lean_inc(x_82); -lean_dec(x_7); -lean_inc(x_14); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_82); -lean_inc(x_1); -x_84 = l_Lean_Meta_isExprDefEq(x_1, x_82, x_11, x_12, x_13, x_14, x_15); -if (lean_obj_tag(x_84) == 0) -{ -lean_object* x_85; uint8_t x_86; -x_85 = lean_ctor_get(x_84, 0); -lean_inc(x_85); -x_86 = lean_unbox(x_85); -lean_dec(x_85); -if (x_86 == 0) -{ -lean_object* x_87; lean_object* x_88; lean_object* x_89; -x_87 = lean_ctor_get(x_84, 1); -lean_inc(x_87); -lean_dec(x_84); -x_88 = lean_box(0); -lean_inc(x_14); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_1); -lean_inc(x_4); -x_89 = l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1(x_2, x_4, x_1, x_82, x_88, x_10, x_11, x_12, x_13, x_14, x_87); -if (lean_obj_tag(x_89) == 0) +block_35: { -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_90, 0); -lean_inc(x_91); -if (lean_obj_tag(x_91) == 0) +lean_object* x_23; +x_23 = lean_ctor_get(x_21, 0); +lean_inc(x_23); +if (lean_obj_tag(x_23) == 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_dec(x_83); +uint8_t x_24; +lean_dec(x_19); lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_4); lean_dec(x_1); -x_92 = lean_ctor_get(x_89, 1); -lean_inc(x_92); -if (lean_is_exclusive(x_89)) { - lean_ctor_release(x_89, 0); - lean_ctor_release(x_89, 1); - x_93 = x_89; +x_24 = !lean_is_exclusive(x_21); +if (x_24 == 0) +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_25 = lean_ctor_get(x_21, 0); +lean_dec(x_25); +x_26 = lean_ctor_get(x_23, 0); +lean_inc(x_26); +lean_dec(x_23); +lean_ctor_set(x_21, 0, x_26); +if (lean_is_scalar(x_20)) { + x_27 = lean_alloc_ctor(0, 2, 0); } else { - lean_dec_ref(x_89); - x_93 = lean_box(0); + x_27 = x_20; + lean_ctor_set_tag(x_27, 0); } -x_94 = lean_ctor_get(x_90, 1); -lean_inc(x_94); -if (lean_is_exclusive(x_90)) { - lean_ctor_release(x_90, 0); - lean_ctor_release(x_90, 1); - x_95 = x_90; -} else { - lean_dec_ref(x_90); - x_95 = lean_box(0); +lean_ctor_set(x_27, 0, x_21); +lean_ctor_set(x_27, 1, x_22); +return x_27; } -x_96 = lean_ctor_get(x_91, 0); -lean_inc(x_96); -lean_dec(x_91); -if (lean_is_scalar(x_95)) { - x_97 = lean_alloc_ctor(0, 2, 0); +else +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_28 = lean_ctor_get(x_21, 1); +lean_inc(x_28); +lean_dec(x_21); +x_29 = lean_ctor_get(x_23, 0); +lean_inc(x_29); +lean_dec(x_23); +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_28); +if (lean_is_scalar(x_20)) { + x_31 = lean_alloc_ctor(0, 2, 0); } else { - x_97 = x_95; + x_31 = x_20; + lean_ctor_set_tag(x_31, 0); } -lean_ctor_set(x_97, 0, x_96); -lean_ctor_set(x_97, 1, x_94); -if (lean_is_scalar(x_93)) { - x_98 = lean_alloc_ctor(0, 2, 0); -} else { - x_98 = x_93; +lean_ctor_set(x_31, 0, x_30); +lean_ctor_set(x_31, 1, x_22); +return x_31; } -lean_ctor_set(x_98, 0, x_97); -lean_ctor_set(x_98, 1, x_92); -return x_98; } else { -lean_object* x_99; lean_object* x_100; lean_object* x_101; -x_99 = lean_ctor_get(x_89, 1); -lean_inc(x_99); -lean_dec(x_89); -x_100 = lean_ctor_get(x_90, 1); -lean_inc(x_100); -lean_dec(x_90); -x_101 = lean_ctor_get(x_91, 0); -lean_inc(x_101); -lean_dec(x_91); -x_7 = x_83; -x_8 = x_101; +lean_object* x_32; lean_object* x_33; +lean_dec(x_20); +x_32 = lean_ctor_get(x_21, 1); +lean_inc(x_32); +lean_dec(x_21); +x_33 = lean_ctor_get(x_23, 0); +lean_inc(x_33); +lean_dec(x_23); +x_7 = x_19; +x_8 = x_33; x_9 = lean_box(0); -x_10 = x_100; -x_15 = x_99; +x_10 = x_32; +x_15 = x_22; goto _start; } } -else -{ -lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; -lean_dec(x_83); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_4); -lean_dec(x_1); -x_103 = lean_ctor_get(x_89, 0); -lean_inc(x_103); -x_104 = lean_ctor_get(x_89, 1); -lean_inc(x_104); -if (lean_is_exclusive(x_89)) { - lean_ctor_release(x_89, 0); - lean_ctor_release(x_89, 1); - x_105 = x_89; -} else { - lean_dec_ref(x_89); - x_105 = lean_box(0); -} -if (lean_is_scalar(x_105)) { - x_106 = lean_alloc_ctor(1, 2, 0); -} else { - x_106 = x_105; -} -lean_ctor_set(x_106, 0, x_103); -lean_ctor_set(x_106, 1, x_104); -return x_106; -} -} -else -{ -lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; -lean_dec(x_83); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_4); -x_107 = lean_ctor_get(x_84, 1); -lean_inc(x_107); -if (lean_is_exclusive(x_84)) { - lean_ctor_release(x_84, 0); - lean_ctor_release(x_84, 1); - x_108 = x_84; -} else { - lean_dec_ref(x_84); - x_108 = lean_box(0); -} -x_109 = lean_ctor_get(x_10, 0); -lean_inc(x_109); -x_110 = lean_ctor_get(x_10, 1); -lean_inc(x_110); -x_111 = lean_ctor_get(x_10, 2); -lean_inc(x_111); -if (lean_is_exclusive(x_10)) { - lean_ctor_release(x_10, 0); - lean_ctor_release(x_10, 1); - lean_ctor_release(x_10, 2); - x_112 = x_10; -} else { - lean_dec_ref(x_10); - x_112 = lean_box(0); -} -lean_inc(x_82); -x_113 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_110, x_1, x_82); -if (lean_is_scalar(x_112)) { - x_114 = lean_alloc_ctor(0, 3, 0); -} else { - x_114 = x_112; -} -lean_ctor_set(x_114, 0, x_109); -lean_ctor_set(x_114, 1, x_113); -lean_ctor_set(x_114, 2, x_111); -x_115 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_115, 0, x_82); -x_116 = lean_box(0); -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_117); -lean_ctor_set(x_118, 1, x_114); -if (lean_is_scalar(x_108)) { - x_119 = lean_alloc_ctor(0, 2, 0); -} else { - x_119 = x_108; -} -lean_ctor_set(x_119, 0, x_118); -lean_ctor_set(x_119, 1, x_107); -return x_119; -} -} -else -{ -lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; -lean_dec(x_83); -lean_dec(x_82); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_4); -lean_dec(x_1); -x_120 = lean_ctor_get(x_84, 0); -lean_inc(x_120); -x_121 = lean_ctor_get(x_84, 1); -lean_inc(x_121); -if (lean_is_exclusive(x_84)) { - lean_ctor_release(x_84, 0); - lean_ctor_release(x_84, 1); - x_122 = x_84; -} else { - lean_dec_ref(x_84); - x_122 = lean_box(0); -} -if (lean_is_scalar(x_122)) { - x_123 = lean_alloc_ctor(1, 2, 0); -} else { - x_123 = x_122; -} -lean_ctor_set(x_123, 0, x_120); -lean_ctor_set(x_123, 1, x_121); -return x_123; -} -} } } } @@ -3787,6 +4110,263 @@ return x_27; static lean_object* _init_l_Lean_Meta_Grind_Canon_canonElemCore___lambda__2___closed__1() { _start: { +lean_object* x_1; +x_1 = lean_mk_string_unchecked("(", 1, 1); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_Canon_canonElemCore___lambda__2___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_Canon_canonElemCore___lambda__2___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Grind_Canon_canonElemCore___lambda__2___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(", ", 2, 2); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_Canon_canonElemCore___lambda__2___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_Canon_canonElemCore___lambda__2___closed__3; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Grind_Canon_canonElemCore___lambda__2___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(") ↦ ", 6, 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_Canon_canonElemCore___lambda__2___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_Canon_canonElemCore___lambda__2___closed__5; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_canonElemCore___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_13 = l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___closed__3; +x_14 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_Canon_canonElemCore___spec__4(x_13, x_7, x_8, x_9, x_10, x_11, x_12); +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +x_16 = lean_ctor_get(x_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; lean_object* x_20; lean_object* x_21; +lean_dec(x_5); +lean_dec(x_4); +x_18 = lean_ctor_get(x_14, 1); +lean_inc(x_18); +lean_dec(x_14); +x_19 = lean_ctor_get(x_15, 1); +lean_inc(x_19); +lean_dec(x_15); +x_20 = lean_box(0); +x_21 = l_Lean_Meta_Grind_Canon_canonElemCore___lambda__1(x_1, x_2, x_3, x_20, x_19, x_8, x_9, x_10, x_11, x_18); +return x_21; +} +else +{ +uint8_t x_22; +x_22 = !lean_is_exclusive(x_14); +if (x_22 == 0) +{ +lean_object* x_23; lean_object* x_24; uint8_t x_25; +x_23 = lean_ctor_get(x_14, 1); +x_24 = lean_ctor_get(x_14, 0); +lean_dec(x_24); +x_25 = !lean_is_exclusive(x_15); +if (x_25 == 0) +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_26 = lean_ctor_get(x_15, 1); +x_27 = lean_ctor_get(x_15, 0); +lean_dec(x_27); +x_28 = l_Lean_MessageData_ofExpr(x_4); +x_29 = l_Lean_Meta_Grind_Canon_canonElemCore___lambda__2___closed__2; +lean_ctor_set_tag(x_15, 7); +lean_ctor_set(x_15, 1, x_28); +lean_ctor_set(x_15, 0, x_29); +x_30 = l_Lean_Meta_Grind_Canon_canonElemCore___lambda__2___closed__4; +lean_ctor_set_tag(x_14, 7); +lean_ctor_set(x_14, 1, x_30); +x_31 = l___private_Init_Data_Repr_0__Nat_reprFast(x_5); +x_32 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_32, 0, x_31); +x_33 = l_Lean_MessageData_ofFormat(x_32); +x_34 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_34, 0, x_14); +lean_ctor_set(x_34, 1, x_33); +x_35 = l_Lean_Meta_Grind_Canon_canonElemCore___lambda__2___closed__6; +x_36 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_36, 0, x_34); +lean_ctor_set(x_36, 1, x_35); +lean_inc(x_1); +x_37 = l_Lean_MessageData_ofExpr(x_1); +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_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1___closed__10; +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_addTrace___at_Lean_Meta_Grind_Canon_canonElemCore___spec__5(x_13, x_40, x_26, x_8, x_9, x_10, x_11, x_23); +x_42 = lean_ctor_get(x_41, 0); +lean_inc(x_42); +x_43 = lean_ctor_get(x_41, 1); +lean_inc(x_43); +lean_dec(x_41); +x_44 = lean_ctor_get(x_42, 0); +lean_inc(x_44); +x_45 = lean_ctor_get(x_42, 1); +lean_inc(x_45); +lean_dec(x_42); +x_46 = l_Lean_Meta_Grind_Canon_canonElemCore___lambda__1(x_1, x_2, x_3, x_44, x_45, x_8, x_9, x_10, x_11, x_43); +lean_dec(x_44); +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_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; 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_47 = lean_ctor_get(x_15, 1); +lean_inc(x_47); +lean_dec(x_15); +x_48 = l_Lean_MessageData_ofExpr(x_4); +x_49 = l_Lean_Meta_Grind_Canon_canonElemCore___lambda__2___closed__2; +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_Meta_Grind_Canon_canonElemCore___lambda__2___closed__4; +lean_ctor_set_tag(x_14, 7); +lean_ctor_set(x_14, 1, x_51); +lean_ctor_set(x_14, 0, x_50); +x_52 = l___private_Init_Data_Repr_0__Nat_reprFast(x_5); +x_53 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_53, 0, x_52); +x_54 = l_Lean_MessageData_ofFormat(x_53); +x_55 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_55, 0, x_14); +lean_ctor_set(x_55, 1, x_54); +x_56 = l_Lean_Meta_Grind_Canon_canonElemCore___lambda__2___closed__6; +x_57 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_57, 0, x_55); +lean_ctor_set(x_57, 1, x_56); +lean_inc(x_1); +x_58 = l_Lean_MessageData_ofExpr(x_1); +x_59 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_59, 0, x_57); +lean_ctor_set(x_59, 1, x_58); +x_60 = l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1___closed__10; +x_61 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_61, 0, x_59); +lean_ctor_set(x_61, 1, x_60); +x_62 = l_Lean_addTrace___at_Lean_Meta_Grind_Canon_canonElemCore___spec__5(x_13, x_61, x_47, x_8, x_9, x_10, x_11, x_23); +x_63 = lean_ctor_get(x_62, 0); +lean_inc(x_63); +x_64 = lean_ctor_get(x_62, 1); +lean_inc(x_64); +lean_dec(x_62); +x_65 = lean_ctor_get(x_63, 0); +lean_inc(x_65); +x_66 = lean_ctor_get(x_63, 1); +lean_inc(x_66); +lean_dec(x_63); +x_67 = l_Lean_Meta_Grind_Canon_canonElemCore___lambda__1(x_1, x_2, x_3, x_65, x_66, x_8, x_9, x_10, x_11, x_64); +lean_dec(x_65); +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; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; +x_68 = lean_ctor_get(x_14, 1); +lean_inc(x_68); +lean_dec(x_14); +x_69 = lean_ctor_get(x_15, 1); +lean_inc(x_69); +if (lean_is_exclusive(x_15)) { + lean_ctor_release(x_15, 0); + lean_ctor_release(x_15, 1); + x_70 = x_15; +} else { + lean_dec_ref(x_15); + x_70 = lean_box(0); +} +x_71 = l_Lean_MessageData_ofExpr(x_4); +x_72 = l_Lean_Meta_Grind_Canon_canonElemCore___lambda__2___closed__2; +if (lean_is_scalar(x_70)) { + x_73 = lean_alloc_ctor(7, 2, 0); +} else { + x_73 = x_70; + lean_ctor_set_tag(x_73, 7); +} +lean_ctor_set(x_73, 0, x_72); +lean_ctor_set(x_73, 1, x_71); +x_74 = l_Lean_Meta_Grind_Canon_canonElemCore___lambda__2___closed__4; +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___private_Init_Data_Repr_0__Nat_reprFast(x_5); +x_77 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_77, 0, x_76); +x_78 = l_Lean_MessageData_ofFormat(x_77); +x_79 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_79, 0, x_75); +lean_ctor_set(x_79, 1, x_78); +x_80 = l_Lean_Meta_Grind_Canon_canonElemCore___lambda__2___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); +lean_inc(x_1); +x_82 = l_Lean_MessageData_ofExpr(x_1); +x_83 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_83, 0, x_81); +lean_ctor_set(x_83, 1, x_82); +x_84 = l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1___closed__10; +x_85 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_85, 0, x_83); +lean_ctor_set(x_85, 1, x_84); +x_86 = l_Lean_addTrace___at_Lean_Meta_Grind_Canon_canonElemCore___spec__5(x_13, x_85, x_69, x_8, x_9, x_10, x_11, x_68); +x_87 = lean_ctor_get(x_86, 0); +lean_inc(x_87); +x_88 = lean_ctor_get(x_86, 1); +lean_inc(x_88); +lean_dec(x_86); +x_89 = lean_ctor_get(x_87, 0); +lean_inc(x_89); +x_90 = lean_ctor_get(x_87, 1); +lean_inc(x_90); +lean_dec(x_87); +x_91 = l_Lean_Meta_Grind_Canon_canonElemCore___lambda__1(x_1, x_2, x_3, x_89, x_90, x_8, x_9, x_10, x_11, x_88); +lean_dec(x_89); +return x_91; +} +} +} +} +static lean_object* _init_l_Lean_Meta_Grind_Canon_canonElemCore___lambda__3___closed__1() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); x_2 = lean_box(0); @@ -3796,10 +4376,12 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_canonElemCore___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_canonElemCore___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; +lean_inc(x_2); +lean_inc(x_1); x_13 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_13, 0, x_1); lean_ctor_set(x_13, 1, x_2); @@ -3827,7 +4409,7 @@ goto block_45; block_45: { lean_object* x_18; lean_object* x_19; -x_18 = l_Lean_Meta_Grind_Canon_canonElemCore___lambda__2___closed__1; +x_18 = l_Lean_Meta_Grind_Canon_canonElemCore___lambda__3___closed__1; lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); @@ -3855,7 +4437,7 @@ 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_Meta_Grind_Canon_canonElemCore___lambda__1(x_4, x_17, x_13, x_25, x_24, x_8, x_9, x_10, x_11, x_23); +x_26 = l_Lean_Meta_Grind_Canon_canonElemCore___lambda__2(x_4, x_17, x_13, x_1, x_2, x_25, x_24, x_8, x_9, x_10, x_11, x_23); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -3872,6 +4454,8 @@ lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); x_27 = !lean_is_exclusive(x_19); if (x_27 == 0) { @@ -3949,6 +4533,8 @@ lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); x_41 = !lean_is_exclusive(x_19); if (x_41 == 0) { @@ -3983,7 +4569,7 @@ if (lean_obj_tag(x_12) == 0) lean_object* x_13; lean_object* x_14; x_13 = lean_box(0); lean_inc(x_5); -x_14 = l_Lean_Meta_Grind_Canon_canonElemCore___lambda__2(x_1, x_2, x_5, x_3, x_4, x_13, x_5, x_6, x_7, x_8, x_9, x_10); +x_14 = l_Lean_Meta_Grind_Canon_canonElemCore___lambda__3(x_1, x_2, x_5, x_3, x_4, x_13, x_5, x_6, x_7, x_8, x_9, x_10); return x_14; } else @@ -4099,6 +4685,19 @@ lean_dec(x_5); return x_13; } } +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___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_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +return x_9; +} +} LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___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: { @@ -4183,10 +4782,23 @@ return x_11; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_canonElemCore___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { +lean_object* x_13; +x_13 = l_Lean_Meta_Grind_Canon_canonElemCore___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_6); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_canonElemCore___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ uint8_t x_13; lean_object* x_14; x_13 = lean_unbox(x_5); lean_dec(x_5); -x_14 = l_Lean_Meta_Grind_Canon_canonElemCore___lambda__2(x_1, x_2, x_3, x_4, x_13, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +x_14 = l_Lean_Meta_Grind_Canon_canonElemCore___lambda__3(x_1, x_2, x_3, x_4, x_13, x_6, x_7, x_8, x_9, x_10, x_11, x_12); lean_dec(x_6); return x_14; } @@ -4220,10 +4832,10 @@ lean_ctor_set_uint8(x_11, 9, x_14); x_15 = 2; x_16 = lean_uint64_shift_right(x_13, x_15); x_17 = lean_uint64_shift_left(x_16, x_15); -x_18 = l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1___closed__9; +x_18 = l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1___closed__11; x_19 = lean_uint64_lor(x_17, x_18); lean_ctor_set_uint64(x_5, sizeof(void*)*7, x_19); -x_20 = 0; +x_20 = 1; x_21 = l_Lean_Meta_Grind_Canon_canonElemCore(x_1, x_2, x_3, x_20, x_4, x_5, x_6, x_7, x_8, x_9); if (lean_obj_tag(x_21) == 0) { @@ -4352,11 +4964,11 @@ lean_ctor_set_uint8(x_57, 16, x_55); x_58 = 2; x_59 = lean_uint64_shift_right(x_39, x_58); x_60 = lean_uint64_shift_left(x_59, x_58); -x_61 = l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1___closed__9; +x_61 = l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1___closed__11; x_62 = lean_uint64_lor(x_60, x_61); lean_ctor_set(x_5, 0, x_57); lean_ctor_set_uint64(x_5, sizeof(void*)*7, x_62); -x_63 = 0; +x_63 = 1; x_64 = l_Lean_Meta_Grind_Canon_canonElemCore(x_1, x_2, x_3, x_63, x_4, x_5, x_6, x_7, x_8, x_9); if (lean_obj_tag(x_64) == 0) { @@ -4497,7 +5109,7 @@ lean_ctor_set_uint8(x_106, 16, x_103); x_107 = 2; x_108 = lean_uint64_shift_right(x_78, x_107); x_109 = lean_uint64_shift_left(x_108, x_107); -x_110 = l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1___closed__9; +x_110 = l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1___closed__11; x_111 = lean_uint64_lor(x_109, x_110); x_112 = lean_alloc_ctor(0, 7, 11); lean_ctor_set(x_112, 0, x_106); @@ -4511,7 +5123,7 @@ lean_ctor_set_uint64(x_112, sizeof(void*)*7, x_111); lean_ctor_set_uint8(x_112, sizeof(void*)*7 + 8, x_79); lean_ctor_set_uint8(x_112, sizeof(void*)*7 + 9, x_86); lean_ctor_set_uint8(x_112, sizeof(void*)*7 + 10, x_87); -x_113 = 0; +x_113 = 1; x_114 = l_Lean_Meta_Grind_Canon_canonElemCore(x_1, x_2, x_3, x_113, x_4, x_112, x_6, x_7, x_8, x_9); if (lean_obj_tag(x_114) == 0) { @@ -4614,7 +5226,7 @@ x_17 = lean_uint64_shift_left(x_16, x_15); x_18 = l_Lean_Meta_Grind_Canon_canonInst___closed__1; x_19 = lean_uint64_lor(x_17, x_18); lean_ctor_set_uint64(x_5, sizeof(void*)*7, x_19); -x_20 = 1; +x_20 = 0; x_21 = l_Lean_Meta_Grind_Canon_canonElemCore(x_1, x_2, x_3, x_20, x_4, x_5, x_6, x_7, x_8, x_9); if (lean_obj_tag(x_21) == 0) { @@ -4747,7 +5359,7 @@ x_61 = l_Lean_Meta_Grind_Canon_canonInst___closed__1; x_62 = lean_uint64_lor(x_60, x_61); lean_ctor_set(x_5, 0, x_57); lean_ctor_set_uint64(x_5, sizeof(void*)*7, x_62); -x_63 = 1; +x_63 = 0; x_64 = l_Lean_Meta_Grind_Canon_canonElemCore(x_1, x_2, x_3, x_63, x_4, x_5, x_6, x_7, x_8, x_9); if (lean_obj_tag(x_64) == 0) { @@ -4902,7 +5514,7 @@ lean_ctor_set_uint64(x_112, sizeof(void*)*7, x_111); lean_ctor_set_uint8(x_112, sizeof(void*)*7 + 8, x_79); lean_ctor_set_uint8(x_112, sizeof(void*)*7 + 9, x_86); lean_ctor_set_uint8(x_112, sizeof(void*)*7 + 10, x_87); -x_113 = 1; +x_113 = 0; x_114 = l_Lean_Meta_Grind_Canon_canonElemCore(x_1, x_2, x_3, x_113, x_4, x_112, x_6, x_7, x_8, x_9); if (lean_obj_tag(x_114) == 0) { @@ -4974,952 +5586,3545 @@ return x_126; } } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Canon_0__Lean_Meta_Grind_Canon_ShouldCanonResult_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___private_Lean_Meta_Tactic_Grind_Canon_0__Lean_Meta_Grind_Canon_ShouldCanonResult_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___private_Lean_Meta_Tactic_Grind_Canon_0__Lean_Meta_Grind_Canon_ShouldCanonResult_toCtorIdx(x_2); -return x_3; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Canon_0__Lean_Meta_Grind_Canon_ShouldCanonResult_noConfusion___rarg___lambda__1(lean_object* x_1) { -_start: -{ -lean_inc(x_1); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Canon_0__Lean_Meta_Grind_Canon_ShouldCanonResult_noConfusion___rarg___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Canon_0__Lean_Meta_Grind_Canon_ShouldCanonResult_noConfusion___rarg___lambda__1___boxed), 1, 0); -return x_1; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Canon_0__Lean_Meta_Grind_Canon_ShouldCanonResult_noConfusion___rarg(uint8_t x_1, uint8_t x_2, lean_object* x_3) { -_start: -{ -lean_object* x_4; -x_4 = l___private_Lean_Meta_Tactic_Grind_Canon_0__Lean_Meta_Grind_Canon_ShouldCanonResult_noConfusion___rarg___closed__1; -return x_4; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Canon_0__Lean_Meta_Grind_Canon_ShouldCanonResult_noConfusion(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Canon_0__Lean_Meta_Grind_Canon_ShouldCanonResult_noConfusion___rarg___boxed), 3, 0); -return x_2; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Canon_0__Lean_Meta_Grind_Canon_ShouldCanonResult_noConfusion___rarg___lambda__1___boxed(lean_object* x_1) { +static uint64_t _init_l_Lean_Meta_Grind_Canon_canonImplicit___closed__1() { _start: { -lean_object* x_2; -x_2 = l___private_Lean_Meta_Tactic_Grind_Canon_0__Lean_Meta_Grind_Canon_ShouldCanonResult_noConfusion___rarg___lambda__1(x_1); -lean_dec(x_1); +uint8_t x_1; uint64_t x_2; +x_1 = 2; +x_2 = l_Lean_Meta_TransparencyMode_toUInt64(x_1); return x_2; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Canon_0__Lean_Meta_Grind_Canon_ShouldCanonResult_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___private_Lean_Meta_Tactic_Grind_Canon_0__Lean_Meta_Grind_Canon_ShouldCanonResult_noConfusion___rarg(x_4, x_5, x_3); -return x_6; -} -} -static uint8_t _init_l_Lean_Meta_Grind_Canon_instInhabitedShouldCanonResult() { -_start: -{ -uint8_t x_1; -x_1 = 0; -return x_1; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_shouldCanon___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_canonImplicit(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { -lean_object* x_8; -x_8 = l_Lean_Meta_isTypeFormer(x_1, x_3, x_4, x_5, x_6, x_7); -if (lean_obj_tag(x_8) == 0) -{ -lean_object* x_9; uint8_t x_10; -x_9 = lean_ctor_get(x_8, 0); -lean_inc(x_9); -x_10 = lean_unbox(x_9); -lean_dec(x_9); +uint8_t x_10; +x_10 = !lean_is_exclusive(x_5); if (x_10 == 0) { -uint8_t x_11; -x_11 = !lean_is_exclusive(x_8); -if (x_11 == 0) +lean_object* x_11; uint8_t x_12; +x_11 = lean_ctor_get(x_5, 0); +x_12 = !lean_is_exclusive(x_11); +if (x_12 == 0) { -lean_object* x_12; uint8_t x_13; lean_object* x_14; -x_12 = lean_ctor_get(x_8, 0); -lean_dec(x_12); -x_13 = 2; -x_14 = lean_box(x_13); -lean_ctor_set(x_8, 0, x_14); -return x_8; -} -else +uint64_t x_13; uint8_t x_14; uint64_t x_15; uint64_t x_16; uint64_t x_17; uint64_t x_18; uint64_t x_19; uint8_t x_20; lean_object* x_21; +x_13 = lean_ctor_get_uint64(x_5, sizeof(void*)*7); +x_14 = 2; +lean_ctor_set_uint8(x_11, 9, x_14); +x_15 = 2; +x_16 = lean_uint64_shift_right(x_13, x_15); +x_17 = lean_uint64_shift_left(x_16, x_15); +x_18 = l_Lean_Meta_Grind_Canon_canonImplicit___closed__1; +x_19 = lean_uint64_lor(x_17, x_18); +lean_ctor_set_uint64(x_5, sizeof(void*)*7, x_19); +x_20 = 2; +x_21 = l_Lean_Meta_Grind_Canon_canonElemCore(x_1, x_2, x_3, x_20, x_4, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_21) == 0) { -lean_object* x_15; uint8_t x_16; lean_object* x_17; lean_object* x_18; -x_15 = lean_ctor_get(x_8, 1); -lean_inc(x_15); -lean_dec(x_8); -x_16 = 2; -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_22; +x_22 = !lean_is_exclusive(x_21); +if (x_22 == 0) { -uint8_t x_19; -x_19 = !lean_is_exclusive(x_8); -if (x_19 == 0) +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_20; uint8_t x_21; lean_object* x_22; -x_20 = lean_ctor_get(x_8, 0); -lean_dec(x_20); -x_21 = 0; -x_22 = lean_box(x_21); -lean_ctor_set(x_8, 0, x_22); -return x_8; +return x_21; } else { -lean_object* x_23; uint8_t x_24; lean_object* x_25; lean_object* x_26; -x_23 = lean_ctor_get(x_8, 1); -lean_inc(x_23); -lean_dec(x_8); -x_24 = 0; -x_25 = lean_box(x_24); -x_26 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_26, 0, x_25); -lean_ctor_set(x_26, 1, x_23); -return x_26; -} -} +lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_25 = lean_ctor_get(x_23, 0); +x_26 = lean_ctor_get(x_23, 1); +lean_inc(x_26); +lean_inc(x_25); +lean_dec(x_23); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +lean_ctor_set(x_21, 0, x_27); +return x_21; } -else -{ -uint8_t x_27; -x_27 = !lean_is_exclusive(x_8); -if (x_27 == 0) -{ -return x_8; } else { -lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_28 = lean_ctor_get(x_8, 0); -x_29 = lean_ctor_get(x_8, 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; +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_8); -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_30 = lean_ctor_get(x_28, 0); +lean_inc(x_30); +x_31 = lean_ctor_get(x_28, 1); +lean_inc(x_31); +if (lean_is_exclusive(x_28)) { + lean_ctor_release(x_28, 0); + lean_ctor_release(x_28, 1); + x_32 = x_28; +} else { + lean_dec_ref(x_28); + x_32 = lean_box(0); } +if (lean_is_scalar(x_32)) { + x_33 = lean_alloc_ctor(0, 2, 0); +} else { + x_33 = x_32; } +lean_ctor_set(x_33, 0, x_30); +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_29); +return x_34; } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_shouldCanon(lean_object* x_1, lean_object* x_2, lean_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_array_get_size(x_1); -x_10 = lean_nat_dec_lt(x_2, x_9); -lean_dec(x_9); -if (x_10 == 0) -{ -lean_object* x_11; lean_object* x_12; -x_11 = lean_box(0); -x_12 = l_Lean_Meta_Grind_Canon_shouldCanon___lambda__1(x_3, x_11, x_4, x_5, x_6, x_7, x_8); -return x_12; } else { -lean_object* x_13; uint8_t x_14; -x_13 = lean_array_fget(x_1, x_2); -x_14 = l_Lean_Meta_ParamInfo_isInstImplicit(x_13); -if (x_14 == 0) -{ -uint8_t x_15; -x_15 = lean_ctor_get_uint8(x_13, sizeof(void*)*1 + 2); -lean_dec(x_13); -if (x_15 == 0) -{ -lean_object* x_16; lean_object* x_17; -x_16 = lean_box(0); -x_17 = l_Lean_Meta_Grind_Canon_shouldCanon___lambda__1(x_3, x_16, x_4, x_5, x_6, x_7, x_8); -return x_17; -} -else +uint8_t x_35; +x_35 = !lean_is_exclusive(x_21); +if (x_35 == 0) { -uint8_t x_18; lean_object* x_19; lean_object* x_20; -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -x_18 = 2; -x_19 = lean_box(x_18); -x_20 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_8); -return x_20; -} +return x_21; } else { -uint8_t x_21; lean_object* x_22; lean_object* x_23; -lean_dec(x_13); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -x_21 = 1; -x_22 = lean_box(x_21); -x_23 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_23, 0, x_22); -lean_ctor_set(x_23, 1, x_8); -return x_23; -} +lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_21, 0); +x_37 = lean_ctor_get(x_21, 1); +lean_inc(x_37); +lean_inc(x_36); +lean_dec(x_21); +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_Grind_Canon_shouldCanon___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: +else { -lean_object* x_8; +uint64_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; uint8_t x_52; uint8_t x_53; uint8_t x_54; uint8_t x_55; uint8_t x_56; lean_object* x_57; uint64_t x_58; uint64_t x_59; uint64_t x_60; uint64_t x_61; uint64_t x_62; uint8_t x_63; lean_object* x_64; +x_39 = lean_ctor_get_uint64(x_5, sizeof(void*)*7); +x_40 = lean_ctor_get_uint8(x_11, 0); +x_41 = lean_ctor_get_uint8(x_11, 1); +x_42 = lean_ctor_get_uint8(x_11, 2); +x_43 = lean_ctor_get_uint8(x_11, 3); +x_44 = lean_ctor_get_uint8(x_11, 4); +x_45 = lean_ctor_get_uint8(x_11, 5); +x_46 = lean_ctor_get_uint8(x_11, 6); +x_47 = lean_ctor_get_uint8(x_11, 7); +x_48 = lean_ctor_get_uint8(x_11, 8); +x_49 = lean_ctor_get_uint8(x_11, 10); +x_50 = lean_ctor_get_uint8(x_11, 11); +x_51 = lean_ctor_get_uint8(x_11, 12); +x_52 = lean_ctor_get_uint8(x_11, 13); +x_53 = lean_ctor_get_uint8(x_11, 14); +x_54 = lean_ctor_get_uint8(x_11, 15); +x_55 = lean_ctor_get_uint8(x_11, 16); +lean_dec(x_11); +x_56 = 2; +x_57 = lean_alloc_ctor(0, 0, 17); +lean_ctor_set_uint8(x_57, 0, x_40); +lean_ctor_set_uint8(x_57, 1, x_41); +lean_ctor_set_uint8(x_57, 2, x_42); +lean_ctor_set_uint8(x_57, 3, x_43); +lean_ctor_set_uint8(x_57, 4, x_44); +lean_ctor_set_uint8(x_57, 5, x_45); +lean_ctor_set_uint8(x_57, 6, x_46); +lean_ctor_set_uint8(x_57, 7, x_47); +lean_ctor_set_uint8(x_57, 8, x_48); +lean_ctor_set_uint8(x_57, 9, x_56); +lean_ctor_set_uint8(x_57, 10, x_49); +lean_ctor_set_uint8(x_57, 11, x_50); +lean_ctor_set_uint8(x_57, 12, x_51); +lean_ctor_set_uint8(x_57, 13, x_52); +lean_ctor_set_uint8(x_57, 14, x_53); +lean_ctor_set_uint8(x_57, 15, x_54); +lean_ctor_set_uint8(x_57, 16, x_55); +x_58 = 2; +x_59 = lean_uint64_shift_right(x_39, x_58); +x_60 = lean_uint64_shift_left(x_59, x_58); +x_61 = l_Lean_Meta_Grind_Canon_canonImplicit___closed__1; +x_62 = lean_uint64_lor(x_60, x_61); +lean_ctor_set(x_5, 0, x_57); +lean_ctor_set_uint64(x_5, sizeof(void*)*7, x_62); +x_63 = 2; +x_64 = l_Lean_Meta_Grind_Canon_canonElemCore(x_1, x_2, x_3, x_63, x_4, x_5, x_6, x_7, x_8, x_9); +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; +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); +} +x_68 = lean_ctor_get(x_65, 0); +lean_inc(x_68); +x_69 = lean_ctor_get(x_65, 1); +lean_inc(x_69); +if (lean_is_exclusive(x_65)) { + lean_ctor_release(x_65, 0); + lean_ctor_release(x_65, 1); + x_70 = x_65; +} else { + lean_dec_ref(x_65); + x_70 = lean_box(0); +} +if (lean_is_scalar(x_70)) { + x_71 = 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); +if (lean_is_scalar(x_67)) { + x_72 = lean_alloc_ctor(0, 2, 0); +} else { + x_72 = x_67; +} +lean_ctor_set(x_72, 0, x_71); +lean_ctor_set(x_72, 1, x_66); +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_64, 0); +lean_inc(x_73); +x_74 = lean_ctor_get(x_64, 1); +lean_inc(x_74); +if (lean_is_exclusive(x_64)) { + lean_ctor_release(x_64, 0); + lean_ctor_release(x_64, 1); + x_75 = x_64; +} else { + lean_dec_ref(x_64); + 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; +} +} +} +else +{ +lean_object* x_77; uint64_t 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; 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; uint8_t x_99; uint8_t x_100; uint8_t x_101; uint8_t x_102; uint8_t x_103; lean_object* x_104; uint8_t x_105; lean_object* x_106; uint64_t x_107; uint64_t x_108; uint64_t x_109; uint64_t x_110; uint64_t x_111; lean_object* x_112; uint8_t x_113; lean_object* x_114; +x_77 = lean_ctor_get(x_5, 0); +x_78 = lean_ctor_get_uint64(x_5, sizeof(void*)*7); +x_79 = lean_ctor_get_uint8(x_5, sizeof(void*)*7 + 8); +x_80 = lean_ctor_get(x_5, 1); +x_81 = lean_ctor_get(x_5, 2); +x_82 = lean_ctor_get(x_5, 3); +x_83 = lean_ctor_get(x_5, 4); +x_84 = lean_ctor_get(x_5, 5); +x_85 = lean_ctor_get(x_5, 6); +x_86 = lean_ctor_get_uint8(x_5, sizeof(void*)*7 + 9); +x_87 = lean_ctor_get_uint8(x_5, sizeof(void*)*7 + 10); +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_inc(x_77); +lean_dec(x_5); +x_88 = lean_ctor_get_uint8(x_77, 0); +x_89 = lean_ctor_get_uint8(x_77, 1); +x_90 = lean_ctor_get_uint8(x_77, 2); +x_91 = lean_ctor_get_uint8(x_77, 3); +x_92 = lean_ctor_get_uint8(x_77, 4); +x_93 = lean_ctor_get_uint8(x_77, 5); +x_94 = lean_ctor_get_uint8(x_77, 6); +x_95 = lean_ctor_get_uint8(x_77, 7); +x_96 = lean_ctor_get_uint8(x_77, 8); +x_97 = lean_ctor_get_uint8(x_77, 10); +x_98 = lean_ctor_get_uint8(x_77, 11); +x_99 = lean_ctor_get_uint8(x_77, 12); +x_100 = lean_ctor_get_uint8(x_77, 13); +x_101 = lean_ctor_get_uint8(x_77, 14); +x_102 = lean_ctor_get_uint8(x_77, 15); +x_103 = lean_ctor_get_uint8(x_77, 16); +if (lean_is_exclusive(x_77)) { + x_104 = x_77; +} else { + lean_dec_ref(x_77); + x_104 = lean_box(0); +} +x_105 = 2; +if (lean_is_scalar(x_104)) { + x_106 = lean_alloc_ctor(0, 0, 17); +} else { + x_106 = x_104; +} +lean_ctor_set_uint8(x_106, 0, x_88); +lean_ctor_set_uint8(x_106, 1, x_89); +lean_ctor_set_uint8(x_106, 2, x_90); +lean_ctor_set_uint8(x_106, 3, x_91); +lean_ctor_set_uint8(x_106, 4, x_92); +lean_ctor_set_uint8(x_106, 5, x_93); +lean_ctor_set_uint8(x_106, 6, x_94); +lean_ctor_set_uint8(x_106, 7, x_95); +lean_ctor_set_uint8(x_106, 8, x_96); +lean_ctor_set_uint8(x_106, 9, x_105); +lean_ctor_set_uint8(x_106, 10, x_97); +lean_ctor_set_uint8(x_106, 11, x_98); +lean_ctor_set_uint8(x_106, 12, x_99); +lean_ctor_set_uint8(x_106, 13, x_100); +lean_ctor_set_uint8(x_106, 14, x_101); +lean_ctor_set_uint8(x_106, 15, x_102); +lean_ctor_set_uint8(x_106, 16, x_103); +x_107 = 2; +x_108 = lean_uint64_shift_right(x_78, x_107); +x_109 = lean_uint64_shift_left(x_108, x_107); +x_110 = l_Lean_Meta_Grind_Canon_canonImplicit___closed__1; +x_111 = lean_uint64_lor(x_109, x_110); +x_112 = lean_alloc_ctor(0, 7, 11); +lean_ctor_set(x_112, 0, x_106); +lean_ctor_set(x_112, 1, x_80); +lean_ctor_set(x_112, 2, x_81); +lean_ctor_set(x_112, 3, x_82); +lean_ctor_set(x_112, 4, x_83); +lean_ctor_set(x_112, 5, x_84); +lean_ctor_set(x_112, 6, x_85); +lean_ctor_set_uint64(x_112, sizeof(void*)*7, x_111); +lean_ctor_set_uint8(x_112, sizeof(void*)*7 + 8, x_79); +lean_ctor_set_uint8(x_112, sizeof(void*)*7 + 9, x_86); +lean_ctor_set_uint8(x_112, sizeof(void*)*7 + 10, x_87); +x_113 = 2; +x_114 = l_Lean_Meta_Grind_Canon_canonElemCore(x_1, x_2, x_3, x_113, x_4, x_112, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_114) == 0) +{ +lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; +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_114); + x_117 = lean_box(0); +} +x_118 = lean_ctor_get(x_115, 0); +lean_inc(x_118); +x_119 = lean_ctor_get(x_115, 1); +lean_inc(x_119); +if (lean_is_exclusive(x_115)) { + lean_ctor_release(x_115, 0); + lean_ctor_release(x_115, 1); + x_120 = x_115; +} else { + lean_dec_ref(x_115); + x_120 = lean_box(0); +} +if (lean_is_scalar(x_120)) { + x_121 = lean_alloc_ctor(0, 2, 0); +} else { + x_121 = x_120; +} +lean_ctor_set(x_121, 0, x_118); +lean_ctor_set(x_121, 1, x_119); +if (lean_is_scalar(x_117)) { + x_122 = lean_alloc_ctor(0, 2, 0); +} else { + x_122 = x_117; +} +lean_ctor_set(x_122, 0, x_121); +lean_ctor_set(x_122, 1, x_116); +return x_122; +} +else +{ +lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; +x_123 = lean_ctor_get(x_114, 0); +lean_inc(x_123); +x_124 = lean_ctor_get(x_114, 1); +lean_inc(x_124); +if (lean_is_exclusive(x_114)) { + lean_ctor_release(x_114, 0); + lean_ctor_release(x_114, 1); + x_125 = x_114; +} else { + lean_dec_ref(x_114); + x_125 = lean_box(0); +} +if (lean_is_scalar(x_125)) { + x_126 = lean_alloc_ctor(1, 2, 0); +} else { + x_126 = x_125; +} +lean_ctor_set(x_126, 0, x_123); +lean_ctor_set(x_126, 1, x_124); +return x_126; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Canon_0__Lean_Meta_Grind_Canon_ShouldCanonResult_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; +} +case 2: +{ +lean_object* x_4; +x_4 = lean_unsigned_to_nat(2u); +return x_4; +} +default: +{ +lean_object* x_5; +x_5 = lean_unsigned_to_nat(3u); +return x_5; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Canon_0__Lean_Meta_Grind_Canon_ShouldCanonResult_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___private_Lean_Meta_Tactic_Grind_Canon_0__Lean_Meta_Grind_Canon_ShouldCanonResult_toCtorIdx(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Canon_0__Lean_Meta_Grind_Canon_ShouldCanonResult_noConfusion___rarg(uint8_t x_1, uint8_t x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Lean_Meta_Grind_Canon_CanonElemKind_noConfusion___rarg___closed__1; +return x_4; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Canon_0__Lean_Meta_Grind_Canon_ShouldCanonResult_noConfusion(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Canon_0__Lean_Meta_Grind_Canon_ShouldCanonResult_noConfusion___rarg___boxed), 3, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Canon_0__Lean_Meta_Grind_Canon_ShouldCanonResult_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___private_Lean_Meta_Tactic_Grind_Canon_0__Lean_Meta_Grind_Canon_ShouldCanonResult_noConfusion___rarg(x_4, x_5, x_3); +return x_6; +} +} +static uint8_t _init_l_Lean_Meta_Grind_Canon_instInhabitedShouldCanonResult() { +_start: +{ +uint8_t x_1; +x_1 = 0; +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_Canon_instReprShouldCanonResult___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("canonType", 9, 9); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_Canon_instReprShouldCanonResult___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_Canon_instReprShouldCanonResult___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_Meta_Grind_Canon_instReprShouldCanonResult___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("canonInst", 9, 9); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_Canon_instReprShouldCanonResult___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_Canon_instReprShouldCanonResult___closed__3; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Grind_Canon_instReprShouldCanonResult___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("canonImplicit", 13, 13); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_Canon_instReprShouldCanonResult___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_Canon_instReprShouldCanonResult___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_Grind_Canon_instReprShouldCanonResult___closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("visit", 5, 5); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_Canon_instReprShouldCanonResult___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_Canon_instReprShouldCanonResult___closed__7; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_instReprShouldCanonResult(uint8_t x_1, lean_object* x_2) { +_start: +{ +switch (x_1) { +case 0: +{ +lean_object* x_3; +x_3 = l_Lean_Meta_Grind_Canon_instReprShouldCanonResult___closed__2; +return x_3; +} +case 1: +{ +lean_object* x_4; +x_4 = l_Lean_Meta_Grind_Canon_instReprShouldCanonResult___closed__4; +return x_4; +} +case 2: +{ +lean_object* x_5; +x_5 = l_Lean_Meta_Grind_Canon_instReprShouldCanonResult___closed__6; +return x_5; +} +default: +{ +lean_object* x_6; +x_6 = l_Lean_Meta_Grind_Canon_instReprShouldCanonResult___closed__8; +return x_6; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_instReprShouldCanonResult___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = lean_unbox(x_1); +lean_dec(x_1); +x_4 = l_Lean_Meta_Grind_Canon_instReprShouldCanonResult(x_3, x_2); +lean_dec(x_2); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_shouldCanon___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_1); +x_8 = l_Lean_Meta_isProp(x_1, x_3, x_4, x_5, x_6, x_7); +if (lean_obj_tag(x_8) == 0) +{ +lean_object* x_9; uint8_t x_10; +x_9 = lean_ctor_get(x_8, 0); +lean_inc(x_9); +x_10 = lean_unbox(x_9); +lean_dec(x_9); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; +x_11 = lean_ctor_get(x_8, 1); +lean_inc(x_11); +lean_dec(x_8); +x_12 = l_Lean_Meta_isTypeFormer(x_1, x_3, x_4, x_5, x_6, x_11); +if (lean_obj_tag(x_12) == 0) +{ +lean_object* x_13; uint8_t x_14; +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_unbox(x_13); +lean_dec(x_13); +if (x_14 == 0) +{ +uint8_t x_15; +x_15 = !lean_is_exclusive(x_12); +if (x_15 == 0) +{ +lean_object* x_16; uint8_t x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_12, 0); +lean_dec(x_16); +x_17 = 3; +x_18 = lean_box(x_17); +lean_ctor_set(x_12, 0, x_18); +return x_12; +} +else +{ +lean_object* x_19; uint8_t x_20; lean_object* x_21; lean_object* x_22; +x_19 = lean_ctor_get(x_12, 1); +lean_inc(x_19); +lean_dec(x_12); +x_20 = 3; +x_21 = lean_box(x_20); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_19); +return x_22; +} +} +else +{ +uint8_t x_23; +x_23 = !lean_is_exclusive(x_12); +if (x_23 == 0) +{ +lean_object* x_24; uint8_t x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_12, 0); +lean_dec(x_24); +x_25 = 0; +x_26 = lean_box(x_25); +lean_ctor_set(x_12, 0, x_26); +return x_12; +} +else +{ +lean_object* x_27; uint8_t x_28; lean_object* x_29; lean_object* x_30; +x_27 = lean_ctor_get(x_12, 1); +lean_inc(x_27); +lean_dec(x_12); +x_28 = 0; +x_29 = lean_box(x_28); +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_27); +return x_30; +} +} +} +else +{ +uint8_t x_31; +x_31 = !lean_is_exclusive(x_12); +if (x_31 == 0) +{ +return x_12; +} +else +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_12, 0); +x_33 = lean_ctor_get(x_12, 1); +lean_inc(x_33); +lean_inc(x_32); +lean_dec(x_12); +x_34 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +return x_34; +} +} +} +else +{ +uint8_t x_35; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_35 = !lean_is_exclusive(x_8); +if (x_35 == 0) +{ +lean_object* x_36; uint8_t x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_8, 0); +lean_dec(x_36); +x_37 = 3; +x_38 = lean_box(x_37); +lean_ctor_set(x_8, 0, x_38); +return x_8; +} +else +{ +lean_object* x_39; uint8_t x_40; lean_object* x_41; lean_object* x_42; +x_39 = lean_ctor_get(x_8, 1); +lean_inc(x_39); +lean_dec(x_8); +x_40 = 3; +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); +return x_42; +} +} +} +else +{ +uint8_t x_43; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_43 = !lean_is_exclusive(x_8); +if (x_43 == 0) +{ +return x_8; +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_8, 0); +x_45 = lean_ctor_get(x_8, 1); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_8); +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_EXPORT lean_object* l_Lean_Meta_Grind_Canon_shouldCanon(lean_object* x_1, lean_object* x_2, lean_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_array_get_size(x_1); +x_10 = lean_nat_dec_lt(x_2, x_9); +lean_dec(x_9); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; +x_11 = lean_box(0); +x_12 = l_Lean_Meta_Grind_Canon_shouldCanon___lambda__1(x_3, x_11, x_4, x_5, x_6, x_7, x_8); +return x_12; +} +else +{ +lean_object* x_13; uint8_t x_14; +x_13 = lean_array_fget(x_1, x_2); +x_14 = l_Lean_Meta_ParamInfo_isInstImplicit(x_13); +if (x_14 == 0) +{ +uint8_t x_15; +x_15 = lean_ctor_get_uint8(x_13, sizeof(void*)*1 + 2); +if (x_15 == 0) +{ +uint8_t x_16; +x_16 = l_Lean_Meta_ParamInfo_isImplicit(x_13); +lean_dec(x_13); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; +x_17 = lean_box(0); +x_18 = l_Lean_Meta_Grind_Canon_shouldCanon___lambda__1(x_3, x_17, x_4, x_5, x_6, x_7, x_8); +return x_18; +} +else +{ +lean_object* x_19; +x_19 = l_Lean_Meta_isTypeFormer(x_3, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_19) == 0) +{ +lean_object* x_20; uint8_t x_21; +x_20 = lean_ctor_get(x_19, 0); +lean_inc(x_20); +x_21 = lean_unbox(x_20); +lean_dec(x_20); +if (x_21 == 0) +{ +uint8_t x_22; +x_22 = !lean_is_exclusive(x_19); +if (x_22 == 0) +{ +lean_object* x_23; uint8_t x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_19, 0); +lean_dec(x_23); +x_24 = 2; +x_25 = lean_box(x_24); +lean_ctor_set(x_19, 0, x_25); +return x_19; +} +else +{ +lean_object* x_26; uint8_t x_27; lean_object* x_28; lean_object* x_29; +x_26 = lean_ctor_get(x_19, 1); +lean_inc(x_26); +lean_dec(x_19); +x_27 = 2; +x_28 = lean_box(x_27); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_26); +return x_29; +} +} +else +{ +uint8_t x_30; +x_30 = !lean_is_exclusive(x_19); +if (x_30 == 0) +{ +lean_object* x_31; uint8_t x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_19, 0); +lean_dec(x_31); +x_32 = 0; +x_33 = lean_box(x_32); +lean_ctor_set(x_19, 0, x_33); +return x_19; +} +else +{ +lean_object* x_34; uint8_t x_35; lean_object* x_36; lean_object* x_37; +x_34 = lean_ctor_get(x_19, 1); +lean_inc(x_34); +lean_dec(x_19); +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_34); +return x_37; +} +} +} +else +{ +uint8_t x_38; +x_38 = !lean_is_exclusive(x_19); +if (x_38 == 0) +{ +return x_19; +} +else +{ +lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_39 = lean_ctor_get(x_19, 0); +x_40 = lean_ctor_get(x_19, 1); +lean_inc(x_40); +lean_inc(x_39); +lean_dec(x_19); +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_42; lean_object* x_43; lean_object* x_44; +lean_dec(x_13); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_42 = 3; +x_43 = lean_box(x_42); +x_44 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_44, 0, x_43); +lean_ctor_set(x_44, 1, x_8); +return x_44; +} +} +else +{ +uint8_t x_45; lean_object* x_46; lean_object* x_47; +lean_dec(x_13); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_45 = 1; +x_46 = lean_box(x_45); +x_47 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_47, 0, x_46); +lean_ctor_set(x_47, 1, x_8); +return x_47; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_shouldCanon___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; x_8 = l_Lean_Meta_Grind_Canon_shouldCanon___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7); lean_dec(x_2); -return x_8; +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_shouldCanon___boxed(lean_object* x_1, lean_object* x_2, lean_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_Grind_Canon_shouldCanon(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 uint8_t l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +uint8_t x_3; +x_3 = 0; +return x_3; +} +else +{ +lean_object* x_4; lean_object* x_5; size_t x_6; size_t x_7; uint8_t x_8; +x_4 = lean_ctor_get(x_2, 0); +x_5 = lean_ctor_get(x_2, 2); +x_6 = lean_ptr_addr(x_4); +x_7 = lean_ptr_addr(x_1); +x_8 = lean_usize_dec_eq(x_6, x_7); +if (x_8 == 0) +{ +x_2 = x_5; +goto _start; +} +else +{ +uint8_t x_10; +x_10 = 1; +return x_10; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +if (lean_obj_tag(x_3) == 0) +{ +lean_dec(x_1); +return x_2; +} +else +{ +uint8_t x_4; +x_4 = !lean_is_exclusive(x_3); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint64_t x_9; uint64_t x_10; uint64_t x_11; uint64_t x_12; uint64_t x_13; uint64_t x_14; uint64_t x_15; size_t x_16; size_t x_17; size_t x_18; size_t x_19; size_t x_20; lean_object* x_21; lean_object* x_22; +x_5 = lean_ctor_get(x_3, 0); +x_6 = lean_ctor_get(x_3, 2); +x_7 = lean_array_get_size(x_2); +lean_inc(x_1); +lean_inc(x_5); +x_8 = lean_apply_1(x_1, x_5); +x_9 = lean_unbox_uint64(x_8); +lean_dec(x_8); +x_10 = 32; +x_11 = lean_uint64_shift_right(x_9, x_10); +x_12 = lean_uint64_xor(x_9, x_11); +x_13 = 16; +x_14 = lean_uint64_shift_right(x_12, x_13); +x_15 = lean_uint64_xor(x_12, x_14); +x_16 = lean_uint64_to_usize(x_15); +x_17 = lean_usize_of_nat(x_7); +lean_dec(x_7); +x_18 = 1; +x_19 = lean_usize_sub(x_17, x_18); +x_20 = lean_usize_land(x_16, x_19); +x_21 = lean_array_uget(x_2, x_20); +lean_ctor_set(x_3, 2, x_21); +x_22 = lean_array_uset(x_2, x_20, x_3); +x_2 = x_22; +x_3 = x_6; +goto _start; +} +else +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; uint64_t x_29; uint64_t x_30; uint64_t x_31; uint64_t x_32; uint64_t x_33; uint64_t x_34; uint64_t x_35; size_t x_36; size_t x_37; size_t x_38; size_t x_39; size_t x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_24 = lean_ctor_get(x_3, 0); +x_25 = lean_ctor_get(x_3, 1); +x_26 = lean_ctor_get(x_3, 2); +lean_inc(x_26); +lean_inc(x_25); +lean_inc(x_24); +lean_dec(x_3); +x_27 = lean_array_get_size(x_2); +lean_inc(x_1); +lean_inc(x_24); +x_28 = lean_apply_1(x_1, x_24); +x_29 = lean_unbox_uint64(x_28); +lean_dec(x_28); +x_30 = 32; +x_31 = lean_uint64_shift_right(x_29, x_30); +x_32 = lean_uint64_xor(x_29, x_31); +x_33 = 16; +x_34 = lean_uint64_shift_right(x_32, x_33); +x_35 = lean_uint64_xor(x_32, x_34); +x_36 = lean_uint64_to_usize(x_35); +x_37 = lean_usize_of_nat(x_27); +lean_dec(x_27); +x_38 = 1; +x_39 = lean_usize_sub(x_37, x_38); +x_40 = lean_usize_land(x_36, x_39); +x_41 = lean_array_uget(x_2, x_40); +x_42 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_42, 0, x_24); +lean_ctor_set(x_42, 1, x_25); +lean_ctor_set(x_42, 2, x_41); +x_43 = lean_array_uset(x_2, x_40, x_42); +x_2 = x_43; +x_3 = x_26; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__4___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__5(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +return x_1; +} +else +{ +uint8_t x_3; +x_3 = !lean_is_exclusive(x_2); +if (x_3 == 0) +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; size_t x_7; uint64_t x_8; uint64_t 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; size_t x_17; size_t x_18; size_t x_19; size_t x_20; size_t x_21; lean_object* x_22; lean_object* x_23; +x_4 = lean_ctor_get(x_2, 0); +x_5 = lean_ctor_get(x_2, 2); +x_6 = lean_array_get_size(x_1); +x_7 = lean_ptr_addr(x_4); +x_8 = lean_usize_to_uint64(x_7); +x_9 = 11; +x_10 = lean_uint64_mix_hash(x_8, x_9); +x_11 = 32; +x_12 = lean_uint64_shift_right(x_10, x_11); +x_13 = lean_uint64_xor(x_10, x_12); +x_14 = 16; +x_15 = lean_uint64_shift_right(x_13, x_14); +x_16 = lean_uint64_xor(x_13, x_15); +x_17 = lean_uint64_to_usize(x_16); +x_18 = lean_usize_of_nat(x_6); +lean_dec(x_6); +x_19 = 1; +x_20 = lean_usize_sub(x_18, x_19); +x_21 = lean_usize_land(x_17, x_20); +x_22 = lean_array_uget(x_1, x_21); +lean_ctor_set(x_2, 2, x_22); +x_23 = lean_array_uset(x_1, x_21, x_2); +x_1 = x_23; +x_2 = x_5; +goto _start; +} +else +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; size_t x_29; uint64_t x_30; uint64_t x_31; uint64_t x_32; uint64_t x_33; uint64_t x_34; uint64_t x_35; uint64_t x_36; uint64_t x_37; uint64_t x_38; size_t x_39; size_t x_40; size_t x_41; size_t x_42; size_t x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_25 = lean_ctor_get(x_2, 0); +x_26 = lean_ctor_get(x_2, 1); +x_27 = lean_ctor_get(x_2, 2); +lean_inc(x_27); +lean_inc(x_26); +lean_inc(x_25); +lean_dec(x_2); +x_28 = lean_array_get_size(x_1); +x_29 = lean_ptr_addr(x_25); +x_30 = lean_usize_to_uint64(x_29); +x_31 = 11; +x_32 = lean_uint64_mix_hash(x_30, x_31); +x_33 = 32; +x_34 = lean_uint64_shift_right(x_32, x_33); +x_35 = lean_uint64_xor(x_32, x_34); +x_36 = 16; +x_37 = lean_uint64_shift_right(x_35, x_36); +x_38 = lean_uint64_xor(x_35, x_37); +x_39 = lean_uint64_to_usize(x_38); +x_40 = lean_usize_of_nat(x_28); +lean_dec(x_28); +x_41 = 1; +x_42 = lean_usize_sub(x_40, x_41); +x_43 = lean_usize_land(x_39, x_42); +x_44 = lean_array_uget(x_1, x_43); +x_45 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_45, 0, x_25); +lean_ctor_set(x_45, 1, x_26); +lean_ctor_set(x_45, 2, x_44); +x_46 = lean_array_uset(x_1, x_43, x_45); +x_1 = x_46; +x_2 = x_27; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand_go___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; uint8_t x_5; +x_4 = lean_array_get_size(x_2); +x_5 = lean_nat_dec_lt(x_1, x_4); +lean_dec(x_4); +if (x_5 == 0) +{ +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +else +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_6 = lean_array_fget(x_2, x_1); +x_7 = lean_box(0); +x_8 = lean_array_fset(x_2, x_1, x_7); +x_9 = l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__4___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__5(x_3, x_6); +x_10 = lean_unsigned_to_nat(1u); +x_11 = lean_nat_add(x_1, x_10); +lean_dec(x_1); +x_1 = x_11; +x_2 = x_8; +x_3 = x_9; +goto _start; +} +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__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; +x_2 = lean_array_get_size(x_1); +x_3 = lean_unsigned_to_nat(2u); +x_4 = lean_nat_mul(x_2, x_3); +lean_dec(x_2); +x_5 = lean_box(0); +x_6 = lean_mk_array(x_4, x_5); +x_7 = lean_unsigned_to_nat(0u); +x_8 = l_Std_DHashMap_Internal_Raw_u2080_expand_go___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__3(x_7, x_1, x_6); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +if (lean_obj_tag(x_3) == 0) +{ +lean_object* x_4; +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(0); +return x_4; +} +else +{ +uint8_t x_5; +x_5 = !lean_is_exclusive(x_3); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x_9; size_t x_10; uint8_t x_11; +x_6 = lean_ctor_get(x_3, 0); +x_7 = lean_ctor_get(x_3, 1); +x_8 = lean_ctor_get(x_3, 2); +x_9 = lean_ptr_addr(x_6); +x_10 = lean_ptr_addr(x_1); +x_11 = lean_usize_dec_eq(x_9, x_10); +if (x_11 == 0) +{ +lean_object* x_12; +x_12 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__6(x_1, x_2, x_8); +lean_ctor_set(x_3, 2, x_12); +return x_3; +} +else +{ +lean_dec(x_7); +lean_dec(x_6); +lean_ctor_set(x_3, 1, x_2); +lean_ctor_set(x_3, 0, x_1); +return x_3; +} +} +else +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; size_t x_16; size_t x_17; uint8_t x_18; +x_13 = lean_ctor_get(x_3, 0); +x_14 = lean_ctor_get(x_3, 1); +x_15 = lean_ctor_get(x_3, 2); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_dec(x_3); +x_16 = lean_ptr_addr(x_13); +x_17 = lean_ptr_addr(x_1); +x_18 = lean_usize_dec_eq(x_16, x_17); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; +x_19 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__6(x_1, x_2, x_15); +x_20 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_20, 0, x_13); +lean_ctor_set(x_20, 1, x_14); +lean_ctor_set(x_20, 2, x_19); +return x_20; +} +else +{ +lean_object* x_21; +lean_dec(x_14); +lean_dec(x_13); +x_21 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_21, 0, x_1); +lean_ctor_set(x_21, 1, x_2); +lean_ctor_set(x_21, 2, x_15); +return x_21; +} +} +} +} +} +static lean_object* _init_l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__7___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_instMonadMetaM; +x_2 = l_StateT_instMonad___rarg(x_1); +return x_2; +} +} +static lean_object* _init_l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__7___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__7___closed__1; +x_2 = l_ReaderT_instMonad___rarg(x_1); +return x_2; +} +} +static lean_object* _init_l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__7___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__7___closed__2; +x_2 = l_Lean_instInhabitedExpr; +x_3 = l_instInhabitedOfMonad___rarg(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_9 = l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__7___closed__3; +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; +} +} +LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_Canon_canonImpl_visit___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) { +_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_isTracingEnabledForCore(x_1, x_9, x_8); +x_11 = !lean_is_exclusive(x_10); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; +x_12 = lean_ctor_get(x_10, 0); +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_3); +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(0, 2, 0); +lean_ctor_set(x_16, 0, x_14); +lean_ctor_set(x_16, 1, 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); +return x_17; +} +} +} +LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Meta_Grind_Canon_canonImpl_visit___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: +{ +lean_object* x_10; lean_object* x_11; uint8_t x_12; +x_10 = lean_ctor_get(x_7, 5); +x_11 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_2, x_5, x_6, x_7, x_8, x_9); +x_12 = !lean_is_exclusive(x_11); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_13 = lean_ctor_get(x_11, 0); +x_14 = lean_ctor_get(x_11, 1); +x_15 = lean_st_ref_take(x_8, x_14); +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_16, 3); +lean_inc(x_17); +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, 1); +x_20 = lean_ctor_get(x_15, 0); +lean_dec(x_20); +x_21 = !lean_is_exclusive(x_16); +if (x_21 == 0) +{ +lean_object* x_22; uint8_t x_23; +x_22 = lean_ctor_get(x_16, 3); +lean_dec(x_22); +x_23 = !lean_is_exclusive(x_17); +if (x_23 == 0) +{ +lean_object* x_24; double 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; +x_24 = lean_ctor_get(x_17, 0); +x_25 = l_Lean_addTrace___at_Lean_Meta_Grind_Canon_canonElemCore___spec__5___closed__1; +x_26 = 0; +x_27 = l_Lean_addTrace___at_Lean_Meta_Grind_Canon_canonElemCore___spec__5___closed__2; +x_28 = lean_alloc_ctor(0, 2, 17); +lean_ctor_set(x_28, 0, x_1); +lean_ctor_set(x_28, 1, x_27); +lean_ctor_set_float(x_28, sizeof(void*)*2, x_25); +lean_ctor_set_float(x_28, sizeof(void*)*2 + 8, x_25); +lean_ctor_set_uint8(x_28, sizeof(void*)*2 + 16, x_26); +x_29 = l_Lean_addTrace___at_Lean_Meta_Grind_Canon_canonElemCore___spec__5___closed__3; +x_30 = lean_alloc_ctor(9, 3, 0); +lean_ctor_set(x_30, 0, x_28); +lean_ctor_set(x_30, 1, x_13); +lean_ctor_set(x_30, 2, x_29); +lean_inc(x_10); +lean_ctor_set(x_15, 1, x_30); +lean_ctor_set(x_15, 0, x_10); +x_31 = l_Lean_PersistentArray_push___rarg(x_24, x_15); +lean_ctor_set(x_17, 0, x_31); +x_32 = lean_st_ref_set(x_8, x_16, x_19); +x_33 = !lean_is_exclusive(x_32); +if (x_33 == 0) +{ +lean_object* x_34; lean_object* x_35; +x_34 = lean_ctor_get(x_32, 0); +lean_dec(x_34); +x_35 = lean_box(0); +lean_ctor_set(x_11, 1, x_4); +lean_ctor_set(x_11, 0, x_35); +lean_ctor_set(x_32, 0, x_11); +return x_32; +} +else +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_32, 1); +lean_inc(x_36); +lean_dec(x_32); +x_37 = lean_box(0); +lean_ctor_set(x_11, 1, x_4); +lean_ctor_set(x_11, 0, x_37); +x_38 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_38, 0, x_11); +lean_ctor_set(x_38, 1, x_36); +return x_38; +} +} +else +{ +uint64_t x_39; lean_object* x_40; double 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; +x_39 = lean_ctor_get_uint64(x_17, sizeof(void*)*1); +x_40 = lean_ctor_get(x_17, 0); +lean_inc(x_40); +lean_dec(x_17); +x_41 = l_Lean_addTrace___at_Lean_Meta_Grind_Canon_canonElemCore___spec__5___closed__1; +x_42 = 0; +x_43 = l_Lean_addTrace___at_Lean_Meta_Grind_Canon_canonElemCore___spec__5___closed__2; +x_44 = lean_alloc_ctor(0, 2, 17); +lean_ctor_set(x_44, 0, x_1); +lean_ctor_set(x_44, 1, x_43); +lean_ctor_set_float(x_44, sizeof(void*)*2, x_41); +lean_ctor_set_float(x_44, sizeof(void*)*2 + 8, x_41); +lean_ctor_set_uint8(x_44, sizeof(void*)*2 + 16, x_42); +x_45 = l_Lean_addTrace___at_Lean_Meta_Grind_Canon_canonElemCore___spec__5___closed__3; +x_46 = lean_alloc_ctor(9, 3, 0); +lean_ctor_set(x_46, 0, x_44); +lean_ctor_set(x_46, 1, x_13); +lean_ctor_set(x_46, 2, x_45); +lean_inc(x_10); +lean_ctor_set(x_15, 1, x_46); +lean_ctor_set(x_15, 0, x_10); +x_47 = l_Lean_PersistentArray_push___rarg(x_40, x_15); +x_48 = lean_alloc_ctor(0, 1, 8); +lean_ctor_set(x_48, 0, x_47); +lean_ctor_set_uint64(x_48, sizeof(void*)*1, x_39); +lean_ctor_set(x_16, 3, x_48); +x_49 = lean_st_ref_set(x_8, x_16, x_19); +x_50 = lean_ctor_get(x_49, 1); +lean_inc(x_50); +if (lean_is_exclusive(x_49)) { + lean_ctor_release(x_49, 0); + lean_ctor_release(x_49, 1); + x_51 = x_49; +} else { + lean_dec_ref(x_49); + x_51 = lean_box(0); +} +x_52 = lean_box(0); +lean_ctor_set(x_11, 1, x_4); +lean_ctor_set(x_11, 0, x_52); +if (lean_is_scalar(x_51)) { + x_53 = lean_alloc_ctor(0, 2, 0); +} else { + x_53 = x_51; +} +lean_ctor_set(x_53, 0, x_11); +lean_ctor_set(x_53, 1, x_50); +return x_53; +} +} +else +{ +lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; uint64_t x_61; lean_object* x_62; lean_object* x_63; double x_64; uint8_t 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_54 = lean_ctor_get(x_16, 0); +x_55 = lean_ctor_get(x_16, 1); +x_56 = lean_ctor_get(x_16, 2); +x_57 = lean_ctor_get(x_16, 4); +x_58 = lean_ctor_get(x_16, 5); +x_59 = lean_ctor_get(x_16, 6); +x_60 = lean_ctor_get(x_16, 7); +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_dec(x_16); +x_61 = lean_ctor_get_uint64(x_17, sizeof(void*)*1); +x_62 = lean_ctor_get(x_17, 0); +lean_inc(x_62); +if (lean_is_exclusive(x_17)) { + lean_ctor_release(x_17, 0); + x_63 = x_17; +} else { + lean_dec_ref(x_17); + x_63 = lean_box(0); +} +x_64 = l_Lean_addTrace___at_Lean_Meta_Grind_Canon_canonElemCore___spec__5___closed__1; +x_65 = 0; +x_66 = l_Lean_addTrace___at_Lean_Meta_Grind_Canon_canonElemCore___spec__5___closed__2; +x_67 = lean_alloc_ctor(0, 2, 17); +lean_ctor_set(x_67, 0, x_1); +lean_ctor_set(x_67, 1, x_66); +lean_ctor_set_float(x_67, sizeof(void*)*2, x_64); +lean_ctor_set_float(x_67, sizeof(void*)*2 + 8, x_64); +lean_ctor_set_uint8(x_67, sizeof(void*)*2 + 16, x_65); +x_68 = l_Lean_addTrace___at_Lean_Meta_Grind_Canon_canonElemCore___spec__5___closed__3; +x_69 = lean_alloc_ctor(9, 3, 0); +lean_ctor_set(x_69, 0, x_67); +lean_ctor_set(x_69, 1, x_13); +lean_ctor_set(x_69, 2, x_68); +lean_inc(x_10); +lean_ctor_set(x_15, 1, x_69); +lean_ctor_set(x_15, 0, x_10); +x_70 = l_Lean_PersistentArray_push___rarg(x_62, x_15); +if (lean_is_scalar(x_63)) { + x_71 = lean_alloc_ctor(0, 1, 8); +} else { + x_71 = x_63; +} +lean_ctor_set(x_71, 0, x_70); +lean_ctor_set_uint64(x_71, sizeof(void*)*1, x_61); +x_72 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_72, 0, x_54); +lean_ctor_set(x_72, 1, x_55); +lean_ctor_set(x_72, 2, x_56); +lean_ctor_set(x_72, 3, x_71); +lean_ctor_set(x_72, 4, x_57); +lean_ctor_set(x_72, 5, x_58); +lean_ctor_set(x_72, 6, x_59); +lean_ctor_set(x_72, 7, x_60); +x_73 = lean_st_ref_set(x_8, x_72, x_19); +x_74 = lean_ctor_get(x_73, 1); +lean_inc(x_74); +if (lean_is_exclusive(x_73)) { + lean_ctor_release(x_73, 0); + lean_ctor_release(x_73, 1); + x_75 = x_73; +} else { + lean_dec_ref(x_73); + x_75 = lean_box(0); +} +x_76 = lean_box(0); +lean_ctor_set(x_11, 1, x_4); +lean_ctor_set(x_11, 0, x_76); +if (lean_is_scalar(x_75)) { + x_77 = lean_alloc_ctor(0, 2, 0); +} else { + x_77 = x_75; +} +lean_ctor_set(x_77, 0, x_11); +lean_ctor_set(x_77, 1, x_74); +return x_77; +} +} +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; uint64_t x_87; lean_object* x_88; lean_object* x_89; double 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; +x_78 = lean_ctor_get(x_15, 1); +lean_inc(x_78); +lean_dec(x_15); +x_79 = lean_ctor_get(x_16, 0); +lean_inc(x_79); +x_80 = lean_ctor_get(x_16, 1); +lean_inc(x_80); +x_81 = lean_ctor_get(x_16, 2); +lean_inc(x_81); +x_82 = lean_ctor_get(x_16, 4); +lean_inc(x_82); +x_83 = lean_ctor_get(x_16, 5); +lean_inc(x_83); +x_84 = lean_ctor_get(x_16, 6); +lean_inc(x_84); +x_85 = lean_ctor_get(x_16, 7); +lean_inc(x_85); +if (lean_is_exclusive(x_16)) { + lean_ctor_release(x_16, 0); + lean_ctor_release(x_16, 1); + lean_ctor_release(x_16, 2); + lean_ctor_release(x_16, 3); + lean_ctor_release(x_16, 4); + lean_ctor_release(x_16, 5); + lean_ctor_release(x_16, 6); + lean_ctor_release(x_16, 7); + x_86 = x_16; +} else { + lean_dec_ref(x_16); + x_86 = lean_box(0); +} +x_87 = lean_ctor_get_uint64(x_17, sizeof(void*)*1); +x_88 = lean_ctor_get(x_17, 0); +lean_inc(x_88); +if (lean_is_exclusive(x_17)) { + lean_ctor_release(x_17, 0); + x_89 = x_17; +} else { + lean_dec_ref(x_17); + x_89 = lean_box(0); +} +x_90 = l_Lean_addTrace___at_Lean_Meta_Grind_Canon_canonElemCore___spec__5___closed__1; +x_91 = 0; +x_92 = l_Lean_addTrace___at_Lean_Meta_Grind_Canon_canonElemCore___spec__5___closed__2; +x_93 = lean_alloc_ctor(0, 2, 17); +lean_ctor_set(x_93, 0, x_1); +lean_ctor_set(x_93, 1, x_92); +lean_ctor_set_float(x_93, sizeof(void*)*2, x_90); +lean_ctor_set_float(x_93, sizeof(void*)*2 + 8, x_90); +lean_ctor_set_uint8(x_93, sizeof(void*)*2 + 16, x_91); +x_94 = l_Lean_addTrace___at_Lean_Meta_Grind_Canon_canonElemCore___spec__5___closed__3; +x_95 = lean_alloc_ctor(9, 3, 0); +lean_ctor_set(x_95, 0, x_93); +lean_ctor_set(x_95, 1, x_13); +lean_ctor_set(x_95, 2, x_94); +lean_inc(x_10); +x_96 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_96, 0, x_10); +lean_ctor_set(x_96, 1, x_95); +x_97 = l_Lean_PersistentArray_push___rarg(x_88, x_96); +if (lean_is_scalar(x_89)) { + x_98 = lean_alloc_ctor(0, 1, 8); +} else { + x_98 = x_89; +} +lean_ctor_set(x_98, 0, x_97); +lean_ctor_set_uint64(x_98, sizeof(void*)*1, x_87); +if (lean_is_scalar(x_86)) { + x_99 = lean_alloc_ctor(0, 8, 0); +} else { + x_99 = x_86; +} +lean_ctor_set(x_99, 0, x_79); +lean_ctor_set(x_99, 1, x_80); +lean_ctor_set(x_99, 2, x_81); +lean_ctor_set(x_99, 3, x_98); +lean_ctor_set(x_99, 4, x_82); +lean_ctor_set(x_99, 5, x_83); +lean_ctor_set(x_99, 6, x_84); +lean_ctor_set(x_99, 7, x_85); +x_100 = lean_st_ref_set(x_8, x_99, x_78); +x_101 = lean_ctor_get(x_100, 1); +lean_inc(x_101); +if (lean_is_exclusive(x_100)) { + lean_ctor_release(x_100, 0); + lean_ctor_release(x_100, 1); + x_102 = x_100; +} else { + lean_dec_ref(x_100); + x_102 = lean_box(0); +} +x_103 = lean_box(0); +lean_ctor_set(x_11, 1, x_4); +lean_ctor_set(x_11, 0, x_103); +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_11); +lean_ctor_set(x_104, 1, x_101); +return x_104; +} +} +else +{ +lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; uint64_t x_120; lean_object* x_121; lean_object* x_122; double 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_105 = lean_ctor_get(x_11, 0); +x_106 = lean_ctor_get(x_11, 1); +lean_inc(x_106); +lean_inc(x_105); +lean_dec(x_11); +x_107 = lean_st_ref_take(x_8, x_106); +x_108 = lean_ctor_get(x_107, 0); +lean_inc(x_108); +x_109 = lean_ctor_get(x_108, 3); +lean_inc(x_109); +x_110 = lean_ctor_get(x_107, 1); +lean_inc(x_110); +if (lean_is_exclusive(x_107)) { + lean_ctor_release(x_107, 0); + lean_ctor_release(x_107, 1); + x_111 = x_107; +} else { + lean_dec_ref(x_107); + x_111 = lean_box(0); +} +x_112 = lean_ctor_get(x_108, 0); +lean_inc(x_112); +x_113 = lean_ctor_get(x_108, 1); +lean_inc(x_113); +x_114 = lean_ctor_get(x_108, 2); +lean_inc(x_114); +x_115 = lean_ctor_get(x_108, 4); +lean_inc(x_115); +x_116 = lean_ctor_get(x_108, 5); +lean_inc(x_116); +x_117 = lean_ctor_get(x_108, 6); +lean_inc(x_117); +x_118 = lean_ctor_get(x_108, 7); +lean_inc(x_118); +if (lean_is_exclusive(x_108)) { + lean_ctor_release(x_108, 0); + lean_ctor_release(x_108, 1); + lean_ctor_release(x_108, 2); + lean_ctor_release(x_108, 3); + lean_ctor_release(x_108, 4); + lean_ctor_release(x_108, 5); + lean_ctor_release(x_108, 6); + lean_ctor_release(x_108, 7); + x_119 = x_108; +} else { + lean_dec_ref(x_108); + x_119 = lean_box(0); +} +x_120 = lean_ctor_get_uint64(x_109, sizeof(void*)*1); +x_121 = lean_ctor_get(x_109, 0); +lean_inc(x_121); +if (lean_is_exclusive(x_109)) { + lean_ctor_release(x_109, 0); + x_122 = x_109; +} else { + lean_dec_ref(x_109); + x_122 = lean_box(0); +} +x_123 = l_Lean_addTrace___at_Lean_Meta_Grind_Canon_canonElemCore___spec__5___closed__1; +x_124 = 0; +x_125 = l_Lean_addTrace___at_Lean_Meta_Grind_Canon_canonElemCore___spec__5___closed__2; +x_126 = lean_alloc_ctor(0, 2, 17); +lean_ctor_set(x_126, 0, x_1); +lean_ctor_set(x_126, 1, x_125); +lean_ctor_set_float(x_126, sizeof(void*)*2, x_123); +lean_ctor_set_float(x_126, sizeof(void*)*2 + 8, x_123); +lean_ctor_set_uint8(x_126, sizeof(void*)*2 + 16, x_124); +x_127 = l_Lean_addTrace___at_Lean_Meta_Grind_Canon_canonElemCore___spec__5___closed__3; +x_128 = lean_alloc_ctor(9, 3, 0); +lean_ctor_set(x_128, 0, x_126); +lean_ctor_set(x_128, 1, x_105); +lean_ctor_set(x_128, 2, x_127); +lean_inc(x_10); +if (lean_is_scalar(x_111)) { + x_129 = lean_alloc_ctor(0, 2, 0); +} else { + x_129 = x_111; +} +lean_ctor_set(x_129, 0, x_10); +lean_ctor_set(x_129, 1, x_128); +x_130 = l_Lean_PersistentArray_push___rarg(x_121, x_129); +if (lean_is_scalar(x_122)) { + x_131 = lean_alloc_ctor(0, 1, 8); +} else { + x_131 = x_122; +} +lean_ctor_set(x_131, 0, x_130); +lean_ctor_set_uint64(x_131, sizeof(void*)*1, x_120); +if (lean_is_scalar(x_119)) { + x_132 = lean_alloc_ctor(0, 8, 0); +} else { + x_132 = x_119; +} +lean_ctor_set(x_132, 0, x_112); +lean_ctor_set(x_132, 1, x_113); +lean_ctor_set(x_132, 2, x_114); +lean_ctor_set(x_132, 3, x_131); +lean_ctor_set(x_132, 4, x_115); +lean_ctor_set(x_132, 5, x_116); +lean_ctor_set(x_132, 6, x_117); +lean_ctor_set(x_132, 7, x_118); +x_133 = lean_st_ref_set(x_8, x_132, x_110); +x_134 = lean_ctor_get(x_133, 1); +lean_inc(x_134); +if (lean_is_exclusive(x_133)) { + lean_ctor_release(x_133, 0); + lean_ctor_release(x_133, 1); + x_135 = x_133; +} else { + lean_dec_ref(x_133); + x_135 = lean_box(0); +} +x_136 = lean_box(0); +x_137 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_137, 0, x_136); +lean_ctor_set(x_137, 1, x_4); +if (lean_is_scalar(x_135)) { + x_138 = lean_alloc_ctor(0, 2, 0); +} else { + x_138 = x_135; +} +lean_ctor_set(x_138, 0, x_137); +lean_ctor_set(x_138, 1, x_134); +return x_138; +} +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +size_t x_13; size_t x_14; uint8_t x_15; +x_13 = lean_ptr_addr(x_1); +x_14 = lean_ptr_addr(x_5); +x_15 = lean_usize_dec_eq(x_13, x_14); +if (x_15 == 0) +{ +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; +x_16 = lean_array_fset(x_3, x_2, x_5); +x_17 = 1; +x_18 = lean_box(x_17); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_16); +lean_ctor_set(x_19, 1, x_18); +x_20 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_20, 0, 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_7); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_12); +return x_22; +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +lean_dec(x_5); +x_23 = lean_box(x_4); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_3); +lean_ctor_set(x_24, 1, x_23); +x_25 = 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_7); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_12); +return x_27; +} +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, uint8_t x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +lean_object* x_15; +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_3); +x_15 = l_Lean_Meta_Grind_Canon_shouldCanon(x_1, x_2, x_3, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_15) == 0) +{ +lean_object* x_16; uint8_t x_17; +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +x_17 = lean_unbox(x_16); +lean_dec(x_16); +switch (x_17) { +case 0: +{ +lean_object* x_18; lean_object* x_19; uint64_t 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; uint8_t x_28; +x_18 = lean_ctor_get(x_10, 0); +lean_inc(x_18); +x_19 = lean_ctor_get(x_15, 1); +lean_inc(x_19); +lean_dec(x_15); +x_20 = lean_ctor_get_uint64(x_10, sizeof(void*)*7); +x_21 = lean_ctor_get_uint8(x_10, sizeof(void*)*7 + 8); +x_22 = lean_ctor_get(x_10, 1); +lean_inc(x_22); +x_23 = lean_ctor_get(x_10, 2); +lean_inc(x_23); +x_24 = lean_ctor_get(x_10, 3); +lean_inc(x_24); +x_25 = lean_ctor_get(x_10, 4); +lean_inc(x_25); +x_26 = lean_ctor_get(x_10, 5); +lean_inc(x_26); +x_27 = lean_ctor_get(x_10, 6); +lean_inc(x_27); +x_28 = !lean_is_exclusive(x_18); +if (x_28 == 0) +{ +uint8_t x_29; uint8_t x_30; uint8_t x_31; uint64_t x_32; uint64_t x_33; uint64_t x_34; uint64_t x_35; uint64_t x_36; lean_object* x_37; uint8_t x_38; lean_object* x_39; +x_29 = lean_ctor_get_uint8(x_10, sizeof(void*)*7 + 9); +x_30 = lean_ctor_get_uint8(x_10, sizeof(void*)*7 + 10); +x_31 = 1; +lean_ctor_set_uint8(x_18, 9, x_31); +x_32 = 2; +x_33 = lean_uint64_shift_right(x_20, x_32); +x_34 = lean_uint64_shift_left(x_33, x_32); +x_35 = l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1___closed__11; +x_36 = lean_uint64_lor(x_34, x_35); +x_37 = lean_alloc_ctor(0, 7, 11); +lean_ctor_set(x_37, 0, x_18); +lean_ctor_set(x_37, 1, x_22); +lean_ctor_set(x_37, 2, x_23); +lean_ctor_set(x_37, 3, x_24); +lean_ctor_set(x_37, 4, x_25); +lean_ctor_set(x_37, 5, x_26); +lean_ctor_set(x_37, 6, x_27); +lean_ctor_set_uint64(x_37, sizeof(void*)*7, x_36); +lean_ctor_set_uint8(x_37, sizeof(void*)*7 + 8, x_21); +lean_ctor_set_uint8(x_37, sizeof(void*)*7 + 9, x_29); +lean_ctor_set_uint8(x_37, sizeof(void*)*7 + 10, x_30); +x_38 = 1; +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_3); +lean_inc(x_2); +x_39 = l_Lean_Meta_Grind_Canon_canonElemCore(x_4, x_2, x_3, x_38, x_9, x_37, x_11, x_12, x_13, x_19); +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; +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_ctor_get(x_40, 0); +lean_inc(x_42); +x_43 = lean_ctor_get(x_40, 1); +lean_inc(x_43); +lean_dec(x_40); +x_44 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10___lambda__1(x_3, x_2, x_5, x_6, x_42, x_8, x_43, x_10, x_11, x_12, x_13, x_41); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_8); +lean_dec(x_2); +lean_dec(x_3); +return x_44; +} +else +{ +uint8_t x_45; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_8); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +x_45 = !lean_is_exclusive(x_39); +if (x_45 == 0) +{ +return x_39; +} +else +{ +lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_46 = lean_ctor_get(x_39, 0); +x_47 = lean_ctor_get(x_39, 1); +lean_inc(x_47); +lean_inc(x_46); +lean_dec(x_39); +x_48 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_48, 0, x_46); +lean_ctor_set(x_48, 1, x_47); +return x_48; +} +} +} +else +{ +uint8_t x_49; uint8_t x_50; uint8_t x_51; uint8_t x_52; uint8_t x_53; uint8_t x_54; uint8_t x_55; uint8_t x_56; uint8_t x_57; uint8_t x_58; uint8_t x_59; uint8_t x_60; uint8_t x_61; uint8_t x_62; uint8_t x_63; uint8_t x_64; uint8_t x_65; uint8_t x_66; uint8_t x_67; lean_object* x_68; uint64_t x_69; uint64_t x_70; uint64_t x_71; uint64_t x_72; uint64_t x_73; lean_object* x_74; uint8_t x_75; lean_object* x_76; +x_49 = lean_ctor_get_uint8(x_10, sizeof(void*)*7 + 9); +x_50 = lean_ctor_get_uint8(x_10, sizeof(void*)*7 + 10); +x_51 = lean_ctor_get_uint8(x_18, 0); +x_52 = lean_ctor_get_uint8(x_18, 1); +x_53 = lean_ctor_get_uint8(x_18, 2); +x_54 = lean_ctor_get_uint8(x_18, 3); +x_55 = lean_ctor_get_uint8(x_18, 4); +x_56 = lean_ctor_get_uint8(x_18, 5); +x_57 = lean_ctor_get_uint8(x_18, 6); +x_58 = lean_ctor_get_uint8(x_18, 7); +x_59 = lean_ctor_get_uint8(x_18, 8); +x_60 = lean_ctor_get_uint8(x_18, 10); +x_61 = lean_ctor_get_uint8(x_18, 11); +x_62 = lean_ctor_get_uint8(x_18, 12); +x_63 = lean_ctor_get_uint8(x_18, 13); +x_64 = lean_ctor_get_uint8(x_18, 14); +x_65 = lean_ctor_get_uint8(x_18, 15); +x_66 = lean_ctor_get_uint8(x_18, 16); +lean_dec(x_18); +x_67 = 1; +x_68 = lean_alloc_ctor(0, 0, 17); +lean_ctor_set_uint8(x_68, 0, x_51); +lean_ctor_set_uint8(x_68, 1, x_52); +lean_ctor_set_uint8(x_68, 2, x_53); +lean_ctor_set_uint8(x_68, 3, x_54); +lean_ctor_set_uint8(x_68, 4, x_55); +lean_ctor_set_uint8(x_68, 5, x_56); +lean_ctor_set_uint8(x_68, 6, x_57); +lean_ctor_set_uint8(x_68, 7, x_58); +lean_ctor_set_uint8(x_68, 8, x_59); +lean_ctor_set_uint8(x_68, 9, x_67); +lean_ctor_set_uint8(x_68, 10, x_60); +lean_ctor_set_uint8(x_68, 11, x_61); +lean_ctor_set_uint8(x_68, 12, x_62); +lean_ctor_set_uint8(x_68, 13, x_63); +lean_ctor_set_uint8(x_68, 14, x_64); +lean_ctor_set_uint8(x_68, 15, x_65); +lean_ctor_set_uint8(x_68, 16, x_66); +x_69 = 2; +x_70 = lean_uint64_shift_right(x_20, x_69); +x_71 = lean_uint64_shift_left(x_70, x_69); +x_72 = l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1___closed__11; +x_73 = lean_uint64_lor(x_71, x_72); +x_74 = lean_alloc_ctor(0, 7, 11); +lean_ctor_set(x_74, 0, x_68); +lean_ctor_set(x_74, 1, x_22); +lean_ctor_set(x_74, 2, x_23); +lean_ctor_set(x_74, 3, x_24); +lean_ctor_set(x_74, 4, x_25); +lean_ctor_set(x_74, 5, x_26); +lean_ctor_set(x_74, 6, x_27); +lean_ctor_set_uint64(x_74, sizeof(void*)*7, x_73); +lean_ctor_set_uint8(x_74, sizeof(void*)*7 + 8, x_21); +lean_ctor_set_uint8(x_74, sizeof(void*)*7 + 9, x_49); +lean_ctor_set_uint8(x_74, sizeof(void*)*7 + 10, x_50); +x_75 = 1; +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_3); +lean_inc(x_2); +x_76 = l_Lean_Meta_Grind_Canon_canonElemCore(x_4, x_2, x_3, x_75, x_9, x_74, x_11, x_12, x_13, x_19); +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; +x_77 = lean_ctor_get(x_76, 0); +lean_inc(x_77); +x_78 = lean_ctor_get(x_76, 1); +lean_inc(x_78); +lean_dec(x_76); +x_79 = lean_ctor_get(x_77, 0); +lean_inc(x_79); +x_80 = lean_ctor_get(x_77, 1); +lean_inc(x_80); +lean_dec(x_77); +x_81 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10___lambda__1(x_3, x_2, x_5, x_6, x_79, x_8, x_80, x_10, x_11, x_12, x_13, x_78); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_8); +lean_dec(x_2); +lean_dec(x_3); +return x_81; +} +else +{ +lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_8); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +x_82 = lean_ctor_get(x_76, 0); +lean_inc(x_82); +x_83 = lean_ctor_get(x_76, 1); +lean_inc(x_83); +if (lean_is_exclusive(x_76)) { + lean_ctor_release(x_76, 0); + lean_ctor_release(x_76, 1); + x_84 = x_76; +} else { + lean_dec_ref(x_76); + x_84 = lean_box(0); +} +if (lean_is_scalar(x_84)) { + x_85 = lean_alloc_ctor(1, 2, 0); +} else { + x_85 = x_84; +} +lean_ctor_set(x_85, 0, x_82); +lean_ctor_set(x_85, 1, x_83); +return x_85; +} +} +} +case 1: +{ +lean_object* x_86; lean_object* x_87; uint64_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; uint8_t x_96; +x_86 = lean_ctor_get(x_10, 0); +lean_inc(x_86); +x_87 = lean_ctor_get(x_15, 1); +lean_inc(x_87); +lean_dec(x_15); +x_88 = lean_ctor_get_uint64(x_10, sizeof(void*)*7); +x_89 = lean_ctor_get_uint8(x_10, sizeof(void*)*7 + 8); +x_90 = lean_ctor_get(x_10, 1); +lean_inc(x_90); +x_91 = lean_ctor_get(x_10, 2); +lean_inc(x_91); +x_92 = lean_ctor_get(x_10, 3); +lean_inc(x_92); +x_93 = lean_ctor_get(x_10, 4); +lean_inc(x_93); +x_94 = lean_ctor_get(x_10, 5); +lean_inc(x_94); +x_95 = lean_ctor_get(x_10, 6); +lean_inc(x_95); +x_96 = !lean_is_exclusive(x_86); +if (x_96 == 0) +{ +uint8_t x_97; uint8_t x_98; uint8_t x_99; uint64_t x_100; uint64_t x_101; uint64_t x_102; uint64_t x_103; uint64_t x_104; lean_object* x_105; uint8_t x_106; lean_object* x_107; +x_97 = lean_ctor_get_uint8(x_10, sizeof(void*)*7 + 9); +x_98 = lean_ctor_get_uint8(x_10, sizeof(void*)*7 + 10); +x_99 = 3; +lean_ctor_set_uint8(x_86, 9, x_99); +x_100 = 2; +x_101 = lean_uint64_shift_right(x_88, x_100); +x_102 = lean_uint64_shift_left(x_101, x_100); +x_103 = l_Lean_Meta_Grind_Canon_canonInst___closed__1; +x_104 = lean_uint64_lor(x_102, x_103); +x_105 = lean_alloc_ctor(0, 7, 11); +lean_ctor_set(x_105, 0, x_86); +lean_ctor_set(x_105, 1, x_90); +lean_ctor_set(x_105, 2, x_91); +lean_ctor_set(x_105, 3, x_92); +lean_ctor_set(x_105, 4, x_93); +lean_ctor_set(x_105, 5, x_94); +lean_ctor_set(x_105, 6, x_95); +lean_ctor_set_uint64(x_105, sizeof(void*)*7, x_104); +lean_ctor_set_uint8(x_105, sizeof(void*)*7 + 8, x_89); +lean_ctor_set_uint8(x_105, sizeof(void*)*7 + 9, x_97); +lean_ctor_set_uint8(x_105, sizeof(void*)*7 + 10, x_98); +x_106 = 0; +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_3); +lean_inc(x_2); +x_107 = l_Lean_Meta_Grind_Canon_canonElemCore(x_4, x_2, x_3, x_106, x_9, x_105, x_11, x_12, x_13, x_87); +if (lean_obj_tag(x_107) == 0) +{ +lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; +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_108, 1); +lean_inc(x_111); +lean_dec(x_108); +x_112 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10___lambda__1(x_3, x_2, x_5, x_6, x_110, x_8, x_111, x_10, x_11, x_12, x_13, x_109); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_8); +lean_dec(x_2); +lean_dec(x_3); +return x_112; +} +else +{ +uint8_t x_113; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_8); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +x_113 = !lean_is_exclusive(x_107); +if (x_113 == 0) +{ +return x_107; +} +else +{ +lean_object* x_114; lean_object* x_115; lean_object* x_116; +x_114 = lean_ctor_get(x_107, 0); +x_115 = lean_ctor_get(x_107, 1); +lean_inc(x_115); +lean_inc(x_114); +lean_dec(x_107); +x_116 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_116, 0, x_114); +lean_ctor_set(x_116, 1, x_115); +return x_116; +} +} +} +else +{ +uint8_t x_117; uint8_t x_118; uint8_t x_119; uint8_t x_120; uint8_t x_121; uint8_t x_122; uint8_t x_123; uint8_t x_124; uint8_t x_125; uint8_t x_126; uint8_t x_127; uint8_t x_128; uint8_t x_129; uint8_t x_130; uint8_t x_131; uint8_t x_132; uint8_t x_133; uint8_t x_134; uint8_t x_135; lean_object* x_136; uint64_t x_137; uint64_t x_138; uint64_t x_139; uint64_t x_140; uint64_t x_141; lean_object* x_142; uint8_t x_143; lean_object* x_144; +x_117 = lean_ctor_get_uint8(x_10, sizeof(void*)*7 + 9); +x_118 = lean_ctor_get_uint8(x_10, sizeof(void*)*7 + 10); +x_119 = lean_ctor_get_uint8(x_86, 0); +x_120 = lean_ctor_get_uint8(x_86, 1); +x_121 = lean_ctor_get_uint8(x_86, 2); +x_122 = lean_ctor_get_uint8(x_86, 3); +x_123 = lean_ctor_get_uint8(x_86, 4); +x_124 = lean_ctor_get_uint8(x_86, 5); +x_125 = lean_ctor_get_uint8(x_86, 6); +x_126 = lean_ctor_get_uint8(x_86, 7); +x_127 = lean_ctor_get_uint8(x_86, 8); +x_128 = lean_ctor_get_uint8(x_86, 10); +x_129 = lean_ctor_get_uint8(x_86, 11); +x_130 = lean_ctor_get_uint8(x_86, 12); +x_131 = lean_ctor_get_uint8(x_86, 13); +x_132 = lean_ctor_get_uint8(x_86, 14); +x_133 = lean_ctor_get_uint8(x_86, 15); +x_134 = lean_ctor_get_uint8(x_86, 16); +lean_dec(x_86); +x_135 = 3; +x_136 = lean_alloc_ctor(0, 0, 17); +lean_ctor_set_uint8(x_136, 0, x_119); +lean_ctor_set_uint8(x_136, 1, x_120); +lean_ctor_set_uint8(x_136, 2, x_121); +lean_ctor_set_uint8(x_136, 3, x_122); +lean_ctor_set_uint8(x_136, 4, x_123); +lean_ctor_set_uint8(x_136, 5, x_124); +lean_ctor_set_uint8(x_136, 6, x_125); +lean_ctor_set_uint8(x_136, 7, x_126); +lean_ctor_set_uint8(x_136, 8, x_127); +lean_ctor_set_uint8(x_136, 9, x_135); +lean_ctor_set_uint8(x_136, 10, x_128); +lean_ctor_set_uint8(x_136, 11, x_129); +lean_ctor_set_uint8(x_136, 12, x_130); +lean_ctor_set_uint8(x_136, 13, x_131); +lean_ctor_set_uint8(x_136, 14, x_132); +lean_ctor_set_uint8(x_136, 15, x_133); +lean_ctor_set_uint8(x_136, 16, x_134); +x_137 = 2; +x_138 = lean_uint64_shift_right(x_88, x_137); +x_139 = lean_uint64_shift_left(x_138, x_137); +x_140 = l_Lean_Meta_Grind_Canon_canonInst___closed__1; +x_141 = lean_uint64_lor(x_139, x_140); +x_142 = lean_alloc_ctor(0, 7, 11); +lean_ctor_set(x_142, 0, x_136); +lean_ctor_set(x_142, 1, x_90); +lean_ctor_set(x_142, 2, x_91); +lean_ctor_set(x_142, 3, x_92); +lean_ctor_set(x_142, 4, x_93); +lean_ctor_set(x_142, 5, x_94); +lean_ctor_set(x_142, 6, x_95); +lean_ctor_set_uint64(x_142, sizeof(void*)*7, x_141); +lean_ctor_set_uint8(x_142, sizeof(void*)*7 + 8, x_89); +lean_ctor_set_uint8(x_142, sizeof(void*)*7 + 9, x_117); +lean_ctor_set_uint8(x_142, sizeof(void*)*7 + 10, x_118); +x_143 = 0; +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_3); +lean_inc(x_2); +x_144 = l_Lean_Meta_Grind_Canon_canonElemCore(x_4, x_2, x_3, x_143, x_9, x_142, x_11, x_12, x_13, x_87); +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; +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); +x_148 = lean_ctor_get(x_145, 1); +lean_inc(x_148); +lean_dec(x_145); +x_149 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10___lambda__1(x_3, x_2, x_5, x_6, x_147, x_8, x_148, x_10, x_11, x_12, x_13, x_146); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_8); +lean_dec(x_2); +lean_dec(x_3); +return x_149; +} +else +{ +lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_8); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +x_150 = lean_ctor_get(x_144, 0); +lean_inc(x_150); +x_151 = lean_ctor_get(x_144, 1); +lean_inc(x_151); +if (lean_is_exclusive(x_144)) { + lean_ctor_release(x_144, 0); + lean_ctor_release(x_144, 1); + x_152 = x_144; +} else { + lean_dec_ref(x_144); + x_152 = lean_box(0); +} +if (lean_is_scalar(x_152)) { + x_153 = lean_alloc_ctor(1, 2, 0); +} else { + x_153 = x_152; +} +lean_ctor_set(x_153, 0, x_150); +lean_ctor_set(x_153, 1, x_151); +return x_153; +} +} +} +case 2: +{ +lean_object* x_154; lean_object* x_155; +x_154 = lean_ctor_get(x_15, 1); +lean_inc(x_154); +lean_dec(x_15); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_8); +lean_inc(x_3); +x_155 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_3, x_8, x_9, x_10, x_11, x_12, x_13, x_154); +if (lean_obj_tag(x_155) == 0) +{ +lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; uint64_t x_161; uint8_t x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; uint8_t x_169; +x_156 = lean_ctor_get(x_155, 0); +lean_inc(x_156); +x_157 = lean_ctor_get(x_10, 0); +lean_inc(x_157); +x_158 = lean_ctor_get(x_155, 1); +lean_inc(x_158); +lean_dec(x_155); +x_159 = lean_ctor_get(x_156, 0); +lean_inc(x_159); +x_160 = lean_ctor_get(x_156, 1); +lean_inc(x_160); +lean_dec(x_156); +x_161 = lean_ctor_get_uint64(x_10, sizeof(void*)*7); +x_162 = lean_ctor_get_uint8(x_10, sizeof(void*)*7 + 8); +x_163 = lean_ctor_get(x_10, 1); +lean_inc(x_163); +x_164 = lean_ctor_get(x_10, 2); +lean_inc(x_164); +x_165 = lean_ctor_get(x_10, 3); +lean_inc(x_165); +x_166 = lean_ctor_get(x_10, 4); +lean_inc(x_166); +x_167 = lean_ctor_get(x_10, 5); +lean_inc(x_167); +x_168 = lean_ctor_get(x_10, 6); +lean_inc(x_168); +x_169 = !lean_is_exclusive(x_157); +if (x_169 == 0) +{ +uint8_t x_170; uint8_t x_171; uint8_t x_172; uint64_t x_173; uint64_t x_174; uint64_t x_175; uint64_t x_176; uint64_t x_177; lean_object* x_178; uint8_t x_179; lean_object* x_180; +x_170 = lean_ctor_get_uint8(x_10, sizeof(void*)*7 + 9); +x_171 = lean_ctor_get_uint8(x_10, sizeof(void*)*7 + 10); +x_172 = 2; +lean_ctor_set_uint8(x_157, 9, x_172); +x_173 = 2; +x_174 = lean_uint64_shift_right(x_161, x_173); +x_175 = lean_uint64_shift_left(x_174, x_173); +x_176 = l_Lean_Meta_Grind_Canon_canonImplicit___closed__1; +x_177 = lean_uint64_lor(x_175, x_176); +x_178 = lean_alloc_ctor(0, 7, 11); +lean_ctor_set(x_178, 0, x_157); +lean_ctor_set(x_178, 1, x_163); +lean_ctor_set(x_178, 2, x_164); +lean_ctor_set(x_178, 3, x_165); +lean_ctor_set(x_178, 4, x_166); +lean_ctor_set(x_178, 5, x_167); +lean_ctor_set(x_178, 6, x_168); +lean_ctor_set_uint64(x_178, sizeof(void*)*7, x_177); +lean_ctor_set_uint8(x_178, sizeof(void*)*7 + 8, x_162); +lean_ctor_set_uint8(x_178, sizeof(void*)*7 + 9, x_170); +lean_ctor_set_uint8(x_178, sizeof(void*)*7 + 10, x_171); +x_179 = 2; +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_2); +x_180 = l_Lean_Meta_Grind_Canon_canonElemCore(x_4, x_2, x_159, x_179, x_160, x_178, x_11, x_12, x_13, x_158); +if (lean_obj_tag(x_180) == 0) +{ +lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; +x_181 = lean_ctor_get(x_180, 0); +lean_inc(x_181); +x_182 = lean_ctor_get(x_180, 1); +lean_inc(x_182); +lean_dec(x_180); +x_183 = lean_ctor_get(x_181, 0); +lean_inc(x_183); +x_184 = lean_ctor_get(x_181, 1); +lean_inc(x_184); +lean_dec(x_181); +x_185 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10___lambda__1(x_3, x_2, x_5, x_6, x_183, x_8, x_184, x_10, x_11, x_12, x_13, x_182); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_8); +lean_dec(x_2); +lean_dec(x_3); +return x_185; +} +else +{ +uint8_t x_186; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_8); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +x_186 = !lean_is_exclusive(x_180); +if (x_186 == 0) +{ +return x_180; } +else +{ +lean_object* x_187; lean_object* x_188; lean_object* x_189; +x_187 = lean_ctor_get(x_180, 0); +x_188 = lean_ctor_get(x_180, 1); +lean_inc(x_188); +lean_inc(x_187); +lean_dec(x_180); +x_189 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_189, 0, x_187); +lean_ctor_set(x_189, 1, x_188); +return x_189; } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_shouldCanon___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: +} +} +else { -lean_object* x_9; -x_9 = l_Lean_Meta_Grind_Canon_shouldCanon(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +uint8_t x_190; uint8_t x_191; uint8_t x_192; uint8_t x_193; uint8_t x_194; uint8_t x_195; uint8_t x_196; uint8_t x_197; uint8_t x_198; uint8_t x_199; uint8_t x_200; uint8_t x_201; uint8_t x_202; uint8_t x_203; uint8_t x_204; uint8_t x_205; uint8_t x_206; uint8_t x_207; uint8_t x_208; lean_object* x_209; uint64_t x_210; uint64_t x_211; uint64_t x_212; uint64_t x_213; uint64_t x_214; lean_object* x_215; uint8_t x_216; lean_object* x_217; +x_190 = lean_ctor_get_uint8(x_10, sizeof(void*)*7 + 9); +x_191 = lean_ctor_get_uint8(x_10, sizeof(void*)*7 + 10); +x_192 = lean_ctor_get_uint8(x_157, 0); +x_193 = lean_ctor_get_uint8(x_157, 1); +x_194 = lean_ctor_get_uint8(x_157, 2); +x_195 = lean_ctor_get_uint8(x_157, 3); +x_196 = lean_ctor_get_uint8(x_157, 4); +x_197 = lean_ctor_get_uint8(x_157, 5); +x_198 = lean_ctor_get_uint8(x_157, 6); +x_199 = lean_ctor_get_uint8(x_157, 7); +x_200 = lean_ctor_get_uint8(x_157, 8); +x_201 = lean_ctor_get_uint8(x_157, 10); +x_202 = lean_ctor_get_uint8(x_157, 11); +x_203 = lean_ctor_get_uint8(x_157, 12); +x_204 = lean_ctor_get_uint8(x_157, 13); +x_205 = lean_ctor_get_uint8(x_157, 14); +x_206 = lean_ctor_get_uint8(x_157, 15); +x_207 = lean_ctor_get_uint8(x_157, 16); +lean_dec(x_157); +x_208 = 2; +x_209 = lean_alloc_ctor(0, 0, 17); +lean_ctor_set_uint8(x_209, 0, x_192); +lean_ctor_set_uint8(x_209, 1, x_193); +lean_ctor_set_uint8(x_209, 2, x_194); +lean_ctor_set_uint8(x_209, 3, x_195); +lean_ctor_set_uint8(x_209, 4, x_196); +lean_ctor_set_uint8(x_209, 5, x_197); +lean_ctor_set_uint8(x_209, 6, x_198); +lean_ctor_set_uint8(x_209, 7, x_199); +lean_ctor_set_uint8(x_209, 8, x_200); +lean_ctor_set_uint8(x_209, 9, x_208); +lean_ctor_set_uint8(x_209, 10, x_201); +lean_ctor_set_uint8(x_209, 11, x_202); +lean_ctor_set_uint8(x_209, 12, x_203); +lean_ctor_set_uint8(x_209, 13, x_204); +lean_ctor_set_uint8(x_209, 14, x_205); +lean_ctor_set_uint8(x_209, 15, x_206); +lean_ctor_set_uint8(x_209, 16, x_207); +x_210 = 2; +x_211 = lean_uint64_shift_right(x_161, x_210); +x_212 = lean_uint64_shift_left(x_211, x_210); +x_213 = l_Lean_Meta_Grind_Canon_canonImplicit___closed__1; +x_214 = lean_uint64_lor(x_212, x_213); +x_215 = lean_alloc_ctor(0, 7, 11); +lean_ctor_set(x_215, 0, x_209); +lean_ctor_set(x_215, 1, x_163); +lean_ctor_set(x_215, 2, x_164); +lean_ctor_set(x_215, 3, x_165); +lean_ctor_set(x_215, 4, x_166); +lean_ctor_set(x_215, 5, x_167); +lean_ctor_set(x_215, 6, x_168); +lean_ctor_set_uint64(x_215, sizeof(void*)*7, x_214); +lean_ctor_set_uint8(x_215, sizeof(void*)*7 + 8, x_162); +lean_ctor_set_uint8(x_215, sizeof(void*)*7 + 9, x_190); +lean_ctor_set_uint8(x_215, sizeof(void*)*7 + 10, x_191); +x_216 = 2; +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_2); +x_217 = l_Lean_Meta_Grind_Canon_canonElemCore(x_4, x_2, x_159, x_216, x_160, x_215, x_11, x_12, x_13, x_158); +if (lean_obj_tag(x_217) == 0) +{ +lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; +x_218 = lean_ctor_get(x_217, 0); +lean_inc(x_218); +x_219 = lean_ctor_get(x_217, 1); +lean_inc(x_219); +lean_dec(x_217); +x_220 = lean_ctor_get(x_218, 0); +lean_inc(x_220); +x_221 = lean_ctor_get(x_218, 1); +lean_inc(x_221); +lean_dec(x_218); +x_222 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10___lambda__1(x_3, x_2, x_5, x_6, x_220, x_8, x_221, x_10, x_11, x_12, x_13, x_219); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_8); lean_dec(x_2); -lean_dec(x_1); -return x_9; +lean_dec(x_3); +return x_222; +} +else +{ +lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_8); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +x_223 = lean_ctor_get(x_217, 0); +lean_inc(x_223); +x_224 = lean_ctor_get(x_217, 1); +lean_inc(x_224); +if (lean_is_exclusive(x_217)) { + lean_ctor_release(x_217, 0); + lean_ctor_release(x_217, 1); + x_225 = x_217; +} else { + lean_dec_ref(x_217); + x_225 = lean_box(0); +} +if (lean_is_scalar(x_225)) { + x_226 = lean_alloc_ctor(1, 2, 0); +} else { + x_226 = x_225; +} +lean_ctor_set(x_226, 0, x_223); +lean_ctor_set(x_226, 1, x_224); +return x_226; +} +} +} +else +{ +uint8_t x_227; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_8); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_227 = !lean_is_exclusive(x_155); +if (x_227 == 0) +{ +return x_155; +} +else +{ +lean_object* x_228; lean_object* x_229; lean_object* x_230; +x_228 = lean_ctor_get(x_155, 0); +x_229 = lean_ctor_get(x_155, 1); +lean_inc(x_229); +lean_inc(x_228); +lean_dec(x_155); +x_230 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_230, 0, x_228); +lean_ctor_set(x_230, 1, x_229); +return x_230; +} +} +} +default: +{ +lean_object* x_231; lean_object* x_232; +lean_dec(x_4); +x_231 = lean_ctor_get(x_15, 1); +lean_inc(x_231); +lean_dec(x_15); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_8); +lean_inc(x_3); +x_232 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_3, x_8, x_9, x_10, x_11, x_12, x_13, x_231); +if (lean_obj_tag(x_232) == 0) +{ +lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; +x_233 = lean_ctor_get(x_232, 0); +lean_inc(x_233); +x_234 = lean_ctor_get(x_232, 1); +lean_inc(x_234); +lean_dec(x_232); +x_235 = lean_ctor_get(x_233, 0); +lean_inc(x_235); +x_236 = lean_ctor_get(x_233, 1); +lean_inc(x_236); +lean_dec(x_233); +x_237 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10___lambda__1(x_3, x_2, x_5, x_6, x_235, x_8, x_236, x_10, x_11, x_12, x_13, x_234); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_8); +lean_dec(x_2); +lean_dec(x_3); +return x_237; +} +else +{ +uint8_t x_238; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_8); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +x_238 = !lean_is_exclusive(x_232); +if (x_238 == 0) +{ +return x_232; +} +else +{ +lean_object* x_239; lean_object* x_240; lean_object* x_241; +x_239 = lean_ctor_get(x_232, 0); +x_240 = lean_ctor_get(x_232, 1); +lean_inc(x_240); +lean_inc(x_239); +lean_dec(x_232); +x_241 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_241, 0, x_239); +lean_ctor_set(x_241, 1, x_240); +return x_241; +} +} +} +} +} +else +{ +uint8_t x_242; +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); +lean_dec(x_3); +lean_dec(x_2); +x_242 = !lean_is_exclusive(x_15); +if (x_242 == 0) +{ +return x_15; +} +else +{ +lean_object* x_243; lean_object* x_244; lean_object* x_245; +x_243 = lean_ctor_get(x_15, 0); +x_244 = lean_ctor_get(x_15, 1); +lean_inc(x_244); +lean_inc(x_243); +lean_dec(x_15); +x_245 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_245, 0, x_243); +lean_ctor_set(x_245, 1, x_244); +return x_245; +} +} } } -static lean_object* _init_l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__1___closed__1() { +static lean_object* _init_l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10___lambda__3___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("[", 1, 1); +return x_1; +} +} +static lean_object* _init_l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10___lambda__3___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_instMonadMetaM; -x_2 = l_StateT_instMonad___rarg(x_1); +x_1 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10___lambda__3___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10___lambda__3___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("]: ", 3, 3); +return x_1; +} +} +static lean_object* _init_l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10___lambda__3___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10___lambda__3___closed__3; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10___lambda__3___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(" : ", 3, 3); +return x_1; +} +} +static lean_object* _init_l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10___lambda__3___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10___lambda__3___closed__5; +x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__1___closed__2() { -_start: +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10___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; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; +x_12 = lean_ctor_get(x_1, 0); +lean_inc(x_12); +x_13 = lean_ctor_get(x_1, 1); +lean_inc(x_13); +lean_dec(x_1); +x_14 = lean_array_fget(x_12, x_2); +x_15 = l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___closed__3; +x_16 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8(x_15, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +x_19 = lean_unbox(x_18); +lean_dec(x_18); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; lean_object* x_24; +x_20 = lean_ctor_get(x_16, 1); +lean_inc(x_20); +lean_dec(x_16); +x_21 = lean_ctor_get(x_17, 1); +lean_inc(x_21); +lean_dec(x_17); +x_22 = lean_box(0); +x_23 = lean_unbox(x_13); +lean_dec(x_13); +x_24 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10___lambda__2(x_3, x_2, x_14, x_4, x_12, x_23, x_22, x_5, x_21, x_7, x_8, x_9, x_10, x_20); +return x_24; +} +else +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_25 = lean_ctor_get(x_16, 1); +lean_inc(x_25); +if (lean_is_exclusive(x_16)) { + lean_ctor_release(x_16, 0); + lean_ctor_release(x_16, 1); + x_26 = x_16; +} else { + lean_dec_ref(x_16); + x_26 = lean_box(0); +} +x_27 = lean_ctor_get(x_17, 1); +lean_inc(x_27); +if (lean_is_exclusive(x_17)) { + lean_ctor_release(x_17, 0); + lean_ctor_release(x_17, 1); + x_28 = x_17; +} else { + lean_dec_ref(x_17); + x_28 = lean_box(0); +} +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_14); +x_29 = l_Lean_Meta_Grind_Canon_shouldCanon(x_3, x_2, x_14, x_7, x_8, x_9, x_10, x_25); +if (lean_obj_tag(x_29) == 0) +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_30 = lean_ctor_get(x_29, 0); +lean_inc(x_30); +x_31 = lean_ctor_get(x_29, 1); +lean_inc(x_31); +lean_dec(x_29); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_14); +x_32 = lean_infer_type(x_14, x_7, x_8, x_9, x_10, x_31); +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; uint8_t x_57; +x_33 = lean_ctor_get(x_32, 0); +lean_inc(x_33); +x_34 = lean_ctor_get(x_32, 1); +lean_inc(x_34); +lean_dec(x_32); +lean_inc(x_14); +x_35 = l_Lean_MessageData_ofExpr(x_14); +x_36 = l_Lean_MessageData_ofExpr(x_33); +x_57 = lean_unbox(x_30); +lean_dec(x_30); +switch (x_57) { +case 0: +{ +lean_object* x_58; +x_58 = l_Lean_Meta_Grind_Canon_instReprShouldCanonResult___closed__2; +x_37 = x_58; +goto block_56; +} +case 1: +{ +lean_object* x_59; +x_59 = l_Lean_Meta_Grind_Canon_instReprShouldCanonResult___closed__4; +x_37 = x_59; +goto block_56; +} +case 2: +{ +lean_object* x_60; +x_60 = l_Lean_Meta_Grind_Canon_instReprShouldCanonResult___closed__6; +x_37 = x_60; +goto block_56; +} +default: +{ +lean_object* x_61; +x_61 = l_Lean_Meta_Grind_Canon_instReprShouldCanonResult___closed__8; +x_37 = x_61; +goto block_56; +} +} +block_56: +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; 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; lean_object* x_55; +x_38 = l_Lean_MessageData_ofFormat(x_37); +x_39 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10___lambda__3___closed__2; +if (lean_is_scalar(x_28)) { + x_40 = lean_alloc_ctor(7, 2, 0); +} else { + x_40 = x_28; + lean_ctor_set_tag(x_40, 7); +} +lean_ctor_set(x_40, 0, x_39); +lean_ctor_set(x_40, 1, x_38); +x_41 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10___lambda__3___closed__4; +if (lean_is_scalar(x_26)) { + x_42 = lean_alloc_ctor(7, 2, 0); +} else { + x_42 = x_26; + lean_ctor_set_tag(x_42, 7); +} +lean_ctor_set(x_42, 0, x_40); +lean_ctor_set(x_42, 1, x_41); +x_43 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_43, 0, x_42); +lean_ctor_set(x_43, 1, x_35); +x_44 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10___lambda__3___closed__6; +x_45 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +x_46 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_36); +x_47 = l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1___closed__10; +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_Grind_Canon_canonImpl_visit___spec__9(x_15, x_48, x_5, x_27, x_7, x_8, x_9, x_10, x_34); +x_50 = lean_ctor_get(x_49, 0); +lean_inc(x_50); +x_51 = lean_ctor_get(x_49, 1); +lean_inc(x_51); +lean_dec(x_49); +x_52 = lean_ctor_get(x_50, 0); +lean_inc(x_52); +x_53 = lean_ctor_get(x_50, 1); +lean_inc(x_53); +lean_dec(x_50); +x_54 = lean_unbox(x_13); +lean_dec(x_13); +x_55 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10___lambda__2(x_3, x_2, x_14, x_4, x_12, x_54, x_52, x_5, x_53, x_7, x_8, x_9, x_10, x_51); +lean_dec(x_52); +return x_55; +} +} +else { -lean_object* x_1; lean_object* x_2; -x_1 = l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__1___closed__1; -x_2 = l_ReaderT_instMonad___rarg(x_1); -return x_2; -} -} -static lean_object* _init_l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__1___closed__3() { -_start: +uint8_t x_62; +lean_dec(x_30); +lean_dec(x_28); +lean_dec(x_27); +lean_dec(x_26); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +x_62 = !lean_is_exclusive(x_32); +if (x_62 == 0) { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__1___closed__2; -x_2 = l_Lean_instInhabitedExpr; -x_3 = l_instInhabitedOfMonad___rarg(x_1, x_2); -return x_3; -} +return x_32; } -LEAN_EXPORT lean_object* l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___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_Meta_Grind_Canon_canonImpl_visit___spec__1___closed__3; -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; +lean_object* x_63; lean_object* x_64; lean_object* x_65; +x_63 = lean_ctor_get(x_32, 0); +x_64 = lean_ctor_get(x_32, 1); +lean_inc(x_64); +lean_inc(x_63); +lean_dec(x_32); +x_65 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_65, 0, x_63); +lean_ctor_set(x_65, 1, x_64); +return x_65; } } -LEAN_EXPORT uint8_t l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__2(lean_object* x_1, lean_object* x_2) { -_start: -{ -if (lean_obj_tag(x_2) == 0) -{ -uint8_t x_3; -x_3 = 0; -return x_3; } else { -lean_object* x_4; lean_object* x_5; size_t x_6; size_t x_7; uint8_t x_8; -x_4 = lean_ctor_get(x_2, 0); -x_5 = lean_ctor_get(x_2, 2); -x_6 = lean_ptr_addr(x_4); -x_7 = lean_ptr_addr(x_1); -x_8 = lean_usize_dec_eq(x_6, x_7); -if (x_8 == 0) +uint8_t x_66; +lean_dec(x_28); +lean_dec(x_27); +lean_dec(x_26); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +x_66 = !lean_is_exclusive(x_29); +if (x_66 == 0) { -x_2 = x_5; -goto _start; +return x_29; } else { -uint8_t x_10; -x_10 = 1; -return x_10; +lean_object* x_67; lean_object* x_68; lean_object* x_69; +x_67 = lean_ctor_get(x_29, 0); +x_68 = lean_ctor_get(x_29, 1); +lean_inc(x_68); +lean_inc(x_67); +lean_dec(x_29); +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_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -if (lean_obj_tag(x_3) == 0) -{ -lean_dec(x_1); -return x_2; } -else +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19, lean_object* x_20) { +_start: { -uint8_t x_4; -x_4 = !lean_is_exclusive(x_3); -if (x_4 == 0) +if (lean_obj_tag(x_13) == 0) { -lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint64_t 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; size_t x_16; size_t x_17; size_t x_18; size_t x_19; size_t x_20; lean_object* x_21; lean_object* x_22; -x_5 = lean_ctor_get(x_3, 0); -x_6 = lean_ctor_get(x_3, 2); -x_7 = lean_array_get_size(x_2); -lean_inc(x_1); -lean_inc(x_5); -x_8 = lean_apply_1(x_1, x_5); -x_9 = lean_unbox_uint64(x_8); +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); lean_dec(x_8); -x_10 = 32; -x_11 = lean_uint64_shift_right(x_9, x_10); -x_12 = lean_uint64_xor(x_9, x_11); -x_13 = 16; -x_14 = lean_uint64_shift_right(x_12, x_13); -x_15 = lean_uint64_xor(x_12, x_14); -x_16 = lean_uint64_to_usize(x_15); -x_17 = lean_usize_of_nat(x_7); lean_dec(x_7); -x_18 = 1; -x_19 = lean_usize_sub(x_17, x_18); -x_20 = lean_usize_land(x_16, x_19); -x_21 = lean_array_uget(x_2, x_20); -lean_ctor_set(x_3, 2, x_21); -x_22 = lean_array_uset(x_2, x_20, x_3); -x_2 = x_22; -x_3 = x_6; -goto _start; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +x_21 = lean_ctor_get(x_13, 0); +lean_inc(x_21); +lean_dec(x_13); +x_22 = lean_ctor_get(x_1, 0); +lean_inc(x_22); +lean_dec(x_1); +x_23 = lean_ctor_get(x_22, 1); +lean_inc(x_23); +lean_dec(x_22); +x_24 = lean_apply_9(x_23, lean_box(0), x_21, x_14, x_15, x_16, x_17, x_18, x_19, x_20); +return x_24; } else { -lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; uint64_t x_29; uint64_t x_30; uint64_t x_31; uint64_t x_32; uint64_t x_33; uint64_t x_34; uint64_t x_35; size_t x_36; size_t x_37; size_t x_38; size_t x_39; size_t x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; -x_24 = lean_ctor_get(x_3, 0); -x_25 = lean_ctor_get(x_3, 1); -x_26 = lean_ctor_get(x_3, 2); -lean_inc(x_26); +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_25 = lean_ctor_get(x_13, 0); lean_inc(x_25); -lean_inc(x_24); -lean_dec(x_3); -x_27 = lean_array_get_size(x_2); -lean_inc(x_1); -lean_inc(x_24); -x_28 = lean_apply_1(x_1, x_24); -x_29 = lean_unbox_uint64(x_28); -lean_dec(x_28); -x_30 = 32; -x_31 = lean_uint64_shift_right(x_29, x_30); -x_32 = lean_uint64_xor(x_29, x_31); -x_33 = 16; -x_34 = lean_uint64_shift_right(x_32, x_33); -x_35 = lean_uint64_xor(x_32, x_34); -x_36 = lean_uint64_to_usize(x_35); -x_37 = lean_usize_of_nat(x_27); -lean_dec(x_27); -x_38 = 1; -x_39 = lean_usize_sub(x_37, x_38); -x_40 = lean_usize_land(x_36, x_39); -x_41 = lean_array_uget(x_2, x_40); -x_42 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_42, 0, x_24); -lean_ctor_set(x_42, 1, x_25); -lean_ctor_set(x_42, 2, x_41); -x_43 = lean_array_uset(x_2, x_40, x_42); -x_2 = x_43; -x_3 = x_26; -goto _start; -} +lean_dec(x_13); +x_26 = lean_ctor_get(x_2, 2); +lean_inc(x_26); +x_27 = lean_nat_add(x_3, x_26); +lean_dec(x_26); +x_28 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_2, x_25, x_27, lean_box(0), lean_box(0), x_14, x_15, x_16, x_17, x_18, x_19, x_20); +return x_28; } } } -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__5___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__6(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19, lean_object* x_20, lean_object* x_21, lean_object* x_22) { _start: { -if (lean_obj_tag(x_2) == 0) -{ -return x_1; -} -else -{ -uint8_t x_3; -x_3 = !lean_is_exclusive(x_2); -if (x_3 == 0) +lean_object* x_23; uint8_t x_24; +x_23 = lean_ctor_get(x_11, 1); +lean_inc(x_23); +x_24 = lean_nat_dec_lt(x_13, x_23); +lean_dec(x_23); +if (x_24 == 0) { -lean_object* x_4; lean_object* x_5; lean_object* x_6; size_t x_7; uint64_t x_8; uint64_t 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; size_t x_17; size_t x_18; size_t x_19; size_t x_20; size_t x_21; lean_object* x_22; lean_object* x_23; -x_4 = lean_ctor_get(x_2, 0); -x_5 = lean_ctor_get(x_2, 2); -x_6 = lean_array_get_size(x_1); -x_7 = lean_ptr_addr(x_4); -x_8 = lean_usize_to_uint64(x_7); -x_9 = 11; -x_10 = lean_uint64_mix_hash(x_8, x_9); -x_11 = 32; -x_12 = lean_uint64_shift_right(x_10, x_11); -x_13 = lean_uint64_xor(x_10, x_12); -x_14 = 16; -x_15 = lean_uint64_shift_right(x_13, x_14); -x_16 = lean_uint64_xor(x_13, x_15); -x_17 = lean_uint64_to_usize(x_16); -x_18 = lean_usize_of_nat(x_6); +lean_object* x_25; lean_object* x_26; lean_object* x_27; +lean_dec(x_13); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); lean_dec(x_6); -x_19 = 1; -x_20 = lean_usize_sub(x_18, x_19); -x_21 = lean_usize_land(x_17, x_20); -x_22 = lean_array_uget(x_1, x_21); -lean_ctor_set(x_2, 2, x_22); -x_23 = lean_array_uset(x_1, x_21, x_2); -x_1 = x_23; -x_2 = x_5; -goto _start; +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_25 = lean_ctor_get(x_1, 0); +lean_inc(x_25); +lean_dec(x_1); +x_26 = lean_ctor_get(x_25, 1); +lean_inc(x_26); +lean_dec(x_25); +x_27 = lean_apply_9(x_26, lean_box(0), x_12, x_16, x_17, x_18, x_19, x_20, x_21, x_22); +return x_27; } else { -lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; size_t x_29; uint64_t x_30; uint64_t x_31; uint64_t x_32; uint64_t x_33; uint64_t x_34; uint64_t x_35; uint64_t x_36; uint64_t x_37; uint64_t x_38; size_t x_39; size_t x_40; size_t x_41; size_t x_42; size_t x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; -x_25 = lean_ctor_get(x_2, 0); -x_26 = lean_ctor_get(x_2, 1); -x_27 = lean_ctor_get(x_2, 2); -lean_inc(x_27); -lean_inc(x_26); -lean_inc(x_25); -lean_dec(x_2); -x_28 = lean_array_get_size(x_1); -x_29 = lean_ptr_addr(x_25); -x_30 = lean_usize_to_uint64(x_29); -x_31 = 11; -x_32 = lean_uint64_mix_hash(x_30, x_31); -x_33 = 32; -x_34 = lean_uint64_shift_right(x_32, x_33); -x_35 = lean_uint64_xor(x_32, x_34); -x_36 = 16; -x_37 = lean_uint64_shift_right(x_35, x_36); -x_38 = lean_uint64_xor(x_35, x_37); -x_39 = lean_uint64_to_usize(x_38); -x_40 = lean_usize_of_nat(x_28); -lean_dec(x_28); -x_41 = 1; -x_42 = lean_usize_sub(x_40, x_41); -x_43 = lean_usize_land(x_39, x_42); -x_44 = lean_array_uget(x_1, x_43); -x_45 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_45, 0, x_25); -lean_ctor_set(x_45, 1, x_26); -lean_ctor_set(x_45, 2, x_44); -x_46 = lean_array_uset(x_1, x_43, x_45); -x_1 = x_46; -x_2 = x_27; -goto _start; -} +lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_28 = lean_ctor_get(x_1, 1); +lean_inc(x_28); +lean_inc(x_5); +lean_inc(x_8); +lean_inc(x_13); +x_29 = lean_alloc_closure((void*)(l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10___lambda__3___boxed), 11, 4); +lean_closure_set(x_29, 0, x_12); +lean_closure_set(x_29, 1, x_13); +lean_closure_set(x_29, 2, x_8); +lean_closure_set(x_29, 3, x_5); +x_30 = lean_alloc_closure((void*)(l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10___lambda__4___boxed), 20, 12); +lean_closure_set(x_30, 0, x_1); +lean_closure_set(x_30, 1, x_11); +lean_closure_set(x_30, 2, x_13); +lean_closure_set(x_30, 3, x_2); +lean_closure_set(x_30, 4, x_3); +lean_closure_set(x_30, 5, x_4); +lean_closure_set(x_30, 6, x_5); +lean_closure_set(x_30, 7, x_6); +lean_closure_set(x_30, 8, x_7); +lean_closure_set(x_30, 9, x_8); +lean_closure_set(x_30, 10, x_9); +lean_closure_set(x_30, 11, x_10); +x_31 = lean_apply_11(x_28, lean_box(0), lean_box(0), x_29, x_30, x_16, x_17, x_18, x_19, x_20, x_21, x_22); +return x_31; } } } -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand_go___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__11___closed__1() { _start: { -lean_object* x_4; uint8_t x_5; -x_4 = lean_array_get_size(x_2); -x_5 = lean_nat_dec_lt(x_1, x_4); -lean_dec(x_4); -if (x_5 == 0) +lean_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)); +return x_1; +} +} +static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__11___closed__2() { +_start: { -lean_dec(x_2); -lean_dec(x_1); -return x_3; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean", 4, 4); +return x_1; } -else +} +static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__11___closed__3() { +_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; -x_6 = lean_array_fget(x_2, x_1); -x_7 = lean_box(0); -x_8 = lean_array_fset(x_2, x_1, x_7); -x_9 = l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__5___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__6(x_3, x_6); -x_10 = lean_unsigned_to_nat(1u); -x_11 = lean_nat_add(x_1, x_10); -lean_dec(x_1); -x_1 = x_11; -x_2 = x_8; -x_3 = x_9; -goto _start; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Grind", 5, 5); +return x_1; } } +static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__11___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("nestedProof", 11, 11); +return x_1; +} } -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__3(lean_object* x_1) { +static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__11___closed__5() { _start: { -lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; -x_2 = lean_array_get_size(x_1); -x_3 = lean_unsigned_to_nat(2u); -x_4 = lean_nat_mul(x_2, x_3); -lean_dec(x_2); -x_5 = lean_box(0); -x_6 = lean_mk_array(x_4, x_5); -x_7 = lean_unsigned_to_nat(0u); -x_8 = l_Std_DHashMap_Internal_Raw_u2080_expand_go___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__4(x_7, x_1, x_6); -return x_8; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__11___closed__2; +x_2 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__11___closed__3; +x_3 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__11___closed__4; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; } } -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__11(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { _start: { -if (lean_obj_tag(x_3) == 0) +switch (lean_obj_tag(x_6)) { +case 0: { -lean_object* x_4; -lean_dec(x_2); -lean_dec(x_1); -x_4 = lean_box(0); -return x_4; +lean_object* x_16; lean_object* x_95; lean_object* x_150; uint8_t x_151; +lean_dec(x_8); +x_150 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__11___closed__5; +x_151 = l_Lean_Expr_isConstOf(x_6, x_150); +if (x_151 == 0) +{ +lean_object* x_152; +x_152 = lean_box(0); +x_95 = x_152; +goto block_149; } else { -uint8_t x_5; -x_5 = !lean_is_exclusive(x_3); -if (x_5 == 0) +lean_object* x_153; lean_object* x_154; uint8_t x_155; +x_153 = lean_array_get_size(x_7); +x_154 = lean_unsigned_to_nat(2u); +x_155 = lean_nat_dec_eq(x_153, x_154); +if (x_155 == 0) { -lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x_9; size_t x_10; uint8_t x_11; -x_6 = lean_ctor_get(x_3, 0); -x_7 = lean_ctor_get(x_3, 1); -x_8 = lean_ctor_get(x_3, 2); -x_9 = lean_ptr_addr(x_6); -x_10 = lean_ptr_addr(x_1); -x_11 = lean_usize_dec_eq(x_9, x_10); -if (x_11 == 0) +lean_object* x_156; +lean_dec(x_153); +x_156 = lean_box(0); +x_95 = x_156; +goto block_149; +} +else { -lean_object* x_12; -x_12 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__7(x_1, x_2, x_8); -lean_ctor_set(x_3, 2, x_12); -return x_3; +lean_object* x_157; uint8_t x_158; +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_157 = lean_unsigned_to_nat(0u); +x_158 = lean_nat_dec_lt(x_157, x_153); +lean_dec(x_153); +if (x_158 == 0) +{ +lean_object* x_159; lean_object* x_160; +x_159 = l_Lean_instInhabitedExpr; +x_160 = l_outOfBounds___rarg(x_159); +x_16 = x_160; +goto block_94; } else { -lean_dec(x_7); -lean_dec(x_6); -lean_ctor_set(x_3, 1, x_2); -lean_ctor_set(x_3, 0, x_1); -return x_3; +lean_object* x_161; +x_161 = lean_array_fget(x_7, x_157); +x_16 = x_161; +goto block_94; } } -else +} +block_94: { -lean_object* x_13; lean_object* x_14; lean_object* x_15; size_t x_16; size_t x_17; uint8_t x_18; -x_13 = lean_ctor_get(x_3, 0); -x_14 = lean_ctor_get(x_3, 1); -x_15 = lean_ctor_get(x_3, 2); -lean_inc(x_15); -lean_inc(x_14); -lean_inc(x_13); -lean_dec(x_3); -x_16 = lean_ptr_addr(x_13); -x_17 = lean_ptr_addr(x_1); -x_18 = lean_usize_dec_eq(x_16, x_17); +lean_object* x_17; +lean_inc(x_16); +x_17 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_16, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +if (lean_obj_tag(x_17) == 0) +{ +uint8_t x_18; +x_18 = !lean_is_exclusive(x_17); if (x_18 == 0) { -lean_object* x_19; lean_object* x_20; -x_19 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__7(x_1, x_2, x_15); -x_20 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_20, 0, x_13); -lean_ctor_set(x_20, 1, x_14); -lean_ctor_set(x_20, 2, x_19); -return x_20; +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_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_21 = lean_ctor_get(x_19, 0); +x_22 = lean_ctor_get(x_19, 1); +x_23 = lean_ctor_get(x_22, 2); +lean_inc(x_23); +lean_inc(x_23); +x_24 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_23, x_21); +if (lean_obj_tag(x_24) == 0) +{ +size_t x_25; size_t x_26; uint8_t x_27; +x_25 = lean_ptr_addr(x_16); +lean_dec(x_16); +x_26 = lean_ptr_addr(x_21); +x_27 = lean_usize_dec_eq(x_25, x_26); +if (x_27 == 0) +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; +lean_dec(x_1); +x_28 = lean_unsigned_to_nat(0u); +lean_inc(x_21); +x_29 = lean_array_set(x_7, x_28, x_21); +x_30 = l_Lean_mkAppN(x_6, x_29); +lean_dec(x_29); +x_31 = lean_ctor_get(x_22, 0); +lean_inc(x_31); +x_32 = lean_ctor_get(x_22, 1); +lean_inc(x_32); +lean_dec(x_22); +lean_inc(x_30); +x_33 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_23, x_21, x_30); +x_34 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_34, 0, x_31); +lean_ctor_set(x_34, 1, x_32); +lean_ctor_set(x_34, 2, x_33); +lean_ctor_set(x_19, 1, x_34); +lean_ctor_set(x_19, 0, x_30); +return x_17; } else { -lean_object* x_21; -lean_dec(x_14); -lean_dec(x_13); -x_21 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_21, 0, x_1); -lean_ctor_set(x_21, 1, x_2); -lean_ctor_set(x_21, 2, x_15); -return x_21; -} +lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; +lean_dec(x_7); +lean_dec(x_6); +x_35 = lean_ctor_get(x_22, 0); +lean_inc(x_35); +x_36 = lean_ctor_get(x_22, 1); +lean_inc(x_36); +lean_dec(x_22); +lean_inc(x_1); +x_37 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_23, x_21, x_1); +x_38 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_38, 0, x_35); +lean_ctor_set(x_38, 1, x_36); +lean_ctor_set(x_38, 2, x_37); +lean_ctor_set(x_19, 1, x_38); +lean_ctor_set(x_19, 0, x_1); +return x_17; } } +else +{ +lean_object* x_39; +lean_dec(x_23); +lean_dec(x_21); +lean_dec(x_16); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_1); +x_39 = lean_ctor_get(x_24, 0); +lean_inc(x_39); +lean_dec(x_24); +lean_ctor_set(x_19, 0, x_39); +return x_17; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: +else { -size_t x_13; size_t x_14; uint8_t x_15; -x_13 = lean_ptr_addr(x_1); -x_14 = lean_ptr_addr(x_5); -x_15 = lean_usize_dec_eq(x_13, x_14); -if (x_15 == 0) +lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_40 = lean_ctor_get(x_19, 0); +x_41 = lean_ctor_get(x_19, 1); +lean_inc(x_41); +lean_inc(x_40); +lean_dec(x_19); +x_42 = lean_ctor_get(x_41, 2); +lean_inc(x_42); +lean_inc(x_42); +x_43 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_42, x_40); +if (lean_obj_tag(x_43) == 0) { -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; -x_16 = lean_array_set(x_3, x_2, x_5); -x_17 = 1; -x_18 = lean_box(x_17); -x_19 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_19, 0, x_16); -lean_ctor_set(x_19, 1, x_18); -x_20 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_20, 0, 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_7); -x_22 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_22, 0, x_21); -lean_ctor_set(x_22, 1, x_12); -return x_22; +size_t x_44; size_t x_45; uint8_t x_46; +x_44 = lean_ptr_addr(x_16); +lean_dec(x_16); +x_45 = lean_ptr_addr(x_40); +x_46 = lean_usize_dec_eq(x_44, x_45); +if (x_46 == 0) +{ +lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; +lean_dec(x_1); +x_47 = lean_unsigned_to_nat(0u); +lean_inc(x_40); +x_48 = lean_array_set(x_7, x_47, x_40); +x_49 = l_Lean_mkAppN(x_6, x_48); +lean_dec(x_48); +x_50 = lean_ctor_get(x_41, 0); +lean_inc(x_50); +x_51 = lean_ctor_get(x_41, 1); +lean_inc(x_51); +lean_dec(x_41); +lean_inc(x_49); +x_52 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_42, x_40, x_49); +x_53 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_53, 0, x_50); +lean_ctor_set(x_53, 1, x_51); +lean_ctor_set(x_53, 2, x_52); +x_54 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_54, 0, x_49); +lean_ctor_set(x_54, 1, x_53); +lean_ctor_set(x_17, 0, x_54); +return x_17; } else { -lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; -lean_dec(x_5); -x_23 = lean_box(x_4); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_3); -lean_ctor_set(x_24, 1, x_23); -x_25 = 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_7); -x_27 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_27, 0, x_26); -lean_ctor_set(x_27, 1, x_12); -return x_27; -} +lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; +lean_dec(x_7); +lean_dec(x_6); +x_55 = lean_ctor_get(x_41, 0); +lean_inc(x_55); +x_56 = lean_ctor_get(x_41, 1); +lean_inc(x_56); +lean_dec(x_41); +lean_inc(x_1); +x_57 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_42, x_40, x_1); +x_58 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_58, 0, x_55); +lean_ctor_set(x_58, 1, x_56); +lean_ctor_set(x_58, 2, x_57); +x_59 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_59, 0, x_1); +lean_ctor_set(x_59, 1, x_58); +lean_ctor_set(x_17, 0, x_59); +return x_17; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___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_ctor_get(x_4, 1); -x_17 = lean_nat_dec_lt(x_6, x_16); -if (x_17 == 0) +else { -lean_object* x_18; lean_object* x_19; -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_9); +lean_object* x_60; lean_object* x_61; +lean_dec(x_42); +lean_dec(x_40); +lean_dec(x_16); +lean_dec(x_7); lean_dec(x_6); lean_dec(x_1); -x_18 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_18, 0, x_5); -lean_ctor_set(x_18, 1, x_10); -x_19 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_19, 0, x_18); -lean_ctor_set(x_19, 1, x_15); -return x_19; +x_60 = lean_ctor_get(x_43, 0); +lean_inc(x_60); +lean_dec(x_43); +x_61 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_61, 0, x_60); +lean_ctor_set(x_61, 1, x_41); +lean_ctor_set(x_17, 0, x_61); +return x_17; +} +} +} +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_17, 0); +x_63 = lean_ctor_get(x_17, 1); +lean_inc(x_63); +lean_inc(x_62); +lean_dec(x_17); +x_64 = lean_ctor_get(x_62, 0); +lean_inc(x_64); +x_65 = lean_ctor_get(x_62, 1); +lean_inc(x_65); +if (lean_is_exclusive(x_62)) { + lean_ctor_release(x_62, 0); + lean_ctor_release(x_62, 1); + x_66 = x_62; +} else { + lean_dec_ref(x_62); + x_66 = lean_box(0); } -else +x_67 = lean_ctor_get(x_65, 2); +lean_inc(x_67); +lean_inc(x_67); +x_68 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_67, x_64); +if (lean_obj_tag(x_68) == 0) { -lean_object* x_20; lean_object* x_21; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_202; uint8_t x_203; -x_29 = lean_ctor_get(x_5, 0); -lean_inc(x_29); -x_30 = lean_ctor_get(x_5, 1); -lean_inc(x_30); -lean_dec(x_5); -x_202 = lean_array_get_size(x_29); -x_203 = lean_nat_dec_lt(x_6, x_202); -lean_dec(x_202); -if (x_203 == 0) +size_t x_69; size_t x_70; uint8_t x_71; +x_69 = lean_ptr_addr(x_16); +lean_dec(x_16); +x_70 = lean_ptr_addr(x_64); +x_71 = lean_usize_dec_eq(x_69, x_70); +if (x_71 == 0) { -lean_object* x_204; lean_object* x_205; -x_204 = l_Lean_instInhabitedExpr; -x_205 = l_outOfBounds___rarg(x_204); -x_31 = x_205; -goto block_201; +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_dec(x_1); +x_72 = lean_unsigned_to_nat(0u); +lean_inc(x_64); +x_73 = lean_array_set(x_7, x_72, x_64); +x_74 = l_Lean_mkAppN(x_6, x_73); +lean_dec(x_73); +x_75 = lean_ctor_get(x_65, 0); +lean_inc(x_75); +x_76 = lean_ctor_get(x_65, 1); +lean_inc(x_76); +lean_dec(x_65); +lean_inc(x_74); +x_77 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_67, x_64, x_74); +x_78 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_78, 0, x_75); +lean_ctor_set(x_78, 1, x_76); +lean_ctor_set(x_78, 2, x_77); +if (lean_is_scalar(x_66)) { + x_79 = lean_alloc_ctor(0, 2, 0); +} else { + x_79 = x_66; +} +lean_ctor_set(x_79, 0, x_74); +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_63); +return x_80; } else { -lean_object* x_206; -x_206 = lean_array_fget(x_29, x_6); -x_31 = x_206; -goto block_201; +lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; +lean_dec(x_7); +lean_dec(x_6); +x_81 = lean_ctor_get(x_65, 0); +lean_inc(x_81); +x_82 = lean_ctor_get(x_65, 1); +lean_inc(x_82); +lean_dec(x_65); +lean_inc(x_1); +x_83 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_67, x_64, x_1); +x_84 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_84, 0, x_81); +lean_ctor_set(x_84, 1, x_82); +lean_ctor_set(x_84, 2, x_83); +if (lean_is_scalar(x_66)) { + x_85 = lean_alloc_ctor(0, 2, 0); +} else { + x_85 = x_66; +} +lean_ctor_set(x_85, 0, x_1); +lean_ctor_set(x_85, 1, x_84); +x_86 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_86, 0, x_85); +lean_ctor_set(x_86, 1, x_63); +return x_86; } -block_28: +} +else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_22 = lean_ctor_get(x_20, 0); -lean_inc(x_22); -x_23 = lean_ctor_get(x_20, 1); -lean_inc(x_23); -lean_dec(x_20); -x_24 = lean_ctor_get(x_22, 0); -lean_inc(x_24); -lean_dec(x_22); -x_25 = lean_ctor_get(x_4, 2); -x_26 = lean_nat_add(x_6, x_25); +lean_object* x_87; lean_object* x_88; lean_object* x_89; +lean_dec(x_67); +lean_dec(x_64); +lean_dec(x_16); +lean_dec(x_7); lean_dec(x_6); -x_5 = x_24; -x_6 = x_26; -x_7 = lean_box(0); -x_8 = lean_box(0); -x_10 = x_23; -x_15 = x_21; -goto _start; +lean_dec(x_1); +x_87 = lean_ctor_get(x_68, 0); +lean_inc(x_87); +lean_dec(x_68); +if (lean_is_scalar(x_66)) { + x_88 = lean_alloc_ctor(0, 2, 0); +} else { + x_88 = x_66; +} +lean_ctor_set(x_88, 0, x_87); +lean_ctor_set(x_88, 1, x_65); +x_89 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_89, 0, x_88); +lean_ctor_set(x_89, 1, x_63); +return x_89; +} } -block_201: -{ -lean_object* x_32; -lean_inc(x_14); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_31); -x_32 = l_Lean_Meta_Grind_Canon_shouldCanon(x_2, x_6, x_31, x_11, x_12, x_13, x_14, x_15); -if (lean_obj_tag(x_32) == 0) -{ -lean_object* x_33; uint8_t x_34; -x_33 = lean_ctor_get(x_32, 0); -lean_inc(x_33); -x_34 = lean_unbox(x_33); -lean_dec(x_33); -switch (x_34) { -case 0: -{ -lean_object* x_35; lean_object* x_36; uint64_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; lean_object* x_44; uint8_t x_45; -x_35 = lean_ctor_get(x_11, 0); -lean_inc(x_35); -x_36 = lean_ctor_get(x_32, 1); -lean_inc(x_36); -lean_dec(x_32); -x_37 = lean_ctor_get_uint64(x_11, sizeof(void*)*7); -x_38 = lean_ctor_get_uint8(x_11, sizeof(void*)*7 + 8); -x_39 = lean_ctor_get(x_11, 1); -lean_inc(x_39); -x_40 = lean_ctor_get(x_11, 2); -lean_inc(x_40); -x_41 = lean_ctor_get(x_11, 3); -lean_inc(x_41); -x_42 = lean_ctor_get(x_11, 4); -lean_inc(x_42); -x_43 = lean_ctor_get(x_11, 5); -lean_inc(x_43); -x_44 = lean_ctor_get(x_11, 6); -lean_inc(x_44); -x_45 = !lean_is_exclusive(x_35); -if (x_45 == 0) -{ -uint8_t x_46; uint8_t x_47; uint8_t x_48; uint64_t x_49; uint64_t x_50; uint64_t x_51; uint64_t x_52; uint64_t x_53; lean_object* x_54; uint8_t x_55; lean_object* x_56; -x_46 = lean_ctor_get_uint8(x_11, sizeof(void*)*7 + 9); -x_47 = lean_ctor_get_uint8(x_11, sizeof(void*)*7 + 10); -x_48 = 1; -lean_ctor_set_uint8(x_35, 9, x_48); -x_49 = 2; -x_50 = lean_uint64_shift_right(x_37, x_49); -x_51 = lean_uint64_shift_left(x_50, x_49); -x_52 = l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1___closed__9; -x_53 = lean_uint64_lor(x_51, x_52); -x_54 = lean_alloc_ctor(0, 7, 11); -lean_ctor_set(x_54, 0, x_35); -lean_ctor_set(x_54, 1, x_39); -lean_ctor_set(x_54, 2, x_40); -lean_ctor_set(x_54, 3, x_41); -lean_ctor_set(x_54, 4, x_42); -lean_ctor_set(x_54, 5, x_43); -lean_ctor_set(x_54, 6, x_44); -lean_ctor_set_uint64(x_54, sizeof(void*)*7, x_53); -lean_ctor_set_uint8(x_54, sizeof(void*)*7 + 8, x_38); -lean_ctor_set_uint8(x_54, sizeof(void*)*7 + 9, x_46); -lean_ctor_set_uint8(x_54, sizeof(void*)*7 + 10, x_47); -x_55 = 0; -lean_inc(x_14); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_31); -lean_inc(x_6); -lean_inc(x_1); -x_56 = l_Lean_Meta_Grind_Canon_canonElemCore(x_1, x_6, x_31, x_55, x_10, x_54, x_12, x_13, x_14, x_36); -if (lean_obj_tag(x_56) == 0) -{ -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; -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_inc(x_59); -x_60 = lean_ctor_get(x_57, 1); -lean_inc(x_60); -lean_dec(x_57); -x_61 = lean_unbox(x_30); -lean_dec(x_30); -x_62 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8___lambda__1(x_31, x_6, x_29, x_61, x_59, x_9, x_60, x_11, x_12, x_13, x_14, x_58); -lean_dec(x_31); -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_20 = x_63; -x_21 = x_64; -goto block_28; } else { -uint8_t x_65; -lean_dec(x_31); -lean_dec(x_30); -lean_dec(x_29); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_9); +uint8_t x_90; +lean_dec(x_16); +lean_dec(x_7); lean_dec(x_6); lean_dec(x_1); -x_65 = !lean_is_exclusive(x_56); -if (x_65 == 0) +x_90 = !lean_is_exclusive(x_17); +if (x_90 == 0) { -return x_56; +return x_17; } else { -lean_object* x_66; lean_object* x_67; lean_object* x_68; -x_66 = lean_ctor_get(x_56, 0); -x_67 = lean_ctor_get(x_56, 1); -lean_inc(x_67); -lean_inc(x_66); -lean_dec(x_56); -x_68 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_68, 0, x_66); -lean_ctor_set(x_68, 1, x_67); -return x_68; -} -} -} -else -{ -uint8_t x_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; uint8_t x_83; uint8_t x_84; uint8_t x_85; uint8_t x_86; uint8_t x_87; lean_object* x_88; uint64_t x_89; uint64_t x_90; uint64_t x_91; uint64_t x_92; uint64_t x_93; lean_object* x_94; uint8_t x_95; lean_object* x_96; -x_69 = lean_ctor_get_uint8(x_11, sizeof(void*)*7 + 9); -x_70 = lean_ctor_get_uint8(x_11, sizeof(void*)*7 + 10); -x_71 = lean_ctor_get_uint8(x_35, 0); -x_72 = lean_ctor_get_uint8(x_35, 1); -x_73 = lean_ctor_get_uint8(x_35, 2); -x_74 = lean_ctor_get_uint8(x_35, 3); -x_75 = lean_ctor_get_uint8(x_35, 4); -x_76 = lean_ctor_get_uint8(x_35, 5); -x_77 = lean_ctor_get_uint8(x_35, 6); -x_78 = lean_ctor_get_uint8(x_35, 7); -x_79 = lean_ctor_get_uint8(x_35, 8); -x_80 = lean_ctor_get_uint8(x_35, 10); -x_81 = lean_ctor_get_uint8(x_35, 11); -x_82 = lean_ctor_get_uint8(x_35, 12); -x_83 = lean_ctor_get_uint8(x_35, 13); -x_84 = lean_ctor_get_uint8(x_35, 14); -x_85 = lean_ctor_get_uint8(x_35, 15); -x_86 = lean_ctor_get_uint8(x_35, 16); -lean_dec(x_35); -x_87 = 1; -x_88 = lean_alloc_ctor(0, 0, 17); -lean_ctor_set_uint8(x_88, 0, x_71); -lean_ctor_set_uint8(x_88, 1, x_72); -lean_ctor_set_uint8(x_88, 2, x_73); -lean_ctor_set_uint8(x_88, 3, x_74); -lean_ctor_set_uint8(x_88, 4, x_75); -lean_ctor_set_uint8(x_88, 5, x_76); -lean_ctor_set_uint8(x_88, 6, x_77); -lean_ctor_set_uint8(x_88, 7, x_78); -lean_ctor_set_uint8(x_88, 8, x_79); -lean_ctor_set_uint8(x_88, 9, x_87); -lean_ctor_set_uint8(x_88, 10, x_80); -lean_ctor_set_uint8(x_88, 11, x_81); -lean_ctor_set_uint8(x_88, 12, x_82); -lean_ctor_set_uint8(x_88, 13, x_83); -lean_ctor_set_uint8(x_88, 14, x_84); -lean_ctor_set_uint8(x_88, 15, x_85); -lean_ctor_set_uint8(x_88, 16, x_86); -x_89 = 2; -x_90 = lean_uint64_shift_right(x_37, x_89); -x_91 = lean_uint64_shift_left(x_90, x_89); -x_92 = l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1___closed__9; -x_93 = lean_uint64_lor(x_91, x_92); -x_94 = lean_alloc_ctor(0, 7, 11); -lean_ctor_set(x_94, 0, x_88); -lean_ctor_set(x_94, 1, x_39); -lean_ctor_set(x_94, 2, x_40); -lean_ctor_set(x_94, 3, x_41); -lean_ctor_set(x_94, 4, x_42); -lean_ctor_set(x_94, 5, x_43); -lean_ctor_set(x_94, 6, x_44); -lean_ctor_set_uint64(x_94, sizeof(void*)*7, x_93); -lean_ctor_set_uint8(x_94, sizeof(void*)*7 + 8, x_38); -lean_ctor_set_uint8(x_94, sizeof(void*)*7 + 9, x_69); -lean_ctor_set_uint8(x_94, sizeof(void*)*7 + 10, x_70); -x_95 = 0; +lean_object* x_91; lean_object* x_92; lean_object* x_93; +x_91 = lean_ctor_get(x_17, 0); +x_92 = lean_ctor_get(x_17, 1); +lean_inc(x_92); +lean_inc(x_91); +lean_dec(x_17); +x_93 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_93, 0, x_91); +lean_ctor_set(x_93, 1, x_92); +return x_93; +} +} +} +block_149: +{ +lean_object* x_96; +lean_dec(x_95); lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); -lean_inc(x_31); +lean_inc(x_11); lean_inc(x_6); -lean_inc(x_1); -x_96 = l_Lean_Meta_Grind_Canon_canonElemCore(x_1, x_6, x_31, x_95, x_10, x_94, x_12, x_13, x_14, x_36); +x_96 = l_Lean_Meta_getFunInfo(x_6, x_11, x_12, x_13, x_14, x_15); if (lean_obj_tag(x_96) == 0) { -lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; uint8_t x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; +lean_object* x_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; uint8_t x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; x_97 = lean_ctor_get(x_96, 0); lean_inc(x_97); x_98 = lean_ctor_get(x_96, 1); @@ -5927,5105 +9132,7744 @@ lean_inc(x_98); lean_dec(x_96); x_99 = lean_ctor_get(x_97, 0); lean_inc(x_99); -x_100 = lean_ctor_get(x_97, 1); -lean_inc(x_100); lean_dec(x_97); -x_101 = lean_unbox(x_30); -lean_dec(x_30); -x_102 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8___lambda__1(x_31, x_6, x_29, x_101, x_99, x_9, x_100, x_11, x_12, x_13, x_14, x_98); -lean_dec(x_31); -x_103 = lean_ctor_get(x_102, 0); +x_100 = lean_array_get_size(x_7); +x_101 = lean_unsigned_to_nat(0u); +x_102 = lean_unsigned_to_nat(1u); +lean_inc(x_100); +x_103 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_103, 0, x_101); +lean_ctor_set(x_103, 1, x_100); +lean_ctor_set(x_103, 2, x_102); +x_104 = 0; +x_105 = lean_box(x_104); +lean_inc(x_7); +x_106 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_106, 0, x_7); +lean_ctor_set(x_106, 1, x_105); +x_107 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__11___closed__1; lean_inc(x_103); -x_104 = lean_ctor_get(x_102, 1); -lean_inc(x_104); -lean_dec(x_102); -x_20 = x_103; -x_21 = x_104; -goto block_28; +lean_inc(x_6); +x_108 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10(x_2, x_3, x_4, x_5, x_6, x_7, x_107, x_99, x_100, x_103, x_103, x_106, x_101, lean_box(0), lean_box(0), x_9, x_10, x_11, x_12, x_13, x_14, x_98); +if (lean_obj_tag(x_108) == 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_109, 0); +lean_inc(x_110); +x_111 = lean_ctor_get(x_110, 1); +lean_inc(x_111); +x_112 = lean_unbox(x_111); +lean_dec(x_111); +if (x_112 == 0) +{ +uint8_t x_113; +lean_dec(x_110); +lean_dec(x_6); +x_113 = !lean_is_exclusive(x_108); +if (x_113 == 0) +{ +lean_object* x_114; uint8_t x_115; +x_114 = lean_ctor_get(x_108, 0); +lean_dec(x_114); +x_115 = !lean_is_exclusive(x_109); +if (x_115 == 0) +{ +lean_object* x_116; +x_116 = lean_ctor_get(x_109, 0); +lean_dec(x_116); +lean_ctor_set(x_109, 0, x_1); +return x_108; } else { -lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; -lean_dec(x_31); -lean_dec(x_30); -lean_dec(x_29); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_9); -lean_dec(x_6); +lean_object* x_117; lean_object* x_118; +x_117 = lean_ctor_get(x_109, 1); +lean_inc(x_117); +lean_dec(x_109); +x_118 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_118, 0, x_1); +lean_ctor_set(x_118, 1, x_117); +lean_ctor_set(x_108, 0, x_118); +return x_108; +} +} +else +{ +lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; +x_119 = lean_ctor_get(x_108, 1); +lean_inc(x_119); +lean_dec(x_108); +x_120 = lean_ctor_get(x_109, 1); +lean_inc(x_120); +if (lean_is_exclusive(x_109)) { + lean_ctor_release(x_109, 0); + lean_ctor_release(x_109, 1); + x_121 = x_109; +} else { + lean_dec_ref(x_109); + x_121 = lean_box(0); +} +if (lean_is_scalar(x_121)) { + x_122 = lean_alloc_ctor(0, 2, 0); +} else { + x_122 = x_121; +} +lean_ctor_set(x_122, 0, x_1); +lean_ctor_set(x_122, 1, x_120); +x_123 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_123, 0, x_122); +lean_ctor_set(x_123, 1, x_119); +return x_123; +} +} +else +{ +uint8_t x_124; lean_dec(x_1); -x_105 = lean_ctor_get(x_96, 0); -lean_inc(x_105); -x_106 = lean_ctor_get(x_96, 1); -lean_inc(x_106); -if (lean_is_exclusive(x_96)) { - lean_ctor_release(x_96, 0); - lean_ctor_release(x_96, 1); - x_107 = x_96; +x_124 = !lean_is_exclusive(x_108); +if (x_124 == 0) +{ +lean_object* x_125; uint8_t x_126; +x_125 = lean_ctor_get(x_108, 0); +lean_dec(x_125); +x_126 = !lean_is_exclusive(x_109); +if (x_126 == 0) +{ +lean_object* x_127; lean_object* x_128; lean_object* x_129; +x_127 = lean_ctor_get(x_109, 0); +lean_dec(x_127); +x_128 = lean_ctor_get(x_110, 0); +lean_inc(x_128); +lean_dec(x_110); +x_129 = l_Lean_mkAppN(x_6, x_128); +lean_dec(x_128); +lean_ctor_set(x_109, 0, x_129); +return x_108; +} +else +{ +lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; +x_130 = lean_ctor_get(x_109, 1); +lean_inc(x_130); +lean_dec(x_109); +x_131 = lean_ctor_get(x_110, 0); +lean_inc(x_131); +lean_dec(x_110); +x_132 = l_Lean_mkAppN(x_6, x_131); +lean_dec(x_131); +x_133 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_133, 0, x_132); +lean_ctor_set(x_133, 1, x_130); +lean_ctor_set(x_108, 0, x_133); +return x_108; +} +} +else +{ +lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; +x_134 = lean_ctor_get(x_108, 1); +lean_inc(x_134); +lean_dec(x_108); +x_135 = lean_ctor_get(x_109, 1); +lean_inc(x_135); +if (lean_is_exclusive(x_109)) { + lean_ctor_release(x_109, 0); + lean_ctor_release(x_109, 1); + x_136 = x_109; } else { - lean_dec_ref(x_96); - x_107 = lean_box(0); + lean_dec_ref(x_109); + x_136 = lean_box(0); } -if (lean_is_scalar(x_107)) { - x_108 = lean_alloc_ctor(1, 2, 0); +x_137 = lean_ctor_get(x_110, 0); +lean_inc(x_137); +lean_dec(x_110); +x_138 = l_Lean_mkAppN(x_6, x_137); +lean_dec(x_137); +if (lean_is_scalar(x_136)) { + x_139 = lean_alloc_ctor(0, 2, 0); } else { - x_108 = x_107; + x_139 = x_136; } -lean_ctor_set(x_108, 0, x_105); -lean_ctor_set(x_108, 1, x_106); +lean_ctor_set(x_139, 0, x_138); +lean_ctor_set(x_139, 1, x_135); +x_140 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_140, 0, x_139); +lean_ctor_set(x_140, 1, x_134); +return x_140; +} +} +} +else +{ +uint8_t x_141; +lean_dec(x_6); +lean_dec(x_1); +x_141 = !lean_is_exclusive(x_108); +if (x_141 == 0) +{ return x_108; } +else +{ +lean_object* x_142; lean_object* x_143; lean_object* x_144; +x_142 = lean_ctor_get(x_108, 0); +x_143 = lean_ctor_get(x_108, 1); +lean_inc(x_143); +lean_inc(x_142); +lean_dec(x_108); +x_144 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_144, 0, x_142); +lean_ctor_set(x_144, 1, x_143); +return x_144; } } -case 1: +} +else { -lean_object* x_109; lean_object* x_110; uint64_t x_111; uint8_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; uint8_t x_119; -x_109 = lean_ctor_get(x_11, 0); -lean_inc(x_109); -x_110 = lean_ctor_get(x_32, 1); -lean_inc(x_110); -lean_dec(x_32); -x_111 = lean_ctor_get_uint64(x_11, sizeof(void*)*7); -x_112 = lean_ctor_get_uint8(x_11, sizeof(void*)*7 + 8); -x_113 = lean_ctor_get(x_11, 1); -lean_inc(x_113); -x_114 = lean_ctor_get(x_11, 2); -lean_inc(x_114); -x_115 = lean_ctor_get(x_11, 3); -lean_inc(x_115); -x_116 = lean_ctor_get(x_11, 4); -lean_inc(x_116); -x_117 = lean_ctor_get(x_11, 5); -lean_inc(x_117); -x_118 = lean_ctor_get(x_11, 6); -lean_inc(x_118); -x_119 = !lean_is_exclusive(x_109); -if (x_119 == 0) -{ -uint8_t x_120; uint8_t x_121; uint8_t x_122; uint64_t x_123; uint64_t x_124; uint64_t x_125; uint64_t x_126; uint64_t x_127; lean_object* x_128; uint8_t x_129; lean_object* x_130; -x_120 = lean_ctor_get_uint8(x_11, sizeof(void*)*7 + 9); -x_121 = lean_ctor_get_uint8(x_11, sizeof(void*)*7 + 10); -x_122 = 3; -lean_ctor_set_uint8(x_109, 9, x_122); -x_123 = 2; -x_124 = lean_uint64_shift_right(x_111, x_123); -x_125 = lean_uint64_shift_left(x_124, x_123); -x_126 = l_Lean_Meta_Grind_Canon_canonInst___closed__1; -x_127 = lean_uint64_lor(x_125, x_126); -x_128 = lean_alloc_ctor(0, 7, 11); -lean_ctor_set(x_128, 0, x_109); -lean_ctor_set(x_128, 1, x_113); -lean_ctor_set(x_128, 2, x_114); -lean_ctor_set(x_128, 3, x_115); -lean_ctor_set(x_128, 4, x_116); -lean_ctor_set(x_128, 5, x_117); -lean_ctor_set(x_128, 6, x_118); -lean_ctor_set_uint64(x_128, sizeof(void*)*7, x_127); -lean_ctor_set_uint8(x_128, sizeof(void*)*7 + 8, x_112); -lean_ctor_set_uint8(x_128, sizeof(void*)*7 + 9, x_120); -lean_ctor_set_uint8(x_128, sizeof(void*)*7 + 10, x_121); -x_129 = 1; -lean_inc(x_14); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_31); -lean_inc(x_6); -lean_inc(x_1); -x_130 = l_Lean_Meta_Grind_Canon_canonElemCore(x_1, x_6, x_31, x_129, x_10, x_128, x_12, x_13, x_14, x_110); -if (lean_obj_tag(x_130) == 0) +uint8_t x_145; +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_145 = !lean_is_exclusive(x_96); +if (x_145 == 0) +{ +return x_96; +} +else +{ +lean_object* x_146; lean_object* x_147; lean_object* x_148; +x_146 = lean_ctor_get(x_96, 0); +x_147 = lean_ctor_get(x_96, 1); +lean_inc(x_147); +lean_inc(x_146); +lean_dec(x_96); +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; +} +} +} +} +case 1: { -lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; uint8_t x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; -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_ctor_get(x_131, 0); -lean_inc(x_133); -x_134 = lean_ctor_get(x_131, 1); -lean_inc(x_134); -lean_dec(x_131); -x_135 = lean_unbox(x_30); -lean_dec(x_30); -x_136 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8___lambda__1(x_31, x_6, x_29, x_135, x_133, x_9, x_134, x_11, x_12, x_13, x_14, x_132); -lean_dec(x_31); -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_20 = x_137; -x_21 = x_138; -goto block_28; +lean_object* x_162; lean_object* x_241; lean_object* x_296; uint8_t x_297; +lean_dec(x_8); +x_296 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__11___closed__5; +x_297 = l_Lean_Expr_isConstOf(x_6, x_296); +if (x_297 == 0) +{ +lean_object* x_298; +x_298 = lean_box(0); +x_241 = x_298; +goto block_295; } else { -uint8_t x_139; -lean_dec(x_31); -lean_dec(x_30); -lean_dec(x_29); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_9); -lean_dec(x_6); -lean_dec(x_1); -x_139 = !lean_is_exclusive(x_130); -if (x_139 == 0) +lean_object* x_299; lean_object* x_300; uint8_t x_301; +x_299 = lean_array_get_size(x_7); +x_300 = lean_unsigned_to_nat(2u); +x_301 = lean_nat_dec_eq(x_299, x_300); +if (x_301 == 0) { -return x_130; +lean_object* x_302; +lean_dec(x_299); +x_302 = lean_box(0); +x_241 = x_302; +goto block_295; } else { -lean_object* x_140; lean_object* x_141; lean_object* x_142; -x_140 = lean_ctor_get(x_130, 0); -x_141 = lean_ctor_get(x_130, 1); -lean_inc(x_141); -lean_inc(x_140); -lean_dec(x_130); -x_142 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_142, 0, x_140); -lean_ctor_set(x_142, 1, x_141); -return x_142; -} -} -} -else -{ -uint8_t x_143; uint8_t x_144; uint8_t x_145; uint8_t x_146; uint8_t x_147; uint8_t x_148; uint8_t x_149; uint8_t 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; lean_object* x_162; uint64_t x_163; uint64_t x_164; uint64_t x_165; uint64_t x_166; uint64_t x_167; lean_object* x_168; uint8_t x_169; lean_object* x_170; -x_143 = lean_ctor_get_uint8(x_11, sizeof(void*)*7 + 9); -x_144 = lean_ctor_get_uint8(x_11, sizeof(void*)*7 + 10); -x_145 = lean_ctor_get_uint8(x_109, 0); -x_146 = lean_ctor_get_uint8(x_109, 1); -x_147 = lean_ctor_get_uint8(x_109, 2); -x_148 = lean_ctor_get_uint8(x_109, 3); -x_149 = lean_ctor_get_uint8(x_109, 4); -x_150 = lean_ctor_get_uint8(x_109, 5); -x_151 = lean_ctor_get_uint8(x_109, 6); -x_152 = lean_ctor_get_uint8(x_109, 7); -x_153 = lean_ctor_get_uint8(x_109, 8); -x_154 = lean_ctor_get_uint8(x_109, 10); -x_155 = lean_ctor_get_uint8(x_109, 11); -x_156 = lean_ctor_get_uint8(x_109, 12); -x_157 = lean_ctor_get_uint8(x_109, 13); -x_158 = lean_ctor_get_uint8(x_109, 14); -x_159 = lean_ctor_get_uint8(x_109, 15); -x_160 = lean_ctor_get_uint8(x_109, 16); -lean_dec(x_109); -x_161 = 3; -x_162 = lean_alloc_ctor(0, 0, 17); -lean_ctor_set_uint8(x_162, 0, x_145); -lean_ctor_set_uint8(x_162, 1, x_146); -lean_ctor_set_uint8(x_162, 2, x_147); -lean_ctor_set_uint8(x_162, 3, x_148); -lean_ctor_set_uint8(x_162, 4, x_149); -lean_ctor_set_uint8(x_162, 5, x_150); -lean_ctor_set_uint8(x_162, 6, x_151); -lean_ctor_set_uint8(x_162, 7, x_152); -lean_ctor_set_uint8(x_162, 8, x_153); -lean_ctor_set_uint8(x_162, 9, x_161); -lean_ctor_set_uint8(x_162, 10, x_154); -lean_ctor_set_uint8(x_162, 11, x_155); -lean_ctor_set_uint8(x_162, 12, x_156); -lean_ctor_set_uint8(x_162, 13, x_157); -lean_ctor_set_uint8(x_162, 14, x_158); -lean_ctor_set_uint8(x_162, 15, x_159); -lean_ctor_set_uint8(x_162, 16, x_160); -x_163 = 2; -x_164 = lean_uint64_shift_right(x_111, x_163); -x_165 = lean_uint64_shift_left(x_164, x_163); -x_166 = l_Lean_Meta_Grind_Canon_canonInst___closed__1; -x_167 = lean_uint64_lor(x_165, x_166); -x_168 = lean_alloc_ctor(0, 7, 11); -lean_ctor_set(x_168, 0, x_162); -lean_ctor_set(x_168, 1, x_113); -lean_ctor_set(x_168, 2, x_114); -lean_ctor_set(x_168, 3, x_115); -lean_ctor_set(x_168, 4, x_116); -lean_ctor_set(x_168, 5, x_117); -lean_ctor_set(x_168, 6, x_118); -lean_ctor_set_uint64(x_168, sizeof(void*)*7, x_167); -lean_ctor_set_uint8(x_168, sizeof(void*)*7 + 8, x_112); -lean_ctor_set_uint8(x_168, sizeof(void*)*7 + 9, x_143); -lean_ctor_set_uint8(x_168, sizeof(void*)*7 + 10, x_144); -x_169 = 1; -lean_inc(x_14); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_31); -lean_inc(x_6); -lean_inc(x_1); -x_170 = l_Lean_Meta_Grind_Canon_canonElemCore(x_1, x_6, x_31, x_169, x_10, x_168, x_12, x_13, x_14, x_110); +lean_object* x_303; uint8_t x_304; +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_303 = lean_unsigned_to_nat(0u); +x_304 = lean_nat_dec_lt(x_303, x_299); +lean_dec(x_299); +if (x_304 == 0) +{ +lean_object* x_305; lean_object* x_306; +x_305 = l_Lean_instInhabitedExpr; +x_306 = l_outOfBounds___rarg(x_305); +x_162 = x_306; +goto block_240; +} +else +{ +lean_object* x_307; +x_307 = lean_array_fget(x_7, x_303); +x_162 = x_307; +goto block_240; +} +} +} +block_240: +{ +lean_object* x_163; +lean_inc(x_162); +x_163 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_162, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +if (lean_obj_tag(x_163) == 0) +{ +uint8_t x_164; +x_164 = !lean_is_exclusive(x_163); +if (x_164 == 0) +{ +lean_object* x_165; uint8_t x_166; +x_165 = lean_ctor_get(x_163, 0); +x_166 = !lean_is_exclusive(x_165); +if (x_166 == 0) +{ +lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; +x_167 = lean_ctor_get(x_165, 0); +x_168 = lean_ctor_get(x_165, 1); +x_169 = lean_ctor_get(x_168, 2); +lean_inc(x_169); +lean_inc(x_169); +x_170 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_169, x_167); if (lean_obj_tag(x_170) == 0) { -lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; uint8_t x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; -x_171 = lean_ctor_get(x_170, 0); -lean_inc(x_171); -x_172 = lean_ctor_get(x_170, 1); -lean_inc(x_172); -lean_dec(x_170); -x_173 = lean_ctor_get(x_171, 0); -lean_inc(x_173); -x_174 = lean_ctor_get(x_171, 1); -lean_inc(x_174); -lean_dec(x_171); -x_175 = lean_unbox(x_30); -lean_dec(x_30); -x_176 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8___lambda__1(x_31, x_6, x_29, x_175, x_173, x_9, x_174, x_11, x_12, x_13, x_14, x_172); -lean_dec(x_31); -x_177 = lean_ctor_get(x_176, 0); +size_t x_171; size_t x_172; uint8_t x_173; +x_171 = lean_ptr_addr(x_162); +lean_dec(x_162); +x_172 = lean_ptr_addr(x_167); +x_173 = lean_usize_dec_eq(x_171, x_172); +if (x_173 == 0) +{ +lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; +lean_dec(x_1); +x_174 = lean_unsigned_to_nat(0u); +lean_inc(x_167); +x_175 = lean_array_set(x_7, x_174, x_167); +x_176 = l_Lean_mkAppN(x_6, x_175); +lean_dec(x_175); +x_177 = lean_ctor_get(x_168, 0); lean_inc(x_177); -x_178 = lean_ctor_get(x_176, 1); +x_178 = lean_ctor_get(x_168, 1); lean_inc(x_178); -lean_dec(x_176); -x_20 = x_177; -x_21 = x_178; -goto block_28; +lean_dec(x_168); +lean_inc(x_176); +x_179 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_169, x_167, x_176); +x_180 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_180, 0, x_177); +lean_ctor_set(x_180, 1, x_178); +lean_ctor_set(x_180, 2, x_179); +lean_ctor_set(x_165, 1, x_180); +lean_ctor_set(x_165, 0, x_176); +return x_163; } else { -lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; -lean_dec(x_31); -lean_dec(x_30); -lean_dec(x_29); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_9); +lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_1); -x_179 = lean_ctor_get(x_170, 0); -lean_inc(x_179); -x_180 = lean_ctor_get(x_170, 1); -lean_inc(x_180); -if (lean_is_exclusive(x_170)) { - lean_ctor_release(x_170, 0); - lean_ctor_release(x_170, 1); - x_181 = x_170; -} else { - lean_dec_ref(x_170); - x_181 = lean_box(0); -} -if (lean_is_scalar(x_181)) { - x_182 = lean_alloc_ctor(1, 2, 0); -} else { - x_182 = x_181; +x_181 = lean_ctor_get(x_168, 0); +lean_inc(x_181); +x_182 = lean_ctor_get(x_168, 1); +lean_inc(x_182); +lean_dec(x_168); +lean_inc(x_1); +x_183 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_169, x_167, x_1); +x_184 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_184, 0, x_181); +lean_ctor_set(x_184, 1, x_182); +lean_ctor_set(x_184, 2, x_183); +lean_ctor_set(x_165, 1, x_184); +lean_ctor_set(x_165, 0, x_1); +return x_163; } -lean_ctor_set(x_182, 0, x_179); -lean_ctor_set(x_182, 1, x_180); -return x_182; } +else +{ +lean_object* x_185; +lean_dec(x_169); +lean_dec(x_167); +lean_dec(x_162); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_1); +x_185 = lean_ctor_get(x_170, 0); +lean_inc(x_185); +lean_dec(x_170); +lean_ctor_set(x_165, 0, x_185); +return x_163; } } -default: -{ -lean_object* x_183; lean_object* x_184; -x_183 = lean_ctor_get(x_32, 1); -lean_inc(x_183); -lean_dec(x_32); -lean_inc(x_14); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_9); -lean_inc(x_31); -x_184 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_31, x_9, x_10, x_11, x_12, x_13, x_14, x_183); -if (lean_obj_tag(x_184) == 0) +else { -lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; uint8_t x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; -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_ctor_get(x_185, 0); +lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; +x_186 = lean_ctor_get(x_165, 0); +x_187 = lean_ctor_get(x_165, 1); lean_inc(x_187); -x_188 = lean_ctor_get(x_185, 1); +lean_inc(x_186); +lean_dec(x_165); +x_188 = lean_ctor_get(x_187, 2); lean_inc(x_188); -lean_dec(x_185); -x_189 = lean_unbox(x_30); -lean_dec(x_30); -x_190 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8___lambda__1(x_31, x_6, x_29, x_189, x_187, x_9, x_188, x_11, x_12, x_13, x_14, x_186); -lean_dec(x_31); -x_191 = lean_ctor_get(x_190, 0); -lean_inc(x_191); -x_192 = lean_ctor_get(x_190, 1); -lean_inc(x_192); -lean_dec(x_190); -x_20 = x_191; -x_21 = x_192; -goto block_28; +lean_inc(x_188); +x_189 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_188, x_186); +if (lean_obj_tag(x_189) == 0) +{ +size_t x_190; size_t x_191; uint8_t x_192; +x_190 = lean_ptr_addr(x_162); +lean_dec(x_162); +x_191 = lean_ptr_addr(x_186); +x_192 = lean_usize_dec_eq(x_190, x_191); +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_dec(x_1); +x_193 = lean_unsigned_to_nat(0u); +lean_inc(x_186); +x_194 = lean_array_set(x_7, x_193, x_186); +x_195 = l_Lean_mkAppN(x_6, x_194); +lean_dec(x_194); +x_196 = lean_ctor_get(x_187, 0); +lean_inc(x_196); +x_197 = lean_ctor_get(x_187, 1); +lean_inc(x_197); +lean_dec(x_187); +lean_inc(x_195); +x_198 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_188, x_186, x_195); +x_199 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_199, 0, x_196); +lean_ctor_set(x_199, 1, x_197); +lean_ctor_set(x_199, 2, x_198); +x_200 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_200, 0, x_195); +lean_ctor_set(x_200, 1, x_199); +lean_ctor_set(x_163, 0, x_200); +return x_163; } else { -uint8_t x_193; -lean_dec(x_31); -lean_dec(x_30); -lean_dec(x_29); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_9); +lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; +lean_dec(x_7); +lean_dec(x_6); +x_201 = lean_ctor_get(x_187, 0); +lean_inc(x_201); +x_202 = lean_ctor_get(x_187, 1); +lean_inc(x_202); +lean_dec(x_187); +lean_inc(x_1); +x_203 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_188, x_186, x_1); +x_204 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_204, 0, x_201); +lean_ctor_set(x_204, 1, x_202); +lean_ctor_set(x_204, 2, x_203); +x_205 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_205, 0, x_1); +lean_ctor_set(x_205, 1, x_204); +lean_ctor_set(x_163, 0, x_205); +return x_163; +} +} +else +{ +lean_object* x_206; lean_object* x_207; +lean_dec(x_188); +lean_dec(x_186); +lean_dec(x_162); +lean_dec(x_7); lean_dec(x_6); lean_dec(x_1); -x_193 = !lean_is_exclusive(x_184); -if (x_193 == 0) +x_206 = lean_ctor_get(x_189, 0); +lean_inc(x_206); +lean_dec(x_189); +x_207 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_207, 0, x_206); +lean_ctor_set(x_207, 1, x_187); +lean_ctor_set(x_163, 0, x_207); +return x_163; +} +} +} +else { -return x_184; +lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; +x_208 = lean_ctor_get(x_163, 0); +x_209 = lean_ctor_get(x_163, 1); +lean_inc(x_209); +lean_inc(x_208); +lean_dec(x_163); +x_210 = lean_ctor_get(x_208, 0); +lean_inc(x_210); +x_211 = lean_ctor_get(x_208, 1); +lean_inc(x_211); +if (lean_is_exclusive(x_208)) { + lean_ctor_release(x_208, 0); + lean_ctor_release(x_208, 1); + x_212 = x_208; +} else { + lean_dec_ref(x_208); + x_212 = lean_box(0); +} +x_213 = lean_ctor_get(x_211, 2); +lean_inc(x_213); +lean_inc(x_213); +x_214 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_213, x_210); +if (lean_obj_tag(x_214) == 0) +{ +size_t x_215; size_t x_216; uint8_t x_217; +x_215 = lean_ptr_addr(x_162); +lean_dec(x_162); +x_216 = lean_ptr_addr(x_210); +x_217 = lean_usize_dec_eq(x_215, x_216); +if (x_217 == 0) +{ +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_dec(x_1); +x_218 = lean_unsigned_to_nat(0u); +lean_inc(x_210); +x_219 = lean_array_set(x_7, x_218, x_210); +x_220 = l_Lean_mkAppN(x_6, x_219); +lean_dec(x_219); +x_221 = lean_ctor_get(x_211, 0); +lean_inc(x_221); +x_222 = lean_ctor_get(x_211, 1); +lean_inc(x_222); +lean_dec(x_211); +lean_inc(x_220); +x_223 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_213, x_210, x_220); +x_224 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_224, 0, x_221); +lean_ctor_set(x_224, 1, x_222); +lean_ctor_set(x_224, 2, x_223); +if (lean_is_scalar(x_212)) { + x_225 = lean_alloc_ctor(0, 2, 0); +} else { + x_225 = x_212; +} +lean_ctor_set(x_225, 0, x_220); +lean_ctor_set(x_225, 1, x_224); +x_226 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_226, 0, x_225); +lean_ctor_set(x_226, 1, x_209); +return x_226; } else { -lean_object* x_194; lean_object* x_195; lean_object* x_196; -x_194 = lean_ctor_get(x_184, 0); -x_195 = lean_ctor_get(x_184, 1); -lean_inc(x_195); -lean_inc(x_194); -lean_dec(x_184); -x_196 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_196, 0, x_194); -lean_ctor_set(x_196, 1, x_195); -return x_196; +lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; +lean_dec(x_7); +lean_dec(x_6); +x_227 = lean_ctor_get(x_211, 0); +lean_inc(x_227); +x_228 = lean_ctor_get(x_211, 1); +lean_inc(x_228); +lean_dec(x_211); +lean_inc(x_1); +x_229 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_213, x_210, x_1); +x_230 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_230, 0, x_227); +lean_ctor_set(x_230, 1, x_228); +lean_ctor_set(x_230, 2, x_229); +if (lean_is_scalar(x_212)) { + x_231 = lean_alloc_ctor(0, 2, 0); +} else { + x_231 = x_212; } +lean_ctor_set(x_231, 0, x_1); +lean_ctor_set(x_231, 1, x_230); +x_232 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_232, 0, x_231); +lean_ctor_set(x_232, 1, x_209); +return x_232; +} +} +else +{ +lean_object* x_233; lean_object* x_234; lean_object* x_235; +lean_dec(x_213); +lean_dec(x_210); +lean_dec(x_162); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_1); +x_233 = lean_ctor_get(x_214, 0); +lean_inc(x_233); +lean_dec(x_214); +if (lean_is_scalar(x_212)) { + x_234 = lean_alloc_ctor(0, 2, 0); +} else { + x_234 = x_212; } +lean_ctor_set(x_234, 0, x_233); +lean_ctor_set(x_234, 1, x_211); +x_235 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_235, 0, x_234); +lean_ctor_set(x_235, 1, x_209); +return x_235; } } } else { -uint8_t x_197; -lean_dec(x_31); -lean_dec(x_30); -lean_dec(x_29); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); +uint8_t x_236; +lean_dec(x_162); +lean_dec(x_7); lean_dec(x_6); lean_dec(x_1); -x_197 = !lean_is_exclusive(x_32); -if (x_197 == 0) +x_236 = !lean_is_exclusive(x_163); +if (x_236 == 0) { -return x_32; +return x_163; } else { -lean_object* x_198; lean_object* x_199; lean_object* x_200; -x_198 = lean_ctor_get(x_32, 0); -x_199 = lean_ctor_get(x_32, 1); -lean_inc(x_199); -lean_inc(x_198); -lean_dec(x_32); -x_200 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_200, 0, x_198); -lean_ctor_set(x_200, 1, x_199); -return x_200; -} +lean_object* x_237; lean_object* x_238; lean_object* x_239; +x_237 = lean_ctor_get(x_163, 0); +x_238 = lean_ctor_get(x_163, 1); +lean_inc(x_238); +lean_inc(x_237); +lean_dec(x_163); +x_239 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_239, 0, x_237); +lean_ctor_set(x_239, 1, x_238); +return x_239; } } } -} -} -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -lean_object* x_10; uint8_t x_11; -x_10 = lean_st_ref_take(x_3, x_9); -x_11 = !lean_is_exclusive(x_10); -if (x_11 == 0) -{ -lean_object* x_12; uint8_t x_13; -x_12 = lean_ctor_get(x_10, 0); -x_13 = !lean_is_exclusive(x_12); -if (x_13 == 0) +block_295: { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; size_t x_18; uint64_t x_19; uint64_t x_20; uint64_t x_21; uint64_t x_22; uint64_t x_23; uint64_t x_24; uint64_t x_25; uint64_t x_26; uint64_t x_27; size_t x_28; size_t x_29; size_t x_30; size_t x_31; size_t x_32; lean_object* x_33; uint8_t x_34; -x_14 = lean_ctor_get(x_10, 1); -x_15 = lean_ctor_get(x_12, 0); -x_16 = lean_ctor_get(x_12, 1); -x_17 = lean_array_get_size(x_16); -x_18 = lean_ptr_addr(x_1); -x_19 = lean_usize_to_uint64(x_18); -x_20 = 11; -x_21 = lean_uint64_mix_hash(x_19, x_20); -x_22 = 32; -x_23 = lean_uint64_shift_right(x_21, x_22); -x_24 = lean_uint64_xor(x_21, x_23); -x_25 = 16; -x_26 = lean_uint64_shift_right(x_24, x_25); -x_27 = lean_uint64_xor(x_24, x_26); -x_28 = lean_uint64_to_usize(x_27); -x_29 = lean_usize_of_nat(x_17); -lean_dec(x_17); -x_30 = 1; -x_31 = lean_usize_sub(x_29, x_30); -x_32 = lean_usize_land(x_28, x_31); -x_33 = lean_array_uget(x_16, x_32); -x_34 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__2(x_1, x_33); -if (x_34 == 0) +lean_object* x_242; +lean_dec(x_241); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_6); +x_242 = l_Lean_Meta_getFunInfo(x_6, x_11, x_12, x_13, x_14, x_15); +if (lean_obj_tag(x_242) == 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; -x_35 = lean_unsigned_to_nat(1u); -x_36 = lean_nat_add(x_15, x_35); -lean_dec(x_15); -lean_inc(x_2); -x_37 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_37, 0, x_1); -lean_ctor_set(x_37, 1, x_2); -lean_ctor_set(x_37, 2, x_33); -x_38 = lean_array_uset(x_16, x_32, x_37); -x_39 = lean_unsigned_to_nat(4u); -x_40 = lean_nat_mul(x_36, x_39); -x_41 = lean_unsigned_to_nat(3u); -x_42 = lean_nat_div(x_40, x_41); -lean_dec(x_40); -x_43 = lean_array_get_size(x_38); -x_44 = lean_nat_dec_le(x_42, x_43); -lean_dec(x_43); -lean_dec(x_42); -if (x_44 == 0) +lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; uint8_t x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; +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 = lean_ctor_get(x_243, 0); +lean_inc(x_245); +lean_dec(x_243); +x_246 = lean_array_get_size(x_7); +x_247 = lean_unsigned_to_nat(0u); +x_248 = lean_unsigned_to_nat(1u); +lean_inc(x_246); +x_249 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_249, 0, x_247); +lean_ctor_set(x_249, 1, x_246); +lean_ctor_set(x_249, 2, x_248); +x_250 = 0; +x_251 = lean_box(x_250); +lean_inc(x_7); +x_252 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_252, 0, x_7); +lean_ctor_set(x_252, 1, x_251); +x_253 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__11___closed__1; +lean_inc(x_249); +lean_inc(x_6); +x_254 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10(x_2, x_3, x_4, x_5, x_6, x_7, x_253, x_245, x_246, x_249, x_249, x_252, x_247, lean_box(0), lean_box(0), x_9, x_10, x_11, x_12, x_13, x_14, x_244); +if (lean_obj_tag(x_254) == 0) { -lean_object* x_45; lean_object* x_46; uint8_t x_47; -x_45 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__3(x_38); -lean_ctor_set(x_12, 1, x_45); -lean_ctor_set(x_12, 0, x_36); -x_46 = lean_st_ref_set(x_3, x_12, x_14); -x_47 = !lean_is_exclusive(x_46); -if (x_47 == 0) +lean_object* x_255; lean_object* x_256; lean_object* x_257; uint8_t x_258; +x_255 = lean_ctor_get(x_254, 0); +lean_inc(x_255); +x_256 = lean_ctor_get(x_255, 0); +lean_inc(x_256); +x_257 = lean_ctor_get(x_256, 1); +lean_inc(x_257); +x_258 = lean_unbox(x_257); +lean_dec(x_257); +if (x_258 == 0) { -lean_object* x_48; -x_48 = lean_ctor_get(x_46, 0); -lean_dec(x_48); -lean_ctor_set(x_10, 1, x_4); -lean_ctor_set(x_10, 0, x_2); -lean_ctor_set(x_46, 0, x_10); -return x_46; +uint8_t x_259; +lean_dec(x_256); +lean_dec(x_6); +x_259 = !lean_is_exclusive(x_254); +if (x_259 == 0) +{ +lean_object* x_260; uint8_t x_261; +x_260 = lean_ctor_get(x_254, 0); +lean_dec(x_260); +x_261 = !lean_is_exclusive(x_255); +if (x_261 == 0) +{ +lean_object* x_262; +x_262 = lean_ctor_get(x_255, 0); +lean_dec(x_262); +lean_ctor_set(x_255, 0, x_1); +return x_254; } else { -lean_object* x_49; lean_object* x_50; -x_49 = lean_ctor_get(x_46, 1); -lean_inc(x_49); -lean_dec(x_46); -lean_ctor_set(x_10, 1, x_4); -lean_ctor_set(x_10, 0, x_2); -x_50 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_50, 0, x_10); -lean_ctor_set(x_50, 1, x_49); -return x_50; +lean_object* x_263; lean_object* x_264; +x_263 = lean_ctor_get(x_255, 1); +lean_inc(x_263); +lean_dec(x_255); +x_264 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_264, 0, x_1); +lean_ctor_set(x_264, 1, x_263); +lean_ctor_set(x_254, 0, x_264); +return x_254; } } else { -lean_object* x_51; uint8_t x_52; -lean_ctor_set(x_12, 1, x_38); -lean_ctor_set(x_12, 0, x_36); -x_51 = lean_st_ref_set(x_3, x_12, x_14); -x_52 = !lean_is_exclusive(x_51); -if (x_52 == 0) +lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; +x_265 = lean_ctor_get(x_254, 1); +lean_inc(x_265); +lean_dec(x_254); +x_266 = lean_ctor_get(x_255, 1); +lean_inc(x_266); +if (lean_is_exclusive(x_255)) { + lean_ctor_release(x_255, 0); + lean_ctor_release(x_255, 1); + x_267 = x_255; +} else { + lean_dec_ref(x_255); + x_267 = lean_box(0); +} +if (lean_is_scalar(x_267)) { + x_268 = lean_alloc_ctor(0, 2, 0); +} else { + x_268 = x_267; +} +lean_ctor_set(x_268, 0, x_1); +lean_ctor_set(x_268, 1, x_266); +x_269 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_269, 0, x_268); +lean_ctor_set(x_269, 1, x_265); +return x_269; +} +} +else { -lean_object* x_53; -x_53 = lean_ctor_get(x_51, 0); -lean_dec(x_53); -lean_ctor_set(x_10, 1, x_4); -lean_ctor_set(x_10, 0, x_2); -lean_ctor_set(x_51, 0, x_10); -return x_51; +uint8_t x_270; +lean_dec(x_1); +x_270 = !lean_is_exclusive(x_254); +if (x_270 == 0) +{ +lean_object* x_271; uint8_t x_272; +x_271 = lean_ctor_get(x_254, 0); +lean_dec(x_271); +x_272 = !lean_is_exclusive(x_255); +if (x_272 == 0) +{ +lean_object* x_273; lean_object* x_274; lean_object* x_275; +x_273 = lean_ctor_get(x_255, 0); +lean_dec(x_273); +x_274 = lean_ctor_get(x_256, 0); +lean_inc(x_274); +lean_dec(x_256); +x_275 = l_Lean_mkAppN(x_6, x_274); +lean_dec(x_274); +lean_ctor_set(x_255, 0, x_275); +return x_254; } else { -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_ctor_set(x_10, 1, x_4); -lean_ctor_set(x_10, 0, x_2); -x_55 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_55, 0, x_10); -lean_ctor_set(x_55, 1, x_54); -return x_55; +lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; +x_276 = lean_ctor_get(x_255, 1); +lean_inc(x_276); +lean_dec(x_255); +x_277 = lean_ctor_get(x_256, 0); +lean_inc(x_277); +lean_dec(x_256); +x_278 = l_Lean_mkAppN(x_6, x_277); +lean_dec(x_277); +x_279 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_279, 0, x_278); +lean_ctor_set(x_279, 1, x_276); +lean_ctor_set(x_254, 0, x_279); +return x_254; +} +} +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_286; +x_280 = lean_ctor_get(x_254, 1); +lean_inc(x_280); +lean_dec(x_254); +x_281 = lean_ctor_get(x_255, 1); +lean_inc(x_281); +if (lean_is_exclusive(x_255)) { + lean_ctor_release(x_255, 0); + lean_ctor_release(x_255, 1); + x_282 = x_255; +} else { + lean_dec_ref(x_255); + x_282 = lean_box(0); +} +x_283 = lean_ctor_get(x_256, 0); +lean_inc(x_283); +lean_dec(x_256); +x_284 = l_Lean_mkAppN(x_6, x_283); +lean_dec(x_283); +if (lean_is_scalar(x_282)) { + x_285 = lean_alloc_ctor(0, 2, 0); +} else { + x_285 = x_282; +} +lean_ctor_set(x_285, 0, x_284); +lean_ctor_set(x_285, 1, x_281); +x_286 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_286, 0, x_285); +lean_ctor_set(x_286, 1, x_280); +return x_286; } } } else { -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_56 = lean_box(0); -x_57 = lean_array_uset(x_16, x_32, x_56); -lean_inc(x_2); -x_58 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__7(x_1, x_2, x_33); -x_59 = lean_array_uset(x_57, x_32, x_58); -lean_ctor_set(x_12, 1, x_59); -x_60 = lean_st_ref_set(x_3, x_12, x_14); -x_61 = !lean_is_exclusive(x_60); -if (x_61 == 0) +uint8_t x_287; +lean_dec(x_6); +lean_dec(x_1); +x_287 = !lean_is_exclusive(x_254); +if (x_287 == 0) { -lean_object* x_62; -x_62 = lean_ctor_get(x_60, 0); -lean_dec(x_62); -lean_ctor_set(x_10, 1, x_4); -lean_ctor_set(x_10, 0, x_2); -lean_ctor_set(x_60, 0, x_10); -return x_60; +return x_254; } else { -lean_object* x_63; lean_object* x_64; -x_63 = lean_ctor_get(x_60, 1); -lean_inc(x_63); -lean_dec(x_60); -lean_ctor_set(x_10, 1, x_4); -lean_ctor_set(x_10, 0, x_2); -x_64 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_64, 0, x_10); -lean_ctor_set(x_64, 1, x_63); -return x_64; +lean_object* x_288; lean_object* x_289; lean_object* x_290; +x_288 = lean_ctor_get(x_254, 0); +x_289 = lean_ctor_get(x_254, 1); +lean_inc(x_289); +lean_inc(x_288); +lean_dec(x_254); +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 { -lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; size_t x_69; uint64_t x_70; uint64_t x_71; uint64_t x_72; uint64_t x_73; uint64_t x_74; uint64_t x_75; uint64_t x_76; uint64_t x_77; uint64_t x_78; size_t x_79; size_t x_80; size_t x_81; size_t x_82; size_t x_83; lean_object* x_84; uint8_t x_85; -x_65 = lean_ctor_get(x_10, 1); -x_66 = lean_ctor_get(x_12, 0); -x_67 = lean_ctor_get(x_12, 1); -lean_inc(x_67); -lean_inc(x_66); +uint8_t x_291; +lean_dec(x_14); +lean_dec(x_13); lean_dec(x_12); -x_68 = lean_array_get_size(x_67); -x_69 = lean_ptr_addr(x_1); -x_70 = lean_usize_to_uint64(x_69); -x_71 = 11; -x_72 = lean_uint64_mix_hash(x_70, x_71); -x_73 = 32; -x_74 = lean_uint64_shift_right(x_72, x_73); -x_75 = lean_uint64_xor(x_72, x_74); -x_76 = 16; -x_77 = lean_uint64_shift_right(x_75, x_76); -x_78 = lean_uint64_xor(x_75, x_77); -x_79 = lean_uint64_to_usize(x_78); -x_80 = lean_usize_of_nat(x_68); -lean_dec(x_68); -x_81 = 1; -x_82 = lean_usize_sub(x_80, x_81); -x_83 = lean_usize_land(x_79, x_82); -x_84 = lean_array_uget(x_67, x_83); -x_85 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__2(x_1, x_84); -if (x_85 == 0) +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); +lean_dec(x_1); +x_291 = !lean_is_exclusive(x_242); +if (x_291 == 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; uint8_t x_95; -x_86 = lean_unsigned_to_nat(1u); -x_87 = lean_nat_add(x_66, x_86); -lean_dec(x_66); -lean_inc(x_2); -x_88 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_88, 0, x_1); -lean_ctor_set(x_88, 1, x_2); -lean_ctor_set(x_88, 2, x_84); -x_89 = lean_array_uset(x_67, x_83, x_88); -x_90 = lean_unsigned_to_nat(4u); -x_91 = lean_nat_mul(x_87, x_90); -x_92 = lean_unsigned_to_nat(3u); -x_93 = lean_nat_div(x_91, x_92); -lean_dec(x_91); -x_94 = lean_array_get_size(x_89); -x_95 = lean_nat_dec_le(x_93, x_94); -lean_dec(x_94); -lean_dec(x_93); -if (x_95 == 0) +return x_242; +} +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; -x_96 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__3(x_89); -x_97 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_97, 0, x_87); -lean_ctor_set(x_97, 1, x_96); -x_98 = lean_st_ref_set(x_3, x_97, x_65); -x_99 = lean_ctor_get(x_98, 1); -lean_inc(x_99); -if (lean_is_exclusive(x_98)) { - lean_ctor_release(x_98, 0); - lean_ctor_release(x_98, 1); - x_100 = x_98; -} else { - lean_dec_ref(x_98); - x_100 = lean_box(0); +lean_object* x_292; lean_object* x_293; lean_object* x_294; +x_292 = lean_ctor_get(x_242, 0); +x_293 = lean_ctor_get(x_242, 1); +lean_inc(x_293); +lean_inc(x_292); +lean_dec(x_242); +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_ctor_set(x_10, 1, x_4); -lean_ctor_set(x_10, 0, x_2); -if (lean_is_scalar(x_100)) { - x_101 = lean_alloc_ctor(0, 2, 0); -} else { - x_101 = x_100; } -lean_ctor_set(x_101, 0, x_10); -lean_ctor_set(x_101, 1, x_99); -return x_101; +} +case 2: +{ +lean_object* x_308; lean_object* x_387; lean_object* x_442; uint8_t x_443; +lean_dec(x_8); +x_442 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__11___closed__5; +x_443 = l_Lean_Expr_isConstOf(x_6, x_442); +if (x_443 == 0) +{ +lean_object* x_444; +x_444 = lean_box(0); +x_387 = x_444; +goto block_441; } else { -lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; -x_102 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_102, 0, x_87); -lean_ctor_set(x_102, 1, x_89); -x_103 = lean_st_ref_set(x_3, x_102, x_65); -x_104 = lean_ctor_get(x_103, 1); -lean_inc(x_104); -if (lean_is_exclusive(x_103)) { - lean_ctor_release(x_103, 0); - lean_ctor_release(x_103, 1); - x_105 = x_103; -} else { - lean_dec_ref(x_103); - x_105 = lean_box(0); +lean_object* x_445; lean_object* x_446; uint8_t x_447; +x_445 = lean_array_get_size(x_7); +x_446 = lean_unsigned_to_nat(2u); +x_447 = lean_nat_dec_eq(x_445, x_446); +if (x_447 == 0) +{ +lean_object* x_448; +lean_dec(x_445); +x_448 = lean_box(0); +x_387 = x_448; +goto block_441; } -lean_ctor_set(x_10, 1, x_4); -lean_ctor_set(x_10, 0, x_2); -if (lean_is_scalar(x_105)) { - x_106 = lean_alloc_ctor(0, 2, 0); -} else { - x_106 = x_105; +else +{ +lean_object* x_449; uint8_t x_450; +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_449 = lean_unsigned_to_nat(0u); +x_450 = lean_nat_dec_lt(x_449, x_445); +lean_dec(x_445); +if (x_450 == 0) +{ +lean_object* x_451; lean_object* x_452; +x_451 = l_Lean_instInhabitedExpr; +x_452 = l_outOfBounds___rarg(x_451); +x_308 = x_452; +goto block_386; +} +else +{ +lean_object* x_453; +x_453 = lean_array_fget(x_7, x_449); +x_308 = x_453; +goto block_386; +} +} +} +block_386: +{ +lean_object* x_309; +lean_inc(x_308); +x_309 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_308, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +if (lean_obj_tag(x_309) == 0) +{ +uint8_t x_310; +x_310 = !lean_is_exclusive(x_309); +if (x_310 == 0) +{ +lean_object* x_311; uint8_t x_312; +x_311 = lean_ctor_get(x_309, 0); +x_312 = !lean_is_exclusive(x_311); +if (x_312 == 0) +{ +lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; +x_313 = lean_ctor_get(x_311, 0); +x_314 = lean_ctor_get(x_311, 1); +x_315 = lean_ctor_get(x_314, 2); +lean_inc(x_315); +lean_inc(x_315); +x_316 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_315, x_313); +if (lean_obj_tag(x_316) == 0) +{ +size_t x_317; size_t x_318; uint8_t x_319; +x_317 = lean_ptr_addr(x_308); +lean_dec(x_308); +x_318 = lean_ptr_addr(x_313); +x_319 = lean_usize_dec_eq(x_317, x_318); +if (x_319 == 0) +{ +lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; +lean_dec(x_1); +x_320 = lean_unsigned_to_nat(0u); +lean_inc(x_313); +x_321 = lean_array_set(x_7, x_320, x_313); +x_322 = l_Lean_mkAppN(x_6, x_321); +lean_dec(x_321); +x_323 = lean_ctor_get(x_314, 0); +lean_inc(x_323); +x_324 = lean_ctor_get(x_314, 1); +lean_inc(x_324); +lean_dec(x_314); +lean_inc(x_322); +x_325 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_315, x_313, x_322); +x_326 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_326, 0, x_323); +lean_ctor_set(x_326, 1, x_324); +lean_ctor_set(x_326, 2, x_325); +lean_ctor_set(x_311, 1, x_326); +lean_ctor_set(x_311, 0, x_322); +return x_309; +} +else +{ +lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; +lean_dec(x_7); +lean_dec(x_6); +x_327 = lean_ctor_get(x_314, 0); +lean_inc(x_327); +x_328 = lean_ctor_get(x_314, 1); +lean_inc(x_328); +lean_dec(x_314); +lean_inc(x_1); +x_329 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_315, x_313, x_1); +x_330 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_330, 0, x_327); +lean_ctor_set(x_330, 1, x_328); +lean_ctor_set(x_330, 2, x_329); +lean_ctor_set(x_311, 1, x_330); +lean_ctor_set(x_311, 0, x_1); +return x_309; +} +} +else +{ +lean_object* x_331; +lean_dec(x_315); +lean_dec(x_313); +lean_dec(x_308); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_1); +x_331 = lean_ctor_get(x_316, 0); +lean_inc(x_331); +lean_dec(x_316); +lean_ctor_set(x_311, 0, x_331); +return x_309; +} +} +else +{ +lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; +x_332 = lean_ctor_get(x_311, 0); +x_333 = lean_ctor_get(x_311, 1); +lean_inc(x_333); +lean_inc(x_332); +lean_dec(x_311); +x_334 = lean_ctor_get(x_333, 2); +lean_inc(x_334); +lean_inc(x_334); +x_335 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_334, x_332); +if (lean_obj_tag(x_335) == 0) +{ +size_t x_336; size_t x_337; uint8_t x_338; +x_336 = lean_ptr_addr(x_308); +lean_dec(x_308); +x_337 = lean_ptr_addr(x_332); +x_338 = lean_usize_dec_eq(x_336, x_337); +if (x_338 == 0) +{ +lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; +lean_dec(x_1); +x_339 = lean_unsigned_to_nat(0u); +lean_inc(x_332); +x_340 = lean_array_set(x_7, x_339, x_332); +x_341 = l_Lean_mkAppN(x_6, x_340); +lean_dec(x_340); +x_342 = lean_ctor_get(x_333, 0); +lean_inc(x_342); +x_343 = lean_ctor_get(x_333, 1); +lean_inc(x_343); +lean_dec(x_333); +lean_inc(x_341); +x_344 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_334, x_332, x_341); +x_345 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_345, 0, x_342); +lean_ctor_set(x_345, 1, x_343); +lean_ctor_set(x_345, 2, x_344); +x_346 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_346, 0, x_341); +lean_ctor_set(x_346, 1, x_345); +lean_ctor_set(x_309, 0, x_346); +return x_309; +} +else +{ +lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; +lean_dec(x_7); +lean_dec(x_6); +x_347 = lean_ctor_get(x_333, 0); +lean_inc(x_347); +x_348 = lean_ctor_get(x_333, 1); +lean_inc(x_348); +lean_dec(x_333); +lean_inc(x_1); +x_349 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_334, x_332, x_1); +x_350 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_350, 0, x_347); +lean_ctor_set(x_350, 1, x_348); +lean_ctor_set(x_350, 2, x_349); +x_351 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_351, 0, x_1); +lean_ctor_set(x_351, 1, x_350); +lean_ctor_set(x_309, 0, x_351); +return x_309; +} +} +else +{ +lean_object* x_352; lean_object* x_353; +lean_dec(x_334); +lean_dec(x_332); +lean_dec(x_308); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_1); +x_352 = lean_ctor_get(x_335, 0); +lean_inc(x_352); +lean_dec(x_335); +x_353 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_353, 0, x_352); +lean_ctor_set(x_353, 1, x_333); +lean_ctor_set(x_309, 0, x_353); +return x_309; } -lean_ctor_set(x_106, 0, x_10); -lean_ctor_set(x_106, 1, x_104); -return x_106; } } else { -lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; -x_107 = lean_box(0); -x_108 = lean_array_uset(x_67, x_83, x_107); -lean_inc(x_2); -x_109 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__7(x_1, x_2, x_84); -x_110 = lean_array_uset(x_108, x_83, x_109); -x_111 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_111, 0, x_66); -lean_ctor_set(x_111, 1, x_110); -x_112 = lean_st_ref_set(x_3, x_111, x_65); -x_113 = lean_ctor_get(x_112, 1); -lean_inc(x_113); -if (lean_is_exclusive(x_112)) { - lean_ctor_release(x_112, 0); - lean_ctor_release(x_112, 1); - x_114 = x_112; +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_354 = lean_ctor_get(x_309, 0); +x_355 = lean_ctor_get(x_309, 1); +lean_inc(x_355); +lean_inc(x_354); +lean_dec(x_309); +x_356 = lean_ctor_get(x_354, 0); +lean_inc(x_356); +x_357 = lean_ctor_get(x_354, 1); +lean_inc(x_357); +if (lean_is_exclusive(x_354)) { + lean_ctor_release(x_354, 0); + lean_ctor_release(x_354, 1); + x_358 = x_354; } else { - lean_dec_ref(x_112); - x_114 = lean_box(0); -} -lean_ctor_set(x_10, 1, x_4); -lean_ctor_set(x_10, 0, x_2); -if (lean_is_scalar(x_114)) { - x_115 = lean_alloc_ctor(0, 2, 0); + lean_dec_ref(x_354); + x_358 = lean_box(0); +} +x_359 = lean_ctor_get(x_357, 2); +lean_inc(x_359); +lean_inc(x_359); +x_360 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_359, x_356); +if (lean_obj_tag(x_360) == 0) +{ +size_t x_361; size_t x_362; uint8_t x_363; +x_361 = lean_ptr_addr(x_308); +lean_dec(x_308); +x_362 = lean_ptr_addr(x_356); +x_363 = lean_usize_dec_eq(x_361, x_362); +if (x_363 == 0) +{ +lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; lean_object* x_369; lean_object* x_370; lean_object* x_371; lean_object* x_372; +lean_dec(x_1); +x_364 = lean_unsigned_to_nat(0u); +lean_inc(x_356); +x_365 = lean_array_set(x_7, x_364, x_356); +x_366 = l_Lean_mkAppN(x_6, x_365); +lean_dec(x_365); +x_367 = lean_ctor_get(x_357, 0); +lean_inc(x_367); +x_368 = lean_ctor_get(x_357, 1); +lean_inc(x_368); +lean_dec(x_357); +lean_inc(x_366); +x_369 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_359, x_356, x_366); +x_370 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_370, 0, x_367); +lean_ctor_set(x_370, 1, x_368); +lean_ctor_set(x_370, 2, x_369); +if (lean_is_scalar(x_358)) { + x_371 = lean_alloc_ctor(0, 2, 0); } else { - x_115 = x_114; + x_371 = x_358; } -lean_ctor_set(x_115, 0, x_10); -lean_ctor_set(x_115, 1, x_113); -return x_115; +lean_ctor_set(x_371, 0, x_366); +lean_ctor_set(x_371, 1, x_370); +x_372 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_372, 0, x_371); +lean_ctor_set(x_372, 1, x_355); +return x_372; +} +else +{ +lean_object* x_373; lean_object* x_374; lean_object* x_375; lean_object* x_376; lean_object* x_377; lean_object* x_378; +lean_dec(x_7); +lean_dec(x_6); +x_373 = lean_ctor_get(x_357, 0); +lean_inc(x_373); +x_374 = lean_ctor_get(x_357, 1); +lean_inc(x_374); +lean_dec(x_357); +lean_inc(x_1); +x_375 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_359, x_356, x_1); +x_376 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_376, 0, x_373); +lean_ctor_set(x_376, 1, x_374); +lean_ctor_set(x_376, 2, x_375); +if (lean_is_scalar(x_358)) { + x_377 = lean_alloc_ctor(0, 2, 0); +} else { + x_377 = x_358; } +lean_ctor_set(x_377, 0, x_1); +lean_ctor_set(x_377, 1, x_376); +x_378 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_378, 0, x_377); +lean_ctor_set(x_378, 1, x_355); +return x_378; } } 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; size_t x_122; uint64_t x_123; uint64_t x_124; uint64_t x_125; uint64_t x_126; uint64_t x_127; uint64_t x_128; uint64_t x_129; uint64_t x_130; uint64_t x_131; size_t x_132; size_t x_133; size_t x_134; size_t x_135; size_t x_136; lean_object* x_137; uint8_t x_138; -x_116 = lean_ctor_get(x_10, 0); -x_117 = lean_ctor_get(x_10, 1); -lean_inc(x_117); -lean_inc(x_116); -lean_dec(x_10); -x_118 = lean_ctor_get(x_116, 0); -lean_inc(x_118); -x_119 = lean_ctor_get(x_116, 1); -lean_inc(x_119); -if (lean_is_exclusive(x_116)) { - lean_ctor_release(x_116, 0); - lean_ctor_release(x_116, 1); - x_120 = x_116; +lean_object* x_379; lean_object* x_380; lean_object* x_381; +lean_dec(x_359); +lean_dec(x_356); +lean_dec(x_308); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_1); +x_379 = lean_ctor_get(x_360, 0); +lean_inc(x_379); +lean_dec(x_360); +if (lean_is_scalar(x_358)) { + x_380 = lean_alloc_ctor(0, 2, 0); } else { - lean_dec_ref(x_116); - x_120 = lean_box(0); + x_380 = x_358; } -x_121 = lean_array_get_size(x_119); -x_122 = lean_ptr_addr(x_1); -x_123 = lean_usize_to_uint64(x_122); -x_124 = 11; -x_125 = lean_uint64_mix_hash(x_123, x_124); -x_126 = 32; -x_127 = lean_uint64_shift_right(x_125, x_126); -x_128 = lean_uint64_xor(x_125, x_127); -x_129 = 16; -x_130 = lean_uint64_shift_right(x_128, x_129); -x_131 = lean_uint64_xor(x_128, x_130); -x_132 = lean_uint64_to_usize(x_131); -x_133 = lean_usize_of_nat(x_121); -lean_dec(x_121); -x_134 = 1; -x_135 = lean_usize_sub(x_133, x_134); -x_136 = lean_usize_land(x_132, x_135); -x_137 = lean_array_uget(x_119, x_136); -x_138 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__2(x_1, x_137); -if (x_138 == 0) +lean_ctor_set(x_380, 0, x_379); +lean_ctor_set(x_380, 1, x_357); +x_381 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_381, 0, x_380); +lean_ctor_set(x_381, 1, x_355); +return x_381; +} +} +} +else { -lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; uint8_t x_148; -x_139 = lean_unsigned_to_nat(1u); -x_140 = lean_nat_add(x_118, x_139); -lean_dec(x_118); -lean_inc(x_2); -x_141 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_141, 0, x_1); -lean_ctor_set(x_141, 1, x_2); -lean_ctor_set(x_141, 2, x_137); -x_142 = lean_array_uset(x_119, x_136, x_141); -x_143 = lean_unsigned_to_nat(4u); -x_144 = lean_nat_mul(x_140, x_143); -x_145 = lean_unsigned_to_nat(3u); -x_146 = lean_nat_div(x_144, x_145); -lean_dec(x_144); -x_147 = lean_array_get_size(x_142); -x_148 = lean_nat_dec_le(x_146, x_147); -lean_dec(x_147); -lean_dec(x_146); -if (x_148 == 0) +uint8_t x_382; +lean_dec(x_308); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_1); +x_382 = !lean_is_exclusive(x_309); +if (x_382 == 0) { -lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; -x_149 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__3(x_142); -if (lean_is_scalar(x_120)) { - x_150 = lean_alloc_ctor(0, 2, 0); -} else { - x_150 = x_120; +return x_309; } -lean_ctor_set(x_150, 0, x_140); -lean_ctor_set(x_150, 1, x_149); -x_151 = lean_st_ref_set(x_3, x_150, x_117); -x_152 = lean_ctor_get(x_151, 1); -lean_inc(x_152); -if (lean_is_exclusive(x_151)) { - lean_ctor_release(x_151, 0); - lean_ctor_release(x_151, 1); - x_153 = x_151; -} else { - lean_dec_ref(x_151); - x_153 = lean_box(0); +else +{ +lean_object* x_383; lean_object* x_384; lean_object* x_385; +x_383 = lean_ctor_get(x_309, 0); +x_384 = lean_ctor_get(x_309, 1); +lean_inc(x_384); +lean_inc(x_383); +lean_dec(x_309); +x_385 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_385, 0, x_383); +lean_ctor_set(x_385, 1, x_384); +return x_385; } -x_154 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_154, 0, x_2); -lean_ctor_set(x_154, 1, x_4); -if (lean_is_scalar(x_153)) { - x_155 = lean_alloc_ctor(0, 2, 0); -} else { - x_155 = x_153; } -lean_ctor_set(x_155, 0, x_154); -lean_ctor_set(x_155, 1, x_152); -return x_155; } -else +block_441: { -lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; -if (lean_is_scalar(x_120)) { - x_156 = lean_alloc_ctor(0, 2, 0); -} else { - x_156 = x_120; +lean_object* x_388; +lean_dec(x_387); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_6); +x_388 = l_Lean_Meta_getFunInfo(x_6, x_11, x_12, x_13, x_14, x_15); +if (lean_obj_tag(x_388) == 0) +{ +lean_object* x_389; lean_object* x_390; lean_object* x_391; lean_object* x_392; lean_object* x_393; lean_object* x_394; lean_object* x_395; uint8_t x_396; lean_object* x_397; lean_object* x_398; lean_object* x_399; lean_object* x_400; +x_389 = lean_ctor_get(x_388, 0); +lean_inc(x_389); +x_390 = lean_ctor_get(x_388, 1); +lean_inc(x_390); +lean_dec(x_388); +x_391 = lean_ctor_get(x_389, 0); +lean_inc(x_391); +lean_dec(x_389); +x_392 = lean_array_get_size(x_7); +x_393 = lean_unsigned_to_nat(0u); +x_394 = lean_unsigned_to_nat(1u); +lean_inc(x_392); +x_395 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_395, 0, x_393); +lean_ctor_set(x_395, 1, x_392); +lean_ctor_set(x_395, 2, x_394); +x_396 = 0; +x_397 = lean_box(x_396); +lean_inc(x_7); +x_398 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_398, 0, x_7); +lean_ctor_set(x_398, 1, x_397); +x_399 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__11___closed__1; +lean_inc(x_395); +lean_inc(x_6); +x_400 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10(x_2, x_3, x_4, x_5, x_6, x_7, x_399, x_391, x_392, x_395, x_395, x_398, x_393, lean_box(0), lean_box(0), x_9, x_10, x_11, x_12, x_13, x_14, x_390); +if (lean_obj_tag(x_400) == 0) +{ +lean_object* x_401; lean_object* x_402; lean_object* x_403; uint8_t x_404; +x_401 = lean_ctor_get(x_400, 0); +lean_inc(x_401); +x_402 = lean_ctor_get(x_401, 0); +lean_inc(x_402); +x_403 = lean_ctor_get(x_402, 1); +lean_inc(x_403); +x_404 = lean_unbox(x_403); +lean_dec(x_403); +if (x_404 == 0) +{ +uint8_t x_405; +lean_dec(x_402); +lean_dec(x_6); +x_405 = !lean_is_exclusive(x_400); +if (x_405 == 0) +{ +lean_object* x_406; uint8_t x_407; +x_406 = lean_ctor_get(x_400, 0); +lean_dec(x_406); +x_407 = !lean_is_exclusive(x_401); +if (x_407 == 0) +{ +lean_object* x_408; +x_408 = lean_ctor_get(x_401, 0); +lean_dec(x_408); +lean_ctor_set(x_401, 0, x_1); +return x_400; } -lean_ctor_set(x_156, 0, x_140); -lean_ctor_set(x_156, 1, x_142); -x_157 = lean_st_ref_set(x_3, x_156, x_117); -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_object* x_409; lean_object* x_410; +x_409 = lean_ctor_get(x_401, 1); +lean_inc(x_409); +lean_dec(x_401); +x_410 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_410, 0, x_1); +lean_ctor_set(x_410, 1, x_409); +lean_ctor_set(x_400, 0, x_410); +return x_400; +} +} +else +{ +lean_object* x_411; lean_object* x_412; lean_object* x_413; lean_object* x_414; lean_object* x_415; +x_411 = lean_ctor_get(x_400, 1); +lean_inc(x_411); +lean_dec(x_400); +x_412 = lean_ctor_get(x_401, 1); +lean_inc(x_412); +if (lean_is_exclusive(x_401)) { + lean_ctor_release(x_401, 0); + lean_ctor_release(x_401, 1); + x_413 = x_401; } else { - lean_dec_ref(x_157); - x_159 = lean_box(0); + lean_dec_ref(x_401); + x_413 = lean_box(0); } -x_160 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_160, 0, x_2); -lean_ctor_set(x_160, 1, x_4); -if (lean_is_scalar(x_159)) { - x_161 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_413)) { + x_414 = lean_alloc_ctor(0, 2, 0); } else { - x_161 = x_159; + x_414 = x_413; } -lean_ctor_set(x_161, 0, x_160); -lean_ctor_set(x_161, 1, x_158); -return x_161; +lean_ctor_set(x_414, 0, x_1); +lean_ctor_set(x_414, 1, x_412); +x_415 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_415, 0, x_414); +lean_ctor_set(x_415, 1, x_411); +return x_415; } } else { -lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; -x_162 = lean_box(0); -x_163 = lean_array_uset(x_119, x_136, x_162); -lean_inc(x_2); -x_164 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__7(x_1, x_2, x_137); -x_165 = lean_array_uset(x_163, x_136, x_164); -if (lean_is_scalar(x_120)) { - x_166 = lean_alloc_ctor(0, 2, 0); -} else { - x_166 = x_120; +uint8_t x_416; +lean_dec(x_1); +x_416 = !lean_is_exclusive(x_400); +if (x_416 == 0) +{ +lean_object* x_417; uint8_t x_418; +x_417 = lean_ctor_get(x_400, 0); +lean_dec(x_417); +x_418 = !lean_is_exclusive(x_401); +if (x_418 == 0) +{ +lean_object* x_419; lean_object* x_420; lean_object* x_421; +x_419 = lean_ctor_get(x_401, 0); +lean_dec(x_419); +x_420 = lean_ctor_get(x_402, 0); +lean_inc(x_420); +lean_dec(x_402); +x_421 = l_Lean_mkAppN(x_6, x_420); +lean_dec(x_420); +lean_ctor_set(x_401, 0, x_421); +return x_400; +} +else +{ +lean_object* x_422; lean_object* x_423; lean_object* x_424; lean_object* x_425; +x_422 = lean_ctor_get(x_401, 1); +lean_inc(x_422); +lean_dec(x_401); +x_423 = lean_ctor_get(x_402, 0); +lean_inc(x_423); +lean_dec(x_402); +x_424 = l_Lean_mkAppN(x_6, x_423); +lean_dec(x_423); +x_425 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_425, 0, x_424); +lean_ctor_set(x_425, 1, x_422); +lean_ctor_set(x_400, 0, x_425); +return x_400; } -lean_ctor_set(x_166, 0, x_118); -lean_ctor_set(x_166, 1, x_165); -x_167 = lean_st_ref_set(x_3, x_166, x_117); -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); } -x_170 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_170, 0, x_2); -lean_ctor_set(x_170, 1, x_4); -if (lean_is_scalar(x_169)) { - x_171 = lean_alloc_ctor(0, 2, 0); +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; +x_426 = lean_ctor_get(x_400, 1); +lean_inc(x_426); +lean_dec(x_400); +x_427 = lean_ctor_get(x_401, 1); +lean_inc(x_427); +if (lean_is_exclusive(x_401)) { + lean_ctor_release(x_401, 0); + lean_ctor_release(x_401, 1); + x_428 = x_401; } else { - x_171 = x_169; -} -lean_ctor_set(x_171, 0, x_170); -lean_ctor_set(x_171, 1, x_168); -return x_171; + lean_dec_ref(x_401); + x_428 = lean_box(0); +} +x_429 = lean_ctor_get(x_402, 0); +lean_inc(x_429); +lean_dec(x_402); +x_430 = l_Lean_mkAppN(x_6, x_429); +lean_dec(x_429); +if (lean_is_scalar(x_428)) { + x_431 = lean_alloc_ctor(0, 2, 0); +} else { + x_431 = x_428; } +lean_ctor_set(x_431, 0, x_430); +lean_ctor_set(x_431, 1, x_427); +x_432 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_432, 0, x_431); +lean_ctor_set(x_432, 1, x_426); +return x_432; } } } -static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__1() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean", 4, 4); -return x_1; -} +uint8_t x_433; +lean_dec(x_6); +lean_dec(x_1); +x_433 = !lean_is_exclusive(x_400); +if (x_433 == 0) +{ +return x_400; } -static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__2() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Grind", 5, 5); -return x_1; +lean_object* x_434; lean_object* x_435; lean_object* x_436; +x_434 = lean_ctor_get(x_400, 0); +x_435 = lean_ctor_get(x_400, 1); +lean_inc(x_435); +lean_inc(x_434); +lean_dec(x_400); +x_436 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_436, 0, x_434); +lean_ctor_set(x_436, 1, x_435); +return x_436; } } -static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__3() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("nestedProof", 11, 11); -return x_1; } +else +{ +uint8_t x_437; +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_437 = !lean_is_exclusive(x_388); +if (x_437 == 0) +{ +return x_388; } -static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__4() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__1; -x_2 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__2; -x_3 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__3; -x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); -return x_4; +lean_object* x_438; lean_object* x_439; lean_object* x_440; +x_438 = lean_ctor_get(x_388, 0); +x_439 = lean_ctor_get(x_388, 1); +lean_inc(x_439); +lean_inc(x_438); +lean_dec(x_388); +x_440 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_440, 0, x_438); +lean_ctor_set(x_440, 1, x_439); +return x_440; } } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___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) { -_start: -{ -switch (lean_obj_tag(x_2)) { -case 0: +} +} +case 3: { -lean_object* x_12; lean_object* x_43; lean_object* x_77; uint8_t x_78; -lean_dec(x_4); -x_77 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__4; -x_78 = l_Lean_Expr_isConstOf(x_2, x_77); -if (x_78 == 0) +lean_object* x_454; lean_object* x_533; lean_object* x_588; uint8_t x_589; +lean_dec(x_8); +x_588 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__11___closed__5; +x_589 = l_Lean_Expr_isConstOf(x_6, x_588); +if (x_589 == 0) { -lean_object* x_79; -x_79 = lean_box(0); -x_43 = x_79; -goto block_76; +lean_object* x_590; +x_590 = lean_box(0); +x_533 = x_590; +goto block_587; } else { -lean_object* x_80; lean_object* x_81; uint8_t x_82; -x_80 = lean_array_get_size(x_3); -x_81 = lean_unsigned_to_nat(2u); -x_82 = lean_nat_dec_eq(x_80, x_81); -if (x_82 == 0) +lean_object* x_591; lean_object* x_592; uint8_t x_593; +x_591 = lean_array_get_size(x_7); +x_592 = lean_unsigned_to_nat(2u); +x_593 = lean_nat_dec_eq(x_591, x_592); +if (x_593 == 0) { -lean_object* x_83; -lean_dec(x_80); -x_83 = lean_box(0); -x_43 = x_83; -goto block_76; +lean_object* x_594; +lean_dec(x_591); +x_594 = lean_box(0); +x_533 = x_594; +goto block_587; } else { -lean_object* x_84; uint8_t x_85; -x_84 = lean_unsigned_to_nat(0u); -x_85 = lean_nat_dec_lt(x_84, x_80); -lean_dec(x_80); -if (x_85 == 0) +lean_object* x_595; uint8_t x_596; +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_595 = lean_unsigned_to_nat(0u); +x_596 = lean_nat_dec_lt(x_595, x_591); +lean_dec(x_591); +if (x_596 == 0) { -lean_object* x_86; lean_object* x_87; -x_86 = l_Lean_instInhabitedExpr; -x_87 = l_outOfBounds___rarg(x_86); -x_12 = x_87; -goto block_42; +lean_object* x_597; lean_object* x_598; +x_597 = l_Lean_instInhabitedExpr; +x_598 = l_outOfBounds___rarg(x_597); +x_454 = x_598; +goto block_532; } else { -lean_object* x_88; -x_88 = lean_array_fget(x_3, x_84); -x_12 = x_88; -goto block_42; +lean_object* x_599; +x_599 = lean_array_fget(x_7, x_595); +x_454 = x_599; +goto block_532; } } } -block_42: +block_532: { -lean_object* x_13; -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_5); -lean_inc(x_12); -x_13 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_12, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_13) == 0) +lean_object* x_455; +lean_inc(x_454); +x_455 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_454, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +if (lean_obj_tag(x_455) == 0) +{ +uint8_t x_456; +x_456 = !lean_is_exclusive(x_455); +if (x_456 == 0) +{ +lean_object* x_457; uint8_t x_458; +x_457 = lean_ctor_get(x_455, 0); +x_458 = !lean_is_exclusive(x_457); +if (x_458 == 0) +{ +lean_object* x_459; lean_object* x_460; lean_object* x_461; lean_object* x_462; +x_459 = lean_ctor_get(x_457, 0); +x_460 = lean_ctor_get(x_457, 1); +x_461 = lean_ctor_get(x_460, 2); +lean_inc(x_461); +lean_inc(x_461); +x_462 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_461, x_459); +if (lean_obj_tag(x_462) == 0) { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_14 = lean_ctor_get(x_13, 0); -lean_inc(x_14); -x_15 = lean_ctor_get(x_13, 1); -lean_inc(x_15); -lean_dec(x_13); -x_16 = lean_ctor_get(x_14, 0); -lean_inc(x_16); -x_17 = lean_ctor_get(x_14, 1); -lean_inc(x_17); -lean_dec(x_14); -x_18 = lean_ctor_get(x_17, 2); -lean_inc(x_18); -lean_inc(x_18); -x_19 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_18, x_16); -if (lean_obj_tag(x_19) == 0) +size_t x_463; size_t x_464; uint8_t x_465; +x_463 = lean_ptr_addr(x_454); +lean_dec(x_454); +x_464 = lean_ptr_addr(x_459); +x_465 = lean_usize_dec_eq(x_463, x_464); +if (x_465 == 0) { -size_t x_20; size_t x_21; uint8_t x_22; -x_20 = lean_ptr_addr(x_12); -lean_dec(x_12); -x_21 = lean_ptr_addr(x_16); -x_22 = lean_usize_dec_eq(x_20, x_21); -if (x_22 == 0) +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_dec(x_1); +x_466 = lean_unsigned_to_nat(0u); +lean_inc(x_459); +x_467 = lean_array_set(x_7, x_466, x_459); +x_468 = l_Lean_mkAppN(x_6, x_467); +lean_dec(x_467); +x_469 = lean_ctor_get(x_460, 0); +lean_inc(x_469); +x_470 = lean_ctor_get(x_460, 1); +lean_inc(x_470); +lean_dec(x_460); +lean_inc(x_468); +x_471 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_461, x_459, x_468); +x_472 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_472, 0, x_469); +lean_ctor_set(x_472, 1, x_470); +lean_ctor_set(x_472, 2, x_471); +lean_ctor_set(x_457, 1, x_472); +lean_ctor_set(x_457, 0, x_468); +return x_455; +} +else +{ +lean_object* x_473; lean_object* x_474; lean_object* x_475; lean_object* x_476; +lean_dec(x_7); +lean_dec(x_6); +x_473 = lean_ctor_get(x_460, 0); +lean_inc(x_473); +x_474 = lean_ctor_get(x_460, 1); +lean_inc(x_474); +lean_dec(x_460); +lean_inc(x_1); +x_475 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_461, x_459, x_1); +x_476 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_476, 0, x_473); +lean_ctor_set(x_476, 1, x_474); +lean_ctor_set(x_476, 2, x_475); +lean_ctor_set(x_457, 1, x_476); +lean_ctor_set(x_457, 0, x_1); +return x_455; +} +} +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; -x_23 = lean_unsigned_to_nat(0u); -lean_inc(x_16); -x_24 = lean_array_set(x_3, x_23, x_16); -x_25 = l_Lean_mkAppN(x_2, x_24); -lean_dec(x_24); -x_26 = lean_ctor_get(x_17, 0); -lean_inc(x_26); -x_27 = lean_ctor_get(x_17, 1); -lean_inc(x_27); -lean_dec(x_17); -lean_inc(x_25); -x_28 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_18, x_16, x_25); -x_29 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_29, 0, x_26); -lean_ctor_set(x_29, 1, x_27); -lean_ctor_set(x_29, 2, x_28); -x_30 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_25, x_5, x_29, x_7, x_8, x_9, x_10, x_15); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); +lean_object* x_477; +lean_dec(x_461); +lean_dec(x_459); +lean_dec(x_454); lean_dec(x_7); -lean_dec(x_5); -return x_30; +lean_dec(x_6); +lean_dec(x_1); +x_477 = lean_ctor_get(x_462, 0); +lean_inc(x_477); +lean_dec(x_462); +lean_ctor_set(x_457, 0, x_477); +return x_455; +} } else { -lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; -lean_dec(x_3); -lean_dec(x_2); -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); -lean_inc(x_1); -x_33 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_18, x_16, x_1); -x_34 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_34, 0, x_31); -lean_ctor_set(x_34, 1, x_32); -lean_ctor_set(x_34, 2, x_33); +lean_object* x_478; lean_object* x_479; lean_object* x_480; lean_object* x_481; +x_478 = lean_ctor_get(x_457, 0); +x_479 = lean_ctor_get(x_457, 1); +lean_inc(x_479); +lean_inc(x_478); +lean_dec(x_457); +x_480 = lean_ctor_get(x_479, 2); +lean_inc(x_480); +lean_inc(x_480); +x_481 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_480, x_478); +if (lean_obj_tag(x_481) == 0) +{ +size_t x_482; size_t x_483; uint8_t x_484; +x_482 = lean_ptr_addr(x_454); +lean_dec(x_454); +x_483 = lean_ptr_addr(x_478); +x_484 = lean_usize_dec_eq(x_482, x_483); +if (x_484 == 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_dec(x_1); +x_485 = lean_unsigned_to_nat(0u); +lean_inc(x_478); +x_486 = lean_array_set(x_7, x_485, x_478); +x_487 = l_Lean_mkAppN(x_6, x_486); +lean_dec(x_486); +x_488 = lean_ctor_get(x_479, 0); +lean_inc(x_488); +x_489 = lean_ctor_get(x_479, 1); +lean_inc(x_489); +lean_dec(x_479); +lean_inc(x_487); +x_490 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_480, x_478, x_487); +x_491 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_491, 0, x_488); +lean_ctor_set(x_491, 1, x_489); +lean_ctor_set(x_491, 2, x_490); +x_492 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_492, 0, x_487); +lean_ctor_set(x_492, 1, x_491); +lean_ctor_set(x_455, 0, x_492); +return x_455; +} +else +{ +lean_object* x_493; lean_object* x_494; lean_object* x_495; lean_object* x_496; lean_object* x_497; +lean_dec(x_7); +lean_dec(x_6); +x_493 = lean_ctor_get(x_479, 0); +lean_inc(x_493); +x_494 = lean_ctor_get(x_479, 1); +lean_inc(x_494); +lean_dec(x_479); lean_inc(x_1); -x_35 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_34, x_7, x_8, x_9, x_10, x_15); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); +x_495 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_480, x_478, x_1); +x_496 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_496, 0, x_493); +lean_ctor_set(x_496, 1, x_494); +lean_ctor_set(x_496, 2, x_495); +x_497 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_497, 0, x_1); +lean_ctor_set(x_497, 1, x_496); +lean_ctor_set(x_455, 0, x_497); +return x_455; +} +} +else +{ +lean_object* x_498; lean_object* x_499; +lean_dec(x_480); +lean_dec(x_478); +lean_dec(x_454); lean_dec(x_7); -lean_dec(x_5); -return x_35; +lean_dec(x_6); +lean_dec(x_1); +x_498 = lean_ctor_get(x_481, 0); +lean_inc(x_498); +lean_dec(x_481); +x_499 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_499, 0, x_498); +lean_ctor_set(x_499, 1, x_479); +lean_ctor_set(x_455, 0, x_499); +return x_455; +} } } else { -lean_object* x_36; lean_object* x_37; -lean_dec(x_18); -lean_dec(x_16); -lean_dec(x_12); -lean_dec(x_3); -lean_dec(x_2); -x_36 = lean_ctor_get(x_19, 0); -lean_inc(x_36); -lean_dec(x_19); -x_37 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_36, x_5, x_17, x_7, x_8, x_9, x_10, x_15); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); +lean_object* x_500; lean_object* x_501; lean_object* x_502; lean_object* x_503; lean_object* x_504; lean_object* x_505; lean_object* x_506; +x_500 = lean_ctor_get(x_455, 0); +x_501 = lean_ctor_get(x_455, 1); +lean_inc(x_501); +lean_inc(x_500); +lean_dec(x_455); +x_502 = lean_ctor_get(x_500, 0); +lean_inc(x_502); +x_503 = lean_ctor_get(x_500, 1); +lean_inc(x_503); +if (lean_is_exclusive(x_500)) { + lean_ctor_release(x_500, 0); + lean_ctor_release(x_500, 1); + x_504 = x_500; +} else { + lean_dec_ref(x_500); + x_504 = lean_box(0); +} +x_505 = lean_ctor_get(x_503, 2); +lean_inc(x_505); +lean_inc(x_505); +x_506 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_505, x_502); +if (lean_obj_tag(x_506) == 0) +{ +size_t x_507; size_t x_508; uint8_t x_509; +x_507 = lean_ptr_addr(x_454); +lean_dec(x_454); +x_508 = lean_ptr_addr(x_502); +x_509 = lean_usize_dec_eq(x_507, x_508); +if (x_509 == 0) +{ +lean_object* x_510; lean_object* x_511; lean_object* x_512; lean_object* x_513; lean_object* x_514; lean_object* x_515; lean_object* x_516; lean_object* x_517; lean_object* x_518; +lean_dec(x_1); +x_510 = lean_unsigned_to_nat(0u); +lean_inc(x_502); +x_511 = lean_array_set(x_7, x_510, x_502); +x_512 = l_Lean_mkAppN(x_6, x_511); +lean_dec(x_511); +x_513 = lean_ctor_get(x_503, 0); +lean_inc(x_513); +x_514 = lean_ctor_get(x_503, 1); +lean_inc(x_514); +lean_dec(x_503); +lean_inc(x_512); +x_515 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_505, x_502, x_512); +x_516 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_516, 0, x_513); +lean_ctor_set(x_516, 1, x_514); +lean_ctor_set(x_516, 2, x_515); +if (lean_is_scalar(x_504)) { + x_517 = lean_alloc_ctor(0, 2, 0); +} else { + x_517 = x_504; +} +lean_ctor_set(x_517, 0, x_512); +lean_ctor_set(x_517, 1, x_516); +x_518 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_518, 0, x_517); +lean_ctor_set(x_518, 1, x_501); +return x_518; +} +else +{ +lean_object* x_519; lean_object* x_520; lean_object* x_521; lean_object* x_522; lean_object* x_523; lean_object* x_524; lean_dec(x_7); -lean_dec(x_5); -return x_37; +lean_dec(x_6); +x_519 = lean_ctor_get(x_503, 0); +lean_inc(x_519); +x_520 = lean_ctor_get(x_503, 1); +lean_inc(x_520); +lean_dec(x_503); +lean_inc(x_1); +x_521 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_505, x_502, x_1); +x_522 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_522, 0, x_519); +lean_ctor_set(x_522, 1, x_520); +lean_ctor_set(x_522, 2, x_521); +if (lean_is_scalar(x_504)) { + x_523 = lean_alloc_ctor(0, 2, 0); +} else { + x_523 = x_504; +} +lean_ctor_set(x_523, 0, x_1); +lean_ctor_set(x_523, 1, x_522); +x_524 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_524, 0, x_523); +lean_ctor_set(x_524, 1, x_501); +return x_524; } } else { -uint8_t x_38; -lean_dec(x_12); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); +lean_object* x_525; lean_object* x_526; lean_object* x_527; +lean_dec(x_505); +lean_dec(x_502); +lean_dec(x_454); lean_dec(x_7); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); +lean_dec(x_6); lean_dec(x_1); -x_38 = !lean_is_exclusive(x_13); -if (x_38 == 0) +x_525 = lean_ctor_get(x_506, 0); +lean_inc(x_525); +lean_dec(x_506); +if (lean_is_scalar(x_504)) { + x_526 = lean_alloc_ctor(0, 2, 0); +} else { + x_526 = x_504; +} +lean_ctor_set(x_526, 0, x_525); +lean_ctor_set(x_526, 1, x_503); +x_527 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_527, 0, x_526); +lean_ctor_set(x_527, 1, x_501); +return x_527; +} +} +} +else { -return x_13; +uint8_t x_528; +lean_dec(x_454); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_1); +x_528 = !lean_is_exclusive(x_455); +if (x_528 == 0) +{ +return x_455; } else { -lean_object* x_39; lean_object* x_40; lean_object* x_41; -x_39 = lean_ctor_get(x_13, 0); -x_40 = lean_ctor_get(x_13, 1); -lean_inc(x_40); -lean_inc(x_39); -lean_dec(x_13); -x_41 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_41, 0, x_39); -lean_ctor_set(x_41, 1, x_40); -return x_41; +lean_object* x_529; lean_object* x_530; lean_object* x_531; +x_529 = lean_ctor_get(x_455, 0); +x_530 = lean_ctor_get(x_455, 1); +lean_inc(x_530); +lean_inc(x_529); +lean_dec(x_455); +x_531 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_531, 0, x_529); +lean_ctor_set(x_531, 1, x_530); +return x_531; } } } -block_76: +block_587: { -lean_object* x_44; -lean_dec(x_43); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); +lean_object* x_534; +lean_dec(x_533); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_6); +x_534 = l_Lean_Meta_getFunInfo(x_6, x_11, x_12, x_13, x_14, x_15); +if (lean_obj_tag(x_534) == 0) +{ +lean_object* x_535; lean_object* x_536; lean_object* x_537; lean_object* x_538; lean_object* x_539; lean_object* x_540; lean_object* x_541; uint8_t x_542; lean_object* x_543; lean_object* x_544; lean_object* x_545; lean_object* x_546; +x_535 = lean_ctor_get(x_534, 0); +lean_inc(x_535); +x_536 = lean_ctor_get(x_534, 1); +lean_inc(x_536); +lean_dec(x_534); +x_537 = lean_ctor_get(x_535, 0); +lean_inc(x_537); +lean_dec(x_535); +x_538 = lean_array_get_size(x_7); +x_539 = lean_unsigned_to_nat(0u); +x_540 = lean_unsigned_to_nat(1u); +lean_inc(x_538); +x_541 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_541, 0, x_539); +lean_ctor_set(x_541, 1, x_538); +lean_ctor_set(x_541, 2, x_540); +x_542 = 0; +x_543 = lean_box(x_542); lean_inc(x_7); -lean_inc(x_2); -x_44 = l_Lean_Meta_getFunInfo(x_2, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_44) == 0) +x_544 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_544, 0, x_7); +lean_ctor_set(x_544, 1, x_543); +x_545 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__11___closed__1; +lean_inc(x_541); +lean_inc(x_6); +x_546 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10(x_2, x_3, x_4, x_5, x_6, x_7, x_545, x_537, x_538, x_541, x_541, x_544, x_539, lean_box(0), lean_box(0), x_9, x_10, x_11, x_12, x_13, x_14, x_536); +if (lean_obj_tag(x_546) == 0) +{ +lean_object* x_547; lean_object* x_548; lean_object* x_549; uint8_t x_550; +x_547 = lean_ctor_get(x_546, 0); +lean_inc(x_547); +x_548 = lean_ctor_get(x_547, 0); +lean_inc(x_548); +x_549 = lean_ctor_get(x_548, 1); +lean_inc(x_549); +x_550 = lean_unbox(x_549); +lean_dec(x_549); +if (x_550 == 0) { -lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; uint8_t x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; -x_45 = lean_ctor_get(x_44, 0); -lean_inc(x_45); -x_46 = lean_ctor_get(x_44, 1); -lean_inc(x_46); -lean_dec(x_44); -x_47 = lean_ctor_get(x_45, 0); -lean_inc(x_47); -lean_dec(x_45); -x_48 = lean_array_get_size(x_3); -x_49 = lean_unsigned_to_nat(0u); -x_50 = lean_unsigned_to_nat(1u); -x_51 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_51, 0, x_49); -lean_ctor_set(x_51, 1, x_48); -lean_ctor_set(x_51, 2, x_50); -x_52 = 0; -x_53 = lean_box(x_52); -x_54 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_54, 0, x_3); -lean_ctor_set(x_54, 1, x_53); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_5); -lean_inc(x_2); -x_55 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8(x_2, x_47, x_51, x_51, x_54, x_49, lean_box(0), lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10, x_46); -lean_dec(x_51); -lean_dec(x_47); -if (lean_obj_tag(x_55) == 0) +uint8_t x_551; +lean_dec(x_548); +lean_dec(x_6); +x_551 = !lean_is_exclusive(x_546); +if (x_551 == 0) { -lean_object* x_56; lean_object* x_57; lean_object* x_58; uint8_t x_59; -x_56 = lean_ctor_get(x_55, 0); -lean_inc(x_56); -x_57 = lean_ctor_get(x_56, 0); -lean_inc(x_57); -x_58 = lean_ctor_get(x_57, 1); -lean_inc(x_58); -x_59 = lean_unbox(x_58); -lean_dec(x_58); -if (x_59 == 0) +lean_object* x_552; uint8_t x_553; +x_552 = lean_ctor_get(x_546, 0); +lean_dec(x_552); +x_553 = !lean_is_exclusive(x_547); +if (x_553 == 0) { -lean_object* x_60; lean_object* x_61; lean_object* x_62; -lean_dec(x_57); -lean_dec(x_2); -x_60 = lean_ctor_get(x_55, 1); -lean_inc(x_60); -lean_dec(x_55); -x_61 = lean_ctor_get(x_56, 1); -lean_inc(x_61); -lean_dec(x_56); -lean_inc(x_1); -x_62 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_61, x_7, x_8, x_9, x_10, x_60); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_62; +lean_object* x_554; +x_554 = lean_ctor_get(x_547, 0); +lean_dec(x_554); +lean_ctor_set(x_547, 0, x_1); +return x_546; } else { -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_55, 1); -lean_inc(x_63); -lean_dec(x_55); -x_64 = lean_ctor_get(x_56, 1); -lean_inc(x_64); -lean_dec(x_56); -x_65 = lean_ctor_get(x_57, 0); -lean_inc(x_65); -lean_dec(x_57); -x_66 = l_Lean_mkAppN(x_2, x_65); -lean_dec(x_65); -x_67 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_66, x_5, x_64, x_7, x_8, x_9, x_10, x_63); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_67; +lean_object* x_555; lean_object* x_556; +x_555 = lean_ctor_get(x_547, 1); +lean_inc(x_555); +lean_dec(x_547); +x_556 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_556, 0, x_1); +lean_ctor_set(x_556, 1, x_555); +lean_ctor_set(x_546, 0, x_556); +return x_546; +} +} +else +{ +lean_object* x_557; lean_object* x_558; lean_object* x_559; lean_object* x_560; lean_object* x_561; +x_557 = lean_ctor_get(x_546, 1); +lean_inc(x_557); +lean_dec(x_546); +x_558 = lean_ctor_get(x_547, 1); +lean_inc(x_558); +if (lean_is_exclusive(x_547)) { + lean_ctor_release(x_547, 0); + lean_ctor_release(x_547, 1); + x_559 = x_547; +} else { + lean_dec_ref(x_547); + x_559 = lean_box(0); +} +if (lean_is_scalar(x_559)) { + x_560 = lean_alloc_ctor(0, 2, 0); +} else { + x_560 = x_559; +} +lean_ctor_set(x_560, 0, x_1); +lean_ctor_set(x_560, 1, x_558); +x_561 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_561, 0, x_560); +lean_ctor_set(x_561, 1, x_557); +return x_561; +} +} +else +{ +uint8_t x_562; +lean_dec(x_1); +x_562 = !lean_is_exclusive(x_546); +if (x_562 == 0) +{ +lean_object* x_563; uint8_t x_564; +x_563 = lean_ctor_get(x_546, 0); +lean_dec(x_563); +x_564 = !lean_is_exclusive(x_547); +if (x_564 == 0) +{ +lean_object* x_565; lean_object* x_566; lean_object* x_567; +x_565 = lean_ctor_get(x_547, 0); +lean_dec(x_565); +x_566 = lean_ctor_get(x_548, 0); +lean_inc(x_566); +lean_dec(x_548); +x_567 = l_Lean_mkAppN(x_6, x_566); +lean_dec(x_566); +lean_ctor_set(x_547, 0, x_567); +return x_546; +} +else +{ +lean_object* x_568; lean_object* x_569; lean_object* x_570; lean_object* x_571; +x_568 = lean_ctor_get(x_547, 1); +lean_inc(x_568); +lean_dec(x_547); +x_569 = lean_ctor_get(x_548, 0); +lean_inc(x_569); +lean_dec(x_548); +x_570 = l_Lean_mkAppN(x_6, x_569); +lean_dec(x_569); +x_571 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_571, 0, x_570); +lean_ctor_set(x_571, 1, x_568); +lean_ctor_set(x_546, 0, x_571); +return x_546; +} +} +else +{ +lean_object* x_572; lean_object* x_573; lean_object* x_574; lean_object* x_575; lean_object* x_576; lean_object* x_577; lean_object* x_578; +x_572 = lean_ctor_get(x_546, 1); +lean_inc(x_572); +lean_dec(x_546); +x_573 = lean_ctor_get(x_547, 1); +lean_inc(x_573); +if (lean_is_exclusive(x_547)) { + lean_ctor_release(x_547, 0); + lean_ctor_release(x_547, 1); + x_574 = x_547; +} else { + lean_dec_ref(x_547); + x_574 = lean_box(0); +} +x_575 = lean_ctor_get(x_548, 0); +lean_inc(x_575); +lean_dec(x_548); +x_576 = l_Lean_mkAppN(x_6, x_575); +lean_dec(x_575); +if (lean_is_scalar(x_574)) { + x_577 = lean_alloc_ctor(0, 2, 0); +} else { + x_577 = x_574; +} +lean_ctor_set(x_577, 0, x_576); +lean_ctor_set(x_577, 1, x_573); +x_578 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_578, 0, x_577); +lean_ctor_set(x_578, 1, x_572); +return x_578; +} } } else { -uint8_t x_68; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -lean_dec(x_2); +uint8_t x_579; +lean_dec(x_6); lean_dec(x_1); -x_68 = !lean_is_exclusive(x_55); -if (x_68 == 0) +x_579 = !lean_is_exclusive(x_546); +if (x_579 == 0) { -return x_55; +return x_546; } else { -lean_object* x_69; lean_object* x_70; lean_object* x_71; -x_69 = lean_ctor_get(x_55, 0); -x_70 = lean_ctor_get(x_55, 1); -lean_inc(x_70); -lean_inc(x_69); -lean_dec(x_55); -x_71 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_71, 0, x_69); -lean_ctor_set(x_71, 1, x_70); -return x_71; +lean_object* x_580; lean_object* x_581; lean_object* x_582; +x_580 = lean_ctor_get(x_546, 0); +x_581 = lean_ctor_get(x_546, 1); +lean_inc(x_581); +lean_inc(x_580); +lean_dec(x_546); +x_582 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_582, 0, x_580); +lean_ctor_set(x_582, 1, x_581); +return x_582; } } } else { -uint8_t x_72; +uint8_t x_583; +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); -lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); +lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_72 = !lean_is_exclusive(x_44); -if (x_72 == 0) +x_583 = !lean_is_exclusive(x_534); +if (x_583 == 0) { -return x_44; +return x_534; } else { -lean_object* x_73; lean_object* x_74; lean_object* x_75; -x_73 = lean_ctor_get(x_44, 0); -x_74 = lean_ctor_get(x_44, 1); -lean_inc(x_74); -lean_inc(x_73); -lean_dec(x_44); -x_75 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_75, 0, x_73); -lean_ctor_set(x_75, 1, x_74); -return x_75; +lean_object* x_584; lean_object* x_585; lean_object* x_586; +x_584 = lean_ctor_get(x_534, 0); +x_585 = lean_ctor_get(x_534, 1); +lean_inc(x_585); +lean_inc(x_584); +lean_dec(x_534); +x_586 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_586, 0, x_584); +lean_ctor_set(x_586, 1, x_585); +return x_586; } } } } -case 1: +case 4: { -lean_object* x_89; lean_object* x_120; lean_object* x_154; uint8_t x_155; -lean_dec(x_4); -x_154 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__4; -x_155 = l_Lean_Expr_isConstOf(x_2, x_154); -if (x_155 == 0) +lean_object* x_600; lean_object* x_679; lean_object* x_734; uint8_t x_735; +lean_dec(x_8); +x_734 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__11___closed__5; +x_735 = l_Lean_Expr_isConstOf(x_6, x_734); +if (x_735 == 0) { -lean_object* x_156; -x_156 = lean_box(0); -x_120 = x_156; -goto block_153; +lean_object* x_736; +x_736 = lean_box(0); +x_679 = x_736; +goto block_733; } else { -lean_object* x_157; lean_object* x_158; uint8_t x_159; -x_157 = lean_array_get_size(x_3); -x_158 = lean_unsigned_to_nat(2u); -x_159 = lean_nat_dec_eq(x_157, x_158); -if (x_159 == 0) +lean_object* x_737; lean_object* x_738; uint8_t x_739; +x_737 = lean_array_get_size(x_7); +x_738 = lean_unsigned_to_nat(2u); +x_739 = lean_nat_dec_eq(x_737, x_738); +if (x_739 == 0) { -lean_object* x_160; -lean_dec(x_157); -x_160 = lean_box(0); -x_120 = x_160; -goto block_153; +lean_object* x_740; +lean_dec(x_737); +x_740 = lean_box(0); +x_679 = x_740; +goto block_733; } else { -lean_object* x_161; uint8_t x_162; -x_161 = lean_unsigned_to_nat(0u); -x_162 = lean_nat_dec_lt(x_161, x_157); -lean_dec(x_157); -if (x_162 == 0) +lean_object* x_741; uint8_t x_742; +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_741 = lean_unsigned_to_nat(0u); +x_742 = lean_nat_dec_lt(x_741, x_737); +lean_dec(x_737); +if (x_742 == 0) { -lean_object* x_163; lean_object* x_164; -x_163 = l_Lean_instInhabitedExpr; -x_164 = l_outOfBounds___rarg(x_163); -x_89 = x_164; -goto block_119; +lean_object* x_743; lean_object* x_744; +x_743 = l_Lean_instInhabitedExpr; +x_744 = l_outOfBounds___rarg(x_743); +x_600 = x_744; +goto block_678; } else { -lean_object* x_165; -x_165 = lean_array_fget(x_3, x_161); -x_89 = x_165; -goto block_119; +lean_object* x_745; +x_745 = lean_array_fget(x_7, x_741); +x_600 = x_745; +goto block_678; } } } -block_119: +block_678: { -lean_object* x_90; -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_5); -lean_inc(x_89); -x_90 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_89, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_90) == 0) +lean_object* x_601; +lean_inc(x_600); +x_601 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_600, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +if (lean_obj_tag(x_601) == 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); -lean_dec(x_90); -x_93 = lean_ctor_get(x_91, 0); -lean_inc(x_93); -x_94 = lean_ctor_get(x_91, 1); -lean_inc(x_94); -lean_dec(x_91); -x_95 = lean_ctor_get(x_94, 2); -lean_inc(x_95); -lean_inc(x_95); -x_96 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_95, x_93); -if (lean_obj_tag(x_96) == 0) +uint8_t x_602; +x_602 = !lean_is_exclusive(x_601); +if (x_602 == 0) { -size_t x_97; size_t x_98; uint8_t x_99; -x_97 = lean_ptr_addr(x_89); -lean_dec(x_89); -x_98 = lean_ptr_addr(x_93); -x_99 = lean_usize_dec_eq(x_97, x_98); -if (x_99 == 0) +lean_object* x_603; uint8_t x_604; +x_603 = lean_ctor_get(x_601, 0); +x_604 = !lean_is_exclusive(x_603); +if (x_604 == 0) { -lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; -x_100 = lean_unsigned_to_nat(0u); -lean_inc(x_93); -x_101 = lean_array_set(x_3, x_100, x_93); -x_102 = l_Lean_mkAppN(x_2, x_101); -lean_dec(x_101); -x_103 = lean_ctor_get(x_94, 0); -lean_inc(x_103); -x_104 = lean_ctor_get(x_94, 1); -lean_inc(x_104); -lean_dec(x_94); -lean_inc(x_102); -x_105 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_95, x_93, x_102); -x_106 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_106, 0, x_103); -lean_ctor_set(x_106, 1, x_104); -lean_ctor_set(x_106, 2, x_105); -x_107 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_102, x_5, x_106, x_7, x_8, x_9, x_10, x_92); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_107; +lean_object* x_605; lean_object* x_606; lean_object* x_607; lean_object* x_608; +x_605 = lean_ctor_get(x_603, 0); +x_606 = lean_ctor_get(x_603, 1); +x_607 = lean_ctor_get(x_606, 2); +lean_inc(x_607); +lean_inc(x_607); +x_608 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_607, x_605); +if (lean_obj_tag(x_608) == 0) +{ +size_t x_609; size_t x_610; uint8_t x_611; +x_609 = lean_ptr_addr(x_600); +lean_dec(x_600); +x_610 = lean_ptr_addr(x_605); +x_611 = lean_usize_dec_eq(x_609, x_610); +if (x_611 == 0) +{ +lean_object* x_612; lean_object* x_613; lean_object* x_614; lean_object* x_615; lean_object* x_616; lean_object* x_617; lean_object* x_618; +lean_dec(x_1); +x_612 = lean_unsigned_to_nat(0u); +lean_inc(x_605); +x_613 = lean_array_set(x_7, x_612, x_605); +x_614 = l_Lean_mkAppN(x_6, x_613); +lean_dec(x_613); +x_615 = lean_ctor_get(x_606, 0); +lean_inc(x_615); +x_616 = lean_ctor_get(x_606, 1); +lean_inc(x_616); +lean_dec(x_606); +lean_inc(x_614); +x_617 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_607, x_605, x_614); +x_618 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_618, 0, x_615); +lean_ctor_set(x_618, 1, x_616); +lean_ctor_set(x_618, 2, x_617); +lean_ctor_set(x_603, 1, x_618); +lean_ctor_set(x_603, 0, x_614); +return x_601; } else { -lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; -lean_dec(x_3); -lean_dec(x_2); -x_108 = lean_ctor_get(x_94, 0); -lean_inc(x_108); -x_109 = lean_ctor_get(x_94, 1); -lean_inc(x_109); -lean_dec(x_94); +lean_object* x_619; lean_object* x_620; lean_object* x_621; lean_object* x_622; +lean_dec(x_7); +lean_dec(x_6); +x_619 = lean_ctor_get(x_606, 0); +lean_inc(x_619); +x_620 = lean_ctor_get(x_606, 1); +lean_inc(x_620); +lean_dec(x_606); lean_inc(x_1); -x_110 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_95, x_93, x_1); -x_111 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_111, 0, x_108); -lean_ctor_set(x_111, 1, x_109); -lean_ctor_set(x_111, 2, x_110); +x_621 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_607, x_605, x_1); +x_622 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_622, 0, x_619); +lean_ctor_set(x_622, 1, x_620); +lean_ctor_set(x_622, 2, x_621); +lean_ctor_set(x_603, 1, x_622); +lean_ctor_set(x_603, 0, x_1); +return x_601; +} +} +else +{ +lean_object* x_623; +lean_dec(x_607); +lean_dec(x_605); +lean_dec(x_600); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_1); +x_623 = lean_ctor_get(x_608, 0); +lean_inc(x_623); +lean_dec(x_608); +lean_ctor_set(x_603, 0, x_623); +return x_601; +} +} +else +{ +lean_object* x_624; lean_object* x_625; lean_object* x_626; lean_object* x_627; +x_624 = lean_ctor_get(x_603, 0); +x_625 = lean_ctor_get(x_603, 1); +lean_inc(x_625); +lean_inc(x_624); +lean_dec(x_603); +x_626 = lean_ctor_get(x_625, 2); +lean_inc(x_626); +lean_inc(x_626); +x_627 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_626, x_624); +if (lean_obj_tag(x_627) == 0) +{ +size_t x_628; size_t x_629; uint8_t x_630; +x_628 = lean_ptr_addr(x_600); +lean_dec(x_600); +x_629 = lean_ptr_addr(x_624); +x_630 = lean_usize_dec_eq(x_628, x_629); +if (x_630 == 0) +{ +lean_object* x_631; lean_object* x_632; lean_object* x_633; lean_object* x_634; lean_object* x_635; lean_object* x_636; lean_object* x_637; lean_object* x_638; +lean_dec(x_1); +x_631 = lean_unsigned_to_nat(0u); +lean_inc(x_624); +x_632 = lean_array_set(x_7, x_631, x_624); +x_633 = l_Lean_mkAppN(x_6, x_632); +lean_dec(x_632); +x_634 = lean_ctor_get(x_625, 0); +lean_inc(x_634); +x_635 = lean_ctor_get(x_625, 1); +lean_inc(x_635); +lean_dec(x_625); +lean_inc(x_633); +x_636 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_626, x_624, x_633); +x_637 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_637, 0, x_634); +lean_ctor_set(x_637, 1, x_635); +lean_ctor_set(x_637, 2, x_636); +x_638 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_638, 0, x_633); +lean_ctor_set(x_638, 1, x_637); +lean_ctor_set(x_601, 0, x_638); +return x_601; +} +else +{ +lean_object* x_639; lean_object* x_640; lean_object* x_641; lean_object* x_642; lean_object* x_643; +lean_dec(x_7); +lean_dec(x_6); +x_639 = lean_ctor_get(x_625, 0); +lean_inc(x_639); +x_640 = lean_ctor_get(x_625, 1); +lean_inc(x_640); +lean_dec(x_625); lean_inc(x_1); -x_112 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_111, x_7, x_8, x_9, x_10, x_92); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); +x_641 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_626, x_624, x_1); +x_642 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_642, 0, x_639); +lean_ctor_set(x_642, 1, x_640); +lean_ctor_set(x_642, 2, x_641); +x_643 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_643, 0, x_1); +lean_ctor_set(x_643, 1, x_642); +lean_ctor_set(x_601, 0, x_643); +return x_601; +} +} +else +{ +lean_object* x_644; lean_object* x_645; +lean_dec(x_626); +lean_dec(x_624); +lean_dec(x_600); lean_dec(x_7); -lean_dec(x_5); -return x_112; +lean_dec(x_6); +lean_dec(x_1); +x_644 = lean_ctor_get(x_627, 0); +lean_inc(x_644); +lean_dec(x_627); +x_645 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_645, 0, x_644); +lean_ctor_set(x_645, 1, x_625); +lean_ctor_set(x_601, 0, x_645); +return x_601; +} } } else { -lean_object* x_113; lean_object* x_114; -lean_dec(x_95); -lean_dec(x_93); -lean_dec(x_89); -lean_dec(x_3); -lean_dec(x_2); -x_113 = lean_ctor_get(x_96, 0); -lean_inc(x_113); -lean_dec(x_96); -x_114 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_113, x_5, x_94, x_7, x_8, x_9, x_10, x_92); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); +lean_object* x_646; lean_object* x_647; lean_object* x_648; lean_object* x_649; lean_object* x_650; lean_object* x_651; lean_object* x_652; +x_646 = lean_ctor_get(x_601, 0); +x_647 = lean_ctor_get(x_601, 1); +lean_inc(x_647); +lean_inc(x_646); +lean_dec(x_601); +x_648 = lean_ctor_get(x_646, 0); +lean_inc(x_648); +x_649 = lean_ctor_get(x_646, 1); +lean_inc(x_649); +if (lean_is_exclusive(x_646)) { + lean_ctor_release(x_646, 0); + lean_ctor_release(x_646, 1); + x_650 = x_646; +} else { + lean_dec_ref(x_646); + x_650 = lean_box(0); +} +x_651 = lean_ctor_get(x_649, 2); +lean_inc(x_651); +lean_inc(x_651); +x_652 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_651, x_648); +if (lean_obj_tag(x_652) == 0) +{ +size_t x_653; size_t x_654; uint8_t x_655; +x_653 = lean_ptr_addr(x_600); +lean_dec(x_600); +x_654 = lean_ptr_addr(x_648); +x_655 = lean_usize_dec_eq(x_653, x_654); +if (x_655 == 0) +{ +lean_object* x_656; lean_object* x_657; lean_object* x_658; lean_object* x_659; lean_object* x_660; lean_object* x_661; lean_object* x_662; lean_object* x_663; lean_object* x_664; +lean_dec(x_1); +x_656 = lean_unsigned_to_nat(0u); +lean_inc(x_648); +x_657 = lean_array_set(x_7, x_656, x_648); +x_658 = l_Lean_mkAppN(x_6, x_657); +lean_dec(x_657); +x_659 = lean_ctor_get(x_649, 0); +lean_inc(x_659); +x_660 = lean_ctor_get(x_649, 1); +lean_inc(x_660); +lean_dec(x_649); +lean_inc(x_658); +x_661 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_651, x_648, x_658); +x_662 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_662, 0, x_659); +lean_ctor_set(x_662, 1, x_660); +lean_ctor_set(x_662, 2, x_661); +if (lean_is_scalar(x_650)) { + x_663 = lean_alloc_ctor(0, 2, 0); +} else { + x_663 = x_650; +} +lean_ctor_set(x_663, 0, x_658); +lean_ctor_set(x_663, 1, x_662); +x_664 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_664, 0, x_663); +lean_ctor_set(x_664, 1, x_647); +return x_664; +} +else +{ +lean_object* x_665; lean_object* x_666; lean_object* x_667; lean_object* x_668; lean_object* x_669; lean_object* x_670; lean_dec(x_7); -lean_dec(x_5); -return x_114; +lean_dec(x_6); +x_665 = lean_ctor_get(x_649, 0); +lean_inc(x_665); +x_666 = lean_ctor_get(x_649, 1); +lean_inc(x_666); +lean_dec(x_649); +lean_inc(x_1); +x_667 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_651, x_648, x_1); +x_668 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_668, 0, x_665); +lean_ctor_set(x_668, 1, x_666); +lean_ctor_set(x_668, 2, x_667); +if (lean_is_scalar(x_650)) { + x_669 = lean_alloc_ctor(0, 2, 0); +} else { + x_669 = x_650; +} +lean_ctor_set(x_669, 0, x_1); +lean_ctor_set(x_669, 1, x_668); +x_670 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_670, 0, x_669); +lean_ctor_set(x_670, 1, x_647); +return x_670; } } else { -uint8_t x_115; -lean_dec(x_89); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); +lean_object* x_671; lean_object* x_672; lean_object* x_673; +lean_dec(x_651); +lean_dec(x_648); +lean_dec(x_600); lean_dec(x_7); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); +lean_dec(x_6); lean_dec(x_1); -x_115 = !lean_is_exclusive(x_90); -if (x_115 == 0) +x_671 = lean_ctor_get(x_652, 0); +lean_inc(x_671); +lean_dec(x_652); +if (lean_is_scalar(x_650)) { + x_672 = lean_alloc_ctor(0, 2, 0); +} else { + x_672 = x_650; +} +lean_ctor_set(x_672, 0, x_671); +lean_ctor_set(x_672, 1, x_649); +x_673 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_673, 0, x_672); +lean_ctor_set(x_673, 1, x_647); +return x_673; +} +} +} +else +{ +uint8_t x_674; +lean_dec(x_600); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_1); +x_674 = !lean_is_exclusive(x_601); +if (x_674 == 0) { -return x_90; +return x_601; } else { -lean_object* x_116; lean_object* x_117; lean_object* x_118; -x_116 = lean_ctor_get(x_90, 0); -x_117 = lean_ctor_get(x_90, 1); -lean_inc(x_117); -lean_inc(x_116); -lean_dec(x_90); -x_118 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_118, 0, x_116); -lean_ctor_set(x_118, 1, x_117); -return x_118; +lean_object* x_675; lean_object* x_676; lean_object* x_677; +x_675 = lean_ctor_get(x_601, 0); +x_676 = lean_ctor_get(x_601, 1); +lean_inc(x_676); +lean_inc(x_675); +lean_dec(x_601); +x_677 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_677, 0, x_675); +lean_ctor_set(x_677, 1, x_676); +return x_677; } } } -block_153: +block_733: { -lean_object* x_121; -lean_dec(x_120); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_2); -x_121 = l_Lean_Meta_getFunInfo(x_2, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_121) == 0) +lean_object* x_680; +lean_dec(x_679); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_6); +x_680 = l_Lean_Meta_getFunInfo(x_6, x_11, x_12, x_13, x_14, x_15); +if (lean_obj_tag(x_680) == 0) { -lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; uint8_t x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; -x_122 = lean_ctor_get(x_121, 0); -lean_inc(x_122); -x_123 = lean_ctor_get(x_121, 1); -lean_inc(x_123); -lean_dec(x_121); -x_124 = lean_ctor_get(x_122, 0); -lean_inc(x_124); -lean_dec(x_122); -x_125 = lean_array_get_size(x_3); -x_126 = lean_unsigned_to_nat(0u); -x_127 = lean_unsigned_to_nat(1u); -x_128 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_128, 0, x_126); -lean_ctor_set(x_128, 1, x_125); -lean_ctor_set(x_128, 2, x_127); -x_129 = 0; -x_130 = lean_box(x_129); -x_131 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_131, 0, x_3); -lean_ctor_set(x_131, 1, x_130); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); +lean_object* x_681; lean_object* x_682; lean_object* x_683; lean_object* x_684; lean_object* x_685; lean_object* x_686; lean_object* x_687; uint8_t x_688; lean_object* x_689; lean_object* x_690; lean_object* x_691; lean_object* x_692; +x_681 = lean_ctor_get(x_680, 0); +lean_inc(x_681); +x_682 = lean_ctor_get(x_680, 1); +lean_inc(x_682); +lean_dec(x_680); +x_683 = lean_ctor_get(x_681, 0); +lean_inc(x_683); +lean_dec(x_681); +x_684 = lean_array_get_size(x_7); +x_685 = lean_unsigned_to_nat(0u); +x_686 = lean_unsigned_to_nat(1u); +lean_inc(x_684); +x_687 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_687, 0, x_685); +lean_ctor_set(x_687, 1, x_684); +lean_ctor_set(x_687, 2, x_686); +x_688 = 0; +x_689 = lean_box(x_688); lean_inc(x_7); -lean_inc(x_5); -lean_inc(x_2); -x_132 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8(x_2, x_124, x_128, x_128, x_131, x_126, lean_box(0), lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10, x_123); -lean_dec(x_128); -lean_dec(x_124); -if (lean_obj_tag(x_132) == 0) +x_690 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_690, 0, x_7); +lean_ctor_set(x_690, 1, x_689); +x_691 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__11___closed__1; +lean_inc(x_687); +lean_inc(x_6); +x_692 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10(x_2, x_3, x_4, x_5, x_6, x_7, x_691, x_683, x_684, x_687, x_687, x_690, x_685, lean_box(0), lean_box(0), x_9, x_10, x_11, x_12, x_13, x_14, x_682); +if (lean_obj_tag(x_692) == 0) +{ +lean_object* x_693; lean_object* x_694; lean_object* x_695; uint8_t x_696; +x_693 = lean_ctor_get(x_692, 0); +lean_inc(x_693); +x_694 = lean_ctor_get(x_693, 0); +lean_inc(x_694); +x_695 = lean_ctor_get(x_694, 1); +lean_inc(x_695); +x_696 = lean_unbox(x_695); +lean_dec(x_695); +if (x_696 == 0) { -lean_object* x_133; lean_object* x_134; lean_object* x_135; uint8_t x_136; -x_133 = lean_ctor_get(x_132, 0); -lean_inc(x_133); -x_134 = lean_ctor_get(x_133, 0); -lean_inc(x_134); -x_135 = lean_ctor_get(x_134, 1); -lean_inc(x_135); -x_136 = lean_unbox(x_135); -lean_dec(x_135); -if (x_136 == 0) +uint8_t x_697; +lean_dec(x_694); +lean_dec(x_6); +x_697 = !lean_is_exclusive(x_692); +if (x_697 == 0) +{ +lean_object* x_698; uint8_t x_699; +x_698 = lean_ctor_get(x_692, 0); +lean_dec(x_698); +x_699 = !lean_is_exclusive(x_693); +if (x_699 == 0) +{ +lean_object* x_700; +x_700 = lean_ctor_get(x_693, 0); +lean_dec(x_700); +lean_ctor_set(x_693, 0, x_1); +return x_692; +} +else +{ +lean_object* x_701; lean_object* x_702; +x_701 = lean_ctor_get(x_693, 1); +lean_inc(x_701); +lean_dec(x_693); +x_702 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_702, 0, x_1); +lean_ctor_set(x_702, 1, x_701); +lean_ctor_set(x_692, 0, x_702); +return x_692; +} +} +else +{ +lean_object* x_703; lean_object* x_704; lean_object* x_705; lean_object* x_706; lean_object* x_707; +x_703 = lean_ctor_get(x_692, 1); +lean_inc(x_703); +lean_dec(x_692); +x_704 = lean_ctor_get(x_693, 1); +lean_inc(x_704); +if (lean_is_exclusive(x_693)) { + lean_ctor_release(x_693, 0); + lean_ctor_release(x_693, 1); + x_705 = x_693; +} else { + lean_dec_ref(x_693); + x_705 = lean_box(0); +} +if (lean_is_scalar(x_705)) { + x_706 = lean_alloc_ctor(0, 2, 0); +} else { + x_706 = x_705; +} +lean_ctor_set(x_706, 0, x_1); +lean_ctor_set(x_706, 1, x_704); +x_707 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_707, 0, x_706); +lean_ctor_set(x_707, 1, x_703); +return x_707; +} +} +else { -lean_object* x_137; lean_object* x_138; lean_object* x_139; -lean_dec(x_134); -lean_dec(x_2); -x_137 = lean_ctor_get(x_132, 1); -lean_inc(x_137); -lean_dec(x_132); -x_138 = lean_ctor_get(x_133, 1); -lean_inc(x_138); -lean_dec(x_133); -lean_inc(x_1); -x_139 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_138, x_7, x_8, x_9, x_10, x_137); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_139; +uint8_t x_708; +lean_dec(x_1); +x_708 = !lean_is_exclusive(x_692); +if (x_708 == 0) +{ +lean_object* x_709; uint8_t x_710; +x_709 = lean_ctor_get(x_692, 0); +lean_dec(x_709); +x_710 = !lean_is_exclusive(x_693); +if (x_710 == 0) +{ +lean_object* x_711; lean_object* x_712; lean_object* x_713; +x_711 = lean_ctor_get(x_693, 0); +lean_dec(x_711); +x_712 = lean_ctor_get(x_694, 0); +lean_inc(x_712); +lean_dec(x_694); +x_713 = l_Lean_mkAppN(x_6, x_712); +lean_dec(x_712); +lean_ctor_set(x_693, 0, x_713); +return x_692; } else { -lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; -x_140 = lean_ctor_get(x_132, 1); -lean_inc(x_140); -lean_dec(x_132); -x_141 = lean_ctor_get(x_133, 1); -lean_inc(x_141); -lean_dec(x_133); -x_142 = lean_ctor_get(x_134, 0); -lean_inc(x_142); -lean_dec(x_134); -x_143 = l_Lean_mkAppN(x_2, x_142); -lean_dec(x_142); -x_144 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_143, x_5, x_141, x_7, x_8, x_9, x_10, x_140); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_144; +lean_object* x_714; lean_object* x_715; lean_object* x_716; lean_object* x_717; +x_714 = lean_ctor_get(x_693, 1); +lean_inc(x_714); +lean_dec(x_693); +x_715 = lean_ctor_get(x_694, 0); +lean_inc(x_715); +lean_dec(x_694); +x_716 = l_Lean_mkAppN(x_6, x_715); +lean_dec(x_715); +x_717 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_717, 0, x_716); +lean_ctor_set(x_717, 1, x_714); +lean_ctor_set(x_692, 0, x_717); +return x_692; } } else { -uint8_t x_145; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -lean_dec(x_2); +lean_object* x_718; lean_object* x_719; lean_object* x_720; lean_object* x_721; lean_object* x_722; lean_object* x_723; lean_object* x_724; +x_718 = lean_ctor_get(x_692, 1); +lean_inc(x_718); +lean_dec(x_692); +x_719 = lean_ctor_get(x_693, 1); +lean_inc(x_719); +if (lean_is_exclusive(x_693)) { + lean_ctor_release(x_693, 0); + lean_ctor_release(x_693, 1); + x_720 = x_693; +} else { + lean_dec_ref(x_693); + x_720 = lean_box(0); +} +x_721 = lean_ctor_get(x_694, 0); +lean_inc(x_721); +lean_dec(x_694); +x_722 = l_Lean_mkAppN(x_6, x_721); +lean_dec(x_721); +if (lean_is_scalar(x_720)) { + x_723 = lean_alloc_ctor(0, 2, 0); +} else { + x_723 = x_720; +} +lean_ctor_set(x_723, 0, x_722); +lean_ctor_set(x_723, 1, x_719); +x_724 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_724, 0, x_723); +lean_ctor_set(x_724, 1, x_718); +return x_724; +} +} +} +else +{ +uint8_t x_725; +lean_dec(x_6); lean_dec(x_1); -x_145 = !lean_is_exclusive(x_132); -if (x_145 == 0) +x_725 = !lean_is_exclusive(x_692); +if (x_725 == 0) { -return x_132; +return x_692; } else { -lean_object* x_146; lean_object* x_147; lean_object* x_148; -x_146 = lean_ctor_get(x_132, 0); -x_147 = lean_ctor_get(x_132, 1); -lean_inc(x_147); -lean_inc(x_146); -lean_dec(x_132); -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; +lean_object* x_726; lean_object* x_727; lean_object* x_728; +x_726 = lean_ctor_get(x_692, 0); +x_727 = lean_ctor_get(x_692, 1); +lean_inc(x_727); +lean_inc(x_726); +lean_dec(x_692); +x_728 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_728, 0, x_726); +lean_ctor_set(x_728, 1, x_727); +return x_728; } } } else { -uint8_t x_149; +uint8_t x_729; +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); -lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); +lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_149 = !lean_is_exclusive(x_121); -if (x_149 == 0) +x_729 = !lean_is_exclusive(x_680); +if (x_729 == 0) { -return x_121; +return x_680; } else { -lean_object* x_150; lean_object* x_151; lean_object* x_152; -x_150 = lean_ctor_get(x_121, 0); -x_151 = lean_ctor_get(x_121, 1); -lean_inc(x_151); -lean_inc(x_150); -lean_dec(x_121); -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_730; lean_object* x_731; lean_object* x_732; +x_730 = lean_ctor_get(x_680, 0); +x_731 = lean_ctor_get(x_680, 1); +lean_inc(x_731); +lean_inc(x_730); +lean_dec(x_680); +x_732 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_732, 0, x_730); +lean_ctor_set(x_732, 1, x_731); +return x_732; } } } } -case 2: +case 5: { -lean_object* x_166; lean_object* x_197; lean_object* x_231; uint8_t x_232; -lean_dec(x_4); -x_231 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__4; -x_232 = l_Lean_Expr_isConstOf(x_2, x_231); -if (x_232 == 0) +lean_object* x_746; lean_object* x_747; lean_object* x_748; lean_object* x_749; lean_object* x_750; +x_746 = lean_ctor_get(x_6, 0); +lean_inc(x_746); +x_747 = lean_ctor_get(x_6, 1); +lean_inc(x_747); +lean_dec(x_6); +x_748 = lean_array_set(x_7, x_8, x_747); +x_749 = lean_unsigned_to_nat(1u); +x_750 = lean_nat_sub(x_8, x_749); +lean_dec(x_8); +x_6 = x_746; +x_7 = x_748; +x_8 = x_750; +goto _start; +} +case 6: +{ +lean_object* x_752; lean_object* x_831; lean_object* x_886; uint8_t x_887; +lean_dec(x_8); +x_886 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__11___closed__5; +x_887 = l_Lean_Expr_isConstOf(x_6, x_886); +if (x_887 == 0) { -lean_object* x_233; -x_233 = lean_box(0); -x_197 = x_233; -goto block_230; +lean_object* x_888; +x_888 = lean_box(0); +x_831 = x_888; +goto block_885; } else { -lean_object* x_234; lean_object* x_235; uint8_t x_236; -x_234 = lean_array_get_size(x_3); -x_235 = lean_unsigned_to_nat(2u); -x_236 = lean_nat_dec_eq(x_234, x_235); -if (x_236 == 0) +lean_object* x_889; lean_object* x_890; uint8_t x_891; +x_889 = lean_array_get_size(x_7); +x_890 = lean_unsigned_to_nat(2u); +x_891 = lean_nat_dec_eq(x_889, x_890); +if (x_891 == 0) { -lean_object* x_237; -lean_dec(x_234); -x_237 = lean_box(0); -x_197 = x_237; -goto block_230; +lean_object* x_892; +lean_dec(x_889); +x_892 = lean_box(0); +x_831 = x_892; +goto block_885; } else { -lean_object* x_238; uint8_t x_239; -x_238 = lean_unsigned_to_nat(0u); -x_239 = lean_nat_dec_lt(x_238, x_234); -lean_dec(x_234); -if (x_239 == 0) +lean_object* x_893; uint8_t x_894; +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_893 = lean_unsigned_to_nat(0u); +x_894 = lean_nat_dec_lt(x_893, x_889); +lean_dec(x_889); +if (x_894 == 0) { -lean_object* x_240; lean_object* x_241; -x_240 = l_Lean_instInhabitedExpr; -x_241 = l_outOfBounds___rarg(x_240); -x_166 = x_241; -goto block_196; +lean_object* x_895; lean_object* x_896; +x_895 = l_Lean_instInhabitedExpr; +x_896 = l_outOfBounds___rarg(x_895); +x_752 = x_896; +goto block_830; } else { -lean_object* x_242; -x_242 = lean_array_fget(x_3, x_238); -x_166 = x_242; -goto block_196; +lean_object* x_897; +x_897 = lean_array_fget(x_7, x_893); +x_752 = x_897; +goto block_830; } } } -block_196: +block_830: { -lean_object* x_167; -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_5); -lean_inc(x_166); -x_167 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_166, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_167) == 0) +lean_object* x_753; +lean_inc(x_752); +x_753 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_752, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +if (lean_obj_tag(x_753) == 0) { -lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; -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_168, 0); -lean_inc(x_170); -x_171 = lean_ctor_get(x_168, 1); -lean_inc(x_171); -lean_dec(x_168); -x_172 = lean_ctor_get(x_171, 2); -lean_inc(x_172); -lean_inc(x_172); -x_173 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_172, x_170); -if (lean_obj_tag(x_173) == 0) -{ -size_t x_174; size_t x_175; uint8_t x_176; -x_174 = lean_ptr_addr(x_166); -lean_dec(x_166); -x_175 = lean_ptr_addr(x_170); -x_176 = lean_usize_dec_eq(x_174, x_175); -if (x_176 == 0) -{ -lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; -x_177 = lean_unsigned_to_nat(0u); -lean_inc(x_170); -x_178 = lean_array_set(x_3, x_177, x_170); -x_179 = l_Lean_mkAppN(x_2, x_178); -lean_dec(x_178); -x_180 = lean_ctor_get(x_171, 0); -lean_inc(x_180); -x_181 = lean_ctor_get(x_171, 1); -lean_inc(x_181); -lean_dec(x_171); -lean_inc(x_179); -x_182 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_172, x_170, x_179); -x_183 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_183, 0, x_180); -lean_ctor_set(x_183, 1, x_181); -lean_ctor_set(x_183, 2, x_182); -x_184 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_179, x_5, x_183, x_7, x_8, x_9, x_10, x_169); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_184; +uint8_t x_754; +x_754 = !lean_is_exclusive(x_753); +if (x_754 == 0) +{ +lean_object* x_755; uint8_t x_756; +x_755 = lean_ctor_get(x_753, 0); +x_756 = !lean_is_exclusive(x_755); +if (x_756 == 0) +{ +lean_object* x_757; lean_object* x_758; lean_object* x_759; lean_object* x_760; +x_757 = lean_ctor_get(x_755, 0); +x_758 = lean_ctor_get(x_755, 1); +x_759 = lean_ctor_get(x_758, 2); +lean_inc(x_759); +lean_inc(x_759); +x_760 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_759, x_757); +if (lean_obj_tag(x_760) == 0) +{ +size_t x_761; size_t x_762; uint8_t x_763; +x_761 = lean_ptr_addr(x_752); +lean_dec(x_752); +x_762 = lean_ptr_addr(x_757); +x_763 = lean_usize_dec_eq(x_761, x_762); +if (x_763 == 0) +{ +lean_object* x_764; lean_object* x_765; lean_object* x_766; lean_object* x_767; lean_object* x_768; lean_object* x_769; lean_object* x_770; +lean_dec(x_1); +x_764 = lean_unsigned_to_nat(0u); +lean_inc(x_757); +x_765 = lean_array_set(x_7, x_764, x_757); +x_766 = l_Lean_mkAppN(x_6, x_765); +lean_dec(x_765); +x_767 = lean_ctor_get(x_758, 0); +lean_inc(x_767); +x_768 = lean_ctor_get(x_758, 1); +lean_inc(x_768); +lean_dec(x_758); +lean_inc(x_766); +x_769 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_759, x_757, x_766); +x_770 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_770, 0, x_767); +lean_ctor_set(x_770, 1, x_768); +lean_ctor_set(x_770, 2, x_769); +lean_ctor_set(x_755, 1, x_770); +lean_ctor_set(x_755, 0, x_766); +return x_753; } else { -lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; -lean_dec(x_3); -lean_dec(x_2); -x_185 = lean_ctor_get(x_171, 0); -lean_inc(x_185); -x_186 = lean_ctor_get(x_171, 1); -lean_inc(x_186); -lean_dec(x_171); +lean_object* x_771; lean_object* x_772; lean_object* x_773; lean_object* x_774; +lean_dec(x_7); +lean_dec(x_6); +x_771 = lean_ctor_get(x_758, 0); +lean_inc(x_771); +x_772 = lean_ctor_get(x_758, 1); +lean_inc(x_772); +lean_dec(x_758); lean_inc(x_1); -x_187 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_172, x_170, x_1); -x_188 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_188, 0, x_185); -lean_ctor_set(x_188, 1, x_186); -lean_ctor_set(x_188, 2, x_187); +x_773 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_759, x_757, x_1); +x_774 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_774, 0, x_771); +lean_ctor_set(x_774, 1, x_772); +lean_ctor_set(x_774, 2, x_773); +lean_ctor_set(x_755, 1, x_774); +lean_ctor_set(x_755, 0, x_1); +return x_753; +} +} +else +{ +lean_object* x_775; +lean_dec(x_759); +lean_dec(x_757); +lean_dec(x_752); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_1); +x_775 = lean_ctor_get(x_760, 0); +lean_inc(x_775); +lean_dec(x_760); +lean_ctor_set(x_755, 0, x_775); +return x_753; +} +} +else +{ +lean_object* x_776; lean_object* x_777; lean_object* x_778; lean_object* x_779; +x_776 = lean_ctor_get(x_755, 0); +x_777 = lean_ctor_get(x_755, 1); +lean_inc(x_777); +lean_inc(x_776); +lean_dec(x_755); +x_778 = lean_ctor_get(x_777, 2); +lean_inc(x_778); +lean_inc(x_778); +x_779 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_778, x_776); +if (lean_obj_tag(x_779) == 0) +{ +size_t x_780; size_t x_781; uint8_t x_782; +x_780 = lean_ptr_addr(x_752); +lean_dec(x_752); +x_781 = lean_ptr_addr(x_776); +x_782 = lean_usize_dec_eq(x_780, x_781); +if (x_782 == 0) +{ +lean_object* x_783; lean_object* x_784; lean_object* x_785; lean_object* x_786; lean_object* x_787; lean_object* x_788; lean_object* x_789; lean_object* x_790; +lean_dec(x_1); +x_783 = lean_unsigned_to_nat(0u); +lean_inc(x_776); +x_784 = lean_array_set(x_7, x_783, x_776); +x_785 = l_Lean_mkAppN(x_6, x_784); +lean_dec(x_784); +x_786 = lean_ctor_get(x_777, 0); +lean_inc(x_786); +x_787 = lean_ctor_get(x_777, 1); +lean_inc(x_787); +lean_dec(x_777); +lean_inc(x_785); +x_788 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_778, x_776, x_785); +x_789 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_789, 0, x_786); +lean_ctor_set(x_789, 1, x_787); +lean_ctor_set(x_789, 2, x_788); +x_790 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_790, 0, x_785); +lean_ctor_set(x_790, 1, x_789); +lean_ctor_set(x_753, 0, x_790); +return x_753; +} +else +{ +lean_object* x_791; lean_object* x_792; lean_object* x_793; lean_object* x_794; lean_object* x_795; +lean_dec(x_7); +lean_dec(x_6); +x_791 = lean_ctor_get(x_777, 0); +lean_inc(x_791); +x_792 = lean_ctor_get(x_777, 1); +lean_inc(x_792); +lean_dec(x_777); lean_inc(x_1); -x_189 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_188, x_7, x_8, x_9, x_10, x_169); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); +x_793 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_778, x_776, x_1); +x_794 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_794, 0, x_791); +lean_ctor_set(x_794, 1, x_792); +lean_ctor_set(x_794, 2, x_793); +x_795 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_795, 0, x_1); +lean_ctor_set(x_795, 1, x_794); +lean_ctor_set(x_753, 0, x_795); +return x_753; +} +} +else +{ +lean_object* x_796; lean_object* x_797; +lean_dec(x_778); +lean_dec(x_776); +lean_dec(x_752); lean_dec(x_7); -lean_dec(x_5); -return x_189; +lean_dec(x_6); +lean_dec(x_1); +x_796 = lean_ctor_get(x_779, 0); +lean_inc(x_796); +lean_dec(x_779); +x_797 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_797, 0, x_796); +lean_ctor_set(x_797, 1, x_777); +lean_ctor_set(x_753, 0, x_797); +return x_753; +} } } else { -lean_object* x_190; lean_object* x_191; -lean_dec(x_172); -lean_dec(x_170); -lean_dec(x_166); -lean_dec(x_3); -lean_dec(x_2); -x_190 = lean_ctor_get(x_173, 0); -lean_inc(x_190); -lean_dec(x_173); -x_191 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_190, x_5, x_171, x_7, x_8, x_9, x_10, x_169); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); +lean_object* x_798; lean_object* x_799; lean_object* x_800; lean_object* x_801; lean_object* x_802; lean_object* x_803; lean_object* x_804; +x_798 = lean_ctor_get(x_753, 0); +x_799 = lean_ctor_get(x_753, 1); +lean_inc(x_799); +lean_inc(x_798); +lean_dec(x_753); +x_800 = lean_ctor_get(x_798, 0); +lean_inc(x_800); +x_801 = lean_ctor_get(x_798, 1); +lean_inc(x_801); +if (lean_is_exclusive(x_798)) { + lean_ctor_release(x_798, 0); + lean_ctor_release(x_798, 1); + x_802 = x_798; +} else { + lean_dec_ref(x_798); + x_802 = lean_box(0); +} +x_803 = lean_ctor_get(x_801, 2); +lean_inc(x_803); +lean_inc(x_803); +x_804 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_803, x_800); +if (lean_obj_tag(x_804) == 0) +{ +size_t x_805; size_t x_806; uint8_t x_807; +x_805 = lean_ptr_addr(x_752); +lean_dec(x_752); +x_806 = lean_ptr_addr(x_800); +x_807 = lean_usize_dec_eq(x_805, x_806); +if (x_807 == 0) +{ +lean_object* x_808; lean_object* x_809; lean_object* x_810; lean_object* x_811; lean_object* x_812; lean_object* x_813; lean_object* x_814; lean_object* x_815; lean_object* x_816; +lean_dec(x_1); +x_808 = lean_unsigned_to_nat(0u); +lean_inc(x_800); +x_809 = lean_array_set(x_7, x_808, x_800); +x_810 = l_Lean_mkAppN(x_6, x_809); +lean_dec(x_809); +x_811 = lean_ctor_get(x_801, 0); +lean_inc(x_811); +x_812 = lean_ctor_get(x_801, 1); +lean_inc(x_812); +lean_dec(x_801); +lean_inc(x_810); +x_813 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_803, x_800, x_810); +x_814 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_814, 0, x_811); +lean_ctor_set(x_814, 1, x_812); +lean_ctor_set(x_814, 2, x_813); +if (lean_is_scalar(x_802)) { + x_815 = lean_alloc_ctor(0, 2, 0); +} else { + x_815 = x_802; +} +lean_ctor_set(x_815, 0, x_810); +lean_ctor_set(x_815, 1, x_814); +x_816 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_816, 0, x_815); +lean_ctor_set(x_816, 1, x_799); +return x_816; +} +else +{ +lean_object* x_817; lean_object* x_818; lean_object* x_819; lean_object* x_820; lean_object* x_821; lean_object* x_822; lean_dec(x_7); -lean_dec(x_5); -return x_191; +lean_dec(x_6); +x_817 = lean_ctor_get(x_801, 0); +lean_inc(x_817); +x_818 = lean_ctor_get(x_801, 1); +lean_inc(x_818); +lean_dec(x_801); +lean_inc(x_1); +x_819 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_803, x_800, x_1); +x_820 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_820, 0, x_817); +lean_ctor_set(x_820, 1, x_818); +lean_ctor_set(x_820, 2, x_819); +if (lean_is_scalar(x_802)) { + x_821 = lean_alloc_ctor(0, 2, 0); +} else { + x_821 = x_802; +} +lean_ctor_set(x_821, 0, x_1); +lean_ctor_set(x_821, 1, x_820); +x_822 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_822, 0, x_821); +lean_ctor_set(x_822, 1, x_799); +return x_822; } } else { -uint8_t x_192; -lean_dec(x_166); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); +lean_object* x_823; lean_object* x_824; lean_object* x_825; +lean_dec(x_803); +lean_dec(x_800); +lean_dec(x_752); lean_dec(x_7); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); +lean_dec(x_6); lean_dec(x_1); -x_192 = !lean_is_exclusive(x_167); -if (x_192 == 0) +x_823 = lean_ctor_get(x_804, 0); +lean_inc(x_823); +lean_dec(x_804); +if (lean_is_scalar(x_802)) { + x_824 = lean_alloc_ctor(0, 2, 0); +} else { + x_824 = x_802; +} +lean_ctor_set(x_824, 0, x_823); +lean_ctor_set(x_824, 1, x_801); +x_825 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_825, 0, x_824); +lean_ctor_set(x_825, 1, x_799); +return x_825; +} +} +} +else +{ +uint8_t x_826; +lean_dec(x_752); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_1); +x_826 = !lean_is_exclusive(x_753); +if (x_826 == 0) { -return x_167; +return x_753; } else { -lean_object* x_193; lean_object* x_194; lean_object* x_195; -x_193 = lean_ctor_get(x_167, 0); -x_194 = lean_ctor_get(x_167, 1); -lean_inc(x_194); -lean_inc(x_193); -lean_dec(x_167); -x_195 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_195, 0, x_193); -lean_ctor_set(x_195, 1, x_194); -return x_195; +lean_object* x_827; lean_object* x_828; lean_object* x_829; +x_827 = lean_ctor_get(x_753, 0); +x_828 = lean_ctor_get(x_753, 1); +lean_inc(x_828); +lean_inc(x_827); +lean_dec(x_753); +x_829 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_829, 0, x_827); +lean_ctor_set(x_829, 1, x_828); +return x_829; } } } -block_230: +block_885: { -lean_object* x_198; -lean_dec(x_197); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_2); -x_198 = l_Lean_Meta_getFunInfo(x_2, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_198) == 0) -{ -lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; uint8_t x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; -x_199 = lean_ctor_get(x_198, 0); -lean_inc(x_199); -x_200 = lean_ctor_get(x_198, 1); -lean_inc(x_200); -lean_dec(x_198); -x_201 = lean_ctor_get(x_199, 0); -lean_inc(x_201); -lean_dec(x_199); -x_202 = lean_array_get_size(x_3); -x_203 = lean_unsigned_to_nat(0u); -x_204 = lean_unsigned_to_nat(1u); -x_205 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_205, 0, x_203); -lean_ctor_set(x_205, 1, x_202); -lean_ctor_set(x_205, 2, x_204); -x_206 = 0; -x_207 = lean_box(x_206); -x_208 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_208, 0, x_3); -lean_ctor_set(x_208, 1, x_207); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); +lean_object* x_832; +lean_dec(x_831); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_6); +x_832 = l_Lean_Meta_getFunInfo(x_6, x_11, x_12, x_13, x_14, x_15); +if (lean_obj_tag(x_832) == 0) +{ +lean_object* x_833; lean_object* x_834; lean_object* x_835; lean_object* x_836; lean_object* x_837; lean_object* x_838; lean_object* x_839; uint8_t x_840; lean_object* x_841; lean_object* x_842; lean_object* x_843; lean_object* x_844; +x_833 = lean_ctor_get(x_832, 0); +lean_inc(x_833); +x_834 = lean_ctor_get(x_832, 1); +lean_inc(x_834); +lean_dec(x_832); +x_835 = lean_ctor_get(x_833, 0); +lean_inc(x_835); +lean_dec(x_833); +x_836 = lean_array_get_size(x_7); +x_837 = lean_unsigned_to_nat(0u); +x_838 = lean_unsigned_to_nat(1u); +lean_inc(x_836); +x_839 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_839, 0, x_837); +lean_ctor_set(x_839, 1, x_836); +lean_ctor_set(x_839, 2, x_838); +x_840 = 0; +x_841 = lean_box(x_840); lean_inc(x_7); -lean_inc(x_5); -lean_inc(x_2); -x_209 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8(x_2, x_201, x_205, x_205, x_208, x_203, lean_box(0), lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10, x_200); -lean_dec(x_205); -lean_dec(x_201); -if (lean_obj_tag(x_209) == 0) +x_842 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_842, 0, x_7); +lean_ctor_set(x_842, 1, x_841); +x_843 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__11___closed__1; +lean_inc(x_839); +lean_inc(x_6); +x_844 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10(x_2, x_3, x_4, x_5, x_6, x_7, x_843, x_835, x_836, x_839, x_839, x_842, x_837, lean_box(0), lean_box(0), x_9, x_10, x_11, x_12, x_13, x_14, x_834); +if (lean_obj_tag(x_844) == 0) { -lean_object* x_210; lean_object* x_211; lean_object* x_212; uint8_t x_213; -x_210 = lean_ctor_get(x_209, 0); -lean_inc(x_210); -x_211 = lean_ctor_get(x_210, 0); -lean_inc(x_211); -x_212 = lean_ctor_get(x_211, 1); -lean_inc(x_212); -x_213 = lean_unbox(x_212); -lean_dec(x_212); -if (x_213 == 0) +lean_object* x_845; lean_object* x_846; lean_object* x_847; uint8_t x_848; +x_845 = lean_ctor_get(x_844, 0); +lean_inc(x_845); +x_846 = lean_ctor_get(x_845, 0); +lean_inc(x_846); +x_847 = lean_ctor_get(x_846, 1); +lean_inc(x_847); +x_848 = lean_unbox(x_847); +lean_dec(x_847); +if (x_848 == 0) { -lean_object* x_214; lean_object* x_215; lean_object* x_216; -lean_dec(x_211); -lean_dec(x_2); -x_214 = lean_ctor_get(x_209, 1); -lean_inc(x_214); -lean_dec(x_209); -x_215 = lean_ctor_get(x_210, 1); -lean_inc(x_215); -lean_dec(x_210); -lean_inc(x_1); -x_216 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_215, x_7, x_8, x_9, x_10, x_214); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_216; +uint8_t x_849; +lean_dec(x_846); +lean_dec(x_6); +x_849 = !lean_is_exclusive(x_844); +if (x_849 == 0) +{ +lean_object* x_850; uint8_t x_851; +x_850 = lean_ctor_get(x_844, 0); +lean_dec(x_850); +x_851 = !lean_is_exclusive(x_845); +if (x_851 == 0) +{ +lean_object* x_852; +x_852 = lean_ctor_get(x_845, 0); +lean_dec(x_852); +lean_ctor_set(x_845, 0, x_1); +return x_844; +} +else +{ +lean_object* x_853; lean_object* x_854; +x_853 = lean_ctor_get(x_845, 1); +lean_inc(x_853); +lean_dec(x_845); +x_854 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_854, 0, x_1); +lean_ctor_set(x_854, 1, x_853); +lean_ctor_set(x_844, 0, x_854); +return x_844; +} +} +else +{ +lean_object* x_855; lean_object* x_856; lean_object* x_857; lean_object* x_858; lean_object* x_859; +x_855 = lean_ctor_get(x_844, 1); +lean_inc(x_855); +lean_dec(x_844); +x_856 = lean_ctor_get(x_845, 1); +lean_inc(x_856); +if (lean_is_exclusive(x_845)) { + lean_ctor_release(x_845, 0); + lean_ctor_release(x_845, 1); + x_857 = x_845; +} else { + lean_dec_ref(x_845); + x_857 = lean_box(0); +} +if (lean_is_scalar(x_857)) { + x_858 = lean_alloc_ctor(0, 2, 0); +} else { + x_858 = x_857; +} +lean_ctor_set(x_858, 0, x_1); +lean_ctor_set(x_858, 1, x_856); +x_859 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_859, 0, x_858); +lean_ctor_set(x_859, 1, x_855); +return x_859; +} } else { -lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; -x_217 = lean_ctor_get(x_209, 1); -lean_inc(x_217); -lean_dec(x_209); -x_218 = lean_ctor_get(x_210, 1); -lean_inc(x_218); -lean_dec(x_210); -x_219 = lean_ctor_get(x_211, 0); -lean_inc(x_219); -lean_dec(x_211); -x_220 = l_Lean_mkAppN(x_2, x_219); -lean_dec(x_219); -x_221 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_220, x_5, x_218, x_7, x_8, x_9, x_10, x_217); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_221; +uint8_t x_860; +lean_dec(x_1); +x_860 = !lean_is_exclusive(x_844); +if (x_860 == 0) +{ +lean_object* x_861; uint8_t x_862; +x_861 = lean_ctor_get(x_844, 0); +lean_dec(x_861); +x_862 = !lean_is_exclusive(x_845); +if (x_862 == 0) +{ +lean_object* x_863; lean_object* x_864; lean_object* x_865; +x_863 = lean_ctor_get(x_845, 0); +lean_dec(x_863); +x_864 = lean_ctor_get(x_846, 0); +lean_inc(x_864); +lean_dec(x_846); +x_865 = l_Lean_mkAppN(x_6, x_864); +lean_dec(x_864); +lean_ctor_set(x_845, 0, x_865); +return x_844; +} +else +{ +lean_object* x_866; lean_object* x_867; lean_object* x_868; lean_object* x_869; +x_866 = lean_ctor_get(x_845, 1); +lean_inc(x_866); +lean_dec(x_845); +x_867 = lean_ctor_get(x_846, 0); +lean_inc(x_867); +lean_dec(x_846); +x_868 = l_Lean_mkAppN(x_6, x_867); +lean_dec(x_867); +x_869 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_869, 0, x_868); +lean_ctor_set(x_869, 1, x_866); +lean_ctor_set(x_844, 0, x_869); +return x_844; +} +} +else +{ +lean_object* x_870; lean_object* x_871; lean_object* x_872; lean_object* x_873; lean_object* x_874; lean_object* x_875; lean_object* x_876; +x_870 = lean_ctor_get(x_844, 1); +lean_inc(x_870); +lean_dec(x_844); +x_871 = lean_ctor_get(x_845, 1); +lean_inc(x_871); +if (lean_is_exclusive(x_845)) { + lean_ctor_release(x_845, 0); + lean_ctor_release(x_845, 1); + x_872 = x_845; +} else { + lean_dec_ref(x_845); + x_872 = lean_box(0); +} +x_873 = lean_ctor_get(x_846, 0); +lean_inc(x_873); +lean_dec(x_846); +x_874 = l_Lean_mkAppN(x_6, x_873); +lean_dec(x_873); +if (lean_is_scalar(x_872)) { + x_875 = lean_alloc_ctor(0, 2, 0); +} else { + x_875 = x_872; +} +lean_ctor_set(x_875, 0, x_874); +lean_ctor_set(x_875, 1, x_871); +x_876 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_876, 0, x_875); +lean_ctor_set(x_876, 1, x_870); +return x_876; +} } } else { -uint8_t x_222; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -lean_dec(x_2); +uint8_t x_877; +lean_dec(x_6); lean_dec(x_1); -x_222 = !lean_is_exclusive(x_209); -if (x_222 == 0) +x_877 = !lean_is_exclusive(x_844); +if (x_877 == 0) { -return x_209; +return x_844; } else { -lean_object* x_223; lean_object* x_224; lean_object* x_225; -x_223 = lean_ctor_get(x_209, 0); -x_224 = lean_ctor_get(x_209, 1); -lean_inc(x_224); -lean_inc(x_223); -lean_dec(x_209); -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_878; lean_object* x_879; lean_object* x_880; +x_878 = lean_ctor_get(x_844, 0); +x_879 = lean_ctor_get(x_844, 1); +lean_inc(x_879); +lean_inc(x_878); +lean_dec(x_844); +x_880 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_880, 0, x_878); +lean_ctor_set(x_880, 1, x_879); +return x_880; } } } else { -uint8_t x_226; +uint8_t x_881; +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); -lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); +lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_226 = !lean_is_exclusive(x_198); -if (x_226 == 0) +x_881 = !lean_is_exclusive(x_832); +if (x_881 == 0) { -return x_198; +return x_832; } else { -lean_object* x_227; lean_object* x_228; lean_object* x_229; -x_227 = lean_ctor_get(x_198, 0); -x_228 = lean_ctor_get(x_198, 1); -lean_inc(x_228); -lean_inc(x_227); -lean_dec(x_198); -x_229 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_229, 0, x_227); -lean_ctor_set(x_229, 1, x_228); -return x_229; +lean_object* x_882; lean_object* x_883; lean_object* x_884; +x_882 = lean_ctor_get(x_832, 0); +x_883 = lean_ctor_get(x_832, 1); +lean_inc(x_883); +lean_inc(x_882); +lean_dec(x_832); +x_884 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_884, 0, x_882); +lean_ctor_set(x_884, 1, x_883); +return x_884; } } } } -case 3: +case 7: { -lean_object* x_243; lean_object* x_274; lean_object* x_308; uint8_t x_309; -lean_dec(x_4); -x_308 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__4; -x_309 = l_Lean_Expr_isConstOf(x_2, x_308); -if (x_309 == 0) +lean_object* x_898; lean_object* x_977; lean_object* x_1032; uint8_t x_1033; +lean_dec(x_8); +x_1032 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__11___closed__5; +x_1033 = l_Lean_Expr_isConstOf(x_6, x_1032); +if (x_1033 == 0) { -lean_object* x_310; -x_310 = lean_box(0); -x_274 = x_310; -goto block_307; +lean_object* x_1034; +x_1034 = lean_box(0); +x_977 = x_1034; +goto block_1031; } else { -lean_object* x_311; lean_object* x_312; uint8_t x_313; -x_311 = lean_array_get_size(x_3); -x_312 = lean_unsigned_to_nat(2u); -x_313 = lean_nat_dec_eq(x_311, x_312); -if (x_313 == 0) +lean_object* x_1035; lean_object* x_1036; uint8_t x_1037; +x_1035 = lean_array_get_size(x_7); +x_1036 = lean_unsigned_to_nat(2u); +x_1037 = lean_nat_dec_eq(x_1035, x_1036); +if (x_1037 == 0) { -lean_object* x_314; -lean_dec(x_311); -x_314 = lean_box(0); -x_274 = x_314; -goto block_307; +lean_object* x_1038; +lean_dec(x_1035); +x_1038 = lean_box(0); +x_977 = x_1038; +goto block_1031; } else { -lean_object* x_315; uint8_t x_316; -x_315 = lean_unsigned_to_nat(0u); -x_316 = lean_nat_dec_lt(x_315, x_311); -lean_dec(x_311); -if (x_316 == 0) +lean_object* x_1039; uint8_t x_1040; +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_1039 = lean_unsigned_to_nat(0u); +x_1040 = lean_nat_dec_lt(x_1039, x_1035); +lean_dec(x_1035); +if (x_1040 == 0) { -lean_object* x_317; lean_object* x_318; -x_317 = l_Lean_instInhabitedExpr; -x_318 = l_outOfBounds___rarg(x_317); -x_243 = x_318; -goto block_273; +lean_object* x_1041; lean_object* x_1042; +x_1041 = l_Lean_instInhabitedExpr; +x_1042 = l_outOfBounds___rarg(x_1041); +x_898 = x_1042; +goto block_976; } else { -lean_object* x_319; -x_319 = lean_array_fget(x_3, x_315); -x_243 = x_319; -goto block_273; +lean_object* x_1043; +x_1043 = lean_array_fget(x_7, x_1039); +x_898 = x_1043; +goto block_976; } } } -block_273: +block_976: { -lean_object* x_244; -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_5); -lean_inc(x_243); -x_244 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_243, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_244) == 0) +lean_object* x_899; +lean_inc(x_898); +x_899 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_898, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +if (lean_obj_tag(x_899) == 0) { -lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; -x_245 = lean_ctor_get(x_244, 0); -lean_inc(x_245); -x_246 = lean_ctor_get(x_244, 1); -lean_inc(x_246); -lean_dec(x_244); -x_247 = lean_ctor_get(x_245, 0); -lean_inc(x_247); -x_248 = lean_ctor_get(x_245, 1); -lean_inc(x_248); -lean_dec(x_245); -x_249 = lean_ctor_get(x_248, 2); -lean_inc(x_249); -lean_inc(x_249); -x_250 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_249, x_247); -if (lean_obj_tag(x_250) == 0) +uint8_t x_900; +x_900 = !lean_is_exclusive(x_899); +if (x_900 == 0) { -size_t x_251; size_t x_252; uint8_t x_253; -x_251 = lean_ptr_addr(x_243); -lean_dec(x_243); -x_252 = lean_ptr_addr(x_247); -x_253 = lean_usize_dec_eq(x_251, x_252); -if (x_253 == 0) -{ -lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; -x_254 = lean_unsigned_to_nat(0u); -lean_inc(x_247); -x_255 = lean_array_set(x_3, x_254, x_247); -x_256 = l_Lean_mkAppN(x_2, x_255); -lean_dec(x_255); -x_257 = lean_ctor_get(x_248, 0); -lean_inc(x_257); -x_258 = lean_ctor_get(x_248, 1); -lean_inc(x_258); -lean_dec(x_248); -lean_inc(x_256); -x_259 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_249, x_247, x_256); -x_260 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_260, 0, x_257); -lean_ctor_set(x_260, 1, x_258); -lean_ctor_set(x_260, 2, x_259); -x_261 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_256, x_5, x_260, x_7, x_8, x_9, x_10, x_246); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); +lean_object* x_901; uint8_t x_902; +x_901 = lean_ctor_get(x_899, 0); +x_902 = !lean_is_exclusive(x_901); +if (x_902 == 0) +{ +lean_object* x_903; lean_object* x_904; lean_object* x_905; lean_object* x_906; +x_903 = lean_ctor_get(x_901, 0); +x_904 = lean_ctor_get(x_901, 1); +x_905 = lean_ctor_get(x_904, 2); +lean_inc(x_905); +lean_inc(x_905); +x_906 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_905, x_903); +if (lean_obj_tag(x_906) == 0) +{ +size_t x_907; size_t x_908; uint8_t x_909; +x_907 = lean_ptr_addr(x_898); +lean_dec(x_898); +x_908 = lean_ptr_addr(x_903); +x_909 = lean_usize_dec_eq(x_907, x_908); +if (x_909 == 0) +{ +lean_object* x_910; lean_object* x_911; lean_object* x_912; lean_object* x_913; lean_object* x_914; lean_object* x_915; lean_object* x_916; +lean_dec(x_1); +x_910 = lean_unsigned_to_nat(0u); +lean_inc(x_903); +x_911 = lean_array_set(x_7, x_910, x_903); +x_912 = l_Lean_mkAppN(x_6, x_911); +lean_dec(x_911); +x_913 = lean_ctor_get(x_904, 0); +lean_inc(x_913); +x_914 = lean_ctor_get(x_904, 1); +lean_inc(x_914); +lean_dec(x_904); +lean_inc(x_912); +x_915 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_905, x_903, x_912); +x_916 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_916, 0, x_913); +lean_ctor_set(x_916, 1, x_914); +lean_ctor_set(x_916, 2, x_915); +lean_ctor_set(x_901, 1, x_916); +lean_ctor_set(x_901, 0, x_912); +return x_899; +} +else +{ +lean_object* x_917; lean_object* x_918; lean_object* x_919; lean_object* x_920; lean_dec(x_7); -lean_dec(x_5); -return x_261; +lean_dec(x_6); +x_917 = lean_ctor_get(x_904, 0); +lean_inc(x_917); +x_918 = lean_ctor_get(x_904, 1); +lean_inc(x_918); +lean_dec(x_904); +lean_inc(x_1); +x_919 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_905, x_903, x_1); +x_920 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_920, 0, x_917); +lean_ctor_set(x_920, 1, x_918); +lean_ctor_set(x_920, 2, x_919); +lean_ctor_set(x_901, 1, x_920); +lean_ctor_set(x_901, 0, x_1); +return x_899; +} } else { -lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; -lean_dec(x_3); -lean_dec(x_2); -x_262 = lean_ctor_get(x_248, 0); -lean_inc(x_262); -x_263 = lean_ctor_get(x_248, 1); -lean_inc(x_263); -lean_dec(x_248); -lean_inc(x_1); -x_264 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_249, x_247, x_1); -x_265 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_265, 0, x_262); -lean_ctor_set(x_265, 1, x_263); -lean_ctor_set(x_265, 2, x_264); +lean_object* x_921; +lean_dec(x_905); +lean_dec(x_903); +lean_dec(x_898); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_1); +x_921 = lean_ctor_get(x_906, 0); +lean_inc(x_921); +lean_dec(x_906); +lean_ctor_set(x_901, 0, x_921); +return x_899; +} +} +else +{ +lean_object* x_922; lean_object* x_923; lean_object* x_924; lean_object* x_925; +x_922 = lean_ctor_get(x_901, 0); +x_923 = lean_ctor_get(x_901, 1); +lean_inc(x_923); +lean_inc(x_922); +lean_dec(x_901); +x_924 = lean_ctor_get(x_923, 2); +lean_inc(x_924); +lean_inc(x_924); +x_925 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_924, x_922); +if (lean_obj_tag(x_925) == 0) +{ +size_t x_926; size_t x_927; uint8_t x_928; +x_926 = lean_ptr_addr(x_898); +lean_dec(x_898); +x_927 = lean_ptr_addr(x_922); +x_928 = lean_usize_dec_eq(x_926, x_927); +if (x_928 == 0) +{ +lean_object* x_929; lean_object* x_930; lean_object* x_931; lean_object* x_932; lean_object* x_933; lean_object* x_934; lean_object* x_935; lean_object* x_936; +lean_dec(x_1); +x_929 = lean_unsigned_to_nat(0u); +lean_inc(x_922); +x_930 = lean_array_set(x_7, x_929, x_922); +x_931 = l_Lean_mkAppN(x_6, x_930); +lean_dec(x_930); +x_932 = lean_ctor_get(x_923, 0); +lean_inc(x_932); +x_933 = lean_ctor_get(x_923, 1); +lean_inc(x_933); +lean_dec(x_923); +lean_inc(x_931); +x_934 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_924, x_922, x_931); +x_935 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_935, 0, x_932); +lean_ctor_set(x_935, 1, x_933); +lean_ctor_set(x_935, 2, x_934); +x_936 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_936, 0, x_931); +lean_ctor_set(x_936, 1, x_935); +lean_ctor_set(x_899, 0, x_936); +return x_899; +} +else +{ +lean_object* x_937; lean_object* x_938; lean_object* x_939; lean_object* x_940; lean_object* x_941; +lean_dec(x_7); +lean_dec(x_6); +x_937 = lean_ctor_get(x_923, 0); +lean_inc(x_937); +x_938 = lean_ctor_get(x_923, 1); +lean_inc(x_938); +lean_dec(x_923); lean_inc(x_1); -x_266 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_265, x_7, x_8, x_9, x_10, x_246); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); +x_939 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_924, x_922, x_1); +x_940 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_940, 0, x_937); +lean_ctor_set(x_940, 1, x_938); +lean_ctor_set(x_940, 2, x_939); +x_941 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_941, 0, x_1); +lean_ctor_set(x_941, 1, x_940); +lean_ctor_set(x_899, 0, x_941); +return x_899; +} +} +else +{ +lean_object* x_942; lean_object* x_943; +lean_dec(x_924); +lean_dec(x_922); +lean_dec(x_898); lean_dec(x_7); -lean_dec(x_5); -return x_266; +lean_dec(x_6); +lean_dec(x_1); +x_942 = lean_ctor_get(x_925, 0); +lean_inc(x_942); +lean_dec(x_925); +x_943 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_943, 0, x_942); +lean_ctor_set(x_943, 1, x_923); +lean_ctor_set(x_899, 0, x_943); +return x_899; +} +} +} +else +{ +lean_object* x_944; lean_object* x_945; lean_object* x_946; lean_object* x_947; lean_object* x_948; lean_object* x_949; lean_object* x_950; +x_944 = lean_ctor_get(x_899, 0); +x_945 = lean_ctor_get(x_899, 1); +lean_inc(x_945); +lean_inc(x_944); +lean_dec(x_899); +x_946 = lean_ctor_get(x_944, 0); +lean_inc(x_946); +x_947 = lean_ctor_get(x_944, 1); +lean_inc(x_947); +if (lean_is_exclusive(x_944)) { + lean_ctor_release(x_944, 0); + lean_ctor_release(x_944, 1); + x_948 = x_944; +} else { + lean_dec_ref(x_944); + x_948 = lean_box(0); +} +x_949 = lean_ctor_get(x_947, 2); +lean_inc(x_949); +lean_inc(x_949); +x_950 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_949, x_946); +if (lean_obj_tag(x_950) == 0) +{ +size_t x_951; size_t x_952; uint8_t x_953; +x_951 = lean_ptr_addr(x_898); +lean_dec(x_898); +x_952 = lean_ptr_addr(x_946); +x_953 = lean_usize_dec_eq(x_951, x_952); +if (x_953 == 0) +{ +lean_object* x_954; lean_object* x_955; lean_object* x_956; lean_object* x_957; lean_object* x_958; lean_object* x_959; lean_object* x_960; lean_object* x_961; lean_object* x_962; +lean_dec(x_1); +x_954 = lean_unsigned_to_nat(0u); +lean_inc(x_946); +x_955 = lean_array_set(x_7, x_954, x_946); +x_956 = l_Lean_mkAppN(x_6, x_955); +lean_dec(x_955); +x_957 = lean_ctor_get(x_947, 0); +lean_inc(x_957); +x_958 = lean_ctor_get(x_947, 1); +lean_inc(x_958); +lean_dec(x_947); +lean_inc(x_956); +x_959 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_949, x_946, x_956); +x_960 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_960, 0, x_957); +lean_ctor_set(x_960, 1, x_958); +lean_ctor_set(x_960, 2, x_959); +if (lean_is_scalar(x_948)) { + x_961 = lean_alloc_ctor(0, 2, 0); +} else { + x_961 = x_948; } +lean_ctor_set(x_961, 0, x_956); +lean_ctor_set(x_961, 1, x_960); +x_962 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_962, 0, x_961); +lean_ctor_set(x_962, 1, x_945); +return x_962; } else { -lean_object* x_267; lean_object* x_268; -lean_dec(x_249); -lean_dec(x_247); -lean_dec(x_243); -lean_dec(x_3); -lean_dec(x_2); -x_267 = lean_ctor_get(x_250, 0); -lean_inc(x_267); -lean_dec(x_250); -x_268 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_267, x_5, x_248, x_7, x_8, x_9, x_10, x_246); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); +lean_object* x_963; lean_object* x_964; lean_object* x_965; lean_object* x_966; lean_object* x_967; lean_object* x_968; lean_dec(x_7); -lean_dec(x_5); -return x_268; +lean_dec(x_6); +x_963 = lean_ctor_get(x_947, 0); +lean_inc(x_963); +x_964 = lean_ctor_get(x_947, 1); +lean_inc(x_964); +lean_dec(x_947); +lean_inc(x_1); +x_965 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_949, x_946, x_1); +x_966 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_966, 0, x_963); +lean_ctor_set(x_966, 1, x_964); +lean_ctor_set(x_966, 2, x_965); +if (lean_is_scalar(x_948)) { + x_967 = lean_alloc_ctor(0, 2, 0); +} else { + x_967 = x_948; +} +lean_ctor_set(x_967, 0, x_1); +lean_ctor_set(x_967, 1, x_966); +x_968 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_968, 0, x_967); +lean_ctor_set(x_968, 1, x_945); +return x_968; } } else { -uint8_t x_269; -lean_dec(x_243); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); +lean_object* x_969; lean_object* x_970; lean_object* x_971; +lean_dec(x_949); +lean_dec(x_946); +lean_dec(x_898); lean_dec(x_7); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); +lean_dec(x_6); +lean_dec(x_1); +x_969 = lean_ctor_get(x_950, 0); +lean_inc(x_969); +lean_dec(x_950); +if (lean_is_scalar(x_948)) { + x_970 = lean_alloc_ctor(0, 2, 0); +} else { + x_970 = x_948; +} +lean_ctor_set(x_970, 0, x_969); +lean_ctor_set(x_970, 1, x_947); +x_971 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_971, 0, x_970); +lean_ctor_set(x_971, 1, x_945); +return x_971; +} +} +} +else +{ +uint8_t x_972; +lean_dec(x_898); +lean_dec(x_7); +lean_dec(x_6); lean_dec(x_1); -x_269 = !lean_is_exclusive(x_244); -if (x_269 == 0) +x_972 = !lean_is_exclusive(x_899); +if (x_972 == 0) { -return x_244; +return x_899; } else { -lean_object* x_270; lean_object* x_271; lean_object* x_272; -x_270 = lean_ctor_get(x_244, 0); -x_271 = lean_ctor_get(x_244, 1); -lean_inc(x_271); -lean_inc(x_270); -lean_dec(x_244); -x_272 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_272, 0, x_270); -lean_ctor_set(x_272, 1, x_271); -return x_272; +lean_object* x_973; lean_object* x_974; lean_object* x_975; +x_973 = lean_ctor_get(x_899, 0); +x_974 = lean_ctor_get(x_899, 1); +lean_inc(x_974); +lean_inc(x_973); +lean_dec(x_899); +x_975 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_975, 0, x_973); +lean_ctor_set(x_975, 1, x_974); +return x_975; } } } -block_307: -{ -lean_object* x_275; -lean_dec(x_274); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_2); -x_275 = l_Lean_Meta_getFunInfo(x_2, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_275) == 0) +block_1031: { -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; uint8_t x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; -x_276 = lean_ctor_get(x_275, 0); -lean_inc(x_276); -x_277 = lean_ctor_get(x_275, 1); -lean_inc(x_277); -lean_dec(x_275); -x_278 = lean_ctor_get(x_276, 0); -lean_inc(x_278); -lean_dec(x_276); -x_279 = lean_array_get_size(x_3); -x_280 = lean_unsigned_to_nat(0u); -x_281 = lean_unsigned_to_nat(1u); -x_282 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_282, 0, x_280); -lean_ctor_set(x_282, 1, x_279); -lean_ctor_set(x_282, 2, x_281); -x_283 = 0; -x_284 = lean_box(x_283); -x_285 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_285, 0, x_3); -lean_ctor_set(x_285, 1, x_284); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); +lean_object* x_978; +lean_dec(x_977); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_6); +x_978 = l_Lean_Meta_getFunInfo(x_6, x_11, x_12, x_13, x_14, x_15); +if (lean_obj_tag(x_978) == 0) +{ +lean_object* x_979; lean_object* x_980; lean_object* x_981; lean_object* x_982; lean_object* x_983; lean_object* x_984; lean_object* x_985; uint8_t x_986; lean_object* x_987; lean_object* x_988; lean_object* x_989; lean_object* x_990; +x_979 = lean_ctor_get(x_978, 0); +lean_inc(x_979); +x_980 = lean_ctor_get(x_978, 1); +lean_inc(x_980); +lean_dec(x_978); +x_981 = lean_ctor_get(x_979, 0); +lean_inc(x_981); +lean_dec(x_979); +x_982 = lean_array_get_size(x_7); +x_983 = lean_unsigned_to_nat(0u); +x_984 = lean_unsigned_to_nat(1u); +lean_inc(x_982); +x_985 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_985, 0, x_983); +lean_ctor_set(x_985, 1, x_982); +lean_ctor_set(x_985, 2, x_984); +x_986 = 0; +x_987 = lean_box(x_986); lean_inc(x_7); -lean_inc(x_5); -lean_inc(x_2); -x_286 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8(x_2, x_278, x_282, x_282, x_285, x_280, lean_box(0), lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10, x_277); -lean_dec(x_282); -lean_dec(x_278); -if (lean_obj_tag(x_286) == 0) -{ -lean_object* x_287; lean_object* x_288; lean_object* x_289; uint8_t x_290; -x_287 = lean_ctor_get(x_286, 0); -lean_inc(x_287); -x_288 = lean_ctor_get(x_287, 0); -lean_inc(x_288); -x_289 = lean_ctor_get(x_288, 1); -lean_inc(x_289); -x_290 = lean_unbox(x_289); -lean_dec(x_289); -if (x_290 == 0) +x_988 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_988, 0, x_7); +lean_ctor_set(x_988, 1, x_987); +x_989 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__11___closed__1; +lean_inc(x_985); +lean_inc(x_6); +x_990 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10(x_2, x_3, x_4, x_5, x_6, x_7, x_989, x_981, x_982, x_985, x_985, x_988, x_983, lean_box(0), lean_box(0), x_9, x_10, x_11, x_12, x_13, x_14, x_980); +if (lean_obj_tag(x_990) == 0) +{ +lean_object* x_991; lean_object* x_992; lean_object* x_993; uint8_t x_994; +x_991 = lean_ctor_get(x_990, 0); +lean_inc(x_991); +x_992 = lean_ctor_get(x_991, 0); +lean_inc(x_992); +x_993 = lean_ctor_get(x_992, 1); +lean_inc(x_993); +x_994 = lean_unbox(x_993); +lean_dec(x_993); +if (x_994 == 0) +{ +uint8_t x_995; +lean_dec(x_992); +lean_dec(x_6); +x_995 = !lean_is_exclusive(x_990); +if (x_995 == 0) +{ +lean_object* x_996; uint8_t x_997; +x_996 = lean_ctor_get(x_990, 0); +lean_dec(x_996); +x_997 = !lean_is_exclusive(x_991); +if (x_997 == 0) +{ +lean_object* x_998; +x_998 = lean_ctor_get(x_991, 0); +lean_dec(x_998); +lean_ctor_set(x_991, 0, x_1); +return x_990; +} +else +{ +lean_object* x_999; lean_object* x_1000; +x_999 = lean_ctor_get(x_991, 1); +lean_inc(x_999); +lean_dec(x_991); +x_1000 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1000, 0, x_1); +lean_ctor_set(x_1000, 1, x_999); +lean_ctor_set(x_990, 0, x_1000); +return x_990; +} +} +else +{ +lean_object* x_1001; lean_object* x_1002; lean_object* x_1003; lean_object* x_1004; lean_object* x_1005; +x_1001 = lean_ctor_get(x_990, 1); +lean_inc(x_1001); +lean_dec(x_990); +x_1002 = lean_ctor_get(x_991, 1); +lean_inc(x_1002); +if (lean_is_exclusive(x_991)) { + lean_ctor_release(x_991, 0); + lean_ctor_release(x_991, 1); + x_1003 = x_991; +} else { + lean_dec_ref(x_991); + x_1003 = lean_box(0); +} +if (lean_is_scalar(x_1003)) { + x_1004 = lean_alloc_ctor(0, 2, 0); +} else { + x_1004 = x_1003; +} +lean_ctor_set(x_1004, 0, x_1); +lean_ctor_set(x_1004, 1, x_1002); +x_1005 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1005, 0, x_1004); +lean_ctor_set(x_1005, 1, x_1001); +return x_1005; +} +} +else { -lean_object* x_291; lean_object* x_292; lean_object* x_293; -lean_dec(x_288); -lean_dec(x_2); -x_291 = lean_ctor_get(x_286, 1); -lean_inc(x_291); -lean_dec(x_286); -x_292 = lean_ctor_get(x_287, 1); -lean_inc(x_292); -lean_dec(x_287); -lean_inc(x_1); -x_293 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_292, x_7, x_8, x_9, x_10, x_291); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_293; -} -else -{ -lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; -x_294 = lean_ctor_get(x_286, 1); -lean_inc(x_294); -lean_dec(x_286); -x_295 = lean_ctor_get(x_287, 1); -lean_inc(x_295); -lean_dec(x_287); -x_296 = lean_ctor_get(x_288, 0); -lean_inc(x_296); -lean_dec(x_288); -x_297 = l_Lean_mkAppN(x_2, x_296); -lean_dec(x_296); -x_298 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_297, x_5, x_295, x_7, x_8, x_9, x_10, x_294); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_298; +uint8_t x_1006; +lean_dec(x_1); +x_1006 = !lean_is_exclusive(x_990); +if (x_1006 == 0) +{ +lean_object* x_1007; uint8_t x_1008; +x_1007 = lean_ctor_get(x_990, 0); +lean_dec(x_1007); +x_1008 = !lean_is_exclusive(x_991); +if (x_1008 == 0) +{ +lean_object* x_1009; lean_object* x_1010; lean_object* x_1011; +x_1009 = lean_ctor_get(x_991, 0); +lean_dec(x_1009); +x_1010 = lean_ctor_get(x_992, 0); +lean_inc(x_1010); +lean_dec(x_992); +x_1011 = l_Lean_mkAppN(x_6, x_1010); +lean_dec(x_1010); +lean_ctor_set(x_991, 0, x_1011); +return x_990; +} +else +{ +lean_object* x_1012; lean_object* x_1013; lean_object* x_1014; lean_object* x_1015; +x_1012 = lean_ctor_get(x_991, 1); +lean_inc(x_1012); +lean_dec(x_991); +x_1013 = lean_ctor_get(x_992, 0); +lean_inc(x_1013); +lean_dec(x_992); +x_1014 = l_Lean_mkAppN(x_6, x_1013); +lean_dec(x_1013); +x_1015 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1015, 0, x_1014); +lean_ctor_set(x_1015, 1, x_1012); +lean_ctor_set(x_990, 0, x_1015); +return x_990; +} +} +else +{ +lean_object* x_1016; lean_object* x_1017; lean_object* x_1018; lean_object* x_1019; lean_object* x_1020; lean_object* x_1021; lean_object* x_1022; +x_1016 = lean_ctor_get(x_990, 1); +lean_inc(x_1016); +lean_dec(x_990); +x_1017 = lean_ctor_get(x_991, 1); +lean_inc(x_1017); +if (lean_is_exclusive(x_991)) { + lean_ctor_release(x_991, 0); + lean_ctor_release(x_991, 1); + x_1018 = x_991; +} else { + lean_dec_ref(x_991); + x_1018 = lean_box(0); +} +x_1019 = lean_ctor_get(x_992, 0); +lean_inc(x_1019); +lean_dec(x_992); +x_1020 = l_Lean_mkAppN(x_6, x_1019); +lean_dec(x_1019); +if (lean_is_scalar(x_1018)) { + x_1021 = lean_alloc_ctor(0, 2, 0); +} else { + x_1021 = x_1018; +} +lean_ctor_set(x_1021, 0, x_1020); +lean_ctor_set(x_1021, 1, x_1017); +x_1022 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1022, 0, x_1021); +lean_ctor_set(x_1022, 1, x_1016); +return x_1022; +} } } else { -uint8_t x_299; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -lean_dec(x_2); +uint8_t x_1023; +lean_dec(x_6); lean_dec(x_1); -x_299 = !lean_is_exclusive(x_286); -if (x_299 == 0) +x_1023 = !lean_is_exclusive(x_990); +if (x_1023 == 0) { -return x_286; +return x_990; } else { -lean_object* x_300; lean_object* x_301; lean_object* x_302; -x_300 = lean_ctor_get(x_286, 0); -x_301 = lean_ctor_get(x_286, 1); -lean_inc(x_301); -lean_inc(x_300); -lean_dec(x_286); -x_302 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_302, 0, x_300); -lean_ctor_set(x_302, 1, x_301); -return x_302; +lean_object* x_1024; lean_object* x_1025; lean_object* x_1026; +x_1024 = lean_ctor_get(x_990, 0); +x_1025 = lean_ctor_get(x_990, 1); +lean_inc(x_1025); +lean_inc(x_1024); +lean_dec(x_990); +x_1026 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1026, 0, x_1024); +lean_ctor_set(x_1026, 1, x_1025); +return x_1026; } } } else { -uint8_t x_303; +uint8_t x_1027; +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); -lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); +lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_303 = !lean_is_exclusive(x_275); -if (x_303 == 0) +x_1027 = !lean_is_exclusive(x_978); +if (x_1027 == 0) { -return x_275; +return x_978; } else { -lean_object* x_304; lean_object* x_305; lean_object* x_306; -x_304 = lean_ctor_get(x_275, 0); -x_305 = lean_ctor_get(x_275, 1); -lean_inc(x_305); -lean_inc(x_304); -lean_dec(x_275); -x_306 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_306, 0, x_304); -lean_ctor_set(x_306, 1, x_305); -return x_306; +lean_object* x_1028; lean_object* x_1029; lean_object* x_1030; +x_1028 = lean_ctor_get(x_978, 0); +x_1029 = lean_ctor_get(x_978, 1); +lean_inc(x_1029); +lean_inc(x_1028); +lean_dec(x_978); +x_1030 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1030, 0, x_1028); +lean_ctor_set(x_1030, 1, x_1029); +return x_1030; } } } } -case 4: +case 8: { -lean_object* x_320; lean_object* x_351; lean_object* x_385; uint8_t x_386; -lean_dec(x_4); -x_385 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__4; -x_386 = l_Lean_Expr_isConstOf(x_2, x_385); -if (x_386 == 0) +lean_object* x_1044; lean_object* x_1123; lean_object* x_1178; uint8_t x_1179; +lean_dec(x_8); +x_1178 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__11___closed__5; +x_1179 = l_Lean_Expr_isConstOf(x_6, x_1178); +if (x_1179 == 0) { -lean_object* x_387; -x_387 = lean_box(0); -x_351 = x_387; -goto block_384; +lean_object* x_1180; +x_1180 = lean_box(0); +x_1123 = x_1180; +goto block_1177; } else { -lean_object* x_388; lean_object* x_389; uint8_t x_390; -x_388 = lean_array_get_size(x_3); -x_389 = lean_unsigned_to_nat(2u); -x_390 = lean_nat_dec_eq(x_388, x_389); -if (x_390 == 0) +lean_object* x_1181; lean_object* x_1182; uint8_t x_1183; +x_1181 = lean_array_get_size(x_7); +x_1182 = lean_unsigned_to_nat(2u); +x_1183 = lean_nat_dec_eq(x_1181, x_1182); +if (x_1183 == 0) { -lean_object* x_391; -lean_dec(x_388); -x_391 = lean_box(0); -x_351 = x_391; -goto block_384; +lean_object* x_1184; +lean_dec(x_1181); +x_1184 = lean_box(0); +x_1123 = x_1184; +goto block_1177; } else { -lean_object* x_392; uint8_t x_393; -x_392 = lean_unsigned_to_nat(0u); -x_393 = lean_nat_dec_lt(x_392, x_388); -lean_dec(x_388); -if (x_393 == 0) +lean_object* x_1185; uint8_t x_1186; +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_1185 = lean_unsigned_to_nat(0u); +x_1186 = lean_nat_dec_lt(x_1185, x_1181); +lean_dec(x_1181); +if (x_1186 == 0) { -lean_object* x_394; lean_object* x_395; -x_394 = l_Lean_instInhabitedExpr; -x_395 = l_outOfBounds___rarg(x_394); -x_320 = x_395; -goto block_350; +lean_object* x_1187; lean_object* x_1188; +x_1187 = l_Lean_instInhabitedExpr; +x_1188 = l_outOfBounds___rarg(x_1187); +x_1044 = x_1188; +goto block_1122; } else { -lean_object* x_396; -x_396 = lean_array_fget(x_3, x_392); -x_320 = x_396; -goto block_350; +lean_object* x_1189; +x_1189 = lean_array_fget(x_7, x_1185); +x_1044 = x_1189; +goto block_1122; } } } -block_350: +block_1122: { -lean_object* x_321; -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_5); -lean_inc(x_320); -x_321 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_320, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_321) == 0) +lean_object* x_1045; +lean_inc(x_1044); +x_1045 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_1044, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +if (lean_obj_tag(x_1045) == 0) { -lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; -x_322 = lean_ctor_get(x_321, 0); -lean_inc(x_322); -x_323 = lean_ctor_get(x_321, 1); -lean_inc(x_323); -lean_dec(x_321); -x_324 = lean_ctor_get(x_322, 0); -lean_inc(x_324); -x_325 = lean_ctor_get(x_322, 1); -lean_inc(x_325); -lean_dec(x_322); -x_326 = lean_ctor_get(x_325, 2); -lean_inc(x_326); -lean_inc(x_326); -x_327 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_326, x_324); -if (lean_obj_tag(x_327) == 0) -{ -size_t x_328; size_t x_329; uint8_t x_330; -x_328 = lean_ptr_addr(x_320); -lean_dec(x_320); -x_329 = lean_ptr_addr(x_324); -x_330 = lean_usize_dec_eq(x_328, x_329); -if (x_330 == 0) -{ -lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; -x_331 = lean_unsigned_to_nat(0u); -lean_inc(x_324); -x_332 = lean_array_set(x_3, x_331, x_324); -x_333 = l_Lean_mkAppN(x_2, x_332); -lean_dec(x_332); -x_334 = lean_ctor_get(x_325, 0); -lean_inc(x_334); -x_335 = lean_ctor_get(x_325, 1); -lean_inc(x_335); -lean_dec(x_325); -lean_inc(x_333); -x_336 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_326, x_324, x_333); -x_337 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_337, 0, x_334); -lean_ctor_set(x_337, 1, x_335); -lean_ctor_set(x_337, 2, x_336); -x_338 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_333, x_5, x_337, x_7, x_8, x_9, x_10, x_323); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); +uint8_t x_1046; +x_1046 = !lean_is_exclusive(x_1045); +if (x_1046 == 0) +{ +lean_object* x_1047; uint8_t x_1048; +x_1047 = lean_ctor_get(x_1045, 0); +x_1048 = !lean_is_exclusive(x_1047); +if (x_1048 == 0) +{ +lean_object* x_1049; lean_object* x_1050; lean_object* x_1051; lean_object* x_1052; +x_1049 = lean_ctor_get(x_1047, 0); +x_1050 = lean_ctor_get(x_1047, 1); +x_1051 = lean_ctor_get(x_1050, 2); +lean_inc(x_1051); +lean_inc(x_1051); +x_1052 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_1051, x_1049); +if (lean_obj_tag(x_1052) == 0) +{ +size_t x_1053; size_t x_1054; uint8_t x_1055; +x_1053 = lean_ptr_addr(x_1044); +lean_dec(x_1044); +x_1054 = lean_ptr_addr(x_1049); +x_1055 = lean_usize_dec_eq(x_1053, x_1054); +if (x_1055 == 0) +{ +lean_object* x_1056; lean_object* x_1057; lean_object* x_1058; lean_object* x_1059; lean_object* x_1060; lean_object* x_1061; lean_object* x_1062; +lean_dec(x_1); +x_1056 = lean_unsigned_to_nat(0u); +lean_inc(x_1049); +x_1057 = lean_array_set(x_7, x_1056, x_1049); +x_1058 = l_Lean_mkAppN(x_6, x_1057); +lean_dec(x_1057); +x_1059 = lean_ctor_get(x_1050, 0); +lean_inc(x_1059); +x_1060 = lean_ctor_get(x_1050, 1); +lean_inc(x_1060); +lean_dec(x_1050); +lean_inc(x_1058); +x_1061 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_1051, x_1049, x_1058); +x_1062 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1062, 0, x_1059); +lean_ctor_set(x_1062, 1, x_1060); +lean_ctor_set(x_1062, 2, x_1061); +lean_ctor_set(x_1047, 1, x_1062); +lean_ctor_set(x_1047, 0, x_1058); +return x_1045; +} +else +{ +lean_object* x_1063; lean_object* x_1064; lean_object* x_1065; lean_object* x_1066; lean_dec(x_7); -lean_dec(x_5); -return x_338; +lean_dec(x_6); +x_1063 = lean_ctor_get(x_1050, 0); +lean_inc(x_1063); +x_1064 = lean_ctor_get(x_1050, 1); +lean_inc(x_1064); +lean_dec(x_1050); +lean_inc(x_1); +x_1065 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_1051, x_1049, x_1); +x_1066 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1066, 0, x_1063); +lean_ctor_set(x_1066, 1, x_1064); +lean_ctor_set(x_1066, 2, x_1065); +lean_ctor_set(x_1047, 1, x_1066); +lean_ctor_set(x_1047, 0, x_1); +return x_1045; +} } else { -lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; -lean_dec(x_3); -lean_dec(x_2); -x_339 = lean_ctor_get(x_325, 0); -lean_inc(x_339); -x_340 = lean_ctor_get(x_325, 1); -lean_inc(x_340); -lean_dec(x_325); -lean_inc(x_1); -x_341 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_326, x_324, x_1); -x_342 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_342, 0, x_339); -lean_ctor_set(x_342, 1, x_340); -lean_ctor_set(x_342, 2, x_341); +lean_object* x_1067; +lean_dec(x_1051); +lean_dec(x_1049); +lean_dec(x_1044); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_1); +x_1067 = lean_ctor_get(x_1052, 0); +lean_inc(x_1067); +lean_dec(x_1052); +lean_ctor_set(x_1047, 0, x_1067); +return x_1045; +} +} +else +{ +lean_object* x_1068; lean_object* x_1069; lean_object* x_1070; lean_object* x_1071; +x_1068 = lean_ctor_get(x_1047, 0); +x_1069 = lean_ctor_get(x_1047, 1); +lean_inc(x_1069); +lean_inc(x_1068); +lean_dec(x_1047); +x_1070 = lean_ctor_get(x_1069, 2); +lean_inc(x_1070); +lean_inc(x_1070); +x_1071 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_1070, x_1068); +if (lean_obj_tag(x_1071) == 0) +{ +size_t x_1072; size_t x_1073; uint8_t x_1074; +x_1072 = lean_ptr_addr(x_1044); +lean_dec(x_1044); +x_1073 = lean_ptr_addr(x_1068); +x_1074 = lean_usize_dec_eq(x_1072, x_1073); +if (x_1074 == 0) +{ +lean_object* x_1075; lean_object* x_1076; lean_object* x_1077; lean_object* x_1078; lean_object* x_1079; lean_object* x_1080; lean_object* x_1081; lean_object* x_1082; +lean_dec(x_1); +x_1075 = lean_unsigned_to_nat(0u); +lean_inc(x_1068); +x_1076 = lean_array_set(x_7, x_1075, x_1068); +x_1077 = l_Lean_mkAppN(x_6, x_1076); +lean_dec(x_1076); +x_1078 = lean_ctor_get(x_1069, 0); +lean_inc(x_1078); +x_1079 = lean_ctor_get(x_1069, 1); +lean_inc(x_1079); +lean_dec(x_1069); +lean_inc(x_1077); +x_1080 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_1070, x_1068, x_1077); +x_1081 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1081, 0, x_1078); +lean_ctor_set(x_1081, 1, x_1079); +lean_ctor_set(x_1081, 2, x_1080); +x_1082 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1082, 0, x_1077); +lean_ctor_set(x_1082, 1, x_1081); +lean_ctor_set(x_1045, 0, x_1082); +return x_1045; +} +else +{ +lean_object* x_1083; lean_object* x_1084; lean_object* x_1085; lean_object* x_1086; lean_object* x_1087; +lean_dec(x_7); +lean_dec(x_6); +x_1083 = lean_ctor_get(x_1069, 0); +lean_inc(x_1083); +x_1084 = lean_ctor_get(x_1069, 1); +lean_inc(x_1084); +lean_dec(x_1069); lean_inc(x_1); -x_343 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_342, x_7, x_8, x_9, x_10, x_323); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); +x_1085 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_1070, x_1068, x_1); +x_1086 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1086, 0, x_1083); +lean_ctor_set(x_1086, 1, x_1084); +lean_ctor_set(x_1086, 2, x_1085); +x_1087 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1087, 0, x_1); +lean_ctor_set(x_1087, 1, x_1086); +lean_ctor_set(x_1045, 0, x_1087); +return x_1045; +} +} +else +{ +lean_object* x_1088; lean_object* x_1089; +lean_dec(x_1070); +lean_dec(x_1068); +lean_dec(x_1044); lean_dec(x_7); -lean_dec(x_5); -return x_343; +lean_dec(x_6); +lean_dec(x_1); +x_1088 = lean_ctor_get(x_1071, 0); +lean_inc(x_1088); +lean_dec(x_1071); +x_1089 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1089, 0, x_1088); +lean_ctor_set(x_1089, 1, x_1069); +lean_ctor_set(x_1045, 0, x_1089); +return x_1045; +} +} +} +else +{ +lean_object* x_1090; lean_object* x_1091; lean_object* x_1092; lean_object* x_1093; lean_object* x_1094; lean_object* x_1095; lean_object* x_1096; +x_1090 = lean_ctor_get(x_1045, 0); +x_1091 = lean_ctor_get(x_1045, 1); +lean_inc(x_1091); +lean_inc(x_1090); +lean_dec(x_1045); +x_1092 = lean_ctor_get(x_1090, 0); +lean_inc(x_1092); +x_1093 = lean_ctor_get(x_1090, 1); +lean_inc(x_1093); +if (lean_is_exclusive(x_1090)) { + lean_ctor_release(x_1090, 0); + lean_ctor_release(x_1090, 1); + x_1094 = x_1090; +} else { + lean_dec_ref(x_1090); + x_1094 = lean_box(0); +} +x_1095 = lean_ctor_get(x_1093, 2); +lean_inc(x_1095); +lean_inc(x_1095); +x_1096 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_1095, x_1092); +if (lean_obj_tag(x_1096) == 0) +{ +size_t x_1097; size_t x_1098; uint8_t x_1099; +x_1097 = lean_ptr_addr(x_1044); +lean_dec(x_1044); +x_1098 = lean_ptr_addr(x_1092); +x_1099 = lean_usize_dec_eq(x_1097, x_1098); +if (x_1099 == 0) +{ +lean_object* x_1100; lean_object* x_1101; lean_object* x_1102; lean_object* x_1103; lean_object* x_1104; lean_object* x_1105; lean_object* x_1106; lean_object* x_1107; lean_object* x_1108; +lean_dec(x_1); +x_1100 = lean_unsigned_to_nat(0u); +lean_inc(x_1092); +x_1101 = lean_array_set(x_7, x_1100, x_1092); +x_1102 = l_Lean_mkAppN(x_6, x_1101); +lean_dec(x_1101); +x_1103 = lean_ctor_get(x_1093, 0); +lean_inc(x_1103); +x_1104 = lean_ctor_get(x_1093, 1); +lean_inc(x_1104); +lean_dec(x_1093); +lean_inc(x_1102); +x_1105 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_1095, x_1092, x_1102); +x_1106 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1106, 0, x_1103); +lean_ctor_set(x_1106, 1, x_1104); +lean_ctor_set(x_1106, 2, x_1105); +if (lean_is_scalar(x_1094)) { + x_1107 = lean_alloc_ctor(0, 2, 0); +} else { + x_1107 = x_1094; } +lean_ctor_set(x_1107, 0, x_1102); +lean_ctor_set(x_1107, 1, x_1106); +x_1108 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1108, 0, x_1107); +lean_ctor_set(x_1108, 1, x_1091); +return x_1108; } else { -lean_object* x_344; lean_object* x_345; -lean_dec(x_326); -lean_dec(x_324); -lean_dec(x_320); -lean_dec(x_3); -lean_dec(x_2); -x_344 = lean_ctor_get(x_327, 0); -lean_inc(x_344); -lean_dec(x_327); -x_345 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_344, x_5, x_325, x_7, x_8, x_9, x_10, x_323); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); +lean_object* x_1109; lean_object* x_1110; lean_object* x_1111; lean_object* x_1112; lean_object* x_1113; lean_object* x_1114; lean_dec(x_7); -lean_dec(x_5); -return x_345; +lean_dec(x_6); +x_1109 = lean_ctor_get(x_1093, 0); +lean_inc(x_1109); +x_1110 = lean_ctor_get(x_1093, 1); +lean_inc(x_1110); +lean_dec(x_1093); +lean_inc(x_1); +x_1111 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_1095, x_1092, x_1); +x_1112 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1112, 0, x_1109); +lean_ctor_set(x_1112, 1, x_1110); +lean_ctor_set(x_1112, 2, x_1111); +if (lean_is_scalar(x_1094)) { + x_1113 = lean_alloc_ctor(0, 2, 0); +} else { + x_1113 = x_1094; +} +lean_ctor_set(x_1113, 0, x_1); +lean_ctor_set(x_1113, 1, x_1112); +x_1114 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1114, 0, x_1113); +lean_ctor_set(x_1114, 1, x_1091); +return x_1114; } } else { -uint8_t x_346; -lean_dec(x_320); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); +lean_object* x_1115; lean_object* x_1116; lean_object* x_1117; +lean_dec(x_1095); +lean_dec(x_1092); +lean_dec(x_1044); lean_dec(x_7); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); +lean_dec(x_6); +lean_dec(x_1); +x_1115 = lean_ctor_get(x_1096, 0); +lean_inc(x_1115); +lean_dec(x_1096); +if (lean_is_scalar(x_1094)) { + x_1116 = lean_alloc_ctor(0, 2, 0); +} else { + x_1116 = x_1094; +} +lean_ctor_set(x_1116, 0, x_1115); +lean_ctor_set(x_1116, 1, x_1093); +x_1117 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1117, 0, x_1116); +lean_ctor_set(x_1117, 1, x_1091); +return x_1117; +} +} +} +else +{ +uint8_t x_1118; +lean_dec(x_1044); +lean_dec(x_7); +lean_dec(x_6); lean_dec(x_1); -x_346 = !lean_is_exclusive(x_321); -if (x_346 == 0) +x_1118 = !lean_is_exclusive(x_1045); +if (x_1118 == 0) { -return x_321; +return x_1045; } else { -lean_object* x_347; lean_object* x_348; lean_object* x_349; -x_347 = lean_ctor_get(x_321, 0); -x_348 = lean_ctor_get(x_321, 1); -lean_inc(x_348); -lean_inc(x_347); -lean_dec(x_321); -x_349 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_349, 0, x_347); -lean_ctor_set(x_349, 1, x_348); -return x_349; +lean_object* x_1119; lean_object* x_1120; lean_object* x_1121; +x_1119 = lean_ctor_get(x_1045, 0); +x_1120 = lean_ctor_get(x_1045, 1); +lean_inc(x_1120); +lean_inc(x_1119); +lean_dec(x_1045); +x_1121 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1121, 0, x_1119); +lean_ctor_set(x_1121, 1, x_1120); +return x_1121; +} +} +} +block_1177: +{ +lean_object* x_1124; +lean_dec(x_1123); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_6); +x_1124 = l_Lean_Meta_getFunInfo(x_6, x_11, x_12, x_13, x_14, x_15); +if (lean_obj_tag(x_1124) == 0) +{ +lean_object* x_1125; lean_object* x_1126; lean_object* x_1127; lean_object* x_1128; lean_object* x_1129; lean_object* x_1130; lean_object* x_1131; uint8_t x_1132; lean_object* x_1133; lean_object* x_1134; lean_object* x_1135; lean_object* x_1136; +x_1125 = lean_ctor_get(x_1124, 0); +lean_inc(x_1125); +x_1126 = lean_ctor_get(x_1124, 1); +lean_inc(x_1126); +lean_dec(x_1124); +x_1127 = lean_ctor_get(x_1125, 0); +lean_inc(x_1127); +lean_dec(x_1125); +x_1128 = lean_array_get_size(x_7); +x_1129 = lean_unsigned_to_nat(0u); +x_1130 = lean_unsigned_to_nat(1u); +lean_inc(x_1128); +x_1131 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1131, 0, x_1129); +lean_ctor_set(x_1131, 1, x_1128); +lean_ctor_set(x_1131, 2, x_1130); +x_1132 = 0; +x_1133 = lean_box(x_1132); +lean_inc(x_7); +x_1134 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1134, 0, x_7); +lean_ctor_set(x_1134, 1, x_1133); +x_1135 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__11___closed__1; +lean_inc(x_1131); +lean_inc(x_6); +x_1136 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10(x_2, x_3, x_4, x_5, x_6, x_7, x_1135, x_1127, x_1128, x_1131, x_1131, x_1134, x_1129, lean_box(0), lean_box(0), x_9, x_10, x_11, x_12, x_13, x_14, x_1126); +if (lean_obj_tag(x_1136) == 0) +{ +lean_object* x_1137; lean_object* x_1138; lean_object* x_1139; uint8_t x_1140; +x_1137 = lean_ctor_get(x_1136, 0); +lean_inc(x_1137); +x_1138 = lean_ctor_get(x_1137, 0); +lean_inc(x_1138); +x_1139 = lean_ctor_get(x_1138, 1); +lean_inc(x_1139); +x_1140 = lean_unbox(x_1139); +lean_dec(x_1139); +if (x_1140 == 0) +{ +uint8_t x_1141; +lean_dec(x_1138); +lean_dec(x_6); +x_1141 = !lean_is_exclusive(x_1136); +if (x_1141 == 0) +{ +lean_object* x_1142; uint8_t x_1143; +x_1142 = lean_ctor_get(x_1136, 0); +lean_dec(x_1142); +x_1143 = !lean_is_exclusive(x_1137); +if (x_1143 == 0) +{ +lean_object* x_1144; +x_1144 = lean_ctor_get(x_1137, 0); +lean_dec(x_1144); +lean_ctor_set(x_1137, 0, x_1); +return x_1136; +} +else +{ +lean_object* x_1145; lean_object* x_1146; +x_1145 = lean_ctor_get(x_1137, 1); +lean_inc(x_1145); +lean_dec(x_1137); +x_1146 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1146, 0, x_1); +lean_ctor_set(x_1146, 1, x_1145); +lean_ctor_set(x_1136, 0, x_1146); +return x_1136; +} +} +else +{ +lean_object* x_1147; lean_object* x_1148; lean_object* x_1149; lean_object* x_1150; lean_object* x_1151; +x_1147 = lean_ctor_get(x_1136, 1); +lean_inc(x_1147); +lean_dec(x_1136); +x_1148 = lean_ctor_get(x_1137, 1); +lean_inc(x_1148); +if (lean_is_exclusive(x_1137)) { + lean_ctor_release(x_1137, 0); + lean_ctor_release(x_1137, 1); + x_1149 = x_1137; +} else { + lean_dec_ref(x_1137); + x_1149 = lean_box(0); } +if (lean_is_scalar(x_1149)) { + x_1150 = lean_alloc_ctor(0, 2, 0); +} else { + x_1150 = x_1149; } +lean_ctor_set(x_1150, 0, x_1); +lean_ctor_set(x_1150, 1, x_1148); +x_1151 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1151, 0, x_1150); +lean_ctor_set(x_1151, 1, x_1147); +return x_1151; } -block_384: -{ -lean_object* x_352; -lean_dec(x_351); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_2); -x_352 = l_Lean_Meta_getFunInfo(x_2, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_352) == 0) -{ -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; uint8_t 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_355 = lean_ctor_get(x_353, 0); -lean_inc(x_355); -lean_dec(x_353); -x_356 = lean_array_get_size(x_3); -x_357 = lean_unsigned_to_nat(0u); -x_358 = lean_unsigned_to_nat(1u); -x_359 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_359, 0, x_357); -lean_ctor_set(x_359, 1, x_356); -lean_ctor_set(x_359, 2, x_358); -x_360 = 0; -x_361 = lean_box(x_360); -x_362 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_362, 0, x_3); -lean_ctor_set(x_362, 1, x_361); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_5); -lean_inc(x_2); -x_363 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8(x_2, x_355, x_359, x_359, x_362, x_357, lean_box(0), lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10, x_354); -lean_dec(x_359); -lean_dec(x_355); -if (lean_obj_tag(x_363) == 0) -{ -lean_object* x_364; lean_object* x_365; lean_object* x_366; uint8_t x_367; -x_364 = lean_ctor_get(x_363, 0); -lean_inc(x_364); -x_365 = lean_ctor_get(x_364, 0); -lean_inc(x_365); -x_366 = lean_ctor_get(x_365, 1); -lean_inc(x_366); -x_367 = lean_unbox(x_366); -lean_dec(x_366); -if (x_367 == 0) -{ -lean_object* x_368; lean_object* x_369; lean_object* x_370; -lean_dec(x_365); -lean_dec(x_2); -x_368 = lean_ctor_get(x_363, 1); -lean_inc(x_368); -lean_dec(x_363); -x_369 = lean_ctor_get(x_364, 1); -lean_inc(x_369); -lean_dec(x_364); -lean_inc(x_1); -x_370 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_369, x_7, x_8, x_9, x_10, x_368); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_370; } else { -lean_object* x_371; lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; -x_371 = lean_ctor_get(x_363, 1); -lean_inc(x_371); -lean_dec(x_363); -x_372 = lean_ctor_get(x_364, 1); -lean_inc(x_372); -lean_dec(x_364); -x_373 = lean_ctor_get(x_365, 0); -lean_inc(x_373); -lean_dec(x_365); -x_374 = l_Lean_mkAppN(x_2, x_373); -lean_dec(x_373); -x_375 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_374, x_5, x_372, x_7, x_8, x_9, x_10, x_371); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_375; +uint8_t x_1152; +lean_dec(x_1); +x_1152 = !lean_is_exclusive(x_1136); +if (x_1152 == 0) +{ +lean_object* x_1153; uint8_t x_1154; +x_1153 = lean_ctor_get(x_1136, 0); +lean_dec(x_1153); +x_1154 = !lean_is_exclusive(x_1137); +if (x_1154 == 0) +{ +lean_object* x_1155; lean_object* x_1156; lean_object* x_1157; +x_1155 = lean_ctor_get(x_1137, 0); +lean_dec(x_1155); +x_1156 = lean_ctor_get(x_1138, 0); +lean_inc(x_1156); +lean_dec(x_1138); +x_1157 = l_Lean_mkAppN(x_6, x_1156); +lean_dec(x_1156); +lean_ctor_set(x_1137, 0, x_1157); +return x_1136; +} +else +{ +lean_object* x_1158; lean_object* x_1159; lean_object* x_1160; lean_object* x_1161; +x_1158 = lean_ctor_get(x_1137, 1); +lean_inc(x_1158); +lean_dec(x_1137); +x_1159 = lean_ctor_get(x_1138, 0); +lean_inc(x_1159); +lean_dec(x_1138); +x_1160 = l_Lean_mkAppN(x_6, x_1159); +lean_dec(x_1159); +x_1161 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1161, 0, x_1160); +lean_ctor_set(x_1161, 1, x_1158); +lean_ctor_set(x_1136, 0, x_1161); +return x_1136; +} +} +else +{ +lean_object* x_1162; lean_object* x_1163; lean_object* x_1164; lean_object* x_1165; lean_object* x_1166; lean_object* x_1167; lean_object* x_1168; +x_1162 = lean_ctor_get(x_1136, 1); +lean_inc(x_1162); +lean_dec(x_1136); +x_1163 = lean_ctor_get(x_1137, 1); +lean_inc(x_1163); +if (lean_is_exclusive(x_1137)) { + lean_ctor_release(x_1137, 0); + lean_ctor_release(x_1137, 1); + x_1164 = x_1137; +} else { + lean_dec_ref(x_1137); + x_1164 = lean_box(0); +} +x_1165 = lean_ctor_get(x_1138, 0); +lean_inc(x_1165); +lean_dec(x_1138); +x_1166 = l_Lean_mkAppN(x_6, x_1165); +lean_dec(x_1165); +if (lean_is_scalar(x_1164)) { + x_1167 = lean_alloc_ctor(0, 2, 0); +} else { + x_1167 = x_1164; +} +lean_ctor_set(x_1167, 0, x_1166); +lean_ctor_set(x_1167, 1, x_1163); +x_1168 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1168, 0, x_1167); +lean_ctor_set(x_1168, 1, x_1162); +return x_1168; +} } } else { -uint8_t x_376; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -lean_dec(x_2); +uint8_t x_1169; +lean_dec(x_6); lean_dec(x_1); -x_376 = !lean_is_exclusive(x_363); -if (x_376 == 0) +x_1169 = !lean_is_exclusive(x_1136); +if (x_1169 == 0) { -return x_363; +return x_1136; } else { -lean_object* x_377; lean_object* x_378; lean_object* x_379; -x_377 = lean_ctor_get(x_363, 0); -x_378 = lean_ctor_get(x_363, 1); -lean_inc(x_378); -lean_inc(x_377); -lean_dec(x_363); -x_379 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_379, 0, x_377); -lean_ctor_set(x_379, 1, x_378); -return x_379; +lean_object* x_1170; lean_object* x_1171; lean_object* x_1172; +x_1170 = lean_ctor_get(x_1136, 0); +x_1171 = lean_ctor_get(x_1136, 1); +lean_inc(x_1171); +lean_inc(x_1170); +lean_dec(x_1136); +x_1172 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1172, 0, x_1170); +lean_ctor_set(x_1172, 1, x_1171); +return x_1172; } } } else { -uint8_t x_380; +uint8_t x_1173; +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); -lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); +lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_380 = !lean_is_exclusive(x_352); -if (x_380 == 0) +x_1173 = !lean_is_exclusive(x_1124); +if (x_1173 == 0) { -return x_352; +return x_1124; } else { -lean_object* x_381; lean_object* x_382; lean_object* x_383; -x_381 = lean_ctor_get(x_352, 0); -x_382 = lean_ctor_get(x_352, 1); -lean_inc(x_382); -lean_inc(x_381); -lean_dec(x_352); -x_383 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_383, 0, x_381); -lean_ctor_set(x_383, 1, x_382); -return x_383; -} +lean_object* x_1174; lean_object* x_1175; lean_object* x_1176; +x_1174 = lean_ctor_get(x_1124, 0); +x_1175 = lean_ctor_get(x_1124, 1); +lean_inc(x_1175); +lean_inc(x_1174); +lean_dec(x_1124); +x_1176 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1176, 0, x_1174); +lean_ctor_set(x_1176, 1, x_1175); +return x_1176; } } } -case 5: -{ -lean_object* x_397; lean_object* x_398; lean_object* x_399; lean_object* x_400; lean_object* x_401; -x_397 = lean_ctor_get(x_2, 0); -lean_inc(x_397); -x_398 = lean_ctor_get(x_2, 1); -lean_inc(x_398); -lean_dec(x_2); -x_399 = lean_array_set(x_3, x_4, x_398); -x_400 = lean_unsigned_to_nat(1u); -x_401 = lean_nat_sub(x_4, x_400); -lean_dec(x_4); -x_2 = x_397; -x_3 = x_399; -x_4 = x_401; -goto _start; } -case 6: +case 9: { -lean_object* x_403; lean_object* x_434; lean_object* x_468; uint8_t x_469; -lean_dec(x_4); -x_468 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__4; -x_469 = l_Lean_Expr_isConstOf(x_2, x_468); -if (x_469 == 0) +lean_object* x_1190; lean_object* x_1269; lean_object* x_1324; uint8_t x_1325; +lean_dec(x_8); +x_1324 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__11___closed__5; +x_1325 = l_Lean_Expr_isConstOf(x_6, x_1324); +if (x_1325 == 0) { -lean_object* x_470; -x_470 = lean_box(0); -x_434 = x_470; -goto block_467; +lean_object* x_1326; +x_1326 = lean_box(0); +x_1269 = x_1326; +goto block_1323; } else { -lean_object* x_471; lean_object* x_472; uint8_t x_473; -x_471 = lean_array_get_size(x_3); -x_472 = lean_unsigned_to_nat(2u); -x_473 = lean_nat_dec_eq(x_471, x_472); -if (x_473 == 0) +lean_object* x_1327; lean_object* x_1328; uint8_t x_1329; +x_1327 = lean_array_get_size(x_7); +x_1328 = lean_unsigned_to_nat(2u); +x_1329 = lean_nat_dec_eq(x_1327, x_1328); +if (x_1329 == 0) { -lean_object* x_474; -lean_dec(x_471); -x_474 = lean_box(0); -x_434 = x_474; -goto block_467; +lean_object* x_1330; +lean_dec(x_1327); +x_1330 = lean_box(0); +x_1269 = x_1330; +goto block_1323; } else { -lean_object* x_475; uint8_t x_476; -x_475 = lean_unsigned_to_nat(0u); -x_476 = lean_nat_dec_lt(x_475, x_471); -lean_dec(x_471); -if (x_476 == 0) +lean_object* x_1331; uint8_t x_1332; +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_1331 = lean_unsigned_to_nat(0u); +x_1332 = lean_nat_dec_lt(x_1331, x_1327); +lean_dec(x_1327); +if (x_1332 == 0) { -lean_object* x_477; lean_object* x_478; -x_477 = l_Lean_instInhabitedExpr; -x_478 = l_outOfBounds___rarg(x_477); -x_403 = x_478; -goto block_433; +lean_object* x_1333; lean_object* x_1334; +x_1333 = l_Lean_instInhabitedExpr; +x_1334 = l_outOfBounds___rarg(x_1333); +x_1190 = x_1334; +goto block_1268; } else { -lean_object* x_479; -x_479 = lean_array_fget(x_3, x_475); -x_403 = x_479; -goto block_433; +lean_object* x_1335; +x_1335 = lean_array_fget(x_7, x_1331); +x_1190 = x_1335; +goto block_1268; } } } -block_433: +block_1268: { -lean_object* x_404; -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_5); -lean_inc(x_403); -x_404 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_403, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_404) == 0) -{ -lean_object* x_405; lean_object* x_406; lean_object* x_407; lean_object* x_408; lean_object* x_409; lean_object* x_410; -x_405 = lean_ctor_get(x_404, 0); -lean_inc(x_405); -x_406 = lean_ctor_get(x_404, 1); -lean_inc(x_406); -lean_dec(x_404); -x_407 = lean_ctor_get(x_405, 0); -lean_inc(x_407); -x_408 = lean_ctor_get(x_405, 1); -lean_inc(x_408); -lean_dec(x_405); -x_409 = lean_ctor_get(x_408, 2); -lean_inc(x_409); -lean_inc(x_409); -x_410 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_409, x_407); -if (lean_obj_tag(x_410) == 0) +lean_object* x_1191; +lean_inc(x_1190); +x_1191 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_1190, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +if (lean_obj_tag(x_1191) == 0) { -size_t x_411; size_t x_412; uint8_t x_413; -x_411 = lean_ptr_addr(x_403); -lean_dec(x_403); -x_412 = lean_ptr_addr(x_407); -x_413 = lean_usize_dec_eq(x_411, x_412); -if (x_413 == 0) -{ -lean_object* x_414; lean_object* x_415; lean_object* x_416; lean_object* x_417; lean_object* x_418; lean_object* x_419; lean_object* x_420; lean_object* x_421; -x_414 = lean_unsigned_to_nat(0u); -lean_inc(x_407); -x_415 = lean_array_set(x_3, x_414, x_407); -x_416 = l_Lean_mkAppN(x_2, x_415); -lean_dec(x_415); -x_417 = lean_ctor_get(x_408, 0); -lean_inc(x_417); -x_418 = lean_ctor_get(x_408, 1); -lean_inc(x_418); -lean_dec(x_408); -lean_inc(x_416); -x_419 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_409, x_407, x_416); -x_420 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_420, 0, x_417); -lean_ctor_set(x_420, 1, x_418); -lean_ctor_set(x_420, 2, x_419); -x_421 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_416, x_5, x_420, x_7, x_8, x_9, x_10, x_406); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); +uint8_t x_1192; +x_1192 = !lean_is_exclusive(x_1191); +if (x_1192 == 0) +{ +lean_object* x_1193; uint8_t x_1194; +x_1193 = lean_ctor_get(x_1191, 0); +x_1194 = !lean_is_exclusive(x_1193); +if (x_1194 == 0) +{ +lean_object* x_1195; lean_object* x_1196; lean_object* x_1197; lean_object* x_1198; +x_1195 = lean_ctor_get(x_1193, 0); +x_1196 = lean_ctor_get(x_1193, 1); +x_1197 = lean_ctor_get(x_1196, 2); +lean_inc(x_1197); +lean_inc(x_1197); +x_1198 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_1197, x_1195); +if (lean_obj_tag(x_1198) == 0) +{ +size_t x_1199; size_t x_1200; uint8_t x_1201; +x_1199 = lean_ptr_addr(x_1190); +lean_dec(x_1190); +x_1200 = lean_ptr_addr(x_1195); +x_1201 = lean_usize_dec_eq(x_1199, x_1200); +if (x_1201 == 0) +{ +lean_object* x_1202; lean_object* x_1203; lean_object* x_1204; lean_object* x_1205; lean_object* x_1206; lean_object* x_1207; lean_object* x_1208; +lean_dec(x_1); +x_1202 = lean_unsigned_to_nat(0u); +lean_inc(x_1195); +x_1203 = lean_array_set(x_7, x_1202, x_1195); +x_1204 = l_Lean_mkAppN(x_6, x_1203); +lean_dec(x_1203); +x_1205 = lean_ctor_get(x_1196, 0); +lean_inc(x_1205); +x_1206 = lean_ctor_get(x_1196, 1); +lean_inc(x_1206); +lean_dec(x_1196); +lean_inc(x_1204); +x_1207 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_1197, x_1195, x_1204); +x_1208 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1208, 0, x_1205); +lean_ctor_set(x_1208, 1, x_1206); +lean_ctor_set(x_1208, 2, x_1207); +lean_ctor_set(x_1193, 1, x_1208); +lean_ctor_set(x_1193, 0, x_1204); +return x_1191; +} +else +{ +lean_object* x_1209; lean_object* x_1210; lean_object* x_1211; lean_object* x_1212; lean_dec(x_7); -lean_dec(x_5); -return x_421; +lean_dec(x_6); +x_1209 = lean_ctor_get(x_1196, 0); +lean_inc(x_1209); +x_1210 = lean_ctor_get(x_1196, 1); +lean_inc(x_1210); +lean_dec(x_1196); +lean_inc(x_1); +x_1211 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_1197, x_1195, x_1); +x_1212 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1212, 0, x_1209); +lean_ctor_set(x_1212, 1, x_1210); +lean_ctor_set(x_1212, 2, x_1211); +lean_ctor_set(x_1193, 1, x_1212); +lean_ctor_set(x_1193, 0, x_1); +return x_1191; +} } else { -lean_object* x_422; lean_object* x_423; lean_object* x_424; lean_object* x_425; lean_object* x_426; -lean_dec(x_3); -lean_dec(x_2); -x_422 = lean_ctor_get(x_408, 0); -lean_inc(x_422); -x_423 = lean_ctor_get(x_408, 1); -lean_inc(x_423); -lean_dec(x_408); -lean_inc(x_1); -x_424 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_409, x_407, x_1); -x_425 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_425, 0, x_422); -lean_ctor_set(x_425, 1, x_423); -lean_ctor_set(x_425, 2, x_424); +lean_object* x_1213; +lean_dec(x_1197); +lean_dec(x_1195); +lean_dec(x_1190); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_1); +x_1213 = lean_ctor_get(x_1198, 0); +lean_inc(x_1213); +lean_dec(x_1198); +lean_ctor_set(x_1193, 0, x_1213); +return x_1191; +} +} +else +{ +lean_object* x_1214; lean_object* x_1215; lean_object* x_1216; lean_object* x_1217; +x_1214 = lean_ctor_get(x_1193, 0); +x_1215 = lean_ctor_get(x_1193, 1); +lean_inc(x_1215); +lean_inc(x_1214); +lean_dec(x_1193); +x_1216 = lean_ctor_get(x_1215, 2); +lean_inc(x_1216); +lean_inc(x_1216); +x_1217 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_1216, x_1214); +if (lean_obj_tag(x_1217) == 0) +{ +size_t x_1218; size_t x_1219; uint8_t x_1220; +x_1218 = lean_ptr_addr(x_1190); +lean_dec(x_1190); +x_1219 = lean_ptr_addr(x_1214); +x_1220 = lean_usize_dec_eq(x_1218, x_1219); +if (x_1220 == 0) +{ +lean_object* x_1221; lean_object* x_1222; lean_object* x_1223; lean_object* x_1224; lean_object* x_1225; lean_object* x_1226; lean_object* x_1227; lean_object* x_1228; +lean_dec(x_1); +x_1221 = lean_unsigned_to_nat(0u); +lean_inc(x_1214); +x_1222 = lean_array_set(x_7, x_1221, x_1214); +x_1223 = l_Lean_mkAppN(x_6, x_1222); +lean_dec(x_1222); +x_1224 = lean_ctor_get(x_1215, 0); +lean_inc(x_1224); +x_1225 = lean_ctor_get(x_1215, 1); +lean_inc(x_1225); +lean_dec(x_1215); +lean_inc(x_1223); +x_1226 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_1216, x_1214, x_1223); +x_1227 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1227, 0, x_1224); +lean_ctor_set(x_1227, 1, x_1225); +lean_ctor_set(x_1227, 2, x_1226); +x_1228 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1228, 0, x_1223); +lean_ctor_set(x_1228, 1, x_1227); +lean_ctor_set(x_1191, 0, x_1228); +return x_1191; +} +else +{ +lean_object* x_1229; lean_object* x_1230; lean_object* x_1231; lean_object* x_1232; lean_object* x_1233; +lean_dec(x_7); +lean_dec(x_6); +x_1229 = lean_ctor_get(x_1215, 0); +lean_inc(x_1229); +x_1230 = lean_ctor_get(x_1215, 1); +lean_inc(x_1230); +lean_dec(x_1215); lean_inc(x_1); -x_426 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_425, x_7, x_8, x_9, x_10, x_406); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); +x_1231 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_1216, x_1214, x_1); +x_1232 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1232, 0, x_1229); +lean_ctor_set(x_1232, 1, x_1230); +lean_ctor_set(x_1232, 2, x_1231); +x_1233 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1233, 0, x_1); +lean_ctor_set(x_1233, 1, x_1232); +lean_ctor_set(x_1191, 0, x_1233); +return x_1191; +} +} +else +{ +lean_object* x_1234; lean_object* x_1235; +lean_dec(x_1216); +lean_dec(x_1214); +lean_dec(x_1190); lean_dec(x_7); -lean_dec(x_5); -return x_426; +lean_dec(x_6); +lean_dec(x_1); +x_1234 = lean_ctor_get(x_1217, 0); +lean_inc(x_1234); +lean_dec(x_1217); +x_1235 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1235, 0, x_1234); +lean_ctor_set(x_1235, 1, x_1215); +lean_ctor_set(x_1191, 0, x_1235); +return x_1191; +} +} +} +else +{ +lean_object* x_1236; lean_object* x_1237; lean_object* x_1238; lean_object* x_1239; lean_object* x_1240; lean_object* x_1241; lean_object* x_1242; +x_1236 = lean_ctor_get(x_1191, 0); +x_1237 = lean_ctor_get(x_1191, 1); +lean_inc(x_1237); +lean_inc(x_1236); +lean_dec(x_1191); +x_1238 = lean_ctor_get(x_1236, 0); +lean_inc(x_1238); +x_1239 = lean_ctor_get(x_1236, 1); +lean_inc(x_1239); +if (lean_is_exclusive(x_1236)) { + lean_ctor_release(x_1236, 0); + lean_ctor_release(x_1236, 1); + x_1240 = x_1236; +} else { + lean_dec_ref(x_1236); + x_1240 = lean_box(0); +} +x_1241 = lean_ctor_get(x_1239, 2); +lean_inc(x_1241); +lean_inc(x_1241); +x_1242 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_1241, x_1238); +if (lean_obj_tag(x_1242) == 0) +{ +size_t x_1243; size_t x_1244; uint8_t x_1245; +x_1243 = lean_ptr_addr(x_1190); +lean_dec(x_1190); +x_1244 = lean_ptr_addr(x_1238); +x_1245 = lean_usize_dec_eq(x_1243, x_1244); +if (x_1245 == 0) +{ +lean_object* x_1246; lean_object* x_1247; lean_object* x_1248; lean_object* x_1249; lean_object* x_1250; lean_object* x_1251; lean_object* x_1252; lean_object* x_1253; lean_object* x_1254; +lean_dec(x_1); +x_1246 = lean_unsigned_to_nat(0u); +lean_inc(x_1238); +x_1247 = lean_array_set(x_7, x_1246, x_1238); +x_1248 = l_Lean_mkAppN(x_6, x_1247); +lean_dec(x_1247); +x_1249 = lean_ctor_get(x_1239, 0); +lean_inc(x_1249); +x_1250 = lean_ctor_get(x_1239, 1); +lean_inc(x_1250); +lean_dec(x_1239); +lean_inc(x_1248); +x_1251 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_1241, x_1238, x_1248); +x_1252 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1252, 0, x_1249); +lean_ctor_set(x_1252, 1, x_1250); +lean_ctor_set(x_1252, 2, x_1251); +if (lean_is_scalar(x_1240)) { + x_1253 = lean_alloc_ctor(0, 2, 0); +} else { + x_1253 = x_1240; } +lean_ctor_set(x_1253, 0, x_1248); +lean_ctor_set(x_1253, 1, x_1252); +x_1254 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1254, 0, x_1253); +lean_ctor_set(x_1254, 1, x_1237); +return x_1254; } else { -lean_object* x_427; lean_object* x_428; -lean_dec(x_409); -lean_dec(x_407); -lean_dec(x_403); -lean_dec(x_3); -lean_dec(x_2); -x_427 = lean_ctor_get(x_410, 0); -lean_inc(x_427); -lean_dec(x_410); -x_428 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_427, x_5, x_408, x_7, x_8, x_9, x_10, x_406); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); +lean_object* x_1255; lean_object* x_1256; lean_object* x_1257; lean_object* x_1258; lean_object* x_1259; lean_object* x_1260; lean_dec(x_7); -lean_dec(x_5); -return x_428; +lean_dec(x_6); +x_1255 = lean_ctor_get(x_1239, 0); +lean_inc(x_1255); +x_1256 = lean_ctor_get(x_1239, 1); +lean_inc(x_1256); +lean_dec(x_1239); +lean_inc(x_1); +x_1257 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_1241, x_1238, x_1); +x_1258 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1258, 0, x_1255); +lean_ctor_set(x_1258, 1, x_1256); +lean_ctor_set(x_1258, 2, x_1257); +if (lean_is_scalar(x_1240)) { + x_1259 = lean_alloc_ctor(0, 2, 0); +} else { + x_1259 = x_1240; +} +lean_ctor_set(x_1259, 0, x_1); +lean_ctor_set(x_1259, 1, x_1258); +x_1260 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1260, 0, x_1259); +lean_ctor_set(x_1260, 1, x_1237); +return x_1260; } } else { -uint8_t x_429; -lean_dec(x_403); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); +lean_object* x_1261; lean_object* x_1262; lean_object* x_1263; +lean_dec(x_1241); +lean_dec(x_1238); +lean_dec(x_1190); lean_dec(x_7); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); +lean_dec(x_6); +lean_dec(x_1); +x_1261 = lean_ctor_get(x_1242, 0); +lean_inc(x_1261); +lean_dec(x_1242); +if (lean_is_scalar(x_1240)) { + x_1262 = lean_alloc_ctor(0, 2, 0); +} else { + x_1262 = x_1240; +} +lean_ctor_set(x_1262, 0, x_1261); +lean_ctor_set(x_1262, 1, x_1239); +x_1263 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1263, 0, x_1262); +lean_ctor_set(x_1263, 1, x_1237); +return x_1263; +} +} +} +else +{ +uint8_t x_1264; +lean_dec(x_1190); +lean_dec(x_7); +lean_dec(x_6); lean_dec(x_1); -x_429 = !lean_is_exclusive(x_404); -if (x_429 == 0) +x_1264 = !lean_is_exclusive(x_1191); +if (x_1264 == 0) { -return x_404; +return x_1191; } else { -lean_object* x_430; lean_object* x_431; lean_object* x_432; -x_430 = lean_ctor_get(x_404, 0); -x_431 = lean_ctor_get(x_404, 1); -lean_inc(x_431); -lean_inc(x_430); -lean_dec(x_404); -x_432 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_432, 0, x_430); -lean_ctor_set(x_432, 1, x_431); -return x_432; +lean_object* x_1265; lean_object* x_1266; lean_object* x_1267; +x_1265 = lean_ctor_get(x_1191, 0); +x_1266 = lean_ctor_get(x_1191, 1); +lean_inc(x_1266); +lean_inc(x_1265); +lean_dec(x_1191); +x_1267 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1267, 0, x_1265); +lean_ctor_set(x_1267, 1, x_1266); +return x_1267; } } } -block_467: +block_1323: { -lean_object* x_435; -lean_dec(x_434); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_2); -x_435 = l_Lean_Meta_getFunInfo(x_2, x_7, x_8, x_9, x_10, x_11); -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; uint8_t x_443; lean_object* x_444; lean_object* x_445; lean_object* x_446; -x_436 = lean_ctor_get(x_435, 0); -lean_inc(x_436); -x_437 = lean_ctor_get(x_435, 1); -lean_inc(x_437); -lean_dec(x_435); -x_438 = lean_ctor_get(x_436, 0); -lean_inc(x_438); -lean_dec(x_436); -x_439 = lean_array_get_size(x_3); -x_440 = lean_unsigned_to_nat(0u); -x_441 = lean_unsigned_to_nat(1u); -x_442 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_442, 0, x_440); -lean_ctor_set(x_442, 1, x_439); -lean_ctor_set(x_442, 2, x_441); -x_443 = 0; -x_444 = lean_box(x_443); -x_445 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_445, 0, x_3); -lean_ctor_set(x_445, 1, x_444); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); +lean_object* x_1270; +lean_dec(x_1269); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_6); +x_1270 = l_Lean_Meta_getFunInfo(x_6, x_11, x_12, x_13, x_14, x_15); +if (lean_obj_tag(x_1270) == 0) +{ +lean_object* x_1271; lean_object* x_1272; lean_object* x_1273; lean_object* x_1274; lean_object* x_1275; lean_object* x_1276; lean_object* x_1277; uint8_t x_1278; lean_object* x_1279; lean_object* x_1280; lean_object* x_1281; lean_object* x_1282; +x_1271 = lean_ctor_get(x_1270, 0); +lean_inc(x_1271); +x_1272 = lean_ctor_get(x_1270, 1); +lean_inc(x_1272); +lean_dec(x_1270); +x_1273 = lean_ctor_get(x_1271, 0); +lean_inc(x_1273); +lean_dec(x_1271); +x_1274 = lean_array_get_size(x_7); +x_1275 = lean_unsigned_to_nat(0u); +x_1276 = lean_unsigned_to_nat(1u); +lean_inc(x_1274); +x_1277 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1277, 0, x_1275); +lean_ctor_set(x_1277, 1, x_1274); +lean_ctor_set(x_1277, 2, x_1276); +x_1278 = 0; +x_1279 = lean_box(x_1278); lean_inc(x_7); -lean_inc(x_5); -lean_inc(x_2); -x_446 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8(x_2, x_438, x_442, x_442, x_445, x_440, lean_box(0), lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10, x_437); -lean_dec(x_442); -lean_dec(x_438); -if (lean_obj_tag(x_446) == 0) -{ -lean_object* x_447; lean_object* x_448; lean_object* x_449; uint8_t x_450; -x_447 = lean_ctor_get(x_446, 0); -lean_inc(x_447); -x_448 = lean_ctor_get(x_447, 0); -lean_inc(x_448); -x_449 = lean_ctor_get(x_448, 1); -lean_inc(x_449); -x_450 = lean_unbox(x_449); -lean_dec(x_449); -if (x_450 == 0) -{ -lean_object* x_451; lean_object* x_452; lean_object* x_453; -lean_dec(x_448); -lean_dec(x_2); -x_451 = lean_ctor_get(x_446, 1); -lean_inc(x_451); -lean_dec(x_446); -x_452 = lean_ctor_get(x_447, 1); -lean_inc(x_452); -lean_dec(x_447); -lean_inc(x_1); -x_453 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_452, x_7, x_8, x_9, x_10, x_451); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_453; +x_1280 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1280, 0, x_7); +lean_ctor_set(x_1280, 1, x_1279); +x_1281 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__11___closed__1; +lean_inc(x_1277); +lean_inc(x_6); +x_1282 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10(x_2, x_3, x_4, x_5, x_6, x_7, x_1281, x_1273, x_1274, x_1277, x_1277, x_1280, x_1275, lean_box(0), lean_box(0), x_9, x_10, x_11, x_12, x_13, x_14, x_1272); +if (lean_obj_tag(x_1282) == 0) +{ +lean_object* x_1283; lean_object* x_1284; lean_object* x_1285; uint8_t x_1286; +x_1283 = lean_ctor_get(x_1282, 0); +lean_inc(x_1283); +x_1284 = lean_ctor_get(x_1283, 0); +lean_inc(x_1284); +x_1285 = lean_ctor_get(x_1284, 1); +lean_inc(x_1285); +x_1286 = lean_unbox(x_1285); +lean_dec(x_1285); +if (x_1286 == 0) +{ +uint8_t x_1287; +lean_dec(x_1284); +lean_dec(x_6); +x_1287 = !lean_is_exclusive(x_1282); +if (x_1287 == 0) +{ +lean_object* x_1288; uint8_t x_1289; +x_1288 = lean_ctor_get(x_1282, 0); +lean_dec(x_1288); +x_1289 = !lean_is_exclusive(x_1283); +if (x_1289 == 0) +{ +lean_object* x_1290; +x_1290 = lean_ctor_get(x_1283, 0); +lean_dec(x_1290); +lean_ctor_set(x_1283, 0, x_1); +return x_1282; +} +else +{ +lean_object* x_1291; lean_object* x_1292; +x_1291 = lean_ctor_get(x_1283, 1); +lean_inc(x_1291); +lean_dec(x_1283); +x_1292 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1292, 0, x_1); +lean_ctor_set(x_1292, 1, x_1291); +lean_ctor_set(x_1282, 0, x_1292); +return x_1282; +} +} +else +{ +lean_object* x_1293; lean_object* x_1294; lean_object* x_1295; lean_object* x_1296; lean_object* x_1297; +x_1293 = lean_ctor_get(x_1282, 1); +lean_inc(x_1293); +lean_dec(x_1282); +x_1294 = lean_ctor_get(x_1283, 1); +lean_inc(x_1294); +if (lean_is_exclusive(x_1283)) { + lean_ctor_release(x_1283, 0); + lean_ctor_release(x_1283, 1); + x_1295 = x_1283; +} else { + lean_dec_ref(x_1283); + x_1295 = lean_box(0); +} +if (lean_is_scalar(x_1295)) { + x_1296 = lean_alloc_ctor(0, 2, 0); +} else { + x_1296 = x_1295; +} +lean_ctor_set(x_1296, 0, x_1); +lean_ctor_set(x_1296, 1, x_1294); +x_1297 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1297, 0, x_1296); +lean_ctor_set(x_1297, 1, x_1293); +return x_1297; +} } else { -lean_object* x_454; lean_object* x_455; lean_object* x_456; lean_object* x_457; lean_object* x_458; -x_454 = lean_ctor_get(x_446, 1); -lean_inc(x_454); -lean_dec(x_446); -x_455 = lean_ctor_get(x_447, 1); -lean_inc(x_455); -lean_dec(x_447); -x_456 = lean_ctor_get(x_448, 0); -lean_inc(x_456); -lean_dec(x_448); -x_457 = l_Lean_mkAppN(x_2, x_456); -lean_dec(x_456); -x_458 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_457, x_5, x_455, x_7, x_8, x_9, x_10, x_454); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_458; +uint8_t x_1298; +lean_dec(x_1); +x_1298 = !lean_is_exclusive(x_1282); +if (x_1298 == 0) +{ +lean_object* x_1299; uint8_t x_1300; +x_1299 = lean_ctor_get(x_1282, 0); +lean_dec(x_1299); +x_1300 = !lean_is_exclusive(x_1283); +if (x_1300 == 0) +{ +lean_object* x_1301; lean_object* x_1302; lean_object* x_1303; +x_1301 = lean_ctor_get(x_1283, 0); +lean_dec(x_1301); +x_1302 = lean_ctor_get(x_1284, 0); +lean_inc(x_1302); +lean_dec(x_1284); +x_1303 = l_Lean_mkAppN(x_6, x_1302); +lean_dec(x_1302); +lean_ctor_set(x_1283, 0, x_1303); +return x_1282; +} +else +{ +lean_object* x_1304; lean_object* x_1305; lean_object* x_1306; lean_object* x_1307; +x_1304 = lean_ctor_get(x_1283, 1); +lean_inc(x_1304); +lean_dec(x_1283); +x_1305 = lean_ctor_get(x_1284, 0); +lean_inc(x_1305); +lean_dec(x_1284); +x_1306 = l_Lean_mkAppN(x_6, x_1305); +lean_dec(x_1305); +x_1307 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1307, 0, x_1306); +lean_ctor_set(x_1307, 1, x_1304); +lean_ctor_set(x_1282, 0, x_1307); +return x_1282; +} +} +else +{ +lean_object* x_1308; lean_object* x_1309; lean_object* x_1310; lean_object* x_1311; lean_object* x_1312; lean_object* x_1313; lean_object* x_1314; +x_1308 = lean_ctor_get(x_1282, 1); +lean_inc(x_1308); +lean_dec(x_1282); +x_1309 = lean_ctor_get(x_1283, 1); +lean_inc(x_1309); +if (lean_is_exclusive(x_1283)) { + lean_ctor_release(x_1283, 0); + lean_ctor_release(x_1283, 1); + x_1310 = x_1283; +} else { + lean_dec_ref(x_1283); + x_1310 = lean_box(0); +} +x_1311 = lean_ctor_get(x_1284, 0); +lean_inc(x_1311); +lean_dec(x_1284); +x_1312 = l_Lean_mkAppN(x_6, x_1311); +lean_dec(x_1311); +if (lean_is_scalar(x_1310)) { + x_1313 = lean_alloc_ctor(0, 2, 0); +} else { + x_1313 = x_1310; +} +lean_ctor_set(x_1313, 0, x_1312); +lean_ctor_set(x_1313, 1, x_1309); +x_1314 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1314, 0, x_1313); +lean_ctor_set(x_1314, 1, x_1308); +return x_1314; +} } } else { -uint8_t x_459; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -lean_dec(x_2); +uint8_t x_1315; +lean_dec(x_6); lean_dec(x_1); -x_459 = !lean_is_exclusive(x_446); -if (x_459 == 0) +x_1315 = !lean_is_exclusive(x_1282); +if (x_1315 == 0) { -return x_446; +return x_1282; } else { -lean_object* x_460; lean_object* x_461; lean_object* x_462; -x_460 = lean_ctor_get(x_446, 0); -x_461 = lean_ctor_get(x_446, 1); -lean_inc(x_461); -lean_inc(x_460); -lean_dec(x_446); -x_462 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_462, 0, x_460); -lean_ctor_set(x_462, 1, x_461); -return x_462; +lean_object* x_1316; lean_object* x_1317; lean_object* x_1318; +x_1316 = lean_ctor_get(x_1282, 0); +x_1317 = lean_ctor_get(x_1282, 1); +lean_inc(x_1317); +lean_inc(x_1316); +lean_dec(x_1282); +x_1318 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1318, 0, x_1316); +lean_ctor_set(x_1318, 1, x_1317); +return x_1318; } } } else { -uint8_t x_463; +uint8_t x_1319; +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); -lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); +lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_463 = !lean_is_exclusive(x_435); -if (x_463 == 0) +x_1319 = !lean_is_exclusive(x_1270); +if (x_1319 == 0) { -return x_435; +return x_1270; } else { -lean_object* x_464; lean_object* x_465; lean_object* x_466; -x_464 = lean_ctor_get(x_435, 0); -x_465 = lean_ctor_get(x_435, 1); -lean_inc(x_465); -lean_inc(x_464); -lean_dec(x_435); -x_466 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_466, 0, x_464); -lean_ctor_set(x_466, 1, x_465); -return x_466; +lean_object* x_1320; lean_object* x_1321; lean_object* x_1322; +x_1320 = lean_ctor_get(x_1270, 0); +x_1321 = lean_ctor_get(x_1270, 1); +lean_inc(x_1321); +lean_inc(x_1320); +lean_dec(x_1270); +x_1322 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1322, 0, x_1320); +lean_ctor_set(x_1322, 1, x_1321); +return x_1322; } } } } -case 7: +case 10: { -lean_object* x_480; lean_object* x_511; lean_object* x_545; uint8_t x_546; -lean_dec(x_4); -x_545 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__4; -x_546 = l_Lean_Expr_isConstOf(x_2, x_545); -if (x_546 == 0) +lean_object* x_1336; lean_object* x_1415; lean_object* x_1470; uint8_t x_1471; +lean_dec(x_8); +x_1470 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__11___closed__5; +x_1471 = l_Lean_Expr_isConstOf(x_6, x_1470); +if (x_1471 == 0) { -lean_object* x_547; -x_547 = lean_box(0); -x_511 = x_547; -goto block_544; +lean_object* x_1472; +x_1472 = lean_box(0); +x_1415 = x_1472; +goto block_1469; } else { -lean_object* x_548; lean_object* x_549; uint8_t x_550; -x_548 = lean_array_get_size(x_3); -x_549 = lean_unsigned_to_nat(2u); -x_550 = lean_nat_dec_eq(x_548, x_549); -if (x_550 == 0) +lean_object* x_1473; lean_object* x_1474; uint8_t x_1475; +x_1473 = lean_array_get_size(x_7); +x_1474 = lean_unsigned_to_nat(2u); +x_1475 = lean_nat_dec_eq(x_1473, x_1474); +if (x_1475 == 0) { -lean_object* x_551; -lean_dec(x_548); -x_551 = lean_box(0); -x_511 = x_551; -goto block_544; +lean_object* x_1476; +lean_dec(x_1473); +x_1476 = lean_box(0); +x_1415 = x_1476; +goto block_1469; } else { -lean_object* x_552; uint8_t x_553; -x_552 = lean_unsigned_to_nat(0u); -x_553 = lean_nat_dec_lt(x_552, x_548); -lean_dec(x_548); -if (x_553 == 0) +lean_object* x_1477; uint8_t x_1478; +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_1477 = lean_unsigned_to_nat(0u); +x_1478 = lean_nat_dec_lt(x_1477, x_1473); +lean_dec(x_1473); +if (x_1478 == 0) { -lean_object* x_554; lean_object* x_555; -x_554 = l_Lean_instInhabitedExpr; -x_555 = l_outOfBounds___rarg(x_554); -x_480 = x_555; -goto block_510; +lean_object* x_1479; lean_object* x_1480; +x_1479 = l_Lean_instInhabitedExpr; +x_1480 = l_outOfBounds___rarg(x_1479); +x_1336 = x_1480; +goto block_1414; } else { -lean_object* x_556; -x_556 = lean_array_fget(x_3, x_552); -x_480 = x_556; -goto block_510; +lean_object* x_1481; +x_1481 = lean_array_fget(x_7, x_1477); +x_1336 = x_1481; +goto block_1414; } } } -block_510: +block_1414: { -lean_object* x_481; -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_5); -lean_inc(x_480); -x_481 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_480, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_481) == 0) +lean_object* x_1337; +lean_inc(x_1336); +x_1337 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_1336, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +if (lean_obj_tag(x_1337) == 0) { -lean_object* x_482; lean_object* x_483; lean_object* x_484; lean_object* x_485; lean_object* x_486; lean_object* x_487; -x_482 = lean_ctor_get(x_481, 0); -lean_inc(x_482); -x_483 = lean_ctor_get(x_481, 1); -lean_inc(x_483); -lean_dec(x_481); -x_484 = lean_ctor_get(x_482, 0); -lean_inc(x_484); -x_485 = lean_ctor_get(x_482, 1); -lean_inc(x_485); -lean_dec(x_482); -x_486 = lean_ctor_get(x_485, 2); -lean_inc(x_486); -lean_inc(x_486); -x_487 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_486, x_484); -if (lean_obj_tag(x_487) == 0) -{ -size_t x_488; size_t x_489; uint8_t x_490; -x_488 = lean_ptr_addr(x_480); -lean_dec(x_480); -x_489 = lean_ptr_addr(x_484); -x_490 = lean_usize_dec_eq(x_488, x_489); -if (x_490 == 0) -{ -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; -x_491 = lean_unsigned_to_nat(0u); -lean_inc(x_484); -x_492 = lean_array_set(x_3, x_491, x_484); -x_493 = l_Lean_mkAppN(x_2, x_492); -lean_dec(x_492); -x_494 = lean_ctor_get(x_485, 0); -lean_inc(x_494); -x_495 = lean_ctor_get(x_485, 1); -lean_inc(x_495); -lean_dec(x_485); -lean_inc(x_493); -x_496 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_486, x_484, x_493); -x_497 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_497, 0, x_494); -lean_ctor_set(x_497, 1, x_495); -lean_ctor_set(x_497, 2, x_496); -x_498 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_493, x_5, x_497, x_7, x_8, x_9, x_10, x_483); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); +uint8_t x_1338; +x_1338 = !lean_is_exclusive(x_1337); +if (x_1338 == 0) +{ +lean_object* x_1339; uint8_t x_1340; +x_1339 = lean_ctor_get(x_1337, 0); +x_1340 = !lean_is_exclusive(x_1339); +if (x_1340 == 0) +{ +lean_object* x_1341; lean_object* x_1342; lean_object* x_1343; lean_object* x_1344; +x_1341 = lean_ctor_get(x_1339, 0); +x_1342 = lean_ctor_get(x_1339, 1); +x_1343 = lean_ctor_get(x_1342, 2); +lean_inc(x_1343); +lean_inc(x_1343); +x_1344 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_1343, x_1341); +if (lean_obj_tag(x_1344) == 0) +{ +size_t x_1345; size_t x_1346; uint8_t x_1347; +x_1345 = lean_ptr_addr(x_1336); +lean_dec(x_1336); +x_1346 = lean_ptr_addr(x_1341); +x_1347 = lean_usize_dec_eq(x_1345, x_1346); +if (x_1347 == 0) +{ +lean_object* x_1348; lean_object* x_1349; lean_object* x_1350; lean_object* x_1351; lean_object* x_1352; lean_object* x_1353; lean_object* x_1354; +lean_dec(x_1); +x_1348 = lean_unsigned_to_nat(0u); +lean_inc(x_1341); +x_1349 = lean_array_set(x_7, x_1348, x_1341); +x_1350 = l_Lean_mkAppN(x_6, x_1349); +lean_dec(x_1349); +x_1351 = lean_ctor_get(x_1342, 0); +lean_inc(x_1351); +x_1352 = lean_ctor_get(x_1342, 1); +lean_inc(x_1352); +lean_dec(x_1342); +lean_inc(x_1350); +x_1353 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_1343, x_1341, x_1350); +x_1354 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1354, 0, x_1351); +lean_ctor_set(x_1354, 1, x_1352); +lean_ctor_set(x_1354, 2, x_1353); +lean_ctor_set(x_1339, 1, x_1354); +lean_ctor_set(x_1339, 0, x_1350); +return x_1337; +} +else +{ +lean_object* x_1355; lean_object* x_1356; lean_object* x_1357; lean_object* x_1358; lean_dec(x_7); -lean_dec(x_5); -return x_498; +lean_dec(x_6); +x_1355 = lean_ctor_get(x_1342, 0); +lean_inc(x_1355); +x_1356 = lean_ctor_get(x_1342, 1); +lean_inc(x_1356); +lean_dec(x_1342); +lean_inc(x_1); +x_1357 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_1343, x_1341, x_1); +x_1358 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1358, 0, x_1355); +lean_ctor_set(x_1358, 1, x_1356); +lean_ctor_set(x_1358, 2, x_1357); +lean_ctor_set(x_1339, 1, x_1358); +lean_ctor_set(x_1339, 0, x_1); +return x_1337; +} +} +else +{ +lean_object* x_1359; +lean_dec(x_1343); +lean_dec(x_1341); +lean_dec(x_1336); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_1); +x_1359 = lean_ctor_get(x_1344, 0); +lean_inc(x_1359); +lean_dec(x_1344); +lean_ctor_set(x_1339, 0, x_1359); +return x_1337; +} +} +else +{ +lean_object* x_1360; lean_object* x_1361; lean_object* x_1362; lean_object* x_1363; +x_1360 = lean_ctor_get(x_1339, 0); +x_1361 = lean_ctor_get(x_1339, 1); +lean_inc(x_1361); +lean_inc(x_1360); +lean_dec(x_1339); +x_1362 = lean_ctor_get(x_1361, 2); +lean_inc(x_1362); +lean_inc(x_1362); +x_1363 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_1362, x_1360); +if (lean_obj_tag(x_1363) == 0) +{ +size_t x_1364; size_t x_1365; uint8_t x_1366; +x_1364 = lean_ptr_addr(x_1336); +lean_dec(x_1336); +x_1365 = lean_ptr_addr(x_1360); +x_1366 = lean_usize_dec_eq(x_1364, x_1365); +if (x_1366 == 0) +{ +lean_object* x_1367; lean_object* x_1368; lean_object* x_1369; lean_object* x_1370; lean_object* x_1371; lean_object* x_1372; lean_object* x_1373; lean_object* x_1374; +lean_dec(x_1); +x_1367 = lean_unsigned_to_nat(0u); +lean_inc(x_1360); +x_1368 = lean_array_set(x_7, x_1367, x_1360); +x_1369 = l_Lean_mkAppN(x_6, x_1368); +lean_dec(x_1368); +x_1370 = lean_ctor_get(x_1361, 0); +lean_inc(x_1370); +x_1371 = lean_ctor_get(x_1361, 1); +lean_inc(x_1371); +lean_dec(x_1361); +lean_inc(x_1369); +x_1372 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_1362, x_1360, x_1369); +x_1373 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1373, 0, x_1370); +lean_ctor_set(x_1373, 1, x_1371); +lean_ctor_set(x_1373, 2, x_1372); +x_1374 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1374, 0, x_1369); +lean_ctor_set(x_1374, 1, x_1373); +lean_ctor_set(x_1337, 0, x_1374); +return x_1337; +} +else +{ +lean_object* x_1375; lean_object* x_1376; lean_object* x_1377; lean_object* x_1378; lean_object* x_1379; +lean_dec(x_7); +lean_dec(x_6); +x_1375 = lean_ctor_get(x_1361, 0); +lean_inc(x_1375); +x_1376 = lean_ctor_get(x_1361, 1); +lean_inc(x_1376); +lean_dec(x_1361); +lean_inc(x_1); +x_1377 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_1362, x_1360, x_1); +x_1378 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1378, 0, x_1375); +lean_ctor_set(x_1378, 1, x_1376); +lean_ctor_set(x_1378, 2, x_1377); +x_1379 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1379, 0, x_1); +lean_ctor_set(x_1379, 1, x_1378); +lean_ctor_set(x_1337, 0, x_1379); +return x_1337; +} +} +else +{ +lean_object* x_1380; lean_object* x_1381; +lean_dec(x_1362); +lean_dec(x_1360); +lean_dec(x_1336); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_1); +x_1380 = lean_ctor_get(x_1363, 0); +lean_inc(x_1380); +lean_dec(x_1363); +x_1381 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1381, 0, x_1380); +lean_ctor_set(x_1381, 1, x_1361); +lean_ctor_set(x_1337, 0, x_1381); +return x_1337; +} +} +} +else +{ +lean_object* x_1382; lean_object* x_1383; lean_object* x_1384; lean_object* x_1385; lean_object* x_1386; lean_object* x_1387; lean_object* x_1388; +x_1382 = lean_ctor_get(x_1337, 0); +x_1383 = lean_ctor_get(x_1337, 1); +lean_inc(x_1383); +lean_inc(x_1382); +lean_dec(x_1337); +x_1384 = lean_ctor_get(x_1382, 0); +lean_inc(x_1384); +x_1385 = lean_ctor_get(x_1382, 1); +lean_inc(x_1385); +if (lean_is_exclusive(x_1382)) { + lean_ctor_release(x_1382, 0); + lean_ctor_release(x_1382, 1); + x_1386 = x_1382; +} else { + lean_dec_ref(x_1382); + x_1386 = lean_box(0); +} +x_1387 = lean_ctor_get(x_1385, 2); +lean_inc(x_1387); +lean_inc(x_1387); +x_1388 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_1387, x_1384); +if (lean_obj_tag(x_1388) == 0) +{ +size_t x_1389; size_t x_1390; uint8_t x_1391; +x_1389 = lean_ptr_addr(x_1336); +lean_dec(x_1336); +x_1390 = lean_ptr_addr(x_1384); +x_1391 = lean_usize_dec_eq(x_1389, x_1390); +if (x_1391 == 0) +{ +lean_object* x_1392; lean_object* x_1393; lean_object* x_1394; lean_object* x_1395; lean_object* x_1396; lean_object* x_1397; lean_object* x_1398; lean_object* x_1399; lean_object* x_1400; +lean_dec(x_1); +x_1392 = lean_unsigned_to_nat(0u); +lean_inc(x_1384); +x_1393 = lean_array_set(x_7, x_1392, x_1384); +x_1394 = l_Lean_mkAppN(x_6, x_1393); +lean_dec(x_1393); +x_1395 = lean_ctor_get(x_1385, 0); +lean_inc(x_1395); +x_1396 = lean_ctor_get(x_1385, 1); +lean_inc(x_1396); +lean_dec(x_1385); +lean_inc(x_1394); +x_1397 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_1387, x_1384, x_1394); +x_1398 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1398, 0, x_1395); +lean_ctor_set(x_1398, 1, x_1396); +lean_ctor_set(x_1398, 2, x_1397); +if (lean_is_scalar(x_1386)) { + x_1399 = lean_alloc_ctor(0, 2, 0); +} else { + x_1399 = x_1386; +} +lean_ctor_set(x_1399, 0, x_1394); +lean_ctor_set(x_1399, 1, x_1398); +x_1400 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1400, 0, x_1399); +lean_ctor_set(x_1400, 1, x_1383); +return x_1400; } else { -lean_object* x_499; lean_object* x_500; lean_object* x_501; lean_object* x_502; lean_object* x_503; -lean_dec(x_3); -lean_dec(x_2); -x_499 = lean_ctor_get(x_485, 0); -lean_inc(x_499); -x_500 = lean_ctor_get(x_485, 1); -lean_inc(x_500); -lean_dec(x_485); -lean_inc(x_1); -x_501 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_486, x_484, x_1); -x_502 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_502, 0, x_499); -lean_ctor_set(x_502, 1, x_500); -lean_ctor_set(x_502, 2, x_501); -lean_inc(x_1); -x_503 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_502, x_7, x_8, x_9, x_10, x_483); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); +lean_object* x_1401; lean_object* x_1402; lean_object* x_1403; lean_object* x_1404; lean_object* x_1405; lean_object* x_1406; lean_dec(x_7); -lean_dec(x_5); -return x_503; +lean_dec(x_6); +x_1401 = lean_ctor_get(x_1385, 0); +lean_inc(x_1401); +x_1402 = lean_ctor_get(x_1385, 1); +lean_inc(x_1402); +lean_dec(x_1385); +lean_inc(x_1); +x_1403 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_1387, x_1384, x_1); +x_1404 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1404, 0, x_1401); +lean_ctor_set(x_1404, 1, x_1402); +lean_ctor_set(x_1404, 2, x_1403); +if (lean_is_scalar(x_1386)) { + x_1405 = lean_alloc_ctor(0, 2, 0); +} else { + x_1405 = x_1386; +} +lean_ctor_set(x_1405, 0, x_1); +lean_ctor_set(x_1405, 1, x_1404); +x_1406 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1406, 0, x_1405); +lean_ctor_set(x_1406, 1, x_1383); +return x_1406; } } else { -lean_object* x_504; lean_object* x_505; -lean_dec(x_486); -lean_dec(x_484); -lean_dec(x_480); -lean_dec(x_3); -lean_dec(x_2); -x_504 = lean_ctor_get(x_487, 0); -lean_inc(x_504); -lean_dec(x_487); -x_505 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_504, x_5, x_485, x_7, x_8, x_9, x_10, x_483); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); +lean_object* x_1407; lean_object* x_1408; lean_object* x_1409; +lean_dec(x_1387); +lean_dec(x_1384); +lean_dec(x_1336); lean_dec(x_7); -lean_dec(x_5); -return x_505; +lean_dec(x_6); +lean_dec(x_1); +x_1407 = lean_ctor_get(x_1388, 0); +lean_inc(x_1407); +lean_dec(x_1388); +if (lean_is_scalar(x_1386)) { + x_1408 = lean_alloc_ctor(0, 2, 0); +} else { + x_1408 = x_1386; +} +lean_ctor_set(x_1408, 0, x_1407); +lean_ctor_set(x_1408, 1, x_1385); +x_1409 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1409, 0, x_1408); +lean_ctor_set(x_1409, 1, x_1383); +return x_1409; +} } } else { -uint8_t x_506; -lean_dec(x_480); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); +uint8_t x_1410; +lean_dec(x_1336); lean_dec(x_7); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); +lean_dec(x_6); lean_dec(x_1); -x_506 = !lean_is_exclusive(x_481); -if (x_506 == 0) +x_1410 = !lean_is_exclusive(x_1337); +if (x_1410 == 0) { -return x_481; +return x_1337; } else { -lean_object* x_507; lean_object* x_508; lean_object* x_509; -x_507 = lean_ctor_get(x_481, 0); -x_508 = lean_ctor_get(x_481, 1); -lean_inc(x_508); -lean_inc(x_507); -lean_dec(x_481); -x_509 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_509, 0, x_507); -lean_ctor_set(x_509, 1, x_508); -return x_509; +lean_object* x_1411; lean_object* x_1412; lean_object* x_1413; +x_1411 = lean_ctor_get(x_1337, 0); +x_1412 = lean_ctor_get(x_1337, 1); +lean_inc(x_1412); +lean_inc(x_1411); +lean_dec(x_1337); +x_1413 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1413, 0, x_1411); +lean_ctor_set(x_1413, 1, x_1412); +return x_1413; } } } -block_544: -{ -lean_object* x_512; -lean_dec(x_511); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_2); -x_512 = l_Lean_Meta_getFunInfo(x_2, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_512) == 0) +block_1469: { -lean_object* x_513; lean_object* x_514; lean_object* x_515; lean_object* x_516; lean_object* x_517; lean_object* x_518; lean_object* x_519; uint8_t x_520; lean_object* x_521; lean_object* x_522; lean_object* x_523; -x_513 = lean_ctor_get(x_512, 0); -lean_inc(x_513); -x_514 = lean_ctor_get(x_512, 1); -lean_inc(x_514); -lean_dec(x_512); -x_515 = lean_ctor_get(x_513, 0); -lean_inc(x_515); -lean_dec(x_513); -x_516 = lean_array_get_size(x_3); -x_517 = lean_unsigned_to_nat(0u); -x_518 = lean_unsigned_to_nat(1u); -x_519 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_519, 0, x_517); -lean_ctor_set(x_519, 1, x_516); -lean_ctor_set(x_519, 2, x_518); -x_520 = 0; -x_521 = lean_box(x_520); -x_522 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_522, 0, x_3); -lean_ctor_set(x_522, 1, x_521); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); +lean_object* x_1416; +lean_dec(x_1415); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_6); +x_1416 = l_Lean_Meta_getFunInfo(x_6, x_11, x_12, x_13, x_14, x_15); +if (lean_obj_tag(x_1416) == 0) +{ +lean_object* x_1417; lean_object* x_1418; lean_object* x_1419; lean_object* x_1420; lean_object* x_1421; lean_object* x_1422; lean_object* x_1423; uint8_t x_1424; lean_object* x_1425; lean_object* x_1426; lean_object* x_1427; lean_object* x_1428; +x_1417 = lean_ctor_get(x_1416, 0); +lean_inc(x_1417); +x_1418 = lean_ctor_get(x_1416, 1); +lean_inc(x_1418); +lean_dec(x_1416); +x_1419 = lean_ctor_get(x_1417, 0); +lean_inc(x_1419); +lean_dec(x_1417); +x_1420 = lean_array_get_size(x_7); +x_1421 = lean_unsigned_to_nat(0u); +x_1422 = lean_unsigned_to_nat(1u); +lean_inc(x_1420); +x_1423 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1423, 0, x_1421); +lean_ctor_set(x_1423, 1, x_1420); +lean_ctor_set(x_1423, 2, x_1422); +x_1424 = 0; +x_1425 = lean_box(x_1424); lean_inc(x_7); -lean_inc(x_5); -lean_inc(x_2); -x_523 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8(x_2, x_515, x_519, x_519, x_522, x_517, lean_box(0), lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10, x_514); -lean_dec(x_519); -lean_dec(x_515); -if (lean_obj_tag(x_523) == 0) -{ -lean_object* x_524; lean_object* x_525; lean_object* x_526; uint8_t x_527; -x_524 = lean_ctor_get(x_523, 0); -lean_inc(x_524); -x_525 = lean_ctor_get(x_524, 0); -lean_inc(x_525); -x_526 = lean_ctor_get(x_525, 1); -lean_inc(x_526); -x_527 = lean_unbox(x_526); -lean_dec(x_526); -if (x_527 == 0) -{ -lean_object* x_528; lean_object* x_529; lean_object* x_530; -lean_dec(x_525); -lean_dec(x_2); -x_528 = lean_ctor_get(x_523, 1); -lean_inc(x_528); -lean_dec(x_523); -x_529 = lean_ctor_get(x_524, 1); -lean_inc(x_529); -lean_dec(x_524); -lean_inc(x_1); -x_530 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_529, x_7, x_8, x_9, x_10, x_528); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_530; +x_1426 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1426, 0, x_7); +lean_ctor_set(x_1426, 1, x_1425); +x_1427 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__11___closed__1; +lean_inc(x_1423); +lean_inc(x_6); +x_1428 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10(x_2, x_3, x_4, x_5, x_6, x_7, x_1427, x_1419, x_1420, x_1423, x_1423, x_1426, x_1421, lean_box(0), lean_box(0), x_9, x_10, x_11, x_12, x_13, x_14, x_1418); +if (lean_obj_tag(x_1428) == 0) +{ +lean_object* x_1429; lean_object* x_1430; lean_object* x_1431; uint8_t x_1432; +x_1429 = lean_ctor_get(x_1428, 0); +lean_inc(x_1429); +x_1430 = lean_ctor_get(x_1429, 0); +lean_inc(x_1430); +x_1431 = lean_ctor_get(x_1430, 1); +lean_inc(x_1431); +x_1432 = lean_unbox(x_1431); +lean_dec(x_1431); +if (x_1432 == 0) +{ +uint8_t x_1433; +lean_dec(x_1430); +lean_dec(x_6); +x_1433 = !lean_is_exclusive(x_1428); +if (x_1433 == 0) +{ +lean_object* x_1434; uint8_t x_1435; +x_1434 = lean_ctor_get(x_1428, 0); +lean_dec(x_1434); +x_1435 = !lean_is_exclusive(x_1429); +if (x_1435 == 0) +{ +lean_object* x_1436; +x_1436 = lean_ctor_get(x_1429, 0); +lean_dec(x_1436); +lean_ctor_set(x_1429, 0, x_1); +return x_1428; +} +else +{ +lean_object* x_1437; lean_object* x_1438; +x_1437 = lean_ctor_get(x_1429, 1); +lean_inc(x_1437); +lean_dec(x_1429); +x_1438 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1438, 0, x_1); +lean_ctor_set(x_1438, 1, x_1437); +lean_ctor_set(x_1428, 0, x_1438); +return x_1428; +} +} +else +{ +lean_object* x_1439; lean_object* x_1440; lean_object* x_1441; lean_object* x_1442; lean_object* x_1443; +x_1439 = lean_ctor_get(x_1428, 1); +lean_inc(x_1439); +lean_dec(x_1428); +x_1440 = lean_ctor_get(x_1429, 1); +lean_inc(x_1440); +if (lean_is_exclusive(x_1429)) { + lean_ctor_release(x_1429, 0); + lean_ctor_release(x_1429, 1); + x_1441 = x_1429; +} else { + lean_dec_ref(x_1429); + x_1441 = lean_box(0); +} +if (lean_is_scalar(x_1441)) { + x_1442 = lean_alloc_ctor(0, 2, 0); +} else { + x_1442 = x_1441; +} +lean_ctor_set(x_1442, 0, x_1); +lean_ctor_set(x_1442, 1, x_1440); +x_1443 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1443, 0, x_1442); +lean_ctor_set(x_1443, 1, x_1439); +return x_1443; +} } else { -lean_object* x_531; lean_object* x_532; lean_object* x_533; lean_object* x_534; lean_object* x_535; -x_531 = lean_ctor_get(x_523, 1); -lean_inc(x_531); -lean_dec(x_523); -x_532 = lean_ctor_get(x_524, 1); -lean_inc(x_532); -lean_dec(x_524); -x_533 = lean_ctor_get(x_525, 0); -lean_inc(x_533); -lean_dec(x_525); -x_534 = l_Lean_mkAppN(x_2, x_533); -lean_dec(x_533); -x_535 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_534, x_5, x_532, x_7, x_8, x_9, x_10, x_531); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_535; +uint8_t x_1444; +lean_dec(x_1); +x_1444 = !lean_is_exclusive(x_1428); +if (x_1444 == 0) +{ +lean_object* x_1445; uint8_t x_1446; +x_1445 = lean_ctor_get(x_1428, 0); +lean_dec(x_1445); +x_1446 = !lean_is_exclusive(x_1429); +if (x_1446 == 0) +{ +lean_object* x_1447; lean_object* x_1448; lean_object* x_1449; +x_1447 = lean_ctor_get(x_1429, 0); +lean_dec(x_1447); +x_1448 = lean_ctor_get(x_1430, 0); +lean_inc(x_1448); +lean_dec(x_1430); +x_1449 = l_Lean_mkAppN(x_6, x_1448); +lean_dec(x_1448); +lean_ctor_set(x_1429, 0, x_1449); +return x_1428; +} +else +{ +lean_object* x_1450; lean_object* x_1451; lean_object* x_1452; lean_object* x_1453; +x_1450 = lean_ctor_get(x_1429, 1); +lean_inc(x_1450); +lean_dec(x_1429); +x_1451 = lean_ctor_get(x_1430, 0); +lean_inc(x_1451); +lean_dec(x_1430); +x_1452 = l_Lean_mkAppN(x_6, x_1451); +lean_dec(x_1451); +x_1453 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1453, 0, x_1452); +lean_ctor_set(x_1453, 1, x_1450); +lean_ctor_set(x_1428, 0, x_1453); +return x_1428; +} +} +else +{ +lean_object* x_1454; lean_object* x_1455; lean_object* x_1456; lean_object* x_1457; lean_object* x_1458; lean_object* x_1459; lean_object* x_1460; +x_1454 = lean_ctor_get(x_1428, 1); +lean_inc(x_1454); +lean_dec(x_1428); +x_1455 = lean_ctor_get(x_1429, 1); +lean_inc(x_1455); +if (lean_is_exclusive(x_1429)) { + lean_ctor_release(x_1429, 0); + lean_ctor_release(x_1429, 1); + x_1456 = x_1429; +} else { + lean_dec_ref(x_1429); + x_1456 = lean_box(0); +} +x_1457 = lean_ctor_get(x_1430, 0); +lean_inc(x_1457); +lean_dec(x_1430); +x_1458 = l_Lean_mkAppN(x_6, x_1457); +lean_dec(x_1457); +if (lean_is_scalar(x_1456)) { + x_1459 = lean_alloc_ctor(0, 2, 0); +} else { + x_1459 = x_1456; +} +lean_ctor_set(x_1459, 0, x_1458); +lean_ctor_set(x_1459, 1, x_1455); +x_1460 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1460, 0, x_1459); +lean_ctor_set(x_1460, 1, x_1454); +return x_1460; +} } } else { -uint8_t x_536; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -lean_dec(x_2); +uint8_t x_1461; +lean_dec(x_6); lean_dec(x_1); -x_536 = !lean_is_exclusive(x_523); -if (x_536 == 0) +x_1461 = !lean_is_exclusive(x_1428); +if (x_1461 == 0) { -return x_523; +return x_1428; } else { -lean_object* x_537; lean_object* x_538; lean_object* x_539; -x_537 = lean_ctor_get(x_523, 0); -x_538 = lean_ctor_get(x_523, 1); -lean_inc(x_538); -lean_inc(x_537); -lean_dec(x_523); -x_539 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_539, 0, x_537); -lean_ctor_set(x_539, 1, x_538); -return x_539; +lean_object* x_1462; lean_object* x_1463; lean_object* x_1464; +x_1462 = lean_ctor_get(x_1428, 0); +x_1463 = lean_ctor_get(x_1428, 1); +lean_inc(x_1463); +lean_inc(x_1462); +lean_dec(x_1428); +x_1464 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1464, 0, x_1462); +lean_ctor_set(x_1464, 1, x_1463); +return x_1464; } } } else { -uint8_t x_540; +uint8_t x_1465; +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); -lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); +lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_540 = !lean_is_exclusive(x_512); -if (x_540 == 0) +x_1465 = !lean_is_exclusive(x_1416); +if (x_1465 == 0) { -return x_512; +return x_1416; } else { -lean_object* x_541; lean_object* x_542; lean_object* x_543; -x_541 = lean_ctor_get(x_512, 0); -x_542 = lean_ctor_get(x_512, 1); -lean_inc(x_542); -lean_inc(x_541); -lean_dec(x_512); -x_543 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_543, 0, x_541); -lean_ctor_set(x_543, 1, x_542); -return x_543; +lean_object* x_1466; lean_object* x_1467; lean_object* x_1468; +x_1466 = lean_ctor_get(x_1416, 0); +x_1467 = lean_ctor_get(x_1416, 1); +lean_inc(x_1467); +lean_inc(x_1466); +lean_dec(x_1416); +x_1468 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1468, 0, x_1466); +lean_ctor_set(x_1468, 1, x_1467); +return x_1468; } } } } -case 8: +default: { -lean_object* x_557; lean_object* x_588; lean_object* x_622; uint8_t x_623; -lean_dec(x_4); -x_622 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__4; -x_623 = l_Lean_Expr_isConstOf(x_2, x_622); -if (x_623 == 0) +lean_object* x_1482; lean_object* x_1561; lean_object* x_1616; uint8_t x_1617; +lean_dec(x_8); +x_1616 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__11___closed__5; +x_1617 = l_Lean_Expr_isConstOf(x_6, x_1616); +if (x_1617 == 0) { -lean_object* x_624; -x_624 = lean_box(0); -x_588 = x_624; -goto block_621; +lean_object* x_1618; +x_1618 = lean_box(0); +x_1561 = x_1618; +goto block_1615; } else { -lean_object* x_625; lean_object* x_626; uint8_t x_627; -x_625 = lean_array_get_size(x_3); -x_626 = lean_unsigned_to_nat(2u); -x_627 = lean_nat_dec_eq(x_625, x_626); -if (x_627 == 0) +lean_object* x_1619; lean_object* x_1620; uint8_t x_1621; +x_1619 = lean_array_get_size(x_7); +x_1620 = lean_unsigned_to_nat(2u); +x_1621 = lean_nat_dec_eq(x_1619, x_1620); +if (x_1621 == 0) { -lean_object* x_628; -lean_dec(x_625); -x_628 = lean_box(0); -x_588 = x_628; -goto block_621; +lean_object* x_1622; +lean_dec(x_1619); +x_1622 = lean_box(0); +x_1561 = x_1622; +goto block_1615; } else { -lean_object* x_629; uint8_t x_630; -x_629 = lean_unsigned_to_nat(0u); -x_630 = lean_nat_dec_lt(x_629, x_625); -lean_dec(x_625); -if (x_630 == 0) +lean_object* x_1623; uint8_t x_1624; +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_1623 = lean_unsigned_to_nat(0u); +x_1624 = lean_nat_dec_lt(x_1623, x_1619); +lean_dec(x_1619); +if (x_1624 == 0) { -lean_object* x_631; lean_object* x_632; -x_631 = l_Lean_instInhabitedExpr; -x_632 = l_outOfBounds___rarg(x_631); -x_557 = x_632; -goto block_587; +lean_object* x_1625; lean_object* x_1626; +x_1625 = l_Lean_instInhabitedExpr; +x_1626 = l_outOfBounds___rarg(x_1625); +x_1482 = x_1626; +goto block_1560; } else { -lean_object* x_633; -x_633 = lean_array_fget(x_3, x_629); -x_557 = x_633; -goto block_587; +lean_object* x_1627; +x_1627 = lean_array_fget(x_7, x_1623); +x_1482 = x_1627; +goto block_1560; } } } -block_587: +block_1560: { -lean_object* x_558; -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_5); -lean_inc(x_557); -x_558 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_557, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_558) == 0) -{ -lean_object* x_559; lean_object* x_560; lean_object* x_561; lean_object* x_562; lean_object* x_563; lean_object* x_564; -x_559 = lean_ctor_get(x_558, 0); -lean_inc(x_559); -x_560 = lean_ctor_get(x_558, 1); -lean_inc(x_560); -lean_dec(x_558); -x_561 = lean_ctor_get(x_559, 0); -lean_inc(x_561); -x_562 = lean_ctor_get(x_559, 1); -lean_inc(x_562); -lean_dec(x_559); -x_563 = lean_ctor_get(x_562, 2); -lean_inc(x_563); -lean_inc(x_563); -x_564 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_563, x_561); -if (lean_obj_tag(x_564) == 0) -{ -size_t x_565; size_t x_566; uint8_t x_567; -x_565 = lean_ptr_addr(x_557); -lean_dec(x_557); -x_566 = lean_ptr_addr(x_561); -x_567 = lean_usize_dec_eq(x_565, x_566); -if (x_567 == 0) -{ -lean_object* x_568; lean_object* x_569; lean_object* x_570; lean_object* x_571; lean_object* x_572; lean_object* x_573; lean_object* x_574; lean_object* x_575; -x_568 = lean_unsigned_to_nat(0u); -lean_inc(x_561); -x_569 = lean_array_set(x_3, x_568, x_561); -x_570 = l_Lean_mkAppN(x_2, x_569); -lean_dec(x_569); -x_571 = lean_ctor_get(x_562, 0); -lean_inc(x_571); -x_572 = lean_ctor_get(x_562, 1); -lean_inc(x_572); -lean_dec(x_562); -lean_inc(x_570); -x_573 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_563, x_561, x_570); -x_574 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_574, 0, x_571); -lean_ctor_set(x_574, 1, x_572); -lean_ctor_set(x_574, 2, x_573); -x_575 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_570, x_5, x_574, x_7, x_8, x_9, x_10, x_560); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); +lean_object* x_1483; +lean_inc(x_1482); +x_1483 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_1482, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +if (lean_obj_tag(x_1483) == 0) +{ +uint8_t x_1484; +x_1484 = !lean_is_exclusive(x_1483); +if (x_1484 == 0) +{ +lean_object* x_1485; uint8_t x_1486; +x_1485 = lean_ctor_get(x_1483, 0); +x_1486 = !lean_is_exclusive(x_1485); +if (x_1486 == 0) +{ +lean_object* x_1487; lean_object* x_1488; lean_object* x_1489; lean_object* x_1490; +x_1487 = lean_ctor_get(x_1485, 0); +x_1488 = lean_ctor_get(x_1485, 1); +x_1489 = lean_ctor_get(x_1488, 2); +lean_inc(x_1489); +lean_inc(x_1489); +x_1490 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_1489, x_1487); +if (lean_obj_tag(x_1490) == 0) +{ +size_t x_1491; size_t x_1492; uint8_t x_1493; +x_1491 = lean_ptr_addr(x_1482); +lean_dec(x_1482); +x_1492 = lean_ptr_addr(x_1487); +x_1493 = lean_usize_dec_eq(x_1491, x_1492); +if (x_1493 == 0) +{ +lean_object* x_1494; lean_object* x_1495; lean_object* x_1496; lean_object* x_1497; lean_object* x_1498; lean_object* x_1499; lean_object* x_1500; +lean_dec(x_1); +x_1494 = lean_unsigned_to_nat(0u); +lean_inc(x_1487); +x_1495 = lean_array_set(x_7, x_1494, x_1487); +x_1496 = l_Lean_mkAppN(x_6, x_1495); +lean_dec(x_1495); +x_1497 = lean_ctor_get(x_1488, 0); +lean_inc(x_1497); +x_1498 = lean_ctor_get(x_1488, 1); +lean_inc(x_1498); +lean_dec(x_1488); +lean_inc(x_1496); +x_1499 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_1489, x_1487, x_1496); +x_1500 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1500, 0, x_1497); +lean_ctor_set(x_1500, 1, x_1498); +lean_ctor_set(x_1500, 2, x_1499); +lean_ctor_set(x_1485, 1, x_1500); +lean_ctor_set(x_1485, 0, x_1496); +return x_1483; +} +else +{ +lean_object* x_1501; lean_object* x_1502; lean_object* x_1503; lean_object* x_1504; lean_dec(x_7); -lean_dec(x_5); -return x_575; +lean_dec(x_6); +x_1501 = lean_ctor_get(x_1488, 0); +lean_inc(x_1501); +x_1502 = lean_ctor_get(x_1488, 1); +lean_inc(x_1502); +lean_dec(x_1488); +lean_inc(x_1); +x_1503 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_1489, x_1487, x_1); +x_1504 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1504, 0, x_1501); +lean_ctor_set(x_1504, 1, x_1502); +lean_ctor_set(x_1504, 2, x_1503); +lean_ctor_set(x_1485, 1, x_1504); +lean_ctor_set(x_1485, 0, x_1); +return x_1483; +} } else { -lean_object* x_576; lean_object* x_577; lean_object* x_578; lean_object* x_579; lean_object* x_580; -lean_dec(x_3); -lean_dec(x_2); -x_576 = lean_ctor_get(x_562, 0); -lean_inc(x_576); -x_577 = lean_ctor_get(x_562, 1); -lean_inc(x_577); -lean_dec(x_562); -lean_inc(x_1); -x_578 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_563, x_561, x_1); -x_579 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_579, 0, x_576); -lean_ctor_set(x_579, 1, x_577); -lean_ctor_set(x_579, 2, x_578); +lean_object* x_1505; +lean_dec(x_1489); +lean_dec(x_1487); +lean_dec(x_1482); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_1); +x_1505 = lean_ctor_get(x_1490, 0); +lean_inc(x_1505); +lean_dec(x_1490); +lean_ctor_set(x_1485, 0, x_1505); +return x_1483; +} +} +else +{ +lean_object* x_1506; lean_object* x_1507; lean_object* x_1508; lean_object* x_1509; +x_1506 = lean_ctor_get(x_1485, 0); +x_1507 = lean_ctor_get(x_1485, 1); +lean_inc(x_1507); +lean_inc(x_1506); +lean_dec(x_1485); +x_1508 = lean_ctor_get(x_1507, 2); +lean_inc(x_1508); +lean_inc(x_1508); +x_1509 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_1508, x_1506); +if (lean_obj_tag(x_1509) == 0) +{ +size_t x_1510; size_t x_1511; uint8_t x_1512; +x_1510 = lean_ptr_addr(x_1482); +lean_dec(x_1482); +x_1511 = lean_ptr_addr(x_1506); +x_1512 = lean_usize_dec_eq(x_1510, x_1511); +if (x_1512 == 0) +{ +lean_object* x_1513; lean_object* x_1514; lean_object* x_1515; lean_object* x_1516; lean_object* x_1517; lean_object* x_1518; lean_object* x_1519; lean_object* x_1520; +lean_dec(x_1); +x_1513 = lean_unsigned_to_nat(0u); +lean_inc(x_1506); +x_1514 = lean_array_set(x_7, x_1513, x_1506); +x_1515 = l_Lean_mkAppN(x_6, x_1514); +lean_dec(x_1514); +x_1516 = lean_ctor_get(x_1507, 0); +lean_inc(x_1516); +x_1517 = lean_ctor_get(x_1507, 1); +lean_inc(x_1517); +lean_dec(x_1507); +lean_inc(x_1515); +x_1518 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_1508, x_1506, x_1515); +x_1519 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1519, 0, x_1516); +lean_ctor_set(x_1519, 1, x_1517); +lean_ctor_set(x_1519, 2, x_1518); +x_1520 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1520, 0, x_1515); +lean_ctor_set(x_1520, 1, x_1519); +lean_ctor_set(x_1483, 0, x_1520); +return x_1483; +} +else +{ +lean_object* x_1521; lean_object* x_1522; lean_object* x_1523; lean_object* x_1524; lean_object* x_1525; +lean_dec(x_7); +lean_dec(x_6); +x_1521 = lean_ctor_get(x_1507, 0); +lean_inc(x_1521); +x_1522 = lean_ctor_get(x_1507, 1); +lean_inc(x_1522); +lean_dec(x_1507); lean_inc(x_1); -x_580 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_579, x_7, x_8, x_9, x_10, x_560); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); +x_1523 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_1508, x_1506, x_1); +x_1524 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1524, 0, x_1521); +lean_ctor_set(x_1524, 1, x_1522); +lean_ctor_set(x_1524, 2, x_1523); +x_1525 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1525, 0, x_1); +lean_ctor_set(x_1525, 1, x_1524); +lean_ctor_set(x_1483, 0, x_1525); +return x_1483; +} +} +else +{ +lean_object* x_1526; lean_object* x_1527; +lean_dec(x_1508); +lean_dec(x_1506); +lean_dec(x_1482); lean_dec(x_7); -lean_dec(x_5); -return x_580; +lean_dec(x_6); +lean_dec(x_1); +x_1526 = lean_ctor_get(x_1509, 0); +lean_inc(x_1526); +lean_dec(x_1509); +x_1527 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1527, 0, x_1526); +lean_ctor_set(x_1527, 1, x_1507); +lean_ctor_set(x_1483, 0, x_1527); +return x_1483; +} +} +} +else +{ +lean_object* x_1528; lean_object* x_1529; lean_object* x_1530; lean_object* x_1531; lean_object* x_1532; lean_object* x_1533; lean_object* x_1534; +x_1528 = lean_ctor_get(x_1483, 0); +x_1529 = lean_ctor_get(x_1483, 1); +lean_inc(x_1529); +lean_inc(x_1528); +lean_dec(x_1483); +x_1530 = lean_ctor_get(x_1528, 0); +lean_inc(x_1530); +x_1531 = lean_ctor_get(x_1528, 1); +lean_inc(x_1531); +if (lean_is_exclusive(x_1528)) { + lean_ctor_release(x_1528, 0); + lean_ctor_release(x_1528, 1); + x_1532 = x_1528; +} else { + lean_dec_ref(x_1528); + x_1532 = lean_box(0); +} +x_1533 = lean_ctor_get(x_1531, 2); +lean_inc(x_1533); +lean_inc(x_1533); +x_1534 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_1533, x_1530); +if (lean_obj_tag(x_1534) == 0) +{ +size_t x_1535; size_t x_1536; uint8_t x_1537; +x_1535 = lean_ptr_addr(x_1482); +lean_dec(x_1482); +x_1536 = lean_ptr_addr(x_1530); +x_1537 = lean_usize_dec_eq(x_1535, x_1536); +if (x_1537 == 0) +{ +lean_object* x_1538; lean_object* x_1539; lean_object* x_1540; lean_object* x_1541; lean_object* x_1542; lean_object* x_1543; lean_object* x_1544; lean_object* x_1545; lean_object* x_1546; +lean_dec(x_1); +x_1538 = lean_unsigned_to_nat(0u); +lean_inc(x_1530); +x_1539 = lean_array_set(x_7, x_1538, x_1530); +x_1540 = l_Lean_mkAppN(x_6, x_1539); +lean_dec(x_1539); +x_1541 = lean_ctor_get(x_1531, 0); +lean_inc(x_1541); +x_1542 = lean_ctor_get(x_1531, 1); +lean_inc(x_1542); +lean_dec(x_1531); +lean_inc(x_1540); +x_1543 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_1533, x_1530, x_1540); +x_1544 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1544, 0, x_1541); +lean_ctor_set(x_1544, 1, x_1542); +lean_ctor_set(x_1544, 2, x_1543); +if (lean_is_scalar(x_1532)) { + x_1545 = lean_alloc_ctor(0, 2, 0); +} else { + x_1545 = x_1532; } +lean_ctor_set(x_1545, 0, x_1540); +lean_ctor_set(x_1545, 1, x_1544); +x_1546 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1546, 0, x_1545); +lean_ctor_set(x_1546, 1, x_1529); +return x_1546; } else { -lean_object* x_581; lean_object* x_582; -lean_dec(x_563); -lean_dec(x_561); -lean_dec(x_557); -lean_dec(x_3); -lean_dec(x_2); -x_581 = lean_ctor_get(x_564, 0); -lean_inc(x_581); -lean_dec(x_564); -x_582 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_581, x_5, x_562, x_7, x_8, x_9, x_10, x_560); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); +lean_object* x_1547; lean_object* x_1548; lean_object* x_1549; lean_object* x_1550; lean_object* x_1551; lean_object* x_1552; lean_dec(x_7); -lean_dec(x_5); -return x_582; +lean_dec(x_6); +x_1547 = lean_ctor_get(x_1531, 0); +lean_inc(x_1547); +x_1548 = lean_ctor_get(x_1531, 1); +lean_inc(x_1548); +lean_dec(x_1531); +lean_inc(x_1); +x_1549 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_1533, x_1530, x_1); +x_1550 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1550, 0, x_1547); +lean_ctor_set(x_1550, 1, x_1548); +lean_ctor_set(x_1550, 2, x_1549); +if (lean_is_scalar(x_1532)) { + x_1551 = lean_alloc_ctor(0, 2, 0); +} else { + x_1551 = x_1532; +} +lean_ctor_set(x_1551, 0, x_1); +lean_ctor_set(x_1551, 1, x_1550); +x_1552 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1552, 0, x_1551); +lean_ctor_set(x_1552, 1, x_1529); +return x_1552; } } else { -uint8_t x_583; -lean_dec(x_557); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); +lean_object* x_1553; lean_object* x_1554; lean_object* x_1555; +lean_dec(x_1533); +lean_dec(x_1530); +lean_dec(x_1482); lean_dec(x_7); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); +lean_dec(x_6); lean_dec(x_1); -x_583 = !lean_is_exclusive(x_558); -if (x_583 == 0) +x_1553 = lean_ctor_get(x_1534, 0); +lean_inc(x_1553); +lean_dec(x_1534); +if (lean_is_scalar(x_1532)) { + x_1554 = lean_alloc_ctor(0, 2, 0); +} else { + x_1554 = x_1532; +} +lean_ctor_set(x_1554, 0, x_1553); +lean_ctor_set(x_1554, 1, x_1531); +x_1555 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1555, 0, x_1554); +lean_ctor_set(x_1555, 1, x_1529); +return x_1555; +} +} +} +else +{ +uint8_t x_1556; +lean_dec(x_1482); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_1); +x_1556 = !lean_is_exclusive(x_1483); +if (x_1556 == 0) { -return x_558; +return x_1483; } else { -lean_object* x_584; lean_object* x_585; lean_object* x_586; -x_584 = lean_ctor_get(x_558, 0); -x_585 = lean_ctor_get(x_558, 1); -lean_inc(x_585); -lean_inc(x_584); -lean_dec(x_558); -x_586 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_586, 0, x_584); -lean_ctor_set(x_586, 1, x_585); -return x_586; +lean_object* x_1557; lean_object* x_1558; lean_object* x_1559; +x_1557 = lean_ctor_get(x_1483, 0); +x_1558 = lean_ctor_get(x_1483, 1); +lean_inc(x_1558); +lean_inc(x_1557); +lean_dec(x_1483); +x_1559 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1559, 0, x_1557); +lean_ctor_set(x_1559, 1, x_1558); +return x_1559; } } } -block_621: +block_1615: { -lean_object* x_589; -lean_dec(x_588); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_2); -x_589 = l_Lean_Meta_getFunInfo(x_2, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_589) == 0) -{ -lean_object* x_590; lean_object* x_591; lean_object* x_592; lean_object* x_593; lean_object* x_594; lean_object* x_595; lean_object* x_596; uint8_t x_597; lean_object* x_598; lean_object* x_599; lean_object* x_600; -x_590 = lean_ctor_get(x_589, 0); -lean_inc(x_590); -x_591 = lean_ctor_get(x_589, 1); -lean_inc(x_591); -lean_dec(x_589); -x_592 = lean_ctor_get(x_590, 0); -lean_inc(x_592); -lean_dec(x_590); -x_593 = lean_array_get_size(x_3); -x_594 = lean_unsigned_to_nat(0u); -x_595 = lean_unsigned_to_nat(1u); -x_596 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_596, 0, x_594); -lean_ctor_set(x_596, 1, x_593); -lean_ctor_set(x_596, 2, x_595); -x_597 = 0; -x_598 = lean_box(x_597); -x_599 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_599, 0, x_3); -lean_ctor_set(x_599, 1, x_598); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); +lean_object* x_1562; +lean_dec(x_1561); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_6); +x_1562 = l_Lean_Meta_getFunInfo(x_6, x_11, x_12, x_13, x_14, x_15); +if (lean_obj_tag(x_1562) == 0) +{ +lean_object* x_1563; lean_object* x_1564; lean_object* x_1565; lean_object* x_1566; lean_object* x_1567; lean_object* x_1568; lean_object* x_1569; uint8_t x_1570; lean_object* x_1571; lean_object* x_1572; lean_object* x_1573; lean_object* x_1574; +x_1563 = lean_ctor_get(x_1562, 0); +lean_inc(x_1563); +x_1564 = lean_ctor_get(x_1562, 1); +lean_inc(x_1564); +lean_dec(x_1562); +x_1565 = lean_ctor_get(x_1563, 0); +lean_inc(x_1565); +lean_dec(x_1563); +x_1566 = lean_array_get_size(x_7); +x_1567 = lean_unsigned_to_nat(0u); +x_1568 = lean_unsigned_to_nat(1u); +lean_inc(x_1566); +x_1569 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1569, 0, x_1567); +lean_ctor_set(x_1569, 1, x_1566); +lean_ctor_set(x_1569, 2, x_1568); +x_1570 = 0; +x_1571 = lean_box(x_1570); lean_inc(x_7); -lean_inc(x_5); -lean_inc(x_2); -x_600 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8(x_2, x_592, x_596, x_596, x_599, x_594, lean_box(0), lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10, x_591); -lean_dec(x_596); -lean_dec(x_592); -if (lean_obj_tag(x_600) == 0) -{ -lean_object* x_601; lean_object* x_602; lean_object* x_603; uint8_t x_604; -x_601 = lean_ctor_get(x_600, 0); -lean_inc(x_601); -x_602 = lean_ctor_get(x_601, 0); -lean_inc(x_602); -x_603 = lean_ctor_get(x_602, 1); -lean_inc(x_603); -x_604 = lean_unbox(x_603); -lean_dec(x_603); -if (x_604 == 0) -{ -lean_object* x_605; lean_object* x_606; lean_object* x_607; -lean_dec(x_602); -lean_dec(x_2); -x_605 = lean_ctor_get(x_600, 1); -lean_inc(x_605); -lean_dec(x_600); -x_606 = lean_ctor_get(x_601, 1); -lean_inc(x_606); -lean_dec(x_601); -lean_inc(x_1); -x_607 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_606, x_7, x_8, x_9, x_10, x_605); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_607; +x_1572 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1572, 0, x_7); +lean_ctor_set(x_1572, 1, x_1571); +x_1573 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__11___closed__1; +lean_inc(x_1569); +lean_inc(x_6); +x_1574 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10(x_2, x_3, x_4, x_5, x_6, x_7, x_1573, x_1565, x_1566, x_1569, x_1569, x_1572, x_1567, lean_box(0), lean_box(0), x_9, x_10, x_11, x_12, x_13, x_14, x_1564); +if (lean_obj_tag(x_1574) == 0) +{ +lean_object* x_1575; lean_object* x_1576; lean_object* x_1577; uint8_t x_1578; +x_1575 = lean_ctor_get(x_1574, 0); +lean_inc(x_1575); +x_1576 = lean_ctor_get(x_1575, 0); +lean_inc(x_1576); +x_1577 = lean_ctor_get(x_1576, 1); +lean_inc(x_1577); +x_1578 = lean_unbox(x_1577); +lean_dec(x_1577); +if (x_1578 == 0) +{ +uint8_t x_1579; +lean_dec(x_1576); +lean_dec(x_6); +x_1579 = !lean_is_exclusive(x_1574); +if (x_1579 == 0) +{ +lean_object* x_1580; uint8_t x_1581; +x_1580 = lean_ctor_get(x_1574, 0); +lean_dec(x_1580); +x_1581 = !lean_is_exclusive(x_1575); +if (x_1581 == 0) +{ +lean_object* x_1582; +x_1582 = lean_ctor_get(x_1575, 0); +lean_dec(x_1582); +lean_ctor_set(x_1575, 0, x_1); +return x_1574; +} +else +{ +lean_object* x_1583; lean_object* x_1584; +x_1583 = lean_ctor_get(x_1575, 1); +lean_inc(x_1583); +lean_dec(x_1575); +x_1584 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1584, 0, x_1); +lean_ctor_set(x_1584, 1, x_1583); +lean_ctor_set(x_1574, 0, x_1584); +return x_1574; +} +} +else +{ +lean_object* x_1585; lean_object* x_1586; lean_object* x_1587; lean_object* x_1588; lean_object* x_1589; +x_1585 = lean_ctor_get(x_1574, 1); +lean_inc(x_1585); +lean_dec(x_1574); +x_1586 = lean_ctor_get(x_1575, 1); +lean_inc(x_1586); +if (lean_is_exclusive(x_1575)) { + lean_ctor_release(x_1575, 0); + lean_ctor_release(x_1575, 1); + x_1587 = x_1575; +} else { + lean_dec_ref(x_1575); + x_1587 = lean_box(0); +} +if (lean_is_scalar(x_1587)) { + x_1588 = lean_alloc_ctor(0, 2, 0); +} else { + x_1588 = x_1587; +} +lean_ctor_set(x_1588, 0, x_1); +lean_ctor_set(x_1588, 1, x_1586); +x_1589 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1589, 0, x_1588); +lean_ctor_set(x_1589, 1, x_1585); +return x_1589; +} } else { -lean_object* x_608; lean_object* x_609; lean_object* x_610; lean_object* x_611; lean_object* x_612; -x_608 = lean_ctor_get(x_600, 1); -lean_inc(x_608); -lean_dec(x_600); -x_609 = lean_ctor_get(x_601, 1); -lean_inc(x_609); -lean_dec(x_601); -x_610 = lean_ctor_get(x_602, 0); -lean_inc(x_610); -lean_dec(x_602); -x_611 = l_Lean_mkAppN(x_2, x_610); -lean_dec(x_610); -x_612 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_611, x_5, x_609, x_7, x_8, x_9, x_10, x_608); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_612; +uint8_t x_1590; +lean_dec(x_1); +x_1590 = !lean_is_exclusive(x_1574); +if (x_1590 == 0) +{ +lean_object* x_1591; uint8_t x_1592; +x_1591 = lean_ctor_get(x_1574, 0); +lean_dec(x_1591); +x_1592 = !lean_is_exclusive(x_1575); +if (x_1592 == 0) +{ +lean_object* x_1593; lean_object* x_1594; lean_object* x_1595; +x_1593 = lean_ctor_get(x_1575, 0); +lean_dec(x_1593); +x_1594 = lean_ctor_get(x_1576, 0); +lean_inc(x_1594); +lean_dec(x_1576); +x_1595 = l_Lean_mkAppN(x_6, x_1594); +lean_dec(x_1594); +lean_ctor_set(x_1575, 0, x_1595); +return x_1574; +} +else +{ +lean_object* x_1596; lean_object* x_1597; lean_object* x_1598; lean_object* x_1599; +x_1596 = lean_ctor_get(x_1575, 1); +lean_inc(x_1596); +lean_dec(x_1575); +x_1597 = lean_ctor_get(x_1576, 0); +lean_inc(x_1597); +lean_dec(x_1576); +x_1598 = l_Lean_mkAppN(x_6, x_1597); +lean_dec(x_1597); +x_1599 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1599, 0, x_1598); +lean_ctor_set(x_1599, 1, x_1596); +lean_ctor_set(x_1574, 0, x_1599); +return x_1574; +} +} +else +{ +lean_object* x_1600; lean_object* x_1601; lean_object* x_1602; lean_object* x_1603; lean_object* x_1604; lean_object* x_1605; lean_object* x_1606; +x_1600 = lean_ctor_get(x_1574, 1); +lean_inc(x_1600); +lean_dec(x_1574); +x_1601 = lean_ctor_get(x_1575, 1); +lean_inc(x_1601); +if (lean_is_exclusive(x_1575)) { + lean_ctor_release(x_1575, 0); + lean_ctor_release(x_1575, 1); + x_1602 = x_1575; +} else { + lean_dec_ref(x_1575); + x_1602 = lean_box(0); +} +x_1603 = lean_ctor_get(x_1576, 0); +lean_inc(x_1603); +lean_dec(x_1576); +x_1604 = l_Lean_mkAppN(x_6, x_1603); +lean_dec(x_1603); +if (lean_is_scalar(x_1602)) { + x_1605 = lean_alloc_ctor(0, 2, 0); +} else { + x_1605 = x_1602; +} +lean_ctor_set(x_1605, 0, x_1604); +lean_ctor_set(x_1605, 1, x_1601); +x_1606 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1606, 0, x_1605); +lean_ctor_set(x_1606, 1, x_1600); +return x_1606; +} } } else { -uint8_t x_613; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -lean_dec(x_2); +uint8_t x_1607; +lean_dec(x_6); lean_dec(x_1); -x_613 = !lean_is_exclusive(x_600); -if (x_613 == 0) +x_1607 = !lean_is_exclusive(x_1574); +if (x_1607 == 0) { -return x_600; +return x_1574; } else { -lean_object* x_614; lean_object* x_615; lean_object* x_616; -x_614 = lean_ctor_get(x_600, 0); -x_615 = lean_ctor_get(x_600, 1); -lean_inc(x_615); -lean_inc(x_614); -lean_dec(x_600); -x_616 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_616, 0, x_614); -lean_ctor_set(x_616, 1, x_615); -return x_616; +lean_object* x_1608; lean_object* x_1609; lean_object* x_1610; +x_1608 = lean_ctor_get(x_1574, 0); +x_1609 = lean_ctor_get(x_1574, 1); +lean_inc(x_1609); +lean_inc(x_1608); +lean_dec(x_1574); +x_1610 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1610, 0, x_1608); +lean_ctor_set(x_1610, 1, x_1609); +return x_1610; } } } else { -uint8_t x_617; +uint8_t x_1611; +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); -lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); +lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_617 = !lean_is_exclusive(x_589); -if (x_617 == 0) +x_1611 = !lean_is_exclusive(x_1562); +if (x_1611 == 0) { -return x_589; +return x_1562; } else { -lean_object* x_618; lean_object* x_619; lean_object* x_620; -x_618 = lean_ctor_get(x_589, 0); -x_619 = lean_ctor_get(x_589, 1); -lean_inc(x_619); -lean_inc(x_618); -lean_dec(x_589); -x_620 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_620, 0, x_618); -lean_ctor_set(x_620, 1, x_619); -return x_620; +lean_object* x_1612; lean_object* x_1613; lean_object* x_1614; +x_1612 = lean_ctor_get(x_1562, 0); +x_1613 = lean_ctor_get(x_1562, 1); +lean_inc(x_1613); +lean_inc(x_1612); +lean_dec(x_1562); +x_1614 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1614, 0, x_1612); +lean_ctor_set(x_1614, 1, x_1613); +return x_1614; } } } } -case 9: -{ -lean_object* x_634; lean_object* x_665; lean_object* x_699; uint8_t x_700; -lean_dec(x_4); -x_699 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__4; -x_700 = l_Lean_Expr_isConstOf(x_2, x_699); -if (x_700 == 0) -{ -lean_object* x_701; -x_701 = lean_box(0); -x_665 = x_701; -goto block_698; } -else +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__12(lean_object* x_1, lean_object* x_2) { +_start: { -lean_object* x_702; lean_object* x_703; uint8_t x_704; -x_702 = lean_array_get_size(x_3); -x_703 = lean_unsigned_to_nat(2u); -x_704 = lean_nat_dec_eq(x_702, x_703); -if (x_704 == 0) +if (lean_obj_tag(x_2) == 0) { -lean_object* x_705; -lean_dec(x_702); -x_705 = lean_box(0); -x_665 = x_705; -goto block_698; +lean_object* x_3; +x_3 = lean_box(0); +return x_3; } else { -lean_object* x_706; uint8_t x_707; -x_706 = lean_unsigned_to_nat(0u); -x_707 = lean_nat_dec_lt(x_706, x_702); -lean_dec(x_702); -if (x_707 == 0) +lean_object* x_4; lean_object* x_5; lean_object* x_6; size_t x_7; size_t x_8; uint8_t x_9; +x_4 = lean_ctor_get(x_2, 0); +x_5 = lean_ctor_get(x_2, 1); +x_6 = lean_ctor_get(x_2, 2); +x_7 = lean_ptr_addr(x_4); +x_8 = lean_ptr_addr(x_1); +x_9 = lean_usize_dec_eq(x_7, x_8); +if (x_9 == 0) { -lean_object* x_708; lean_object* x_709; -x_708 = l_Lean_instInhabitedExpr; -x_709 = l_outOfBounds___rarg(x_708); -x_634 = x_709; -goto block_664; +x_2 = x_6; +goto _start; } else { -lean_object* x_710; -x_710 = lean_array_fget(x_3, x_706); -x_634 = x_710; -goto block_664; +lean_object* x_11; +lean_inc(x_5); +x_11 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_11, 0, x_5); +return x_11; } } } -block_664: -{ -lean_object* x_635; -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_5); -lean_inc(x_634); -x_635 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_634, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_635) == 0) -{ -lean_object* x_636; lean_object* x_637; lean_object* x_638; lean_object* x_639; lean_object* x_640; lean_object* x_641; -x_636 = lean_ctor_get(x_635, 0); -lean_inc(x_636); -x_637 = lean_ctor_get(x_635, 1); -lean_inc(x_637); -lean_dec(x_635); -x_638 = lean_ctor_get(x_636, 0); -lean_inc(x_638); -x_639 = lean_ctor_get(x_636, 1); -lean_inc(x_639); -lean_dec(x_636); -x_640 = lean_ctor_get(x_639, 2); -lean_inc(x_640); -lean_inc(x_640); -x_641 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_640, x_638); -if (lean_obj_tag(x_641) == 0) -{ -size_t x_642; size_t x_643; uint8_t x_644; -x_642 = lean_ptr_addr(x_634); -lean_dec(x_634); -x_643 = lean_ptr_addr(x_638); -x_644 = lean_usize_dec_eq(x_642, x_643); -if (x_644 == 0) -{ -lean_object* x_645; lean_object* x_646; lean_object* x_647; lean_object* x_648; lean_object* x_649; lean_object* x_650; lean_object* x_651; lean_object* x_652; -x_645 = lean_unsigned_to_nat(0u); -lean_inc(x_638); -x_646 = lean_array_set(x_3, x_645, x_638); -x_647 = l_Lean_mkAppN(x_2, x_646); -lean_dec(x_646); -x_648 = lean_ctor_get(x_639, 0); -lean_inc(x_648); -x_649 = lean_ctor_get(x_639, 1); -lean_inc(x_649); -lean_dec(x_639); -lean_inc(x_647); -x_650 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_640, x_638, x_647); -x_651 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_651, 0, x_648); -lean_ctor_set(x_651, 1, x_649); -lean_ctor_set(x_651, 2, x_650); -x_652 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_647, x_5, x_651, x_7, x_8, x_9, x_10, x_637); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_652; } -else +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_canonImpl_visit___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_653; lean_object* x_654; lean_object* x_655; lean_object* x_656; lean_object* x_657; -lean_dec(x_3); -lean_dec(x_2); -x_653 = lean_ctor_get(x_639, 0); -lean_inc(x_653); -x_654 = lean_ctor_get(x_639, 1); -lean_inc(x_654); -lean_dec(x_639); -lean_inc(x_1); -x_655 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_640, x_638, x_1); -x_656 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_656, 0, x_653); -lean_ctor_set(x_656, 1, x_654); -lean_ctor_set(x_656, 2, x_655); -lean_inc(x_1); -x_657 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_656, x_7, x_8, x_9, x_10, x_637); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_657; -} +lean_object* x_10; uint8_t x_11; +x_10 = lean_st_ref_take(x_3, x_9); +x_11 = !lean_is_exclusive(x_10); +if (x_11 == 0) +{ +lean_object* x_12; uint8_t x_13; +x_12 = lean_ctor_get(x_10, 0); +x_13 = !lean_is_exclusive(x_12); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; size_t x_18; uint64_t x_19; uint64_t x_20; uint64_t x_21; uint64_t x_22; uint64_t x_23; uint64_t x_24; uint64_t x_25; uint64_t x_26; uint64_t x_27; size_t x_28; size_t x_29; size_t x_30; size_t x_31; size_t x_32; lean_object* x_33; uint8_t x_34; +x_14 = lean_ctor_get(x_10, 1); +x_15 = lean_ctor_get(x_12, 0); +x_16 = lean_ctor_get(x_12, 1); +x_17 = lean_array_get_size(x_16); +x_18 = lean_ptr_addr(x_1); +x_19 = lean_usize_to_uint64(x_18); +x_20 = 11; +x_21 = lean_uint64_mix_hash(x_19, x_20); +x_22 = 32; +x_23 = lean_uint64_shift_right(x_21, x_22); +x_24 = lean_uint64_xor(x_21, x_23); +x_25 = 16; +x_26 = lean_uint64_shift_right(x_24, x_25); +x_27 = lean_uint64_xor(x_24, x_26); +x_28 = lean_uint64_to_usize(x_27); +x_29 = lean_usize_of_nat(x_17); +lean_dec(x_17); +x_30 = 1; +x_31 = lean_usize_sub(x_29, x_30); +x_32 = lean_usize_land(x_28, x_31); +x_33 = lean_array_uget(x_16, x_32); +x_34 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__1(x_1, x_33); +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; lean_object* x_43; uint8_t x_44; +x_35 = lean_unsigned_to_nat(1u); +x_36 = lean_nat_add(x_15, x_35); +lean_dec(x_15); +lean_inc(x_2); +x_37 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_37, 0, x_1); +lean_ctor_set(x_37, 1, x_2); +lean_ctor_set(x_37, 2, x_33); +x_38 = lean_array_uset(x_16, x_32, x_37); +x_39 = lean_unsigned_to_nat(4u); +x_40 = lean_nat_mul(x_36, x_39); +x_41 = lean_unsigned_to_nat(3u); +x_42 = lean_nat_div(x_40, x_41); +lean_dec(x_40); +x_43 = lean_array_get_size(x_38); +x_44 = lean_nat_dec_le(x_42, x_43); +lean_dec(x_43); +lean_dec(x_42); +if (x_44 == 0) +{ +lean_object* x_45; lean_object* x_46; uint8_t x_47; +x_45 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__2(x_38); +lean_ctor_set(x_12, 1, x_45); +lean_ctor_set(x_12, 0, x_36); +x_46 = lean_st_ref_set(x_3, x_12, x_14); +x_47 = !lean_is_exclusive(x_46); +if (x_47 == 0) +{ +lean_object* x_48; +x_48 = lean_ctor_get(x_46, 0); +lean_dec(x_48); +lean_ctor_set(x_10, 1, x_4); +lean_ctor_set(x_10, 0, x_2); +lean_ctor_set(x_46, 0, x_10); +return x_46; } else { -lean_object* x_658; lean_object* x_659; -lean_dec(x_640); -lean_dec(x_638); -lean_dec(x_634); -lean_dec(x_3); -lean_dec(x_2); -x_658 = lean_ctor_get(x_641, 0); -lean_inc(x_658); -lean_dec(x_641); -x_659 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_658, x_5, x_639, x_7, x_8, x_9, x_10, x_637); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_659; +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_ctor_set(x_10, 1, x_4); +lean_ctor_set(x_10, 0, x_2); +x_50 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_50, 0, x_10); +lean_ctor_set(x_50, 1, x_49); +return x_50; } } else { -uint8_t x_660; -lean_dec(x_634); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_660 = !lean_is_exclusive(x_635); -if (x_660 == 0) +lean_object* x_51; uint8_t x_52; +lean_ctor_set(x_12, 1, x_38); +lean_ctor_set(x_12, 0, x_36); +x_51 = lean_st_ref_set(x_3, x_12, x_14); +x_52 = !lean_is_exclusive(x_51); +if (x_52 == 0) { -return x_635; +lean_object* x_53; +x_53 = lean_ctor_get(x_51, 0); +lean_dec(x_53); +lean_ctor_set(x_10, 1, x_4); +lean_ctor_set(x_10, 0, x_2); +lean_ctor_set(x_51, 0, x_10); +return x_51; } else { -lean_object* x_661; lean_object* x_662; lean_object* x_663; -x_661 = lean_ctor_get(x_635, 0); -x_662 = lean_ctor_get(x_635, 1); -lean_inc(x_662); -lean_inc(x_661); -lean_dec(x_635); -x_663 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_663, 0, x_661); -lean_ctor_set(x_663, 1, x_662); -return x_663; +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_ctor_set(x_10, 1, x_4); +lean_ctor_set(x_10, 0, x_2); +x_55 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_55, 0, x_10); +lean_ctor_set(x_55, 1, x_54); +return x_55; } } } -block_698: +else { -lean_object* x_666; -lean_dec(x_665); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_2); -x_666 = l_Lean_Meta_getFunInfo(x_2, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_666) == 0) -{ -lean_object* x_667; lean_object* x_668; lean_object* x_669; lean_object* x_670; lean_object* x_671; lean_object* x_672; lean_object* x_673; uint8_t x_674; lean_object* x_675; lean_object* x_676; lean_object* x_677; -x_667 = lean_ctor_get(x_666, 0); -lean_inc(x_667); -x_668 = lean_ctor_get(x_666, 1); -lean_inc(x_668); -lean_dec(x_666); -x_669 = lean_ctor_get(x_667, 0); -lean_inc(x_669); -lean_dec(x_667); -x_670 = lean_array_get_size(x_3); -x_671 = lean_unsigned_to_nat(0u); -x_672 = lean_unsigned_to_nat(1u); -x_673 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_673, 0, x_671); -lean_ctor_set(x_673, 1, x_670); -lean_ctor_set(x_673, 2, x_672); -x_674 = 0; -x_675 = lean_box(x_674); -x_676 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_676, 0, x_3); -lean_ctor_set(x_676, 1, x_675); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_5); +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_56 = lean_box(0); +x_57 = lean_array_uset(x_16, x_32, x_56); lean_inc(x_2); -x_677 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8(x_2, x_669, x_673, x_673, x_676, x_671, lean_box(0), lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10, x_668); -lean_dec(x_673); -lean_dec(x_669); -if (lean_obj_tag(x_677) == 0) -{ -lean_object* x_678; lean_object* x_679; lean_object* x_680; uint8_t x_681; -x_678 = lean_ctor_get(x_677, 0); -lean_inc(x_678); -x_679 = lean_ctor_get(x_678, 0); -lean_inc(x_679); -x_680 = lean_ctor_get(x_679, 1); -lean_inc(x_680); -x_681 = lean_unbox(x_680); -lean_dec(x_680); -if (x_681 == 0) +x_58 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__6(x_1, x_2, x_33); +x_59 = lean_array_uset(x_57, x_32, x_58); +lean_ctor_set(x_12, 1, x_59); +x_60 = lean_st_ref_set(x_3, x_12, x_14); +x_61 = !lean_is_exclusive(x_60); +if (x_61 == 0) { -lean_object* x_682; lean_object* x_683; lean_object* x_684; -lean_dec(x_679); -lean_dec(x_2); -x_682 = lean_ctor_get(x_677, 1); -lean_inc(x_682); -lean_dec(x_677); -x_683 = lean_ctor_get(x_678, 1); -lean_inc(x_683); -lean_dec(x_678); -lean_inc(x_1); -x_684 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_683, x_7, x_8, x_9, x_10, x_682); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_684; +lean_object* x_62; +x_62 = lean_ctor_get(x_60, 0); +lean_dec(x_62); +lean_ctor_set(x_10, 1, x_4); +lean_ctor_set(x_10, 0, x_2); +lean_ctor_set(x_60, 0, x_10); +return x_60; +} +else +{ +lean_object* x_63; lean_object* x_64; +x_63 = lean_ctor_get(x_60, 1); +lean_inc(x_63); +lean_dec(x_60); +lean_ctor_set(x_10, 1, x_4); +lean_ctor_set(x_10, 0, x_2); +x_64 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_64, 0, x_10); +lean_ctor_set(x_64, 1, x_63); +return x_64; +} +} } else { -lean_object* x_685; lean_object* x_686; lean_object* x_687; lean_object* x_688; lean_object* x_689; -x_685 = lean_ctor_get(x_677, 1); -lean_inc(x_685); -lean_dec(x_677); -x_686 = lean_ctor_get(x_678, 1); -lean_inc(x_686); -lean_dec(x_678); -x_687 = lean_ctor_get(x_679, 0); -lean_inc(x_687); -lean_dec(x_679); -x_688 = l_Lean_mkAppN(x_2, x_687); -lean_dec(x_687); -x_689 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_688, x_5, x_686, x_7, x_8, x_9, x_10, x_685); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_689; +lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; size_t x_69; uint64_t x_70; uint64_t x_71; uint64_t x_72; uint64_t x_73; uint64_t x_74; uint64_t x_75; uint64_t x_76; uint64_t x_77; uint64_t x_78; size_t x_79; size_t x_80; size_t x_81; size_t x_82; size_t x_83; lean_object* x_84; uint8_t x_85; +x_65 = lean_ctor_get(x_10, 1); +x_66 = lean_ctor_get(x_12, 0); +x_67 = lean_ctor_get(x_12, 1); +lean_inc(x_67); +lean_inc(x_66); +lean_dec(x_12); +x_68 = lean_array_get_size(x_67); +x_69 = lean_ptr_addr(x_1); +x_70 = lean_usize_to_uint64(x_69); +x_71 = 11; +x_72 = lean_uint64_mix_hash(x_70, x_71); +x_73 = 32; +x_74 = lean_uint64_shift_right(x_72, x_73); +x_75 = lean_uint64_xor(x_72, x_74); +x_76 = 16; +x_77 = lean_uint64_shift_right(x_75, x_76); +x_78 = lean_uint64_xor(x_75, x_77); +x_79 = lean_uint64_to_usize(x_78); +x_80 = lean_usize_of_nat(x_68); +lean_dec(x_68); +x_81 = 1; +x_82 = lean_usize_sub(x_80, x_81); +x_83 = lean_usize_land(x_79, x_82); +x_84 = lean_array_uget(x_67, x_83); +x_85 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__1(x_1, x_84); +if (x_85 == 0) +{ +lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; uint8_t x_95; +x_86 = lean_unsigned_to_nat(1u); +x_87 = lean_nat_add(x_66, x_86); +lean_dec(x_66); +lean_inc(x_2); +x_88 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_88, 0, x_1); +lean_ctor_set(x_88, 1, x_2); +lean_ctor_set(x_88, 2, x_84); +x_89 = lean_array_uset(x_67, x_83, x_88); +x_90 = lean_unsigned_to_nat(4u); +x_91 = lean_nat_mul(x_87, x_90); +x_92 = lean_unsigned_to_nat(3u); +x_93 = lean_nat_div(x_91, x_92); +lean_dec(x_91); +x_94 = lean_array_get_size(x_89); +x_95 = lean_nat_dec_le(x_93, x_94); +lean_dec(x_94); +lean_dec(x_93); +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; +x_96 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__2(x_89); +x_97 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_97, 0, x_87); +lean_ctor_set(x_97, 1, x_96); +x_98 = lean_st_ref_set(x_3, x_97, x_65); +x_99 = lean_ctor_get(x_98, 1); +lean_inc(x_99); +if (lean_is_exclusive(x_98)) { + lean_ctor_release(x_98, 0); + lean_ctor_release(x_98, 1); + x_100 = x_98; +} else { + lean_dec_ref(x_98); + x_100 = lean_box(0); } +lean_ctor_set(x_10, 1, x_4); +lean_ctor_set(x_10, 0, x_2); +if (lean_is_scalar(x_100)) { + x_101 = lean_alloc_ctor(0, 2, 0); +} else { + x_101 = x_100; } -else -{ -uint8_t x_690; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -lean_dec(x_2); -lean_dec(x_1); -x_690 = !lean_is_exclusive(x_677); -if (x_690 == 0) -{ -return x_677; +lean_ctor_set(x_101, 0, x_10); +lean_ctor_set(x_101, 1, x_99); +return x_101; } else { -lean_object* x_691; lean_object* x_692; lean_object* x_693; -x_691 = lean_ctor_get(x_677, 0); -x_692 = lean_ctor_get(x_677, 1); -lean_inc(x_692); -lean_inc(x_691); -lean_dec(x_677); -x_693 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_693, 0, x_691); -lean_ctor_set(x_693, 1, x_692); -return x_693; +lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; +x_102 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_102, 0, x_87); +lean_ctor_set(x_102, 1, x_89); +x_103 = lean_st_ref_set(x_3, x_102, x_65); +x_104 = lean_ctor_get(x_103, 1); +lean_inc(x_104); +if (lean_is_exclusive(x_103)) { + lean_ctor_release(x_103, 0); + lean_ctor_release(x_103, 1); + x_105 = x_103; +} else { + lean_dec_ref(x_103); + x_105 = lean_box(0); } +lean_ctor_set(x_10, 1, x_4); +lean_ctor_set(x_10, 0, x_2); +if (lean_is_scalar(x_105)) { + x_106 = lean_alloc_ctor(0, 2, 0); +} else { + x_106 = x_105; } +lean_ctor_set(x_106, 0, x_10); +lean_ctor_set(x_106, 1, x_104); +return x_106; } -else -{ -uint8_t x_694; -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_694 = !lean_is_exclusive(x_666); -if (x_694 == 0) -{ -return x_666; } else { -lean_object* x_695; lean_object* x_696; lean_object* x_697; -x_695 = lean_ctor_get(x_666, 0); -x_696 = lean_ctor_get(x_666, 1); -lean_inc(x_696); -lean_inc(x_695); -lean_dec(x_666); -x_697 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_697, 0, x_695); -lean_ctor_set(x_697, 1, x_696); -return x_697; +lean_object* x_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_107 = lean_box(0); +x_108 = lean_array_uset(x_67, x_83, x_107); +lean_inc(x_2); +x_109 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__6(x_1, x_2, x_84); +x_110 = lean_array_uset(x_108, x_83, x_109); +x_111 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_111, 0, x_66); +lean_ctor_set(x_111, 1, x_110); +x_112 = lean_st_ref_set(x_3, x_111, x_65); +x_113 = lean_ctor_get(x_112, 1); +lean_inc(x_113); +if (lean_is_exclusive(x_112)) { + lean_ctor_release(x_112, 0); + lean_ctor_release(x_112, 1); + x_114 = x_112; +} else { + lean_dec_ref(x_112); + x_114 = lean_box(0); } +lean_ctor_set(x_10, 1, x_4); +lean_ctor_set(x_10, 0, x_2); +if (lean_is_scalar(x_114)) { + x_115 = lean_alloc_ctor(0, 2, 0); +} else { + x_115 = x_114; } +lean_ctor_set(x_115, 0, x_10); +lean_ctor_set(x_115, 1, x_113); +return x_115; } } -case 10: -{ -lean_object* x_711; lean_object* x_742; lean_object* x_776; uint8_t x_777; -lean_dec(x_4); -x_776 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__4; -x_777 = l_Lean_Expr_isConstOf(x_2, x_776); -if (x_777 == 0) -{ -lean_object* x_778; -x_778 = lean_box(0); -x_742 = x_778; -goto block_775; } else { -lean_object* x_779; lean_object* x_780; uint8_t x_781; -x_779 = lean_array_get_size(x_3); -x_780 = lean_unsigned_to_nat(2u); -x_781 = lean_nat_dec_eq(x_779, x_780); -if (x_781 == 0) -{ -lean_object* x_782; -lean_dec(x_779); -x_782 = lean_box(0); -x_742 = x_782; -goto block_775; +lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; size_t x_122; uint64_t x_123; uint64_t x_124; uint64_t x_125; uint64_t x_126; uint64_t x_127; uint64_t x_128; uint64_t x_129; uint64_t x_130; uint64_t x_131; size_t x_132; size_t x_133; size_t x_134; size_t x_135; size_t x_136; lean_object* x_137; uint8_t x_138; +x_116 = lean_ctor_get(x_10, 0); +x_117 = lean_ctor_get(x_10, 1); +lean_inc(x_117); +lean_inc(x_116); +lean_dec(x_10); +x_118 = lean_ctor_get(x_116, 0); +lean_inc(x_118); +x_119 = lean_ctor_get(x_116, 1); +lean_inc(x_119); +if (lean_is_exclusive(x_116)) { + lean_ctor_release(x_116, 0); + lean_ctor_release(x_116, 1); + x_120 = x_116; +} else { + lean_dec_ref(x_116); + x_120 = lean_box(0); } -else -{ -lean_object* x_783; uint8_t x_784; -x_783 = lean_unsigned_to_nat(0u); -x_784 = lean_nat_dec_lt(x_783, x_779); -lean_dec(x_779); -if (x_784 == 0) +x_121 = lean_array_get_size(x_119); +x_122 = lean_ptr_addr(x_1); +x_123 = lean_usize_to_uint64(x_122); +x_124 = 11; +x_125 = lean_uint64_mix_hash(x_123, x_124); +x_126 = 32; +x_127 = lean_uint64_shift_right(x_125, x_126); +x_128 = lean_uint64_xor(x_125, x_127); +x_129 = 16; +x_130 = lean_uint64_shift_right(x_128, x_129); +x_131 = lean_uint64_xor(x_128, x_130); +x_132 = lean_uint64_to_usize(x_131); +x_133 = lean_usize_of_nat(x_121); +lean_dec(x_121); +x_134 = 1; +x_135 = lean_usize_sub(x_133, x_134); +x_136 = lean_usize_land(x_132, x_135); +x_137 = lean_array_uget(x_119, x_136); +x_138 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__1(x_1, x_137); +if (x_138 == 0) { -lean_object* x_785; lean_object* x_786; -x_785 = l_Lean_instInhabitedExpr; -x_786 = l_outOfBounds___rarg(x_785); -x_711 = x_786; -goto block_741; -} -else +lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; uint8_t x_148; +x_139 = lean_unsigned_to_nat(1u); +x_140 = lean_nat_add(x_118, x_139); +lean_dec(x_118); +lean_inc(x_2); +x_141 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_141, 0, x_1); +lean_ctor_set(x_141, 1, x_2); +lean_ctor_set(x_141, 2, x_137); +x_142 = lean_array_uset(x_119, x_136, x_141); +x_143 = lean_unsigned_to_nat(4u); +x_144 = lean_nat_mul(x_140, x_143); +x_145 = lean_unsigned_to_nat(3u); +x_146 = lean_nat_div(x_144, x_145); +lean_dec(x_144); +x_147 = lean_array_get_size(x_142); +x_148 = lean_nat_dec_le(x_146, x_147); +lean_dec(x_147); +lean_dec(x_146); +if (x_148 == 0) { -lean_object* x_787; -x_787 = lean_array_fget(x_3, x_783); -x_711 = x_787; -goto block_741; -} -} +lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; +x_149 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__2(x_142); +if (lean_is_scalar(x_120)) { + x_150 = lean_alloc_ctor(0, 2, 0); +} else { + x_150 = x_120; } -block_741: -{ -lean_object* x_712; -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_5); -lean_inc(x_711); -x_712 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_711, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_712) == 0) -{ -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_713 = lean_ctor_get(x_712, 0); -lean_inc(x_713); -x_714 = lean_ctor_get(x_712, 1); -lean_inc(x_714); -lean_dec(x_712); -x_715 = lean_ctor_get(x_713, 0); -lean_inc(x_715); -x_716 = lean_ctor_get(x_713, 1); -lean_inc(x_716); -lean_dec(x_713); -x_717 = lean_ctor_get(x_716, 2); -lean_inc(x_717); -lean_inc(x_717); -x_718 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_717, x_715); -if (lean_obj_tag(x_718) == 0) -{ -size_t x_719; size_t x_720; uint8_t x_721; -x_719 = lean_ptr_addr(x_711); -lean_dec(x_711); -x_720 = lean_ptr_addr(x_715); -x_721 = lean_usize_dec_eq(x_719, x_720); -if (x_721 == 0) -{ -lean_object* x_722; lean_object* x_723; lean_object* x_724; lean_object* x_725; lean_object* x_726; lean_object* x_727; lean_object* x_728; lean_object* x_729; -x_722 = lean_unsigned_to_nat(0u); -lean_inc(x_715); -x_723 = lean_array_set(x_3, x_722, x_715); -x_724 = l_Lean_mkAppN(x_2, x_723); -lean_dec(x_723); -x_725 = lean_ctor_get(x_716, 0); -lean_inc(x_725); -x_726 = lean_ctor_get(x_716, 1); -lean_inc(x_726); -lean_dec(x_716); -lean_inc(x_724); -x_727 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_717, x_715, x_724); -x_728 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_728, 0, x_725); -lean_ctor_set(x_728, 1, x_726); -lean_ctor_set(x_728, 2, x_727); -x_729 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_724, x_5, x_728, x_7, x_8, x_9, x_10, x_714); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_729; +lean_ctor_set(x_150, 0, x_140); +lean_ctor_set(x_150, 1, x_149); +x_151 = lean_st_ref_set(x_3, x_150, x_117); +x_152 = lean_ctor_get(x_151, 1); +lean_inc(x_152); +if (lean_is_exclusive(x_151)) { + lean_ctor_release(x_151, 0); + lean_ctor_release(x_151, 1); + x_153 = x_151; +} else { + lean_dec_ref(x_151); + x_153 = lean_box(0); } -else -{ -lean_object* x_730; lean_object* x_731; lean_object* x_732; lean_object* x_733; lean_object* x_734; -lean_dec(x_3); -lean_dec(x_2); -x_730 = lean_ctor_get(x_716, 0); -lean_inc(x_730); -x_731 = lean_ctor_get(x_716, 1); -lean_inc(x_731); -lean_dec(x_716); -lean_inc(x_1); -x_732 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_717, x_715, x_1); -x_733 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_733, 0, x_730); -lean_ctor_set(x_733, 1, x_731); -lean_ctor_set(x_733, 2, x_732); -lean_inc(x_1); -x_734 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_733, x_7, x_8, x_9, x_10, x_714); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_734; +x_154 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_154, 0, x_2); +lean_ctor_set(x_154, 1, x_4); +if (lean_is_scalar(x_153)) { + x_155 = lean_alloc_ctor(0, 2, 0); +} else { + x_155 = x_153; } +lean_ctor_set(x_155, 0, x_154); +lean_ctor_set(x_155, 1, x_152); +return x_155; } else { -lean_object* x_735; lean_object* x_736; -lean_dec(x_717); -lean_dec(x_715); -lean_dec(x_711); -lean_dec(x_3); -lean_dec(x_2); -x_735 = lean_ctor_get(x_718, 0); -lean_inc(x_735); -lean_dec(x_718); -x_736 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_735, x_5, x_716, x_7, x_8, x_9, x_10, x_714); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_736; +lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; +if (lean_is_scalar(x_120)) { + x_156 = lean_alloc_ctor(0, 2, 0); +} else { + x_156 = x_120; } +lean_ctor_set(x_156, 0, x_140); +lean_ctor_set(x_156, 1, x_142); +x_157 = lean_st_ref_set(x_3, x_156, x_117); +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_157); + x_159 = lean_box(0); +} +x_160 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_160, 0, x_2); +lean_ctor_set(x_160, 1, x_4); +if (lean_is_scalar(x_159)) { + x_161 = lean_alloc_ctor(0, 2, 0); +} else { + x_161 = x_159; +} +lean_ctor_set(x_161, 0, x_160); +lean_ctor_set(x_161, 1, x_158); +return x_161; } -else -{ -uint8_t x_737; -lean_dec(x_711); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_737 = !lean_is_exclusive(x_712); -if (x_737 == 0) -{ -return x_712; } else { -lean_object* x_738; lean_object* x_739; lean_object* x_740; -x_738 = lean_ctor_get(x_712, 0); -x_739 = lean_ctor_get(x_712, 1); -lean_inc(x_739); -lean_inc(x_738); -lean_dec(x_712); -x_740 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_740, 0, x_738); -lean_ctor_set(x_740, 1, x_739); -return x_740; +lean_object* x_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_162 = lean_box(0); +x_163 = lean_array_uset(x_119, x_136, x_162); +lean_inc(x_2); +x_164 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__6(x_1, x_2, x_137); +x_165 = lean_array_uset(x_163, x_136, x_164); +if (lean_is_scalar(x_120)) { + x_166 = lean_alloc_ctor(0, 2, 0); +} else { + x_166 = x_120; +} +lean_ctor_set(x_166, 0, x_118); +lean_ctor_set(x_166, 1, x_165); +x_167 = lean_st_ref_set(x_3, x_166, x_117); +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); +} +x_170 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_170, 0, x_2); +lean_ctor_set(x_170, 1, x_4); +if (lean_is_scalar(x_169)) { + x_171 = lean_alloc_ctor(0, 2, 0); +} else { + x_171 = x_169; +} +lean_ctor_set(x_171, 0, x_170); +lean_ctor_set(x_171, 1, x_168); +return x_171; } } } -block_775: +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: { -lean_object* x_743; -lean_dec(x_742); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_2); -x_743 = l_Lean_Meta_getFunInfo(x_2, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_743) == 0) -{ -lean_object* x_744; lean_object* x_745; lean_object* x_746; lean_object* x_747; lean_object* x_748; lean_object* x_749; lean_object* x_750; uint8_t x_751; lean_object* x_752; lean_object* x_753; lean_object* x_754; -x_744 = lean_ctor_get(x_743, 0); -lean_inc(x_744); -x_745 = lean_ctor_get(x_743, 1); -lean_inc(x_745); -lean_dec(x_743); -x_746 = lean_ctor_get(x_744, 0); -lean_inc(x_746); -lean_dec(x_744); -x_747 = lean_array_get_size(x_3); -x_748 = lean_unsigned_to_nat(0u); -x_749 = lean_unsigned_to_nat(1u); -x_750 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_750, 0, x_748); -lean_ctor_set(x_750, 1, x_747); -lean_ctor_set(x_750, 2, x_749); -x_751 = 0; -x_752 = lean_box(x_751); -x_753 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_753, 0, x_3); -lean_ctor_set(x_753, 1, x_752); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_5); +lean_object* x_15; size_t x_16; size_t x_17; uint8_t x_18; +lean_inc(x_3); lean_inc(x_2); -x_754 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8(x_2, x_746, x_750, x_750, x_753, x_748, lean_box(0), lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10, x_745); -lean_dec(x_750); -lean_dec(x_746); -if (lean_obj_tag(x_754) == 0) -{ -lean_object* x_755; lean_object* x_756; lean_object* x_757; uint8_t x_758; -x_755 = lean_ctor_get(x_754, 0); -lean_inc(x_755); -x_756 = lean_ctor_get(x_755, 0); -lean_inc(x_756); -x_757 = lean_ctor_get(x_756, 1); -lean_inc(x_757); -x_758 = lean_unbox(x_757); -lean_dec(x_757); -if (x_758 == 0) -{ -lean_object* x_759; lean_object* x_760; lean_object* x_761; -lean_dec(x_756); -lean_dec(x_2); -x_759 = lean_ctor_get(x_754, 1); -lean_inc(x_759); -lean_dec(x_754); -x_760 = lean_ctor_get(x_755, 1); -lean_inc(x_760); -lean_dec(x_755); lean_inc(x_1); -x_761 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_760, x_7, x_8, x_9, x_10, x_759); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_761; -} -else +x_15 = l_Lean_Expr_forallE___override(x_1, x_2, x_3, x_4); +x_16 = lean_ptr_addr(x_2); +lean_dec(x_2); +x_17 = lean_ptr_addr(x_5); +x_18 = lean_usize_dec_eq(x_16, x_17); +if (x_18 == 0) { -lean_object* x_762; lean_object* x_763; lean_object* x_764; lean_object* x_765; lean_object* x_766; -x_762 = lean_ctor_get(x_754, 1); -lean_inc(x_762); -lean_dec(x_754); -x_763 = lean_ctor_get(x_755, 1); -lean_inc(x_763); -lean_dec(x_755); -x_764 = lean_ctor_get(x_756, 0); -lean_inc(x_764); -lean_dec(x_756); -x_765 = l_Lean_mkAppN(x_2, x_764); -lean_dec(x_764); -x_766 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_765, x_5, x_763, x_7, x_8, x_9, x_10, x_762); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_766; -} +lean_object* x_19; lean_object* x_20; +lean_dec(x_15); +lean_dec(x_3); +x_19 = l_Lean_Expr_forallE___override(x_1, x_5, x_7, x_4); +x_20 = lean_apply_8(x_6, x_19, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +return x_20; } else { -uint8_t x_767; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -lean_dec(x_2); -lean_dec(x_1); -x_767 = !lean_is_exclusive(x_754); -if (x_767 == 0) +size_t x_21; size_t x_22; uint8_t x_23; +x_21 = lean_ptr_addr(x_3); +lean_dec(x_3); +x_22 = lean_ptr_addr(x_7); +x_23 = lean_usize_dec_eq(x_21, x_22); +if (x_23 == 0) { -return x_754; +lean_object* x_24; lean_object* x_25; +lean_dec(x_15); +x_24 = l_Lean_Expr_forallE___override(x_1, x_5, x_7, x_4); +x_25 = lean_apply_8(x_6, x_24, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +return x_25; } else { -lean_object* x_768; lean_object* x_769; lean_object* x_770; -x_768 = lean_ctor_get(x_754, 0); -x_769 = lean_ctor_get(x_754, 1); -lean_inc(x_769); -lean_inc(x_768); -lean_dec(x_754); -x_770 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_770, 0, x_768); -lean_ctor_set(x_770, 1, x_769); -return x_770; -} -} +uint8_t x_26; +x_26 = l___private_Lean_Expr_0__Lean_beqBinderInfo____x40_Lean_Expr___hyg_406_(x_4, x_4); +if (x_26 == 0) +{ +lean_object* x_27; lean_object* x_28; +lean_dec(x_15); +x_27 = l_Lean_Expr_forallE___override(x_1, x_5, x_7, x_4); +x_28 = lean_apply_8(x_6, x_27, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +return x_28; } else { -uint8_t x_771; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); +lean_object* x_29; 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_771 = !lean_is_exclusive(x_743); -if (x_771 == 0) -{ -return x_743; +x_29 = lean_apply_8(x_6, x_15, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +return x_29; } -else -{ -lean_object* x_772; lean_object* x_773; lean_object* x_774; -x_772 = lean_ctor_get(x_743, 0); -x_773 = lean_ctor_get(x_743, 1); -lean_inc(x_773); -lean_inc(x_772); -lean_dec(x_743); -x_774 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_774, 0, x_772); -lean_ctor_set(x_774, 1, x_773); -return x_774; } } } } -default: -{ -lean_object* x_788; lean_object* x_819; lean_object* x_853; uint8_t x_854; -lean_dec(x_4); -x_853 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__4; -x_854 = l_Lean_Expr_isConstOf(x_2, x_853); -if (x_854 == 0) +static lean_object* _init_l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__3___closed__1() { +_start: { -lean_object* x_855; -x_855 = lean_box(0); -x_819 = x_855; -goto block_852; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.Meta.Tactic.Grind.Canon", 28, 28); +return x_1; } -else -{ -lean_object* x_856; lean_object* x_857; uint8_t x_858; -x_856 = lean_array_get_size(x_3); -x_857 = lean_unsigned_to_nat(2u); -x_858 = lean_nat_dec_eq(x_856, x_857); -if (x_858 == 0) -{ -lean_object* x_859; -lean_dec(x_856); -x_859 = lean_box(0); -x_819 = x_859; -goto block_852; } -else +static lean_object* _init_l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__3___closed__2() { +_start: { -lean_object* x_860; uint8_t x_861; -x_860 = lean_unsigned_to_nat(0u); -x_861 = lean_nat_dec_lt(x_860, x_856); -lean_dec(x_856); -if (x_861 == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.Meta.Grind.Canon.canonImpl.visit", 37, 37); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__3___closed__3() { +_start: { -lean_object* x_862; lean_object* x_863; -x_862 = l_Lean_instInhabitedExpr; -x_863 = l_outOfBounds___rarg(x_862); -x_788 = x_863; -goto block_818; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("unreachable code has been reached", 33, 33); +return x_1; } -else +} +static lean_object* _init_l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__3___closed__4() { +_start: { -lean_object* x_864; -x_864 = lean_array_fget(x_3, x_860); -x_788 = x_864; -goto block_818; +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_Grind_Canon_canonImpl_visit___lambda__3___closed__1; +x_2 = l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__3___closed__2; +x_3 = lean_unsigned_to_nat(190u); +x_4 = lean_unsigned_to_nat(13u); +x_5 = l_Lean_Meta_Grind_Canon_canonImpl_visit___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; +} } +static lean_object* _init_l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__3___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_levelZero; +x_2 = l_Lean_Expr_sort___override(x_1); +return x_2; } } -block_818: +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +switch (lean_obj_tag(x_1)) { +case 5: { -lean_object* x_789; +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_14 = lean_unsigned_to_nat(0u); +x_15 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_14); +x_16 = l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__3___closed__5; +lean_inc(x_15); +x_17 = lean_mk_array(x_15, x_16); +x_18 = lean_unsigned_to_nat(1u); +x_19 = lean_nat_sub(x_15, x_18); +lean_dec(x_15); +lean_inc(x_12); +lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); -lean_inc(x_8); lean_inc(x_7); -lean_inc(x_5); -lean_inc(x_788); -x_789 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_788, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_789) == 0) -{ -lean_object* x_790; lean_object* x_791; lean_object* x_792; lean_object* x_793; lean_object* x_794; lean_object* x_795; -x_790 = lean_ctor_get(x_789, 0); -lean_inc(x_790); -x_791 = lean_ctor_get(x_789, 1); -lean_inc(x_791); -lean_dec(x_789); -x_792 = lean_ctor_get(x_790, 0); -lean_inc(x_792); -x_793 = lean_ctor_get(x_790, 1); -lean_inc(x_793); -lean_dec(x_790); -x_794 = lean_ctor_get(x_793, 2); -lean_inc(x_794); -lean_inc(x_794); -x_795 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_794, x_792); -if (lean_obj_tag(x_795) == 0) -{ -size_t x_796; size_t x_797; uint8_t x_798; -x_796 = lean_ptr_addr(x_788); -lean_dec(x_788); -x_797 = lean_ptr_addr(x_792); -x_798 = lean_usize_dec_eq(x_796, x_797); -if (x_798 == 0) -{ -lean_object* x_799; lean_object* x_800; lean_object* x_801; lean_object* x_802; lean_object* x_803; lean_object* x_804; lean_object* x_805; lean_object* x_806; -x_799 = lean_unsigned_to_nat(0u); -lean_inc(x_792); -x_800 = lean_array_set(x_3, x_799, x_792); -x_801 = l_Lean_mkAppN(x_2, x_800); -lean_dec(x_800); -x_802 = lean_ctor_get(x_793, 0); -lean_inc(x_802); -x_803 = lean_ctor_get(x_793, 1); -lean_inc(x_803); -lean_dec(x_793); -lean_inc(x_801); -x_804 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_794, x_792, x_801); -x_805 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_805, 0, x_802); -lean_ctor_set(x_805, 1, x_803); -lean_ctor_set(x_805, 2, x_804); -x_806 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_801, x_5, x_805, x_7, x_8, x_9, x_10, x_791); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_806; -} -else -{ -lean_object* x_807; lean_object* x_808; lean_object* x_809; lean_object* x_810; lean_object* x_811; -lean_dec(x_3); -lean_dec(x_2); -x_807 = lean_ctor_get(x_793, 0); -lean_inc(x_807); -x_808 = lean_ctor_get(x_793, 1); -lean_inc(x_808); -lean_dec(x_793); -lean_inc(x_1); -x_809 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_794, x_792, x_1); -x_810 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_810, 0, x_807); -lean_ctor_set(x_810, 1, x_808); -lean_ctor_set(x_810, 2, x_809); -lean_inc(x_1); -x_811 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_810, x_7, x_8, x_9, x_10, x_791); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_811; -} -} -else +lean_inc_n(x_1, 2); +x_20 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__11(x_1, x_2, x_3, x_4, x_5, x_1, x_17, x_19, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_20) == 0) { -lean_object* x_812; lean_object* x_813; -lean_dec(x_794); -lean_dec(x_792); -lean_dec(x_788); -lean_dec(x_3); -lean_dec(x_2); -x_812 = lean_ctor_get(x_795, 0); -lean_inc(x_812); -lean_dec(x_795); -x_813 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_812, x_5, x_793, x_7, x_8, x_9, x_10, x_791); +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_21 = lean_ctor_get(x_20, 0); +lean_inc(x_21); +x_22 = lean_ctor_get(x_20, 1); +lean_inc(x_22); +lean_dec(x_20); +x_23 = lean_ctor_get(x_21, 0); +lean_inc(x_23); +x_24 = lean_ctor_get(x_21, 1); +lean_inc(x_24); +lean_dec(x_21); +x_25 = l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__1(x_1, x_23, x_7, x_24, x_9, x_10, x_11, x_12, x_22); +lean_dec(x_12); +lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); -lean_dec(x_8); lean_dec(x_7); -lean_dec(x_5); -return x_813; -} +return x_25; } else { -uint8_t x_814; -lean_dec(x_788); +uint8_t x_26; +lean_dec(x_12); +lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); -lean_dec(x_8); lean_dec(x_7); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); lean_dec(x_1); -x_814 = !lean_is_exclusive(x_789); -if (x_814 == 0) +x_26 = !lean_is_exclusive(x_20); +if (x_26 == 0) { -return x_789; +return x_20; } else { -lean_object* x_815; lean_object* x_816; lean_object* x_817; -x_815 = lean_ctor_get(x_789, 0); -x_816 = lean_ctor_get(x_789, 1); -lean_inc(x_816); -lean_inc(x_815); -lean_dec(x_789); -x_817 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_817, 0, x_815); -lean_ctor_set(x_817, 1, x_816); -return x_817; +lean_object* x_27; lean_object* x_28; lean_object* x_29; +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_alloc_ctor(1, 2, 0); +lean_ctor_set(x_29, 0, x_27); +lean_ctor_set(x_29, 1, x_28); +return x_29; } } } -block_852: +case 7: { -lean_object* x_820; -lean_dec(x_819); +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_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_30 = lean_ctor_get(x_1, 0); +lean_inc(x_30); +x_31 = lean_ctor_get(x_1, 1); +lean_inc(x_31); +x_32 = lean_ctor_get(x_1, 2); +lean_inc(x_32); +x_33 = lean_ctor_get_uint8(x_1, sizeof(void*)*3 + 8); +x_34 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__1___boxed), 9, 1); +lean_closure_set(x_34, 0, 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_2); -x_820 = l_Lean_Meta_getFunInfo(x_2, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_820) == 0) -{ -lean_object* x_821; lean_object* x_822; lean_object* x_823; lean_object* x_824; lean_object* x_825; lean_object* x_826; lean_object* x_827; uint8_t x_828; lean_object* x_829; lean_object* x_830; lean_object* x_831; -x_821 = lean_ctor_get(x_820, 0); -lean_inc(x_821); -x_822 = lean_ctor_get(x_820, 1); -lean_inc(x_822); -lean_dec(x_820); -x_823 = lean_ctor_get(x_821, 0); -lean_inc(x_823); -lean_dec(x_821); -x_824 = lean_array_get_size(x_3); -x_825 = lean_unsigned_to_nat(0u); -x_826 = lean_unsigned_to_nat(1u); -x_827 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_827, 0, x_825); -lean_ctor_set(x_827, 1, x_824); -lean_ctor_set(x_827, 2, x_826); -x_828 = 0; -x_829 = lean_box(x_828); -x_830 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_830, 0, x_3); -lean_ctor_set(x_830, 1, x_829); +lean_inc(x_31); +x_35 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_31, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_35) == 0) +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; uint8_t x_40; +x_36 = lean_ctor_get(x_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, 0); +lean_inc(x_38); +x_39 = lean_ctor_get(x_36, 1); +lean_inc(x_39); +lean_dec(x_36); +x_40 = l_Lean_Expr_hasLooseBVars(x_32); +if (x_40 == 0) +{ +lean_object* x_41; +lean_inc(x_12); +lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); -lean_inc(x_8); lean_inc(x_7); -lean_inc(x_5); -lean_inc(x_2); -x_831 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8(x_2, x_823, x_827, x_827, x_830, x_825, lean_box(0), lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10, x_822); -lean_dec(x_827); -lean_dec(x_823); -if (lean_obj_tag(x_831) == 0) -{ -lean_object* x_832; lean_object* x_833; lean_object* x_834; uint8_t x_835; -x_832 = lean_ctor_get(x_831, 0); -lean_inc(x_832); -x_833 = lean_ctor_get(x_832, 0); -lean_inc(x_833); -x_834 = lean_ctor_get(x_833, 1); -lean_inc(x_834); -x_835 = lean_unbox(x_834); -lean_dec(x_834); -if (x_835 == 0) +lean_inc(x_32); +x_41 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_32, x_7, x_39, x_9, x_10, x_11, x_12, x_37); +if (lean_obj_tag(x_41) == 0) { -lean_object* x_836; lean_object* x_837; lean_object* x_838; -lean_dec(x_833); -lean_dec(x_2); -x_836 = lean_ctor_get(x_831, 1); -lean_inc(x_836); -lean_dec(x_831); -x_837 = lean_ctor_get(x_832, 1); -lean_inc(x_837); -lean_dec(x_832); -lean_inc(x_1); -x_838 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_837, x_7, x_8, x_9, x_10, x_836); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_838; +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_42 = lean_ctor_get(x_41, 0); +lean_inc(x_42); +x_43 = lean_ctor_get(x_41, 1); +lean_inc(x_43); +lean_dec(x_41); +x_44 = lean_ctor_get(x_42, 0); +lean_inc(x_44); +x_45 = lean_ctor_get(x_42, 1); +lean_inc(x_45); +lean_dec(x_42); +x_46 = l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__2(x_30, x_31, x_32, x_33, x_38, x_34, x_44, x_7, x_45, x_9, x_10, x_11, x_12, x_43); +return x_46; } else { -lean_object* x_839; lean_object* x_840; lean_object* x_841; lean_object* x_842; lean_object* x_843; -x_839 = lean_ctor_get(x_831, 1); -lean_inc(x_839); -lean_dec(x_831); -x_840 = lean_ctor_get(x_832, 1); -lean_inc(x_840); -lean_dec(x_832); -x_841 = lean_ctor_get(x_833, 0); -lean_inc(x_841); -lean_dec(x_833); -x_842 = l_Lean_mkAppN(x_2, x_841); -lean_dec(x_841); -x_843 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_842, x_5, x_840, x_7, x_8, x_9, x_10, x_839); +uint8_t x_47; +lean_dec(x_38); +lean_dec(x_34); +lean_dec(x_32); +lean_dec(x_31); +lean_dec(x_30); +lean_dec(x_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); -return x_843; -} +x_47 = !lean_is_exclusive(x_41); +if (x_47 == 0) +{ +return x_41; } else { -uint8_t x_844; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -lean_dec(x_2); -lean_dec(x_1); -x_844 = !lean_is_exclusive(x_831); -if (x_844 == 0) -{ -return x_831; +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_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 { -lean_object* x_845; lean_object* x_846; lean_object* x_847; -x_845 = lean_ctor_get(x_831, 0); -x_846 = lean_ctor_get(x_831, 1); -lean_inc(x_846); -lean_inc(x_845); -lean_dec(x_831); -x_847 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_847, 0, x_845); -lean_ctor_set(x_847, 1, x_846); -return x_847; -} +lean_object* x_51; +lean_inc(x_32); +x_51 = l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__2(x_30, x_31, x_32, x_33, x_38, x_34, x_32, x_7, x_39, x_9, x_10, x_11, x_12, x_37); +return x_51; } } else { -uint8_t x_848; +uint8_t x_52; +lean_dec(x_34); +lean_dec(x_32); +lean_dec(x_31); +lean_dec(x_30); +lean_dec(x_12); +lean_dec(x_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_848 = !lean_is_exclusive(x_820); -if (x_848 == 0) +x_52 = !lean_is_exclusive(x_35); +if (x_52 == 0) { -return x_820; +return x_35; } else { -lean_object* x_849; lean_object* x_850; lean_object* x_851; -x_849 = lean_ctor_get(x_820, 0); -x_850 = lean_ctor_get(x_820, 1); -lean_inc(x_850); -lean_inc(x_849); -lean_dec(x_820); -x_851 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_851, 0, x_849); -lean_ctor_set(x_851, 1, x_850); -return x_851; -} -} -} -} +lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_53 = lean_ctor_get(x_35, 0); +x_54 = lean_ctor_get(x_35, 1); +lean_inc(x_54); +lean_inc(x_53); +lean_dec(x_35); +x_55 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_55, 0, x_53); +lean_ctor_set(x_55, 1, x_54); +return x_55; } } } -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10(lean_object* x_1, lean_object* x_2) { -_start: +default: { -if (lean_obj_tag(x_2) == 0) +lean_object* x_56; lean_object* x_57; +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_56 = l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__3___closed__4; +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_7); +x_57 = l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__7(x_56, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_57) == 0) { -lean_object* x_3; -x_3 = lean_box(0); -return x_3; +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_inc(x_59); +lean_dec(x_57); +x_60 = lean_ctor_get(x_58, 0); +lean_inc(x_60); +x_61 = lean_ctor_get(x_58, 1); +lean_inc(x_61); +lean_dec(x_58); +x_62 = l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__1(x_1, x_60, x_7, x_61, x_9, x_10, x_11, x_12, x_59); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_7); +return x_62; } else { -lean_object* x_4; lean_object* x_5; lean_object* x_6; size_t x_7; size_t x_8; uint8_t x_9; -x_4 = lean_ctor_get(x_2, 0); -x_5 = lean_ctor_get(x_2, 1); -x_6 = lean_ctor_get(x_2, 2); -x_7 = lean_ptr_addr(x_4); -x_8 = lean_ptr_addr(x_1); -x_9 = lean_usize_dec_eq(x_7, x_8); -if (x_9 == 0) +uint8_t x_63; +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_7); +lean_dec(x_1); +x_63 = !lean_is_exclusive(x_57); +if (x_63 == 0) { -x_2 = x_6; -goto _start; +return x_57; } else { -lean_object* x_11; -lean_inc(x_5); -x_11 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_11, 0, x_5); -return x_11; -} -} -} -} -static lean_object* _init_l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__1___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_levelZero; -x_2 = l_Lean_Expr_sort___override(x_1); -return x_2; +lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_64 = lean_ctor_get(x_57, 0); +x_65 = lean_ctor_get(x_57, 1); +lean_inc(x_65); +lean_inc(x_64); +lean_dec(x_57); +x_66 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_66, 0, x_64); +lean_ctor_set(x_66, 1, x_65); +return x_66; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_10 = lean_unsigned_to_nat(0u); -x_11 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_10); -x_12 = l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__1___closed__1; -lean_inc(x_11); -x_13 = lean_mk_array(x_11, x_12); -x_14 = lean_unsigned_to_nat(1u); -x_15 = lean_nat_sub(x_11, x_14); -lean_dec(x_11); -lean_inc(x_1); -x_16 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9(x_1, x_1, x_13, x_15, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -return x_16; } } -static lean_object* _init_l_Lean_Meta_Grind_Canon_canonImpl_visit___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean.Meta.Tactic.Grind.Canon", 28, 28); -return x_1; } } -static lean_object* _init_l_Lean_Meta_Grind_Canon_canonImpl_visit___closed__2() { +static lean_object* _init_l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__4___closed__1() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean.Meta.Grind.Canon.canonImpl.visit", 37, 37); -return x_1; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_instMonadMetaM; +x_2 = lean_alloc_closure((void*)(l_StateT_lift___rarg), 4, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; } } -static lean_object* _init_l_Lean_Meta_Grind_Canon_canonImpl_visit___closed__3() { +static lean_object* _init_l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__4___closed__2() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("unreachable code has been reached", 33, 33); +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)); return x_1; } } -static lean_object* _init_l_Lean_Meta_Grind_Canon_canonImpl_visit___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_Meta_Grind_Canon_canonImpl_visit___closed__1; -x_2 = l_Lean_Meta_Grind_Canon_canonImpl_visit___closed__2; -x_3 = lean_unsigned_to_nat(115u); -x_4 = lean_unsigned_to_nat(18u); -x_5 = l_Lean_Meta_Grind_Canon_canonImpl_visit___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_Meta_Grind_Canon_canonImpl_visit(lean_object* x_1, lean_object* x_2, lean_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_Grind_Canon_canonImpl_visit___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { -switch (lean_obj_tag(x_1)) { -case 0: -{ -lean_object* x_9; lean_object* x_10; -lean_dec(x_1); -x_9 = l_Lean_Meta_Grind_Canon_canonImpl_visit___closed__4; -x_10 = l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__1(x_9, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -return x_10; -} -case 5: -{ -lean_object* x_11; uint8_t x_12; -x_11 = lean_st_ref_get(x_2, x_8); -x_12 = !lean_is_exclusive(x_11); -if (x_12 == 0) +lean_object* x_10; uint8_t x_11; +x_10 = lean_st_ref_get(x_3, x_9); +x_11 = !lean_is_exclusive(x_10); +if (x_11 == 0) { -lean_object* x_13; uint8_t x_14; -x_13 = lean_ctor_get(x_11, 0); -x_14 = !lean_is_exclusive(x_13); -if (x_14 == 0) +lean_object* x_12; uint8_t x_13; +x_12 = lean_ctor_get(x_10, 0); +x_13 = !lean_is_exclusive(x_12); +if (x_13 == 0) { -lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; size_t x_19; uint64_t x_20; uint64_t x_21; uint64_t x_22; uint64_t x_23; uint64_t x_24; uint64_t x_25; uint64_t x_26; uint64_t x_27; uint64_t x_28; size_t x_29; size_t x_30; size_t x_31; size_t x_32; size_t x_33; lean_object* x_34; lean_object* x_35; -x_15 = lean_ctor_get(x_11, 1); -x_16 = lean_ctor_get(x_13, 1); -x_17 = lean_ctor_get(x_13, 0); +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; size_t x_18; uint64_t x_19; uint64_t x_20; uint64_t x_21; uint64_t x_22; uint64_t x_23; uint64_t x_24; uint64_t x_25; uint64_t x_26; uint64_t x_27; size_t x_28; size_t x_29; size_t x_30; size_t x_31; size_t x_32; lean_object* x_33; lean_object* x_34; +x_14 = lean_ctor_get(x_10, 1); +x_15 = lean_ctor_get(x_12, 1); +x_16 = lean_ctor_get(x_12, 0); +lean_dec(x_16); +x_17 = lean_array_get_size(x_15); +x_18 = lean_ptr_addr(x_1); +x_19 = lean_usize_to_uint64(x_18); +x_20 = 11; +x_21 = lean_uint64_mix_hash(x_19, x_20); +x_22 = 32; +x_23 = lean_uint64_shift_right(x_21, x_22); +x_24 = lean_uint64_xor(x_21, x_23); +x_25 = 16; +x_26 = lean_uint64_shift_right(x_24, x_25); +x_27 = lean_uint64_xor(x_24, x_26); +x_28 = lean_uint64_to_usize(x_27); +x_29 = lean_usize_of_nat(x_17); lean_dec(x_17); -x_18 = lean_array_get_size(x_16); -x_19 = lean_ptr_addr(x_1); -x_20 = lean_usize_to_uint64(x_19); -x_21 = 11; -x_22 = lean_uint64_mix_hash(x_20, x_21); -x_23 = 32; -x_24 = lean_uint64_shift_right(x_22, x_23); -x_25 = lean_uint64_xor(x_22, x_24); -x_26 = 16; -x_27 = lean_uint64_shift_right(x_25, x_26); -x_28 = lean_uint64_xor(x_25, x_27); -x_29 = lean_uint64_to_usize(x_28); -x_30 = lean_usize_of_nat(x_18); -lean_dec(x_18); -x_31 = 1; -x_32 = lean_usize_sub(x_30, x_31); -x_33 = lean_usize_land(x_29, x_32); -x_34 = lean_array_uget(x_16, x_33); -lean_dec(x_16); -x_35 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10(x_1, x_34); -lean_dec(x_34); -if (lean_obj_tag(x_35) == 0) +x_30 = 1; +x_31 = lean_usize_sub(x_29, x_30); +x_32 = lean_usize_land(x_28, x_31); +x_33 = lean_array_uget(x_15, x_32); +lean_dec(x_15); +x_34 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__12(x_1, x_33); +lean_dec(x_33); +if (lean_obj_tag(x_34) == 0) { -lean_object* x_36; lean_object* x_37; -lean_free_object(x_13); -lean_free_object(x_11); -x_36 = lean_box(0); -x_37 = l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__1(x_1, x_36, x_2, x_3, x_4, x_5, x_6, x_7, x_15); -return x_37; +lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; +lean_free_object(x_12); +lean_free_object(x_10); +x_35 = l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__7___closed__2; +x_36 = l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__4___closed__1; +x_37 = l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__4___closed__2; +x_38 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__11___closed__1; +x_39 = lean_box(0); +x_40 = l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__3(x_1, x_35, x_36, x_37, x_38, x_39, x_3, x_4, x_5, x_6, x_7, x_8, x_14); +return x_40; } else { -lean_object* x_38; +lean_object* x_41; +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_2); +lean_dec(x_3); lean_dec(x_1); -x_38 = lean_ctor_get(x_35, 0); -lean_inc(x_38); -lean_dec(x_35); -lean_ctor_set(x_13, 1, x_3); -lean_ctor_set(x_13, 0, x_38); -return x_11; +x_41 = lean_ctor_get(x_34, 0); +lean_inc(x_41); +lean_dec(x_34); +lean_ctor_set(x_12, 1, x_4); +lean_ctor_set(x_12, 0, x_41); +return x_10; } } else { -lean_object* x_39; lean_object* x_40; lean_object* x_41; size_t x_42; uint64_t x_43; uint64_t x_44; uint64_t x_45; uint64_t x_46; uint64_t x_47; uint64_t x_48; uint64_t x_49; uint64_t x_50; uint64_t x_51; size_t x_52; size_t x_53; size_t x_54; size_t x_55; size_t x_56; lean_object* x_57; lean_object* x_58; -x_39 = lean_ctor_get(x_11, 1); -x_40 = lean_ctor_get(x_13, 1); -lean_inc(x_40); -lean_dec(x_13); -x_41 = lean_array_get_size(x_40); -x_42 = lean_ptr_addr(x_1); -x_43 = lean_usize_to_uint64(x_42); -x_44 = 11; -x_45 = lean_uint64_mix_hash(x_43, x_44); -x_46 = 32; -x_47 = lean_uint64_shift_right(x_45, x_46); -x_48 = lean_uint64_xor(x_45, x_47); -x_49 = 16; +lean_object* x_42; lean_object* x_43; lean_object* x_44; size_t x_45; uint64_t x_46; uint64_t x_47; uint64_t x_48; uint64_t x_49; uint64_t x_50; uint64_t x_51; uint64_t x_52; uint64_t x_53; uint64_t x_54; size_t x_55; size_t x_56; size_t x_57; size_t x_58; size_t x_59; lean_object* x_60; lean_object* x_61; +x_42 = lean_ctor_get(x_10, 1); +x_43 = lean_ctor_get(x_12, 1); +lean_inc(x_43); +lean_dec(x_12); +x_44 = lean_array_get_size(x_43); +x_45 = lean_ptr_addr(x_1); +x_46 = lean_usize_to_uint64(x_45); +x_47 = 11; +x_48 = lean_uint64_mix_hash(x_46, x_47); +x_49 = 32; x_50 = lean_uint64_shift_right(x_48, x_49); x_51 = lean_uint64_xor(x_48, x_50); -x_52 = lean_uint64_to_usize(x_51); -x_53 = lean_usize_of_nat(x_41); -lean_dec(x_41); -x_54 = 1; -x_55 = lean_usize_sub(x_53, x_54); -x_56 = lean_usize_land(x_52, x_55); -x_57 = lean_array_uget(x_40, x_56); -lean_dec(x_40); -x_58 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10(x_1, x_57); -lean_dec(x_57); -if (lean_obj_tag(x_58) == 0) -{ -lean_object* x_59; lean_object* x_60; -lean_free_object(x_11); -x_59 = lean_box(0); -x_60 = l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__1(x_1, x_59, x_2, x_3, x_4, x_5, x_6, x_7, x_39); -return x_60; +x_52 = 16; +x_53 = lean_uint64_shift_right(x_51, x_52); +x_54 = lean_uint64_xor(x_51, x_53); +x_55 = lean_uint64_to_usize(x_54); +x_56 = lean_usize_of_nat(x_44); +lean_dec(x_44); +x_57 = 1; +x_58 = lean_usize_sub(x_56, x_57); +x_59 = lean_usize_land(x_55, x_58); +x_60 = lean_array_uget(x_43, x_59); +lean_dec(x_43); +x_61 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__12(x_1, x_60); +lean_dec(x_60); +if (lean_obj_tag(x_61) == 0) +{ +lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; +lean_free_object(x_10); +x_62 = l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__7___closed__2; +x_63 = l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__4___closed__1; +x_64 = l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__4___closed__2; +x_65 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__11___closed__1; +x_66 = lean_box(0); +x_67 = l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__3(x_1, x_62, x_63, x_64, x_65, x_66, x_3, x_4, x_5, x_6, x_7, x_8, x_42); +return x_67; } else { -lean_object* x_61; lean_object* x_62; +lean_object* x_68; lean_object* x_69; +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_2); +lean_dec(x_3); lean_dec(x_1); -x_61 = lean_ctor_get(x_58, 0); -lean_inc(x_61); -lean_dec(x_58); -x_62 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_62, 0, x_61); -lean_ctor_set(x_62, 1, x_3); -lean_ctor_set(x_11, 0, x_62); -return x_11; +x_68 = lean_ctor_get(x_61, 0); +lean_inc(x_68); +lean_dec(x_61); +x_69 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_69, 0, x_68); +lean_ctor_set(x_69, 1, x_4); +lean_ctor_set(x_10, 0, x_69); +return x_10; } } } 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; uint64_t x_69; uint64_t x_70; uint64_t x_71; uint64_t x_72; uint64_t x_73; uint64_t x_74; uint64_t x_75; uint64_t x_76; uint64_t x_77; size_t x_78; size_t x_79; size_t x_80; size_t x_81; size_t x_82; lean_object* x_83; lean_object* x_84; -x_63 = lean_ctor_get(x_11, 0); -x_64 = lean_ctor_get(x_11, 1); -lean_inc(x_64); -lean_inc(x_63); -lean_dec(x_11); -x_65 = lean_ctor_get(x_63, 1); -lean_inc(x_65); -if (lean_is_exclusive(x_63)) { - lean_ctor_release(x_63, 0); - lean_ctor_release(x_63, 1); - x_66 = x_63; +lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; size_t x_75; uint64_t x_76; uint64_t x_77; uint64_t x_78; uint64_t x_79; uint64_t x_80; uint64_t x_81; uint64_t x_82; uint64_t x_83; uint64_t x_84; size_t x_85; size_t x_86; size_t x_87; size_t x_88; size_t x_89; lean_object* x_90; lean_object* x_91; +x_70 = lean_ctor_get(x_10, 0); +x_71 = lean_ctor_get(x_10, 1); +lean_inc(x_71); +lean_inc(x_70); +lean_dec(x_10); +x_72 = lean_ctor_get(x_70, 1); +lean_inc(x_72); +if (lean_is_exclusive(x_70)) { + lean_ctor_release(x_70, 0); + lean_ctor_release(x_70, 1); + x_73 = x_70; } else { - lean_dec_ref(x_63); - x_66 = lean_box(0); -} -x_67 = lean_array_get_size(x_65); -x_68 = lean_ptr_addr(x_1); -x_69 = lean_usize_to_uint64(x_68); -x_70 = 11; -x_71 = lean_uint64_mix_hash(x_69, x_70); -x_72 = 32; -x_73 = lean_uint64_shift_right(x_71, x_72); -x_74 = lean_uint64_xor(x_71, x_73); -x_75 = 16; -x_76 = lean_uint64_shift_right(x_74, x_75); -x_77 = lean_uint64_xor(x_74, x_76); -x_78 = lean_uint64_to_usize(x_77); -x_79 = lean_usize_of_nat(x_67); -lean_dec(x_67); -x_80 = 1; -x_81 = lean_usize_sub(x_79, x_80); -x_82 = lean_usize_land(x_78, x_81); -x_83 = lean_array_uget(x_65, x_82); -lean_dec(x_65); -x_84 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10(x_1, x_83); -lean_dec(x_83); -if (lean_obj_tag(x_84) == 0) + lean_dec_ref(x_70); + x_73 = lean_box(0); +} +x_74 = lean_array_get_size(x_72); +x_75 = lean_ptr_addr(x_1); +x_76 = lean_usize_to_uint64(x_75); +x_77 = 11; +x_78 = lean_uint64_mix_hash(x_76, x_77); +x_79 = 32; +x_80 = lean_uint64_shift_right(x_78, x_79); +x_81 = lean_uint64_xor(x_78, x_80); +x_82 = 16; +x_83 = lean_uint64_shift_right(x_81, x_82); +x_84 = lean_uint64_xor(x_81, x_83); +x_85 = lean_uint64_to_usize(x_84); +x_86 = lean_usize_of_nat(x_74); +lean_dec(x_74); +x_87 = 1; +x_88 = lean_usize_sub(x_86, x_87); +x_89 = lean_usize_land(x_85, x_88); +x_90 = lean_array_uget(x_72, x_89); +lean_dec(x_72); +x_91 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__12(x_1, x_90); +lean_dec(x_90); +if (lean_obj_tag(x_91) == 0) { -lean_object* x_85; lean_object* x_86; -lean_dec(x_66); -x_85 = lean_box(0); -x_86 = l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__1(x_1, x_85, x_2, x_3, x_4, x_5, x_6, x_7, x_64); -return x_86; +lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; +lean_dec(x_73); +x_92 = l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__7___closed__2; +x_93 = l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__4___closed__1; +x_94 = l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__4___closed__2; +x_95 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__11___closed__1; +x_96 = lean_box(0); +x_97 = l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__3(x_1, x_92, x_93, x_94, x_95, x_96, x_3, x_4, x_5, x_6, x_7, x_8, x_71); +return x_97; } else { -lean_object* x_87; lean_object* x_88; lean_object* x_89; +lean_object* x_98; lean_object* x_99; lean_object* x_100; +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_2); +lean_dec(x_3); lean_dec(x_1); -x_87 = lean_ctor_get(x_84, 0); -lean_inc(x_87); -lean_dec(x_84); -if (lean_is_scalar(x_66)) { - x_88 = lean_alloc_ctor(0, 2, 0); +x_98 = lean_ctor_get(x_91, 0); +lean_inc(x_98); +lean_dec(x_91); +if (lean_is_scalar(x_73)) { + x_99 = lean_alloc_ctor(0, 2, 0); } else { - x_88 = x_66; + x_99 = x_73; } -lean_ctor_set(x_88, 0, x_87); -lean_ctor_set(x_88, 1, x_3); -x_89 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_89, 0, x_88); -lean_ctor_set(x_89, 1, x_64); -return x_89; +lean_ctor_set(x_99, 0, x_98); +lean_ctor_set(x_99, 1, x_4); +x_100 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_100, 0, x_99); +lean_ctor_set(x_100, 1, x_71); +return x_100; } } } -default: +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_canonImpl_visit(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +uint8_t x_9; +x_9 = l_Lean_Expr_isApp(x_1); +if (x_9 == 0) +{ +uint8_t x_10; +x_10 = l_Lean_Expr_isForall(x_1); +if (x_10 == 0) { -lean_object* x_90; lean_object* x_91; +lean_object* x_11; lean_object* x_12; lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); -x_90 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_90, 0, x_1); -lean_ctor_set(x_90, 1, x_3); -x_91 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_91, 0, x_90); -lean_ctor_set(x_91, 1, x_8); -return x_91; +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_1); +lean_ctor_set(x_11, 1, x_3); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_8); +return x_12; +} +else +{ +lean_object* x_13; lean_object* x_14; +x_13 = lean_box(0); +x_14 = l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__4(x_1, x_13, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +return x_14; +} } +else +{ +lean_object* x_15; lean_object* x_16; +x_15 = lean_box(0); +x_16 = l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__4(x_1, x_15, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +return x_16; } } } -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__2___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__2(x_1, x_2); +x_3 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_Canon_canonImpl_visit___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_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___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, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_Canon_canonImpl_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) { +_start: +{ +lean_object* x_9; +x_9 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Meta_Grind_Canon_canonImpl_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) { +_start: +{ +lean_object* x_10; +x_10 = l_Lean_addTrace___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9(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_3); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_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_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8___lambda__1(x_1, x_2, x_3, x_13, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +x_14 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10___lambda__1(x_1, x_2, x_3, x_13, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -11036,22 +16880,101 @@ lean_dec(x_1); return x_14; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_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, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { _start: { -lean_object* x_16; -x_16 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___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); -lean_dec(x_4); +uint8_t x_15; lean_object* x_16; +x_15 = lean_unbox(x_6); +lean_dec(x_6); +x_16 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10___lambda__2(x_1, x_2, x_3, x_4, x_5, x_15, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +lean_dec(x_7); +lean_dec(x_1); +return x_16; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10___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_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10___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_3); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10___lambda__4___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_object* x_19 = _args[18]; +lean_object* x_20 = _args[19]; +_start: +{ +lean_object* x_21; +x_21 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20); lean_dec(x_3); +return x_21; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10___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_object* x_19 = _args[18]; +lean_object* x_20 = _args[19]; +lean_object* x_21 = _args[20]; +lean_object* x_22 = _args[21]; +_start: +{ +lean_object* x_23; +x_23 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22); +return x_23; +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__12___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__12(x_1, x_2); lean_dec(x_2); -return x_16; +lean_dec(x_1); +return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_canonImpl_visit___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_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_10 = l_Lean_Meta_Grind_Canon_canonImpl_visit___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); @@ -11060,21 +16983,30 @@ lean_dec(x_3); return x_10; } } -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { _start: { -lean_object* x_3; -x_3 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10(x_1, x_2); -lean_dec(x_2); -lean_dec(x_1); -return x_3; +uint8_t x_15; lean_object* x_16; +x_15 = lean_unbox(x_4); +lean_dec(x_4); +x_16 = l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__2(x_1, x_2, x_3, x_15, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +return x_16; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_canonImpl_visit___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: +{ +lean_object* x_14; +x_14 = l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_6); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { lean_object* x_10; -x_10 = l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_10 = l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__4(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; } @@ -11202,12 +17134,155 @@ x_8 = l_Lean_Meta_Grind_Canon_canonImpl(x_1, x_2, x_3, x_4, x_5, x_6, x_7); return x_8; } } +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_canon___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_Meta_Grind_Canon_canonImpl(x_1, x_3, x_4, x_5, x_6, x_7, x_8); +return x_9; +} +} LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_canon(lean_object* x_1, 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_Grind_Canon_canonImpl(x_1, x_2, x_3, x_4, x_5, x_6, x_7); -return x_8; +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; +x_8 = l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___closed__3; +x_9 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_Canon_canonElemCore___spec__4(x_8, x_2, x_3, x_4, x_5, x_6, x_7); +x_10 = lean_ctor_get(x_9, 0); +lean_inc(x_10); +x_11 = lean_ctor_get(x_10, 0); +lean_inc(x_11); +x_12 = lean_unbox(x_11); +lean_dec(x_11); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_13 = lean_ctor_get(x_9, 1); +lean_inc(x_13); +lean_dec(x_9); +x_14 = lean_ctor_get(x_10, 1); +lean_inc(x_14); +lean_dec(x_10); +x_15 = l_Lean_Meta_Grind_Canon_canonImpl(x_1, x_14, x_3, x_4, x_5, x_6, x_13); +return x_15; +} +else +{ +uint8_t x_16; +x_16 = !lean_is_exclusive(x_9); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; uint8_t x_19; +x_17 = lean_ctor_get(x_9, 1); +x_18 = lean_ctor_get(x_9, 0); +lean_dec(x_18); +x_19 = !lean_is_exclusive(x_10); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_20 = lean_ctor_get(x_10, 1); +x_21 = lean_ctor_get(x_10, 0); +lean_dec(x_21); +lean_inc(x_1); +x_22 = l_Lean_MessageData_ofExpr(x_1); +x_23 = l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1___closed__10; +lean_ctor_set_tag(x_10, 7); +lean_ctor_set(x_10, 1, x_22); +lean_ctor_set(x_10, 0, x_23); +lean_ctor_set_tag(x_9, 7); +lean_ctor_set(x_9, 1, x_23); +x_24 = l_Lean_addTrace___at_Lean_Meta_Grind_Canon_canonElemCore___spec__5(x_8, x_9, x_20, x_3, x_4, x_5, x_6, x_17); +x_25 = lean_ctor_get(x_24, 0); +lean_inc(x_25); +x_26 = lean_ctor_get(x_24, 1); +lean_inc(x_26); +lean_dec(x_24); +x_27 = lean_ctor_get(x_25, 1); +lean_inc(x_27); +lean_dec(x_25); +x_28 = l_Lean_Meta_Grind_Canon_canonImpl(x_1, x_27, x_3, x_4, x_5, x_6, x_26); +return x_28; +} +else +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_29 = lean_ctor_get(x_10, 1); +lean_inc(x_29); +lean_dec(x_10); +lean_inc(x_1); +x_30 = l_Lean_MessageData_ofExpr(x_1); +x_31 = l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1___closed__10; +x_32 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_30); +lean_ctor_set_tag(x_9, 7); +lean_ctor_set(x_9, 1, x_31); +lean_ctor_set(x_9, 0, x_32); +x_33 = l_Lean_addTrace___at_Lean_Meta_Grind_Canon_canonElemCore___spec__5(x_8, x_9, x_29, x_3, x_4, x_5, x_6, x_17); +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, 1); +lean_inc(x_36); +lean_dec(x_34); +x_37 = l_Lean_Meta_Grind_Canon_canonImpl(x_1, x_36, x_3, x_4, x_5, x_6, 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; +x_38 = lean_ctor_get(x_9, 1); +lean_inc(x_38); +lean_dec(x_9); +x_39 = lean_ctor_get(x_10, 1); +lean_inc(x_39); +if (lean_is_exclusive(x_10)) { + lean_ctor_release(x_10, 0); + lean_ctor_release(x_10, 1); + x_40 = x_10; +} else { + lean_dec_ref(x_10); + x_40 = lean_box(0); +} +lean_inc(x_1); +x_41 = l_Lean_MessageData_ofExpr(x_1); +x_42 = l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1___closed__10; +if (lean_is_scalar(x_40)) { + x_43 = lean_alloc_ctor(7, 2, 0); +} else { + x_43 = x_40; + lean_ctor_set_tag(x_43, 7); +} +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_43); +lean_ctor_set(x_44, 1, x_42); +x_45 = l_Lean_addTrace___at_Lean_Meta_Grind_Canon_canonElemCore___spec__5(x_8, x_44, x_39, x_3, x_4, x_5, x_6, x_38); +x_46 = lean_ctor_get(x_45, 0); +lean_inc(x_46); +x_47 = lean_ctor_get(x_45, 1); +lean_inc(x_47); +lean_dec(x_45); +x_48 = lean_ctor_get(x_46, 1); +lean_inc(x_48); +lean_dec(x_46); +x_49 = l_Lean_Meta_Grind_Canon_canonImpl(x_1, x_48, x_3, x_4, x_5, x_6, x_47); +return x_49; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_canon___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; +x_9 = l_Lean_Meta_Grind_Canon_canon___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_2); +return x_9; } } lean_object* initialize_Init_Grind_Util(uint8_t builtin, lean_object*); @@ -11247,6 +17322,18 @@ l_Lean_Meta_Grind_Canon_instInhabitedState___closed__3 = _init_l_Lean_Meta_Grind lean_mark_persistent(l_Lean_Meta_Grind_Canon_instInhabitedState___closed__3); l_Lean_Meta_Grind_Canon_instInhabitedState = _init_l_Lean_Meta_Grind_Canon_instInhabitedState(); lean_mark_persistent(l_Lean_Meta_Grind_Canon_instInhabitedState); +l_Lean_Meta_Grind_Canon_CanonElemKind_noConfusion___rarg___closed__1 = _init_l_Lean_Meta_Grind_Canon_CanonElemKind_noConfusion___rarg___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_Canon_CanonElemKind_noConfusion___rarg___closed__1); +l_Lean_Meta_Grind_Canon_instBEqCanonElemKind___closed__1 = _init_l_Lean_Meta_Grind_Canon_instBEqCanonElemKind___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_Canon_instBEqCanonElemKind___closed__1); +l_Lean_Meta_Grind_Canon_instBEqCanonElemKind = _init_l_Lean_Meta_Grind_Canon_instBEqCanonElemKind(); +lean_mark_persistent(l_Lean_Meta_Grind_Canon_instBEqCanonElemKind); +l_Lean_Meta_Grind_Canon_CanonElemKind_explain___closed__1 = _init_l_Lean_Meta_Grind_Canon_CanonElemKind_explain___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_Canon_CanonElemKind_explain___closed__1); +l_Lean_Meta_Grind_Canon_CanonElemKind_explain___closed__2 = _init_l_Lean_Meta_Grind_Canon_CanonElemKind_explain___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_Canon_CanonElemKind_explain___closed__2); +l_Lean_Meta_Grind_Canon_CanonElemKind_explain___closed__3 = _init_l_Lean_Meta_Grind_Canon_CanonElemKind_explain___closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind_Canon_CanonElemKind_explain___closed__3); l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_Canon_canonElemCore___spec__2___closed__1 = _init_l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_Canon_canonElemCore___spec__2___closed__1(); l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_Canon_canonElemCore___spec__2___closed__2 = _init_l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_Canon_canonElemCore___spec__2___closed__2(); l_Lean_addTrace___at_Lean_Meta_Grind_Canon_canonElemCore___spec__5___closed__1 = _init_l_Lean_addTrace___at_Lean_Meta_Grind_Canon_canonElemCore___spec__5___closed__1(); @@ -11273,36 +17360,99 @@ lean_mark_persistent(l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemC l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1___closed__8 = _init_l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1___closed__8(); lean_mark_persistent(l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1___closed__8); l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1___closed__9 = _init_l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1___closed__9(); +lean_mark_persistent(l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1___closed__9); +l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1___closed__10 = _init_l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1___closed__10(); +lean_mark_persistent(l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1___closed__10); +l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1___closed__11 = _init_l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1___closed__11(); +l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___closed__1 = _init_l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___closed__1(); +lean_mark_persistent(l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___closed__1); +l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___closed__2 = _init_l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___closed__2(); +lean_mark_persistent(l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___closed__2); +l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___closed__3 = _init_l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___closed__3(); +lean_mark_persistent(l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___closed__3); +l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___closed__4 = _init_l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___closed__4(); +lean_mark_persistent(l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___closed__4); +l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___closed__5 = _init_l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___closed__5(); +lean_mark_persistent(l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___closed__5); +l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___closed__6 = _init_l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___closed__6(); +lean_mark_persistent(l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___closed__6); +l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___closed__7 = _init_l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___closed__7(); +lean_mark_persistent(l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___closed__7); l_Lean_Meta_Grind_Canon_canonElemCore___lambda__2___closed__1 = _init_l_Lean_Meta_Grind_Canon_canonElemCore___lambda__2___closed__1(); lean_mark_persistent(l_Lean_Meta_Grind_Canon_canonElemCore___lambda__2___closed__1); +l_Lean_Meta_Grind_Canon_canonElemCore___lambda__2___closed__2 = _init_l_Lean_Meta_Grind_Canon_canonElemCore___lambda__2___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_Canon_canonElemCore___lambda__2___closed__2); +l_Lean_Meta_Grind_Canon_canonElemCore___lambda__2___closed__3 = _init_l_Lean_Meta_Grind_Canon_canonElemCore___lambda__2___closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind_Canon_canonElemCore___lambda__2___closed__3); +l_Lean_Meta_Grind_Canon_canonElemCore___lambda__2___closed__4 = _init_l_Lean_Meta_Grind_Canon_canonElemCore___lambda__2___closed__4(); +lean_mark_persistent(l_Lean_Meta_Grind_Canon_canonElemCore___lambda__2___closed__4); +l_Lean_Meta_Grind_Canon_canonElemCore___lambda__2___closed__5 = _init_l_Lean_Meta_Grind_Canon_canonElemCore___lambda__2___closed__5(); +lean_mark_persistent(l_Lean_Meta_Grind_Canon_canonElemCore___lambda__2___closed__5); +l_Lean_Meta_Grind_Canon_canonElemCore___lambda__2___closed__6 = _init_l_Lean_Meta_Grind_Canon_canonElemCore___lambda__2___closed__6(); +lean_mark_persistent(l_Lean_Meta_Grind_Canon_canonElemCore___lambda__2___closed__6); +l_Lean_Meta_Grind_Canon_canonElemCore___lambda__3___closed__1 = _init_l_Lean_Meta_Grind_Canon_canonElemCore___lambda__3___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_Canon_canonElemCore___lambda__3___closed__1); l_Lean_Meta_Grind_Canon_canonInst___closed__1 = _init_l_Lean_Meta_Grind_Canon_canonInst___closed__1(); -l___private_Lean_Meta_Tactic_Grind_Canon_0__Lean_Meta_Grind_Canon_ShouldCanonResult_noConfusion___rarg___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Canon_0__Lean_Meta_Grind_Canon_ShouldCanonResult_noConfusion___rarg___closed__1(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Canon_0__Lean_Meta_Grind_Canon_ShouldCanonResult_noConfusion___rarg___closed__1); +l_Lean_Meta_Grind_Canon_canonImplicit___closed__1 = _init_l_Lean_Meta_Grind_Canon_canonImplicit___closed__1(); l_Lean_Meta_Grind_Canon_instInhabitedShouldCanonResult = _init_l_Lean_Meta_Grind_Canon_instInhabitedShouldCanonResult(); -l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__1___closed__1 = _init_l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__1___closed__1(); -lean_mark_persistent(l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__1___closed__1); -l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__1___closed__2 = _init_l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__1___closed__2(); -lean_mark_persistent(l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__1___closed__2); -l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__1___closed__3 = _init_l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__1___closed__3(); -lean_mark_persistent(l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__1___closed__3); -l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__1 = _init_l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__1(); -lean_mark_persistent(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__1); -l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__2 = _init_l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__2(); -lean_mark_persistent(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__2); -l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__3 = _init_l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__3(); -lean_mark_persistent(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__3); -l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__4 = _init_l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__4(); -lean_mark_persistent(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__4); -l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__1___closed__1 = _init_l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__1___closed__1(); -lean_mark_persistent(l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__1___closed__1); -l_Lean_Meta_Grind_Canon_canonImpl_visit___closed__1 = _init_l_Lean_Meta_Grind_Canon_canonImpl_visit___closed__1(); -lean_mark_persistent(l_Lean_Meta_Grind_Canon_canonImpl_visit___closed__1); -l_Lean_Meta_Grind_Canon_canonImpl_visit___closed__2 = _init_l_Lean_Meta_Grind_Canon_canonImpl_visit___closed__2(); -lean_mark_persistent(l_Lean_Meta_Grind_Canon_canonImpl_visit___closed__2); -l_Lean_Meta_Grind_Canon_canonImpl_visit___closed__3 = _init_l_Lean_Meta_Grind_Canon_canonImpl_visit___closed__3(); -lean_mark_persistent(l_Lean_Meta_Grind_Canon_canonImpl_visit___closed__3); -l_Lean_Meta_Grind_Canon_canonImpl_visit___closed__4 = _init_l_Lean_Meta_Grind_Canon_canonImpl_visit___closed__4(); -lean_mark_persistent(l_Lean_Meta_Grind_Canon_canonImpl_visit___closed__4); +l_Lean_Meta_Grind_Canon_instReprShouldCanonResult___closed__1 = _init_l_Lean_Meta_Grind_Canon_instReprShouldCanonResult___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_Canon_instReprShouldCanonResult___closed__1); +l_Lean_Meta_Grind_Canon_instReprShouldCanonResult___closed__2 = _init_l_Lean_Meta_Grind_Canon_instReprShouldCanonResult___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_Canon_instReprShouldCanonResult___closed__2); +l_Lean_Meta_Grind_Canon_instReprShouldCanonResult___closed__3 = _init_l_Lean_Meta_Grind_Canon_instReprShouldCanonResult___closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind_Canon_instReprShouldCanonResult___closed__3); +l_Lean_Meta_Grind_Canon_instReprShouldCanonResult___closed__4 = _init_l_Lean_Meta_Grind_Canon_instReprShouldCanonResult___closed__4(); +lean_mark_persistent(l_Lean_Meta_Grind_Canon_instReprShouldCanonResult___closed__4); +l_Lean_Meta_Grind_Canon_instReprShouldCanonResult___closed__5 = _init_l_Lean_Meta_Grind_Canon_instReprShouldCanonResult___closed__5(); +lean_mark_persistent(l_Lean_Meta_Grind_Canon_instReprShouldCanonResult___closed__5); +l_Lean_Meta_Grind_Canon_instReprShouldCanonResult___closed__6 = _init_l_Lean_Meta_Grind_Canon_instReprShouldCanonResult___closed__6(); +lean_mark_persistent(l_Lean_Meta_Grind_Canon_instReprShouldCanonResult___closed__6); +l_Lean_Meta_Grind_Canon_instReprShouldCanonResult___closed__7 = _init_l_Lean_Meta_Grind_Canon_instReprShouldCanonResult___closed__7(); +lean_mark_persistent(l_Lean_Meta_Grind_Canon_instReprShouldCanonResult___closed__7); +l_Lean_Meta_Grind_Canon_instReprShouldCanonResult___closed__8 = _init_l_Lean_Meta_Grind_Canon_instReprShouldCanonResult___closed__8(); +lean_mark_persistent(l_Lean_Meta_Grind_Canon_instReprShouldCanonResult___closed__8); +l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__7___closed__1 = _init_l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__7___closed__1(); +lean_mark_persistent(l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__7___closed__1); +l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__7___closed__2 = _init_l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__7___closed__2(); +lean_mark_persistent(l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__7___closed__2); +l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__7___closed__3 = _init_l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__7___closed__3(); +lean_mark_persistent(l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__7___closed__3); +l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10___lambda__3___closed__1 = _init_l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10___lambda__3___closed__1(); +lean_mark_persistent(l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10___lambda__3___closed__1); +l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10___lambda__3___closed__2 = _init_l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10___lambda__3___closed__2(); +lean_mark_persistent(l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10___lambda__3___closed__2); +l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10___lambda__3___closed__3 = _init_l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10___lambda__3___closed__3(); +lean_mark_persistent(l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10___lambda__3___closed__3); +l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10___lambda__3___closed__4 = _init_l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10___lambda__3___closed__4(); +lean_mark_persistent(l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10___lambda__3___closed__4); +l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10___lambda__3___closed__5 = _init_l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10___lambda__3___closed__5(); +lean_mark_persistent(l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10___lambda__3___closed__5); +l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10___lambda__3___closed__6 = _init_l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10___lambda__3___closed__6(); +lean_mark_persistent(l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10___lambda__3___closed__6); +l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__11___closed__1 = _init_l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__11___closed__1(); +lean_mark_persistent(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__11___closed__1); +l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__11___closed__2 = _init_l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__11___closed__2(); +lean_mark_persistent(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__11___closed__2); +l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__11___closed__3 = _init_l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__11___closed__3(); +lean_mark_persistent(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__11___closed__3); +l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__11___closed__4 = _init_l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__11___closed__4(); +lean_mark_persistent(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__11___closed__4); +l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__11___closed__5 = _init_l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__11___closed__5(); +lean_mark_persistent(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__11___closed__5); +l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__3___closed__1 = _init_l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__3___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__3___closed__1); +l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__3___closed__2 = _init_l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__3___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__3___closed__2); +l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__3___closed__3 = _init_l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__3___closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__3___closed__3); +l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__3___closed__4 = _init_l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__3___closed__4(); +lean_mark_persistent(l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__3___closed__4); +l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__3___closed__5 = _init_l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__3___closed__5(); +lean_mark_persistent(l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__3___closed__5); +l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__4___closed__1 = _init_l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__4___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__4___closed__1); +l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__4___closed__2 = _init_l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__4___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__4___closed__2); l_Lean_Meta_Grind_Canon_canonImpl___closed__1 = _init_l_Lean_Meta_Grind_Canon_canonImpl___closed__1(); lean_mark_persistent(l_Lean_Meta_Grind_Canon_canonImpl___closed__1); return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Lean/Meta/Tactic/Grind/Cases.c b/stage0/stdlib/Lean/Meta/Tactic/Grind/Cases.c index 85a332d356ac..381b145afbf4 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Grind/Cases.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Grind/Cases.c @@ -13,65 +13,113 @@ #ifdef __cplusplus extern "C" { #endif +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___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_EXPORT lean_object* l_Lean_Meta_Grind_cases___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_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__8___boxed(lean_object**); lean_object* l_Array_mapMUnsafe_map___at_Lean_LocalContext_getFVars___spec__1(size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_getConstInfo___at_Lean_Meta_mkConstWithFreshMVarLevels___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_throwTacticEx___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkAppN(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___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_whnf(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___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_Meta_Grind_cases___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_Lean_Meta_Grind_cases___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___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_array_push(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__6___boxed(lean_object**); static lean_object* l_Lean_Meta_Grind_cases_throwInductiveExpected___closed__2; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___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_EXPORT lean_object* l_Lean_Meta_Grind_cases___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*); extern lean_object* l_Lean_casesOnSuffix; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___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_MVarId_getTag(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___lambda__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___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*); lean_object* l_Lean_Meta_RecursorInfo_numIndices(lean_object*); +static lean_object* l_Lean_Meta_Grind_cases___lambda__13___closed__4; +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__13___boxed(lean_object**); lean_object* l_Lean_stringToMessageData(lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__7___boxed(lean_object**); static lean_object* l_Lean_Meta_Grind_cases_throwInductiveExpected___closed__4; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___lambda__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___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* 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_Meta_Grind_cases___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_Grind_cases___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__11___boxed(lean_object**); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static size_t l_Lean_Meta_Grind_cases___lambda__13___closed__3; +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__4___boxed(lean_object**); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_generalizeIndices_x27(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__1___closed__4; lean_object* l_Lean_MessageData_ofFormat(lean_object*); lean_object* l_Lean_MVarId_withContext___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___lambda__13(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_to_list(lean_object*); -lean_object* l_Lean_FVarId_getType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_cases___lambda__1___closed__1; +lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__12(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_cases_throwInductiveExpected___closed__1; lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); -lean_object* l_Lean_MVarId_clear(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__1___closed__2; -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___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_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___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_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__10___boxed(lean_object**); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__5___boxed(lean_object**); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___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* l_Lean_MVarId_assert(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___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_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___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*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___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* l_Lean_MessageData_ofExpr(lean_object*); -lean_object* l_Lean_Meta_getMajorTypeIndices(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__10(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_cases___lambda__12___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___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_Grind_cases___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* l_Lean_Meta_mkRecursorAppPrefix(lean_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_Grind_cases___lambda__13___closed__1; lean_object* l_Lean_Meta_mkRecursorInfo(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_app___override(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__9___boxed(lean_object**); lean_object* l_Lean_Meta_RecursorInfo_numMinors(lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_cases_throwInductiveExpected___closed__6; -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___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_Grind_cases___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_Name_mkStr2(lean_object*, lean_object*); lean_object* l_Lean_indentExpr(lean_object*); -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___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_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__1___boxed(lean_object**); static lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__1___closed__1; lean_object* l_Lean_Expr_getAppFn(lean_object*); lean_object* lean_array_mk(lean_object*); -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___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* l_Lean_MVarId_tryClearMany(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___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*); +static lean_object* l_Lean_Meta_Grind_cases___lambda__13___closed__2; lean_object* l_Lean_Expr_fvar___override(lean_object*); size_t lean_array_size(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___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_EXPORT lean_object* l_Lean_Meta_Grind_cases___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_Meta_mkFreshExprSyntheticOpaqueMVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases_throwInductiveExpected___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_Grind_cases___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___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_Grind_cases___lambda__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_cases_throwInductiveExpected___closed__3; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___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_infer_type(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_intro1Core(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__2___boxed(lean_object**); static lean_object* l_Lean_Meta_Grind_cases_throwInductiveExpected___closed__5; lean_object* lean_nat_add(lean_object*, lean_object*); -lean_object* l_Lean_Meta_generalizeIndices(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__12___boxed(lean_object**); static lean_object* l_Lean_Meta_Grind_cases_throwInductiveExpected___closed__7; static lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__1___closed__3; lean_object* l_Lean_Expr_mvarId_x21(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases_throwInductiveExpected(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__11(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___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_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__3___boxed(lean_object**); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* _init_l_Lean_Meta_Grind_cases_throwInductiveExpected___closed__1() { _start: { @@ -135,29 +183,28 @@ return x_2; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases_throwInductiveExpected(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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; -x_10 = l_Lean_Expr_fvar___override(x_2); -x_11 = l_Lean_MessageData_ofExpr(x_10); -x_12 = l_Lean_Meta_Grind_cases_throwInductiveExpected___closed__5; -x_13 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_13, 0, x_12); -lean_ctor_set(x_13, 1, x_11); -x_14 = l_Lean_Meta_Grind_cases_throwInductiveExpected___closed__7; -x_15 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_15, 0, x_13); -lean_ctor_set(x_15, 1, x_14); -x_16 = l_Lean_indentExpr(x_4); +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_10 = l_Lean_MessageData_ofExpr(x_2); +x_11 = l_Lean_Meta_Grind_cases_throwInductiveExpected___closed__5; +x_12 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_10); +x_13 = l_Lean_Meta_Grind_cases_throwInductiveExpected___closed__7; +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_indentExpr(x_4); +x_16 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_16, 0, x_14); +lean_ctor_set(x_16, 1, x_15); x_17 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_17, 0, x_15); -lean_ctor_set(x_17, 1, x_16); -x_18 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_17, 0, x_16); +lean_ctor_set(x_17, 1, x_13); +x_18 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_18, 0, x_17); -lean_ctor_set(x_18, 1, x_14); -x_19 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_19, 0, x_18); -x_20 = l_Lean_Meta_Grind_cases_throwInductiveExpected___closed__3; -x_21 = l_Lean_Meta_throwTacticEx___rarg(x_20, x_1, x_19, x_5, x_6, x_7, x_8, x_9); -return x_21; +x_19 = l_Lean_Meta_Grind_cases_throwInductiveExpected___closed__3; +x_20 = l_Lean_Meta_throwTacticEx___rarg(x_19, x_1, x_18, x_5, x_6, x_7, x_8, x_9); +return x_20; } } LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases_throwInductiveExpected___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { @@ -172,25 +219,6 @@ lean_dec(x_5); return x_10; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___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) { -_start: -{ -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_10 = lean_array_push(x_3, x_4); -x_11 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_11, 0, x_1); -lean_ctor_set(x_11, 1, x_2); -x_12 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_12, 0, x_10); -lean_ctor_set(x_12, 1, x_11); -x_13 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_13, 0, x_12); -x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_13); -lean_ctor_set(x_14, 1, x_9); -return x_14; -} -} static lean_object* _init_l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__1___closed__1() { _start: { @@ -228,986 +256,10850 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___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_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18) { _start: { -lean_object* x_16; uint8_t x_17; -x_16 = lean_ctor_get(x_6, 1); -x_17 = lean_nat_dec_lt(x_8, x_16); -if (x_17 == 0) +lean_object* x_19; uint8_t x_20; +x_19 = lean_ctor_get(x_9, 1); +x_20 = lean_nat_dec_lt(x_11, x_19); +if (x_20 == 0) { -lean_object* x_18; +lean_object* x_21; +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); lean_dec(x_11); -lean_dec(x_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_18 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_18, 0, x_7); -lean_ctor_set(x_18, 1, x_15); -return x_18; +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_19; lean_object* x_20; 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_7, 0); -lean_inc(x_26); -x_27 = lean_ctor_get(x_7, 1); -lean_inc(x_27); -lean_dec(x_7); -x_28 = lean_ctor_get(x_27, 0); -lean_inc(x_28); -x_29 = lean_ctor_get(x_27, 1); +lean_object* x_22; lean_object* x_23; lean_object* x_29; 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_10, 0); lean_inc(x_29); -lean_dec(x_27); -lean_inc(x_14); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -x_30 = lean_whnf(x_29, x_11, x_12, x_13, x_14, x_15); -if (lean_obj_tag(x_30) == 0) -{ -lean_object* x_31; +x_30 = lean_ctor_get(x_10, 1); +lean_inc(x_30); +lean_dec(x_10); x_31 = lean_ctor_get(x_30, 0); lean_inc(x_31); -if (lean_obj_tag(x_31) == 7) -{ -lean_object* x_32; lean_object* x_33; lean_object* x_34; 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_32 = lean_ctor_get(x_30, 1); lean_inc(x_32); lean_dec(x_30); -x_33 = lean_ctor_get(x_31, 1); +x_33 = lean_ctor_get(x_32, 0); lean_inc(x_33); -x_34 = lean_ctor_get(x_31, 2); +x_34 = lean_ctor_get(x_32, 1); lean_inc(x_34); -lean_dec(x_31); -lean_inc(x_11); -lean_inc(x_3); -x_35 = l_Lean_Meta_mkFreshExprSyntheticOpaqueMVar(x_33, x_3, x_11, x_12, x_13, x_14, x_32); +lean_dec(x_32); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_35 = lean_whnf(x_34, x_14, x_15, x_16, x_17, x_18); +if (lean_obj_tag(x_35) == 0) +{ +lean_object* x_36; x_36 = lean_ctor_get(x_35, 0); lean_inc(x_36); +if (lean_obj_tag(x_36) == 7) +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; uint8_t x_41; x_37 = lean_ctor_get(x_35, 1); lean_inc(x_37); lean_dec(x_35); -lean_inc(x_36); -x_38 = l_Lean_Expr_app___override(x_28, x_36); -x_39 = l_Lean_Expr_mvarId_x21(x_36); +x_38 = lean_ctor_get(x_36, 1); +lean_inc(x_38); +x_39 = lean_ctor_get(x_36, 2); +lean_inc(x_39); lean_dec(x_36); +x_40 = lean_unsigned_to_nat(1u); +x_41 = lean_nat_dec_lt(x_40, x_7); +if (x_41 == 0) +{ +lean_object* x_42; uint8_t x_43; lean_inc(x_14); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_2); -x_40 = l_Lean_MVarId_clear(x_39, x_2, x_11, x_12, x_13, x_14, x_37); -if (lean_obj_tag(x_40) == 0) +lean_inc(x_1); +x_42 = l_Lean_Meta_mkFreshExprSyntheticOpaqueMVar(x_38, x_1, x_14, x_15, x_16, x_17, x_37); +x_43 = !lean_is_exclusive(x_42); +if (x_43 == 0) { -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_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_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__1___lambda__1(x_38, x_34, x_26, x_41, x_11, x_12, x_13, x_14, x_42); -x_44 = lean_ctor_get(x_43, 0); +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_44 = lean_ctor_get(x_42, 0); +x_45 = lean_ctor_get(x_42, 1); lean_inc(x_44); -x_45 = lean_ctor_get(x_43, 1); -lean_inc(x_45); -lean_dec(x_43); -x_19 = x_44; -x_20 = x_45; -goto block_25; -} -else -{ -uint8_t x_46; -lean_dec(x_38); -lean_dec(x_34); -lean_dec(x_26); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_8); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_46 = !lean_is_exclusive(x_40); -if (x_46 == 0) -{ -return x_40; -} -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; -} -} -} -else +x_46 = l_Lean_Expr_app___override(x_33, x_44); +x_47 = l_Lean_Expr_mvarId_x21(x_44); +lean_dec(x_44); +lean_inc(x_3); +lean_inc(x_5); +x_48 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_49 = l_Lean_MVarId_tryClearMany(x_47, x_48, x_14, x_15, x_16, x_17, x_45); +lean_dec(x_48); +if (lean_obj_tag(x_49) == 0) { -lean_object* x_50; lean_object* x_51; lean_object* x_52; uint8_t x_53; -lean_dec(x_31); -lean_dec(x_28); -lean_dec(x_26); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -x_50 = lean_ctor_get(x_30, 1); +lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_50 = lean_ctor_get(x_49, 0); lean_inc(x_50); -lean_dec(x_30); -x_51 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__1___closed__4; -x_52 = l_Lean_Meta_throwTacticEx___rarg(x_4, x_1, x_51, x_11, x_12, x_13, x_14, x_50); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -x_53 = !lean_is_exclusive(x_52); -if (x_53 == 0) -{ -return x_52; -} -else -{ -lean_object* x_54; lean_object* x_55; lean_object* x_56; -x_54 = lean_ctor_get(x_52, 0); -x_55 = lean_ctor_get(x_52, 1); -lean_inc(x_55); -lean_inc(x_54); -lean_dec(x_52); -x_56 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_56, 0, x_54); -lean_ctor_set(x_56, 1, x_55); -return x_56; -} -} +x_51 = lean_ctor_get(x_49, 1); +lean_inc(x_51); +lean_dec(x_49); +x_52 = lean_array_push(x_31, x_50); +x_53 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +lean_ctor_set(x_42, 1, x_39); +lean_ctor_set(x_42, 0, x_46); +x_54 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_54, 0, x_52); +lean_ctor_set(x_54, 1, x_42); +x_55 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_55, 0, x_53); +lean_ctor_set(x_55, 1, x_54); +x_56 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_56, 0, x_55); +x_22 = x_56; +x_23 = x_51; +goto block_28; } else { uint8_t x_57; -lean_dec(x_28); -lean_dec(x_26); +lean_dec(x_46); +lean_free_object(x_42); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); lean_dec(x_11); -lean_dec(x_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_57 = !lean_is_exclusive(x_30); +x_57 = !lean_is_exclusive(x_49); if (x_57 == 0) { -return x_30; +return x_49; } else { lean_object* x_58; lean_object* x_59; lean_object* x_60; -x_58 = lean_ctor_get(x_30, 0); -x_59 = lean_ctor_get(x_30, 1); +x_58 = lean_ctor_get(x_49, 0); +x_59 = lean_ctor_get(x_49, 1); lean_inc(x_59); lean_inc(x_58); -lean_dec(x_30); +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; } } -block_25: -{ -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_19, 0); -lean_inc(x_21); -lean_dec(x_19); -x_22 = lean_ctor_get(x_6, 2); -x_23 = lean_nat_add(x_8, x_22); -lean_dec(x_8); -x_7 = x_21; -x_8 = x_23; -x_9 = lean_box(0); -x_10 = lean_box(0); -x_15 = x_20; -goto _start; -} -} } -} -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___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) { -_start: +else { -lean_object* x_15; uint8_t x_16; -x_15 = lean_ctor_get(x_5, 1); -x_16 = lean_nat_dec_lt(x_7, x_15); -if (x_16 == 0) +lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_61 = lean_ctor_get(x_42, 0); +x_62 = lean_ctor_get(x_42, 1); +lean_inc(x_62); +lean_inc(x_61); +lean_dec(x_42); +lean_inc(x_61); +x_63 = l_Lean_Expr_app___override(x_33, x_61); +x_64 = l_Lean_Expr_mvarId_x21(x_61); +lean_dec(x_61); +lean_inc(x_3); +lean_inc(x_5); +x_65 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_66 = l_Lean_MVarId_tryClearMany(x_64, x_65, x_14, x_15, x_16, x_17, x_62); +lean_dec(x_65); +if (lean_obj_tag(x_66) == 0) { -lean_object* x_17; -lean_dec(x_13); -lean_dec(x_12); +lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; +x_67 = 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_array_push(x_31, x_67); +x_70 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +x_71 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_71, 0, x_63); +lean_ctor_set(x_71, 1, x_39); +x_72 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_72, 0, x_69); +lean_ctor_set(x_72, 1, x_71); +x_73 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_73, 0, x_70); +lean_ctor_set(x_73, 1, x_72); +x_74 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_74, 0, x_73); +x_22 = x_74; +x_23 = x_68; +goto block_28; +} +else +{ +lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; +lean_dec(x_63); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +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_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); lean_dec(x_1); -x_17 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_17, 0, x_6); -lean_ctor_set(x_17, 1, x_14); -return x_17; +x_75 = lean_ctor_get(x_66, 0); +lean_inc(x_75); +x_76 = lean_ctor_get(x_66, 1); +lean_inc(x_76); +if (lean_is_exclusive(x_66)) { + lean_ctor_release(x_66, 0); + lean_ctor_release(x_66, 1); + x_77 = x_66; +} else { + lean_dec_ref(x_66); + x_77 = lean_box(0); +} +if (lean_is_scalar(x_77)) { + x_78 = lean_alloc_ctor(1, 2, 0); +} else { + x_78 = x_77; +} +lean_ctor_set(x_78, 0, x_75); +lean_ctor_set(x_78, 1, x_76); +return x_78; +} +} } else { -lean_object* x_18; lean_object* x_19; 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_6, 0); -lean_inc(x_25); -x_26 = lean_ctor_get(x_6, 1); -lean_inc(x_26); -lean_dec(x_6); -x_27 = lean_ctor_get(x_26, 0); -lean_inc(x_27); -x_28 = lean_ctor_get(x_26, 1); -lean_inc(x_28); -lean_dec(x_26); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -x_29 = lean_whnf(x_28, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_29) == 0) +lean_object* x_79; lean_object* x_80; uint8_t x_81; +lean_inc(x_29); +lean_inc(x_1); +x_79 = l_Lean_Name_num___override(x_1, x_29); +lean_inc(x_14); +x_80 = l_Lean_Meta_mkFreshExprSyntheticOpaqueMVar(x_38, x_79, x_14, x_15, x_16, x_17, x_37); +x_81 = !lean_is_exclusive(x_80); +if (x_81 == 0) { -lean_object* x_30; -x_30 = lean_ctor_get(x_29, 0); -lean_inc(x_30); -if (lean_obj_tag(x_30) == 7) +lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; +x_82 = lean_ctor_get(x_80, 0); +x_83 = lean_ctor_get(x_80, 1); +lean_inc(x_82); +x_84 = l_Lean_Expr_app___override(x_33, x_82); +x_85 = l_Lean_Expr_mvarId_x21(x_82); +lean_dec(x_82); +lean_inc(x_3); +lean_inc(x_5); +x_86 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_87 = l_Lean_MVarId_tryClearMany(x_85, x_86, x_14, x_15, x_16, x_17, x_83); +lean_dec(x_86); +if (lean_obj_tag(x_87) == 0) { -lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; -x_31 = lean_ctor_get(x_29, 1); -lean_inc(x_31); +lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; +x_88 = lean_ctor_get(x_87, 0); +lean_inc(x_88); +x_89 = lean_ctor_get(x_87, 1); +lean_inc(x_89); +lean_dec(x_87); +x_90 = lean_array_push(x_31, x_88); +x_91 = lean_nat_add(x_29, x_40); lean_dec(x_29); -x_32 = lean_ctor_get(x_30, 1); -lean_inc(x_32); -x_33 = lean_ctor_get(x_30, 2); -lean_inc(x_33); -lean_dec(x_30); -lean_inc(x_10); -lean_inc(x_1); -x_34 = l_Lean_Meta_mkFreshExprSyntheticOpaqueMVar(x_32, x_1, x_10, x_11, x_12, x_13, x_31); -x_35 = lean_ctor_get(x_34, 0); -lean_inc(x_35); -x_36 = lean_ctor_get(x_34, 1); -lean_inc(x_36); -lean_dec(x_34); -lean_inc(x_35); -x_37 = l_Lean_Expr_app___override(x_27, x_35); -x_38 = l_Lean_Expr_mvarId_x21(x_35); -lean_dec(x_35); -x_39 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__1___lambda__1(x_37, x_33, x_25, x_38, x_10, x_11, x_12, x_13, x_36); -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_18 = x_40; -x_19 = x_41; -goto block_24; +lean_ctor_set(x_80, 1, x_39); +lean_ctor_set(x_80, 0, x_84); +x_92 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_92, 0, x_90); +lean_ctor_set(x_92, 1, x_80); +x_93 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_93, 0, x_91); +lean_ctor_set(x_93, 1, x_92); +x_94 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_94, 0, x_93); +x_22 = x_94; +x_23 = x_89; +goto block_28; } else { -lean_object* x_42; lean_object* x_43; lean_object* x_44; uint8_t x_45; -lean_dec(x_30); -lean_dec(x_27); -lean_dec(x_25); -lean_dec(x_7); -lean_dec(x_1); -x_42 = lean_ctor_get(x_29, 1); -lean_inc(x_42); +uint8_t x_95; +lean_dec(x_84); +lean_free_object(x_80); +lean_dec(x_39); +lean_dec(x_31); lean_dec(x_29); -x_43 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__1___closed__4; -x_44 = l_Lean_Meta_throwTacticEx___rarg(x_3, x_2, x_43, x_10, x_11, x_12, x_13, x_42); -lean_dec(x_13); -lean_dec(x_12); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); lean_dec(x_11); -lean_dec(x_10); -x_45 = !lean_is_exclusive(x_44); -if (x_45 == 0) +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_95 = !lean_is_exclusive(x_87); +if (x_95 == 0) { -return x_44; +return x_87; } 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_96; lean_object* x_97; lean_object* x_98; +x_96 = lean_ctor_get(x_87, 0); +x_97 = lean_ctor_get(x_87, 1); +lean_inc(x_97); +lean_inc(x_96); +lean_dec(x_87); +x_98 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_98, 0, x_96); +lean_ctor_set(x_98, 1, x_97); +return x_98; } } } else { -uint8_t x_49; -lean_dec(x_27); -lean_dec(x_25); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_7); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_49 = !lean_is_exclusive(x_29); -if (x_49 == 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_80, 0); +x_100 = lean_ctor_get(x_80, 1); +lean_inc(x_100); +lean_inc(x_99); +lean_dec(x_80); +lean_inc(x_99); +x_101 = l_Lean_Expr_app___override(x_33, x_99); +x_102 = l_Lean_Expr_mvarId_x21(x_99); +lean_dec(x_99); +lean_inc(x_3); +lean_inc(x_5); +x_103 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_104 = l_Lean_MVarId_tryClearMany(x_102, x_103, x_14, x_15, x_16, x_17, x_100); +lean_dec(x_103); +if (lean_obj_tag(x_104) == 0) { -return x_29; +lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; +x_105 = lean_ctor_get(x_104, 0); +lean_inc(x_105); +x_106 = lean_ctor_get(x_104, 1); +lean_inc(x_106); +lean_dec(x_104); +x_107 = lean_array_push(x_31, x_105); +x_108 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +x_109 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_109, 0, x_101); +lean_ctor_set(x_109, 1, x_39); +x_110 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_110, 0, x_107); +lean_ctor_set(x_110, 1, x_109); +x_111 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_111, 0, x_108); +lean_ctor_set(x_111, 1, x_110); +x_112 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_112, 0, x_111); +x_22 = x_112; +x_23 = x_106; +goto block_28; } else { -lean_object* x_50; lean_object* x_51; lean_object* x_52; -x_50 = lean_ctor_get(x_29, 0); -x_51 = lean_ctor_get(x_29, 1); -lean_inc(x_51); -lean_inc(x_50); +lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; +lean_dec(x_101); +lean_dec(x_39); +lean_dec(x_31); lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_113 = lean_ctor_get(x_104, 0); +lean_inc(x_113); +x_114 = lean_ctor_get(x_104, 1); +lean_inc(x_114); +if (lean_is_exclusive(x_104)) { + lean_ctor_release(x_104, 0); + lean_ctor_release(x_104, 1); + x_115 = x_104; +} else { + lean_dec_ref(x_104); + x_115 = lean_box(0); +} +if (lean_is_scalar(x_115)) { + x_116 = lean_alloc_ctor(1, 2, 0); +} else { + x_116 = x_115; +} +lean_ctor_set(x_116, 0, x_113); +lean_ctor_set(x_116, 1, x_114); +return x_116; +} +} +} +} +else +{ +lean_object* x_117; lean_object* x_118; lean_object* x_119; uint8_t x_120; +lean_dec(x_36); +lean_dec(x_33); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_11); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +x_117 = lean_ctor_get(x_35, 1); +lean_inc(x_117); +lean_dec(x_35); +x_118 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__1___closed__4; +x_119 = l_Lean_Meta_throwTacticEx___rarg(x_6, x_4, x_118, x_14, x_15, x_16, x_17, x_117); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +x_120 = !lean_is_exclusive(x_119); +if (x_120 == 0) +{ +return x_119; +} +else +{ +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_124; +lean_dec(x_33); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_124 = !lean_is_exclusive(x_35); +if (x_124 == 0) +{ +return x_35; +} +else +{ +lean_object* x_125; lean_object* x_126; lean_object* x_127; +x_125 = lean_ctor_get(x_35, 0); +x_126 = lean_ctor_get(x_35, 1); +lean_inc(x_126); +lean_inc(x_125); +lean_dec(x_35); +x_127 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_127, 0, x_125); +lean_ctor_set(x_127, 1, x_126); +return x_127; +} +} +block_28: +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_22, 0); +lean_inc(x_24); +lean_dec(x_22); +x_25 = lean_ctor_get(x_9, 2); +x_26 = lean_nat_add(x_11, x_25); +lean_dec(x_11); +x_10 = x_24; +x_11 = x_26; +x_12 = lean_box(0); +x_13 = lean_box(0); +x_18 = x_23; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18) { +_start: +{ +lean_object* x_19; uint8_t x_20; +x_19 = lean_ctor_get(x_9, 1); +x_20 = lean_nat_dec_lt(x_11, x_19); +if (x_20 == 0) +{ +lean_object* x_21; +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +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; lean_object* x_23; lean_object* x_29; 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_10, 0); +lean_inc(x_29); +x_30 = lean_ctor_get(x_10, 1); +lean_inc(x_30); +lean_dec(x_10); +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_32, 0); +lean_inc(x_33); +x_34 = lean_ctor_get(x_32, 1); +lean_inc(x_34); +lean_dec(x_32); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_35 = lean_whnf(x_34, x_14, x_15, x_16, x_17, x_18); +if (lean_obj_tag(x_35) == 0) +{ +lean_object* x_36; +x_36 = lean_ctor_get(x_35, 0); +lean_inc(x_36); +if (lean_obj_tag(x_36) == 7) +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; uint8_t x_41; +x_37 = lean_ctor_get(x_35, 1); +lean_inc(x_37); +lean_dec(x_35); +x_38 = lean_ctor_get(x_36, 1); +lean_inc(x_38); +x_39 = lean_ctor_get(x_36, 2); +lean_inc(x_39); +lean_dec(x_36); +x_40 = lean_unsigned_to_nat(1u); +x_41 = lean_nat_dec_lt(x_40, x_7); +if (x_41 == 0) +{ +lean_object* x_42; uint8_t x_43; +lean_inc(x_14); +lean_inc(x_2); +x_42 = l_Lean_Meta_mkFreshExprSyntheticOpaqueMVar(x_38, x_2, x_14, x_15, x_16, x_17, x_37); +x_43 = !lean_is_exclusive(x_42); +if (x_43 == 0) +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_44 = lean_ctor_get(x_42, 0); +x_45 = lean_ctor_get(x_42, 1); +lean_inc(x_44); +x_46 = l_Lean_Expr_app___override(x_33, x_44); +x_47 = l_Lean_Expr_mvarId_x21(x_44); +lean_dec(x_44); +lean_inc(x_4); +lean_inc(x_5); +x_48 = lean_array_push(x_5, x_4); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_49 = l_Lean_MVarId_tryClearMany(x_47, x_48, x_14, x_15, x_16, x_17, x_45); +lean_dec(x_48); +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; +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_array_push(x_31, x_50); +x_53 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +lean_ctor_set(x_42, 1, x_39); +lean_ctor_set(x_42, 0, x_46); +x_54 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_54, 0, x_52); +lean_ctor_set(x_54, 1, x_42); +x_55 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_55, 0, x_53); +lean_ctor_set(x_55, 1, x_54); +x_56 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_56, 0, x_55); +x_22 = x_56; +x_23 = x_51; +goto block_28; +} +else +{ +uint8_t x_57; +lean_dec(x_46); +lean_free_object(x_42); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_57 = !lean_is_exclusive(x_49); +if (x_57 == 0) +{ +return x_49; +} +else +{ +lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_58 = lean_ctor_get(x_49, 0); +x_59 = lean_ctor_get(x_49, 1); +lean_inc(x_59); +lean_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; +} +} +} +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; +x_61 = lean_ctor_get(x_42, 0); +x_62 = lean_ctor_get(x_42, 1); +lean_inc(x_62); +lean_inc(x_61); +lean_dec(x_42); +lean_inc(x_61); +x_63 = l_Lean_Expr_app___override(x_33, x_61); +x_64 = l_Lean_Expr_mvarId_x21(x_61); +lean_dec(x_61); +lean_inc(x_4); +lean_inc(x_5); +x_65 = lean_array_push(x_5, x_4); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_66 = l_Lean_MVarId_tryClearMany(x_64, x_65, x_14, x_15, x_16, x_17, x_62); +lean_dec(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; +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_array_push(x_31, x_67); +x_70 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +x_71 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_71, 0, x_63); +lean_ctor_set(x_71, 1, x_39); +x_72 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_72, 0, x_69); +lean_ctor_set(x_72, 1, x_71); +x_73 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_73, 0, x_70); +lean_ctor_set(x_73, 1, x_72); +x_74 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_74, 0, x_73); +x_22 = x_74; +x_23 = x_68; +goto block_28; +} +else +{ +lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; +lean_dec(x_63); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_75 = lean_ctor_get(x_66, 0); +lean_inc(x_75); +x_76 = lean_ctor_get(x_66, 1); +lean_inc(x_76); +if (lean_is_exclusive(x_66)) { + lean_ctor_release(x_66, 0); + lean_ctor_release(x_66, 1); + x_77 = x_66; +} else { + lean_dec_ref(x_66); + x_77 = lean_box(0); +} +if (lean_is_scalar(x_77)) { + x_78 = lean_alloc_ctor(1, 2, 0); +} else { + x_78 = x_77; +} +lean_ctor_set(x_78, 0, x_75); +lean_ctor_set(x_78, 1, x_76); +return x_78; +} +} +} +else +{ +lean_object* x_79; lean_object* x_80; uint8_t x_81; +lean_inc(x_29); +lean_inc(x_2); +x_79 = l_Lean_Name_num___override(x_2, x_29); +lean_inc(x_14); +x_80 = l_Lean_Meta_mkFreshExprSyntheticOpaqueMVar(x_38, x_79, x_14, x_15, x_16, x_17, x_37); +x_81 = !lean_is_exclusive(x_80); +if (x_81 == 0) +{ +lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; +x_82 = lean_ctor_get(x_80, 0); +x_83 = lean_ctor_get(x_80, 1); +lean_inc(x_82); +x_84 = l_Lean_Expr_app___override(x_33, x_82); +x_85 = l_Lean_Expr_mvarId_x21(x_82); +lean_dec(x_82); +lean_inc(x_4); +lean_inc(x_5); +x_86 = lean_array_push(x_5, x_4); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_87 = l_Lean_MVarId_tryClearMany(x_85, x_86, x_14, x_15, x_16, x_17, x_83); +lean_dec(x_86); +if (lean_obj_tag(x_87) == 0) +{ +lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; +x_88 = lean_ctor_get(x_87, 0); +lean_inc(x_88); +x_89 = lean_ctor_get(x_87, 1); +lean_inc(x_89); +lean_dec(x_87); +x_90 = lean_array_push(x_31, x_88); +x_91 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +lean_ctor_set(x_80, 1, x_39); +lean_ctor_set(x_80, 0, x_84); +x_92 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_92, 0, x_90); +lean_ctor_set(x_92, 1, x_80); +x_93 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_93, 0, x_91); +lean_ctor_set(x_93, 1, x_92); +x_94 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_94, 0, x_93); +x_22 = x_94; +x_23 = x_89; +goto block_28; +} +else +{ +uint8_t x_95; +lean_dec(x_84); +lean_free_object(x_80); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_95 = !lean_is_exclusive(x_87); +if (x_95 == 0) +{ +return x_87; +} +else +{ +lean_object* x_96; lean_object* x_97; lean_object* x_98; +x_96 = lean_ctor_get(x_87, 0); +x_97 = lean_ctor_get(x_87, 1); +lean_inc(x_97); +lean_inc(x_96); +lean_dec(x_87); +x_98 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_98, 0, x_96); +lean_ctor_set(x_98, 1, x_97); +return x_98; +} +} +} +else +{ +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_80, 0); +x_100 = lean_ctor_get(x_80, 1); +lean_inc(x_100); +lean_inc(x_99); +lean_dec(x_80); +lean_inc(x_99); +x_101 = l_Lean_Expr_app___override(x_33, x_99); +x_102 = l_Lean_Expr_mvarId_x21(x_99); +lean_dec(x_99); +lean_inc(x_4); +lean_inc(x_5); +x_103 = lean_array_push(x_5, x_4); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_104 = l_Lean_MVarId_tryClearMany(x_102, x_103, x_14, x_15, x_16, x_17, x_100); +lean_dec(x_103); +if (lean_obj_tag(x_104) == 0) +{ +lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; +x_105 = lean_ctor_get(x_104, 0); +lean_inc(x_105); +x_106 = lean_ctor_get(x_104, 1); +lean_inc(x_106); +lean_dec(x_104); +x_107 = lean_array_push(x_31, x_105); +x_108 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +x_109 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_109, 0, x_101); +lean_ctor_set(x_109, 1, x_39); +x_110 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_110, 0, x_107); +lean_ctor_set(x_110, 1, x_109); +x_111 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_111, 0, x_108); +lean_ctor_set(x_111, 1, x_110); +x_112 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_112, 0, x_111); +x_22 = x_112; +x_23 = x_106; +goto block_28; +} +else +{ +lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; +lean_dec(x_101); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_113 = lean_ctor_get(x_104, 0); +lean_inc(x_113); +x_114 = lean_ctor_get(x_104, 1); +lean_inc(x_114); +if (lean_is_exclusive(x_104)) { + lean_ctor_release(x_104, 0); + lean_ctor_release(x_104, 1); + x_115 = x_104; +} else { + lean_dec_ref(x_104); + x_115 = lean_box(0); +} +if (lean_is_scalar(x_115)) { + x_116 = lean_alloc_ctor(1, 2, 0); +} else { + x_116 = x_115; +} +lean_ctor_set(x_116, 0, x_113); +lean_ctor_set(x_116, 1, x_114); +return x_116; +} +} +} +} +else +{ +lean_object* x_117; lean_object* x_118; lean_object* x_119; uint8_t x_120; +lean_dec(x_36); +lean_dec(x_33); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_11); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +x_117 = lean_ctor_get(x_35, 1); +lean_inc(x_117); +lean_dec(x_35); +x_118 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__1___closed__4; +x_119 = l_Lean_Meta_throwTacticEx___rarg(x_6, x_1, x_118, x_14, x_15, x_16, x_17, x_117); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +x_120 = !lean_is_exclusive(x_119); +if (x_120 == 0) +{ +return x_119; +} +else +{ +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_124; +lean_dec(x_33); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_124 = !lean_is_exclusive(x_35); +if (x_124 == 0) +{ +return x_35; +} +else +{ +lean_object* x_125; lean_object* x_126; lean_object* x_127; +x_125 = lean_ctor_get(x_35, 0); +x_126 = lean_ctor_get(x_35, 1); +lean_inc(x_126); +lean_inc(x_125); +lean_dec(x_35); +x_127 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_127, 0, x_125); +lean_ctor_set(x_127, 1, x_126); +return x_127; +} +} +block_28: +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_22, 0); +lean_inc(x_24); +lean_dec(x_22); +x_25 = lean_ctor_get(x_9, 2); +x_26 = lean_nat_add(x_11, x_25); +lean_dec(x_11); +x_10 = x_24; +x_11 = x_26; +x_12 = lean_box(0); +x_13 = lean_box(0); +x_18 = x_23; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18) { +_start: +{ +lean_object* x_19; uint8_t x_20; +x_19 = lean_ctor_get(x_9, 1); +x_20 = lean_nat_dec_lt(x_11, x_19); +if (x_20 == 0) +{ +lean_object* x_21; +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +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; lean_object* x_23; lean_object* x_29; 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_10, 0); +lean_inc(x_29); +x_30 = lean_ctor_get(x_10, 1); +lean_inc(x_30); +lean_dec(x_10); +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_32, 0); +lean_inc(x_33); +x_34 = lean_ctor_get(x_32, 1); +lean_inc(x_34); +lean_dec(x_32); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_35 = lean_whnf(x_34, x_14, x_15, x_16, x_17, x_18); +if (lean_obj_tag(x_35) == 0) +{ +lean_object* x_36; +x_36 = lean_ctor_get(x_35, 0); +lean_inc(x_36); +if (lean_obj_tag(x_36) == 7) +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; uint8_t x_41; +x_37 = lean_ctor_get(x_35, 1); +lean_inc(x_37); +lean_dec(x_35); +x_38 = lean_ctor_get(x_36, 1); +lean_inc(x_38); +x_39 = lean_ctor_get(x_36, 2); +lean_inc(x_39); +lean_dec(x_36); +x_40 = lean_unsigned_to_nat(1u); +x_41 = lean_nat_dec_lt(x_40, x_7); +if (x_41 == 0) +{ +lean_object* x_42; uint8_t x_43; +lean_inc(x_14); +lean_inc(x_1); +x_42 = l_Lean_Meta_mkFreshExprSyntheticOpaqueMVar(x_38, x_1, x_14, x_15, x_16, x_17, x_37); +x_43 = !lean_is_exclusive(x_42); +if (x_43 == 0) +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_44 = lean_ctor_get(x_42, 0); +x_45 = lean_ctor_get(x_42, 1); +lean_inc(x_44); +x_46 = l_Lean_Expr_app___override(x_33, x_44); +x_47 = l_Lean_Expr_mvarId_x21(x_44); +lean_dec(x_44); +lean_inc(x_3); +lean_inc(x_5); +x_48 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_49 = l_Lean_MVarId_tryClearMany(x_47, x_48, x_14, x_15, x_16, x_17, x_45); +lean_dec(x_48); +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; +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_array_push(x_31, x_50); +x_53 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +lean_ctor_set(x_42, 1, x_39); +lean_ctor_set(x_42, 0, x_46); +x_54 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_54, 0, x_52); +lean_ctor_set(x_54, 1, x_42); +x_55 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_55, 0, x_53); +lean_ctor_set(x_55, 1, x_54); +x_56 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_56, 0, x_55); +x_22 = x_56; +x_23 = x_51; +goto block_28; +} +else +{ +uint8_t x_57; +lean_dec(x_46); +lean_free_object(x_42); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_57 = !lean_is_exclusive(x_49); +if (x_57 == 0) +{ +return x_49; +} +else +{ +lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_58 = lean_ctor_get(x_49, 0); +x_59 = lean_ctor_get(x_49, 1); +lean_inc(x_59); +lean_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; +} +} +} +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; +x_61 = lean_ctor_get(x_42, 0); +x_62 = lean_ctor_get(x_42, 1); +lean_inc(x_62); +lean_inc(x_61); +lean_dec(x_42); +lean_inc(x_61); +x_63 = l_Lean_Expr_app___override(x_33, x_61); +x_64 = l_Lean_Expr_mvarId_x21(x_61); +lean_dec(x_61); +lean_inc(x_3); +lean_inc(x_5); +x_65 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_66 = l_Lean_MVarId_tryClearMany(x_64, x_65, x_14, x_15, x_16, x_17, x_62); +lean_dec(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; +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_array_push(x_31, x_67); +x_70 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +x_71 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_71, 0, x_63); +lean_ctor_set(x_71, 1, x_39); +x_72 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_72, 0, x_69); +lean_ctor_set(x_72, 1, x_71); +x_73 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_73, 0, x_70); +lean_ctor_set(x_73, 1, x_72); +x_74 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_74, 0, x_73); +x_22 = x_74; +x_23 = x_68; +goto block_28; +} +else +{ +lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; +lean_dec(x_63); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_75 = lean_ctor_get(x_66, 0); +lean_inc(x_75); +x_76 = lean_ctor_get(x_66, 1); +lean_inc(x_76); +if (lean_is_exclusive(x_66)) { + lean_ctor_release(x_66, 0); + lean_ctor_release(x_66, 1); + x_77 = x_66; +} else { + lean_dec_ref(x_66); + x_77 = lean_box(0); +} +if (lean_is_scalar(x_77)) { + x_78 = lean_alloc_ctor(1, 2, 0); +} else { + x_78 = x_77; +} +lean_ctor_set(x_78, 0, x_75); +lean_ctor_set(x_78, 1, x_76); +return x_78; +} +} +} +else +{ +lean_object* x_79; lean_object* x_80; uint8_t x_81; +lean_inc(x_29); +lean_inc(x_1); +x_79 = l_Lean_Name_num___override(x_1, x_29); +lean_inc(x_14); +x_80 = l_Lean_Meta_mkFreshExprSyntheticOpaqueMVar(x_38, x_79, x_14, x_15, x_16, x_17, x_37); +x_81 = !lean_is_exclusive(x_80); +if (x_81 == 0) +{ +lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; +x_82 = lean_ctor_get(x_80, 0); +x_83 = lean_ctor_get(x_80, 1); +lean_inc(x_82); +x_84 = l_Lean_Expr_app___override(x_33, x_82); +x_85 = l_Lean_Expr_mvarId_x21(x_82); +lean_dec(x_82); +lean_inc(x_3); +lean_inc(x_5); +x_86 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_87 = l_Lean_MVarId_tryClearMany(x_85, x_86, x_14, x_15, x_16, x_17, x_83); +lean_dec(x_86); +if (lean_obj_tag(x_87) == 0) +{ +lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; +x_88 = lean_ctor_get(x_87, 0); +lean_inc(x_88); +x_89 = lean_ctor_get(x_87, 1); +lean_inc(x_89); +lean_dec(x_87); +x_90 = lean_array_push(x_31, x_88); +x_91 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +lean_ctor_set(x_80, 1, x_39); +lean_ctor_set(x_80, 0, x_84); +x_92 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_92, 0, x_90); +lean_ctor_set(x_92, 1, x_80); +x_93 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_93, 0, x_91); +lean_ctor_set(x_93, 1, x_92); +x_94 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_94, 0, x_93); +x_22 = x_94; +x_23 = x_89; +goto block_28; +} +else +{ +uint8_t x_95; +lean_dec(x_84); +lean_free_object(x_80); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_95 = !lean_is_exclusive(x_87); +if (x_95 == 0) +{ +return x_87; +} +else +{ +lean_object* x_96; lean_object* x_97; lean_object* x_98; +x_96 = lean_ctor_get(x_87, 0); +x_97 = lean_ctor_get(x_87, 1); +lean_inc(x_97); +lean_inc(x_96); +lean_dec(x_87); +x_98 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_98, 0, x_96); +lean_ctor_set(x_98, 1, x_97); +return x_98; +} +} +} +else +{ +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_80, 0); +x_100 = lean_ctor_get(x_80, 1); +lean_inc(x_100); +lean_inc(x_99); +lean_dec(x_80); +lean_inc(x_99); +x_101 = l_Lean_Expr_app___override(x_33, x_99); +x_102 = l_Lean_Expr_mvarId_x21(x_99); +lean_dec(x_99); +lean_inc(x_3); +lean_inc(x_5); +x_103 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_104 = l_Lean_MVarId_tryClearMany(x_102, x_103, x_14, x_15, x_16, x_17, x_100); +lean_dec(x_103); +if (lean_obj_tag(x_104) == 0) +{ +lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; +x_105 = lean_ctor_get(x_104, 0); +lean_inc(x_105); +x_106 = lean_ctor_get(x_104, 1); +lean_inc(x_106); +lean_dec(x_104); +x_107 = lean_array_push(x_31, x_105); +x_108 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +x_109 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_109, 0, x_101); +lean_ctor_set(x_109, 1, x_39); +x_110 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_110, 0, x_107); +lean_ctor_set(x_110, 1, x_109); +x_111 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_111, 0, x_108); +lean_ctor_set(x_111, 1, x_110); +x_112 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_112, 0, x_111); +x_22 = x_112; +x_23 = x_106; +goto block_28; +} +else +{ +lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; +lean_dec(x_101); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_113 = lean_ctor_get(x_104, 0); +lean_inc(x_113); +x_114 = lean_ctor_get(x_104, 1); +lean_inc(x_114); +if (lean_is_exclusive(x_104)) { + lean_ctor_release(x_104, 0); + lean_ctor_release(x_104, 1); + x_115 = x_104; +} else { + lean_dec_ref(x_104); + x_115 = lean_box(0); +} +if (lean_is_scalar(x_115)) { + x_116 = lean_alloc_ctor(1, 2, 0); +} else { + x_116 = x_115; +} +lean_ctor_set(x_116, 0, x_113); +lean_ctor_set(x_116, 1, x_114); +return x_116; +} +} +} +} +else +{ +lean_object* x_117; lean_object* x_118; lean_object* x_119; uint8_t x_120; +lean_dec(x_36); +lean_dec(x_33); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_11); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +x_117 = lean_ctor_get(x_35, 1); +lean_inc(x_117); +lean_dec(x_35); +x_118 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__1___closed__4; +x_119 = l_Lean_Meta_throwTacticEx___rarg(x_6, x_4, x_118, x_14, x_15, x_16, x_17, x_117); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +x_120 = !lean_is_exclusive(x_119); +if (x_120 == 0) +{ +return x_119; +} +else +{ +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_124; +lean_dec(x_33); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_124 = !lean_is_exclusive(x_35); +if (x_124 == 0) +{ +return x_35; +} +else +{ +lean_object* x_125; lean_object* x_126; lean_object* x_127; +x_125 = lean_ctor_get(x_35, 0); +x_126 = lean_ctor_get(x_35, 1); +lean_inc(x_126); +lean_inc(x_125); +lean_dec(x_35); +x_127 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_127, 0, x_125); +lean_ctor_set(x_127, 1, x_126); +return x_127; +} +} +block_28: +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_22, 0); +lean_inc(x_24); +lean_dec(x_22); +x_25 = lean_ctor_get(x_9, 2); +x_26 = lean_nat_add(x_11, x_25); +lean_dec(x_11); +x_10 = x_24; +x_11 = x_26; +x_12 = lean_box(0); +x_13 = lean_box(0); +x_18 = x_23; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18) { +_start: +{ +lean_object* x_19; uint8_t x_20; +x_19 = lean_ctor_get(x_9, 1); +x_20 = lean_nat_dec_lt(x_11, x_19); +if (x_20 == 0) +{ +lean_object* x_21; +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +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; lean_object* x_23; lean_object* x_29; 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_10, 0); +lean_inc(x_29); +x_30 = lean_ctor_get(x_10, 1); +lean_inc(x_30); +lean_dec(x_10); +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_32, 0); +lean_inc(x_33); +x_34 = lean_ctor_get(x_32, 1); +lean_inc(x_34); +lean_dec(x_32); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_35 = lean_whnf(x_34, x_14, x_15, x_16, x_17, x_18); +if (lean_obj_tag(x_35) == 0) +{ +lean_object* x_36; +x_36 = lean_ctor_get(x_35, 0); +lean_inc(x_36); +if (lean_obj_tag(x_36) == 7) +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; uint8_t x_41; +x_37 = lean_ctor_get(x_35, 1); +lean_inc(x_37); +lean_dec(x_35); +x_38 = lean_ctor_get(x_36, 1); +lean_inc(x_38); +x_39 = lean_ctor_get(x_36, 2); +lean_inc(x_39); +lean_dec(x_36); +x_40 = lean_unsigned_to_nat(1u); +x_41 = lean_nat_dec_lt(x_40, x_7); +if (x_41 == 0) +{ +lean_object* x_42; uint8_t x_43; +lean_inc(x_14); +lean_inc(x_1); +x_42 = l_Lean_Meta_mkFreshExprSyntheticOpaqueMVar(x_38, x_1, x_14, x_15, x_16, x_17, x_37); +x_43 = !lean_is_exclusive(x_42); +if (x_43 == 0) +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_44 = lean_ctor_get(x_42, 0); +x_45 = lean_ctor_get(x_42, 1); +lean_inc(x_44); +x_46 = l_Lean_Expr_app___override(x_33, x_44); +x_47 = l_Lean_Expr_mvarId_x21(x_44); +lean_dec(x_44); +lean_inc(x_3); +lean_inc(x_5); +x_48 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_49 = l_Lean_MVarId_tryClearMany(x_47, x_48, x_14, x_15, x_16, x_17, x_45); +lean_dec(x_48); +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; +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_array_push(x_31, x_50); +x_53 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +lean_ctor_set(x_42, 1, x_39); +lean_ctor_set(x_42, 0, x_46); +x_54 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_54, 0, x_52); +lean_ctor_set(x_54, 1, x_42); +x_55 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_55, 0, x_53); +lean_ctor_set(x_55, 1, x_54); +x_56 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_56, 0, x_55); +x_22 = x_56; +x_23 = x_51; +goto block_28; +} +else +{ +uint8_t x_57; +lean_dec(x_46); +lean_free_object(x_42); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_57 = !lean_is_exclusive(x_49); +if (x_57 == 0) +{ +return x_49; +} +else +{ +lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_58 = lean_ctor_get(x_49, 0); +x_59 = lean_ctor_get(x_49, 1); +lean_inc(x_59); +lean_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; +} +} +} +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; +x_61 = lean_ctor_get(x_42, 0); +x_62 = lean_ctor_get(x_42, 1); +lean_inc(x_62); +lean_inc(x_61); +lean_dec(x_42); +lean_inc(x_61); +x_63 = l_Lean_Expr_app___override(x_33, x_61); +x_64 = l_Lean_Expr_mvarId_x21(x_61); +lean_dec(x_61); +lean_inc(x_3); +lean_inc(x_5); +x_65 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_66 = l_Lean_MVarId_tryClearMany(x_64, x_65, x_14, x_15, x_16, x_17, x_62); +lean_dec(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; +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_array_push(x_31, x_67); +x_70 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +x_71 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_71, 0, x_63); +lean_ctor_set(x_71, 1, x_39); +x_72 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_72, 0, x_69); +lean_ctor_set(x_72, 1, x_71); +x_73 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_73, 0, x_70); +lean_ctor_set(x_73, 1, x_72); +x_74 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_74, 0, x_73); +x_22 = x_74; +x_23 = x_68; +goto block_28; +} +else +{ +lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; +lean_dec(x_63); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_75 = lean_ctor_get(x_66, 0); +lean_inc(x_75); +x_76 = lean_ctor_get(x_66, 1); +lean_inc(x_76); +if (lean_is_exclusive(x_66)) { + lean_ctor_release(x_66, 0); + lean_ctor_release(x_66, 1); + x_77 = x_66; +} else { + lean_dec_ref(x_66); + x_77 = lean_box(0); +} +if (lean_is_scalar(x_77)) { + x_78 = lean_alloc_ctor(1, 2, 0); +} else { + x_78 = x_77; +} +lean_ctor_set(x_78, 0, x_75); +lean_ctor_set(x_78, 1, x_76); +return x_78; +} +} +} +else +{ +lean_object* x_79; lean_object* x_80; uint8_t x_81; +lean_inc(x_29); +lean_inc(x_1); +x_79 = l_Lean_Name_num___override(x_1, x_29); +lean_inc(x_14); +x_80 = l_Lean_Meta_mkFreshExprSyntheticOpaqueMVar(x_38, x_79, x_14, x_15, x_16, x_17, x_37); +x_81 = !lean_is_exclusive(x_80); +if (x_81 == 0) +{ +lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; +x_82 = lean_ctor_get(x_80, 0); +x_83 = lean_ctor_get(x_80, 1); +lean_inc(x_82); +x_84 = l_Lean_Expr_app___override(x_33, x_82); +x_85 = l_Lean_Expr_mvarId_x21(x_82); +lean_dec(x_82); +lean_inc(x_3); +lean_inc(x_5); +x_86 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_87 = l_Lean_MVarId_tryClearMany(x_85, x_86, x_14, x_15, x_16, x_17, x_83); +lean_dec(x_86); +if (lean_obj_tag(x_87) == 0) +{ +lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; +x_88 = lean_ctor_get(x_87, 0); +lean_inc(x_88); +x_89 = lean_ctor_get(x_87, 1); +lean_inc(x_89); +lean_dec(x_87); +x_90 = lean_array_push(x_31, x_88); +x_91 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +lean_ctor_set(x_80, 1, x_39); +lean_ctor_set(x_80, 0, x_84); +x_92 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_92, 0, x_90); +lean_ctor_set(x_92, 1, x_80); +x_93 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_93, 0, x_91); +lean_ctor_set(x_93, 1, x_92); +x_94 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_94, 0, x_93); +x_22 = x_94; +x_23 = x_89; +goto block_28; +} +else +{ +uint8_t x_95; +lean_dec(x_84); +lean_free_object(x_80); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_95 = !lean_is_exclusive(x_87); +if (x_95 == 0) +{ +return x_87; +} +else +{ +lean_object* x_96; lean_object* x_97; lean_object* x_98; +x_96 = lean_ctor_get(x_87, 0); +x_97 = lean_ctor_get(x_87, 1); +lean_inc(x_97); +lean_inc(x_96); +lean_dec(x_87); +x_98 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_98, 0, x_96); +lean_ctor_set(x_98, 1, x_97); +return x_98; +} +} +} +else +{ +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_80, 0); +x_100 = lean_ctor_get(x_80, 1); +lean_inc(x_100); +lean_inc(x_99); +lean_dec(x_80); +lean_inc(x_99); +x_101 = l_Lean_Expr_app___override(x_33, x_99); +x_102 = l_Lean_Expr_mvarId_x21(x_99); +lean_dec(x_99); +lean_inc(x_3); +lean_inc(x_5); +x_103 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_104 = l_Lean_MVarId_tryClearMany(x_102, x_103, x_14, x_15, x_16, x_17, x_100); +lean_dec(x_103); +if (lean_obj_tag(x_104) == 0) +{ +lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; +x_105 = lean_ctor_get(x_104, 0); +lean_inc(x_105); +x_106 = lean_ctor_get(x_104, 1); +lean_inc(x_106); +lean_dec(x_104); +x_107 = lean_array_push(x_31, x_105); +x_108 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +x_109 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_109, 0, x_101); +lean_ctor_set(x_109, 1, x_39); +x_110 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_110, 0, x_107); +lean_ctor_set(x_110, 1, x_109); +x_111 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_111, 0, x_108); +lean_ctor_set(x_111, 1, x_110); +x_112 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_112, 0, x_111); +x_22 = x_112; +x_23 = x_106; +goto block_28; +} +else +{ +lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; +lean_dec(x_101); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_113 = lean_ctor_get(x_104, 0); +lean_inc(x_113); +x_114 = lean_ctor_get(x_104, 1); +lean_inc(x_114); +if (lean_is_exclusive(x_104)) { + lean_ctor_release(x_104, 0); + lean_ctor_release(x_104, 1); + x_115 = x_104; +} else { + lean_dec_ref(x_104); + x_115 = lean_box(0); +} +if (lean_is_scalar(x_115)) { + x_116 = lean_alloc_ctor(1, 2, 0); +} else { + x_116 = x_115; +} +lean_ctor_set(x_116, 0, x_113); +lean_ctor_set(x_116, 1, x_114); +return x_116; +} +} +} +} +else +{ +lean_object* x_117; lean_object* x_118; lean_object* x_119; uint8_t x_120; +lean_dec(x_36); +lean_dec(x_33); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_11); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +x_117 = lean_ctor_get(x_35, 1); +lean_inc(x_117); +lean_dec(x_35); +x_118 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__1___closed__4; +x_119 = l_Lean_Meta_throwTacticEx___rarg(x_6, x_4, x_118, x_14, x_15, x_16, x_17, x_117); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +x_120 = !lean_is_exclusive(x_119); +if (x_120 == 0) +{ +return x_119; +} +else +{ +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_124; +lean_dec(x_33); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_124 = !lean_is_exclusive(x_35); +if (x_124 == 0) +{ +return x_35; +} +else +{ +lean_object* x_125; lean_object* x_126; lean_object* x_127; +x_125 = lean_ctor_get(x_35, 0); +x_126 = lean_ctor_get(x_35, 1); +lean_inc(x_126); +lean_inc(x_125); +lean_dec(x_35); +x_127 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_127, 0, x_125); +lean_ctor_set(x_127, 1, x_126); +return x_127; +} +} +block_28: +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_22, 0); +lean_inc(x_24); +lean_dec(x_22); +x_25 = lean_ctor_get(x_9, 2); +x_26 = lean_nat_add(x_11, x_25); +lean_dec(x_11); +x_10 = x_24; +x_11 = x_26; +x_12 = lean_box(0); +x_13 = lean_box(0); +x_18 = x_23; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18) { +_start: +{ +lean_object* x_19; uint8_t x_20; +x_19 = lean_ctor_get(x_9, 1); +x_20 = lean_nat_dec_lt(x_11, x_19); +if (x_20 == 0) +{ +lean_object* x_21; +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +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; lean_object* x_23; lean_object* x_29; 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_10, 0); +lean_inc(x_29); +x_30 = lean_ctor_get(x_10, 1); +lean_inc(x_30); +lean_dec(x_10); +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_32, 0); +lean_inc(x_33); +x_34 = lean_ctor_get(x_32, 1); +lean_inc(x_34); +lean_dec(x_32); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_35 = lean_whnf(x_34, x_14, x_15, x_16, x_17, x_18); +if (lean_obj_tag(x_35) == 0) +{ +lean_object* x_36; +x_36 = lean_ctor_get(x_35, 0); +lean_inc(x_36); +if (lean_obj_tag(x_36) == 7) +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; uint8_t x_41; +x_37 = lean_ctor_get(x_35, 1); +lean_inc(x_37); +lean_dec(x_35); +x_38 = lean_ctor_get(x_36, 1); +lean_inc(x_38); +x_39 = lean_ctor_get(x_36, 2); +lean_inc(x_39); +lean_dec(x_36); +x_40 = lean_unsigned_to_nat(1u); +x_41 = lean_nat_dec_lt(x_40, x_7); +if (x_41 == 0) +{ +lean_object* x_42; uint8_t x_43; +lean_inc(x_14); +lean_inc(x_1); +x_42 = l_Lean_Meta_mkFreshExprSyntheticOpaqueMVar(x_38, x_1, x_14, x_15, x_16, x_17, x_37); +x_43 = !lean_is_exclusive(x_42); +if (x_43 == 0) +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_44 = lean_ctor_get(x_42, 0); +x_45 = lean_ctor_get(x_42, 1); +lean_inc(x_44); +x_46 = l_Lean_Expr_app___override(x_33, x_44); +x_47 = l_Lean_Expr_mvarId_x21(x_44); +lean_dec(x_44); +lean_inc(x_3); +lean_inc(x_5); +x_48 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_49 = l_Lean_MVarId_tryClearMany(x_47, x_48, x_14, x_15, x_16, x_17, x_45); +lean_dec(x_48); +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; +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_array_push(x_31, x_50); +x_53 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +lean_ctor_set(x_42, 1, x_39); +lean_ctor_set(x_42, 0, x_46); +x_54 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_54, 0, x_52); +lean_ctor_set(x_54, 1, x_42); +x_55 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_55, 0, x_53); +lean_ctor_set(x_55, 1, x_54); +x_56 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_56, 0, x_55); +x_22 = x_56; +x_23 = x_51; +goto block_28; +} +else +{ +uint8_t x_57; +lean_dec(x_46); +lean_free_object(x_42); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_57 = !lean_is_exclusive(x_49); +if (x_57 == 0) +{ +return x_49; +} +else +{ +lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_58 = lean_ctor_get(x_49, 0); +x_59 = lean_ctor_get(x_49, 1); +lean_inc(x_59); +lean_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; +} +} +} +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; +x_61 = lean_ctor_get(x_42, 0); +x_62 = lean_ctor_get(x_42, 1); +lean_inc(x_62); +lean_inc(x_61); +lean_dec(x_42); +lean_inc(x_61); +x_63 = l_Lean_Expr_app___override(x_33, x_61); +x_64 = l_Lean_Expr_mvarId_x21(x_61); +lean_dec(x_61); +lean_inc(x_3); +lean_inc(x_5); +x_65 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_66 = l_Lean_MVarId_tryClearMany(x_64, x_65, x_14, x_15, x_16, x_17, x_62); +lean_dec(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; +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_array_push(x_31, x_67); +x_70 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +x_71 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_71, 0, x_63); +lean_ctor_set(x_71, 1, x_39); +x_72 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_72, 0, x_69); +lean_ctor_set(x_72, 1, x_71); +x_73 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_73, 0, x_70); +lean_ctor_set(x_73, 1, x_72); +x_74 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_74, 0, x_73); +x_22 = x_74; +x_23 = x_68; +goto block_28; +} +else +{ +lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; +lean_dec(x_63); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_75 = lean_ctor_get(x_66, 0); +lean_inc(x_75); +x_76 = lean_ctor_get(x_66, 1); +lean_inc(x_76); +if (lean_is_exclusive(x_66)) { + lean_ctor_release(x_66, 0); + lean_ctor_release(x_66, 1); + x_77 = x_66; +} else { + lean_dec_ref(x_66); + x_77 = lean_box(0); +} +if (lean_is_scalar(x_77)) { + x_78 = lean_alloc_ctor(1, 2, 0); +} else { + x_78 = x_77; +} +lean_ctor_set(x_78, 0, x_75); +lean_ctor_set(x_78, 1, x_76); +return x_78; +} +} +} +else +{ +lean_object* x_79; lean_object* x_80; uint8_t x_81; +lean_inc(x_29); +lean_inc(x_1); +x_79 = l_Lean_Name_num___override(x_1, x_29); +lean_inc(x_14); +x_80 = l_Lean_Meta_mkFreshExprSyntheticOpaqueMVar(x_38, x_79, x_14, x_15, x_16, x_17, x_37); +x_81 = !lean_is_exclusive(x_80); +if (x_81 == 0) +{ +lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; +x_82 = lean_ctor_get(x_80, 0); +x_83 = lean_ctor_get(x_80, 1); +lean_inc(x_82); +x_84 = l_Lean_Expr_app___override(x_33, x_82); +x_85 = l_Lean_Expr_mvarId_x21(x_82); +lean_dec(x_82); +lean_inc(x_3); +lean_inc(x_5); +x_86 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_87 = l_Lean_MVarId_tryClearMany(x_85, x_86, x_14, x_15, x_16, x_17, x_83); +lean_dec(x_86); +if (lean_obj_tag(x_87) == 0) +{ +lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; +x_88 = lean_ctor_get(x_87, 0); +lean_inc(x_88); +x_89 = lean_ctor_get(x_87, 1); +lean_inc(x_89); +lean_dec(x_87); +x_90 = lean_array_push(x_31, x_88); +x_91 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +lean_ctor_set(x_80, 1, x_39); +lean_ctor_set(x_80, 0, x_84); +x_92 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_92, 0, x_90); +lean_ctor_set(x_92, 1, x_80); +x_93 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_93, 0, x_91); +lean_ctor_set(x_93, 1, x_92); +x_94 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_94, 0, x_93); +x_22 = x_94; +x_23 = x_89; +goto block_28; +} +else +{ +uint8_t x_95; +lean_dec(x_84); +lean_free_object(x_80); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_95 = !lean_is_exclusive(x_87); +if (x_95 == 0) +{ +return x_87; +} +else +{ +lean_object* x_96; lean_object* x_97; lean_object* x_98; +x_96 = lean_ctor_get(x_87, 0); +x_97 = lean_ctor_get(x_87, 1); +lean_inc(x_97); +lean_inc(x_96); +lean_dec(x_87); +x_98 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_98, 0, x_96); +lean_ctor_set(x_98, 1, x_97); +return x_98; +} +} +} +else +{ +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_80, 0); +x_100 = lean_ctor_get(x_80, 1); +lean_inc(x_100); +lean_inc(x_99); +lean_dec(x_80); +lean_inc(x_99); +x_101 = l_Lean_Expr_app___override(x_33, x_99); +x_102 = l_Lean_Expr_mvarId_x21(x_99); +lean_dec(x_99); +lean_inc(x_3); +lean_inc(x_5); +x_103 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_104 = l_Lean_MVarId_tryClearMany(x_102, x_103, x_14, x_15, x_16, x_17, x_100); +lean_dec(x_103); +if (lean_obj_tag(x_104) == 0) +{ +lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; +x_105 = lean_ctor_get(x_104, 0); +lean_inc(x_105); +x_106 = lean_ctor_get(x_104, 1); +lean_inc(x_106); +lean_dec(x_104); +x_107 = lean_array_push(x_31, x_105); +x_108 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +x_109 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_109, 0, x_101); +lean_ctor_set(x_109, 1, x_39); +x_110 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_110, 0, x_107); +lean_ctor_set(x_110, 1, x_109); +x_111 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_111, 0, x_108); +lean_ctor_set(x_111, 1, x_110); +x_112 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_112, 0, x_111); +x_22 = x_112; +x_23 = x_106; +goto block_28; +} +else +{ +lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; +lean_dec(x_101); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_113 = lean_ctor_get(x_104, 0); +lean_inc(x_113); +x_114 = lean_ctor_get(x_104, 1); +lean_inc(x_114); +if (lean_is_exclusive(x_104)) { + lean_ctor_release(x_104, 0); + lean_ctor_release(x_104, 1); + x_115 = x_104; +} else { + lean_dec_ref(x_104); + x_115 = lean_box(0); +} +if (lean_is_scalar(x_115)) { + x_116 = lean_alloc_ctor(1, 2, 0); +} else { + x_116 = x_115; +} +lean_ctor_set(x_116, 0, x_113); +lean_ctor_set(x_116, 1, x_114); +return x_116; +} +} +} +} +else +{ +lean_object* x_117; lean_object* x_118; lean_object* x_119; uint8_t x_120; +lean_dec(x_36); +lean_dec(x_33); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_11); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +x_117 = lean_ctor_get(x_35, 1); +lean_inc(x_117); +lean_dec(x_35); +x_118 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__1___closed__4; +x_119 = l_Lean_Meta_throwTacticEx___rarg(x_6, x_4, x_118, x_14, x_15, x_16, x_17, x_117); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +x_120 = !lean_is_exclusive(x_119); +if (x_120 == 0) +{ +return x_119; +} +else +{ +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_124; +lean_dec(x_33); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_124 = !lean_is_exclusive(x_35); +if (x_124 == 0) +{ +return x_35; +} +else +{ +lean_object* x_125; lean_object* x_126; lean_object* x_127; +x_125 = lean_ctor_get(x_35, 0); +x_126 = lean_ctor_get(x_35, 1); +lean_inc(x_126); +lean_inc(x_125); +lean_dec(x_35); +x_127 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_127, 0, x_125); +lean_ctor_set(x_127, 1, x_126); +return x_127; +} +} +block_28: +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_22, 0); +lean_inc(x_24); +lean_dec(x_22); +x_25 = lean_ctor_get(x_9, 2); +x_26 = lean_nat_add(x_11, x_25); +lean_dec(x_11); +x_10 = x_24; +x_11 = x_26; +x_12 = lean_box(0); +x_13 = lean_box(0); +x_18 = x_23; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18) { +_start: +{ +lean_object* x_19; uint8_t x_20; +x_19 = lean_ctor_get(x_9, 1); +x_20 = lean_nat_dec_lt(x_11, x_19); +if (x_20 == 0) +{ +lean_object* x_21; +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +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; lean_object* x_23; lean_object* x_29; 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_10, 0); +lean_inc(x_29); +x_30 = lean_ctor_get(x_10, 1); +lean_inc(x_30); +lean_dec(x_10); +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_32, 0); +lean_inc(x_33); +x_34 = lean_ctor_get(x_32, 1); +lean_inc(x_34); +lean_dec(x_32); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_35 = lean_whnf(x_34, x_14, x_15, x_16, x_17, x_18); +if (lean_obj_tag(x_35) == 0) +{ +lean_object* x_36; +x_36 = lean_ctor_get(x_35, 0); +lean_inc(x_36); +if (lean_obj_tag(x_36) == 7) +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; uint8_t x_41; +x_37 = lean_ctor_get(x_35, 1); +lean_inc(x_37); +lean_dec(x_35); +x_38 = lean_ctor_get(x_36, 1); +lean_inc(x_38); +x_39 = lean_ctor_get(x_36, 2); +lean_inc(x_39); +lean_dec(x_36); +x_40 = lean_unsigned_to_nat(1u); +x_41 = lean_nat_dec_lt(x_40, x_7); +if (x_41 == 0) +{ +lean_object* x_42; uint8_t x_43; +lean_inc(x_14); +lean_inc(x_1); +x_42 = l_Lean_Meta_mkFreshExprSyntheticOpaqueMVar(x_38, x_1, x_14, x_15, x_16, x_17, x_37); +x_43 = !lean_is_exclusive(x_42); +if (x_43 == 0) +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_44 = lean_ctor_get(x_42, 0); +x_45 = lean_ctor_get(x_42, 1); +lean_inc(x_44); +x_46 = l_Lean_Expr_app___override(x_33, x_44); +x_47 = l_Lean_Expr_mvarId_x21(x_44); +lean_dec(x_44); +lean_inc(x_3); +lean_inc(x_5); +x_48 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_49 = l_Lean_MVarId_tryClearMany(x_47, x_48, x_14, x_15, x_16, x_17, x_45); +lean_dec(x_48); +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; +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_array_push(x_31, x_50); +x_53 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +lean_ctor_set(x_42, 1, x_39); +lean_ctor_set(x_42, 0, x_46); +x_54 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_54, 0, x_52); +lean_ctor_set(x_54, 1, x_42); +x_55 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_55, 0, x_53); +lean_ctor_set(x_55, 1, x_54); +x_56 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_56, 0, x_55); +x_22 = x_56; +x_23 = x_51; +goto block_28; +} +else +{ +uint8_t x_57; +lean_dec(x_46); +lean_free_object(x_42); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_57 = !lean_is_exclusive(x_49); +if (x_57 == 0) +{ +return x_49; +} +else +{ +lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_58 = lean_ctor_get(x_49, 0); +x_59 = lean_ctor_get(x_49, 1); +lean_inc(x_59); +lean_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; +} +} +} +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; +x_61 = lean_ctor_get(x_42, 0); +x_62 = lean_ctor_get(x_42, 1); +lean_inc(x_62); +lean_inc(x_61); +lean_dec(x_42); +lean_inc(x_61); +x_63 = l_Lean_Expr_app___override(x_33, x_61); +x_64 = l_Lean_Expr_mvarId_x21(x_61); +lean_dec(x_61); +lean_inc(x_3); +lean_inc(x_5); +x_65 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_66 = l_Lean_MVarId_tryClearMany(x_64, x_65, x_14, x_15, x_16, x_17, x_62); +lean_dec(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; +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_array_push(x_31, x_67); +x_70 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +x_71 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_71, 0, x_63); +lean_ctor_set(x_71, 1, x_39); +x_72 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_72, 0, x_69); +lean_ctor_set(x_72, 1, x_71); +x_73 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_73, 0, x_70); +lean_ctor_set(x_73, 1, x_72); +x_74 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_74, 0, x_73); +x_22 = x_74; +x_23 = x_68; +goto block_28; +} +else +{ +lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; +lean_dec(x_63); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_75 = lean_ctor_get(x_66, 0); +lean_inc(x_75); +x_76 = lean_ctor_get(x_66, 1); +lean_inc(x_76); +if (lean_is_exclusive(x_66)) { + lean_ctor_release(x_66, 0); + lean_ctor_release(x_66, 1); + x_77 = x_66; +} else { + lean_dec_ref(x_66); + x_77 = lean_box(0); +} +if (lean_is_scalar(x_77)) { + x_78 = lean_alloc_ctor(1, 2, 0); +} else { + x_78 = x_77; +} +lean_ctor_set(x_78, 0, x_75); +lean_ctor_set(x_78, 1, x_76); +return x_78; +} +} +} +else +{ +lean_object* x_79; lean_object* x_80; uint8_t x_81; +lean_inc(x_29); +lean_inc(x_1); +x_79 = l_Lean_Name_num___override(x_1, x_29); +lean_inc(x_14); +x_80 = l_Lean_Meta_mkFreshExprSyntheticOpaqueMVar(x_38, x_79, x_14, x_15, x_16, x_17, x_37); +x_81 = !lean_is_exclusive(x_80); +if (x_81 == 0) +{ +lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; +x_82 = lean_ctor_get(x_80, 0); +x_83 = lean_ctor_get(x_80, 1); +lean_inc(x_82); +x_84 = l_Lean_Expr_app___override(x_33, x_82); +x_85 = l_Lean_Expr_mvarId_x21(x_82); +lean_dec(x_82); +lean_inc(x_3); +lean_inc(x_5); +x_86 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_87 = l_Lean_MVarId_tryClearMany(x_85, x_86, x_14, x_15, x_16, x_17, x_83); +lean_dec(x_86); +if (lean_obj_tag(x_87) == 0) +{ +lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; +x_88 = lean_ctor_get(x_87, 0); +lean_inc(x_88); +x_89 = lean_ctor_get(x_87, 1); +lean_inc(x_89); +lean_dec(x_87); +x_90 = lean_array_push(x_31, x_88); +x_91 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +lean_ctor_set(x_80, 1, x_39); +lean_ctor_set(x_80, 0, x_84); +x_92 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_92, 0, x_90); +lean_ctor_set(x_92, 1, x_80); +x_93 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_93, 0, x_91); +lean_ctor_set(x_93, 1, x_92); +x_94 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_94, 0, x_93); +x_22 = x_94; +x_23 = x_89; +goto block_28; +} +else +{ +uint8_t x_95; +lean_dec(x_84); +lean_free_object(x_80); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_95 = !lean_is_exclusive(x_87); +if (x_95 == 0) +{ +return x_87; +} +else +{ +lean_object* x_96; lean_object* x_97; lean_object* x_98; +x_96 = lean_ctor_get(x_87, 0); +x_97 = lean_ctor_get(x_87, 1); +lean_inc(x_97); +lean_inc(x_96); +lean_dec(x_87); +x_98 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_98, 0, x_96); +lean_ctor_set(x_98, 1, x_97); +return x_98; +} +} +} +else +{ +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_80, 0); +x_100 = lean_ctor_get(x_80, 1); +lean_inc(x_100); +lean_inc(x_99); +lean_dec(x_80); +lean_inc(x_99); +x_101 = l_Lean_Expr_app___override(x_33, x_99); +x_102 = l_Lean_Expr_mvarId_x21(x_99); +lean_dec(x_99); +lean_inc(x_3); +lean_inc(x_5); +x_103 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_104 = l_Lean_MVarId_tryClearMany(x_102, x_103, x_14, x_15, x_16, x_17, x_100); +lean_dec(x_103); +if (lean_obj_tag(x_104) == 0) +{ +lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; +x_105 = lean_ctor_get(x_104, 0); +lean_inc(x_105); +x_106 = lean_ctor_get(x_104, 1); +lean_inc(x_106); +lean_dec(x_104); +x_107 = lean_array_push(x_31, x_105); +x_108 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +x_109 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_109, 0, x_101); +lean_ctor_set(x_109, 1, x_39); +x_110 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_110, 0, x_107); +lean_ctor_set(x_110, 1, x_109); +x_111 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_111, 0, x_108); +lean_ctor_set(x_111, 1, x_110); +x_112 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_112, 0, x_111); +x_22 = x_112; +x_23 = x_106; +goto block_28; +} +else +{ +lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; +lean_dec(x_101); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_113 = lean_ctor_get(x_104, 0); +lean_inc(x_113); +x_114 = lean_ctor_get(x_104, 1); +lean_inc(x_114); +if (lean_is_exclusive(x_104)) { + lean_ctor_release(x_104, 0); + lean_ctor_release(x_104, 1); + x_115 = x_104; +} else { + lean_dec_ref(x_104); + x_115 = lean_box(0); +} +if (lean_is_scalar(x_115)) { + x_116 = lean_alloc_ctor(1, 2, 0); +} else { + x_116 = x_115; +} +lean_ctor_set(x_116, 0, x_113); +lean_ctor_set(x_116, 1, x_114); +return x_116; +} +} +} +} +else +{ +lean_object* x_117; lean_object* x_118; lean_object* x_119; uint8_t x_120; +lean_dec(x_36); +lean_dec(x_33); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_11); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +x_117 = lean_ctor_get(x_35, 1); +lean_inc(x_117); +lean_dec(x_35); +x_118 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__1___closed__4; +x_119 = l_Lean_Meta_throwTacticEx___rarg(x_6, x_4, x_118, x_14, x_15, x_16, x_17, x_117); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +x_120 = !lean_is_exclusive(x_119); +if (x_120 == 0) +{ +return x_119; +} +else +{ +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_124; +lean_dec(x_33); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_124 = !lean_is_exclusive(x_35); +if (x_124 == 0) +{ +return x_35; +} +else +{ +lean_object* x_125; lean_object* x_126; lean_object* x_127; +x_125 = lean_ctor_get(x_35, 0); +x_126 = lean_ctor_get(x_35, 1); +lean_inc(x_126); +lean_inc(x_125); +lean_dec(x_35); +x_127 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_127, 0, x_125); +lean_ctor_set(x_127, 1, x_126); +return x_127; +} +} +block_28: +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_22, 0); +lean_inc(x_24); +lean_dec(x_22); +x_25 = lean_ctor_get(x_9, 2); +x_26 = lean_nat_add(x_11, x_25); +lean_dec(x_11); +x_10 = x_24; +x_11 = x_26; +x_12 = lean_box(0); +x_13 = lean_box(0); +x_18 = x_23; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18) { +_start: +{ +lean_object* x_19; uint8_t x_20; +x_19 = lean_ctor_get(x_9, 1); +x_20 = lean_nat_dec_lt(x_11, x_19); +if (x_20 == 0) +{ +lean_object* x_21; +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +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; lean_object* x_23; lean_object* x_29; 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_10, 0); +lean_inc(x_29); +x_30 = lean_ctor_get(x_10, 1); +lean_inc(x_30); +lean_dec(x_10); +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_32, 0); +lean_inc(x_33); +x_34 = lean_ctor_get(x_32, 1); +lean_inc(x_34); +lean_dec(x_32); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_35 = lean_whnf(x_34, x_14, x_15, x_16, x_17, x_18); +if (lean_obj_tag(x_35) == 0) +{ +lean_object* x_36; +x_36 = lean_ctor_get(x_35, 0); +lean_inc(x_36); +if (lean_obj_tag(x_36) == 7) +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; uint8_t x_41; +x_37 = lean_ctor_get(x_35, 1); +lean_inc(x_37); +lean_dec(x_35); +x_38 = lean_ctor_get(x_36, 1); +lean_inc(x_38); +x_39 = lean_ctor_get(x_36, 2); +lean_inc(x_39); +lean_dec(x_36); +x_40 = lean_unsigned_to_nat(1u); +x_41 = lean_nat_dec_lt(x_40, x_7); +if (x_41 == 0) +{ +lean_object* x_42; uint8_t x_43; +lean_inc(x_14); +lean_inc(x_1); +x_42 = l_Lean_Meta_mkFreshExprSyntheticOpaqueMVar(x_38, x_1, x_14, x_15, x_16, x_17, x_37); +x_43 = !lean_is_exclusive(x_42); +if (x_43 == 0) +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_44 = lean_ctor_get(x_42, 0); +x_45 = lean_ctor_get(x_42, 1); +lean_inc(x_44); +x_46 = l_Lean_Expr_app___override(x_33, x_44); +x_47 = l_Lean_Expr_mvarId_x21(x_44); +lean_dec(x_44); +lean_inc(x_3); +lean_inc(x_5); +x_48 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_49 = l_Lean_MVarId_tryClearMany(x_47, x_48, x_14, x_15, x_16, x_17, x_45); +lean_dec(x_48); +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; +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_array_push(x_31, x_50); +x_53 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +lean_ctor_set(x_42, 1, x_39); +lean_ctor_set(x_42, 0, x_46); +x_54 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_54, 0, x_52); +lean_ctor_set(x_54, 1, x_42); +x_55 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_55, 0, x_53); +lean_ctor_set(x_55, 1, x_54); +x_56 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_56, 0, x_55); +x_22 = x_56; +x_23 = x_51; +goto block_28; +} +else +{ +uint8_t x_57; +lean_dec(x_46); +lean_free_object(x_42); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_57 = !lean_is_exclusive(x_49); +if (x_57 == 0) +{ +return x_49; +} +else +{ +lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_58 = lean_ctor_get(x_49, 0); +x_59 = lean_ctor_get(x_49, 1); +lean_inc(x_59); +lean_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; +} +} +} +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; +x_61 = lean_ctor_get(x_42, 0); +x_62 = lean_ctor_get(x_42, 1); +lean_inc(x_62); +lean_inc(x_61); +lean_dec(x_42); +lean_inc(x_61); +x_63 = l_Lean_Expr_app___override(x_33, x_61); +x_64 = l_Lean_Expr_mvarId_x21(x_61); +lean_dec(x_61); +lean_inc(x_3); +lean_inc(x_5); +x_65 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_66 = l_Lean_MVarId_tryClearMany(x_64, x_65, x_14, x_15, x_16, x_17, x_62); +lean_dec(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; +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_array_push(x_31, x_67); +x_70 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +x_71 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_71, 0, x_63); +lean_ctor_set(x_71, 1, x_39); +x_72 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_72, 0, x_69); +lean_ctor_set(x_72, 1, x_71); +x_73 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_73, 0, x_70); +lean_ctor_set(x_73, 1, x_72); +x_74 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_74, 0, x_73); +x_22 = x_74; +x_23 = x_68; +goto block_28; +} +else +{ +lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; +lean_dec(x_63); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_75 = lean_ctor_get(x_66, 0); +lean_inc(x_75); +x_76 = lean_ctor_get(x_66, 1); +lean_inc(x_76); +if (lean_is_exclusive(x_66)) { + lean_ctor_release(x_66, 0); + lean_ctor_release(x_66, 1); + x_77 = x_66; +} else { + lean_dec_ref(x_66); + x_77 = lean_box(0); +} +if (lean_is_scalar(x_77)) { + x_78 = lean_alloc_ctor(1, 2, 0); +} else { + x_78 = x_77; +} +lean_ctor_set(x_78, 0, x_75); +lean_ctor_set(x_78, 1, x_76); +return x_78; +} +} +} +else +{ +lean_object* x_79; lean_object* x_80; uint8_t x_81; +lean_inc(x_29); +lean_inc(x_1); +x_79 = l_Lean_Name_num___override(x_1, x_29); +lean_inc(x_14); +x_80 = l_Lean_Meta_mkFreshExprSyntheticOpaqueMVar(x_38, x_79, x_14, x_15, x_16, x_17, x_37); +x_81 = !lean_is_exclusive(x_80); +if (x_81 == 0) +{ +lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; +x_82 = lean_ctor_get(x_80, 0); +x_83 = lean_ctor_get(x_80, 1); +lean_inc(x_82); +x_84 = l_Lean_Expr_app___override(x_33, x_82); +x_85 = l_Lean_Expr_mvarId_x21(x_82); +lean_dec(x_82); +lean_inc(x_3); +lean_inc(x_5); +x_86 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_87 = l_Lean_MVarId_tryClearMany(x_85, x_86, x_14, x_15, x_16, x_17, x_83); +lean_dec(x_86); +if (lean_obj_tag(x_87) == 0) +{ +lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; +x_88 = lean_ctor_get(x_87, 0); +lean_inc(x_88); +x_89 = lean_ctor_get(x_87, 1); +lean_inc(x_89); +lean_dec(x_87); +x_90 = lean_array_push(x_31, x_88); +x_91 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +lean_ctor_set(x_80, 1, x_39); +lean_ctor_set(x_80, 0, x_84); +x_92 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_92, 0, x_90); +lean_ctor_set(x_92, 1, x_80); +x_93 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_93, 0, x_91); +lean_ctor_set(x_93, 1, x_92); +x_94 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_94, 0, x_93); +x_22 = x_94; +x_23 = x_89; +goto block_28; +} +else +{ +uint8_t x_95; +lean_dec(x_84); +lean_free_object(x_80); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_95 = !lean_is_exclusive(x_87); +if (x_95 == 0) +{ +return x_87; +} +else +{ +lean_object* x_96; lean_object* x_97; lean_object* x_98; +x_96 = lean_ctor_get(x_87, 0); +x_97 = lean_ctor_get(x_87, 1); +lean_inc(x_97); +lean_inc(x_96); +lean_dec(x_87); +x_98 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_98, 0, x_96); +lean_ctor_set(x_98, 1, x_97); +return x_98; +} +} +} +else +{ +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_80, 0); +x_100 = lean_ctor_get(x_80, 1); +lean_inc(x_100); +lean_inc(x_99); +lean_dec(x_80); +lean_inc(x_99); +x_101 = l_Lean_Expr_app___override(x_33, x_99); +x_102 = l_Lean_Expr_mvarId_x21(x_99); +lean_dec(x_99); +lean_inc(x_3); +lean_inc(x_5); +x_103 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_104 = l_Lean_MVarId_tryClearMany(x_102, x_103, x_14, x_15, x_16, x_17, x_100); +lean_dec(x_103); +if (lean_obj_tag(x_104) == 0) +{ +lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; +x_105 = lean_ctor_get(x_104, 0); +lean_inc(x_105); +x_106 = lean_ctor_get(x_104, 1); +lean_inc(x_106); +lean_dec(x_104); +x_107 = lean_array_push(x_31, x_105); +x_108 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +x_109 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_109, 0, x_101); +lean_ctor_set(x_109, 1, x_39); +x_110 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_110, 0, x_107); +lean_ctor_set(x_110, 1, x_109); +x_111 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_111, 0, x_108); +lean_ctor_set(x_111, 1, x_110); +x_112 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_112, 0, x_111); +x_22 = x_112; +x_23 = x_106; +goto block_28; +} +else +{ +lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; +lean_dec(x_101); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_113 = lean_ctor_get(x_104, 0); +lean_inc(x_113); +x_114 = lean_ctor_get(x_104, 1); +lean_inc(x_114); +if (lean_is_exclusive(x_104)) { + lean_ctor_release(x_104, 0); + lean_ctor_release(x_104, 1); + x_115 = x_104; +} else { + lean_dec_ref(x_104); + x_115 = lean_box(0); +} +if (lean_is_scalar(x_115)) { + x_116 = lean_alloc_ctor(1, 2, 0); +} else { + x_116 = x_115; +} +lean_ctor_set(x_116, 0, x_113); +lean_ctor_set(x_116, 1, x_114); +return x_116; +} +} +} +} +else +{ +lean_object* x_117; lean_object* x_118; lean_object* x_119; uint8_t x_120; +lean_dec(x_36); +lean_dec(x_33); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_11); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +x_117 = lean_ctor_get(x_35, 1); +lean_inc(x_117); +lean_dec(x_35); +x_118 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__1___closed__4; +x_119 = l_Lean_Meta_throwTacticEx___rarg(x_6, x_4, x_118, x_14, x_15, x_16, x_17, x_117); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +x_120 = !lean_is_exclusive(x_119); +if (x_120 == 0) +{ +return x_119; +} +else +{ +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_124; +lean_dec(x_33); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_124 = !lean_is_exclusive(x_35); +if (x_124 == 0) +{ +return x_35; +} +else +{ +lean_object* x_125; lean_object* x_126; lean_object* x_127; +x_125 = lean_ctor_get(x_35, 0); +x_126 = lean_ctor_get(x_35, 1); +lean_inc(x_126); +lean_inc(x_125); +lean_dec(x_35); +x_127 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_127, 0, x_125); +lean_ctor_set(x_127, 1, x_126); +return x_127; +} +} +block_28: +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_22, 0); +lean_inc(x_24); +lean_dec(x_22); +x_25 = lean_ctor_get(x_9, 2); +x_26 = lean_nat_add(x_11, x_25); +lean_dec(x_11); +x_10 = x_24; +x_11 = x_26; +x_12 = lean_box(0); +x_13 = lean_box(0); +x_18 = x_23; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__8(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18) { +_start: +{ +lean_object* x_19; uint8_t x_20; +x_19 = lean_ctor_get(x_9, 1); +x_20 = lean_nat_dec_lt(x_11, x_19); +if (x_20 == 0) +{ +lean_object* x_21; +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +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; lean_object* x_23; lean_object* x_29; 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_10, 0); +lean_inc(x_29); +x_30 = lean_ctor_get(x_10, 1); +lean_inc(x_30); +lean_dec(x_10); +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_32, 0); +lean_inc(x_33); +x_34 = lean_ctor_get(x_32, 1); +lean_inc(x_34); +lean_dec(x_32); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_35 = lean_whnf(x_34, x_14, x_15, x_16, x_17, x_18); +if (lean_obj_tag(x_35) == 0) +{ +lean_object* x_36; +x_36 = lean_ctor_get(x_35, 0); +lean_inc(x_36); +if (lean_obj_tag(x_36) == 7) +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; uint8_t x_41; +x_37 = lean_ctor_get(x_35, 1); +lean_inc(x_37); +lean_dec(x_35); +x_38 = lean_ctor_get(x_36, 1); +lean_inc(x_38); +x_39 = lean_ctor_get(x_36, 2); +lean_inc(x_39); +lean_dec(x_36); +x_40 = lean_unsigned_to_nat(1u); +x_41 = lean_nat_dec_lt(x_40, x_7); +if (x_41 == 0) +{ +lean_object* x_42; uint8_t x_43; +lean_inc(x_14); +lean_inc(x_1); +x_42 = l_Lean_Meta_mkFreshExprSyntheticOpaqueMVar(x_38, x_1, x_14, x_15, x_16, x_17, x_37); +x_43 = !lean_is_exclusive(x_42); +if (x_43 == 0) +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_44 = lean_ctor_get(x_42, 0); +x_45 = lean_ctor_get(x_42, 1); +lean_inc(x_44); +x_46 = l_Lean_Expr_app___override(x_33, x_44); +x_47 = l_Lean_Expr_mvarId_x21(x_44); +lean_dec(x_44); +lean_inc(x_3); +lean_inc(x_5); +x_48 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_49 = l_Lean_MVarId_tryClearMany(x_47, x_48, x_14, x_15, x_16, x_17, x_45); +lean_dec(x_48); +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; +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_array_push(x_31, x_50); +x_53 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +lean_ctor_set(x_42, 1, x_39); +lean_ctor_set(x_42, 0, x_46); +x_54 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_54, 0, x_52); +lean_ctor_set(x_54, 1, x_42); +x_55 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_55, 0, x_53); +lean_ctor_set(x_55, 1, x_54); +x_56 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_56, 0, x_55); +x_22 = x_56; +x_23 = x_51; +goto block_28; +} +else +{ +uint8_t x_57; +lean_dec(x_46); +lean_free_object(x_42); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_57 = !lean_is_exclusive(x_49); +if (x_57 == 0) +{ +return x_49; +} +else +{ +lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_58 = lean_ctor_get(x_49, 0); +x_59 = lean_ctor_get(x_49, 1); +lean_inc(x_59); +lean_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; +} +} +} +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; +x_61 = lean_ctor_get(x_42, 0); +x_62 = lean_ctor_get(x_42, 1); +lean_inc(x_62); +lean_inc(x_61); +lean_dec(x_42); +lean_inc(x_61); +x_63 = l_Lean_Expr_app___override(x_33, x_61); +x_64 = l_Lean_Expr_mvarId_x21(x_61); +lean_dec(x_61); +lean_inc(x_3); +lean_inc(x_5); +x_65 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_66 = l_Lean_MVarId_tryClearMany(x_64, x_65, x_14, x_15, x_16, x_17, x_62); +lean_dec(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; +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_array_push(x_31, x_67); +x_70 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +x_71 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_71, 0, x_63); +lean_ctor_set(x_71, 1, x_39); +x_72 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_72, 0, x_69); +lean_ctor_set(x_72, 1, x_71); +x_73 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_73, 0, x_70); +lean_ctor_set(x_73, 1, x_72); +x_74 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_74, 0, x_73); +x_22 = x_74; +x_23 = x_68; +goto block_28; +} +else +{ +lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; +lean_dec(x_63); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_75 = lean_ctor_get(x_66, 0); +lean_inc(x_75); +x_76 = lean_ctor_get(x_66, 1); +lean_inc(x_76); +if (lean_is_exclusive(x_66)) { + lean_ctor_release(x_66, 0); + lean_ctor_release(x_66, 1); + x_77 = x_66; +} else { + lean_dec_ref(x_66); + x_77 = lean_box(0); +} +if (lean_is_scalar(x_77)) { + x_78 = lean_alloc_ctor(1, 2, 0); +} else { + x_78 = x_77; +} +lean_ctor_set(x_78, 0, x_75); +lean_ctor_set(x_78, 1, x_76); +return x_78; +} +} +} +else +{ +lean_object* x_79; lean_object* x_80; uint8_t x_81; +lean_inc(x_29); +lean_inc(x_1); +x_79 = l_Lean_Name_num___override(x_1, x_29); +lean_inc(x_14); +x_80 = l_Lean_Meta_mkFreshExprSyntheticOpaqueMVar(x_38, x_79, x_14, x_15, x_16, x_17, x_37); +x_81 = !lean_is_exclusive(x_80); +if (x_81 == 0) +{ +lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; +x_82 = lean_ctor_get(x_80, 0); +x_83 = lean_ctor_get(x_80, 1); +lean_inc(x_82); +x_84 = l_Lean_Expr_app___override(x_33, x_82); +x_85 = l_Lean_Expr_mvarId_x21(x_82); +lean_dec(x_82); +lean_inc(x_3); +lean_inc(x_5); +x_86 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_87 = l_Lean_MVarId_tryClearMany(x_85, x_86, x_14, x_15, x_16, x_17, x_83); +lean_dec(x_86); +if (lean_obj_tag(x_87) == 0) +{ +lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; +x_88 = lean_ctor_get(x_87, 0); +lean_inc(x_88); +x_89 = lean_ctor_get(x_87, 1); +lean_inc(x_89); +lean_dec(x_87); +x_90 = lean_array_push(x_31, x_88); +x_91 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +lean_ctor_set(x_80, 1, x_39); +lean_ctor_set(x_80, 0, x_84); +x_92 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_92, 0, x_90); +lean_ctor_set(x_92, 1, x_80); +x_93 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_93, 0, x_91); +lean_ctor_set(x_93, 1, x_92); +x_94 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_94, 0, x_93); +x_22 = x_94; +x_23 = x_89; +goto block_28; +} +else +{ +uint8_t x_95; +lean_dec(x_84); +lean_free_object(x_80); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_95 = !lean_is_exclusive(x_87); +if (x_95 == 0) +{ +return x_87; +} +else +{ +lean_object* x_96; lean_object* x_97; lean_object* x_98; +x_96 = lean_ctor_get(x_87, 0); +x_97 = lean_ctor_get(x_87, 1); +lean_inc(x_97); +lean_inc(x_96); +lean_dec(x_87); +x_98 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_98, 0, x_96); +lean_ctor_set(x_98, 1, x_97); +return x_98; +} +} +} +else +{ +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_80, 0); +x_100 = lean_ctor_get(x_80, 1); +lean_inc(x_100); +lean_inc(x_99); +lean_dec(x_80); +lean_inc(x_99); +x_101 = l_Lean_Expr_app___override(x_33, x_99); +x_102 = l_Lean_Expr_mvarId_x21(x_99); +lean_dec(x_99); +lean_inc(x_3); +lean_inc(x_5); +x_103 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_104 = l_Lean_MVarId_tryClearMany(x_102, x_103, x_14, x_15, x_16, x_17, x_100); +lean_dec(x_103); +if (lean_obj_tag(x_104) == 0) +{ +lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; +x_105 = lean_ctor_get(x_104, 0); +lean_inc(x_105); +x_106 = lean_ctor_get(x_104, 1); +lean_inc(x_106); +lean_dec(x_104); +x_107 = lean_array_push(x_31, x_105); +x_108 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +x_109 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_109, 0, x_101); +lean_ctor_set(x_109, 1, x_39); +x_110 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_110, 0, x_107); +lean_ctor_set(x_110, 1, x_109); +x_111 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_111, 0, x_108); +lean_ctor_set(x_111, 1, x_110); +x_112 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_112, 0, x_111); +x_22 = x_112; +x_23 = x_106; +goto block_28; +} +else +{ +lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; +lean_dec(x_101); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_113 = lean_ctor_get(x_104, 0); +lean_inc(x_113); +x_114 = lean_ctor_get(x_104, 1); +lean_inc(x_114); +if (lean_is_exclusive(x_104)) { + lean_ctor_release(x_104, 0); + lean_ctor_release(x_104, 1); + x_115 = x_104; +} else { + lean_dec_ref(x_104); + x_115 = lean_box(0); +} +if (lean_is_scalar(x_115)) { + x_116 = lean_alloc_ctor(1, 2, 0); +} else { + x_116 = x_115; +} +lean_ctor_set(x_116, 0, x_113); +lean_ctor_set(x_116, 1, x_114); +return x_116; +} +} +} +} +else +{ +lean_object* x_117; lean_object* x_118; lean_object* x_119; uint8_t x_120; +lean_dec(x_36); +lean_dec(x_33); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_11); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +x_117 = lean_ctor_get(x_35, 1); +lean_inc(x_117); +lean_dec(x_35); +x_118 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__1___closed__4; +x_119 = l_Lean_Meta_throwTacticEx___rarg(x_6, x_4, x_118, x_14, x_15, x_16, x_17, x_117); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +x_120 = !lean_is_exclusive(x_119); +if (x_120 == 0) +{ +return x_119; +} +else +{ +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_124; +lean_dec(x_33); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_124 = !lean_is_exclusive(x_35); +if (x_124 == 0) +{ +return x_35; +} +else +{ +lean_object* x_125; lean_object* x_126; lean_object* x_127; +x_125 = lean_ctor_get(x_35, 0); +x_126 = lean_ctor_get(x_35, 1); +lean_inc(x_126); +lean_inc(x_125); +lean_dec(x_35); +x_127 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_127, 0, x_125); +lean_ctor_set(x_127, 1, x_126); +return x_127; +} +} +block_28: +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_22, 0); +lean_inc(x_24); +lean_dec(x_22); +x_25 = lean_ctor_get(x_9, 2); +x_26 = lean_nat_add(x_11, x_25); +lean_dec(x_11); +x_10 = x_24; +x_11 = x_26; +x_12 = lean_box(0); +x_13 = lean_box(0); +x_18 = x_23; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__9(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18) { +_start: +{ +lean_object* x_19; uint8_t x_20; +x_19 = lean_ctor_get(x_9, 1); +x_20 = lean_nat_dec_lt(x_11, x_19); +if (x_20 == 0) +{ +lean_object* x_21; +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +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; lean_object* x_23; lean_object* x_29; 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_10, 0); +lean_inc(x_29); +x_30 = lean_ctor_get(x_10, 1); +lean_inc(x_30); +lean_dec(x_10); +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_32, 0); +lean_inc(x_33); +x_34 = lean_ctor_get(x_32, 1); +lean_inc(x_34); +lean_dec(x_32); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_35 = lean_whnf(x_34, x_14, x_15, x_16, x_17, x_18); +if (lean_obj_tag(x_35) == 0) +{ +lean_object* x_36; +x_36 = lean_ctor_get(x_35, 0); +lean_inc(x_36); +if (lean_obj_tag(x_36) == 7) +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; uint8_t x_41; +x_37 = lean_ctor_get(x_35, 1); +lean_inc(x_37); +lean_dec(x_35); +x_38 = lean_ctor_get(x_36, 1); +lean_inc(x_38); +x_39 = lean_ctor_get(x_36, 2); +lean_inc(x_39); +lean_dec(x_36); +x_40 = lean_unsigned_to_nat(1u); +x_41 = lean_nat_dec_lt(x_40, x_7); +if (x_41 == 0) +{ +lean_object* x_42; uint8_t x_43; +lean_inc(x_14); +lean_inc(x_1); +x_42 = l_Lean_Meta_mkFreshExprSyntheticOpaqueMVar(x_38, x_1, x_14, x_15, x_16, x_17, x_37); +x_43 = !lean_is_exclusive(x_42); +if (x_43 == 0) +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_44 = lean_ctor_get(x_42, 0); +x_45 = lean_ctor_get(x_42, 1); +lean_inc(x_44); +x_46 = l_Lean_Expr_app___override(x_33, x_44); +x_47 = l_Lean_Expr_mvarId_x21(x_44); +lean_dec(x_44); +lean_inc(x_3); +lean_inc(x_5); +x_48 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_49 = l_Lean_MVarId_tryClearMany(x_47, x_48, x_14, x_15, x_16, x_17, x_45); +lean_dec(x_48); +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; +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_array_push(x_31, x_50); +x_53 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +lean_ctor_set(x_42, 1, x_39); +lean_ctor_set(x_42, 0, x_46); +x_54 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_54, 0, x_52); +lean_ctor_set(x_54, 1, x_42); +x_55 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_55, 0, x_53); +lean_ctor_set(x_55, 1, x_54); +x_56 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_56, 0, x_55); +x_22 = x_56; +x_23 = x_51; +goto block_28; +} +else +{ +uint8_t x_57; +lean_dec(x_46); +lean_free_object(x_42); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_57 = !lean_is_exclusive(x_49); +if (x_57 == 0) +{ +return x_49; +} +else +{ +lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_58 = lean_ctor_get(x_49, 0); +x_59 = lean_ctor_get(x_49, 1); +lean_inc(x_59); +lean_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; +} +} +} +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; +x_61 = lean_ctor_get(x_42, 0); +x_62 = lean_ctor_get(x_42, 1); +lean_inc(x_62); +lean_inc(x_61); +lean_dec(x_42); +lean_inc(x_61); +x_63 = l_Lean_Expr_app___override(x_33, x_61); +x_64 = l_Lean_Expr_mvarId_x21(x_61); +lean_dec(x_61); +lean_inc(x_3); +lean_inc(x_5); +x_65 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_66 = l_Lean_MVarId_tryClearMany(x_64, x_65, x_14, x_15, x_16, x_17, x_62); +lean_dec(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; +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_array_push(x_31, x_67); +x_70 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +x_71 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_71, 0, x_63); +lean_ctor_set(x_71, 1, x_39); +x_72 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_72, 0, x_69); +lean_ctor_set(x_72, 1, x_71); +x_73 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_73, 0, x_70); +lean_ctor_set(x_73, 1, x_72); +x_74 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_74, 0, x_73); +x_22 = x_74; +x_23 = x_68; +goto block_28; +} +else +{ +lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; +lean_dec(x_63); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_75 = lean_ctor_get(x_66, 0); +lean_inc(x_75); +x_76 = lean_ctor_get(x_66, 1); +lean_inc(x_76); +if (lean_is_exclusive(x_66)) { + lean_ctor_release(x_66, 0); + lean_ctor_release(x_66, 1); + x_77 = x_66; +} else { + lean_dec_ref(x_66); + x_77 = lean_box(0); +} +if (lean_is_scalar(x_77)) { + x_78 = lean_alloc_ctor(1, 2, 0); +} else { + x_78 = x_77; +} +lean_ctor_set(x_78, 0, x_75); +lean_ctor_set(x_78, 1, x_76); +return x_78; +} +} +} +else +{ +lean_object* x_79; lean_object* x_80; uint8_t x_81; +lean_inc(x_29); +lean_inc(x_1); +x_79 = l_Lean_Name_num___override(x_1, x_29); +lean_inc(x_14); +x_80 = l_Lean_Meta_mkFreshExprSyntheticOpaqueMVar(x_38, x_79, x_14, x_15, x_16, x_17, x_37); +x_81 = !lean_is_exclusive(x_80); +if (x_81 == 0) +{ +lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; +x_82 = lean_ctor_get(x_80, 0); +x_83 = lean_ctor_get(x_80, 1); +lean_inc(x_82); +x_84 = l_Lean_Expr_app___override(x_33, x_82); +x_85 = l_Lean_Expr_mvarId_x21(x_82); +lean_dec(x_82); +lean_inc(x_3); +lean_inc(x_5); +x_86 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_87 = l_Lean_MVarId_tryClearMany(x_85, x_86, x_14, x_15, x_16, x_17, x_83); +lean_dec(x_86); +if (lean_obj_tag(x_87) == 0) +{ +lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; +x_88 = lean_ctor_get(x_87, 0); +lean_inc(x_88); +x_89 = lean_ctor_get(x_87, 1); +lean_inc(x_89); +lean_dec(x_87); +x_90 = lean_array_push(x_31, x_88); +x_91 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +lean_ctor_set(x_80, 1, x_39); +lean_ctor_set(x_80, 0, x_84); +x_92 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_92, 0, x_90); +lean_ctor_set(x_92, 1, x_80); +x_93 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_93, 0, x_91); +lean_ctor_set(x_93, 1, x_92); +x_94 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_94, 0, x_93); +x_22 = x_94; +x_23 = x_89; +goto block_28; +} +else +{ +uint8_t x_95; +lean_dec(x_84); +lean_free_object(x_80); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_95 = !lean_is_exclusive(x_87); +if (x_95 == 0) +{ +return x_87; +} +else +{ +lean_object* x_96; lean_object* x_97; lean_object* x_98; +x_96 = lean_ctor_get(x_87, 0); +x_97 = lean_ctor_get(x_87, 1); +lean_inc(x_97); +lean_inc(x_96); +lean_dec(x_87); +x_98 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_98, 0, x_96); +lean_ctor_set(x_98, 1, x_97); +return x_98; +} +} +} +else +{ +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_80, 0); +x_100 = lean_ctor_get(x_80, 1); +lean_inc(x_100); +lean_inc(x_99); +lean_dec(x_80); +lean_inc(x_99); +x_101 = l_Lean_Expr_app___override(x_33, x_99); +x_102 = l_Lean_Expr_mvarId_x21(x_99); +lean_dec(x_99); +lean_inc(x_3); +lean_inc(x_5); +x_103 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_104 = l_Lean_MVarId_tryClearMany(x_102, x_103, x_14, x_15, x_16, x_17, x_100); +lean_dec(x_103); +if (lean_obj_tag(x_104) == 0) +{ +lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; +x_105 = lean_ctor_get(x_104, 0); +lean_inc(x_105); +x_106 = lean_ctor_get(x_104, 1); +lean_inc(x_106); +lean_dec(x_104); +x_107 = lean_array_push(x_31, x_105); +x_108 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +x_109 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_109, 0, x_101); +lean_ctor_set(x_109, 1, x_39); +x_110 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_110, 0, x_107); +lean_ctor_set(x_110, 1, x_109); +x_111 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_111, 0, x_108); +lean_ctor_set(x_111, 1, x_110); +x_112 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_112, 0, x_111); +x_22 = x_112; +x_23 = x_106; +goto block_28; +} +else +{ +lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; +lean_dec(x_101); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_113 = lean_ctor_get(x_104, 0); +lean_inc(x_113); +x_114 = lean_ctor_get(x_104, 1); +lean_inc(x_114); +if (lean_is_exclusive(x_104)) { + lean_ctor_release(x_104, 0); + lean_ctor_release(x_104, 1); + x_115 = x_104; +} else { + lean_dec_ref(x_104); + x_115 = lean_box(0); +} +if (lean_is_scalar(x_115)) { + x_116 = lean_alloc_ctor(1, 2, 0); +} else { + x_116 = x_115; +} +lean_ctor_set(x_116, 0, x_113); +lean_ctor_set(x_116, 1, x_114); +return x_116; +} +} +} +} +else +{ +lean_object* x_117; lean_object* x_118; lean_object* x_119; uint8_t x_120; +lean_dec(x_36); +lean_dec(x_33); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_11); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +x_117 = lean_ctor_get(x_35, 1); +lean_inc(x_117); +lean_dec(x_35); +x_118 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__1___closed__4; +x_119 = l_Lean_Meta_throwTacticEx___rarg(x_6, x_4, x_118, x_14, x_15, x_16, x_17, x_117); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +x_120 = !lean_is_exclusive(x_119); +if (x_120 == 0) +{ +return x_119; +} +else +{ +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_124; +lean_dec(x_33); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_124 = !lean_is_exclusive(x_35); +if (x_124 == 0) +{ +return x_35; +} +else +{ +lean_object* x_125; lean_object* x_126; lean_object* x_127; +x_125 = lean_ctor_get(x_35, 0); +x_126 = lean_ctor_get(x_35, 1); +lean_inc(x_126); +lean_inc(x_125); +lean_dec(x_35); +x_127 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_127, 0, x_125); +lean_ctor_set(x_127, 1, x_126); +return x_127; +} +} +block_28: +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_22, 0); +lean_inc(x_24); +lean_dec(x_22); +x_25 = lean_ctor_get(x_9, 2); +x_26 = lean_nat_add(x_11, x_25); +lean_dec(x_11); +x_10 = x_24; +x_11 = x_26; +x_12 = lean_box(0); +x_13 = lean_box(0); +x_18 = x_23; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__10(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18) { +_start: +{ +lean_object* x_19; uint8_t x_20; +x_19 = lean_ctor_get(x_9, 1); +x_20 = lean_nat_dec_lt(x_11, x_19); +if (x_20 == 0) +{ +lean_object* x_21; +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +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; lean_object* x_23; lean_object* x_29; 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_10, 0); +lean_inc(x_29); +x_30 = lean_ctor_get(x_10, 1); +lean_inc(x_30); +lean_dec(x_10); +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_32, 0); +lean_inc(x_33); +x_34 = lean_ctor_get(x_32, 1); +lean_inc(x_34); +lean_dec(x_32); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_35 = lean_whnf(x_34, x_14, x_15, x_16, x_17, x_18); +if (lean_obj_tag(x_35) == 0) +{ +lean_object* x_36; +x_36 = lean_ctor_get(x_35, 0); +lean_inc(x_36); +if (lean_obj_tag(x_36) == 7) +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; uint8_t x_41; +x_37 = lean_ctor_get(x_35, 1); +lean_inc(x_37); +lean_dec(x_35); +x_38 = lean_ctor_get(x_36, 1); +lean_inc(x_38); +x_39 = lean_ctor_get(x_36, 2); +lean_inc(x_39); +lean_dec(x_36); +x_40 = lean_unsigned_to_nat(1u); +x_41 = lean_nat_dec_lt(x_40, x_7); +if (x_41 == 0) +{ +lean_object* x_42; uint8_t x_43; +lean_inc(x_14); +lean_inc(x_1); +x_42 = l_Lean_Meta_mkFreshExprSyntheticOpaqueMVar(x_38, x_1, x_14, x_15, x_16, x_17, x_37); +x_43 = !lean_is_exclusive(x_42); +if (x_43 == 0) +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_44 = lean_ctor_get(x_42, 0); +x_45 = lean_ctor_get(x_42, 1); +lean_inc(x_44); +x_46 = l_Lean_Expr_app___override(x_33, x_44); +x_47 = l_Lean_Expr_mvarId_x21(x_44); +lean_dec(x_44); +lean_inc(x_3); +lean_inc(x_5); +x_48 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_49 = l_Lean_MVarId_tryClearMany(x_47, x_48, x_14, x_15, x_16, x_17, x_45); +lean_dec(x_48); +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; +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_array_push(x_31, x_50); +x_53 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +lean_ctor_set(x_42, 1, x_39); +lean_ctor_set(x_42, 0, x_46); +x_54 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_54, 0, x_52); +lean_ctor_set(x_54, 1, x_42); +x_55 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_55, 0, x_53); +lean_ctor_set(x_55, 1, x_54); +x_56 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_56, 0, x_55); +x_22 = x_56; +x_23 = x_51; +goto block_28; +} +else +{ +uint8_t x_57; +lean_dec(x_46); +lean_free_object(x_42); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_57 = !lean_is_exclusive(x_49); +if (x_57 == 0) +{ +return x_49; +} +else +{ +lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_58 = lean_ctor_get(x_49, 0); +x_59 = lean_ctor_get(x_49, 1); +lean_inc(x_59); +lean_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; +} +} +} +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; +x_61 = lean_ctor_get(x_42, 0); +x_62 = lean_ctor_get(x_42, 1); +lean_inc(x_62); +lean_inc(x_61); +lean_dec(x_42); +lean_inc(x_61); +x_63 = l_Lean_Expr_app___override(x_33, x_61); +x_64 = l_Lean_Expr_mvarId_x21(x_61); +lean_dec(x_61); +lean_inc(x_3); +lean_inc(x_5); +x_65 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_66 = l_Lean_MVarId_tryClearMany(x_64, x_65, x_14, x_15, x_16, x_17, x_62); +lean_dec(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; +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_array_push(x_31, x_67); +x_70 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +x_71 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_71, 0, x_63); +lean_ctor_set(x_71, 1, x_39); +x_72 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_72, 0, x_69); +lean_ctor_set(x_72, 1, x_71); +x_73 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_73, 0, x_70); +lean_ctor_set(x_73, 1, x_72); +x_74 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_74, 0, x_73); +x_22 = x_74; +x_23 = x_68; +goto block_28; +} +else +{ +lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; +lean_dec(x_63); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_75 = lean_ctor_get(x_66, 0); +lean_inc(x_75); +x_76 = lean_ctor_get(x_66, 1); +lean_inc(x_76); +if (lean_is_exclusive(x_66)) { + lean_ctor_release(x_66, 0); + lean_ctor_release(x_66, 1); + x_77 = x_66; +} else { + lean_dec_ref(x_66); + x_77 = lean_box(0); +} +if (lean_is_scalar(x_77)) { + x_78 = lean_alloc_ctor(1, 2, 0); +} else { + x_78 = x_77; +} +lean_ctor_set(x_78, 0, x_75); +lean_ctor_set(x_78, 1, x_76); +return x_78; +} +} +} +else +{ +lean_object* x_79; lean_object* x_80; uint8_t x_81; +lean_inc(x_29); +lean_inc(x_1); +x_79 = l_Lean_Name_num___override(x_1, x_29); +lean_inc(x_14); +x_80 = l_Lean_Meta_mkFreshExprSyntheticOpaqueMVar(x_38, x_79, x_14, x_15, x_16, x_17, x_37); +x_81 = !lean_is_exclusive(x_80); +if (x_81 == 0) +{ +lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; +x_82 = lean_ctor_get(x_80, 0); +x_83 = lean_ctor_get(x_80, 1); +lean_inc(x_82); +x_84 = l_Lean_Expr_app___override(x_33, x_82); +x_85 = l_Lean_Expr_mvarId_x21(x_82); +lean_dec(x_82); +lean_inc(x_3); +lean_inc(x_5); +x_86 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_87 = l_Lean_MVarId_tryClearMany(x_85, x_86, x_14, x_15, x_16, x_17, x_83); +lean_dec(x_86); +if (lean_obj_tag(x_87) == 0) +{ +lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; +x_88 = lean_ctor_get(x_87, 0); +lean_inc(x_88); +x_89 = lean_ctor_get(x_87, 1); +lean_inc(x_89); +lean_dec(x_87); +x_90 = lean_array_push(x_31, x_88); +x_91 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +lean_ctor_set(x_80, 1, x_39); +lean_ctor_set(x_80, 0, x_84); +x_92 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_92, 0, x_90); +lean_ctor_set(x_92, 1, x_80); +x_93 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_93, 0, x_91); +lean_ctor_set(x_93, 1, x_92); +x_94 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_94, 0, x_93); +x_22 = x_94; +x_23 = x_89; +goto block_28; +} +else +{ +uint8_t x_95; +lean_dec(x_84); +lean_free_object(x_80); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_95 = !lean_is_exclusive(x_87); +if (x_95 == 0) +{ +return x_87; +} +else +{ +lean_object* x_96; lean_object* x_97; lean_object* x_98; +x_96 = lean_ctor_get(x_87, 0); +x_97 = lean_ctor_get(x_87, 1); +lean_inc(x_97); +lean_inc(x_96); +lean_dec(x_87); +x_98 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_98, 0, x_96); +lean_ctor_set(x_98, 1, x_97); +return x_98; +} +} +} +else +{ +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_80, 0); +x_100 = lean_ctor_get(x_80, 1); +lean_inc(x_100); +lean_inc(x_99); +lean_dec(x_80); +lean_inc(x_99); +x_101 = l_Lean_Expr_app___override(x_33, x_99); +x_102 = l_Lean_Expr_mvarId_x21(x_99); +lean_dec(x_99); +lean_inc(x_3); +lean_inc(x_5); +x_103 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_104 = l_Lean_MVarId_tryClearMany(x_102, x_103, x_14, x_15, x_16, x_17, x_100); +lean_dec(x_103); +if (lean_obj_tag(x_104) == 0) +{ +lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; +x_105 = lean_ctor_get(x_104, 0); +lean_inc(x_105); +x_106 = lean_ctor_get(x_104, 1); +lean_inc(x_106); +lean_dec(x_104); +x_107 = lean_array_push(x_31, x_105); +x_108 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +x_109 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_109, 0, x_101); +lean_ctor_set(x_109, 1, x_39); +x_110 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_110, 0, x_107); +lean_ctor_set(x_110, 1, x_109); +x_111 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_111, 0, x_108); +lean_ctor_set(x_111, 1, x_110); +x_112 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_112, 0, x_111); +x_22 = x_112; +x_23 = x_106; +goto block_28; +} +else +{ +lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; +lean_dec(x_101); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_113 = lean_ctor_get(x_104, 0); +lean_inc(x_113); +x_114 = lean_ctor_get(x_104, 1); +lean_inc(x_114); +if (lean_is_exclusive(x_104)) { + lean_ctor_release(x_104, 0); + lean_ctor_release(x_104, 1); + x_115 = x_104; +} else { + lean_dec_ref(x_104); + x_115 = lean_box(0); +} +if (lean_is_scalar(x_115)) { + x_116 = lean_alloc_ctor(1, 2, 0); +} else { + x_116 = x_115; +} +lean_ctor_set(x_116, 0, x_113); +lean_ctor_set(x_116, 1, x_114); +return x_116; +} +} +} +} +else +{ +lean_object* x_117; lean_object* x_118; lean_object* x_119; uint8_t x_120; +lean_dec(x_36); +lean_dec(x_33); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_11); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +x_117 = lean_ctor_get(x_35, 1); +lean_inc(x_117); +lean_dec(x_35); +x_118 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__1___closed__4; +x_119 = l_Lean_Meta_throwTacticEx___rarg(x_6, x_4, x_118, x_14, x_15, x_16, x_17, x_117); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +x_120 = !lean_is_exclusive(x_119); +if (x_120 == 0) +{ +return x_119; +} +else +{ +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_124; +lean_dec(x_33); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_124 = !lean_is_exclusive(x_35); +if (x_124 == 0) +{ +return x_35; +} +else +{ +lean_object* x_125; lean_object* x_126; lean_object* x_127; +x_125 = lean_ctor_get(x_35, 0); +x_126 = lean_ctor_get(x_35, 1); +lean_inc(x_126); +lean_inc(x_125); +lean_dec(x_35); +x_127 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_127, 0, x_125); +lean_ctor_set(x_127, 1, x_126); +return x_127; +} +} +block_28: +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_22, 0); +lean_inc(x_24); +lean_dec(x_22); +x_25 = lean_ctor_get(x_9, 2); +x_26 = lean_nat_add(x_11, x_25); +lean_dec(x_11); +x_10 = x_24; +x_11 = x_26; +x_12 = lean_box(0); +x_13 = lean_box(0); +x_18 = x_23; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__11(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18) { +_start: +{ +lean_object* x_19; uint8_t x_20; +x_19 = lean_ctor_get(x_9, 1); +x_20 = lean_nat_dec_lt(x_11, x_19); +if (x_20 == 0) +{ +lean_object* x_21; +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +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; lean_object* x_23; lean_object* x_29; 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_10, 0); +lean_inc(x_29); +x_30 = lean_ctor_get(x_10, 1); +lean_inc(x_30); +lean_dec(x_10); +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_32, 0); +lean_inc(x_33); +x_34 = lean_ctor_get(x_32, 1); +lean_inc(x_34); +lean_dec(x_32); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_35 = lean_whnf(x_34, x_14, x_15, x_16, x_17, x_18); +if (lean_obj_tag(x_35) == 0) +{ +lean_object* x_36; +x_36 = lean_ctor_get(x_35, 0); +lean_inc(x_36); +if (lean_obj_tag(x_36) == 7) +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; uint8_t x_41; +x_37 = lean_ctor_get(x_35, 1); +lean_inc(x_37); +lean_dec(x_35); +x_38 = lean_ctor_get(x_36, 1); +lean_inc(x_38); +x_39 = lean_ctor_get(x_36, 2); +lean_inc(x_39); +lean_dec(x_36); +x_40 = lean_unsigned_to_nat(1u); +x_41 = lean_nat_dec_lt(x_40, x_7); +if (x_41 == 0) +{ +lean_object* x_42; uint8_t x_43; +lean_inc(x_14); +lean_inc(x_1); +x_42 = l_Lean_Meta_mkFreshExprSyntheticOpaqueMVar(x_38, x_1, x_14, x_15, x_16, x_17, x_37); +x_43 = !lean_is_exclusive(x_42); +if (x_43 == 0) +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_44 = lean_ctor_get(x_42, 0); +x_45 = lean_ctor_get(x_42, 1); +lean_inc(x_44); +x_46 = l_Lean_Expr_app___override(x_33, x_44); +x_47 = l_Lean_Expr_mvarId_x21(x_44); +lean_dec(x_44); +lean_inc(x_3); +lean_inc(x_5); +x_48 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_49 = l_Lean_MVarId_tryClearMany(x_47, x_48, x_14, x_15, x_16, x_17, x_45); +lean_dec(x_48); +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; +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_array_push(x_31, x_50); +x_53 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +lean_ctor_set(x_42, 1, x_39); +lean_ctor_set(x_42, 0, x_46); +x_54 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_54, 0, x_52); +lean_ctor_set(x_54, 1, x_42); +x_55 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_55, 0, x_53); +lean_ctor_set(x_55, 1, x_54); +x_56 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_56, 0, x_55); +x_22 = x_56; +x_23 = x_51; +goto block_28; +} +else +{ +uint8_t x_57; +lean_dec(x_46); +lean_free_object(x_42); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_57 = !lean_is_exclusive(x_49); +if (x_57 == 0) +{ +return x_49; +} +else +{ +lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_58 = lean_ctor_get(x_49, 0); +x_59 = lean_ctor_get(x_49, 1); +lean_inc(x_59); +lean_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; +} +} +} +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; +x_61 = lean_ctor_get(x_42, 0); +x_62 = lean_ctor_get(x_42, 1); +lean_inc(x_62); +lean_inc(x_61); +lean_dec(x_42); +lean_inc(x_61); +x_63 = l_Lean_Expr_app___override(x_33, x_61); +x_64 = l_Lean_Expr_mvarId_x21(x_61); +lean_dec(x_61); +lean_inc(x_3); +lean_inc(x_5); +x_65 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_66 = l_Lean_MVarId_tryClearMany(x_64, x_65, x_14, x_15, x_16, x_17, x_62); +lean_dec(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; +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_array_push(x_31, x_67); +x_70 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +x_71 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_71, 0, x_63); +lean_ctor_set(x_71, 1, x_39); +x_72 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_72, 0, x_69); +lean_ctor_set(x_72, 1, x_71); +x_73 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_73, 0, x_70); +lean_ctor_set(x_73, 1, x_72); +x_74 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_74, 0, x_73); +x_22 = x_74; +x_23 = x_68; +goto block_28; +} +else +{ +lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; +lean_dec(x_63); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_75 = lean_ctor_get(x_66, 0); +lean_inc(x_75); +x_76 = lean_ctor_get(x_66, 1); +lean_inc(x_76); +if (lean_is_exclusive(x_66)) { + lean_ctor_release(x_66, 0); + lean_ctor_release(x_66, 1); + x_77 = x_66; +} else { + lean_dec_ref(x_66); + x_77 = lean_box(0); +} +if (lean_is_scalar(x_77)) { + x_78 = lean_alloc_ctor(1, 2, 0); +} else { + x_78 = x_77; +} +lean_ctor_set(x_78, 0, x_75); +lean_ctor_set(x_78, 1, x_76); +return x_78; +} +} +} +else +{ +lean_object* x_79; lean_object* x_80; uint8_t x_81; +lean_inc(x_29); +lean_inc(x_1); +x_79 = l_Lean_Name_num___override(x_1, x_29); +lean_inc(x_14); +x_80 = l_Lean_Meta_mkFreshExprSyntheticOpaqueMVar(x_38, x_79, x_14, x_15, x_16, x_17, x_37); +x_81 = !lean_is_exclusive(x_80); +if (x_81 == 0) +{ +lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; +x_82 = lean_ctor_get(x_80, 0); +x_83 = lean_ctor_get(x_80, 1); +lean_inc(x_82); +x_84 = l_Lean_Expr_app___override(x_33, x_82); +x_85 = l_Lean_Expr_mvarId_x21(x_82); +lean_dec(x_82); +lean_inc(x_3); +lean_inc(x_5); +x_86 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_87 = l_Lean_MVarId_tryClearMany(x_85, x_86, x_14, x_15, x_16, x_17, x_83); +lean_dec(x_86); +if (lean_obj_tag(x_87) == 0) +{ +lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; +x_88 = lean_ctor_get(x_87, 0); +lean_inc(x_88); +x_89 = lean_ctor_get(x_87, 1); +lean_inc(x_89); +lean_dec(x_87); +x_90 = lean_array_push(x_31, x_88); +x_91 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +lean_ctor_set(x_80, 1, x_39); +lean_ctor_set(x_80, 0, x_84); +x_92 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_92, 0, x_90); +lean_ctor_set(x_92, 1, x_80); +x_93 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_93, 0, x_91); +lean_ctor_set(x_93, 1, x_92); +x_94 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_94, 0, x_93); +x_22 = x_94; +x_23 = x_89; +goto block_28; +} +else +{ +uint8_t x_95; +lean_dec(x_84); +lean_free_object(x_80); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_95 = !lean_is_exclusive(x_87); +if (x_95 == 0) +{ +return x_87; +} +else +{ +lean_object* x_96; lean_object* x_97; lean_object* x_98; +x_96 = lean_ctor_get(x_87, 0); +x_97 = lean_ctor_get(x_87, 1); +lean_inc(x_97); +lean_inc(x_96); +lean_dec(x_87); +x_98 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_98, 0, x_96); +lean_ctor_set(x_98, 1, x_97); +return x_98; +} +} +} +else +{ +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_80, 0); +x_100 = lean_ctor_get(x_80, 1); +lean_inc(x_100); +lean_inc(x_99); +lean_dec(x_80); +lean_inc(x_99); +x_101 = l_Lean_Expr_app___override(x_33, x_99); +x_102 = l_Lean_Expr_mvarId_x21(x_99); +lean_dec(x_99); +lean_inc(x_3); +lean_inc(x_5); +x_103 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_104 = l_Lean_MVarId_tryClearMany(x_102, x_103, x_14, x_15, x_16, x_17, x_100); +lean_dec(x_103); +if (lean_obj_tag(x_104) == 0) +{ +lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; +x_105 = lean_ctor_get(x_104, 0); +lean_inc(x_105); +x_106 = lean_ctor_get(x_104, 1); +lean_inc(x_106); +lean_dec(x_104); +x_107 = lean_array_push(x_31, x_105); +x_108 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +x_109 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_109, 0, x_101); +lean_ctor_set(x_109, 1, x_39); +x_110 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_110, 0, x_107); +lean_ctor_set(x_110, 1, x_109); +x_111 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_111, 0, x_108); +lean_ctor_set(x_111, 1, x_110); +x_112 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_112, 0, x_111); +x_22 = x_112; +x_23 = x_106; +goto block_28; +} +else +{ +lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; +lean_dec(x_101); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_113 = lean_ctor_get(x_104, 0); +lean_inc(x_113); +x_114 = lean_ctor_get(x_104, 1); +lean_inc(x_114); +if (lean_is_exclusive(x_104)) { + lean_ctor_release(x_104, 0); + lean_ctor_release(x_104, 1); + x_115 = x_104; +} else { + lean_dec_ref(x_104); + x_115 = lean_box(0); +} +if (lean_is_scalar(x_115)) { + x_116 = lean_alloc_ctor(1, 2, 0); +} else { + x_116 = x_115; +} +lean_ctor_set(x_116, 0, x_113); +lean_ctor_set(x_116, 1, x_114); +return x_116; +} +} +} +} +else +{ +lean_object* x_117; lean_object* x_118; lean_object* x_119; uint8_t x_120; +lean_dec(x_36); +lean_dec(x_33); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_11); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +x_117 = lean_ctor_get(x_35, 1); +lean_inc(x_117); +lean_dec(x_35); +x_118 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__1___closed__4; +x_119 = l_Lean_Meta_throwTacticEx___rarg(x_6, x_4, x_118, x_14, x_15, x_16, x_17, x_117); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +x_120 = !lean_is_exclusive(x_119); +if (x_120 == 0) +{ +return x_119; +} +else +{ +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_124; +lean_dec(x_33); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_124 = !lean_is_exclusive(x_35); +if (x_124 == 0) +{ +return x_35; +} +else +{ +lean_object* x_125; lean_object* x_126; lean_object* x_127; +x_125 = lean_ctor_get(x_35, 0); +x_126 = lean_ctor_get(x_35, 1); +lean_inc(x_126); +lean_inc(x_125); +lean_dec(x_35); +x_127 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_127, 0, x_125); +lean_ctor_set(x_127, 1, x_126); +return x_127; +} +} +block_28: +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_22, 0); +lean_inc(x_24); +lean_dec(x_22); +x_25 = lean_ctor_get(x_9, 2); +x_26 = lean_nat_add(x_11, x_25); +lean_dec(x_11); +x_10 = x_24; +x_11 = x_26; +x_12 = lean_box(0); +x_13 = lean_box(0); +x_18 = x_23; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__12(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18) { +_start: +{ +lean_object* x_19; uint8_t x_20; +x_19 = lean_ctor_get(x_9, 1); +x_20 = lean_nat_dec_lt(x_11, x_19); +if (x_20 == 0) +{ +lean_object* x_21; +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +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; lean_object* x_23; lean_object* x_29; 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_10, 0); +lean_inc(x_29); +x_30 = lean_ctor_get(x_10, 1); +lean_inc(x_30); +lean_dec(x_10); +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_32, 0); +lean_inc(x_33); +x_34 = lean_ctor_get(x_32, 1); +lean_inc(x_34); +lean_dec(x_32); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_35 = lean_whnf(x_34, x_14, x_15, x_16, x_17, x_18); +if (lean_obj_tag(x_35) == 0) +{ +lean_object* x_36; +x_36 = lean_ctor_get(x_35, 0); +lean_inc(x_36); +if (lean_obj_tag(x_36) == 7) +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; uint8_t x_41; +x_37 = lean_ctor_get(x_35, 1); +lean_inc(x_37); +lean_dec(x_35); +x_38 = lean_ctor_get(x_36, 1); +lean_inc(x_38); +x_39 = lean_ctor_get(x_36, 2); +lean_inc(x_39); +lean_dec(x_36); +x_40 = lean_unsigned_to_nat(1u); +x_41 = lean_nat_dec_lt(x_40, x_7); +if (x_41 == 0) +{ +lean_object* x_42; uint8_t x_43; +lean_inc(x_14); +lean_inc(x_1); +x_42 = l_Lean_Meta_mkFreshExprSyntheticOpaqueMVar(x_38, x_1, x_14, x_15, x_16, x_17, x_37); +x_43 = !lean_is_exclusive(x_42); +if (x_43 == 0) +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_44 = lean_ctor_get(x_42, 0); +x_45 = lean_ctor_get(x_42, 1); +lean_inc(x_44); +x_46 = l_Lean_Expr_app___override(x_33, x_44); +x_47 = l_Lean_Expr_mvarId_x21(x_44); +lean_dec(x_44); +lean_inc(x_3); +lean_inc(x_5); +x_48 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_49 = l_Lean_MVarId_tryClearMany(x_47, x_48, x_14, x_15, x_16, x_17, x_45); +lean_dec(x_48); +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; +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_array_push(x_31, x_50); +x_53 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +lean_ctor_set(x_42, 1, x_39); +lean_ctor_set(x_42, 0, x_46); +x_54 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_54, 0, x_52); +lean_ctor_set(x_54, 1, x_42); +x_55 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_55, 0, x_53); +lean_ctor_set(x_55, 1, x_54); +x_56 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_56, 0, x_55); +x_22 = x_56; +x_23 = x_51; +goto block_28; +} +else +{ +uint8_t x_57; +lean_dec(x_46); +lean_free_object(x_42); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_57 = !lean_is_exclusive(x_49); +if (x_57 == 0) +{ +return x_49; +} +else +{ +lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_58 = lean_ctor_get(x_49, 0); +x_59 = lean_ctor_get(x_49, 1); +lean_inc(x_59); +lean_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; +} +} +} +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; +x_61 = lean_ctor_get(x_42, 0); +x_62 = lean_ctor_get(x_42, 1); +lean_inc(x_62); +lean_inc(x_61); +lean_dec(x_42); +lean_inc(x_61); +x_63 = l_Lean_Expr_app___override(x_33, x_61); +x_64 = l_Lean_Expr_mvarId_x21(x_61); +lean_dec(x_61); +lean_inc(x_3); +lean_inc(x_5); +x_65 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_66 = l_Lean_MVarId_tryClearMany(x_64, x_65, x_14, x_15, x_16, x_17, x_62); +lean_dec(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; +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_array_push(x_31, x_67); +x_70 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +x_71 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_71, 0, x_63); +lean_ctor_set(x_71, 1, x_39); +x_72 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_72, 0, x_69); +lean_ctor_set(x_72, 1, x_71); +x_73 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_73, 0, x_70); +lean_ctor_set(x_73, 1, x_72); +x_74 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_74, 0, x_73); +x_22 = x_74; +x_23 = x_68; +goto block_28; +} +else +{ +lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; +lean_dec(x_63); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_75 = lean_ctor_get(x_66, 0); +lean_inc(x_75); +x_76 = lean_ctor_get(x_66, 1); +lean_inc(x_76); +if (lean_is_exclusive(x_66)) { + lean_ctor_release(x_66, 0); + lean_ctor_release(x_66, 1); + x_77 = x_66; +} else { + lean_dec_ref(x_66); + x_77 = lean_box(0); +} +if (lean_is_scalar(x_77)) { + x_78 = lean_alloc_ctor(1, 2, 0); +} else { + x_78 = x_77; +} +lean_ctor_set(x_78, 0, x_75); +lean_ctor_set(x_78, 1, x_76); +return x_78; +} +} +} +else +{ +lean_object* x_79; lean_object* x_80; uint8_t x_81; +lean_inc(x_29); +lean_inc(x_1); +x_79 = l_Lean_Name_num___override(x_1, x_29); +lean_inc(x_14); +x_80 = l_Lean_Meta_mkFreshExprSyntheticOpaqueMVar(x_38, x_79, x_14, x_15, x_16, x_17, x_37); +x_81 = !lean_is_exclusive(x_80); +if (x_81 == 0) +{ +lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; +x_82 = lean_ctor_get(x_80, 0); +x_83 = lean_ctor_get(x_80, 1); +lean_inc(x_82); +x_84 = l_Lean_Expr_app___override(x_33, x_82); +x_85 = l_Lean_Expr_mvarId_x21(x_82); +lean_dec(x_82); +lean_inc(x_3); +lean_inc(x_5); +x_86 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_87 = l_Lean_MVarId_tryClearMany(x_85, x_86, x_14, x_15, x_16, x_17, x_83); +lean_dec(x_86); +if (lean_obj_tag(x_87) == 0) +{ +lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; +x_88 = lean_ctor_get(x_87, 0); +lean_inc(x_88); +x_89 = lean_ctor_get(x_87, 1); +lean_inc(x_89); +lean_dec(x_87); +x_90 = lean_array_push(x_31, x_88); +x_91 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +lean_ctor_set(x_80, 1, x_39); +lean_ctor_set(x_80, 0, x_84); +x_92 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_92, 0, x_90); +lean_ctor_set(x_92, 1, x_80); +x_93 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_93, 0, x_91); +lean_ctor_set(x_93, 1, x_92); +x_94 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_94, 0, x_93); +x_22 = x_94; +x_23 = x_89; +goto block_28; +} +else +{ +uint8_t x_95; +lean_dec(x_84); +lean_free_object(x_80); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_95 = !lean_is_exclusive(x_87); +if (x_95 == 0) +{ +return x_87; +} +else +{ +lean_object* x_96; lean_object* x_97; lean_object* x_98; +x_96 = lean_ctor_get(x_87, 0); +x_97 = lean_ctor_get(x_87, 1); +lean_inc(x_97); +lean_inc(x_96); +lean_dec(x_87); +x_98 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_98, 0, x_96); +lean_ctor_set(x_98, 1, x_97); +return x_98; +} +} +} +else +{ +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_80, 0); +x_100 = lean_ctor_get(x_80, 1); +lean_inc(x_100); +lean_inc(x_99); +lean_dec(x_80); +lean_inc(x_99); +x_101 = l_Lean_Expr_app___override(x_33, x_99); +x_102 = l_Lean_Expr_mvarId_x21(x_99); +lean_dec(x_99); +lean_inc(x_3); +lean_inc(x_5); +x_103 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_104 = l_Lean_MVarId_tryClearMany(x_102, x_103, x_14, x_15, x_16, x_17, x_100); +lean_dec(x_103); +if (lean_obj_tag(x_104) == 0) +{ +lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; +x_105 = lean_ctor_get(x_104, 0); +lean_inc(x_105); +x_106 = lean_ctor_get(x_104, 1); +lean_inc(x_106); +lean_dec(x_104); +x_107 = lean_array_push(x_31, x_105); +x_108 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +x_109 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_109, 0, x_101); +lean_ctor_set(x_109, 1, x_39); +x_110 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_110, 0, x_107); +lean_ctor_set(x_110, 1, x_109); +x_111 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_111, 0, x_108); +lean_ctor_set(x_111, 1, x_110); +x_112 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_112, 0, x_111); +x_22 = x_112; +x_23 = x_106; +goto block_28; +} +else +{ +lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; +lean_dec(x_101); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_113 = lean_ctor_get(x_104, 0); +lean_inc(x_113); +x_114 = lean_ctor_get(x_104, 1); +lean_inc(x_114); +if (lean_is_exclusive(x_104)) { + lean_ctor_release(x_104, 0); + lean_ctor_release(x_104, 1); + x_115 = x_104; +} else { + lean_dec_ref(x_104); + x_115 = lean_box(0); +} +if (lean_is_scalar(x_115)) { + x_116 = lean_alloc_ctor(1, 2, 0); +} else { + x_116 = x_115; +} +lean_ctor_set(x_116, 0, x_113); +lean_ctor_set(x_116, 1, x_114); +return x_116; +} +} +} +} +else +{ +lean_object* x_117; lean_object* x_118; lean_object* x_119; uint8_t x_120; +lean_dec(x_36); +lean_dec(x_33); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_11); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +x_117 = lean_ctor_get(x_35, 1); +lean_inc(x_117); +lean_dec(x_35); +x_118 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__1___closed__4; +x_119 = l_Lean_Meta_throwTacticEx___rarg(x_6, x_4, x_118, x_14, x_15, x_16, x_17, x_117); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +x_120 = !lean_is_exclusive(x_119); +if (x_120 == 0) +{ +return x_119; +} +else +{ +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_124; +lean_dec(x_33); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_124 = !lean_is_exclusive(x_35); +if (x_124 == 0) +{ +return x_35; +} +else +{ +lean_object* x_125; lean_object* x_126; lean_object* x_127; +x_125 = lean_ctor_get(x_35, 0); +x_126 = lean_ctor_get(x_35, 1); +lean_inc(x_126); +lean_inc(x_125); +lean_dec(x_35); +x_127 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_127, 0, x_125); +lean_ctor_set(x_127, 1, x_126); +return x_127; +} +} +block_28: +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_22, 0); +lean_inc(x_24); +lean_dec(x_22); +x_25 = lean_ctor_get(x_9, 2); +x_26 = lean_nat_add(x_11, x_25); +lean_dec(x_11); +x_10 = x_24; +x_11 = x_26; +x_12 = lean_box(0); +x_13 = lean_box(0); +x_18 = x_23; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___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, 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; uint8_t x_20; +x_19 = lean_ctor_get(x_9, 1); +x_20 = lean_nat_dec_lt(x_11, x_19); +if (x_20 == 0) +{ +lean_object* x_21; +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +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; lean_object* x_23; lean_object* x_29; 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_10, 0); +lean_inc(x_29); +x_30 = lean_ctor_get(x_10, 1); +lean_inc(x_30); +lean_dec(x_10); +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_32, 0); +lean_inc(x_33); +x_34 = lean_ctor_get(x_32, 1); +lean_inc(x_34); +lean_dec(x_32); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_35 = lean_whnf(x_34, x_14, x_15, x_16, x_17, x_18); +if (lean_obj_tag(x_35) == 0) +{ +lean_object* x_36; +x_36 = lean_ctor_get(x_35, 0); +lean_inc(x_36); +if (lean_obj_tag(x_36) == 7) +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; uint8_t x_41; +x_37 = lean_ctor_get(x_35, 1); +lean_inc(x_37); +lean_dec(x_35); +x_38 = lean_ctor_get(x_36, 1); +lean_inc(x_38); +x_39 = lean_ctor_get(x_36, 2); +lean_inc(x_39); +lean_dec(x_36); +x_40 = lean_unsigned_to_nat(1u); +x_41 = lean_nat_dec_lt(x_40, x_7); +if (x_41 == 0) +{ +lean_object* x_42; uint8_t x_43; +lean_inc(x_14); +lean_inc(x_1); +x_42 = l_Lean_Meta_mkFreshExprSyntheticOpaqueMVar(x_38, x_1, x_14, x_15, x_16, x_17, x_37); +x_43 = !lean_is_exclusive(x_42); +if (x_43 == 0) +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_44 = lean_ctor_get(x_42, 0); +x_45 = lean_ctor_get(x_42, 1); +lean_inc(x_44); +x_46 = l_Lean_Expr_app___override(x_33, x_44); +x_47 = l_Lean_Expr_mvarId_x21(x_44); +lean_dec(x_44); +lean_inc(x_4); +lean_inc(x_5); +x_48 = lean_array_push(x_5, x_4); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_49 = l_Lean_MVarId_tryClearMany(x_47, x_48, x_14, x_15, x_16, x_17, x_45); +lean_dec(x_48); +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; +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_array_push(x_31, x_50); +x_53 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +lean_ctor_set(x_42, 1, x_39); +lean_ctor_set(x_42, 0, x_46); +x_54 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_54, 0, x_52); +lean_ctor_set(x_54, 1, x_42); +x_55 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_55, 0, x_53); +lean_ctor_set(x_55, 1, x_54); +x_56 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_56, 0, x_55); +x_22 = x_56; +x_23 = x_51; +goto block_28; +} +else +{ +uint8_t x_57; +lean_dec(x_46); +lean_free_object(x_42); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_57 = !lean_is_exclusive(x_49); +if (x_57 == 0) +{ +return x_49; +} +else +{ +lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_58 = lean_ctor_get(x_49, 0); +x_59 = lean_ctor_get(x_49, 1); +lean_inc(x_59); +lean_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; +} +} +} +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; +x_61 = lean_ctor_get(x_42, 0); +x_62 = lean_ctor_get(x_42, 1); +lean_inc(x_62); +lean_inc(x_61); +lean_dec(x_42); +lean_inc(x_61); +x_63 = l_Lean_Expr_app___override(x_33, x_61); +x_64 = l_Lean_Expr_mvarId_x21(x_61); +lean_dec(x_61); +lean_inc(x_4); +lean_inc(x_5); +x_65 = lean_array_push(x_5, x_4); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_66 = l_Lean_MVarId_tryClearMany(x_64, x_65, x_14, x_15, x_16, x_17, x_62); +lean_dec(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; +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_array_push(x_31, x_67); +x_70 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +x_71 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_71, 0, x_63); +lean_ctor_set(x_71, 1, x_39); +x_72 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_72, 0, x_69); +lean_ctor_set(x_72, 1, x_71); +x_73 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_73, 0, x_70); +lean_ctor_set(x_73, 1, x_72); +x_74 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_74, 0, x_73); +x_22 = x_74; +x_23 = x_68; +goto block_28; +} +else +{ +lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; +lean_dec(x_63); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_75 = lean_ctor_get(x_66, 0); +lean_inc(x_75); +x_76 = lean_ctor_get(x_66, 1); +lean_inc(x_76); +if (lean_is_exclusive(x_66)) { + lean_ctor_release(x_66, 0); + lean_ctor_release(x_66, 1); + x_77 = x_66; +} else { + lean_dec_ref(x_66); + x_77 = lean_box(0); +} +if (lean_is_scalar(x_77)) { + x_78 = lean_alloc_ctor(1, 2, 0); +} else { + x_78 = x_77; +} +lean_ctor_set(x_78, 0, x_75); +lean_ctor_set(x_78, 1, x_76); +return x_78; +} +} +} +else +{ +lean_object* x_79; lean_object* x_80; uint8_t x_81; +lean_inc(x_29); +lean_inc(x_1); +x_79 = l_Lean_Name_num___override(x_1, x_29); +lean_inc(x_14); +x_80 = l_Lean_Meta_mkFreshExprSyntheticOpaqueMVar(x_38, x_79, x_14, x_15, x_16, x_17, x_37); +x_81 = !lean_is_exclusive(x_80); +if (x_81 == 0) +{ +lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; +x_82 = lean_ctor_get(x_80, 0); +x_83 = lean_ctor_get(x_80, 1); +lean_inc(x_82); +x_84 = l_Lean_Expr_app___override(x_33, x_82); +x_85 = l_Lean_Expr_mvarId_x21(x_82); +lean_dec(x_82); +lean_inc(x_4); +lean_inc(x_5); +x_86 = lean_array_push(x_5, x_4); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_87 = l_Lean_MVarId_tryClearMany(x_85, x_86, x_14, x_15, x_16, x_17, x_83); +lean_dec(x_86); +if (lean_obj_tag(x_87) == 0) +{ +lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; +x_88 = lean_ctor_get(x_87, 0); +lean_inc(x_88); +x_89 = lean_ctor_get(x_87, 1); +lean_inc(x_89); +lean_dec(x_87); +x_90 = lean_array_push(x_31, x_88); +x_91 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +lean_ctor_set(x_80, 1, x_39); +lean_ctor_set(x_80, 0, x_84); +x_92 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_92, 0, x_90); +lean_ctor_set(x_92, 1, x_80); +x_93 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_93, 0, x_91); +lean_ctor_set(x_93, 1, x_92); +x_94 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_94, 0, x_93); +x_22 = x_94; +x_23 = x_89; +goto block_28; +} +else +{ +uint8_t x_95; +lean_dec(x_84); +lean_free_object(x_80); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_95 = !lean_is_exclusive(x_87); +if (x_95 == 0) +{ +return x_87; +} +else +{ +lean_object* x_96; lean_object* x_97; lean_object* x_98; +x_96 = lean_ctor_get(x_87, 0); +x_97 = lean_ctor_get(x_87, 1); +lean_inc(x_97); +lean_inc(x_96); +lean_dec(x_87); +x_98 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_98, 0, x_96); +lean_ctor_set(x_98, 1, x_97); +return x_98; +} +} +} +else +{ +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_80, 0); +x_100 = lean_ctor_get(x_80, 1); +lean_inc(x_100); +lean_inc(x_99); +lean_dec(x_80); +lean_inc(x_99); +x_101 = l_Lean_Expr_app___override(x_33, x_99); +x_102 = l_Lean_Expr_mvarId_x21(x_99); +lean_dec(x_99); +lean_inc(x_4); +lean_inc(x_5); +x_103 = lean_array_push(x_5, x_4); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_104 = l_Lean_MVarId_tryClearMany(x_102, x_103, x_14, x_15, x_16, x_17, x_100); +lean_dec(x_103); +if (lean_obj_tag(x_104) == 0) +{ +lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; +x_105 = lean_ctor_get(x_104, 0); +lean_inc(x_105); +x_106 = lean_ctor_get(x_104, 1); +lean_inc(x_106); +lean_dec(x_104); +x_107 = lean_array_push(x_31, x_105); +x_108 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +x_109 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_109, 0, x_101); +lean_ctor_set(x_109, 1, x_39); +x_110 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_110, 0, x_107); +lean_ctor_set(x_110, 1, x_109); +x_111 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_111, 0, x_108); +lean_ctor_set(x_111, 1, x_110); +x_112 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_112, 0, x_111); +x_22 = x_112; +x_23 = x_106; +goto block_28; +} +else +{ +lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; +lean_dec(x_101); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_113 = lean_ctor_get(x_104, 0); +lean_inc(x_113); +x_114 = lean_ctor_get(x_104, 1); +lean_inc(x_114); +if (lean_is_exclusive(x_104)) { + lean_ctor_release(x_104, 0); + lean_ctor_release(x_104, 1); + x_115 = x_104; +} else { + lean_dec_ref(x_104); + x_115 = lean_box(0); +} +if (lean_is_scalar(x_115)) { + x_116 = lean_alloc_ctor(1, 2, 0); +} else { + x_116 = x_115; +} +lean_ctor_set(x_116, 0, x_113); +lean_ctor_set(x_116, 1, x_114); +return x_116; +} +} +} +} +else +{ +lean_object* x_117; lean_object* x_118; lean_object* x_119; uint8_t x_120; +lean_dec(x_36); +lean_dec(x_33); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_11); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_117 = lean_ctor_get(x_35, 1); +lean_inc(x_117); +lean_dec(x_35); +x_118 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__1___closed__4; +x_119 = l_Lean_Meta_throwTacticEx___rarg(x_6, x_3, x_118, x_14, x_15, x_16, x_17, x_117); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +x_120 = !lean_is_exclusive(x_119); +if (x_120 == 0) +{ +return x_119; +} +else +{ +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_124; +lean_dec(x_33); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_124 = !lean_is_exclusive(x_35); +if (x_124 == 0) +{ +return x_35; +} +else +{ +lean_object* x_125; lean_object* x_126; lean_object* x_127; +x_125 = lean_ctor_get(x_35, 0); +x_126 = lean_ctor_get(x_35, 1); +lean_inc(x_126); +lean_inc(x_125); +lean_dec(x_35); +x_127 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_127, 0, x_125); +lean_ctor_set(x_127, 1, x_126); +return x_127; +} +} +block_28: +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_22, 0); +lean_inc(x_24); +lean_dec(x_22); +x_25 = lean_ctor_get(x_9, 2); +x_26 = lean_nat_add(x_11, x_25); +lean_dec(x_11); +x_10 = x_24; +x_11 = x_26; +x_12 = lean_box(0); +x_13 = lean_box(0); +x_18 = x_23; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_13 = l_Lean_Meta_mkRecursorAppPrefix(x_1, x_2, x_3, x_4, x_5, 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; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +x_16 = l_Lean_mkAppN(x_14, x_5); +lean_inc(x_3); +x_17 = l_Lean_Expr_fvar___override(x_3); +x_18 = l_Lean_Expr_app___override(x_16, x_17); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_18); +x_19 = lean_infer_type(x_18, x_8, x_9, x_10, x_11, x_15); +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; +x_20 = lean_ctor_get(x_19, 0); +lean_inc(x_20); +x_21 = lean_ctor_get(x_19, 1); +lean_inc(x_21); +lean_dec(x_19); +x_22 = l_Lean_Meta_RecursorInfo_numMinors(x_4); +x_23 = lean_unsigned_to_nat(0u); +x_24 = lean_unsigned_to_nat(1u); +lean_inc(x_22); +x_25 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_22); +lean_ctor_set(x_25, 2, x_24); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_18); +lean_ctor_set(x_26, 1, x_20); +lean_inc(x_6); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_6); +lean_ctor_set(x_27, 1, x_26); +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_24); +lean_ctor_set(x_28, 1, x_27); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_1); +x_29 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__1(x_7, x_4, x_3, x_1, x_6, x_2, x_22, x_25, x_25, x_28, x_23, lean_box(0), lean_box(0), x_8, x_9, x_10, x_11, x_21); +lean_dec(x_25); +lean_dec(x_22); +lean_dec(x_4); +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; uint8_t x_37; +x_30 = lean_ctor_get(x_29, 0); +lean_inc(x_30); +x_31 = lean_ctor_get(x_30, 1); +lean_inc(x_31); +lean_dec(x_30); +x_32 = lean_ctor_get(x_31, 1); +lean_inc(x_32); +x_33 = lean_ctor_get(x_29, 1); +lean_inc(x_33); +lean_dec(x_29); +x_34 = lean_ctor_get(x_31, 0); +lean_inc(x_34); +lean_dec(x_31); +x_35 = lean_ctor_get(x_32, 0); +lean_inc(x_35); +lean_dec(x_32); +x_36 = l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_1, x_35, 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); +x_37 = !lean_is_exclusive(x_36); +if (x_37 == 0) +{ +lean_object* x_38; lean_object* x_39; +x_38 = lean_ctor_get(x_36, 0); +lean_dec(x_38); +x_39 = lean_array_to_list(x_34); +lean_ctor_set(x_36, 0, x_39); +return x_36; +} +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_array_to_list(x_34); +x_42 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_40); +return x_42; +} +} +else +{ +uint8_t x_43; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_1); +x_43 = !lean_is_exclusive(x_29); +if (x_43 == 0) +{ +return x_29; +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_29, 0); +x_45 = lean_ctor_get(x_29, 1); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_29); +x_46 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_46, 0, x_44); +lean_ctor_set(x_46, 1, x_45); +return x_46; +} +} +} +else +{ +uint8_t x_47; +lean_dec(x_18); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_47 = !lean_is_exclusive(x_19); +if (x_47 == 0) +{ +return x_19; +} +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_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; +} +} +} +else +{ +uint8_t x_51; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_51 = !lean_is_exclusive(x_13); +if (x_51 == 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_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; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_13 = l_Lean_Meta_mkRecursorAppPrefix(x_1, x_2, x_3, x_4, x_5, 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; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +x_16 = l_Lean_mkAppN(x_14, x_5); +lean_inc(x_3); +x_17 = l_Lean_Expr_fvar___override(x_3); +x_18 = l_Lean_Expr_app___override(x_16, x_17); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_18); +x_19 = lean_infer_type(x_18, x_8, x_9, x_10, x_11, x_15); +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; +x_20 = lean_ctor_get(x_19, 0); +lean_inc(x_20); +x_21 = lean_ctor_get(x_19, 1); +lean_inc(x_21); +lean_dec(x_19); +x_22 = l_Lean_Meta_RecursorInfo_numMinors(x_4); +x_23 = lean_unsigned_to_nat(0u); +x_24 = lean_unsigned_to_nat(1u); +lean_inc(x_22); +x_25 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_22); +lean_ctor_set(x_25, 2, x_24); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_18); +lean_ctor_set(x_26, 1, x_20); +lean_inc(x_6); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_6); +lean_ctor_set(x_27, 1, x_26); +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_24); +lean_ctor_set(x_28, 1, x_27); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_1); +x_29 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__3(x_7, x_4, x_3, x_1, x_6, x_2, x_22, x_25, x_25, x_28, x_23, lean_box(0), lean_box(0), x_8, x_9, x_10, x_11, x_21); +lean_dec(x_25); +lean_dec(x_22); +lean_dec(x_4); +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; uint8_t x_37; +x_30 = lean_ctor_get(x_29, 0); +lean_inc(x_30); +x_31 = lean_ctor_get(x_30, 1); +lean_inc(x_31); +lean_dec(x_30); +x_32 = lean_ctor_get(x_31, 1); +lean_inc(x_32); +x_33 = lean_ctor_get(x_29, 1); +lean_inc(x_33); +lean_dec(x_29); +x_34 = lean_ctor_get(x_31, 0); +lean_inc(x_34); +lean_dec(x_31); +x_35 = lean_ctor_get(x_32, 0); +lean_inc(x_35); +lean_dec(x_32); +x_36 = l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_1, x_35, 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); +x_37 = !lean_is_exclusive(x_36); +if (x_37 == 0) +{ +lean_object* x_38; lean_object* x_39; +x_38 = lean_ctor_get(x_36, 0); +lean_dec(x_38); +x_39 = lean_array_to_list(x_34); +lean_ctor_set(x_36, 0, x_39); +return x_36; +} +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_array_to_list(x_34); +x_42 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_40); +return x_42; +} +} +else +{ +uint8_t x_43; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_1); +x_43 = !lean_is_exclusive(x_29); +if (x_43 == 0) +{ +return x_29; +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_29, 0); +x_45 = lean_ctor_get(x_29, 1); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_29); +x_46 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_46, 0, x_44); +lean_ctor_set(x_46, 1, x_45); +return x_46; +} +} +} +else +{ +uint8_t x_47; +lean_dec(x_18); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_47 = !lean_is_exclusive(x_19); +if (x_47 == 0) +{ +return x_19; +} +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_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; +} +} +} +else +{ +uint8_t x_51; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_51 = !lean_is_exclusive(x_13); +if (x_51 == 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_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; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___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; +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_13 = l_Lean_Meta_mkRecursorAppPrefix(x_1, x_2, x_3, x_4, x_5, 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; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +x_16 = l_Lean_mkAppN(x_14, x_5); +lean_inc(x_3); +x_17 = l_Lean_Expr_fvar___override(x_3); +x_18 = l_Lean_Expr_app___override(x_16, x_17); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_18); +x_19 = lean_infer_type(x_18, x_8, x_9, x_10, x_11, x_15); +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; +x_20 = lean_ctor_get(x_19, 0); +lean_inc(x_20); +x_21 = lean_ctor_get(x_19, 1); +lean_inc(x_21); +lean_dec(x_19); +x_22 = l_Lean_Meta_RecursorInfo_numMinors(x_4); +x_23 = lean_unsigned_to_nat(0u); +x_24 = lean_unsigned_to_nat(1u); +lean_inc(x_22); +x_25 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_22); +lean_ctor_set(x_25, 2, x_24); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_18); +lean_ctor_set(x_26, 1, x_20); +lean_inc(x_6); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_6); +lean_ctor_set(x_27, 1, x_26); +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_24); +lean_ctor_set(x_28, 1, x_27); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_1); +x_29 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__4(x_7, x_4, x_3, x_1, x_6, x_2, x_22, x_25, x_25, x_28, x_23, lean_box(0), lean_box(0), x_8, x_9, x_10, x_11, x_21); +lean_dec(x_25); +lean_dec(x_22); +lean_dec(x_4); +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; uint8_t x_37; +x_30 = lean_ctor_get(x_29, 0); +lean_inc(x_30); +x_31 = lean_ctor_get(x_30, 1); +lean_inc(x_31); +lean_dec(x_30); +x_32 = lean_ctor_get(x_31, 1); +lean_inc(x_32); +x_33 = lean_ctor_get(x_29, 1); +lean_inc(x_33); +lean_dec(x_29); +x_34 = lean_ctor_get(x_31, 0); +lean_inc(x_34); +lean_dec(x_31); +x_35 = lean_ctor_get(x_32, 0); +lean_inc(x_35); +lean_dec(x_32); +x_36 = l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_1, x_35, 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); +x_37 = !lean_is_exclusive(x_36); +if (x_37 == 0) +{ +lean_object* x_38; lean_object* x_39; +x_38 = lean_ctor_get(x_36, 0); +lean_dec(x_38); +x_39 = lean_array_to_list(x_34); +lean_ctor_set(x_36, 0, x_39); +return x_36; +} +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_array_to_list(x_34); +x_42 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_40); +return x_42; +} +} +else +{ +uint8_t x_43; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_1); +x_43 = !lean_is_exclusive(x_29); +if (x_43 == 0) +{ +return x_29; +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_29, 0); +x_45 = lean_ctor_get(x_29, 1); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_29); +x_46 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_46, 0, x_44); +lean_ctor_set(x_46, 1, x_45); +return x_46; +} +} +} +else +{ +uint8_t x_47; +lean_dec(x_18); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_47 = !lean_is_exclusive(x_19); +if (x_47 == 0) +{ +return x_19; +} +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_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; +} +} +} +else +{ +uint8_t x_51; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_51 = !lean_is_exclusive(x_13); +if (x_51 == 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_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; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___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) { +_start: +{ +lean_object* x_13; +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_13 = l_Lean_Meta_mkRecursorAppPrefix(x_1, x_2, x_3, x_4, x_5, 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; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +x_16 = l_Lean_mkAppN(x_14, x_5); +lean_inc(x_3); +x_17 = l_Lean_Expr_fvar___override(x_3); +x_18 = l_Lean_Expr_app___override(x_16, x_17); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_18); +x_19 = lean_infer_type(x_18, x_8, x_9, x_10, x_11, x_15); +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; +x_20 = lean_ctor_get(x_19, 0); +lean_inc(x_20); +x_21 = lean_ctor_get(x_19, 1); +lean_inc(x_21); +lean_dec(x_19); +x_22 = l_Lean_Meta_RecursorInfo_numMinors(x_4); +x_23 = lean_unsigned_to_nat(0u); +x_24 = lean_unsigned_to_nat(1u); +lean_inc(x_22); +x_25 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_22); +lean_ctor_set(x_25, 2, x_24); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_18); +lean_ctor_set(x_26, 1, x_20); +lean_inc(x_6); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_6); +lean_ctor_set(x_27, 1, x_26); +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_24); +lean_ctor_set(x_28, 1, x_27); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_1); +x_29 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__5(x_7, x_4, x_3, x_1, x_6, x_2, x_22, x_25, x_25, x_28, x_23, lean_box(0), lean_box(0), x_8, x_9, x_10, x_11, x_21); +lean_dec(x_25); +lean_dec(x_22); +lean_dec(x_4); +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; uint8_t x_37; +x_30 = lean_ctor_get(x_29, 0); +lean_inc(x_30); +x_31 = lean_ctor_get(x_30, 1); +lean_inc(x_31); +lean_dec(x_30); +x_32 = lean_ctor_get(x_31, 1); +lean_inc(x_32); +x_33 = lean_ctor_get(x_29, 1); +lean_inc(x_33); +lean_dec(x_29); +x_34 = lean_ctor_get(x_31, 0); +lean_inc(x_34); +lean_dec(x_31); +x_35 = lean_ctor_get(x_32, 0); +lean_inc(x_35); +lean_dec(x_32); +x_36 = l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_1, x_35, 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); +x_37 = !lean_is_exclusive(x_36); +if (x_37 == 0) +{ +lean_object* x_38; lean_object* x_39; +x_38 = lean_ctor_get(x_36, 0); +lean_dec(x_38); +x_39 = lean_array_to_list(x_34); +lean_ctor_set(x_36, 0, x_39); +return x_36; +} +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_array_to_list(x_34); +x_42 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_40); +return x_42; +} +} +else +{ +uint8_t x_43; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_1); +x_43 = !lean_is_exclusive(x_29); +if (x_43 == 0) +{ +return x_29; +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_29, 0); +x_45 = lean_ctor_get(x_29, 1); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_29); +x_46 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_46, 0, x_44); +lean_ctor_set(x_46, 1, x_45); +return x_46; +} +} +} +else +{ +uint8_t x_47; +lean_dec(x_18); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_47 = !lean_is_exclusive(x_19); +if (x_47 == 0) +{ +return x_19; +} +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_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; +} +} +} +else +{ +uint8_t x_51; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_51 = !lean_is_exclusive(x_13); +if (x_51 == 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_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; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___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) { +_start: +{ +lean_object* x_13; +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_13 = l_Lean_Meta_mkRecursorAppPrefix(x_1, x_2, x_3, x_4, x_5, 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; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +x_16 = l_Lean_mkAppN(x_14, x_5); +lean_inc(x_3); +x_17 = l_Lean_Expr_fvar___override(x_3); +x_18 = l_Lean_Expr_app___override(x_16, x_17); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_18); +x_19 = lean_infer_type(x_18, x_8, x_9, x_10, x_11, x_15); +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; +x_20 = lean_ctor_get(x_19, 0); +lean_inc(x_20); +x_21 = lean_ctor_get(x_19, 1); +lean_inc(x_21); +lean_dec(x_19); +x_22 = l_Lean_Meta_RecursorInfo_numMinors(x_4); +x_23 = lean_unsigned_to_nat(0u); +x_24 = lean_unsigned_to_nat(1u); +lean_inc(x_22); +x_25 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_22); +lean_ctor_set(x_25, 2, x_24); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_18); +lean_ctor_set(x_26, 1, x_20); +lean_inc(x_6); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_6); +lean_ctor_set(x_27, 1, x_26); +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_24); +lean_ctor_set(x_28, 1, x_27); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_1); +x_29 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__6(x_7, x_4, x_3, x_1, x_6, x_2, x_22, x_25, x_25, x_28, x_23, lean_box(0), lean_box(0), x_8, x_9, x_10, x_11, x_21); +lean_dec(x_25); +lean_dec(x_22); +lean_dec(x_4); +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; uint8_t x_37; +x_30 = lean_ctor_get(x_29, 0); +lean_inc(x_30); +x_31 = lean_ctor_get(x_30, 1); +lean_inc(x_31); +lean_dec(x_30); +x_32 = lean_ctor_get(x_31, 1); +lean_inc(x_32); +x_33 = lean_ctor_get(x_29, 1); +lean_inc(x_33); +lean_dec(x_29); +x_34 = lean_ctor_get(x_31, 0); +lean_inc(x_34); +lean_dec(x_31); +x_35 = lean_ctor_get(x_32, 0); +lean_inc(x_35); +lean_dec(x_32); +x_36 = l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_1, x_35, 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); +x_37 = !lean_is_exclusive(x_36); +if (x_37 == 0) +{ +lean_object* x_38; lean_object* x_39; +x_38 = lean_ctor_get(x_36, 0); +lean_dec(x_38); +x_39 = lean_array_to_list(x_34); +lean_ctor_set(x_36, 0, x_39); +return x_36; +} +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_array_to_list(x_34); +x_42 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_40); +return x_42; +} +} +else +{ +uint8_t x_43; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_1); +x_43 = !lean_is_exclusive(x_29); +if (x_43 == 0) +{ +return x_29; +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_29, 0); +x_45 = lean_ctor_get(x_29, 1); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_29); +x_46 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_46, 0, x_44); +lean_ctor_set(x_46, 1, x_45); +return x_46; +} +} +} +else +{ +uint8_t x_47; +lean_dec(x_18); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_47 = !lean_is_exclusive(x_19); +if (x_47 == 0) +{ +return x_19; +} +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_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; +} +} +} +else +{ +uint8_t x_51; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_51 = !lean_is_exclusive(x_13); +if (x_51 == 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_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; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___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) { +_start: +{ +lean_object* x_13; +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_13 = l_Lean_Meta_mkRecursorAppPrefix(x_1, x_2, x_3, x_4, x_5, 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; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +x_16 = l_Lean_mkAppN(x_14, x_5); +lean_inc(x_3); +x_17 = l_Lean_Expr_fvar___override(x_3); +x_18 = l_Lean_Expr_app___override(x_16, x_17); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_18); +x_19 = lean_infer_type(x_18, x_8, x_9, x_10, x_11, x_15); +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; +x_20 = lean_ctor_get(x_19, 0); +lean_inc(x_20); +x_21 = lean_ctor_get(x_19, 1); +lean_inc(x_21); +lean_dec(x_19); +x_22 = l_Lean_Meta_RecursorInfo_numMinors(x_4); +x_23 = lean_unsigned_to_nat(0u); +x_24 = lean_unsigned_to_nat(1u); +lean_inc(x_22); +x_25 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_22); +lean_ctor_set(x_25, 2, x_24); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_18); +lean_ctor_set(x_26, 1, x_20); +lean_inc(x_6); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_6); +lean_ctor_set(x_27, 1, x_26); +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_24); +lean_ctor_set(x_28, 1, x_27); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_1); +x_29 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__7(x_7, x_4, x_3, x_1, x_6, x_2, x_22, x_25, x_25, x_28, x_23, lean_box(0), lean_box(0), x_8, x_9, x_10, x_11, x_21); +lean_dec(x_25); +lean_dec(x_22); +lean_dec(x_4); +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; uint8_t x_37; +x_30 = lean_ctor_get(x_29, 0); +lean_inc(x_30); +x_31 = lean_ctor_get(x_30, 1); +lean_inc(x_31); +lean_dec(x_30); +x_32 = lean_ctor_get(x_31, 1); +lean_inc(x_32); +x_33 = lean_ctor_get(x_29, 1); +lean_inc(x_33); +lean_dec(x_29); +x_34 = lean_ctor_get(x_31, 0); +lean_inc(x_34); +lean_dec(x_31); +x_35 = lean_ctor_get(x_32, 0); +lean_inc(x_35); +lean_dec(x_32); +x_36 = l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_1, x_35, 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); +x_37 = !lean_is_exclusive(x_36); +if (x_37 == 0) +{ +lean_object* x_38; lean_object* x_39; +x_38 = lean_ctor_get(x_36, 0); +lean_dec(x_38); +x_39 = lean_array_to_list(x_34); +lean_ctor_set(x_36, 0, x_39); +return x_36; +} +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_array_to_list(x_34); +x_42 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_40); +return x_42; +} +} +else +{ +uint8_t x_43; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_1); +x_43 = !lean_is_exclusive(x_29); +if (x_43 == 0) +{ +return x_29; +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_29, 0); +x_45 = lean_ctor_get(x_29, 1); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_29); +x_46 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_46, 0, x_44); +lean_ctor_set(x_46, 1, x_45); +return x_46; +} +} +} +else +{ +uint8_t x_47; +lean_dec(x_18); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_47 = !lean_is_exclusive(x_19); +if (x_47 == 0) +{ +return x_19; +} +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_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; +} +} +} +else +{ +uint8_t x_51; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_51 = !lean_is_exclusive(x_13); +if (x_51 == 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_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; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___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_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_13 = l_Lean_Meta_mkRecursorAppPrefix(x_1, x_2, x_3, x_4, x_5, 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; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +x_16 = l_Lean_mkAppN(x_14, x_5); +lean_inc(x_3); +x_17 = l_Lean_Expr_fvar___override(x_3); +x_18 = l_Lean_Expr_app___override(x_16, x_17); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_18); +x_19 = lean_infer_type(x_18, x_8, x_9, x_10, x_11, x_15); +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; +x_20 = lean_ctor_get(x_19, 0); +lean_inc(x_20); +x_21 = lean_ctor_get(x_19, 1); +lean_inc(x_21); +lean_dec(x_19); +x_22 = l_Lean_Meta_RecursorInfo_numMinors(x_4); +x_23 = lean_unsigned_to_nat(0u); +x_24 = lean_unsigned_to_nat(1u); +lean_inc(x_22); +x_25 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_22); +lean_ctor_set(x_25, 2, x_24); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_18); +lean_ctor_set(x_26, 1, x_20); +lean_inc(x_6); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_6); +lean_ctor_set(x_27, 1, x_26); +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_24); +lean_ctor_set(x_28, 1, x_27); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_1); +x_29 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__8(x_7, x_4, x_3, x_1, x_6, x_2, x_22, x_25, x_25, x_28, x_23, lean_box(0), lean_box(0), x_8, x_9, x_10, x_11, x_21); +lean_dec(x_25); +lean_dec(x_22); +lean_dec(x_4); +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; uint8_t x_37; +x_30 = lean_ctor_get(x_29, 0); +lean_inc(x_30); +x_31 = lean_ctor_get(x_30, 1); +lean_inc(x_31); +lean_dec(x_30); +x_32 = lean_ctor_get(x_31, 1); +lean_inc(x_32); +x_33 = lean_ctor_get(x_29, 1); +lean_inc(x_33); +lean_dec(x_29); +x_34 = lean_ctor_get(x_31, 0); +lean_inc(x_34); +lean_dec(x_31); +x_35 = lean_ctor_get(x_32, 0); +lean_inc(x_35); +lean_dec(x_32); +x_36 = l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_1, x_35, 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); +x_37 = !lean_is_exclusive(x_36); +if (x_37 == 0) +{ +lean_object* x_38; lean_object* x_39; +x_38 = lean_ctor_get(x_36, 0); +lean_dec(x_38); +x_39 = lean_array_to_list(x_34); +lean_ctor_set(x_36, 0, x_39); +return x_36; +} +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_array_to_list(x_34); +x_42 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_40); +return x_42; +} +} +else +{ +uint8_t x_43; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_1); +x_43 = !lean_is_exclusive(x_29); +if (x_43 == 0) +{ +return x_29; +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_29, 0); +x_45 = lean_ctor_get(x_29, 1); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_29); +x_46 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_46, 0, x_44); +lean_ctor_set(x_46, 1, x_45); +return x_46; +} +} +} +else +{ +uint8_t x_47; +lean_dec(x_18); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_47 = !lean_is_exclusive(x_19); +if (x_47 == 0) +{ +return x_19; +} +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_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; +} +} +} +else +{ +uint8_t x_51; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_51 = !lean_is_exclusive(x_13); +if (x_51 == 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_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; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___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_13; +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_13 = l_Lean_Meta_mkRecursorAppPrefix(x_1, x_2, x_3, x_4, x_5, 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; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +x_16 = l_Lean_mkAppN(x_14, x_5); +lean_inc(x_3); +x_17 = l_Lean_Expr_fvar___override(x_3); +x_18 = l_Lean_Expr_app___override(x_16, x_17); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_18); +x_19 = lean_infer_type(x_18, x_8, x_9, x_10, x_11, x_15); +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; +x_20 = lean_ctor_get(x_19, 0); +lean_inc(x_20); +x_21 = lean_ctor_get(x_19, 1); +lean_inc(x_21); +lean_dec(x_19); +x_22 = l_Lean_Meta_RecursorInfo_numMinors(x_4); +x_23 = lean_unsigned_to_nat(0u); +x_24 = lean_unsigned_to_nat(1u); +lean_inc(x_22); +x_25 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_22); +lean_ctor_set(x_25, 2, x_24); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_18); +lean_ctor_set(x_26, 1, x_20); +lean_inc(x_6); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_6); +lean_ctor_set(x_27, 1, x_26); +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_24); +lean_ctor_set(x_28, 1, x_27); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_1); +x_29 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__9(x_7, x_4, x_3, x_1, x_6, x_2, x_22, x_25, x_25, x_28, x_23, lean_box(0), lean_box(0), x_8, x_9, x_10, x_11, x_21); +lean_dec(x_25); +lean_dec(x_22); +lean_dec(x_4); +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; uint8_t x_37; +x_30 = lean_ctor_get(x_29, 0); +lean_inc(x_30); +x_31 = lean_ctor_get(x_30, 1); +lean_inc(x_31); +lean_dec(x_30); +x_32 = lean_ctor_get(x_31, 1); +lean_inc(x_32); +x_33 = lean_ctor_get(x_29, 1); +lean_inc(x_33); +lean_dec(x_29); +x_34 = lean_ctor_get(x_31, 0); +lean_inc(x_34); +lean_dec(x_31); +x_35 = lean_ctor_get(x_32, 0); +lean_inc(x_35); +lean_dec(x_32); +x_36 = l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_1, x_35, 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); +x_37 = !lean_is_exclusive(x_36); +if (x_37 == 0) +{ +lean_object* x_38; lean_object* x_39; +x_38 = lean_ctor_get(x_36, 0); +lean_dec(x_38); +x_39 = lean_array_to_list(x_34); +lean_ctor_set(x_36, 0, x_39); +return x_36; +} +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_array_to_list(x_34); +x_42 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_40); +return x_42; +} +} +else +{ +uint8_t x_43; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_1); +x_43 = !lean_is_exclusive(x_29); +if (x_43 == 0) +{ +return x_29; +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_29, 0); +x_45 = lean_ctor_get(x_29, 1); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_29); +x_46 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_46, 0, x_44); +lean_ctor_set(x_46, 1, x_45); +return x_46; +} +} +} +else +{ +uint8_t x_47; +lean_dec(x_18); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_47 = !lean_is_exclusive(x_19); +if (x_47 == 0) +{ +return x_19; +} +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_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; +} +} +} +else +{ +uint8_t x_51; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_51 = !lean_is_exclusive(x_13); +if (x_51 == 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_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; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___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_13; +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_13 = l_Lean_Meta_mkRecursorAppPrefix(x_1, x_2, x_3, x_4, x_5, 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; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +x_16 = l_Lean_mkAppN(x_14, x_5); +lean_inc(x_3); +x_17 = l_Lean_Expr_fvar___override(x_3); +x_18 = l_Lean_Expr_app___override(x_16, x_17); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_18); +x_19 = lean_infer_type(x_18, x_8, x_9, x_10, x_11, x_15); +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; +x_20 = lean_ctor_get(x_19, 0); +lean_inc(x_20); +x_21 = lean_ctor_get(x_19, 1); +lean_inc(x_21); +lean_dec(x_19); +x_22 = l_Lean_Meta_RecursorInfo_numMinors(x_4); +x_23 = lean_unsigned_to_nat(0u); +x_24 = lean_unsigned_to_nat(1u); +lean_inc(x_22); +x_25 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_22); +lean_ctor_set(x_25, 2, x_24); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_18); +lean_ctor_set(x_26, 1, x_20); +lean_inc(x_6); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_6); +lean_ctor_set(x_27, 1, x_26); +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_24); +lean_ctor_set(x_28, 1, x_27); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_1); +x_29 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__10(x_7, x_4, x_3, x_1, x_6, x_2, x_22, x_25, x_25, x_28, x_23, lean_box(0), lean_box(0), x_8, x_9, x_10, x_11, x_21); +lean_dec(x_25); +lean_dec(x_22); +lean_dec(x_4); +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; uint8_t x_37; +x_30 = lean_ctor_get(x_29, 0); +lean_inc(x_30); +x_31 = lean_ctor_get(x_30, 1); +lean_inc(x_31); +lean_dec(x_30); +x_32 = lean_ctor_get(x_31, 1); +lean_inc(x_32); +x_33 = lean_ctor_get(x_29, 1); +lean_inc(x_33); +lean_dec(x_29); +x_34 = lean_ctor_get(x_31, 0); +lean_inc(x_34); +lean_dec(x_31); +x_35 = lean_ctor_get(x_32, 0); +lean_inc(x_35); +lean_dec(x_32); +x_36 = l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_1, x_35, 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); +x_37 = !lean_is_exclusive(x_36); +if (x_37 == 0) +{ +lean_object* x_38; lean_object* x_39; +x_38 = lean_ctor_get(x_36, 0); +lean_dec(x_38); +x_39 = lean_array_to_list(x_34); +lean_ctor_set(x_36, 0, x_39); +return x_36; +} +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_array_to_list(x_34); +x_42 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_40); +return x_42; +} +} +else +{ +uint8_t x_43; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_1); +x_43 = !lean_is_exclusive(x_29); +if (x_43 == 0) +{ +return x_29; +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_29, 0); +x_45 = lean_ctor_get(x_29, 1); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_29); +x_46 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_46, 0, x_44); +lean_ctor_set(x_46, 1, x_45); +return x_46; +} +} +} +else +{ +uint8_t x_47; +lean_dec(x_18); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_47 = !lean_is_exclusive(x_19); +if (x_47 == 0) +{ +return x_19; +} +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_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; +} +} +} +else +{ +uint8_t x_51; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_51 = !lean_is_exclusive(x_13); +if (x_51 == 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_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; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___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) { +_start: +{ +lean_object* x_13; +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_13 = l_Lean_Meta_mkRecursorAppPrefix(x_1, x_2, x_3, x_4, x_5, 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; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +x_16 = l_Lean_mkAppN(x_14, x_5); +lean_inc(x_3); +x_17 = l_Lean_Expr_fvar___override(x_3); +x_18 = l_Lean_Expr_app___override(x_16, x_17); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_18); +x_19 = lean_infer_type(x_18, x_8, x_9, x_10, x_11, x_15); +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; +x_20 = lean_ctor_get(x_19, 0); +lean_inc(x_20); +x_21 = lean_ctor_get(x_19, 1); +lean_inc(x_21); +lean_dec(x_19); +x_22 = l_Lean_Meta_RecursorInfo_numMinors(x_4); +x_23 = lean_unsigned_to_nat(0u); +x_24 = lean_unsigned_to_nat(1u); +lean_inc(x_22); +x_25 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_22); +lean_ctor_set(x_25, 2, x_24); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_18); +lean_ctor_set(x_26, 1, x_20); +lean_inc(x_6); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_6); +lean_ctor_set(x_27, 1, x_26); +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_24); +lean_ctor_set(x_28, 1, x_27); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_1); +x_29 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__11(x_7, x_4, x_3, x_1, x_6, x_2, x_22, x_25, x_25, x_28, x_23, lean_box(0), lean_box(0), x_8, x_9, x_10, x_11, x_21); +lean_dec(x_25); +lean_dec(x_22); +lean_dec(x_4); +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; uint8_t x_37; +x_30 = lean_ctor_get(x_29, 0); +lean_inc(x_30); +x_31 = lean_ctor_get(x_30, 1); +lean_inc(x_31); +lean_dec(x_30); +x_32 = lean_ctor_get(x_31, 1); +lean_inc(x_32); +x_33 = lean_ctor_get(x_29, 1); +lean_inc(x_33); +lean_dec(x_29); +x_34 = lean_ctor_get(x_31, 0); +lean_inc(x_34); +lean_dec(x_31); +x_35 = lean_ctor_get(x_32, 0); +lean_inc(x_35); +lean_dec(x_32); +x_36 = l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_1, x_35, 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); +x_37 = !lean_is_exclusive(x_36); +if (x_37 == 0) +{ +lean_object* x_38; lean_object* x_39; +x_38 = lean_ctor_get(x_36, 0); +lean_dec(x_38); +x_39 = lean_array_to_list(x_34); +lean_ctor_set(x_36, 0, x_39); +return x_36; +} +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_array_to_list(x_34); +x_42 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_40); +return x_42; +} +} +else +{ +uint8_t x_43; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_1); +x_43 = !lean_is_exclusive(x_29); +if (x_43 == 0) +{ +return x_29; +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_29, 0); +x_45 = lean_ctor_get(x_29, 1); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_29); +x_46 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_46, 0, x_44); +lean_ctor_set(x_46, 1, x_45); +return x_46; +} +} +} +else +{ +uint8_t x_47; +lean_dec(x_18); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_47 = !lean_is_exclusive(x_19); +if (x_47 == 0) +{ +return x_19; +} +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_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; +} +} +} +else +{ +uint8_t x_51; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_51 = !lean_is_exclusive(x_13); +if (x_51 == 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_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; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___lambda__11(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_13 = l_Lean_Meta_mkRecursorAppPrefix(x_1, x_2, x_3, x_4, x_5, 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; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +x_16 = l_Lean_mkAppN(x_14, x_5); +lean_inc(x_3); +x_17 = l_Lean_Expr_fvar___override(x_3); +x_18 = l_Lean_Expr_app___override(x_16, x_17); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_18); +x_19 = lean_infer_type(x_18, x_8, x_9, x_10, x_11, x_15); +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; +x_20 = lean_ctor_get(x_19, 0); +lean_inc(x_20); +x_21 = lean_ctor_get(x_19, 1); +lean_inc(x_21); +lean_dec(x_19); +x_22 = l_Lean_Meta_RecursorInfo_numMinors(x_4); +x_23 = lean_unsigned_to_nat(0u); +x_24 = lean_unsigned_to_nat(1u); +lean_inc(x_22); +x_25 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_22); +lean_ctor_set(x_25, 2, x_24); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_18); +lean_ctor_set(x_26, 1, x_20); +lean_inc(x_6); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_6); +lean_ctor_set(x_27, 1, x_26); +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_24); +lean_ctor_set(x_28, 1, x_27); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_1); +x_29 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__12(x_7, x_4, x_3, x_1, x_6, x_2, x_22, x_25, x_25, x_28, x_23, lean_box(0), lean_box(0), x_8, x_9, x_10, x_11, x_21); +lean_dec(x_25); +lean_dec(x_22); +lean_dec(x_4); +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; uint8_t x_37; +x_30 = lean_ctor_get(x_29, 0); +lean_inc(x_30); +x_31 = lean_ctor_get(x_30, 1); +lean_inc(x_31); +lean_dec(x_30); +x_32 = lean_ctor_get(x_31, 1); +lean_inc(x_32); +x_33 = lean_ctor_get(x_29, 1); +lean_inc(x_33); +lean_dec(x_29); +x_34 = lean_ctor_get(x_31, 0); +lean_inc(x_34); +lean_dec(x_31); +x_35 = lean_ctor_get(x_32, 0); +lean_inc(x_35); +lean_dec(x_32); +x_36 = l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_1, x_35, 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); +x_37 = !lean_is_exclusive(x_36); +if (x_37 == 0) +{ +lean_object* x_38; lean_object* x_39; +x_38 = lean_ctor_get(x_36, 0); +lean_dec(x_38); +x_39 = lean_array_to_list(x_34); +lean_ctor_set(x_36, 0, x_39); +return x_36; +} +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_array_to_list(x_34); +x_42 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_40); +return x_42; +} +} +else +{ +uint8_t x_43; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_1); +x_43 = !lean_is_exclusive(x_29); +if (x_43 == 0) +{ +return x_29; +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_29, 0); +x_45 = lean_ctor_get(x_29, 1); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_29); +x_46 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_46, 0, x_44); +lean_ctor_set(x_46, 1, x_45); +return x_46; +} +} +} +else +{ +uint8_t x_47; +lean_dec(x_18); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_47 = !lean_is_exclusive(x_19); +if (x_47 == 0) +{ +return x_19; +} +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_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; +} +} +} +else +{ +uint8_t x_51; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_51 = !lean_is_exclusive(x_13); +if (x_51 == 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_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; +} +} +} +} +static lean_object* _init_l_Lean_Meta_Grind_cases___lambda__12___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(0); +x_2 = lean_array_mk(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___lambda__12(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_13 = l_Lean_Meta_mkRecursorAppPrefix(x_1, x_2, x_3, x_4, x_5, 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; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +x_16 = l_Lean_mkAppN(x_14, x_5); +lean_inc(x_3); +x_17 = l_Lean_Expr_fvar___override(x_3); +x_18 = l_Lean_Expr_app___override(x_16, x_17); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_18); +x_19 = lean_infer_type(x_18, x_8, x_9, x_10, x_11, x_15); +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; +x_20 = lean_ctor_get(x_19, 0); +lean_inc(x_20); +x_21 = lean_ctor_get(x_19, 1); +lean_inc(x_21); +lean_dec(x_19); +x_22 = l_Lean_Meta_RecursorInfo_numMinors(x_4); +x_23 = lean_unsigned_to_nat(0u); +x_24 = lean_unsigned_to_nat(1u); +lean_inc(x_22); +x_25 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_22); +lean_ctor_set(x_25, 2, x_24); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_18); +lean_ctor_set(x_26, 1, x_20); +x_27 = l_Lean_Meta_Grind_cases___lambda__12___closed__1; +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_27); +lean_ctor_set(x_28, 1, x_26); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_24); +lean_ctor_set(x_29, 1, x_28); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_1); +x_30 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__13(x_6, x_4, x_1, x_3, x_7, x_2, x_22, x_25, x_25, x_29, x_23, lean_box(0), lean_box(0), x_8, x_9, x_10, x_11, x_21); +lean_dec(x_25); +lean_dec(x_22); +lean_dec(x_4); +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; uint8_t x_38; +x_31 = lean_ctor_get(x_30, 0); +lean_inc(x_31); +x_32 = lean_ctor_get(x_31, 1); +lean_inc(x_32); +lean_dec(x_31); +x_33 = lean_ctor_get(x_32, 1); +lean_inc(x_33); +x_34 = lean_ctor_get(x_30, 1); +lean_inc(x_34); +lean_dec(x_30); +x_35 = lean_ctor_get(x_32, 0); +lean_inc(x_35); +lean_dec(x_32); +x_36 = lean_ctor_get(x_33, 0); +lean_inc(x_36); +lean_dec(x_33); +x_37 = l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_1, x_36, x_8, x_9, x_10, x_11, x_34); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +x_38 = !lean_is_exclusive(x_37); +if (x_38 == 0) +{ +lean_object* x_39; lean_object* x_40; +x_39 = lean_ctor_get(x_37, 0); +lean_dec(x_39); +x_40 = lean_array_to_list(x_35); +lean_ctor_set(x_37, 0, x_40); +return x_37; +} +else +{ +lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_41 = lean_ctor_get(x_37, 1); +lean_inc(x_41); +lean_dec(x_37); +x_42 = lean_array_to_list(x_35); +x_43 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_43, 0, x_42); +lean_ctor_set(x_43, 1, x_41); +return x_43; +} +} +else +{ +uint8_t x_44; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_1); +x_44 = !lean_is_exclusive(x_30); +if (x_44 == 0) +{ +return x_30; +} +else +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_45 = lean_ctor_get(x_30, 0); +x_46 = lean_ctor_get(x_30, 1); +lean_inc(x_46); +lean_inc(x_45); +lean_dec(x_30); +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_48; +lean_dec(x_18); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_48 = !lean_is_exclusive(x_19); +if (x_48 == 0) +{ +return x_19; +} +else +{ +lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_49 = lean_ctor_get(x_19, 0); +x_50 = lean_ctor_get(x_19, 1); +lean_inc(x_50); +lean_inc(x_49); +lean_dec(x_19); +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_52; +lean_dec(x_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_52 = !lean_is_exclusive(x_13); +if (x_52 == 0) +{ +return x_13; +} +else +{ +lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_53 = lean_ctor_get(x_13, 0); +x_54 = lean_ctor_get(x_13, 1); +lean_inc(x_54); +lean_inc(x_53); +lean_dec(x_13); +x_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; +} +} +} +} +static lean_object* _init_l_Lean_Meta_Grind_cases___lambda__13___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("x", 1, 1); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_cases___lambda__13___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_Grind_cases___lambda__13___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static size_t _init_l_Lean_Meta_Grind_cases___lambda__13___closed__3() { +_start: +{ +lean_object* x_1; size_t x_2; +x_1 = l_Lean_Meta_Grind_cases___lambda__12___closed__1; +x_2 = lean_array_size(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Grind_cases___lambda__13___closed__4() { +_start: +{ +size_t x_1; size_t x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l_Lean_Meta_Grind_cases___lambda__13___closed__3; +x_3 = l_Lean_Meta_Grind_cases___lambda__12___closed__1; +x_4 = l_Array_mapMUnsafe_map___at_Lean_LocalContext_getFVars___spec__1(x_2, x_1, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___lambda__13(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; +lean_inc(x_1); +x_8 = l_Lean_MVarId_getTag(x_1, x_3, x_4, x_5, x_6, x_7); +if (lean_obj_tag(x_8) == 0) +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_9 = lean_ctor_get(x_8, 0); +lean_inc(x_9); +x_10 = lean_ctor_get(x_8, 1); +lean_inc(x_10); +lean_dec(x_8); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +x_11 = lean_infer_type(x_2, x_3, x_4, x_5, x_6, 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); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_14 = lean_whnf(x_12, x_3, x_4, x_5, x_6, x_13); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +x_16 = lean_ctor_get(x_14, 1); +lean_inc(x_16); +lean_dec(x_14); +x_17 = l_Lean_Expr_getAppFn(x_15); +if (lean_obj_tag(x_17) == 4) +{ +lean_object* x_18; lean_object* x_19; +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +lean_dec(x_17); +lean_inc(x_18); +x_19 = l_Lean_getConstInfo___at_Lean_Meta_mkConstWithFreshMVarLevels___spec__1(x_18, x_3, x_4, x_5, x_6, x_16); +if (lean_obj_tag(x_19) == 0) +{ +lean_object* x_20; +x_20 = lean_ctor_get(x_19, 0); +lean_inc(x_20); +if (lean_obj_tag(x_20) == 5) +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +lean_dec(x_20); +x_21 = lean_ctor_get(x_19, 1); +lean_inc(x_21); +lean_dec(x_19); +x_22 = l_Lean_casesOnSuffix; +x_23 = l_Lean_Name_str___override(x_18, x_22); +x_24 = lean_box(0); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_25 = l_Lean_Meta_mkRecursorInfo(x_23, x_24, x_3, x_4, x_5, x_6, x_21); +if (lean_obj_tag(x_25) == 0) +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; uint8_t 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_Meta_RecursorInfo_numIndices(x_26); +x_29 = lean_unsigned_to_nat(0u); +x_30 = lean_nat_dec_lt(x_29, x_28); +lean_dec(x_28); +if (x_30 == 0) +{ +switch (lean_obj_tag(x_2)) { +case 0: +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_31 = l_Lean_Meta_Grind_cases___lambda__13___closed__2; +x_32 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_31, x_5, x_6, x_27); +x_33 = lean_ctor_get(x_32, 0); +lean_inc(x_33); +x_34 = lean_ctor_get(x_32, 1); +lean_inc(x_34); +lean_dec(x_32); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_35 = l_Lean_MVarId_assert(x_1, x_33, x_15, x_2, x_3, x_4, x_5, x_6, x_34); +if (lean_obj_tag(x_35) == 0) +{ +lean_object* x_36; lean_object* x_37; uint8_t 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 = 0; +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_39 = l_Lean_Meta_intro1Core(x_36, x_38, x_3, x_4, x_5, x_6, x_37); +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; lean_object* x_48; +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_ctor_get(x_40, 0); +lean_inc(x_42); +x_43 = lean_ctor_get(x_40, 1); +lean_inc(x_43); +lean_dec(x_40); +x_44 = l_Lean_Meta_Grind_cases_throwInductiveExpected___closed__3; +x_45 = l_Lean_Meta_Grind_cases___lambda__13___closed__4; +x_46 = l_Lean_Meta_Grind_cases___lambda__12___closed__1; +lean_inc(x_43); +x_47 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_cases___lambda__1___boxed), 12, 7); +lean_closure_set(x_47, 0, x_43); +lean_closure_set(x_47, 1, x_44); +lean_closure_set(x_47, 2, x_42); +lean_closure_set(x_47, 3, x_26); +lean_closure_set(x_47, 4, x_45); +lean_closure_set(x_47, 5, x_46); +lean_closure_set(x_47, 6, x_9); +x_48 = l_Lean_MVarId_withContext___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___spec__2___rarg(x_43, x_47, x_3, x_4, x_5, x_6, x_41); +return x_48; +} +else +{ +uint8_t x_49; +lean_dec(x_26); +lean_dec(x_9); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_49 = !lean_is_exclusive(x_39); +if (x_49 == 0) +{ +return x_39; +} +else +{ +lean_object* x_50; lean_object* x_51; lean_object* x_52; +x_50 = lean_ctor_get(x_39, 0); +x_51 = lean_ctor_get(x_39, 1); +lean_inc(x_51); +lean_inc(x_50); +lean_dec(x_39); 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_24: +} +else +{ +uint8_t x_53; +lean_dec(x_26); +lean_dec(x_9); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_53 = !lean_is_exclusive(x_35); +if (x_53 == 0) +{ +return x_35; +} +else +{ +lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_54 = lean_ctor_get(x_35, 0); +x_55 = lean_ctor_get(x_35, 1); +lean_inc(x_55); +lean_inc(x_54); +lean_dec(x_35); +x_56 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_56, 0, x_54); +lean_ctor_set(x_56, 1, x_55); +return x_56; +} +} +} +case 1: +{ +lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; +lean_dec(x_15); +x_57 = lean_ctor_get(x_2, 0); +lean_inc(x_57); +lean_dec(x_2); +x_58 = l_Lean_Meta_Grind_cases_throwInductiveExpected___closed__3; +x_59 = l_Lean_Meta_Grind_cases___lambda__13___closed__4; +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_26); +lean_inc(x_57); +lean_inc(x_1); +x_60 = l_Lean_Meta_mkRecursorAppPrefix(x_1, x_58, x_57, x_26, x_59, x_3, x_4, x_5, x_6, x_27); +if (lean_obj_tag(x_60) == 0) +{ +lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_61 = lean_ctor_get(x_60, 0); +lean_inc(x_61); +x_62 = lean_ctor_get(x_60, 1); +lean_inc(x_62); +lean_dec(x_60); +x_63 = l_Lean_mkAppN(x_61, x_59); +lean_inc(x_57); +x_64 = l_Lean_Expr_fvar___override(x_57); +x_65 = l_Lean_Expr_app___override(x_63, x_64); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_65); +x_66 = lean_infer_type(x_65, x_3, x_4, x_5, x_6, x_62); +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; +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_Meta_RecursorInfo_numMinors(x_26); +x_70 = lean_unsigned_to_nat(1u); +lean_inc(x_69); +x_71 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_71, 0, x_29); +lean_ctor_set(x_71, 1, x_69); +lean_ctor_set(x_71, 2, x_70); +x_72 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_72, 0, x_65); +lean_ctor_set(x_72, 1, x_67); +x_73 = l_Lean_Meta_Grind_cases___lambda__12___closed__1; +x_74 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_74, 0, x_73); +lean_ctor_set(x_74, 1, x_72); +x_75 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_75, 0, x_70); +lean_ctor_set(x_75, 1, x_74); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_1); +x_76 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__2(x_1, x_9, x_26, x_57, x_73, x_58, x_69, x_71, x_71, x_75, x_29, lean_box(0), lean_box(0), x_3, x_4, x_5, x_6, x_68); +lean_dec(x_71); +lean_dec(x_69); +lean_dec(x_26); +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; uint8_t x_84; +x_77 = lean_ctor_get(x_76, 0); +lean_inc(x_77); +x_78 = lean_ctor_get(x_77, 1); +lean_inc(x_78); +lean_dec(x_77); +x_79 = lean_ctor_get(x_78, 1); +lean_inc(x_79); +x_80 = lean_ctor_get(x_76, 1); +lean_inc(x_80); +lean_dec(x_76); +x_81 = lean_ctor_get(x_78, 0); +lean_inc(x_81); +lean_dec(x_78); +x_82 = lean_ctor_get(x_79, 0); +lean_inc(x_82); +lean_dec(x_79); +x_83 = l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_1, x_82, x_3, x_4, x_5, x_6, x_80); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_84 = !lean_is_exclusive(x_83); +if (x_84 == 0) +{ +lean_object* x_85; lean_object* x_86; +x_85 = lean_ctor_get(x_83, 0); +lean_dec(x_85); +x_86 = lean_array_to_list(x_81); +lean_ctor_set(x_83, 0, x_86); +return x_83; +} +else +{ +lean_object* x_87; lean_object* x_88; lean_object* x_89; +x_87 = lean_ctor_get(x_83, 1); +lean_inc(x_87); +lean_dec(x_83); +x_88 = lean_array_to_list(x_81); +x_89 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_89, 0, x_88); +lean_ctor_set(x_89, 1, x_87); +return x_89; +} +} +else +{ +uint8_t x_90; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_90 = !lean_is_exclusive(x_76); +if (x_90 == 0) +{ +return x_76; +} +else +{ +lean_object* x_91; lean_object* x_92; lean_object* x_93; +x_91 = lean_ctor_get(x_76, 0); +x_92 = lean_ctor_get(x_76, 1); +lean_inc(x_92); +lean_inc(x_91); +lean_dec(x_76); +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 +{ +uint8_t x_94; +lean_dec(x_65); +lean_dec(x_57); +lean_dec(x_26); +lean_dec(x_9); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_94 = !lean_is_exclusive(x_66); +if (x_94 == 0) +{ +return x_66; +} +else +{ +lean_object* x_95; lean_object* x_96; lean_object* x_97; +x_95 = lean_ctor_get(x_66, 0); +x_96 = lean_ctor_get(x_66, 1); +lean_inc(x_96); +lean_inc(x_95); +lean_dec(x_66); +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; +} +} +} +else +{ +uint8_t x_98; +lean_dec(x_57); +lean_dec(x_26); +lean_dec(x_9); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_98 = !lean_is_exclusive(x_60); +if (x_98 == 0) +{ +return x_60; +} +else +{ +lean_object* x_99; lean_object* x_100; lean_object* x_101; +x_99 = lean_ctor_get(x_60, 0); +x_100 = lean_ctor_get(x_60, 1); +lean_inc(x_100); +lean_inc(x_99); +lean_dec(x_60); +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; +} +} +} +case 2: +{ +lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; +x_102 = l_Lean_Meta_Grind_cases___lambda__13___closed__2; +x_103 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_102, x_5, x_6, x_27); +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_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_106 = l_Lean_MVarId_assert(x_1, x_104, x_15, x_2, x_3, x_4, x_5, x_6, x_105); +if (lean_obj_tag(x_106) == 0) +{ +lean_object* x_107; lean_object* x_108; uint8_t x_109; lean_object* x_110; +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; +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_110 = l_Lean_Meta_intro1Core(x_107, x_109, x_3, x_4, x_5, x_6, x_108); +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; +x_111 = lean_ctor_get(x_110, 0); +lean_inc(x_111); +x_112 = lean_ctor_get(x_110, 1); +lean_inc(x_112); +lean_dec(x_110); +x_113 = lean_ctor_get(x_111, 0); +lean_inc(x_113); +x_114 = lean_ctor_get(x_111, 1); +lean_inc(x_114); +lean_dec(x_111); +x_115 = l_Lean_Meta_Grind_cases_throwInductiveExpected___closed__3; +x_116 = l_Lean_Meta_Grind_cases___lambda__13___closed__4; +x_117 = l_Lean_Meta_Grind_cases___lambda__12___closed__1; +lean_inc(x_114); +x_118 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_cases___lambda__2___boxed), 12, 7); +lean_closure_set(x_118, 0, x_114); +lean_closure_set(x_118, 1, x_115); +lean_closure_set(x_118, 2, x_113); +lean_closure_set(x_118, 3, x_26); +lean_closure_set(x_118, 4, x_116); +lean_closure_set(x_118, 5, x_117); +lean_closure_set(x_118, 6, x_9); +x_119 = l_Lean_MVarId_withContext___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___spec__2___rarg(x_114, x_118, x_3, x_4, x_5, x_6, x_112); +return x_119; +} +else +{ +uint8_t x_120; +lean_dec(x_26); +lean_dec(x_9); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_120 = !lean_is_exclusive(x_110); +if (x_120 == 0) +{ +return x_110; +} +else +{ +lean_object* x_121; lean_object* x_122; lean_object* x_123; +x_121 = lean_ctor_get(x_110, 0); +x_122 = lean_ctor_get(x_110, 1); +lean_inc(x_122); +lean_inc(x_121); +lean_dec(x_110); +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_26); +lean_dec(x_9); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_124 = !lean_is_exclusive(x_106); +if (x_124 == 0) +{ +return x_106; +} +else +{ +lean_object* x_125; lean_object* x_126; lean_object* x_127; +x_125 = lean_ctor_get(x_106, 0); +x_126 = lean_ctor_get(x_106, 1); +lean_inc(x_126); +lean_inc(x_125); +lean_dec(x_106); +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; +} +} +} +case 3: +{ +lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; +x_128 = l_Lean_Meta_Grind_cases___lambda__13___closed__2; +x_129 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_128, x_5, x_6, x_27); +x_130 = lean_ctor_get(x_129, 0); +lean_inc(x_130); +x_131 = lean_ctor_get(x_129, 1); +lean_inc(x_131); +lean_dec(x_129); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_132 = l_Lean_MVarId_assert(x_1, x_130, x_15, x_2, x_3, x_4, x_5, x_6, x_131); +if (lean_obj_tag(x_132) == 0) +{ +lean_object* x_133; lean_object* x_134; uint8_t x_135; lean_object* x_136; +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; +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_136 = l_Lean_Meta_intro1Core(x_133, x_135, x_3, x_4, x_5, x_6, x_134); +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_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; +x_137 = lean_ctor_get(x_136, 0); +lean_inc(x_137); +x_138 = lean_ctor_get(x_136, 1); +lean_inc(x_138); +lean_dec(x_136); +x_139 = lean_ctor_get(x_137, 0); +lean_inc(x_139); +x_140 = lean_ctor_get(x_137, 1); +lean_inc(x_140); +lean_dec(x_137); +x_141 = l_Lean_Meta_Grind_cases_throwInductiveExpected___closed__3; +x_142 = l_Lean_Meta_Grind_cases___lambda__13___closed__4; +x_143 = l_Lean_Meta_Grind_cases___lambda__12___closed__1; +lean_inc(x_140); +x_144 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_cases___lambda__3___boxed), 12, 7); +lean_closure_set(x_144, 0, x_140); +lean_closure_set(x_144, 1, x_141); +lean_closure_set(x_144, 2, x_139); +lean_closure_set(x_144, 3, x_26); +lean_closure_set(x_144, 4, x_142); +lean_closure_set(x_144, 5, x_143); +lean_closure_set(x_144, 6, x_9); +x_145 = l_Lean_MVarId_withContext___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___spec__2___rarg(x_140, x_144, x_3, x_4, x_5, x_6, x_138); +return x_145; +} +else +{ +uint8_t x_146; +lean_dec(x_26); +lean_dec(x_9); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_146 = !lean_is_exclusive(x_136); +if (x_146 == 0) +{ +return x_136; +} +else +{ +lean_object* x_147; lean_object* x_148; lean_object* x_149; +x_147 = lean_ctor_get(x_136, 0); +x_148 = lean_ctor_get(x_136, 1); +lean_inc(x_148); +lean_inc(x_147); +lean_dec(x_136); +x_149 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_149, 0, x_147); +lean_ctor_set(x_149, 1, x_148); +return x_149; +} +} +} +else +{ +uint8_t x_150; +lean_dec(x_26); +lean_dec(x_9); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_150 = !lean_is_exclusive(x_132); +if (x_150 == 0) +{ +return x_132; +} +else +{ +lean_object* x_151; lean_object* x_152; lean_object* x_153; +x_151 = lean_ctor_get(x_132, 0); +x_152 = lean_ctor_get(x_132, 1); +lean_inc(x_152); +lean_inc(x_151); +lean_dec(x_132); +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; +} +} +} +case 4: +{ +lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; +x_154 = l_Lean_Meta_Grind_cases___lambda__13___closed__2; +x_155 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_154, x_5, x_6, x_27); +x_156 = lean_ctor_get(x_155, 0); +lean_inc(x_156); +x_157 = lean_ctor_get(x_155, 1); +lean_inc(x_157); +lean_dec(x_155); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_158 = l_Lean_MVarId_assert(x_1, x_156, x_15, x_2, x_3, x_4, x_5, x_6, x_157); +if (lean_obj_tag(x_158) == 0) +{ +lean_object* x_159; lean_object* x_160; uint8_t 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); +lean_dec(x_158); +x_161 = 0; +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_162 = l_Lean_Meta_intro1Core(x_159, x_161, x_3, x_4, x_5, x_6, x_160); +if (lean_obj_tag(x_162) == 0) +{ +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_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 = lean_ctor_get(x_163, 0); +lean_inc(x_165); +x_166 = lean_ctor_get(x_163, 1); +lean_inc(x_166); +lean_dec(x_163); +x_167 = l_Lean_Meta_Grind_cases_throwInductiveExpected___closed__3; +x_168 = l_Lean_Meta_Grind_cases___lambda__13___closed__4; +x_169 = l_Lean_Meta_Grind_cases___lambda__12___closed__1; +lean_inc(x_166); +x_170 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_cases___lambda__4___boxed), 12, 7); +lean_closure_set(x_170, 0, x_166); +lean_closure_set(x_170, 1, x_167); +lean_closure_set(x_170, 2, x_165); +lean_closure_set(x_170, 3, x_26); +lean_closure_set(x_170, 4, x_168); +lean_closure_set(x_170, 5, x_169); +lean_closure_set(x_170, 6, x_9); +x_171 = l_Lean_MVarId_withContext___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___spec__2___rarg(x_166, x_170, x_3, x_4, x_5, x_6, x_164); +return x_171; +} +else +{ +uint8_t x_172; +lean_dec(x_26); +lean_dec(x_9); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_172 = !lean_is_exclusive(x_162); +if (x_172 == 0) +{ +return x_162; +} +else +{ +lean_object* x_173; lean_object* x_174; lean_object* x_175; +x_173 = lean_ctor_get(x_162, 0); +x_174 = lean_ctor_get(x_162, 1); +lean_inc(x_174); +lean_inc(x_173); +lean_dec(x_162); +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_176; +lean_dec(x_26); +lean_dec(x_9); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_176 = !lean_is_exclusive(x_158); +if (x_176 == 0) +{ +return x_158; +} +else +{ +lean_object* x_177; lean_object* x_178; lean_object* x_179; +x_177 = lean_ctor_get(x_158, 0); +x_178 = lean_ctor_get(x_158, 1); +lean_inc(x_178); +lean_inc(x_177); +lean_dec(x_158); +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; +} +} +} +case 5: +{ +lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; +x_180 = l_Lean_Meta_Grind_cases___lambda__13___closed__2; +x_181 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_180, x_5, x_6, x_27); +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_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_184 = l_Lean_MVarId_assert(x_1, x_182, x_15, x_2, x_3, x_4, x_5, x_6, x_183); +if (lean_obj_tag(x_184) == 0) +{ +lean_object* x_185; lean_object* x_186; uint8_t x_187; lean_object* x_188; +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 = 0; +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_188 = l_Lean_Meta_intro1Core(x_185, x_187, x_3, x_4, x_5, x_6, x_186); +if (lean_obj_tag(x_188) == 0) +{ +lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; +x_189 = lean_ctor_get(x_188, 0); +lean_inc(x_189); +x_190 = lean_ctor_get(x_188, 1); +lean_inc(x_190); +lean_dec(x_188); +x_191 = lean_ctor_get(x_189, 0); +lean_inc(x_191); +x_192 = lean_ctor_get(x_189, 1); +lean_inc(x_192); +lean_dec(x_189); +x_193 = l_Lean_Meta_Grind_cases_throwInductiveExpected___closed__3; +x_194 = l_Lean_Meta_Grind_cases___lambda__13___closed__4; +x_195 = l_Lean_Meta_Grind_cases___lambda__12___closed__1; +lean_inc(x_192); +x_196 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_cases___lambda__5___boxed), 12, 7); +lean_closure_set(x_196, 0, x_192); +lean_closure_set(x_196, 1, x_193); +lean_closure_set(x_196, 2, x_191); +lean_closure_set(x_196, 3, x_26); +lean_closure_set(x_196, 4, x_194); +lean_closure_set(x_196, 5, x_195); +lean_closure_set(x_196, 6, x_9); +x_197 = l_Lean_MVarId_withContext___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___spec__2___rarg(x_192, x_196, x_3, x_4, x_5, x_6, x_190); +return x_197; +} +else +{ +uint8_t x_198; +lean_dec(x_26); +lean_dec(x_9); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_198 = !lean_is_exclusive(x_188); +if (x_198 == 0) +{ +return x_188; +} +else { -lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_20 = lean_ctor_get(x_18, 0); -lean_inc(x_20); -lean_dec(x_18); -x_21 = lean_ctor_get(x_5, 2); -x_22 = lean_nat_add(x_7, x_21); -lean_dec(x_7); -x_6 = x_20; -x_7 = x_22; -x_8 = lean_box(0); -x_9 = lean_box(0); -x_14 = x_19; -goto _start; +lean_object* x_199; lean_object* x_200; lean_object* x_201; +x_199 = lean_ctor_get(x_188, 0); +x_200 = lean_ctor_get(x_188, 1); +lean_inc(x_200); +lean_inc(x_199); +lean_dec(x_188); +x_201 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_201, 0, x_199); +lean_ctor_set(x_201, 1, x_200); +return x_201; } } } +else +{ +uint8_t x_202; +lean_dec(x_26); +lean_dec(x_9); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_202 = !lean_is_exclusive(x_184); +if (x_202 == 0) +{ +return x_184; } -static lean_object* _init_l_Lean_Meta_Grind_cases___lambda__1___closed__1() { -_start: +else { -lean_object* x_1; lean_object* x_2; -x_1 = lean_box(0); -x_2 = lean_array_mk(x_1); -return x_2; +lean_object* x_203; lean_object* x_204; lean_object* x_205; +x_203 = lean_ctor_get(x_184, 0); +x_204 = lean_ctor_get(x_184, 1); +lean_inc(x_204); +lean_inc(x_203); +lean_dec(x_184); +x_205 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_205, 0, x_203); +lean_ctor_set(x_205, 1, x_204); +return x_205; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___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: +} +case 6: { -lean_object* x_12; -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); +lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; +x_206 = l_Lean_Meta_Grind_cases___lambda__13___closed__2; +x_207 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_206, x_5, x_6, x_27); +x_208 = lean_ctor_get(x_207, 0); +lean_inc(x_208); +x_209 = lean_ctor_get(x_207, 1); +lean_inc(x_209); +lean_dec(x_207); +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_12 = l_Lean_Meta_mkRecursorAppPrefix(x_1, x_2, x_3, x_4, x_5, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_12) == 0) +x_210 = l_Lean_MVarId_assert(x_1, x_208, x_15, x_2, x_3, x_4, x_5, x_6, x_209); +if (lean_obj_tag(x_210) == 0) { -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_13 = lean_ctor_get(x_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_mkAppN(x_13, x_5); -x_16 = l_Lean_Expr_fvar___override(x_3); -x_17 = l_Lean_Expr_app___override(x_15, x_16); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_17); -x_18 = lean_infer_type(x_17, x_7, x_8, x_9, x_10, x_14); -if (lean_obj_tag(x_18) == 0) +lean_object* x_211; lean_object* x_212; uint8_t x_213; lean_object* x_214; +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; +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_214 = l_Lean_Meta_intro1Core(x_211, x_213, x_3, x_4, x_5, x_6, x_212); +if (lean_obj_tag(x_214) == 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; -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_Meta_RecursorInfo_numMinors(x_4); -lean_dec(x_4); -x_22 = lean_unsigned_to_nat(0u); -x_23 = lean_unsigned_to_nat(1u); -x_24 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_24, 0, x_22); -lean_ctor_set(x_24, 1, x_21); -lean_ctor_set(x_24, 2, x_23); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_17); -lean_ctor_set(x_25, 1, x_19); -x_26 = l_Lean_Meta_Grind_cases___lambda__1___closed__1; -x_27 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_27, 0, x_26); -lean_ctor_set(x_27, 1, x_25); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_1); -x_28 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__2(x_6, x_1, x_2, x_24, x_24, x_27, x_22, lean_box(0), lean_box(0), x_7, x_8, x_9, x_10, x_20); -lean_dec(x_24); -if (lean_obj_tag(x_28) == 0) +lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; +x_215 = lean_ctor_get(x_214, 0); +lean_inc(x_215); +x_216 = lean_ctor_get(x_214, 1); +lean_inc(x_216); +lean_dec(x_214); +x_217 = lean_ctor_get(x_215, 0); +lean_inc(x_217); +x_218 = lean_ctor_get(x_215, 1); +lean_inc(x_218); +lean_dec(x_215); +x_219 = l_Lean_Meta_Grind_cases_throwInductiveExpected___closed__3; +x_220 = l_Lean_Meta_Grind_cases___lambda__13___closed__4; +x_221 = l_Lean_Meta_Grind_cases___lambda__12___closed__1; +lean_inc(x_218); +x_222 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_cases___lambda__6___boxed), 12, 7); +lean_closure_set(x_222, 0, x_218); +lean_closure_set(x_222, 1, x_219); +lean_closure_set(x_222, 2, x_217); +lean_closure_set(x_222, 3, x_26); +lean_closure_set(x_222, 4, x_220); +lean_closure_set(x_222, 5, x_221); +lean_closure_set(x_222, 6, x_9); +x_223 = l_Lean_MVarId_withContext___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___spec__2___rarg(x_218, x_222, x_3, x_4, x_5, x_6, x_216); +return x_223; +} +else { -lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; -x_29 = lean_ctor_get(x_28, 0); -lean_inc(x_29); -x_30 = lean_ctor_get(x_29, 1); -lean_inc(x_30); -x_31 = lean_ctor_get(x_28, 1); -lean_inc(x_31); -lean_dec(x_28); -x_32 = lean_ctor_get(x_29, 0); -lean_inc(x_32); -lean_dec(x_29); -x_33 = lean_ctor_get(x_30, 0); -lean_inc(x_33); -lean_dec(x_30); -x_34 = l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_1, x_33, x_7, x_8, x_9, x_10, x_31); -lean_dec(x_10); +uint8_t x_224; +lean_dec(x_26); lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -x_35 = !lean_is_exclusive(x_34); -if (x_35 == 0) +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_224 = !lean_is_exclusive(x_214); +if (x_224 == 0) { -lean_object* x_36; lean_object* x_37; -x_36 = lean_ctor_get(x_34, 0); -lean_dec(x_36); -x_37 = lean_array_to_list(x_32); -lean_ctor_set(x_34, 0, x_37); -return x_34; +return x_214; } 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_array_to_list(x_32); -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_225; lean_object* x_226; lean_object* x_227; +x_225 = lean_ctor_get(x_214, 0); +x_226 = lean_ctor_get(x_214, 1); +lean_inc(x_226); +lean_inc(x_225); +lean_dec(x_214); +x_227 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_227, 0, x_225); +lean_ctor_set(x_227, 1, x_226); +return x_227; +} } } else { -uint8_t x_41; -lean_dec(x_10); +uint8_t x_228; +lean_dec(x_26); lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_1); -x_41 = !lean_is_exclusive(x_28); -if (x_41 == 0) +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_228 = !lean_is_exclusive(x_210); +if (x_228 == 0) { -return x_28; +return x_210; } else { -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_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; +lean_object* x_229; lean_object* x_230; lean_object* x_231; +x_229 = lean_ctor_get(x_210, 0); +x_230 = lean_ctor_get(x_210, 1); +lean_inc(x_230); +lean_inc(x_229); +lean_dec(x_210); +x_231 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_231, 0, x_229); +lean_ctor_set(x_231, 1, x_230); +return x_231; +} } } +case 7: +{ +lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; +x_232 = l_Lean_Meta_Grind_cases___lambda__13___closed__2; +x_233 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_232, x_5, x_6, x_27); +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); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_236 = l_Lean_MVarId_assert(x_1, x_234, x_15, x_2, x_3, x_4, x_5, x_6, x_235); +if (lean_obj_tag(x_236) == 0) +{ +lean_object* x_237; lean_object* x_238; uint8_t x_239; lean_object* x_240; +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 = 0; +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_240 = l_Lean_Meta_intro1Core(x_237, x_239, x_3, x_4, x_5, x_6, x_238); +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; +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_243 = lean_ctor_get(x_241, 0); +lean_inc(x_243); +x_244 = lean_ctor_get(x_241, 1); +lean_inc(x_244); +lean_dec(x_241); +x_245 = l_Lean_Meta_Grind_cases_throwInductiveExpected___closed__3; +x_246 = l_Lean_Meta_Grind_cases___lambda__13___closed__4; +x_247 = l_Lean_Meta_Grind_cases___lambda__12___closed__1; +lean_inc(x_244); +x_248 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_cases___lambda__7___boxed), 12, 7); +lean_closure_set(x_248, 0, x_244); +lean_closure_set(x_248, 1, x_245); +lean_closure_set(x_248, 2, x_243); +lean_closure_set(x_248, 3, x_26); +lean_closure_set(x_248, 4, x_246); +lean_closure_set(x_248, 5, x_247); +lean_closure_set(x_248, 6, x_9); +x_249 = l_Lean_MVarId_withContext___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___spec__2___rarg(x_244, x_248, x_3, x_4, x_5, x_6, x_242); +return x_249; } else { -uint8_t x_45; -lean_dec(x_17); -lean_dec(x_10); +uint8_t x_250; +lean_dec(x_26); 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_18); -if (x_45 == 0) +lean_dec(x_3); +x_250 = !lean_is_exclusive(x_240); +if (x_250 == 0) { -return x_18; +return x_240; } else { -lean_object* x_46; lean_object* x_47; lean_object* x_48; -x_46 = lean_ctor_get(x_18, 0); -x_47 = lean_ctor_get(x_18, 1); -lean_inc(x_47); -lean_inc(x_46); -lean_dec(x_18); -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_251; lean_object* x_252; lean_object* x_253; +x_251 = lean_ctor_get(x_240, 0); +x_252 = lean_ctor_get(x_240, 1); +lean_inc(x_252); +lean_inc(x_251); +lean_dec(x_240); +x_253 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_253, 0, x_251); +lean_ctor_set(x_253, 1, x_252); +return x_253; } } } else { -uint8_t x_49; -lean_dec(x_10); +uint8_t x_254; +lean_dec(x_26); 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_12); -if (x_49 == 0) +x_254 = !lean_is_exclusive(x_236); +if (x_254 == 0) { -return x_12; +return x_236; } else { -lean_object* x_50; lean_object* x_51; lean_object* x_52; -x_50 = lean_ctor_get(x_12, 0); -x_51 = lean_ctor_get(x_12, 1); -lean_inc(x_51); -lean_inc(x_50); -lean_dec(x_12); -x_52 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_52, 0, x_50); -lean_ctor_set(x_52, 1, x_51); -return x_52; -} +lean_object* x_255; lean_object* x_256; lean_object* x_257; +x_255 = lean_ctor_get(x_236, 0); +x_256 = lean_ctor_get(x_236, 1); +lean_inc(x_256); +lean_inc(x_255); +lean_dec(x_236); +x_257 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_257, 0, x_255); +lean_ctor_set(x_257, 1, x_256); +return x_257; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: -{ -lean_object* x_8; -lean_inc(x_1); -x_8 = l_Lean_MVarId_getTag(x_1, x_3, x_4, x_5, x_6, x_7); -if (lean_obj_tag(x_8) == 0) +case 8: { -lean_object* x_9; lean_object* x_10; lean_object* x_11; -x_9 = lean_ctor_get(x_8, 0); -lean_inc(x_9); -x_10 = lean_ctor_get(x_8, 1); -lean_inc(x_10); -lean_dec(x_8); +lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; +x_258 = l_Lean_Meta_Grind_cases___lambda__13___closed__2; +x_259 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_258, x_5, x_6, x_27); +x_260 = lean_ctor_get(x_259, 0); +lean_inc(x_260); +x_261 = lean_ctor_get(x_259, 1); +lean_inc(x_261); +lean_dec(x_259); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); lean_inc(x_3); -lean_inc(x_2); -x_11 = l_Lean_FVarId_getType(x_2, x_3, x_4, x_5, x_6, x_10); -if (lean_obj_tag(x_11) == 0) +x_262 = l_Lean_MVarId_assert(x_1, x_260, x_15, x_2, x_3, x_4, x_5, x_6, x_261); +if (lean_obj_tag(x_262) == 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); +lean_object* x_263; lean_object* x_264; uint8_t x_265; lean_object* x_266; +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 = 0; lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -x_14 = lean_whnf(x_12, x_3, x_4, x_5, x_6, x_13); -if (lean_obj_tag(x_14) == 0) +x_266 = l_Lean_Meta_intro1Core(x_263, x_265, x_3, x_4, x_5, x_6, x_264); +if (lean_obj_tag(x_266) == 0) { -lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = lean_ctor_get(x_14, 0); -lean_inc(x_15); -x_16 = lean_ctor_get(x_14, 1); -lean_inc(x_16); -lean_dec(x_14); -x_17 = l_Lean_Expr_getAppFn(x_15); -if (lean_obj_tag(x_17) == 4) +lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; +x_267 = lean_ctor_get(x_266, 0); +lean_inc(x_267); +x_268 = lean_ctor_get(x_266, 1); +lean_inc(x_268); +lean_dec(x_266); +x_269 = lean_ctor_get(x_267, 0); +lean_inc(x_269); +x_270 = lean_ctor_get(x_267, 1); +lean_inc(x_270); +lean_dec(x_267); +x_271 = l_Lean_Meta_Grind_cases_throwInductiveExpected___closed__3; +x_272 = l_Lean_Meta_Grind_cases___lambda__13___closed__4; +x_273 = l_Lean_Meta_Grind_cases___lambda__12___closed__1; +lean_inc(x_270); +x_274 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_cases___lambda__8___boxed), 12, 7); +lean_closure_set(x_274, 0, x_270); +lean_closure_set(x_274, 1, x_271); +lean_closure_set(x_274, 2, x_269); +lean_closure_set(x_274, 3, x_26); +lean_closure_set(x_274, 4, x_272); +lean_closure_set(x_274, 5, x_273); +lean_closure_set(x_274, 6, x_9); +x_275 = l_Lean_MVarId_withContext___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___spec__2___rarg(x_270, x_274, x_3, x_4, x_5, x_6, x_268); +return x_275; +} +else { -lean_object* x_18; lean_object* x_19; -x_18 = lean_ctor_get(x_17, 0); -lean_inc(x_18); -lean_dec(x_17); -lean_inc(x_18); -x_19 = l_Lean_getConstInfo___at_Lean_Meta_mkConstWithFreshMVarLevels___spec__1(x_18, x_3, x_4, x_5, x_6, x_16); -if (lean_obj_tag(x_19) == 0) +uint8_t x_276; +lean_dec(x_26); +lean_dec(x_9); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_276 = !lean_is_exclusive(x_266); +if (x_276 == 0) { -lean_object* x_20; -x_20 = lean_ctor_get(x_19, 0); -lean_inc(x_20); -if (lean_obj_tag(x_20) == 5) +return x_266; +} +else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; -lean_dec(x_20); -x_21 = lean_ctor_get(x_19, 1); -lean_inc(x_21); -lean_dec(x_19); -x_22 = l_Lean_casesOnSuffix; -x_23 = l_Lean_Name_str___override(x_18, x_22); -x_24 = lean_box(0); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -x_25 = l_Lean_Meta_mkRecursorInfo(x_23, x_24, x_3, x_4, x_5, x_6, x_21); -if (lean_obj_tag(x_25) == 0) +lean_object* x_277; lean_object* x_278; lean_object* x_279; +x_277 = lean_ctor_get(x_266, 0); +x_278 = lean_ctor_get(x_266, 1); +lean_inc(x_278); +lean_inc(x_277); +lean_dec(x_266); +x_279 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_279, 0, x_277); +lean_ctor_set(x_279, 1, x_278); +return x_279; +} +} +} +else { -lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; uint8_t 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_Meta_RecursorInfo_numIndices(x_26); -x_29 = lean_unsigned_to_nat(0u); -x_30 = lean_nat_dec_lt(x_29, x_28); -lean_dec(x_28); -if (x_30 == 0) +uint8_t x_280; +lean_dec(x_26); +lean_dec(x_9); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_280 = !lean_is_exclusive(x_262); +if (x_280 == 0) { -lean_object* x_31; lean_object* x_32; -x_31 = l_Lean_Meta_Grind_cases_throwInductiveExpected___closed__3; -lean_inc(x_3); -lean_inc(x_26); -lean_inc(x_1); -x_32 = l_Lean_Meta_getMajorTypeIndices(x_1, x_31, x_26, x_15, x_3, x_4, x_5, x_6, x_27); -if (lean_obj_tag(x_32) == 0) +return x_262; +} +else { -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_281; lean_object* x_282; lean_object* x_283; +x_281 = lean_ctor_get(x_262, 0); +x_282 = lean_ctor_get(x_262, 1); +lean_inc(x_282); +lean_inc(x_281); +lean_dec(x_262); +x_283 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_283, 0, x_281); +lean_ctor_set(x_283, 1, x_282); +return x_283; +} +} +} +case 9: +{ +lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; +x_284 = l_Lean_Meta_Grind_cases___lambda__13___closed__2; +x_285 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_284, x_5, x_6, x_27); +x_286 = lean_ctor_get(x_285, 0); +lean_inc(x_286); +x_287 = lean_ctor_get(x_285, 1); +lean_inc(x_287); +lean_dec(x_285); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -lean_inc(x_26); -lean_inc(x_2); -lean_inc(x_1); -x_35 = l_Lean_Meta_mkRecursorAppPrefix(x_1, x_31, x_2, x_26, x_33, x_3, x_4, x_5, x_6, x_34); -if (lean_obj_tag(x_35) == 0) +x_288 = l_Lean_MVarId_assert(x_1, x_286, x_15, x_2, x_3, x_4, x_5, x_6, x_287); +if (lean_obj_tag(x_288) == 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_mkAppN(x_36, x_33); -lean_dec(x_33); -lean_inc(x_2); -x_39 = l_Lean_Expr_fvar___override(x_2); -x_40 = l_Lean_Expr_app___override(x_38, x_39); +lean_object* x_289; lean_object* x_290; uint8_t x_291; lean_object* x_292; +x_289 = lean_ctor_get(x_288, 0); +lean_inc(x_289); +x_290 = lean_ctor_get(x_288, 1); +lean_inc(x_290); +lean_dec(x_288); +x_291 = 0; lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -lean_inc(x_40); -x_41 = lean_infer_type(x_40, x_3, x_4, x_5, x_6, x_37); -if (lean_obj_tag(x_41) == 0) +x_292 = l_Lean_Meta_intro1Core(x_289, x_291, x_3, x_4, x_5, x_6, x_290); +if (lean_obj_tag(x_292) == 0) +{ +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_293 = lean_ctor_get(x_292, 0); +lean_inc(x_293); +x_294 = lean_ctor_get(x_292, 1); +lean_inc(x_294); +lean_dec(x_292); +x_295 = lean_ctor_get(x_293, 0); +lean_inc(x_295); +x_296 = lean_ctor_get(x_293, 1); +lean_inc(x_296); +lean_dec(x_293); +x_297 = l_Lean_Meta_Grind_cases_throwInductiveExpected___closed__3; +x_298 = l_Lean_Meta_Grind_cases___lambda__13___closed__4; +x_299 = l_Lean_Meta_Grind_cases___lambda__12___closed__1; +lean_inc(x_296); +x_300 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_cases___lambda__9___boxed), 12, 7); +lean_closure_set(x_300, 0, x_296); +lean_closure_set(x_300, 1, x_297); +lean_closure_set(x_300, 2, x_295); +lean_closure_set(x_300, 3, x_26); +lean_closure_set(x_300, 4, x_298); +lean_closure_set(x_300, 5, x_299); +lean_closure_set(x_300, 6, x_9); +x_301 = l_Lean_MVarId_withContext___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___spec__2___rarg(x_296, x_300, x_3, x_4, x_5, x_6, x_294); +return x_301; +} +else +{ +uint8_t x_302; +lean_dec(x_26); +lean_dec(x_9); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_302 = !lean_is_exclusive(x_292); +if (x_302 == 0) +{ +return x_292; +} +else { -lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; -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_RecursorInfo_numMinors(x_26); -lean_dec(x_26); -x_45 = lean_unsigned_to_nat(1u); -x_46 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_46, 0, x_29); -lean_ctor_set(x_46, 1, x_44); -lean_ctor_set(x_46, 2, x_45); -x_47 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_47, 0, x_40); -lean_ctor_set(x_47, 1, x_42); -x_48 = l_Lean_Meta_Grind_cases___lambda__1___closed__1; -x_49 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_49, 0, x_48); -lean_ctor_set(x_49, 1, x_47); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -lean_inc(x_1); -x_50 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__1(x_1, x_2, x_9, x_31, x_46, x_46, x_49, x_29, lean_box(0), lean_box(0), x_3, x_4, x_5, x_6, x_43); -lean_dec(x_46); -if (lean_obj_tag(x_50) == 0) +lean_object* x_303; lean_object* x_304; lean_object* x_305; +x_303 = lean_ctor_get(x_292, 0); +x_304 = lean_ctor_get(x_292, 1); +lean_inc(x_304); +lean_inc(x_303); +lean_dec(x_292); +x_305 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_305, 0, x_303); +lean_ctor_set(x_305, 1, x_304); +return x_305; +} +} +} +else { -lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; uint8_t x_57; -x_51 = lean_ctor_get(x_50, 0); -lean_inc(x_51); -x_52 = lean_ctor_get(x_51, 1); -lean_inc(x_52); -x_53 = lean_ctor_get(x_50, 1); -lean_inc(x_53); -lean_dec(x_50); -x_54 = lean_ctor_get(x_51, 0); -lean_inc(x_54); -lean_dec(x_51); -x_55 = lean_ctor_get(x_52, 0); -lean_inc(x_55); -lean_dec(x_52); -x_56 = l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_1, x_55, x_3, x_4, x_5, x_6, x_53); +uint8_t x_306; +lean_dec(x_26); +lean_dec(x_9); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_57 = !lean_is_exclusive(x_56); -if (x_57 == 0) +x_306 = !lean_is_exclusive(x_288); +if (x_306 == 0) { -lean_object* x_58; lean_object* x_59; -x_58 = lean_ctor_get(x_56, 0); -lean_dec(x_58); -x_59 = lean_array_to_list(x_54); -lean_ctor_set(x_56, 0, x_59); -return x_56; +return x_288; } 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_array_to_list(x_54); -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_307; lean_object* x_308; lean_object* x_309; +x_307 = lean_ctor_get(x_288, 0); +x_308 = lean_ctor_get(x_288, 1); +lean_inc(x_308); +lean_inc(x_307); +lean_dec(x_288); +x_309 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_309, 0, x_307); +lean_ctor_set(x_309, 1, x_308); +return x_309; } } +} +case 10: +{ +lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; +x_310 = l_Lean_Meta_Grind_cases___lambda__13___closed__2; +x_311 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_310, x_5, x_6, x_27); +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); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_314 = l_Lean_MVarId_assert(x_1, x_312, x_15, x_2, x_3, x_4, x_5, x_6, x_313); +if (lean_obj_tag(x_314) == 0) +{ +lean_object* x_315; lean_object* x_316; uint8_t x_317; lean_object* x_318; +x_315 = lean_ctor_get(x_314, 0); +lean_inc(x_315); +x_316 = lean_ctor_get(x_314, 1); +lean_inc(x_316); +lean_dec(x_314); +x_317 = 0; +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_318 = l_Lean_Meta_intro1Core(x_315, x_317, x_3, x_4, x_5, x_6, x_316); +if (lean_obj_tag(x_318) == 0) +{ +lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; +x_319 = lean_ctor_get(x_318, 0); +lean_inc(x_319); +x_320 = lean_ctor_get(x_318, 1); +lean_inc(x_320); +lean_dec(x_318); +x_321 = lean_ctor_get(x_319, 0); +lean_inc(x_321); +x_322 = lean_ctor_get(x_319, 1); +lean_inc(x_322); +lean_dec(x_319); +x_323 = l_Lean_Meta_Grind_cases_throwInductiveExpected___closed__3; +x_324 = l_Lean_Meta_Grind_cases___lambda__13___closed__4; +x_325 = l_Lean_Meta_Grind_cases___lambda__12___closed__1; +lean_inc(x_322); +x_326 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_cases___lambda__10___boxed), 12, 7); +lean_closure_set(x_326, 0, x_322); +lean_closure_set(x_326, 1, x_323); +lean_closure_set(x_326, 2, x_321); +lean_closure_set(x_326, 3, x_26); +lean_closure_set(x_326, 4, x_324); +lean_closure_set(x_326, 5, x_325); +lean_closure_set(x_326, 6, x_9); +x_327 = l_Lean_MVarId_withContext___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___spec__2___rarg(x_322, x_326, x_3, x_4, x_5, x_6, x_320); +return x_327; +} else { -uint8_t x_63; +uint8_t x_328; +lean_dec(x_26); +lean_dec(x_9); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -lean_dec(x_1); -x_63 = !lean_is_exclusive(x_50); -if (x_63 == 0) +x_328 = !lean_is_exclusive(x_318); +if (x_328 == 0) { -return x_50; +return x_318; } else { -lean_object* x_64; lean_object* x_65; lean_object* x_66; -x_64 = lean_ctor_get(x_50, 0); -x_65 = lean_ctor_get(x_50, 1); -lean_inc(x_65); -lean_inc(x_64); -lean_dec(x_50); -x_66 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_66, 0, x_64); -lean_ctor_set(x_66, 1, x_65); -return x_66; +lean_object* x_329; lean_object* x_330; lean_object* x_331; +x_329 = lean_ctor_get(x_318, 0); +x_330 = lean_ctor_get(x_318, 1); +lean_inc(x_330); +lean_inc(x_329); +lean_dec(x_318); +x_331 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_331, 0, x_329); +lean_ctor_set(x_331, 1, x_330); +return x_331; } } } else { -uint8_t x_67; -lean_dec(x_40); +uint8_t x_332; lean_dec(x_26); lean_dec(x_9); 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_41); -if (x_67 == 0) +x_332 = !lean_is_exclusive(x_314); +if (x_332 == 0) { -return x_41; +return x_314; } else { -lean_object* x_68; lean_object* x_69; lean_object* x_70; -x_68 = lean_ctor_get(x_41, 0); -x_69 = lean_ctor_get(x_41, 1); -lean_inc(x_69); -lean_inc(x_68); -lean_dec(x_41); -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_333; lean_object* x_334; lean_object* x_335; +x_333 = lean_ctor_get(x_314, 0); +x_334 = lean_ctor_get(x_314, 1); +lean_inc(x_334); +lean_inc(x_333); +lean_dec(x_314); +x_335 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_335, 0, x_333); +lean_ctor_set(x_335, 1, x_334); +return x_335; } } } +default: +{ +lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; +x_336 = l_Lean_Meta_Grind_cases___lambda__13___closed__2; +x_337 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_336, x_5, x_6, x_27); +x_338 = lean_ctor_get(x_337, 0); +lean_inc(x_338); +x_339 = lean_ctor_get(x_337, 1); +lean_inc(x_339); +lean_dec(x_337); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_340 = l_Lean_MVarId_assert(x_1, x_338, x_15, x_2, x_3, x_4, x_5, x_6, x_339); +if (lean_obj_tag(x_340) == 0) +{ +lean_object* x_341; lean_object* x_342; uint8_t x_343; lean_object* x_344; +x_341 = lean_ctor_get(x_340, 0); +lean_inc(x_341); +x_342 = lean_ctor_get(x_340, 1); +lean_inc(x_342); +lean_dec(x_340); +x_343 = 0; +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_344 = l_Lean_Meta_intro1Core(x_341, x_343, x_3, x_4, x_5, x_6, x_342); +if (lean_obj_tag(x_344) == 0) +{ +lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; lean_object* x_353; +x_345 = lean_ctor_get(x_344, 0); +lean_inc(x_345); +x_346 = lean_ctor_get(x_344, 1); +lean_inc(x_346); +lean_dec(x_344); +x_347 = lean_ctor_get(x_345, 0); +lean_inc(x_347); +x_348 = lean_ctor_get(x_345, 1); +lean_inc(x_348); +lean_dec(x_345); +x_349 = l_Lean_Meta_Grind_cases_throwInductiveExpected___closed__3; +x_350 = l_Lean_Meta_Grind_cases___lambda__13___closed__4; +x_351 = l_Lean_Meta_Grind_cases___lambda__12___closed__1; +lean_inc(x_348); +x_352 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_cases___lambda__11___boxed), 12, 7); +lean_closure_set(x_352, 0, x_348); +lean_closure_set(x_352, 1, x_349); +lean_closure_set(x_352, 2, x_347); +lean_closure_set(x_352, 3, x_26); +lean_closure_set(x_352, 4, x_350); +lean_closure_set(x_352, 5, x_351); +lean_closure_set(x_352, 6, x_9); +x_353 = l_Lean_MVarId_withContext___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___spec__2___rarg(x_348, x_352, x_3, x_4, x_5, x_6, x_346); +return x_353; +} else { -uint8_t x_71; -lean_dec(x_33); +uint8_t x_354; lean_dec(x_26); lean_dec(x_9); 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_71 = !lean_is_exclusive(x_35); -if (x_71 == 0) +x_354 = !lean_is_exclusive(x_344); +if (x_354 == 0) { -return x_35; +return x_344; } else { -lean_object* x_72; lean_object* x_73; lean_object* x_74; -x_72 = lean_ctor_get(x_35, 0); -x_73 = lean_ctor_get(x_35, 1); -lean_inc(x_73); -lean_inc(x_72); -lean_dec(x_35); -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_355; lean_object* x_356; lean_object* x_357; +x_355 = lean_ctor_get(x_344, 0); +x_356 = lean_ctor_get(x_344, 1); +lean_inc(x_356); +lean_inc(x_355); +lean_dec(x_344); +x_357 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_357, 0, x_355); +lean_ctor_set(x_357, 1, x_356); +return x_357; } } } else { -uint8_t x_75; +uint8_t x_358; lean_dec(x_26); lean_dec(x_9); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_75 = !lean_is_exclusive(x_32); -if (x_75 == 0) +x_358 = !lean_is_exclusive(x_340); +if (x_358 == 0) { -return x_32; +return x_340; } else { -lean_object* x_76; lean_object* x_77; lean_object* x_78; -x_76 = lean_ctor_get(x_32, 0); -x_77 = lean_ctor_get(x_32, 1); -lean_inc(x_77); -lean_inc(x_76); -lean_dec(x_32); -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_359; lean_object* x_360; lean_object* x_361; +x_359 = lean_ctor_get(x_340, 0); +x_360 = lean_ctor_get(x_340, 1); +lean_inc(x_360); +lean_inc(x_359); +lean_dec(x_340); +x_361 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_361, 0, x_359); +lean_ctor_set(x_361, 1, x_360); +return x_361; +} +} } } } else { -lean_object* x_79; +lean_object* x_362; lean_dec(x_15); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -x_79 = l_Lean_Meta_generalizeIndices(x_1, x_2, x_3, x_4, x_5, x_6, x_27); -if (lean_obj_tag(x_79) == 0) +x_362 = l_Lean_Meta_generalizeIndices_x27(x_1, x_2, x_24, x_3, x_4, x_5, x_6, x_27); +if (lean_obj_tag(x_362) == 0) { -lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; size_t x_85; size_t x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; -x_80 = lean_ctor_get(x_79, 0); -lean_inc(x_80); -x_81 = lean_ctor_get(x_79, 1); -lean_inc(x_81); -lean_dec(x_79); -x_82 = lean_ctor_get(x_80, 0); -lean_inc(x_82); -x_83 = lean_ctor_get(x_80, 2); -lean_inc(x_83); -x_84 = lean_ctor_get(x_80, 1); -lean_inc(x_84); -lean_dec(x_80); -x_85 = lean_array_size(x_84); -x_86 = 0; -x_87 = l_Array_mapMUnsafe_map___at_Lean_LocalContext_getFVars___spec__1(x_85, x_86, x_84); -x_88 = l_Lean_Meta_Grind_cases_throwInductiveExpected___closed__3; -lean_inc(x_82); -x_89 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_cases___lambda__1___boxed), 11, 6); -lean_closure_set(x_89, 0, x_82); -lean_closure_set(x_89, 1, x_88); -lean_closure_set(x_89, 2, x_83); -lean_closure_set(x_89, 3, x_26); -lean_closure_set(x_89, 4, x_87); -lean_closure_set(x_89, 5, x_9); -x_90 = l_Lean_MVarId_withContext___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___spec__2___rarg(x_82, x_89, x_3, x_4, x_5, x_6, x_81); -return x_90; +lean_object* x_363; lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; size_t x_368; size_t x_369; lean_object* x_370; lean_object* x_371; lean_object* x_372; lean_object* x_373; +x_363 = lean_ctor_get(x_362, 0); +lean_inc(x_363); +x_364 = lean_ctor_get(x_362, 1); +lean_inc(x_364); +lean_dec(x_362); +x_365 = lean_ctor_get(x_363, 0); +lean_inc(x_365); +x_366 = lean_ctor_get(x_363, 2); +lean_inc(x_366); +x_367 = lean_ctor_get(x_363, 1); +lean_inc(x_367); +lean_dec(x_363); +x_368 = lean_array_size(x_367); +x_369 = 0; +lean_inc(x_367); +x_370 = l_Array_mapMUnsafe_map___at_Lean_LocalContext_getFVars___spec__1(x_368, x_369, x_367); +x_371 = l_Lean_Meta_Grind_cases_throwInductiveExpected___closed__3; +lean_inc(x_365); +x_372 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_cases___lambda__12___boxed), 12, 7); +lean_closure_set(x_372, 0, x_365); +lean_closure_set(x_372, 1, x_371); +lean_closure_set(x_372, 2, x_366); +lean_closure_set(x_372, 3, x_26); +lean_closure_set(x_372, 4, x_370); +lean_closure_set(x_372, 5, x_9); +lean_closure_set(x_372, 6, x_367); +x_373 = l_Lean_MVarId_withContext___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___spec__2___rarg(x_365, x_372, x_3, x_4, x_5, x_6, x_364); +return x_373; } else { -uint8_t x_91; +uint8_t x_374; lean_dec(x_26); lean_dec(x_9); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_91 = !lean_is_exclusive(x_79); -if (x_91 == 0) +x_374 = !lean_is_exclusive(x_362); +if (x_374 == 0) { -return x_79; +return x_362; } else { -lean_object* x_92; lean_object* x_93; lean_object* x_94; -x_92 = lean_ctor_get(x_79, 0); -x_93 = lean_ctor_get(x_79, 1); -lean_inc(x_93); -lean_inc(x_92); -lean_dec(x_79); -x_94 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_94, 0, x_92); -lean_ctor_set(x_94, 1, x_93); -return x_94; +lean_object* x_375; lean_object* x_376; lean_object* x_377; +x_375 = lean_ctor_get(x_362, 0); +x_376 = lean_ctor_get(x_362, 1); +lean_inc(x_376); +lean_inc(x_375); +lean_dec(x_362); +x_377 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_377, 0, x_375); +lean_ctor_set(x_377, 1, x_376); +return x_377; } } } } else { -uint8_t x_95; +uint8_t x_378; lean_dec(x_15); lean_dec(x_9); lean_dec(x_6); @@ -1216,46 +11108,46 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_95 = !lean_is_exclusive(x_25); -if (x_95 == 0) +x_378 = !lean_is_exclusive(x_25); +if (x_378 == 0) { return x_25; } else { -lean_object* x_96; lean_object* x_97; lean_object* x_98; -x_96 = lean_ctor_get(x_25, 0); -x_97 = lean_ctor_get(x_25, 1); -lean_inc(x_97); -lean_inc(x_96); +lean_object* x_379; lean_object* x_380; lean_object* x_381; +x_379 = lean_ctor_get(x_25, 0); +x_380 = lean_ctor_get(x_25, 1); +lean_inc(x_380); +lean_inc(x_379); lean_dec(x_25); -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; +x_381 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_381, 0, x_379); +lean_ctor_set(x_381, 1, x_380); +return x_381; } } } else { -lean_object* x_99; lean_object* x_100; +lean_object* x_382; lean_object* x_383; lean_dec(x_20); lean_dec(x_18); lean_dec(x_9); -x_99 = lean_ctor_get(x_19, 1); -lean_inc(x_99); +x_382 = lean_ctor_get(x_19, 1); +lean_inc(x_382); lean_dec(x_19); -x_100 = l_Lean_Meta_Grind_cases_throwInductiveExpected(x_1, x_2, lean_box(0), x_15, x_3, x_4, x_5, x_6, x_99); +x_383 = l_Lean_Meta_Grind_cases_throwInductiveExpected(x_1, x_2, lean_box(0), x_15, x_3, x_4, x_5, x_6, x_382); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -return x_100; +return x_383; } } else { -uint8_t x_101; +uint8_t x_384; lean_dec(x_18); lean_dec(x_15); lean_dec(x_9); @@ -1265,42 +11157,42 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_101 = !lean_is_exclusive(x_19); -if (x_101 == 0) +x_384 = !lean_is_exclusive(x_19); +if (x_384 == 0) { return x_19; } else { -lean_object* x_102; lean_object* x_103; lean_object* x_104; -x_102 = lean_ctor_get(x_19, 0); -x_103 = lean_ctor_get(x_19, 1); -lean_inc(x_103); -lean_inc(x_102); +lean_object* x_385; lean_object* x_386; lean_object* x_387; +x_385 = lean_ctor_get(x_19, 0); +x_386 = lean_ctor_get(x_19, 1); +lean_inc(x_386); +lean_inc(x_385); lean_dec(x_19); -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; +x_387 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_387, 0, x_385); +lean_ctor_set(x_387, 1, x_386); +return x_387; } } } else { -lean_object* x_105; +lean_object* x_388; lean_dec(x_17); lean_dec(x_9); -x_105 = l_Lean_Meta_Grind_cases_throwInductiveExpected(x_1, x_2, lean_box(0), x_15, x_3, x_4, x_5, x_6, x_16); +x_388 = l_Lean_Meta_Grind_cases_throwInductiveExpected(x_1, x_2, lean_box(0), x_15, x_3, x_4, x_5, x_6, x_16); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -return x_105; +return x_388; } } else { -uint8_t x_106; +uint8_t x_389; lean_dec(x_9); lean_dec(x_6); lean_dec(x_5); @@ -1308,29 +11200,29 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_106 = !lean_is_exclusive(x_14); -if (x_106 == 0) +x_389 = !lean_is_exclusive(x_14); +if (x_389 == 0) { return x_14; } else { -lean_object* x_107; lean_object* x_108; lean_object* x_109; -x_107 = lean_ctor_get(x_14, 0); -x_108 = lean_ctor_get(x_14, 1); -lean_inc(x_108); -lean_inc(x_107); +lean_object* x_390; lean_object* x_391; lean_object* x_392; +x_390 = lean_ctor_get(x_14, 0); +x_391 = lean_ctor_get(x_14, 1); +lean_inc(x_391); +lean_inc(x_390); lean_dec(x_14); -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_392 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_392, 0, x_390); +lean_ctor_set(x_392, 1, x_391); +return x_392; } } } else { -uint8_t x_110; +uint8_t x_393; lean_dec(x_9); lean_dec(x_6); lean_dec(x_5); @@ -1338,52 +11230,52 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_110 = !lean_is_exclusive(x_11); -if (x_110 == 0) +x_393 = !lean_is_exclusive(x_11); +if (x_393 == 0) { return x_11; } else { -lean_object* x_111; lean_object* x_112; lean_object* x_113; -x_111 = lean_ctor_get(x_11, 0); -x_112 = lean_ctor_get(x_11, 1); -lean_inc(x_112); -lean_inc(x_111); +lean_object* x_394; lean_object* x_395; lean_object* x_396; +x_394 = lean_ctor_get(x_11, 0); +x_395 = lean_ctor_get(x_11, 1); +lean_inc(x_395); +lean_inc(x_394); lean_dec(x_11); -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_396 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_396, 0, x_394); +lean_ctor_set(x_396, 1, x_395); +return x_396; } } } else { -uint8_t x_114; +uint8_t x_397; 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_114 = !lean_is_exclusive(x_8); -if (x_114 == 0) +x_397 = !lean_is_exclusive(x_8); +if (x_397 == 0) { return x_8; } else { -lean_object* x_115; lean_object* x_116; lean_object* x_117; -x_115 = lean_ctor_get(x_8, 0); -x_116 = lean_ctor_get(x_8, 1); -lean_inc(x_116); -lean_inc(x_115); +lean_object* x_398; lean_object* x_399; lean_object* x_400; +x_398 = lean_ctor_get(x_8, 0); +x_399 = lean_ctor_get(x_8, 1); +lean_inc(x_399); +lean_inc(x_398); lean_dec(x_8); -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; +x_400 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_400, 0, x_398); +lean_ctor_set(x_400, 1, x_399); +return x_400; } } } @@ -1393,52 +11285,509 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases(lean_object* x_1, lean_object* { lean_object* x_8; lean_object* x_9; lean_inc(x_1); -x_8 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_cases___lambda__2), 7, 2); +x_8 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_cases___lambda__13), 7, 2); lean_closure_set(x_8, 0, x_1); lean_closure_set(x_8, 1, x_2); x_9 = l_Lean_MVarId_withContext___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___spec__2___rarg(x_1, x_8, x_3, x_4, x_5, x_6, x_7); return x_9; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___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_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___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]; _start: { -lean_object* x_10; -x_10 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__1___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_object* x_19; +x_19 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_6); +lean_dec(x_2); +return x_19; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__2___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +lean_object* x_18 = _args[17]; +_start: +{ +lean_object* x_19; +x_19 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___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); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_3); +return x_19; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__3___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +lean_object* x_18 = _args[17]; +_start: +{ +lean_object* x_19; +x_19 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_2); +return x_19; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__4___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +lean_object* x_18 = _args[17]; +_start: +{ +lean_object* x_19; +x_19 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_2); +return x_19; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__5___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +lean_object* x_18 = _args[17]; +_start: +{ +lean_object* x_19; +x_19 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__5(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_2); +return x_19; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__6___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +lean_object* x_18 = _args[17]; +_start: +{ +lean_object* x_19; +x_19 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__6(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_2); +return x_19; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__7___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +lean_object* x_18 = _args[17]; +_start: +{ +lean_object* x_19; +x_19 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__7(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_2); +return x_19; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__8___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +lean_object* x_18 = _args[17]; +_start: +{ +lean_object* x_19; +x_19 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__8(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_2); +return x_19; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__9___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +lean_object* x_18 = _args[17]; +_start: +{ +lean_object* x_19; +x_19 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__9(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_2); +return x_19; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__10___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +lean_object* x_18 = _args[17]; +_start: +{ +lean_object* x_19; +x_19 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__10(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_2); +return x_19; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__11___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +lean_object* x_18 = _args[17]; +_start: +{ +lean_object* x_19; +x_19 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__11(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_2); +return x_19; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__12___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +lean_object* x_18 = _args[17]; +_start: +{ +lean_object* x_19; +x_19 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__12(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_2); +return x_19; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__13___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +lean_object* x_18 = _args[17]; +_start: +{ +lean_object* x_19; +x_19 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___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, x_14, x_15, x_16, x_17, x_18); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_2); +return x_19; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___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_13; +x_13 = l_Lean_Meta_Grind_cases___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_10; +return x_13; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -lean_object* x_16; -x_16 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___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); -lean_dec(x_6); +lean_object* x_13; +x_13 = l_Lean_Meta_Grind_cases___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); lean_dec(x_5); -return x_16; +return x_13; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -lean_object* x_15; -x_15 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___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); +lean_object* x_13; +x_13 = l_Lean_Meta_Grind_cases___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); lean_dec(x_5); -lean_dec(x_4); -return x_15; +return x_13; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +x_13 = l_Lean_Meta_Grind_cases___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); +lean_dec(x_5); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___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) { +_start: +{ +lean_object* x_13; +x_13 = l_Lean_Meta_Grind_cases___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); +lean_dec(x_5); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___lambda__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +x_13 = l_Lean_Meta_Grind_cases___lambda__6(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_5); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___lambda__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +x_13 = l_Lean_Meta_Grind_cases___lambda__7(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_5); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___lambda__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +x_13 = l_Lean_Meta_Grind_cases___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); +lean_dec(x_5); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___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: +{ +lean_object* x_13; +x_13 = l_Lean_Meta_Grind_cases___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); +lean_dec(x_5); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___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) { +_start: +{ +lean_object* x_13; +x_13 = l_Lean_Meta_Grind_cases___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); +lean_dec(x_5); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___lambda__11___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +x_13 = l_Lean_Meta_Grind_cases___lambda__11(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_5); +return x_13; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___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_Meta_Grind_cases___lambda__12___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -lean_object* x_12; -x_12 = l_Lean_Meta_Grind_cases___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_object* x_13; +x_13 = l_Lean_Meta_Grind_cases___lambda__12(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); lean_dec(x_5); -return x_12; +return x_13; } } lean_object* initialize_Lean_Meta_Tactic_Cases(uint8_t builtin, lean_object*); @@ -1472,8 +11821,15 @@ l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__1___closed__3 = _i lean_mark_persistent(l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__1___closed__3); l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__1___closed__4 = _init_l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__1___closed__4(); lean_mark_persistent(l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__1___closed__4); -l_Lean_Meta_Grind_cases___lambda__1___closed__1 = _init_l_Lean_Meta_Grind_cases___lambda__1___closed__1(); -lean_mark_persistent(l_Lean_Meta_Grind_cases___lambda__1___closed__1); +l_Lean_Meta_Grind_cases___lambda__12___closed__1 = _init_l_Lean_Meta_Grind_cases___lambda__12___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_cases___lambda__12___closed__1); +l_Lean_Meta_Grind_cases___lambda__13___closed__1 = _init_l_Lean_Meta_Grind_cases___lambda__13___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_cases___lambda__13___closed__1); +l_Lean_Meta_Grind_cases___lambda__13___closed__2 = _init_l_Lean_Meta_Grind_cases___lambda__13___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_cases___lambda__13___closed__2); +l_Lean_Meta_Grind_cases___lambda__13___closed__3 = _init_l_Lean_Meta_Grind_cases___lambda__13___closed__3(); +l_Lean_Meta_Grind_cases___lambda__13___closed__4 = _init_l_Lean_Meta_Grind_cases___lambda__13___closed__4(); +lean_mark_persistent(l_Lean_Meta_Grind_cases___lambda__13___closed__4); return lean_io_result_mk_ok(lean_box(0)); } #ifdef __cplusplus diff --git a/stage0/stdlib/Lean/Meta/Tactic/Grind/CasesMatch.c b/stage0/stdlib/Lean/Meta/Tactic/Grind/CasesMatch.c new file mode 100644 index 000000000000..eed01a9a4bfc --- /dev/null +++ b/stage0/stdlib/Lean/Meta/Tactic/Grind/CasesMatch.c @@ -0,0 +1,1154 @@ +// Lean compiler output +// Module: Lean.Meta.Tactic.Grind.CasesMatch +// Imports: Lean.Meta.Tactic.Util Lean.Meta.Tactic.Cases Lean.Meta.Match.MatcherApp +#include +#if defined(__clang__) +#pragma clang diagnostic ignored "-Wunused-parameter" +#pragma clang diagnostic ignored "-Wunused-label" +#elif defined(__GNUC__) && !defined(__CLANG__) +#pragma GCC diagnostic ignored "-Wunused-parameter" +#pragma GCC diagnostic ignored "-Wunused-label" +#pragma GCC diagnostic ignored "-Wunused-but-set-variable" +#endif +#ifdef __cplusplus +extern "C" { +#endif +lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_casesMatch_updateTags(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_throwTacticEx___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_mkAppN(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_casesMatch_updateTags___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Expr_sort___override(lean_object*); +static lean_object* l_Lean_Meta_Grind_casesMatch_mkMotiveAndRefls___closed__2; +static lean_object* l_Lean_Meta_Grind_casesMatch___lambda__2___closed__1; +lean_object* lean_array_fget(lean_object*, lean_object*); +lean_object* l_Lean_MVarId_getTag(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_matchMatcherApp_x3f___at___private_Lean_Meta_Tactic_Split_0__Lean_Meta_Split_generalizeMatchDiscrs_mkNewTarget___spec__1(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_casesMatch___lambda__2___closed__6; +lean_object* l_Lean_stringToMessageData(lean_object*); +lean_object* l_Lean_Level_ofNat(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_casesMatch_updateTags___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_casesMatch_mkMotiveAndRefls___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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_Meta_Grind_casesMatch___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_Grind_casesMatch_mkMotiveAndRefls___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_forallBoundedTelescope___at_Lean_Meta_arrowDomainsN___spec__6___rarg(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_MVarId_getType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_casesMatch___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_Grind_casesMatch___lambda__2___closed__5; +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_getLevel(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_outOfBounds___rarg(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_casesMatch(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_array_to_list(lean_object*); +LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Meta_Grind_casesMatch___spec__1(lean_object*, lean_object*); +lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); +extern lean_object* l_Lean_instInhabitedExpr; +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_casesMatch_updateTags___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_casesMatch_mkMotiveAndRefls___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_casesMatch_mkMotiveAndRefls___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_Grind_casesMatch___lambda__2___closed__7; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_casesMatch_mkMotiveAndRefls(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_forallMetaBoundedTelescope(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Expr_app___override(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_casesMatch_mkMotiveAndRefls___lambda__1(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*); +uint8_t lean_nat_dec_lt(lean_object*, lean_object*); +lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); +lean_object* l_Lean_indentExpr(lean_object*); +lean_object* l_Lean_Meta_mkForallFVars(lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_array_set(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_withNewEqs___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_get_match_equations_for(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Expr_getAppFn(lean_object*); +lean_object* l_List_reverse___rarg(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_casesMatch_mkMotiveAndRefls___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +size_t lean_usize_add(size_t, size_t); +lean_object* lean_array_uget(lean_object*, size_t); +size_t lean_array_size(lean_object*); +static lean_object* l_Lean_Meta_Grind_casesMatch___lambda__2___closed__4; +lean_object* l_Lean_MVarId_setTag(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_array_get_size(lean_object*); +static lean_object* l_Lean_Meta_Grind_casesMatch___lambda__2___closed__2; +lean_object* lean_infer_type(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_Grind_casesMatch___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_mkLambdaFVars(lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_nat_add(lean_object*, lean_object*); +lean_object* l_Lean_Expr_mvarId_x21(lean_object*); +static lean_object* l_Lean_Meta_Grind_casesMatch___lambda__2___closed__3; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_casesMatch_mkMotiveAndRefls___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l_Lean_MVarId_getType(x_1, 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; uint8_t x_14; uint8_t x_15; lean_object* x_16; +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 = 1; +x_15 = 1; +x_16 = l_Lean_Meta_mkForallFVars(x_3, x_11, x_13, x_14, x_15, x_5, x_6, x_7, x_8, x_12); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; +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_Meta_mkLambdaFVars(x_2, x_17, x_13, x_14, x_13, x_15, x_5, x_6, x_7, x_8, x_18); +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; +x_21 = lean_ctor_get(x_19, 0); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_4); +lean_ctor_set(x_19, 0, x_22); +return x_19; +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +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_dec(x_19); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_4); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set(x_26, 1, x_24); +return x_26; +} +} +else +{ +uint8_t x_27; +lean_dec(x_4); +x_27 = !lean_is_exclusive(x_19); +if (x_27 == 0) +{ +return x_19; +} +else +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_28 = lean_ctor_get(x_19, 0); +x_29 = lean_ctor_get(x_19, 1); +lean_inc(x_29); +lean_inc(x_28); +lean_dec(x_19); +x_30 = lean_alloc_ctor(1, 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_31; +lean_dec(x_4); +x_31 = !lean_is_exclusive(x_16); +if (x_31 == 0) +{ +return x_16; +} +else +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_16, 0); +x_33 = lean_ctor_get(x_16, 1); +lean_inc(x_33); +lean_inc(x_32); +lean_dec(x_16); +x_34 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +return x_34; +} +} +} +else +{ +uint8_t x_35; +lean_dec(x_4); +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_Meta_Grind_casesMatch_mkMotiveAndRefls___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; lean_object* x_11; +lean_inc(x_3); +x_10 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_casesMatch_mkMotiveAndRefls___lambda__1___boxed), 9, 2); +lean_closure_set(x_10, 0, x_1); +lean_closure_set(x_10, 1, x_3); +x_11 = l_Lean_Meta_withNewEqs___rarg(x_2, x_3, x_10, x_5, x_6, x_7, x_8, x_9); +return x_11; +} +} +static lean_object* _init_l_Lean_Meta_Grind_casesMatch_mkMotiveAndRefls___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = l_Lean_Level_ofNat(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Grind_casesMatch_mkMotiveAndRefls___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_casesMatch_mkMotiveAndRefls___closed__1; +x_2 = l_Lean_Expr_sort___override(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_casesMatch_mkMotiveAndRefls(lean_object* x_1, lean_object* x_2, lean_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_Expr_getAppFn(x_2); +x_10 = lean_ctor_get(x_3, 4); +lean_inc(x_10); +x_11 = l_Lean_mkAppN(x_9, x_10); +lean_dec(x_10); +x_12 = l_Lean_Meta_Grind_casesMatch_mkMotiveAndRefls___closed__2; +x_13 = l_Lean_Expr_app___override(x_11, x_12); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +x_14 = lean_infer_type(x_13, x_4, x_5, x_6, x_7, x_8); +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; uint8_t x_21; lean_object* x_22; +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +x_16 = lean_ctor_get(x_14, 1); +lean_inc(x_16); +lean_dec(x_14); +x_17 = lean_ctor_get(x_3, 6); +lean_inc(x_17); +lean_dec(x_3); +x_18 = lean_array_get_size(x_17); +x_19 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_19, 0, x_18); +x_20 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_casesMatch_mkMotiveAndRefls___lambda__2___boxed), 9, 2); +lean_closure_set(x_20, 0, x_1); +lean_closure_set(x_20, 1, x_17); +x_21 = 0; +x_22 = l_Lean_Meta_forallBoundedTelescope___at_Lean_Meta_arrowDomainsN___spec__6___rarg(x_15, x_19, x_20, x_21, x_4, x_5, x_6, x_7, x_16); +return x_22; +} +else +{ +uint8_t x_23; +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_23 = !lean_is_exclusive(x_14); +if (x_23 == 0) +{ +return x_14; +} +else +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_14, 0); +x_25 = lean_ctor_get(x_14, 1); +lean_inc(x_25); +lean_inc(x_24); +lean_dec(x_14); +x_26 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_26, 0, x_24); +lean_ctor_set(x_26, 1, x_25); +return x_26; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_casesMatch_mkMotiveAndRefls___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l_Lean_Meta_Grind_casesMatch_mkMotiveAndRefls___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_3); +lean_dec(x_2); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_casesMatch_mkMotiveAndRefls___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_Meta_Grind_casesMatch_mkMotiveAndRefls___lambda__2(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_Meta_Grind_casesMatch_mkMotiveAndRefls___boxed(lean_object* x_1, lean_object* x_2, lean_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_Grind_casesMatch_mkMotiveAndRefls(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_2); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_casesMatch_updateTags___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: +{ +uint8_t x_13; +x_13 = lean_usize_dec_lt(x_6, x_5); +if (x_13 == 0) +{ +lean_object* x_14; +lean_dec(x_2); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_7); +lean_ctor_set(x_14, 1, x_12); +return x_14; +} +else +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; size_t x_22; size_t x_23; +x_15 = lean_array_uget(x_4, x_6); +x_16 = l_Lean_Expr_mvarId_x21(x_15); +lean_dec(x_15); +lean_inc(x_7); +lean_inc(x_2); +x_17 = l_Lean_Name_num___override(x_2, x_7); +x_18 = l_Lean_MVarId_setTag(x_16, x_17, x_8, x_9, x_10, x_11, x_12); +x_19 = lean_ctor_get(x_18, 1); +lean_inc(x_19); +lean_dec(x_18); +x_20 = lean_unsigned_to_nat(1u); +x_21 = lean_nat_add(x_7, x_20); +lean_dec(x_7); +x_22 = 1; +x_23 = lean_usize_add(x_6, x_22); +x_6 = x_23; +x_7 = x_21; +x_12 = x_19; +goto _start; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_casesMatch_updateTags(lean_object* x_1, 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_MVarId_getTag(x_1, x_3, x_4, x_5, x_6, x_7); +if (lean_obj_tag(x_8) == 0) +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_9 = lean_ctor_get(x_8, 0); +lean_inc(x_9); +x_10 = lean_ctor_get(x_8, 1); +lean_inc(x_10); +lean_dec(x_8); +x_11 = lean_array_get_size(x_2); +x_12 = lean_unsigned_to_nat(1u); +x_13 = lean_nat_dec_eq(x_11, x_12); +if (x_13 == 0) +{ +lean_object* x_14; size_t x_15; size_t x_16; lean_object* x_17; uint8_t x_18; +lean_dec(x_11); +x_14 = lean_box(0); +x_15 = lean_array_size(x_2); +x_16 = 0; +x_17 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_casesMatch_updateTags___spec__1(x_2, x_9, x_14, x_2, x_15, x_16, x_12, x_3, x_4, x_5, x_6, x_10); +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; uint8_t x_25; +x_24 = lean_unsigned_to_nat(0u); +x_25 = lean_nat_dec_lt(x_24, x_11); +lean_dec(x_11); +if (x_25 == 0) +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_26 = l_Lean_instInhabitedExpr; +x_27 = l_outOfBounds___rarg(x_26); +x_28 = l_Lean_Expr_mvarId_x21(x_27); +lean_dec(x_27); +x_29 = l_Lean_MVarId_setTag(x_28, x_9, x_3, x_4, x_5, x_6, x_10); +return x_29; +} +else +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_30 = lean_array_fget(x_2, x_24); +x_31 = l_Lean_Expr_mvarId_x21(x_30); +lean_dec(x_30); +x_32 = l_Lean_MVarId_setTag(x_31, x_9, x_3, x_4, x_5, x_6, x_10); +return x_32; +} +} +} +else +{ +uint8_t x_33; +x_33 = !lean_is_exclusive(x_8); +if (x_33 == 0) +{ +return x_8; +} +else +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_8, 0); +x_35 = lean_ctor_get(x_8, 1); +lean_inc(x_35); +lean_inc(x_34); +lean_dec(x_8); +x_36 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_36, 0, x_34); +lean_ctor_set(x_36, 1, x_35); +return x_36; +} +} +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_casesMatch_updateTags___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +size_t x_13; size_t x_14; lean_object* x_15; +x_13 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_14 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_15 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_casesMatch_updateTags___spec__1(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_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +return x_15; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_casesMatch_updateTags___boxed(lean_object* x_1, 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_Grind_casesMatch_updateTags(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_8; +} +} +LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Meta_Grind_casesMatch___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_Expr_mvarId_x21(x_5); +lean_dec(x_5); +lean_ctor_set(x_1, 1, x_2); +lean_ctor_set(x_1, 0, x_7); +{ +lean_object* _tmp_0 = x_6; +lean_object* _tmp_1 = x_1; +x_1 = _tmp_0; +x_2 = _tmp_1; +} +goto _start; +} +else +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_9 = lean_ctor_get(x_1, 0); +x_10 = lean_ctor_get(x_1, 1); +lean_inc(x_10); +lean_inc(x_9); +lean_dec(x_1); +x_11 = l_Lean_Expr_mvarId_x21(x_9); +lean_dec(x_9); +x_12 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_2); +x_1 = x_10; +x_2 = x_12; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_casesMatch___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; +x_12 = lean_ctor_get(x_1, 0); +lean_inc(x_12); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +x_13 = lean_get_match_equations_for(x_12, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_13) == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_14 = lean_ctor_get(x_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, 1); +lean_inc(x_16); +lean_dec(x_14); +x_17 = lean_array_to_list(x_5); +x_18 = l_Lean_Expr_const___override(x_16, x_17); +x_19 = lean_ctor_get(x_1, 4); +lean_inc(x_19); +x_20 = l_Lean_mkAppN(x_18, x_19); +lean_dec(x_19); +x_21 = l_Lean_Expr_app___override(x_20, x_2); +x_22 = lean_ctor_get(x_1, 6); +lean_inc(x_22); +x_23 = l_Lean_mkAppN(x_21, x_22); +lean_dec(x_22); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_23); +x_24 = lean_infer_type(x_23, x_7, x_8, x_9, x_10, x_15); +if (lean_obj_tag(x_24) == 0) +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; lean_object* x_30; +x_25 = lean_ctor_get(x_24, 0); +lean_inc(x_25); +x_26 = lean_ctor_get(x_24, 1); +lean_inc(x_26); +lean_dec(x_24); +x_27 = lean_ctor_get(x_1, 8); +lean_inc(x_27); +lean_dec(x_1); +x_28 = lean_array_get_size(x_27); +lean_dec(x_27); +x_29 = 2; +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +x_30 = l_Lean_Meta_forallMetaBoundedTelescope(x_25, x_28, x_29, x_7, x_8, x_9, x_10, x_26); +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; +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 = l_Lean_mkAppN(x_23, x_33); +x_35 = l_Lean_mkAppN(x_34, x_3); +lean_inc(x_4); +x_36 = l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_4, x_35, x_7, x_8, x_9, x_10, x_32); +x_37 = lean_ctor_get(x_36, 1); +lean_inc(x_37); +lean_dec(x_36); +x_38 = l_Lean_Meta_Grind_casesMatch_updateTags(x_4, x_33, x_7, x_8, x_9, x_10, x_37); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +if (lean_obj_tag(x_38) == 0) +{ +uint8_t x_39; +x_39 = !lean_is_exclusive(x_38); +if (x_39 == 0) +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_40 = lean_ctor_get(x_38, 0); +lean_dec(x_40); +x_41 = lean_array_to_list(x_33); +x_42 = lean_box(0); +x_43 = l_List_mapTR_loop___at_Lean_Meta_Grind_casesMatch___spec__1(x_41, x_42); +lean_ctor_set(x_38, 0, x_43); +return x_38; +} +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_38, 1); +lean_inc(x_44); +lean_dec(x_38); +x_45 = lean_array_to_list(x_33); +x_46 = lean_box(0); +x_47 = l_List_mapTR_loop___at_Lean_Meta_Grind_casesMatch___spec__1(x_45, 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_44); +return x_48; +} +} +else +{ +uint8_t x_49; +lean_dec(x_33); +x_49 = !lean_is_exclusive(x_38); +if (x_49 == 0) +{ +return x_38; +} +else +{ +lean_object* x_50; lean_object* x_51; lean_object* x_52; +x_50 = lean_ctor_get(x_38, 0); +x_51 = lean_ctor_get(x_38, 1); +lean_inc(x_51); +lean_inc(x_50); +lean_dec(x_38); +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_53; +lean_dec(x_23); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_4); +x_53 = !lean_is_exclusive(x_30); +if (x_53 == 0) +{ +return x_30; +} +else +{ +lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_54 = lean_ctor_get(x_30, 0); +x_55 = lean_ctor_get(x_30, 1); +lean_inc(x_55); +lean_inc(x_54); +lean_dec(x_30); +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_57; +lean_dec(x_23); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_4); +lean_dec(x_1); +x_57 = !lean_is_exclusive(x_24); +if (x_57 == 0) +{ +return x_24; +} +else +{ +lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_58 = lean_ctor_get(x_24, 0); +x_59 = lean_ctor_get(x_24, 1); +lean_inc(x_59); +lean_inc(x_58); +lean_dec(x_24); +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_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_61 = !lean_is_exclusive(x_13); +if (x_61 == 0) +{ +return x_13; +} +else +{ +lean_object* x_62; lean_object* x_63; lean_object* x_64; +x_62 = lean_ctor_get(x_13, 0); +x_63 = lean_ctor_get(x_13, 1); +lean_inc(x_63); +lean_inc(x_62); +lean_dec(x_13); +x_64 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_64, 0, x_62); +lean_ctor_set(x_64, 1, x_63); +return x_64; +} +} +} +} +static lean_object* _init_l_Lean_Meta_Grind_casesMatch___lambda__2___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("grind", 5, 5); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_casesMatch___lambda__2___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("casesMatch", 10, 10); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_casesMatch___lambda__2___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_casesMatch___lambda__2___closed__1; +x_2 = l_Lean_Meta_Grind_casesMatch___lambda__2___closed__2; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind_casesMatch___lambda__2___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("`match`-expression expected", 27, 27); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_casesMatch___lambda__2___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_casesMatch___lambda__2___closed__4; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Grind_casesMatch___lambda__2___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("", 0, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_casesMatch___lambda__2___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_casesMatch___lambda__2___closed__6; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_casesMatch___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; lean_object* x_9; +x_8 = 0; +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_1); +x_9 = l_Lean_Meta_matchMatcherApp_x3f___at___private_Lean_Meta_Tactic_Split_0__Lean_Meta_Split_generalizeMatchDiscrs_mkNewTarget___spec__1(x_1, x_8, 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) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_11 = lean_ctor_get(x_9, 1); +lean_inc(x_11); +lean_dec(x_9); +x_12 = l_Lean_indentExpr(x_1); +x_13 = l_Lean_Meta_Grind_casesMatch___lambda__2___closed__5; +x_14 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_12); +x_15 = l_Lean_Meta_Grind_casesMatch___lambda__2___closed__7; +x_16 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_16, 0, x_14); +lean_ctor_set(x_16, 1, x_15); +x_17 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_17, 0, x_16); +x_18 = l_Lean_Meta_Grind_casesMatch___lambda__2___closed__3; +x_19 = l_Lean_Meta_throwTacticEx___rarg(x_18, x_2, x_17, x_3, x_4, x_5, x_6, x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_19; +} +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_9, 1); +lean_inc(x_20); +lean_dec(x_9); +x_21 = lean_ctor_get(x_10, 0); +lean_inc(x_21); +lean_dec(x_10); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_21); +lean_inc(x_2); +x_22 = l_Lean_Meta_Grind_casesMatch_mkMotiveAndRefls(x_2, x_1, x_21, x_3, x_4, x_5, x_6, x_20); +lean_dec(x_1); +if (lean_obj_tag(x_22) == 0) +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +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); +x_26 = lean_ctor_get(x_23, 1); +lean_inc(x_26); +lean_dec(x_23); +lean_inc(x_2); +x_27 = l_Lean_MVarId_getType(x_2, x_3, x_4, x_5, x_6, x_24); +if (lean_obj_tag(x_27) == 0) +{ +lean_object* x_28; +x_28 = lean_ctor_get(x_21, 2); +lean_inc(x_28); +if (lean_obj_tag(x_28) == 0) +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_29 = lean_ctor_get(x_27, 1); +lean_inc(x_29); +lean_dec(x_27); +x_30 = lean_ctor_get(x_21, 1); +lean_inc(x_30); +x_31 = lean_box(0); +x_32 = l_Lean_Meta_Grind_casesMatch___lambda__1(x_21, x_25, x_26, x_2, x_30, x_31, x_3, x_4, x_5, x_6, x_29); +lean_dec(x_26); +return x_32; +} +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_27, 0); +lean_inc(x_33); +x_34 = lean_ctor_get(x_27, 1); +lean_inc(x_34); +lean_dec(x_27); +x_35 = lean_ctor_get(x_21, 1); +lean_inc(x_35); +x_36 = lean_ctor_get(x_28, 0); +lean_inc(x_36); +lean_dec(x_28); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_37 = l_Lean_Meta_getLevel(x_33, x_3, x_4, x_5, x_6, x_34); +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; +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_array_set(x_35, x_36, x_38); +lean_dec(x_36); +x_41 = lean_box(0); +x_42 = l_Lean_Meta_Grind_casesMatch___lambda__1(x_21, x_25, x_26, x_2, x_40, x_41, x_3, x_4, x_5, x_6, x_39); +lean_dec(x_26); +return x_42; +} +else +{ +uint8_t x_43; +lean_dec(x_36); +lean_dec(x_35); +lean_dec(x_26); +lean_dec(x_25); +lean_dec(x_21); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_43 = !lean_is_exclusive(x_37); +if (x_43 == 0) +{ +return x_37; +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_37, 0); +x_45 = lean_ctor_get(x_37, 1); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_37); +x_46 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_46, 0, x_44); +lean_ctor_set(x_46, 1, x_45); +return x_46; +} +} +} +} +else +{ +uint8_t x_47; +lean_dec(x_26); +lean_dec(x_25); +lean_dec(x_21); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_47 = !lean_is_exclusive(x_27); +if (x_47 == 0) +{ +return x_27; +} +else +{ +lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_48 = lean_ctor_get(x_27, 0); +x_49 = lean_ctor_get(x_27, 1); +lean_inc(x_49); +lean_inc(x_48); +lean_dec(x_27); +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_51; +lean_dec(x_21); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_51 = !lean_is_exclusive(x_22); +if (x_51 == 0) +{ +return x_22; +} +else +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; +x_52 = lean_ctor_get(x_22, 0); +x_53 = lean_ctor_get(x_22, 1); +lean_inc(x_53); +lean_inc(x_52); +lean_dec(x_22); +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_55; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_55 = !lean_is_exclusive(x_9); +if (x_55 == 0) +{ +return x_9; +} +else +{ +lean_object* x_56; lean_object* x_57; lean_object* x_58; +x_56 = lean_ctor_get(x_9, 0); +x_57 = lean_ctor_get(x_9, 1); +lean_inc(x_57); +lean_inc(x_56); +lean_dec(x_9); +x_58 = 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_EXPORT lean_object* l_Lean_Meta_Grind_casesMatch(lean_object* x_1, 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_inc(x_1); +x_8 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_casesMatch___lambda__2), 7, 2); +lean_closure_set(x_8, 0, x_2); +lean_closure_set(x_8, 1, x_1); +x_9 = l_Lean_MVarId_withContext___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___spec__2___rarg(x_1, x_8, x_3, x_4, x_5, x_6, x_7); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_casesMatch___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_Meta_Grind_casesMatch___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_3); +return x_12; +} +} +lean_object* initialize_Lean_Meta_Tactic_Util(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Meta_Tactic_Cases(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Meta_Match_MatcherApp(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Lean_Meta_Tactic_Grind_CasesMatch(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_Tactic_Util(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Lean_Meta_Tactic_Cases(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Lean_Meta_Match_MatcherApp(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l_Lean_Meta_Grind_casesMatch_mkMotiveAndRefls___closed__1 = _init_l_Lean_Meta_Grind_casesMatch_mkMotiveAndRefls___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_casesMatch_mkMotiveAndRefls___closed__1); +l_Lean_Meta_Grind_casesMatch_mkMotiveAndRefls___closed__2 = _init_l_Lean_Meta_Grind_casesMatch_mkMotiveAndRefls___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_casesMatch_mkMotiveAndRefls___closed__2); +l_Lean_Meta_Grind_casesMatch___lambda__2___closed__1 = _init_l_Lean_Meta_Grind_casesMatch___lambda__2___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_casesMatch___lambda__2___closed__1); +l_Lean_Meta_Grind_casesMatch___lambda__2___closed__2 = _init_l_Lean_Meta_Grind_casesMatch___lambda__2___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_casesMatch___lambda__2___closed__2); +l_Lean_Meta_Grind_casesMatch___lambda__2___closed__3 = _init_l_Lean_Meta_Grind_casesMatch___lambda__2___closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind_casesMatch___lambda__2___closed__3); +l_Lean_Meta_Grind_casesMatch___lambda__2___closed__4 = _init_l_Lean_Meta_Grind_casesMatch___lambda__2___closed__4(); +lean_mark_persistent(l_Lean_Meta_Grind_casesMatch___lambda__2___closed__4); +l_Lean_Meta_Grind_casesMatch___lambda__2___closed__5 = _init_l_Lean_Meta_Grind_casesMatch___lambda__2___closed__5(); +lean_mark_persistent(l_Lean_Meta_Grind_casesMatch___lambda__2___closed__5); +l_Lean_Meta_Grind_casesMatch___lambda__2___closed__6 = _init_l_Lean_Meta_Grind_casesMatch___lambda__2___closed__6(); +lean_mark_persistent(l_Lean_Meta_Grind_casesMatch___lambda__2___closed__6); +l_Lean_Meta_Grind_casesMatch___lambda__2___closed__7 = _init_l_Lean_Meta_Grind_casesMatch___lambda__2___closed__7(); +lean_mark_persistent(l_Lean_Meta_Grind_casesMatch___lambda__2___closed__7); +return lean_io_result_mk_ok(lean_box(0)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Lean/Meta/Tactic/Grind/Combinators.c b/stage0/stdlib/Lean/Meta/Tactic/Grind/Combinators.c new file mode 100644 index 000000000000..a1cc45a523d7 --- /dev/null +++ b/stage0/stdlib/Lean/Meta/Tactic/Grind/Combinators.c @@ -0,0 +1,1388 @@ +// Lean compiler output +// Module: Lean.Meta.Tactic.Grind.Combinators +// Imports: Lean.Meta.Tactic.Grind.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 +lean_object* l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___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_Meta_Grind_toGrindTactic___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_toGrindTactic___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_Meta_Grind_toGrindTactic___closed__2; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instOrElseGrindTactic(lean_object*, lean_object*, 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_Grind_toGrindTactic___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_Grind_GrindTactic_iterate_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* l_List_appendTR___rarg(lean_object*, lean_object*); +lean_object* lean_st_ref_get(lean_object*, lean_object*); +lean_object* lean_st_mk_ref(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_toGrindTactic___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_Grind_GrindTactic_iterate(lean_object*, 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_Grind_GrindTactic_x27_toGrindTactic(lean_object*, 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_Grind_GrindTactic_andThen(lean_object*, lean_object*, 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_Grind_toGrindTactic___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___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_Meta_Grind_toGrindTactic(lean_object*, 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_Grind_toGrindTactic___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_Grind_GrindTactic_try(lean_object*, 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_Grind_applyToAll_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* l_List_reverse___rarg(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GrindTactic_orElse(lean_object*, lean_object*, 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_Grind_applyToAll(lean_object*, 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_Grind_instAndThenGrindTactic(lean_object*, lean_object*, 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_Grind_GrindTactic_try(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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_2); +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) +{ +lean_object* x_12; +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +if (lean_obj_tag(x_12) == 0) +{ +uint8_t x_13; +x_13 = !lean_is_exclusive(x_11); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_14 = lean_ctor_get(x_11, 0); +lean_dec(x_14); +x_15 = lean_box(0); +x_16 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_16, 0, x_2); +lean_ctor_set(x_16, 1, x_15); +x_17 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_17, 0, x_16); +lean_ctor_set(x_11, 0, x_17); +return x_11; +} +else +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_18 = lean_ctor_get(x_11, 1); +lean_inc(x_18); +lean_dec(x_11); +x_19 = lean_box(0); +x_20 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_20, 0, x_2); +lean_ctor_set(x_20, 1, x_19); +x_21 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_21, 0, x_20); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_18); +return x_22; +} +} +else +{ +uint8_t x_23; +lean_dec(x_2); +x_23 = !lean_is_exclusive(x_11); +if (x_23 == 0) +{ +lean_object* x_24; uint8_t x_25; +x_24 = lean_ctor_get(x_11, 0); +lean_dec(x_24); +x_25 = !lean_is_exclusive(x_12); +if (x_25 == 0) +{ +return x_11; +} +else +{ +lean_object* x_26; lean_object* x_27; +x_26 = lean_ctor_get(x_12, 0); +lean_inc(x_26); +lean_dec(x_12); +x_27 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_11, 0, x_27); +return x_11; +} +} +else +{ +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_11, 1); +lean_inc(x_28); +lean_dec(x_11); +x_29 = lean_ctor_get(x_12, 0); +lean_inc(x_29); +if (lean_is_exclusive(x_12)) { + lean_ctor_release(x_12, 0); + x_30 = x_12; +} else { + lean_dec_ref(x_12); + x_30 = lean_box(0); +} +if (lean_is_scalar(x_30)) { + x_31 = lean_alloc_ctor(1, 1, 0); +} else { + x_31 = x_30; +} +lean_ctor_set(x_31, 0, x_29); +x_32 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_28); +return x_32; +} +} +} +else +{ +uint8_t x_33; +lean_dec(x_2); +x_33 = !lean_is_exclusive(x_11); +if (x_33 == 0) +{ +return x_11; +} +else +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_11, 0); +x_35 = lean_ctor_get(x_11, 1); +lean_inc(x_35); +lean_inc(x_34); +lean_dec(x_11); +x_36 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_36, 0, x_34); +lean_ctor_set(x_36, 1, x_35); +return x_36; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_applyToAll_go(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_12; lean_object* x_13; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_12 = l_List_reverse___rarg(x_3); +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_11); +return x_13; +} +else +{ +uint8_t x_14; +x_14 = !lean_is_exclusive(x_2); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_ctor_get(x_2, 0); +x_16 = lean_ctor_get(x_2, 1); +lean_inc(x_1); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_15); +x_17 = lean_apply_9(x_1, x_15, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_17) == 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_19; +x_19 = lean_ctor_get(x_17, 1); +lean_inc(x_19); +lean_dec(x_17); +lean_ctor_set(x_2, 1, x_3); +{ +lean_object* _tmp_1 = x_16; +lean_object* _tmp_2 = x_2; +lean_object* _tmp_10 = x_19; +x_2 = _tmp_1; +x_3 = _tmp_2; +x_11 = _tmp_10; +} +goto _start; +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +lean_free_object(x_2); +lean_dec(x_15); +x_21 = lean_ctor_get(x_17, 1); +lean_inc(x_21); +lean_dec(x_17); +x_22 = lean_ctor_get(x_18, 0); +lean_inc(x_22); +lean_dec(x_18); +x_23 = l_List_appendTR___rarg(x_22, x_3); +x_2 = x_16; +x_3 = x_23; +x_11 = x_21; +goto _start; +} +} +else +{ +uint8_t x_25; +lean_free_object(x_2); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_25 = !lean_is_exclusive(x_17); +if (x_25 == 0) +{ +return x_17; +} +else +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_26 = lean_ctor_get(x_17, 0); +x_27 = lean_ctor_get(x_17, 1); +lean_inc(x_27); +lean_inc(x_26); +lean_dec(x_17); +x_28 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_28, 0, x_26); +lean_ctor_set(x_28, 1, x_27); +return x_28; +} +} +} +else +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_29 = lean_ctor_get(x_2, 0); +x_30 = lean_ctor_get(x_2, 1); +lean_inc(x_30); +lean_inc(x_29); +lean_dec(x_2); +lean_inc(x_1); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_29); +x_31 = lean_apply_9(x_1, x_29, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_31) == 0) +{ +lean_object* x_32; +x_32 = lean_ctor_get(x_31, 0); +lean_inc(x_32); +if (lean_obj_tag(x_32) == 0) +{ +lean_object* x_33; lean_object* x_34; +x_33 = lean_ctor_get(x_31, 1); +lean_inc(x_33); +lean_dec(x_31); +x_34 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_34, 0, x_29); +lean_ctor_set(x_34, 1, x_3); +x_2 = x_30; +x_3 = x_34; +x_11 = x_33; +goto _start; +} +else +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; +lean_dec(x_29); +x_36 = lean_ctor_get(x_31, 1); +lean_inc(x_36); +lean_dec(x_31); +x_37 = lean_ctor_get(x_32, 0); +lean_inc(x_37); +lean_dec(x_32); +x_38 = l_List_appendTR___rarg(x_37, x_3); +x_2 = x_30; +x_3 = x_38; +x_11 = x_36; +goto _start; +} +} +else +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +lean_dec(x_30); +lean_dec(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_1); +x_40 = lean_ctor_get(x_31, 0); +lean_inc(x_40); +x_41 = lean_ctor_get(x_31, 1); +lean_inc(x_41); +if (lean_is_exclusive(x_31)) { + lean_ctor_release(x_31, 0); + lean_ctor_release(x_31, 1); + x_42 = x_31; +} else { + lean_dec_ref(x_31); + x_42 = lean_box(0); +} +if (lean_is_scalar(x_42)) { + x_43 = lean_alloc_ctor(1, 2, 0); +} else { + x_43 = x_42; +} +lean_ctor_set(x_43, 0, x_40); +lean_ctor_set(x_43, 1, x_41); +return x_43; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_applyToAll(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; lean_object* x_12; +x_11 = lean_box(0); +x_12 = l_Lean_Meta_Grind_applyToAll_go(x_1, x_2, 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_Meta_Grind_GrindTactic_andThen(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, 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_9(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_12) == 0) +{ +lean_object* x_13; +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +if (lean_obj_tag(x_13) == 0) +{ +uint8_t x_14; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +x_14 = !lean_is_exclusive(x_12); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; +x_15 = lean_ctor_get(x_12, 0); +lean_dec(x_15); +x_16 = lean_box(0); +lean_ctor_set(x_12, 0, x_16); +return x_12; +} +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_17 = lean_ctor_get(x_12, 1); +lean_inc(x_17); +lean_dec(x_12); +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_20; uint8_t x_21; +x_20 = lean_ctor_get(x_12, 1); +lean_inc(x_20); +lean_dec(x_12); +x_21 = !lean_is_exclusive(x_13); +if (x_21 == 0) +{ +lean_object* x_22; lean_object* x_23; +x_22 = lean_ctor_get(x_13, 0); +x_23 = l_Lean_Meta_Grind_applyToAll(x_2, x_22, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_20); +if (lean_obj_tag(x_23) == 0) +{ +uint8_t x_24; +x_24 = !lean_is_exclusive(x_23); +if (x_24 == 0) +{ +lean_object* x_25; +x_25 = lean_ctor_get(x_23, 0); +lean_ctor_set(x_13, 0, x_25); +lean_ctor_set(x_23, 0, x_13); +return x_23; +} +else +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_26 = lean_ctor_get(x_23, 0); +x_27 = lean_ctor_get(x_23, 1); +lean_inc(x_27); +lean_inc(x_26); +lean_dec(x_23); +lean_ctor_set(x_13, 0, x_26); +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_13); +lean_ctor_set(x_28, 1, x_27); +return x_28; +} +} +else +{ +uint8_t x_29; +lean_free_object(x_13); +x_29 = !lean_is_exclusive(x_23); +if (x_29 == 0) +{ +return x_23; +} +else +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_30 = lean_ctor_get(x_23, 0); +x_31 = lean_ctor_get(x_23, 1); +lean_inc(x_31); +lean_inc(x_30); +lean_dec(x_23); +x_32 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_32, 0, x_30); +lean_ctor_set(x_32, 1, x_31); +return x_32; +} +} +} +else +{ +lean_object* x_33; lean_object* x_34; +x_33 = lean_ctor_get(x_13, 0); +lean_inc(x_33); +lean_dec(x_13); +x_34 = l_Lean_Meta_Grind_applyToAll(x_2, x_33, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_20); +if (lean_obj_tag(x_34) == 0) +{ +lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; +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); +} +x_38 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_38, 0, x_35); +if (lean_is_scalar(x_37)) { + x_39 = lean_alloc_ctor(0, 2, 0); +} else { + x_39 = x_37; +} +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_36); +return x_39; +} +else +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_40 = lean_ctor_get(x_34, 0); +lean_inc(x_40); +x_41 = lean_ctor_get(x_34, 1); +lean_inc(x_41); +if (lean_is_exclusive(x_34)) { + lean_ctor_release(x_34, 0); + lean_ctor_release(x_34, 1); + x_42 = x_34; +} else { + lean_dec_ref(x_34); + x_42 = lean_box(0); +} +if (lean_is_scalar(x_42)) { + x_43 = lean_alloc_ctor(1, 2, 0); +} else { + x_43 = x_42; +} +lean_ctor_set(x_43, 0, x_40); +lean_ctor_set(x_43, 1, x_41); +return x_43; +} +} +} +} +else +{ +uint8_t x_44; +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); +x_44 = !lean_is_exclusive(x_12); +if (x_44 == 0) +{ +return x_12; +} +else +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_45 = lean_ctor_get(x_12, 0); +x_46 = lean_ctor_get(x_12, 1); +lean_inc(x_46); +lean_inc(x_45); +lean_dec(x_12); +x_47 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_47, 0, x_45); +lean_ctor_set(x_47, 1, x_46); +return x_47; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instAndThenGrindTactic(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, 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 = lean_box(0); +x_13 = lean_apply_1(x_2, x_12); +x_14 = l_Lean_Meta_Grind_GrindTactic_andThen(x_1, x_13, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GrindTactic_iterate_go(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_12; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_3); +lean_ctor_set(x_12, 1, x_11); +return x_12; +} +else +{ +uint8_t x_13; +x_13 = !lean_is_exclusive(x_2); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_14 = lean_ctor_get(x_2, 0); +x_15 = lean_ctor_get(x_2, 1); +lean_inc(x_1); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_14); +x_16 = lean_apply_9(x_1, x_14, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_16) == 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_18; +x_18 = lean_ctor_get(x_16, 1); +lean_inc(x_18); +lean_dec(x_16); +lean_ctor_set(x_2, 1, x_3); +{ +lean_object* _tmp_1 = x_15; +lean_object* _tmp_2 = x_2; +lean_object* _tmp_10 = x_18; +x_2 = _tmp_1; +x_3 = _tmp_2; +x_11 = _tmp_10; +} +goto _start; +} +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; +lean_free_object(x_2); +lean_dec(x_14); +x_20 = lean_ctor_get(x_16, 1); +lean_inc(x_20); +lean_dec(x_16); +x_21 = lean_ctor_get(x_17, 0); +lean_inc(x_21); +lean_dec(x_17); +x_22 = l_List_appendTR___rarg(x_21, x_15); +x_2 = x_22; +x_11 = x_20; +goto _start; +} +} +else +{ +uint8_t x_24; +lean_free_object(x_2); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_24 = !lean_is_exclusive(x_16); +if (x_24 == 0) +{ +return x_16; +} +else +{ +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_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; +} +} +} +else +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_28 = lean_ctor_get(x_2, 0); +x_29 = lean_ctor_get(x_2, 1); +lean_inc(x_29); +lean_inc(x_28); +lean_dec(x_2); +lean_inc(x_1); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_28); +x_30 = lean_apply_9(x_1, x_28, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_30) == 0) +{ +lean_object* x_31; +x_31 = lean_ctor_get(x_30, 0); +lean_inc(x_31); +if (lean_obj_tag(x_31) == 0) +{ +lean_object* x_32; lean_object* x_33; +x_32 = lean_ctor_get(x_30, 1); +lean_inc(x_32); +lean_dec(x_30); +x_33 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_33, 0, x_28); +lean_ctor_set(x_33, 1, x_3); +x_2 = x_29; +x_3 = x_33; +x_11 = x_32; +goto _start; +} +else +{ +lean_object* x_35; lean_object* x_36; lean_object* x_37; +lean_dec(x_28); +x_35 = lean_ctor_get(x_30, 1); +lean_inc(x_35); +lean_dec(x_30); +x_36 = lean_ctor_get(x_31, 0); +lean_inc(x_36); +lean_dec(x_31); +x_37 = l_List_appendTR___rarg(x_36, x_29); +x_2 = x_37; +x_11 = x_35; +goto _start; +} +} +else +{ +lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; +lean_dec(x_29); +lean_dec(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_3); +lean_dec(x_1); +x_39 = lean_ctor_get(x_30, 0); +lean_inc(x_39); +x_40 = lean_ctor_get(x_30, 1); +lean_inc(x_40); +if (lean_is_exclusive(x_30)) { + lean_ctor_release(x_30, 0); + lean_ctor_release(x_30, 1); + x_41 = x_30; +} else { + lean_dec_ref(x_30); + x_41 = lean_box(0); +} +if (lean_is_scalar(x_41)) { + x_42 = lean_alloc_ctor(1, 2, 0); +} else { + x_42 = x_41; +} +lean_ctor_set(x_42, 0, x_39); +lean_ctor_set(x_42, 1, x_40); +return x_42; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GrindTactic_iterate(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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_box(0); +x_12 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_12, 0, x_2); +lean_ctor_set(x_12, 1, x_11); +x_13 = l_Lean_Meta_Grind_GrindTactic_iterate_go(x_1, x_12, x_11, x_3, x_4, x_5, 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) +{ +lean_object* x_15; lean_object* x_16; +x_15 = lean_ctor_get(x_13, 0); +x_16 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_16, 0, x_15); +lean_ctor_set(x_13, 0, x_16); +return x_13; +} +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_17 = lean_ctor_get(x_13, 0); +x_18 = lean_ctor_get(x_13, 1); +lean_inc(x_18); +lean_inc(x_17); +lean_dec(x_13); +x_19 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_19, 0, x_17); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_19); +lean_ctor_set(x_20, 1, x_18); +return x_20; +} +} +else +{ +uint8_t x_21; +x_21 = !lean_is_exclusive(x_13); +if (x_21 == 0) +{ +return x_13; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_22 = lean_ctor_get(x_13, 0); +x_23 = lean_ctor_get(x_13, 1); +lean_inc(x_23); +lean_inc(x_22); +lean_dec(x_13); +x_24 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_24, 0, x_22); +lean_ctor_set(x_24, 1, x_23); +return x_24; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GrindTactic_orElse(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_12 = lean_apply_9(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_12) == 0) +{ +lean_object* x_13; +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; +x_14 = lean_ctor_get(x_12, 1); +lean_inc(x_14); +lean_dec(x_12); +x_15 = lean_apply_9(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_14); +return x_15; +} +else +{ +uint8_t x_16; +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_16 = !lean_is_exclusive(x_12); +if (x_16 == 0) +{ +lean_object* x_17; uint8_t x_18; +x_17 = lean_ctor_get(x_12, 0); +lean_dec(x_17); +x_18 = !lean_is_exclusive(x_13); +if (x_18 == 0) +{ +return x_12; +} +else +{ +lean_object* x_19; lean_object* x_20; +x_19 = lean_ctor_get(x_13, 0); +lean_inc(x_19); +lean_dec(x_13); +x_20 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_20, 0, x_19); +lean_ctor_set(x_12, 0, x_20); +return x_12; +} +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_21 = lean_ctor_get(x_12, 1); +lean_inc(x_21); +lean_dec(x_12); +x_22 = lean_ctor_get(x_13, 0); +lean_inc(x_22); +if (lean_is_exclusive(x_13)) { + lean_ctor_release(x_13, 0); + x_23 = x_13; +} else { + lean_dec_ref(x_13); + x_23 = lean_box(0); +} +if (lean_is_scalar(x_23)) { + x_24 = lean_alloc_ctor(1, 1, 0); +} else { + x_24 = x_23; +} +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_21); +return x_25; +} +} +} +else +{ +uint8_t 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_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_26 = !lean_is_exclusive(x_12); +if (x_26 == 0) +{ +return x_12; +} +else +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_12, 0); +x_28 = lean_ctor_get(x_12, 1); +lean_inc(x_28); +lean_inc(x_27); +lean_dec(x_12); +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_Grind_instOrElseGrindTactic(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, 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 = lean_box(0); +x_13 = lean_apply_1(x_2, x_12); +x_14 = l_Lean_Meta_Grind_GrindTactic_andThen(x_1, x_13, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_toGrindTactic___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; uint8_t x_11; +x_10 = lean_st_mk_ref(x_1, x_9); +x_11 = !lean_is_exclusive(x_10); +if (x_11 == 0) +{ +return x_10; +} +else +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_12 = lean_ctor_get(x_10, 0); +x_13 = lean_ctor_get(x_10, 1); +lean_inc(x_13); +lean_inc(x_12); +lean_dec(x_10); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_12); +lean_ctor_set(x_14, 1, x_13); +return x_14; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_toGrindTactic___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +lean_inc(x_2); +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) +{ +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 = lean_st_ref_get(x_2, x_12); +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_15 = lean_ctor_get(x_13, 1); +x_16 = lean_st_ref_get(x_2, x_15); +lean_dec(x_2); +x_17 = !lean_is_exclusive(x_16); +if (x_17 == 0) +{ +lean_object* x_18; +x_18 = lean_ctor_get(x_16, 0); +lean_ctor_set(x_13, 1, x_18); +lean_ctor_set(x_16, 0, x_13); +return x_16; +} +else +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_16, 0); +x_20 = lean_ctor_get(x_16, 1); +lean_inc(x_20); +lean_inc(x_19); +lean_dec(x_16); +lean_ctor_set(x_13, 1, x_19); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_13); +lean_ctor_set(x_21, 1, x_20); +return x_21; +} +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_22 = lean_ctor_get(x_13, 0); +x_23 = lean_ctor_get(x_13, 1); +lean_inc(x_23); +lean_inc(x_22); +lean_dec(x_13); +x_24 = lean_st_ref_get(x_2, x_23); +lean_dec(x_2); +x_25 = lean_ctor_get(x_24, 0); +lean_inc(x_25); +x_26 = lean_ctor_get(x_24, 1); +lean_inc(x_26); +if (lean_is_exclusive(x_24)) { + lean_ctor_release(x_24, 0); + lean_ctor_release(x_24, 1); + x_27 = x_24; +} else { + lean_dec_ref(x_24); + x_27 = lean_box(0); +} +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_22); +lean_ctor_set(x_28, 1, x_25); +if (lean_is_scalar(x_27)) { + x_29 = lean_alloc_ctor(0, 2, 0); +} else { + x_29 = x_27; +} +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_26); +return x_29; +} +} +else +{ +uint8_t x_30; +lean_dec(x_2); +x_30 = !lean_is_exclusive(x_11); +if (x_30 == 0) +{ +return x_11; +} +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_11, 0); +x_32 = lean_ctor_get(x_11, 1); +lean_inc(x_32); +lean_inc(x_31); +lean_dec(x_11); +x_33 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_33, 0, x_31); +lean_ctor_set(x_33, 1, x_32); +return x_33; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_toGrindTactic___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; +x_10 = !lean_is_exclusive(x_1); +if (x_10 == 0) +{ +lean_object* x_11; +x_11 = lean_ctor_get(x_1, 1); +lean_dec(x_11); +lean_ctor_set(x_1, 1, x_9); +return x_1; +} +else +{ +lean_object* x_12; lean_object* x_13; +x_12 = lean_ctor_get(x_1, 0); +lean_inc(x_12); +lean_dec(x_1); +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_9); +return x_13; +} +} +} +static lean_object* _init_l_Lean_Meta_Grind_toGrindTactic___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_toGrindTactic___lambda__3___boxed), 9, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_toGrindTactic___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(0); +x_2 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_toGrindTactic(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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; +x_11 = lean_ctor_get(x_2, 0); +lean_inc(x_11); +x_12 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_toGrindTactic___lambda__1___boxed), 9, 1); +lean_closure_set(x_12, 0, x_2); +x_13 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_toGrindTactic___lambda__2), 10, 1); +lean_closure_set(x_13, 0, x_1); +x_14 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); +lean_closure_set(x_14, 0, x_12); +lean_closure_set(x_14, 1, x_13); +x_15 = l_Lean_Meta_Grind_toGrindTactic___closed__1; +x_16 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); +lean_closure_set(x_16, 0, x_14); +lean_closure_set(x_16, 1, x_15); +x_17 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(x_11, x_16, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_17) == 0) +{ +lean_object* x_18; uint8_t x_19; +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +x_19 = lean_ctor_get_uint8(x_18, sizeof(void*)*20); +if (x_19 == 0) +{ +uint8_t x_20; +x_20 = !lean_is_exclusive(x_17); +if (x_20 == 0) +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_21 = lean_ctor_get(x_17, 0); +lean_dec(x_21); +x_22 = lean_box(0); +x_23 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_23, 0, x_18); +lean_ctor_set(x_23, 1, x_22); +x_24 = lean_alloc_ctor(1, 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_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_25 = lean_ctor_get(x_17, 1); +lean_inc(x_25); +lean_dec(x_17); +x_26 = lean_box(0); +x_27 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_27, 0, x_18); +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_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_25); +return x_29; +} +} +else +{ +uint8_t x_30; +lean_dec(x_18); +x_30 = !lean_is_exclusive(x_17); +if (x_30 == 0) +{ +lean_object* x_31; lean_object* x_32; +x_31 = lean_ctor_get(x_17, 0); +lean_dec(x_31); +x_32 = l_Lean_Meta_Grind_toGrindTactic___closed__2; +lean_ctor_set(x_17, 0, x_32); +return x_17; +} +else +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_33 = lean_ctor_get(x_17, 1); +lean_inc(x_33); +lean_dec(x_17); +x_34 = l_Lean_Meta_Grind_toGrindTactic___closed__2; +x_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_34); +lean_ctor_set(x_35, 1, x_33); +return x_35; +} +} +} +else +{ +uint8_t x_36; +x_36 = !lean_is_exclusive(x_17); +if (x_36 == 0) +{ +return x_17; +} +else +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_37 = lean_ctor_get(x_17, 0); +x_38 = lean_ctor_get(x_17, 1); +lean_inc(x_38); +lean_inc(x_37); +lean_dec(x_17); +x_39 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_39, 0, x_37); +lean_ctor_set(x_39, 1, x_38); +return x_39; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_toGrindTactic___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l_Lean_Meta_Grind_toGrindTactic___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_toGrindTactic___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l_Lean_Meta_Grind_toGrindTactic___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GrindTactic_x27_toGrindTactic(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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) +{ +return x_11; +} +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_11, 0); +x_21 = lean_ctor_get(x_11, 1); +lean_inc(x_21); +lean_inc(x_20); +lean_dec(x_11); +x_22 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_22, 0, x_20); +lean_ctor_set(x_22, 1, x_21); +return x_22; +} +} +} +} +lean_object* initialize_Lean_Meta_Tactic_Grind_Types(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Lean_Meta_Tactic_Grind_Combinators(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_Tactic_Grind_Types(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l_Lean_Meta_Grind_toGrindTactic___closed__1 = _init_l_Lean_Meta_Grind_toGrindTactic___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_toGrindTactic___closed__1); +l_Lean_Meta_Grind_toGrindTactic___closed__2 = _init_l_Lean_Meta_Grind_toGrindTactic___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_toGrindTactic___closed__2); +return lean_io_result_mk_ok(lean_box(0)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Lean/Meta/Tactic/Grind/Core.c b/stage0/stdlib/Lean/Meta/Tactic/Grind/Core.c index 80d6a7357d90..a061c68ec77f 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Grind/Core.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Grind/Core.c @@ -70,6 +70,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Gr LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addEq(lean_object*, lean_object*, lean_object*, 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_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_updateMT___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_closeGoalWithTrueEqFalse___closed__4; +lean_object* l_Lean_Meta_Grind_isCongrRoot(lean_object*, 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_Grind_add_goEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Level_ofNat(lean_object*); lean_object* l_Lean_Expr_appArg_x21(lean_object*); @@ -102,6 +103,7 @@ LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tact LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_eraseAux___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__2(lean_object*, lean_object*, size_t, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_closeGoalWithValuesEq___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqCore_processTodo(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_Grind_addNewEq(lean_object*, lean_object*, lean_object*, lean_object*, 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_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_closeGoal(lean_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_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, 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*); @@ -120,7 +122,6 @@ lean_object* l_Lean_Meta_mkHEq(lean_object*, lean_object*, lean_object*, lean_ob 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_Meta_Grind_add___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* l_Lean_Expr_appFnCleanup(lean_object*, lean_object*); -lean_object* l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___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* l_Lean_Meta_Grind_addCongrTable(lean_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_Grind_Core_0__Lean_Meta_Grind_resetNewEqs(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_Grind_Core_0__Lean_Meta_Grind_closeGoalWithValuesEq___closed__15; @@ -146,6 +147,7 @@ lean_object* l_Lean_Meta_Grind_getParents(lean_object*, lean_object*, lean_objec static lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_resetNewEqs___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqCore(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___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqCore_processTodo___closed__1; +lean_object* l_Lean_Meta_Grind_updateLastTag(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_Grind_Core_0__Lean_Meta_Grind_closeGoalWithValuesEq___closed__10; lean_object* l_Array_eraseIdx___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; @@ -155,6 +157,7 @@ static lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_a static lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_reinsertParents___spec__1___closed__2; lean_object* l_Lean_MessageData_ofExpr(lean_object*); lean_object* l_Lean_Meta_Grind_copyParentsTo(lean_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_Grind_updateLastTag___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_Grind_addHypothesis(lean_object*, lean_object*, 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_Grind_add___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___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqCore___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -177,6 +180,7 @@ uint8_t lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__2(lean_object*, 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*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_closeGoalWithTrueEqFalse(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_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__4; +lean_object* l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___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___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_closeGoalWithValuesEq___closed__6; lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_reinsertParents(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -192,7 +196,6 @@ static lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_a LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_eraseAux___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_invertTrans_go___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_add_go___closed__2; -lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_getENode(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__4; lean_object* l_Lean_Meta_instantiateMVarsIfMVarApp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1011,6 +1014,7 @@ LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_erase___at___private_Lean_Meta lean_object* x_4; uint64_t x_5; size_t x_6; lean_object* x_7; x_4 = lean_ctor_get(x_1, 1); lean_inc(x_4); +lean_inc(x_3); x_5 = l_Lean_Meta_Grind_congrHash(x_4, x_3); x_6 = lean_uint64_to_usize(x_5); x_7 = l_Lean_PersistentHashMap_eraseAux___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__2(x_1, x_2, x_6, x_3); @@ -1030,7 +1034,7 @@ return x_2; LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; 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_12; lean_object* x_13; lean_object* x_14; 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; uint8_t x_37; x_12 = lean_st_ref_take(x_3, x_11); x_13 = lean_ctor_get(x_12, 0); lean_inc(x_13); @@ -1049,7 +1053,7 @@ x_19 = lean_ctor_get(x_13, 4); lean_inc(x_19); x_20 = lean_ctor_get(x_13, 5); lean_inc(x_20); -x_21 = lean_ctor_get_uint8(x_13, sizeof(void*)*14); +x_21 = lean_ctor_get_uint8(x_13, sizeof(void*)*20); x_22 = lean_ctor_get(x_13, 6); lean_inc(x_22); x_23 = lean_ctor_get(x_13, 7); @@ -1066,105 +1070,135 @@ x_28 = lean_ctor_get(x_13, 12); lean_inc(x_28); x_29 = lean_ctor_get(x_13, 13); lean_inc(x_29); +x_30 = lean_ctor_get(x_13, 14); +lean_inc(x_30); +x_31 = lean_ctor_get(x_13, 15); +lean_inc(x_31); +x_32 = lean_ctor_get(x_13, 16); +lean_inc(x_32); +x_33 = lean_ctor_get(x_13, 17); +lean_inc(x_33); +x_34 = lean_ctor_get(x_13, 18); +lean_inc(x_34); +x_35 = lean_ctor_get(x_13, 19); +lean_inc(x_35); lean_inc(x_13); -x_30 = l_Lean_PersistentHashMap_erase___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__1(x_13, x_18, x_1); -x_31 = !lean_is_exclusive(x_13); -if (x_31 == 0) +x_36 = l_Lean_PersistentHashMap_erase___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__1(x_13, x_18, x_1); +x_37 = !lean_is_exclusive(x_13); +if (x_37 == 0) { -lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; uint8_t x_47; -x_32 = lean_ctor_get(x_13, 13); -lean_dec(x_32); -x_33 = lean_ctor_get(x_13, 12); -lean_dec(x_33); -x_34 = lean_ctor_get(x_13, 11); -lean_dec(x_34); -x_35 = lean_ctor_get(x_13, 10); -lean_dec(x_35); -x_36 = lean_ctor_get(x_13, 9); -lean_dec(x_36); -x_37 = lean_ctor_get(x_13, 8); -lean_dec(x_37); -x_38 = lean_ctor_get(x_13, 7); +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; uint8_t x_59; +x_38 = lean_ctor_get(x_13, 19); lean_dec(x_38); -x_39 = lean_ctor_get(x_13, 6); +x_39 = lean_ctor_get(x_13, 18); lean_dec(x_39); -x_40 = lean_ctor_get(x_13, 5); +x_40 = lean_ctor_get(x_13, 17); lean_dec(x_40); -x_41 = lean_ctor_get(x_13, 4); +x_41 = lean_ctor_get(x_13, 16); lean_dec(x_41); -x_42 = lean_ctor_get(x_13, 3); +x_42 = lean_ctor_get(x_13, 15); lean_dec(x_42); -x_43 = lean_ctor_get(x_13, 2); +x_43 = lean_ctor_get(x_13, 14); lean_dec(x_43); -x_44 = lean_ctor_get(x_13, 1); +x_44 = lean_ctor_get(x_13, 13); lean_dec(x_44); -x_45 = lean_ctor_get(x_13, 0); +x_45 = lean_ctor_get(x_13, 12); lean_dec(x_45); -lean_ctor_set(x_13, 3, x_30); -x_46 = lean_st_ref_set(x_3, x_13, x_14); -x_47 = !lean_is_exclusive(x_46); -if (x_47 == 0) -{ -lean_object* x_48; lean_object* x_49; -x_48 = lean_ctor_get(x_46, 0); +x_46 = lean_ctor_get(x_13, 11); +lean_dec(x_46); +x_47 = lean_ctor_get(x_13, 10); +lean_dec(x_47); +x_48 = lean_ctor_get(x_13, 9); lean_dec(x_48); -x_49 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___lambda__1___closed__1; -lean_ctor_set(x_46, 0, x_49); -return x_46; +x_49 = lean_ctor_get(x_13, 8); +lean_dec(x_49); +x_50 = lean_ctor_get(x_13, 7); +lean_dec(x_50); +x_51 = lean_ctor_get(x_13, 6); +lean_dec(x_51); +x_52 = lean_ctor_get(x_13, 5); +lean_dec(x_52); +x_53 = lean_ctor_get(x_13, 4); +lean_dec(x_53); +x_54 = lean_ctor_get(x_13, 3); +lean_dec(x_54); +x_55 = lean_ctor_get(x_13, 2); +lean_dec(x_55); +x_56 = lean_ctor_get(x_13, 1); +lean_dec(x_56); +x_57 = lean_ctor_get(x_13, 0); +lean_dec(x_57); +lean_ctor_set(x_13, 3, x_36); +x_58 = lean_st_ref_set(x_3, x_13, x_14); +x_59 = !lean_is_exclusive(x_58); +if (x_59 == 0) +{ +lean_object* x_60; lean_object* x_61; +x_60 = lean_ctor_get(x_58, 0); +lean_dec(x_60); +x_61 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___lambda__1___closed__1; +lean_ctor_set(x_58, 0, x_61); +return x_58; } else { -lean_object* x_50; lean_object* x_51; lean_object* x_52; -x_50 = lean_ctor_get(x_46, 1); -lean_inc(x_50); -lean_dec(x_46); -x_51 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___lambda__1___closed__1; -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_62; lean_object* x_63; lean_object* x_64; +x_62 = lean_ctor_get(x_58, 1); +lean_inc(x_62); +lean_dec(x_58); +x_63 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___lambda__1___closed__1; +x_64 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_64, 0, x_63); +lean_ctor_set(x_64, 1, x_62); +return x_64; } } else { -lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; +lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_dec(x_13); -x_53 = lean_alloc_ctor(0, 14, 1); -lean_ctor_set(x_53, 0, x_15); -lean_ctor_set(x_53, 1, x_16); -lean_ctor_set(x_53, 2, x_17); -lean_ctor_set(x_53, 3, x_30); -lean_ctor_set(x_53, 4, x_19); -lean_ctor_set(x_53, 5, x_20); -lean_ctor_set(x_53, 6, x_22); -lean_ctor_set(x_53, 7, x_23); -lean_ctor_set(x_53, 8, x_24); -lean_ctor_set(x_53, 9, x_25); -lean_ctor_set(x_53, 10, x_26); -lean_ctor_set(x_53, 11, x_27); -lean_ctor_set(x_53, 12, x_28); -lean_ctor_set(x_53, 13, x_29); -lean_ctor_set_uint8(x_53, sizeof(void*)*14, x_21); -x_54 = lean_st_ref_set(x_3, x_53, x_14); -x_55 = lean_ctor_get(x_54, 1); -lean_inc(x_55); -if (lean_is_exclusive(x_54)) { - lean_ctor_release(x_54, 0); - lean_ctor_release(x_54, 1); - x_56 = x_54; +x_65 = lean_alloc_ctor(0, 20, 1); +lean_ctor_set(x_65, 0, x_15); +lean_ctor_set(x_65, 1, x_16); +lean_ctor_set(x_65, 2, x_17); +lean_ctor_set(x_65, 3, x_36); +lean_ctor_set(x_65, 4, x_19); +lean_ctor_set(x_65, 5, x_20); +lean_ctor_set(x_65, 6, x_22); +lean_ctor_set(x_65, 7, x_23); +lean_ctor_set(x_65, 8, x_24); +lean_ctor_set(x_65, 9, x_25); +lean_ctor_set(x_65, 10, x_26); +lean_ctor_set(x_65, 11, x_27); +lean_ctor_set(x_65, 12, x_28); +lean_ctor_set(x_65, 13, x_29); +lean_ctor_set(x_65, 14, x_30); +lean_ctor_set(x_65, 15, x_31); +lean_ctor_set(x_65, 16, x_32); +lean_ctor_set(x_65, 17, x_33); +lean_ctor_set(x_65, 18, x_34); +lean_ctor_set(x_65, 19, x_35); +lean_ctor_set_uint8(x_65, sizeof(void*)*20, x_21); +x_66 = lean_st_ref_set(x_3, x_65, x_14); +x_67 = lean_ctor_get(x_66, 1); +lean_inc(x_67); +if (lean_is_exclusive(x_66)) { + lean_ctor_release(x_66, 0); + lean_ctor_release(x_66, 1); + x_68 = x_66; } else { - lean_dec_ref(x_54); - x_56 = lean_box(0); + lean_dec_ref(x_66); + x_68 = lean_box(0); } -x_57 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___lambda__1___closed__1; -if (lean_is_scalar(x_56)) { - x_58 = lean_alloc_ctor(0, 2, 0); +x_69 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___lambda__1___closed__1; +if (lean_is_scalar(x_68)) { + x_70 = lean_alloc_ctor(0, 2, 0); } else { - x_58 = x_56; + x_70 = x_68; } -lean_ctor_set(x_58, 0, x_57); -lean_ctor_set(x_58, 1, x_55); -return x_58; +lean_ctor_set(x_70, 0, x_69); +lean_ctor_set(x_70, 1, x_67); +return x_70; } } } @@ -1252,7 +1286,7 @@ return x_13; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; x_14 = lean_ctor_get(x_1, 0); lean_inc(x_14); x_15 = lean_ctor_get(x_1, 1); @@ -1261,202 +1295,308 @@ x_16 = lean_ctor_get(x_1, 3); lean_inc(x_16); lean_dec(x_1); x_17 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4(x_14, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -x_18 = !lean_is_exclusive(x_17); -if (x_18 == 0) +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; uint8_t x_24; -x_19 = lean_ctor_get(x_17, 1); +lean_object* x_18; +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +if (lean_obj_tag(x_18) == 0) +{ +uint8_t x_19; +lean_dec(x_16); +lean_dec(x_15); +x_19 = !lean_is_exclusive(x_17); +if (x_19 == 0) +{ +lean_object* x_20; x_20 = lean_ctor_get(x_17, 0); lean_dec(x_20); -x_21 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__4; -x_22 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_21, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_19); -x_23 = lean_ctor_get(x_22, 0); +return x_17; +} +else +{ +lean_object* x_21; lean_object* x_22; +x_21 = lean_ctor_get(x_17, 1); +lean_inc(x_21); +lean_dec(x_17); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_18); +lean_ctor_set(x_22, 1, x_21); +return x_22; +} +} +else +{ +lean_object* x_23; uint8_t x_24; lean_object* x_25; uint8_t x_81; +lean_dec(x_18); +x_23 = lean_ctor_get(x_17, 1); lean_inc(x_23); -x_24 = lean_unbox(x_23); -lean_dec(x_23); +lean_dec(x_17); +x_81 = l_Lean_Expr_isApp(x_15); +if (x_81 == 0) +{ +x_24 = x_81; +x_25 = x_23; +goto block_80; +} +else +{ +lean_object* x_82; +lean_inc(x_15); +x_82 = l_Lean_Meta_Grind_isCongrRoot(x_15, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_23); +if (lean_obj_tag(x_82) == 0) +{ +lean_object* x_83; lean_object* x_84; uint8_t x_85; +x_83 = lean_ctor_get(x_82, 0); +lean_inc(x_83); +x_84 = lean_ctor_get(x_82, 1); +lean_inc(x_84); +lean_dec(x_82); +x_85 = lean_unbox(x_83); +lean_dec(x_83); +x_24 = x_85; +x_25 = x_84; +goto block_80; +} +else +{ +uint8_t x_86; +lean_dec(x_16); +lean_dec(x_15); +x_86 = !lean_is_exclusive(x_82); +if (x_86 == 0) +{ +return x_82; +} +else +{ +lean_object* x_87; lean_object* x_88; lean_object* x_89; +x_87 = lean_ctor_get(x_82, 0); +x_88 = lean_ctor_get(x_82, 1); +lean_inc(x_88); +lean_inc(x_87); +lean_dec(x_82); +x_89 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_89, 0, x_87); +lean_ctor_set(x_89, 1, x_88); +return x_89; +} +} +} +block_80: +{ 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_free_object(x_17); -x_25 = lean_ctor_get(x_22, 1); -lean_inc(x_25); -lean_dec(x_22); +lean_object* x_26; +lean_dec(x_15); x_26 = lean_box(0); -x_27 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___lambda__1(x_15, x_26, x_3, x_4, 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_ctor_get(x_27, 1); -lean_inc(x_29); -lean_dec(x_27); -x_30 = lean_ctor_get(x_28, 0); -lean_inc(x_30); -lean_dec(x_28); x_1 = x_16; -x_2 = x_30; -x_11 = x_29; +x_2 = x_26; +x_11 = x_25; goto _start; } else { -uint8_t x_32; -x_32 = !lean_is_exclusive(x_22); -if (x_32 == 0) +lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; +x_28 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__4; +x_29 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_28, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_25); +x_30 = lean_ctor_get(x_29, 0); +lean_inc(x_30); +x_31 = lean_unbox(x_30); +lean_dec(x_30); +if (x_31 == 0) { -lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; 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_22, 1); -x_34 = lean_ctor_get(x_22, 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_29, 1); +lean_inc(x_32); +lean_dec(x_29); +x_33 = lean_box(0); +x_34 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___lambda__1(x_15, x_33, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_32); +x_35 = lean_ctor_get(x_34, 0); +lean_inc(x_35); +x_36 = lean_ctor_get(x_34, 1); +lean_inc(x_36); lean_dec(x_34); -lean_inc(x_15); -x_35 = l_Lean_MessageData_ofExpr(x_15); -x_36 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__6; -lean_ctor_set_tag(x_22, 7); -lean_ctor_set(x_22, 1, x_35); -lean_ctor_set(x_22, 0, x_36); -x_37 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; -lean_ctor_set_tag(x_17, 7); -lean_ctor_set(x_17, 1, x_37); -lean_ctor_set(x_17, 0, x_22); -x_38 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_21, x_17, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, 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); -x_41 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___lambda__1(x_15, x_39, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_40); -lean_dec(x_39); -x_42 = lean_ctor_get(x_41, 0); -lean_inc(x_42); -x_43 = lean_ctor_get(x_41, 1); -lean_inc(x_43); -lean_dec(x_41); -x_44 = lean_ctor_get(x_42, 0); -lean_inc(x_44); -lean_dec(x_42); +x_37 = lean_ctor_get(x_35, 0); +lean_inc(x_37); +lean_dec(x_35); x_1 = x_16; -x_2 = x_44; -x_11 = x_43; +x_2 = x_37; +x_11 = x_36; goto _start; } else { -lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; -x_46 = lean_ctor_get(x_22, 1); -lean_inc(x_46); -lean_dec(x_22); +uint8_t x_39; +x_39 = !lean_is_exclusive(x_29); +if (x_39 == 0) +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_40 = lean_ctor_get(x_29, 1); +x_41 = lean_ctor_get(x_29, 0); +lean_dec(x_41); +x_42 = l_Lean_Meta_Grind_updateLastTag(x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_40); +if (lean_obj_tag(x_42) == 0) +{ +lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; +x_43 = lean_ctor_get(x_42, 1); +lean_inc(x_43); +lean_dec(x_42); lean_inc(x_15); -x_47 = l_Lean_MessageData_ofExpr(x_15); -x_48 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__6; -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_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; -lean_ctor_set_tag(x_17, 7); -lean_ctor_set(x_17, 1, x_50); -lean_ctor_set(x_17, 0, x_49); -x_51 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_21, x_17, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_46); +x_44 = l_Lean_MessageData_ofExpr(x_15); +x_45 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__6; +lean_ctor_set_tag(x_29, 7); +lean_ctor_set(x_29, 1, x_44); +lean_ctor_set(x_29, 0, x_45); +x_46 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; +x_47 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_47, 0, x_29); +lean_ctor_set(x_47, 1, x_46); +x_48 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_28, x_47, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_43); +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_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___lambda__1(x_15, x_49, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_50); +lean_dec(x_49); x_52 = lean_ctor_get(x_51, 0); lean_inc(x_52); x_53 = lean_ctor_get(x_51, 1); lean_inc(x_53); lean_dec(x_51); -x_54 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___lambda__1(x_15, x_52, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_53); +x_54 = lean_ctor_get(x_52, 0); +lean_inc(x_54); lean_dec(x_52); -x_55 = lean_ctor_get(x_54, 0); -lean_inc(x_55); -x_56 = lean_ctor_get(x_54, 1); -lean_inc(x_56); -lean_dec(x_54); -x_57 = lean_ctor_get(x_55, 0); -lean_inc(x_57); -lean_dec(x_55); x_1 = x_16; -x_2 = x_57; -x_11 = x_56; +x_2 = x_54; +x_11 = x_53; goto _start; } +else +{ +uint8_t x_56; +lean_free_object(x_29); +lean_dec(x_16); +lean_dec(x_15); +x_56 = !lean_is_exclusive(x_42); +if (x_56 == 0) +{ +return x_42; +} +else +{ +lean_object* x_57; lean_object* x_58; lean_object* x_59; +x_57 = lean_ctor_get(x_42, 0); +x_58 = lean_ctor_get(x_42, 1); +lean_inc(x_58); +lean_inc(x_57); +lean_dec(x_42); +x_59 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_59, 0, x_57); +lean_ctor_set(x_59, 1, x_58); +return x_59; +} } } else { -lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; uint8_t x_63; -x_59 = lean_ctor_get(x_17, 1); -lean_inc(x_59); -lean_dec(x_17); -x_60 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__4; -x_61 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_60, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_59); -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_60; lean_object* x_61; +x_60 = lean_ctor_get(x_29, 1); +lean_inc(x_60); +lean_dec(x_29); +x_61 = l_Lean_Meta_Grind_updateLastTag(x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_60); +if (lean_obj_tag(x_61) == 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; -x_64 = lean_ctor_get(x_61, 1); -lean_inc(x_64); +lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; +x_62 = lean_ctor_get(x_61, 1); +lean_inc(x_62); lean_dec(x_61); -x_65 = lean_box(0); -x_66 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___lambda__1(x_15, x_65, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_64); -x_67 = lean_ctor_get(x_66, 0); -lean_inc(x_67); -x_68 = lean_ctor_get(x_66, 1); -lean_inc(x_68); -lean_dec(x_66); -x_69 = lean_ctor_get(x_67, 0); +lean_inc(x_15); +x_63 = l_Lean_MessageData_ofExpr(x_15); +x_64 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__6; +x_65 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_65, 0, x_64); +lean_ctor_set(x_65, 1, x_63); +x_66 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; +x_67 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_67, 0, x_65); +lean_ctor_set(x_67, 1, x_66); +x_68 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_28, x_67, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_62); +x_69 = lean_ctor_get(x_68, 0); lean_inc(x_69); -lean_dec(x_67); +x_70 = lean_ctor_get(x_68, 1); +lean_inc(x_70); +lean_dec(x_68); +x_71 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___lambda__1(x_15, x_69, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_70); +lean_dec(x_69); +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); +lean_dec(x_72); x_1 = x_16; -x_2 = x_69; -x_11 = x_68; +x_2 = x_74; +x_11 = x_73; goto _start; } 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; 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_71 = lean_ctor_get(x_61, 1); -lean_inc(x_71); +lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; +lean_dec(x_16); +lean_dec(x_15); +x_76 = lean_ctor_get(x_61, 0); +lean_inc(x_76); +x_77 = lean_ctor_get(x_61, 1); +lean_inc(x_77); if (lean_is_exclusive(x_61)) { lean_ctor_release(x_61, 0); lean_ctor_release(x_61, 1); - x_72 = x_61; + x_78 = x_61; } else { lean_dec_ref(x_61); - x_72 = lean_box(0); + x_78 = lean_box(0); } -lean_inc(x_15); -x_73 = l_Lean_MessageData_ofExpr(x_15); -x_74 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__6; -if (lean_is_scalar(x_72)) { - x_75 = lean_alloc_ctor(7, 2, 0); +if (lean_is_scalar(x_78)) { + x_79 = lean_alloc_ctor(1, 2, 0); } else { - x_75 = x_72; - lean_ctor_set_tag(x_75, 7); + x_79 = x_78; } -lean_ctor_set(x_75, 0, x_74); -lean_ctor_set(x_75, 1, x_73); -x_76 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; -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_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_60, x_77, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_71); -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 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___lambda__1(x_15, x_79, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_80); -lean_dec(x_79); -x_82 = lean_ctor_get(x_81, 0); -lean_inc(x_82); -x_83 = lean_ctor_get(x_81, 1); -lean_inc(x_83); -lean_dec(x_81); -x_84 = lean_ctor_get(x_82, 0); -lean_inc(x_84); -lean_dec(x_82); -x_1 = x_16; -x_2 = x_84; -x_11 = x_83; -goto _start; +lean_ctor_set(x_79, 0, x_76); +lean_ctor_set(x_79, 1, x_77); +return x_79; +} +} +} +} +} +} +} +else +{ +uint8_t x_90; +lean_dec(x_16); +lean_dec(x_15); +x_90 = !lean_is_exclusive(x_17); +if (x_90 == 0) +{ +return x_17; +} +else +{ +lean_object* x_91; lean_object* x_92; lean_object* x_93; +x_91 = lean_ctor_get(x_17, 0); +x_92 = lean_ctor_get(x_17, 1); +lean_inc(x_92); +lean_inc(x_91); +lean_dec(x_17); +x_93 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_93, 0, x_91); +lean_ctor_set(x_93, 1, x_92); +return x_93; } } } @@ -1465,7 +1605,7 @@ goto _start; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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; +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; x_11 = l_Lean_Meta_Grind_getParentsAndReset(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); x_12 = lean_ctor_get(x_11, 0); lean_inc(x_12); @@ -1475,6 +1615,9 @@ lean_dec(x_11); x_14 = lean_box(0); lean_inc(x_12); x_15 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4(x_12, x_14, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_13); +if (lean_obj_tag(x_15) == 0) +{ +uint8_t x_16; x_16 = !lean_is_exclusive(x_15); if (x_16 == 0) { @@ -1496,6 +1639,30 @@ lean_ctor_set(x_19, 1, x_18); return x_19; } } +else +{ +uint8_t x_20; +lean_dec(x_12); +x_20 = !lean_is_exclusive(x_15); +if (x_20 == 0) +{ +return x_15; +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_15, 0); +x_22 = lean_ctor_get(x_15, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_15); +x_23 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_23, 0, x_21); +lean_ctor_set(x_23, 1, x_22); +return x_23; +} +} +} } LEAN_EXPORT lean_object* l_Array_indexOfAux___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: @@ -1718,89 +1885,42 @@ return x_22; } else { -lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; +lean_object* x_23; uint8_t x_24; lean_object* x_25; uint8_t x_104; lean_dec(x_18); x_23 = lean_ctor_get(x_17, 1); lean_inc(x_23); lean_dec(x_17); -x_24 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__4; -x_25 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_24, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_23); -x_26 = lean_ctor_get(x_25, 0); -lean_inc(x_26); -x_27 = lean_unbox(x_26); -lean_dec(x_26); -if (x_27 == 0) -{ -lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_28 = lean_ctor_get(x_25, 1); -lean_inc(x_28); -lean_dec(x_25); -x_29 = lean_box(0); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -x_30 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_reinsertParents___spec__1___lambda__1(x_15, x_29, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_28); -if (lean_obj_tag(x_30) == 0) -{ -lean_object* x_31; -x_31 = lean_ctor_get(x_30, 0); -lean_inc(x_31); -if (lean_obj_tag(x_31) == 0) -{ -uint8_t x_32; -lean_dec(x_16); -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_32 = !lean_is_exclusive(x_30); -if (x_32 == 0) +x_104 = l_Lean_Expr_isApp(x_15); +if (x_104 == 0) { -lean_object* x_33; -x_33 = lean_ctor_get(x_30, 0); -lean_dec(x_33); -return x_30; +x_24 = x_104; +x_25 = x_23; +goto block_103; } else { -lean_object* x_34; lean_object* x_35; -x_34 = lean_ctor_get(x_30, 1); -lean_inc(x_34); -lean_dec(x_30); -x_35 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_35, 0, x_31); -lean_ctor_set(x_35, 1, x_34); -return x_35; -} -} -else +lean_object* x_105; +lean_inc(x_15); +x_105 = l_Lean_Meta_Grind_isCongrRoot(x_15, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_23); +if (lean_obj_tag(x_105) == 0) { -lean_object* x_36; lean_object* x_37; -x_36 = lean_ctor_get(x_30, 1); -lean_inc(x_36); -lean_dec(x_30); -x_37 = lean_ctor_get(x_31, 0); -lean_inc(x_37); -lean_dec(x_31); -x_1 = x_16; -x_2 = x_37; -x_11 = x_36; -goto _start; -} +lean_object* x_106; lean_object* x_107; uint8_t x_108; +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_unbox(x_106); +lean_dec(x_106); +x_24 = x_108; +x_25 = x_107; +goto block_103; } else { -uint8_t x_39; +uint8_t x_109; lean_dec(x_16); +lean_dec(x_15); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -1809,52 +1929,54 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_39 = !lean_is_exclusive(x_30); -if (x_39 == 0) +x_109 = !lean_is_exclusive(x_105); +if (x_109 == 0) { -return x_30; +return x_105; } else { -lean_object* x_40; lean_object* x_41; lean_object* x_42; -x_40 = lean_ctor_get(x_30, 0); -x_41 = lean_ctor_get(x_30, 1); -lean_inc(x_41); -lean_inc(x_40); -lean_dec(x_30); -x_42 = lean_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_110; lean_object* x_111; lean_object* x_112; +x_110 = lean_ctor_get(x_105, 0); +x_111 = lean_ctor_get(x_105, 1); +lean_inc(x_111); +lean_inc(x_110); +lean_dec(x_105); +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; } } } +block_103: +{ +if (x_24 == 0) +{ +lean_object* x_26; +lean_dec(x_15); +x_26 = lean_box(0); +x_1 = x_16; +x_2 = x_26; +x_11 = x_25; +goto _start; +} else { -uint8_t x_43; -x_43 = !lean_is_exclusive(x_25); -if (x_43 == 0) +lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; +x_28 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__4; +x_29 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_28, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_25); +x_30 = lean_ctor_get(x_29, 0); +lean_inc(x_30); +x_31 = lean_unbox(x_30); +lean_dec(x_30); +if (x_31 == 0) { -lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; 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_25, 1); -x_45 = lean_ctor_get(x_25, 0); -lean_dec(x_45); -lean_inc(x_15); -x_46 = l_Lean_MessageData_ofExpr(x_15); -x_47 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_reinsertParents___spec__1___closed__2; -lean_ctor_set_tag(x_25, 7); -lean_ctor_set(x_25, 1, x_46); -lean_ctor_set(x_25, 0, x_47); -x_48 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; -x_49 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_49, 0, x_25); -lean_ctor_set(x_49, 1, x_48); -x_50 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_24, x_49, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_44); -x_51 = lean_ctor_get(x_50, 0); -lean_inc(x_51); -x_52 = lean_ctor_get(x_50, 1); -lean_inc(x_52); -lean_dec(x_50); +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_29, 1); +lean_inc(x_32); +lean_dec(x_29); +x_33 = lean_box(0); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); @@ -1863,16 +1985,15 @@ lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -x_53 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_reinsertParents___spec__1___lambda__1(x_15, x_51, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_52); -lean_dec(x_51); -if (lean_obj_tag(x_53) == 0) +x_34 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_reinsertParents___spec__1___lambda__1(x_15, x_33, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_32); +if (lean_obj_tag(x_34) == 0) { -lean_object* x_54; -x_54 = lean_ctor_get(x_53, 0); -lean_inc(x_54); -if (lean_obj_tag(x_54) == 0) +lean_object* x_35; +x_35 = lean_ctor_get(x_34, 0); +lean_inc(x_35); +if (lean_obj_tag(x_35) == 0) { -uint8_t x_55; +uint8_t x_36; lean_dec(x_16); lean_dec(x_10); lean_dec(x_9); @@ -1882,44 +2003,44 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_55 = !lean_is_exclusive(x_53); -if (x_55 == 0) +x_36 = !lean_is_exclusive(x_34); +if (x_36 == 0) { -lean_object* x_56; -x_56 = lean_ctor_get(x_53, 0); -lean_dec(x_56); -return x_53; +lean_object* x_37; +x_37 = lean_ctor_get(x_34, 0); +lean_dec(x_37); +return x_34; } else { -lean_object* x_57; lean_object* x_58; -x_57 = lean_ctor_get(x_53, 1); -lean_inc(x_57); -lean_dec(x_53); -x_58 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_58, 0, x_54); -lean_ctor_set(x_58, 1, x_57); -return x_58; +lean_object* x_38; lean_object* x_39; +x_38 = lean_ctor_get(x_34, 1); +lean_inc(x_38); +lean_dec(x_34); +x_39 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_39, 0, x_35); +lean_ctor_set(x_39, 1, x_38); +return x_39; } } else { -lean_object* x_59; lean_object* x_60; -x_59 = lean_ctor_get(x_53, 1); -lean_inc(x_59); -lean_dec(x_53); -x_60 = lean_ctor_get(x_54, 0); -lean_inc(x_60); -lean_dec(x_54); +lean_object* x_40; lean_object* x_41; +x_40 = lean_ctor_get(x_34, 1); +lean_inc(x_40); +lean_dec(x_34); +x_41 = lean_ctor_get(x_35, 0); +lean_inc(x_41); +lean_dec(x_35); x_1 = x_16; -x_2 = x_60; -x_11 = x_59; +x_2 = x_41; +x_11 = x_40; goto _start; } } else { -uint8_t x_62; +uint8_t x_43; lean_dec(x_16); lean_dec(x_10); lean_dec(x_9); @@ -1929,48 +2050,59 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_62 = !lean_is_exclusive(x_53); -if (x_62 == 0) +x_43 = !lean_is_exclusive(x_34); +if (x_43 == 0) { -return x_53; +return x_34; } else { -lean_object* x_63; lean_object* x_64; lean_object* x_65; -x_63 = lean_ctor_get(x_53, 0); -x_64 = lean_ctor_get(x_53, 1); -lean_inc(x_64); -lean_inc(x_63); -lean_dec(x_53); -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_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_34, 0); +x_45 = lean_ctor_get(x_34, 1); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_34); +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_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_66 = lean_ctor_get(x_25, 1); -lean_inc(x_66); -lean_dec(x_25); +uint8_t x_47; +x_47 = !lean_is_exclusive(x_29); +if (x_47 == 0) +{ +lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_48 = lean_ctor_get(x_29, 1); +x_49 = lean_ctor_get(x_29, 0); +lean_dec(x_49); +x_50 = l_Lean_Meta_Grind_updateLastTag(x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_48); +if (lean_obj_tag(x_50) == 0) +{ +lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; +x_51 = lean_ctor_get(x_50, 1); +lean_inc(x_51); +lean_dec(x_50); lean_inc(x_15); -x_67 = l_Lean_MessageData_ofExpr(x_15); -x_68 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_reinsertParents___spec__1___closed__2; -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_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; -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_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_24, x_71, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_66); -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_52 = l_Lean_MessageData_ofExpr(x_15); +x_53 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_reinsertParents___spec__1___closed__2; +lean_ctor_set_tag(x_29, 7); +lean_ctor_set(x_29, 1, x_52); +lean_ctor_set(x_29, 0, x_53); +x_54 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; +x_55 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_55, 0, x_29); +lean_ctor_set(x_55, 1, x_54); +x_56 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_28, x_55, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_51); +x_57 = lean_ctor_get(x_56, 0); +lean_inc(x_57); +x_58 = lean_ctor_get(x_56, 1); +lean_inc(x_58); +lean_dec(x_56); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); @@ -1979,16 +2111,16 @@ lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -x_75 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_reinsertParents___spec__1___lambda__1(x_15, x_73, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_74); -lean_dec(x_73); -if (lean_obj_tag(x_75) == 0) +x_59 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_reinsertParents___spec__1___lambda__1(x_15, x_57, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_58); +lean_dec(x_57); +if (lean_obj_tag(x_59) == 0) { -lean_object* x_76; -x_76 = lean_ctor_get(x_75, 0); -lean_inc(x_76); -if (lean_obj_tag(x_76) == 0) +lean_object* x_60; +x_60 = lean_ctor_get(x_59, 0); +lean_inc(x_60); +if (lean_obj_tag(x_60) == 0) { -lean_object* x_77; lean_object* x_78; lean_object* x_79; +uint8_t x_61; lean_dec(x_16); lean_dec(x_10); lean_dec(x_9); @@ -1998,43 +2130,44 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_77 = lean_ctor_get(x_75, 1); -lean_inc(x_77); -if (lean_is_exclusive(x_75)) { - lean_ctor_release(x_75, 0); - lean_ctor_release(x_75, 1); - x_78 = x_75; -} else { - lean_dec_ref(x_75); - x_78 = lean_box(0); +x_61 = !lean_is_exclusive(x_59); +if (x_61 == 0) +{ +lean_object* x_62; +x_62 = lean_ctor_get(x_59, 0); +lean_dec(x_62); +return x_59; } -if (lean_is_scalar(x_78)) { - x_79 = lean_alloc_ctor(0, 2, 0); -} else { - x_79 = x_78; +else +{ +lean_object* x_63; lean_object* x_64; +x_63 = lean_ctor_get(x_59, 1); +lean_inc(x_63); +lean_dec(x_59); +x_64 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_64, 0, x_60); +lean_ctor_set(x_64, 1, x_63); +return x_64; } -lean_ctor_set(x_79, 0, x_76); -lean_ctor_set(x_79, 1, x_77); -return x_79; } else { -lean_object* x_80; lean_object* x_81; -x_80 = lean_ctor_get(x_75, 1); -lean_inc(x_80); -lean_dec(x_75); -x_81 = lean_ctor_get(x_76, 0); -lean_inc(x_81); -lean_dec(x_76); +lean_object* x_65; lean_object* x_66; +x_65 = lean_ctor_get(x_59, 1); +lean_inc(x_65); +lean_dec(x_59); +x_66 = lean_ctor_get(x_60, 0); +lean_inc(x_66); +lean_dec(x_60); x_1 = x_16; -x_2 = x_81; -x_11 = x_80; +x_2 = x_66; +x_11 = x_65; goto _start; } } else { -lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; +uint8_t x_68; lean_dec(x_16); lean_dec(x_10); lean_dec(x_9); @@ -2044,26 +2177,219 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_83 = lean_ctor_get(x_75, 0); -lean_inc(x_83); -x_84 = lean_ctor_get(x_75, 1); -lean_inc(x_84); -if (lean_is_exclusive(x_75)) { - lean_ctor_release(x_75, 0); - lean_ctor_release(x_75, 1); - x_85 = x_75; +x_68 = !lean_is_exclusive(x_59); +if (x_68 == 0) +{ +return x_59; +} +else +{ +lean_object* x_69; lean_object* x_70; lean_object* x_71; +x_69 = lean_ctor_get(x_59, 0); +x_70 = lean_ctor_get(x_59, 1); +lean_inc(x_70); +lean_inc(x_69); +lean_dec(x_59); +x_71 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_71, 0, x_69); +lean_ctor_set(x_71, 1, x_70); +return x_71; +} +} +} +else +{ +uint8_t x_72; +lean_free_object(x_29); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_72 = !lean_is_exclusive(x_50); +if (x_72 == 0) +{ +return x_50; +} +else +{ +lean_object* x_73; lean_object* x_74; lean_object* x_75; +x_73 = lean_ctor_get(x_50, 0); +x_74 = lean_ctor_get(x_50, 1); +lean_inc(x_74); +lean_inc(x_73); +lean_dec(x_50); +x_75 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_75, 0, x_73); +lean_ctor_set(x_75, 1, x_74); +return x_75; +} +} +} +else +{ +lean_object* x_76; lean_object* x_77; +x_76 = lean_ctor_get(x_29, 1); +lean_inc(x_76); +lean_dec(x_29); +x_77 = l_Lean_Meta_Grind_updateLastTag(x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_76); +if (lean_obj_tag(x_77) == 0) +{ +lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; +x_78 = lean_ctor_get(x_77, 1); +lean_inc(x_78); +lean_dec(x_77); +lean_inc(x_15); +x_79 = l_Lean_MessageData_ofExpr(x_15); +x_80 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_reinsertParents___spec__1___closed__2; +x_81 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_81, 0, x_80); +lean_ctor_set(x_81, 1, x_79); +x_82 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; +x_83 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_83, 0, x_81); +lean_ctor_set(x_83, 1, x_82); +x_84 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_28, x_83, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_78); +x_85 = lean_ctor_get(x_84, 0); +lean_inc(x_85); +x_86 = lean_ctor_get(x_84, 1); +lean_inc(x_86); +lean_dec(x_84); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_87 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_reinsertParents___spec__1___lambda__1(x_15, x_85, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_86); +lean_dec(x_85); +if (lean_obj_tag(x_87) == 0) +{ +lean_object* x_88; +x_88 = lean_ctor_get(x_87, 0); +lean_inc(x_88); +if (lean_obj_tag(x_88) == 0) +{ +lean_object* x_89; lean_object* x_90; lean_object* x_91; +lean_dec(x_16); +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_89 = lean_ctor_get(x_87, 1); +lean_inc(x_89); +if (lean_is_exclusive(x_87)) { + lean_ctor_release(x_87, 0); + lean_ctor_release(x_87, 1); + x_90 = x_87; +} else { + lean_dec_ref(x_87); + x_90 = lean_box(0); +} +if (lean_is_scalar(x_90)) { + x_91 = lean_alloc_ctor(0, 2, 0); +} else { + x_91 = x_90; +} +lean_ctor_set(x_91, 0, x_88); +lean_ctor_set(x_91, 1, x_89); +return x_91; +} +else +{ +lean_object* x_92; lean_object* x_93; +x_92 = lean_ctor_get(x_87, 1); +lean_inc(x_92); +lean_dec(x_87); +x_93 = lean_ctor_get(x_88, 0); +lean_inc(x_93); +lean_dec(x_88); +x_1 = x_16; +x_2 = x_93; +x_11 = x_92; +goto _start; +} +} +else +{ +lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; +lean_dec(x_16); +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_95 = lean_ctor_get(x_87, 0); +lean_inc(x_95); +x_96 = lean_ctor_get(x_87, 1); +lean_inc(x_96); +if (lean_is_exclusive(x_87)) { + lean_ctor_release(x_87, 0); + lean_ctor_release(x_87, 1); + x_97 = x_87; +} else { + lean_dec_ref(x_87); + x_97 = lean_box(0); +} +if (lean_is_scalar(x_97)) { + x_98 = lean_alloc_ctor(1, 2, 0); +} else { + x_98 = x_97; +} +lean_ctor_set(x_98, 0, x_95); +lean_ctor_set(x_98, 1, x_96); +return x_98; +} +} +else +{ +lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_99 = lean_ctor_get(x_77, 0); +lean_inc(x_99); +x_100 = lean_ctor_get(x_77, 1); +lean_inc(x_100); +if (lean_is_exclusive(x_77)) { + lean_ctor_release(x_77, 0); + lean_ctor_release(x_77, 1); + x_101 = x_77; } else { - lean_dec_ref(x_75); - x_85 = lean_box(0); + lean_dec_ref(x_77); + x_101 = lean_box(0); } -if (lean_is_scalar(x_85)) { - x_86 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_101)) { + x_102 = lean_alloc_ctor(1, 2, 0); } else { - x_86 = x_85; + x_102 = x_101; +} +lean_ctor_set(x_102, 0, x_99); +lean_ctor_set(x_102, 1, x_100); +return x_102; +} } -lean_ctor_set(x_86, 0, x_83); -lean_ctor_set(x_86, 1, x_84); -return x_86; } } } @@ -2071,7 +2397,7 @@ return x_86; } else { -uint8_t x_87; +uint8_t x_113; lean_dec(x_16); lean_dec(x_15); lean_dec(x_10); @@ -2082,23 +2408,23 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_87 = !lean_is_exclusive(x_17); -if (x_87 == 0) +x_113 = !lean_is_exclusive(x_17); +if (x_113 == 0) { return x_17; } else { -lean_object* x_88; lean_object* x_89; lean_object* x_90; -x_88 = lean_ctor_get(x_17, 0); -x_89 = lean_ctor_get(x_17, 1); -lean_inc(x_89); -lean_inc(x_88); +lean_object* x_114; lean_object* x_115; lean_object* x_116; +x_114 = lean_ctor_get(x_17, 0); +x_115 = lean_ctor_get(x_17, 1); +lean_inc(x_115); +lean_inc(x_114); lean_dec(x_17); -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; +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; } } } @@ -2266,10 +2592,6 @@ x_27 = lean_ctor_get(x_25, 1); lean_inc(x_27); lean_dec(x_25); x_28 = l_Lean_Meta_Grind_closeGoal(x_26, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_27); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); return x_28; } else @@ -2591,10 +2913,6 @@ x_27 = lean_ctor_get(x_25, 1); lean_inc(x_27); lean_dec(x_25); x_28 = l_Lean_Meta_Grind_closeGoal(x_26, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_27); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); return x_28; } else @@ -4363,7 +4681,7 @@ x_29 = l_Lean_Meta_Grind_setENode(x_1, x_2, x_12, x_13, x_14, x_15, x_16, x_17, x_30 = !lean_is_exclusive(x_29); if (x_30 == 0) { -lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; uint8_t x_37; lean_object* x_38; uint8_t x_39; +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; uint8_t x_37; lean_object* x_38; x_31 = lean_ctor_get(x_29, 1); x_32 = lean_ctor_get(x_29, 0); lean_dec(x_32); @@ -4377,14 +4695,16 @@ x_36 = lean_ctor_get_uint8(x_6, sizeof(void*)*10 + 3); x_37 = lean_ctor_get_uint8(x_6, sizeof(void*)*10 + 4); lean_dec(x_6); x_38 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents(x_33, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_31); -x_39 = !lean_is_exclusive(x_38); -if (x_39 == 0) +if (lean_obj_tag(x_38) == 0) { -lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; -x_40 = lean_ctor_get(x_38, 0); -x_41 = lean_ctor_get(x_38, 1); -x_42 = lean_ctor_get(x_7, 2); -lean_inc(x_42); +lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_39 = lean_ctor_get(x_38, 0); +lean_inc(x_39); +x_40 = lean_ctor_get(x_38, 1); +lean_inc(x_40); +lean_dec(x_38); +x_41 = lean_ctor_get(x_7, 2); +lean_inc(x_41); lean_dec(x_7); lean_inc(x_19); lean_inc(x_18); @@ -4394,167 +4714,172 @@ lean_inc(x_15); lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); -lean_inc(x_42); +lean_inc(x_41); lean_inc(x_1); -x_43 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_updateRoots_loop(x_1, x_42, x_1, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_41); -if (lean_obj_tag(x_43) == 0) +x_42 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_updateRoots_loop(x_1, x_41, x_1, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_40); +if (lean_obj_tag(x_42) == 0) { -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); +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); lean_inc(x_8); -x_45 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_8, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_44); -x_46 = lean_ctor_get(x_45, 0); -lean_inc(x_46); -x_47 = lean_unbox(x_46); -lean_dec(x_46); -if (x_47 == 0) +x_44 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_8, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_43); +x_45 = lean_ctor_get(x_44, 0); +lean_inc(x_45); +x_46 = lean_unbox(x_45); +lean_dec(x_45); +if (x_46 == 0) { -lean_object* x_48; lean_object* x_49; lean_object* x_50; -lean_free_object(x_38); +lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_free_object(x_29); lean_dec(x_8); lean_dec(x_1); -x_48 = lean_ctor_get(x_45, 1); -lean_inc(x_48); -lean_dec(x_45); -x_49 = lean_box(0); -x_50 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__2(x_40, x_33, x_9, x_24, x_35, x_34, x_42, x_10, x_37, x_36, x_49, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_48); +x_47 = lean_ctor_get(x_44, 1); +lean_inc(x_47); +lean_dec(x_44); +x_48 = lean_box(0); +x_49 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__2(x_39, x_33, x_9, x_24, x_35, x_34, x_41, x_10, x_37, x_36, x_48, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_47); lean_dec(x_35); -return x_50; +return x_49; } else { -uint8_t x_51; -x_51 = !lean_is_exclusive(x_45); -if (x_51 == 0) -{ -lean_object* x_52; lean_object* x_53; lean_object* x_54; uint8_t x_55; -x_52 = lean_ctor_get(x_45, 1); -x_53 = lean_ctor_get(x_45, 0); -lean_dec(x_53); -x_54 = l_Lean_Meta_Grind_ppENodeRef(x_1, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_52); -x_55 = !lean_is_exclusive(x_54); -if (x_55 == 0) +uint8_t x_50; +x_50 = !lean_is_exclusive(x_44); +if (x_50 == 0) { -lean_object* x_56; lean_object* x_57; lean_object* x_58; uint8_t x_59; -x_56 = lean_ctor_get(x_54, 0); -x_57 = lean_ctor_get(x_54, 1); -x_58 = l_Lean_Meta_Grind_ppENodeRef(x_42, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_57); -x_59 = !lean_is_exclusive(x_58); -if (x_59 == 0) +lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_51 = lean_ctor_get(x_44, 1); +x_52 = lean_ctor_get(x_44, 0); +lean_dec(x_52); +x_53 = l_Lean_Meta_Grind_updateLastTag(x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_51); +if (lean_obj_tag(x_53) == 0) { -lean_object* x_60; lean_object* x_61; lean_object* x_62; -x_60 = lean_ctor_get(x_58, 0); -x_61 = lean_ctor_get(x_58, 1); -x_62 = l_Lean_Meta_Grind_getRoot(x_1, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_61); -if (lean_obj_tag(x_62) == 0) +lean_object* x_54; lean_object* x_55; uint8_t x_56; +x_54 = lean_ctor_get(x_53, 1); +lean_inc(x_54); +lean_dec(x_53); +x_55 = l_Lean_Meta_Grind_ppENodeRef(x_1, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_54); +x_56 = !lean_is_exclusive(x_55); +if (x_56 == 0) { -lean_object* x_63; lean_object* x_64; lean_object* x_65; uint8_t 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_57; lean_object* x_58; lean_object* x_59; uint8_t x_60; +x_57 = lean_ctor_get(x_55, 0); +x_58 = lean_ctor_get(x_55, 1); +x_59 = l_Lean_Meta_Grind_ppENodeRef(x_41, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_58); +x_60 = !lean_is_exclusive(x_59); +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 = lean_ctor_get(x_59, 1); +x_63 = l_Lean_Meta_Grind_getRoot(x_1, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_62); +if (lean_obj_tag(x_63) == 0) +{ +lean_object* x_64; lean_object* x_65; lean_object* x_66; uint8_t x_67; +x_64 = lean_ctor_get(x_63, 0); lean_inc(x_64); -lean_dec(x_62); -x_65 = l_Lean_Meta_Grind_ppENodeRef(x_63, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_64); +x_65 = lean_ctor_get(x_63, 1); +lean_inc(x_65); lean_dec(x_63); -x_66 = !lean_is_exclusive(x_65); -if (x_66 == 0) +x_66 = l_Lean_Meta_Grind_ppENodeRef(x_64, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_65); +lean_dec(x_64); +x_67 = !lean_is_exclusive(x_66); +if (x_67 == 0) { -lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; -x_67 = lean_ctor_get(x_65, 0); -x_68 = lean_ctor_get(x_65, 1); -x_69 = l_Lean_MessageData_ofFormat(x_56); -x_70 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; -lean_ctor_set_tag(x_65, 7); -lean_ctor_set(x_65, 1, x_69); -lean_ctor_set(x_65, 0, x_70); -x_71 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__2; -lean_ctor_set_tag(x_58, 7); -lean_ctor_set(x_58, 1, x_71); -lean_ctor_set(x_58, 0, x_65); -x_72 = l_Lean_MessageData_ofFormat(x_60); -lean_ctor_set_tag(x_54, 7); -lean_ctor_set(x_54, 1, x_72); -lean_ctor_set(x_54, 0, x_58); -x_73 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__4; -lean_ctor_set_tag(x_45, 7); -lean_ctor_set(x_45, 1, x_73); -lean_ctor_set(x_45, 0, x_54); -x_74 = l_Lean_MessageData_ofFormat(x_67); -lean_ctor_set_tag(x_38, 7); -lean_ctor_set(x_38, 1, x_74); -lean_ctor_set(x_38, 0, x_45); +lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; +x_68 = lean_ctor_get(x_66, 0); +x_69 = lean_ctor_get(x_66, 1); +x_70 = l_Lean_MessageData_ofFormat(x_57); +x_71 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; +lean_ctor_set_tag(x_66, 7); +lean_ctor_set(x_66, 1, x_70); +lean_ctor_set(x_66, 0, x_71); +x_72 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__2; +lean_ctor_set_tag(x_59, 7); +lean_ctor_set(x_59, 1, x_72); +lean_ctor_set(x_59, 0, x_66); +x_73 = l_Lean_MessageData_ofFormat(x_61); +lean_ctor_set_tag(x_55, 7); +lean_ctor_set(x_55, 1, x_73); +lean_ctor_set(x_55, 0, x_59); +x_74 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__4; +lean_ctor_set_tag(x_44, 7); +lean_ctor_set(x_44, 1, x_74); +lean_ctor_set(x_44, 0, x_55); +x_75 = l_Lean_MessageData_ofFormat(x_68); lean_ctor_set_tag(x_29, 7); -lean_ctor_set(x_29, 1, x_70); -lean_ctor_set(x_29, 0, x_38); -x_75 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_8, x_29, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_68); -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___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__2(x_40, x_33, x_9, x_24, x_35, x_34, x_42, x_10, x_37, x_36, x_76, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_77); -lean_dec(x_76); +lean_ctor_set(x_29, 1, x_75); +lean_ctor_set(x_29, 0, x_44); +x_76 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_76, 0, x_29); +lean_ctor_set(x_76, 1, x_71); +x_77 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_8, x_76, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_69); +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 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__2(x_39, x_33, x_9, x_24, x_35, x_34, x_41, x_10, x_37, x_36, x_78, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_79); +lean_dec(x_78); lean_dec(x_35); -return x_78; +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; 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_79 = lean_ctor_get(x_65, 0); -x_80 = lean_ctor_get(x_65, 1); -lean_inc(x_80); -lean_inc(x_79); -lean_dec(x_65); -x_81 = l_Lean_MessageData_ofFormat(x_56); -x_82 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; -x_83 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_83, 0, x_82); -lean_ctor_set(x_83, 1, x_81); -x_84 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__2; -lean_ctor_set_tag(x_58, 7); -lean_ctor_set(x_58, 1, x_84); -lean_ctor_set(x_58, 0, x_83); -x_85 = l_Lean_MessageData_ofFormat(x_60); -lean_ctor_set_tag(x_54, 7); -lean_ctor_set(x_54, 1, x_85); -lean_ctor_set(x_54, 0, x_58); -x_86 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__4; -lean_ctor_set_tag(x_45, 7); -lean_ctor_set(x_45, 1, x_86); -lean_ctor_set(x_45, 0, x_54); -x_87 = l_Lean_MessageData_ofFormat(x_79); -lean_ctor_set_tag(x_38, 7); -lean_ctor_set(x_38, 1, x_87); -lean_ctor_set(x_38, 0, x_45); +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_81 = lean_ctor_get(x_66, 0); +x_82 = lean_ctor_get(x_66, 1); +lean_inc(x_82); +lean_inc(x_81); +lean_dec(x_66); +x_83 = l_Lean_MessageData_ofFormat(x_57); +x_84 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; +x_85 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_85, 0, x_84); +lean_ctor_set(x_85, 1, x_83); +x_86 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__2; +lean_ctor_set_tag(x_59, 7); +lean_ctor_set(x_59, 1, x_86); +lean_ctor_set(x_59, 0, x_85); +x_87 = l_Lean_MessageData_ofFormat(x_61); +lean_ctor_set_tag(x_55, 7); +lean_ctor_set(x_55, 1, x_87); +lean_ctor_set(x_55, 0, x_59); +x_88 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__4; +lean_ctor_set_tag(x_44, 7); +lean_ctor_set(x_44, 1, x_88); +lean_ctor_set(x_44, 0, x_55); +x_89 = l_Lean_MessageData_ofFormat(x_81); lean_ctor_set_tag(x_29, 7); -lean_ctor_set(x_29, 1, x_82); -lean_ctor_set(x_29, 0, x_38); -x_88 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_8, x_29, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_80); -x_89 = lean_ctor_get(x_88, 0); -lean_inc(x_89); -x_90 = lean_ctor_get(x_88, 1); -lean_inc(x_90); -lean_dec(x_88); -x_91 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__2(x_40, x_33, x_9, x_24, x_35, x_34, x_42, x_10, x_37, x_36, x_89, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_90); -lean_dec(x_89); +lean_ctor_set(x_29, 1, x_89); +lean_ctor_set(x_29, 0, x_44); +x_90 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_90, 0, x_29); +lean_ctor_set(x_90, 1, x_84); +x_91 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_8, x_90, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_82); +x_92 = lean_ctor_get(x_91, 0); +lean_inc(x_92); +x_93 = lean_ctor_get(x_91, 1); +lean_inc(x_93); +lean_dec(x_91); +x_94 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__2(x_39, x_33, x_9, x_24, x_35, x_34, x_41, x_10, x_37, x_36, x_92, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_93); +lean_dec(x_92); lean_dec(x_35); -return x_91; +return x_94; } } else { -uint8_t x_92; -lean_free_object(x_58); -lean_dec(x_60); -lean_free_object(x_54); -lean_dec(x_56); -lean_free_object(x_45); -lean_dec(x_42); -lean_free_object(x_38); -lean_dec(x_40); +uint8_t x_95; +lean_free_object(x_59); +lean_dec(x_61); +lean_free_object(x_55); +lean_dec(x_57); +lean_free_object(x_44); +lean_dec(x_41); +lean_dec(x_39); lean_dec(x_35); lean_dec(x_34); lean_dec(x_33); @@ -4570,107 +4895,106 @@ lean_dec(x_13); lean_dec(x_12); lean_dec(x_9); lean_dec(x_8); -x_92 = !lean_is_exclusive(x_62); -if (x_92 == 0) +x_95 = !lean_is_exclusive(x_63); +if (x_95 == 0) { -return x_62; +return x_63; } else { -lean_object* x_93; lean_object* x_94; lean_object* x_95; -x_93 = lean_ctor_get(x_62, 0); -x_94 = lean_ctor_get(x_62, 1); -lean_inc(x_94); -lean_inc(x_93); -lean_dec(x_62); -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_96; lean_object* x_97; lean_object* x_98; +x_96 = lean_ctor_get(x_63, 0); +x_97 = lean_ctor_get(x_63, 1); +lean_inc(x_97); +lean_inc(x_96); +lean_dec(x_63); +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 { -lean_object* x_96; lean_object* x_97; lean_object* x_98; -x_96 = lean_ctor_get(x_58, 0); -x_97 = lean_ctor_get(x_58, 1); -lean_inc(x_97); -lean_inc(x_96); -lean_dec(x_58); -x_98 = l_Lean_Meta_Grind_getRoot(x_1, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, 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_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_dec(x_98); -x_101 = l_Lean_Meta_Grind_ppENodeRef(x_99, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_100); -lean_dec(x_99); +lean_inc(x_99); +lean_dec(x_59); +x_101 = l_Lean_Meta_Grind_getRoot(x_1, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_100); +if (lean_obj_tag(x_101) == 0) +{ +lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; x_102 = lean_ctor_get(x_101, 0); lean_inc(x_102); x_103 = lean_ctor_get(x_101, 1); lean_inc(x_103); -if (lean_is_exclusive(x_101)) { - lean_ctor_release(x_101, 0); - lean_ctor_release(x_101, 1); - x_104 = x_101; +lean_dec(x_101); +x_104 = l_Lean_Meta_Grind_ppENodeRef(x_102, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_103); +lean_dec(x_102); +x_105 = lean_ctor_get(x_104, 0); +lean_inc(x_105); +x_106 = lean_ctor_get(x_104, 1); +lean_inc(x_106); +if (lean_is_exclusive(x_104)) { + lean_ctor_release(x_104, 0); + lean_ctor_release(x_104, 1); + x_107 = x_104; } else { - lean_dec_ref(x_101); - x_104 = lean_box(0); + lean_dec_ref(x_104); + x_107 = lean_box(0); } -x_105 = l_Lean_MessageData_ofFormat(x_56); -x_106 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; -if (lean_is_scalar(x_104)) { - x_107 = lean_alloc_ctor(7, 2, 0); +x_108 = l_Lean_MessageData_ofFormat(x_57); +x_109 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; +if (lean_is_scalar(x_107)) { + x_110 = lean_alloc_ctor(7, 2, 0); } else { - x_107 = x_104; - lean_ctor_set_tag(x_107, 7); -} -lean_ctor_set(x_107, 0, x_106); -lean_ctor_set(x_107, 1, x_105); -x_108 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__2; -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_MessageData_ofFormat(x_96); -lean_ctor_set_tag(x_54, 7); -lean_ctor_set(x_54, 1, x_110); -lean_ctor_set(x_54, 0, x_109); -x_111 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__4; -lean_ctor_set_tag(x_45, 7); -lean_ctor_set(x_45, 1, x_111); -lean_ctor_set(x_45, 0, x_54); -x_112 = l_Lean_MessageData_ofFormat(x_102); -lean_ctor_set_tag(x_38, 7); -lean_ctor_set(x_38, 1, x_112); -lean_ctor_set(x_38, 0, x_45); + x_110 = x_107; + lean_ctor_set_tag(x_110, 7); +} +lean_ctor_set(x_110, 0, x_109); +lean_ctor_set(x_110, 1, x_108); +x_111 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__2; +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_MessageData_ofFormat(x_99); +lean_ctor_set_tag(x_55, 7); +lean_ctor_set(x_55, 1, x_113); +lean_ctor_set(x_55, 0, x_112); +x_114 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__4; +lean_ctor_set_tag(x_44, 7); +lean_ctor_set(x_44, 1, x_114); +lean_ctor_set(x_44, 0, x_55); +x_115 = l_Lean_MessageData_ofFormat(x_105); lean_ctor_set_tag(x_29, 7); -lean_ctor_set(x_29, 1, x_106); -lean_ctor_set(x_29, 0, x_38); -x_113 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_8, x_29, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_103); -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___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__2(x_40, x_33, x_9, x_24, x_35, x_34, x_42, x_10, x_37, x_36, x_114, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_115); -lean_dec(x_114); +lean_ctor_set(x_29, 1, x_115); +lean_ctor_set(x_29, 0, x_44); +x_116 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_116, 0, x_29); +lean_ctor_set(x_116, 1, x_109); +x_117 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_8, x_116, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_106); +x_118 = lean_ctor_get(x_117, 0); +lean_inc(x_118); +x_119 = lean_ctor_get(x_117, 1); +lean_inc(x_119); +lean_dec(x_117); +x_120 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__2(x_39, x_33, x_9, x_24, x_35, x_34, x_41, x_10, x_37, x_36, x_118, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_119); +lean_dec(x_118); lean_dec(x_35); -return x_116; +return x_120; } else { -lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; -lean_dec(x_96); -lean_free_object(x_54); -lean_dec(x_56); -lean_free_object(x_45); -lean_dec(x_42); -lean_free_object(x_38); -lean_dec(x_40); +lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; +lean_dec(x_99); +lean_free_object(x_55); +lean_dec(x_57); +lean_free_object(x_44); +lean_dec(x_41); +lean_dec(x_39); lean_dec(x_35); lean_dec(x_34); lean_dec(x_33); @@ -4686,128 +5010,127 @@ lean_dec(x_13); lean_dec(x_12); lean_dec(x_9); lean_dec(x_8); -x_117 = lean_ctor_get(x_98, 0); -lean_inc(x_117); -x_118 = lean_ctor_get(x_98, 1); -lean_inc(x_118); -if (lean_is_exclusive(x_98)) { - lean_ctor_release(x_98, 0); - lean_ctor_release(x_98, 1); - x_119 = x_98; +x_121 = lean_ctor_get(x_101, 0); +lean_inc(x_121); +x_122 = lean_ctor_get(x_101, 1); +lean_inc(x_122); +if (lean_is_exclusive(x_101)) { + lean_ctor_release(x_101, 0); + lean_ctor_release(x_101, 1); + x_123 = x_101; } else { - lean_dec_ref(x_98); - x_119 = lean_box(0); + lean_dec_ref(x_101); + x_123 = lean_box(0); } -if (lean_is_scalar(x_119)) { - x_120 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_123)) { + x_124 = lean_alloc_ctor(1, 2, 0); } else { - x_120 = x_119; + x_124 = x_123; } -lean_ctor_set(x_120, 0, x_117); -lean_ctor_set(x_120, 1, x_118); -return x_120; +lean_ctor_set(x_124, 0, x_121); +lean_ctor_set(x_124, 1, x_122); +return x_124; } } } else { -lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; -x_121 = lean_ctor_get(x_54, 0); -x_122 = lean_ctor_get(x_54, 1); -lean_inc(x_122); -lean_inc(x_121); -lean_dec(x_54); -x_123 = l_Lean_Meta_Grind_ppENodeRef(x_42, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_122); -x_124 = lean_ctor_get(x_123, 0); -lean_inc(x_124); -x_125 = lean_ctor_get(x_123, 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; lean_object* x_131; +x_125 = lean_ctor_get(x_55, 0); +x_126 = lean_ctor_get(x_55, 1); +lean_inc(x_126); lean_inc(x_125); -if (lean_is_exclusive(x_123)) { - lean_ctor_release(x_123, 0); - lean_ctor_release(x_123, 1); - x_126 = x_123; -} else { - lean_dec_ref(x_123); - x_126 = lean_box(0); -} -x_127 = l_Lean_Meta_Grind_getRoot(x_1, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_125); -if (lean_obj_tag(x_127) == 0) -{ -lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; +lean_dec(x_55); +x_127 = l_Lean_Meta_Grind_ppENodeRef(x_41, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, 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_Meta_Grind_ppENodeRef(x_128, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_129); -lean_dec(x_128); -x_131 = lean_ctor_get(x_130, 0); -lean_inc(x_131); -x_132 = lean_ctor_get(x_130, 1); +if (lean_is_exclusive(x_127)) { + lean_ctor_release(x_127, 0); + lean_ctor_release(x_127, 1); + x_130 = x_127; +} else { + lean_dec_ref(x_127); + x_130 = lean_box(0); +} +x_131 = l_Lean_Meta_Grind_getRoot(x_1, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_129); +if (lean_obj_tag(x_131) == 0) +{ +lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; +x_132 = lean_ctor_get(x_131, 0); lean_inc(x_132); -if (lean_is_exclusive(x_130)) { - lean_ctor_release(x_130, 0); - lean_ctor_release(x_130, 1); - x_133 = x_130; +x_133 = lean_ctor_get(x_131, 1); +lean_inc(x_133); +lean_dec(x_131); +x_134 = l_Lean_Meta_Grind_ppENodeRef(x_132, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_133); +lean_dec(x_132); +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_130); - x_133 = lean_box(0); + lean_dec_ref(x_134); + x_137 = lean_box(0); } -x_134 = l_Lean_MessageData_ofFormat(x_121); -x_135 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; -if (lean_is_scalar(x_133)) { - x_136 = lean_alloc_ctor(7, 2, 0); +x_138 = l_Lean_MessageData_ofFormat(x_125); +x_139 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; +if (lean_is_scalar(x_137)) { + x_140 = lean_alloc_ctor(7, 2, 0); } else { - x_136 = x_133; - lean_ctor_set_tag(x_136, 7); -} -lean_ctor_set(x_136, 0, x_135); -lean_ctor_set(x_136, 1, x_134); -x_137 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__2; -if (lean_is_scalar(x_126)) { - x_138 = lean_alloc_ctor(7, 2, 0); + x_140 = x_137; + lean_ctor_set_tag(x_140, 7); +} +lean_ctor_set(x_140, 0, x_139); +lean_ctor_set(x_140, 1, x_138); +x_141 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__2; +if (lean_is_scalar(x_130)) { + x_142 = lean_alloc_ctor(7, 2, 0); } else { - x_138 = x_126; - lean_ctor_set_tag(x_138, 7); -} -lean_ctor_set(x_138, 0, x_136); -lean_ctor_set(x_138, 1, x_137); -x_139 = l_Lean_MessageData_ofFormat(x_124); -x_140 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_140, 0, x_138); -lean_ctor_set(x_140, 1, x_139); -x_141 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__4; -lean_ctor_set_tag(x_45, 7); -lean_ctor_set(x_45, 1, x_141); -lean_ctor_set(x_45, 0, x_140); -x_142 = l_Lean_MessageData_ofFormat(x_131); -lean_ctor_set_tag(x_38, 7); -lean_ctor_set(x_38, 1, x_142); -lean_ctor_set(x_38, 0, x_45); + x_142 = x_130; + lean_ctor_set_tag(x_142, 7); +} +lean_ctor_set(x_142, 0, x_140); +lean_ctor_set(x_142, 1, x_141); +x_143 = l_Lean_MessageData_ofFormat(x_128); +x_144 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_144, 0, x_142); +lean_ctor_set(x_144, 1, x_143); +x_145 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__4; +lean_ctor_set_tag(x_44, 7); +lean_ctor_set(x_44, 1, x_145); +lean_ctor_set(x_44, 0, x_144); +x_146 = l_Lean_MessageData_ofFormat(x_135); lean_ctor_set_tag(x_29, 7); -lean_ctor_set(x_29, 1, x_135); -lean_ctor_set(x_29, 0, x_38); -x_143 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_8, x_29, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_132); -x_144 = lean_ctor_get(x_143, 0); -lean_inc(x_144); -x_145 = lean_ctor_get(x_143, 1); -lean_inc(x_145); -lean_dec(x_143); -x_146 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__2(x_40, x_33, x_9, x_24, x_35, x_34, x_42, x_10, x_37, x_36, x_144, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_145); -lean_dec(x_144); +lean_ctor_set(x_29, 1, x_146); +lean_ctor_set(x_29, 0, x_44); +x_147 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_147, 0, x_29); +lean_ctor_set(x_147, 1, x_139); +x_148 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_8, x_147, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_136); +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 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__2(x_39, x_33, x_9, x_24, x_35, x_34, x_41, x_10, x_37, x_36, x_149, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_150); +lean_dec(x_149); lean_dec(x_35); -return x_146; +return x_151; } else { -lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; -lean_dec(x_126); -lean_dec(x_124); -lean_dec(x_121); -lean_free_object(x_45); -lean_dec(x_42); -lean_free_object(x_38); -lean_dec(x_40); +lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; +lean_dec(x_130); +lean_dec(x_128); +lean_dec(x_125); +lean_free_object(x_44); +lean_dec(x_41); +lean_dec(x_39); lean_dec(x_35); lean_dec(x_34); lean_dec(x_33); @@ -4823,72 +5146,85 @@ lean_dec(x_13); lean_dec(x_12); lean_dec(x_9); lean_dec(x_8); -x_147 = lean_ctor_get(x_127, 0); -lean_inc(x_147); -x_148 = lean_ctor_get(x_127, 1); -lean_inc(x_148); -if (lean_is_exclusive(x_127)) { - lean_ctor_release(x_127, 0); - lean_ctor_release(x_127, 1); - x_149 = x_127; +x_152 = lean_ctor_get(x_131, 0); +lean_inc(x_152); +x_153 = lean_ctor_get(x_131, 1); +lean_inc(x_153); +if (lean_is_exclusive(x_131)) { + lean_ctor_release(x_131, 0); + lean_ctor_release(x_131, 1); + x_154 = x_131; } else { - lean_dec_ref(x_127); - x_149 = lean_box(0); + lean_dec_ref(x_131); + x_154 = lean_box(0); } -if (lean_is_scalar(x_149)) { - x_150 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_154)) { + x_155 = lean_alloc_ctor(1, 2, 0); } else { - x_150 = x_149; + x_155 = x_154; } -lean_ctor_set(x_150, 0, x_147); -lean_ctor_set(x_150, 1, x_148); -return x_150; +lean_ctor_set(x_155, 0, x_152); +lean_ctor_set(x_155, 1, x_153); +return x_155; } } } else { -lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; -x_151 = lean_ctor_get(x_45, 1); -lean_inc(x_151); -lean_dec(x_45); -x_152 = l_Lean_Meta_Grind_ppENodeRef(x_1, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_151); -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_152); - x_155 = lean_box(0); +uint8_t x_156; +lean_free_object(x_44); +lean_dec(x_41); +lean_dec(x_39); +lean_dec(x_35); +lean_dec(x_34); +lean_dec(x_33); +lean_free_object(x_29); +lean_dec(x_24); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_1); +x_156 = !lean_is_exclusive(x_53); +if (x_156 == 0) +{ +return x_53; } -x_156 = l_Lean_Meta_Grind_ppENodeRef(x_42, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_154); -x_157 = lean_ctor_get(x_156, 0); -lean_inc(x_157); -x_158 = lean_ctor_get(x_156, 1); +else +{ +lean_object* x_157; lean_object* x_158; lean_object* x_159; +x_157 = lean_ctor_get(x_53, 0); +x_158 = lean_ctor_get(x_53, 1); lean_inc(x_158); -if (lean_is_exclusive(x_156)) { - lean_ctor_release(x_156, 0); - lean_ctor_release(x_156, 1); - x_159 = x_156; -} else { - lean_dec_ref(x_156); - x_159 = lean_box(0); +lean_inc(x_157); +lean_dec(x_53); +x_159 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_159, 0, x_157); +lean_ctor_set(x_159, 1, x_158); +return x_159; +} } -x_160 = l_Lean_Meta_Grind_getRoot(x_1, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_158); -if (lean_obj_tag(x_160) == 0) +} +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; -x_161 = lean_ctor_get(x_160, 0); -lean_inc(x_161); -x_162 = lean_ctor_get(x_160, 1); +lean_object* x_160; lean_object* x_161; +x_160 = lean_ctor_get(x_44, 1); +lean_inc(x_160); +lean_dec(x_44); +x_161 = l_Lean_Meta_Grind_updateLastTag(x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_160); +if (lean_obj_tag(x_161) == 0) +{ +lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; +x_162 = lean_ctor_get(x_161, 1); lean_inc(x_162); -lean_dec(x_160); -x_163 = l_Lean_Meta_Grind_ppENodeRef(x_161, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_162); lean_dec(x_161); +x_163 = l_Lean_Meta_Grind_ppENodeRef(x_1, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_162); x_164 = lean_ctor_get(x_163, 0); lean_inc(x_164); x_165 = lean_ctor_get(x_163, 1); @@ -4901,66 +5237,101 @@ if (lean_is_exclusive(x_163)) { lean_dec_ref(x_163); x_166 = lean_box(0); } -x_167 = l_Lean_MessageData_ofFormat(x_153); -x_168 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; -if (lean_is_scalar(x_166)) { - x_169 = lean_alloc_ctor(7, 2, 0); +x_167 = l_Lean_Meta_Grind_ppENodeRef(x_41, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_165); +x_168 = lean_ctor_get(x_167, 0); +lean_inc(x_168); +x_169 = lean_ctor_get(x_167, 1); +lean_inc(x_169); +if (lean_is_exclusive(x_167)) { + lean_ctor_release(x_167, 0); + lean_ctor_release(x_167, 1); + x_170 = x_167; +} else { + lean_dec_ref(x_167); + x_170 = lean_box(0); +} +x_171 = l_Lean_Meta_Grind_getRoot(x_1, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_169); +if (lean_obj_tag(x_171) == 0) +{ +lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; +x_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_Grind_ppENodeRef(x_172, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_173); +lean_dec(x_172); +x_175 = lean_ctor_get(x_174, 0); +lean_inc(x_175); +x_176 = lean_ctor_get(x_174, 1); +lean_inc(x_176); +if (lean_is_exclusive(x_174)) { + lean_ctor_release(x_174, 0); + lean_ctor_release(x_174, 1); + x_177 = x_174; } else { - x_169 = x_166; - lean_ctor_set_tag(x_169, 7); -} -lean_ctor_set(x_169, 0, x_168); -lean_ctor_set(x_169, 1, x_167); -x_170 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__2; -if (lean_is_scalar(x_159)) { - x_171 = lean_alloc_ctor(7, 2, 0); + lean_dec_ref(x_174); + x_177 = lean_box(0); +} +x_178 = l_Lean_MessageData_ofFormat(x_164); +x_179 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; +if (lean_is_scalar(x_177)) { + x_180 = lean_alloc_ctor(7, 2, 0); +} else { + x_180 = x_177; + lean_ctor_set_tag(x_180, 7); +} +lean_ctor_set(x_180, 0, x_179); +lean_ctor_set(x_180, 1, x_178); +x_181 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__2; +if (lean_is_scalar(x_170)) { + x_182 = lean_alloc_ctor(7, 2, 0); } else { - x_171 = x_159; - lean_ctor_set_tag(x_171, 7); -} -lean_ctor_set(x_171, 0, x_169); -lean_ctor_set(x_171, 1, x_170); -x_172 = l_Lean_MessageData_ofFormat(x_157); -if (lean_is_scalar(x_155)) { - x_173 = lean_alloc_ctor(7, 2, 0); + x_182 = x_170; + lean_ctor_set_tag(x_182, 7); +} +lean_ctor_set(x_182, 0, x_180); +lean_ctor_set(x_182, 1, x_181); +x_183 = l_Lean_MessageData_ofFormat(x_168); +if (lean_is_scalar(x_166)) { + x_184 = lean_alloc_ctor(7, 2, 0); } else { - x_173 = x_155; - lean_ctor_set_tag(x_173, 7); -} -lean_ctor_set(x_173, 0, x_171); -lean_ctor_set(x_173, 1, x_172); -x_174 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__4; -x_175 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_175, 0, x_173); -lean_ctor_set(x_175, 1, x_174); -x_176 = l_Lean_MessageData_ofFormat(x_164); -lean_ctor_set_tag(x_38, 7); -lean_ctor_set(x_38, 1, x_176); -lean_ctor_set(x_38, 0, x_175); + x_184 = x_166; + lean_ctor_set_tag(x_184, 7); +} +lean_ctor_set(x_184, 0, x_182); +lean_ctor_set(x_184, 1, x_183); +x_185 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__4; +x_186 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_186, 0, x_184); +lean_ctor_set(x_186, 1, x_185); +x_187 = l_Lean_MessageData_ofFormat(x_175); lean_ctor_set_tag(x_29, 7); -lean_ctor_set(x_29, 1, x_168); -lean_ctor_set(x_29, 0, x_38); -x_177 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_8, x_29, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_165); -x_178 = lean_ctor_get(x_177, 0); -lean_inc(x_178); -x_179 = lean_ctor_get(x_177, 1); -lean_inc(x_179); -lean_dec(x_177); -x_180 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__2(x_40, x_33, x_9, x_24, x_35, x_34, x_42, x_10, x_37, x_36, x_178, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_179); -lean_dec(x_178); +lean_ctor_set(x_29, 1, x_187); +lean_ctor_set(x_29, 0, x_186); +x_188 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_188, 0, x_29); +lean_ctor_set(x_188, 1, x_179); +x_189 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_8, x_188, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_176); +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___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__2(x_39, x_33, x_9, x_24, x_35, x_34, x_41, x_10, x_37, x_36, x_190, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_191); +lean_dec(x_190); lean_dec(x_35); -return x_180; +return x_192; } else { -lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; -lean_dec(x_159); -lean_dec(x_157); -lean_dec(x_155); -lean_dec(x_153); -lean_dec(x_42); -lean_free_object(x_38); -lean_dec(x_40); +lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; +lean_dec(x_170); +lean_dec(x_168); +lean_dec(x_166); +lean_dec(x_164); +lean_dec(x_41); +lean_dec(x_39); lean_dec(x_35); lean_dec(x_34); lean_dec(x_33); @@ -4976,36 +5347,33 @@ lean_dec(x_13); lean_dec(x_12); lean_dec(x_9); lean_dec(x_8); -x_181 = lean_ctor_get(x_160, 0); -lean_inc(x_181); -x_182 = lean_ctor_get(x_160, 1); -lean_inc(x_182); -if (lean_is_exclusive(x_160)) { - lean_ctor_release(x_160, 0); - lean_ctor_release(x_160, 1); - x_183 = x_160; +x_193 = lean_ctor_get(x_171, 0); +lean_inc(x_193); +x_194 = lean_ctor_get(x_171, 1); +lean_inc(x_194); +if (lean_is_exclusive(x_171)) { + lean_ctor_release(x_171, 0); + lean_ctor_release(x_171, 1); + x_195 = x_171; } else { - lean_dec_ref(x_160); - x_183 = lean_box(0); + lean_dec_ref(x_171); + x_195 = lean_box(0); } -if (lean_is_scalar(x_183)) { - x_184 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_195)) { + x_196 = lean_alloc_ctor(1, 2, 0); } else { - x_184 = x_183; -} -lean_ctor_set(x_184, 0, x_181); -lean_ctor_set(x_184, 1, x_182); -return x_184; -} + x_196 = x_195; } +lean_ctor_set(x_196, 0, x_193); +lean_ctor_set(x_196, 1, x_194); +return x_196; } } else { -uint8_t x_185; -lean_dec(x_42); -lean_free_object(x_38); -lean_dec(x_40); +lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; +lean_dec(x_41); +lean_dec(x_39); lean_dec(x_35); lean_dec(x_34); lean_dec(x_33); @@ -5022,201 +5390,35 @@ lean_dec(x_12); lean_dec(x_9); lean_dec(x_8); lean_dec(x_1); -x_185 = !lean_is_exclusive(x_43); -if (x_185 == 0) -{ -return x_43; -} -else -{ -lean_object* x_186; lean_object* x_187; lean_object* x_188; -x_186 = lean_ctor_get(x_43, 0); -x_187 = lean_ctor_get(x_43, 1); -lean_inc(x_187); -lean_inc(x_186); -lean_dec(x_43); -x_188 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_188, 0, x_186); -lean_ctor_set(x_188, 1, x_187); -return x_188; -} -} -} -else -{ -lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; -x_189 = lean_ctor_get(x_38, 0); -x_190 = lean_ctor_get(x_38, 1); -lean_inc(x_190); -lean_inc(x_189); -lean_dec(x_38); -x_191 = lean_ctor_get(x_7, 2); -lean_inc(x_191); -lean_dec(x_7); -lean_inc(x_19); -lean_inc(x_18); -lean_inc(x_17); -lean_inc(x_16); -lean_inc(x_15); -lean_inc(x_14); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_191); -lean_inc(x_1); -x_192 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_updateRoots_loop(x_1, x_191, x_1, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_190); -if (lean_obj_tag(x_192) == 0) -{ -lean_object* x_193; lean_object* x_194; lean_object* x_195; uint8_t x_196; -x_193 = lean_ctor_get(x_192, 1); -lean_inc(x_193); -lean_dec(x_192); -lean_inc(x_8); -x_194 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_8, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_193); -x_195 = lean_ctor_get(x_194, 0); -lean_inc(x_195); -x_196 = lean_unbox(x_195); -lean_dec(x_195); -if (x_196 == 0) -{ -lean_object* x_197; lean_object* x_198; lean_object* x_199; -lean_free_object(x_29); -lean_dec(x_8); -lean_dec(x_1); -x_197 = lean_ctor_get(x_194, 1); +x_197 = lean_ctor_get(x_161, 0); lean_inc(x_197); -lean_dec(x_194); -x_198 = lean_box(0); -x_199 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__2(x_189, x_33, x_9, x_24, x_35, x_34, x_191, x_10, x_37, x_36, x_198, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_197); -lean_dec(x_35); -return x_199; -} -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_object* x_210; -x_200 = lean_ctor_get(x_194, 1); -lean_inc(x_200); -if (lean_is_exclusive(x_194)) { - lean_ctor_release(x_194, 0); - lean_ctor_release(x_194, 1); - x_201 = x_194; +x_198 = lean_ctor_get(x_161, 1); +lean_inc(x_198); +if (lean_is_exclusive(x_161)) { + lean_ctor_release(x_161, 0); + lean_ctor_release(x_161, 1); + x_199 = x_161; } else { - lean_dec_ref(x_194); - x_201 = lean_box(0); + lean_dec_ref(x_161); + x_199 = lean_box(0); } -x_202 = l_Lean_Meta_Grind_ppENodeRef(x_1, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_200); -x_203 = lean_ctor_get(x_202, 0); -lean_inc(x_203); -x_204 = lean_ctor_get(x_202, 1); -lean_inc(x_204); -if (lean_is_exclusive(x_202)) { - lean_ctor_release(x_202, 0); - lean_ctor_release(x_202, 1); - x_205 = x_202; +if (lean_is_scalar(x_199)) { + x_200 = lean_alloc_ctor(1, 2, 0); } else { - lean_dec_ref(x_202); - x_205 = lean_box(0); + x_200 = x_199; +} +lean_ctor_set(x_200, 0, x_197); +lean_ctor_set(x_200, 1, x_198); +return x_200; } -x_206 = l_Lean_Meta_Grind_ppENodeRef(x_191, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_204); -x_207 = lean_ctor_get(x_206, 0); -lean_inc(x_207); -x_208 = lean_ctor_get(x_206, 1); -lean_inc(x_208); -if (lean_is_exclusive(x_206)) { - lean_ctor_release(x_206, 0); - lean_ctor_release(x_206, 1); - x_209 = x_206; -} else { - lean_dec_ref(x_206); - x_209 = lean_box(0); } -x_210 = l_Lean_Meta_Grind_getRoot(x_1, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_208); -if (lean_obj_tag(x_210) == 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; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; -x_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_Lean_Meta_Grind_ppENodeRef(x_211, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_212); -lean_dec(x_211); -x_214 = lean_ctor_get(x_213, 0); -lean_inc(x_214); -x_215 = lean_ctor_get(x_213, 1); -lean_inc(x_215); -if (lean_is_exclusive(x_213)) { - lean_ctor_release(x_213, 0); - lean_ctor_release(x_213, 1); - x_216 = x_213; -} else { - lean_dec_ref(x_213); - x_216 = lean_box(0); } -x_217 = l_Lean_MessageData_ofFormat(x_203); -x_218 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; -if (lean_is_scalar(x_216)) { - x_219 = lean_alloc_ctor(7, 2, 0); -} else { - x_219 = x_216; - lean_ctor_set_tag(x_219, 7); -} -lean_ctor_set(x_219, 0, x_218); -lean_ctor_set(x_219, 1, x_217); -x_220 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__2; -if (lean_is_scalar(x_209)) { - x_221 = lean_alloc_ctor(7, 2, 0); -} else { - x_221 = x_209; - lean_ctor_set_tag(x_221, 7); -} -lean_ctor_set(x_221, 0, x_219); -lean_ctor_set(x_221, 1, x_220); -x_222 = l_Lean_MessageData_ofFormat(x_207); -if (lean_is_scalar(x_205)) { - x_223 = lean_alloc_ctor(7, 2, 0); -} else { - x_223 = x_205; - lean_ctor_set_tag(x_223, 7); -} -lean_ctor_set(x_223, 0, x_221); -lean_ctor_set(x_223, 1, x_222); -x_224 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__4; -if (lean_is_scalar(x_201)) { - x_225 = lean_alloc_ctor(7, 2, 0); -} else { - x_225 = x_201; - lean_ctor_set_tag(x_225, 7); -} -lean_ctor_set(x_225, 0, x_223); -lean_ctor_set(x_225, 1, x_224); -x_226 = l_Lean_MessageData_ofFormat(x_214); -x_227 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_227, 0, x_225); -lean_ctor_set(x_227, 1, x_226); -lean_ctor_set_tag(x_29, 7); -lean_ctor_set(x_29, 1, x_218); -lean_ctor_set(x_29, 0, x_227); -x_228 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_8, x_29, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_215); -x_229 = lean_ctor_get(x_228, 0); -lean_inc(x_229); -x_230 = lean_ctor_get(x_228, 1); -lean_inc(x_230); -lean_dec(x_228); -x_231 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__2(x_189, x_33, x_9, x_24, x_35, x_34, x_191, x_10, x_37, x_36, x_229, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_230); -lean_dec(x_229); -lean_dec(x_35); -return x_231; } else { -lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; -lean_dec(x_209); -lean_dec(x_207); -lean_dec(x_205); -lean_dec(x_203); -lean_dec(x_201); -lean_dec(x_191); -lean_dec(x_189); +uint8_t x_201; +lean_dec(x_41); +lean_dec(x_39); lean_dec(x_35); lean_dec(x_34); lean_dec(x_33); @@ -5232,34 +5434,30 @@ lean_dec(x_13); lean_dec(x_12); lean_dec(x_9); lean_dec(x_8); -x_232 = lean_ctor_get(x_210, 0); -lean_inc(x_232); -x_233 = lean_ctor_get(x_210, 1); -lean_inc(x_233); -if (lean_is_exclusive(x_210)) { - lean_ctor_release(x_210, 0); - lean_ctor_release(x_210, 1); - x_234 = x_210; -} else { - lean_dec_ref(x_210); - 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_dec(x_1); +x_201 = !lean_is_exclusive(x_42); +if (x_201 == 0) +{ +return x_42; } -lean_ctor_set(x_235, 0, x_232); -lean_ctor_set(x_235, 1, x_233); -return x_235; +else +{ +lean_object* x_202; lean_object* x_203; lean_object* x_204; +x_202 = lean_ctor_get(x_42, 0); +x_203 = lean_ctor_get(x_42, 1); +lean_inc(x_203); +lean_inc(x_202); +lean_dec(x_42); +x_204 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_204, 0, x_202); +lean_ctor_set(x_204, 1, x_203); +return x_204; } } } else { -lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; -lean_dec(x_191); -lean_dec(x_189); +uint8_t x_205; lean_dec(x_35); lean_dec(x_34); lean_dec(x_33); @@ -5275,60 +5473,54 @@ lean_dec(x_13); lean_dec(x_12); lean_dec(x_9); lean_dec(x_8); +lean_dec(x_7); lean_dec(x_1); -x_236 = lean_ctor_get(x_192, 0); -lean_inc(x_236); -x_237 = lean_ctor_get(x_192, 1); -lean_inc(x_237); -if (lean_is_exclusive(x_192)) { - lean_ctor_release(x_192, 0); - lean_ctor_release(x_192, 1); - x_238 = x_192; -} else { - lean_dec_ref(x_192); - x_238 = lean_box(0); -} -if (lean_is_scalar(x_238)) { - x_239 = lean_alloc_ctor(1, 2, 0); -} else { - x_239 = x_238; +x_205 = !lean_is_exclusive(x_38); +if (x_205 == 0) +{ +return x_38; } -lean_ctor_set(x_239, 0, x_236); -lean_ctor_set(x_239, 1, x_237); -return x_239; +else +{ +lean_object* x_206; lean_object* x_207; lean_object* x_208; +x_206 = lean_ctor_get(x_38, 0); +x_207 = lean_ctor_get(x_38, 1); +lean_inc(x_207); +lean_inc(x_206); +lean_dec(x_38); +x_208 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_208, 0, x_206); +lean_ctor_set(x_208, 1, x_207); +return x_208; } } } else { -lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; uint8_t x_244; uint8_t x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; -x_240 = lean_ctor_get(x_29, 1); -lean_inc(x_240); +lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; uint8_t x_213; uint8_t x_214; lean_object* x_215; +x_209 = lean_ctor_get(x_29, 1); +lean_inc(x_209); lean_dec(x_29); -x_241 = lean_ctor_get(x_6, 0); -lean_inc(x_241); -x_242 = lean_ctor_get(x_6, 1); -lean_inc(x_242); -x_243 = lean_ctor_get(x_6, 6); -lean_inc(x_243); -x_244 = lean_ctor_get_uint8(x_6, sizeof(void*)*10 + 3); -x_245 = lean_ctor_get_uint8(x_6, sizeof(void*)*10 + 4); +x_210 = lean_ctor_get(x_6, 0); +lean_inc(x_210); +x_211 = lean_ctor_get(x_6, 1); +lean_inc(x_211); +x_212 = lean_ctor_get(x_6, 6); +lean_inc(x_212); +x_213 = lean_ctor_get_uint8(x_6, sizeof(void*)*10 + 3); +x_214 = lean_ctor_get_uint8(x_6, sizeof(void*)*10 + 4); lean_dec(x_6); -x_246 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents(x_241, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_240); -x_247 = lean_ctor_get(x_246, 0); -lean_inc(x_247); -x_248 = lean_ctor_get(x_246, 1); -lean_inc(x_248); -if (lean_is_exclusive(x_246)) { - lean_ctor_release(x_246, 0); - lean_ctor_release(x_246, 1); - x_249 = x_246; -} else { - lean_dec_ref(x_246); - x_249 = lean_box(0); -} -x_250 = lean_ctor_get(x_7, 2); -lean_inc(x_250); +x_215 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents(x_210, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_209); +if (lean_obj_tag(x_215) == 0) +{ +lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; +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_7, 2); +lean_inc(x_218); lean_dec(x_7); lean_inc(x_19); lean_inc(x_18); @@ -5338,171 +5530,171 @@ lean_inc(x_15); lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); -lean_inc(x_250); +lean_inc(x_218); lean_inc(x_1); -x_251 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_updateRoots_loop(x_1, x_250, x_1, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_248); -if (lean_obj_tag(x_251) == 0) +x_219 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_updateRoots_loop(x_1, x_218, x_1, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_217); +if (lean_obj_tag(x_219) == 0) { -lean_object* x_252; lean_object* x_253; lean_object* x_254; uint8_t x_255; -x_252 = lean_ctor_get(x_251, 1); -lean_inc(x_252); -lean_dec(x_251); +lean_object* x_220; lean_object* x_221; lean_object* x_222; uint8_t x_223; +x_220 = lean_ctor_get(x_219, 1); +lean_inc(x_220); +lean_dec(x_219); lean_inc(x_8); -x_253 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_8, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_252); -x_254 = lean_ctor_get(x_253, 0); -lean_inc(x_254); -x_255 = lean_unbox(x_254); -lean_dec(x_254); -if (x_255 == 0) -{ -lean_object* x_256; lean_object* x_257; lean_object* x_258; -lean_dec(x_249); +x_221 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_8, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_220); +x_222 = lean_ctor_get(x_221, 0); +lean_inc(x_222); +x_223 = lean_unbox(x_222); +lean_dec(x_222); +if (x_223 == 0) +{ +lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_dec(x_8); lean_dec(x_1); -x_256 = lean_ctor_get(x_253, 1); -lean_inc(x_256); -lean_dec(x_253); -x_257 = lean_box(0); -x_258 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__2(x_247, x_241, x_9, x_24, x_243, x_242, x_250, x_10, x_245, x_244, x_257, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_256); -lean_dec(x_243); -return x_258; -} -else -{ -lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; -x_259 = lean_ctor_get(x_253, 1); -lean_inc(x_259); -if (lean_is_exclusive(x_253)) { - lean_ctor_release(x_253, 0); - lean_ctor_release(x_253, 1); - x_260 = x_253; +x_224 = lean_ctor_get(x_221, 1); +lean_inc(x_224); +lean_dec(x_221); +x_225 = lean_box(0); +x_226 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__2(x_216, x_210, x_9, x_24, x_212, x_211, x_218, x_10, x_214, x_213, x_225, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_224); +lean_dec(x_212); +return x_226; +} +else +{ +lean_object* x_227; lean_object* x_228; lean_object* x_229; +x_227 = lean_ctor_get(x_221, 1); +lean_inc(x_227); +if (lean_is_exclusive(x_221)) { + lean_ctor_release(x_221, 0); + lean_ctor_release(x_221, 1); + x_228 = x_221; } else { - lean_dec_ref(x_253); - x_260 = lean_box(0); + lean_dec_ref(x_221); + x_228 = lean_box(0); } -x_261 = l_Lean_Meta_Grind_ppENodeRef(x_1, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_259); -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; +x_229 = l_Lean_Meta_Grind_updateLastTag(x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_227); +if (lean_obj_tag(x_229) == 0) +{ +lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; +x_230 = lean_ctor_get(x_229, 1); +lean_inc(x_230); +lean_dec(x_229); +x_231 = l_Lean_Meta_Grind_ppENodeRef(x_1, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_230); +x_232 = lean_ctor_get(x_231, 0); +lean_inc(x_232); +x_233 = lean_ctor_get(x_231, 1); +lean_inc(x_233); +if (lean_is_exclusive(x_231)) { + lean_ctor_release(x_231, 0); + lean_ctor_release(x_231, 1); + x_234 = x_231; } else { - lean_dec_ref(x_261); - x_264 = lean_box(0); + lean_dec_ref(x_231); + x_234 = lean_box(0); } -x_265 = l_Lean_Meta_Grind_ppENodeRef(x_250, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_263); -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; +x_235 = l_Lean_Meta_Grind_ppENodeRef(x_218, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_233); +x_236 = lean_ctor_get(x_235, 0); +lean_inc(x_236); +x_237 = lean_ctor_get(x_235, 1); +lean_inc(x_237); +if (lean_is_exclusive(x_235)) { + lean_ctor_release(x_235, 0); + lean_ctor_release(x_235, 1); + x_238 = x_235; } else { - lean_dec_ref(x_265); - x_268 = lean_box(0); + lean_dec_ref(x_235); + x_238 = lean_box(0); } -x_269 = l_Lean_Meta_Grind_getRoot(x_1, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_267); -if (lean_obj_tag(x_269) == 0) +x_239 = l_Lean_Meta_Grind_getRoot(x_1, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_237); +if (lean_obj_tag(x_239) == 0) { -lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; 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; -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_272 = l_Lean_Meta_Grind_ppENodeRef(x_270, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_271); -lean_dec(x_270); -x_273 = lean_ctor_get(x_272, 0); -lean_inc(x_273); -x_274 = lean_ctor_get(x_272, 1); -lean_inc(x_274); -if (lean_is_exclusive(x_272)) { - lean_ctor_release(x_272, 0); - lean_ctor_release(x_272, 1); - x_275 = x_272; +lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; +x_240 = lean_ctor_get(x_239, 0); +lean_inc(x_240); +x_241 = lean_ctor_get(x_239, 1); +lean_inc(x_241); +lean_dec(x_239); +x_242 = l_Lean_Meta_Grind_ppENodeRef(x_240, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_241); +lean_dec(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); +if (lean_is_exclusive(x_242)) { + lean_ctor_release(x_242, 0); + lean_ctor_release(x_242, 1); + x_245 = x_242; } else { - lean_dec_ref(x_272); - x_275 = lean_box(0); + lean_dec_ref(x_242); + x_245 = lean_box(0); } -x_276 = l_Lean_MessageData_ofFormat(x_262); -x_277 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; -if (lean_is_scalar(x_275)) { - x_278 = lean_alloc_ctor(7, 2, 0); +x_246 = l_Lean_MessageData_ofFormat(x_232); +x_247 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; +if (lean_is_scalar(x_245)) { + x_248 = lean_alloc_ctor(7, 2, 0); } else { - x_278 = x_275; - lean_ctor_set_tag(x_278, 7); + x_248 = x_245; + lean_ctor_set_tag(x_248, 7); } -lean_ctor_set(x_278, 0, x_277); -lean_ctor_set(x_278, 1, x_276); -x_279 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__2; -if (lean_is_scalar(x_268)) { - x_280 = 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_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__2; +if (lean_is_scalar(x_238)) { + x_250 = lean_alloc_ctor(7, 2, 0); } else { - x_280 = x_268; - lean_ctor_set_tag(x_280, 7); + x_250 = x_238; + lean_ctor_set_tag(x_250, 7); } -lean_ctor_set(x_280, 0, x_278); -lean_ctor_set(x_280, 1, x_279); -x_281 = l_Lean_MessageData_ofFormat(x_266); -if (lean_is_scalar(x_264)) { - x_282 = lean_alloc_ctor(7, 2, 0); -} else { - x_282 = x_264; - lean_ctor_set_tag(x_282, 7); -} -lean_ctor_set(x_282, 0, x_280); -lean_ctor_set(x_282, 1, x_281); -x_283 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__4; -if (lean_is_scalar(x_260)) { - x_284 = 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_ofFormat(x_236); +if (lean_is_scalar(x_234)) { + x_252 = lean_alloc_ctor(7, 2, 0); } else { - x_284 = x_260; - lean_ctor_set_tag(x_284, 7); -} -lean_ctor_set(x_284, 0, x_282); -lean_ctor_set(x_284, 1, x_283); -x_285 = l_Lean_MessageData_ofFormat(x_273); -if (lean_is_scalar(x_249)) { - x_286 = lean_alloc_ctor(7, 2, 0); + x_252 = x_234; + lean_ctor_set_tag(x_252, 7); +} +lean_ctor_set(x_252, 0, x_250); +lean_ctor_set(x_252, 1, x_251); +x_253 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__4; +if (lean_is_scalar(x_228)) { + x_254 = lean_alloc_ctor(7, 2, 0); } else { - x_286 = x_249; - lean_ctor_set_tag(x_286, 7); -} -lean_ctor_set(x_286, 0, x_284); -lean_ctor_set(x_286, 1, x_285); -x_287 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_287, 0, x_286); -lean_ctor_set(x_287, 1, x_277); -x_288 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_8, x_287, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_274); -x_289 = lean_ctor_get(x_288, 0); -lean_inc(x_289); -x_290 = lean_ctor_get(x_288, 1); -lean_inc(x_290); -lean_dec(x_288); -x_291 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__2(x_247, x_241, x_9, x_24, x_243, x_242, x_250, x_10, x_245, x_244, x_289, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_290); -lean_dec(x_289); -lean_dec(x_243); -return x_291; -} -else -{ -lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; -lean_dec(x_268); -lean_dec(x_266); -lean_dec(x_264); -lean_dec(x_262); -lean_dec(x_260); -lean_dec(x_250); -lean_dec(x_249); -lean_dec(x_247); -lean_dec(x_243); -lean_dec(x_242); -lean_dec(x_241); + x_254 = x_228; + lean_ctor_set_tag(x_254, 7); +} +lean_ctor_set(x_254, 0, x_252); +lean_ctor_set(x_254, 1, x_253); +x_255 = l_Lean_MessageData_ofFormat(x_243); +x_256 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_256, 0, x_254); +lean_ctor_set(x_256, 1, x_255); +x_257 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_257, 0, x_256); +lean_ctor_set(x_257, 1, x_247); +x_258 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_8, x_257, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, 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 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__2(x_216, x_210, x_9, x_24, x_212, x_211, x_218, x_10, x_214, x_213, x_259, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_260); +lean_dec(x_259); +lean_dec(x_212); +return x_261; +} +else +{ +lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; +lean_dec(x_238); +lean_dec(x_236); +lean_dec(x_234); +lean_dec(x_232); +lean_dec(x_228); +lean_dec(x_218); +lean_dec(x_216); +lean_dec(x_212); +lean_dec(x_211); +lean_dec(x_210); lean_dec(x_24); lean_dec(x_19); lean_dec(x_18); @@ -5514,38 +5706,37 @@ lean_dec(x_13); lean_dec(x_12); lean_dec(x_9); lean_dec(x_8); -x_292 = lean_ctor_get(x_269, 0); -lean_inc(x_292); -x_293 = lean_ctor_get(x_269, 1); -lean_inc(x_293); -if (lean_is_exclusive(x_269)) { - lean_ctor_release(x_269, 0); - lean_ctor_release(x_269, 1); - x_294 = x_269; +x_262 = lean_ctor_get(x_239, 0); +lean_inc(x_262); +x_263 = lean_ctor_get(x_239, 1); +lean_inc(x_263); +if (lean_is_exclusive(x_239)) { + lean_ctor_release(x_239, 0); + lean_ctor_release(x_239, 1); + x_264 = x_239; } else { - lean_dec_ref(x_269); - x_294 = lean_box(0); + lean_dec_ref(x_239); + x_264 = lean_box(0); } -if (lean_is_scalar(x_294)) { - x_295 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_264)) { + x_265 = lean_alloc_ctor(1, 2, 0); } else { - x_295 = x_294; -} -lean_ctor_set(x_295, 0, x_292); -lean_ctor_set(x_295, 1, x_293); -return x_295; + 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_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; -lean_dec(x_250); -lean_dec(x_249); -lean_dec(x_247); -lean_dec(x_243); -lean_dec(x_242); -lean_dec(x_241); +lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; +lean_dec(x_228); +lean_dec(x_218); +lean_dec(x_216); +lean_dec(x_212); +lean_dec(x_211); +lean_dec(x_210); lean_dec(x_24); lean_dec(x_19); lean_dec(x_18); @@ -5558,110 +5749,190 @@ lean_dec(x_12); lean_dec(x_9); lean_dec(x_8); lean_dec(x_1); -x_296 = lean_ctor_get(x_251, 0); -lean_inc(x_296); -x_297 = lean_ctor_get(x_251, 1); -lean_inc(x_297); -if (lean_is_exclusive(x_251)) { - lean_ctor_release(x_251, 0); - lean_ctor_release(x_251, 1); - x_298 = x_251; -} else { - lean_dec_ref(x_251); - x_298 = lean_box(0); +x_266 = lean_ctor_get(x_229, 0); +lean_inc(x_266); +x_267 = lean_ctor_get(x_229, 1); +lean_inc(x_267); +if (lean_is_exclusive(x_229)) { + lean_ctor_release(x_229, 0); + lean_ctor_release(x_229, 1); + x_268 = x_229; +} else { + lean_dec_ref(x_229); + x_268 = lean_box(0); } -if (lean_is_scalar(x_298)) { - x_299 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_268)) { + x_269 = lean_alloc_ctor(1, 2, 0); } else { - x_299 = x_298; + x_269 = x_268; } -lean_ctor_set(x_299, 0, x_296); -lean_ctor_set(x_299, 1, x_297); -return x_299; +lean_ctor_set(x_269, 0, x_266); +lean_ctor_set(x_269, 1, x_267); +return x_269; } } } else { -lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; uint8_t x_305; uint8_t x_306; uint8_t x_307; uint8_t 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; uint8_t x_321; uint8_t x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; -x_300 = lean_ctor_get(x_2, 0); -x_301 = lean_ctor_get(x_2, 1); -x_302 = lean_ctor_get(x_2, 2); -x_303 = lean_ctor_get(x_2, 3); -x_304 = lean_ctor_get(x_2, 6); -x_305 = lean_ctor_get_uint8(x_2, sizeof(void*)*10 + 1); -x_306 = lean_ctor_get_uint8(x_2, sizeof(void*)*10 + 2); -x_307 = lean_ctor_get_uint8(x_2, sizeof(void*)*10 + 3); -x_308 = lean_ctor_get_uint8(x_2, sizeof(void*)*10 + 4); -x_309 = lean_ctor_get(x_2, 7); -x_310 = lean_ctor_get(x_2, 8); -x_311 = lean_ctor_get(x_2, 9); -lean_inc(x_311); -lean_inc(x_310); -lean_inc(x_309); -lean_inc(x_304); -lean_inc(x_303); -lean_inc(x_302); -lean_inc(x_301); -lean_inc(x_300); +lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; +lean_dec(x_218); +lean_dec(x_216); +lean_dec(x_212); +lean_dec(x_211); +lean_dec(x_210); +lean_dec(x_24); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_1); +x_270 = lean_ctor_get(x_219, 0); +lean_inc(x_270); +x_271 = lean_ctor_get(x_219, 1); +lean_inc(x_271); +if (lean_is_exclusive(x_219)) { + lean_ctor_release(x_219, 0); + lean_ctor_release(x_219, 1); + x_272 = x_219; +} else { + lean_dec_ref(x_219); + x_272 = lean_box(0); +} +if (lean_is_scalar(x_272)) { + x_273 = lean_alloc_ctor(1, 2, 0); +} else { + x_273 = x_272; +} +lean_ctor_set(x_273, 0, x_270); +lean_ctor_set(x_273, 1, x_271); +return x_273; +} +} +else +{ +lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; +lean_dec(x_212); +lean_dec(x_211); +lean_dec(x_210); +lean_dec(x_24); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_1); +x_274 = lean_ctor_get(x_215, 0); +lean_inc(x_274); +x_275 = lean_ctor_get(x_215, 1); +lean_inc(x_275); +if (lean_is_exclusive(x_215)) { + lean_ctor_release(x_215, 0); + lean_ctor_release(x_215, 1); + x_276 = x_215; +} else { + lean_dec_ref(x_215); + x_276 = lean_box(0); +} +if (lean_is_scalar(x_276)) { + x_277 = lean_alloc_ctor(1, 2, 0); +} else { + x_277 = x_276; +} +lean_ctor_set(x_277, 0, x_274); +lean_ctor_set(x_277, 1, x_275); +return x_277; +} +} +} +else +{ +lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; uint8_t x_283; uint8_t x_284; uint8_t x_285; uint8_t 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; uint8_t x_299; uint8_t x_300; lean_object* x_301; +x_278 = lean_ctor_get(x_2, 0); +x_279 = lean_ctor_get(x_2, 1); +x_280 = lean_ctor_get(x_2, 2); +x_281 = lean_ctor_get(x_2, 3); +x_282 = lean_ctor_get(x_2, 6); +x_283 = lean_ctor_get_uint8(x_2, sizeof(void*)*10 + 1); +x_284 = lean_ctor_get_uint8(x_2, sizeof(void*)*10 + 2); +x_285 = lean_ctor_get_uint8(x_2, sizeof(void*)*10 + 3); +x_286 = lean_ctor_get_uint8(x_2, sizeof(void*)*10 + 4); +x_287 = lean_ctor_get(x_2, 7); +x_288 = lean_ctor_get(x_2, 8); +x_289 = lean_ctor_get(x_2, 9); +lean_inc(x_289); +lean_inc(x_288); +lean_inc(x_287); +lean_inc(x_282); +lean_inc(x_281); +lean_inc(x_280); +lean_inc(x_279); +lean_inc(x_278); lean_dec(x_2); -x_312 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_312, 0, x_3); -x_313 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_313, 0, x_4); -lean_inc(x_302); -x_314 = lean_alloc_ctor(0, 10, 5); -lean_ctor_set(x_314, 0, x_300); -lean_ctor_set(x_314, 1, x_301); -lean_ctor_set(x_314, 2, x_302); -lean_ctor_set(x_314, 3, x_303); -lean_ctor_set(x_314, 4, x_312); -lean_ctor_set(x_314, 5, x_313); -lean_ctor_set(x_314, 6, x_304); -lean_ctor_set(x_314, 7, x_309); -lean_ctor_set(x_314, 8, x_310); -lean_ctor_set(x_314, 9, x_311); -lean_ctor_set_uint8(x_314, sizeof(void*)*10, x_5); -lean_ctor_set_uint8(x_314, sizeof(void*)*10 + 1, x_305); -lean_ctor_set_uint8(x_314, sizeof(void*)*10 + 2, x_306); -lean_ctor_set_uint8(x_314, sizeof(void*)*10 + 3, x_307); -lean_ctor_set_uint8(x_314, sizeof(void*)*10 + 4, x_308); +x_290 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_290, 0, x_3); +x_291 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_291, 0, x_4); +lean_inc(x_280); +x_292 = lean_alloc_ctor(0, 10, 5); +lean_ctor_set(x_292, 0, x_278); +lean_ctor_set(x_292, 1, x_279); +lean_ctor_set(x_292, 2, x_280); +lean_ctor_set(x_292, 3, x_281); +lean_ctor_set(x_292, 4, x_290); +lean_ctor_set(x_292, 5, x_291); +lean_ctor_set(x_292, 6, x_282); +lean_ctor_set(x_292, 7, x_287); +lean_ctor_set(x_292, 8, x_288); +lean_ctor_set(x_292, 9, x_289); +lean_ctor_set_uint8(x_292, sizeof(void*)*10, x_5); +lean_ctor_set_uint8(x_292, sizeof(void*)*10 + 1, x_283); +lean_ctor_set_uint8(x_292, sizeof(void*)*10 + 2, x_284); +lean_ctor_set_uint8(x_292, sizeof(void*)*10 + 3, x_285); +lean_ctor_set_uint8(x_292, sizeof(void*)*10 + 4, x_286); lean_inc(x_1); -x_315 = l_Lean_Meta_Grind_setENode(x_1, x_314, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_22); -x_316 = lean_ctor_get(x_315, 1); -lean_inc(x_316); -if (lean_is_exclusive(x_315)) { - lean_ctor_release(x_315, 0); - lean_ctor_release(x_315, 1); - x_317 = x_315; +x_293 = l_Lean_Meta_Grind_setENode(x_1, x_292, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_22); +x_294 = lean_ctor_get(x_293, 1); +lean_inc(x_294); +if (lean_is_exclusive(x_293)) { + lean_ctor_release(x_293, 0); + lean_ctor_release(x_293, 1); + x_295 = x_293; } else { - lean_dec_ref(x_315); - x_317 = lean_box(0); + lean_dec_ref(x_293); + x_295 = lean_box(0); } -x_318 = lean_ctor_get(x_6, 0); -lean_inc(x_318); -x_319 = lean_ctor_get(x_6, 1); -lean_inc(x_319); -x_320 = lean_ctor_get(x_6, 6); -lean_inc(x_320); -x_321 = lean_ctor_get_uint8(x_6, sizeof(void*)*10 + 3); -x_322 = lean_ctor_get_uint8(x_6, sizeof(void*)*10 + 4); +x_296 = lean_ctor_get(x_6, 0); +lean_inc(x_296); +x_297 = lean_ctor_get(x_6, 1); +lean_inc(x_297); +x_298 = lean_ctor_get(x_6, 6); +lean_inc(x_298); +x_299 = lean_ctor_get_uint8(x_6, sizeof(void*)*10 + 3); +x_300 = lean_ctor_get_uint8(x_6, sizeof(void*)*10 + 4); lean_dec(x_6); -x_323 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents(x_318, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_316); -x_324 = lean_ctor_get(x_323, 0); -lean_inc(x_324); -x_325 = lean_ctor_get(x_323, 1); -lean_inc(x_325); -if (lean_is_exclusive(x_323)) { - lean_ctor_release(x_323, 0); - lean_ctor_release(x_323, 1); - x_326 = x_323; -} else { - lean_dec_ref(x_323); - x_326 = lean_box(0); -} -x_327 = lean_ctor_get(x_7, 2); -lean_inc(x_327); +x_301 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents(x_296, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_294); +if (lean_obj_tag(x_301) == 0) +{ +lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; +x_302 = lean_ctor_get(x_301, 0); +lean_inc(x_302); +x_303 = lean_ctor_get(x_301, 1); +lean_inc(x_303); +lean_dec(x_301); +x_304 = lean_ctor_get(x_7, 2); +lean_inc(x_304); lean_dec(x_7); lean_inc(x_19); lean_inc(x_18); @@ -5671,179 +5942,222 @@ lean_inc(x_15); lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); -lean_inc(x_327); +lean_inc(x_304); lean_inc(x_1); -x_328 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_updateRoots_loop(x_1, x_327, x_1, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_325); -if (lean_obj_tag(x_328) == 0) +x_305 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_updateRoots_loop(x_1, x_304, x_1, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_303); +if (lean_obj_tag(x_305) == 0) { -lean_object* x_329; lean_object* x_330; lean_object* x_331; uint8_t x_332; -x_329 = lean_ctor_get(x_328, 1); -lean_inc(x_329); -lean_dec(x_328); +lean_object* x_306; lean_object* x_307; lean_object* x_308; uint8_t x_309; +x_306 = lean_ctor_get(x_305, 1); +lean_inc(x_306); +lean_dec(x_305); lean_inc(x_8); -x_330 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_8, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_329); -x_331 = lean_ctor_get(x_330, 0); -lean_inc(x_331); -x_332 = lean_unbox(x_331); -lean_dec(x_331); -if (x_332 == 0) -{ -lean_object* x_333; lean_object* x_334; lean_object* x_335; -lean_dec(x_326); -lean_dec(x_317); +x_307 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_8, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_306); +x_308 = lean_ctor_get(x_307, 0); +lean_inc(x_308); +x_309 = lean_unbox(x_308); +lean_dec(x_308); +if (x_309 == 0) +{ +lean_object* x_310; lean_object* x_311; lean_object* x_312; +lean_dec(x_295); lean_dec(x_8); lean_dec(x_1); -x_333 = lean_ctor_get(x_330, 1); -lean_inc(x_333); -lean_dec(x_330); -x_334 = lean_box(0); -x_335 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__2(x_324, x_318, x_9, x_302, x_320, x_319, x_327, x_10, x_322, x_321, x_334, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_333); -lean_dec(x_320); -return x_335; +x_310 = lean_ctor_get(x_307, 1); +lean_inc(x_310); +lean_dec(x_307); +x_311 = lean_box(0); +x_312 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__2(x_302, x_296, x_9, x_280, x_298, x_297, x_304, x_10, x_300, x_299, x_311, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_310); +lean_dec(x_298); +return x_312; } else { -lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; -x_336 = lean_ctor_get(x_330, 1); -lean_inc(x_336); -if (lean_is_exclusive(x_330)) { - lean_ctor_release(x_330, 0); - lean_ctor_release(x_330, 1); - x_337 = x_330; -} else { - lean_dec_ref(x_330); - x_337 = lean_box(0); -} -x_338 = l_Lean_Meta_Grind_ppENodeRef(x_1, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_336); -x_339 = lean_ctor_get(x_338, 0); -lean_inc(x_339); -x_340 = lean_ctor_get(x_338, 1); -lean_inc(x_340); -if (lean_is_exclusive(x_338)) { - lean_ctor_release(x_338, 0); - lean_ctor_release(x_338, 1); - x_341 = x_338; +lean_object* x_313; lean_object* x_314; lean_object* x_315; +x_313 = lean_ctor_get(x_307, 1); +lean_inc(x_313); +if (lean_is_exclusive(x_307)) { + lean_ctor_release(x_307, 0); + lean_ctor_release(x_307, 1); + x_314 = x_307; } else { - lean_dec_ref(x_338); - x_341 = lean_box(0); -} -x_342 = l_Lean_Meta_Grind_ppENodeRef(x_327, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_340); -x_343 = lean_ctor_get(x_342, 0); -lean_inc(x_343); -x_344 = lean_ctor_get(x_342, 1); -lean_inc(x_344); -if (lean_is_exclusive(x_342)) { - lean_ctor_release(x_342, 0); - lean_ctor_release(x_342, 1); - x_345 = x_342; -} else { - lean_dec_ref(x_342); - x_345 = lean_box(0); + lean_dec_ref(x_307); + x_314 = lean_box(0); } -x_346 = l_Lean_Meta_Grind_getRoot(x_1, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_344); -if (lean_obj_tag(x_346) == 0) +x_315 = l_Lean_Meta_Grind_updateLastTag(x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_313); +if (lean_obj_tag(x_315) == 0) { -lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; lean_object* x_353; lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; -x_347 = lean_ctor_get(x_346, 0); -lean_inc(x_347); -x_348 = lean_ctor_get(x_346, 1); -lean_inc(x_348); -lean_dec(x_346); -x_349 = l_Lean_Meta_Grind_ppENodeRef(x_347, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_348); -lean_dec(x_347); -x_350 = lean_ctor_get(x_349, 0); -lean_inc(x_350); -x_351 = lean_ctor_get(x_349, 1); -lean_inc(x_351); -if (lean_is_exclusive(x_349)) { - lean_ctor_release(x_349, 0); - lean_ctor_release(x_349, 1); - x_352 = x_349; +lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; +x_316 = lean_ctor_get(x_315, 1); +lean_inc(x_316); +lean_dec(x_315); +x_317 = l_Lean_Meta_Grind_ppENodeRef(x_1, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_316); +x_318 = lean_ctor_get(x_317, 0); +lean_inc(x_318); +x_319 = lean_ctor_get(x_317, 1); +lean_inc(x_319); +if (lean_is_exclusive(x_317)) { + lean_ctor_release(x_317, 0); + lean_ctor_release(x_317, 1); + x_320 = x_317; } else { - lean_dec_ref(x_349); - x_352 = lean_box(0); + lean_dec_ref(x_317); + x_320 = lean_box(0); +} +x_321 = l_Lean_Meta_Grind_ppENodeRef(x_304, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_319); +x_322 = lean_ctor_get(x_321, 0); +lean_inc(x_322); +x_323 = lean_ctor_get(x_321, 1); +lean_inc(x_323); +if (lean_is_exclusive(x_321)) { + lean_ctor_release(x_321, 0); + lean_ctor_release(x_321, 1); + x_324 = x_321; +} else { + lean_dec_ref(x_321); + x_324 = lean_box(0); } -x_353 = l_Lean_MessageData_ofFormat(x_339); -x_354 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; -if (lean_is_scalar(x_352)) { - x_355 = lean_alloc_ctor(7, 2, 0); +x_325 = l_Lean_Meta_Grind_getRoot(x_1, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_323); +if (lean_obj_tag(x_325) == 0) +{ +lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; +x_326 = lean_ctor_get(x_325, 0); +lean_inc(x_326); +x_327 = lean_ctor_get(x_325, 1); +lean_inc(x_327); +lean_dec(x_325); +x_328 = l_Lean_Meta_Grind_ppENodeRef(x_326, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_327); +lean_dec(x_326); +x_329 = lean_ctor_get(x_328, 0); +lean_inc(x_329); +x_330 = lean_ctor_get(x_328, 1); +lean_inc(x_330); +if (lean_is_exclusive(x_328)) { + lean_ctor_release(x_328, 0); + lean_ctor_release(x_328, 1); + x_331 = x_328; } else { - x_355 = x_352; - lean_ctor_set_tag(x_355, 7); + lean_dec_ref(x_328); + x_331 = lean_box(0); } -lean_ctor_set(x_355, 0, x_354); -lean_ctor_set(x_355, 1, x_353); -x_356 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__2; -if (lean_is_scalar(x_345)) { - x_357 = lean_alloc_ctor(7, 2, 0); +x_332 = l_Lean_MessageData_ofFormat(x_318); +x_333 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; +if (lean_is_scalar(x_331)) { + x_334 = lean_alloc_ctor(7, 2, 0); } else { - x_357 = x_345; - lean_ctor_set_tag(x_357, 7); -} -lean_ctor_set(x_357, 0, x_355); -lean_ctor_set(x_357, 1, x_356); -x_358 = l_Lean_MessageData_ofFormat(x_343); -if (lean_is_scalar(x_341)) { - x_359 = lean_alloc_ctor(7, 2, 0); + x_334 = x_331; + lean_ctor_set_tag(x_334, 7); +} +lean_ctor_set(x_334, 0, x_333); +lean_ctor_set(x_334, 1, x_332); +x_335 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__2; +if (lean_is_scalar(x_324)) { + x_336 = lean_alloc_ctor(7, 2, 0); } else { - x_359 = x_341; - lean_ctor_set_tag(x_359, 7); -} -lean_ctor_set(x_359, 0, x_357); -lean_ctor_set(x_359, 1, x_358); -x_360 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__4; -if (lean_is_scalar(x_337)) { - x_361 = lean_alloc_ctor(7, 2, 0); + x_336 = x_324; + lean_ctor_set_tag(x_336, 7); +} +lean_ctor_set(x_336, 0, x_334); +lean_ctor_set(x_336, 1, x_335); +x_337 = l_Lean_MessageData_ofFormat(x_322); +if (lean_is_scalar(x_320)) { + x_338 = lean_alloc_ctor(7, 2, 0); } else { - x_361 = x_337; - lean_ctor_set_tag(x_361, 7); -} -lean_ctor_set(x_361, 0, x_359); -lean_ctor_set(x_361, 1, x_360); -x_362 = l_Lean_MessageData_ofFormat(x_350); -if (lean_is_scalar(x_326)) { - x_363 = lean_alloc_ctor(7, 2, 0); + x_338 = x_320; + lean_ctor_set_tag(x_338, 7); +} +lean_ctor_set(x_338, 0, x_336); +lean_ctor_set(x_338, 1, x_337); +x_339 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__4; +if (lean_is_scalar(x_314)) { + x_340 = lean_alloc_ctor(7, 2, 0); } else { - x_363 = x_326; - lean_ctor_set_tag(x_363, 7); -} -lean_ctor_set(x_363, 0, x_361); -lean_ctor_set(x_363, 1, x_362); -if (lean_is_scalar(x_317)) { - x_364 = lean_alloc_ctor(7, 2, 0); + x_340 = x_314; + lean_ctor_set_tag(x_340, 7); +} +lean_ctor_set(x_340, 0, x_338); +lean_ctor_set(x_340, 1, x_339); +x_341 = l_Lean_MessageData_ofFormat(x_329); +if (lean_is_scalar(x_295)) { + x_342 = lean_alloc_ctor(7, 2, 0); } else { - x_364 = x_317; - lean_ctor_set_tag(x_364, 7); -} -lean_ctor_set(x_364, 0, x_363); -lean_ctor_set(x_364, 1, x_354); -x_365 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_8, x_364, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_351); -x_366 = lean_ctor_get(x_365, 0); -lean_inc(x_366); -x_367 = lean_ctor_get(x_365, 1); -lean_inc(x_367); -lean_dec(x_365); -x_368 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__2(x_324, x_318, x_9, x_302, x_320, x_319, x_327, x_10, x_322, x_321, x_366, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_367); -lean_dec(x_366); -lean_dec(x_320); -return x_368; + x_342 = x_295; + lean_ctor_set_tag(x_342, 7); +} +lean_ctor_set(x_342, 0, x_340); +lean_ctor_set(x_342, 1, x_341); +x_343 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_343, 0, x_342); +lean_ctor_set(x_343, 1, x_333); +x_344 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_8, x_343, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_330); +x_345 = lean_ctor_get(x_344, 0); +lean_inc(x_345); +x_346 = lean_ctor_get(x_344, 1); +lean_inc(x_346); +lean_dec(x_344); +x_347 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__2(x_302, x_296, x_9, x_280, x_298, x_297, x_304, x_10, x_300, x_299, x_345, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_346); +lean_dec(x_345); +lean_dec(x_298); +return x_347; } else { -lean_object* x_369; lean_object* x_370; lean_object* x_371; lean_object* x_372; -lean_dec(x_345); -lean_dec(x_343); -lean_dec(x_341); -lean_dec(x_339); -lean_dec(x_337); -lean_dec(x_327); -lean_dec(x_326); +lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_dec(x_324); +lean_dec(x_322); lean_dec(x_320); -lean_dec(x_319); lean_dec(x_318); -lean_dec(x_317); +lean_dec(x_314); +lean_dec(x_304); +lean_dec(x_302); +lean_dec(x_298); +lean_dec(x_297); +lean_dec(x_296); +lean_dec(x_295); +lean_dec(x_280); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_9); +lean_dec(x_8); +x_348 = lean_ctor_get(x_325, 0); +lean_inc(x_348); +x_349 = lean_ctor_get(x_325, 1); +lean_inc(x_349); +if (lean_is_exclusive(x_325)) { + lean_ctor_release(x_325, 0); + lean_ctor_release(x_325, 1); + x_350 = x_325; +} else { + lean_dec_ref(x_325); + x_350 = lean_box(0); +} +if (lean_is_scalar(x_350)) { + x_351 = lean_alloc_ctor(1, 2, 0); +} else { + x_351 = x_350; +} +lean_ctor_set(x_351, 0, x_348); +lean_ctor_set(x_351, 1, x_349); +return x_351; +} +} +else +{ +lean_object* x_352; lean_object* x_353; lean_object* x_354; lean_object* x_355; +lean_dec(x_314); +lean_dec(x_304); lean_dec(x_302); +lean_dec(x_298); +lean_dec(x_297); +lean_dec(x_296); +lean_dec(x_295); +lean_dec(x_280); lean_dec(x_19); lean_dec(x_18); lean_dec(x_17); @@ -5854,40 +6168,40 @@ lean_dec(x_13); lean_dec(x_12); lean_dec(x_9); lean_dec(x_8); -x_369 = lean_ctor_get(x_346, 0); -lean_inc(x_369); -x_370 = lean_ctor_get(x_346, 1); -lean_inc(x_370); -if (lean_is_exclusive(x_346)) { - lean_ctor_release(x_346, 0); - lean_ctor_release(x_346, 1); - x_371 = x_346; +lean_dec(x_1); +x_352 = lean_ctor_get(x_315, 0); +lean_inc(x_352); +x_353 = lean_ctor_get(x_315, 1); +lean_inc(x_353); +if (lean_is_exclusive(x_315)) { + lean_ctor_release(x_315, 0); + lean_ctor_release(x_315, 1); + x_354 = x_315; } else { - lean_dec_ref(x_346); - x_371 = lean_box(0); + lean_dec_ref(x_315); + x_354 = lean_box(0); } -if (lean_is_scalar(x_371)) { - x_372 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_354)) { + x_355 = lean_alloc_ctor(1, 2, 0); } else { - x_372 = x_371; + x_355 = x_354; } -lean_ctor_set(x_372, 0, x_369); -lean_ctor_set(x_372, 1, x_370); -return x_372; +lean_ctor_set(x_355, 0, x_352); +lean_ctor_set(x_355, 1, x_353); +return x_355; } } } else { -lean_object* x_373; lean_object* x_374; lean_object* x_375; lean_object* x_376; -lean_dec(x_327); -lean_dec(x_326); -lean_dec(x_324); -lean_dec(x_320); -lean_dec(x_319); -lean_dec(x_318); -lean_dec(x_317); +lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; +lean_dec(x_304); lean_dec(x_302); +lean_dec(x_298); +lean_dec(x_297); +lean_dec(x_296); +lean_dec(x_295); +lean_dec(x_280); lean_dec(x_19); lean_dec(x_18); lean_dec(x_17); @@ -5899,32 +6213,74 @@ lean_dec(x_12); lean_dec(x_9); lean_dec(x_8); lean_dec(x_1); -x_373 = lean_ctor_get(x_328, 0); -lean_inc(x_373); -x_374 = lean_ctor_get(x_328, 1); -lean_inc(x_374); -if (lean_is_exclusive(x_328)) { - lean_ctor_release(x_328, 0); - lean_ctor_release(x_328, 1); - x_375 = x_328; +x_356 = lean_ctor_get(x_305, 0); +lean_inc(x_356); +x_357 = lean_ctor_get(x_305, 1); +lean_inc(x_357); +if (lean_is_exclusive(x_305)) { + lean_ctor_release(x_305, 0); + lean_ctor_release(x_305, 1); + x_358 = x_305; } else { - lean_dec_ref(x_328); - x_375 = lean_box(0); + lean_dec_ref(x_305); + x_358 = lean_box(0); +} +if (lean_is_scalar(x_358)) { + x_359 = lean_alloc_ctor(1, 2, 0); +} else { + x_359 = x_358; +} +lean_ctor_set(x_359, 0, x_356); +lean_ctor_set(x_359, 1, x_357); +return x_359; +} +} +else +{ +lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; +lean_dec(x_298); +lean_dec(x_297); +lean_dec(x_296); +lean_dec(x_295); +lean_dec(x_280); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_1); +x_360 = lean_ctor_get(x_301, 0); +lean_inc(x_360); +x_361 = lean_ctor_get(x_301, 1); +lean_inc(x_361); +if (lean_is_exclusive(x_301)) { + lean_ctor_release(x_301, 0); + lean_ctor_release(x_301, 1); + x_362 = x_301; +} else { + lean_dec_ref(x_301); + x_362 = lean_box(0); } -if (lean_is_scalar(x_375)) { - x_376 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_362)) { + x_363 = lean_alloc_ctor(1, 2, 0); } else { - x_376 = x_375; + x_363 = x_362; } -lean_ctor_set(x_376, 0, x_373); -lean_ctor_set(x_376, 1, x_374); -return x_376; +lean_ctor_set(x_363, 0, x_360); +lean_ctor_set(x_363, 1, x_361); +return x_363; } } } else { -uint8_t x_377; +uint8_t x_364; lean_dec(x_19); lean_dec(x_18); lean_dec(x_17); @@ -5941,23 +6297,23 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_377 = !lean_is_exclusive(x_21); -if (x_377 == 0) +x_364 = !lean_is_exclusive(x_21); +if (x_364 == 0) { return x_21; } else { -lean_object* x_378; lean_object* x_379; lean_object* x_380; -x_378 = lean_ctor_get(x_21, 0); -x_379 = lean_ctor_get(x_21, 1); -lean_inc(x_379); -lean_inc(x_378); +lean_object* x_365; lean_object* x_366; lean_object* x_367; +x_365 = lean_ctor_get(x_21, 0); +x_366 = lean_ctor_get(x_21, 1); +lean_inc(x_366); +lean_inc(x_365); lean_dec(x_21); -x_380 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_380, 0, x_378); -lean_ctor_set(x_380, 1, x_379); -return x_380; +x_367 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_367, 0, x_365); +lean_ctor_set(x_367, 1, x_366); +return x_367; } } } @@ -6011,7 +6367,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Gr { lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; x_19 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__1; -x_20 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_19, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +x_20 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_19, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); x_21 = lean_ctor_get(x_20, 0); lean_inc(x_21); x_22 = lean_unbox(x_21); @@ -6032,209 +6388,302 @@ uint8_t x_26; x_26 = !lean_is_exclusive(x_20); if (x_26 == 0) { -lean_object* x_27; lean_object* x_28; lean_object* x_29; uint8_t x_30; +lean_object* x_27; lean_object* x_28; lean_object* x_29; x_27 = lean_ctor_get(x_20, 1); x_28 = lean_ctor_get(x_20, 0); lean_dec(x_28); -x_29 = l_Lean_Meta_Grind_ppENodeRef(x_3, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_27); -x_30 = !lean_is_exclusive(x_29); -if (x_30 == 0) +x_29 = l_Lean_Meta_Grind_updateLastTag(x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_27); +if (lean_obj_tag(x_29) == 0) { -lean_object* x_31; lean_object* x_32; lean_object* x_33; uint8_t x_34; -x_31 = lean_ctor_get(x_29, 0); -x_32 = lean_ctor_get(x_29, 1); -x_33 = l_Lean_Meta_Grind_ppENodeRef(x_4, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_32); -x_34 = !lean_is_exclusive(x_33); -if (x_34 == 0) +lean_object* x_30; lean_object* x_31; uint8_t x_32; +x_30 = lean_ctor_get(x_29, 1); +lean_inc(x_30); +lean_dec(x_29); +x_31 = l_Lean_Meta_Grind_ppENodeRef(x_3, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_30); +x_32 = !lean_is_exclusive(x_31); +if (x_32 == 0) { -lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; -x_35 = lean_ctor_get(x_33, 0); -x_36 = lean_ctor_get(x_33, 1); -x_37 = l_Lean_MessageData_ofFormat(x_31); -x_38 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__3; -lean_ctor_set_tag(x_33, 7); -lean_ctor_set(x_33, 1, x_37); -lean_ctor_set(x_33, 0, x_38); -x_39 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__5; -lean_ctor_set_tag(x_29, 7); -lean_ctor_set(x_29, 1, x_39); -lean_ctor_set(x_29, 0, x_33); -x_40 = l_Lean_MessageData_ofFormat(x_35); +lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; +x_33 = lean_ctor_get(x_31, 0); +x_34 = lean_ctor_get(x_31, 1); +x_35 = l_Lean_Meta_Grind_ppENodeRef(x_4, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_34); +x_36 = !lean_is_exclusive(x_35); +if (x_36 == 0) +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; 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_35, 0); +x_38 = lean_ctor_get(x_35, 1); +x_39 = l_Lean_MessageData_ofFormat(x_33); +x_40 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__3; +lean_ctor_set_tag(x_35, 7); +lean_ctor_set(x_35, 1, x_39); +lean_ctor_set(x_35, 0, x_40); +x_41 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__5; +lean_ctor_set_tag(x_31, 7); +lean_ctor_set(x_31, 1, x_41); +lean_ctor_set(x_31, 0, x_35); +x_42 = l_Lean_MessageData_ofFormat(x_37); lean_ctor_set_tag(x_20, 7); -lean_ctor_set(x_20, 1, x_40); -lean_ctor_set(x_20, 0, x_29); -x_41 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; -x_42 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_42, 0, x_20); -lean_ctor_set(x_42, 1, x_41); -x_43 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_19, x_42, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_36); -x_44 = lean_ctor_get(x_43, 0); -lean_inc(x_44); -x_45 = lean_ctor_get(x_43, 1); -lean_inc(x_45); -lean_dec(x_43); -x_46 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3(x_3, x_5, x_4, x_1, x_9, x_7, x_6, x_19, x_8, x_2, x_44, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_45); -lean_dec(x_44); -return x_46; +lean_ctor_set(x_20, 1, x_42); +lean_ctor_set(x_20, 0, x_31); +x_43 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; +x_44 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_44, 0, x_20); +lean_ctor_set(x_44, 1, x_43); +x_45 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_19, x_44, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_38); +x_46 = lean_ctor_get(x_45, 0); +lean_inc(x_46); +x_47 = lean_ctor_get(x_45, 1); +lean_inc(x_47); +lean_dec(x_45); +x_48 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3(x_3, x_5, x_4, x_1, x_9, x_7, x_6, x_19, x_8, x_2, x_46, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_47); +lean_dec(x_46); +return x_48; } 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_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; -x_47 = lean_ctor_get(x_33, 0); -x_48 = lean_ctor_get(x_33, 1); -lean_inc(x_48); -lean_inc(x_47); -lean_dec(x_33); -x_49 = l_Lean_MessageData_ofFormat(x_31); -x_50 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__3; -x_51 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_51, 0, x_50); -lean_ctor_set(x_51, 1, x_49); -x_52 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__5; -lean_ctor_set_tag(x_29, 7); -lean_ctor_set(x_29, 1, x_52); -lean_ctor_set(x_29, 0, x_51); -x_53 = l_Lean_MessageData_ofFormat(x_47); +lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; 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_35, 0); +x_50 = lean_ctor_get(x_35, 1); +lean_inc(x_50); +lean_inc(x_49); +lean_dec(x_35); +x_51 = l_Lean_MessageData_ofFormat(x_33); +x_52 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__3; +x_53 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_53, 0, x_52); +lean_ctor_set(x_53, 1, x_51); +x_54 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__5; +lean_ctor_set_tag(x_31, 7); +lean_ctor_set(x_31, 1, x_54); +lean_ctor_set(x_31, 0, x_53); +x_55 = l_Lean_MessageData_ofFormat(x_49); lean_ctor_set_tag(x_20, 7); -lean_ctor_set(x_20, 1, x_53); -lean_ctor_set(x_20, 0, x_29); -x_54 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; -x_55 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_55, 0, x_20); -lean_ctor_set(x_55, 1, x_54); -x_56 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_19, x_55, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_48); -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___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3(x_3, x_5, x_4, x_1, x_9, x_7, x_6, x_19, x_8, x_2, x_57, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_58); -lean_dec(x_57); -return x_59; +lean_ctor_set(x_20, 1, x_55); +lean_ctor_set(x_20, 0, x_31); +x_56 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; +x_57 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_57, 0, x_20); +lean_ctor_set(x_57, 1, x_56); +x_58 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_19, x_57, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_50); +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___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3(x_3, x_5, x_4, x_1, x_9, x_7, x_6, x_19, x_8, x_2, x_59, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_60); +lean_dec(x_59); +return x_61; } } else { -lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; 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_60 = lean_ctor_get(x_29, 0); -x_61 = lean_ctor_get(x_29, 1); -lean_inc(x_61); -lean_inc(x_60); -lean_dec(x_29); -x_62 = l_Lean_Meta_Grind_ppENodeRef(x_4, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_61); -x_63 = lean_ctor_get(x_62, 0); +lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; +x_62 = lean_ctor_get(x_31, 0); +x_63 = lean_ctor_get(x_31, 1); lean_inc(x_63); -x_64 = lean_ctor_get(x_62, 1); -lean_inc(x_64); -if (lean_is_exclusive(x_62)) { - lean_ctor_release(x_62, 0); - lean_ctor_release(x_62, 1); - x_65 = x_62; +lean_inc(x_62); +lean_dec(x_31); +x_64 = l_Lean_Meta_Grind_ppENodeRef(x_4, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, 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); +if (lean_is_exclusive(x_64)) { + lean_ctor_release(x_64, 0); + lean_ctor_release(x_64, 1); + x_67 = x_64; } else { - lean_dec_ref(x_62); - x_65 = lean_box(0); + lean_dec_ref(x_64); + x_67 = lean_box(0); } -x_66 = l_Lean_MessageData_ofFormat(x_60); -x_67 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__3; -if (lean_is_scalar(x_65)) { - x_68 = lean_alloc_ctor(7, 2, 0); +x_68 = l_Lean_MessageData_ofFormat(x_62); +x_69 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__3; +if (lean_is_scalar(x_67)) { + x_70 = lean_alloc_ctor(7, 2, 0); } else { - x_68 = x_65; - lean_ctor_set_tag(x_68, 7); -} -lean_ctor_set(x_68, 0, x_67); -lean_ctor_set(x_68, 1, x_66); -x_69 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__5; -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_ofFormat(x_63); + x_70 = x_67; + lean_ctor_set_tag(x_70, 7); +} +lean_ctor_set(x_70, 0, x_69); +lean_ctor_set(x_70, 1, x_68); +x_71 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__5; +x_72 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_72, 0, x_70); +lean_ctor_set(x_72, 1, x_71); +x_73 = l_Lean_MessageData_ofFormat(x_65); lean_ctor_set_tag(x_20, 7); -lean_ctor_set(x_20, 1, x_71); -lean_ctor_set(x_20, 0, x_70); -x_72 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; -x_73 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_73, 0, x_20); -lean_ctor_set(x_73, 1, x_72); -x_74 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_19, x_73, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_64); -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___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3(x_3, x_5, x_4, x_1, x_9, x_7, x_6, x_19, x_8, x_2, x_75, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_76); -lean_dec(x_75); -return x_77; +lean_ctor_set(x_20, 1, x_73); +lean_ctor_set(x_20, 0, x_72); +x_74 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; +x_75 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_75, 0, x_20); +lean_ctor_set(x_75, 1, x_74); +x_76 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_19, x_75, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_66); +x_77 = lean_ctor_get(x_76, 0); +lean_inc(x_77); +x_78 = lean_ctor_get(x_76, 1); +lean_inc(x_78); +lean_dec(x_76); +x_79 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3(x_3, x_5, x_4, x_1, x_9, x_7, x_6, x_19, x_8, x_2, x_77, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_78); +lean_dec(x_77); +return x_79; } } else { -lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; 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_78 = lean_ctor_get(x_20, 1); -lean_inc(x_78); -lean_dec(x_20); -x_79 = l_Lean_Meta_Grind_ppENodeRef(x_3, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_78); -x_80 = lean_ctor_get(x_79, 0); -lean_inc(x_80); -x_81 = lean_ctor_get(x_79, 1); -lean_inc(x_81); -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_79); - x_82 = lean_box(0); -} -x_83 = l_Lean_Meta_Grind_ppENodeRef(x_4, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_81); -x_84 = lean_ctor_get(x_83, 0); -lean_inc(x_84); -x_85 = lean_ctor_get(x_83, 1); -lean_inc(x_85); -if (lean_is_exclusive(x_83)) { - lean_ctor_release(x_83, 0); - lean_ctor_release(x_83, 1); - x_86 = x_83; -} else { - lean_dec_ref(x_83); - x_86 = lean_box(0); -} -x_87 = l_Lean_MessageData_ofFormat(x_80); -x_88 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__3; -if (lean_is_scalar(x_86)) { - x_89 = lean_alloc_ctor(7, 2, 0); -} else { - x_89 = x_86; - lean_ctor_set_tag(x_89, 7); -} -lean_ctor_set(x_89, 0, x_88); -lean_ctor_set(x_89, 1, x_87); -x_90 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__5; -if (lean_is_scalar(x_82)) { - x_91 = lean_alloc_ctor(7, 2, 0); -} else { - x_91 = x_82; - lean_ctor_set_tag(x_91, 7); -} -lean_ctor_set(x_91, 0, x_89); -lean_ctor_set(x_91, 1, x_90); -x_92 = l_Lean_MessageData_ofFormat(x_84); -x_93 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_93, 0, x_91); -lean_ctor_set(x_93, 1, x_92); -x_94 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; -x_95 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_95, 0, x_93); -lean_ctor_set(x_95, 1, x_94); -x_96 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_19, x_95, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_85); -x_97 = lean_ctor_get(x_96, 0); -lean_inc(x_97); -x_98 = lean_ctor_get(x_96, 1); -lean_inc(x_98); -lean_dec(x_96); -x_99 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3(x_3, x_5, x_4, x_1, x_9, x_7, x_6, x_19, x_8, x_2, x_97, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_98); -lean_dec(x_97); -return x_99; +uint8_t x_80; +lean_free_object(x_20); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_80 = !lean_is_exclusive(x_29); +if (x_80 == 0) +{ +return x_29; +} +else +{ +lean_object* x_81; lean_object* x_82; lean_object* x_83; +x_81 = lean_ctor_get(x_29, 0); +x_82 = lean_ctor_get(x_29, 1); +lean_inc(x_82); +lean_inc(x_81); +lean_dec(x_29); +x_83 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_83, 0, x_81); +lean_ctor_set(x_83, 1, x_82); +return x_83; +} +} +} +else +{ +lean_object* x_84; lean_object* x_85; +x_84 = lean_ctor_get(x_20, 1); +lean_inc(x_84); +lean_dec(x_20); +x_85 = l_Lean_Meta_Grind_updateLastTag(x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_84); +if (lean_obj_tag(x_85) == 0) +{ +lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; 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_86 = lean_ctor_get(x_85, 1); +lean_inc(x_86); +lean_dec(x_85); +x_87 = l_Lean_Meta_Grind_ppENodeRef(x_3, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_86); +x_88 = lean_ctor_get(x_87, 0); +lean_inc(x_88); +x_89 = lean_ctor_get(x_87, 1); +lean_inc(x_89); +if (lean_is_exclusive(x_87)) { + lean_ctor_release(x_87, 0); + lean_ctor_release(x_87, 1); + x_90 = x_87; +} else { + lean_dec_ref(x_87); + x_90 = lean_box(0); +} +x_91 = l_Lean_Meta_Grind_ppENodeRef(x_4, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_89); +x_92 = lean_ctor_get(x_91, 0); +lean_inc(x_92); +x_93 = lean_ctor_get(x_91, 1); +lean_inc(x_93); +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_91); + x_94 = lean_box(0); +} +x_95 = l_Lean_MessageData_ofFormat(x_88); +x_96 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__3; +if (lean_is_scalar(x_94)) { + x_97 = lean_alloc_ctor(7, 2, 0); +} else { + x_97 = x_94; + lean_ctor_set_tag(x_97, 7); +} +lean_ctor_set(x_97, 0, x_96); +lean_ctor_set(x_97, 1, x_95); +x_98 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__5; +if (lean_is_scalar(x_90)) { + x_99 = lean_alloc_ctor(7, 2, 0); +} else { + x_99 = x_90; + lean_ctor_set_tag(x_99, 7); +} +lean_ctor_set(x_99, 0, x_97); +lean_ctor_set(x_99, 1, x_98); +x_100 = l_Lean_MessageData_ofFormat(x_92); +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_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; +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_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_19, x_103, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_93); +x_105 = lean_ctor_get(x_104, 0); +lean_inc(x_105); +x_106 = lean_ctor_get(x_104, 1); +lean_inc(x_106); +lean_dec(x_104); +x_107 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3(x_3, x_5, x_4, x_1, x_9, x_7, x_6, x_19, x_8, x_2, x_105, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_106); +lean_dec(x_105); +return x_107; +} +else +{ +lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_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_108 = lean_ctor_get(x_85, 0); +lean_inc(x_108); +x_109 = lean_ctor_get(x_85, 1); +lean_inc(x_109); +if (lean_is_exclusive(x_85)) { + lean_ctor_release(x_85, 0); + lean_ctor_release(x_85, 1); + x_110 = x_85; +} else { + lean_dec_ref(x_85); + x_110 = lean_box(0); +} +if (lean_is_scalar(x_110)) { + x_111 = lean_alloc_ctor(1, 2, 0); +} else { + x_111 = x_110; +} +lean_ctor_set(x_111, 0, x_108); +lean_ctor_set(x_111, 1, x_109); +return x_111; +} } } } @@ -6395,7 +6844,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Gr { lean_object* x_11; lean_object* x_12; uint8_t x_13; x_11 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__1; -x_12 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_11, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_12 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_11, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); x_13 = !lean_is_exclusive(x_12); if (x_13 == 0) { @@ -6416,6 +6865,13 @@ return x_19; else { lean_object* x_20; +x_20 = l_Lean_Meta_Grind_updateLastTag(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_15); +if (lean_obj_tag(x_20) == 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); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); @@ -6424,36 +6880,36 @@ lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); -x_20 = l_Lean_Meta_Grind_ppState(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_15); -if (lean_obj_tag(x_20) == 0) +x_22 = l_Lean_Meta_Grind_ppState(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_21); +if (lean_obj_tag(x_22) == 0) { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_21 = lean_ctor_get(x_20, 0); -lean_inc(x_21); -x_22 = lean_ctor_get(x_20, 1); -lean_inc(x_22); -lean_dec(x_20); -x_23 = l_Lean_MessageData_ofFormat(x_21); -x_24 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__2___closed__3; +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +x_24 = lean_ctor_get(x_22, 1); +lean_inc(x_24); +lean_dec(x_22); +x_25 = l_Lean_MessageData_ofFormat(x_23); +x_26 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__2___closed__3; lean_ctor_set_tag(x_12, 7); -lean_ctor_set(x_12, 1, x_23); -lean_ctor_set(x_12, 0, x_24); -x_25 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; -x_26 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_26, 0, x_12); -lean_ctor_set(x_26, 1, x_25); -x_27 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_11, x_26, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_22); -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_apply_10(x_16, x_28, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_29); -return x_30; +lean_ctor_set(x_12, 1, x_25); +lean_ctor_set(x_12, 0, x_26); +x_27 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; +x_28 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_28, 0, x_12); +lean_ctor_set(x_28, 1, x_27); +x_29 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_11, x_28, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_24); +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_apply_10(x_16, x_30, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_31); +return x_32; } else { -uint8_t x_31; +uint8_t x_33; lean_free_object(x_12); lean_dec(x_9); lean_dec(x_8); @@ -6463,48 +6919,87 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_31 = !lean_is_exclusive(x_20); -if (x_31 == 0) +x_33 = !lean_is_exclusive(x_22); +if (x_33 == 0) +{ +return x_22; +} +else +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_22, 0); +x_35 = lean_ctor_get(x_22, 1); +lean_inc(x_35); +lean_inc(x_34); +lean_dec(x_22); +x_36 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_36, 0, x_34); +lean_ctor_set(x_36, 1, x_35); +return x_36; +} +} +} +else +{ +uint8_t x_37; +lean_free_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_2); +x_37 = !lean_is_exclusive(x_20); +if (x_37 == 0) { return x_20; } else { -lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_32 = lean_ctor_get(x_20, 0); -x_33 = lean_ctor_get(x_20, 1); -lean_inc(x_33); -lean_inc(x_32); +lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_38 = lean_ctor_get(x_20, 0); +x_39 = lean_ctor_get(x_20, 1); +lean_inc(x_39); +lean_inc(x_38); lean_dec(x_20); -x_34 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_34, 0, x_32); -lean_ctor_set(x_34, 1, x_33); -return x_34; +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_35; lean_object* x_36; lean_object* x_37; uint8_t x_38; -x_35 = lean_ctor_get(x_12, 0); -x_36 = lean_ctor_get(x_12, 1); -lean_inc(x_36); -lean_inc(x_35); +lean_object* x_41; lean_object* x_42; lean_object* x_43; uint8_t x_44; +x_41 = lean_ctor_get(x_12, 0); +x_42 = lean_ctor_get(x_12, 1); +lean_inc(x_42); +lean_inc(x_41); lean_dec(x_12); -x_37 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__2___closed__1; -x_38 = lean_unbox(x_35); -lean_dec(x_35); -if (x_38 == 0) +x_43 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__2___closed__1; +x_44 = lean_unbox(x_41); +lean_dec(x_41); +if (x_44 == 0) { -lean_object* x_39; lean_object* x_40; -x_39 = lean_box(0); -x_40 = lean_apply_10(x_37, x_39, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_36); -return x_40; +lean_object* x_45; lean_object* x_46; +x_45 = lean_box(0); +x_46 = lean_apply_10(x_43, x_45, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_42); +return x_46; } else { -lean_object* x_41; +lean_object* x_47; +x_47 = l_Lean_Meta_Grind_updateLastTag(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_42); +if (lean_obj_tag(x_47) == 0) +{ +lean_object* x_48; lean_object* x_49; +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); @@ -6513,36 +7008,69 @@ lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); -x_41 = l_Lean_Meta_Grind_ppState(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_36); -if (lean_obj_tag(x_41) == 0) +x_49 = l_Lean_Meta_Grind_ppState(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_48); +if (lean_obj_tag(x_49) == 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 = l_Lean_MessageData_ofFormat(x_42); -x_45 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__2___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_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; -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_Grind_addCongrTable___spec__9(x_11, x_48, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_43); +lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; x_50 = 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_apply_10(x_37, x_50, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_51); -return x_52; +x_52 = l_Lean_MessageData_ofFormat(x_50); +x_53 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__2___closed__3; +x_54 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_54, 0, x_53); +lean_ctor_set(x_54, 1, x_52); +x_55 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___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 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_11, x_56, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_51); +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_apply_10(x_43, x_58, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_59); +return x_60; +} +else +{ +lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; +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_ctor_get(x_49, 0); +lean_inc(x_61); +x_62 = lean_ctor_get(x_49, 1); +lean_inc(x_62); +if (lean_is_exclusive(x_49)) { + lean_ctor_release(x_49, 0); + lean_ctor_release(x_49, 1); + x_63 = x_49; +} else { + lean_dec_ref(x_49); + x_63 = lean_box(0); +} +if (lean_is_scalar(x_63)) { + x_64 = lean_alloc_ctor(1, 2, 0); +} else { + x_64 = x_63; +} +lean_ctor_set(x_64, 0, x_61); +lean_ctor_set(x_64, 1, x_62); +return x_64; +} } else { -lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; +lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -6551,26 +7079,26 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_53 = lean_ctor_get(x_41, 0); -lean_inc(x_53); -x_54 = lean_ctor_get(x_41, 1); -lean_inc(x_54); -if (lean_is_exclusive(x_41)) { - lean_ctor_release(x_41, 0); - lean_ctor_release(x_41, 1); - x_55 = x_41; +x_65 = lean_ctor_get(x_47, 0); +lean_inc(x_65); +x_66 = lean_ctor_get(x_47, 1); +lean_inc(x_66); +if (lean_is_exclusive(x_47)) { + lean_ctor_release(x_47, 0); + lean_ctor_release(x_47, 1); + x_67 = x_47; } else { - lean_dec_ref(x_41); - x_55 = lean_box(0); + lean_dec_ref(x_47); + x_67 = lean_box(0); } -if (lean_is_scalar(x_55)) { - x_56 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_67)) { + x_68 = lean_alloc_ctor(1, 2, 0); } else { - x_56 = x_55; + x_68 = x_67; } -lean_ctor_set(x_56, 0, x_53); -lean_ctor_set(x_56, 1, x_54); -return x_56; +lean_ctor_set(x_68, 0, x_65); +lean_ctor_set(x_68, 1, x_66); +return x_68; } } } @@ -7446,8 +7974,19 @@ return x_41; } else { -lean_object* x_42; lean_object* x_43; uint8_t x_44; uint8_t x_45; lean_object* x_46; lean_object* x_47; +lean_object* x_42; +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); x_42 = l_Lean_Meta_Grind_markAsInconsistent(x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_35); +if (lean_obj_tag(x_42) == 0) +{ +lean_object* x_43; uint8_t x_44; uint8_t x_45; lean_object* x_46; lean_object* x_47; x_43 = lean_ctor_get(x_42, 1); lean_inc(x_43); lean_dec(x_42); @@ -7457,27 +7996,114 @@ x_46 = lean_box(0); x_47 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__6(x_19, x_29, x_3, x_4, x_5, x_6, x_1, x_2, x_44, x_45, x_46, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_43); return x_47; } +else +{ +uint8_t x_48; +lean_dec(x_29); +lean_dec(x_19); +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_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_48 = !lean_is_exclusive(x_42); +if (x_48 == 0) +{ +return x_42; } else { -lean_object* x_48; lean_object* x_49; uint8_t x_50; uint8_t x_51; lean_object* x_52; lean_object* x_53; -lean_dec(x_21); -x_48 = l_Lean_Meta_Grind_markAsInconsistent(x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_35); -x_49 = lean_ctor_get(x_48, 1); +lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_49 = lean_ctor_get(x_42, 0); +x_50 = lean_ctor_get(x_42, 1); +lean_inc(x_50); lean_inc(x_49); -lean_dec(x_48); -x_50 = 1; -x_51 = 0; -x_52 = lean_box(0); -x_53 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__6(x_19, x_29, x_3, x_4, x_5, x_6, x_1, x_2, x_50, x_51, x_52, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_49); -return x_53; +lean_dec(x_42); +x_51 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_51, 0, x_49); +lean_ctor_set(x_51, 1, x_50); +return x_51; +} +} +} +} +else +{ +lean_object* x_52; +lean_dec(x_21); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +x_52 = l_Lean_Meta_Grind_markAsInconsistent(x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_35); +if (lean_obj_tag(x_52) == 0) +{ +lean_object* x_53; uint8_t 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); +lean_dec(x_52); +x_54 = 1; +x_55 = 0; +x_56 = lean_box(0); +x_57 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__6(x_19, x_29, x_3, x_4, x_5, x_6, x_1, x_2, x_54, x_55, x_56, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_53); +return x_57; +} +else +{ +uint8_t x_58; +lean_dec(x_29); +lean_dec(x_19); +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_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_58 = !lean_is_exclusive(x_52); +if (x_58 == 0) +{ +return x_52; +} +else +{ +lean_object* x_59; lean_object* x_60; lean_object* x_61; +x_59 = lean_ctor_get(x_52, 0); +x_60 = lean_ctor_get(x_52, 1); +lean_inc(x_60); +lean_inc(x_59); +lean_dec(x_52); +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_54; +uint8_t x_62; lean_dec(x_21); lean_dec(x_19); lean_dec(x_17); @@ -7494,29 +8120,29 @@ lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_54 = !lean_is_exclusive(x_22); -if (x_54 == 0) +x_62 = !lean_is_exclusive(x_22); +if (x_62 == 0) { return x_22; } else { -lean_object* x_55; lean_object* x_56; lean_object* x_57; -x_55 = lean_ctor_get(x_22, 0); -x_56 = lean_ctor_get(x_22, 1); -lean_inc(x_56); -lean_inc(x_55); +lean_object* x_63; lean_object* x_64; lean_object* x_65; +x_63 = lean_ctor_get(x_22, 0); +x_64 = lean_ctor_get(x_22, 1); +lean_inc(x_64); +lean_inc(x_63); lean_dec(x_22); -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; +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_58; +uint8_t x_66; lean_dec(x_17); lean_dec(x_15); lean_dec(x_14); @@ -7531,23 +8157,23 @@ lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_58 = !lean_is_exclusive(x_18); -if (x_58 == 0) +x_66 = !lean_is_exclusive(x_18); +if (x_66 == 0) { return x_18; } else { -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_object* x_67; lean_object* x_68; lean_object* x_69; +x_67 = lean_ctor_get(x_18, 0); +x_68 = lean_ctor_get(x_18, 1); +lean_inc(x_68); +lean_inc(x_67); 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; +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; } } } @@ -7575,7 +8201,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Gr { lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; x_17 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__8___closed__2; -x_18 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_17, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +x_18 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_17, x_8, x_9, 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); @@ -7592,53 +8218,50 @@ return x_23; } else { -if (x_4 == 0) -{ -uint8_t x_24; -x_24 = !lean_is_exclusive(x_18); -if (x_24 == 0) +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_18, 1); +lean_inc(x_24); +if (lean_is_exclusive(x_18)) { + lean_ctor_release(x_18, 0); + lean_ctor_release(x_18, 1); + x_25 = x_18; +} else { + lean_dec_ref(x_18); + x_25 = lean_box(0); +} +x_26 = l_Lean_Meta_Grind_updateLastTag(x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_24); +if (lean_obj_tag(x_26) == 0) { -lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_25 = lean_ctor_get(x_18, 1); -x_26 = lean_ctor_get(x_18, 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); +if (x_4 == 0) +{ +lean_object* x_39; lean_inc(x_15); lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); lean_inc(x_6); lean_inc(x_5); -x_27 = l_Lean_Meta_mkEq(x_5, x_6, x_12, x_13, x_14, x_15, x_25); -if (lean_obj_tag(x_27) == 0) +x_39 = l_Lean_Meta_mkEq(x_5, x_6, x_12, x_13, x_14, x_15, x_27); +if (lean_obj_tag(x_39) == 0) { -lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_28 = lean_ctor_get(x_27, 0); -lean_inc(x_28); -x_29 = lean_ctor_get(x_27, 1); -lean_inc(x_29); -lean_dec(x_27); -x_30 = l_Lean_MessageData_ofExpr(x_28); -x_31 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; -lean_ctor_set_tag(x_18, 7); -lean_ctor_set(x_18, 1, x_30); -lean_ctor_set(x_18, 0, x_31); -x_32 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_32, 0, x_18); -lean_ctor_set(x_32, 1, x_31); -x_33 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_17, x_32, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_29); -x_34 = lean_ctor_get(x_33, 0); -lean_inc(x_34); -x_35 = lean_ctor_get(x_33, 1); -lean_inc(x_35); -lean_dec(x_33); -x_36 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__7(x_1, x_2, x_3, x_4, x_5, x_6, x_34, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_35); -lean_dec(x_34); -return x_36; +lean_object* x_40; lean_object* x_41; +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_28 = x_40; +x_29 = x_41; +goto block_38; } else { -uint8_t x_37; -lean_free_object(x_18); +uint8_t x_42; +lean_dec(x_25); lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); @@ -7652,151 +8275,52 @@ lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_37 = !lean_is_exclusive(x_27); -if (x_37 == 0) -{ -return x_27; -} -else +x_42 = !lean_is_exclusive(x_39); +if (x_42 == 0) { -lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_38 = lean_ctor_get(x_27, 0); -x_39 = lean_ctor_get(x_27, 1); -lean_inc(x_39); -lean_inc(x_38); -lean_dec(x_27); -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; -} -} +return x_39; } else { -lean_object* x_41; lean_object* x_42; -x_41 = lean_ctor_get(x_18, 1); -lean_inc(x_41); -lean_dec(x_18); -lean_inc(x_15); -lean_inc(x_14); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_6); -lean_inc(x_5); -x_42 = l_Lean_Meta_mkEq(x_5, x_6, x_12, x_13, x_14, x_15, x_41); -if (lean_obj_tag(x_42) == 0) -{ -lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; -x_43 = lean_ctor_get(x_42, 0); -lean_inc(x_43); -x_44 = lean_ctor_get(x_42, 1); +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_inc(x_44); -lean_dec(x_42); -x_45 = l_Lean_MessageData_ofExpr(x_43); -x_46 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; -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 = 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_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_17, x_48, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_44); -x_50 = lean_ctor_get(x_49, 0); -lean_inc(x_50); -x_51 = lean_ctor_get(x_49, 1); -lean_inc(x_51); -lean_dec(x_49); -x_52 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__7(x_1, x_2, x_3, x_4, x_5, x_6, x_50, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_51); -lean_dec(x_50); -return x_52; -} -else -{ -lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; -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_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_53 = lean_ctor_get(x_42, 0); -lean_inc(x_53); -x_54 = lean_ctor_get(x_42, 1); -lean_inc(x_54); -if (lean_is_exclusive(x_42)) { - lean_ctor_release(x_42, 0); - lean_ctor_release(x_42, 1); - x_55 = x_42; -} else { - lean_dec_ref(x_42); - 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_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; } } } else { -uint8_t x_57; -x_57 = !lean_is_exclusive(x_18); -if (x_57 == 0) -{ -lean_object* x_58; lean_object* x_59; lean_object* x_60; -x_58 = lean_ctor_get(x_18, 1); -x_59 = lean_ctor_get(x_18, 0); -lean_dec(x_59); +lean_object* x_46; lean_inc(x_15); lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); lean_inc(x_6); lean_inc(x_5); -x_60 = l_Lean_Meta_mkHEq(x_5, x_6, x_12, x_13, x_14, x_15, x_58); -if (lean_obj_tag(x_60) == 0) +x_46 = l_Lean_Meta_mkHEq(x_5, x_6, x_12, x_13, x_14, x_15, x_27); +if (lean_obj_tag(x_46) == 0) { -lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; -x_61 = lean_ctor_get(x_60, 0); -lean_inc(x_61); -x_62 = lean_ctor_get(x_60, 1); -lean_inc(x_62); -lean_dec(x_60); -x_63 = l_Lean_MessageData_ofExpr(x_61); -x_64 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; -lean_ctor_set_tag(x_18, 7); -lean_ctor_set(x_18, 1, x_63); -lean_ctor_set(x_18, 0, x_64); -x_65 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_65, 0, x_18); -lean_ctor_set(x_65, 1, x_64); -x_66 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_17, x_65, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, 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 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__7(x_1, x_2, x_3, x_4, x_5, x_6, x_67, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_68); -lean_dec(x_67); -return x_69; +lean_object* x_47; lean_object* x_48; +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_28 = x_47; +x_29 = x_48; +goto block_38; } else { -uint8_t x_70; -lean_free_object(x_18); +uint8_t x_49; +lean_dec(x_25); lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); @@ -7810,68 +8334,57 @@ lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_70 = !lean_is_exclusive(x_60); -if (x_70 == 0) +x_49 = !lean_is_exclusive(x_46); +if (x_49 == 0) { -return x_60; +return x_46; } else { -lean_object* x_71; lean_object* x_72; lean_object* x_73; -x_71 = lean_ctor_get(x_60, 0); -x_72 = lean_ctor_get(x_60, 1); -lean_inc(x_72); -lean_inc(x_71); -lean_dec(x_60); -x_73 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_73, 0, x_71); -lean_ctor_set(x_73, 1, x_72); -return x_73; -} +lean_object* x_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_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 -{ -lean_object* x_74; lean_object* x_75; -x_74 = lean_ctor_get(x_18, 1); -lean_inc(x_74); -lean_dec(x_18); -lean_inc(x_15); -lean_inc(x_14); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_6); -lean_inc(x_5); -x_75 = l_Lean_Meta_mkHEq(x_5, x_6, x_12, x_13, x_14, x_15, x_74); -if (lean_obj_tag(x_75) == 0) +} +block_38: { -lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; -x_76 = lean_ctor_get(x_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_MessageData_ofExpr(x_76); -x_79 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; -x_80 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_80, 0, x_79); -lean_ctor_set(x_80, 1, x_78); -x_81 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_81, 0, x_80); -lean_ctor_set(x_81, 1, x_79); -x_82 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_17, x_81, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_77); -x_83 = lean_ctor_get(x_82, 0); -lean_inc(x_83); -x_84 = lean_ctor_get(x_82, 1); -lean_inc(x_84); -lean_dec(x_82); -x_85 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__7(x_1, x_2, x_3, x_4, x_5, x_6, x_83, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_84); -lean_dec(x_83); -return x_85; +lean_object* x_30; lean_object* x_31; 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_Lean_MessageData_ofExpr(x_28); +x_31 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; +if (lean_is_scalar(x_25)) { + x_32 = lean_alloc_ctor(7, 2, 0); +} else { + x_32 = x_25; + lean_ctor_set_tag(x_32, 7); +} +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_30); +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_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_17, x_33, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_29); +x_35 = lean_ctor_get(x_34, 0); +lean_inc(x_35); +x_36 = lean_ctor_get(x_34, 1); +lean_inc(x_36); +lean_dec(x_34); +x_37 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__7(x_1, x_2, x_3, x_4, x_5, x_6, x_35, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_36); +lean_dec(x_35); +return x_37; +} } else { -lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; +uint8_t x_53; +lean_dec(x_25); lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); @@ -7885,27 +8398,23 @@ lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_86 = lean_ctor_get(x_75, 0); -lean_inc(x_86); -x_87 = lean_ctor_get(x_75, 1); -lean_inc(x_87); -if (lean_is_exclusive(x_75)) { - lean_ctor_release(x_75, 0); - lean_ctor_release(x_75, 1); - x_88 = x_75; -} else { - lean_dec_ref(x_75); - x_88 = lean_box(0); -} -if (lean_is_scalar(x_88)) { - x_89 = lean_alloc_ctor(1, 2, 0); -} else { - x_89 = x_88; -} -lean_ctor_set(x_89, 0, x_86); -lean_ctor_set(x_89, 1, x_87); -return x_89; +x_53 = !lean_is_exclusive(x_26); +if (x_53 == 0) +{ +return x_26; } +else +{ +lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_54 = lean_ctor_get(x_26, 0); +x_55 = lean_ctor_get(x_26, 1); +lean_inc(x_55); +lean_inc(x_54); +lean_dec(x_26); +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; } } } @@ -8009,7 +8518,7 @@ lean_dec(x_18); lean_dec(x_15); lean_dec(x_3); x_25 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__1; -x_26 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_25, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_19); +x_26 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_25, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_19); x_27 = !lean_is_exclusive(x_26); if (x_27 == 0) { @@ -8031,232 +8540,315 @@ return x_33; } else { -lean_object* x_34; uint8_t x_35; -x_34 = l_Lean_Meta_Grind_ppENodeRef(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_29); +lean_object* x_34; +x_34 = l_Lean_Meta_Grind_updateLastTag(x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_29); +if (lean_obj_tag(x_34) == 0) +{ +lean_object* x_35; lean_object* x_36; uint8_t x_37; +x_35 = lean_ctor_get(x_34, 1); +lean_inc(x_35); +lean_dec(x_34); +x_36 = l_Lean_Meta_Grind_ppENodeRef(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_35); lean_dec(x_1); -x_35 = !lean_is_exclusive(x_34); -if (x_35 == 0) +x_37 = !lean_is_exclusive(x_36); +if (x_37 == 0) { -lean_object* x_36; lean_object* x_37; lean_object* x_38; uint8_t x_39; -x_36 = lean_ctor_get(x_34, 0); -x_37 = lean_ctor_get(x_34, 1); -x_38 = l_Lean_Meta_Grind_ppENodeRef(x_2, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_37); +lean_object* x_38; lean_object* x_39; lean_object* x_40; uint8_t x_41; +x_38 = lean_ctor_get(x_36, 0); +x_39 = lean_ctor_get(x_36, 1); +x_40 = l_Lean_Meta_Grind_ppENodeRef(x_2, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_39); lean_dec(x_2); -x_39 = !lean_is_exclusive(x_38); -if (x_39 == 0) +x_41 = !lean_is_exclusive(x_40); +if (x_41 == 0) { -lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; -x_40 = lean_ctor_get(x_38, 0); -x_41 = lean_ctor_get(x_38, 1); -x_42 = l_Lean_MessageData_ofFormat(x_36); -x_43 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; -lean_ctor_set_tag(x_38, 7); -lean_ctor_set(x_38, 1, x_42); -lean_ctor_set(x_38, 0, x_43); -x_44 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__3; -lean_ctor_set_tag(x_34, 7); -lean_ctor_set(x_34, 1, x_44); -lean_ctor_set(x_34, 0, x_38); -x_45 = l_Lean_MessageData_ofFormat(x_40); +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_42 = lean_ctor_get(x_40, 0); +x_43 = lean_ctor_get(x_40, 1); +x_44 = l_Lean_MessageData_ofFormat(x_38); +x_45 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; +lean_ctor_set_tag(x_40, 7); +lean_ctor_set(x_40, 1, x_44); +lean_ctor_set(x_40, 0, x_45); +x_46 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__3; +lean_ctor_set_tag(x_36, 7); +lean_ctor_set(x_36, 1, x_46); +lean_ctor_set(x_36, 0, x_40); +x_47 = l_Lean_MessageData_ofFormat(x_42); lean_ctor_set_tag(x_26, 7); -lean_ctor_set(x_26, 1, x_45); -lean_ctor_set(x_26, 0, x_34); -x_46 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__5; -x_47 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_47, 0, x_26); -lean_ctor_set(x_47, 1, x_46); -x_48 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_25, x_47, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_41); -x_49 = lean_ctor_get(x_48, 0); -lean_inc(x_49); -x_50 = lean_ctor_get(x_48, 1); -lean_inc(x_50); -lean_dec(x_48); -x_51 = lean_apply_10(x_30, x_49, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_50); -return x_51; +lean_ctor_set(x_26, 1, x_47); +lean_ctor_set(x_26, 0, x_36); +x_48 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__5; +x_49 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_49, 0, x_26); +lean_ctor_set(x_49, 1, x_48); +x_50 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_25, x_49, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_43); +x_51 = lean_ctor_get(x_50, 0); +lean_inc(x_51); +x_52 = lean_ctor_get(x_50, 1); +lean_inc(x_52); +lean_dec(x_50); +x_53 = lean_apply_10(x_30, x_51, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_52); +return x_53; } else { -lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; -x_52 = lean_ctor_get(x_38, 0); -x_53 = lean_ctor_get(x_38, 1); -lean_inc(x_53); -lean_inc(x_52); -lean_dec(x_38); -x_54 = l_Lean_MessageData_ofFormat(x_36); -x_55 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; -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___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__3; -lean_ctor_set_tag(x_34, 7); -lean_ctor_set(x_34, 1, x_57); -lean_ctor_set(x_34, 0, x_56); -x_58 = l_Lean_MessageData_ofFormat(x_52); +lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_54 = lean_ctor_get(x_40, 0); +x_55 = lean_ctor_get(x_40, 1); +lean_inc(x_55); +lean_inc(x_54); +lean_dec(x_40); +x_56 = l_Lean_MessageData_ofFormat(x_38); +x_57 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; +x_58 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_58, 0, x_57); +lean_ctor_set(x_58, 1, x_56); +x_59 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__3; +lean_ctor_set_tag(x_36, 7); +lean_ctor_set(x_36, 1, x_59); +lean_ctor_set(x_36, 0, x_58); +x_60 = l_Lean_MessageData_ofFormat(x_54); lean_ctor_set_tag(x_26, 7); -lean_ctor_set(x_26, 1, x_58); -lean_ctor_set(x_26, 0, x_34); -x_59 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__5; -x_60 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_60, 0, x_26); -lean_ctor_set(x_60, 1, x_59); -x_61 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_25, x_60, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_53); -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_apply_10(x_30, x_62, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_63); -return x_64; +lean_ctor_set(x_26, 1, x_60); +lean_ctor_set(x_26, 0, x_36); +x_61 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__5; +x_62 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_62, 0, x_26); +lean_ctor_set(x_62, 1, x_61); +x_63 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_25, x_62, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_55); +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_apply_10(x_30, x_64, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_65); +return x_66; } } else { -lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; -x_65 = lean_ctor_get(x_34, 0); -x_66 = lean_ctor_get(x_34, 1); -lean_inc(x_66); -lean_inc(x_65); -lean_dec(x_34); -x_67 = l_Lean_Meta_Grind_ppENodeRef(x_2, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_66); -lean_dec(x_2); -x_68 = lean_ctor_get(x_67, 0); +lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; 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_36, 0); +x_68 = lean_ctor_get(x_36, 1); lean_inc(x_68); -x_69 = lean_ctor_get(x_67, 1); -lean_inc(x_69); -if (lean_is_exclusive(x_67)) { - lean_ctor_release(x_67, 0); - lean_ctor_release(x_67, 1); - x_70 = x_67; +lean_inc(x_67); +lean_dec(x_36); +x_69 = l_Lean_Meta_Grind_ppENodeRef(x_2, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_68); +lean_dec(x_2); +x_70 = lean_ctor_get(x_69, 0); +lean_inc(x_70); +x_71 = lean_ctor_get(x_69, 1); +lean_inc(x_71); +if (lean_is_exclusive(x_69)) { + lean_ctor_release(x_69, 0); + lean_ctor_release(x_69, 1); + x_72 = x_69; } else { - lean_dec_ref(x_67); - x_70 = lean_box(0); + lean_dec_ref(x_69); + x_72 = lean_box(0); } -x_71 = l_Lean_MessageData_ofFormat(x_65); -x_72 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; -if (lean_is_scalar(x_70)) { - x_73 = lean_alloc_ctor(7, 2, 0); +x_73 = l_Lean_MessageData_ofFormat(x_67); +x_74 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; +if (lean_is_scalar(x_72)) { + x_75 = lean_alloc_ctor(7, 2, 0); } else { - x_73 = x_70; - lean_ctor_set_tag(x_73, 7); + x_75 = x_72; + lean_ctor_set_tag(x_75, 7); } -lean_ctor_set(x_73, 0, x_72); -lean_ctor_set(x_73, 1, x_71); -x_74 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__3; -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_MessageData_ofFormat(x_68); +lean_ctor_set(x_75, 0, x_74); +lean_ctor_set(x_75, 1, x_73); +x_76 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__3; +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_ofFormat(x_70); lean_ctor_set_tag(x_26, 7); -lean_ctor_set(x_26, 1, x_76); -lean_ctor_set(x_26, 0, x_75); -x_77 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__5; -x_78 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_78, 0, x_26); -lean_ctor_set(x_78, 1, x_77); -x_79 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_25, x_78, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_69); -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_apply_10(x_30, x_80, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_81); -return x_82; +lean_ctor_set(x_26, 1, x_78); +lean_ctor_set(x_26, 0, x_77); +x_79 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__5; +x_80 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_80, 0, x_26); +lean_ctor_set(x_80, 1, x_79); +x_81 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_25, x_80, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_71); +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_apply_10(x_30, x_82, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_83); +return x_84; +} +} +else +{ +uint8_t x_85; +lean_free_object(x_26); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_2); +lean_dec(x_1); +x_85 = !lean_is_exclusive(x_34); +if (x_85 == 0) +{ +return x_34; +} +else +{ +lean_object* x_86; lean_object* x_87; lean_object* x_88; +x_86 = lean_ctor_get(x_34, 0); +x_87 = lean_ctor_get(x_34, 1); +lean_inc(x_87); +lean_inc(x_86); +lean_dec(x_34); +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 { -lean_object* x_83; lean_object* x_84; lean_object* x_85; uint8_t x_86; -x_83 = lean_ctor_get(x_26, 0); -x_84 = lean_ctor_get(x_26, 1); -lean_inc(x_84); -lean_inc(x_83); +lean_object* x_89; lean_object* x_90; lean_object* x_91; uint8_t x_92; +x_89 = lean_ctor_get(x_26, 0); +x_90 = lean_ctor_get(x_26, 1); +lean_inc(x_90); +lean_inc(x_89); lean_dec(x_26); -x_85 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__1; -x_86 = lean_unbox(x_83); -lean_dec(x_83); -if (x_86 == 0) +x_91 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__1; +x_92 = lean_unbox(x_89); +lean_dec(x_89); +if (x_92 == 0) { -lean_object* x_87; lean_object* x_88; +lean_object* x_93; lean_object* x_94; lean_dec(x_2); lean_dec(x_1); -x_87 = lean_box(0); -x_88 = lean_apply_10(x_85, x_87, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_84); -return x_88; +x_93 = lean_box(0); +x_94 = lean_apply_10(x_91, x_93, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_90); +return x_94; } else { -lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; -x_89 = l_Lean_Meta_Grind_ppENodeRef(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_84); +lean_object* x_95; +x_95 = l_Lean_Meta_Grind_updateLastTag(x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_90); +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; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; +x_96 = lean_ctor_get(x_95, 1); +lean_inc(x_96); +lean_dec(x_95); +x_97 = l_Lean_Meta_Grind_ppENodeRef(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_96); lean_dec(x_1); -x_90 = lean_ctor_get(x_89, 0); -lean_inc(x_90); -x_91 = lean_ctor_get(x_89, 1); -lean_inc(x_91); -if (lean_is_exclusive(x_89)) { - lean_ctor_release(x_89, 0); - lean_ctor_release(x_89, 1); - x_92 = x_89; +x_98 = lean_ctor_get(x_97, 0); +lean_inc(x_98); +x_99 = lean_ctor_get(x_97, 1); +lean_inc(x_99); +if (lean_is_exclusive(x_97)) { + lean_ctor_release(x_97, 0); + lean_ctor_release(x_97, 1); + x_100 = x_97; } else { - lean_dec_ref(x_89); - x_92 = lean_box(0); + lean_dec_ref(x_97); + x_100 = lean_box(0); } -x_93 = l_Lean_Meta_Grind_ppENodeRef(x_2, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_91); +x_101 = l_Lean_Meta_Grind_ppENodeRef(x_2, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_99); lean_dec(x_2); -x_94 = lean_ctor_get(x_93, 0); -lean_inc(x_94); -x_95 = lean_ctor_get(x_93, 1); -lean_inc(x_95); -if (lean_is_exclusive(x_93)) { - lean_ctor_release(x_93, 0); - lean_ctor_release(x_93, 1); - x_96 = x_93; +x_102 = lean_ctor_get(x_101, 0); +lean_inc(x_102); +x_103 = lean_ctor_get(x_101, 1); +lean_inc(x_103); +if (lean_is_exclusive(x_101)) { + lean_ctor_release(x_101, 0); + lean_ctor_release(x_101, 1); + x_104 = x_101; } else { - lean_dec_ref(x_93); - x_96 = lean_box(0); + lean_dec_ref(x_101); + x_104 = lean_box(0); } -x_97 = l_Lean_MessageData_ofFormat(x_90); -x_98 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; -if (lean_is_scalar(x_96)) { - x_99 = lean_alloc_ctor(7, 2, 0); +x_105 = l_Lean_MessageData_ofFormat(x_98); +x_106 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; +if (lean_is_scalar(x_104)) { + x_107 = lean_alloc_ctor(7, 2, 0); } else { - x_99 = x_96; - lean_ctor_set_tag(x_99, 7); + x_107 = x_104; + lean_ctor_set_tag(x_107, 7); +} +lean_ctor_set(x_107, 0, x_106); +lean_ctor_set(x_107, 1, x_105); +x_108 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__3; +if (lean_is_scalar(x_100)) { + x_109 = lean_alloc_ctor(7, 2, 0); +} else { + x_109 = x_100; + lean_ctor_set_tag(x_109, 7); +} +lean_ctor_set(x_109, 0, x_107); +lean_ctor_set(x_109, 1, x_108); +x_110 = l_Lean_MessageData_ofFormat(x_102); +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 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__5; +x_113 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_113, 0, x_111); +lean_ctor_set(x_113, 1, x_112); +x_114 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_25, x_113, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_103); +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_apply_10(x_91, x_115, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_116); +return x_117; } -lean_ctor_set(x_99, 0, x_98); -lean_ctor_set(x_99, 1, x_97); -x_100 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__3; -if (lean_is_scalar(x_92)) { - x_101 = lean_alloc_ctor(7, 2, 0); +else +{ +lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_2); +lean_dec(x_1); +x_118 = lean_ctor_get(x_95, 0); +lean_inc(x_118); +x_119 = lean_ctor_get(x_95, 1); +lean_inc(x_119); +if (lean_is_exclusive(x_95)) { + lean_ctor_release(x_95, 0); + lean_ctor_release(x_95, 1); + x_120 = x_95; } else { - x_101 = x_92; - lean_ctor_set_tag(x_101, 7); + lean_dec_ref(x_95); + x_120 = lean_box(0); +} +if (lean_is_scalar(x_120)) { + x_121 = lean_alloc_ctor(1, 2, 0); +} else { + x_121 = x_120; +} +lean_ctor_set(x_121, 0, x_118); +lean_ctor_set(x_121, 1, x_119); +return x_121; } -lean_ctor_set(x_101, 0, x_99); -lean_ctor_set(x_101, 1, x_100); -x_102 = l_Lean_MessageData_ofFormat(x_94); -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___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__5; -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_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_25, x_105, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_95); -x_107 = lean_ctor_get(x_106, 0); -lean_inc(x_107); -x_108 = lean_ctor_get(x_106, 1); -lean_inc(x_108); -lean_dec(x_106); -x_109 = lean_apply_10(x_85, x_107, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_108); -return x_109; } } } } else { -uint8_t x_110; +uint8_t x_122; lean_dec(x_15); lean_dec(x_12); lean_dec(x_11); @@ -8269,29 +8861,29 @@ lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_110 = !lean_is_exclusive(x_17); -if (x_110 == 0) +x_122 = !lean_is_exclusive(x_17); +if (x_122 == 0) { return x_17; } else { -lean_object* x_111; lean_object* x_112; lean_object* x_113; -x_111 = lean_ctor_get(x_17, 0); -x_112 = lean_ctor_get(x_17, 1); -lean_inc(x_112); -lean_inc(x_111); +lean_object* x_123; lean_object* x_124; lean_object* x_125; +x_123 = lean_ctor_get(x_17, 0); +x_124 = lean_ctor_get(x_17, 1); +lean_inc(x_124); +lean_inc(x_123); lean_dec(x_17); -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_125 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_125, 0, x_123); +lean_ctor_set(x_125, 1, x_124); +return x_125; } } } else { -uint8_t x_114; +uint8_t x_126; lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); @@ -8303,23 +8895,23 @@ lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_114 = !lean_is_exclusive(x_14); -if (x_114 == 0) +x_126 = !lean_is_exclusive(x_14); +if (x_126 == 0) { return x_14; } else { -lean_object* x_115; lean_object* x_116; lean_object* x_117; -x_115 = lean_ctor_get(x_14, 0); -x_116 = lean_ctor_get(x_14, 1); -lean_inc(x_116); -lean_inc(x_115); +lean_object* x_127; lean_object* x_128; lean_object* x_129; +x_127 = lean_ctor_get(x_14, 0); +x_128 = lean_ctor_get(x_14, 1); +lean_inc(x_128); +lean_inc(x_127); lean_dec(x_14); -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; +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; } } } @@ -8514,13 +9106,13 @@ return x_22; } else { -lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; 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_23 = lean_ctor_get(x_11, 0); x_24 = lean_ctor_get(x_11, 1); x_25 = lean_ctor_get(x_11, 2); x_26 = lean_ctor_get(x_11, 3); x_27 = lean_ctor_get(x_11, 4); -x_28 = lean_ctor_get_uint8(x_11, sizeof(void*)*14); +x_28 = lean_ctor_get_uint8(x_11, sizeof(void*)*20); x_29 = lean_ctor_get(x_11, 6); x_30 = lean_ctor_get(x_11, 7); x_31 = lean_ctor_get(x_11, 8); @@ -8529,6 +9121,18 @@ x_33 = lean_ctor_get(x_11, 10); x_34 = lean_ctor_get(x_11, 11); x_35 = lean_ctor_get(x_11, 12); x_36 = lean_ctor_get(x_11, 13); +x_37 = lean_ctor_get(x_11, 14); +x_38 = lean_ctor_get(x_11, 15); +x_39 = lean_ctor_get(x_11, 16); +x_40 = lean_ctor_get(x_11, 17); +x_41 = lean_ctor_get(x_11, 18); +x_42 = lean_ctor_get(x_11, 19); +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_inc(x_34); @@ -8543,43 +9147,49 @@ lean_inc(x_25); lean_inc(x_24); lean_inc(x_23); lean_dec(x_11); -x_37 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_resetNewEqs___closed__1; -x_38 = lean_alloc_ctor(0, 14, 1); -lean_ctor_set(x_38, 0, x_23); -lean_ctor_set(x_38, 1, x_24); -lean_ctor_set(x_38, 2, x_25); -lean_ctor_set(x_38, 3, x_26); -lean_ctor_set(x_38, 4, x_27); -lean_ctor_set(x_38, 5, x_37); -lean_ctor_set(x_38, 6, x_29); -lean_ctor_set(x_38, 7, x_30); -lean_ctor_set(x_38, 8, x_31); -lean_ctor_set(x_38, 9, x_32); -lean_ctor_set(x_38, 10, x_33); -lean_ctor_set(x_38, 11, x_34); -lean_ctor_set(x_38, 12, x_35); -lean_ctor_set(x_38, 13, x_36); -lean_ctor_set_uint8(x_38, sizeof(void*)*14, x_28); -x_39 = lean_st_ref_set(x_1, x_38, x_12); -x_40 = lean_ctor_get(x_39, 1); -lean_inc(x_40); -if (lean_is_exclusive(x_39)) { - lean_ctor_release(x_39, 0); - lean_ctor_release(x_39, 1); - x_41 = x_39; +x_43 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_resetNewEqs___closed__1; +x_44 = lean_alloc_ctor(0, 20, 1); +lean_ctor_set(x_44, 0, x_23); +lean_ctor_set(x_44, 1, x_24); +lean_ctor_set(x_44, 2, x_25); +lean_ctor_set(x_44, 3, x_26); +lean_ctor_set(x_44, 4, x_27); +lean_ctor_set(x_44, 5, x_43); +lean_ctor_set(x_44, 6, x_29); +lean_ctor_set(x_44, 7, x_30); +lean_ctor_set(x_44, 8, x_31); +lean_ctor_set(x_44, 9, x_32); +lean_ctor_set(x_44, 10, x_33); +lean_ctor_set(x_44, 11, x_34); +lean_ctor_set(x_44, 12, x_35); +lean_ctor_set(x_44, 13, x_36); +lean_ctor_set(x_44, 14, x_37); +lean_ctor_set(x_44, 15, x_38); +lean_ctor_set(x_44, 16, x_39); +lean_ctor_set(x_44, 17, x_40); +lean_ctor_set(x_44, 18, x_41); +lean_ctor_set(x_44, 19, x_42); +lean_ctor_set_uint8(x_44, sizeof(void*)*20, x_28); +x_45 = lean_st_ref_set(x_1, x_44, x_12); +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_39); - x_41 = lean_box(0); + lean_dec_ref(x_45); + x_47 = lean_box(0); } -x_42 = lean_box(0); -if (lean_is_scalar(x_41)) { - x_43 = lean_alloc_ctor(0, 2, 0); +x_48 = lean_box(0); +if (lean_is_scalar(x_47)) { + x_49 = lean_alloc_ctor(0, 2, 0); } else { - x_43 = x_41; + x_49 = x_47; } -lean_ctor_set(x_43, 0, x_42); -lean_ctor_set(x_43, 1, x_40); -return x_43; +lean_ctor_set(x_49, 0, x_48); +lean_ctor_set(x_49, 1, x_46); +return x_49; } } } @@ -8671,14 +9281,14 @@ 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; 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; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; +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; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; 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_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); x_32 = lean_ctor_get(x_17, 5); -x_33 = lean_ctor_get_uint8(x_17, sizeof(void*)*14); +x_33 = lean_ctor_get_uint8(x_17, sizeof(void*)*20); x_34 = lean_ctor_get(x_17, 6); x_35 = lean_ctor_get(x_17, 7); x_36 = lean_ctor_get(x_17, 8); @@ -8687,6 +9297,18 @@ x_38 = lean_ctor_get(x_17, 10); x_39 = lean_ctor_get(x_17, 11); x_40 = lean_ctor_get(x_17, 12); x_41 = lean_ctor_get(x_17, 13); +x_42 = lean_ctor_get(x_17, 14); +x_43 = lean_ctor_get(x_17, 15); +x_44 = lean_ctor_get(x_17, 16); +x_45 = lean_ctor_get(x_17, 17); +x_46 = lean_ctor_get(x_17, 18); +x_47 = lean_ctor_get(x_17, 19); +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); @@ -8702,164 +9324,194 @@ lean_inc(x_29); lean_inc(x_28); lean_inc(x_27); lean_dec(x_17); -x_42 = lean_array_pop(x_32); -x_43 = lean_alloc_ctor(0, 14, 1); -lean_ctor_set(x_43, 0, x_27); -lean_ctor_set(x_43, 1, x_28); -lean_ctor_set(x_43, 2, x_29); -lean_ctor_set(x_43, 3, x_30); -lean_ctor_set(x_43, 4, x_31); -lean_ctor_set(x_43, 5, x_42); -lean_ctor_set(x_43, 6, x_34); -lean_ctor_set(x_43, 7, x_35); -lean_ctor_set(x_43, 8, x_36); -lean_ctor_set(x_43, 9, x_37); -lean_ctor_set(x_43, 10, x_38); -lean_ctor_set(x_43, 11, x_39); -lean_ctor_set(x_43, 12, x_40); -lean_ctor_set(x_43, 13, x_41); -lean_ctor_set_uint8(x_43, sizeof(void*)*14, x_33); -x_44 = lean_st_ref_set(x_1, x_43, x_18); -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; +x_48 = lean_array_pop(x_32); +x_49 = lean_alloc_ctor(0, 20, 1); +lean_ctor_set(x_49, 0, x_27); +lean_ctor_set(x_49, 1, x_28); +lean_ctor_set(x_49, 2, x_29); +lean_ctor_set(x_49, 3, x_30); +lean_ctor_set(x_49, 4, x_31); +lean_ctor_set(x_49, 5, x_48); +lean_ctor_set(x_49, 6, x_34); +lean_ctor_set(x_49, 7, x_35); +lean_ctor_set(x_49, 8, x_36); +lean_ctor_set(x_49, 9, x_37); +lean_ctor_set(x_49, 10, x_38); +lean_ctor_set(x_49, 11, x_39); +lean_ctor_set(x_49, 12, x_40); +lean_ctor_set(x_49, 13, x_41); +lean_ctor_set(x_49, 14, x_42); +lean_ctor_set(x_49, 15, x_43); +lean_ctor_set(x_49, 16, x_44); +lean_ctor_set(x_49, 17, x_45); +lean_ctor_set(x_49, 18, x_46); +lean_ctor_set(x_49, 19, x_47); +lean_ctor_set_uint8(x_49, sizeof(void*)*20, x_33); +x_50 = lean_st_ref_set(x_1, x_49, x_18); +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_44); - x_46 = lean_box(0); + lean_dec_ref(x_50); + x_52 = lean_box(0); } -if (lean_is_scalar(x_46)) { - x_47 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_52)) { + x_53 = lean_alloc_ctor(0, 2, 0); } else { - x_47 = x_46; + x_53 = x_52; } -lean_ctor_set(x_47, 0, x_15); -lean_ctor_set(x_47, 1, x_45); -return x_47; +lean_ctor_set(x_53, 0, x_15); +lean_ctor_set(x_53, 1, x_51); +return x_53; } } } else { -lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; -x_48 = lean_ctor_get(x_10, 0); -x_49 = lean_ctor_get(x_10, 1); -lean_inc(x_49); -lean_inc(x_48); +lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_54 = lean_ctor_get(x_10, 0); +x_55 = lean_ctor_get(x_10, 1); +lean_inc(x_55); +lean_inc(x_54); lean_dec(x_10); -x_50 = lean_ctor_get(x_48, 5); -lean_inc(x_50); -lean_dec(x_48); -x_51 = l_Array_back_x3f___rarg(x_50); -lean_dec(x_50); -if (lean_obj_tag(x_51) == 0) +x_56 = lean_ctor_get(x_54, 5); +lean_inc(x_56); +lean_dec(x_54); +x_57 = l_Array_back_x3f___rarg(x_56); +lean_dec(x_56); +if (lean_obj_tag(x_57) == 0) { -lean_object* x_52; -x_52 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_52, 0, x_51); -lean_ctor_set(x_52, 1, x_49); -return x_52; +lean_object* x_58; +x_58 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_58, 0, x_57); +lean_ctor_set(x_58, 1, x_55); +return x_58; } else { -lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; uint8_t x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; -x_53 = lean_st_ref_take(x_1, x_49); -x_54 = lean_ctor_get(x_53, 0); -lean_inc(x_54); -x_55 = lean_ctor_get(x_53, 1); -lean_inc(x_55); -lean_dec(x_53); -x_56 = lean_ctor_get(x_54, 0); -lean_inc(x_56); -x_57 = lean_ctor_get(x_54, 1); -lean_inc(x_57); -x_58 = lean_ctor_get(x_54, 2); -lean_inc(x_58); -x_59 = lean_ctor_get(x_54, 3); -lean_inc(x_59); -x_60 = lean_ctor_get(x_54, 4); +lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; uint8_t x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; 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_59 = lean_st_ref_take(x_1, x_55); +x_60 = lean_ctor_get(x_59, 0); lean_inc(x_60); -x_61 = lean_ctor_get(x_54, 5); +x_61 = lean_ctor_get(x_59, 1); lean_inc(x_61); -x_62 = lean_ctor_get_uint8(x_54, sizeof(void*)*14); -x_63 = lean_ctor_get(x_54, 6); +lean_dec(x_59); +x_62 = lean_ctor_get(x_60, 0); +lean_inc(x_62); +x_63 = lean_ctor_get(x_60, 1); lean_inc(x_63); -x_64 = lean_ctor_get(x_54, 7); +x_64 = lean_ctor_get(x_60, 2); lean_inc(x_64); -x_65 = lean_ctor_get(x_54, 8); +x_65 = lean_ctor_get(x_60, 3); lean_inc(x_65); -x_66 = lean_ctor_get(x_54, 9); +x_66 = lean_ctor_get(x_60, 4); lean_inc(x_66); -x_67 = lean_ctor_get(x_54, 10); +x_67 = lean_ctor_get(x_60, 5); lean_inc(x_67); -x_68 = lean_ctor_get(x_54, 11); -lean_inc(x_68); -x_69 = lean_ctor_get(x_54, 12); +x_68 = lean_ctor_get_uint8(x_60, sizeof(void*)*20); +x_69 = lean_ctor_get(x_60, 6); lean_inc(x_69); -x_70 = lean_ctor_get(x_54, 13); +x_70 = lean_ctor_get(x_60, 7); lean_inc(x_70); -if (lean_is_exclusive(x_54)) { - lean_ctor_release(x_54, 0); - lean_ctor_release(x_54, 1); - lean_ctor_release(x_54, 2); - lean_ctor_release(x_54, 3); - lean_ctor_release(x_54, 4); - lean_ctor_release(x_54, 5); - lean_ctor_release(x_54, 6); - lean_ctor_release(x_54, 7); - lean_ctor_release(x_54, 8); - lean_ctor_release(x_54, 9); - lean_ctor_release(x_54, 10); - lean_ctor_release(x_54, 11); - lean_ctor_release(x_54, 12); - lean_ctor_release(x_54, 13); - x_71 = x_54; +x_71 = lean_ctor_get(x_60, 8); +lean_inc(x_71); +x_72 = lean_ctor_get(x_60, 9); +lean_inc(x_72); +x_73 = lean_ctor_get(x_60, 10); +lean_inc(x_73); +x_74 = lean_ctor_get(x_60, 11); +lean_inc(x_74); +x_75 = lean_ctor_get(x_60, 12); +lean_inc(x_75); +x_76 = lean_ctor_get(x_60, 13); +lean_inc(x_76); +x_77 = lean_ctor_get(x_60, 14); +lean_inc(x_77); +x_78 = lean_ctor_get(x_60, 15); +lean_inc(x_78); +x_79 = lean_ctor_get(x_60, 16); +lean_inc(x_79); +x_80 = lean_ctor_get(x_60, 17); +lean_inc(x_80); +x_81 = lean_ctor_get(x_60, 18); +lean_inc(x_81); +x_82 = lean_ctor_get(x_60, 19); +lean_inc(x_82); +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); + lean_ctor_release(x_60, 5); + lean_ctor_release(x_60, 6); + lean_ctor_release(x_60, 7); + lean_ctor_release(x_60, 8); + lean_ctor_release(x_60, 9); + lean_ctor_release(x_60, 10); + lean_ctor_release(x_60, 11); + lean_ctor_release(x_60, 12); + lean_ctor_release(x_60, 13); + lean_ctor_release(x_60, 14); + lean_ctor_release(x_60, 15); + lean_ctor_release(x_60, 16); + lean_ctor_release(x_60, 17); + lean_ctor_release(x_60, 18); + lean_ctor_release(x_60, 19); + x_83 = x_60; } else { - lean_dec_ref(x_54); - x_71 = lean_box(0); + lean_dec_ref(x_60); + x_83 = lean_box(0); } -x_72 = lean_array_pop(x_61); -if (lean_is_scalar(x_71)) { - x_73 = lean_alloc_ctor(0, 14, 1); +x_84 = lean_array_pop(x_67); +if (lean_is_scalar(x_83)) { + x_85 = lean_alloc_ctor(0, 20, 1); } else { - x_73 = x_71; -} -lean_ctor_set(x_73, 0, x_56); -lean_ctor_set(x_73, 1, x_57); -lean_ctor_set(x_73, 2, x_58); -lean_ctor_set(x_73, 3, x_59); -lean_ctor_set(x_73, 4, x_60); -lean_ctor_set(x_73, 5, x_72); -lean_ctor_set(x_73, 6, x_63); -lean_ctor_set(x_73, 7, x_64); -lean_ctor_set(x_73, 8, x_65); -lean_ctor_set(x_73, 9, x_66); -lean_ctor_set(x_73, 10, x_67); -lean_ctor_set(x_73, 11, x_68); -lean_ctor_set(x_73, 12, x_69); -lean_ctor_set(x_73, 13, x_70); -lean_ctor_set_uint8(x_73, sizeof(void*)*14, x_62); -x_74 = lean_st_ref_set(x_1, x_73, x_55); -x_75 = lean_ctor_get(x_74, 1); -lean_inc(x_75); -if (lean_is_exclusive(x_74)) { - lean_ctor_release(x_74, 0); - lean_ctor_release(x_74, 1); - x_76 = x_74; + x_85 = x_83; +} +lean_ctor_set(x_85, 0, x_62); +lean_ctor_set(x_85, 1, x_63); +lean_ctor_set(x_85, 2, x_64); +lean_ctor_set(x_85, 3, x_65); +lean_ctor_set(x_85, 4, x_66); +lean_ctor_set(x_85, 5, x_84); +lean_ctor_set(x_85, 6, x_69); +lean_ctor_set(x_85, 7, x_70); +lean_ctor_set(x_85, 8, x_71); +lean_ctor_set(x_85, 9, x_72); +lean_ctor_set(x_85, 10, x_73); +lean_ctor_set(x_85, 11, x_74); +lean_ctor_set(x_85, 12, x_75); +lean_ctor_set(x_85, 13, x_76); +lean_ctor_set(x_85, 14, x_77); +lean_ctor_set(x_85, 15, x_78); +lean_ctor_set(x_85, 16, x_79); +lean_ctor_set(x_85, 17, x_80); +lean_ctor_set(x_85, 18, x_81); +lean_ctor_set(x_85, 19, x_82); +lean_ctor_set_uint8(x_85, sizeof(void*)*20, x_68); +x_86 = lean_st_ref_set(x_1, x_85, x_61); +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_74); - x_76 = lean_box(0); + lean_dec_ref(x_86); + x_88 = lean_box(0); } -if (lean_is_scalar(x_76)) { - x_77 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_88)) { + x_89 = lean_alloc_ctor(0, 2, 0); } else { - x_77 = x_76; + x_89 = x_88; } -lean_ctor_set(x_77, 0, x_51); -lean_ctor_set(x_77, 1, x_75); -return x_77; +lean_ctor_set(x_89, 0, x_57); +lean_ctor_set(x_89, 1, x_87); +return x_89; } } } @@ -9204,6 +9856,117 @@ x_14 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqCore(x_1, return x_14; } } +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addNewEq(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, 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); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_1); +x_14 = l_Lean_Meta_Grind_internalize(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; lean_object* x_16; +x_15 = lean_ctor_get(x_14, 1); +lean_inc(x_15); +lean_dec(x_14); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_2); +x_16 = l_Lean_Meta_Grind_internalize(x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_15); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_17; uint8_t x_18; lean_object* x_19; +x_17 = lean_ctor_get(x_16, 1); +lean_inc(x_17); +lean_dec(x_16); +x_18 = 0; +x_19 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqCore(x_1, x_2, x_3, x_18, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_17); +return x_19; +} +else +{ +uint8_t x_20; +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_20 = !lean_is_exclusive(x_16); +if (x_20 == 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_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; +} +} +} +else +{ +uint8_t 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_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_24 = !lean_is_exclusive(x_14); +if (x_24 == 0) +{ +return x_14; +} +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_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_EXPORT lean_object* l_Lean_Meta_Grind_add_goEq(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, uint8_t x_6, uint8_t x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, 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: { @@ -9978,7 +10741,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Grind_add(lean_object* x_1, lean_object* x_ { lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; x_13 = l_Lean_Meta_Grind_add___closed__2; -x_14 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +x_14 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); x_15 = lean_ctor_get(x_14, 0); lean_inc(x_15); x_16 = lean_unbox(x_15); @@ -9999,53 +10762,138 @@ uint8_t x_20; x_20 = !lean_is_exclusive(x_14); if (x_20 == 0) { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +lean_object* x_21; lean_object* x_22; lean_object* x_23; x_21 = lean_ctor_get(x_14, 1); x_22 = lean_ctor_get(x_14, 0); lean_dec(x_22); +x_23 = l_Lean_Meta_Grind_updateLastTag(x_4, x_5, 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; 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_23, 1); +lean_inc(x_24); +lean_dec(x_23); lean_inc(x_1); -x_23 = l_Lean_MessageData_ofExpr(x_1); -x_24 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; +x_25 = l_Lean_MessageData_ofExpr(x_1); +x_26 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; lean_ctor_set_tag(x_14, 7); -lean_ctor_set(x_14, 1, x_23); -lean_ctor_set(x_14, 0, x_24); -x_25 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_25, 0, x_14); -lean_ctor_set(x_25, 1, x_24); -x_26 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_13, x_25, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_21); -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_Grind_add___lambda__4(x_1, x_2, x_3, x_27, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_28); -lean_dec(x_27); -return x_29; +lean_ctor_set(x_14, 1, x_25); +lean_ctor_set(x_14, 0, x_26); +x_27 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_27, 0, x_14); +lean_ctor_set(x_27, 1, x_26); +x_28 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_13, x_27, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, 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 = l_Lean_Meta_Grind_add___lambda__4(x_1, x_2, x_3, x_29, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_30); +lean_dec(x_29); +return x_31; } else { -lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; -x_30 = lean_ctor_get(x_14, 1); -lean_inc(x_30); +uint8_t x_32; +lean_free_object(x_14); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_32 = !lean_is_exclusive(x_23); +if (x_32 == 0) +{ +return x_23; +} +else +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_33 = lean_ctor_get(x_23, 0); +x_34 = lean_ctor_get(x_23, 1); +lean_inc(x_34); +lean_inc(x_33); +lean_dec(x_23); +x_35 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_35, 0, x_33); +lean_ctor_set(x_35, 1, x_34); +return x_35; +} +} +} +else +{ +lean_object* x_36; lean_object* x_37; +x_36 = lean_ctor_get(x_14, 1); +lean_inc(x_36); lean_dec(x_14); +x_37 = l_Lean_Meta_Grind_updateLastTag(x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_36); +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, 1); +lean_inc(x_38); +lean_dec(x_37); lean_inc(x_1); -x_31 = l_Lean_MessageData_ofExpr(x_1); -x_32 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; -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 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_34, 0, x_33); -lean_ctor_set(x_34, 1, x_32); -x_35 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_13, x_34, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, 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_Grind_add___lambda__4(x_1, x_2, x_3, x_36, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_37); -lean_dec(x_36); -return x_38; +x_39 = l_Lean_MessageData_ofExpr(x_1); +x_40 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; +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 = 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_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_13, x_42, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_38); +x_44 = lean_ctor_get(x_43, 0); +lean_inc(x_44); +x_45 = lean_ctor_get(x_43, 1); +lean_inc(x_45); +lean_dec(x_43); +x_46 = l_Lean_Meta_Grind_add___lambda__4(x_1, x_2, x_3, x_44, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_45); +lean_dec(x_44); +return x_46; +} +else +{ +lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* 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_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_47 = lean_ctor_get(x_37, 0); +lean_inc(x_47); +x_48 = lean_ctor_get(x_37, 1); +lean_inc(x_48); +if (lean_is_exclusive(x_37)) { + lean_ctor_release(x_37, 0); + lean_ctor_release(x_37, 1); + x_49 = x_37; +} else { + lean_dec_ref(x_37); + x_49 = lean_box(0); +} +if (lean_is_scalar(x_49)) { + x_50 = lean_alloc_ctor(1, 2, 0); +} else { + x_50 = x_49; +} +lean_ctor_set(x_50, 0, x_47); +lean_ctor_set(x_50, 1, x_48); +return x_50; +} } } } diff --git a/stage0/stdlib/Lean/Meta/Tactic/Grind/Ctor.c b/stage0/stdlib/Lean/Meta/Tactic/Grind/Ctor.c index 45be339f7150..89db94455359 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Grind/Ctor.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Grind/Ctor.c @@ -20,14 +20,11 @@ static lean_object* l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_p lean_object* l_Lean_ConstantInfo_type(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_isExprDefEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint64_t lean_uint64_lor(uint64_t, uint64_t); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___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_Meta_Grind_propagateCtor(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_isApp(lean_object*); -lean_object* l_Lean_PersistentArray_push___rarg(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___closed__3; -static lean_object* l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___spec__2___closed__2; lean_object* lean_mk_array(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__5___closed__6; static lean_object* l_Lean_Meta_Grind_propagateCtor___lambda__3___closed__2; @@ -42,7 +39,6 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateCtor___lambda__3___boxed(lea LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_getFalseExpr___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___closed__1; -lean_object* lean_st_ref_take(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateCtor___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 lean_expr_eqv(lean_object*, lean_object*); uint64_t lean_uint64_shift_right(uint64_t, uint64_t); @@ -53,27 +49,26 @@ lean_object* l_Lean_Expr_appArg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___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_Lean_Meta_Grind_propagateCtor___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 double l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___spec__2___closed__1; lean_object* l_Lean_Expr_appFnCleanup(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___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___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__5___closed__7; lean_object* l_Lean_Meta_whnfD(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getConstInfo___at_Lean_Meta_Grind_propagateCtor___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_Name_str___override(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___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_getConstInfo___at_Lean_Meta_Grind_propagateCtor___spec__1___closed__1; -static lean_object* l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___spec__2___closed__3; +lean_object* l_Lean_Meta_Grind_updateLastTag(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofConstName(lean_object*, uint8_t); uint8_t l_Lean_Environment_contains(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___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___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__5___closed__1; -double l_Float_ofScientific(lean_object*, uint8_t, lean_object*); +lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___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_throwError___at_Lean_Meta_Grind_propagateCtor___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_grind_mk_eq_proof(lean_object*, 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_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__3___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs(lean_object*, lean_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_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___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_Grind_propagateCtor___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_Grind_shareCommon(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_Grind_updateLastTag___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateCtor___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* l_Lean_indentExpr(lean_object*); @@ -88,12 +83,10 @@ uint64_t lean_uint64_shift_left(uint64_t, uint64_t); static lean_object* l_Lean_getConstInfo___at_Lean_Meta_Grind_propagateCtor___spec__1___closed__3; lean_object* l_Lean_Meta_instantiateMVarsIfMVarApp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__3___closed__1; -lean_object* lean_array_mk(lean_object*); LEAN_EXPORT lean_object* l_Lean_getConstInfo___at_Lean_Meta_Grind_propagateCtor___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_isTracingEnabledForCore(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__5___closed__8; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___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_Meta_Grind_propagateCtor___lambda__3___closed__1; -lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateCtor___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_getConstInfo___at_Lean_Meta_Grind_propagateCtor___spec__1___closed__2; static lean_object* l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___closed__2; @@ -102,394 +95,46 @@ uint64_t l_Lean_Meta_TransparencyMode_toUInt64(uint8_t); static lean_object* l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___closed__4; static lean_object* l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__5___closed__2; static uint64_t l_Lean_Meta_Grind_propagateCtor___closed__1; -LEAN_EXPORT lean_object* l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___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_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_11; lean_object* x_12; uint8_t x_13; -x_11 = lean_ctor_get(x_8, 2); -x_12 = l_Lean_isTracingEnabledForCore(x_1, x_11, x_10); -x_13 = !lean_is_exclusive(x_12); -if (x_13 == 0) -{ -return x_12; -} -else -{ -lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_14 = lean_ctor_get(x_12, 0); -x_15 = lean_ctor_get(x_12, 1); -lean_inc(x_15); -lean_inc(x_14); -lean_dec(x_12); -x_16 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_16, 0, x_14); -lean_ctor_set(x_16, 1, x_15); -return x_16; -} -} -} -static double _init_l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___spec__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_addTrace___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___spec__2___closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("", 0, 0); -return x_1; -} -} -static lean_object* _init_l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___spec__2___closed__3() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = lean_box(0); -x_2 = lean_array_mk(x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___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___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -lean_object* x_12; 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_9, 5); -x_13 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_2, x_7, x_8, x_9, x_10, x_11); +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; lean_object* x_20; +x_13 = l_Lean_Meta_Grind_shareCommon(x_2, 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_ctor_get(x_13, 1); lean_inc(x_15); lean_dec(x_13); -x_16 = lean_st_ref_take(x_10, x_15); +x_16 = l_Lean_Meta_Grind_shareCommon(x_3, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_15); x_17 = lean_ctor_get(x_16, 0); lean_inc(x_17); -x_18 = lean_ctor_get(x_17, 3); +x_18 = lean_ctor_get(x_16, 1); lean_inc(x_18); -x_19 = !lean_is_exclusive(x_16); -if (x_19 == 0) -{ -lean_object* x_20; lean_object* x_21; uint8_t x_22; -x_20 = lean_ctor_get(x_16, 1); -x_21 = lean_ctor_get(x_16, 0); -lean_dec(x_21); -x_22 = !lean_is_exclusive(x_17); -if (x_22 == 0) -{ -lean_object* x_23; uint8_t x_24; -x_23 = lean_ctor_get(x_17, 3); -lean_dec(x_23); -x_24 = !lean_is_exclusive(x_18); -if (x_24 == 0) -{ -lean_object* x_25; double 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; uint8_t x_34; -x_25 = lean_ctor_get(x_18, 0); -x_26 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___spec__2___closed__1; -x_27 = 0; -x_28 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___spec__2___closed__2; -x_29 = lean_alloc_ctor(0, 2, 17); -lean_ctor_set(x_29, 0, x_1); -lean_ctor_set(x_29, 1, x_28); -lean_ctor_set_float(x_29, sizeof(void*)*2, x_26); -lean_ctor_set_float(x_29, sizeof(void*)*2 + 8, x_26); -lean_ctor_set_uint8(x_29, sizeof(void*)*2 + 16, x_27); -x_30 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___spec__2___closed__3; -x_31 = lean_alloc_ctor(9, 3, 0); -lean_ctor_set(x_31, 0, x_29); -lean_ctor_set(x_31, 1, x_14); -lean_ctor_set(x_31, 2, x_30); -lean_inc(x_12); -lean_ctor_set(x_16, 1, x_31); -lean_ctor_set(x_16, 0, x_12); -x_32 = l_Lean_PersistentArray_push___rarg(x_25, x_16); -lean_ctor_set(x_18, 0, x_32); -x_33 = lean_st_ref_set(x_10, x_17, x_20); -x_34 = !lean_is_exclusive(x_33); -if (x_34 == 0) -{ -lean_object* x_35; lean_object* x_36; -x_35 = lean_ctor_get(x_33, 0); -lean_dec(x_35); -x_36 = lean_box(0); -lean_ctor_set(x_33, 0, x_36); -return x_33; -} -else -{ -lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_37 = lean_ctor_get(x_33, 1); -lean_inc(x_37); -lean_dec(x_33); -x_38 = lean_box(0); -x_39 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_39, 0, x_38); -lean_ctor_set(x_39, 1, x_37); -return x_39; -} -} -else -{ -uint64_t x_40; lean_object* x_41; double 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; -x_40 = lean_ctor_get_uint64(x_18, sizeof(void*)*1); -x_41 = lean_ctor_get(x_18, 0); -lean_inc(x_41); -lean_dec(x_18); -x_42 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___spec__2___closed__1; -x_43 = 0; -x_44 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___spec__2___closed__2; -x_45 = lean_alloc_ctor(0, 2, 17); -lean_ctor_set(x_45, 0, x_1); -lean_ctor_set(x_45, 1, x_44); -lean_ctor_set_float(x_45, sizeof(void*)*2, x_42); -lean_ctor_set_float(x_45, sizeof(void*)*2 + 8, x_42); -lean_ctor_set_uint8(x_45, sizeof(void*)*2 + 16, x_43); -x_46 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___spec__2___closed__3; -x_47 = lean_alloc_ctor(9, 3, 0); -lean_ctor_set(x_47, 0, x_45); -lean_ctor_set(x_47, 1, x_14); -lean_ctor_set(x_47, 2, x_46); -lean_inc(x_12); -lean_ctor_set(x_16, 1, x_47); -lean_ctor_set(x_16, 0, x_12); -x_48 = l_Lean_PersistentArray_push___rarg(x_41, x_16); -x_49 = lean_alloc_ctor(0, 1, 8); -lean_ctor_set(x_49, 0, x_48); -lean_ctor_set_uint64(x_49, sizeof(void*)*1, x_40); -lean_ctor_set(x_17, 3, x_49); -x_50 = lean_st_ref_set(x_10, x_17, x_20); -x_51 = lean_ctor_get(x_50, 1); -lean_inc(x_51); -if (lean_is_exclusive(x_50)) { - lean_ctor_release(x_50, 0); - lean_ctor_release(x_50, 1); - x_52 = x_50; -} else { - lean_dec_ref(x_50); - x_52 = lean_box(0); -} -x_53 = lean_box(0); -if (lean_is_scalar(x_52)) { - x_54 = lean_alloc_ctor(0, 2, 0); -} else { - x_54 = x_52; -} -lean_ctor_set(x_54, 0, x_53); -lean_ctor_set(x_54, 1, x_51); -return x_54; -} -} -else -{ -lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; uint64_t x_62; lean_object* x_63; lean_object* x_64; double x_65; uint8_t x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; -x_55 = lean_ctor_get(x_17, 0); -x_56 = lean_ctor_get(x_17, 1); -x_57 = lean_ctor_get(x_17, 2); -x_58 = lean_ctor_get(x_17, 4); -x_59 = lean_ctor_get(x_17, 5); -x_60 = lean_ctor_get(x_17, 6); -x_61 = lean_ctor_get(x_17, 7); -lean_inc(x_61); -lean_inc(x_60); -lean_inc(x_59); -lean_inc(x_58); -lean_inc(x_57); -lean_inc(x_56); -lean_inc(x_55); -lean_dec(x_17); -x_62 = lean_ctor_get_uint64(x_18, sizeof(void*)*1); -x_63 = lean_ctor_get(x_18, 0); -lean_inc(x_63); -if (lean_is_exclusive(x_18)) { - lean_ctor_release(x_18, 0); - x_64 = x_18; -} else { - lean_dec_ref(x_18); - x_64 = lean_box(0); -} -x_65 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___spec__2___closed__1; -x_66 = 0; -x_67 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___spec__2___closed__2; -x_68 = lean_alloc_ctor(0, 2, 17); -lean_ctor_set(x_68, 0, x_1); -lean_ctor_set(x_68, 1, x_67); -lean_ctor_set_float(x_68, sizeof(void*)*2, x_65); -lean_ctor_set_float(x_68, sizeof(void*)*2 + 8, x_65); -lean_ctor_set_uint8(x_68, sizeof(void*)*2 + 16, x_66); -x_69 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___spec__2___closed__3; -x_70 = lean_alloc_ctor(9, 3, 0); -lean_ctor_set(x_70, 0, x_68); -lean_ctor_set(x_70, 1, x_14); -lean_ctor_set(x_70, 2, x_69); -lean_inc(x_12); -lean_ctor_set(x_16, 1, x_70); -lean_ctor_set(x_16, 0, x_12); -x_71 = l_Lean_PersistentArray_push___rarg(x_63, x_16); -if (lean_is_scalar(x_64)) { - x_72 = lean_alloc_ctor(0, 1, 8); -} else { - x_72 = x_64; -} -lean_ctor_set(x_72, 0, x_71); -lean_ctor_set_uint64(x_72, sizeof(void*)*1, x_62); -x_73 = lean_alloc_ctor(0, 8, 0); -lean_ctor_set(x_73, 0, x_55); -lean_ctor_set(x_73, 1, x_56); -lean_ctor_set(x_73, 2, x_57); -lean_ctor_set(x_73, 3, x_72); -lean_ctor_set(x_73, 4, x_58); -lean_ctor_set(x_73, 5, x_59); -lean_ctor_set(x_73, 6, x_60); -lean_ctor_set(x_73, 7, x_61); -x_74 = lean_st_ref_set(x_10, x_73, x_20); -x_75 = lean_ctor_get(x_74, 1); -lean_inc(x_75); -if (lean_is_exclusive(x_74)) { - lean_ctor_release(x_74, 0); - lean_ctor_release(x_74, 1); - x_76 = x_74; -} else { - lean_dec_ref(x_74); - x_76 = lean_box(0); -} -x_77 = lean_box(0); -if (lean_is_scalar(x_76)) { - x_78 = lean_alloc_ctor(0, 2, 0); -} else { - x_78 = x_76; -} -lean_ctor_set(x_78, 0, x_77); -lean_ctor_set(x_78, 1, x_75); -return x_78; -} -} -else -{ -lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; uint64_t x_88; lean_object* x_89; lean_object* x_90; double x_91; uint8_t x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; -x_79 = lean_ctor_get(x_16, 1); -lean_inc(x_79); lean_dec(x_16); -x_80 = lean_ctor_get(x_17, 0); -lean_inc(x_80); -x_81 = lean_ctor_get(x_17, 1); -lean_inc(x_81); -x_82 = lean_ctor_get(x_17, 2); -lean_inc(x_82); -x_83 = lean_ctor_get(x_17, 4); -lean_inc(x_83); -x_84 = lean_ctor_get(x_17, 5); -lean_inc(x_84); -x_85 = lean_ctor_get(x_17, 6); -lean_inc(x_85); -x_86 = lean_ctor_get(x_17, 7); -lean_inc(x_86); -if (lean_is_exclusive(x_17)) { - lean_ctor_release(x_17, 0); - lean_ctor_release(x_17, 1); - lean_ctor_release(x_17, 2); - lean_ctor_release(x_17, 3); - lean_ctor_release(x_17, 4); - lean_ctor_release(x_17, 5); - lean_ctor_release(x_17, 6); - lean_ctor_release(x_17, 7); - x_87 = x_17; -} else { - lean_dec_ref(x_17); - x_87 = lean_box(0); -} -x_88 = lean_ctor_get_uint64(x_18, sizeof(void*)*1); -x_89 = lean_ctor_get(x_18, 0); -lean_inc(x_89); -if (lean_is_exclusive(x_18)) { - lean_ctor_release(x_18, 0); - x_90 = x_18; -} else { - lean_dec_ref(x_18); - x_90 = lean_box(0); -} -x_91 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___spec__2___closed__1; -x_92 = 0; -x_93 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___spec__2___closed__2; -x_94 = lean_alloc_ctor(0, 2, 17); -lean_ctor_set(x_94, 0, x_1); -lean_ctor_set(x_94, 1, x_93); -lean_ctor_set_float(x_94, sizeof(void*)*2, x_91); -lean_ctor_set_float(x_94, sizeof(void*)*2 + 8, x_91); -lean_ctor_set_uint8(x_94, sizeof(void*)*2 + 16, x_92); -x_95 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___spec__2___closed__3; -x_96 = lean_alloc_ctor(9, 3, 0); -lean_ctor_set(x_96, 0, x_94); -lean_ctor_set(x_96, 1, x_14); -lean_ctor_set(x_96, 2, x_95); -lean_inc(x_12); -x_97 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_97, 0, x_12); -lean_ctor_set(x_97, 1, x_96); -x_98 = l_Lean_PersistentArray_push___rarg(x_89, x_97); -if (lean_is_scalar(x_90)) { - x_99 = lean_alloc_ctor(0, 1, 8); -} else { - x_99 = x_90; -} -lean_ctor_set(x_99, 0, x_98); -lean_ctor_set_uint64(x_99, sizeof(void*)*1, x_88); -if (lean_is_scalar(x_87)) { - x_100 = lean_alloc_ctor(0, 8, 0); -} else { - x_100 = x_87; -} -lean_ctor_set(x_100, 0, x_80); -lean_ctor_set(x_100, 1, x_81); -lean_ctor_set(x_100, 2, x_82); -lean_ctor_set(x_100, 3, x_99); -lean_ctor_set(x_100, 4, x_83); -lean_ctor_set(x_100, 5, x_84); -lean_ctor_set(x_100, 6, x_85); -lean_ctor_set(x_100, 7, x_86); -x_101 = lean_st_ref_set(x_10, x_100, x_79); -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; -} else { - lean_dec_ref(x_101); - x_103 = lean_box(0); -} -x_104 = lean_box(0); -if (lean_is_scalar(x_103)) { - x_105 = lean_alloc_ctor(0, 2, 0); -} else { - x_105 = x_103; -} -lean_ctor_set(x_105, 0, x_104); -lean_ctor_set(x_105, 1, x_102); -return x_105; -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -uint8_t x_13; lean_object* x_14; -x_13 = 1; -x_14 = l_Lean_Meta_Grind_pushEqCore(x_2, x_3, x_1, x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -return x_14; +x_19 = 1; +x_20 = l_Lean_Meta_Grind_pushEqCore(x_14, x_17, x_1, x_19, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_18); +return x_20; } } LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -uint8_t x_13; lean_object* x_14; -x_13 = 0; -x_14 = l_Lean_Meta_Grind_pushEqCore(x_2, x_3, x_1, x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -return x_14; +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; lean_object* x_20; +x_13 = l_Lean_Meta_Grind_shareCommon(x_2, 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_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +x_16 = l_Lean_Meta_Grind_shareCommon(x_3, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_15); +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +x_18 = lean_ctor_get(x_16, 1); +lean_inc(x_18); +lean_dec(x_16); +x_19 = 0; +x_20 = l_Lean_Meta_Grind_pushEqCore(x_14, x_17, x_1, x_19, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_18); +return x_20; } } static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__3___closed__1() { @@ -637,8 +282,16 @@ return x_2; static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__5___closed__7() { _start: { +lean_object* x_1; +x_1 = lean_mk_string_unchecked("", 0, 0); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__5___closed__8() { +_start: +{ lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___spec__2___closed__2; +x_1 = l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__5___closed__7; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } @@ -648,7 +301,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Gr { lean_object* x_12; lean_object* x_13; uint8_t x_14; x_12 = l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__5___closed__3; -x_13 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___spec__1(x_12, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_13 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_12, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); x_14 = !lean_is_exclusive(x_13); if (x_14 == 0) { @@ -669,65 +322,146 @@ 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; -x_21 = l_Lean_indentExpr(x_1); -x_22 = l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__5___closed__6; +lean_object* x_21; +x_21 = l_Lean_Meta_Grind_updateLastTag(x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, 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; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_22 = lean_ctor_get(x_21, 1); +lean_inc(x_22); +lean_dec(x_21); +x_23 = l_Lean_indentExpr(x_1); +x_24 = l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__5___closed__6; lean_ctor_set_tag(x_13, 7); -lean_ctor_set(x_13, 1, x_21); -lean_ctor_set(x_13, 0, x_22); -x_23 = l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__5___closed__7; -x_24 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_24, 0, x_13); -lean_ctor_set(x_24, 1, x_23); -x_25 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___spec__2(x_12, x_24, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_16); -x_26 = lean_ctor_get(x_25, 0); -lean_inc(x_26); -x_27 = lean_ctor_get(x_25, 1); -lean_inc(x_27); -lean_dec(x_25); -x_28 = lean_apply_10(x_17, x_26, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_27); -return x_28; +lean_ctor_set(x_13, 1, x_23); +lean_ctor_set(x_13, 0, x_24); +x_25 = l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__5___closed__8; +x_26 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_26, 0, x_13); +lean_ctor_set(x_26, 1, x_25); +x_27 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_12, x_26, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_22); +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_apply_10(x_17, x_28, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_29); +return x_30; } +else +{ +uint8_t x_31; +lean_free_object(x_13); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_31 = !lean_is_exclusive(x_21); +if (x_31 == 0) +{ +return x_21; } else { -lean_object* x_29; lean_object* x_30; lean_object* x_31; uint8_t x_32; -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_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_21, 0); +x_33 = lean_ctor_get(x_21, 1); +lean_inc(x_33); +lean_inc(x_32); +lean_dec(x_21); +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_35; lean_object* x_36; lean_object* x_37; uint8_t x_38; +x_35 = lean_ctor_get(x_13, 0); +x_36 = lean_ctor_get(x_13, 1); +lean_inc(x_36); +lean_inc(x_35); lean_dec(x_13); -x_31 = l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__5___closed__4; -x_32 = lean_unbox(x_29); -lean_dec(x_29); -if (x_32 == 0) +x_37 = l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__5___closed__4; +x_38 = lean_unbox(x_35); +lean_dec(x_35); +if (x_38 == 0) { -lean_object* x_33; lean_object* x_34; +lean_object* x_39; lean_object* x_40; lean_dec(x_1); -x_33 = lean_box(0); -x_34 = lean_apply_10(x_31, x_33, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_30); -return x_34; +x_39 = lean_box(0); +x_40 = lean_apply_10(x_37, x_39, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_36); +return x_40; } 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; -x_35 = l_Lean_indentExpr(x_1); -x_36 = l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__5___closed__6; -x_37 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_37, 0, x_36); -lean_ctor_set(x_37, 1, x_35); -x_38 = l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__5___closed__7; -x_39 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_39, 0, x_37); -lean_ctor_set(x_39, 1, x_38); -x_40 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___spec__2(x_12, x_39, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_30); -x_41 = lean_ctor_get(x_40, 0); -lean_inc(x_41); -x_42 = lean_ctor_get(x_40, 1); +lean_object* x_41; +x_41 = l_Lean_Meta_Grind_updateLastTag(x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_36); +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; +x_42 = lean_ctor_get(x_41, 1); lean_inc(x_42); -lean_dec(x_40); -x_43 = lean_apply_10(x_31, x_41, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_42); -return x_43; +lean_dec(x_41); +x_43 = l_Lean_indentExpr(x_1); +x_44 = l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__5___closed__6; +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___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__5___closed__8; +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_Meta_Grind_updateLastTag___spec__2(x_12, x_47, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_42); +x_49 = lean_ctor_get(x_48, 0); +lean_inc(x_49); +x_50 = lean_ctor_get(x_48, 1); +lean_inc(x_50); +lean_dec(x_48); +x_51 = lean_apply_10(x_37, x_49, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_50); +return x_51; +} +else +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; +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_52 = lean_ctor_get(x_41, 0); +lean_inc(x_52); +x_53 = lean_ctor_get(x_41, 1); +lean_inc(x_53); +if (lean_is_exclusive(x_41)) { + lean_ctor_release(x_41, 0); + lean_ctor_release(x_41, 1); + x_54 = x_41; +} else { + lean_dec_ref(x_41); + x_54 = lean_box(0); +} +if (lean_is_scalar(x_54)) { + x_55 = lean_alloc_ctor(1, 2, 0); +} else { + x_55 = x_54; +} +lean_ctor_set(x_55, 0, x_52); +lean_ctor_set(x_55, 1, x_53); +return x_55; +} } } } @@ -870,10 +604,9 @@ return x_42; } else { -uint8_t x_43; lean_object* x_44; +lean_object* x_43; lean_dec(x_1); -x_43 = 1; -x_44 = l_Lean_Meta_Grind_pushEqCore(x_31, x_19, x_2, x_43, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_14); +x_43 = l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__1(x_2, x_31, x_19, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_14); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -882,18 +615,17 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -return x_44; +return x_43; } } } else { -uint8_t x_45; lean_object* x_46; +lean_object* x_44; lean_dec(x_32); lean_dec(x_31); lean_dec(x_1); -x_45 = 0; -x_46 = l_Lean_Meta_Grind_pushEqCore(x_24, x_19, x_2, x_45, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_14); +x_44 = l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__2(x_2, x_24, x_19, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_14); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -902,53 +634,21 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -return x_46; +return x_44; } } } else { -lean_object* x_47; +lean_object* x_45; lean_dec(x_25); lean_dec(x_1); -x_47 = l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__3(x_2, x_24, x_19, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_14); -return x_47; -} -} -} +x_45 = l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__3(x_2, x_24, x_19, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_14); +return x_45; } } -LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___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_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___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_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_11; } } -LEAN_EXPORT lean_object* l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -lean_object* x_12; -x_12 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -return x_12; -} } LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___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: @@ -1209,10 +909,6 @@ x_21 = lean_ctor_get(x_19, 1); lean_inc(x_21); lean_dec(x_19); x_22 = l_Lean_Meta_Grind_closeGoal(x_20, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_21); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); return x_22; } else @@ -2294,11 +1990,6 @@ _G_initialized = true; res = initialize_Lean_Meta_Tactic_Grind_Types(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___spec__2___closed__1 = _init_l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___spec__2___closed__1(); -l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___spec__2___closed__2 = _init_l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___spec__2___closed__2(); -lean_mark_persistent(l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___spec__2___closed__2); -l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___spec__2___closed__3 = _init_l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___spec__2___closed__3(); -lean_mark_persistent(l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___spec__2___closed__3); l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__3___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__3___closed__1(); lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__3___closed__1); l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__3___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__3___closed__2(); @@ -2317,6 +2008,8 @@ l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lam lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__5___closed__6); l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__5___closed__7 = _init_l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__5___closed__7(); lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__5___closed__7); +l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__5___closed__8 = _init_l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__5___closed__8(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__5___closed__8); l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___closed__1(); lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___closed__1); l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___closed__2(); diff --git a/stage0/stdlib/Lean/Meta/Tactic/Grind/DoNotSimp.c b/stage0/stdlib/Lean/Meta/Tactic/Grind/DoNotSimp.c new file mode 100644 index 000000000000..2e56424bb556 --- /dev/null +++ b/stage0/stdlib/Lean/Meta/Tactic/Grind/DoNotSimp.c @@ -0,0 +1,608 @@ +// Lean compiler output +// Module: Lean.Meta.Tactic.Grind.DoNotSimp +// Imports: Init.Grind.Util Init.Simproc Lean.Meta.Tactic.Simp.Simproc +#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___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__5; +static lean_object* l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__2; +uint8_t l_Lean_Expr_isApp(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markAsDoNotSimp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_eraseDoNotSimp(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_reduceDoNotSimp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_markAsDoNotSimp___closed__4; +lean_object* l_Lean_Expr_cleanupAnnotations(lean_object*); +static lean_object* l_Lean_Meta_Grind_reduceDoNotSimp___lambda__2___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_reduceDoNotSimp___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_eraseDoNotSimp___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_reduceDoNotSimp___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_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__3; +static lean_object* l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__8; +lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_eraseDoNotSimp___lambda__3___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addDoNotSimp___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_reduceDoNotSimp___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_Grind_addDoNotSimp(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_eraseDoNotSimp___closed__2; +lean_object* l_Lean_Expr_appArg(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_Expr_appFnCleanup(lean_object*, lean_object*); +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_Grind_reduceDoNotSimp___closed__1; +static lean_object* l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__6; +static lean_object* l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_eraseDoNotSimp___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_eraseDoNotSimp___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Core_transform___at_Lean_Core_betaReduce___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__4; +static lean_object* l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__7; +lean_object* l_Lean_Meta_Simp_registerBuiltinDSimproc(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_markAsDoNotSimp___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_reduceDoNotSimp___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Expr_isConstOf(lean_object*, lean_object*); +static lean_object* l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__9; +static lean_object* l_Lean_Meta_Grind_markAsDoNotSimp___closed__2; +lean_object* l_Lean_Meta_instantiateMVarsIfMVarApp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_array_mk(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_eraseDoNotSimp___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_eraseDoNotSimp___closed__1; +static lean_object* l_Lean_Meta_Grind_markAsDoNotSimp___closed__3; +lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_eraseDoNotSimp___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_eraseDoNotSimp___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_eraseDoNotSimp___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* _init_l_Lean_Meta_Grind_markAsDoNotSimp___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean", 4, 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_markAsDoNotSimp___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Grind", 5, 5); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_markAsDoNotSimp___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("doNotSimp", 9, 9); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_markAsDoNotSimp___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Meta_Grind_markAsDoNotSimp___closed__1; +x_2 = l_Lean_Meta_Grind_markAsDoNotSimp___closed__2; +x_3 = l_Lean_Meta_Grind_markAsDoNotSimp___closed__3; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markAsDoNotSimp(lean_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; +x_7 = lean_box(0); +x_8 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_8, 0, x_1); +lean_ctor_set(x_8, 1, x_7); +x_9 = lean_array_mk(x_8); +x_10 = l_Lean_Meta_Grind_markAsDoNotSimp___closed__4; +x_11 = l_Lean_Meta_mkAppM(x_10, x_9, x_2, x_3, x_4, x_5, x_6); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_reduceDoNotSimp___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; lean_object* x_12; +x_11 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_11, 0, x_1); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_10); +return x_12; +} +} +static lean_object* _init_l_Lean_Meta_Grind_reduceDoNotSimp___lambda__2___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(0); +x_2 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_reduceDoNotSimp___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; lean_object* x_11; +x_10 = l_Lean_Meta_Grind_reduceDoNotSimp___lambda__2___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); +return x_11; +} +} +static lean_object* _init_l_Lean_Meta_Grind_reduceDoNotSimp___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_reduceDoNotSimp___lambda__2___boxed), 9, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_reduceDoNotSimp(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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; uint8_t x_15; +lean_inc(x_1); +x_10 = l_Lean_Meta_instantiateMVarsIfMVarApp(x_1, x_5, x_6, x_7, x_8, x_9); +x_11 = lean_ctor_get(x_10, 0); +lean_inc(x_11); +x_12 = lean_ctor_get(x_10, 1); +lean_inc(x_12); +lean_dec(x_10); +x_13 = l_Lean_Meta_Grind_reduceDoNotSimp___closed__1; +x_14 = l_Lean_Expr_cleanupAnnotations(x_11); +x_15 = l_Lean_Expr_isApp(x_14); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; +lean_dec(x_14); +lean_dec(x_1); +x_16 = lean_box(0); +x_17 = lean_apply_9(x_13, x_16, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_12); +return x_17; +} +else +{ +lean_object* x_18; uint8_t x_19; +x_18 = l_Lean_Expr_appFnCleanup(x_14, lean_box(0)); +x_19 = l_Lean_Expr_isApp(x_18); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; +lean_dec(x_18); +lean_dec(x_1); +x_20 = lean_box(0); +x_21 = lean_apply_9(x_13, x_20, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_12); +return x_21; +} +else +{ +lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_22 = l_Lean_Expr_appFnCleanup(x_18, lean_box(0)); +x_23 = l_Lean_Meta_Grind_markAsDoNotSimp___closed__4; +x_24 = l_Lean_Expr_isConstOf(x_22, x_23); +lean_dec(x_22); +if (x_24 == 0) +{ +lean_object* x_25; lean_object* x_26; +lean_dec(x_1); +x_25 = lean_box(0); +x_26 = lean_apply_9(x_13, x_25, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_12); +return x_26; +} +else +{ +lean_object* x_27; lean_object* x_28; +x_27 = lean_box(0); +x_28 = l_Lean_Meta_Grind_reduceDoNotSimp___lambda__1(x_1, x_27, 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); +lean_dec(x_2); +return x_28; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_reduceDoNotSimp___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_Meta_Grind_reduceDoNotSimp___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_reduceDoNotSimp___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_Meta_Grind_reduceDoNotSimp___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_10; +} +} +static lean_object* _init_l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Meta", 4, 4); +return x_1; +} +} +static lean_object* _init_l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("reduceDoNotSimp", 15, 15); +return x_1; +} +} +static lean_object* _init_l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____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_Meta_Grind_markAsDoNotSimp___closed__1; +x_2 = l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__1; +x_3 = l_Lean_Meta_Grind_markAsDoNotSimp___closed__2; +x_4 = l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__2; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_markAsDoNotSimp___closed__4; +x_2 = lean_unsigned_to_nat(2u); +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_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = lean_box(3); +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(3); +x_2 = l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__5; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__4; +x_2 = l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____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___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__7; +x_2 = lean_array_mk(x_1); +return x_2; +} +} +static lean_object* _init_l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__9() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_reduceDoNotSimp), 9, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153_(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_2 = l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__3; +x_3 = l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__8; +x_4 = l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__9; +x_5 = l_Lean_Meta_Simp_registerBuiltinDSimproc(x_2, x_3, x_4, x_1); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addDoNotSimp(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_object* x_7; +x_5 = l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__3; +x_6 = 0; +x_7 = l_Lean_Meta_Simp_Simprocs_add(x_1, x_5, x_6, x_2, x_3, x_4); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addDoNotSimp___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_Grind_addDoNotSimp(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_Meta_Grind_eraseDoNotSimp___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; +x_5 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_5, 0, x_1); +x_6 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_6, 0, x_5); +x_7 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_7, 0, x_6); +lean_ctor_set(x_7, 1, x_4); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_eraseDoNotSimp___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; +x_6 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_6, 0, x_1); +x_7 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_7, 0, x_6); +x_8 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_8, 0, x_7); +lean_ctor_set(x_8, 1, x_5); +return x_8; +} +} +static lean_object* _init_l_Lean_Meta_Grind_eraseDoNotSimp___lambda__3___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_eraseDoNotSimp___lambda__1___boxed), 4, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_eraseDoNotSimp___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; uint8_t x_7; +x_5 = l_Lean_Meta_Grind_eraseDoNotSimp___lambda__3___closed__1; +lean_inc(x_1); +x_6 = l_Lean_Expr_cleanupAnnotations(x_1); +x_7 = l_Lean_Expr_isApp(x_6); +if (x_7 == 0) +{ +lean_object* x_8; lean_object* x_9; +lean_dec(x_6); +x_8 = lean_box(0); +x_9 = l_Lean_Meta_Grind_eraseDoNotSimp___lambda__2(x_1, x_8, x_2, x_3, x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_9; +} +else +{ +lean_object* x_10; lean_object* x_11; uint8_t x_12; +x_10 = l_Lean_Expr_appArg(x_6, lean_box(0)); +x_11 = l_Lean_Expr_appFnCleanup(x_6, lean_box(0)); +x_12 = l_Lean_Expr_isApp(x_11); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; +lean_dec(x_11); +lean_dec(x_10); +x_13 = lean_box(0); +x_14 = l_Lean_Meta_Grind_eraseDoNotSimp___lambda__2(x_1, x_13, x_2, x_3, x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_14; +} +else +{ +lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_15 = l_Lean_Expr_appFnCleanup(x_11, lean_box(0)); +x_16 = l_Lean_Meta_Grind_markAsDoNotSimp___closed__4; +x_17 = l_Lean_Expr_isConstOf(x_15, x_16); +lean_dec(x_15); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; +lean_dec(x_10); +x_18 = lean_box(0); +x_19 = l_Lean_Meta_Grind_eraseDoNotSimp___lambda__2(x_1, x_18, x_2, x_3, x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_19; +} +else +{ +lean_object* x_20; +lean_dec(x_1); +x_20 = lean_apply_4(x_5, x_10, x_2, x_3, x_4); +return x_20; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_eraseDoNotSimp___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; lean_object* x_6; +x_5 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_5, 0, x_1); +x_6 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_6, 0, x_5); +lean_ctor_set(x_6, 1, x_4); +return x_6; +} +} +static lean_object* _init_l_Lean_Meta_Grind_eraseDoNotSimp___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_eraseDoNotSimp___lambda__3), 4, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_eraseDoNotSimp___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_eraseDoNotSimp___lambda__4___boxed), 4, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_eraseDoNotSimp(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_Meta_Grind_eraseDoNotSimp___closed__1; +x_6 = l_Lean_Meta_Grind_eraseDoNotSimp___closed__2; +x_7 = l_Lean_Core_transform___at_Lean_Core_betaReduce___spec__1(x_1, x_5, x_6, x_2, x_3, x_4); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_eraseDoNotSimp___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_Grind_eraseDoNotSimp___lambda__1(x_1, x_2, x_3, x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_eraseDoNotSimp___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; +x_6 = l_Lean_Meta_Grind_eraseDoNotSimp___lambda__2(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_eraseDoNotSimp___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l_Lean_Meta_Grind_eraseDoNotSimp___lambda__4(x_1, x_2, x_3, x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_5; +} +} +lean_object* initialize_Init_Grind_Util(uint8_t builtin, lean_object*); +lean_object* initialize_Init_Simproc(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Meta_Tactic_Simp_Simproc(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Lean_Meta_Tactic_Grind_DoNotSimp(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_Grind_Util(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Init_Simproc(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Lean_Meta_Tactic_Simp_Simproc(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l_Lean_Meta_Grind_markAsDoNotSimp___closed__1 = _init_l_Lean_Meta_Grind_markAsDoNotSimp___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_markAsDoNotSimp___closed__1); +l_Lean_Meta_Grind_markAsDoNotSimp___closed__2 = _init_l_Lean_Meta_Grind_markAsDoNotSimp___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_markAsDoNotSimp___closed__2); +l_Lean_Meta_Grind_markAsDoNotSimp___closed__3 = _init_l_Lean_Meta_Grind_markAsDoNotSimp___closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind_markAsDoNotSimp___closed__3); +l_Lean_Meta_Grind_markAsDoNotSimp___closed__4 = _init_l_Lean_Meta_Grind_markAsDoNotSimp___closed__4(); +lean_mark_persistent(l_Lean_Meta_Grind_markAsDoNotSimp___closed__4); +l_Lean_Meta_Grind_reduceDoNotSimp___lambda__2___closed__1 = _init_l_Lean_Meta_Grind_reduceDoNotSimp___lambda__2___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_reduceDoNotSimp___lambda__2___closed__1); +l_Lean_Meta_Grind_reduceDoNotSimp___closed__1 = _init_l_Lean_Meta_Grind_reduceDoNotSimp___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_reduceDoNotSimp___closed__1); +l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__1 = _init_l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__1); +l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__2 = _init_l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__2(); +lean_mark_persistent(l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__2); +l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__3 = _init_l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__3(); +lean_mark_persistent(l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__3); +l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__4 = _init_l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__4(); +lean_mark_persistent(l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__4); +l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__5 = _init_l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__5(); +lean_mark_persistent(l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__5); +l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__6 = _init_l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__6(); +lean_mark_persistent(l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__6); +l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__7 = _init_l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__7(); +lean_mark_persistent(l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__7); +l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__8 = _init_l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__8(); +lean_mark_persistent(l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__8); +l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__9 = _init_l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__9(); +lean_mark_persistent(l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__9); +if (builtin) {res = l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153_(lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +}l_Lean_Meta_Grind_eraseDoNotSimp___lambda__3___closed__1 = _init_l_Lean_Meta_Grind_eraseDoNotSimp___lambda__3___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_eraseDoNotSimp___lambda__3___closed__1); +l_Lean_Meta_Grind_eraseDoNotSimp___closed__1 = _init_l_Lean_Meta_Grind_eraseDoNotSimp___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_eraseDoNotSimp___closed__1); +l_Lean_Meta_Grind_eraseDoNotSimp___closed__2 = _init_l_Lean_Meta_Grind_eraseDoNotSimp___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_eraseDoNotSimp___closed__2); +return lean_io_result_mk_ok(lean_box(0)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Lean/Meta/Tactic/Grind/EMatch.c b/stage0/stdlib/Lean/Meta/Tactic/Grind/EMatch.c index b573249175a5..33e0caa7696d 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Grind/EMatch.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Grind/EMatch.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Lean.Meta.Tactic.Grind.EMatch -// Imports: Lean.Meta.Tactic.Grind.Types Lean.Meta.Tactic.Grind.Intro +// Imports: Lean.Meta.Tactic.Grind.Types Lean.Meta.Tactic.Grind.Intro Lean.Meta.Tactic.Grind.DoNotSimp #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -14,7 +14,7 @@ extern "C" { #endif lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___boxed(lean_object*, lean_object*, lean_object*, lean_object*, 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_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___spec__1___closed__2; LEAN_EXPORT lean_object* l_Lean_mkConstWithLevelParams___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___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*); @@ -23,6 +23,8 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Grind_EMatch_ematchTheorem(lean_object*, le LEAN_EXPORT lean_object* l_Lean_Meta_Grind_EMatch_ematchTheorem___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_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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*); extern lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__5; +lean_object* l_Lean_mkNatLit(lean_object*); static lean_object* l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2; lean_object* lean_mk_empty_array_with_capacity(lean_object*); lean_object* l_Lean_mkAppN(lean_object*, lean_object*); @@ -33,22 +35,31 @@ lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(lean_object*, l lean_object* l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___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___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___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_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___boxed(lean_object**); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___boxed(lean_object**); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematchAndAssert___lambda__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_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__8; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_EMatch_instInhabitedContext; LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___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_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__3; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_ConstantInfo_levelParams(lean_object*); +lean_object* l_Lean_mkAppB(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___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___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__2; +lean_object* l_Lean_Meta_isOffset_x3f(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_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_main___spec__1___boxed(lean_object**); static lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___spec__1___closed__6; uint64_t lean_uint64_lor(uint64_t, uint64_t); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_withInitApp___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_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__3; uint8_t l_Lean_Expr_isApp(lean_object*); +lean_object* l_Lean_Meta_Grind_markAsDoNotSimp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___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* l_Lean_Meta_trySynthInstance(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___lambda__1___closed__3; lean_object* l_Lean_MessageData_ofList(lean_object*); -lean_object* l_Lean_Meta_Grind_iterate(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_push___rarg(lean_object*, lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); lean_object* l_Array_toSubarray___rarg(lean_object*, lean_object*, lean_object*); @@ -60,67 +71,74 @@ static lean_object* l_Lean_Meta_Grind_ematch___closed__1; uint8_t lean_usize_dec_eq(size_t, size_t); static lean_object* l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5___closed__4; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_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*); +uint8_t l_Lean_Meta_Match_isMatchEqnTheorem(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___closed__4; lean_object* l_instInhabitedReaderT___rarg___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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*); static lean_object* l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5___closed__1; lean_object* l_Lean_Expr_bvar___override(lean_object*); lean_object* lean_array_fget(lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_getMaxGeneration___rarg(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___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__8___closed__1; extern lean_object* l_Lean_Meta_Grind_grind_debug_proofs; -static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__4; +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_List_mapTR_loop___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__2(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processOffset___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_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_Grind_EMatch_ematchTheorems___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* lean_environment_find(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_EMatch_ematchTheorems___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static size_t l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematchAndAssert(lean_object*, lean_object*, 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_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___boxed(lean_object**); LEAN_EXPORT lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_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___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__3; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_Choice_updateGen(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_ReaderT_read___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___spec__1___closed__3; -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematchAndAssert_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_Expr_cleanupAnnotations(lean_object*); static lean_object* l_Lean_Meta_Grind_EMatch_ematchTheorem___lambda__1___closed__2; -static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__3; lean_object* l_Lean_stringToMessageData(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_EMatch_M_run_x27___rarg___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Meta_Grind_assertNext_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_evalNat(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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*); +lean_object* l_Lean_mkApp4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_List_head_x21___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___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*); static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___closed__1; -static lean_object* l_Lean_Meta_Grind_ematchAndAssert_x3f___closed__1; static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___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___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__2; LEAN_EXPORT lean_object* l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_appArg_x21(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematchAndAssert_x3f___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___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__3; uint8_t l_Lean_Expr_isBVar(lean_object*); lean_object* l_List_mapTR_loop___at_Lean_mkConstWithLevelParams___spec__1(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__8___closed__2; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__1; static lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___spec__1___closed__5; LEAN_EXPORT lean_object* l_Lean_mkConstWithLevelParams___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___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*); extern lean_object* l_instInhabitedPUnit; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematchAndAssert___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_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_hasMVar(lean_object*); lean_object* l_Lean_Meta_Grind_EMatchTheorem_getProofWithFreshMVarLevels(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_Grind_checkMaxEmatchExceeded(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forMAux___at_Lean_Meta_Grind_EMatch_ematchTheorems___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___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__5; lean_object* l_List_appendTR___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_ReaderT_read___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem(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_Lean_Meta_Grind_ematch___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___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___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___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___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___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__4; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__2; lean_object* lean_checked_assign(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Array_reverse___rarg(lean_object*); lean_object* lean_st_ref_take(lean_object*, lean_object*); static lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___spec__1___closed__4; +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___spec__1___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_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*); static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__7; static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__3; @@ -128,101 +146,119 @@ uint8_t lean_expr_eqv(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withNewMCtxDepth___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__12___rarg(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_EMatch_instInhabitedState; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__10(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___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*); uint64_t lean_uint64_shift_right(uint64_t, uint64_t); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices(lean_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_ofSyntax(lean_object*); static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___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___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MVarId_getType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withNewMCtxDepth___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__12___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_x3f___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1(size_t, size_t, lean_object*); static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__8; +extern lean_object* l___private_Lean_Expr_0__Lean_natAddFn; lean_object* l_Lean_Meta_Grind_checkMaxInstancesExceeded(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3___closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematch___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_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__2; +lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDeclImp___rarg(lean_object*, uint8_t, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_groundPattern_x3f(lean_object*); static lean_object* l_Lean_Meta_Grind_EMatch_instInhabitedCnstr___closed__1; static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__9; -static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__3; -lean_object* l_Lean_MessageData_ofFormat(lean_object*); lean_object* l_Lean_PersistentArray_append___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_EMatch_instInhabitedCnstr___closed__2; 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_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_main(lean_object*, lean_object*, lean_object*, lean_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_Grind_EMatch_instInhabitedChoice___closed__2; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__4; +lean_object* l_Lean_Expr_appArg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___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* l_outOfBounds___rarg(lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_ematchAndAssert_x3f___closed__2; LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_main___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_ematch___closed__2; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___lambda__1___closed__1; lean_object* lean_st_ref_get(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__7; static lean_object* l_Lean_Meta_Grind_EMatch_instInhabitedContext___closed__1; LEAN_EXPORT lean_object* l_Lean_PersistentArray_forMAux___at_Lean_Meta_Grind_EMatch_ematchTheorems___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*); uint8_t l_List_isEmpty___rarg(lean_object*); lean_object* lean_st_mk_ref(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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*, lean_object*); lean_object* lean_array_to_list(lean_object*); +lean_object* l_Lean_Expr_appFnCleanup(lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___lambda__1___closed__3; static lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___spec__1___closed__1; LEAN_EXPORT lean_object* l_Lean_PersistentArray_forM___at_Lean_Meta_Grind_EMatch_ematchTheorems___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_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_withInitApp(lean_object*); lean_object* l_Lean_Expr_toHeadIndex(lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__4; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType(lean_object*, lean_object*, lean_object*, lean_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_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData(lean_object*); +lean_object* l_Lean_Meta_Grind_isOffsetPattern_x3f(lean_object*); static lean_object* l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5___closed__2; -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___closed__2; 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_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__10; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_constName_x21(lean_object*); extern lean_object* l_Lean_instInhabitedExpr; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__4___boxed(lean_object**); lean_object* l_Lean_MessageData_ofConst(lean_object*); +static lean_object* l_Lean_Meta_Grind_ematchAndAssert___closed__1; LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_EMatch_instInhabitedCnstr; uint8_t l_Lean_BinderInfo_isInstImplicit(uint8_t); static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__5; -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___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_Grind_ematchAndAssert___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_Meta_Grind_EMatch_ematchTheorem___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___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___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__6; LEAN_EXPORT lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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_Lean_Meta_Grind_EMatch_M_run_x27___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_Grind_ematchAndAssert_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_Grind_ematchAndAssert_x3f___lambda__4___closed__1; static lean_object* l_Lean_Meta_Grind_EMatch_instInhabitedChoice___closed__1; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__2; LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_main___spec__1___lambda__1___boxed(lean_object**); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3___closed__2; uint8_t l_Lean_Meta_Grind_isSameExpr_unsafe__1(lean_object*, lean_object*); static uint64_t l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___closed__2; +lean_object* l_Lean_Meta_Grind_GrindTactic_iterate(lean_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_Grind_EMatch_0__Lean_Meta_Grind_EMatch_eqvFunctions___boxed(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_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_main___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematch___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_Meta_Grind_EMatch_instInhabitedContext___closed__2; static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___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_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_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_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___closed__3; lean_object* l_Lean_Expr_appFn_x21(lean_object*); +lean_object* l_Lean_Meta_Grind_updateLastTag(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_ematch___closed__3; static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__4; extern lean_object* l_Lean_Meta_instMonadMetaM; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__5; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_EMatch_ematchTheorem___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_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__9___boxed(lean_object**); static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___spec__1___closed__1; uint8_t l_Lean_PersistentHashMap_contains___at_Lean_MVarId_isAssigned___spec__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematchAndAssert_x3f___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofConstName(lean_object*, uint8_t); lean_object* l_OptionT_instMonad___rarg(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematchAndAssert_x3f___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___boxed(lean_object**); lean_object* l_Lean_MessageData_ofExpr(lean_object*); static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__4; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___lambda__1___closed__4; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance(lean_object*, lean_object*, lean_object*, lean_object*, lean_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_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_List_tail_x21___rarg(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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*); static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__6; @@ -230,47 +266,56 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_Grind_EMatch_e LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); double l_Float_ofScientific(lean_object*, uint8_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__1; +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___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 lean_object* l_Lean_Meta_Grind_EMatch_ematchTheorems(lean_object*, lean_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_forallMetaBoundedTelescope(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_Grind_EMatch_ematchTheorems___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_getNext(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_Grind_ENode_isCongrRoot(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematchAndAssert___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_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5___closed__3; lean_object* l_Lean_PersistentHashMap_find_x3f___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___boxed(lean_object*, lean_object*, lean_object*, 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_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__4; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_EMatch_M_run_x27(lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__5; +lean_object* l_Lean_Expr_constLevels_x21(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematchAndAssert___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* l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___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_Meta_withLocalDecl___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___spec__1(lean_object*); LEAN_EXPORT lean_object* l_ReaderT_bind___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__11___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematch___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___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___boxed(lean_object**); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematchAndAssert_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_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___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___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_x3f___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_eqvFunctions(lean_object*, lean_object*); +lean_object* l_Lean_Meta_Grind_shareCommon(lean_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_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___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_Grind_addTheoremInstance(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___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*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_unassigned___closed__1; LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__6___boxed(lean_object**); -static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__7; static lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__2___closed__2; +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processOffset___spec__1___boxed(lean_object**); lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); lean_object* l_Lean_indentExpr(lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__8; +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processOffset___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static double l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__1; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__7___boxed(lean_object**); static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__1; lean_object* l_Lean_Meta_Grind_markTheoremInstance(lean_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_mkHasTypeButIsExpectedMsg(lean_object*, lean_object*, lean_object*, 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_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___lambda__2___closed__1; +lean_object* l_Lean_Meta_mkForallFVars(lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_pushChoice(lean_object*, lean_object*, 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_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__2; lean_object* lean_panic_fn(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processOffset(lean_object*, lean_object*, lean_object*, lean_object*, lean_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_Grind_isEqv(lean_object*, 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_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__2___closed__3; -static lean_object* l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3___closed__1; lean_object* lean_nat_sub(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_getAppFn(lean_object*); @@ -278,48 +323,50 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_ static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__1; LEAN_EXPORT lean_object* l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___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*); uint64_t lean_uint64_shift_left(uint64_t, uint64_t); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematchAndAssert___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___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___lambda__1___closed__2; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_getENode(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_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___closed__2; -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematchAndAssert_x3f___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_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_List_reverse___rarg(lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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*); static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__5; -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_mk(lean_object*); lean_object* l_Lean_instantiateMVarsCore(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withNewMCtxDepth___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__12(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematch(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_Grind_EMatch_ematchTheorems___spec__3(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__4(lean_object*, lean_object*); lean_object* l_Lean_isTracingEnabledForCore(lean_object*, lean_object*, lean_object*); size_t lean_usize_add(size_t, size_t); LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___spec__1___boxed(lean_object**); -static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch(lean_object*, lean_object*, lean_object*, lean_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_object* l_Lean_Expr_fvar___override(lean_object*); size_t lean_array_size(lean_object*); +static lean_object* l_Lean_Meta_Grind_ematchAndAssert___lambda__4___closed__1; lean_object* l_instInhabitedOfMonad___rarg(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__5___closed__1; lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematchAndAssert___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_ReaderT_bind___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__11(lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_unassigned___closed__3; -static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__6; static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___lambda__1___closed__1; -static size_t l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__1; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__7; lean_object* lean_string_append(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__5___boxed(lean_object**); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_Grind_canon(lean_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_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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_array_get_size(lean_object*); +static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__11; +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___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_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___lambda__1___closed__2; LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_EMatch_instInhabitedContext___closed__3; +lean_object* l_Lean_Meta_Grind_Origin_key(lean_object*); lean_object* lean_infer_type(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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*); uint8_t lean_nat_dec_le(lean_object*, lean_object*); @@ -329,10 +376,12 @@ LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Meta_Grind_EMatch_ematchThe lean_object* l_Lean_Expr_bvarIdx_x21(lean_object*); lean_object* l_Lean_Meta_mkLambdaFVars(lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint64_t l_Lean_Meta_TransparencyMode_toUInt64(uint8_t); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__4___closed__1; lean_object* lean_nat_add(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_isAssigned___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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*); uint8_t l_Lean_Expr_isConst(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_pushChoice___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___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___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_unassigned; uint8_t l_Lean_Expr_isFVar(lean_object*); @@ -342,18 +391,24 @@ LEAN_EXPORT lean_object* l_Lean_PersistentArray_forM___at_Lean_Meta_Grind_EMatch static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_unassigned___closed__2; lean_object* lean_array_uset(lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_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_object* l_Lean_MessageData_ofName(lean_object*); static lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__2___closed__1; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__6; LEAN_EXPORT lean_object* l_Lean_MVarId_isAssigned___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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_expr_instantiate1(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__6; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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_object*, lean_object*); +lean_object* l_Lean_Meta_Grind_assertNext(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_isProof(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___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___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__5; lean_object* l_ReaderT_instMonad___rarg(lean_object*); +lean_object* l_Lean_Meta_Grind_internalize(lean_object*, lean_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_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue(lean_object*, lean_object*, lean_object*, lean_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_Grind_ematchAndAssert___closed__2; LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___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*); static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__4; -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematchAndAssert_x3f___lambda__3(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_Grind_EMatch_instInhabitedCnstr___closed__1() { _start: { @@ -586,13 +641,15 @@ return x_6; static lean_object* _init_l_Lean_Meta_Grind_EMatch_instInhabitedContext___closed__3() { _start: { -uint8_t x_1; lean_object* x_2; lean_object* x_3; +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = 0; x_2 = l_Lean_Meta_Grind_EMatch_instInhabitedContext___closed__2; -x_3 = lean_alloc_ctor(0, 1, 1); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_1); -return x_3; +x_3 = l_Lean_Meta_Grind_EMatch_instInhabitedCnstr___closed__1; +x_4 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_4, 0, x_2); +lean_ctor_set(x_4, 1, x_3); +lean_ctor_set_uint8(x_4, sizeof(void*)*2, x_1); +return x_4; } } static lean_object* _init_l_Lean_Meta_Grind_EMatch_instInhabitedContext() { @@ -614,13 +671,15 @@ return x_1; static lean_object* _init_l_Lean_Meta_Grind_EMatch_M_run_x27___rarg___closed__1() { _start: { -uint8_t x_1; lean_object* x_2; lean_object* x_3; +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = 1; x_2 = l_Lean_Meta_Grind_EMatch_instInhabitedContext___closed__2; -x_3 = lean_alloc_ctor(0, 1, 1); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_1); -return x_3; +x_3 = l_Lean_Meta_Grind_EMatch_instInhabitedCnstr___closed__1; +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_Grind_EMatch_M_run_x27___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) { @@ -701,6 +760,44 @@ x_2 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_EMatch_M_run_x27___rarg), 10, return x_2; } } +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_withInitApp___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +uint8_t x_14; +x_14 = !lean_is_exclusive(x_3); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; +x_15 = lean_ctor_get(x_3, 1); +lean_dec(x_15); +lean_ctor_set(x_3, 1, x_1); +x_16 = lean_apply_11(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_16; +} +else +{ +uint8_t x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_17 = lean_ctor_get_uint8(x_3, sizeof(void*)*2); +x_18 = lean_ctor_get(x_3, 0); +lean_inc(x_18); +lean_dec(x_3); +x_19 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_1); +lean_ctor_set_uint8(x_19, sizeof(void*)*2, x_17); +x_20 = lean_apply_11(x_2, x_19, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_20; +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_withInitApp(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_withInitApp___rarg), 13, 0); +return x_2; +} +} static lean_object* _init_l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___spec__1___closed__1() { _start: { @@ -796,7 +893,7 @@ static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___closed__1; x_2 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___closed__2; -x_3 = lean_unsigned_to_nat(80u); +x_3 = lean_unsigned_to_nat(88u); x_4 = lean_unsigned_to_nat(4u); x_5 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -1063,88 +1160,106 @@ x_4 = lean_box(x_3); return x_4; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_x3f___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__1() { _start: { -lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; -x_14 = l_Lean_Expr_appArg_x21(x_1); -x_15 = l_Lean_Expr_appArg_x21(x_2); -x_16 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare; -x_17 = lean_expr_eqv(x_14, x_16); -if (x_17 == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("assertion violation: ", 21, 21); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__2() { +_start: { -uint8_t x_18; -x_18 = l_Lean_Expr_isBVar(x_14); -if (x_18 == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("!isPatternDontCare pArg\n ", 28, 28); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__3() { +_start: { -lean_object* x_19; -x_19 = l_Lean_Meta_Grind_groundPattern_x3f(x_14); -if (lean_obj_tag(x_19) == 0) +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__2; +x_3 = lean_string_append(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__4() { +_start: { -lean_object* x_20; uint8_t x_21; -x_20 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_20, 0, x_14); -lean_ctor_set(x_20, 1, x_15); -x_21 = !lean_is_exclusive(x_3); -if (x_21 == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("_private.Lean.Meta.Tactic.Grind.EMatch.0.Lean.Meta.Grind.EMatch.matchArg\?", 73, 73); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__5() { +_start: { -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_3, 0); -x_23 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_23, 0, x_20); -lean_ctor_set(x_23, 1, x_22); -lean_ctor_set(x_3, 0, x_23); -x_24 = l_Lean_Expr_appFn_x21(x_1); -x_25 = l_Lean_Expr_appFn_x21(x_2); -x_26 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_x3f(x_3, x_24, x_25, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -lean_dec(x_25); -lean_dec(x_24); -return x_26; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__4; +x_3 = lean_unsigned_to_nat(109u); +x_4 = lean_unsigned_to_nat(4u); +x_5 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__3; +x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); +return x_6; } -else +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__6() { +_start: { -lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_27 = lean_ctor_get(x_3, 0); -x_28 = lean_ctor_get(x_3, 1); -x_29 = lean_ctor_get(x_3, 2); -lean_inc(x_29); -lean_inc(x_28); -lean_inc(x_27); -lean_dec(x_3); -x_30 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_30, 0, x_20); -lean_ctor_set(x_30, 1, x_27); -x_31 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_31, 0, x_30); -lean_ctor_set(x_31, 1, x_28); -lean_ctor_set(x_31, 2, x_29); -x_32 = l_Lean_Expr_appFn_x21(x_1); -x_33 = l_Lean_Expr_appFn_x21(x_2); -x_34 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_x3f(x_31, x_32, x_33, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -lean_dec(x_33); -lean_dec(x_32); -return x_34; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Option.isNone <| isOffsetPattern\? pArg\n ", 43, 43); +return x_1; } } -else +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__7() { +_start: { -lean_object* x_35; lean_object* x_36; -lean_dec(x_14); -x_35 = lean_ctor_get(x_19, 0); -lean_inc(x_35); -lean_dec(x_19); -x_36 = l_Lean_Meta_Grind_isEqv(x_35, x_15, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -if (lean_obj_tag(x_36) == 0) +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__6; +x_3 = lean_string_append(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__8() { +_start: { -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_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__4; +x_3 = lean_unsigned_to_nat(108u); +x_4 = lean_unsigned_to_nat(4u); +x_5 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__7; +x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); +return x_6; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_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, lean_object* x_11, lean_object* x_12) { +_start: { -uint8_t x_39; -lean_dec(x_12); +lean_object* x_13; uint8_t x_14; +x_13 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare; +x_14 = lean_expr_eqv(x_2, x_13); +if (x_14 == 0) +{ +uint8_t x_15; +x_15 = l_Lean_Expr_isBVar(x_2); +if (x_15 == 0) +{ +lean_object* x_16; +x_16 = l_Lean_Meta_Grind_groundPattern_x3f(x_2); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_17; +lean_inc(x_2); +x_17 = l_Lean_Meta_Grind_isOffsetPattern_x3f(x_2); +if (lean_obj_tag(x_17) == 0) +{ +lean_object* x_18; uint8_t x_19; lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -1152,48 +1267,75 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_3); -x_39 = !lean_is_exclusive(x_36); -if (x_39 == 0) +lean_dec(x_4); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_2); +lean_ctor_set(x_18, 1, x_3); +x_19 = !lean_is_exclusive(x_1); +if (x_19 == 0) { -lean_object* x_40; lean_object* x_41; -x_40 = lean_ctor_get(x_36, 0); -lean_dec(x_40); -x_41 = lean_box(0); -lean_ctor_set(x_36, 0, x_41); -return x_36; +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_20 = lean_ctor_get(x_1, 0); +x_21 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_21, 0, x_18); +lean_ctor_set(x_21, 1, x_20); +lean_ctor_set(x_1, 0, x_21); +x_22 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_22, 0, x_1); +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_12); +return x_23; } else { -lean_object* x_42; lean_object* x_43; lean_object* x_44; -x_42 = lean_ctor_get(x_36, 1); -lean_inc(x_42); -lean_dec(x_36); -x_43 = lean_box(0); -x_44 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_44, 0, x_43); -lean_ctor_set(x_44, 1, x_42); -return x_44; +lean_object* x_24; 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_1, 0); +x_25 = lean_ctor_get(x_1, 1); +x_26 = lean_ctor_get(x_1, 2); +lean_inc(x_26); +lean_inc(x_25); +lean_inc(x_24); +lean_dec(x_1); +x_27 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_27, 0, x_18); +lean_ctor_set(x_27, 1, x_24); +x_28 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_28, 0, x_27); +lean_ctor_set(x_28, 1, x_25); +lean_ctor_set(x_28, 2, x_26); +x_29 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_29, 0, 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_12); +return x_30; } } else { -lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; -x_45 = lean_ctor_get(x_36, 1); -lean_inc(x_45); -lean_dec(x_36); -x_46 = l_Lean_Expr_appFn_x21(x_1); -x_47 = l_Lean_Expr_appFn_x21(x_2); -x_48 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_x3f(x_3, x_46, x_47, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_45); -lean_dec(x_47); -lean_dec(x_46); -return x_48; -} -} -else +uint8_t x_31; +lean_dec(x_2); +x_31 = !lean_is_exclusive(x_17); +if (x_31 == 0) { -uint8_t x_49; -lean_dec(x_12); +lean_object* x_32; uint8_t x_33; +x_32 = lean_ctor_get(x_17, 0); +x_33 = !lean_is_exclusive(x_32); +if (x_33 == 0) +{ +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_34); +x_36 = l_Lean_Meta_Grind_isOffsetPattern_x3f(x_34); +if (lean_obj_tag(x_36) == 0) +{ +uint8_t x_37; +x_37 = lean_expr_eqv(x_34, x_13); +if (x_37 == 0) +{ +lean_object* x_38; uint8_t x_39; lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -1201,52 +1343,96 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_3); -x_49 = !lean_is_exclusive(x_36); -if (x_49 == 0) +lean_dec(x_4); +x_38 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_38, 0, x_34); +lean_ctor_set(x_38, 1, x_35); +lean_ctor_set(x_38, 2, x_3); +x_39 = !lean_is_exclusive(x_1); +if (x_39 == 0) { -return x_36; +lean_object* x_40; lean_object* x_41; +x_40 = lean_ctor_get(x_1, 0); +lean_ctor_set_tag(x_32, 1); +lean_ctor_set(x_32, 1, x_40); +lean_ctor_set(x_32, 0, x_38); +lean_ctor_set(x_1, 0, x_32); +lean_ctor_set(x_17, 0, x_1); +x_41 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_41, 0, x_17); +lean_ctor_set(x_41, 1, x_12); +return x_41; } else { -lean_object* x_50; lean_object* x_51; lean_object* x_52; -x_50 = lean_ctor_get(x_36, 0); -x_51 = lean_ctor_get(x_36, 1); -lean_inc(x_51); -lean_inc(x_50); -lean_dec(x_36); -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_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_42 = lean_ctor_get(x_1, 0); +x_43 = lean_ctor_get(x_1, 1); +x_44 = lean_ctor_get(x_1, 2); +lean_inc(x_44); +lean_inc(x_43); +lean_inc(x_42); +lean_dec(x_1); +lean_ctor_set_tag(x_32, 1); +lean_ctor_set(x_32, 1, x_42); +lean_ctor_set(x_32, 0, x_38); +x_45 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_45, 0, x_32); +lean_ctor_set(x_45, 1, x_43); +lean_ctor_set(x_45, 2, x_44); +lean_ctor_set(x_17, 0, x_45); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_17); +lean_ctor_set(x_46, 1, x_12); +return x_46; } } +else +{ +lean_object* x_47; lean_object* x_48; +lean_free_object(x_32); +lean_dec(x_35); +lean_dec(x_34); +lean_free_object(x_17); +lean_dec(x_3); +lean_dec(x_1); +x_47 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__5; +x_48 = l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___spec__1(x_47, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_48; } } else { -lean_object* x_53; lean_object* x_54; -x_53 = l_Lean_Expr_bvarIdx_x21(x_14); -lean_dec(x_14); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -x_54 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f(x_3, x_53, x_15, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -lean_dec(x_53); -if (lean_obj_tag(x_54) == 0) +lean_object* x_49; lean_object* x_50; +lean_dec(x_36); +lean_free_object(x_32); +lean_dec(x_35); +lean_dec(x_34); +lean_free_object(x_17); +lean_dec(x_3); +lean_dec(x_1); +x_49 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__8; +x_50 = l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___spec__1(x_49, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_50; +} +} +else { -lean_object* x_55; -x_55 = lean_ctor_get(x_54, 0); -lean_inc(x_55); -if (lean_obj_tag(x_55) == 0) +lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_51 = lean_ctor_get(x_32, 0); +x_52 = lean_ctor_get(x_32, 1); +lean_inc(x_52); +lean_inc(x_51); +lean_dec(x_32); +lean_inc(x_51); +x_53 = l_Lean_Meta_Grind_isOffsetPattern_x3f(x_51); +if (lean_obj_tag(x_53) == 0) { -uint8_t x_56; -lean_dec(x_12); +uint8_t x_54; +x_54 = lean_expr_eqv(x_51, x_13); +if (x_54 == 0) +{ +lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -1254,50 +1440,98 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -x_56 = !lean_is_exclusive(x_54); -if (x_56 == 0) -{ -lean_object* x_57; lean_object* x_58; -x_57 = lean_ctor_get(x_54, 0); -lean_dec(x_57); -x_58 = lean_box(0); -lean_ctor_set(x_54, 0, x_58); -return x_54; +lean_dec(x_4); +x_55 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_55, 0, x_51); +lean_ctor_set(x_55, 1, x_52); +lean_ctor_set(x_55, 2, x_3); +x_56 = lean_ctor_get(x_1, 0); +lean_inc(x_56); +x_57 = lean_ctor_get(x_1, 1); +lean_inc(x_57); +x_58 = lean_ctor_get(x_1, 2); +lean_inc(x_58); +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + lean_ctor_release(x_1, 1); + lean_ctor_release(x_1, 2); + x_59 = x_1; +} else { + lean_dec_ref(x_1); + x_59 = lean_box(0); } -else -{ -lean_object* x_59; lean_object* x_60; lean_object* x_61; -x_59 = lean_ctor_get(x_54, 1); -lean_inc(x_59); -lean_dec(x_54); -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_59); -return x_61; +x_60 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_60, 0, x_55); +lean_ctor_set(x_60, 1, x_56); +if (lean_is_scalar(x_59)) { + x_61 = lean_alloc_ctor(0, 3, 0); +} else { + x_61 = x_59; } +lean_ctor_set(x_61, 0, x_60); +lean_ctor_set(x_61, 1, x_57); +lean_ctor_set(x_61, 2, x_58); +lean_ctor_set(x_17, 0, x_61); +x_62 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_62, 0, x_17); +lean_ctor_set(x_62, 1, x_12); +return x_62; } else { -lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; -x_62 = lean_ctor_get(x_54, 1); -lean_inc(x_62); -lean_dec(x_54); -x_63 = lean_ctor_get(x_55, 0); -lean_inc(x_63); -lean_dec(x_55); -x_64 = l_Lean_Expr_appFn_x21(x_1); -x_65 = l_Lean_Expr_appFn_x21(x_2); -x_66 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_x3f(x_63, x_64, x_65, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_62); -lean_dec(x_65); -lean_dec(x_64); +lean_object* x_63; lean_object* x_64; +lean_dec(x_52); +lean_dec(x_51); +lean_free_object(x_17); +lean_dec(x_3); +lean_dec(x_1); +x_63 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__5; +x_64 = l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___spec__1(x_63, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_64; +} +} +else +{ +lean_object* x_65; lean_object* x_66; +lean_dec(x_53); +lean_dec(x_52); +lean_dec(x_51); +lean_free_object(x_17); +lean_dec(x_3); +lean_dec(x_1); +x_65 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__8; +x_66 = l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___spec__1(x_65, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); return x_66; } } +} else { -uint8_t x_67; -lean_dec(x_12); +lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; +x_67 = lean_ctor_get(x_17, 0); +lean_inc(x_67); +lean_dec(x_17); +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_67); + x_70 = lean_box(0); +} +lean_inc(x_68); +x_71 = l_Lean_Meta_Grind_isOffsetPattern_x3f(x_68); +if (lean_obj_tag(x_71) == 0) +{ +uint8_t x_72; +x_72 = lean_expr_eqv(x_68, x_13); +if (x_72 == 0) +{ +lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -1305,49 +1539,88 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -x_67 = !lean_is_exclusive(x_54); -if (x_67 == 0) -{ -return x_54; +lean_dec(x_4); +x_73 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_73, 0, x_68); +lean_ctor_set(x_73, 1, x_69); +lean_ctor_set(x_73, 2, x_3); +x_74 = lean_ctor_get(x_1, 0); +lean_inc(x_74); +x_75 = lean_ctor_get(x_1, 1); +lean_inc(x_75); +x_76 = lean_ctor_get(x_1, 2); +lean_inc(x_76); +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + lean_ctor_release(x_1, 1); + lean_ctor_release(x_1, 2); + x_77 = x_1; +} else { + lean_dec_ref(x_1); + x_77 = lean_box(0); } -else -{ -lean_object* x_68; lean_object* x_69; lean_object* x_70; -x_68 = lean_ctor_get(x_54, 0); -x_69 = lean_ctor_get(x_54, 1); -lean_inc(x_69); -lean_inc(x_68); -lean_dec(x_54); -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; +if (lean_is_scalar(x_70)) { + x_78 = lean_alloc_ctor(1, 2, 0); +} else { + x_78 = x_70; + lean_ctor_set_tag(x_78, 1); +} +lean_ctor_set(x_78, 0, x_73); +lean_ctor_set(x_78, 1, x_74); +if (lean_is_scalar(x_77)) { + x_79 = lean_alloc_ctor(0, 3, 0); +} else { + x_79 = x_77; } +lean_ctor_set(x_79, 0, x_78); +lean_ctor_set(x_79, 1, x_75); +lean_ctor_set(x_79, 2, x_76); +x_80 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_80, 0, x_79); +x_81 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_81, 0, x_80); +lean_ctor_set(x_81, 1, x_12); +return x_81; } +else +{ +lean_object* x_82; lean_object* x_83; +lean_dec(x_70); +lean_dec(x_69); +lean_dec(x_68); +lean_dec(x_3); +lean_dec(x_1); +x_82 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__5; +x_83 = l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___spec__1(x_82, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_83; } } else { -lean_object* x_71; lean_object* x_72; lean_object* x_73; -lean_dec(x_15); -lean_dec(x_14); -x_71 = l_Lean_Expr_appFn_x21(x_1); -x_72 = l_Lean_Expr_appFn_x21(x_2); -x_73 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_x3f(x_3, x_71, x_72, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -lean_dec(x_72); +lean_object* x_84; lean_object* x_85; lean_dec(x_71); -return x_73; +lean_dec(x_70); +lean_dec(x_69); +lean_dec(x_68); +lean_dec(x_3); +lean_dec(x_1); +x_84 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__8; +x_85 = l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___spec__1(x_84, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_85; } } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_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, lean_object* x_11, lean_object* x_12) { -_start: +} +else { -uint8_t x_13; -x_13 = l_Lean_Expr_isApp(x_2); -if (x_13 == 0) +uint8_t x_86; +lean_dec(x_2); +x_86 = !lean_is_exclusive(x_16); +if (x_86 == 0) { -lean_object* x_14; lean_object* x_15; +lean_object* x_87; lean_object* x_88; +x_87 = lean_ctor_get(x_16, 0); +x_88 = l_Lean_Meta_Grind_isEqv(x_87, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -1356,694 +1629,604 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_14 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_14, 0, x_1); -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; -} -else +if (lean_obj_tag(x_88) == 0) { -lean_object* x_16; lean_object* x_17; -x_16 = lean_box(0); -x_17 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_x3f___lambda__1(x_2, x_3, x_1, x_16, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -return x_17; -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_x3f___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { -_start: +lean_object* x_89; uint8_t x_90; +x_89 = lean_ctor_get(x_88, 0); +lean_inc(x_89); +x_90 = lean_unbox(x_89); +lean_dec(x_89); +if (x_90 == 0) { -lean_object* x_14; -x_14 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_x3f___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -lean_dec(x_4); -lean_dec(x_2); +uint8_t x_91; +lean_free_object(x_16); lean_dec(x_1); -return x_14; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_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, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -lean_object* x_13; -x_13 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_x3f(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_3); -lean_dec(x_2); -return x_13; -} -} -LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { -_start: -{ -lean_object* x_15; -x_15 = l_Lean_Meta_Grind_getNext(x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_15) == 0) -{ -uint8_t x_16; -x_16 = !lean_is_exclusive(x_15); -if (x_16 == 0) -{ -lean_object* x_17; uint8_t x_18; -x_17 = lean_ctor_get(x_15, 0); -x_18 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_17, x_1); -if (x_18 == 0) +x_91 = !lean_is_exclusive(x_88); +if (x_91 == 0) { -lean_object* x_19; -x_19 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_19, 0, x_17); -lean_ctor_set(x_15, 0, x_19); -return x_15; +lean_object* x_92; lean_object* x_93; +x_92 = lean_ctor_get(x_88, 0); +lean_dec(x_92); +x_93 = lean_box(0); +lean_ctor_set(x_88, 0, x_93); +return x_88; } else { -lean_object* x_20; -x_20 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_20, 0, x_17); -lean_ctor_set(x_15, 0, x_20); -return x_15; +lean_object* x_94; lean_object* x_95; lean_object* x_96; +x_94 = lean_ctor_get(x_88, 1); +lean_inc(x_94); +lean_dec(x_88); +x_95 = lean_box(0); +x_96 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_96, 0, x_95); +lean_ctor_set(x_96, 1, x_94); +return x_96; } } else { -lean_object* x_21; lean_object* x_22; uint8_t 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 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_21, x_1); -if (x_23 == 0) +uint8_t x_97; +x_97 = !lean_is_exclusive(x_88); +if (x_97 == 0) { -lean_object* x_24; lean_object* x_25; -x_24 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_24, 0, x_21); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_24); -lean_ctor_set(x_25, 1, x_22); -return x_25; +lean_object* x_98; +x_98 = lean_ctor_get(x_88, 0); +lean_dec(x_98); +lean_ctor_set(x_16, 0, x_1); +lean_ctor_set(x_88, 0, x_16); +return x_88; } else { -lean_object* x_26; lean_object* x_27; -x_26 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_26, 0, 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); -return x_27; +lean_object* x_99; lean_object* x_100; +x_99 = lean_ctor_get(x_88, 1); +lean_inc(x_99); +lean_dec(x_88); +lean_ctor_set(x_16, 0, x_1); +x_100 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_100, 0, x_16); +lean_ctor_set(x_100, 1, x_99); +return x_100; } } } else { -uint8_t x_28; -x_28 = !lean_is_exclusive(x_15); -if (x_28 == 0) +uint8_t x_101; +lean_free_object(x_16); +lean_dec(x_1); +x_101 = !lean_is_exclusive(x_88); +if (x_101 == 0) { -return x_15; +return x_88; } else { -lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_29 = lean_ctor_get(x_15, 0); -x_30 = lean_ctor_get(x_15, 1); -lean_inc(x_30); -lean_inc(x_29); -lean_dec(x_15); -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_102; lean_object* x_103; lean_object* x_104; +x_102 = lean_ctor_get(x_88, 0); +x_103 = lean_ctor_get(x_88, 1); +lean_inc(x_103); +lean_inc(x_102); +lean_dec(x_88); +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_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18) { -_start: +else { -lean_object* x_19; lean_object* x_35; -lean_inc(x_7); -x_35 = l_Lean_Meta_Grind_getENode(x_7, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); -if (lean_obj_tag(x_35) == 0) +lean_object* x_105; lean_object* x_106; +x_105 = lean_ctor_get(x_16, 0); +lean_inc(x_105); +lean_dec(x_16); +x_106 = l_Lean_Meta_Grind_isEqv(x_105, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +if (lean_obj_tag(x_106) == 0) { -lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; uint8_t x_109; -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, 8); -lean_inc(x_38); -x_109 = lean_nat_dec_le(x_38, x_4); -if (x_109 == 0) +lean_object* x_107; uint8_t x_108; +x_107 = lean_ctor_get(x_106, 0); +lean_inc(x_107); +x_108 = lean_unbox(x_107); +lean_dec(x_107); +if (x_108 == 0) { -lean_object* x_110; lean_object* x_111; -lean_dec(x_38); -lean_dec(x_36); -x_110 = lean_box(0); -x_111 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_3, x_7, x_110, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_37); -x_19 = x_111; -goto block_34; +lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; +lean_dec(x_1); +x_109 = lean_ctor_get(x_106, 1); +lean_inc(x_109); +if (lean_is_exclusive(x_106)) { + lean_ctor_release(x_106, 0); + lean_ctor_release(x_106, 1); + x_110 = x_106; +} else { + lean_dec_ref(x_106); + x_110 = lean_box(0); +} +x_111 = lean_box(0); +if (lean_is_scalar(x_110)) { + x_112 = lean_alloc_ctor(0, 2, 0); +} else { + x_112 = x_110; +} +lean_ctor_set(x_112, 0, x_111); +lean_ctor_set(x_112, 1, x_109); +return x_112; } else { -uint8_t x_112; -x_112 = lean_ctor_get_uint8(x_36, sizeof(void*)*10 + 4); -if (x_112 == 0) -{ -lean_object* x_113; uint8_t x_114; -x_113 = lean_ctor_get(x_36, 3); +lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; +x_113 = lean_ctor_get(x_106, 1); lean_inc(x_113); -lean_dec(x_36); -x_114 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_7, x_113); -lean_dec(x_113); -if (x_114 == 0) -{ -lean_object* x_115; lean_object* x_116; -lean_dec(x_38); -x_115 = lean_box(0); -x_116 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_3, x_7, x_115, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_37); -x_19 = x_116; -goto block_34; +if (lean_is_exclusive(x_106)) { + lean_ctor_release(x_106, 0); + lean_ctor_release(x_106, 1); + x_114 = x_106; +} else { + lean_dec_ref(x_106); + x_114 = lean_box(0); +} +x_115 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_115, 0, x_1); +if (lean_is_scalar(x_114)) { + x_116 = lean_alloc_ctor(0, 2, 0); +} else { + x_116 = x_114; +} +lean_ctor_set(x_116, 0, x_115); +lean_ctor_set(x_116, 1, x_113); +return x_116; } -else -{ -lean_object* x_117; uint8_t x_118; -x_117 = l_Lean_Expr_getAppFn(x_7); -x_118 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_eqvFunctions(x_5, x_117); -lean_dec(x_117); -if (x_118 == 0) -{ -lean_object* x_119; lean_object* x_120; -lean_dec(x_38); -x_119 = lean_box(0); -x_120 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_3, x_7, x_119, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_37); -x_19 = x_120; -goto block_34; } else { -lean_object* x_121; lean_object* x_122; uint8_t x_123; -x_121 = lean_unsigned_to_nat(0u); -x_122 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_7, x_121); -x_123 = lean_nat_dec_eq(x_122, x_6); -lean_dec(x_122); -if (x_123 == 0) -{ -lean_object* x_124; lean_object* x_125; -lean_dec(x_38); -x_124 = lean_box(0); -x_125 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_3, x_7, x_124, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_37); -x_19 = x_125; -goto block_34; +lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; +lean_dec(x_1); +x_117 = lean_ctor_get(x_106, 0); +lean_inc(x_117); +x_118 = lean_ctor_get(x_106, 1); +lean_inc(x_118); +if (lean_is_exclusive(x_106)) { + lean_ctor_release(x_106, 0); + lean_ctor_release(x_106, 1); + x_119 = x_106; +} else { + lean_dec_ref(x_106); + x_119 = lean_box(0); } -else -{ -lean_object* x_126; -x_126 = lean_box(0); -x_39 = x_126; -goto block_108; +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_127; uint8_t x_128; -lean_dec(x_36); -x_127 = l_Lean_Expr_getAppFn(x_7); -x_128 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_eqvFunctions(x_5, x_127); -lean_dec(x_127); -if (x_128 == 0) -{ -lean_object* x_129; lean_object* x_130; -lean_dec(x_38); -x_129 = lean_box(0); -x_130 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_3, x_7, x_129, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_37); -x_19 = x_130; -goto block_34; } else { -lean_object* x_131; lean_object* x_132; uint8_t x_133; -x_131 = lean_unsigned_to_nat(0u); -x_132 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_7, x_131); -x_133 = lean_nat_dec_eq(x_132, x_6); -lean_dec(x_132); -if (x_133 == 0) -{ -lean_object* x_134; lean_object* x_135; -lean_dec(x_38); -x_134 = lean_box(0); -x_135 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_3, x_7, x_134, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_37); -x_19 = x_135; -goto block_34; +lean_object* x_121; lean_object* x_122; +x_121 = l_Lean_Expr_bvarIdx_x21(x_2); +lean_dec(x_2); +x_122 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f(x_1, x_121, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_121); +return x_122; +} } else { -lean_object* x_136; -x_136 = lean_box(0); -x_39 = x_136; -goto block_108; -} -} +lean_object* x_123; lean_object* x_124; +lean_dec(x_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_123 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_123, 0, x_1); +x_124 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_124, 0, x_123); +lean_ctor_set(x_124, 1, x_12); +return x_124; } } -block_108: -{ -lean_object* x_40; -lean_dec(x_39); -lean_inc(x_17); -lean_inc(x_16); -lean_inc(x_15); -lean_inc(x_14); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_1); -x_40 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_x3f(x_1, x_2, x_7, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_37); -if (lean_obj_tag(x_40) == 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_42; lean_object* x_43; lean_object* x_44; -lean_dec(x_38); -x_42 = lean_ctor_get(x_40, 1); -lean_inc(x_42); -lean_dec(x_40); -x_43 = lean_box(0); -x_44 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_3, x_7, x_43, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_42); -x_19 = x_44; -goto block_34; } -else -{ -lean_object* x_45; lean_object* x_46; uint8_t x_47; -x_45 = lean_ctor_get(x_41, 0); -lean_inc(x_45); -lean_dec(x_41); -x_46 = lean_ctor_get(x_40, 1); -lean_inc(x_46); -lean_dec(x_40); -x_47 = !lean_is_exclusive(x_45); -if (x_47 == 0) +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_Choice_updateGen(lean_object* x_1, lean_object* x_2) { +_start: { -lean_object* x_48; uint8_t x_49; -x_48 = lean_ctor_get(x_45, 1); -x_49 = lean_nat_dec_le(x_38, x_48); -if (x_49 == 0) +uint8_t x_3; +x_3 = !lean_is_exclusive(x_1); +if (x_3 == 0) { -lean_object* x_50; uint8_t x_51; -lean_dec(x_48); -lean_ctor_set(x_45, 1, x_38); -x_50 = lean_st_ref_take(x_9, x_46); -x_51 = !lean_is_exclusive(x_50); -if (x_51 == 0) +lean_object* x_4; uint8_t x_5; +x_4 = lean_ctor_get(x_1, 1); +x_5 = lean_nat_dec_le(x_2, x_4); +if (x_5 == 0) { -lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; -x_52 = lean_ctor_get(x_50, 0); -x_53 = lean_ctor_get(x_50, 1); -lean_ctor_set_tag(x_50, 1); -lean_ctor_set(x_50, 1, x_52); -lean_ctor_set(x_50, 0, x_45); -x_54 = lean_st_ref_set(x_9, x_50, x_53); -x_55 = lean_ctor_get(x_54, 1); -lean_inc(x_55); -lean_dec(x_54); -x_56 = lean_box(0); -x_57 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_3, x_7, x_56, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_55); -x_19 = x_57; -goto block_34; +lean_dec(x_4); +lean_ctor_set(x_1, 1, x_2); +return x_1; } 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; -x_58 = lean_ctor_get(x_50, 0); -x_59 = lean_ctor_get(x_50, 1); -lean_inc(x_59); -lean_inc(x_58); -lean_dec(x_50); -x_60 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_60, 0, x_45); -lean_ctor_set(x_60, 1, x_58); -x_61 = lean_st_ref_set(x_9, x_60, x_59); -x_62 = lean_ctor_get(x_61, 1); -lean_inc(x_62); -lean_dec(x_61); -x_63 = lean_box(0); -x_64 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_3, x_7, x_63, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_62); -x_19 = x_64; -goto block_34; +lean_dec(x_2); +return x_1; } } else { -lean_object* x_65; uint8_t x_66; -lean_dec(x_38); -x_65 = lean_st_ref_take(x_9, x_46); -x_66 = !lean_is_exclusive(x_65); -if (x_66 == 0) +lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; +x_6 = lean_ctor_get(x_1, 0); +x_7 = lean_ctor_get(x_1, 1); +x_8 = lean_ctor_get(x_1, 2); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_dec(x_1); +x_9 = lean_nat_dec_le(x_2, x_7); +if (x_9 == 0) { -lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; -x_67 = lean_ctor_get(x_65, 0); -x_68 = lean_ctor_get(x_65, 1); -lean_ctor_set_tag(x_65, 1); -lean_ctor_set(x_65, 1, x_67); -lean_ctor_set(x_65, 0, x_45); -x_69 = lean_st_ref_set(x_9, x_65, x_68); -x_70 = lean_ctor_get(x_69, 1); -lean_inc(x_70); -lean_dec(x_69); -x_71 = lean_box(0); -x_72 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_3, x_7, x_71, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_70); -x_19 = x_72; -goto block_34; +lean_object* x_10; +lean_dec(x_7); +x_10 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_10, 0, x_6); +lean_ctor_set(x_10, 1, x_2); +lean_ctor_set(x_10, 2, x_8); +return x_10; } 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; -x_73 = lean_ctor_get(x_65, 0); -x_74 = lean_ctor_get(x_65, 1); -lean_inc(x_74); -lean_inc(x_73); -lean_dec(x_65); -x_75 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_75, 0, x_45); -lean_ctor_set(x_75, 1, x_73); -x_76 = lean_st_ref_set(x_9, x_75, x_74); -x_77 = lean_ctor_get(x_76, 1); -lean_inc(x_77); -lean_dec(x_76); -x_78 = lean_box(0); -x_79 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_3, x_7, x_78, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_77); -x_19 = x_79; -goto block_34; +lean_object* x_11; +lean_dec(x_2); +x_11 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_11, 0, x_6); +lean_ctor_set(x_11, 1, x_7); +lean_ctor_set(x_11, 2, x_8); +return x_11; } } } -else +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_pushChoice(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, 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_80; lean_object* x_81; lean_object* x_82; uint8_t x_83; -x_80 = lean_ctor_get(x_45, 0); -x_81 = lean_ctor_get(x_45, 1); -x_82 = lean_ctor_get(x_45, 2); -lean_inc(x_82); -lean_inc(x_81); -lean_inc(x_80); -lean_dec(x_45); -x_83 = lean_nat_dec_le(x_38, x_81); -if (x_83 == 0) +lean_object* x_13; uint8_t x_14; +x_13 = lean_st_ref_take(x_3, x_12); +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 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_dec(x_81); -x_84 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_84, 0, x_80); -lean_ctor_set(x_84, 1, x_38); -lean_ctor_set(x_84, 2, x_82); -x_85 = lean_st_ref_take(x_9, x_46); -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); +lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_15 = lean_ctor_get(x_13, 0); +x_16 = lean_ctor_get(x_13, 1); +lean_ctor_set_tag(x_13, 1); +lean_ctor_set(x_13, 1, x_15); +lean_ctor_set(x_13, 0, x_1); +x_17 = lean_st_ref_set(x_3, x_13, x_16); +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; } -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); +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; } -lean_ctor_set(x_89, 0, x_84); -lean_ctor_set(x_89, 1, x_86); -x_90 = lean_st_ref_set(x_9, x_89, x_87); -x_91 = lean_ctor_get(x_90, 1); -lean_inc(x_91); -lean_dec(x_90); -x_92 = lean_box(0); -x_93 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_3, x_7, x_92, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_91); -x_19 = x_93; -goto block_34; } 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_dec(x_38); -x_94 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_94, 0, x_80); -lean_ctor_set(x_94, 1, x_81); -lean_ctor_set(x_94, 2, x_82); -x_95 = lean_st_ref_take(x_9, x_46); -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; +lean_object* x_24; lean_object* x_25; 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_13, 0); +x_25 = lean_ctor_get(x_13, 1); +lean_inc(x_25); +lean_inc(x_24); +lean_dec(x_13); +x_26 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_26, 0, x_1); +lean_ctor_set(x_26, 1, x_24); +x_27 = lean_st_ref_set(x_3, x_26, x_25); +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_95); - x_98 = lean_box(0); + lean_dec_ref(x_27); + x_29 = lean_box(0); } -if (lean_is_scalar(x_98)) { - x_99 = lean_alloc_ctor(1, 2, 0); +x_30 = lean_box(0); +if (lean_is_scalar(x_29)) { + x_31 = lean_alloc_ctor(0, 2, 0); } else { - x_99 = x_98; - lean_ctor_set_tag(x_99, 1); + x_31 = x_29; } -lean_ctor_set(x_99, 0, x_94); -lean_ctor_set(x_99, 1, x_96); -x_100 = lean_st_ref_set(x_9, x_99, x_97); -x_101 = lean_ctor_get(x_100, 1); -lean_inc(x_101); -lean_dec(x_100); -x_102 = lean_box(0); -x_103 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_3, x_7, x_102, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_101); -x_19 = x_103; -goto block_34; +lean_ctor_set(x_31, 0, x_30); +lean_ctor_set(x_31, 1, x_28); +return x_31; } } } +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_pushChoice___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +x_13 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_pushChoice(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_13; +} } -else +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_x3f___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: { -uint8_t x_104; -lean_dec(x_38); +lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_14 = l_Lean_Expr_appArg_x21(x_1); +x_15 = l_Lean_Expr_appArg_x21(x_2); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +x_16 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f(x_3, x_14, x_15, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_17; +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +if (lean_obj_tag(x_17) == 0) +{ +uint8_t 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); -x_104 = !lean_is_exclusive(x_40); -if (x_104 == 0) +lean_dec(x_6); +lean_dec(x_5); +x_18 = !lean_is_exclusive(x_16); +if (x_18 == 0) { -x_19 = x_40; -goto block_34; +lean_object* x_19; lean_object* x_20; +x_19 = lean_ctor_get(x_16, 0); +lean_dec(x_19); +x_20 = lean_box(0); +lean_ctor_set(x_16, 0, x_20); +return x_16; } else { -lean_object* x_105; lean_object* x_106; lean_object* x_107; -x_105 = lean_ctor_get(x_40, 0); -x_106 = lean_ctor_get(x_40, 1); -lean_inc(x_106); -lean_inc(x_105); -lean_dec(x_40); -x_107 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_107, 0, x_105); -lean_ctor_set(x_107, 1, x_106); -x_19 = x_107; -goto block_34; +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_16, 1); +lean_inc(x_21); +lean_dec(x_16); +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; +x_24 = lean_ctor_get(x_16, 1); +lean_inc(x_24); +lean_dec(x_16); +x_25 = lean_ctor_get(x_17, 0); +lean_inc(x_25); +lean_dec(x_17); +x_26 = l_Lean_Expr_appFn_x21(x_1); +x_27 = l_Lean_Expr_appFn_x21(x_2); +x_28 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_x3f(x_25, x_26, x_27, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_24); +lean_dec(x_27); +lean_dec(x_26); +return x_28; } } else { -uint8_t x_137; +uint8_t x_29; +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_7); -x_137 = !lean_is_exclusive(x_35); -if (x_137 == 0) +lean_dec(x_6); +lean_dec(x_5); +x_29 = !lean_is_exclusive(x_16); +if (x_29 == 0) { -x_19 = x_35; -goto block_34; +return x_16; } else { -lean_object* x_138; lean_object* x_139; lean_object* x_140; -x_138 = lean_ctor_get(x_35, 0); -x_139 = lean_ctor_get(x_35, 1); -lean_inc(x_139); -lean_inc(x_138); -lean_dec(x_35); -x_140 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_140, 0, x_138); -lean_ctor_set(x_140, 1, x_139); -x_19 = x_140; -goto block_34; +lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_30 = lean_ctor_get(x_16, 0); +x_31 = lean_ctor_get(x_16, 1); +lean_inc(x_31); +lean_inc(x_30); +lean_dec(x_16); +x_32 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_32, 0, x_30); +lean_ctor_set(x_32, 1, x_31); +return x_32; } } -block_34: -{ -if (lean_obj_tag(x_19) == 0) +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_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, lean_object* x_11, lean_object* x_12) { +_start: { -lean_object* x_20; -x_20 = lean_ctor_get(x_19, 0); -lean_inc(x_20); -if (lean_obj_tag(x_20) == 0) +uint8_t x_13; +x_13 = l_Lean_Expr_isApp(x_2); +if (x_13 == 0) { -uint8_t x_21; -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_object* x_14; lean_object* x_15; lean_dec(x_11); lean_dec(x_10); -lean_dec(x_1); -x_21 = !lean_is_exclusive(x_19); -if (x_21 == 0) -{ -lean_object* x_22; lean_object* x_23; -x_22 = lean_ctor_get(x_19, 0); -lean_dec(x_22); -x_23 = lean_ctor_get(x_20, 0); -lean_inc(x_23); -lean_dec(x_20); -lean_ctor_set(x_19, 0, x_23); -return x_19; +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_14 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_14, 0, x_1); +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; } else { -lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_24 = lean_ctor_get(x_19, 1); -lean_inc(x_24); -lean_dec(x_19); -x_25 = lean_ctor_get(x_20, 0); -lean_inc(x_25); -lean_dec(x_20); -x_26 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_26, 0, x_25); -lean_ctor_set(x_26, 1, x_24); -return x_26; -} +lean_object* x_16; lean_object* x_17; +x_16 = lean_box(0); +x_17 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_x3f___lambda__1(x_2, x_3, x_1, x_16, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_17; } -else -{ -lean_object* x_27; lean_object* x_28; -x_27 = lean_ctor_get(x_19, 1); -lean_inc(x_27); -lean_dec(x_19); -x_28 = lean_ctor_get(x_20, 0); -lean_inc(x_28); -lean_dec(x_20); -x_7 = x_28; -x_18 = x_27; -goto _start; } } -else +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_x3f___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: { -uint8_t x_30; -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); +lean_object* x_14; +x_14 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_x3f___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_4); +lean_dec(x_2); lean_dec(x_1); -x_30 = !lean_is_exclusive(x_19); -if (x_30 == 0) -{ -return x_19; +return x_14; } -else +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_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, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: { -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; +lean_object* x_13; +x_13 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_x3f(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_3); +lean_dec(x_2); +return x_13; } } +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +lean_object* x_15; +x_15 = l_Lean_Meta_Grind_getNext(x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_15) == 0) +{ +uint8_t x_16; +x_16 = !lean_is_exclusive(x_15); +if (x_16 == 0) +{ +lean_object* x_17; uint8_t x_18; +x_17 = lean_ctor_get(x_15, 0); +x_18 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_17, x_1); +if (x_18 == 0) +{ +lean_object* x_19; +x_19 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_19, 0, x_17); +lean_ctor_set(x_15, 0, x_19); +return x_15; } +else +{ +lean_object* x_20; +x_20 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_20, 0, x_17); +lean_ctor_set(x_15, 0, x_20); +return x_15; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { -_start: +else { -lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_15 = l_Lean_Meta_Grind_getMaxGeneration___rarg(x_8, x_9, x_10, x_11, x_12, x_13, x_14); -x_16 = lean_ctor_get(x_15, 0); -lean_inc(x_16); -x_17 = lean_ctor_get(x_15, 1); -lean_inc(x_17); +lean_object* x_21; lean_object* x_22; uint8_t 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_18 = l_Lean_Expr_getAppFn(x_2); -x_19 = lean_unsigned_to_nat(0u); -x_20 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_2, x_19); -lean_inc(x_3); -x_21 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1(x_1, x_2, x_3, x_16, x_18, x_20, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_17); -lean_dec(x_20); -lean_dec(x_18); -lean_dec(x_16); -lean_dec(x_3); -if (lean_obj_tag(x_21) == 0) -{ -uint8_t x_22; -x_22 = !lean_is_exclusive(x_21); -if (x_22 == 0) +x_23 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_21, x_1); +if (x_23 == 0) { -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; +lean_object* x_24; lean_object* x_25; +x_24 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_24, 0, x_21); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_22); +return x_25; } 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 = lean_box(0); +lean_object* x_26; lean_object* x_27; +x_26 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_26, 0, 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_25); +lean_ctor_set(x_27, 1, x_22); return x_27; } } +} else { uint8_t x_28; -x_28 = !lean_is_exclusive(x_21); +x_28 = !lean_is_exclusive(x_15); if (x_28 == 0) { -return x_21; +return x_15; } 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); +x_29 = lean_ctor_get(x_15, 0); +x_30 = lean_ctor_get(x_15, 1); lean_inc(x_30); lean_inc(x_29); -lean_dec(x_21); +lean_dec(x_15); x_31 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_31, 0, x_29); lean_ctor_set(x_31, 1, x_30); @@ -2052,563 +2235,577 @@ return x_31; } } } -LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18) { _start: { -lean_object* x_15; -x_15 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +lean_object* x_19; lean_object* x_20; lean_object* x_26; +lean_inc(x_7); +x_26 = l_Lean_Meta_Grind_getENode(x_7, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +if (lean_obj_tag(x_26) == 0) +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_81; +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, 8); +lean_inc(x_29); +x_81 = lean_nat_dec_lt(x_29, x_4); +if (x_81 == 0) +{ +lean_object* x_82; lean_object* x_83; +lean_dec(x_29); +lean_dec(x_27); +x_82 = lean_box(0); +x_83 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_3, x_7, x_82, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_28); +if (lean_obj_tag(x_83) == 0) +{ +lean_object* x_84; lean_object* x_85; +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_19 = x_84; +x_20 = x_85; +goto block_25; +} +else +{ +uint8_t x_86; +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); lean_dec(x_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); -return x_15; +x_86 = !lean_is_exclusive(x_83); +if (x_86 == 0) +{ +return x_83; } +else +{ +lean_object* x_87; lean_object* x_88; lean_object* x_89; +x_87 = lean_ctor_get(x_83, 0); +x_88 = lean_ctor_get(x_83, 1); +lean_inc(x_88); +lean_inc(x_87); +lean_dec(x_83); +x_89 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_89, 0, x_87); +lean_ctor_set(x_89, 1, x_88); +return x_89; } -LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___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]; -_start: +} +} +else { -lean_object* x_19; -x_19 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_19; +uint8_t x_90; +x_90 = lean_ctor_get_uint8(x_27, sizeof(void*)*10 + 4); +if (x_90 == 0) +{ +uint8_t x_91; +x_91 = l_Lean_Meta_Grind_ENode_isCongrRoot(x_27); +lean_dec(x_27); +if (x_91 == 0) +{ +lean_object* x_92; lean_object* x_93; +lean_dec(x_29); +x_92 = lean_box(0); +x_93 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_3, x_7, x_92, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_28); +if (lean_obj_tag(x_93) == 0) +{ +lean_object* x_94; lean_object* x_95; +x_94 = lean_ctor_get(x_93, 0); +lean_inc(x_94); +x_95 = lean_ctor_get(x_93, 1); +lean_inc(x_95); +lean_dec(x_93); +x_19 = x_94; +x_20 = x_95; +goto block_25; } +else +{ +uint8_t x_96; +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_1); +x_96 = !lean_is_exclusive(x_93); +if (x_96 == 0) +{ +return x_93; } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { -_start: +else { -lean_object* x_15; -x_15 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_2); -return x_15; +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; } } -static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___spec__1___closed__1() { -_start: +} +else { -lean_object* x_1; lean_object* x_2; -x_1 = lean_box(0); -x_2 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; +lean_object* x_100; +x_100 = lean_box(0); +x_30 = x_100; +goto block_80; } } -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19, lean_object* x_20) { -_start: +else { -if (lean_obj_tag(x_7) == 0) +lean_object* x_101; +lean_dec(x_27); +x_101 = lean_box(0); +x_30 = x_101; +goto block_80; +} +} +block_80: { -lean_object* x_21; -lean_dec(x_19); -lean_dec(x_18); +lean_object* x_31; uint8_t x_32; +lean_dec(x_30); +x_31 = l_Lean_Expr_getAppFn(x_7); +x_32 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_eqvFunctions(x_5, x_31); +lean_dec(x_31); +if (x_32 == 0) +{ +lean_object* x_33; lean_object* x_34; +lean_dec(x_29); +x_33 = lean_box(0); +x_34 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_3, x_7, x_33, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_28); +if (lean_obj_tag(x_34) == 0) +{ +lean_object* x_35; lean_object* x_36; +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_19 = x_35; +x_20 = x_36; +goto block_25; +} +else +{ +uint8_t x_37; lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); lean_dec(x_1); -x_21 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_21, 0, x_8); -lean_ctor_set(x_21, 1, x_20); -return x_21; +x_37 = !lean_is_exclusive(x_34); +if (x_37 == 0) +{ +return x_34; } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_30; -lean_dec(x_8); -x_22 = lean_ctor_get(x_7, 0); -lean_inc(x_22); -x_23 = lean_ctor_get(x_7, 1); -lean_inc(x_23); -if (lean_is_exclusive(x_7)) { - lean_ctor_release(x_7, 0); - lean_ctor_release(x_7, 1); - x_24 = x_7; -} else { - lean_dec_ref(x_7); - x_24 = lean_box(0); +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_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; +} } -lean_inc(x_22); -x_30 = l_Lean_Meta_Grind_getENode(x_22, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20); -if (lean_obj_tag(x_30) == 0) -{ -lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_83; -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, 8); -lean_inc(x_33); -x_83 = lean_nat_dec_le(x_33, x_4); -if (x_83 == 0) -{ -lean_object* x_84; -lean_dec(x_33); -lean_dec(x_31); -lean_dec(x_24); -lean_dec(x_22); -x_84 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___spec__1___closed__1; -x_25 = x_84; -x_26 = x_32; -goto block_29; } else { -uint8_t x_85; -x_85 = lean_ctor_get_uint8(x_31, sizeof(void*)*10 + 4); -if (x_85 == 0) +lean_object* x_41; lean_object* x_42; uint8_t x_43; +x_41 = lean_unsigned_to_nat(0u); +x_42 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_7, x_41); +x_43 = lean_nat_dec_eq(x_42, x_6); +lean_dec(x_42); +if (x_43 == 0) { -lean_object* x_86; uint8_t x_87; -x_86 = lean_ctor_get(x_31, 3); -lean_inc(x_86); -lean_dec(x_31); -x_87 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_86, x_22); -lean_dec(x_86); -if (x_87 == 0) +lean_object* x_44; lean_object* x_45; +lean_dec(x_29); +x_44 = lean_box(0); +x_45 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_3, x_7, x_44, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_28); +if (lean_obj_tag(x_45) == 0) { -lean_object* x_88; -lean_dec(x_33); -lean_dec(x_24); -lean_dec(x_22); -x_88 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___spec__1___closed__1; -x_25 = x_88; -x_26 = x_32; -goto block_29; +lean_object* x_46; lean_object* x_47; +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_19 = x_46; +x_20 = x_47; +goto block_25; } else { -lean_object* x_89; -x_89 = lean_box(0); -x_34 = x_89; -goto block_82; -} +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_dec(x_10); +lean_dec(x_1); +x_48 = !lean_is_exclusive(x_45); +if (x_48 == 0) +{ +return x_45; } else { -lean_object* x_90; -lean_dec(x_31); -x_90 = lean_box(0); -x_34 = x_90; -goto block_82; +lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_49 = lean_ctor_get(x_45, 0); +x_50 = lean_ctor_get(x_45, 1); +lean_inc(x_50); +lean_inc(x_49); +lean_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; } } -block_82: +} +else { -lean_object* x_35; -lean_dec(x_34); -lean_inc(x_19); -lean_inc(x_18); +lean_object* x_52; lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); lean_inc(x_1); -x_35 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_x3f(x_1, x_2, x_22, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_32); -lean_dec(x_22); -if (lean_obj_tag(x_35) == 0) +x_52 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_x3f(x_1, x_2, x_7, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_28); +if (lean_obj_tag(x_52) == 0) { -lean_object* x_36; -x_36 = lean_ctor_get(x_35, 0); -lean_inc(x_36); -if (lean_obj_tag(x_36) == 0) +lean_object* x_53; +x_53 = lean_ctor_get(x_52, 0); +lean_inc(x_53); +if (lean_obj_tag(x_53) == 0) { -lean_object* x_37; lean_object* x_38; -lean_dec(x_33); -lean_dec(x_24); -x_37 = lean_ctor_get(x_35, 1); -lean_inc(x_37); -lean_dec(x_35); -x_38 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___spec__1___closed__1; -x_25 = x_38; -x_26 = x_37; -goto block_29; +lean_object* x_54; lean_object* x_55; lean_object* x_56; +lean_dec(x_29); +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_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_3, x_7, x_55, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_54); +if (lean_obj_tag(x_56) == 0) +{ +lean_object* x_57; lean_object* x_58; +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_19 = x_57; +x_20 = x_58; +goto block_25; } else { -lean_object* x_39; lean_object* x_40; uint8_t x_41; -x_39 = lean_ctor_get(x_36, 0); -lean_inc(x_39); -lean_dec(x_36); -x_40 = lean_ctor_get(x_35, 1); -lean_inc(x_40); -lean_dec(x_35); -x_41 = !lean_is_exclusive(x_39); -if (x_41 == 0) -{ -lean_object* x_42; uint8_t x_43; -x_42 = lean_ctor_get(x_39, 1); -x_43 = lean_nat_dec_le(x_33, x_42); -if (x_43 == 0) +uint8_t x_59; +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_1); +x_59 = !lean_is_exclusive(x_56); +if (x_59 == 0) { -lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; -lean_dec(x_42); -lean_ctor_set(x_39, 1, x_33); -x_44 = lean_st_ref_take(x_11, x_40); -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); -if (lean_is_scalar(x_24)) { - x_47 = lean_alloc_ctor(1, 2, 0); -} else { - x_47 = x_24; -} -lean_ctor_set(x_47, 0, x_39); -lean_ctor_set(x_47, 1, x_45); -x_48 = lean_st_ref_set(x_11, x_47, x_46); -x_49 = lean_ctor_get(x_48, 1); -lean_inc(x_49); -lean_dec(x_48); -x_50 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___spec__1___closed__1; -x_25 = x_50; -x_26 = x_49; -goto block_29; +return x_56; } 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_dec(x_33); -x_51 = lean_st_ref_take(x_11, x_40); -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); -if (lean_is_scalar(x_24)) { - x_54 = lean_alloc_ctor(1, 2, 0); -} else { - x_54 = x_24; +lean_object* x_60; lean_object* x_61; lean_object* x_62; +x_60 = lean_ctor_get(x_56, 0); +x_61 = lean_ctor_get(x_56, 1); +lean_inc(x_61); +lean_inc(x_60); +lean_dec(x_56); +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_ctor_set(x_54, 0, x_39); -lean_ctor_set(x_54, 1, x_52); -x_55 = lean_st_ref_set(x_11, x_54, x_53); -x_56 = lean_ctor_get(x_55, 1); -lean_inc(x_56); -lean_dec(x_55); -x_57 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___spec__1___closed__1; -x_25 = x_57; -x_26 = x_56; -goto block_29; } } else { -lean_object* x_58; lean_object* x_59; lean_object* x_60; uint8_t x_61; -x_58 = lean_ctor_get(x_39, 0); -x_59 = lean_ctor_get(x_39, 1); -x_60 = lean_ctor_get(x_39, 2); -lean_inc(x_60); -lean_inc(x_59); -lean_inc(x_58); -lean_dec(x_39); -x_61 = lean_nat_dec_le(x_33, x_59); -if (x_61 == 0) -{ -lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; -lean_dec(x_59); -x_62 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_62, 0, x_58); -lean_ctor_set(x_62, 1, x_33); -lean_ctor_set(x_62, 2, x_60); -x_63 = lean_st_ref_take(x_11, x_40); -x_64 = lean_ctor_get(x_63, 0); +lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; +x_63 = lean_ctor_get(x_52, 1); +lean_inc(x_63); +lean_dec(x_52); +x_64 = lean_ctor_get(x_53, 0); lean_inc(x_64); -x_65 = lean_ctor_get(x_63, 1); -lean_inc(x_65); -lean_dec(x_63); -if (lean_is_scalar(x_24)) { - x_66 = lean_alloc_ctor(1, 2, 0); -} else { - x_66 = x_24; -} -lean_ctor_set(x_66, 0, x_62); -lean_ctor_set(x_66, 1, x_64); -x_67 = lean_st_ref_set(x_11, x_66, x_65); -x_68 = lean_ctor_get(x_67, 1); +lean_dec(x_53); +x_65 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_Choice_updateGen(x_64, x_29); +x_66 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_pushChoice(x_65, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_63); +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_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_3, x_7, x_67, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_68); lean_dec(x_67); -x_69 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___spec__1___closed__1; -x_25 = x_69; -x_26 = x_68; -goto block_29; +if (lean_obj_tag(x_69) == 0) +{ +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_inc(x_71); +lean_dec(x_69); +x_19 = x_70; +x_20 = x_71; +goto block_25; } 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_33); -x_70 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_70, 0, x_58); -lean_ctor_set(x_70, 1, x_59); -lean_ctor_set(x_70, 2, x_60); -x_71 = lean_st_ref_take(x_11, x_40); -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); -if (lean_is_scalar(x_24)) { - x_74 = lean_alloc_ctor(1, 2, 0); -} else { - x_74 = x_24; +uint8_t x_72; +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_1); +x_72 = !lean_is_exclusive(x_69); +if (x_72 == 0) +{ +return x_69; } -lean_ctor_set(x_74, 0, x_70); -lean_ctor_set(x_74, 1, x_72); -x_75 = lean_st_ref_set(x_11, x_74, x_73); -x_76 = lean_ctor_get(x_75, 1); -lean_inc(x_76); -lean_dec(x_75); -x_77 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___spec__1___closed__1; -x_25 = x_77; -x_26 = x_76; -goto block_29; +else +{ +lean_object* x_73; lean_object* x_74; lean_object* x_75; +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 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_75, 0, x_73); +lean_ctor_set(x_75, 1, x_74); +return x_75; } } } } else { -uint8_t x_78; -lean_dec(x_33); -lean_dec(x_24); -lean_dec(x_23); -lean_dec(x_19); -lean_dec(x_18); +uint8_t x_76; +lean_dec(x_29); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_7); lean_dec(x_1); -x_78 = !lean_is_exclusive(x_35); -if (x_78 == 0) +x_76 = !lean_is_exclusive(x_52); +if (x_76 == 0) { -return x_35; +return x_52; } 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_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_object* x_77; lean_object* x_78; lean_object* x_79; +x_77 = lean_ctor_get(x_52, 0); +x_78 = lean_ctor_get(x_52, 1); +lean_inc(x_78); +lean_inc(x_77); +lean_dec(x_52); +x_79 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_79, 0, x_77); +lean_ctor_set(x_79, 1, x_78); +return x_79; +} +} } } } } else { -uint8_t x_91; -lean_dec(x_24); -lean_dec(x_23); -lean_dec(x_22); -lean_dec(x_19); -lean_dec(x_18); +uint8_t x_102; lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_7); lean_dec(x_1); -x_91 = !lean_is_exclusive(x_30); -if (x_91 == 0) +x_102 = !lean_is_exclusive(x_26); +if (x_102 == 0) { -return x_30; +return x_26; } else { -lean_object* x_92; lean_object* x_93; lean_object* x_94; -x_92 = lean_ctor_get(x_30, 0); -x_93 = lean_ctor_get(x_30, 1); -lean_inc(x_93); -lean_inc(x_92); -lean_dec(x_30); -x_94 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_94, 0, x_92); -lean_ctor_set(x_94, 1, x_93); -return x_94; +lean_object* x_103; lean_object* x_104; lean_object* x_105; +x_103 = lean_ctor_get(x_26, 0); +x_104 = lean_ctor_get(x_26, 1); +lean_inc(x_104); +lean_inc(x_103); +lean_dec(x_26); +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; } } -block_29: +block_25: { -lean_object* x_27; -x_27 = lean_ctor_get(x_25, 0); -lean_inc(x_27); -lean_dec(x_25); +if (lean_obj_tag(x_19) == 0) +{ +lean_object* x_21; lean_object* x_22; +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_1); +x_21 = lean_ctor_get(x_19, 0); +lean_inc(x_21); +lean_dec(x_19); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_20); +return x_22; +} +else +{ +lean_object* x_23; +x_23 = lean_ctor_get(x_19, 0); +lean_inc(x_23); +lean_dec(x_19); x_7 = x_23; -x_8 = x_27; -x_9 = lean_box(0); -x_20 = x_26; +x_18 = x_20; goto _start; } } } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_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_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { _start: { -lean_object* x_14; uint8_t x_15; -x_14 = lean_st_ref_get(x_5, x_13); -x_15 = !lean_is_exclusive(x_14); -if (x_15 == 0) -{ -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_16 = lean_ctor_get(x_14, 0); -x_17 = lean_ctor_get(x_14, 1); -x_18 = lean_ctor_get(x_16, 4); -lean_inc(x_18); +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_15 = l_Lean_Meta_Grind_getMaxGeneration___rarg(x_8, x_9, x_10, x_11, x_12, x_13, x_14); +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_15, 1); +lean_inc(x_17); +lean_dec(x_15); +x_18 = l_Lean_Expr_getAppFn(x_2); +x_19 = lean_unsigned_to_nat(0u); +x_20 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_2, x_19); +lean_inc(x_3); +x_21 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1(x_1, x_2, x_3, x_16, x_18, x_20, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_17); +lean_dec(x_20); +lean_dec(x_18); lean_dec(x_16); -lean_inc(x_2); -x_19 = l_Lean_Expr_toHeadIndex(x_2); -x_20 = l_Lean_PersistentHashMap_find_x3f___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__1(x_18, x_19); -lean_dec(x_19); -if (lean_obj_tag(x_20) == 0) +lean_dec(x_3); +if (lean_obj_tag(x_21) == 0) { -lean_object* 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_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_2); -lean_dec(x_1); -x_21 = lean_box(0); -lean_ctor_set(x_14, 0, x_21); -return x_14; -} -else +uint8_t x_22; +x_22 = !lean_is_exclusive(x_21); +if (x_22 == 0) { -lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; -lean_free_object(x_14); -x_22 = lean_ctor_get(x_20, 0); -lean_inc(x_22); -lean_dec(x_20); -x_23 = l_Lean_Meta_Grind_getMaxGeneration___rarg(x_7, x_8, x_9, x_10, x_11, x_12, x_17); -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -x_25 = lean_ctor_get(x_23, 1); -lean_inc(x_25); +lean_object* x_23; lean_object* x_24; +x_23 = lean_ctor_get(x_21, 0); lean_dec(x_23); -x_26 = lean_box(0); -x_27 = lean_box(0); -lean_inc(x_22); -x_28 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___spec__1(x_1, x_2, x_22, x_24, x_26, x_22, x_22, x_27, lean_box(0), x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_25); -lean_dec(x_24); -lean_dec(x_22); -lean_dec(x_2); -if (lean_obj_tag(x_28) == 0) -{ -uint8_t x_29; -x_29 = !lean_is_exclusive(x_28); -if (x_29 == 0) -{ -lean_object* x_30; -x_30 = lean_ctor_get(x_28, 0); -lean_dec(x_30); -lean_ctor_set(x_28, 0, x_27); -return x_28; +x_24 = lean_box(0); +lean_ctor_set(x_21, 0, x_24); +return x_21; } else { -lean_object* x_31; lean_object* x_32; -x_31 = lean_ctor_get(x_28, 1); -lean_inc(x_31); -lean_dec(x_28); -x_32 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_32, 0, x_27); -lean_ctor_set(x_32, 1, x_31); -return x_32; +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_33; -x_33 = !lean_is_exclusive(x_28); -if (x_33 == 0) +uint8_t x_28; +x_28 = !lean_is_exclusive(x_21); +if (x_28 == 0) { -return x_28; +return x_21; } else { -lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_34 = lean_ctor_get(x_28, 0); -x_35 = lean_ctor_get(x_28, 1); -lean_inc(x_35); -lean_inc(x_34); -lean_dec(x_28); -x_36 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_36, 0, x_34); -lean_ctor_set(x_36, 1, x_35); -return x_36; +lean_object* 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 -{ -lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; -x_37 = lean_ctor_get(x_14, 0); -x_38 = lean_ctor_get(x_14, 1); -lean_inc(x_38); -lean_inc(x_37); -lean_dec(x_14); -x_39 = lean_ctor_get(x_37, 4); -lean_inc(x_39); -lean_dec(x_37); -lean_inc(x_2); -x_40 = l_Lean_Expr_toHeadIndex(x_2); -x_41 = l_Lean_PersistentHashMap_find_x3f___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__1(x_39, x_40); -lean_dec(x_40); -if (lean_obj_tag(x_41) == 0) +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: { -lean_object* x_42; lean_object* x_43; +lean_object* x_15; +x_15 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); @@ -2617,84 +2814,13 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_2); +lean_dec(x_4); +lean_dec(x_3); lean_dec(x_1); -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 -{ -lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; -x_44 = lean_ctor_get(x_41, 0); -lean_inc(x_44); -lean_dec(x_41); -x_45 = l_Lean_Meta_Grind_getMaxGeneration___rarg(x_7, x_8, x_9, x_10, x_11, x_12, x_38); -x_46 = lean_ctor_get(x_45, 0); -lean_inc(x_46); -x_47 = lean_ctor_get(x_45, 1); -lean_inc(x_47); -lean_dec(x_45); -x_48 = lean_box(0); -x_49 = lean_box(0); -lean_inc(x_44); -x_50 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___spec__1(x_1, x_2, x_44, x_46, x_48, x_44, x_44, x_49, lean_box(0), x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_47); -lean_dec(x_46); -lean_dec(x_44); -lean_dec(x_2); -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, 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); -} -if (lean_is_scalar(x_52)) { - x_53 = lean_alloc_ctor(0, 2, 0); -} else { - x_53 = x_52; -} -lean_ctor_set(x_53, 0, x_49); -lean_ctor_set(x_53, 1, x_51); -return x_53; -} -else -{ -lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; -x_54 = lean_ctor_get(x_50, 0); -lean_inc(x_54); -x_55 = lean_ctor_get(x_50, 1); -lean_inc(x_55); -if (lean_is_exclusive(x_50)) { - lean_ctor_release(x_50, 0); - lean_ctor_release(x_50, 1); - x_56 = x_50; -} else { - lean_dec_ref(x_50); - x_56 = lean_box(0); -} -if (lean_is_scalar(x_56)) { - x_57 = lean_alloc_ctor(1, 2, 0); -} else { - x_57 = x_56; -} -lean_ctor_set(x_57, 0, x_54); -lean_ctor_set(x_57, 1, x_55); -return x_57; -} -} -} +return x_15; } } -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___spec__1___boxed(lean_object** _args) { +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___boxed(lean_object** _args) { lean_object* x_1 = _args[0]; lean_object* x_2 = _args[1]; lean_object* x_3 = _args[2]; @@ -2713,4546 +2839,3890 @@ lean_object* x_15 = _args[14]; lean_object* x_16 = _args[15]; lean_object* x_17 = _args[16]; lean_object* x_18 = _args[17]; -lean_object* x_19 = _args[18]; -lean_object* x_20 = _args[19]; _start: { -lean_object* x_21; -x_21 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20); -lean_dec(x_11); -lean_dec(x_10); +lean_object* x_19; +x_19 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -return x_21; +return x_19; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_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_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { _start: { -lean_object* x_14; -x_14 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_object* x_15; +x_15 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); -return x_14; +lean_dec(x_2); +return x_15; } } -LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___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_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processOffset___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: { -uint8_t x_13; -x_13 = l_Lean_Expr_hasMVar(x_1); -if (x_13 == 0) -{ -lean_object* x_14; -x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_1); -lean_ctor_set(x_14, 1, x_12); -return x_14; -} -else +lean_object* x_18; lean_object* x_34; +lean_inc(x_6); +x_34 = l_Lean_Meta_Grind_getENode(x_6, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +if (lean_obj_tag(x_34) == 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; uint8_t x_25; -x_15 = lean_st_ref_get(x_9, x_12); -x_16 = lean_ctor_get(x_15, 0); +lean_object* x_35; lean_object* x_36; lean_object* x_37; uint8_t x_38; +x_35 = lean_ctor_get(x_34, 0); +lean_inc(x_35); +x_36 = lean_ctor_get(x_34, 1); +lean_inc(x_36); +lean_dec(x_34); +x_37 = lean_ctor_get(x_35, 8); +lean_inc(x_37); +lean_dec(x_35); +x_38 = lean_nat_dec_lt(x_37, x_5); +if (x_38 == 0) +{ +lean_object* x_39; lean_object* x_40; +lean_dec(x_37); +x_39 = lean_box(0); +x_40 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_4, x_6, x_39, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_36); +x_18 = x_40; +goto block_33; +} +else +{ +lean_object* x_41; 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_instantiateMVarsCore(x_18, x_1); -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_st_ref_take(x_9, x_17); -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_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_6); +x_41 = l_Lean_Meta_isOffset_x3f(x_6, x_13, x_14, x_15, x_16, x_36); +if (lean_obj_tag(x_41) == 0) { -lean_object* x_26; lean_object* x_27; uint8_t x_28; -x_26 = lean_ctor_get(x_23, 0); -lean_dec(x_26); -lean_ctor_set(x_23, 0, x_21); -x_27 = lean_st_ref_set(x_9, x_23, x_24); -x_28 = !lean_is_exclusive(x_27); -if (x_28 == 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_29; -x_29 = lean_ctor_get(x_27, 0); -lean_dec(x_29); -lean_ctor_set(x_27, 0, x_20); -return x_27; +lean_object* x_43; lean_object* x_44; +x_43 = lean_ctor_get(x_41, 1); +lean_inc(x_43); +lean_dec(x_41); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_6); +x_44 = l_Lean_Meta_evalNat(x_6, x_13, x_14, x_15, x_16, x_43); +if (lean_obj_tag(x_44) == 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_46; lean_object* x_47; lean_object* x_48; +lean_dec(x_37); +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_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_4, x_6, x_47, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_46); +x_18 = x_48; +goto block_33; } 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_20); -lean_ctor_set(x_31, 1, x_30); -return x_31; +lean_object* x_49; lean_object* x_50; uint8_t x_51; +x_49 = lean_ctor_get(x_44, 1); +lean_inc(x_49); +lean_dec(x_44); +x_50 = lean_ctor_get(x_45, 0); +lean_inc(x_50); +lean_dec(x_45); +x_51 = lean_nat_dec_le(x_3, x_50); +if (x_51 == 0) +{ +lean_object* x_52; lean_object* x_53; +lean_dec(x_50); +lean_dec(x_37); +x_52 = lean_box(0); +x_53 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_4, x_6, x_52, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_49); +x_18 = x_53; +goto block_33; } +else +{ +lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_54 = lean_nat_sub(x_50, x_3); +lean_dec(x_50); +x_55 = l_Lean_mkNatLit(x_54); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +x_56 = l_Lean_Meta_Grind_canon(x_55, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_49); +if (lean_obj_tag(x_56) == 0) +{ +lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; +x_57 = lean_ctor_get(x_56, 0); +lean_inc(x_57); +x_58 = lean_ctor_get(x_56, 1); +lean_inc(x_58); +lean_dec(x_56); +x_59 = l_Lean_Meta_Grind_shareCommon(x_57, x_10, x_11, x_12, x_13, x_14, x_15, x_16, 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_unsigned_to_nat(1u); +x_63 = lean_nat_add(x_37, x_62); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_60); +x_64 = l_Lean_Meta_Grind_internalize(x_60, x_63, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_61); +if (lean_obj_tag(x_64) == 0) +{ +lean_object* x_65; lean_object* x_66; +x_65 = lean_ctor_get(x_64, 1); +lean_inc(x_65); +lean_dec(x_64); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_2); +lean_inc(x_1); +x_66 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f(x_1, x_2, x_60, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_65); +if (lean_obj_tag(x_66) == 0) +{ +lean_object* x_67; +x_67 = lean_ctor_get(x_66, 0); +lean_inc(x_67); +if (lean_obj_tag(x_67) == 0) +{ +lean_object* x_68; lean_object* x_69; lean_object* x_70; +lean_dec(x_37); +x_68 = lean_ctor_get(x_66, 1); +lean_inc(x_68); +lean_dec(x_66); +x_69 = lean_box(0); +x_70 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_4, x_6, x_69, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_68); +x_18 = x_70; +goto block_33; } else { -lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_32 = lean_ctor_get(x_23, 1); -x_33 = lean_ctor_get(x_23, 2); -x_34 = lean_ctor_get(x_23, 3); -x_35 = lean_ctor_get(x_23, 4); -lean_inc(x_35); -lean_inc(x_34); -lean_inc(x_33); -lean_inc(x_32); -lean_dec(x_23); -x_36 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_36, 0, x_21); -lean_ctor_set(x_36, 1, x_32); -lean_ctor_set(x_36, 2, x_33); -lean_ctor_set(x_36, 3, x_34); -lean_ctor_set(x_36, 4, x_35); -x_37 = lean_st_ref_set(x_9, x_36, x_24); -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); +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_71 = lean_ctor_get(x_66, 1); +lean_inc(x_71); +lean_dec(x_66); +x_72 = lean_ctor_get(x_67, 0); +lean_inc(x_72); +lean_dec(x_67); +x_73 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_Choice_updateGen(x_72, x_37); +x_74 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_pushChoice(x_73, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_71); +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_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_4, x_6, x_75, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_76); +lean_dec(x_75); +x_18 = x_77; +goto block_33; } -if (lean_is_scalar(x_39)) { - x_40 = lean_alloc_ctor(0, 2, 0); -} else { - x_40 = x_39; } -lean_ctor_set(x_40, 0, x_20); -lean_ctor_set(x_40, 1, x_38); -return x_40; +else +{ +uint8_t x_78; +lean_dec(x_37); +lean_dec(x_6); +x_78 = !lean_is_exclusive(x_66); +if (x_78 == 0) +{ +x_18 = x_66; +goto block_33; } +else +{ +lean_object* x_79; lean_object* x_80; lean_object* x_81; +x_79 = lean_ctor_get(x_66, 0); +x_80 = lean_ctor_get(x_66, 1); +lean_inc(x_80); +lean_inc(x_79); +lean_dec(x_66); +x_81 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_81, 0, x_79); +lean_ctor_set(x_81, 1, x_80); +x_18 = x_81; +goto block_33; } } } -LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: +else { -lean_object* x_13; lean_object* x_14; uint8_t x_15; -x_13 = lean_ctor_get(x_10, 2); -x_14 = l_Lean_isTracingEnabledForCore(x_1, x_13, x_12); -x_15 = !lean_is_exclusive(x_14); -if (x_15 == 0) +uint8_t x_82; +lean_dec(x_60); +lean_dec(x_37); +lean_dec(x_6); +x_82 = !lean_is_exclusive(x_64); +if (x_82 == 0) { -return x_14; +x_18 = x_64; +goto block_33; } else { -lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_16 = lean_ctor_get(x_14, 0); -x_17 = lean_ctor_get(x_14, 1); -lean_inc(x_17); -lean_inc(x_16); -lean_dec(x_14); -x_18 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_18, 0, x_16); -lean_ctor_set(x_18, 1, x_17); -return x_18; +lean_object* x_83; lean_object* x_84; lean_object* x_85; +x_83 = lean_ctor_get(x_64, 0); +x_84 = lean_ctor_get(x_64, 1); +lean_inc(x_84); +lean_inc(x_83); +lean_dec(x_64); +x_85 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_85, 0, x_83); +lean_ctor_set(x_85, 1, x_84); +x_18 = x_85; +goto block_33; } } } -LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___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) { -_start: +else { -lean_object* x_13; lean_object* x_14; uint8_t x_15; -x_13 = lean_ctor_get(x_10, 5); -x_14 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_1, x_8, x_9, x_10, x_11, x_12); -x_15 = !lean_is_exclusive(x_14); -if (x_15 == 0) +uint8_t x_86; +lean_dec(x_37); +lean_dec(x_6); +x_86 = !lean_is_exclusive(x_56); +if (x_86 == 0) { -lean_object* x_16; lean_object* x_17; -x_16 = lean_ctor_get(x_14, 0); -lean_inc(x_13); -x_17 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_17, 0, x_13); -lean_ctor_set(x_17, 1, x_16); -lean_ctor_set_tag(x_14, 1); -lean_ctor_set(x_14, 0, x_17); -return x_14; +x_18 = x_56; +goto block_33; } else { -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_18 = lean_ctor_get(x_14, 0); -x_19 = lean_ctor_get(x_14, 1); -lean_inc(x_19); -lean_inc(x_18); -lean_dec(x_14); -lean_inc(x_13); -x_20 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_20, 0, x_13); -lean_ctor_set(x_20, 1, x_18); -x_21 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_21, 0, x_20); -lean_ctor_set(x_21, 1, x_19); -return x_21; +lean_object* x_87; lean_object* x_88; lean_object* x_89; +x_87 = lean_ctor_get(x_56, 0); +x_88 = lean_ctor_get(x_56, 1); +lean_inc(x_88); +lean_inc(x_87); +lean_dec(x_56); +x_89 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_89, 0, x_87); +lean_ctor_set(x_89, 1, x_88); +x_18 = x_89; +goto block_33; } } } -static lean_object* _init_l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("unknown constant '", 18, 18); -return x_1; } } -static lean_object* _init_l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5___closed__2() { -_start: +else { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5___closed__1; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; -} +uint8_t x_90; +lean_dec(x_37); +lean_dec(x_6); +x_90 = !lean_is_exclusive(x_44); +if (x_90 == 0) +{ +x_18 = x_44; +goto block_33; } -static lean_object* _init_l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5___closed__3() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("'", 1, 1); -return x_1; +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_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); +x_18 = x_93; +goto block_33; } } -static lean_object* _init_l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5___closed__4() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5___closed__3; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; } +else +{ +lean_object* x_94; lean_object* x_95; uint8_t x_96; +x_94 = lean_ctor_get(x_42, 0); +lean_inc(x_94); +lean_dec(x_42); +x_95 = lean_ctor_get(x_41, 1); +lean_inc(x_95); +lean_dec(x_41); +x_96 = !lean_is_exclusive(x_94); +if (x_96 == 0) +{ +lean_object* x_97; lean_object* x_98; uint8_t x_99; +x_97 = lean_ctor_get(x_94, 0); +x_98 = lean_ctor_get(x_94, 1); +x_99 = lean_nat_dec_lt(x_98, x_3); +if (x_99 == 0) +{ +uint8_t x_100; +lean_free_object(x_94); +x_100 = lean_nat_dec_eq(x_98, x_3); +if (x_100 == 0) +{ +uint8_t x_101; +x_101 = lean_nat_dec_lt(x_3, x_98); +if (x_101 == 0) +{ +lean_object* x_102; lean_object* x_103; +lean_dec(x_98); +lean_dec(x_97); +lean_dec(x_37); +x_102 = lean_box(0); +x_103 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_4, x_6, x_102, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_95); +x_18 = x_103; +goto block_33; } -LEAN_EXPORT lean_object* l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: +else { -lean_object* x_13; uint8_t x_14; -x_13 = lean_st_ref_get(x_11, x_12); -x_14 = !lean_is_exclusive(x_13); -if (x_14 == 0) +lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; +x_104 = lean_nat_sub(x_98, x_3); +lean_dec(x_98); +x_105 = l_Lean_mkNatLit(x_104); +x_106 = l___private_Lean_Expr_0__Lean_natAddFn; +x_107 = l_Lean_mkAppB(x_106, x_97, x_105); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +x_108 = l_Lean_Meta_Grind_canon(x_107, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_95); +if (lean_obj_tag(x_108) == 0) { -lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_15 = lean_ctor_get(x_13, 0); -x_16 = lean_ctor_get(x_13, 1); -x_17 = lean_ctor_get(x_15, 0); -lean_inc(x_17); -lean_dec(x_15); +lean_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 = l_Lean_Meta_Grind_shareCommon(x_109, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_110); +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_unsigned_to_nat(1u); +x_115 = lean_nat_add(x_37, x_114); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_112); +x_116 = l_Lean_Meta_Grind_internalize(x_112, x_115, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_113); +if (lean_obj_tag(x_116) == 0) +{ +lean_object* x_117; lean_object* x_118; +x_117 = lean_ctor_get(x_116, 1); +lean_inc(x_117); +lean_dec(x_116); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_2); lean_inc(x_1); -x_18 = lean_environment_find(x_17, x_1); -if (lean_obj_tag(x_18) == 0) +x_118 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f(x_1, x_2, x_112, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_117); +if (lean_obj_tag(x_118) == 0) { -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_free_object(x_13); -x_19 = 0; -x_20 = l_Lean_MessageData_ofConstName(x_1, x_19); -x_21 = l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5___closed__2; -x_22 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_22, 0, x_21); -lean_ctor_set(x_22, 1, x_20); -x_23 = l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5___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_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__6(x_24, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_16); -return x_25; +lean_object* x_119; +x_119 = lean_ctor_get(x_118, 0); +lean_inc(x_119); +if (lean_obj_tag(x_119) == 0) +{ +lean_object* x_120; lean_object* x_121; lean_object* x_122; +lean_dec(x_37); +x_120 = lean_ctor_get(x_118, 1); +lean_inc(x_120); +lean_dec(x_118); +x_121 = lean_box(0); +x_122 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_4, x_6, x_121, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_120); +x_18 = x_122; +goto block_33; } else { -lean_object* x_26; -lean_dec(x_1); -x_26 = lean_ctor_get(x_18, 0); -lean_inc(x_26); -lean_dec(x_18); -lean_ctor_set(x_13, 0, x_26); -return x_13; +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_123 = lean_ctor_get(x_118, 1); +lean_inc(x_123); +lean_dec(x_118); +x_124 = lean_ctor_get(x_119, 0); +lean_inc(x_124); +lean_dec(x_119); +x_125 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_Choice_updateGen(x_124, x_37); +x_126 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_pushChoice(x_125, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_123); +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_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_4, x_6, x_127, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_128); +lean_dec(x_127); +x_18 = x_129; +goto block_33; } } else { -lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; -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 = lean_ctor_get(x_27, 0); -lean_inc(x_29); -lean_dec(x_27); -lean_inc(x_1); -x_30 = lean_environment_find(x_29, x_1); -if (lean_obj_tag(x_30) == 0) +uint8_t x_130; +lean_dec(x_37); +lean_dec(x_6); +x_130 = !lean_is_exclusive(x_118); +if (x_130 == 0) { -uint8_t x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; -x_31 = 0; -x_32 = l_Lean_MessageData_ofConstName(x_1, x_31); -x_33 = l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5___closed__2; -x_34 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_34, 0, x_33); -lean_ctor_set(x_34, 1, x_32); -x_35 = l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5___closed__4; -x_36 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_36, 0, x_34); -lean_ctor_set(x_36, 1, x_35); -x_37 = l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__6(x_36, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_28); -return x_37; +x_18 = x_118; +goto block_33; } else { -lean_object* x_38; lean_object* x_39; -lean_dec(x_1); -x_38 = lean_ctor_get(x_30, 0); -lean_inc(x_38); -lean_dec(x_30); -x_39 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_39, 0, x_38); -lean_ctor_set(x_39, 1, x_28); -return x_39; -} +lean_object* x_131; lean_object* x_132; lean_object* x_133; +x_131 = lean_ctor_get(x_118, 0); +x_132 = lean_ctor_get(x_118, 1); +lean_inc(x_132); +lean_inc(x_131); +lean_dec(x_118); +x_133 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_133, 0, x_131); +lean_ctor_set(x_133, 1, x_132); +x_18 = x_133; +goto block_33; } } } -LEAN_EXPORT lean_object* l_Lean_mkConstWithLevelParams___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -lean_object* x_13; -lean_inc(x_1); -x_13 = l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___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); -if (lean_obj_tag(x_13) == 0) +else { -uint8_t x_14; -x_14 = !lean_is_exclusive(x_13); -if (x_14 == 0) +uint8_t x_134; +lean_dec(x_112); +lean_dec(x_37); +lean_dec(x_6); +x_134 = !lean_is_exclusive(x_116); +if (x_134 == 0) { -lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_15 = lean_ctor_get(x_13, 0); -x_16 = l_Lean_ConstantInfo_levelParams(x_15); -lean_dec(x_15); -x_17 = lean_box(0); -x_18 = l_List_mapTR_loop___at_Lean_mkConstWithLevelParams___spec__1(x_16, x_17); -x_19 = l_Lean_Expr_const___override(x_1, x_18); -lean_ctor_set(x_13, 0, x_19); -return x_13; +x_18 = x_116; +goto block_33; } else { -lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_20 = lean_ctor_get(x_13, 0); -x_21 = lean_ctor_get(x_13, 1); -lean_inc(x_21); -lean_inc(x_20); -lean_dec(x_13); -x_22 = l_Lean_ConstantInfo_levelParams(x_20); -lean_dec(x_20); -x_23 = lean_box(0); -x_24 = l_List_mapTR_loop___at_Lean_mkConstWithLevelParams___spec__1(x_22, x_23); -x_25 = l_Lean_Expr_const___override(x_1, x_24); -x_26 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_26, 0, x_25); -lean_ctor_set(x_26, 1, x_21); -return x_26; +lean_object* x_135; lean_object* x_136; lean_object* x_137; +x_135 = lean_ctor_get(x_116, 0); +x_136 = lean_ctor_get(x_116, 1); +lean_inc(x_136); +lean_inc(x_135); +lean_dec(x_116); +x_137 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_137, 0, x_135); +lean_ctor_set(x_137, 1, x_136); +x_18 = x_137; +goto block_33; +} } } else { -uint8_t x_27; -lean_dec(x_1); -x_27 = !lean_is_exclusive(x_13); -if (x_27 == 0) +uint8_t x_138; +lean_dec(x_37); +lean_dec(x_6); +x_138 = !lean_is_exclusive(x_108); +if (x_138 == 0) { -return x_13; +x_18 = x_108; +goto block_33; } else { -lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_28 = lean_ctor_get(x_13, 0); -x_29 = lean_ctor_get(x_13, 1); -lean_inc(x_29); -lean_inc(x_28); -lean_dec(x_13); -x_30 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_30, 0, x_28); -lean_ctor_set(x_30, 1, x_29); -return x_30; -} -} +lean_object* x_139; lean_object* x_140; lean_object* x_141; +x_139 = lean_ctor_get(x_108, 0); +x_140 = lean_ctor_get(x_108, 1); +lean_inc(x_140); +lean_inc(x_139); +lean_dec(x_108); +x_141 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_141, 0, x_139); +lean_ctor_set(x_141, 1, x_140); +x_18 = x_141; +goto block_33; } } -static lean_object* _init_l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("[unknown]", 9, 9); -return x_1; } } -static lean_object* _init_l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3___closed__2() { -_start: +else { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3___closed__1; -x_2 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; -} +lean_object* x_142; +lean_dec(x_98); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_2); +lean_inc(x_1); +x_142 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f(x_1, x_2, x_97, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_95); +if (lean_obj_tag(x_142) == 0) +{ +lean_object* x_143; +x_143 = lean_ctor_get(x_142, 0); +lean_inc(x_143); +if (lean_obj_tag(x_143) == 0) +{ +lean_object* x_144; lean_object* x_145; lean_object* x_146; +lean_dec(x_37); +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_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_4, x_6, x_145, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_144); +x_18 = x_146; +goto block_33; } -static lean_object* _init_l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3___closed__3() { -_start: +else { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3___closed__2; -x_2 = l_Lean_MessageData_ofFormat(x_1); -return x_2; +lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; +x_147 = lean_ctor_get(x_142, 1); +lean_inc(x_147); +lean_dec(x_142); +x_148 = lean_ctor_get(x_143, 0); +lean_inc(x_148); +lean_dec(x_143); +x_149 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_Choice_updateGen(x_148, x_37); +x_150 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_pushChoice(x_149, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_147); +x_151 = lean_ctor_get(x_150, 0); +lean_inc(x_151); +x_152 = lean_ctor_get(x_150, 1); +lean_inc(x_152); +lean_dec(x_150); +x_153 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_4, x_6, x_151, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_152); +lean_dec(x_151); +x_18 = x_153; +goto block_33; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -switch (lean_obj_tag(x_1)) { -case 0: -{ -lean_object* x_13; lean_object* x_14; -x_13 = lean_ctor_get(x_1, 0); -lean_inc(x_13); -lean_dec(x_1); -x_14 = l_Lean_mkConstWithLevelParams___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__4(x_13, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_14) == 0) +else { -uint8_t x_15; -x_15 = !lean_is_exclusive(x_14); -if (x_15 == 0) +uint8_t x_154; +lean_dec(x_37); +lean_dec(x_6); +x_154 = !lean_is_exclusive(x_142); +if (x_154 == 0) { -lean_object* x_16; lean_object* x_17; -x_16 = lean_ctor_get(x_14, 0); -x_17 = l_Lean_MessageData_ofConst(x_16); -lean_ctor_set(x_14, 0, x_17); -return x_14; +x_18 = x_142; +goto block_33; } else { -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_18 = lean_ctor_get(x_14, 0); -x_19 = lean_ctor_get(x_14, 1); -lean_inc(x_19); -lean_inc(x_18); -lean_dec(x_14); -x_20 = l_Lean_MessageData_ofConst(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_object* x_155; lean_object* x_156; lean_object* x_157; +x_155 = lean_ctor_get(x_142, 0); +x_156 = lean_ctor_get(x_142, 1); +lean_inc(x_156); +lean_inc(x_155); +lean_dec(x_142); +x_157 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_157, 0, x_155); +lean_ctor_set(x_157, 1, x_156); +x_18 = x_157; +goto block_33; +} +} } } else { -uint8_t x_22; -x_22 = !lean_is_exclusive(x_14); -if (x_22 == 0) -{ -return x_14; +lean_object* x_158; lean_object* x_159; lean_object* x_160; uint8_t x_161; +lean_inc(x_1); +x_158 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_Choice_updateGen(x_1, x_37); +x_159 = lean_nat_sub(x_3, x_98); +lean_dec(x_98); +lean_inc(x_2); +x_160 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_160, 0, x_2); +lean_ctor_set(x_160, 1, x_159); +lean_ctor_set(x_160, 2, x_97); +x_161 = !lean_is_exclusive(x_158); +if (x_161 == 0) +{ +lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; +x_162 = lean_ctor_get(x_158, 0); +lean_ctor_set_tag(x_94, 1); +lean_ctor_set(x_94, 1, x_162); +lean_ctor_set(x_94, 0, x_160); +lean_ctor_set(x_158, 0, x_94); +x_163 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_pushChoice(x_158, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_95); +x_164 = lean_ctor_get(x_163, 0); +lean_inc(x_164); +x_165 = lean_ctor_get(x_163, 1); +lean_inc(x_165); +lean_dec(x_163); +x_166 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_4, x_6, x_164, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_165); +lean_dec(x_164); +x_18 = x_166; +goto block_33; } else { -lean_object* x_23; lean_object* x_24; lean_object* x_25; -x_23 = lean_ctor_get(x_14, 0); -x_24 = lean_ctor_get(x_14, 1); -lean_inc(x_24); -lean_inc(x_23); -lean_dec(x_14); -x_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_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; +x_167 = lean_ctor_get(x_158, 0); +x_168 = lean_ctor_get(x_158, 1); +x_169 = lean_ctor_get(x_158, 2); +lean_inc(x_169); +lean_inc(x_168); +lean_inc(x_167); +lean_dec(x_158); +lean_ctor_set_tag(x_94, 1); +lean_ctor_set(x_94, 1, x_167); +lean_ctor_set(x_94, 0, x_160); +x_170 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_170, 0, x_94); +lean_ctor_set(x_170, 1, x_168); +lean_ctor_set(x_170, 2, x_169); +x_171 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_pushChoice(x_170, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_95); +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_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_4, x_6, x_172, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_173); +lean_dec(x_172); +x_18 = x_174; +goto block_33; } } } -case 1: +else { -lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_26 = lean_ctor_get(x_1, 0); -lean_inc(x_26); -lean_dec(x_1); -x_27 = l_Lean_Expr_fvar___override(x_26); -x_28 = l_Lean_MessageData_ofExpr(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_12); -return x_29; +lean_object* x_175; lean_object* x_176; uint8_t x_177; +x_175 = lean_ctor_get(x_94, 0); +x_176 = lean_ctor_get(x_94, 1); +lean_inc(x_176); +lean_inc(x_175); +lean_dec(x_94); +x_177 = lean_nat_dec_lt(x_176, x_3); +if (x_177 == 0) +{ +uint8_t x_178; +x_178 = lean_nat_dec_eq(x_176, x_3); +if (x_178 == 0) +{ +uint8_t x_179; +x_179 = lean_nat_dec_lt(x_3, x_176); +if (x_179 == 0) +{ +lean_object* x_180; lean_object* x_181; +lean_dec(x_176); +lean_dec(x_175); +lean_dec(x_37); +x_180 = lean_box(0); +x_181 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_4, x_6, x_180, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_95); +x_18 = x_181; +goto block_33; } -case 2: +else { -uint8_t x_30; -x_30 = !lean_is_exclusive(x_1); -if (x_30 == 0) +lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; +x_182 = lean_nat_sub(x_176, x_3); +lean_dec(x_176); +x_183 = l_Lean_mkNatLit(x_182); +x_184 = l___private_Lean_Expr_0__Lean_natAddFn; +x_185 = l_Lean_mkAppB(x_184, x_175, x_183); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +x_186 = l_Lean_Meta_Grind_canon(x_185, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_95); +if (lean_obj_tag(x_186) == 0) { -lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_31 = lean_ctor_get(x_1, 1); -x_32 = lean_ctor_get(x_1, 0); -lean_dec(x_32); -x_33 = l_Lean_MessageData_ofSyntax(x_31); -lean_ctor_set_tag(x_1, 0); -lean_ctor_set(x_1, 1, x_12); -lean_ctor_set(x_1, 0, x_33); -return x_1; +lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; +x_187 = lean_ctor_get(x_186, 0); +lean_inc(x_187); +x_188 = lean_ctor_get(x_186, 1); +lean_inc(x_188); +lean_dec(x_186); +x_189 = l_Lean_Meta_Grind_shareCommon(x_187, x_10, x_11, x_12, x_13, x_14, x_15, x_16, 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 = lean_unsigned_to_nat(1u); +x_193 = lean_nat_add(x_37, x_192); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_190); +x_194 = l_Lean_Meta_Grind_internalize(x_190, x_193, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_191); +if (lean_obj_tag(x_194) == 0) +{ +lean_object* x_195; lean_object* x_196; +x_195 = lean_ctor_get(x_194, 1); +lean_inc(x_195); +lean_dec(x_194); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_2); +lean_inc(x_1); +x_196 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f(x_1, x_2, x_190, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_195); +if (lean_obj_tag(x_196) == 0) +{ +lean_object* x_197; +x_197 = lean_ctor_get(x_196, 0); +lean_inc(x_197); +if (lean_obj_tag(x_197) == 0) +{ +lean_object* x_198; lean_object* x_199; lean_object* x_200; +lean_dec(x_37); +x_198 = lean_ctor_get(x_196, 1); +lean_inc(x_198); +lean_dec(x_196); +x_199 = lean_box(0); +x_200 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_4, x_6, x_199, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_198); +x_18 = x_200; +goto block_33; } else { -lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_34 = lean_ctor_get(x_1, 1); -lean_inc(x_34); -lean_dec(x_1); -x_35 = l_Lean_MessageData_ofSyntax(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_12); -return 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; +x_201 = lean_ctor_get(x_196, 1); +lean_inc(x_201); +lean_dec(x_196); +x_202 = lean_ctor_get(x_197, 0); +lean_inc(x_202); +lean_dec(x_197); +x_203 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_Choice_updateGen(x_202, x_37); +x_204 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_pushChoice(x_203, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_201); +x_205 = lean_ctor_get(x_204, 0); +lean_inc(x_205); +x_206 = lean_ctor_get(x_204, 1); +lean_inc(x_206); +lean_dec(x_204); +x_207 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_4, x_6, x_205, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_206); +lean_dec(x_205); +x_18 = x_207; +goto block_33; } } -default: +else { -lean_object* x_37; lean_object* x_38; -x_37 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3___closed__3; -x_38 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_38, 0, x_37); -lean_ctor_set(x_38, 1, x_12); -return x_38; +lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; +lean_dec(x_37); +lean_dec(x_6); +x_208 = lean_ctor_get(x_196, 0); +lean_inc(x_208); +x_209 = lean_ctor_get(x_196, 1); +lean_inc(x_209); +if (lean_is_exclusive(x_196)) { + lean_ctor_release(x_196, 0); + lean_ctor_release(x_196, 1); + x_210 = x_196; +} else { + lean_dec_ref(x_196); + x_210 = lean_box(0); } +if (lean_is_scalar(x_210)) { + x_211 = lean_alloc_ctor(1, 2, 0); +} else { + x_211 = x_210; } +lean_ctor_set(x_211, 0, x_208); +lean_ctor_set(x_211, 1, x_209); +x_18 = x_211; +goto block_33; } } -static double _init_l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__1() { -_start: +else { -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_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; +lean_dec(x_190); +lean_dec(x_37); +lean_dec(x_6); +x_212 = lean_ctor_get(x_194, 0); +lean_inc(x_212); +x_213 = lean_ctor_get(x_194, 1); +lean_inc(x_213); +if (lean_is_exclusive(x_194)) { + lean_ctor_release(x_194, 0); + lean_ctor_release(x_194, 1); + x_214 = x_194; +} else { + lean_dec_ref(x_194); + x_214 = lean_box(0); } +if (lean_is_scalar(x_214)) { + x_215 = lean_alloc_ctor(1, 2, 0); +} else { + x_215 = x_214; } -static lean_object* _init_l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2() { -_start: +lean_ctor_set(x_215, 0, x_212); +lean_ctor_set(x_215, 1, x_213); +x_18 = x_215; +goto block_33; +} +} +else { -lean_object* x_1; lean_object* x_2; -x_1 = lean_box(0); -x_2 = lean_array_mk(x_1); -return x_2; +lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; +lean_dec(x_37); +lean_dec(x_6); +x_216 = lean_ctor_get(x_186, 0); +lean_inc(x_216); +x_217 = lean_ctor_get(x_186, 1); +lean_inc(x_217); +if (lean_is_exclusive(x_186)) { + lean_ctor_release(x_186, 0); + lean_ctor_release(x_186, 1); + x_218 = x_186; +} else { + lean_dec_ref(x_186); + x_218 = lean_box(0); } +if (lean_is_scalar(x_218)) { + x_219 = lean_alloc_ctor(1, 2, 0); +} else { + x_219 = x_218; } -LEAN_EXPORT lean_object* l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { -_start: +lean_ctor_set(x_219, 0, x_216); +lean_ctor_set(x_219, 1, x_217); +x_18 = x_219; +goto block_33; +} +} +} +else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; -x_14 = lean_ctor_get(x_11, 5); -x_15 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_2, x_9, x_10, x_11, x_12, x_13); -x_16 = lean_ctor_get(x_15, 0); +lean_object* x_220; +lean_dec(x_176); lean_inc(x_16); -x_17 = lean_ctor_get(x_15, 1); -lean_inc(x_17); -lean_dec(x_15); -x_18 = lean_st_ref_take(x_12, x_17); -x_19 = lean_ctor_get(x_18, 0); -lean_inc(x_19); -x_20 = lean_ctor_get(x_19, 3); -lean_inc(x_20); -x_21 = !lean_is_exclusive(x_18); -if (x_21 == 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); +lean_inc(x_9); +lean_inc(x_2); +lean_inc(x_1); +x_220 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f(x_1, x_2, x_175, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_95); +if (lean_obj_tag(x_220) == 0) { -lean_object* x_22; lean_object* x_23; uint8_t x_24; -x_22 = lean_ctor_get(x_18, 1); -x_23 = lean_ctor_get(x_18, 0); -lean_dec(x_23); -x_24 = !lean_is_exclusive(x_19); -if (x_24 == 0) -{ -lean_object* x_25; uint8_t x_26; -x_25 = lean_ctor_get(x_19, 3); -lean_dec(x_25); -x_26 = !lean_is_exclusive(x_20); -if (x_26 == 0) -{ -lean_object* x_27; double 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; uint8_t x_36; -x_27 = lean_ctor_get(x_20, 0); -x_28 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__1; -x_29 = 0; -x_30 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__1; -x_31 = lean_alloc_ctor(0, 2, 17); -lean_ctor_set(x_31, 0, x_1); -lean_ctor_set(x_31, 1, x_30); -lean_ctor_set_float(x_31, sizeof(void*)*2, x_28); -lean_ctor_set_float(x_31, sizeof(void*)*2 + 8, x_28); -lean_ctor_set_uint8(x_31, sizeof(void*)*2 + 16, x_29); -x_32 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2; -x_33 = lean_alloc_ctor(9, 3, 0); -lean_ctor_set(x_33, 0, x_31); -lean_ctor_set(x_33, 1, x_16); -lean_ctor_set(x_33, 2, x_32); -lean_inc(x_14); -lean_ctor_set(x_18, 1, x_33); -lean_ctor_set(x_18, 0, x_14); -x_34 = l_Lean_PersistentArray_push___rarg(x_27, x_18); -lean_ctor_set(x_20, 0, x_34); -x_35 = lean_st_ref_set(x_12, x_19, x_22); -x_36 = !lean_is_exclusive(x_35); -if (x_36 == 0) +lean_object* x_221; +x_221 = lean_ctor_get(x_220, 0); +lean_inc(x_221); +if (lean_obj_tag(x_221) == 0) { -lean_object* x_37; lean_object* x_38; -x_37 = lean_ctor_get(x_35, 0); +lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_dec(x_37); -x_38 = lean_box(0); -lean_ctor_set(x_35, 0, x_38); -return x_35; +x_222 = lean_ctor_get(x_220, 1); +lean_inc(x_222); +lean_dec(x_220); +x_223 = lean_box(0); +x_224 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_4, x_6, x_223, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_222); +x_18 = x_224; +goto block_33; } else { -lean_object* x_39; lean_object* x_40; lean_object* x_41; -x_39 = lean_ctor_get(x_35, 1); -lean_inc(x_39); -lean_dec(x_35); -x_40 = lean_box(0); -x_41 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_41, 0, x_40); -lean_ctor_set(x_41, 1, x_39); -return x_41; +lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; +x_225 = lean_ctor_get(x_220, 1); +lean_inc(x_225); +lean_dec(x_220); +x_226 = lean_ctor_get(x_221, 0); +lean_inc(x_226); +lean_dec(x_221); +x_227 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_Choice_updateGen(x_226, x_37); +x_228 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_pushChoice(x_227, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_225); +x_229 = lean_ctor_get(x_228, 0); +lean_inc(x_229); +x_230 = lean_ctor_get(x_228, 1); +lean_inc(x_230); +lean_dec(x_228); +x_231 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_4, x_6, x_229, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_230); +lean_dec(x_229); +x_18 = x_231; +goto block_33; } } else { -uint64_t x_42; lean_object* x_43; double 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; -x_42 = lean_ctor_get_uint64(x_20, sizeof(void*)*1); -x_43 = lean_ctor_get(x_20, 0); -lean_inc(x_43); -lean_dec(x_20); -x_44 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__1; -x_45 = 0; -x_46 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__1; -x_47 = lean_alloc_ctor(0, 2, 17); -lean_ctor_set(x_47, 0, x_1); -lean_ctor_set(x_47, 1, x_46); -lean_ctor_set_float(x_47, sizeof(void*)*2, x_44); -lean_ctor_set_float(x_47, sizeof(void*)*2 + 8, x_44); -lean_ctor_set_uint8(x_47, sizeof(void*)*2 + 16, x_45); -x_48 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2; -x_49 = lean_alloc_ctor(9, 3, 0); -lean_ctor_set(x_49, 0, x_47); -lean_ctor_set(x_49, 1, x_16); -lean_ctor_set(x_49, 2, x_48); -lean_inc(x_14); -lean_ctor_set(x_18, 1, x_49); -lean_ctor_set(x_18, 0, x_14); -x_50 = l_Lean_PersistentArray_push___rarg(x_43, x_18); -x_51 = lean_alloc_ctor(0, 1, 8); -lean_ctor_set(x_51, 0, x_50); -lean_ctor_set_uint64(x_51, sizeof(void*)*1, x_42); -lean_ctor_set(x_19, 3, x_51); -x_52 = lean_st_ref_set(x_12, x_19, x_22); -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_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; +lean_dec(x_37); +lean_dec(x_6); +x_232 = lean_ctor_get(x_220, 0); +lean_inc(x_232); +x_233 = lean_ctor_get(x_220, 1); +lean_inc(x_233); +if (lean_is_exclusive(x_220)) { + lean_ctor_release(x_220, 0); + lean_ctor_release(x_220, 1); + x_234 = x_220; } else { - lean_dec_ref(x_52); - x_54 = lean_box(0); + lean_dec_ref(x_220); + x_234 = lean_box(0); } -x_55 = lean_box(0); -if (lean_is_scalar(x_54)) { - x_56 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_234)) { + x_235 = lean_alloc_ctor(1, 2, 0); } else { - x_56 = x_54; + x_235 = x_234; +} +lean_ctor_set(x_235, 0, x_232); +lean_ctor_set(x_235, 1, x_233); +x_18 = x_235; +goto block_33; } -lean_ctor_set(x_56, 0, x_55); -lean_ctor_set(x_56, 1, x_53); -return x_56; } } else { -lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; uint64_t x_64; lean_object* x_65; lean_object* x_66; double x_67; uint8_t x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; -x_57 = lean_ctor_get(x_19, 0); -x_58 = lean_ctor_get(x_19, 1); -x_59 = lean_ctor_get(x_19, 2); -x_60 = lean_ctor_get(x_19, 4); -x_61 = lean_ctor_get(x_19, 5); -x_62 = lean_ctor_get(x_19, 6); -x_63 = lean_ctor_get(x_19, 7); -lean_inc(x_63); -lean_inc(x_62); -lean_inc(x_61); -lean_inc(x_60); -lean_inc(x_59); -lean_inc(x_58); -lean_inc(x_57); -lean_dec(x_19); -x_64 = lean_ctor_get_uint64(x_20, sizeof(void*)*1); -x_65 = lean_ctor_get(x_20, 0); -lean_inc(x_65); -if (lean_is_exclusive(x_20)) { - lean_ctor_release(x_20, 0); - x_66 = x_20; +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_inc(x_1); +x_236 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_Choice_updateGen(x_1, x_37); +x_237 = lean_nat_sub(x_3, x_176); +lean_dec(x_176); +lean_inc(x_2); +x_238 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_238, 0, x_2); +lean_ctor_set(x_238, 1, x_237); +lean_ctor_set(x_238, 2, x_175); +x_239 = lean_ctor_get(x_236, 0); +lean_inc(x_239); +x_240 = lean_ctor_get(x_236, 1); +lean_inc(x_240); +x_241 = lean_ctor_get(x_236, 2); +lean_inc(x_241); +if (lean_is_exclusive(x_236)) { + lean_ctor_release(x_236, 0); + lean_ctor_release(x_236, 1); + lean_ctor_release(x_236, 2); + x_242 = x_236; } else { - lean_dec_ref(x_20); - x_66 = lean_box(0); + lean_dec_ref(x_236); + x_242 = lean_box(0); } -x_67 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__1; -x_68 = 0; -x_69 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__1; -x_70 = lean_alloc_ctor(0, 2, 17); -lean_ctor_set(x_70, 0, x_1); -lean_ctor_set(x_70, 1, x_69); -lean_ctor_set_float(x_70, sizeof(void*)*2, x_67); -lean_ctor_set_float(x_70, sizeof(void*)*2 + 8, x_67); -lean_ctor_set_uint8(x_70, sizeof(void*)*2 + 16, x_68); -x_71 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2; -x_72 = lean_alloc_ctor(9, 3, 0); -lean_ctor_set(x_72, 0, x_70); -lean_ctor_set(x_72, 1, x_16); -lean_ctor_set(x_72, 2, x_71); -lean_inc(x_14); -lean_ctor_set(x_18, 1, x_72); -lean_ctor_set(x_18, 0, x_14); -x_73 = l_Lean_PersistentArray_push___rarg(x_65, x_18); -if (lean_is_scalar(x_66)) { - x_74 = lean_alloc_ctor(0, 1, 8); +x_243 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_243, 0, x_238); +lean_ctor_set(x_243, 1, x_239); +if (lean_is_scalar(x_242)) { + x_244 = lean_alloc_ctor(0, 3, 0); } else { - x_74 = x_66; + x_244 = x_242; } -lean_ctor_set(x_74, 0, x_73); -lean_ctor_set_uint64(x_74, sizeof(void*)*1, x_64); -x_75 = lean_alloc_ctor(0, 8, 0); -lean_ctor_set(x_75, 0, x_57); -lean_ctor_set(x_75, 1, x_58); -lean_ctor_set(x_75, 2, x_59); -lean_ctor_set(x_75, 3, x_74); -lean_ctor_set(x_75, 4, x_60); -lean_ctor_set(x_75, 5, x_61); -lean_ctor_set(x_75, 6, x_62); -lean_ctor_set(x_75, 7, x_63); -x_76 = lean_st_ref_set(x_12, x_75, x_22); -x_77 = lean_ctor_get(x_76, 1); -lean_inc(x_77); -if (lean_is_exclusive(x_76)) { - lean_ctor_release(x_76, 0); - lean_ctor_release(x_76, 1); - x_78 = x_76; -} else { - lean_dec_ref(x_76); - x_78 = lean_box(0); +lean_ctor_set(x_244, 0, x_243); +lean_ctor_set(x_244, 1, x_240); +lean_ctor_set(x_244, 2, x_241); +x_245 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_pushChoice(x_244, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_95); +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_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_4, x_6, x_246, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_247); +lean_dec(x_246); +x_18 = x_248; +goto block_33; } -x_79 = lean_box(0); -if (lean_is_scalar(x_78)) { - x_80 = lean_alloc_ctor(0, 2, 0); -} else { - x_80 = x_78; } -lean_ctor_set(x_80, 0, x_79); -lean_ctor_set(x_80, 1, x_77); -return x_80; } } else { -lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; uint64_t x_90; lean_object* x_91; lean_object* x_92; double x_93; uint8_t x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; -x_81 = lean_ctor_get(x_18, 1); -lean_inc(x_81); -lean_dec(x_18); -x_82 = lean_ctor_get(x_19, 0); -lean_inc(x_82); -x_83 = lean_ctor_get(x_19, 1); -lean_inc(x_83); -x_84 = lean_ctor_get(x_19, 2); -lean_inc(x_84); -x_85 = lean_ctor_get(x_19, 4); -lean_inc(x_85); -x_86 = lean_ctor_get(x_19, 5); -lean_inc(x_86); -x_87 = lean_ctor_get(x_19, 6); -lean_inc(x_87); -x_88 = lean_ctor_get(x_19, 7); -lean_inc(x_88); -if (lean_is_exclusive(x_19)) { - lean_ctor_release(x_19, 0); - lean_ctor_release(x_19, 1); - lean_ctor_release(x_19, 2); - lean_ctor_release(x_19, 3); - lean_ctor_release(x_19, 4); - lean_ctor_release(x_19, 5); - lean_ctor_release(x_19, 6); - lean_ctor_release(x_19, 7); - x_89 = x_19; -} else { - lean_dec_ref(x_19); - x_89 = lean_box(0); +uint8_t x_249; +lean_dec(x_37); +lean_dec(x_6); +x_249 = !lean_is_exclusive(x_41); +if (x_249 == 0) +{ +x_18 = x_41; +goto block_33; } -x_90 = lean_ctor_get_uint64(x_20, sizeof(void*)*1); -x_91 = lean_ctor_get(x_20, 0); -lean_inc(x_91); -if (lean_is_exclusive(x_20)) { - lean_ctor_release(x_20, 0); - x_92 = x_20; -} else { - lean_dec_ref(x_20); - x_92 = lean_box(0); +else +{ +lean_object* x_250; lean_object* x_251; lean_object* x_252; +x_250 = lean_ctor_get(x_41, 0); +x_251 = lean_ctor_get(x_41, 1); +lean_inc(x_251); +lean_inc(x_250); +lean_dec(x_41); +x_252 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_252, 0, x_250); +lean_ctor_set(x_252, 1, x_251); +x_18 = x_252; +goto block_33; } -x_93 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__1; -x_94 = 0; -x_95 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__1; -x_96 = lean_alloc_ctor(0, 2, 17); -lean_ctor_set(x_96, 0, x_1); -lean_ctor_set(x_96, 1, x_95); -lean_ctor_set_float(x_96, sizeof(void*)*2, x_93); -lean_ctor_set_float(x_96, sizeof(void*)*2 + 8, x_93); -lean_ctor_set_uint8(x_96, sizeof(void*)*2 + 16, x_94); -x_97 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2; -x_98 = lean_alloc_ctor(9, 3, 0); -lean_ctor_set(x_98, 0, x_96); -lean_ctor_set(x_98, 1, x_16); -lean_ctor_set(x_98, 2, x_97); -lean_inc(x_14); -x_99 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_99, 0, x_14); -lean_ctor_set(x_99, 1, x_98); -x_100 = l_Lean_PersistentArray_push___rarg(x_91, x_99); -if (lean_is_scalar(x_92)) { - x_101 = lean_alloc_ctor(0, 1, 8); -} else { - x_101 = x_92; } -lean_ctor_set(x_101, 0, x_100); -lean_ctor_set_uint64(x_101, sizeof(void*)*1, x_90); -if (lean_is_scalar(x_89)) { - x_102 = lean_alloc_ctor(0, 8, 0); -} else { - x_102 = x_89; -} -lean_ctor_set(x_102, 0, x_82); -lean_ctor_set(x_102, 1, x_83); -lean_ctor_set(x_102, 2, x_84); -lean_ctor_set(x_102, 3, x_101); -lean_ctor_set(x_102, 4, x_85); -lean_ctor_set(x_102, 5, x_86); -lean_ctor_set(x_102, 6, x_87); -lean_ctor_set(x_102, 7, x_88); -x_103 = lean_st_ref_set(x_12, x_102, x_81); -x_104 = lean_ctor_get(x_103, 1); -lean_inc(x_104); -if (lean_is_exclusive(x_103)) { - lean_ctor_release(x_103, 0); - lean_ctor_release(x_103, 1); - x_105 = x_103; -} else { - lean_dec_ref(x_103); - x_105 = lean_box(0); -} -x_106 = lean_box(0); -if (lean_is_scalar(x_105)) { - x_107 = lean_alloc_ctor(0, 2, 0); -} else { - x_107 = x_105; -} -lean_ctor_set(x_107, 0, x_106); -lean_ctor_set(x_107, 1, x_104); -return x_107; } } -} -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___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; -x_16 = l_Lean_Meta_Grind_addTheoremInstance(x_1, x_2, x_3, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); -return x_16; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("grind", 5, 5); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("ematch", 6, 6); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__3() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("instance", 8, 8); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__4() { -_start: +uint8_t x_253; +lean_dec(x_6); +x_253 = !lean_is_exclusive(x_34); +if (x_253 == 0) { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__1; -x_2 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__2; -x_3 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__3; -x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); -return x_4; -} +x_18 = x_34; +goto block_33; } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__5() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked(": ", 2, 2); -return x_1; +lean_object* x_254; lean_object* x_255; lean_object* x_256; +x_254 = lean_ctor_get(x_34, 0); +x_255 = lean_ctor_get(x_34, 1); +lean_inc(x_255); +lean_inc(x_254); +lean_dec(x_34); +x_256 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_256, 0, x_254); +lean_ctor_set(x_256, 1, x_255); +x_18 = x_256; +goto block_33; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__6() { -_start: +block_33: { -lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__5; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { -_start: +if (lean_obj_tag(x_18) == 0) { -lean_object* x_16; -lean_inc(x_14); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_1); -x_16 = lean_infer_type(x_1, x_11, x_12, x_13, x_14, x_15); -if (lean_obj_tag(x_16) == 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_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -x_18 = lean_ctor_get(x_16, 1); -lean_inc(x_18); +uint8_t x_20; lean_dec(x_16); -x_19 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__4; -x_20 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__2(x_19, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_18); -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_dec(x_3); -x_23 = lean_ctor_get(x_20, 1); -lean_inc(x_23); -lean_dec(x_20); -x_24 = l_Lean_Meta_Grind_addTheoremInstance(x_1, x_17, x_2, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_23); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -return x_24; -} -else -{ -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; -x_26 = lean_ctor_get(x_20, 1); -x_27 = lean_ctor_get(x_20, 0); -lean_dec(x_27); -x_28 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_3, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_26); -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; 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_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_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; -lean_ctor_set_tag(x_20, 7); -lean_ctor_set(x_20, 1, x_29); -lean_ctor_set(x_20, 0, x_31); -x_32 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__6; -x_33 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_33, 0, x_20); -lean_ctor_set(x_33, 1, x_32); -lean_inc(x_17); -x_34 = l_Lean_MessageData_ofExpr(x_17); -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 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_36, 0, x_35); -lean_ctor_set(x_36, 1, x_31); -x_37 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_19, x_36, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_30); -x_38 = lean_ctor_get(x_37, 1); -lean_inc(x_38); -lean_dec(x_37); -x_39 = l_Lean_Meta_Grind_addTheoremInstance(x_1, x_17, x_2, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_38); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -return x_39; -} -else -{ -uint8_t x_40; -lean_free_object(x_20); -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_2); lean_dec(x_1); -x_40 = !lean_is_exclusive(x_28); -if (x_40 == 0) +x_20 = !lean_is_exclusive(x_18); +if (x_20 == 0) { -return x_28; +lean_object* x_21; lean_object* x_22; +x_21 = lean_ctor_get(x_18, 0); +lean_dec(x_21); +x_22 = lean_ctor_get(x_19, 0); +lean_inc(x_22); +lean_dec(x_19); +lean_ctor_set(x_18, 0, x_22); +return x_18; } else { -lean_object* x_41; lean_object* x_42; lean_object* x_43; -x_41 = lean_ctor_get(x_28, 0); -x_42 = lean_ctor_get(x_28, 1); -lean_inc(x_42); -lean_inc(x_41); -lean_dec(x_28); -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_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_18, 1); +lean_inc(x_23); +lean_dec(x_18); +x_24 = lean_ctor_get(x_19, 0); +lean_inc(x_24); +lean_dec(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_23); +return x_25; } } else { -lean_object* x_44; lean_object* x_45; -x_44 = lean_ctor_get(x_20, 1); -lean_inc(x_44); -lean_dec(x_20); -x_45 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_3, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, 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; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; -x_46 = lean_ctor_get(x_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_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___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_46); -x_50 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__6; -x_51 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_51, 0, x_49); -lean_ctor_set(x_51, 1, x_50); -lean_inc(x_17); -x_52 = l_Lean_MessageData_ofExpr(x_17); -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_48); -x_55 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_19, x_54, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_47); -x_56 = lean_ctor_get(x_55, 1); -lean_inc(x_56); -lean_dec(x_55); -x_57 = l_Lean_Meta_Grind_addTheoremInstance(x_1, x_17, x_2, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_56); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -return x_57; +lean_object* x_26; lean_object* x_27; +x_26 = lean_ctor_get(x_18, 1); +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_6 = x_27; +x_17 = x_26; +goto _start; +} } else { -lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; -lean_dec(x_17); +uint8_t x_29; +lean_dec(x_16); +lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); lean_dec(x_2); lean_dec(x_1); -x_58 = lean_ctor_get(x_45, 0); -lean_inc(x_58); -x_59 = lean_ctor_get(x_45, 1); -lean_inc(x_59); -if (lean_is_exclusive(x_45)) { - lean_ctor_release(x_45, 0); - lean_ctor_release(x_45, 1); - x_60 = x_45; -} else { - lean_dec_ref(x_45); - x_60 = lean_box(0); +x_29 = !lean_is_exclusive(x_18); +if (x_29 == 0) +{ +return x_18; } -if (lean_is_scalar(x_60)) { - x_61 = lean_alloc_ctor(1, 2, 0); -} else { - x_61 = x_60; +else +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_30 = lean_ctor_get(x_18, 0); +x_31 = lean_ctor_get(x_18, 1); +lean_inc(x_31); +lean_inc(x_30); +lean_dec(x_18); +x_32 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_32, 0, x_30); +lean_ctor_set(x_32, 1, x_31); +return x_32; } -lean_ctor_set(x_61, 0, x_58); -lean_ctor_set(x_61, 1, x_59); -return x_61; } } } } -else +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processOffset(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +_start: { -uint8_t x_62; -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_62 = !lean_is_exclusive(x_16); -if (x_62 == 0) -{ -return x_16; -} -else -{ -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_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; -} -} -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = l_Lean_Meta_Grind_grind_debug_proofs; -return x_1; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, 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; -x_15 = l_Lean_instantiateMVars___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__1(x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -x_16 = lean_ctor_get(x_15, 0); -lean_inc(x_16); -x_17 = lean_ctor_get(x_15, 1); +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_16 = l_Lean_Meta_Grind_getMaxGeneration___rarg(x_9, x_10, x_11, x_12, x_13, x_14, x_15); +x_17 = lean_ctor_get(x_16, 0); lean_inc(x_17); -lean_dec(x_15); -x_18 = lean_ctor_get(x_12, 2); +x_18 = lean_ctor_get(x_16, 1); lean_inc(x_18); -x_19 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___closed__1; -x_20 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_18, x_19); -lean_dec(x_18); +lean_dec(x_16); +lean_inc(x_4); +x_19 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processOffset___spec__1(x_1, x_2, x_3, x_4, x_17, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_18); +lean_dec(x_17); +lean_dec(x_4); +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; -x_21 = lean_box(0); -x_22 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2(x_16, x_3, x_1, x_21, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_17); -return 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_23; -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_16); -x_23 = l_Lean_Meta_check(x_16, x_10, x_11, x_12, x_13, x_17); -if (lean_obj_tag(x_23) == 0) -{ -lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -x_25 = lean_ctor_get(x_23, 1); -lean_inc(x_25); -lean_dec(x_23); -x_26 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2(x_16, x_3, x_1, x_24, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_25); -lean_dec(x_24); -return x_26; +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 { -uint8_t x_27; -lean_dec(x_16); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_3); -lean_dec(x_1); -x_27 = !lean_is_exclusive(x_23); -if (x_27 == 0) +uint8_t x_26; +x_26 = !lean_is_exclusive(x_19); +if (x_26 == 0) { -return x_23; +return x_19; } else { -lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_28 = lean_ctor_get(x_23, 0); -x_29 = lean_ctor_get(x_23, 1); -lean_inc(x_29); +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_19, 0); +x_28 = lean_ctor_get(x_19, 1); lean_inc(x_28); -lean_dec(x_23); -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_19); +x_29 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_29, 0, x_27); +lean_ctor_set(x_29, 1, x_28); +return x_29; } } } } -LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processOffset___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]; _start: { -lean_object* x_13; -x_13 = l_Lean_instantiateMVars___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); +lean_object* x_18; +x_18 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processOffset___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); lean_dec(x_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_13; +return x_18; } } -LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processOffset___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, 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_13; -x_13 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); +lean_object* x_16; +x_16 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processOffset(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_6); lean_dec(x_5); -lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); -return x_13; +return x_16; } } -LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___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) { +static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___spec__1___closed__1() { _start: { -lean_object* x_13; -x_13 = l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__6(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_13; +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_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___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_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19, lean_object* x_20) { _start: { -lean_object* x_13; -x_13 = l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___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); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_13; -} -} -LEAN_EXPORT lean_object* l_Lean_mkConstWithLevelParams___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: +if (lean_obj_tag(x_7) == 0) { -lean_object* x_13; -x_13 = l_Lean_mkConstWithLevelParams___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_13; -} +lean_object* x_21; +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_1); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_8); +lean_ctor_set(x_21, 1, x_20); +return x_21; } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: +else { -lean_object* x_13; -x_13 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_30; 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_13; -} +x_22 = lean_ctor_get(x_7, 0); +lean_inc(x_22); +x_23 = lean_ctor_get(x_7, 1); +lean_inc(x_23); +if (lean_is_exclusive(x_7)) { + lean_ctor_release(x_7, 0); + lean_ctor_release(x_7, 1); + x_24 = x_7; +} else { + lean_dec_ref(x_7); + x_24 = lean_box(0); } -LEAN_EXPORT lean_object* l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { -_start: +lean_inc(x_22); +x_30 = l_Lean_Meta_Grind_getENode(x_22, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20); +if (lean_obj_tag(x_30) == 0) { -lean_object* x_14; -x_14 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -return x_14; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { -_start: +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_83; +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, 8); +lean_inc(x_33); +x_83 = lean_nat_dec_lt(x_33, x_4); +if (x_83 == 0) { -lean_object* x_16; -x_16 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -return x_16; +lean_object* x_84; +lean_dec(x_33); +lean_dec(x_31); +lean_dec(x_24); +lean_dec(x_22); +x_84 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___spec__1___closed__1; +x_25 = x_84; +x_26 = x_32; +goto block_29; } +else +{ +uint8_t x_85; +x_85 = lean_ctor_get_uint8(x_31, sizeof(void*)*10 + 4); +if (x_85 == 0) +{ +uint8_t x_86; +x_86 = l_Lean_Meta_Grind_ENode_isCongrRoot(x_31); +lean_dec(x_31); +if (x_86 == 0) +{ +lean_object* x_87; +lean_dec(x_33); +lean_dec(x_24); +lean_dec(x_22); +x_87 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___spec__1___closed__1; +x_25 = x_87; +x_26 = x_32; +goto block_29; } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { -_start: +else { -lean_object* x_16; -x_16 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -return x_16; +lean_object* x_88; +x_88 = lean_box(0); +x_34 = x_88; +goto block_82; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { -_start: +else { -lean_object* x_15; -x_15 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -return x_15; +lean_object* x_89; +lean_dec(x_31); +x_89 = lean_box(0); +x_34 = x_89; +goto block_82; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem_synthesizeInstance(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: +block_82: { -lean_object* x_8; lean_object* x_9; -x_8 = lean_box(0); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -x_9 = l_Lean_Meta_trySynthInstance(x_2, x_8, x_3, x_4, x_5, x_6, x_7); -if (lean_obj_tag(x_9) == 0) +lean_object* x_35; +lean_dec(x_34); +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_1); +x_35 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_x3f(x_1, x_2, x_22, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_32); +lean_dec(x_22); +if (lean_obj_tag(x_35) == 0) { -lean_object* x_10; -x_10 = lean_ctor_get(x_9, 0); -lean_inc(x_10); -if (lean_obj_tag(x_10) == 1) +lean_object* x_36; +x_36 = lean_ctor_get(x_35, 0); +lean_inc(x_36); +if (lean_obj_tag(x_36) == 0) { -lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_11 = lean_ctor_get(x_9, 1); -lean_inc(x_11); -lean_dec(x_9); -x_12 = lean_ctor_get(x_10, 0); -lean_inc(x_12); -lean_dec(x_10); -x_13 = l_Lean_Meta_isExprDefEq(x_1, x_12, x_3, x_4, x_5, x_6, x_11); -return x_13; +lean_object* x_37; lean_object* x_38; +lean_dec(x_33); +lean_dec(x_24); +x_37 = lean_ctor_get(x_35, 1); +lean_inc(x_37); +lean_dec(x_35); +x_38 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___spec__1___closed__1; +x_25 = x_38; +x_26 = x_37; +goto block_29; } else { -uint8_t x_14; -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_14 = !lean_is_exclusive(x_9); -if (x_14 == 0) +lean_object* x_39; lean_object* x_40; uint8_t x_41; +x_39 = lean_ctor_get(x_36, 0); +lean_inc(x_39); +lean_dec(x_36); +x_40 = lean_ctor_get(x_35, 1); +lean_inc(x_40); +lean_dec(x_35); +x_41 = !lean_is_exclusive(x_39); +if (x_41 == 0) { -lean_object* x_15; uint8_t x_16; lean_object* x_17; -x_15 = lean_ctor_get(x_9, 0); -lean_dec(x_15); -x_16 = 0; -x_17 = lean_box(x_16); -lean_ctor_set(x_9, 0, x_17); -return x_9; +lean_object* x_42; uint8_t x_43; +x_42 = lean_ctor_get(x_39, 1); +x_43 = lean_nat_dec_le(x_33, x_42); +if (x_43 == 0) +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; +lean_dec(x_42); +lean_ctor_set(x_39, 1, x_33); +x_44 = lean_st_ref_take(x_11, x_40); +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); +if (lean_is_scalar(x_24)) { + x_47 = lean_alloc_ctor(1, 2, 0); +} else { + x_47 = x_24; +} +lean_ctor_set(x_47, 0, x_39); +lean_ctor_set(x_47, 1, x_45); +x_48 = lean_st_ref_set(x_11, x_47, x_46); +x_49 = lean_ctor_get(x_48, 1); +lean_inc(x_49); +lean_dec(x_48); +x_50 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___spec__1___closed__1; +x_25 = x_50; +x_26 = x_49; +goto block_29; } else { -lean_object* x_18; uint8_t x_19; lean_object* x_20; lean_object* x_21; -x_18 = lean_ctor_get(x_9, 1); -lean_inc(x_18); -lean_dec(x_9); -x_19 = 0; -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_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; +lean_dec(x_33); +x_51 = lean_st_ref_take(x_11, x_40); +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); +if (lean_is_scalar(x_24)) { + x_54 = lean_alloc_ctor(1, 2, 0); +} else { + x_54 = x_24; } +lean_ctor_set(x_54, 0, x_39); +lean_ctor_set(x_54, 1, x_52); +x_55 = lean_st_ref_set(x_11, x_54, x_53); +x_56 = lean_ctor_get(x_55, 1); +lean_inc(x_56); +lean_dec(x_55); +x_57 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___spec__1___closed__1; +x_25 = x_57; +x_26 = x_56; +goto block_29; } } else { -uint8_t x_22; -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_1); -x_22 = !lean_is_exclusive(x_9); -if (x_22 == 0) +lean_object* x_58; lean_object* x_59; lean_object* x_60; uint8_t x_61; +x_58 = lean_ctor_get(x_39, 0); +x_59 = lean_ctor_get(x_39, 1); +x_60 = lean_ctor_get(x_39, 2); +lean_inc(x_60); +lean_inc(x_59); +lean_inc(x_58); +lean_dec(x_39); +x_61 = lean_nat_dec_le(x_33, x_59); +if (x_61 == 0) { -return x_9; +lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; +lean_dec(x_59); +x_62 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_62, 0, x_58); +lean_ctor_set(x_62, 1, x_33); +lean_ctor_set(x_62, 2, x_60); +x_63 = lean_st_ref_take(x_11, x_40); +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); +if (lean_is_scalar(x_24)) { + x_66 = lean_alloc_ctor(1, 2, 0); +} else { + x_66 = x_24; +} +lean_ctor_set(x_66, 0, x_62); +lean_ctor_set(x_66, 1, x_64); +x_67 = lean_st_ref_set(x_11, x_66, x_65); +x_68 = lean_ctor_get(x_67, 1); +lean_inc(x_68); +lean_dec(x_67); +x_69 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___spec__1___closed__1; +x_25 = x_69; +x_26 = x_68; +goto block_29; } else { -lean_object* x_23; lean_object* x_24; lean_object* x_25; -x_23 = lean_ctor_get(x_9, 0); -x_24 = lean_ctor_get(x_9, 1); -lean_inc(x_24); -lean_inc(x_23); -lean_dec(x_9); -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_70; 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_33); +x_70 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_70, 0, x_58); +lean_ctor_set(x_70, 1, x_59); +lean_ctor_set(x_70, 2, x_60); +x_71 = lean_st_ref_take(x_11, x_40); +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); +if (lean_is_scalar(x_24)) { + x_74 = lean_alloc_ctor(1, 2, 0); +} else { + x_74 = x_24; } -LEAN_EXPORT lean_object* l_ReaderT_read___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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) { -_start: -{ -lean_object* x_12; -x_12 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_12, 0, x_1); -lean_ctor_set(x_12, 1, x_11); -return x_12; +lean_ctor_set(x_74, 0, x_70); +lean_ctor_set(x_74, 1, x_72); +x_75 = lean_st_ref_set(x_11, x_74, x_73); +x_76 = lean_ctor_get(x_75, 1); +lean_inc(x_76); +lean_dec(x_75); +x_77 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___spec__1___closed__1; +x_25 = x_77; +x_26 = x_76; +goto block_29; } } -static lean_object* _init_l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__2___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___spec__1___closed__4; -x_2 = l_ReaderT_instMonad___rarg(x_1); -return x_2; } } -static lean_object* _init_l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__2___closed__2() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__2___closed__1; -x_2 = l_instInhabitedPUnit; -x_3 = l_instInhabitedOfMonad___rarg(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__2___closed__3() { -_start: +uint8_t x_78; +lean_dec(x_33); +lean_dec(x_24); +lean_dec(x_23); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_1); +x_78 = !lean_is_exclusive(x_35); +if (x_78 == 0) { -lean_object* x_1; lean_object* x_2; -x_1 = l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__2___closed__2; -x_2 = lean_alloc_closure((void*)(l_instInhabitedReaderT___rarg___boxed), 2, 1); -lean_closure_set(x_2, 0, x_1); -return x_2; -} +return x_35; } -LEAN_EXPORT lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: +else { -lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_13 = l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__2___closed__3; -x_14 = lean_panic_fn(x_13, x_1); -x_15 = lean_apply_11(x_14, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -return x_15; +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_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; } } -static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___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; } } -static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___lambda__1___closed__2() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___lambda__1___closed__1; -x_2 = lean_box(0); -x_3 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set(x_3, 1, x_2); -return x_3; -} -} -static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___lambda__1___closed__3() { -_start: +uint8_t x_90; +lean_dec(x_24); +lean_dec(x_23); +lean_dec(x_22); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_1); +x_90 = !lean_is_exclusive(x_30); +if (x_90 == 0) { -lean_object* x_1; lean_object* x_2; -x_1 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___lambda__1___closed__2; -x_2 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; -} +return x_30; } -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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) { -_start: +else { -lean_object* x_13; lean_object* x_14; -x_13 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___lambda__1___closed__3; -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_91; lean_object* x_92; lean_object* x_93; +x_91 = lean_ctor_get(x_30, 0); +x_92 = lean_ctor_get(x_30, 1); +lean_inc(x_92); +lean_inc(x_91); +lean_dec(x_30); +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; } } -static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__1() { -_start: +block_29: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("issues", 6, 6); -return x_1; +lean_object* x_27; +x_27 = lean_ctor_get(x_25, 0); +lean_inc(x_27); +lean_dec(x_25); +x_7 = x_23; +x_8 = x_27; +x_9 = lean_box(0); +x_20 = x_26; +goto _start; } } -static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__1; -x_2 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__1; -x_3 = l_Lean_Name_mkStr2(x_1, x_2); -return x_3; } } -static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__3() { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___lambda__1___boxed), 12, 0); -return x_1; -} -} -static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__4() { -_start: +lean_object* x_14; uint8_t x_15; +x_14 = lean_st_ref_get(x_5, x_13); +x_15 = !lean_is_exclusive(x_14); +if (x_15 == 0) { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("type error constructing proof for ", 34, 34); -return x_1; -} -} -static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__5() { -_start: +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_16 = lean_ctor_get(x_14, 0); +x_17 = lean_ctor_get(x_14, 1); +x_18 = lean_ctor_get(x_16, 4); +lean_inc(x_18); +lean_dec(x_16); +lean_inc(x_2); +x_19 = l_Lean_Expr_toHeadIndex(x_2); +x_20 = l_Lean_PersistentHashMap_find_x3f___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__1(x_18, x_19); +lean_dec(x_19); +if (lean_obj_tag(x_20) == 0) { -lean_object* x_1; lean_object* x_2; -x_1 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__4; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; -} +lean_object* 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_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_2); +lean_dec(x_1); +x_21 = lean_box(0); +lean_ctor_set(x_14, 0, x_21); +return x_14; } -static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__6() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("\nwhen assigning metavariable ", 29, 29); -return x_1; -} -} -static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__7() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__6; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; -} -} -static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__8() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked(" with ", 6, 6); -return x_1; -} -} -static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__9() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__8; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19, lean_object* x_20, lean_object* x_21, lean_object* x_22, lean_object* x_23) { -_start: -{ -lean_object* x_24; uint8_t x_25; -x_24 = lean_ctor_get(x_8, 1); -x_25 = lean_nat_dec_lt(x_10, x_24); -if (x_25 == 0) -{ -lean_object* x_26; -lean_dec(x_22); -lean_dec(x_21); +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +lean_free_object(x_14); +x_22 = lean_ctor_get(x_20, 0); +lean_inc(x_22); lean_dec(x_20); -lean_dec(x_19); -lean_dec(x_18); -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_10); -lean_dec(x_7); -lean_dec(x_1); -x_26 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_26, 0, x_9); -lean_ctor_set(x_26, 1, x_23); -return x_26; -} -else -{ -lean_object* x_27; lean_object* x_28; lean_object* x_36; lean_object* x_165; lean_object* x_166; lean_object* x_167; uint8_t x_168; -lean_dec(x_9); -x_165 = lean_nat_sub(x_3, x_10); -x_166 = lean_unsigned_to_nat(1u); -x_167 = lean_nat_sub(x_165, x_166); -lean_dec(x_165); -x_168 = lean_nat_dec_lt(x_167, x_4); -if (x_168 == 0) -{ -lean_object* x_169; lean_object* x_170; -lean_dec(x_167); -x_169 = l_Lean_instInhabitedExpr; -x_170 = l_outOfBounds___rarg(x_169); -x_36 = x_170; -goto block_164; -} -else -{ -lean_object* x_171; -x_171 = lean_array_fget(x_2, x_167); -lean_dec(x_167); -x_36 = x_171; -goto block_164; -} -block_35: +x_23 = l_Lean_Meta_Grind_getMaxGeneration___rarg(x_7, x_8, x_9, x_10, x_11, x_12, x_17); +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +x_25 = lean_ctor_get(x_23, 1); +lean_inc(x_25); +lean_dec(x_23); +x_26 = lean_box(0); +x_27 = lean_box(0); +lean_inc(x_22); +x_28 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___spec__1(x_1, x_2, x_22, x_24, x_26, x_22, x_22, x_27, lean_box(0), x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_25); +lean_dec(x_24); +lean_dec(x_22); +lean_dec(x_2); +if (lean_obj_tag(x_28) == 0) { -if (lean_obj_tag(x_27) == 0) +uint8_t x_29; +x_29 = !lean_is_exclusive(x_28); +if (x_29 == 0) { -lean_object* x_29; lean_object* x_30; -lean_dec(x_22); -lean_dec(x_21); -lean_dec(x_20); -lean_dec(x_19); -lean_dec(x_18); -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_10); -lean_dec(x_7); -lean_dec(x_1); -x_29 = lean_ctor_get(x_27, 0); -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_object* x_30; +x_30 = lean_ctor_get(x_28, 0); +lean_dec(x_30); +lean_ctor_set(x_28, 0, 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); +lean_object* x_31; lean_object* x_32; +x_31 = lean_ctor_get(x_28, 1); lean_inc(x_31); -lean_dec(x_27); -x_32 = lean_ctor_get(x_8, 2); -x_33 = lean_nat_add(x_10, x_32); -lean_dec(x_10); -x_9 = x_31; -x_10 = x_33; -x_11 = lean_box(0); -x_12 = lean_box(0); -x_23 = x_28; -goto _start; -} -} -block_164: -{ -lean_object* x_37; uint8_t x_38; -x_37 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_unassigned; -x_38 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_36, x_37); -if (x_38 == 0) -{ -lean_object* x_39; uint8_t x_40; lean_object* x_41; lean_object* x_130; lean_object* x_131; -x_39 = lean_array_fget(x_5, x_10); -x_130 = l_Lean_Expr_mvarId_x21(x_39); -lean_inc(x_130); -x_131 = l_Lean_MVarId_getType(x_130, x_19, x_20, x_21, x_22, x_23); -if (lean_obj_tag(x_131) == 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_inc(x_22); -lean_inc(x_21); -lean_inc(x_20); -lean_inc(x_19); -lean_inc(x_36); -x_134 = lean_infer_type(x_36, x_19, x_20, x_21, x_22, x_133); -if (lean_obj_tag(x_134) == 0) -{ -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_22); -lean_inc(x_21); -lean_inc(x_20); -lean_inc(x_19); -x_137 = l_Lean_Meta_isExprDefEq(x_132, x_135, x_19, x_20, x_21, x_22, x_136); -if (lean_obj_tag(x_137) == 0) -{ -lean_object* x_138; uint8_t x_139; -x_138 = lean_ctor_get(x_137, 0); -lean_inc(x_138); -x_139 = lean_unbox(x_138); -if (x_139 == 0) -{ -lean_object* x_140; uint8_t x_141; -lean_dec(x_130); -x_140 = lean_ctor_get(x_137, 1); -lean_inc(x_140); -lean_dec(x_137); -x_141 = lean_unbox(x_138); -lean_dec(x_138); -x_40 = x_141; -x_41 = x_140; -goto block_129; +lean_dec(x_28); +x_32 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_32, 0, x_27); +lean_ctor_set(x_32, 1, x_31); +return x_32; } -else -{ -lean_object* x_142; lean_object* x_143; -lean_dec(x_138); -x_142 = lean_ctor_get(x_137, 1); -lean_inc(x_142); -lean_dec(x_137); -lean_inc(x_22); -lean_inc(x_21); -lean_inc(x_20); -lean_inc(x_19); -lean_inc(x_36); -x_143 = lean_checked_assign(x_130, x_36, x_19, x_20, x_21, x_22, x_142); -if (lean_obj_tag(x_143) == 0) -{ -lean_object* x_144; lean_object* x_145; uint8_t x_146; -x_144 = lean_ctor_get(x_143, 0); -lean_inc(x_144); -x_145 = lean_ctor_get(x_143, 1); -lean_inc(x_145); -lean_dec(x_143); -x_146 = lean_unbox(x_144); -lean_dec(x_144); -x_40 = x_146; -x_41 = x_145; -goto block_129; } else { -uint8_t x_147; -lean_dec(x_39); -lean_dec(x_36); -lean_dec(x_22); -lean_dec(x_21); -lean_dec(x_20); -lean_dec(x_19); -lean_dec(x_18); -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_10); -lean_dec(x_7); -lean_dec(x_1); -x_147 = !lean_is_exclusive(x_143); -if (x_147 == 0) +uint8_t x_33; +x_33 = !lean_is_exclusive(x_28); +if (x_33 == 0) { -return x_143; +return x_28; } else { -lean_object* x_148; lean_object* x_149; lean_object* x_150; -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); -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_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_28, 0); +x_35 = lean_ctor_get(x_28, 1); +lean_inc(x_35); +lean_inc(x_34); +lean_dec(x_28); +x_36 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_36, 0, x_34); +lean_ctor_set(x_36, 1, x_35); +return x_36; } } } } else { -uint8_t x_151; -lean_dec(x_130); -lean_dec(x_39); -lean_dec(x_36); -lean_dec(x_22); -lean_dec(x_21); -lean_dec(x_20); -lean_dec(x_19); -lean_dec(x_18); -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_15); +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_37 = lean_ctor_get(x_14, 0); +x_38 = lean_ctor_get(x_14, 1); +lean_inc(x_38); +lean_inc(x_37); lean_dec(x_14); -lean_dec(x_13); +x_39 = lean_ctor_get(x_37, 4); +lean_inc(x_39); +lean_dec(x_37); +lean_inc(x_2); +x_40 = l_Lean_Expr_toHeadIndex(x_2); +x_41 = l_Lean_PersistentHashMap_find_x3f___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__1(x_39, x_40); +lean_dec(x_40); +if (lean_obj_tag(x_41) == 0) +{ +lean_object* x_42; lean_object* 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_2); lean_dec(x_1); -x_151 = !lean_is_exclusive(x_137); -if (x_151 == 0) -{ -return x_137; +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 { -lean_object* x_152; lean_object* x_153; lean_object* x_154; -x_152 = lean_ctor_get(x_137, 0); -x_153 = lean_ctor_get(x_137, 1); -lean_inc(x_153); -lean_inc(x_152); -lean_dec(x_137); -x_154 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_154, 0, x_152); -lean_ctor_set(x_154, 1, x_153); -return x_154; +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_44 = lean_ctor_get(x_41, 0); +lean_inc(x_44); +lean_dec(x_41); +x_45 = l_Lean_Meta_Grind_getMaxGeneration___rarg(x_7, x_8, x_9, x_10, x_11, x_12, x_38); +x_46 = lean_ctor_get(x_45, 0); +lean_inc(x_46); +x_47 = lean_ctor_get(x_45, 1); +lean_inc(x_47); +lean_dec(x_45); +x_48 = lean_box(0); +x_49 = lean_box(0); +lean_inc(x_44); +x_50 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___spec__1(x_1, x_2, x_44, x_46, x_48, x_44, x_44, x_49, lean_box(0), x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_47); +lean_dec(x_46); +lean_dec(x_44); +lean_dec(x_2); +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, 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); } +if (lean_is_scalar(x_52)) { + x_53 = lean_alloc_ctor(0, 2, 0); +} else { + x_53 = x_52; } +lean_ctor_set(x_53, 0, x_49); +lean_ctor_set(x_53, 1, x_51); +return x_53; } else { -uint8_t x_155; -lean_dec(x_132); -lean_dec(x_130); -lean_dec(x_39); -lean_dec(x_36); -lean_dec(x_22); -lean_dec(x_21); -lean_dec(x_20); -lean_dec(x_19); -lean_dec(x_18); -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_10); -lean_dec(x_7); -lean_dec(x_1); -x_155 = !lean_is_exclusive(x_134); -if (x_155 == 0) -{ -return x_134; +lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_54 = lean_ctor_get(x_50, 0); +lean_inc(x_54); +x_55 = lean_ctor_get(x_50, 1); +lean_inc(x_55); +if (lean_is_exclusive(x_50)) { + lean_ctor_release(x_50, 0); + lean_ctor_release(x_50, 1); + x_56 = x_50; +} else { + lean_dec_ref(x_50); + x_56 = lean_box(0); } -else -{ -lean_object* x_156; lean_object* x_157; lean_object* x_158; -x_156 = lean_ctor_get(x_134, 0); -x_157 = lean_ctor_get(x_134, 1); -lean_inc(x_157); -lean_inc(x_156); -lean_dec(x_134); -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; +if (lean_is_scalar(x_56)) { + x_57 = lean_alloc_ctor(1, 2, 0); +} else { + x_57 = x_56; } +lean_ctor_set(x_57, 0, x_54); +lean_ctor_set(x_57, 1, x_55); +return x_57; } } -else +} +} +} +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___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_object* x_19 = _args[18]; +lean_object* x_20 = _args[19]; +_start: { -uint8_t x_159; -lean_dec(x_130); -lean_dec(x_39); -lean_dec(x_36); -lean_dec(x_22); -lean_dec(x_21); -lean_dec(x_20); -lean_dec(x_19); -lean_dec(x_18); -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); +lean_object* x_21; +x_21 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20); +lean_dec(x_11); lean_dec(x_10); -lean_dec(x_7); -lean_dec(x_1); -x_159 = !lean_is_exclusive(x_131); -if (x_159 == 0) +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_21; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, 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: { -return x_131; +lean_object* x_14; +x_14 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_4); +lean_dec(x_3); +return x_14; } -else +} +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___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, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: { -lean_object* x_160; lean_object* x_161; lean_object* x_162; -x_160 = lean_ctor_get(x_131, 0); -x_161 = lean_ctor_get(x_131, 1); -lean_inc(x_161); -lean_inc(x_160); -lean_dec(x_131); -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_14; +x_14 = lean_apply_12(x_1, x_8, x_2, x_3, x_4, x_5, x_6, x_7, x_9, x_10, x_11, x_12, x_13); +return x_14; } } -block_129: +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___spec__1___rarg(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +_start: { -if (x_40 == 0) +lean_object* x_17; lean_object* x_18; +x_17 = lean_alloc_closure((void*)(l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___spec__1___rarg___lambda__1), 13, 7); +lean_closure_set(x_17, 0, x_4); +lean_closure_set(x_17, 1, x_6); +lean_closure_set(x_17, 2, x_7); +lean_closure_set(x_17, 3, x_8); +lean_closure_set(x_17, 4, x_9); +lean_closure_set(x_17, 5, x_10); +lean_closure_set(x_17, 6, x_11); +x_18 = l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDeclImp___rarg(x_1, x_2, x_3, x_17, x_5, x_12, x_13, x_14, x_15, x_16); +if (lean_obj_tag(x_18) == 0) { -lean_object* x_42; lean_object* x_43; uint8_t x_44; -x_42 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__2; -x_43 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__2(x_42, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_41); -x_44 = !lean_is_exclusive(x_43); -if (x_44 == 0) +uint8_t x_19; +x_19 = !lean_is_exclusive(x_18); +if (x_19 == 0) { -lean_object* x_45; lean_object* x_46; lean_object* x_47; uint8_t x_48; -x_45 = lean_ctor_get(x_43, 0); -x_46 = lean_ctor_get(x_43, 1); -x_47 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__3; -x_48 = lean_unbox(x_45); -lean_dec(x_45); -if (x_48 == 0) +return x_18; +} +else { -lean_object* x_49; lean_object* x_50; -lean_free_object(x_43); -lean_dec(x_39); -lean_dec(x_36); -x_49 = lean_box(0); -lean_inc(x_22); +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_18, 0); +x_21 = lean_ctor_get(x_18, 1); lean_inc(x_21); lean_inc(x_20); -lean_inc(x_19); -lean_inc(x_18); -lean_inc(x_17); -lean_inc(x_16); -lean_inc(x_15); -lean_inc(x_14); -lean_inc(x_13); -x_50 = lean_apply_12(x_47, x_49, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_46); -if (lean_obj_tag(x_50) == 0) +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 { -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_inc(x_52); -lean_dec(x_50); -x_27 = x_51; -x_28 = x_52; -goto block_35; +uint8_t x_23; +x_23 = !lean_is_exclusive(x_18); +if (x_23 == 0) +{ +return x_18; } else { -uint8_t x_53; -lean_dec(x_22); -lean_dec(x_21); -lean_dec(x_20); -lean_dec(x_19); +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_18, 0); +x_25 = lean_ctor_get(x_18, 1); +lean_inc(x_25); +lean_inc(x_24); lean_dec(x_18); -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_10); -lean_dec(x_7); -lean_dec(x_1); -x_53 = !lean_is_exclusive(x_50); -if (x_53 == 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); +return x_26; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___spec__1(lean_object* x_1) { +_start: { -return x_50; +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___spec__1___rarg___boxed), 16, 0); +return x_2; } -else +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___lambda__1___closed__1() { +_start: { -lean_object* x_54; lean_object* x_55; lean_object* x_56; -x_54 = lean_ctor_get(x_50, 0); -x_55 = lean_ctor_get(x_50, 1); -lean_inc(x_55); -lean_inc(x_54); -lean_dec(x_50); -x_56 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_56, 0, x_54); -lean_ctor_set(x_56, 1, x_55); -return x_56; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean", 4, 4); +return x_1; } } +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___lambda__1___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Grind", 5, 5); +return x_1; } -else +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___lambda__1___closed__3() { +_start: { -lean_object* x_57; lean_object* x_58; -x_57 = lean_ctor_get(x_1, 5); -lean_inc(x_57); -x_58 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_57, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_46); -if (lean_obj_tag(x_58) == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("EqMatch", 7, 7); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___lambda__1___closed__4() { +_start: { -lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; 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_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_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__5; -lean_ctor_set_tag(x_43, 7); -lean_ctor_set(x_43, 1, x_59); -lean_ctor_set(x_43, 0, x_61); -x_62 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__7; -x_63 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_63, 0, x_43); -lean_ctor_set(x_63, 1, x_62); -x_64 = l_Lean_MessageData_ofExpr(x_39); -x_65 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_65, 0, x_63); -lean_ctor_set(x_65, 1, x_64); -x_66 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__9; -x_67 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_67, 0, x_65); -lean_ctor_set(x_67, 1, x_66); -x_68 = l_Lean_indentExpr(x_36); -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 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; -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_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_42, x_71, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_60); -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); -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_inc(x_13); -x_75 = lean_apply_12(x_47, x_73, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_74); -if (lean_obj_tag(x_75) == 0) -{ -lean_object* x_76; lean_object* x_77; -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_27 = x_76; -x_28 = x_77; -goto block_35; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___lambda__1___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___lambda__1___closed__2; +x_3 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___lambda__1___closed__3; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; } -else +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +_start: { -uint8_t x_78; -lean_dec(x_22); -lean_dec(x_21); -lean_dec(x_20); -lean_dec(x_19); -lean_dec(x_18); -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_10); -lean_dec(x_7); -lean_dec(x_1); -x_78 = !lean_is_exclusive(x_75); -if (x_78 == 0) +lean_object* x_17; +x_17 = l_Lean_Meta_Grind_markAsDoNotSimp(x_4, x_12, x_13, x_14, x_15, x_16); +if (lean_obj_tag(x_17) == 0) { -return x_75; +uint8_t x_18; +x_18 = !lean_is_exclusive(x_17); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_19 = lean_ctor_get(x_17, 0); +x_20 = l_Lean_Expr_constLevels_x21(x_2); +x_21 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___lambda__1___closed__4; +x_22 = l_Lean_Expr_const___override(x_21, x_20); +x_23 = l_Lean_mkApp4(x_22, x_3, x_19, x_5, x_1); +lean_ctor_set(x_17, 0, x_23); +return x_17; } else { -lean_object* x_79; lean_object* x_80; lean_object* x_81; -x_79 = lean_ctor_get(x_75, 0); -x_80 = lean_ctor_get(x_75, 1); -lean_inc(x_80); -lean_inc(x_79); -lean_dec(x_75); -x_81 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_81, 0, x_79); -lean_ctor_set(x_81, 1, x_80); -return x_81; -} +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_24 = lean_ctor_get(x_17, 0); +x_25 = lean_ctor_get(x_17, 1); +lean_inc(x_25); +lean_inc(x_24); +lean_dec(x_17); +x_26 = l_Lean_Expr_constLevels_x21(x_2); +x_27 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___lambda__1___closed__4; +x_28 = l_Lean_Expr_const___override(x_27, x_26); +x_29 = l_Lean_mkApp4(x_28, x_3, x_24, x_5, x_1); +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_25); +return x_30; } } else { -uint8_t x_82; -lean_free_object(x_43); -lean_dec(x_39); -lean_dec(x_36); -lean_dec(x_22); -lean_dec(x_21); -lean_dec(x_20); -lean_dec(x_19); -lean_dec(x_18); -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_10); -lean_dec(x_7); +uint8_t x_31; +lean_dec(x_5); +lean_dec(x_3); lean_dec(x_1); -x_82 = !lean_is_exclusive(x_58); -if (x_82 == 0) +x_31 = !lean_is_exclusive(x_17); +if (x_31 == 0) { -return x_58; +return x_17; } else { -lean_object* x_83; lean_object* x_84; lean_object* x_85; -x_83 = lean_ctor_get(x_58, 0); -x_84 = lean_ctor_get(x_58, 1); -lean_inc(x_84); -lean_inc(x_83); -lean_dec(x_58); -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_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_17, 0); +x_33 = lean_ctor_get(x_17, 1); +lean_inc(x_33); +lean_inc(x_32); +lean_dec(x_17); +x_34 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +return x_34; } } } } -else +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: { -lean_object* x_86; lean_object* x_87; lean_object* x_88; uint8_t x_89; -x_86 = lean_ctor_get(x_43, 0); -x_87 = lean_ctor_get(x_43, 1); -lean_inc(x_87); -lean_inc(x_86); -lean_dec(x_43); -x_88 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__3; -x_89 = lean_unbox(x_86); -lean_dec(x_86); -if (x_89 == 0) +lean_object* x_14; +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; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: { -lean_object* x_90; lean_object* x_91; -lean_dec(x_39); -lean_dec(x_36); -x_90 = lean_box(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_inc(x_14); +lean_object* x_15; lean_object* x_16; +x_15 = lean_expr_instantiate1(x_1, x_3); lean_inc(x_13); -x_91 = lean_apply_12(x_88, x_90, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_87); -if (lean_obj_tag(x_91) == 0) -{ -lean_object* x_92; lean_object* x_93; -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_27 = x_92; -x_28 = x_93; -goto block_35; -} -else +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +x_16 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType(x_15, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_16) == 0) { -lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; -lean_dec(x_22); -lean_dec(x_21); -lean_dec(x_20); -lean_dec(x_19); -lean_dec(x_18); -lean_dec(x_17); +lean_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; +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_dec(x_15); -lean_dec(x_14); +x_19 = lean_box(0); +x_20 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_20, 0, x_3); +lean_ctor_set(x_20, 1, x_19); +x_21 = lean_array_mk(x_20); +x_22 = 0; +x_23 = 1; +x_24 = 1; +x_25 = l_Lean_Meta_mkForallFVars(x_21, x_17, x_22, x_23, x_24, x_10, x_11, x_12, x_13, x_18); lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); lean_dec(x_10); -lean_dec(x_7); -lean_dec(x_1); -x_94 = lean_ctor_get(x_91, 0); -lean_inc(x_94); -x_95 = lean_ctor_get(x_91, 1); -lean_inc(x_95); -if (lean_is_exclusive(x_91)) { - lean_ctor_release(x_91, 0); - lean_ctor_release(x_91, 1); - x_96 = x_91; -} else { - lean_dec_ref(x_91); - x_96 = lean_box(0); -} -if (lean_is_scalar(x_96)) { - x_97 = lean_alloc_ctor(1, 2, 0); -} else { - x_97 = x_96; -} -lean_ctor_set(x_97, 0, x_94); -lean_ctor_set(x_97, 1, x_95); -return x_97; -} +lean_dec(x_21); +return x_25; } else { -lean_object* x_98; lean_object* x_99; -x_98 = lean_ctor_get(x_1, 5); -lean_inc(x_98); -x_99 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_98, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_87); -if (lean_obj_tag(x_99) == 0) -{ -lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; -x_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_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__5; -x_103 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_103, 0, x_102); -lean_ctor_set(x_103, 1, x_100); -x_104 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__7; -x_105 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_105, 0, x_103); -lean_ctor_set(x_105, 1, x_104); -x_106 = l_Lean_MessageData_ofExpr(x_39); -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_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__9; -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_indentExpr(x_36); -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 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; -x_113 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_113, 0, x_111); -lean_ctor_set(x_113, 1, x_112); -x_114 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_42, x_113, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_101); -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); -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_inc(x_13); -x_117 = lean_apply_12(x_88, x_115, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_116); -if (lean_obj_tag(x_117) == 0) +uint8_t x_26; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_3); +x_26 = !lean_is_exclusive(x_16); +if (x_26 == 0) { -lean_object* x_118; lean_object* x_119; -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_27 = x_118; -x_28 = x_119; -goto block_35; +return x_16; } else { -lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; -lean_dec(x_22); -lean_dec(x_21); -lean_dec(x_20); -lean_dec(x_19); -lean_dec(x_18); -lean_dec(x_17); +lean_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); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_10); -lean_dec(x_7); -lean_dec(x_1); -x_120 = lean_ctor_get(x_117, 0); -lean_inc(x_120); -x_121 = lean_ctor_get(x_117, 1); -lean_inc(x_121); -if (lean_is_exclusive(x_117)) { - lean_ctor_release(x_117, 0); - lean_ctor_release(x_117, 1); - x_122 = x_117; -} else { - lean_dec_ref(x_117); - x_122 = lean_box(0); +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; } -if (lean_is_scalar(x_122)) { - x_123 = lean_alloc_ctor(1, 2, 0); -} else { - x_123 = x_122; } -lean_ctor_set(x_123, 0, x_120); -lean_ctor_set(x_123, 1, x_121); -return x_123; } } -else +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___closed__1() { +_start: { -lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; -lean_dec(x_39); -lean_dec(x_36); -lean_dec(x_22); -lean_dec(x_21); -lean_dec(x_20); -lean_dec(x_19); -lean_dec(x_18); -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_10); -lean_dec(x_7); -lean_dec(x_1); -x_124 = lean_ctor_get(x_99, 0); -lean_inc(x_124); -x_125 = lean_ctor_get(x_99, 1); -lean_inc(x_125); -if (lean_is_exclusive(x_99)) { - lean_ctor_release(x_99, 0); - lean_ctor_release(x_99, 1); - x_126 = x_99; -} else { - lean_dec_ref(x_99); - x_126 = lean_box(0); -} -if (lean_is_scalar(x_126)) { - x_127 = lean_alloc_ctor(1, 2, 0); -} else { - x_127 = x_126; -} -lean_ctor_set(x_127, 0, x_124); -lean_ctor_set(x_127, 1, x_125); -return x_127; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Eq", 2, 2); +return x_1; } } +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } } -else +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, 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_128; -lean_dec(x_39); -lean_dec(x_36); -lean_inc(x_7); -x_128 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_128, 0, x_7); -x_27 = x_128; -x_28 = x_41; -goto block_35; -} +if (lean_obj_tag(x_1) == 7) +{ +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_1, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_1, 1); +lean_inc(x_15); +x_16 = lean_ctor_get(x_1, 2); +lean_inc(x_16); +x_17 = lean_ctor_get_uint8(x_1, sizeof(void*)*3 + 8); +lean_dec(x_1); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +x_18 = l_Lean_Meta_Grind_markAsDoNotSimp(x_15, x_9, x_10, x_11, x_12, x_13); +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; +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_18, 1); +lean_inc(x_20); +lean_dec(x_18); +x_21 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___lambda__3___boxed), 14, 2); +lean_closure_set(x_21, 0, x_16); +lean_closure_set(x_21, 1, x_2); +x_22 = 0; +x_23 = l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___spec__1___rarg(x_14, x_17, x_19, x_21, x_22, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_20); +return x_23; } +else +{ +uint8_t x_24; +lean_dec(x_16); +lean_dec(x_14); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_24 = !lean_is_exclusive(x_18); +if (x_24 == 0) +{ +return x_18; } else { -lean_object* x_163; -lean_dec(x_36); -lean_inc(x_7); -x_163 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_163, 0, x_7); -x_27 = x_163; -x_28 = x_23; -goto block_35; +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_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_28; uint8_t x_29; +lean_inc(x_1); +x_28 = l_Lean_Expr_cleanupAnnotations(x_1); +x_29 = l_Lean_Expr_isApp(x_28); +if (x_29 == 0) +{ +lean_object* x_30; +lean_dec(x_28); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_1); +lean_ctor_set(x_30, 1, x_13); +return x_30; } +else +{ +lean_object* x_31; lean_object* x_32; uint8_t x_33; +x_31 = l_Lean_Expr_appArg(x_28, lean_box(0)); +x_32 = l_Lean_Expr_appFnCleanup(x_28, lean_box(0)); +x_33 = l_Lean_Expr_isApp(x_32); +if (x_33 == 0) +{ +lean_object* x_34; +lean_dec(x_32); +lean_dec(x_31); +lean_dec(x_12); +lean_dec(x_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_34 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_34, 0, x_1); +lean_ctor_set(x_34, 1, x_13); +return x_34; } -LEAN_EXPORT lean_object* l_Lean_MVarId_isAssigned___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: +else { -lean_object* x_13; uint8_t x_14; -x_13 = lean_st_ref_get(x_9, x_12); -x_14 = !lean_is_exclusive(x_13); -if (x_14 == 0) +lean_object* x_35; lean_object* x_36; uint8_t x_37; +x_35 = l_Lean_Expr_appArg(x_32, lean_box(0)); +x_36 = l_Lean_Expr_appFnCleanup(x_32, lean_box(0)); +x_37 = l_Lean_Expr_isApp(x_36); +if (x_37 == 0) { -lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; lean_object* x_19; -x_15 = lean_ctor_get(x_13, 0); -x_16 = lean_ctor_get(x_15, 0); -lean_inc(x_16); -lean_dec(x_15); -x_17 = lean_ctor_get(x_16, 7); -lean_inc(x_17); -lean_dec(x_16); -x_18 = l_Lean_PersistentHashMap_contains___at_Lean_MVarId_isAssigned___spec__1(x_17, x_1); -x_19 = lean_box(x_18); -lean_ctor_set(x_13, 0, x_19); -return x_13; +lean_object* x_38; +lean_dec(x_36); +lean_dec(x_35); +lean_dec(x_31); +lean_dec(x_12); +lean_dec(x_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_38 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_38, 0, x_1); +lean_ctor_set(x_38, 1, x_13); +return x_38; } else { -lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; lean_object* x_25; lean_object* x_26; -x_20 = lean_ctor_get(x_13, 0); -x_21 = lean_ctor_get(x_13, 1); -lean_inc(x_21); -lean_inc(x_20); -lean_dec(x_13); -x_22 = lean_ctor_get(x_20, 0); -lean_inc(x_22); -lean_dec(x_20); -x_23 = lean_ctor_get(x_22, 7); -lean_inc(x_23); -lean_dec(x_22); -x_24 = l_Lean_PersistentHashMap_contains___at_Lean_MVarId_isAssigned___spec__1(x_23, x_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_21); -return x_26; +lean_object* x_39; lean_object* x_40; lean_object* x_41; uint8_t x_42; +x_39 = l_Lean_Expr_appArg(x_36, lean_box(0)); +x_40 = l_Lean_Expr_appFnCleanup(x_36, lean_box(0)); +x_41 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___closed__2; +x_42 = l_Lean_Expr_isConstOf(x_40, x_41); +if (x_42 == 0) +{ +lean_object* x_43; +lean_dec(x_40); +lean_dec(x_39); +lean_dec(x_35); +lean_dec(x_31); +lean_dec(x_12); +lean_dec(x_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_43 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_43, 0, x_1); +lean_ctor_set(x_43, 1, x_13); +return x_43; +} +else +{ +lean_object* x_44; +lean_dec(x_1); +x_44 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___lambda__1(x_2, x_40, x_39, x_35, x_31, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_40); +return x_44; } } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___spec__1___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { _start: { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_14 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___lambda__1___closed__1; -x_15 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_15, 0, x_14); -lean_ctor_set(x_15, 1, x_1); -x_16 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_16, 0, 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; +uint8_t x_17; uint8_t x_18; lean_object* x_19; +x_17 = lean_unbox(x_2); +lean_dec(x_2); +x_18 = lean_unbox(x_5); +lean_dec(x_5); +x_19 = l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___spec__1___rarg(x_1, x_17, x_3, x_4, x_18, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +return x_19; } } -static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___closed__1() { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("failed to synthesize instance when instantiating ", 49, 49); -return x_1; +lean_object* x_17; +x_17 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_2); +return x_17; } } -static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___closed__2() { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___closed__1; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; +lean_object* x_14; +x_14 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_14; } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, size_t x_6, size_t x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___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_20; -x_20 = lean_usize_dec_lt(x_7, x_6); -if (x_20 == 0) -{ -lean_object* x_21; -lean_dec(x_18); -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_3); +lean_object* x_15; +x_15 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); lean_dec(x_1); -x_21 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_21, 0, x_8); -lean_ctor_set(x_21, 1, x_19); -return x_21; +return x_15; } -else -{ -lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_32; -x_22 = lean_array_uget(x_5, x_7); -x_32 = !lean_is_exclusive(x_8); -if (x_32 == 0) +} +LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___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: { -lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; uint8_t x_38; -x_33 = lean_ctor_get(x_8, 1); -x_34 = lean_ctor_get(x_8, 0); -lean_dec(x_34); -x_35 = lean_ctor_get(x_33, 0); -lean_inc(x_35); -x_36 = lean_ctor_get(x_33, 1); -lean_inc(x_36); -x_37 = lean_ctor_get(x_33, 2); -lean_inc(x_37); -x_38 = lean_nat_dec_lt(x_36, x_37); -if (x_38 == 0) +uint8_t x_13; +x_13 = l_Lean_Expr_hasMVar(x_1); +if (x_13 == 0) { -lean_object* x_39; -lean_dec(x_37); -lean_dec(x_36); -lean_dec(x_35); -lean_dec(x_22); -lean_inc(x_3); -lean_ctor_set(x_8, 0, x_3); -x_39 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_39, 0, x_8); -x_23 = x_39; -x_24 = x_19; -goto block_31; +lean_object* x_14; +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_1); +lean_ctor_set(x_14, 1, x_12); +return x_14; } else { -uint8_t x_40; -x_40 = !lean_is_exclusive(x_33); -if (x_40 == 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; uint8_t x_25; +x_15 = lean_st_ref_get(x_9, 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 = lean_ctor_get(x_16, 0); +lean_inc(x_18); +lean_dec(x_16); +x_19 = l_Lean_instantiateMVarsCore(x_18, x_1); +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_st_ref_take(x_9, x_17); +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_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; uint8_t x_50; -x_41 = lean_ctor_get(x_33, 2); -lean_dec(x_41); -x_42 = lean_ctor_get(x_33, 1); -lean_dec(x_42); -x_43 = lean_ctor_get(x_33, 0); -lean_dec(x_43); -x_44 = lean_array_fget(x_35, x_36); -x_45 = lean_unbox(x_44); -lean_dec(x_44); -x_46 = lean_unsigned_to_nat(1u); -x_47 = lean_nat_add(x_36, x_46); -lean_dec(x_36); -lean_ctor_set(x_33, 1, x_47); -x_48 = l_Lean_Expr_mvarId_x21(x_22); -x_49 = l_Lean_MVarId_isAssigned___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__4(x_48, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19); -lean_dec(x_48); -x_50 = !lean_is_exclusive(x_49); -if (x_50 == 0) +lean_object* x_26; lean_object* x_27; uint8_t x_28; +x_26 = lean_ctor_get(x_23, 0); +lean_dec(x_26); +lean_ctor_set(x_23, 0, x_21); +x_27 = lean_st_ref_set(x_9, x_23, x_24); +x_28 = !lean_is_exclusive(x_27); +if (x_28 == 0) { -lean_object* x_51; lean_object* x_52; uint8_t x_53; -x_51 = lean_ctor_get(x_49, 0); -x_52 = lean_ctor_get(x_49, 1); -x_53 = l_Lean_BinderInfo_isInstImplicit(x_45); -if (x_53 == 0) +lean_object* x_29; +x_29 = lean_ctor_get(x_27, 0); +lean_dec(x_29); +lean_ctor_set(x_27, 0, x_20); +return x_27; +} +else { -lean_object* x_54; -lean_free_object(x_49); -lean_dec(x_51); -lean_dec(x_22); -lean_inc(x_3); -lean_ctor_set(x_8, 0, x_3); -x_54 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_54, 0, x_8); -x_23 = x_54; -x_24 = x_52; -goto block_31; +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_20); +lean_ctor_set(x_31, 1, x_30); +return x_31; +} } else { -uint8_t x_55; -x_55 = lean_unbox(x_51); -lean_dec(x_51); -if (x_55 == 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; +x_32 = lean_ctor_get(x_23, 1); +x_33 = lean_ctor_get(x_23, 2); +x_34 = lean_ctor_get(x_23, 3); +x_35 = lean_ctor_get(x_23, 4); +lean_inc(x_35); +lean_inc(x_34); +lean_inc(x_33); +lean_inc(x_32); +lean_dec(x_23); +x_36 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_36, 0, x_21); +lean_ctor_set(x_36, 1, x_32); +lean_ctor_set(x_36, 2, x_33); +lean_ctor_set(x_36, 3, x_34); +lean_ctor_set(x_36, 4, x_35); +x_37 = lean_st_ref_set(x_9, x_36, x_24); +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); +} +if (lean_is_scalar(x_39)) { + x_40 = lean_alloc_ctor(0, 2, 0); +} else { + x_40 = x_39; +} +lean_ctor_set(x_40, 0, x_20); +lean_ctor_set(x_40, 1, x_38); +return x_40; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: { -lean_object* x_56; -lean_inc(x_18); -lean_inc(x_17); -lean_inc(x_16); -lean_inc(x_15); -lean_inc(x_22); -x_56 = lean_infer_type(x_22, x_15, x_16, x_17, x_18, x_52); -if (lean_obj_tag(x_56) == 0) +lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_13 = lean_ctor_get(x_10, 2); +x_14 = l_Lean_isTracingEnabledForCore(x_1, x_13, x_12); +x_15 = !lean_is_exclusive(x_14); +if (x_15 == 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_56, 1); -lean_inc(x_58); -lean_dec(x_56); -lean_inc(x_18); +return x_14; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_14, 0); +x_17 = lean_ctor_get(x_14, 1); lean_inc(x_17); lean_inc(x_16); -lean_inc(x_15); -lean_inc(x_57); -x_59 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem_synthesizeInstance(x_22, x_57, x_15, x_16, x_17, x_18, 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; lean_object* x_64; lean_object* x_65; uint8_t x_66; -lean_free_object(x_8); -x_62 = lean_ctor_get(x_59, 1); -lean_inc(x_62); -lean_dec(x_59); -x_63 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__2; -x_64 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__2(x_63, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_62); -x_65 = lean_ctor_get(x_64, 0); -lean_inc(x_65); -x_66 = lean_unbox(x_65); -lean_dec(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_dec(x_57); -lean_free_object(x_49); -x_67 = lean_ctor_get(x_64, 1); -lean_inc(x_67); -lean_dec(x_64); -x_68 = lean_box(0); -x_69 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___lambda__1(x_33, x_68, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_67); -x_70 = lean_ctor_get(x_69, 0); -lean_inc(x_70); -x_71 = lean_ctor_get(x_69, 1); -lean_inc(x_71); -lean_dec(x_69); -x_23 = x_70; -x_24 = x_71; -goto block_31; +lean_dec(x_14); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_16); +lean_ctor_set(x_18, 1, x_17); +return x_18; } -else +} +} +LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___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) { +_start: { -uint8_t x_72; -x_72 = !lean_is_exclusive(x_64); -if (x_72 == 0) +lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_13 = lean_ctor_get(x_10, 5); +x_14 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_1, x_8, x_9, x_10, x_11, x_12); +x_15 = !lean_is_exclusive(x_14); +if (x_15 == 0) { -lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; -x_73 = lean_ctor_get(x_64, 1); -x_74 = lean_ctor_get(x_64, 0); -lean_dec(x_74); -x_75 = lean_ctor_get(x_1, 5); -lean_inc(x_75); -x_76 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_75, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_73); -if (lean_obj_tag(x_76) == 0) -{ -lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; -x_77 = lean_ctor_get(x_76, 0); -lean_inc(x_77); -x_78 = lean_ctor_get(x_76, 1); -lean_inc(x_78); -lean_dec(x_76); -x_79 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___closed__2; -lean_ctor_set_tag(x_64, 7); -lean_ctor_set(x_64, 1, x_77); -lean_ctor_set(x_64, 0, x_79); -x_80 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; -lean_ctor_set_tag(x_49, 7); -lean_ctor_set(x_49, 1, x_80); -lean_ctor_set(x_49, 0, x_64); -x_81 = l_Lean_indentExpr(x_57); -x_82 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_82, 0, x_49); -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_80); -x_84 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_63, x_83, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_78); -x_85 = lean_ctor_get(x_84, 0); -lean_inc(x_85); -x_86 = lean_ctor_get(x_84, 1); -lean_inc(x_86); -lean_dec(x_84); -x_87 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___lambda__1(x_33, x_85, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_86); -lean_dec(x_85); -x_88 = lean_ctor_get(x_87, 0); -lean_inc(x_88); -x_89 = lean_ctor_get(x_87, 1); -lean_inc(x_89); -lean_dec(x_87); -x_23 = x_88; -x_24 = x_89; -goto block_31; -} -else -{ -uint8_t x_90; -lean_free_object(x_64); -lean_dec(x_57); -lean_free_object(x_49); -lean_dec(x_33); -lean_dec(x_18); -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_3); -lean_dec(x_1); -x_90 = !lean_is_exclusive(x_76); -if (x_90 == 0) -{ -return x_76; +lean_object* x_16; lean_object* x_17; +x_16 = lean_ctor_get(x_14, 0); +lean_inc(x_13); +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_13); +lean_ctor_set(x_17, 1, x_16); +lean_ctor_set_tag(x_14, 1); +lean_ctor_set(x_14, 0, x_17); +return x_14; } else { -lean_object* x_91; lean_object* x_92; lean_object* x_93; -x_91 = lean_ctor_get(x_76, 0); -x_92 = lean_ctor_get(x_76, 1); -lean_inc(x_92); -lean_inc(x_91); -lean_dec(x_76); -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_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_18 = lean_ctor_get(x_14, 0); +x_19 = lean_ctor_get(x_14, 1); +lean_inc(x_19); +lean_inc(x_18); +lean_dec(x_14); +lean_inc(x_13); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_13); +lean_ctor_set(x_20, 1, x_18); +x_21 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_21, 0, x_20); +lean_ctor_set(x_21, 1, x_19); +return x_21; } } -else -{ -lean_object* x_94; lean_object* x_95; lean_object* x_96; -x_94 = lean_ctor_get(x_64, 1); -lean_inc(x_94); -lean_dec(x_64); -x_95 = lean_ctor_get(x_1, 5); -lean_inc(x_95); -x_96 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_95, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_94); -if (lean_obj_tag(x_96) == 0) -{ -lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; 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_97 = lean_ctor_get(x_96, 0); -lean_inc(x_97); -x_98 = lean_ctor_get(x_96, 1); -lean_inc(x_98); -lean_dec(x_96); -x_99 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___closed__2; -x_100 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_100, 0, x_99); -lean_ctor_set(x_100, 1, x_97); -x_101 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; -lean_ctor_set_tag(x_49, 7); -lean_ctor_set(x_49, 1, x_101); -lean_ctor_set(x_49, 0, x_100); -x_102 = l_Lean_indentExpr(x_57); -x_103 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_103, 0, x_49); -lean_ctor_set(x_103, 1, x_102); -x_104 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_104, 0, x_103); -lean_ctor_set(x_104, 1, x_101); -x_105 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_63, x_104, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_98); -x_106 = lean_ctor_get(x_105, 0); -lean_inc(x_106); -x_107 = lean_ctor_get(x_105, 1); -lean_inc(x_107); -lean_dec(x_105); -x_108 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___lambda__1(x_33, x_106, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_107); -lean_dec(x_106); -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_23 = x_109; -x_24 = x_110; -goto block_31; } -else +static lean_object* _init_l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5___closed__1() { +_start: { -lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; -lean_dec(x_57); -lean_free_object(x_49); -lean_dec(x_33); -lean_dec(x_18); -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_3); -lean_dec(x_1); -x_111 = lean_ctor_get(x_96, 0); -lean_inc(x_111); -x_112 = lean_ctor_get(x_96, 1); -lean_inc(x_112); -if (lean_is_exclusive(x_96)) { - lean_ctor_release(x_96, 0); - lean_ctor_release(x_96, 1); - x_113 = x_96; -} else { - lean_dec_ref(x_96); - x_113 = lean_box(0); +lean_object* x_1; +x_1 = lean_mk_string_unchecked("unknown constant '", 18, 18); +return x_1; } -if (lean_is_scalar(x_113)) { - x_114 = lean_alloc_ctor(1, 2, 0); -} else { - x_114 = x_113; } -lean_ctor_set(x_114, 0, x_111); -lean_ctor_set(x_114, 1, x_112); -return x_114; +static lean_object* _init_l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } } +static lean_object* _init_l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("'", 1, 1); +return x_1; } } -else +static lean_object* _init_l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5___closed__4() { +_start: { -lean_object* x_115; lean_object* x_116; -lean_dec(x_57); -lean_free_object(x_49); -x_115 = lean_ctor_get(x_59, 1); -lean_inc(x_115); -lean_dec(x_59); -lean_inc(x_3); -lean_ctor_set(x_8, 0, x_3); -x_116 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_116, 0, x_8); -x_23 = x_116; -x_24 = x_115; -goto block_31; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5___closed__3; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } } -else +LEAN_EXPORT lean_object* l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: { -uint8_t x_117; -lean_dec(x_57); -lean_free_object(x_49); -lean_dec(x_33); -lean_free_object(x_8); -lean_dec(x_18); -lean_dec(x_17); -lean_dec(x_16); +lean_object* x_13; uint8_t x_14; +x_13 = lean_st_ref_get(x_11, x_12); +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; +x_15 = lean_ctor_get(x_13, 0); +x_16 = lean_ctor_get(x_13, 1); +x_17 = lean_ctor_get(x_15, 0); +lean_inc(x_17); lean_dec(x_15); -lean_dec(x_3); -lean_dec(x_1); -x_117 = !lean_is_exclusive(x_59); -if (x_117 == 0) +lean_inc(x_1); +x_18 = lean_environment_find(x_17, x_1); +if (lean_obj_tag(x_18) == 0) { -return x_59; +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_free_object(x_13); +x_19 = 0; +x_20 = l_Lean_MessageData_ofConstName(x_1, x_19); +x_21 = l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5___closed__2; +x_22 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_20); +x_23 = l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5___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_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__6(x_24, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_16); +return x_25; } else { -lean_object* x_118; lean_object* x_119; lean_object* x_120; -x_118 = lean_ctor_get(x_59, 0); -x_119 = lean_ctor_get(x_59, 1); -lean_inc(x_119); -lean_inc(x_118); -lean_dec(x_59); -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; -} +lean_object* x_26; +lean_dec(x_1); +x_26 = lean_ctor_get(x_18, 0); +lean_inc(x_26); +lean_dec(x_18); +lean_ctor_set(x_13, 0, x_26); +return x_13; } } else { -uint8_t x_121; -lean_free_object(x_49); -lean_dec(x_33); -lean_free_object(x_8); -lean_dec(x_22); -lean_dec(x_18); -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_3); -lean_dec(x_1); -x_121 = !lean_is_exclusive(x_56); -if (x_121 == 0) +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +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 = lean_ctor_get(x_27, 0); +lean_inc(x_29); +lean_dec(x_27); +lean_inc(x_1); +x_30 = lean_environment_find(x_29, x_1); +if (lean_obj_tag(x_30) == 0) { -return x_56; +uint8_t x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_31 = 0; +x_32 = l_Lean_MessageData_ofConstName(x_1, x_31); +x_33 = l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5___closed__2; +x_34 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_34, 0, x_33); +lean_ctor_set(x_34, 1, x_32); +x_35 = l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5___closed__4; +x_36 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_36, 0, x_34); +lean_ctor_set(x_36, 1, x_35); +x_37 = l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__6(x_36, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_28); +return x_37; } else { -lean_object* x_122; lean_object* x_123; lean_object* x_124; -x_122 = lean_ctor_get(x_56, 0); -x_123 = lean_ctor_get(x_56, 1); -lean_inc(x_123); -lean_inc(x_122); -lean_dec(x_56); -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; +lean_object* x_38; lean_object* x_39; +lean_dec(x_1); +x_38 = lean_ctor_get(x_30, 0); +lean_inc(x_38); +lean_dec(x_30); +x_39 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_28); +return x_39; } } } -else +} +LEAN_EXPORT lean_object* l_Lean_mkConstWithLevelParams___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: { -lean_object* x_125; -lean_free_object(x_49); -lean_dec(x_22); -lean_inc(x_3); -lean_ctor_set(x_8, 0, x_3); -x_125 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_125, 0, x_8); -x_23 = x_125; -x_24 = x_52; -goto block_31; +lean_object* x_13; +lean_inc(x_1); +x_13 = l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___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); +if (lean_obj_tag(x_13) == 0) +{ +uint8_t x_14; +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_15 = lean_ctor_get(x_13, 0); +x_16 = l_Lean_ConstantInfo_levelParams(x_15); +lean_dec(x_15); +x_17 = lean_box(0); +x_18 = l_List_mapTR_loop___at_Lean_mkConstWithLevelParams___spec__1(x_16, x_17); +x_19 = l_Lean_Expr_const___override(x_1, x_18); +lean_ctor_set(x_13, 0, x_19); +return x_13; } +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_20 = lean_ctor_get(x_13, 0); +x_21 = lean_ctor_get(x_13, 1); +lean_inc(x_21); +lean_inc(x_20); +lean_dec(x_13); +x_22 = l_Lean_ConstantInfo_levelParams(x_20); +lean_dec(x_20); +x_23 = lean_box(0); +x_24 = l_List_mapTR_loop___at_Lean_mkConstWithLevelParams___spec__1(x_22, x_23); +x_25 = l_Lean_Expr_const___override(x_1, x_24); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set(x_26, 1, x_21); +return x_26; } } else { -lean_object* x_126; lean_object* x_127; uint8_t x_128; -x_126 = lean_ctor_get(x_49, 0); -x_127 = lean_ctor_get(x_49, 1); -lean_inc(x_127); -lean_inc(x_126); -lean_dec(x_49); -x_128 = l_Lean_BinderInfo_isInstImplicit(x_45); -if (x_128 == 0) +uint8_t x_27; +lean_dec(x_1); +x_27 = !lean_is_exclusive(x_13); +if (x_27 == 0) { -lean_object* x_129; -lean_dec(x_126); -lean_dec(x_22); -lean_inc(x_3); -lean_ctor_set(x_8, 0, x_3); -x_129 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_129, 0, x_8); -x_23 = x_129; -x_24 = x_127; -goto block_31; +return x_13; } else { -uint8_t x_130; -x_130 = lean_unbox(x_126); -lean_dec(x_126); -if (x_130 == 0) -{ -lean_object* x_131; -lean_inc(x_18); -lean_inc(x_17); -lean_inc(x_16); -lean_inc(x_15); -lean_inc(x_22); -x_131 = lean_infer_type(x_22, x_15, x_16, x_17, x_18, x_127); -if (lean_obj_tag(x_131) == 0) +lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_28 = lean_ctor_get(x_13, 0); +x_29 = lean_ctor_get(x_13, 1); +lean_inc(x_29); +lean_inc(x_28); +lean_dec(x_13); +x_30 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_30, 0, x_28); +lean_ctor_set(x_30, 1, x_29); +return x_30; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: { -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_inc(x_18); -lean_inc(x_17); -lean_inc(x_16); -lean_inc(x_15); -lean_inc(x_132); -x_134 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem_synthesizeInstance(x_22, x_132, x_15, x_16, x_17, x_18, x_133); -if (lean_obj_tag(x_134) == 0) +switch (lean_obj_tag(x_1)) { +case 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_13; lean_object* x_14; +x_13 = lean_ctor_get(x_1, 0); +lean_inc(x_13); +lean_dec(x_1); +x_14 = l_Lean_mkConstWithLevelParams___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__4(x_13, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_14) == 0) { -lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; uint8_t x_141; -lean_free_object(x_8); -x_137 = lean_ctor_get(x_134, 1); -lean_inc(x_137); -lean_dec(x_134); -x_138 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__2; -x_139 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__2(x_138, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_137); -x_140 = lean_ctor_get(x_139, 0); -lean_inc(x_140); -x_141 = lean_unbox(x_140); -lean_dec(x_140); -if (x_141 == 0) +uint8_t x_15; +x_15 = !lean_is_exclusive(x_14); +if (x_15 == 0) { -lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; -lean_dec(x_132); -x_142 = lean_ctor_get(x_139, 1); -lean_inc(x_142); -lean_dec(x_139); -x_143 = lean_box(0); -x_144 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___lambda__1(x_33, x_143, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_142); -x_145 = lean_ctor_get(x_144, 0); -lean_inc(x_145); -x_146 = lean_ctor_get(x_144, 1); -lean_inc(x_146); -lean_dec(x_144); -x_23 = x_145; -x_24 = x_146; -goto block_31; +lean_object* x_16; lean_object* x_17; +x_16 = lean_ctor_get(x_14, 0); +x_17 = l_Lean_MessageData_ofConst(x_16); +lean_ctor_set(x_14, 0, x_17); +return x_14; } else { -lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; -x_147 = lean_ctor_get(x_139, 1); -lean_inc(x_147); -if (lean_is_exclusive(x_139)) { - lean_ctor_release(x_139, 0); - lean_ctor_release(x_139, 1); - x_148 = x_139; -} else { - lean_dec_ref(x_139); - x_148 = lean_box(0); -} -x_149 = lean_ctor_get(x_1, 5); -lean_inc(x_149); -x_150 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_149, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_147); -if (lean_obj_tag(x_150) == 0) -{ -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; -x_151 = lean_ctor_get(x_150, 0); -lean_inc(x_151); -x_152 = lean_ctor_get(x_150, 1); -lean_inc(x_152); -lean_dec(x_150); -x_153 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___closed__2; -if (lean_is_scalar(x_148)) { - x_154 = lean_alloc_ctor(7, 2, 0); -} else { - x_154 = x_148; - lean_ctor_set_tag(x_154, 7); +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_18 = lean_ctor_get(x_14, 0); +x_19 = lean_ctor_get(x_14, 1); +lean_inc(x_19); +lean_inc(x_18); +lean_dec(x_14); +x_20 = l_Lean_MessageData_ofConst(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_ctor_set(x_154, 0, x_153); -lean_ctor_set(x_154, 1, x_151); -x_155 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; -x_156 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_156, 0, x_154); -lean_ctor_set(x_156, 1, x_155); -x_157 = l_Lean_indentExpr(x_132); -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 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_159, 0, x_158); -lean_ctor_set(x_159, 1, x_155); -x_160 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_138, x_159, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_152); -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_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___lambda__1(x_33, x_161, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_162); -lean_dec(x_161); -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_23 = x_164; -x_24 = x_165; -goto block_31; } else { -lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; -lean_dec(x_148); -lean_dec(x_132); -lean_dec(x_33); -lean_dec(x_18); -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_3); -lean_dec(x_1); -x_166 = lean_ctor_get(x_150, 0); -lean_inc(x_166); -x_167 = lean_ctor_get(x_150, 1); -lean_inc(x_167); -if (lean_is_exclusive(x_150)) { - lean_ctor_release(x_150, 0); - lean_ctor_release(x_150, 1); - x_168 = x_150; -} else { - lean_dec_ref(x_150); - x_168 = lean_box(0); -} -if (lean_is_scalar(x_168)) { - x_169 = lean_alloc_ctor(1, 2, 0); -} else { - x_169 = x_168; +uint8_t x_22; +x_22 = !lean_is_exclusive(x_14); +if (x_22 == 0) +{ +return x_14; } -lean_ctor_set(x_169, 0, x_166); -lean_ctor_set(x_169, 1, x_167); -return x_169; +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_14, 0); +x_24 = lean_ctor_get(x_14, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_14); +x_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 +case 1: { -lean_object* x_170; lean_object* x_171; -lean_dec(x_132); -x_170 = lean_ctor_get(x_134, 1); -lean_inc(x_170); -lean_dec(x_134); -lean_inc(x_3); -lean_ctor_set(x_8, 0, x_3); -x_171 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_171, 0, x_8); -x_23 = x_171; -x_24 = x_170; -goto block_31; +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_26 = lean_ctor_get(x_1, 0); +lean_inc(x_26); +lean_dec(x_1); +x_27 = l_Lean_Expr_fvar___override(x_26); +x_28 = l_Lean_MessageData_ofExpr(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_12); +return x_29; } +case 2: +{ +uint8_t x_30; +x_30 = !lean_is_exclusive(x_1); +if (x_30 == 0) +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_1, 1); +x_32 = lean_ctor_get(x_1, 0); +lean_dec(x_32); +x_33 = l_Lean_MessageData_ofSyntax(x_31); +lean_ctor_set_tag(x_1, 0); +lean_ctor_set(x_1, 1, x_12); +lean_ctor_set(x_1, 0, x_33); +return x_1; } else { -lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; -lean_dec(x_132); -lean_dec(x_33); -lean_free_object(x_8); -lean_dec(x_18); -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_3); +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_1, 1); +lean_inc(x_34); lean_dec(x_1); -x_172 = lean_ctor_get(x_134, 0); -lean_inc(x_172); -x_173 = lean_ctor_get(x_134, 1); -lean_inc(x_173); -if (lean_is_exclusive(x_134)) { - lean_ctor_release(x_134, 0); - lean_ctor_release(x_134, 1); - x_174 = x_134; -} else { - lean_dec_ref(x_134); - x_174 = lean_box(0); +x_35 = l_Lean_MessageData_ofSyntax(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_12); +return x_36; } -if (lean_is_scalar(x_174)) { - x_175 = lean_alloc_ctor(1, 2, 0); -} else { - x_175 = x_174; } -lean_ctor_set(x_175, 0, x_172); -lean_ctor_set(x_175, 1, x_173); -return x_175; -} -} -else +default: { -lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; -lean_dec(x_33); -lean_free_object(x_8); -lean_dec(x_22); -lean_dec(x_18); -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_3); +lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_37 = lean_ctor_get(x_1, 0); +lean_inc(x_37); lean_dec(x_1); -x_176 = lean_ctor_get(x_131, 0); -lean_inc(x_176); -x_177 = lean_ctor_get(x_131, 1); -lean_inc(x_177); -if (lean_is_exclusive(x_131)) { - lean_ctor_release(x_131, 0); - lean_ctor_release(x_131, 1); - x_178 = x_131; -} else { - lean_dec_ref(x_131); - x_178 = lean_box(0); +x_38 = l_Lean_MessageData_ofName(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_12); +return x_39; } -if (lean_is_scalar(x_178)) { - x_179 = lean_alloc_ctor(1, 2, 0); -} else { - x_179 = x_178; } -lean_ctor_set(x_179, 0, x_176); -lean_ctor_set(x_179, 1, x_177); -return x_179; } } -else +static double _init_l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__1() { +_start: { -lean_object* x_180; -lean_dec(x_22); -lean_inc(x_3); -lean_ctor_set(x_8, 0, x_3); -x_180 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_180, 0, x_8); -x_23 = x_180; -x_24 = x_127; -goto block_31; -} -} +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; } } -else +static lean_object* _init_l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2() { +_start: { -lean_object* x_181; uint8_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; lean_object* x_190; uint8_t x_191; -lean_dec(x_33); -x_181 = lean_array_fget(x_35, x_36); -x_182 = lean_unbox(x_181); -lean_dec(x_181); -x_183 = lean_unsigned_to_nat(1u); -x_184 = lean_nat_add(x_36, x_183); -lean_dec(x_36); -x_185 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_185, 0, x_35); -lean_ctor_set(x_185, 1, x_184); -lean_ctor_set(x_185, 2, x_37); -x_186 = l_Lean_Expr_mvarId_x21(x_22); -x_187 = l_Lean_MVarId_isAssigned___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__4(x_186, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19); -lean_dec(x_186); -x_188 = lean_ctor_get(x_187, 0); -lean_inc(x_188); -x_189 = lean_ctor_get(x_187, 1); -lean_inc(x_189); -if (lean_is_exclusive(x_187)) { - lean_ctor_release(x_187, 0); - lean_ctor_release(x_187, 1); - x_190 = x_187; -} else { - lean_dec_ref(x_187); - x_190 = lean_box(0); +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(0); +x_2 = lean_array_mk(x_1); +return x_2; } -x_191 = l_Lean_BinderInfo_isInstImplicit(x_182); -if (x_191 == 0) -{ -lean_object* x_192; -lean_dec(x_190); -lean_dec(x_188); -lean_dec(x_22); -lean_inc(x_3); -lean_ctor_set(x_8, 1, x_185); -lean_ctor_set(x_8, 0, x_3); -x_192 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_192, 0, x_8); -x_23 = x_192; -x_24 = x_189; -goto block_31; } -else -{ -uint8_t x_193; -x_193 = lean_unbox(x_188); -lean_dec(x_188); -if (x_193 == 0) +LEAN_EXPORT lean_object* l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: { -lean_object* x_194; -lean_inc(x_18); -lean_inc(x_17); +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_14 = lean_ctor_get(x_11, 5); +x_15 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_2, x_9, x_10, x_11, x_12, x_13); +x_16 = lean_ctor_get(x_15, 0); lean_inc(x_16); -lean_inc(x_15); -lean_inc(x_22); -x_194 = lean_infer_type(x_22, x_15, x_16, x_17, x_18, x_189); -if (lean_obj_tag(x_194) == 0) -{ -lean_object* x_195; lean_object* x_196; lean_object* x_197; -x_195 = lean_ctor_get(x_194, 0); -lean_inc(x_195); -x_196 = lean_ctor_get(x_194, 1); -lean_inc(x_196); -lean_dec(x_194); -lean_inc(x_18); +x_17 = lean_ctor_get(x_15, 1); lean_inc(x_17); -lean_inc(x_16); -lean_inc(x_15); -lean_inc(x_195); -x_197 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem_synthesizeInstance(x_22, x_195, x_15, x_16, x_17, x_18, x_196); -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_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; uint8_t x_204; -lean_free_object(x_8); -x_200 = lean_ctor_get(x_197, 1); -lean_inc(x_200); -lean_dec(x_197); -x_201 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__2; -x_202 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__2(x_201, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_200); -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_dec(x_15); +x_18 = lean_st_ref_take(x_12, x_17); +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_19, 3); +lean_inc(x_20); +x_21 = !lean_is_exclusive(x_18); +if (x_21 == 0) { -lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; -lean_dec(x_195); -lean_dec(x_190); -x_205 = lean_ctor_get(x_202, 1); -lean_inc(x_205); -lean_dec(x_202); -x_206 = lean_box(0); -x_207 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___lambda__1(x_185, x_206, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_205); -x_208 = lean_ctor_get(x_207, 0); -lean_inc(x_208); -x_209 = lean_ctor_get(x_207, 1); -lean_inc(x_209); -lean_dec(x_207); -x_23 = x_208; -x_24 = x_209; -goto block_31; -} -else +lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_22 = lean_ctor_get(x_18, 1); +x_23 = lean_ctor_get(x_18, 0); +lean_dec(x_23); +x_24 = !lean_is_exclusive(x_19); +if (x_24 == 0) { -lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; -x_210 = lean_ctor_get(x_202, 1); -lean_inc(x_210); -if (lean_is_exclusive(x_202)) { - lean_ctor_release(x_202, 0); - lean_ctor_release(x_202, 1); - x_211 = x_202; -} else { - lean_dec_ref(x_202); - x_211 = lean_box(0); -} -x_212 = lean_ctor_get(x_1, 5); -lean_inc(x_212); -x_213 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_212, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_210); -if (lean_obj_tag(x_213) == 0) +lean_object* x_25; uint8_t x_26; +x_25 = lean_ctor_get(x_19, 3); +lean_dec(x_25); +x_26 = !lean_is_exclusive(x_20); +if (x_26 == 0) { -lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; -x_214 = lean_ctor_get(x_213, 0); -lean_inc(x_214); -x_215 = lean_ctor_get(x_213, 1); -lean_inc(x_215); -lean_dec(x_213); -x_216 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___closed__2; -if (lean_is_scalar(x_211)) { - x_217 = lean_alloc_ctor(7, 2, 0); -} else { - x_217 = x_211; - lean_ctor_set_tag(x_217, 7); -} -lean_ctor_set(x_217, 0, x_216); -lean_ctor_set(x_217, 1, x_214); -x_218 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; -if (lean_is_scalar(x_190)) { - x_219 = lean_alloc_ctor(7, 2, 0); -} else { - x_219 = x_190; - lean_ctor_set_tag(x_219, 7); -} -lean_ctor_set(x_219, 0, x_217); -lean_ctor_set(x_219, 1, x_218); -x_220 = l_Lean_indentExpr(x_195); -x_221 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_221, 0, x_219); -lean_ctor_set(x_221, 1, x_220); -x_222 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_222, 0, x_221); -lean_ctor_set(x_222, 1, x_218); -x_223 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_201, x_222, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_215); -x_224 = lean_ctor_get(x_223, 0); -lean_inc(x_224); -x_225 = lean_ctor_get(x_223, 1); -lean_inc(x_225); -lean_dec(x_223); -x_226 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___lambda__1(x_185, x_224, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_225); -lean_dec(x_224); -x_227 = lean_ctor_get(x_226, 0); -lean_inc(x_227); -x_228 = lean_ctor_get(x_226, 1); -lean_inc(x_228); -lean_dec(x_226); -x_23 = x_227; -x_24 = x_228; -goto block_31; -} -else +lean_object* x_27; double 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; uint8_t x_36; +x_27 = lean_ctor_get(x_20, 0); +x_28 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__1; +x_29 = 0; +x_30 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__1; +x_31 = lean_alloc_ctor(0, 2, 17); +lean_ctor_set(x_31, 0, x_1); +lean_ctor_set(x_31, 1, x_30); +lean_ctor_set_float(x_31, sizeof(void*)*2, x_28); +lean_ctor_set_float(x_31, sizeof(void*)*2 + 8, x_28); +lean_ctor_set_uint8(x_31, sizeof(void*)*2 + 16, x_29); +x_32 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2; +x_33 = lean_alloc_ctor(9, 3, 0); +lean_ctor_set(x_33, 0, x_31); +lean_ctor_set(x_33, 1, x_16); +lean_ctor_set(x_33, 2, x_32); +lean_inc(x_14); +lean_ctor_set(x_18, 1, x_33); +lean_ctor_set(x_18, 0, x_14); +x_34 = l_Lean_PersistentArray_push___rarg(x_27, x_18); +lean_ctor_set(x_20, 0, x_34); +x_35 = lean_st_ref_set(x_12, x_19, x_22); +x_36 = !lean_is_exclusive(x_35); +if (x_36 == 0) { -lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; -lean_dec(x_211); -lean_dec(x_195); -lean_dec(x_190); -lean_dec(x_185); -lean_dec(x_18); -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_3); -lean_dec(x_1); -x_229 = lean_ctor_get(x_213, 0); -lean_inc(x_229); -x_230 = lean_ctor_get(x_213, 1); -lean_inc(x_230); -if (lean_is_exclusive(x_213)) { - lean_ctor_release(x_213, 0); - lean_ctor_release(x_213, 1); - x_231 = x_213; -} else { - lean_dec_ref(x_213); - x_231 = lean_box(0); -} -if (lean_is_scalar(x_231)) { - x_232 = lean_alloc_ctor(1, 2, 0); -} else { - x_232 = x_231; -} -lean_ctor_set(x_232, 0, x_229); -lean_ctor_set(x_232, 1, x_230); -return x_232; -} -} +lean_object* x_37; lean_object* x_38; +x_37 = lean_ctor_get(x_35, 0); +lean_dec(x_37); +x_38 = lean_box(0); +lean_ctor_set(x_35, 0, x_38); +return x_35; } else { -lean_object* x_233; lean_object* x_234; -lean_dec(x_195); -lean_dec(x_190); -x_233 = lean_ctor_get(x_197, 1); -lean_inc(x_233); -lean_dec(x_197); -lean_inc(x_3); -lean_ctor_set(x_8, 1, x_185); -lean_ctor_set(x_8, 0, x_3); -x_234 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_234, 0, x_8); -x_23 = x_234; -x_24 = x_233; -goto block_31; +lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_39 = lean_ctor_get(x_35, 1); +lean_inc(x_39); +lean_dec(x_35); +x_40 = lean_box(0); +x_41 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_41, 0, x_40); +lean_ctor_set(x_41, 1, x_39); +return x_41; } } else { -lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; -lean_dec(x_195); -lean_dec(x_190); -lean_dec(x_185); -lean_free_object(x_8); -lean_dec(x_18); -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_3); -lean_dec(x_1); -x_235 = lean_ctor_get(x_197, 0); -lean_inc(x_235); -x_236 = lean_ctor_get(x_197, 1); -lean_inc(x_236); -if (lean_is_exclusive(x_197)) { - lean_ctor_release(x_197, 0); - lean_ctor_release(x_197, 1); - x_237 = x_197; +uint64_t x_42; lean_object* x_43; double 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; +x_42 = lean_ctor_get_uint64(x_20, sizeof(void*)*1); +x_43 = lean_ctor_get(x_20, 0); +lean_inc(x_43); +lean_dec(x_20); +x_44 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__1; +x_45 = 0; +x_46 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__1; +x_47 = lean_alloc_ctor(0, 2, 17); +lean_ctor_set(x_47, 0, x_1); +lean_ctor_set(x_47, 1, x_46); +lean_ctor_set_float(x_47, sizeof(void*)*2, x_44); +lean_ctor_set_float(x_47, sizeof(void*)*2 + 8, x_44); +lean_ctor_set_uint8(x_47, sizeof(void*)*2 + 16, x_45); +x_48 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2; +x_49 = lean_alloc_ctor(9, 3, 0); +lean_ctor_set(x_49, 0, x_47); +lean_ctor_set(x_49, 1, x_16); +lean_ctor_set(x_49, 2, x_48); +lean_inc(x_14); +lean_ctor_set(x_18, 1, x_49); +lean_ctor_set(x_18, 0, x_14); +x_50 = l_Lean_PersistentArray_push___rarg(x_43, x_18); +x_51 = lean_alloc_ctor(0, 1, 8); +lean_ctor_set(x_51, 0, x_50); +lean_ctor_set_uint64(x_51, sizeof(void*)*1, x_42); +lean_ctor_set(x_19, 3, x_51); +x_52 = lean_st_ref_set(x_12, x_19, x_22); +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_197); - x_237 = lean_box(0); + lean_dec_ref(x_52); + x_54 = lean_box(0); } -if (lean_is_scalar(x_237)) { - x_238 = lean_alloc_ctor(1, 2, 0); +x_55 = lean_box(0); +if (lean_is_scalar(x_54)) { + x_56 = lean_alloc_ctor(0, 2, 0); } else { - x_238 = x_237; + x_56 = x_54; } -lean_ctor_set(x_238, 0, x_235); -lean_ctor_set(x_238, 1, x_236); -return x_238; +lean_ctor_set(x_56, 0, x_55); +lean_ctor_set(x_56, 1, x_53); +return x_56; } } else { -lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; -lean_dec(x_190); -lean_dec(x_185); -lean_free_object(x_8); -lean_dec(x_22); -lean_dec(x_18); -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_3); -lean_dec(x_1); -x_239 = lean_ctor_get(x_194, 0); -lean_inc(x_239); -x_240 = lean_ctor_get(x_194, 1); -lean_inc(x_240); -if (lean_is_exclusive(x_194)) { - lean_ctor_release(x_194, 0); - lean_ctor_release(x_194, 1); - x_241 = x_194; +lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; uint64_t x_64; lean_object* x_65; lean_object* x_66; double x_67; uint8_t x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; +x_57 = lean_ctor_get(x_19, 0); +x_58 = lean_ctor_get(x_19, 1); +x_59 = lean_ctor_get(x_19, 2); +x_60 = lean_ctor_get(x_19, 4); +x_61 = lean_ctor_get(x_19, 5); +x_62 = lean_ctor_get(x_19, 6); +x_63 = lean_ctor_get(x_19, 7); +lean_inc(x_63); +lean_inc(x_62); +lean_inc(x_61); +lean_inc(x_60); +lean_inc(x_59); +lean_inc(x_58); +lean_inc(x_57); +lean_dec(x_19); +x_64 = lean_ctor_get_uint64(x_20, sizeof(void*)*1); +x_65 = lean_ctor_get(x_20, 0); +lean_inc(x_65); +if (lean_is_exclusive(x_20)) { + lean_ctor_release(x_20, 0); + x_66 = x_20; } else { - lean_dec_ref(x_194); - x_241 = lean_box(0); + lean_dec_ref(x_20); + x_66 = lean_box(0); } -if (lean_is_scalar(x_241)) { - x_242 = lean_alloc_ctor(1, 2, 0); +x_67 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__1; +x_68 = 0; +x_69 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__1; +x_70 = lean_alloc_ctor(0, 2, 17); +lean_ctor_set(x_70, 0, x_1); +lean_ctor_set(x_70, 1, x_69); +lean_ctor_set_float(x_70, sizeof(void*)*2, x_67); +lean_ctor_set_float(x_70, sizeof(void*)*2 + 8, x_67); +lean_ctor_set_uint8(x_70, sizeof(void*)*2 + 16, x_68); +x_71 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2; +x_72 = lean_alloc_ctor(9, 3, 0); +lean_ctor_set(x_72, 0, x_70); +lean_ctor_set(x_72, 1, x_16); +lean_ctor_set(x_72, 2, x_71); +lean_inc(x_14); +lean_ctor_set(x_18, 1, x_72); +lean_ctor_set(x_18, 0, x_14); +x_73 = l_Lean_PersistentArray_push___rarg(x_65, x_18); +if (lean_is_scalar(x_66)) { + x_74 = lean_alloc_ctor(0, 1, 8); } else { - x_242 = x_241; -} -lean_ctor_set(x_242, 0, x_239); -lean_ctor_set(x_242, 1, x_240); -return x_242; -} -} -else -{ -lean_object* x_243; -lean_dec(x_190); -lean_dec(x_22); -lean_inc(x_3); -lean_ctor_set(x_8, 1, x_185); -lean_ctor_set(x_8, 0, x_3); -x_243 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_243, 0, x_8); -x_23 = x_243; -x_24 = x_189; -goto block_31; -} + x_74 = x_66; } +lean_ctor_set(x_74, 0, x_73); +lean_ctor_set_uint64(x_74, sizeof(void*)*1, x_64); +x_75 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_75, 0, x_57); +lean_ctor_set(x_75, 1, x_58); +lean_ctor_set(x_75, 2, x_59); +lean_ctor_set(x_75, 3, x_74); +lean_ctor_set(x_75, 4, x_60); +lean_ctor_set(x_75, 5, x_61); +lean_ctor_set(x_75, 6, x_62); +lean_ctor_set(x_75, 7, x_63); +x_76 = lean_st_ref_set(x_12, x_75, x_22); +x_77 = lean_ctor_get(x_76, 1); +lean_inc(x_77); +if (lean_is_exclusive(x_76)) { + lean_ctor_release(x_76, 0); + lean_ctor_release(x_76, 1); + x_78 = x_76; +} else { + lean_dec_ref(x_76); + x_78 = lean_box(0); } +x_79 = lean_box(0); +if (lean_is_scalar(x_78)) { + x_80 = lean_alloc_ctor(0, 2, 0); +} else { + x_80 = x_78; } +lean_ctor_set(x_80, 0, x_79); +lean_ctor_set(x_80, 1, x_77); +return x_80; } -else -{ -lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; uint8_t x_248; -x_244 = lean_ctor_get(x_8, 1); -lean_inc(x_244); -lean_dec(x_8); -x_245 = lean_ctor_get(x_244, 0); -lean_inc(x_245); -x_246 = lean_ctor_get(x_244, 1); -lean_inc(x_246); -x_247 = lean_ctor_get(x_244, 2); -lean_inc(x_247); -x_248 = lean_nat_dec_lt(x_246, x_247); -if (x_248 == 0) -{ -lean_object* x_249; lean_object* x_250; -lean_dec(x_247); -lean_dec(x_246); -lean_dec(x_245); -lean_dec(x_22); -lean_inc(x_3); -x_249 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_249, 0, x_3); -lean_ctor_set(x_249, 1, x_244); -x_250 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_250, 0, x_249); -x_23 = x_250; -x_24 = x_19; -goto block_31; } else { -lean_object* x_251; lean_object* x_252; uint8_t 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; uint8_t x_262; -if (lean_is_exclusive(x_244)) { - lean_ctor_release(x_244, 0); - lean_ctor_release(x_244, 1); - lean_ctor_release(x_244, 2); - x_251 = x_244; +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; uint64_t x_90; lean_object* x_91; lean_object* x_92; double x_93; uint8_t x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; +x_81 = lean_ctor_get(x_18, 1); +lean_inc(x_81); +lean_dec(x_18); +x_82 = lean_ctor_get(x_19, 0); +lean_inc(x_82); +x_83 = lean_ctor_get(x_19, 1); +lean_inc(x_83); +x_84 = lean_ctor_get(x_19, 2); +lean_inc(x_84); +x_85 = lean_ctor_get(x_19, 4); +lean_inc(x_85); +x_86 = lean_ctor_get(x_19, 5); +lean_inc(x_86); +x_87 = lean_ctor_get(x_19, 6); +lean_inc(x_87); +x_88 = lean_ctor_get(x_19, 7); +lean_inc(x_88); +if (lean_is_exclusive(x_19)) { + lean_ctor_release(x_19, 0); + lean_ctor_release(x_19, 1); + lean_ctor_release(x_19, 2); + lean_ctor_release(x_19, 3); + lean_ctor_release(x_19, 4); + lean_ctor_release(x_19, 5); + lean_ctor_release(x_19, 6); + lean_ctor_release(x_19, 7); + x_89 = x_19; } else { - lean_dec_ref(x_244); - x_251 = lean_box(0); + lean_dec_ref(x_19); + x_89 = lean_box(0); } -x_252 = lean_array_fget(x_245, x_246); -x_253 = lean_unbox(x_252); -lean_dec(x_252); -x_254 = lean_unsigned_to_nat(1u); -x_255 = lean_nat_add(x_246, x_254); -lean_dec(x_246); -if (lean_is_scalar(x_251)) { - x_256 = lean_alloc_ctor(0, 3, 0); +x_90 = lean_ctor_get_uint64(x_20, sizeof(void*)*1); +x_91 = lean_ctor_get(x_20, 0); +lean_inc(x_91); +if (lean_is_exclusive(x_20)) { + lean_ctor_release(x_20, 0); + x_92 = x_20; } else { - x_256 = x_251; + lean_dec_ref(x_20); + x_92 = lean_box(0); } -lean_ctor_set(x_256, 0, x_245); -lean_ctor_set(x_256, 1, x_255); -lean_ctor_set(x_256, 2, x_247); -x_257 = l_Lean_Expr_mvarId_x21(x_22); -x_258 = l_Lean_MVarId_isAssigned___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__4(x_257, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19); -lean_dec(x_257); -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; +x_93 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__1; +x_94 = 0; +x_95 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__1; +x_96 = lean_alloc_ctor(0, 2, 17); +lean_ctor_set(x_96, 0, x_1); +lean_ctor_set(x_96, 1, x_95); +lean_ctor_set_float(x_96, sizeof(void*)*2, x_93); +lean_ctor_set_float(x_96, sizeof(void*)*2 + 8, x_93); +lean_ctor_set_uint8(x_96, sizeof(void*)*2 + 16, x_94); +x_97 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2; +x_98 = lean_alloc_ctor(9, 3, 0); +lean_ctor_set(x_98, 0, x_96); +lean_ctor_set(x_98, 1, x_16); +lean_ctor_set(x_98, 2, x_97); +lean_inc(x_14); +x_99 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_99, 0, x_14); +lean_ctor_set(x_99, 1, x_98); +x_100 = l_Lean_PersistentArray_push___rarg(x_91, x_99); +if (lean_is_scalar(x_92)) { + x_101 = lean_alloc_ctor(0, 1, 8); } else { - lean_dec_ref(x_258); - x_261 = lean_box(0); -} -x_262 = l_Lean_BinderInfo_isInstImplicit(x_253); -if (x_262 == 0) -{ -lean_object* x_263; lean_object* x_264; -lean_dec(x_261); -lean_dec(x_259); -lean_dec(x_22); -lean_inc(x_3); -x_263 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_263, 0, x_3); -lean_ctor_set(x_263, 1, x_256); -x_264 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_264, 0, x_263); -x_23 = x_264; -x_24 = x_260; -goto block_31; -} -else -{ -uint8_t x_265; -x_265 = lean_unbox(x_259); -lean_dec(x_259); -if (x_265 == 0) -{ -lean_object* x_266; -lean_inc(x_18); -lean_inc(x_17); -lean_inc(x_16); -lean_inc(x_15); -lean_inc(x_22); -x_266 = lean_infer_type(x_22, x_15, x_16, x_17, x_18, x_260); -if (lean_obj_tag(x_266) == 0) -{ -lean_object* x_267; lean_object* x_268; lean_object* x_269; -x_267 = lean_ctor_get(x_266, 0); -lean_inc(x_267); -x_268 = lean_ctor_get(x_266, 1); -lean_inc(x_268); -lean_dec(x_266); -lean_inc(x_18); -lean_inc(x_17); -lean_inc(x_16); -lean_inc(x_15); -lean_inc(x_267); -x_269 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem_synthesizeInstance(x_22, x_267, x_15, x_16, x_17, x_18, x_268); -if (lean_obj_tag(x_269) == 0) -{ -lean_object* x_270; uint8_t x_271; -x_270 = lean_ctor_get(x_269, 0); -lean_inc(x_270); -x_271 = lean_unbox(x_270); -lean_dec(x_270); -if (x_271 == 0) -{ -lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; uint8_t x_276; -x_272 = lean_ctor_get(x_269, 1); -lean_inc(x_272); -lean_dec(x_269); -x_273 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__2; -x_274 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__2(x_273, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_272); -x_275 = lean_ctor_get(x_274, 0); -lean_inc(x_275); -x_276 = lean_unbox(x_275); -lean_dec(x_275); -if (x_276 == 0) -{ -lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; -lean_dec(x_267); -lean_dec(x_261); -x_277 = lean_ctor_get(x_274, 1); -lean_inc(x_277); -lean_dec(x_274); -x_278 = lean_box(0); -x_279 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___lambda__1(x_256, x_278, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_277); -x_280 = lean_ctor_get(x_279, 0); -lean_inc(x_280); -x_281 = lean_ctor_get(x_279, 1); -lean_inc(x_281); -lean_dec(x_279); -x_23 = x_280; -x_24 = x_281; -goto block_31; + x_101 = x_92; } -else -{ -lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; -x_282 = lean_ctor_get(x_274, 1); -lean_inc(x_282); -if (lean_is_exclusive(x_274)) { - lean_ctor_release(x_274, 0); - lean_ctor_release(x_274, 1); - x_283 = x_274; +lean_ctor_set(x_101, 0, x_100); +lean_ctor_set_uint64(x_101, sizeof(void*)*1, x_90); +if (lean_is_scalar(x_89)) { + x_102 = lean_alloc_ctor(0, 8, 0); } else { - lean_dec_ref(x_274); - x_283 = lean_box(0); + x_102 = x_89; } -x_284 = lean_ctor_get(x_1, 5); -lean_inc(x_284); -x_285 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_284, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_282); -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, 0); -lean_inc(x_286); -x_287 = lean_ctor_get(x_285, 1); -lean_inc(x_287); -lean_dec(x_285); -x_288 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___closed__2; -if (lean_is_scalar(x_283)) { - x_289 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_102, 0, x_82); +lean_ctor_set(x_102, 1, x_83); +lean_ctor_set(x_102, 2, x_84); +lean_ctor_set(x_102, 3, x_101); +lean_ctor_set(x_102, 4, x_85); +lean_ctor_set(x_102, 5, x_86); +lean_ctor_set(x_102, 6, x_87); +lean_ctor_set(x_102, 7, x_88); +x_103 = lean_st_ref_set(x_12, x_102, x_81); +x_104 = lean_ctor_get(x_103, 1); +lean_inc(x_104); +if (lean_is_exclusive(x_103)) { + lean_ctor_release(x_103, 0); + lean_ctor_release(x_103, 1); + x_105 = x_103; } else { - x_289 = x_283; - lean_ctor_set_tag(x_289, 7); + lean_dec_ref(x_103); + x_105 = lean_box(0); } -lean_ctor_set(x_289, 0, x_288); -lean_ctor_set(x_289, 1, x_286); -x_290 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; -if (lean_is_scalar(x_261)) { - x_291 = lean_alloc_ctor(7, 2, 0); +x_106 = lean_box(0); +if (lean_is_scalar(x_105)) { + x_107 = lean_alloc_ctor(0, 2, 0); } else { - x_291 = x_261; - lean_ctor_set_tag(x_291, 7); + x_107 = x_105; } -lean_ctor_set(x_291, 0, x_289); -lean_ctor_set(x_291, 1, x_290); -x_292 = l_Lean_indentExpr(x_267); -x_293 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_293, 0, x_291); -lean_ctor_set(x_293, 1, x_292); -x_294 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_294, 0, x_293); -lean_ctor_set(x_294, 1, x_290); -x_295 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_273, x_294, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_287); -x_296 = lean_ctor_get(x_295, 0); -lean_inc(x_296); -x_297 = lean_ctor_get(x_295, 1); -lean_inc(x_297); -lean_dec(x_295); -x_298 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___lambda__1(x_256, x_296, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_297); -lean_dec(x_296); -x_299 = lean_ctor_get(x_298, 0); -lean_inc(x_299); -x_300 = lean_ctor_get(x_298, 1); -lean_inc(x_300); -lean_dec(x_298); -x_23 = x_299; -x_24 = x_300; -goto block_31; -} -else -{ -lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; -lean_dec(x_283); -lean_dec(x_267); -lean_dec(x_261); -lean_dec(x_256); -lean_dec(x_18); -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_3); -lean_dec(x_1); -x_301 = lean_ctor_get(x_285, 0); -lean_inc(x_301); -x_302 = lean_ctor_get(x_285, 1); -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; -} else { - lean_dec_ref(x_285); - x_303 = lean_box(0); -} -if (lean_is_scalar(x_303)) { - x_304 = lean_alloc_ctor(1, 2, 0); -} else { - x_304 = x_303; -} -lean_ctor_set(x_304, 0, x_301); -lean_ctor_set(x_304, 1, x_302); -return x_304; +lean_ctor_set(x_107, 0, x_106); +lean_ctor_set(x_107, 1, x_104); +return x_107; } } } -else +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___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_305; lean_object* x_306; lean_object* x_307; -lean_dec(x_267); -lean_dec(x_261); -x_305 = lean_ctor_get(x_269, 1); -lean_inc(x_305); -lean_dec(x_269); -lean_inc(x_3); -x_306 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_306, 0, x_3); -lean_ctor_set(x_306, 1, x_256); -x_307 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_307, 0, x_306); -x_23 = x_307; -x_24 = x_305; -goto block_31; +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_unsigned_to_nat(1u); +x_17 = lean_nat_add(x_1, x_16); +x_18 = l_Lean_Meta_Grind_addTheoremInstance(x_2, x_3, x_17, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +return x_18; } } -else +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__1() { +_start: { -lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; -lean_dec(x_267); -lean_dec(x_261); -lean_dec(x_256); -lean_dec(x_18); -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_3); -lean_dec(x_1); -x_308 = lean_ctor_get(x_269, 0); -lean_inc(x_308); -x_309 = lean_ctor_get(x_269, 1); -lean_inc(x_309); -if (lean_is_exclusive(x_269)) { - lean_ctor_release(x_269, 0); - lean_ctor_release(x_269, 1); - x_310 = x_269; -} else { - lean_dec_ref(x_269); - x_310 = lean_box(0); +lean_object* x_1; +x_1 = lean_mk_string_unchecked("grind", 5, 5); +return x_1; } -if (lean_is_scalar(x_310)) { - x_311 = lean_alloc_ctor(1, 2, 0); -} else { - x_311 = x_310; } -lean_ctor_set(x_311, 0, x_308); -lean_ctor_set(x_311, 1, x_309); -return x_311; +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("ematch", 6, 6); +return x_1; } } -else +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__3() { +_start: { -lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; -lean_dec(x_261); -lean_dec(x_256); -lean_dec(x_22); -lean_dec(x_18); -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_3); -lean_dec(x_1); -x_312 = lean_ctor_get(x_266, 0); -lean_inc(x_312); -x_313 = lean_ctor_get(x_266, 1); -lean_inc(x_313); -if (lean_is_exclusive(x_266)) { - lean_ctor_release(x_266, 0); - lean_ctor_release(x_266, 1); - x_314 = x_266; -} else { - lean_dec_ref(x_266); - x_314 = lean_box(0); +lean_object* x_1; +x_1 = lean_mk_string_unchecked("instance", 8, 8); +return x_1; } -if (lean_is_scalar(x_314)) { - x_315 = lean_alloc_ctor(1, 2, 0); -} else { - x_315 = x_314; } -lean_ctor_set(x_315, 0, x_312); -lean_ctor_set(x_315, 1, x_313); -return x_315; +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__2; +x_3 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__3; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; } } -else +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__5() { +_start: { -lean_object* x_316; lean_object* x_317; -lean_dec(x_261); -lean_dec(x_22); -lean_inc(x_3); -x_316 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_316, 0, x_3); -lean_ctor_set(x_316, 1, x_256); -x_317 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_317, 0, x_316); -x_23 = x_317; -x_24 = x_260; -goto block_31; +lean_object* x_1; +x_1 = lean_mk_string_unchecked(": ", 2, 2); +return x_1; } } +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__5; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } } -block_31: +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +_start: { -if (lean_obj_tag(x_23) == 0) +lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_17 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__4; +x_18 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__2(x_17, x_6, x_7, x_8, x_9, 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_25; lean_object* x_26; -lean_dec(x_18); -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_15); +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_dec(x_3); -lean_dec(x_1); -x_25 = lean_ctor_get(x_23, 0); -lean_inc(x_25); -lean_dec(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_24); -return x_26; +x_21 = lean_ctor_get(x_18, 1); +lean_inc(x_21); +lean_dec(x_18); +x_22 = lean_box(0); +x_23 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__1(x_1, x_2, x_4, x_22, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_21); +return x_23; } else { -lean_object* x_27; size_t x_28; size_t x_29; -x_27 = lean_ctor_get(x_23, 0); -lean_inc(x_27); -lean_dec(x_23); -x_28 = 1; -x_29 = lean_usize_add(x_7, x_28); -x_7 = x_29; -x_8 = x_27; -x_19 = x_24; -goto _start; +uint8_t x_24; +x_24 = !lean_is_exclusive(x_18); +if (x_24 == 0) +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_25 = lean_ctor_get(x_18, 1); +x_26 = lean_ctor_get(x_18, 0); +lean_dec(x_26); +x_27 = l_Lean_Meta_Grind_updateLastTag(x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_25); +if (lean_obj_tag(x_27) == 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_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_3, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, 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; lean_object* x_39; lean_object* x_40; lean_object* x_41; +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_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; +lean_ctor_set_tag(x_18, 7); +lean_ctor_set(x_18, 1, x_30); +lean_ctor_set(x_18, 0, x_32); +x_33 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__6; +x_34 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_34, 0, x_18); +lean_ctor_set(x_34, 1, x_33); +lean_inc(x_4); +x_35 = l_Lean_MessageData_ofExpr(x_4); +x_36 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_36, 0, x_34); +lean_ctor_set(x_36, 1, x_35); +x_37 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_37, 0, x_36); +lean_ctor_set(x_37, 1, x_32); +x_38 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_17, x_37, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_31); +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___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__1(x_1, x_2, x_4, x_39, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_40); +lean_dec(x_39); +return x_41; } +else +{ +uint8_t x_42; +lean_free_object(x_18); +lean_dec(x_4); +lean_dec(x_2); +x_42 = !lean_is_exclusive(x_29); +if (x_42 == 0) +{ +return x_29; } +else +{ +lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_43 = lean_ctor_get(x_29, 0); +x_44 = lean_ctor_get(x_29, 1); +lean_inc(x_44); +lean_inc(x_43); +lean_dec(x_29); +x_45 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +return x_45; } } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, size_t x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18) { -_start: -{ -uint8_t x_19; -x_19 = lean_usize_dec_lt(x_6, x_5); -if (x_19 == 0) +else { -lean_object* x_20; -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); +uint8_t x_46; +lean_free_object(x_18); +lean_dec(x_4); lean_dec(x_3); -x_20 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_20, 0, x_7); -lean_ctor_set(x_20, 1, x_18); -return x_20; +lean_dec(x_2); +x_46 = !lean_is_exclusive(x_27); +if (x_46 == 0) +{ +return x_27; } else { -lean_object* x_21; lean_object* x_22; -lean_dec(x_7); -x_21 = lean_array_uget(x_4, x_6); -lean_inc(x_17); -lean_inc(x_16); -lean_inc(x_15); -lean_inc(x_14); -lean_inc(x_21); -x_22 = l_Lean_Meta_isProof(x_21, x_14, x_15, x_16, x_17, x_18); -if (lean_obj_tag(x_22) == 0) +lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_47 = lean_ctor_get(x_27, 0); +x_48 = lean_ctor_get(x_27, 1); +lean_inc(x_48); +lean_inc(x_47); +lean_dec(x_27); +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_23; uint8_t x_24; -x_23 = lean_ctor_get(x_22, 0); -lean_inc(x_23); -x_24 = lean_unbox(x_23); -lean_dec(x_23); -if (x_24 == 0) +lean_object* x_50; lean_object* x_51; +x_50 = lean_ctor_get(x_18, 1); +lean_inc(x_50); +lean_dec(x_18); +x_51 = l_Lean_Meta_Grind_updateLastTag(x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_50); +if (lean_obj_tag(x_51) == 0) { -uint8_t x_25; -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_3); -x_25 = !lean_is_exclusive(x_22); -if (x_25 == 0) +lean_object* x_52; lean_object* x_53; +x_52 = lean_ctor_get(x_51, 1); +lean_inc(x_52); +lean_dec(x_51); +x_53 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_3, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_52); +if (lean_obj_tag(x_53) == 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_22, 0); -lean_dec(x_26); -x_27 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_27, 0, x_21); -x_28 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_28, 0, x_27); -x_29 = lean_box(0); -x_30 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_30, 0, x_28); -lean_ctor_set(x_30, 1, x_29); -lean_ctor_set(x_22, 0, x_30); -return x_22; +lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_54 = lean_ctor_get(x_53, 0); +lean_inc(x_54); +x_55 = lean_ctor_get(x_53, 1); +lean_inc(x_55); +lean_dec(x_53); +x_56 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; +x_57 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_57, 0, x_56); +lean_ctor_set(x_57, 1, x_54); +x_58 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__6; +x_59 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_59, 0, x_57); +lean_ctor_set(x_59, 1, x_58); +lean_inc(x_4); +x_60 = l_Lean_MessageData_ofExpr(x_4); +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 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_62, 0, x_61); +lean_ctor_set(x_62, 1, x_56); +x_63 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_17, x_62, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_55); +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___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__1(x_1, x_2, x_4, x_64, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_65); +lean_dec(x_64); +return x_66; } else { -lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_31 = lean_ctor_get(x_22, 1); -lean_inc(x_31); -lean_dec(x_22); -x_32 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_32, 0, x_21); -x_33 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_33, 0, x_32); -x_34 = lean_box(0); -x_35 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -x_36 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_36, 0, x_35); -lean_ctor_set(x_36, 1, x_31); -return x_36; -} +lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; +lean_dec(x_4); +lean_dec(x_2); +x_67 = lean_ctor_get(x_53, 0); +lean_inc(x_67); +x_68 = lean_ctor_get(x_53, 1); +lean_inc(x_68); +if (lean_is_exclusive(x_53)) { + lean_ctor_release(x_53, 0); + lean_ctor_release(x_53, 1); + x_69 = x_53; +} else { + lean_dec_ref(x_53); + x_69 = lean_box(0); } -else -{ -lean_object* x_37; size_t x_38; size_t x_39; -lean_dec(x_21); -x_37 = lean_ctor_get(x_22, 1); -lean_inc(x_37); -lean_dec(x_22); -x_38 = 1; -x_39 = lean_usize_add(x_6, x_38); -lean_inc(x_3); -{ -size_t _tmp_5 = x_39; -lean_object* _tmp_6 = x_3; -lean_object* _tmp_17 = x_37; -x_6 = _tmp_5; -x_7 = _tmp_6; -x_18 = _tmp_17; +if (lean_is_scalar(x_69)) { + x_70 = lean_alloc_ctor(1, 2, 0); +} else { + x_70 = x_69; } -goto _start; +lean_ctor_set(x_70, 0, x_67); +lean_ctor_set(x_70, 1, x_68); +return x_70; } } else { -uint8_t x_41; -lean_dec(x_21); -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); +lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; +lean_dec(x_4); lean_dec(x_3); -x_41 = !lean_is_exclusive(x_22); -if (x_41 == 0) -{ -return x_22; +lean_dec(x_2); +x_71 = lean_ctor_get(x_51, 0); +lean_inc(x_71); +x_72 = lean_ctor_get(x_51, 1); +lean_inc(x_72); +if (lean_is_exclusive(x_51)) { + lean_ctor_release(x_51, 0); + lean_ctor_release(x_51, 1); + x_73 = x_51; +} else { + lean_dec_ref(x_51); + x_73 = lean_box(0); } -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; +if (lean_is_scalar(x_73)) { + x_74 = lean_alloc_ctor(1, 2, 0); +} else { + x_74 = x_73; } +lean_ctor_set(x_74, 0, x_71); +lean_ctor_set(x_74, 1, x_72); +return x_74; } } } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, size_t x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18) { +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___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) { _start: { -uint8_t x_19; -x_19 = lean_usize_dec_lt(x_6, x_5); -if (x_19 == 0) +lean_object* x_16; +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_1); +x_16 = lean_infer_type(x_1, x_11, x_12, x_13, x_14, x_15); +if (lean_obj_tag(x_16) == 0) { -lean_object* x_20; -lean_dec(x_17); +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +x_18 = lean_ctor_get(x_16, 1); +lean_inc(x_18); lean_dec(x_16); -lean_dec(x_15); +x_19 = lean_st_ref_get(x_14, x_18); +x_20 = lean_ctor_get(x_19, 0); +lean_inc(x_20); +x_21 = lean_ctor_get(x_19, 1); +lean_inc(x_21); +lean_dec(x_19); +x_22 = lean_ctor_get(x_20, 0); +lean_inc(x_22); +lean_dec(x_20); +x_23 = l_Lean_Meta_Grind_Origin_key(x_3); +x_24 = l_Lean_Meta_Match_isMatchEqnTheorem(x_22, x_23); +lean_dec(x_23); +lean_dec(x_22); +if (x_24 == 0) +{ +lean_object* x_25; lean_object* x_26; +x_25 = lean_box(0); +x_26 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2(x_2, x_1, x_3, x_17, x_25, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_21); lean_dec(x_14); -lean_dec(x_3); -x_20 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_20, 0, x_7); -lean_ctor_set(x_20, 1, x_18); -return x_20; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +return x_26; } else { -lean_object* x_21; lean_object* x_22; -lean_dec(x_7); -x_21 = lean_array_uget(x_4, x_6); -lean_inc(x_17); -lean_inc(x_16); -lean_inc(x_15); +lean_object* x_27; lean_object* x_28; +x_27 = lean_ctor_get(x_5, 1); +lean_inc(x_27); lean_inc(x_14); -lean_inc(x_21); -x_22 = l_Lean_Meta_isProof(x_21, x_14, x_15, x_16, x_17, x_18); -if (lean_obj_tag(x_22) == 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); +lean_inc(x_6); +lean_inc(x_5); +x_28 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType(x_17, x_27, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_21); +if (lean_obj_tag(x_28) == 0) { -lean_object* x_23; uint8_t x_24; -x_23 = lean_ctor_get(x_22, 0); -lean_inc(x_23); -x_24 = lean_unbox(x_23); -lean_dec(x_23); -if (x_24 == 0) +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_29 = lean_ctor_get(x_28, 0); +lean_inc(x_29); +x_30 = lean_ctor_get(x_28, 1); +lean_inc(x_30); +lean_dec(x_28); +x_31 = lean_box(0); +x_32 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2(x_2, x_1, x_3, x_29, x_31, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_30); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +return x_32; +} +else { -uint8_t x_25; -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_15); +uint8_t x_33; lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); lean_dec(x_3); -x_25 = !lean_is_exclusive(x_22); -if (x_25 == 0) +lean_dec(x_1); +x_33 = !lean_is_exclusive(x_28); +if (x_33 == 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_22, 0); -lean_dec(x_26); -x_27 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_27, 0, x_21); -x_28 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_28, 0, x_27); -x_29 = lean_box(0); -x_30 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_30, 0, x_28); -lean_ctor_set(x_30, 1, x_29); -lean_ctor_set(x_22, 0, x_30); -return x_22; +return x_28; } else { -lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_31 = lean_ctor_get(x_22, 1); -lean_inc(x_31); -lean_dec(x_22); -x_32 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_32, 0, x_21); -x_33 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_33, 0, x_32); -x_34 = lean_box(0); -x_35 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -x_36 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_36, 0, x_35); -lean_ctor_set(x_36, 1, x_31); -return x_36; -} -} -else -{ -lean_object* x_37; size_t x_38; size_t x_39; -lean_dec(x_21); -x_37 = lean_ctor_get(x_22, 1); -lean_inc(x_37); -lean_dec(x_22); -x_38 = 1; -x_39 = lean_usize_add(x_6, x_38); -lean_inc(x_3); -{ -size_t _tmp_5 = x_39; -lean_object* _tmp_6 = x_3; -lean_object* _tmp_17 = x_37; -x_6 = _tmp_5; -x_7 = _tmp_6; -x_18 = _tmp_17; +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_28, 0); +x_35 = lean_ctor_get(x_28, 1); +lean_inc(x_35); +lean_inc(x_34); +lean_dec(x_28); +x_36 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_36, 0, x_34); +lean_ctor_set(x_36, 1, x_35); +return x_36; +} } -goto _start; } } else { -uint8_t x_41; -lean_dec(x_21); -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_15); +uint8_t x_37; lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); lean_dec(x_3); -x_41 = !lean_is_exclusive(x_22); -if (x_41 == 0) +lean_dec(x_1); +x_37 = !lean_is_exclusive(x_16); +if (x_37 == 0) { -return x_22; +return x_16; } 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_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_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; } } } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__8(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___closed__1() { _start: { -uint8_t x_16; -x_16 = lean_usize_dec_eq(x_2, x_3); -if (x_16 == 0) +lean_object* x_1; +x_1 = l_Lean_Meta_Grind_grind_debug_proofs; +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, 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_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; -x_17 = lean_array_uget(x_1, x_2); -x_18 = l_Lean_Expr_mvarId_x21(x_17); -x_19 = l_Lean_MVarId_isAssigned___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__4(x_18, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_15 = l_Lean_instantiateMVars___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__1(x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_15, 1); +lean_inc(x_17); +lean_dec(x_15); +x_18 = lean_ctor_get(x_12, 2); +lean_inc(x_18); +x_19 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___closed__1; +x_20 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_18, x_19); lean_dec(x_18); -x_20 = lean_ctor_get(x_19, 0); -lean_inc(x_20); -x_21 = lean_unbox(x_20); -lean_dec(x_20); -if (x_21 == 0) +if (x_20 == 0) { -lean_object* x_22; lean_object* x_23; size_t x_24; size_t x_25; -x_22 = lean_ctor_get(x_19, 1); -lean_inc(x_22); -lean_dec(x_19); -x_23 = lean_array_push(x_4, x_17); -x_24 = 1; -x_25 = lean_usize_add(x_2, x_24); -x_2 = x_25; -x_4 = x_23; -x_15 = x_22; -goto _start; +lean_object* x_21; lean_object* x_22; +x_21 = lean_box(0); +x_22 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__3(x_16, x_3, x_1, x_21, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_17); +return x_22; } else { -lean_object* x_27; size_t x_28; size_t x_29; -lean_dec(x_17); -x_27 = lean_ctor_get(x_19, 1); -lean_inc(x_27); -lean_dec(x_19); -x_28 = 1; -x_29 = lean_usize_add(x_2, x_28); -x_2 = x_29; -x_15 = x_27; -goto _start; +lean_object* x_23; +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_16); +x_23 = l_Lean_Meta_check(x_16, x_10, x_11, x_12, x_13, x_17); +if (lean_obj_tag(x_23) == 0) +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +x_25 = lean_ctor_get(x_23, 1); +lean_inc(x_25); +lean_dec(x_23); +x_26 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__3(x_16, x_3, x_1, x_24, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_25); +lean_dec(x_24); +return x_26; } +else +{ +uint8_t x_27; +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_1); +x_27 = !lean_is_exclusive(x_23); +if (x_27 == 0) +{ +return x_23; } else { -lean_object* x_31; -x_31 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_31, 0, x_4); -lean_ctor_set(x_31, 1, x_15); -return x_31; +lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_28 = lean_ctor_get(x_23, 0); +x_29 = lean_ctor_get(x_23, 1); +lean_inc(x_29); +lean_inc(x_28); +lean_dec(x_23); +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_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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_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_instantiateMVars___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -uint8_t x_19; -x_19 = lean_usize_dec_lt(x_6, x_5); -if (x_19 == 0) -{ -lean_object* x_20; -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); +lean_object* x_13; +x_13 = l_Lean_instantiateMVars___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); lean_dec(x_3); -x_20 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_20, 0, x_7); -lean_ctor_set(x_20, 1, x_18); -return x_20; +lean_dec(x_2); +return x_13; } -else +} +LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: { -lean_object* x_21; lean_object* x_22; +lean_object* x_13; +x_13 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_7); -x_21 = lean_array_uget(x_4, x_6); -lean_inc(x_17); -lean_inc(x_16); -lean_inc(x_15); -lean_inc(x_14); -lean_inc(x_21); -x_22 = l_Lean_Meta_isProof(x_21, x_14, x_15, x_16, x_17, x_18); -if (lean_obj_tag(x_22) == 0) -{ -lean_object* x_23; uint8_t x_24; -x_23 = lean_ctor_get(x_22, 0); -lean_inc(x_23); -x_24 = lean_unbox(x_23); -lean_dec(x_23); -if (x_24 == 0) -{ -uint8_t x_25; -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); lean_dec(x_3); -x_25 = !lean_is_exclusive(x_22); -if (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_22, 0); -lean_dec(x_26); -x_27 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_27, 0, x_21); -x_28 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_28, 0, x_27); -x_29 = lean_box(0); -x_30 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_30, 0, x_28); -lean_ctor_set(x_30, 1, x_29); -lean_ctor_set(x_22, 0, x_30); -return x_22; -} -else -{ -lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_31 = lean_ctor_get(x_22, 1); -lean_inc(x_31); -lean_dec(x_22); -x_32 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_32, 0, x_21); -x_33 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_33, 0, x_32); -x_34 = lean_box(0); -x_35 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -x_36 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_36, 0, x_35); -lean_ctor_set(x_36, 1, x_31); -return x_36; -} -} -else -{ -lean_object* x_37; size_t x_38; size_t x_39; -lean_dec(x_21); -x_37 = lean_ctor_get(x_22, 1); -lean_inc(x_37); -lean_dec(x_22); -x_38 = 1; -x_39 = lean_usize_add(x_6, x_38); -lean_inc(x_3); -{ -size_t _tmp_5 = x_39; -lean_object* _tmp_6 = x_3; -lean_object* _tmp_17 = x_37; -x_6 = _tmp_5; -x_7 = _tmp_6; -x_18 = _tmp_17; -} -goto _start; +lean_dec(x_2); +return x_13; } } -else +LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___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) { +_start: { -uint8_t x_41; -lean_dec(x_21); -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); +lean_object* x_13; +x_13 = l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__6(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_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_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_dec(x_2); +return x_13; } } -LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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_object* x_12, lean_object* x_13, lean_object* x_14) { +LEAN_EXPORT lean_object* l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -uint8_t x_15; -x_15 = lean_usize_dec_eq(x_2, x_3); -if (x_15 == 0) -{ -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; -x_16 = lean_array_uget(x_1, x_2); -x_17 = l_Lean_Expr_mvarId_x21(x_16); -lean_dec(x_16); -x_18 = l_Lean_MVarId_isAssigned___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__4(x_17, 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_17); -x_19 = lean_ctor_get(x_18, 0); -lean_inc(x_19); -x_20 = lean_unbox(x_19); -lean_dec(x_19); -if (x_20 == 0) -{ -uint8_t x_21; -x_21 = !lean_is_exclusive(x_18); -if (x_21 == 0) -{ -lean_object* x_22; uint8_t x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_18, 0); -lean_dec(x_22); -x_23 = 1; -x_24 = lean_box(x_23); -lean_ctor_set(x_18, 0, x_24); -return x_18; -} -else -{ -lean_object* x_25; uint8_t x_26; lean_object* x_27; lean_object* x_28; -x_25 = lean_ctor_get(x_18, 1); -lean_inc(x_25); -lean_dec(x_18); -x_26 = 1; -x_27 = lean_box(x_26); -x_28 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_28, 0, x_27); -lean_ctor_set(x_28, 1, x_25); -return x_28; +lean_object* x_13; +x_13 = l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___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); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_13; } } -else +LEAN_EXPORT lean_object* l_Lean_mkConstWithLevelParams___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: { -lean_object* x_29; size_t x_30; size_t x_31; -x_29 = lean_ctor_get(x_18, 1); -lean_inc(x_29); -lean_dec(x_18); -x_30 = 1; -x_31 = lean_usize_add(x_2, x_30); -x_2 = x_31; -x_14 = x_29; -goto _start; +lean_object* x_13; +x_13 = l_Lean_mkConstWithLevelParams___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_13; } } -else +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: { -uint8_t x_33; lean_object* x_34; lean_object* x_35; -x_33 = 0; -x_34 = lean_box(x_33); -x_35 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_35, 0, x_34); -lean_ctor_set(x_35, 1, x_14); -return x_35; -} +lean_object* x_13; +x_13 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_13; } } -LEAN_EXPORT lean_object* l_ReaderT_bind___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__11___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { lean_object* x_14; -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -x_14 = lean_apply_11(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -if (lean_obj_tag(x_14) == 0) -{ -lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = lean_ctor_get(x_14, 0); -lean_inc(x_15); -x_16 = lean_ctor_get(x_14, 1); -lean_inc(x_16); -lean_dec(x_14); -x_17 = lean_apply_12(x_2, x_15, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_16); -return x_17; -} -else -{ -uint8_t x_18; +x_14 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); @@ -7263,2110 +6733,7184 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); -x_18 = !lean_is_exclusive(x_14); -if (x_18 == 0) -{ return x_14; } -else -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_19 = lean_ctor_get(x_14, 0); -x_20 = lean_ctor_get(x_14, 1); -lean_inc(x_20); -lean_inc(x_19); -lean_dec(x_14); -x_21 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_21, 0, x_19); -lean_ctor_set(x_21, 1, x_20); -return x_21; -} -} -} } -LEAN_EXPORT lean_object* l_ReaderT_bind___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__11(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { _start: { -lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l_ReaderT_bind___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__11___rarg), 13, 0); -return x_3; +lean_object* x_16; +x_16 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +return x_16; } } -LEAN_EXPORT lean_object* l_Lean_Meta_withNewMCtxDepth___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__12___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_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { _start: { -lean_object* x_14; lean_object* x_15; -x_14 = lean_apply_6(x_1, x_3, x_4, x_5, x_6, x_7, x_8); -x_15 = l___private_Lean_Meta_Basic_0__Lean_Meta_withNewMCtxDepthImp___rarg(x_2, x_14, x_9, x_10, x_11, x_12, x_13); -if (lean_obj_tag(x_15) == 0) -{ -uint8_t x_16; -x_16 = !lean_is_exclusive(x_15); -if (x_16 == 0) +lean_object* x_17; +x_17 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_1); +return x_17; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +_start: { -return x_15; +lean_object* x_16; +x_16 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +lean_dec(x_4); +lean_dec(x_2); +return x_16; } -else +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, 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_object* x_18; lean_object* x_19; -x_17 = lean_ctor_get(x_15, 0); -x_18 = lean_ctor_get(x_15, 1); -lean_inc(x_18); -lean_inc(x_17); -lean_dec(x_15); -x_19 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_19, 0, x_17); -lean_ctor_set(x_19, 1, x_18); -return x_19; +lean_object* x_15; +x_15 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance(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); +return x_15; } } -else +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem_synthesizeInstance(lean_object* x_1, 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_20; -x_20 = !lean_is_exclusive(x_15); -if (x_20 == 0) +lean_object* x_8; lean_object* x_9; +x_8 = lean_box(0); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_9 = l_Lean_Meta_trySynthInstance(x_2, x_8, x_3, x_4, x_5, x_6, x_7); +if (lean_obj_tag(x_9) == 0) { -return x_15; +lean_object* x_10; +x_10 = lean_ctor_get(x_9, 0); +lean_inc(x_10); +if (lean_obj_tag(x_10) == 1) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_11 = lean_ctor_get(x_9, 1); +lean_inc(x_11); +lean_dec(x_9); +x_12 = lean_ctor_get(x_10, 0); +lean_inc(x_12); +lean_dec(x_10); +x_13 = l_Lean_Meta_isExprDefEq(x_1, x_12, x_3, x_4, x_5, x_6, x_11); +return x_13; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_15, 0); -x_22 = lean_ctor_get(x_15, 1); -lean_inc(x_22); -lean_inc(x_21); +uint8_t x_14; +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_14 = !lean_is_exclusive(x_9); +if (x_14 == 0) +{ +lean_object* x_15; uint8_t x_16; lean_object* x_17; +x_15 = lean_ctor_get(x_9, 0); 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; -} -} -} +x_16 = 0; +x_17 = lean_box(x_16); +lean_ctor_set(x_9, 0, x_17); +return x_9; } -LEAN_EXPORT lean_object* l_Lean_Meta_withNewMCtxDepth___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__12(lean_object* x_1) { -_start: +else { -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_Meta_withNewMCtxDepth___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__12___rarg___boxed), 13, 0); -return x_2; +lean_object* x_18; uint8_t x_19; lean_object* x_20; lean_object* x_21; +x_18 = lean_ctor_get(x_9, 1); +lean_inc(x_18); +lean_dec(x_9); +x_19 = 0; +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_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { -_start: -{ -lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; uint8_t x_21; uint8_t x_22; lean_object* x_23; -x_17 = l_Lean_instantiateMVars___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__1(x_1, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); -x_18 = lean_ctor_get(x_17, 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 = 1; -x_22 = 0; -x_23 = l_Lean_Meta_mkLambdaFVars(x_2, x_18, x_20, x_21, x_20, x_22, x_12, x_13, x_14, x_15, x_19); -if (lean_obj_tag(x_23) == 0) -{ -lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -x_25 = lean_ctor_get(x_23, 1); -lean_inc(x_25); -lean_dec(x_23); -x_26 = lean_ctor_get(x_3, 5); -lean_inc(x_26); -lean_dec(x_3); -x_27 = lean_ctor_get(x_4, 1); -lean_inc(x_27); -lean_dec(x_4); -x_28 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance(x_26, x_24, x_27, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_25); -return x_28; } else { -uint8_t x_29; -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); +uint8_t x_22; +lean_dec(x_6); +lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_29 = !lean_is_exclusive(x_23); -if (x_29 == 0) +lean_dec(x_1); +x_22 = !lean_is_exclusive(x_9); +if (x_22 == 0) { -return x_23; +return x_9; } else { -lean_object* x_30; lean_object* x_31; lean_object* x_32; -x_30 = lean_ctor_get(x_23, 0); -x_31 = lean_ctor_get(x_23, 1); -lean_inc(x_31); -lean_inc(x_30); -lean_dec(x_23); -x_32 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_32, 0, x_30); -lean_ctor_set(x_32, 1, x_31); -return x_32; +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_9, 0); +x_24 = lean_ctor_get(x_9, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_9); +x_25 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_24); +return x_25; } } } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l_ReaderT_read___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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) { _start: { -lean_object* x_14; -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; +lean_object* x_12; +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_1); +lean_ctor_set(x_12, 1, x_11); +return x_12; } } -static size_t _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__1() { +static lean_object* _init_l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__2___closed__1() { _start: { -lean_object* x_1; size_t x_2; -x_1 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2; -x_2 = lean_array_size(x_1); +lean_object* x_1; lean_object* x_2; +x_1 = l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___spec__1___closed__4; +x_2 = l_ReaderT_instMonad___rarg(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__2() { +static lean_object* _init_l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__2___closed__2() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("failed to instantiate ", 22, 22); -return x_1; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__2___closed__1; +x_2 = l_instInhabitedPUnit; +x_3 = l_instInhabitedOfMonad___rarg(x_1, x_2); +return x_3; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__3() { +static lean_object* _init_l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__2___closed__3() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__2; -x_2 = l_Lean_stringToMessageData(x_1); +x_1 = l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__2___closed__2; +x_2 = lean_alloc_closure((void*)(l_instInhabitedReaderT___rarg___boxed), 2, 1); +lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__4() { +LEAN_EXPORT lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked(", failed to instantiate non propositional argument with type", 60, 60); -return x_1; +lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_13 = l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__2___closed__3; +x_14 = lean_panic_fn(x_13, x_1); +x_15 = lean_apply_11(x_14, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_15; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__5() { +static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___lambda__1___closed__1() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__4; -x_2 = l_Lean_stringToMessageData(x_1); +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___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, size_t x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19, lean_object* x_20, lean_object* x_21) { +static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___lambda__1___closed__2() { _start: { -lean_object* x_22; uint8_t x_23; uint8_t x_24; lean_object* x_25; -x_22 = lean_unsigned_to_nat(0u); -x_23 = lean_nat_dec_lt(x_22, x_1); -if (x_23 == 0) -{ -uint8_t x_292; -x_292 = 1; -x_24 = x_292; -x_25 = x_21; -goto block_291; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___lambda__1___closed__1; +x_2 = lean_box(0); +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; } -else -{ -size_t x_293; lean_object* x_294; lean_object* x_295; uint8_t x_296; -x_293 = lean_usize_of_nat(x_1); -x_294 = l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__10(x_3, x_8, x_293, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21); -x_295 = lean_ctor_get(x_294, 0); -lean_inc(x_295); -x_296 = lean_unbox(x_295); -lean_dec(x_295); -if (x_296 == 0) -{ -lean_object* x_297; uint8_t x_298; -x_297 = lean_ctor_get(x_294, 1); -lean_inc(x_297); -lean_dec(x_294); -x_298 = 1; -x_24 = x_298; -x_25 = x_297; -goto block_291; } -else +static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___lambda__1___closed__3() { +_start: { -lean_object* x_299; uint8_t x_300; -x_299 = lean_ctor_get(x_294, 1); -lean_inc(x_299); -lean_dec(x_294); -x_300 = 0; -x_24 = x_300; -x_25 = x_299; -goto block_291; +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___lambda__1___closed__2; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; } } -block_291: -{ -if (x_24 == 0) -{ -lean_object* x_26; -x_26 = l_Lean_mkAppN(x_2, x_3); -if (x_23 == 0) -{ -lean_object* x_27; lean_object* x_28; lean_object* x_101; size_t x_102; lean_object* x_103; -x_101 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2; -x_102 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__1; -lean_inc(x_20); -lean_inc(x_19); -lean_inc(x_18); -lean_inc(x_17); -lean_inc(x_7); -x_103 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__6(x_6, x_101, x_7, x_101, x_102, x_8, x_7, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_25); -if (lean_obj_tag(x_103) == 0) -{ -lean_object* x_104; lean_object* x_105; -x_104 = lean_ctor_get(x_103, 0); -lean_inc(x_104); -x_105 = lean_ctor_get(x_104, 0); -lean_inc(x_105); -lean_dec(x_104); -if (lean_obj_tag(x_105) == 0) +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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) { +_start: { -lean_object* x_106; -x_106 = lean_ctor_get(x_103, 1); -lean_inc(x_106); -lean_dec(x_103); -x_27 = x_9; -x_28 = x_106; -goto block_100; +lean_object* x_13; lean_object* x_14; +x_13 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___lambda__1___closed__3; +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 +} +static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__1() { +_start: { -lean_object* x_107; lean_object* x_108; -lean_dec(x_9); -x_107 = lean_ctor_get(x_103, 1); -lean_inc(x_107); -lean_dec(x_103); -x_108 = lean_ctor_get(x_105, 0); -lean_inc(x_108); -lean_dec(x_105); -x_27 = x_108; -x_28 = x_107; -goto block_100; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("issues", 6, 6); +return x_1; } } -else -{ -uint8_t x_109; -lean_dec(x_26); -lean_dec(x_20); -lean_dec(x_19); -lean_dec(x_18); -lean_dec(x_17); -lean_dec(x_9); -lean_dec(x_5); -lean_dec(x_4); -x_109 = !lean_is_exclusive(x_103); -if (x_109 == 0) +static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__2() { +_start: { -return x_103; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__1; +x_2 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__1; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; } -else +} +static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__3() { +_start: { -lean_object* x_110; lean_object* x_111; lean_object* x_112; -x_110 = lean_ctor_get(x_103, 0); -x_111 = lean_ctor_get(x_103, 1); -lean_inc(x_111); -lean_inc(x_110); -lean_dec(x_103); -x_112 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_112, 0, x_110); -lean_ctor_set(x_112, 1, x_111); -return x_112; +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___lambda__1___boxed), 12, 0); +return x_1; } } -block_100: +static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__4() { +_start: { -if (lean_obj_tag(x_27) == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("type error constructing proof for ", 34, 34); +return x_1; +} +} +static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__5() { +_start: { -lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_29 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2; -x_30 = lean_box(0); -x_31 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__1(x_26, x_29, x_4, x_5, x_30, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_28); -return x_31; +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__4; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } -else +} +static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__6() { +_start: { -lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; -x_32 = lean_ctor_get(x_27, 0); -lean_inc(x_32); -lean_dec(x_27); -x_33 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__2; -x_34 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__2(x_33, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_28); -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_1; +x_1 = lean_mk_string_unchecked("\nwhen assigning metavariable ", 29, 29); +return x_1; +} +} +static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__7() { +_start: { -lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; -lean_dec(x_32); -x_37 = lean_ctor_get(x_34, 1); -lean_inc(x_37); -lean_dec(x_34); -x_38 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2; -x_39 = lean_box(0); -x_40 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__1(x_26, x_38, x_4, x_5, x_39, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_37); -return x_40; +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__6; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } -else +} +static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__8() { +_start: { -uint8_t x_41; -x_41 = !lean_is_exclusive(x_34); -if (x_41 == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked(" with ", 6, 6); +return x_1; +} +} +static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__9() { +_start: { -lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; -x_42 = lean_ctor_get(x_34, 1); -x_43 = lean_ctor_get(x_34, 0); -lean_dec(x_43); -x_44 = lean_ctor_get(x_4, 5); -lean_inc(x_44); -x_45 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_44, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_42); -if (lean_obj_tag(x_45) == 0) +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__8; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__10() { +_start: { -lean_object* x_46; lean_object* x_47; lean_object* x_48; -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); -lean_inc(x_20); -lean_inc(x_19); -lean_inc(x_18); -lean_inc(x_17); -x_48 = lean_infer_type(x_32, x_17, x_18, x_19, x_20, x_47); -if (lean_obj_tag(x_48) == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("\n", 1, 1); +return x_1; +} +} +static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__11() { +_start: { -lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; 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_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___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__3; -lean_ctor_set_tag(x_34, 7); -lean_ctor_set(x_34, 1, x_46); -lean_ctor_set(x_34, 0, x_51); -x_52 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__5; -x_53 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_53, 0, x_34); -lean_ctor_set(x_53, 1, x_52); -x_54 = l_Lean_indentExpr(x_49); -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_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; -x_57 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_57, 0, x_55); -lean_ctor_set(x_57, 1, x_56); -x_58 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_33, x_57, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_50); -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_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2; -x_62 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__1(x_26, x_61, x_4, x_5, x_59, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_60); -lean_dec(x_59); -return x_62; +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__10; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } -else +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19, lean_object* x_20, lean_object* x_21, lean_object* x_22, lean_object* x_23) { +_start: { -uint8_t x_63; -lean_dec(x_46); -lean_free_object(x_34); -lean_dec(x_26); +lean_object* x_24; uint8_t x_25; +x_24 = lean_ctor_get(x_8, 1); +x_25 = lean_nat_dec_lt(x_10, x_24); +if (x_25 == 0) +{ +lean_object* x_26; +lean_dec(x_22); +lean_dec(x_21); lean_dec(x_20); lean_dec(x_19); lean_dec(x_18); lean_dec(x_17); -lean_dec(x_5); -lean_dec(x_4); -x_63 = !lean_is_exclusive(x_48); -if (x_63 == 0) -{ -return x_48; +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_10); +lean_dec(x_7); +lean_dec(x_1); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_9); +lean_ctor_set(x_26, 1, x_23); +return x_26; } else { -lean_object* x_64; lean_object* x_65; lean_object* x_66; -x_64 = lean_ctor_get(x_48, 0); -x_65 = lean_ctor_get(x_48, 1); -lean_inc(x_65); -lean_inc(x_64); -lean_dec(x_48); -x_66 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_66, 0, x_64); -lean_ctor_set(x_66, 1, x_65); -return x_66; -} -} +lean_object* x_27; lean_object* x_28; lean_object* x_36; lean_object* x_197; lean_object* x_198; lean_object* x_199; uint8_t x_200; +lean_dec(x_9); +x_197 = lean_nat_sub(x_3, x_10); +x_198 = lean_unsigned_to_nat(1u); +x_199 = lean_nat_sub(x_197, x_198); +lean_dec(x_197); +x_200 = lean_nat_dec_lt(x_199, x_4); +if (x_200 == 0) +{ +lean_object* x_201; lean_object* x_202; +lean_dec(x_199); +x_201 = l_Lean_instInhabitedExpr; +x_202 = l_outOfBounds___rarg(x_201); +x_36 = x_202; +goto block_196; } else { -uint8_t x_67; -lean_free_object(x_34); -lean_dec(x_32); -lean_dec(x_26); +lean_object* x_203; +x_203 = lean_array_fget(x_2, x_199); +lean_dec(x_199); +x_36 = x_203; +goto block_196; +} +block_35: +{ +if (lean_obj_tag(x_27) == 0) +{ +lean_object* x_29; lean_object* x_30; +lean_dec(x_22); +lean_dec(x_21); lean_dec(x_20); lean_dec(x_19); lean_dec(x_18); lean_dec(x_17); -lean_dec(x_5); -lean_dec(x_4); -x_67 = !lean_is_exclusive(x_45); -if (x_67 == 0) -{ -return x_45; +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_10); +lean_dec(x_7); +lean_dec(x_1); +x_29 = lean_ctor_get(x_27, 0); +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; } else { -lean_object* x_68; lean_object* x_69; lean_object* x_70; -x_68 = lean_ctor_get(x_45, 0); -x_69 = lean_ctor_get(x_45, 1); -lean_inc(x_69); -lean_inc(x_68); -lean_dec(x_45); -x_70 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_70, 0, x_68); -lean_ctor_set(x_70, 1, x_69); -return x_70; +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_27, 0); +lean_inc(x_31); +lean_dec(x_27); +x_32 = lean_ctor_get(x_8, 2); +x_33 = lean_nat_add(x_10, x_32); +lean_dec(x_10); +x_9 = x_31; +x_10 = x_33; +x_11 = lean_box(0); +x_12 = lean_box(0); +x_23 = x_28; +goto _start; } } +block_196: +{ +lean_object* x_37; uint8_t x_38; +x_37 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_unassigned; +x_38 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_36, x_37); +if (x_38 == 0) +{ +lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_39 = lean_array_fget(x_5, x_10); +x_40 = l_Lean_Expr_mvarId_x21(x_39); +lean_inc(x_40); +x_41 = l_Lean_MVarId_getType(x_40, x_19, x_20, x_21, x_22, x_23); +if (lean_obj_tag(x_41) == 0) +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_42 = lean_ctor_get(x_41, 0); +lean_inc(x_42); +x_43 = lean_ctor_get(x_41, 1); +lean_inc(x_43); +lean_dec(x_41); +lean_inc(x_22); +lean_inc(x_21); +lean_inc(x_20); +lean_inc(x_19); +lean_inc(x_36); +x_44 = lean_infer_type(x_36, x_19, x_20, x_21, x_22, x_43); +if (lean_obj_tag(x_44) == 0) +{ +lean_object* x_45; lean_object* x_46; uint8_t x_47; lean_object* x_48; lean_object* x_169; +x_45 = lean_ctor_get(x_44, 0); +lean_inc(x_45); +x_46 = lean_ctor_get(x_44, 1); +lean_inc(x_46); +lean_dec(x_44); +lean_inc(x_22); +lean_inc(x_21); +lean_inc(x_20); +lean_inc(x_19); +lean_inc(x_45); +lean_inc(x_42); +x_169 = l_Lean_Meta_isExprDefEq(x_42, x_45, x_19, x_20, x_21, x_22, x_46); +if (lean_obj_tag(x_169) == 0) +{ +lean_object* x_170; uint8_t x_171; +x_170 = lean_ctor_get(x_169, 0); +lean_inc(x_170); +x_171 = lean_unbox(x_170); +if (x_171 == 0) +{ +lean_object* x_172; uint8_t x_173; +lean_dec(x_40); +x_172 = lean_ctor_get(x_169, 1); +lean_inc(x_172); +lean_dec(x_169); +x_173 = lean_unbox(x_170); +lean_dec(x_170); +x_47 = x_173; +x_48 = x_172; +goto block_168; } else { -lean_object* x_71; lean_object* x_72; lean_object* x_73; -x_71 = lean_ctor_get(x_34, 1); -lean_inc(x_71); -lean_dec(x_34); -x_72 = lean_ctor_get(x_4, 5); -lean_inc(x_72); -x_73 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_72, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_71); -if (lean_obj_tag(x_73) == 0) -{ -lean_object* x_74; lean_object* x_75; lean_object* x_76; -x_74 = lean_ctor_get(x_73, 0); -lean_inc(x_74); -x_75 = lean_ctor_get(x_73, 1); -lean_inc(x_75); -lean_dec(x_73); +lean_object* x_174; lean_object* x_175; +lean_dec(x_170); +x_174 = lean_ctor_get(x_169, 1); +lean_inc(x_174); +lean_dec(x_169); +lean_inc(x_22); +lean_inc(x_21); lean_inc(x_20); lean_inc(x_19); -lean_inc(x_18); -lean_inc(x_17); -x_76 = lean_infer_type(x_32, x_17, x_18, x_19, x_20, x_75); -if (lean_obj_tag(x_76) == 0) +lean_inc(x_36); +x_175 = lean_checked_assign(x_40, x_36, x_19, x_20, x_21, x_22, x_174); +if (lean_obj_tag(x_175) == 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; -x_77 = lean_ctor_get(x_76, 0); -lean_inc(x_77); -x_78 = lean_ctor_get(x_76, 1); -lean_inc(x_78); -lean_dec(x_76); -x_79 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__3; -x_80 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_80, 0, x_79); -lean_ctor_set(x_80, 1, x_74); -x_81 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__5; -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_indentExpr(x_77); -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_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; -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_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_33, x_86, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_78); -x_88 = lean_ctor_get(x_87, 0); -lean_inc(x_88); -x_89 = lean_ctor_get(x_87, 1); -lean_inc(x_89); -lean_dec(x_87); -x_90 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2; -x_91 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__1(x_26, x_90, x_4, x_5, x_88, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_89); -lean_dec(x_88); -return x_91; +lean_object* x_176; lean_object* x_177; uint8_t x_178; +x_176 = lean_ctor_get(x_175, 0); +lean_inc(x_176); +x_177 = lean_ctor_get(x_175, 1); +lean_inc(x_177); +lean_dec(x_175); +x_178 = lean_unbox(x_176); +lean_dec(x_176); +x_47 = x_178; +x_48 = x_177; +goto block_168; } else { -lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; -lean_dec(x_74); -lean_dec(x_26); +uint8_t x_179; +lean_dec(x_45); +lean_dec(x_42); +lean_dec(x_39); +lean_dec(x_36); +lean_dec(x_22); +lean_dec(x_21); lean_dec(x_20); lean_dec(x_19); lean_dec(x_18); lean_dec(x_17); -lean_dec(x_5); -lean_dec(x_4); -x_92 = lean_ctor_get(x_76, 0); -lean_inc(x_92); -x_93 = lean_ctor_get(x_76, 1); -lean_inc(x_93); -if (lean_is_exclusive(x_76)) { - lean_ctor_release(x_76, 0); - lean_ctor_release(x_76, 1); - x_94 = x_76; -} else { - lean_dec_ref(x_76); - x_94 = lean_box(0); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_10); +lean_dec(x_7); +lean_dec(x_1); +x_179 = !lean_is_exclusive(x_175); +if (x_179 == 0) +{ +return x_175; +} +else +{ +lean_object* x_180; lean_object* x_181; lean_object* x_182; +x_180 = lean_ctor_get(x_175, 0); +x_181 = lean_ctor_get(x_175, 1); +lean_inc(x_181); +lean_inc(x_180); +lean_dec(x_175); +x_182 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_182, 0, x_180); +lean_ctor_set(x_182, 1, x_181); +return x_182; } -if (lean_is_scalar(x_94)) { - x_95 = lean_alloc_ctor(1, 2, 0); -} else { - x_95 = x_94; } -lean_ctor_set(x_95, 0, x_92); -lean_ctor_set(x_95, 1, x_93); -return x_95; } } else { -lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; -lean_dec(x_32); -lean_dec(x_26); +uint8_t x_183; +lean_dec(x_45); +lean_dec(x_42); +lean_dec(x_40); +lean_dec(x_39); +lean_dec(x_36); +lean_dec(x_22); +lean_dec(x_21); lean_dec(x_20); lean_dec(x_19); lean_dec(x_18); lean_dec(x_17); -lean_dec(x_5); -lean_dec(x_4); -x_96 = lean_ctor_get(x_73, 0); -lean_inc(x_96); -x_97 = lean_ctor_get(x_73, 1); -lean_inc(x_97); -if (lean_is_exclusive(x_73)) { - lean_ctor_release(x_73, 0); - lean_ctor_release(x_73, 1); - x_98 = x_73; -} else { - lean_dec_ref(x_73); - x_98 = lean_box(0); -} -if (lean_is_scalar(x_98)) { - x_99 = lean_alloc_ctor(1, 2, 0); -} else { - x_99 = x_98; -} -lean_ctor_set(x_99, 0, x_96); -lean_ctor_set(x_99, 1, x_97); -return x_99; -} -} -} +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_10); +lean_dec(x_7); +lean_dec(x_1); +x_183 = !lean_is_exclusive(x_169); +if (x_183 == 0) +{ +return x_169; } +else +{ +lean_object* x_184; lean_object* x_185; lean_object* x_186; +x_184 = lean_ctor_get(x_169, 0); +x_185 = lean_ctor_get(x_169, 1); +lean_inc(x_185); +lean_inc(x_184); +lean_dec(x_169); +x_186 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_186, 0, x_184); +lean_ctor_set(x_186, 1, x_185); +return x_186; } } -else +block_168: { -uint8_t x_113; -x_113 = lean_nat_dec_le(x_1, x_1); -if (x_113 == 0) +if (x_47 == 0) +{ +lean_object* x_49; lean_object* x_50; uint8_t x_51; +x_49 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__2; +x_50 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__2(x_49, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_48); +x_51 = !lean_is_exclusive(x_50); +if (x_51 == 0) +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; uint8_t x_55; +x_52 = lean_ctor_get(x_50, 0); +x_53 = lean_ctor_get(x_50, 1); +x_54 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__3; +x_55 = lean_unbox(x_52); +lean_dec(x_52); +if (x_55 == 0) { -lean_object* x_114; lean_object* x_115; lean_object* x_188; size_t x_189; lean_object* x_190; -x_188 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2; -x_189 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__1; +lean_object* x_56; lean_object* x_57; +lean_free_object(x_50); +lean_dec(x_45); +lean_dec(x_42); +lean_dec(x_39); +lean_dec(x_36); +x_56 = lean_box(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_7); -x_190 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__7(x_6, x_188, x_7, x_188, x_189, x_8, x_7, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_25); -if (lean_obj_tag(x_190) == 0) -{ -lean_object* x_191; lean_object* x_192; -x_191 = lean_ctor_get(x_190, 0); -lean_inc(x_191); -x_192 = lean_ctor_get(x_191, 0); -lean_inc(x_192); -lean_dec(x_191); -if (lean_obj_tag(x_192) == 0) -{ -lean_object* x_193; -x_193 = lean_ctor_get(x_190, 1); -lean_inc(x_193); -lean_dec(x_190); -x_114 = x_9; -x_115 = x_193; -goto block_187; -} -else +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +x_57 = lean_apply_12(x_54, x_56, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_53); +if (lean_obj_tag(x_57) == 0) { -lean_object* x_194; lean_object* x_195; -lean_dec(x_9); -x_194 = lean_ctor_get(x_190, 1); -lean_inc(x_194); -lean_dec(x_190); -x_195 = lean_ctor_get(x_192, 0); -lean_inc(x_195); -lean_dec(x_192); -x_114 = x_195; -x_115 = x_194; -goto block_187; -} +lean_object* x_58; lean_object* x_59; +x_58 = lean_ctor_get(x_57, 0); +lean_inc(x_58); +x_59 = lean_ctor_get(x_57, 1); +lean_inc(x_59); +lean_dec(x_57); +x_27 = x_58; +x_28 = x_59; +goto block_35; } else { -uint8_t x_196; -lean_dec(x_26); +uint8_t x_60; +lean_dec(x_22); +lean_dec(x_21); lean_dec(x_20); lean_dec(x_19); lean_dec(x_18); lean_dec(x_17); -lean_dec(x_9); -lean_dec(x_5); -lean_dec(x_4); -x_196 = !lean_is_exclusive(x_190); -if (x_196 == 0) +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_10); +lean_dec(x_7); +lean_dec(x_1); +x_60 = !lean_is_exclusive(x_57); +if (x_60 == 0) { -return x_190; +return x_57; } else { -lean_object* x_197; lean_object* x_198; lean_object* x_199; -x_197 = lean_ctor_get(x_190, 0); -x_198 = lean_ctor_get(x_190, 1); -lean_inc(x_198); -lean_inc(x_197); -lean_dec(x_190); -x_199 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_199, 0, x_197); -lean_ctor_set(x_199, 1, x_198); -return x_199; +lean_object* x_61; lean_object* x_62; lean_object* x_63; +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(1, 2, 0); +lean_ctor_set(x_63, 0, x_61); +lean_ctor_set(x_63, 1, x_62); +return x_63; } } -block_187: -{ -if (lean_obj_tag(x_114) == 0) -{ -lean_object* x_116; lean_object* x_117; lean_object* x_118; -x_116 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2; -x_117 = lean_box(0); -x_118 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__1(x_26, x_116, x_4, x_5, x_117, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_115); -return x_118; } else { -lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; uint8_t x_123; -x_119 = lean_ctor_get(x_114, 0); -lean_inc(x_119); -lean_dec(x_114); -x_120 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__2; -x_121 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__2(x_120, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_115); -x_122 = lean_ctor_get(x_121, 0); -lean_inc(x_122); -x_123 = lean_unbox(x_122); -lean_dec(x_122); -if (x_123 == 0) +lean_object* x_64; +x_64 = l_Lean_Meta_Grind_updateLastTag(x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_53); +if (lean_obj_tag(x_64) == 0) { -lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; -lean_dec(x_119); -x_124 = lean_ctor_get(x_121, 1); -lean_inc(x_124); -lean_dec(x_121); -x_125 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2; -x_126 = lean_box(0); -x_127 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__1(x_26, x_125, x_4, x_5, x_126, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_124); -return x_127; -} -else +lean_object* x_65; lean_object* x_66; lean_object* x_67; +x_65 = lean_ctor_get(x_64, 1); +lean_inc(x_65); +lean_dec(x_64); +x_66 = lean_ctor_get(x_1, 5); +lean_inc(x_66); +x_67 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_66, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_65); +if (lean_obj_tag(x_67) == 0) { -uint8_t x_128; -x_128 = !lean_is_exclusive(x_121); -if (x_128 == 0) +lean_object* x_68; lean_object* x_69; lean_object* x_70; +x_68 = lean_ctor_get(x_67, 0); +lean_inc(x_68); +x_69 = lean_ctor_get(x_67, 1); +lean_inc(x_69); +lean_dec(x_67); +lean_inc(x_22); +lean_inc(x_21); +lean_inc(x_20); +lean_inc(x_19); +x_70 = l_Lean_Meta_mkHasTypeButIsExpectedMsg(x_45, x_42, x_19, x_20, x_21, x_22, x_69); +if (lean_obj_tag(x_70) == 0) { -lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; -x_129 = lean_ctor_get(x_121, 1); -x_130 = lean_ctor_get(x_121, 0); -lean_dec(x_130); -x_131 = lean_ctor_get(x_4, 5); -lean_inc(x_131); -x_132 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_131, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_129); -if (lean_obj_tag(x_132) == 0) -{ -lean_object* x_133; lean_object* x_134; lean_object* x_135; -x_133 = lean_ctor_get(x_132, 0); -lean_inc(x_133); -x_134 = lean_ctor_get(x_132, 1); -lean_inc(x_134); -lean_dec(x_132); +lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; +x_71 = lean_ctor_get(x_70, 0); +lean_inc(x_71); +x_72 = lean_ctor_get(x_70, 1); +lean_inc(x_72); +lean_dec(x_70); +x_73 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__5; +lean_ctor_set_tag(x_50, 7); +lean_ctor_set(x_50, 1, x_68); +lean_ctor_set(x_50, 0, x_73); +x_74 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__7; +x_75 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_75, 0, x_50); +lean_ctor_set(x_75, 1, x_74); +x_76 = l_Lean_MessageData_ofExpr(x_39); +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_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__9; +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_indentExpr(x_36); +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_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__11; +x_83 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_83, 0, x_81); +lean_ctor_set(x_83, 1, x_82); +x_84 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_84, 0, x_83); +lean_ctor_set(x_84, 1, x_71); +x_85 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; +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_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_49, x_86, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_72); +x_88 = lean_ctor_get(x_87, 0); +lean_inc(x_88); +x_89 = lean_ctor_get(x_87, 1); +lean_inc(x_89); +lean_dec(x_87); +lean_inc(x_22); +lean_inc(x_21); lean_inc(x_20); lean_inc(x_19); lean_inc(x_18); lean_inc(x_17); -x_135 = lean_infer_type(x_119, x_17, x_18, x_19, x_20, x_134); -if (lean_obj_tag(x_135) == 0) +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +x_90 = lean_apply_12(x_54, x_88, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_89); +if (lean_obj_tag(x_90) == 0) { -lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; -x_136 = lean_ctor_get(x_135, 0); -lean_inc(x_136); -x_137 = lean_ctor_get(x_135, 1); -lean_inc(x_137); -lean_dec(x_135); -x_138 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__3; -lean_ctor_set_tag(x_121, 7); -lean_ctor_set(x_121, 1, x_133); -lean_ctor_set(x_121, 0, x_138); -x_139 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__5; -x_140 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_140, 0, x_121); -lean_ctor_set(x_140, 1, x_139); -x_141 = l_Lean_indentExpr(x_136); -x_142 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_142, 0, x_140); -lean_ctor_set(x_142, 1, x_141); -x_143 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; -x_144 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_144, 0, x_142); -lean_ctor_set(x_144, 1, x_143); -x_145 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_120, x_144, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, 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_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2; -x_149 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__1(x_26, x_148, x_4, x_5, x_146, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_147); -lean_dec(x_146); -return x_149; +lean_object* x_91; lean_object* x_92; +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_27 = x_91; +x_28 = x_92; +goto block_35; } else { -uint8_t x_150; -lean_dec(x_133); -lean_free_object(x_121); -lean_dec(x_26); +uint8_t x_93; +lean_dec(x_22); +lean_dec(x_21); lean_dec(x_20); lean_dec(x_19); lean_dec(x_18); lean_dec(x_17); -lean_dec(x_5); -lean_dec(x_4); -x_150 = !lean_is_exclusive(x_135); -if (x_150 == 0) +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_10); +lean_dec(x_7); +lean_dec(x_1); +x_93 = !lean_is_exclusive(x_90); +if (x_93 == 0) { -return x_135; +return x_90; } else { -lean_object* x_151; lean_object* x_152; lean_object* x_153; -x_151 = lean_ctor_get(x_135, 0); -x_152 = lean_ctor_get(x_135, 1); -lean_inc(x_152); -lean_inc(x_151); -lean_dec(x_135); -x_153 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_153, 0, x_151); -lean_ctor_set(x_153, 1, x_152); -return x_153; +lean_object* x_94; lean_object* x_95; lean_object* x_96; +x_94 = lean_ctor_get(x_90, 0); +x_95 = lean_ctor_get(x_90, 1); +lean_inc(x_95); +lean_inc(x_94); +lean_dec(x_90); +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_154; -lean_free_object(x_121); -lean_dec(x_119); -lean_dec(x_26); +uint8_t x_97; +lean_dec(x_68); +lean_free_object(x_50); +lean_dec(x_39); +lean_dec(x_36); +lean_dec(x_22); +lean_dec(x_21); lean_dec(x_20); lean_dec(x_19); lean_dec(x_18); lean_dec(x_17); -lean_dec(x_5); -lean_dec(x_4); -x_154 = !lean_is_exclusive(x_132); -if (x_154 == 0) +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_10); +lean_dec(x_7); +lean_dec(x_1); +x_97 = !lean_is_exclusive(x_70); +if (x_97 == 0) { -return x_132; +return x_70; } else { -lean_object* x_155; lean_object* x_156; lean_object* x_157; -x_155 = lean_ctor_get(x_132, 0); -x_156 = lean_ctor_get(x_132, 1); -lean_inc(x_156); -lean_inc(x_155); -lean_dec(x_132); -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_98; lean_object* x_99; lean_object* x_100; +x_98 = lean_ctor_get(x_70, 0); +x_99 = lean_ctor_get(x_70, 1); +lean_inc(x_99); +lean_inc(x_98); +lean_dec(x_70); +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_158; lean_object* x_159; lean_object* x_160; -x_158 = lean_ctor_get(x_121, 1); -lean_inc(x_158); -lean_dec(x_121); -x_159 = lean_ctor_get(x_4, 5); -lean_inc(x_159); -x_160 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_159, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_158); -if (lean_obj_tag(x_160) == 0) -{ -lean_object* x_161; lean_object* x_162; lean_object* x_163; -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); -lean_inc(x_20); -lean_inc(x_19); -lean_inc(x_18); -lean_inc(x_17); -x_163 = lean_infer_type(x_119, x_17, x_18, x_19, x_20, x_162); -if (lean_obj_tag(x_163) == 0) -{ -lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; -x_164 = lean_ctor_get(x_163, 0); -lean_inc(x_164); -x_165 = lean_ctor_get(x_163, 1); -lean_inc(x_165); -lean_dec(x_163); -x_166 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__3; -x_167 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_167, 0, x_166); -lean_ctor_set(x_167, 1, x_161); -x_168 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__5; -x_169 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_169, 0, x_167); -lean_ctor_set(x_169, 1, x_168); -x_170 = l_Lean_indentExpr(x_164); -x_171 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_171, 0, x_169); -lean_ctor_set(x_171, 1, x_170); -x_172 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; -x_173 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_173, 0, x_171); -lean_ctor_set(x_173, 1, x_172); -x_174 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_120, x_173, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_165); -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_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2; -x_178 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__1(x_26, x_177, x_4, x_5, x_175, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_176); -lean_dec(x_175); -return x_178; } else { -lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; -lean_dec(x_161); -lean_dec(x_26); +uint8_t x_101; +lean_free_object(x_50); +lean_dec(x_45); +lean_dec(x_42); +lean_dec(x_39); +lean_dec(x_36); +lean_dec(x_22); +lean_dec(x_21); lean_dec(x_20); lean_dec(x_19); lean_dec(x_18); lean_dec(x_17); -lean_dec(x_5); -lean_dec(x_4); -x_179 = lean_ctor_get(x_163, 0); -lean_inc(x_179); -x_180 = lean_ctor_get(x_163, 1); -lean_inc(x_180); -if (lean_is_exclusive(x_163)) { - lean_ctor_release(x_163, 0); - lean_ctor_release(x_163, 1); - x_181 = x_163; -} else { - lean_dec_ref(x_163); - x_181 = lean_box(0); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_10); +lean_dec(x_7); +lean_dec(x_1); +x_101 = !lean_is_exclusive(x_67); +if (x_101 == 0) +{ +return x_67; } -if (lean_is_scalar(x_181)) { - x_182 = lean_alloc_ctor(1, 2, 0); -} else { - x_182 = x_181; +else +{ +lean_object* x_102; lean_object* x_103; lean_object* x_104; +x_102 = lean_ctor_get(x_67, 0); +x_103 = lean_ctor_get(x_67, 1); +lean_inc(x_103); +lean_inc(x_102); +lean_dec(x_67); +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_ctor_set(x_182, 0, x_179); -lean_ctor_set(x_182, 1, x_180); -return x_182; } } else { -lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; -lean_dec(x_119); -lean_dec(x_26); +uint8_t x_105; +lean_free_object(x_50); +lean_dec(x_45); +lean_dec(x_42); +lean_dec(x_39); +lean_dec(x_36); +lean_dec(x_22); +lean_dec(x_21); lean_dec(x_20); lean_dec(x_19); lean_dec(x_18); lean_dec(x_17); -lean_dec(x_5); -lean_dec(x_4); -x_183 = lean_ctor_get(x_160, 0); -lean_inc(x_183); -x_184 = lean_ctor_get(x_160, 1); -lean_inc(x_184); -if (lean_is_exclusive(x_160)) { - lean_ctor_release(x_160, 0); - lean_ctor_release(x_160, 1); - x_185 = x_160; -} else { - lean_dec_ref(x_160); - x_185 = lean_box(0); -} -if (lean_is_scalar(x_185)) { - x_186 = lean_alloc_ctor(1, 2, 0); -} else { - x_186 = x_185; -} -lean_ctor_set(x_186, 0, x_183); -lean_ctor_set(x_186, 1, x_184); -return x_186; -} +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_10); +lean_dec(x_7); +lean_dec(x_1); +x_105 = !lean_is_exclusive(x_64); +if (x_105 == 0) +{ +return x_64; } +else +{ +lean_object* x_106; lean_object* x_107; lean_object* x_108; +x_106 = lean_ctor_get(x_64, 0); +x_107 = lean_ctor_get(x_64, 1); +lean_inc(x_107); +lean_inc(x_106); +lean_dec(x_64); +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 { -size_t x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; size_t x_206; lean_object* x_207; lean_object* x_208; lean_object* x_277; -x_200 = lean_usize_of_nat(x_1); -x_201 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2; -x_202 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__8(x_3, x_8, x_200, x_201, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_25); -x_203 = lean_ctor_get(x_202, 0); -lean_inc(x_203); -x_204 = lean_ctor_get(x_202, 1); -lean_inc(x_204); -if (lean_is_exclusive(x_202)) { - lean_ctor_release(x_202, 0); - lean_ctor_release(x_202, 1); - x_205 = x_202; -} else { - lean_dec_ref(x_202); - x_205 = lean_box(0); -} -x_206 = lean_array_size(x_203); +lean_object* x_109; lean_object* x_110; lean_object* x_111; uint8_t x_112; +x_109 = lean_ctor_get(x_50, 0); +x_110 = lean_ctor_get(x_50, 1); +lean_inc(x_110); +lean_inc(x_109); +lean_dec(x_50); +x_111 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__3; +x_112 = lean_unbox(x_109); +lean_dec(x_109); +if (x_112 == 0) +{ +lean_object* x_113; lean_object* x_114; +lean_dec(x_45); +lean_dec(x_42); +lean_dec(x_39); +lean_dec(x_36); +x_113 = lean_box(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_7); -x_277 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__9(x_6, x_203, x_7, x_203, x_206, x_8, x_7, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_204); -if (lean_obj_tag(x_277) == 0) -{ -lean_object* x_278; lean_object* x_279; -x_278 = lean_ctor_get(x_277, 0); -lean_inc(x_278); -x_279 = lean_ctor_get(x_278, 0); -lean_inc(x_279); -lean_dec(x_278); -if (lean_obj_tag(x_279) == 0) -{ -lean_object* x_280; -x_280 = lean_ctor_get(x_277, 1); -lean_inc(x_280); -lean_dec(x_277); -x_207 = x_9; -x_208 = x_280; -goto block_276; -} -else +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +x_114 = lean_apply_12(x_111, x_113, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_110); +if (lean_obj_tag(x_114) == 0) { -lean_object* x_281; lean_object* x_282; -lean_dec(x_9); -x_281 = lean_ctor_get(x_277, 1); -lean_inc(x_281); -lean_dec(x_277); -x_282 = lean_ctor_get(x_279, 0); -lean_inc(x_282); -lean_dec(x_279); -x_207 = x_282; -x_208 = x_281; -goto block_276; -} +lean_object* x_115; lean_object* x_116; +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_27 = x_115; +x_28 = x_116; +goto block_35; } else { -uint8_t x_283; -lean_dec(x_205); -lean_dec(x_203); -lean_dec(x_26); +lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; +lean_dec(x_22); +lean_dec(x_21); lean_dec(x_20); lean_dec(x_19); lean_dec(x_18); lean_dec(x_17); -lean_dec(x_9); -lean_dec(x_5); -lean_dec(x_4); -x_283 = !lean_is_exclusive(x_277); -if (x_283 == 0) -{ -return x_277; -} -else -{ -lean_object* x_284; lean_object* x_285; lean_object* x_286; -x_284 = lean_ctor_get(x_277, 0); -x_285 = lean_ctor_get(x_277, 1); -lean_inc(x_285); -lean_inc(x_284); -lean_dec(x_277); -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; +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_10); +lean_dec(x_7); +lean_dec(x_1); +x_117 = lean_ctor_get(x_114, 0); +lean_inc(x_117); +x_118 = lean_ctor_get(x_114, 1); +lean_inc(x_118); +if (lean_is_exclusive(x_114)) { + lean_ctor_release(x_114, 0); + lean_ctor_release(x_114, 1); + x_119 = x_114; +} else { + lean_dec_ref(x_114); + x_119 = lean_box(0); } +if (lean_is_scalar(x_119)) { + x_120 = lean_alloc_ctor(1, 2, 0); +} else { + x_120 = x_119; } -block_276: -{ -if (lean_obj_tag(x_207) == 0) -{ -lean_object* x_209; lean_object* x_210; -lean_dec(x_205); -x_209 = lean_box(0); -x_210 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__1(x_26, x_203, x_4, x_5, x_209, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_208); -lean_dec(x_203); -return x_210; +lean_ctor_set(x_120, 0, x_117); +lean_ctor_set(x_120, 1, x_118); +return x_120; } -else -{ -lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; uint8_t x_215; -x_211 = lean_ctor_get(x_207, 0); -lean_inc(x_211); -lean_dec(x_207); -x_212 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__2; -x_213 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__2(x_212, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_208); -x_214 = lean_ctor_get(x_213, 0); -lean_inc(x_214); -x_215 = lean_unbox(x_214); -lean_dec(x_214); -if (x_215 == 0) -{ -lean_object* x_216; lean_object* x_217; lean_object* x_218; -lean_dec(x_211); -lean_dec(x_205); -x_216 = lean_ctor_get(x_213, 1); -lean_inc(x_216); -lean_dec(x_213); -x_217 = lean_box(0); -x_218 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__1(x_26, x_203, x_4, x_5, x_217, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_216); -lean_dec(x_203); -return x_218; } else { -uint8_t x_219; -x_219 = !lean_is_exclusive(x_213); -if (x_219 == 0) +lean_object* x_121; +x_121 = l_Lean_Meta_Grind_updateLastTag(x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_110); +if (lean_obj_tag(x_121) == 0) { -lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; -x_220 = lean_ctor_get(x_213, 1); -x_221 = lean_ctor_get(x_213, 0); -lean_dec(x_221); -x_222 = lean_ctor_get(x_4, 5); -lean_inc(x_222); -x_223 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_222, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_220); -if (lean_obj_tag(x_223) == 0) +lean_object* x_122; lean_object* x_123; lean_object* x_124; +x_122 = lean_ctor_get(x_121, 1); +lean_inc(x_122); +lean_dec(x_121); +x_123 = lean_ctor_get(x_1, 5); +lean_inc(x_123); +x_124 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_123, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_122); +if (lean_obj_tag(x_124) == 0) { -lean_object* x_224; lean_object* x_225; lean_object* x_226; -x_224 = lean_ctor_get(x_223, 0); -lean_inc(x_224); -x_225 = lean_ctor_get(x_223, 1); -lean_inc(x_225); -lean_dec(x_223); +lean_object* x_125; lean_object* x_126; lean_object* x_127; +x_125 = lean_ctor_get(x_124, 0); +lean_inc(x_125); +x_126 = lean_ctor_get(x_124, 1); +lean_inc(x_126); +lean_dec(x_124); +lean_inc(x_22); +lean_inc(x_21); +lean_inc(x_20); +lean_inc(x_19); +x_127 = l_Lean_Meta_mkHasTypeButIsExpectedMsg(x_45, x_42, x_19, x_20, x_21, x_22, x_126); +if (lean_obj_tag(x_127) == 0) +{ +lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; +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_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__5; +x_131 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_131, 0, x_130); +lean_ctor_set(x_131, 1, x_125); +x_132 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__7; +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_39); +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_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__9; +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_36); +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_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__11; +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 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_142, 0, x_141); +lean_ctor_set(x_142, 1, x_128); +x_143 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; +x_144 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_144, 0, x_142); +lean_ctor_set(x_144, 1, x_143); +x_145 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_49, x_144, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_129); +x_146 = lean_ctor_get(x_145, 0); +lean_inc(x_146); +x_147 = lean_ctor_get(x_145, 1); +lean_inc(x_147); +lean_dec(x_145); +lean_inc(x_22); +lean_inc(x_21); lean_inc(x_20); lean_inc(x_19); lean_inc(x_18); lean_inc(x_17); -x_226 = lean_infer_type(x_211, x_17, x_18, x_19, x_20, x_225); -if (lean_obj_tag(x_226) == 0) +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +x_148 = lean_apply_12(x_111, x_146, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_147); +if (lean_obj_tag(x_148) == 0) { -lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; -x_227 = lean_ctor_get(x_226, 0); -lean_inc(x_227); -x_228 = lean_ctor_get(x_226, 1); -lean_inc(x_228); -lean_dec(x_226); -x_229 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__3; -lean_ctor_set_tag(x_213, 7); -lean_ctor_set(x_213, 1, x_224); -lean_ctor_set(x_213, 0, x_229); -x_230 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__5; -if (lean_is_scalar(x_205)) { - x_231 = lean_alloc_ctor(7, 2, 0); -} else { - x_231 = x_205; - lean_ctor_set_tag(x_231, 7); -} -lean_ctor_set(x_231, 0, x_213); -lean_ctor_set(x_231, 1, x_230); -x_232 = l_Lean_indentExpr(x_227); -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_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; -x_235 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_235, 0, x_233); -lean_ctor_set(x_235, 1, x_234); -x_236 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_212, x_235, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_228); -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 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__1(x_26, x_203, x_4, x_5, x_237, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_238); -lean_dec(x_237); -lean_dec(x_203); -return x_239; +lean_object* x_149; lean_object* x_150; +x_149 = lean_ctor_get(x_148, 0); +lean_inc(x_149); +x_150 = lean_ctor_get(x_148, 1); +lean_inc(x_150); +lean_dec(x_148); +x_27 = x_149; +x_28 = x_150; +goto block_35; } else { -uint8_t x_240; -lean_dec(x_224); -lean_free_object(x_213); -lean_dec(x_205); -lean_dec(x_203); -lean_dec(x_26); +lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; +lean_dec(x_22); +lean_dec(x_21); lean_dec(x_20); lean_dec(x_19); lean_dec(x_18); lean_dec(x_17); -lean_dec(x_5); -lean_dec(x_4); -x_240 = !lean_is_exclusive(x_226); -if (x_240 == 0) -{ -return x_226; +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_10); +lean_dec(x_7); +lean_dec(x_1); +x_151 = lean_ctor_get(x_148, 0); +lean_inc(x_151); +x_152 = lean_ctor_get(x_148, 1); +lean_inc(x_152); +if (lean_is_exclusive(x_148)) { + lean_ctor_release(x_148, 0); + lean_ctor_release(x_148, 1); + x_153 = x_148; +} else { + lean_dec_ref(x_148); + x_153 = lean_box(0); } -else -{ -lean_object* x_241; lean_object* x_242; lean_object* x_243; -x_241 = lean_ctor_get(x_226, 0); -x_242 = lean_ctor_get(x_226, 1); -lean_inc(x_242); -lean_inc(x_241); -lean_dec(x_226); -x_243 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_243, 0, x_241); -lean_ctor_set(x_243, 1, x_242); -return x_243; +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 { -uint8_t x_244; -lean_free_object(x_213); -lean_dec(x_211); -lean_dec(x_205); -lean_dec(x_203); -lean_dec(x_26); +lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; +lean_dec(x_125); +lean_dec(x_39); +lean_dec(x_36); +lean_dec(x_22); +lean_dec(x_21); lean_dec(x_20); lean_dec(x_19); lean_dec(x_18); lean_dec(x_17); -lean_dec(x_5); -lean_dec(x_4); -x_244 = !lean_is_exclusive(x_223); -if (x_244 == 0) -{ -return x_223; -} -else -{ -lean_object* x_245; lean_object* x_246; lean_object* x_247; -x_245 = lean_ctor_get(x_223, 0); -x_246 = lean_ctor_get(x_223, 1); -lean_inc(x_246); -lean_inc(x_245); -lean_dec(x_223); -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; -} -} +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_10); +lean_dec(x_7); +lean_dec(x_1); +x_155 = lean_ctor_get(x_127, 0); +lean_inc(x_155); +x_156 = lean_ctor_get(x_127, 1); +lean_inc(x_156); +if (lean_is_exclusive(x_127)) { + lean_ctor_release(x_127, 0); + lean_ctor_release(x_127, 1); + x_157 = x_127; +} else { + lean_dec_ref(x_127); + x_157 = lean_box(0); } -else -{ -lean_object* x_248; lean_object* x_249; lean_object* x_250; -x_248 = lean_ctor_get(x_213, 1); -lean_inc(x_248); -lean_dec(x_213); -x_249 = lean_ctor_get(x_4, 5); -lean_inc(x_249); -x_250 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_249, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_248); -if (lean_obj_tag(x_250) == 0) -{ -lean_object* x_251; lean_object* x_252; lean_object* x_253; -x_251 = lean_ctor_get(x_250, 0); -lean_inc(x_251); -x_252 = lean_ctor_get(x_250, 1); -lean_inc(x_252); -lean_dec(x_250); -lean_inc(x_20); -lean_inc(x_19); -lean_inc(x_18); -lean_inc(x_17); -x_253 = lean_infer_type(x_211, x_17, x_18, x_19, x_20, x_252); -if (lean_obj_tag(x_253) == 0) -{ -lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; -x_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_253); -x_256 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__3; -x_257 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_257, 0, x_256); -lean_ctor_set(x_257, 1, x_251); -x_258 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__5; -if (lean_is_scalar(x_205)) { - x_259 = lean_alloc_ctor(7, 2, 0); +if (lean_is_scalar(x_157)) { + x_158 = lean_alloc_ctor(1, 2, 0); } else { - x_259 = x_205; - lean_ctor_set_tag(x_259, 7); + x_158 = x_157; +} +lean_ctor_set(x_158, 0, x_155); +lean_ctor_set(x_158, 1, x_156); +return x_158; } -lean_ctor_set(x_259, 0, x_257); -lean_ctor_set(x_259, 1, x_258); -x_260 = l_Lean_indentExpr(x_254); -x_261 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_261, 0, x_259); -lean_ctor_set(x_261, 1, x_260); -x_262 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; -x_263 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_263, 0, x_261); -lean_ctor_set(x_263, 1, x_262); -x_264 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_212, x_263, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_255); -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___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__1(x_26, x_203, x_4, x_5, x_265, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_266); -lean_dec(x_265); -lean_dec(x_203); -return x_267; } else { -lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; -lean_dec(x_251); -lean_dec(x_205); -lean_dec(x_203); -lean_dec(x_26); +lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; +lean_dec(x_45); +lean_dec(x_42); +lean_dec(x_39); +lean_dec(x_36); +lean_dec(x_22); +lean_dec(x_21); lean_dec(x_20); lean_dec(x_19); lean_dec(x_18); lean_dec(x_17); -lean_dec(x_5); -lean_dec(x_4); -x_268 = lean_ctor_get(x_253, 0); -lean_inc(x_268); -x_269 = lean_ctor_get(x_253, 1); -lean_inc(x_269); -if (lean_is_exclusive(x_253)) { - lean_ctor_release(x_253, 0); - lean_ctor_release(x_253, 1); - x_270 = x_253; +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_10); +lean_dec(x_7); +lean_dec(x_1); +x_159 = lean_ctor_get(x_124, 0); +lean_inc(x_159); +x_160 = lean_ctor_get(x_124, 1); +lean_inc(x_160); +if (lean_is_exclusive(x_124)) { + lean_ctor_release(x_124, 0); + lean_ctor_release(x_124, 1); + x_161 = x_124; } else { - lean_dec_ref(x_253); - x_270 = lean_box(0); + lean_dec_ref(x_124); + x_161 = lean_box(0); } -if (lean_is_scalar(x_270)) { - x_271 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_161)) { + x_162 = lean_alloc_ctor(1, 2, 0); } else { - x_271 = x_270; + x_162 = x_161; } -lean_ctor_set(x_271, 0, x_268); -lean_ctor_set(x_271, 1, x_269); -return x_271; +lean_ctor_set(x_162, 0, x_159); +lean_ctor_set(x_162, 1, x_160); +return x_162; } } else { -lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; -lean_dec(x_211); -lean_dec(x_205); -lean_dec(x_203); -lean_dec(x_26); +lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; +lean_dec(x_45); +lean_dec(x_42); +lean_dec(x_39); +lean_dec(x_36); +lean_dec(x_22); +lean_dec(x_21); lean_dec(x_20); lean_dec(x_19); lean_dec(x_18); lean_dec(x_17); -lean_dec(x_5); -lean_dec(x_4); -x_272 = lean_ctor_get(x_250, 0); -lean_inc(x_272); -x_273 = lean_ctor_get(x_250, 1); -lean_inc(x_273); -if (lean_is_exclusive(x_250)) { - lean_ctor_release(x_250, 0); - lean_ctor_release(x_250, 1); - x_274 = x_250; +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_10); +lean_dec(x_7); +lean_dec(x_1); +x_163 = lean_ctor_get(x_121, 0); +lean_inc(x_163); +x_164 = lean_ctor_get(x_121, 1); +lean_inc(x_164); +if (lean_is_exclusive(x_121)) { + lean_ctor_release(x_121, 0); + lean_ctor_release(x_121, 1); + x_165 = x_121; } else { - lean_dec_ref(x_250); - x_274 = lean_box(0); + lean_dec_ref(x_121); + x_165 = lean_box(0); } -if (lean_is_scalar(x_274)) { - x_275 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_165)) { + x_166 = lean_alloc_ctor(1, 2, 0); } else { - x_275 = x_274; -} -lean_ctor_set(x_275, 0, x_272); -lean_ctor_set(x_275, 1, x_273); -return x_275; + 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; +lean_dec(x_45); +lean_dec(x_42); +lean_dec(x_39); +lean_dec(x_36); +lean_inc(x_7); +x_167 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_167, 0, x_7); +x_27 = x_167; +x_28 = x_48; +goto block_35; } } } else { -lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; -lean_dec(x_9); +uint8_t x_187; +lean_dec(x_42); +lean_dec(x_40); +lean_dec(x_39); +lean_dec(x_36); +lean_dec(x_22); +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_10); lean_dec(x_7); -x_287 = lean_ctor_get(x_4, 5); -lean_inc(x_287); -lean_dec(x_4); -x_288 = l_Lean_mkAppN(x_2, x_3); -x_289 = lean_ctor_get(x_5, 1); -lean_inc(x_289); -lean_dec(x_5); -x_290 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance(x_287, x_288, x_289, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_25); -return x_290; +lean_dec(x_1); +x_187 = !lean_is_exclusive(x_44); +if (x_187 == 0) +{ +return x_44; } +else +{ +lean_object* x_188; lean_object* x_189; lean_object* x_190; +x_188 = lean_ctor_get(x_44, 0); +x_189 = lean_ctor_get(x_44, 1); +lean_inc(x_189); +lean_inc(x_188); +lean_dec(x_44); +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_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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, 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: +else { -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_21 = lean_array_get_size(x_1); -x_22 = lean_unsigned_to_nat(0u); -x_23 = l_Array_toSubarray___rarg(x_1, x_22, x_21); -x_24 = lean_box(0); -lean_inc(x_2); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_2); -lean_ctor_set(x_25, 1, x_23); -x_26 = lean_array_size(x_3); -x_27 = 0; -lean_inc(x_19); -lean_inc(x_18); -lean_inc(x_17); -lean_inc(x_16); -lean_inc(x_2); -lean_inc(x_4); -x_28 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5(x_4, x_3, x_2, x_24, x_3, x_26, x_27, x_25, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20); -if (lean_obj_tag(x_28) == 0) -{ -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); -if (lean_obj_tag(x_30) == 0) -{ -lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_31 = lean_ctor_get(x_28, 1); -lean_inc(x_31); -lean_dec(x_28); -x_32 = lean_box(0); -x_33 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3(x_5, x_6, x_3, x_4, x_7, x_24, x_8, x_27, x_2, x_32, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_31); -return x_33; -} -else -{ -uint8_t x_34; +uint8_t x_191; +lean_dec(x_40); +lean_dec(x_39); +lean_dec(x_36); +lean_dec(x_22); +lean_dec(x_21); +lean_dec(x_20); lean_dec(x_19); lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); -lean_dec(x_8); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_10); lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_4); -lean_dec(x_2); -x_34 = !lean_is_exclusive(x_28); -if (x_34 == 0) +lean_dec(x_1); +x_191 = !lean_is_exclusive(x_41); +if (x_191 == 0) { -lean_object* x_35; lean_object* x_36; -x_35 = lean_ctor_get(x_28, 0); -lean_dec(x_35); -x_36 = lean_ctor_get(x_30, 0); -lean_inc(x_36); -lean_dec(x_30); -lean_ctor_set(x_28, 0, x_36); -return x_28; +return x_41; } else { -lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_37 = lean_ctor_get(x_28, 1); -lean_inc(x_37); -lean_dec(x_28); -x_38 = lean_ctor_get(x_30, 0); -lean_inc(x_38); -lean_dec(x_30); -x_39 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_39, 0, x_38); -lean_ctor_set(x_39, 1, x_37); -return x_39; +lean_object* x_192; lean_object* x_193; lean_object* x_194; +x_192 = lean_ctor_get(x_41, 0); +x_193 = lean_ctor_get(x_41, 1); +lean_inc(x_193); +lean_inc(x_192); +lean_dec(x_41); +x_194 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_194, 0, x_192); +lean_ctor_set(x_194, 1, x_193); +return x_194; } } } else { -uint8_t x_40; -lean_dec(x_19); -lean_dec(x_18); -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_4); -lean_dec(x_2); -x_40 = !lean_is_exclusive(x_28); -if (x_40 == 0) -{ -return x_28; +lean_object* x_195; +lean_dec(x_36); +lean_inc(x_7); +x_195 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_195, 0, x_7); +x_27 = x_195; +x_28 = x_23; +goto block_35; } -else -{ -lean_object* x_41; lean_object* x_42; lean_object* x_43; -x_41 = lean_ctor_get(x_28, 0); -x_42 = lean_ctor_get(x_28, 1); -lean_inc(x_42); -lean_inc(x_41); -lean_dec(x_28); -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; } } } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__5___closed__1() { +LEAN_EXPORT lean_object* l_Lean_MVarId_isAssigned___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -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___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19, lean_object* x_20) { -_start: +lean_object* x_13; uint8_t x_14; +x_13 = lean_st_ref_get(x_9, x_12); +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 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; -x_21 = lean_array_get_size(x_1); -x_22 = lean_unsigned_to_nat(0u); -x_23 = lean_unsigned_to_nat(1u); -lean_inc(x_21); -x_24 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_24, 0, x_22); -lean_ctor_set(x_24, 1, x_21); -lean_ctor_set(x_24, 2, x_23); -x_25 = lean_box(0); -x_26 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__5___closed__1; -lean_inc(x_19); -lean_inc(x_18); -lean_inc(x_17); +lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; lean_object* x_19; +x_15 = lean_ctor_get(x_13, 0); +x_16 = lean_ctor_get(x_15, 0); lean_inc(x_16); -lean_inc(x_15); -lean_inc(x_14); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_2); -x_27 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3(x_2, x_3, x_4, x_5, x_1, x_24, x_26, x_24, x_26, x_22, lean_box(0), lean_box(0), x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20); -lean_dec(x_24); -if (lean_obj_tag(x_27) == 0) -{ -lean_object* x_28; lean_object* x_29; -x_28 = lean_ctor_get(x_27, 0); -lean_inc(x_28); -x_29 = lean_ctor_get(x_28, 0); -lean_inc(x_29); -lean_dec(x_28); -if (lean_obj_tag(x_29) == 0) -{ -lean_object* x_30; lean_object* x_31; lean_object* x_32; -x_30 = lean_ctor_get(x_27, 1); -lean_inc(x_30); -lean_dec(x_27); -x_31 = lean_box(0); -x_32 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__4(x_6, x_25, x_1, x_2, x_21, x_7, x_8, x_26, x_31, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_30); 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_21); -return x_32; +x_17 = lean_ctor_get(x_16, 7); +lean_inc(x_17); +lean_dec(x_16); +x_18 = l_Lean_PersistentHashMap_contains___at_Lean_MVarId_isAssigned___spec__1(x_17, x_1); +x_19 = lean_box(x_18); +lean_ctor_set(x_13, 0, x_19); +return x_13; } else { -uint8_t x_33; -lean_dec(x_21); -lean_dec(x_19); -lean_dec(x_18); -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; lean_object* x_25; lean_object* x_26; +x_20 = lean_ctor_get(x_13, 0); +x_21 = lean_ctor_get(x_13, 1); +lean_inc(x_21); +lean_inc(x_20); lean_dec(x_13); -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_2); -x_33 = !lean_is_exclusive(x_27); -if (x_33 == 0) +x_22 = lean_ctor_get(x_20, 0); +lean_inc(x_22); +lean_dec(x_20); +x_23 = lean_ctor_get(x_22, 7); +lean_inc(x_23); +lean_dec(x_22); +x_24 = l_Lean_PersistentHashMap_contains___at_Lean_MVarId_isAssigned___spec__1(x_23, x_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_21); +return x_26; +} +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: { -lean_object* x_34; lean_object* x_35; -x_34 = lean_ctor_get(x_27, 0); -lean_dec(x_34); -x_35 = lean_ctor_get(x_29, 0); -lean_inc(x_35); -lean_dec(x_29); -lean_ctor_set(x_27, 0, x_35); -return x_27; +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_14 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___lambda__1___closed__1; +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_14); +lean_ctor_set(x_15, 1, x_1); +x_16 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_16, 0, 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; } -else +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___closed__1() { +_start: { -lean_object* x_36; lean_object* x_37; lean_object* x_38; -x_36 = lean_ctor_get(x_27, 1); -lean_inc(x_36); -lean_dec(x_27); -x_37 = lean_ctor_get(x_29, 0); -lean_inc(x_37); -lean_dec(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_36); -return x_38; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("failed to synthesize instance when instantiating ", 49, 49); +return x_1; } } +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} } -else +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, size_t x_6, size_t x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, 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_39; -lean_dec(x_21); -lean_dec(x_19); +uint8_t x_20; +x_20 = lean_usize_dec_lt(x_7, x_6); +if (x_20 == 0) +{ +lean_object* x_21; lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_2); -x_39 = !lean_is_exclusive(x_27); -if (x_39 == 0) -{ -return x_27; +lean_dec(x_3); +lean_dec(x_1); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_8); +lean_ctor_set(x_21, 1, x_19); +return x_21; } else { -lean_object* x_40; lean_object* x_41; lean_object* x_42; -x_40 = lean_ctor_get(x_27, 0); -x_41 = lean_ctor_get(x_27, 1); -lean_inc(x_41); -lean_inc(x_40); -lean_dec(x_27); -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_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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) { -_start: +lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_32; +x_22 = lean_array_uget(x_5, x_7); +x_32 = !lean_is_exclusive(x_8); +if (x_32 == 0) { -lean_object* x_13; lean_object* x_14; -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; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__1() { -_start: +lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; uint8_t x_38; +x_33 = lean_ctor_get(x_8, 1); +x_34 = lean_ctor_get(x_8, 0); +lean_dec(x_34); +x_35 = lean_ctor_get(x_33, 0); +lean_inc(x_35); +x_36 = lean_ctor_get(x_33, 1); +lean_inc(x_36); +x_37 = lean_ctor_get(x_33, 2); +lean_inc(x_37); +x_38 = lean_nat_dec_lt(x_36, x_37); +if (x_38 == 0) { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("assertion violation: ", 21, 21); -return x_1; -} +lean_object* x_39; +lean_dec(x_37); +lean_dec(x_36); +lean_dec(x_35); +lean_dec(x_22); +lean_inc(x_3); +lean_ctor_set(x_8, 0, x_3); +x_39 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_39, 0, x_8); +x_23 = x_39; +x_24 = x_19; +goto block_31; } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__2() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("c.assignment.size == numParams\n ", 33, 33); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__3() { -_start: +uint8_t x_40; +x_40 = !lean_is_exclusive(x_33); +if (x_40 == 0) { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__1; -x_2 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__2; -x_3 = lean_string_append(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__4() { -_start: +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; uint8_t x_50; +x_41 = lean_ctor_get(x_33, 2); +lean_dec(x_41); +x_42 = lean_ctor_get(x_33, 1); +lean_dec(x_42); +x_43 = lean_ctor_get(x_33, 0); +lean_dec(x_43); +x_44 = lean_array_fget(x_35, x_36); +x_45 = lean_unbox(x_44); +lean_dec(x_44); +x_46 = lean_unsigned_to_nat(1u); +x_47 = lean_nat_add(x_36, x_46); +lean_dec(x_36); +lean_ctor_set(x_33, 1, x_47); +x_48 = l_Lean_Expr_mvarId_x21(x_22); +x_49 = l_Lean_MVarId_isAssigned___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__4(x_48, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19); +lean_dec(x_48); +x_50 = !lean_is_exclusive(x_49); +if (x_50 == 0) { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("_private.Lean.Meta.Tactic.Grind.EMatch.0.Lean.Meta.Grind.EMatch.instantiateTheorem", 82, 82); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__5() { -_start: +lean_object* x_51; lean_object* x_52; uint8_t x_53; +x_51 = lean_ctor_get(x_49, 0); +x_52 = lean_ctor_get(x_49, 1); +x_53 = l_Lean_BinderInfo_isInstImplicit(x_45); +if (x_53 == 0) { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_1 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___closed__1; -x_2 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__4; -x_3 = lean_unsigned_to_nat(172u); -x_4 = lean_unsigned_to_nat(2u); -x_5 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___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_54; +lean_free_object(x_49); +lean_dec(x_51); +lean_dec(x_22); +lean_inc(x_3); +lean_ctor_set(x_8, 0, x_3); +x_54 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_54, 0, x_8); +x_23 = x_54; +x_24 = x_52; +goto block_31; } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__6() { -_start: +else { -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___boxed), 12, 0); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__7() { -_start: +uint8_t x_55; +x_55 = lean_unbox(x_51); +lean_dec(x_51); +if (x_55 == 0) { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("unexpected number of parameters at ", 35, 35); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__8() { -_start: +lean_object* x_56; +lean_inc(x_18); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_22); +x_56 = lean_infer_type(x_22, x_15, x_16, x_17, x_18, x_52); +if (lean_obj_tag(x_56) == 0) { -lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__7; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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_object* x_15) { -_start: +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_56, 1); +lean_inc(x_58); +lean_dec(x_56); +lean_inc(x_18); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_57); +x_59 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem_synthesizeInstance(x_22, x_57, x_15, x_16, x_17, x_18, x_58); +if (lean_obj_tag(x_59) == 0) { -lean_object* x_16; -lean_inc(x_1); -x_16 = l_Lean_Meta_Grind_EMatchTheorem_getProofWithFreshMVarLevels(x_1, x_11, x_12, x_13, x_14, x_15); -if (lean_obj_tag(x_16) == 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_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -x_18 = lean_ctor_get(x_16, 1); -lean_inc(x_18); -lean_dec(x_16); -x_19 = lean_ctor_get(x_1, 2); -lean_inc(x_19); -x_20 = lean_array_get_size(x_2); -x_21 = lean_nat_dec_eq(x_20, x_19); -if (x_21 == 0) +lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; uint8_t x_66; +lean_free_object(x_8); +x_62 = lean_ctor_get(x_59, 1); +lean_inc(x_62); +lean_dec(x_59); +x_63 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__2; +x_64 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__2(x_63, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_62); +x_65 = lean_ctor_get(x_64, 0); +lean_inc(x_65); +x_66 = lean_unbox(x_65); +lean_dec(x_65); +if (x_66 == 0) { -lean_object* x_22; lean_object* x_23; -lean_dec(x_20); -lean_dec(x_19); -lean_dec(x_17); -lean_dec(x_3); -lean_dec(x_1); -x_22 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__5; -x_23 = l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__2(x_22, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_18); -return x_23; +lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; +lean_dec(x_57); +lean_free_object(x_49); +x_67 = lean_ctor_get(x_64, 1); +lean_inc(x_67); +lean_dec(x_64); +x_68 = lean_box(0); +x_69 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___lambda__1(x_33, x_68, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_67); +x_70 = lean_ctor_get(x_69, 0); +lean_inc(x_70); +x_71 = lean_ctor_get(x_69, 1); +lean_inc(x_71); +lean_dec(x_69); +x_23 = x_70; +x_24 = x_71; +goto block_31; } else { -lean_object* x_24; -lean_inc(x_14); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_17); -x_24 = lean_infer_type(x_17, x_11, x_12, x_13, x_14, x_18); -if (lean_obj_tag(x_24) == 0) +uint8_t x_72; +x_72 = !lean_is_exclusive(x_64); +if (x_72 == 0) { -lean_object* x_25; lean_object* x_26; uint8_t x_27; lean_object* x_28; -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 = 0; -lean_inc(x_14); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_19); -x_28 = l_Lean_Meta_forallMetaBoundedTelescope(x_25, x_19, x_27, x_11, x_12, x_13, x_14, x_26); -if (lean_obj_tag(x_28) == 0) +lean_object* x_73; lean_object* x_74; lean_object* x_75; +x_73 = lean_ctor_get(x_64, 1); +x_74 = lean_ctor_get(x_64, 0); +lean_dec(x_74); +x_75 = l_Lean_Meta_Grind_updateLastTag(x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_73); +if (lean_obj_tag(x_75) == 0) { -lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; -x_29 = lean_ctor_get(x_28, 0); -lean_inc(x_29); -x_30 = lean_ctor_get(x_29, 1); -lean_inc(x_30); -x_31 = lean_ctor_get(x_28, 1); -lean_inc(x_31); -lean_dec(x_28); -x_32 = lean_ctor_get(x_29, 0); -lean_inc(x_32); -lean_dec(x_29); -x_33 = !lean_is_exclusive(x_30); -if (x_33 == 0) +lean_object* x_76; lean_object* x_77; lean_object* x_78; +x_76 = lean_ctor_get(x_75, 1); +lean_inc(x_76); +lean_dec(x_75); +x_77 = lean_ctor_get(x_1, 5); +lean_inc(x_77); +x_78 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_77, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_76); +if (lean_obj_tag(x_78) == 0) { -lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; -x_34 = lean_ctor_get(x_30, 0); -x_35 = lean_ctor_get(x_30, 1); -lean_dec(x_35); -x_36 = lean_array_get_size(x_32); -x_37 = lean_nat_dec_eq(x_36, x_19); -lean_dec(x_36); -if (x_37 == 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_object* x_89; lean_object* x_90; lean_object* x_91; +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 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___closed__2; +lean_ctor_set_tag(x_64, 7); +lean_ctor_set(x_64, 1, x_79); +lean_ctor_set(x_64, 0, x_81); +x_82 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; +lean_ctor_set_tag(x_49, 7); +lean_ctor_set(x_49, 1, x_82); +lean_ctor_set(x_49, 0, x_64); +x_83 = l_Lean_indentExpr(x_57); +x_84 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_84, 0, x_49); +lean_ctor_set(x_84, 1, x_83); +x_85 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_85, 0, x_84); +lean_ctor_set(x_85, 1, x_82); +x_86 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_63, x_85, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_80); +x_87 = lean_ctor_get(x_86, 0); +lean_inc(x_87); +x_88 = lean_ctor_get(x_86, 1); +lean_inc(x_88); +lean_dec(x_86); +x_89 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___lambda__1(x_33, x_87, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_88); +lean_dec(x_87); +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_23 = x_90; +x_24 = x_91; +goto block_31; +} +else { -lean_object* x_38; lean_object* x_39; uint8_t x_40; -lean_dec(x_34); -lean_dec(x_32); -lean_dec(x_20); -lean_dec(x_19); +uint8_t x_92; +lean_free_object(x_64); +lean_dec(x_57); +lean_free_object(x_49); +lean_dec(x_33); +lean_dec(x_18); lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); lean_dec(x_3); -x_38 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__2; -x_39 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__2(x_38, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_31); -x_40 = !lean_is_exclusive(x_39); -if (x_40 == 0) +lean_dec(x_1); +x_92 = !lean_is_exclusive(x_78); +if (x_92 == 0) { -lean_object* x_41; lean_object* x_42; lean_object* x_43; uint8_t x_44; -x_41 = lean_ctor_get(x_39, 0); -x_42 = lean_ctor_get(x_39, 1); -x_43 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__6; -x_44 = lean_unbox(x_41); -lean_dec(x_41); -if (x_44 == 0) +return x_78; +} +else { -lean_object* x_45; lean_object* x_46; -lean_free_object(x_39); -lean_free_object(x_30); +lean_object* x_93; lean_object* x_94; lean_object* x_95; +x_93 = lean_ctor_get(x_78, 0); +x_94 = lean_ctor_get(x_78, 1); +lean_inc(x_94); +lean_inc(x_93); +lean_dec(x_78); +x_95 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_95, 0, x_93); +lean_ctor_set(x_95, 1, x_94); +return x_95; +} +} +} +else +{ +uint8_t x_96; +lean_free_object(x_64); +lean_dec(x_57); +lean_free_object(x_49); +lean_dec(x_33); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_3); lean_dec(x_1); -x_45 = lean_box(0); -x_46 = lean_apply_12(x_43, x_45, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_42); -return x_46; +x_96 = !lean_is_exclusive(x_75); +if (x_96 == 0) +{ +return x_75; } else { -lean_object* x_47; lean_object* x_48; -x_47 = lean_ctor_get(x_1, 5); -lean_inc(x_47); +lean_object* x_97; lean_object* x_98; lean_object* x_99; +x_97 = lean_ctor_get(x_75, 0); +x_98 = lean_ctor_get(x_75, 1); +lean_inc(x_98); +lean_inc(x_97); +lean_dec(x_75); +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 +{ +lean_object* x_100; lean_object* x_101; +x_100 = lean_ctor_get(x_64, 1); +lean_inc(x_100); +lean_dec(x_64); +x_101 = l_Lean_Meta_Grind_updateLastTag(x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_100); +if (lean_obj_tag(x_101) == 0) +{ +lean_object* x_102; lean_object* x_103; lean_object* x_104; +x_102 = lean_ctor_get(x_101, 1); +lean_inc(x_102); +lean_dec(x_101); +x_103 = lean_ctor_get(x_1, 5); +lean_inc(x_103); +x_104 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_103, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_102); +if (lean_obj_tag(x_104) == 0) +{ +lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; 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_105 = lean_ctor_get(x_104, 0); +lean_inc(x_105); +x_106 = lean_ctor_get(x_104, 1); +lean_inc(x_106); +lean_dec(x_104); +x_107 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___closed__2; +x_108 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_108, 0, x_107); +lean_ctor_set(x_108, 1, x_105); +x_109 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; +lean_ctor_set_tag(x_49, 7); +lean_ctor_set(x_49, 1, x_109); +lean_ctor_set(x_49, 0, x_108); +x_110 = l_Lean_indentExpr(x_57); +x_111 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_111, 0, x_49); +lean_ctor_set(x_111, 1, x_110); +x_112 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_112, 0, x_111); +lean_ctor_set(x_112, 1, x_109); +x_113 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_63, x_112, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_106); +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_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___lambda__1(x_33, x_114, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_115); +lean_dec(x_114); +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_23 = x_117; +x_24 = x_118; +goto block_31; +} +else +{ +lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; +lean_dec(x_57); +lean_free_object(x_49); +lean_dec(x_33); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_3); lean_dec(x_1); -x_48 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_47, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_42); -if (lean_obj_tag(x_48) == 0) +x_119 = lean_ctor_get(x_104, 0); +lean_inc(x_119); +x_120 = lean_ctor_get(x_104, 1); +lean_inc(x_120); +if (lean_is_exclusive(x_104)) { + lean_ctor_release(x_104, 0); + lean_ctor_release(x_104, 1); + x_121 = x_104; +} else { + lean_dec_ref(x_104); + x_121 = lean_box(0); +} +if (lean_is_scalar(x_121)) { + x_122 = lean_alloc_ctor(1, 2, 0); +} else { + x_122 = x_121; +} +lean_ctor_set(x_122, 0, x_119); +lean_ctor_set(x_122, 1, x_120); +return x_122; +} +} +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_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___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__8; -lean_ctor_set_tag(x_39, 7); -lean_ctor_set(x_39, 1, x_49); -lean_ctor_set(x_39, 0, x_51); -x_52 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; -lean_ctor_set_tag(x_30, 7); -lean_ctor_set(x_30, 1, x_52); -lean_ctor_set(x_30, 0, x_39); -x_53 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_38, x_30, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, 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); -x_56 = lean_apply_12(x_43, x_54, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_55); -return x_56; +lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; +lean_dec(x_57); +lean_free_object(x_49); +lean_dec(x_33); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_3); +lean_dec(x_1); +x_123 = lean_ctor_get(x_101, 0); +lean_inc(x_123); +x_124 = lean_ctor_get(x_101, 1); +lean_inc(x_124); +if (lean_is_exclusive(x_101)) { + lean_ctor_release(x_101, 0); + lean_ctor_release(x_101, 1); + x_125 = x_101; +} else { + lean_dec_ref(x_101); + x_125 = lean_box(0); +} +if (lean_is_scalar(x_125)) { + x_126 = lean_alloc_ctor(1, 2, 0); +} else { + x_126 = x_125; +} +lean_ctor_set(x_126, 0, x_123); +lean_ctor_set(x_126, 1, x_124); +return x_126; +} +} +} } else { -uint8_t x_57; -lean_free_object(x_39); -lean_free_object(x_30); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -x_57 = !lean_is_exclusive(x_48); -if (x_57 == 0) +lean_object* x_127; lean_object* x_128; +lean_dec(x_57); +lean_free_object(x_49); +x_127 = lean_ctor_get(x_59, 1); +lean_inc(x_127); +lean_dec(x_59); +lean_inc(x_3); +lean_ctor_set(x_8, 0, x_3); +x_128 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_128, 0, x_8); +x_23 = x_128; +x_24 = x_127; +goto block_31; +} +} +else { -return x_48; +uint8_t x_129; +lean_dec(x_57); +lean_free_object(x_49); +lean_dec(x_33); +lean_free_object(x_8); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_3); +lean_dec(x_1); +x_129 = !lean_is_exclusive(x_59); +if (x_129 == 0) +{ +return x_59; } else { -lean_object* x_58; lean_object* x_59; lean_object* x_60; -x_58 = lean_ctor_get(x_48, 0); -x_59 = lean_ctor_get(x_48, 1); -lean_inc(x_59); -lean_inc(x_58); -lean_dec(x_48); -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_130; lean_object* x_131; lean_object* x_132; +x_130 = lean_ctor_get(x_59, 0); +x_131 = lean_ctor_get(x_59, 1); +lean_inc(x_131); +lean_inc(x_130); +lean_dec(x_59); +x_132 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_132, 0, x_130); +lean_ctor_set(x_132, 1, x_131); +return x_132; +} +} +} +else +{ +uint8_t x_133; +lean_free_object(x_49); +lean_dec(x_33); +lean_free_object(x_8); +lean_dec(x_22); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_3); +lean_dec(x_1); +x_133 = !lean_is_exclusive(x_56); +if (x_133 == 0) +{ +return x_56; +} +else +{ +lean_object* x_134; lean_object* x_135; lean_object* x_136; +x_134 = lean_ctor_get(x_56, 0); +x_135 = lean_ctor_get(x_56, 1); +lean_inc(x_135); +lean_inc(x_134); +lean_dec(x_56); +x_136 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_136, 0, x_134); +lean_ctor_set(x_136, 1, x_135); +return x_136; +} +} +} +else +{ +lean_object* x_137; +lean_free_object(x_49); +lean_dec(x_22); +lean_inc(x_3); +lean_ctor_set(x_8, 0, x_3); +x_137 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_137, 0, x_8); +x_23 = x_137; +x_24 = x_52; +goto block_31; +} +} +} +else +{ +lean_object* x_138; lean_object* x_139; uint8_t x_140; +x_138 = lean_ctor_get(x_49, 0); +x_139 = lean_ctor_get(x_49, 1); +lean_inc(x_139); +lean_inc(x_138); +lean_dec(x_49); +x_140 = l_Lean_BinderInfo_isInstImplicit(x_45); +if (x_140 == 0) +{ +lean_object* x_141; +lean_dec(x_138); +lean_dec(x_22); +lean_inc(x_3); +lean_ctor_set(x_8, 0, x_3); +x_141 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_141, 0, x_8); +x_23 = x_141; +x_24 = x_139; +goto block_31; +} +else +{ +uint8_t x_142; +x_142 = lean_unbox(x_138); +lean_dec(x_138); +if (x_142 == 0) +{ +lean_object* x_143; +lean_inc(x_18); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_22); +x_143 = lean_infer_type(x_22, x_15, x_16, x_17, x_18, x_139); +if (lean_obj_tag(x_143) == 0) +{ +lean_object* x_144; lean_object* x_145; lean_object* x_146; +x_144 = lean_ctor_get(x_143, 0); +lean_inc(x_144); +x_145 = lean_ctor_get(x_143, 1); +lean_inc(x_145); +lean_dec(x_143); +lean_inc(x_18); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_144); +x_146 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem_synthesizeInstance(x_22, x_144, x_15, x_16, x_17, x_18, x_145); +if (lean_obj_tag(x_146) == 0) +{ +lean_object* x_147; uint8_t x_148; +x_147 = lean_ctor_get(x_146, 0); +lean_inc(x_147); +x_148 = lean_unbox(x_147); +lean_dec(x_147); +if (x_148 == 0) +{ +lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; uint8_t x_153; +lean_free_object(x_8); +x_149 = lean_ctor_get(x_146, 1); +lean_inc(x_149); +lean_dec(x_146); +x_150 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__2; +x_151 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__2(x_150, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_149); +x_152 = lean_ctor_get(x_151, 0); +lean_inc(x_152); +x_153 = lean_unbox(x_152); +lean_dec(x_152); +if (x_153 == 0) +{ +lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; +lean_dec(x_144); +x_154 = lean_ctor_get(x_151, 1); +lean_inc(x_154); +lean_dec(x_151); +x_155 = lean_box(0); +x_156 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___lambda__1(x_33, x_155, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_154); +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_23 = x_157; +x_24 = x_158; +goto block_31; +} +else +{ +lean_object* x_159; lean_object* x_160; lean_object* x_161; +x_159 = lean_ctor_get(x_151, 1); +lean_inc(x_159); +if (lean_is_exclusive(x_151)) { + lean_ctor_release(x_151, 0); + lean_ctor_release(x_151, 1); + x_160 = x_151; +} else { + lean_dec_ref(x_151); + x_160 = lean_box(0); +} +x_161 = l_Lean_Meta_Grind_updateLastTag(x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_159); +if (lean_obj_tag(x_161) == 0) +{ +lean_object* x_162; lean_object* x_163; lean_object* x_164; +x_162 = lean_ctor_get(x_161, 1); +lean_inc(x_162); +lean_dec(x_161); +x_163 = lean_ctor_get(x_1, 5); +lean_inc(x_163); +x_164 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_163, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_162); +if (lean_obj_tag(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; +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_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___closed__2; +if (lean_is_scalar(x_160)) { + x_168 = lean_alloc_ctor(7, 2, 0); +} else { + x_168 = x_160; + lean_ctor_set_tag(x_168, 7); +} +lean_ctor_set(x_168, 0, x_167); +lean_ctor_set(x_168, 1, x_165); +x_169 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; +x_170 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_170, 0, x_168); +lean_ctor_set(x_170, 1, x_169); +x_171 = l_Lean_indentExpr(x_144); +x_172 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_172, 0, x_170); +lean_ctor_set(x_172, 1, x_171); +x_173 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_173, 0, x_172); +lean_ctor_set(x_173, 1, x_169); +x_174 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_150, x_173, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_166); +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_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___lambda__1(x_33, x_175, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_176); +lean_dec(x_175); +x_178 = lean_ctor_get(x_177, 0); +lean_inc(x_178); +x_179 = lean_ctor_get(x_177, 1); +lean_inc(x_179); +lean_dec(x_177); +x_23 = x_178; +x_24 = x_179; +goto block_31; +} +else +{ +lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; +lean_dec(x_160); +lean_dec(x_144); +lean_dec(x_33); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_3); +lean_dec(x_1); +x_180 = lean_ctor_get(x_164, 0); +lean_inc(x_180); +x_181 = lean_ctor_get(x_164, 1); +lean_inc(x_181); +if (lean_is_exclusive(x_164)) { + lean_ctor_release(x_164, 0); + lean_ctor_release(x_164, 1); + x_182 = x_164; +} else { + lean_dec_ref(x_164); + x_182 = lean_box(0); +} +if (lean_is_scalar(x_182)) { + x_183 = lean_alloc_ctor(1, 2, 0); +} else { + x_183 = x_182; +} +lean_ctor_set(x_183, 0, x_180); +lean_ctor_set(x_183, 1, x_181); +return x_183; +} +} +else +{ +lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; +lean_dec(x_160); +lean_dec(x_144); +lean_dec(x_33); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_3); +lean_dec(x_1); +x_184 = lean_ctor_get(x_161, 0); +lean_inc(x_184); +x_185 = lean_ctor_get(x_161, 1); +lean_inc(x_185); +if (lean_is_exclusive(x_161)) { + lean_ctor_release(x_161, 0); + lean_ctor_release(x_161, 1); + x_186 = x_161; +} else { + lean_dec_ref(x_161); + x_186 = lean_box(0); +} +if (lean_is_scalar(x_186)) { + x_187 = lean_alloc_ctor(1, 2, 0); +} else { + x_187 = x_186; +} +lean_ctor_set(x_187, 0, x_184); +lean_ctor_set(x_187, 1, x_185); +return x_187; +} +} +} +else +{ +lean_object* x_188; lean_object* x_189; +lean_dec(x_144); +x_188 = lean_ctor_get(x_146, 1); +lean_inc(x_188); +lean_dec(x_146); +lean_inc(x_3); +lean_ctor_set(x_8, 0, x_3); +x_189 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_189, 0, x_8); +x_23 = x_189; +x_24 = x_188; +goto block_31; +} +} +else +{ +lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; +lean_dec(x_144); +lean_dec(x_33); +lean_free_object(x_8); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_3); +lean_dec(x_1); +x_190 = lean_ctor_get(x_146, 0); +lean_inc(x_190); +x_191 = lean_ctor_get(x_146, 1); +lean_inc(x_191); +if (lean_is_exclusive(x_146)) { + lean_ctor_release(x_146, 0); + lean_ctor_release(x_146, 1); + x_192 = x_146; +} else { + lean_dec_ref(x_146); + x_192 = lean_box(0); +} +if (lean_is_scalar(x_192)) { + x_193 = lean_alloc_ctor(1, 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; +} +} +else +{ +lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; +lean_dec(x_33); +lean_free_object(x_8); +lean_dec(x_22); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_3); +lean_dec(x_1); +x_194 = lean_ctor_get(x_143, 0); +lean_inc(x_194); +x_195 = lean_ctor_get(x_143, 1); +lean_inc(x_195); +if (lean_is_exclusive(x_143)) { + lean_ctor_release(x_143, 0); + lean_ctor_release(x_143, 1); + x_196 = x_143; +} else { + lean_dec_ref(x_143); + x_196 = lean_box(0); +} +if (lean_is_scalar(x_196)) { + x_197 = lean_alloc_ctor(1, 2, 0); +} else { + x_197 = x_196; +} +lean_ctor_set(x_197, 0, x_194); +lean_ctor_set(x_197, 1, x_195); +return x_197; +} +} +else +{ +lean_object* x_198; +lean_dec(x_22); +lean_inc(x_3); +lean_ctor_set(x_8, 0, x_3); +x_198 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_198, 0, x_8); +x_23 = x_198; +x_24 = x_139; +goto block_31; +} +} +} +} +else +{ +lean_object* x_199; uint8_t x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; uint8_t x_209; +lean_dec(x_33); +x_199 = lean_array_fget(x_35, x_36); +x_200 = lean_unbox(x_199); +lean_dec(x_199); +x_201 = lean_unsigned_to_nat(1u); +x_202 = lean_nat_add(x_36, x_201); +lean_dec(x_36); +x_203 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_203, 0, x_35); +lean_ctor_set(x_203, 1, x_202); +lean_ctor_set(x_203, 2, x_37); +x_204 = l_Lean_Expr_mvarId_x21(x_22); +x_205 = l_Lean_MVarId_isAssigned___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__4(x_204, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19); +lean_dec(x_204); +x_206 = lean_ctor_get(x_205, 0); +lean_inc(x_206); +x_207 = lean_ctor_get(x_205, 1); +lean_inc(x_207); +if (lean_is_exclusive(x_205)) { + lean_ctor_release(x_205, 0); + lean_ctor_release(x_205, 1); + x_208 = x_205; +} else { + lean_dec_ref(x_205); + x_208 = lean_box(0); +} +x_209 = l_Lean_BinderInfo_isInstImplicit(x_200); +if (x_209 == 0) +{ +lean_object* x_210; +lean_dec(x_208); +lean_dec(x_206); +lean_dec(x_22); +lean_inc(x_3); +lean_ctor_set(x_8, 1, x_203); +lean_ctor_set(x_8, 0, x_3); +x_210 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_210, 0, x_8); +x_23 = x_210; +x_24 = x_207; +goto block_31; +} +else +{ +uint8_t x_211; +x_211 = lean_unbox(x_206); +lean_dec(x_206); +if (x_211 == 0) +{ +lean_object* x_212; +lean_inc(x_18); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_22); +x_212 = lean_infer_type(x_22, x_15, x_16, x_17, x_18, x_207); +if (lean_obj_tag(x_212) == 0) +{ +lean_object* x_213; lean_object* x_214; lean_object* x_215; +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_18); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_213); +x_215 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem_synthesizeInstance(x_22, x_213, x_15, x_16, x_17, x_18, x_214); +if (lean_obj_tag(x_215) == 0) +{ +lean_object* x_216; uint8_t x_217; +x_216 = lean_ctor_get(x_215, 0); +lean_inc(x_216); +x_217 = lean_unbox(x_216); +lean_dec(x_216); +if (x_217 == 0) +{ +lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; uint8_t x_222; +lean_free_object(x_8); +x_218 = lean_ctor_get(x_215, 1); +lean_inc(x_218); +lean_dec(x_215); +x_219 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__2; +x_220 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__2(x_219, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_218); +x_221 = lean_ctor_get(x_220, 0); +lean_inc(x_221); +x_222 = lean_unbox(x_221); +lean_dec(x_221); +if (x_222 == 0) +{ +lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; +lean_dec(x_213); +lean_dec(x_208); +x_223 = lean_ctor_get(x_220, 1); +lean_inc(x_223); +lean_dec(x_220); +x_224 = lean_box(0); +x_225 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___lambda__1(x_203, x_224, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_223); +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_23 = x_226; +x_24 = x_227; +goto block_31; +} +else +{ +lean_object* x_228; lean_object* x_229; lean_object* x_230; +x_228 = lean_ctor_get(x_220, 1); +lean_inc(x_228); +if (lean_is_exclusive(x_220)) { + lean_ctor_release(x_220, 0); + lean_ctor_release(x_220, 1); + x_229 = x_220; +} else { + lean_dec_ref(x_220); + x_229 = lean_box(0); +} +x_230 = l_Lean_Meta_Grind_updateLastTag(x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_228); +if (lean_obj_tag(x_230) == 0) +{ +lean_object* x_231; lean_object* x_232; lean_object* x_233; +x_231 = lean_ctor_get(x_230, 1); +lean_inc(x_231); +lean_dec(x_230); +x_232 = lean_ctor_get(x_1, 5); +lean_inc(x_232); +x_233 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_232, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_231); +if (lean_obj_tag(x_233) == 0) +{ +lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; 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; +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 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___closed__2; +if (lean_is_scalar(x_229)) { + x_237 = lean_alloc_ctor(7, 2, 0); +} else { + x_237 = x_229; + lean_ctor_set_tag(x_237, 7); +} +lean_ctor_set(x_237, 0, x_236); +lean_ctor_set(x_237, 1, x_234); +x_238 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; +if (lean_is_scalar(x_208)) { + x_239 = lean_alloc_ctor(7, 2, 0); +} else { + x_239 = x_208; + lean_ctor_set_tag(x_239, 7); +} +lean_ctor_set(x_239, 0, x_237); +lean_ctor_set(x_239, 1, x_238); +x_240 = l_Lean_indentExpr(x_213); +x_241 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_241, 0, x_239); +lean_ctor_set(x_241, 1, x_240); +x_242 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_242, 0, x_241); +lean_ctor_set(x_242, 1, x_238); +x_243 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_219, x_242, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_235); +x_244 = lean_ctor_get(x_243, 0); +lean_inc(x_244); +x_245 = lean_ctor_get(x_243, 1); +lean_inc(x_245); +lean_dec(x_243); +x_246 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___lambda__1(x_203, x_244, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_245); +lean_dec(x_244); +x_247 = lean_ctor_get(x_246, 0); +lean_inc(x_247); +x_248 = lean_ctor_get(x_246, 1); +lean_inc(x_248); +lean_dec(x_246); +x_23 = x_247; +x_24 = x_248; +goto block_31; +} +else +{ +lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; +lean_dec(x_229); +lean_dec(x_213); +lean_dec(x_208); +lean_dec(x_203); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_3); +lean_dec(x_1); +x_249 = lean_ctor_get(x_233, 0); +lean_inc(x_249); +x_250 = lean_ctor_get(x_233, 1); +lean_inc(x_250); +if (lean_is_exclusive(x_233)) { + lean_ctor_release(x_233, 0); + lean_ctor_release(x_233, 1); + x_251 = x_233; +} else { + lean_dec_ref(x_233); + x_251 = lean_box(0); +} +if (lean_is_scalar(x_251)) { + x_252 = lean_alloc_ctor(1, 2, 0); +} else { + x_252 = x_251; +} +lean_ctor_set(x_252, 0, x_249); +lean_ctor_set(x_252, 1, x_250); +return x_252; +} +} +else +{ +lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; +lean_dec(x_229); +lean_dec(x_213); +lean_dec(x_208); +lean_dec(x_203); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_3); +lean_dec(x_1); +x_253 = lean_ctor_get(x_230, 0); +lean_inc(x_253); +x_254 = lean_ctor_get(x_230, 1); +lean_inc(x_254); +if (lean_is_exclusive(x_230)) { + lean_ctor_release(x_230, 0); + lean_ctor_release(x_230, 1); + x_255 = x_230; +} else { + lean_dec_ref(x_230); + x_255 = lean_box(0); +} +if (lean_is_scalar(x_255)) { + x_256 = lean_alloc_ctor(1, 2, 0); +} else { + x_256 = x_255; +} +lean_ctor_set(x_256, 0, x_253); +lean_ctor_set(x_256, 1, x_254); +return x_256; +} +} +} +else +{ +lean_object* x_257; lean_object* x_258; +lean_dec(x_213); +lean_dec(x_208); +x_257 = lean_ctor_get(x_215, 1); +lean_inc(x_257); +lean_dec(x_215); +lean_inc(x_3); +lean_ctor_set(x_8, 1, x_203); +lean_ctor_set(x_8, 0, x_3); +x_258 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_258, 0, x_8); +x_23 = x_258; +x_24 = x_257; +goto block_31; +} +} +else +{ +lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; +lean_dec(x_213); +lean_dec(x_208); +lean_dec(x_203); +lean_free_object(x_8); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_3); +lean_dec(x_1); +x_259 = lean_ctor_get(x_215, 0); +lean_inc(x_259); +x_260 = lean_ctor_get(x_215, 1); +lean_inc(x_260); +if (lean_is_exclusive(x_215)) { + lean_ctor_release(x_215, 0); + lean_ctor_release(x_215, 1); + x_261 = x_215; +} else { + lean_dec_ref(x_215); + x_261 = lean_box(0); +} +if (lean_is_scalar(x_261)) { + x_262 = lean_alloc_ctor(1, 2, 0); +} else { + x_262 = x_261; +} +lean_ctor_set(x_262, 0, x_259); +lean_ctor_set(x_262, 1, x_260); +return x_262; +} +} +else +{ +lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; +lean_dec(x_208); +lean_dec(x_203); +lean_free_object(x_8); +lean_dec(x_22); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_3); +lean_dec(x_1); +x_263 = lean_ctor_get(x_212, 0); +lean_inc(x_263); +x_264 = lean_ctor_get(x_212, 1); +lean_inc(x_264); +if (lean_is_exclusive(x_212)) { + lean_ctor_release(x_212, 0); + lean_ctor_release(x_212, 1); + x_265 = x_212; +} else { + lean_dec_ref(x_212); + x_265 = lean_box(0); +} +if (lean_is_scalar(x_265)) { + x_266 = lean_alloc_ctor(1, 2, 0); +} else { + x_266 = x_265; +} +lean_ctor_set(x_266, 0, x_263); +lean_ctor_set(x_266, 1, x_264); +return x_266; +} +} +else +{ +lean_object* x_267; +lean_dec(x_208); +lean_dec(x_22); +lean_inc(x_3); +lean_ctor_set(x_8, 1, x_203); +lean_ctor_set(x_8, 0, x_3); +x_267 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_267, 0, x_8); +x_23 = x_267; +x_24 = x_207; +goto block_31; +} +} +} +} +} +else +{ +lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; uint8_t x_272; +x_268 = lean_ctor_get(x_8, 1); +lean_inc(x_268); +lean_dec(x_8); +x_269 = lean_ctor_get(x_268, 0); +lean_inc(x_269); +x_270 = lean_ctor_get(x_268, 1); +lean_inc(x_270); +x_271 = lean_ctor_get(x_268, 2); +lean_inc(x_271); +x_272 = lean_nat_dec_lt(x_270, x_271); +if (x_272 == 0) +{ +lean_object* x_273; lean_object* x_274; +lean_dec(x_271); +lean_dec(x_270); +lean_dec(x_269); +lean_dec(x_22); +lean_inc(x_3); +x_273 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_273, 0, x_3); +lean_ctor_set(x_273, 1, x_268); +x_274 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_274, 0, x_273); +x_23 = x_274; +x_24 = x_19; +goto block_31; +} +else +{ +lean_object* x_275; lean_object* x_276; uint8_t x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; uint8_t x_286; +if (lean_is_exclusive(x_268)) { + lean_ctor_release(x_268, 0); + lean_ctor_release(x_268, 1); + lean_ctor_release(x_268, 2); + x_275 = x_268; +} else { + lean_dec_ref(x_268); + x_275 = lean_box(0); +} +x_276 = lean_array_fget(x_269, x_270); +x_277 = lean_unbox(x_276); +lean_dec(x_276); +x_278 = lean_unsigned_to_nat(1u); +x_279 = lean_nat_add(x_270, x_278); +lean_dec(x_270); +if (lean_is_scalar(x_275)) { + x_280 = lean_alloc_ctor(0, 3, 0); +} else { + x_280 = x_275; +} +lean_ctor_set(x_280, 0, x_269); +lean_ctor_set(x_280, 1, x_279); +lean_ctor_set(x_280, 2, x_271); +x_281 = l_Lean_Expr_mvarId_x21(x_22); +x_282 = l_Lean_MVarId_isAssigned___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__4(x_281, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19); +lean_dec(x_281); +x_283 = lean_ctor_get(x_282, 0); +lean_inc(x_283); +x_284 = lean_ctor_get(x_282, 1); +lean_inc(x_284); +if (lean_is_exclusive(x_282)) { + lean_ctor_release(x_282, 0); + lean_ctor_release(x_282, 1); + x_285 = x_282; +} else { + lean_dec_ref(x_282); + x_285 = lean_box(0); +} +x_286 = l_Lean_BinderInfo_isInstImplicit(x_277); +if (x_286 == 0) +{ +lean_object* x_287; lean_object* x_288; +lean_dec(x_285); +lean_dec(x_283); +lean_dec(x_22); +lean_inc(x_3); +x_287 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_287, 0, x_3); +lean_ctor_set(x_287, 1, x_280); +x_288 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_288, 0, x_287); +x_23 = x_288; +x_24 = x_284; +goto block_31; +} +else +{ +uint8_t x_289; +x_289 = lean_unbox(x_283); +lean_dec(x_283); +if (x_289 == 0) +{ +lean_object* x_290; +lean_inc(x_18); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_22); +x_290 = lean_infer_type(x_22, x_15, x_16, x_17, x_18, x_284); +if (lean_obj_tag(x_290) == 0) +{ +lean_object* x_291; lean_object* x_292; lean_object* x_293; +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); +lean_inc(x_18); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_291); +x_293 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem_synthesizeInstance(x_22, x_291, x_15, x_16, x_17, x_18, x_292); +if (lean_obj_tag(x_293) == 0) +{ +lean_object* x_294; uint8_t x_295; +x_294 = lean_ctor_get(x_293, 0); +lean_inc(x_294); +x_295 = lean_unbox(x_294); +lean_dec(x_294); +if (x_295 == 0) +{ +lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; uint8_t x_300; +x_296 = lean_ctor_get(x_293, 1); +lean_inc(x_296); +lean_dec(x_293); +x_297 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__2; +x_298 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__2(x_297, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_296); +x_299 = lean_ctor_get(x_298, 0); +lean_inc(x_299); +x_300 = lean_unbox(x_299); +lean_dec(x_299); +if (x_300 == 0) +{ +lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; +lean_dec(x_291); +lean_dec(x_285); +x_301 = lean_ctor_get(x_298, 1); +lean_inc(x_301); +lean_dec(x_298); +x_302 = lean_box(0); +x_303 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___lambda__1(x_280, x_302, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_301); +x_304 = lean_ctor_get(x_303, 0); +lean_inc(x_304); +x_305 = lean_ctor_get(x_303, 1); +lean_inc(x_305); +lean_dec(x_303); +x_23 = x_304; +x_24 = x_305; +goto block_31; +} +else +{ +lean_object* x_306; lean_object* x_307; lean_object* x_308; +x_306 = lean_ctor_get(x_298, 1); +lean_inc(x_306); +if (lean_is_exclusive(x_298)) { + lean_ctor_release(x_298, 0); + lean_ctor_release(x_298, 1); + x_307 = x_298; +} else { + lean_dec_ref(x_298); + x_307 = lean_box(0); +} +x_308 = l_Lean_Meta_Grind_updateLastTag(x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_306); +if (lean_obj_tag(x_308) == 0) +{ +lean_object* x_309; lean_object* x_310; lean_object* x_311; +x_309 = lean_ctor_get(x_308, 1); +lean_inc(x_309); +lean_dec(x_308); +x_310 = lean_ctor_get(x_1, 5); +lean_inc(x_310); +x_311 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_310, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_309); +if (lean_obj_tag(x_311) == 0) +{ +lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; +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_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___closed__2; +if (lean_is_scalar(x_307)) { + x_315 = lean_alloc_ctor(7, 2, 0); +} else { + x_315 = x_307; + lean_ctor_set_tag(x_315, 7); +} +lean_ctor_set(x_315, 0, x_314); +lean_ctor_set(x_315, 1, x_312); +x_316 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; +if (lean_is_scalar(x_285)) { + x_317 = lean_alloc_ctor(7, 2, 0); +} else { + x_317 = x_285; + lean_ctor_set_tag(x_317, 7); +} +lean_ctor_set(x_317, 0, x_315); +lean_ctor_set(x_317, 1, x_316); +x_318 = l_Lean_indentExpr(x_291); +x_319 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_319, 0, x_317); +lean_ctor_set(x_319, 1, x_318); +x_320 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_320, 0, x_319); +lean_ctor_set(x_320, 1, x_316); +x_321 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_297, x_320, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_313); +x_322 = lean_ctor_get(x_321, 0); +lean_inc(x_322); +x_323 = lean_ctor_get(x_321, 1); +lean_inc(x_323); +lean_dec(x_321); +x_324 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___lambda__1(x_280, x_322, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_323); +lean_dec(x_322); +x_325 = lean_ctor_get(x_324, 0); +lean_inc(x_325); +x_326 = lean_ctor_get(x_324, 1); +lean_inc(x_326); +lean_dec(x_324); +x_23 = x_325; +x_24 = x_326; +goto block_31; +} +else +{ +lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; +lean_dec(x_307); +lean_dec(x_291); +lean_dec(x_285); +lean_dec(x_280); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_3); +lean_dec(x_1); +x_327 = lean_ctor_get(x_311, 0); +lean_inc(x_327); +x_328 = lean_ctor_get(x_311, 1); +lean_inc(x_328); +if (lean_is_exclusive(x_311)) { + lean_ctor_release(x_311, 0); + lean_ctor_release(x_311, 1); + x_329 = x_311; +} else { + lean_dec_ref(x_311); + x_329 = lean_box(0); +} +if (lean_is_scalar(x_329)) { + x_330 = lean_alloc_ctor(1, 2, 0); +} else { + x_330 = x_329; +} +lean_ctor_set(x_330, 0, x_327); +lean_ctor_set(x_330, 1, x_328); +return x_330; +} +} +else +{ +lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; +lean_dec(x_307); +lean_dec(x_291); +lean_dec(x_285); +lean_dec(x_280); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_3); +lean_dec(x_1); +x_331 = lean_ctor_get(x_308, 0); +lean_inc(x_331); +x_332 = lean_ctor_get(x_308, 1); +lean_inc(x_332); +if (lean_is_exclusive(x_308)) { + lean_ctor_release(x_308, 0); + lean_ctor_release(x_308, 1); + x_333 = x_308; +} else { + lean_dec_ref(x_308); + x_333 = lean_box(0); +} +if (lean_is_scalar(x_333)) { + x_334 = lean_alloc_ctor(1, 2, 0); +} else { + x_334 = 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_335; lean_object* x_336; lean_object* x_337; +lean_dec(x_291); +lean_dec(x_285); +x_335 = lean_ctor_get(x_293, 1); +lean_inc(x_335); +lean_dec(x_293); +lean_inc(x_3); +x_336 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_336, 0, x_3); +lean_ctor_set(x_336, 1, x_280); +x_337 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_337, 0, x_336); +x_23 = x_337; +x_24 = x_335; +goto block_31; +} +} +else +{ +lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; +lean_dec(x_291); +lean_dec(x_285); +lean_dec(x_280); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_3); +lean_dec(x_1); +x_338 = lean_ctor_get(x_293, 0); +lean_inc(x_338); +x_339 = lean_ctor_get(x_293, 1); +lean_inc(x_339); +if (lean_is_exclusive(x_293)) { + lean_ctor_release(x_293, 0); + lean_ctor_release(x_293, 1); + x_340 = x_293; +} else { + lean_dec_ref(x_293); + x_340 = lean_box(0); +} +if (lean_is_scalar(x_340)) { + x_341 = lean_alloc_ctor(1, 2, 0); +} else { + x_341 = x_340; +} +lean_ctor_set(x_341, 0, x_338); +lean_ctor_set(x_341, 1, x_339); +return x_341; +} +} +else +{ +lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; +lean_dec(x_285); +lean_dec(x_280); +lean_dec(x_22); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_3); +lean_dec(x_1); +x_342 = lean_ctor_get(x_290, 0); +lean_inc(x_342); +x_343 = lean_ctor_get(x_290, 1); +lean_inc(x_343); +if (lean_is_exclusive(x_290)) { + lean_ctor_release(x_290, 0); + lean_ctor_release(x_290, 1); + x_344 = x_290; +} else { + lean_dec_ref(x_290); + x_344 = lean_box(0); +} +if (lean_is_scalar(x_344)) { + x_345 = lean_alloc_ctor(1, 2, 0); +} else { + x_345 = x_344; +} +lean_ctor_set(x_345, 0, x_342); +lean_ctor_set(x_345, 1, x_343); +return x_345; +} +} +else +{ +lean_object* x_346; lean_object* x_347; +lean_dec(x_285); +lean_dec(x_22); +lean_inc(x_3); +x_346 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_346, 0, x_3); +lean_ctor_set(x_346, 1, x_280); +x_347 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_347, 0, x_346); +x_23 = x_347; +x_24 = x_284; +goto block_31; +} +} +} +} +block_31: +{ +if (lean_obj_tag(x_23) == 0) +{ +lean_object* x_25; lean_object* x_26; +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_3); +lean_dec(x_1); +x_25 = lean_ctor_get(x_23, 0); +lean_inc(x_25); +lean_dec(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_24); +return x_26; +} +else +{ +lean_object* x_27; size_t x_28; size_t x_29; +x_27 = lean_ctor_get(x_23, 0); +lean_inc(x_27); +lean_dec(x_23); +x_28 = 1; +x_29 = lean_usize_add(x_7, x_28); +x_7 = x_29; +x_8 = x_27; +x_19 = x_24; +goto _start; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, size_t x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18) { +_start: +{ +uint8_t x_19; +x_19 = lean_usize_dec_lt(x_6, x_5); +if (x_19 == 0) +{ +lean_object* x_20; +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_3); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_7); +lean_ctor_set(x_20, 1, x_18); +return x_20; +} +else +{ +lean_object* x_21; lean_object* x_22; +lean_dec(x_7); +x_21 = lean_array_uget(x_4, x_6); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_21); +x_22 = l_Lean_Meta_isProof(x_21, x_14, x_15, x_16, x_17, x_18); +if (lean_obj_tag(x_22) == 0) +{ +lean_object* x_23; uint8_t x_24; +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +x_24 = lean_unbox(x_23); +lean_dec(x_23); +if (x_24 == 0) +{ +uint8_t x_25; +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_3); +x_25 = !lean_is_exclusive(x_22); +if (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_22, 0); +lean_dec(x_26); +x_27 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_27, 0, x_21); +x_28 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_28, 0, x_27); +x_29 = lean_box(0); +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_28); +lean_ctor_set(x_30, 1, x_29); +lean_ctor_set(x_22, 0, x_30); +return x_22; +} +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_31 = lean_ctor_get(x_22, 1); +lean_inc(x_31); +lean_dec(x_22); +x_32 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_32, 0, x_21); +x_33 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_33, 0, x_32); +x_34 = lean_box(0); +x_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_33); +lean_ctor_set(x_35, 1, x_34); +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set(x_36, 1, x_31); +return x_36; +} +} +else +{ +lean_object* x_37; size_t x_38; size_t x_39; +lean_dec(x_21); +x_37 = lean_ctor_get(x_22, 1); +lean_inc(x_37); +lean_dec(x_22); +x_38 = 1; +x_39 = lean_usize_add(x_6, x_38); +lean_inc(x_3); +{ +size_t _tmp_5 = x_39; +lean_object* _tmp_6 = x_3; +lean_object* _tmp_17 = x_37; +x_6 = _tmp_5; +x_7 = _tmp_6; +x_18 = _tmp_17; +} +goto _start; +} +} +else +{ +uint8_t x_41; +lean_dec(x_21); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_3); +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_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, size_t x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18) { +_start: +{ +uint8_t x_19; +x_19 = lean_usize_dec_lt(x_6, x_5); +if (x_19 == 0) +{ +lean_object* x_20; +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_3); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_7); +lean_ctor_set(x_20, 1, x_18); +return x_20; +} +else +{ +lean_object* x_21; lean_object* x_22; +lean_dec(x_7); +x_21 = lean_array_uget(x_4, x_6); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_21); +x_22 = l_Lean_Meta_isProof(x_21, x_14, x_15, x_16, x_17, x_18); +if (lean_obj_tag(x_22) == 0) +{ +lean_object* x_23; uint8_t x_24; +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +x_24 = lean_unbox(x_23); +lean_dec(x_23); +if (x_24 == 0) +{ +uint8_t x_25; +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_3); +x_25 = !lean_is_exclusive(x_22); +if (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_22, 0); +lean_dec(x_26); +x_27 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_27, 0, x_21); +x_28 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_28, 0, x_27); +x_29 = lean_box(0); +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_28); +lean_ctor_set(x_30, 1, x_29); +lean_ctor_set(x_22, 0, x_30); +return x_22; +} +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_31 = lean_ctor_get(x_22, 1); +lean_inc(x_31); +lean_dec(x_22); +x_32 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_32, 0, x_21); +x_33 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_33, 0, x_32); +x_34 = lean_box(0); +x_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_33); +lean_ctor_set(x_35, 1, x_34); +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set(x_36, 1, x_31); +return x_36; +} +} +else +{ +lean_object* x_37; size_t x_38; size_t x_39; +lean_dec(x_21); +x_37 = lean_ctor_get(x_22, 1); +lean_inc(x_37); +lean_dec(x_22); +x_38 = 1; +x_39 = lean_usize_add(x_6, x_38); +lean_inc(x_3); +{ +size_t _tmp_5 = x_39; +lean_object* _tmp_6 = x_3; +lean_object* _tmp_17 = x_37; +x_6 = _tmp_5; +x_7 = _tmp_6; +x_18 = _tmp_17; +} +goto _start; +} +} +else +{ +uint8_t x_41; +lean_dec(x_21); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_3); +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_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__8(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +_start: +{ +uint8_t x_16; +x_16 = lean_usize_dec_eq(x_2, x_3); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_17 = lean_array_uget(x_1, x_2); +x_18 = l_Lean_Expr_mvarId_x21(x_17); +x_19 = l_Lean_MVarId_isAssigned___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__4(x_18, 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_18); +x_20 = lean_ctor_get(x_19, 0); +lean_inc(x_20); +x_21 = lean_unbox(x_20); +lean_dec(x_20); +if (x_21 == 0) +{ +lean_object* x_22; lean_object* x_23; size_t x_24; size_t x_25; +x_22 = lean_ctor_get(x_19, 1); +lean_inc(x_22); +lean_dec(x_19); +x_23 = lean_array_push(x_4, x_17); +x_24 = 1; +x_25 = lean_usize_add(x_2, x_24); +x_2 = x_25; +x_4 = x_23; +x_15 = x_22; +goto _start; +} +else +{ +lean_object* x_27; size_t x_28; size_t x_29; +lean_dec(x_17); +x_27 = lean_ctor_get(x_19, 1); +lean_inc(x_27); +lean_dec(x_19); +x_28 = 1; +x_29 = lean_usize_add(x_2, x_28); +x_2 = x_29; +x_15 = x_27; +goto _start; +} +} +else +{ +lean_object* x_31; +x_31 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_31, 0, x_4); +lean_ctor_set(x_31, 1, x_15); +return x_31; +} +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18) { +_start: +{ +uint8_t x_19; +x_19 = lean_usize_dec_lt(x_6, x_5); +if (x_19 == 0) +{ +lean_object* x_20; +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_3); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_7); +lean_ctor_set(x_20, 1, x_18); +return x_20; +} +else +{ +lean_object* x_21; lean_object* x_22; +lean_dec(x_7); +x_21 = lean_array_uget(x_4, x_6); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_21); +x_22 = l_Lean_Meta_isProof(x_21, x_14, x_15, x_16, x_17, x_18); +if (lean_obj_tag(x_22) == 0) +{ +lean_object* x_23; uint8_t x_24; +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +x_24 = lean_unbox(x_23); +lean_dec(x_23); +if (x_24 == 0) +{ +uint8_t x_25; +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_3); +x_25 = !lean_is_exclusive(x_22); +if (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_22, 0); +lean_dec(x_26); +x_27 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_27, 0, x_21); +x_28 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_28, 0, x_27); +x_29 = lean_box(0); +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_28); +lean_ctor_set(x_30, 1, x_29); +lean_ctor_set(x_22, 0, x_30); +return x_22; +} +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_31 = lean_ctor_get(x_22, 1); +lean_inc(x_31); +lean_dec(x_22); +x_32 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_32, 0, x_21); +x_33 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_33, 0, x_32); +x_34 = lean_box(0); +x_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_33); +lean_ctor_set(x_35, 1, x_34); +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set(x_36, 1, x_31); +return x_36; +} +} +else +{ +lean_object* x_37; size_t x_38; size_t x_39; +lean_dec(x_21); +x_37 = lean_ctor_get(x_22, 1); +lean_inc(x_37); +lean_dec(x_22); +x_38 = 1; +x_39 = lean_usize_add(x_6, x_38); +lean_inc(x_3); +{ +size_t _tmp_5 = x_39; +lean_object* _tmp_6 = x_3; +lean_object* _tmp_17 = x_37; +x_6 = _tmp_5; +x_7 = _tmp_6; +x_18 = _tmp_17; +} +goto _start; +} +} +else +{ +uint8_t x_41; +lean_dec(x_21); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_3); +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_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +uint8_t x_15; +x_15 = lean_usize_dec_eq(x_2, x_3); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_16 = lean_array_uget(x_1, x_2); +x_17 = l_Lean_Expr_mvarId_x21(x_16); +lean_dec(x_16); +x_18 = l_Lean_MVarId_isAssigned___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__4(x_17, 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_17); +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_unbox(x_19); +lean_dec(x_19); +if (x_20 == 0) +{ +uint8_t x_21; +x_21 = !lean_is_exclusive(x_18); +if (x_21 == 0) +{ +lean_object* x_22; uint8_t x_23; lean_object* x_24; +x_22 = lean_ctor_get(x_18, 0); +lean_dec(x_22); +x_23 = 1; +x_24 = lean_box(x_23); +lean_ctor_set(x_18, 0, x_24); +return x_18; +} +else +{ +lean_object* x_25; uint8_t x_26; lean_object* x_27; lean_object* x_28; +x_25 = lean_ctor_get(x_18, 1); +lean_inc(x_25); +lean_dec(x_18); +x_26 = 1; +x_27 = lean_box(x_26); +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_27); +lean_ctor_set(x_28, 1, x_25); +return x_28; +} +} +else +{ +lean_object* x_29; size_t x_30; size_t x_31; +x_29 = lean_ctor_get(x_18, 1); +lean_inc(x_29); +lean_dec(x_18); +x_30 = 1; +x_31 = lean_usize_add(x_2, x_30); +x_2 = x_31; +x_14 = x_29; +goto _start; +} +} +else +{ +uint8_t x_33; lean_object* x_34; lean_object* x_35; +x_33 = 0; +x_34 = lean_box(x_33); +x_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_34); +lean_ctor_set(x_35, 1, x_14); +return x_35; +} +} +} +LEAN_EXPORT lean_object* l_ReaderT_bind___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__11___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +lean_object* x_14; +lean_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); +lean_inc(x_3); +x_14 = lean_apply_11(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +x_16 = lean_ctor_get(x_14, 1); +lean_inc(x_16); +lean_dec(x_14); +x_17 = lean_apply_12(x_2, x_15, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_16); +return x_17; +} +else +{ +uint8_t x_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_2); +x_18 = !lean_is_exclusive(x_14); +if (x_18 == 0) +{ +return x_14; +} +else +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_14, 0); +x_20 = lean_ctor_get(x_14, 1); +lean_inc(x_20); +lean_inc(x_19); +lean_dec(x_14); +x_21 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_21, 0, x_19); +lean_ctor_set(x_21, 1, x_20); +return x_21; +} +} +} +} +LEAN_EXPORT lean_object* l_ReaderT_bind___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__11(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_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__11___rarg), 13, 0); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_withNewMCtxDepth___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__12___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_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +lean_object* x_14; lean_object* x_15; +x_14 = lean_apply_6(x_1, x_3, x_4, x_5, x_6, x_7, x_8); +x_15 = l___private_Lean_Meta_Basic_0__Lean_Meta_withNewMCtxDepthImp___rarg(x_2, x_14, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_15) == 0) +{ +uint8_t x_16; +x_16 = !lean_is_exclusive(x_15); +if (x_16 == 0) +{ +return x_15; +} +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_17 = lean_ctor_get(x_15, 0); +x_18 = lean_ctor_get(x_15, 1); +lean_inc(x_18); +lean_inc(x_17); +lean_dec(x_15); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_17); +lean_ctor_set(x_19, 1, x_18); +return x_19; +} +} +else +{ +uint8_t x_20; +x_20 = !lean_is_exclusive(x_15); +if (x_20 == 0) +{ +return x_15; +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_15, 0); +x_22 = lean_ctor_get(x_15, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_15); +x_23 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_23, 0, x_21); +lean_ctor_set(x_23, 1, x_22); +return x_23; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_withNewMCtxDepth___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__12(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_Meta_withNewMCtxDepth___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__12___rarg___boxed), 13, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +_start: +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; uint8_t x_21; uint8_t x_22; lean_object* x_23; +x_17 = l_Lean_instantiateMVars___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__1(x_1, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +x_18 = lean_ctor_get(x_17, 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 = 1; +x_22 = 0; +x_23 = l_Lean_Meta_mkLambdaFVars(x_2, x_18, x_20, x_21, x_20, x_22, x_12, x_13, x_14, x_15, x_19); +if (lean_obj_tag(x_23) == 0) +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +x_25 = lean_ctor_get(x_23, 1); +lean_inc(x_25); +lean_dec(x_23); +x_26 = lean_ctor_get(x_3, 5); +lean_inc(x_26); +lean_dec(x_3); +x_27 = lean_ctor_get(x_4, 1); +x_28 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance(x_26, x_24, x_27, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_25); +return x_28; +} +else +{ +uint8_t x_29; +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_3); +x_29 = !lean_is_exclusive(x_23); +if (x_29 == 0) +{ +return x_23; +} +else +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_30 = lean_ctor_get(x_23, 0); +x_31 = lean_ctor_get(x_23, 1); +lean_inc(x_31); +lean_inc(x_30); +lean_dec(x_23); +x_32 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_32, 0, x_30); +lean_ctor_set(x_32, 1, x_31); +return x_32; +} +} +} +} +static size_t _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__1() { +_start: +{ +lean_object* x_1; size_t x_2; +x_1 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2; +x_2 = lean_array_size(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("failed to instantiate ", 22, 22); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__2; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(", failed to instantiate non propositional argument with type", 60, 60); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__4; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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, 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, lean_object* x_19, lean_object* x_20, lean_object* x_21) { +_start: +{ +lean_object* x_22; lean_object* x_23; uint8_t x_24; uint8_t x_25; lean_object* x_26; +x_22 = l_Lean_mkAppN(x_1, x_2); +x_23 = lean_unsigned_to_nat(0u); +x_24 = lean_nat_dec_lt(x_23, x_3); +if (x_24 == 0) +{ +uint8_t x_327; +x_327 = 1; +x_25 = x_327; +x_26 = x_21; +goto block_326; +} +else +{ +size_t x_328; lean_object* x_329; lean_object* x_330; uint8_t x_331; +x_328 = lean_usize_of_nat(x_3); +x_329 = l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__10(x_2, x_8, x_328, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21); +x_330 = lean_ctor_get(x_329, 0); +lean_inc(x_330); +x_331 = lean_unbox(x_330); +lean_dec(x_330); +if (x_331 == 0) +{ +lean_object* x_332; uint8_t x_333; +x_332 = lean_ctor_get(x_329, 1); +lean_inc(x_332); +lean_dec(x_329); +x_333 = 1; +x_25 = x_333; +x_26 = x_332; +goto block_326; +} +else +{ +lean_object* x_334; uint8_t x_335; +x_334 = lean_ctor_get(x_329, 1); +lean_inc(x_334); +lean_dec(x_329); +x_335 = 0; +x_25 = x_335; +x_26 = x_334; +goto block_326; +} +} +block_326: +{ +if (x_25 == 0) +{ +if (x_24 == 0) +{ +lean_object* x_27; lean_object* x_28; lean_object* x_113; size_t x_114; lean_object* x_115; +x_113 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2; +x_114 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__1; +lean_inc(x_20); +lean_inc(x_19); +lean_inc(x_18); +lean_inc(x_17); +lean_inc(x_7); +x_115 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__6(x_6, x_113, x_7, x_113, x_114, x_8, x_7, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_26); +if (lean_obj_tag(x_115) == 0) +{ +lean_object* x_116; lean_object* x_117; +x_116 = lean_ctor_get(x_115, 0); +lean_inc(x_116); +x_117 = lean_ctor_get(x_116, 0); +lean_inc(x_117); +lean_dec(x_116); +if (lean_obj_tag(x_117) == 0) +{ +lean_object* x_118; +x_118 = lean_ctor_get(x_115, 1); +lean_inc(x_118); +lean_dec(x_115); +x_27 = x_9; +x_28 = x_118; +goto block_112; +} +else +{ +lean_object* x_119; lean_object* x_120; +lean_dec(x_9); +x_119 = lean_ctor_get(x_115, 1); +lean_inc(x_119); +lean_dec(x_115); +x_120 = lean_ctor_get(x_117, 0); +lean_inc(x_120); +lean_dec(x_117); +x_27 = x_120; +x_28 = x_119; +goto block_112; +} +} +else +{ +uint8_t x_121; +lean_dec(x_22); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_9); +lean_dec(x_4); +x_121 = !lean_is_exclusive(x_115); +if (x_121 == 0) +{ +return x_115; +} +else +{ +lean_object* x_122; lean_object* x_123; lean_object* x_124; +x_122 = lean_ctor_get(x_115, 0); +x_123 = lean_ctor_get(x_115, 1); +lean_inc(x_123); +lean_inc(x_122); +lean_dec(x_115); +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; +} +} +block_112: +{ +if (lean_obj_tag(x_27) == 0) +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_29 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2; +x_30 = lean_box(0); +x_31 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__1(x_22, x_29, x_4, x_5, x_30, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_28); +return x_31; +} +else +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; +x_32 = lean_ctor_get(x_27, 0); +lean_inc(x_32); +lean_dec(x_27); +x_33 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__2; +x_34 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__2(x_33, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_28); +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_object* x_40; +lean_dec(x_32); +x_37 = lean_ctor_get(x_34, 1); +lean_inc(x_37); +lean_dec(x_34); +x_38 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2; +x_39 = lean_box(0); +x_40 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__1(x_22, x_38, x_4, x_5, x_39, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_37); +return x_40; +} +else +{ +uint8_t x_41; +x_41 = !lean_is_exclusive(x_34); +if (x_41 == 0) +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_42 = lean_ctor_get(x_34, 1); +x_43 = lean_ctor_get(x_34, 0); +lean_dec(x_43); +x_44 = l_Lean_Meta_Grind_updateLastTag(x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_42); +if (lean_obj_tag(x_44) == 0) +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_45 = lean_ctor_get(x_44, 1); +lean_inc(x_45); +lean_dec(x_44); +x_46 = lean_ctor_get(x_4, 5); +lean_inc(x_46); +x_47 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_46, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_45); +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, 0); +lean_inc(x_48); +x_49 = lean_ctor_get(x_47, 1); +lean_inc(x_49); +lean_dec(x_47); +lean_inc(x_20); +lean_inc(x_19); +lean_inc(x_18); +lean_inc(x_17); +x_50 = lean_infer_type(x_32, x_17, x_18, x_19, x_20, x_49); +if (lean_obj_tag(x_50) == 0) +{ +lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; +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___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__3; +lean_ctor_set_tag(x_34, 7); +lean_ctor_set(x_34, 1, x_48); +lean_ctor_set(x_34, 0, x_53); +x_54 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__5; +x_55 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_55, 0, x_34); +lean_ctor_set(x_55, 1, x_54); +x_56 = l_Lean_indentExpr(x_51); +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_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; +x_59 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_59, 0, x_57); +lean_ctor_set(x_59, 1, x_58); +x_60 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_33, x_59, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_52); +x_61 = lean_ctor_get(x_60, 0); +lean_inc(x_61); +x_62 = lean_ctor_get(x_60, 1); +lean_inc(x_62); +lean_dec(x_60); +x_63 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2; +x_64 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__1(x_22, x_63, x_4, x_5, x_61, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_62); +lean_dec(x_61); +return x_64; +} +else +{ +uint8_t x_65; +lean_dec(x_48); +lean_free_object(x_34); +lean_dec(x_22); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_4); +x_65 = !lean_is_exclusive(x_50); +if (x_65 == 0) +{ +return x_50; +} +else +{ +lean_object* x_66; lean_object* x_67; lean_object* x_68; +x_66 = lean_ctor_get(x_50, 0); +x_67 = lean_ctor_get(x_50, 1); +lean_inc(x_67); +lean_inc(x_66); +lean_dec(x_50); +x_68 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_68, 0, x_66); +lean_ctor_set(x_68, 1, x_67); +return x_68; +} +} +} +else +{ +uint8_t x_69; +lean_free_object(x_34); +lean_dec(x_32); +lean_dec(x_22); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_4); +x_69 = !lean_is_exclusive(x_47); +if (x_69 == 0) +{ +return x_47; +} +else +{ +lean_object* x_70; lean_object* x_71; lean_object* x_72; +x_70 = lean_ctor_get(x_47, 0); +x_71 = lean_ctor_get(x_47, 1); +lean_inc(x_71); +lean_inc(x_70); +lean_dec(x_47); +x_72 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_72, 0, x_70); +lean_ctor_set(x_72, 1, x_71); +return x_72; +} +} +} +else +{ +uint8_t x_73; +lean_free_object(x_34); +lean_dec(x_32); +lean_dec(x_22); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_4); +x_73 = !lean_is_exclusive(x_44); +if (x_73 == 0) +{ +return x_44; +} +else +{ +lean_object* x_74; lean_object* x_75; lean_object* x_76; +x_74 = lean_ctor_get(x_44, 0); +x_75 = lean_ctor_get(x_44, 1); +lean_inc(x_75); +lean_inc(x_74); +lean_dec(x_44); +x_76 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_76, 0, x_74); +lean_ctor_set(x_76, 1, x_75); +return x_76; +} +} +} +else +{ +lean_object* x_77; lean_object* x_78; +x_77 = lean_ctor_get(x_34, 1); +lean_inc(x_77); +lean_dec(x_34); +x_78 = l_Lean_Meta_Grind_updateLastTag(x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_77); +if (lean_obj_tag(x_78) == 0) +{ +lean_object* x_79; lean_object* x_80; lean_object* x_81; +x_79 = lean_ctor_get(x_78, 1); +lean_inc(x_79); +lean_dec(x_78); +x_80 = lean_ctor_get(x_4, 5); +lean_inc(x_80); +x_81 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_80, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_79); +if (lean_obj_tag(x_81) == 0) +{ +lean_object* x_82; lean_object* x_83; lean_object* x_84; +x_82 = lean_ctor_get(x_81, 0); +lean_inc(x_82); +x_83 = lean_ctor_get(x_81, 1); +lean_inc(x_83); +lean_dec(x_81); +lean_inc(x_20); +lean_inc(x_19); +lean_inc(x_18); +lean_inc(x_17); +x_84 = lean_infer_type(x_32, x_17, x_18, x_19, x_20, x_83); +if (lean_obj_tag(x_84) == 0) +{ +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_85 = lean_ctor_get(x_84, 0); +lean_inc(x_85); +x_86 = lean_ctor_get(x_84, 1); +lean_inc(x_86); +lean_dec(x_84); +x_87 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__3; +x_88 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_88, 0, x_87); +lean_ctor_set(x_88, 1, x_82); +x_89 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__5; +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_indentExpr(x_85); +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_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; +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_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_33, x_94, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_86); +x_96 = lean_ctor_get(x_95, 0); +lean_inc(x_96); +x_97 = lean_ctor_get(x_95, 1); +lean_inc(x_97); +lean_dec(x_95); +x_98 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2; +x_99 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__1(x_22, x_98, x_4, x_5, x_96, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_97); +lean_dec(x_96); +return x_99; +} +else +{ +lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; +lean_dec(x_82); +lean_dec(x_22); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_4); +x_100 = lean_ctor_get(x_84, 0); +lean_inc(x_100); +x_101 = lean_ctor_get(x_84, 1); +lean_inc(x_101); +if (lean_is_exclusive(x_84)) { + lean_ctor_release(x_84, 0); + lean_ctor_release(x_84, 1); + x_102 = x_84; +} else { + lean_dec_ref(x_84); + 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(x_103, 0, x_100); +lean_ctor_set(x_103, 1, x_101); +return x_103; +} +} +else +{ +lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; +lean_dec(x_32); +lean_dec(x_22); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_4); +x_104 = lean_ctor_get(x_81, 0); +lean_inc(x_104); +x_105 = lean_ctor_get(x_81, 1); +lean_inc(x_105); +if (lean_is_exclusive(x_81)) { + lean_ctor_release(x_81, 0); + lean_ctor_release(x_81, 1); + x_106 = x_81; +} else { + lean_dec_ref(x_81); + 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; +} +} +else +{ +lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; +lean_dec(x_32); +lean_dec(x_22); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_4); +x_108 = lean_ctor_get(x_78, 0); +lean_inc(x_108); +x_109 = lean_ctor_get(x_78, 1); +lean_inc(x_109); +if (lean_is_exclusive(x_78)) { + lean_ctor_release(x_78, 0); + lean_ctor_release(x_78, 1); + x_110 = x_78; +} else { + lean_dec_ref(x_78); + x_110 = lean_box(0); +} +if (lean_is_scalar(x_110)) { + x_111 = lean_alloc_ctor(1, 2, 0); +} else { + x_111 = x_110; +} +lean_ctor_set(x_111, 0, x_108); +lean_ctor_set(x_111, 1, x_109); +return x_111; +} +} +} +} +} +} +else +{ +uint8_t x_125; +x_125 = lean_nat_dec_le(x_3, x_3); +if (x_125 == 0) +{ +lean_object* x_126; lean_object* x_127; lean_object* x_212; size_t x_213; lean_object* x_214; +x_212 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2; +x_213 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__1; +lean_inc(x_20); +lean_inc(x_19); +lean_inc(x_18); +lean_inc(x_17); +lean_inc(x_7); +x_214 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__7(x_6, x_212, x_7, x_212, x_213, x_8, x_7, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_26); +if (lean_obj_tag(x_214) == 0) +{ +lean_object* x_215; lean_object* x_216; +x_215 = lean_ctor_get(x_214, 0); +lean_inc(x_215); +x_216 = lean_ctor_get(x_215, 0); +lean_inc(x_216); +lean_dec(x_215); +if (lean_obj_tag(x_216) == 0) +{ +lean_object* x_217; +x_217 = lean_ctor_get(x_214, 1); +lean_inc(x_217); +lean_dec(x_214); +x_126 = x_9; +x_127 = x_217; +goto block_211; +} +else +{ +lean_object* x_218; lean_object* x_219; +lean_dec(x_9); +x_218 = lean_ctor_get(x_214, 1); +lean_inc(x_218); +lean_dec(x_214); +x_219 = lean_ctor_get(x_216, 0); +lean_inc(x_219); +lean_dec(x_216); +x_126 = x_219; +x_127 = x_218; +goto block_211; +} +} +else +{ +uint8_t x_220; +lean_dec(x_22); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_9); +lean_dec(x_4); +x_220 = !lean_is_exclusive(x_214); +if (x_220 == 0) +{ +return x_214; +} +else +{ +lean_object* x_221; lean_object* x_222; lean_object* x_223; +x_221 = lean_ctor_get(x_214, 0); +x_222 = lean_ctor_get(x_214, 1); +lean_inc(x_222); +lean_inc(x_221); +lean_dec(x_214); +x_223 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_223, 0, x_221); +lean_ctor_set(x_223, 1, x_222); +return x_223; +} +} +block_211: +{ +if (lean_obj_tag(x_126) == 0) +{ +lean_object* x_128; lean_object* x_129; lean_object* x_130; +x_128 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2; +x_129 = lean_box(0); +x_130 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__1(x_22, x_128, x_4, x_5, x_129, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_127); +return x_130; +} +else +{ +lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; uint8_t x_135; +x_131 = lean_ctor_get(x_126, 0); +lean_inc(x_131); +lean_dec(x_126); +x_132 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__2; +x_133 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__2(x_132, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_127); +x_134 = lean_ctor_get(x_133, 0); +lean_inc(x_134); +x_135 = lean_unbox(x_134); +lean_dec(x_134); +if (x_135 == 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_133, 1); +lean_inc(x_136); +lean_dec(x_133); +x_137 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2; +x_138 = lean_box(0); +x_139 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__1(x_22, x_137, x_4, x_5, x_138, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_136); +return x_139; +} +else +{ +uint8_t x_140; +x_140 = !lean_is_exclusive(x_133); +if (x_140 == 0) +{ +lean_object* x_141; lean_object* x_142; lean_object* x_143; +x_141 = lean_ctor_get(x_133, 1); +x_142 = lean_ctor_get(x_133, 0); +lean_dec(x_142); +x_143 = l_Lean_Meta_Grind_updateLastTag(x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_141); +if (lean_obj_tag(x_143) == 0) +{ +lean_object* x_144; lean_object* x_145; lean_object* x_146; +x_144 = lean_ctor_get(x_143, 1); +lean_inc(x_144); +lean_dec(x_143); +x_145 = lean_ctor_get(x_4, 5); +lean_inc(x_145); +x_146 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_145, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_144); +if (lean_obj_tag(x_146) == 0) +{ +lean_object* x_147; lean_object* x_148; lean_object* x_149; +x_147 = lean_ctor_get(x_146, 0); +lean_inc(x_147); +x_148 = lean_ctor_get(x_146, 1); +lean_inc(x_148); +lean_dec(x_146); +lean_inc(x_20); +lean_inc(x_19); +lean_inc(x_18); +lean_inc(x_17); +x_149 = lean_infer_type(x_131, x_17, x_18, x_19, x_20, x_148); +if (lean_obj_tag(x_149) == 0) +{ +lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; +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_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__3; +lean_ctor_set_tag(x_133, 7); +lean_ctor_set(x_133, 1, x_147); +lean_ctor_set(x_133, 0, x_152); +x_153 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__5; +x_154 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_154, 0, x_133); +lean_ctor_set(x_154, 1, x_153); +x_155 = l_Lean_indentExpr(x_150); +x_156 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_156, 0, x_154); +lean_ctor_set(x_156, 1, x_155); +x_157 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; +x_158 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_158, 0, x_156); +lean_ctor_set(x_158, 1, x_157); +x_159 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_132, x_158, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_151); +x_160 = lean_ctor_get(x_159, 0); +lean_inc(x_160); +x_161 = lean_ctor_get(x_159, 1); +lean_inc(x_161); +lean_dec(x_159); +x_162 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2; +x_163 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__1(x_22, x_162, x_4, x_5, x_160, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_161); +lean_dec(x_160); +return x_163; +} +else +{ +uint8_t x_164; +lean_dec(x_147); +lean_free_object(x_133); +lean_dec(x_22); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_4); +x_164 = !lean_is_exclusive(x_149); +if (x_164 == 0) +{ +return x_149; +} +else +{ +lean_object* x_165; lean_object* x_166; lean_object* x_167; +x_165 = lean_ctor_get(x_149, 0); +x_166 = lean_ctor_get(x_149, 1); +lean_inc(x_166); +lean_inc(x_165); +lean_dec(x_149); +x_167 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_167, 0, x_165); +lean_ctor_set(x_167, 1, x_166); +return x_167; +} +} +} +else +{ +uint8_t x_168; +lean_free_object(x_133); +lean_dec(x_131); +lean_dec(x_22); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_4); +x_168 = !lean_is_exclusive(x_146); +if (x_168 == 0) +{ +return x_146; +} +else +{ +lean_object* x_169; lean_object* x_170; lean_object* x_171; +x_169 = lean_ctor_get(x_146, 0); +x_170 = lean_ctor_get(x_146, 1); +lean_inc(x_170); +lean_inc(x_169); +lean_dec(x_146); +x_171 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_171, 0, x_169); +lean_ctor_set(x_171, 1, x_170); +return x_171; +} +} +} +else +{ +uint8_t x_172; +lean_free_object(x_133); +lean_dec(x_131); +lean_dec(x_22); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_4); +x_172 = !lean_is_exclusive(x_143); +if (x_172 == 0) +{ +return x_143; +} +else +{ +lean_object* x_173; lean_object* x_174; lean_object* x_175; +x_173 = lean_ctor_get(x_143, 0); +x_174 = lean_ctor_get(x_143, 1); +lean_inc(x_174); +lean_inc(x_173); +lean_dec(x_143); +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 +{ +lean_object* x_176; lean_object* x_177; +x_176 = lean_ctor_get(x_133, 1); +lean_inc(x_176); +lean_dec(x_133); +x_177 = l_Lean_Meta_Grind_updateLastTag(x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_176); +if (lean_obj_tag(x_177) == 0) +{ +lean_object* x_178; lean_object* x_179; lean_object* x_180; +x_178 = lean_ctor_get(x_177, 1); +lean_inc(x_178); +lean_dec(x_177); +x_179 = lean_ctor_get(x_4, 5); +lean_inc(x_179); +x_180 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_179, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_178); +if (lean_obj_tag(x_180) == 0) +{ +lean_object* x_181; lean_object* x_182; lean_object* x_183; +x_181 = lean_ctor_get(x_180, 0); +lean_inc(x_181); +x_182 = lean_ctor_get(x_180, 1); +lean_inc(x_182); +lean_dec(x_180); +lean_inc(x_20); +lean_inc(x_19); +lean_inc(x_18); +lean_inc(x_17); +x_183 = lean_infer_type(x_131, x_17, x_18, x_19, x_20, x_182); +if (lean_obj_tag(x_183) == 0) +{ +lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; +x_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 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__3; +x_187 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_187, 0, x_186); +lean_ctor_set(x_187, 1, x_181); +x_188 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__5; +x_189 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_189, 0, x_187); +lean_ctor_set(x_189, 1, x_188); +x_190 = l_Lean_indentExpr(x_184); +x_191 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_191, 0, x_189); +lean_ctor_set(x_191, 1, x_190); +x_192 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; +x_193 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_193, 0, x_191); +lean_ctor_set(x_193, 1, x_192); +x_194 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_132, x_193, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_185); +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_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2; +x_198 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__1(x_22, x_197, x_4, x_5, x_195, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_196); +lean_dec(x_195); +return x_198; +} +else +{ +lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; +lean_dec(x_181); +lean_dec(x_22); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_4); +x_199 = lean_ctor_get(x_183, 0); +lean_inc(x_199); +x_200 = lean_ctor_get(x_183, 1); +lean_inc(x_200); +if (lean_is_exclusive(x_183)) { + lean_ctor_release(x_183, 0); + lean_ctor_release(x_183, 1); + x_201 = x_183; +} else { + lean_dec_ref(x_183); + x_201 = lean_box(0); +} +if (lean_is_scalar(x_201)) { + x_202 = lean_alloc_ctor(1, 2, 0); +} else { + x_202 = x_201; +} +lean_ctor_set(x_202, 0, x_199); +lean_ctor_set(x_202, 1, x_200); +return x_202; +} +} +else +{ +lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; +lean_dec(x_131); +lean_dec(x_22); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_4); +x_203 = lean_ctor_get(x_180, 0); +lean_inc(x_203); +x_204 = lean_ctor_get(x_180, 1); +lean_inc(x_204); +if (lean_is_exclusive(x_180)) { + lean_ctor_release(x_180, 0); + lean_ctor_release(x_180, 1); + x_205 = x_180; +} else { + lean_dec_ref(x_180); + x_205 = lean_box(0); +} +if (lean_is_scalar(x_205)) { + x_206 = lean_alloc_ctor(1, 2, 0); +} else { + x_206 = x_205; +} +lean_ctor_set(x_206, 0, x_203); +lean_ctor_set(x_206, 1, x_204); +return x_206; +} +} +else +{ +lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; +lean_dec(x_131); +lean_dec(x_22); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_4); +x_207 = lean_ctor_get(x_177, 0); +lean_inc(x_207); +x_208 = lean_ctor_get(x_177, 1); +lean_inc(x_208); +if (lean_is_exclusive(x_177)) { + lean_ctor_release(x_177, 0); + lean_ctor_release(x_177, 1); + x_209 = x_177; +} else { + lean_dec_ref(x_177); + 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 +{ +size_t x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; size_t x_230; lean_object* x_231; lean_object* x_232; lean_object* x_313; +x_224 = lean_usize_of_nat(x_3); +x_225 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2; +x_226 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__8(x_2, x_8, x_224, x_225, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_26); +x_227 = lean_ctor_get(x_226, 0); +lean_inc(x_227); +x_228 = lean_ctor_get(x_226, 1); +lean_inc(x_228); +if (lean_is_exclusive(x_226)) { + lean_ctor_release(x_226, 0); + lean_ctor_release(x_226, 1); + x_229 = x_226; +} else { + lean_dec_ref(x_226); + x_229 = lean_box(0); +} +x_230 = lean_array_size(x_227); +lean_inc(x_20); +lean_inc(x_19); +lean_inc(x_18); +lean_inc(x_17); +lean_inc(x_7); +x_313 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__9(x_6, x_227, x_7, x_227, x_230, x_8, x_7, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_228); +if (lean_obj_tag(x_313) == 0) +{ +lean_object* x_314; lean_object* x_315; +x_314 = lean_ctor_get(x_313, 0); +lean_inc(x_314); +x_315 = lean_ctor_get(x_314, 0); +lean_inc(x_315); +lean_dec(x_314); +if (lean_obj_tag(x_315) == 0) +{ +lean_object* x_316; +x_316 = lean_ctor_get(x_313, 1); +lean_inc(x_316); +lean_dec(x_313); +x_231 = x_9; +x_232 = x_316; +goto block_312; +} +else +{ +lean_object* x_317; lean_object* x_318; +lean_dec(x_9); +x_317 = lean_ctor_get(x_313, 1); +lean_inc(x_317); +lean_dec(x_313); +x_318 = lean_ctor_get(x_315, 0); +lean_inc(x_318); +lean_dec(x_315); +x_231 = x_318; +x_232 = x_317; +goto block_312; +} +} +else +{ +uint8_t x_319; +lean_dec(x_229); +lean_dec(x_227); +lean_dec(x_22); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_9); +lean_dec(x_4); +x_319 = !lean_is_exclusive(x_313); +if (x_319 == 0) +{ +return x_313; +} +else +{ +lean_object* x_320; lean_object* x_321; lean_object* x_322; +x_320 = lean_ctor_get(x_313, 0); +x_321 = lean_ctor_get(x_313, 1); +lean_inc(x_321); +lean_inc(x_320); +lean_dec(x_313); +x_322 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_322, 0, x_320); +lean_ctor_set(x_322, 1, x_321); +return x_322; +} +} +block_312: +{ +if (lean_obj_tag(x_231) == 0) +{ +lean_object* x_233; lean_object* x_234; +lean_dec(x_229); +x_233 = lean_box(0); +x_234 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__1(x_22, x_227, x_4, x_5, x_233, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_232); +lean_dec(x_227); +return x_234; +} +else +{ +lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; uint8_t x_239; +x_235 = lean_ctor_get(x_231, 0); +lean_inc(x_235); +lean_dec(x_231); +x_236 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__2; +x_237 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__2(x_236, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_232); +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_240; lean_object* x_241; lean_object* x_242; +lean_dec(x_235); +lean_dec(x_229); +x_240 = lean_ctor_get(x_237, 1); +lean_inc(x_240); +lean_dec(x_237); +x_241 = lean_box(0); +x_242 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__1(x_22, x_227, x_4, x_5, x_241, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_240); +lean_dec(x_227); +return x_242; +} +else +{ +uint8_t x_243; +x_243 = !lean_is_exclusive(x_237); +if (x_243 == 0) +{ +lean_object* x_244; lean_object* x_245; lean_object* x_246; +x_244 = lean_ctor_get(x_237, 1); +x_245 = lean_ctor_get(x_237, 0); +lean_dec(x_245); +x_246 = l_Lean_Meta_Grind_updateLastTag(x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_244); +if (lean_obj_tag(x_246) == 0) +{ +lean_object* x_247; lean_object* x_248; lean_object* x_249; +x_247 = lean_ctor_get(x_246, 1); +lean_inc(x_247); +lean_dec(x_246); +x_248 = lean_ctor_get(x_4, 5); +lean_inc(x_248); +x_249 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_248, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_247); +if (lean_obj_tag(x_249) == 0) +{ +lean_object* x_250; lean_object* x_251; lean_object* x_252; +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); +lean_inc(x_20); +lean_inc(x_19); +lean_inc(x_18); +lean_inc(x_17); +x_252 = lean_infer_type(x_235, x_17, x_18, x_19, x_20, x_251); +if (lean_obj_tag(x_252) == 0) +{ +lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; +x_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_255 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__3; +lean_ctor_set_tag(x_237, 7); +lean_ctor_set(x_237, 1, x_250); +lean_ctor_set(x_237, 0, x_255); +x_256 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__5; +if (lean_is_scalar(x_229)) { + x_257 = lean_alloc_ctor(7, 2, 0); +} else { + x_257 = x_229; + lean_ctor_set_tag(x_257, 7); +} +lean_ctor_set(x_257, 0, x_237); +lean_ctor_set(x_257, 1, x_256); +x_258 = l_Lean_indentExpr(x_253); +x_259 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_259, 0, x_257); +lean_ctor_set(x_259, 1, x_258); +x_260 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; +x_261 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_261, 0, x_259); +lean_ctor_set(x_261, 1, x_260); +x_262 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_236, x_261, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_254); +x_263 = lean_ctor_get(x_262, 0); +lean_inc(x_263); +x_264 = lean_ctor_get(x_262, 1); +lean_inc(x_264); +lean_dec(x_262); +x_265 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__1(x_22, x_227, x_4, x_5, x_263, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_264); +lean_dec(x_263); +lean_dec(x_227); +return x_265; +} +else +{ +uint8_t x_266; +lean_dec(x_250); +lean_free_object(x_237); +lean_dec(x_229); +lean_dec(x_227); +lean_dec(x_22); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_4); +x_266 = !lean_is_exclusive(x_252); +if (x_266 == 0) +{ +return x_252; +} +else +{ +lean_object* x_267; lean_object* x_268; lean_object* x_269; +x_267 = lean_ctor_get(x_252, 0); +x_268 = lean_ctor_get(x_252, 1); +lean_inc(x_268); +lean_inc(x_267); +lean_dec(x_252); +x_269 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_269, 0, x_267); +lean_ctor_set(x_269, 1, x_268); +return x_269; +} +} +} +else +{ +uint8_t x_270; +lean_free_object(x_237); +lean_dec(x_235); +lean_dec(x_229); +lean_dec(x_227); +lean_dec(x_22); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_4); +x_270 = !lean_is_exclusive(x_249); +if (x_270 == 0) +{ +return x_249; +} +else +{ +lean_object* x_271; lean_object* x_272; lean_object* x_273; +x_271 = lean_ctor_get(x_249, 0); +x_272 = lean_ctor_get(x_249, 1); +lean_inc(x_272); +lean_inc(x_271); +lean_dec(x_249); +x_273 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_273, 0, x_271); +lean_ctor_set(x_273, 1, x_272); +return x_273; +} +} +} +else +{ +uint8_t x_274; +lean_free_object(x_237); +lean_dec(x_235); +lean_dec(x_229); +lean_dec(x_227); +lean_dec(x_22); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_4); +x_274 = !lean_is_exclusive(x_246); +if (x_274 == 0) +{ +return x_246; +} +else +{ +lean_object* x_275; lean_object* x_276; lean_object* x_277; +x_275 = lean_ctor_get(x_246, 0); +x_276 = lean_ctor_get(x_246, 1); +lean_inc(x_276); +lean_inc(x_275); +lean_dec(x_246); +x_277 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_277, 0, x_275); +lean_ctor_set(x_277, 1, x_276); +return x_277; +} +} +} +else +{ +lean_object* x_278; lean_object* x_279; +x_278 = lean_ctor_get(x_237, 1); +lean_inc(x_278); +lean_dec(x_237); +x_279 = l_Lean_Meta_Grind_updateLastTag(x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_278); +if (lean_obj_tag(x_279) == 0) +{ +lean_object* x_280; lean_object* x_281; lean_object* x_282; +x_280 = lean_ctor_get(x_279, 1); +lean_inc(x_280); +lean_dec(x_279); +x_281 = lean_ctor_get(x_4, 5); +lean_inc(x_281); +x_282 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_281, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_280); +if (lean_obj_tag(x_282) == 0) +{ +lean_object* x_283; lean_object* x_284; lean_object* x_285; +x_283 = lean_ctor_get(x_282, 0); +lean_inc(x_283); +x_284 = lean_ctor_get(x_282, 1); +lean_inc(x_284); +lean_dec(x_282); +lean_inc(x_20); +lean_inc(x_19); +lean_inc(x_18); +lean_inc(x_17); +x_285 = lean_infer_type(x_235, x_17, x_18, x_19, x_20, x_284); +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; +x_286 = lean_ctor_get(x_285, 0); +lean_inc(x_286); +x_287 = lean_ctor_get(x_285, 1); +lean_inc(x_287); +lean_dec(x_285); +x_288 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__3; +x_289 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_289, 0, x_288); +lean_ctor_set(x_289, 1, x_283); +x_290 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__5; +if (lean_is_scalar(x_229)) { + x_291 = lean_alloc_ctor(7, 2, 0); +} else { + x_291 = x_229; + lean_ctor_set_tag(x_291, 7); +} +lean_ctor_set(x_291, 0, x_289); +lean_ctor_set(x_291, 1, x_290); +x_292 = l_Lean_indentExpr(x_286); +x_293 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_293, 0, x_291); +lean_ctor_set(x_293, 1, x_292); +x_294 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; +x_295 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_295, 0, x_293); +lean_ctor_set(x_295, 1, x_294); +x_296 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_236, x_295, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_287); +x_297 = lean_ctor_get(x_296, 0); +lean_inc(x_297); +x_298 = lean_ctor_get(x_296, 1); +lean_inc(x_298); +lean_dec(x_296); +x_299 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__1(x_22, x_227, x_4, x_5, x_297, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_298); +lean_dec(x_297); +lean_dec(x_227); +return x_299; +} +else +{ +lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; +lean_dec(x_283); +lean_dec(x_229); +lean_dec(x_227); +lean_dec(x_22); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_4); +x_300 = lean_ctor_get(x_285, 0); +lean_inc(x_300); +x_301 = lean_ctor_get(x_285, 1); +lean_inc(x_301); +if (lean_is_exclusive(x_285)) { + lean_ctor_release(x_285, 0); + lean_ctor_release(x_285, 1); + x_302 = x_285; +} else { + lean_dec_ref(x_285); + x_302 = lean_box(0); +} +if (lean_is_scalar(x_302)) { + x_303 = lean_alloc_ctor(1, 2, 0); +} else { + x_303 = x_302; +} +lean_ctor_set(x_303, 0, x_300); +lean_ctor_set(x_303, 1, x_301); +return x_303; +} +} +else +{ +lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; +lean_dec(x_235); +lean_dec(x_229); +lean_dec(x_227); +lean_dec(x_22); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_4); +x_304 = lean_ctor_get(x_282, 0); +lean_inc(x_304); +x_305 = lean_ctor_get(x_282, 1); +lean_inc(x_305); +if (lean_is_exclusive(x_282)) { + lean_ctor_release(x_282, 0); + lean_ctor_release(x_282, 1); + x_306 = x_282; +} else { + lean_dec_ref(x_282); + 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_dec(x_235); +lean_dec(x_229); +lean_dec(x_227); +lean_dec(x_22); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_4); +x_308 = lean_ctor_get(x_279, 0); +lean_inc(x_308); +x_309 = lean_ctor_get(x_279, 1); +lean_inc(x_309); +if (lean_is_exclusive(x_279)) { + lean_ctor_release(x_279, 0); + lean_ctor_release(x_279, 1); + x_310 = x_279; +} else { + lean_dec_ref(x_279); + x_310 = lean_box(0); +} +if (lean_is_scalar(x_310)) { + x_311 = lean_alloc_ctor(1, 2, 0); +} else { + x_311 = x_310; +} +lean_ctor_set(x_311, 0, x_308); +lean_ctor_set(x_311, 1, x_309); +return x_311; +} +} +} +} +} +} +} +} +else +{ +lean_object* x_323; lean_object* x_324; lean_object* x_325; +lean_dec(x_9); +lean_dec(x_7); +x_323 = lean_ctor_get(x_4, 5); +lean_inc(x_323); +lean_dec(x_4); +x_324 = lean_ctor_get(x_5, 1); +x_325 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance(x_323, x_22, x_324, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_26); +return x_325; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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, lean_object* x_19, lean_object* x_20) { +_start: +{ +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_21 = lean_array_get_size(x_1); +x_22 = lean_unsigned_to_nat(0u); +x_23 = l_Array_toSubarray___rarg(x_1, x_22, x_21); +x_24 = lean_box(0); +lean_inc(x_2); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_2); +lean_ctor_set(x_25, 1, x_23); +x_26 = lean_array_size(x_3); +x_27 = 0; +lean_inc(x_19); +lean_inc(x_18); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_2); +lean_inc(x_4); +x_28 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5(x_4, x_3, x_2, x_24, x_3, x_26, x_27, x_25, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20); +if (lean_obj_tag(x_28) == 0) +{ +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); +if (lean_obj_tag(x_30) == 0) +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_28, 1); +lean_inc(x_31); +lean_dec(x_28); +x_32 = lean_box(0); +x_33 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2(x_5, x_3, x_6, x_4, x_7, x_24, x_8, x_27, x_2, x_32, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_31); +return x_33; +} +else +{ +uint8_t x_34; +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_8); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +x_34 = !lean_is_exclusive(x_28); +if (x_34 == 0) +{ +lean_object* x_35; lean_object* x_36; +x_35 = lean_ctor_get(x_28, 0); +lean_dec(x_35); +x_36 = lean_ctor_get(x_30, 0); +lean_inc(x_36); +lean_dec(x_30); +lean_ctor_set(x_28, 0, x_36); +return x_28; +} +else +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_37 = lean_ctor_get(x_28, 1); +lean_inc(x_37); +lean_dec(x_28); +x_38 = lean_ctor_get(x_30, 0); +lean_inc(x_38); +lean_dec(x_30); +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_40; +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_8); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +x_40 = !lean_is_exclusive(x_28); +if (x_40 == 0) +{ +return x_28; +} +else +{ +lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_41 = lean_ctor_get(x_28, 0); +x_42 = lean_ctor_get(x_28, 1); +lean_inc(x_42); +lean_inc(x_41); +lean_dec(x_28); +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; +} +} +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__4___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = lean_box(0); +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19, lean_object* x_20) { +_start: +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_21 = lean_array_get_size(x_1); +x_22 = lean_unsigned_to_nat(0u); +x_23 = lean_unsigned_to_nat(1u); +lean_inc(x_21); +x_24 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_24, 0, x_22); +lean_ctor_set(x_24, 1, x_21); +lean_ctor_set(x_24, 2, x_23); +x_25 = lean_box(0); +x_26 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__4___closed__1; +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_inc(x_2); +x_27 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3(x_2, x_3, x_4, x_5, x_1, x_24, x_26, x_24, x_26, x_22, lean_box(0), lean_box(0), x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20); +lean_dec(x_24); +if (lean_obj_tag(x_27) == 0) +{ +lean_object* x_28; lean_object* x_29; +x_28 = lean_ctor_get(x_27, 0); +lean_inc(x_28); +x_29 = lean_ctor_get(x_28, 0); +lean_inc(x_29); +lean_dec(x_28); +if (lean_obj_tag(x_29) == 0) +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_30 = lean_ctor_get(x_27, 1); +lean_inc(x_30); +lean_dec(x_27); +x_31 = lean_box(0); +x_32 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3(x_6, x_25, x_1, x_2, x_7, x_21, x_8, x_26, x_31, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_30); +lean_dec(x_21); +return x_32; +} +else +{ +uint8_t x_33; +lean_dec(x_21); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_2); +x_33 = !lean_is_exclusive(x_27); +if (x_33 == 0) +{ +lean_object* x_34; lean_object* x_35; +x_34 = lean_ctor_get(x_27, 0); +lean_dec(x_34); +x_35 = lean_ctor_get(x_29, 0); +lean_inc(x_35); +lean_dec(x_29); +lean_ctor_set(x_27, 0, x_35); +return x_27; +} +else +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_27, 1); +lean_inc(x_36); +lean_dec(x_27); +x_37 = lean_ctor_get(x_29, 0); +lean_inc(x_37); +lean_dec(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_36); +return x_38; +} +} +} +else +{ +uint8_t x_39; +lean_dec(x_21); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_2); +x_39 = !lean_is_exclusive(x_27); +if (x_39 == 0) +{ +return x_27; +} +else +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_40 = lean_ctor_get(x_27, 0); +x_41 = lean_ctor_get(x_27, 1); +lean_inc(x_41); +lean_inc(x_40); +lean_dec(x_27); +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_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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) { +_start: +{ +lean_object* x_13; lean_object* x_14; +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; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("c.assignment.size == numParams\n ", 33, 33); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__1; +x_3 = lean_string_append(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("_private.Lean.Meta.Tactic.Grind.EMatch.0.Lean.Meta.Grind.EMatch.instantiateTheorem", 82, 82); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__3; +x_3 = lean_unsigned_to_nat(240u); +x_4 = lean_unsigned_to_nat(2u); +x_5 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__2; +x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); +return x_6; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__5___boxed), 12, 0); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("unexpected number of parameters at ", 35, 35); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__6; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +_start: +{ +lean_object* x_16; +lean_inc(x_1); +x_16 = l_Lean_Meta_Grind_EMatchTheorem_getProofWithFreshMVarLevels(x_1, x_11, x_12, x_13, x_14, x_15); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +x_18 = lean_ctor_get(x_16, 1); +lean_inc(x_18); +lean_dec(x_16); +x_19 = lean_ctor_get(x_1, 2); +lean_inc(x_19); +x_20 = lean_array_get_size(x_2); +x_21 = lean_nat_dec_eq(x_20, x_19); +if (x_21 == 0) +{ +lean_object* x_22; lean_object* x_23; +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_17); +lean_dec(x_1); +x_22 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__4; +x_23 = l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__2(x_22, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_18); +return x_23; +} +else +{ +lean_object* x_24; +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_17); +x_24 = lean_infer_type(x_17, x_11, x_12, x_13, x_14, x_18); +if (lean_obj_tag(x_24) == 0) +{ +lean_object* x_25; lean_object* x_26; uint8_t x_27; lean_object* x_28; +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 = 0; +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_19); +x_28 = l_Lean_Meta_forallMetaBoundedTelescope(x_25, x_19, x_27, x_11, x_12, x_13, x_14, x_26); +if (lean_obj_tag(x_28) == 0) +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; +x_29 = lean_ctor_get(x_28, 0); +lean_inc(x_29); +x_30 = lean_ctor_get(x_29, 1); +lean_inc(x_30); +x_31 = lean_ctor_get(x_28, 1); +lean_inc(x_31); +lean_dec(x_28); +x_32 = lean_ctor_get(x_29, 0); +lean_inc(x_32); +lean_dec(x_29); +x_33 = !lean_is_exclusive(x_30); +if (x_33 == 0) +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; +x_34 = lean_ctor_get(x_30, 0); +x_35 = lean_ctor_get(x_30, 1); +lean_dec(x_35); +x_36 = lean_array_get_size(x_32); +x_37 = lean_nat_dec_eq(x_36, x_19); +lean_dec(x_36); +if (x_37 == 0) +{ +lean_object* x_38; lean_object* x_39; uint8_t x_40; +lean_dec(x_34); +lean_dec(x_32); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_17); +x_38 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__2; +x_39 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__2(x_38, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_31); +x_40 = !lean_is_exclusive(x_39); +if (x_40 == 0) +{ +lean_object* x_41; lean_object* x_42; lean_object* x_43; uint8_t x_44; +x_41 = lean_ctor_get(x_39, 0); +x_42 = lean_ctor_get(x_39, 1); +x_43 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__5; +x_44 = lean_unbox(x_41); +lean_dec(x_41); +if (x_44 == 0) +{ +lean_object* x_45; lean_object* x_46; +lean_free_object(x_39); +lean_free_object(x_30); +lean_dec(x_1); +x_45 = lean_box(0); +x_46 = lean_apply_12(x_43, x_45, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_42); +return x_46; +} +else +{ +lean_object* x_47; +x_47 = l_Lean_Meta_Grind_updateLastTag(x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_42); +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_1, 5); +lean_inc(x_49); +lean_dec(x_1); +x_50 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_49, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_48); +if (lean_obj_tag(x_50) == 0) +{ +lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; +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___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__7; +lean_ctor_set_tag(x_39, 7); +lean_ctor_set(x_39, 1, x_51); +lean_ctor_set(x_39, 0, x_53); +x_54 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; +lean_ctor_set_tag(x_30, 7); +lean_ctor_set(x_30, 1, x_54); +lean_ctor_set(x_30, 0, x_39); +x_55 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_38, x_30, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_52); +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_apply_12(x_43, x_56, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_57); +return x_58; +} +else +{ +uint8_t x_59; +lean_free_object(x_39); +lean_free_object(x_30); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +x_59 = !lean_is_exclusive(x_50); +if (x_59 == 0) +{ +return x_50; +} +else +{ +lean_object* x_60; lean_object* x_61; lean_object* x_62; +x_60 = lean_ctor_get(x_50, 0); +x_61 = lean_ctor_get(x_50, 1); +lean_inc(x_61); +lean_inc(x_60); +lean_dec(x_50); +x_62 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_62, 0, x_60); +lean_ctor_set(x_62, 1, x_61); +return x_62; +} +} +} +else +{ +uint8_t x_63; +lean_free_object(x_39); +lean_free_object(x_30); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_1); +x_63 = !lean_is_exclusive(x_47); +if (x_63 == 0) +{ +return x_47; +} +else +{ +lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_64 = lean_ctor_get(x_47, 0); +x_65 = lean_ctor_get(x_47, 1); +lean_inc(x_65); +lean_inc(x_64); +lean_dec(x_47); +x_66 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_66, 0, x_64); +lean_ctor_set(x_66, 1, x_65); +return x_66; +} +} +} +} +else +{ +lean_object* x_67; lean_object* x_68; lean_object* x_69; uint8_t x_70; +x_67 = lean_ctor_get(x_39, 0); +x_68 = lean_ctor_get(x_39, 1); +lean_inc(x_68); +lean_inc(x_67); +lean_dec(x_39); +x_69 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__5; +x_70 = lean_unbox(x_67); +lean_dec(x_67); +if (x_70 == 0) +{ +lean_object* x_71; lean_object* x_72; +lean_free_object(x_30); +lean_dec(x_1); +x_71 = lean_box(0); +x_72 = lean_apply_12(x_69, x_71, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_68); +return x_72; +} +else +{ +lean_object* x_73; +x_73 = l_Lean_Meta_Grind_updateLastTag(x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_68); +if (lean_obj_tag(x_73) == 0) +{ +lean_object* x_74; lean_object* x_75; lean_object* x_76; +x_74 = lean_ctor_get(x_73, 1); +lean_inc(x_74); +lean_dec(x_73); +x_75 = lean_ctor_get(x_1, 5); +lean_inc(x_75); +lean_dec(x_1); +x_76 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_75, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_74); +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_ctor_get(x_76, 0); +lean_inc(x_77); +x_78 = lean_ctor_get(x_76, 1); +lean_inc(x_78); +lean_dec(x_76); +x_79 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__7; +x_80 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_80, 0, x_79); +lean_ctor_set(x_80, 1, x_77); +x_81 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; +lean_ctor_set_tag(x_30, 7); +lean_ctor_set(x_30, 1, x_81); +lean_ctor_set(x_30, 0, x_80); +x_82 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_38, x_30, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_78); +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_apply_12(x_69, x_83, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_84); +return x_85; +} +else +{ +lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; +lean_free_object(x_30); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +x_86 = lean_ctor_get(x_76, 0); +lean_inc(x_86); +x_87 = lean_ctor_get(x_76, 1); +lean_inc(x_87); +if (lean_is_exclusive(x_76)) { + lean_ctor_release(x_76, 0); + lean_ctor_release(x_76, 1); + x_88 = x_76; +} else { + lean_dec_ref(x_76); + x_88 = lean_box(0); +} +if (lean_is_scalar(x_88)) { + x_89 = lean_alloc_ctor(1, 2, 0); +} else { + x_89 = x_88; +} +lean_ctor_set(x_89, 0, x_86); +lean_ctor_set(x_89, 1, x_87); +return x_89; +} +} +else +{ +lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; +lean_free_object(x_30); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_1); +x_90 = lean_ctor_get(x_73, 0); +lean_inc(x_90); +x_91 = lean_ctor_get(x_73, 1); +lean_inc(x_91); +if (lean_is_exclusive(x_73)) { + lean_ctor_release(x_73, 0); + lean_ctor_release(x_73, 1); + x_92 = x_73; +} else { + lean_dec_ref(x_73); + 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; +} +} +} +} +else +{ +lean_object* x_94; lean_object* x_95; +lean_free_object(x_30); +x_94 = lean_box(0); +x_95 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__4(x_32, x_1, x_2, x_19, x_20, x_34, x_17, x_3, x_94, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_31); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_32); +return x_95; +} +} +else +{ +lean_object* x_96; lean_object* x_97; uint8_t x_98; +x_96 = lean_ctor_get(x_30, 0); +lean_inc(x_96); +lean_dec(x_30); +x_97 = lean_array_get_size(x_32); +x_98 = lean_nat_dec_eq(x_97, x_19); +lean_dec(x_97); +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; uint8_t x_105; +lean_dec(x_96); +lean_dec(x_32); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_17); +x_99 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__2; +x_100 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__2(x_99, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_31); +x_101 = lean_ctor_get(x_100, 0); +lean_inc(x_101); +x_102 = lean_ctor_get(x_100, 1); +lean_inc(x_102); +if (lean_is_exclusive(x_100)) { + lean_ctor_release(x_100, 0); + lean_ctor_release(x_100, 1); + x_103 = x_100; +} else { + lean_dec_ref(x_100); + x_103 = lean_box(0); +} +x_104 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__5; +x_105 = lean_unbox(x_101); +lean_dec(x_101); +if (x_105 == 0) +{ +lean_object* x_106; lean_object* x_107; +lean_dec(x_103); +lean_dec(x_1); +x_106 = lean_box(0); +x_107 = lean_apply_12(x_104, x_106, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_102); +return x_107; +} +else +{ +lean_object* x_108; +x_108 = l_Lean_Meta_Grind_updateLastTag(x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_102); +if (lean_obj_tag(x_108) == 0) +{ +lean_object* x_109; lean_object* x_110; lean_object* x_111; +x_109 = lean_ctor_get(x_108, 1); +lean_inc(x_109); +lean_dec(x_108); +x_110 = lean_ctor_get(x_1, 5); +lean_inc(x_110); +lean_dec(x_1); +x_111 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_110, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_109); +if (lean_obj_tag(x_111) == 0) +{ +lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; +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___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__7; +if (lean_is_scalar(x_103)) { + x_115 = lean_alloc_ctor(7, 2, 0); +} else { + x_115 = x_103; + lean_ctor_set_tag(x_115, 7); +} +lean_ctor_set(x_115, 0, x_114); +lean_ctor_set(x_115, 1, x_112); +x_116 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; +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_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_99, x_117, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_113); +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_apply_12(x_104, x_119, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_120); +return x_121; +} +else +{ +lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; +lean_dec(x_103); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +x_122 = lean_ctor_get(x_111, 0); +lean_inc(x_122); +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_111); + x_124 = lean_box(0); +} +if (lean_is_scalar(x_124)) { + x_125 = lean_alloc_ctor(1, 2, 0); +} else { + x_125 = x_124; +} +lean_ctor_set(x_125, 0, x_122); +lean_ctor_set(x_125, 1, x_123); +return x_125; +} +} +else +{ +lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; +lean_dec(x_103); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_1); +x_126 = lean_ctor_get(x_108, 0); +lean_inc(x_126); +x_127 = lean_ctor_get(x_108, 1); +lean_inc(x_127); +if (lean_is_exclusive(x_108)) { + lean_ctor_release(x_108, 0); + lean_ctor_release(x_108, 1); + x_128 = x_108; +} else { + lean_dec_ref(x_108); + x_128 = lean_box(0); +} +if (lean_is_scalar(x_128)) { + x_129 = lean_alloc_ctor(1, 2, 0); +} else { + x_129 = x_128; +} +lean_ctor_set(x_129, 0, x_126); +lean_ctor_set(x_129, 1, x_127); +return x_129; +} +} +} +else +{ +lean_object* x_130; lean_object* x_131; +x_130 = lean_box(0); +x_131 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__4(x_32, x_1, x_2, x_19, x_20, x_96, x_17, x_3, x_130, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_31); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_32); +return x_131; +} +} +} +else +{ +uint8_t x_132; +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_17); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_1); +x_132 = !lean_is_exclusive(x_28); +if (x_132 == 0) +{ +return x_28; +} +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); +lean_inc(x_134); +lean_inc(x_133); +lean_dec(x_28); +x_135 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_135, 0, x_133); +lean_ctor_set(x_135, 1, x_134); +return x_135; +} +} +} +else +{ +uint8_t x_136; +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_17); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_1); +x_136 = !lean_is_exclusive(x_24); +if (x_136 == 0) +{ +return x_24; +} +else +{ +lean_object* x_137; lean_object* x_138; lean_object* x_139; +x_137 = lean_ctor_get(x_24, 0); +x_138 = lean_ctor_get(x_24, 1); +lean_inc(x_138); +lean_inc(x_137); +lean_dec(x_24); +x_139 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_139, 0, x_137); +lean_ctor_set(x_139, 1, x_138); +return x_139; +} +} +} +} +else +{ +uint8_t x_140; +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_1); +x_140 = !lean_is_exclusive(x_16); +if (x_140 == 0) +{ +return x_16; +} +else +{ +lean_object* x_141; lean_object* x_142; lean_object* x_143; +x_141 = lean_ctor_get(x_16, 0); +x_142 = lean_ctor_get(x_16, 1); +lean_inc(x_142); +lean_inc(x_141); +lean_dec(x_16); +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; +} +} +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("assignment", 10, 10); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__2; +x_3 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__3; +x_4 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__1; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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_object* x_15) { +_start: +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; +x_16 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__2; +x_17 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__2(x_16, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +x_19 = lean_unbox(x_18); +lean_dec(x_18); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_17, 1); +lean_inc(x_20); +lean_dec(x_17); +x_21 = lean_box(0); +x_22 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6(x_1, x_2, x_3, x_21, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_20); +lean_dec(x_2); +return x_22; +} +else +{ +uint8_t x_23; +x_23 = !lean_is_exclusive(x_17); +if (x_23 == 0) +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_17, 1); +x_25 = lean_ctor_get(x_17, 0); +lean_dec(x_25); +x_26 = l_Lean_Meta_Grind_updateLastTag(x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, 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, 1); +lean_inc(x_27); +lean_dec(x_26); +x_28 = lean_ctor_get(x_1, 5); +lean_inc(x_28); +x_29 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_28, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_27); +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; +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_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; +lean_ctor_set_tag(x_17, 7); +lean_ctor_set(x_17, 1, x_30); +lean_ctor_set(x_17, 0, x_32); +x_33 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__6; +x_34 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_34, 0, x_17); +lean_ctor_set(x_34, 1, x_33); +lean_inc(x_2); +x_35 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData(x_2); +x_36 = lean_array_to_list(x_35); +x_37 = lean_box(0); +x_38 = l_List_mapTR_loop___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__2(x_36, x_37); +x_39 = l_Lean_MessageData_ofList(x_38); +x_40 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_40, 0, x_34); +lean_ctor_set(x_40, 1, x_39); +x_41 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_41, 0, x_40); +lean_ctor_set(x_41, 1, x_32); +x_42 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_16, x_41, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_31); +x_43 = lean_ctor_get(x_42, 0); +lean_inc(x_43); +x_44 = lean_ctor_get(x_42, 1); +lean_inc(x_44); +lean_dec(x_42); +x_45 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6(x_1, x_2, x_3, x_43, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_44); +lean_dec(x_43); +lean_dec(x_2); +return x_45; +} +else +{ +uint8_t x_46; +lean_free_object(x_17); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_2); +lean_dec(x_1); +x_46 = !lean_is_exclusive(x_29); +if (x_46 == 0) +{ +return x_29; +} +else +{ +lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_47 = lean_ctor_get(x_29, 0); +x_48 = lean_ctor_get(x_29, 1); +lean_inc(x_48); +lean_inc(x_47); +lean_dec(x_29); +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_50; +lean_free_object(x_17); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_2); +lean_dec(x_1); +x_50 = !lean_is_exclusive(x_26); +if (x_50 == 0) +{ +return x_26; +} +else +{ +lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_51 = lean_ctor_get(x_26, 0); +x_52 = lean_ctor_get(x_26, 1); +lean_inc(x_52); +lean_inc(x_51); +lean_dec(x_26); +x_53 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_53, 0, x_51); +lean_ctor_set(x_53, 1, x_52); +return x_53; +} +} +} +else +{ +lean_object* x_54; lean_object* x_55; +x_54 = lean_ctor_get(x_17, 1); +lean_inc(x_54); +lean_dec(x_17); +x_55 = l_Lean_Meta_Grind_updateLastTag(x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_54); +if (lean_obj_tag(x_55) == 0) +{ +lean_object* x_56; lean_object* x_57; lean_object* x_58; +x_56 = lean_ctor_get(x_55, 1); +lean_inc(x_56); +lean_dec(x_55); +x_57 = lean_ctor_get(x_1, 5); +lean_inc(x_57); +x_58 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_57, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_56); +if (lean_obj_tag(x_58) == 0) +{ +lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; 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_59 = lean_ctor_get(x_58, 0); +lean_inc(x_59); +x_60 = lean_ctor_get(x_58, 1); +lean_inc(x_60); +lean_dec(x_58); +x_61 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; +x_62 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_62, 0, x_61); +lean_ctor_set(x_62, 1, x_59); +x_63 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__6; +x_64 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_64, 0, x_62); +lean_ctor_set(x_64, 1, x_63); +lean_inc(x_2); +x_65 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData(x_2); +x_66 = lean_array_to_list(x_65); +x_67 = lean_box(0); +x_68 = l_List_mapTR_loop___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__2(x_66, x_67); +x_69 = l_Lean_MessageData_ofList(x_68); +x_70 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_70, 0, x_64); +lean_ctor_set(x_70, 1, x_69); +x_71 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_71, 0, x_70); +lean_ctor_set(x_71, 1, x_61); +x_72 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_16, x_71, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_60); +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___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6(x_1, x_2, x_3, x_73, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_74); +lean_dec(x_73); +lean_dec(x_2); +return x_75; +} +else +{ +lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_2); +lean_dec(x_1); +x_76 = lean_ctor_get(x_58, 0); +lean_inc(x_76); +x_77 = lean_ctor_get(x_58, 1); +lean_inc(x_77); +if (lean_is_exclusive(x_58)) { + lean_ctor_release(x_58, 0); + lean_ctor_release(x_58, 1); + x_78 = x_58; +} else { + lean_dec_ref(x_58); + 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 +{ +lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_2); +lean_dec(x_1); +x_80 = lean_ctor_get(x_55, 0); +lean_inc(x_80); +x_81 = lean_ctor_get(x_55, 1); +lean_inc(x_81); +if (lean_is_exclusive(x_55)) { + lean_ctor_release(x_55, 0); + lean_ctor_release(x_55, 1); + x_82 = x_55; +} else { + lean_dec_ref(x_55); + 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; +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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) { +_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; +x_14 = lean_ctor_get(x_2, 0); +lean_inc(x_14); +lean_dec(x_2); +x_15 = lean_ctor_get(x_14, 1); +lean_inc(x_15); +x_16 = lean_ctor_get(x_1, 2); +lean_inc(x_16); +lean_inc(x_16); +x_17 = l_Lean_Meta_Grind_markTheoremInstance(x_15, x_16, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +x_19 = lean_unbox(x_18); +lean_dec(x_18); +if (x_19 == 0) +{ +uint8_t x_20; +lean_dec(x_16); +lean_dec(x_14); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_20 = !lean_is_exclusive(x_17); +if (x_20 == 0) +{ +lean_object* x_21; lean_object* x_22; +x_21 = lean_ctor_get(x_17, 0); +lean_dec(x_21); +x_22 = lean_box(0); +lean_ctor_set(x_17, 0, x_22); +return x_17; +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_17, 1); +lean_inc(x_23); +lean_dec(x_17); +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_26; lean_object* x_27; lean_object* x_28; +x_26 = lean_ctor_get(x_17, 1); +lean_inc(x_26); +lean_dec(x_17); +x_27 = lean_box(0); +x_28 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7(x_14, x_16, x_1, x_27, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_26); +lean_dec(x_1); +return x_28; +} +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_ReaderT_read___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__1___boxed), 11, 0); +return x_1; +} +} +static uint64_t _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___closed__2() { +_start: +{ +uint8_t x_1; uint64_t x_2; +x_1 = 1; +x_2 = l_Lean_Meta_TransparencyMode_toUInt64(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_13 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__8), 13, 1); +lean_closure_set(x_13, 0, x_1); +x_14 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___closed__1; +x_15 = lean_alloc_closure((void*)(l_ReaderT_bind___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__11___rarg), 13, 2); +lean_closure_set(x_15, 0, x_14); +lean_closure_set(x_15, 1, x_13); +x_16 = !lean_is_exclusive(x_8); +if (x_16 == 0) +{ +lean_object* x_17; uint8_t x_18; +x_17 = lean_ctor_get(x_8, 0); +x_18 = !lean_is_exclusive(x_17); +if (x_18 == 0) +{ +uint64_t x_19; uint8_t x_20; uint64_t x_21; uint64_t x_22; uint64_t x_23; uint64_t x_24; uint64_t x_25; uint8_t x_26; lean_object* x_27; +x_19 = lean_ctor_get_uint64(x_8, sizeof(void*)*7); +x_20 = 1; +lean_ctor_set_uint8(x_17, 9, x_20); +x_21 = 2; +x_22 = lean_uint64_shift_right(x_19, x_21); +x_23 = lean_uint64_shift_left(x_22, x_21); +x_24 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___closed__2; +x_25 = lean_uint64_lor(x_23, x_24); +lean_ctor_set_uint64(x_8, sizeof(void*)*7, x_25); +x_26 = 0; +x_27 = l_Lean_Meta_withNewMCtxDepth___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__12___rarg(x_15, x_26, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_27) == 0) +{ +uint8_t x_28; +x_28 = !lean_is_exclusive(x_27); +if (x_28 == 0) +{ +return x_27; +} +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_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; +} +} +else +{ +uint8_t x_32; +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 +{ +uint64_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; 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; uint64_t x_55; uint64_t x_56; uint64_t x_57; uint64_t x_58; uint64_t x_59; uint8_t x_60; lean_object* x_61; +x_36 = lean_ctor_get_uint64(x_8, sizeof(void*)*7); +x_37 = lean_ctor_get_uint8(x_17, 0); +x_38 = lean_ctor_get_uint8(x_17, 1); +x_39 = lean_ctor_get_uint8(x_17, 2); +x_40 = lean_ctor_get_uint8(x_17, 3); +x_41 = lean_ctor_get_uint8(x_17, 4); +x_42 = lean_ctor_get_uint8(x_17, 5); +x_43 = lean_ctor_get_uint8(x_17, 6); +x_44 = lean_ctor_get_uint8(x_17, 7); +x_45 = lean_ctor_get_uint8(x_17, 8); +x_46 = lean_ctor_get_uint8(x_17, 10); +x_47 = lean_ctor_get_uint8(x_17, 11); +x_48 = lean_ctor_get_uint8(x_17, 12); +x_49 = lean_ctor_get_uint8(x_17, 13); +x_50 = lean_ctor_get_uint8(x_17, 14); +x_51 = lean_ctor_get_uint8(x_17, 15); +x_52 = lean_ctor_get_uint8(x_17, 16); +lean_dec(x_17); +x_53 = 1; +x_54 = lean_alloc_ctor(0, 0, 17); +lean_ctor_set_uint8(x_54, 0, x_37); +lean_ctor_set_uint8(x_54, 1, x_38); +lean_ctor_set_uint8(x_54, 2, x_39); +lean_ctor_set_uint8(x_54, 3, x_40); +lean_ctor_set_uint8(x_54, 4, x_41); +lean_ctor_set_uint8(x_54, 5, x_42); +lean_ctor_set_uint8(x_54, 6, x_43); +lean_ctor_set_uint8(x_54, 7, x_44); +lean_ctor_set_uint8(x_54, 8, x_45); +lean_ctor_set_uint8(x_54, 9, x_53); +lean_ctor_set_uint8(x_54, 10, x_46); +lean_ctor_set_uint8(x_54, 11, x_47); +lean_ctor_set_uint8(x_54, 12, x_48); +lean_ctor_set_uint8(x_54, 13, x_49); +lean_ctor_set_uint8(x_54, 14, x_50); +lean_ctor_set_uint8(x_54, 15, x_51); +lean_ctor_set_uint8(x_54, 16, x_52); +x_55 = 2; +x_56 = lean_uint64_shift_right(x_36, x_55); +x_57 = lean_uint64_shift_left(x_56, x_55); +x_58 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___closed__2; +x_59 = lean_uint64_lor(x_57, x_58); +lean_ctor_set(x_8, 0, x_54); +lean_ctor_set_uint64(x_8, sizeof(void*)*7, x_59); +x_60 = 0; +x_61 = l_Lean_Meta_withNewMCtxDepth___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__12___rarg(x_15, x_60, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_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_61, 0); +lean_inc(x_66); +x_67 = lean_ctor_get(x_61, 1); +lean_inc(x_67); +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_61); + x_68 = lean_box(0); +} +if (lean_is_scalar(x_68)) { + x_69 = lean_alloc_ctor(1, 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; uint64_t x_71; 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; 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; 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; lean_object* x_97; uint8_t x_98; lean_object* x_99; uint64_t x_100; uint64_t x_101; uint64_t x_102; uint64_t x_103; uint64_t x_104; lean_object* x_105; uint8_t x_106; lean_object* x_107; +x_70 = lean_ctor_get(x_8, 0); +x_71 = lean_ctor_get_uint64(x_8, sizeof(void*)*7); +x_72 = lean_ctor_get_uint8(x_8, sizeof(void*)*7 + 8); +x_73 = lean_ctor_get(x_8, 1); +x_74 = lean_ctor_get(x_8, 2); +x_75 = lean_ctor_get(x_8, 3); +x_76 = lean_ctor_get(x_8, 4); +x_77 = lean_ctor_get(x_8, 5); +x_78 = lean_ctor_get(x_8, 6); +x_79 = lean_ctor_get_uint8(x_8, sizeof(void*)*7 + 9); +x_80 = lean_ctor_get_uint8(x_8, sizeof(void*)*7 + 10); +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_70); +lean_dec(x_8); +x_81 = lean_ctor_get_uint8(x_70, 0); +x_82 = lean_ctor_get_uint8(x_70, 1); +x_83 = lean_ctor_get_uint8(x_70, 2); +x_84 = lean_ctor_get_uint8(x_70, 3); +x_85 = lean_ctor_get_uint8(x_70, 4); +x_86 = lean_ctor_get_uint8(x_70, 5); +x_87 = lean_ctor_get_uint8(x_70, 6); +x_88 = lean_ctor_get_uint8(x_70, 7); +x_89 = lean_ctor_get_uint8(x_70, 8); +x_90 = lean_ctor_get_uint8(x_70, 10); +x_91 = lean_ctor_get_uint8(x_70, 11); +x_92 = lean_ctor_get_uint8(x_70, 12); +x_93 = lean_ctor_get_uint8(x_70, 13); +x_94 = lean_ctor_get_uint8(x_70, 14); +x_95 = lean_ctor_get_uint8(x_70, 15); +x_96 = lean_ctor_get_uint8(x_70, 16); +if (lean_is_exclusive(x_70)) { + x_97 = x_70; +} else { + lean_dec_ref(x_70); + x_97 = lean_box(0); +} +x_98 = 1; +if (lean_is_scalar(x_97)) { + x_99 = lean_alloc_ctor(0, 0, 17); +} else { + x_99 = x_97; +} +lean_ctor_set_uint8(x_99, 0, x_81); +lean_ctor_set_uint8(x_99, 1, x_82); +lean_ctor_set_uint8(x_99, 2, x_83); +lean_ctor_set_uint8(x_99, 3, x_84); +lean_ctor_set_uint8(x_99, 4, x_85); +lean_ctor_set_uint8(x_99, 5, x_86); +lean_ctor_set_uint8(x_99, 6, x_87); +lean_ctor_set_uint8(x_99, 7, x_88); +lean_ctor_set_uint8(x_99, 8, x_89); +lean_ctor_set_uint8(x_99, 9, x_98); +lean_ctor_set_uint8(x_99, 10, x_90); +lean_ctor_set_uint8(x_99, 11, x_91); +lean_ctor_set_uint8(x_99, 12, x_92); +lean_ctor_set_uint8(x_99, 13, x_93); +lean_ctor_set_uint8(x_99, 14, x_94); +lean_ctor_set_uint8(x_99, 15, x_95); +lean_ctor_set_uint8(x_99, 16, x_96); +x_100 = 2; +x_101 = lean_uint64_shift_right(x_71, x_100); +x_102 = lean_uint64_shift_left(x_101, x_100); +x_103 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___closed__2; +x_104 = lean_uint64_lor(x_102, x_103); +x_105 = lean_alloc_ctor(0, 7, 11); +lean_ctor_set(x_105, 0, x_99); +lean_ctor_set(x_105, 1, x_73); +lean_ctor_set(x_105, 2, x_74); +lean_ctor_set(x_105, 3, x_75); +lean_ctor_set(x_105, 4, x_76); +lean_ctor_set(x_105, 5, x_77); +lean_ctor_set(x_105, 6, x_78); +lean_ctor_set_uint64(x_105, sizeof(void*)*7, x_104); +lean_ctor_set_uint8(x_105, sizeof(void*)*7 + 8, x_72); +lean_ctor_set_uint8(x_105, sizeof(void*)*7 + 9, x_79); +lean_ctor_set_uint8(x_105, sizeof(void*)*7 + 10, x_80); +x_106 = 0; +x_107 = l_Lean_Meta_withNewMCtxDepth___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__12___rarg(x_15, x_106, x_2, x_3, x_4, x_5, x_6, x_7, x_105, x_9, x_10, x_11, x_12); +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); +if (lean_is_exclusive(x_107)) { + lean_ctor_release(x_107, 0); + lean_ctor_release(x_107, 1); + x_110 = x_107; +} else { + lean_dec_ref(x_107); + x_110 = lean_box(0); +} +if (lean_is_scalar(x_110)) { + x_111 = lean_alloc_ctor(0, 2, 0); +} else { + x_111 = x_110; +} +lean_ctor_set(x_111, 0, x_108); +lean_ctor_set(x_111, 1, x_109); +return x_111; +} +else +{ +lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; +x_112 = lean_ctor_get(x_107, 0); +lean_inc(x_112); +x_113 = lean_ctor_get(x_107, 1); +lean_inc(x_113); +if (lean_is_exclusive(x_107)) { + lean_ctor_release(x_107, 0); + lean_ctor_release(x_107, 1); + x_114 = x_107; +} else { + lean_dec_ref(x_107); + x_114 = lean_box(0); } +if (lean_is_scalar(x_114)) { + x_115 = lean_alloc_ctor(1, 2, 0); +} else { + x_115 = x_114; } +lean_ctor_set(x_115, 0, x_112); +lean_ctor_set(x_115, 1, x_113); +return x_115; } } -else -{ -lean_object* x_61; lean_object* x_62; lean_object* x_63; uint8_t x_64; -x_61 = lean_ctor_get(x_39, 0); -x_62 = lean_ctor_get(x_39, 1); -lean_inc(x_62); -lean_inc(x_61); -lean_dec(x_39); -x_63 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__6; -x_64 = lean_unbox(x_61); -lean_dec(x_61); -if (x_64 == 0) -{ -lean_object* x_65; lean_object* x_66; -lean_free_object(x_30); -lean_dec(x_1); -x_65 = lean_box(0); -x_66 = lean_apply_12(x_63, x_65, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_62); -return x_66; } -else -{ -lean_object* x_67; lean_object* x_68; -x_67 = lean_ctor_get(x_1, 5); -lean_inc(x_67); -lean_dec(x_1); -x_68 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_67, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_62); -if (lean_obj_tag(x_68) == 0) -{ -lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; -x_69 = lean_ctor_get(x_68, 0); -lean_inc(x_69); -x_70 = lean_ctor_get(x_68, 1); -lean_inc(x_70); -lean_dec(x_68); -x_71 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__8; -x_72 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_72, 0, x_71); -lean_ctor_set(x_72, 1, x_69); -x_73 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; -lean_ctor_set_tag(x_30, 7); -lean_ctor_set(x_30, 1, x_73); -lean_ctor_set(x_30, 0, x_72); -x_74 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_38, x_30, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, 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_apply_12(x_63, x_75, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_76); -return x_77; } -else +LEAN_EXPORT lean_object* l_ReaderT_read___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: { -lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; -lean_free_object(x_30); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); +lean_object* x_12; +x_12 = l_ReaderT_read___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -x_78 = lean_ctor_get(x_68, 0); -lean_inc(x_78); -x_79 = lean_ctor_get(x_68, 1); -lean_inc(x_79); -if (lean_is_exclusive(x_68)) { - lean_ctor_release(x_68, 0); - lean_ctor_release(x_68, 1); - x_80 = x_68; -} else { - lean_dec_ref(x_68); - x_80 = lean_box(0); -} -if (lean_is_scalar(x_80)) { - x_81 = lean_alloc_ctor(1, 2, 0); -} else { - x_81 = x_80; -} -lean_ctor_set(x_81, 0, x_78); -lean_ctor_set(x_81, 1, x_79); -return x_81; -} -} +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_12; } } -else +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: { -lean_object* x_82; lean_object* x_83; -lean_free_object(x_30); -x_82 = lean_box(0); -x_83 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__5(x_32, x_1, x_2, x_19, x_20, x_34, x_17, x_3, x_82, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_31); -lean_dec(x_20); -lean_dec(x_19); -lean_dec(x_32); -return x_83; +lean_object* x_13; +x_13 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_13; } } -else -{ -lean_object* x_84; lean_object* x_85; uint8_t x_86; -x_84 = lean_ctor_get(x_30, 0); -lean_inc(x_84); -lean_dec(x_30); -x_85 = lean_array_get_size(x_32); -x_86 = lean_nat_dec_eq(x_85, x_19); -lean_dec(x_85); -if (x_86 == 0) +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +lean_object* x_18 = _args[17]; +lean_object* x_19 = _args[18]; +lean_object* x_20 = _args[19]; +lean_object* x_21 = _args[20]; +lean_object* x_22 = _args[21]; +lean_object* x_23 = _args[22]; +_start: { -lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; uint8_t x_93; -lean_dec(x_84); -lean_dec(x_32); -lean_dec(x_20); -lean_dec(x_19); -lean_dec(x_17); +lean_object* x_24; +x_24 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_23); +lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); lean_dec(x_3); -x_87 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__2; -x_88 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__2(x_87, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_31); -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); +lean_dec(x_2); +return x_24; } -x_92 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__6; -x_93 = lean_unbox(x_89); -lean_dec(x_89); -if (x_93 == 0) -{ -lean_object* x_94; lean_object* x_95; -lean_dec(x_91); -lean_dec(x_1); -x_94 = lean_box(0); -x_95 = lean_apply_12(x_92, x_94, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_90); -return x_95; } -else +LEAN_EXPORT lean_object* l_Lean_MVarId_isAssigned___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: { -lean_object* x_96; lean_object* x_97; -x_96 = lean_ctor_get(x_1, 5); -lean_inc(x_96); +lean_object* x_13; +x_13 = l_Lean_MVarId_isAssigned___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_11); +lean_dec(x_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_97 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_96, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_90); -if (lean_obj_tag(x_97) == 0) -{ -lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; -x_98 = lean_ctor_get(x_97, 0); -lean_inc(x_98); -x_99 = lean_ctor_get(x_97, 1); -lean_inc(x_99); -lean_dec(x_97); -x_100 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__8; -if (lean_is_scalar(x_91)) { - x_101 = lean_alloc_ctor(7, 2, 0); -} else { - x_101 = x_91; - lean_ctor_set_tag(x_101, 7); +return x_13; } -lean_ctor_set(x_101, 0, x_100); -lean_ctor_set(x_101, 1, x_98); -x_102 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; -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_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_87, x_103, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_99); -x_105 = lean_ctor_get(x_104, 0); -lean_inc(x_105); -x_106 = lean_ctor_get(x_104, 1); -lean_inc(x_106); -lean_dec(x_104); -x_107 = lean_apply_12(x_92, x_105, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_106); -return x_107; } -else +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: { -lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; -lean_dec(x_91); -lean_dec(x_14); -lean_dec(x_13); +lean_object* x_14; +x_14 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); @@ -9375,121 +13919,139 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -x_108 = lean_ctor_get(x_97, 0); -lean_inc(x_108); -x_109 = lean_ctor_get(x_97, 1); -lean_inc(x_109); -if (lean_is_exclusive(x_97)) { - lean_ctor_release(x_97, 0); - lean_ctor_release(x_97, 1); - x_110 = x_97; -} else { - lean_dec_ref(x_97); - x_110 = lean_box(0); -} -if (lean_is_scalar(x_110)) { - x_111 = lean_alloc_ctor(1, 2, 0); -} else { - x_111 = x_110; -} -lean_ctor_set(x_111, 0, x_108); -lean_ctor_set(x_111, 1, x_109); -return x_111; -} +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_14; } } -else +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +lean_object* x_18 = _args[17]; +lean_object* x_19 = _args[18]; +_start: { -lean_object* x_112; lean_object* x_113; -x_112 = lean_box(0); -x_113 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__5(x_32, x_1, x_2, x_19, x_20, x_84, x_17, x_3, x_112, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_31); -lean_dec(x_20); -lean_dec(x_19); -lean_dec(x_32); -return x_113; -} +size_t x_20; size_t x_21; lean_object* x_22; +x_20 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_21 = lean_unbox_usize(x_7); +lean_dec(x_7); +x_22 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5(x_1, x_2, x_3, x_4, x_5, x_20, x_21, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +return x_22; } } -else +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__6___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +lean_object* x_18 = _args[17]; +_start: { -uint8_t x_114; -lean_dec(x_20); -lean_dec(x_19); -lean_dec(x_17); -lean_dec(x_14); +size_t x_19; size_t x_20; lean_object* x_21; +x_19 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_20 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_21 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__6(x_1, x_2, x_3, x_4, x_19, x_20, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); +lean_dec(x_4); +lean_dec(x_2); lean_dec(x_1); -x_114 = !lean_is_exclusive(x_28); -if (x_114 == 0) -{ -return x_28; -} -else -{ -lean_object* x_115; lean_object* x_116; lean_object* x_117; -x_115 = lean_ctor_get(x_28, 0); -x_116 = lean_ctor_get(x_28, 1); -lean_inc(x_116); -lean_inc(x_115); -lean_dec(x_28); -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; -} +return x_21; } } -else +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__7___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +lean_object* x_18 = _args[17]; +_start: { -uint8_t x_118; -lean_dec(x_20); -lean_dec(x_19); -lean_dec(x_17); -lean_dec(x_14); +size_t x_19; size_t x_20; lean_object* x_21; +x_19 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_20 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_21 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__7(x_1, x_2, x_3, x_4, x_19, x_20, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); +lean_dec(x_4); +lean_dec(x_2); lean_dec(x_1); -x_118 = !lean_is_exclusive(x_24); -if (x_118 == 0) -{ -return x_24; -} -else -{ -lean_object* x_119; lean_object* x_120; lean_object* x_121; -x_119 = lean_ctor_get(x_24, 0); -x_120 = lean_ctor_get(x_24, 1); -lean_inc(x_120); -lean_inc(x_119); -lean_dec(x_24); -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; -} -} +return x_21; } } -else +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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) { +_start: { -uint8_t x_122; +size_t x_16; size_t x_17; lean_object* x_18; +x_16 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_17 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_18 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__8(x_1, x_16, x_17, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); @@ -9500,211 +14062,58 @@ 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_122 = !lean_is_exclusive(x_16); -if (x_122 == 0) -{ -return x_16; -} -else -{ -lean_object* x_123; lean_object* x_124; lean_object* x_125; -x_123 = lean_ctor_get(x_16, 0); -x_124 = lean_ctor_get(x_16, 1); -lean_inc(x_124); -lean_inc(x_123); -lean_dec(x_16); -x_125 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_125, 0, x_123); -lean_ctor_set(x_125, 1, x_124); -return x_125; -} -} -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__8___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("assignment", 10, 10); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__8___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_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__1; -x_2 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__2; -x_3 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__3; -x_4 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__8___closed__1; -x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); -return x_5; +return x_18; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__9___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +lean_object* x_18 = _args[17]; _start: { -lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; -x_16 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__8___closed__2; -x_17 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__2(x_16, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); -x_18 = lean_ctor_get(x_17, 0); -lean_inc(x_18); -x_19 = lean_unbox(x_18); -lean_dec(x_18); -if (x_19 == 0) -{ -lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_20 = lean_ctor_get(x_17, 1); -lean_inc(x_20); -lean_dec(x_17); -x_21 = lean_box(0); -x_22 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7(x_1, x_2, x_3, x_21, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_20); -lean_dec(x_2); -return x_22; -} -else -{ -uint8_t x_23; -x_23 = !lean_is_exclusive(x_17); -if (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_17, 1); -x_25 = lean_ctor_get(x_17, 0); -lean_dec(x_25); -x_26 = lean_ctor_get(x_1, 5); -lean_inc(x_26); -x_27 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_26, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_24); -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; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; -x_28 = lean_ctor_get(x_27, 0); -lean_inc(x_28); -x_29 = lean_ctor_get(x_27, 1); -lean_inc(x_29); -lean_dec(x_27); -x_30 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; -lean_ctor_set_tag(x_17, 7); -lean_ctor_set(x_17, 1, x_28); -lean_ctor_set(x_17, 0, x_30); -x_31 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__6; -x_32 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_32, 0, x_17); -lean_ctor_set(x_32, 1, x_31); -lean_inc(x_2); -x_33 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData(x_2); -x_34 = lean_array_to_list(x_33); -x_35 = lean_box(0); -x_36 = l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__4(x_34, x_35); -x_37 = l_Lean_MessageData_ofList(x_36); -x_38 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_38, 0, x_32); -lean_ctor_set(x_38, 1, x_37); -x_39 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_39, 0, x_38); -lean_ctor_set(x_39, 1, x_30); -x_40 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_16, x_39, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_29); -x_41 = lean_ctor_get(x_40, 0); -lean_inc(x_41); -x_42 = lean_ctor_get(x_40, 1); -lean_inc(x_42); -lean_dec(x_40); -x_43 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7(x_1, x_2, x_3, x_41, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_42); -lean_dec(x_41); -lean_dec(x_2); -return x_43; -} -else -{ -uint8_t x_44; -lean_free_object(x_17); -lean_dec(x_14); +size_t x_19; size_t x_20; lean_object* x_21; +x_19 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_20 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_21 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__9(x_1, x_2, x_3, x_4, x_19, x_20, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); +lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_44 = !lean_is_exclusive(x_27); -if (x_44 == 0) -{ -return x_27; -} -else -{ -lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_45 = lean_ctor_get(x_27, 0); -x_46 = lean_ctor_get(x_27, 1); -lean_inc(x_46); -lean_inc(x_45); -lean_dec(x_27); -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; -} -} +return x_21; } -else -{ -lean_object* x_48; lean_object* x_49; lean_object* x_50; -x_48 = lean_ctor_get(x_17, 1); -lean_inc(x_48); -lean_dec(x_17); -x_49 = lean_ctor_get(x_1, 5); -lean_inc(x_49); -x_50 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_49, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_48); -if (lean_obj_tag(x_50) == 0) -{ -lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; 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_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_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; -x_54 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_54, 0, x_53); -lean_ctor_set(x_54, 1, x_51); -x_55 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__6; -x_56 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_56, 0, x_54); -lean_ctor_set(x_56, 1, x_55); -lean_inc(x_2); -x_57 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData(x_2); -x_58 = lean_array_to_list(x_57); -x_59 = lean_box(0); -x_60 = l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__4(x_58, x_59); -x_61 = l_Lean_MessageData_ofList(x_60); -x_62 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_62, 0, x_56); -lean_ctor_set(x_62, 1, x_61); -x_63 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_63, 0, x_62); -lean_ctor_set(x_63, 1, x_53); -x_64 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_16, x_63, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_52); -x_65 = lean_ctor_get(x_64, 0); -lean_inc(x_65); -x_66 = lean_ctor_get(x_64, 1); -lean_inc(x_66); -lean_dec(x_64); -x_67 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7(x_1, x_2, x_3, x_65, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_66); -lean_dec(x_65); -lean_dec(x_2); -return x_67; } -else +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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) { +_start: { -lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; -lean_dec(x_14); +size_t x_15; size_t x_16; lean_object* x_17; +x_15 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_16 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_17 = l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__10(x_1, x_15, x_16, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); @@ -9714,669 +14123,821 @@ 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_4); lean_dec(x_1); -x_68 = lean_ctor_get(x_50, 0); -lean_inc(x_68); -x_69 = lean_ctor_get(x_50, 1); -lean_inc(x_69); -if (lean_is_exclusive(x_50)) { - lean_ctor_release(x_50, 0); - lean_ctor_release(x_50, 1); - x_70 = x_50; -} else { - lean_dec_ref(x_50); - x_70 = lean_box(0); -} -if (lean_is_scalar(x_70)) { - x_71 = lean_alloc_ctor(1, 2, 0); -} else { - x_71 = x_70; +return x_17; } -lean_ctor_set(x_71, 0, x_68); -lean_ctor_set(x_71, 1, x_69); -return x_71; } +LEAN_EXPORT lean_object* l_Lean_Meta_withNewMCtxDepth___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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, 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_2); +lean_dec(x_2); +x_15 = l_Lean_Meta_withNewMCtxDepth___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__12___rarg(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); +return x_15; } } +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +_start: +{ +lean_object* x_17; +x_17 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +return x_17; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +lean_object* x_18 = _args[17]; +lean_object* x_19 = _args[18]; +lean_object* x_20 = _args[19]; +lean_object* x_21 = _args[20]; _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; -x_14 = lean_ctor_get(x_2, 0); -lean_inc(x_14); +size_t x_22; lean_object* x_23; +x_22 = lean_unbox_usize(x_8); +lean_dec(x_8); +x_23 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_22, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21); +lean_dec(x_10); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); lean_dec(x_2); -x_15 = lean_ctor_get(x_14, 1); -lean_inc(x_15); -x_16 = lean_ctor_get(x_1, 2); -lean_inc(x_16); -lean_inc(x_16); -x_17 = l_Lean_Meta_Grind_markTheoremInstance(x_15, x_16, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -x_18 = lean_ctor_get(x_17, 0); -lean_inc(x_18); -x_19 = lean_unbox(x_18); -lean_dec(x_18); -if (x_19 == 0) +return x_23; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +lean_object* x_18 = _args[17]; +lean_object* x_19 = _args[18]; +lean_object* x_20 = _args[19]; +_start: { -uint8_t x_20; -lean_dec(x_16); -lean_dec(x_14); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); +lean_object* x_21; +x_21 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20); lean_dec(x_9); -lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); +lean_dec(x_3); +return x_21; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__4___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_object* x_19 = _args[18]; +lean_object* x_20 = _args[19]; +_start: +{ +lean_object* x_21; +x_21 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_20 = !lean_is_exclusive(x_17); -if (x_20 == 0) -{ -lean_object* x_21; lean_object* x_22; -x_21 = lean_ctor_get(x_17, 0); -lean_dec(x_21); -x_22 = lean_box(0); -lean_ctor_set(x_17, 0, x_22); -return x_17; -} -else -{ -lean_object* x_23; lean_object* x_24; lean_object* x_25; -x_23 = lean_ctor_get(x_17, 1); -lean_inc(x_23); -lean_dec(x_17); -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; +return x_21; } } -else +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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) { +_start: { -lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_26 = lean_ctor_get(x_17, 1); -lean_inc(x_26); -lean_dec(x_17); -x_27 = lean_box(0); -x_28 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__8(x_14, x_16, x_1, x_27, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_26); -return x_28; -} +lean_object* x_13; +x_13 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_13; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___closed__1() { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { _start: { -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_ReaderT_read___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__1___boxed), 11, 0); -return x_1; +lean_object* x_16; +x_16 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_16; } } -static uint64_t _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___closed__2() { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { _start: { -uint8_t x_1; uint64_t x_2; -x_1 = 1; -x_2 = l_Lean_Meta_TransparencyMode_toUInt64(x_1); -return x_2; +lean_object* x_16; +x_16 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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); +lean_dec(x_4); +lean_dec(x_3); +return x_16; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_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_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___spec__1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { _start: { -lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; -x_13 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__9), 13, 1); -lean_closure_set(x_13, 0, x_1); -x_14 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___closed__1; -x_15 = lean_alloc_closure((void*)(l_ReaderT_bind___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__11___rarg), 13, 2); -lean_closure_set(x_15, 0, x_14); -lean_closure_set(x_15, 1, x_13); -x_16 = !lean_is_exclusive(x_8); -if (x_16 == 0) -{ -lean_object* x_17; uint8_t x_18; -x_17 = lean_ctor_get(x_8, 0); -x_18 = !lean_is_exclusive(x_17); -if (x_18 == 0) -{ -uint64_t x_19; uint8_t x_20; uint64_t x_21; uint64_t x_22; uint64_t x_23; uint64_t x_24; uint64_t x_25; uint8_t x_26; lean_object* x_27; -x_19 = lean_ctor_get_uint64(x_8, sizeof(void*)*7); -x_20 = 1; -lean_ctor_set_uint8(x_17, 9, x_20); -x_21 = 2; -x_22 = lean_uint64_shift_right(x_19, x_21); -x_23 = lean_uint64_shift_left(x_22, x_21); -x_24 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___closed__2; -x_25 = lean_uint64_lor(x_23, x_24); -lean_ctor_set_uint64(x_8, sizeof(void*)*7, x_25); -x_26 = 0; -x_27 = l_Lean_Meta_withNewMCtxDepth___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__12___rarg(x_15, x_26, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_27) == 0) +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_15 = lean_st_ref_take(x_5, x_14); +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_15, 1); +lean_inc(x_17); +lean_dec(x_15); +x_18 = l_Lean_Meta_Grind_EMatch_instInhabitedChoice; +x_19 = l_List_head_x21___rarg(x_18, x_16); +x_20 = l_List_tail_x21___rarg(x_16); +lean_dec(x_16); +x_21 = lean_st_ref_set(x_5, x_20, x_17); +x_22 = !lean_is_exclusive(x_21); +if (x_22 == 0) { -uint8_t x_28; -x_28 = !lean_is_exclusive(x_27); +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; +x_23 = lean_ctor_get(x_21, 1); +x_24 = lean_ctor_get(x_21, 0); +lean_dec(x_24); +x_25 = lean_ctor_get(x_19, 0); +lean_inc(x_25); +x_26 = lean_ctor_get(x_19, 1); +lean_inc(x_26); +x_27 = lean_ctor_get(x_19, 2); +lean_inc(x_27); +x_28 = lean_nat_dec_lt(x_26, x_1); if (x_28 == 0) { -return x_27; +lean_object* x_29; +lean_dec(x_27); +lean_dec(x_26); +lean_dec(x_25); +lean_dec(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_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_29 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_29, 0, x_2); +lean_ctor_set(x_21, 0, x_29); +return x_21; } 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_inc(x_29); +lean_free_object(x_21); +if (lean_obj_tag(x_25) == 0) +{ +lean_object* x_30; 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; -} -} -else +lean_dec(x_26); +x_30 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem(x_19, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_23); +if (lean_obj_tag(x_30) == 0) { -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_30); +if (x_31 == 0) { -return x_27; +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_2); +lean_ctor_set(x_30, 0, x_33); +return x_30; } 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_30, 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_dec(x_30); +x_35 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_35, 0, x_2); +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set(x_36, 1, x_34); +return x_36; } } else { -uint64_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; 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; uint64_t x_55; uint64_t x_56; uint64_t x_57; uint64_t x_58; uint64_t x_59; uint8_t x_60; lean_object* x_61; -x_36 = lean_ctor_get_uint64(x_8, sizeof(void*)*7); -x_37 = lean_ctor_get_uint8(x_17, 0); -x_38 = lean_ctor_get_uint8(x_17, 1); -x_39 = lean_ctor_get_uint8(x_17, 2); -x_40 = lean_ctor_get_uint8(x_17, 3); -x_41 = lean_ctor_get_uint8(x_17, 4); -x_42 = lean_ctor_get_uint8(x_17, 5); -x_43 = lean_ctor_get_uint8(x_17, 6); -x_44 = lean_ctor_get_uint8(x_17, 7); -x_45 = lean_ctor_get_uint8(x_17, 8); -x_46 = lean_ctor_get_uint8(x_17, 10); -x_47 = lean_ctor_get_uint8(x_17, 11); -x_48 = lean_ctor_get_uint8(x_17, 12); -x_49 = lean_ctor_get_uint8(x_17, 13); -x_50 = lean_ctor_get_uint8(x_17, 14); -x_51 = lean_ctor_get_uint8(x_17, 15); -x_52 = lean_ctor_get_uint8(x_17, 16); -lean_dec(x_17); -x_53 = 1; -x_54 = lean_alloc_ctor(0, 0, 17); -lean_ctor_set_uint8(x_54, 0, x_37); -lean_ctor_set_uint8(x_54, 1, x_38); -lean_ctor_set_uint8(x_54, 2, x_39); -lean_ctor_set_uint8(x_54, 3, x_40); -lean_ctor_set_uint8(x_54, 4, x_41); -lean_ctor_set_uint8(x_54, 5, x_42); -lean_ctor_set_uint8(x_54, 6, x_43); -lean_ctor_set_uint8(x_54, 7, x_44); -lean_ctor_set_uint8(x_54, 8, x_45); -lean_ctor_set_uint8(x_54, 9, x_53); -lean_ctor_set_uint8(x_54, 10, x_46); -lean_ctor_set_uint8(x_54, 11, x_47); -lean_ctor_set_uint8(x_54, 12, x_48); -lean_ctor_set_uint8(x_54, 13, x_49); -lean_ctor_set_uint8(x_54, 14, x_50); -lean_ctor_set_uint8(x_54, 15, x_51); -lean_ctor_set_uint8(x_54, 16, x_52); -x_55 = 2; -x_56 = lean_uint64_shift_right(x_36, x_55); -x_57 = lean_uint64_shift_left(x_56, x_55); -x_58 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___closed__2; -x_59 = lean_uint64_lor(x_57, x_58); -lean_ctor_set(x_8, 0, x_54); -lean_ctor_set_uint64(x_8, sizeof(void*)*7, x_59); -x_60 = 0; -x_61 = l_Lean_Meta_withNewMCtxDepth___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__12___rarg(x_15, x_60, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_61) == 0) +uint8_t x_37; +lean_dec(x_2); +x_37 = !lean_is_exclusive(x_30); +if (x_37 == 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; +return x_30; } else { -lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; -x_66 = lean_ctor_get(x_61, 0); -lean_inc(x_66); -x_67 = lean_ctor_get(x_61, 1); -lean_inc(x_67); -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_61); - x_68 = lean_box(0); -} -if (lean_is_scalar(x_68)) { - x_69 = lean_alloc_ctor(1, 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; +lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_38 = lean_ctor_get(x_30, 0); +x_39 = lean_ctor_get(x_30, 1); +lean_inc(x_39); +lean_inc(x_38); +lean_dec(x_30); +x_40 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_40, 0, x_38); +lean_ctor_set(x_40, 1, x_39); +return x_40; } } } else { -lean_object* x_70; uint64_t x_71; 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; 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; 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; lean_object* x_97; uint8_t x_98; lean_object* x_99; uint64_t x_100; uint64_t x_101; uint64_t x_102; uint64_t x_103; uint64_t x_104; lean_object* x_105; uint8_t x_106; lean_object* x_107; -x_70 = lean_ctor_get(x_8, 0); -x_71 = lean_ctor_get_uint64(x_8, sizeof(void*)*7); -x_72 = lean_ctor_get_uint8(x_8, sizeof(void*)*7 + 8); -x_73 = lean_ctor_get(x_8, 1); -x_74 = lean_ctor_get(x_8, 2); -x_75 = lean_ctor_get(x_8, 3); -x_76 = lean_ctor_get(x_8, 4); -x_77 = lean_ctor_get(x_8, 5); -x_78 = lean_ctor_get(x_8, 6); -x_79 = lean_ctor_get_uint8(x_8, sizeof(void*)*7 + 9); -x_80 = lean_ctor_get_uint8(x_8, sizeof(void*)*7 + 10); -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_70); -lean_dec(x_8); -x_81 = lean_ctor_get_uint8(x_70, 0); -x_82 = lean_ctor_get_uint8(x_70, 1); -x_83 = lean_ctor_get_uint8(x_70, 2); -x_84 = lean_ctor_get_uint8(x_70, 3); -x_85 = lean_ctor_get_uint8(x_70, 4); -x_86 = lean_ctor_get_uint8(x_70, 5); -x_87 = lean_ctor_get_uint8(x_70, 6); -x_88 = lean_ctor_get_uint8(x_70, 7); -x_89 = lean_ctor_get_uint8(x_70, 8); -x_90 = lean_ctor_get_uint8(x_70, 10); -x_91 = lean_ctor_get_uint8(x_70, 11); -x_92 = lean_ctor_get_uint8(x_70, 12); -x_93 = lean_ctor_get_uint8(x_70, 13); -x_94 = lean_ctor_get_uint8(x_70, 14); -x_95 = lean_ctor_get_uint8(x_70, 15); -x_96 = lean_ctor_get_uint8(x_70, 16); -if (lean_is_exclusive(x_70)) { - x_97 = x_70; -} else { - lean_dec_ref(x_70); - x_97 = lean_box(0); -} -x_98 = 1; -if (lean_is_scalar(x_97)) { - x_99 = lean_alloc_ctor(0, 0, 17); -} else { - x_99 = x_97; -} -lean_ctor_set_uint8(x_99, 0, x_81); -lean_ctor_set_uint8(x_99, 1, x_82); -lean_ctor_set_uint8(x_99, 2, x_83); -lean_ctor_set_uint8(x_99, 3, x_84); -lean_ctor_set_uint8(x_99, 4, x_85); -lean_ctor_set_uint8(x_99, 5, x_86); -lean_ctor_set_uint8(x_99, 6, x_87); -lean_ctor_set_uint8(x_99, 7, x_88); -lean_ctor_set_uint8(x_99, 8, x_89); -lean_ctor_set_uint8(x_99, 9, x_98); -lean_ctor_set_uint8(x_99, 10, x_90); -lean_ctor_set_uint8(x_99, 11, x_91); -lean_ctor_set_uint8(x_99, 12, x_92); -lean_ctor_set_uint8(x_99, 13, x_93); -lean_ctor_set_uint8(x_99, 14, x_94); -lean_ctor_set_uint8(x_99, 15, x_95); -lean_ctor_set_uint8(x_99, 16, x_96); -x_100 = 2; -x_101 = lean_uint64_shift_right(x_71, x_100); -x_102 = lean_uint64_shift_left(x_101, x_100); -x_103 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___closed__2; -x_104 = lean_uint64_lor(x_102, x_103); -x_105 = lean_alloc_ctor(0, 7, 11); -lean_ctor_set(x_105, 0, x_99); -lean_ctor_set(x_105, 1, x_73); -lean_ctor_set(x_105, 2, x_74); -lean_ctor_set(x_105, 3, x_75); -lean_ctor_set(x_105, 4, x_76); -lean_ctor_set(x_105, 5, x_77); -lean_ctor_set(x_105, 6, x_78); -lean_ctor_set_uint64(x_105, sizeof(void*)*7, x_104); -lean_ctor_set_uint8(x_105, sizeof(void*)*7 + 8, x_72); -lean_ctor_set_uint8(x_105, sizeof(void*)*7 + 9, x_79); -lean_ctor_set_uint8(x_105, sizeof(void*)*7 + 10, x_80); -x_106 = 0; -x_107 = l_Lean_Meta_withNewMCtxDepth___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__12___rarg(x_15, x_106, x_2, x_3, x_4, x_5, x_6, x_7, x_105, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_107) == 0) +uint8_t x_41; +x_41 = !lean_is_exclusive(x_19); +if (x_41 == 0) +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_42 = lean_ctor_get(x_19, 2); +lean_dec(x_42); +x_43 = lean_ctor_get(x_19, 1); +lean_dec(x_43); +x_44 = lean_ctor_get(x_19, 0); +lean_dec(x_44); +x_45 = lean_ctor_get(x_25, 0); +lean_inc(x_45); +switch (lean_obj_tag(x_45)) { +case 0: +{ +lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_46 = lean_ctor_get(x_25, 1); +lean_inc(x_46); +lean_dec(x_25); +x_47 = lean_ctor_get(x_45, 0); +lean_inc(x_47); +x_48 = lean_ctor_get(x_45, 1); +lean_inc(x_48); +lean_dec(x_45); +lean_ctor_set(x_19, 0, x_46); +x_49 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch(x_19, x_47, x_48, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_23); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_47); +if (lean_obj_tag(x_49) == 0) +{ +uint8_t x_50; +x_50 = !lean_is_exclusive(x_49); +if (x_50 == 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); -if (lean_is_exclusive(x_107)) { - lean_ctor_release(x_107, 0); - lean_ctor_release(x_107, 1); - x_110 = x_107; -} else { - lean_dec_ref(x_107); - x_110 = lean_box(0); -} -if (lean_is_scalar(x_110)) { - x_111 = lean_alloc_ctor(0, 2, 0); -} else { - x_111 = x_110; -} -lean_ctor_set(x_111, 0, x_108); -lean_ctor_set(x_111, 1, x_109); -return x_111; +lean_object* x_51; lean_object* x_52; +x_51 = lean_ctor_get(x_49, 0); +lean_dec(x_51); +x_52 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_52, 0, x_2); +lean_ctor_set(x_49, 0, x_52); +return x_49; } else { -lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; -x_112 = lean_ctor_get(x_107, 0); -lean_inc(x_112); -x_113 = lean_ctor_get(x_107, 1); -lean_inc(x_113); -if (lean_is_exclusive(x_107)) { - lean_ctor_release(x_107, 0); - lean_ctor_release(x_107, 1); - x_114 = x_107; -} else { - lean_dec_ref(x_107); - x_114 = lean_box(0); +lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_53 = lean_ctor_get(x_49, 1); +lean_inc(x_53); +lean_dec(x_49); +x_54 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_54, 0, x_2); +x_55 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_55, 0, x_54); +lean_ctor_set(x_55, 1, x_53); +return x_55; } -if (lean_is_scalar(x_114)) { - x_115 = lean_alloc_ctor(1, 2, 0); -} else { - x_115 = x_114; } -lean_ctor_set(x_115, 0, x_112); -lean_ctor_set(x_115, 1, x_113); -return x_115; +else +{ +uint8_t x_56; +lean_dec(x_2); +x_56 = !lean_is_exclusive(x_49); +if (x_56 == 0) +{ +return x_49; } +else +{ +lean_object* x_57; lean_object* x_58; lean_object* x_59; +x_57 = lean_ctor_get(x_49, 0); +x_58 = lean_ctor_get(x_49, 1); +lean_inc(x_58); +lean_inc(x_57); +lean_dec(x_49); +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_EXPORT lean_object* l_ReaderT_read___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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: +case 1: { -lean_object* x_12; -x_12 = l_ReaderT_read___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); +lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; +x_60 = lean_ctor_get(x_25, 1); +lean_inc(x_60); +lean_dec(x_25); +x_61 = lean_ctor_get(x_45, 0); +lean_inc(x_61); +x_62 = lean_ctor_get(x_45, 1); +lean_inc(x_62); +x_63 = lean_ctor_get(x_45, 2); +lean_inc(x_63); +lean_dec(x_45); +lean_ctor_set(x_19, 0, x_60); +x_64 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processOffset(x_19, x_61, x_62, x_63, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_23); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_12; +lean_dec(x_62); +if (lean_obj_tag(x_64) == 0) +{ +uint8_t x_65; +x_65 = !lean_is_exclusive(x_64); +if (x_65 == 0) +{ +lean_object* x_66; lean_object* x_67; +x_66 = lean_ctor_get(x_64, 0); +lean_dec(x_66); +x_67 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_67, 0, x_2); +lean_ctor_set(x_64, 0, x_67); +return x_64; +} +else +{ +lean_object* x_68; lean_object* x_69; lean_object* x_70; +x_68 = lean_ctor_get(x_64, 1); +lean_inc(x_68); +lean_dec(x_64); +x_69 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_69, 0, x_2); +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_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: +else { -lean_object* x_13; -x_13 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_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); +uint8_t x_71; lean_dec(x_2); -lean_dec(x_1); -return x_13; +x_71 = !lean_is_exclusive(x_64); +if (x_71 == 0) +{ +return x_64; } +else +{ +lean_object* x_72; lean_object* x_73; lean_object* x_74; +x_72 = lean_ctor_get(x_64, 0); +x_73 = lean_ctor_get(x_64, 1); +lean_inc(x_73); +lean_inc(x_72); +lean_dec(x_64); +x_74 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_74, 0, x_72); +lean_ctor_set(x_74, 1, x_73); +return x_74; } -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___boxed(lean_object** _args) { -lean_object* x_1 = _args[0]; -lean_object* x_2 = _args[1]; -lean_object* x_3 = _args[2]; -lean_object* x_4 = _args[3]; -lean_object* x_5 = _args[4]; -lean_object* x_6 = _args[5]; -lean_object* x_7 = _args[6]; -lean_object* x_8 = _args[7]; -lean_object* x_9 = _args[8]; -lean_object* x_10 = _args[9]; -lean_object* x_11 = _args[10]; -lean_object* x_12 = _args[11]; -lean_object* x_13 = _args[12]; -lean_object* x_14 = _args[13]; -lean_object* x_15 = _args[14]; -lean_object* x_16 = _args[15]; -lean_object* x_17 = _args[16]; -lean_object* x_18 = _args[17]; -lean_object* x_19 = _args[18]; -lean_object* x_20 = _args[19]; -lean_object* x_21 = _args[20]; -lean_object* x_22 = _args[21]; -lean_object* x_23 = _args[22]; -_start: +} +} +default: { -lean_object* x_24; -x_24 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_23); -lean_dec(x_8); -lean_dec(x_6); +lean_object* x_75; uint8_t x_76; +x_75 = lean_ctor_get(x_25, 1); +lean_inc(x_75); +lean_dec(x_25); +x_76 = !lean_is_exclusive(x_45); +if (x_76 == 0) +{ +lean_object* x_77; lean_object* x_78; +x_77 = lean_ctor_get(x_45, 0); +lean_ctor_set(x_19, 0, x_75); +x_78 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue(x_19, x_77, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_23); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_24; +if (lean_obj_tag(x_78) == 0) +{ +uint8_t x_79; +x_79 = !lean_is_exclusive(x_78); +if (x_79 == 0) +{ +lean_object* x_80; +x_80 = lean_ctor_get(x_78, 0); +lean_dec(x_80); +lean_ctor_set_tag(x_45, 1); +lean_ctor_set(x_45, 0, x_2); +lean_ctor_set(x_78, 0, x_45); +return x_78; +} +else +{ +lean_object* x_81; lean_object* x_82; +x_81 = lean_ctor_get(x_78, 1); +lean_inc(x_81); +lean_dec(x_78); +lean_ctor_set_tag(x_45, 1); +lean_ctor_set(x_45, 0, x_2); +x_82 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_82, 0, x_45); +lean_ctor_set(x_82, 1, x_81); +return x_82; } } -LEAN_EXPORT lean_object* l_Lean_MVarId_isAssigned___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: +else { -lean_object* x_13; -x_13 = l_Lean_MVarId_isAssigned___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_11); -lean_dec(x_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); +uint8_t x_83; +lean_free_object(x_45); lean_dec(x_2); -lean_dec(x_1); -return x_13; +x_83 = !lean_is_exclusive(x_78); +if (x_83 == 0) +{ +return x_78; +} +else +{ +lean_object* x_84; lean_object* x_85; lean_object* x_86; +x_84 = lean_ctor_get(x_78, 0); +x_85 = lean_ctor_get(x_78, 1); +lean_inc(x_85); +lean_inc(x_84); +lean_dec(x_78); +x_86 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_86, 0, x_84); +lean_ctor_set(x_86, 1, x_85); +return x_86; } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { -_start: +} +else { -lean_object* x_14; -x_14 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); +lean_object* x_87; lean_object* x_88; +x_87 = lean_ctor_get(x_45, 0); +lean_inc(x_87); +lean_dec(x_45); +lean_ctor_set(x_19, 0, x_75); +x_88 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue(x_19, x_87, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_23); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); +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, 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 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_91, 0, x_2); +if (lean_is_scalar(x_90)) { + x_92 = lean_alloc_ctor(0, 2, 0); +} else { + x_92 = x_90; +} +lean_ctor_set(x_92, 0, x_91); +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_dec(x_2); -return x_14; +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_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___boxed(lean_object** _args) { -lean_object* x_1 = _args[0]; -lean_object* x_2 = _args[1]; -lean_object* x_3 = _args[2]; -lean_object* x_4 = _args[3]; -lean_object* x_5 = _args[4]; -lean_object* x_6 = _args[5]; -lean_object* x_7 = _args[6]; -lean_object* x_8 = _args[7]; -lean_object* x_9 = _args[8]; -lean_object* x_10 = _args[9]; -lean_object* x_11 = _args[10]; -lean_object* x_12 = _args[11]; -lean_object* x_13 = _args[12]; -lean_object* x_14 = _args[13]; -lean_object* x_15 = _args[14]; -lean_object* x_16 = _args[15]; -lean_object* x_17 = _args[16]; -lean_object* x_18 = _args[17]; -lean_object* x_19 = _args[18]; -_start: +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_dec(x_19); +x_97 = lean_ctor_get(x_25, 0); +lean_inc(x_97); +switch (lean_obj_tag(x_97)) { +case 0: { -size_t x_20; size_t x_21; lean_object* x_22; -x_20 = lean_unbox_usize(x_6); -lean_dec(x_6); -x_21 = lean_unbox_usize(x_7); -lean_dec(x_7); -x_22 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5(x_1, x_2, x_3, x_4, x_5, x_20, x_21, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); +lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; +x_98 = lean_ctor_get(x_25, 1); +lean_inc(x_98); +lean_dec(x_25); +x_99 = lean_ctor_get(x_97, 0); +lean_inc(x_99); +x_100 = lean_ctor_get(x_97, 1); +lean_inc(x_100); +lean_dec(x_97); +x_101 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_101, 0, x_98); +lean_ctor_set(x_101, 1, x_26); +lean_ctor_set(x_101, 2, x_27); +x_102 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch(x_101, x_99, x_100, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_23); lean_dec(x_5); lean_dec(x_4); +lean_dec(x_99); +if (lean_obj_tag(x_102) == 0) +{ +lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; +x_103 = lean_ctor_get(x_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_102); + x_104 = lean_box(0); +} +x_105 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_105, 0, x_2); +if (lean_is_scalar(x_104)) { + x_106 = lean_alloc_ctor(0, 2, 0); +} else { + x_106 = x_104; +} +lean_ctor_set(x_106, 0, x_105); +lean_ctor_set(x_106, 1, x_103); +return x_106; +} +else +{ +lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_dec(x_2); -return x_22; +x_107 = lean_ctor_get(x_102, 0); +lean_inc(x_107); +x_108 = lean_ctor_get(x_102, 1); +lean_inc(x_108); +if (lean_is_exclusive(x_102)) { + lean_ctor_release(x_102, 0); + lean_ctor_release(x_102, 1); + x_109 = x_102; +} else { + lean_dec_ref(x_102); + 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_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__6___boxed(lean_object** _args) { -lean_object* x_1 = _args[0]; -lean_object* x_2 = _args[1]; -lean_object* x_3 = _args[2]; -lean_object* x_4 = _args[3]; -lean_object* x_5 = _args[4]; -lean_object* x_6 = _args[5]; -lean_object* x_7 = _args[6]; -lean_object* x_8 = _args[7]; -lean_object* x_9 = _args[8]; -lean_object* x_10 = _args[9]; -lean_object* x_11 = _args[10]; -lean_object* x_12 = _args[11]; -lean_object* x_13 = _args[12]; -lean_object* x_14 = _args[13]; -lean_object* x_15 = _args[14]; -lean_object* x_16 = _args[15]; -lean_object* x_17 = _args[16]; -lean_object* x_18 = _args[17]; -_start: +lean_ctor_set(x_110, 0, x_107); +lean_ctor_set(x_110, 1, x_108); +return x_110; +} +} +case 1: { -size_t x_19; size_t x_20; lean_object* x_21; -x_19 = lean_unbox_usize(x_5); +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_25, 1); +lean_inc(x_111); +lean_dec(x_25); +x_112 = lean_ctor_get(x_97, 0); +lean_inc(x_112); +x_113 = lean_ctor_get(x_97, 1); +lean_inc(x_113); +x_114 = lean_ctor_get(x_97, 2); +lean_inc(x_114); +lean_dec(x_97); +x_115 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_115, 0, x_111); +lean_ctor_set(x_115, 1, x_26); +lean_ctor_set(x_115, 2, x_27); +x_116 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processOffset(x_115, x_112, x_113, x_114, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_23); lean_dec(x_5); -x_20 = lean_unbox_usize(x_6); -lean_dec(x_6); -x_21 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__6(x_1, x_2, x_3, x_4, x_19, x_20, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); lean_dec(x_4); +lean_dec(x_113); +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 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_119, 0, x_2); +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; +} +else +{ +lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_dec(x_2); -lean_dec(x_1); -return x_21; +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_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__7___boxed(lean_object** _args) { -lean_object* x_1 = _args[0]; -lean_object* x_2 = _args[1]; -lean_object* x_3 = _args[2]; -lean_object* x_4 = _args[3]; -lean_object* x_5 = _args[4]; -lean_object* x_6 = _args[5]; -lean_object* x_7 = _args[6]; -lean_object* x_8 = _args[7]; -lean_object* x_9 = _args[8]; -lean_object* x_10 = _args[9]; -lean_object* x_11 = _args[10]; -lean_object* x_12 = _args[11]; -lean_object* x_13 = _args[12]; -lean_object* x_14 = _args[13]; -lean_object* x_15 = _args[14]; -lean_object* x_16 = _args[15]; -lean_object* x_17 = _args[16]; -lean_object* x_18 = _args[17]; -_start: +lean_ctor_set(x_124, 0, x_121); +lean_ctor_set(x_124, 1, x_122); +return x_124; +} +} +default: { -size_t x_19; size_t x_20; lean_object* x_21; -x_19 = lean_unbox_usize(x_5); +lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; +x_125 = lean_ctor_get(x_25, 1); +lean_inc(x_125); +lean_dec(x_25); +x_126 = lean_ctor_get(x_97, 0); +lean_inc(x_126); +if (lean_is_exclusive(x_97)) { + lean_ctor_release(x_97, 0); + x_127 = x_97; +} else { + lean_dec_ref(x_97); + x_127 = lean_box(0); +} +x_128 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_128, 0, x_125); +lean_ctor_set(x_128, 1, x_26); +lean_ctor_set(x_128, 2, x_27); +x_129 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue(x_128, x_126, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_23); lean_dec(x_5); -x_20 = lean_unbox_usize(x_6); -lean_dec(x_6); -x_21 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__7(x_1, x_2, x_3, x_4, x_19, x_20, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); lean_dec(x_4); -lean_dec(x_2); -lean_dec(x_1); -return x_21; +if (lean_obj_tag(x_129) == 0) +{ +lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; +x_130 = lean_ctor_get(x_129, 1); +lean_inc(x_130); +if (lean_is_exclusive(x_129)) { + lean_ctor_release(x_129, 0); + lean_ctor_release(x_129, 1); + x_131 = x_129; +} else { + lean_dec_ref(x_129); + x_131 = lean_box(0); } +if (lean_is_scalar(x_127)) { + x_132 = lean_alloc_ctor(1, 1, 0); +} else { + x_132 = x_127; + lean_ctor_set_tag(x_132, 1); } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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) { -_start: +lean_ctor_set(x_132, 0, x_2); +if (lean_is_scalar(x_131)) { + x_133 = lean_alloc_ctor(0, 2, 0); +} else { + x_133 = x_131; +} +lean_ctor_set(x_133, 0, x_132); +lean_ctor_set(x_133, 1, x_130); +return x_133; +} +else { -size_t x_16; size_t x_17; lean_object* x_18; -x_16 = lean_unbox_usize(x_2); +lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; +lean_dec(x_127); lean_dec(x_2); -x_17 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_18 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__8(x_1, x_16, x_17, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); -lean_dec(x_14); +x_134 = lean_ctor_get(x_129, 0); +lean_inc(x_134); +x_135 = lean_ctor_get(x_129, 1); +lean_inc(x_135); +if (lean_is_exclusive(x_129)) { + lean_ctor_release(x_129, 0); + lean_ctor_release(x_129, 1); + x_136 = x_129; +} else { + lean_dec_ref(x_129); + x_136 = lean_box(0); +} +if (lean_is_scalar(x_136)) { + x_137 = lean_alloc_ctor(1, 2, 0); +} else { + x_137 = x_136; +} +lean_ctor_set(x_137, 0, x_134); +lean_ctor_set(x_137, 1, x_135); +return x_137; +} +} +} +} +} +} +} +else +{ +lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; uint8_t x_142; +x_138 = lean_ctor_get(x_21, 1); +lean_inc(x_138); +lean_dec(x_21); +x_139 = lean_ctor_get(x_19, 0); +lean_inc(x_139); +x_140 = lean_ctor_get(x_19, 1); +lean_inc(x_140); +x_141 = lean_ctor_get(x_19, 2); +lean_inc(x_141); +x_142 = lean_nat_dec_lt(x_140, x_1); +if (x_142 == 0) +{ +lean_object* x_143; lean_object* x_144; +lean_dec(x_141); +lean_dec(x_140); +lean_dec(x_139); +lean_dec(x_19); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); @@ -10386,306 +14947,363 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_1); -return x_18; +lean_dec(x_4); +x_143 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_143, 0, x_2); +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 +{ +if (lean_obj_tag(x_139) == 0) +{ +lean_object* x_145; +lean_dec(x_141); +lean_dec(x_140); +x_145 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem(x_19, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_138); +if (lean_obj_tag(x_145) == 0) +{ +lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; +x_146 = lean_ctor_get(x_145, 1); +lean_inc(x_146); +if (lean_is_exclusive(x_145)) { + lean_ctor_release(x_145, 0); + lean_ctor_release(x_145, 1); + x_147 = x_145; +} else { + lean_dec_ref(x_145); + x_147 = lean_box(0); } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__9___boxed(lean_object** _args) { -lean_object* x_1 = _args[0]; -lean_object* x_2 = _args[1]; -lean_object* x_3 = _args[2]; -lean_object* x_4 = _args[3]; -lean_object* x_5 = _args[4]; -lean_object* x_6 = _args[5]; -lean_object* x_7 = _args[6]; -lean_object* x_8 = _args[7]; -lean_object* x_9 = _args[8]; -lean_object* x_10 = _args[9]; -lean_object* x_11 = _args[10]; -lean_object* x_12 = _args[11]; -lean_object* x_13 = _args[12]; -lean_object* x_14 = _args[13]; -lean_object* x_15 = _args[14]; -lean_object* x_16 = _args[15]; -lean_object* x_17 = _args[16]; -lean_object* x_18 = _args[17]; -_start: +x_148 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_148, 0, x_2); +if (lean_is_scalar(x_147)) { + x_149 = lean_alloc_ctor(0, 2, 0); +} else { + x_149 = x_147; +} +lean_ctor_set(x_149, 0, x_148); +lean_ctor_set(x_149, 1, x_146); +return x_149; +} +else { -size_t x_19; size_t x_20; lean_object* x_21; -x_19 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_20 = lean_unbox_usize(x_6); -lean_dec(x_6); -x_21 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__9(x_1, x_2, x_3, x_4, x_19, x_20, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_4); +lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_dec(x_2); -lean_dec(x_1); -return x_21; +x_150 = lean_ctor_get(x_145, 0); +lean_inc(x_150); +x_151 = lean_ctor_get(x_145, 1); +lean_inc(x_151); +if (lean_is_exclusive(x_145)) { + lean_ctor_release(x_145, 0); + lean_ctor_release(x_145, 1); + x_152 = x_145; +} else { + lean_dec_ref(x_145); + x_152 = lean_box(0); +} +if (lean_is_scalar(x_152)) { + x_153 = lean_alloc_ctor(1, 2, 0); +} else { + x_153 = x_152; } +lean_ctor_set(x_153, 0, x_150); +lean_ctor_set(x_153, 1, x_151); +return x_153; } -LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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) { -_start: +} +else { -size_t x_15; size_t x_16; lean_object* x_17; -x_15 = lean_unbox_usize(x_2); -lean_dec(x_2); -x_16 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_17 = l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__10(x_1, x_15, x_16, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); +lean_object* x_154; lean_object* x_155; +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_154 = x_19; +} else { + lean_dec_ref(x_19); + x_154 = lean_box(0); +} +x_155 = lean_ctor_get(x_139, 0); +lean_inc(x_155); +switch (lean_obj_tag(x_155)) { +case 0: +{ +lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; +x_156 = lean_ctor_get(x_139, 1); +lean_inc(x_156); +lean_dec(x_139); +x_157 = lean_ctor_get(x_155, 0); +lean_inc(x_157); +x_158 = lean_ctor_get(x_155, 1); +lean_inc(x_158); +lean_dec(x_155); +if (lean_is_scalar(x_154)) { + x_159 = lean_alloc_ctor(0, 3, 0); +} else { + x_159 = x_154; +} +lean_ctor_set(x_159, 0, x_156); +lean_ctor_set(x_159, 1, x_140); +lean_ctor_set(x_159, 2, x_141); +x_160 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch(x_159, x_157, x_158, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_138); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_1); -return x_17; +lean_dec(x_157); +if (lean_obj_tag(x_160) == 0) +{ +lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; +x_161 = lean_ctor_get(x_160, 1); +lean_inc(x_161); +if (lean_is_exclusive(x_160)) { + lean_ctor_release(x_160, 0); + lean_ctor_release(x_160, 1); + x_162 = x_160; +} else { + lean_dec_ref(x_160); + x_162 = lean_box(0); +} +x_163 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_163, 0, x_2); +if (lean_is_scalar(x_162)) { + x_164 = lean_alloc_ctor(0, 2, 0); +} else { + x_164 = x_162; } +lean_ctor_set(x_164, 0, x_163); +lean_ctor_set(x_164, 1, x_161); +return x_164; } -LEAN_EXPORT lean_object* l_Lean_Meta_withNewMCtxDepth___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { -_start: +else { -uint8_t x_14; lean_object* x_15; -x_14 = lean_unbox(x_2); +lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_dec(x_2); -x_15 = l_Lean_Meta_withNewMCtxDepth___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__12___rarg(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); -return x_15; +x_165 = lean_ctor_get(x_160, 0); +lean_inc(x_165); +x_166 = lean_ctor_get(x_160, 1); +lean_inc(x_166); +if (lean_is_exclusive(x_160)) { + lean_ctor_release(x_160, 0); + lean_ctor_release(x_160, 1); + x_167 = x_160; +} else { + lean_dec_ref(x_160); + x_167 = lean_box(0); } +if (lean_is_scalar(x_167)) { + x_168 = lean_alloc_ctor(1, 2, 0); +} else { + x_168 = x_167; } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { -_start: -{ -lean_object* x_17; -x_17 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_2); -return x_17; +lean_ctor_set(x_168, 0, x_165); +lean_ctor_set(x_168, 1, x_166); +return x_168; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { -_start: +case 1: { -lean_object* x_14; -x_14 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); +lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; +x_169 = lean_ctor_get(x_139, 1); +lean_inc(x_169); +lean_dec(x_139); +x_170 = lean_ctor_get(x_155, 0); +lean_inc(x_170); +x_171 = lean_ctor_get(x_155, 1); +lean_inc(x_171); +x_172 = lean_ctor_get(x_155, 2); +lean_inc(x_172); +lean_dec(x_155); +if (lean_is_scalar(x_154)) { + x_173 = lean_alloc_ctor(0, 3, 0); +} else { + x_173 = x_154; +} +lean_ctor_set(x_173, 0, x_169); +lean_ctor_set(x_173, 1, x_140); +lean_ctor_set(x_173, 2, x_141); +x_174 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processOffset(x_173, x_170, x_171, x_172, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_138); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_14; +lean_dec(x_171); +if (lean_obj_tag(x_174) == 0) +{ +lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; +x_175 = lean_ctor_get(x_174, 1); +lean_inc(x_175); +if (lean_is_exclusive(x_174)) { + lean_ctor_release(x_174, 0); + lean_ctor_release(x_174, 1); + x_176 = x_174; +} else { + lean_dec_ref(x_174); + x_176 = lean_box(0); } +x_177 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_177, 0, x_2); +if (lean_is_scalar(x_176)) { + x_178 = lean_alloc_ctor(0, 2, 0); +} else { + x_178 = x_176; } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___boxed(lean_object** _args) { -lean_object* x_1 = _args[0]; -lean_object* x_2 = _args[1]; -lean_object* x_3 = _args[2]; -lean_object* x_4 = _args[3]; -lean_object* x_5 = _args[4]; -lean_object* x_6 = _args[5]; -lean_object* x_7 = _args[6]; -lean_object* x_8 = _args[7]; -lean_object* x_9 = _args[8]; -lean_object* x_10 = _args[9]; -lean_object* x_11 = _args[10]; -lean_object* x_12 = _args[11]; -lean_object* x_13 = _args[12]; -lean_object* x_14 = _args[13]; -lean_object* x_15 = _args[14]; -lean_object* x_16 = _args[15]; -lean_object* x_17 = _args[16]; -lean_object* x_18 = _args[17]; -lean_object* x_19 = _args[18]; -lean_object* x_20 = _args[19]; -lean_object* x_21 = _args[20]; -_start: +lean_ctor_set(x_178, 0, x_177); +lean_ctor_set(x_178, 1, x_175); +return x_178; +} +else { -size_t x_22; lean_object* x_23; -x_22 = lean_unbox_usize(x_8); -lean_dec(x_8); -x_23 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_22, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_6); -lean_dec(x_3); -lean_dec(x_1); -return x_23; +lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; +lean_dec(x_2); +x_179 = lean_ctor_get(x_174, 0); +lean_inc(x_179); +x_180 = lean_ctor_get(x_174, 1); +lean_inc(x_180); +if (lean_is_exclusive(x_174)) { + lean_ctor_release(x_174, 0); + lean_ctor_release(x_174, 1); + x_181 = x_174; +} else { + lean_dec_ref(x_174); + x_181 = lean_box(0); } +if (lean_is_scalar(x_181)) { + x_182 = lean_alloc_ctor(1, 2, 0); +} else { + x_182 = x_181; } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__4___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_object* x_19 = _args[18]; -lean_object* x_20 = _args[19]; -_start: -{ -lean_object* x_21; -x_21 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20); -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_5); -lean_dec(x_3); -return x_21; +lean_ctor_set(x_182, 0, x_179); +lean_ctor_set(x_182, 1, x_180); +return x_182; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__5___boxed(lean_object** _args) { -lean_object* x_1 = _args[0]; -lean_object* x_2 = _args[1]; -lean_object* x_3 = _args[2]; -lean_object* x_4 = _args[3]; -lean_object* x_5 = _args[4]; -lean_object* x_6 = _args[5]; -lean_object* x_7 = _args[6]; -lean_object* x_8 = _args[7]; -lean_object* x_9 = _args[8]; -lean_object* x_10 = _args[9]; -lean_object* x_11 = _args[10]; -lean_object* x_12 = _args[11]; -lean_object* x_13 = _args[12]; -lean_object* x_14 = _args[13]; -lean_object* x_15 = _args[14]; -lean_object* x_16 = _args[15]; -lean_object* x_17 = _args[16]; -lean_object* x_18 = _args[17]; -lean_object* x_19 = _args[18]; -lean_object* x_20 = _args[19]; -_start: +default: { -lean_object* x_21; -x_21 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__5(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20); -lean_dec(x_9); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_1); -return x_21; +lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; +x_183 = lean_ctor_get(x_139, 1); +lean_inc(x_183); +lean_dec(x_139); +x_184 = lean_ctor_get(x_155, 0); +lean_inc(x_184); +if (lean_is_exclusive(x_155)) { + lean_ctor_release(x_155, 0); + x_185 = x_155; +} else { + lean_dec_ref(x_155); + x_185 = lean_box(0); } +if (lean_is_scalar(x_154)) { + x_186 = lean_alloc_ctor(0, 3, 0); +} else { + x_186 = x_154; } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -lean_object* x_13; -x_13 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); +lean_ctor_set(x_186, 0, x_183); +lean_ctor_set(x_186, 1, x_140); +lean_ctor_set(x_186, 2, x_141); +x_187 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue(x_186, x_184, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_138); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_13; +if (lean_obj_tag(x_187) == 0) +{ +lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; +x_188 = lean_ctor_get(x_187, 1); +lean_inc(x_188); +if (lean_is_exclusive(x_187)) { + lean_ctor_release(x_187, 0); + lean_ctor_release(x_187, 1); + x_189 = x_187; +} else { + lean_dec_ref(x_187); + x_189 = lean_box(0); } +if (lean_is_scalar(x_185)) { + x_190 = lean_alloc_ctor(1, 1, 0); +} else { + x_190 = x_185; + lean_ctor_set_tag(x_190, 1); } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { -_start: +lean_ctor_set(x_190, 0, x_2); +if (lean_is_scalar(x_189)) { + x_191 = lean_alloc_ctor(0, 2, 0); +} else { + x_191 = x_189; +} +lean_ctor_set(x_191, 0, x_190); +lean_ctor_set(x_191, 1, x_188); +return x_191; +} +else { -lean_object* x_16; -x_16 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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); -lean_dec(x_4); +lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; +lean_dec(x_185); lean_dec(x_2); -return x_16; +x_192 = lean_ctor_get(x_187, 0); +lean_inc(x_192); +x_193 = lean_ctor_get(x_187, 1); +lean_inc(x_193); +if (lean_is_exclusive(x_187)) { + lean_ctor_release(x_187, 0); + lean_ctor_release(x_187, 1); + x_194 = x_187; +} else { + lean_dec_ref(x_187); + x_194 = lean_box(0); +} +if (lean_is_scalar(x_194)) { + x_195 = lean_alloc_ctor(1, 2, 0); +} else { + x_195 = x_194; +} +lean_ctor_set(x_195, 0, x_192); +lean_ctor_set(x_195, 1, x_193); +return x_195; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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, lean_object* x_15) { -_start: -{ -lean_object* x_16; -x_16 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___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); -lean_dec(x_4); -return x_16; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -lean_object* x_13; -x_13 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -return x_13; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___lambda__2___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___lambda__1___boxed), 12, 0); -return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { _start: { -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; 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_13 = lean_st_ref_take(x_3, 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_Meta_Grind_EMatch_instInhabitedChoice; -x_17 = l_List_head_x21___rarg(x_16, x_14); -x_18 = l_List_tail_x21___rarg(x_14); -lean_dec(x_14); -x_19 = lean_st_ref_set(x_3, x_18, x_15); -x_20 = lean_ctor_get(x_19, 1); -lean_inc(x_20); -lean_dec(x_19); -x_21 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___lambda__2___closed__1; -x_22 = lean_ctor_get(x_17, 0); +lean_object* x_15; uint8_t x_16; +lean_dec(x_3); +x_15 = lean_st_ref_get(x_5, x_14); +x_16 = !lean_is_exclusive(x_15); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; uint8_t x_19; +x_17 = lean_ctor_get(x_15, 0); +x_18 = lean_ctor_get(x_15, 1); +x_19 = l_List_isEmpty___rarg(x_17); +lean_dec(x_17); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; +lean_free_object(x_15); +x_20 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__2; +x_21 = l_Lean_Core_checkSystem(x_20, x_12, x_13, x_18); +if (lean_obj_tag(x_21) == 0) +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; +x_22 = lean_ctor_get(x_21, 1); lean_inc(x_22); -if (lean_obj_tag(x_22) == 0) +lean_dec(x_21); +x_23 = l_Lean_Meta_Grind_checkMaxInstancesExceeded(x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_22); +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +x_25 = lean_unbox(x_24); +lean_dec(x_24); +if (x_25 == 0) { -lean_object* x_23; +lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_26 = lean_ctor_get(x_23, 1); +lean_inc(x_26); +lean_dec(x_23); +x_27 = lean_box(0); +lean_inc(x_13); +lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); @@ -10694,23 +15312,18 @@ 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_23 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem(x_17, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_20); -if (lean_obj_tag(x_23) == 0) +x_28 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___spec__1___lambda__1(x_1, x_2, x_27, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_26); +if (lean_obj_tag(x_28) == 0) { -lean_object* x_24; lean_object* x_25; lean_object* x_26; -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_apply_12(x_21, x_24, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_25); -return x_26; -} -else +lean_object* x_29; +x_29 = lean_ctor_get(x_28, 0); +lean_inc(x_29); +if (lean_obj_tag(x_29) == 0) { -uint8_t x_27; +uint8_t x_30; +lean_dec(x_13); +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -10719,75 +15332,53 @@ 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_27 = !lean_is_exclusive(x_23); -if (x_27 == 0) +x_30 = !lean_is_exclusive(x_28); +if (x_30 == 0) { -return x_23; +lean_object* x_31; lean_object* x_32; +x_31 = lean_ctor_get(x_28, 0); +lean_dec(x_31); +x_32 = lean_ctor_get(x_29, 0); +lean_inc(x_32); +lean_dec(x_29); +lean_ctor_set(x_28, 0, x_32); +return x_28; } else { -lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_28 = lean_ctor_get(x_23, 0); -x_29 = lean_ctor_get(x_23, 1); -lean_inc(x_29); -lean_inc(x_28); -lean_dec(x_23); -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_33; lean_object* x_34; lean_object* x_35; +x_33 = lean_ctor_get(x_28, 1); +lean_inc(x_33); +lean_dec(x_28); +x_34 = lean_ctor_get(x_29, 0); +lean_inc(x_34); +lean_dec(x_29); +x_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_34); +lean_ctor_set(x_35, 1, x_33); +return x_35; } } else { -lean_object* x_31; -x_31 = lean_ctor_get(x_22, 0); -lean_inc(x_31); -if (lean_obj_tag(x_31) == 0) -{ -uint8_t x_32; -x_32 = !lean_is_exclusive(x_17); -if (x_32 == 0) -{ -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_17, 0); -lean_dec(x_33); -x_34 = lean_ctor_get(x_22, 1); -lean_inc(x_34); -lean_dec(x_22); -x_35 = lean_ctor_get(x_31, 0); -lean_inc(x_35); -x_36 = lean_ctor_get(x_31, 1); +lean_object* x_36; lean_object* x_37; +x_36 = lean_ctor_get(x_28, 1); lean_inc(x_36); -lean_dec(x_31); -lean_ctor_set(x_17, 0, x_34); -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_37 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch(x_17, x_35, x_36, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_20); -lean_dec(x_35); -if (lean_obj_tag(x_37) == 0) -{ -lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_38 = lean_ctor_get(x_37, 0); -lean_inc(x_38); -x_39 = lean_ctor_get(x_37, 1); -lean_inc(x_39); -lean_dec(x_37); -x_40 = lean_apply_12(x_21, x_38, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_39); -return x_40; +lean_dec(x_28); +x_37 = lean_ctor_get(x_29, 0); +lean_inc(x_37); +lean_dec(x_29); +x_3 = x_37; +x_14 = x_36; +goto _start; +} } else { -uint8_t x_41; +uint8_t x_39; +lean_dec(x_13); +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -10796,72 +15387,70 @@ 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_41 = !lean_is_exclusive(x_37); -if (x_41 == 0) +x_39 = !lean_is_exclusive(x_28); +if (x_39 == 0) { -return x_37; +return x_28; } else { -lean_object* x_42; lean_object* x_43; lean_object* x_44; -x_42 = lean_ctor_get(x_37, 0); -x_43 = lean_ctor_get(x_37, 1); -lean_inc(x_43); -lean_inc(x_42); -lean_dec(x_37); -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_40; lean_object* x_41; lean_object* x_42; +x_40 = lean_ctor_get(x_28, 0); +x_41 = lean_ctor_get(x_28, 1); +lean_inc(x_41); +lean_inc(x_40); +lean_dec(x_28); +x_42 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_42, 0, x_40); +lean_ctor_set(x_42, 1, x_41); +return x_42; } } } else { -lean_object* x_45; 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_45 = lean_ctor_get(x_17, 1); -x_46 = lean_ctor_get(x_17, 2); -lean_inc(x_46); -lean_inc(x_45); -lean_dec(x_17); -x_47 = lean_ctor_get(x_22, 1); -lean_inc(x_47); -lean_dec(x_22); -x_48 = lean_ctor_get(x_31, 0); -lean_inc(x_48); -x_49 = lean_ctor_get(x_31, 1); -lean_inc(x_49); -lean_dec(x_31); -x_50 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_50, 0, x_47); -lean_ctor_set(x_50, 1, x_45); -lean_ctor_set(x_50, 2, x_46); -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_51 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch(x_50, x_48, x_49, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_20); -lean_dec(x_48); -if (lean_obj_tag(x_51) == 0) +uint8_t x_43; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +x_43 = !lean_is_exclusive(x_23); +if (x_43 == 0) { -lean_object* x_52; lean_object* x_53; lean_object* x_54; -x_52 = lean_ctor_get(x_51, 0); -lean_inc(x_52); -x_53 = lean_ctor_get(x_51, 1); -lean_inc(x_53); -lean_dec(x_51); -x_54 = lean_apply_12(x_21, x_52, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_53); -return x_54; +lean_object* x_44; lean_object* x_45; +x_44 = lean_ctor_get(x_23, 0); +lean_dec(x_44); +x_45 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___lambda__1___closed__2; +lean_ctor_set(x_23, 0, x_45); +return x_23; +} +else +{ +lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_46 = lean_ctor_get(x_23, 1); +lean_inc(x_46); +lean_dec(x_23); +x_47 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___lambda__1___closed__2; +x_48 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_48, 0, x_47); +lean_ctor_set(x_48, 1, x_46); +return x_48; +} +} } else { -lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; +uint8_t x_49; +lean_dec(x_13); +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -10870,47 +15459,78 @@ 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_55 = lean_ctor_get(x_51, 0); -lean_inc(x_55); -x_56 = lean_ctor_get(x_51, 1); -lean_inc(x_56); -if (lean_is_exclusive(x_51)) { - lean_ctor_release(x_51, 0); - lean_ctor_release(x_51, 1); - x_57 = x_51; -} else { - lean_dec_ref(x_51); - x_57 = lean_box(0); +x_49 = !lean_is_exclusive(x_21); +if (x_49 == 0) +{ +return x_21; +} +else +{ +lean_object* x_50; lean_object* x_51; lean_object* x_52; +x_50 = lean_ctor_get(x_21, 0); +x_51 = lean_ctor_get(x_21, 1); +lean_inc(x_51); +lean_inc(x_50); +lean_dec(x_21); +x_52 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_52, 0, x_50); +lean_ctor_set(x_52, 1, x_51); +return x_52; } -if (lean_is_scalar(x_57)) { - x_58 = lean_alloc_ctor(1, 2, 0); -} else { - x_58 = x_57; } -lean_ctor_set(x_58, 0, x_55); -lean_ctor_set(x_58, 1, x_56); -return x_58; } +else +{ +lean_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_ctor_set(x_15, 0, x_2); +return x_15; } } else { -uint8_t x_59; -x_59 = !lean_is_exclusive(x_17); -if (x_59 == 0) +lean_object* x_53; lean_object* x_54; uint8_t x_55; +x_53 = lean_ctor_get(x_15, 0); +x_54 = lean_ctor_get(x_15, 1); +lean_inc(x_54); +lean_inc(x_53); +lean_dec(x_15); +x_55 = l_List_isEmpty___rarg(x_53); +lean_dec(x_53); +if (x_55 == 0) +{ +lean_object* x_56; lean_object* x_57; +x_56 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__2; +x_57 = l_Lean_Core_checkSystem(x_56, x_12, x_13, x_54); +if (lean_obj_tag(x_57) == 0) +{ +lean_object* x_58; lean_object* x_59; lean_object* x_60; uint8_t x_61; +x_58 = lean_ctor_get(x_57, 1); +lean_inc(x_58); +lean_dec(x_57); +x_59 = l_Lean_Meta_Grind_checkMaxInstancesExceeded(x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_58); +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_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; -x_60 = lean_ctor_get(x_17, 0); -lean_dec(x_60); -x_61 = lean_ctor_get(x_22, 1); -lean_inc(x_61); -lean_dec(x_22); -x_62 = lean_ctor_get(x_31, 0); +lean_object* x_62; lean_object* x_63; lean_object* x_64; +x_62 = lean_ctor_get(x_59, 1); lean_inc(x_62); -lean_dec(x_31); -lean_ctor_set(x_17, 0, x_61); +lean_dec(x_59); +x_63 = lean_box(0); +lean_inc(x_13); +lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); @@ -10919,21 +15539,18 @@ lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -x_63 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue(x_17, x_62, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_20); -if (lean_obj_tag(x_63) == 0) +lean_inc(x_2); +x_64 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___spec__1___lambda__1(x_1, x_2, x_63, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_62); +if (lean_obj_tag(x_64) == 0) { -lean_object* x_64; lean_object* x_65; lean_object* x_66; -x_64 = lean_ctor_get(x_63, 0); -lean_inc(x_64); -x_65 = lean_ctor_get(x_63, 1); +lean_object* x_65; +x_65 = lean_ctor_get(x_64, 0); lean_inc(x_65); -lean_dec(x_63); -x_66 = lean_apply_12(x_21, x_64, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_65); -return x_66; -} -else +if (lean_obj_tag(x_65) == 0) { -uint8_t x_67; +lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; +lean_dec(x_13); +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -10942,69 +15559,48 @@ 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_67 = !lean_is_exclusive(x_63); -if (x_67 == 0) -{ -return x_63; +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); } -else -{ -lean_object* x_68; lean_object* x_69; lean_object* x_70; -x_68 = lean_ctor_get(x_63, 0); -x_69 = lean_ctor_get(x_63, 1); -lean_inc(x_69); +x_68 = lean_ctor_get(x_65, 0); lean_inc(x_68); -lean_dec(x_63); -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_dec(x_65); +if (lean_is_scalar(x_67)) { + x_69 = lean_alloc_ctor(0, 2, 0); +} else { + x_69 = x_67; } +lean_ctor_set(x_69, 0, x_68); +lean_ctor_set(x_69, 1, x_66); +return x_69; } else { -lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; -x_71 = lean_ctor_get(x_17, 1); -x_72 = lean_ctor_get(x_17, 2); -lean_inc(x_72); +lean_object* x_70; lean_object* x_71; +x_70 = lean_ctor_get(x_64, 1); +lean_inc(x_70); +lean_dec(x_64); +x_71 = lean_ctor_get(x_65, 0); lean_inc(x_71); -lean_dec(x_17); -x_73 = lean_ctor_get(x_22, 1); -lean_inc(x_73); -lean_dec(x_22); -x_74 = lean_ctor_get(x_31, 0); -lean_inc(x_74); -lean_dec(x_31); -x_75 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_75, 0, x_73); -lean_ctor_set(x_75, 1, x_71); -lean_ctor_set(x_75, 2, x_72); -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_76 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue(x_75, x_74, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_20); -if (lean_obj_tag(x_76) == 0) -{ -lean_object* x_77; lean_object* x_78; lean_object* x_79; -x_77 = lean_ctor_get(x_76, 0); -lean_inc(x_77); -x_78 = lean_ctor_get(x_76, 1); -lean_inc(x_78); -lean_dec(x_76); -x_79 = lean_apply_12(x_21, x_77, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_78); -return x_79; +lean_dec(x_65); +x_3 = x_71; +x_14 = x_70; +goto _start; +} } else { -lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; +lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; +lean_dec(x_13); +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -11013,86 +15609,35 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); lean_dec(x_2); -x_80 = lean_ctor_get(x_76, 0); -lean_inc(x_80); -x_81 = lean_ctor_get(x_76, 1); -lean_inc(x_81); -if (lean_is_exclusive(x_76)) { - lean_ctor_release(x_76, 0); - lean_ctor_release(x_76, 1); - x_82 = x_76; +x_73 = lean_ctor_get(x_64, 0); +lean_inc(x_73); +x_74 = lean_ctor_get(x_64, 1); +lean_inc(x_74); +if (lean_is_exclusive(x_64)) { + lean_ctor_release(x_64, 0); + lean_ctor_release(x_64, 1); + x_75 = x_64; } else { - lean_dec_ref(x_76); - x_82 = lean_box(0); + lean_dec_ref(x_64); + x_75 = lean_box(0); } -if (lean_is_scalar(x_82)) { - x_83 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_75)) { + x_76 = 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; -} -} -} -} -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___lambda__2___boxed), 12, 0); -return x_1; + 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___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, 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_st_ref_get(x_2, x_11); -x_13 = !lean_is_exclusive(x_12); -if (x_13 == 0) -{ -lean_object* x_14; lean_object* x_15; uint8_t x_16; -x_14 = lean_ctor_get(x_12, 0); -x_15 = lean_ctor_get(x_12, 1); -x_16 = l_List_isEmpty___rarg(x_14); -lean_dec(x_14); -if (x_16 == 0) -{ -lean_object* x_17; lean_object* x_18; -lean_free_object(x_12); -x_17 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__2; -x_18 = l_Lean_Core_checkSystem(x_17, x_9, x_10, x_15); -if (lean_obj_tag(x_18) == 0) -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; -x_19 = lean_ctor_get(x_18, 1); -lean_inc(x_19); -lean_dec(x_18); -x_20 = l_Lean_Meta_Grind_checkMaxInstancesExceeded(x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_19); -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; -x_23 = lean_ctor_get(x_20, 1); -lean_inc(x_23); -lean_dec(x_20); -x_24 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___closed__1; -x_25 = lean_box(0); -x_26 = lean_apply_12(x_24, x_25, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_23); -return x_26; } else { -uint8_t x_27; +lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -11100,36 +15645,34 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -x_27 = !lean_is_exclusive(x_20); -if (x_27 == 0) -{ -lean_object* x_28; lean_object* x_29; -x_28 = lean_ctor_get(x_20, 0); -lean_dec(x_28); -x_29 = lean_box(0); -lean_ctor_set(x_20, 0, x_29); -return x_20; +x_77 = lean_ctor_get(x_59, 1); +lean_inc(x_77); +if (lean_is_exclusive(x_59)) { + lean_ctor_release(x_59, 0); + lean_ctor_release(x_59, 1); + x_78 = x_59; +} else { + lean_dec_ref(x_59); + x_78 = lean_box(0); } -else -{ -lean_object* x_30; lean_object* x_31; lean_object* x_32; -x_30 = lean_ctor_get(x_20, 1); -lean_inc(x_30); -lean_dec(x_20); -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; +x_79 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___lambda__1___closed__2; +if (lean_is_scalar(x_78)) { + x_80 = lean_alloc_ctor(0, 2, 0); +} else { + x_80 = x_78; } +lean_ctor_set(x_80, 0, x_79); +lean_ctor_set(x_80, 1, x_77); +return x_80; } } else { -uint8_t x_33; +lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -11137,32 +15680,35 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -x_33 = !lean_is_exclusive(x_18); -if (x_33 == 0) -{ -return x_18; +x_81 = lean_ctor_get(x_57, 0); +lean_inc(x_81); +x_82 = lean_ctor_get(x_57, 1); +lean_inc(x_82); +if (lean_is_exclusive(x_57)) { + lean_ctor_release(x_57, 0); + lean_ctor_release(x_57, 1); + x_83 = x_57; +} else { + lean_dec_ref(x_57); + x_83 = lean_box(0); } -else -{ -lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_34 = lean_ctor_get(x_18, 0); -x_35 = lean_ctor_get(x_18, 1); -lean_inc(x_35); -lean_inc(x_34); -lean_dec(x_18); -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; +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; } } else { -lean_object* x_37; +lean_object* x_85; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -11170,158 +15716,133 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_37 = lean_box(0); -lean_ctor_set(x_12, 0, x_37); -return x_12; +x_85 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_85, 0, x_2); +lean_ctor_set(x_85, 1, x_54); +return x_85; } } -else +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, 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_38; lean_object* x_39; uint8_t x_40; -x_38 = lean_ctor_get(x_12, 0); -x_39 = lean_ctor_get(x_12, 1); -lean_inc(x_39); -lean_inc(x_38); +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_12 = l_Lean_Meta_Grind_getMaxGeneration___rarg(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_ctor_get(x_12, 1); +lean_inc(x_14); lean_dec(x_12); -x_40 = l_List_isEmpty___rarg(x_38); -lean_dec(x_38); -if (x_40 == 0) +x_15 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__4___closed__1; +x_16 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___spec__1(x_13, x_15, x_15, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_14); +lean_dec(x_13); +if (lean_obj_tag(x_16) == 0) { -lean_object* x_41; lean_object* x_42; -x_41 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__2; -x_42 = l_Lean_Core_checkSystem(x_41, x_9, x_10, x_39); -if (lean_obj_tag(x_42) == 0) +lean_object* x_17; lean_object* x_18; +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +lean_dec(x_17); +if (lean_obj_tag(x_18) == 0) { -lean_object* x_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 = l_Lean_Meta_Grind_checkMaxInstancesExceeded(x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_43); -x_45 = lean_ctor_get(x_44, 0); -lean_inc(x_45); -x_46 = lean_unbox(x_45); -lean_dec(x_45); -if (x_46 == 0) +uint8_t x_19; +x_19 = !lean_is_exclusive(x_16); +if (x_19 == 0) { -lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; -x_47 = lean_ctor_get(x_44, 1); -lean_inc(x_47); -lean_dec(x_44); -x_48 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___closed__1; -x_49 = lean_box(0); -x_50 = lean_apply_12(x_48, x_49, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_47); -return x_50; +lean_object* x_20; lean_object* x_21; +x_20 = lean_ctor_get(x_16, 0); +lean_dec(x_20); +x_21 = lean_box(0); +lean_ctor_set(x_16, 0, x_21); +return x_16; } else { -lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; -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_51 = lean_ctor_get(x_44, 1); -lean_inc(x_51); -if (lean_is_exclusive(x_44)) { - lean_ctor_release(x_44, 0); - lean_ctor_release(x_44, 1); - x_52 = x_44; -} else { - lean_dec_ref(x_44); - 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_object* x_22; lean_object* x_23; lean_object* x_24; +x_22 = lean_ctor_get(x_16, 1); +lean_inc(x_22); +lean_dec(x_16); +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_ctor_set(x_54, 0, x_53); -lean_ctor_set(x_54, 1, x_51); -return x_54; } +else +{ +uint8_t x_25; +x_25 = !lean_is_exclusive(x_16); +if (x_25 == 0) +{ +lean_object* x_26; lean_object* x_27; +x_26 = lean_ctor_get(x_16, 0); +lean_dec(x_26); +x_27 = lean_ctor_get(x_18, 0); +lean_inc(x_27); +lean_dec(x_18); +lean_ctor_set(x_16, 0, x_27); +return x_16; } else { -lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_55 = lean_ctor_get(x_42, 0); -lean_inc(x_55); -x_56 = lean_ctor_get(x_42, 1); -lean_inc(x_56); -if (lean_is_exclusive(x_42)) { - lean_ctor_release(x_42, 0); - lean_ctor_release(x_42, 1); - x_57 = x_42; -} else { - lean_dec_ref(x_42); - x_57 = lean_box(0); +lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_28 = lean_ctor_get(x_16, 1); +lean_inc(x_28); +lean_dec(x_16); +x_29 = lean_ctor_get(x_18, 0); +lean_inc(x_29); +lean_dec(x_18); +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; } -if (lean_is_scalar(x_57)) { - x_58 = lean_alloc_ctor(1, 2, 0); -} else { - x_58 = x_57; } -lean_ctor_set(x_58, 0, x_55); -lean_ctor_set(x_58, 1, x_56); -return x_58; } +else +{ +uint8_t x_31; +x_31 = !lean_is_exclusive(x_16); +if (x_31 == 0) +{ +return x_16; } else { -lean_object* x_59; lean_object* x_60; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_59 = lean_box(0); -x_60 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_60, 0, x_59); -lean_ctor_set(x_60, 1, x_39); -return x_60; +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_16, 0); +x_33 = lean_ctor_get(x_16, 1); +lean_inc(x_33); +lean_inc(x_32); +lean_dec(x_16); +x_34 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +return x_34; } } } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___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 lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___spec__1___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { _start: { -lean_object* x_13; -x_13 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___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_object* x_15; +x_15 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___spec__1___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +lean_dec(x_3); lean_dec(x_1); -return x_13; +return x_15; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { _start: { -lean_object* x_13; -x_13 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_object* x_15; +x_15 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); lean_dec(x_1); -return x_13; +return x_15; } } LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_main___spec__1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, uint8_t x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19) { @@ -11336,20 +15857,17 @@ uint8_t x_21; 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_89; +lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_123; x_22 = lean_ctor_get(x_20, 0); x_23 = lean_ctor_get(x_20, 1); -x_89 = lean_ctor_get_uint8(x_22, sizeof(void*)*10 + 4); -if (x_89 == 0) +x_123 = lean_ctor_get_uint8(x_22, sizeof(void*)*10 + 4); +if (x_123 == 0) { -lean_object* x_90; uint8_t x_91; -x_90 = lean_ctor_get(x_22, 3); -lean_inc(x_90); -x_91 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_90, x_1); -lean_dec(x_90); -if (x_91 == 0) +uint8_t x_124; +x_124 = l_Lean_Meta_Grind_ENode_isCongrRoot(x_22); +if (x_124 == 0) { -lean_object* x_92; +lean_object* x_125; lean_dec(x_22); lean_dec(x_18); lean_dec(x_17); @@ -11364,31 +15882,31 @@ lean_dec(x_9); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_92 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_92, 0, x_5); -lean_ctor_set(x_20, 0, x_92); +x_125 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_125, 0, x_5); +lean_ctor_set(x_20, 0, x_125); return x_20; } else { if (x_6 == 0) { -lean_object* x_93; +lean_object* x_126; lean_free_object(x_20); -x_93 = lean_box(0); -x_24 = x_93; -goto block_88; +x_126 = lean_box(0); +x_24 = x_126; +goto block_122; } else { -lean_object* x_94; uint8_t x_95; -x_94 = lean_ctor_get(x_22, 9); -lean_inc(x_94); -x_95 = lean_nat_dec_eq(x_94, x_7); -lean_dec(x_94); -if (x_95 == 0) +lean_object* x_127; uint8_t x_128; +x_127 = lean_ctor_get(x_22, 9); +lean_inc(x_127); +x_128 = lean_nat_dec_eq(x_127, x_7); +lean_dec(x_127); +if (x_128 == 0) { -lean_object* x_96; +lean_object* x_129; lean_dec(x_22); lean_dec(x_18); lean_dec(x_17); @@ -11403,18 +15921,18 @@ lean_dec(x_9); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_96 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_96, 0, x_5); -lean_ctor_set(x_20, 0, x_96); +x_129 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_129, 0, x_5); +lean_ctor_set(x_20, 0, x_129); return x_20; } else { -lean_object* x_97; +lean_object* x_130; lean_free_object(x_20); -x_97 = lean_box(0); -x_24 = x_97; -goto block_88; +x_130 = lean_box(0); +x_24 = x_130; +goto block_122; } } } @@ -11423,22 +15941,22 @@ else { if (x_6 == 0) { -lean_object* x_98; +lean_object* x_131; lean_free_object(x_20); -x_98 = lean_box(0); -x_24 = x_98; -goto block_88; +x_131 = lean_box(0); +x_24 = x_131; +goto block_122; } else { -lean_object* x_99; uint8_t x_100; -x_99 = lean_ctor_get(x_22, 9); -lean_inc(x_99); -x_100 = lean_nat_dec_eq(x_99, x_7); -lean_dec(x_99); -if (x_100 == 0) +lean_object* x_132; uint8_t x_133; +x_132 = lean_ctor_get(x_22, 9); +lean_inc(x_132); +x_133 = lean_nat_dec_eq(x_132, x_7); +lean_dec(x_132); +if (x_133 == 0) { -lean_object* x_101; +lean_object* x_134; lean_dec(x_22); lean_dec(x_18); lean_dec(x_17); @@ -11453,24 +15971,24 @@ lean_dec(x_9); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_101 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_101, 0, x_5); -lean_ctor_set(x_20, 0, x_101); +x_134 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_134, 0, x_5); +lean_ctor_set(x_20, 0, x_134); return x_20; } else { -lean_object* x_102; +lean_object* x_135; lean_free_object(x_20); -x_102 = lean_box(0); -x_24 = x_102; -goto block_88; +x_135 = lean_box(0); +x_24 = x_135; +goto block_122; } } } -block_88: +block_122: { -lean_object* x_25; lean_object* x_26; lean_object* x_27; +lean_object* x_25; lean_object* x_26; uint8_t x_27; lean_dec(x_24); x_25 = lean_ctor_get(x_22, 8); lean_inc(x_25); @@ -11479,6 +15997,14 @@ x_26 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_26, 0, x_2); lean_ctor_set(x_26, 1, x_25); lean_ctor_set(x_26, 2, x_3); +x_27 = !lean_is_exclusive(x_9); +if (x_27 == 0) +{ +lean_object* x_28; lean_object* x_29; +x_28 = lean_ctor_get(x_9, 1); +lean_dec(x_28); +lean_inc(x_1); +lean_ctor_set(x_9, 1, x_1); lean_inc(x_18); lean_inc(x_17); lean_inc(x_16); @@ -11487,16 +16013,17 @@ lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); -x_27 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_x3f(x_26, x_4, x_1, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_23); +x_29 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_x3f(x_26, x_4, x_1, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_23); lean_dec(x_1); -if (lean_obj_tag(x_27) == 0) +if (lean_obj_tag(x_29) == 0) { -lean_object* x_28; -x_28 = lean_ctor_get(x_27, 0); -lean_inc(x_28); -if (lean_obj_tag(x_28) == 0) +lean_object* x_30; +x_30 = lean_ctor_get(x_29, 0); +lean_inc(x_30); +if (lean_obj_tag(x_30) == 0) { -uint8_t x_29; +uint8_t x_31; +lean_dec(x_9); lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); @@ -11506,263 +16033,454 @@ lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); -lean_dec(x_9); -x_29 = !lean_is_exclusive(x_27); -if (x_29 == 0) +x_31 = !lean_is_exclusive(x_29); +if (x_31 == 0) { -lean_object* x_30; lean_object* x_31; -x_30 = lean_ctor_get(x_27, 0); -lean_dec(x_30); -x_31 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_31, 0, x_5); -lean_ctor_set(x_27, 0, x_31); -return x_27; +lean_object* x_32; lean_object* x_33; +x_32 = lean_ctor_get(x_29, 0); +lean_dec(x_32); +x_33 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_33, 0, x_5); +lean_ctor_set(x_29, 0, x_33); +return x_29; } else { -lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_32 = lean_ctor_get(x_27, 1); -lean_inc(x_32); -lean_dec(x_27); -x_33 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_33, 0, x_5); -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; +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_29, 1); +lean_inc(x_34); +lean_dec(x_29); +x_35 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_35, 0, x_5); +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_35; uint8_t x_36; -x_35 = lean_ctor_get(x_27, 1); -lean_inc(x_35); -lean_dec(x_27); -x_36 = !lean_is_exclusive(x_28); -if (x_36 == 0) +lean_object* x_37; uint8_t x_38; +x_37 = lean_ctor_get(x_29, 1); +lean_inc(x_37); +lean_dec(x_29); +x_38 = !lean_is_exclusive(x_30); +if (x_38 == 0) { -lean_object* x_37; lean_object* x_38; uint8_t x_39; -x_37 = lean_ctor_get(x_28, 0); -x_38 = lean_st_ref_take(x_10, x_35); -x_39 = !lean_is_exclusive(x_38); -if (x_39 == 0) +lean_object* x_39; lean_object* x_40; uint8_t x_41; +x_39 = lean_ctor_get(x_30, 0); +x_40 = lean_st_ref_take(x_10, x_37); +x_41 = !lean_is_exclusive(x_40); +if (x_41 == 0) { -lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; -x_40 = lean_ctor_get(x_38, 1); -x_41 = lean_ctor_get(x_38, 0); -lean_dec(x_41); -x_42 = lean_box(0); -lean_ctor_set_tag(x_38, 1); -lean_ctor_set(x_38, 1, x_42); -lean_ctor_set(x_38, 0, x_37); -x_43 = lean_st_ref_set(x_10, x_38, x_40); -x_44 = lean_ctor_get(x_43, 1); -lean_inc(x_44); +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_42 = lean_ctor_get(x_40, 1); +x_43 = lean_ctor_get(x_40, 0); lean_dec(x_43); -x_45 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices(x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_44); -if (lean_obj_tag(x_45) == 0) +x_44 = lean_box(0); +lean_ctor_set_tag(x_40, 1); +lean_ctor_set(x_40, 1, x_44); +lean_ctor_set(x_40, 0, x_39); +x_45 = lean_st_ref_set(x_10, x_40, x_42); +x_46 = lean_ctor_get(x_45, 1); +lean_inc(x_46); +lean_dec(x_45); +x_47 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices(x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_46); +if (lean_obj_tag(x_47) == 0) { -uint8_t x_46; -x_46 = !lean_is_exclusive(x_45); -if (x_46 == 0) +uint8_t x_48; +x_48 = !lean_is_exclusive(x_47); +if (x_48 == 0) { -lean_object* x_47; -x_47 = lean_ctor_get(x_45, 0); +lean_object* x_49; +x_49 = lean_ctor_get(x_47, 0); +lean_dec(x_49); +lean_ctor_set(x_30, 0, x_5); +lean_ctor_set(x_47, 0, x_30); +return x_47; +} +else +{ +lean_object* x_50; lean_object* x_51; +x_50 = lean_ctor_get(x_47, 1); +lean_inc(x_50); lean_dec(x_47); -lean_ctor_set(x_28, 0, x_5); -lean_ctor_set(x_45, 0, x_28); -return x_45; +lean_ctor_set(x_30, 0, x_5); +x_51 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_51, 0, x_30); +lean_ctor_set(x_51, 1, x_50); +return x_51; +} +} +else +{ +uint8_t x_52; +lean_free_object(x_30); +lean_dec(x_5); +x_52 = !lean_is_exclusive(x_47); +if (x_52 == 0) +{ +return x_47; +} +else +{ +lean_object* x_53; lean_object* x_54; lean_object* x_55; +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_alloc_ctor(1, 2, 0); +lean_ctor_set(x_55, 0, x_53); +lean_ctor_set(x_55, 1, x_54); +return x_55; +} +} +} +else +{ +lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; +x_56 = lean_ctor_get(x_40, 1); +lean_inc(x_56); +lean_dec(x_40); +x_57 = lean_box(0); +x_58 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_58, 0, x_39); +lean_ctor_set(x_58, 1, x_57); +x_59 = lean_st_ref_set(x_10, x_58, x_56); +x_60 = lean_ctor_get(x_59, 1); +lean_inc(x_60); +lean_dec(x_59); +x_61 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices(x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_60); +if (lean_obj_tag(x_61) == 0) +{ +lean_object* x_62; lean_object* x_63; lean_object* x_64; +x_62 = lean_ctor_get(x_61, 1); +lean_inc(x_62); +if (lean_is_exclusive(x_61)) { + lean_ctor_release(x_61, 0); + lean_ctor_release(x_61, 1); + x_63 = x_61; +} else { + lean_dec_ref(x_61); + x_63 = lean_box(0); +} +lean_ctor_set(x_30, 0, x_5); +if (lean_is_scalar(x_63)) { + x_64 = lean_alloc_ctor(0, 2, 0); +} else { + x_64 = x_63; +} +lean_ctor_set(x_64, 0, x_30); +lean_ctor_set(x_64, 1, x_62); +return x_64; +} +else +{ +lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; +lean_free_object(x_30); +lean_dec(x_5); +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; +} +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; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; +x_69 = lean_ctor_get(x_30, 0); +lean_inc(x_69); +lean_dec(x_30); +x_70 = lean_st_ref_take(x_10, x_37); +x_71 = lean_ctor_get(x_70, 1); +lean_inc(x_71); +if (lean_is_exclusive(x_70)) { + lean_ctor_release(x_70, 0); + lean_ctor_release(x_70, 1); + x_72 = x_70; +} else { + lean_dec_ref(x_70); + x_72 = lean_box(0); +} +x_73 = lean_box(0); +if (lean_is_scalar(x_72)) { + x_74 = lean_alloc_ctor(1, 2, 0); +} else { + x_74 = x_72; + lean_ctor_set_tag(x_74, 1); +} +lean_ctor_set(x_74, 0, x_69); +lean_ctor_set(x_74, 1, x_73); +x_75 = lean_st_ref_set(x_10, x_74, x_71); +x_76 = lean_ctor_get(x_75, 1); +lean_inc(x_76); +lean_dec(x_75); +x_77 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices(x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_76); +if (lean_obj_tag(x_77) == 0) +{ +lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; +x_78 = lean_ctor_get(x_77, 1); +lean_inc(x_78); +if (lean_is_exclusive(x_77)) { + lean_ctor_release(x_77, 0); + lean_ctor_release(x_77, 1); + x_79 = x_77; +} else { + lean_dec_ref(x_77); + x_79 = lean_box(0); +} +x_80 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_80, 0, x_5); +if (lean_is_scalar(x_79)) { + x_81 = lean_alloc_ctor(0, 2, 0); +} else { + x_81 = x_79; +} +lean_ctor_set(x_81, 0, x_80); +lean_ctor_set(x_81, 1, x_78); +return x_81; } else { -lean_object* x_48; lean_object* x_49; -x_48 = lean_ctor_get(x_45, 1); -lean_inc(x_48); -lean_dec(x_45); -lean_ctor_set(x_28, 0, x_5); -x_49 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_49, 0, x_28); -lean_ctor_set(x_49, 1, x_48); -return x_49; +lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; +lean_dec(x_5); +x_82 = lean_ctor_get(x_77, 0); +lean_inc(x_82); +x_83 = lean_ctor_get(x_77, 1); +lean_inc(x_83); +if (lean_is_exclusive(x_77)) { + lean_ctor_release(x_77, 0); + lean_ctor_release(x_77, 1); + x_84 = x_77; +} else { + lean_dec_ref(x_77); + x_84 = lean_box(0); +} +if (lean_is_scalar(x_84)) { + x_85 = lean_alloc_ctor(1, 2, 0); +} else { + x_85 = x_84; +} +lean_ctor_set(x_85, 0, x_82); +lean_ctor_set(x_85, 1, x_83); +return x_85; +} +} } } else { -uint8_t x_50; -lean_free_object(x_28); +uint8_t x_86; +lean_dec(x_9); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); lean_dec(x_5); -x_50 = !lean_is_exclusive(x_45); -if (x_50 == 0) +x_86 = !lean_is_exclusive(x_29); +if (x_86 == 0) { -return x_45; +return x_29; } else { -lean_object* x_51; lean_object* x_52; lean_object* x_53; -x_51 = lean_ctor_get(x_45, 0); -x_52 = lean_ctor_get(x_45, 1); -lean_inc(x_52); -lean_inc(x_51); -lean_dec(x_45); -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_87; lean_object* x_88; lean_object* x_89; +x_87 = lean_ctor_get(x_29, 0); +x_88 = lean_ctor_get(x_29, 1); +lean_inc(x_88); +lean_inc(x_87); +lean_dec(x_29); +x_89 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_89, 0, x_87); +lean_ctor_set(x_89, 1, x_88); +return x_89; } } } else { -lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; -x_54 = lean_ctor_get(x_38, 1); -lean_inc(x_54); -lean_dec(x_38); -x_55 = lean_box(0); -x_56 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_56, 0, x_37); -lean_ctor_set(x_56, 1, x_55); -x_57 = lean_st_ref_set(x_10, x_56, x_54); -x_58 = lean_ctor_get(x_57, 1); -lean_inc(x_58); -lean_dec(x_57); -x_59 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices(x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_58); -if (lean_obj_tag(x_59) == 0) +uint8_t x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; +x_90 = lean_ctor_get_uint8(x_9, sizeof(void*)*2); +x_91 = lean_ctor_get(x_9, 0); +lean_inc(x_91); +lean_dec(x_9); +lean_inc(x_1); +x_92 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_92, 0, x_91); +lean_ctor_set(x_92, 1, x_1); +lean_ctor_set_uint8(x_92, sizeof(void*)*2, x_90); +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); +x_93 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_x3f(x_26, x_4, x_1, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_23); +lean_dec(x_1); +if (lean_obj_tag(x_93) == 0) { -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; +lean_object* x_94; +x_94 = lean_ctor_get(x_93, 0); +lean_inc(x_94); +if (lean_obj_tag(x_94) == 0) +{ +lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; +lean_dec(x_92); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +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_59); - x_61 = lean_box(0); + lean_dec_ref(x_93); + x_96 = lean_box(0); } -lean_ctor_set(x_28, 0, x_5); -if (lean_is_scalar(x_61)) { - x_62 = lean_alloc_ctor(0, 2, 0); +x_97 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_97, 0, x_5); +if (lean_is_scalar(x_96)) { + x_98 = lean_alloc_ctor(0, 2, 0); } else { - x_62 = x_61; + x_98 = x_96; } -lean_ctor_set(x_62, 0, x_28); -lean_ctor_set(x_62, 1, x_60); -return x_62; +lean_ctor_set(x_98, 0, x_97); +lean_ctor_set(x_98, 1, x_95); +return x_98; } else { -lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; -lean_free_object(x_28); -lean_dec(x_5); -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; +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_99 = lean_ctor_get(x_93, 1); +lean_inc(x_99); +lean_dec(x_93); +x_100 = lean_ctor_get(x_94, 0); +lean_inc(x_100); +if (lean_is_exclusive(x_94)) { + lean_ctor_release(x_94, 0); + x_101 = x_94; } else { - lean_dec_ref(x_59); - x_65 = lean_box(0); + lean_dec_ref(x_94); + x_101 = lean_box(0); } -if (lean_is_scalar(x_65)) { - x_66 = lean_alloc_ctor(1, 2, 0); +x_102 = lean_st_ref_take(x_10, x_99); +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 { - x_66 = x_65; -} -lean_ctor_set(x_66, 0, x_63); -lean_ctor_set(x_66, 1, x_64); -return x_66; -} -} + lean_dec_ref(x_102); + x_104 = lean_box(0); } -else -{ -lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; -x_67 = lean_ctor_get(x_28, 0); -lean_inc(x_67); -lean_dec(x_28); -x_68 = lean_st_ref_take(x_10, x_35); -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; +x_105 = lean_box(0); +if (lean_is_scalar(x_104)) { + x_106 = lean_alloc_ctor(1, 2, 0); } else { - lean_dec_ref(x_68); - x_70 = lean_box(0); + x_106 = x_104; + lean_ctor_set_tag(x_106, 1); } -x_71 = lean_box(0); -if (lean_is_scalar(x_70)) { - x_72 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_106, 0, x_100); +lean_ctor_set(x_106, 1, x_105); +x_107 = lean_st_ref_set(x_10, x_106, x_103); +x_108 = lean_ctor_get(x_107, 1); +lean_inc(x_108); +lean_dec(x_107); +x_109 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices(x_92, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_108); +if (lean_obj_tag(x_109) == 0) +{ +lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; +x_110 = lean_ctor_get(x_109, 1); +lean_inc(x_110); +if (lean_is_exclusive(x_109)) { + lean_ctor_release(x_109, 0); + lean_ctor_release(x_109, 1); + x_111 = x_109; } else { - x_72 = x_70; - lean_ctor_set_tag(x_72, 1); + lean_dec_ref(x_109); + x_111 = lean_box(0); } -lean_ctor_set(x_72, 0, x_67); -lean_ctor_set(x_72, 1, x_71); -x_73 = lean_st_ref_set(x_10, x_72, x_69); -x_74 = lean_ctor_get(x_73, 1); -lean_inc(x_74); -lean_dec(x_73); -x_75 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices(x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_74); -if (lean_obj_tag(x_75) == 0) -{ -lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; -x_76 = lean_ctor_get(x_75, 1); -lean_inc(x_76); -if (lean_is_exclusive(x_75)) { - lean_ctor_release(x_75, 0); - lean_ctor_release(x_75, 1); - x_77 = x_75; +if (lean_is_scalar(x_101)) { + x_112 = lean_alloc_ctor(1, 1, 0); } else { - lean_dec_ref(x_75); - x_77 = lean_box(0); + x_112 = x_101; } -x_78 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_78, 0, x_5); -if (lean_is_scalar(x_77)) { - x_79 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_112, 0, x_5); +if (lean_is_scalar(x_111)) { + x_113 = lean_alloc_ctor(0, 2, 0); } else { - x_79 = x_77; + x_113 = x_111; } -lean_ctor_set(x_79, 0, x_78); -lean_ctor_set(x_79, 1, x_76); -return x_79; +lean_ctor_set(x_113, 0, x_112); +lean_ctor_set(x_113, 1, x_110); +return x_113; } else { -lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; +lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; +lean_dec(x_101); lean_dec(x_5); -x_80 = lean_ctor_get(x_75, 0); -lean_inc(x_80); -x_81 = lean_ctor_get(x_75, 1); -lean_inc(x_81); -if (lean_is_exclusive(x_75)) { - lean_ctor_release(x_75, 0); - lean_ctor_release(x_75, 1); - x_82 = x_75; +x_114 = lean_ctor_get(x_109, 0); +lean_inc(x_114); +x_115 = lean_ctor_get(x_109, 1); +lean_inc(x_115); +if (lean_is_exclusive(x_109)) { + lean_ctor_release(x_109, 0); + lean_ctor_release(x_109, 1); + x_116 = x_109; } else { - lean_dec_ref(x_75); - x_82 = lean_box(0); + lean_dec_ref(x_109); + x_116 = lean_box(0); } -if (lean_is_scalar(x_82)) { - x_83 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_116)) { + x_117 = 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; + x_117 = x_116; } +lean_ctor_set(x_117, 0, x_114); +lean_ctor_set(x_117, 1, x_115); +return x_117; } } } else { -uint8_t x_84; +lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; +lean_dec(x_92); lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); @@ -11772,49 +16490,48 @@ lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); -lean_dec(x_9); lean_dec(x_5); -x_84 = !lean_is_exclusive(x_27); -if (x_84 == 0) -{ -return x_27; +x_118 = lean_ctor_get(x_93, 0); +lean_inc(x_118); +x_119 = lean_ctor_get(x_93, 1); +lean_inc(x_119); +if (lean_is_exclusive(x_93)) { + lean_ctor_release(x_93, 0); + lean_ctor_release(x_93, 1); + x_120 = x_93; +} else { + lean_dec_ref(x_93); + x_120 = lean_box(0); } -else -{ -lean_object* x_85; lean_object* x_86; lean_object* x_87; -x_85 = lean_ctor_get(x_27, 0); -x_86 = lean_ctor_get(x_27, 1); -lean_inc(x_86); -lean_inc(x_85); -lean_dec(x_27); -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; +if (lean_is_scalar(x_120)) { + x_121 = lean_alloc_ctor(1, 2, 0); +} else { + x_121 = 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_103; lean_object* x_104; lean_object* x_105; uint8_t x_138; -x_103 = lean_ctor_get(x_20, 0); -x_104 = lean_ctor_get(x_20, 1); -lean_inc(x_104); -lean_inc(x_103); +lean_object* x_136; lean_object* x_137; lean_object* x_138; uint8_t x_175; +x_136 = lean_ctor_get(x_20, 0); +x_137 = lean_ctor_get(x_20, 1); +lean_inc(x_137); +lean_inc(x_136); lean_dec(x_20); -x_138 = lean_ctor_get_uint8(x_103, sizeof(void*)*10 + 4); -if (x_138 == 0) +x_175 = lean_ctor_get_uint8(x_136, sizeof(void*)*10 + 4); +if (x_175 == 0) { -lean_object* x_139; uint8_t x_140; -x_139 = lean_ctor_get(x_103, 3); -lean_inc(x_139); -x_140 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_139, x_1); -lean_dec(x_139); -if (x_140 == 0) +uint8_t x_176; +x_176 = l_Lean_Meta_Grind_ENode_isCongrRoot(x_136); +if (x_176 == 0) { -lean_object* x_141; lean_object* x_142; -lean_dec(x_103); +lean_object* x_177; lean_object* x_178; +lean_dec(x_136); lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); @@ -11828,33 +16545,33 @@ lean_dec(x_9); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_141 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_141, 0, x_5); -x_142 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_142, 0, x_141); -lean_ctor_set(x_142, 1, x_104); -return x_142; +x_177 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_177, 0, x_5); +x_178 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_178, 0, x_177); +lean_ctor_set(x_178, 1, x_137); +return x_178; } else { if (x_6 == 0) { -lean_object* x_143; -x_143 = lean_box(0); -x_105 = x_143; -goto block_137; +lean_object* x_179; +x_179 = lean_box(0); +x_138 = x_179; +goto block_174; } else { -lean_object* x_144; uint8_t x_145; -x_144 = lean_ctor_get(x_103, 9); -lean_inc(x_144); -x_145 = lean_nat_dec_eq(x_144, x_7); -lean_dec(x_144); -if (x_145 == 0) +lean_object* x_180; uint8_t x_181; +x_180 = lean_ctor_get(x_136, 9); +lean_inc(x_180); +x_181 = lean_nat_dec_eq(x_180, x_7); +lean_dec(x_180); +if (x_181 == 0) { -lean_object* x_146; lean_object* x_147; -lean_dec(x_103); +lean_object* x_182; lean_object* x_183; +lean_dec(x_136); lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); @@ -11868,19 +16585,19 @@ lean_dec(x_9); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_146 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_146, 0, x_5); -x_147 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_147, 0, x_146); -lean_ctor_set(x_147, 1, x_104); -return x_147; +x_182 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_182, 0, x_5); +x_183 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_183, 0, x_182); +lean_ctor_set(x_183, 1, x_137); +return x_183; } else { -lean_object* x_148; -x_148 = lean_box(0); -x_105 = x_148; -goto block_137; +lean_object* x_184; +x_184 = lean_box(0); +x_138 = x_184; +goto block_174; } } } @@ -11889,22 +16606,22 @@ else { if (x_6 == 0) { -lean_object* x_149; -x_149 = lean_box(0); -x_105 = x_149; -goto block_137; +lean_object* x_185; +x_185 = lean_box(0); +x_138 = x_185; +goto block_174; } else { -lean_object* x_150; uint8_t x_151; -x_150 = lean_ctor_get(x_103, 9); -lean_inc(x_150); -x_151 = lean_nat_dec_eq(x_150, x_7); -lean_dec(x_150); -if (x_151 == 0) +lean_object* x_186; uint8_t x_187; +x_186 = lean_ctor_get(x_136, 9); +lean_inc(x_186); +x_187 = lean_nat_dec_eq(x_186, x_7); +lean_dec(x_186); +if (x_187 == 0) { -lean_object* x_152; lean_object* x_153; -lean_dec(x_103); +lean_object* x_188; lean_object* x_189; +lean_dec(x_136); lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); @@ -11918,33 +16635,53 @@ lean_dec(x_9); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_152 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_152, 0, x_5); -x_153 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_153, 0, x_152); -lean_ctor_set(x_153, 1, x_104); -return x_153; +x_188 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_188, 0, x_5); +x_189 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_189, 0, x_188); +lean_ctor_set(x_189, 1, x_137); +return x_189; } else { -lean_object* x_154; -x_154 = lean_box(0); -x_105 = x_154; -goto block_137; +lean_object* x_190; +x_190 = lean_box(0); +x_138 = x_190; +goto block_174; } } } -block_137: +block_174: { -lean_object* x_106; lean_object* x_107; lean_object* x_108; -lean_dec(x_105); -x_106 = lean_ctor_get(x_103, 8); -lean_inc(x_106); -lean_dec(x_103); -x_107 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_107, 0, x_2); -lean_ctor_set(x_107, 1, x_106); -lean_ctor_set(x_107, 2, x_3); +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_object* x_145; +lean_dec(x_138); +x_139 = lean_ctor_get(x_136, 8); +lean_inc(x_139); +lean_dec(x_136); +x_140 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_140, 0, x_2); +lean_ctor_set(x_140, 1, x_139); +lean_ctor_set(x_140, 2, x_3); +x_141 = lean_ctor_get_uint8(x_9, sizeof(void*)*2); +x_142 = lean_ctor_get(x_9, 0); +lean_inc(x_142); +if (lean_is_exclusive(x_9)) { + lean_ctor_release(x_9, 0); + lean_ctor_release(x_9, 1); + x_143 = x_9; +} else { + lean_dec_ref(x_9); + x_143 = lean_box(0); +} +lean_inc(x_1); +if (lean_is_scalar(x_143)) { + x_144 = lean_alloc_ctor(0, 2, 1); +} else { + x_144 = x_143; +} +lean_ctor_set(x_144, 0, x_142); +lean_ctor_set(x_144, 1, x_1); +lean_ctor_set_uint8(x_144, sizeof(void*)*2, x_141); lean_inc(x_18); lean_inc(x_17); lean_inc(x_16); @@ -11953,16 +16690,17 @@ lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); -x_108 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_x3f(x_107, x_4, x_1, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_104); +x_145 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_x3f(x_140, x_4, x_1, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_137); lean_dec(x_1); -if (lean_obj_tag(x_108) == 0) +if (lean_obj_tag(x_145) == 0) { -lean_object* x_109; -x_109 = lean_ctor_get(x_108, 0); -lean_inc(x_109); -if (lean_obj_tag(x_109) == 0) +lean_object* x_146; +x_146 = lean_ctor_get(x_145, 0); +lean_inc(x_146); +if (lean_obj_tag(x_146) == 0) { -lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; +lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; +lean_dec(x_144); lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); @@ -11972,127 +16710,127 @@ lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); -lean_dec(x_9); -x_110 = lean_ctor_get(x_108, 1); -lean_inc(x_110); -if (lean_is_exclusive(x_108)) { - lean_ctor_release(x_108, 0); - lean_ctor_release(x_108, 1); - x_111 = x_108; +x_147 = lean_ctor_get(x_145, 1); +lean_inc(x_147); +if (lean_is_exclusive(x_145)) { + lean_ctor_release(x_145, 0); + lean_ctor_release(x_145, 1); + x_148 = x_145; } else { - lean_dec_ref(x_108); - x_111 = lean_box(0); + lean_dec_ref(x_145); + x_148 = lean_box(0); } -x_112 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_112, 0, x_5); -if (lean_is_scalar(x_111)) { - x_113 = lean_alloc_ctor(0, 2, 0); +x_149 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_149, 0, x_5); +if (lean_is_scalar(x_148)) { + x_150 = lean_alloc_ctor(0, 2, 0); } else { - x_113 = x_111; + x_150 = x_148; } -lean_ctor_set(x_113, 0, x_112); -lean_ctor_set(x_113, 1, x_110); -return x_113; +lean_ctor_set(x_150, 0, x_149); +lean_ctor_set(x_150, 1, x_147); +return x_150; } else { -lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; -x_114 = lean_ctor_get(x_108, 1); -lean_inc(x_114); -lean_dec(x_108); -x_115 = lean_ctor_get(x_109, 0); -lean_inc(x_115); -if (lean_is_exclusive(x_109)) { - lean_ctor_release(x_109, 0); - x_116 = x_109; +lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; +x_151 = lean_ctor_get(x_145, 1); +lean_inc(x_151); +lean_dec(x_145); +x_152 = lean_ctor_get(x_146, 0); +lean_inc(x_152); +if (lean_is_exclusive(x_146)) { + lean_ctor_release(x_146, 0); + x_153 = x_146; } else { - lean_dec_ref(x_109); - x_116 = lean_box(0); + lean_dec_ref(x_146); + x_153 = lean_box(0); } -x_117 = lean_st_ref_take(x_10, x_114); -x_118 = lean_ctor_get(x_117, 1); -lean_inc(x_118); -if (lean_is_exclusive(x_117)) { - lean_ctor_release(x_117, 0); - lean_ctor_release(x_117, 1); - x_119 = x_117; +x_154 = lean_st_ref_take(x_10, x_151); +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; } else { - lean_dec_ref(x_117); - x_119 = lean_box(0); + lean_dec_ref(x_154); + x_156 = lean_box(0); } -x_120 = lean_box(0); -if (lean_is_scalar(x_119)) { - x_121 = lean_alloc_ctor(1, 2, 0); +x_157 = lean_box(0); +if (lean_is_scalar(x_156)) { + x_158 = lean_alloc_ctor(1, 2, 0); } else { - x_121 = x_119; - lean_ctor_set_tag(x_121, 1); + x_158 = x_156; + lean_ctor_set_tag(x_158, 1); } -lean_ctor_set(x_121, 0, x_115); -lean_ctor_set(x_121, 1, x_120); -x_122 = lean_st_ref_set(x_10, x_121, x_118); -x_123 = lean_ctor_get(x_122, 1); -lean_inc(x_123); -lean_dec(x_122); -x_124 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices(x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_123); -if (lean_obj_tag(x_124) == 0) +lean_ctor_set(x_158, 0, x_152); +lean_ctor_set(x_158, 1, x_157); +x_159 = lean_st_ref_set(x_10, x_158, x_155); +x_160 = lean_ctor_get(x_159, 1); +lean_inc(x_160); +lean_dec(x_159); +x_161 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices(x_144, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_160); +if (lean_obj_tag(x_161) == 0) { -lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; -x_125 = lean_ctor_get(x_124, 1); -lean_inc(x_125); -if (lean_is_exclusive(x_124)) { - lean_ctor_release(x_124, 0); - lean_ctor_release(x_124, 1); - x_126 = x_124; +lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; +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_124); - x_126 = lean_box(0); + lean_dec_ref(x_161); + x_163 = lean_box(0); } -if (lean_is_scalar(x_116)) { - x_127 = lean_alloc_ctor(1, 1, 0); +if (lean_is_scalar(x_153)) { + x_164 = lean_alloc_ctor(1, 1, 0); } else { - x_127 = x_116; + x_164 = x_153; } -lean_ctor_set(x_127, 0, x_5); -if (lean_is_scalar(x_126)) { - x_128 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_164, 0, x_5); +if (lean_is_scalar(x_163)) { + x_165 = lean_alloc_ctor(0, 2, 0); } else { - x_128 = x_126; + x_165 = x_163; } -lean_ctor_set(x_128, 0, x_127); -lean_ctor_set(x_128, 1, x_125); -return x_128; +lean_ctor_set(x_165, 0, x_164); +lean_ctor_set(x_165, 1, x_162); +return x_165; } else { -lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; -lean_dec(x_116); +lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; +lean_dec(x_153); lean_dec(x_5); -x_129 = lean_ctor_get(x_124, 0); -lean_inc(x_129); -x_130 = lean_ctor_get(x_124, 1); -lean_inc(x_130); -if (lean_is_exclusive(x_124)) { - lean_ctor_release(x_124, 0); - lean_ctor_release(x_124, 1); - x_131 = x_124; +x_166 = lean_ctor_get(x_161, 0); +lean_inc(x_166); +x_167 = lean_ctor_get(x_161, 1); +lean_inc(x_167); +if (lean_is_exclusive(x_161)) { + lean_ctor_release(x_161, 0); + lean_ctor_release(x_161, 1); + x_168 = x_161; } else { - lean_dec_ref(x_124); - x_131 = lean_box(0); + lean_dec_ref(x_161); + x_168 = lean_box(0); } -if (lean_is_scalar(x_131)) { - x_132 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_168)) { + x_169 = lean_alloc_ctor(1, 2, 0); } else { - x_132 = x_131; + x_169 = x_168; } -lean_ctor_set(x_132, 0, x_129); -lean_ctor_set(x_132, 1, x_130); -return x_132; +lean_ctor_set(x_169, 0, x_166); +lean_ctor_set(x_169, 1, x_167); +return x_169; } } } else { -lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; +lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; +lean_dec(x_144); lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); @@ -12102,35 +16840,34 @@ lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); -lean_dec(x_9); lean_dec(x_5); -x_133 = lean_ctor_get(x_108, 0); -lean_inc(x_133); -x_134 = lean_ctor_get(x_108, 1); -lean_inc(x_134); -if (lean_is_exclusive(x_108)) { - lean_ctor_release(x_108, 0); - lean_ctor_release(x_108, 1); - x_135 = x_108; +x_170 = lean_ctor_get(x_145, 0); +lean_inc(x_170); +x_171 = lean_ctor_get(x_145, 1); +lean_inc(x_171); +if (lean_is_exclusive(x_145)) { + lean_ctor_release(x_145, 0); + lean_ctor_release(x_145, 1); + x_172 = x_145; } else { - lean_dec_ref(x_108); - x_135 = lean_box(0); + lean_dec_ref(x_145); + x_172 = lean_box(0); } -if (lean_is_scalar(x_135)) { - x_136 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_172)) { + x_173 = lean_alloc_ctor(1, 2, 0); } else { - x_136 = x_135; + x_173 = x_172; } -lean_ctor_set(x_136, 0, x_133); -lean_ctor_set(x_136, 1, x_134); -return x_136; +lean_ctor_set(x_173, 0, x_170); +lean_ctor_set(x_173, 1, x_171); +return x_173; } } } } else { -uint8_t x_155; +uint8_t x_191; lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); @@ -12145,23 +16882,23 @@ lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_155 = !lean_is_exclusive(x_20); -if (x_155 == 0) +x_191 = !lean_is_exclusive(x_20); +if (x_191 == 0) { return x_20; } else { -lean_object* x_156; lean_object* x_157; lean_object* x_158; -x_156 = lean_ctor_get(x_20, 0); -x_157 = lean_ctor_get(x_20, 1); -lean_inc(x_157); -lean_inc(x_156); +lean_object* x_192; lean_object* x_193; lean_object* x_194; +x_192 = lean_ctor_get(x_20, 0); +x_193 = lean_ctor_get(x_20, 1); +lean_inc(x_193); +lean_inc(x_192); lean_dec(x_20); -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_194 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_194, 0, x_192); +lean_ctor_set(x_194, 1, x_193); +return x_194; } } } @@ -12422,7 +17159,7 @@ lean_inc(x_24); lean_dec(x_23); x_25 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_unassigned; x_26 = lean_mk_array(x_24, x_25); -x_27 = lean_ctor_get_uint8(x_3, sizeof(void*)*1); +x_27 = lean_ctor_get_uint8(x_3, sizeof(void*)*2); x_28 = lean_st_ref_get(x_5, x_17); x_29 = lean_ctor_get(x_28, 0); lean_inc(x_29); @@ -12433,7 +17170,7 @@ x_31 = lean_ctor_get(x_29, 6); lean_inc(x_31); lean_dec(x_29); x_32 = lean_box(0); -x_33 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__5___closed__1; +x_33 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__4___closed__1; lean_inc(x_22); x_34 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_main___spec__1(x_1, x_2, x_22, x_26, x_27, x_31, x_32, x_33, x_22, x_22, x_33, lean_box(0), x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_30); lean_dec(x_31); @@ -12577,7 +17314,7 @@ lean_inc(x_62); lean_dec(x_61); x_63 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_unassigned; x_64 = lean_mk_array(x_62, x_63); -x_65 = lean_ctor_get_uint8(x_3, sizeof(void*)*1); +x_65 = lean_ctor_get_uint8(x_3, sizeof(void*)*2); x_66 = lean_st_ref_get(x_5, x_54); x_67 = lean_ctor_get(x_66, 0); lean_inc(x_67); @@ -12588,7 +17325,7 @@ x_69 = lean_ctor_get(x_67, 6); lean_inc(x_69); lean_dec(x_67); x_70 = lean_box(0); -x_71 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__5___closed__1; +x_71 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__4___closed__1; lean_inc(x_60); x_72 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_main___spec__1(x_1, x_2, x_60, x_64, x_65, x_69, x_70, x_71, x_60, x_60, x_71, lean_box(0), x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_68); lean_dec(x_69); @@ -12767,7 +17504,7 @@ if (x_4 == 0) lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_1, 0); x_6 = lean_ctor_get(x_1, 1); -x_7 = lean_alloc_ctor(1, 1, 0); +x_7 = lean_alloc_ctor(2, 1, 0); lean_ctor_set(x_7, 0, x_5); lean_ctor_set(x_1, 1, x_2); lean_ctor_set(x_1, 0, x_7); @@ -12787,7 +17524,7 @@ x_10 = lean_ctor_get(x_1, 1); lean_inc(x_10); lean_inc(x_9); lean_dec(x_1); -x_11 = lean_alloc_ctor(1, 1, 0); +x_11 = lean_alloc_ctor(2, 1, 0); lean_ctor_set(x_11, 0, x_9); x_12 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_12, 0, x_11); @@ -12855,7 +17592,7 @@ lean_object* x_24; lean_object* x_25; x_24 = lean_ctor_get(x_23, 1); lean_inc(x_24); lean_dec(x_23); -x_25 = lean_alloc_ctor(1, 1, 0); +x_25 = lean_alloc_ctor(2, 1, 0); lean_ctor_set(x_25, 0, x_17); lean_ctor_set(x_1, 1, x_2); lean_ctor_set(x_1, 0, x_25); @@ -12938,7 +17675,7 @@ lean_object* x_38; lean_object* x_39; lean_object* x_40; x_38 = lean_ctor_get(x_37, 1); lean_inc(x_38); lean_dec(x_37); -x_39 = lean_alloc_ctor(1, 1, 0); +x_39 = lean_alloc_ctor(2, 1, 0); lean_ctor_set(x_39, 0, x_31); x_40 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_40, 0, x_39); @@ -13003,7 +17740,7 @@ static lean_object* _init_l_Lean_Meta_Grind_EMatch_ematchTheorem___lambda__1___c lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___closed__1; x_2 = l_Lean_Meta_Grind_EMatch_ematchTheorem___lambda__1___closed__1; -x_3 = lean_unsigned_to_nat(242u); +x_3 = lean_unsigned_to_nat(315u); x_4 = lean_unsigned_to_nat(22u); x_5 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -13020,7 +17757,7 @@ x_15 = !lean_is_exclusive(x_3); if (x_15 == 0) { uint8_t x_16; lean_object* x_17; -x_16 = lean_ctor_get_uint8(x_3, sizeof(void*)*1); +x_16 = lean_ctor_get_uint8(x_3, sizeof(void*)*2); x_17 = lean_ctor_get(x_3, 0); lean_dec(x_17); lean_ctor_set(x_3, 0, x_1); @@ -13072,54 +17809,57 @@ return x_29; } else { -uint8_t x_30; lean_object* x_31; -x_30 = lean_ctor_get_uint8(x_3, sizeof(void*)*1); +uint8_t x_30; lean_object* x_31; lean_object* x_32; +x_30 = lean_ctor_get_uint8(x_3, sizeof(void*)*2); +x_31 = lean_ctor_get(x_3, 1); +lean_inc(x_31); lean_dec(x_3); -x_31 = lean_alloc_ctor(0, 1, 1); -lean_ctor_set(x_31, 0, x_1); -lean_ctor_set_uint8(x_31, sizeof(void*)*1, x_30); +x_32 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_32, 0, x_1); +lean_ctor_set(x_32, 1, x_31); +lean_ctor_set_uint8(x_32, sizeof(void*)*2, x_30); if (lean_obj_tag(x_14) == 0) { -lean_object* x_32; lean_object* x_33; -x_32 = l_Lean_Meta_Grind_EMatch_ematchTheorem___lambda__1___closed__2; -x_33 = l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__2(x_32, x_31, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -return x_33; +lean_object* x_33; lean_object* x_34; +x_33 = l_Lean_Meta_Grind_EMatch_ematchTheorem___lambda__1___closed__2; +x_34 = l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__2(x_33, x_32, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_34; } else { -lean_object* x_34; -x_34 = lean_ctor_get(x_14, 1); -lean_inc(x_34); -if (lean_obj_tag(x_34) == 0) -{ -lean_object* x_35; lean_object* x_36; lean_object* x_37; -x_35 = lean_ctor_get(x_14, 0); +lean_object* x_35; +x_35 = lean_ctor_get(x_14, 1); lean_inc(x_35); +if (lean_obj_tag(x_35) == 0) +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_14, 0); +lean_inc(x_36); lean_dec(x_14); -x_36 = lean_box(0); -x_37 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_main(x_35, x_36, x_31, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -return x_37; +x_37 = lean_box(0); +x_38 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_main(x_36, x_37, x_32, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_38; } else { if (x_30 == 0) { -lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; -x_38 = lean_ctor_get(x_14, 0); -lean_inc(x_38); +lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_39 = lean_ctor_get(x_14, 0); +lean_inc(x_39); lean_dec(x_14); -x_39 = lean_box(0); -x_40 = l_List_mapTR_loop___at_Lean_Meta_Grind_EMatch_ematchTheorem_tryAll___spec__1(x_34, x_39); -x_41 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_main(x_38, x_40, x_31, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -return x_41; +x_40 = lean_box(0); +x_41 = l_List_mapTR_loop___at_Lean_Meta_Grind_EMatch_ematchTheorem_tryAll___spec__1(x_35, x_40); +x_42 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_main(x_39, x_41, x_32, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_42; } else { -lean_object* x_42; lean_object* x_43; -lean_dec(x_34); -x_42 = lean_box(0); -x_43 = l_Lean_Meta_Grind_EMatch_ematchTheorem_tryAll(x_14, x_42, x_31, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -return x_43; +lean_object* x_43; lean_object* x_44; +lean_dec(x_35); +x_43 = lean_box(0); +x_44 = l_Lean_Meta_Grind_EMatch_ematchTheorem_tryAll(x_14, x_43, x_32, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_44; } } } @@ -13757,22 +18497,25 @@ if (x_13 == 0) { uint8_t x_14; lean_object* x_15; x_14 = 1; -lean_ctor_set_uint8(x_2, sizeof(void*)*1, x_14); +lean_ctor_set_uint8(x_2, sizeof(void*)*2, x_14); x_15 = l_Lean_PersistentArray_forM___at_Lean_Meta_Grind_EMatch_ematchTheorems___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); return x_15; } else { -lean_object* x_16; uint8_t x_17; lean_object* x_18; lean_object* x_19; +lean_object* x_16; lean_object* x_17; uint8_t x_18; lean_object* x_19; lean_object* x_20; x_16 = lean_ctor_get(x_2, 0); +x_17 = lean_ctor_get(x_2, 1); +lean_inc(x_17); lean_inc(x_16); lean_dec(x_2); -x_17 = 1; -x_18 = lean_alloc_ctor(0, 1, 1); -lean_ctor_set(x_18, 0, x_16); -lean_ctor_set_uint8(x_18, sizeof(void*)*1, x_17); -x_19 = l_Lean_PersistentArray_forM___at_Lean_Meta_Grind_EMatch_ematchTheorems___spec__1(x_1, x_18, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -return x_19; +x_18 = 1; +x_19 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_19, 0, x_16); +lean_ctor_set(x_19, 1, x_17); +lean_ctor_set_uint8(x_19, sizeof(void*)*2, x_18); +x_20 = l_Lean_PersistentArray_forM___at_Lean_Meta_Grind_EMatch_ematchTheorems___spec__1(x_1, x_19, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_20; } } } @@ -13785,22 +18528,25 @@ if (x_14 == 0) { uint8_t x_15; lean_object* x_16; x_15 = 0; -lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_15); +lean_ctor_set_uint8(x_3, sizeof(void*)*2, x_15); x_16 = l_Lean_PersistentArray_forM___at_Lean_Meta_Grind_EMatch_ematchTheorems___spec__1(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); return x_16; } else { -lean_object* x_17; uint8_t x_18; lean_object* x_19; lean_object* x_20; +lean_object* x_17; lean_object* x_18; uint8_t x_19; lean_object* x_20; lean_object* x_21; x_17 = lean_ctor_get(x_3, 0); +x_18 = lean_ctor_get(x_3, 1); +lean_inc(x_18); lean_inc(x_17); lean_dec(x_3); -x_18 = 0; -x_19 = lean_alloc_ctor(0, 1, 1); -lean_ctor_set(x_19, 0, x_17); -lean_ctor_set_uint8(x_19, sizeof(void*)*1, x_18); -x_20 = l_Lean_PersistentArray_forM___at_Lean_Meta_Grind_EMatch_ematchTheorems___spec__1(x_1, x_19, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -return x_20; +x_19 = 0; +x_20 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_20, 0, x_17); +lean_ctor_set(x_20, 1, x_18); +lean_ctor_set_uint8(x_20, sizeof(void*)*2, x_19); +x_21 = l_Lean_PersistentArray_forM___at_Lean_Meta_Grind_EMatch_ematchTheorems___spec__1(x_1, x_20, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_21; } } } @@ -13843,74 +18589,107 @@ return x_5; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematch(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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_Meta_Grind_checkMaxInstancesExceeded(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -x_11 = lean_ctor_get(x_10, 0); -lean_inc(x_11); -x_12 = lean_unbox(x_11); -lean_dec(x_11); -if (x_12 == 0) +uint8_t x_10; lean_object* x_11; lean_object* x_84; lean_object* x_85; uint8_t x_86; +x_84 = l_Lean_Meta_Grind_checkMaxInstancesExceeded(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_85 = lean_ctor_get(x_84, 0); +lean_inc(x_85); +x_86 = lean_unbox(x_85); +if (x_86 == 0) { -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; -x_13 = lean_ctor_get(x_10, 1); +lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; uint8_t x_91; +lean_dec(x_85); +x_87 = lean_ctor_get(x_84, 1); +lean_inc(x_87); +lean_dec(x_84); +x_88 = l_Lean_Meta_Grind_checkMaxEmatchExceeded(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_87); +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_unbox(x_89); +lean_dec(x_89); +x_10 = x_91; +x_11 = x_90; +goto block_83; +} +else +{ +lean_object* x_92; uint8_t x_93; +x_92 = lean_ctor_get(x_84, 1); +lean_inc(x_92); +lean_dec(x_84); +x_93 = lean_unbox(x_85); +lean_dec(x_85); +x_10 = x_93; +x_11 = x_92; +goto block_83; +} +block_83: +{ +if (x_10 == 0) +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_12 = lean_st_ref_get(x_1, x_11); +x_13 = lean_ctor_get(x_12, 0); lean_inc(x_13); -lean_dec(x_10); -x_14 = lean_st_ref_get(x_1, x_13); -x_15 = lean_ctor_get(x_14, 0); -lean_inc(x_15); -x_16 = lean_ctor_get(x_14, 1); +x_14 = lean_ctor_get(x_12, 1); +lean_inc(x_14); +lean_dec(x_12); +x_15 = lean_st_ref_get(x_1, x_14); +x_16 = lean_ctor_get(x_15, 0); lean_inc(x_16); -lean_dec(x_14); -x_17 = lean_st_ref_get(x_1, x_16); -x_18 = lean_ctor_get(x_17, 0); +x_17 = lean_ctor_get(x_15, 1); +lean_inc(x_17); +lean_dec(x_15); +x_18 = lean_ctor_get(x_13, 8); lean_inc(x_18); -x_19 = lean_ctor_get(x_17, 1); +lean_dec(x_13); +x_19 = lean_ctor_get(x_16, 9); lean_inc(x_19); -lean_dec(x_17); -x_20 = lean_ctor_get(x_15, 8); -lean_inc(x_20); -lean_dec(x_15); -x_21 = lean_ctor_get(x_18, 9); -lean_inc(x_21); -lean_dec(x_18); -x_22 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_ematch___lambda__1___boxed), 12, 1); +lean_dec(x_16); +x_20 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_ematch___lambda__1___boxed), 12, 1); +lean_closure_set(x_20, 0, x_18); +x_21 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_ematch___lambda__2___boxed), 13, 1); +lean_closure_set(x_21, 0, x_19); +x_22 = lean_alloc_closure((void*)(l_ReaderT_bind___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__11___rarg), 13, 2); lean_closure_set(x_22, 0, x_20); -x_23 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_ematch___lambda__2___boxed), 13, 1); -lean_closure_set(x_23, 0, x_21); -x_24 = lean_alloc_closure((void*)(l_ReaderT_bind___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__11___rarg), 13, 2); -lean_closure_set(x_24, 0, x_22); -lean_closure_set(x_24, 1, x_23); +lean_closure_set(x_22, 1, x_21); lean_inc(x_1); -x_25 = l_Lean_Meta_Grind_EMatch_M_run_x27___rarg(x_24, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_19); -if (lean_obj_tag(x_25) == 0) +x_23 = l_Lean_Meta_Grind_EMatch_M_run_x27___rarg(x_22, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_17); +if (lean_obj_tag(x_23) == 0) { -lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; uint8_t x_30; -x_26 = lean_ctor_get(x_25, 1); +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; +x_24 = lean_ctor_get(x_23, 1); +lean_inc(x_24); +lean_dec(x_23); +x_25 = lean_st_ref_take(x_1, 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_27 = lean_st_ref_take(x_1, x_26); -x_28 = lean_ctor_get(x_27, 0); -lean_inc(x_28); -x_29 = lean_ctor_get(x_27, 1); -lean_inc(x_29); -lean_dec(x_27); -x_30 = !lean_is_exclusive(x_28); -if (x_30 == 0) +x_28 = !lean_is_exclusive(x_26); +if (x_28 == 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; uint8_t x_39; -x_31 = lean_ctor_get(x_28, 6); -x_32 = lean_ctor_get(x_28, 8); -x_33 = lean_ctor_get(x_28, 9); -x_34 = lean_unsigned_to_nat(1u); -x_35 = lean_nat_add(x_31, x_34); +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; uint8_t x_39; +x_29 = lean_ctor_get(x_26, 6); +x_30 = lean_ctor_get(x_26, 8); +x_31 = lean_ctor_get(x_26, 9); +x_32 = lean_ctor_get(x_26, 12); +x_33 = lean_unsigned_to_nat(1u); +x_34 = lean_nat_add(x_29, x_33); +lean_dec(x_29); +x_35 = l_Lean_PersistentArray_append___rarg(x_30, x_31); lean_dec(x_31); -x_36 = l_Lean_PersistentArray_append___rarg(x_32, x_33); -lean_dec(x_33); +x_36 = lean_nat_add(x_32, x_33); +lean_dec(x_32); x_37 = l_Lean_Meta_Grind_ematch___closed__3; -lean_ctor_set(x_28, 9, x_37); -lean_ctor_set(x_28, 8, x_36); -lean_ctor_set(x_28, 6, x_35); -x_38 = lean_st_ref_set(x_1, x_28, x_29); +lean_ctor_set(x_26, 12, x_36); +lean_ctor_set(x_26, 9, x_37); +lean_ctor_set(x_26, 8, x_35); +lean_ctor_set(x_26, 6, x_34); +x_38 = lean_st_ref_set(x_1, x_26, x_27); lean_dec(x_1); x_39 = !lean_is_exclusive(x_38); if (x_39 == 0) @@ -13937,22 +18716,34 @@ return x_44; } else { -lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; 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; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; -x_45 = lean_ctor_get(x_28, 0); -x_46 = lean_ctor_get(x_28, 1); -x_47 = lean_ctor_get(x_28, 2); -x_48 = lean_ctor_get(x_28, 3); -x_49 = lean_ctor_get(x_28, 4); -x_50 = lean_ctor_get(x_28, 5); -x_51 = lean_ctor_get_uint8(x_28, sizeof(void*)*14); -x_52 = lean_ctor_get(x_28, 6); -x_53 = lean_ctor_get(x_28, 7); -x_54 = lean_ctor_get(x_28, 8); -x_55 = lean_ctor_get(x_28, 9); -x_56 = lean_ctor_get(x_28, 10); -x_57 = lean_ctor_get(x_28, 11); -x_58 = lean_ctor_get(x_28, 12); -x_59 = lean_ctor_get(x_28, 13); +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; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; +x_45 = lean_ctor_get(x_26, 0); +x_46 = lean_ctor_get(x_26, 1); +x_47 = lean_ctor_get(x_26, 2); +x_48 = lean_ctor_get(x_26, 3); +x_49 = lean_ctor_get(x_26, 4); +x_50 = lean_ctor_get(x_26, 5); +x_51 = lean_ctor_get_uint8(x_26, sizeof(void*)*20); +x_52 = lean_ctor_get(x_26, 6); +x_53 = lean_ctor_get(x_26, 7); +x_54 = lean_ctor_get(x_26, 8); +x_55 = lean_ctor_get(x_26, 9); +x_56 = lean_ctor_get(x_26, 10); +x_57 = lean_ctor_get(x_26, 11); +x_58 = lean_ctor_get(x_26, 12); +x_59 = lean_ctor_get(x_26, 13); +x_60 = lean_ctor_get(x_26, 14); +x_61 = lean_ctor_get(x_26, 15); +x_62 = lean_ctor_get(x_26, 16); +x_63 = lean_ctor_get(x_26, 17); +x_64 = lean_ctor_get(x_26, 18); +x_65 = lean_ctor_get(x_26, 19); +lean_inc(x_65); +lean_inc(x_64); +lean_inc(x_63); +lean_inc(x_62); +lean_inc(x_61); +lean_inc(x_60); lean_inc(x_59); lean_inc(x_58); lean_inc(x_57); @@ -13967,79 +18758,87 @@ lean_inc(x_48); lean_inc(x_47); lean_inc(x_46); lean_inc(x_45); -lean_dec(x_28); -x_60 = lean_unsigned_to_nat(1u); -x_61 = lean_nat_add(x_52, x_60); +lean_dec(x_26); +x_66 = lean_unsigned_to_nat(1u); +x_67 = lean_nat_add(x_52, x_66); lean_dec(x_52); -x_62 = l_Lean_PersistentArray_append___rarg(x_54, x_55); +x_68 = l_Lean_PersistentArray_append___rarg(x_54, x_55); lean_dec(x_55); -x_63 = l_Lean_Meta_Grind_ematch___closed__3; -x_64 = lean_alloc_ctor(0, 14, 1); -lean_ctor_set(x_64, 0, x_45); -lean_ctor_set(x_64, 1, x_46); -lean_ctor_set(x_64, 2, x_47); -lean_ctor_set(x_64, 3, x_48); -lean_ctor_set(x_64, 4, x_49); -lean_ctor_set(x_64, 5, x_50); -lean_ctor_set(x_64, 6, x_61); -lean_ctor_set(x_64, 7, x_53); -lean_ctor_set(x_64, 8, x_62); -lean_ctor_set(x_64, 9, x_63); -lean_ctor_set(x_64, 10, x_56); -lean_ctor_set(x_64, 11, x_57); -lean_ctor_set(x_64, 12, x_58); -lean_ctor_set(x_64, 13, x_59); -lean_ctor_set_uint8(x_64, sizeof(void*)*14, x_51); -x_65 = lean_st_ref_set(x_1, x_64, x_29); -lean_dec(x_1); -x_66 = lean_ctor_get(x_65, 1); -lean_inc(x_66); -if (lean_is_exclusive(x_65)) { - lean_ctor_release(x_65, 0); - lean_ctor_release(x_65, 1); - x_67 = x_65; +x_69 = lean_nat_add(x_58, x_66); +lean_dec(x_58); +x_70 = l_Lean_Meta_Grind_ematch___closed__3; +x_71 = lean_alloc_ctor(0, 20, 1); +lean_ctor_set(x_71, 0, x_45); +lean_ctor_set(x_71, 1, x_46); +lean_ctor_set(x_71, 2, x_47); +lean_ctor_set(x_71, 3, x_48); +lean_ctor_set(x_71, 4, x_49); +lean_ctor_set(x_71, 5, x_50); +lean_ctor_set(x_71, 6, x_67); +lean_ctor_set(x_71, 7, x_53); +lean_ctor_set(x_71, 8, x_68); +lean_ctor_set(x_71, 9, x_70); +lean_ctor_set(x_71, 10, x_56); +lean_ctor_set(x_71, 11, x_57); +lean_ctor_set(x_71, 12, x_69); +lean_ctor_set(x_71, 13, x_59); +lean_ctor_set(x_71, 14, x_60); +lean_ctor_set(x_71, 15, x_61); +lean_ctor_set(x_71, 16, x_62); +lean_ctor_set(x_71, 17, x_63); +lean_ctor_set(x_71, 18, x_64); +lean_ctor_set(x_71, 19, x_65); +lean_ctor_set_uint8(x_71, sizeof(void*)*20, x_51); +x_72 = lean_st_ref_set(x_1, x_71, x_27); +lean_dec(x_1); +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_65); - x_67 = lean_box(0); + lean_dec_ref(x_72); + x_74 = lean_box(0); } -x_68 = lean_box(0); -if (lean_is_scalar(x_67)) { - x_69 = lean_alloc_ctor(0, 2, 0); +x_75 = lean_box(0); +if (lean_is_scalar(x_74)) { + x_76 = lean_alloc_ctor(0, 2, 0); } else { - x_69 = x_67; + x_76 = x_74; } -lean_ctor_set(x_69, 0, x_68); -lean_ctor_set(x_69, 1, x_66); -return x_69; +lean_ctor_set(x_76, 0, x_75); +lean_ctor_set(x_76, 1, x_73); +return x_76; } } else { -uint8_t x_70; +uint8_t x_77; lean_dec(x_1); -x_70 = !lean_is_exclusive(x_25); -if (x_70 == 0) +x_77 = !lean_is_exclusive(x_23); +if (x_77 == 0) { -return x_25; +return x_23; } else { -lean_object* x_71; lean_object* x_72; lean_object* x_73; -x_71 = lean_ctor_get(x_25, 0); -x_72 = lean_ctor_get(x_25, 1); -lean_inc(x_72); -lean_inc(x_71); -lean_dec(x_25); -x_73 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_73, 0, x_71); -lean_ctor_set(x_73, 1, x_72); -return x_73; +lean_object* x_78; lean_object* x_79; lean_object* x_80; +x_78 = lean_ctor_get(x_23, 0); +x_79 = lean_ctor_get(x_23, 1); +lean_inc(x_79); +lean_inc(x_78); +lean_dec(x_23); +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 { -uint8_t x_74; +lean_object* x_81; lean_object* x_82; lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -14048,27 +18847,11 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_74 = !lean_is_exclusive(x_10); -if (x_74 == 0) -{ -lean_object* x_75; lean_object* x_76; -x_75 = lean_ctor_get(x_10, 0); -lean_dec(x_75); -x_76 = lean_box(0); -lean_ctor_set(x_10, 0, x_76); -return x_10; -} -else -{ -lean_object* x_77; lean_object* x_78; lean_object* x_79; -x_77 = lean_ctor_get(x_10, 1); -lean_inc(x_77); -lean_dec(x_10); -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; +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_11); +return x_82; } } } @@ -14092,7 +18875,7 @@ lean_dec(x_1); return x_14; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematchAndAssert_x3f___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematchAndAssert___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; uint8_t x_11; @@ -14117,7 +18900,7 @@ return x_14; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematchAndAssert_x3f___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematchAndAssert___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; @@ -14221,7 +19004,7 @@ return x_32; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematchAndAssert_x3f___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematchAndAssert___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; @@ -14247,90 +19030,40 @@ return x_13; } } } -static lean_object* _init_l_Lean_Meta_Grind_ematchAndAssert_x3f___lambda__4___closed__1() { +static lean_object* _init_l_Lean_Meta_Grind_ematchAndAssert___lambda__4___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_assertNext_x3f), 9, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_assertNext), 9, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematchAndAssert_x3f___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematchAndAssert___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { lean_object* x_11; lean_object* x_12; -x_11 = l_Lean_Meta_Grind_ematchAndAssert_x3f___lambda__4___closed__1; -x_12 = l_Lean_Meta_Grind_iterate(x_1, 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) -{ -uint8_t x_13; -x_13 = !lean_is_exclusive(x_12); -if (x_13 == 0) -{ -lean_object* x_14; lean_object* x_15; -x_14 = lean_ctor_get(x_12, 0); -x_15 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_15, 0, x_14); -lean_ctor_set(x_12, 0, x_15); -return x_12; -} -else -{ -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_16 = lean_ctor_get(x_12, 0); -x_17 = lean_ctor_get(x_12, 1); -lean_inc(x_17); -lean_inc(x_16); -lean_dec(x_12); -x_18 = lean_alloc_ctor(1, 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; -} -} -else -{ -uint8_t x_20; -x_20 = !lean_is_exclusive(x_12); -if (x_20 == 0) -{ +x_11 = l_Lean_Meta_Grind_ematchAndAssert___lambda__4___closed__1; +x_12 = l_Lean_Meta_Grind_GrindTactic_iterate(x_11, x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); 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; -} -} -} } -static lean_object* _init_l_Lean_Meta_Grind_ematchAndAssert_x3f___closed__1() { +static lean_object* _init_l_Lean_Meta_Grind_ematchAndAssert___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_ematchAndAssert_x3f___lambda__2), 9, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_ematchAndAssert___lambda__2), 9, 0); return x_1; } } -static lean_object* _init_l_Lean_Meta_Grind_ematchAndAssert_x3f___closed__2() { +static lean_object* _init_l_Lean_Meta_Grind_ematchAndAssert___closed__2() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_ematchAndAssert_x3f___lambda__3___boxed), 9, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_ematchAndAssert___lambda__3___boxed), 9, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematchAndAssert_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_EXPORT lean_object* l_Lean_Meta_Grind_ematchAndAssert(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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; @@ -14338,13 +19071,13 @@ x_10 = lean_ctor_get(x_1, 11); lean_inc(x_10); x_11 = lean_ctor_get(x_1, 0); lean_inc(x_11); -x_12 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_ematchAndAssert_x3f___lambda__1___boxed), 9, 1); +x_12 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_ematchAndAssert___lambda__1___boxed), 9, 1); lean_closure_set(x_12, 0, x_1); -x_13 = l_Lean_Meta_Grind_ematchAndAssert_x3f___closed__1; +x_13 = l_Lean_Meta_Grind_ematchAndAssert___closed__1; x_14 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); lean_closure_set(x_14, 0, x_12); lean_closure_set(x_14, 1, x_13); -x_15 = l_Lean_Meta_Grind_ematchAndAssert_x3f___closed__2; +x_15 = l_Lean_Meta_Grind_ematchAndAssert___closed__2; x_16 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); lean_closure_set(x_16, 0, x_14); lean_closure_set(x_16, 1, x_15); @@ -14374,8 +19107,8 @@ if (x_22 == 0) { lean_object* x_23; lean_object* x_24; lean_free_object(x_17); -x_23 = lean_box(0); -x_24 = l_Lean_Meta_Grind_ematchAndAssert_x3f___lambda__4(x_19, x_23, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_20); +x_23 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_assertNext), 9, 0); +x_24 = l_Lean_Meta_Grind_GrindTactic_iterate(x_23, x_19, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_20); return x_24; } else @@ -14410,8 +19143,8 @@ lean_dec(x_28); if (x_29 == 0) { lean_object* x_30; lean_object* x_31; -x_30 = lean_box(0); -x_31 = l_Lean_Meta_Grind_ematchAndAssert_x3f___lambda__4(x_26, x_30, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_27); +x_30 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_assertNext), 9, 0); +x_31 = l_Lean_Meta_Grind_GrindTactic_iterate(x_30, x_26, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_27); return x_31; } else @@ -14465,11 +19198,11 @@ return x_37; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematchAndAssert_x3f___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematchAndAssert___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { lean_object* x_10; -x_10 = l_Lean_Meta_Grind_ematchAndAssert_x3f___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_10 = l_Lean_Meta_Grind_ematchAndAssert___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); @@ -14480,11 +19213,11 @@ lean_dec(x_2); return x_10; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematchAndAssert_x3f___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematchAndAssert___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { lean_object* x_10; -x_10 = l_Lean_Meta_Grind_ematchAndAssert_x3f___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_10 = l_Lean_Meta_Grind_ematchAndAssert___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); @@ -14495,11 +19228,11 @@ lean_dec(x_2); return x_10; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematchAndAssert_x3f___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematchAndAssert___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_11; -x_11 = l_Lean_Meta_Grind_ematchAndAssert_x3f___lambda__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_Meta_Grind_ematchAndAssert___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_2); return x_11; } @@ -14508,7 +19241,7 @@ static lean_object* _init_l_Lean_Meta_Grind_ematchStar___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_ematchAndAssert_x3f), 9, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_ematchAndAssert), 9, 0); return x_1; } } @@ -14517,12 +19250,13 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematchStar(lean_object* x_1, lean_obj { lean_object* x_10; lean_object* x_11; x_10 = l_Lean_Meta_Grind_ematchStar___closed__1; -x_11 = l_Lean_Meta_Grind_iterate(x_1, x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_11 = l_Lean_Meta_Grind_GrindTactic_iterate(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* initialize_Lean_Meta_Tactic_Grind_Types(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_Tactic_Grind_Intro(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Meta_Tactic_Grind_DoNotSimp(uint8_t builtin, lean_object*); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Lean_Meta_Tactic_Grind_EMatch(uint8_t builtin, lean_object* w) { lean_object * res; @@ -14534,6 +19268,9 @@ lean_dec_ref(res); res = initialize_Lean_Meta_Tactic_Grind_Intro(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +res = initialize_Lean_Meta_Tactic_Grind_DoNotSimp(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); l_Lean_Meta_Grind_EMatch_instInhabitedCnstr___closed__1 = _init_l_Lean_Meta_Grind_EMatch_instInhabitedCnstr___closed__1(); lean_mark_persistent(l_Lean_Meta_Grind_EMatch_instInhabitedCnstr___closed__1); l_Lean_Meta_Grind_EMatch_instInhabitedCnstr___closed__2 = _init_l_Lean_Meta_Grind_EMatch_instInhabitedCnstr___closed__2(); @@ -14594,8 +19331,36 @@ l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f__ lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___closed__3); l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___closed__4(); lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___closed__4); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__1); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__2); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__3); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__4(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__4); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__5 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__5(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__5); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__6 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__6(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__6); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__7 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__7(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__7); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__8 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__8(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__8); l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___spec__1___closed__1 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___spec__1___closed__1(); lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___spec__1___closed__1); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___lambda__1___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___lambda__1___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___lambda__1___closed__1); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___lambda__1___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___lambda__1___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___lambda__1___closed__2); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___lambda__1___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___lambda__1___closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___lambda__1___closed__3); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___lambda__1___closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___lambda__1___closed__4(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___lambda__1___closed__4); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___closed__1); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___closed__2); l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5___closed__1 = _init_l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5___closed__1(); lean_mark_persistent(l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5___closed__1); l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5___closed__2 = _init_l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5___closed__2(); @@ -14604,12 +19369,6 @@ l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Gr lean_mark_persistent(l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5___closed__3); l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5___closed__4 = _init_l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5___closed__4(); lean_mark_persistent(l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5___closed__4); -l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3___closed__1 = _init_l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3___closed__1(); -lean_mark_persistent(l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3___closed__1); -l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3___closed__2 = _init_l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3___closed__2(); -lean_mark_persistent(l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3___closed__2); -l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3___closed__3 = _init_l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3___closed__3(); -lean_mark_persistent(l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3___closed__3); l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__1 = _init_l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__1(); l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2 = _init_l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2(); lean_mark_persistent(l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2); @@ -14657,48 +19416,46 @@ l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_ lean_mark_persistent(l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__8); l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__9 = _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__9(); lean_mark_persistent(l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__9); +l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__10 = _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__10(); +lean_mark_persistent(l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__10); +l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__11 = _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__11(); +lean_mark_persistent(l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__11); l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___closed__1 = _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___closed__1(); lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___closed__1); l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___closed__2 = _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___closed__2(); lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___closed__2); -l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__1(); -l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__2(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__2); -l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__3(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__3); -l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__4(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__4); -l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__5 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__5(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__5); -l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__5___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__5___closed__1(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__5___closed__1); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__1(); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__2); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__3); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__4(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__4); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__5 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__5(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__5); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__4___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__4___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__4___closed__1); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__1); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__2); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__3); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__4(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__4); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__5 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__5(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__5); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__6 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__6(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__6); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__7 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__7(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__7); l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__1(); lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__1); l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__2(); lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__2); -l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__3(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__3); -l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__4(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__4); -l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__5 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__5(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__5); -l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__6 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__6(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__6); -l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__7 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__7(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__7); -l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__8 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__8(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__8); -l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__8___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__8___closed__1(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__8___closed__1); -l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__8___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__8___closed__2(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__8___closed__2); l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___closed__1(); lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___closed__1); l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___closed__2(); -l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___lambda__2___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___lambda__2___closed__1(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___lambda__2___closed__1); -l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___closed__1(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___closed__1); l_Lean_Meta_Grind_EMatch_ematchTheorem___lambda__1___closed__1 = _init_l_Lean_Meta_Grind_EMatch_ematchTheorem___lambda__1___closed__1(); lean_mark_persistent(l_Lean_Meta_Grind_EMatch_ematchTheorem___lambda__1___closed__1); l_Lean_Meta_Grind_EMatch_ematchTheorem___lambda__1___closed__2 = _init_l_Lean_Meta_Grind_EMatch_ematchTheorem___lambda__1___closed__2(); @@ -14709,12 +19466,12 @@ l_Lean_Meta_Grind_ematch___closed__2 = _init_l_Lean_Meta_Grind_ematch___closed__ lean_mark_persistent(l_Lean_Meta_Grind_ematch___closed__2); l_Lean_Meta_Grind_ematch___closed__3 = _init_l_Lean_Meta_Grind_ematch___closed__3(); lean_mark_persistent(l_Lean_Meta_Grind_ematch___closed__3); -l_Lean_Meta_Grind_ematchAndAssert_x3f___lambda__4___closed__1 = _init_l_Lean_Meta_Grind_ematchAndAssert_x3f___lambda__4___closed__1(); -lean_mark_persistent(l_Lean_Meta_Grind_ematchAndAssert_x3f___lambda__4___closed__1); -l_Lean_Meta_Grind_ematchAndAssert_x3f___closed__1 = _init_l_Lean_Meta_Grind_ematchAndAssert_x3f___closed__1(); -lean_mark_persistent(l_Lean_Meta_Grind_ematchAndAssert_x3f___closed__1); -l_Lean_Meta_Grind_ematchAndAssert_x3f___closed__2 = _init_l_Lean_Meta_Grind_ematchAndAssert_x3f___closed__2(); -lean_mark_persistent(l_Lean_Meta_Grind_ematchAndAssert_x3f___closed__2); +l_Lean_Meta_Grind_ematchAndAssert___lambda__4___closed__1 = _init_l_Lean_Meta_Grind_ematchAndAssert___lambda__4___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_ematchAndAssert___lambda__4___closed__1); +l_Lean_Meta_Grind_ematchAndAssert___closed__1 = _init_l_Lean_Meta_Grind_ematchAndAssert___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_ematchAndAssert___closed__1); +l_Lean_Meta_Grind_ematchAndAssert___closed__2 = _init_l_Lean_Meta_Grind_ematchAndAssert___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_ematchAndAssert___closed__2); l_Lean_Meta_Grind_ematchStar___closed__1 = _init_l_Lean_Meta_Grind_ematchStar___closed__1(); lean_mark_persistent(l_Lean_Meta_Grind_ematchStar___closed__1); return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Lean/Meta/Tactic/Grind/EMatchTheorem.c b/stage0/stdlib/Lean/Meta/Tactic/Grind/EMatchTheorem.c index 3186aa6411c1..f05bbee30afe 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Grind/EMatchTheorem.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Grind/EMatchTheorem.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Lean.Meta.Tactic.Grind.EMatchTheorem -// Imports: Lean.HeadIndex Lean.PrettyPrinter Lean.Util.FoldConsts Lean.Util.CollectFVars Lean.Meta.Basic Lean.Meta.InferType +// Imports: Init.Grind.Util Init.Grind.Tactics Lean.HeadIndex Lean.PrettyPrinter Lean.Util.FoldConsts Lean.Util.CollectFVars Lean.Meta.Basic Lean.Meta.InferType Lean.Meta.Eqns Lean.Meta.Tactic.Grind.Util #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -13,185 +13,278 @@ #ifdef __cplusplus extern "C" { #endif -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__2___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addEMatchTheorem___lambda__1___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchEqTheoremCore(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_Meta_Grind_mkEMatchTheoremWithKind_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__2___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ematchTheoremsExt; +static lean_object* l_Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2___closed__1; lean_object* l___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2170_(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__21; lean_object* l_Lean_Name_reprPrec(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__10(lean_object*, size_t, lean_object*); static lean_object* l_Lean_Meta_Grind_ppPattern___closed__4; -static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__15; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collectPatterns_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getProofFor___closed__1; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___closed__1; +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__12; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare; -static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__10; +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__15; +static lean_object* l_Lean_Meta_Grind_preprocessPattern___lambda__1___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__6(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instHashableOrigin___boxed(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___boxed(lean_object*, 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_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___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_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__5(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__3; lean_object* l_Lean_getConstInfo___at_Lean_Meta_mkConstWithFreshMVarLevels___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__6(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__3; static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___lambda__1___closed__1; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute___closed__4; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__4; static lean_object* l_Lean_Meta_Grind_addEMatchTheorem___closed__1; static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__6; lean_object* lean_mk_empty_array_with_capacity(lean_object*); lean_object* l_Lean_mkAppN(lean_object*, lean_object*); size_t lean_usize_shift_right(size_t, size_t); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addEMatchTheorem___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_Meta_Grind_mkGroundPattern___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Origin_pp___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__3___boxed(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_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__2; +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__8; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___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_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__4; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___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_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_detectOffsets(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__1; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__4___closed__1; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_detectOffsets___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_EMatchTheorems_contains___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_Meta_Grind_NormalizePattern_main___spec__1(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_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___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_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__8___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -uint8_t l_List_all___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__15; -static lean_object* l_Lean_Meta_Grind_Origin_pp___rarg___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_EMatchTheorem_getProofWithFreshMVarLevels___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__1___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_NormalizePattern_getPatternSupportMask(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___closed__2; +static lean_object* l_Lean_Meta_Grind_isOffsetPattern_x3f___closed__1; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___spec__1___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___spec__2(lean_object*, lean_object*, lean_object*, size_t, size_t); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchTheoremWithKind_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*); static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__19; uint8_t lean_usize_dec_le(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addEMatchTheorem(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_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__7(size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Array_contains___at_Lean_Meta_arrowDomainsN___spec__2(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__5; LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__4(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_ConstantInfo_levelParams(lean_object*); lean_object* l_Lean_indentD(lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__1; LEAN_EXPORT uint8_t l_Lean_Meta_Grind_isPatternDontCare(lean_object*); +uint8_t l_Lean_Exception_isInterrupt(lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__3; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__6___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__2; uint64_t lean_uint64_of_nat(lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__6(lean_object*, size_t, size_t, lean_object*, lean_object*); +lean_object* l_Lean_mkAppB(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_isOffset_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___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_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collectPatterns_x3f___closed__1; +lean_object* l_Lean_Meta_Grind_unfoldReducible___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_mkOffsetPattern___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instInhabitedOrigin; size_t lean_uint64_to_usize(uint64_t); -static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__11; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___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_Array_mapMUnsafe_map___at_Lean_Meta_Grind_NormalizePattern_getPatternSupportMask___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___spec__1___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__3; +static lean_object* l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___closed__4; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__1(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_detectOffsets___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__18___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__19(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isApp(lean_object*); -LEAN_EXPORT lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__2(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__13; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr___closed__2; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__3___closed__1; lean_object* l_Lean_Expr_sort___override(lean_object*); lean_object* l_Lean_MessageData_ofList(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_PersistentArray_push___rarg(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3___closed__2; static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___closed__2; lean_object* lean_array_push(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_addEMatchTheorem___closed__5; lean_object* l_Array_toSubarray___rarg(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_addEMatchTheorem___closed__6; +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__19; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect(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_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__5; size_t lean_usize_mul(size_t, size_t); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toCtorIdx___boxed(lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr___closed__3; lean_object* l_Lean_RBNode_fold___at_Lean_RBMap_size___spec__1___rarg(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute___closed__3; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchTheoremWithKind_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*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBTree_ofList___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__4(lean_object*); -static lean_object* l_Lean_Meta_Grind_Origin_pp___rarg___closed__3; lean_object* lean_mk_array(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__5; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchEqTheoremCore___lambda__1(uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_Grind_foldProjs___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__17; uint8_t lean_usize_dec_eq(size_t, size_t); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___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_PersistentHashMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____spec__1___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51_(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__10; static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__1___closed__1; LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand_go___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__3(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279_(lean_object*); LEAN_EXPORT lean_object* l_List_mapTR_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__8(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr(lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instBEqOrigin__1___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__12(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_ppPattern___closed__8; -static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__16; static lean_object* l_Lean_Meta_Grind_mkGroundPattern___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_detectOffsets___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_instInhabitedOrigin___closed__1; lean_object* l_Lean_Expr_bvar___override(lean_object*); LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___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_array_fget(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__2(lean_object*, 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_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___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_Grind_addEMatchTheorem___closed__11; -static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__17; -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_mkEMatchEqTheoremCore___lambda__3___closed__1; LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); -static size_t l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__1; +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__17; static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__1___closed__2; -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__2(lean_object*, lean_object*); lean_object* l_Lean_Expr_fvarId_x21(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__10(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); -static size_t l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__2; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__7; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_noConfusion___rarg(uint8_t, uint8_t, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_Grind_NormalizePattern_getPatternSupportMask___spec__1(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3(lean_object*, size_t, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__4; +LEAN_EXPORT lean_object* l_Array_filterMapM___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getPropTypes___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__9(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); lean_object* l_Lean_Meta_forallTelescopeReducing___at_Lean_Meta_getParamNames___spec__2___rarg(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_EMatchTheorems_contains___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instInhabitedEMatchTheorems; LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isAppOf(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getProofFor(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Origin_pp(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666_(lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__6; +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collectPatterns_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 uint8_t l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_EMatchTheorems_contains___spec__1(lean_object*, lean_object*); lean_object* l_Nat_nextPowerOfTwo_go(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchTheoremWithKind_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* l_Lean_Expr_cleanupAnnotations(lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__7; lean_object* l_Lean_stringToMessageData(lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr___spec__2(uint8_t, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_NormalizePattern_main___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instReprOrigin; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__9; LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___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_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___lambda__1___closed__2; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getPropTypes___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_instReprOrigin___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___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_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__3___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__4(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_detectOffsets___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_instInhabitedEMatchTheorem___closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instInhabitedEMatchTheorem; 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_Grind_addEMatchTheorem___closed__8; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchTheoremCore___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_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__1___closed__1; LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___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_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__2___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_addEMatchTheorem___closed__9; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_noConfusion___rarg___lambda__1(lean_object*); +LEAN_EXPORT uint8_t l_Lean_Meta_Grind_EMatchTheorems_isErased(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_NormalizePattern_getPatternSupportMask___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getPropTypes(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__7; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_ppPattern___closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isPatternDontCare___boxed(lean_object*); +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__5; uint8_t l_Lean_Expr_isBVar(lean_object*); lean_object* l_List_mapTR_loop___at_Lean_mkConstWithLevelParams___spec__1(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_addEMatchTheorem___closed__2; -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__15(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__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_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__4; +static lean_object* l_Lean_Meta_Grind_mkOffsetPattern___closed__5; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr___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_Grind_mkGroundPattern___closed__2; LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__2(lean_object*); uint8_t l_Lean_Expr_hasMVar(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_EMatchTheorem_getProofWithFreshMVarLevels(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Origin_key___boxed(lean_object*); -LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__3(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__7; static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare___closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_EMatchTheorems_insert(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__1; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__2; lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__5(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__14___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Origin_pp___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__20(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__6; +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__20; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__6___closed__3; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__10(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_Meta_Grind_NormalizePattern_main___closed__3; static lean_object* l_Lean_Meta_Grind_NormalizePattern_main___closed__4; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__2(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3___closed__3; +static lean_object* l_Lean_Meta_Grind_mkOffsetPattern___closed__2; size_t lean_usize_of_nat(lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__6; static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3___closed__1; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__11(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_PersistentHashMap_isUnaryNode___rarg(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_preprocessPattern___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___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_expr_abstract(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_beqOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_811_(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___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_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__14___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__15___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_mapTR_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__6___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__7(lean_object*, lean_object*, size_t, size_t, lean_object*); -LEAN_EXPORT lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___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_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__14; static lean_object* l_Lean_Meta_Grind_ppPattern___closed__7; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_preprocessPattern(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__11; +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__3; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__2; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Core_transform___at_Lean_Meta_Grind_unfoldReducible___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_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_foundBVar___boxed(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_Grind_EMatchTheorems_insert___spec__4(size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_ConstantInfo_isTheorem(lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldrM___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__2___boxed(lean_object*, lean_object*); lean_object* lean_st_ref_take(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__2; -static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__13; +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_Grind_mkEMatchEqTheoremCore___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, 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*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getProofFor___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_EMatchTheorems_retrieve_x3f___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f_go___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__18; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__1; uint8_t lean_expr_eqv(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_registerSimpleScopedEnvExtension___rarg(lean_object*, lean_object*); @@ -199,174 +292,311 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lea lean_object* l_Lean_Meta_forallBoundedTelescope___at_Lean_Meta_arrowDomainsN___spec__6___rarg(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___closed__3; uint64_t lean_uint64_shift_right(uint64_t, uint64_t); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_nat_to_int(lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand_go___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__2(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofSyntax(lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___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*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__8; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__7; lean_object* lean_nat_div(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_NormalizePattern_main___closed__5; -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addEMatchTheorem___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_Grind_addEMatchTheorem___lambda__3___closed__4; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go_goArg(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__5; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFunMask(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__10; +static lean_object* l_Lean_Meta_Grind_preprocessPattern___closed__2; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_groundPattern_x3f(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_beqOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_811____boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isPatternFnCandidate(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_RBNode_findCore___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__2(lean_object*, lean_object*); -static lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__2___closed__2; -LEAN_EXPORT uint8_t l_Lean_Meta_Grind_addEMatchTheorem___lambda__1(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchTheorem(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_preprocessPattern___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__2(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Syntax_getKind(lean_object*); lean_object* l_Lean_MessageData_ofFormat(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFunMask___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__5; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__1; +static lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__1___closed__2; +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__10; +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___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_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__7; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__9(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_Expr_appArg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isOffsetPattern_x3f___lambda__2___boxed(lean_object*); lean_object* l_Lean_FVarId_getDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_outOfBounds___rarg(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___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*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__7___closed__3; static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___closed__1; +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__1; LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchEqTheoremCore___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___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_Grind_getEMatchTheorems(lean_object*); +lean_object* l_Lean_Meta_getEqnsFor_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__8; LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___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_st_ref_get(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__18; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_noConfusion___rarg___lambda__1___boxed(lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___closed__4; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__8; static lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__2___closed__1; static lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__3___closed__2; +static lean_object* l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedEMatchTheorems___spec__1___closed__1; +lean_object* l_Lean_RBNode_findCore___at_Lean_Meta_Closure_process___spec__1(lean_object*, lean_object*); uint8_t l_List_isEmpty___rarg(lean_object*); lean_object* lean_st_mk_ref(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__16; +static lean_object* l_Lean_Meta_Grind_preprocessPattern___lambda__1___closed__2; LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isForbidden(lean_object*); LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___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_array_to_list(lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__6; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__4(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Expr_appFnCleanup(lean_object*, lean_object*); +lean_object* l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_hasLooseBVars(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchEqTheoremCore___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__9; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_beqTheoremKind____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_6540____boxed(lean_object*, lean_object*); +lean_object* l_Lean_ScopedEnvExtension_modifyState___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___closed__3; uint8_t l_List_elem___at_Lean_Meta_Occurrences_contains___spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instBEqOrigin; static lean_object* l_Lean_Meta_Grind_ppPattern___closed__2; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute___closed__6; lean_object* l_Lean_Expr_toHeadIndex(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___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_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__15; +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__13; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__9___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__6___closed__2; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isOffsetPattern_x3f(lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_EMatchTheorems_retrieve_x3f___spec__2___boxed(lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_levelZero; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchTheoremWithKind_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*); static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__3; +lean_object* l_panic___at_String_fromUTF8_x21___spec__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_NormalizePattern_main(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_ppPattern___closed__11; +static lean_object* l_Lean_Meta_Grind_mkOffsetPattern___closed__4; lean_object* l_Lean_mkConstWithLevelParams___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__9; lean_object* l_Lean_Expr_constName_x21(lean_object*); +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__6; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___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; -static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFunMask___closed__1; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__6___closed__1; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__12; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFn_x3f(lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__5; -static lean_object* l_Lean_Meta_Grind_addEMatchTheorem___lambda__2___closed__1; +static lean_object* l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__5; LEAN_EXPORT uint8_t l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute___closed__1; lean_object* l_Lean_MessageData_ofConst(lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____spec__1(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getEMatchTheorems___boxed(lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__9; lean_object* l_Lean_FVarId_getType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_grind_normalize(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__19; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___closed__2; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__6; lean_object* l_Lean_mkAnnotation(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__9; static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__13; static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_foundBVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_preprocessPattern___closed__3; LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__1(lean_object*); -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_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_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_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_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__6; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____boxed(lean_object*, lean_object*); uint8_t lean_name_eq(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__2___boxed(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_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__4(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); lean_object* l_Lean_Expr_instantiateLevelParamsArray(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___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_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr___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*); +uint8_t l_Lean_Expr_isArrow(lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___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_EXPORT lean_object* l_Lean_ScopedEnvExtension_add___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr___spec__1(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__12___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__13(lean_object*, lean_object*, size_t, size_t, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__7; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute___closed__5; +LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___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_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__4; +lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__2; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___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___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_Grind_foldProjs___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_NormalizePattern_getPatternSupportMask___closed__1; static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__2; -static lean_object* l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__6; +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__14; +static lean_object* l_Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2___closed__6; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instBEqTheoremKind; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames; -static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__3; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__5; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure(uint8_t); +static lean_object* l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__3; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchEqTheorem(lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__9; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__7___closed__1; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__5; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchEqTheoremCore___lambda__2(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_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__1___closed__1; lean_object* l_Lean_CollectFVars_main(lean_object*, lean_object*); extern lean_object* l_Lean_Meta_instMonadMetaM; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Origin_pp___rarg___lambda__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_indexOfAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__8___boxed(lean_object*, lean_object*, lean_object*); +lean_object* l_Array_eraseIdx___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__10; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___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___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFunMask___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_instInhabitedEMatchTheorem___closed__1; -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_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*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_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*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___closed__1; -static lean_object* l_Lean_Meta_Grind_Origin_key___closed__2; +static lean_object* l_Lean_Meta_Grind_mkEMatchEqTheoremCore___lambda__3___closed__2; static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__7; LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___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_usize_to_nat(size_t); lean_object* l_Lean_Meta_instInhabitedMetaM___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__2; LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___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*); -static lean_object* l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____spec__1___closed__2; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); +lean_object* l_Array_append___rarg(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__14; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___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_Meta_Grind_EMatchTheorems_insert___closed__2; static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___closed__1; lean_object* l_Lean_MessageData_ofExpr(lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____spec__1; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__3(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__4; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr___closed__1; +static lean_object* l_Lean_Meta_Grind_mkOffsetPattern___closed__3; +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__4; +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_Grind_mkEMatchEqTheoremCore___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_beqTheoremKind____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_6540_(uint8_t, uint8_t); +double l_Float_ofScientific(lean_object*, uint8_t, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___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_Meta_Grind_isOffsetPattern_x3f___lambda__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__12(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__4; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f_go___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__14___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__15(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__2(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Origin_pp___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_EMatchTheorems_isErased___boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__12; +LEAN_EXPORT lean_object* l_Array_indexOfAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__8(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collectPatterns_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__11; +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getPropTypes___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___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isForbidden___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addEMatchTheorem___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_panic___at_Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchEqTheoremCore___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__20; +LEAN_EXPORT uint64_t l_Lean_Meta_Grind_instHashableOrigin(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go_goArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__6___closed__2; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__4; +LEAN_EXPORT lean_object* l_Array_filterMapM___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getPropTypes___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___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_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__8; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkGroundPattern(lean_object*); static lean_object* l_Lean_Meta_Grind_ppPattern___closed__9; LEAN_EXPORT lean_object* l_List_mapTR_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__3(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldrM___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__2(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__1; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__4; static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__5; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr(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_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3___closed__4; -static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__14; +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isAtomicPattern(lean_object*); -LEAN_EXPORT lean_object* l_panic___at_Lean_Meta_Grind_addEMatchTheorem___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___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_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__7(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__16; -lean_object* l_Lean_PersistentHashMap_instInhabited(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__1(lean_object*); +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__18; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__3___closed__2; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__10___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__11(lean_object*, lean_object*, size_t, size_t, lean_object*); -static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__20; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_NormalizePattern_normalizePattern(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_EMatchTheorems_contains___spec__2(lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___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_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__8; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___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_Grind_mkEMatchEqTheoremCore___lambda__4(uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr___lambda__1(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__11; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___boxed(lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__4; +static size_t l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__1; +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__3; +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__6; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collectPatterns_x3f___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_Grind_isOffsetPattern_x3f___closed__2; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f_go(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__14; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); -static lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__2___closed__1; +LEAN_EXPORT uint8_t l_Lean_Meta_Grind_EMatchTheorems_contains(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__10; +lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute___closed__2; +lean_object* l_Lean_mkRawNatLit(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toCtorIdx(uint8_t); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collectPatterns_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_id___rarg___boxed(lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__3; +static lean_object* l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___closed__2; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__13___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_noConfusion(lean_object*); +lean_object* l_Lean_indentExpr(lean_object*); +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__6___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchEqTheorem___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_erase___at_Lean_Meta_Grind_EMatchTheorems_retrieve_x3f___spec__1(lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_mkEmptyEntries(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__6(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_erase___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__6(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__20___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ppPattern(lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___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_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_EMatchTheorems_insert___closed__1; LEAN_EXPORT uint8_t l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_foundBVar___spec__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__5(lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Expr_isConstOf(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___lambda__2___closed__2; lean_object* lean_array_set(lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3; uint64_t l_Lean_Name_hash___override(lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__3; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_EMatchTheorems_retrieve_x3f___spec__2(lean_object*, size_t, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__14; 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_Grind_mkEMatchTheoremCore(lean_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_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Array_contains___at_Lean_registerInternalExceptionId___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__3(lean_object*, lean_object*, lean_object*); @@ -374,1148 +604,1039 @@ uint64_t lean_uint64_xor(uint64_t, uint64_t); extern lean_object* l_Id_instMonad; lean_object* l_Repr_addAppParen(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Syntax_isNone(lean_object*); lean_object* lean_panic_fn(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___lambda__2___closed__1; lean_object* l_Lean_annotation_x3f(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__6; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__2; +LEAN_EXPORT uint8_t l_Lean_Meta_Grind_instInhabitedTheoremKind; +static uint64_t l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__2; lean_object* lean_nat_sub(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___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_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_ScopedEnvExtension_getState___rarg(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_getAppFn(lean_object*); -static lean_object* l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchEqTheoremCore___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_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__18; static lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__3___closed__1; lean_object* lean_nat_mul(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getPropTypes___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_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__11(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); lean_object* l_Lean_Meta_isTypeFormer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_panic___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__1___closed__2; +static lean_object* l_Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2___closed__2; +lean_object* l_Array_indexOfAux___at_Lean_MetavarContext_setMVarUserName___spec__3(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_ScopedEnvExtension_addCore___rarg(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__2; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__8; LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__4; +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__11; lean_object* l_Lean_PersistentHashMap_mkCollisionNode___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___closed__6; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__1(lean_object*, lean_object*); +lean_object* l_Lean_Meta_instantiateMVarsIfMVarApp(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_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___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_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__7(lean_object*, size_t, lean_object*); uint8_t l_Lean_LocalDecl_binderInfo(lean_object*); lean_object* l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__12; static lean_object* l_Lean_Meta_Grind_NormalizePattern_main___closed__1; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_EMatchTheorems_contains___spec__2___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_EMatchTheorems_insert___closed__3; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_noConfusion___rarg___closed__1; +lean_object* l_Lean_registerBuiltinAttribute(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__4; static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__17; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__8; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__5___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*); -static lean_object* l_Lean_Meta_Grind_Origin_key___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isPatternFnCandidate___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_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__19; -static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__11; +static lean_object* l_Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2___closed__3; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__3; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__7; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3(lean_object*, size_t, size_t, lean_object*, lean_object*); size_t lean_usize_sub(size_t, size_t); lean_object* lean_array_mk(lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getProofFor___closed__2; +uint8_t l_Lean_Syntax_structEq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___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_throwError___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getProofFor___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_foundBVar___spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___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_Grind_ppPattern___closed__1; -static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__10; -static lean_object* l_Lean_Meta_Grind_addEMatchTheorem___closed__10; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__7___closed__4; +static lean_object* l_Lean_Meta_Grind_instBEqOrigin___closed__1; +LEAN_EXPORT lean_object* l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___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_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__4(lean_object*, lean_object*); +lean_object* l_Lean_isTracingEnabledForCore(lean_object*, lean_object*, lean_object*); size_t lean_usize_add(size_t, size_t); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___closed__3; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedEMatchTheorems___spec__1; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__7; LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__2; +static double l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___spec__2___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f_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_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__6; +LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___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___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___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_array_uget(lean_object*, size_t); lean_object* l_Lean_Expr_fvar___override(lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___closed__1; size_t lean_array_size(lean_object*); -static lean_object* l_Lean_Meta_Grind_addEMatchTheorem___closed__12; +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__5; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__10___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_panic___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__1___closed__1; -static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__13; lean_object* l_instInhabitedOfMonad___rarg(lean_object*, lean_object*); -static lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__6___closed__1; -static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__3; lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_EMatchTheorems_retrieve_x3f(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_PersistentHashMap_erase___at_Lean_Meta_Grind_EMatchTheorems_retrieve_x3f___spec__1___boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__3; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_erase___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__6___boxed(lean_object*, lean_object*); size_t lean_usize_shift_left(size_t, size_t); +lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___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 uint8_t l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isGroundPattern(lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__18___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__19___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__7___closed__2; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___boxed(lean_object*); uint8_t l___private_Lean_HeadIndex_0__Lean_beqHeadIndex____x40_Lean_HeadIndex___hyg_69_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFn_x3f___boxed(lean_object*); lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__9; +static lean_object* l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__15; +static lean_object* l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_NormalizePattern_getPatternSupportMask___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_preprocessPattern___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__3(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_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_instBEqTheoremKind___closed__1; lean_object* lean_string_append(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_ppPattern___closed__10; static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__2; lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____boxed(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Lean_Meta_Grind_instBEqOrigin__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_ScopedEnvExtension_add___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__2(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__17; -static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__12; -extern lean_object* l_Lean_Name_instBEq; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_detectOffsets___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_Grind_EMatchTheorems_erase(lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); -static lean_object* l_Lean_Meta_Grind_addEMatchTheorem___closed__3; -static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__16; +LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getProofFor___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_EMatchTheorems_contains___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__7___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__3; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute(uint8_t); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkOffsetPattern(lean_object*, lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1; +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__13; static lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Origin_key(lean_object*); lean_object* lean_infer_type(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isGroundPattern___boxed(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_Meta_Grind_isOffsetPattern_x3f___closed__3; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collectPatterns_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 lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__12___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__13___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFunMask___spec__1(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedEMatchTheorems___spec__1___closed__2; LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__8___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___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*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__10; static lean_object* l_Lean_Meta_Grind_ppPattern___closed__5; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_detectOffsets___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_EMatchTheorems_contains___spec__1___boxed(lean_object*, lean_object*); lean_object* lean_nat_add(lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(lean_object*); lean_object* l_Lean_MessageData_bracket(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___lambda__1___closed__1; uint8_t l_Lean_Expr_isConst(lean_object*); LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___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*); -static lean_object* l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__5; +uint8_t l_Lean_Exception_isRuntime(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchEqTheoremCore___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__2; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__17(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addEMatchEqTheorem(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint64_t l_Lean_HeadIndex_hash(lean_object*); uint8_t l_Lean_Expr_isFVar(lean_object*); static lean_object* l_Lean_Meta_Grind_EMatchTheorems_insert___closed__4; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getEMatchTheorems___rarg___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFunMask___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_ppPattern___closed__6; lean_object* l_Lean_Meta_mkConstWithFreshMVarLevels(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_Origin_pp___rarg___closed__1; +static lean_object* l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___closed__5; +static size_t l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__2; +uint64_t l___private_Lean_Meta_Basic_0__Lean_Meta_Config_toKey(lean_object*); +lean_object* l_Lean_Meta_Grind_unfoldReducible___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__8(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_groundPattern_x3f___boxed(lean_object*); lean_object* lean_array_uset(lean_object*, size_t, lean_object*); +LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____spec__2(lean_object*, lean_object*, size_t, size_t); lean_object* l_Lean_MessageData_ofName(lean_object*); +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__16; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__14(size_t, 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___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__12; -static lean_object* l_Lean_Meta_Grind_addEMatchTheorem___closed__4; lean_object* l___private_Init_Data_Repr_0__Nat_reprFast(lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare___closed__1; -static lean_object* l_Lean_Meta_Grind_addEMatchTheorem___closed__7; +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__8; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__13(lean_object*, size_t, size_t, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_instInhabitedEMatchTheorem___closed__2; -static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__7; +static lean_object* l_Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2___closed__5; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_detectOffsets___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__9; lean_object* l_Lean_Meta_isProof(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___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_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614_(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchTheoremCore___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* l_ReaderT_instMonad___rarg(lean_object*); size_t lean_usize_land(size_t, size_t); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isOffsetPattern_x3f___lambda__2(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__1; -extern lean_object* l_Lean_instHashableName; +static lean_object* l_Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2___closed__4; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getEMatchTheorems___rarg(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___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*); uint8_t l_Array_isEmpty___rarg(lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__10___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___closed__3; lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_openAbstractMVarsResult___spec__1(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* _init_l_Lean_Meta_Grind_instInhabitedOrigin___closed__1() { +static lean_object* _init_l_Lean_Meta_Grind_mkOffsetPattern___closed__1() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = lean_box(0); -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_unchecked("Lean", 4, 4); +return x_1; } } -static lean_object* _init_l_Lean_Meta_Grind_instInhabitedOrigin() { +static lean_object* _init_l_Lean_Meta_Grind_mkOffsetPattern___closed__2() { _start: { lean_object* x_1; -x_1 = l_Lean_Meta_Grind_instInhabitedOrigin___closed__1; +x_1 = lean_mk_string_unchecked("Grind", 5, 5); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__1() { +static lean_object* _init_l_Lean_Meta_Grind_mkOffsetPattern___closed__3() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean.Meta.Grind.Origin.decl", 27, 27); +x_1 = lean_mk_string_unchecked("offset", 6, 6); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__2() { +static lean_object* _init_l_Lean_Meta_Grind_mkOffsetPattern___closed__4() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__1; -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; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Meta_Grind_mkOffsetPattern___closed__1; +x_2 = l_Lean_Meta_Grind_mkOffsetPattern___closed__2; +x_3 = l_Lean_Meta_Grind_mkOffsetPattern___closed__3; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__3() { +static lean_object* _init_l_Lean_Meta_Grind_mkOffsetPattern___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__2; -x_2 = lean_box(1); -x_3 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set(x_3, 1, x_2); +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind_mkOffsetPattern___closed__4; +x_3 = l_Lean_Expr_const___override(x_2, x_1); return x_3; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__4() { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkOffsetPattern(lean_object* x_1, lean_object* x_2) { _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; +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Lean_mkRawNatLit(x_2); +x_4 = l_Lean_Meta_Grind_mkOffsetPattern___closed__5; +x_5 = l_Lean_mkAppB(x_4, x_1, x_3); +return x_5; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__5() { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_detectOffsets___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = lean_unsigned_to_nat(1u); -x_2 = lean_nat_to_int(x_1); -return x_2; +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_9 = l_Lean_Meta_Grind_mkOffsetPattern(x_1, x_2); +x_10 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_10, 0, x_9); +x_11 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_11, 0, 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_8); +return x_12; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__6() { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_detectOffsets___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean.Meta.Grind.Origin.fvar", 27, 27); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__7() { -_start: +uint8_t x_8; +x_8 = lean_expr_eqv(x_2, x_1); +if (x_8 == 0) { -lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__6; -x_2 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; +switch (lean_obj_tag(x_2)) { +case 6: +{ +lean_object* x_9; lean_object* x_10; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_9 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_9, 0, x_2); +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; } +case 7: +{ +lean_object* x_11; lean_object* x_12; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_11 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_11, 0, x_2); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_7); +return x_12; } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__8() { -_start: +case 8: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__7; -x_2 = lean_box(1); -x_3 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set(x_3, 1, x_2); -return x_3; +lean_object* x_13; lean_object* x_14; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_13 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_13, 0, x_2); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_7); +return x_14; } +default: +{ +lean_object* x_15; +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +x_15 = l_Lean_Meta_isOffset_x3f(x_2, x_3, x_4, x_5, x_6, x_7); +if (lean_obj_tag(x_15) == 0) +{ +lean_object* x_16; +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +if (lean_obj_tag(x_16) == 0) +{ +uint8_t x_17; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_17 = !lean_is_exclusive(x_15); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_18 = lean_ctor_get(x_15, 0); +lean_dec(x_18); +x_19 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_19, 0, x_2); +x_20 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_20, 0, x_19); +lean_ctor_set(x_15, 0, x_20); +return x_15; } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__9() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean.Meta.Grind.Origin.stx", 26, 26); -return x_1; +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_21 = lean_ctor_get(x_15, 1); +lean_inc(x_21); +lean_dec(x_15); +x_22 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_22, 0, x_2); +x_23 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_23, 0, x_22); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set(x_24, 1, x_21); +return x_24; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__10() { -_start: +else { -lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__9; -x_2 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; -} +uint8_t x_25; +lean_dec(x_2); +x_25 = !lean_is_exclusive(x_16); +if (x_25 == 0) +{ +uint8_t x_26; +x_26 = !lean_is_exclusive(x_15); +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; uint8_t x_33; +x_27 = lean_ctor_get(x_16, 0); +x_28 = lean_ctor_get(x_15, 1); +x_29 = lean_ctor_get(x_15, 0); +lean_dec(x_29); +x_30 = lean_ctor_get(x_27, 0); +lean_inc(x_30); +x_31 = lean_ctor_get(x_27, 1); +lean_inc(x_31); +lean_dec(x_27); +x_32 = lean_unsigned_to_nat(0u); +x_33 = lean_nat_dec_eq(x_31, x_32); +if (x_33 == 0) +{ +lean_object* x_34; lean_object* x_35; +lean_free_object(x_15); +lean_free_object(x_16); +x_34 = lean_box(0); +x_35 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_detectOffsets___lambda__1(x_30, x_31, x_34, x_3, x_4, x_5, x_6, x_28); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_35; } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__11() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__10; -x_2 = lean_box(1); -x_3 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set(x_3, 1, x_2); -return x_3; +lean_object* x_36; +lean_dec(x_31); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_ctor_set(x_16, 0, x_30); +x_36 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_36, 0, x_16); +lean_ctor_set(x_15, 0, x_36); +return x_15; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__12() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean.Meta.Grind.Origin.other", 28, 28); -return x_1; -} +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; uint8_t x_42; +x_37 = lean_ctor_get(x_16, 0); +x_38 = lean_ctor_get(x_15, 1); +lean_inc(x_38); +lean_dec(x_15); +x_39 = lean_ctor_get(x_37, 0); +lean_inc(x_39); +x_40 = lean_ctor_get(x_37, 1); +lean_inc(x_40); +lean_dec(x_37); +x_41 = lean_unsigned_to_nat(0u); +x_42 = lean_nat_dec_eq(x_40, x_41); +if (x_42 == 0) +{ +lean_object* x_43; lean_object* x_44; +lean_free_object(x_16); +x_43 = lean_box(0); +x_44 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_detectOffsets___lambda__1(x_39, x_40, x_43, x_3, x_4, x_5, x_6, x_38); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_44; } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__13() { -_start: +else { -lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__12; -x_2 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; +lean_object* x_45; lean_object* x_46; +lean_dec(x_40); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_ctor_set(x_16, 0, x_39); +x_45 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_45, 0, x_16); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_38); +return x_46; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__14() { -_start: +} +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__4; -x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__13; -x_3 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set(x_3, 1, x_2); -return x_3; +lean_object* x_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; +x_47 = lean_ctor_get(x_16, 0); +lean_inc(x_47); +lean_dec(x_16); +x_48 = lean_ctor_get(x_15, 1); +lean_inc(x_48); +if (lean_is_exclusive(x_15)) { + lean_ctor_release(x_15, 0); + lean_ctor_release(x_15, 1); + x_49 = x_15; +} else { + lean_dec_ref(x_15); + x_49 = lean_box(0); } +x_50 = lean_ctor_get(x_47, 0); +lean_inc(x_50); +x_51 = lean_ctor_get(x_47, 1); +lean_inc(x_51); +lean_dec(x_47); +x_52 = lean_unsigned_to_nat(0u); +x_53 = lean_nat_dec_eq(x_51, x_52); +if (x_53 == 0) +{ +lean_object* x_54; lean_object* x_55; +lean_dec(x_49); +x_54 = lean_box(0); +x_55 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_detectOffsets___lambda__1(x_50, x_51, x_54, x_3, x_4, x_5, x_6, x_48); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_55; } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__15() { -_start: +else { -lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__14; -x_2 = 0; -x_3 = lean_alloc_ctor(6, 1, 1); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); -return x_3; +lean_object* x_56; lean_object* x_57; lean_object* x_58; +lean_dec(x_51); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_56 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_56, 0, x_50); +x_57 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_57, 0, x_56); +if (lean_is_scalar(x_49)) { + x_58 = lean_alloc_ctor(0, 2, 0); +} else { + x_58 = x_49; } +lean_ctor_set(x_58, 0, x_57); +lean_ctor_set(x_58, 1, x_48); +return x_58; } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__16() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__5; -x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__13; -x_3 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set(x_3, 1, x_2); -return x_3; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__17() { -_start: -{ -lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__16; -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_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51_(lean_object* x_1, lean_object* x_2) { -_start: -{ -switch (lean_obj_tag(x_1)) { -case 0: -{ -lean_object* x_3; lean_object* x_4; uint8_t x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; -x_3 = lean_ctor_get(x_1, 0); -lean_inc(x_3); -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_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____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_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____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); -x_11 = 0; -x_12 = lean_alloc_ctor(6, 1, 1); -lean_ctor_set(x_12, 0, x_10); -lean_ctor_set_uint8(x_12, sizeof(void*)*1, x_11); -x_13 = l_Repr_addAppParen(x_12, x_2); -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_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____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); -x_16 = 0; -x_17 = lean_alloc_ctor(6, 1, 1); -lean_ctor_set(x_17, 0, x_15); -lean_ctor_set_uint8(x_17, sizeof(void*)*1, x_16); -x_18 = l_Repr_addAppParen(x_17, x_2); -return x_18; -} -} -case 1: -{ -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_1, 0); -lean_inc(x_19); -lean_dec(x_1); -x_20 = lean_unsigned_to_nat(1024u); -x_21 = lean_nat_dec_le(x_20, x_2); -x_22 = l_Lean_Name_reprPrec(x_19, x_20); -x_23 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__8; -x_24 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_24, 0, x_23); -lean_ctor_set(x_24, 1, x_22); -if (x_21 == 0) -{ -lean_object* x_25; lean_object* x_26; uint8_t x_27; lean_object* x_28; lean_object* x_29; -x_25 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__4; -x_26 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_26, 0, x_25); -lean_ctor_set(x_26, 1, x_24); -x_27 = 0; -x_28 = lean_alloc_ctor(6, 1, 1); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set_uint8(x_28, sizeof(void*)*1, x_27); -x_29 = l_Repr_addAppParen(x_28, x_2); -return x_29; -} -else -{ -lean_object* x_30; lean_object* x_31; uint8_t x_32; lean_object* x_33; lean_object* x_34; -x_30 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__5; -x_31 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_31, 0, x_30); -lean_ctor_set(x_31, 1, x_24); -x_32 = 0; -x_33 = lean_alloc_ctor(6, 1, 1); -lean_ctor_set(x_33, 0, x_31); -lean_ctor_set_uint8(x_33, sizeof(void*)*1, x_32); -x_34 = l_Repr_addAppParen(x_33, x_2); -return x_34; -} -} -case 2: -{ -uint8_t x_35; -x_35 = !lean_is_exclusive(x_1); -if (x_35 == 0) -{ -lean_object* x_36; lean_object* x_37; lean_object* x_38; uint8_t x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; -x_36 = lean_ctor_get(x_1, 0); -x_37 = lean_ctor_get(x_1, 1); -x_38 = lean_unsigned_to_nat(1024u); -x_39 = lean_nat_dec_le(x_38, x_2); -x_40 = l_Lean_Name_reprPrec(x_36, x_38); -x_41 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__11; -lean_ctor_set_tag(x_1, 5); -lean_ctor_set(x_1, 1, x_40); -lean_ctor_set(x_1, 0, x_41); -x_42 = lean_box(1); -x_43 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_43, 0, x_1); -lean_ctor_set(x_43, 1, x_42); -x_44 = l___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2170_(x_37, x_38); -x_45 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_45, 0, x_43); -lean_ctor_set(x_45, 1, x_44); -if (x_39 == 0) -{ -lean_object* x_46; lean_object* x_47; uint8_t x_48; lean_object* x_49; lean_object* x_50; -x_46 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__4; -x_47 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_47, 0, x_46); -lean_ctor_set(x_47, 1, x_45); -x_48 = 0; -x_49 = lean_alloc_ctor(6, 1, 1); -lean_ctor_set(x_49, 0, x_47); -lean_ctor_set_uint8(x_49, sizeof(void*)*1, x_48); -x_50 = l_Repr_addAppParen(x_49, x_2); -return x_50; -} -else -{ -lean_object* x_51; lean_object* x_52; uint8_t x_53; lean_object* x_54; lean_object* x_55; -x_51 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____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_45); -x_53 = 0; -x_54 = lean_alloc_ctor(6, 1, 1); -lean_ctor_set(x_54, 0, x_52); -lean_ctor_set_uint8(x_54, sizeof(void*)*1, x_53); -x_55 = l_Repr_addAppParen(x_54, x_2); -return x_55; -} } else { -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; -x_56 = lean_ctor_get(x_1, 0); -x_57 = lean_ctor_get(x_1, 1); -lean_inc(x_57); -lean_inc(x_56); -lean_dec(x_1); -x_58 = lean_unsigned_to_nat(1024u); -x_59 = lean_nat_dec_le(x_58, x_2); -x_60 = l_Lean_Name_reprPrec(x_56, x_58); -x_61 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__11; -x_62 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_62, 0, x_61); -lean_ctor_set(x_62, 1, x_60); -x_63 = lean_box(1); -x_64 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_64, 0, x_62); -lean_ctor_set(x_64, 1, x_63); -x_65 = l___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2170_(x_57, x_58); -x_66 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_66, 0, x_64); -lean_ctor_set(x_66, 1, x_65); +uint8_t x_59; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_59 = !lean_is_exclusive(x_15); if (x_59 == 0) { -lean_object* x_67; lean_object* x_68; uint8_t x_69; lean_object* x_70; lean_object* x_71; -x_67 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__4; -x_68 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_68, 0, x_67); -lean_ctor_set(x_68, 1, x_66); -x_69 = 0; -x_70 = lean_alloc_ctor(6, 1, 1); -lean_ctor_set(x_70, 0, x_68); -lean_ctor_set_uint8(x_70, sizeof(void*)*1, x_69); -x_71 = l_Repr_addAppParen(x_70, x_2); -return x_71; +return x_15; } else { -lean_object* x_72; lean_object* x_73; uint8_t x_74; lean_object* x_75; lean_object* x_76; -x_72 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__5; -x_73 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_73, 0, x_72); -lean_ctor_set(x_73, 1, x_66); -x_74 = 0; -x_75 = lean_alloc_ctor(6, 1, 1); -lean_ctor_set(x_75, 0, x_73); -lean_ctor_set_uint8(x_75, sizeof(void*)*1, x_74); -x_76 = l_Repr_addAppParen(x_75, x_2); -return x_76; +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_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; +} } } } -default: -{ -lean_object* x_77; uint8_t x_78; -x_77 = lean_unsigned_to_nat(1024u); -x_78 = lean_nat_dec_le(x_77, x_2); -if (x_78 == 0) -{ -lean_object* x_79; lean_object* x_80; -x_79 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__15; -x_80 = l_Repr_addAppParen(x_79, x_2); -return x_80; } else { -lean_object* x_81; lean_object* x_82; -x_81 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__17; -x_82 = l_Repr_addAppParen(x_81, x_2); -return x_82; -} -} +lean_object* x_63; lean_object* x_64; lean_object* x_65; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_63 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_63, 0, x_2); +x_64 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_64, 0, x_63); +x_65 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_65, 0, x_64); +lean_ctor_set(x_65, 1, x_7); +return x_65; } } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_detectOffsets___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { -lean_object* x_3; -x_3 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51_(x_1, x_2); -lean_dec(x_2); -return x_3; +lean_object* x_7; lean_object* x_8; +x_7 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_7, 0, x_1); +x_8 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_8, 0, x_7); +lean_ctor_set(x_8, 1, x_6); +return x_8; } } -static lean_object* _init_l_Lean_Meta_Grind_instReprOrigin___closed__1() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_detectOffsets___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_detectOffsets___lambda__3___boxed), 6, 0); return x_1; } } -static lean_object* _init_l_Lean_Meta_Grind_instReprOrigin() { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_detectOffsets(lean_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_Meta_Grind_instReprOrigin___closed__1; -return x_1; +lean_object* x_7; lean_object* x_8; lean_object* x_9; +lean_inc(x_1); +x_7 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_detectOffsets___lambda__2___boxed), 7, 1); +lean_closure_set(x_7, 0, x_1); +x_8 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_detectOffsets___closed__1; +x_9 = l_Lean_Core_transform___at_Lean_Meta_Grind_unfoldReducible___spec__1(x_1, x_7, x_8, x_2, x_3, x_4, x_5, x_6); +return x_9; } } -static lean_object* _init_l_Lean_Meta_Grind_Origin_key___closed__1() { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_detectOffsets___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("other", 5, 5); -return x_1; +lean_object* x_9; +x_9 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_detectOffsets___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_9; } } -static lean_object* _init_l_Lean_Meta_Grind_Origin_key___closed__2() { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_detectOffsets___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Meta_Grind_Origin_key___closed__1; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; +lean_object* x_8; +x_8 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_detectOffsets___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_1); +return x_8; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Origin_key(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_detectOffsets___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { -if (lean_obj_tag(x_1) == 3) -{ -lean_object* x_2; -x_2 = l_Lean_Meta_Grind_Origin_key___closed__2; -return x_2; -} -else -{ -lean_object* x_3; -x_3 = lean_ctor_get(x_1, 0); -lean_inc(x_3); -return x_3; -} +lean_object* x_7; +x_7 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_detectOffsets___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_7; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Origin_key___boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isOffsetPattern_x3f___lambda__1(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_2; -x_2 = l_Lean_Meta_Grind_Origin_key(x_1); -lean_dec(x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Origin_pp___rarg___lambda__1(lean_object* x_1, lean_object* x_2) { -_start: +if (lean_obj_tag(x_2) == 9) { -lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_3 = lean_ctor_get(x_1, 0); +lean_object* x_3; +x_3 = lean_ctor_get(x_2, 0); lean_inc(x_3); -lean_dec(x_1); -x_4 = lean_ctor_get(x_3, 1); -lean_inc(x_4); -lean_dec(x_3); -x_5 = l_Lean_MessageData_ofConst(x_2); -x_6 = lean_apply_2(x_4, lean_box(0), x_5); -return x_6; -} -} -static lean_object* _init_l_Lean_Meta_Grind_Origin_pp___rarg___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("[unknown]", 9, 9); -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_Grind_Origin_pp___rarg___closed__2() { -_start: +lean_dec(x_2); +if (lean_obj_tag(x_3) == 0) { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_Grind_Origin_pp___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_Meta_Grind_Origin_pp___rarg___closed__3() { -_start: +uint8_t x_4; +x_4 = !lean_is_exclusive(x_3); +if (x_4 == 0) { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_Grind_Origin_pp___rarg___closed__2; -x_2 = l_Lean_MessageData_ofFormat(x_1); -return x_2; -} +lean_object* x_5; lean_object* x_6; +x_5 = lean_ctor_get(x_3, 0); +x_6 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_6, 0, x_1); +lean_ctor_set(x_6, 1, x_5); +lean_ctor_set_tag(x_3, 1); +lean_ctor_set(x_3, 0, x_6); +return x_3; } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Origin_pp___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -switch (lean_obj_tag(x_4)) { -case 0: +else { -lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; -x_5 = lean_ctor_get(x_4, 0); -lean_inc(x_5); -lean_dec(x_4); -x_6 = lean_ctor_get(x_1, 1); -lean_inc(x_6); -lean_inc(x_1); -x_7 = l_Lean_mkConstWithLevelParams___rarg(x_1, x_2, x_3, x_5); -x_8 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_Origin_pp___rarg___lambda__1), 2, 1); -lean_closure_set(x_8, 0, x_1); -x_9 = lean_apply_4(x_6, lean_box(0), lean_box(0), x_7, x_8); +lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_7 = lean_ctor_get(x_3, 0); +lean_inc(x_7); +lean_dec(x_3); +x_8 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_8, 0, x_1); +lean_ctor_set(x_8, 1, x_7); +x_9 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_9, 0, x_8); return x_9; } -case 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_dec(x_3); -lean_dec(x_2); -x_10 = lean_ctor_get(x_4, 0); -lean_inc(x_10); -lean_dec(x_4); -x_11 = lean_ctor_get(x_1, 0); -lean_inc(x_11); -lean_dec(x_1); -x_12 = lean_ctor_get(x_11, 1); -lean_inc(x_12); -lean_dec(x_11); -x_13 = l_Lean_Expr_fvar___override(x_10); -x_14 = l_Lean_MessageData_ofExpr(x_13); -x_15 = lean_apply_2(x_12, lean_box(0), x_14); -return x_15; } -case 2: +else { -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +lean_object* x_10; lean_dec(x_3); -lean_dec(x_2); -x_16 = lean_ctor_get(x_4, 1); -lean_inc(x_16); -lean_dec(x_4); -x_17 = lean_ctor_get(x_1, 0); -lean_inc(x_17); lean_dec(x_1); -x_18 = lean_ctor_get(x_17, 1); -lean_inc(x_18); -lean_dec(x_17); -x_19 = l_Lean_MessageData_ofSyntax(x_16); -x_20 = lean_apply_2(x_18, lean_box(0), x_19); -return x_20; +x_10 = lean_box(0); +return x_10; } -default: +} +else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; -lean_dec(x_3); +lean_object* x_11; lean_dec(x_2); -x_21 = lean_ctor_get(x_1, 0); -lean_inc(x_21); lean_dec(x_1); -x_22 = lean_ctor_get(x_21, 1); -lean_inc(x_22); -lean_dec(x_21); -x_23 = l_Lean_Meta_Grind_Origin_pp___rarg___closed__3; -x_24 = lean_apply_2(x_22, lean_box(0), x_23); -return x_24; -} +x_11 = lean_box(0); +return x_11; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Origin_pp(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isOffsetPattern_x3f___lambda__2(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_Origin_pp___rarg), 4, 0); -return x_2; -} -} -static lean_object* _init_l_Lean_Meta_Grind_instInhabitedEMatchTheorem___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = lean_unsigned_to_nat(0u); -x_2 = lean_mk_empty_array_with_capacity(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Meta_Grind_instInhabitedEMatchTheorem___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = lean_unsigned_to_nat(0u); -x_2 = l_Lean_Expr_bvar___override(x_1); +x_2 = lean_box(0); return x_2; } } -static lean_object* _init_l_Lean_Meta_Grind_instInhabitedEMatchTheorem___closed__3() { +static lean_object* _init_l_Lean_Meta_Grind_isOffsetPattern_x3f___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 = lean_box(0); -x_2 = l_Lean_Meta_Grind_instInhabitedEMatchTheorem___closed__1; -x_3 = l_Lean_Meta_Grind_instInhabitedEMatchTheorem___closed__2; -x_4 = lean_unsigned_to_nat(0u); -x_5 = l_Lean_Meta_Grind_instInhabitedOrigin___closed__1; -x_6 = lean_alloc_ctor(0, 6, 0); -lean_ctor_set(x_6, 0, x_2); -lean_ctor_set(x_6, 1, x_3); -lean_ctor_set(x_6, 2, x_4); -lean_ctor_set(x_6, 3, x_1); -lean_ctor_set(x_6, 4, x_1); -lean_ctor_set(x_6, 5, x_5); -return x_6; +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_isOffsetPattern_x3f___lambda__1), 2, 0); +return x_1; } } -static lean_object* _init_l_Lean_Meta_Grind_instInhabitedEMatchTheorem() { +static lean_object* _init_l_Lean_Meta_Grind_isOffsetPattern_x3f___closed__2() { _start: { lean_object* x_1; -x_1 = l_Lean_Meta_Grind_instInhabitedEMatchTheorem___closed__3; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_isOffsetPattern_x3f___lambda__2___boxed), 1, 0); return x_1; } } -static lean_object* _init_l_panic___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__1___closed__1() { +static lean_object* _init_l_Lean_Meta_Grind_isOffsetPattern_x3f___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Name_instBEq; -x_2 = l_Lean_instHashableName; -x_3 = l_Lean_PersistentHashMap_instInhabited(lean_box(0), lean_box(0), x_1, x_2); +x_1 = l_Lean_Meta_Grind_isOffsetPattern_x3f___closed__2; +x_2 = lean_box(0); +x_3 = lean_apply_1(x_1, x_2); return x_3; } } -static lean_object* _init_l_panic___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__1___closed__2() { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isOffsetPattern_x3f(lean_object* x_1) { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Id_instMonad; -x_2 = l_panic___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__1___closed__1; -x_3 = l_instInhabitedOfMonad___rarg(x_1, x_2); -return x_3; -} -} -LEAN_EXPORT lean_object* l_panic___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__1(lean_object* x_1) { -_start: +lean_object* x_2; lean_object* x_3; uint8_t x_4; +x_2 = l_Lean_Meta_Grind_isOffsetPattern_x3f___closed__1; +x_3 = l_Lean_Expr_cleanupAnnotations(x_1); +x_4 = l_Lean_Expr_isApp(x_3); +if (x_4 == 0) { -lean_object* x_2; lean_object* x_3; -x_2 = l_panic___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__1___closed__2; -x_3 = lean_panic_fn(x_2, x_1); -return x_3; -} +lean_object* x_5; +lean_dec(x_3); +x_5 = l_Lean_Meta_Grind_isOffsetPattern_x3f___closed__3; +return x_5; } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__4(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; 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_6; lean_object* x_7; uint8_t x_8; +x_6 = l_Lean_Expr_appArg(x_3, lean_box(0)); +x_7 = l_Lean_Expr_appFnCleanup(x_3, lean_box(0)); +x_8 = l_Lean_Expr_isApp(x_7); +if (x_8 == 0) { -lean_object* x_8; -lean_dec(x_4); -x_8 = lean_box(0); -return x_8; +lean_object* x_9; +lean_dec(x_7); +lean_dec(x_6); +x_9 = l_Lean_Meta_Grind_isOffsetPattern_x3f___closed__3; +return x_9; } 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_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_10 = l_Lean_Expr_appArg(x_7, lean_box(0)); +x_11 = l_Lean_Expr_appFnCleanup(x_7, lean_box(0)); +x_12 = l_Lean_Meta_Grind_mkOffsetPattern___closed__4; +x_13 = l_Lean_Expr_isConstOf(x_11, x_12); +lean_dec(x_11); +if (x_13 == 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; +lean_object* x_14; +lean_dec(x_10); +lean_dec(x_6); +x_14 = l_Lean_Meta_Grind_isOffsetPattern_x3f___closed__3; +return x_14; } 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); +lean_object* x_15; +x_15 = lean_apply_2(x_2, x_10, x_6); return x_15; } } } } -static size_t _init_l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__1() { +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isOffsetPattern_x3f___lambda__2___boxed(lean_object* x_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; +lean_object* x_2; +x_2 = l_Lean_Meta_Grind_isOffsetPattern_x3f___lambda__2(x_1); +lean_dec(x_1); +return x_2; } } -static size_t _init_l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__2() { +static lean_object* _init_l_Lean_Meta_Grind_preprocessPattern___lambda__1___closed__1() { _start: { -size_t x_1; size_t x_2; size_t x_3; -x_1 = 1; -x_2 = l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__1; -x_3 = lean_usize_sub(x_2, x_1); -return x_3; +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_foldProjs___lambda__3___boxed), 6, 0); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3(lean_object* x_1, size_t x_2, lean_object* x_3) { +static lean_object* _init_l_Lean_Meta_Grind_preprocessPattern___lambda__1___closed__2() { _start: { -if (lean_obj_tag(x_1) == 0) +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_foldProjs___lambda__2), 6, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_preprocessPattern___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_4; -x_4 = !lean_is_exclusive(x_1); -if (x_4 == 0) +lean_object* x_7; +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +x_7 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_detectOffsets(x_1, x_2, x_3, x_4, x_5, x_6); +if (lean_obj_tag(x_7) == 0) { -lean_object* x_5; size_t x_6; size_t x_7; size_t x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; -x_5 = lean_ctor_get(x_1, 0); -x_6 = 5; -x_7 = l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__2; -x_8 = lean_usize_land(x_2, x_7); -x_9 = lean_usize_to_nat(x_8); -x_10 = lean_box(2); -x_11 = lean_array_get(x_10, x_5, x_9); -lean_dec(x_9); -lean_dec(x_5); -switch (lean_obj_tag(x_11)) { -case 0: +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; lean_object* x_13; +x_8 = lean_ctor_get(x_7, 0); +lean_inc(x_8); +x_9 = lean_ctor_get(x_7, 1); +lean_inc(x_9); +lean_dec(x_7); +x_10 = l_Lean_Meta_Grind_preprocessPattern___lambda__1___closed__1; +x_11 = l_Lean_Meta_Grind_preprocessPattern___lambda__1___closed__2; +x_12 = 0; +x_13 = l_Lean_Meta_transform___at_Lean_Meta_zetaReduce___spec__1(x_8, x_10, x_11, x_12, x_12, x_2, x_3, x_4, x_5, x_9); +if (lean_obj_tag(x_13) == 0) { -lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_12 = lean_ctor_get(x_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_name_eq(x_3, x_12); -lean_dec(x_12); +uint8_t x_14; +x_14 = !lean_is_exclusive(x_13); if (x_14 == 0) { -lean_object* x_15; -lean_dec(x_13); -lean_free_object(x_1); -x_15 = lean_box(0); -return x_15; +return x_13; } else { -lean_ctor_set_tag(x_1, 1); -lean_ctor_set(x_1, 0, x_13); -return x_1; -} -} -case 1: -{ -lean_object* x_16; size_t x_17; -lean_free_object(x_1); -x_16 = lean_ctor_get(x_11, 0); +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_dec(x_11); -x_17 = lean_usize_shift_right(x_2, x_6); -x_1 = x_16; -x_2 = x_17; -goto _start; -} -default: -{ -lean_object* x_19; -lean_free_object(x_1); -x_19 = lean_box(0); -return x_19; -} +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 { -lean_object* x_20; size_t x_21; size_t x_22; size_t x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_20 = lean_ctor_get(x_1, 0); -lean_inc(x_20); -lean_dec(x_1); -x_21 = 5; -x_22 = l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__2; -x_23 = lean_usize_land(x_2, x_22); -x_24 = lean_usize_to_nat(x_23); -x_25 = lean_box(2); -x_26 = lean_array_get(x_25, x_20, x_24); -lean_dec(x_24); -lean_dec(x_20); -switch (lean_obj_tag(x_26)) { -case 0: -{ -lean_object* x_27; lean_object* x_28; uint8_t x_29; -x_27 = lean_ctor_get(x_26, 0); -lean_inc(x_27); -x_28 = lean_ctor_get(x_26, 1); -lean_inc(x_28); -lean_dec(x_26); -x_29 = lean_name_eq(x_3, x_27); -lean_dec(x_27); -if (x_29 == 0) +uint8_t x_18; +x_18 = !lean_is_exclusive(x_13); +if (x_18 == 0) { -lean_object* x_30; -lean_dec(x_28); -x_30 = lean_box(0); -return x_30; +return x_13; } else { -lean_object* x_31; -x_31 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_31, 0, x_28); -return x_31; +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_13, 0); +x_20 = lean_ctor_get(x_13, 1); +lean_inc(x_20); +lean_inc(x_19); +lean_dec(x_13); +x_21 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_21, 0, x_19); +lean_ctor_set(x_21, 1, x_20); +return x_21; } } -case 1: -{ -lean_object* x_32; size_t x_33; -x_32 = lean_ctor_get(x_26, 0); -lean_inc(x_32); -lean_dec(x_26); -x_33 = lean_usize_shift_right(x_2, x_21); -x_1 = x_32; -x_2 = x_33; -goto _start; } -default: +else { -lean_object* x_35; -x_35 = lean_box(0); -return x_35; -} -} -} +uint8_t x_22; +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_22 = !lean_is_exclusive(x_7); +if (x_22 == 0) +{ +return x_7; } else { -lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_36 = lean_ctor_get(x_1, 0); -lean_inc(x_36); -x_37 = lean_ctor_get(x_1, 1); -lean_inc(x_37); -lean_dec(x_1); -x_38 = lean_unsigned_to_nat(0u); -x_39 = l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__4(x_36, x_37, lean_box(0), x_38, x_3); -lean_dec(x_37); -lean_dec(x_36); -return x_39; +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_7, 0); +x_24 = lean_ctor_get(x_7, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_7); +x_25 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_24); +return x_25; } } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__2(lean_object* x_1, lean_object* x_2) { +} +static lean_object* _init_l_Lean_Meta_Grind_preprocessPattern___closed__1() { _start: { -uint64_t x_3; size_t x_4; lean_object* x_5; -x_3 = l_Lean_Name_hash___override(x_2); -x_4 = lean_uint64_to_usize(x_3); -x_5 = l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3(x_1, x_4, x_2); -return x_5; +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_unfoldReducible___lambda__2), 6, 0); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__7(size_t 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_Grind_preprocessPattern___closed__2() { _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; +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_unfoldReducible___lambda__3___boxed), 6, 0); +return x_1; } -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_Grind_EMatchTheorems_insert___spec__6(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; } +static lean_object* _init_l_Lean_Meta_Grind_preprocessPattern___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_preprocessPattern___lambda__1), 6, 0); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__8(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_preprocessPattern(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_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_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_8 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_1, x_3, x_4, x_5, x_6, x_7); +x_9 = lean_ctor_get(x_8, 0); +lean_inc(x_9); +x_10 = lean_ctor_get(x_8, 1); +lean_inc(x_10); +lean_dec(x_8); +x_11 = l_Lean_Meta_Grind_preprocessPattern___closed__1; +x_12 = l_Lean_Meta_Grind_preprocessPattern___closed__2; lean_inc(x_6); -x_7 = lean_array_get_size(x_5); -x_8 = lean_nat_dec_lt(x_2, x_7); -lean_dec(x_7); -if (x_8 == 0) +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_13 = l_Lean_Core_transform___at_Lean_Meta_Grind_unfoldReducible___spec__1(x_9, x_11, x_12, x_3, x_4, x_5, x_6, x_10); +if (lean_obj_tag(x_13) == 0) { -uint8_t x_9; -lean_dec(x_2); -x_9 = !lean_is_exclusive(x_1); -if (x_9 == 0) +lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +x_16 = l_Lean_Meta_Grind_preprocessPattern___closed__3; +if (x_2 == 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; +lean_object* x_17; +x_17 = lean_apply_6(x_16, x_14, x_3, x_4, x_5, x_6, x_15); +return x_17; } 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; -} +lean_object* x_18; +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_18 = lean_grind_normalize(x_14, x_3, x_4, x_5, x_6, x_15); +if (lean_obj_tag(x_18) == 0) +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_18, 1); +lean_inc(x_20); +lean_dec(x_18); +x_21 = lean_apply_6(x_16, x_19, x_3, x_4, x_5, x_6, x_20); +return x_21; } 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; +uint8_t x_22; 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; +lean_dec(x_4); +lean_dec(x_3); +x_22 = !lean_is_exclusive(x_18); +if (x_22 == 0) +{ +return x_18; } 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; +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_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; +uint8_t x_26; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_26 = !lean_is_exclusive(x_13); +if (x_26 == 0) +{ +return x_13; } 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_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 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_29, 0, x_27); lean_ctor_set(x_29, 1, x_28); @@ -1524,866 +1645,792 @@ return x_29; } } } +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_preprocessPattern___boxed(lean_object* x_1, 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_Grind_preprocessPattern(x_1, x_8, x_3, x_4, x_5, x_6, x_7); +return x_9; +} +} +static lean_object* _init_l_Lean_Meta_Grind_instInhabitedOrigin___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(0); +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_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__6___closed__1() { +static lean_object* _init_l_Lean_Meta_Grind_instInhabitedOrigin() { _start: { lean_object* x_1; -x_1 = l_Lean_PersistentHashMap_mkEmptyEntries(lean_box(0), lean_box(0)); +x_1 = l_Lean_Meta_Grind_instInhabitedOrigin___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__6(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5) { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__1() { _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_Grind_EMatchTheorems_insert___spec__3___closed__2; -x_11 = lean_usize_land(x_2, x_10); -x_12 = lean_usize_to_nat(x_11); -x_13 = lean_array_get_size(x_7); -x_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); +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.Meta.Grind.Origin.decl", 27, 27); return x_1; } -else +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__2() { +_start: { -lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = lean_array_fget(x_7, x_12); -x_16 = lean_box(0); -x_17 = lean_array_fset(x_7, x_12, x_16); -switch (lean_obj_tag(x_15)) { -case 0: +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__1; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__3() { +_start: { -uint8_t x_18; -x_18 = !lean_is_exclusive(x_15); -if (x_18 == 0) +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__2; +x_2 = lean_box(1); +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__4() { +_start: { -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_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_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__5() { +_start: { -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; +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; } -else +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__6() { +_start: { -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); +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.Meta.Grind.Origin.fvar", 27, 27); return x_1; } } -else +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__7() { +_start: { -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_1; lean_object* x_2; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__6; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__8() { +_start: { -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; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__7; +x_2 = lean_box(1); +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; } -else +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__9() { +_start: { -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); +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.Meta.Grind.Origin.stx", 26, 26); return x_1; } } -} -case 1: -{ -uint8_t x_34; -x_34 = !lean_is_exclusive(x_15); -if (x_34 == 0) +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__10() { +_start: { -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_Grind_EMatchTheorems_insert___spec__6(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; +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__9; +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___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__11() { +_start: { -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_Grind_EMatchTheorems_insert___spec__6(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; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__10; +x_2 = lean_box(1); +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; } } -default: +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__12() { +_start: { -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); +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.Meta.Grind.Origin.local", 28, 28); return x_1; } } +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__12; +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___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__14() { +_start: { -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_Grind_EMatchTheorems_insert___spec__3___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; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__13; +x_2 = lean_box(1); +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; } -else +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614_(lean_object* x_1, lean_object* x_2) { +_start: { -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)) { +switch (lean_obj_tag(x_1)) { 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_3; lean_object* x_4; uint8_t x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_3 = lean_ctor_get(x_1, 0); +lean_inc(x_3); +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_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____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_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; +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_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____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); +x_11 = 0; +x_12 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_12, 0, x_10); +lean_ctor_set_uint8(x_12, sizeof(void*)*1, x_11); +x_13 = l_Repr_addAppParen(x_12, x_2); +return x_13; } 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; +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_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____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); +x_16 = 0; +x_17 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_17, 0, x_15); +lean_ctor_set_uint8(x_17, sizeof(void*)*1, x_16); +x_18 = l_Repr_addAppParen(x_17, x_2); +return x_18; } } 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_Grind_EMatchTheorems_insert___spec__6(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_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_1, 0); +lean_inc(x_19); +lean_dec(x_1); +x_20 = lean_unsigned_to_nat(1024u); +x_21 = lean_nat_dec_le(x_20, x_2); +x_22 = l_Lean_Name_reprPrec(x_19, x_20); +x_23 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__8; +x_24 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set(x_24, 1, x_22); +if (x_21 == 0) { -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; -} -} -} -} +lean_object* x_25; lean_object* x_26; uint8_t x_27; lean_object* x_28; lean_object* x_29; +x_25 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__4; +x_26 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set(x_26, 1, x_24); +x_27 = 0; +x_28 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_28, 0, x_26); +lean_ctor_set_uint8(x_28, sizeof(void*)*1, x_27); +x_29 = l_Repr_addAppParen(x_28, x_2); +return x_29; } 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_Grind_EMatchTheorems_insert___spec__8(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_30; lean_object* x_31; uint8_t x_32; lean_object* x_33; lean_object* x_34; +x_30 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__5; +x_31 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_31, 0, x_30); +lean_ctor_set(x_31, 1, x_24); +x_32 = 0; +x_33 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_33, 0, x_31); +lean_ctor_set_uint8(x_33, sizeof(void*)*1, x_32); +x_34 = l_Repr_addAppParen(x_33, x_2); +return x_34; +} +} +case 2: { -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) +uint8_t x_35; +x_35 = !lean_is_exclusive(x_1); +if (x_35 == 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_Grind_EMatchTheorems_insert___spec__6___closed__1; -x_93 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__7(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 +lean_object* x_36; lean_object* x_37; lean_object* x_38; uint8_t x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_36 = lean_ctor_get(x_1, 0); +x_37 = lean_ctor_get(x_1, 1); +x_38 = lean_unsigned_to_nat(1024u); +x_39 = lean_nat_dec_le(x_38, x_2); +x_40 = l_Lean_Name_reprPrec(x_36, x_38); +x_41 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__11; +lean_ctor_set_tag(x_1, 5); +lean_ctor_set(x_1, 1, x_40); +lean_ctor_set(x_1, 0, x_41); +x_42 = lean_box(1); +x_43 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_43, 0, x_1); +lean_ctor_set(x_43, 1, x_42); +x_44 = l___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2170_(x_37, x_38); +x_45 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +if (x_39 == 0) { -return x_84; -} +lean_object* x_46; lean_object* x_47; uint8_t x_48; lean_object* x_49; lean_object* x_50; +x_46 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__4; +x_47 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_47, 0, x_46); +lean_ctor_set(x_47, 1, x_45); +x_48 = 0; +x_49 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_49, 0, x_47); +lean_ctor_set_uint8(x_49, sizeof(void*)*1, x_48); +x_50 = l_Repr_addAppParen(x_49, x_2); +return x_50; } else { -return x_84; +lean_object* x_51; lean_object* x_52; uint8_t x_53; lean_object* x_54; lean_object* x_55; +x_51 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____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_45); +x_53 = 0; +x_54 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_54, 0, x_52); +lean_ctor_set_uint8(x_54, sizeof(void*)*1, x_53); +x_55 = l_Repr_addAppParen(x_54, x_2); +return x_55; } } 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_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; +x_56 = lean_ctor_get(x_1, 0); +x_57 = lean_ctor_get(x_1, 1); +lean_inc(x_57); +lean_inc(x_56); 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_Grind_EMatchTheorems_insert___spec__8(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) +x_58 = lean_unsigned_to_nat(1024u); +x_59 = lean_nat_dec_le(x_58, x_2); +x_60 = l_Lean_Name_reprPrec(x_56, x_58); +x_61 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__11; +x_62 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_62, 0, x_61); +lean_ctor_set(x_62, 1, x_60); +x_63 = lean_box(1); +x_64 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_64, 0, x_62); +lean_ctor_set(x_64, 1, x_63); +x_65 = l___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2170_(x_57, x_58); +x_66 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_66, 0, x_64); +lean_ctor_set(x_66, 1, x_65); +if (x_59 == 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_Grind_EMatchTheorems_insert___spec__6___closed__1; -x_107 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__7(x_3, x_104, x_105, lean_box(0), x_97, x_106); -lean_dec(x_105); -lean_dec(x_104); -return x_107; +lean_object* x_67; lean_object* x_68; uint8_t x_69; lean_object* x_70; lean_object* x_71; +x_67 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__4; +x_68 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_68, 0, x_67); +lean_ctor_set(x_68, 1, x_66); +x_69 = 0; +x_70 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_70, 0, x_68); +lean_ctor_set_uint8(x_70, sizeof(void*)*1, x_69); +x_71 = l_Repr_addAppParen(x_70, x_2); +return x_71; } else { -return x_98; +lean_object* x_72; lean_object* x_73; uint8_t x_74; lean_object* x_75; lean_object* x_76; +x_72 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__5; +x_73 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_73, 0, x_72); +lean_ctor_set(x_73, 1, x_66); +x_74 = 0; +x_75 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_75, 0, x_73); +lean_ctor_set_uint8(x_75, sizeof(void*)*1, x_74); +x_76 = l_Repr_addAppParen(x_75, x_2); +return x_76; } } -else +} +default: { -return x_98; +lean_object* x_77; lean_object* x_78; uint8_t x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; +x_77 = lean_ctor_get(x_1, 0); +lean_inc(x_77); +lean_dec(x_1); +x_78 = lean_unsigned_to_nat(1024u); +x_79 = lean_nat_dec_le(x_78, x_2); +x_80 = l_Lean_Name_reprPrec(x_77, x_78); +x_81 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__14; +x_82 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_82, 0, x_81); +lean_ctor_set(x_82, 1, x_80); +if (x_79 == 0) +{ +lean_object* x_83; lean_object* x_84; uint8_t x_85; lean_object* x_86; lean_object* x_87; +x_83 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__4; +x_84 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_84, 0, x_83); +lean_ctor_set(x_84, 1, x_82); +x_85 = 0; +x_86 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_86, 0, x_84); +lean_ctor_set_uint8(x_86, sizeof(void*)*1, x_85); +x_87 = l_Repr_addAppParen(x_86, x_2); +return x_87; +} +else +{ +lean_object* x_88; lean_object* x_89; uint8_t x_90; lean_object* x_91; lean_object* x_92; +x_88 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__5; +x_89 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_89, 0, x_88); +lean_ctor_set(x_89, 1, x_82); +x_90 = 0; +x_91 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_91, 0, x_89); +lean_ctor_set_uint8(x_91, sizeof(void*)*1, x_90); +x_92 = l_Repr_addAppParen(x_91, x_2); +return x_92; } } } } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____boxed(lean_object* x_1, lean_object* x_2) { _start: { -uint64_t x_4; size_t x_5; size_t x_6; lean_object* x_7; -x_4 = l_Lean_Name_hash___override(x_2); -x_5 = lean_uint64_to_usize(x_4); -x_6 = 1; -x_7 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__6(x_1, x_5, x_6, x_2, x_3); -return x_7; +lean_object* x_3; +x_3 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614_(x_1, x_2); +lean_dec(x_2); +return x_3; } } -static lean_object* _init_l_Lean_Meta_Grind_EMatchTheorems_insert___closed__1() { +static lean_object* _init_l_Lean_Meta_Grind_instReprOrigin___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean.Meta.Tactic.Grind.EMatchTheorem", 36, 36); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____boxed), 2, 0); return x_1; } } -static lean_object* _init_l_Lean_Meta_Grind_EMatchTheorems_insert___closed__2() { +static lean_object* _init_l_Lean_Meta_Grind_instReprOrigin() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean.Meta.Grind.EMatchTheorems.insert", 37, 37); +x_1 = l_Lean_Meta_Grind_instReprOrigin___closed__1; return x_1; } } -static lean_object* _init_l_Lean_Meta_Grind_EMatchTheorems_insert___closed__3() { +LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_beqOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_811_(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("unreachable code has been reached", 33, 33); -return x_1; -} +switch (lean_obj_tag(x_1)) { +case 0: +{ +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_3; lean_object* x_4; uint8_t x_5; +x_3 = lean_ctor_get(x_1, 0); +lean_inc(x_3); +lean_dec(x_1); +x_4 = lean_ctor_get(x_2, 0); +lean_inc(x_4); +lean_dec(x_2); +x_5 = lean_name_eq(x_3, x_4); +lean_dec(x_4); +lean_dec(x_3); +return x_5; } -static lean_object* _init_l_Lean_Meta_Grind_EMatchTheorems_insert___closed__4() { -_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 = l_Lean_Meta_Grind_EMatchTheorems_insert___closed__1; -x_2 = l_Lean_Meta_Grind_EMatchTheorems_insert___closed__2; -x_3 = lean_unsigned_to_nat(64u); -x_4 = lean_unsigned_to_nat(6u); -x_5 = l_Lean_Meta_Grind_EMatchTheorems_insert___closed__3; -x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); +uint8_t x_6; +lean_dec(x_2); +lean_dec(x_1); +x_6 = 0; return x_6; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_EMatchTheorems_insert(lean_object* x_1, lean_object* x_2) { -_start: +case 1: { -lean_object* x_3; -x_3 = lean_ctor_get(x_2, 4); -lean_inc(x_3); -if (lean_obj_tag(x_3) == 0) +if (lean_obj_tag(x_2) == 1) { -lean_object* x_4; lean_object* x_5; -lean_dec(x_2); +lean_object* x_7; lean_object* x_8; uint8_t x_9; +x_7 = lean_ctor_get(x_1, 0); +lean_inc(x_7); lean_dec(x_1); -x_4 = l_Lean_Meta_Grind_EMatchTheorems_insert___closed__4; -x_5 = l_panic___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__1(x_4); -return x_5; +x_8 = lean_ctor_get(x_2, 0); +lean_inc(x_8); +lean_dec(x_2); +x_9 = lean_name_eq(x_7, x_8); +lean_dec(x_8); +lean_dec(x_7); +return x_9; } else { -lean_object* x_6; -x_6 = lean_ctor_get(x_3, 0); -lean_inc(x_6); -if (lean_obj_tag(x_6) == 2) -{ -uint8_t x_7; -x_7 = !lean_is_exclusive(x_2); -if (x_7 == 0) +uint8_t x_10; +lean_dec(x_2); +lean_dec(x_1); +x_10 = 0; +return x_10; +} +} +case 2: { -lean_object* x_8; uint8_t x_9; -x_8 = lean_ctor_get(x_2, 4); -lean_dec(x_8); -x_9 = !lean_is_exclusive(x_3); -if (x_9 == 0) +if (lean_obj_tag(x_2) == 2) { -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_10 = lean_ctor_get(x_3, 1); -x_11 = lean_ctor_get(x_3, 0); -lean_dec(x_11); -x_12 = lean_ctor_get(x_6, 0); +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_11 = lean_ctor_get(x_1, 0); +lean_inc(x_11); +x_12 = lean_ctor_get(x_1, 1); lean_inc(x_12); -lean_dec(x_6); -lean_ctor_set(x_2, 4, x_10); -lean_inc(x_1); -x_13 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__2(x_1, x_12); -if (lean_obj_tag(x_13) == 0) +lean_dec(x_1); +x_13 = lean_ctor_get(x_2, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_2, 1); +lean_inc(x_14); +lean_dec(x_2); +x_15 = lean_name_eq(x_11, x_13); +lean_dec(x_13); +lean_dec(x_11); +if (x_15 == 0) { -lean_object* x_14; lean_object* x_15; -x_14 = lean_box(0); -lean_ctor_set(x_3, 1, x_14); -lean_ctor_set(x_3, 0, x_2); -x_15 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__5(x_1, x_12, x_3); -return x_15; +uint8_t x_16; +lean_dec(x_14); +lean_dec(x_12); +x_16 = 0; +return x_16; } else { -lean_object* x_16; lean_object* x_17; -x_16 = lean_ctor_get(x_13, 0); -lean_inc(x_16); -lean_dec(x_13); -lean_ctor_set(x_3, 1, x_16); -lean_ctor_set(x_3, 0, x_2); -x_17 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__5(x_1, x_12, x_3); +uint8_t x_17; +x_17 = l_Lean_Syntax_structEq(x_12, x_14); return x_17; } } else { -lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_18 = lean_ctor_get(x_3, 1); -lean_inc(x_18); -lean_dec(x_3); -x_19 = lean_ctor_get(x_6, 0); -lean_inc(x_19); -lean_dec(x_6); -lean_ctor_set(x_2, 4, x_18); -lean_inc(x_1); -x_20 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__2(x_1, x_19); -if (lean_obj_tag(x_20) == 0) -{ -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_box(0); -x_22 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_22, 0, x_2); -lean_ctor_set(x_22, 1, x_21); -x_23 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__5(x_1, x_19, x_22); -return x_23; +uint8_t x_18; +lean_dec(x_2); +lean_dec(x_1); +x_18 = 0; +return x_18; } -else +} +default: { -lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_24 = lean_ctor_get(x_20, 0); -lean_inc(x_24); +if (lean_obj_tag(x_2) == 3) +{ +lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_19 = lean_ctor_get(x_1, 0); +lean_inc(x_19); +lean_dec(x_1); +x_20 = lean_ctor_get(x_2, 0); +lean_inc(x_20); +lean_dec(x_2); +x_21 = lean_name_eq(x_19, x_20); lean_dec(x_20); -x_25 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_25, 0, x_2); -lean_ctor_set(x_25, 1, x_24); -x_26 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__5(x_1, x_19, x_25); -return x_26; -} -} +lean_dec(x_19); +return x_21; } else { -lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_27 = lean_ctor_get(x_2, 0); -x_28 = lean_ctor_get(x_2, 1); -x_29 = lean_ctor_get(x_2, 2); -x_30 = lean_ctor_get(x_2, 3); -x_31 = lean_ctor_get(x_2, 5); -lean_inc(x_31); -lean_inc(x_30); -lean_inc(x_29); -lean_inc(x_28); -lean_inc(x_27); +uint8_t x_22; lean_dec(x_2); -x_32 = lean_ctor_get(x_3, 1); -lean_inc(x_32); -if (lean_is_exclusive(x_3)) { - lean_ctor_release(x_3, 0); - lean_ctor_release(x_3, 1); - x_33 = x_3; -} else { - lean_dec_ref(x_3); - x_33 = lean_box(0); +lean_dec(x_1); +x_22 = 0; +return x_22; } -x_34 = lean_ctor_get(x_6, 0); -lean_inc(x_34); -lean_dec(x_6); -x_35 = lean_alloc_ctor(0, 6, 0); -lean_ctor_set(x_35, 0, x_27); -lean_ctor_set(x_35, 1, x_28); -lean_ctor_set(x_35, 2, x_29); -lean_ctor_set(x_35, 3, x_30); -lean_ctor_set(x_35, 4, x_32); -lean_ctor_set(x_35, 5, x_31); -lean_inc(x_1); -x_36 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__2(x_1, x_34); -if (lean_obj_tag(x_36) == 0) -{ -lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_37 = lean_box(0); -if (lean_is_scalar(x_33)) { - x_38 = lean_alloc_ctor(1, 2, 0); -} else { - x_38 = x_33; } -lean_ctor_set(x_38, 0, x_35); -lean_ctor_set(x_38, 1, x_37); -x_39 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__5(x_1, x_34, x_38); -return x_39; } -else -{ -lean_object* x_40; lean_object* x_41; lean_object* x_42; -x_40 = lean_ctor_get(x_36, 0); -lean_inc(x_40); -lean_dec(x_36); -if (lean_is_scalar(x_33)) { - x_41 = lean_alloc_ctor(1, 2, 0); -} else { - x_41 = x_33; } -lean_ctor_set(x_41, 0, x_35); -lean_ctor_set(x_41, 1, x_40); -x_42 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__5(x_1, x_34, x_41); -return x_42; } +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_beqOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_811____boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_beqOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_811_(x_1, x_2); +x_4 = lean_box(x_3); +return x_4; } } -else -{ -lean_object* x_43; lean_object* x_44; -lean_dec(x_6); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_43 = l_Lean_Meta_Grind_EMatchTheorems_insert___closed__4; -x_44 = l_panic___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__1(x_43); -return x_44; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +static lean_object* _init_l_Lean_Meta_Grind_instBEqOrigin___closed__1() { _start: { -lean_object* x_6; -x_6 = l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__4(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_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_beqOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_811____boxed), 2, 0); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +static lean_object* _init_l_Lean_Meta_Grind_instBEqOrigin() { _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_Grind_EMatchTheorems_insert___spec__3(x_1, x_4, x_3); -lean_dec(x_3); -return x_5; +lean_object* x_1; +x_1 = l_Lean_Meta_Grind_instBEqOrigin___closed__1; +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__2___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Origin_key(lean_object* x_1) { _start: { -lean_object* x_3; -x_3 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__2(x_1, x_2); -lean_dec(x_2); -return x_3; +lean_object* x_2; +x_2 = lean_ctor_get(x_1, 0); +lean_inc(x_2); +return x_2; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_EMatchTheorems_insert___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_Lean_Meta_Grind_Origin_key___boxed(lean_object* x_1) { _start: { -size_t x_7; lean_object* x_8; -x_7 = lean_unbox_usize(x_1); +lean_object* x_2; +x_2 = l_Lean_Meta_Grind_Origin_key(x_1); lean_dec(x_1); -x_8 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__7(x_7, x_2, x_3, x_4, x_5, x_6); -lean_dec(x_3); -lean_dec(x_2); -return x_8; +return x_2; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Origin_pp___rarg___lambda__1(lean_object* x_1, lean_object* x_2) { _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_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_3 = lean_ctor_get(x_1, 0); +lean_inc(x_3); +lean_dec(x_1); +x_4 = lean_ctor_get(x_3, 1); +lean_inc(x_4); lean_dec(x_3); -x_8 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__6(x_1, x_6, x_7, x_4, x_5); -return x_8; +x_5 = l_Lean_MessageData_ofConst(x_2); +x_6 = lean_apply_2(x_4, lean_box(0), x_5); +return x_6; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_EMatchTheorem_getProofWithFreshMVarLevels(lean_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_Grind_Origin_pp___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -lean_object* x_7; uint8_t x_8; -x_7 = lean_ctor_get(x_1, 1); -lean_inc(x_7); -x_8 = l_Lean_Expr_isConst(x_7); -if (x_8 == 0) -{ -lean_object* x_9; uint8_t x_10; -x_9 = lean_ctor_get(x_1, 0); -lean_inc(x_9); -lean_dec(x_1); -x_10 = l_Array_isEmpty___rarg(x_9); -if (x_10 == 0) +switch (lean_obj_tag(x_4)) { +case 0: { -size_t x_11; size_t x_12; lean_object* x_13; uint8_t x_14; -x_11 = lean_array_size(x_9); -x_12 = 0; -lean_inc(x_9); -x_13 = l_Array_mapMUnsafe_map___at_Lean_Meta_openAbstractMVarsResult___spec__1(x_11, x_12, x_9, x_2, x_3, x_4, x_5, x_6); -x_14 = !lean_is_exclusive(x_13); -if (x_14 == 0) +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_5 = lean_ctor_get(x_4, 0); +lean_inc(x_5); +lean_dec(x_4); +x_6 = lean_ctor_get(x_1, 1); +lean_inc(x_6); +lean_inc(x_1); +x_7 = l_Lean_mkConstWithLevelParams___rarg(x_1, x_2, x_3, x_5); +x_8 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_Origin_pp___rarg___lambda__1), 2, 1); +lean_closure_set(x_8, 0, x_1); +x_9 = lean_apply_4(x_6, lean_box(0), lean_box(0), x_7, x_8); +return x_9; +} +case 1: { -lean_object* x_15; lean_object* x_16; -x_15 = lean_ctor_get(x_13, 0); -x_16 = l_Lean_Expr_instantiateLevelParamsArray(x_7, x_9, x_15); -lean_dec(x_7); -lean_ctor_set(x_13, 0, x_16); -return x_13; +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +lean_dec(x_3); +lean_dec(x_2); +x_10 = lean_ctor_get(x_4, 0); +lean_inc(x_10); +lean_dec(x_4); +x_11 = lean_ctor_get(x_1, 0); +lean_inc(x_11); +lean_dec(x_1); +x_12 = lean_ctor_get(x_11, 1); +lean_inc(x_12); +lean_dec(x_11); +x_13 = l_Lean_Expr_fvar___override(x_10); +x_14 = l_Lean_MessageData_ofExpr(x_13); +x_15 = lean_apply_2(x_12, lean_box(0), x_14); +return x_15; } -else +case 2: { -lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_17 = lean_ctor_get(x_13, 0); -x_18 = lean_ctor_get(x_13, 1); -lean_inc(x_18); +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +lean_dec(x_3); +lean_dec(x_2); +x_16 = lean_ctor_get(x_4, 1); +lean_inc(x_16); +lean_dec(x_4); +x_17 = lean_ctor_get(x_1, 0); lean_inc(x_17); -lean_dec(x_13); -x_19 = l_Lean_Expr_instantiateLevelParamsArray(x_7, x_9, x_17); -lean_dec(x_7); -x_20 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_18); +lean_dec(x_1); +x_18 = lean_ctor_get(x_17, 1); +lean_inc(x_18); +lean_dec(x_17); +x_19 = l_Lean_MessageData_ofSyntax(x_16); +x_20 = lean_apply_2(x_18, lean_box(0), x_19); return x_20; } -} -else -{ -lean_object* x_21; -lean_dec(x_9); -x_21 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_21, 0, x_7); -lean_ctor_set(x_21, 1, x_6); -return x_21; -} -} -else +default: { -lean_object* x_22; uint8_t 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); +lean_dec(x_2); +x_21 = lean_ctor_get(x_4, 0); +lean_inc(x_21); +lean_dec(x_4); x_22 = lean_ctor_get(x_1, 0); lean_inc(x_22); lean_dec(x_1); -x_23 = l_Array_isEmpty___rarg(x_22); -if (x_23 == 0) -{ -size_t x_24; size_t x_25; lean_object* x_26; uint8_t x_27; -x_24 = lean_array_size(x_22); -x_25 = 0; -lean_inc(x_22); -x_26 = l_Array_mapMUnsafe_map___at_Lean_Meta_openAbstractMVarsResult___spec__1(x_24, x_25, x_22, x_2, x_3, x_4, x_5, x_6); -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); -x_29 = l_Lean_Expr_instantiateLevelParamsArray(x_7, x_22, x_28); -lean_dec(x_7); -lean_ctor_set(x_26, 0, x_29); -return x_26; +x_23 = lean_ctor_get(x_22, 1); +lean_inc(x_23); +lean_dec(x_22); +x_24 = l_Lean_MessageData_ofName(x_21); +x_25 = lean_apply_2(x_23, lean_box(0), x_24); +return x_25; } -else -{ -lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_30 = lean_ctor_get(x_26, 0); -x_31 = lean_ctor_get(x_26, 1); -lean_inc(x_31); -lean_inc(x_30); -lean_dec(x_26); -x_32 = l_Lean_Expr_instantiateLevelParamsArray(x_7, x_22, x_30); -lean_dec(x_7); -x_33 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_33, 0, x_32); -lean_ctor_set(x_33, 1, x_31); -return x_33; } } -else +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Origin_pp(lean_object* x_1) { +_start: { -lean_object* x_34; lean_object* x_35; -lean_dec(x_22); -x_34 = l_Lean_Expr_constName_x21(x_7); -lean_inc(x_34); -x_35 = l_Lean_getConstInfo___at_Lean_Meta_mkConstWithFreshMVarLevels___spec__1(x_34, x_2, x_3, x_4, x_5, x_6); -if (lean_obj_tag(x_35) == 0) +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_Origin_pp___rarg), 4, 0); +return x_2; +} +} +LEAN_EXPORT uint8_t l_Lean_Meta_Grind_instBEqOrigin__1(lean_object* x_1, lean_object* x_2) { +_start: { -uint8_t x_36; -x_36 = !lean_is_exclusive(x_35); -if (x_36 == 0) +lean_object* x_3; lean_object* x_4; uint8_t x_5; +x_3 = l_Lean_Meta_Grind_Origin_key(x_1); +x_4 = l_Lean_Meta_Grind_Origin_key(x_2); +x_5 = lean_name_eq(x_3, x_4); +lean_dec(x_4); +lean_dec(x_3); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instBEqOrigin__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: { -lean_object* x_37; lean_object* x_38; lean_object* x_39; uint8_t x_40; -x_37 = lean_ctor_get(x_35, 0); -x_38 = lean_ctor_get(x_35, 1); -x_39 = l_Lean_ConstantInfo_levelParams(x_37); -lean_dec(x_37); -x_40 = l_List_isEmpty___rarg(x_39); -lean_dec(x_39); -if (x_40 == 0) +uint8_t x_3; lean_object* x_4; +x_3 = l_Lean_Meta_Grind_instBEqOrigin__1(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT uint64_t l_Lean_Meta_Grind_instHashableOrigin(lean_object* x_1) { +_start: { -lean_object* x_41; -lean_free_object(x_35); -lean_dec(x_7); -x_41 = l_Lean_Meta_mkConstWithFreshMVarLevels(x_34, x_2, x_3, x_4, x_5, x_38); -return x_41; +lean_object* x_2; uint64_t x_3; +x_2 = l_Lean_Meta_Grind_Origin_key(x_1); +x_3 = l_Lean_Name_hash___override(x_2); +lean_dec(x_2); +return x_3; } -else +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instHashableOrigin___boxed(lean_object* x_1) { +_start: { -lean_dec(x_34); -lean_ctor_set(x_35, 0, x_7); -return x_35; +uint64_t x_2; lean_object* x_3; +x_2 = l_Lean_Meta_Grind_instHashableOrigin(x_1); +lean_dec(x_1); +x_3 = lean_box_uint64(x_2); +return x_3; } } -else +static lean_object* _init_l_Lean_Meta_Grind_instInhabitedEMatchTheorem___closed__1() { +_start: { -lean_object* x_42; lean_object* x_43; lean_object* x_44; uint8_t x_45; -x_42 = lean_ctor_get(x_35, 0); -x_43 = lean_ctor_get(x_35, 1); -lean_inc(x_43); -lean_inc(x_42); -lean_dec(x_35); -x_44 = l_Lean_ConstantInfo_levelParams(x_42); -lean_dec(x_42); -x_45 = l_List_isEmpty___rarg(x_44); -lean_dec(x_44); -if (x_45 == 0) -{ -lean_object* x_46; -lean_dec(x_7); -x_46 = l_Lean_Meta_mkConstWithFreshMVarLevels(x_34, x_2, x_3, x_4, x_5, x_43); -return x_46; -} -else -{ -lean_object* x_47; -lean_dec(x_34); -x_47 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_47, 0, x_7); -lean_ctor_set(x_47, 1, x_43); -return x_47; -} -} +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; } -else -{ -uint8_t x_48; -lean_dec(x_34); -lean_dec(x_7); -x_48 = !lean_is_exclusive(x_35); -if (x_48 == 0) -{ -return x_35; } -else +static lean_object* _init_l_Lean_Meta_Grind_instInhabitedEMatchTheorem___closed__2() { +_start: { -lean_object* x_49; lean_object* x_50; lean_object* x_51; -x_49 = lean_ctor_get(x_35, 0); -x_50 = lean_ctor_get(x_35, 1); -lean_inc(x_50); -lean_inc(x_49); -lean_dec(x_35); -x_51 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_51, 0, x_49); -lean_ctor_set(x_51, 1, x_50); -return x_51; -} -} +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = l_Lean_Expr_bvar___override(x_1); +return x_2; } } +static lean_object* _init_l_Lean_Meta_Grind_instInhabitedEMatchTheorem___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind_instInhabitedEMatchTheorem___closed__1; +x_3 = l_Lean_Meta_Grind_instInhabitedEMatchTheorem___closed__2; +x_4 = lean_unsigned_to_nat(0u); +x_5 = l_Lean_Meta_Grind_instInhabitedOrigin___closed__1; +x_6 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_6, 0, x_2); +lean_ctor_set(x_6, 1, x_3); +lean_ctor_set(x_6, 2, x_4); +lean_ctor_set(x_6, 3, x_1); +lean_ctor_set(x_6, 4, x_1); +lean_ctor_set(x_6, 5, x_5); +return x_6; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_EMatchTheorem_getProofWithFreshMVarLevels___boxed(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_Grind_instInhabitedEMatchTheorem() { _start: { -lean_object* x_7; -x_7 = l_Lean_Meta_Grind_EMatchTheorem_getProofWithFreshMVarLevels(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 = l_Lean_Meta_Grind_instInhabitedEMatchTheorem___closed__3; +return x_1; } } -static lean_object* _init_l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____spec__1___closed__1() { +static lean_object* _init_l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedEMatchTheorems___spec__1___closed__1() { _start: { lean_object* x_1; @@ -2391,15772 +2438,27350 @@ 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_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____spec__1___closed__2() { +static lean_object* _init_l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedEMatchTheorems___spec__1___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____spec__1___closed__1; +x_1 = l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedEMatchTheorems___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_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____spec__1() { -_start: -{ -lean_object* x_1; -x_1 = l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____spec__1___closed__2; -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__1() { +static lean_object* _init_l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedEMatchTheorems___spec__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("_private", 8, 8); +x_1 = l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedEMatchTheorems___spec__1___closed__2; return x_1; } } -static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__2() { +static lean_object* _init_l_Lean_Meta_Grind_instInhabitedEMatchTheorems() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__1; -x_3 = l_Lean_Name_str___override(x_1, x_2); +x_1 = l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedEMatchTheorems___spec__1___closed__2; +x_2 = l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedEMatchTheorems___spec__1; +x_3 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +lean_ctor_set(x_3, 2, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__3() { +static lean_object* _init_l_panic___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__1___closed__1() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean", 4, 4); -return x_1; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Id_instMonad; +x_2 = l_Lean_Meta_Grind_instInhabitedEMatchTheorems; +x_3 = l_instInhabitedOfMonad___rarg(x_1, x_2); +return x_3; } } -static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__4() { +LEAN_EXPORT lean_object* l_panic___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__1(lean_object* x_1) { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__2; -x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__3; -x_3 = l_Lean_Name_str___override(x_1, x_2); +lean_object* x_2; lean_object* x_3; +x_2 = l_panic___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__1___closed__1; +x_3 = lean_panic_fn(x_2, x_1); return x_3; } } -static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__5() { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__4(size_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Meta", 4, 4); -return x_1; -} +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; } -static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__6() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__4; -x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__5; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; +lean_object* x_9; lean_object* x_10; lean_object* x_11; uint64_t x_12; size_t x_13; size_t x_14; size_t x_15; size_t x_16; size_t x_17; size_t x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_9 = lean_array_fget(x_2, x_5); +x_10 = lean_array_fget(x_3, x_5); +x_11 = l_Lean_Meta_Grind_Origin_key(x_9); +x_12 = l_Lean_Name_hash___override(x_11); +lean_dec(x_11); +x_13 = lean_uint64_to_usize(x_12); +x_14 = 1; +x_15 = lean_usize_sub(x_1, x_14); +x_16 = 5; +x_17 = lean_usize_mul(x_16, x_15); +x_18 = lean_usize_shift_right(x_13, x_17); +x_19 = lean_unsigned_to_nat(1u); +x_20 = lean_nat_add(x_5, x_19); +lean_dec(x_5); +x_21 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3(x_6, x_18, x_1, x_9, x_10); +x_4 = lean_box(0); +x_5 = x_20; +x_6 = x_21; +goto _start; } } -static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__7() { +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__5(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_unchecked("Tactic", 6, 6); +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; } -} -static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__8() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__6; -x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__7; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; +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; } } -static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__9() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Grind", 5, 5); -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__10() { -_start: +lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_17 = lean_array_fget(x_5, x_2); +x_18 = l_Lean_Meta_Grind_Origin_key(x_3); +x_19 = l_Lean_Meta_Grind_Origin_key(x_17); +lean_dec(x_17); +x_20 = lean_name_eq(x_18, x_19); +lean_dec(x_19); +lean_dec(x_18); +if (x_20 == 0) { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__8; -x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__9; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; -} +lean_object* x_21; lean_object* x_22; +lean_dec(x_6); +lean_dec(x_5); +x_21 = lean_unsigned_to_nat(1u); +x_22 = lean_nat_add(x_2, x_21); +lean_dec(x_2); +x_2 = x_22; +goto _start; } -static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__11() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("EMatchTheorem", 13, 13); +uint8_t x_24; +x_24 = !lean_is_exclusive(x_1); +if (x_24 == 0) +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_25 = lean_ctor_get(x_1, 1); +lean_dec(x_25); +x_26 = lean_ctor_get(x_1, 0); +lean_dec(x_26); +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); +lean_ctor_set(x_1, 1, x_28); +lean_ctor_set(x_1, 0, x_27); return x_1; } -} -static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__12() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__10; -x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__11; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; -} +lean_object* x_29; lean_object* x_30; lean_object* x_31; +lean_dec(x_1); +x_29 = lean_array_fset(x_5, x_2, x_3); +x_30 = lean_array_fset(x_6, x_2, x_4); +lean_dec(x_2); +x_31 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_31, 0, x_29); +lean_ctor_set(x_31, 1, x_30); +return x_31; } -static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__13() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__12; -x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Name_num___override(x_1, x_2); -return x_3; } } -static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__14() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__13; -x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__3; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; } } -static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__15() { +static size_t _init_l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__1() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__14; -x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__5; -x_3 = l_Lean_Name_str___override(x_1, x_2); +size_t x_1; size_t x_2; size_t x_3; +x_1 = 1; +x_2 = 5; +x_3 = lean_usize_shift_left(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__16() { +static size_t _init_l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__2() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__15; -x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__9; -x_3 = l_Lean_Name_str___override(x_1, x_2); +size_t x_1; size_t x_2; size_t x_3; +x_1 = 1; +x_2 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__1; +x_3 = lean_usize_sub(x_2, x_1); return x_3; } } -static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__17() { +static lean_object* _init_l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__3() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("ematchTheoremsExt", 17, 17); +x_1 = l_Lean_PersistentHashMap_mkEmptyEntries(lean_box(0), lean_box(0)); return x_1; } } -static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__18() { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3(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 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__16; -x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__17; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__19() { -_start: +if (lean_obj_tag(x_1) == 0) { -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_EMatchTheorems_insert), 2, 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_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__2; +x_11 = lean_usize_land(x_2, x_10); +x_12 = lean_usize_to_nat(x_11); +x_13 = lean_array_get_size(x_7); +x_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; } -} -static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__20() { -_start: +else { -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_id___rarg___boxed), 1, 0); +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_array_fget(x_7, x_12); +x_16 = lean_box(0); +x_17 = lean_array_fset(x_7, x_12, x_16); +switch (lean_obj_tag(x_15)) { +case 0: +{ +uint8_t x_18; +x_18 = !lean_is_exclusive(x_15); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; +x_19 = lean_ctor_get(x_15, 0); +x_20 = lean_ctor_get(x_15, 1); +x_21 = l_Lean_Meta_Grind_Origin_key(x_4); +x_22 = l_Lean_Meta_Grind_Origin_key(x_19); +x_23 = lean_name_eq(x_21, x_22); +lean_dec(x_22); +lean_dec(x_21); +if (x_23 == 0) +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +lean_free_object(x_15); +x_24 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_19, x_20, x_4, x_5); +x_25 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_25, 0, x_24); +x_26 = lean_array_fset(x_17, x_12, x_25); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_26); return x_1; } -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687_(lean_object* x_1) { -_start: +else { -lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__18; -x_3 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__19; -x_4 = l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____spec__1; -x_5 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__20; -x_6 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_6, 0, x_2); -lean_ctor_set(x_6, 1, x_3); -lean_ctor_set(x_6, 2, x_4); -lean_ctor_set(x_6, 3, x_5); -x_7 = l_Lean_registerSimpleScopedEnvExtension___rarg(x_6, x_1); -return x_7; +lean_object* x_27; +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_27 = lean_array_fset(x_17, x_12, x_15); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_27); +return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__1() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Eq", 2, 2); +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_15, 0); +x_29 = lean_ctor_get(x_15, 1); +lean_inc(x_29); +lean_inc(x_28); +lean_dec(x_15); +x_30 = l_Lean_Meta_Grind_Origin_key(x_4); +x_31 = l_Lean_Meta_Grind_Origin_key(x_28); +x_32 = lean_name_eq(x_30, x_31); +lean_dec(x_31); +lean_dec(x_30); +if (x_32 == 0) +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_33 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_28, x_29, x_4, x_5); +x_34 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_34, 0, x_33); +x_35 = lean_array_fset(x_17, x_12, x_34); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_35); return x_1; } -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___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_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__1; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; +lean_object* x_36; lean_object* x_37; +lean_dec(x_29); +lean_dec(x_28); +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_4); +lean_ctor_set(x_36, 1, x_5); +x_37 = lean_array_fset(x_17, x_12, x_36); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_37); +return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__3() { -_start: +} +case 1: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("HEq", 3, 3); +uint8_t x_38; +x_38 = !lean_is_exclusive(x_15); +if (x_38 == 0) +{ +lean_object* x_39; size_t x_40; size_t x_41; lean_object* x_42; lean_object* x_43; +x_39 = lean_ctor_get(x_15, 0); +x_40 = lean_usize_shift_right(x_2, x_9); +x_41 = lean_usize_add(x_3, x_8); +x_42 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3(x_39, x_40, x_41, x_4, x_5); +lean_ctor_set(x_15, 0, x_42); +x_43 = lean_array_fset(x_17, x_12, x_15); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_43); return x_1; } -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__4() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__3; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; +lean_object* x_44; size_t x_45; size_t x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_44 = lean_ctor_get(x_15, 0); +lean_inc(x_44); +lean_dec(x_15); +x_45 = lean_usize_shift_right(x_2, x_9); +x_46 = lean_usize_add(x_3, x_8); +x_47 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3(x_44, x_45, x_46, x_4, x_5); +x_48 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_48, 0, x_47); +x_49 = lean_array_fset(x_17, x_12, x_48); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_49); +return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__5() { -_start: +default: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Iff", 3, 3); +lean_object* x_50; lean_object* x_51; +x_50 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_50, 0, x_4); +lean_ctor_set(x_50, 1, x_5); +x_51 = lean_array_fset(x_17, x_12, x_50); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_51); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___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_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__5; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__7() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("And", 3, 3); -return x_1; -} +lean_object* x_52; size_t x_53; size_t x_54; size_t x_55; size_t x_56; lean_object* x_57; lean_object* x_58; uint8_t x_59; +x_52 = lean_ctor_get(x_1, 0); +lean_inc(x_52); +lean_dec(x_1); +x_53 = 1; +x_54 = 5; +x_55 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__2; +x_56 = lean_usize_land(x_2, x_55); +x_57 = lean_usize_to_nat(x_56); +x_58 = lean_array_get_size(x_52); +x_59 = lean_nat_dec_lt(x_57, x_58); +lean_dec(x_58); +if (x_59 == 0) +{ +lean_object* x_60; +lean_dec(x_57); +lean_dec(x_5); +lean_dec(x_4); +x_60 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_60, 0, x_52); +return x_60; } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__8() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__7; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; +lean_object* x_61; lean_object* x_62; lean_object* x_63; +x_61 = lean_array_fget(x_52, x_57); +x_62 = lean_box(0); +x_63 = lean_array_fset(x_52, x_57, x_62); +switch (lean_obj_tag(x_61)) { +case 0: +{ +lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; uint8_t x_69; +x_64 = lean_ctor_get(x_61, 0); +lean_inc(x_64); +x_65 = lean_ctor_get(x_61, 1); +lean_inc(x_65); +if (lean_is_exclusive(x_61)) { + lean_ctor_release(x_61, 0); + lean_ctor_release(x_61, 1); + x_66 = x_61; +} else { + lean_dec_ref(x_61); + x_66 = lean_box(0); } +x_67 = l_Lean_Meta_Grind_Origin_key(x_4); +x_68 = l_Lean_Meta_Grind_Origin_key(x_64); +x_69 = lean_name_eq(x_67, x_68); +lean_dec(x_68); +lean_dec(x_67); +if (x_69 == 0) +{ +lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; +lean_dec(x_66); +x_70 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_64, x_65, x_4, x_5); +x_71 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_71, 0, x_70); +x_72 = lean_array_fset(x_63, x_57, x_71); +lean_dec(x_57); +x_73 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_73, 0, x_72); +return x_73; } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__9() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Or", 2, 2); -return x_1; +lean_object* x_74; lean_object* x_75; lean_object* x_76; +lean_dec(x_65); +lean_dec(x_64); +if (lean_is_scalar(x_66)) { + x_74 = lean_alloc_ctor(0, 2, 0); +} else { + x_74 = x_66; } +lean_ctor_set(x_74, 0, x_4); +lean_ctor_set(x_74, 1, x_5); +x_75 = lean_array_fset(x_63, x_57, x_74); +lean_dec(x_57); +x_76 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_76, 0, x_75); +return x_76; } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__10() { -_start: +} +case 1: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__9; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; +lean_object* x_77; lean_object* x_78; size_t x_79; size_t x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; +x_77 = lean_ctor_get(x_61, 0); +lean_inc(x_77); +if (lean_is_exclusive(x_61)) { + lean_ctor_release(x_61, 0); + x_78 = x_61; +} else { + lean_dec_ref(x_61); + x_78 = lean_box(0); +} +x_79 = lean_usize_shift_right(x_2, x_54); +x_80 = lean_usize_add(x_3, x_53); +x_81 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3(x_77, x_79, x_80, x_4, x_5); +if (lean_is_scalar(x_78)) { + x_82 = lean_alloc_ctor(1, 1, 0); +} else { + x_82 = x_78; } +lean_ctor_set(x_82, 0, x_81); +x_83 = lean_array_fset(x_63, x_57, x_82); +lean_dec(x_57); +x_84 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_84, 0, x_83); +return x_84; } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__11() { -_start: +default: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Not", 3, 3); -return x_1; +lean_object* x_85; lean_object* x_86; lean_object* x_87; +x_85 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_85, 0, x_4); +lean_ctor_set(x_85, 1, x_5); +x_86 = lean_array_fset(x_63, x_57, x_85); +lean_dec(x_57); +x_87 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_87, 0, x_86); +return x_87; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___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_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__11; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___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_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__12; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set(x_3, 1, x_1); -return x_3; } +else +{ +uint8_t x_88; +x_88 = !lean_is_exclusive(x_1); +if (x_88 == 0) +{ +lean_object* x_89; lean_object* x_90; size_t x_91; uint8_t x_92; +x_89 = lean_unsigned_to_nat(0u); +x_90 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__5(x_1, x_89, x_4, x_5); +x_91 = 7; +x_92 = lean_usize_dec_le(x_91, x_3); +if (x_92 == 0) +{ +lean_object* x_93; lean_object* x_94; uint8_t x_95; +x_93 = l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(x_90); +x_94 = lean_unsigned_to_nat(4u); +x_95 = lean_nat_dec_lt(x_93, x_94); +lean_dec(x_93); +if (x_95 == 0) +{ +lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; +x_96 = lean_ctor_get(x_90, 0); +lean_inc(x_96); +x_97 = lean_ctor_get(x_90, 1); +lean_inc(x_97); +lean_dec(x_90); +x_98 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__3; +x_99 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__4(x_3, x_96, x_97, lean_box(0), x_89, x_98); +lean_dec(x_97); +lean_dec(x_96); +return x_99; } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__14() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__10; -x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__13; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set(x_3, 1, x_2); -return x_3; +return x_90; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__15() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__8; -x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__14; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set(x_3, 1, x_2); -return x_3; +return x_90; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__16() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__6; -x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__15; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set(x_3, 1, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__17() { -_start: +lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; size_t x_105; uint8_t x_106; +x_100 = lean_ctor_get(x_1, 0); +x_101 = lean_ctor_get(x_1, 1); +lean_inc(x_101); +lean_inc(x_100); +lean_dec(x_1); +x_102 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_102, 0, x_100); +lean_ctor_set(x_102, 1, x_101); +x_103 = lean_unsigned_to_nat(0u); +x_104 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__5(x_102, x_103, x_4, x_5); +x_105 = 7; +x_106 = lean_usize_dec_le(x_105, x_3); +if (x_106 == 0) +{ +lean_object* x_107; lean_object* x_108; uint8_t x_109; +x_107 = l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(x_104); +x_108 = lean_unsigned_to_nat(4u); +x_109 = lean_nat_dec_lt(x_107, x_108); +lean_dec(x_107); +if (x_109 == 0) { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__4; -x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__16; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set(x_3, 1, x_2); -return x_3; -} +lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; +x_110 = lean_ctor_get(x_104, 0); +lean_inc(x_110); +x_111 = lean_ctor_get(x_104, 1); +lean_inc(x_111); +lean_dec(x_104); +x_112 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__3; +x_113 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__4(x_3, x_110, x_111, lean_box(0), x_103, x_112); +lean_dec(x_111); +lean_dec(x_110); +return x_113; } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__18() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__2; -x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__17; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set(x_3, 1, x_2); -return x_3; +return x_104; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__19() { -_start: +else { -lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__18; -x_2 = lean_array_mk(x_1); -return x_2; -} +return x_104; } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames() { -_start: -{ -lean_object* x_1; -x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__19; -return x_1; } } -LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isForbidden(lean_object* x_1) { -_start: -{ -lean_object* x_2; uint8_t x_3; -x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames; -x_3 = l_Array_contains___at_Lean_registerInternalExceptionId___spec__1(x_2, x_1); -return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isForbidden___boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -uint8_t x_2; lean_object* x_3; -x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isForbidden(x_1); -lean_dec(x_1); -x_3 = lean_box(x_2); -return x_3; +lean_object* x_4; uint64_t x_5; size_t x_6; size_t x_7; lean_object* x_8; +x_4 = l_Lean_Meta_Grind_Origin_key(x_2); +x_5 = l_Lean_Name_hash___override(x_4); +lean_dec(x_4); +x_6 = lean_uint64_to_usize(x_5); +x_7 = 1; +x_8 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3(x_1, x_6, x_7, x_2, x_3); +return x_8; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare___closed__1() { +LEAN_EXPORT lean_object* l_Array_indexOfAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__8(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("[grind_dontcare]", 16, 16); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare___closed__2() { -_start: +lean_object* x_4; uint8_t x_5; +x_4 = lean_array_get_size(x_1); +x_5 = lean_nat_dec_lt(x_3, x_4); +lean_dec(x_4); +if (x_5 == 0) { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare___closed__1; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; +lean_object* x_6; +lean_dec(x_3); +x_6 = lean_box(0); +return x_6; } +else +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; +x_7 = lean_array_fget(x_1, x_3); +x_8 = l_Lean_Meta_Grind_Origin_key(x_7); +lean_dec(x_7); +x_9 = l_Lean_Meta_Grind_Origin_key(x_2); +x_10 = lean_name_eq(x_8, x_9); +lean_dec(x_9); +lean_dec(x_8); +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_3, x_11); +lean_dec(x_3); +x_3 = x_12; +goto _start; } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare___closed__3() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare___closed__2; -x_3 = l_Lean_Expr_const___override(x_2, x_1); -return x_3; +lean_object* x_14; +x_14 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_14, 0, x_3); +return x_14; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare() { -_start: -{ -lean_object* x_1; -x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare___closed__3; -return x_1; } } -static lean_object* _init_l_Lean_Meta_Grind_mkGroundPattern___closed__1() { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__7(lean_object* x_1, size_t x_2, lean_object* x_3) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("grind", 5, 5); +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_4; size_t x_5; size_t x_6; size_t x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_4 = lean_ctor_get(x_1, 0); +lean_inc(x_4); +x_5 = 5; +x_6 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___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); +switch (lean_obj_tag(x_10)) { +case 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); +lean_dec(x_10); +x_12 = l_Lean_Meta_Grind_Origin_key(x_3); +x_13 = l_Lean_Meta_Grind_Origin_key(x_11); +lean_dec(x_11); +x_14 = lean_name_eq(x_12, x_13); +lean_dec(x_13); +lean_dec(x_12); +if (x_14 == 0) +{ +lean_dec(x_8); +lean_dec(x_4); return x_1; } -} -static lean_object* _init_l_Lean_Meta_Grind_mkGroundPattern___closed__2() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("ground_pat", 10, 10); +uint8_t x_15; +x_15 = !lean_is_exclusive(x_1); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; +x_16 = lean_ctor_get(x_1, 0); +lean_dec(x_16); +x_17 = lean_array_set(x_4, x_8, x_9); +lean_dec(x_8); +lean_ctor_set(x_1, 0, x_17); return x_1; } -} -static lean_object* _init_l_Lean_Meta_Grind_mkGroundPattern___closed__3() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_Grind_mkGroundPattern___closed__1; -x_2 = l_Lean_Meta_Grind_mkGroundPattern___closed__2; -x_3 = l_Lean_Name_mkStr2(x_1, x_2); -return x_3; -} +lean_object* x_18; lean_object* x_19; +lean_dec(x_1); +x_18 = lean_array_set(x_4, x_8, x_9); +lean_dec(x_8); +x_19 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_19, 0, x_18); +return x_19; } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkGroundPattern(lean_object* x_1) { -_start: -{ -lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_Meta_Grind_mkGroundPattern___closed__3; -x_3 = l_Lean_mkAnnotation(x_2, x_1); -return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_groundPattern_x3f(lean_object* x_1) { -_start: +case 1: { -lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_Meta_Grind_mkGroundPattern___closed__3; -x_3 = l_Lean_annotation_x3f(x_2, x_1); -return x_3; +uint8_t x_20; +x_20 = !lean_is_exclusive(x_1); +if (x_20 == 0) +{ +lean_object* x_21; uint8_t x_22; +x_21 = lean_ctor_get(x_1, 0); +lean_dec(x_21); +x_22 = !lean_is_exclusive(x_10); +if (x_22 == 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_10, 0); +x_24 = lean_array_set(x_4, x_8, x_9); +x_25 = lean_usize_shift_right(x_2, x_5); +x_26 = l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__7(x_23, x_25, x_3); +lean_inc(x_26); +x_27 = l_Lean_PersistentHashMap_isUnaryNode___rarg(x_26); +if (lean_obj_tag(x_27) == 0) +{ +lean_object* x_28; +lean_ctor_set(x_10, 0, x_26); +x_28 = lean_array_set(x_24, x_8, x_10); +lean_dec(x_8); +lean_ctor_set(x_1, 0, x_28); +return x_1; } +else +{ +lean_object* x_29; uint8_t x_30; +lean_dec(x_26); +lean_free_object(x_10); +x_29 = lean_ctor_get(x_27, 0); +lean_inc(x_29); +lean_dec(x_27); +x_30 = !lean_is_exclusive(x_29); +if (x_30 == 0) +{ +lean_object* x_31; +x_31 = lean_array_set(x_24, x_8, x_29); +lean_dec(x_8); +lean_ctor_set(x_1, 0, x_31); +return x_1; } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_groundPattern_x3f___boxed(lean_object* x_1) { -_start: +else { -lean_object* x_2; -x_2 = l_Lean_Meta_Grind_groundPattern_x3f(x_1); -lean_dec(x_1); -return x_2; +lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_32 = lean_ctor_get(x_29, 0); +x_33 = lean_ctor_get(x_29, 1); +lean_inc(x_33); +lean_inc(x_32); +lean_dec(x_29); +x_34 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +x_35 = lean_array_set(x_24, x_8, x_34); +lean_dec(x_8); +lean_ctor_set(x_1, 0, x_35); +return x_1; } } -LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isGroundPattern(lean_object* x_1) { -_start: +} +else { -lean_object* x_2; -x_2 = l_Lean_Meta_Grind_groundPattern_x3f(x_1); -if (lean_obj_tag(x_2) == 0) +lean_object* x_36; lean_object* x_37; size_t x_38; lean_object* x_39; lean_object* x_40; +x_36 = lean_ctor_get(x_10, 0); +lean_inc(x_36); +lean_dec(x_10); +x_37 = lean_array_set(x_4, x_8, x_9); +x_38 = lean_usize_shift_right(x_2, x_5); +x_39 = l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__7(x_36, x_38, x_3); +lean_inc(x_39); +x_40 = l_Lean_PersistentHashMap_isUnaryNode___rarg(x_39); +if (lean_obj_tag(x_40) == 0) { -uint8_t x_3; -x_3 = 0; -return x_3; +lean_object* x_41; lean_object* x_42; +x_41 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_41, 0, x_39); +x_42 = lean_array_set(x_37, x_8, x_41); +lean_dec(x_8); +lean_ctor_set(x_1, 0, x_42); +return x_1; } else { -uint8_t x_4; -lean_dec(x_2); -x_4 = 1; -return x_4; +lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; +lean_dec(x_39); +x_43 = lean_ctor_get(x_40, 0); +lean_inc(x_43); +lean_dec(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); +if (lean_is_exclusive(x_43)) { + lean_ctor_release(x_43, 0); + lean_ctor_release(x_43, 1); + x_46 = x_43; +} else { + lean_dec_ref(x_43); + x_46 = lean_box(0); +} +if (lean_is_scalar(x_46)) { + x_47 = lean_alloc_ctor(0, 2, 0); +} else { + x_47 = x_46; +} +lean_ctor_set(x_47, 0, x_44); +lean_ctor_set(x_47, 1, x_45); +x_48 = lean_array_set(x_37, x_8, x_47); +lean_dec(x_8); +lean_ctor_set(x_1, 0, x_48); +return x_1; } } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isGroundPattern___boxed(lean_object* x_1) { -_start: +else { -uint8_t x_2; lean_object* x_3; -x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isGroundPattern(x_1); +lean_object* x_49; lean_object* x_50; lean_object* x_51; size_t x_52; lean_object* x_53; lean_object* x_54; lean_dec(x_1); -x_3 = lean_box(x_2); -return x_3; -} +x_49 = lean_ctor_get(x_10, 0); +lean_inc(x_49); +if (lean_is_exclusive(x_10)) { + lean_ctor_release(x_10, 0); + x_50 = x_10; +} else { + lean_dec_ref(x_10); + x_50 = lean_box(0); } -LEAN_EXPORT uint8_t l_Lean_Meta_Grind_isPatternDontCare(lean_object* x_1) { -_start: +x_51 = lean_array_set(x_4, x_8, x_9); +x_52 = lean_usize_shift_right(x_2, x_5); +x_53 = l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__7(x_49, x_52, x_3); +lean_inc(x_53); +x_54 = l_Lean_PersistentHashMap_isUnaryNode___rarg(x_53); +if (lean_obj_tag(x_54) == 0) { -lean_object* x_2; uint8_t x_3; -x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare; -x_3 = lean_expr_eqv(x_1, x_2); -return x_3; +lean_object* x_55; lean_object* x_56; lean_object* x_57; +if (lean_is_scalar(x_50)) { + x_55 = lean_alloc_ctor(1, 1, 0); +} else { + x_55 = x_50; } +lean_ctor_set(x_55, 0, x_53); +x_56 = lean_array_set(x_51, x_8, x_55); +lean_dec(x_8); +x_57 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_57, 0, x_56); +return x_57; } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isPatternDontCare___boxed(lean_object* x_1) { -_start: +else { -uint8_t x_2; lean_object* x_3; -x_2 = l_Lean_Meta_Grind_isPatternDontCare(x_1); -lean_dec(x_1); -x_3 = lean_box(x_2); -return x_3; +lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; +lean_dec(x_53); +lean_dec(x_50); +x_58 = lean_ctor_get(x_54, 0); +lean_inc(x_58); +lean_dec(x_54); +x_59 = lean_ctor_get(x_58, 0); +lean_inc(x_59); +x_60 = lean_ctor_get(x_58, 1); +lean_inc(x_60); +if (lean_is_exclusive(x_58)) { + lean_ctor_release(x_58, 0); + lean_ctor_release(x_58, 1); + x_61 = x_58; +} else { + lean_dec_ref(x_58); + x_61 = lean_box(0); } +if (lean_is_scalar(x_61)) { + x_62 = lean_alloc_ctor(0, 2, 0); +} else { + x_62 = x_61; } -LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isAtomicPattern(lean_object* x_1) { -_start: -{ -uint8_t x_2; -x_2 = l_Lean_Expr_isBVar(x_1); -if (x_2 == 0) -{ -lean_object* x_3; uint8_t x_4; -x_3 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare; -x_4 = lean_expr_eqv(x_1, x_3); -if (x_4 == 0) -{ -uint8_t x_5; -x_5 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isGroundPattern(x_1); -return x_5; +lean_ctor_set(x_62, 0, x_59); +lean_ctor_set(x_62, 1, x_60); +x_63 = lean_array_set(x_51, x_8, x_62); +lean_dec(x_8); +x_64 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_64, 0, x_63); +return x_64; } -else -{ -uint8_t x_6; -x_6 = 1; -return x_6; } } -else +default: { -uint8_t x_7; -x_7 = 1; -return x_7; +lean_dec(x_8); +lean_dec(x_4); +return x_1; } } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isAtomicPattern___boxed(lean_object* x_1) { -_start: +else { -uint8_t x_2; lean_object* x_3; -x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isAtomicPattern(x_1); -lean_dec(x_1); -x_3 = lean_box(x_2); -return x_3; -} -} -static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___closed__1() { -_start: +lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; +x_65 = lean_ctor_get(x_1, 0); +lean_inc(x_65); +x_66 = lean_ctor_get(x_1, 1); +lean_inc(x_66); +x_67 = lean_unsigned_to_nat(0u); +x_68 = l_Array_indexOfAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__8(x_65, x_3, x_67); +if (lean_obj_tag(x_68) == 0) { -lean_object* x_1; -x_1 = lean_mk_string_unchecked(" ", 1, 1); +lean_dec(x_66); +lean_dec(x_65); return x_1; } -} -static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___closed__2() { -_start: +else { -lean_object* x_1; lean_object* x_2; -x_1 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___closed__1; -x_2 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; -} -} -static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___closed__3() { -_start: +uint8_t x_69; +x_69 = !lean_is_exclusive(x_1); +if (x_69 == 0) { -lean_object* x_1; lean_object* x_2; -x_1 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___closed__2; -x_2 = l_Lean_MessageData_ofFormat(x_1); -return x_2; -} +lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; +x_70 = lean_ctor_get(x_1, 1); +lean_dec(x_70); +x_71 = lean_ctor_get(x_1, 0); +lean_dec(x_71); +x_72 = lean_ctor_get(x_68, 0); +lean_inc(x_72); +lean_dec(x_68); +lean_inc(x_72); +x_73 = l_Array_eraseIdx___rarg(x_65, x_72, lean_box(0)); +x_74 = l_Array_eraseIdx___rarg(x_66, x_72, lean_box(0)); +lean_ctor_set(x_1, 1, x_74); +lean_ctor_set(x_1, 0, x_73); +return x_1; } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1(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; -x_4 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___closed__3; -x_5 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_5, 0, x_2); -lean_ctor_set(x_5, 1, x_4); -x_6 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_6, 0, x_5); -lean_ctor_set(x_6, 1, x_1); -x_7 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_7, 0, x_6); -return x_7; -} +lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; +lean_dec(x_1); +x_75 = lean_ctor_get(x_68, 0); +lean_inc(x_75); +lean_dec(x_68); +lean_inc(x_75); +x_76 = l_Array_eraseIdx___rarg(x_65, x_75, lean_box(0)); +x_77 = l_Array_eraseIdx___rarg(x_66, x_75, lean_box(0)); +x_78 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_78, 0, x_76); +lean_ctor_set(x_78, 1, x_77); +return x_78; } -static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___boxed), 3, 0); -return x_1; } } -static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("(", 1, 1); -return x_1; } } -static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3() { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_erase___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__6(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked(")", 1, 1); -return x_1; +lean_object* x_3; uint64_t x_4; size_t x_5; lean_object* x_6; +x_3 = l_Lean_Meta_Grind_Origin_key(x_2); +x_4 = l_Lean_Name_hash___override(x_3); +lean_dec(x_3); +x_5 = lean_uint64_to_usize(x_4); +x_6 = l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__7(x_1, x_5, x_2); +return x_6; } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___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_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__11(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_5, x_4); +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) { -return x_6; +lean_object* x_8; +lean_dec(x_4); +x_8 = lean_box(0); +return x_8; } else { -lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; -x_8 = lean_array_uget(x_3, x_5); -lean_inc(x_8); -x_9 = l_Lean_Meta_Grind_ppPattern(x_8); -x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1; -x_11 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isAtomicPattern(x_8); -lean_dec(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 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2; -x_13 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3; -x_14 = l_Lean_MessageData_bracket(x_12, x_9, x_13); -x_15 = lean_box(0); -x_16 = lean_apply_3(x_10, x_14, x_6, x_15); -if (lean_obj_tag(x_16) == 0) -{ -lean_object* x_17; -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -lean_dec(x_16); -return x_17; -} -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_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_5, x_19); -x_5 = x_20; -x_6 = x_18; +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_22; lean_object* x_23; -x_22 = lean_box(0); -x_23 = lean_apply_3(x_10, x_9, x_6, x_22); -if (lean_obj_tag(x_23) == 0) -{ -lean_object* x_24; -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -lean_dec(x_23); -return x_24; -} else { -lean_object* x_25; size_t x_26; size_t x_27; -x_25 = lean_ctor_get(x_23, 0); -lean_inc(x_25); -lean_dec(x_23); -x_26 = 1; -x_27 = lean_usize_add(x_5, x_26); -x_5 = x_27; -x_6 = x_25; -goto _start; -} +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; } } } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__10(lean_object* x_1, size_t x_2, lean_object* x_3) { _start: { -uint8_t x_7; -x_7 = lean_usize_dec_lt(x_5, x_4); -if (x_7 == 0) -{ -return x_6; -} -else +if (lean_obj_tag(x_1) == 0) { -lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; -x_8 = lean_array_uget(x_3, x_5); -lean_inc(x_8); -x_9 = l_Lean_Meta_Grind_ppPattern(x_8); -x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1; -x_11 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isAtomicPattern(x_8); -lean_dec(x_8); -if (x_11 == 0) +uint8_t x_4; +x_4 = !lean_is_exclusive(x_1); +if (x_4 == 0) { -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_12 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2; -x_13 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3; -x_14 = l_Lean_MessageData_bracket(x_12, x_9, x_13); -x_15 = lean_box(0); -x_16 = lean_apply_3(x_10, x_14, x_6, x_15); -if (lean_obj_tag(x_16) == 0) +lean_object* x_5; size_t x_6; size_t x_7; size_t x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_5 = lean_ctor_get(x_1, 0); +x_6 = 5; +x_7 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__2; +x_8 = lean_usize_land(x_2, x_7); +x_9 = lean_usize_to_nat(x_8); +x_10 = lean_box(2); +x_11 = lean_array_get(x_10, x_5, x_9); +lean_dec(x_9); +lean_dec(x_5); +switch (lean_obj_tag(x_11)) { +case 0: { -lean_object* x_17; -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -lean_dec(x_16); -return x_17; -} -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_name_eq(x_3, x_12); +lean_dec(x_12); +if (x_14 == 0) { -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_5, x_19); -x_5 = x_20; -x_6 = x_18; -goto _start; -} +lean_object* x_15; +lean_dec(x_13); +lean_free_object(x_1); +x_15 = lean_box(0); +return x_15; } else { -lean_object* x_22; lean_object* x_23; -x_22 = lean_box(0); -x_23 = lean_apply_3(x_10, x_9, x_6, x_22); -if (lean_obj_tag(x_23) == 0) -{ -lean_object* x_24; -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -lean_dec(x_23); -return x_24; +lean_ctor_set_tag(x_1, 1); +lean_ctor_set(x_1, 0, x_13); +return x_1; } -else +} +case 1: { -lean_object* x_25; size_t x_26; size_t x_27; -x_25 = lean_ctor_get(x_23, 0); -lean_inc(x_25); -lean_dec(x_23); -x_26 = 1; -x_27 = lean_usize_add(x_5, x_26); -x_5 = x_27; -x_6 = x_25; +lean_object* x_16; size_t x_17; +lean_free_object(x_1); +x_16 = lean_ctor_get(x_11, 0); +lean_inc(x_16); +lean_dec(x_11); +x_17 = lean_usize_shift_right(x_2, x_6); +x_1 = x_16; +x_2 = x_17; goto _start; } +default: +{ +lean_object* x_19; +lean_free_object(x_1); +x_19 = lean_box(0); +return x_19; } } } -} -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { -_start: -{ -uint8_t x_7; -x_7 = lean_usize_dec_lt(x_5, x_4); -if (x_7 == 0) -{ -return x_6; -} else { -lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; -x_8 = lean_array_uget(x_3, x_5); -lean_inc(x_8); -x_9 = l_Lean_Meta_Grind_ppPattern(x_8); -x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1; -x_11 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isAtomicPattern(x_8); -lean_dec(x_8); -if (x_11 == 0) +lean_object* x_20; size_t x_21; size_t x_22; size_t x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_20 = lean_ctor_get(x_1, 0); +lean_inc(x_20); +lean_dec(x_1); +x_21 = 5; +x_22 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__2; +x_23 = lean_usize_land(x_2, x_22); +x_24 = lean_usize_to_nat(x_23); +x_25 = lean_box(2); +x_26 = lean_array_get(x_25, x_20, x_24); +lean_dec(x_24); +lean_dec(x_20); +switch (lean_obj_tag(x_26)) { +case 0: { -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_12 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2; -x_13 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3; -x_14 = l_Lean_MessageData_bracket(x_12, x_9, x_13); -x_15 = lean_box(0); -x_16 = lean_apply_3(x_10, x_14, x_6, x_15); -if (lean_obj_tag(x_16) == 0) +lean_object* x_27; lean_object* x_28; uint8_t x_29; +x_27 = lean_ctor_get(x_26, 0); +lean_inc(x_27); +x_28 = lean_ctor_get(x_26, 1); +lean_inc(x_28); +lean_dec(x_26); +x_29 = lean_name_eq(x_3, x_27); +lean_dec(x_27); +if (x_29 == 0) { -lean_object* x_17; -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -lean_dec(x_16); -return x_17; +lean_object* x_30; +lean_dec(x_28); +x_30 = lean_box(0); +return x_30; } else { -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_5, x_19); -x_5 = x_20; -x_6 = x_18; -goto _start; +lean_object* x_31; +x_31 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_31, 0, x_28); +return x_31; } } -else +case 1: { -lean_object* x_22; lean_object* x_23; -x_22 = lean_box(0); -x_23 = lean_apply_3(x_10, x_9, x_6, x_22); -if (lean_obj_tag(x_23) == 0) +lean_object* x_32; size_t x_33; +x_32 = lean_ctor_get(x_26, 0); +lean_inc(x_32); +lean_dec(x_26); +x_33 = lean_usize_shift_right(x_2, x_21); +x_1 = x_32; +x_2 = x_33; +goto _start; +} +default: { -lean_object* x_24; -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -lean_dec(x_23); -return x_24; +lean_object* x_35; +x_35 = lean_box(0); +return x_35; +} +} +} } else { -lean_object* x_25; size_t x_26; size_t x_27; -x_25 = lean_ctor_get(x_23, 0); -lean_inc(x_25); -lean_dec(x_23); -x_26 = 1; -x_27 = lean_usize_add(x_5, x_26); -x_5 = x_27; -x_6 = x_25; -goto _start; +lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_36 = lean_ctor_get(x_1, 0); +lean_inc(x_36); +x_37 = lean_ctor_get(x_1, 1); +lean_inc(x_37); +lean_dec(x_1); +x_38 = lean_unsigned_to_nat(0u); +x_39 = l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__11(x_36, x_37, lean_box(0), x_38, x_3); +lean_dec(x_37); +lean_dec(x_36); +return x_39; } } } +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__9(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint64_t x_3; size_t x_4; lean_object* x_5; +x_3 = l_Lean_Name_hash___override(x_2); +x_4 = lean_uint64_to_usize(x_3); +x_5 = l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__10(x_1, x_4, x_2); +return x_5; } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__14(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: { -uint8_t x_7; -x_7 = lean_usize_dec_lt(x_5, x_4); -if (x_7 == 0) +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_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; -x_8 = lean_array_uget(x_3, x_5); -lean_inc(x_8); -x_9 = l_Lean_Meta_Grind_ppPattern(x_8); -x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1; -x_11 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isAtomicPattern(x_8); -lean_dec(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 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2; -x_13 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3; -x_14 = l_Lean_MessageData_bracket(x_12, x_9, x_13); -x_15 = lean_box(0); -x_16 = lean_apply_3(x_10, x_14, x_6, x_15); -if (lean_obj_tag(x_16) == 0) -{ -lean_object* x_17; -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -lean_dec(x_16); -return x_17; +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_Grind_EMatchTheorems_insert___spec__13(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_Grind_EMatchTheorems_insert___spec__15(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; 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_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_5, x_19); -x_5 = x_20; -x_6 = x_18; -goto _start; +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_22; lean_object* x_23; -x_22 = lean_box(0); -x_23 = lean_apply_3(x_10, x_9, x_6, x_22); -if (lean_obj_tag(x_23) == 0) +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_24; -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -lean_dec(x_23); -return x_24; +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 { -lean_object* x_25; size_t x_26; size_t x_27; -x_25 = lean_ctor_get(x_23, 0); -lean_inc(x_25); +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_26 = 1; -x_27 = lean_usize_add(x_5, x_26); -x_5 = x_27; -x_6 = x_25; -goto _start; +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; } } } } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__13(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5) { _start: { -uint8_t x_7; -x_7 = lean_usize_dec_lt(x_5, x_4); -if (x_7 == 0) +if (lean_obj_tag(x_1) == 0) { -return x_6; +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_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__2; +x_11 = lean_usize_land(x_2, x_10); +x_12 = lean_usize_to_nat(x_11); +x_13 = lean_array_get_size(x_7); +x_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_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; -x_8 = lean_array_uget(x_3, x_5); -lean_inc(x_8); -x_9 = l_Lean_Meta_Grind_ppPattern(x_8); -x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1; -x_11 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isAtomicPattern(x_8); -lean_dec(x_8); -if (x_11 == 0) +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_array_fget(x_7, x_12); +x_16 = lean_box(0); +x_17 = lean_array_fset(x_7, x_12, x_16); +switch (lean_obj_tag(x_15)) { +case 0: { -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_12 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2; -x_13 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3; -x_14 = l_Lean_MessageData_bracket(x_12, x_9, x_13); -x_15 = lean_box(0); -x_16 = lean_apply_3(x_10, x_14, x_6, x_15); -if (lean_obj_tag(x_16) == 0) +uint8_t x_18; +x_18 = !lean_is_exclusive(x_15); +if (x_18 == 0) { -lean_object* x_17; -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -lean_dec(x_16); -return x_17; +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_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_5, x_19); -x_5 = x_20; -x_6 = x_18; -goto _start; +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_22; lean_object* x_23; -x_22 = lean_box(0); -x_23 = lean_apply_3(x_10, x_9, x_6, x_22); -if (lean_obj_tag(x_23) == 0) +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_24; -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -lean_dec(x_23); -return x_24; +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_25; size_t x_26; size_t x_27; -x_25 = lean_ctor_get(x_23, 0); -lean_inc(x_25); -lean_dec(x_23); -x_26 = 1; -x_27 = lean_usize_add(x_5, x_26); -x_5 = x_27; -x_6 = x_25; -goto _start; -} -} +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; } } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { -_start: +case 1: { -uint8_t x_7; -x_7 = lean_usize_dec_lt(x_5, x_4); -if (x_7 == 0) +uint8_t x_34; +x_34 = !lean_is_exclusive(x_15); +if (x_34 == 0) { -return x_6; +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_Grind_EMatchTheorems_insert___spec__13(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_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; -x_8 = lean_array_uget(x_3, x_5); -lean_inc(x_8); -x_9 = l_Lean_Meta_Grind_ppPattern(x_8); -x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1; -x_11 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isAtomicPattern(x_8); -lean_dec(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 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2; -x_13 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3; -x_14 = l_Lean_MessageData_bracket(x_12, x_9, x_13); -x_15 = lean_box(0); -x_16 = lean_apply_3(x_10, x_14, x_6, x_15); -if (lean_obj_tag(x_16) == 0) -{ -lean_object* x_17; -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -lean_dec(x_16); -return x_17; +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_Grind_EMatchTheorems_insert___spec__13(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; } -else +} +default: { -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_5, x_19); -x_5 = x_20; -x_6 = x_18; -goto _start; +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_22; lean_object* x_23; -x_22 = lean_box(0); -x_23 = lean_apply_3(x_10, x_9, x_6, x_22); -if (lean_obj_tag(x_23) == 0) +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_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___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_24; -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -lean_dec(x_23); -return x_24; +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_25; size_t x_26; size_t x_27; -x_25 = lean_ctor_get(x_23, 0); -lean_inc(x_25); -lean_dec(x_23); -x_26 = 1; -x_27 = lean_usize_add(x_5, x_26); -x_5 = x_27; -x_6 = x_25; -goto _start; +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; } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { -_start: +case 1: { -uint8_t x_7; -x_7 = lean_usize_dec_lt(x_5, x_4); -if (x_7 == 0) +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_Grind_EMatchTheorems_insert___spec__13(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: { -return x_6; +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 { -lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; -x_8 = lean_array_uget(x_3, x_5); -lean_inc(x_8); -x_9 = l_Lean_Meta_Grind_ppPattern(x_8); -x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1; -x_11 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isAtomicPattern(x_8); -lean_dec(x_8); -if (x_11 == 0) +uint8_t x_82; +x_82 = !lean_is_exclusive(x_1); +if (x_82 == 0) { -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_12 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2; -x_13 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3; -x_14 = l_Lean_MessageData_bracket(x_12, x_9, x_13); -x_15 = lean_box(0); -x_16 = lean_apply_3(x_10, x_14, x_6, x_15); -if (lean_obj_tag(x_16) == 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_Grind_EMatchTheorems_insert___spec__15(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_17; -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -lean_dec(x_16); -return x_17; +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_Grind_EMatchTheorems_insert___spec__3___closed__3; +x_93 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__14(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 { -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_5, x_19); -x_5 = x_20; -x_6 = x_18; -goto _start; +return x_84; } } else { -lean_object* x_22; lean_object* x_23; -x_22 = lean_box(0); -x_23 = lean_apply_3(x_10, x_9, x_6, x_22); -if (lean_obj_tag(x_23) == 0) -{ -lean_object* x_24; -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -lean_dec(x_23); -return x_24; +return x_84; +} } else { -lean_object* x_25; size_t x_26; size_t x_27; -x_25 = lean_ctor_get(x_23, 0); -lean_inc(x_25); -lean_dec(x_23); -x_26 = 1; -x_27 = lean_usize_add(x_5, x_26); -x_5 = x_27; -x_6 = x_25; -goto _start; -} -} -} -} -} -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__8(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { -_start: +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_Grind_EMatchTheorems_insert___spec__15(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) { -uint8_t x_7; -x_7 = lean_usize_dec_lt(x_5, x_4); -if (x_7 == 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) { -return x_6; +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_Grind_EMatchTheorems_insert___spec__3___closed__3; +x_107 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__14(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 { -lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; -x_8 = lean_array_uget(x_3, x_5); -lean_inc(x_8); -x_9 = l_Lean_Meta_Grind_ppPattern(x_8); -x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1; -x_11 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isAtomicPattern(x_8); -lean_dec(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 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2; -x_13 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3; -x_14 = l_Lean_MessageData_bracket(x_12, x_9, x_13); -x_15 = lean_box(0); -x_16 = lean_apply_3(x_10, x_14, x_6, x_15); -if (lean_obj_tag(x_16) == 0) -{ -lean_object* x_17; -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -lean_dec(x_16); -return x_17; +return x_98; +} } else { -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_5, x_19); -x_5 = x_20; -x_6 = x_18; -goto _start; +return x_98; } } -else +} +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__12(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -lean_object* x_22; lean_object* x_23; -x_22 = lean_box(0); -x_23 = lean_apply_3(x_10, x_9, x_6, x_22); -if (lean_obj_tag(x_23) == 0) +uint64_t x_4; size_t x_5; size_t x_6; lean_object* x_7; +x_4 = l_Lean_Name_hash___override(x_2); +x_5 = lean_uint64_to_usize(x_4); +x_6 = 1; +x_7 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__13(x_1, x_5, x_6, x_2, x_3); +return x_7; +} +} +static lean_object* _init_l_Lean_Meta_Grind_EMatchTheorems_insert___closed__1() { +_start: { -lean_object* x_24; -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -lean_dec(x_23); -return x_24; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.Meta.Tactic.Grind.EMatchTheorem", 36, 36); +return x_1; } -else +} +static lean_object* _init_l_Lean_Meta_Grind_EMatchTheorems_insert___closed__2() { +_start: { -lean_object* x_25; size_t x_26; size_t x_27; -x_25 = lean_ctor_get(x_23, 0); -lean_inc(x_25); -lean_dec(x_23); -x_26 = 1; -x_27 = lean_usize_add(x_5, x_26); -x_5 = x_27; -x_6 = x_25; -goto _start; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.Meta.Grind.EMatchTheorems.insert", 37, 37); +return x_1; +} } +static lean_object* _init_l_Lean_Meta_Grind_EMatchTheorems_insert___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("unreachable code has been reached", 33, 33); +return x_1; } } +static lean_object* _init_l_Lean_Meta_Grind_EMatchTheorems_insert___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_Meta_Grind_EMatchTheorems_insert___closed__1; +x_2 = l_Lean_Meta_Grind_EMatchTheorems_insert___closed__2; +x_3 = lean_unsigned_to_nat(120u); +x_4 = lean_unsigned_to_nat(6u); +x_5 = l_Lean_Meta_Grind_EMatchTheorems_insert___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_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__9(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_EMatchTheorems_insert(lean_object* x_1, lean_object* x_2) { _start: { -uint8_t x_7; -x_7 = lean_usize_dec_lt(x_5, x_4); -if (x_7 == 0) +lean_object* x_3; +x_3 = lean_ctor_get(x_2, 4); +lean_inc(x_3); +if (lean_obj_tag(x_3) == 0) { -return x_6; +lean_object* x_4; lean_object* x_5; +lean_dec(x_2); +lean_dec(x_1); +x_4 = l_Lean_Meta_Grind_EMatchTheorems_insert___closed__4; +x_5 = l_panic___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__1(x_4); +return x_5; } else { -lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; -x_8 = lean_array_uget(x_3, x_5); +lean_object* x_6; +x_6 = lean_ctor_get(x_3, 0); +lean_inc(x_6); +if (lean_obj_tag(x_6) == 2) +{ +uint8_t x_7; +x_7 = !lean_is_exclusive(x_2); +if (x_7 == 0) +{ +lean_object* x_8; lean_object* x_9; uint8_t x_10; +x_8 = lean_ctor_get(x_2, 5); +x_9 = lean_ctor_get(x_2, 4); +lean_dec(x_9); +x_10 = !lean_is_exclusive(x_3); +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_3, 1); +x_12 = lean_ctor_get(x_3, 0); +lean_dec(x_12); +x_13 = lean_ctor_get(x_6, 0); +lean_inc(x_13); +lean_dec(x_6); lean_inc(x_8); -x_9 = l_Lean_Meta_Grind_ppPattern(x_8); -x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1; -x_11 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isAtomicPattern(x_8); -lean_dec(x_8); -if (x_11 == 0) +lean_ctor_set(x_2, 4, x_11); +x_14 = !lean_is_exclusive(x_1); +if (x_14 == 0) { -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_12 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2; -x_13 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3; -x_14 = l_Lean_MessageData_bracket(x_12, x_9, x_13); -x_15 = lean_box(0); -x_16 = lean_apply_3(x_10, x_14, x_6, x_15); -if (lean_obj_tag(x_16) == 0) +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_15 = lean_ctor_get(x_1, 0); +x_16 = lean_ctor_get(x_1, 1); +x_17 = lean_ctor_get(x_1, 2); +x_18 = lean_box(0); +lean_inc(x_8); +x_19 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__2(x_16, x_8, x_18); +x_20 = l_Lean_PersistentHashMap_erase___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__6(x_17, x_8); +lean_dec(x_8); +lean_inc(x_15); +x_21 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__9(x_15, x_13); +if (lean_obj_tag(x_21) == 0) { -lean_object* x_17; -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -lean_dec(x_16); -return x_17; +lean_object* x_22; lean_object* x_23; +x_22 = lean_box(0); +lean_ctor_set(x_3, 1, x_22); +lean_ctor_set(x_3, 0, x_2); +x_23 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__12(x_15, x_13, x_3); +lean_ctor_set(x_1, 2, x_20); +lean_ctor_set(x_1, 1, x_19); +lean_ctor_set(x_1, 0, x_23); +return x_1; } else { -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_5, x_19); -x_5 = x_20; -x_6 = x_18; -goto _start; +lean_object* x_24; lean_object* x_25; +x_24 = lean_ctor_get(x_21, 0); +lean_inc(x_24); +lean_dec(x_21); +lean_ctor_set(x_3, 1, x_24); +lean_ctor_set(x_3, 0, x_2); +x_25 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__12(x_15, x_13, x_3); +lean_ctor_set(x_1, 2, x_20); +lean_ctor_set(x_1, 1, x_19); +lean_ctor_set(x_1, 0, x_25); +return x_1; } } else { -lean_object* x_22; lean_object* x_23; -x_22 = lean_box(0); -x_23 = lean_apply_3(x_10, x_9, x_6, x_22); -if (lean_obj_tag(x_23) == 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_1, 0); +x_27 = lean_ctor_get(x_1, 1); +x_28 = lean_ctor_get(x_1, 2); +lean_inc(x_28); +lean_inc(x_27); +lean_inc(x_26); +lean_dec(x_1); +x_29 = lean_box(0); +lean_inc(x_8); +x_30 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__2(x_27, x_8, x_29); +x_31 = l_Lean_PersistentHashMap_erase___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__6(x_28, x_8); +lean_dec(x_8); +lean_inc(x_26); +x_32 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__9(x_26, x_13); +if (lean_obj_tag(x_32) == 0) { -lean_object* x_24; -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -lean_dec(x_23); -return x_24; +lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_33 = lean_box(0); +lean_ctor_set(x_3, 1, x_33); +lean_ctor_set(x_3, 0, x_2); +x_34 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__12(x_26, x_13, x_3); +x_35 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_35, 0, x_34); +lean_ctor_set(x_35, 1, x_30); +lean_ctor_set(x_35, 2, x_31); +return x_35; } else { -lean_object* x_25; size_t x_26; size_t x_27; -x_25 = lean_ctor_get(x_23, 0); -lean_inc(x_25); -lean_dec(x_23); -x_26 = 1; -x_27 = lean_usize_add(x_5, x_26); -x_5 = x_27; -x_6 = x_25; -goto _start; -} -} +lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_32, 0); +lean_inc(x_36); +lean_dec(x_32); +lean_ctor_set(x_3, 1, x_36); +lean_ctor_set(x_3, 0, x_2); +x_37 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__12(x_26, x_13, x_3); +x_38 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_38, 0, x_37); +lean_ctor_set(x_38, 1, x_30); +lean_ctor_set(x_38, 2, x_31); +return x_38; } } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__10(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { -_start: +else { -uint8_t x_7; -x_7 = lean_usize_dec_lt(x_5, x_4); -if (x_7 == 0) -{ -return x_6; +lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; 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_39 = lean_ctor_get(x_3, 1); +lean_inc(x_39); +lean_dec(x_3); +x_40 = lean_ctor_get(x_6, 0); +lean_inc(x_40); +lean_dec(x_6); +lean_inc(x_8); +lean_ctor_set(x_2, 4, x_39); +x_41 = lean_ctor_get(x_1, 0); +lean_inc(x_41); +x_42 = lean_ctor_get(x_1, 1); +lean_inc(x_42); +x_43 = lean_ctor_get(x_1, 2); +lean_inc(x_43); +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + lean_ctor_release(x_1, 1); + lean_ctor_release(x_1, 2); + x_44 = x_1; +} else { + lean_dec_ref(x_1); + x_44 = lean_box(0); } -else -{ -lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; -x_8 = lean_array_uget(x_3, x_5); +x_45 = lean_box(0); lean_inc(x_8); -x_9 = l_Lean_Meta_Grind_ppPattern(x_8); -x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1; -x_11 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isAtomicPattern(x_8); +x_46 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__2(x_42, x_8, x_45); +x_47 = l_Lean_PersistentHashMap_erase___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__6(x_43, x_8); lean_dec(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 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2; -x_13 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3; -x_14 = l_Lean_MessageData_bracket(x_12, x_9, x_13); -x_15 = lean_box(0); -x_16 = lean_apply_3(x_10, x_14, x_6, x_15); -if (lean_obj_tag(x_16) == 0) +lean_inc(x_41); +x_48 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__9(x_41, x_40); +if (lean_obj_tag(x_48) == 0) { -lean_object* x_17; -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -lean_dec(x_16); -return x_17; +lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; +x_49 = lean_box(0); +x_50 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_50, 0, x_2); +lean_ctor_set(x_50, 1, x_49); +x_51 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__12(x_41, x_40, x_50); +if (lean_is_scalar(x_44)) { + x_52 = lean_alloc_ctor(0, 3, 0); +} else { + x_52 = x_44; +} +lean_ctor_set(x_52, 0, x_51); +lean_ctor_set(x_52, 1, x_46); +lean_ctor_set(x_52, 2, x_47); +return x_52; } else { -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_5, x_19); -x_5 = x_20; -x_6 = x_18; -goto _start; +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 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_54, 0, x_2); +lean_ctor_set(x_54, 1, x_53); +x_55 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__12(x_41, x_40, x_54); +if (lean_is_scalar(x_44)) { + x_56 = lean_alloc_ctor(0, 3, 0); +} else { + x_56 = x_44; +} +lean_ctor_set(x_56, 0, x_55); +lean_ctor_set(x_56, 1, x_46); +lean_ctor_set(x_56, 2, x_47); +return x_56; } } -else -{ -lean_object* x_22; lean_object* x_23; -x_22 = lean_box(0); -x_23 = lean_apply_3(x_10, x_9, x_6, x_22); -if (lean_obj_tag(x_23) == 0) -{ -lean_object* x_24; -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -lean_dec(x_23); -return x_24; } else { -lean_object* x_25; size_t x_26; size_t x_27; -x_25 = lean_ctor_get(x_23, 0); -lean_inc(x_25); -lean_dec(x_23); -x_26 = 1; -x_27 = lean_usize_add(x_5, x_26); -x_5 = x_27; -x_6 = x_25; -goto _start; -} +lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; 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_57 = lean_ctor_get(x_2, 0); +x_58 = lean_ctor_get(x_2, 1); +x_59 = lean_ctor_get(x_2, 2); +x_60 = lean_ctor_get(x_2, 3); +x_61 = lean_ctor_get(x_2, 5); +lean_inc(x_61); +lean_inc(x_60); +lean_inc(x_59); +lean_inc(x_58); +lean_inc(x_57); +lean_dec(x_2); +x_62 = lean_ctor_get(x_3, 1); +lean_inc(x_62); +if (lean_is_exclusive(x_3)) { + lean_ctor_release(x_3, 0); + lean_ctor_release(x_3, 1); + x_63 = x_3; +} else { + lean_dec_ref(x_3); + x_63 = lean_box(0); } +x_64 = lean_ctor_get(x_6, 0); +lean_inc(x_64); +lean_dec(x_6); +lean_inc(x_61); +x_65 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_65, 0, x_57); +lean_ctor_set(x_65, 1, x_58); +lean_ctor_set(x_65, 2, x_59); +lean_ctor_set(x_65, 3, x_60); +lean_ctor_set(x_65, 4, x_62); +lean_ctor_set(x_65, 5, x_61); +x_66 = lean_ctor_get(x_1, 0); +lean_inc(x_66); +x_67 = lean_ctor_get(x_1, 1); +lean_inc(x_67); +x_68 = lean_ctor_get(x_1, 2); +lean_inc(x_68); +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + lean_ctor_release(x_1, 1); + lean_ctor_release(x_1, 2); + x_69 = x_1; +} else { + lean_dec_ref(x_1); + x_69 = lean_box(0); } +x_70 = lean_box(0); +lean_inc(x_61); +x_71 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__2(x_67, x_61, x_70); +x_72 = l_Lean_PersistentHashMap_erase___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__6(x_68, x_61); +lean_dec(x_61); +lean_inc(x_66); +x_73 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__9(x_66, x_64); +if (lean_obj_tag(x_73) == 0) +{ +lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; +x_74 = lean_box(0); +if (lean_is_scalar(x_63)) { + x_75 = lean_alloc_ctor(1, 2, 0); +} else { + x_75 = x_63; } +lean_ctor_set(x_75, 0, x_65); +lean_ctor_set(x_75, 1, x_74); +x_76 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__12(x_66, x_64, x_75); +if (lean_is_scalar(x_69)) { + x_77 = lean_alloc_ctor(0, 3, 0); +} else { + x_77 = x_69; } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__11(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { -_start: -{ -uint8_t x_7; -x_7 = lean_usize_dec_lt(x_5, x_4); -if (x_7 == 0) -{ -return x_6; +lean_ctor_set(x_77, 0, x_76); +lean_ctor_set(x_77, 1, x_71); +lean_ctor_set(x_77, 2, x_72); +return x_77; } else { -lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; -x_8 = lean_array_uget(x_3, x_5); -lean_inc(x_8); -x_9 = l_Lean_Meta_Grind_ppPattern(x_8); -x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1; -x_11 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isAtomicPattern(x_8); -lean_dec(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 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2; -x_13 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3; -x_14 = l_Lean_MessageData_bracket(x_12, x_9, x_13); -x_15 = lean_box(0); -x_16 = lean_apply_3(x_10, x_14, x_6, x_15); -if (lean_obj_tag(x_16) == 0) -{ -lean_object* x_17; -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -lean_dec(x_16); -return x_17; +lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; +x_78 = lean_ctor_get(x_73, 0); +lean_inc(x_78); +lean_dec(x_73); +if (lean_is_scalar(x_63)) { + x_79 = lean_alloc_ctor(1, 2, 0); +} else { + x_79 = x_63; } -else -{ -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_5, x_19); -x_5 = x_20; -x_6 = x_18; -goto _start; +lean_ctor_set(x_79, 0, x_65); +lean_ctor_set(x_79, 1, x_78); +x_80 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__12(x_66, x_64, x_79); +if (lean_is_scalar(x_69)) { + x_81 = lean_alloc_ctor(0, 3, 0); +} else { + x_81 = x_69; +} +lean_ctor_set(x_81, 0, x_80); +lean_ctor_set(x_81, 1, x_71); +lean_ctor_set(x_81, 2, x_72); +return x_81; } } -else -{ -lean_object* x_22; lean_object* x_23; -x_22 = lean_box(0); -x_23 = lean_apply_3(x_10, x_9, x_6, x_22); -if (lean_obj_tag(x_23) == 0) -{ -lean_object* x_24; -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -lean_dec(x_23); -return x_24; } else { -lean_object* x_25; size_t x_26; size_t x_27; -x_25 = lean_ctor_get(x_23, 0); -lean_inc(x_25); -lean_dec(x_23); -x_26 = 1; -x_27 = lean_usize_add(x_5, x_26); -x_5 = x_27; -x_6 = x_25; -goto _start; -} +lean_object* x_82; lean_object* x_83; +lean_dec(x_6); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_82 = l_Lean_Meta_Grind_EMatchTheorems_insert___closed__4; +x_83 = l_panic___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__1(x_82); +return x_83; } } } } -static lean_object* _init_l_Lean_Meta_Grind_ppPattern___closed__1() { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("#", 1, 1); -return x_1; +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_Grind_EMatchTheorems_insert___spec__4(x_7, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_3); +lean_dec(x_2); +return x_8; } } -static lean_object* _init_l_Lean_Meta_Grind_ppPattern___closed__2() { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___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_1; lean_object* x_2; -x_1 = l_Lean_Meta_Grind_ppPattern___closed__1; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; +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_Grind_EMatchTheorems_insert___spec__3(x_1, x_6, x_7, x_4, x_5); +return x_8; } } -static lean_object* _init_l_Lean_Meta_Grind_ppPattern___closed__3() { +LEAN_EXPORT lean_object* l_Array_indexOfAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("", 0, 0); -return x_1; +lean_object* x_4; +x_4 = l_Array_indexOfAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__8(x_1, x_2, x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_4; } } -static lean_object* _init_l_Lean_Meta_Grind_ppPattern___closed__4() { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__7___boxed(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_Meta_Grind_ppPattern___closed__3; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; +size_t x_4; lean_object* x_5; +x_4 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_5 = l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__7(x_1, x_4, x_3); +lean_dec(x_3); +return x_5; } } -static lean_object* _init_l_Lean_Meta_Grind_ppPattern___closed__5() { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_erase___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__6___boxed(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_levelZero; -x_2 = l_Lean_Expr_sort___override(x_1); -return x_2; +lean_object* x_3; +x_3 = l_Lean_PersistentHashMap_erase___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__6(x_1, x_2); +lean_dec(x_2); +return x_3; } } -static lean_object* _init_l_Lean_Meta_Grind_ppPattern___closed__6() { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__11___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; -x_1 = lean_mk_string_unchecked("\?", 1, 1); -return x_1; +lean_object* x_6; +x_6 = l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__11(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; } } -static lean_object* _init_l_Lean_Meta_Grind_ppPattern___closed__7() { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__10___boxed(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_Meta_Grind_ppPattern___closed__6; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; +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_Grind_EMatchTheorems_insert___spec__10(x_1, x_4, x_3); +lean_dec(x_3); +return x_5; } } -static lean_object* _init_l_Lean_Meta_Grind_ppPattern___closed__8() { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__9___boxed(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("`[", 2, 2); -return x_1; +lean_object* x_3; +x_3 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__9(x_1, x_2); +lean_dec(x_2); +return x_3; } } -static lean_object* _init_l_Lean_Meta_Grind_ppPattern___closed__9() { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__14___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_Grind_ppPattern___closed__8; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; +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_Grind_EMatchTheorems_insert___spec__14(x_7, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_3); +lean_dec(x_2); +return x_8; } } -static lean_object* _init_l_Lean_Meta_Grind_ppPattern___closed__10() { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__13___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("]", 1, 1); -return x_1; +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_Grind_EMatchTheorems_insert___spec__13(x_1, x_6, x_7, x_4, x_5); +return x_8; } } -static lean_object* _init_l_Lean_Meta_Grind_ppPattern___closed__11() { +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_EMatchTheorems_contains___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_1; lean_object* x_2; -x_1 = l_Lean_Meta_Grind_ppPattern___closed__10; -x_2 = l_Lean_stringToMessageData(x_1); -return 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) +{ +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; uint8_t x_12; +x_9 = lean_array_fget(x_1, x_4); +x_10 = l_Lean_Meta_Grind_Origin_key(x_5); +x_11 = l_Lean_Meta_Grind_Origin_key(x_9); +lean_dec(x_9); +x_12 = lean_name_eq(x_10, x_11); +lean_dec(x_11); +lean_dec(x_10); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; +x_13 = lean_unsigned_to_nat(1u); +x_14 = lean_nat_add(x_4, x_13); +lean_dec(x_4); +x_3 = lean_box(0); +x_4 = x_14; +goto _start; } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ppPattern(lean_object* x_1) { -_start: +else { -lean_object* x_2; -x_2 = l_Lean_Meta_Grind_groundPattern_x3f(x_1); -if (lean_obj_tag(x_2) == 0) +uint8_t x_16; +lean_dec(x_4); +x_16 = 1; +return x_16; +} +} +} +} +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_EMatchTheorems_contains___spec__2(lean_object* x_1, size_t x_2, lean_object* x_3) { +_start: { -lean_object* x_3; uint8_t x_4; -x_3 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare; -x_4 = lean_expr_eqv(x_1, x_3); -if (x_4 == 0) +if (lean_obj_tag(x_1) == 0) { -switch (lean_obj_tag(x_1)) { +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_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___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_5; lean_object* x_6; 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_ctor_get(x_1, 0); -lean_inc(x_5); -lean_dec(x_1); -x_6 = l___private_Init_Data_Repr_0__Nat_reprFast(x_5); -x_7 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_7, 0, x_6); -x_8 = l_Lean_MessageData_ofFormat(x_7); -x_9 = l_Lean_Meta_Grind_ppPattern___closed__2; -x_10 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_10, 0, x_9); -lean_ctor_set(x_10, 1, x_8); -x_11 = l_Lean_Meta_Grind_ppPattern___closed__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; +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); +lean_dec(x_10); +x_12 = l_Lean_Meta_Grind_Origin_key(x_3); +x_13 = l_Lean_Meta_Grind_Origin_key(x_11); +lean_dec(x_11); +x_14 = lean_name_eq(x_12, x_13); +lean_dec(x_13); +lean_dec(x_12); +return x_14; } case 1: { -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; size_t x_26; size_t x_27; lean_object* x_28; -x_13 = l_Lean_Expr_getAppFn(x_1); -x_14 = l_Lean_MessageData_ofExpr(x_13); -x_15 = l_Lean_Meta_Grind_ppPattern___closed__4; -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 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_17, 0, x_16); -lean_ctor_set(x_17, 1, x_15); -x_18 = lean_unsigned_to_nat(0u); -x_19 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_18); -x_20 = l_Lean_Meta_Grind_ppPattern___closed__5; +lean_object* x_15; size_t x_16; +x_15 = lean_ctor_get(x_10, 0); +lean_inc(x_15); +lean_dec(x_10); +x_16 = lean_usize_shift_right(x_2, x_5); +x_1 = x_15; +x_2 = x_16; +goto _start; +} +default: +{ +uint8_t x_18; +x_18 = 0; +return x_18; +} +} +} +else +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_19 = lean_ctor_get(x_1, 0); lean_inc(x_19); -x_21 = lean_mk_array(x_19, x_20); -x_22 = lean_unsigned_to_nat(1u); -x_23 = lean_nat_sub(x_19, x_22); +x_20 = lean_ctor_get(x_1, 1); +lean_inc(x_20); +lean_dec(x_1); +x_21 = lean_unsigned_to_nat(0u); +x_22 = l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_EMatchTheorems_contains___spec__3(x_19, x_20, lean_box(0), x_21, x_3); +lean_dec(x_20); lean_dec(x_19); -x_24 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_21, x_23); -x_25 = lean_box(0); -x_26 = lean_array_size(x_24); -x_27 = 0; -x_28 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1(x_24, x_25, x_24, x_26, x_27, x_17); -lean_dec(x_24); -return x_28; +return x_22; } -case 2: +} +} +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_EMatchTheorems_contains___spec__1(lean_object* x_1, lean_object* x_2) { +_start: { -lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; 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_29 = l_Lean_Expr_getAppFn(x_1); -x_30 = l_Lean_MessageData_ofExpr(x_29); -x_31 = l_Lean_Meta_Grind_ppPattern___closed__4; -x_32 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_32, 0, x_31); -lean_ctor_set(x_32, 1, x_30); -x_33 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_33, 0, x_32); -lean_ctor_set(x_33, 1, x_31); -x_34 = lean_unsigned_to_nat(0u); -x_35 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_34); -x_36 = l_Lean_Meta_Grind_ppPattern___closed__5; -lean_inc(x_35); -x_37 = lean_mk_array(x_35, x_36); -x_38 = lean_unsigned_to_nat(1u); -x_39 = lean_nat_sub(x_35, x_38); -lean_dec(x_35); -x_40 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_37, x_39); -x_41 = lean_box(0); -x_42 = lean_array_size(x_40); -x_43 = 0; -x_44 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__2(x_40, x_41, x_40, x_42, x_43, x_33); -lean_dec(x_40); -return x_44; +lean_object* x_3; uint64_t x_4; size_t x_5; uint8_t x_6; +x_3 = l_Lean_Meta_Grind_Origin_key(x_2); +x_4 = l_Lean_Name_hash___override(x_3); +lean_dec(x_3); +x_5 = lean_uint64_to_usize(x_4); +x_6 = l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_EMatchTheorems_contains___spec__2(x_1, x_5, x_2); +return x_6; } -case 3: +} +LEAN_EXPORT uint8_t l_Lean_Meta_Grind_EMatchTheorems_contains(lean_object* x_1, lean_object* x_2) { +_start: { -lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; size_t x_58; size_t x_59; lean_object* x_60; -x_45 = l_Lean_Expr_getAppFn(x_1); -x_46 = l_Lean_MessageData_ofExpr(x_45); -x_47 = l_Lean_Meta_Grind_ppPattern___closed__4; -x_48 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_48, 0, x_47); -lean_ctor_set(x_48, 1, x_46); -x_49 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_49, 0, x_48); -lean_ctor_set(x_49, 1, x_47); -x_50 = lean_unsigned_to_nat(0u); -x_51 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_50); -x_52 = l_Lean_Meta_Grind_ppPattern___closed__5; -lean_inc(x_51); -x_53 = lean_mk_array(x_51, x_52); -x_54 = lean_unsigned_to_nat(1u); -x_55 = lean_nat_sub(x_51, x_54); -lean_dec(x_51); -x_56 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_53, x_55); -x_57 = lean_box(0); -x_58 = lean_array_size(x_56); -x_59 = 0; -x_60 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__3(x_56, x_57, x_56, x_58, x_59, x_49); -lean_dec(x_56); -return x_60; +lean_object* x_3; uint8_t x_4; +x_3 = lean_ctor_get(x_1, 1); +lean_inc(x_3); +lean_dec(x_1); +x_4 = l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_EMatchTheorems_contains___spec__1(x_3, x_2); +return x_4; } -case 4: +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_EMatchTheorems_contains___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_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; size_t x_74; size_t x_75; lean_object* x_76; -x_61 = l_Lean_Expr_getAppFn(x_1); -x_62 = l_Lean_MessageData_ofExpr(x_61); -x_63 = l_Lean_Meta_Grind_ppPattern___closed__4; -x_64 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_64, 0, x_63); -lean_ctor_set(x_64, 1, x_62); -x_65 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_65, 0, x_64); -lean_ctor_set(x_65, 1, x_63); -x_66 = lean_unsigned_to_nat(0u); -x_67 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_66); -x_68 = l_Lean_Meta_Grind_ppPattern___closed__5; -lean_inc(x_67); -x_69 = lean_mk_array(x_67, x_68); -x_70 = lean_unsigned_to_nat(1u); -x_71 = lean_nat_sub(x_67, x_70); -lean_dec(x_67); -x_72 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_69, x_71); -x_73 = lean_box(0); -x_74 = lean_array_size(x_72); -x_75 = 0; -x_76 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__4(x_72, x_73, x_72, x_74, x_75, x_65); -lean_dec(x_72); -return x_76; +uint8_t x_6; lean_object* x_7; +x_6 = l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_EMatchTheorems_contains___spec__3(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_5); +lean_dec(x_2); +lean_dec(x_1); +x_7 = lean_box(x_6); +return x_7; } -case 5: +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_EMatchTheorems_contains___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_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; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; size_t x_90; size_t x_91; lean_object* x_92; -x_77 = l_Lean_Expr_getAppFn(x_1); -x_78 = l_Lean_MessageData_ofExpr(x_77); -x_79 = l_Lean_Meta_Grind_ppPattern___closed__4; -x_80 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_80, 0, x_79); -lean_ctor_set(x_80, 1, x_78); -x_81 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_81, 0, x_80); -lean_ctor_set(x_81, 1, x_79); -x_82 = lean_unsigned_to_nat(0u); -x_83 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_82); -x_84 = l_Lean_Meta_Grind_ppPattern___closed__5; -lean_inc(x_83); -x_85 = lean_mk_array(x_83, x_84); -x_86 = lean_unsigned_to_nat(1u); -x_87 = lean_nat_sub(x_83, x_86); -lean_dec(x_83); -x_88 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_85, x_87); -x_89 = lean_box(0); -x_90 = lean_array_size(x_88); -x_91 = 0; -x_92 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__5(x_88, x_89, x_88, x_90, x_91, x_81); -lean_dec(x_88); -return x_92; +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_Meta_Grind_EMatchTheorems_contains___spec__2(x_1, x_4, x_3); +lean_dec(x_3); +x_6 = lean_box(x_5); +return x_6; } -case 6: -{ -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; size_t x_106; size_t x_107; lean_object* x_108; -x_93 = l_Lean_Expr_getAppFn(x_1); -x_94 = l_Lean_MessageData_ofExpr(x_93); -x_95 = l_Lean_Meta_Grind_ppPattern___closed__4; -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_alloc_ctor(7, 2, 0); -lean_ctor_set(x_97, 0, x_96); -lean_ctor_set(x_97, 1, x_95); -x_98 = lean_unsigned_to_nat(0u); -x_99 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_98); -x_100 = l_Lean_Meta_Grind_ppPattern___closed__5; -lean_inc(x_99); -x_101 = lean_mk_array(x_99, x_100); -x_102 = lean_unsigned_to_nat(1u); -x_103 = lean_nat_sub(x_99, x_102); -lean_dec(x_99); -x_104 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_101, x_103); -x_105 = lean_box(0); -x_106 = lean_array_size(x_104); -x_107 = 0; -x_108 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__6(x_104, x_105, x_104, x_106, x_107, x_97); -lean_dec(x_104); -return x_108; } -case 7: +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_EMatchTheorems_contains___spec__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: { -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; size_t x_122; size_t x_123; lean_object* x_124; -x_109 = l_Lean_Expr_getAppFn(x_1); -x_110 = l_Lean_MessageData_ofExpr(x_109); -x_111 = l_Lean_Meta_Grind_ppPattern___closed__4; -x_112 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_112, 0, x_111); -lean_ctor_set(x_112, 1, x_110); -x_113 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_113, 0, x_112); -lean_ctor_set(x_113, 1, x_111); -x_114 = lean_unsigned_to_nat(0u); -x_115 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_114); -x_116 = l_Lean_Meta_Grind_ppPattern___closed__5; -lean_inc(x_115); -x_117 = lean_mk_array(x_115, x_116); -x_118 = lean_unsigned_to_nat(1u); -x_119 = lean_nat_sub(x_115, x_118); -lean_dec(x_115); -x_120 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_117, x_119); -x_121 = lean_box(0); -x_122 = lean_array_size(x_120); -x_123 = 0; -x_124 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__7(x_120, x_121, x_120, x_122, x_123, x_113); -lean_dec(x_120); -return x_124; +uint8_t x_3; lean_object* x_4; +x_3 = l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_EMatchTheorems_contains___spec__1(x_1, x_2); +lean_dec(x_2); +x_4 = lean_box(x_3); +return x_4; } -case 8: -{ -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; size_t x_138; size_t x_139; lean_object* x_140; -x_125 = l_Lean_Expr_getAppFn(x_1); -x_126 = l_Lean_MessageData_ofExpr(x_125); -x_127 = l_Lean_Meta_Grind_ppPattern___closed__4; -x_128 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_128, 0, x_127); -lean_ctor_set(x_128, 1, x_126); -x_129 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_129, 0, x_128); -lean_ctor_set(x_129, 1, x_127); -x_130 = lean_unsigned_to_nat(0u); -x_131 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_130); -x_132 = l_Lean_Meta_Grind_ppPattern___closed__5; -lean_inc(x_131); -x_133 = lean_mk_array(x_131, x_132); -x_134 = lean_unsigned_to_nat(1u); -x_135 = lean_nat_sub(x_131, x_134); -lean_dec(x_131); -x_136 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_133, x_135); -x_137 = lean_box(0); -x_138 = lean_array_size(x_136); -x_139 = 0; -x_140 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__8(x_136, x_137, x_136, x_138, x_139, x_129); -lean_dec(x_136); -return x_140; } -case 9: +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_EMatchTheorems_contains___boxed(lean_object* x_1, lean_object* x_2) { +_start: { -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; size_t x_154; size_t x_155; lean_object* x_156; -x_141 = l_Lean_Expr_getAppFn(x_1); -x_142 = l_Lean_MessageData_ofExpr(x_141); -x_143 = l_Lean_Meta_Grind_ppPattern___closed__4; -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 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_145, 0, x_144); -lean_ctor_set(x_145, 1, x_143); -x_146 = lean_unsigned_to_nat(0u); -x_147 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_146); -x_148 = l_Lean_Meta_Grind_ppPattern___closed__5; -lean_inc(x_147); -x_149 = lean_mk_array(x_147, x_148); -x_150 = lean_unsigned_to_nat(1u); -x_151 = lean_nat_sub(x_147, x_150); -lean_dec(x_147); -x_152 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_149, x_151); -x_153 = lean_box(0); -x_154 = lean_array_size(x_152); -x_155 = 0; -x_156 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__9(x_152, x_153, x_152, x_154, x_155, x_145); -lean_dec(x_152); -return x_156; +uint8_t x_3; lean_object* x_4; +x_3 = l_Lean_Meta_Grind_EMatchTheorems_contains(x_1, x_2); +lean_dec(x_2); +x_4 = lean_box(x_3); +return x_4; } -case 10: -{ -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; size_t x_170; size_t x_171; lean_object* x_172; -x_157 = l_Lean_Expr_getAppFn(x_1); -x_158 = l_Lean_MessageData_ofExpr(x_157); -x_159 = l_Lean_Meta_Grind_ppPattern___closed__4; -x_160 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_160, 0, x_159); -lean_ctor_set(x_160, 1, x_158); -x_161 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_161, 0, x_160); -lean_ctor_set(x_161, 1, x_159); -x_162 = lean_unsigned_to_nat(0u); -x_163 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_162); -x_164 = l_Lean_Meta_Grind_ppPattern___closed__5; -lean_inc(x_163); -x_165 = lean_mk_array(x_163, x_164); -x_166 = lean_unsigned_to_nat(1u); -x_167 = lean_nat_sub(x_163, x_166); -lean_dec(x_163); -x_168 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_165, x_167); -x_169 = lean_box(0); -x_170 = lean_array_size(x_168); -x_171 = 0; -x_172 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__10(x_168, x_169, x_168, x_170, x_171, x_161); -lean_dec(x_168); -return x_172; } -default: +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_EMatchTheorems_erase(lean_object* x_1, lean_object* x_2) { +_start: { -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; size_t x_186; size_t x_187; lean_object* x_188; -x_173 = l_Lean_Expr_getAppFn(x_1); -x_174 = l_Lean_MessageData_ofExpr(x_173); -x_175 = l_Lean_Meta_Grind_ppPattern___closed__4; -x_176 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_176, 0, x_175); -lean_ctor_set(x_176, 1, x_174); -x_177 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_177, 0, x_176); -lean_ctor_set(x_177, 1, x_175); -x_178 = lean_unsigned_to_nat(0u); -x_179 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_178); -x_180 = l_Lean_Meta_Grind_ppPattern___closed__5; -lean_inc(x_179); -x_181 = lean_mk_array(x_179, x_180); -x_182 = lean_unsigned_to_nat(1u); -x_183 = lean_nat_sub(x_179, x_182); -lean_dec(x_179); -x_184 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_181, x_183); -x_185 = lean_box(0); -x_186 = lean_array_size(x_184); -x_187 = 0; -x_188 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__11(x_184, x_185, x_184, x_186, x_187, x_177); -lean_dec(x_184); -return x_188; -} -} -} -else +uint8_t x_3; +x_3 = !lean_is_exclusive(x_1); +if (x_3 == 0) { -lean_object* x_189; -lean_dec(x_1); -x_189 = l_Lean_Meta_Grind_ppPattern___closed__7; -return x_189; -} +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_4 = lean_ctor_get(x_1, 1); +x_5 = lean_ctor_get(x_1, 2); +x_6 = l_Lean_PersistentHashMap_erase___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__6(x_4, x_2); +x_7 = lean_box(0); +x_8 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__2(x_5, x_2, x_7); +lean_ctor_set(x_1, 2, x_8); +lean_ctor_set(x_1, 1, x_6); +return x_1; } 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_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_9 = lean_ctor_get(x_1, 0); +x_10 = lean_ctor_get(x_1, 1); +x_11 = lean_ctor_get(x_1, 2); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); lean_dec(x_1); -x_190 = lean_ctor_get(x_2, 0); -lean_inc(x_190); -lean_dec(x_2); -x_191 = l_Lean_MessageData_ofExpr(x_190); -x_192 = l_Lean_Meta_Grind_ppPattern___closed__9; -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_Meta_Grind_ppPattern___closed__11; -x_195 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_195, 0, x_193); -lean_ctor_set(x_195, 1, x_194); -return x_195; +x_12 = l_Lean_PersistentHashMap_erase___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__6(x_10, x_2); +x_13 = lean_box(0); +x_14 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__2(x_11, x_2, x_13); +x_15 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_15, 0, x_9); +lean_ctor_set(x_15, 1, x_12); +lean_ctor_set(x_15, 2, x_14); +return x_15; } } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT uint8_t l_Lean_Meta_Grind_EMatchTheorems_isErased(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_4; -x_4 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1(x_1, x_2, x_3); -lean_dec(x_3); +lean_object* x_3; uint8_t x_4; +x_3 = lean_ctor_get(x_1, 2); +lean_inc(x_3); +lean_dec(x_1); +x_4 = l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_EMatchTheorems_contains___spec__1(x_3, x_2); return x_4; } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___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_Meta_Grind_EMatchTheorems_isErased___boxed(lean_object* x_1, lean_object* x_2) { _start: { -size_t x_7; size_t x_8; lean_object* x_9; -x_7 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_8 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_9 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1(x_1, x_2, x_3, x_7, x_8, x_6); -lean_dec(x_3); +uint8_t x_3; lean_object* x_4; +x_3 = l_Lean_Meta_Grind_EMatchTheorems_isErased(x_1, x_2); lean_dec(x_2); -lean_dec(x_1); -return x_9; +x_4 = lean_box(x_3); +return x_4; } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___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_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_EMatchTheorems_retrieve_x3f___spec__2(lean_object* x_1, size_t x_2, lean_object* x_3) { _start: { -size_t x_7; size_t x_8; lean_object* x_9; -x_7 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_8 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_9 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__2(x_1, x_2, x_3, x_7, x_8, x_6); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_9; -} -} -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___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: +if (lean_obj_tag(x_1) == 0) { -size_t x_7; size_t x_8; lean_object* x_9; -x_7 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_8 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_9 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__3(x_1, x_2, x_3, x_7, x_8, x_6); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_9; -} -} -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: +lean_object* x_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); +x_5 = 5; +x_6 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___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); +switch (lean_obj_tag(x_10)) { +case 0: { -size_t x_7; size_t x_8; lean_object* x_9; -x_7 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_8 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_9 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__4(x_1, x_2, x_3, x_7, x_8, x_6); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_9; -} -} -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: +lean_object* x_11; uint8_t x_12; +x_11 = lean_ctor_get(x_10, 0); +lean_inc(x_11); +lean_dec(x_10); +x_12 = lean_name_eq(x_3, x_11); +lean_dec(x_11); +if (x_12 == 0) { -size_t x_7; size_t x_8; lean_object* x_9; -x_7 = lean_unbox_usize(x_4); +lean_dec(x_8); lean_dec(x_4); -x_8 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_9 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__5(x_1, x_2, x_3, x_7, x_8, x_6); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_9; +return x_1; } +else +{ +uint8_t x_13; +x_13 = !lean_is_exclusive(x_1); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; +x_14 = lean_ctor_get(x_1, 0); +lean_dec(x_14); +x_15 = lean_array_set(x_4, x_8, x_9); +lean_dec(x_8); +lean_ctor_set(x_1, 0, x_15); +return x_1; } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___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: +else { -size_t x_7; size_t x_8; lean_object* x_9; -x_7 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_8 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_9 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__6(x_1, x_2, x_3, x_7, x_8, x_6); -lean_dec(x_3); -lean_dec(x_2); +lean_object* x_16; lean_object* x_17; lean_dec(x_1); -return x_9; +x_16 = lean_array_set(x_4, x_8, x_9); +lean_dec(x_8); +x_17 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_17, 0, x_16); +return x_17; } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: +} +case 1: { -size_t x_7; size_t x_8; lean_object* x_9; -x_7 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_8 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_9 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__7(x_1, x_2, x_3, x_7, x_8, x_6); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_9; +uint8_t x_18; +x_18 = !lean_is_exclusive(x_1); +if (x_18 == 0) +{ +lean_object* x_19; uint8_t x_20; +x_19 = lean_ctor_get(x_1, 0); +lean_dec(x_19); +x_20 = !lean_is_exclusive(x_10); +if (x_20 == 0) +{ +lean_object* x_21; lean_object* x_22; size_t x_23; lean_object* x_24; lean_object* x_25; +x_21 = lean_ctor_get(x_10, 0); +x_22 = lean_array_set(x_4, x_8, x_9); +x_23 = lean_usize_shift_right(x_2, x_5); +x_24 = l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_EMatchTheorems_retrieve_x3f___spec__2(x_21, x_23, x_3); +lean_inc(x_24); +x_25 = l_Lean_PersistentHashMap_isUnaryNode___rarg(x_24); +if (lean_obj_tag(x_25) == 0) +{ +lean_object* x_26; +lean_ctor_set(x_10, 0, x_24); +x_26 = lean_array_set(x_22, x_8, x_10); +lean_dec(x_8); +lean_ctor_set(x_1, 0, x_26); +return x_1; } +else +{ +lean_object* x_27; uint8_t x_28; +lean_dec(x_24); +lean_free_object(x_10); +x_27 = lean_ctor_get(x_25, 0); +lean_inc(x_27); +lean_dec(x_25); +x_28 = !lean_is_exclusive(x_27); +if (x_28 == 0) +{ +lean_object* x_29; +x_29 = lean_array_set(x_22, x_8, x_27); +lean_dec(x_8); +lean_ctor_set(x_1, 0, x_29); +return x_1; } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: +else { -size_t x_7; size_t x_8; lean_object* x_9; -x_7 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_8 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_9 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__8(x_1, x_2, x_3, x_7, x_8, x_6); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_9; +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_30 = lean_ctor_get(x_27, 0); +x_31 = lean_ctor_get(x_27, 1); +lean_inc(x_31); +lean_inc(x_30); +lean_dec(x_27); +x_32 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_32, 0, x_30); +lean_ctor_set(x_32, 1, x_31); +x_33 = lean_array_set(x_22, x_8, x_32); +lean_dec(x_8); +lean_ctor_set(x_1, 0, x_33); +return x_1; } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___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) { -_start: -{ -size_t x_7; size_t x_8; lean_object* x_9; -x_7 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_8 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_9 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__9(x_1, x_2, x_3, x_7, x_8, x_6); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_9; } +else +{ +lean_object* x_34; lean_object* x_35; size_t x_36; lean_object* x_37; lean_object* x_38; +x_34 = lean_ctor_get(x_10, 0); +lean_inc(x_34); +lean_dec(x_10); +x_35 = lean_array_set(x_4, x_8, x_9); +x_36 = lean_usize_shift_right(x_2, x_5); +x_37 = l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_EMatchTheorems_retrieve_x3f___spec__2(x_34, x_36, x_3); +lean_inc(x_37); +x_38 = l_Lean_PersistentHashMap_isUnaryNode___rarg(x_37); +if (lean_obj_tag(x_38) == 0) +{ +lean_object* x_39; lean_object* x_40; +x_39 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_39, 0, x_37); +x_40 = lean_array_set(x_35, x_8, x_39); +lean_dec(x_8); +lean_ctor_set(x_1, 0, x_40); +return x_1; } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___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: +else { -size_t x_7; size_t x_8; lean_object* x_9; -x_7 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_8 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_9 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__10(x_1, x_2, x_3, x_7, x_8, x_6); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_9; +lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; +lean_dec(x_37); +x_41 = lean_ctor_get(x_38, 0); +lean_inc(x_41); +lean_dec(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_43); +if (lean_is_exclusive(x_41)) { + lean_ctor_release(x_41, 0); + lean_ctor_release(x_41, 1); + x_44 = x_41; +} else { + lean_dec_ref(x_41); + x_44 = lean_box(0); } +if (lean_is_scalar(x_44)) { + x_45 = lean_alloc_ctor(0, 2, 0); +} else { + x_45 = x_44; } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___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: -{ -size_t x_7; size_t x_8; lean_object* x_9; -x_7 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_8 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_9 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__11(x_1, x_2, x_3, x_7, x_8, x_6); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_9; +lean_ctor_set(x_45, 0, x_42); +lean_ctor_set(x_45, 1, x_43); +x_46 = lean_array_set(x_35, x_8, x_45); +lean_dec(x_8); +lean_ctor_set(x_1, 0, x_46); +return x_1; } } -LEAN_EXPORT uint8_t l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__1(lean_object* x_1, lean_object* x_2) { -_start: -{ -if (lean_obj_tag(x_2) == 0) -{ -uint8_t x_3; -x_3 = 0; -return x_3; } else { -lean_object* x_4; lean_object* x_5; uint8_t x_6; -x_4 = lean_ctor_get(x_2, 0); -x_5 = lean_ctor_get(x_2, 2); -x_6 = l___private_Lean_HeadIndex_0__Lean_beqHeadIndex____x40_Lean_HeadIndex___hyg_69_(x_4, x_1); -if (x_6 == 0) +lean_object* x_47; lean_object* x_48; lean_object* x_49; size_t x_50; lean_object* x_51; lean_object* x_52; +lean_dec(x_1); +x_47 = lean_ctor_get(x_10, 0); +lean_inc(x_47); +if (lean_is_exclusive(x_10)) { + lean_ctor_release(x_10, 0); + x_48 = x_10; +} else { + lean_dec_ref(x_10); + x_48 = lean_box(0); +} +x_49 = lean_array_set(x_4, x_8, x_9); +x_50 = lean_usize_shift_right(x_2, x_5); +x_51 = l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_EMatchTheorems_retrieve_x3f___spec__2(x_47, x_50, x_3); +lean_inc(x_51); +x_52 = l_Lean_PersistentHashMap_isUnaryNode___rarg(x_51); +if (lean_obj_tag(x_52) == 0) { -x_2 = x_5; -goto _start; +lean_object* x_53; lean_object* x_54; lean_object* x_55; +if (lean_is_scalar(x_48)) { + x_53 = lean_alloc_ctor(1, 1, 0); +} else { + x_53 = x_48; +} +lean_ctor_set(x_53, 0, x_51); +x_54 = lean_array_set(x_49, x_8, x_53); +lean_dec(x_8); +x_55 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_55, 0, x_54); +return x_55; } else { -uint8_t x_8; -x_8 = 1; -return x_8; +lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; +lean_dec(x_51); +lean_dec(x_48); +x_56 = lean_ctor_get(x_52, 0); +lean_inc(x_56); +lean_dec(x_52); +x_57 = lean_ctor_get(x_56, 0); +lean_inc(x_57); +x_58 = lean_ctor_get(x_56, 1); +lean_inc(x_58); +if (lean_is_exclusive(x_56)) { + lean_ctor_release(x_56, 0); + lean_ctor_release(x_56, 1); + x_59 = x_56; +} else { + lean_dec_ref(x_56); + x_59 = lean_box(0); } +if (lean_is_scalar(x_59)) { + x_60 = lean_alloc_ctor(0, 2, 0); +} else { + x_60 = x_59; } +lean_ctor_set(x_60, 0, x_57); +lean_ctor_set(x_60, 1, x_58); +x_61 = lean_array_set(x_49, x_8, x_60); +lean_dec(x_8); +x_62 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_62, 0, x_61); +return x_62; } } -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__4(lean_object* x_1, lean_object* x_2) { -_start: -{ -if (lean_obj_tag(x_2) == 0) +} +default: { +lean_dec(x_8); +lean_dec(x_4); return x_1; } +} +} else { -uint8_t x_3; -x_3 = !lean_is_exclusive(x_2); -if (x_3 == 0) +lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_63 = lean_ctor_get(x_1, 0); +lean_inc(x_63); +x_64 = lean_ctor_get(x_1, 1); +lean_inc(x_64); +x_65 = lean_unsigned_to_nat(0u); +x_66 = l_Array_indexOfAux___at_Lean_MetavarContext_setMVarUserName___spec__3(x_63, x_3, x_65); +if (lean_obj_tag(x_66) == 0) { -lean_object* x_4; lean_object* x_5; lean_object* x_6; uint64_t x_7; uint64_t x_8; uint64_t x_9; uint64_t x_10; uint64_t x_11; uint64_t x_12; uint64_t x_13; size_t x_14; size_t x_15; size_t x_16; size_t x_17; size_t x_18; lean_object* x_19; lean_object* x_20; -x_4 = lean_ctor_get(x_2, 0); -x_5 = lean_ctor_get(x_2, 2); -x_6 = lean_array_get_size(x_1); -x_7 = l_Lean_HeadIndex_hash(x_4); -x_8 = 32; -x_9 = lean_uint64_shift_right(x_7, x_8); -x_10 = lean_uint64_xor(x_7, x_9); -x_11 = 16; -x_12 = lean_uint64_shift_right(x_10, x_11); -x_13 = lean_uint64_xor(x_10, x_12); -x_14 = lean_uint64_to_usize(x_13); -x_15 = lean_usize_of_nat(x_6); -lean_dec(x_6); -x_16 = 1; -x_17 = lean_usize_sub(x_15, x_16); -x_18 = lean_usize_land(x_14, x_17); -x_19 = lean_array_uget(x_1, x_18); -lean_ctor_set(x_2, 2, x_19); -x_20 = lean_array_uset(x_1, x_18, x_2); -x_1 = x_20; -x_2 = x_5; -goto _start; +lean_dec(x_64); +lean_dec(x_63); +return x_1; } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint64_t x_26; uint64_t x_27; uint64_t x_28; uint64_t x_29; uint64_t x_30; uint64_t x_31; uint64_t x_32; size_t x_33; size_t x_34; size_t x_35; size_t x_36; size_t x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_22 = lean_ctor_get(x_2, 0); -x_23 = lean_ctor_get(x_2, 1); -x_24 = lean_ctor_get(x_2, 2); -lean_inc(x_24); -lean_inc(x_23); -lean_inc(x_22); -lean_dec(x_2); -x_25 = lean_array_get_size(x_1); -x_26 = l_Lean_HeadIndex_hash(x_22); -x_27 = 32; -x_28 = lean_uint64_shift_right(x_26, x_27); -x_29 = lean_uint64_xor(x_26, x_28); -x_30 = 16; -x_31 = lean_uint64_shift_right(x_29, x_30); -x_32 = lean_uint64_xor(x_29, x_31); -x_33 = lean_uint64_to_usize(x_32); -x_34 = lean_usize_of_nat(x_25); -lean_dec(x_25); -x_35 = 1; -x_36 = lean_usize_sub(x_34, x_35); -x_37 = lean_usize_land(x_33, x_36); -x_38 = lean_array_uget(x_1, x_37); -x_39 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_39, 0, x_22); -lean_ctor_set(x_39, 1, x_23); -lean_ctor_set(x_39, 2, x_38); -x_40 = lean_array_uset(x_1, x_37, x_39); -x_1 = x_40; -x_2 = x_24; -goto _start; +uint8_t x_67; +x_67 = !lean_is_exclusive(x_1); +if (x_67 == 0) +{ +lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; +x_68 = lean_ctor_get(x_1, 1); +lean_dec(x_68); +x_69 = lean_ctor_get(x_1, 0); +lean_dec(x_69); +x_70 = lean_ctor_get(x_66, 0); +lean_inc(x_70); +lean_dec(x_66); +lean_inc(x_70); +x_71 = l_Array_eraseIdx___rarg(x_63, x_70, lean_box(0)); +x_72 = l_Array_eraseIdx___rarg(x_64, x_70, lean_box(0)); +lean_ctor_set(x_1, 1, x_72); +lean_ctor_set(x_1, 0, x_71); +return x_1; +} +else +{ +lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; +lean_dec(x_1); +x_73 = lean_ctor_get(x_66, 0); +lean_inc(x_73); +lean_dec(x_66); +lean_inc(x_73); +x_74 = l_Array_eraseIdx___rarg(x_63, x_73, lean_box(0)); +x_75 = l_Array_eraseIdx___rarg(x_64, x_73, lean_box(0)); +x_76 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_76, 0, x_74); +lean_ctor_set(x_76, 1, x_75); +return x_76; } } } } -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand_go___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_erase___at_Lean_Meta_Grind_EMatchTheorems_retrieve_x3f___spec__1(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_4; uint8_t x_5; -x_4 = lean_array_get_size(x_2); -x_5 = lean_nat_dec_lt(x_1, x_4); +uint64_t x_3; size_t x_4; lean_object* x_5; +x_3 = l_Lean_Name_hash___override(x_2); +x_4 = lean_uint64_to_usize(x_3); +x_5 = l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_EMatchTheorems_retrieve_x3f___spec__2(x_1, x_4, x_2); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_EMatchTheorems_retrieve_x3f(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = !lean_is_exclusive(x_1); +if (x_3 == 0) +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_4 = lean_ctor_get(x_1, 0); +x_5 = lean_ctor_get(x_1, 1); +x_6 = lean_ctor_get(x_1, 2); +lean_inc(x_4); +x_7 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__9(x_4, x_2); +if (lean_obj_tag(x_7) == 0) +{ +lean_object* x_8; +lean_free_object(x_1); +lean_dec(x_6); +lean_dec(x_5); lean_dec(x_4); -if (x_5 == 0) +x_8 = lean_box(0); +return x_8; +} +else { -lean_dec(x_2); -lean_dec(x_1); -return x_3; +uint8_t x_9; +x_9 = !lean_is_exclusive(x_7); +if (x_9 == 0) +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_10 = lean_ctor_get(x_7, 0); +x_11 = l_Lean_PersistentHashMap_erase___at_Lean_Meta_Grind_EMatchTheorems_retrieve_x3f___spec__1(x_4, x_2); +lean_ctor_set(x_1, 0, x_11); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_10); +lean_ctor_set(x_12, 1, x_1); +lean_ctor_set(x_7, 0, x_12); +return x_7; } else { -lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; -x_6 = lean_array_fget(x_2, x_1); -x_7 = lean_box(0); -x_8 = lean_array_fset(x_2, x_1, x_7); -x_9 = l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__4(x_3, x_6); -x_10 = lean_unsigned_to_nat(1u); -x_11 = lean_nat_add(x_1, x_10); +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_13 = lean_ctor_get(x_7, 0); +lean_inc(x_13); +lean_dec(x_7); +x_14 = l_Lean_PersistentHashMap_erase___at_Lean_Meta_Grind_EMatchTheorems_retrieve_x3f___spec__1(x_4, x_2); +lean_ctor_set(x_1, 0, x_14); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_13); +lean_ctor_set(x_15, 1, x_1); +x_16 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_16, 0, x_15); +return x_16; +} +} +} +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_17 = lean_ctor_get(x_1, 0); +x_18 = lean_ctor_get(x_1, 1); +x_19 = lean_ctor_get(x_1, 2); +lean_inc(x_19); +lean_inc(x_18); +lean_inc(x_17); lean_dec(x_1); -x_1 = x_11; -x_2 = x_8; -x_3 = x_9; -goto _start; +lean_inc(x_17); +x_20 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__9(x_17, x_2); +if (lean_obj_tag(x_20) == 0) +{ +lean_object* x_21; +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +x_21 = lean_box(0); +return x_21; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_22 = lean_ctor_get(x_20, 0); +lean_inc(x_22); +if (lean_is_exclusive(x_20)) { + lean_ctor_release(x_20, 0); + x_23 = x_20; +} else { + lean_dec_ref(x_20); + x_23 = lean_box(0); +} +x_24 = l_Lean_PersistentHashMap_erase___at_Lean_Meta_Grind_EMatchTheorems_retrieve_x3f___spec__1(x_17, x_2); +x_25 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_18); +lean_ctor_set(x_25, 2, x_19); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_22); +lean_ctor_set(x_26, 1, x_25); +if (lean_is_scalar(x_23)) { + x_27 = lean_alloc_ctor(1, 1, 0); +} else { + x_27 = x_23; } +lean_ctor_set(x_27, 0, x_26); +return x_27; } } -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__2(lean_object* x_1) { +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_EMatchTheorems_retrieve_x3f___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; -x_2 = lean_array_get_size(x_1); -x_3 = lean_unsigned_to_nat(2u); -x_4 = lean_nat_mul(x_2, x_3); +size_t x_4; lean_object* x_5; +x_4 = lean_unbox_usize(x_2); lean_dec(x_2); -x_5 = lean_box(0); -x_6 = lean_mk_array(x_4, x_5); -x_7 = lean_unsigned_to_nat(0u); -x_8 = l_Std_DHashMap_Internal_Raw_u2080_expand_go___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__3(x_7, x_1, x_6); -return x_8; +x_5 = l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_EMatchTheorems_retrieve_x3f___spec__2(x_1, x_4, x_3); +lean_dec(x_3); +return x_5; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_erase___at_Lean_Meta_Grind_EMatchTheorems_retrieve_x3f___spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; -x_8 = lean_st_ref_get(x_2, x_7); -x_9 = lean_ctor_get(x_8, 0); -lean_inc(x_9); -x_10 = lean_ctor_get(x_9, 1); -lean_inc(x_10); -lean_dec(x_9); -x_11 = !lean_is_exclusive(x_8); -if (x_11 == 0) +lean_object* x_3; +x_3 = l_Lean_PersistentHashMap_erase___at_Lean_Meta_Grind_EMatchTheorems_retrieve_x3f___spec__1(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_EMatchTheorems_retrieve_x3f___boxed(lean_object* x_1, lean_object* x_2) { +_start: { -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint64_t x_16; uint64_t x_17; uint64_t x_18; uint64_t x_19; uint64_t x_20; uint64_t x_21; uint64_t x_22; size_t x_23; size_t x_24; size_t x_25; size_t x_26; size_t x_27; lean_object* x_28; uint8_t x_29; -x_12 = lean_ctor_get(x_8, 1); -x_13 = lean_ctor_get(x_8, 0); -lean_dec(x_13); -x_14 = lean_ctor_get(x_10, 1); -lean_inc(x_14); -lean_dec(x_10); -x_15 = lean_array_get_size(x_14); -x_16 = l_Lean_HeadIndex_hash(x_1); -x_17 = 32; -x_18 = lean_uint64_shift_right(x_16, x_17); -x_19 = lean_uint64_xor(x_16, x_18); -x_20 = 16; -x_21 = lean_uint64_shift_right(x_19, x_20); -x_22 = lean_uint64_xor(x_19, x_21); -x_23 = lean_uint64_to_usize(x_22); -x_24 = lean_usize_of_nat(x_15); -lean_dec(x_15); -x_25 = 1; -x_26 = lean_usize_sub(x_24, x_25); -x_27 = lean_usize_land(x_23, x_26); -x_28 = lean_array_uget(x_14, x_27); -lean_dec(x_14); -x_29 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__1(x_1, x_28); -lean_dec(x_28); -if (x_29 == 0) +lean_object* x_3; +x_3 = l_Lean_Meta_Grind_EMatchTheorems_retrieve_x3f(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_EMatchTheorem_getProofWithFreshMVarLevels(lean_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_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; -lean_free_object(x_8); -x_30 = lean_st_ref_take(x_2, x_12); -x_31 = lean_ctor_get(x_30, 0); -lean_inc(x_31); -x_32 = lean_ctor_get(x_30, 1); -lean_inc(x_32); -lean_dec(x_30); -x_33 = !lean_is_exclusive(x_31); -if (x_33 == 0) +lean_object* x_7; uint8_t x_8; +x_7 = lean_ctor_get(x_1, 1); +lean_inc(x_7); +x_8 = l_Lean_Expr_isConst(x_7); +if (x_8 == 0) { -lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; -x_34 = lean_ctor_get(x_31, 0); -x_35 = lean_ctor_get(x_31, 1); -lean_inc(x_1); -x_36 = lean_array_push(x_34, x_1); -x_37 = !lean_is_exclusive(x_35); -if (x_37 == 0) +lean_object* x_9; uint8_t x_10; +x_9 = lean_ctor_get(x_1, 0); +lean_inc(x_9); +lean_dec(x_1); +x_10 = l_Array_isEmpty___rarg(x_9); +if (x_10 == 0) { -lean_object* x_38; lean_object* x_39; lean_object* x_40; size_t x_41; size_t x_42; size_t x_43; lean_object* x_44; uint8_t x_45; -x_38 = lean_ctor_get(x_35, 0); -x_39 = lean_ctor_get(x_35, 1); -x_40 = lean_array_get_size(x_39); -x_41 = lean_usize_of_nat(x_40); -lean_dec(x_40); -x_42 = lean_usize_sub(x_41, x_25); -x_43 = lean_usize_land(x_23, x_42); -x_44 = lean_array_uget(x_39, x_43); -x_45 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__1(x_1, x_44); -if (x_45 == 0) -{ -lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; uint8_t x_56; -x_46 = lean_unsigned_to_nat(1u); -x_47 = lean_nat_add(x_38, x_46); -lean_dec(x_38); -x_48 = lean_box(0); -x_49 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_49, 0, x_1); -lean_ctor_set(x_49, 1, x_48); -lean_ctor_set(x_49, 2, x_44); -x_50 = lean_array_uset(x_39, x_43, x_49); -x_51 = lean_unsigned_to_nat(4u); -x_52 = lean_nat_mul(x_47, x_51); -x_53 = lean_unsigned_to_nat(3u); -x_54 = lean_nat_div(x_52, x_53); -lean_dec(x_52); -x_55 = lean_array_get_size(x_50); -x_56 = lean_nat_dec_le(x_54, x_55); -lean_dec(x_55); -lean_dec(x_54); -if (x_56 == 0) +size_t x_11; size_t x_12; lean_object* x_13; uint8_t x_14; +x_11 = lean_array_size(x_9); +x_12 = 0; +lean_inc(x_9); +x_13 = l_Array_mapMUnsafe_map___at_Lean_Meta_openAbstractMVarsResult___spec__1(x_11, x_12, x_9, x_2, x_3, x_4, x_5, x_6); +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) { -lean_object* x_57; lean_object* x_58; uint8_t x_59; -x_57 = l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__2(x_50); -lean_ctor_set(x_35, 1, x_57); -lean_ctor_set(x_35, 0, x_47); -lean_ctor_set(x_31, 0, x_36); -x_58 = lean_st_ref_set(x_2, x_31, x_32); -x_59 = !lean_is_exclusive(x_58); -if (x_59 == 0) +lean_object* x_15; lean_object* x_16; +x_15 = lean_ctor_get(x_13, 0); +x_16 = l_Lean_Expr_instantiateLevelParamsArray(x_7, x_9, x_15); +lean_dec(x_7); +lean_ctor_set(x_13, 0, x_16); +return x_13; +} +else { -lean_object* x_60; -x_60 = lean_ctor_get(x_58, 0); -lean_dec(x_60); -lean_ctor_set(x_58, 0, x_48); -return x_58; +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_17 = lean_ctor_get(x_13, 0); +x_18 = lean_ctor_get(x_13, 1); +lean_inc(x_18); +lean_inc(x_17); +lean_dec(x_13); +x_19 = l_Lean_Expr_instantiateLevelParamsArray(x_7, x_9, x_17); +lean_dec(x_7); +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_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(0, 2, 0); -lean_ctor_set(x_62, 0, x_48); -lean_ctor_set(x_62, 1, x_61); -return x_62; +lean_object* x_21; +lean_dec(x_9); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_7); +lean_ctor_set(x_21, 1, x_6); +return x_21; } } else { -lean_object* x_63; uint8_t x_64; -lean_ctor_set(x_35, 1, x_50); -lean_ctor_set(x_35, 0, x_47); -lean_ctor_set(x_31, 0, x_36); -x_63 = lean_st_ref_set(x_2, x_31, x_32); -x_64 = !lean_is_exclusive(x_63); -if (x_64 == 0) +lean_object* x_22; uint8_t x_23; +x_22 = lean_ctor_get(x_1, 0); +lean_inc(x_22); +lean_dec(x_1); +x_23 = l_Array_isEmpty___rarg(x_22); +if (x_23 == 0) { -lean_object* x_65; -x_65 = lean_ctor_get(x_63, 0); -lean_dec(x_65); -lean_ctor_set(x_63, 0, x_48); -return x_63; +size_t x_24; size_t x_25; lean_object* x_26; uint8_t x_27; +x_24 = lean_array_size(x_22); +x_25 = 0; +lean_inc(x_22); +x_26 = l_Array_mapMUnsafe_map___at_Lean_Meta_openAbstractMVarsResult___spec__1(x_24, x_25, x_22, x_2, x_3, x_4, x_5, x_6); +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); +x_29 = l_Lean_Expr_instantiateLevelParamsArray(x_7, x_22, x_28); +lean_dec(x_7); +lean_ctor_set(x_26, 0, x_29); +return x_26; } 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(0, 2, 0); -lean_ctor_set(x_67, 0, x_48); -lean_ctor_set(x_67, 1, x_66); -return x_67; +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_30 = lean_ctor_get(x_26, 0); +x_31 = lean_ctor_get(x_26, 1); +lean_inc(x_31); +lean_inc(x_30); +lean_dec(x_26); +x_32 = l_Lean_Expr_instantiateLevelParamsArray(x_7, x_22, x_30); +lean_dec(x_7); +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_31); +return x_33; +} +} +else +{ +lean_object* x_34; lean_object* x_35; +lean_dec(x_22); +x_34 = l_Lean_Expr_constName_x21(x_7); +lean_inc(x_34); +x_35 = l_Lean_getConstInfo___at_Lean_Meta_mkConstWithFreshMVarLevels___spec__1(x_34, x_2, x_3, x_4, x_5, x_6); +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; lean_object* x_39; uint8_t x_40; +x_37 = lean_ctor_get(x_35, 0); +x_38 = lean_ctor_get(x_35, 1); +x_39 = l_Lean_ConstantInfo_levelParams(x_37); +lean_dec(x_37); +x_40 = l_List_isEmpty___rarg(x_39); +lean_dec(x_39); +if (x_40 == 0) +{ +lean_object* x_41; +lean_free_object(x_35); +lean_dec(x_7); +x_41 = l_Lean_Meta_mkConstWithFreshMVarLevels(x_34, x_2, x_3, x_4, x_5, x_38); +return x_41; } +else +{ +lean_dec(x_34); +lean_ctor_set(x_35, 0, x_7); +return x_35; } } else { -lean_object* x_68; uint8_t x_69; +lean_object* x_42; lean_object* x_43; lean_object* x_44; uint8_t x_45; +x_42 = lean_ctor_get(x_35, 0); +x_43 = lean_ctor_get(x_35, 1); +lean_inc(x_43); +lean_inc(x_42); +lean_dec(x_35); +x_44 = l_Lean_ConstantInfo_levelParams(x_42); +lean_dec(x_42); +x_45 = l_List_isEmpty___rarg(x_44); lean_dec(x_44); -lean_dec(x_1); -lean_ctor_set(x_31, 0, x_36); -x_68 = lean_st_ref_set(x_2, x_31, x_32); -x_69 = !lean_is_exclusive(x_68); -if (x_69 == 0) +if (x_45 == 0) { -lean_object* x_70; lean_object* x_71; -x_70 = lean_ctor_get(x_68, 0); -lean_dec(x_70); -x_71 = lean_box(0); -lean_ctor_set(x_68, 0, x_71); -return x_68; +lean_object* x_46; +lean_dec(x_7); +x_46 = l_Lean_Meta_mkConstWithFreshMVarLevels(x_34, x_2, x_3, x_4, x_5, x_43); +return x_46; } else { -lean_object* x_72; lean_object* x_73; lean_object* x_74; -x_72 = lean_ctor_get(x_68, 1); -lean_inc(x_72); -lean_dec(x_68); -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_72); -return x_74; +lean_object* x_47; +lean_dec(x_34); +x_47 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_47, 0, x_7); +lean_ctor_set(x_47, 1, x_43); +return x_47; } } } else { -lean_object* x_75; lean_object* x_76; lean_object* x_77; size_t x_78; size_t x_79; size_t x_80; lean_object* x_81; uint8_t x_82; -x_75 = lean_ctor_get(x_35, 0); -x_76 = lean_ctor_get(x_35, 1); -lean_inc(x_76); -lean_inc(x_75); -lean_dec(x_35); -x_77 = lean_array_get_size(x_76); -x_78 = lean_usize_of_nat(x_77); -lean_dec(x_77); -x_79 = lean_usize_sub(x_78, x_25); -x_80 = lean_usize_land(x_23, x_79); -x_81 = lean_array_uget(x_76, x_80); -x_82 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__1(x_1, x_81); -if (x_82 == 0) +uint8_t x_48; +lean_dec(x_34); +lean_dec(x_7); +x_48 = !lean_is_exclusive(x_35); +if (x_48 == 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; uint8_t x_93; -x_83 = lean_unsigned_to_nat(1u); -x_84 = lean_nat_add(x_75, x_83); -lean_dec(x_75); -x_85 = lean_box(0); -x_86 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_86, 0, x_1); -lean_ctor_set(x_86, 1, x_85); -lean_ctor_set(x_86, 2, x_81); -x_87 = lean_array_uset(x_76, x_80, x_86); -x_88 = lean_unsigned_to_nat(4u); -x_89 = lean_nat_mul(x_84, x_88); -x_90 = lean_unsigned_to_nat(3u); -x_91 = lean_nat_div(x_89, x_90); -lean_dec(x_89); -x_92 = lean_array_get_size(x_87); -x_93 = lean_nat_dec_le(x_91, x_92); -lean_dec(x_92); -lean_dec(x_91); -if (x_93 == 0) +return x_35; +} +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; -x_94 = l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__2(x_87); -x_95 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_95, 0, x_84); -lean_ctor_set(x_95, 1, x_94); -lean_ctor_set(x_31, 1, x_95); -lean_ctor_set(x_31, 0, x_36); -x_96 = lean_st_ref_set(x_2, x_31, x_32); -x_97 = lean_ctor_get(x_96, 1); -lean_inc(x_97); -if (lean_is_exclusive(x_96)) { - lean_ctor_release(x_96, 0); - lean_ctor_release(x_96, 1); - x_98 = x_96; -} else { - lean_dec_ref(x_96); - x_98 = lean_box(0); +lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_49 = lean_ctor_get(x_35, 0); +x_50 = lean_ctor_get(x_35, 1); +lean_inc(x_50); +lean_inc(x_49); +lean_dec(x_35); +x_51 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_51, 0, x_49); +lean_ctor_set(x_51, 1, x_50); +return x_51; } -if (lean_is_scalar(x_98)) { - x_99 = lean_alloc_ctor(0, 2, 0); -} else { - x_99 = x_98; } -lean_ctor_set(x_99, 0, x_85); -lean_ctor_set(x_99, 1, x_97); -return x_99; } -else -{ -lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; -x_100 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_100, 0, x_84); -lean_ctor_set(x_100, 1, x_87); -lean_ctor_set(x_31, 1, x_100); -lean_ctor_set(x_31, 0, x_36); -x_101 = lean_st_ref_set(x_2, x_31, x_32); -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; -} else { - lean_dec_ref(x_101); - x_103 = lean_box(0); } -if (lean_is_scalar(x_103)) { - x_104 = lean_alloc_ctor(0, 2, 0); -} else { - x_104 = x_103; } -lean_ctor_set(x_104, 0, x_85); -lean_ctor_set(x_104, 1, x_102); -return x_104; } +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_EMatchTheorem_getProofWithFreshMVarLevels___boxed(lean_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_Grind_EMatchTheorem_getProofWithFreshMVarLevels(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; } -else +} +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__1() { +_start: { -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_dec(x_81); -lean_dec(x_1); -x_105 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_105, 0, x_75); -lean_ctor_set(x_105, 1, x_76); -lean_ctor_set(x_31, 1, x_105); -lean_ctor_set(x_31, 0, x_36); -x_106 = lean_st_ref_set(x_2, x_31, x_32); -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); +lean_object* x_1; +x_1 = lean_mk_string_unchecked("_private", 8, 8); +return x_1; } -x_109 = lean_box(0); -if (lean_is_scalar(x_108)) { - x_110 = lean_alloc_ctor(0, 2, 0); -} else { - x_110 = x_108; } -lean_ctor_set(x_110, 0, x_109); -lean_ctor_set(x_110, 1, x_107); -return x_110; +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____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_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} } +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__2; +x_2 = l_Lean_Meta_Grind_mkOffsetPattern___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } } -else +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__4() { +_start: { -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; size_t x_119; size_t x_120; size_t x_121; lean_object* x_122; uint8_t x_123; -x_111 = lean_ctor_get(x_31, 0); -x_112 = lean_ctor_get(x_31, 1); -x_113 = lean_ctor_get(x_31, 2); -lean_inc(x_113); -lean_inc(x_112); -lean_inc(x_111); -lean_dec(x_31); -lean_inc(x_1); -x_114 = lean_array_push(x_111, x_1); -x_115 = lean_ctor_get(x_112, 0); -lean_inc(x_115); -x_116 = lean_ctor_get(x_112, 1); -lean_inc(x_116); -if (lean_is_exclusive(x_112)) { - lean_ctor_release(x_112, 0); - lean_ctor_release(x_112, 1); - x_117 = x_112; -} else { - lean_dec_ref(x_112); - x_117 = lean_box(0); +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Meta", 4, 4); +return x_1; } -x_118 = lean_array_get_size(x_116); -x_119 = lean_usize_of_nat(x_118); -lean_dec(x_118); -x_120 = lean_usize_sub(x_119, x_25); -x_121 = lean_usize_land(x_23, x_120); -x_122 = lean_array_uget(x_116, x_121); -x_123 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__1(x_1, x_122); -if (x_123 == 0) +} +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__5() { +_start: { -lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; uint8_t x_134; -x_124 = lean_unsigned_to_nat(1u); -x_125 = lean_nat_add(x_115, x_124); -lean_dec(x_115); -x_126 = lean_box(0); -x_127 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_127, 0, x_1); -lean_ctor_set(x_127, 1, x_126); -lean_ctor_set(x_127, 2, x_122); -x_128 = lean_array_uset(x_116, x_121, x_127); -x_129 = lean_unsigned_to_nat(4u); -x_130 = lean_nat_mul(x_125, x_129); -x_131 = lean_unsigned_to_nat(3u); -x_132 = lean_nat_div(x_130, x_131); -lean_dec(x_130); -x_133 = lean_array_get_size(x_128); -x_134 = lean_nat_dec_le(x_132, x_133); -lean_dec(x_133); -lean_dec(x_132); -if (x_134 == 0) +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__3; +x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__4; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__6() { +_start: { -lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; -x_135 = l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__2(x_128); -if (lean_is_scalar(x_117)) { - x_136 = lean_alloc_ctor(0, 2, 0); -} else { - x_136 = x_117; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Tactic", 6, 6); +return x_1; } -lean_ctor_set(x_136, 0, x_125); -lean_ctor_set(x_136, 1, x_135); -x_137 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_137, 0, x_114); -lean_ctor_set(x_137, 1, x_136); -lean_ctor_set(x_137, 2, x_113); -x_138 = lean_st_ref_set(x_2, x_137, x_32); -x_139 = lean_ctor_get(x_138, 1); -lean_inc(x_139); -if (lean_is_exclusive(x_138)) { - lean_ctor_release(x_138, 0); - lean_ctor_release(x_138, 1); - x_140 = x_138; -} else { - lean_dec_ref(x_138); - x_140 = lean_box(0); } -if (lean_is_scalar(x_140)) { - x_141 = lean_alloc_ctor(0, 2, 0); -} else { - x_141 = x_140; +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__5; +x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__6; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } -lean_ctor_set(x_141, 0, x_126); -lean_ctor_set(x_141, 1, x_139); -return x_141; } -else +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__8() { +_start: { -lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; -if (lean_is_scalar(x_117)) { - x_142 = lean_alloc_ctor(0, 2, 0); -} else { - x_142 = x_117; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__7; +x_2 = l_Lean_Meta_Grind_mkOffsetPattern___closed__2; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } -lean_ctor_set(x_142, 0, x_125); -lean_ctor_set(x_142, 1, x_128); -x_143 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_143, 0, x_114); -lean_ctor_set(x_143, 1, x_142); -lean_ctor_set(x_143, 2, x_113); -x_144 = lean_st_ref_set(x_2, x_143, x_32); -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; +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__9() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("EMatchTheorem", 13, 13); +return x_1; } -lean_ctor_set(x_147, 0, x_126); -lean_ctor_set(x_147, 1, x_145); -return x_147; +} +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__8; +x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__9; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } } -else +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__11() { +_start: { -lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; -lean_dec(x_122); -lean_dec(x_1); -if (lean_is_scalar(x_117)) { - x_148 = lean_alloc_ctor(0, 2, 0); -} else { - x_148 = x_117; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__10; +x_2 = lean_unsigned_to_nat(0u); +x_3 = l_Lean_Name_num___override(x_1, x_2); +return x_3; } -lean_ctor_set(x_148, 0, x_115); -lean_ctor_set(x_148, 1, x_116); -x_149 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_149, 0, x_114); -lean_ctor_set(x_149, 1, x_148); -lean_ctor_set(x_149, 2, x_113); -x_150 = lean_st_ref_set(x_2, x_149, x_32); -x_151 = lean_ctor_get(x_150, 1); -lean_inc(x_151); -if (lean_is_exclusive(x_150)) { - lean_ctor_release(x_150, 0); - lean_ctor_release(x_150, 1); - x_152 = x_150; -} else { - lean_dec_ref(x_150); - x_152 = lean_box(0); } -x_153 = lean_box(0); -if (lean_is_scalar(x_152)) { - x_154 = lean_alloc_ctor(0, 2, 0); -} else { - x_154 = x_152; +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__11; +x_2 = l_Lean_Meta_Grind_mkOffsetPattern___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } -lean_ctor_set(x_154, 0, x_153); -lean_ctor_set(x_154, 1, x_151); -return x_154; } +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__12; +x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__4; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } } -else +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__14() { +_start: { -lean_object* x_155; -lean_dec(x_1); -x_155 = lean_box(0); -lean_ctor_set(x_8, 0, x_155); -return x_8; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__13; +x_2 = l_Lean_Meta_Grind_mkOffsetPattern___closed__2; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } } -else +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__15() { +_start: { -lean_object* x_156; lean_object* x_157; lean_object* x_158; uint64_t x_159; uint64_t x_160; uint64_t x_161; uint64_t x_162; uint64_t x_163; uint64_t x_164; uint64_t x_165; size_t x_166; size_t x_167; size_t x_168; size_t x_169; size_t x_170; lean_object* x_171; uint8_t x_172; -x_156 = lean_ctor_get(x_8, 1); -lean_inc(x_156); -lean_dec(x_8); -x_157 = lean_ctor_get(x_10, 1); -lean_inc(x_157); -lean_dec(x_10); -x_158 = lean_array_get_size(x_157); -x_159 = l_Lean_HeadIndex_hash(x_1); -x_160 = 32; -x_161 = lean_uint64_shift_right(x_159, x_160); -x_162 = lean_uint64_xor(x_159, x_161); -x_163 = 16; -x_164 = lean_uint64_shift_right(x_162, x_163); -x_165 = lean_uint64_xor(x_162, x_164); -x_166 = lean_uint64_to_usize(x_165); -x_167 = lean_usize_of_nat(x_158); -lean_dec(x_158); -x_168 = 1; -x_169 = lean_usize_sub(x_167, x_168); -x_170 = lean_usize_land(x_166, x_169); -x_171 = lean_array_uget(x_157, x_170); -lean_dec(x_157); -x_172 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__1(x_1, x_171); -lean_dec(x_171); -if (x_172 == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("ematchTheoremsExt", 17, 17); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__16() { +_start: { -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; size_t x_185; size_t x_186; size_t x_187; lean_object* x_188; uint8_t x_189; -x_173 = lean_st_ref_take(x_2, x_156); -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); -if (lean_is_exclusive(x_174)) { - lean_ctor_release(x_174, 0); - lean_ctor_release(x_174, 1); - lean_ctor_release(x_174, 2); - x_179 = x_174; -} else { - lean_dec_ref(x_174); - x_179 = lean_box(0); +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__14; +x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__15; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } -lean_inc(x_1); -x_180 = lean_array_push(x_176, x_1); -x_181 = lean_ctor_get(x_177, 0); -lean_inc(x_181); -x_182 = lean_ctor_get(x_177, 1); -lean_inc(x_182); -if (lean_is_exclusive(x_177)) { - lean_ctor_release(x_177, 0); - lean_ctor_release(x_177, 1); - x_183 = x_177; -} else { - lean_dec_ref(x_177); - x_183 = lean_box(0); } -x_184 = lean_array_get_size(x_182); -x_185 = lean_usize_of_nat(x_184); -lean_dec(x_184); -x_186 = lean_usize_sub(x_185, x_168); -x_187 = lean_usize_land(x_166, x_186); -x_188 = lean_array_uget(x_182, x_187); -x_189 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__1(x_1, x_188); -if (x_189 == 0) +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__17() { +_start: { -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; uint8_t x_200; -x_190 = lean_unsigned_to_nat(1u); -x_191 = lean_nat_add(x_181, x_190); -lean_dec(x_181); -x_192 = lean_box(0); -x_193 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_193, 0, x_1); -lean_ctor_set(x_193, 1, x_192); -lean_ctor_set(x_193, 2, x_188); -x_194 = lean_array_uset(x_182, x_187, x_193); -x_195 = lean_unsigned_to_nat(4u); -x_196 = lean_nat_mul(x_191, x_195); -x_197 = lean_unsigned_to_nat(3u); -x_198 = lean_nat_div(x_196, x_197); -lean_dec(x_196); -x_199 = lean_array_get_size(x_194); -x_200 = lean_nat_dec_le(x_198, x_199); -lean_dec(x_199); -lean_dec(x_198); -if (x_200 == 0) +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedEMatchTheorems___spec__1___closed__2; +x_2 = l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedEMatchTheorems___spec__1; +x_3 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +lean_ctor_set(x_3, 2, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__18() { +_start: { -lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; -x_201 = l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__2(x_194); -if (lean_is_scalar(x_183)) { - x_202 = lean_alloc_ctor(0, 2, 0); -} else { - x_202 = x_183; +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_EMatchTheorems_insert), 2, 0); +return x_1; } -lean_ctor_set(x_202, 0, x_191); -lean_ctor_set(x_202, 1, x_201); -if (lean_is_scalar(x_179)) { - x_203 = lean_alloc_ctor(0, 3, 0); -} else { - x_203 = x_179; } -lean_ctor_set(x_203, 0, x_180); -lean_ctor_set(x_203, 1, x_202); -lean_ctor_set(x_203, 2, x_178); -x_204 = lean_st_ref_set(x_2, x_203, x_175); -x_205 = lean_ctor_get(x_204, 1); -lean_inc(x_205); -if (lean_is_exclusive(x_204)) { - lean_ctor_release(x_204, 0); - lean_ctor_release(x_204, 1); - x_206 = x_204; -} else { - lean_dec_ref(x_204); - x_206 = lean_box(0); +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__19() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_id___rarg___boxed), 1, 0); +return x_1; } -if (lean_is_scalar(x_206)) { - x_207 = lean_alloc_ctor(0, 2, 0); -} else { - x_207 = x_206; } -lean_ctor_set(x_207, 0, x_192); -lean_ctor_set(x_207, 1, x_205); -return x_207; +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__20() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__16; +x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__18; +x_3 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__17; +x_4 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__19; +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; } -else +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666_(lean_object* x_1) { +_start: { -lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; -if (lean_is_scalar(x_183)) { - x_208 = lean_alloc_ctor(0, 2, 0); -} else { - x_208 = x_183; +lean_object* x_2; lean_object* x_3; +x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__20; +x_3 = l_Lean_registerSimpleScopedEnvExtension___rarg(x_2, x_1); +return x_3; } -lean_ctor_set(x_208, 0, x_191); -lean_ctor_set(x_208, 1, x_194); -if (lean_is_scalar(x_179)) { - x_209 = lean_alloc_ctor(0, 3, 0); -} else { - x_209 = x_179; } -lean_ctor_set(x_209, 0, x_180); -lean_ctor_set(x_209, 1, x_208); -lean_ctor_set(x_209, 2, x_178); -x_210 = lean_st_ref_set(x_2, x_209, x_175); -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_210); - x_212 = lean_box(0); +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Eq", 2, 2); +return x_1; } -if (lean_is_scalar(x_212)) { - x_213 = lean_alloc_ctor(0, 2, 0); -} else { - x_213 = x_212; } -lean_ctor_set(x_213, 0, x_192); -lean_ctor_set(x_213, 1, x_211); -return x_213; +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } } -else +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__3() { +_start: { -lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; -lean_dec(x_188); -lean_dec(x_1); -if (lean_is_scalar(x_183)) { - x_214 = lean_alloc_ctor(0, 2, 0); -} else { - x_214 = x_183; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("HEq", 3, 3); +return x_1; } -lean_ctor_set(x_214, 0, x_181); -lean_ctor_set(x_214, 1, x_182); -if (lean_is_scalar(x_179)) { - x_215 = lean_alloc_ctor(0, 3, 0); -} else { - x_215 = x_179; } -lean_ctor_set(x_215, 0, x_180); -lean_ctor_set(x_215, 1, x_214); -lean_ctor_set(x_215, 2, x_178); -x_216 = lean_st_ref_set(x_2, x_215, x_175); -x_217 = lean_ctor_get(x_216, 1); -lean_inc(x_217); -if (lean_is_exclusive(x_216)) { - lean_ctor_release(x_216, 0); - lean_ctor_release(x_216, 1); - x_218 = x_216; -} else { - lean_dec_ref(x_216); - x_218 = lean_box(0); +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___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_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__3; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } -x_219 = lean_box(0); -if (lean_is_scalar(x_218)) { - x_220 = lean_alloc_ctor(0, 2, 0); -} else { - x_220 = x_218; } -lean_ctor_set(x_220, 0, x_219); -lean_ctor_set(x_220, 1, x_217); -return x_220; +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Iff", 3, 3); +return x_1; } } -else +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__6() { +_start: { -lean_object* x_221; lean_object* x_222; -lean_dec(x_1); -x_221 = lean_box(0); -x_222 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_222, 0, x_221); -lean_ctor_set(x_222, 1, x_156); -return x_222; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__5; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } } +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("And", 3, 3); +return x_1; } } -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__1___boxed(lean_object* x_1, lean_object* x_2) { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__8() { _start: { -uint8_t x_3; lean_object* x_4; -x_3 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__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_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__7; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__9() { _start: { -lean_object* x_8; -x_8 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol(x_1, x_2, x_3, x_4, x_5, x_6, x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_8; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Or", 2, 2); +return x_1; } } -LEAN_EXPORT uint8_t l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_foundBVar___spec__1(lean_object* x_1, lean_object* x_2) { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__10() { _start: { -if (lean_obj_tag(x_2) == 0) -{ -uint8_t x_3; -x_3 = 0; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__9; +x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } -else -{ -lean_object* x_4; lean_object* x_5; uint8_t x_6; -x_4 = lean_ctor_get(x_2, 0); -x_5 = lean_ctor_get(x_2, 2); -x_6 = lean_nat_dec_eq(x_4, x_1); -if (x_6 == 0) -{ -x_2 = x_5; -goto _start; } -else +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__11() { +_start: { -uint8_t x_8; -x_8 = 1; -return x_8; -} -} +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Not", 3, 3); +return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_foundBVar(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__12() { _start: { -lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; -x_8 = lean_st_ref_get(x_2, x_7); -x_9 = lean_ctor_get(x_8, 0); -lean_inc(x_9); -x_10 = lean_ctor_get(x_9, 2); -lean_inc(x_10); -lean_dec(x_9); -x_11 = !lean_is_exclusive(x_8); -if (x_11 == 0) -{ -lean_object* x_12; lean_object* x_13; lean_object* x_14; uint64_t x_15; uint64_t x_16; uint64_t x_17; uint64_t x_18; uint64_t x_19; uint64_t x_20; uint64_t x_21; size_t x_22; size_t x_23; size_t x_24; size_t x_25; size_t x_26; lean_object* x_27; uint8_t x_28; lean_object* x_29; -x_12 = lean_ctor_get(x_8, 0); -lean_dec(x_12); -x_13 = lean_ctor_get(x_10, 1); -lean_inc(x_13); -lean_dec(x_10); -x_14 = lean_array_get_size(x_13); -x_15 = lean_uint64_of_nat(x_1); -x_16 = 32; -x_17 = lean_uint64_shift_right(x_15, x_16); -x_18 = lean_uint64_xor(x_15, x_17); -x_19 = 16; -x_20 = lean_uint64_shift_right(x_18, x_19); -x_21 = lean_uint64_xor(x_18, x_20); -x_22 = lean_uint64_to_usize(x_21); -x_23 = lean_usize_of_nat(x_14); -lean_dec(x_14); -x_24 = 1; -x_25 = lean_usize_sub(x_23, x_24); -x_26 = lean_usize_land(x_22, x_25); -x_27 = lean_array_uget(x_13, x_26); -lean_dec(x_13); -x_28 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_foundBVar___spec__1(x_1, x_27); -lean_dec(x_27); -x_29 = lean_box(x_28); -lean_ctor_set(x_8, 0, x_29); -return x_8; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__11; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } -else -{ -lean_object* x_30; lean_object* x_31; lean_object* x_32; uint64_t x_33; uint64_t x_34; uint64_t x_35; uint64_t x_36; uint64_t x_37; uint64_t x_38; uint64_t x_39; size_t x_40; size_t x_41; size_t x_42; size_t x_43; size_t x_44; lean_object* x_45; uint8_t x_46; lean_object* x_47; lean_object* x_48; -x_30 = lean_ctor_get(x_8, 1); -lean_inc(x_30); -lean_dec(x_8); -x_31 = lean_ctor_get(x_10, 1); -lean_inc(x_31); -lean_dec(x_10); -x_32 = lean_array_get_size(x_31); -x_33 = lean_uint64_of_nat(x_1); -x_34 = 32; -x_35 = lean_uint64_shift_right(x_33, x_34); -x_36 = lean_uint64_xor(x_33, x_35); -x_37 = 16; -x_38 = lean_uint64_shift_right(x_36, x_37); -x_39 = lean_uint64_xor(x_36, x_38); -x_40 = lean_uint64_to_usize(x_39); -x_41 = lean_usize_of_nat(x_32); -lean_dec(x_32); -x_42 = 1; -x_43 = lean_usize_sub(x_41, x_42); -x_44 = lean_usize_land(x_40, x_43); -x_45 = lean_array_uget(x_31, x_44); -lean_dec(x_31); -x_46 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_foundBVar___spec__1(x_1, x_45); -lean_dec(x_45); -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_30); -return x_48; } +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___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_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__12; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; } } -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_foundBVar___spec__1___boxed(lean_object* x_1, lean_object* x_2) { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__14() { _start: { -uint8_t x_3; lean_object* x_4; -x_3 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_foundBVar___spec__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_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__10; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__13; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_foundBVar___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__15() { _start: { -lean_object* x_8; -x_8 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_foundBVar(x_1, x_2, x_3, x_4, x_5, x_6, x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_8; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__8; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__14; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; } } -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__16() { _start: { -if (lean_obj_tag(x_3) == 0) -{ -lean_dec(x_1); -return x_2; -} -else -{ -uint8_t x_4; -x_4 = !lean_is_exclusive(x_3); -if (x_4 == 0) -{ -lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint64_t x_9; uint64_t x_10; uint64_t x_11; uint64_t x_12; uint64_t x_13; uint64_t x_14; uint64_t x_15; size_t x_16; size_t x_17; size_t x_18; size_t x_19; size_t x_20; lean_object* x_21; lean_object* x_22; -x_5 = lean_ctor_get(x_3, 0); -x_6 = lean_ctor_get(x_3, 2); -x_7 = lean_array_get_size(x_2); -lean_inc(x_1); -lean_inc(x_5); -x_8 = lean_apply_1(x_1, x_5); -x_9 = lean_unbox_uint64(x_8); -lean_dec(x_8); -x_10 = 32; -x_11 = lean_uint64_shift_right(x_9, x_10); -x_12 = lean_uint64_xor(x_9, x_11); -x_13 = 16; -x_14 = lean_uint64_shift_right(x_12, x_13); -x_15 = lean_uint64_xor(x_12, x_14); -x_16 = lean_uint64_to_usize(x_15); -x_17 = lean_usize_of_nat(x_7); -lean_dec(x_7); -x_18 = 1; -x_19 = lean_usize_sub(x_17, x_18); -x_20 = lean_usize_land(x_16, x_19); -x_21 = lean_array_uget(x_2, x_20); -lean_ctor_set(x_3, 2, x_21); -x_22 = lean_array_uset(x_2, x_20, x_3); -x_2 = x_22; -x_3 = x_6; -goto _start; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__6; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__15; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; } -else +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__17() { +_start: { -lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; uint64_t x_29; uint64_t x_30; uint64_t x_31; uint64_t x_32; uint64_t x_33; uint64_t x_34; uint64_t x_35; size_t x_36; size_t x_37; size_t x_38; size_t x_39; size_t x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; -x_24 = lean_ctor_get(x_3, 0); -x_25 = lean_ctor_get(x_3, 1); -x_26 = lean_ctor_get(x_3, 2); -lean_inc(x_26); -lean_inc(x_25); -lean_inc(x_24); -lean_dec(x_3); -x_27 = lean_array_get_size(x_2); -lean_inc(x_1); -lean_inc(x_24); -x_28 = lean_apply_1(x_1, x_24); -x_29 = lean_unbox_uint64(x_28); -lean_dec(x_28); -x_30 = 32; -x_31 = lean_uint64_shift_right(x_29, x_30); -x_32 = lean_uint64_xor(x_29, x_31); -x_33 = 16; -x_34 = lean_uint64_shift_right(x_32, x_33); -x_35 = lean_uint64_xor(x_32, x_34); -x_36 = lean_uint64_to_usize(x_35); -x_37 = lean_usize_of_nat(x_27); -lean_dec(x_27); -x_38 = 1; -x_39 = lean_usize_sub(x_37, x_38); -x_40 = lean_usize_land(x_36, x_39); -x_41 = lean_array_uget(x_2, x_40); -x_42 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_42, 0, x_24); -lean_ctor_set(x_42, 1, x_25); -lean_ctor_set(x_42, 2, x_41); -x_43 = lean_array_uset(x_2, x_40, x_42); -x_2 = x_43; -x_3 = x_26; -goto _start; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__4; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__16; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; } } +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__18() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__2; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__17; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; } } -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__3___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__4(lean_object* x_1, lean_object* x_2) { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__19() { _start: { -if (lean_obj_tag(x_2) == 0) +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__18; +x_2 = lean_array_mk(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames() { +_start: { +lean_object* x_1; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__19; return x_1; } -else -{ -uint8_t x_3; -x_3 = !lean_is_exclusive(x_2); -if (x_3 == 0) +} +LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isForbidden(lean_object* x_1) { +_start: { -lean_object* x_4; lean_object* x_5; lean_object* x_6; uint64_t x_7; uint64_t x_8; uint64_t x_9; uint64_t x_10; uint64_t x_11; uint64_t x_12; uint64_t x_13; size_t x_14; size_t x_15; size_t x_16; size_t x_17; size_t x_18; lean_object* x_19; lean_object* x_20; -x_4 = lean_ctor_get(x_2, 0); -x_5 = lean_ctor_get(x_2, 2); -x_6 = lean_array_get_size(x_1); -x_7 = lean_uint64_of_nat(x_4); -x_8 = 32; -x_9 = lean_uint64_shift_right(x_7, x_8); -x_10 = lean_uint64_xor(x_7, x_9); -x_11 = 16; -x_12 = lean_uint64_shift_right(x_10, x_11); -x_13 = lean_uint64_xor(x_10, x_12); -x_14 = lean_uint64_to_usize(x_13); -x_15 = lean_usize_of_nat(x_6); -lean_dec(x_6); -x_16 = 1; -x_17 = lean_usize_sub(x_15, x_16); -x_18 = lean_usize_land(x_14, x_17); -x_19 = lean_array_uget(x_1, x_18); -lean_ctor_set(x_2, 2, x_19); -x_20 = lean_array_uset(x_1, x_18, x_2); -x_1 = x_20; -x_2 = x_5; -goto _start; +lean_object* x_2; uint8_t x_3; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames; +x_3 = l_Array_contains___at_Lean_registerInternalExceptionId___spec__1(x_2, x_1); +return x_3; } -else +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isForbidden___boxed(lean_object* x_1) { +_start: { -lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint64_t x_26; uint64_t x_27; uint64_t x_28; uint64_t x_29; uint64_t x_30; uint64_t x_31; uint64_t x_32; size_t x_33; size_t x_34; size_t x_35; size_t x_36; size_t x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_22 = lean_ctor_get(x_2, 0); -x_23 = lean_ctor_get(x_2, 1); -x_24 = lean_ctor_get(x_2, 2); -lean_inc(x_24); -lean_inc(x_23); -lean_inc(x_22); -lean_dec(x_2); -x_25 = lean_array_get_size(x_1); -x_26 = lean_uint64_of_nat(x_22); -x_27 = 32; -x_28 = lean_uint64_shift_right(x_26, x_27); -x_29 = lean_uint64_xor(x_26, x_28); -x_30 = 16; -x_31 = lean_uint64_shift_right(x_29, x_30); -x_32 = lean_uint64_xor(x_29, x_31); -x_33 = lean_uint64_to_usize(x_32); -x_34 = lean_usize_of_nat(x_25); -lean_dec(x_25); -x_35 = 1; -x_36 = lean_usize_sub(x_34, x_35); -x_37 = lean_usize_land(x_33, x_36); -x_38 = lean_array_uget(x_1, x_37); -x_39 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_39, 0, x_22); -lean_ctor_set(x_39, 1, x_23); -lean_ctor_set(x_39, 2, x_38); -x_40 = lean_array_uset(x_1, x_37, x_39); -x_1 = x_40; -x_2 = x_24; -goto _start; +uint8_t x_2; lean_object* x_3; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isForbidden(x_1); +lean_dec(x_1); +x_3 = lean_box(x_2); +return x_3; } } +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("[grind_dontcare]", 16, 16); +return x_1; } } -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand_go___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare___closed__2() { _start: { -lean_object* x_4; uint8_t x_5; -x_4 = lean_array_get_size(x_2); -x_5 = lean_nat_dec_lt(x_1, x_4); -lean_dec(x_4); -if (x_5 == 0) +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare___closed__3() { +_start: { -lean_dec(x_2); -lean_dec(x_1); +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare___closed__2; +x_3 = l_Lean_Expr_const___override(x_2, x_1); return x_3; } -else +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare() { +_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; -x_6 = lean_array_fget(x_2, x_1); -x_7 = lean_box(0); -x_8 = lean_array_fset(x_2, x_1, x_7); -x_9 = l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__3___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__4(x_3, x_6); -x_10 = lean_unsigned_to_nat(1u); -x_11 = lean_nat_add(x_1, x_10); -lean_dec(x_1); -x_1 = x_11; -x_2 = x_8; -x_3 = x_9; -goto _start; +lean_object* x_1; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare___closed__3; +return x_1; } } +static lean_object* _init_l_Lean_Meta_Grind_mkGroundPattern___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("grind", 5, 5); +return x_1; +} } -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__1(lean_object* x_1) { +static lean_object* _init_l_Lean_Meta_Grind_mkGroundPattern___closed__2() { _start: { -lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; -x_2 = lean_array_get_size(x_1); -x_3 = lean_unsigned_to_nat(2u); -x_4 = lean_nat_mul(x_2, x_3); -lean_dec(x_2); -x_5 = lean_box(0); -x_6 = lean_mk_array(x_4, x_5); -x_7 = lean_unsigned_to_nat(0u); -x_8 = l_Std_DHashMap_Internal_Raw_u2080_expand_go___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__2(x_7, x_1, x_6); -return x_8; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("ground_pat", 10, 10); +return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +static lean_object* _init_l_Lean_Meta_Grind_mkGroundPattern___closed__3() { _start: { -lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_8 = lean_st_ref_take(x_2, x_7); -x_9 = lean_ctor_get(x_8, 0); -lean_inc(x_9); -x_10 = lean_ctor_get(x_9, 2); -lean_inc(x_10); -x_11 = lean_ctor_get(x_8, 1); -lean_inc(x_11); -lean_dec(x_8); -x_12 = !lean_is_exclusive(x_9); -if (x_12 == 0) +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_mkGroundPattern___closed__1; +x_2 = l_Lean_Meta_Grind_mkGroundPattern___closed__2; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkGroundPattern(lean_object* x_1) { +_start: { -lean_object* x_13; uint8_t x_14; -x_13 = lean_ctor_get(x_9, 2); -lean_dec(x_13); -x_14 = !lean_is_exclusive(x_10); -if (x_14 == 0) +lean_object* x_2; lean_object* x_3; +x_2 = l_Lean_Meta_Grind_mkGroundPattern___closed__3; +x_3 = l_Lean_mkAnnotation(x_2, x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_groundPattern_x3f(lean_object* x_1) { +_start: { -lean_object* x_15; lean_object* x_16; lean_object* x_17; uint64_t x_18; uint64_t x_19; uint64_t x_20; uint64_t x_21; uint64_t x_22; uint64_t x_23; uint64_t x_24; size_t x_25; size_t x_26; size_t x_27; size_t x_28; size_t x_29; lean_object* x_30; uint8_t x_31; -x_15 = lean_ctor_get(x_10, 0); -x_16 = lean_ctor_get(x_10, 1); -x_17 = lean_array_get_size(x_16); -x_18 = lean_uint64_of_nat(x_1); -x_19 = 32; -x_20 = lean_uint64_shift_right(x_18, x_19); -x_21 = lean_uint64_xor(x_18, x_20); -x_22 = 16; -x_23 = lean_uint64_shift_right(x_21, x_22); -x_24 = lean_uint64_xor(x_21, x_23); -x_25 = lean_uint64_to_usize(x_24); -x_26 = lean_usize_of_nat(x_17); -lean_dec(x_17); -x_27 = 1; -x_28 = lean_usize_sub(x_26, x_27); -x_29 = lean_usize_land(x_25, x_28); -x_30 = lean_array_uget(x_16, x_29); -x_31 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_foundBVar___spec__1(x_1, x_30); -if (x_31 == 0) +lean_object* x_2; lean_object* x_3; +x_2 = l_Lean_Meta_Grind_mkGroundPattern___closed__3; +x_3 = l_Lean_annotation_x3f(x_2, x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_groundPattern_x3f___boxed(lean_object* x_1) { +_start: { -lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; 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_32 = lean_unsigned_to_nat(1u); -x_33 = lean_nat_add(x_15, x_32); -lean_dec(x_15); -x_34 = lean_box(0); -x_35 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_35, 0, x_1); -lean_ctor_set(x_35, 1, x_34); -lean_ctor_set(x_35, 2, x_30); -x_36 = lean_array_uset(x_16, x_29, x_35); -x_37 = lean_unsigned_to_nat(4u); -x_38 = lean_nat_mul(x_33, x_37); -x_39 = lean_unsigned_to_nat(3u); -x_40 = lean_nat_div(x_38, x_39); -lean_dec(x_38); -x_41 = lean_array_get_size(x_36); -x_42 = lean_nat_dec_le(x_40, x_41); -lean_dec(x_41); -lean_dec(x_40); -if (x_42 == 0) +lean_object* x_2; +x_2 = l_Lean_Meta_Grind_groundPattern_x3f(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isGroundPattern(lean_object* x_1) { +_start: { -lean_object* x_43; lean_object* x_44; uint8_t x_45; -x_43 = l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__1(x_36); -lean_ctor_set(x_10, 1, x_43); -lean_ctor_set(x_10, 0, x_33); -x_44 = lean_st_ref_set(x_2, x_9, x_11); -x_45 = !lean_is_exclusive(x_44); -if (x_45 == 0) +lean_object* x_2; +x_2 = l_Lean_Meta_Grind_groundPattern_x3f(x_1); +if (lean_obj_tag(x_2) == 0) { -lean_object* x_46; -x_46 = lean_ctor_get(x_44, 0); -lean_dec(x_46); -lean_ctor_set(x_44, 0, x_34); -return x_44; +uint8_t x_3; +x_3 = 0; +return x_3; } else { -lean_object* x_47; lean_object* x_48; -x_47 = lean_ctor_get(x_44, 1); -lean_inc(x_47); -lean_dec(x_44); -x_48 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_48, 0, x_34); -lean_ctor_set(x_48, 1, x_47); -return x_48; +uint8_t x_4; +lean_dec(x_2); +x_4 = 1; +return x_4; } } -else +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isGroundPattern___boxed(lean_object* x_1) { +_start: { -lean_object* x_49; uint8_t x_50; -lean_ctor_set(x_10, 1, x_36); -lean_ctor_set(x_10, 0, x_33); -x_49 = lean_st_ref_set(x_2, x_9, x_11); -x_50 = !lean_is_exclusive(x_49); -if (x_50 == 0) +uint8_t x_2; lean_object* x_3; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isGroundPattern(x_1); +lean_dec(x_1); +x_3 = lean_box(x_2); +return x_3; +} +} +LEAN_EXPORT uint8_t l_Lean_Meta_Grind_isPatternDontCare(lean_object* x_1) { +_start: { -lean_object* x_51; -x_51 = lean_ctor_get(x_49, 0); -lean_dec(x_51); -lean_ctor_set(x_49, 0, x_34); -return x_49; +lean_object* x_2; uint8_t x_3; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare; +x_3 = lean_expr_eqv(x_1, x_2); +return x_3; } -else +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isPatternDontCare___boxed(lean_object* x_1) { +_start: { -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_34); -lean_ctor_set(x_53, 1, x_52); -return x_53; +uint8_t x_2; lean_object* x_3; +x_2 = l_Lean_Meta_Grind_isPatternDontCare(x_1); +lean_dec(x_1); +x_3 = lean_box(x_2); +return x_3; } } +LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isAtomicPattern(lean_object* x_1) { +_start: +{ +uint8_t x_2; +x_2 = l_Lean_Expr_isBVar(x_1); +if (x_2 == 0) +{ +lean_object* x_3; uint8_t x_4; +x_3 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare; +x_4 = lean_expr_eqv(x_1, x_3); +if (x_4 == 0) +{ +uint8_t x_5; +x_5 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isGroundPattern(x_1); +return x_5; } else { -lean_object* x_54; uint8_t x_55; -lean_dec(x_30); -lean_dec(x_1); -x_54 = lean_st_ref_set(x_2, x_9, x_11); -x_55 = !lean_is_exclusive(x_54); -if (x_55 == 0) -{ -lean_object* x_56; lean_object* x_57; -x_56 = lean_ctor_get(x_54, 0); -lean_dec(x_56); -x_57 = lean_box(0); -lean_ctor_set(x_54, 0, x_57); -return x_54; +uint8_t x_6; +x_6 = 1; +return x_6; +} } else { -lean_object* x_58; lean_object* x_59; lean_object* x_60; -x_58 = lean_ctor_get(x_54, 1); -lean_inc(x_58); -lean_dec(x_54); -x_59 = lean_box(0); -x_60 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_60, 0, x_59); -lean_ctor_set(x_60, 1, x_58); -return x_60; +uint8_t x_7; +x_7 = 1; +return x_7; } } } -else +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isAtomicPattern___boxed(lean_object* x_1) { +_start: { -lean_object* x_61; lean_object* x_62; lean_object* x_63; uint64_t x_64; uint64_t x_65; uint64_t x_66; uint64_t x_67; uint64_t x_68; uint64_t x_69; uint64_t x_70; size_t x_71; size_t x_72; size_t x_73; size_t x_74; size_t x_75; lean_object* x_76; uint8_t x_77; -x_61 = lean_ctor_get(x_10, 0); -x_62 = lean_ctor_get(x_10, 1); -lean_inc(x_62); -lean_inc(x_61); -lean_dec(x_10); -x_63 = lean_array_get_size(x_62); -x_64 = lean_uint64_of_nat(x_1); -x_65 = 32; -x_66 = lean_uint64_shift_right(x_64, x_65); -x_67 = lean_uint64_xor(x_64, x_66); -x_68 = 16; -x_69 = lean_uint64_shift_right(x_67, x_68); -x_70 = lean_uint64_xor(x_67, x_69); -x_71 = lean_uint64_to_usize(x_70); -x_72 = lean_usize_of_nat(x_63); -lean_dec(x_63); -x_73 = 1; -x_74 = lean_usize_sub(x_72, x_73); -x_75 = lean_usize_land(x_71, x_74); -x_76 = lean_array_uget(x_62, x_75); -x_77 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_foundBVar___spec__1(x_1, x_76); -if (x_77 == 0) +uint8_t x_2; lean_object* x_3; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isAtomicPattern(x_1); +lean_dec(x_1); +x_3 = lean_box(x_2); +return x_3; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___closed__1() { +_start: { -lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; uint8_t x_88; -x_78 = lean_unsigned_to_nat(1u); -x_79 = lean_nat_add(x_61, x_78); -lean_dec(x_61); -x_80 = lean_box(0); -x_81 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_81, 0, x_1); -lean_ctor_set(x_81, 1, x_80); -lean_ctor_set(x_81, 2, x_76); -x_82 = lean_array_uset(x_62, x_75, x_81); -x_83 = lean_unsigned_to_nat(4u); -x_84 = lean_nat_mul(x_79, x_83); -x_85 = lean_unsigned_to_nat(3u); -x_86 = lean_nat_div(x_84, x_85); -lean_dec(x_84); -x_87 = lean_array_get_size(x_82); -x_88 = lean_nat_dec_le(x_86, x_87); -lean_dec(x_87); -lean_dec(x_86); -if (x_88 == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked(" ", 1, 1); +return x_1; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___closed__2() { +_start: { -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 = l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__1(x_82); -x_90 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_90, 0, x_79); -lean_ctor_set(x_90, 1, x_89); -lean_ctor_set(x_9, 2, x_90); -x_91 = lean_st_ref_set(x_2, x_9, x_11); -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); +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___closed__1; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; } -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; +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___closed__2; +x_2 = l_Lean_MessageData_ofFormat(x_1); +return x_2; } -else +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; -x_95 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_95, 0, x_79); -lean_ctor_set(x_95, 1, x_82); -lean_ctor_set(x_9, 2, x_95); -x_96 = lean_st_ref_set(x_2, x_9, x_11); -x_97 = lean_ctor_get(x_96, 1); -lean_inc(x_97); -if (lean_is_exclusive(x_96)) { - lean_ctor_release(x_96, 0); - lean_ctor_release(x_96, 1); - x_98 = x_96; -} else { - lean_dec_ref(x_96); - x_98 = lean_box(0); +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_4 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___closed__3; +x_5 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_5, 0, x_2); +lean_ctor_set(x_5, 1, x_4); +x_6 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_6, 0, x_5); +lean_ctor_set(x_6, 1, x_1); +x_7 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_7, 0, x_6); +return x_7; } -if (lean_is_scalar(x_98)) { - x_99 = lean_alloc_ctor(0, 2, 0); -} else { - x_99 = x_98; } -lean_ctor_set(x_99, 0, x_80); -lean_ctor_set(x_99, 1, x_97); -return x_99; +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___boxed), 3, 0); +return x_1; } } -else +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2() { +_start: { -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_dec(x_76); -lean_dec(x_1); -x_100 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_100, 0, x_61); -lean_ctor_set(x_100, 1, x_62); -lean_ctor_set(x_9, 2, x_100); -x_101 = lean_st_ref_set(x_2, x_9, x_11); -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; -} else { - lean_dec_ref(x_101); - x_103 = lean_box(0); +lean_object* x_1; +x_1 = lean_mk_string_unchecked("(", 1, 1); +return x_1; } -x_104 = lean_box(0); -if (lean_is_scalar(x_103)) { - x_105 = lean_alloc_ctor(0, 2, 0); -} else { - x_105 = x_103; } -lean_ctor_set(x_105, 0, x_104); -lean_ctor_set(x_105, 1, x_102); -return x_105; +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(")", 1, 1); +return x_1; } } +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___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) { +_start: +{ +uint8_t x_7; +x_7 = lean_usize_dec_lt(x_5, x_4); +if (x_7 == 0) +{ +return x_6; } 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; uint64_t x_112; uint64_t x_113; uint64_t x_114; uint64_t x_115; uint64_t x_116; uint64_t x_117; uint64_t x_118; size_t x_119; size_t x_120; size_t x_121; size_t x_122; size_t x_123; lean_object* x_124; uint8_t x_125; -x_106 = lean_ctor_get(x_9, 0); -x_107 = lean_ctor_get(x_9, 1); -lean_inc(x_107); -lean_inc(x_106); -lean_dec(x_9); -x_108 = lean_ctor_get(x_10, 0); -lean_inc(x_108); -x_109 = lean_ctor_get(x_10, 1); -lean_inc(x_109); -if (lean_is_exclusive(x_10)) { - lean_ctor_release(x_10, 0); - lean_ctor_release(x_10, 1); - x_110 = x_10; -} else { - lean_dec_ref(x_10); - x_110 = lean_box(0); -} -x_111 = lean_array_get_size(x_109); -x_112 = lean_uint64_of_nat(x_1); -x_113 = 32; -x_114 = lean_uint64_shift_right(x_112, x_113); -x_115 = lean_uint64_xor(x_112, x_114); -x_116 = 16; -x_117 = lean_uint64_shift_right(x_115, x_116); -x_118 = lean_uint64_xor(x_115, x_117); -x_119 = lean_uint64_to_usize(x_118); -x_120 = lean_usize_of_nat(x_111); -lean_dec(x_111); -x_121 = 1; -x_122 = lean_usize_sub(x_120, x_121); -x_123 = lean_usize_land(x_119, x_122); -x_124 = lean_array_uget(x_109, x_123); -x_125 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_foundBVar___spec__1(x_1, x_124); -if (x_125 == 0) +lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_8 = lean_array_uget(x_3, x_5); +lean_inc(x_8); +x_9 = l_Lean_Meta_Grind_ppPattern(x_8); +x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1; +x_11 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isAtomicPattern(x_8); +lean_dec(x_8); +if (x_11 == 0) { -lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; uint8_t x_136; -x_126 = lean_unsigned_to_nat(1u); -x_127 = lean_nat_add(x_108, x_126); -lean_dec(x_108); -x_128 = lean_box(0); -x_129 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_129, 0, x_1); -lean_ctor_set(x_129, 1, x_128); -lean_ctor_set(x_129, 2, x_124); -x_130 = lean_array_uset(x_109, x_123, x_129); -x_131 = lean_unsigned_to_nat(4u); -x_132 = lean_nat_mul(x_127, x_131); -x_133 = lean_unsigned_to_nat(3u); -x_134 = lean_nat_div(x_132, x_133); -lean_dec(x_132); -x_135 = lean_array_get_size(x_130); -x_136 = lean_nat_dec_le(x_134, x_135); -lean_dec(x_135); -lean_dec(x_134); -if (x_136 == 0) +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_12 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2; +x_13 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3; +x_14 = l_Lean_MessageData_bracket(x_12, x_9, x_13); +x_15 = lean_box(0); +x_16 = lean_apply_3(x_10, x_14, x_6, x_15); +if (lean_obj_tag(x_16) == 0) { -lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; -x_137 = l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__1(x_130); -if (lean_is_scalar(x_110)) { - x_138 = lean_alloc_ctor(0, 2, 0); -} else { - x_138 = x_110; -} -lean_ctor_set(x_138, 0, x_127); -lean_ctor_set(x_138, 1, x_137); -x_139 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_139, 0, x_106); -lean_ctor_set(x_139, 1, x_107); -lean_ctor_set(x_139, 2, x_138); -x_140 = lean_st_ref_set(x_2, x_139, x_11); -x_141 = lean_ctor_get(x_140, 1); -lean_inc(x_141); -if (lean_is_exclusive(x_140)) { - lean_ctor_release(x_140, 0); - lean_ctor_release(x_140, 1); - x_142 = x_140; -} else { - lean_dec_ref(x_140); - x_142 = lean_box(0); +lean_object* x_17; +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +lean_dec(x_16); +return x_17; } -if (lean_is_scalar(x_142)) { - x_143 = lean_alloc_ctor(0, 2, 0); -} else { - x_143 = x_142; +else +{ +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_5, x_19); +x_5 = x_20; +x_6 = x_18; +goto _start; } -lean_ctor_set(x_143, 0, x_128); -lean_ctor_set(x_143, 1, x_141); -return x_143; } 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; -if (lean_is_scalar(x_110)) { - x_144 = lean_alloc_ctor(0, 2, 0); -} else { - x_144 = x_110; +lean_object* x_22; lean_object* x_23; +x_22 = lean_box(0); +x_23 = lean_apply_3(x_10, x_9, x_6, x_22); +if (lean_obj_tag(x_23) == 0) +{ +lean_object* x_24; +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +lean_dec(x_23); +return x_24; } -lean_ctor_set(x_144, 0, x_127); -lean_ctor_set(x_144, 1, x_130); -x_145 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_145, 0, x_106); -lean_ctor_set(x_145, 1, x_107); -lean_ctor_set(x_145, 2, x_144); -x_146 = lean_st_ref_set(x_2, x_145, x_11); -x_147 = lean_ctor_get(x_146, 1); -lean_inc(x_147); -if (lean_is_exclusive(x_146)) { - lean_ctor_release(x_146, 0); - lean_ctor_release(x_146, 1); - x_148 = x_146; -} else { - lean_dec_ref(x_146); - x_148 = lean_box(0); +else +{ +lean_object* x_25; size_t x_26; size_t x_27; +x_25 = lean_ctor_get(x_23, 0); +lean_inc(x_25); +lean_dec(x_23); +x_26 = 1; +x_27 = lean_usize_add(x_5, x_26); +x_5 = x_27; +x_6 = x_25; +goto _start; } -if (lean_is_scalar(x_148)) { - x_149 = lean_alloc_ctor(0, 2, 0); -} else { - x_149 = x_148; } -lean_ctor_set(x_149, 0, x_128); -lean_ctor_set(x_149, 1, x_147); -return x_149; } } +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { +_start: +{ +uint8_t x_7; +x_7 = lean_usize_dec_lt(x_5, x_4); +if (x_7 == 0) +{ +return x_6; +} else { -lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; -lean_dec(x_124); -lean_dec(x_1); -if (lean_is_scalar(x_110)) { - x_150 = lean_alloc_ctor(0, 2, 0); -} else { - x_150 = x_110; +lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_8 = lean_array_uget(x_3, x_5); +lean_inc(x_8); +x_9 = l_Lean_Meta_Grind_ppPattern(x_8); +x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1; +x_11 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isAtomicPattern(x_8); +lean_dec(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 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2; +x_13 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3; +x_14 = l_Lean_MessageData_bracket(x_12, x_9, x_13); +x_15 = lean_box(0); +x_16 = lean_apply_3(x_10, x_14, x_6, x_15); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_17; +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +lean_dec(x_16); +return x_17; } -lean_ctor_set(x_150, 0, x_108); -lean_ctor_set(x_150, 1, x_109); -x_151 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_151, 0, x_106); -lean_ctor_set(x_151, 1, x_107); -lean_ctor_set(x_151, 2, x_150); -x_152 = lean_st_ref_set(x_2, x_151, x_11); -x_153 = lean_ctor_get(x_152, 1); -lean_inc(x_153); -if (lean_is_exclusive(x_152)) { - lean_ctor_release(x_152, 0); - lean_ctor_release(x_152, 1); - x_154 = x_152; -} else { - lean_dec_ref(x_152); - x_154 = lean_box(0); +else +{ +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_5, x_19); +x_5 = x_20; +x_6 = x_18; +goto _start; } -x_155 = lean_box(0); -if (lean_is_scalar(x_154)) { - x_156 = lean_alloc_ctor(0, 2, 0); -} else { - x_156 = x_154; } -lean_ctor_set(x_156, 0, x_155); -lean_ctor_set(x_156, 1, x_153); -return x_156; +else +{ +lean_object* x_22; lean_object* x_23; +x_22 = lean_box(0); +x_23 = lean_apply_3(x_10, x_9, x_6, x_22); +if (lean_obj_tag(x_23) == 0) +{ +lean_object* x_24; +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +lean_dec(x_23); +return x_24; } +else +{ +lean_object* x_25; size_t x_26; size_t x_27; +x_25 = lean_ctor_get(x_23, 0); +lean_inc(x_25); +lean_dec(x_23); +x_26 = 1; +x_27 = lean_usize_add(x_5, x_26); +x_5 = x_27; +x_6 = x_25; +goto _start; } } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: -{ -lean_object* x_8; -x_8 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar(x_1, x_2, x_3, x_4, x_5, x_6, x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_8; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFn_x3f(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { _start: { -uint8_t x_2; -x_2 = l_Lean_Expr_isApp(x_1); -if (x_2 == 0) +uint8_t x_7; +x_7 = lean_usize_dec_lt(x_5, x_4); +if (x_7 == 0) { -lean_object* x_3; -x_3 = lean_box(0); -return x_3; +return x_6; } else { -lean_object* x_4; -x_4 = l_Lean_Expr_getAppFn(x_1); -switch (lean_obj_tag(x_4)) { -case 1: -{ -lean_object* x_5; -x_5 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_5, 0, x_4); -return x_5; -} -case 4: +lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_8 = lean_array_uget(x_3, x_5); +lean_inc(x_8); +x_9 = l_Lean_Meta_Grind_ppPattern(x_8); +x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1; +x_11 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isAtomicPattern(x_8); +lean_dec(x_8); +if (x_11 == 0) { -lean_object* x_6; lean_object* x_7; uint8_t x_8; -x_6 = lean_ctor_get(x_4, 0); -lean_inc(x_6); -x_7 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames; -x_8 = l_Array_contains___at_Lean_registerInternalExceptionId___spec__1(x_7, x_6); -lean_dec(x_6); -if (x_8 == 0) +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_12 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2; +x_13 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3; +x_14 = l_Lean_MessageData_bracket(x_12, x_9, x_13); +x_15 = lean_box(0); +x_16 = lean_apply_3(x_10, x_14, x_6, x_15); +if (lean_obj_tag(x_16) == 0) { -lean_object* x_9; -x_9 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_9, 0, x_4); -return x_9; +lean_object* x_17; +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +lean_dec(x_16); +return x_17; } else { -lean_object* x_10; -lean_dec(x_4); -x_10 = lean_box(0); -return x_10; +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_5, x_19); +x_5 = x_20; +x_6 = x_18; +goto _start; } } -default: +else { -lean_object* x_11; -lean_dec(x_4); -x_11 = lean_box(0); -return x_11; -} +lean_object* x_22; lean_object* x_23; +x_22 = lean_box(0); +x_23 = lean_apply_3(x_10, x_9, x_6, x_22); +if (lean_obj_tag(x_23) == 0) +{ +lean_object* x_24; +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +lean_dec(x_23); +return x_24; } +else +{ +lean_object* x_25; size_t x_26; size_t x_27; +x_25 = lean_ctor_get(x_23, 0); +lean_inc(x_25); +lean_dec(x_23); +x_26 = 1; +x_27 = lean_usize_add(x_5, x_26); +x_5 = x_27; +x_6 = x_25; +goto _start; } } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFn_x3f___boxed(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFn_x3f(x_1); -lean_dec(x_1); -return x_2; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFunMask___spec__1(size_t x_1, size_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { _start: { -uint8_t x_9; -x_9 = lean_usize_dec_lt(x_2, x_1); -if (x_9 == 0) +uint8_t x_7; +x_7 = lean_usize_dec_lt(x_5, x_4); +if (x_7 == 0) { -lean_object* x_10; -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -x_10 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_10, 0, x_3); -lean_ctor_set(x_10, 1, x_8); -return x_10; +return x_6; } else { -lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; lean_object* x_15; uint8_t x_22; lean_object* x_23; lean_object* x_38; -x_11 = lean_array_uget(x_3, x_2); -x_12 = lean_unsigned_to_nat(0u); -x_13 = lean_array_uset(x_3, x_2, x_12); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_11); -x_38 = l_Lean_Meta_isTypeFormer(x_11, x_4, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_38) == 0) +lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_8 = lean_array_uget(x_3, x_5); +lean_inc(x_8); +x_9 = l_Lean_Meta_Grind_ppPattern(x_8); +x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1; +x_11 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isAtomicPattern(x_8); +lean_dec(x_8); +if (x_11 == 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); -if (x_40 == 0) +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_12 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2; +x_13 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3; +x_14 = l_Lean_MessageData_bracket(x_12, x_9, x_13); +x_15 = lean_box(0); +x_16 = lean_apply_3(x_10, x_14, x_6, x_15); +if (lean_obj_tag(x_16) == 0) { -lean_object* x_41; lean_object* x_42; -lean_dec(x_39); -x_41 = lean_ctor_get(x_38, 1); -lean_inc(x_41); -lean_dec(x_38); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_11); -x_42 = l_Lean_Meta_isProof(x_11, x_4, x_5, x_6, x_7, x_41); -if (lean_obj_tag(x_42) == 0) +lean_object* x_17; +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +lean_dec(x_16); +return x_17; +} +else { -lean_object* x_43; lean_object* x_44; uint8_t x_45; -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_unbox(x_43); -lean_dec(x_43); -x_22 = x_45; -x_23 = x_44; -goto block_37; +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_5, x_19); +x_5 = x_20; +x_6 = x_18; +goto _start; +} } else { -uint8_t x_46; -lean_dec(x_13); -lean_dec(x_11); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -x_46 = !lean_is_exclusive(x_42); -if (x_46 == 0) +lean_object* x_22; lean_object* x_23; +x_22 = lean_box(0); +x_23 = lean_apply_3(x_10, x_9, x_6, x_22); +if (lean_obj_tag(x_23) == 0) { -return x_42; +lean_object* x_24; +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +lean_dec(x_23); +return x_24; } else { -lean_object* x_47; lean_object* x_48; lean_object* x_49; -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_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_25; size_t x_26; size_t x_27; +x_25 = lean_ctor_get(x_23, 0); +lean_inc(x_25); +lean_dec(x_23); +x_26 = 1; +x_27 = lean_usize_add(x_5, x_26); +x_5 = x_27; +x_6 = x_25; +goto _start; } } } -else -{ -lean_object* x_50; uint8_t x_51; -x_50 = lean_ctor_get(x_38, 1); -lean_inc(x_50); -lean_dec(x_38); -x_51 = lean_unbox(x_39); -lean_dec(x_39); -x_22 = x_51; -x_23 = x_50; -goto block_37; } } -else +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { +_start: { -uint8_t x_52; -lean_dec(x_13); -lean_dec(x_11); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -x_52 = !lean_is_exclusive(x_38); -if (x_52 == 0) +uint8_t x_7; +x_7 = lean_usize_dec_lt(x_5, x_4); +if (x_7 == 0) { -return x_38; +return x_6; } else { -lean_object* x_53; lean_object* x_54; lean_object* x_55; -x_53 = lean_ctor_get(x_38, 0); -x_54 = lean_ctor_get(x_38, 1); -lean_inc(x_54); -lean_inc(x_53); -lean_dec(x_38); -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; -} -} -block_21: -{ -size_t x_16; size_t x_17; lean_object* x_18; lean_object* x_19; -x_16 = 1; -x_17 = lean_usize_add(x_2, x_16); -x_18 = lean_box(x_14); -x_19 = lean_array_uset(x_13, x_2, x_18); -x_2 = x_17; -x_3 = x_19; -x_8 = x_15; -goto _start; -} -block_37: -{ -if (x_22 == 0) -{ -lean_object* x_24; lean_object* x_25; -x_24 = l_Lean_Expr_fvarId_x21(x_11); -lean_dec(x_11); -lean_inc(x_4); -x_25 = l_Lean_FVarId_getDecl(x_24, x_4, x_5, x_6, x_7, x_23); -if (lean_obj_tag(x_25) == 0) +lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_8 = lean_array_uget(x_3, x_5); +lean_inc(x_8); +x_9 = l_Lean_Meta_Grind_ppPattern(x_8); +x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1; +x_11 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isAtomicPattern(x_8); +lean_dec(x_8); +if (x_11 == 0) { -lean_object* x_26; lean_object* x_27; uint8_t 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_LocalDecl_binderInfo(x_26); -lean_dec(x_26); -x_29 = lean_box(x_28); -if (lean_obj_tag(x_29) == 3) +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_12 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2; +x_13 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3; +x_14 = l_Lean_MessageData_bracket(x_12, x_9, x_13); +x_15 = lean_box(0); +x_16 = lean_apply_3(x_10, x_14, x_6, x_15); +if (lean_obj_tag(x_16) == 0) { -uint8_t x_30; -x_30 = 1; -x_14 = x_30; -x_15 = x_27; -goto block_21; +lean_object* x_17; +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +lean_dec(x_16); +return x_17; } else { -uint8_t x_31; -lean_dec(x_29); -x_31 = 0; -x_14 = x_31; -x_15 = x_27; -goto block_21; +lean_object* x_18; size_t x_19; size_t x_20; +x_18 = lean_ctor_get(x_16, 0); +lean_inc(x_18); +lean_dec(x_16); +x_19 = 1; +x_20 = lean_usize_add(x_5, x_19); +x_5 = x_20; +x_6 = x_18; +goto _start; } } else { -uint8_t x_32; -lean_dec(x_13); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -x_32 = !lean_is_exclusive(x_25); -if (x_32 == 0) -{ -return x_25; -} -else +lean_object* x_22; lean_object* x_23; +x_22 = lean_box(0); +x_23 = lean_apply_3(x_10, x_9, x_6, x_22); +if (lean_obj_tag(x_23) == 0) { -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_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_object* x_24; +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +lean_dec(x_23); +return x_24; } else { -uint8_t x_36; -lean_dec(x_11); -x_36 = 1; -x_14 = x_36; -x_15 = x_23; -goto block_21; +lean_object* x_25; size_t x_26; size_t x_27; +x_25 = lean_ctor_get(x_23, 0); +lean_inc(x_25); +lean_dec(x_23); +x_26 = 1; +x_27 = lean_usize_add(x_5, x_26); +x_5 = x_27; +x_6 = x_25; +goto _start; } } } } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFunMask___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { _start: { -size_t x_8; size_t x_9; lean_object* x_10; -x_8 = lean_array_size(x_1); -x_9 = 0; -x_10 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFunMask___spec__1(x_8, x_9, x_1, x_3, x_4, x_5, x_6, x_7); -return x_10; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFunMask___closed__1() { -_start: +uint8_t x_7; +x_7 = lean_usize_dec_lt(x_5, x_4); +if (x_7 == 0) { -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFunMask___lambda__1___boxed), 7, 0); -return x_1; -} +return x_6; } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFunMask(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: -{ -lean_object* x_8; -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -x_8 = lean_infer_type(x_1, x_3, x_4, x_5, x_6, x_7); -if (lean_obj_tag(x_8) == 0) +else { -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; -x_9 = lean_ctor_get(x_8, 0); -lean_inc(x_9); -x_10 = lean_ctor_get(x_8, 1); -lean_inc(x_10); +lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_8 = lean_array_uget(x_3, x_5); +lean_inc(x_8); +x_9 = l_Lean_Meta_Grind_ppPattern(x_8); +x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1; +x_11 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isAtomicPattern(x_8); lean_dec(x_8); -x_11 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_11, 0, x_2); -x_12 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFunMask___closed__1; -x_13 = 0; -x_14 = l_Lean_Meta_forallBoundedTelescope___at_Lean_Meta_arrowDomainsN___spec__6___rarg(x_9, x_11, x_12, x_13, x_3, x_4, x_5, x_6, x_10); -return x_14; +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 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2; +x_13 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3; +x_14 = l_Lean_MessageData_bracket(x_12, x_9, x_13); +x_15 = lean_box(0); +x_16 = lean_apply_3(x_10, x_14, x_6, x_15); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_17; +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +lean_dec(x_16); +return x_17; } else { -uint8_t x_15; -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_8); -if (x_15 == 0) +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_5, x_19); +x_5 = x_20; +x_6 = x_18; +goto _start; +} +} +else { -return x_8; +lean_object* x_22; lean_object* x_23; +x_22 = lean_box(0); +x_23 = lean_apply_3(x_10, x_9, x_6, x_22); +if (lean_obj_tag(x_23) == 0) +{ +lean_object* x_24; +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +lean_dec(x_23); +return x_24; } else { -lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_16 = lean_ctor_get(x_8, 0); -x_17 = lean_ctor_get(x_8, 1); -lean_inc(x_17); -lean_inc(x_16); -lean_dec(x_8); -x_18 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_18, 0, x_16); -lean_ctor_set(x_18, 1, x_17); -return x_18; +lean_object* x_25; size_t x_26; size_t x_27; +x_25 = lean_ctor_get(x_23, 0); +lean_inc(x_25); +lean_dec(x_23); +x_26 = 1; +x_27 = lean_usize_add(x_5, x_26); +x_5 = x_27; +x_6 = x_25; +goto _start; } } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFunMask___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { _start: { -size_t x_9; size_t x_10; lean_object* x_11; -x_9 = lean_unbox_usize(x_1); -lean_dec(x_1); -x_10 = lean_unbox_usize(x_2); -lean_dec(x_2); -x_11 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFunMask___spec__1(x_9, x_10, x_3, x_4, x_5, x_6, x_7, x_8); -return x_11; +uint8_t x_7; +x_7 = lean_usize_dec_lt(x_5, x_4); +if (x_7 == 0) +{ +return x_6; } +else +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_8 = lean_array_uget(x_3, x_5); +lean_inc(x_8); +x_9 = l_Lean_Meta_Grind_ppPattern(x_8); +x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1; +x_11 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isAtomicPattern(x_8); +lean_dec(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 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2; +x_13 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3; +x_14 = l_Lean_MessageData_bracket(x_12, x_9, x_13); +x_15 = lean_box(0); +x_16 = lean_apply_3(x_10, x_14, x_6, x_15); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_17; +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +lean_dec(x_16); +return x_17; } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFunMask___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: +else { -lean_object* x_8; -x_8 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFunMask___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7); -lean_dec(x_2); -return x_8; +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_5, x_19); +x_5 = x_20; +x_6 = x_18; +goto _start; } } -LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: +else { -lean_object* x_8; lean_object* x_9; uint8_t x_10; -x_8 = lean_ctor_get(x_5, 5); -x_9 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_1, x_3, x_4, x_5, x_6, x_7); -x_10 = !lean_is_exclusive(x_9); -if (x_10 == 0) +lean_object* x_22; lean_object* x_23; +x_22 = lean_box(0); +x_23 = lean_apply_3(x_10, x_9, x_6, x_22); +if (lean_obj_tag(x_23) == 0) { -lean_object* x_11; lean_object* x_12; -x_11 = lean_ctor_get(x_9, 0); -lean_inc(x_8); -x_12 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_12, 0, x_8); -lean_ctor_set(x_12, 1, x_11); -lean_ctor_set_tag(x_9, 1); -lean_ctor_set(x_9, 0, x_12); -return x_9; +lean_object* x_24; +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +lean_dec(x_23); +return x_24; } 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); -lean_inc(x_8); -x_15 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_15, 0, x_8); -lean_ctor_set(x_15, 1, x_13); -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_25; size_t x_26; size_t x_27; +x_25 = lean_ctor_get(x_23, 0); +lean_inc(x_25); +lean_dec(x_23); +x_26 = 1; +x_27 = lean_usize_add(x_5, x_26); +x_5 = x_27; +x_6 = x_25; +goto _start; } } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__2___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_10 = lean_array_set(x_2, x_1, x_3); -x_11 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_11, 0, 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_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__8(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { _start: { -lean_object* x_14; uint8_t x_15; -x_14 = lean_ctor_get(x_3, 1); -x_15 = lean_nat_dec_lt(x_5, x_14); -if (x_15 == 0) +uint8_t x_7; +x_7 = lean_usize_dec_lt(x_5, x_4); +if (x_7 == 0) { -lean_object* 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_5); -x_16 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_16, 0, x_4); -lean_ctor_set(x_16, 1, x_13); -return x_16; +return x_6; } else { -lean_object* x_17; lean_object* x_35; uint8_t x_36; lean_object* x_37; uint8_t x_38; lean_object* x_39; -x_35 = lean_array_get_size(x_4); -x_36 = lean_nat_dec_lt(x_5, x_35); -lean_dec(x_35); -x_37 = lean_array_get_size(x_1); -x_38 = lean_nat_dec_lt(x_5, x_37); -lean_dec(x_37); -if (x_36 == 0) +lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_8 = lean_array_uget(x_3, x_5); +lean_inc(x_8); +x_9 = l_Lean_Meta_Grind_ppPattern(x_8); +x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1; +x_11 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isAtomicPattern(x_8); +lean_dec(x_8); +if (x_11 == 0) { -lean_object* x_93; lean_object* x_94; -x_93 = l_Lean_instInhabitedExpr; -x_94 = l_outOfBounds___rarg(x_93); -x_39 = x_94; -goto block_92; -} -else +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_12 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2; +x_13 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3; +x_14 = l_Lean_MessageData_bracket(x_12, x_9, x_13); +x_15 = lean_box(0); +x_16 = lean_apply_3(x_10, x_14, x_6, x_15); +if (lean_obj_tag(x_16) == 0) { -lean_object* x_95; -x_95 = lean_array_fget(x_4, x_5); -x_39 = x_95; -goto block_92; +lean_object* x_17; +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +lean_dec(x_16); +return x_17; } -block_34: -{ -if (lean_obj_tag(x_17) == 0) +else { -lean_object* x_18; -x_18 = lean_ctor_get(x_17, 0); +lean_object* x_18; size_t x_19; size_t x_20; +x_18 = lean_ctor_get(x_16, 0); lean_inc(x_18); -if (lean_obj_tag(x_18) == 0) -{ -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_5); -x_19 = !lean_is_exclusive(x_17); -if (x_19 == 0) -{ -lean_object* x_20; lean_object* x_21; -x_20 = lean_ctor_get(x_17, 0); -lean_dec(x_20); -x_21 = lean_ctor_get(x_18, 0); -lean_inc(x_21); -lean_dec(x_18); -lean_ctor_set(x_17, 0, x_21); -return x_17; +lean_dec(x_16); +x_19 = 1; +x_20 = lean_usize_add(x_5, x_19); +x_5 = x_20; +x_6 = x_18; +goto _start; +} } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_17, 1); -lean_inc(x_22); -lean_dec(x_17); -x_23 = lean_ctor_get(x_18, 0); -lean_inc(x_23); -lean_dec(x_18); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_23); -lean_ctor_set(x_24, 1, x_22); +lean_object* x_22; lean_object* x_23; +x_22 = lean_box(0); +x_23 = lean_apply_3(x_10, x_9, x_6, x_22); +if (lean_obj_tag(x_23) == 0) +{ +lean_object* x_24; +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +lean_dec(x_23); return x_24; } -} else { -lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_25 = lean_ctor_get(x_17, 1); +lean_object* x_25; size_t x_26; size_t x_27; +x_25 = lean_ctor_get(x_23, 0); 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_3, 2); -x_28 = lean_nat_add(x_5, x_27); -lean_dec(x_5); -x_4 = x_26; -x_5 = x_28; -x_6 = lean_box(0); -x_7 = lean_box(0); -x_13 = x_25; +lean_dec(x_23); +x_26 = 1; +x_27 = lean_usize_add(x_5, x_26); +x_5 = x_27; +x_6 = x_25; goto _start; } } -else -{ -uint8_t x_30; -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_5); -x_30 = !lean_is_exclusive(x_17); -if (x_30 == 0) -{ -return x_17; -} -else -{ -lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_31 = lean_ctor_get(x_17, 0); -x_32 = lean_ctor_get(x_17, 1); -lean_inc(x_32); -lean_inc(x_31); -lean_dec(x_17); -x_33 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_33, 0, x_31); -lean_ctor_set(x_33, 1, x_32); -return x_33; } } } -block_92: +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__9(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { +_start: { -uint8_t x_40; uint8_t x_41; -x_40 = l_Lean_Expr_hasLooseBVars(x_39); -if (x_38 == 0) +uint8_t x_7; +x_7 = lean_usize_dec_lt(x_5, x_4); +if (x_7 == 0) { -uint8_t x_89; -x_89 = 0; -x_41 = x_89; -goto block_88; +return x_6; } else { -lean_object* x_90; uint8_t x_91; -x_90 = lean_array_fget(x_1, x_5); -x_91 = lean_unbox(x_90); -lean_dec(x_90); -x_41 = x_91; -goto block_88; -} -block_88: -{ -if (x_40 == 0) +lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_8 = lean_array_uget(x_3, x_5); +lean_inc(x_8); +x_9 = l_Lean_Meta_Grind_ppPattern(x_8); +x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1; +x_11 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isAtomicPattern(x_8); +lean_dec(x_8); +if (x_11 == 0) { -uint8_t x_42; -x_42 = l_Lean_Expr_hasMVar(x_39); -if (x_42 == 0) +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_12 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2; +x_13 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3; +x_14 = l_Lean_MessageData_bracket(x_12, x_9, x_13); +x_15 = lean_box(0); +x_16 = lean_apply_3(x_10, x_14, x_6, x_15); +if (lean_obj_tag(x_16) == 0) { -lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; -x_43 = l_Lean_Meta_Grind_mkGroundPattern(x_39); -x_44 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__2___lambda__1(x_5, x_4, x_43, x_8, x_9, x_10, x_11, x_12, x_13); -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 = lean_ctor_get(x_3, 2); -x_49 = lean_nat_add(x_5, x_48); -lean_dec(x_5); -x_4 = x_47; -x_5 = x_49; -x_6 = lean_box(0); -x_7 = lean_box(0); -x_13 = x_46; -goto _start; +lean_object* x_17; +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +lean_dec(x_16); +return x_17; } 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_dec(x_39); -x_51 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare; -x_52 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__2___lambda__1(x_5, x_4, x_51, x_8, x_9, x_10, x_11, x_12, x_13); -x_53 = lean_ctor_get(x_52, 0); -lean_inc(x_53); -x_54 = lean_ctor_get(x_52, 1); -lean_inc(x_54); -lean_dec(x_52); -x_55 = lean_ctor_get(x_53, 0); -lean_inc(x_55); -lean_dec(x_53); -x_56 = lean_ctor_get(x_3, 2); -x_57 = lean_nat_add(x_5, x_56); -lean_dec(x_5); -x_4 = x_55; -x_5 = x_57; -x_6 = lean_box(0); -x_7 = lean_box(0); -x_13 = x_54; +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_5, x_19); +x_5 = x_20; +x_6 = x_18; goto _start; } } else { -if (lean_obj_tag(x_39) == 0) -{ -lean_object* x_59; lean_object* x_60; -x_59 = lean_ctor_get(x_39, 0); -lean_inc(x_59); -x_60 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_foundBVar(x_59, x_8, x_9, x_10, x_11, x_12, x_13); -if (x_41 == 0) +lean_object* x_22; lean_object* x_23; +x_22 = lean_box(0); +x_23 = lean_apply_3(x_10, x_9, x_6, x_22); +if (lean_obj_tag(x_23) == 0) { -lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; -x_61 = lean_ctor_get(x_60, 1); -lean_inc(x_61); -lean_dec(x_60); -x_62 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar(x_59, x_8, x_9, x_10, x_11, x_12, x_61); -x_63 = lean_ctor_get(x_62, 1); -lean_inc(x_63); -lean_dec(x_62); -x_64 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__2___lambda__1(x_5, x_4, x_39, x_8, x_9, x_10, x_11, x_12, x_63); -x_17 = x_64; -goto block_34; +lean_object* x_24; +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +lean_dec(x_23); +return x_24; } else { -lean_object* x_65; uint8_t x_66; -x_65 = lean_ctor_get(x_60, 0); -lean_inc(x_65); -x_66 = lean_unbox(x_65); -lean_dec(x_65); -if (x_66 == 0) -{ -lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; -x_67 = lean_ctor_get(x_60, 1); -lean_inc(x_67); -lean_dec(x_60); -x_68 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar(x_59, x_8, x_9, x_10, x_11, x_12, x_67); -x_69 = lean_ctor_get(x_68, 1); -lean_inc(x_69); -lean_dec(x_68); -x_70 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__2___lambda__1(x_5, x_4, x_39, x_8, x_9, x_10, x_11, x_12, x_69); -x_17 = x_70; -goto block_34; +lean_object* x_25; size_t x_26; size_t x_27; +x_25 = lean_ctor_get(x_23, 0); +lean_inc(x_25); +lean_dec(x_23); +x_26 = 1; +x_27 = lean_usize_add(x_5, x_26); +x_5 = x_27; +x_6 = x_25; +goto _start; } -else -{ -lean_object* x_71; lean_object* x_72; lean_object* x_73; -lean_dec(x_59); -lean_dec(x_39); -x_71 = lean_ctor_get(x_60, 1); -lean_inc(x_71); -lean_dec(x_60); -x_72 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare; -x_73 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__2___lambda__1(x_5, x_4, x_72, x_8, x_9, x_10, x_11, x_12, x_71); -x_17 = x_73; -goto block_34; } } } -else -{ -if (x_41 == 0) +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__10(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { +_start: { -lean_object* x_74; -x_74 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFn_x3f(x_39); -if (lean_obj_tag(x_74) == 0) +uint8_t x_7; +x_7 = lean_usize_dec_lt(x_5, x_4); +if (x_7 == 0) { -lean_object* x_75; lean_object* x_76; -lean_dec(x_39); -x_75 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare; -x_76 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__2___lambda__1(x_5, x_4, x_75, x_8, x_9, x_10, x_11, x_12, x_13); -x_17 = x_76; -goto block_34; +return x_6; } else { -uint8_t x_77; lean_object* x_78; -lean_dec(x_74); -x_77 = 0; -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); +lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_8 = lean_array_uget(x_3, x_5); lean_inc(x_8); -x_78 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go(x_39, x_77, x_8, x_9, x_10, x_11, x_12, x_13); -if (lean_obj_tag(x_78) == 0) -{ -lean_object* x_79; lean_object* x_80; lean_object* x_81; -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 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__2___lambda__1(x_5, x_4, x_79, x_8, x_9, x_10, x_11, x_12, x_80); -x_17 = x_81; -goto block_34; -} -else +x_9 = l_Lean_Meta_Grind_ppPattern(x_8); +x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1; +x_11 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isAtomicPattern(x_8); +lean_dec(x_8); +if (x_11 == 0) { -uint8_t x_82; -lean_dec(x_4); -x_82 = !lean_is_exclusive(x_78); -if (x_82 == 0) +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_12 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2; +x_13 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3; +x_14 = l_Lean_MessageData_bracket(x_12, x_9, x_13); +x_15 = lean_box(0); +x_16 = lean_apply_3(x_10, x_14, x_6, x_15); +if (lean_obj_tag(x_16) == 0) { -x_17 = x_78; -goto block_34; +lean_object* x_17; +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +lean_dec(x_16); +return x_17; } else { -lean_object* x_83; lean_object* x_84; lean_object* x_85; -x_83 = lean_ctor_get(x_78, 0); -x_84 = lean_ctor_get(x_78, 1); -lean_inc(x_84); -lean_inc(x_83); -lean_dec(x_78); -x_85 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_85, 0, x_83); -lean_ctor_set(x_85, 1, x_84); -x_17 = x_85; -goto block_34; -} -} +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_5, x_19); +x_5 = x_20; +x_6 = x_18; +goto _start; } } else { -lean_object* x_86; lean_object* x_87; -lean_dec(x_39); -x_86 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare; -x_87 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__2___lambda__1(x_5, x_4, x_86, x_8, x_9, x_10, x_11, x_12, x_13); -x_17 = x_87; -goto block_34; -} -} +lean_object* x_22; lean_object* x_23; +x_22 = lean_box(0); +x_23 = lean_apply_3(x_10, x_9, x_6, x_22); +if (lean_obj_tag(x_23) == 0) +{ +lean_object* x_24; +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +lean_dec(x_23); +return x_24; } +else +{ +lean_object* x_25; size_t x_26; size_t x_27; +x_25 = lean_ctor_get(x_23, 0); +lean_inc(x_25); +lean_dec(x_23); +x_26 = 1; +x_27 = lean_usize_add(x_5, x_26); +x_5 = x_27; +x_6 = x_25; +goto _start; } } } } } -static lean_object* _init_l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__3___closed__1() { +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__11(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_instMonadMetaM; -x_2 = l_ReaderT_instMonad___rarg(x_1); -return x_2; -} -} -static lean_object* _init_l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__3___closed__2() { -_start: +uint8_t x_7; +x_7 = lean_usize_dec_lt(x_5, x_4); +if (x_7 == 0) { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__3___closed__1; -x_2 = l_Lean_instInhabitedExpr; -x_3 = l_instInhabitedOfMonad___rarg(x_1, x_2); -return x_3; +return x_6; } +else +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_8 = lean_array_uget(x_3, x_5); +lean_inc(x_8); +x_9 = l_Lean_Meta_Grind_ppPattern(x_8); +x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1; +x_11 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isAtomicPattern(x_8); +lean_dec(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 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2; +x_13 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3; +x_14 = l_Lean_MessageData_bracket(x_12, x_9, x_13); +x_15 = lean_box(0); +x_16 = lean_apply_3(x_10, x_14, x_6, x_15); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_17; +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +lean_dec(x_16); +return x_17; } -LEAN_EXPORT lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_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) { -_start: +else { -lean_object* x_8; lean_object* x_9; lean_object* x_10; -x_8 = l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__3___closed__2; -x_9 = lean_panic_fn(x_8, x_1); -x_10 = lean_apply_6(x_9, x_2, x_3, x_4, x_5, x_6, x_7); -return x_10; +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_5, x_19); +x_5 = x_20; +x_6 = x_18; +goto _start; } } -LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_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) { -_start: +else { -lean_object* x_8; lean_object* x_9; uint8_t x_10; -x_8 = lean_ctor_get(x_5, 5); -x_9 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_1, x_3, x_4, x_5, x_6, x_7); -x_10 = !lean_is_exclusive(x_9); -if (x_10 == 0) +lean_object* x_22; lean_object* x_23; +x_22 = lean_box(0); +x_23 = lean_apply_3(x_10, x_9, x_6, x_22); +if (lean_obj_tag(x_23) == 0) { -lean_object* x_11; lean_object* x_12; -x_11 = lean_ctor_get(x_9, 0); -lean_inc(x_8); -x_12 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_12, 0, x_8); -lean_ctor_set(x_12, 1, x_11); -lean_ctor_set_tag(x_9, 1); -lean_ctor_set(x_9, 0, x_12); -return x_9; +lean_object* x_24; +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +lean_dec(x_23); +return x_24; } 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); -lean_inc(x_8); -x_15 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_15, 0, x_8); -lean_ctor_set(x_15, 1, x_13); -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_25; size_t x_26; size_t x_27; +x_25 = lean_ctor_get(x_23, 0); +lean_inc(x_25); +lean_dec(x_23); +x_26 = 1; +x_27 = lean_usize_add(x_5, x_26); +x_5 = x_27; +x_6 = x_25; +goto _start; } } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__1() { +} +} +static lean_object* _init_l_Lean_Meta_Grind_ppPattern___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("invalid pattern, (non-forbidden) application expected", 53, 53); +x_1 = lean_mk_string_unchecked("#", 1, 1); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__2() { +static lean_object* _init_l_Lean_Meta_Grind_ppPattern___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__1; +x_1 = l_Lean_Meta_Grind_ppPattern___closed__1; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__3() { +static lean_object* _init_l_Lean_Meta_Grind_ppPattern___closed__3() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("assertion violation: ", 21, 21); +x_1 = lean_mk_string_unchecked("", 0, 0); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__4() { +static lean_object* _init_l_Lean_Meta_Grind_ppPattern___closed__4() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("f.isConst || f.isFVar\n ", 24, 24); -return x_1; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_ppPattern___closed__3; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__5() { +static lean_object* _init_l_Lean_Meta_Grind_ppPattern___closed__5() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__3; -x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__4; -x_3 = lean_string_append(x_1, x_2); -return x_3; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_levelZero; +x_2 = l_Lean_Expr_sort___override(x_1); +return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__6() { +static lean_object* _init_l_Lean_Meta_Grind_ppPattern___closed__6() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("_private.Lean.Meta.Tactic.Grind.EMatchTheorem.0.Lean.Meta.Grind.NormalizePattern.go", 83, 83); +x_1 = lean_mk_string_unchecked("\?", 1, 1); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__7() { +static lean_object* _init_l_Lean_Meta_Grind_ppPattern___closed__7() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_1 = l_Lean_Meta_Grind_EMatchTheorems_insert___closed__1; -x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__6; -x_3 = lean_unsigned_to_nat(177u); -x_4 = lean_unsigned_to_nat(2u); -x_5 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__5; -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; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_ppPattern___closed__6; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_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) { +static lean_object* _init_l_Lean_Meta_Grind_ppPattern___closed__8() { _start: { -lean_object* x_9; -x_9 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFn_x3f(x_1); -if (lean_obj_tag(x_9) == 0) -{ -lean_object* x_10; lean_object* x_11; -lean_dec(x_1); -x_10 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__2; -x_11 = l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__1(x_10, 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_11; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("`[", 2, 2); +return x_1; } -else -{ -lean_object* x_12; lean_object* x_13; uint8_t x_46; -x_12 = lean_ctor_get(x_9, 0); -lean_inc(x_12); -lean_dec(x_9); -x_46 = l_Lean_Expr_isConst(x_12); -if (x_46 == 0) -{ -uint8_t x_47; -x_47 = l_Lean_Expr_isFVar(x_12); -if (x_47 == 0) +} +static lean_object* _init_l_Lean_Meta_Grind_ppPattern___closed__9() { +_start: { -lean_object* x_48; lean_object* x_49; -lean_dec(x_12); -lean_dec(x_1); -x_48 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__7; -x_49 = l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__3(x_48, x_3, x_4, x_5, x_6, x_7, x_8); -return x_49; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_ppPattern___closed__8; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } -else +} +static lean_object* _init_l_Lean_Meta_Grind_ppPattern___closed__10() { +_start: { -lean_object* x_50; -x_50 = lean_box(0); -x_13 = x_50; -goto block_45; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("]", 1, 1); +return x_1; } } -else +static lean_object* _init_l_Lean_Meta_Grind_ppPattern___closed__11() { +_start: { -lean_object* x_51; -x_51 = lean_box(0); -x_13 = x_51; -goto block_45; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_ppPattern___closed__10; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } -block_45: +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ppPattern(lean_object* x_1) { +_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_dec(x_13); -lean_inc(x_12); -x_14 = l_Lean_Expr_toHeadIndex(x_12); -x_15 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol(x_14, x_3, x_4, x_5, x_6, x_7, x_8); -x_16 = lean_ctor_get(x_15, 1); -lean_inc(x_16); +lean_object* x_2; +x_2 = l_Lean_Meta_Grind_groundPattern_x3f(x_1); +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_3; uint8_t x_4; +x_3 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare; +x_4 = lean_expr_eqv(x_1, x_3); +if (x_4 == 0) +{ +switch (lean_obj_tag(x_1)) { +case 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; +x_5 = lean_ctor_get(x_1, 0); +lean_inc(x_5); +lean_dec(x_1); +x_6 = l___private_Init_Data_Repr_0__Nat_reprFast(x_5); +x_7 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_7, 0, x_6); +x_8 = l_Lean_MessageData_ofFormat(x_7); +x_9 = l_Lean_Meta_Grind_ppPattern___closed__2; +x_10 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_10, 0, x_9); +lean_ctor_set(x_10, 1, x_8); +x_11 = l_Lean_Meta_Grind_ppPattern___closed__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; +} +case 1: +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; size_t x_26; size_t x_27; lean_object* x_28; +x_13 = l_Lean_Expr_getAppFn(x_1); +x_14 = l_Lean_MessageData_ofExpr(x_13); +x_15 = l_Lean_Meta_Grind_ppPattern___closed__4; +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 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_17, 0, x_16); +lean_ctor_set(x_17, 1, x_15); +x_18 = lean_unsigned_to_nat(0u); +x_19 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_18); +x_20 = l_Lean_Meta_Grind_ppPattern___closed__5; +lean_inc(x_19); +x_21 = lean_mk_array(x_19, x_20); +x_22 = lean_unsigned_to_nat(1u); +x_23 = lean_nat_sub(x_19, x_22); +lean_dec(x_19); +x_24 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_21, x_23); +x_25 = lean_box(0); +x_26 = lean_array_size(x_24); +x_27 = 0; +x_28 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1(x_24, x_25, x_24, x_26, x_27, x_17); +lean_dec(x_24); +return x_28; +} +case 2: +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; 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_29 = l_Lean_Expr_getAppFn(x_1); +x_30 = l_Lean_MessageData_ofExpr(x_29); +x_31 = l_Lean_Meta_Grind_ppPattern___closed__4; +x_32 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_30); +x_33 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_31); +x_34 = lean_unsigned_to_nat(0u); +x_35 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_34); +x_36 = l_Lean_Meta_Grind_ppPattern___closed__5; +lean_inc(x_35); +x_37 = lean_mk_array(x_35, x_36); +x_38 = lean_unsigned_to_nat(1u); +x_39 = lean_nat_sub(x_35, x_38); +lean_dec(x_35); +x_40 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_37, x_39); +x_41 = lean_box(0); +x_42 = lean_array_size(x_40); +x_43 = 0; +x_44 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__2(x_40, x_41, x_40, x_42, x_43, x_33); +lean_dec(x_40); +return x_44; +} +case 3: +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; size_t x_58; size_t x_59; lean_object* x_60; +x_45 = l_Lean_Expr_getAppFn(x_1); +x_46 = l_Lean_MessageData_ofExpr(x_45); +x_47 = l_Lean_Meta_Grind_ppPattern___closed__4; +x_48 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_48, 0, x_47); +lean_ctor_set(x_48, 1, x_46); +x_49 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_49, 0, x_48); +lean_ctor_set(x_49, 1, x_47); +x_50 = lean_unsigned_to_nat(0u); +x_51 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_50); +x_52 = l_Lean_Meta_Grind_ppPattern___closed__5; +lean_inc(x_51); +x_53 = lean_mk_array(x_51, x_52); +x_54 = lean_unsigned_to_nat(1u); +x_55 = lean_nat_sub(x_51, x_54); +lean_dec(x_51); +x_56 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_53, x_55); +x_57 = lean_box(0); +x_58 = lean_array_size(x_56); +x_59 = 0; +x_60 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__3(x_56, x_57, x_56, x_58, x_59, x_49); +lean_dec(x_56); +return x_60; +} +case 4: +{ +lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; size_t x_74; size_t x_75; lean_object* x_76; +x_61 = l_Lean_Expr_getAppFn(x_1); +x_62 = l_Lean_MessageData_ofExpr(x_61); +x_63 = l_Lean_Meta_Grind_ppPattern___closed__4; +x_64 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_64, 0, x_63); +lean_ctor_set(x_64, 1, x_62); +x_65 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_65, 0, x_64); +lean_ctor_set(x_65, 1, x_63); +x_66 = lean_unsigned_to_nat(0u); +x_67 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_66); +x_68 = l_Lean_Meta_Grind_ppPattern___closed__5; +lean_inc(x_67); +x_69 = lean_mk_array(x_67, x_68); +x_70 = lean_unsigned_to_nat(1u); +x_71 = lean_nat_sub(x_67, x_70); +lean_dec(x_67); +x_72 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_69, x_71); +x_73 = lean_box(0); +x_74 = lean_array_size(x_72); +x_75 = 0; +x_76 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__4(x_72, x_73, x_72, x_74, x_75, x_65); +lean_dec(x_72); +return x_76; +} +case 5: +{ +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; size_t x_90; size_t x_91; lean_object* x_92; +x_77 = l_Lean_Expr_getAppFn(x_1); +x_78 = l_Lean_MessageData_ofExpr(x_77); +x_79 = l_Lean_Meta_Grind_ppPattern___closed__4; +x_80 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_80, 0, x_79); +lean_ctor_set(x_80, 1, x_78); +x_81 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_81, 0, x_80); +lean_ctor_set(x_81, 1, x_79); +x_82 = lean_unsigned_to_nat(0u); +x_83 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_82); +x_84 = l_Lean_Meta_Grind_ppPattern___closed__5; +lean_inc(x_83); +x_85 = lean_mk_array(x_83, x_84); +x_86 = lean_unsigned_to_nat(1u); +x_87 = lean_nat_sub(x_83, x_86); +lean_dec(x_83); +x_88 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_85, x_87); +x_89 = lean_box(0); +x_90 = lean_array_size(x_88); +x_91 = 0; +x_92 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__5(x_88, x_89, x_88, x_90, x_91, x_81); +lean_dec(x_88); +return x_92; +} +case 6: +{ +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; size_t x_106; size_t x_107; lean_object* x_108; +x_93 = l_Lean_Expr_getAppFn(x_1); +x_94 = l_Lean_MessageData_ofExpr(x_93); +x_95 = l_Lean_Meta_Grind_ppPattern___closed__4; +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_alloc_ctor(7, 2, 0); +lean_ctor_set(x_97, 0, x_96); +lean_ctor_set(x_97, 1, x_95); +x_98 = lean_unsigned_to_nat(0u); +x_99 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_98); +x_100 = l_Lean_Meta_Grind_ppPattern___closed__5; +lean_inc(x_99); +x_101 = lean_mk_array(x_99, x_100); +x_102 = lean_unsigned_to_nat(1u); +x_103 = lean_nat_sub(x_99, x_102); +lean_dec(x_99); +x_104 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_101, x_103); +x_105 = lean_box(0); +x_106 = lean_array_size(x_104); +x_107 = 0; +x_108 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__6(x_104, x_105, x_104, x_106, x_107, x_97); +lean_dec(x_104); +return x_108; +} +case 7: +{ +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; size_t x_122; size_t x_123; lean_object* x_124; +x_109 = l_Lean_Expr_getAppFn(x_1); +x_110 = l_Lean_MessageData_ofExpr(x_109); +x_111 = l_Lean_Meta_Grind_ppPattern___closed__4; +x_112 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_112, 0, x_111); +lean_ctor_set(x_112, 1, x_110); +x_113 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_113, 0, x_112); +lean_ctor_set(x_113, 1, x_111); +x_114 = lean_unsigned_to_nat(0u); +x_115 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_114); +x_116 = l_Lean_Meta_Grind_ppPattern___closed__5; +lean_inc(x_115); +x_117 = lean_mk_array(x_115, x_116); +x_118 = lean_unsigned_to_nat(1u); +x_119 = lean_nat_sub(x_115, x_118); +lean_dec(x_115); +x_120 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_117, x_119); +x_121 = lean_box(0); +x_122 = lean_array_size(x_120); +x_123 = 0; +x_124 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__7(x_120, x_121, x_120, x_122, x_123, x_113); +lean_dec(x_120); +return x_124; +} +case 8: +{ +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; size_t x_138; size_t x_139; lean_object* x_140; +x_125 = l_Lean_Expr_getAppFn(x_1); +x_126 = l_Lean_MessageData_ofExpr(x_125); +x_127 = l_Lean_Meta_Grind_ppPattern___closed__4; +x_128 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_128, 0, x_127); +lean_ctor_set(x_128, 1, x_126); +x_129 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_129, 0, x_128); +lean_ctor_set(x_129, 1, x_127); +x_130 = lean_unsigned_to_nat(0u); +x_131 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_130); +x_132 = l_Lean_Meta_Grind_ppPattern___closed__5; +lean_inc(x_131); +x_133 = lean_mk_array(x_131, x_132); +x_134 = lean_unsigned_to_nat(1u); +x_135 = lean_nat_sub(x_131, x_134); +lean_dec(x_131); +x_136 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_133, x_135); +x_137 = lean_box(0); +x_138 = lean_array_size(x_136); +x_139 = 0; +x_140 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__8(x_136, x_137, x_136, x_138, x_139, x_129); +lean_dec(x_136); +return x_140; +} +case 9: +{ +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; size_t x_154; size_t x_155; lean_object* x_156; +x_141 = l_Lean_Expr_getAppFn(x_1); +x_142 = l_Lean_MessageData_ofExpr(x_141); +x_143 = l_Lean_Meta_Grind_ppPattern___closed__4; +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 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_145, 0, x_144); +lean_ctor_set(x_145, 1, x_143); +x_146 = lean_unsigned_to_nat(0u); +x_147 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_146); +x_148 = l_Lean_Meta_Grind_ppPattern___closed__5; +lean_inc(x_147); +x_149 = lean_mk_array(x_147, x_148); +x_150 = lean_unsigned_to_nat(1u); +x_151 = lean_nat_sub(x_147, x_150); +lean_dec(x_147); +x_152 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_149, x_151); +x_153 = lean_box(0); +x_154 = lean_array_size(x_152); +x_155 = 0; +x_156 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__9(x_152, x_153, x_152, x_154, x_155, x_145); +lean_dec(x_152); +return x_156; +} +case 10: +{ +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; size_t x_170; size_t x_171; lean_object* x_172; +x_157 = l_Lean_Expr_getAppFn(x_1); +x_158 = l_Lean_MessageData_ofExpr(x_157); +x_159 = l_Lean_Meta_Grind_ppPattern___closed__4; +x_160 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_160, 0, x_159); +lean_ctor_set(x_160, 1, x_158); +x_161 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_161, 0, x_160); +lean_ctor_set(x_161, 1, x_159); +x_162 = lean_unsigned_to_nat(0u); +x_163 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_162); +x_164 = l_Lean_Meta_Grind_ppPattern___closed__5; +lean_inc(x_163); +x_165 = lean_mk_array(x_163, x_164); +x_166 = lean_unsigned_to_nat(1u); +x_167 = lean_nat_sub(x_163, x_166); +lean_dec(x_163); +x_168 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_165, x_167); +x_169 = lean_box(0); +x_170 = lean_array_size(x_168); +x_171 = 0; +x_172 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__10(x_168, x_169, x_168, x_170, x_171, x_161); +lean_dec(x_168); +return x_172; +} +default: +{ +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; size_t x_186; size_t x_187; lean_object* x_188; +x_173 = l_Lean_Expr_getAppFn(x_1); +x_174 = l_Lean_MessageData_ofExpr(x_173); +x_175 = l_Lean_Meta_Grind_ppPattern___closed__4; +x_176 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_176, 0, x_175); +lean_ctor_set(x_176, 1, x_174); +x_177 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_177, 0, x_176); +lean_ctor_set(x_177, 1, x_175); +x_178 = lean_unsigned_to_nat(0u); +x_179 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_178); +x_180 = l_Lean_Meta_Grind_ppPattern___closed__5; +lean_inc(x_179); +x_181 = lean_mk_array(x_179, x_180); +x_182 = lean_unsigned_to_nat(1u); +x_183 = lean_nat_sub(x_179, x_182); +lean_dec(x_179); +x_184 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_181, x_183); +x_185 = lean_box(0); +x_186 = lean_array_size(x_184); +x_187 = 0; +x_188 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__11(x_184, x_185, x_184, x_186, x_187, x_177); +lean_dec(x_184); +return x_188; +} +} +} +else +{ +lean_object* x_189; +lean_dec(x_1); +x_189 = l_Lean_Meta_Grind_ppPattern___closed__7; +return x_189; +} +} +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_dec(x_1); +x_190 = lean_ctor_get(x_2, 0); +lean_inc(x_190); +lean_dec(x_2); +x_191 = l_Lean_MessageData_ofExpr(x_190); +x_192 = l_Lean_Meta_Grind_ppPattern___closed__9; +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_Meta_Grind_ppPattern___closed__11; +x_195 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_195, 0, x_193); +lean_ctor_set(x_195, 1, x_194); +return x_195; +} +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1(x_1, x_2, x_3); +lean_dec(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___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_7; size_t x_8; lean_object* x_9; +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_9 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1(x_1, x_2, x_3, x_7, x_8, x_6); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +size_t x_7; size_t x_8; lean_object* x_9; +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_9 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__2(x_1, x_2, x_3, x_7, x_8, x_6); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +size_t x_7; size_t x_8; lean_object* x_9; +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_9 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__3(x_1, x_2, x_3, x_7, x_8, x_6); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +size_t x_7; size_t x_8; lean_object* x_9; +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_9 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__4(x_1, x_2, x_3, x_7, x_8, x_6); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +size_t x_7; size_t x_8; lean_object* x_9; +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_9 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__5(x_1, x_2, x_3, x_7, x_8, x_6); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___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; size_t x_8; lean_object* x_9; +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_9 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__6(x_1, x_2, x_3, x_7, x_8, x_6); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +size_t x_7; size_t x_8; lean_object* x_9; +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_9 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__7(x_1, x_2, x_3, x_7, x_8, x_6); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +size_t x_7; size_t x_8; lean_object* x_9; +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_9 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__8(x_1, x_2, x_3, x_7, x_8, x_6); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___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) { +_start: +{ +size_t x_7; size_t x_8; lean_object* x_9; +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_9 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__9(x_1, x_2, x_3, x_7, x_8, x_6); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___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: +{ +size_t x_7; size_t x_8; lean_object* x_9; +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_9 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__10(x_1, x_2, x_3, x_7, x_8, x_6); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___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: +{ +size_t x_7; size_t x_8; lean_object* x_9; +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_9 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__11(x_1, x_2, x_3, x_7, x_8, x_6); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_9; +} +} +LEAN_EXPORT uint8_t l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +uint8_t x_3; +x_3 = 0; +return x_3; +} +else +{ +lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_4 = lean_ctor_get(x_2, 0); +x_5 = lean_ctor_get(x_2, 2); +x_6 = l___private_Lean_HeadIndex_0__Lean_beqHeadIndex____x40_Lean_HeadIndex___hyg_69_(x_4, x_1); +if (x_6 == 0) +{ +x_2 = x_5; +goto _start; +} +else +{ +uint8_t x_8; +x_8 = 1; +return x_8; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__4(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +return x_1; +} +else +{ +uint8_t x_3; +x_3 = !lean_is_exclusive(x_2); +if (x_3 == 0) +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; uint64_t x_7; uint64_t x_8; uint64_t x_9; uint64_t x_10; uint64_t x_11; uint64_t x_12; uint64_t x_13; size_t x_14; size_t x_15; size_t x_16; size_t x_17; size_t x_18; lean_object* x_19; lean_object* x_20; +x_4 = lean_ctor_get(x_2, 0); +x_5 = lean_ctor_get(x_2, 2); +x_6 = lean_array_get_size(x_1); +x_7 = l_Lean_HeadIndex_hash(x_4); +x_8 = 32; +x_9 = lean_uint64_shift_right(x_7, x_8); +x_10 = lean_uint64_xor(x_7, x_9); +x_11 = 16; +x_12 = lean_uint64_shift_right(x_10, x_11); +x_13 = lean_uint64_xor(x_10, x_12); +x_14 = lean_uint64_to_usize(x_13); +x_15 = lean_usize_of_nat(x_6); +lean_dec(x_6); +x_16 = 1; +x_17 = lean_usize_sub(x_15, x_16); +x_18 = lean_usize_land(x_14, x_17); +x_19 = lean_array_uget(x_1, x_18); +lean_ctor_set(x_2, 2, x_19); +x_20 = lean_array_uset(x_1, x_18, x_2); +x_1 = x_20; +x_2 = x_5; +goto _start; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint64_t x_26; uint64_t x_27; uint64_t x_28; uint64_t x_29; uint64_t x_30; uint64_t x_31; uint64_t x_32; size_t x_33; size_t x_34; size_t x_35; size_t x_36; size_t x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_22 = lean_ctor_get(x_2, 0); +x_23 = lean_ctor_get(x_2, 1); +x_24 = lean_ctor_get(x_2, 2); +lean_inc(x_24); +lean_inc(x_23); +lean_inc(x_22); +lean_dec(x_2); +x_25 = lean_array_get_size(x_1); +x_26 = l_Lean_HeadIndex_hash(x_22); +x_27 = 32; +x_28 = lean_uint64_shift_right(x_26, x_27); +x_29 = lean_uint64_xor(x_26, x_28); +x_30 = 16; +x_31 = lean_uint64_shift_right(x_29, x_30); +x_32 = lean_uint64_xor(x_29, x_31); +x_33 = lean_uint64_to_usize(x_32); +x_34 = lean_usize_of_nat(x_25); +lean_dec(x_25); +x_35 = 1; +x_36 = lean_usize_sub(x_34, x_35); +x_37 = lean_usize_land(x_33, x_36); +x_38 = lean_array_uget(x_1, x_37); +x_39 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_39, 0, x_22); +lean_ctor_set(x_39, 1, x_23); +lean_ctor_set(x_39, 2, x_38); +x_40 = lean_array_uset(x_1, x_37, x_39); +x_1 = x_40; +x_2 = x_24; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand_go___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; uint8_t x_5; +x_4 = lean_array_get_size(x_2); +x_5 = lean_nat_dec_lt(x_1, x_4); +lean_dec(x_4); +if (x_5 == 0) +{ +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +else +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_6 = lean_array_fget(x_2, x_1); +x_7 = lean_box(0); +x_8 = lean_array_fset(x_2, x_1, x_7); +x_9 = l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__4(x_3, x_6); +x_10 = lean_unsigned_to_nat(1u); +x_11 = lean_nat_add(x_1, x_10); +lean_dec(x_1); +x_1 = x_11; +x_2 = x_8; +x_3 = x_9; +goto _start; +} +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__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; +x_2 = lean_array_get_size(x_1); +x_3 = lean_unsigned_to_nat(2u); +x_4 = lean_nat_mul(x_2, x_3); +lean_dec(x_2); +x_5 = lean_box(0); +x_6 = lean_mk_array(x_4, x_5); +x_7 = lean_unsigned_to_nat(0u); +x_8 = l_Std_DHashMap_Internal_Raw_u2080_expand_go___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__3(x_7, x_1, x_6); +return x_8; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_8 = lean_st_ref_get(x_2, x_7); +x_9 = lean_ctor_get(x_8, 0); +lean_inc(x_9); +x_10 = lean_ctor_get(x_9, 1); +lean_inc(x_10); +lean_dec(x_9); +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; uint64_t x_16; uint64_t x_17; uint64_t x_18; uint64_t x_19; uint64_t x_20; uint64_t x_21; uint64_t x_22; size_t x_23; size_t x_24; size_t x_25; size_t x_26; size_t x_27; lean_object* x_28; uint8_t x_29; +x_12 = lean_ctor_get(x_8, 1); +x_13 = lean_ctor_get(x_8, 0); +lean_dec(x_13); +x_14 = lean_ctor_get(x_10, 1); +lean_inc(x_14); +lean_dec(x_10); +x_15 = lean_array_get_size(x_14); +x_16 = l_Lean_HeadIndex_hash(x_1); +x_17 = 32; +x_18 = lean_uint64_shift_right(x_16, x_17); +x_19 = lean_uint64_xor(x_16, x_18); +x_20 = 16; +x_21 = lean_uint64_shift_right(x_19, x_20); +x_22 = lean_uint64_xor(x_19, x_21); +x_23 = lean_uint64_to_usize(x_22); +x_24 = lean_usize_of_nat(x_15); +lean_dec(x_15); +x_25 = 1; +x_26 = lean_usize_sub(x_24, x_25); +x_27 = lean_usize_land(x_23, x_26); +x_28 = lean_array_uget(x_14, x_27); +lean_dec(x_14); +x_29 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__1(x_1, x_28); +lean_dec(x_28); +if (x_29 == 0) +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; +lean_free_object(x_8); +x_30 = lean_st_ref_take(x_2, x_12); +x_31 = lean_ctor_get(x_30, 0); +lean_inc(x_31); +x_32 = lean_ctor_get(x_30, 1); +lean_inc(x_32); +lean_dec(x_30); +x_33 = !lean_is_exclusive(x_31); +if (x_33 == 0) +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; +x_34 = lean_ctor_get(x_31, 0); +x_35 = lean_ctor_get(x_31, 1); +lean_inc(x_1); +x_36 = lean_array_push(x_34, x_1); +x_37 = !lean_is_exclusive(x_35); +if (x_37 == 0) +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; size_t x_41; size_t x_42; size_t x_43; lean_object* x_44; uint8_t x_45; +x_38 = lean_ctor_get(x_35, 0); +x_39 = lean_ctor_get(x_35, 1); +x_40 = lean_array_get_size(x_39); +x_41 = lean_usize_of_nat(x_40); +lean_dec(x_40); +x_42 = lean_usize_sub(x_41, x_25); +x_43 = lean_usize_land(x_23, x_42); +x_44 = lean_array_uget(x_39, x_43); +x_45 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__1(x_1, x_44); +if (x_45 == 0) +{ +lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; uint8_t x_56; +x_46 = lean_unsigned_to_nat(1u); +x_47 = lean_nat_add(x_38, x_46); +lean_dec(x_38); +x_48 = lean_box(0); +x_49 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_49, 0, x_1); +lean_ctor_set(x_49, 1, x_48); +lean_ctor_set(x_49, 2, x_44); +x_50 = lean_array_uset(x_39, x_43, x_49); +x_51 = lean_unsigned_to_nat(4u); +x_52 = lean_nat_mul(x_47, x_51); +x_53 = lean_unsigned_to_nat(3u); +x_54 = lean_nat_div(x_52, x_53); +lean_dec(x_52); +x_55 = lean_array_get_size(x_50); +x_56 = lean_nat_dec_le(x_54, x_55); +lean_dec(x_55); +lean_dec(x_54); +if (x_56 == 0) +{ +lean_object* x_57; lean_object* x_58; uint8_t x_59; +x_57 = l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__2(x_50); +lean_ctor_set(x_35, 1, x_57); +lean_ctor_set(x_35, 0, x_47); +lean_ctor_set(x_31, 0, x_36); +x_58 = lean_st_ref_set(x_2, x_31, x_32); +x_59 = !lean_is_exclusive(x_58); +if (x_59 == 0) +{ +lean_object* x_60; +x_60 = lean_ctor_get(x_58, 0); +lean_dec(x_60); +lean_ctor_set(x_58, 0, x_48); +return x_58; +} +else +{ +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(0, 2, 0); +lean_ctor_set(x_62, 0, x_48); +lean_ctor_set(x_62, 1, x_61); +return x_62; +} +} +else +{ +lean_object* x_63; uint8_t x_64; +lean_ctor_set(x_35, 1, x_50); +lean_ctor_set(x_35, 0, x_47); +lean_ctor_set(x_31, 0, x_36); +x_63 = lean_st_ref_set(x_2, x_31, x_32); +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(x_63, 0, x_48); +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(0, 2, 0); +lean_ctor_set(x_67, 0, x_48); +lean_ctor_set(x_67, 1, x_66); +return x_67; +} +} +} +else +{ +lean_object* x_68; uint8_t x_69; +lean_dec(x_44); +lean_dec(x_1); +lean_ctor_set(x_31, 0, x_36); +x_68 = lean_st_ref_set(x_2, x_31, x_32); +x_69 = !lean_is_exclusive(x_68); +if (x_69 == 0) +{ +lean_object* x_70; lean_object* x_71; +x_70 = lean_ctor_get(x_68, 0); +lean_dec(x_70); +x_71 = lean_box(0); +lean_ctor_set(x_68, 0, x_71); +return x_68; +} +else +{ +lean_object* x_72; lean_object* x_73; lean_object* x_74; +x_72 = lean_ctor_get(x_68, 1); +lean_inc(x_72); +lean_dec(x_68); +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_72); +return x_74; +} +} +} +else +{ +lean_object* x_75; lean_object* x_76; lean_object* x_77; size_t x_78; size_t x_79; size_t x_80; lean_object* x_81; uint8_t x_82; +x_75 = lean_ctor_get(x_35, 0); +x_76 = lean_ctor_get(x_35, 1); +lean_inc(x_76); +lean_inc(x_75); +lean_dec(x_35); +x_77 = lean_array_get_size(x_76); +x_78 = lean_usize_of_nat(x_77); +lean_dec(x_77); +x_79 = lean_usize_sub(x_78, x_25); +x_80 = lean_usize_land(x_23, x_79); +x_81 = lean_array_uget(x_76, x_80); +x_82 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__1(x_1, x_81); +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; uint8_t x_93; +x_83 = lean_unsigned_to_nat(1u); +x_84 = lean_nat_add(x_75, x_83); +lean_dec(x_75); +x_85 = lean_box(0); +x_86 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_86, 0, x_1); +lean_ctor_set(x_86, 1, x_85); +lean_ctor_set(x_86, 2, x_81); +x_87 = lean_array_uset(x_76, x_80, x_86); +x_88 = lean_unsigned_to_nat(4u); +x_89 = lean_nat_mul(x_84, x_88); +x_90 = lean_unsigned_to_nat(3u); +x_91 = lean_nat_div(x_89, x_90); +lean_dec(x_89); +x_92 = lean_array_get_size(x_87); +x_93 = lean_nat_dec_le(x_91, x_92); +lean_dec(x_92); +lean_dec(x_91); +if (x_93 == 0) +{ +lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; +x_94 = l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__2(x_87); +x_95 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_95, 0, x_84); +lean_ctor_set(x_95, 1, x_94); +lean_ctor_set(x_31, 1, x_95); +lean_ctor_set(x_31, 0, x_36); +x_96 = lean_st_ref_set(x_2, x_31, x_32); +x_97 = lean_ctor_get(x_96, 1); +lean_inc(x_97); +if (lean_is_exclusive(x_96)) { + lean_ctor_release(x_96, 0); + lean_ctor_release(x_96, 1); + x_98 = x_96; +} else { + lean_dec_ref(x_96); + x_98 = lean_box(0); +} +if (lean_is_scalar(x_98)) { + x_99 = lean_alloc_ctor(0, 2, 0); +} else { + x_99 = x_98; +} +lean_ctor_set(x_99, 0, x_85); +lean_ctor_set(x_99, 1, x_97); +return x_99; +} +else +{ +lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; +x_100 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_100, 0, x_84); +lean_ctor_set(x_100, 1, x_87); +lean_ctor_set(x_31, 1, x_100); +lean_ctor_set(x_31, 0, x_36); +x_101 = lean_st_ref_set(x_2, x_31, x_32); +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; +} else { + lean_dec_ref(x_101); + x_103 = lean_box(0); +} +if (lean_is_scalar(x_103)) { + x_104 = lean_alloc_ctor(0, 2, 0); +} else { + x_104 = x_103; +} +lean_ctor_set(x_104, 0, x_85); +lean_ctor_set(x_104, 1, x_102); +return x_104; +} +} +else +{ +lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; +lean_dec(x_81); +lean_dec(x_1); +x_105 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_105, 0, x_75); +lean_ctor_set(x_105, 1, x_76); +lean_ctor_set(x_31, 1, x_105); +lean_ctor_set(x_31, 0, x_36); +x_106 = lean_st_ref_set(x_2, x_31, x_32); +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); +} +x_109 = lean_box(0); +if (lean_is_scalar(x_108)) { + x_110 = lean_alloc_ctor(0, 2, 0); +} else { + x_110 = x_108; +} +lean_ctor_set(x_110, 0, x_109); +lean_ctor_set(x_110, 1, x_107); +return x_110; +} +} +} +else +{ +lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; size_t x_119; size_t x_120; size_t x_121; lean_object* x_122; uint8_t x_123; +x_111 = lean_ctor_get(x_31, 0); +x_112 = lean_ctor_get(x_31, 1); +x_113 = lean_ctor_get(x_31, 2); +lean_inc(x_113); +lean_inc(x_112); +lean_inc(x_111); +lean_dec(x_31); +lean_inc(x_1); +x_114 = lean_array_push(x_111, x_1); +x_115 = lean_ctor_get(x_112, 0); +lean_inc(x_115); +x_116 = lean_ctor_get(x_112, 1); +lean_inc(x_116); +if (lean_is_exclusive(x_112)) { + lean_ctor_release(x_112, 0); + lean_ctor_release(x_112, 1); + x_117 = x_112; +} else { + lean_dec_ref(x_112); + x_117 = lean_box(0); +} +x_118 = lean_array_get_size(x_116); +x_119 = lean_usize_of_nat(x_118); +lean_dec(x_118); +x_120 = lean_usize_sub(x_119, x_25); +x_121 = lean_usize_land(x_23, x_120); +x_122 = lean_array_uget(x_116, x_121); +x_123 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__1(x_1, x_122); +if (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; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; uint8_t x_134; +x_124 = lean_unsigned_to_nat(1u); +x_125 = lean_nat_add(x_115, x_124); +lean_dec(x_115); +x_126 = lean_box(0); +x_127 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_127, 0, x_1); +lean_ctor_set(x_127, 1, x_126); +lean_ctor_set(x_127, 2, x_122); +x_128 = lean_array_uset(x_116, x_121, x_127); +x_129 = lean_unsigned_to_nat(4u); +x_130 = lean_nat_mul(x_125, x_129); +x_131 = lean_unsigned_to_nat(3u); +x_132 = lean_nat_div(x_130, x_131); +lean_dec(x_130); +x_133 = lean_array_get_size(x_128); +x_134 = lean_nat_dec_le(x_132, x_133); +lean_dec(x_133); +lean_dec(x_132); +if (x_134 == 0) +{ +lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; +x_135 = l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__2(x_128); +if (lean_is_scalar(x_117)) { + x_136 = lean_alloc_ctor(0, 2, 0); +} else { + x_136 = x_117; +} +lean_ctor_set(x_136, 0, x_125); +lean_ctor_set(x_136, 1, x_135); +x_137 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_137, 0, x_114); +lean_ctor_set(x_137, 1, x_136); +lean_ctor_set(x_137, 2, x_113); +x_138 = lean_st_ref_set(x_2, x_137, x_32); +x_139 = lean_ctor_get(x_138, 1); +lean_inc(x_139); +if (lean_is_exclusive(x_138)) { + lean_ctor_release(x_138, 0); + lean_ctor_release(x_138, 1); + x_140 = x_138; +} else { + lean_dec_ref(x_138); + x_140 = lean_box(0); +} +if (lean_is_scalar(x_140)) { + x_141 = lean_alloc_ctor(0, 2, 0); +} else { + x_141 = x_140; +} +lean_ctor_set(x_141, 0, x_126); +lean_ctor_set(x_141, 1, x_139); +return x_141; +} +else +{ +lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; +if (lean_is_scalar(x_117)) { + x_142 = lean_alloc_ctor(0, 2, 0); +} else { + x_142 = x_117; +} +lean_ctor_set(x_142, 0, x_125); +lean_ctor_set(x_142, 1, x_128); +x_143 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_143, 0, x_114); +lean_ctor_set(x_143, 1, x_142); +lean_ctor_set(x_143, 2, x_113); +x_144 = lean_st_ref_set(x_2, x_143, x_32); +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_126); +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_object* x_152; lean_object* x_153; lean_object* x_154; +lean_dec(x_122); +lean_dec(x_1); +if (lean_is_scalar(x_117)) { + x_148 = lean_alloc_ctor(0, 2, 0); +} else { + x_148 = x_117; +} +lean_ctor_set(x_148, 0, x_115); +lean_ctor_set(x_148, 1, x_116); +x_149 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_149, 0, x_114); +lean_ctor_set(x_149, 1, x_148); +lean_ctor_set(x_149, 2, x_113); +x_150 = lean_st_ref_set(x_2, x_149, x_32); +x_151 = lean_ctor_get(x_150, 1); +lean_inc(x_151); +if (lean_is_exclusive(x_150)) { + lean_ctor_release(x_150, 0); + lean_ctor_release(x_150, 1); + x_152 = x_150; +} else { + lean_dec_ref(x_150); + x_152 = lean_box(0); +} +x_153 = lean_box(0); +if (lean_is_scalar(x_152)) { + x_154 = lean_alloc_ctor(0, 2, 0); +} else { + x_154 = x_152; +} +lean_ctor_set(x_154, 0, x_153); +lean_ctor_set(x_154, 1, x_151); +return x_154; +} +} +} +else +{ +lean_object* x_155; +lean_dec(x_1); +x_155 = lean_box(0); +lean_ctor_set(x_8, 0, x_155); +return x_8; +} +} +else +{ +lean_object* x_156; lean_object* x_157; lean_object* x_158; uint64_t x_159; uint64_t x_160; uint64_t x_161; uint64_t x_162; uint64_t x_163; uint64_t x_164; uint64_t x_165; size_t x_166; size_t x_167; size_t x_168; size_t x_169; size_t x_170; lean_object* x_171; uint8_t x_172; +x_156 = lean_ctor_get(x_8, 1); +lean_inc(x_156); +lean_dec(x_8); +x_157 = lean_ctor_get(x_10, 1); +lean_inc(x_157); +lean_dec(x_10); +x_158 = lean_array_get_size(x_157); +x_159 = l_Lean_HeadIndex_hash(x_1); +x_160 = 32; +x_161 = lean_uint64_shift_right(x_159, x_160); +x_162 = lean_uint64_xor(x_159, x_161); +x_163 = 16; +x_164 = lean_uint64_shift_right(x_162, x_163); +x_165 = lean_uint64_xor(x_162, x_164); +x_166 = lean_uint64_to_usize(x_165); +x_167 = lean_usize_of_nat(x_158); +lean_dec(x_158); +x_168 = 1; +x_169 = lean_usize_sub(x_167, x_168); +x_170 = lean_usize_land(x_166, x_169); +x_171 = lean_array_uget(x_157, x_170); +lean_dec(x_157); +x_172 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__1(x_1, x_171); +lean_dec(x_171); +if (x_172 == 0) +{ +lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; size_t x_185; size_t x_186; size_t x_187; lean_object* x_188; uint8_t x_189; +x_173 = lean_st_ref_take(x_2, x_156); +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); +if (lean_is_exclusive(x_174)) { + lean_ctor_release(x_174, 0); + lean_ctor_release(x_174, 1); + lean_ctor_release(x_174, 2); + x_179 = x_174; +} else { + lean_dec_ref(x_174); + x_179 = lean_box(0); +} +lean_inc(x_1); +x_180 = lean_array_push(x_176, x_1); +x_181 = lean_ctor_get(x_177, 0); +lean_inc(x_181); +x_182 = lean_ctor_get(x_177, 1); +lean_inc(x_182); +if (lean_is_exclusive(x_177)) { + lean_ctor_release(x_177, 0); + lean_ctor_release(x_177, 1); + x_183 = x_177; +} else { + lean_dec_ref(x_177); + x_183 = lean_box(0); +} +x_184 = lean_array_get_size(x_182); +x_185 = lean_usize_of_nat(x_184); +lean_dec(x_184); +x_186 = lean_usize_sub(x_185, x_168); +x_187 = lean_usize_land(x_166, x_186); +x_188 = lean_array_uget(x_182, x_187); +x_189 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__1(x_1, x_188); +if (x_189 == 0) +{ +lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; uint8_t x_200; +x_190 = lean_unsigned_to_nat(1u); +x_191 = lean_nat_add(x_181, x_190); +lean_dec(x_181); +x_192 = lean_box(0); +x_193 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_193, 0, x_1); +lean_ctor_set(x_193, 1, x_192); +lean_ctor_set(x_193, 2, x_188); +x_194 = lean_array_uset(x_182, x_187, x_193); +x_195 = lean_unsigned_to_nat(4u); +x_196 = lean_nat_mul(x_191, x_195); +x_197 = lean_unsigned_to_nat(3u); +x_198 = lean_nat_div(x_196, x_197); +lean_dec(x_196); +x_199 = lean_array_get_size(x_194); +x_200 = lean_nat_dec_le(x_198, x_199); +lean_dec(x_199); +lean_dec(x_198); +if (x_200 == 0) +{ +lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; +x_201 = l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__2(x_194); +if (lean_is_scalar(x_183)) { + x_202 = lean_alloc_ctor(0, 2, 0); +} else { + x_202 = x_183; +} +lean_ctor_set(x_202, 0, x_191); +lean_ctor_set(x_202, 1, x_201); +if (lean_is_scalar(x_179)) { + x_203 = lean_alloc_ctor(0, 3, 0); +} else { + x_203 = x_179; +} +lean_ctor_set(x_203, 0, x_180); +lean_ctor_set(x_203, 1, x_202); +lean_ctor_set(x_203, 2, x_178); +x_204 = lean_st_ref_set(x_2, x_203, x_175); +x_205 = lean_ctor_get(x_204, 1); +lean_inc(x_205); +if (lean_is_exclusive(x_204)) { + lean_ctor_release(x_204, 0); + lean_ctor_release(x_204, 1); + x_206 = x_204; +} else { + lean_dec_ref(x_204); + x_206 = lean_box(0); +} +if (lean_is_scalar(x_206)) { + x_207 = lean_alloc_ctor(0, 2, 0); +} else { + x_207 = x_206; +} +lean_ctor_set(x_207, 0, x_192); +lean_ctor_set(x_207, 1, x_205); +return x_207; +} +else +{ +lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; +if (lean_is_scalar(x_183)) { + x_208 = lean_alloc_ctor(0, 2, 0); +} else { + x_208 = x_183; +} +lean_ctor_set(x_208, 0, x_191); +lean_ctor_set(x_208, 1, x_194); +if (lean_is_scalar(x_179)) { + x_209 = lean_alloc_ctor(0, 3, 0); +} else { + x_209 = x_179; +} +lean_ctor_set(x_209, 0, x_180); +lean_ctor_set(x_209, 1, x_208); +lean_ctor_set(x_209, 2, x_178); +x_210 = lean_st_ref_set(x_2, x_209, x_175); +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_210); + x_212 = lean_box(0); +} +if (lean_is_scalar(x_212)) { + x_213 = lean_alloc_ctor(0, 2, 0); +} else { + x_213 = x_212; +} +lean_ctor_set(x_213, 0, x_192); +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_dec(x_188); +lean_dec(x_1); +if (lean_is_scalar(x_183)) { + x_214 = lean_alloc_ctor(0, 2, 0); +} else { + x_214 = x_183; +} +lean_ctor_set(x_214, 0, x_181); +lean_ctor_set(x_214, 1, x_182); +if (lean_is_scalar(x_179)) { + x_215 = lean_alloc_ctor(0, 3, 0); +} else { + x_215 = x_179; +} +lean_ctor_set(x_215, 0, x_180); +lean_ctor_set(x_215, 1, x_214); +lean_ctor_set(x_215, 2, x_178); +x_216 = lean_st_ref_set(x_2, x_215, x_175); +x_217 = lean_ctor_get(x_216, 1); +lean_inc(x_217); +if (lean_is_exclusive(x_216)) { + lean_ctor_release(x_216, 0); + lean_ctor_release(x_216, 1); + x_218 = x_216; +} else { + lean_dec_ref(x_216); + x_218 = lean_box(0); +} +x_219 = lean_box(0); +if (lean_is_scalar(x_218)) { + x_220 = lean_alloc_ctor(0, 2, 0); +} else { + x_220 = x_218; +} +lean_ctor_set(x_220, 0, x_219); +lean_ctor_set(x_220, 1, x_217); +return x_220; +} +} +else +{ +lean_object* x_221; lean_object* x_222; +lean_dec(x_1); +x_221 = lean_box(0); +x_222 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_222, 0, x_221); +lean_ctor_set(x_222, 1, x_156); +return x_222; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__1(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; +x_8 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_8; +} +} +LEAN_EXPORT uint8_t l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_foundBVar___spec__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +uint8_t x_3; +x_3 = 0; +return x_3; +} +else +{ +lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_4 = lean_ctor_get(x_2, 0); +x_5 = lean_ctor_get(x_2, 2); +x_6 = lean_nat_dec_eq(x_4, x_1); +if (x_6 == 0) +{ +x_2 = x_5; +goto _start; +} +else +{ +uint8_t x_8; +x_8 = 1; +return x_8; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_foundBVar(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_8 = lean_st_ref_get(x_2, x_7); +x_9 = lean_ctor_get(x_8, 0); +lean_inc(x_9); +x_10 = lean_ctor_get(x_9, 2); +lean_inc(x_10); +lean_dec(x_9); +x_11 = !lean_is_exclusive(x_8); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; uint64_t x_15; uint64_t x_16; uint64_t x_17; uint64_t x_18; uint64_t x_19; uint64_t x_20; uint64_t x_21; size_t x_22; size_t x_23; size_t x_24; size_t x_25; size_t x_26; lean_object* x_27; uint8_t x_28; lean_object* x_29; +x_12 = lean_ctor_get(x_8, 0); +lean_dec(x_12); +x_13 = lean_ctor_get(x_10, 1); +lean_inc(x_13); +lean_dec(x_10); +x_14 = lean_array_get_size(x_13); +x_15 = lean_uint64_of_nat(x_1); +x_16 = 32; +x_17 = lean_uint64_shift_right(x_15, x_16); +x_18 = lean_uint64_xor(x_15, x_17); +x_19 = 16; +x_20 = lean_uint64_shift_right(x_18, x_19); +x_21 = lean_uint64_xor(x_18, x_20); +x_22 = lean_uint64_to_usize(x_21); +x_23 = lean_usize_of_nat(x_14); +lean_dec(x_14); +x_24 = 1; +x_25 = lean_usize_sub(x_23, x_24); +x_26 = lean_usize_land(x_22, x_25); +x_27 = lean_array_uget(x_13, x_26); +lean_dec(x_13); +x_28 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_foundBVar___spec__1(x_1, x_27); +lean_dec(x_27); +x_29 = lean_box(x_28); +lean_ctor_set(x_8, 0, x_29); +return x_8; +} +else +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; uint64_t x_33; uint64_t x_34; uint64_t x_35; uint64_t x_36; uint64_t x_37; uint64_t x_38; uint64_t x_39; size_t x_40; size_t x_41; size_t x_42; size_t x_43; size_t x_44; lean_object* x_45; uint8_t x_46; lean_object* x_47; lean_object* x_48; +x_30 = lean_ctor_get(x_8, 1); +lean_inc(x_30); +lean_dec(x_8); +x_31 = lean_ctor_get(x_10, 1); +lean_inc(x_31); +lean_dec(x_10); +x_32 = lean_array_get_size(x_31); +x_33 = lean_uint64_of_nat(x_1); +x_34 = 32; +x_35 = lean_uint64_shift_right(x_33, x_34); +x_36 = lean_uint64_xor(x_33, x_35); +x_37 = 16; +x_38 = lean_uint64_shift_right(x_36, x_37); +x_39 = lean_uint64_xor(x_36, x_38); +x_40 = lean_uint64_to_usize(x_39); +x_41 = lean_usize_of_nat(x_32); +lean_dec(x_32); +x_42 = 1; +x_43 = lean_usize_sub(x_41, x_42); +x_44 = lean_usize_land(x_40, x_43); +x_45 = lean_array_uget(x_31, x_44); +lean_dec(x_31); +x_46 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_foundBVar___spec__1(x_1, x_45); +lean_dec(x_45); +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_30); +return x_48; +} +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_foundBVar___spec__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_foundBVar___spec__1(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_foundBVar___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; +x_8 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_foundBVar(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +if (lean_obj_tag(x_3) == 0) +{ +lean_dec(x_1); +return x_2; +} +else +{ +uint8_t x_4; +x_4 = !lean_is_exclusive(x_3); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint64_t x_9; uint64_t x_10; uint64_t x_11; uint64_t x_12; uint64_t x_13; uint64_t x_14; uint64_t x_15; size_t x_16; size_t x_17; size_t x_18; size_t x_19; size_t x_20; lean_object* x_21; lean_object* x_22; +x_5 = lean_ctor_get(x_3, 0); +x_6 = lean_ctor_get(x_3, 2); +x_7 = lean_array_get_size(x_2); +lean_inc(x_1); +lean_inc(x_5); +x_8 = lean_apply_1(x_1, x_5); +x_9 = lean_unbox_uint64(x_8); +lean_dec(x_8); +x_10 = 32; +x_11 = lean_uint64_shift_right(x_9, x_10); +x_12 = lean_uint64_xor(x_9, x_11); +x_13 = 16; +x_14 = lean_uint64_shift_right(x_12, x_13); +x_15 = lean_uint64_xor(x_12, x_14); +x_16 = lean_uint64_to_usize(x_15); +x_17 = lean_usize_of_nat(x_7); +lean_dec(x_7); +x_18 = 1; +x_19 = lean_usize_sub(x_17, x_18); +x_20 = lean_usize_land(x_16, x_19); +x_21 = lean_array_uget(x_2, x_20); +lean_ctor_set(x_3, 2, x_21); +x_22 = lean_array_uset(x_2, x_20, x_3); +x_2 = x_22; +x_3 = x_6; +goto _start; +} +else +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; uint64_t x_29; uint64_t x_30; uint64_t x_31; uint64_t x_32; uint64_t x_33; uint64_t x_34; uint64_t x_35; size_t x_36; size_t x_37; size_t x_38; size_t x_39; size_t x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_24 = lean_ctor_get(x_3, 0); +x_25 = lean_ctor_get(x_3, 1); +x_26 = lean_ctor_get(x_3, 2); +lean_inc(x_26); +lean_inc(x_25); +lean_inc(x_24); +lean_dec(x_3); +x_27 = lean_array_get_size(x_2); +lean_inc(x_1); +lean_inc(x_24); +x_28 = lean_apply_1(x_1, x_24); +x_29 = lean_unbox_uint64(x_28); +lean_dec(x_28); +x_30 = 32; +x_31 = lean_uint64_shift_right(x_29, x_30); +x_32 = lean_uint64_xor(x_29, x_31); +x_33 = 16; +x_34 = lean_uint64_shift_right(x_32, x_33); +x_35 = lean_uint64_xor(x_32, x_34); +x_36 = lean_uint64_to_usize(x_35); +x_37 = lean_usize_of_nat(x_27); +lean_dec(x_27); +x_38 = 1; +x_39 = lean_usize_sub(x_37, x_38); +x_40 = lean_usize_land(x_36, x_39); +x_41 = lean_array_uget(x_2, x_40); +x_42 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_42, 0, x_24); +lean_ctor_set(x_42, 1, x_25); +lean_ctor_set(x_42, 2, x_41); +x_43 = lean_array_uset(x_2, x_40, x_42); +x_2 = x_43; +x_3 = x_26; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__3___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__4(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +return x_1; +} +else +{ +uint8_t x_3; +x_3 = !lean_is_exclusive(x_2); +if (x_3 == 0) +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; uint64_t x_7; uint64_t x_8; uint64_t x_9; uint64_t x_10; uint64_t x_11; uint64_t x_12; uint64_t x_13; size_t x_14; size_t x_15; size_t x_16; size_t x_17; size_t x_18; lean_object* x_19; lean_object* x_20; +x_4 = lean_ctor_get(x_2, 0); +x_5 = lean_ctor_get(x_2, 2); +x_6 = lean_array_get_size(x_1); +x_7 = lean_uint64_of_nat(x_4); +x_8 = 32; +x_9 = lean_uint64_shift_right(x_7, x_8); +x_10 = lean_uint64_xor(x_7, x_9); +x_11 = 16; +x_12 = lean_uint64_shift_right(x_10, x_11); +x_13 = lean_uint64_xor(x_10, x_12); +x_14 = lean_uint64_to_usize(x_13); +x_15 = lean_usize_of_nat(x_6); +lean_dec(x_6); +x_16 = 1; +x_17 = lean_usize_sub(x_15, x_16); +x_18 = lean_usize_land(x_14, x_17); +x_19 = lean_array_uget(x_1, x_18); +lean_ctor_set(x_2, 2, x_19); +x_20 = lean_array_uset(x_1, x_18, x_2); +x_1 = x_20; +x_2 = x_5; +goto _start; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint64_t x_26; uint64_t x_27; uint64_t x_28; uint64_t x_29; uint64_t x_30; uint64_t x_31; uint64_t x_32; size_t x_33; size_t x_34; size_t x_35; size_t x_36; size_t x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_22 = lean_ctor_get(x_2, 0); +x_23 = lean_ctor_get(x_2, 1); +x_24 = lean_ctor_get(x_2, 2); +lean_inc(x_24); +lean_inc(x_23); +lean_inc(x_22); +lean_dec(x_2); +x_25 = lean_array_get_size(x_1); +x_26 = lean_uint64_of_nat(x_22); +x_27 = 32; +x_28 = lean_uint64_shift_right(x_26, x_27); +x_29 = lean_uint64_xor(x_26, x_28); +x_30 = 16; +x_31 = lean_uint64_shift_right(x_29, x_30); +x_32 = lean_uint64_xor(x_29, x_31); +x_33 = lean_uint64_to_usize(x_32); +x_34 = lean_usize_of_nat(x_25); +lean_dec(x_25); +x_35 = 1; +x_36 = lean_usize_sub(x_34, x_35); +x_37 = lean_usize_land(x_33, x_36); +x_38 = lean_array_uget(x_1, x_37); +x_39 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_39, 0, x_22); +lean_ctor_set(x_39, 1, x_23); +lean_ctor_set(x_39, 2, x_38); +x_40 = lean_array_uset(x_1, x_37, x_39); +x_1 = x_40; +x_2 = x_24; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand_go___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; uint8_t x_5; +x_4 = lean_array_get_size(x_2); +x_5 = lean_nat_dec_lt(x_1, x_4); +lean_dec(x_4); +if (x_5 == 0) +{ +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +else +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_6 = lean_array_fget(x_2, x_1); +x_7 = lean_box(0); +x_8 = lean_array_fset(x_2, x_1, x_7); +x_9 = l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__3___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__4(x_3, x_6); +x_10 = lean_unsigned_to_nat(1u); +x_11 = lean_nat_add(x_1, x_10); +lean_dec(x_1); +x_1 = x_11; +x_2 = x_8; +x_3 = x_9; +goto _start; +} +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__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; lean_object* x_7; lean_object* x_8; +x_2 = lean_array_get_size(x_1); +x_3 = lean_unsigned_to_nat(2u); +x_4 = lean_nat_mul(x_2, x_3); +lean_dec(x_2); +x_5 = lean_box(0); +x_6 = lean_mk_array(x_4, x_5); +x_7 = lean_unsigned_to_nat(0u); +x_8 = l_Std_DHashMap_Internal_Raw_u2080_expand_go___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__2(x_7, x_1, x_6); +return x_8; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar(lean_object* x_1, 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; +x_8 = lean_st_ref_take(x_2, x_7); +x_9 = lean_ctor_get(x_8, 0); +lean_inc(x_9); +x_10 = lean_ctor_get(x_9, 2); +lean_inc(x_10); +x_11 = lean_ctor_get(x_8, 1); +lean_inc(x_11); +lean_dec(x_8); +x_12 = !lean_is_exclusive(x_9); +if (x_12 == 0) +{ +lean_object* x_13; uint8_t x_14; +x_13 = lean_ctor_get(x_9, 2); +lean_dec(x_13); +x_14 = !lean_is_exclusive(x_10); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; uint64_t x_18; uint64_t x_19; uint64_t x_20; uint64_t x_21; uint64_t x_22; uint64_t x_23; uint64_t x_24; size_t x_25; size_t x_26; size_t x_27; size_t x_28; size_t x_29; lean_object* x_30; uint8_t x_31; +x_15 = lean_ctor_get(x_10, 0); +x_16 = lean_ctor_get(x_10, 1); +x_17 = lean_array_get_size(x_16); +x_18 = lean_uint64_of_nat(x_1); +x_19 = 32; +x_20 = lean_uint64_shift_right(x_18, x_19); +x_21 = lean_uint64_xor(x_18, x_20); +x_22 = 16; +x_23 = lean_uint64_shift_right(x_21, x_22); +x_24 = lean_uint64_xor(x_21, x_23); +x_25 = lean_uint64_to_usize(x_24); +x_26 = lean_usize_of_nat(x_17); +lean_dec(x_17); +x_27 = 1; +x_28 = lean_usize_sub(x_26, x_27); +x_29 = lean_usize_land(x_25, x_28); +x_30 = lean_array_uget(x_16, x_29); +x_31 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_foundBVar___spec__1(x_1, x_30); +if (x_31 == 0) +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; uint8_t x_42; +x_32 = lean_unsigned_to_nat(1u); +x_33 = lean_nat_add(x_15, x_32); +lean_dec(x_15); +x_34 = lean_box(0); +x_35 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_35, 0, x_1); +lean_ctor_set(x_35, 1, x_34); +lean_ctor_set(x_35, 2, x_30); +x_36 = lean_array_uset(x_16, x_29, x_35); +x_37 = lean_unsigned_to_nat(4u); +x_38 = lean_nat_mul(x_33, x_37); +x_39 = lean_unsigned_to_nat(3u); +x_40 = lean_nat_div(x_38, x_39); +lean_dec(x_38); +x_41 = lean_array_get_size(x_36); +x_42 = lean_nat_dec_le(x_40, x_41); +lean_dec(x_41); +lean_dec(x_40); +if (x_42 == 0) +{ +lean_object* x_43; lean_object* x_44; uint8_t x_45; +x_43 = l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__1(x_36); +lean_ctor_set(x_10, 1, x_43); +lean_ctor_set(x_10, 0, x_33); +x_44 = lean_st_ref_set(x_2, x_9, x_11); +x_45 = !lean_is_exclusive(x_44); +if (x_45 == 0) +{ +lean_object* x_46; +x_46 = lean_ctor_get(x_44, 0); +lean_dec(x_46); +lean_ctor_set(x_44, 0, x_34); +return x_44; +} +else +{ +lean_object* x_47; lean_object* x_48; +x_47 = lean_ctor_get(x_44, 1); +lean_inc(x_47); +lean_dec(x_44); +x_48 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_48, 0, x_34); +lean_ctor_set(x_48, 1, x_47); +return x_48; +} +} +else +{ +lean_object* x_49; uint8_t x_50; +lean_ctor_set(x_10, 1, x_36); +lean_ctor_set(x_10, 0, x_33); +x_49 = lean_st_ref_set(x_2, x_9, x_11); +x_50 = !lean_is_exclusive(x_49); +if (x_50 == 0) +{ +lean_object* x_51; +x_51 = lean_ctor_get(x_49, 0); +lean_dec(x_51); +lean_ctor_set(x_49, 0, x_34); +return x_49; +} +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_34); +lean_ctor_set(x_53, 1, x_52); +return x_53; +} +} +} +else +{ +lean_object* x_54; uint8_t x_55; +lean_dec(x_30); +lean_dec(x_1); +x_54 = lean_st_ref_set(x_2, x_9, x_11); +x_55 = !lean_is_exclusive(x_54); +if (x_55 == 0) +{ +lean_object* x_56; lean_object* x_57; +x_56 = lean_ctor_get(x_54, 0); +lean_dec(x_56); +x_57 = lean_box(0); +lean_ctor_set(x_54, 0, x_57); +return x_54; +} +else +{ +lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_58 = lean_ctor_get(x_54, 1); +lean_inc(x_58); +lean_dec(x_54); +x_59 = lean_box(0); +x_60 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_60, 0, x_59); +lean_ctor_set(x_60, 1, x_58); +return x_60; +} +} +} +else +{ +lean_object* x_61; lean_object* x_62; lean_object* x_63; uint64_t x_64; uint64_t x_65; uint64_t x_66; uint64_t x_67; uint64_t x_68; uint64_t x_69; uint64_t x_70; size_t x_71; size_t x_72; size_t x_73; size_t x_74; size_t x_75; lean_object* x_76; uint8_t x_77; +x_61 = lean_ctor_get(x_10, 0); +x_62 = lean_ctor_get(x_10, 1); +lean_inc(x_62); +lean_inc(x_61); +lean_dec(x_10); +x_63 = lean_array_get_size(x_62); +x_64 = lean_uint64_of_nat(x_1); +x_65 = 32; +x_66 = lean_uint64_shift_right(x_64, x_65); +x_67 = lean_uint64_xor(x_64, x_66); +x_68 = 16; +x_69 = lean_uint64_shift_right(x_67, x_68); +x_70 = lean_uint64_xor(x_67, x_69); +x_71 = lean_uint64_to_usize(x_70); +x_72 = lean_usize_of_nat(x_63); +lean_dec(x_63); +x_73 = 1; +x_74 = lean_usize_sub(x_72, x_73); +x_75 = lean_usize_land(x_71, x_74); +x_76 = lean_array_uget(x_62, x_75); +x_77 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_foundBVar___spec__1(x_1, x_76); +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; uint8_t x_88; +x_78 = lean_unsigned_to_nat(1u); +x_79 = lean_nat_add(x_61, x_78); +lean_dec(x_61); +x_80 = lean_box(0); +x_81 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_81, 0, x_1); +lean_ctor_set(x_81, 1, x_80); +lean_ctor_set(x_81, 2, x_76); +x_82 = lean_array_uset(x_62, x_75, x_81); +x_83 = lean_unsigned_to_nat(4u); +x_84 = lean_nat_mul(x_79, x_83); +x_85 = lean_unsigned_to_nat(3u); +x_86 = lean_nat_div(x_84, x_85); +lean_dec(x_84); +x_87 = lean_array_get_size(x_82); +x_88 = lean_nat_dec_le(x_86, x_87); +lean_dec(x_87); +lean_dec(x_86); +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; +x_89 = l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__1(x_82); +x_90 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_90, 0, x_79); +lean_ctor_set(x_90, 1, x_89); +lean_ctor_set(x_9, 2, x_90); +x_91 = lean_st_ref_set(x_2, x_9, x_11); +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_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; +x_95 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_95, 0, x_79); +lean_ctor_set(x_95, 1, x_82); +lean_ctor_set(x_9, 2, x_95); +x_96 = lean_st_ref_set(x_2, x_9, x_11); +x_97 = lean_ctor_get(x_96, 1); +lean_inc(x_97); +if (lean_is_exclusive(x_96)) { + lean_ctor_release(x_96, 0); + lean_ctor_release(x_96, 1); + x_98 = x_96; +} else { + lean_dec_ref(x_96); + x_98 = lean_box(0); +} +if (lean_is_scalar(x_98)) { + x_99 = lean_alloc_ctor(0, 2, 0); +} else { + x_99 = x_98; +} +lean_ctor_set(x_99, 0, x_80); +lean_ctor_set(x_99, 1, x_97); +return x_99; +} +} +else +{ +lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; +lean_dec(x_76); +lean_dec(x_1); +x_100 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_100, 0, x_61); +lean_ctor_set(x_100, 1, x_62); +lean_ctor_set(x_9, 2, x_100); +x_101 = lean_st_ref_set(x_2, x_9, x_11); +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; +} else { + lean_dec_ref(x_101); + x_103 = lean_box(0); +} +x_104 = lean_box(0); +if (lean_is_scalar(x_103)) { + x_105 = lean_alloc_ctor(0, 2, 0); +} else { + x_105 = x_103; +} +lean_ctor_set(x_105, 0, x_104); +lean_ctor_set(x_105, 1, x_102); +return x_105; +} +} +} +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; uint64_t x_112; uint64_t x_113; uint64_t x_114; uint64_t x_115; uint64_t x_116; uint64_t x_117; uint64_t x_118; size_t x_119; size_t x_120; size_t x_121; size_t x_122; size_t x_123; lean_object* x_124; uint8_t x_125; +x_106 = lean_ctor_get(x_9, 0); +x_107 = lean_ctor_get(x_9, 1); +lean_inc(x_107); +lean_inc(x_106); +lean_dec(x_9); +x_108 = lean_ctor_get(x_10, 0); +lean_inc(x_108); +x_109 = lean_ctor_get(x_10, 1); +lean_inc(x_109); +if (lean_is_exclusive(x_10)) { + lean_ctor_release(x_10, 0); + lean_ctor_release(x_10, 1); + x_110 = x_10; +} else { + lean_dec_ref(x_10); + x_110 = lean_box(0); +} +x_111 = lean_array_get_size(x_109); +x_112 = lean_uint64_of_nat(x_1); +x_113 = 32; +x_114 = lean_uint64_shift_right(x_112, x_113); +x_115 = lean_uint64_xor(x_112, x_114); +x_116 = 16; +x_117 = lean_uint64_shift_right(x_115, x_116); +x_118 = lean_uint64_xor(x_115, x_117); +x_119 = lean_uint64_to_usize(x_118); +x_120 = lean_usize_of_nat(x_111); +lean_dec(x_111); +x_121 = 1; +x_122 = lean_usize_sub(x_120, x_121); +x_123 = lean_usize_land(x_119, x_122); +x_124 = lean_array_uget(x_109, x_123); +x_125 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_foundBVar___spec__1(x_1, x_124); +if (x_125 == 0) +{ +lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; uint8_t x_136; +x_126 = lean_unsigned_to_nat(1u); +x_127 = lean_nat_add(x_108, x_126); +lean_dec(x_108); +x_128 = lean_box(0); +x_129 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_129, 0, x_1); +lean_ctor_set(x_129, 1, x_128); +lean_ctor_set(x_129, 2, x_124); +x_130 = lean_array_uset(x_109, x_123, x_129); +x_131 = lean_unsigned_to_nat(4u); +x_132 = lean_nat_mul(x_127, x_131); +x_133 = lean_unsigned_to_nat(3u); +x_134 = lean_nat_div(x_132, x_133); +lean_dec(x_132); +x_135 = lean_array_get_size(x_130); +x_136 = lean_nat_dec_le(x_134, x_135); +lean_dec(x_135); +lean_dec(x_134); +if (x_136 == 0) +{ +lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; +x_137 = l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__1(x_130); +if (lean_is_scalar(x_110)) { + x_138 = lean_alloc_ctor(0, 2, 0); +} else { + x_138 = x_110; +} +lean_ctor_set(x_138, 0, x_127); +lean_ctor_set(x_138, 1, x_137); +x_139 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_139, 0, x_106); +lean_ctor_set(x_139, 1, x_107); +lean_ctor_set(x_139, 2, x_138); +x_140 = lean_st_ref_set(x_2, x_139, x_11); +x_141 = lean_ctor_get(x_140, 1); +lean_inc(x_141); +if (lean_is_exclusive(x_140)) { + lean_ctor_release(x_140, 0); + lean_ctor_release(x_140, 1); + x_142 = x_140; +} else { + lean_dec_ref(x_140); + x_142 = lean_box(0); +} +if (lean_is_scalar(x_142)) { + x_143 = lean_alloc_ctor(0, 2, 0); +} else { + x_143 = x_142; +} +lean_ctor_set(x_143, 0, x_128); +lean_ctor_set(x_143, 1, x_141); +return x_143; +} +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; +if (lean_is_scalar(x_110)) { + x_144 = lean_alloc_ctor(0, 2, 0); +} else { + x_144 = x_110; +} +lean_ctor_set(x_144, 0, x_127); +lean_ctor_set(x_144, 1, x_130); +x_145 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_145, 0, x_106); +lean_ctor_set(x_145, 1, x_107); +lean_ctor_set(x_145, 2, x_144); +x_146 = lean_st_ref_set(x_2, x_145, x_11); +x_147 = lean_ctor_get(x_146, 1); +lean_inc(x_147); +if (lean_is_exclusive(x_146)) { + lean_ctor_release(x_146, 0); + lean_ctor_release(x_146, 1); + x_148 = x_146; +} else { + lean_dec_ref(x_146); + x_148 = lean_box(0); +} +if (lean_is_scalar(x_148)) { + x_149 = lean_alloc_ctor(0, 2, 0); +} else { + x_149 = x_148; +} +lean_ctor_set(x_149, 0, x_128); +lean_ctor_set(x_149, 1, x_147); +return x_149; +} +} +else +{ +lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; +lean_dec(x_124); +lean_dec(x_1); +if (lean_is_scalar(x_110)) { + x_150 = lean_alloc_ctor(0, 2, 0); +} else { + x_150 = x_110; +} +lean_ctor_set(x_150, 0, x_108); +lean_ctor_set(x_150, 1, x_109); +x_151 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_151, 0, x_106); +lean_ctor_set(x_151, 1, x_107); +lean_ctor_set(x_151, 2, x_150); +x_152 = lean_st_ref_set(x_2, x_151, x_11); +x_153 = lean_ctor_get(x_152, 1); +lean_inc(x_153); +if (lean_is_exclusive(x_152)) { + lean_ctor_release(x_152, 0); + lean_ctor_release(x_152, 1); + x_154 = x_152; +} else { + lean_dec_ref(x_152); + x_154 = lean_box(0); +} +x_155 = lean_box(0); +if (lean_is_scalar(x_154)) { + x_156 = lean_alloc_ctor(0, 2, 0); +} else { + x_156 = x_154; +} +lean_ctor_set(x_156, 0, x_155); +lean_ctor_set(x_156, 1, x_153); +return x_156; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; +x_8 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_8; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFn_x3f(lean_object* x_1) { +_start: +{ +uint8_t x_2; +x_2 = l_Lean_Expr_isApp(x_1); +if (x_2 == 0) +{ +lean_object* x_3; +x_3 = lean_box(0); +return x_3; +} +else +{ +lean_object* x_4; +x_4 = l_Lean_Expr_getAppFn(x_1); +switch (lean_obj_tag(x_4)) { +case 1: +{ +lean_object* x_5; +x_5 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_5, 0, x_4); +return x_5; +} +case 4: +{ +lean_object* x_6; lean_object* x_7; uint8_t x_8; +x_6 = lean_ctor_get(x_4, 0); +lean_inc(x_6); +x_7 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames; +x_8 = l_Array_contains___at_Lean_registerInternalExceptionId___spec__1(x_7, x_6); +lean_dec(x_6); +if (x_8 == 0) +{ +lean_object* x_9; +x_9 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_9, 0, x_4); +return x_9; +} +else +{ +lean_object* x_10; +lean_dec(x_4); +x_10 = lean_box(0); +return x_10; +} +} +default: +{ +lean_object* x_11; +lean_dec(x_4); +x_11 = lean_box(0); +return x_11; +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFn_x3f___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFn_x3f(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_Grind_NormalizePattern_getPatternSupportMask___spec__1(size_t x_1, size_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +uint8_t x_9; +x_9 = lean_usize_dec_lt(x_2, x_1); +if (x_9 == 0) +{ +lean_object* x_10; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_10 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_10, 0, x_3); +lean_ctor_set(x_10, 1, x_8); +return x_10; +} +else +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; lean_object* x_15; uint8_t x_22; lean_object* x_23; lean_object* x_38; +x_11 = lean_array_uget(x_3, x_2); +x_12 = lean_unsigned_to_nat(0u); +x_13 = lean_array_uset(x_3, x_2, x_12); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_11); +x_38 = l_Lean_Meta_isProp(x_11, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_38) == 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_41; lean_object* x_42; +x_41 = lean_ctor_get(x_38, 1); +lean_inc(x_41); +lean_dec(x_38); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_11); +x_42 = l_Lean_Meta_isTypeFormer(x_11, x_4, x_5, x_6, x_7, x_41); +if (lean_obj_tag(x_42) == 0) +{ +lean_object* x_43; uint8_t x_44; +x_43 = lean_ctor_get(x_42, 0); +lean_inc(x_43); +x_44 = lean_unbox(x_43); +if (x_44 == 0) +{ +lean_object* x_45; lean_object* x_46; +lean_dec(x_43); +x_45 = lean_ctor_get(x_42, 1); +lean_inc(x_45); +lean_dec(x_42); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_11); +x_46 = l_Lean_Meta_isProof(x_11, x_4, x_5, x_6, x_7, x_45); +if (lean_obj_tag(x_46) == 0) +{ +lean_object* x_47; lean_object* x_48; uint8_t 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 = lean_unbox(x_47); +lean_dec(x_47); +x_22 = x_49; +x_23 = x_48; +goto block_37; +} +else +{ +uint8_t x_50; +lean_dec(x_13); +lean_dec(x_11); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_50 = !lean_is_exclusive(x_46); +if (x_50 == 0) +{ +return x_46; +} +else +{ +lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_51 = lean_ctor_get(x_46, 0); +x_52 = lean_ctor_get(x_46, 1); +lean_inc(x_52); +lean_inc(x_51); +lean_dec(x_46); +x_53 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_53, 0, x_51); +lean_ctor_set(x_53, 1, x_52); +return x_53; +} +} +} +else +{ +lean_object* x_54; uint8_t x_55; +x_54 = lean_ctor_get(x_42, 1); +lean_inc(x_54); +lean_dec(x_42); +x_55 = lean_unbox(x_43); +lean_dec(x_43); +x_22 = x_55; +x_23 = x_54; +goto block_37; +} +} +else +{ +uint8_t x_56; +lean_dec(x_13); +lean_dec(x_11); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_56 = !lean_is_exclusive(x_42); +if (x_56 == 0) +{ +return x_42; +} +else +{ +lean_object* x_57; lean_object* x_58; lean_object* x_59; +x_57 = lean_ctor_get(x_42, 0); +x_58 = lean_ctor_get(x_42, 1); +lean_inc(x_58); +lean_inc(x_57); +lean_dec(x_42); +x_59 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_59, 0, x_57); +lean_ctor_set(x_59, 1, x_58); +return x_59; +} +} +} +else +{ +lean_object* x_60; size_t x_61; size_t x_62; uint8_t x_63; lean_object* x_64; lean_object* x_65; +lean_dec(x_11); +x_60 = lean_ctor_get(x_38, 1); +lean_inc(x_60); +lean_dec(x_38); +x_61 = 1; +x_62 = lean_usize_add(x_2, x_61); +x_63 = 0; +x_64 = lean_box(x_63); +x_65 = lean_array_uset(x_13, x_2, x_64); +x_2 = x_62; +x_3 = x_65; +x_8 = x_60; +goto _start; +} +} +else +{ +uint8_t x_67; +lean_dec(x_13); +lean_dec(x_11); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_67 = !lean_is_exclusive(x_38); +if (x_67 == 0) +{ +return x_38; +} +else +{ +lean_object* x_68; lean_object* x_69; lean_object* x_70; +x_68 = lean_ctor_get(x_38, 0); +x_69 = lean_ctor_get(x_38, 1); +lean_inc(x_69); +lean_inc(x_68); +lean_dec(x_38); +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; +} +} +block_21: +{ +size_t x_16; size_t x_17; lean_object* x_18; lean_object* x_19; +x_16 = 1; +x_17 = lean_usize_add(x_2, x_16); +x_18 = lean_box(x_14); +x_19 = lean_array_uset(x_13, x_2, x_18); +x_2 = x_17; +x_3 = x_19; +x_8 = x_15; +goto _start; +} +block_37: +{ +if (x_22 == 0) +{ +lean_object* x_24; lean_object* x_25; +x_24 = l_Lean_Expr_fvarId_x21(x_11); +lean_dec(x_11); +lean_inc(x_4); +x_25 = l_Lean_FVarId_getDecl(x_24, x_4, x_5, x_6, x_7, x_23); +if (lean_obj_tag(x_25) == 0) +{ +lean_object* x_26; lean_object* x_27; uint8_t 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_LocalDecl_binderInfo(x_26); +lean_dec(x_26); +x_29 = lean_box(x_28); +if (lean_obj_tag(x_29) == 3) +{ +uint8_t x_30; +x_30 = 1; +x_14 = x_30; +x_15 = x_27; +goto block_21; +} +else +{ +uint8_t x_31; +lean_dec(x_29); +x_31 = 0; +x_14 = x_31; +x_15 = x_27; +goto block_21; +} +} +else +{ +uint8_t x_32; +lean_dec(x_13); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_32 = !lean_is_exclusive(x_25); +if (x_32 == 0) +{ +return x_25; +} +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_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; +} +} +} +else +{ +uint8_t x_36; +lean_dec(x_11); +x_36 = 1; +x_14 = x_36; +x_15 = x_23; +goto block_21; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_NormalizePattern_getPatternSupportMask___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +size_t x_8; size_t x_9; lean_object* x_10; +x_8 = lean_array_size(x_1); +x_9 = 0; +x_10 = l_Array_mapMUnsafe_map___at_Lean_Meta_Grind_NormalizePattern_getPatternSupportMask___spec__1(x_8, x_9, x_1, x_3, x_4, x_5, x_6, x_7); +return x_10; +} +} +static lean_object* _init_l_Lean_Meta_Grind_NormalizePattern_getPatternSupportMask___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_NormalizePattern_getPatternSupportMask___lambda__1___boxed), 7, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_NormalizePattern_getPatternSupportMask(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_8 = lean_infer_type(x_1, x_3, x_4, x_5, x_6, x_7); +if (lean_obj_tag(x_8) == 0) +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; +x_9 = lean_ctor_get(x_8, 0); +lean_inc(x_9); +x_10 = lean_ctor_get(x_8, 1); +lean_inc(x_10); +lean_dec(x_8); +x_11 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_11, 0, x_2); +x_12 = l_Lean_Meta_Grind_NormalizePattern_getPatternSupportMask___closed__1; +x_13 = 0; +x_14 = l_Lean_Meta_forallBoundedTelescope___at_Lean_Meta_arrowDomainsN___spec__6___rarg(x_9, x_11, x_12, x_13, x_3, x_4, x_5, x_6, x_10); +return x_14; +} +else +{ +uint8_t x_15; +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_8); +if (x_15 == 0) +{ +return x_8; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_8, 0); +x_17 = lean_ctor_get(x_8, 1); +lean_inc(x_17); +lean_inc(x_16); +lean_dec(x_8); +x_18 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_18, 0, x_16); +lean_ctor_set(x_18, 1, x_17); +return x_18; +} +} +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_Grind_NormalizePattern_getPatternSupportMask___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +size_t x_9; size_t x_10; lean_object* x_11; +x_9 = lean_unbox_usize(x_1); +lean_dec(x_1); +x_10 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_11 = l_Array_mapMUnsafe_map___at_Lean_Meta_Grind_NormalizePattern_getPatternSupportMask___spec__1(x_9, x_10, x_3, x_4, x_5, x_6, x_7, x_8); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_NormalizePattern_getPatternSupportMask___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; +x_8 = l_Lean_Meta_Grind_NormalizePattern_getPatternSupportMask___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_2); +return x_8; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go_goArg(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +uint8_t x_9; +x_9 = l_Lean_Expr_hasLooseBVars(x_1); +if (x_9 == 0) +{ +uint8_t x_10; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_10 = l_Lean_Expr_hasMVar(x_1); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; +x_11 = l_Lean_Meta_Grind_mkGroundPattern(x_1); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_8); +return x_12; +} +else +{ +lean_object* x_13; lean_object* x_14; +lean_dec(x_1); +x_13 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare; +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_8); +return x_14; +} +} +else +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_15; lean_object* x_16; +x_15 = lean_ctor_get(x_1, 0); +lean_inc(x_15); +x_16 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_foundBVar(x_15, x_3, x_4, x_5, x_6, x_7, x_8); +if (x_2 == 0) +{ +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___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar(x_15, x_3, x_4, x_5, x_6, x_7, x_17); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_19 = !lean_is_exclusive(x_18); +if (x_19 == 0) +{ +lean_object* x_20; +x_20 = lean_ctor_get(x_18, 0); +lean_dec(x_20); +lean_ctor_set(x_18, 0, x_1); +return x_18; +} +else +{ +lean_object* x_21; lean_object* x_22; +x_21 = lean_ctor_get(x_18, 1); +lean_inc(x_21); +lean_dec(x_18); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_1); +lean_ctor_set(x_22, 1, x_21); +return x_22; +} +} +else +{ +lean_object* x_23; uint8_t x_24; +x_23 = lean_ctor_get(x_16, 0); +lean_inc(x_23); +x_24 = lean_unbox(x_23); +lean_dec(x_23); +if (x_24 == 0) +{ +lean_object* x_25; lean_object* x_26; uint8_t x_27; +x_25 = lean_ctor_get(x_16, 1); +lean_inc(x_25); +lean_dec(x_16); +x_26 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar(x_15, x_3, x_4, x_5, x_6, x_7, x_25); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +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(x_26, 0, x_1); +return x_26; +} +else +{ +lean_object* x_29; lean_object* x_30; +x_29 = lean_ctor_get(x_26, 1); +lean_inc(x_29); +lean_dec(x_26); +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_1); +lean_ctor_set(x_30, 1, x_29); +return x_30; +} +} +else +{ +uint8_t x_31; +lean_dec(x_15); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_31 = !lean_is_exclusive(x_16); +if (x_31 == 0) +{ +lean_object* x_32; lean_object* x_33; +x_32 = lean_ctor_get(x_16, 0); +lean_dec(x_32); +x_33 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare; +lean_ctor_set(x_16, 0, x_33); +return x_16; +} +else +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_16, 1); +lean_inc(x_34); +lean_dec(x_16); +x_35 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare; +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 +{ +if (x_2 == 0) +{ +lean_object* x_37; +x_37 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFn_x3f(x_1); +if (lean_obj_tag(x_37) == 0) +{ +lean_object* x_38; lean_object* x_39; +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_38 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare; +x_39 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_8); +return x_39; +} +else +{ +uint8_t x_40; lean_object* x_41; +lean_dec(x_37); +x_40 = 0; +x_41 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go(x_1, x_40, x_3, x_4, x_5, x_6, x_7, x_8); +return x_41; +} +} +else +{ +lean_object* x_42; lean_object* x_43; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_42 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare; +x_43 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_43, 0, x_42); +lean_ctor_set(x_43, 1, x_8); +return x_43; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; lean_object* x_9; uint8_t x_10; +x_8 = lean_ctor_get(x_5, 5); +x_9 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_1, x_3, x_4, x_5, x_6, x_7); +x_10 = !lean_is_exclusive(x_9); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; +x_11 = lean_ctor_get(x_9, 0); +lean_inc(x_8); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_8); +lean_ctor_set(x_12, 1, x_11); +lean_ctor_set_tag(x_9, 1); +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); +lean_inc(x_8); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_8); +lean_ctor_set(x_15, 1, x_13); +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_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, 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_ctor_get(x_5, 1); +x_17 = lean_nat_dec_lt(x_7, x_16); +if (x_17 == 0) +{ +lean_object* x_18; +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_7); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_6); +lean_ctor_set(x_18, 1, x_15); +return x_18; +} +else +{ +lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_19 = lean_array_fget(x_6, x_7); +x_20 = lean_array_get_size(x_3); +x_21 = lean_nat_dec_lt(x_7, x_20); +lean_dec(x_20); +if (x_21 == 0) +{ +uint8_t x_22; lean_object* x_23; +x_22 = 0; +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +x_23 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go_goArg(x_19, x_22, x_10, x_11, x_12, x_13, x_14, x_15); +if (lean_obj_tag(x_23) == 0) +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +x_25 = lean_ctor_get(x_23, 1); +lean_inc(x_25); +lean_dec(x_23); +x_26 = lean_array_fset(x_6, x_7, x_24); +x_27 = lean_ctor_get(x_5, 2); +x_28 = lean_nat_add(x_7, x_27); +lean_dec(x_7); +x_6 = x_26; +x_7 = x_28; +x_8 = lean_box(0); +x_9 = lean_box(0); +x_15 = x_25; +goto _start; +} +else +{ +uint8_t x_30; +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_7); +lean_dec(x_6); +x_30 = !lean_is_exclusive(x_23); +if (x_30 == 0) +{ +return x_23; +} +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_23, 0); +x_32 = lean_ctor_get(x_23, 1); +lean_inc(x_32); +lean_inc(x_31); +lean_dec(x_23); +x_33 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_33, 0, x_31); +lean_ctor_set(x_33, 1, x_32); +return x_33; +} +} +} +else +{ +lean_object* x_34; uint8_t x_35; lean_object* x_36; +x_34 = lean_array_fget(x_3, x_7); +x_35 = lean_unbox(x_34); +lean_dec(x_34); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +x_36 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go_goArg(x_19, x_35, x_10, x_11, x_12, x_13, x_14, x_15); +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; +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_fset(x_6, x_7, x_37); +x_40 = lean_ctor_get(x_5, 2); +x_41 = lean_nat_add(x_7, x_40); +lean_dec(x_7); +x_6 = x_39; +x_7 = x_41; +x_8 = lean_box(0); +x_9 = lean_box(0); +x_15 = x_38; +goto _start; +} +else +{ +uint8_t x_43; +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_7); +lean_dec(x_6); +x_43 = !lean_is_exclusive(x_36); +if (x_43 == 0) +{ +return x_36; +} +else +{ +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; +} +} +} +} +} +} +static lean_object* _init_l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__3___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_instMonadMetaM; +x_2 = l_ReaderT_instMonad___rarg(x_1); +return x_2; +} +} +static lean_object* _init_l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__3___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__3___closed__1; +x_2 = l_Lean_instInhabitedExpr; +x_3 = l_instInhabitedOfMonad___rarg(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_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) { +_start: +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_8 = l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__3___closed__2; +x_9 = lean_panic_fn(x_8, x_1); +x_10 = lean_apply_6(x_9, x_2, x_3, x_4, x_5, x_6, x_7); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_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) { +_start: +{ +lean_object* x_8; lean_object* x_9; uint8_t x_10; +x_8 = lean_ctor_get(x_5, 5); +x_9 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_1, x_3, x_4, x_5, x_6, x_7); +x_10 = !lean_is_exclusive(x_9); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; +x_11 = lean_ctor_get(x_9, 0); +lean_inc(x_8); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_8); +lean_ctor_set(x_12, 1, x_11); +lean_ctor_set_tag(x_9, 1); +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); +lean_inc(x_8); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_8); +lean_ctor_set(x_15, 1, x_13); +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; +} +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("invalid pattern, (non-forbidden) application expected", 53, 53); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("assertion violation: ", 21, 21); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("f.isConst || f.isFVar\n ", 24, 24); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__3; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__4; +x_3 = lean_string_append(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("_private.Lean.Meta.Tactic.Grind.EMatchTheorem.0.Lean.Meta.Grind.NormalizePattern.go", 83, 83); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l_Lean_Meta_Grind_EMatchTheorems_insert___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__6; +x_3 = lean_unsigned_to_nat(267u); +x_4 = lean_unsigned_to_nat(2u); +x_5 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__5; +x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); +return x_6; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_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; +x_9 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFn_x3f(x_1); +if (lean_obj_tag(x_9) == 0) +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_10 = l_Lean_indentExpr(x_1); +x_11 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__2; +x_12 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_10); +x_13 = l_Lean_Meta_Grind_ppPattern___closed__4; +x_14 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_14, 0, x_12); +lean_ctor_set(x_14, 1, x_13); +x_15 = l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__1(x_14, 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_15; +} +else +{ +lean_object* x_16; lean_object* x_17; uint8_t x_50; +x_16 = lean_ctor_get(x_9, 0); +lean_inc(x_16); +lean_dec(x_9); +x_50 = l_Lean_Expr_isConst(x_16); +if (x_50 == 0) +{ +uint8_t x_51; +x_51 = l_Lean_Expr_isFVar(x_16); +if (x_51 == 0) +{ +lean_object* x_52; lean_object* x_53; +lean_dec(x_16); +lean_dec(x_1); +x_52 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__7; +x_53 = l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__3(x_52, x_3, x_4, x_5, x_6, x_7, x_8); +return x_53; +} +else +{ +lean_object* x_54; +x_54 = lean_box(0); +x_17 = x_54; +goto block_49; +} +} +else +{ +lean_object* x_55; +x_55 = lean_box(0); +x_17 = x_55; +goto block_49; +} +block_49: +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +lean_dec(x_17); +lean_inc(x_16); +x_18 = l_Lean_Expr_toHeadIndex(x_16); +x_19 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol(x_18, x_3, x_4, x_5, x_6, x_7, x_8); +x_20 = lean_ctor_get(x_19, 1); +lean_inc(x_20); +lean_dec(x_19); +x_21 = lean_unsigned_to_nat(0u); +x_22 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_21); +x_23 = l_Lean_Meta_Grind_ppPattern___closed__5; +lean_inc(x_22); +x_24 = lean_mk_array(x_22, x_23); +x_25 = lean_unsigned_to_nat(1u); +x_26 = lean_nat_sub(x_22, x_25); +lean_dec(x_22); +lean_inc(x_1); +x_27 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_24, x_26); +x_28 = lean_array_get_size(x_27); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_28); +lean_inc(x_16); +x_29 = l_Lean_Meta_Grind_NormalizePattern_getPatternSupportMask(x_16, x_28, x_4, x_5, x_6, x_7, x_20); +if (lean_obj_tag(x_29) == 0) +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_30 = lean_ctor_get(x_29, 0); +lean_inc(x_30); +x_31 = lean_ctor_get(x_29, 1); +lean_inc(x_31); +lean_dec(x_29); +lean_inc(x_28); +x_32 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_32, 0, x_21); +lean_ctor_set(x_32, 1, x_28); +lean_ctor_set(x_32, 2, x_25); +x_33 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__2(x_1, x_28, x_30, x_32, x_32, x_27, x_21, lean_box(0), lean_box(0), x_3, x_4, x_5, x_6, x_7, x_31); +lean_dec(x_32); +lean_dec(x_30); +lean_dec(x_28); +lean_dec(x_1); +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; lean_object* x_36; +x_35 = lean_ctor_get(x_33, 0); +x_36 = l_Lean_mkAppN(x_16, x_35); +lean_dec(x_35); +lean_ctor_set(x_33, 0, x_36); +return x_33; +} +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); +x_39 = l_Lean_mkAppN(x_16, x_37); +lean_dec(x_37); +x_40 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_40, 0, x_39); +lean_ctor_set(x_40, 1, x_38); +return x_40; +} +} +else +{ +uint8_t x_41; +lean_dec(x_16); +x_41 = !lean_is_exclusive(x_33); +if (x_41 == 0) +{ +return x_33; +} +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; +} +} +} +else +{ +uint8_t x_45; +lean_dec(x_28); +lean_dec(x_27); +lean_dec(x_16); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_45 = !lean_is_exclusive(x_29); +if (x_45 == 0) +{ +return x_29; +} +else +{ +lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_46 = lean_ctor_get(x_29, 0); +x_47 = lean_ctor_get(x_29, 1); +lean_inc(x_47); +lean_inc(x_46); +lean_dec(x_29); +x_48 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_48, 0, x_46); +lean_ctor_set(x_48, 1, x_47); +return x_48; +} +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_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) { +_start: +{ +lean_object* x_9; +lean_inc(x_1); +x_9 = l_Lean_Meta_Grind_isOffsetPattern_x3f(x_1); +if (lean_obj_tag(x_9) == 0) +{ +lean_object* x_10; lean_object* x_11; +x_10 = lean_box(0); +x_11 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1(x_1, x_10, x_3, x_4, x_5, x_6, x_7, x_8); +return x_11; +} +else +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; +lean_dec(x_1); +x_12 = lean_ctor_get(x_9, 0); +lean_inc(x_12); +lean_dec(x_9); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_12, 1); +lean_inc(x_14); +lean_dec(x_12); +x_15 = 0; +x_16 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go_goArg(x_13, x_15, x_3, x_4, x_5, x_6, x_7, 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; uint8_t x_20; +x_18 = lean_ctor_get(x_16, 0); +x_19 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare; +x_20 = lean_expr_eqv(x_18, x_19); +if (x_20 == 0) +{ +lean_object* x_21; +x_21 = l_Lean_Meta_Grind_mkOffsetPattern(x_18, x_14); +lean_ctor_set(x_16, 0, x_21); +return x_16; +} +else +{ +lean_dec(x_18); +lean_dec(x_14); +lean_ctor_set(x_16, 0, x_19); +return x_16; +} +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; +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 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare; +x_25 = lean_expr_eqv(x_22, x_24); +if (x_25 == 0) +{ +lean_object* x_26; lean_object* x_27; +x_26 = l_Lean_Meta_Grind_mkOffsetPattern(x_22, x_14); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_23); +return x_27; +} +else +{ +lean_object* x_28; +lean_dec(x_22); +lean_dec(x_14); +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_24); +lean_ctor_set(x_28, 1, x_23); +return x_28; +} +} +} +else +{ +uint8_t x_29; +lean_dec(x_14); +x_29 = !lean_is_exclusive(x_16); +if (x_29 == 0) +{ +return x_16; +} +else +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_30 = lean_ctor_get(x_16, 0); +x_31 = lean_ctor_get(x_16, 1); +lean_inc(x_31); +lean_inc(x_30); +lean_dec(x_16); +x_32 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_32, 0, x_30); +lean_ctor_set(x_32, 1, x_31); +return x_32; +} +} +} +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("invalid pattern, it does not have pattern variables", 51, 51); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +if (x_2 == 0) +{ +lean_object* x_9; lean_object* x_10; +x_9 = lean_box(0); +x_10 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__2(x_1, x_9, x_3, x_4, x_5, x_6, x_7, x_8); +return x_10; +} +else +{ +uint8_t x_11; +x_11 = l_Lean_Expr_hasLooseBVars(x_1); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; uint8_t x_14; +lean_dec(x_1); +x_12 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___closed__2; +x_13 = l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__4(x_12, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) +{ +return x_13; +} +else +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_ctor_get(x_13, 0); +x_16 = lean_ctor_get(x_13, 1); +lean_inc(x_16); +lean_inc(x_15); +lean_dec(x_13); +x_17 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_17, 0, x_15); +lean_ctor_set(x_17, 1, x_16); +return x_17; +} +} +else +{ +lean_object* x_18; lean_object* x_19; +x_18 = lean_box(0); +x_19 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__2(x_1, x_18, x_3, x_4, x_5, x_6, x_7, x_8); +return x_19; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go_goArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +uint8_t x_9; lean_object* x_10; +x_9 = lean_unbox(x_2); +lean_dec(x_2); +x_10 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go_goArg(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_throwError___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_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, lean_object* x_7) { +_start: +{ +lean_object* x_8; +x_8 = l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +_start: +{ +lean_object* x_16; +x_16 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_16; +} +} +LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_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) { +_start: +{ +lean_object* x_8; +x_8 = l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_8; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_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) { +_start: +{ +lean_object* x_9; +x_9 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_2); +return x_9; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_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) { +_start: +{ +lean_object* x_9; +x_9 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_2); +return x_9; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_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) { +_start: +{ +uint8_t x_9; lean_object* x_10; +x_9 = lean_unbox(x_2); +lean_dec(x_2); +x_10 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go(x_1, x_9, x_3, x_4, x_5, x_6, x_7, x_8); +return x_10; +} +} +LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_Meta_Grind_NormalizePattern_main___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_9; lean_object* x_10; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_9 = l_List_reverse___rarg(x_2); +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_1); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; uint8_t x_14; lean_object* x_15; +x_12 = lean_ctor_get(x_1, 0); +x_13 = lean_ctor_get(x_1, 1); +x_14 = 0; +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_15 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go(x_12, 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; +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_15, 1); +lean_inc(x_17); +lean_dec(x_15); +lean_ctor_set(x_1, 1, x_2); +lean_ctor_set(x_1, 0, x_16); +{ +lean_object* _tmp_0 = x_13; +lean_object* _tmp_1 = x_1; +lean_object* _tmp_7 = x_17; +x_1 = _tmp_0; +x_2 = _tmp_1; +x_8 = _tmp_7; +} +goto _start; +} +else +{ +uint8_t x_19; +lean_free_object(x_1); +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_19 = !lean_is_exclusive(x_15); +if (x_19 == 0) +{ +return x_15; +} +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_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; +} +} +} +else +{ +lean_object* x_23; lean_object* x_24; uint8_t x_25; lean_object* x_26; +x_23 = lean_ctor_get(x_1, 0); +x_24 = lean_ctor_get(x_1, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_1); +x_25 = 0; +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_26 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go(x_23, x_25, x_3, x_4, x_5, x_6, x_7, x_8); +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 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_29, 0, x_27); +lean_ctor_set(x_29, 1, x_2); +x_1 = x_24; +x_2 = x_29; +x_8 = x_28; +goto _start; +} +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; +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_2); +x_31 = lean_ctor_get(x_26, 0); +lean_inc(x_31); +x_32 = lean_ctor_get(x_26, 1); +lean_inc(x_32); +if (lean_is_exclusive(x_26)) { + lean_ctor_release(x_26, 0); + lean_ctor_release(x_26, 1); + x_33 = x_26; +} else { + lean_dec_ref(x_26); + x_33 = lean_box(0); +} +if (lean_is_scalar(x_33)) { + x_34 = lean_alloc_ctor(1, 2, 0); +} else { + x_34 = x_33; +} +lean_ctor_set(x_34, 0, x_31); +lean_ctor_set(x_34, 1, x_32); +return x_34; +} +} +} +} +} +static lean_object* _init_l_Lean_Meta_Grind_NormalizePattern_main___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(0); +x_2 = lean_array_mk(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Grind_NormalizePattern_main___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_unsigned_to_nat(10u); +x_2 = lean_unsigned_to_nat(1u); +x_3 = l_Nat_nextPowerOfTwo_go(x_1, x_2, lean_box(0)); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind_NormalizePattern_main___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_Grind_NormalizePattern_main___closed__2; +x_3 = lean_mk_array(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind_NormalizePattern_main___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_unsigned_to_nat(0u); +x_2 = l_Lean_Meta_Grind_NormalizePattern_main___closed__3; +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind_NormalizePattern_main___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_NormalizePattern_main___closed__1; +x_2 = l_Lean_Meta_Grind_NormalizePattern_main___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_2); +lean_ctor_set(x_3, 2, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_NormalizePattern_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) { +_start: +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; +x_7 = lean_box(0); +x_8 = l_Lean_Meta_Grind_NormalizePattern_main___closed__5; +x_9 = lean_st_mk_ref(x_8, 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; +x_11 = lean_ctor_get(x_9, 0); +x_12 = lean_ctor_get(x_9, 1); +lean_inc(x_11); +x_13 = l_List_mapM_loop___at_Lean_Meta_Grind_NormalizePattern_main___spec__1(x_1, x_7, x_11, x_2, x_3, x_4, x_5, 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_st_ref_get(x_11, x_15); +lean_dec(x_11); +x_17 = !lean_is_exclusive(x_16); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_18 = lean_ctor_get(x_16, 0); +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_array_to_list(x_19); +x_21 = lean_ctor_get(x_18, 2); +lean_inc(x_21); +lean_dec(x_18); +lean_ctor_set(x_9, 1, x_21); +lean_ctor_set(x_9, 0, x_20); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_14); +lean_ctor_set(x_22, 1, x_9); +lean_ctor_set(x_16, 0, x_22); +return x_16; +} +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; +x_23 = lean_ctor_get(x_16, 0); +x_24 = lean_ctor_get(x_16, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_16); +x_25 = lean_ctor_get(x_23, 0); +lean_inc(x_25); +x_26 = lean_array_to_list(x_25); +x_27 = lean_ctor_get(x_23, 2); +lean_inc(x_27); +lean_dec(x_23); +lean_ctor_set(x_9, 1, x_27); +lean_ctor_set(x_9, 0, x_26); +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_14); +lean_ctor_set(x_28, 1, x_9); +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 +{ +uint8_t x_30; +lean_free_object(x_9); +lean_dec(x_11); +x_30 = !lean_is_exclusive(x_13); +if (x_30 == 0) +{ +return x_13; +} +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_13, 0); +x_32 = lean_ctor_get(x_13, 1); +lean_inc(x_32); +lean_inc(x_31); +lean_dec(x_13); +x_33 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_33, 0, x_31); +lean_ctor_set(x_33, 1, x_32); +return x_33; +} +} +} +else +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_9, 0); +x_35 = lean_ctor_get(x_9, 1); +lean_inc(x_35); +lean_inc(x_34); +lean_dec(x_9); +lean_inc(x_34); +x_36 = l_List_mapM_loop___at_Lean_Meta_Grind_NormalizePattern_main___spec__1(x_1, x_7, x_34, x_2, x_3, x_4, x_5, x_35); +if (lean_obj_tag(x_36) == 0) +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_37 = lean_ctor_get(x_36, 0); +lean_inc(x_37); +x_38 = lean_ctor_get(x_36, 1); +lean_inc(x_38); +lean_dec(x_36); +x_39 = lean_st_ref_get(x_34, x_38); +lean_dec(x_34); +x_40 = lean_ctor_get(x_39, 0); +lean_inc(x_40); +x_41 = lean_ctor_get(x_39, 1); +lean_inc(x_41); +if (lean_is_exclusive(x_39)) { + lean_ctor_release(x_39, 0); + lean_ctor_release(x_39, 1); + x_42 = x_39; +} else { + lean_dec_ref(x_39); + x_42 = lean_box(0); +} +x_43 = lean_ctor_get(x_40, 0); +lean_inc(x_43); +x_44 = lean_array_to_list(x_43); +x_45 = lean_ctor_get(x_40, 2); +lean_inc(x_45); +lean_dec(x_40); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_44); +lean_ctor_set(x_46, 1, x_45); +x_47 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_47, 0, x_37); +lean_ctor_set(x_47, 1, x_46); +if (lean_is_scalar(x_42)) { + x_48 = lean_alloc_ctor(0, 2, 0); +} else { + x_48 = x_42; +} +lean_ctor_set(x_48, 0, x_47); +lean_ctor_set(x_48, 1, x_41); +return x_48; +} +else +{ +lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; +lean_dec(x_34); +x_49 = lean_ctor_get(x_36, 0); +lean_inc(x_49); +x_50 = lean_ctor_get(x_36, 1); +lean_inc(x_50); +if (lean_is_exclusive(x_36)) { + lean_ctor_release(x_36, 0); + lean_ctor_release(x_36, 1); + x_51 = x_36; +} else { + lean_dec_ref(x_36); + x_51 = lean_box(0); +} +if (lean_is_scalar(x_51)) { + x_52 = lean_alloc_ctor(1, 2, 0); +} else { + x_52 = x_51; +} +lean_ctor_set(x_52, 0, x_49); +lean_ctor_set(x_52, 1, x_50); +return x_52; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_NormalizePattern_normalizePattern(lean_object* x_1, 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 = 0; +x_9 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go(x_1, x_8, x_2, x_3, x_4, x_5, x_6, x_7); +return x_9; +} +} +LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___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) { +_start: +{ +uint8_t x_7; +x_7 = lean_usize_dec_eq(x_5, x_6); +if (x_7 == 0) +{ +lean_object* x_8; lean_object* x_9; +x_8 = lean_array_uget(x_4, x_5); +x_9 = l_Lean_RBNode_findCore___at_Lean_Meta_Closure_process___spec__1(x_1, x_8); +if (lean_obj_tag(x_9) == 0) +{ +size_t x_10; size_t x_11; +lean_dec(x_8); +x_10 = 1; +x_11 = lean_usize_add(x_5, x_10); +x_5 = x_11; +goto _start; +} +else +{ +lean_object* x_13; +lean_dec(x_9); +x_13 = l_Lean_RBNode_findCore___at_Lean_Meta_Closure_process___spec__1(x_2, x_8); +lean_dec(x_8); +if (lean_obj_tag(x_13) == 0) +{ +uint8_t x_14; +x_14 = 1; +return x_14; +} +else +{ +size_t x_15; size_t x_16; +lean_dec(x_13); +x_15 = 1; +x_16 = lean_usize_add(x_5, x_15); +x_5 = x_16; +goto _start; +} +} +} +else +{ +uint8_t x_18; +x_18 = 0; +return x_18; +} +} +} +LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___spec__1___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___spec__2(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; +x_7 = lean_array_uget(x_3, x_4); +x_8 = l_Lean_RBNode_findCore___at_Lean_Meta_Closure_process___spec__1(x_1, x_7); +if (lean_obj_tag(x_8) == 0) +{ +size_t x_9; size_t x_10; +lean_dec(x_7); +x_9 = 1; +x_10 = lean_usize_add(x_4, x_9); +x_4 = x_10; +goto _start; +} +else +{ +lean_object* x_12; +lean_dec(x_8); +x_12 = l_Lean_RBNode_findCore___at_Lean_Meta_Closure_process___spec__1(x_2, x_7); +lean_dec(x_7); +if (lean_obj_tag(x_12) == 0) +{ +uint8_t x_13; +x_13 = 1; +return x_13; +} +else +{ +size_t x_14; size_t x_15; +lean_dec(x_12); +x_14 = 1; +x_15 = lean_usize_add(x_4, x_14); +x_4 = x_15; +goto _start; +} +} +} +else +{ +uint8_t x_17; +x_17 = 0; +return x_17; +} +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; +x_3 = l_Lean_Meta_Grind_NormalizePattern_main___closed__1; +x_4 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_4, 0, x_2); +lean_ctor_set(x_4, 1, x_1); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; +x_4 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___closed__1; +x_5 = l_Lean_CollectFVars_main(x_3, x_4); +x_6 = lean_ctor_get(x_5, 2); +lean_inc(x_6); +lean_dec(x_5); +x_7 = lean_array_get_size(x_6); +x_8 = lean_unsigned_to_nat(0u); +x_9 = lean_nat_dec_lt(x_8, x_7); +if (x_9 == 0) +{ +uint8_t x_10; +lean_dec(x_7); +lean_dec(x_6); +x_10 = 1; +return x_10; +} +else +{ +size_t x_11; size_t x_12; uint8_t x_13; +x_11 = 0; +x_12 = lean_usize_of_nat(x_7); +lean_dec(x_7); +x_13 = l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___spec__1___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___spec__2(x_1, x_2, x_6, x_11, x_12); +lean_dec(x_6); +if (x_13 == 0) +{ +uint8_t x_14; +x_14 = 1; +return x_14; +} +else +{ +uint8_t x_15; +x_15 = 0; +return x_15; +} +} +} +} +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___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_7; size_t x_8; uint8_t x_9; lean_object* x_10; +x_7 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_8 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_9 = l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___spec__1(x_1, x_2, x_3, x_4, x_7, x_8); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_10 = lean_box(x_9); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___spec__1___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +size_t x_6; size_t x_7; 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___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___spec__1___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___spec__2(x_1, x_2, x_3, x_6, x_7); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_9 = lean_box(x_8); +return x_9; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___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___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars(x_1, x_2, x_3); +lean_dec(x_2); +lean_dec(x_1); +x_5 = lean_box(x_4); +return x_5; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__1___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; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__1___closed__1; +x_2 = lean_box(0); +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___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, size_t x_7, size_t x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +uint8_t x_15; +x_15 = lean_usize_dec_lt(x_8, x_7); +if (x_15 == 0) +{ +lean_object* x_16; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_5); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_9); +lean_ctor_set(x_16, 1, x_14); +return x_16; +} +else +{ +lean_object* x_17; lean_object* x_18; +lean_dec(x_9); +x_17 = lean_array_uget(x_6, x_8); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +x_18 = lean_infer_type(x_17, x_10, x_11, x_12, x_13, x_14); +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_ctor_get(x_18, 1); +x_22 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars(x_1, x_2, x_20); +if (x_22 == 0) +{ +lean_object* x_23; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_5); +x_23 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__1___closed__2; +lean_ctor_set(x_18, 0, x_23); +return x_18; +} +else +{ +size_t x_24; size_t x_25; +lean_free_object(x_18); +x_24 = 1; +x_25 = lean_usize_add(x_8, x_24); +lean_inc(x_5); +{ +size_t _tmp_7 = x_25; +lean_object* _tmp_8 = x_5; +lean_object* _tmp_13 = x_21; +x_8 = _tmp_7; +x_9 = _tmp_8; +x_14 = _tmp_13; +} +goto _start; +} +} +else +{ +lean_object* x_27; lean_object* x_28; uint8_t x_29; +x_27 = lean_ctor_get(x_18, 0); +x_28 = lean_ctor_get(x_18, 1); +lean_inc(x_28); +lean_inc(x_27); +lean_dec(x_18); +x_29 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars(x_1, x_2, x_27); +if (x_29 == 0) +{ +lean_object* x_30; lean_object* x_31; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_5); +x_30 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__1___closed__2; +x_31 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_31, 0, x_30); +lean_ctor_set(x_31, 1, x_28); +return x_31; +} +else +{ +size_t x_32; size_t x_33; +x_32 = 1; +x_33 = lean_usize_add(x_8, x_32); +lean_inc(x_5); +{ +size_t _tmp_7 = x_33; +lean_object* _tmp_8 = x_5; +lean_object* _tmp_13 = x_28; +x_8 = _tmp_7; +x_9 = _tmp_8; +x_14 = _tmp_13; +} +goto _start; +} +} +} +else +{ +uint8_t x_35; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_5); +x_35 = !lean_is_exclusive(x_18); +if (x_35 == 0) +{ +return x_18; +} +else +{ +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_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; +} +} +} +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("semiOutParam", 12, 12); +return x_1; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("outParam", 8, 8); +return x_1; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__3; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___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, size_t x_7, size_t x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +uint8_t x_15; +x_15 = lean_usize_dec_lt(x_8, x_7); +if (x_15 == 0) +{ +lean_object* x_16; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_4); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_9); +lean_ctor_set(x_16, 1, x_14); +return x_16; +} +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_27; +x_17 = lean_array_uget(x_6, x_8); +x_27 = !lean_is_exclusive(x_9); +if (x_27 == 0) +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; +x_28 = lean_ctor_get(x_9, 1); +x_29 = lean_ctor_get(x_9, 0); +lean_dec(x_29); +x_30 = lean_ctor_get(x_28, 0); +lean_inc(x_30); +x_31 = lean_ctor_get(x_28, 1); +lean_inc(x_31); +x_32 = lean_ctor_get(x_28, 2); +lean_inc(x_32); +x_33 = lean_nat_dec_lt(x_31, x_32); +if (x_33 == 0) +{ +lean_object* x_34; +lean_dec(x_32); +lean_dec(x_31); +lean_dec(x_30); +lean_dec(x_17); +lean_inc(x_4); +lean_ctor_set(x_9, 0, x_4); +x_34 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_34, 0, x_9); +x_18 = x_34; +x_19 = x_14; +goto block_26; +} +else +{ +uint8_t x_35; +x_35 = !lean_is_exclusive(x_28); +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; +x_36 = lean_ctor_get(x_28, 2); +lean_dec(x_36); +x_37 = lean_ctor_get(x_28, 1); +lean_dec(x_37); +x_38 = lean_ctor_get(x_28, 0); +lean_dec(x_38); +x_39 = lean_array_fget(x_30, x_31); +x_40 = lean_unsigned_to_nat(1u); +x_41 = lean_nat_add(x_31, x_40); +lean_dec(x_31); +lean_ctor_set(x_28, 1, x_41); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +x_42 = lean_infer_type(x_17, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(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_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_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__2; +x_46 = l_Lean_Expr_isAppOf(x_43, x_45); +if (x_46 == 0) +{ +lean_object* x_47; uint8_t x_48; +x_47 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__4; +x_48 = l_Lean_Expr_isAppOf(x_43, x_47); +lean_dec(x_43); +if (x_48 == 0) +{ +uint8_t x_49; +x_49 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars(x_1, x_2, x_39); +if (x_49 == 0) +{ +lean_object* x_50; lean_object* x_51; +x_50 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__1___closed__1; +lean_ctor_set(x_9, 0, x_50); +x_51 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_51, 0, x_9); +x_18 = x_51; +x_19 = x_44; +goto block_26; +} +else +{ +lean_object* x_52; +lean_inc(x_4); +lean_ctor_set(x_9, 0, x_4); +x_52 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_52, 0, x_9); +x_18 = x_52; +x_19 = x_44; +goto block_26; +} +} +else +{ +lean_object* x_53; +lean_dec(x_39); +lean_inc(x_4); +lean_ctor_set(x_9, 0, x_4); +x_53 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_53, 0, x_9); +x_18 = x_53; +x_19 = x_44; +goto block_26; +} +} +else +{ +lean_object* x_54; +lean_dec(x_43); +lean_dec(x_39); +lean_inc(x_4); +lean_ctor_set(x_9, 0, x_4); +x_54 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_54, 0, x_9); +x_18 = x_54; +x_19 = x_44; +goto block_26; +} +} +else +{ +uint8_t x_55; +lean_dec(x_28); +lean_dec(x_39); +lean_free_object(x_9); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_4); +x_55 = !lean_is_exclusive(x_42); +if (x_55 == 0) +{ +return x_42; +} +else +{ +lean_object* x_56; lean_object* x_57; lean_object* x_58; +x_56 = lean_ctor_get(x_42, 0); +x_57 = lean_ctor_get(x_42, 1); +lean_inc(x_57); +lean_inc(x_56); +lean_dec(x_42); +x_58 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_58, 0, x_56); +lean_ctor_set(x_58, 1, x_57); +return x_58; +} +} +} +else +{ +lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; +lean_dec(x_28); +x_59 = lean_array_fget(x_30, x_31); +x_60 = lean_unsigned_to_nat(1u); +x_61 = lean_nat_add(x_31, x_60); +lean_dec(x_31); +x_62 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_62, 0, x_30); +lean_ctor_set(x_62, 1, x_61); +lean_ctor_set(x_62, 2, x_32); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +x_63 = lean_infer_type(x_17, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_63) == 0) +{ +lean_object* x_64; lean_object* x_65; lean_object* x_66; uint8_t 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 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__2; +x_67 = l_Lean_Expr_isAppOf(x_64, x_66); +if (x_67 == 0) +{ +lean_object* x_68; uint8_t x_69; +x_68 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__4; +x_69 = l_Lean_Expr_isAppOf(x_64, x_68); +lean_dec(x_64); +if (x_69 == 0) +{ +uint8_t x_70; +x_70 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars(x_1, x_2, x_59); +if (x_70 == 0) +{ +lean_object* x_71; lean_object* x_72; +x_71 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__1___closed__1; +lean_ctor_set(x_9, 1, x_62); +lean_ctor_set(x_9, 0, x_71); +x_72 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_72, 0, x_9); +x_18 = x_72; +x_19 = x_65; +goto block_26; +} +else +{ +lean_object* x_73; +lean_inc(x_4); +lean_ctor_set(x_9, 1, x_62); +lean_ctor_set(x_9, 0, x_4); +x_73 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_73, 0, x_9); +x_18 = x_73; +x_19 = x_65; +goto block_26; +} +} +else +{ +lean_object* x_74; +lean_dec(x_59); +lean_inc(x_4); +lean_ctor_set(x_9, 1, x_62); +lean_ctor_set(x_9, 0, x_4); +x_74 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_74, 0, x_9); +x_18 = x_74; +x_19 = x_65; +goto block_26; +} +} +else +{ +lean_object* x_75; +lean_dec(x_64); +lean_dec(x_59); +lean_inc(x_4); +lean_ctor_set(x_9, 1, x_62); +lean_ctor_set(x_9, 0, x_4); +x_75 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_75, 0, x_9); +x_18 = x_75; +x_19 = x_65; +goto block_26; +} +} +else +{ +lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; +lean_dec(x_62); +lean_dec(x_59); +lean_free_object(x_9); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_4); +x_76 = lean_ctor_get(x_63, 0); +lean_inc(x_76); +x_77 = lean_ctor_get(x_63, 1); +lean_inc(x_77); +if (lean_is_exclusive(x_63)) { + lean_ctor_release(x_63, 0); + lean_ctor_release(x_63, 1); + x_78 = x_63; +} else { + lean_dec_ref(x_63); + 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 +{ +lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; uint8_t x_84; +x_80 = lean_ctor_get(x_9, 1); +lean_inc(x_80); +lean_dec(x_9); +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_80, 2); +lean_inc(x_83); +x_84 = lean_nat_dec_lt(x_82, x_83); +if (x_84 == 0) +{ +lean_object* x_85; lean_object* x_86; +lean_dec(x_83); +lean_dec(x_82); +lean_dec(x_81); +lean_dec(x_17); +lean_inc(x_4); +x_85 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_85, 0, x_4); +lean_ctor_set(x_85, 1, x_80); +x_86 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_86, 0, x_85); +x_18 = x_86; +x_19 = x_14; +goto block_26; +} +else +{ +lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; +if (lean_is_exclusive(x_80)) { + lean_ctor_release(x_80, 0); + lean_ctor_release(x_80, 1); + lean_ctor_release(x_80, 2); + x_87 = x_80; +} else { + lean_dec_ref(x_80); + x_87 = lean_box(0); +} +x_88 = lean_array_fget(x_81, x_82); +x_89 = lean_unsigned_to_nat(1u); +x_90 = lean_nat_add(x_82, x_89); +lean_dec(x_82); +if (lean_is_scalar(x_87)) { + x_91 = lean_alloc_ctor(0, 3, 0); +} else { + x_91 = x_87; +} +lean_ctor_set(x_91, 0, x_81); +lean_ctor_set(x_91, 1, x_90); +lean_ctor_set(x_91, 2, x_83); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +x_92 = lean_infer_type(x_17, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_92) == 0) +{ +lean_object* x_93; lean_object* x_94; lean_object* x_95; uint8_t x_96; +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_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__2; +x_96 = l_Lean_Expr_isAppOf(x_93, x_95); +if (x_96 == 0) +{ +lean_object* x_97; uint8_t x_98; +x_97 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__4; +x_98 = l_Lean_Expr_isAppOf(x_93, x_97); +lean_dec(x_93); +if (x_98 == 0) +{ +uint8_t x_99; +x_99 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars(x_1, x_2, x_88); +if (x_99 == 0) +{ +lean_object* x_100; lean_object* x_101; lean_object* x_102; +x_100 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__1___closed__1; +x_101 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_101, 0, x_100); +lean_ctor_set(x_101, 1, x_91); +x_102 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_102, 0, x_101); +x_18 = x_102; +x_19 = x_94; +goto block_26; +} +else +{ +lean_object* x_103; lean_object* x_104; +lean_inc(x_4); +x_103 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_103, 0, x_4); +lean_ctor_set(x_103, 1, x_91); +x_104 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_104, 0, x_103); +x_18 = x_104; +x_19 = x_94; +goto block_26; +} +} +else +{ +lean_object* x_105; lean_object* x_106; +lean_dec(x_88); +lean_inc(x_4); +x_105 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_105, 0, x_4); +lean_ctor_set(x_105, 1, x_91); +x_106 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_106, 0, x_105); +x_18 = x_106; +x_19 = x_94; +goto block_26; +} +} +else +{ +lean_object* x_107; lean_object* x_108; +lean_dec(x_93); +lean_dec(x_88); +lean_inc(x_4); +x_107 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_107, 0, x_4); +lean_ctor_set(x_107, 1, x_91); +x_108 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_108, 0, x_107); +x_18 = x_108; +x_19 = x_94; +goto block_26; +} +} +else +{ +lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; +lean_dec(x_91); +lean_dec(x_88); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_4); +x_109 = lean_ctor_get(x_92, 0); +lean_inc(x_109); +x_110 = lean_ctor_get(x_92, 1); +lean_inc(x_110); +if (lean_is_exclusive(x_92)) { + lean_ctor_release(x_92, 0); + lean_ctor_release(x_92, 1); + x_111 = x_92; +} else { + lean_dec_ref(x_92); + x_111 = lean_box(0); +} +if (lean_is_scalar(x_111)) { + x_112 = lean_alloc_ctor(1, 2, 0); +} else { + x_112 = x_111; +} +lean_ctor_set(x_112, 0, x_109); +lean_ctor_set(x_112, 1, x_110); +return x_112; +} +} +} +block_26: +{ +if (lean_obj_tag(x_18) == 0) +{ +lean_object* x_20; lean_object* x_21; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_4); +x_20 = lean_ctor_get(x_18, 0); +lean_inc(x_20); +lean_dec(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; +} +else +{ +lean_object* x_22; size_t x_23; size_t x_24; +x_22 = lean_ctor_get(x_18, 0); +lean_inc(x_22); +lean_dec(x_18); +x_23 = 1; +x_24 = lean_usize_add(x_8, x_23); +x_8 = x_24; +x_9 = x_22; +x_14 = x_19; +goto _start; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___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) { +_start: +{ +uint8_t x_7; lean_object* x_8; lean_object* x_9; +x_7 = 1; +x_8 = lean_box(x_7); +x_9 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_9, 0, x_8); +lean_ctor_set(x_9, 1, x_6); +return x_9; +} +} +static lean_object* _init_l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__2___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__1___boxed), 6, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, size_t x_6, lean_object* x_7, 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; size_t x_18; lean_object* x_19; +x_14 = lean_array_get_size(x_1); +x_15 = lean_unsigned_to_nat(0u); +x_16 = l_Array_toSubarray___rarg(x_1, x_15, x_14); +lean_inc(x_2); +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_2); +lean_ctor_set(x_17, 1, x_16); +x_18 = lean_array_size(x_7); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +x_19 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2(x_3, x_4, x_5, x_2, x_7, x_7, x_18, x_6, x_17, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_19) == 0) +{ +lean_object* x_20; lean_object* x_21; +x_20 = lean_ctor_get(x_19, 0); +lean_inc(x_20); +x_21 = lean_ctor_get(x_20, 0); +lean_inc(x_21); +lean_dec(x_20); +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_19, 1); +lean_inc(x_22); +lean_dec(x_19); +x_23 = l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__2___closed__1; +x_24 = lean_box(0); +x_25 = lean_apply_6(x_23, x_24, x_9, x_10, x_11, x_12, x_22); +return x_25; +} +else +{ +uint8_t x_26; +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +x_26 = !lean_is_exclusive(x_19); +if (x_26 == 0) +{ +lean_object* x_27; lean_object* x_28; +x_27 = lean_ctor_get(x_19, 0); +lean_dec(x_27); +x_28 = lean_ctor_get(x_21, 0); +lean_inc(x_28); +lean_dec(x_21); +lean_ctor_set(x_19, 0, x_28); +return x_19; +} +else +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_29 = lean_ctor_get(x_19, 1); +lean_inc(x_29); +lean_dec(x_19); +x_30 = lean_ctor_get(x_21, 0); +lean_inc(x_30); +lean_dec(x_21); +x_31 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_31, 0, x_30); +lean_ctor_set(x_31, 1, x_29); +return x_31; +} +} +} +else +{ +uint8_t x_32; +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +x_32 = !lean_is_exclusive(x_19); +if (x_32 == 0) +{ +return x_19; +} +else +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_33 = lean_ctor_get(x_19, 0); +x_34 = lean_ctor_get(x_19, 1); +lean_inc(x_34); +lean_inc(x_33); +lean_dec(x_19); +x_35 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_35, 0, x_33); +lean_ctor_set(x_35, 1, x_34); +return x_35; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, size_t x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +lean_object* x_15; +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +x_15 = lean_infer_type(x_1, x_10, x_11, x_12, x_13, x_14); +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; uint8_t x_23; lean_object* x_24; +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_unsigned_to_nat(0u); +x_19 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_2, x_18); +x_20 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_20, 0, x_19); +x_21 = lean_box_usize(x_8); +x_22 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__2___boxed), 13, 6); +lean_closure_set(x_22, 0, x_3); +lean_closure_set(x_22, 1, x_4); +lean_closure_set(x_22, 2, x_5); +lean_closure_set(x_22, 3, x_6); +lean_closure_set(x_22, 4, x_7); +lean_closure_set(x_22, 5, x_21); +x_23 = 0; +x_24 = l_Lean_Meta_forallBoundedTelescope___at_Lean_Meta_arrowDomainsN___spec__6___rarg(x_16, x_20, x_22, x_23, x_10, x_11, x_12, x_13, x_17); +return x_24; +} +else +{ +uint8_t x_25; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_25 = !lean_is_exclusive(x_15); +if (x_25 == 0) +{ +return x_15; +} +else +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_26 = lean_ctor_get(x_15, 0); +x_27 = lean_ctor_get(x_15, 1); +lean_inc(x_27); +lean_inc(x_26); +lean_dec(x_15); +x_28 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_28, 0, x_26); +lean_ctor_set(x_28, 1, x_27); +return x_28; +} +} +} +} +static lean_object* _init_l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = lean_box(0); +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +if (lean_obj_tag(x_5) == 5) +{ +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_5, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_5, 1); +lean_inc(x_14); +lean_dec(x_5); +x_15 = lean_array_set(x_6, x_7, x_14); +x_16 = lean_unsigned_to_nat(1u); +x_17 = lean_nat_sub(x_7, x_16); +lean_dec(x_7); +x_5 = x_13; +x_6 = x_15; +x_7 = x_17; +goto _start; +} +else +{ +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_dec(x_7); +x_19 = lean_box(0); +x_20 = lean_box(0); +x_21 = lean_array_size(x_3); +x_22 = 0; +x_23 = l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___closed__1; +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +x_24 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__1(x_1, x_2, x_3, x_19, x_23, x_3, x_21, x_22, x_23, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_24) == 0) +{ +lean_object* x_25; lean_object* x_26; +x_25 = lean_ctor_get(x_24, 0); +lean_inc(x_25); +x_26 = lean_ctor_get(x_25, 0); +lean_inc(x_26); +lean_dec(x_25); +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_24, 1); +lean_inc(x_27); +lean_dec(x_24); +x_28 = lean_box(0); +x_29 = l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__3(x_5, x_4, x_6, x_20, x_1, x_2, x_19, x_22, x_28, x_8, x_9, x_10, x_11, x_27); +return x_29; +} +else +{ +uint8_t x_30; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_2); +lean_dec(x_1); +x_30 = !lean_is_exclusive(x_24); +if (x_30 == 0) +{ +lean_object* x_31; lean_object* x_32; +x_31 = lean_ctor_get(x_24, 0); +lean_dec(x_31); +x_32 = lean_ctor_get(x_26, 0); +lean_inc(x_32); +lean_dec(x_26); +lean_ctor_set(x_24, 0, x_32); +return x_24; +} +else +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_33 = lean_ctor_get(x_24, 1); +lean_inc(x_33); +lean_dec(x_24); +x_34 = lean_ctor_get(x_26, 0); +lean_inc(x_34); +lean_dec(x_26); +x_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_34); +lean_ctor_set(x_35, 1, x_33); +return x_35; +} +} +} +else +{ +uint8_t x_36; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_2); +lean_dec(x_1); +x_36 = !lean_is_exclusive(x_24); +if (x_36 == 0) +{ +return x_24; +} +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_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_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_10 = lean_unsigned_to_nat(0u); +x_11 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_4, x_10); +x_12 = l_Lean_Meta_Grind_ppPattern___closed__5; +lean_inc(x_11); +x_13 = lean_mk_array(x_11, x_12); +x_14 = lean_unsigned_to_nat(1u); +x_15 = lean_nat_sub(x_11, x_14); +lean_dec(x_11); +lean_inc(x_4); +x_16 = l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3(x_1, x_2, x_3, x_4, x_4, x_13, x_15, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_4); +return x_16; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; uint8_t x_10; lean_object* x_11; +x_9 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___lambda__1___boxed), 9, 2); +lean_closure_set(x_9, 0, x_1); +lean_closure_set(x_9, 1, x_2); +x_10 = 0; +x_11 = l_Lean_Meta_forallTelescopeReducing___at_Lean_Meta_getParamNames___spec__2___rarg(x_3, x_9, x_10, x_4, x_5, x_6, x_7, x_8); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +size_t x_15; size_t x_16; lean_object* x_17; +x_15 = lean_unbox_usize(x_7); +lean_dec(x_7); +x_16 = lean_unbox_usize(x_8); +lean_dec(x_8); +x_17 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_15, x_16, x_9, x_10, x_11, x_12, x_13, x_14); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_17; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +size_t x_15; size_t x_16; lean_object* x_17; +x_15 = lean_unbox_usize(x_7); +lean_dec(x_7); +x_16 = lean_unbox_usize(x_8); +lean_dec(x_8); +x_17 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_15, x_16, x_9, x_10, x_11, x_12, x_13, x_14); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_17; +} +} +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___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) { +_start: +{ +lean_object* x_7; +x_7 = l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +size_t x_14; lean_object* x_15; +x_14 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_15 = l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__2(x_1, x_2, x_3, x_4, x_5, x_14, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_15; +} +} +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +size_t x_15; lean_object* x_16; +x_15 = lean_unbox_usize(x_8); +lean_dec(x_8); +x_16 = l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_15, x_9, x_10, x_11, x_12, x_13, x_14); +lean_dec(x_9); +lean_dec(x_2); +return x_16; +} +} +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +x_13 = l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_4); +lean_dec(x_3); +return x_13; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___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_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___lambda__1(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_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_instInhabitedMetaM___boxed), 5, 1); +lean_closure_set(x_1, 0, lean_box(0)); +return x_1; +} +} +LEAN_EXPORT lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___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; lean_object* x_9; +x_7 = l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__1___closed__1; +x_8 = lean_panic_fn(x_7, x_1); +x_9 = lean_apply_5(x_8, x_2, x_3, x_4, x_5, x_6); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldrM___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__2(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +lean_inc(x_1); +return x_1; +} +else +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_3 = lean_ctor_get(x_2, 0); +x_4 = lean_ctor_get(x_2, 2); +x_5 = l_Std_DHashMap_Internal_AssocList_foldrM___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__2(x_1, x_4); +lean_inc(x_3); +x_6 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_6, 0, x_3); +lean_ctor_set(x_6, 1, x_5); +return x_6; +} +} +} +LEAN_EXPORT lean_object* l_List_mapTR_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__3(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_Expr_fvarId_x21(x_5); +lean_dec(x_5); +lean_ctor_set(x_1, 1, x_2); +lean_ctor_set(x_1, 0, x_7); +{ +lean_object* _tmp_0 = x_6; +lean_object* _tmp_1 = x_1; +x_1 = _tmp_0; +x_2 = _tmp_1; +} +goto _start; +} +else +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_9 = lean_ctor_get(x_1, 0); +x_10 = lean_ctor_get(x_1, 1); +lean_inc(x_10); +lean_inc(x_9); +lean_dec(x_1); +x_11 = l_Lean_Expr_fvarId_x21(x_9); +lean_dec(x_9); +x_12 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_2); +x_1 = x_10; +x_2 = x_12; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_RBTree_ofList___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__4(lean_object* x_1) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_2; +x_2 = lean_box(0); +return x_2; +} +else +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_3 = lean_ctor_get(x_1, 0); +lean_inc(x_3); +x_4 = lean_ctor_get(x_1, 1); +lean_inc(x_4); +lean_dec(x_1); +x_5 = l_Lean_RBTree_ofList___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__4(x_4); +x_6 = lean_box(0); +x_7 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_5, x_3, x_6); +return x_7; +} +} +} +LEAN_EXPORT lean_object* l_List_mapTR_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__5(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_4) == 0) +{ +lean_object* x_6; +x_6 = l_List_reverse___rarg(x_5); +return x_6; +} +else +{ +uint8_t x_7; +x_7 = !lean_is_exclusive(x_4); +if (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; +x_8 = lean_ctor_get(x_4, 0); +x_9 = lean_ctor_get(x_4, 1); +x_10 = lean_nat_sub(x_1, x_8); +lean_dec(x_8); +x_11 = lean_unsigned_to_nat(1u); +x_12 = lean_nat_sub(x_10, x_11); +lean_dec(x_10); +x_13 = lean_nat_dec_lt(x_12, x_3); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; +lean_dec(x_12); +x_14 = l_Lean_instInhabitedExpr; +x_15 = l_outOfBounds___rarg(x_14); +x_16 = l_Lean_Expr_fvarId_x21(x_15); lean_dec(x_15); -x_17 = lean_unsigned_to_nat(0u); -x_18 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_17); -x_19 = l_Lean_Meta_Grind_ppPattern___closed__5; +lean_ctor_set(x_4, 1, x_5); +lean_ctor_set(x_4, 0, x_16); +{ +lean_object* _tmp_3 = x_9; +lean_object* _tmp_4 = x_4; +x_4 = _tmp_3; +x_5 = _tmp_4; +} +goto _start; +} +else +{ +lean_object* x_18; lean_object* x_19; +x_18 = lean_array_fget(x_2, x_12); +lean_dec(x_12); +x_19 = l_Lean_Expr_fvarId_x21(x_18); +lean_dec(x_18); +lean_ctor_set(x_4, 1, x_5); +lean_ctor_set(x_4, 0, x_19); +{ +lean_object* _tmp_3 = x_9; +lean_object* _tmp_4 = x_4; +x_4 = _tmp_3; +x_5 = _tmp_4; +} +goto _start; +} +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; +x_21 = lean_ctor_get(x_4, 0); +x_22 = lean_ctor_get(x_4, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_4); +x_23 = lean_nat_sub(x_1, x_21); +lean_dec(x_21); +x_24 = lean_unsigned_to_nat(1u); +x_25 = lean_nat_sub(x_23, x_24); +lean_dec(x_23); +x_26 = lean_nat_dec_lt(x_25, x_3); +if (x_26 == 0) +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +lean_dec(x_25); +x_27 = l_Lean_instInhabitedExpr; +x_28 = l_outOfBounds___rarg(x_27); +x_29 = l_Lean_Expr_fvarId_x21(x_28); +lean_dec(x_28); +x_30 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_5); +x_4 = x_22; +x_5 = x_30; +goto _start; +} +else +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_array_fget(x_2, x_25); +lean_dec(x_25); +x_33 = l_Lean_Expr_fvarId_x21(x_32); +lean_dec(x_32); +x_34 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_34, 0, x_33); +lean_ctor_set(x_34, 1, x_5); +x_4 = x_22; +x_5 = x_34; +goto _start; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { +_start: +{ +uint8_t x_7; +x_7 = lean_usize_dec_eq(x_4, x_5); +if (x_7 == 0) +{ +lean_object* x_8; size_t x_9; size_t x_10; lean_object* x_11; +x_8 = lean_array_uget(x_3, x_4); +x_9 = 1; +x_10 = lean_usize_add(x_4, x_9); +x_11 = l_Lean_RBNode_findCore___at_Lean_Meta_Closure_process___spec__1(x_2, x_8); +if (lean_obj_tag(x_11) == 0) +{ +lean_dec(x_8); +x_4 = x_10; +goto _start; +} +else +{ +lean_object* x_13; lean_object* x_14; +lean_dec(x_11); +x_13 = lean_box(0); +x_14 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_6, x_8, x_13); +x_4 = x_10; +x_6 = x_14; +goto _start; +} +} +else +{ +return x_6; +} +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__6___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__7(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; size_t x_8; size_t x_9; lean_object* x_10; +x_7 = lean_array_uget(x_2, x_3); +x_8 = 1; +x_9 = lean_usize_add(x_3, x_8); +x_10 = l_Lean_RBNode_findCore___at_Lean_Meta_Closure_process___spec__1(x_1, x_7); +if (lean_obj_tag(x_10) == 0) +{ +lean_dec(x_7); +x_3 = x_9; +goto _start; +} +else +{ +lean_object* x_12; lean_object* x_13; +lean_dec(x_10); +x_12 = lean_box(0); +x_13 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_5, x_7, x_12); +x_3 = x_9; +x_5 = x_13; +goto _start; +} +} +else +{ +return x_5; +} +} +} +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___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) { +_start: +{ +if (lean_obj_tag(x_7) == 0) +{ +lean_object* x_15; +lean_dec(x_10); +lean_dec(x_1); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_8); +lean_ctor_set(x_15, 1, x_14); +return x_15; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_7, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_7, 1); +lean_inc(x_17); +lean_dec(x_7); +lean_inc(x_10); +x_18 = l_Lean_FVarId_getType(x_16, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_18) == 0) +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_18, 1); +lean_inc(x_20); +lean_dec(x_18); +x_21 = lean_box(0); +lean_inc(x_1); +x_22 = lean_array_mk(x_1); +x_23 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; +x_24 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set(x_24, 1, x_21); +lean_ctor_set(x_24, 2, x_22); +x_25 = l_Lean_CollectFVars_main(x_19, x_24); +x_26 = lean_ctor_get(x_25, 2); +lean_inc(x_26); +lean_dec(x_25); +x_27 = lean_array_get_size(x_26); +x_28 = lean_unsigned_to_nat(0u); +x_29 = lean_nat_dec_lt(x_28, x_27); +if (x_29 == 0) +{ +lean_dec(x_27); +lean_dec(x_26); +x_7 = x_17; +x_9 = lean_box(0); +x_14 = x_20; +goto _start; +} +else +{ +uint8_t x_31; +x_31 = lean_nat_dec_le(x_27, x_27); +if (x_31 == 0) +{ +lean_dec(x_27); +lean_dec(x_26); +x_7 = x_17; +x_9 = lean_box(0); +x_14 = x_20; +goto _start; +} +else +{ +size_t x_33; size_t x_34; lean_object* x_35; +x_33 = 0; +x_34 = lean_usize_of_nat(x_27); +lean_dec(x_27); +x_35 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__6___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__7(x_3, x_26, x_33, x_34, x_8); +lean_dec(x_26); +x_7 = x_17; +x_8 = x_35; +x_9 = lean_box(0); +x_14 = x_20; +goto _start; +} +} +} +else +{ +uint8_t x_37; +lean_dec(x_17); +lean_dec(x_10); +lean_dec(x_8); +lean_dec(x_1); +x_37 = !lean_is_exclusive(x_18); +if (x_37 == 0) +{ +return x_18; +} +else +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_38 = lean_ctor_get(x_18, 0); +x_39 = lean_ctor_get(x_18, 1); +lean_inc(x_39); +lean_inc(x_38); +lean_dec(x_18); +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_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__8___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__9(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +if (lean_obj_tag(x_6) == 0) +{ +lean_object* x_14; +lean_dec(x_9); +lean_dec(x_1); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_7); +lean_ctor_set(x_14, 1, x_13); +return x_14; +} +else +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_ctor_get(x_6, 0); +lean_inc(x_15); +x_16 = lean_ctor_get(x_6, 1); +lean_inc(x_16); +lean_dec(x_6); +lean_inc(x_9); +x_17 = l_Lean_FVarId_getType(x_15, x_9, x_10, x_11, x_12, 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_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_17, 0); lean_inc(x_18); -x_20 = lean_mk_array(x_18, x_19); -x_21 = lean_unsigned_to_nat(1u); -x_22 = lean_nat_sub(x_18, x_21); +x_19 = lean_ctor_get(x_17, 1); +lean_inc(x_19); +lean_dec(x_17); +x_20 = lean_box(0); +lean_inc(x_1); +x_21 = lean_array_mk(x_1); +x_22 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; +x_23 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_20); +lean_ctor_set(x_23, 2, x_21); +x_24 = l_Lean_CollectFVars_main(x_18, x_23); +x_25 = lean_ctor_get(x_24, 2); +lean_inc(x_25); +lean_dec(x_24); +x_26 = lean_array_get_size(x_25); +x_27 = lean_unsigned_to_nat(0u); +x_28 = lean_nat_dec_lt(x_27, x_26); +if (x_28 == 0) +{ +lean_dec(x_26); +lean_dec(x_25); +x_6 = x_16; +x_8 = lean_box(0); +x_13 = x_19; +goto _start; +} +else +{ +uint8_t x_30; +x_30 = lean_nat_dec_le(x_26, x_26); +if (x_30 == 0) +{ +lean_dec(x_26); +lean_dec(x_25); +x_6 = x_16; +x_8 = lean_box(0); +x_13 = x_19; +goto _start; +} +else +{ +size_t x_32; size_t x_33; lean_object* x_34; +x_32 = 0; +x_33 = lean_usize_of_nat(x_26); +lean_dec(x_26); +x_34 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__6___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__7(x_2, x_25, x_32, x_33, x_7); +lean_dec(x_25); +x_6 = x_16; +x_7 = x_34; +x_8 = lean_box(0); +x_13 = x_19; +goto _start; +} +} +} +else +{ +uint8_t x_36; +lean_dec(x_16); +lean_dec(x_9); +lean_dec(x_7); +lean_dec(x_1); +x_36 = !lean_is_exclusive(x_17); +if (x_36 == 0) +{ +return x_17; +} +else +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_37 = lean_ctor_get(x_17, 0); +x_38 = lean_ctor_get(x_17, 1); +lean_inc(x_38); +lean_inc(x_37); +lean_dec(x_17); +x_39 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_39, 0, x_37); +lean_ctor_set(x_39, 1, x_38); +return x_39; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__10(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { +_start: +{ +uint8_t x_7; +x_7 = lean_usize_dec_eq(x_4, x_5); +if (x_7 == 0) +{ +lean_object* x_8; size_t x_9; size_t x_10; lean_object* x_11; +x_8 = lean_array_uget(x_3, x_4); +x_9 = 1; +x_10 = lean_usize_add(x_4, x_9); +x_11 = l_Lean_RBNode_findCore___at_Lean_Meta_Closure_process___spec__1(x_2, x_8); +if (lean_obj_tag(x_11) == 0) +{ +lean_dec(x_8); +x_4 = x_10; +goto _start; +} +else +{ +lean_object* x_13; lean_object* x_14; +lean_dec(x_11); +x_13 = lean_box(0); +x_14 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_6, x_8, x_13); +x_4 = x_10; +x_6 = x_14; +goto _start; +} +} +else +{ +return x_6; +} +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__10___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__11(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5) { +_start: +{ +uint8_t x_6; +x_6 = lean_usize_dec_eq(x_3, x_4); +if (x_6 == 0) +{ +lean_object* x_7; size_t x_8; size_t x_9; lean_object* x_10; +x_7 = lean_array_uget(x_2, x_3); +x_8 = 1; +x_9 = lean_usize_add(x_3, x_8); +x_10 = l_Lean_RBNode_findCore___at_Lean_Meta_Closure_process___spec__1(x_1, x_7); +if (lean_obj_tag(x_10) == 0) +{ +lean_dec(x_7); +x_3 = x_9; +goto _start; +} +else +{ +lean_object* x_12; lean_object* x_13; +lean_dec(x_10); +x_12 = lean_box(0); +x_13 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_5, x_7, x_12); +x_3 = x_9; +x_5 = x_13; +goto _start; +} +} +else +{ +return x_5; +} +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__12(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { +_start: +{ +uint8_t x_7; +x_7 = lean_usize_dec_eq(x_4, x_5); +if (x_7 == 0) +{ +lean_object* x_8; size_t x_9; size_t x_10; lean_object* x_11; +x_8 = lean_array_uget(x_3, x_4); +x_9 = 1; +x_10 = lean_usize_add(x_4, x_9); +x_11 = l_Lean_RBNode_findCore___at_Lean_Meta_Closure_process___spec__1(x_2, x_8); +if (lean_obj_tag(x_11) == 0) +{ +lean_dec(x_8); +x_4 = x_10; +goto _start; +} +else +{ +lean_object* x_13; lean_object* x_14; +lean_dec(x_11); +x_13 = lean_box(0); +x_14 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_6, x_8, x_13); +x_4 = x_10; +x_6 = x_14; +goto _start; +} +} +else +{ +return x_6; +} +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__12___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__13(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5) { +_start: +{ +uint8_t x_6; +x_6 = lean_usize_dec_eq(x_3, x_4); +if (x_6 == 0) +{ +lean_object* x_7; size_t x_8; size_t x_9; lean_object* x_10; +x_7 = lean_array_uget(x_2, x_3); +x_8 = 1; +x_9 = lean_usize_add(x_3, x_8); +x_10 = l_Lean_RBNode_findCore___at_Lean_Meta_Closure_process___spec__1(x_1, x_7); +if (lean_obj_tag(x_10) == 0) +{ +lean_dec(x_7); +x_3 = x_9; +goto _start; +} +else +{ +lean_object* x_12; lean_object* x_13; +lean_dec(x_10); +x_12 = lean_box(0); +x_13 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_5, x_7, x_12); +x_3 = x_9; +x_5 = x_13; +goto _start; +} +} +else +{ +return x_5; +} +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__14(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, size_t x_7, size_t x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +uint8_t x_15; +x_15 = lean_usize_dec_lt(x_8, x_7); +if (x_15 == 0) +{ +lean_object* x_16; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_4); +lean_dec(x_2); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_9); +lean_ctor_set(x_16, 1, x_14); +return x_16; +} +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_25; +x_17 = lean_array_uget(x_6, x_8); +x_25 = !lean_is_exclusive(x_9); +if (x_25 == 0) +{ +lean_object* x_26; uint8_t x_27; +x_26 = lean_ctor_get(x_9, 1); +x_27 = !lean_is_exclusive(x_26); +if (x_27 == 0) +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_28 = lean_ctor_get(x_9, 0); +x_29 = lean_ctor_get(x_26, 0); +x_30 = lean_ctor_get(x_26, 1); +x_31 = l_Lean_Expr_fvarId_x21(x_17); +x_32 = l_Lean_RBNode_findCore___at_Lean_Meta_Closure_process___spec__1(x_30, x_31); +if (lean_obj_tag(x_32) == 0) +{ +lean_object* x_33; +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +x_33 = lean_infer_type(x_17, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_33) == 0) +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_33, 0); +lean_inc(x_34); +x_35 = lean_ctor_get(x_33, 1); +lean_inc(x_35); +lean_dec(x_33); +x_36 = l_Lean_RBNode_findCore___at_Lean_Meta_Closure_process___spec__1(x_28, x_31); +if (lean_obj_tag(x_36) == 0) +{ +lean_object* x_37; +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_34); +x_37 = l_Lean_Meta_isProp(x_34, x_10, x_11, x_12, x_13, x_35); +if (lean_obj_tag(x_37) == 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_40; lean_object* x_41; +x_40 = lean_ctor_get(x_37, 1); +lean_inc(x_40); +lean_dec(x_37); +lean_inc(x_10); +lean_inc(x_31); +x_41 = l_Lean_FVarId_getDecl(x_31, x_10, x_11, x_12, x_13, x_40); +if (lean_obj_tag(x_41) == 0) +{ +lean_object* x_42; lean_object* x_43; uint8_t 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 = l_Lean_LocalDecl_binderInfo(x_42); +lean_dec(x_42); +x_45 = lean_box(x_44); +if (lean_obj_tag(x_45) == 3) +{ +lean_object* x_46; +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_34); +lean_inc(x_28); +lean_inc(x_4); +x_46 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized(x_4, x_28, x_34, x_10, x_11, x_12, x_13, x_43); +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); +lean_dec(x_47); +if (x_48 == 0) +{ +lean_object* x_49; lean_object* x_50; +lean_dec(x_34); +lean_dec(x_31); +x_49 = lean_ctor_get(x_46, 1); +lean_inc(x_49); +lean_dec(x_46); +x_50 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_50, 0, x_9); +x_18 = x_50; +x_19 = x_49; +goto block_24; +} +else +{ +lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; uint8_t x_62; lean_object* x_63; +lean_dec(x_29); +x_51 = lean_ctor_get(x_46, 1); +lean_inc(x_51); +lean_dec(x_46); +x_52 = lean_box(0); +lean_inc(x_31); +x_53 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_28, x_31, x_52); +x_54 = lean_box(0); +lean_inc(x_2); +x_55 = lean_array_mk(x_2); +x_56 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; +x_57 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_57, 0, x_56); +lean_ctor_set(x_57, 1, x_54); +lean_ctor_set(x_57, 2, x_55); +x_58 = l_Lean_CollectFVars_main(x_34, x_57); +x_59 = lean_ctor_get(x_58, 2); +lean_inc(x_59); +lean_dec(x_58); +x_60 = lean_array_get_size(x_59); +x_61 = lean_unsigned_to_nat(0u); +x_62 = lean_nat_dec_lt(x_61, x_60); +x_63 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_30, x_31, x_52); +if (x_62 == 0) +{ +uint8_t x_64; lean_object* x_65; lean_object* x_66; +lean_dec(x_60); +lean_dec(x_59); +x_64 = 1; +x_65 = lean_box(x_64); +lean_ctor_set(x_26, 1, x_63); +lean_ctor_set(x_26, 0, x_65); +lean_ctor_set(x_9, 0, x_53); +x_66 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_66, 0, x_9); +x_18 = x_66; +x_19 = x_51; +goto block_24; +} +else +{ +uint8_t x_67; +x_67 = lean_nat_dec_le(x_60, x_60); +if (x_67 == 0) +{ +uint8_t x_68; lean_object* x_69; lean_object* x_70; +lean_dec(x_60); +lean_dec(x_59); +x_68 = 1; +x_69 = lean_box(x_68); +lean_ctor_set(x_26, 1, x_63); +lean_ctor_set(x_26, 0, x_69); +lean_ctor_set(x_9, 0, x_53); +x_70 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_70, 0, x_9); +x_18 = x_70; +x_19 = x_51; +goto block_24; +} +else +{ +size_t x_71; size_t x_72; lean_object* x_73; uint8_t x_74; lean_object* x_75; lean_object* x_76; +x_71 = 0; +x_72 = lean_usize_of_nat(x_60); +lean_dec(x_60); +x_73 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__10___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__11(x_4, x_59, x_71, x_72, x_53); +lean_dec(x_59); +x_74 = 1; +x_75 = lean_box(x_74); +lean_ctor_set(x_26, 1, x_63); +lean_ctor_set(x_26, 0, x_75); +lean_ctor_set(x_9, 0, x_73); +x_76 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_76, 0, x_9); +x_18 = x_76; +x_19 = x_51; +goto block_24; +} +} +} +} +else +{ +uint8_t x_77; +lean_dec(x_34); +lean_dec(x_31); +lean_free_object(x_26); +lean_dec(x_30); +lean_dec(x_29); +lean_free_object(x_9); +lean_dec(x_28); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_4); +lean_dec(x_2); +x_77 = !lean_is_exclusive(x_46); +if (x_77 == 0) +{ +return x_46; +} +else +{ +lean_object* x_78; lean_object* x_79; lean_object* x_80; +x_78 = lean_ctor_get(x_46, 0); +x_79 = lean_ctor_get(x_46, 1); +lean_inc(x_79); +lean_inc(x_78); +lean_dec(x_46); +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_81; +lean_dec(x_45); +lean_dec(x_34); +lean_dec(x_31); +x_81 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_81, 0, x_9); +x_18 = x_81; +x_19 = x_43; +goto block_24; +} +} +else +{ +uint8_t x_82; +lean_dec(x_34); +lean_dec(x_31); +lean_free_object(x_26); +lean_dec(x_30); +lean_dec(x_29); +lean_free_object(x_9); +lean_dec(x_28); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_4); +lean_dec(x_2); +x_82 = !lean_is_exclusive(x_41); +if (x_82 == 0) +{ +return x_41; +} +else +{ +lean_object* x_83; lean_object* x_84; lean_object* x_85; +x_83 = lean_ctor_get(x_41, 0); +x_84 = lean_ctor_get(x_41, 1); +lean_inc(x_84); +lean_inc(x_83); +lean_dec(x_41); +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_86; uint8_t x_87; +x_86 = lean_ctor_get(x_37, 1); +lean_inc(x_86); +lean_dec(x_37); +x_87 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars(x_4, x_28, x_34); +if (x_87 == 0) +{ +lean_object* x_88; +lean_dec(x_31); +x_88 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_88, 0, x_9); +x_18 = x_88; +x_19 = x_86; +goto block_24; +} +else +{ +lean_object* x_89; lean_object* x_90; lean_object* x_91; uint8_t x_92; lean_object* x_93; lean_object* x_94; +lean_dec(x_29); +x_89 = lean_box(0); +lean_inc(x_31); +x_90 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_28, x_31, x_89); +x_91 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_30, x_31, x_89); +x_92 = 1; +x_93 = lean_box(x_92); +lean_ctor_set(x_26, 1, x_91); +lean_ctor_set(x_26, 0, x_93); +lean_ctor_set(x_9, 0, x_90); +x_94 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_94, 0, x_9); +x_18 = x_94; +x_19 = x_86; +goto block_24; +} +} +} +else +{ +uint8_t x_95; +lean_dec(x_34); +lean_dec(x_31); +lean_free_object(x_26); +lean_dec(x_30); +lean_dec(x_29); +lean_free_object(x_9); +lean_dec(x_28); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_4); +lean_dec(x_2); +x_95 = !lean_is_exclusive(x_37); +if (x_95 == 0) +{ +return x_37; +} +else +{ +lean_object* x_96; lean_object* x_97; lean_object* x_98; +x_96 = lean_ctor_get(x_37, 0); +x_97 = lean_ctor_get(x_37, 1); +lean_inc(x_97); +lean_inc(x_96); +lean_dec(x_37); +x_98 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_98, 0, x_96); +lean_ctor_set(x_98, 1, x_97); +return x_98; +} +} +} +else +{ +uint8_t x_99; +lean_dec(x_29); +x_99 = !lean_is_exclusive(x_36); +if (x_99 == 0) +{ +lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; uint8_t x_109; lean_object* x_110; lean_object* x_111; +x_100 = lean_ctor_get(x_36, 0); +lean_dec(x_100); +x_101 = lean_box(0); +lean_inc(x_2); +x_102 = lean_array_mk(x_2); +x_103 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; +x_104 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_104, 0, x_103); +lean_ctor_set(x_104, 1, x_101); +lean_ctor_set(x_104, 2, x_102); +x_105 = l_Lean_CollectFVars_main(x_34, x_104); +x_106 = lean_ctor_get(x_105, 2); +lean_inc(x_106); +lean_dec(x_105); +x_107 = lean_array_get_size(x_106); +x_108 = lean_unsigned_to_nat(0u); +x_109 = lean_nat_dec_lt(x_108, x_107); +x_110 = lean_box(0); +x_111 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_30, x_31, x_110); +if (x_109 == 0) +{ +uint8_t x_112; lean_object* x_113; +lean_dec(x_107); +lean_dec(x_106); +x_112 = 1; +x_113 = lean_box(x_112); +lean_ctor_set(x_26, 1, x_111); +lean_ctor_set(x_26, 0, x_113); +lean_ctor_set(x_36, 0, x_9); +x_18 = x_36; +x_19 = x_35; +goto block_24; +} +else +{ +uint8_t x_114; +x_114 = lean_nat_dec_le(x_107, x_107); +if (x_114 == 0) +{ +uint8_t x_115; lean_object* x_116; +lean_dec(x_107); +lean_dec(x_106); +x_115 = 1; +x_116 = lean_box(x_115); +lean_ctor_set(x_26, 1, x_111); +lean_ctor_set(x_26, 0, x_116); +lean_ctor_set(x_36, 0, x_9); +x_18 = x_36; +x_19 = x_35; +goto block_24; +} +else +{ +size_t x_117; size_t x_118; lean_object* x_119; uint8_t x_120; lean_object* x_121; +x_117 = 0; +x_118 = lean_usize_of_nat(x_107); +lean_dec(x_107); +x_119 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__12___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__13(x_4, x_106, x_117, x_118, x_28); +lean_dec(x_106); +x_120 = 1; +x_121 = lean_box(x_120); +lean_ctor_set(x_26, 1, x_111); +lean_ctor_set(x_26, 0, x_121); +lean_ctor_set(x_9, 0, x_119); +lean_ctor_set(x_36, 0, x_9); +x_18 = x_36; +x_19 = x_35; +goto block_24; +} +} +} +else +{ +lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; uint8_t x_130; lean_object* x_131; lean_object* x_132; +lean_dec(x_36); +x_122 = lean_box(0); +lean_inc(x_2); +x_123 = lean_array_mk(x_2); +x_124 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; +x_125 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_125, 0, x_124); +lean_ctor_set(x_125, 1, x_122); +lean_ctor_set(x_125, 2, x_123); +x_126 = l_Lean_CollectFVars_main(x_34, x_125); +x_127 = lean_ctor_get(x_126, 2); +lean_inc(x_127); +lean_dec(x_126); +x_128 = lean_array_get_size(x_127); +x_129 = lean_unsigned_to_nat(0u); +x_130 = lean_nat_dec_lt(x_129, x_128); +x_131 = lean_box(0); +x_132 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_30, x_31, x_131); +if (x_130 == 0) +{ +uint8_t x_133; lean_object* x_134; lean_object* x_135; +lean_dec(x_128); +lean_dec(x_127); +x_133 = 1; +x_134 = lean_box(x_133); +lean_ctor_set(x_26, 1, x_132); +lean_ctor_set(x_26, 0, x_134); +x_135 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_135, 0, x_9); +x_18 = x_135; +x_19 = x_35; +goto block_24; +} +else +{ +uint8_t x_136; +x_136 = lean_nat_dec_le(x_128, x_128); +if (x_136 == 0) +{ +uint8_t x_137; lean_object* x_138; lean_object* x_139; +lean_dec(x_128); +lean_dec(x_127); +x_137 = 1; +x_138 = lean_box(x_137); +lean_ctor_set(x_26, 1, x_132); +lean_ctor_set(x_26, 0, x_138); +x_139 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_139, 0, x_9); +x_18 = x_139; +x_19 = x_35; +goto block_24; +} +else +{ +size_t x_140; size_t x_141; lean_object* x_142; uint8_t x_143; lean_object* x_144; lean_object* x_145; +x_140 = 0; +x_141 = lean_usize_of_nat(x_128); +lean_dec(x_128); +x_142 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__12___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__13(x_4, x_127, x_140, x_141, x_28); +lean_dec(x_127); +x_143 = 1; +x_144 = lean_box(x_143); +lean_ctor_set(x_26, 1, x_132); +lean_ctor_set(x_26, 0, x_144); +lean_ctor_set(x_9, 0, x_142); +x_145 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_145, 0, x_9); +x_18 = x_145; +x_19 = x_35; +goto block_24; +} +} +} +} +} +else +{ +uint8_t x_146; +lean_dec(x_31); +lean_free_object(x_26); +lean_dec(x_30); +lean_dec(x_29); +lean_free_object(x_9); +lean_dec(x_28); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_4); +lean_dec(x_2); +x_146 = !lean_is_exclusive(x_33); +if (x_146 == 0) +{ +return x_33; +} +else +{ +lean_object* x_147; lean_object* x_148; lean_object* x_149; +x_147 = lean_ctor_get(x_33, 0); +x_148 = lean_ctor_get(x_33, 1); +lean_inc(x_148); +lean_inc(x_147); +lean_dec(x_33); +x_149 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_149, 0, x_147); +lean_ctor_set(x_149, 1, x_148); +return x_149; +} +} +} +else +{ +uint8_t x_150; +lean_dec(x_31); +lean_dec(x_17); +x_150 = !lean_is_exclusive(x_32); +if (x_150 == 0) +{ +lean_object* x_151; +x_151 = lean_ctor_get(x_32, 0); +lean_dec(x_151); +lean_ctor_set(x_32, 0, x_9); +x_18 = x_32; +x_19 = x_14; +goto block_24; +} +else +{ +lean_object* x_152; +lean_dec(x_32); +x_152 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_152, 0, x_9); +x_18 = x_152; +x_19 = x_14; +goto block_24; +} +} +} +else +{ +lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; +x_153 = lean_ctor_get(x_9, 0); +x_154 = lean_ctor_get(x_26, 0); +x_155 = lean_ctor_get(x_26, 1); +lean_inc(x_155); +lean_inc(x_154); +lean_dec(x_26); +x_156 = l_Lean_Expr_fvarId_x21(x_17); +x_157 = l_Lean_RBNode_findCore___at_Lean_Meta_Closure_process___spec__1(x_155, x_156); +if (lean_obj_tag(x_157) == 0) +{ +lean_object* x_158; +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +x_158 = lean_infer_type(x_17, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_158) == 0) +{ +lean_object* x_159; lean_object* x_160; lean_object* x_161; +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_RBNode_findCore___at_Lean_Meta_Closure_process___spec__1(x_153, x_156); +if (lean_obj_tag(x_161) == 0) +{ +lean_object* x_162; +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_159); +x_162 = l_Lean_Meta_isProp(x_159, x_10, x_11, x_12, x_13, x_160); +if (lean_obj_tag(x_162) == 0) +{ +lean_object* x_163; uint8_t x_164; +x_163 = lean_ctor_get(x_162, 0); +lean_inc(x_163); +x_164 = lean_unbox(x_163); +lean_dec(x_163); +if (x_164 == 0) +{ +lean_object* x_165; lean_object* x_166; +x_165 = lean_ctor_get(x_162, 1); +lean_inc(x_165); +lean_dec(x_162); +lean_inc(x_10); +lean_inc(x_156); +x_166 = l_Lean_FVarId_getDecl(x_156, x_10, x_11, x_12, x_13, x_165); +if (lean_obj_tag(x_166) == 0) +{ +lean_object* x_167; lean_object* x_168; uint8_t x_169; lean_object* x_170; +x_167 = lean_ctor_get(x_166, 0); +lean_inc(x_167); +x_168 = lean_ctor_get(x_166, 1); +lean_inc(x_168); +lean_dec(x_166); +x_169 = l_Lean_LocalDecl_binderInfo(x_167); +lean_dec(x_167); +x_170 = lean_box(x_169); +if (lean_obj_tag(x_170) == 3) +{ +lean_object* x_171; +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_159); +lean_inc(x_153); +lean_inc(x_4); +x_171 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized(x_4, x_153, x_159, x_10, x_11, x_12, x_13, x_168); +if (lean_obj_tag(x_171) == 0) +{ +lean_object* x_172; uint8_t x_173; +x_172 = lean_ctor_get(x_171, 0); +lean_inc(x_172); +x_173 = lean_unbox(x_172); +lean_dec(x_172); +if (x_173 == 0) +{ +lean_object* x_174; lean_object* x_175; lean_object* x_176; +lean_dec(x_159); +lean_dec(x_156); +x_174 = lean_ctor_get(x_171, 1); +lean_inc(x_174); +lean_dec(x_171); +x_175 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_175, 0, x_154); +lean_ctor_set(x_175, 1, x_155); +lean_ctor_set(x_9, 1, x_175); +x_176 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_176, 0, x_9); +x_18 = x_176; +x_19 = x_174; +goto block_24; +} +else +{ +lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; uint8_t x_188; lean_object* x_189; +lean_dec(x_154); +x_177 = lean_ctor_get(x_171, 1); +lean_inc(x_177); +lean_dec(x_171); +x_178 = lean_box(0); +lean_inc(x_156); +x_179 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_153, x_156, x_178); +x_180 = lean_box(0); +lean_inc(x_2); +x_181 = lean_array_mk(x_2); +x_182 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; +x_183 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_183, 0, x_182); +lean_ctor_set(x_183, 1, x_180); +lean_ctor_set(x_183, 2, x_181); +x_184 = l_Lean_CollectFVars_main(x_159, x_183); +x_185 = lean_ctor_get(x_184, 2); +lean_inc(x_185); +lean_dec(x_184); +x_186 = lean_array_get_size(x_185); +x_187 = lean_unsigned_to_nat(0u); +x_188 = lean_nat_dec_lt(x_187, x_186); +x_189 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_155, x_156, x_178); +if (x_188 == 0) +{ +uint8_t x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; +lean_dec(x_186); +lean_dec(x_185); +x_190 = 1; +x_191 = lean_box(x_190); +x_192 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_192, 0, x_191); +lean_ctor_set(x_192, 1, x_189); +lean_ctor_set(x_9, 1, x_192); +lean_ctor_set(x_9, 0, x_179); +x_193 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_193, 0, x_9); +x_18 = x_193; +x_19 = x_177; +goto block_24; +} +else +{ +uint8_t x_194; +x_194 = lean_nat_dec_le(x_186, x_186); +if (x_194 == 0) +{ +uint8_t x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; +lean_dec(x_186); +lean_dec(x_185); +x_195 = 1; +x_196 = lean_box(x_195); +x_197 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_197, 0, x_196); +lean_ctor_set(x_197, 1, x_189); +lean_ctor_set(x_9, 1, x_197); +lean_ctor_set(x_9, 0, x_179); +x_198 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_198, 0, x_9); +x_18 = x_198; +x_19 = x_177; +goto block_24; +} +else +{ +size_t x_199; size_t x_200; lean_object* x_201; uint8_t x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; +x_199 = 0; +x_200 = lean_usize_of_nat(x_186); +lean_dec(x_186); +x_201 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__10___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__11(x_4, x_185, x_199, x_200, x_179); +lean_dec(x_185); +x_202 = 1; +x_203 = lean_box(x_202); +x_204 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_204, 0, x_203); +lean_ctor_set(x_204, 1, x_189); +lean_ctor_set(x_9, 1, x_204); +lean_ctor_set(x_9, 0, x_201); +x_205 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_205, 0, x_9); +x_18 = x_205; +x_19 = x_177; +goto block_24; +} +} +} +} +else +{ +lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; +lean_dec(x_159); +lean_dec(x_156); +lean_dec(x_155); +lean_dec(x_154); +lean_free_object(x_9); +lean_dec(x_153); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_4); +lean_dec(x_2); +x_206 = lean_ctor_get(x_171, 0); +lean_inc(x_206); +x_207 = lean_ctor_get(x_171, 1); +lean_inc(x_207); +if (lean_is_exclusive(x_171)) { + lean_ctor_release(x_171, 0); + lean_ctor_release(x_171, 1); + x_208 = x_171; +} else { + lean_dec_ref(x_171); + x_208 = lean_box(0); +} +if (lean_is_scalar(x_208)) { + x_209 = lean_alloc_ctor(1, 2, 0); +} else { + x_209 = x_208; +} +lean_ctor_set(x_209, 0, x_206); +lean_ctor_set(x_209, 1, x_207); +return x_209; +} +} +else +{ +lean_object* x_210; lean_object* x_211; +lean_dec(x_170); +lean_dec(x_159); +lean_dec(x_156); +x_210 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_210, 0, x_154); +lean_ctor_set(x_210, 1, x_155); +lean_ctor_set(x_9, 1, x_210); +x_211 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_211, 0, x_9); +x_18 = x_211; +x_19 = x_168; +goto block_24; +} +} +else +{ +lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; +lean_dec(x_159); +lean_dec(x_156); +lean_dec(x_155); +lean_dec(x_154); +lean_free_object(x_9); +lean_dec(x_153); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_4); +lean_dec(x_2); +x_212 = lean_ctor_get(x_166, 0); +lean_inc(x_212); +x_213 = lean_ctor_get(x_166, 1); +lean_inc(x_213); +if (lean_is_exclusive(x_166)) { + lean_ctor_release(x_166, 0); + lean_ctor_release(x_166, 1); + x_214 = x_166; +} else { + lean_dec_ref(x_166); + x_214 = lean_box(0); +} +if (lean_is_scalar(x_214)) { + x_215 = lean_alloc_ctor(1, 2, 0); +} else { + x_215 = x_214; +} +lean_ctor_set(x_215, 0, x_212); +lean_ctor_set(x_215, 1, x_213); +return x_215; +} +} +else +{ +lean_object* x_216; uint8_t x_217; +x_216 = lean_ctor_get(x_162, 1); +lean_inc(x_216); +lean_dec(x_162); +x_217 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars(x_4, x_153, x_159); +if (x_217 == 0) +{ +lean_object* x_218; lean_object* x_219; +lean_dec(x_156); +x_218 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_218, 0, x_154); +lean_ctor_set(x_218, 1, x_155); +lean_ctor_set(x_9, 1, x_218); +x_219 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_219, 0, x_9); +x_18 = x_219; +x_19 = x_216; +goto block_24; +} +else +{ +lean_object* x_220; lean_object* x_221; lean_object* x_222; uint8_t x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; +lean_dec(x_154); +x_220 = lean_box(0); +lean_inc(x_156); +x_221 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_153, x_156, x_220); +x_222 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_155, x_156, x_220); +x_223 = 1; +x_224 = lean_box(x_223); +x_225 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_225, 0, x_224); +lean_ctor_set(x_225, 1, x_222); +lean_ctor_set(x_9, 1, x_225); +lean_ctor_set(x_9, 0, x_221); +x_226 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_226, 0, x_9); +x_18 = x_226; +x_19 = x_216; +goto block_24; +} +} +} +else +{ +lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; +lean_dec(x_159); +lean_dec(x_156); +lean_dec(x_155); +lean_dec(x_154); +lean_free_object(x_9); +lean_dec(x_153); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_4); +lean_dec(x_2); +x_227 = lean_ctor_get(x_162, 0); +lean_inc(x_227); +x_228 = lean_ctor_get(x_162, 1); +lean_inc(x_228); +if (lean_is_exclusive(x_162)) { + lean_ctor_release(x_162, 0); + lean_ctor_release(x_162, 1); + x_229 = x_162; +} else { + lean_dec_ref(x_162); + x_229 = lean_box(0); +} +if (lean_is_scalar(x_229)) { + x_230 = lean_alloc_ctor(1, 2, 0); +} else { + x_230 = x_229; +} +lean_ctor_set(x_230, 0, x_227); +lean_ctor_set(x_230, 1, x_228); +return x_230; +} +} +else +{ +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; uint8_t x_240; lean_object* x_241; lean_object* x_242; +lean_dec(x_154); +if (lean_is_exclusive(x_161)) { + lean_ctor_release(x_161, 0); + x_231 = x_161; +} else { + lean_dec_ref(x_161); + x_231 = lean_box(0); +} +x_232 = lean_box(0); +lean_inc(x_2); +x_233 = lean_array_mk(x_2); +x_234 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; +x_235 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_235, 0, x_234); +lean_ctor_set(x_235, 1, x_232); +lean_ctor_set(x_235, 2, x_233); +x_236 = l_Lean_CollectFVars_main(x_159, x_235); +x_237 = lean_ctor_get(x_236, 2); +lean_inc(x_237); +lean_dec(x_236); +x_238 = lean_array_get_size(x_237); +x_239 = lean_unsigned_to_nat(0u); +x_240 = lean_nat_dec_lt(x_239, x_238); +x_241 = lean_box(0); +x_242 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_155, x_156, x_241); +if (x_240 == 0) +{ +uint8_t x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; +lean_dec(x_238); +lean_dec(x_237); +x_243 = 1; +x_244 = lean_box(x_243); +x_245 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_245, 0, x_244); +lean_ctor_set(x_245, 1, x_242); +lean_ctor_set(x_9, 1, x_245); +if (lean_is_scalar(x_231)) { + x_246 = lean_alloc_ctor(1, 1, 0); +} else { + x_246 = x_231; +} +lean_ctor_set(x_246, 0, x_9); +x_18 = x_246; +x_19 = x_160; +goto block_24; +} +else +{ +uint8_t x_247; +x_247 = lean_nat_dec_le(x_238, x_238); +if (x_247 == 0) +{ +uint8_t x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; +lean_dec(x_238); +lean_dec(x_237); +x_248 = 1; +x_249 = lean_box(x_248); +x_250 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_250, 0, x_249); +lean_ctor_set(x_250, 1, x_242); +lean_ctor_set(x_9, 1, x_250); +if (lean_is_scalar(x_231)) { + x_251 = lean_alloc_ctor(1, 1, 0); +} else { + x_251 = x_231; +} +lean_ctor_set(x_251, 0, x_9); +x_18 = x_251; +x_19 = x_160; +goto block_24; +} +else +{ +size_t x_252; size_t x_253; lean_object* x_254; uint8_t x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; +x_252 = 0; +x_253 = lean_usize_of_nat(x_238); +lean_dec(x_238); +x_254 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__12___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__13(x_4, x_237, x_252, x_253, x_153); +lean_dec(x_237); +x_255 = 1; +x_256 = lean_box(x_255); +x_257 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_257, 0, x_256); +lean_ctor_set(x_257, 1, x_242); +lean_ctor_set(x_9, 1, x_257); +lean_ctor_set(x_9, 0, x_254); +if (lean_is_scalar(x_231)) { + x_258 = lean_alloc_ctor(1, 1, 0); +} else { + x_258 = x_231; +} +lean_ctor_set(x_258, 0, x_9); +x_18 = x_258; +x_19 = x_160; +goto block_24; +} +} +} +} +else +{ +lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; +lean_dec(x_156); +lean_dec(x_155); +lean_dec(x_154); +lean_free_object(x_9); +lean_dec(x_153); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_4); +lean_dec(x_2); +x_259 = lean_ctor_get(x_158, 0); +lean_inc(x_259); +x_260 = lean_ctor_get(x_158, 1); +lean_inc(x_260); +if (lean_is_exclusive(x_158)) { + lean_ctor_release(x_158, 0); + lean_ctor_release(x_158, 1); + x_261 = x_158; +} else { + lean_dec_ref(x_158); + x_261 = lean_box(0); +} +if (lean_is_scalar(x_261)) { + x_262 = lean_alloc_ctor(1, 2, 0); +} else { + x_262 = x_261; +} +lean_ctor_set(x_262, 0, x_259); +lean_ctor_set(x_262, 1, x_260); +return x_262; +} +} +else +{ +lean_object* x_263; lean_object* x_264; lean_object* x_265; +lean_dec(x_156); +lean_dec(x_17); +if (lean_is_exclusive(x_157)) { + lean_ctor_release(x_157, 0); + x_263 = x_157; +} else { + lean_dec_ref(x_157); + x_263 = lean_box(0); +} +x_264 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_264, 0, x_154); +lean_ctor_set(x_264, 1, x_155); +lean_ctor_set(x_9, 1, x_264); +if (lean_is_scalar(x_263)) { + x_265 = lean_alloc_ctor(1, 1, 0); +} else { + x_265 = x_263; +} +lean_ctor_set(x_265, 0, x_9); +x_18 = x_265; +x_19 = x_14; +goto block_24; +} +} +} +else +{ +lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; +x_266 = lean_ctor_get(x_9, 1); +x_267 = lean_ctor_get(x_9, 0); +lean_inc(x_266); +lean_inc(x_267); +lean_dec(x_9); +x_268 = lean_ctor_get(x_266, 0); +lean_inc(x_268); +x_269 = lean_ctor_get(x_266, 1); +lean_inc(x_269); +if (lean_is_exclusive(x_266)) { + lean_ctor_release(x_266, 0); + lean_ctor_release(x_266, 1); + x_270 = x_266; +} else { + lean_dec_ref(x_266); + x_270 = lean_box(0); +} +x_271 = l_Lean_Expr_fvarId_x21(x_17); +x_272 = l_Lean_RBNode_findCore___at_Lean_Meta_Closure_process___spec__1(x_269, x_271); +if (lean_obj_tag(x_272) == 0) +{ +lean_object* x_273; +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +x_273 = lean_infer_type(x_17, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_273) == 0) +{ +lean_object* x_274; lean_object* x_275; lean_object* x_276; +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 = l_Lean_RBNode_findCore___at_Lean_Meta_Closure_process___spec__1(x_267, x_271); +if (lean_obj_tag(x_276) == 0) +{ +lean_object* x_277; +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_274); +x_277 = l_Lean_Meta_isProp(x_274, x_10, x_11, x_12, x_13, x_275); +if (lean_obj_tag(x_277) == 0) +{ +lean_object* x_278; uint8_t x_279; +x_278 = lean_ctor_get(x_277, 0); +lean_inc(x_278); +x_279 = lean_unbox(x_278); +lean_dec(x_278); +if (x_279 == 0) +{ +lean_object* x_280; lean_object* x_281; +x_280 = lean_ctor_get(x_277, 1); +lean_inc(x_280); +lean_dec(x_277); +lean_inc(x_10); +lean_inc(x_271); +x_281 = l_Lean_FVarId_getDecl(x_271, x_10, x_11, x_12, x_13, x_280); +if (lean_obj_tag(x_281) == 0) +{ +lean_object* x_282; lean_object* x_283; uint8_t x_284; lean_object* x_285; +x_282 = lean_ctor_get(x_281, 0); +lean_inc(x_282); +x_283 = lean_ctor_get(x_281, 1); +lean_inc(x_283); +lean_dec(x_281); +x_284 = l_Lean_LocalDecl_binderInfo(x_282); +lean_dec(x_282); +x_285 = lean_box(x_284); +if (lean_obj_tag(x_285) == 3) +{ +lean_object* x_286; +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_274); +lean_inc(x_267); +lean_inc(x_4); +x_286 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized(x_4, x_267, x_274, x_10, x_11, x_12, x_13, x_283); +if (lean_obj_tag(x_286) == 0) +{ +lean_object* x_287; uint8_t x_288; +x_287 = lean_ctor_get(x_286, 0); +lean_inc(x_287); +x_288 = lean_unbox(x_287); +lean_dec(x_287); +if (x_288 == 0) +{ +lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; +lean_dec(x_274); +lean_dec(x_271); +x_289 = lean_ctor_get(x_286, 1); +lean_inc(x_289); +lean_dec(x_286); +if (lean_is_scalar(x_270)) { + x_290 = lean_alloc_ctor(0, 2, 0); +} else { + x_290 = x_270; +} +lean_ctor_set(x_290, 0, x_268); +lean_ctor_set(x_290, 1, x_269); +x_291 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_291, 0, x_267); +lean_ctor_set(x_291, 1, x_290); +x_292 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_292, 0, x_291); +x_18 = x_292; +x_19 = x_289; +goto block_24; +} +else +{ +lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; uint8_t x_304; lean_object* x_305; +lean_dec(x_268); +x_293 = lean_ctor_get(x_286, 1); +lean_inc(x_293); +lean_dec(x_286); +x_294 = lean_box(0); +lean_inc(x_271); +x_295 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_267, x_271, x_294); +x_296 = lean_box(0); +lean_inc(x_2); +x_297 = lean_array_mk(x_2); +x_298 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; +x_299 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_299, 0, x_298); +lean_ctor_set(x_299, 1, x_296); +lean_ctor_set(x_299, 2, x_297); +x_300 = l_Lean_CollectFVars_main(x_274, x_299); +x_301 = lean_ctor_get(x_300, 2); +lean_inc(x_301); +lean_dec(x_300); +x_302 = lean_array_get_size(x_301); +x_303 = lean_unsigned_to_nat(0u); +x_304 = lean_nat_dec_lt(x_303, x_302); +x_305 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_269, x_271, x_294); +if (x_304 == 0) +{ +uint8_t x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; +lean_dec(x_302); +lean_dec(x_301); +x_306 = 1; +x_307 = lean_box(x_306); +if (lean_is_scalar(x_270)) { + x_308 = lean_alloc_ctor(0, 2, 0); +} else { + x_308 = x_270; +} +lean_ctor_set(x_308, 0, x_307); +lean_ctor_set(x_308, 1, x_305); +x_309 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_309, 0, x_295); +lean_ctor_set(x_309, 1, x_308); +x_310 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_310, 0, x_309); +x_18 = x_310; +x_19 = x_293; +goto block_24; +} +else +{ +uint8_t x_311; +x_311 = lean_nat_dec_le(x_302, x_302); +if (x_311 == 0) +{ +uint8_t x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; +lean_dec(x_302); +lean_dec(x_301); +x_312 = 1; +x_313 = lean_box(x_312); +if (lean_is_scalar(x_270)) { + x_314 = lean_alloc_ctor(0, 2, 0); +} else { + x_314 = x_270; +} +lean_ctor_set(x_314, 0, x_313); +lean_ctor_set(x_314, 1, x_305); +x_315 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_315, 0, x_295); +lean_ctor_set(x_315, 1, x_314); +x_316 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_316, 0, x_315); +x_18 = x_316; +x_19 = x_293; +goto block_24; +} +else +{ +size_t x_317; size_t x_318; lean_object* x_319; uint8_t x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; +x_317 = 0; +x_318 = lean_usize_of_nat(x_302); +lean_dec(x_302); +x_319 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__10___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__11(x_4, x_301, x_317, x_318, x_295); +lean_dec(x_301); +x_320 = 1; +x_321 = lean_box(x_320); +if (lean_is_scalar(x_270)) { + x_322 = lean_alloc_ctor(0, 2, 0); +} else { + x_322 = x_270; +} +lean_ctor_set(x_322, 0, x_321); +lean_ctor_set(x_322, 1, x_305); +x_323 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_323, 0, x_319); +lean_ctor_set(x_323, 1, x_322); +x_324 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_324, 0, x_323); +x_18 = x_324; +x_19 = x_293; +goto block_24; +} +} +} +} +else +{ +lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; +lean_dec(x_274); +lean_dec(x_271); +lean_dec(x_270); +lean_dec(x_269); +lean_dec(x_268); +lean_dec(x_267); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_4); +lean_dec(x_2); +x_325 = lean_ctor_get(x_286, 0); +lean_inc(x_325); +x_326 = lean_ctor_get(x_286, 1); +lean_inc(x_326); +if (lean_is_exclusive(x_286)) { + lean_ctor_release(x_286, 0); + lean_ctor_release(x_286, 1); + x_327 = x_286; +} else { + lean_dec_ref(x_286); + x_327 = lean_box(0); +} +if (lean_is_scalar(x_327)) { + x_328 = lean_alloc_ctor(1, 2, 0); +} else { + x_328 = x_327; +} +lean_ctor_set(x_328, 0, x_325); +lean_ctor_set(x_328, 1, x_326); +return x_328; +} +} +else +{ +lean_object* x_329; lean_object* x_330; lean_object* x_331; +lean_dec(x_285); +lean_dec(x_274); +lean_dec(x_271); +if (lean_is_scalar(x_270)) { + x_329 = lean_alloc_ctor(0, 2, 0); +} else { + x_329 = x_270; +} +lean_ctor_set(x_329, 0, x_268); +lean_ctor_set(x_329, 1, x_269); +x_330 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_330, 0, x_267); +lean_ctor_set(x_330, 1, x_329); +x_331 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_331, 0, x_330); +x_18 = x_331; +x_19 = x_283; +goto block_24; +} +} +else +{ +lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; +lean_dec(x_274); +lean_dec(x_271); +lean_dec(x_270); +lean_dec(x_269); +lean_dec(x_268); +lean_dec(x_267); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_4); +lean_dec(x_2); +x_332 = lean_ctor_get(x_281, 0); +lean_inc(x_332); +x_333 = lean_ctor_get(x_281, 1); +lean_inc(x_333); +if (lean_is_exclusive(x_281)) { + lean_ctor_release(x_281, 0); + lean_ctor_release(x_281, 1); + x_334 = x_281; +} else { + lean_dec_ref(x_281); + x_334 = lean_box(0); +} +if (lean_is_scalar(x_334)) { + x_335 = lean_alloc_ctor(1, 2, 0); +} else { + x_335 = x_334; +} +lean_ctor_set(x_335, 0, x_332); +lean_ctor_set(x_335, 1, x_333); +return x_335; +} +} +else +{ +lean_object* x_336; uint8_t x_337; +x_336 = lean_ctor_get(x_277, 1); +lean_inc(x_336); +lean_dec(x_277); +x_337 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars(x_4, x_267, x_274); +if (x_337 == 0) +{ +lean_object* x_338; lean_object* x_339; lean_object* x_340; +lean_dec(x_271); +if (lean_is_scalar(x_270)) { + x_338 = lean_alloc_ctor(0, 2, 0); +} else { + x_338 = x_270; +} +lean_ctor_set(x_338, 0, x_268); +lean_ctor_set(x_338, 1, x_269); +x_339 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_339, 0, x_267); +lean_ctor_set(x_339, 1, x_338); +x_340 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_340, 0, x_339); +x_18 = x_340; +x_19 = x_336; +goto block_24; +} +else +{ +lean_object* x_341; lean_object* x_342; lean_object* x_343; uint8_t x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; +lean_dec(x_268); +x_341 = lean_box(0); +lean_inc(x_271); +x_342 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_267, x_271, x_341); +x_343 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_269, x_271, x_341); +x_344 = 1; +x_345 = lean_box(x_344); +if (lean_is_scalar(x_270)) { + x_346 = lean_alloc_ctor(0, 2, 0); +} else { + x_346 = x_270; +} +lean_ctor_set(x_346, 0, x_345); +lean_ctor_set(x_346, 1, x_343); +x_347 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_347, 0, x_342); +lean_ctor_set(x_347, 1, x_346); +x_348 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_348, 0, x_347); +x_18 = x_348; +x_19 = x_336; +goto block_24; +} +} +} +else +{ +lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; +lean_dec(x_274); +lean_dec(x_271); +lean_dec(x_270); +lean_dec(x_269); +lean_dec(x_268); +lean_dec(x_267); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_4); +lean_dec(x_2); +x_349 = lean_ctor_get(x_277, 0); +lean_inc(x_349); +x_350 = lean_ctor_get(x_277, 1); +lean_inc(x_350); +if (lean_is_exclusive(x_277)) { + lean_ctor_release(x_277, 0); + lean_ctor_release(x_277, 1); + x_351 = x_277; +} else { + lean_dec_ref(x_277); + x_351 = lean_box(0); +} +if (lean_is_scalar(x_351)) { + x_352 = lean_alloc_ctor(1, 2, 0); +} else { + x_352 = x_351; +} +lean_ctor_set(x_352, 0, x_349); +lean_ctor_set(x_352, 1, x_350); +return x_352; +} +} +else +{ +lean_object* x_353; lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; uint8_t x_362; lean_object* x_363; lean_object* x_364; +lean_dec(x_268); +if (lean_is_exclusive(x_276)) { + lean_ctor_release(x_276, 0); + x_353 = x_276; +} else { + lean_dec_ref(x_276); + x_353 = lean_box(0); +} +x_354 = lean_box(0); +lean_inc(x_2); +x_355 = lean_array_mk(x_2); +x_356 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; +x_357 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_357, 0, x_356); +lean_ctor_set(x_357, 1, x_354); +lean_ctor_set(x_357, 2, x_355); +x_358 = l_Lean_CollectFVars_main(x_274, x_357); +x_359 = lean_ctor_get(x_358, 2); +lean_inc(x_359); +lean_dec(x_358); +x_360 = lean_array_get_size(x_359); +x_361 = lean_unsigned_to_nat(0u); +x_362 = lean_nat_dec_lt(x_361, x_360); +x_363 = lean_box(0); +x_364 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_269, x_271, x_363); +if (x_362 == 0) +{ +uint8_t x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; lean_object* x_369; +lean_dec(x_360); +lean_dec(x_359); +x_365 = 1; +x_366 = lean_box(x_365); +if (lean_is_scalar(x_270)) { + x_367 = lean_alloc_ctor(0, 2, 0); +} else { + x_367 = x_270; +} +lean_ctor_set(x_367, 0, x_366); +lean_ctor_set(x_367, 1, x_364); +x_368 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_368, 0, x_267); +lean_ctor_set(x_368, 1, x_367); +if (lean_is_scalar(x_353)) { + x_369 = lean_alloc_ctor(1, 1, 0); +} else { + x_369 = x_353; +} +lean_ctor_set(x_369, 0, x_368); +x_18 = x_369; +x_19 = x_275; +goto block_24; +} +else +{ +uint8_t x_370; +x_370 = lean_nat_dec_le(x_360, x_360); +if (x_370 == 0) +{ +uint8_t x_371; lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; +lean_dec(x_360); +lean_dec(x_359); +x_371 = 1; +x_372 = lean_box(x_371); +if (lean_is_scalar(x_270)) { + x_373 = lean_alloc_ctor(0, 2, 0); +} else { + x_373 = x_270; +} +lean_ctor_set(x_373, 0, x_372); +lean_ctor_set(x_373, 1, x_364); +x_374 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_374, 0, x_267); +lean_ctor_set(x_374, 1, x_373); +if (lean_is_scalar(x_353)) { + x_375 = lean_alloc_ctor(1, 1, 0); +} else { + x_375 = x_353; +} +lean_ctor_set(x_375, 0, x_374); +x_18 = x_375; +x_19 = x_275; +goto block_24; +} +else +{ +size_t x_376; size_t x_377; lean_object* x_378; uint8_t x_379; lean_object* x_380; lean_object* x_381; lean_object* x_382; lean_object* x_383; +x_376 = 0; +x_377 = lean_usize_of_nat(x_360); +lean_dec(x_360); +x_378 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__12___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__13(x_4, x_359, x_376, x_377, x_267); +lean_dec(x_359); +x_379 = 1; +x_380 = lean_box(x_379); +if (lean_is_scalar(x_270)) { + x_381 = lean_alloc_ctor(0, 2, 0); +} else { + x_381 = x_270; +} +lean_ctor_set(x_381, 0, x_380); +lean_ctor_set(x_381, 1, x_364); +x_382 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_382, 0, x_378); +lean_ctor_set(x_382, 1, x_381); +if (lean_is_scalar(x_353)) { + x_383 = lean_alloc_ctor(1, 1, 0); +} else { + x_383 = x_353; +} +lean_ctor_set(x_383, 0, x_382); +x_18 = x_383; +x_19 = x_275; +goto block_24; +} +} +} +} +else +{ +lean_object* x_384; lean_object* x_385; lean_object* x_386; lean_object* x_387; +lean_dec(x_271); +lean_dec(x_270); +lean_dec(x_269); +lean_dec(x_268); +lean_dec(x_267); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_4); +lean_dec(x_2); +x_384 = lean_ctor_get(x_273, 0); +lean_inc(x_384); +x_385 = lean_ctor_get(x_273, 1); +lean_inc(x_385); +if (lean_is_exclusive(x_273)) { + lean_ctor_release(x_273, 0); + lean_ctor_release(x_273, 1); + x_386 = x_273; +} else { + lean_dec_ref(x_273); + x_386 = lean_box(0); +} +if (lean_is_scalar(x_386)) { + x_387 = lean_alloc_ctor(1, 2, 0); +} else { + x_387 = x_386; +} +lean_ctor_set(x_387, 0, x_384); +lean_ctor_set(x_387, 1, x_385); +return x_387; +} +} +else +{ +lean_object* x_388; lean_object* x_389; lean_object* x_390; lean_object* x_391; +lean_dec(x_271); +lean_dec(x_17); +if (lean_is_exclusive(x_272)) { + lean_ctor_release(x_272, 0); + x_388 = x_272; +} else { + lean_dec_ref(x_272); + x_388 = lean_box(0); +} +if (lean_is_scalar(x_270)) { + x_389 = lean_alloc_ctor(0, 2, 0); +} else { + x_389 = x_270; +} +lean_ctor_set(x_389, 0, x_268); +lean_ctor_set(x_389, 1, x_269); +x_390 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_390, 0, x_267); +lean_ctor_set(x_390, 1, x_389); +if (lean_is_scalar(x_388)) { + x_391 = lean_alloc_ctor(1, 1, 0); +} else { + x_391 = x_388; +} +lean_ctor_set(x_391, 0, x_390); +x_18 = x_391; +x_19 = x_14; +goto block_24; +} +} +block_24: +{ +lean_object* x_20; size_t x_21; size_t x_22; +x_20 = lean_ctor_get(x_18, 0); +lean_inc(x_20); lean_dec(x_18); -x_23 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_20, x_22); -x_24 = lean_array_get_size(x_23); -lean_inc(x_7); +x_21 = 1; +x_22 = lean_usize_add(x_8, x_21); +x_8 = x_22; +x_9 = x_20; +x_14 = x_19; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__14___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__15(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, size_t x_6, size_t x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +uint8_t x_14; +x_14 = lean_usize_dec_lt(x_7, x_6); +if (x_14 == 0) +{ +lean_object* x_15; +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_3); +lean_dec(x_2); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_8); +lean_ctor_set(x_15, 1, x_13); +return x_15; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_24; +x_16 = lean_array_uget(x_5, x_7); +x_24 = !lean_is_exclusive(x_8); +if (x_24 == 0) +{ +lean_object* x_25; uint8_t x_26; +x_25 = lean_ctor_get(x_8, 1); +x_26 = !lean_is_exclusive(x_25); +if (x_26 == 0) +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_27 = lean_ctor_get(x_8, 0); +x_28 = lean_ctor_get(x_25, 0); +x_29 = lean_ctor_get(x_25, 1); +x_30 = l_Lean_Expr_fvarId_x21(x_16); +x_31 = l_Lean_RBNode_findCore___at_Lean_Meta_Closure_process___spec__1(x_29, x_30); +if (lean_obj_tag(x_31) == 0) +{ +lean_object* x_32; +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +x_32 = lean_infer_type(x_16, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_32) == 0) +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_33 = lean_ctor_get(x_32, 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_RBNode_findCore___at_Lean_Meta_Closure_process___spec__1(x_27, x_30); +if (lean_obj_tag(x_35) == 0) +{ +lean_object* x_36; +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_33); +x_36 = l_Lean_Meta_isProp(x_33, x_9, x_10, x_11, x_12, x_34); +if (lean_obj_tag(x_36) == 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_39; lean_object* x_40; +x_39 = lean_ctor_get(x_36, 1); +lean_inc(x_39); +lean_dec(x_36); +lean_inc(x_9); +lean_inc(x_30); +x_40 = l_Lean_FVarId_getDecl(x_30, x_9, x_10, x_11, x_12, x_39); +if (lean_obj_tag(x_40) == 0) +{ +lean_object* x_41; lean_object* x_42; uint8_t 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 = l_Lean_LocalDecl_binderInfo(x_41); +lean_dec(x_41); +x_44 = lean_box(x_43); +if (lean_obj_tag(x_44) == 3) +{ +lean_object* x_45; +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_33); +lean_inc(x_27); +lean_inc(x_3); +x_45 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized(x_3, x_27, x_33, x_9, x_10, x_11, x_12, x_42); +if (lean_obj_tag(x_45) == 0) +{ +lean_object* x_46; uint8_t x_47; +x_46 = lean_ctor_get(x_45, 0); +lean_inc(x_46); +x_47 = lean_unbox(x_46); +lean_dec(x_46); +if (x_47 == 0) +{ +lean_object* x_48; lean_object* x_49; +lean_dec(x_33); +lean_dec(x_30); +x_48 = lean_ctor_get(x_45, 1); +lean_inc(x_48); +lean_dec(x_45); +x_49 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_49, 0, x_8); +x_17 = x_49; +x_18 = x_48; +goto block_23; +} +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; uint8_t x_61; lean_object* x_62; +lean_dec(x_28); +x_50 = lean_ctor_get(x_45, 1); +lean_inc(x_50); +lean_dec(x_45); +x_51 = lean_box(0); +lean_inc(x_30); +x_52 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_27, x_30, x_51); +x_53 = lean_box(0); +lean_inc(x_2); +x_54 = lean_array_mk(x_2); +x_55 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; +x_56 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_56, 0, x_55); +lean_ctor_set(x_56, 1, x_53); +lean_ctor_set(x_56, 2, x_54); +x_57 = l_Lean_CollectFVars_main(x_33, x_56); +x_58 = lean_ctor_get(x_57, 2); +lean_inc(x_58); +lean_dec(x_57); +x_59 = lean_array_get_size(x_58); +x_60 = lean_unsigned_to_nat(0u); +x_61 = lean_nat_dec_lt(x_60, x_59); +x_62 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_29, x_30, x_51); +if (x_61 == 0) +{ +uint8_t x_63; lean_object* x_64; lean_object* x_65; +lean_dec(x_59); +lean_dec(x_58); +x_63 = 1; +x_64 = lean_box(x_63); +lean_ctor_set(x_25, 1, x_62); +lean_ctor_set(x_25, 0, x_64); +lean_ctor_set(x_8, 0, x_52); +x_65 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_65, 0, x_8); +x_17 = x_65; +x_18 = x_50; +goto block_23; +} +else +{ +uint8_t x_66; +x_66 = lean_nat_dec_le(x_59, x_59); +if (x_66 == 0) +{ +uint8_t x_67; lean_object* x_68; lean_object* x_69; +lean_dec(x_59); +lean_dec(x_58); +x_67 = 1; +x_68 = lean_box(x_67); +lean_ctor_set(x_25, 1, x_62); +lean_ctor_set(x_25, 0, x_68); +lean_ctor_set(x_8, 0, x_52); +x_69 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_69, 0, x_8); +x_17 = x_69; +x_18 = x_50; +goto block_23; +} +else +{ +size_t x_70; size_t x_71; lean_object* x_72; uint8_t x_73; lean_object* x_74; lean_object* x_75; +x_70 = 0; +x_71 = lean_usize_of_nat(x_59); +lean_dec(x_59); +x_72 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__10___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__11(x_3, x_58, x_70, x_71, x_52); +lean_dec(x_58); +x_73 = 1; +x_74 = lean_box(x_73); +lean_ctor_set(x_25, 1, x_62); +lean_ctor_set(x_25, 0, x_74); +lean_ctor_set(x_8, 0, x_72); +x_75 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_75, 0, x_8); +x_17 = x_75; +x_18 = x_50; +goto block_23; +} +} +} +} +else +{ +uint8_t x_76; +lean_dec(x_33); +lean_dec(x_30); +lean_free_object(x_25); +lean_dec(x_29); +lean_dec(x_28); +lean_free_object(x_8); +lean_dec(x_27); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_3); +lean_dec(x_2); +x_76 = !lean_is_exclusive(x_45); +if (x_76 == 0) +{ +return x_45; +} +else +{ +lean_object* x_77; lean_object* x_78; lean_object* x_79; +x_77 = lean_ctor_get(x_45, 0); +x_78 = lean_ctor_get(x_45, 1); +lean_inc(x_78); +lean_inc(x_77); +lean_dec(x_45); +x_79 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_79, 0, x_77); +lean_ctor_set(x_79, 1, x_78); +return x_79; +} +} +} +else +{ +lean_object* x_80; +lean_dec(x_44); +lean_dec(x_33); +lean_dec(x_30); +x_80 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_80, 0, x_8); +x_17 = x_80; +x_18 = x_42; +goto block_23; +} +} +else +{ +uint8_t x_81; +lean_dec(x_33); +lean_dec(x_30); +lean_free_object(x_25); +lean_dec(x_29); +lean_dec(x_28); +lean_free_object(x_8); +lean_dec(x_27); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_3); +lean_dec(x_2); +x_81 = !lean_is_exclusive(x_40); +if (x_81 == 0) +{ +return x_40; +} +else +{ +lean_object* x_82; lean_object* x_83; lean_object* x_84; +x_82 = lean_ctor_get(x_40, 0); +x_83 = lean_ctor_get(x_40, 1); +lean_inc(x_83); +lean_inc(x_82); +lean_dec(x_40); +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_85; uint8_t x_86; +x_85 = lean_ctor_get(x_36, 1); +lean_inc(x_85); +lean_dec(x_36); +x_86 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars(x_3, x_27, x_33); +if (x_86 == 0) +{ +lean_object* x_87; +lean_dec(x_30); +x_87 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_87, 0, x_8); +x_17 = x_87; +x_18 = x_85; +goto block_23; +} +else +{ +lean_object* x_88; lean_object* x_89; lean_object* x_90; uint8_t x_91; lean_object* x_92; lean_object* x_93; +lean_dec(x_28); +x_88 = lean_box(0); +lean_inc(x_30); +x_89 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_27, x_30, x_88); +x_90 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_29, x_30, x_88); +x_91 = 1; +x_92 = lean_box(x_91); +lean_ctor_set(x_25, 1, x_90); +lean_ctor_set(x_25, 0, x_92); +lean_ctor_set(x_8, 0, x_89); +x_93 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_93, 0, x_8); +x_17 = x_93; +x_18 = x_85; +goto block_23; +} +} +} +else +{ +uint8_t x_94; +lean_dec(x_33); +lean_dec(x_30); +lean_free_object(x_25); +lean_dec(x_29); +lean_dec(x_28); +lean_free_object(x_8); +lean_dec(x_27); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_3); +lean_dec(x_2); +x_94 = !lean_is_exclusive(x_36); +if (x_94 == 0) +{ +return x_36; +} +else +{ +lean_object* x_95; lean_object* x_96; lean_object* x_97; +x_95 = lean_ctor_get(x_36, 0); +x_96 = lean_ctor_get(x_36, 1); +lean_inc(x_96); +lean_inc(x_95); +lean_dec(x_36); +x_97 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_97, 0, x_95); +lean_ctor_set(x_97, 1, x_96); +return x_97; +} +} +} +else +{ +uint8_t x_98; +lean_dec(x_28); +x_98 = !lean_is_exclusive(x_35); +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; uint8_t x_108; lean_object* x_109; lean_object* x_110; +x_99 = lean_ctor_get(x_35, 0); +lean_dec(x_99); +x_100 = lean_box(0); +lean_inc(x_2); +x_101 = lean_array_mk(x_2); +x_102 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; +x_103 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_103, 0, x_102); +lean_ctor_set(x_103, 1, x_100); +lean_ctor_set(x_103, 2, x_101); +x_104 = l_Lean_CollectFVars_main(x_33, x_103); +x_105 = lean_ctor_get(x_104, 2); +lean_inc(x_105); +lean_dec(x_104); +x_106 = lean_array_get_size(x_105); +x_107 = lean_unsigned_to_nat(0u); +x_108 = lean_nat_dec_lt(x_107, x_106); +x_109 = lean_box(0); +x_110 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_29, x_30, x_109); +if (x_108 == 0) +{ +uint8_t x_111; lean_object* x_112; +lean_dec(x_106); +lean_dec(x_105); +x_111 = 1; +x_112 = lean_box(x_111); +lean_ctor_set(x_25, 1, x_110); +lean_ctor_set(x_25, 0, x_112); +lean_ctor_set(x_35, 0, x_8); +x_17 = x_35; +x_18 = x_34; +goto block_23; +} +else +{ +uint8_t x_113; +x_113 = lean_nat_dec_le(x_106, x_106); +if (x_113 == 0) +{ +uint8_t x_114; lean_object* x_115; +lean_dec(x_106); +lean_dec(x_105); +x_114 = 1; +x_115 = lean_box(x_114); +lean_ctor_set(x_25, 1, x_110); +lean_ctor_set(x_25, 0, x_115); +lean_ctor_set(x_35, 0, x_8); +x_17 = x_35; +x_18 = x_34; +goto block_23; +} +else +{ +size_t x_116; size_t x_117; lean_object* x_118; uint8_t x_119; lean_object* x_120; +x_116 = 0; +x_117 = lean_usize_of_nat(x_106); +lean_dec(x_106); +x_118 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__12___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__13(x_3, x_105, x_116, x_117, x_27); +lean_dec(x_105); +x_119 = 1; +x_120 = lean_box(x_119); +lean_ctor_set(x_25, 1, x_110); +lean_ctor_set(x_25, 0, x_120); +lean_ctor_set(x_8, 0, x_118); +lean_ctor_set(x_35, 0, x_8); +x_17 = x_35; +x_18 = x_34; +goto block_23; +} +} +} +else +{ +lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; uint8_t x_129; lean_object* x_130; lean_object* x_131; +lean_dec(x_35); +x_121 = lean_box(0); +lean_inc(x_2); +x_122 = lean_array_mk(x_2); +x_123 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; +x_124 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_124, 0, x_123); +lean_ctor_set(x_124, 1, x_121); +lean_ctor_set(x_124, 2, x_122); +x_125 = l_Lean_CollectFVars_main(x_33, x_124); +x_126 = lean_ctor_get(x_125, 2); +lean_inc(x_126); +lean_dec(x_125); +x_127 = lean_array_get_size(x_126); +x_128 = lean_unsigned_to_nat(0u); +x_129 = lean_nat_dec_lt(x_128, x_127); +x_130 = lean_box(0); +x_131 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_29, x_30, x_130); +if (x_129 == 0) +{ +uint8_t x_132; lean_object* x_133; lean_object* x_134; +lean_dec(x_127); +lean_dec(x_126); +x_132 = 1; +x_133 = lean_box(x_132); +lean_ctor_set(x_25, 1, x_131); +lean_ctor_set(x_25, 0, x_133); +x_134 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_134, 0, x_8); +x_17 = x_134; +x_18 = x_34; +goto block_23; +} +else +{ +uint8_t x_135; +x_135 = lean_nat_dec_le(x_127, x_127); +if (x_135 == 0) +{ +uint8_t x_136; lean_object* x_137; lean_object* x_138; +lean_dec(x_127); +lean_dec(x_126); +x_136 = 1; +x_137 = lean_box(x_136); +lean_ctor_set(x_25, 1, x_131); +lean_ctor_set(x_25, 0, x_137); +x_138 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_138, 0, x_8); +x_17 = x_138; +x_18 = x_34; +goto block_23; +} +else +{ +size_t x_139; size_t x_140; lean_object* x_141; uint8_t x_142; lean_object* x_143; lean_object* x_144; +x_139 = 0; +x_140 = lean_usize_of_nat(x_127); +lean_dec(x_127); +x_141 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__12___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__13(x_3, x_126, x_139, x_140, x_27); +lean_dec(x_126); +x_142 = 1; +x_143 = lean_box(x_142); +lean_ctor_set(x_25, 1, x_131); +lean_ctor_set(x_25, 0, x_143); +lean_ctor_set(x_8, 0, x_141); +x_144 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_144, 0, x_8); +x_17 = x_144; +x_18 = x_34; +goto block_23; +} +} +} +} +} +else +{ +uint8_t x_145; +lean_dec(x_30); +lean_free_object(x_25); +lean_dec(x_29); +lean_dec(x_28); +lean_free_object(x_8); +lean_dec(x_27); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_3); +lean_dec(x_2); +x_145 = !lean_is_exclusive(x_32); +if (x_145 == 0) +{ +return x_32; +} +else +{ +lean_object* x_146; lean_object* x_147; lean_object* x_148; +x_146 = lean_ctor_get(x_32, 0); +x_147 = lean_ctor_get(x_32, 1); +lean_inc(x_147); +lean_inc(x_146); +lean_dec(x_32); +x_148 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_148, 0, x_146); +lean_ctor_set(x_148, 1, x_147); +return x_148; +} +} +} +else +{ +uint8_t x_149; +lean_dec(x_30); +lean_dec(x_16); +x_149 = !lean_is_exclusive(x_31); +if (x_149 == 0) +{ +lean_object* x_150; +x_150 = lean_ctor_get(x_31, 0); +lean_dec(x_150); +lean_ctor_set(x_31, 0, x_8); +x_17 = x_31; +x_18 = x_13; +goto block_23; +} +else +{ +lean_object* x_151; +lean_dec(x_31); +x_151 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_151, 0, x_8); +x_17 = x_151; +x_18 = x_13; +goto block_23; +} +} +} +else +{ +lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; +x_152 = lean_ctor_get(x_8, 0); +x_153 = lean_ctor_get(x_25, 0); +x_154 = lean_ctor_get(x_25, 1); +lean_inc(x_154); +lean_inc(x_153); +lean_dec(x_25); +x_155 = l_Lean_Expr_fvarId_x21(x_16); +x_156 = l_Lean_RBNode_findCore___at_Lean_Meta_Closure_process___spec__1(x_154, x_155); +if (lean_obj_tag(x_156) == 0) +{ +lean_object* x_157; +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +x_157 = lean_infer_type(x_16, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_157) == 0) +{ +lean_object* x_158; lean_object* x_159; lean_object* x_160; +x_158 = lean_ctor_get(x_157, 0); +lean_inc(x_158); +x_159 = lean_ctor_get(x_157, 1); +lean_inc(x_159); +lean_dec(x_157); +x_160 = l_Lean_RBNode_findCore___at_Lean_Meta_Closure_process___spec__1(x_152, x_155); +if (lean_obj_tag(x_160) == 0) +{ +lean_object* x_161; +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_158); +x_161 = l_Lean_Meta_isProp(x_158, x_9, x_10, x_11, x_12, x_159); +if (lean_obj_tag(x_161) == 0) +{ +lean_object* x_162; uint8_t x_163; +x_162 = lean_ctor_get(x_161, 0); +lean_inc(x_162); +x_163 = lean_unbox(x_162); +lean_dec(x_162); +if (x_163 == 0) +{ +lean_object* x_164; lean_object* x_165; +x_164 = lean_ctor_get(x_161, 1); +lean_inc(x_164); +lean_dec(x_161); +lean_inc(x_9); +lean_inc(x_155); +x_165 = l_Lean_FVarId_getDecl(x_155, x_9, x_10, x_11, x_12, x_164); +if (lean_obj_tag(x_165) == 0) +{ +lean_object* x_166; lean_object* x_167; uint8_t x_168; lean_object* x_169; +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_Lean_LocalDecl_binderInfo(x_166); +lean_dec(x_166); +x_169 = lean_box(x_168); +if (lean_obj_tag(x_169) == 3) +{ +lean_object* x_170; +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_158); +lean_inc(x_152); +lean_inc(x_3); +x_170 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized(x_3, x_152, x_158, x_9, x_10, x_11, x_12, x_167); +if (lean_obj_tag(x_170) == 0) +{ +lean_object* x_171; uint8_t x_172; +x_171 = lean_ctor_get(x_170, 0); +lean_inc(x_171); +x_172 = lean_unbox(x_171); +lean_dec(x_171); +if (x_172 == 0) +{ +lean_object* x_173; lean_object* x_174; lean_object* x_175; +lean_dec(x_158); +lean_dec(x_155); +x_173 = lean_ctor_get(x_170, 1); +lean_inc(x_173); +lean_dec(x_170); +x_174 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_174, 0, x_153); +lean_ctor_set(x_174, 1, x_154); +lean_ctor_set(x_8, 1, x_174); +x_175 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_175, 0, x_8); +x_17 = x_175; +x_18 = x_173; +goto block_23; +} +else +{ +lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; uint8_t x_187; lean_object* x_188; +lean_dec(x_153); +x_176 = lean_ctor_get(x_170, 1); +lean_inc(x_176); +lean_dec(x_170); +x_177 = lean_box(0); +lean_inc(x_155); +x_178 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_152, x_155, x_177); +x_179 = lean_box(0); +lean_inc(x_2); +x_180 = lean_array_mk(x_2); +x_181 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; +x_182 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_182, 0, x_181); +lean_ctor_set(x_182, 1, x_179); +lean_ctor_set(x_182, 2, x_180); +x_183 = l_Lean_CollectFVars_main(x_158, x_182); +x_184 = lean_ctor_get(x_183, 2); +lean_inc(x_184); +lean_dec(x_183); +x_185 = lean_array_get_size(x_184); +x_186 = lean_unsigned_to_nat(0u); +x_187 = lean_nat_dec_lt(x_186, x_185); +x_188 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_154, x_155, x_177); +if (x_187 == 0) +{ +uint8_t x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; +lean_dec(x_185); +lean_dec(x_184); +x_189 = 1; +x_190 = lean_box(x_189); +x_191 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_191, 0, x_190); +lean_ctor_set(x_191, 1, x_188); +lean_ctor_set(x_8, 1, x_191); +lean_ctor_set(x_8, 0, x_178); +x_192 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_192, 0, x_8); +x_17 = x_192; +x_18 = x_176; +goto block_23; +} +else +{ +uint8_t x_193; +x_193 = lean_nat_dec_le(x_185, x_185); +if (x_193 == 0) +{ +uint8_t x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; +lean_dec(x_185); +lean_dec(x_184); +x_194 = 1; +x_195 = lean_box(x_194); +x_196 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_196, 0, x_195); +lean_ctor_set(x_196, 1, x_188); +lean_ctor_set(x_8, 1, x_196); +lean_ctor_set(x_8, 0, x_178); +x_197 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_197, 0, x_8); +x_17 = x_197; +x_18 = x_176; +goto block_23; +} +else +{ +size_t x_198; size_t x_199; lean_object* x_200; uint8_t x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; +x_198 = 0; +x_199 = lean_usize_of_nat(x_185); +lean_dec(x_185); +x_200 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__10___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__11(x_3, x_184, x_198, x_199, x_178); +lean_dec(x_184); +x_201 = 1; +x_202 = lean_box(x_201); +x_203 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_203, 0, x_202); +lean_ctor_set(x_203, 1, x_188); +lean_ctor_set(x_8, 1, x_203); +lean_ctor_set(x_8, 0, x_200); +x_204 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_204, 0, x_8); +x_17 = x_204; +x_18 = x_176; +goto block_23; +} +} +} +} +else +{ +lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; +lean_dec(x_158); +lean_dec(x_155); +lean_dec(x_154); +lean_dec(x_153); +lean_free_object(x_8); +lean_dec(x_152); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_3); +lean_dec(x_2); +x_205 = lean_ctor_get(x_170, 0); +lean_inc(x_205); +x_206 = lean_ctor_get(x_170, 1); +lean_inc(x_206); +if (lean_is_exclusive(x_170)) { + lean_ctor_release(x_170, 0); + lean_ctor_release(x_170, 1); + x_207 = x_170; +} else { + lean_dec_ref(x_170); + x_207 = lean_box(0); +} +if (lean_is_scalar(x_207)) { + x_208 = lean_alloc_ctor(1, 2, 0); +} else { + x_208 = x_207; +} +lean_ctor_set(x_208, 0, x_205); +lean_ctor_set(x_208, 1, x_206); +return x_208; +} +} +else +{ +lean_object* x_209; lean_object* x_210; +lean_dec(x_169); +lean_dec(x_158); +lean_dec(x_155); +x_209 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_209, 0, x_153); +lean_ctor_set(x_209, 1, x_154); +lean_ctor_set(x_8, 1, x_209); +x_210 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_210, 0, x_8); +x_17 = x_210; +x_18 = x_167; +goto block_23; +} +} +else +{ +lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; +lean_dec(x_158); +lean_dec(x_155); +lean_dec(x_154); +lean_dec(x_153); +lean_free_object(x_8); +lean_dec(x_152); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_3); +lean_dec(x_2); +x_211 = lean_ctor_get(x_165, 0); +lean_inc(x_211); +x_212 = lean_ctor_get(x_165, 1); +lean_inc(x_212); +if (lean_is_exclusive(x_165)) { + lean_ctor_release(x_165, 0); + lean_ctor_release(x_165, 1); + x_213 = x_165; +} else { + lean_dec_ref(x_165); + 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_215; uint8_t x_216; +x_215 = lean_ctor_get(x_161, 1); +lean_inc(x_215); +lean_dec(x_161); +x_216 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars(x_3, x_152, x_158); +if (x_216 == 0) +{ +lean_object* x_217; lean_object* x_218; +lean_dec(x_155); +x_217 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_217, 0, x_153); +lean_ctor_set(x_217, 1, x_154); +lean_ctor_set(x_8, 1, x_217); +x_218 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_218, 0, x_8); +x_17 = x_218; +x_18 = x_215; +goto block_23; +} +else +{ +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; +lean_dec(x_153); +x_219 = lean_box(0); +lean_inc(x_155); +x_220 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_152, x_155, x_219); +x_221 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_154, x_155, x_219); +x_222 = 1; +x_223 = lean_box(x_222); +x_224 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_224, 0, x_223); +lean_ctor_set(x_224, 1, x_221); +lean_ctor_set(x_8, 1, x_224); +lean_ctor_set(x_8, 0, x_220); +x_225 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_225, 0, x_8); +x_17 = x_225; +x_18 = x_215; +goto block_23; +} +} +} +else +{ +lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; +lean_dec(x_158); +lean_dec(x_155); +lean_dec(x_154); +lean_dec(x_153); +lean_free_object(x_8); +lean_dec(x_152); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_3); +lean_dec(x_2); +x_226 = lean_ctor_get(x_161, 0); +lean_inc(x_226); +x_227 = lean_ctor_get(x_161, 1); +lean_inc(x_227); +if (lean_is_exclusive(x_161)) { + lean_ctor_release(x_161, 0); + lean_ctor_release(x_161, 1); + x_228 = x_161; +} else { + lean_dec_ref(x_161); + x_228 = lean_box(0); +} +if (lean_is_scalar(x_228)) { + x_229 = lean_alloc_ctor(1, 2, 0); +} else { + x_229 = x_228; +} +lean_ctor_set(x_229, 0, x_226); +lean_ctor_set(x_229, 1, x_227); +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; uint8_t x_239; lean_object* x_240; lean_object* x_241; +lean_dec(x_153); +if (lean_is_exclusive(x_160)) { + lean_ctor_release(x_160, 0); + x_230 = x_160; +} else { + lean_dec_ref(x_160); + x_230 = lean_box(0); +} +x_231 = lean_box(0); +lean_inc(x_2); +x_232 = lean_array_mk(x_2); +x_233 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; +x_234 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_234, 0, x_233); +lean_ctor_set(x_234, 1, x_231); +lean_ctor_set(x_234, 2, x_232); +x_235 = l_Lean_CollectFVars_main(x_158, x_234); +x_236 = lean_ctor_get(x_235, 2); +lean_inc(x_236); +lean_dec(x_235); +x_237 = lean_array_get_size(x_236); +x_238 = lean_unsigned_to_nat(0u); +x_239 = lean_nat_dec_lt(x_238, x_237); +x_240 = lean_box(0); +x_241 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_154, x_155, x_240); +if (x_239 == 0) +{ +uint8_t x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; +lean_dec(x_237); +lean_dec(x_236); +x_242 = 1; +x_243 = lean_box(x_242); +x_244 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_244, 0, x_243); +lean_ctor_set(x_244, 1, x_241); +lean_ctor_set(x_8, 1, x_244); +if (lean_is_scalar(x_230)) { + x_245 = lean_alloc_ctor(1, 1, 0); +} else { + x_245 = x_230; +} +lean_ctor_set(x_245, 0, x_8); +x_17 = x_245; +x_18 = x_159; +goto block_23; +} +else +{ +uint8_t x_246; +x_246 = lean_nat_dec_le(x_237, x_237); +if (x_246 == 0) +{ +uint8_t x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; +lean_dec(x_237); +lean_dec(x_236); +x_247 = 1; +x_248 = lean_box(x_247); +x_249 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_249, 0, x_248); +lean_ctor_set(x_249, 1, x_241); +lean_ctor_set(x_8, 1, x_249); +if (lean_is_scalar(x_230)) { + x_250 = lean_alloc_ctor(1, 1, 0); +} else { + x_250 = x_230; +} +lean_ctor_set(x_250, 0, x_8); +x_17 = x_250; +x_18 = x_159; +goto block_23; +} +else +{ +size_t x_251; size_t x_252; lean_object* x_253; uint8_t x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; +x_251 = 0; +x_252 = lean_usize_of_nat(x_237); +lean_dec(x_237); +x_253 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__12___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__13(x_3, x_236, x_251, x_252, x_152); +lean_dec(x_236); +x_254 = 1; +x_255 = lean_box(x_254); +x_256 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_256, 0, x_255); +lean_ctor_set(x_256, 1, x_241); +lean_ctor_set(x_8, 1, x_256); +lean_ctor_set(x_8, 0, x_253); +if (lean_is_scalar(x_230)) { + x_257 = lean_alloc_ctor(1, 1, 0); +} else { + x_257 = x_230; +} +lean_ctor_set(x_257, 0, x_8); +x_17 = x_257; +x_18 = x_159; +goto block_23; +} +} +} +} +else +{ +lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; +lean_dec(x_155); +lean_dec(x_154); +lean_dec(x_153); +lean_free_object(x_8); +lean_dec(x_152); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_3); +lean_dec(x_2); +x_258 = lean_ctor_get(x_157, 0); +lean_inc(x_258); +x_259 = lean_ctor_get(x_157, 1); +lean_inc(x_259); +if (lean_is_exclusive(x_157)) { + lean_ctor_release(x_157, 0); + lean_ctor_release(x_157, 1); + x_260 = x_157; +} else { + lean_dec_ref(x_157); + x_260 = lean_box(0); +} +if (lean_is_scalar(x_260)) { + x_261 = lean_alloc_ctor(1, 2, 0); +} else { + x_261 = x_260; +} +lean_ctor_set(x_261, 0, x_258); +lean_ctor_set(x_261, 1, x_259); +return x_261; +} +} +else +{ +lean_object* x_262; lean_object* x_263; lean_object* x_264; +lean_dec(x_155); +lean_dec(x_16); +if (lean_is_exclusive(x_156)) { + lean_ctor_release(x_156, 0); + x_262 = x_156; +} else { + lean_dec_ref(x_156); + x_262 = lean_box(0); +} +x_263 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_263, 0, x_153); +lean_ctor_set(x_263, 1, x_154); +lean_ctor_set(x_8, 1, x_263); +if (lean_is_scalar(x_262)) { + x_264 = lean_alloc_ctor(1, 1, 0); +} else { + x_264 = x_262; +} +lean_ctor_set(x_264, 0, x_8); +x_17 = x_264; +x_18 = x_13; +goto block_23; +} +} +} +else +{ +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_265 = lean_ctor_get(x_8, 1); +x_266 = lean_ctor_get(x_8, 0); +lean_inc(x_265); +lean_inc(x_266); +lean_dec(x_8); +x_267 = lean_ctor_get(x_265, 0); +lean_inc(x_267); +x_268 = lean_ctor_get(x_265, 1); +lean_inc(x_268); +if (lean_is_exclusive(x_265)) { + lean_ctor_release(x_265, 0); + lean_ctor_release(x_265, 1); + x_269 = x_265; +} else { + lean_dec_ref(x_265); + x_269 = lean_box(0); +} +x_270 = l_Lean_Expr_fvarId_x21(x_16); +x_271 = l_Lean_RBNode_findCore___at_Lean_Meta_Closure_process___spec__1(x_268, x_270); +if (lean_obj_tag(x_271) == 0) +{ +lean_object* x_272; +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +x_272 = lean_infer_type(x_16, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_272) == 0) +{ +lean_object* x_273; lean_object* x_274; lean_object* x_275; +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_275 = l_Lean_RBNode_findCore___at_Lean_Meta_Closure_process___spec__1(x_266, x_270); +if (lean_obj_tag(x_275) == 0) +{ +lean_object* x_276; +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_273); +x_276 = l_Lean_Meta_isProp(x_273, x_9, x_10, x_11, x_12, x_274); +if (lean_obj_tag(x_276) == 0) +{ +lean_object* x_277; uint8_t x_278; +x_277 = lean_ctor_get(x_276, 0); +lean_inc(x_277); +x_278 = lean_unbox(x_277); +lean_dec(x_277); +if (x_278 == 0) +{ +lean_object* x_279; lean_object* x_280; +x_279 = lean_ctor_get(x_276, 1); +lean_inc(x_279); +lean_dec(x_276); +lean_inc(x_9); +lean_inc(x_270); +x_280 = l_Lean_FVarId_getDecl(x_270, x_9, x_10, x_11, x_12, x_279); +if (lean_obj_tag(x_280) == 0) +{ +lean_object* x_281; lean_object* x_282; uint8_t x_283; lean_object* x_284; +x_281 = lean_ctor_get(x_280, 0); +lean_inc(x_281); +x_282 = lean_ctor_get(x_280, 1); +lean_inc(x_282); +lean_dec(x_280); +x_283 = l_Lean_LocalDecl_binderInfo(x_281); +lean_dec(x_281); +x_284 = lean_box(x_283); +if (lean_obj_tag(x_284) == 3) +{ +lean_object* x_285; +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_273); +lean_inc(x_266); +lean_inc(x_3); +x_285 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized(x_3, x_266, x_273, x_9, x_10, x_11, x_12, x_282); +if (lean_obj_tag(x_285) == 0) +{ +lean_object* x_286; uint8_t x_287; +x_286 = lean_ctor_get(x_285, 0); +lean_inc(x_286); +x_287 = lean_unbox(x_286); +lean_dec(x_286); +if (x_287 == 0) +{ +lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; +lean_dec(x_273); +lean_dec(x_270); +x_288 = lean_ctor_get(x_285, 1); +lean_inc(x_288); +lean_dec(x_285); +if (lean_is_scalar(x_269)) { + x_289 = lean_alloc_ctor(0, 2, 0); +} else { + x_289 = x_269; +} +lean_ctor_set(x_289, 0, x_267); +lean_ctor_set(x_289, 1, x_268); +x_290 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_290, 0, x_266); +lean_ctor_set(x_290, 1, x_289); +x_291 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_291, 0, x_290); +x_17 = x_291; +x_18 = x_288; +goto block_23; +} +else +{ +lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; uint8_t x_303; lean_object* x_304; +lean_dec(x_267); +x_292 = lean_ctor_get(x_285, 1); +lean_inc(x_292); +lean_dec(x_285); +x_293 = lean_box(0); +lean_inc(x_270); +x_294 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_266, x_270, x_293); +x_295 = lean_box(0); +lean_inc(x_2); +x_296 = lean_array_mk(x_2); +x_297 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; +x_298 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_298, 0, x_297); +lean_ctor_set(x_298, 1, x_295); +lean_ctor_set(x_298, 2, x_296); +x_299 = l_Lean_CollectFVars_main(x_273, x_298); +x_300 = lean_ctor_get(x_299, 2); +lean_inc(x_300); +lean_dec(x_299); +x_301 = lean_array_get_size(x_300); +x_302 = lean_unsigned_to_nat(0u); +x_303 = lean_nat_dec_lt(x_302, x_301); +x_304 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_268, x_270, x_293); +if (x_303 == 0) +{ +uint8_t x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; +lean_dec(x_301); +lean_dec(x_300); +x_305 = 1; +x_306 = lean_box(x_305); +if (lean_is_scalar(x_269)) { + x_307 = lean_alloc_ctor(0, 2, 0); +} else { + x_307 = x_269; +} +lean_ctor_set(x_307, 0, x_306); +lean_ctor_set(x_307, 1, x_304); +x_308 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_308, 0, x_294); +lean_ctor_set(x_308, 1, x_307); +x_309 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_309, 0, x_308); +x_17 = x_309; +x_18 = x_292; +goto block_23; +} +else +{ +uint8_t x_310; +x_310 = lean_nat_dec_le(x_301, x_301); +if (x_310 == 0) +{ +uint8_t x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; +lean_dec(x_301); +lean_dec(x_300); +x_311 = 1; +x_312 = lean_box(x_311); +if (lean_is_scalar(x_269)) { + x_313 = lean_alloc_ctor(0, 2, 0); +} else { + x_313 = x_269; +} +lean_ctor_set(x_313, 0, x_312); +lean_ctor_set(x_313, 1, x_304); +x_314 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_314, 0, x_294); +lean_ctor_set(x_314, 1, x_313); +x_315 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_315, 0, x_314); +x_17 = x_315; +x_18 = x_292; +goto block_23; +} +else +{ +size_t x_316; size_t x_317; lean_object* x_318; uint8_t x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; +x_316 = 0; +x_317 = lean_usize_of_nat(x_301); +lean_dec(x_301); +x_318 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__10___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__11(x_3, x_300, x_316, x_317, x_294); +lean_dec(x_300); +x_319 = 1; +x_320 = lean_box(x_319); +if (lean_is_scalar(x_269)) { + x_321 = lean_alloc_ctor(0, 2, 0); +} else { + x_321 = x_269; +} +lean_ctor_set(x_321, 0, x_320); +lean_ctor_set(x_321, 1, x_304); +x_322 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_322, 0, x_318); +lean_ctor_set(x_322, 1, x_321); +x_323 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_323, 0, x_322); +x_17 = x_323; +x_18 = x_292; +goto block_23; +} +} +} +} +else +{ +lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; +lean_dec(x_273); +lean_dec(x_270); +lean_dec(x_269); +lean_dec(x_268); +lean_dec(x_267); +lean_dec(x_266); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_3); +lean_dec(x_2); +x_324 = lean_ctor_get(x_285, 0); +lean_inc(x_324); +x_325 = lean_ctor_get(x_285, 1); +lean_inc(x_325); +if (lean_is_exclusive(x_285)) { + lean_ctor_release(x_285, 0); + lean_ctor_release(x_285, 1); + x_326 = x_285; +} else { + lean_dec_ref(x_285); + x_326 = lean_box(0); +} +if (lean_is_scalar(x_326)) { + x_327 = lean_alloc_ctor(1, 2, 0); +} else { + x_327 = x_326; +} +lean_ctor_set(x_327, 0, x_324); +lean_ctor_set(x_327, 1, x_325); +return x_327; +} +} +else +{ +lean_object* x_328; lean_object* x_329; lean_object* x_330; +lean_dec(x_284); +lean_dec(x_273); +lean_dec(x_270); +if (lean_is_scalar(x_269)) { + x_328 = lean_alloc_ctor(0, 2, 0); +} else { + x_328 = x_269; +} +lean_ctor_set(x_328, 0, x_267); +lean_ctor_set(x_328, 1, x_268); +x_329 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_329, 0, x_266); +lean_ctor_set(x_329, 1, x_328); +x_330 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_330, 0, x_329); +x_17 = x_330; +x_18 = x_282; +goto block_23; +} +} +else +{ +lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; +lean_dec(x_273); +lean_dec(x_270); +lean_dec(x_269); +lean_dec(x_268); +lean_dec(x_267); +lean_dec(x_266); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_3); +lean_dec(x_2); +x_331 = lean_ctor_get(x_280, 0); +lean_inc(x_331); +x_332 = lean_ctor_get(x_280, 1); +lean_inc(x_332); +if (lean_is_exclusive(x_280)) { + lean_ctor_release(x_280, 0); + lean_ctor_release(x_280, 1); + x_333 = x_280; +} else { + lean_dec_ref(x_280); + x_333 = lean_box(0); +} +if (lean_is_scalar(x_333)) { + x_334 = lean_alloc_ctor(1, 2, 0); +} else { + x_334 = 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_335; uint8_t x_336; +x_335 = lean_ctor_get(x_276, 1); +lean_inc(x_335); +lean_dec(x_276); +x_336 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars(x_3, x_266, x_273); +if (x_336 == 0) +{ +lean_object* x_337; lean_object* x_338; lean_object* x_339; +lean_dec(x_270); +if (lean_is_scalar(x_269)) { + x_337 = lean_alloc_ctor(0, 2, 0); +} else { + x_337 = x_269; +} +lean_ctor_set(x_337, 0, x_267); +lean_ctor_set(x_337, 1, x_268); +x_338 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_338, 0, x_266); +lean_ctor_set(x_338, 1, x_337); +x_339 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_339, 0, x_338); +x_17 = x_339; +x_18 = x_335; +goto block_23; +} +else +{ +lean_object* x_340; lean_object* x_341; lean_object* x_342; uint8_t x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; +lean_dec(x_267); +x_340 = lean_box(0); +lean_inc(x_270); +x_341 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_266, x_270, x_340); +x_342 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_268, x_270, x_340); +x_343 = 1; +x_344 = lean_box(x_343); +if (lean_is_scalar(x_269)) { + x_345 = lean_alloc_ctor(0, 2, 0); +} else { + x_345 = x_269; +} +lean_ctor_set(x_345, 0, x_344); +lean_ctor_set(x_345, 1, x_342); +x_346 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_346, 0, x_341); +lean_ctor_set(x_346, 1, x_345); +x_347 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_347, 0, x_346); +x_17 = x_347; +x_18 = x_335; +goto block_23; +} +} +} +else +{ +lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; +lean_dec(x_273); +lean_dec(x_270); +lean_dec(x_269); +lean_dec(x_268); +lean_dec(x_267); +lean_dec(x_266); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_3); +lean_dec(x_2); +x_348 = lean_ctor_get(x_276, 0); +lean_inc(x_348); +x_349 = lean_ctor_get(x_276, 1); +lean_inc(x_349); +if (lean_is_exclusive(x_276)) { + lean_ctor_release(x_276, 0); + lean_ctor_release(x_276, 1); + x_350 = x_276; +} else { + lean_dec_ref(x_276); + x_350 = lean_box(0); +} +if (lean_is_scalar(x_350)) { + x_351 = lean_alloc_ctor(1, 2, 0); +} else { + x_351 = x_350; +} +lean_ctor_set(x_351, 0, x_348); +lean_ctor_set(x_351, 1, x_349); +return x_351; +} +} +else +{ +lean_object* x_352; lean_object* x_353; lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; uint8_t x_361; lean_object* x_362; lean_object* x_363; +lean_dec(x_267); +if (lean_is_exclusive(x_275)) { + lean_ctor_release(x_275, 0); + x_352 = x_275; +} else { + lean_dec_ref(x_275); + x_352 = lean_box(0); +} +x_353 = lean_box(0); +lean_inc(x_2); +x_354 = lean_array_mk(x_2); +x_355 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; +x_356 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_356, 0, x_355); +lean_ctor_set(x_356, 1, x_353); +lean_ctor_set(x_356, 2, x_354); +x_357 = l_Lean_CollectFVars_main(x_273, x_356); +x_358 = lean_ctor_get(x_357, 2); +lean_inc(x_358); +lean_dec(x_357); +x_359 = lean_array_get_size(x_358); +x_360 = lean_unsigned_to_nat(0u); +x_361 = lean_nat_dec_lt(x_360, x_359); +x_362 = lean_box(0); +x_363 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_268, x_270, x_362); +if (x_361 == 0) +{ +uint8_t x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; +lean_dec(x_359); +lean_dec(x_358); +x_364 = 1; +x_365 = lean_box(x_364); +if (lean_is_scalar(x_269)) { + x_366 = lean_alloc_ctor(0, 2, 0); +} else { + x_366 = x_269; +} +lean_ctor_set(x_366, 0, x_365); +lean_ctor_set(x_366, 1, x_363); +x_367 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_367, 0, x_266); +lean_ctor_set(x_367, 1, x_366); +if (lean_is_scalar(x_352)) { + x_368 = lean_alloc_ctor(1, 1, 0); +} else { + x_368 = x_352; +} +lean_ctor_set(x_368, 0, x_367); +x_17 = x_368; +x_18 = x_274; +goto block_23; +} +else +{ +uint8_t x_369; +x_369 = lean_nat_dec_le(x_359, x_359); +if (x_369 == 0) +{ +uint8_t x_370; lean_object* x_371; lean_object* x_372; lean_object* x_373; lean_object* x_374; +lean_dec(x_359); +lean_dec(x_358); +x_370 = 1; +x_371 = lean_box(x_370); +if (lean_is_scalar(x_269)) { + x_372 = lean_alloc_ctor(0, 2, 0); +} else { + x_372 = x_269; +} +lean_ctor_set(x_372, 0, x_371); +lean_ctor_set(x_372, 1, x_363); +x_373 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_373, 0, x_266); +lean_ctor_set(x_373, 1, x_372); +if (lean_is_scalar(x_352)) { + x_374 = lean_alloc_ctor(1, 1, 0); +} else { + x_374 = x_352; +} +lean_ctor_set(x_374, 0, x_373); +x_17 = x_374; +x_18 = x_274; +goto block_23; +} +else +{ +size_t x_375; size_t x_376; lean_object* x_377; uint8_t x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; lean_object* x_382; +x_375 = 0; +x_376 = lean_usize_of_nat(x_359); +lean_dec(x_359); +x_377 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__12___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__13(x_3, x_358, x_375, x_376, x_266); +lean_dec(x_358); +x_378 = 1; +x_379 = lean_box(x_378); +if (lean_is_scalar(x_269)) { + x_380 = lean_alloc_ctor(0, 2, 0); +} else { + x_380 = x_269; +} +lean_ctor_set(x_380, 0, x_379); +lean_ctor_set(x_380, 1, x_363); +x_381 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_381, 0, x_377); +lean_ctor_set(x_381, 1, x_380); +if (lean_is_scalar(x_352)) { + x_382 = lean_alloc_ctor(1, 1, 0); +} else { + x_382 = x_352; +} +lean_ctor_set(x_382, 0, x_381); +x_17 = x_382; +x_18 = x_274; +goto block_23; +} +} +} +} +else +{ +lean_object* x_383; lean_object* x_384; lean_object* x_385; lean_object* x_386; +lean_dec(x_270); +lean_dec(x_269); +lean_dec(x_268); +lean_dec(x_267); +lean_dec(x_266); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_3); +lean_dec(x_2); +x_383 = lean_ctor_get(x_272, 0); +lean_inc(x_383); +x_384 = lean_ctor_get(x_272, 1); +lean_inc(x_384); +if (lean_is_exclusive(x_272)) { + lean_ctor_release(x_272, 0); + lean_ctor_release(x_272, 1); + x_385 = x_272; +} else { + lean_dec_ref(x_272); + x_385 = lean_box(0); +} +if (lean_is_scalar(x_385)) { + x_386 = lean_alloc_ctor(1, 2, 0); +} else { + x_386 = x_385; +} +lean_ctor_set(x_386, 0, x_383); +lean_ctor_set(x_386, 1, x_384); +return x_386; +} +} +else +{ +lean_object* x_387; lean_object* x_388; lean_object* x_389; lean_object* x_390; +lean_dec(x_270); +lean_dec(x_16); +if (lean_is_exclusive(x_271)) { + lean_ctor_release(x_271, 0); + x_387 = x_271; +} else { + lean_dec_ref(x_271); + x_387 = lean_box(0); +} +if (lean_is_scalar(x_269)) { + x_388 = lean_alloc_ctor(0, 2, 0); +} else { + x_388 = x_269; +} +lean_ctor_set(x_388, 0, x_267); +lean_ctor_set(x_388, 1, x_268); +x_389 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_389, 0, x_266); +lean_ctor_set(x_389, 1, x_388); +if (lean_is_scalar(x_387)) { + x_390 = lean_alloc_ctor(1, 1, 0); +} else { + x_390 = x_387; +} +lean_ctor_set(x_390, 0, x_389); +x_17 = x_390; +x_18 = x_13; +goto block_23; +} +} +block_23: +{ +lean_object* x_19; size_t x_20; size_t x_21; +x_19 = lean_ctor_get(x_17, 0); +lean_inc(x_19); +lean_dec(x_17); +x_20 = 1; +x_21 = lean_usize_add(x_7, x_20); +x_7 = x_21; +x_8 = x_19; +x_13 = x_18; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___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: +{ +if (x_1 == 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_box(x_1); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_2); +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_3); +lean_ctor_set(x_13, 1, x_12); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_4); +lean_ctor_set(x_14, 1, x_13); +x_15 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_15, 0, x_14); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_15); +lean_ctor_set(x_16, 1, x_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; +x_17 = lean_box(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_2); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_3); +lean_ctor_set(x_19, 1, x_18); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_4); +lean_ctor_set(x_20, 1, x_19); +x_21 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_21, 0, x_20); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_10); +return x_22; +} +} +} +static lean_object* _init_l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___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_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___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, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; uint8_t x_14; +x_13 = lean_ctor_get(x_7, 1); +lean_inc(x_13); +lean_dec(x_7); +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) +{ +lean_object* x_15; uint8_t x_16; +x_15 = lean_ctor_get(x_13, 1); +x_16 = !lean_is_exclusive(x_15); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; uint8_t x_19; lean_object* x_20; size_t x_21; size_t x_22; lean_object* x_23; +x_17 = lean_ctor_get(x_15, 0); +lean_dec(x_17); +x_18 = lean_box(0); +x_19 = 0; +x_20 = lean_box(x_19); +lean_ctor_set(x_15, 0, x_20); +x_21 = lean_array_size(x_2); +x_22 = 0; +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_5); +lean_inc(x_3); +x_23 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__14___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__15(x_2, x_3, x_5, x_18, x_2, x_21, x_22, x_13, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_23) == 0) +{ +lean_object* x_24; lean_object* x_25; uint8_t x_26; +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_is_exclusive(x_23); +if (x_26 == 0) +{ +lean_object* x_27; lean_object* x_28; uint8_t x_29; +x_27 = lean_ctor_get(x_23, 1); +x_28 = lean_ctor_get(x_23, 0); +lean_dec(x_28); +x_29 = !lean_is_exclusive(x_24); +if (x_29 == 0) +{ +lean_object* x_30; lean_object* x_31; uint8_t x_32; +x_30 = lean_ctor_get(x_24, 0); +x_31 = lean_ctor_get(x_24, 1); +lean_dec(x_31); +x_32 = !lean_is_exclusive(x_25); +if (x_32 == 0) +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; +x_33 = lean_ctor_get(x_25, 0); +x_34 = lean_ctor_get(x_25, 1); +x_35 = lean_unsigned_to_nat(0u); +x_36 = l_Lean_RBNode_fold___at_Lean_RBMap_size___spec__1___rarg(x_35, x_30); +x_37 = lean_nat_dec_eq(x_36, x_1); +lean_dec(x_36); +if (x_37 == 0) +{ +lean_object* x_38; uint8_t x_39; lean_object* x_40; lean_object* x_41; +lean_free_object(x_25); +lean_free_object(x_24); +lean_free_object(x_23); +x_38 = lean_box(0); +x_39 = lean_unbox(x_33); +lean_dec(x_33); +lean_inc(x_6); +x_40 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___lambda__1(x_39, x_34, x_30, x_6, x_38, x_8, x_9, x_10, x_11, x_27); +x_41 = lean_ctor_get(x_40, 0); +lean_inc(x_41); +if (lean_obj_tag(x_41) == 0) +{ +uint8_t x_42; +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_3); +x_42 = !lean_is_exclusive(x_40); +if (x_42 == 0) +{ +lean_object* x_43; lean_object* x_44; +x_43 = lean_ctor_get(x_40, 0); +lean_dec(x_43); +x_44 = lean_ctor_get(x_41, 0); +lean_inc(x_44); +lean_dec(x_41); +lean_ctor_set(x_40, 0, x_44); +return x_40; +} +else +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_45 = lean_ctor_get(x_40, 1); +lean_inc(x_45); +lean_dec(x_40); +x_46 = lean_ctor_get(x_41, 0); +lean_inc(x_46); +lean_dec(x_41); +x_47 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_47, 0, x_46); +lean_ctor_set(x_47, 1, x_45); +return x_47; +} +} +else +{ +lean_object* x_48; lean_object* x_49; +x_48 = lean_ctor_get(x_40, 1); +lean_inc(x_48); +lean_dec(x_40); +x_49 = lean_ctor_get(x_41, 0); +lean_inc(x_49); +lean_dec(x_41); +x_7 = x_49; +x_12 = x_48; +goto _start; +} +} +else +{ +lean_object* x_51; lean_object* x_52; +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_3); +x_51 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___closed__1; +x_52 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_52, 0, x_51); +lean_ctor_set(x_52, 1, x_24); +lean_ctor_set(x_23, 0, x_52); +return x_23; +} +} +else +{ +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_25, 0); +x_54 = lean_ctor_get(x_25, 1); +lean_inc(x_54); +lean_inc(x_53); +lean_dec(x_25); +x_55 = lean_unsigned_to_nat(0u); +x_56 = l_Lean_RBNode_fold___at_Lean_RBMap_size___spec__1___rarg(x_55, x_30); +x_57 = lean_nat_dec_eq(x_56, x_1); +lean_dec(x_56); +if (x_57 == 0) +{ +lean_object* x_58; uint8_t x_59; lean_object* x_60; lean_object* x_61; +lean_free_object(x_24); +lean_free_object(x_23); +x_58 = lean_box(0); +x_59 = lean_unbox(x_53); +lean_dec(x_53); +lean_inc(x_6); +x_60 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___lambda__1(x_59, x_54, x_30, x_6, x_58, x_8, x_9, x_10, x_11, x_27); +x_61 = lean_ctor_get(x_60, 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; +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_3); +x_62 = lean_ctor_get(x_60, 1); +lean_inc(x_62); +if (lean_is_exclusive(x_60)) { + lean_ctor_release(x_60, 0); + lean_ctor_release(x_60, 1); + x_63 = x_60; +} else { + lean_dec_ref(x_60); + x_63 = lean_box(0); +} +x_64 = lean_ctor_get(x_61, 0); +lean_inc(x_64); +lean_dec(x_61); +if (lean_is_scalar(x_63)) { + x_65 = lean_alloc_ctor(0, 2, 0); +} else { + x_65 = x_63; +} +lean_ctor_set(x_65, 0, x_64); +lean_ctor_set(x_65, 1, x_62); +return x_65; +} +else +{ +lean_object* x_66; lean_object* x_67; +x_66 = lean_ctor_get(x_60, 1); +lean_inc(x_66); +lean_dec(x_60); +x_67 = lean_ctor_get(x_61, 0); +lean_inc(x_67); +lean_dec(x_61); +x_7 = x_67; +x_12 = x_66; +goto _start; +} +} +else +{ +lean_object* x_69; lean_object* x_70; lean_object* x_71; +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_3); +x_69 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_69, 0, x_53); +lean_ctor_set(x_69, 1, x_54); +lean_ctor_set(x_24, 1, x_69); +x_70 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___closed__1; +x_71 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_71, 0, x_70); +lean_ctor_set(x_71, 1, x_24); +lean_ctor_set(x_23, 0, x_71); +return x_23; +} +} +} +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; uint8_t x_78; +x_72 = lean_ctor_get(x_24, 0); +lean_inc(x_72); +lean_dec(x_24); +x_73 = lean_ctor_get(x_25, 0); +lean_inc(x_73); +x_74 = lean_ctor_get(x_25, 1); +lean_inc(x_74); +if (lean_is_exclusive(x_25)) { + lean_ctor_release(x_25, 0); + lean_ctor_release(x_25, 1); + x_75 = x_25; +} else { + lean_dec_ref(x_25); + x_75 = lean_box(0); +} +x_76 = lean_unsigned_to_nat(0u); +x_77 = l_Lean_RBNode_fold___at_Lean_RBMap_size___spec__1___rarg(x_76, x_72); +x_78 = lean_nat_dec_eq(x_77, x_1); +lean_dec(x_77); +if (x_78 == 0) +{ +lean_object* x_79; uint8_t x_80; lean_object* x_81; lean_object* x_82; +lean_dec(x_75); +lean_free_object(x_23); +x_79 = lean_box(0); +x_80 = lean_unbox(x_73); +lean_dec(x_73); +lean_inc(x_6); +x_81 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___lambda__1(x_80, x_74, x_72, x_6, x_79, x_8, x_9, x_10, x_11, x_27); +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; lean_object* x_85; lean_object* x_86; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +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; +} else { + lean_dec_ref(x_81); + x_84 = lean_box(0); +} +x_85 = lean_ctor_get(x_82, 0); +lean_inc(x_85); +lean_dec(x_82); +if (lean_is_scalar(x_84)) { + x_86 = lean_alloc_ctor(0, 2, 0); +} else { + x_86 = x_84; +} +lean_ctor_set(x_86, 0, x_85); +lean_ctor_set(x_86, 1, x_83); +return x_86; +} +else +{ +lean_object* x_87; lean_object* x_88; +x_87 = lean_ctor_get(x_81, 1); +lean_inc(x_87); +lean_dec(x_81); +x_88 = lean_ctor_get(x_82, 0); +lean_inc(x_88); +lean_dec(x_82); +x_7 = x_88; +x_12 = x_87; +goto _start; +} +} +else +{ +lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; +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_3); +if (lean_is_scalar(x_75)) { + x_90 = lean_alloc_ctor(0, 2, 0); +} else { + x_90 = x_75; +} +lean_ctor_set(x_90, 0, x_73); +lean_ctor_set(x_90, 1, x_74); +x_91 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_91, 0, x_72); +lean_ctor_set(x_91, 1, x_90); +x_92 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___closed__1; +x_93 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_93, 0, x_92); +lean_ctor_set(x_93, 1, x_91); +lean_ctor_set(x_23, 0, x_93); +return x_23; +} +} +} +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; uint8_t x_102; +x_94 = lean_ctor_get(x_23, 1); +lean_inc(x_94); +lean_dec(x_23); +x_95 = lean_ctor_get(x_24, 0); +lean_inc(x_95); +if (lean_is_exclusive(x_24)) { + lean_ctor_release(x_24, 0); + lean_ctor_release(x_24, 1); + x_96 = x_24; +} else { + lean_dec_ref(x_24); + x_96 = lean_box(0); +} +x_97 = lean_ctor_get(x_25, 0); +lean_inc(x_97); +x_98 = lean_ctor_get(x_25, 1); +lean_inc(x_98); +if (lean_is_exclusive(x_25)) { + lean_ctor_release(x_25, 0); + lean_ctor_release(x_25, 1); + x_99 = x_25; +} else { + lean_dec_ref(x_25); + x_99 = lean_box(0); +} +x_100 = lean_unsigned_to_nat(0u); +x_101 = l_Lean_RBNode_fold___at_Lean_RBMap_size___spec__1___rarg(x_100, x_95); +x_102 = lean_nat_dec_eq(x_101, x_1); +lean_dec(x_101); +if (x_102 == 0) +{ +lean_object* x_103; uint8_t x_104; lean_object* x_105; lean_object* x_106; +lean_dec(x_99); +lean_dec(x_96); +x_103 = lean_box(0); +x_104 = lean_unbox(x_97); +lean_dec(x_97); lean_inc(x_6); +x_105 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___lambda__1(x_104, x_98, x_95, x_6, x_103, x_8, x_9, x_10, x_11, x_94); +x_106 = lean_ctor_get(x_105, 0); +lean_inc(x_106); +if (lean_obj_tag(x_106) == 0) +{ +lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; +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_3); +x_107 = lean_ctor_get(x_105, 1); +lean_inc(x_107); +if (lean_is_exclusive(x_105)) { + lean_ctor_release(x_105, 0); + lean_ctor_release(x_105, 1); + x_108 = x_105; +} else { + lean_dec_ref(x_105); + x_108 = lean_box(0); +} +x_109 = lean_ctor_get(x_106, 0); +lean_inc(x_109); +lean_dec(x_106); +if (lean_is_scalar(x_108)) { + x_110 = lean_alloc_ctor(0, 2, 0); +} else { + x_110 = x_108; +} +lean_ctor_set(x_110, 0, x_109); +lean_ctor_set(x_110, 1, x_107); +return x_110; +} +else +{ +lean_object* x_111; lean_object* x_112; +x_111 = lean_ctor_get(x_105, 1); +lean_inc(x_111); +lean_dec(x_105); +x_112 = lean_ctor_get(x_106, 0); +lean_inc(x_112); +lean_dec(x_106); +x_7 = x_112; +x_12 = x_111; +goto _start; +} +} +else +{ +lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; +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_3); +if (lean_is_scalar(x_99)) { + x_114 = lean_alloc_ctor(0, 2, 0); +} else { + x_114 = x_99; +} +lean_ctor_set(x_114, 0, x_97); +lean_ctor_set(x_114, 1, x_98); +if (lean_is_scalar(x_96)) { + x_115 = lean_alloc_ctor(0, 2, 0); +} else { + x_115 = x_96; +} +lean_ctor_set(x_115, 0, x_95); +lean_ctor_set(x_115, 1, x_114); +x_116 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___closed__1; +x_117 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_117, 0, x_116); +lean_ctor_set(x_117, 1, x_115); +x_118 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_118, 0, x_117); +lean_ctor_set(x_118, 1, x_94); +return x_118; +} +} +} +else +{ +uint8_t x_119; +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_3); +x_119 = !lean_is_exclusive(x_23); +if (x_119 == 0) +{ +return x_23; +} +else +{ +lean_object* x_120; lean_object* x_121; lean_object* x_122; +x_120 = lean_ctor_get(x_23, 0); +x_121 = lean_ctor_get(x_23, 1); +lean_inc(x_121); +lean_inc(x_120); +lean_dec(x_23); +x_122 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_122, 0, x_120); +lean_ctor_set(x_122, 1, x_121); +return x_122; +} +} +} +else +{ +lean_object* x_123; lean_object* x_124; uint8_t x_125; lean_object* x_126; lean_object* x_127; size_t x_128; size_t x_129; lean_object* x_130; +x_123 = lean_ctor_get(x_15, 1); +lean_inc(x_123); +lean_dec(x_15); +x_124 = lean_box(0); +x_125 = 0; +x_126 = lean_box(x_125); +x_127 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_127, 0, x_126); +lean_ctor_set(x_127, 1, x_123); +lean_ctor_set(x_13, 1, x_127); +x_128 = lean_array_size(x_2); +x_129 = 0; +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_24); -lean_inc(x_12); -x_25 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFunMask(x_12, x_24, x_4, x_5, x_6, x_7, x_16); -if (lean_obj_tag(x_25) == 0) +lean_inc(x_3); +x_130 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__14___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__15(x_2, x_3, x_5, x_124, x_2, x_128, x_129, x_13, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_130) == 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 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_28, 0, x_17); -lean_ctor_set(x_28, 1, x_24); -lean_ctor_set(x_28, 2, x_21); -x_29 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__2(x_26, x_28, x_28, x_23, x_17, lean_box(0), lean_box(0), x_3, x_4, x_5, x_6, x_7, x_27); -lean_dec(x_28); -lean_dec(x_26); -if (lean_obj_tag(x_29) == 0) +lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; uint8_t x_142; +x_131 = lean_ctor_get(x_130, 0); +lean_inc(x_131); +x_132 = lean_ctor_get(x_131, 1); +lean_inc(x_132); +x_133 = lean_ctor_get(x_130, 1); +lean_inc(x_133); +if (lean_is_exclusive(x_130)) { + lean_ctor_release(x_130, 0); + lean_ctor_release(x_130, 1); + x_134 = x_130; +} else { + lean_dec_ref(x_130); + x_134 = lean_box(0); +} +x_135 = lean_ctor_get(x_131, 0); +lean_inc(x_135); +if (lean_is_exclusive(x_131)) { + lean_ctor_release(x_131, 0); + lean_ctor_release(x_131, 1); + x_136 = x_131; +} else { + lean_dec_ref(x_131); + x_136 = lean_box(0); +} +x_137 = lean_ctor_get(x_132, 0); +lean_inc(x_137); +x_138 = lean_ctor_get(x_132, 1); +lean_inc(x_138); +if (lean_is_exclusive(x_132)) { + lean_ctor_release(x_132, 0); + lean_ctor_release(x_132, 1); + x_139 = x_132; +} else { + lean_dec_ref(x_132); + x_139 = lean_box(0); +} +x_140 = lean_unsigned_to_nat(0u); +x_141 = l_Lean_RBNode_fold___at_Lean_RBMap_size___spec__1___rarg(x_140, x_135); +x_142 = lean_nat_dec_eq(x_141, x_1); +lean_dec(x_141); +if (x_142 == 0) { -uint8_t x_30; -x_30 = !lean_is_exclusive(x_29); -if (x_30 == 0) +lean_object* x_143; uint8_t x_144; lean_object* x_145; lean_object* x_146; +lean_dec(x_139); +lean_dec(x_136); +lean_dec(x_134); +x_143 = lean_box(0); +x_144 = lean_unbox(x_137); +lean_dec(x_137); +lean_inc(x_6); +x_145 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___lambda__1(x_144, x_138, x_135, x_6, x_143, x_8, x_9, x_10, x_11, x_133); +x_146 = lean_ctor_get(x_145, 0); +lean_inc(x_146); +if (lean_obj_tag(x_146) == 0) { -lean_object* x_31; lean_object* x_32; -x_31 = lean_ctor_get(x_29, 0); -x_32 = l_Lean_mkAppN(x_12, x_31); -lean_dec(x_31); -lean_ctor_set(x_29, 0, x_32); -return x_29; +lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; +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_3); +x_147 = lean_ctor_get(x_145, 1); +lean_inc(x_147); +if (lean_is_exclusive(x_145)) { + lean_ctor_release(x_145, 0); + lean_ctor_release(x_145, 1); + x_148 = x_145; +} else { + lean_dec_ref(x_145); + x_148 = lean_box(0); } -else -{ -lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_33 = lean_ctor_get(x_29, 0); -x_34 = lean_ctor_get(x_29, 1); -lean_inc(x_34); -lean_inc(x_33); -lean_dec(x_29); -x_35 = l_Lean_mkAppN(x_12, x_33); -lean_dec(x_33); -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_149 = lean_ctor_get(x_146, 0); +lean_inc(x_149); +lean_dec(x_146); +if (lean_is_scalar(x_148)) { + x_150 = lean_alloc_ctor(0, 2, 0); +} else { + x_150 = x_148; } +lean_ctor_set(x_150, 0, x_149); +lean_ctor_set(x_150, 1, x_147); +return x_150; } else { -uint8_t x_37; -lean_dec(x_12); -x_37 = !lean_is_exclusive(x_29); -if (x_37 == 0) -{ -return x_29; +lean_object* x_151; lean_object* x_152; +x_151 = lean_ctor_get(x_145, 1); +lean_inc(x_151); +lean_dec(x_145); +x_152 = lean_ctor_get(x_146, 0); +lean_inc(x_152); +lean_dec(x_146); +x_7 = x_152; +x_12 = x_151; +goto _start; +} } else { -lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_38 = lean_ctor_get(x_29, 0); -x_39 = lean_ctor_get(x_29, 1); -lean_inc(x_39); -lean_inc(x_38); -lean_dec(x_29); -x_40 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_40, 0, x_38); -lean_ctor_set(x_40, 1, x_39); -return x_40; +lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; +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_3); +if (lean_is_scalar(x_139)) { + x_154 = lean_alloc_ctor(0, 2, 0); +} else { + x_154 = x_139; +} +lean_ctor_set(x_154, 0, x_137); +lean_ctor_set(x_154, 1, x_138); +if (lean_is_scalar(x_136)) { + x_155 = lean_alloc_ctor(0, 2, 0); +} else { + x_155 = x_136; +} +lean_ctor_set(x_155, 0, x_135); +lean_ctor_set(x_155, 1, x_154); +x_156 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___closed__1; +x_157 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_157, 0, x_156); +lean_ctor_set(x_157, 1, x_155); +if (lean_is_scalar(x_134)) { + x_158 = lean_alloc_ctor(0, 2, 0); +} else { + x_158 = x_134; } +lean_ctor_set(x_158, 0, x_157); +lean_ctor_set(x_158, 1, x_133); +return x_158; } } else { -uint8_t x_41; -lean_dec(x_24); -lean_dec(x_23); -lean_dec(x_12); -lean_dec(x_7); +lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; +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_41 = !lean_is_exclusive(x_25); -if (x_41 == 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_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_159 = lean_ctor_get(x_130, 0); +lean_inc(x_159); +x_160 = lean_ctor_get(x_130, 1); +lean_inc(x_160); +if (lean_is_exclusive(x_130)) { + lean_ctor_release(x_130, 0); + lean_ctor_release(x_130, 1); + x_161 = x_130; +} else { + lean_dec_ref(x_130); + x_161 = lean_box(0); } +if (lean_is_scalar(x_161)) { + x_162 = lean_alloc_ctor(1, 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; } } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___closed__1() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("invalid pattern, it does not have pattern variables", 51, 51); -return x_1; +lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; uint8_t x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; size_t x_172; size_t x_173; lean_object* x_174; +x_163 = lean_ctor_get(x_13, 1); +x_164 = lean_ctor_get(x_13, 0); +lean_inc(x_163); +lean_inc(x_164); +lean_dec(x_13); +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; +} else { + lean_dec_ref(x_163); + x_166 = lean_box(0); } +x_167 = lean_box(0); +x_168 = 0; +x_169 = lean_box(x_168); +if (lean_is_scalar(x_166)) { + x_170 = lean_alloc_ctor(0, 2, 0); +} else { + x_170 = x_166; } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___closed__2() { -_start: +lean_ctor_set(x_170, 0, x_169); +lean_ctor_set(x_170, 1, x_165); +x_171 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_171, 0, x_164); +lean_ctor_set(x_171, 1, x_170); +x_172 = lean_array_size(x_2); +x_173 = 0; +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_5); +lean_inc(x_3); +x_174 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__14___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__15(x_2, x_3, x_5, x_167, x_2, x_172, x_173, x_171, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_174) == 0) { -lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___closed__1; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; +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; uint8_t x_186; +x_175 = lean_ctor_get(x_174, 0); +lean_inc(x_175); +x_176 = lean_ctor_get(x_175, 1); +lean_inc(x_176); +x_177 = lean_ctor_get(x_174, 1); +lean_inc(x_177); +if (lean_is_exclusive(x_174)) { + lean_ctor_release(x_174, 0); + lean_ctor_release(x_174, 1); + x_178 = x_174; +} else { + lean_dec_ref(x_174); + x_178 = lean_box(0); } +x_179 = lean_ctor_get(x_175, 0); +lean_inc(x_179); +if (lean_is_exclusive(x_175)) { + lean_ctor_release(x_175, 0); + lean_ctor_release(x_175, 1); + x_180 = x_175; +} else { + lean_dec_ref(x_175); + x_180 = lean_box(0); } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -if (x_2 == 0) -{ -lean_object* x_9; lean_object* x_10; -x_9 = lean_box(0); -x_10 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1(x_1, x_9, x_3, x_4, x_5, x_6, x_7, x_8); -return x_10; +x_181 = lean_ctor_get(x_176, 0); +lean_inc(x_181); +x_182 = lean_ctor_get(x_176, 1); +lean_inc(x_182); +if (lean_is_exclusive(x_176)) { + lean_ctor_release(x_176, 0); + lean_ctor_release(x_176, 1); + x_183 = x_176; +} else { + lean_dec_ref(x_176); + x_183 = lean_box(0); } -else +x_184 = lean_unsigned_to_nat(0u); +x_185 = l_Lean_RBNode_fold___at_Lean_RBMap_size___spec__1___rarg(x_184, x_179); +x_186 = lean_nat_dec_eq(x_185, x_1); +lean_dec(x_185); +if (x_186 == 0) { -uint8_t x_11; -x_11 = l_Lean_Expr_hasLooseBVars(x_1); -if (x_11 == 0) +lean_object* x_187; uint8_t x_188; lean_object* x_189; lean_object* x_190; +lean_dec(x_183); +lean_dec(x_180); +lean_dec(x_178); +x_187 = lean_box(0); +x_188 = lean_unbox(x_181); +lean_dec(x_181); +lean_inc(x_6); +x_189 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___lambda__1(x_188, x_182, x_179, x_6, x_187, x_8, x_9, x_10, x_11, x_177); +x_190 = lean_ctor_get(x_189, 0); +lean_inc(x_190); +if (lean_obj_tag(x_190) == 0) { -lean_object* x_12; lean_object* x_13; uint8_t x_14; -lean_dec(x_1); -x_12 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___closed__2; -x_13 = l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__4(x_12, x_3, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_7); +lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; +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_14 = !lean_is_exclusive(x_13); -if (x_14 == 0) -{ -return x_13; +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; +} else { + lean_dec_ref(x_189); + x_192 = lean_box(0); } -else -{ -lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = lean_ctor_get(x_13, 0); -x_16 = lean_ctor_get(x_13, 1); -lean_inc(x_16); -lean_inc(x_15); -lean_dec(x_13); -x_17 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_17, 0, x_15); -lean_ctor_set(x_17, 1, x_16); -return x_17; +x_193 = lean_ctor_get(x_190, 0); +lean_inc(x_193); +lean_dec(x_190); +if (lean_is_scalar(x_192)) { + x_194 = lean_alloc_ctor(0, 2, 0); +} else { + x_194 = x_192; } +lean_ctor_set(x_194, 0, x_193); +lean_ctor_set(x_194, 1, x_191); +return x_194; } else { -lean_object* x_18; lean_object* x_19; -x_18 = lean_box(0); -x_19 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1(x_1, x_18, x_3, x_4, x_5, x_6, x_7, x_8); -return x_19; -} -} +lean_object* x_195; lean_object* x_196; +x_195 = lean_ctor_get(x_189, 1); +lean_inc(x_195); +lean_dec(x_189); +x_196 = lean_ctor_get(x_190, 0); +lean_inc(x_196); +lean_dec(x_190); +x_7 = x_196; +x_12 = x_195; +goto _start; } } -LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_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, lean_object* x_7) { -_start: +else { -lean_object* x_8; -x_8 = l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); -return x_8; -} +if (lean_is_scalar(x_183)) { + x_198 = lean_alloc_ctor(0, 2, 0); +} else { + x_198 = x_183; } -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__2___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -lean_object* x_10; -x_10 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__2___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_1); -return x_10; +lean_ctor_set(x_198, 0, x_181); +lean_ctor_set(x_198, 1, x_182); +if (lean_is_scalar(x_180)) { + x_199 = lean_alloc_ctor(0, 2, 0); +} else { + x_199 = x_180; } +lean_ctor_set(x_199, 0, x_179); +lean_ctor_set(x_199, 1, x_198); +x_200 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___closed__1; +x_201 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_201, 0, x_200); +lean_ctor_set(x_201, 1, x_199); +if (lean_is_scalar(x_178)) { + x_202 = lean_alloc_ctor(0, 2, 0); +} else { + x_202 = x_178; } -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { -_start: -{ -lean_object* x_14; -x_14 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_14; +lean_ctor_set(x_202, 0, x_201); +lean_ctor_set(x_202, 1, x_177); +return x_202; } } -LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_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) { -_start: +else { -lean_object* x_8; -x_8 = l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); -return x_8; +x_203 = lean_ctor_get(x_174, 0); +lean_inc(x_203); +x_204 = lean_ctor_get(x_174, 1); +lean_inc(x_204); +if (lean_is_exclusive(x_174)) { + lean_ctor_release(x_174, 0); + lean_ctor_release(x_174, 1); + x_205 = x_174; +} else { + lean_dec_ref(x_174); + x_205 = lean_box(0); } +if (lean_is_scalar(x_205)) { + x_206 = lean_alloc_ctor(1, 2, 0); +} else { + x_206 = x_205; } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_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) { -_start: -{ -lean_object* x_9; -x_9 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_2); -return x_9; +lean_ctor_set(x_206, 0, x_203); +lean_ctor_set(x_206, 1, x_204); +return x_206; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_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) { -_start: -{ -uint8_t x_9; lean_object* x_10; -x_9 = lean_unbox(x_2); -lean_dec(x_2); -x_10 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go(x_1, x_9, x_3, x_4, x_5, x_6, x_7, x_8); -return x_10; } } -LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_Meta_Grind_NormalizePattern_main___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__17(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -if (lean_obj_tag(x_1) == 0) -{ -lean_object* x_9; lean_object* x_10; -lean_dec(x_7); +lean_object* x_12; uint8_t x_13; +x_12 = lean_ctor_get(x_6, 1); +lean_inc(x_12); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -x_9 = l_List_reverse___rarg(x_2); -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 +x_13 = !lean_is_exclusive(x_12); +if (x_13 == 0) { -uint8_t x_11; -x_11 = !lean_is_exclusive(x_1); -if (x_11 == 0) +lean_object* x_14; uint8_t x_15; +x_14 = lean_ctor_get(x_12, 1); +x_15 = !lean_is_exclusive(x_14); +if (x_15 == 0) { -lean_object* x_12; lean_object* x_13; uint8_t x_14; lean_object* x_15; -x_12 = lean_ctor_get(x_1, 0); -x_13 = lean_ctor_get(x_1, 1); -x_14 = 0; +lean_object* x_16; lean_object* x_17; uint8_t x_18; lean_object* x_19; size_t x_20; size_t x_21; lean_object* x_22; +x_16 = lean_ctor_get(x_14, 0); +lean_dec(x_16); +x_17 = lean_box(0); +x_18 = 0; +x_19 = lean_box(x_18); +lean_ctor_set(x_14, 0, x_19); +x_20 = lean_array_size(x_2); +x_21 = 0; +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -x_15 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go(x_12, x_14, x_3, x_4, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_15) == 0) +x_22 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__14___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__15(x_2, x_3, x_4, x_17, x_2, x_20, x_21, x_12, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_22) == 0) { -lean_object* x_16; lean_object* x_17; -x_16 = lean_ctor_get(x_15, 0); -lean_inc(x_16); -x_17 = lean_ctor_get(x_15, 1); -lean_inc(x_17); -lean_dec(x_15); -lean_ctor_set(x_1, 1, x_2); -lean_ctor_set(x_1, 0, x_16); +lean_object* x_23; lean_object* x_24; uint8_t x_25; +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +x_24 = lean_ctor_get(x_23, 1); +lean_inc(x_24); +x_25 = !lean_is_exclusive(x_22); +if (x_25 == 0) { -lean_object* _tmp_0 = x_13; -lean_object* _tmp_1 = x_1; -lean_object* _tmp_7 = x_17; -x_1 = _tmp_0; -x_2 = _tmp_1; -x_8 = _tmp_7; -} -goto _start; -} -else +lean_object* x_26; lean_object* x_27; uint8_t x_28; +x_26 = lean_ctor_get(x_22, 1); +x_27 = lean_ctor_get(x_22, 0); +lean_dec(x_27); +x_28 = !lean_is_exclusive(x_23); +if (x_28 == 0) { -uint8_t x_19; -lean_free_object(x_1); -lean_dec(x_13); +lean_object* x_29; lean_object* x_30; uint8_t x_31; +x_29 = lean_ctor_get(x_23, 0); +x_30 = lean_ctor_get(x_23, 1); +lean_dec(x_30); +x_31 = !lean_is_exclusive(x_24); +if (x_31 == 0) +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; +x_32 = lean_ctor_get(x_24, 0); +x_33 = lean_ctor_get(x_24, 1); +x_34 = lean_unsigned_to_nat(0u); +x_35 = l_Lean_RBNode_fold___at_Lean_RBMap_size___spec__1___rarg(x_34, x_29); +x_36 = lean_nat_dec_eq(x_35, x_1); +lean_dec(x_35); +if (x_36 == 0) +{ +lean_object* x_37; uint8_t x_38; lean_object* x_39; lean_object* x_40; +lean_free_object(x_24); +lean_free_object(x_23); +lean_free_object(x_22); +x_37 = lean_box(0); +x_38 = lean_unbox(x_32); +lean_dec(x_32); +lean_inc(x_5); +x_39 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___lambda__1(x_38, x_33, x_29, x_5, x_37, x_7, x_8, x_9, x_10, x_26); +x_40 = lean_ctor_get(x_39, 0); +lean_inc(x_40); +if (lean_obj_tag(x_40) == 0) +{ +uint8_t x_41; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_7); -lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); -x_19 = !lean_is_exclusive(x_15); -if (x_19 == 0) +x_41 = !lean_is_exclusive(x_39); +if (x_41 == 0) { -return x_15; +lean_object* x_42; lean_object* x_43; +x_42 = lean_ctor_get(x_39, 0); +lean_dec(x_42); +x_43 = lean_ctor_get(x_40, 0); +lean_inc(x_43); +lean_dec(x_40); +lean_ctor_set(x_39, 0, x_43); +return x_39; } 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_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_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_39, 1); +lean_inc(x_44); +lean_dec(x_39); +x_45 = lean_ctor_get(x_40, 0); +lean_inc(x_45); +lean_dec(x_40); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_44); +return x_46; } } else { -lean_object* x_23; lean_object* x_24; uint8_t x_25; lean_object* x_26; -x_23 = lean_ctor_get(x_1, 0); -x_24 = lean_ctor_get(x_1, 1); -lean_inc(x_24); -lean_inc(x_23); -lean_dec(x_1); -x_25 = 0; -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -x_26 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go(x_23, x_25, x_3, x_4, x_5, x_6, x_7, x_8); -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 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_29, 0, x_27); -lean_ctor_set(x_29, 1, x_2); -x_1 = x_24; -x_2 = x_29; -x_8 = x_28; +lean_object* x_47; lean_object* x_48; +x_47 = lean_ctor_get(x_39, 1); +lean_inc(x_47); +lean_dec(x_39); +x_48 = lean_ctor_get(x_40, 0); +lean_inc(x_48); +lean_dec(x_40); +x_6 = x_48; +x_11 = x_47; goto _start; } +} else { -lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; -lean_dec(x_24); +lean_object* x_50; lean_object* x_51; +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_31 = lean_ctor_get(x_26, 0); -lean_inc(x_31); -x_32 = lean_ctor_get(x_26, 1); -lean_inc(x_32); -if (lean_is_exclusive(x_26)) { - lean_ctor_release(x_26, 0); - lean_ctor_release(x_26, 1); - x_33 = x_26; -} else { - lean_dec_ref(x_26); - x_33 = lean_box(0); -} -if (lean_is_scalar(x_33)) { - x_34 = lean_alloc_ctor(1, 2, 0); -} else { - x_34 = x_33; -} -lean_ctor_set(x_34, 0, x_31); -lean_ctor_set(x_34, 1, x_32); -return x_34; -} -} -} -} -} -static lean_object* _init_l_Lean_Meta_Grind_NormalizePattern_main___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = lean_box(0); -x_2 = lean_array_mk(x_1); -return x_2; +x_50 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___closed__1; +x_51 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_51, 0, x_50); +lean_ctor_set(x_51, 1, x_23); +lean_ctor_set(x_22, 0, x_51); +return x_22; } } -static lean_object* _init_l_Lean_Meta_Grind_NormalizePattern_main___closed__2() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(10u); -x_2 = lean_unsigned_to_nat(1u); -x_3 = l_Nat_nextPowerOfTwo_go(x_1, x_2, lean_box(0)); -return x_3; -} -} -static lean_object* _init_l_Lean_Meta_Grind_NormalizePattern_main___closed__3() { -_start: +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_24, 0); +x_53 = lean_ctor_get(x_24, 1); +lean_inc(x_53); +lean_inc(x_52); +lean_dec(x_24); +x_54 = lean_unsigned_to_nat(0u); +x_55 = l_Lean_RBNode_fold___at_Lean_RBMap_size___spec__1___rarg(x_54, x_29); +x_56 = lean_nat_dec_eq(x_55, x_1); +lean_dec(x_55); +if (x_56 == 0) { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Meta_Grind_NormalizePattern_main___closed__2; -x_3 = lean_mk_array(x_2, x_1); -return x_3; -} -} -static lean_object* _init_l_Lean_Meta_Grind_NormalizePattern_main___closed__4() { -_start: +lean_object* x_57; uint8_t x_58; lean_object* x_59; lean_object* x_60; +lean_free_object(x_23); +lean_free_object(x_22); +x_57 = lean_box(0); +x_58 = lean_unbox(x_52); +lean_dec(x_52); +lean_inc(x_5); +x_59 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___lambda__1(x_58, x_53, x_29, x_5, x_57, x_7, x_8, x_9, x_10, x_26); +x_60 = lean_ctor_get(x_59, 0); +lean_inc(x_60); +if (lean_obj_tag(x_60) == 0) { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(0u); -x_2 = l_Lean_Meta_Grind_NormalizePattern_main___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_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; +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_61 = lean_ctor_get(x_59, 1); +lean_inc(x_61); +if (lean_is_exclusive(x_59)) { + lean_ctor_release(x_59, 0); + lean_ctor_release(x_59, 1); + x_62 = x_59; +} else { + lean_dec_ref(x_59); + x_62 = lean_box(0); } -static lean_object* _init_l_Lean_Meta_Grind_NormalizePattern_main___closed__5() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_Grind_NormalizePattern_main___closed__1; -x_2 = l_Lean_Meta_Grind_NormalizePattern_main___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_2); -lean_ctor_set(x_3, 2, x_2); -return x_3; +x_63 = lean_ctor_get(x_60, 0); +lean_inc(x_63); +lean_dec(x_60); +if (lean_is_scalar(x_62)) { + x_64 = lean_alloc_ctor(0, 2, 0); +} else { + x_64 = x_62; } +lean_ctor_set(x_64, 0, x_63); +lean_ctor_set(x_64, 1, x_61); +return x_64; } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_NormalizePattern_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) { -_start: -{ -lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; -x_7 = lean_box(0); -x_8 = l_Lean_Meta_Grind_NormalizePattern_main___closed__5; -x_9 = lean_st_mk_ref(x_8, 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; -x_11 = lean_ctor_get(x_9, 0); -x_12 = lean_ctor_get(x_9, 1); -lean_inc(x_11); -x_13 = l_List_mapM_loop___at_Lean_Meta_Grind_NormalizePattern_main___spec__1(x_1, x_7, x_11, x_2, x_3, x_4, x_5, 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_st_ref_get(x_11, x_15); -lean_dec(x_11); -x_17 = !lean_is_exclusive(x_16); -if (x_17 == 0) +else { -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_18 = lean_ctor_get(x_16, 0); -x_19 = lean_ctor_get(x_18, 0); -lean_inc(x_19); -x_20 = lean_array_to_list(x_19); -x_21 = lean_ctor_get(x_18, 2); -lean_inc(x_21); -lean_dec(x_18); -lean_ctor_set(x_9, 1, x_21); -lean_ctor_set(x_9, 0, x_20); -x_22 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_22, 0, x_14); -lean_ctor_set(x_22, 1, x_9); -lean_ctor_set(x_16, 0, x_22); -return x_16; +lean_object* x_65; lean_object* x_66; +x_65 = lean_ctor_get(x_59, 1); +lean_inc(x_65); +lean_dec(x_59); +x_66 = lean_ctor_get(x_60, 0); +lean_inc(x_66); +lean_dec(x_60); +x_6 = x_66; +x_11 = x_65; +goto _start; +} } 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; -x_23 = lean_ctor_get(x_16, 0); -x_24 = lean_ctor_get(x_16, 1); -lean_inc(x_24); -lean_inc(x_23); -lean_dec(x_16); -x_25 = lean_ctor_get(x_23, 0); -lean_inc(x_25); -x_26 = lean_array_to_list(x_25); -x_27 = lean_ctor_get(x_23, 2); -lean_inc(x_27); -lean_dec(x_23); -lean_ctor_set(x_9, 1, x_27); -lean_ctor_set(x_9, 0, x_26); -x_28 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_28, 0, x_14); -lean_ctor_set(x_28, 1, x_9); -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; +lean_object* x_68; lean_object* x_69; lean_object* x_70; +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_68 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_68, 0, x_52); +lean_ctor_set(x_68, 1, x_53); +lean_ctor_set(x_23, 1, x_68); +x_69 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___closed__1; +x_70 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_70, 0, x_69); +lean_ctor_set(x_70, 1, x_23); +lean_ctor_set(x_22, 0, x_70); +return x_22; +} } } else { -uint8_t x_30; -lean_free_object(x_9); -lean_dec(x_11); -x_30 = !lean_is_exclusive(x_13); -if (x_30 == 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; uint8_t x_77; +x_71 = lean_ctor_get(x_23, 0); +lean_inc(x_71); +lean_dec(x_23); +x_72 = lean_ctor_get(x_24, 0); +lean_inc(x_72); +x_73 = lean_ctor_get(x_24, 1); +lean_inc(x_73); +if (lean_is_exclusive(x_24)) { + lean_ctor_release(x_24, 0); + lean_ctor_release(x_24, 1); + x_74 = x_24; +} else { + lean_dec_ref(x_24); + x_74 = lean_box(0); +} +x_75 = lean_unsigned_to_nat(0u); +x_76 = l_Lean_RBNode_fold___at_Lean_RBMap_size___spec__1___rarg(x_75, x_71); +x_77 = lean_nat_dec_eq(x_76, x_1); +lean_dec(x_76); +if (x_77 == 0) { -return x_13; +lean_object* x_78; uint8_t x_79; lean_object* x_80; lean_object* x_81; +lean_dec(x_74); +lean_free_object(x_22); +x_78 = lean_box(0); +x_79 = lean_unbox(x_72); +lean_dec(x_72); +lean_inc(x_5); +x_80 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___lambda__1(x_79, x_73, x_71, x_5, x_78, x_7, x_8, x_9, x_10, x_26); +x_81 = lean_ctor_get(x_80, 0); +lean_inc(x_81); +if (lean_obj_tag(x_81) == 0) +{ +lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_82 = lean_ctor_get(x_80, 1); +lean_inc(x_82); +if (lean_is_exclusive(x_80)) { + lean_ctor_release(x_80, 0); + lean_ctor_release(x_80, 1); + x_83 = x_80; +} else { + lean_dec_ref(x_80); + x_83 = lean_box(0); +} +x_84 = lean_ctor_get(x_81, 0); +lean_inc(x_84); +lean_dec(x_81); +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_31; lean_object* x_32; lean_object* x_33; -x_31 = lean_ctor_get(x_13, 0); -x_32 = lean_ctor_get(x_13, 1); -lean_inc(x_32); -lean_inc(x_31); -lean_dec(x_13); -x_33 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_33, 0, x_31); -lean_ctor_set(x_33, 1, x_32); -return x_33; -} +lean_object* x_86; lean_object* x_87; +x_86 = lean_ctor_get(x_80, 1); +lean_inc(x_86); +lean_dec(x_80); +x_87 = lean_ctor_get(x_81, 0); +lean_inc(x_87); +lean_dec(x_81); +x_6 = x_87; +x_11 = x_86; +goto _start; } } else { -lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_34 = lean_ctor_get(x_9, 0); -x_35 = lean_ctor_get(x_9, 1); -lean_inc(x_35); -lean_inc(x_34); +lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; +lean_dec(x_10); lean_dec(x_9); -lean_inc(x_34); -x_36 = l_List_mapM_loop___at_Lean_Meta_Grind_NormalizePattern_main___spec__1(x_1, x_7, x_34, x_2, x_3, x_4, x_5, x_35); -if (lean_obj_tag(x_36) == 0) +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +if (lean_is_scalar(x_74)) { + x_89 = lean_alloc_ctor(0, 2, 0); +} else { + x_89 = x_74; +} +lean_ctor_set(x_89, 0, x_72); +lean_ctor_set(x_89, 1, x_73); +x_90 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_90, 0, x_71); +lean_ctor_set(x_90, 1, x_89); +x_91 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___closed__1; +x_92 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_92, 0, x_91); +lean_ctor_set(x_92, 1, x_90); +lean_ctor_set(x_22, 0, x_92); +return x_22; +} +} +} +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_36, 0); -lean_inc(x_37); -x_38 = lean_ctor_get(x_36, 1); -lean_inc(x_38); -lean_dec(x_36); -x_39 = lean_st_ref_get(x_34, x_38); -lean_dec(x_34); -x_40 = lean_ctor_get(x_39, 0); -lean_inc(x_40); -x_41 = lean_ctor_get(x_39, 1); -lean_inc(x_41); -if (lean_is_exclusive(x_39)) { - lean_ctor_release(x_39, 0); - lean_ctor_release(x_39, 1); - x_42 = x_39; +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; +x_93 = lean_ctor_get(x_22, 1); +lean_inc(x_93); +lean_dec(x_22); +x_94 = lean_ctor_get(x_23, 0); +lean_inc(x_94); +if (lean_is_exclusive(x_23)) { + lean_ctor_release(x_23, 0); + lean_ctor_release(x_23, 1); + x_95 = x_23; } else { - lean_dec_ref(x_39); - x_42 = lean_box(0); + lean_dec_ref(x_23); + x_95 = lean_box(0); } -x_43 = lean_ctor_get(x_40, 0); -lean_inc(x_43); -x_44 = lean_array_to_list(x_43); -x_45 = lean_ctor_get(x_40, 2); -lean_inc(x_45); -lean_dec(x_40); -x_46 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_45); -x_47 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_47, 0, x_37); -lean_ctor_set(x_47, 1, x_46); -if (lean_is_scalar(x_42)) { - x_48 = lean_alloc_ctor(0, 2, 0); +x_96 = lean_ctor_get(x_24, 0); +lean_inc(x_96); +x_97 = lean_ctor_get(x_24, 1); +lean_inc(x_97); +if (lean_is_exclusive(x_24)) { + lean_ctor_release(x_24, 0); + lean_ctor_release(x_24, 1); + x_98 = x_24; +} else { + lean_dec_ref(x_24); + x_98 = lean_box(0); +} +x_99 = lean_unsigned_to_nat(0u); +x_100 = l_Lean_RBNode_fold___at_Lean_RBMap_size___spec__1___rarg(x_99, x_94); +x_101 = lean_nat_dec_eq(x_100, x_1); +lean_dec(x_100); +if (x_101 == 0) +{ +lean_object* x_102; uint8_t x_103; lean_object* x_104; lean_object* x_105; +lean_dec(x_98); +lean_dec(x_95); +x_102 = lean_box(0); +x_103 = lean_unbox(x_96); +lean_dec(x_96); +lean_inc(x_5); +x_104 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___lambda__1(x_103, x_97, x_94, x_5, x_102, x_7, x_8, x_9, x_10, x_93); +x_105 = lean_ctor_get(x_104, 0); +lean_inc(x_105); +if (lean_obj_tag(x_105) == 0) +{ +lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; +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_106 = lean_ctor_get(x_104, 1); +lean_inc(x_106); +if (lean_is_exclusive(x_104)) { + lean_ctor_release(x_104, 0); + lean_ctor_release(x_104, 1); + x_107 = x_104; } else { - x_48 = x_42; + lean_dec_ref(x_104); + x_107 = lean_box(0); } -lean_ctor_set(x_48, 0, x_47); -lean_ctor_set(x_48, 1, x_41); -return x_48; +x_108 = lean_ctor_get(x_105, 0); +lean_inc(x_108); +lean_dec(x_105); +if (lean_is_scalar(x_107)) { + x_109 = lean_alloc_ctor(0, 2, 0); +} else { + x_109 = x_107; +} +lean_ctor_set(x_109, 0, x_108); +lean_ctor_set(x_109, 1, x_106); +return x_109; } else { -lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; -lean_dec(x_34); -x_49 = lean_ctor_get(x_36, 0); -lean_inc(x_49); -x_50 = lean_ctor_get(x_36, 1); -lean_inc(x_50); -if (lean_is_exclusive(x_36)) { - lean_ctor_release(x_36, 0); - lean_ctor_release(x_36, 1); - x_51 = x_36; -} else { - lean_dec_ref(x_36); - x_51 = lean_box(0); +lean_object* x_110; lean_object* x_111; +x_110 = lean_ctor_get(x_104, 1); +lean_inc(x_110); +lean_dec(x_104); +x_111 = lean_ctor_get(x_105, 0); +lean_inc(x_111); +lean_dec(x_105); +x_6 = x_111; +x_11 = x_110; +goto _start; } -if (lean_is_scalar(x_51)) { - x_52 = lean_alloc_ctor(1, 2, 0); -} else { - x_52 = x_51; } -lean_ctor_set(x_52, 0, x_49); -lean_ctor_set(x_52, 1, x_50); -return x_52; +else +{ +lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +if (lean_is_scalar(x_98)) { + x_113 = lean_alloc_ctor(0, 2, 0); +} else { + x_113 = x_98; } +lean_ctor_set(x_113, 0, x_96); +lean_ctor_set(x_113, 1, x_97); +if (lean_is_scalar(x_95)) { + x_114 = lean_alloc_ctor(0, 2, 0); +} else { + x_114 = x_95; } +lean_ctor_set(x_114, 0, x_94); +lean_ctor_set(x_114, 1, x_113); +x_115 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___closed__1; +x_116 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_116, 0, x_115); +lean_ctor_set(x_116, 1, x_114); +x_117 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_117, 0, x_116); +lean_ctor_set(x_117, 1, x_93); +return x_117; } } -LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___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) { -_start: -{ -uint8_t x_7; -x_7 = lean_usize_dec_eq(x_5, x_6); -if (x_7 == 0) -{ -lean_object* x_8; lean_object* x_9; -x_8 = lean_array_uget(x_4, x_5); -x_9 = l_Lean_RBNode_findCore___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__2(x_1, x_8); -if (lean_obj_tag(x_9) == 0) -{ -size_t x_10; size_t x_11; -lean_dec(x_8); -x_10 = 1; -x_11 = lean_usize_add(x_5, x_10); -x_5 = x_11; -goto _start; } else { -lean_object* x_13; +uint8_t x_118; +lean_dec(x_10); lean_dec(x_9); -x_13 = l_Lean_RBNode_findCore___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__2(x_2, x_8); lean_dec(x_8); -if (lean_obj_tag(x_13) == 0) +lean_dec(x_7); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_118 = !lean_is_exclusive(x_22); +if (x_118 == 0) { -uint8_t x_14; -x_14 = 1; -return x_14; +return x_22; } else { -size_t x_15; size_t x_16; -lean_dec(x_13); -x_15 = 1; -x_16 = lean_usize_add(x_5, x_15); -x_5 = x_16; -goto _start; +lean_object* x_119; lean_object* x_120; lean_object* x_121; +x_119 = lean_ctor_get(x_22, 0); +x_120 = lean_ctor_get(x_22, 1); +lean_inc(x_120); +lean_inc(x_119); +lean_dec(x_22); +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_18; -x_18 = 0; -return x_18; +lean_object* x_122; lean_object* x_123; uint8_t x_124; lean_object* x_125; lean_object* x_126; size_t x_127; size_t x_128; lean_object* x_129; +x_122 = lean_ctor_get(x_14, 1); +lean_inc(x_122); +lean_dec(x_14); +x_123 = lean_box(0); +x_124 = 0; +x_125 = lean_box(x_124); +x_126 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_126, 0, x_125); +lean_ctor_set(x_126, 1, x_122); +lean_ctor_set(x_12, 1, x_126); +x_127 = lean_array_size(x_2); +x_128 = 0; +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_4); +lean_inc(x_3); +x_129 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__14___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__15(x_2, x_3, x_4, x_123, x_2, x_127, x_128, x_12, x_7, x_8, x_9, x_10, x_11); +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; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; uint8_t x_141; +x_130 = lean_ctor_get(x_129, 0); +lean_inc(x_130); +x_131 = lean_ctor_get(x_130, 1); +lean_inc(x_131); +x_132 = lean_ctor_get(x_129, 1); +lean_inc(x_132); +if (lean_is_exclusive(x_129)) { + lean_ctor_release(x_129, 0); + lean_ctor_release(x_129, 1); + x_133 = x_129; +} else { + lean_dec_ref(x_129); + x_133 = lean_box(0); } +x_134 = lean_ctor_get(x_130, 0); +lean_inc(x_134); +if (lean_is_exclusive(x_130)) { + lean_ctor_release(x_130, 0); + lean_ctor_release(x_130, 1); + x_135 = x_130; +} else { + lean_dec_ref(x_130); + x_135 = lean_box(0); } +x_136 = lean_ctor_get(x_131, 0); +lean_inc(x_136); +x_137 = lean_ctor_get(x_131, 1); +lean_inc(x_137); +if (lean_is_exclusive(x_131)) { + lean_ctor_release(x_131, 0); + lean_ctor_release(x_131, 1); + x_138 = x_131; +} else { + lean_dec_ref(x_131); + x_138 = lean_box(0); } -LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___spec__1___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___spec__2(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; -x_7 = lean_array_uget(x_3, x_4); -x_8 = l_Lean_RBNode_findCore___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__2(x_1, x_7); -if (lean_obj_tag(x_8) == 0) +x_139 = lean_unsigned_to_nat(0u); +x_140 = l_Lean_RBNode_fold___at_Lean_RBMap_size___spec__1___rarg(x_139, x_134); +x_141 = lean_nat_dec_eq(x_140, x_1); +lean_dec(x_140); +if (x_141 == 0) { -size_t x_9; size_t x_10; -lean_dec(x_7); -x_9 = 1; -x_10 = lean_usize_add(x_4, x_9); -x_4 = x_10; -goto _start; -} -else +lean_object* x_142; uint8_t x_143; lean_object* x_144; lean_object* x_145; +lean_dec(x_138); +lean_dec(x_135); +lean_dec(x_133); +x_142 = lean_box(0); +x_143 = lean_unbox(x_136); +lean_dec(x_136); +lean_inc(x_5); +x_144 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___lambda__1(x_143, x_137, x_134, x_5, x_142, x_7, x_8, x_9, x_10, x_132); +x_145 = lean_ctor_get(x_144, 0); +lean_inc(x_145); +if (lean_obj_tag(x_145) == 0) { -lean_object* x_12; +lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; +lean_dec(x_10); +lean_dec(x_9); lean_dec(x_8); -x_12 = l_Lean_RBNode_findCore___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__2(x_2, x_7); lean_dec(x_7); -if (lean_obj_tag(x_12) == 0) -{ -uint8_t x_13; -x_13 = 1; -return x_13; +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +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; +} else { + lean_dec_ref(x_144); + x_147 = lean_box(0); +} +x_148 = lean_ctor_get(x_145, 0); +lean_inc(x_148); +lean_dec(x_145); +if (lean_is_scalar(x_147)) { + x_149 = lean_alloc_ctor(0, 2, 0); +} else { + x_149 = x_147; +} +lean_ctor_set(x_149, 0, x_148); +lean_ctor_set(x_149, 1, x_146); +return x_149; } else { -size_t x_14; size_t x_15; -lean_dec(x_12); -x_14 = 1; -x_15 = lean_usize_add(x_4, x_14); -x_4 = x_15; +lean_object* x_150; lean_object* x_151; +x_150 = lean_ctor_get(x_144, 1); +lean_inc(x_150); +lean_dec(x_144); +x_151 = lean_ctor_get(x_145, 0); +lean_inc(x_151); +lean_dec(x_145); +x_6 = x_151; +x_11 = x_150; goto _start; } } -} else { -uint8_t x_17; -x_17 = 0; -return x_17; -} +lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; +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); +if (lean_is_scalar(x_138)) { + x_153 = lean_alloc_ctor(0, 2, 0); +} else { + x_153 = x_138; } +lean_ctor_set(x_153, 0, x_136); +lean_ctor_set(x_153, 1, x_137); +if (lean_is_scalar(x_135)) { + x_154 = lean_alloc_ctor(0, 2, 0); +} else { + x_154 = x_135; } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = lean_box(0); -x_2 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; -x_3 = l_Lean_Meta_Grind_NormalizePattern_main___closed__1; -x_4 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_4, 0, x_2); -lean_ctor_set(x_4, 1, x_1); -lean_ctor_set(x_4, 2, x_3); -return x_4; +lean_ctor_set(x_154, 0, x_134); +lean_ctor_set(x_154, 1, x_153); +x_155 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___closed__1; +x_156 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_156, 0, x_155); +lean_ctor_set(x_156, 1, x_154); +if (lean_is_scalar(x_133)) { + x_157 = lean_alloc_ctor(0, 2, 0); +} else { + x_157 = x_133; } +lean_ctor_set(x_157, 0, x_156); +lean_ctor_set(x_157, 1, x_132); +return x_157; } -LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; -x_4 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___closed__1; -x_5 = l_Lean_CollectFVars_main(x_3, x_4); -x_6 = lean_ctor_get(x_5, 2); -lean_inc(x_6); -lean_dec(x_5); -x_7 = lean_array_get_size(x_6); -x_8 = lean_unsigned_to_nat(0u); -x_9 = lean_nat_dec_lt(x_8, x_7); -if (x_9 == 0) -{ -uint8_t x_10; -lean_dec(x_7); -lean_dec(x_6); -x_10 = 1; -return x_10; } else { -size_t x_11; size_t x_12; uint8_t x_13; -x_11 = 0; -x_12 = lean_usize_of_nat(x_7); +lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_7); -x_13 = l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___spec__1___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___spec__2(x_1, x_2, x_6, x_11, x_12); -lean_dec(x_6); -if (x_13 == 0) -{ -uint8_t x_14; -x_14 = 1; -return x_14; -} -else -{ -uint8_t x_15; -x_15 = 0; -return x_15; -} -} -} -} -LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___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_7; size_t x_8; uint8_t x_9; lean_object* x_10; -x_7 = lean_unbox_usize(x_5); lean_dec(x_5); -x_8 = lean_unbox_usize(x_6); -lean_dec(x_6); -x_9 = l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___spec__1(x_1, x_2, x_3, x_4, x_7, x_8); lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_10 = lean_box(x_9); -return x_10; -} +x_158 = lean_ctor_get(x_129, 0); +lean_inc(x_158); +x_159 = lean_ctor_get(x_129, 1); +lean_inc(x_159); +if (lean_is_exclusive(x_129)) { + lean_ctor_release(x_129, 0); + lean_ctor_release(x_129, 1); + x_160 = x_129; +} else { + lean_dec_ref(x_129); + x_160 = lean_box(0); } -LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___spec__1___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: -{ -size_t x_6; size_t x_7; 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___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___spec__1___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___spec__2(x_1, x_2, x_3, x_6, x_7); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_9 = lean_box(x_8); -return x_9; +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; } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___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___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars(x_1, x_2, x_3); -lean_dec(x_2); -lean_dec(x_1); -x_5 = lean_box(x_4); -return x_5; } } -static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__1___closed__1() { -_start: +else { -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_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; uint8_t x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; size_t x_171; size_t x_172; lean_object* x_173; +x_162 = lean_ctor_get(x_12, 1); +x_163 = lean_ctor_get(x_12, 0); +lean_inc(x_162); +lean_inc(x_163); +lean_dec(x_12); +x_164 = lean_ctor_get(x_162, 1); +lean_inc(x_164); +if (lean_is_exclusive(x_162)) { + lean_ctor_release(x_162, 0); + lean_ctor_release(x_162, 1); + x_165 = x_162; +} else { + lean_dec_ref(x_162); + x_165 = lean_box(0); } +x_166 = lean_box(0); +x_167 = 0; +x_168 = lean_box(x_167); +if (lean_is_scalar(x_165)) { + x_169 = lean_alloc_ctor(0, 2, 0); +} else { + x_169 = x_165; } -static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__1___closed__2() { -_start: +lean_ctor_set(x_169, 0, x_168); +lean_ctor_set(x_169, 1, x_164); +x_170 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_170, 0, x_163); +lean_ctor_set(x_170, 1, x_169); +x_171 = lean_array_size(x_2); +x_172 = 0; +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_4); +lean_inc(x_3); +x_173 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__14___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__15(x_2, x_3, x_4, x_166, x_2, x_171, x_172, x_170, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_173) == 0) { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__1___closed__1; -x_2 = lean_box(0); -x_3 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set(x_3, 1, x_2); -return x_3; -} +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; uint8_t x_185; +x_174 = lean_ctor_get(x_173, 0); +lean_inc(x_174); +x_175 = lean_ctor_get(x_174, 1); +lean_inc(x_175); +x_176 = lean_ctor_get(x_173, 1); +lean_inc(x_176); +if (lean_is_exclusive(x_173)) { + lean_ctor_release(x_173, 0); + lean_ctor_release(x_173, 1); + x_177 = x_173; +} else { + lean_dec_ref(x_173); + x_177 = lean_box(0); } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___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, size_t x_7, size_t x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { -_start: -{ -uint8_t x_15; -x_15 = lean_usize_dec_lt(x_8, x_7); -if (x_15 == 0) -{ -lean_object* x_16; -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_5); -x_16 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_16, 0, x_9); -lean_ctor_set(x_16, 1, x_14); -return x_16; +x_178 = lean_ctor_get(x_174, 0); +lean_inc(x_178); +if (lean_is_exclusive(x_174)) { + lean_ctor_release(x_174, 0); + lean_ctor_release(x_174, 1); + x_179 = x_174; +} else { + lean_dec_ref(x_174); + x_179 = lean_box(0); } -else -{ -lean_object* x_17; lean_object* x_18; -lean_dec(x_9); -x_17 = lean_array_uget(x_6, x_8); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -x_18 = lean_infer_type(x_17, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_18) == 0) -{ -uint8_t x_19; -x_19 = !lean_is_exclusive(x_18); -if (x_19 == 0) +x_180 = lean_ctor_get(x_175, 0); +lean_inc(x_180); +x_181 = lean_ctor_get(x_175, 1); +lean_inc(x_181); +if (lean_is_exclusive(x_175)) { + lean_ctor_release(x_175, 0); + lean_ctor_release(x_175, 1); + x_182 = x_175; +} else { + lean_dec_ref(x_175); + x_182 = lean_box(0); +} +x_183 = lean_unsigned_to_nat(0u); +x_184 = l_Lean_RBNode_fold___at_Lean_RBMap_size___spec__1___rarg(x_183, x_178); +x_185 = lean_nat_dec_eq(x_184, x_1); +lean_dec(x_184); +if (x_185 == 0) { -lean_object* x_20; lean_object* x_21; uint8_t x_22; -x_20 = lean_ctor_get(x_18, 0); -x_21 = lean_ctor_get(x_18, 1); -x_22 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars(x_1, x_2, x_20); -if (x_22 == 0) +lean_object* x_186; uint8_t x_187; lean_object* x_188; lean_object* x_189; +lean_dec(x_182); +lean_dec(x_179); +lean_dec(x_177); +x_186 = lean_box(0); +x_187 = lean_unbox(x_180); +lean_dec(x_180); +lean_inc(x_5); +x_188 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___lambda__1(x_187, x_181, x_178, x_5, x_186, x_7, x_8, x_9, x_10, x_176); +x_189 = lean_ctor_get(x_188, 0); +lean_inc(x_189); +if (lean_obj_tag(x_189) == 0) { -lean_object* x_23; -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); +lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); lean_dec(x_5); -x_23 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__1___closed__2; -lean_ctor_set(x_18, 0, x_23); -return x_18; +lean_dec(x_4); +lean_dec(x_3); +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); +} +x_192 = lean_ctor_get(x_189, 0); +lean_inc(x_192); +lean_dec(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); +return x_193; } else { -size_t x_24; size_t x_25; -lean_free_object(x_18); -x_24 = 1; -x_25 = lean_usize_add(x_8, x_24); -lean_inc(x_5); -{ -size_t _tmp_7 = x_25; -lean_object* _tmp_8 = x_5; -lean_object* _tmp_13 = x_21; -x_8 = _tmp_7; -x_9 = _tmp_8; -x_14 = _tmp_13; -} +lean_object* x_194; lean_object* x_195; +x_194 = lean_ctor_get(x_188, 1); +lean_inc(x_194); +lean_dec(x_188); +x_195 = lean_ctor_get(x_189, 0); +lean_inc(x_195); +lean_dec(x_189); +x_6 = x_195; +x_11 = x_194; goto _start; } } else { -lean_object* x_27; lean_object* x_28; uint8_t x_29; -x_27 = lean_ctor_get(x_18, 0); -x_28 = lean_ctor_get(x_18, 1); -lean_inc(x_28); -lean_inc(x_27); -lean_dec(x_18); -x_29 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars(x_1, x_2, x_27); -if (x_29 == 0) -{ -lean_object* x_30; lean_object* x_31; -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); +lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); lean_dec(x_5); -x_30 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__1___closed__2; -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; +lean_dec(x_4); +lean_dec(x_3); +if (lean_is_scalar(x_182)) { + x_197 = lean_alloc_ctor(0, 2, 0); +} else { + x_197 = x_182; } -else -{ -size_t x_32; size_t x_33; -x_32 = 1; -x_33 = lean_usize_add(x_8, x_32); -lean_inc(x_5); -{ -size_t _tmp_7 = x_33; -lean_object* _tmp_8 = x_5; -lean_object* _tmp_13 = x_28; -x_8 = _tmp_7; -x_9 = _tmp_8; -x_14 = _tmp_13; +lean_ctor_set(x_197, 0, x_180); +lean_ctor_set(x_197, 1, x_181); +if (lean_is_scalar(x_179)) { + x_198 = lean_alloc_ctor(0, 2, 0); +} else { + x_198 = x_179; } -goto _start; +lean_ctor_set(x_198, 0, x_178); +lean_ctor_set(x_198, 1, x_197); +x_199 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___closed__1; +x_200 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_200, 0, x_199); +lean_ctor_set(x_200, 1, x_198); +if (lean_is_scalar(x_177)) { + x_201 = lean_alloc_ctor(0, 2, 0); +} else { + x_201 = x_177; } +lean_ctor_set(x_201, 0, x_200); +lean_ctor_set(x_201, 1, x_176); +return x_201; } } else { -uint8_t x_35; -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); +lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); lean_dec(x_5); -x_35 = !lean_is_exclusive(x_18); -if (x_35 == 0) -{ -return x_18; +lean_dec(x_4); +lean_dec(x_3); +x_202 = lean_ctor_get(x_173, 0); +lean_inc(x_202); +x_203 = lean_ctor_get(x_173, 1); +lean_inc(x_203); +if (lean_is_exclusive(x_173)) { + lean_ctor_release(x_173, 0); + lean_ctor_release(x_173, 1); + x_204 = x_173; +} else { + lean_dec_ref(x_173); + x_204 = lean_box(0); } -else -{ -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_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; +if (lean_is_scalar(x_204)) { + x_205 = lean_alloc_ctor(1, 2, 0); +} else { + x_205 = x_204; } +lean_ctor_set(x_205, 0, x_202); +lean_ctor_set(x_205, 1, x_203); +return x_205; } } } } -static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__1() { +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___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) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("semiOutParam", 12, 12); -return x_1; -} -} -static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__2() { -_start: +lean_object* x_15; uint8_t x_16; +x_15 = lean_ctor_get(x_5, 1); +x_16 = lean_nat_dec_lt(x_7, x_15); +if (x_16 == 0) { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__1; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; +lean_object* x_17; +lean_dec(x_7); +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_6); +lean_ctor_set(x_17, 1, x_14); +return x_17; } +else +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_18 = lean_array_fget(x_1, x_7); +x_19 = l_Lean_Expr_fvarId_x21(x_18); +lean_dec(x_18); +x_20 = l_Lean_RBNode_findCore___at_Lean_Meta_Closure_process___spec__1(x_3, x_19); +lean_dec(x_19); +if (lean_obj_tag(x_20) == 0) +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +lean_inc(x_7); +x_21 = lean_array_push(x_6, x_7); +x_22 = lean_ctor_get(x_5, 2); +x_23 = lean_nat_add(x_7, x_22); +lean_dec(x_7); +x_6 = x_21; +x_7 = x_23; +x_8 = lean_box(0); +x_9 = lean_box(0); +goto _start; } -static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__3() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("outParam", 8, 8); -return x_1; +lean_object* x_25; lean_object* x_26; +lean_dec(x_20); +x_25 = lean_ctor_get(x_5, 2); +x_26 = lean_nat_add(x_7, x_25); +lean_dec(x_7); +x_7 = x_26; +x_8 = lean_box(0); +x_9 = lean_box(0); +goto _start; } } -static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__4() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__3; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___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, 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_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__18___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__19(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { -uint8_t x_15; -x_15 = lean_usize_dec_lt(x_8, x_7); +lean_object* x_14; uint8_t x_15; +x_14 = lean_ctor_get(x_4, 1); +x_15 = lean_nat_dec_lt(x_6, x_14); if (x_15 == 0) { lean_object* x_16; -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_4); +lean_dec(x_6); x_16 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_16, 0, x_9); -lean_ctor_set(x_16, 1, x_14); +lean_ctor_set(x_16, 0, x_5); +lean_ctor_set(x_16, 1, x_13); return x_16; } else { -lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_27; -x_17 = lean_array_uget(x_6, x_8); -x_27 = !lean_is_exclusive(x_9); -if (x_27 == 0) -{ -lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; -x_28 = lean_ctor_get(x_9, 1); -x_29 = lean_ctor_get(x_9, 0); -lean_dec(x_29); -x_30 = lean_ctor_get(x_28, 0); -lean_inc(x_30); -x_31 = lean_ctor_get(x_28, 1); -lean_inc(x_31); -x_32 = lean_ctor_get(x_28, 2); -lean_inc(x_32); -x_33 = lean_nat_dec_lt(x_31, x_32); -if (x_33 == 0) -{ -lean_object* x_34; -lean_dec(x_32); -lean_dec(x_31); -lean_dec(x_30); +lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_17 = lean_array_fget(x_1, x_6); +x_18 = l_Lean_Expr_fvarId_x21(x_17); lean_dec(x_17); -lean_inc(x_4); -lean_ctor_set(x_9, 0, x_4); -x_34 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_34, 0, x_9); -x_18 = x_34; -x_19 = x_14; -goto block_26; +x_19 = l_Lean_RBNode_findCore___at_Lean_Meta_Closure_process___spec__1(x_2, x_18); +lean_dec(x_18); +if (lean_obj_tag(x_19) == 0) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; +lean_inc(x_6); +x_20 = lean_array_push(x_5, x_6); +x_21 = lean_ctor_get(x_4, 2); +x_22 = lean_nat_add(x_6, x_21); +lean_dec(x_6); +x_5 = x_20; +x_6 = x_22; +x_7 = lean_box(0); +x_8 = lean_box(0); +goto _start; } else { -uint8_t x_35; -x_35 = !lean_is_exclusive(x_28); -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; -x_36 = lean_ctor_get(x_28, 2); -lean_dec(x_36); -x_37 = lean_ctor_get(x_28, 1); -lean_dec(x_37); -x_38 = lean_ctor_get(x_28, 0); -lean_dec(x_38); -x_39 = lean_array_fget(x_30, x_31); -x_40 = lean_unsigned_to_nat(1u); -x_41 = lean_nat_add(x_31, x_40); -lean_dec(x_31); -lean_ctor_set(x_28, 1, x_41); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -x_42 = lean_infer_type(x_17, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(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_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_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__2; -x_46 = l_Lean_Expr_isAppOf(x_43, x_45); -if (x_46 == 0) -{ -lean_object* x_47; uint8_t x_48; -x_47 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__4; -x_48 = l_Lean_Expr_isAppOf(x_43, x_47); -lean_dec(x_43); -if (x_48 == 0) +lean_object* x_24; lean_object* x_25; +lean_dec(x_19); +x_24 = lean_ctor_get(x_4, 2); +x_25 = lean_nat_add(x_6, x_24); +lean_dec(x_6); +x_6 = x_25; +x_7 = lean_box(0); +x_8 = lean_box(0); +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__20(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { +_start: { -uint8_t x_49; -x_49 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars(x_1, x_2, x_39); -if (x_49 == 0) +uint8_t x_5; +x_5 = lean_usize_dec_eq(x_2, x_3); +if (x_5 == 0) { -lean_object* x_50; lean_object* x_51; -x_50 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__1___closed__1; -lean_ctor_set(x_9, 0, x_50); -x_51 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_51, 0, x_9); -x_18 = x_51; -x_19 = x_44; -goto block_26; +size_t x_6; size_t x_7; lean_object* x_8; lean_object* x_9; +x_6 = 1; +x_7 = lean_usize_sub(x_2, x_6); +x_8 = lean_array_uget(x_1, x_7); +x_9 = l_Std_DHashMap_Internal_AssocList_foldrM___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__2(x_4, x_8); +lean_dec(x_8); +lean_dec(x_4); +x_2 = x_7; +x_4 = x_9; +goto _start; } else { -lean_object* x_52; -lean_inc(x_4); -lean_ctor_set(x_9, 0, x_4); -x_52 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_52, 0, x_9); -x_18 = x_52; -x_19 = x_44; -goto block_26; +return x_4; +} +} } +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_11 = lean_array_mk(x_1); +x_12 = lean_unsigned_to_nat(0u); +x_13 = lean_unsigned_to_nat(1u); +x_14 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_14, 0, x_12); +lean_ctor_set(x_14, 1, x_2); +lean_ctor_set(x_14, 2, x_13); +x_15 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__18___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__19(x_3, x_4, x_14, x_14, x_11, x_12, lean_box(0), lean_box(0), x_6, x_7, x_8, x_9, x_10); +lean_dec(x_14); +x_16 = !lean_is_exclusive(x_15); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_17 = lean_ctor_get(x_15, 0); +x_18 = lean_array_to_list(x_17); +x_19 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_15, 0, x_19); +return x_15; } else { -lean_object* x_53; -lean_dec(x_39); -lean_inc(x_4); -lean_ctor_set(x_9, 0, x_4); -x_53 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_53, 0, x_9); -x_18 = x_53; -x_19 = x_44; -goto block_26; +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_20 = lean_ctor_get(x_15, 0); +x_21 = lean_ctor_get(x_15, 1); +lean_inc(x_21); +lean_inc(x_20); +lean_dec(x_15); +x_22 = lean_array_to_list(x_20); +x_23 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_23, 0, 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 +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: { -lean_object* x_54; -lean_dec(x_43); -lean_dec(x_39); +lean_object* x_14; uint8_t x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_14 = lean_box(0); +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_1); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_7); +lean_ctor_set(x_18, 1, 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); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); lean_inc(x_4); -lean_ctor_set(x_9, 0, x_4); -x_54 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_54, 0, x_9); -x_18 = x_54; -x_19 = x_44; -goto block_26; -} +x_20 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__17(x_2, x_3, x_4, x_5, x_14, x_19, x_9, x_10, x_11, x_12, x_13); +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_21, 1); +lean_inc(x_22); +x_23 = lean_ctor_get(x_21, 0); +lean_inc(x_23); +lean_dec(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; +x_24 = lean_ctor_get(x_20, 1); +lean_inc(x_24); +lean_dec(x_20); +x_25 = lean_ctor_get(x_22, 0); +lean_inc(x_25); +lean_dec(x_22); +x_26 = lean_box(0); +x_27 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__1(x_4, x_6, x_3, x_25, x_26, x_9, x_10, x_11, x_12, x_24); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_25); +return x_27; } else { -uint8_t x_55; -lean_dec(x_28); -lean_dec(x_39); -lean_free_object(x_9); -lean_dec(x_13); +uint8_t x_28; +lean_dec(x_22); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_6); lean_dec(x_4); -x_55 = !lean_is_exclusive(x_42); -if (x_55 == 0) +x_28 = !lean_is_exclusive(x_20); +if (x_28 == 0) { -return x_42; +lean_object* x_29; lean_object* x_30; +x_29 = lean_ctor_get(x_20, 0); +lean_dec(x_29); +x_30 = lean_ctor_get(x_23, 0); +lean_inc(x_30); +lean_dec(x_23); +lean_ctor_set(x_20, 0, x_30); +return x_20; } else { -lean_object* x_56; lean_object* x_57; lean_object* x_58; -x_56 = lean_ctor_get(x_42, 0); -x_57 = lean_ctor_get(x_42, 1); -lean_inc(x_57); -lean_inc(x_56); -lean_dec(x_42); -x_58 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_58, 0, x_56); -lean_ctor_set(x_58, 1, x_57); -return x_58; +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_20, 1); +lean_inc(x_31); +lean_dec(x_20); +x_32 = lean_ctor_get(x_23, 0); +lean_inc(x_32); +lean_dec(x_23); +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_31); +return x_33; } } } else { -lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; -lean_dec(x_28); -x_59 = lean_array_fget(x_30, x_31); -x_60 = lean_unsigned_to_nat(1u); -x_61 = lean_nat_add(x_31, x_60); -lean_dec(x_31); -x_62 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_62, 0, x_30); -lean_ctor_set(x_62, 1, x_61); -lean_ctor_set(x_62, 2, x_32); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -x_63 = lean_infer_type(x_17, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_63) == 0) -{ -lean_object* x_64; lean_object* x_65; lean_object* x_66; uint8_t 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 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__2; -x_67 = l_Lean_Expr_isAppOf(x_64, x_66); -if (x_67 == 0) -{ -lean_object* x_68; uint8_t x_69; -x_68 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__4; -x_69 = l_Lean_Expr_isAppOf(x_64, x_68); -lean_dec(x_64); -if (x_69 == 0) -{ -uint8_t x_70; -x_70 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars(x_1, x_2, x_59); -if (x_70 == 0) +uint8_t x_34; +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_6); +lean_dec(x_4); +x_34 = !lean_is_exclusive(x_20); +if (x_34 == 0) { -lean_object* x_71; lean_object* x_72; -x_71 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__1___closed__1; -lean_ctor_set(x_9, 1, x_62); -lean_ctor_set(x_9, 0, x_71); -x_72 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_72, 0, x_9); -x_18 = x_72; -x_19 = x_65; -goto block_26; +return x_20; } else { -lean_object* x_73; -lean_inc(x_4); -lean_ctor_set(x_9, 1, x_62); -lean_ctor_set(x_9, 0, x_4); -x_73 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_73, 0, x_9); -x_18 = x_73; -x_19 = x_65; -goto block_26; +lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_35 = lean_ctor_get(x_20, 0); +x_36 = lean_ctor_get(x_20, 1); +lean_inc(x_36); +lean_inc(x_35); +lean_dec(x_20); +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_74; -lean_dec(x_59); -lean_inc(x_4); -lean_ctor_set(x_9, 1, x_62); -lean_ctor_set(x_9, 0, x_4); -x_74 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_74, 0, x_9); -x_18 = x_74; -x_19 = x_65; -goto block_26; } } -else +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3___closed__1() { +_start: { -lean_object* x_75; -lean_dec(x_64); -lean_dec(x_59); -lean_inc(x_4); -lean_ctor_set(x_9, 1, x_62); -lean_ctor_set(x_9, 0, x_4); -x_75 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_75, 0, x_9); -x_18 = x_75; -x_19 = x_65; -goto block_26; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("numParams == xs.size\n ", 25, 25); +return x_1; } } -else +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3___closed__2() { +_start: { -lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; -lean_dec(x_62); -lean_dec(x_59); -lean_free_object(x_9); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_4); -x_76 = lean_ctor_get(x_63, 0); -lean_inc(x_76); -x_77 = lean_ctor_get(x_63, 1); -lean_inc(x_77); -if (lean_is_exclusive(x_63)) { - lean_ctor_release(x_63, 0); - lean_ctor_release(x_63, 1); - x_78 = x_63; -} else { - lean_dec_ref(x_63); - 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; -} -} +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__3; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3___closed__1; +x_3 = lean_string_append(x_1, x_2); +return x_3; } } -else -{ -lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; uint8_t x_84; -x_80 = lean_ctor_get(x_9, 1); -lean_inc(x_80); -lean_dec(x_9); -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_80, 2); -lean_inc(x_83); -x_84 = lean_nat_dec_lt(x_82, x_83); -if (x_84 == 0) +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3___closed__3() { +_start: { -lean_object* x_85; lean_object* x_86; -lean_dec(x_83); -lean_dec(x_82); -lean_dec(x_81); -lean_dec(x_17); -lean_inc(x_4); -x_85 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_85, 0, x_4); -lean_ctor_set(x_85, 1, x_80); -x_86 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_86, 0, x_85); -x_18 = x_86; -x_19 = x_14; -goto block_26; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("_private.Lean.Meta.Tactic.Grind.EMatchTheorem.0.Lean.Meta.Grind.checkCoverage", 77, 77); +return x_1; } -else +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3___closed__4() { +_start: { -lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; -if (lean_is_exclusive(x_80)) { - lean_ctor_release(x_80, 0); - lean_ctor_release(x_80, 1); - lean_ctor_release(x_80, 2); - x_87 = x_80; -} else { - lean_dec_ref(x_80); - x_87 = lean_box(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_Lean_Meta_Grind_EMatchTheorems_insert___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3___closed__3; +x_3 = lean_unsigned_to_nat(358u); +x_4 = lean_unsigned_to_nat(4u); +x_5 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3___closed__2; +x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); +return x_6; } -x_88 = lean_array_fget(x_81, x_82); -x_89 = lean_unsigned_to_nat(1u); -x_90 = lean_nat_add(x_82, x_89); -lean_dec(x_82); -if (lean_is_scalar(x_87)) { - x_91 = lean_alloc_ctor(0, 3, 0); -} else { - x_91 = x_87; } -lean_ctor_set(x_91, 0, x_81); -lean_ctor_set(x_91, 1, x_90); -lean_ctor_set(x_91, 2, x_83); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -x_92 = lean_infer_type(x_17, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_92) == 0) +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___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_93; lean_object* x_94; lean_object* x_95; uint8_t x_96; -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_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__2; -x_96 = l_Lean_Expr_isAppOf(x_93, x_95); -if (x_96 == 0) +lean_object* x_10; uint8_t x_11; +x_10 = lean_array_get_size(x_3); +x_11 = lean_nat_dec_eq(x_1, x_10); +if (x_11 == 0) { -lean_object* x_97; uint8_t x_98; -x_97 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__4; -x_98 = l_Lean_Expr_isAppOf(x_93, x_97); -lean_dec(x_93); -if (x_98 == 0) +lean_object* x_12; lean_object* x_13; +lean_dec(x_10); +lean_dec(x_3); +x_12 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3___closed__4; +x_13 = l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__1(x_12, x_5, x_6, x_7, x_8, x_9); +return x_13; +} +else { -uint8_t x_99; -x_99 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars(x_1, x_2, x_88); -if (x_99 == 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_48; uint8_t x_49; +x_14 = lean_box(0); +x_15 = lean_ctor_get(x_2, 1); +x_16 = lean_array_get_size(x_15); +lean_inc(x_3); +x_17 = lean_array_to_list(x_3); +x_18 = l_List_mapTR_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__3(x_17, x_14); +x_19 = l_Lean_RBTree_ofList___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__4(x_18); +x_20 = lean_box(0); +x_48 = lean_unsigned_to_nat(0u); +x_49 = lean_nat_dec_lt(x_48, x_16); +if (x_49 == 0) { -lean_object* x_100; lean_object* x_101; lean_object* x_102; -x_100 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__1___closed__1; -x_101 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_101, 0, x_100); -lean_ctor_set(x_101, 1, x_91); -x_102 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_102, 0, x_101); -x_18 = x_102; -x_19 = x_94; -goto block_26; +lean_dec(x_16); +x_21 = x_14; +goto block_47; } else { -lean_object* x_103; lean_object* x_104; -lean_inc(x_4); -x_103 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_103, 0, x_4); -lean_ctor_set(x_103, 1, x_91); -x_104 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_104, 0, x_103); -x_18 = x_104; -x_19 = x_94; -goto block_26; +size_t x_50; size_t x_51; lean_object* x_52; +x_50 = lean_usize_of_nat(x_16); +lean_dec(x_16); +x_51 = 0; +x_52 = l_Array_foldrMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__20(x_15, x_50, x_51, x_14); +x_21 = x_52; +goto block_47; } +block_47: +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_22 = l_List_mapTR_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__5(x_1, x_3, x_10, x_21, x_14); +lean_inc(x_22); +x_23 = l_Lean_RBTree_ofList___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__4(x_22); +lean_inc(x_5); +lean_inc(x_23); +lean_inc(x_22); +x_24 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__8___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__9(x_14, x_19, x_20, x_22, x_22, x_22, x_23, lean_box(0), x_5, x_6, x_7, x_8, x_9); +lean_dec(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; lean_object* x_28; lean_object* x_29; uint8_t x_30; +x_26 = lean_ctor_get(x_24, 0); +x_27 = lean_ctor_get(x_24, 1); +x_28 = lean_unsigned_to_nat(0u); +x_29 = l_Lean_RBNode_fold___at_Lean_RBMap_size___spec__1___rarg(x_28, x_26); +x_30 = lean_nat_dec_eq(x_29, x_1); +lean_dec(x_29); +if (x_30 == 0) +{ +lean_object* x_31; lean_object* x_32; +lean_free_object(x_24); +x_31 = lean_box(0); +x_32 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__2(x_23, x_1, x_3, x_14, x_19, x_10, x_26, x_31, x_5, x_6, x_7, x_8, x_27); +lean_dec(x_3); +return x_32; } else { -lean_object* x_105; lean_object* x_106; -lean_dec(x_88); -lean_inc(x_4); -x_105 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_105, 0, x_4); -lean_ctor_set(x_105, 1, x_91); -x_106 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_106, 0, x_105); -x_18 = x_106; -x_19 = x_94; -goto block_26; +lean_object* x_33; +lean_dec(x_26); +lean_dec(x_23); +lean_dec(x_19); +lean_dec(x_10); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +x_33 = lean_box(0); +lean_ctor_set(x_24, 0, x_33); +return x_24; } } else { -lean_object* x_107; lean_object* x_108; -lean_dec(x_93); -lean_dec(x_88); -lean_inc(x_4); -x_107 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_107, 0, x_4); -lean_ctor_set(x_107, 1, x_91); -x_108 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_108, 0, x_107); -x_18 = x_108; -x_19 = x_94; -goto block_26; -} +lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; uint8_t x_38; +x_34 = lean_ctor_get(x_24, 0); +x_35 = lean_ctor_get(x_24, 1); +lean_inc(x_35); +lean_inc(x_34); +lean_dec(x_24); +x_36 = lean_unsigned_to_nat(0u); +x_37 = l_Lean_RBNode_fold___at_Lean_RBMap_size___spec__1___rarg(x_36, x_34); +x_38 = lean_nat_dec_eq(x_37, x_1); +lean_dec(x_37); +if (x_38 == 0) +{ +lean_object* x_39; lean_object* x_40; +x_39 = lean_box(0); +x_40 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__2(x_23, x_1, x_3, x_14, x_19, x_10, x_34, x_39, x_5, x_6, x_7, x_8, x_35); +lean_dec(x_3); +return x_40; } else { -lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; -lean_dec(x_91); -lean_dec(x_88); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); +lean_object* x_41; lean_object* x_42; +lean_dec(x_34); +lean_dec(x_23); +lean_dec(x_19); lean_dec(x_10); -lean_dec(x_4); -x_109 = lean_ctor_get(x_92, 0); -lean_inc(x_109); -x_110 = lean_ctor_get(x_92, 1); -lean_inc(x_110); -if (lean_is_exclusive(x_92)) { - lean_ctor_release(x_92, 0); - lean_ctor_release(x_92, 1); - x_111 = x_92; -} else { - lean_dec_ref(x_92); - x_111 = lean_box(0); -} -if (lean_is_scalar(x_111)) { - x_112 = lean_alloc_ctor(1, 2, 0); -} else { - x_112 = x_111; -} -lean_ctor_set(x_112, 0, x_109); -lean_ctor_set(x_112, 1, x_110); -return x_112; -} +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +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_35); +return x_42; } } -block_26: -{ -if (lean_obj_tag(x_18) == 0) -{ -lean_object* x_20; lean_object* x_21; -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_4); -x_20 = lean_ctor_get(x_18, 0); -lean_inc(x_20); -lean_dec(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; } else { -lean_object* x_22; size_t x_23; size_t x_24; -x_22 = lean_ctor_get(x_18, 0); -lean_inc(x_22); -lean_dec(x_18); -x_23 = 1; -x_24 = lean_usize_add(x_8, x_23); -x_8 = x_24; -x_9 = x_22; -x_14 = x_19; -goto _start; +uint8_t x_43; +lean_dec(x_23); +lean_dec(x_19); +lean_dec(x_10); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +x_43 = !lean_is_exclusive(x_24); +if (x_43 == 0) +{ +return x_24; } +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_24, 0); +x_45 = lean_ctor_get(x_24, 1); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_24); +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_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___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) { -_start: -{ -uint8_t x_7; lean_object* x_8; lean_object* x_9; -x_7 = 1; -x_8 = lean_box(x_7); -x_9 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_9, 0, x_8); -lean_ctor_set(x_9, 1, x_6); -return x_9; } } -static lean_object* _init_l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__2___closed__1() { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__1___boxed), 6, 0); -return x_1; -} -} -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, size_t x_6, lean_object* x_7, 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; +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +x_10 = lean_infer_type(x_1, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_10) == 0) { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; size_t x_18; lean_object* x_19; -x_14 = lean_array_get_size(x_1); -x_15 = lean_unsigned_to_nat(0u); -x_16 = l_Array_toSubarray___rarg(x_1, x_15, x_14); -lean_inc(x_2); -x_17 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_17, 0, x_2); -lean_ctor_set(x_17, 1, x_16); -x_18 = lean_array_size(x_7); -lean_inc(x_12); +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_ctor_get(x_10, 0); lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -x_19 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2(x_3, x_4, x_5, x_2, x_7, x_7, x_18, x_6, x_17, x_9, x_10, x_11, x_12, x_13); -if (lean_obj_tag(x_19) == 0) -{ -lean_object* x_20; lean_object* x_21; -x_20 = lean_ctor_get(x_19, 0); -lean_inc(x_20); -x_21 = lean_ctor_get(x_20, 0); -lean_inc(x_21); -lean_dec(x_20); -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_19, 1); -lean_inc(x_22); -lean_dec(x_19); -x_23 = l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__2___closed__1; -x_24 = lean_box(0); -x_25 = lean_apply_6(x_23, x_24, x_9, x_10, x_11, x_12, x_22); -return x_25; +x_12 = lean_ctor_get(x_10, 1); +lean_inc(x_12); +lean_dec(x_10); +lean_inc(x_2); +x_13 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_13, 0, x_2); +x_14 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3___boxed), 9, 2); +lean_closure_set(x_14, 0, x_2); +lean_closure_set(x_14, 1, x_3); +x_15 = 0; +x_16 = l_Lean_Meta_forallBoundedTelescope___at_Lean_Meta_arrowDomainsN___spec__6___rarg(x_11, x_13, x_14, x_15, x_5, x_6, x_7, x_8, x_12); +return x_16; } else { -uint8_t x_26; -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -x_26 = !lean_is_exclusive(x_19); -if (x_26 == 0) +uint8_t x_17; +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_17 = !lean_is_exclusive(x_10); +if (x_17 == 0) { -lean_object* x_27; lean_object* x_28; -x_27 = lean_ctor_get(x_19, 0); -lean_dec(x_27); -x_28 = lean_ctor_get(x_21, 0); -lean_inc(x_28); -lean_dec(x_21); -lean_ctor_set(x_19, 0, x_28); -return x_19; +return x_10; } else { -lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_29 = lean_ctor_get(x_19, 1); -lean_inc(x_29); -lean_dec(x_19); -x_30 = lean_ctor_get(x_21, 0); -lean_inc(x_30); -lean_dec(x_21); -x_31 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_31, 0, x_30); -lean_ctor_set(x_31, 1, x_29); -return x_31; +lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_18 = lean_ctor_get(x_10, 0); +x_19 = lean_ctor_get(x_10, 1); +lean_inc(x_19); +lean_inc(x_18); +lean_dec(x_10); +x_20 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_20, 0, x_18); +lean_ctor_set(x_20, 1, x_19); +return x_20; } } } -else +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage(lean_object* x_1, lean_object* x_2, lean_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_32; -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); +lean_object* x_9; uint8_t x_10; +x_9 = lean_ctor_get(x_3, 0); +lean_inc(x_9); +x_10 = lean_nat_dec_eq(x_9, x_2); lean_dec(x_9); -x_32 = !lean_is_exclusive(x_19); -if (x_32 == 0) +if (x_10 == 0) { -return x_19; +lean_object* x_11; lean_object* x_12; +x_11 = lean_box(0); +x_12 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__4(x_1, x_2, x_3, x_11, x_4, x_5, x_6, x_7, x_8); +return x_12; } else { -lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_33 = lean_ctor_get(x_19, 0); -x_34 = lean_ctor_get(x_19, 1); -lean_inc(x_34); -lean_inc(x_33); -lean_dec(x_19); -x_35 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -return x_35; +lean_object* x_13; lean_object* x_14; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_13 = lean_box(0); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_8); +return x_14; } } } +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldrM___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__2___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_DHashMap_Internal_AssocList_foldrM___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__2(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, size_t x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +} +LEAN_EXPORT lean_object* l_List_mapTR_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__5___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_15; -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -x_15 = lean_infer_type(x_1, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_15) == 0) +lean_object* x_6; +x_6 = l_List_mapTR_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__5(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: { -lean_object* x_16; 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; -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_unsigned_to_nat(0u); -x_19 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_2, x_18); -x_20 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_20, 0, x_19); -x_21 = lean_box_usize(x_8); -x_22 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__2___boxed), 13, 6); -lean_closure_set(x_22, 0, x_3); -lean_closure_set(x_22, 1, x_4); -lean_closure_set(x_22, 2, x_5); -lean_closure_set(x_22, 3, x_6); -lean_closure_set(x_22, 4, x_7); -lean_closure_set(x_22, 5, x_21); -x_23 = 0; -x_24 = l_Lean_Meta_forallBoundedTelescope___at_Lean_Meta_arrowDomainsN___spec__6___rarg(x_16, x_20, x_22, x_23, x_10, x_11, x_12, x_13, x_17); -return x_24; +size_t x_7; size_t x_8; lean_object* x_9; +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_9 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__6(x_1, x_2, x_3, x_7, x_8, x_6); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_9; } -else +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__6___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__7___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_25; +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___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__6___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__7(x_1, x_2, x_6, x_7, x_5); +lean_dec(x_2); +lean_dec(x_1); +return x_8; +} +} +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___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) { +_start: +{ +lean_object* x_15; +x_15 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___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); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_25 = !lean_is_exclusive(x_15); -if (x_25 == 0) -{ +lean_dec(x_2); return x_15; } -else -{ -lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_26 = lean_ctor_get(x_15, 0); -x_27 = lean_ctor_get(x_15, 1); -lean_inc(x_27); -lean_inc(x_26); -lean_dec(x_15); -x_28 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set(x_28, 1, x_27); -return x_28; -} -} } -} -static lean_object* _init_l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___closed__1() { +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__8___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__9___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { -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_object* x_14; +x_14 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__8___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__9(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_14; } } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___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: { -if (lean_obj_tag(x_5) == 5) -{ -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_5, 0); -lean_inc(x_13); -x_14 = lean_ctor_get(x_5, 1); -lean_inc(x_14); +size_t x_7; size_t x_8; lean_object* x_9; +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = lean_unbox_usize(x_5); lean_dec(x_5); -x_15 = lean_array_set(x_6, x_7, x_14); -x_16 = lean_unsigned_to_nat(1u); -x_17 = lean_nat_sub(x_7, x_16); -lean_dec(x_7); -x_5 = x_13; -x_6 = x_15; -x_7 = x_17; -goto _start; +x_9 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__10(x_1, x_2, x_3, x_7, x_8, x_6); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_9; } -else -{ -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_dec(x_7); -x_19 = lean_box(0); -x_20 = lean_box(0); -x_21 = lean_array_size(x_3); -x_22 = 0; -x_23 = l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___closed__1; -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -x_24 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__1(x_1, x_2, x_3, x_19, x_23, x_3, x_21, x_22, x_23, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_24) == 0) -{ -lean_object* x_25; lean_object* x_26; -x_25 = lean_ctor_get(x_24, 0); -lean_inc(x_25); -x_26 = lean_ctor_get(x_25, 0); -lean_inc(x_26); -lean_dec(x_25); -if (lean_obj_tag(x_26) == 0) +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__10___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__11___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_27; lean_object* x_28; lean_object* x_29; -x_27 = lean_ctor_get(x_24, 1); -lean_inc(x_27); -lean_dec(x_24); -x_28 = lean_box(0); -x_29 = l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__3(x_5, x_4, x_6, x_20, x_1, x_2, x_19, x_22, x_28, x_8, x_9, x_10, x_11, x_27); -return x_29; +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___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__10___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__11(x_1, x_2, x_6, x_7, x_5); +lean_dec(x_2); +lean_dec(x_1); +return x_8; } -else +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___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: { -uint8_t x_30; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_6); +size_t x_7; size_t x_8; lean_object* x_9; +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = lean_unbox_usize(x_5); lean_dec(x_5); +x_9 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__12(x_1, x_2, x_3, x_7, x_8, x_6); +lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_30 = !lean_is_exclusive(x_24); -if (x_30 == 0) -{ -lean_object* x_31; lean_object* x_32; -x_31 = lean_ctor_get(x_24, 0); -lean_dec(x_31); -x_32 = lean_ctor_get(x_26, 0); -lean_inc(x_32); -lean_dec(x_26); -lean_ctor_set(x_24, 0, x_32); -return x_24; +return x_9; } -else -{ -lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_33 = lean_ctor_get(x_24, 1); -lean_inc(x_33); -lean_dec(x_24); -x_34 = lean_ctor_get(x_26, 0); -lean_inc(x_34); -lean_dec(x_26); -x_35 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_35, 0, x_34); -lean_ctor_set(x_35, 1, x_33); -return x_35; } +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__12___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__13___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +size_t x_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___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__12___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__13(x_1, x_2, x_6, x_7, x_5); +lean_dec(x_2); +lean_dec(x_1); +return x_8; } } -else +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__14___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: { -uint8_t x_36; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); +size_t x_15; size_t x_16; lean_object* x_17; +x_15 = lean_unbox_usize(x_7); +lean_dec(x_7); +x_16 = lean_unbox_usize(x_8); lean_dec(x_8); +x_17 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__14(x_1, x_2, x_3, x_4, x_5, x_6, x_15, x_16, x_9, x_10, x_11, x_12, x_13, x_14); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_2); +lean_dec(x_3); lean_dec(x_1); -x_36 = !lean_is_exclusive(x_24); -if (x_36 == 0) -{ -return x_24; -} -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_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; -} -} -} +return x_17; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__14___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__15___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_10 = lean_unsigned_to_nat(0u); -x_11 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_4, x_10); -x_12 = l_Lean_Meta_Grind_ppPattern___closed__5; -lean_inc(x_11); -x_13 = lean_mk_array(x_11, x_12); -x_14 = lean_unsigned_to_nat(1u); -x_15 = lean_nat_sub(x_11, x_14); -lean_dec(x_11); -lean_inc(x_4); -x_16 = l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3(x_1, x_2, x_3, x_4, x_4, x_13, x_15, x_5, x_6, x_7, x_8, x_9); +size_t x_14; size_t x_15; lean_object* x_16; +x_14 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_15 = lean_unbox_usize(x_7); +lean_dec(x_7); +x_16 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__14___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__15(x_1, x_2, x_3, x_4, x_5, x_14, x_15, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_5); lean_dec(x_4); +lean_dec(x_1); return x_16; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized(lean_object* x_1, lean_object* x_2, lean_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_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___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_9; uint8_t x_10; lean_object* x_11; -x_9 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___lambda__1___boxed), 9, 2); -lean_closure_set(x_9, 0, x_1); -lean_closure_set(x_9, 1, x_2); -x_10 = 0; -x_11 = l_Lean_Meta_forallTelescopeReducing___at_Lean_Meta_getParamNames___spec__2___rarg(x_3, x_9, x_10, x_4, x_5, x_6, x_7, x_8); -return x_11; +uint8_t x_11; lean_object* x_12; +x_11 = lean_unbox(x_1); +lean_dec(x_1); +x_12 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___lambda__1(x_11, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +return x_12; } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -size_t x_15; size_t x_16; lean_object* x_17; -x_15 = lean_unbox_usize(x_7); -lean_dec(x_7); -x_16 = lean_unbox_usize(x_8); -lean_dec(x_8); -x_17 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_15, x_16, x_9, x_10, x_11, x_12, x_13, x_14); -lean_dec(x_6); +lean_object* x_13; +x_13 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); lean_dec(x_4); -lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -return x_17; +return x_13; +} +} +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__17___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; +x_12 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__17(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_2); +lean_dec(x_1); +return x_12; } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___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) { _start: { -size_t x_15; size_t x_16; lean_object* x_17; -x_15 = lean_unbox_usize(x_7); -lean_dec(x_7); -x_16 = lean_unbox_usize(x_8); -lean_dec(x_8); -x_17 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_15, x_16, x_9, x_10, x_11, x_12, x_13, x_14); -lean_dec(x_6); +lean_object* x_15; +x_15 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___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); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); lean_dec(x_5); +lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -return x_17; +return x_15; } } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___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_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__18___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__19___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { -lean_object* x_7; -x_7 = l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6); -lean_dec(x_5); +lean_object* x_14; +x_14 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__18___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__19(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__20___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_foldrMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__20(x_1, x_5, x_6, x_4); +lean_dec(x_1); return x_7; } } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___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: { -size_t x_14; lean_object* x_15; -x_14 = lean_unbox_usize(x_6); -lean_dec(x_6); -x_15 = l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__2(x_1, x_2, x_3, x_4, x_5, x_14, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_object* x_11; +x_11 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); +lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -return x_15; +return x_11; } } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { -size_t x_15; lean_object* x_16; -x_15 = lean_unbox_usize(x_8); +lean_object* x_14; +x_14 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); lean_dec(x_8); -x_16 = l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_15, x_9, x_10, x_11, x_12, x_13, x_14); -lean_dec(x_9); +lean_dec(x_3); lean_dec(x_2); -return x_16; +return x_14; } } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___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_13; -x_13 = l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_object* x_10; +x_10 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); lean_dec(x_4); -lean_dec(x_3); -return x_13; +lean_dec(x_2); +lean_dec(x_1); +return x_10; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___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_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { lean_object* x_10; -x_10 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -lean_dec(x_3); +x_10 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_4); return x_10; } } -static lean_object* _init_l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__1___closed__1() { +static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___lambda__1___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Meta_instInhabitedMetaM___boxed), 5, 1); -lean_closure_set(x_1, 0, lean_box(0)); +x_1 = lean_mk_string_unchecked(" : ", 3, 3); return x_1; } } -LEAN_EXPORT lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___lambda__1___closed__2() { _start: { -lean_object* x_7; lean_object* x_8; lean_object* x_9; -x_7 = l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__1___closed__1; -x_8 = lean_panic_fn(x_7, x_1); -x_9 = lean_apply_5(x_8, x_2, x_3, x_4, x_5, x_6); -return x_9; +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___lambda__1___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } } -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldrM___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__2(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___lambda__1(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -if (lean_obj_tag(x_2) == 0) -{ +lean_object* x_11; lean_inc(x_1); -return x_1; +x_11 = lean_infer_type(x_1, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_11) == 0) +{ +uint8_t x_12; +x_12 = !lean_is_exclusive(x_11); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; 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 = l_Lean_MessageData_ofExpr(x_1); +lean_inc(x_2); +x_15 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_15, 0, x_2); +lean_ctor_set(x_15, 1, x_14); +x_16 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___lambda__1___closed__2; +x_17 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_17, 0, x_15); +lean_ctor_set(x_17, 1, x_16); +x_18 = l_Lean_MessageData_ofExpr(x_13); +x_19 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_19, 0, x_17); +lean_ctor_set(x_19, 1, x_18); +x_20 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_20, 0, x_19); +lean_ctor_set(x_20, 1, x_2); +x_21 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_21, 0, x_4); +lean_ctor_set(x_21, 1, x_20); +x_22 = lean_box(x_3); +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_21); +x_24 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set(x_11, 0, x_24); +return x_11; } else { -lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_3 = lean_ctor_get(x_2, 0); -x_4 = lean_ctor_get(x_2, 2); -x_5 = l_Std_DHashMap_Internal_AssocList_foldrM___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__2(x_1, x_4); -lean_inc(x_3); -x_6 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_6, 0, x_3); -lean_ctor_set(x_6, 1, x_5); -return x_6; -} +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; 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 = l_Lean_MessageData_ofExpr(x_1); +lean_inc(x_2); +x_28 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_28, 0, x_2); +lean_ctor_set(x_28, 1, x_27); +x_29 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___lambda__1___closed__2; +x_30 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_30, 0, x_28); +lean_ctor_set(x_30, 1, x_29); +x_31 = l_Lean_MessageData_ofExpr(x_25); +x_32 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_32, 0, x_30); +lean_ctor_set(x_32, 1, x_31); +x_33 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_2); +x_34 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_34, 0, x_4); +lean_ctor_set(x_34, 1, x_33); +x_35 = lean_box(x_3); +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set(x_36, 1, x_34); +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_26); +return x_38; } } -LEAN_EXPORT lean_object* l_List_mapTR_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__3(lean_object* x_1, lean_object* x_2) { -_start: +else { -if (lean_obj_tag(x_1) == 0) +uint8_t x_39; +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_39 = !lean_is_exclusive(x_11); +if (x_39 == 0) { -lean_object* x_3; -x_3 = l_List_reverse___rarg(x_2); -return x_3; +return x_11; } 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_Expr_fvarId_x21(x_5); -lean_dec(x_5); -lean_ctor_set(x_1, 1, x_2); -lean_ctor_set(x_1, 0, x_7); +lean_object* x_40; lean_object* x_41; lean_object* x_42; +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_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); +return x_42; +} +} +} +} +static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___closed__1() { +_start: { -lean_object* _tmp_0 = x_6; -lean_object* _tmp_1 = x_1; -x_1 = _tmp_0; -x_2 = _tmp_1; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("\n", 1, 1); +return x_1; } -goto _start; } -else +static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___closed__2() { +_start: { -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_9 = lean_ctor_get(x_1, 0); -x_10 = lean_ctor_get(x_1, 1); -lean_inc(x_10); -lean_inc(x_9); -lean_dec(x_1); -x_11 = l_Lean_Expr_fvarId_x21(x_9); -lean_dec(x_9); -x_12 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_12, 0, x_11); -lean_ctor_set(x_12, 1, x_2); -x_1 = x_10; -x_2 = x_12; -goto _start; +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___closed__1; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; } } +static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___closed__2; +x_2 = l_Lean_MessageData_ofFormat(x_1); +return x_2; } } -LEAN_EXPORT lean_object* l_Lean_RBTree_ofList___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__4(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { _start: { -if (lean_obj_tag(x_1) == 0) +lean_object* x_15; uint8_t x_16; +x_15 = lean_ctor_get(x_5, 1); +x_16 = lean_nat_dec_lt(x_7, x_15); +if (x_16 == 0) { -lean_object* x_2; -x_2 = lean_box(0); -return x_2; +lean_object* x_17; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_7); +lean_dec(x_3); +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_6); +lean_ctor_set(x_17, 1, x_14); +return x_17; } else { -lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_3 = lean_ctor_get(x_1, 0); -lean_inc(x_3); -x_4 = lean_ctor_get(x_1, 1); -lean_inc(x_4); -lean_dec(x_1); -x_5 = l_Lean_RBTree_ofList___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__4(x_4); -x_6 = lean_box(0); -x_7 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_5, x_3, x_6); -return x_7; -} -} -} -LEAN_EXPORT lean_object* l_List_mapTR_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: +uint8_t x_18; +x_18 = !lean_is_exclusive(x_6); +if (x_18 == 0) { -if (lean_obj_tag(x_4) == 0) +lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_19 = lean_ctor_get(x_6, 0); +x_20 = lean_ctor_get(x_6, 1); +x_21 = l_List_elem___at_Lean_Meta_Occurrences_contains___spec__1(x_7, x_1); +if (x_21 == 0) { -lean_object* x_6; -x_6 = l_List_reverse___rarg(x_5); -return x_6; +lean_object* x_22; lean_object* x_23; +x_22 = lean_ctor_get(x_5, 2); +x_23 = lean_nat_add(x_7, x_22); +lean_dec(x_7); +x_7 = x_23; +x_8 = lean_box(0); +x_9 = lean_box(0); +goto _start; } else { -uint8_t x_7; -x_7 = !lean_is_exclusive(x_4); -if (x_7 == 0) +lean_object* x_25; uint8_t x_26; +lean_free_object(x_6); +x_25 = lean_array_fget(x_2, x_7); +x_26 = lean_unbox(x_19); +if (x_26 == 0) { -lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; -x_8 = lean_ctor_get(x_4, 0); -x_9 = lean_ctor_get(x_4, 1); -x_10 = lean_nat_sub(x_1, x_8); -lean_dec(x_8); -x_11 = lean_unsigned_to_nat(1u); -x_12 = lean_nat_sub(x_10, x_11); -lean_dec(x_10); -x_13 = lean_nat_dec_lt(x_12, x_3); -if (x_13 == 0) +lean_object* x_27; lean_object* x_28; lean_object* x_29; uint8_t x_30; lean_object* x_31; +x_27 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___closed__3; +x_28 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_28, 0, x_20); +lean_ctor_set(x_28, 1, x_27); +x_29 = lean_box(0); +x_30 = lean_unbox(x_19); +lean_dec(x_19); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_3); +x_31 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___lambda__1(x_25, x_3, x_30, x_28, x_29, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_31) == 0) { -lean_object* x_14; lean_object* x_15; lean_object* x_16; +lean_object* x_32; +x_32 = lean_ctor_get(x_31, 0); +lean_inc(x_32); +if (lean_obj_tag(x_32) == 0) +{ +uint8_t x_33; +lean_dec(x_13); lean_dec(x_12); -x_14 = l_Lean_instInhabitedExpr; -x_15 = l_outOfBounds___rarg(x_14); -x_16 = l_Lean_Expr_fvarId_x21(x_15); -lean_dec(x_15); -lean_ctor_set(x_4, 1, x_5); -lean_ctor_set(x_4, 0, x_16); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_7); +lean_dec(x_3); +x_33 = !lean_is_exclusive(x_31); +if (x_33 == 0) { -lean_object* _tmp_3 = x_9; -lean_object* _tmp_4 = x_4; -x_4 = _tmp_3; -x_5 = _tmp_4; +lean_object* x_34; lean_object* x_35; +x_34 = lean_ctor_get(x_31, 0); +lean_dec(x_34); +x_35 = lean_ctor_get(x_32, 0); +lean_inc(x_35); +lean_dec(x_32); +lean_ctor_set(x_31, 0, x_35); +return x_31; +} +else +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_31, 1); +lean_inc(x_36); +lean_dec(x_31); +x_37 = lean_ctor_get(x_32, 0); +lean_inc(x_37); +lean_dec(x_32); +x_38 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_38, 0, x_37); +lean_ctor_set(x_38, 1, x_36); +return x_38; } +} +else +{ +lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_39 = lean_ctor_get(x_31, 1); +lean_inc(x_39); +lean_dec(x_31); +x_40 = lean_ctor_get(x_32, 0); +lean_inc(x_40); +lean_dec(x_32); +x_41 = lean_ctor_get(x_5, 2); +x_42 = lean_nat_add(x_7, x_41); +lean_dec(x_7); +x_6 = x_40; +x_7 = x_42; +x_8 = lean_box(0); +x_9 = lean_box(0); +x_14 = x_39; goto _start; } +} else { -lean_object* x_18; lean_object* x_19; -x_18 = lean_array_fget(x_2, x_12); +uint8_t x_44; +lean_dec(x_13); lean_dec(x_12); -x_19 = l_Lean_Expr_fvarId_x21(x_18); -lean_dec(x_18); -lean_ctor_set(x_4, 1, x_5); -lean_ctor_set(x_4, 0, x_19); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_7); +lean_dec(x_3); +x_44 = !lean_is_exclusive(x_31); +if (x_44 == 0) { -lean_object* _tmp_3 = x_9; -lean_object* _tmp_4 = x_4; -x_4 = _tmp_3; -x_5 = _tmp_4; +return x_31; +} +else +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_45 = lean_ctor_get(x_31, 0); +x_46 = lean_ctor_get(x_31, 1); +lean_inc(x_46); +lean_inc(x_45); +lean_dec(x_31); +x_47 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_47, 0, x_45); +lean_ctor_set(x_47, 1, x_46); +return x_47; } -goto _start; } } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; -x_21 = lean_ctor_get(x_4, 0); -x_22 = lean_ctor_get(x_4, 1); -lean_inc(x_22); -lean_inc(x_21); -lean_dec(x_4); -x_23 = lean_nat_sub(x_1, x_21); -lean_dec(x_21); -x_24 = lean_unsigned_to_nat(1u); -x_25 = lean_nat_sub(x_23, x_24); -lean_dec(x_23); -x_26 = lean_nat_dec_lt(x_25, x_3); -if (x_26 == 0) +uint8_t x_48; lean_object* x_49; lean_object* x_50; +lean_dec(x_19); +x_48 = 0; +x_49 = lean_box(0); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_3); +x_50 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___lambda__1(x_25, x_3, x_48, x_20, x_49, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_50) == 0) { -lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; -lean_dec(x_25); -x_27 = l_Lean_instInhabitedExpr; -x_28 = l_outOfBounds___rarg(x_27); -x_29 = l_Lean_Expr_fvarId_x21(x_28); -lean_dec(x_28); -x_30 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_30, 0, x_29); -lean_ctor_set(x_30, 1, x_5); -x_4 = x_22; -x_5 = x_30; -goto _start; +lean_object* x_51; +x_51 = lean_ctor_get(x_50, 0); +lean_inc(x_51); +if (lean_obj_tag(x_51) == 0) +{ +uint8_t x_52; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_7); +lean_dec(x_3); +x_52 = !lean_is_exclusive(x_50); +if (x_52 == 0) +{ +lean_object* x_53; lean_object* x_54; +x_53 = lean_ctor_get(x_50, 0); +lean_dec(x_53); +x_54 = lean_ctor_get(x_51, 0); +lean_inc(x_54); +lean_dec(x_51); +lean_ctor_set(x_50, 0, x_54); +return x_50; } else { -lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_32 = lean_array_fget(x_2, x_25); -lean_dec(x_25); -x_33 = l_Lean_Expr_fvarId_x21(x_32); -lean_dec(x_32); -x_34 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_34, 0, x_33); -lean_ctor_set(x_34, 1, x_5); -x_4 = x_22; -x_5 = x_34; -goto _start; -} -} -} +lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_55 = lean_ctor_get(x_50, 1); +lean_inc(x_55); +lean_dec(x_50); +x_56 = lean_ctor_get(x_51, 0); +lean_inc(x_56); +lean_dec(x_51); +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; } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { -_start: -{ -uint8_t x_7; -x_7 = lean_usize_dec_eq(x_4, x_5); -if (x_7 == 0) -{ -lean_object* x_8; size_t x_9; size_t x_10; lean_object* x_11; -x_8 = lean_array_uget(x_3, x_4); -x_9 = 1; -x_10 = lean_usize_add(x_4, x_9); -x_11 = l_Lean_RBNode_findCore___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__2(x_2, x_8); -if (lean_obj_tag(x_11) == 0) +else { -lean_dec(x_8); -x_4 = x_10; +lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; +x_58 = lean_ctor_get(x_50, 1); +lean_inc(x_58); +lean_dec(x_50); +x_59 = lean_ctor_get(x_51, 0); +lean_inc(x_59); +lean_dec(x_51); +x_60 = lean_ctor_get(x_5, 2); +x_61 = lean_nat_add(x_7, x_60); +lean_dec(x_7); +x_6 = x_59; +x_7 = x_61; +x_8 = lean_box(0); +x_9 = lean_box(0); +x_14 = x_58; goto _start; } +} else { -lean_object* x_13; lean_object* x_14; +uint8_t x_63; +lean_dec(x_13); +lean_dec(x_12); lean_dec(x_11); -x_13 = lean_box(0); -x_14 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_6, x_8, x_13); -x_4 = x_10; -x_6 = x_14; -goto _start; -} +lean_dec(x_10); +lean_dec(x_7); +lean_dec(x_3); +x_63 = !lean_is_exclusive(x_50); +if (x_63 == 0) +{ +return x_50; } else { -return x_6; +lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_64 = lean_ctor_get(x_50, 0); +x_65 = lean_ctor_get(x_50, 1); +lean_inc(x_65); +lean_inc(x_64); +lean_dec(x_50); +x_66 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_66, 0, x_64); +lean_ctor_set(x_66, 1, x_65); +return x_66; } } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__6___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__7(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) +} +} +else { -lean_object* x_7; size_t x_8; size_t x_9; lean_object* x_10; -x_7 = lean_array_uget(x_2, x_3); -x_8 = 1; -x_9 = lean_usize_add(x_3, x_8); -x_10 = l_Lean_RBNode_findCore___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__2(x_1, x_7); -if (lean_obj_tag(x_10) == 0) +lean_object* x_67; lean_object* x_68; uint8_t x_69; +x_67 = lean_ctor_get(x_6, 0); +x_68 = lean_ctor_get(x_6, 1); +lean_inc(x_68); +lean_inc(x_67); +lean_dec(x_6); +x_69 = l_List_elem___at_Lean_Meta_Occurrences_contains___spec__1(x_7, x_1); +if (x_69 == 0) { +lean_object* x_70; lean_object* x_71; lean_object* x_72; +x_70 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_70, 0, x_67); +lean_ctor_set(x_70, 1, x_68); +x_71 = lean_ctor_get(x_5, 2); +x_72 = lean_nat_add(x_7, x_71); lean_dec(x_7); -x_3 = x_9; +x_6 = x_70; +x_7 = x_72; +x_8 = lean_box(0); +x_9 = lean_box(0); goto _start; } else { -lean_object* x_12; lean_object* x_13; -lean_dec(x_10); -x_12 = lean_box(0); -x_13 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_5, x_7, x_12); -x_3 = x_9; -x_5 = x_13; -goto _start; -} -} -else +lean_object* x_74; uint8_t x_75; +x_74 = lean_array_fget(x_2, x_7); +x_75 = lean_unbox(x_67); +if (x_75 == 0) { -return x_5; -} -} -} -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___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) { -_start: +lean_object* x_76; lean_object* x_77; lean_object* x_78; uint8_t x_79; lean_object* x_80; +x_76 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___closed__3; +x_77 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_77, 0, x_68); +lean_ctor_set(x_77, 1, x_76); +x_78 = lean_box(0); +x_79 = lean_unbox(x_67); +lean_dec(x_67); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_3); +x_80 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___lambda__1(x_74, x_3, x_79, x_77, x_78, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_80) == 0) { -if (lean_obj_tag(x_7) == 0) +lean_object* x_81; +x_81 = lean_ctor_get(x_80, 0); +lean_inc(x_81); +if (lean_obj_tag(x_81) == 0) { -lean_object* x_15; +lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); lean_dec(x_10); -lean_dec(x_1); -x_15 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_15, 0, x_8); -lean_ctor_set(x_15, 1, x_14); -return x_15; +lean_dec(x_7); +lean_dec(x_3); +x_82 = lean_ctor_get(x_80, 1); +lean_inc(x_82); +if (lean_is_exclusive(x_80)) { + lean_ctor_release(x_80, 0); + lean_ctor_release(x_80, 1); + x_83 = x_80; +} else { + lean_dec_ref(x_80); + x_83 = lean_box(0); +} +x_84 = lean_ctor_get(x_81, 0); +lean_inc(x_84); +lean_dec(x_81); +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_16; lean_object* x_17; lean_object* x_18; -x_16 = lean_ctor_get(x_7, 0); -lean_inc(x_16); -x_17 = lean_ctor_get(x_7, 1); -lean_inc(x_17); +lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; +x_86 = lean_ctor_get(x_80, 1); +lean_inc(x_86); +lean_dec(x_80); +x_87 = lean_ctor_get(x_81, 0); +lean_inc(x_87); +lean_dec(x_81); +x_88 = lean_ctor_get(x_5, 2); +x_89 = lean_nat_add(x_7, x_88); lean_dec(x_7); -lean_inc(x_10); -x_18 = l_Lean_FVarId_getType(x_16, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_18) == 0) -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; -x_19 = lean_ctor_get(x_18, 0); -lean_inc(x_19); -x_20 = lean_ctor_get(x_18, 1); -lean_inc(x_20); -lean_dec(x_18); -x_21 = lean_box(0); -lean_inc(x_1); -x_22 = lean_array_mk(x_1); -x_23 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; -x_24 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_24, 0, x_23); -lean_ctor_set(x_24, 1, x_21); -lean_ctor_set(x_24, 2, x_22); -x_25 = l_Lean_CollectFVars_main(x_19, x_24); -x_26 = lean_ctor_get(x_25, 2); -lean_inc(x_26); -lean_dec(x_25); -x_27 = lean_array_get_size(x_26); -x_28 = lean_unsigned_to_nat(0u); -x_29 = lean_nat_dec_lt(x_28, x_27); -if (x_29 == 0) -{ -lean_dec(x_27); -lean_dec(x_26); -x_7 = x_17; +x_6 = x_87; +x_7 = x_89; +x_8 = lean_box(0); x_9 = lean_box(0); -x_14 = x_20; +x_14 = x_86; goto _start; } +} +else +{ +lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_7); +lean_dec(x_3); +x_91 = lean_ctor_get(x_80, 0); +lean_inc(x_91); +x_92 = lean_ctor_get(x_80, 1); +lean_inc(x_92); +if (lean_is_exclusive(x_80)) { + lean_ctor_release(x_80, 0); + lean_ctor_release(x_80, 1); + x_93 = x_80; +} else { + lean_dec_ref(x_80); + x_93 = lean_box(0); +} +if (lean_is_scalar(x_93)) { + x_94 = lean_alloc_ctor(1, 2, 0); +} else { + x_94 = x_93; +} +lean_ctor_set(x_94, 0, x_91); +lean_ctor_set(x_94, 1, x_92); +return x_94; +} +} else { -uint8_t x_31; -x_31 = lean_nat_dec_le(x_27, x_27); -if (x_31 == 0) +uint8_t x_95; lean_object* x_96; lean_object* x_97; +lean_dec(x_67); +x_95 = 0; +x_96 = lean_box(0); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_3); +x_97 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___lambda__1(x_74, x_3, x_95, x_68, x_96, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_97) == 0) { -lean_dec(x_27); -lean_dec(x_26); -x_7 = x_17; -x_9 = lean_box(0); -x_14 = x_20; -goto _start; +lean_object* x_98; +x_98 = lean_ctor_get(x_97, 0); +lean_inc(x_98); +if (lean_obj_tag(x_98) == 0) +{ +lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_7); +lean_dec(x_3); +x_99 = lean_ctor_get(x_97, 1); +lean_inc(x_99); +if (lean_is_exclusive(x_97)) { + lean_ctor_release(x_97, 0); + lean_ctor_release(x_97, 1); + x_100 = x_97; +} else { + lean_dec_ref(x_97); + x_100 = lean_box(0); +} +x_101 = lean_ctor_get(x_98, 0); +lean_inc(x_101); +lean_dec(x_98); +if (lean_is_scalar(x_100)) { + x_102 = lean_alloc_ctor(0, 2, 0); +} else { + x_102 = x_100; +} +lean_ctor_set(x_102, 0, x_101); +lean_ctor_set(x_102, 1, x_99); +return x_102; } else { -size_t x_33; size_t x_34; lean_object* x_35; -x_33 = 0; -x_34 = lean_usize_of_nat(x_27); -lean_dec(x_27); -x_35 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__6___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__7(x_3, x_26, x_33, x_34, x_8); -lean_dec(x_26); -x_7 = x_17; -x_8 = x_35; +lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; +x_103 = lean_ctor_get(x_97, 1); +lean_inc(x_103); +lean_dec(x_97); +x_104 = lean_ctor_get(x_98, 0); +lean_inc(x_104); +lean_dec(x_98); +x_105 = lean_ctor_get(x_5, 2); +x_106 = lean_nat_add(x_7, x_105); +lean_dec(x_7); +x_6 = x_104; +x_7 = x_106; +x_8 = lean_box(0); x_9 = lean_box(0); -x_14 = x_20; +x_14 = x_103; goto _start; } } -} else { -uint8_t x_37; -lean_dec(x_17); +lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); lean_dec(x_10); -lean_dec(x_8); -lean_dec(x_1); -x_37 = !lean_is_exclusive(x_18); -if (x_37 == 0) -{ -return x_18; +lean_dec(x_7); +lean_dec(x_3); +x_108 = lean_ctor_get(x_97, 0); +lean_inc(x_108); +x_109 = lean_ctor_get(x_97, 1); +lean_inc(x_109); +if (lean_is_exclusive(x_97)) { + lean_ctor_release(x_97, 0); + lean_ctor_release(x_97, 1); + x_110 = x_97; +} else { + lean_dec_ref(x_97); + x_110 = lean_box(0); } -else -{ -lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_38 = lean_ctor_get(x_18, 0); -x_39 = lean_ctor_get(x_18, 1); -lean_inc(x_39); -lean_inc(x_38); -lean_dec(x_18); -x_40 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_40, 0, x_38); -lean_ctor_set(x_40, 1, x_39); -return x_40; +if (lean_is_scalar(x_110)) { + x_111 = lean_alloc_ctor(1, 2, 0); +} else { + x_111 = x_110; } +lean_ctor_set(x_111, 0, x_108); +lean_ctor_set(x_111, 1, x_109); +return x_111; } } } } -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__8___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__9(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { -_start: -{ -if (lean_obj_tag(x_6) == 0) -{ -lean_object* x_14; -lean_dec(x_9); -lean_dec(x_1); -x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_7); -lean_ctor_set(x_14, 1, x_13); -return x_14; } -else -{ -lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = lean_ctor_get(x_6, 0); -lean_inc(x_15); -x_16 = lean_ctor_get(x_6, 1); -lean_inc(x_16); -lean_dec(x_6); -lean_inc(x_9); -x_17 = l_Lean_FVarId_getType(x_15, x_9, x_10, x_11, x_12, 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_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_17, 0); -lean_inc(x_18); -x_19 = lean_ctor_get(x_17, 1); -lean_inc(x_19); -lean_dec(x_17); -x_20 = lean_box(0); -lean_inc(x_1); -x_21 = lean_array_mk(x_1); -x_22 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; -x_23 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_23, 0, x_22); -lean_ctor_set(x_23, 1, x_20); -lean_ctor_set(x_23, 2, x_21); -x_24 = l_Lean_CollectFVars_main(x_18, x_23); -x_25 = lean_ctor_get(x_24, 2); -lean_inc(x_25); -lean_dec(x_24); -x_26 = lean_array_get_size(x_25); -x_27 = lean_unsigned_to_nat(0u); -x_28 = lean_nat_dec_lt(x_27, x_26); -if (x_28 == 0) -{ -lean_dec(x_26); -lean_dec(x_25); -x_6 = x_16; -x_8 = lean_box(0); -x_13 = x_19; -goto _start; } -else -{ -uint8_t x_30; -x_30 = lean_nat_dec_le(x_26, x_26); -if (x_30 == 0) -{ -lean_dec(x_26); -lean_dec(x_25); -x_6 = x_16; -x_8 = lean_box(0); -x_13 = x_19; -goto _start; } -else +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___lambda__1___closed__1() { +_start: { -size_t x_32; size_t x_33; lean_object* x_34; -x_32 = 0; -x_33 = lean_usize_of_nat(x_26); -lean_dec(x_26); -x_34 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__6___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__7(x_2, x_25, x_32, x_33, x_7); -lean_dec(x_25); -x_6 = x_16; -x_7 = x_34; -x_8 = lean_box(0); -x_13 = x_19; -goto _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_Grind_ppPattern___closed__4; +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_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_9 = lean_array_get_size(x_2); +x_10 = lean_unsigned_to_nat(0u); +x_11 = lean_unsigned_to_nat(1u); +x_12 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_12, 0, x_10); +lean_ctor_set(x_12, 1, x_9); +lean_ctor_set(x_12, 2, x_11); +x_13 = l_Lean_Meta_Grind_ppPattern___closed__4; +x_14 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___lambda__1___closed__1; +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +x_15 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1(x_1, x_2, x_13, x_12, x_12, x_14, x_10, lean_box(0), lean_box(0), x_4, x_5, x_6, x_7, x_8); +lean_dec(x_12); +if (lean_obj_tag(x_15) == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_16 = lean_ctor_get(x_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, 1); +lean_inc(x_18); +lean_dec(x_16); +x_19 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_18, x_4, x_5, x_6, x_7, x_17); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_19; } else { -uint8_t x_36; -lean_dec(x_16); -lean_dec(x_9); +uint8_t x_20; lean_dec(x_7); -lean_dec(x_1); -x_36 = !lean_is_exclusive(x_17); -if (x_36 == 0) +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_20 = !lean_is_exclusive(x_15); +if (x_20 == 0) { -return x_17; +return x_15; } else { -lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_37 = lean_ctor_get(x_17, 0); -x_38 = lean_ctor_get(x_17, 1); -lean_inc(x_38); -lean_inc(x_37); -lean_dec(x_17); -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_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_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__10(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt(lean_object* x_1, lean_object* x_2, lean_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_7; -x_7 = lean_usize_dec_eq(x_4, x_5); -if (x_7 == 0) -{ -lean_object* x_8; size_t x_9; size_t x_10; lean_object* x_11; -x_8 = lean_array_uget(x_3, x_4); -x_9 = 1; -x_10 = lean_usize_add(x_4, x_9); -x_11 = l_Lean_RBNode_findCore___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__2(x_2, x_8); -if (lean_obj_tag(x_11) == 0) +lean_object* x_9; +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +x_9 = lean_infer_type(x_1, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_9) == 0) { -lean_dec(x_8); -x_4 = x_10; -goto _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; +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_alloc_ctor(1, 1, 0); +lean_ctor_set(x_12, 0, x_2); +x_13 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___lambda__1___boxed), 8, 1); +lean_closure_set(x_13, 0, x_3); +x_14 = 0; +x_15 = l_Lean_Meta_forallBoundedTelescope___at_Lean_Meta_arrowDomainsN___spec__6___rarg(x_10, x_12, x_13, x_14, x_4, x_5, x_6, x_7, x_11); +return x_15; } else { -lean_object* x_13; lean_object* x_14; -lean_dec(x_11); -x_13 = lean_box(0); -x_14 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_6, x_8, x_13); -x_4 = x_10; -x_6 = x_14; -goto _start; -} +uint8_t x_16; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_16 = !lean_is_exclusive(x_9); +if (x_16 == 0) +{ +return x_9; } else { -return x_6; +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(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_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__10___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__11(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_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___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) { _start: { -uint8_t x_6; -x_6 = lean_usize_dec_eq(x_3, x_4); -if (x_6 == 0) -{ -lean_object* x_7; size_t x_8; size_t x_9; lean_object* x_10; -x_7 = lean_array_uget(x_2, x_3); -x_8 = 1; -x_9 = lean_usize_add(x_3, x_8); -x_10 = l_Lean_RBNode_findCore___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__2(x_1, x_7); -if (lean_obj_tag(x_10) == 0) -{ -lean_dec(x_7); -x_3 = x_9; -goto _start; +uint8_t x_11; lean_object* x_12; +x_11 = lean_unbox(x_3); +lean_dec(x_3); +x_12 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___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_5); +return x_12; } -else +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: { -lean_object* x_12; lean_object* x_13; -lean_dec(x_10); -x_12 = lean_box(0); -x_13 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_5, x_7, x_12); -x_3 = x_9; -x_5 = x_13; -goto _start; +lean_object* x_15; +x_15 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +return x_15; } } -else +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___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: { -return x_5; -} +lean_object* x_9; +x_9 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_9; } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__12(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__1(lean_object* x_1, lean_object* x_2) { _start: { -uint8_t x_7; -x_7 = lean_usize_dec_eq(x_4, x_5); -if (x_7 == 0) -{ -lean_object* x_8; size_t x_9; size_t x_10; lean_object* x_11; -x_8 = lean_array_uget(x_3, x_4); -x_9 = 1; -x_10 = lean_usize_add(x_4, x_9); -x_11 = l_Lean_RBNode_findCore___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__2(x_2, x_8); -if (lean_obj_tag(x_11) == 0) +if (lean_obj_tag(x_1) == 0) { -lean_dec(x_8); -x_4 = x_10; -goto _start; +lean_object* x_3; +x_3 = l_List_reverse___rarg(x_2); +return x_3; } else { -lean_object* x_13; lean_object* x_14; -lean_dec(x_11); -x_13 = lean_box(0); -x_14 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_6, x_8, x_13); -x_4 = x_10; -x_6 = x_14; -goto _start; +uint8_t x_4; +x_4 = !lean_is_exclusive(x_1); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_5 = lean_ctor_get(x_1, 0); +x_6 = lean_ctor_get(x_1, 1); +x_7 = l_Lean_Meta_Grind_ppPattern(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 { -return x_6; +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_Meta_Grind_ppPattern(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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__12___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__13(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_List_mapTR_loop___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__2(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) +if (lean_obj_tag(x_1) == 0) { -lean_object* x_7; size_t x_8; size_t x_9; lean_object* x_10; -x_7 = lean_array_uget(x_2, x_3); -x_8 = 1; -x_9 = lean_usize_add(x_3, x_8); -x_10 = l_Lean_RBNode_findCore___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__2(x_1, x_7); -if (lean_obj_tag(x_10) == 0) +lean_object* x_3; +x_3 = l_List_reverse___rarg(x_2); +return x_3; +} +else { -lean_dec(x_7); -x_3 = x_9; +uint8_t x_4; +x_4 = !lean_is_exclusive(x_1); +if (x_4 == 0) +{ +lean_object* x_5; +x_5 = lean_ctor_get(x_1, 1); +lean_ctor_set(x_1, 1, x_2); +{ +lean_object* _tmp_0 = x_5; +lean_object* _tmp_1 = x_1; +x_1 = _tmp_0; +x_2 = _tmp_1; +} goto _start; } else { -lean_object* x_12; lean_object* x_13; -lean_dec(x_10); -x_12 = lean_box(0); -x_13 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_5, x_7, x_12); -x_3 = x_9; -x_5 = x_13; +lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_7 = lean_ctor_get(x_1, 0); +x_8 = lean_ctor_get(x_1, 1); +lean_inc(x_8); +lean_inc(x_7); +lean_dec(x_1); +x_9 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_9, 0, x_7); +lean_ctor_set(x_9, 1, x_2); +x_1 = x_8; +x_2 = x_9; goto _start; } } -else -{ -return x_5; -} } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__14(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, size_t x_7, size_t x_8, lean_object* x_9, 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_Grind_Origin_pp___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { -uint8_t x_15; -x_15 = lean_usize_dec_lt(x_8, x_7); -if (x_15 == 0) -{ -lean_object* x_16; -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_4); -lean_dec(x_2); -x_16 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_16, 0, x_9); -lean_ctor_set(x_16, 1, x_14); -return x_16; -} -else -{ -lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_25; -x_17 = lean_array_uget(x_6, x_8); -x_25 = !lean_is_exclusive(x_9); -if (x_25 == 0) -{ -lean_object* x_26; uint8_t x_27; -x_26 = lean_ctor_get(x_9, 1); -x_27 = !lean_is_exclusive(x_26); -if (x_27 == 0) +switch (lean_obj_tag(x_1)) { +case 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_9, 0); -x_29 = lean_ctor_get(x_26, 0); -x_30 = lean_ctor_get(x_26, 1); -x_31 = l_Lean_Expr_fvarId_x21(x_17); -x_32 = l_Lean_RBNode_findCore___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__2(x_30, x_31); -if (lean_obj_tag(x_32) == 0) +lean_object* x_7; lean_object* x_8; +x_7 = lean_ctor_get(x_1, 0); +lean_inc(x_7); +lean_dec(x_1); +x_8 = l_Lean_mkConstWithLevelParams___at_Lean_Meta_registerCoercion___spec__1(x_7, x_2, x_3, x_4, x_5, x_6); +if (lean_obj_tag(x_8) == 0) { -lean_object* x_33; -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -x_33 = lean_infer_type(x_17, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_33) == 0) +uint8_t x_9; +x_9 = !lean_is_exclusive(x_8); +if (x_9 == 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 = l_Lean_RBNode_findCore___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__2(x_28, x_31); -if (lean_obj_tag(x_36) == 0) +lean_object* x_10; lean_object* x_11; +x_10 = lean_ctor_get(x_8, 0); +x_11 = l_Lean_MessageData_ofConst(x_10); +lean_ctor_set(x_8, 0, x_11); +return x_8; +} +else { -lean_object* x_37; +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_12 = lean_ctor_get(x_8, 0); +x_13 = lean_ctor_get(x_8, 1); lean_inc(x_13); lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_34); -x_37 = l_Lean_Meta_isProp(x_34, x_10, x_11, x_12, x_13, x_35); -if (lean_obj_tag(x_37) == 0) +lean_dec(x_8); +x_14 = l_Lean_MessageData_ofConst(x_12); +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_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) +uint8_t x_16; +x_16 = !lean_is_exclusive(x_8); +if (x_16 == 0) { -lean_object* x_40; lean_object* x_41; -x_40 = lean_ctor_get(x_37, 1); -lean_inc(x_40); -lean_dec(x_37); -lean_inc(x_10); -lean_inc(x_31); -x_41 = l_Lean_FVarId_getDecl(x_31, x_10, x_11, x_12, x_13, x_40); -if (lean_obj_tag(x_41) == 0) +return x_8; +} +else { -lean_object* x_42; lean_object* x_43; uint8_t 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 = l_Lean_LocalDecl_binderInfo(x_42); -lean_dec(x_42); -x_45 = lean_box(x_44); -if (lean_obj_tag(x_45) == 3) +lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_17 = lean_ctor_get(x_8, 0); +x_18 = lean_ctor_get(x_8, 1); +lean_inc(x_18); +lean_inc(x_17); +lean_dec(x_8); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_17); +lean_ctor_set(x_19, 1, x_18); +return x_19; +} +} +} +case 1: { -lean_object* x_46; -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_34); -lean_inc(x_28); -lean_inc(x_4); -x_46 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized(x_4, x_28, x_34, x_10, x_11, x_12, x_13, x_43); -if (lean_obj_tag(x_46) == 0) +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); +lean_dec(x_1); +x_21 = l_Lean_Expr_fvar___override(x_20); +x_22 = l_Lean_MessageData_ofExpr(x_21); +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_6); +return x_23; +} +case 2: { -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); -lean_dec(x_47); -if (x_48 == 0) +uint8_t x_24; +x_24 = !lean_is_exclusive(x_1); +if (x_24 == 0) { -lean_object* x_49; lean_object* x_50; -lean_dec(x_34); -lean_dec(x_31); -x_49 = lean_ctor_get(x_46, 1); -lean_inc(x_49); -lean_dec(x_46); -x_50 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_50, 0, x_9); -x_18 = x_50; -x_19 = x_49; -goto block_24; +lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_25 = lean_ctor_get(x_1, 1); +x_26 = lean_ctor_get(x_1, 0); +lean_dec(x_26); +x_27 = l_Lean_MessageData_ofSyntax(x_25); +lean_ctor_set_tag(x_1, 0); +lean_ctor_set(x_1, 1, x_6); +lean_ctor_set(x_1, 0, x_27); +return x_1; } else { -lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; uint8_t x_62; lean_object* x_63; -lean_dec(x_29); -x_51 = lean_ctor_get(x_46, 1); -lean_inc(x_51); -lean_dec(x_46); -x_52 = lean_box(0); +lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_28 = lean_ctor_get(x_1, 1); +lean_inc(x_28); +lean_dec(x_1); +x_29 = l_Lean_MessageData_ofSyntax(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_6); +return x_30; +} +} +default: +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_1, 0); lean_inc(x_31); -x_53 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_28, x_31, x_52); -x_54 = lean_box(0); -lean_inc(x_2); -x_55 = lean_array_mk(x_2); -x_56 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; -x_57 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_57, 0, x_56); -lean_ctor_set(x_57, 1, x_54); -lean_ctor_set(x_57, 2, x_55); -x_58 = l_Lean_CollectFVars_main(x_34, x_57); -x_59 = lean_ctor_get(x_58, 2); -lean_inc(x_59); -lean_dec(x_58); -x_60 = lean_array_get_size(x_59); -x_61 = lean_unsigned_to_nat(0u); -x_62 = lean_nat_dec_lt(x_61, x_60); -x_63 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_30, x_31, x_52); -if (x_62 == 0) +lean_dec(x_1); +x_32 = l_Lean_MessageData_ofName(x_31); +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_6); +return x_33; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: { -uint8_t x_64; lean_object* x_65; lean_object* x_66; -lean_dec(x_60); -lean_dec(x_59); -x_64 = 1; -x_65 = lean_box(x_64); -lean_ctor_set(x_26, 1, x_63); -lean_ctor_set(x_26, 0, x_65); -lean_ctor_set(x_9, 0, x_53); -x_66 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_66, 0, x_9); -x_18 = x_66; -x_19 = x_51; -goto block_24; +lean_object* x_13; lean_object* x_14; +x_13 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_13, 0, x_1); +lean_ctor_set(x_13, 1, x_2); +lean_ctor_set(x_13, 2, x_3); +lean_ctor_set(x_13, 3, x_4); +lean_ctor_set(x_13, 4, x_5); +lean_ctor_set(x_13, 5, x_6); +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 +} +static lean_object* _init_l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___closed__1() { +_start: { -uint8_t x_67; -x_67 = lean_nat_dec_le(x_60, x_60); -if (x_67 == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("invalid pattern(s) for `", 24, 24); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___closed__2() { +_start: { -uint8_t x_68; lean_object* x_69; lean_object* x_70; -lean_dec(x_60); -lean_dec(x_59); -x_68 = 1; -x_69 = lean_box(x_68); -lean_ctor_set(x_26, 1, x_63); -lean_ctor_set(x_26, 0, x_69); -lean_ctor_set(x_9, 0, x_53); -x_70 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_70, 0, x_9); -x_18 = x_70; -x_19 = x_51; -goto block_24; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("`", 1, 1); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___closed__3; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } -else +} +static lean_object* _init_l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___closed__5() { +_start: { -size_t x_71; size_t x_72; lean_object* x_73; uint8_t x_74; lean_object* x_75; lean_object* x_76; -x_71 = 0; -x_72 = lean_usize_of_nat(x_60); -lean_dec(x_60); -x_73 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__10___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__11(x_4, x_59, x_71, x_72, x_53); -lean_dec(x_59); -x_74 = 1; -x_75 = lean_box(x_74); -lean_ctor_set(x_26, 1, x_63); -lean_ctor_set(x_26, 0, x_75); -lean_ctor_set(x_9, 0, x_73); -x_76 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_76, 0, x_9); -x_18 = x_76; -x_19 = x_51; -goto block_24; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("\nthe following theorem parameters cannot be instantiated:", 57, 57); +return x_1; } } +static lean_object* _init_l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___closed__5; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } } -else +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: { -uint8_t x_77; -lean_dec(x_34); -lean_dec(x_31); -lean_free_object(x_26); -lean_dec(x_30); -lean_dec(x_29); -lean_free_object(x_9); -lean_dec(x_28); -lean_dec(x_13); +lean_object* x_14; +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_2); +lean_inc(x_1); +x_14 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage(x_1, x_2, x_3, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +if (lean_obj_tag(x_15) == 0) +{ +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_Meta_Grind_mkEMatchTheoremCore___lambda__1(x_4, x_1, x_2, x_5, x_6, x_7, x_17, x_9, x_10, x_11, x_12, x_16); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); -lean_dec(x_4); -lean_dec(x_2); -x_77 = !lean_is_exclusive(x_46); -if (x_77 == 0) -{ -return x_46; +lean_dec(x_9); +return x_18; } else { -lean_object* x_78; lean_object* x_79; lean_object* x_80; -x_78 = lean_ctor_get(x_46, 0); -x_79 = lean_ctor_get(x_46, 1); -lean_inc(x_79); -lean_inc(x_78); -lean_dec(x_46); -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_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_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_6); +lean_dec(x_4); +x_19 = lean_ctor_get(x_14, 1); +lean_inc(x_19); +lean_dec(x_14); +x_20 = lean_ctor_get(x_15, 0); +lean_inc(x_20); +lean_dec(x_15); +x_21 = lean_box(0); +x_22 = l_List_mapTR_loop___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__1(x_5, x_21); +x_23 = l_List_mapTR_loop___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__2(x_22, x_21); +x_24 = l_Lean_MessageData_ofList(x_23); +x_25 = l_Lean_Meta_Grind_ppPattern___closed__4; +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 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_25); +x_28 = l_Lean_Meta_Grind_Origin_pp___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__3(x_7, x_9, x_10, x_11, x_12, x_19); +if (lean_obj_tag(x_28) == 0) +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_29 = lean_ctor_get(x_28, 0); +lean_inc(x_29); +x_30 = lean_ctor_get(x_28, 1); +lean_inc(x_30); +lean_dec(x_28); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +x_31 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt(x_1, x_2, x_20, x_9, x_10, x_11, x_12, x_30); +if (lean_obj_tag(x_31) == 0) +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; uint8_t x_46; +x_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_Meta_Grind_mkEMatchTheoremCore___lambda__2___closed__2; +x_35 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_35, 0, x_34); +lean_ctor_set(x_35, 1, x_29); +x_36 = l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___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_Lean_indentD(x_27); +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_Meta_Grind_mkEMatchTheoremCore___lambda__2___closed__6; +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_indentD(x_32); +x_43 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_43, 0, x_41); +lean_ctor_set(x_43, 1, x_42); +x_44 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_44, 0, x_43); +lean_ctor_set(x_44, 1, x_25); +x_45 = l_Lean_throwError___at___private_Lean_Meta_InferType_0__Lean_Meta_inferProjType___spec__1(x_44, x_9, x_10, x_11, x_12, x_33); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +x_46 = !lean_is_exclusive(x_45); +if (x_46 == 0) +{ +return x_45; } else { -lean_object* x_81; +lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_47 = lean_ctor_get(x_45, 0); +x_48 = lean_ctor_get(x_45, 1); +lean_inc(x_48); +lean_inc(x_47); lean_dec(x_45); -lean_dec(x_34); -lean_dec(x_31); -x_81 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_81, 0, x_9); -x_18 = x_81; -x_19 = x_43; -goto block_24; +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_82; -lean_dec(x_34); -lean_dec(x_31); -lean_free_object(x_26); -lean_dec(x_30); +uint8_t x_50; lean_dec(x_29); -lean_free_object(x_9); -lean_dec(x_28); -lean_dec(x_13); +lean_dec(x_27); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); -lean_dec(x_4); -lean_dec(x_2); -x_82 = !lean_is_exclusive(x_41); -if (x_82 == 0) +lean_dec(x_9); +x_50 = !lean_is_exclusive(x_31); +if (x_50 == 0) { -return x_41; +return x_31; } else { -lean_object* x_83; lean_object* x_84; lean_object* x_85; -x_83 = lean_ctor_get(x_41, 0); -x_84 = lean_ctor_get(x_41, 1); -lean_inc(x_84); -lean_inc(x_83); -lean_dec(x_41); -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_51; lean_object* x_52; lean_object* x_53; +x_51 = lean_ctor_get(x_31, 0); +x_52 = lean_ctor_get(x_31, 1); +lean_inc(x_52); +lean_inc(x_51); +lean_dec(x_31); +x_53 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_53, 0, x_51); +lean_ctor_set(x_53, 1, x_52); +return x_53; } } } else { -lean_object* x_86; uint8_t x_87; -x_86 = lean_ctor_get(x_37, 1); -lean_inc(x_86); -lean_dec(x_37); -x_87 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars(x_4, x_28, x_34); -if (x_87 == 0) +uint8_t x_54; +lean_dec(x_27); +lean_dec(x_20); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_2); +lean_dec(x_1); +x_54 = !lean_is_exclusive(x_28); +if (x_54 == 0) { -lean_object* x_88; -lean_dec(x_31); -x_88 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_88, 0, x_9); -x_18 = x_88; -x_19 = x_86; -goto block_24; +return x_28; } else { -lean_object* x_89; lean_object* x_90; lean_object* x_91; uint8_t x_92; lean_object* x_93; lean_object* x_94; -lean_dec(x_29); -x_89 = lean_box(0); -lean_inc(x_31); -x_90 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_28, x_31, x_89); -x_91 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_30, x_31, x_89); -x_92 = 1; -x_93 = lean_box(x_92); -lean_ctor_set(x_26, 1, x_91); -lean_ctor_set(x_26, 0, x_93); -lean_ctor_set(x_9, 0, x_90); -x_94 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_94, 0, x_9); -x_18 = x_94; -x_19 = x_86; -goto block_24; +lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_55 = lean_ctor_get(x_28, 0); +x_56 = lean_ctor_get(x_28, 1); +lean_inc(x_56); +lean_inc(x_55); +lean_dec(x_28); +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_95; -lean_dec(x_34); -lean_dec(x_31); -lean_free_object(x_26); -lean_dec(x_30); -lean_dec(x_29); -lean_free_object(x_9); -lean_dec(x_28); -lean_dec(x_13); +uint8_t x_58; lean_dec(x_12); 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_2); -x_95 = !lean_is_exclusive(x_37); -if (x_95 == 0) +lean_dec(x_1); +x_58 = !lean_is_exclusive(x_14); +if (x_58 == 0) +{ +return x_14; +} +else +{ +lean_object* x_59; lean_object* x_60; lean_object* x_61; +x_59 = lean_ctor_get(x_14, 0); +x_60 = lean_ctor_get(x_14, 1); +lean_inc(x_60); +lean_inc(x_59); +lean_dec(x_14); +x_61 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_61, 0, x_59); +lean_ctor_set(x_61, 1, x_60); +return x_61; +} +} +} +} +static lean_object* _init_l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("ematch", 6, 6); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("pattern", 7, 7); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__3() { +_start: { -return x_37; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Meta_Grind_mkGroundPattern___closed__1; +x_2 = l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__1; +x_3 = l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__2; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; } -else +} +static lean_object* _init_l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__4() { +_start: { -lean_object* x_96; lean_object* x_97; lean_object* x_98; -x_96 = lean_ctor_get(x_37, 0); -x_97 = lean_ctor_get(x_37, 1); -lean_inc(x_97); -lean_inc(x_96); -lean_dec(x_37); -x_98 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_98, 0, x_96); -lean_ctor_set(x_98, 1, x_97); -return x_98; +lean_object* x_1; +x_1 = lean_mk_string_unchecked(": ", 2, 2); +return x_1; +} } +static lean_object* _init_l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__4; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } } -else +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchTheoremCore(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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_99; -lean_dec(x_29); -x_99 = !lean_is_exclusive(x_36); -if (x_99 == 0) +lean_object* x_11; +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +x_11 = l_Lean_Meta_Grind_NormalizePattern_main(x_5, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_11) == 0) { -lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; uint8_t x_109; lean_object* x_110; lean_object* x_111; -x_100 = lean_ctor_get(x_36, 0); -lean_dec(x_100); -x_101 = lean_box(0); -lean_inc(x_2); -x_102 = lean_array_mk(x_2); -x_103 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; -x_104 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_104, 0, x_103); -lean_ctor_set(x_104, 1, x_101); -lean_ctor_set(x_104, 2, x_102); -x_105 = l_Lean_CollectFVars_main(x_34, x_104); -x_106 = lean_ctor_get(x_105, 2); -lean_inc(x_106); -lean_dec(x_105); -x_107 = lean_array_get_size(x_106); -x_108 = lean_unsigned_to_nat(0u); -x_109 = lean_nat_dec_lt(x_108, x_107); -x_110 = lean_box(0); -x_111 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_30, x_31, x_110); -if (x_109 == 0) +lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +x_13 = lean_ctor_get(x_12, 1); +lean_inc(x_13); +x_14 = lean_ctor_get(x_11, 1); +lean_inc(x_14); +lean_dec(x_11); +x_15 = !lean_is_exclusive(x_12); +if (x_15 == 0) { -uint8_t x_112; lean_object* x_113; -lean_dec(x_107); -lean_dec(x_106); -x_112 = 1; -x_113 = lean_box(x_112); -lean_ctor_set(x_26, 1, x_111); -lean_ctor_set(x_26, 0, x_113); -lean_ctor_set(x_36, 0, x_9); -x_18 = x_36; -x_19 = x_35; -goto block_24; +lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_16 = lean_ctor_get(x_12, 0); +x_17 = lean_ctor_get(x_12, 1); +lean_dec(x_17); +x_18 = !lean_is_exclusive(x_13); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_19 = lean_ctor_get(x_13, 0); +x_20 = lean_ctor_get(x_13, 1); +x_21 = l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__3; +x_22 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_21, x_6, x_7, x_8, x_9, x_14); +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +x_24 = lean_unbox(x_23); +lean_dec(x_23); +if (x_24 == 0) +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; +lean_free_object(x_13); +lean_free_object(x_12); +x_25 = lean_ctor_get(x_22, 1); +lean_inc(x_25); +lean_dec(x_22); +x_26 = lean_box(0); +x_27 = l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2(x_4, x_3, x_20, x_2, x_16, x_19, x_1, x_26, x_6, x_7, x_8, x_9, x_25); +return x_27; } else { -uint8_t x_114; -x_114 = lean_nat_dec_le(x_107, x_107); -if (x_114 == 0) +uint8_t x_28; +x_28 = !lean_is_exclusive(x_22); +if (x_28 == 0) { -uint8_t x_115; lean_object* x_116; -lean_dec(x_107); -lean_dec(x_106); -x_115 = 1; -x_116 = lean_box(x_115); -lean_ctor_set(x_26, 1, x_111); -lean_ctor_set(x_26, 0, x_116); -lean_ctor_set(x_36, 0, x_9); -x_18 = x_36; -x_19 = x_35; -goto block_24; +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; 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_29 = lean_ctor_get(x_22, 1); +x_30 = lean_ctor_get(x_22, 0); +lean_dec(x_30); +lean_inc(x_4); +x_31 = l_Lean_MessageData_ofConst(x_4); +x_32 = l_Lean_Meta_Grind_ppPattern___closed__4; +lean_ctor_set_tag(x_22, 7); +lean_ctor_set(x_22, 1, x_31); +lean_ctor_set(x_22, 0, x_32); +x_33 = l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__5; +lean_ctor_set_tag(x_13, 7); +lean_ctor_set(x_13, 1, x_33); +lean_ctor_set(x_13, 0, x_22); +x_34 = lean_box(0); +lean_inc(x_16); +x_35 = l_List_mapTR_loop___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__1(x_16, x_34); +x_36 = l_List_mapTR_loop___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__2(x_35, x_34); +x_37 = l_Lean_MessageData_ofList(x_36); +lean_ctor_set_tag(x_12, 7); +lean_ctor_set(x_12, 1, x_37); +lean_ctor_set(x_12, 0, x_13); +x_38 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_38, 0, x_12); +lean_ctor_set(x_38, 1, x_32); +x_39 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_21, x_38, x_6, x_7, x_8, x_9, x_29); +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_Meta_Grind_mkEMatchTheoremCore___lambda__2(x_4, x_3, x_20, x_2, x_16, x_19, x_1, x_40, x_6, x_7, x_8, x_9, x_41); +lean_dec(x_40); +return x_42; } else { -size_t x_117; size_t x_118; lean_object* x_119; uint8_t x_120; lean_object* x_121; -x_117 = 0; -x_118 = lean_usize_of_nat(x_107); -lean_dec(x_107); -x_119 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__12___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__13(x_4, x_106, x_117, x_118, x_28); -lean_dec(x_106); -x_120 = 1; -x_121 = lean_box(x_120); -lean_ctor_set(x_26, 1, x_111); -lean_ctor_set(x_26, 0, x_121); -lean_ctor_set(x_9, 0, x_119); -lean_ctor_set(x_36, 0, x_9); -x_18 = x_36; -x_19 = x_35; -goto block_24; +lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; 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_22, 1); +lean_inc(x_43); +lean_dec(x_22); +lean_inc(x_4); +x_44 = l_Lean_MessageData_ofConst(x_4); +x_45 = l_Lean_Meta_Grind_ppPattern___closed__4; +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_Meta_Grind_mkEMatchTheoremCore___closed__5; +lean_ctor_set_tag(x_13, 7); +lean_ctor_set(x_13, 1, x_47); +lean_ctor_set(x_13, 0, x_46); +x_48 = lean_box(0); +lean_inc(x_16); +x_49 = l_List_mapTR_loop___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__1(x_16, x_48); +x_50 = l_List_mapTR_loop___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__2(x_49, x_48); +x_51 = l_Lean_MessageData_ofList(x_50); +lean_ctor_set_tag(x_12, 7); +lean_ctor_set(x_12, 1, x_51); +lean_ctor_set(x_12, 0, x_13); +x_52 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_52, 0, x_12); +lean_ctor_set(x_52, 1, x_45); +x_53 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_21, x_52, x_6, x_7, x_8, x_9, x_43); +x_54 = lean_ctor_get(x_53, 0); +lean_inc(x_54); +x_55 = lean_ctor_get(x_53, 1); +lean_inc(x_55); +lean_dec(x_53); +x_56 = l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2(x_4, x_3, x_20, x_2, x_16, x_19, x_1, x_54, x_6, x_7, x_8, x_9, x_55); +lean_dec(x_54); +return x_56; } } } else { -lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; uint8_t x_130; lean_object* x_131; lean_object* x_132; -lean_dec(x_36); -x_122 = lean_box(0); -lean_inc(x_2); -x_123 = lean_array_mk(x_2); -x_124 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; -x_125 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_125, 0, x_124); -lean_ctor_set(x_125, 1, x_122); -lean_ctor_set(x_125, 2, x_123); -x_126 = l_Lean_CollectFVars_main(x_34, x_125); -x_127 = lean_ctor_get(x_126, 2); -lean_inc(x_127); -lean_dec(x_126); -x_128 = lean_array_get_size(x_127); -x_129 = lean_unsigned_to_nat(0u); -x_130 = lean_nat_dec_lt(x_129, x_128); -x_131 = lean_box(0); -x_132 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_30, x_31, x_131); -if (x_130 == 0) +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_57 = lean_ctor_get(x_13, 0); +x_58 = lean_ctor_get(x_13, 1); +lean_inc(x_58); +lean_inc(x_57); +lean_dec(x_13); +x_59 = l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__3; +x_60 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_59, x_6, x_7, x_8, x_9, x_14); +x_61 = lean_ctor_get(x_60, 0); +lean_inc(x_61); +x_62 = lean_unbox(x_61); +lean_dec(x_61); +if (x_62 == 0) { -uint8_t x_133; lean_object* x_134; lean_object* x_135; -lean_dec(x_128); -lean_dec(x_127); -x_133 = 1; -x_134 = lean_box(x_133); -lean_ctor_set(x_26, 1, x_132); -lean_ctor_set(x_26, 0, x_134); -x_135 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_135, 0, x_9); -x_18 = x_135; -x_19 = x_35; -goto block_24; +lean_object* x_63; lean_object* x_64; lean_object* x_65; +lean_free_object(x_12); +x_63 = lean_ctor_get(x_60, 1); +lean_inc(x_63); +lean_dec(x_60); +x_64 = lean_box(0); +x_65 = l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2(x_4, x_3, x_58, x_2, x_16, x_57, x_1, x_64, x_6, x_7, x_8, x_9, x_63); +return x_65; } else { -uint8_t x_136; -x_136 = lean_nat_dec_le(x_128, x_128); -if (x_136 == 0) -{ -uint8_t x_137; lean_object* x_138; lean_object* x_139; -lean_dec(x_128); -lean_dec(x_127); -x_137 = 1; -x_138 = lean_box(x_137); -lean_ctor_set(x_26, 1, x_132); -lean_ctor_set(x_26, 0, x_138); -x_139 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_139, 0, x_9); -x_18 = x_139; -x_19 = x_35; -goto block_24; +lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; +x_66 = lean_ctor_get(x_60, 1); +lean_inc(x_66); +if (lean_is_exclusive(x_60)) { + lean_ctor_release(x_60, 0); + lean_ctor_release(x_60, 1); + x_67 = x_60; +} else { + lean_dec_ref(x_60); + x_67 = lean_box(0); +} +lean_inc(x_4); +x_68 = l_Lean_MessageData_ofConst(x_4); +x_69 = l_Lean_Meta_Grind_ppPattern___closed__4; +if (lean_is_scalar(x_67)) { + x_70 = lean_alloc_ctor(7, 2, 0); +} else { + x_70 = x_67; + lean_ctor_set_tag(x_70, 7); +} +lean_ctor_set(x_70, 0, x_69); +lean_ctor_set(x_70, 1, x_68); +x_71 = l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__5; +x_72 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_72, 0, x_70); +lean_ctor_set(x_72, 1, x_71); +x_73 = lean_box(0); +lean_inc(x_16); +x_74 = l_List_mapTR_loop___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__1(x_16, x_73); +x_75 = l_List_mapTR_loop___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__2(x_74, x_73); +x_76 = l_Lean_MessageData_ofList(x_75); +lean_ctor_set_tag(x_12, 7); +lean_ctor_set(x_12, 1, x_76); +lean_ctor_set(x_12, 0, x_72); +x_77 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_77, 0, x_12); +lean_ctor_set(x_77, 1, x_69); +x_78 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_59, x_77, x_6, x_7, x_8, x_9, x_66); +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 = l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2(x_4, x_3, x_58, x_2, x_16, x_57, x_1, x_79, x_6, x_7, x_8, x_9, x_80); +lean_dec(x_79); +return x_81; +} +} } else { -size_t x_140; size_t x_141; lean_object* x_142; uint8_t x_143; lean_object* x_144; lean_object* x_145; -x_140 = 0; -x_141 = lean_usize_of_nat(x_128); -lean_dec(x_128); -x_142 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__12___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__13(x_4, x_127, x_140, x_141, x_28); -lean_dec(x_127); -x_143 = 1; -x_144 = lean_box(x_143); -lean_ctor_set(x_26, 1, x_132); -lean_ctor_set(x_26, 0, x_144); -lean_ctor_set(x_9, 0, x_142); -x_145 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_145, 0, x_9); -x_18 = x_145; -x_19 = x_35; -goto block_24; +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_82 = lean_ctor_get(x_12, 0); +lean_inc(x_82); +lean_dec(x_12); +x_83 = lean_ctor_get(x_13, 0); +lean_inc(x_83); +x_84 = lean_ctor_get(x_13, 1); +lean_inc(x_84); +if (lean_is_exclusive(x_13)) { + lean_ctor_release(x_13, 0); + lean_ctor_release(x_13, 1); + x_85 = x_13; +} else { + lean_dec_ref(x_13); + x_85 = lean_box(0); +} +x_86 = l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__3; +x_87 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_86, x_6, x_7, x_8, x_9, x_14); +x_88 = lean_ctor_get(x_87, 0); +lean_inc(x_88); +x_89 = lean_unbox(x_88); +lean_dec(x_88); +if (x_89 == 0) +{ +lean_object* x_90; lean_object* x_91; lean_object* x_92; +lean_dec(x_85); +x_90 = lean_ctor_get(x_87, 1); +lean_inc(x_90); +lean_dec(x_87); +x_91 = lean_box(0); +x_92 = l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2(x_4, x_3, x_84, x_2, x_82, x_83, x_1, x_91, x_6, x_7, x_8, x_9, x_90); +return x_92; +} +else +{ +lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; 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_93 = lean_ctor_get(x_87, 1); +lean_inc(x_93); +if (lean_is_exclusive(x_87)) { + lean_ctor_release(x_87, 0); + lean_ctor_release(x_87, 1); + x_94 = x_87; +} else { + lean_dec_ref(x_87); + x_94 = lean_box(0); +} +lean_inc(x_4); +x_95 = l_Lean_MessageData_ofConst(x_4); +x_96 = l_Lean_Meta_Grind_ppPattern___closed__4; +if (lean_is_scalar(x_94)) { + x_97 = lean_alloc_ctor(7, 2, 0); +} else { + x_97 = x_94; + lean_ctor_set_tag(x_97, 7); } +lean_ctor_set(x_97, 0, x_96); +lean_ctor_set(x_97, 1, x_95); +x_98 = l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__5; +if (lean_is_scalar(x_85)) { + x_99 = lean_alloc_ctor(7, 2, 0); +} else { + x_99 = x_85; + lean_ctor_set_tag(x_99, 7); } +lean_ctor_set(x_99, 0, x_97); +lean_ctor_set(x_99, 1, x_98); +x_100 = lean_box(0); +lean_inc(x_82); +x_101 = l_List_mapTR_loop___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__1(x_82, x_100); +x_102 = l_List_mapTR_loop___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__2(x_101, x_100); +x_103 = l_Lean_MessageData_ofList(x_102); +x_104 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_104, 0, x_99); +lean_ctor_set(x_104, 1, x_103); +x_105 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_105, 0, x_104); +lean_ctor_set(x_105, 1, x_96); +x_106 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_86, x_105, x_6, x_7, x_8, x_9, x_93); +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 = l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2(x_4, x_3, x_84, x_2, x_82, x_83, x_1, x_107, x_6, x_7, x_8, x_9, x_108); +lean_dec(x_107); +return x_109; } } } else { -uint8_t x_146; -lean_dec(x_31); -lean_free_object(x_26); -lean_dec(x_30); -lean_dec(x_29); -lean_free_object(x_9); -lean_dec(x_28); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); +uint8_t x_110; +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_146 = !lean_is_exclusive(x_33); -if (x_146 == 0) +lean_dec(x_1); +x_110 = !lean_is_exclusive(x_11); +if (x_110 == 0) { -return x_33; +return x_11; } else { -lean_object* x_147; lean_object* x_148; lean_object* x_149; -x_147 = lean_ctor_get(x_33, 0); -x_148 = lean_ctor_get(x_33, 1); -lean_inc(x_148); -lean_inc(x_147); -lean_dec(x_33); -x_149 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_149, 0, x_147); -lean_ctor_set(x_149, 1, x_148); -return x_149; +lean_object* x_111; lean_object* x_112; lean_object* x_113; +x_111 = lean_ctor_get(x_11, 0); +x_112 = lean_ctor_get(x_11, 1); +lean_inc(x_112); +lean_inc(x_111); +lean_dec(x_11); +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_EXPORT lean_object* l_Lean_Meta_Grind_Origin_pp___at_Lean_Meta_Grind_mkEMatchTheoremCore___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_150; -lean_dec(x_31); -lean_dec(x_17); -x_150 = !lean_is_exclusive(x_32); -if (x_150 == 0) +lean_object* x_7; +x_7 = l_Lean_Meta_Grind_Origin_pp___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__3(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchTheoremCore___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_151; -x_151 = lean_ctor_get(x_32, 0); -lean_dec(x_151); -lean_ctor_set(x_32, 0, x_9); -x_18 = x_32; -x_19 = x_14; -goto block_24; +lean_object* x_13; +x_13 = l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +return x_13; } -else +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: { -lean_object* x_152; -lean_dec(x_32); -x_152 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_152, 0, x_9); -x_18 = x_152; -x_19 = x_14; -goto block_24; +lean_object* x_14; +x_14 = l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_8); +return x_14; } } +LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getProofFor___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_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; -x_153 = lean_ctor_get(x_9, 0); -x_154 = lean_ctor_get(x_26, 0); -x_155 = lean_ctor_get(x_26, 1); -lean_inc(x_155); -lean_inc(x_154); -lean_dec(x_26); -x_156 = l_Lean_Expr_fvarId_x21(x_17); -x_157 = l_Lean_RBNode_findCore___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__2(x_155, x_156); -if (lean_obj_tag(x_157) == 0) -{ -lean_object* x_158; -lean_inc(x_13); -lean_inc(x_12); +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); -x_158 = lean_infer_type(x_17, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_158) == 0) +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_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getProofFor___closed__1() { +_start: { -lean_object* x_159; lean_object* x_160; lean_object* x_161; -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_RBNode_findCore___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__2(x_153, x_156); -if (lean_obj_tag(x_161) == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("` is not a theorem", 18, 18); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getProofFor___closed__2() { +_start: { -lean_object* x_162; -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_159); -x_162 = l_Lean_Meta_isProp(x_159, x_10, x_11, x_12, x_13, x_160); -if (lean_obj_tag(x_162) == 0) +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getProofFor___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getProofFor(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: { -lean_object* x_163; uint8_t x_164; -x_163 = lean_ctor_get(x_162, 0); -lean_inc(x_163); -x_164 = lean_unbox(x_163); -lean_dec(x_163); -if (x_164 == 0) +lean_object* x_5; +lean_inc(x_1); +x_5 = l_Lean_getConstInfo___at___private_Lean_Compiler_InlineAttrs_0__Lean_Compiler_isValidMacroInline___spec__1(x_1, x_2, x_3, x_4); +if (lean_obj_tag(x_5) == 0) { -lean_object* x_165; lean_object* x_166; -x_165 = lean_ctor_get(x_162, 1); -lean_inc(x_165); -lean_dec(x_162); -lean_inc(x_10); -lean_inc(x_156); -x_166 = l_Lean_FVarId_getDecl(x_156, x_10, x_11, x_12, x_13, x_165); -if (lean_obj_tag(x_166) == 0) +lean_object* x_6; +x_6 = lean_ctor_get(x_5, 0); +lean_inc(x_6); +if (lean_obj_tag(x_6) == 2) { -lean_object* x_167; lean_object* x_168; uint8_t x_169; lean_object* x_170; -x_167 = lean_ctor_get(x_166, 0); -lean_inc(x_167); -x_168 = lean_ctor_get(x_166, 1); -lean_inc(x_168); -lean_dec(x_166); -x_169 = l_Lean_LocalDecl_binderInfo(x_167); -lean_dec(x_167); -x_170 = lean_box(x_169); -if (lean_obj_tag(x_170) == 3) +uint8_t x_7; +x_7 = !lean_is_exclusive(x_5); +if (x_7 == 0) { -lean_object* x_171; -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_8 = lean_ctor_get(x_5, 0); +lean_dec(x_8); +x_9 = lean_ctor_get(x_6, 0); +lean_inc(x_9); +lean_dec(x_6); +x_10 = lean_ctor_get(x_9, 0); lean_inc(x_10); -lean_inc(x_159); -lean_inc(x_153); -lean_inc(x_4); -x_171 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized(x_4, x_153, x_159, x_10, x_11, x_12, x_13, x_168); -if (lean_obj_tag(x_171) == 0) +lean_dec(x_9); +x_11 = lean_ctor_get(x_10, 1); +lean_inc(x_11); +lean_dec(x_10); +x_12 = lean_box(0); +x_13 = l_List_mapTR_loop___at_Lean_mkConstWithLevelParams___spec__1(x_11, x_12); +x_14 = l_Lean_Expr_const___override(x_1, x_13); +lean_ctor_set(x_5, 0, x_14); +return x_5; +} +else { -lean_object* x_172; uint8_t x_173; -x_172 = lean_ctor_get(x_171, 0); -lean_inc(x_172); -x_173 = lean_unbox(x_172); -lean_dec(x_172); -if (x_173 == 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; +x_15 = lean_ctor_get(x_5, 1); +lean_inc(x_15); +lean_dec(x_5); +x_16 = lean_ctor_get(x_6, 0); +lean_inc(x_16); +lean_dec(x_6); +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +lean_dec(x_16); +x_18 = lean_ctor_get(x_17, 1); +lean_inc(x_18); +lean_dec(x_17); +x_19 = lean_box(0); +x_20 = l_List_mapTR_loop___at_Lean_mkConstWithLevelParams___spec__1(x_18, x_19); +x_21 = l_Lean_Expr_const___override(x_1, x_20); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_15); +return x_22; +} +} +else { -lean_object* x_174; lean_object* x_175; lean_object* x_176; -lean_dec(x_159); -lean_dec(x_156); -x_174 = lean_ctor_get(x_171, 1); -lean_inc(x_174); -lean_dec(x_171); -x_175 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_175, 0, x_154); -lean_ctor_set(x_175, 1, x_155); -lean_ctor_set(x_9, 1, x_175); -x_176 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_176, 0, x_9); -x_18 = x_176; -x_19 = x_174; -goto block_24; +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +lean_dec(x_6); +x_23 = lean_ctor_get(x_5, 1); +lean_inc(x_23); +lean_dec(x_5); +x_24 = l_Lean_MessageData_ofName(x_1); +x_25 = l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___closed__4; +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___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getProofFor___closed__2; +x_28 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_28, 0, x_26); +lean_ctor_set(x_28, 1, x_27); +x_29 = l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getProofFor___spec__1(x_28, x_2, x_3, x_23); +return x_29; +} } else { -lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; uint8_t x_188; lean_object* x_189; -lean_dec(x_154); -x_177 = lean_ctor_get(x_171, 1); -lean_inc(x_177); -lean_dec(x_171); -x_178 = lean_box(0); -lean_inc(x_156); -x_179 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_153, x_156, x_178); -x_180 = lean_box(0); -lean_inc(x_2); -x_181 = lean_array_mk(x_2); -x_182 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; -x_183 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_183, 0, x_182); -lean_ctor_set(x_183, 1, x_180); -lean_ctor_set(x_183, 2, x_181); -x_184 = l_Lean_CollectFVars_main(x_159, x_183); -x_185 = lean_ctor_get(x_184, 2); -lean_inc(x_185); -lean_dec(x_184); -x_186 = lean_array_get_size(x_185); -x_187 = lean_unsigned_to_nat(0u); -x_188 = lean_nat_dec_lt(x_187, x_186); -x_189 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_155, x_156, x_178); -if (x_188 == 0) +uint8_t x_30; +lean_dec(x_1); +x_30 = !lean_is_exclusive(x_5); +if (x_30 == 0) { -uint8_t x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; -lean_dec(x_186); -lean_dec(x_185); -x_190 = 1; -x_191 = lean_box(x_190); -x_192 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_192, 0, x_191); -lean_ctor_set(x_192, 1, x_189); -lean_ctor_set(x_9, 1, x_192); -lean_ctor_set(x_9, 0, x_179); -x_193 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_193, 0, x_9); -x_18 = x_193; -x_19 = x_177; -goto block_24; +return x_5; } else { -uint8_t x_194; -x_194 = lean_nat_dec_le(x_186, x_186); -if (x_194 == 0) -{ -uint8_t x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; -lean_dec(x_186); -lean_dec(x_185); -x_195 = 1; -x_196 = lean_box(x_195); -x_197 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_197, 0, x_196); -lean_ctor_set(x_197, 1, x_189); -lean_ctor_set(x_9, 1, x_197); -lean_ctor_set(x_9, 0, x_179); -x_198 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_198, 0, x_9); -x_18 = x_198; -x_19 = x_177; -goto block_24; +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_5, 0); +x_32 = lean_ctor_get(x_5, 1); +lean_inc(x_32); +lean_inc(x_31); +lean_dec(x_5); +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_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getProofFor___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: { -size_t x_199; size_t x_200; lean_object* x_201; uint8_t x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; -x_199 = 0; -x_200 = lean_usize_of_nat(x_186); -lean_dec(x_186); -x_201 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__10___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__11(x_4, x_185, x_199, x_200, x_179); -lean_dec(x_185); -x_202 = 1; -x_203 = lean_box(x_202); -x_204 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_204, 0, x_203); -lean_ctor_set(x_204, 1, x_189); -lean_ctor_set(x_9, 1, x_204); -lean_ctor_set(x_9, 0, x_201); -x_205 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_205, 0, x_9); -x_18 = x_205; -x_19 = x_177; -goto block_24; +lean_object* x_5; +x_5 = l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getProofFor___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___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getProofFor___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getProofFor(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_Meta_Grind_mkEMatchTheorem(lean_object* x_1, lean_object* x_2, lean_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 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getProofFor(x_1, x_6, x_7, x_8); +if (lean_obj_tag(x_9) == 0) +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +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_alloc_ctor(0, 1, 0); +lean_ctor_set(x_12, 0, x_1); +x_13 = l_Lean_Meta_Grind_NormalizePattern_main___closed__1; +x_14 = l_Lean_Meta_Grind_mkEMatchTheoremCore(x_12, x_13, x_2, x_10, x_3, x_4, x_5, x_6, x_7, x_11); +return x_14; } else { -lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; -lean_dec(x_159); -lean_dec(x_156); -lean_dec(x_155); -lean_dec(x_154); -lean_free_object(x_9); -lean_dec(x_153); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); +uint8_t x_15; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); lean_dec(x_4); +lean_dec(x_3); lean_dec(x_2); -x_206 = lean_ctor_get(x_171, 0); -lean_inc(x_206); -x_207 = lean_ctor_get(x_171, 1); -lean_inc(x_207); -if (lean_is_exclusive(x_171)) { - lean_ctor_release(x_171, 0); - lean_ctor_release(x_171, 1); - x_208 = x_171; -} else { - lean_dec_ref(x_171); - x_208 = lean_box(0); +lean_dec(x_1); +x_15 = !lean_is_exclusive(x_9); +if (x_15 == 0) +{ +return x_9; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_9, 0); +x_17 = lean_ctor_get(x_9, 1); +lean_inc(x_17); +lean_inc(x_16); +lean_dec(x_9); +x_18 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_18, 0, x_16); +lean_ctor_set(x_18, 1, x_17); +return x_18; } -if (lean_is_scalar(x_208)) { - x_209 = lean_alloc_ctor(1, 2, 0); -} else { - x_209 = x_208; } -lean_ctor_set(x_209, 0, x_206); -lean_ctor_set(x_209, 1, x_207); -return x_209; } } +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_Grind_mkEMatchEqTheoremCore___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; +x_7 = lean_ctor_get(x_4, 5); +x_8 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_1, x_2, x_3, x_4, x_5, x_6); +x_9 = !lean_is_exclusive(x_8); +if (x_9 == 0) +{ +lean_object* x_10; lean_object* x_11; +x_10 = lean_ctor_get(x_8, 0); +lean_inc(x_7); +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_7); +lean_ctor_set(x_11, 1, x_10); +lean_ctor_set_tag(x_8, 1); +lean_ctor_set(x_8, 0, x_11); +return x_8; +} else { -lean_object* x_210; lean_object* x_211; -lean_dec(x_170); -lean_dec(x_159); -lean_dec(x_156); -x_210 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_210, 0, x_154); -lean_ctor_set(x_210, 1, x_155); -lean_ctor_set(x_9, 1, x_210); -x_211 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_211, 0, x_9); -x_18 = x_211; -x_19 = x_168; -goto block_24; +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_12 = lean_ctor_get(x_8, 0); +x_13 = lean_ctor_get(x_8, 1); +lean_inc(x_13); +lean_inc(x_12); +lean_dec(x_8); +lean_inc(x_7); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_7); +lean_ctor_set(x_14, 1, x_12); +x_15 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_15, 0, x_14); +lean_ctor_set(x_15, 1, x_13); +return x_15; } } -else +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchEqTheoremCore___lambda__1(uint8_t x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: { -lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; -lean_dec(x_159); -lean_dec(x_156); -lean_dec(x_155); -lean_dec(x_154); -lean_free_object(x_9); -lean_dec(x_153); -lean_dec(x_13); +if (x_1 == 0) +{ +uint8_t x_10; +x_10 = !lean_is_exclusive(x_4); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_11 = lean_ctor_get(x_4, 1); +x_12 = lean_ctor_get(x_4, 0); lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_4); -lean_dec(x_2); -x_212 = lean_ctor_get(x_166, 0); -lean_inc(x_212); -x_213 = lean_ctor_get(x_166, 1); -lean_inc(x_213); -if (lean_is_exclusive(x_166)) { - lean_ctor_release(x_166, 0); - lean_ctor_release(x_166, 1); - x_214 = x_166; -} else { - lean_dec_ref(x_166); - x_214 = lean_box(0); -} -if (lean_is_scalar(x_214)) { - x_215 = lean_alloc_ctor(1, 2, 0); -} else { - x_215 = x_214; +x_13 = l_Lean_Meta_Grind_preprocessPattern(x_11, x_2, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_13) == 0) +{ +uint8_t x_14; +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_15 = lean_ctor_get(x_13, 0); +x_16 = lean_array_get_size(x_3); +x_17 = lean_expr_abstract(x_15, x_3); +lean_dec(x_15); +x_18 = lean_box(0); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_17); +lean_ctor_set(x_19, 1, x_18); +lean_ctor_set(x_4, 1, x_19); +lean_ctor_set(x_4, 0, x_16); +lean_ctor_set(x_13, 0, x_4); +return x_13; } -lean_ctor_set(x_215, 0, x_212); -lean_ctor_set(x_215, 1, x_213); -return x_215; +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_20 = lean_ctor_get(x_13, 0); +x_21 = lean_ctor_get(x_13, 1); +lean_inc(x_21); +lean_inc(x_20); +lean_dec(x_13); +x_22 = lean_array_get_size(x_3); +x_23 = lean_expr_abstract(x_20, x_3); +lean_dec(x_20); +x_24 = lean_box(0); +x_25 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_24); +lean_ctor_set(x_4, 1, x_25); +lean_ctor_set(x_4, 0, x_22); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_4); +lean_ctor_set(x_26, 1, x_21); +return x_26; } } else { -lean_object* x_216; uint8_t x_217; -x_216 = lean_ctor_get(x_162, 1); -lean_inc(x_216); -lean_dec(x_162); -x_217 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars(x_4, x_153, x_159); -if (x_217 == 0) +uint8_t x_27; +lean_free_object(x_4); +x_27 = !lean_is_exclusive(x_13); +if (x_27 == 0) { -lean_object* x_218; lean_object* x_219; -lean_dec(x_156); -x_218 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_218, 0, x_154); -lean_ctor_set(x_218, 1, x_155); -lean_ctor_set(x_9, 1, x_218); -x_219 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_219, 0, x_9); -x_18 = x_219; -x_19 = x_216; -goto block_24; +return x_13; } else { -lean_object* x_220; lean_object* x_221; lean_object* x_222; uint8_t x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; -lean_dec(x_154); -x_220 = lean_box(0); -lean_inc(x_156); -x_221 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_153, x_156, x_220); -x_222 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_155, x_156, x_220); -x_223 = 1; -x_224 = lean_box(x_223); -x_225 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_225, 0, x_224); -lean_ctor_set(x_225, 1, x_222); -lean_ctor_set(x_9, 1, x_225); -lean_ctor_set(x_9, 0, x_221); -x_226 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_226, 0, x_9); -x_18 = x_226; -x_19 = x_216; -goto block_24; +lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_28 = lean_ctor_get(x_13, 0); +x_29 = lean_ctor_get(x_13, 1); +lean_inc(x_29); +lean_inc(x_28); +lean_dec(x_13); +x_30 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_30, 0, x_28); +lean_ctor_set(x_30, 1, x_29); +return x_30; } } } else { -lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; -lean_dec(x_159); -lean_dec(x_156); -lean_dec(x_155); -lean_dec(x_154); -lean_free_object(x_9); -lean_dec(x_153); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); +lean_object* x_31; lean_object* x_32; +x_31 = lean_ctor_get(x_4, 1); +lean_inc(x_31); lean_dec(x_4); -lean_dec(x_2); -x_227 = lean_ctor_get(x_162, 0); -lean_inc(x_227); -x_228 = lean_ctor_get(x_162, 1); -lean_inc(x_228); -if (lean_is_exclusive(x_162)) { - lean_ctor_release(x_162, 0); - lean_ctor_release(x_162, 1); - x_229 = x_162; +x_32 = l_Lean_Meta_Grind_preprocessPattern(x_31, x_2, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_32) == 0) +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_33 = lean_ctor_get(x_32, 0); +lean_inc(x_33); +x_34 = lean_ctor_get(x_32, 1); +lean_inc(x_34); +if (lean_is_exclusive(x_32)) { + lean_ctor_release(x_32, 0); + lean_ctor_release(x_32, 1); + x_35 = x_32; } else { - lean_dec_ref(x_162); - x_229 = lean_box(0); + lean_dec_ref(x_32); + x_35 = lean_box(0); } -if (lean_is_scalar(x_229)) { - x_230 = lean_alloc_ctor(1, 2, 0); +x_36 = lean_array_get_size(x_3); +x_37 = lean_expr_abstract(x_33, x_3); +lean_dec(x_33); +x_38 = lean_box(0); +x_39 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_39, 0, x_37); +lean_ctor_set(x_39, 1, x_38); +x_40 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_40, 0, x_36); +lean_ctor_set(x_40, 1, x_39); +if (lean_is_scalar(x_35)) { + x_41 = lean_alloc_ctor(0, 2, 0); } else { - x_230 = x_229; -} -lean_ctor_set(x_230, 0, x_227); -lean_ctor_set(x_230, 1, x_228); -return x_230; + x_41 = x_35; } +lean_ctor_set(x_41, 0, x_40); +lean_ctor_set(x_41, 1, x_34); +return x_41; } else { -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; uint8_t x_240; lean_object* x_241; lean_object* x_242; -lean_dec(x_154); -if (lean_is_exclusive(x_161)) { - lean_ctor_release(x_161, 0); - x_231 = x_161; +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_42 = lean_ctor_get(x_32, 0); +lean_inc(x_42); +x_43 = lean_ctor_get(x_32, 1); +lean_inc(x_43); +if (lean_is_exclusive(x_32)) { + lean_ctor_release(x_32, 0); + lean_ctor_release(x_32, 1); + x_44 = x_32; } else { - lean_dec_ref(x_161); - x_231 = lean_box(0); + lean_dec_ref(x_32); + x_44 = lean_box(0); } -x_232 = lean_box(0); -lean_inc(x_2); -x_233 = lean_array_mk(x_2); -x_234 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; -x_235 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_235, 0, x_234); -lean_ctor_set(x_235, 1, x_232); -lean_ctor_set(x_235, 2, x_233); -x_236 = l_Lean_CollectFVars_main(x_159, x_235); -x_237 = lean_ctor_get(x_236, 2); -lean_inc(x_237); -lean_dec(x_236); -x_238 = lean_array_get_size(x_237); -x_239 = lean_unsigned_to_nat(0u); -x_240 = lean_nat_dec_lt(x_239, x_238); -x_241 = lean_box(0); -x_242 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_155, x_156, x_241); -if (x_240 == 0) -{ -uint8_t x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; -lean_dec(x_238); -lean_dec(x_237); -x_243 = 1; -x_244 = lean_box(x_243); -x_245 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_245, 0, x_244); -lean_ctor_set(x_245, 1, x_242); -lean_ctor_set(x_9, 1, x_245); -if (lean_is_scalar(x_231)) { - x_246 = lean_alloc_ctor(1, 1, 0); +if (lean_is_scalar(x_44)) { + x_45 = lean_alloc_ctor(1, 2, 0); } else { - x_246 = x_231; + x_45 = x_44; +} +lean_ctor_set(x_45, 0, x_42); +lean_ctor_set(x_45, 1, x_43); +return x_45; +} } -lean_ctor_set(x_246, 0, x_9); -x_18 = x_246; -x_19 = x_160; -goto block_24; } else { -uint8_t x_247; -x_247 = lean_nat_dec_le(x_238, x_238); -if (x_247 == 0) +uint8_t x_46; +x_46 = !lean_is_exclusive(x_4); +if (x_46 == 0) { -uint8_t x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; -lean_dec(x_238); -lean_dec(x_237); -x_248 = 1; -x_249 = lean_box(x_248); -x_250 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_250, 0, x_249); -lean_ctor_set(x_250, 1, x_242); -lean_ctor_set(x_9, 1, x_250); -if (lean_is_scalar(x_231)) { - x_251 = lean_alloc_ctor(1, 1, 0); -} else { - x_251 = x_231; -} -lean_ctor_set(x_251, 0, x_9); -x_18 = x_251; -x_19 = x_160; -goto block_24; +lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_47 = lean_ctor_get(x_4, 0); +x_48 = lean_ctor_get(x_4, 1); +lean_dec(x_48); +x_49 = l_Lean_Meta_Grind_preprocessPattern(x_47, x_2, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_49) == 0) +{ +uint8_t x_50; +x_50 = !lean_is_exclusive(x_49); +if (x_50 == 0) +{ +lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_51 = lean_ctor_get(x_49, 0); +x_52 = lean_array_get_size(x_3); +x_53 = lean_expr_abstract(x_51, x_3); +lean_dec(x_51); +x_54 = lean_box(0); +x_55 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_55, 0, x_53); +lean_ctor_set(x_55, 1, x_54); +lean_ctor_set(x_4, 1, x_55); +lean_ctor_set(x_4, 0, x_52); +lean_ctor_set(x_49, 0, x_4); +return x_49; } else { -size_t x_252; size_t x_253; lean_object* x_254; uint8_t x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; -x_252 = 0; -x_253 = lean_usize_of_nat(x_238); -lean_dec(x_238); -x_254 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__12___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__13(x_4, x_237, x_252, x_253, x_153); -lean_dec(x_237); -x_255 = 1; -x_256 = lean_box(x_255); -x_257 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_257, 0, x_256); -lean_ctor_set(x_257, 1, x_242); -lean_ctor_set(x_9, 1, x_257); -lean_ctor_set(x_9, 0, x_254); -if (lean_is_scalar(x_231)) { - x_258 = lean_alloc_ctor(1, 1, 0); -} else { - x_258 = x_231; +lean_object* x_56; 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_49, 0); +x_57 = lean_ctor_get(x_49, 1); +lean_inc(x_57); +lean_inc(x_56); +lean_dec(x_49); +x_58 = lean_array_get_size(x_3); +x_59 = lean_expr_abstract(x_56, x_3); +lean_dec(x_56); +x_60 = lean_box(0); +x_61 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_61, 0, x_59); +lean_ctor_set(x_61, 1, x_60); +lean_ctor_set(x_4, 1, x_61); +lean_ctor_set(x_4, 0, x_58); +x_62 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_62, 0, x_4); +lean_ctor_set(x_62, 1, x_57); +return x_62; } -lean_ctor_set(x_258, 0, x_9); -x_18 = x_258; -x_19 = x_160; -goto block_24; } +else +{ +uint8_t x_63; +lean_free_object(x_4); +x_63 = !lean_is_exclusive(x_49); +if (x_63 == 0) +{ +return x_49; +} +else +{ +lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_64 = lean_ctor_get(x_49, 0); +x_65 = lean_ctor_get(x_49, 1); +lean_inc(x_65); +lean_inc(x_64); +lean_dec(x_49); +x_66 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_66, 0, x_64); +lean_ctor_set(x_66, 1, x_65); +return x_66; } } } else { -lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; -lean_dec(x_156); -lean_dec(x_155); -lean_dec(x_154); -lean_free_object(x_9); -lean_dec(x_153); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); +lean_object* x_67; lean_object* x_68; +x_67 = lean_ctor_get(x_4, 0); +lean_inc(x_67); lean_dec(x_4); -lean_dec(x_2); -x_259 = lean_ctor_get(x_158, 0); -lean_inc(x_259); -x_260 = lean_ctor_get(x_158, 1); -lean_inc(x_260); -if (lean_is_exclusive(x_158)) { - lean_ctor_release(x_158, 0); - lean_ctor_release(x_158, 1); - x_261 = x_158; -} else { - lean_dec_ref(x_158); - x_261 = lean_box(0); -} -if (lean_is_scalar(x_261)) { - x_262 = lean_alloc_ctor(1, 2, 0); +x_68 = l_Lean_Meta_Grind_preprocessPattern(x_67, x_2, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_68) == 0) +{ +lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; +x_69 = lean_ctor_get(x_68, 0); +lean_inc(x_69); +x_70 = lean_ctor_get(x_68, 1); +lean_inc(x_70); +if (lean_is_exclusive(x_68)) { + lean_ctor_release(x_68, 0); + lean_ctor_release(x_68, 1); + x_71 = x_68; } else { - x_262 = x_261; -} -lean_ctor_set(x_262, 0, x_259); -lean_ctor_set(x_262, 1, x_260); -return x_262; + lean_dec_ref(x_68); + x_71 = lean_box(0); +} +x_72 = lean_array_get_size(x_3); +x_73 = lean_expr_abstract(x_69, x_3); +lean_dec(x_69); +x_74 = lean_box(0); +x_75 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_75, 0, x_73); +lean_ctor_set(x_75, 1, x_74); +x_76 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_76, 0, x_72); +lean_ctor_set(x_76, 1, x_75); +if (lean_is_scalar(x_71)) { + x_77 = lean_alloc_ctor(0, 2, 0); +} else { + x_77 = x_71; } +lean_ctor_set(x_77, 0, x_76); +lean_ctor_set(x_77, 1, x_70); +return x_77; } else { -lean_object* x_263; lean_object* x_264; lean_object* x_265; -lean_dec(x_156); -lean_dec(x_17); -if (lean_is_exclusive(x_157)) { - lean_ctor_release(x_157, 0); - x_263 = x_157; +lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; +x_78 = lean_ctor_get(x_68, 0); +lean_inc(x_78); +x_79 = lean_ctor_get(x_68, 1); +lean_inc(x_79); +if (lean_is_exclusive(x_68)) { + lean_ctor_release(x_68, 0); + lean_ctor_release(x_68, 1); + x_80 = x_68; } else { - lean_dec_ref(x_157); - x_263 = lean_box(0); + lean_dec_ref(x_68); + x_80 = lean_box(0); } -x_264 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_264, 0, x_154); -lean_ctor_set(x_264, 1, x_155); -lean_ctor_set(x_9, 1, x_264); -if (lean_is_scalar(x_263)) { - x_265 = lean_alloc_ctor(1, 1, 0); +if (lean_is_scalar(x_80)) { + x_81 = lean_alloc_ctor(1, 2, 0); } else { - x_265 = x_263; + x_81 = x_80; } -lean_ctor_set(x_265, 0, x_9); -x_18 = x_265; -x_19 = x_14; -goto block_24; +lean_ctor_set(x_81, 0, x_78); +lean_ctor_set(x_81, 1, x_79); +return x_81; } } } -else -{ -lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; -x_266 = lean_ctor_get(x_9, 1); -x_267 = lean_ctor_get(x_9, 0); -lean_inc(x_266); -lean_inc(x_267); -lean_dec(x_9); -x_268 = lean_ctor_get(x_266, 0); -lean_inc(x_268); -x_269 = lean_ctor_get(x_266, 1); -lean_inc(x_269); -if (lean_is_exclusive(x_266)) { - lean_ctor_release(x_266, 0); - lean_ctor_release(x_266, 1); - x_270 = x_266; -} else { - lean_dec_ref(x_266); - x_270 = lean_box(0); } -x_271 = l_Lean_Expr_fvarId_x21(x_17); -x_272 = l_Lean_RBNode_findCore___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__2(x_269, x_271); -if (lean_obj_tag(x_272) == 0) +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchEqTheoremCore___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_273; -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -x_273 = lean_infer_type(x_17, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_273) == 0) +lean_object* x_9; lean_object* x_10; +x_9 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_9, 0, x_2); +lean_ctor_set(x_9, 1, x_3); +x_10 = lean_apply_6(x_1, x_9, x_4, x_5, x_6, x_7, x_8); +return x_10; +} +} +static lean_object* _init_l_Lean_Meta_Grind_mkEMatchEqTheoremCore___lambda__3___closed__1() { +_start: { -lean_object* x_274; lean_object* x_275; lean_object* x_276; -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 = l_Lean_RBNode_findCore___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__2(x_267, x_271); -if (lean_obj_tag(x_276) == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("invalid E-matching equality theorem, conclusion must be an equality", 67, 67); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_mkEMatchEqTheoremCore___lambda__3___closed__2() { +_start: { -lean_object* x_277; -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_274); -x_277 = l_Lean_Meta_isProp(x_274, x_10, x_11, x_12, x_13, x_275); -if (lean_obj_tag(x_277) == 0) +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_mkEMatchEqTheoremCore___lambda__3___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchEqTheoremCore___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_278; uint8_t x_279; -x_278 = lean_ctor_get(x_277, 0); -lean_inc(x_278); -x_279 = lean_unbox(x_278); -lean_dec(x_278); -if (x_279 == 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; uint8_t x_15; +x_9 = l_Lean_indentExpr(x_1); +x_10 = l_Lean_Meta_Grind_mkEMatchEqTheoremCore___lambda__3___closed__2; +x_11 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_11, 0, x_10); +lean_ctor_set(x_11, 1, x_9); +x_12 = l_Lean_Meta_Grind_ppPattern___closed__4; +x_13 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_13, 0, x_11); +lean_ctor_set(x_13, 1, x_12); +x_14 = l_Lean_throwError___at_Lean_Meta_Grind_mkEMatchEqTheoremCore___spec__1(x_13, x_4, x_5, x_6, x_7, x_8); +x_15 = !lean_is_exclusive(x_14); +if (x_15 == 0) { -lean_object* x_280; lean_object* x_281; -x_280 = lean_ctor_get(x_277, 1); -lean_inc(x_280); -lean_dec(x_277); -lean_inc(x_10); -lean_inc(x_271); -x_281 = l_Lean_FVarId_getDecl(x_271, x_10, x_11, x_12, x_13, x_280); -if (lean_obj_tag(x_281) == 0) +return x_14; +} +else { -lean_object* x_282; lean_object* x_283; uint8_t x_284; lean_object* x_285; -x_282 = lean_ctor_get(x_281, 0); -lean_inc(x_282); -x_283 = lean_ctor_get(x_281, 1); -lean_inc(x_283); -lean_dec(x_281); -x_284 = l_Lean_LocalDecl_binderInfo(x_282); -lean_dec(x_282); -x_285 = lean_box(x_284); -if (lean_obj_tag(x_285) == 3) +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_14, 0); +x_17 = lean_ctor_get(x_14, 1); +lean_inc(x_17); +lean_inc(x_16); +lean_dec(x_14); +x_18 = lean_alloc_ctor(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_Grind_mkEMatchEqTheoremCore___lambda__4(uint8_t x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: { -lean_object* x_286; -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_274); -lean_inc(x_267); +lean_object* x_10; 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_box(x_1); +x_11 = lean_box(x_2); +x_12 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_mkEMatchEqTheoremCore___lambda__1___boxed), 9, 3); +lean_closure_set(x_12, 0, x_10); +lean_closure_set(x_12, 1, x_11); +lean_closure_set(x_12, 2, x_3); lean_inc(x_4); -x_286 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized(x_4, x_267, x_274, x_10, x_11, x_12, x_13, x_283); -if (lean_obj_tag(x_286) == 0) -{ -lean_object* x_287; uint8_t x_288; -x_287 = lean_ctor_get(x_286, 0); -lean_inc(x_287); -x_288 = lean_unbox(x_287); -lean_dec(x_287); -if (x_288 == 0) +x_13 = l_Lean_Meta_instantiateMVarsIfMVarApp(x_4, x_5, x_6, 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_16 = l_Lean_Expr_cleanupAnnotations(x_14); +x_17 = l_Lean_Expr_isApp(x_16); +if (x_17 == 0) { -lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; -lean_dec(x_274); -lean_dec(x_271); -x_289 = lean_ctor_get(x_286, 1); -lean_inc(x_289); -lean_dec(x_286); -if (lean_is_scalar(x_270)) { - x_290 = lean_alloc_ctor(0, 2, 0); -} else { - x_290 = x_270; -} -lean_ctor_set(x_290, 0, x_268); -lean_ctor_set(x_290, 1, x_269); -x_291 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_291, 0, x_267); -lean_ctor_set(x_291, 1, x_290); -x_292 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_292, 0, x_291); -x_18 = x_292; -x_19 = x_289; -goto block_24; +lean_object* x_18; lean_object* x_19; +lean_dec(x_16); +x_18 = lean_box(0); +x_19 = l_Lean_Meta_Grind_mkEMatchEqTheoremCore___lambda__3(x_4, x_12, x_18, x_5, x_6, x_7, x_8, x_15); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_12); +return x_19; } else { -lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; uint8_t x_304; lean_object* x_305; -lean_dec(x_268); -x_293 = lean_ctor_get(x_286, 1); -lean_inc(x_293); -lean_dec(x_286); -x_294 = lean_box(0); -lean_inc(x_271); -x_295 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_267, x_271, x_294); -x_296 = lean_box(0); -lean_inc(x_2); -x_297 = lean_array_mk(x_2); -x_298 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; -x_299 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_299, 0, x_298); -lean_ctor_set(x_299, 1, x_296); -lean_ctor_set(x_299, 2, x_297); -x_300 = l_Lean_CollectFVars_main(x_274, x_299); -x_301 = lean_ctor_get(x_300, 2); -lean_inc(x_301); -lean_dec(x_300); -x_302 = lean_array_get_size(x_301); -x_303 = lean_unsigned_to_nat(0u); -x_304 = lean_nat_dec_lt(x_303, x_302); -x_305 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_269, x_271, x_294); -if (x_304 == 0) +lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_20 = l_Lean_Expr_appArg(x_16, lean_box(0)); +x_21 = l_Lean_Expr_appFnCleanup(x_16, lean_box(0)); +x_22 = l_Lean_Expr_isApp(x_21); +if (x_22 == 0) { -uint8_t x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; -lean_dec(x_302); -lean_dec(x_301); -x_306 = 1; -x_307 = lean_box(x_306); -if (lean_is_scalar(x_270)) { - x_308 = lean_alloc_ctor(0, 2, 0); -} else { - x_308 = x_270; +lean_object* x_23; lean_object* x_24; +lean_dec(x_21); +lean_dec(x_20); +x_23 = lean_box(0); +x_24 = l_Lean_Meta_Grind_mkEMatchEqTheoremCore___lambda__3(x_4, x_12, x_23, x_5, x_6, x_7, x_8, x_15); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_12); +return x_24; } -lean_ctor_set(x_308, 0, x_307); -lean_ctor_set(x_308, 1, x_305); -x_309 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_309, 0, x_295); -lean_ctor_set(x_309, 1, x_308); -x_310 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_310, 0, x_309); -x_18 = x_310; -x_19 = x_293; -goto block_24; +else +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; +x_25 = l_Lean_Expr_appArg(x_21, lean_box(0)); +x_26 = l_Lean_Expr_appFnCleanup(x_21, lean_box(0)); +x_27 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__6; +x_28 = l_Lean_Expr_isConstOf(x_26, x_27); +if (x_28 == 0) +{ +uint8_t x_29; +x_29 = l_Lean_Expr_isApp(x_26); +if (x_29 == 0) +{ +lean_object* x_30; lean_object* x_31; +lean_dec(x_26); +lean_dec(x_25); +lean_dec(x_20); +x_30 = lean_box(0); +x_31 = l_Lean_Meta_Grind_mkEMatchEqTheoremCore___lambda__3(x_4, x_12, x_30, x_5, x_6, x_7, x_8, x_15); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_12); +return x_31; } else { -uint8_t x_311; -x_311 = lean_nat_dec_le(x_302, x_302); -if (x_311 == 0) +lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; +x_32 = l_Lean_Expr_appArg(x_26, lean_box(0)); +x_33 = l_Lean_Expr_appFnCleanup(x_26, lean_box(0)); +x_34 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__2; +x_35 = l_Lean_Expr_isConstOf(x_33, x_34); +if (x_35 == 0) { -uint8_t x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; -lean_dec(x_302); -lean_dec(x_301); -x_312 = 1; -x_313 = lean_box(x_312); -if (lean_is_scalar(x_270)) { - x_314 = lean_alloc_ctor(0, 2, 0); -} else { - x_314 = x_270; +uint8_t x_36; +lean_dec(x_25); +x_36 = l_Lean_Expr_isApp(x_33); +if (x_36 == 0) +{ +lean_object* x_37; lean_object* x_38; +lean_dec(x_33); +lean_dec(x_32); +lean_dec(x_20); +x_37 = lean_box(0); +x_38 = l_Lean_Meta_Grind_mkEMatchEqTheoremCore___lambda__3(x_4, x_12, x_37, x_5, x_6, x_7, x_8, x_15); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_12); +return x_38; } -lean_ctor_set(x_314, 0, x_313); -lean_ctor_set(x_314, 1, x_305); -x_315 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_315, 0, x_295); -lean_ctor_set(x_315, 1, x_314); -x_316 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_316, 0, x_315); -x_18 = x_316; -x_19 = x_293; -goto block_24; +else +{ +lean_object* x_39; lean_object* x_40; uint8_t x_41; +x_39 = l_Lean_Expr_appFnCleanup(x_33, lean_box(0)); +x_40 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__4; +x_41 = l_Lean_Expr_isConstOf(x_39, x_40); +lean_dec(x_39); +if (x_41 == 0) +{ +lean_object* x_42; lean_object* x_43; +lean_dec(x_32); +lean_dec(x_20); +x_42 = lean_box(0); +x_43 = l_Lean_Meta_Grind_mkEMatchEqTheoremCore___lambda__3(x_4, x_12, x_42, x_5, x_6, x_7, x_8, x_15); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_12); +return x_43; } else { -size_t x_317; size_t x_318; lean_object* x_319; uint8_t x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; -x_317 = 0; -x_318 = lean_usize_of_nat(x_302); -lean_dec(x_302); -x_319 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__10___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__11(x_4, x_301, x_317, x_318, x_295); -lean_dec(x_301); -x_320 = 1; -x_321 = lean_box(x_320); -if (lean_is_scalar(x_270)) { - x_322 = lean_alloc_ctor(0, 2, 0); -} else { - x_322 = x_270; +lean_object* x_44; +lean_dec(x_4); +x_44 = l_Lean_Meta_Grind_mkEMatchEqTheoremCore___lambda__2(x_12, x_32, x_20, x_5, x_6, x_7, x_8, x_15); +return x_44; } -lean_ctor_set(x_322, 0, x_321); -lean_ctor_set(x_322, 1, x_305); -x_323 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_323, 0, x_319); -lean_ctor_set(x_323, 1, x_322); -x_324 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_324, 0, x_323); -x_18 = x_324; -x_19 = x_293; -goto block_24; } } +else +{ +lean_object* x_45; +lean_dec(x_33); +lean_dec(x_32); +lean_dec(x_4); +x_45 = l_Lean_Meta_Grind_mkEMatchEqTheoremCore___lambda__2(x_12, x_25, x_20, x_5, x_6, x_7, x_8, x_15); +return x_45; +} } } else { -lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; -lean_dec(x_274); -lean_dec(x_271); -lean_dec(x_270); -lean_dec(x_269); -lean_dec(x_268); -lean_dec(x_267); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); +lean_object* x_46; +lean_dec(x_26); lean_dec(x_4); -lean_dec(x_2); -x_325 = lean_ctor_get(x_286, 0); -lean_inc(x_325); -x_326 = lean_ctor_get(x_286, 1); -lean_inc(x_326); -if (lean_is_exclusive(x_286)) { - lean_ctor_release(x_286, 0); - lean_ctor_release(x_286, 1); - x_327 = x_286; -} else { - lean_dec_ref(x_286); - x_327 = lean_box(0); +x_46 = l_Lean_Meta_Grind_mkEMatchEqTheoremCore___lambda__2(x_12, x_25, x_20, x_5, x_6, x_7, x_8, x_15); +return x_46; } -if (lean_is_scalar(x_327)) { - x_328 = lean_alloc_ctor(1, 2, 0); -} else { - x_328 = x_327; } -lean_ctor_set(x_328, 0, x_325); -lean_ctor_set(x_328, 1, x_326); -return x_328; } } +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchEqTheoremCore(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_3); +x_11 = lean_infer_type(x_3, 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; lean_object* x_15; lean_object* x_16; uint8_t x_17; lean_object* x_18; +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +x_13 = lean_ctor_get(x_11, 1); +lean_inc(x_13); +lean_dec(x_11); +x_14 = lean_box(x_5); +x_15 = lean_box(x_4); +x_16 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_mkEMatchEqTheoremCore___lambda__4___boxed), 9, 2); +lean_closure_set(x_16, 0, x_14); +lean_closure_set(x_16, 1, x_15); +x_17 = 0; +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +x_18 = l_Lean_Meta_forallTelescopeReducing___at_Lean_Meta_getParamNames___spec__2___rarg(x_12, x_16, x_17, x_6, x_7, x_8, x_9, x_13); +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; +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); +x_22 = lean_ctor_get(x_19, 1); +lean_inc(x_22); +lean_dec(x_19); +x_23 = l_Lean_Meta_Grind_mkEMatchTheoremCore(x_1, x_2, x_21, x_3, x_22, x_6, x_7, x_8, x_9, x_20); +return x_23; +} +else +{ +uint8_t x_24; +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_24 = !lean_is_exclusive(x_18); +if (x_24 == 0) +{ +return x_18; +} else { -lean_object* x_329; lean_object* x_330; lean_object* x_331; -lean_dec(x_285); -lean_dec(x_274); -lean_dec(x_271); -if (lean_is_scalar(x_270)) { - x_329 = lean_alloc_ctor(0, 2, 0); -} else { - x_329 = x_270; +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_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; } -lean_ctor_set(x_329, 0, x_268); -lean_ctor_set(x_329, 1, x_269); -x_330 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_330, 0, x_267); -lean_ctor_set(x_330, 1, x_329); -x_331 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_331, 0, x_330); -x_18 = x_331; -x_19 = x_283; -goto block_24; } } else { -lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; -lean_dec(x_274); -lean_dec(x_271); -lean_dec(x_270); -lean_dec(x_269); -lean_dec(x_268); -lean_dec(x_267); -lean_dec(x_13); -lean_dec(x_12); +uint8_t x_28; +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_28 = !lean_is_exclusive(x_11); +if (x_28 == 0) +{ +return x_11; +} +else +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_29 = lean_ctor_get(x_11, 0); +x_30 = lean_ctor_get(x_11, 1); +lean_inc(x_30); +lean_inc(x_29); lean_dec(x_11); -lean_dec(x_10); +x_31 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_31, 0, x_29); +lean_ctor_set(x_31, 1, x_30); +return x_31; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_Grind_mkEMatchEqTheoremCore___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_Meta_Grind_mkEMatchEqTheoremCore___spec__1(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchEqTheoremCore___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; uint8_t x_11; lean_object* x_12; +x_10 = lean_unbox(x_1); +lean_dec(x_1); +x_11 = lean_unbox(x_2); +lean_dec(x_2); +x_12 = l_Lean_Meta_Grind_mkEMatchEqTheoremCore___lambda__1(x_10, x_11, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_3); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchEqTheoremCore___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_Lean_Meta_Grind_mkEMatchEqTheoremCore___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchEqTheoremCore___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +uint8_t x_10; uint8_t x_11; lean_object* x_12; +x_10 = lean_unbox(x_1); +lean_dec(x_1); +x_11 = lean_unbox(x_2); lean_dec(x_2); -x_332 = lean_ctor_get(x_281, 0); -lean_inc(x_332); -x_333 = lean_ctor_get(x_281, 1); -lean_inc(x_333); -if (lean_is_exclusive(x_281)) { - lean_ctor_release(x_281, 0); - lean_ctor_release(x_281, 1); - x_334 = x_281; -} else { - lean_dec_ref(x_281); - x_334 = lean_box(0); +x_12 = l_Lean_Meta_Grind_mkEMatchEqTheoremCore___lambda__4(x_10, x_11, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_12; } -if (lean_is_scalar(x_334)) { - x_335 = lean_alloc_ctor(1, 2, 0); -} else { - x_335 = x_334; } -lean_ctor_set(x_335, 0, x_332); -lean_ctor_set(x_335, 1, x_333); -return x_335; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchEqTheoremCore___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +uint8_t x_11; uint8_t x_12; lean_object* x_13; +x_11 = lean_unbox(x_4); +lean_dec(x_4); +x_12 = lean_unbox(x_5); +lean_dec(x_5); +x_13 = l_Lean_Meta_Grind_mkEMatchEqTheoremCore(x_1, x_2, x_3, x_11, x_12, x_6, x_7, x_8, x_9, x_10); +return x_13; } } -else +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchEqTheorem(lean_object* x_1, uint8_t x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: { -lean_object* x_336; uint8_t x_337; -x_336 = lean_ctor_get(x_277, 1); -lean_inc(x_336); -lean_dec(x_277); -x_337 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars(x_4, x_267, x_274); -if (x_337 == 0) +lean_object* x_9; +lean_inc(x_1); +x_9 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getProofFor(x_1, x_6, x_7, x_8); +if (lean_obj_tag(x_9) == 0) { -lean_object* x_338; lean_object* x_339; lean_object* x_340; -lean_dec(x_271); -if (lean_is_scalar(x_270)) { - x_338 = lean_alloc_ctor(0, 2, 0); -} else { - x_338 = x_270; +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_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_alloc_ctor(0, 1, 0); +lean_ctor_set(x_12, 0, x_1); +x_13 = l_Lean_Meta_Grind_NormalizePattern_main___closed__1; +x_14 = l_Lean_Meta_Grind_mkEMatchEqTheoremCore(x_12, x_13, x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_11); +return x_14; } -lean_ctor_set(x_338, 0, x_268); -lean_ctor_set(x_338, 1, x_269); -x_339 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_339, 0, x_267); -lean_ctor_set(x_339, 1, x_338); -x_340 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_340, 0, x_339); -x_18 = x_340; -x_19 = x_336; -goto block_24; +else +{ +uint8_t x_15; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_15 = !lean_is_exclusive(x_9); +if (x_15 == 0) +{ +return x_9; } else { -lean_object* x_341; lean_object* x_342; lean_object* x_343; uint8_t x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; -lean_dec(x_268); -x_341 = lean_box(0); -lean_inc(x_271); -x_342 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_267, x_271, x_341); -x_343 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_269, x_271, x_341); -x_344 = 1; -x_345 = lean_box(x_344); -if (lean_is_scalar(x_270)) { - x_346 = lean_alloc_ctor(0, 2, 0); -} else { - x_346 = x_270; +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_9, 0); +x_17 = lean_ctor_get(x_9, 1); +lean_inc(x_17); +lean_inc(x_16); +lean_dec(x_9); +x_18 = lean_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_ctor_set(x_346, 0, x_345); -lean_ctor_set(x_346, 1, x_343); -x_347 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_347, 0, x_342); -lean_ctor_set(x_347, 1, x_346); -x_348 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_348, 0, x_347); -x_18 = x_348; -x_19 = x_336; -goto block_24; } } } -else +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchEqTheorem___boxed(lean_object* x_1, lean_object* x_2, lean_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_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; -lean_dec(x_274); -lean_dec(x_271); -lean_dec(x_270); -lean_dec(x_269); -lean_dec(x_268); -lean_dec(x_267); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_4); +uint8_t x_9; uint8_t x_10; lean_object* x_11; +x_9 = lean_unbox(x_2); lean_dec(x_2); -x_349 = lean_ctor_get(x_277, 0); -lean_inc(x_349); -x_350 = lean_ctor_get(x_277, 1); -lean_inc(x_350); -if (lean_is_exclusive(x_277)) { - lean_ctor_release(x_277, 0); - lean_ctor_release(x_277, 1); - x_351 = x_277; -} else { - lean_dec_ref(x_277); - x_351 = lean_box(0); -} -if (lean_is_scalar(x_351)) { - x_352 = lean_alloc_ctor(1, 2, 0); -} else { - x_352 = x_351; -} -lean_ctor_set(x_352, 0, x_349); -lean_ctor_set(x_352, 1, x_350); -return x_352; +x_10 = lean_unbox(x_3); +lean_dec(x_3); +x_11 = l_Lean_Meta_Grind_mkEMatchEqTheorem(x_1, x_9, x_10, x_4, x_5, x_6, x_7, x_8); +return x_11; } } -else +static lean_object* _init_l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__1___closed__1() { +_start: { -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; uint8_t x_362; lean_object* x_363; lean_object* x_364; -lean_dec(x_268); -if (lean_is_exclusive(x_276)) { - lean_ctor_release(x_276, 0); - x_353 = x_276; -} else { - lean_dec_ref(x_276); - x_353 = lean_box(0); +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedEMatchTheorems___spec__1___closed__2; +x_2 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_2, 0, x_1); +lean_ctor_set(x_2, 1, x_1); +return x_2; } -x_354 = lean_box(0); -lean_inc(x_2); -x_355 = lean_array_mk(x_2); -x_356 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; -x_357 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_357, 0, x_356); -lean_ctor_set(x_357, 1, x_354); -lean_ctor_set(x_357, 2, x_355); -x_358 = l_Lean_CollectFVars_main(x_274, x_357); -x_359 = lean_ctor_get(x_358, 2); -lean_inc(x_359); -lean_dec(x_358); -x_360 = lean_array_get_size(x_359); -x_361 = lean_unsigned_to_nat(0u); -x_362 = lean_nat_dec_lt(x_361, x_360); -x_363 = lean_box(0); -x_364 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_269, x_271, x_363); -if (x_362 == 0) -{ -uint8_t x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; lean_object* x_369; -lean_dec(x_360); -lean_dec(x_359); -x_365 = 1; -x_366 = lean_box(x_365); -if (lean_is_scalar(x_270)) { - x_367 = lean_alloc_ctor(0, 2, 0); -} else { - x_367 = x_270; } -lean_ctor_set(x_367, 0, x_366); -lean_ctor_set(x_367, 1, x_364); -x_368 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_368, 0, x_267); -lean_ctor_set(x_368, 1, x_367); -if (lean_is_scalar(x_353)) { - x_369 = lean_alloc_ctor(1, 1, 0); -} else { - x_369 = x_353; +static lean_object* _init_l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedEMatchTheorems___spec__1___closed__2; +x_2 = lean_alloc_ctor(0, 6, 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); +lean_ctor_set(x_2, 4, x_1); +lean_ctor_set(x_2, 5, x_1); +return x_2; } -lean_ctor_set(x_369, 0, x_368); -x_18 = x_369; -x_19 = x_275; -goto block_24; } -else +LEAN_EXPORT lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___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, lean_object* x_7, lean_object* x_8) { +_start: { -uint8_t x_370; -x_370 = lean_nat_dec_le(x_360, x_360); -if (x_370 == 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, 6); +lean_inc(x_9); +lean_dec(x_6); +x_10 = lean_st_ref_take(x_7, 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_13 = !lean_is_exclusive(x_11); +if (x_13 == 0) { -uint8_t x_371; lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; -lean_dec(x_360); -lean_dec(x_359); -x_371 = 1; -x_372 = lean_box(x_371); -if (lean_is_scalar(x_270)) { - x_373 = lean_alloc_ctor(0, 2, 0); -} else { - x_373 = x_270; +lean_object* x_14; lean_object* x_15; lean_object* x_16; 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_14 = lean_ctor_get(x_11, 0); +x_15 = lean_ctor_get(x_11, 4); +lean_dec(x_15); +x_16 = l_Lean_ScopedEnvExtension_addCore___rarg(x_14, x_1, x_2, x_3, x_9); +x_17 = l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__1___closed__1; +lean_ctor_set(x_11, 4, x_17); +lean_ctor_set(x_11, 0, x_16); +x_18 = lean_st_ref_set(x_7, x_11, x_12); +x_19 = lean_ctor_get(x_18, 1); +lean_inc(x_19); +lean_dec(x_18); +x_20 = lean_st_ref_take(x_5, x_19); +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; uint8_t x_27; +x_24 = lean_ctor_get(x_21, 1); +lean_dec(x_24); +x_25 = l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__1___closed__2; +lean_ctor_set(x_21, 1, x_25); +x_26 = lean_st_ref_set(x_5, x_21, x_22); +x_27 = !lean_is_exclusive(x_26); +if (x_27 == 0) +{ +lean_object* x_28; lean_object* x_29; +x_28 = lean_ctor_get(x_26, 0); +lean_dec(x_28); +x_29 = lean_box(0); +lean_ctor_set(x_26, 0, x_29); +return x_26; } -lean_ctor_set(x_373, 0, x_372); -lean_ctor_set(x_373, 1, x_364); -x_374 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_374, 0, x_267); -lean_ctor_set(x_374, 1, x_373); -if (lean_is_scalar(x_353)) { - x_375 = lean_alloc_ctor(1, 1, 0); -} else { - x_375 = x_353; +else +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_30 = lean_ctor_get(x_26, 1); +lean_inc(x_30); +lean_dec(x_26); +x_31 = lean_box(0); +x_32 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_30); +return x_32; } -lean_ctor_set(x_375, 0, x_374); -x_18 = x_375; -x_19 = x_275; -goto block_24; } else { -size_t x_376; size_t x_377; lean_object* x_378; uint8_t x_379; lean_object* x_380; lean_object* x_381; lean_object* x_382; lean_object* x_383; -x_376 = 0; -x_377 = lean_usize_of_nat(x_360); -lean_dec(x_360); -x_378 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__12___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__13(x_4, x_359, x_376, x_377, x_267); -lean_dec(x_359); -x_379 = 1; -x_380 = lean_box(x_379); -if (lean_is_scalar(x_270)) { - x_381 = lean_alloc_ctor(0, 2, 0); +lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_33 = lean_ctor_get(x_21, 0); +x_34 = lean_ctor_get(x_21, 2); +x_35 = lean_ctor_get(x_21, 3); +x_36 = lean_ctor_get(x_21, 4); +lean_inc(x_36); +lean_inc(x_35); +lean_inc(x_34); +lean_inc(x_33); +lean_dec(x_21); +x_37 = l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__1___closed__2; +x_38 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_38, 0, x_33); +lean_ctor_set(x_38, 1, x_37); +lean_ctor_set(x_38, 2, x_34); +lean_ctor_set(x_38, 3, x_35); +lean_ctor_set(x_38, 4, x_36); +x_39 = lean_st_ref_set(x_5, x_38, x_22); +x_40 = lean_ctor_get(x_39, 1); +lean_inc(x_40); +if (lean_is_exclusive(x_39)) { + lean_ctor_release(x_39, 0); + lean_ctor_release(x_39, 1); + x_41 = x_39; } else { - x_381 = x_270; + lean_dec_ref(x_39); + x_41 = lean_box(0); } -lean_ctor_set(x_381, 0, x_380); -lean_ctor_set(x_381, 1, x_364); -x_382 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_382, 0, x_378); -lean_ctor_set(x_382, 1, x_381); -if (lean_is_scalar(x_353)) { - x_383 = lean_alloc_ctor(1, 1, 0); +x_42 = lean_box(0); +if (lean_is_scalar(x_41)) { + x_43 = lean_alloc_ctor(0, 2, 0); } else { - x_383 = x_353; -} -lean_ctor_set(x_383, 0, x_382); -x_18 = x_383; -x_19 = x_275; -goto block_24; -} + x_43 = x_41; } +lean_ctor_set(x_43, 0, x_42); +lean_ctor_set(x_43, 1, x_40); +return x_43; } } else { -lean_object* x_384; lean_object* x_385; lean_object* x_386; lean_object* x_387; -lean_dec(x_271); -lean_dec(x_270); -lean_dec(x_269); -lean_dec(x_268); -lean_dec(x_267); -lean_dec(x_13); -lean_dec(x_12); +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; 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_44 = lean_ctor_get(x_11, 0); +x_45 = lean_ctor_get(x_11, 1); +x_46 = lean_ctor_get(x_11, 2); +x_47 = lean_ctor_get(x_11, 3); +x_48 = lean_ctor_get(x_11, 5); +x_49 = lean_ctor_get(x_11, 6); +x_50 = lean_ctor_get(x_11, 7); +lean_inc(x_50); +lean_inc(x_49); +lean_inc(x_48); +lean_inc(x_47); +lean_inc(x_46); +lean_inc(x_45); +lean_inc(x_44); lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_4); -lean_dec(x_2); -x_384 = lean_ctor_get(x_273, 0); -lean_inc(x_384); -x_385 = lean_ctor_get(x_273, 1); -lean_inc(x_385); -if (lean_is_exclusive(x_273)) { - lean_ctor_release(x_273, 0); - lean_ctor_release(x_273, 1); - x_386 = x_273; +x_51 = l_Lean_ScopedEnvExtension_addCore___rarg(x_44, x_1, x_2, x_3, x_9); +x_52 = l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__1___closed__1; +x_53 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_53, 0, x_51); +lean_ctor_set(x_53, 1, x_45); +lean_ctor_set(x_53, 2, x_46); +lean_ctor_set(x_53, 3, x_47); +lean_ctor_set(x_53, 4, x_52); +lean_ctor_set(x_53, 5, x_48); +lean_ctor_set(x_53, 6, x_49); +lean_ctor_set(x_53, 7, x_50); +x_54 = lean_st_ref_set(x_7, x_53, x_12); +x_55 = lean_ctor_get(x_54, 1); +lean_inc(x_55); +lean_dec(x_54); +x_56 = lean_st_ref_take(x_5, x_55); +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_inc(x_59); +x_60 = lean_ctor_get(x_57, 2); +lean_inc(x_60); +x_61 = lean_ctor_get(x_57, 3); +lean_inc(x_61); +x_62 = lean_ctor_get(x_57, 4); +lean_inc(x_62); +if (lean_is_exclusive(x_57)) { + lean_ctor_release(x_57, 0); + lean_ctor_release(x_57, 1); + lean_ctor_release(x_57, 2); + lean_ctor_release(x_57, 3); + lean_ctor_release(x_57, 4); + x_63 = x_57; } else { - lean_dec_ref(x_273); - x_386 = lean_box(0); + lean_dec_ref(x_57); + x_63 = lean_box(0); } -if (lean_is_scalar(x_386)) { - x_387 = lean_alloc_ctor(1, 2, 0); +x_64 = l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__1___closed__2; +if (lean_is_scalar(x_63)) { + x_65 = lean_alloc_ctor(0, 5, 0); } else { - x_387 = x_386; -} -lean_ctor_set(x_387, 0, x_384); -lean_ctor_set(x_387, 1, x_385); -return x_387; -} + x_65 = x_63; } -else -{ -lean_object* x_388; lean_object* x_389; lean_object* x_390; lean_object* x_391; -lean_dec(x_271); -lean_dec(x_17); -if (lean_is_exclusive(x_272)) { - lean_ctor_release(x_272, 0); - x_388 = x_272; +lean_ctor_set(x_65, 0, x_59); +lean_ctor_set(x_65, 1, x_64); +lean_ctor_set(x_65, 2, x_60); +lean_ctor_set(x_65, 3, x_61); +lean_ctor_set(x_65, 4, x_62); +x_66 = lean_st_ref_set(x_5, x_65, x_58); +x_67 = lean_ctor_get(x_66, 1); +lean_inc(x_67); +if (lean_is_exclusive(x_66)) { + lean_ctor_release(x_66, 0); + lean_ctor_release(x_66, 1); + x_68 = x_66; } else { - lean_dec_ref(x_272); - x_388 = lean_box(0); + lean_dec_ref(x_66); + x_68 = lean_box(0); } -if (lean_is_scalar(x_270)) { - x_389 = lean_alloc_ctor(0, 2, 0); +x_69 = lean_box(0); +if (lean_is_scalar(x_68)) { + x_70 = lean_alloc_ctor(0, 2, 0); } else { - x_389 = x_270; + x_70 = x_68; } -lean_ctor_set(x_389, 0, x_268); -lean_ctor_set(x_389, 1, x_269); -x_390 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_390, 0, x_267); -lean_ctor_set(x_390, 1, x_389); -if (lean_is_scalar(x_388)) { - x_391 = lean_alloc_ctor(1, 1, 0); -} else { - x_391 = x_388; +lean_ctor_set(x_70, 0, x_69); +lean_ctor_set(x_70, 1, x_67); +return x_70; } -lean_ctor_set(x_391, 0, x_390); -x_18 = x_391; -x_19 = x_14; -goto block_24; } } -block_24: +static lean_object* _init_l_Lean_Meta_Grind_addEMatchTheorem___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ematchTheoremsExt; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addEMatchTheorem(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +x_9 = l_Lean_Meta_Grind_mkEMatchTheorem(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_9) == 0) +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; +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_Meta_Grind_addEMatchTheorem___closed__1; +x_13 = 0; +x_14 = l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__1(x_12, x_10, x_13, x_4, x_5, x_6, x_7, x_11); +lean_dec(x_7); +lean_dec(x_5); +lean_dec(x_4); +return x_14; +} +else +{ +uint8_t x_15; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_15 = !lean_is_exclusive(x_9); +if (x_15 == 0) +{ +return x_9; +} +else { -lean_object* x_20; size_t x_21; size_t x_22; -x_20 = lean_ctor_get(x_18, 0); -lean_inc(x_20); -lean_dec(x_18); -x_21 = 1; -x_22 = lean_usize_add(x_8, x_21); -x_8 = x_22; -x_9 = x_20; -x_14 = x_19; -goto _start; +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_9, 0); +x_17 = lean_ctor_get(x_9, 1); +lean_inc(x_17); +lean_inc(x_16); +lean_dec(x_9); +x_18 = lean_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_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__14___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__15(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, size_t x_6, size_t x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { -uint8_t x_14; -x_14 = lean_usize_dec_lt(x_7, x_6); -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_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__1(x_1, x_2, x_9, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_7); +lean_dec(x_5); +lean_dec(x_4); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addEMatchEqTheorem(lean_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_15; -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); +uint8_t x_7; lean_object* x_8; +x_7 = 1; +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +x_8 = l_Lean_Meta_Grind_mkEMatchEqTheorem(x_1, x_7, x_7, x_2, x_3, x_4, x_5, x_6); +if (lean_obj_tag(x_8) == 0) +{ +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_8, 0); +lean_inc(x_9); +x_10 = lean_ctor_get(x_8, 1); +lean_inc(x_10); +lean_dec(x_8); +x_11 = l_Lean_Meta_Grind_addEMatchTheorem___closed__1; +x_12 = 0; +x_13 = l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__1(x_11, x_9, x_12, x_2, x_3, x_4, x_5, x_10); +lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); -x_15 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_15, 0, x_8); -lean_ctor_set(x_15, 1, x_13); -return x_15; +return x_13; } else { -lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_24; -x_16 = lean_array_uget(x_5, x_7); -x_24 = !lean_is_exclusive(x_8); -if (x_24 == 0) +uint8_t x_14; +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_14 = !lean_is_exclusive(x_8); +if (x_14 == 0) { -lean_object* x_25; uint8_t x_26; -x_25 = lean_ctor_get(x_8, 1); -x_26 = !lean_is_exclusive(x_25); -if (x_26 == 0) +return x_8; +} +else { -lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_27 = lean_ctor_get(x_8, 0); -x_28 = lean_ctor_get(x_25, 0); -x_29 = lean_ctor_get(x_25, 1); -x_30 = l_Lean_Expr_fvarId_x21(x_16); -x_31 = l_Lean_RBNode_findCore___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__2(x_29, x_30); -if (lean_obj_tag(x_31) == 0) +lean_object* x_15; lean_object* x_16; lean_object* x_17; +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_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_Meta_Grind_getEMatchTheorems___rarg(lean_object* x_1, lean_object* x_2) { +_start: { -lean_object* x_32; -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -x_32 = lean_infer_type(x_16, x_9, x_10, x_11, x_12, x_13); -if (lean_obj_tag(x_32) == 0) +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_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_RBNode_findCore___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__2(x_27, x_30); -if (lean_obj_tag(x_35) == 0) +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_5 = lean_ctor_get(x_3, 0); +x_6 = lean_ctor_get(x_5, 0); +lean_inc(x_6); +lean_dec(x_5); +x_7 = l_Lean_Meta_Grind_instInhabitedEMatchTheorems; +x_8 = l_Lean_Meta_Grind_addEMatchTheorem___closed__1; +x_9 = l_Lean_ScopedEnvExtension_getState___rarg(x_7, x_8, x_6); +lean_dec(x_6); +lean_ctor_set(x_3, 0, x_9); +return x_3; +} +else { -lean_object* x_36; -lean_inc(x_12); +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_10 = lean_ctor_get(x_3, 0); +x_11 = lean_ctor_get(x_3, 1); lean_inc(x_11); lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_33); -x_36 = l_Lean_Meta_isProp(x_33, x_9, x_10, x_11, x_12, x_34); -if (lean_obj_tag(x_36) == 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_39; lean_object* x_40; -x_39 = lean_ctor_get(x_36, 1); -lean_inc(x_39); -lean_dec(x_36); -lean_inc(x_9); -lean_inc(x_30); -x_40 = l_Lean_FVarId_getDecl(x_30, x_9, x_10, x_11, x_12, x_39); -if (lean_obj_tag(x_40) == 0) -{ -lean_object* x_41; lean_object* x_42; uint8_t 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 = l_Lean_LocalDecl_binderInfo(x_41); -lean_dec(x_41); -x_44 = lean_box(x_43); -if (lean_obj_tag(x_44) == 3) -{ -lean_object* x_45; +lean_dec(x_3); +x_12 = lean_ctor_get(x_10, 0); lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_33); -lean_inc(x_27); -lean_inc(x_3); -x_45 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized(x_3, x_27, x_33, x_9, x_10, x_11, x_12, x_42); -if (lean_obj_tag(x_45) == 0) +lean_dec(x_10); +x_13 = l_Lean_Meta_Grind_instInhabitedEMatchTheorems; +x_14 = l_Lean_Meta_Grind_addEMatchTheorem___closed__1; +x_15 = l_Lean_ScopedEnvExtension_getState___rarg(x_13, x_14, x_12); +lean_dec(x_12); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_15); +lean_ctor_set(x_16, 1, x_11); +return x_16; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getEMatchTheorems(lean_object* x_1) { +_start: { -lean_object* x_46; uint8_t x_47; -x_46 = lean_ctor_get(x_45, 0); -lean_inc(x_46); -x_47 = lean_unbox(x_46); -lean_dec(x_46); -if (x_47 == 0) +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_getEMatchTheorems___rarg___boxed), 2, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getEMatchTheorems___rarg___boxed(lean_object* x_1, lean_object* x_2) { +_start: { -lean_object* x_48; lean_object* x_49; -lean_dec(x_33); -lean_dec(x_30); -x_48 = lean_ctor_get(x_45, 1); -lean_inc(x_48); -lean_dec(x_45); -x_49 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_49, 0, x_8); -x_17 = x_49; -x_18 = x_48; -goto block_23; +lean_object* x_3; +x_3 = l_Lean_Meta_Grind_getEMatchTheorems___rarg(x_1, x_2); +lean_dec(x_1); +return x_3; } -else +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getEMatchTheorems___boxed(lean_object* x_1) { +_start: { -lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; 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_dec(x_28); -x_50 = lean_ctor_get(x_45, 1); -lean_inc(x_50); -lean_dec(x_45); -x_51 = lean_box(0); -lean_inc(x_30); -x_52 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_27, x_30, x_51); -x_53 = lean_box(0); -lean_inc(x_2); -x_54 = lean_array_mk(x_2); -x_55 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; -x_56 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_56, 0, x_55); -lean_ctor_set(x_56, 1, x_53); -lean_ctor_set(x_56, 2, x_54); -x_57 = l_Lean_CollectFVars_main(x_33, x_56); -x_58 = lean_ctor_get(x_57, 2); -lean_inc(x_58); -lean_dec(x_57); -x_59 = lean_array_get_size(x_58); -x_60 = lean_unsigned_to_nat(0u); -x_61 = lean_nat_dec_lt(x_60, x_59); -x_62 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_29, x_30, x_51); -if (x_61 == 0) +lean_object* x_2; +x_2 = l_Lean_Meta_Grind_getEMatchTheorems(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_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: { -uint8_t x_63; lean_object* x_64; lean_object* x_65; -lean_dec(x_59); -lean_dec(x_58); -x_63 = 1; -x_64 = lean_box(x_63); -lean_ctor_set(x_25, 1, x_62); -lean_ctor_set(x_25, 0, x_64); -lean_ctor_set(x_8, 0, x_52); -x_65 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_65, 0, x_8); -x_17 = x_65; -x_18 = x_50; -goto block_23; +lean_object* x_3; +x_3 = lean_unsigned_to_nat(1u); +return x_3; } -else +case 2: { -uint8_t x_66; -x_66 = lean_nat_dec_le(x_59, x_59); -if (x_66 == 0) +lean_object* x_4; +x_4 = lean_unsigned_to_nat(2u); +return x_4; +} +case 3: { -uint8_t x_67; lean_object* x_68; lean_object* x_69; -lean_dec(x_59); -lean_dec(x_58); -x_67 = 1; -x_68 = lean_box(x_67); -lean_ctor_set(x_25, 1, x_62); -lean_ctor_set(x_25, 0, x_68); -lean_ctor_set(x_8, 0, x_52); -x_69 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_69, 0, x_8); -x_17 = x_69; -x_18 = x_50; -goto block_23; +lean_object* x_5; +x_5 = lean_unsigned_to_nat(3u); +return x_5; } -else +case 4: { -size_t x_70; size_t x_71; lean_object* x_72; uint8_t x_73; lean_object* x_74; lean_object* x_75; -x_70 = 0; -x_71 = lean_usize_of_nat(x_59); -lean_dec(x_59); -x_72 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__10___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__11(x_3, x_58, x_70, x_71, x_52); -lean_dec(x_58); -x_73 = 1; -x_74 = lean_box(x_73); -lean_ctor_set(x_25, 1, x_62); -lean_ctor_set(x_25, 0, x_74); -lean_ctor_set(x_8, 0, x_72); -x_75 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_75, 0, x_8); -x_17 = x_75; -x_18 = x_50; -goto block_23; +lean_object* x_6; +x_6 = lean_unsigned_to_nat(4u); +return x_6; +} +default: +{ +lean_object* x_7; +x_7 = lean_unsigned_to_nat(5u); +return x_7; } } } } -else +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toCtorIdx___boxed(lean_object* x_1) { +_start: { -uint8_t x_76; -lean_dec(x_33); -lean_dec(x_30); -lean_free_object(x_25); -lean_dec(x_29); -lean_dec(x_28); -lean_free_object(x_8); -lean_dec(x_27); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_3); -lean_dec(x_2); -x_76 = !lean_is_exclusive(x_45); -if (x_76 == 0) +uint8_t x_2; lean_object* x_3; +x_2 = lean_unbox(x_1); +lean_dec(x_1); +x_3 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toCtorIdx(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_noConfusion___rarg___lambda__1(lean_object* x_1) { +_start: { -return x_45; +lean_inc(x_1); +return x_1; } -else +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_noConfusion___rarg___closed__1() { +_start: { -lean_object* x_77; lean_object* x_78; lean_object* x_79; -x_77 = lean_ctor_get(x_45, 0); -x_78 = lean_ctor_get(x_45, 1); -lean_inc(x_78); -lean_inc(x_77); -lean_dec(x_45); -x_79 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_79, 0, x_77); -lean_ctor_set(x_79, 1, x_78); -return x_79; +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_noConfusion___rarg___lambda__1___boxed), 1, 0); +return x_1; +} } +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_noConfusion___rarg(uint8_t x_1, uint8_t x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_noConfusion___rarg___closed__1; +return x_4; } } -else +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_noConfusion(lean_object* x_1) { +_start: { -lean_object* x_80; -lean_dec(x_44); -lean_dec(x_33); -lean_dec(x_30); -x_80 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_80, 0, x_8); -x_17 = x_80; -x_18 = x_42; -goto block_23; +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_noConfusion___rarg___boxed), 3, 0); +return x_2; } } -else +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_noConfusion___rarg___lambda__1___boxed(lean_object* x_1) { +_start: { -uint8_t x_81; -lean_dec(x_33); -lean_dec(x_30); -lean_free_object(x_25); -lean_dec(x_29); -lean_dec(x_28); -lean_free_object(x_8); -lean_dec(x_27); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_3); +lean_object* x_2; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_noConfusion___rarg___lambda__1(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_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_81 = !lean_is_exclusive(x_40); -if (x_81 == 0) +x_6 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_noConfusion___rarg(x_4, x_5, x_3); +return x_6; +} +} +static uint8_t _init_l_Lean_Meta_Grind_instInhabitedTheoremKind() { +_start: { -return x_40; +uint8_t x_1; +x_1 = 0; +return x_1; } -else +} +LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_beqTheoremKind____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_6540_(uint8_t x_1, uint8_t x_2) { +_start: { -lean_object* x_82; lean_object* x_83; lean_object* x_84; -x_82 = lean_ctor_get(x_40, 0); -x_83 = lean_ctor_get(x_40, 1); -lean_inc(x_83); -lean_inc(x_82); -lean_dec(x_40); -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_3; lean_object* x_4; uint8_t x_5; +x_3 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toCtorIdx(x_1); +x_4 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_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_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_beqTheoremKind____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_6540____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_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_beqTheoremKind____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_6540_(x_3, x_4); +x_6 = lean_box(x_5); +return x_6; } } -else +static lean_object* _init_l_Lean_Meta_Grind_instBEqTheoremKind___closed__1() { +_start: { -lean_object* x_85; uint8_t x_86; -x_85 = lean_ctor_get(x_36, 1); -lean_inc(x_85); -lean_dec(x_36); -x_86 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars(x_3, x_27, x_33); -if (x_86 == 0) +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_beqTheoremKind____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_6540____boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_instBEqTheoremKind() { +_start: { -lean_object* x_87; -lean_dec(x_30); -x_87 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_87, 0, x_8); -x_17 = x_87; -x_18 = x_85; -goto block_23; +lean_object* x_1; +x_1 = l_Lean_Meta_Grind_instBEqTheoremKind___closed__1; +return x_1; } -else +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute___closed__1() { +_start: { -lean_object* x_88; lean_object* x_89; lean_object* x_90; uint8_t x_91; lean_object* x_92; lean_object* x_93; -lean_dec(x_28); -x_88 = lean_box(0); -lean_inc(x_30); -x_89 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_27, x_30, x_88); -x_90 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_29, x_30, x_88); -x_91 = 1; -x_92 = lean_box(x_91); -lean_ctor_set(x_25, 1, x_90); -lean_ctor_set(x_25, 0, x_92); -lean_ctor_set(x_8, 0, x_89); -x_93 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_93, 0, x_8); -x_17 = x_93; -x_18 = x_85; -goto block_23; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("[grind =]", 9, 9); +return x_1; } } +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("[grind =_]", 10, 10); +return x_1; } -else +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute___closed__3() { +_start: { -uint8_t x_94; -lean_dec(x_33); -lean_dec(x_30); -lean_free_object(x_25); -lean_dec(x_29); -lean_dec(x_28); -lean_free_object(x_8); -lean_dec(x_27); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_3); -lean_dec(x_2); -x_94 = !lean_is_exclusive(x_36); -if (x_94 == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("[grind _=_]", 11, 11); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute___closed__4() { +_start: { -return x_36; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("[grind →]", 11, 9); +return x_1; +} } -else +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute___closed__5() { +_start: { -lean_object* x_95; lean_object* x_96; lean_object* x_97; -x_95 = lean_ctor_get(x_36, 0); -x_96 = lean_ctor_get(x_36, 1); -lean_inc(x_96); -lean_inc(x_95); -lean_dec(x_36); -x_97 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_97, 0, x_95); -lean_ctor_set(x_97, 1, x_96); -return x_97; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("[grind ←]", 11, 9); +return x_1; } } +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("[grind]", 7, 7); +return x_1; +} } -else +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute(uint8_t x_1) { +_start: { -uint8_t x_98; -lean_dec(x_28); -x_98 = !lean_is_exclusive(x_35); -if (x_98 == 0) +switch (x_1) { +case 0: { -lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; uint8_t x_108; lean_object* x_109; lean_object* x_110; -x_99 = lean_ctor_get(x_35, 0); -lean_dec(x_99); -x_100 = lean_box(0); -lean_inc(x_2); -x_101 = lean_array_mk(x_2); -x_102 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; -x_103 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_103, 0, x_102); -lean_ctor_set(x_103, 1, x_100); -lean_ctor_set(x_103, 2, x_101); -x_104 = l_Lean_CollectFVars_main(x_33, x_103); -x_105 = lean_ctor_get(x_104, 2); -lean_inc(x_105); -lean_dec(x_104); -x_106 = lean_array_get_size(x_105); -x_107 = lean_unsigned_to_nat(0u); -x_108 = lean_nat_dec_lt(x_107, x_106); -x_109 = lean_box(0); -x_110 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_29, x_30, x_109); -if (x_108 == 0) +lean_object* x_2; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute___closed__1; +return x_2; +} +case 1: { -uint8_t x_111; lean_object* x_112; -lean_dec(x_106); -lean_dec(x_105); -x_111 = 1; -x_112 = lean_box(x_111); -lean_ctor_set(x_25, 1, x_110); -lean_ctor_set(x_25, 0, x_112); -lean_ctor_set(x_35, 0, x_8); -x_17 = x_35; -x_18 = x_34; -goto block_23; +lean_object* x_3; +x_3 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute___closed__2; +return x_3; } -else +case 2: { -uint8_t x_113; -x_113 = lean_nat_dec_le(x_106, x_106); -if (x_113 == 0) +lean_object* x_4; +x_4 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute___closed__3; +return x_4; +} +case 3: { -uint8_t x_114; lean_object* x_115; -lean_dec(x_106); -lean_dec(x_105); -x_114 = 1; -x_115 = lean_box(x_114); -lean_ctor_set(x_25, 1, x_110); -lean_ctor_set(x_25, 0, x_115); -lean_ctor_set(x_35, 0, x_8); -x_17 = x_35; -x_18 = x_34; -goto block_23; +lean_object* x_5; +x_5 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute___closed__4; +return x_5; } -else +case 4: { -size_t x_116; size_t x_117; lean_object* x_118; uint8_t x_119; lean_object* x_120; -x_116 = 0; -x_117 = lean_usize_of_nat(x_106); -lean_dec(x_106); -x_118 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__12___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__13(x_3, x_105, x_116, x_117, x_27); -lean_dec(x_105); -x_119 = 1; -x_120 = lean_box(x_119); -lean_ctor_set(x_25, 1, x_110); -lean_ctor_set(x_25, 0, x_120); -lean_ctor_set(x_8, 0, x_118); -lean_ctor_set(x_35, 0, x_8); -x_17 = x_35; -x_18 = x_34; -goto block_23; +lean_object* x_6; +x_6 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute___closed__5; +return x_6; } +default: +{ +lean_object* x_7; +x_7 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute___closed__6; +return x_7; } } -else -{ -lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; uint8_t x_129; lean_object* x_130; lean_object* x_131; -lean_dec(x_35); -x_121 = lean_box(0); -lean_inc(x_2); -x_122 = lean_array_mk(x_2); -x_123 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; -x_124 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_124, 0, x_123); -lean_ctor_set(x_124, 1, x_121); -lean_ctor_set(x_124, 2, x_122); -x_125 = l_Lean_CollectFVars_main(x_33, x_124); -x_126 = lean_ctor_get(x_125, 2); -lean_inc(x_126); -lean_dec(x_125); -x_127 = lean_array_get_size(x_126); -x_128 = lean_unsigned_to_nat(0u); -x_129 = lean_nat_dec_lt(x_128, x_127); -x_130 = lean_box(0); -x_131 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_29, x_30, x_130); -if (x_129 == 0) +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute___boxed(lean_object* x_1) { +_start: { -uint8_t x_132; lean_object* x_133; lean_object* x_134; -lean_dec(x_127); -lean_dec(x_126); -x_132 = 1; -x_133 = lean_box(x_132); -lean_ctor_set(x_25, 1, x_131); -lean_ctor_set(x_25, 0, x_133); -x_134 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_134, 0, x_8); -x_17 = x_134; -x_18 = x_34; -goto block_23; +uint8_t x_2; lean_object* x_3; +x_2 = lean_unbox(x_1); +lean_dec(x_1); +x_3 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute(x_2); +return x_3; } -else +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__1() { +_start: { -uint8_t x_135; -x_135 = lean_nat_dec_le(x_127, x_127); -if (x_135 == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("failed to find pattern in the left-hand side of the theorem's conclusion", 72, 72); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__2() { +_start: { -uint8_t x_136; lean_object* x_137; lean_object* x_138; -lean_dec(x_127); -lean_dec(x_126); -x_136 = 1; -x_137 = lean_box(x_136); -lean_ctor_set(x_25, 1, x_131); -lean_ctor_set(x_25, 0, x_137); -x_138 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_138, 0, x_8); -x_17 = x_138; -x_18 = x_34; -goto block_23; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("failed to find pattern in the right-hand side of the theorem's conclusion", 73, 73); +return x_1; } -else +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__3() { +_start: { -size_t x_139; size_t x_140; lean_object* x_141; uint8_t x_142; lean_object* x_143; lean_object* x_144; -x_139 = 0; -x_140 = lean_usize_of_nat(x_127); -lean_dec(x_127); -x_141 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__12___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__13(x_3, x_126, x_139, x_140, x_27); -lean_dec(x_126); -x_142 = 1; -x_143 = lean_box(x_142); -lean_ctor_set(x_25, 1, x_131); -lean_ctor_set(x_25, 0, x_143); -lean_ctor_set(x_8, 0, x_141); -x_144 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_144, 0, x_8); -x_17 = x_144; -x_18 = x_34; -goto block_23; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("_private.Lean.Meta.Tactic.Grind.EMatchTheorem.0.Lean.Meta.Grind.TheoremKind.explainFailure", 90, 90); +return x_1; +} } +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___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_Meta_Grind_EMatchTheorems_insert___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__3; +x_3 = lean_unsigned_to_nat(517u); +x_4 = lean_unsigned_to_nat(16u); +x_5 = l_Lean_Meta_Grind_EMatchTheorems_insert___closed__3; +x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); +return x_6; } } +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("failed to find patterns in the antecedents of the theorem", 57, 57); +return x_1; } } -else +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__6() { +_start: { -uint8_t x_145; -lean_dec(x_30); -lean_free_object(x_25); -lean_dec(x_29); -lean_dec(x_28); -lean_free_object(x_8); -lean_dec(x_27); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_3); -lean_dec(x_2); -x_145 = !lean_is_exclusive(x_32); -if (x_145 == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("failed to find patterns in the theorem's conclusion", 51, 51); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__7() { +_start: { -return x_32; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("failed to find patterns", 23, 23); +return x_1; } -else +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure(uint8_t x_1) { +_start: { -lean_object* x_146; lean_object* x_147; lean_object* x_148; -x_146 = lean_ctor_get(x_32, 0); -x_147 = lean_ctor_get(x_32, 1); -lean_inc(x_147); -lean_inc(x_146); -lean_dec(x_32); -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; +switch (x_1) { +case 0: +{ +lean_object* x_2; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__1; +return x_2; } +case 1: +{ +lean_object* x_3; +x_3 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__2; +return x_3; } +case 2: +{ +lean_object* x_4; lean_object* x_5; +x_4 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__4; +x_5 = l_panic___at_String_fromUTF8_x21___spec__1(x_4); +return x_5; } -else +case 3: { -uint8_t x_149; -lean_dec(x_30); -lean_dec(x_16); -x_149 = !lean_is_exclusive(x_31); -if (x_149 == 0) +lean_object* x_6; +x_6 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__5; +return x_6; +} +case 4: { -lean_object* x_150; -x_150 = lean_ctor_get(x_31, 0); -lean_dec(x_150); -lean_ctor_set(x_31, 0, x_8); -x_17 = x_31; -x_18 = x_13; -goto block_23; +lean_object* x_7; +x_7 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__6; +return x_7; } -else +default: { -lean_object* x_151; -lean_dec(x_31); -x_151 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_151, 0, x_8); -x_17 = x_151; -x_18 = x_13; -goto block_23; +lean_object* x_8; +x_8 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__7; +return x_8; } } } -else +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___boxed(lean_object* x_1) { +_start: { -lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; -x_152 = lean_ctor_get(x_8, 0); -x_153 = lean_ctor_get(x_25, 0); -x_154 = lean_ctor_get(x_25, 1); -lean_inc(x_154); -lean_inc(x_153); -lean_dec(x_25); -x_155 = l_Lean_Expr_fvarId_x21(x_16); -x_156 = l_Lean_RBNode_findCore___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__2(x_154, x_155); -if (lean_obj_tag(x_156) == 0) +uint8_t x_2; lean_object* x_3; +x_2 = lean_unbox(x_1); +lean_dec(x_1); +x_3 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getPropTypes___spec__2(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: { -lean_object* x_157; -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -x_157 = lean_infer_type(x_16, x_9, x_10, x_11, x_12, x_13); -if (lean_obj_tag(x_157) == 0) +uint8_t x_10; +x_10 = lean_usize_dec_eq(x_2, x_3); +if (x_10 == 0) { -lean_object* x_158; lean_object* x_159; lean_object* 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 = l_Lean_RBNode_findCore___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__2(x_152, x_155); -if (lean_obj_tag(x_160) == 0) +lean_object* x_11; lean_object* x_12; +x_11 = lean_array_uget(x_1, x_2); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +x_12 = lean_infer_type(x_11, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_12) == 0) { -lean_object* x_161; -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_158); -x_161 = l_Lean_Meta_isProp(x_158, x_9, x_10, x_11, x_12, x_159); -if (lean_obj_tag(x_161) == 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_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_13); +x_15 = l_Lean_Meta_isProp(x_13, x_5, x_6, x_7, x_8, x_14); +if (lean_obj_tag(x_15) == 0) { -lean_object* x_162; uint8_t x_163; -x_162 = lean_ctor_get(x_161, 0); -lean_inc(x_162); -x_163 = lean_unbox(x_162); -lean_dec(x_162); -if (x_163 == 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_164; lean_object* x_165; -x_164 = lean_ctor_get(x_161, 1); -lean_inc(x_164); -lean_dec(x_161); -lean_inc(x_9); -lean_inc(x_155); -x_165 = l_Lean_FVarId_getDecl(x_155, x_9, x_10, x_11, x_12, x_164); -if (lean_obj_tag(x_165) == 0) +lean_object* x_18; size_t x_19; size_t x_20; +lean_dec(x_13); +x_18 = lean_ctor_get(x_15, 1); +lean_inc(x_18); +lean_dec(x_15); +x_19 = 1; +x_20 = lean_usize_add(x_2, x_19); +x_2 = x_20; +x_9 = x_18; +goto _start; +} +else { -lean_object* x_166; lean_object* x_167; uint8_t x_168; lean_object* x_169; -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_Lean_LocalDecl_binderInfo(x_166); -lean_dec(x_166); -x_169 = lean_box(x_168); -if (lean_obj_tag(x_169) == 3) +lean_object* x_22; lean_object* x_23; size_t x_24; size_t x_25; +x_22 = lean_ctor_get(x_15, 1); +lean_inc(x_22); +lean_dec(x_15); +x_23 = lean_array_push(x_4, x_13); +x_24 = 1; +x_25 = lean_usize_add(x_2, x_24); +x_2 = x_25; +x_4 = x_23; +x_9 = x_22; +goto _start; +} +} +else { -lean_object* x_170; -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_158); -lean_inc(x_152); -lean_inc(x_3); -x_170 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized(x_3, x_152, x_158, x_9, x_10, x_11, x_12, x_167); -if (lean_obj_tag(x_170) == 0) +uint8_t x_27; +lean_dec(x_13); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_27 = !lean_is_exclusive(x_15); +if (x_27 == 0) { -lean_object* x_171; uint8_t x_172; -x_171 = lean_ctor_get(x_170, 0); -lean_inc(x_171); -x_172 = lean_unbox(x_171); -lean_dec(x_171); -if (x_172 == 0) +return x_15; +} +else { -lean_object* x_173; lean_object* x_174; lean_object* x_175; -lean_dec(x_158); -lean_dec(x_155); -x_173 = lean_ctor_get(x_170, 1); -lean_inc(x_173); -lean_dec(x_170); -x_174 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_174, 0, x_153); -lean_ctor_set(x_174, 1, x_154); -lean_ctor_set(x_8, 1, x_174); -x_175 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_175, 0, x_8); -x_17 = x_175; -x_18 = x_173; -goto block_23; +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_dec(x_15); +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_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; uint8_t x_187; lean_object* x_188; -lean_dec(x_153); -x_176 = lean_ctor_get(x_170, 1); -lean_inc(x_176); -lean_dec(x_170); -x_177 = lean_box(0); -lean_inc(x_155); -x_178 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_152, x_155, x_177); -x_179 = lean_box(0); -lean_inc(x_2); -x_180 = lean_array_mk(x_2); -x_181 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; -x_182 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_182, 0, x_181); -lean_ctor_set(x_182, 1, x_179); -lean_ctor_set(x_182, 2, x_180); -x_183 = l_Lean_CollectFVars_main(x_158, x_182); -x_184 = lean_ctor_get(x_183, 2); -lean_inc(x_184); -lean_dec(x_183); -x_185 = lean_array_get_size(x_184); -x_186 = lean_unsigned_to_nat(0u); -x_187 = lean_nat_dec_lt(x_186, x_185); -x_188 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_154, x_155, x_177); -if (x_187 == 0) +uint8_t x_31; +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_31 = !lean_is_exclusive(x_12); +if (x_31 == 0) { -uint8_t x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; -lean_dec(x_185); -lean_dec(x_184); -x_189 = 1; -x_190 = lean_box(x_189); -x_191 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_191, 0, x_190); -lean_ctor_set(x_191, 1, x_188); -lean_ctor_set(x_8, 1, x_191); -lean_ctor_set(x_8, 0, x_178); -x_192 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_192, 0, x_8); -x_17 = x_192; -x_18 = x_176; -goto block_23; +return x_12; } else { -uint8_t x_193; -x_193 = lean_nat_dec_le(x_185, x_185); -if (x_193 == 0) -{ -uint8_t x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; -lean_dec(x_185); -lean_dec(x_184); -x_194 = 1; -x_195 = lean_box(x_194); -x_196 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_196, 0, x_195); -lean_ctor_set(x_196, 1, x_188); -lean_ctor_set(x_8, 1, x_196); -lean_ctor_set(x_8, 0, x_178); -x_197 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_197, 0, x_8); -x_17 = x_197; -x_18 = x_176; -goto block_23; +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_12, 0); +x_33 = lean_ctor_get(x_12, 1); +lean_inc(x_33); +lean_inc(x_32); +lean_dec(x_12); +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 { -size_t x_198; size_t x_199; lean_object* x_200; uint8_t x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; -x_198 = 0; -x_199 = lean_usize_of_nat(x_185); -lean_dec(x_185); -x_200 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__10___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__11(x_3, x_184, x_198, x_199, x_178); -lean_dec(x_184); -x_201 = 1; -x_202 = lean_box(x_201); -x_203 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_203, 0, x_202); -lean_ctor_set(x_203, 1, x_188); -lean_ctor_set(x_8, 1, x_203); -lean_ctor_set(x_8, 0, x_200); -x_204 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_204, 0, x_8); -x_17 = x_204; -x_18 = x_176; -goto block_23; +lean_object* x_35; +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_4); +lean_ctor_set(x_35, 1, x_9); +return x_35; } } } +LEAN_EXPORT lean_object* l_Array_filterMapM___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getPropTypes___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: +{ +uint8_t x_9; +x_9 = lean_nat_dec_lt(x_2, x_3); +if (x_9 == 0) +{ +lean_object* x_10; lean_object* x_11; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_10 = l_Lean_Meta_Grind_NormalizePattern_main___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_8); +return x_11; } else { -lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; -lean_dec(x_158); -lean_dec(x_155); -lean_dec(x_154); -lean_dec(x_153); -lean_free_object(x_8); -lean_dec(x_152); +lean_object* x_12; uint8_t x_13; +x_12 = lean_array_get_size(x_1); +x_13 = lean_nat_dec_le(x_3, x_12); lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_3); -lean_dec(x_2); -x_205 = lean_ctor_get(x_170, 0); -lean_inc(x_205); -x_206 = lean_ctor_get(x_170, 1); -lean_inc(x_206); -if (lean_is_exclusive(x_170)) { - lean_ctor_release(x_170, 0); - lean_ctor_release(x_170, 1); - x_207 = x_170; -} else { - lean_dec_ref(x_170); - x_207 = lean_box(0); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_14 = l_Lean_Meta_Grind_NormalizePattern_main___closed__1; +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_14); +lean_ctor_set(x_15, 1, x_8); +return x_15; } -if (lean_is_scalar(x_207)) { - x_208 = lean_alloc_ctor(1, 2, 0); -} else { - x_208 = x_207; +else +{ +size_t x_16; size_t x_17; lean_object* x_18; lean_object* x_19; +x_16 = lean_usize_of_nat(x_2); +x_17 = lean_usize_of_nat(x_3); +x_18 = l_Lean_Meta_Grind_NormalizePattern_main___closed__1; +x_19 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getPropTypes___spec__2(x_1, x_16, x_17, x_18, x_4, x_5, x_6, x_7, x_8); +return x_19; } -lean_ctor_set(x_208, 0, x_205); -lean_ctor_set(x_208, 1, x_206); -return x_208; } } -else +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getPropTypes(lean_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_209; lean_object* x_210; -lean_dec(x_169); -lean_dec(x_158); -lean_dec(x_155); -x_209 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_209, 0, x_153); -lean_ctor_set(x_209, 1, x_154); -lean_ctor_set(x_8, 1, x_209); -x_210 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_210, 0, x_8); -x_17 = x_210; -x_18 = x_167; -goto block_23; +lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_7 = lean_array_get_size(x_1); +x_8 = lean_unsigned_to_nat(0u); +x_9 = l_Array_filterMapM___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getPropTypes___spec__1(x_1, x_8, x_7, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_7); +return x_9; } } -else +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getPropTypes___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: { -lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; -lean_dec(x_158); -lean_dec(x_155); -lean_dec(x_154); -lean_dec(x_153); -lean_free_object(x_8); -lean_dec(x_152); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); +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_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getPropTypes___spec__2(x_1, x_10, x_11, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_1); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Array_filterMapM___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getPropTypes___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_Array_filterMapM___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getPropTypes___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); lean_dec(x_3); lean_dec(x_2); -x_211 = lean_ctor_get(x_165, 0); -lean_inc(x_211); -x_212 = lean_ctor_get(x_165, 1); -lean_inc(x_212); -if (lean_is_exclusive(x_165)) { - lean_ctor_release(x_165, 0); - lean_ctor_release(x_165, 1); - x_213 = x_165; -} else { - lean_dec_ref(x_165); - x_213 = lean_box(0); +lean_dec(x_1); +return x_9; } -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; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getPropTypes___boxed(lean_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_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getPropTypes(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_1); +return x_7; } } -else +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isPatternFnCandidate(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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_215; uint8_t x_216; -x_215 = lean_ctor_get(x_161, 1); -lean_inc(x_215); -lean_dec(x_161); -x_216 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars(x_3, x_152, x_158); -if (x_216 == 0) +switch (lean_obj_tag(x_1)) { +case 1: { -lean_object* x_217; lean_object* x_218; -lean_dec(x_155); -x_217 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_217, 0, x_153); -lean_ctor_set(x_217, 1, x_154); -lean_ctor_set(x_8, 1, x_217); -x_218 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_218, 0, x_8); -x_17 = x_218; -x_18 = x_215; -goto block_23; +lean_object* x_10; uint8_t x_11; +x_10 = lean_ctor_get(x_2, 1); +x_11 = l_Array_contains___at_Lean_Meta_arrowDomainsN___spec__2(x_10, x_1); +if (x_11 == 0) +{ +uint8_t x_12; lean_object* x_13; lean_object* x_14; +x_12 = 1; +x_13 = lean_box(x_12); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_9); +return x_14; } else { -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; -lean_dec(x_153); -x_219 = lean_box(0); -lean_inc(x_155); -x_220 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_152, x_155, x_219); -x_221 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_154, x_155, x_219); -x_222 = 1; -x_223 = lean_box(x_222); -x_224 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_224, 0, x_223); -lean_ctor_set(x_224, 1, x_221); -lean_ctor_set(x_8, 1, x_224); -lean_ctor_set(x_8, 0, x_220); -x_225 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_225, 0, x_8); -x_17 = x_225; -x_18 = x_215; -goto block_23; +uint8_t x_15; lean_object* x_16; lean_object* x_17; +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_9); +return x_17; } } +case 4: +{ +lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_18 = lean_ctor_get(x_1, 0); +x_19 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames; +x_20 = l_Array_contains___at_Lean_registerInternalExceptionId___spec__1(x_19, x_18); +if (x_20 == 0) +{ +uint8_t x_21; lean_object* x_22; lean_object* x_23; +x_21 = 1; +x_22 = lean_box(x_21); +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_9); +return x_23; } else { -lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; -lean_dec(x_158); -lean_dec(x_155); -lean_dec(x_154); -lean_dec(x_153); -lean_free_object(x_8); -lean_dec(x_152); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_3); -lean_dec(x_2); -x_226 = lean_ctor_get(x_161, 0); -lean_inc(x_226); -x_227 = lean_ctor_get(x_161, 1); -lean_inc(x_227); -if (lean_is_exclusive(x_161)) { - lean_ctor_release(x_161, 0); - lean_ctor_release(x_161, 1); - x_228 = x_161; -} else { - lean_dec_ref(x_161); - x_228 = lean_box(0); +uint8_t x_24; lean_object* x_25; lean_object* x_26; +x_24 = 0; +x_25 = lean_box(x_24); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set(x_26, 1, x_9); +return x_26; +} +} +default: +{ +uint8_t x_27; lean_object* x_28; lean_object* x_29; +x_27 = 0; +x_28 = lean_box(x_27); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_9); +return x_29; +} } -if (lean_is_scalar(x_228)) { - x_229 = lean_alloc_ctor(1, 2, 0); -} else { - x_229 = x_228; } -lean_ctor_set(x_229, 0, x_226); -lean_ctor_set(x_229, 1, x_227); -return x_229; +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isPatternFnCandidate___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isPatternFnCandidate(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); +lean_dec(x_1); +return x_10; } } +LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___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; +x_10 = lean_ctor_get(x_7, 2); +x_11 = l_Lean_isTracingEnabledForCore(x_1, x_10, x_9); +x_12 = !lean_is_exclusive(x_11); +if (x_12 == 0) +{ +return x_11; +} 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; uint8_t x_239; lean_object* x_240; lean_object* x_241; -lean_dec(x_153); -if (lean_is_exclusive(x_160)) { - lean_ctor_release(x_160, 0); - x_230 = x_160; -} else { - lean_dec_ref(x_160); - x_230 = lean_box(0); +lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_13 = lean_ctor_get(x_11, 0); +x_14 = lean_ctor_get(x_11, 1); +lean_inc(x_14); +lean_inc(x_13); +lean_dec(x_11); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_13); +lean_ctor_set(x_15, 1, x_14); +return x_15; } -x_231 = lean_box(0); -lean_inc(x_2); -x_232 = lean_array_mk(x_2); -x_233 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; -x_234 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_234, 0, x_233); -lean_ctor_set(x_234, 1, x_231); -lean_ctor_set(x_234, 2, x_232); -x_235 = l_Lean_CollectFVars_main(x_158, x_234); -x_236 = lean_ctor_get(x_235, 2); -lean_inc(x_236); -lean_dec(x_235); -x_237 = lean_array_get_size(x_236); -x_238 = lean_unsigned_to_nat(0u); -x_239 = lean_nat_dec_lt(x_238, x_237); -x_240 = lean_box(0); -x_241 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_154, x_155, x_240); -if (x_239 == 0) +} +} +static double _init_l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___spec__2___closed__1() { +_start: { -uint8_t x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; -lean_dec(x_237); -lean_dec(x_236); -x_242 = 1; -x_243 = lean_box(x_242); -x_244 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_244, 0, x_243); -lean_ctor_set(x_244, 1, x_241); -lean_ctor_set(x_8, 1, x_244); -if (lean_is_scalar(x_230)) { - x_245 = lean_alloc_ctor(1, 1, 0); -} else { - x_245 = x_230; +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_ctor_set(x_245, 0, x_8); -x_17 = x_245; -x_18 = x_159; -goto block_23; } -else +LEAN_EXPORT lean_object* l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: { -uint8_t x_246; -x_246 = lean_nat_dec_le(x_237, x_237); -if (x_246 == 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; uint8_t x_18; +x_11 = lean_ctor_get(x_8, 5); +x_12 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_2, x_6, x_7, x_8, x_9, x_10); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_12, 1); +lean_inc(x_14); +lean_dec(x_12); +x_15 = lean_st_ref_take(x_9, x_14); +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_16, 3); +lean_inc(x_17); +x_18 = !lean_is_exclusive(x_15); +if (x_18 == 0) { -uint8_t x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; -lean_dec(x_237); -lean_dec(x_236); -x_247 = 1; -x_248 = lean_box(x_247); -x_249 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_249, 0, x_248); -lean_ctor_set(x_249, 1, x_241); -lean_ctor_set(x_8, 1, x_249); -if (lean_is_scalar(x_230)) { - x_250 = lean_alloc_ctor(1, 1, 0); -} else { - x_250 = x_230; +lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_19 = lean_ctor_get(x_15, 1); +x_20 = lean_ctor_get(x_15, 0); +lean_dec(x_20); +x_21 = !lean_is_exclusive(x_16); +if (x_21 == 0) +{ +lean_object* x_22; uint8_t x_23; +x_22 = lean_ctor_get(x_16, 3); +lean_dec(x_22); +x_23 = !lean_is_exclusive(x_17); +if (x_23 == 0) +{ +lean_object* x_24; double 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; +x_24 = lean_ctor_get(x_17, 0); +x_25 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___spec__2___closed__1; +x_26 = 0; +x_27 = l_Lean_Meta_Grind_ppPattern___closed__3; +x_28 = lean_alloc_ctor(0, 2, 17); +lean_ctor_set(x_28, 0, x_1); +lean_ctor_set(x_28, 1, x_27); +lean_ctor_set_float(x_28, sizeof(void*)*2, x_25); +lean_ctor_set_float(x_28, sizeof(void*)*2 + 8, x_25); +lean_ctor_set_uint8(x_28, sizeof(void*)*2 + 16, x_26); +x_29 = l_Lean_Meta_Grind_NormalizePattern_main___closed__1; +x_30 = lean_alloc_ctor(9, 3, 0); +lean_ctor_set(x_30, 0, x_28); +lean_ctor_set(x_30, 1, x_13); +lean_ctor_set(x_30, 2, x_29); +lean_inc(x_11); +lean_ctor_set(x_15, 1, x_30); +lean_ctor_set(x_15, 0, x_11); +x_31 = l_Lean_PersistentArray_push___rarg(x_24, x_15); +lean_ctor_set(x_17, 0, x_31); +x_32 = lean_st_ref_set(x_9, x_16, x_19); +x_33 = !lean_is_exclusive(x_32); +if (x_33 == 0) +{ +lean_object* x_34; lean_object* x_35; +x_34 = lean_ctor_get(x_32, 0); +lean_dec(x_34); +x_35 = lean_box(0); +lean_ctor_set(x_32, 0, x_35); +return x_32; +} +else +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_32, 1); +lean_inc(x_36); +lean_dec(x_32); +x_37 = lean_box(0); +x_38 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_38, 0, x_37); +lean_ctor_set(x_38, 1, x_36); +return x_38; } -lean_ctor_set(x_250, 0, x_8); -x_17 = x_250; -x_18 = x_159; -goto block_23; } else { -size_t x_251; size_t x_252; lean_object* x_253; uint8_t x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; -x_251 = 0; -x_252 = lean_usize_of_nat(x_237); -lean_dec(x_237); -x_253 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__12___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__13(x_3, x_236, x_251, x_252, x_152); -lean_dec(x_236); -x_254 = 1; -x_255 = lean_box(x_254); -x_256 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_256, 0, x_255); -lean_ctor_set(x_256, 1, x_241); -lean_ctor_set(x_8, 1, x_256); -lean_ctor_set(x_8, 0, x_253); -if (lean_is_scalar(x_230)) { - x_257 = lean_alloc_ctor(1, 1, 0); +uint64_t x_39; lean_object* x_40; double 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; +x_39 = lean_ctor_get_uint64(x_17, sizeof(void*)*1); +x_40 = lean_ctor_get(x_17, 0); +lean_inc(x_40); +lean_dec(x_17); +x_41 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___spec__2___closed__1; +x_42 = 0; +x_43 = l_Lean_Meta_Grind_ppPattern___closed__3; +x_44 = lean_alloc_ctor(0, 2, 17); +lean_ctor_set(x_44, 0, x_1); +lean_ctor_set(x_44, 1, x_43); +lean_ctor_set_float(x_44, sizeof(void*)*2, x_41); +lean_ctor_set_float(x_44, sizeof(void*)*2 + 8, x_41); +lean_ctor_set_uint8(x_44, sizeof(void*)*2 + 16, x_42); +x_45 = l_Lean_Meta_Grind_NormalizePattern_main___closed__1; +x_46 = lean_alloc_ctor(9, 3, 0); +lean_ctor_set(x_46, 0, x_44); +lean_ctor_set(x_46, 1, x_13); +lean_ctor_set(x_46, 2, x_45); +lean_inc(x_11); +lean_ctor_set(x_15, 1, x_46); +lean_ctor_set(x_15, 0, x_11); +x_47 = l_Lean_PersistentArray_push___rarg(x_40, x_15); +x_48 = lean_alloc_ctor(0, 1, 8); +lean_ctor_set(x_48, 0, x_47); +lean_ctor_set_uint64(x_48, sizeof(void*)*1, x_39); +lean_ctor_set(x_16, 3, x_48); +x_49 = lean_st_ref_set(x_9, x_16, x_19); +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 { - x_257 = x_230; -} -lean_ctor_set(x_257, 0, x_8); -x_17 = x_257; -x_18 = x_159; -goto block_23; + lean_dec_ref(x_49); + x_51 = lean_box(0); } +x_52 = lean_box(0); +if (lean_is_scalar(x_51)) { + x_53 = lean_alloc_ctor(0, 2, 0); +} else { + x_53 = x_51; } +lean_ctor_set(x_53, 0, x_52); +lean_ctor_set(x_53, 1, x_50); +return x_53; } } else { -lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; -lean_dec(x_155); -lean_dec(x_154); -lean_dec(x_153); -lean_free_object(x_8); -lean_dec(x_152); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_3); -lean_dec(x_2); -x_258 = lean_ctor_get(x_157, 0); -lean_inc(x_258); -x_259 = lean_ctor_get(x_157, 1); -lean_inc(x_259); -if (lean_is_exclusive(x_157)) { - lean_ctor_release(x_157, 0); - lean_ctor_release(x_157, 1); - x_260 = x_157; +lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; uint64_t x_61; lean_object* x_62; lean_object* x_63; double x_64; uint8_t 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_54 = lean_ctor_get(x_16, 0); +x_55 = lean_ctor_get(x_16, 1); +x_56 = lean_ctor_get(x_16, 2); +x_57 = lean_ctor_get(x_16, 4); +x_58 = lean_ctor_get(x_16, 5); +x_59 = lean_ctor_get(x_16, 6); +x_60 = lean_ctor_get(x_16, 7); +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_dec(x_16); +x_61 = lean_ctor_get_uint64(x_17, sizeof(void*)*1); +x_62 = lean_ctor_get(x_17, 0); +lean_inc(x_62); +if (lean_is_exclusive(x_17)) { + lean_ctor_release(x_17, 0); + x_63 = x_17; } else { - lean_dec_ref(x_157); - x_260 = lean_box(0); + lean_dec_ref(x_17); + x_63 = lean_box(0); } -if (lean_is_scalar(x_260)) { - x_261 = lean_alloc_ctor(1, 2, 0); +x_64 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___spec__2___closed__1; +x_65 = 0; +x_66 = l_Lean_Meta_Grind_ppPattern___closed__3; +x_67 = lean_alloc_ctor(0, 2, 17); +lean_ctor_set(x_67, 0, x_1); +lean_ctor_set(x_67, 1, x_66); +lean_ctor_set_float(x_67, sizeof(void*)*2, x_64); +lean_ctor_set_float(x_67, sizeof(void*)*2 + 8, x_64); +lean_ctor_set_uint8(x_67, sizeof(void*)*2 + 16, x_65); +x_68 = l_Lean_Meta_Grind_NormalizePattern_main___closed__1; +x_69 = lean_alloc_ctor(9, 3, 0); +lean_ctor_set(x_69, 0, x_67); +lean_ctor_set(x_69, 1, x_13); +lean_ctor_set(x_69, 2, x_68); +lean_inc(x_11); +lean_ctor_set(x_15, 1, x_69); +lean_ctor_set(x_15, 0, x_11); +x_70 = l_Lean_PersistentArray_push___rarg(x_62, x_15); +if (lean_is_scalar(x_63)) { + x_71 = lean_alloc_ctor(0, 1, 8); } else { - x_261 = x_260; + x_71 = x_63; } -lean_ctor_set(x_261, 0, x_258); -lean_ctor_set(x_261, 1, x_259); -return x_261; +lean_ctor_set(x_71, 0, x_70); +lean_ctor_set_uint64(x_71, sizeof(void*)*1, x_61); +x_72 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_72, 0, x_54); +lean_ctor_set(x_72, 1, x_55); +lean_ctor_set(x_72, 2, x_56); +lean_ctor_set(x_72, 3, x_71); +lean_ctor_set(x_72, 4, x_57); +lean_ctor_set(x_72, 5, x_58); +lean_ctor_set(x_72, 6, x_59); +lean_ctor_set(x_72, 7, x_60); +x_73 = lean_st_ref_set(x_9, x_72, x_19); +x_74 = lean_ctor_get(x_73, 1); +lean_inc(x_74); +if (lean_is_exclusive(x_73)) { + lean_ctor_release(x_73, 0); + lean_ctor_release(x_73, 1); + x_75 = x_73; +} else { + lean_dec_ref(x_73); + x_75 = lean_box(0); +} +x_76 = lean_box(0); +if (lean_is_scalar(x_75)) { + x_77 = lean_alloc_ctor(0, 2, 0); +} else { + x_77 = x_75; +} +lean_ctor_set(x_77, 0, x_76); +lean_ctor_set(x_77, 1, x_74); +return x_77; } } else { -lean_object* x_262; lean_object* x_263; lean_object* x_264; -lean_dec(x_155); -lean_dec(x_16); -if (lean_is_exclusive(x_156)) { - lean_ctor_release(x_156, 0); - x_262 = x_156; +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; uint64_t x_87; lean_object* x_88; lean_object* x_89; double 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; +x_78 = lean_ctor_get(x_15, 1); +lean_inc(x_78); +lean_dec(x_15); +x_79 = lean_ctor_get(x_16, 0); +lean_inc(x_79); +x_80 = lean_ctor_get(x_16, 1); +lean_inc(x_80); +x_81 = lean_ctor_get(x_16, 2); +lean_inc(x_81); +x_82 = lean_ctor_get(x_16, 4); +lean_inc(x_82); +x_83 = lean_ctor_get(x_16, 5); +lean_inc(x_83); +x_84 = lean_ctor_get(x_16, 6); +lean_inc(x_84); +x_85 = lean_ctor_get(x_16, 7); +lean_inc(x_85); +if (lean_is_exclusive(x_16)) { + lean_ctor_release(x_16, 0); + lean_ctor_release(x_16, 1); + lean_ctor_release(x_16, 2); + lean_ctor_release(x_16, 3); + lean_ctor_release(x_16, 4); + lean_ctor_release(x_16, 5); + lean_ctor_release(x_16, 6); + lean_ctor_release(x_16, 7); + x_86 = x_16; } else { - lean_dec_ref(x_156); - x_262 = lean_box(0); + lean_dec_ref(x_16); + x_86 = lean_box(0); } -x_263 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_263, 0, x_153); -lean_ctor_set(x_263, 1, x_154); -lean_ctor_set(x_8, 1, x_263); -if (lean_is_scalar(x_262)) { - x_264 = lean_alloc_ctor(1, 1, 0); +x_87 = lean_ctor_get_uint64(x_17, sizeof(void*)*1); +x_88 = lean_ctor_get(x_17, 0); +lean_inc(x_88); +if (lean_is_exclusive(x_17)) { + lean_ctor_release(x_17, 0); + x_89 = x_17; +} else { + lean_dec_ref(x_17); + x_89 = lean_box(0); +} +x_90 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___spec__2___closed__1; +x_91 = 0; +x_92 = l_Lean_Meta_Grind_ppPattern___closed__3; +x_93 = lean_alloc_ctor(0, 2, 17); +lean_ctor_set(x_93, 0, x_1); +lean_ctor_set(x_93, 1, x_92); +lean_ctor_set_float(x_93, sizeof(void*)*2, x_90); +lean_ctor_set_float(x_93, sizeof(void*)*2 + 8, x_90); +lean_ctor_set_uint8(x_93, sizeof(void*)*2 + 16, x_91); +x_94 = l_Lean_Meta_Grind_NormalizePattern_main___closed__1; +x_95 = lean_alloc_ctor(9, 3, 0); +lean_ctor_set(x_95, 0, x_93); +lean_ctor_set(x_95, 1, x_13); +lean_ctor_set(x_95, 2, x_94); +lean_inc(x_11); +x_96 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_96, 0, x_11); +lean_ctor_set(x_96, 1, x_95); +x_97 = l_Lean_PersistentArray_push___rarg(x_88, x_96); +if (lean_is_scalar(x_89)) { + x_98 = lean_alloc_ctor(0, 1, 8); } else { - x_264 = x_262; + x_98 = x_89; } -lean_ctor_set(x_264, 0, x_8); -x_17 = x_264; -x_18 = x_13; -goto block_23; +lean_ctor_set(x_98, 0, x_97); +lean_ctor_set_uint64(x_98, sizeof(void*)*1, x_87); +if (lean_is_scalar(x_86)) { + x_99 = lean_alloc_ctor(0, 8, 0); +} else { + x_99 = x_86; +} +lean_ctor_set(x_99, 0, x_79); +lean_ctor_set(x_99, 1, x_80); +lean_ctor_set(x_99, 2, x_81); +lean_ctor_set(x_99, 3, x_98); +lean_ctor_set(x_99, 4, x_82); +lean_ctor_set(x_99, 5, x_83); +lean_ctor_set(x_99, 6, x_84); +lean_ctor_set(x_99, 7, x_85); +x_100 = lean_st_ref_set(x_9, x_99, x_78); +x_101 = lean_ctor_get(x_100, 1); +lean_inc(x_101); +if (lean_is_exclusive(x_100)) { + lean_ctor_release(x_100, 0); + lean_ctor_release(x_100, 1); + x_102 = x_100; +} else { + lean_dec_ref(x_100); + x_102 = lean_box(0); } +x_103 = lean_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 -{ -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_265 = lean_ctor_get(x_8, 1); -x_266 = lean_ctor_get(x_8, 0); -lean_inc(x_265); -lean_inc(x_266); -lean_dec(x_8); -x_267 = lean_ctor_get(x_265, 0); -lean_inc(x_267); -x_268 = lean_ctor_get(x_265, 1); -lean_inc(x_268); -if (lean_is_exclusive(x_265)) { - lean_ctor_release(x_265, 0); - lean_ctor_release(x_265, 1); - x_269 = x_265; -} else { - lean_dec_ref(x_265); - x_269 = lean_box(0); } -x_270 = l_Lean_Expr_fvarId_x21(x_16); -x_271 = l_Lean_RBNode_findCore___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__2(x_268, x_270); -if (lean_obj_tag(x_271) == 0) -{ -lean_object* x_272; -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -x_272 = lean_infer_type(x_16, x_9, x_10, x_11, x_12, x_13); -if (lean_obj_tag(x_272) == 0) +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___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: { -lean_object* x_273; lean_object* x_274; lean_object* x_275; -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_275 = l_Lean_RBNode_findCore___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__2(x_266, x_270); -if (lean_obj_tag(x_275) == 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_st_ref_take(x_5, x_11); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_12, 1); +lean_inc(x_14); +lean_dec(x_12); +x_15 = lean_ctor_get(x_13, 0); +lean_inc(x_15); +lean_dec(x_13); +x_16 = lean_array_push(x_15, x_1); +x_17 = lean_alloc_ctor(0, 1, 1); +lean_ctor_set(x_17, 0, x_16); +lean_ctor_set_uint8(x_17, sizeof(void*)*1, x_2); +x_18 = lean_st_ref_set(x_5, x_17, x_14); +x_19 = !lean_is_exclusive(x_18); +if (x_19 == 0) { -lean_object* x_276; -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_273); -x_276 = l_Lean_Meta_isProp(x_273, x_9, x_10, x_11, x_12, x_274); -if (lean_obj_tag(x_276) == 0) +lean_object* x_20; lean_object* x_21; +x_20 = lean_ctor_get(x_18, 0); +lean_dec(x_20); +x_21 = lean_box(0); +lean_ctor_set(x_18, 0, x_21); +return x_18; +} +else { -lean_object* x_277; uint8_t x_278; -x_277 = lean_ctor_get(x_276, 0); -lean_inc(x_277); -x_278 = lean_unbox(x_277); -lean_dec(x_277); -if (x_278 == 0) +lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_22 = lean_ctor_get(x_18, 1); +lean_inc(x_22); +lean_dec(x_18); +x_23 = lean_box(0); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set(x_24, 1, x_22); +return x_24; +} +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___lambda__2___closed__1() { +_start: { -lean_object* x_279; lean_object* x_280; -x_279 = lean_ctor_get(x_276, 1); -lean_inc(x_279); -lean_dec(x_276); -lean_inc(x_9); -lean_inc(x_270); -x_280 = l_Lean_FVarId_getDecl(x_270, x_9, x_10, x_11, x_12, x_279); -if (lean_obj_tag(x_280) == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("found full coverage", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___lambda__2___closed__2() { +_start: { -lean_object* x_281; lean_object* x_282; uint8_t x_283; lean_object* x_284; -x_281 = lean_ctor_get(x_280, 0); -lean_inc(x_281); -x_282 = lean_ctor_get(x_280, 1); -lean_inc(x_282); -lean_dec(x_280); -x_283 = l_Lean_LocalDecl_binderInfo(x_281); -lean_dec(x_281); -x_284 = lean_box(x_283); -if (lean_obj_tag(x_284) == 3) +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___lambda__2___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___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_285; -lean_inc(x_12); -lean_inc(x_11); +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_12 = lean_st_ref_get(x_6, x_11); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_12, 1); +lean_inc(x_14); +lean_dec(x_12); +x_15 = lean_ctor_get(x_13, 2); +lean_inc(x_15); +lean_dec(x_13); +x_16 = lean_ctor_get(x_4, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_4, 1); +lean_inc(x_17); +x_18 = lean_array_get_size(x_17); +lean_dec(x_17); lean_inc(x_10); lean_inc(x_9); -lean_inc(x_273); -lean_inc(x_266); -lean_inc(x_3); -x_285 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized(x_3, x_266, x_273, x_9, x_10, x_11, x_12, x_282); -if (lean_obj_tag(x_285) == 0) +lean_inc(x_8); +lean_inc(x_7); +x_19 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage(x_16, x_18, x_15, x_7, x_8, x_9, x_10, x_14); +if (lean_obj_tag(x_19) == 0) { -lean_object* x_286; uint8_t x_287; -x_286 = lean_ctor_get(x_285, 0); -lean_inc(x_286); -x_287 = lean_unbox(x_286); -lean_dec(x_286); -if (x_287 == 0) +lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_20 = lean_ctor_get(x_19, 0); +lean_inc(x_20); +x_21 = lean_ctor_get(x_19, 1); +lean_inc(x_21); +lean_dec(x_19); +if (lean_obj_tag(x_20) == 0) { -lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; -lean_dec(x_273); -lean_dec(x_270); -x_288 = lean_ctor_get(x_285, 1); -lean_inc(x_288); -lean_dec(x_285); -if (lean_is_scalar(x_269)) { - x_289 = lean_alloc_ctor(0, 2, 0); -} else { - x_289 = x_269; -} -lean_ctor_set(x_289, 0, x_267); -lean_ctor_set(x_289, 1, x_268); -x_290 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_290, 0, x_266); -lean_ctor_set(x_290, 1, x_289); -x_291 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_291, 0, x_290); -x_17 = x_291; -x_18 = x_288; -goto block_23; +uint8_t x_38; +x_38 = 1; +x_22 = x_38; +goto block_37; } else { -lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; uint8_t x_303; lean_object* x_304; -lean_dec(x_267); -x_292 = lean_ctor_get(x_285, 1); -lean_inc(x_292); -lean_dec(x_285); -x_293 = lean_box(0); -lean_inc(x_270); -x_294 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_266, x_270, x_293); -x_295 = lean_box(0); -lean_inc(x_2); -x_296 = lean_array_mk(x_2); -x_297 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; -x_298 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_298, 0, x_297); -lean_ctor_set(x_298, 1, x_295); -lean_ctor_set(x_298, 2, x_296); -x_299 = l_Lean_CollectFVars_main(x_273, x_298); -x_300 = lean_ctor_get(x_299, 2); -lean_inc(x_300); -lean_dec(x_299); -x_301 = lean_array_get_size(x_300); -x_302 = lean_unsigned_to_nat(0u); -x_303 = lean_nat_dec_lt(x_302, x_301); -x_304 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_268, x_270, x_293); -if (x_303 == 0) -{ -uint8_t x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; -lean_dec(x_301); -lean_dec(x_300); -x_305 = 1; -x_306 = lean_box(x_305); -if (lean_is_scalar(x_269)) { - x_307 = lean_alloc_ctor(0, 2, 0); -} else { - x_307 = x_269; -} -lean_ctor_set(x_307, 0, x_306); -lean_ctor_set(x_307, 1, x_304); -x_308 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_308, 0, x_294); -lean_ctor_set(x_308, 1, x_307); -x_309 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_309, 0, x_308); -x_17 = x_309; -x_18 = x_292; -goto block_23; +uint8_t x_39; +lean_dec(x_20); +x_39 = 0; +x_22 = x_39; +goto block_37; } -else -{ -uint8_t x_310; -x_310 = lean_nat_dec_le(x_301, x_301); -if (x_310 == 0) +block_37: { -uint8_t x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; -lean_dec(x_301); -lean_dec(x_300); -x_311 = 1; -x_312 = lean_box(x_311); -if (lean_is_scalar(x_269)) { - x_313 = lean_alloc_ctor(0, 2, 0); -} else { - x_313 = x_269; -} -lean_ctor_set(x_313, 0, x_312); -lean_ctor_set(x_313, 1, x_304); -x_314 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_314, 0, x_294); -lean_ctor_set(x_314, 1, x_313); -x_315 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_315, 0, x_314); -x_17 = x_315; -x_18 = x_292; -goto block_23; +if (x_22 == 0) +{ +lean_object* x_23; lean_object* x_24; +lean_dec(x_2); +x_23 = lean_box(0); +x_24 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___lambda__1(x_1, x_22, x_23, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_21); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_4); +return x_24; } else { -size_t x_316; size_t x_317; lean_object* x_318; uint8_t x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; -x_316 = 0; -x_317 = lean_usize_of_nat(x_301); -lean_dec(x_301); -x_318 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__10___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__11(x_3, x_300, x_316, x_317, x_294); -lean_dec(x_300); -x_319 = 1; -x_320 = lean_box(x_319); -if (lean_is_scalar(x_269)) { - x_321 = lean_alloc_ctor(0, 2, 0); -} else { - x_321 = x_269; +lean_object* x_25; lean_object* x_26; uint8_t x_27; +lean_inc(x_2); +x_25 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___spec__1(x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_21); +x_26 = lean_ctor_get(x_25, 0); +lean_inc(x_26); +x_27 = lean_unbox(x_26); +lean_dec(x_26); +if (x_27 == 0) +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; +lean_dec(x_2); +x_28 = lean_ctor_get(x_25, 1); +lean_inc(x_28); +lean_dec(x_25); +x_29 = lean_box(0); +x_30 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___lambda__1(x_1, x_22, x_29, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_28); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_4); +return x_30; } -lean_ctor_set(x_321, 0, x_320); -lean_ctor_set(x_321, 1, x_304); -x_322 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_322, 0, x_318); -lean_ctor_set(x_322, 1, x_321); -x_323 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_323, 0, x_322); -x_17 = x_323; -x_18 = x_292; -goto block_23; +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_31 = lean_ctor_get(x_25, 1); +lean_inc(x_31); +lean_dec(x_25); +x_32 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___lambda__2___closed__2; +x_33 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___spec__2(x_2, x_32, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_31); +x_34 = lean_ctor_get(x_33, 0); +lean_inc(x_34); +x_35 = lean_ctor_get(x_33, 1); +lean_inc(x_35); +lean_dec(x_33); +x_36 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___lambda__1(x_1, x_22, x_34, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_35); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_4); +lean_dec(x_34); +return x_36; } } } } else { -lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; -lean_dec(x_273); -lean_dec(x_270); -lean_dec(x_269); -lean_dec(x_268); -lean_dec(x_267); -lean_dec(x_266); -lean_dec(x_12); -lean_dec(x_11); +uint8_t x_40; lean_dec(x_10); lean_dec(x_9); -lean_dec(x_3); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_4); lean_dec(x_2); -x_324 = lean_ctor_get(x_285, 0); -lean_inc(x_324); -x_325 = lean_ctor_get(x_285, 1); -lean_inc(x_325); -if (lean_is_exclusive(x_285)) { - lean_ctor_release(x_285, 0); - lean_ctor_release(x_285, 1); - x_326 = x_285; -} else { - lean_dec_ref(x_285); - x_326 = lean_box(0); +lean_dec(x_1); +x_40 = !lean_is_exclusive(x_19); +if (x_40 == 0) +{ +return x_19; } -if (lean_is_scalar(x_326)) { - x_327 = lean_alloc_ctor(1, 2, 0); -} else { - x_327 = x_326; +else +{ +lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_41 = lean_ctor_get(x_19, 0); +x_42 = lean_ctor_get(x_19, 1); +lean_inc(x_42); +lean_inc(x_41); +lean_dec(x_19); +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_ctor_set(x_327, 0, x_324); -lean_ctor_set(x_327, 1, x_325); -return x_327; } } -else +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___closed__1() { +_start: { -lean_object* x_328; lean_object* x_329; lean_object* x_330; -lean_dec(x_284); -lean_dec(x_273); -lean_dec(x_270); -if (lean_is_scalar(x_269)) { - x_328 = lean_alloc_ctor(0, 2, 0); -} else { - x_328 = x_269; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("search", 6, 6); +return x_1; } -lean_ctor_set(x_328, 0, x_267); -lean_ctor_set(x_328, 1, x_268); -x_329 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_329, 0, x_266); -lean_ctor_set(x_329, 1, x_328); -x_330 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_330, 0, x_329); -x_17 = x_330; -x_18 = x_282; -goto block_23; } +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___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_Grind_mkGroundPattern___closed__1; +x_2 = l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__1; +x_3 = l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__2; +x_4 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___closed__1; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; } -else +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___closed__3() { +_start: { -lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; -lean_dec(x_273); -lean_dec(x_270); -lean_dec(x_269); -lean_dec(x_268); -lean_dec(x_267); -lean_dec(x_266); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_3); -lean_dec(x_2); -x_331 = lean_ctor_get(x_280, 0); -lean_inc(x_331); -x_332 = lean_ctor_get(x_280, 1); -lean_inc(x_332); -if (lean_is_exclusive(x_280)) { - lean_ctor_release(x_280, 0); - lean_ctor_release(x_280, 1); - x_333 = x_280; -} else { - lean_dec_ref(x_280); - x_333 = lean_box(0); +lean_object* x_1; +x_1 = lean_mk_string_unchecked("found pattern: ", 15, 15); +return x_1; } -if (lean_is_scalar(x_333)) { - x_334 = lean_alloc_ctor(1, 2, 0); -} else { - x_334 = x_333; } -lean_ctor_set(x_334, 0, x_331); -lean_ctor_set(x_334, 1, x_332); -return x_334; +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___closed__3; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } } -else +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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_335; uint8_t x_336; -x_335 = lean_ctor_get(x_276, 1); -lean_inc(x_335); -lean_dec(x_276); -x_336 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars(x_3, x_266, x_273); -if (x_336 == 0) +lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_10 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___closed__2; +x_11 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___spec__1(x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +x_13 = lean_unbox(x_12); +lean_dec(x_12); +if (x_13 == 0) { -lean_object* x_337; lean_object* x_338; lean_object* x_339; -lean_dec(x_270); -if (lean_is_scalar(x_269)) { - x_337 = lean_alloc_ctor(0, 2, 0); -} else { - x_337 = x_269; -} -lean_ctor_set(x_337, 0, x_267); -lean_ctor_set(x_337, 1, x_268); -x_338 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_338, 0, x_266); -lean_ctor_set(x_338, 1, x_337); -x_339 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_339, 0, x_338); -x_17 = x_339; -x_18 = x_335; -goto block_23; +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___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___lambda__2(x_1, x_10, x_15, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_14); +return x_16; } else { -lean_object* x_340; lean_object* x_341; lean_object* x_342; uint8_t x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; -lean_dec(x_267); -x_340 = lean_box(0); -lean_inc(x_270); -x_341 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_266, x_270, x_340); -x_342 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_268, x_270, x_340); -x_343 = 1; -x_344 = lean_box(x_343); -if (lean_is_scalar(x_269)) { - x_345 = lean_alloc_ctor(0, 2, 0); -} else { - x_345 = x_269; +uint8_t x_17; +x_17 = !lean_is_exclusive(x_11); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_18 = lean_ctor_get(x_11, 1); +x_19 = lean_ctor_get(x_11, 0); +lean_dec(x_19); +lean_inc(x_1); +x_20 = l_Lean_Meta_Grind_ppPattern(x_1); +x_21 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___closed__4; +lean_ctor_set_tag(x_11, 7); +lean_ctor_set(x_11, 1, x_20); +lean_ctor_set(x_11, 0, x_21); +x_22 = l_Lean_Meta_Grind_ppPattern___closed__4; +x_23 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_23, 0, x_11); +lean_ctor_set(x_23, 1, x_22); +x_24 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___spec__2(x_10, x_23, x_2, x_3, x_4, x_5, x_6, x_7, x_8, 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 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___lambda__2(x_1, x_10, x_25, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_26); +lean_dec(x_25); +return x_27; } -lean_ctor_set(x_345, 0, x_344); -lean_ctor_set(x_345, 1, x_342); -x_346 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_346, 0, x_341); -lean_ctor_set(x_346, 1, x_345); -x_347 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_347, 0, x_346); -x_17 = x_347; -x_18 = x_335; -goto block_23; +else +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_28 = lean_ctor_get(x_11, 1); +lean_inc(x_28); +lean_dec(x_11); +lean_inc(x_1); +x_29 = l_Lean_Meta_Grind_ppPattern(x_1); +x_30 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___closed__4; +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_Meta_Grind_ppPattern___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_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___spec__2(x_10, x_33, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_28); +x_35 = lean_ctor_get(x_34, 0); +lean_inc(x_35); +x_36 = lean_ctor_get(x_34, 1); +lean_inc(x_36); +lean_dec(x_34); +x_37 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___lambda__2(x_1, x_10, x_35, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_36); +lean_dec(x_35); +return x_37; } } } -else +} +LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___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_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; -lean_dec(x_273); -lean_dec(x_270); -lean_dec(x_269); -lean_dec(x_268); -lean_dec(x_267); -lean_dec(x_266); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); +lean_object* x_10; +x_10 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___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); -x_348 = lean_ctor_get(x_276, 0); -lean_inc(x_348); -x_349 = lean_ctor_get(x_276, 1); -lean_inc(x_349); -if (lean_is_exclusive(x_276)) { - lean_ctor_release(x_276, 0); - lean_ctor_release(x_276, 1); - x_350 = x_276; -} else { - lean_dec_ref(x_276); - x_350 = lean_box(0); +return x_10; } -if (lean_is_scalar(x_350)) { - x_351 = lean_alloc_ctor(1, 2, 0); -} else { - x_351 = x_350; } -lean_ctor_set(x_351, 0, x_348); -lean_ctor_set(x_351, 1, x_349); -return x_351; +LEAN_EXPORT lean_object* l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_11; } } -else +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___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_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; uint8_t x_361; lean_object* x_362; lean_object* x_363; -lean_dec(x_267); -if (lean_is_exclusive(x_275)) { - lean_ctor_release(x_275, 0); - x_352 = x_275; -} else { - lean_dec_ref(x_275); - x_352 = lean_box(0); +uint8_t x_12; lean_object* x_13; +x_12 = lean_unbox(x_2); +lean_dec(x_2); +x_13 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___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_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_13; } -x_353 = lean_box(0); -lean_inc(x_2); -x_354 = lean_array_mk(x_2); -x_355 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; -x_356 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_356, 0, x_355); -lean_ctor_set(x_356, 1, x_353); -lean_ctor_set(x_356, 2, x_354); -x_357 = l_Lean_CollectFVars_main(x_273, x_356); -x_358 = lean_ctor_get(x_357, 2); -lean_inc(x_358); -lean_dec(x_357); -x_359 = lean_array_get_size(x_358); -x_360 = lean_unsigned_to_nat(0u); -x_361 = lean_nat_dec_lt(x_360, x_359); -x_362 = lean_box(0); -x_363 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_268, x_270, x_362); -if (x_361 == 0) +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___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_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; -lean_dec(x_359); -lean_dec(x_358); -x_364 = 1; -x_365 = lean_box(x_364); -if (lean_is_scalar(x_269)) { - x_366 = lean_alloc_ctor(0, 2, 0); -} else { - x_366 = x_269; +lean_object* x_12; +x_12 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +return x_12; } -lean_ctor_set(x_366, 0, x_365); -lean_ctor_set(x_366, 1, x_363); -x_367 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_367, 0, x_266); -lean_ctor_set(x_367, 1, x_366); -if (lean_is_scalar(x_352)) { - x_368 = lean_alloc_ctor(1, 1, 0); -} else { - x_368 = x_352; } -lean_ctor_set(x_368, 0, x_367); -x_17 = x_368; -x_18 = x_274; -goto block_23; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_4); +lean_dec(x_3); +return x_10; } -else +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___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, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: { -uint8_t x_369; -x_369 = lean_nat_dec_le(x_359, x_359); -if (x_369 == 0) +uint8_t x_15; +x_15 = lean_usize_dec_lt(x_5, x_4); +if (x_15 == 0) { -uint8_t x_370; lean_object* x_371; lean_object* x_372; lean_object* x_373; lean_object* x_374; -lean_dec(x_359); -lean_dec(x_358); -x_370 = 1; -x_371 = lean_box(x_370); -if (lean_is_scalar(x_269)) { - x_372 = lean_alloc_ctor(0, 2, 0); -} else { - x_372 = x_269; -} -lean_ctor_set(x_372, 0, x_371); -lean_ctor_set(x_372, 1, x_363); -x_373 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_373, 0, x_266); -lean_ctor_set(x_373, 1, x_372); -if (lean_is_scalar(x_352)) { - x_374 = lean_alloc_ctor(1, 1, 0); -} else { - x_374 = x_352; +lean_object* x_16; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_6); +lean_ctor_set(x_16, 1, x_14); +return x_16; } -lean_ctor_set(x_374, 0, x_373); -x_17 = x_374; -x_18 = x_274; -goto block_23; +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_17 = lean_array_uget(x_3, x_5); +x_18 = lean_ctor_get(x_6, 0); +lean_inc(x_18); +x_19 = lean_ctor_get(x_6, 1); +lean_inc(x_19); +x_20 = lean_ctor_get(x_6, 2); +lean_inc(x_20); +x_21 = lean_nat_dec_lt(x_19, x_20); +if (x_21 == 0) +{ +lean_object* x_22; +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_6); +lean_ctor_set(x_22, 1, x_14); +return x_22; } else { -size_t x_375; size_t x_376; lean_object* x_377; uint8_t x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; lean_object* x_382; -x_375 = 0; -x_376 = lean_usize_of_nat(x_359); -lean_dec(x_359); -x_377 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__12___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__13(x_3, x_358, x_375, x_376, x_266); -lean_dec(x_358); -x_378 = 1; -x_379 = lean_box(x_378); -if (lean_is_scalar(x_269)) { - x_380 = lean_alloc_ctor(0, 2, 0); -} else { - x_380 = x_269; +uint8_t x_23; +x_23 = !lean_is_exclusive(x_6); +if (x_23 == 0) +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; lean_object* x_29; lean_object* x_30; +x_24 = lean_ctor_get(x_6, 2); +lean_dec(x_24); +x_25 = lean_ctor_get(x_6, 1); +lean_dec(x_25); +x_26 = lean_ctor_get(x_6, 0); +lean_dec(x_26); +x_27 = lean_array_fget(x_18, x_19); +x_28 = lean_unbox(x_27); +lean_dec(x_27); +x_29 = lean_unsigned_to_nat(1u); +x_30 = lean_nat_add(x_19, x_29); +lean_dec(x_19); +lean_ctor_set(x_6, 1, x_30); +if (x_28 == 0) +{ +lean_object* x_31; +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_31 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect(x_17, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_31) == 0) +{ +lean_object* x_32; size_t x_33; size_t x_34; +x_32 = lean_ctor_get(x_31, 1); +lean_inc(x_32); +lean_dec(x_31); +x_33 = 1; +x_34 = lean_usize_add(x_5, x_33); +x_5 = x_34; +x_14 = x_32; +goto _start; } -lean_ctor_set(x_380, 0, x_379); -lean_ctor_set(x_380, 1, x_363); -x_381 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_381, 0, x_377); -lean_ctor_set(x_381, 1, x_380); -if (lean_is_scalar(x_352)) { - x_382 = lean_alloc_ctor(1, 1, 0); -} else { - x_382 = x_352; +else +{ +uint8_t x_36; +lean_dec(x_6); +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_31); +if (x_36 == 0) +{ +return x_31; +} +else +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_37 = lean_ctor_get(x_31, 0); +x_38 = lean_ctor_get(x_31, 1); +lean_inc(x_38); +lean_inc(x_37); +lean_dec(x_31); +x_39 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_39, 0, x_37); +lean_ctor_set(x_39, 1, x_38); +return x_39; } -lean_ctor_set(x_382, 0, x_381); -x_17 = x_382; -x_18 = x_274; -goto block_23; } } +else +{ +size_t x_40; size_t x_41; +lean_dec(x_17); +x_40 = 1; +x_41 = lean_usize_add(x_5, x_40); +x_5 = x_41; +goto _start; } } else { -lean_object* x_383; lean_object* x_384; lean_object* x_385; lean_object* x_386; -lean_dec(x_270); -lean_dec(x_269); -lean_dec(x_268); -lean_dec(x_267); -lean_dec(x_266); +lean_object* x_43; uint8_t x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; +lean_dec(x_6); +x_43 = lean_array_fget(x_18, x_19); +x_44 = lean_unbox(x_43); +lean_dec(x_43); +x_45 = lean_unsigned_to_nat(1u); +x_46 = lean_nat_add(x_19, x_45); +lean_dec(x_19); +x_47 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_47, 0, x_18); +lean_ctor_set(x_47, 1, x_46); +lean_ctor_set(x_47, 2, x_20); +if (x_44 == 0) +{ +lean_object* x_48; +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_48 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect(x_17, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_48) == 0) +{ +lean_object* x_49; size_t x_50; size_t x_51; +x_49 = lean_ctor_get(x_48, 1); +lean_inc(x_49); +lean_dec(x_48); +x_50 = 1; +x_51 = lean_usize_add(x_5, x_50); +x_5 = x_51; +x_6 = x_47; +x_14 = x_49; +goto _start; +} +else +{ +lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; +lean_dec(x_47); +lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); -lean_dec(x_3); -lean_dec(x_2); -x_383 = lean_ctor_get(x_272, 0); -lean_inc(x_383); -x_384 = lean_ctor_get(x_272, 1); -lean_inc(x_384); -if (lean_is_exclusive(x_272)) { - lean_ctor_release(x_272, 0); - lean_ctor_release(x_272, 1); - x_385 = x_272; +lean_dec(x_8); +lean_dec(x_7); +x_53 = lean_ctor_get(x_48, 0); +lean_inc(x_53); +x_54 = lean_ctor_get(x_48, 1); +lean_inc(x_54); +if (lean_is_exclusive(x_48)) { + lean_ctor_release(x_48, 0); + lean_ctor_release(x_48, 1); + x_55 = x_48; } else { - lean_dec_ref(x_272); - x_385 = lean_box(0); + lean_dec_ref(x_48); + x_55 = lean_box(0); } -if (lean_is_scalar(x_385)) { - x_386 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_55)) { + x_56 = lean_alloc_ctor(1, 2, 0); } else { - x_386 = x_385; + x_56 = x_55; } -lean_ctor_set(x_386, 0, x_383); -lean_ctor_set(x_386, 1, x_384); -return x_386; +lean_ctor_set(x_56, 0, x_53); +lean_ctor_set(x_56, 1, x_54); +return x_56; } } else { -lean_object* x_387; lean_object* x_388; lean_object* x_389; lean_object* x_390; -lean_dec(x_270); -lean_dec(x_16); -if (lean_is_exclusive(x_271)) { - lean_ctor_release(x_271, 0); - x_387 = x_271; -} else { - lean_dec_ref(x_271); - x_387 = lean_box(0); -} -if (lean_is_scalar(x_269)) { - x_388 = lean_alloc_ctor(0, 2, 0); -} else { - x_388 = x_269; -} -lean_ctor_set(x_388, 0, x_267); -lean_ctor_set(x_388, 1, x_268); -x_389 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_389, 0, x_266); -lean_ctor_set(x_389, 1, x_388); -if (lean_is_scalar(x_387)) { - x_390 = lean_alloc_ctor(1, 1, 0); -} else { - x_390 = x_387; -} -lean_ctor_set(x_390, 0, x_389); -x_17 = x_390; -x_18 = x_13; -goto block_23; -} -} -block_23: -{ -lean_object* x_19; size_t x_20; size_t x_21; -x_19 = lean_ctor_get(x_17, 0); -lean_inc(x_19); +size_t x_57; size_t x_58; lean_dec(x_17); -x_20 = 1; -x_21 = lean_usize_add(x_7, x_20); -x_7 = x_21; -x_8 = x_19; -x_13 = x_18; +x_57 = 1; +x_58 = lean_usize_add(x_5, x_57); +x_5 = x_58; +x_6 = x_47; goto _start; } } } } -LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___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: -{ -if (x_1 == 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_box(x_1); -x_12 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_12, 0, x_11); -lean_ctor_set(x_12, 1, x_2); -x_13 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_13, 0, x_3); -lean_ctor_set(x_13, 1, x_12); -x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_4); -lean_ctor_set(x_14, 1, x_13); -x_15 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_15, 0, x_14); -x_16 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_16, 0, x_15); -lean_ctor_set(x_16, 1, x_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; -x_17 = lean_box(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_2); -x_19 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_19, 0, x_3); -lean_ctor_set(x_19, 1, x_18); -x_20 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_20, 0, x_4); -lean_ctor_set(x_20, 1, x_19); -x_21 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_21, 0, x_20); -x_22 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_22, 0, x_21); -lean_ctor_set(x_22, 1, x_10); -return x_22; -} -} -} -static lean_object* _init_l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___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_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___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, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -lean_object* x_13; uint8_t x_14; -x_13 = lean_ctor_get(x_7, 1); +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_12 = lean_unsigned_to_nat(0u); +x_13 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_12); +x_14 = l_Lean_Meta_Grind_ppPattern___closed__5; lean_inc(x_13); -lean_dec(x_7); -x_14 = !lean_is_exclusive(x_13); -if (x_14 == 0) -{ -lean_object* x_15; uint8_t x_16; -x_15 = lean_ctor_get(x_13, 1); -x_16 = !lean_is_exclusive(x_15); -if (x_16 == 0) -{ -lean_object* x_17; lean_object* x_18; uint8_t x_19; lean_object* x_20; size_t x_21; size_t x_22; lean_object* x_23; -x_17 = lean_ctor_get(x_15, 0); -lean_dec(x_17); -x_18 = lean_box(0); -x_19 = 0; -x_20 = lean_box(x_19); -lean_ctor_set(x_15, 0, x_20); -x_21 = lean_array_size(x_2); -x_22 = 0; -lean_inc(x_11); +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_1, x_15, x_17); +x_19 = lean_array_get_size(x_18); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); -lean_inc(x_5); -lean_inc(x_3); -x_23 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__14___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__15(x_2, x_3, x_5, x_18, x_2, x_21, x_22, x_13, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_23) == 0) +lean_inc(x_7); +x_20 = l_Lean_Meta_Grind_NormalizePattern_getPatternSupportMask(x_2, x_19, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_20) == 0) { -lean_object* x_24; lean_object* x_25; uint8_t x_26; -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_is_exclusive(x_23); -if (x_26 == 0) +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; size_t x_26; size_t x_27; lean_object* x_28; +x_21 = lean_ctor_get(x_20, 0); +lean_inc(x_21); +x_22 = lean_ctor_get(x_20, 1); +lean_inc(x_22); +lean_dec(x_20); +x_23 = lean_array_get_size(x_21); +x_24 = l_Array_toSubarray___rarg(x_21, x_12, x_23); +x_25 = lean_box(0); +x_26 = lean_array_size(x_18); +x_27 = 0; +x_28 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___spec__1(x_18, x_25, x_18, x_26, x_27, x_24, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_22); +lean_dec(x_18); +if (lean_obj_tag(x_28) == 0) { -lean_object* x_27; lean_object* x_28; uint8_t x_29; -x_27 = lean_ctor_get(x_23, 1); -x_28 = lean_ctor_get(x_23, 0); -lean_dec(x_28); -x_29 = !lean_is_exclusive(x_24); +uint8_t x_29; +x_29 = !lean_is_exclusive(x_28); if (x_29 == 0) { -lean_object* x_30; lean_object* x_31; uint8_t x_32; -x_30 = lean_ctor_get(x_24, 0); -x_31 = lean_ctor_get(x_24, 1); -lean_dec(x_31); -x_32 = !lean_is_exclusive(x_25); -if (x_32 == 0) -{ -lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; -x_33 = lean_ctor_get(x_25, 0); -x_34 = lean_ctor_get(x_25, 1); -x_35 = lean_unsigned_to_nat(0u); -x_36 = l_Lean_RBNode_fold___at_Lean_RBMap_size___spec__1___rarg(x_35, x_30); -x_37 = lean_nat_dec_eq(x_36, x_1); -lean_dec(x_36); -if (x_37 == 0) -{ -lean_object* x_38; uint8_t x_39; lean_object* x_40; lean_object* x_41; -lean_free_object(x_25); -lean_free_object(x_24); -lean_free_object(x_23); -x_38 = lean_box(0); -x_39 = lean_unbox(x_33); -lean_dec(x_33); -lean_inc(x_6); -x_40 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___lambda__1(x_39, x_34, x_30, x_6, x_38, x_8, x_9, x_10, x_11, x_27); -x_41 = lean_ctor_get(x_40, 0); -lean_inc(x_41); -if (lean_obj_tag(x_41) == 0) -{ -uint8_t x_42; -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_3); -x_42 = !lean_is_exclusive(x_40); -if (x_42 == 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_43; lean_object* x_44; -x_43 = lean_ctor_get(x_40, 0); -lean_dec(x_43); -x_44 = lean_ctor_get(x_41, 0); -lean_inc(x_44); -lean_dec(x_41); -lean_ctor_set(x_40, 0, x_44); -return x_40; +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_45; lean_object* x_46; lean_object* x_47; -x_45 = lean_ctor_get(x_40, 1); -lean_inc(x_45); -lean_dec(x_40); -x_46 = lean_ctor_get(x_41, 0); -lean_inc(x_46); -lean_dec(x_41); -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; -} +uint8_t x_35; +x_35 = !lean_is_exclusive(x_28); +if (x_35 == 0) +{ +return x_28; } else { -lean_object* x_48; lean_object* x_49; -x_48 = lean_ctor_get(x_40, 1); -lean_inc(x_48); -lean_dec(x_40); -x_49 = lean_ctor_get(x_41, 0); -lean_inc(x_49); -lean_dec(x_41); -x_7 = x_49; -x_12 = x_48; -goto _start; +lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_28, 0); +x_37 = lean_ctor_get(x_28, 1); +lean_inc(x_37); +lean_inc(x_36); +lean_dec(x_28); +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_51; lean_object* x_52; -lean_dec(x_11); +uint8_t x_39; +lean_dec(x_18); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); +lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_3); -x_51 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___closed__1; -x_52 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_52, 0, x_51); -lean_ctor_set(x_52, 1, x_24); -lean_ctor_set(x_23, 0, x_52); -return x_23; -} +lean_dec(x_4); +x_39 = !lean_is_exclusive(x_20); +if (x_39 == 0) +{ +return x_20; } else { -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_25, 0); -x_54 = lean_ctor_get(x_25, 1); -lean_inc(x_54); -lean_inc(x_53); -lean_dec(x_25); -x_55 = lean_unsigned_to_nat(0u); -x_56 = l_Lean_RBNode_fold___at_Lean_RBMap_size___spec__1___rarg(x_55, x_30); -x_57 = lean_nat_dec_eq(x_56, x_1); -lean_dec(x_56); -if (x_57 == 0) -{ -lean_object* x_58; uint8_t x_59; lean_object* x_60; lean_object* x_61; -lean_free_object(x_24); -lean_free_object(x_23); -x_58 = lean_box(0); -x_59 = lean_unbox(x_53); -lean_dec(x_53); -lean_inc(x_6); -x_60 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___lambda__1(x_59, x_54, x_30, x_6, x_58, x_8, x_9, x_10, x_11, x_27); -x_61 = lean_ctor_get(x_60, 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; -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_3); -x_62 = lean_ctor_get(x_60, 1); -lean_inc(x_62); -if (lean_is_exclusive(x_60)) { - lean_ctor_release(x_60, 0); - lean_ctor_release(x_60, 1); - x_63 = x_60; -} else { - lean_dec_ref(x_60); - x_63 = lean_box(0); +lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_40 = lean_ctor_get(x_20, 0); +x_41 = lean_ctor_get(x_20, 1); +lean_inc(x_41); +lean_inc(x_40); +lean_dec(x_20); +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; } -x_64 = lean_ctor_get(x_61, 0); -lean_inc(x_64); -lean_dec(x_61); -if (lean_is_scalar(x_63)) { - x_65 = lean_alloc_ctor(0, 2, 0); -} else { - x_65 = x_63; } -lean_ctor_set(x_65, 0, x_64); -lean_ctor_set(x_65, 1, x_62); -return x_65; } -else -{ -lean_object* x_66; lean_object* x_67; -x_66 = lean_ctor_get(x_60, 1); -lean_inc(x_66); -lean_dec(x_60); -x_67 = lean_ctor_get(x_61, 0); -lean_inc(x_67); -lean_dec(x_61); -x_7 = x_67; -x_12 = x_66; -goto _start; } +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; uint8_t x_12; +x_11 = lean_st_ref_set(x_5, x_1, x_10); +x_12 = !lean_is_exclusive(x_11); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_13 = lean_ctor_get(x_11, 0); +x_14 = lean_box(0); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_13); +lean_ctor_set(x_15, 1, x_14); +lean_ctor_set(x_11, 0, x_15); +return x_11; } else { -lean_object* x_69; lean_object* x_70; lean_object* x_71; +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_11, 0); +x_17 = lean_ctor_get(x_11, 1); +lean_inc(x_17); +lean_inc(x_16); 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_3); -x_69 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_69, 0, x_53); -lean_ctor_set(x_69, 1, x_54); -lean_ctor_set(x_24, 1, x_69); -x_70 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___closed__1; -x_71 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_71, 0, x_70); -lean_ctor_set(x_71, 1, x_24); -lean_ctor_set(x_23, 0, x_71); -return x_23; +x_18 = lean_box(0); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_16); +lean_ctor_set(x_19, 1, x_18); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_19); +lean_ctor_set(x_20, 1, x_17); +return x_20; } } } -else +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__3___closed__1() { +_start: { -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_72 = lean_ctor_get(x_24, 0); -lean_inc(x_72); -lean_dec(x_24); -x_73 = lean_ctor_get(x_25, 0); -lean_inc(x_73); -x_74 = lean_ctor_get(x_25, 1); -lean_inc(x_74); -if (lean_is_exclusive(x_25)) { - lean_ctor_release(x_25, 0); - lean_ctor_release(x_25, 1); - x_75 = x_25; -} else { - lean_dec_ref(x_25); - x_75 = lean_box(0); +lean_object* x_1; +x_1 = lean_mk_string_unchecked("skip, no new variables covered", 30, 30); +return x_1; } -x_76 = lean_unsigned_to_nat(0u); -x_77 = l_Lean_RBNode_fold___at_Lean_RBMap_size___spec__1___rarg(x_76, x_72); -x_78 = lean_nat_dec_eq(x_77, x_1); -lean_dec(x_77); -if (x_78 == 0) +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__3___closed__2() { +_start: { -lean_object* x_79; uint8_t x_80; lean_object* x_81; lean_object* x_82; -lean_dec(x_75); -lean_free_object(x_23); -x_79 = lean_box(0); -x_80 = lean_unbox(x_73); -lean_dec(x_73); -lean_inc(x_6); -x_81 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___lambda__1(x_80, x_74, x_72, x_6, x_79, x_8, x_9, x_10, x_11, x_27); -x_82 = lean_ctor_get(x_81, 0); -lean_inc(x_82); -if (lean_obj_tag(x_82) == 0) +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__3___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: { -lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -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; -} else { - lean_dec_ref(x_81); - x_84 = lean_box(0); +lean_object* x_12; lean_object* x_13; uint8_t x_14; +lean_inc(x_1); +x_12 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___spec__1(x_1, 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_15; lean_object* x_16; lean_object* x_17; +lean_dec(x_1); +x_15 = lean_ctor_get(x_12, 1); +lean_inc(x_15); +lean_dec(x_12); +x_16 = lean_box(0); +x_17 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__2(x_2, x_16, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_15); +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; +x_18 = lean_ctor_get(x_12, 1); +lean_inc(x_18); +lean_dec(x_12); +x_19 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__3___closed__2; +x_20 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___spec__2(x_1, x_19, 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); +x_22 = lean_ctor_get(x_20, 1); +lean_inc(x_22); +lean_dec(x_20); +x_23 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__2(x_2, x_21, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_22); +lean_dec(x_21); +return x_23; } -x_85 = lean_ctor_get(x_82, 0); -lean_inc(x_85); -lean_dec(x_82); -if (lean_is_scalar(x_84)) { - x_86 = lean_alloc_ctor(0, 2, 0); -} else { - x_86 = x_84; } -lean_ctor_set(x_86, 0, x_85); -lean_ctor_set(x_86, 1, x_83); -return x_86; } -else +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__4___closed__1() { +_start: { -lean_object* x_87; lean_object* x_88; -x_87 = lean_ctor_get(x_81, 1); -lean_inc(x_87); -lean_dec(x_81); -x_88 = lean_ctor_get(x_82, 0); -lean_inc(x_88); -lean_dec(x_82); -x_7 = x_88; -x_12 = x_87; -goto _start; +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(0); +x_2 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_2, 0, x_1); +lean_ctor_set(x_2, 1, x_1); +return x_2; } } -else +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___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) { +_start: { -lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; +uint8_t x_13; lean_object* x_14; +x_13 = 0; +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +x_14 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go(x_1, x_13, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; 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_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +x_16 = lean_ctor_get(x_14, 1); +lean_inc(x_16); +lean_dec(x_14); +x_17 = lean_st_ref_get(x_7, x_16); +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +x_19 = lean_ctor_get(x_17, 1); +lean_inc(x_19); +lean_dec(x_17); +x_20 = lean_ctor_get(x_3, 2); +lean_inc(x_20); +x_21 = lean_ctor_get(x_20, 0); +lean_inc(x_21); +lean_dec(x_20); +x_22 = lean_ctor_get(x_18, 2); +lean_inc(x_22); +lean_dec(x_18); +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +lean_dec(x_22); +x_24 = lean_nat_dec_lt(x_21, x_23); +lean_dec(x_23); +lean_dec(x_21); +if (x_24 == 0) +{ +lean_object* x_25; lean_object* x_26; +lean_dec(x_15); +x_25 = lean_box(0); +x_26 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__3(x_2, x_3, x_25, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_19); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); -lean_dec(x_6); +lean_dec(x_7); lean_dec(x_5); +return x_26; +} +else +{ +lean_object* x_27; lean_dec(x_3); -if (lean_is_scalar(x_75)) { - x_90 = lean_alloc_ctor(0, 2, 0); -} else { - x_90 = x_75; +lean_dec(x_2); +x_27 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern(x_15, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_19); +lean_dec(x_7); +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; lean_object* x_30; +x_29 = lean_ctor_get(x_27, 0); +lean_dec(x_29); +x_30 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__4___closed__1; +lean_ctor_set(x_27, 0, x_30); +return x_27; } -lean_ctor_set(x_90, 0, x_73); -lean_ctor_set(x_90, 1, x_74); -x_91 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_91, 0, x_72); -lean_ctor_set(x_91, 1, x_90); -x_92 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___closed__1; -x_93 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_93, 0, x_92); -lean_ctor_set(x_93, 1, x_91); -lean_ctor_set(x_23, 0, x_93); -return x_23; +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_27, 1); +lean_inc(x_31); +lean_dec(x_27); +x_32 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__4___closed__1; +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_31); +return x_33; } } +else +{ +uint8_t x_34; +x_34 = !lean_is_exclusive(x_27); +if (x_34 == 0) +{ +return x_27; } 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; uint8_t x_102; -x_94 = lean_ctor_get(x_23, 1); -lean_inc(x_94); -lean_dec(x_23); -x_95 = lean_ctor_get(x_24, 0); -lean_inc(x_95); -if (lean_is_exclusive(x_24)) { - lean_ctor_release(x_24, 0); - lean_ctor_release(x_24, 1); - x_96 = x_24; -} else { - lean_dec_ref(x_24); - x_96 = lean_box(0); +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; } -x_97 = lean_ctor_get(x_25, 0); -lean_inc(x_97); -x_98 = lean_ctor_get(x_25, 1); -lean_inc(x_98); -if (lean_is_exclusive(x_25)) { - lean_ctor_release(x_25, 0); - lean_ctor_release(x_25, 1); - x_99 = x_25; -} else { - lean_dec_ref(x_25); - x_99 = lean_box(0); } -x_100 = lean_unsigned_to_nat(0u); -x_101 = l_Lean_RBNode_fold___at_Lean_RBMap_size___spec__1___rarg(x_100, x_95); -x_102 = lean_nat_dec_eq(x_101, x_1); -lean_dec(x_101); -if (x_102 == 0) -{ -lean_object* x_103; uint8_t x_104; lean_object* x_105; lean_object* x_106; -lean_dec(x_99); -lean_dec(x_96); -x_103 = lean_box(0); -x_104 = lean_unbox(x_97); -lean_dec(x_97); -lean_inc(x_6); -x_105 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___lambda__1(x_104, x_98, x_95, x_6, x_103, x_8, x_9, x_10, x_11, x_94); -x_106 = lean_ctor_get(x_105, 0); -lean_inc(x_106); -if (lean_obj_tag(x_106) == 0) +} +} +else { -lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; +uint8_t x_38; lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); -lean_dec(x_6); +lean_dec(x_7); lean_dec(x_5); lean_dec(x_3); -x_107 = lean_ctor_get(x_105, 1); -lean_inc(x_107); -if (lean_is_exclusive(x_105)) { - lean_ctor_release(x_105, 0); - lean_ctor_release(x_105, 1); - x_108 = x_105; -} else { - lean_dec_ref(x_105); - x_108 = lean_box(0); +lean_dec(x_2); +x_38 = !lean_is_exclusive(x_14); +if (x_38 == 0) +{ +return x_14; } -x_109 = lean_ctor_get(x_106, 0); -lean_inc(x_109); -lean_dec(x_106); -if (lean_is_scalar(x_108)) { - x_110 = lean_alloc_ctor(0, 2, 0); -} else { - x_110 = x_108; +else +{ +lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_39 = lean_ctor_get(x_14, 0); +x_40 = lean_ctor_get(x_14, 1); +lean_inc(x_40); +lean_inc(x_39); +lean_dec(x_14); +x_41 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_41, 0, x_39); +lean_ctor_set(x_41, 1, x_40); +return x_41; } -lean_ctor_set(x_110, 0, x_109); -lean_ctor_set(x_110, 1, x_107); -return x_110; } -else +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: { -lean_object* x_111; lean_object* x_112; -x_111 = lean_ctor_get(x_105, 1); -lean_inc(x_111); -lean_dec(x_105); -x_112 = lean_ctor_get(x_106, 0); -lean_inc(x_112); -lean_dec(x_106); -x_7 = x_112; -x_12 = x_111; -goto _start; +lean_object* x_10; lean_object* x_11; +x_10 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__4___closed__1; +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_10); +lean_ctor_set(x_11, 1, x_9); +return x_11; } } -else +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__6___closed__1() { +_start: { -lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; -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_3); -if (lean_is_scalar(x_99)) { - x_114 = lean_alloc_ctor(0, 2, 0); -} else { - x_114 = x_99; +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__5___boxed), 9, 0); +return x_1; } -lean_ctor_set(x_114, 0, x_97); -lean_ctor_set(x_114, 1, x_98); -if (lean_is_scalar(x_96)) { - x_115 = lean_alloc_ctor(0, 2, 0); -} else { - x_115 = x_96; } -lean_ctor_set(x_115, 0, x_95); -lean_ctor_set(x_115, 1, x_114); -x_116 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___closed__1; -x_117 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_117, 0, x_116); -lean_ctor_set(x_117, 1, x_115); -x_118 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_118, 0, x_117); -lean_ctor_set(x_118, 1, x_94); -return x_118; +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__6___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("skip, does not contain pattern variables", 40, 40); +return x_1; } } +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__6___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__6___closed__2; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } -else +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___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) { +_start: { -uint8_t x_119; -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_object* x_13; lean_object* x_14; uint8_t x_15; +x_13 = lean_ctor_get(x_5, 1); +lean_inc(x_13); +x_14 = lean_expr_abstract(x_1, x_13); +lean_dec(x_13); +x_15 = l_Lean_Expr_hasLooseBVars(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_20; +lean_dec(x_14); lean_dec(x_3); -x_119 = !lean_is_exclusive(x_23); -if (x_119 == 0) +lean_inc(x_2); +x_16 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___spec__1(x_2, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +x_18 = lean_ctor_get(x_16, 1); +lean_inc(x_18); +lean_dec(x_16); +x_19 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__6___closed__1; +x_20 = lean_unbox(x_17); +lean_dec(x_17); +if (x_20 == 0) { -return x_23; +lean_object* x_21; lean_object* x_22; +lean_dec(x_2); +x_21 = lean_box(0); +x_22 = lean_apply_9(x_19, x_21, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_18); +return x_22; } else { -lean_object* x_120; lean_object* x_121; lean_object* x_122; -x_120 = lean_ctor_get(x_23, 0); -x_121 = lean_ctor_get(x_23, 1); -lean_inc(x_121); -lean_inc(x_120); -lean_dec(x_23); -x_122 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_122, 0, x_120); -lean_ctor_set(x_122, 1, x_121); -return x_122; -} +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_23 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__6___closed__3; +x_24 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___spec__2(x_2, x_23, x_5, x_6, x_7, x_8, x_9, x_10, x_11, 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_apply_9(x_19, x_25, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_26); +return x_27; } } else { -lean_object* x_123; lean_object* x_124; uint8_t x_125; lean_object* x_126; lean_object* x_127; size_t x_128; size_t x_129; lean_object* x_130; -x_123 = lean_ctor_get(x_15, 1); -lean_inc(x_123); -lean_dec(x_15); -x_124 = lean_box(0); -x_125 = 0; -x_126 = lean_box(x_125); -x_127 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_127, 0, x_126); -lean_ctor_set(x_127, 1, x_123); -lean_ctor_set(x_13, 1, x_127); -x_128 = lean_array_size(x_2); -x_129 = 0; -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_5); -lean_inc(x_3); -x_130 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__14___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__15(x_2, x_3, x_5, x_124, x_2, x_128, x_129, x_13, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_130) == 0) -{ -lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; uint8_t x_142; -x_131 = lean_ctor_get(x_130, 0); -lean_inc(x_131); -x_132 = lean_ctor_get(x_131, 1); -lean_inc(x_132); -x_133 = lean_ctor_get(x_130, 1); -lean_inc(x_133); -if (lean_is_exclusive(x_130)) { - lean_ctor_release(x_130, 0); - lean_ctor_release(x_130, 1); - x_134 = x_130; -} else { - lean_dec_ref(x_130); - x_134 = lean_box(0); +lean_object* x_28; lean_object* x_29; +x_28 = lean_box(0); +x_29 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__4(x_14, x_2, x_3, x_28, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_6); +return x_29; } -x_135 = lean_ctor_get(x_131, 0); -lean_inc(x_135); -if (lean_is_exclusive(x_131)) { - lean_ctor_release(x_131, 0); - lean_ctor_release(x_131, 1); - x_136 = x_131; -} else { - lean_dec_ref(x_131); - x_136 = lean_box(0); } -x_137 = lean_ctor_get(x_132, 0); -lean_inc(x_137); -x_138 = lean_ctor_get(x_132, 1); -lean_inc(x_138); -if (lean_is_exclusive(x_132)) { - lean_ctor_release(x_132, 0); - lean_ctor_release(x_132, 1); - x_139 = x_132; -} else { - lean_dec_ref(x_132); - x_139 = lean_box(0); } -x_140 = lean_unsigned_to_nat(0u); -x_141 = l_Lean_RBNode_fold___at_Lean_RBMap_size___spec__1___rarg(x_140, x_135); -x_142 = lean_nat_dec_eq(x_141, x_1); -lean_dec(x_141); -if (x_142 == 0) +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__7___closed__1() { +_start: { -lean_object* x_143; uint8_t x_144; lean_object* x_145; lean_object* x_146; -lean_dec(x_139); -lean_dec(x_136); -lean_dec(x_134); -x_143 = lean_box(0); -x_144 = lean_unbox(x_137); -lean_dec(x_137); -lean_inc(x_6); -x_145 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___lambda__1(x_144, x_138, x_135, x_6, x_143, x_8, x_9, x_10, x_11, x_133); -x_146 = lean_ctor_get(x_145, 0); -lean_inc(x_146); -if (lean_obj_tag(x_146) == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("skip, exception during normalization", 36, 36); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__7___closed__2() { +_start: { -lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; -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_3); -x_147 = lean_ctor_get(x_145, 1); -lean_inc(x_147); -if (lean_is_exclusive(x_145)) { - lean_ctor_release(x_145, 0); - lean_ctor_release(x_145, 1); - x_148 = x_145; -} else { - lean_dec_ref(x_145); - x_148 = lean_box(0); +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__7___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } -x_149 = lean_ctor_get(x_146, 0); -lean_inc(x_149); -lean_dec(x_146); -if (lean_is_scalar(x_148)) { - x_150 = lean_alloc_ctor(0, 2, 0); -} else { - x_150 = x_148; } -lean_ctor_set(x_150, 0, x_149); -lean_ctor_set(x_150, 1, x_147); -return x_150; +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__7___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("candidate: ", 11, 11); +return x_1; } -else +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__7___closed__4() { +_start: { -lean_object* x_151; lean_object* x_152; -x_151 = lean_ctor_get(x_145, 1); -lean_inc(x_151); -lean_dec(x_145); -x_152 = lean_ctor_get(x_146, 0); -lean_inc(x_152); -lean_dec(x_146); -x_7 = x_152; -x_12 = x_151; -goto _start; +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__7___closed__3; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } } -else +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___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) { +_start: { -lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; -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_3); -if (lean_is_scalar(x_139)) { - x_154 = lean_alloc_ctor(0, 2, 0); -} else { - x_154 = x_139; +switch (lean_obj_tag(x_1)) { +case 5: +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_11 = l_Lean_Expr_getAppFn(x_1); +x_12 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isPatternFnCandidate(x_11, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_unbox(x_13); +lean_dec(x_13); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_ctor_get(x_12, 1); +lean_inc(x_15); +lean_dec(x_12); +x_16 = lean_box(0); +x_17 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__1(x_1, x_11, x_16, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_15); +return x_17; } -lean_ctor_set(x_154, 0, x_137); -lean_ctor_set(x_154, 1, x_138); -if (lean_is_scalar(x_136)) { - x_155 = lean_alloc_ctor(0, 2, 0); +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_32; lean_object* x_33; lean_object* x_55; lean_object* x_56; uint8_t x_57; +x_18 = lean_ctor_get(x_12, 1); +lean_inc(x_18); +if (lean_is_exclusive(x_12)) { + lean_ctor_release(x_12, 0); + lean_ctor_release(x_12, 1); + x_19 = x_12; } else { - x_155 = x_136; + lean_dec_ref(x_12); + x_19 = lean_box(0); } -lean_ctor_set(x_155, 0, x_135); -lean_ctor_set(x_155, 1, x_154); -x_156 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___closed__1; -x_157 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_157, 0, x_156); -lean_ctor_set(x_157, 1, x_155); -if (lean_is_scalar(x_134)) { - x_158 = lean_alloc_ctor(0, 2, 0); +x_20 = lean_st_ref_get(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); +if (lean_is_exclusive(x_20)) { + lean_ctor_release(x_20, 0); + lean_ctor_release(x_20, 1); + x_23 = x_20; } else { - x_158 = x_134; -} -lean_ctor_set(x_158, 0, x_157); -lean_ctor_set(x_158, 1, x_133); -return x_158; + lean_dec_ref(x_20); + x_23 = lean_box(0); } +x_24 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___closed__2; +x_55 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___spec__1(x_24, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_22); +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; lean_object* x_60; +x_58 = lean_ctor_get(x_55, 1); +lean_inc(x_58); +lean_dec(x_55); +x_59 = lean_box(0); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_21); +x_60 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__6(x_1, x_24, x_21, x_59, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_58); +if (lean_obj_tag(x_60) == 0) +{ +lean_object* x_61; lean_object* x_62; +lean_dec(x_21); +lean_dec(x_19); +x_61 = lean_ctor_get(x_60, 0); +lean_inc(x_61); +x_62 = lean_ctor_get(x_60, 1); +lean_inc(x_62); +lean_dec(x_60); +x_25 = x_61; +x_26 = x_62; +goto block_31; } else { -lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; -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_3); -x_159 = lean_ctor_get(x_130, 0); -lean_inc(x_159); -x_160 = lean_ctor_get(x_130, 1); -lean_inc(x_160); -if (lean_is_exclusive(x_130)) { - lean_ctor_release(x_130, 0); - lean_ctor_release(x_130, 1); - x_161 = x_130; -} else { - lean_dec_ref(x_130); - x_161 = lean_box(0); -} -if (lean_is_scalar(x_161)) { - x_162 = lean_alloc_ctor(1, 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; +lean_object* x_63; lean_object* x_64; +x_63 = lean_ctor_get(x_60, 0); +lean_inc(x_63); +x_64 = lean_ctor_get(x_60, 1); +lean_inc(x_64); +lean_dec(x_60); +x_32 = x_63; +x_33 = x_64; +goto block_54; } } +else +{ +uint8_t x_65; +x_65 = !lean_is_exclusive(x_55); +if (x_65 == 0) +{ +lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; +x_66 = lean_ctor_get(x_55, 1); +x_67 = lean_ctor_get(x_55, 0); +lean_dec(x_67); +lean_inc(x_1); +x_68 = l_Lean_MessageData_ofExpr(x_1); +x_69 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__7___closed__4; +lean_ctor_set_tag(x_55, 7); +lean_ctor_set(x_55, 1, x_68); +lean_ctor_set(x_55, 0, x_69); +x_70 = l_Lean_Meta_Grind_ppPattern___closed__4; +x_71 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_71, 0, x_55); +lean_ctor_set(x_71, 1, x_70); +x_72 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___spec__2(x_24, x_71, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_66); +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); +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_21); +x_75 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__6(x_1, x_24, x_21, x_73, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_74); +lean_dec(x_73); +if (lean_obj_tag(x_75) == 0) +{ +lean_object* x_76; lean_object* x_77; +lean_dec(x_21); +lean_dec(x_19); +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_25 = x_76; +x_26 = x_77; +goto block_31; } else { -lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; uint8_t x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; size_t x_172; size_t x_173; lean_object* x_174; -x_163 = lean_ctor_get(x_13, 1); -x_164 = lean_ctor_get(x_13, 0); -lean_inc(x_163); -lean_inc(x_164); -lean_dec(x_13); -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; -} else { - lean_dec_ref(x_163); - x_166 = lean_box(0); +lean_object* x_78; lean_object* x_79; +x_78 = lean_ctor_get(x_75, 0); +lean_inc(x_78); +x_79 = lean_ctor_get(x_75, 1); +lean_inc(x_79); +lean_dec(x_75); +x_32 = x_78; +x_33 = x_79; +goto block_54; } -x_167 = lean_box(0); -x_168 = 0; -x_169 = lean_box(x_168); -if (lean_is_scalar(x_166)) { - x_170 = lean_alloc_ctor(0, 2, 0); -} else { - x_170 = x_166; } -lean_ctor_set(x_170, 0, x_169); -lean_ctor_set(x_170, 1, x_165); -x_171 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_171, 0, x_164); -lean_ctor_set(x_171, 1, x_170); -x_172 = lean_array_size(x_2); -x_173 = 0; -lean_inc(x_11); -lean_inc(x_10); +else +{ +lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; +x_80 = lean_ctor_get(x_55, 1); +lean_inc(x_80); +lean_dec(x_55); +lean_inc(x_1); +x_81 = l_Lean_MessageData_ofExpr(x_1); +x_82 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__7___closed__4; +x_83 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_83, 0, x_82); +lean_ctor_set(x_83, 1, x_81); +x_84 = l_Lean_Meta_Grind_ppPattern___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_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___spec__2(x_24, x_85, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_80); +x_87 = lean_ctor_get(x_86, 0); +lean_inc(x_87); +x_88 = lean_ctor_get(x_86, 1); +lean_inc(x_88); +lean_dec(x_86); 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_174 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__14___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__15(x_2, x_3, x_5, x_167, x_2, x_172, x_173, x_171, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_174) == 0) +lean_inc(x_21); +x_89 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__6(x_1, x_24, x_21, x_87, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_88); +lean_dec(x_87); +if (lean_obj_tag(x_89) == 0) { -lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; uint8_t x_186; -x_175 = lean_ctor_get(x_174, 0); -lean_inc(x_175); -x_176 = lean_ctor_get(x_175, 1); -lean_inc(x_176); -x_177 = lean_ctor_get(x_174, 1); -lean_inc(x_177); -if (lean_is_exclusive(x_174)) { - lean_ctor_release(x_174, 0); - lean_ctor_release(x_174, 1); - x_178 = x_174; -} else { - lean_dec_ref(x_174); - x_178 = lean_box(0); +lean_object* x_90; lean_object* x_91; +lean_dec(x_21); +lean_dec(x_19); +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_25 = x_90; +x_26 = x_91; +goto block_31; } -x_179 = lean_ctor_get(x_175, 0); -lean_inc(x_179); -if (lean_is_exclusive(x_175)) { - lean_ctor_release(x_175, 0); - lean_ctor_release(x_175, 1); - x_180 = x_175; -} else { - lean_dec_ref(x_175); - x_180 = lean_box(0); +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_32 = x_92; +x_33 = x_93; +goto block_54; } -x_181 = lean_ctor_get(x_176, 0); -lean_inc(x_181); -x_182 = lean_ctor_get(x_176, 1); -lean_inc(x_182); -if (lean_is_exclusive(x_176)) { - lean_ctor_release(x_176, 0); - lean_ctor_release(x_176, 1); - x_183 = x_176; -} else { - lean_dec_ref(x_176); - x_183 = lean_box(0); } -x_184 = lean_unsigned_to_nat(0u); -x_185 = l_Lean_RBNode_fold___at_Lean_RBMap_size___spec__1___rarg(x_184, x_179); -x_186 = lean_nat_dec_eq(x_185, x_1); -lean_dec(x_185); -if (x_186 == 0) +} +block_31: { -lean_object* x_187; uint8_t x_188; lean_object* x_189; lean_object* x_190; -lean_dec(x_183); -lean_dec(x_180); -lean_dec(x_178); -x_187 = lean_box(0); -x_188 = lean_unbox(x_181); -lean_dec(x_181); -lean_inc(x_6); -x_189 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___lambda__1(x_188, x_182, x_179, x_6, x_187, x_8, x_9, x_10, x_11, x_177); -x_190 = lean_ctor_get(x_189, 0); -lean_inc(x_190); -if (lean_obj_tag(x_190) == 0) +if (lean_obj_tag(x_25) == 0) { -lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; +lean_object* x_27; lean_object* x_28; +lean_dec(x_23); +x_27 = lean_ctor_get(x_25, 0); +lean_inc(x_27); +lean_dec(x_25); +x_28 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__1(x_1, x_11, x_27, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_26); +lean_dec(x_27); +return x_28; +} +else +{ +lean_object* x_29; lean_object* x_30; lean_dec(x_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_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_dec(x_1); +x_29 = lean_ctor_get(x_25, 0); +lean_inc(x_29); +lean_dec(x_25); +if (lean_is_scalar(x_23)) { + x_30 = lean_alloc_ctor(0, 2, 0); } else { - lean_dec_ref(x_189); - x_192 = lean_box(0); + x_30 = x_23; } -x_193 = lean_ctor_get(x_190, 0); -lean_inc(x_193); -lean_dec(x_190); -if (lean_is_scalar(x_192)) { - x_194 = lean_alloc_ctor(0, 2, 0); -} else { - x_194 = x_192; +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_26); +return x_30; } -lean_ctor_set(x_194, 0, x_193); -lean_ctor_set(x_194, 1, x_191); -return x_194; +} +block_54: +{ +uint8_t x_34; +x_34 = l_Lean_Exception_isInterrupt(x_32); +if (x_34 == 0) +{ +uint8_t x_35; +x_35 = l_Lean_Exception_isRuntime(x_32); +if (x_35 == 0) +{ +lean_object* x_36; lean_object* x_37; uint8_t x_38; +lean_dec(x_32); +lean_dec(x_19); +x_36 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___spec__1(x_24, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_33); +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_39; lean_object* x_40; lean_object* x_41; lean_object* 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 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__2(x_21, x_40, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_39); +x_42 = lean_ctor_get(x_41, 0); +lean_inc(x_42); +x_43 = lean_ctor_get(x_41, 1); +lean_inc(x_43); +lean_dec(x_41); +x_25 = x_42; +x_26 = x_43; +goto block_31; } else { -lean_object* x_195; lean_object* x_196; -x_195 = lean_ctor_get(x_189, 1); -lean_inc(x_195); -lean_dec(x_189); -x_196 = lean_ctor_get(x_190, 0); -lean_inc(x_196); -lean_dec(x_190); -x_7 = x_196; -x_12 = x_195; -goto _start; +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_44 = lean_ctor_get(x_36, 1); +lean_inc(x_44); +lean_dec(x_36); +x_45 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__7___closed__2; +x_46 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___spec__2(x_24, x_45, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_44); +x_47 = lean_ctor_get(x_46, 0); +lean_inc(x_47); +x_48 = lean_ctor_get(x_46, 1); +lean_inc(x_48); +lean_dec(x_46); +x_49 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__2(x_21, x_47, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_48); +lean_dec(x_47); +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_25 = x_50; +x_26 = x_51; +goto block_31; } } else { -lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; +lean_object* x_52; +lean_dec(x_23); +lean_dec(x_21); lean_dec(x_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); -if (lean_is_scalar(x_183)) { - x_198 = lean_alloc_ctor(0, 2, 0); -} else { - x_198 = x_183; -} -lean_ctor_set(x_198, 0, x_181); -lean_ctor_set(x_198, 1, x_182); -if (lean_is_scalar(x_180)) { - x_199 = lean_alloc_ctor(0, 2, 0); -} else { - x_199 = x_180; -} -lean_ctor_set(x_199, 0, x_179); -lean_ctor_set(x_199, 1, x_198); -x_200 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___closed__1; -x_201 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_201, 0, x_200); -lean_ctor_set(x_201, 1, x_199); -if (lean_is_scalar(x_178)) { - x_202 = lean_alloc_ctor(0, 2, 0); +lean_dec(x_1); +if (lean_is_scalar(x_19)) { + x_52 = lean_alloc_ctor(1, 2, 0); } else { - x_202 = x_178; + x_52 = x_19; + lean_ctor_set_tag(x_52, 1); } -lean_ctor_set(x_202, 0, x_201); -lean_ctor_set(x_202, 1, x_177); -return x_202; +lean_ctor_set(x_52, 0, x_32); +lean_ctor_set(x_52, 1, x_33); +return x_52; } } else { -lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; +lean_object* x_53; +lean_dec(x_23); +lean_dec(x_21); lean_dec(x_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_203 = lean_ctor_get(x_174, 0); -lean_inc(x_203); -x_204 = lean_ctor_get(x_174, 1); -lean_inc(x_204); -if (lean_is_exclusive(x_174)) { - lean_ctor_release(x_174, 0); - lean_ctor_release(x_174, 1); - x_205 = x_174; -} else { - lean_dec_ref(x_174); - x_205 = lean_box(0); -} -if (lean_is_scalar(x_205)) { - x_206 = lean_alloc_ctor(1, 2, 0); +lean_dec(x_1); +if (lean_is_scalar(x_19)) { + x_53 = lean_alloc_ctor(1, 2, 0); } else { - x_206 = x_205; + x_53 = x_19; + lean_ctor_set_tag(x_53, 1); } -lean_ctor_set(x_206, 0, x_203); -lean_ctor_set(x_206, 1, x_204); -return x_206; +lean_ctor_set(x_53, 0, x_32); +lean_ctor_set(x_53, 1, x_33); +return x_53; } } } } -LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__17(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: +case 7: { -lean_object* x_12; uint8_t x_13; -x_12 = lean_ctor_get(x_6, 1); -lean_inc(x_12); -lean_dec(x_6); -x_13 = !lean_is_exclusive(x_12); -if (x_13 == 0) +lean_object* x_94; lean_object* x_95; uint8_t x_96; +x_94 = lean_ctor_get(x_1, 1); +lean_inc(x_94); +x_95 = lean_ctor_get(x_1, 2); +lean_inc(x_95); +x_96 = l_Lean_Expr_isArrow(x_1); +lean_dec(x_1); +if (x_96 == 0) { -lean_object* x_14; uint8_t x_15; -x_14 = lean_ctor_get(x_12, 1); -x_15 = !lean_is_exclusive(x_14); -if (x_15 == 0) +lean_object* x_97; lean_object* x_98; +lean_dec(x_95); +lean_dec(x_94); +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_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_10); +return x_98; +} +else { -lean_object* x_16; lean_object* x_17; uint8_t x_18; lean_object* x_19; size_t x_20; size_t x_21; lean_object* x_22; -x_16 = lean_ctor_get(x_14, 0); -lean_dec(x_16); -x_17 = lean_box(0); -x_18 = 0; -x_19 = lean_box(x_18); -lean_ctor_set(x_14, 0, x_19); -x_20 = lean_array_size(x_2); -x_21 = 0; -lean_inc(x_10); +lean_object* x_99; lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); -lean_inc(x_4); -lean_inc(x_3); -x_22 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__14___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__15(x_2, x_3, x_4, x_17, x_2, x_20, x_21, x_12, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_22) == 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_is_exclusive(x_22); -if (x_25 == 0) -{ -lean_object* x_26; lean_object* x_27; uint8_t x_28; -x_26 = lean_ctor_get(x_22, 1); -x_27 = lean_ctor_get(x_22, 0); -lean_dec(x_27); -x_28 = !lean_is_exclusive(x_23); -if (x_28 == 0) -{ -lean_object* x_29; lean_object* x_30; uint8_t x_31; -x_29 = lean_ctor_get(x_23, 0); -x_30 = lean_ctor_get(x_23, 1); -lean_dec(x_30); -x_31 = !lean_is_exclusive(x_24); -if (x_31 == 0) -{ -lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; -x_32 = lean_ctor_get(x_24, 0); -x_33 = lean_ctor_get(x_24, 1); -x_34 = lean_unsigned_to_nat(0u); -x_35 = l_Lean_RBNode_fold___at_Lean_RBMap_size___spec__1___rarg(x_34, x_29); -x_36 = lean_nat_dec_eq(x_35, x_1); -lean_dec(x_35); -if (x_36 == 0) +lean_inc(x_6); +lean_inc(x_94); +x_99 = l_Lean_Meta_isProp(x_94, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_99) == 0) { -lean_object* x_37; uint8_t x_38; lean_object* x_39; lean_object* x_40; -lean_free_object(x_24); -lean_free_object(x_23); -lean_free_object(x_22); -x_37 = lean_box(0); -x_38 = lean_unbox(x_32); -lean_dec(x_32); -lean_inc(x_5); -x_39 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___lambda__1(x_38, x_33, x_29, x_5, x_37, x_7, x_8, x_9, x_10, x_26); -x_40 = lean_ctor_get(x_39, 0); -lean_inc(x_40); -if (lean_obj_tag(x_40) == 0) +lean_object* x_100; uint8_t x_101; +x_100 = lean_ctor_get(x_99, 0); +lean_inc(x_100); +x_101 = lean_unbox(x_100); +lean_dec(x_100); +if (x_101 == 0) { -uint8_t x_41; -lean_dec(x_10); +uint8_t x_102; +lean_dec(x_95); +lean_dec(x_94); 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_41 = !lean_is_exclusive(x_39); -if (x_41 == 0) +x_102 = !lean_is_exclusive(x_99); +if (x_102 == 0) { -lean_object* x_42; lean_object* x_43; -x_42 = lean_ctor_get(x_39, 0); -lean_dec(x_42); -x_43 = lean_ctor_get(x_40, 0); -lean_inc(x_43); -lean_dec(x_40); -lean_ctor_set(x_39, 0, x_43); -return x_39; +lean_object* x_103; lean_object* x_104; +x_103 = lean_ctor_get(x_99, 0); +lean_dec(x_103); +x_104 = lean_box(0); +lean_ctor_set(x_99, 0, x_104); +return x_99; } else { -lean_object* x_44; lean_object* x_45; lean_object* x_46; -x_44 = lean_ctor_get(x_39, 1); -lean_inc(x_44); -lean_dec(x_39); -x_45 = lean_ctor_get(x_40, 0); -lean_inc(x_45); -lean_dec(x_40); -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_105; lean_object* x_106; lean_object* x_107; +x_105 = lean_ctor_get(x_99, 1); +lean_inc(x_105); +lean_dec(x_99); +x_106 = lean_box(0); +x_107 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_107, 0, x_106); +lean_ctor_set(x_107, 1, x_105); +return x_107; } } else { -lean_object* x_47; lean_object* x_48; -x_47 = lean_ctor_get(x_39, 1); -lean_inc(x_47); -lean_dec(x_39); -x_48 = lean_ctor_get(x_40, 0); -lean_inc(x_48); -lean_dec(x_40); -x_6 = x_48; -x_11 = x_47; -goto _start; -} -} -else +lean_object* x_108; lean_object* x_109; +x_108 = lean_ctor_get(x_99, 1); +lean_inc(x_108); +lean_dec(x_99); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_95); +x_109 = l_Lean_Meta_isProp(x_95, x_6, x_7, x_8, x_9, x_108); +if (lean_obj_tag(x_109) == 0) { -lean_object* x_50; lean_object* x_51; -lean_dec(x_10); +lean_object* x_110; uint8_t x_111; +x_110 = lean_ctor_get(x_109, 0); +lean_inc(x_110); +x_111 = lean_unbox(x_110); +lean_dec(x_110); +if (x_111 == 0) +{ +uint8_t x_112; +lean_dec(x_95); +lean_dec(x_94); 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_50 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___closed__1; -x_51 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_51, 0, x_50); -lean_ctor_set(x_51, 1, x_23); -lean_ctor_set(x_22, 0, x_51); -return x_22; -} +x_112 = !lean_is_exclusive(x_109); +if (x_112 == 0) +{ +lean_object* x_113; lean_object* x_114; +x_113 = lean_ctor_get(x_109, 0); +lean_dec(x_113); +x_114 = lean_box(0); +lean_ctor_set(x_109, 0, x_114); +return x_109; } else { -lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; uint8_t x_56; -x_52 = lean_ctor_get(x_24, 0); -x_53 = lean_ctor_get(x_24, 1); -lean_inc(x_53); -lean_inc(x_52); -lean_dec(x_24); -x_54 = lean_unsigned_to_nat(0u); -x_55 = l_Lean_RBNode_fold___at_Lean_RBMap_size___spec__1___rarg(x_54, x_29); -x_56 = lean_nat_dec_eq(x_55, x_1); -lean_dec(x_55); -if (x_56 == 0) +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_box(0); +x_117 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_117, 0, x_116); +lean_ctor_set(x_117, 1, x_115); +return x_117; +} +} +else { -lean_object* x_57; uint8_t x_58; lean_object* x_59; lean_object* x_60; -lean_free_object(x_23); -lean_free_object(x_22); -x_57 = lean_box(0); -x_58 = lean_unbox(x_52); -lean_dec(x_52); +lean_object* x_118; lean_object* x_119; +x_118 = lean_ctor_get(x_109, 1); +lean_inc(x_118); +lean_dec(x_109); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); lean_inc(x_5); -x_59 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___lambda__1(x_58, x_53, x_29, x_5, x_57, x_7, x_8, x_9, x_10, x_26); -x_60 = lean_ctor_get(x_59, 0); -lean_inc(x_60); -if (lean_obj_tag(x_60) == 0) +lean_inc(x_4); +lean_inc(x_3); +x_119 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect(x_94, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_118); +if (lean_obj_tag(x_119) == 0) { -lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; -lean_dec(x_10); +lean_object* x_120; lean_object* x_121; +x_120 = lean_ctor_get(x_119, 1); +lean_inc(x_120); +lean_dec(x_119); +x_121 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect(x_95, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_120); +return x_121; +} +else +{ +uint8_t x_122; +lean_dec(x_95); 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_61 = lean_ctor_get(x_59, 1); -lean_inc(x_61); -if (lean_is_exclusive(x_59)) { - lean_ctor_release(x_59, 0); - lean_ctor_release(x_59, 1); - x_62 = x_59; -} else { - lean_dec_ref(x_59); - x_62 = lean_box(0); -} -x_63 = lean_ctor_get(x_60, 0); -lean_inc(x_63); -lean_dec(x_60); -if (lean_is_scalar(x_62)) { - x_64 = lean_alloc_ctor(0, 2, 0); -} else { - x_64 = x_62; -} -lean_ctor_set(x_64, 0, x_63); -lean_ctor_set(x_64, 1, x_61); -return x_64; +x_122 = !lean_is_exclusive(x_119); +if (x_122 == 0) +{ +return x_119; } 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); -x_66 = lean_ctor_get(x_60, 0); -lean_inc(x_66); -lean_dec(x_60); -x_6 = x_66; -x_11 = x_65; -goto _start; +lean_object* x_123; lean_object* x_124; lean_object* x_125; +x_123 = lean_ctor_get(x_119, 0); +x_124 = lean_ctor_get(x_119, 1); +lean_inc(x_124); +lean_inc(x_123); +lean_dec(x_119); +x_125 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_125, 0, x_123); +lean_ctor_set(x_125, 1, x_124); +return x_125; +} +} } } else { -lean_object* x_68; lean_object* x_69; lean_object* x_70; -lean_dec(x_10); +uint8_t x_126; +lean_dec(x_95); +lean_dec(x_94); 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_68 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_68, 0, x_52); -lean_ctor_set(x_68, 1, x_53); -lean_ctor_set(x_23, 1, x_68); -x_69 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___closed__1; -x_70 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_70, 0, x_69); -lean_ctor_set(x_70, 1, x_23); -lean_ctor_set(x_22, 0, x_70); -return x_22; -} -} +x_126 = !lean_is_exclusive(x_109); +if (x_126 == 0) +{ +return x_109; } 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; uint8_t x_77; -x_71 = lean_ctor_get(x_23, 0); -lean_inc(x_71); -lean_dec(x_23); -x_72 = lean_ctor_get(x_24, 0); -lean_inc(x_72); -x_73 = lean_ctor_get(x_24, 1); -lean_inc(x_73); -if (lean_is_exclusive(x_24)) { - lean_ctor_release(x_24, 0); - lean_ctor_release(x_24, 1); - x_74 = x_24; -} else { - lean_dec_ref(x_24); - x_74 = lean_box(0); +lean_object* x_127; lean_object* x_128; lean_object* x_129; +x_127 = lean_ctor_get(x_109, 0); +x_128 = lean_ctor_get(x_109, 1); +lean_inc(x_128); +lean_inc(x_127); +lean_dec(x_109); +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; } -x_75 = lean_unsigned_to_nat(0u); -x_76 = l_Lean_RBNode_fold___at_Lean_RBMap_size___spec__1___rarg(x_75, x_71); -x_77 = lean_nat_dec_eq(x_76, x_1); -lean_dec(x_76); -if (x_77 == 0) -{ -lean_object* x_78; uint8_t x_79; lean_object* x_80; lean_object* x_81; -lean_dec(x_74); -lean_free_object(x_22); -x_78 = lean_box(0); -x_79 = lean_unbox(x_72); -lean_dec(x_72); -lean_inc(x_5); -x_80 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___lambda__1(x_79, x_73, x_71, x_5, x_78, x_7, x_8, x_9, x_10, x_26); -x_81 = lean_ctor_get(x_80, 0); -lean_inc(x_81); -if (lean_obj_tag(x_81) == 0) +} +} +} +else { -lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; -lean_dec(x_10); +uint8_t x_130; +lean_dec(x_95); +lean_dec(x_94); 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_82 = lean_ctor_get(x_80, 1); -lean_inc(x_82); -if (lean_is_exclusive(x_80)) { - lean_ctor_release(x_80, 0); - lean_ctor_release(x_80, 1); - x_83 = x_80; -} else { - lean_dec_ref(x_80); - x_83 = lean_box(0); -} -x_84 = lean_ctor_get(x_81, 0); -lean_inc(x_84); -lean_dec(x_81); -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; +x_130 = !lean_is_exclusive(x_99); +if (x_130 == 0) +{ +return x_99; } else { -lean_object* x_86; lean_object* x_87; -x_86 = lean_ctor_get(x_80, 1); -lean_inc(x_86); -lean_dec(x_80); -x_87 = lean_ctor_get(x_81, 0); -lean_inc(x_87); -lean_dec(x_81); -x_6 = x_87; -x_11 = x_86; -goto _start; +lean_object* x_131; lean_object* x_132; lean_object* x_133; +x_131 = lean_ctor_get(x_99, 0); +x_132 = lean_ctor_get(x_99, 1); +lean_inc(x_132); +lean_inc(x_131); +lean_dec(x_99); +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 +} +} +default: { -lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; -lean_dec(x_10); +lean_object* x_134; lean_object* x_135; lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); +lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -if (lean_is_scalar(x_74)) { - x_89 = lean_alloc_ctor(0, 2, 0); -} else { - x_89 = x_74; +lean_dec(x_1); +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_10); +return x_135; +} } -lean_ctor_set(x_89, 0, x_72); -lean_ctor_set(x_89, 1, x_73); -x_90 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_90, 0, x_71); -lean_ctor_set(x_90, 1, x_89); -x_91 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___closed__1; -x_92 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_92, 0, x_91); -lean_ctor_set(x_92, 1, x_90); -lean_ctor_set(x_22, 0, x_92); -return x_22; } } +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; lean_object* x_11; uint8_t x_12; +x_10 = lean_st_ref_get(x_3, x_9); +x_11 = lean_ctor_get(x_10, 0); +lean_inc(x_11); +x_12 = lean_ctor_get_uint8(x_11, sizeof(void*)*1); +lean_dec(x_11); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_13 = lean_ctor_get(x_10, 1); +lean_inc(x_13); +lean_dec(x_10); +x_14 = lean_box(0); +x_15 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__7(x_1, x_14, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_13); +return x_15; } else { -lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; uint8_t x_101; -x_93 = lean_ctor_get(x_22, 1); -lean_inc(x_93); -lean_dec(x_22); -x_94 = lean_ctor_get(x_23, 0); -lean_inc(x_94); -if (lean_is_exclusive(x_23)) { - lean_ctor_release(x_23, 0); - lean_ctor_release(x_23, 1); - x_95 = x_23; -} else { - lean_dec_ref(x_23); - x_95 = lean_box(0); +uint8_t x_16; +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_16 = !lean_is_exclusive(x_10); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; +x_17 = lean_ctor_get(x_10, 0); +lean_dec(x_17); +x_18 = lean_box(0); +lean_ctor_set(x_10, 0, x_18); +return x_10; } -x_96 = lean_ctor_get(x_24, 0); -lean_inc(x_96); -x_97 = lean_ctor_get(x_24, 1); -lean_inc(x_97); -if (lean_is_exclusive(x_24)) { - lean_ctor_release(x_24, 0); - lean_ctor_release(x_24, 1); - x_98 = x_24; -} else { - lean_dec_ref(x_24); - x_98 = lean_box(0); +else +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_10, 1); +lean_inc(x_19); +lean_dec(x_10); +x_20 = lean_box(0); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_20); +lean_ctor_set(x_21, 1, x_19); +return x_21; } -x_99 = lean_unsigned_to_nat(0u); -x_100 = l_Lean_RBNode_fold___at_Lean_RBMap_size___spec__1___rarg(x_99, x_94); -x_101 = lean_nat_dec_eq(x_100, x_1); -lean_dec(x_100); -if (x_101 == 0) +} +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: { -lean_object* x_102; uint8_t x_103; lean_object* x_104; lean_object* x_105; -lean_dec(x_98); -lean_dec(x_95); -x_102 = lean_box(0); -x_103 = lean_unbox(x_96); -lean_dec(x_96); -lean_inc(x_5); -x_104 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___lambda__1(x_103, x_97, x_94, x_5, x_102, x_7, x_8, x_9, x_10, x_93); -x_105 = lean_ctor_get(x_104, 0); -lean_inc(x_105); -if (lean_obj_tag(x_105) == 0) +size_t x_15; size_t x_16; lean_object* x_17; +x_15 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_16 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_17 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___spec__1(x_1, x_2, x_3, x_15, x_16, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_17; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___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_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; -lean_dec(x_10); +lean_object* x_12; +x_12 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___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_3); +return x_12; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); +lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_106 = lean_ctor_get(x_104, 1); -lean_inc(x_106); -if (lean_is_exclusive(x_104)) { - lean_ctor_release(x_104, 0); - lean_ctor_release(x_104, 1); - x_107 = x_104; -} else { - lean_dec_ref(x_104); - x_107 = lean_box(0); -} -x_108 = lean_ctor_get(x_105, 0); -lean_inc(x_108); -lean_dec(x_105); -if (lean_is_scalar(x_107)) { - x_109 = lean_alloc_ctor(0, 2, 0); -} else { - x_109 = x_107; -} -lean_ctor_set(x_109, 0, x_108); -lean_ctor_set(x_109, 1, x_106); -return x_109; -} -else -{ -lean_object* x_110; lean_object* x_111; -x_110 = lean_ctor_get(x_104, 1); -lean_inc(x_110); -lean_dec(x_104); -x_111 = lean_ctor_get(x_105, 0); -lean_inc(x_111); -lean_dec(x_105); -x_6 = x_111; -x_11 = x_110; -goto _start; +lean_dec(x_2); +return x_11; } } -else +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___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_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; +lean_object* x_12; +x_12 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); +lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -if (lean_is_scalar(x_98)) { - x_113 = lean_alloc_ctor(0, 2, 0); -} else { - x_113 = x_98; -} -lean_ctor_set(x_113, 0, x_96); -lean_ctor_set(x_113, 1, x_97); -if (lean_is_scalar(x_95)) { - x_114 = lean_alloc_ctor(0, 2, 0); -} else { - x_114 = x_95; +return x_12; } -lean_ctor_set(x_114, 0, x_94); -lean_ctor_set(x_114, 1, x_113); -x_115 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___closed__1; -x_116 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_116, 0, x_115); -lean_ctor_set(x_116, 1, x_114); -x_117 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_117, 0, x_116); -lean_ctor_set(x_117, 1, x_93); -return x_117; } +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +x_13 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___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); +lean_dec(x_6); +lean_dec(x_4); +return x_13; } } -else +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: { -uint8_t x_118; -lean_dec(x_10); -lean_dec(x_9); +lean_object* x_10; +x_10 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___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_dec(x_4); lean_dec(x_3); -x_118 = !lean_is_exclusive(x_22); -if (x_118 == 0) +lean_dec(x_2); +lean_dec(x_1); +return x_10; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: { -return x_22; +lean_object* x_13; +x_13 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__6(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_4); +lean_dec(x_1); +return x_13; } -else +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: { -lean_object* x_119; lean_object* x_120; lean_object* x_121; -x_119 = lean_ctor_get(x_22, 0); -x_120 = lean_ctor_get(x_22, 1); -lean_inc(x_120); -lean_inc(x_119); -lean_dec(x_22); -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_11; +x_11 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__7(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_2); +return x_11; } } +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collectPatterns_x3f___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, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +_start: +{ +uint8_t x_16; +x_16 = lean_usize_dec_lt(x_6, x_5); +if (x_16 == 0) +{ +lean_object* x_17; +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_3); +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_7); +lean_ctor_set(x_17, 1, x_15); +return x_17; } else { -lean_object* x_122; lean_object* x_123; uint8_t x_124; lean_object* x_125; lean_object* x_126; size_t x_127; size_t x_128; lean_object* x_129; -x_122 = lean_ctor_get(x_14, 1); -lean_inc(x_122); -lean_dec(x_14); -x_123 = lean_box(0); -x_124 = 0; -x_125 = lean_box(x_124); -x_126 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_126, 0, x_125); -lean_ctor_set(x_126, 1, x_122); -lean_ctor_set(x_12, 1, x_126); -x_127 = lean_array_size(x_2); -x_128 = 0; +lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_28; lean_object* x_29; +lean_dec(x_7); +x_18 = lean_array_uget(x_4, x_6); +x_28 = 1; +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +x_29 = l_Lean_Meta_Grind_preprocessPattern(x_18, x_28, x_11, x_12, x_13, x_14, x_15); +if (lean_obj_tag(x_29) == 0) +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_30 = lean_ctor_get(x_29, 0); +lean_inc(x_30); +x_31 = lean_ctor_get(x_29, 1); +lean_inc(x_31); +lean_dec(x_29); +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_4); +x_32 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect(x_30, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_31); +if (lean_obj_tag(x_32) == 0) +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; +x_33 = lean_ctor_get(x_32, 1); +lean_inc(x_33); +lean_dec(x_32); +x_34 = lean_st_ref_get(x_9, x_33); +x_35 = lean_ctor_get(x_34, 0); +lean_inc(x_35); +x_36 = lean_ctor_get_uint8(x_35, sizeof(void*)*1); +lean_dec(x_35); +if (x_36 == 0) +{ +lean_object* x_37; lean_object* x_38; +x_37 = lean_ctor_get(x_34, 1); +lean_inc(x_37); +lean_dec(x_34); lean_inc(x_3); -x_129 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__14___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__15(x_2, x_3, x_4, x_123, x_2, x_127, x_128, x_12, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_129) == 0) +x_38 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_38, 0, x_3); +x_19 = x_38; +x_20 = x_37; +goto block_27; +} +else { -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; uint8_t x_141; -x_130 = lean_ctor_get(x_129, 0); -lean_inc(x_130); -x_131 = lean_ctor_get(x_130, 1); -lean_inc(x_131); -x_132 = lean_ctor_get(x_129, 1); -lean_inc(x_132); -if (lean_is_exclusive(x_129)) { - lean_ctor_release(x_129, 0); - lean_ctor_release(x_129, 1); - x_133 = x_129; -} else { - lean_dec_ref(x_129); - x_133 = lean_box(0); +lean_object* x_39; lean_object* x_40; uint8_t x_41; +x_39 = lean_ctor_get(x_34, 1); +lean_inc(x_39); +lean_dec(x_34); +x_40 = lean_st_ref_get(x_9, x_39); +x_41 = !lean_is_exclusive(x_40); +if (x_41 == 0) +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_42 = lean_ctor_get(x_40, 0); +x_43 = lean_ctor_get(x_40, 1); +x_44 = lean_ctor_get(x_42, 0); +lean_inc(x_44); +lean_dec(x_42); +x_45 = lean_array_to_list(x_44); +x_46 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_46, 0, x_45); +x_47 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_47, 0, x_46); +x_48 = lean_box(0); +lean_ctor_set(x_40, 1, x_48); +lean_ctor_set(x_40, 0, x_47); +x_49 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_49, 0, x_40); +x_19 = x_49; +x_20 = x_43; +goto block_27; } -x_134 = lean_ctor_get(x_130, 0); -lean_inc(x_134); -if (lean_is_exclusive(x_130)) { - lean_ctor_release(x_130, 0); - lean_ctor_release(x_130, 1); - x_135 = x_130; -} else { - lean_dec_ref(x_130); - x_135 = lean_box(0); +else +{ +lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; +x_50 = lean_ctor_get(x_40, 0); +x_51 = lean_ctor_get(x_40, 1); +lean_inc(x_51); +lean_inc(x_50); +lean_dec(x_40); +x_52 = lean_ctor_get(x_50, 0); +lean_inc(x_52); +lean_dec(x_50); +x_53 = lean_array_to_list(x_52); +x_54 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_54, 0, x_53); +x_55 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_55, 0, x_54); +x_56 = lean_box(0); +x_57 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_57, 0, x_55); +lean_ctor_set(x_57, 1, x_56); +x_58 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_58, 0, x_57); +x_19 = x_58; +x_20 = x_51; +goto block_27; } -x_136 = lean_ctor_get(x_131, 0); -lean_inc(x_136); -x_137 = lean_ctor_get(x_131, 1); -lean_inc(x_137); -if (lean_is_exclusive(x_131)) { - lean_ctor_release(x_131, 0); - lean_ctor_release(x_131, 1); - x_138 = x_131; -} else { - lean_dec_ref(x_131); - x_138 = lean_box(0); } -x_139 = lean_unsigned_to_nat(0u); -x_140 = l_Lean_RBNode_fold___at_Lean_RBMap_size___spec__1___rarg(x_139, x_134); -x_141 = lean_nat_dec_eq(x_140, x_1); -lean_dec(x_140); -if (x_141 == 0) -{ -lean_object* x_142; uint8_t x_143; lean_object* x_144; lean_object* x_145; -lean_dec(x_138); -lean_dec(x_135); -lean_dec(x_133); -x_142 = lean_box(0); -x_143 = lean_unbox(x_136); -lean_dec(x_136); -lean_inc(x_5); -x_144 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___lambda__1(x_143, x_137, x_134, x_5, x_142, x_7, x_8, x_9, x_10, x_132); -x_145 = lean_ctor_get(x_144, 0); -lean_inc(x_145); -if (lean_obj_tag(x_145) == 0) +} +else { -lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; +uint8_t x_59; +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_5); -lean_dec(x_4); lean_dec(x_3); -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; -} else { - lean_dec_ref(x_144); - x_147 = lean_box(0); -} -x_148 = lean_ctor_get(x_145, 0); -lean_inc(x_148); -lean_dec(x_145); -if (lean_is_scalar(x_147)) { - x_149 = lean_alloc_ctor(0, 2, 0); -} else { - x_149 = x_147; -} -lean_ctor_set(x_149, 0, x_148); -lean_ctor_set(x_149, 1, x_146); -return x_149; +x_59 = !lean_is_exclusive(x_32); +if (x_59 == 0) +{ +return x_32; } else { -lean_object* x_150; lean_object* x_151; -x_150 = lean_ctor_get(x_144, 1); -lean_inc(x_150); -lean_dec(x_144); -x_151 = lean_ctor_get(x_145, 0); -lean_inc(x_151); -lean_dec(x_145); -x_6 = x_151; -x_11 = x_150; -goto _start; +lean_object* x_60; lean_object* x_61; lean_object* x_62; +x_60 = lean_ctor_get(x_32, 0); +x_61 = lean_ctor_get(x_32, 1); +lean_inc(x_61); +lean_inc(x_60); +lean_dec(x_32); +x_62 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_62, 0, x_60); +lean_ctor_set(x_62, 1, x_61); +return x_62; +} } } else { -lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; +uint8_t x_63; +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_5); -lean_dec(x_4); lean_dec(x_3); -if (lean_is_scalar(x_138)) { - x_153 = lean_alloc_ctor(0, 2, 0); -} else { - x_153 = x_138; -} -lean_ctor_set(x_153, 0, x_136); -lean_ctor_set(x_153, 1, x_137); -if (lean_is_scalar(x_135)) { - x_154 = lean_alloc_ctor(0, 2, 0); -} else { - x_154 = x_135; -} -lean_ctor_set(x_154, 0, x_134); -lean_ctor_set(x_154, 1, x_153); -x_155 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___closed__1; -x_156 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_156, 0, x_155); -lean_ctor_set(x_156, 1, x_154); -if (lean_is_scalar(x_133)) { - x_157 = lean_alloc_ctor(0, 2, 0); -} else { - x_157 = x_133; +x_63 = !lean_is_exclusive(x_29); +if (x_63 == 0) +{ +return x_29; } -lean_ctor_set(x_157, 0, x_156); -lean_ctor_set(x_157, 1, x_132); -return x_157; +else +{ +lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_64 = lean_ctor_get(x_29, 0); +x_65 = lean_ctor_get(x_29, 1); +lean_inc(x_65); +lean_inc(x_64); +lean_dec(x_29); +x_66 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_66, 0, x_64); +lean_ctor_set(x_66, 1, x_65); +return x_66; } } -else +block_27: { -lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; +if (lean_obj_tag(x_19) == 0) +{ +lean_object* x_21; lean_object* x_22; +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -lean_dec(x_4); lean_dec(x_3); -x_158 = lean_ctor_get(x_129, 0); -lean_inc(x_158); -x_159 = lean_ctor_get(x_129, 1); -lean_inc(x_159); -if (lean_is_exclusive(x_129)) { - lean_ctor_release(x_129, 0); - lean_ctor_release(x_129, 1); - x_160 = x_129; -} else { - lean_dec_ref(x_129); - x_160 = lean_box(0); +x_21 = lean_ctor_get(x_19, 0); +lean_inc(x_21); +lean_dec(x_19); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_20); +return x_22; } -if (lean_is_scalar(x_160)) { - x_161 = lean_alloc_ctor(1, 2, 0); -} else { - x_161 = x_160; +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_6, x_24); +x_6 = x_25; +x_7 = x_23; +x_15 = x_20; +goto _start; } -lean_ctor_set(x_161, 0, x_158); -lean_ctor_set(x_161, 1, x_159); -return x_161; } } } -else +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collectPatterns_x3f___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: { -lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; uint8_t x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; size_t x_171; size_t x_172; lean_object* x_173; -x_162 = lean_ctor_get(x_12, 1); -x_163 = lean_ctor_get(x_12, 0); -lean_inc(x_162); -lean_inc(x_163); -lean_dec(x_12); -x_164 = lean_ctor_get(x_162, 1); -lean_inc(x_164); -if (lean_is_exclusive(x_162)) { - lean_ctor_release(x_162, 0); - lean_ctor_release(x_162, 1); - x_165 = x_162; -} else { - lean_dec_ref(x_162); - x_165 = lean_box(0); +lean_object* x_11; +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_1); +lean_ctor_set(x_11, 1, x_10); +return x_11; } -x_166 = lean_box(0); -x_167 = 0; -x_168 = lean_box(x_167); -if (lean_is_scalar(x_165)) { - x_169 = lean_alloc_ctor(0, 2, 0); -} else { - x_169 = x_165; } -lean_ctor_set(x_169, 0, x_168); -lean_ctor_set(x_169, 1, x_164); -x_170 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_170, 0, x_163); -lean_ctor_set(x_170, 1, x_169); -x_171 = lean_array_size(x_2); -x_172 = 0; -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_4); -lean_inc(x_3); -x_173 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__14___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__15(x_2, x_3, x_4, x_166, x_2, x_171, x_172, x_170, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_173) == 0) +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collectPatterns_x3f___closed__1() { +_start: { -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; uint8_t x_185; -x_174 = lean_ctor_get(x_173, 0); -lean_inc(x_174); -x_175 = lean_ctor_get(x_174, 1); -lean_inc(x_175); -x_176 = lean_ctor_get(x_173, 1); -lean_inc(x_176); -if (lean_is_exclusive(x_173)) { - lean_ctor_release(x_173, 0); - lean_ctor_release(x_173, 1); - x_177 = x_173; -} else { - lean_dec_ref(x_173); - x_177 = lean_box(0); +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_NormalizePattern_main___closed__1; +x_2 = 0; +x_3 = lean_alloc_ctor(0, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; } -x_178 = lean_ctor_get(x_174, 0); -lean_inc(x_178); -if (lean_is_exclusive(x_174)) { - lean_ctor_release(x_174, 0); - lean_ctor_release(x_174, 1); - x_179 = x_174; +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collectPatterns_x3f(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; lean_object* x_10; size_t x_11; size_t x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; +x_9 = lean_box(0); +x_10 = lean_box(0); +x_11 = lean_array_size(x_3); +x_12 = 0; +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_1); +lean_ctor_set(x_13, 1, x_2); +x_14 = l_Lean_Meta_Grind_NormalizePattern_main___closed__5; +x_15 = lean_st_mk_ref(x_14, 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); +if (lean_is_exclusive(x_15)) { + lean_ctor_release(x_15, 0); + lean_ctor_release(x_15, 1); + x_18 = x_15; } else { - lean_dec_ref(x_174); - x_179 = lean_box(0); + lean_dec_ref(x_15); + x_18 = lean_box(0); +} +x_49 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collectPatterns_x3f___closed__1; +x_50 = lean_st_mk_ref(x_49, x_17); +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_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___closed__1; +lean_inc(x_16); +lean_inc(x_51); +x_54 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collectPatterns_x3f___spec__1(x_3, x_9, x_53, x_3, x_11, x_12, x_53, x_13, x_51, x_16, x_4, x_5, x_6, x_7, x_52); +if (lean_obj_tag(x_54) == 0) +{ +lean_object* x_55; lean_object* x_56; +x_55 = lean_ctor_get(x_54, 0); +lean_inc(x_55); +x_56 = lean_ctor_get(x_55, 0); +lean_inc(x_56); +lean_dec(x_55); +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_54, 1); +lean_inc(x_57); +lean_dec(x_54); +x_58 = lean_st_ref_get(x_51, x_57); +lean_dec(x_51); +x_59 = lean_ctor_get(x_58, 1); +lean_inc(x_59); +lean_dec(x_58); +x_19 = x_10; +x_20 = x_59; +goto block_48; +} +else +{ +lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; +x_60 = lean_ctor_get(x_54, 1); +lean_inc(x_60); +lean_dec(x_54); +x_61 = lean_ctor_get(x_56, 0); +lean_inc(x_61); +lean_dec(x_56); +x_62 = lean_st_ref_get(x_51, x_60); +lean_dec(x_51); +x_63 = lean_ctor_get(x_62, 1); +lean_inc(x_63); +lean_dec(x_62); +x_19 = x_61; +x_20 = x_63; +goto block_48; +} +} +else +{ +uint8_t x_64; +lean_dec(x_51); +lean_dec(x_18); +lean_dec(x_16); +x_64 = !lean_is_exclusive(x_54); +if (x_64 == 0) +{ +return x_54; } -x_180 = lean_ctor_get(x_175, 0); -lean_inc(x_180); -x_181 = lean_ctor_get(x_175, 1); -lean_inc(x_181); -if (lean_is_exclusive(x_175)) { - lean_ctor_release(x_175, 0); - lean_ctor_release(x_175, 1); - x_182 = x_175; -} else { - lean_dec_ref(x_175); - x_182 = lean_box(0); +else +{ +lean_object* x_65; lean_object* x_66; lean_object* x_67; +x_65 = lean_ctor_get(x_54, 0); +x_66 = lean_ctor_get(x_54, 1); +lean_inc(x_66); +lean_inc(x_65); +lean_dec(x_54); +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; } -x_183 = lean_unsigned_to_nat(0u); -x_184 = l_Lean_RBNode_fold___at_Lean_RBMap_size___spec__1___rarg(x_183, x_178); -x_185 = lean_nat_dec_eq(x_184, x_1); -lean_dec(x_184); -if (x_185 == 0) +} +block_48: { -lean_object* x_186; uint8_t x_187; lean_object* x_188; lean_object* x_189; -lean_dec(x_182); -lean_dec(x_179); -lean_dec(x_177); -x_186 = lean_box(0); -x_187 = lean_unbox(x_180); -lean_dec(x_180); -lean_inc(x_5); -x_188 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___lambda__1(x_187, x_181, x_178, x_5, x_186, x_7, x_8, x_9, x_10, x_176); -x_189 = lean_ctor_get(x_188, 0); -lean_inc(x_189); -if (lean_obj_tag(x_189) == 0) +lean_object* x_21; +x_21 = lean_st_ref_get(x_16, x_20); +lean_dec(x_16); +if (lean_obj_tag(x_19) == 0) { -lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; -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_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); +uint8_t x_22; +lean_dec(x_18); +x_22 = !lean_is_exclusive(x_21); +if (x_22 == 0) +{ +lean_object* x_23; +x_23 = lean_ctor_get(x_21, 0); +lean_dec(x_23); +lean_ctor_set(x_21, 0, x_10); +return x_21; } -x_192 = lean_ctor_get(x_189, 0); -lean_inc(x_192); -lean_dec(x_189); -if (lean_is_scalar(x_191)) { - x_193 = lean_alloc_ctor(0, 2, 0); +else +{ +lean_object* x_24; lean_object* x_25; +x_24 = lean_ctor_get(x_21, 1); +lean_inc(x_24); +lean_dec(x_21); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_10); +lean_ctor_set(x_25, 1, x_24); +return x_25; +} +} +else +{ +uint8_t x_26; +x_26 = !lean_is_exclusive(x_21); +if (x_26 == 0) +{ +uint8_t x_27; +x_27 = !lean_is_exclusive(x_19); +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_21, 0); +x_29 = lean_ctor_get(x_19, 0); +x_30 = lean_ctor_get(x_28, 0); +lean_inc(x_30); +lean_dec(x_28); +x_31 = lean_array_to_list(x_30); +if (lean_is_scalar(x_18)) { + x_32 = lean_alloc_ctor(0, 2, 0); } else { - x_193 = x_191; + x_32 = x_18; } -lean_ctor_set(x_193, 0, x_192); -lean_ctor_set(x_193, 1, x_190); -return x_193; +lean_ctor_set(x_32, 0, x_29); +lean_ctor_set(x_32, 1, x_31); +lean_ctor_set(x_19, 0, x_32); +lean_ctor_set(x_21, 0, x_19); +return x_21; } else { -lean_object* x_194; lean_object* x_195; -x_194 = lean_ctor_get(x_188, 1); -lean_inc(x_194); -lean_dec(x_188); -x_195 = lean_ctor_get(x_189, 0); -lean_inc(x_195); -lean_dec(x_189); -x_6 = x_195; -x_11 = x_194; -goto _start; +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_21, 0); +x_34 = lean_ctor_get(x_19, 0); +lean_inc(x_34); +lean_dec(x_19); +x_35 = lean_ctor_get(x_33, 0); +lean_inc(x_35); +lean_dec(x_33); +x_36 = lean_array_to_list(x_35); +if (lean_is_scalar(x_18)) { + x_37 = lean_alloc_ctor(0, 2, 0); +} else { + x_37 = x_18; +} +lean_ctor_set(x_37, 0, x_34); +lean_ctor_set(x_37, 1, x_36); +x_38 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_38, 0, x_37); +lean_ctor_set(x_21, 0, x_38); +return x_21; } } else { -lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; -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); -if (lean_is_scalar(x_182)) { - x_197 = lean_alloc_ctor(0, 2, 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; +x_39 = lean_ctor_get(x_21, 0); +x_40 = lean_ctor_get(x_21, 1); +lean_inc(x_40); +lean_inc(x_39); +lean_dec(x_21); +x_41 = lean_ctor_get(x_19, 0); +lean_inc(x_41); +if (lean_is_exclusive(x_19)) { + lean_ctor_release(x_19, 0); + x_42 = x_19; } else { - x_197 = x_182; + lean_dec_ref(x_19); + x_42 = lean_box(0); } -lean_ctor_set(x_197, 0, x_180); -lean_ctor_set(x_197, 1, x_181); -if (lean_is_scalar(x_179)) { - x_198 = lean_alloc_ctor(0, 2, 0); +x_43 = lean_ctor_get(x_39, 0); +lean_inc(x_43); +lean_dec(x_39); +x_44 = lean_array_to_list(x_43); +if (lean_is_scalar(x_18)) { + x_45 = lean_alloc_ctor(0, 2, 0); } else { - x_198 = x_179; + x_45 = x_18; } -lean_ctor_set(x_198, 0, x_178); -lean_ctor_set(x_198, 1, x_197); -x_199 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___closed__1; -x_200 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_200, 0, x_199); -lean_ctor_set(x_200, 1, x_198); -if (lean_is_scalar(x_177)) { - x_201 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_45, 0, x_41); +lean_ctor_set(x_45, 1, x_44); +if (lean_is_scalar(x_42)) { + x_46 = lean_alloc_ctor(1, 1, 0); } else { - x_201 = x_177; + x_46 = x_42; } -lean_ctor_set(x_201, 0, x_200); -lean_ctor_set(x_201, 1, x_176); -return x_201; +lean_ctor_set(x_46, 0, x_45); +x_47 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_47, 0, x_46); +lean_ctor_set(x_47, 1, x_40); +return x_47; } } -else +} +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collectPatterns_x3f___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, 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_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; -lean_dec(x_10); +size_t x_16; size_t x_17; lean_object* x_18; +x_16 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_17 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_18 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collectPatterns_x3f___spec__1(x_1, x_2, x_3, x_4, x_16, x_17, 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_2); +lean_dec(x_1); +return x_18; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collectPatterns_x3f___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collectPatterns_x3f___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); +lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_202 = lean_ctor_get(x_173, 0); -lean_inc(x_202); -x_203 = lean_ctor_get(x_173, 1); -lean_inc(x_203); -if (lean_is_exclusive(x_173)) { - lean_ctor_release(x_173, 0); - lean_ctor_release(x_173, 1); - x_204 = x_173; -} else { - lean_dec_ref(x_173); - x_204 = lean_box(0); +lean_dec(x_2); +return x_11; } -if (lean_is_scalar(x_204)) { - x_205 = lean_alloc_ctor(1, 2, 0); -} else { - x_205 = x_204; } -lean_ctor_set(x_205, 0, x_202); -lean_ctor_set(x_205, 1, x_203); -return x_205; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collectPatterns_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, lean_object* x_8) { +_start: +{ +lean_object* x_9; +x_9 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collectPatterns_x3f(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_Meta_Grind_mkEMatchTheoremWithKind_x3f_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, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_13 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_13, 0, x_1); +lean_ctor_set(x_13, 1, x_2); +lean_ctor_set(x_13, 2, x_3); +lean_ctor_set(x_13, 3, x_4); +lean_ctor_set(x_13, 4, x_5); +lean_ctor_set(x_13, 5, x_6); +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_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___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_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f_go(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -lean_object* x_15; uint8_t x_16; -x_15 = lean_ctor_get(x_5, 1); -x_16 = lean_nat_dec_lt(x_7, x_15); -if (x_16 == 0) -{ -lean_object* x_17; -lean_dec(x_7); -x_17 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_17, 0, x_6); -lean_ctor_set(x_17, 1, x_14); -return x_17; -} -else +lean_object* x_11; +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_4); +lean_inc(x_3); +x_11 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collectPatterns_x3f(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_18; lean_object* x_19; lean_object* x_20; -x_18 = lean_array_fget(x_1, x_7); -x_19 = l_Lean_Expr_fvarId_x21(x_18); -lean_dec(x_18); -x_20 = l_Lean_RBNode_findCore___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__2(x_3, x_19); -lean_dec(x_19); -if (lean_obj_tag(x_20) == 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_21; lean_object* x_22; lean_object* x_23; -lean_inc(x_7); -x_21 = lean_array_push(x_6, x_7); -x_22 = lean_ctor_get(x_5, 2); -x_23 = lean_nat_add(x_7, x_22); +uint8_t x_13; +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_7); -x_6 = x_21; -x_7 = x_23; -x_8 = lean_box(0); -x_9 = lean_box(0); -goto _start; +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_13 = !lean_is_exclusive(x_11); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; +x_14 = lean_ctor_get(x_11, 0); +lean_dec(x_14); +x_15 = lean_box(0); +lean_ctor_set(x_11, 0, x_15); +return x_11; } else { -lean_object* x_25; lean_object* x_26; -lean_dec(x_20); -x_25 = lean_ctor_get(x_5, 2); -x_26 = lean_nat_add(x_7, x_25); -lean_dec(x_7); -x_7 = x_26; -x_8 = lean_box(0); -x_9 = lean_box(0); -goto _start; -} -} +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_11, 1); +lean_inc(x_16); +lean_dec(x_11); +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_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__18___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__19(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { -_start: +else { -lean_object* x_14; uint8_t x_15; -x_14 = lean_ctor_get(x_4, 1); -x_15 = lean_nat_dec_lt(x_6, x_14); -if (x_15 == 0) +lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_19 = lean_ctor_get(x_12, 0); +lean_inc(x_19); +lean_dec(x_12); +x_20 = lean_ctor_get(x_11, 1); +lean_inc(x_20); +lean_dec(x_11); +x_21 = !lean_is_exclusive(x_19); +if (x_21 == 0) { -lean_object* x_16; +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_19, 0); +x_23 = lean_ctor_get(x_19, 1); +x_24 = lean_array_get_size(x_4); +lean_dec(x_4); +x_25 = l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__3; +x_26 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_25, x_6, x_7, x_8, x_9, x_20); +x_27 = lean_ctor_get(x_26, 0); +lean_inc(x_27); +x_28 = lean_unbox(x_27); +lean_dec(x_27); +if (x_28 == 0) +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; +lean_free_object(x_19); +x_29 = lean_ctor_get(x_26, 1); +lean_inc(x_29); +lean_dec(x_26); +x_30 = lean_box(0); +x_31 = l_Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f_go___lambda__1(x_2, x_3, x_24, x_22, x_23, x_1, x_30, x_6, x_7, x_8, x_9, x_29); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); lean_dec(x_6); -x_16 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_16, 0, x_5); -lean_ctor_set(x_16, 1, x_13); -return x_16; +return x_31; } else { -lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_17 = lean_array_fget(x_1, x_6); -x_18 = l_Lean_Expr_fvarId_x21(x_17); -lean_dec(x_17); -x_19 = l_Lean_RBNode_findCore___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__2(x_2, x_18); -lean_dec(x_18); -if (lean_obj_tag(x_19) == 0) +uint8_t x_32; +x_32 = !lean_is_exclusive(x_26); +if (x_32 == 0) { -lean_object* x_20; lean_object* x_21; lean_object* x_22; -lean_inc(x_6); -x_20 = lean_array_push(x_5, x_6); -x_21 = lean_ctor_get(x_4, 2); -x_22 = lean_nat_add(x_6, x_21); +lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_33 = lean_ctor_get(x_26, 1); +x_34 = lean_ctor_get(x_26, 0); +lean_dec(x_34); +lean_inc(x_1); +x_35 = l_Lean_Meta_Grind_Origin_pp___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__3(x_1, x_6, x_7, x_8, x_9, x_33); +if (lean_obj_tag(x_35) == 0) +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; 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_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_Grind_ppPattern___closed__4; +lean_ctor_set_tag(x_26, 7); +lean_ctor_set(x_26, 1, x_36); +lean_ctor_set(x_26, 0, x_38); +x_39 = l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__5; +lean_ctor_set_tag(x_19, 7); +lean_ctor_set(x_19, 1, x_39); +lean_ctor_set(x_19, 0, x_26); +x_40 = lean_box(0); +lean_inc(x_22); +x_41 = l_List_mapTR_loop___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__1(x_22, x_40); +x_42 = l_List_mapTR_loop___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__2(x_41, x_40); +x_43 = l_Lean_MessageData_ofList(x_42); +x_44 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_44, 0, x_19); +lean_ctor_set(x_44, 1, x_43); +x_45 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_45, 0, x_44); +lean_ctor_set(x_45, 1, x_38); +x_46 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_25, x_45, x_6, x_7, x_8, x_9, x_37); +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_Meta_Grind_mkEMatchTheoremWithKind_x3f_go___lambda__1(x_2, x_3, x_24, x_22, x_23, x_1, x_47, x_6, x_7, x_8, x_9, x_48); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); lean_dec(x_6); -x_5 = x_20; -x_6 = x_22; -x_7 = lean_box(0); -x_8 = lean_box(0); -goto _start; +lean_dec(x_47); +return x_49; } else { -lean_object* x_24; lean_object* x_25; -lean_dec(x_19); -x_24 = lean_ctor_get(x_4, 2); -x_25 = lean_nat_add(x_6, x_24); +uint8_t x_50; +lean_free_object(x_26); +lean_dec(x_24); +lean_free_object(x_19); +lean_dec(x_23); +lean_dec(x_22); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); lean_dec(x_6); -x_6 = x_25; -x_7 = lean_box(0); -x_8 = lean_box(0); -goto _start; +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_50 = !lean_is_exclusive(x_35); +if (x_50 == 0) +{ +return x_35; } +else +{ +lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_51 = lean_ctor_get(x_35, 0); +x_52 = lean_ctor_get(x_35, 1); +lean_inc(x_52); +lean_inc(x_51); +lean_dec(x_35); +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_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__20(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { -_start: +else { -uint8_t x_5; -x_5 = lean_usize_dec_eq(x_2, x_3); -if (x_5 == 0) +lean_object* x_54; lean_object* x_55; +x_54 = lean_ctor_get(x_26, 1); +lean_inc(x_54); +lean_dec(x_26); +lean_inc(x_1); +x_55 = l_Lean_Meta_Grind_Origin_pp___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__3(x_1, x_6, x_7, x_8, x_9, x_54); +if (lean_obj_tag(x_55) == 0) { -size_t x_6; size_t x_7; lean_object* x_8; lean_object* x_9; -x_6 = 1; -x_7 = lean_usize_sub(x_2, x_6); -x_8 = lean_array_uget(x_1, x_7); -x_9 = l_Std_DHashMap_Internal_AssocList_foldrM___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__2(x_4, x_8); +lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; 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_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_Meta_Grind_ppPattern___closed__4; +x_59 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_59, 0, x_58); +lean_ctor_set(x_59, 1, x_56); +x_60 = l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__5; +lean_ctor_set_tag(x_19, 7); +lean_ctor_set(x_19, 1, x_60); +lean_ctor_set(x_19, 0, x_59); +x_61 = lean_box(0); +lean_inc(x_22); +x_62 = l_List_mapTR_loop___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__1(x_22, x_61); +x_63 = l_List_mapTR_loop___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__2(x_62, x_61); +x_64 = l_Lean_MessageData_ofList(x_63); +x_65 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_65, 0, x_19); +lean_ctor_set(x_65, 1, x_64); +x_66 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_66, 0, x_65); +lean_ctor_set(x_66, 1, x_58); +x_67 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_25, x_66, x_6, x_7, x_8, x_9, x_57); +x_68 = lean_ctor_get(x_67, 0); +lean_inc(x_68); +x_69 = lean_ctor_get(x_67, 1); +lean_inc(x_69); +lean_dec(x_67); +x_70 = l_Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f_go___lambda__1(x_2, x_3, x_24, x_22, x_23, x_1, x_68, x_6, x_7, x_8, x_9, x_69); +lean_dec(x_9); lean_dec(x_8); -lean_dec(x_4); -x_2 = x_7; -x_4 = x_9; -goto _start; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_68); +return x_70; } else { -return x_4; -} +lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; +lean_dec(x_24); +lean_free_object(x_19); +lean_dec(x_23); +lean_dec(x_22); +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_71 = lean_ctor_get(x_55, 0); +lean_inc(x_71); +x_72 = lean_ctor_get(x_55, 1); +lean_inc(x_72); +if (lean_is_exclusive(x_55)) { + lean_ctor_release(x_55, 0); + lean_ctor_release(x_55, 1); + x_73 = x_55; +} else { + lean_dec_ref(x_55); + 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_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; -x_11 = lean_array_mk(x_1); -x_12 = lean_unsigned_to_nat(0u); -x_13 = lean_unsigned_to_nat(1u); -x_14 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_14, 0, x_12); -lean_ctor_set(x_14, 1, x_2); -lean_ctor_set(x_14, 2, x_13); -x_15 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__18___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__19(x_3, x_4, x_14, x_14, x_11, x_12, lean_box(0), lean_box(0), x_6, x_7, x_8, x_9, x_10); -lean_dec(x_14); -x_16 = !lean_is_exclusive(x_15); -if (x_16 == 0) -{ -lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_17 = lean_ctor_get(x_15, 0); -x_18 = lean_array_to_list(x_17); -x_19 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_19, 0, x_18); -lean_ctor_set(x_15, 0, x_19); -return x_15; +lean_ctor_set(x_74, 0, x_71); +lean_ctor_set(x_74, 1, x_72); +return x_74; } -else -{ -lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_20 = lean_ctor_get(x_15, 0); -x_21 = lean_ctor_get(x_15, 1); -lean_inc(x_21); -lean_inc(x_20); -lean_dec(x_15); -x_22 = lean_array_to_list(x_20); -x_23 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_23, 0, x_22); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_23); -lean_ctor_set(x_24, 1, x_21); -return x_24; } } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { -_start: -{ -lean_object* x_14; 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_14 = lean_box(0); -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_1); -x_18 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_18, 0, x_7); -lean_ctor_set(x_18, 1, 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); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_4); -x_20 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__17(x_2, x_3, x_4, x_5, x_14, x_19, x_9, x_10, x_11, x_12, x_13); -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_21, 1); -lean_inc(x_22); -x_23 = lean_ctor_get(x_21, 0); -lean_inc(x_23); -lean_dec(x_21); -if (lean_obj_tag(x_23) == 0) +else { -lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_24 = lean_ctor_get(x_20, 1); -lean_inc(x_24); -lean_dec(x_20); -x_25 = lean_ctor_get(x_22, 0); -lean_inc(x_25); -lean_dec(x_22); -x_26 = lean_box(0); -x_27 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__1(x_4, x_6, x_3, x_25, x_26, x_9, x_10, x_11, x_12, x_24); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); +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; +x_75 = lean_ctor_get(x_19, 0); +x_76 = lean_ctor_get(x_19, 1); +lean_inc(x_76); +lean_inc(x_75); +lean_dec(x_19); +x_77 = lean_array_get_size(x_4); +lean_dec(x_4); +x_78 = l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__3; +x_79 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_78, x_6, x_7, x_8, x_9, x_20); +x_80 = lean_ctor_get(x_79, 0); +lean_inc(x_80); +x_81 = lean_unbox(x_80); +lean_dec(x_80); +if (x_81 == 0) +{ +lean_object* x_82; lean_object* x_83; lean_object* x_84; +x_82 = lean_ctor_get(x_79, 1); +lean_inc(x_82); +lean_dec(x_79); +x_83 = lean_box(0); +x_84 = l_Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f_go___lambda__1(x_2, x_3, x_77, x_75, x_76, x_1, x_83, x_6, x_7, x_8, x_9, x_82); lean_dec(x_9); -lean_dec(x_25); -return x_27; +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +return x_84; } else { -uint8_t x_28; -lean_dec(x_22); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); +lean_object* x_85; lean_object* x_86; lean_object* x_87; +x_85 = lean_ctor_get(x_79, 1); +lean_inc(x_85); +if (lean_is_exclusive(x_79)) { + lean_ctor_release(x_79, 0); + lean_ctor_release(x_79, 1); + x_86 = x_79; +} else { + lean_dec_ref(x_79); + x_86 = lean_box(0); +} +lean_inc(x_1); +x_87 = l_Lean_Meta_Grind_Origin_pp___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__3(x_1, x_6, x_7, x_8, x_9, x_85); +if (lean_obj_tag(x_87) == 0) +{ +lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; 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_88 = lean_ctor_get(x_87, 0); +lean_inc(x_88); +x_89 = lean_ctor_get(x_87, 1); +lean_inc(x_89); +lean_dec(x_87); +x_90 = l_Lean_Meta_Grind_ppPattern___closed__4; +if (lean_is_scalar(x_86)) { + x_91 = lean_alloc_ctor(7, 2, 0); +} else { + x_91 = x_86; + lean_ctor_set_tag(x_91, 7); +} +lean_ctor_set(x_91, 0, x_90); +lean_ctor_set(x_91, 1, x_88); +x_92 = l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__5; +x_93 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_93, 0, x_91); +lean_ctor_set(x_93, 1, x_92); +x_94 = lean_box(0); +lean_inc(x_75); +x_95 = l_List_mapTR_loop___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__1(x_75, x_94); +x_96 = l_List_mapTR_loop___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__2(x_95, x_94); +x_97 = l_Lean_MessageData_ofList(x_96); +x_98 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_98, 0, x_93); +lean_ctor_set(x_98, 1, x_97); +x_99 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_99, 0, x_98); +lean_ctor_set(x_99, 1, x_90); +x_100 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_78, x_99, x_6, x_7, x_8, x_9, x_89); +x_101 = lean_ctor_get(x_100, 0); +lean_inc(x_101); +x_102 = lean_ctor_get(x_100, 1); +lean_inc(x_102); +lean_dec(x_100); +x_103 = l_Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f_go___lambda__1(x_2, x_3, x_77, x_75, x_76, x_1, x_101, x_6, x_7, x_8, x_9, x_102); lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_4); -x_28 = !lean_is_exclusive(x_20); -if (x_28 == 0) -{ -lean_object* x_29; lean_object* x_30; -x_29 = lean_ctor_get(x_20, 0); -lean_dec(x_29); -x_30 = lean_ctor_get(x_23, 0); -lean_inc(x_30); -lean_dec(x_23); -lean_ctor_set(x_20, 0, x_30); -return x_20; +lean_dec(x_101); +return x_103; } else { -lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_31 = lean_ctor_get(x_20, 1); -lean_inc(x_31); -lean_dec(x_20); -x_32 = lean_ctor_get(x_23, 0); -lean_inc(x_32); -lean_dec(x_23); -x_33 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_33, 0, x_32); -lean_ctor_set(x_33, 1, x_31); -return x_33; +lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; +lean_dec(x_86); +lean_dec(x_77); +lean_dec(x_76); +lean_dec(x_75); +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_104 = lean_ctor_get(x_87, 0); +lean_inc(x_104); +x_105 = lean_ctor_get(x_87, 1); +lean_inc(x_105); +if (lean_is_exclusive(x_87)) { + lean_ctor_release(x_87, 0); + lean_ctor_release(x_87, 1); + x_106 = x_87; +} else { + lean_dec_ref(x_87); + 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; +} +} } } } else { -uint8_t x_34; -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); +uint8_t x_108; lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); lean_dec(x_6); lean_dec(x_4); -x_34 = !lean_is_exclusive(x_20); -if (x_34 == 0) +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_108 = !lean_is_exclusive(x_11); +if (x_108 == 0) { -return x_20; +return x_11; } else { -lean_object* x_35; lean_object* x_36; lean_object* x_37; -x_35 = lean_ctor_get(x_20, 0); -x_36 = lean_ctor_get(x_20, 1); -lean_inc(x_36); -lean_inc(x_35); -lean_dec(x_20); -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_109; lean_object* x_110; lean_object* x_111; +x_109 = lean_ctor_get(x_11, 0); +x_110 = lean_ctor_get(x_11, 1); +lean_inc(x_110); +lean_inc(x_109); +lean_dec(x_11); +x_111 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_111, 0, x_109); +lean_ctor_set(x_111, 1, x_110); +return x_111; } } } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3___closed__1() { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f_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, lean_object* x_11, lean_object* x_12) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("numParams == xs.size\n ", 25, 25); -return x_1; +lean_object* x_13; +x_13 = l_Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f_go___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +return x_13; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3___closed__2() { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f_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_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__3; -x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3___closed__1; -x_3 = lean_string_append(x_1, x_2); -return x_3; +lean_object* x_11; +x_11 = l_Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f_go(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; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3___closed__3() { +LEAN_EXPORT lean_object* l_panic___at_Lean_Meta_Grind_mkEMatchTheoremWithKind_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; lean_object* x_9; +x_7 = l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__1___closed__1; +x_8 = lean_panic_fn(x_7, x_1); +x_9 = lean_apply_5(x_8, x_2, x_3, x_4, x_5, x_6); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; +x_12 = l_Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f_go(x_1, x_2, x_3, x_4, x_5, x_7, x_8, x_9, x_10, x_11); +return x_12; +} +} +static lean_object* _init_l_Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("_private.Lean.Meta.Tactic.Grind.EMatchTheorem.0.Lean.Meta.Grind.checkCoverage", 77, 77); +x_1 = lean_mk_string_unchecked("Lean.Meta.Grind.mkEMatchTheoremWithKind\?", 40, 40); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3___closed__4() { +static lean_object* _init_l_Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_Meta_Grind_EMatchTheorems_insert___closed__1; -x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3___closed__3; -x_3 = lean_unsigned_to_nat(263u); -x_4 = lean_unsigned_to_nat(4u); -x_5 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3___closed__2; +x_2 = l_Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2___closed__1; +x_3 = lean_unsigned_to_nat(617u); +x_4 = lean_unsigned_to_nat(13u); +x_5 = l_Lean_Meta_Grind_EMatchTheorems_insert___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___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___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) { +static lean_object* _init_l_Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2___closed__3() { _start: { -lean_object* x_10; uint8_t x_11; -x_10 = lean_array_get_size(x_3); -x_11 = lean_nat_dec_eq(x_1, x_10); -if (x_11 == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("invalid `grind` forward theorem, theorem `", 42, 42); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2___closed__4() { +_start: { -lean_object* x_12; lean_object* x_13; -lean_dec(x_10); -lean_dec(x_3); -x_12 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3___closed__4; -x_13 = l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__1(x_12, x_5, x_6, x_7, x_8, x_9); -return x_13; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2___closed__3; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("` does not have propositional hypotheses", 40, 40); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2___closed__5; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2(uint8_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; +x_12 = lean_box(x_1); +switch (lean_obj_tag(x_12)) { +case 3: +{ +lean_object* x_13; +lean_dec(x_6); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +x_13 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getPropTypes(x_5, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_13) == 0) +{ +lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +x_16 = l_Array_isEmpty___rarg(x_14); +if (x_16 == 0) +{ +lean_object* x_17; +x_17 = l_Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f_go(x_2, x_3, x_4, x_5, x_14, x_7, x_8, x_9, x_10, x_15); +lean_dec(x_14); +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_48; uint8_t x_49; -x_14 = lean_box(0); -x_15 = lean_ctor_get(x_2, 1); -x_16 = lean_array_get_size(x_15); -lean_inc(x_3); -x_17 = lean_array_to_list(x_3); -x_18 = l_List_mapTR_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__3(x_17, x_14); -x_19 = l_Lean_RBTree_ofList___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__4(x_18); -x_20 = lean_box(0); -x_48 = lean_unsigned_to_nat(0u); -x_49 = lean_nat_dec_lt(x_48, x_16); -if (x_49 == 0) +lean_object* x_18; +lean_dec(x_14); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_18 = l_Lean_Meta_Grind_Origin_pp___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__3(x_2, x_7, x_8, x_9, x_10, x_15); +if (lean_obj_tag(x_18) == 0) { -lean_dec(x_16); -x_21 = x_14; -goto block_47; +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_18, 1); +lean_inc(x_20); +lean_dec(x_18); +x_21 = l_Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2___closed__4; +x_22 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_19); +x_23 = l_Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2___closed__6; +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_Meta_InferType_0__Lean_Meta_inferProjType___spec__1(x_24, x_7, x_8, x_9, x_10, x_20); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +x_26 = !lean_is_exclusive(x_25); +if (x_26 == 0) +{ +return x_25; } else { -size_t x_50; size_t x_51; lean_object* x_52; -x_50 = lean_usize_of_nat(x_16); -lean_dec(x_16); -x_51 = 0; -x_52 = l_Array_foldrMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__20(x_15, x_50, x_51, x_14); -x_21 = x_52; -goto block_47; +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; } -block_47: -{ -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = l_List_mapTR_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__5(x_1, x_3, x_10, x_21, x_14); -lean_inc(x_22); -x_23 = l_Lean_RBTree_ofList___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__4(x_22); -lean_inc(x_5); -lean_inc(x_23); -lean_inc(x_22); -x_24 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__8___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__9(x_14, x_19, x_20, x_22, x_22, x_22, x_23, lean_box(0), x_5, x_6, x_7, x_8, x_9); -lean_dec(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; lean_object* x_28; lean_object* x_29; uint8_t x_30; -x_26 = lean_ctor_get(x_24, 0); -x_27 = lean_ctor_get(x_24, 1); -x_28 = lean_unsigned_to_nat(0u); -x_29 = l_Lean_RBNode_fold___at_Lean_RBMap_size___spec__1___rarg(x_28, x_26); -x_30 = lean_nat_dec_eq(x_29, x_1); -lean_dec(x_29); -if (x_30 == 0) -{ -lean_object* x_31; lean_object* x_32; -lean_free_object(x_24); -x_31 = lean_box(0); -x_32 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__2(x_23, x_1, x_3, x_14, x_19, x_10, x_26, x_31, x_5, x_6, x_7, x_8, x_27); -lean_dec(x_3); -return x_32; } else { -lean_object* x_33; -lean_dec(x_26); -lean_dec(x_23); -lean_dec(x_19); +uint8_t x_30; lean_dec(x_10); +lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -x_33 = lean_box(0); -lean_ctor_set(x_24, 0, x_33); -return x_24; -} +x_30 = !lean_is_exclusive(x_18); +if (x_30 == 0) +{ +return x_18; } else { -lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; uint8_t x_38; -x_34 = lean_ctor_get(x_24, 0); -x_35 = lean_ctor_get(x_24, 1); -lean_inc(x_35); -lean_inc(x_34); -lean_dec(x_24); -x_36 = lean_unsigned_to_nat(0u); -x_37 = l_Lean_RBNode_fold___at_Lean_RBMap_size___spec__1___rarg(x_36, x_34); -x_38 = lean_nat_dec_eq(x_37, x_1); -lean_dec(x_37); -if (x_38 == 0) -{ -lean_object* x_39; lean_object* x_40; -x_39 = lean_box(0); -x_40 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__2(x_23, x_1, x_3, x_14, x_19, x_10, x_34, x_39, x_5, x_6, x_7, x_8, x_35); -lean_dec(x_3); -return x_40; +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_18, 0); +x_32 = lean_ctor_get(x_18, 1); +lean_inc(x_32); +lean_inc(x_31); +lean_dec(x_18); +x_33 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_33, 0, x_31); +lean_ctor_set(x_33, 1, x_32); +return x_33; +} +} +} } else { -lean_object* x_41; lean_object* x_42; -lean_dec(x_34); -lean_dec(x_23); -lean_dec(x_19); +uint8_t x_34; 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_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_35); -return x_42; +lean_dec(x_2); +x_34 = !lean_is_exclusive(x_13); +if (x_34 == 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_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; +} } } +case 4: +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_38 = lean_box(0); +x_39 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_39, 0, x_6); +lean_ctor_set(x_39, 1, x_38); +x_40 = lean_array_mk(x_39); +x_41 = l_Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f_go(x_2, x_3, x_4, x_5, x_40, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_40); +return x_41; +} +case 5: +{ +lean_object* x_42; +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +x_42 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getPropTypes(x_5, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_42) == 0) +{ +lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_43 = lean_ctor_get(x_42, 0); +lean_inc(x_43); +x_44 = lean_ctor_get(x_42, 1); +lean_inc(x_44); +lean_dec(x_42); +x_45 = lean_box(0); +x_46 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_46, 0, x_6); +lean_ctor_set(x_46, 1, x_45); +x_47 = lean_array_mk(x_46); +x_48 = l_Array_append___rarg(x_47, x_43); +lean_dec(x_43); +x_49 = l_Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f_go(x_2, x_3, x_4, x_5, x_48, x_7, x_8, x_9, x_10, x_44); +lean_dec(x_48); +return x_49; } else { -uint8_t x_43; -lean_dec(x_23); -lean_dec(x_19); +uint8_t x_50; lean_dec(x_10); +lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); +lean_dec(x_4); lean_dec(x_3); -x_43 = !lean_is_exclusive(x_24); -if (x_43 == 0) +lean_dec(x_2); +x_50 = !lean_is_exclusive(x_42); +if (x_50 == 0) { -return x_24; +return x_42; } else { -lean_object* x_44; lean_object* x_45; lean_object* x_46; -x_44 = lean_ctor_get(x_24, 0); -x_45 = lean_ctor_get(x_24, 1); -lean_inc(x_45); -lean_inc(x_44); -lean_dec(x_24); -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_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_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: +default: { -lean_object* x_10; +lean_object* x_54; lean_object* x_55; +lean_dec(x_12); +lean_dec(x_6); +x_54 = l_Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2___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_10 = lean_infer_type(x_1, x_5, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_10) == 0) +x_55 = l_panic___at_Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___spec__1(x_54, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_55) == 0) { -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_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_2); -x_13 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_13, 0, x_2); -x_14 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3___boxed), 9, 2); -lean_closure_set(x_14, 0, x_2); -lean_closure_set(x_14, 1, x_3); -x_15 = 0; -x_16 = l_Lean_Meta_forallBoundedTelescope___at_Lean_Meta_arrowDomainsN___spec__6___rarg(x_11, x_13, x_14, x_15, x_5, x_6, x_7, x_8, x_12); -return x_16; +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); +x_58 = l_Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f_go(x_2, x_3, x_4, x_5, x_56, x_7, x_8, x_9, x_10, x_57); +lean_dec(x_56); +return x_58; } else { -uint8_t x_17; +uint8_t x_59; +lean_dec(x_10); +lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_6); lean_dec(x_5); +lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_17 = !lean_is_exclusive(x_10); -if (x_17 == 0) +x_59 = !lean_is_exclusive(x_55); +if (x_59 == 0) { -return x_10; +return x_55; } else { -lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_18 = lean_ctor_get(x_10, 0); -x_19 = lean_ctor_get(x_10, 1); -lean_inc(x_19); -lean_inc(x_18); -lean_dec(x_10); -x_20 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_20, 0, x_18); -lean_ctor_set(x_20, 1, x_19); -return x_20; +lean_object* x_60; lean_object* x_61; lean_object* x_62; +x_60 = lean_ctor_get(x_55, 0); +x_61 = lean_ctor_get(x_55, 1); +lean_inc(x_61); +lean_inc(x_60); +lean_dec(x_55); +x_62 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_62, 0, x_60); +lean_ctor_set(x_62, 1, x_61); +return x_62; } } } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage(lean_object* x_1, lean_object* x_2, lean_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_Grind_mkEMatchTheoremWithKind_x3f___lambda__3(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -lean_object* x_9; uint8_t x_10; -x_9 = lean_ctor_get(x_3, 0); +lean_object* x_11; lean_inc(x_9); -x_10 = lean_nat_dec_eq(x_9, x_2); -lean_dec(x_9); -if (x_10 == 0) +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_1); +x_11 = lean_infer_type(x_1, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_11) == 0) { -lean_object* x_11; lean_object* x_12; -x_11 = lean_box(0); -x_12 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__4(x_1, x_2, x_3, x_11, x_4, x_5, x_6, x_7, x_8); -return x_12; +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_11, 0); +lean_inc(x_12); +x_13 = lean_ctor_get(x_11, 1); +lean_inc(x_13); +lean_dec(x_11); +x_14 = lean_box(x_2); +x_15 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2___boxed), 11, 4); +lean_closure_set(x_15, 0, x_14); +lean_closure_set(x_15, 1, x_3); +lean_closure_set(x_15, 2, x_4); +lean_closure_set(x_15, 3, x_1); +x_16 = 0; +x_17 = l_Lean_Meta_forallTelescopeReducing___at_Lean_Meta_getParamNames___spec__2___rarg(x_12, x_15, x_16, x_6, x_7, x_8, x_9, x_13); +return x_17; } else { -lean_object* x_13; lean_object* x_14; +uint8_t x_18; +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); lean_dec(x_1); -x_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_8); -return x_14; +x_18 = !lean_is_exclusive(x_11); +if (x_18 == 0) +{ +return x_11; } +else +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_11, 0); +x_20 = lean_ctor_get(x_11, 1); +lean_inc(x_20); +lean_inc(x_19); +lean_dec(x_11); +x_21 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_21, 0, x_19); +lean_ctor_set(x_21, 1, x_20); +return x_21; } } -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldrM___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__2___boxed(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; -x_3 = l_Std_DHashMap_Internal_AssocList_foldrM___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__2(x_1, x_2); -lean_dec(x_2); -lean_dec(x_1); -return x_3; } } -LEAN_EXPORT lean_object* l_List_mapTR_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { -lean_object* x_6; -x_6 = l_List_mapTR_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__5(x_1, x_2, x_3, x_4, x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_6; +uint8_t x_10; uint8_t x_11; +x_10 = 0; +x_11 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_beqTheoremKind____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_6540_(x_4, x_10); +if (x_11 == 0) +{ +uint8_t x_12; uint8_t x_13; +x_12 = 1; +x_13 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_beqTheoremKind____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_6540_(x_4, x_12); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; +x_14 = lean_box(0); +x_15 = l_Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__3(x_3, x_4, x_1, x_2, x_14, x_5, x_6, x_7, x_8, x_9); +return x_15; } +else +{ +uint8_t x_16; lean_object* x_17; +x_16 = 0; +x_17 = l_Lean_Meta_Grind_mkEMatchEqTheoremCore(x_1, x_2, x_3, x_16, x_16, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_17) == 0) +{ +uint8_t x_18; +x_18 = !lean_is_exclusive(x_17); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; +x_19 = lean_ctor_get(x_17, 0); +x_20 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_20, 0, x_19); +lean_ctor_set(x_17, 0, x_20); +return x_17; } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___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: +else { -size_t x_7; size_t x_8; lean_object* x_9; -x_7 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_8 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_9 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__6(x_1, x_2, x_3, x_7, x_8, x_6); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_9; +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_21 = lean_ctor_get(x_17, 0); +x_22 = lean_ctor_get(x_17, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_17); +x_23 = lean_alloc_ctor(1, 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_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__6___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: +else { -size_t x_6; size_t x_7; lean_object* x_8; -x_6 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_7 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_8 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__6___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__7(x_1, x_2, x_6, x_7, x_5); -lean_dec(x_2); -lean_dec(x_1); -return x_8; -} +uint8_t x_25; +x_25 = !lean_is_exclusive(x_17); +if (x_25 == 0) +{ +return x_17; } -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___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) { -_start: +else { -lean_object* x_15; -x_15 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___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); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_15; +lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_26 = lean_ctor_get(x_17, 0); +x_27 = lean_ctor_get(x_17, 1); +lean_inc(x_27); +lean_inc(x_26); +lean_dec(x_17); +x_28 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_28, 0, x_26); +lean_ctor_set(x_28, 1, x_27); +return x_28; } } -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__8___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__9___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { -_start: -{ -lean_object* x_14; -x_14 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__8___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__9(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_14; } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___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: +else { -size_t x_7; size_t x_8; lean_object* x_9; -x_7 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_8 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_9 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__10(x_1, x_2, x_3, x_7, x_8, x_6); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_9; -} +uint8_t x_29; uint8_t x_30; lean_object* x_31; +x_29 = 0; +x_30 = 1; +x_31 = l_Lean_Meta_Grind_mkEMatchEqTheoremCore(x_1, x_2, x_3, x_29, x_30, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_31) == 0) +{ +uint8_t x_32; +x_32 = !lean_is_exclusive(x_31); +if (x_32 == 0) +{ +lean_object* x_33; lean_object* x_34; +x_33 = lean_ctor_get(x_31, 0); +x_34 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_34, 0, x_33); +lean_ctor_set(x_31, 0, x_34); +return x_31; } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__10___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__11___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: +else { -size_t x_6; size_t x_7; lean_object* x_8; -x_6 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_7 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_8 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__10___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__11(x_1, x_2, x_6, x_7, x_5); -lean_dec(x_2); -lean_dec(x_1); -return x_8; +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_inc(x_36); +lean_inc(x_35); +lean_dec(x_31); +x_37 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_37, 0, 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_36); +return x_38; } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___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: +else { -size_t x_7; size_t x_8; lean_object* x_9; -x_7 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_8 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_9 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__12(x_1, x_2, x_3, x_7, x_8, x_6); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_9; -} +uint8_t x_39; +x_39 = !lean_is_exclusive(x_31); +if (x_39 == 0) +{ +return x_31; } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__12___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__13___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: +else { -size_t x_6; size_t x_7; lean_object* x_8; -x_6 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_7 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_8 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__12___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__13(x_1, x_2, x_6, x_7, x_5); -lean_dec(x_2); -lean_dec(x_1); -return x_8; +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_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_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__14___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { -_start: -{ -size_t x_15; size_t x_16; lean_object* x_17; -x_15 = lean_unbox_usize(x_7); -lean_dec(x_7); -x_16 = lean_unbox_usize(x_8); -lean_dec(x_8); -x_17 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__14(x_1, x_2, x_3, x_4, x_5, x_6, x_15, x_16, x_9, x_10, x_11, x_12, x_13, x_14); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_1); -return x_17; } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__14___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__15___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -size_t x_14; size_t x_15; lean_object* x_16; -x_14 = lean_unbox_usize(x_6); +lean_object* x_12; +x_12 = l_Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___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); -x_15 = lean_unbox_usize(x_7); -lean_dec(x_7); -x_16 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__14___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__15(x_1, x_2, x_3, x_4, x_5, x_14, x_15, x_8, x_9, x_10, x_11, x_12, x_13); lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_1); -return x_16; +return x_12; } } -LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -uint8_t x_11; lean_object* x_12; -x_11 = lean_unbox(x_1); +uint8_t x_12; lean_object* x_13; +x_12 = lean_unbox(x_1); lean_dec(x_1); -x_12 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___lambda__1(x_11, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); +x_13 = l_Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2(x_12, x_2, x_3, 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_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +uint8_t x_11; lean_object* x_12; +x_11 = lean_unbox(x_2); +lean_dec(x_2); +x_12 = l_Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__3(x_1, x_11, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); lean_dec(x_5); return x_12; } } -LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchTheoremWithKind_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, lean_object* x_8, lean_object* x_9) { _start: { -lean_object* x_13; -x_13 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +uint8_t x_10; lean_object* x_11; +x_10 = lean_unbox(x_4); lean_dec(x_4); -lean_dec(x_2); -lean_dec(x_1); -return x_13; +x_11 = l_Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f(x_1, x_2, x_3, x_10, x_5, x_6, x_7, x_8, x_9); +return x_11; } } -LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__17___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__1() { _start: { -lean_object* x_12; -x_12 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__17(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -lean_dec(x_2); -lean_dec(x_1); -return x_12; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Parser", 6, 6); +return x_1; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___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) { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__2() { _start: { -lean_object* x_15; -x_15 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___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); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_15; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Attr", 4, 4); +return x_1; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__18___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__19___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__3() { _start: { -lean_object* x_14; -x_14 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__18___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__19(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_14; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("grindEq", 7, 7); +return x_1; } } -LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__20___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__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_foldrMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__20(x_1, x_5, x_6, x_4); -lean_dec(x_1); -return x_7; +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_Grind_mkOffsetPattern___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__1; +x_3 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__2; +x_4 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__3; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___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) { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__5() { _start: { -lean_object* x_11; -x_11 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -return x_11; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("grindFwd", 8, 8); +return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___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) { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__6() { _start: { -lean_object* x_14; -x_14 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -return x_14; +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_Grind_mkOffsetPattern___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__1; +x_3 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__2; +x_4 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__5; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___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) { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__7() { _start: { -lean_object* x_10; -x_10 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -lean_dec(x_4); -lean_dec(x_2); -lean_dec(x_1); -return x_10; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("grindEqRhs", 10, 10); +return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___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) { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__8() { _start: { -lean_object* x_10; -x_10 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -lean_dec(x_4); -return x_10; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Lean_Meta_Grind_mkOffsetPattern___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__1; +x_3 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__2; +x_4 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__7; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; } } -static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___lambda__1___closed__1() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__9() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked(" : ", 3, 3); +x_1 = lean_mk_string_unchecked("grindEqBoth", 11, 11); return x_1; } } -static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___lambda__1___closed__2() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__10() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___lambda__1___closed__1; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Lean_Meta_Grind_mkOffsetPattern___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__1; +x_3 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__2; +x_4 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__9; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___lambda__1(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind(lean_object* x_1) { _start: { -lean_object* x_11; -lean_inc(x_1); -x_11 = lean_infer_type(x_1, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_11) == 0) +lean_object* x_2; lean_object* x_3; uint8_t x_4; +x_2 = lean_unsigned_to_nat(1u); +x_3 = l_Lean_Syntax_getArg(x_1, x_2); +x_4 = l_Lean_Syntax_isNone(x_3); +if (x_4 == 0) { -uint8_t x_12; -x_12 = !lean_is_exclusive(x_11); -if (x_12 == 0) +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; +x_5 = lean_unsigned_to_nat(0u); +x_6 = l_Lean_Syntax_getArg(x_3, x_5); +lean_dec(x_3); +x_7 = l_Lean_Syntax_getKind(x_6); +x_8 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__4; +x_9 = lean_name_eq(x_7, x_8); +if (x_9 == 0) { -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_13 = lean_ctor_get(x_11, 0); -x_14 = l_Lean_MessageData_ofExpr(x_1); -lean_inc(x_2); -x_15 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_15, 0, x_2); -lean_ctor_set(x_15, 1, x_14); -x_16 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___lambda__1___closed__2; -x_17 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_17, 0, x_15); -lean_ctor_set(x_17, 1, x_16); -x_18 = l_Lean_MessageData_ofExpr(x_13); -x_19 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_19, 0, x_17); -lean_ctor_set(x_19, 1, x_18); -x_20 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_2); -x_21 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_21, 0, x_4); -lean_ctor_set(x_21, 1, x_20); -x_22 = lean_box(x_3); -x_23 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_23, 0, x_22); -lean_ctor_set(x_23, 1, x_21); -x_24 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_24, 0, x_23); -lean_ctor_set(x_11, 0, x_24); -return x_11; +lean_object* x_10; uint8_t x_11; +x_10 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__6; +x_11 = lean_name_eq(x_7, x_10); +if (x_11 == 0) +{ +lean_object* x_12; uint8_t x_13; +x_12 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__8; +x_13 = lean_name_eq(x_7, x_12); +if (x_13 == 0) +{ +lean_object* x_14; uint8_t x_15; +x_14 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__10; +x_15 = lean_name_eq(x_7, x_14); +lean_dec(x_7); +if (x_15 == 0) +{ +uint8_t x_16; +x_16 = 4; +return x_16; } else { -lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; -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 = l_Lean_MessageData_ofExpr(x_1); -lean_inc(x_2); -x_28 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_28, 0, x_2); -lean_ctor_set(x_28, 1, x_27); -x_29 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___lambda__1___closed__2; -x_30 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_30, 0, x_28); -lean_ctor_set(x_30, 1, x_29); -x_31 = l_Lean_MessageData_ofExpr(x_25); -x_32 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_32, 0, x_30); -lean_ctor_set(x_32, 1, x_31); -x_33 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_33, 0, x_32); -lean_ctor_set(x_33, 1, x_2); -x_34 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_34, 0, x_4); -lean_ctor_set(x_34, 1, x_33); -x_35 = lean_box(x_3); -x_36 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_36, 0, x_35); -lean_ctor_set(x_36, 1, x_34); -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_26); -return x_38; +uint8_t x_17; +x_17 = 2; +return x_17; } } else { -uint8_t x_39; -lean_dec(x_4); -lean_dec(x_2); -lean_dec(x_1); -x_39 = !lean_is_exclusive(x_11); -if (x_39 == 0) -{ -return x_11; +uint8_t x_18; +lean_dec(x_7); +x_18 = 1; +return x_18; +} +} +else +{ +uint8_t x_19; +lean_dec(x_7); +x_19 = 3; +return x_19; +} +} +else +{ +uint8_t x_20; +lean_dec(x_7); +x_20 = 0; +return x_20; +} +} +else +{ +uint8_t x_21; +lean_dec(x_3); +x_21 = 5; +return x_21; +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind(x_1); +lean_dec(x_1); +x_3 = lean_box(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_ScopedEnvExtension_add___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr___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, 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_6, 6); +lean_inc(x_9); +lean_dec(x_6); +x_10 = lean_st_ref_take(x_7, 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_13 = !lean_is_exclusive(x_11); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; +x_14 = lean_ctor_get(x_11, 0); +x_15 = lean_ctor_get(x_11, 4); +lean_dec(x_15); +x_16 = l_Lean_ScopedEnvExtension_addCore___rarg(x_14, x_1, x_2, x_3, x_9); +x_17 = l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__1___closed__1; +lean_ctor_set(x_11, 4, x_17); +lean_ctor_set(x_11, 0, x_16); +x_18 = lean_st_ref_set(x_7, x_11, x_12); +x_19 = lean_ctor_get(x_18, 1); +lean_inc(x_19); +lean_dec(x_18); +x_20 = lean_st_ref_take(x_5, x_19); +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; uint8_t x_27; +x_24 = lean_ctor_get(x_21, 1); +lean_dec(x_24); +x_25 = l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__1___closed__2; +lean_ctor_set(x_21, 1, x_25); +x_26 = lean_st_ref_set(x_5, x_21, x_22); +x_27 = !lean_is_exclusive(x_26); +if (x_27 == 0) +{ +lean_object* x_28; lean_object* x_29; +x_28 = lean_ctor_get(x_26, 0); +lean_dec(x_28); +x_29 = lean_box(0); +lean_ctor_set(x_26, 0, x_29); +return x_26; +} +else +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_30 = lean_ctor_get(x_26, 1); +lean_inc(x_30); +lean_dec(x_26); +x_31 = lean_box(0); +x_32 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_30); +return x_32; +} } else { -lean_object* x_40; lean_object* x_41; lean_object* x_42; -x_40 = lean_ctor_get(x_11, 0); -x_41 = lean_ctor_get(x_11, 1); -lean_inc(x_41); +lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_33 = lean_ctor_get(x_21, 0); +x_34 = lean_ctor_get(x_21, 2); +x_35 = lean_ctor_get(x_21, 3); +x_36 = lean_ctor_get(x_21, 4); +lean_inc(x_36); +lean_inc(x_35); +lean_inc(x_34); +lean_inc(x_33); +lean_dec(x_21); +x_37 = l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__1___closed__2; +x_38 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_38, 0, x_33); +lean_ctor_set(x_38, 1, x_37); +lean_ctor_set(x_38, 2, x_34); +lean_ctor_set(x_38, 3, x_35); +lean_ctor_set(x_38, 4, x_36); +x_39 = lean_st_ref_set(x_5, x_38, x_22); +x_40 = lean_ctor_get(x_39, 1); lean_inc(x_40); -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); -return x_42; +if (lean_is_exclusive(x_39)) { + lean_ctor_release(x_39, 0); + lean_ctor_release(x_39, 1); + x_41 = x_39; +} else { + lean_dec_ref(x_39); + x_41 = lean_box(0); } +x_42 = lean_box(0); +if (lean_is_scalar(x_41)) { + x_43 = lean_alloc_ctor(0, 2, 0); +} else { + x_43 = x_41; } +lean_ctor_set(x_43, 0, x_42); +lean_ctor_set(x_43, 1, x_40); +return x_43; } } -static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___closed__1() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("\n", 1, 1); -return x_1; +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; 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_44 = lean_ctor_get(x_11, 0); +x_45 = lean_ctor_get(x_11, 1); +x_46 = lean_ctor_get(x_11, 2); +x_47 = lean_ctor_get(x_11, 3); +x_48 = lean_ctor_get(x_11, 5); +x_49 = lean_ctor_get(x_11, 6); +x_50 = lean_ctor_get(x_11, 7); +lean_inc(x_50); +lean_inc(x_49); +lean_inc(x_48); +lean_inc(x_47); +lean_inc(x_46); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_11); +x_51 = l_Lean_ScopedEnvExtension_addCore___rarg(x_44, x_1, x_2, x_3, x_9); +x_52 = l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__1___closed__1; +x_53 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_53, 0, x_51); +lean_ctor_set(x_53, 1, x_45); +lean_ctor_set(x_53, 2, x_46); +lean_ctor_set(x_53, 3, x_47); +lean_ctor_set(x_53, 4, x_52); +lean_ctor_set(x_53, 5, x_48); +lean_ctor_set(x_53, 6, x_49); +lean_ctor_set(x_53, 7, x_50); +x_54 = lean_st_ref_set(x_7, x_53, x_12); +x_55 = lean_ctor_get(x_54, 1); +lean_inc(x_55); +lean_dec(x_54); +x_56 = lean_st_ref_take(x_5, x_55); +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_inc(x_59); +x_60 = lean_ctor_get(x_57, 2); +lean_inc(x_60); +x_61 = lean_ctor_get(x_57, 3); +lean_inc(x_61); +x_62 = lean_ctor_get(x_57, 4); +lean_inc(x_62); +if (lean_is_exclusive(x_57)) { + lean_ctor_release(x_57, 0); + lean_ctor_release(x_57, 1); + lean_ctor_release(x_57, 2); + lean_ctor_release(x_57, 3); + lean_ctor_release(x_57, 4); + x_63 = x_57; +} else { + lean_dec_ref(x_57); + x_63 = lean_box(0); } +x_64 = l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__1___closed__2; +if (lean_is_scalar(x_63)) { + x_65 = lean_alloc_ctor(0, 5, 0); +} else { + x_65 = x_63; } -static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___closed__1; -x_2 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; +lean_ctor_set(x_65, 0, x_59); +lean_ctor_set(x_65, 1, x_64); +lean_ctor_set(x_65, 2, x_60); +lean_ctor_set(x_65, 3, x_61); +lean_ctor_set(x_65, 4, x_62); +x_66 = lean_st_ref_set(x_5, x_65, x_58); +x_67 = lean_ctor_get(x_66, 1); +lean_inc(x_67); +if (lean_is_exclusive(x_66)) { + lean_ctor_release(x_66, 0); + lean_ctor_release(x_66, 1); + x_68 = x_66; +} else { + lean_dec_ref(x_66); + x_68 = lean_box(0); +} +x_69 = lean_box(0); +if (lean_is_scalar(x_68)) { + x_70 = lean_alloc_ctor(0, 2, 0); +} else { + x_70 = x_68; } +lean_ctor_set(x_70, 0, x_69); +lean_ctor_set(x_70, 1, x_67); +return x_70; } -static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___closed__3() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___closed__2; -x_2 = l_Lean_MessageData_ofFormat(x_1); -return x_2; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___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_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr___spec__2(uint8_t 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: { -lean_object* x_15; uint8_t x_16; -x_15 = lean_ctor_get(x_5, 1); -x_16 = lean_nat_dec_lt(x_7, x_15); -if (x_16 == 0) +uint8_t x_13; +x_13 = lean_usize_dec_lt(x_6, x_5); +if (x_13 == 0) { -lean_object* x_17; -lean_dec(x_13); -lean_dec(x_12); +lean_object* x_14; lean_dec(x_11); lean_dec(x_10); -lean_dec(x_7); -lean_dec(x_3); -x_17 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_17, 0, x_6); -lean_ctor_set(x_17, 1, x_14); -return x_17; +lean_dec(x_9); +lean_dec(x_8); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_7); +lean_ctor_set(x_14, 1, x_12); +return x_14; } else { -uint8_t x_18; -x_18 = !lean_is_exclusive(x_6); -if (x_18 == 0) -{ -lean_object* x_19; lean_object* x_20; uint8_t x_21; -x_19 = lean_ctor_get(x_6, 0); -x_20 = lean_ctor_get(x_6, 1); -x_21 = l_List_elem___at_Lean_Meta_Occurrences_contains___spec__1(x_7, x_1); -if (x_21 == 0) -{ -lean_object* x_22; lean_object* x_23; -x_22 = lean_ctor_get(x_5, 2); -x_23 = lean_nat_add(x_7, x_22); +lean_object* x_15; uint8_t x_16; lean_object* x_17; lean_dec(x_7); -x_7 = x_23; -x_8 = lean_box(0); -x_9 = lean_box(0); -goto _start; -} -else -{ -lean_object* x_25; uint8_t x_26; -lean_free_object(x_6); -x_25 = lean_array_fget(x_2, x_7); -x_26 = lean_unbox(x_19); -if (x_26 == 0) -{ -lean_object* x_27; lean_object* x_28; lean_object* x_29; uint8_t x_30; lean_object* x_31; -x_27 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___closed__3; -x_28 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_28, 0, x_20); -lean_ctor_set(x_28, 1, x_27); -x_29 = lean_box(0); -x_30 = lean_unbox(x_19); -lean_dec(x_19); -lean_inc(x_13); -lean_inc(x_12); +x_15 = lean_array_uget(x_4, x_6); +x_16 = 1; lean_inc(x_11); lean_inc(x_10); -lean_inc(x_3); -x_31 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___lambda__1(x_25, x_3, x_30, x_28, x_29, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_31) == 0) -{ -lean_object* x_32; -x_32 = lean_ctor_get(x_31, 0); -lean_inc(x_32); -if (lean_obj_tag(x_32) == 0) +lean_inc(x_9); +lean_inc(x_8); +x_17 = l_Lean_Meta_Grind_mkEMatchEqTheorem(x_15, x_16, x_16, x_8, x_9, x_10, x_11, 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; size_t x_23; size_t x_24; lean_object* x_25; +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +x_19 = lean_ctor_get(x_17, 1); +lean_inc(x_19); +lean_dec(x_17); +x_20 = l_Lean_Meta_Grind_addEMatchTheorem___closed__1; +lean_inc(x_10); +x_21 = l_Lean_ScopedEnvExtension_add___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr___spec__1(x_20, x_18, x_1, x_8, x_9, x_10, x_11, x_19); +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_6, x_23); +x_25 = lean_box(0); +x_6 = x_24; +x_7 = x_25; +x_12 = x_22; +goto _start; +} +else { -uint8_t x_33; -lean_dec(x_13); -lean_dec(x_12); +uint8_t x_27; lean_dec(x_11); lean_dec(x_10); -lean_dec(x_7); -lean_dec(x_3); -x_33 = !lean_is_exclusive(x_31); -if (x_33 == 0) +lean_dec(x_9); +lean_dec(x_8); +x_27 = !lean_is_exclusive(x_17); +if (x_27 == 0) { -lean_object* x_34; lean_object* x_35; -x_34 = lean_ctor_get(x_31, 0); -lean_dec(x_34); -x_35 = lean_ctor_get(x_32, 0); -lean_inc(x_35); -lean_dec(x_32); -lean_ctor_set(x_31, 0, x_35); -return x_31; +return x_17; } else { -lean_object* x_36; lean_object* x_37; lean_object* x_38; -x_36 = lean_ctor_get(x_31, 1); -lean_inc(x_36); -lean_dec(x_31); -x_37 = lean_ctor_get(x_32, 0); -lean_inc(x_37); -lean_dec(x_32); -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_28; lean_object* x_29; lean_object* x_30; +x_28 = lean_ctor_get(x_17, 0); +x_29 = lean_ctor_get(x_17, 1); +lean_inc(x_29); +lean_inc(x_28); +lean_dec(x_17); +x_30 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_30, 0, x_28); +lean_ctor_set(x_30, 1, x_29); +return x_30; } } -else -{ -lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; -x_39 = lean_ctor_get(x_31, 1); -lean_inc(x_39); -lean_dec(x_31); -x_40 = lean_ctor_get(x_32, 0); -lean_inc(x_40); -lean_dec(x_32); -x_41 = lean_ctor_get(x_5, 2); -x_42 = lean_nat_add(x_7, x_41); -lean_dec(x_7); -x_6 = x_40; -x_7 = x_42; -x_8 = lean_box(0); -x_9 = lean_box(0); -x_14 = x_39; -goto _start; } } -else +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr___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: { -uint8_t x_44; -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_7); -lean_dec(x_3); -x_44 = !lean_is_exclusive(x_31); -if (x_44 == 0) +lean_object* x_9; size_t x_10; size_t x_11; lean_object* x_12; lean_object* x_13; +x_9 = lean_box(0); +x_10 = lean_array_size(x_1); +x_11 = 0; +x_12 = lean_box(0); +x_13 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr___spec__2(x_2, x_1, x_9, x_1, x_10, x_11, x_12, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_13) == 0) { -return x_31; +uint8_t x_14; +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) +{ +lean_object* x_15; +x_15 = lean_ctor_get(x_13, 0); +lean_dec(x_15); +lean_ctor_set(x_13, 0, x_12); +return x_13; } else { -lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_45 = lean_ctor_get(x_31, 0); -x_46 = lean_ctor_get(x_31, 1); -lean_inc(x_46); -lean_inc(x_45); -lean_dec(x_31); -x_47 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_47, 0, x_45); -lean_ctor_set(x_47, 1, x_46); -return x_47; -} +lean_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_12); +lean_ctor_set(x_17, 1, x_16); +return x_17; } } else { -uint8_t x_48; lean_object* x_49; lean_object* x_50; -lean_dec(x_19); -x_48 = 0; -x_49 = lean_box(0); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_3); -x_50 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___lambda__1(x_25, x_3, x_48, x_20, x_49, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_50) == 0) +uint8_t x_18; +x_18 = !lean_is_exclusive(x_13); +if (x_18 == 0) { -lean_object* x_51; -x_51 = lean_ctor_get(x_50, 0); -lean_inc(x_51); -if (lean_obj_tag(x_51) == 0) +return x_13; +} +else { -uint8_t x_52; +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_7); -lean_dec(x_3); -x_52 = !lean_is_exclusive(x_50); -if (x_52 == 0) +x_21 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_21, 0, x_19); +lean_ctor_set(x_21, 1, x_20); +return x_21; +} +} +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr___closed__1() { +_start: { -lean_object* x_53; lean_object* x_54; -x_53 = lean_ctor_get(x_50, 0); -lean_dec(x_53); -x_54 = lean_ctor_get(x_51, 0); -lean_inc(x_54); -lean_dec(x_51); -lean_ctor_set(x_50, 0, x_54); -return x_50; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("` attribute can only be applied to equational theorems or function definitions", 78, 78); +return x_1; } -else +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr___closed__2() { +_start: { -lean_object* x_55; lean_object* x_56; lean_object* x_57; -x_55 = lean_ctor_get(x_50, 1); -lean_inc(x_55); -lean_dec(x_50); -x_56 = lean_ctor_get(x_51, 0); -lean_inc(x_56); -lean_dec(x_51); -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; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("` is a definition, you must only use the left-hand side for extracting patterns", 79, 79); +return x_1; } } -else +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr___closed__3() { +_start: { -lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; -x_58 = lean_ctor_get(x_50, 1); -lean_inc(x_58); -lean_dec(x_50); -x_59 = lean_ctor_get(x_51, 0); -lean_inc(x_59); -lean_dec(x_51); -x_60 = lean_ctor_get(x_5, 2); -x_61 = lean_nat_add(x_7, x_60); -lean_dec(x_7); -x_6 = x_59; -x_7 = x_61; -x_8 = lean_box(0); -x_9 = lean_box(0); -x_14 = x_58; -goto _start; +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr___closed__2; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } } -else +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr(lean_object* x_1, uint8_t x_2, uint8_t x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: { -uint8_t x_63; -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); +lean_object* x_10; +lean_inc(x_1); +x_10 = l_Lean_getConstInfo___at_Lean_Meta_mkConstWithFreshMVarLevels___spec__1(x_1, 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; +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_isTheorem(x_11); +lean_dec(x_11); +if (x_13 == 0) +{ +lean_object* x_14; +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_1); +x_14 = l_Lean_Meta_getEqnsFor_x3f(x_1, x_5, x_6, x_7, x_8, x_12); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +if (lean_obj_tag(x_15) == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +lean_dec(x_1); +x_16 = lean_ctor_get(x_14, 1); +lean_inc(x_16); +lean_dec(x_14); +x_17 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute(x_3); +x_18 = l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___closed__3; +x_19 = lean_string_append(x_18, x_17); +lean_dec(x_17); +x_20 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr___closed__1; +x_21 = lean_string_append(x_19, x_20); +x_22 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_22, 0, x_21); +x_23 = l_Lean_MessageData_ofFormat(x_22); +x_24 = l_Lean_throwError___at_Lean_Meta_setInlineAttribute___spec__1(x_23, x_5, x_6, x_7, x_8, x_16); +lean_dec(x_8); lean_dec(x_7); -lean_dec(x_3); -x_63 = !lean_is_exclusive(x_50); -if (x_63 == 0) +lean_dec(x_6); +lean_dec(x_5); +return x_24; +} +else +{ +if (x_4 == 0) +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; uint8_t x_32; +lean_dec(x_15); +x_25 = lean_ctor_get(x_14, 1); +lean_inc(x_25); +lean_dec(x_14); +x_26 = l_Lean_MessageData_ofName(x_1); +x_27 = l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___closed__4; +x_28 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_28, 0, x_27); +lean_ctor_set(x_28, 1, x_26); +x_29 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr___closed__3; +x_30 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_30, 0, x_28); +lean_ctor_set(x_30, 1, x_29); +x_31 = l_Lean_throwError___at___private_Lean_Meta_InferType_0__Lean_Meta_inferProjType___spec__1(x_30, x_5, x_6, x_7, x_8, x_25); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +x_32 = !lean_is_exclusive(x_31); +if (x_32 == 0) { -return x_50; +return x_31; } else { -lean_object* x_64; lean_object* x_65; lean_object* x_66; -x_64 = lean_ctor_get(x_50, 0); -x_65 = lean_ctor_get(x_50, 1); -lean_inc(x_65); -lean_inc(x_64); -lean_dec(x_50); -x_66 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_66, 0, x_64); -lean_ctor_set(x_66, 1, x_65); -return x_66; +lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_33 = lean_ctor_get(x_31, 0); +x_34 = lean_ctor_get(x_31, 1); +lean_inc(x_34); +lean_inc(x_33); +lean_dec(x_31); +x_35 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_35, 0, x_33); +lean_ctor_set(x_35, 1, x_34); +return x_35; } } +else +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; +lean_dec(x_1); +x_36 = lean_ctor_get(x_14, 1); +lean_inc(x_36); +lean_dec(x_14); +x_37 = lean_ctor_get(x_15, 0); +lean_inc(x_37); +lean_dec(x_15); +x_38 = lean_box(0); +x_39 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr___lambda__1(x_37, x_2, x_38, x_5, x_6, x_7, x_8, x_36); +lean_dec(x_37); +return x_39; } } } else { -lean_object* x_67; lean_object* x_68; uint8_t x_69; -x_67 = lean_ctor_get(x_6, 0); -x_68 = lean_ctor_get(x_6, 1); -lean_inc(x_68); -lean_inc(x_67); +uint8_t x_40; +lean_dec(x_8); +lean_dec(x_7); lean_dec(x_6); -x_69 = l_List_elem___at_Lean_Meta_Occurrences_contains___spec__1(x_7, x_1); -if (x_69 == 0) +lean_dec(x_5); +lean_dec(x_1); +x_40 = !lean_is_exclusive(x_14); +if (x_40 == 0) { -lean_object* x_70; lean_object* x_71; lean_object* x_72; -x_70 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_70, 0, x_67); -lean_ctor_set(x_70, 1, x_68); -x_71 = lean_ctor_get(x_5, 2); -x_72 = lean_nat_add(x_7, x_71); -lean_dec(x_7); -x_6 = x_70; -x_7 = x_72; -x_8 = lean_box(0); -x_9 = lean_box(0); -goto _start; +return x_14; } else { -lean_object* x_74; uint8_t x_75; -x_74 = lean_array_fget(x_2, x_7); -x_75 = lean_unbox(x_67); -if (x_75 == 0) +lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_41 = lean_ctor_get(x_14, 0); +x_42 = lean_ctor_get(x_14, 1); +lean_inc(x_42); +lean_inc(x_41); +lean_dec(x_14); +x_43 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_43, 0, x_41); +lean_ctor_set(x_43, 1, x_42); +return x_43; +} +} +} +else { -lean_object* x_76; lean_object* x_77; lean_object* x_78; uint8_t x_79; lean_object* x_80; -x_76 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___closed__3; -x_77 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_77, 0, x_68); -lean_ctor_set(x_77, 1, x_76); -x_78 = lean_box(0); -x_79 = lean_unbox(x_67); -lean_dec(x_67); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_3); -x_80 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___lambda__1(x_74, x_3, x_79, x_77, x_78, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_80) == 0) +uint8_t x_44; lean_object* x_45; +x_44 = 1; +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +x_45 = l_Lean_Meta_Grind_mkEMatchEqTheorem(x_1, x_44, x_4, x_5, x_6, x_7, x_8, x_12); +if (lean_obj_tag(x_45) == 0) { -lean_object* x_81; -x_81 = lean_ctor_get(x_80, 0); -lean_inc(x_81); -if (lean_obj_tag(x_81) == 0) +lean_object* x_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 = l_Lean_Meta_Grind_addEMatchTheorem___closed__1; +x_49 = l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__1(x_48, x_46, x_2, x_5, x_6, x_7, x_8, x_47); +lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_5); +return x_49; +} +else { -lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); +uint8_t x_50; +lean_dec(x_8); lean_dec(x_7); -lean_dec(x_3); -x_82 = lean_ctor_get(x_80, 1); -lean_inc(x_82); -if (lean_is_exclusive(x_80)) { - lean_ctor_release(x_80, 0); - lean_ctor_release(x_80, 1); - x_83 = x_80; -} else { - lean_dec_ref(x_80); - x_83 = lean_box(0); +lean_dec(x_6); +lean_dec(x_5); +x_50 = !lean_is_exclusive(x_45); +if (x_50 == 0) +{ +return x_45; +} +else +{ +lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_51 = lean_ctor_get(x_45, 0); +x_52 = lean_ctor_get(x_45, 1); +lean_inc(x_52); +lean_inc(x_51); +lean_dec(x_45); +x_53 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_53, 0, x_51); +lean_ctor_set(x_53, 1, x_52); +return x_53; +} } -x_84 = lean_ctor_get(x_81, 0); -lean_inc(x_84); -lean_dec(x_81); -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; -x_86 = lean_ctor_get(x_80, 1); -lean_inc(x_86); -lean_dec(x_80); -x_87 = lean_ctor_get(x_81, 0); -lean_inc(x_87); -lean_dec(x_81); -x_88 = lean_ctor_get(x_5, 2); -x_89 = lean_nat_add(x_7, x_88); +uint8_t x_54; +lean_dec(x_8); lean_dec(x_7); -x_6 = x_87; -x_7 = x_89; -x_8 = lean_box(0); -x_9 = lean_box(0); -x_14 = x_86; -goto _start; -} +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_1); +x_54 = !lean_is_exclusive(x_10); +if (x_54 == 0) +{ +return x_10; } else { -lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); +lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_55 = lean_ctor_get(x_10, 0); +x_56 = lean_ctor_get(x_10, 1); +lean_inc(x_56); +lean_inc(x_55); lean_dec(x_10); -lean_dec(x_7); -lean_dec(x_3); -x_91 = lean_ctor_get(x_80, 0); -lean_inc(x_91); -x_92 = lean_ctor_get(x_80, 1); -lean_inc(x_92); -if (lean_is_exclusive(x_80)) { - lean_ctor_release(x_80, 0); - lean_ctor_release(x_80, 1); - x_93 = x_80; -} else { - lean_dec_ref(x_80); - x_93 = lean_box(0); +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; } -if (lean_is_scalar(x_93)) { - x_94 = lean_alloc_ctor(1, 2, 0); -} else { - x_94 = x_93; } -lean_ctor_set(x_94, 0, x_91); -lean_ctor_set(x_94, 1, x_92); -return x_94; } } -else +LEAN_EXPORT lean_object* l_Lean_ScopedEnvExtension_add___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: { -uint8_t x_95; lean_object* x_96; lean_object* x_97; -lean_dec(x_67); -x_95 = 0; -x_96 = lean_box(0); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_3); -x_97 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___lambda__1(x_74, x_3, x_95, x_68, x_96, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_97) == 0) +uint8_t x_9; lean_object* x_10; +x_9 = lean_unbox(x_3); +lean_dec(x_3); +x_10 = l_Lean_ScopedEnvExtension_add___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr___spec__1(x_1, x_2, x_9, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_7); +lean_dec(x_5); +lean_dec(x_4); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: { -lean_object* x_98; -x_98 = lean_ctor_get(x_97, 0); -lean_inc(x_98); -if (lean_obj_tag(x_98) == 0) +uint8_t x_13; size_t x_14; size_t x_15; lean_object* x_16; +x_13 = lean_unbox(x_1); +lean_dec(x_1); +x_14 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_15 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_16 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr___spec__2(x_13, x_2, x_3, x_4, x_14, x_15, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_16; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr___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_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_7); +uint8_t x_9; lean_object* x_10; +x_9 = lean_unbox(x_2); +lean_dec(x_2); +x_10 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr___lambda__1(x_1, x_9, x_3, x_4, x_5, x_6, x_7, x_8); lean_dec(x_3); -x_99 = lean_ctor_get(x_97, 1); -lean_inc(x_99); -if (lean_is_exclusive(x_97)) { - lean_ctor_release(x_97, 0); - lean_ctor_release(x_97, 1); - x_100 = x_97; -} else { - lean_dec_ref(x_97); - x_100 = lean_box(0); +lean_dec(x_1); +return x_10; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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 = lean_unbox(x_2); +lean_dec(x_2); +x_11 = lean_unbox(x_3); +lean_dec(x_3); +x_12 = lean_unbox(x_4); +lean_dec(x_4); +x_13 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr(x_1, x_10, x_11, x_12, x_5, x_6, x_7, x_8, x_9); +return x_13; } -x_101 = lean_ctor_get(x_98, 0); -lean_inc(x_101); -lean_dec(x_98); -if (lean_is_scalar(x_100)) { - x_102 = lean_alloc_ctor(0, 2, 0); -} else { - x_102 = x_100; } -lean_ctor_set(x_102, 0, x_101); -lean_ctor_set(x_102, 1, x_99); -return x_102; +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("`@", 2, 2); +return x_1; } -else +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__2() { +_start: { -lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; -x_103 = lean_ctor_get(x_97, 1); -lean_inc(x_103); -lean_dec(x_97); -x_104 = lean_ctor_get(x_98, 0); -lean_inc(x_104); -lean_dec(x_98); -x_105 = lean_ctor_get(x_5, 2); -x_106 = lean_nat_add(x_7, x_105); -lean_dec(x_7); -x_6 = x_104; -x_7 = x_106; -x_8 = lean_box(0); -x_9 = lean_box(0); -x_14 = x_103; -goto _start; +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } } -else +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__3() { +_start: { -lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_7); -lean_dec(x_3); -x_108 = lean_ctor_get(x_97, 0); -lean_inc(x_108); -x_109 = lean_ctor_get(x_97, 1); -lean_inc(x_109); -if (lean_is_exclusive(x_97)) { - lean_ctor_release(x_97, 0); - lean_ctor_release(x_97, 1); - x_110 = x_97; -} else { - lean_dec_ref(x_97); - x_110 = lean_box(0); +lean_object* x_1; +x_1 = lean_mk_string_unchecked(" theorem ", 9, 9); +return x_1; } -if (lean_is_scalar(x_110)) { - x_111 = lean_alloc_ctor(1, 2, 0); -} else { - x_111 = x_110; } -lean_ctor_set(x_111, 0, x_108); -lean_ctor_set(x_111, 1, x_109); -return x_111; +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__3; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } } +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("` ", 2, 2); +return x_1; +} } +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__5; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } } +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(", consider using different options or the `grind_pattern` command", 65, 65); +return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___lambda__1___closed__1() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__8() { _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_Grind_ppPattern___closed__4; -x_3 = lean_box(x_1); -x_4 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_4, 0, x_3); -lean_ctor_set(x_4, 1, x_2); -return x_4; +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__7; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr(lean_object* x_1, uint8_t x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_9 = lean_array_get_size(x_2); -x_10 = lean_unsigned_to_nat(0u); -x_11 = lean_unsigned_to_nat(1u); -x_12 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_12, 0, x_10); -lean_ctor_set(x_12, 1, x_9); -lean_ctor_set(x_12, 2, x_11); -x_13 = l_Lean_Meta_Grind_ppPattern___closed__4; -x_14 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___lambda__1___closed__1; -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -x_15 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1(x_1, x_2, x_13, x_12, x_12, x_14, x_10, lean_box(0), lean_box(0), x_4, x_5, x_6, x_7, x_8); -lean_dec(x_12); +uint8_t x_9; uint8_t x_10; +x_9 = 0; +x_10 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_beqTheoremKind____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_6540_(x_3, x_9); +if (x_10 == 0) +{ +uint8_t x_11; uint8_t x_12; +x_11 = 1; +x_12 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_beqTheoremKind____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_6540_(x_3, x_11); +if (x_12 == 0) +{ +uint8_t x_13; uint8_t x_14; +x_13 = 2; +x_14 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_beqTheoremKind____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_6540_(x_3, x_13); +if (x_14 == 0) +{ +lean_object* x_15; +lean_inc(x_1); +x_15 = l_Lean_getConstInfo___at_Lean_Meta_mkConstWithFreshMVarLevels___spec__1(x_1, 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_16; lean_object* x_17; uint8_t x_18; x_16 = lean_ctor_get(x_15, 0); lean_inc(x_16); x_17 = lean_ctor_get(x_15, 1); lean_inc(x_17); lean_dec(x_15); -x_18 = lean_ctor_get(x_16, 1); -lean_inc(x_18); +x_18 = l_Lean_ConstantInfo_isTheorem(x_16); lean_dec(x_16); -x_19 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_18, x_4, x_5, x_6, x_7, x_17); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -return x_19; -} -else -{ -uint8_t x_20; -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -x_20 = !lean_is_exclusive(x_15); -if (x_20 == 0) +if (x_18 == 0) { -return x_15; +uint8_t x_19; lean_object* x_20; +x_19 = 1; +x_20 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr(x_1, x_2, x_3, x_19, x_4, x_5, x_6, x_7, x_17); +return x_20; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_15, 0); -x_22 = lean_ctor_get(x_15, 1); -lean_inc(x_22); -lean_inc(x_21); -lean_dec(x_15); -x_23 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_23, 0, x_21); -lean_ctor_set(x_23, 1, x_22); -return x_23; -} -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt(lean_object* x_1, lean_object* x_2, lean_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_21; +lean_inc(x_1); +x_21 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getProofFor(x_1, x_6, x_7, x_17); +if (lean_obj_tag(x_21) == 0) { -lean_object* x_9; +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_22 = lean_ctor_get(x_21, 0); +lean_inc(x_22); +x_23 = lean_ctor_get(x_21, 1); +lean_inc(x_23); +lean_dec(x_21); +lean_inc(x_1); +x_24 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_24, 0, x_1); +x_25 = l_Lean_Meta_Grind_NormalizePattern_main___closed__1; lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -x_9 = lean_infer_type(x_1, x_4, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_9) == 0) +x_26 = l_Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f(x_24, x_25, x_22, x_3, x_4, x_5, x_6, x_7, x_23); +if (lean_obj_tag(x_26) == 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 = 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_alloc_ctor(1, 1, 0); -lean_ctor_set(x_12, 0, x_2); -x_13 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___lambda__1___boxed), 8, 1); -lean_closure_set(x_13, 0, x_3); -x_14 = 0; -x_15 = l_Lean_Meta_forallBoundedTelescope___at_Lean_Meta_arrowDomainsN___spec__6___rarg(x_10, x_12, x_13, x_14, x_4, x_5, x_6, x_7, x_11); -return x_15; -} -else +lean_object* x_27; +x_27 = lean_ctor_get(x_26, 0); +lean_inc(x_27); +if (lean_obj_tag(x_27) == 0) { -uint8_t x_16; +lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; 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_28 = lean_ctor_get(x_26, 1); +lean_inc(x_28); +lean_dec(x_26); +x_29 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute(x_3); +x_30 = l_Lean_stringToMessageData(x_29); +lean_dec(x_29); +x_31 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___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___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___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_MessageData_ofName(x_1); +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___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__6; +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___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure(x_3); +x_40 = l_Lean_stringToMessageData(x_39); +lean_dec(x_39); +x_41 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_41, 0, x_38); +lean_ctor_set(x_41, 1, x_40); +x_42 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___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 = l_Lean_throwError___at_Lean_Meta_setInlineAttribute___spec__1(x_43, x_4, x_5, x_6, x_7, x_28); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -x_16 = !lean_is_exclusive(x_9); -if (x_16 == 0) -{ -return x_9; +return x_44; } else { -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(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_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___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) { -_start: -{ -uint8_t x_11; lean_object* x_12; -x_11 = lean_unbox(x_3); -lean_dec(x_3); -x_12 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___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_5); -return x_12; -} -} -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { -_start: -{ -lean_object* x_15; -x_15 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; +lean_dec(x_1); +x_45 = lean_ctor_get(x_26, 1); +lean_inc(x_45); +lean_dec(x_26); +x_46 = lean_ctor_get(x_27, 0); +lean_inc(x_46); +lean_dec(x_27); +x_47 = l_Lean_Meta_Grind_addEMatchTheorem___closed__1; +x_48 = l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__1(x_47, x_46, x_2, x_4, x_5, x_6, x_7, x_45); +lean_dec(x_7); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_2); -lean_dec(x_1); -return x_15; +return x_48; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___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___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_3); -lean_dec(x_2); +uint8_t x_49; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); lean_dec(x_1); -return x_9; -} -} -LEAN_EXPORT lean_object* l_panic___at_Lean_Meta_Grind_addEMatchTheorem___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: +x_49 = !lean_is_exclusive(x_26); +if (x_49 == 0) { -lean_object* x_7; lean_object* x_8; lean_object* x_9; -x_7 = l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__1___closed__1; -x_8 = lean_panic_fn(x_7, x_1); -x_9 = lean_apply_5(x_8, x_2, x_3, x_4, x_5, x_6); -return x_9; -} +return x_26; } -static lean_object* _init_l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__2___closed__1() { -_start: +else { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____spec__1___closed__2; -x_2 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_2, 0, x_1); -lean_ctor_set(x_2, 1, x_1); -return x_2; -} +lean_object* x_50; lean_object* x_51; lean_object* x_52; +x_50 = lean_ctor_get(x_26, 0); +x_51 = lean_ctor_get(x_26, 1); +lean_inc(x_51); +lean_inc(x_50); +lean_dec(x_26); +x_52 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_52, 0, x_50); +lean_ctor_set(x_52, 1, x_51); +return x_52; } -static lean_object* _init_l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__2___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____spec__1___closed__2; -x_2 = lean_alloc_ctor(0, 6, 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); -lean_ctor_set(x_2, 4, x_1); -lean_ctor_set(x_2, 5, x_1); -return x_2; } } -LEAN_EXPORT lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___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) { -_start: +else { -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, 6); -lean_inc(x_9); +uint8_t x_53; +lean_dec(x_7); lean_dec(x_6); -x_10 = lean_st_ref_take(x_7, 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_13 = !lean_is_exclusive(x_11); -if (x_13 == 0) -{ -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; -x_14 = lean_ctor_get(x_11, 0); -x_15 = lean_ctor_get(x_11, 4); -lean_dec(x_15); -x_16 = l_Lean_ScopedEnvExtension_addCore___rarg(x_14, x_1, x_2, x_3, x_9); -x_17 = l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__2___closed__1; -lean_ctor_set(x_11, 4, x_17); -lean_ctor_set(x_11, 0, x_16); -x_18 = lean_st_ref_set(x_7, x_11, x_12); -x_19 = lean_ctor_get(x_18, 1); -lean_inc(x_19); -lean_dec(x_18); -x_20 = lean_st_ref_take(x_5, x_19); -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; uint8_t x_27; -x_24 = lean_ctor_get(x_21, 1); -lean_dec(x_24); -x_25 = l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__2___closed__2; -lean_ctor_set(x_21, 1, x_25); -x_26 = lean_st_ref_set(x_5, x_21, x_22); -x_27 = !lean_is_exclusive(x_26); -if (x_27 == 0) -{ -lean_object* x_28; lean_object* x_29; -x_28 = lean_ctor_get(x_26, 0); -lean_dec(x_28); -x_29 = lean_box(0); -lean_ctor_set(x_26, 0, x_29); -return x_26; -} -else +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_53 = !lean_is_exclusive(x_21); +if (x_53 == 0) { -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; -} +return x_21; } else { -lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; -x_33 = lean_ctor_get(x_21, 0); -x_34 = lean_ctor_get(x_21, 2); -x_35 = lean_ctor_get(x_21, 3); -x_36 = lean_ctor_get(x_21, 4); -lean_inc(x_36); -lean_inc(x_35); -lean_inc(x_34); -lean_inc(x_33); +lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_54 = lean_ctor_get(x_21, 0); +x_55 = lean_ctor_get(x_21, 1); +lean_inc(x_55); +lean_inc(x_54); lean_dec(x_21); -x_37 = l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__2___closed__2; -x_38 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_38, 0, x_33); -lean_ctor_set(x_38, 1, x_37); -lean_ctor_set(x_38, 2, x_34); -lean_ctor_set(x_38, 3, x_35); -lean_ctor_set(x_38, 4, x_36); -x_39 = lean_st_ref_set(x_5, x_38, x_22); -x_40 = lean_ctor_get(x_39, 1); -lean_inc(x_40); -if (lean_is_exclusive(x_39)) { - lean_ctor_release(x_39, 0); - lean_ctor_release(x_39, 1); - x_41 = x_39; -} else { - lean_dec_ref(x_39); - x_41 = lean_box(0); +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_42 = lean_box(0); -if (lean_is_scalar(x_41)) { - x_43 = lean_alloc_ctor(0, 2, 0); -} else { - x_43 = x_41; } -lean_ctor_set(x_43, 0, x_42); -lean_ctor_set(x_43, 1, x_40); -return x_43; } } else { -lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; 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_44 = lean_ctor_get(x_11, 0); -x_45 = lean_ctor_get(x_11, 1); -x_46 = lean_ctor_get(x_11, 2); -x_47 = lean_ctor_get(x_11, 3); -x_48 = lean_ctor_get(x_11, 5); -x_49 = lean_ctor_get(x_11, 6); -x_50 = lean_ctor_get(x_11, 7); -lean_inc(x_50); -lean_inc(x_49); -lean_inc(x_48); -lean_inc(x_47); -lean_inc(x_46); -lean_inc(x_45); -lean_inc(x_44); -lean_dec(x_11); -x_51 = l_Lean_ScopedEnvExtension_addCore___rarg(x_44, x_1, x_2, x_3, x_9); -x_52 = l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__2___closed__1; -x_53 = lean_alloc_ctor(0, 8, 0); -lean_ctor_set(x_53, 0, x_51); -lean_ctor_set(x_53, 1, x_45); -lean_ctor_set(x_53, 2, x_46); -lean_ctor_set(x_53, 3, x_47); -lean_ctor_set(x_53, 4, x_52); -lean_ctor_set(x_53, 5, x_48); -lean_ctor_set(x_53, 6, x_49); -lean_ctor_set(x_53, 7, x_50); -x_54 = lean_st_ref_set(x_7, x_53, x_12); -x_55 = lean_ctor_get(x_54, 1); -lean_inc(x_55); -lean_dec(x_54); -x_56 = lean_st_ref_take(x_5, x_55); -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_inc(x_59); -x_60 = lean_ctor_get(x_57, 2); -lean_inc(x_60); -x_61 = lean_ctor_get(x_57, 3); -lean_inc(x_61); -x_62 = lean_ctor_get(x_57, 4); -lean_inc(x_62); -if (lean_is_exclusive(x_57)) { - lean_ctor_release(x_57, 0); - lean_ctor_release(x_57, 1); - lean_ctor_release(x_57, 2); - lean_ctor_release(x_57, 3); - lean_ctor_release(x_57, 4); - x_63 = x_57; -} else { - lean_dec_ref(x_57); - x_63 = lean_box(0); -} -x_64 = l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__2___closed__2; -if (lean_is_scalar(x_63)) { - x_65 = lean_alloc_ctor(0, 5, 0); -} else { - x_65 = x_63; -} -lean_ctor_set(x_65, 0, x_59); -lean_ctor_set(x_65, 1, x_64); -lean_ctor_set(x_65, 2, x_60); -lean_ctor_set(x_65, 3, x_61); -lean_ctor_set(x_65, 4, x_62); -x_66 = lean_st_ref_set(x_5, x_65, x_58); -x_67 = lean_ctor_get(x_66, 1); -lean_inc(x_67); -if (lean_is_exclusive(x_66)) { - lean_ctor_release(x_66, 0); - lean_ctor_release(x_66, 1); - x_68 = x_66; -} else { - lean_dec_ref(x_66); - x_68 = lean_box(0); -} -x_69 = lean_box(0); -if (lean_is_scalar(x_68)) { - x_70 = lean_alloc_ctor(0, 2, 0); -} else { - x_70 = x_68; +uint8_t x_57; +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_15); +if (x_57 == 0) +{ +return x_15; } -lean_ctor_set(x_70, 0, x_69); -lean_ctor_set(x_70, 1, x_67); -return x_70; +else +{ +lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_58 = lean_ctor_get(x_15, 0); +x_59 = lean_ctor_get(x_15, 1); +lean_inc(x_59); +lean_inc(x_58); +lean_dec(x_15); +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_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__3(lean_object* x_1, lean_object* x_2) { -_start: +else { -if (lean_obj_tag(x_1) == 0) +uint8_t x_61; lean_object* x_62; +x_61 = 1; +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_1); +x_62 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr(x_1, x_2, x_3, x_61, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_62) == 0) { -lean_object* x_3; -x_3 = l_List_reverse___rarg(x_2); -return x_3; +lean_object* x_63; uint8_t x_64; lean_object* x_65; +x_63 = lean_ctor_get(x_62, 1); +lean_inc(x_63); +lean_dec(x_62); +x_64 = 0; +x_65 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr(x_1, x_2, x_3, x_64, x_4, x_5, x_6, x_7, x_63); +return x_65; } else { -uint8_t x_4; -x_4 = !lean_is_exclusive(x_1); -if (x_4 == 0) +uint8_t x_66; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_66 = !lean_is_exclusive(x_62); +if (x_66 == 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_Meta_Grind_ppPattern(x_5); -lean_ctor_set(x_1, 1, x_2); -lean_ctor_set(x_1, 0, x_7); +return x_62; +} +else { -lean_object* _tmp_0 = x_6; -lean_object* _tmp_1 = x_1; -x_1 = _tmp_0; -x_2 = _tmp_1; +lean_object* x_67; lean_object* x_68; lean_object* x_69; +x_67 = lean_ctor_get(x_62, 0); +x_68 = lean_ctor_get(x_62, 1); +lean_inc(x_68); +lean_inc(x_67); +lean_dec(x_62); +x_69 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_69, 0, x_67); +lean_ctor_set(x_69, 1, x_68); +return x_69; +} +} } -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_Meta_Grind_ppPattern(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; +uint8_t x_70; lean_object* x_71; +x_70 = 0; +x_71 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr(x_1, x_2, x_3, x_70, x_4, x_5, x_6, x_7, x_8); +return x_71; } } +else +{ +uint8_t x_72; lean_object* x_73; +x_72 = 1; +x_73 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr(x_1, x_2, x_3, x_72, x_4, x_5, x_6, x_7, x_8); +return x_73; +} } } -LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__4(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { -if (lean_obj_tag(x_1) == 0) -{ -lean_object* x_3; -x_3 = l_List_reverse___rarg(x_2); -return x_3; +uint8_t x_9; uint8_t x_10; lean_object* x_11; +x_9 = lean_unbox(x_2); +lean_dec(x_2); +x_10 = lean_unbox(x_3); +lean_dec(x_3); +x_11 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr(x_1, x_9, x_10, x_4, x_5, x_6, x_7, x_8); +return x_11; } -else -{ -uint8_t x_4; -x_4 = !lean_is_exclusive(x_1); -if (x_4 == 0) +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____spec__1(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { +_start: { -lean_object* x_5; -x_5 = lean_ctor_get(x_1, 1); -lean_ctor_set(x_1, 1, x_2); +uint8_t x_5; +x_5 = lean_usize_dec_eq(x_2, x_3); +if (x_5 == 0) { -lean_object* _tmp_0 = x_5; -lean_object* _tmp_1 = x_1; -x_1 = _tmp_0; -x_2 = _tmp_1; -} +lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x_9; size_t x_10; +x_6 = lean_array_uget(x_1, x_2); +x_7 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_7, 0, x_6); +x_8 = l_Lean_Meta_Grind_EMatchTheorems_erase(x_4, x_7); +x_9 = 1; +x_10 = lean_usize_add(x_2, x_9); +x_2 = x_10; +x_4 = x_8; goto _start; } else { -lean_object* x_7; lean_object* x_8; lean_object* x_9; -x_7 = lean_ctor_get(x_1, 0); -x_8 = lean_ctor_get(x_1, 1); -lean_inc(x_8); -lean_inc(x_7); -lean_dec(x_1); -x_9 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_9, 0, x_7); -lean_ctor_set(x_9, 1, x_2); -x_1 = x_8; -x_2 = x_9; -goto _start; -} +return x_4; } } } -LEAN_EXPORT uint8_t l_Lean_Meta_Grind_addEMatchTheorem___lambda__1(lean_object* x_1) { +LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____spec__2(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4) { _start: { -if (lean_obj_tag(x_1) == 2) +uint8_t x_5; +x_5 = lean_usize_dec_eq(x_3, x_4); +if (x_5 == 0) { -uint8_t x_2; -x_2 = 1; -return x_2; +lean_object* x_6; lean_object* x_7; uint8_t x_8; +x_6 = lean_array_uget(x_2, x_3); +x_7 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_7, 0, x_6); +lean_inc(x_1); +x_8 = l_Lean_Meta_Grind_EMatchTheorems_contains(x_1, x_7); +lean_dec(x_7); +if (x_8 == 0) +{ +uint8_t x_9; +lean_dec(x_1); +x_9 = 1; +return x_9; } else { -uint8_t x_3; -x_3 = 0; -return x_3; +size_t x_10; size_t x_11; +x_10 = 1; +x_11 = lean_usize_add(x_3, x_10); +x_3 = x_11; +goto _start; +} +} +else +{ +uint8_t x_13; +lean_dec(x_1); +x_13 = 0; +return x_13; } } } -static lean_object* _init_l_Lean_Meta_Grind_addEMatchTheorem___lambda__2___closed__1() { +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__1() { _start: { -lean_object* x_1; -x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ematchTheoremsExt; -return x_1; +uint8_t x_1; uint8_t x_2; uint8_t x_3; uint8_t x_4; uint8_t x_5; lean_object* x_6; +x_1 = 0; +x_2 = 1; +x_3 = 1; +x_4 = 0; +x_5 = 2; +x_6 = lean_alloc_ctor(0, 0, 17); +lean_ctor_set_uint8(x_6, 0, x_1); +lean_ctor_set_uint8(x_6, 1, x_1); +lean_ctor_set_uint8(x_6, 2, x_1); +lean_ctor_set_uint8(x_6, 3, x_1); +lean_ctor_set_uint8(x_6, 4, x_1); +lean_ctor_set_uint8(x_6, 5, x_2); +lean_ctor_set_uint8(x_6, 6, x_2); +lean_ctor_set_uint8(x_6, 7, x_1); +lean_ctor_set_uint8(x_6, 8, x_2); +lean_ctor_set_uint8(x_6, 9, x_3); +lean_ctor_set_uint8(x_6, 10, x_4); +lean_ctor_set_uint8(x_6, 11, x_2); +lean_ctor_set_uint8(x_6, 12, x_2); +lean_ctor_set_uint8(x_6, 13, x_2); +lean_ctor_set_uint8(x_6, 14, x_5); +lean_ctor_set_uint8(x_6, 15, x_2); +lean_ctor_set_uint8(x_6, 16, x_2); +return x_6; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addEMatchTheorem___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) { +static uint64_t _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__2() { _start: { -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_array_mk(x_1); -x_14 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_14, 0, x_2); -x_15 = lean_alloc_ctor(0, 6, 0); -lean_ctor_set(x_15, 0, x_13); -lean_ctor_set(x_15, 1, x_3); -lean_ctor_set(x_15, 2, x_4); -lean_ctor_set(x_15, 3, x_5); -lean_ctor_set(x_15, 4, x_6); -lean_ctor_set(x_15, 5, x_14); -x_16 = l_Lean_Meta_Grind_addEMatchTheorem___lambda__2___closed__1; -x_17 = 0; -x_18 = l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__2(x_16, x_15, x_17, x_8, x_9, x_10, x_11, x_12); -return x_18; +lean_object* x_1; uint64_t x_2; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__1; +x_2 = l___private_Lean_Meta_Basic_0__Lean_Meta_Config_toKey(x_1); +return x_2; } } -static lean_object* _init_l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__1() { +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__3() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("invalid pattern(s) for `", 24, 24); -return x_1; +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(32u); +x_2 = lean_mk_empty_array_with_capacity(x_1); +return x_2; } } -static lean_object* _init_l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__2() { +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__4() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__1; -x_2 = l_Lean_stringToMessageData(x_1); +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__3; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__3() { +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__5() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("`", 1, 1); -return x_1; +size_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = 5; +x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__4; +x_3 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__3; +x_4 = lean_unsigned_to_nat(0u); +x_5 = lean_alloc_ctor(0, 4, sizeof(size_t)*1); +lean_ctor_set(x_5, 0, x_2); +lean_ctor_set(x_5, 1, x_3); +lean_ctor_set(x_5, 2, x_4); +lean_ctor_set(x_5, 3, x_4); +lean_ctor_set_usize(x_5, 4, x_1); +return x_5; } } -static lean_object* _init_l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__4() { +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__6() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__3; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedEMatchTheorems___spec__1___closed__2; +x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__5; +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; } } -static lean_object* _init_l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__5() { +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__7() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("\nthe following theorem parameters cannot be instantiated:", 57, 57); -return x_1; +lean_object* x_1; lean_object* x_2; lean_object* x_3; uint64_t x_4; uint8_t x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_1 = lean_box(0); +x_2 = lean_box(0); +x_3 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__1; +x_4 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__2; +x_5 = 0; +x_6 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__6; +x_7 = l_Lean_Meta_Grind_NormalizePattern_main___closed__1; +x_8 = lean_unsigned_to_nat(0u); +x_9 = lean_alloc_ctor(0, 7, 11); +lean_ctor_set(x_9, 0, x_3); +lean_ctor_set(x_9, 1, x_1); +lean_ctor_set(x_9, 2, x_6); +lean_ctor_set(x_9, 3, x_7); +lean_ctor_set(x_9, 4, x_2); +lean_ctor_set(x_9, 5, x_8); +lean_ctor_set(x_9, 6, x_2); +lean_ctor_set_uint64(x_9, sizeof(void*)*7, x_4); +lean_ctor_set_uint8(x_9, sizeof(void*)*7 + 8, x_5); +lean_ctor_set_uint8(x_9, sizeof(void*)*7 + 9, x_5); +lean_ctor_set_uint8(x_9, sizeof(void*)*7 + 10, x_5); +return x_9; +} +} +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_unsigned_to_nat(0u); +x_2 = l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedEMatchTheorems___spec__1___closed__2; +x_3 = lean_alloc_ctor(0, 9, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_1); +lean_ctor_set(x_3, 2, x_1); +lean_ctor_set(x_3, 3, x_2); +lean_ctor_set(x_3, 4, x_2); +lean_ctor_set(x_3, 5, x_2); +lean_ctor_set(x_3, 6, x_2); +lean_ctor_set(x_3, 7, x_2); +lean_ctor_set(x_3, 8, x_2); +return x_3; } } -static lean_object* _init_l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__6() { +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__9() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__5; -x_2 = l_Lean_stringToMessageData(x_1); +x_1 = l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedEMatchTheorems___spec__1___closed__2; +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; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addEMatchTheorem___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) { +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__10() { _start: { -lean_object* x_14; -lean_inc(x_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; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__8; +x_3 = l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__1___closed__2; +x_4 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__5; +x_5 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__9; +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_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____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: +{ +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; +x_7 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind(x_2); +x_8 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__10; +x_9 = lean_st_mk_ref(x_8, x_6); +x_10 = lean_ctor_get(x_9, 0); +lean_inc(x_10); +x_11 = lean_ctor_get(x_9, 1); lean_inc(x_11); +lean_dec(x_9); +x_12 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__7; lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_2); -lean_inc(x_1); -x_14 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage(x_1, x_2, x_3, x_9, x_10, x_11, x_12, x_13); -if (lean_obj_tag(x_14) == 0) +x_13 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr(x_1, x_3, x_7, x_12, x_10, x_4, x_5, x_11); +if (lean_obj_tag(x_13) == 0) { -lean_object* x_15; -x_15 = lean_ctor_get(x_14, 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); -if (lean_obj_tag(x_15) == 0) -{ -lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_16 = lean_ctor_get(x_14, 1); -lean_inc(x_16); -lean_dec(x_14); -x_17 = lean_box(0); -x_18 = l_Lean_Meta_Grind_addEMatchTheorem___lambda__2(x_4, x_5, x_1, x_2, x_6, x_7, x_17, x_9, x_10, x_11, x_12, x_16); -lean_dec(x_12); +lean_dec(x_13); +x_16 = lean_st_ref_get(x_10, x_15); lean_dec(x_10); -lean_dec(x_9); -return x_18; +x_17 = !lean_is_exclusive(x_16); +if (x_17 == 0) +{ +lean_object* x_18; +x_18 = lean_ctor_get(x_16, 0); +lean_dec(x_18); +lean_ctor_set(x_16, 0, x_14); +return x_16; } else { -lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; -lean_dec(x_7); -x_19 = lean_ctor_get(x_14, 1); +lean_object* x_19; lean_object* x_20; +x_19 = lean_ctor_get(x_16, 1); lean_inc(x_19); -lean_dec(x_14); -x_20 = lean_ctor_get(x_15, 0); -lean_inc(x_20); -lean_dec(x_15); -lean_inc(x_4); -x_21 = l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__3(x_6, x_4); -x_22 = l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__4(x_21, x_4); -x_23 = l_Lean_MessageData_ofList(x_22); -x_24 = l_Lean_Meta_Grind_ppPattern___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); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -x_27 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt(x_1, x_2, x_20, x_9, x_10, x_11, x_12, x_19); -if (lean_obj_tag(x_27) == 0) +lean_dec(x_16); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_14); +lean_ctor_set(x_20, 1, x_19); +return x_20; +} +} +else { -lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; 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_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_MessageData_ofName(x_5); -x_31 = l_Lean_Meta_Grind_addEMatchTheorem___lambda__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_Meta_Grind_addEMatchTheorem___lambda__3___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_indentD(x_26); -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_Meta_Grind_addEMatchTheorem___lambda__3___closed__6; -x_38 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_38, 0, x_36); -lean_ctor_set(x_38, 1, x_37); -x_39 = l_Lean_indentD(x_28); -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_alloc_ctor(7, 2, 0); -lean_ctor_set(x_41, 0, x_40); -lean_ctor_set(x_41, 1, x_24); -x_42 = l_Lean_throwError___at___private_Lean_Meta_InferType_0__Lean_Meta_inferProjType___spec__1(x_41, x_9, x_10, x_11, x_12, x_29); -lean_dec(x_12); -lean_dec(x_11); +uint8_t x_21; lean_dec(x_10); -lean_dec(x_9); -x_43 = !lean_is_exclusive(x_42); -if (x_43 == 0) +x_21 = !lean_is_exclusive(x_13); +if (x_21 == 0) { -return x_42; +return x_13; } else { -lean_object* x_44; lean_object* x_45; lean_object* x_46; -x_44 = lean_ctor_get(x_42, 0); -x_45 = lean_ctor_get(x_42, 1); -lean_inc(x_45); -lean_inc(x_44); -lean_dec(x_42); -x_46 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_45); -return x_46; +lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_22 = lean_ctor_get(x_13, 0); +x_23 = lean_ctor_get(x_13, 1); +lean_inc(x_23); +lean_inc(x_22); +lean_dec(x_13); +x_24 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_24, 0, x_22); +lean_ctor_set(x_24, 1, x_23); +return x_24; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__2(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; uint8_t x_5; +x_3 = lean_array_get_size(x_1); +x_4 = lean_unsigned_to_nat(0u); +x_5 = lean_nat_dec_lt(x_4, x_3); +if (x_5 == 0) +{ +lean_dec(x_3); +return x_2; } +else +{ +uint8_t x_6; +x_6 = lean_nat_dec_le(x_3, x_3); +if (x_6 == 0) +{ +lean_dec(x_3); +return x_2; } else { -uint8_t x_47; +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_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____spec__1(x_1, x_7, x_8, x_2); +return x_9; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____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; uint8_t x_11; +x_10 = lean_st_ref_take(x_8, x_9); +x_11 = !lean_is_exclusive(x_10); +if (x_11 == 0) +{ +lean_object* x_12; uint8_t x_13; +x_12 = lean_ctor_get(x_10, 0); +x_13 = !lean_is_exclusive(x_12); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; +x_14 = lean_ctor_get(x_10, 1); +x_15 = lean_ctor_get(x_12, 0); +x_16 = lean_ctor_get(x_12, 4); +lean_dec(x_16); +x_17 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__2___boxed), 2, 1); +lean_closure_set(x_17, 0, x_1); +x_18 = l_Lean_Meta_Grind_addEMatchTheorem___closed__1; +x_19 = l_Lean_ScopedEnvExtension_modifyState___rarg(x_18, x_15, x_17); +lean_inc(x_2); +lean_ctor_set(x_10, 1, x_2); +lean_ctor_set(x_10, 0, x_2); +lean_ctor_set(x_12, 4, x_10); +lean_ctor_set(x_12, 0, x_19); +x_20 = lean_st_ref_set(x_8, x_12, x_14); +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_26; lean_object* x_27; uint8_t x_28; +x_26 = lean_ctor_get(x_23, 1); lean_dec(x_26); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_5); -x_47 = !lean_is_exclusive(x_27); -if (x_47 == 0) +lean_ctor_set(x_23, 1, x_3); +x_27 = lean_st_ref_set(x_6, x_23, x_24); +x_28 = !lean_is_exclusive(x_27); +if (x_28 == 0) { +lean_object* x_29; lean_object* x_30; +x_29 = lean_ctor_get(x_27, 0); +lean_dec(x_29); +x_30 = lean_box(0); +lean_ctor_set(x_27, 0, x_30); return x_27; } else { -lean_object* x_48; lean_object* x_49; lean_object* x_50; -x_48 = lean_ctor_get(x_27, 0); -x_49 = lean_ctor_get(x_27, 1); -lean_inc(x_49); -lean_inc(x_48); +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_27, 1); +lean_inc(x_31); lean_dec(x_27); -x_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_32 = lean_box(0); +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_31); +return x_33; +} +} +else +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_34 = lean_ctor_get(x_23, 0); +x_35 = lean_ctor_get(x_23, 2); +x_36 = lean_ctor_get(x_23, 3); +x_37 = lean_ctor_get(x_23, 4); +lean_inc(x_37); +lean_inc(x_36); +lean_inc(x_35); +lean_inc(x_34); +lean_dec(x_23); +x_38 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_38, 0, x_34); +lean_ctor_set(x_38, 1, x_3); +lean_ctor_set(x_38, 2, x_35); +lean_ctor_set(x_38, 3, x_36); +lean_ctor_set(x_38, 4, x_37); +x_39 = lean_st_ref_set(x_6, x_38, x_24); +x_40 = lean_ctor_get(x_39, 1); +lean_inc(x_40); +if (lean_is_exclusive(x_39)) { + lean_ctor_release(x_39, 0); + lean_ctor_release(x_39, 1); + x_41 = x_39; +} else { + lean_dec_ref(x_39); + x_41 = lean_box(0); } +x_42 = lean_box(0); +if (lean_is_scalar(x_41)) { + x_43 = lean_alloc_ctor(0, 2, 0); +} else { + x_43 = x_41; } +lean_ctor_set(x_43, 0, x_42); +lean_ctor_set(x_43, 1, x_40); +return x_43; } } else { -uint8_t x_51; +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; +x_44 = lean_ctor_get(x_10, 1); +x_45 = lean_ctor_get(x_12, 0); +x_46 = lean_ctor_get(x_12, 1); +x_47 = lean_ctor_get(x_12, 2); +x_48 = lean_ctor_get(x_12, 3); +x_49 = lean_ctor_get(x_12, 5); +x_50 = lean_ctor_get(x_12, 6); +x_51 = lean_ctor_get(x_12, 7); +lean_inc(x_51); +lean_inc(x_50); +lean_inc(x_49); +lean_inc(x_48); +lean_inc(x_47); +lean_inc(x_46); +lean_inc(x_45); lean_dec(x_12); -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_2); -lean_dec(x_1); -x_51 = !lean_is_exclusive(x_14); -if (x_51 == 0) -{ -return x_14; +x_52 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__2___boxed), 2, 1); +lean_closure_set(x_52, 0, x_1); +x_53 = l_Lean_Meta_Grind_addEMatchTheorem___closed__1; +x_54 = l_Lean_ScopedEnvExtension_modifyState___rarg(x_53, x_45, x_52); +lean_inc(x_2); +lean_ctor_set(x_10, 1, x_2); +lean_ctor_set(x_10, 0, x_2); +x_55 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_55, 0, x_54); +lean_ctor_set(x_55, 1, x_46); +lean_ctor_set(x_55, 2, x_47); +lean_ctor_set(x_55, 3, x_48); +lean_ctor_set(x_55, 4, x_10); +lean_ctor_set(x_55, 5, x_49); +lean_ctor_set(x_55, 6, x_50); +lean_ctor_set(x_55, 7, x_51); +x_56 = lean_st_ref_set(x_8, x_55, x_44); +x_57 = lean_ctor_get(x_56, 1); +lean_inc(x_57); +lean_dec(x_56); +x_58 = lean_st_ref_take(x_6, x_57); +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_59, 0); +lean_inc(x_61); +x_62 = lean_ctor_get(x_59, 2); +lean_inc(x_62); +x_63 = lean_ctor_get(x_59, 3); +lean_inc(x_63); +x_64 = lean_ctor_get(x_59, 4); +lean_inc(x_64); +if (lean_is_exclusive(x_59)) { + lean_ctor_release(x_59, 0); + lean_ctor_release(x_59, 1); + lean_ctor_release(x_59, 2); + lean_ctor_release(x_59, 3); + lean_ctor_release(x_59, 4); + x_65 = x_59; +} else { + lean_dec_ref(x_59); + x_65 = lean_box(0); +} +if (lean_is_scalar(x_65)) { + x_66 = lean_alloc_ctor(0, 5, 0); +} else { + x_66 = x_65; +} +lean_ctor_set(x_66, 0, x_61); +lean_ctor_set(x_66, 1, x_3); +lean_ctor_set(x_66, 2, x_62); +lean_ctor_set(x_66, 3, x_63); +lean_ctor_set(x_66, 4, x_64); +x_67 = lean_st_ref_set(x_6, x_66, x_60); +x_68 = lean_ctor_get(x_67, 1); +lean_inc(x_68); +if (lean_is_exclusive(x_67)) { + lean_ctor_release(x_67, 0); + lean_ctor_release(x_67, 1); + x_69 = x_67; +} else { + lean_dec_ref(x_67); + x_69 = lean_box(0); +} +x_70 = lean_box(0); +if (lean_is_scalar(x_69)) { + x_71 = lean_alloc_ctor(0, 2, 0); +} else { + x_71 = x_69; +} +lean_ctor_set(x_71, 0, x_70); +lean_ctor_set(x_71, 1, x_68); +return x_71; +} } else { -lean_object* x_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_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; +lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; +x_72 = lean_ctor_get(x_10, 0); +x_73 = lean_ctor_get(x_10, 1); +lean_inc(x_73); +lean_inc(x_72); +lean_dec(x_10); +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); +x_80 = lean_ctor_get(x_72, 7); +lean_inc(x_80); +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); + lean_ctor_release(x_72, 7); + x_81 = x_72; +} else { + lean_dec_ref(x_72); + x_81 = lean_box(0); +} +x_82 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__2___boxed), 2, 1); +lean_closure_set(x_82, 0, x_1); +x_83 = l_Lean_Meta_Grind_addEMatchTheorem___closed__1; +x_84 = l_Lean_ScopedEnvExtension_modifyState___rarg(x_83, x_74, x_82); +lean_inc(x_2); +x_85 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_85, 0, x_2); +lean_ctor_set(x_85, 1, x_2); +if (lean_is_scalar(x_81)) { + x_86 = lean_alloc_ctor(0, 8, 0); +} else { + x_86 = x_81; +} +lean_ctor_set(x_86, 0, x_84); +lean_ctor_set(x_86, 1, x_75); +lean_ctor_set(x_86, 2, x_76); +lean_ctor_set(x_86, 3, x_77); +lean_ctor_set(x_86, 4, x_85); +lean_ctor_set(x_86, 5, x_78); +lean_ctor_set(x_86, 6, x_79); +lean_ctor_set(x_86, 7, x_80); +x_87 = lean_st_ref_set(x_8, x_86, x_73); +x_88 = lean_ctor_get(x_87, 1); +lean_inc(x_88); +lean_dec(x_87); +x_89 = lean_st_ref_take(x_6, 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, 2); +lean_inc(x_93); +x_94 = lean_ctor_get(x_90, 3); +lean_inc(x_94); +x_95 = lean_ctor_get(x_90, 4); +lean_inc(x_95); +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_96 = x_90; +} else { + lean_dec_ref(x_90); + x_96 = lean_box(0); +} +if (lean_is_scalar(x_96)) { + x_97 = lean_alloc_ctor(0, 5, 0); +} else { + x_97 = x_96; +} +lean_ctor_set(x_97, 0, x_92); +lean_ctor_set(x_97, 1, x_3); +lean_ctor_set(x_97, 2, x_93); +lean_ctor_set(x_97, 3, x_94); +lean_ctor_set(x_97, 4, x_95); +x_98 = lean_st_ref_set(x_6, x_97, x_91); +x_99 = lean_ctor_get(x_98, 1); +lean_inc(x_99); +if (lean_is_exclusive(x_98)) { + lean_ctor_release(x_98, 0); + lean_ctor_release(x_98, 1); + x_100 = x_98; +} else { + lean_dec_ref(x_98); + x_100 = lean_box(0); +} +x_101 = lean_box(0); +if (lean_is_scalar(x_100)) { + x_102 = lean_alloc_ctor(0, 2, 0); +} else { + x_102 = x_100; } +lean_ctor_set(x_102, 0, x_101); +lean_ctor_set(x_102, 1, x_99); +return x_102; } } } -static lean_object* _init_l_Lean_Meta_Grind_addEMatchTheorem___closed__1() { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__4(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("` is not a theorem, you cannot assign patterns to non-theorems for the `grind` tactic", 85, 85); -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_Grind_addEMatchTheorem___closed__2() { -_start: +lean_object* x_3; lean_object* x_4; +x_3 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_3, 0, x_1); +x_4 = l_Lean_Meta_Grind_EMatchTheorems_erase(x_2, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; uint8_t x_11; +x_10 = lean_st_ref_take(x_8, x_9); +x_11 = !lean_is_exclusive(x_10); +if (x_11 == 0) +{ +lean_object* x_12; uint8_t x_13; +x_12 = lean_ctor_get(x_10, 0); +x_13 = !lean_is_exclusive(x_12); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; +x_14 = lean_ctor_get(x_10, 1); +x_15 = lean_ctor_get(x_12, 0); +x_16 = lean_ctor_get(x_12, 4); +lean_dec(x_16); +x_17 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__4), 2, 1); +lean_closure_set(x_17, 0, x_1); +x_18 = l_Lean_Meta_Grind_addEMatchTheorem___closed__1; +x_19 = l_Lean_ScopedEnvExtension_modifyState___rarg(x_18, x_15, x_17); +lean_inc(x_2); +lean_ctor_set(x_10, 1, x_2); +lean_ctor_set(x_10, 0, x_2); +lean_ctor_set(x_12, 4, x_10); +lean_ctor_set(x_12, 0, x_19); +x_20 = lean_st_ref_set(x_8, x_12, x_14); +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_26; lean_object* x_27; uint8_t x_28; +x_26 = lean_ctor_get(x_23, 1); +lean_dec(x_26); +lean_ctor_set(x_23, 1, x_3); +x_27 = lean_st_ref_set(x_6, x_23, x_24); +x_28 = !lean_is_exclusive(x_27); +if (x_28 == 0) { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_Grind_addEMatchTheorem___closed__1; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; -} +lean_object* x_29; lean_object* x_30; +x_29 = lean_ctor_get(x_27, 0); +lean_dec(x_29); +x_30 = lean_box(0); +lean_ctor_set(x_27, 0, x_30); +return x_27; } -static lean_object* _init_l_Lean_Meta_Grind_addEMatchTheorem___closed__3() { -_start: +else { -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_addEMatchTheorem___lambda__1___boxed), 1, 0); -return x_1; +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_27, 1); +lean_inc(x_31); +lean_dec(x_27); +x_32 = lean_box(0); +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_31); +return x_33; } } -static lean_object* _init_l_Lean_Meta_Grind_addEMatchTheorem___closed__4() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("symbols.all fun s => s matches .const _\n ", 42, 42); -return x_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; +x_34 = lean_ctor_get(x_23, 0); +x_35 = lean_ctor_get(x_23, 2); +x_36 = lean_ctor_get(x_23, 3); +x_37 = lean_ctor_get(x_23, 4); +lean_inc(x_37); +lean_inc(x_36); +lean_inc(x_35); +lean_inc(x_34); +lean_dec(x_23); +x_38 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_38, 0, x_34); +lean_ctor_set(x_38, 1, x_3); +lean_ctor_set(x_38, 2, x_35); +lean_ctor_set(x_38, 3, x_36); +lean_ctor_set(x_38, 4, x_37); +x_39 = lean_st_ref_set(x_6, x_38, x_24); +x_40 = lean_ctor_get(x_39, 1); +lean_inc(x_40); +if (lean_is_exclusive(x_39)) { + lean_ctor_release(x_39, 0); + lean_ctor_release(x_39, 1); + x_41 = x_39; +} else { + lean_dec_ref(x_39); + x_41 = lean_box(0); } +x_42 = lean_box(0); +if (lean_is_scalar(x_41)) { + x_43 = lean_alloc_ctor(0, 2, 0); +} else { + x_43 = x_41; } -static lean_object* _init_l_Lean_Meta_Grind_addEMatchTheorem___closed__5() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__3; -x_2 = l_Lean_Meta_Grind_addEMatchTheorem___closed__4; -x_3 = lean_string_append(x_1, x_2); -return x_3; +lean_ctor_set(x_43, 0, x_42); +lean_ctor_set(x_43, 1, x_40); +return x_43; } } -static lean_object* _init_l_Lean_Meta_Grind_addEMatchTheorem___closed__6() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean.Meta.Grind.addEMatchTheorem", 32, 32); -return x_1; +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; +x_44 = lean_ctor_get(x_10, 1); +x_45 = lean_ctor_get(x_12, 0); +x_46 = lean_ctor_get(x_12, 1); +x_47 = lean_ctor_get(x_12, 2); +x_48 = lean_ctor_get(x_12, 3); +x_49 = lean_ctor_get(x_12, 5); +x_50 = lean_ctor_get(x_12, 6); +x_51 = lean_ctor_get(x_12, 7); +lean_inc(x_51); +lean_inc(x_50); +lean_inc(x_49); +lean_inc(x_48); +lean_inc(x_47); +lean_inc(x_46); +lean_inc(x_45); +lean_dec(x_12); +x_52 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__4), 2, 1); +lean_closure_set(x_52, 0, x_1); +x_53 = l_Lean_Meta_Grind_addEMatchTheorem___closed__1; +x_54 = l_Lean_ScopedEnvExtension_modifyState___rarg(x_53, x_45, x_52); +lean_inc(x_2); +lean_ctor_set(x_10, 1, x_2); +lean_ctor_set(x_10, 0, x_2); +x_55 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_55, 0, x_54); +lean_ctor_set(x_55, 1, x_46); +lean_ctor_set(x_55, 2, x_47); +lean_ctor_set(x_55, 3, x_48); +lean_ctor_set(x_55, 4, x_10); +lean_ctor_set(x_55, 5, x_49); +lean_ctor_set(x_55, 6, x_50); +lean_ctor_set(x_55, 7, x_51); +x_56 = lean_st_ref_set(x_8, x_55, x_44); +x_57 = lean_ctor_get(x_56, 1); +lean_inc(x_57); +lean_dec(x_56); +x_58 = lean_st_ref_take(x_6, x_57); +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_59, 0); +lean_inc(x_61); +x_62 = lean_ctor_get(x_59, 2); +lean_inc(x_62); +x_63 = lean_ctor_get(x_59, 3); +lean_inc(x_63); +x_64 = lean_ctor_get(x_59, 4); +lean_inc(x_64); +if (lean_is_exclusive(x_59)) { + lean_ctor_release(x_59, 0); + lean_ctor_release(x_59, 1); + lean_ctor_release(x_59, 2); + lean_ctor_release(x_59, 3); + lean_ctor_release(x_59, 4); + x_65 = x_59; +} else { + lean_dec_ref(x_59); + x_65 = lean_box(0); } +if (lean_is_scalar(x_65)) { + x_66 = lean_alloc_ctor(0, 5, 0); +} else { + x_66 = x_65; +} +lean_ctor_set(x_66, 0, x_61); +lean_ctor_set(x_66, 1, x_3); +lean_ctor_set(x_66, 2, x_62); +lean_ctor_set(x_66, 3, x_63); +lean_ctor_set(x_66, 4, x_64); +x_67 = lean_st_ref_set(x_6, x_66, x_60); +x_68 = lean_ctor_get(x_67, 1); +lean_inc(x_68); +if (lean_is_exclusive(x_67)) { + lean_ctor_release(x_67, 0); + lean_ctor_release(x_67, 1); + x_69 = x_67; +} else { + lean_dec_ref(x_67); + x_69 = lean_box(0); } -static lean_object* _init_l_Lean_Meta_Grind_addEMatchTheorem___closed__7() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_1 = l_Lean_Meta_Grind_EMatchTheorems_insert___closed__1; -x_2 = l_Lean_Meta_Grind_addEMatchTheorem___closed__6; -x_3 = lean_unsigned_to_nat(340u); -x_4 = lean_unsigned_to_nat(2u); -x_5 = l_Lean_Meta_Grind_addEMatchTheorem___closed__5; -x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); -return x_6; +x_70 = lean_box(0); +if (lean_is_scalar(x_69)) { + x_71 = lean_alloc_ctor(0, 2, 0); +} else { + x_71 = x_69; +} +lean_ctor_set(x_71, 0, x_70); +lean_ctor_set(x_71, 1, x_68); +return x_71; } } -static lean_object* _init_l_Lean_Meta_Grind_addEMatchTheorem___closed__8() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("ematch", 6, 6); -return x_1; +lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; +x_72 = lean_ctor_get(x_10, 0); +x_73 = lean_ctor_get(x_10, 1); +lean_inc(x_73); +lean_inc(x_72); +lean_dec(x_10); +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); +x_80 = lean_ctor_get(x_72, 7); +lean_inc(x_80); +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); + lean_ctor_release(x_72, 7); + x_81 = x_72; +} else { + lean_dec_ref(x_72); + x_81 = lean_box(0); +} +x_82 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__4), 2, 1); +lean_closure_set(x_82, 0, x_1); +x_83 = l_Lean_Meta_Grind_addEMatchTheorem___closed__1; +x_84 = l_Lean_ScopedEnvExtension_modifyState___rarg(x_83, x_74, x_82); +lean_inc(x_2); +x_85 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_85, 0, x_2); +lean_ctor_set(x_85, 1, x_2); +if (lean_is_scalar(x_81)) { + x_86 = lean_alloc_ctor(0, 8, 0); +} else { + x_86 = x_81; +} +lean_ctor_set(x_86, 0, x_84); +lean_ctor_set(x_86, 1, x_75); +lean_ctor_set(x_86, 2, x_76); +lean_ctor_set(x_86, 3, x_77); +lean_ctor_set(x_86, 4, x_85); +lean_ctor_set(x_86, 5, x_78); +lean_ctor_set(x_86, 6, x_79); +lean_ctor_set(x_86, 7, x_80); +x_87 = lean_st_ref_set(x_8, x_86, x_73); +x_88 = lean_ctor_get(x_87, 1); +lean_inc(x_88); +lean_dec(x_87); +x_89 = lean_st_ref_take(x_6, 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, 2); +lean_inc(x_93); +x_94 = lean_ctor_get(x_90, 3); +lean_inc(x_94); +x_95 = lean_ctor_get(x_90, 4); +lean_inc(x_95); +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_96 = x_90; +} else { + lean_dec_ref(x_90); + x_96 = lean_box(0); } +if (lean_is_scalar(x_96)) { + x_97 = lean_alloc_ctor(0, 5, 0); +} else { + x_97 = x_96; +} +lean_ctor_set(x_97, 0, x_92); +lean_ctor_set(x_97, 1, x_3); +lean_ctor_set(x_97, 2, x_93); +lean_ctor_set(x_97, 3, x_94); +lean_ctor_set(x_97, 4, x_95); +x_98 = lean_st_ref_set(x_6, x_97, x_91); +x_99 = lean_ctor_get(x_98, 1); +lean_inc(x_99); +if (lean_is_exclusive(x_98)) { + lean_ctor_release(x_98, 0); + lean_ctor_release(x_98, 1); + x_100 = x_98; +} else { + lean_dec_ref(x_98); + x_100 = lean_box(0); } -static lean_object* _init_l_Lean_Meta_Grind_addEMatchTheorem___closed__9() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("pattern", 7, 7); -return x_1; +x_101 = lean_box(0); +if (lean_is_scalar(x_100)) { + x_102 = lean_alloc_ctor(0, 2, 0); +} else { + x_102 = x_100; } +lean_ctor_set(x_102, 0, x_101); +lean_ctor_set(x_102, 1, x_99); +return x_102; } -static lean_object* _init_l_Lean_Meta_Grind_addEMatchTheorem___closed__10() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Meta_Grind_mkGroundPattern___closed__1; -x_2 = l_Lean_Meta_Grind_addEMatchTheorem___closed__8; -x_3 = l_Lean_Meta_Grind_addEMatchTheorem___closed__9; -x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); -return x_4; } } -static lean_object* _init_l_Lean_Meta_Grind_addEMatchTheorem___closed__11() { +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__6___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked(": ", 2, 2); +x_1 = lean_mk_string_unchecked("` is not marked with the `[grind]` attribute", 44, 44); return x_1; } } -static lean_object* _init_l_Lean_Meta_Grind_addEMatchTheorem___closed__12() { +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__6___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_Grind_addEMatchTheorem___closed__11; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__6___closed__1; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addEMatchTheorem(lean_object* x_1, lean_object* x_2, lean_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_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -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; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_27; lean_object* x_28; lean_inc(x_1); -x_9 = l_Lean_getConstInfo___at_Lean_Meta_mkConstWithFreshMVarLevels___spec__1(x_1, x_4, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_9) == 0) +x_5 = l_Lean_MessageData_ofName(x_1); +x_6 = l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___closed__4; +x_7 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_7, 0, x_6); +lean_ctor_set(x_7, 1, x_5); +x_8 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__6___closed__2; +x_9 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_9, 0, x_7); +lean_ctor_set(x_9, 1, x_8); +x_10 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__10; +x_11 = lean_st_mk_ref(x_10, 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_27 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__7; +lean_inc(x_1); +x_28 = l_Lean_getConstInfo___at_Lean_Meta_mkConstWithFreshMVarLevels___spec__1(x_1, x_27, x_12, x_2, x_3, x_13); +if (lean_obj_tag(x_28) == 0) { -lean_object* x_10; -x_10 = lean_ctor_get(x_9, 0); -lean_inc(x_10); -if (lean_obj_tag(x_10) == 2) +lean_object* x_29; lean_object* x_30; uint8_t x_31; +x_29 = lean_ctor_get(x_28, 0); +lean_inc(x_29); +x_30 = lean_ctor_get(x_28, 1); +lean_inc(x_30); +lean_dec(x_28); +x_31 = l_Lean_ConstantInfo_isTheorem(x_29); +lean_dec(x_29); +if (x_31 == 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_9, 1); -lean_inc(x_11); -lean_dec(x_9); -x_12 = lean_ctor_get(x_10, 0); +lean_object* x_32; +lean_inc(x_3); +lean_inc(x_2); lean_inc(x_12); -lean_dec(x_10); -x_13 = lean_ctor_get(x_12, 0); -lean_inc(x_13); +x_32 = l_Lean_Meta_getEqnsFor_x3f(x_1, x_27, x_12, x_2, x_3, x_30); +if (lean_obj_tag(x_32) == 0) +{ +lean_object* x_33; +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; uint8_t x_36; +x_34 = lean_ctor_get(x_32, 1); +lean_inc(x_34); +lean_dec(x_32); +x_35 = l_Lean_throwError___at___private_Lean_Meta_InferType_0__Lean_Meta_inferProjType___spec__1(x_9, x_27, x_12, x_2, x_3, x_34); +lean_dec(x_3); +lean_dec(x_2); lean_dec(x_12); -x_14 = lean_ctor_get(x_13, 1); -lean_inc(x_14); -lean_dec(x_13); -x_15 = lean_box(0); -x_16 = l_List_mapTR_loop___at_Lean_mkConstWithLevelParams___spec__1(x_14, x_15); -lean_inc(x_1); -x_17 = l_Lean_Expr_const___override(x_1, x_16); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -x_18 = l_Lean_Meta_Grind_NormalizePattern_main(x_3, x_4, x_5, x_6, x_7, x_11); -if (lean_obj_tag(x_18) == 0) +x_36 = !lean_is_exclusive(x_35); +if (x_36 == 0) { -lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; -x_19 = lean_ctor_get(x_18, 0); -lean_inc(x_19); -x_20 = lean_ctor_get(x_19, 1); -lean_inc(x_20); -x_21 = lean_ctor_get(x_18, 1); -lean_inc(x_21); -lean_dec(x_18); -x_22 = !lean_is_exclusive(x_19); -if (x_22 == 0) +return x_35; +} +else { -lean_object* x_23; lean_object* x_24; uint8_t x_25; -x_23 = lean_ctor_get(x_19, 0); -x_24 = lean_ctor_get(x_19, 1); -lean_dec(x_24); -x_25 = !lean_is_exclusive(x_20); -if (x_25 == 0) +lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_37 = lean_ctor_get(x_35, 0); +x_38 = lean_ctor_get(x_35, 1); +lean_inc(x_38); +lean_inc(x_37); +lean_dec(x_35); +x_39 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_39, 0, x_37); +lean_ctor_set(x_39, 1, x_38); +return x_39; +} +} +else { -lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; -x_26 = lean_ctor_get(x_20, 0); -x_27 = lean_ctor_get(x_20, 1); -x_28 = l_Lean_Meta_Grind_addEMatchTheorem___closed__3; -lean_inc(x_26); -x_29 = l_List_all___rarg(x_26, x_28); -if (x_29 == 0) +lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; uint8_t x_51; +x_40 = lean_ctor_get(x_32, 1); +lean_inc(x_40); +lean_dec(x_32); +x_41 = lean_ctor_get(x_33, 0); +lean_inc(x_41); +lean_dec(x_33); +x_42 = lean_st_ref_get(x_3, x_40); +x_43 = lean_ctor_get(x_42, 0); +lean_inc(x_43); +x_44 = lean_ctor_get(x_42, 1); +lean_inc(x_44); +lean_dec(x_42); +x_45 = lean_ctor_get(x_43, 0); +lean_inc(x_45); +lean_dec(x_43); +x_46 = l_Lean_Meta_Grind_instInhabitedEMatchTheorems; +x_47 = l_Lean_Meta_Grind_addEMatchTheorem___closed__1; +x_48 = l_Lean_ScopedEnvExtension_getState___rarg(x_46, x_47, x_45); +lean_dec(x_45); +x_49 = lean_array_get_size(x_41); +x_50 = lean_unsigned_to_nat(0u); +x_51 = lean_nat_dec_lt(x_50, x_49); +if (x_51 == 0) { -lean_object* x_30; lean_object* x_31; -lean_free_object(x_20); -lean_dec(x_27); -lean_dec(x_26); -lean_free_object(x_19); -lean_dec(x_23); -lean_dec(x_17); +lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; +lean_dec(x_49); +lean_dec(x_48); +lean_dec(x_9); +x_52 = l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedEMatchTheorems___spec__1___closed__2; +x_53 = l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__1___closed__2; +x_54 = lean_box(0); +x_55 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__3(x_41, x_52, x_53, x_54, x_27, x_12, x_2, x_3, x_44); +lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -x_30 = l_Lean_Meta_Grind_addEMatchTheorem___closed__7; -x_31 = l_panic___at_Lean_Meta_Grind_addEMatchTheorem___spec__1(x_30, x_4, x_5, x_6, x_7, x_21); -return x_31; +x_14 = x_55; +goto block_26; } else { -lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; -x_32 = l_Lean_Meta_Grind_addEMatchTheorem___closed__10; -x_33 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_32, x_4, x_5, x_6, x_7, x_21); -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) +size_t x_56; size_t x_57; uint8_t x_58; +x_56 = 0; +x_57 = lean_usize_of_nat(x_49); +lean_dec(x_49); +x_58 = l_Array_anyMUnsafe_any___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____spec__2(x_48, x_41, x_56, x_57); +if (x_58 == 0) { -lean_object* x_36; lean_object* x_37; lean_object* x_38; -lean_free_object(x_20); -lean_free_object(x_19); -x_36 = lean_ctor_get(x_33, 1); -lean_inc(x_36); -lean_dec(x_33); -x_37 = lean_box(0); -x_38 = l_Lean_Meta_Grind_addEMatchTheorem___lambda__3(x_17, x_2, x_27, x_15, x_1, x_23, x_26, x_37, x_4, x_5, x_6, x_7, x_36); -return x_38; +lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; +lean_dec(x_9); +x_59 = l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedEMatchTheorems___spec__1___closed__2; +x_60 = l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__1___closed__2; +x_61 = lean_box(0); +x_62 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__3(x_41, x_59, x_60, x_61, x_27, x_12, x_2, x_3, x_44); +lean_dec(x_3); +lean_dec(x_2); +x_14 = x_62; +goto block_26; } else { -uint8_t x_39; -x_39 = !lean_is_exclusive(x_33); -if (x_39 == 0) -{ -lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; -x_40 = lean_ctor_get(x_33, 1); -x_41 = lean_ctor_get(x_33, 0); +lean_object* x_63; uint8_t x_64; lean_dec(x_41); -lean_inc(x_1); -x_42 = l_Lean_MessageData_ofName(x_1); -x_43 = l_Lean_Meta_Grind_ppPattern___closed__4; -lean_ctor_set_tag(x_33, 7); -lean_ctor_set(x_33, 1, x_42); -lean_ctor_set(x_33, 0, x_43); -x_44 = l_Lean_Meta_Grind_addEMatchTheorem___closed__12; -lean_ctor_set_tag(x_20, 7); -lean_ctor_set(x_20, 1, x_44); -lean_ctor_set(x_20, 0, x_33); -lean_inc(x_23); -x_45 = l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__3(x_23, x_15); -x_46 = l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__4(x_45, x_15); -x_47 = l_Lean_MessageData_ofList(x_46); -lean_ctor_set_tag(x_19, 7); -lean_ctor_set(x_19, 1, x_47); -lean_ctor_set(x_19, 0, x_20); -x_48 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_48, 0, x_19); -lean_ctor_set(x_48, 1, x_43); -x_49 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_32, x_48, x_4, x_5, x_6, x_7, x_40); -x_50 = lean_ctor_get(x_49, 0); -lean_inc(x_50); -x_51 = lean_ctor_get(x_49, 1); -lean_inc(x_51); -lean_dec(x_49); -x_52 = l_Lean_Meta_Grind_addEMatchTheorem___lambda__3(x_17, x_2, x_27, x_15, x_1, x_23, x_26, x_50, x_4, x_5, x_6, x_7, x_51); -lean_dec(x_50); -return x_52; +x_63 = l_Lean_throwError___at___private_Lean_Meta_InferType_0__Lean_Meta_inferProjType___spec__1(x_9, x_27, x_12, x_2, x_3, x_44); +lean_dec(x_3); +lean_dec(x_2); +x_64 = !lean_is_exclusive(x_63); +if (x_64 == 0) +{ +x_14 = x_63; +goto block_26; } else { -lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; -x_53 = lean_ctor_get(x_33, 1); -lean_inc(x_53); -lean_dec(x_33); -lean_inc(x_1); -x_54 = l_Lean_MessageData_ofName(x_1); -x_55 = l_Lean_Meta_Grind_ppPattern___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 = l_Lean_Meta_Grind_addEMatchTheorem___closed__12; -lean_ctor_set_tag(x_20, 7); -lean_ctor_set(x_20, 1, x_57); -lean_ctor_set(x_20, 0, x_56); -lean_inc(x_23); -x_58 = l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__3(x_23, x_15); -x_59 = l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__4(x_58, x_15); -x_60 = l_Lean_MessageData_ofList(x_59); -lean_ctor_set_tag(x_19, 7); -lean_ctor_set(x_19, 1, x_60); -lean_ctor_set(x_19, 0, x_20); -x_61 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_61, 0, x_19); -lean_ctor_set(x_61, 1, x_55); -x_62 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_32, x_61, x_4, x_5, x_6, x_7, x_53); -x_63 = lean_ctor_get(x_62, 0); -lean_inc(x_63); -x_64 = lean_ctor_get(x_62, 1); -lean_inc(x_64); -lean_dec(x_62); -x_65 = l_Lean_Meta_Grind_addEMatchTheorem___lambda__3(x_17, x_2, x_27, x_15, x_1, x_23, x_26, x_63, x_4, x_5, x_6, x_7, x_64); +lean_object* x_65; lean_object* x_66; lean_object* x_67; +x_65 = lean_ctor_get(x_63, 0); +x_66 = lean_ctor_get(x_63, 1); +lean_inc(x_66); +lean_inc(x_65); lean_dec(x_63); -return x_65; +x_67 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_67, 0, x_65); +lean_ctor_set(x_67, 1, x_66); +x_14 = x_67; +goto block_26; +} } } } } else { -lean_object* x_66; lean_object* x_67; lean_object* x_68; uint8_t x_69; -x_66 = lean_ctor_get(x_20, 0); -x_67 = lean_ctor_get(x_20, 1); -lean_inc(x_67); -lean_inc(x_66); -lean_dec(x_20); -x_68 = l_Lean_Meta_Grind_addEMatchTheorem___closed__3; -lean_inc(x_66); -x_69 = l_List_all___rarg(x_66, x_68); -if (x_69 == 0) -{ -lean_object* x_70; lean_object* x_71; -lean_dec(x_67); -lean_dec(x_66); -lean_free_object(x_19); -lean_dec(x_23); -lean_dec(x_17); +uint8_t x_68; +lean_dec(x_12); +lean_dec(x_9); +lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -x_70 = l_Lean_Meta_Grind_addEMatchTheorem___closed__7; -x_71 = l_panic___at_Lean_Meta_Grind_addEMatchTheorem___spec__1(x_70, x_4, x_5, x_6, x_7, x_21); +x_68 = !lean_is_exclusive(x_32); +if (x_68 == 0) +{ +return x_32; +} +else +{ +lean_object* x_69; lean_object* x_70; lean_object* x_71; +x_69 = lean_ctor_get(x_32, 0); +x_70 = lean_ctor_get(x_32, 1); +lean_inc(x_70); +lean_inc(x_69); +lean_dec(x_32); +x_71 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_71, 0, x_69); +lean_ctor_set(x_71, 1, x_70); return x_71; } +} +} else { -lean_object* x_72; lean_object* x_73; lean_object* x_74; uint8_t x_75; -x_72 = l_Lean_Meta_Grind_addEMatchTheorem___closed__10; -x_73 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_72, x_4, x_5, x_6, x_7, x_21); -x_74 = lean_ctor_get(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; lean_object* x_78; lean_object* x_79; uint8_t x_80; +x_72 = lean_st_ref_get(x_3, x_30); +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_unbox(x_74); -lean_dec(x_74); -if (x_75 == 0) -{ -lean_object* x_76; lean_object* x_77; lean_object* x_78; -lean_free_object(x_19); -x_76 = lean_ctor_get(x_73, 1); -lean_inc(x_76); +lean_dec(x_72); +x_75 = lean_ctor_get(x_73, 0); +lean_inc(x_75); lean_dec(x_73); -x_77 = lean_box(0); -x_78 = l_Lean_Meta_Grind_addEMatchTheorem___lambda__3(x_17, x_2, x_67, x_15, x_1, x_23, x_66, x_77, x_4, x_5, x_6, x_7, x_76); -return x_78; +x_76 = l_Lean_Meta_Grind_instInhabitedEMatchTheorems; +x_77 = l_Lean_Meta_Grind_addEMatchTheorem___closed__1; +x_78 = l_Lean_ScopedEnvExtension_getState___rarg(x_76, x_77, x_75); +lean_dec(x_75); +lean_inc(x_1); +x_79 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_79, 0, x_1); +x_80 = l_Lean_Meta_Grind_EMatchTheorems_contains(x_78, x_79); +lean_dec(x_79); +if (x_80 == 0) +{ +lean_object* x_81; uint8_t x_82; +lean_dec(x_1); +x_81 = l_Lean_throwError___at___private_Lean_Meta_InferType_0__Lean_Meta_inferProjType___spec__1(x_9, x_27, x_12, x_2, x_3, x_74); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_12); +x_82 = !lean_is_exclusive(x_81); +if (x_82 == 0) +{ +return x_81; } else { -lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; -x_79 = lean_ctor_get(x_73, 1); -lean_inc(x_79); -if (lean_is_exclusive(x_73)) { - lean_ctor_release(x_73, 0); - lean_ctor_release(x_73, 1); - x_80 = x_73; -} else { - lean_dec_ref(x_73); - x_80 = lean_box(0); -} -lean_inc(x_1); -x_81 = l_Lean_MessageData_ofName(x_1); -x_82 = l_Lean_Meta_Grind_ppPattern___closed__4; -if (lean_is_scalar(x_80)) { - x_83 = lean_alloc_ctor(7, 2, 0); -} else { - x_83 = x_80; - lean_ctor_set_tag(x_83, 7); -} -lean_ctor_set(x_83, 0, x_82); -lean_ctor_set(x_83, 1, x_81); -x_84 = l_Lean_Meta_Grind_addEMatchTheorem___closed__12; -x_85 = lean_alloc_ctor(7, 2, 0); +lean_object* x_83; lean_object* x_84; lean_object* x_85; +x_83 = lean_ctor_get(x_81, 0); +x_84 = lean_ctor_get(x_81, 1); +lean_inc(x_84); +lean_inc(x_83); +lean_dec(x_81); +x_85 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_85, 0, x_83); lean_ctor_set(x_85, 1, x_84); -lean_inc(x_23); -x_86 = l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__3(x_23, x_15); -x_87 = l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__4(x_86, x_15); -x_88 = l_Lean_MessageData_ofList(x_87); -lean_ctor_set_tag(x_19, 7); -lean_ctor_set(x_19, 1, x_88); -lean_ctor_set(x_19, 0, x_85); -x_89 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_89, 0, x_19); -lean_ctor_set(x_89, 1, x_82); -x_90 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_72, x_89, x_4, x_5, x_6, x_7, x_79); -x_91 = lean_ctor_get(x_90, 0); +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; uint8_t x_93; +lean_dec(x_9); +x_86 = l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedEMatchTheorems___spec__1___closed__2; +x_87 = l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__1___closed__2; +x_88 = lean_box(0); +x_89 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__5(x_1, x_86, x_87, x_88, x_27, x_12, x_2, x_3, x_74); +lean_dec(x_3); +lean_dec(x_2); +x_90 = lean_ctor_get(x_89, 0); +lean_inc(x_90); +x_91 = lean_ctor_get(x_89, 1); lean_inc(x_91); -x_92 = lean_ctor_get(x_90, 1); -lean_inc(x_92); -lean_dec(x_90); -x_93 = l_Lean_Meta_Grind_addEMatchTheorem___lambda__3(x_17, x_2, x_67, x_15, x_1, x_23, x_66, x_91, x_4, x_5, x_6, x_7, x_92); -lean_dec(x_91); -return x_93; +lean_dec(x_89); +x_92 = lean_st_ref_get(x_12, x_91); +lean_dec(x_12); +x_93 = !lean_is_exclusive(x_92); +if (x_93 == 0) +{ +lean_object* x_94; +x_94 = lean_ctor_get(x_92, 0); +lean_dec(x_94); +lean_ctor_set(x_92, 0, x_90); +return x_92; +} +else +{ +lean_object* x_95; lean_object* x_96; +x_95 = lean_ctor_get(x_92, 1); +lean_inc(x_95); +lean_dec(x_92); +x_96 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_96, 0, x_90); +lean_ctor_set(x_96, 1, x_95); +return x_96; } } } } else { -lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; uint8_t x_99; -x_94 = lean_ctor_get(x_19, 0); -lean_inc(x_94); +uint8_t x_97; +lean_dec(x_12); +lean_dec(x_9); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_97 = !lean_is_exclusive(x_28); +if (x_97 == 0) +{ +return x_28; +} +else +{ +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; +} +} +block_26: +{ +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_st_ref_get(x_12, x_16); +lean_dec(x_12); +x_18 = !lean_is_exclusive(x_17); +if (x_18 == 0) +{ +lean_object* x_19; +x_19 = lean_ctor_get(x_17, 0); lean_dec(x_19); -x_95 = lean_ctor_get(x_20, 0); -lean_inc(x_95); -x_96 = lean_ctor_get(x_20, 1); -lean_inc(x_96); -if (lean_is_exclusive(x_20)) { - lean_ctor_release(x_20, 0); - lean_ctor_release(x_20, 1); - x_97 = x_20; -} else { - lean_dec_ref(x_20); - x_97 = lean_box(0); +lean_ctor_set(x_17, 0, x_15); +return x_17; } -x_98 = l_Lean_Meta_Grind_addEMatchTheorem___closed__3; -lean_inc(x_95); -x_99 = l_List_all___rarg(x_95, x_98); -if (x_99 == 0) +else { -lean_object* x_100; lean_object* x_101; -lean_dec(x_97); -lean_dec(x_96); -lean_dec(x_95); -lean_dec(x_94); +lean_object* x_20; lean_object* x_21; +x_20 = lean_ctor_get(x_17, 1); +lean_inc(x_20); lean_dec(x_17); -lean_dec(x_2); -lean_dec(x_1); -x_100 = l_Lean_Meta_Grind_addEMatchTheorem___closed__7; -x_101 = l_panic___at_Lean_Meta_Grind_addEMatchTheorem___spec__1(x_100, x_4, x_5, x_6, x_7, x_21); -return x_101; +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_15); +lean_ctor_set(x_21, 1, x_20); +return x_21; +} +} +else +{ +uint8_t x_22; +lean_dec(x_12); +x_22 = !lean_is_exclusive(x_14); +if (x_22 == 0) +{ +return x_14; +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_14, 0); +x_24 = lean_ctor_get(x_14, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_14); +x_25 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_24); +return x_25; +} +} +} +} +} +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind_mkOffsetPattern___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__1; +x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__4; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } -else -{ -lean_object* x_102; lean_object* x_103; lean_object* x_104; uint8_t x_105; -x_102 = l_Lean_Meta_Grind_addEMatchTheorem___closed__10; -x_103 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_102, x_4, x_5, x_6, x_7, x_21); -x_104 = lean_ctor_get(x_103, 0); -lean_inc(x_104); -x_105 = lean_unbox(x_104); -lean_dec(x_104); -if (x_105 == 0) +} +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__3() { +_start: { -lean_object* x_106; lean_object* x_107; lean_object* x_108; -lean_dec(x_97); -x_106 = lean_ctor_get(x_103, 1); -lean_inc(x_106); -lean_dec(x_103); -x_107 = lean_box(0); -x_108 = l_Lean_Meta_Grind_addEMatchTheorem___lambda__3(x_17, x_2, x_96, x_15, x_1, x_94, x_95, x_107, x_4, x_5, x_6, x_7, x_106); -return x_108; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__2; +x_2 = l_Lean_Meta_Grind_mkOffsetPattern___closed__2; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } -else +} +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__4() { +_start: { -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_ctor_get(x_103, 1); -lean_inc(x_109); -if (lean_is_exclusive(x_103)) { - lean_ctor_release(x_103, 0); - lean_ctor_release(x_103, 1); - x_110 = x_103; -} else { - lean_dec_ref(x_103); - x_110 = lean_box(0); +lean_object* x_1; +x_1 = lean_mk_string_unchecked("initFn", 6, 6); +return x_1; } -lean_inc(x_1); -x_111 = l_Lean_MessageData_ofName(x_1); -x_112 = l_Lean_Meta_Grind_ppPattern___closed__4; -if (lean_is_scalar(x_110)) { - x_113 = lean_alloc_ctor(7, 2, 0); -} else { - x_113 = x_110; - lean_ctor_set_tag(x_113, 7); } -lean_ctor_set(x_113, 0, x_112); -lean_ctor_set(x_113, 1, x_111); -x_114 = l_Lean_Meta_Grind_addEMatchTheorem___closed__12; -if (lean_is_scalar(x_97)) { - x_115 = lean_alloc_ctor(7, 2, 0); -} else { - x_115 = x_97; - lean_ctor_set_tag(x_115, 7); +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__3; +x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__4; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } -lean_ctor_set(x_115, 0, x_113); -lean_ctor_set(x_115, 1, x_114); -lean_inc(x_94); -x_116 = l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__3(x_94, x_15); -x_117 = l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__4(x_116, x_15); -x_118 = l_Lean_MessageData_ofList(x_117); -x_119 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_119, 0, x_115); -lean_ctor_set(x_119, 1, x_118); -x_120 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_120, 0, x_119); -lean_ctor_set(x_120, 1, x_112); -x_121 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_102, x_120, x_4, x_5, x_6, x_7, x_109); -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 = l_Lean_Meta_Grind_addEMatchTheorem___lambda__3(x_17, x_2, x_96, x_15, x_1, x_94, x_95, x_122, x_4, x_5, x_6, x_7, x_123); -lean_dec(x_122); -return x_124; } +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("_@", 2, 2); +return x_1; } } +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__5; +x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__6; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } -else +} +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__8() { +_start: { -uint8_t x_125; -lean_dec(x_17); -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_125 = !lean_is_exclusive(x_18); -if (x_125 == 0) +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__7; +x_2 = l_Lean_Meta_Grind_mkOffsetPattern___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__9() { +_start: { -return x_18; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__8; +x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__4; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } -else +} +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__10() { +_start: { -lean_object* x_126; lean_object* x_127; lean_object* x_128; -x_126 = lean_ctor_get(x_18, 0); -x_127 = lean_ctor_get(x_18, 1); -lean_inc(x_127); -lean_inc(x_126); -lean_dec(x_18); -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_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__9; +x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__6; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} } +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__10; +x_2 = l_Lean_Meta_Grind_mkOffsetPattern___closed__2; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } } -else +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__12() { +_start: { -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_dec(x_10); -lean_dec(x_3); -lean_dec(x_2); -x_129 = lean_ctor_get(x_9, 1); -lean_inc(x_129); -lean_dec(x_9); -x_130 = l_Lean_MessageData_ofName(x_1); -x_131 = l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__4; -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_Meta_Grind_addEMatchTheorem___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_throwError___at_Lean_Meta_setInlineAttribute___spec__1(x_134, x_4, x_5, x_6, x_7, x_129); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -return x_135; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__11; +x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__9; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } } -else +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__13() { +_start: { -uint8_t x_136; -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_136 = !lean_is_exclusive(x_9); -if (x_136 == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("_hyg", 4, 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__14() { +_start: { -return x_9; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__12; +x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__13; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } -else +} +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__15() { +_start: { -lean_object* x_137; lean_object* x_138; lean_object* x_139; -x_137 = lean_ctor_get(x_9, 0); -x_138 = lean_ctor_get(x_9, 1); -lean_inc(x_138); -lean_inc(x_137); -lean_dec(x_9); -x_139 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_139, 0, x_137); -lean_ctor_set(x_139, 1, x_138); -return x_139; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__14; +x_2 = lean_unsigned_to_nat(10279u); +x_3 = l_Lean_Name_num___override(x_1, x_2); +return x_3; } } +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__16() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind_mkGroundPattern___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } } -LEAN_EXPORT lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___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) { +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__17() { _start: { -uint8_t x_9; lean_object* x_10; -x_9 = lean_unbox(x_3); -lean_dec(x_3); -x_10 = l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___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_5); -lean_dec(x_4); -return x_10; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("The `[grind]` attribute is used to annotate declarations.When applied to an equational theorem, `[grind =]`, `[grind =_]`, or `[grind _=_]`will mark the theorem for use in heuristic instantiations by the `grind` tactic,\n using respectively the left-hand side, the right-hand side, or both sides of the theorem.When applied to a function, `[grind =]` automatically annotates the equational theorems associated with that function.When applied to a theorem `[grind ←]` will instantiate the theorem whenever it encounters the conclusion of the theorem\n (that is, it will use the theorem for backwards reasoning).When applied to a theorem `[grind →]` will instantiate the theorem whenever it encounters sufficiently many of the propositional hypotheses\n (that is, it will use the theorem for forwards reasoning).The attribute `[grind]` by itself will effectively try `[grind ←]` (if the conclusion is sufficient for instantiation) and then `[grind →]`.The `grind` tactic utilizes annotated theorems to add instances of matching patterns into the local context during proof search.For example, if a theorem `@[grind =] theorem foo_idempotent : foo (foo x) = foo x` is annotated,`grind` will add an instance of this theorem to the local context whenever it encounters the pattern `foo (foo x)`.", 1310, 1302); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addEMatchTheorem___lambda__1___boxed(lean_object* x_1) { +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__18() { _start: { -uint8_t x_2; lean_object* x_3; -x_2 = l_Lean_Meta_Grind_addEMatchTheorem___lambda__1(x_1); -lean_dec(x_1); -x_3 = lean_box(x_2); -return x_3; +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_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__15; +x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__16; +x_3 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__17; +x_4 = 1; +x_5 = lean_alloc_ctor(0, 3, 1); +lean_ctor_set(x_5, 0, x_1); +lean_ctor_set(x_5, 1, x_2); +lean_ctor_set(x_5, 2, x_3); +lean_ctor_set_uint8(x_5, sizeof(void*)*3, x_4); +return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addEMatchTheorem___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__19() { _start: { -lean_object* x_13; -x_13 = l_Lean_Meta_Grind_addEMatchTheorem___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_11); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -return x_13; +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___boxed), 6, 0); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addEMatchTheorem___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) { +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__20() { _start: { -lean_object* x_14; -x_14 = l_Lean_Meta_Grind_addEMatchTheorem___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -lean_dec(x_8); -return x_14; +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__6), 4, 0); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getEMatchTheorems___rarg(lean_object* x_1, lean_object* x_2) { +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__21() { _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_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__18; +x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__19; +x_3 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__20; +x_4 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279_(lean_object* x_1) { +_start: { -lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; -x_5 = lean_ctor_get(x_3, 0); -x_6 = lean_ctor_get(x_5, 0); -lean_inc(x_6); -lean_dec(x_5); -x_7 = l_panic___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__1___closed__1; -x_8 = l_Lean_Meta_Grind_addEMatchTheorem___lambda__2___closed__1; -x_9 = l_Lean_ScopedEnvExtension_getState___rarg(x_7, x_8, x_6); -lean_dec(x_6); -lean_ctor_set(x_3, 0, x_9); +lean_object* x_2; lean_object* x_3; +x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__21; +x_3 = l_Lean_registerBuiltinAttribute(x_2, x_1); return x_3; } -else +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: { -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_10 = lean_ctor_get(x_3, 0); -x_11 = lean_ctor_get(x_3, 1); -lean_inc(x_11); -lean_inc(x_10); +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_12 = lean_ctor_get(x_10, 0); -lean_inc(x_12); -lean_dec(x_10); -x_13 = l_panic___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__1___closed__1; -x_14 = l_Lean_Meta_Grind_addEMatchTheorem___lambda__2___closed__1; -x_15 = l_Lean_ScopedEnvExtension_getState___rarg(x_13, x_14, x_12); -lean_dec(x_12); -x_16 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_16, 0, x_15); -lean_ctor_set(x_16, 1, x_11); -return x_16; +x_7 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____spec__1(x_1, x_5, x_6, x_4); +lean_dec(x_1); +return x_7; +} } +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____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; uint8_t x_7; lean_object* x_8; +x_5 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_6 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_7 = l_Array_anyMUnsafe_any___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____spec__2(x_1, x_2, x_5, x_6); +lean_dec(x_2); +x_8 = lean_box(x_7); +return x_8; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getEMatchTheorems(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_getEMatchTheorems___rarg___boxed), 2, 0); -return x_2; +uint8_t x_7; lean_object* x_8; +x_7 = lean_unbox(x_3); +lean_dec(x_3); +x_8 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1(x_1, x_2, x_7, x_4, x_5, x_6); +lean_dec(x_2); +return x_8; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getEMatchTheorems___rarg___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__2___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Meta_Grind_getEMatchTheorems___rarg(x_1, x_2); +x_3 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__2(x_1, x_2); lean_dec(x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getEMatchTheorems___boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____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_2; -x_2 = l_Lean_Meta_Grind_getEMatchTheorems(x_1); -lean_dec(x_1); -return x_2; +lean_object* x_10; +x_10 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____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_dec(x_4); +return x_10; } } +lean_object* initialize_Init_Grind_Util(uint8_t builtin, lean_object*); +lean_object* initialize_Init_Grind_Tactics(uint8_t builtin, lean_object*); lean_object* initialize_Lean_HeadIndex(uint8_t builtin, lean_object*); lean_object* initialize_Lean_PrettyPrinter(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Util_FoldConsts(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Util_CollectFVars(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_Basic(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_InferType(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Meta_Eqns(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Meta_Tactic_Grind_Util(uint8_t builtin, lean_object*); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Lean_Meta_Tactic_Grind_EMatchTheorem(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_Grind_Util(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Init_Grind_Tactics(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); res = initialize_Lean_HeadIndex(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); @@ -18175,58 +29800,80 @@ lean_dec_ref(res); res = initialize_Lean_Meta_InferType(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +res = initialize_Lean_Meta_Eqns(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Lean_Meta_Tactic_Grind_Util(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l_Lean_Meta_Grind_mkOffsetPattern___closed__1 = _init_l_Lean_Meta_Grind_mkOffsetPattern___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_mkOffsetPattern___closed__1); +l_Lean_Meta_Grind_mkOffsetPattern___closed__2 = _init_l_Lean_Meta_Grind_mkOffsetPattern___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_mkOffsetPattern___closed__2); +l_Lean_Meta_Grind_mkOffsetPattern___closed__3 = _init_l_Lean_Meta_Grind_mkOffsetPattern___closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind_mkOffsetPattern___closed__3); +l_Lean_Meta_Grind_mkOffsetPattern___closed__4 = _init_l_Lean_Meta_Grind_mkOffsetPattern___closed__4(); +lean_mark_persistent(l_Lean_Meta_Grind_mkOffsetPattern___closed__4); +l_Lean_Meta_Grind_mkOffsetPattern___closed__5 = _init_l_Lean_Meta_Grind_mkOffsetPattern___closed__5(); +lean_mark_persistent(l_Lean_Meta_Grind_mkOffsetPattern___closed__5); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_detectOffsets___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_detectOffsets___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_detectOffsets___closed__1); +l_Lean_Meta_Grind_isOffsetPattern_x3f___closed__1 = _init_l_Lean_Meta_Grind_isOffsetPattern_x3f___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_isOffsetPattern_x3f___closed__1); +l_Lean_Meta_Grind_isOffsetPattern_x3f___closed__2 = _init_l_Lean_Meta_Grind_isOffsetPattern_x3f___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_isOffsetPattern_x3f___closed__2); +l_Lean_Meta_Grind_isOffsetPattern_x3f___closed__3 = _init_l_Lean_Meta_Grind_isOffsetPattern_x3f___closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind_isOffsetPattern_x3f___closed__3); +l_Lean_Meta_Grind_preprocessPattern___lambda__1___closed__1 = _init_l_Lean_Meta_Grind_preprocessPattern___lambda__1___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_preprocessPattern___lambda__1___closed__1); +l_Lean_Meta_Grind_preprocessPattern___lambda__1___closed__2 = _init_l_Lean_Meta_Grind_preprocessPattern___lambda__1___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_preprocessPattern___lambda__1___closed__2); +l_Lean_Meta_Grind_preprocessPattern___closed__1 = _init_l_Lean_Meta_Grind_preprocessPattern___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_preprocessPattern___closed__1); +l_Lean_Meta_Grind_preprocessPattern___closed__2 = _init_l_Lean_Meta_Grind_preprocessPattern___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_preprocessPattern___closed__2); +l_Lean_Meta_Grind_preprocessPattern___closed__3 = _init_l_Lean_Meta_Grind_preprocessPattern___closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind_preprocessPattern___closed__3); l_Lean_Meta_Grind_instInhabitedOrigin___closed__1 = _init_l_Lean_Meta_Grind_instInhabitedOrigin___closed__1(); lean_mark_persistent(l_Lean_Meta_Grind_instInhabitedOrigin___closed__1); l_Lean_Meta_Grind_instInhabitedOrigin = _init_l_Lean_Meta_Grind_instInhabitedOrigin(); lean_mark_persistent(l_Lean_Meta_Grind_instInhabitedOrigin); -l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__1(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__1); -l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__2(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__2); -l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__3(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__3); -l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__4(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__4); -l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__5 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__5(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__5); -l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__6 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__6(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__6); -l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__7 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__7(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__7); -l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__8 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__8(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__8); -l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__9 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__9(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__9); -l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__10 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__10(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__10); -l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__11 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__11(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__11); -l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__12 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__12(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__12); -l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__13 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__13(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__13); -l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__14 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__14(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__14); -l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__15 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__15(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__15); -l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__16 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__16(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__16); -l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__17 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__17(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__17); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__1); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__2); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__3); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__4(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__4); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__5 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__5(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__5); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__6 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__6(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__6); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__7 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__7(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__7); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__8 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__8(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__8); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__9 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__9(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__9); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__10 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__10(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__10); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__11 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__11(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__11); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__12 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__12(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__12); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__13 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__13(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__13); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__14 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__14(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__14); l_Lean_Meta_Grind_instReprOrigin___closed__1 = _init_l_Lean_Meta_Grind_instReprOrigin___closed__1(); lean_mark_persistent(l_Lean_Meta_Grind_instReprOrigin___closed__1); l_Lean_Meta_Grind_instReprOrigin = _init_l_Lean_Meta_Grind_instReprOrigin(); lean_mark_persistent(l_Lean_Meta_Grind_instReprOrigin); -l_Lean_Meta_Grind_Origin_key___closed__1 = _init_l_Lean_Meta_Grind_Origin_key___closed__1(); -lean_mark_persistent(l_Lean_Meta_Grind_Origin_key___closed__1); -l_Lean_Meta_Grind_Origin_key___closed__2 = _init_l_Lean_Meta_Grind_Origin_key___closed__2(); -lean_mark_persistent(l_Lean_Meta_Grind_Origin_key___closed__2); -l_Lean_Meta_Grind_Origin_pp___rarg___closed__1 = _init_l_Lean_Meta_Grind_Origin_pp___rarg___closed__1(); -lean_mark_persistent(l_Lean_Meta_Grind_Origin_pp___rarg___closed__1); -l_Lean_Meta_Grind_Origin_pp___rarg___closed__2 = _init_l_Lean_Meta_Grind_Origin_pp___rarg___closed__2(); -lean_mark_persistent(l_Lean_Meta_Grind_Origin_pp___rarg___closed__2); -l_Lean_Meta_Grind_Origin_pp___rarg___closed__3 = _init_l_Lean_Meta_Grind_Origin_pp___rarg___closed__3(); -lean_mark_persistent(l_Lean_Meta_Grind_Origin_pp___rarg___closed__3); +l_Lean_Meta_Grind_instBEqOrigin___closed__1 = _init_l_Lean_Meta_Grind_instBEqOrigin___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_instBEqOrigin___closed__1); +l_Lean_Meta_Grind_instBEqOrigin = _init_l_Lean_Meta_Grind_instBEqOrigin(); +lean_mark_persistent(l_Lean_Meta_Grind_instBEqOrigin); l_Lean_Meta_Grind_instInhabitedEMatchTheorem___closed__1 = _init_l_Lean_Meta_Grind_instInhabitedEMatchTheorem___closed__1(); lean_mark_persistent(l_Lean_Meta_Grind_instInhabitedEMatchTheorem___closed__1); l_Lean_Meta_Grind_instInhabitedEMatchTheorem___closed__2 = _init_l_Lean_Meta_Grind_instInhabitedEMatchTheorem___closed__2(); @@ -18235,14 +29882,20 @@ l_Lean_Meta_Grind_instInhabitedEMatchTheorem___closed__3 = _init_l_Lean_Meta_Gri lean_mark_persistent(l_Lean_Meta_Grind_instInhabitedEMatchTheorem___closed__3); l_Lean_Meta_Grind_instInhabitedEMatchTheorem = _init_l_Lean_Meta_Grind_instInhabitedEMatchTheorem(); lean_mark_persistent(l_Lean_Meta_Grind_instInhabitedEMatchTheorem); +l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedEMatchTheorems___spec__1___closed__1 = _init_l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedEMatchTheorems___spec__1___closed__1(); +lean_mark_persistent(l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedEMatchTheorems___spec__1___closed__1); +l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedEMatchTheorems___spec__1___closed__2 = _init_l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedEMatchTheorems___spec__1___closed__2(); +lean_mark_persistent(l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedEMatchTheorems___spec__1___closed__2); +l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedEMatchTheorems___spec__1 = _init_l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedEMatchTheorems___spec__1(); +lean_mark_persistent(l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedEMatchTheorems___spec__1); +l_Lean_Meta_Grind_instInhabitedEMatchTheorems = _init_l_Lean_Meta_Grind_instInhabitedEMatchTheorems(); +lean_mark_persistent(l_Lean_Meta_Grind_instInhabitedEMatchTheorems); l_panic___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__1___closed__1 = _init_l_panic___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__1___closed__1(); lean_mark_persistent(l_panic___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__1___closed__1); -l_panic___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__1___closed__2 = _init_l_panic___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__1___closed__2(); -lean_mark_persistent(l_panic___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__1___closed__2); -l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__1 = _init_l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__1(); -l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__2 = _init_l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__2(); -l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__6___closed__1 = _init_l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__6___closed__1(); -lean_mark_persistent(l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__6___closed__1); +l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__1 = _init_l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__1(); +l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__2 = _init_l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__2(); +l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__3 = _init_l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__3(); +lean_mark_persistent(l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__3); l_Lean_Meta_Grind_EMatchTheorems_insert___closed__1 = _init_l_Lean_Meta_Grind_EMatchTheorems_insert___closed__1(); lean_mark_persistent(l_Lean_Meta_Grind_EMatchTheorems_insert___closed__1); l_Lean_Meta_Grind_EMatchTheorems_insert___closed__2 = _init_l_Lean_Meta_Grind_EMatchTheorems_insert___closed__2(); @@ -18251,53 +29904,47 @@ l_Lean_Meta_Grind_EMatchTheorems_insert___closed__3 = _init_l_Lean_Meta_Grind_EM lean_mark_persistent(l_Lean_Meta_Grind_EMatchTheorems_insert___closed__3); l_Lean_Meta_Grind_EMatchTheorems_insert___closed__4 = _init_l_Lean_Meta_Grind_EMatchTheorems_insert___closed__4(); lean_mark_persistent(l_Lean_Meta_Grind_EMatchTheorems_insert___closed__4); -l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____spec__1___closed__1 = _init_l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____spec__1___closed__1(); -lean_mark_persistent(l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____spec__1___closed__1); -l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____spec__1___closed__2 = _init_l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____spec__1___closed__2(); -lean_mark_persistent(l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____spec__1___closed__2); -l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____spec__1 = _init_l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____spec__1(); -lean_mark_persistent(l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____spec__1); -l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__1 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__1(); -lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__1); -l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__2 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__2(); -lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__2); -l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__3 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__3(); -lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__3); -l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__4 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__4(); -lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__4); -l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__5 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__5(); -lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__5); -l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__6 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__6(); -lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__6); -l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__7 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__7(); -lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__7); -l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__8 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__8(); -lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__8); -l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__9 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__9(); -lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__9); -l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__10 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__10(); -lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__10); -l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__11 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__11(); -lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__11); -l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__12 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__12(); -lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__12); -l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__13 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__13(); -lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__13); -l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__14 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__14(); -lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__14); -l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__15 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__15(); -lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__15); -l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__16 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__16(); -lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__16); -l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__17 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__17(); -lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__17); -l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__18 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__18(); -lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__18); -l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__19 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__19(); -lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__19); -l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__20 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__20(); -lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__20); -if (builtin) {res = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687_(lean_io_mk_world()); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__1 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__1); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__2 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__2); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__3 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__3); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__4 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__4(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__4); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__5 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__5(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__5); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__6 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__6(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__6); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__7 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__7(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__7); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__8 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__8(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__8); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__9 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__9(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__9); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__10 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__10(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__10); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__11 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__11(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__11); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__12 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__12(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__12); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__13 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__13(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__13); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__14 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__14(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__14); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__15 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__15(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__15); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__16 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__16(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__16); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__17 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__17(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__17); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__18 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__18(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__18); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__19 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__19(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__19); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__20 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__20(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__20); +if (builtin) {res = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ematchTheoremsExt = lean_io_result_get_value(res); lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ematchTheoremsExt); @@ -18390,8 +30037,8 @@ l_Lean_Meta_Grind_ppPattern___closed__10 = _init_l_Lean_Meta_Grind_ppPattern___c lean_mark_persistent(l_Lean_Meta_Grind_ppPattern___closed__10); l_Lean_Meta_Grind_ppPattern___closed__11 = _init_l_Lean_Meta_Grind_ppPattern___closed__11(); lean_mark_persistent(l_Lean_Meta_Grind_ppPattern___closed__11); -l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFunMask___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFunMask___closed__1(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFunMask___closed__1); +l_Lean_Meta_Grind_NormalizePattern_getPatternSupportMask___closed__1 = _init_l_Lean_Meta_Grind_NormalizePattern_getPatternSupportMask___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_NormalizePattern_getPatternSupportMask___closed__1); l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__3___closed__1 = _init_l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__3___closed__1(); lean_mark_persistent(l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__3___closed__1); l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__3___closed__2 = _init_l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__3___closed__2(); @@ -18466,49 +30113,233 @@ l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0 lean_mark_persistent(l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___closed__3); l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___lambda__1___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___lambda__1___closed__1(); lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___lambda__1___closed__1); -l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__2___closed__1 = _init_l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__2___closed__1(); -lean_mark_persistent(l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__2___closed__1); -l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__2___closed__2 = _init_l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__2___closed__2(); -lean_mark_persistent(l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__2___closed__2); -l_Lean_Meta_Grind_addEMatchTheorem___lambda__2___closed__1 = _init_l_Lean_Meta_Grind_addEMatchTheorem___lambda__2___closed__1(); -lean_mark_persistent(l_Lean_Meta_Grind_addEMatchTheorem___lambda__2___closed__1); -l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__1 = _init_l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__1(); -lean_mark_persistent(l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__1); -l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__2 = _init_l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__2(); -lean_mark_persistent(l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__2); -l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__3 = _init_l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__3(); -lean_mark_persistent(l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__3); -l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__4 = _init_l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__4(); -lean_mark_persistent(l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__4); -l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__5 = _init_l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__5(); -lean_mark_persistent(l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__5); -l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__6 = _init_l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__6(); -lean_mark_persistent(l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__6); +l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___closed__1 = _init_l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___closed__1); +l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___closed__2 = _init_l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___closed__2); +l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___closed__3 = _init_l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___closed__3); +l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___closed__4 = _init_l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___closed__4(); +lean_mark_persistent(l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___closed__4); +l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___closed__5 = _init_l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___closed__5(); +lean_mark_persistent(l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___closed__5); +l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___closed__6 = _init_l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___closed__6(); +lean_mark_persistent(l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___closed__6); +l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__1 = _init_l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__1); +l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__2 = _init_l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__2); +l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__3 = _init_l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__3); +l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__4 = _init_l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__4(); +lean_mark_persistent(l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__4); +l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__5 = _init_l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__5(); +lean_mark_persistent(l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__5); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getProofFor___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getProofFor___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getProofFor___closed__1); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getProofFor___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getProofFor___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getProofFor___closed__2); +l_Lean_Meta_Grind_mkEMatchEqTheoremCore___lambda__3___closed__1 = _init_l_Lean_Meta_Grind_mkEMatchEqTheoremCore___lambda__3___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_mkEMatchEqTheoremCore___lambda__3___closed__1); +l_Lean_Meta_Grind_mkEMatchEqTheoremCore___lambda__3___closed__2 = _init_l_Lean_Meta_Grind_mkEMatchEqTheoremCore___lambda__3___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_mkEMatchEqTheoremCore___lambda__3___closed__2); +l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__1___closed__1 = _init_l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__1___closed__1(); +lean_mark_persistent(l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__1___closed__1); +l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__1___closed__2 = _init_l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__1___closed__2(); +lean_mark_persistent(l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__1___closed__2); l_Lean_Meta_Grind_addEMatchTheorem___closed__1 = _init_l_Lean_Meta_Grind_addEMatchTheorem___closed__1(); lean_mark_persistent(l_Lean_Meta_Grind_addEMatchTheorem___closed__1); -l_Lean_Meta_Grind_addEMatchTheorem___closed__2 = _init_l_Lean_Meta_Grind_addEMatchTheorem___closed__2(); -lean_mark_persistent(l_Lean_Meta_Grind_addEMatchTheorem___closed__2); -l_Lean_Meta_Grind_addEMatchTheorem___closed__3 = _init_l_Lean_Meta_Grind_addEMatchTheorem___closed__3(); -lean_mark_persistent(l_Lean_Meta_Grind_addEMatchTheorem___closed__3); -l_Lean_Meta_Grind_addEMatchTheorem___closed__4 = _init_l_Lean_Meta_Grind_addEMatchTheorem___closed__4(); -lean_mark_persistent(l_Lean_Meta_Grind_addEMatchTheorem___closed__4); -l_Lean_Meta_Grind_addEMatchTheorem___closed__5 = _init_l_Lean_Meta_Grind_addEMatchTheorem___closed__5(); -lean_mark_persistent(l_Lean_Meta_Grind_addEMatchTheorem___closed__5); -l_Lean_Meta_Grind_addEMatchTheorem___closed__6 = _init_l_Lean_Meta_Grind_addEMatchTheorem___closed__6(); -lean_mark_persistent(l_Lean_Meta_Grind_addEMatchTheorem___closed__6); -l_Lean_Meta_Grind_addEMatchTheorem___closed__7 = _init_l_Lean_Meta_Grind_addEMatchTheorem___closed__7(); -lean_mark_persistent(l_Lean_Meta_Grind_addEMatchTheorem___closed__7); -l_Lean_Meta_Grind_addEMatchTheorem___closed__8 = _init_l_Lean_Meta_Grind_addEMatchTheorem___closed__8(); -lean_mark_persistent(l_Lean_Meta_Grind_addEMatchTheorem___closed__8); -l_Lean_Meta_Grind_addEMatchTheorem___closed__9 = _init_l_Lean_Meta_Grind_addEMatchTheorem___closed__9(); -lean_mark_persistent(l_Lean_Meta_Grind_addEMatchTheorem___closed__9); -l_Lean_Meta_Grind_addEMatchTheorem___closed__10 = _init_l_Lean_Meta_Grind_addEMatchTheorem___closed__10(); -lean_mark_persistent(l_Lean_Meta_Grind_addEMatchTheorem___closed__10); -l_Lean_Meta_Grind_addEMatchTheorem___closed__11 = _init_l_Lean_Meta_Grind_addEMatchTheorem___closed__11(); -lean_mark_persistent(l_Lean_Meta_Grind_addEMatchTheorem___closed__11); -l_Lean_Meta_Grind_addEMatchTheorem___closed__12 = _init_l_Lean_Meta_Grind_addEMatchTheorem___closed__12(); -lean_mark_persistent(l_Lean_Meta_Grind_addEMatchTheorem___closed__12); -return lean_io_result_mk_ok(lean_box(0)); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_noConfusion___rarg___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_noConfusion___rarg___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_noConfusion___rarg___closed__1); +l_Lean_Meta_Grind_instInhabitedTheoremKind = _init_l_Lean_Meta_Grind_instInhabitedTheoremKind(); +l_Lean_Meta_Grind_instBEqTheoremKind___closed__1 = _init_l_Lean_Meta_Grind_instBEqTheoremKind___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_instBEqTheoremKind___closed__1); +l_Lean_Meta_Grind_instBEqTheoremKind = _init_l_Lean_Meta_Grind_instBEqTheoremKind(); +lean_mark_persistent(l_Lean_Meta_Grind_instBEqTheoremKind); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute___closed__1); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute___closed__2); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute___closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute___closed__3); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute___closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute___closed__4(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute___closed__4); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute___closed__5 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute___closed__5(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute___closed__5); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute___closed__6 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute___closed__6(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute___closed__6); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__1); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__2); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__3); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__4(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__4); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__5 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__5(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__5); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__6 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__6(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__6); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__7 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__7(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__7); +l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___spec__2___closed__1 = _init_l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___spec__2___closed__1(); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___lambda__2___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___lambda__2___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___lambda__2___closed__1); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___lambda__2___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___lambda__2___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___lambda__2___closed__2); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___closed__1); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___closed__2); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___closed__3); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___closed__4(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___closed__4); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__3___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__3___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__3___closed__1); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__3___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__3___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__3___closed__2); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__4___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__4___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__4___closed__1); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__6___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__6___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__6___closed__1); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__6___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__6___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__6___closed__2); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__6___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__6___closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__6___closed__3); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__7___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__7___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__7___closed__1); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__7___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__7___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__7___closed__2); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__7___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__7___closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__7___closed__3); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__7___closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__7___closed__4(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__7___closed__4); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collectPatterns_x3f___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collectPatterns_x3f___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collectPatterns_x3f___closed__1); +l_Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2___closed__1 = _init_l_Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2___closed__1); +l_Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2___closed__2 = _init_l_Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2___closed__2); +l_Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2___closed__3 = _init_l_Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2___closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2___closed__3); +l_Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2___closed__4 = _init_l_Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2___closed__4(); +lean_mark_persistent(l_Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2___closed__4); +l_Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2___closed__5 = _init_l_Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2___closed__5(); +lean_mark_persistent(l_Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2___closed__5); +l_Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2___closed__6 = _init_l_Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2___closed__6(); +lean_mark_persistent(l_Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2___closed__6); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__1); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__2); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__3); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__4(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__4); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__5 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__5(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__5); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__6 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__6(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__6); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__7 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__7(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__7); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__8 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__8(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__8); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__9 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__9(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__9); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__10 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__10(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__10); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr___closed__1); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr___closed__2); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr___closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr___closed__3); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__1); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__2); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__3); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__4(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__4); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__5 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__5(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__5); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__6 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__6(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__6); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__7 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__7(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__7); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__8 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__8(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__8); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__1 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__1); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__2 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__2(); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__3 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__3); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__4 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__4(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__4); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__5 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__5(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__5); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__6 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__6(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__6); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__7 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__7(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__7); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__8 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__8(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__8); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__9 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__9(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__9); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__10 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__10(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__10); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__6___closed__1 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__6___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__6___closed__1); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__6___closed__2 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__6___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__6___closed__2); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__1 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__1); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__2 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__2); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__3 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__3); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__4 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__4(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__4); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__5 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__5(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__5); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__6 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__6(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__6); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__7 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__7(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__7); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__8 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__8(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__8); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__9 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__9(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__9); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__10 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__10(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__10); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__11 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__11(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__11); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__12 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__12(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__12); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__13 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__13(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__13); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__14 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__14(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__14); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__15 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__15(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__15); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__16 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__16(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__16); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__17 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__17(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__17); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__18 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__18(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__18); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__19 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__19(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__19); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__20 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__20(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__20); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__21 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__21(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__21); +if (builtin) {res = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279_(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/Tactic/Grind/ForallProp.c b/stage0/stdlib/Lean/Meta/Tactic/Grind/ForallProp.c index 65165f8e3b54..05ee49a5c60f 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Grind/ForallProp.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Grind/ForallProp.c @@ -15,461 +15,3106 @@ extern "C" { #endif lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); lean_object* l_Lean_Meta_Simp_Result_getProof(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_propagateForallPropDown___closed__6; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateForallPropUp___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___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_isEqTrueHyp_x3f___lambda__2(lean_object*); +static lean_object* l_Lean_Meta_Grind_propagateForallPropDown___closed__2; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateForallPropUp___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* l_Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_isEqTrueHyp_x3f___lambda__1___boxed(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_addLocalEMatchTheorems___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkAppB(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateForallPropDown(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_isApp(lean_object*); +static lean_object* l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__5; +static lean_object* l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_isEqTrueHyp_x3f___closed__1; +static lean_object* l_Lean_Meta_Grind_propagateForallPropUp___closed__3; +static lean_object* l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_addLocalEMatchTheorems___lambda__1___closed__1; lean_object* l_Lean_Meta_Grind_pushEqCore(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__6; -static lean_object* l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__2; -static lean_object* l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__8; +static lean_object* l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_addLocalEMatchTheorems___lambda__4___closed__1; +static lean_object* l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_addLocalEMatchTheorems___closed__2; +static lean_object* l_Lean_Meta_Grind_propagateForallPropUp___closed__4; +lean_object* l_Lean_Expr_cleanupAnnotations(lean_object*); +lean_object* l_Lean_stringToMessageData(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_isEqTrueHyp_x3f___lambda__1(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateForallPropUp___lambda__1___boxed(lean_object**); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateForallPropUp___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__6; +static lean_object* l_Lean_Meta_Grind_propagateForallPropUp___lambda__1___closed__2; lean_object* l_Lean_Meta_Grind_simp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_Grind_activateTheorem(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__5; -static lean_object* l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__3; -static lean_object* l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__1; +static lean_object* l_Lean_Meta_Grind_propagateForallPropDown___closed__3; +static lean_object* l_Lean_Meta_Grind_propagateForallPropUp___closed__1; +static lean_object* l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_addLocalEMatchTheorems___lambda__1___closed__3; +static lean_object* l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__6; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_addLocalEMatchTheorems___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_st_ref_take(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_isEqTrueHyp_x3f___closed__2; +static lean_object* l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_addLocalEMatchTheorems___lambda__1___closed__4; +static lean_object* l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__2; +static lean_object* l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_addLocalEMatchTheorems___closed__1; +static lean_object* l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__4; +static lean_object* l_Lean_Meta_Grind_propagateForallPropDown___closed__1; +static lean_object* l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_isEqTrueHyp_x3f___closed__4; +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___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_isEqTrueHyp_x3f___closed__3; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___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_get(lean_object*, lean_object*); +lean_object* l_Lean_Expr_appFnCleanup(lean_object*, lean_object*); +uint8_t l_Lean_Expr_hasLooseBVars(lean_object*); +static lean_object* l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__11; +static lean_object* l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__9; +static lean_object* l_Lean_Meta_Grind_propagateForallPropDown___closed__7; +lean_object* l_Lean_Meta_Grind_pushEqTrue(lean_object*, lean_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_Grind_propagateForallPropUp___lambda__3___closed__1; lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_getGeneration(lean_object*, 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_Grind_propagateForallProp___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__4; +lean_object* l_Lean_Meta_Grind_alreadyInternalized(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_mkNot(lean_object*); +static lean_object* l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__3; +lean_object* l_Lean_Meta_Grind_updateLastTag(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__8; +static lean_object* l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__7; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateForallPropUp___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*); +static lean_object* l_Lean_Meta_Grind_propagateForallPropDown___closed__8; +static lean_object* l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__10; +lean_object* l_Lean_MessageData_ofExpr(lean_object*); +static lean_object* l_Lean_Meta_Grind_propagateForallPropDown___closed__5; +lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___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_name_append_index_after(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_addLocalEMatchTheorems___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_Meta_Grind_propagateForallPropUp___lambda__2___closed__3; +lean_object* l_Lean_Meta_Grind_addNewFact(lean_object*, lean_object*, lean_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_Grind_ForallProp_0__Lean_Meta_Grind_addLocalEMatchTheorems___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__7; +uint8_t lean_nat_dec_eq(lean_object*, lean_object*); +lean_object* l_Lean_mkApp3(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_Grind_mkEqFalseProof(lean_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_Grind_propagateForallPropUp___lambda__2___closed__8; +lean_object* l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___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_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__4; +lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_addLocalEMatchTheorems___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_propagateForallPropDown___closed__4; +lean_object* l_Lean_indentExpr(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___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_Grind_propagateForallPropUp___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___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_isEqTrueHyp_x3f(lean_object*); +static lean_object* l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__1; +uint8_t l_Lean_Expr_isConstOf(lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_mkEqTrueProof(lean_object*, 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_Grind_propagateForallProp(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_Grind_ForallProp_0__Lean_Meta_Grind_addLocalEMatchTheorems___lambda__1___closed__2; +static lean_object* l_Lean_Meta_Grind_propagateForallPropDown___closed__9; +static lean_object* l_Lean_Meta_Grind_propagateForallPropUp___closed__2; +lean_object* lean_array_mk(lean_object*); +static lean_object* l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__9; +lean_object* l_Lean_Meta_Grind_isEqFalse(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_Grind_pushEqFalse(lean_object*, lean_object*, 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_Grind_propagateForallPropUp___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___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_isEqTrueHyp_x3f___closed__5; +static lean_object* l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__2; +lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_propagateForallPropUp___lambda__1___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateForallPropUp___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp(lean_object*, lean_object*, lean_object*, 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_Grind_propagateForallPropUp(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_Grind_isEqTrue(lean_object*, 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_Grind_propagateForallProp___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_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_isEqTrueHyp_x3f___lambda__2___boxed(lean_object*); +static lean_object* l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__1; +lean_object* lean_nat_add(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_addLocalEMatchTheorems(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_lam___override(lean_object*, lean_object*, lean_object*, uint8_t); lean_object* l_Lean_mkApp5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_propagateForallPropUp___lambda__3___closed__2; +static lean_object* l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__5; lean_object* lean_expr_instantiate1(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__7; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_addLocalEMatchTheorems___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_internalize(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* _init_l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__1() { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_addLocalEMatchTheorems___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_propagateForallPropDown___closed__10; +static lean_object* _init_l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("of_eq_true", 10, 10); +x_1 = lean_mk_string_unchecked("Lean", 4, 4); return x_1; } } -static lean_object* _init_l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__2() { +static lean_object* _init_l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__2() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Meta_Grind_propagateForallProp___lambda__1___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_unchecked("Grind", 5, 5); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("imp_eq_of_eq_true_right", 23, 23); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__1; +x_2 = l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__2; +x_3 = l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__3; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; } } -static lean_object* _init_l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__3() { +static lean_object* _init_l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__2; +x_2 = l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__4; x_3 = l_Lean_Expr_const___override(x_2, x_1); return x_3; } } -static lean_object* _init_l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__4() { +static lean_object* _init_l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__6() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean", 4, 4); +x_1 = lean_mk_string_unchecked("imp_eq_of_eq_true_left", 22, 22); return x_1; } } -static lean_object* _init_l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__5() { +static lean_object* _init_l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__7() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Grind", 5, 5); -return x_1; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__1; +x_2 = l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__2; +x_3 = l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__6; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__7; +x_3 = l_Lean_Expr_const___override(x_2, x_1); +return x_3; } } -static lean_object* _init_l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__6() { +static lean_object* _init_l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__9() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("forall_propagator", 17, 17); +x_1 = lean_mk_string_unchecked("imp_eq_of_eq_false_left", 23, 23); return x_1; } } -static lean_object* _init_l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__7() { +static lean_object* _init_l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__4; -x_2 = l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__5; -x_3 = l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__6; +x_1 = l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__1; +x_2 = l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__2; +x_3 = l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__9; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__8() { +static lean_object* _init_l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__7; +x_2 = l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__10; x_3 = l_Lean_Expr_const___override(x_2, x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateForallProp___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___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_16; -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_object* x_14; lean_inc(x_1); -x_16 = l_Lean_Meta_Grind_mkEqTrueProof(x_1, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); -if (lean_obj_tag(x_16) == 0) +x_14 = l_Lean_Meta_Grind_isEqFalse(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_14) == 0) { -lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -x_18 = lean_ctor_get(x_16, 1); -lean_inc(x_18); -lean_dec(x_16); -x_19 = l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__3; +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_17; lean_object* x_18; +x_17 = lean_ctor_get(x_14, 1); lean_inc(x_17); +lean_dec(x_14); lean_inc(x_1); -x_20 = l_Lean_mkAppB(x_19, x_1, x_17); -x_21 = lean_expr_instantiate1(x_2, x_20); -lean_dec(x_20); -lean_inc(x_14); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_9); -x_22 = l_Lean_Meta_Grind_simp(x_21, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_18); +x_18 = l_Lean_Meta_Grind_isEqTrue(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_17); +if (lean_obj_tag(x_18) == 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_21; lean_object* x_22; +x_21 = lean_ctor_get(x_18, 1); +lean_inc(x_21); +lean_dec(x_18); +lean_inc(x_2); +x_22 = l_Lean_Meta_Grind_isEqTrue(x_2, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_21); if (lean_obj_tag(x_22) == 0) { -lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +lean_object* x_23; uint8_t x_24; 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_1); -x_25 = l_Lean_Expr_lam___override(x_3, x_1, x_2, x_4); -x_26 = lean_ctor_get(x_23, 0); -lean_inc(x_26); -lean_inc(x_5); -x_27 = l_Lean_Meta_Grind_getGeneration(x_5, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_24); -if (lean_obj_tag(x_27) == 0) +x_24 = lean_unbox(x_23); +lean_dec(x_23); +if (x_24 == 0) +{ +uint8_t x_25; +lean_dec(x_12); +lean_dec(x_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_25 = !lean_is_exclusive(x_22); +if (x_25 == 0) +{ +lean_object* x_26; lean_object* x_27; +x_26 = lean_ctor_get(x_22, 0); +lean_dec(x_26); +x_27 = lean_box(0); +lean_ctor_set(x_22, 0, x_27); +return x_22; +} +else { lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_28 = lean_ctor_get(x_27, 0); +x_28 = lean_ctor_get(x_22, 1); lean_inc(x_28); -x_29 = lean_ctor_get(x_27, 1); -lean_inc(x_29); -lean_dec(x_27); -lean_inc(x_14); -lean_inc(x_13); +lean_dec(x_22); +x_29 = lean_box(0); +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_28); +return x_30; +} +} +else +{ +lean_object* x_31; lean_object* x_32; +x_31 = lean_ctor_get(x_22, 1); +lean_inc(x_31); +lean_dec(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_26); -x_30 = l_Lean_Meta_Grind_internalize(x_26, x_28, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_29); -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); -lean_inc(x_14); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -x_32 = l_Lean_Meta_Simp_Result_getProof(x_23, x_11, x_12, x_13, x_14, x_31); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_2); +x_32 = l_Lean_Meta_Grind_mkEqTrueProof(x_2, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_31); if (lean_obj_tag(x_32) == 0) { -lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; lean_object* x_38; +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_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_Grind_propagateForallProp___lambda__1___closed__8; -lean_inc(x_26); -x_36 = l_Lean_mkApp5(x_35, x_1, x_25, x_26, x_17, x_33); -x_37 = 0; -x_38 = l_Lean_Meta_Grind_pushEqCore(x_5, x_26, x_36, x_37, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_34); -lean_dec(x_14); -lean_dec(x_13); +x_35 = l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__5; +x_36 = l_Lean_mkApp3(x_35, x_1, x_2, x_33); +x_37 = l_Lean_Meta_Grind_pushEqTrue(x_3, x_36, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_34); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -return x_38; +lean_dec(x_6); +lean_dec(x_5); +return x_37; } else { -uint8_t x_39; -lean_dec(x_26); -lean_dec(x_25); -lean_dec(x_17); -lean_dec(x_14); -lean_dec(x_13); +uint8_t x_38; lean_dec(x_12); lean_dec(x_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_39 = !lean_is_exclusive(x_32); -if (x_39 == 0) +x_38 = !lean_is_exclusive(x_32); +if (x_38 == 0) { return x_32; } else { -lean_object* x_40; lean_object* x_41; lean_object* x_42; -x_40 = lean_ctor_get(x_32, 0); -x_41 = lean_ctor_get(x_32, 1); -lean_inc(x_41); +lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_39 = lean_ctor_get(x_32, 0); +x_40 = lean_ctor_get(x_32, 1); lean_inc(x_40); +lean_inc(x_39); lean_dec(x_32); -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; +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_43; -lean_dec(x_26); -lean_dec(x_25); -lean_dec(x_23); -lean_dec(x_17); -lean_dec(x_14); -lean_dec(x_13); +uint8_t x_42; lean_dec(x_12); lean_dec(x_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_43 = !lean_is_exclusive(x_30); -if (x_43 == 0) +x_42 = !lean_is_exclusive(x_22); +if (x_42 == 0) { -return x_30; +return x_22; } 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_inc(x_45); +lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_43 = lean_ctor_get(x_22, 0); +x_44 = lean_ctor_get(x_22, 1); 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_inc(x_43); +lean_dec(x_22); +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_47; -lean_dec(x_26); -lean_dec(x_25); -lean_dec(x_23); -lean_dec(x_17); -lean_dec(x_14); -lean_dec(x_13); +lean_object* x_46; lean_object* x_47; +x_46 = lean_ctor_get(x_18, 1); +lean_inc(x_46); +lean_dec(x_18); +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_1); +x_47 = l_Lean_Meta_Grind_mkEqTrueProof(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_46); +if (lean_obj_tag(x_47) == 0) +{ +lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; uint8_t x_52; lean_object* x_53; +x_48 = lean_ctor_get(x_47, 0); +lean_inc(x_48); +x_49 = lean_ctor_get(x_47, 1); +lean_inc(x_49); +lean_dec(x_47); +x_50 = l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__8; +lean_inc(x_2); +x_51 = l_Lean_mkApp3(x_50, x_1, x_2, x_48); +x_52 = 0; +x_53 = l_Lean_Meta_Grind_pushEqCore(x_3, x_2, x_51, x_52, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_49); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); +lean_dec(x_6); lean_dec(x_5); -lean_dec(x_1); -x_47 = !lean_is_exclusive(x_27); -if (x_47 == 0) -{ -return x_27; -} -else -{ -lean_object* x_48; lean_object* x_49; lean_object* x_50; -x_48 = lean_ctor_get(x_27, 0); -x_49 = lean_ctor_get(x_27, 1); -lean_inc(x_49); -lean_inc(x_48); -lean_dec(x_27); -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; -} -} +return x_53; } else { -uint8_t x_51; -lean_dec(x_17); -lean_dec(x_14); -lean_dec(x_13); +uint8_t x_54; lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); +lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_51 = !lean_is_exclusive(x_22); -if (x_51 == 0) +x_54 = !lean_is_exclusive(x_47); +if (x_54 == 0) { -return x_22; +return x_47; } else { -lean_object* x_52; lean_object* x_53; lean_object* x_54; -x_52 = lean_ctor_get(x_22, 0); -x_53 = lean_ctor_get(x_22, 1); -lean_inc(x_53); -lean_inc(x_52); -lean_dec(x_22); -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_55; lean_object* x_56; lean_object* x_57; +x_55 = lean_ctor_get(x_47, 0); +x_56 = lean_ctor_get(x_47, 1); +lean_inc(x_56); +lean_inc(x_55); +lean_dec(x_47); +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_55; -lean_dec(x_14); -lean_dec(x_13); +uint8_t x_58; lean_dec(x_12); lean_dec(x_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_55 = !lean_is_exclusive(x_16); -if (x_55 == 0) +x_58 = !lean_is_exclusive(x_18); +if (x_58 == 0) { -return x_16; +return x_18; } else { -lean_object* x_56; lean_object* x_57; lean_object* x_58; -x_56 = lean_ctor_get(x_16, 0); -x_57 = lean_ctor_get(x_16, 1); -lean_inc(x_57); -lean_inc(x_56); -lean_dec(x_16); -x_58 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_58, 0, x_56); -lean_ctor_set(x_58, 1, x_57); -return x_58; +lean_object* x_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; } } } -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateForallProp(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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) == 7) +else { -lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; lean_object* x_15; -x_11 = lean_ctor_get(x_1, 0); -lean_inc(x_11); -x_12 = lean_ctor_get(x_1, 1); -lean_inc(x_12); -x_13 = lean_ctor_get(x_1, 2); -lean_inc(x_13); -x_14 = lean_ctor_get_uint8(x_1, sizeof(void*)*3 + 8); +lean_object* x_62; lean_object* x_63; +x_62 = lean_ctor_get(x_14, 1); +lean_inc(x_62); +lean_dec(x_14); lean_inc(x_12); -x_15 = l_Lean_Meta_Grind_isEqTrue(x_12, 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_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_1); +x_63 = l_Lean_Meta_Grind_mkEqFalseProof(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_62); +if (lean_obj_tag(x_63) == 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_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; +x_64 = lean_ctor_get(x_63, 0); +lean_inc(x_64); +x_65 = lean_ctor_get(x_63, 1); +lean_inc(x_65); +lean_dec(x_63); +x_66 = l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__11; +x_67 = l_Lean_mkApp3(x_66, x_1, x_2, x_64); +x_68 = l_Lean_Meta_Grind_pushEqTrue(x_3, x_67, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_65); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +return x_68; +} +else { -uint8_t x_18; -lean_dec(x_13); +uint8_t x_69; lean_dec(x_12); lean_dec(x_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_18 = !lean_is_exclusive(x_15); -if (x_18 == 0) +x_69 = !lean_is_exclusive(x_63); +if (x_69 == 0) { -lean_object* x_19; lean_object* x_20; -x_19 = lean_ctor_get(x_15, 0); -lean_dec(x_19); -x_20 = lean_box(0); -lean_ctor_set(x_15, 0, x_20); -return x_15; +return x_63; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_15, 1); -lean_inc(x_21); -lean_dec(x_15); -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_70; lean_object* x_71; lean_object* x_72; +x_70 = lean_ctor_get(x_63, 0); +x_71 = lean_ctor_get(x_63, 1); +lean_inc(x_71); +lean_inc(x_70); +lean_dec(x_63); +x_72 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_72, 0, x_70); +lean_ctor_set(x_72, 1, x_71); +return x_72; } } -else -{ -lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_24 = lean_ctor_get(x_15, 1); -lean_inc(x_24); -lean_dec(x_15); -x_25 = lean_box(0); -x_26 = l_Lean_Meta_Grind_propagateForallProp___lambda__1(x_12, x_13, x_11, x_14, x_1, x_25, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_24); -return x_26; } } else { -uint8_t x_27; -lean_dec(x_13); +uint8_t x_73; lean_dec(x_12); lean_dec(x_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_73 = !lean_is_exclusive(x_14); +if (x_73 == 0) +{ +return x_14; +} +else +{ +lean_object* x_74; lean_object* x_75; lean_object* x_76; +x_74 = lean_ctor_get(x_14, 0); +x_75 = lean_ctor_get(x_14, 1); +lean_inc(x_75); +lean_inc(x_74); +lean_dec(x_14); +x_76 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_76, 0, x_74); +lean_ctor_set(x_76, 1, x_75); +return x_76; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, 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; uint8_t x_15; +x_13 = l_Lean_Meta_Grind_alreadyInternalized(x_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) +{ +uint8_t 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_16 = !lean_is_exclusive(x_13); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; +x_17 = lean_ctor_get(x_13, 0); +lean_dec(x_17); +x_18 = lean_box(0); +lean_ctor_set(x_13, 0, x_18); +return x_13; +} +else +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_13, 1); +lean_inc(x_19); +lean_dec(x_13); +x_20 = lean_box(0); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_20); +lean_ctor_set(x_21, 1, x_19); +return x_21; +} +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_22 = lean_ctor_get(x_13, 1); +lean_inc(x_22); +lean_dec(x_13); +x_23 = lean_box(0); +x_24 = l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1(x_2, x_3, x_1, x_23, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_22); +return x_24; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +lean_object* x_14; +x_14 = l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_4); +return x_14; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateForallPropUp___lambda__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("forall_propagator", 17, 17); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateForallPropUp___lambda__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__1; +x_2 = l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__2; +x_3 = l_Lean_Meta_Grind_propagateForallPropUp___lambda__1___closed__1; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateForallPropUp___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17) { +_start: +{ +lean_object* x_18; +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +x_18 = l_Lean_Meta_Simp_Result_getProof(x_1, x_13, x_14, x_15, x_16, x_17); +if (lean_obj_tag(x_18) == 0) +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; lean_object* x_25; +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_Meta_Grind_propagateForallPropUp___lambda__1___closed__2; +x_22 = l_Lean_Expr_const___override(x_21, x_2); +lean_inc(x_5); +x_23 = l_Lean_mkApp5(x_22, x_3, x_4, x_5, x_6, x_19); +x_24 = 0; +x_25 = l_Lean_Meta_Grind_pushEqCore(x_7, x_5, x_23, x_24, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_20); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +return x_25; +} +else +{ +uint8_t x_26; +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_26 = !lean_is_exclusive(x_18); +if (x_26 == 0) +{ +return x_18; +} +else +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_18, 0); +x_28 = lean_ctor_get(x_18, 1); +lean_inc(x_28); +lean_inc(x_27); +lean_dec(x_18); +x_29 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_29, 0, x_27); +lean_ctor_set(x_29, 1, x_28); +return x_29; +} +} +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("of_eq_true", 10, 10); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateForallPropUp___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_Meta_Grind_propagateForallPropUp___lambda__2___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__2; +x_3 = l_Lean_Expr_const___override(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("q': ", 4, 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__4; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(" for", 4, 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__6; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__8() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("", 0, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__8; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateForallPropUp___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +_start: +{ +lean_object* x_17; +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_1); +x_17 = l_Lean_Meta_Grind_mkEqTrueProof(x_1, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +if (lean_obj_tag(x_17) == 0) +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +x_19 = lean_ctor_get(x_17, 1); +lean_inc(x_19); +lean_dec(x_17); +x_20 = lean_box(0); +x_21 = l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__3; +lean_inc(x_18); +lean_inc(x_1); +x_22 = l_Lean_mkAppB(x_21, x_1, x_18); +x_23 = lean_expr_instantiate1(x_2, x_22); +lean_dec(x_22); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_10); +x_24 = l_Lean_Meta_Grind_simp(x_23, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_19); +if (lean_obj_tag(x_24) == 0) +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +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_1); +x_27 = l_Lean_Expr_lam___override(x_3, x_1, x_2, x_4); +x_28 = lean_ctor_get(x_25, 0); +lean_inc(x_28); +lean_inc(x_5); +x_29 = l_Lean_Meta_Grind_getGeneration(x_5, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_26); +if (lean_obj_tag(x_29) == 0) +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_30 = lean_ctor_get(x_29, 0); +lean_inc(x_30); +x_31 = lean_ctor_get(x_29, 1); +lean_inc(x_31); +lean_dec(x_29); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_28); +x_32 = l_Lean_Meta_Grind_internalize(x_28, x_30, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_31); +if (lean_obj_tag(x_32) == 0) +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; +x_33 = lean_ctor_get(x_32, 1); +lean_inc(x_33); +lean_dec(x_32); +lean_inc(x_6); +x_34 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_6, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, 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_6); +x_37 = lean_ctor_get(x_34, 1); +lean_inc(x_37); +lean_dec(x_34); +x_38 = lean_box(0); +x_39 = l_Lean_Meta_Grind_propagateForallPropUp___lambda__1(x_25, x_20, x_1, x_27, x_28, x_18, x_5, x_38, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_37); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +return x_39; +} +else +{ +uint8_t x_40; +x_40 = !lean_is_exclusive(x_34); +if (x_40 == 0) +{ +lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_41 = lean_ctor_get(x_34, 1); +x_42 = lean_ctor_get(x_34, 0); +lean_dec(x_42); +x_43 = l_Lean_Meta_Grind_updateLastTag(x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_41); +if (lean_obj_tag(x_43) == 0) +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_44 = lean_ctor_get(x_43, 1); +lean_inc(x_44); +lean_dec(x_43); +lean_inc(x_28); +x_45 = l_Lean_MessageData_ofExpr(x_28); +x_46 = l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__5; +lean_ctor_set_tag(x_34, 7); +lean_ctor_set(x_34, 1, x_45); +lean_ctor_set(x_34, 0, x_46); +x_47 = l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__7; +x_48 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_48, 0, x_34); +lean_ctor_set(x_48, 1, x_47); +lean_inc(x_5); +x_49 = l_Lean_indentExpr(x_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_Meta_Grind_propagateForallPropUp___lambda__2___closed__9; +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_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_6, x_52, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_44); +x_54 = lean_ctor_get(x_53, 0); +lean_inc(x_54); +x_55 = lean_ctor_get(x_53, 1); +lean_inc(x_55); +lean_dec(x_53); +x_56 = l_Lean_Meta_Grind_propagateForallPropUp___lambda__1(x_25, x_20, x_1, x_27, x_28, x_18, x_5, x_54, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_55); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_54); +return x_56; +} +else +{ +uint8_t x_57; +lean_free_object(x_34); +lean_dec(x_28); +lean_dec(x_27); +lean_dec(x_25); +lean_dec(x_18); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_1); +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 +{ +lean_object* x_61; lean_object* x_62; +x_61 = lean_ctor_get(x_34, 1); +lean_inc(x_61); +lean_dec(x_34); +x_62 = l_Lean_Meta_Grind_updateLastTag(x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_61); +if (lean_obj_tag(x_62) == 0) +{ +lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; 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_63 = lean_ctor_get(x_62, 1); +lean_inc(x_63); +lean_dec(x_62); +lean_inc(x_28); +x_64 = l_Lean_MessageData_ofExpr(x_28); +x_65 = l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__5; +x_66 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_66, 0, x_65); +lean_ctor_set(x_66, 1, x_64); +x_67 = l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__7; +x_68 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_68, 0, x_66); +lean_ctor_set(x_68, 1, x_67); +lean_inc(x_5); +x_69 = l_Lean_indentExpr(x_5); +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_Meta_Grind_propagateForallPropUp___lambda__2___closed__9; +x_72 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_72, 0, x_70); +lean_ctor_set(x_72, 1, x_71); +x_73 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_6, x_72, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_63); +x_74 = lean_ctor_get(x_73, 0); +lean_inc(x_74); +x_75 = lean_ctor_get(x_73, 1); +lean_inc(x_75); +lean_dec(x_73); +x_76 = l_Lean_Meta_Grind_propagateForallPropUp___lambda__1(x_25, x_20, x_1, x_27, x_28, x_18, x_5, x_74, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_75); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_74); +return x_76; +} +else +{ +lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; +lean_dec(x_28); +lean_dec(x_27); +lean_dec(x_25); +lean_dec(x_18); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_1); +x_77 = lean_ctor_get(x_62, 0); +lean_inc(x_77); +x_78 = lean_ctor_get(x_62, 1); +lean_inc(x_78); +if (lean_is_exclusive(x_62)) { + lean_ctor_release(x_62, 0); + lean_ctor_release(x_62, 1); + x_79 = x_62; +} else { + lean_dec_ref(x_62); + x_79 = lean_box(0); +} +if (lean_is_scalar(x_79)) { + x_80 = lean_alloc_ctor(1, 2, 0); +} else { + x_80 = x_79; +} +lean_ctor_set(x_80, 0, x_77); +lean_ctor_set(x_80, 1, x_78); +return x_80; +} +} +} +} +else +{ +uint8_t x_81; +lean_dec(x_28); +lean_dec(x_27); +lean_dec(x_25); +lean_dec(x_18); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_1); +x_81 = !lean_is_exclusive(x_32); +if (x_81 == 0) +{ +return x_32; +} +else +{ +lean_object* x_82; lean_object* x_83; lean_object* x_84; +x_82 = lean_ctor_get(x_32, 0); +x_83 = lean_ctor_get(x_32, 1); +lean_inc(x_83); +lean_inc(x_82); +lean_dec(x_32); +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_85; +lean_dec(x_28); +lean_dec(x_27); +lean_dec(x_25); +lean_dec(x_18); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_1); +x_85 = !lean_is_exclusive(x_29); +if (x_85 == 0) +{ +return x_29; +} +else +{ +lean_object* x_86; lean_object* x_87; lean_object* x_88; +x_86 = lean_ctor_get(x_29, 0); +x_87 = lean_ctor_get(x_29, 1); +lean_inc(x_87); +lean_inc(x_86); +lean_dec(x_29); +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_89; +lean_dec(x_18); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_89 = !lean_is_exclusive(x_24); +if (x_89 == 0) +{ +return x_24; +} +else +{ +lean_object* x_90; lean_object* x_91; lean_object* x_92; +x_90 = lean_ctor_get(x_24, 0); +x_91 = lean_ctor_get(x_24, 1); +lean_inc(x_91); +lean_inc(x_90); +lean_dec(x_24); +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_93; +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_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_93 = !lean_is_exclusive(x_17); +if (x_93 == 0) +{ +return x_17; +} +else +{ +lean_object* x_94; lean_object* x_95; lean_object* x_96; +x_94 = lean_ctor_get(x_17, 0); +x_95 = lean_ctor_get(x_17, 1); +lean_inc(x_95); +lean_inc(x_94); +lean_dec(x_17); +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; +} +} +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateForallPropUp___lambda__3___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("isEqTrue, ", 10, 10); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateForallPropUp___lambda__3___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_propagateForallPropUp___lambda__3___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateForallPropUp___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +_start: +{ +lean_object* x_17; lean_object* x_18; uint8_t x_19; +lean_inc(x_1); +x_17 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_1, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +x_19 = lean_unbox(x_18); +lean_dec(x_18); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_17, 1); +lean_inc(x_20); +lean_dec(x_17); +x_21 = lean_box(0); +x_22 = l_Lean_Meta_Grind_propagateForallPropUp___lambda__2(x_2, x_3, x_4, x_5, x_6, x_1, x_21, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_20); +return x_22; +} +else +{ +uint8_t x_23; +x_23 = !lean_is_exclusive(x_17); +if (x_23 == 0) +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_17, 1); +x_25 = lean_ctor_get(x_17, 0); +lean_dec(x_25); +x_26 = l_Lean_Meta_Grind_updateLastTag(x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_24); +if (lean_obj_tag(x_26) == 0) +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_27 = lean_ctor_get(x_26, 1); +lean_inc(x_27); +lean_dec(x_26); +lean_inc(x_6); +x_28 = l_Lean_MessageData_ofExpr(x_6); +x_29 = l_Lean_Meta_Grind_propagateForallPropUp___lambda__3___closed__2; +lean_ctor_set_tag(x_17, 7); +lean_ctor_set(x_17, 1, x_28); +lean_ctor_set(x_17, 0, x_29); +x_30 = l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__9; +x_31 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_31, 0, x_17); +lean_ctor_set(x_31, 1, x_30); +lean_inc(x_1); +x_32 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_1, x_31, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_27); +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_Grind_propagateForallPropUp___lambda__2(x_2, x_3, x_4, x_5, x_6, x_1, x_33, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_34); +lean_dec(x_33); +return x_35; +} +else +{ +uint8_t x_36; +lean_free_object(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_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_36 = !lean_is_exclusive(x_26); +if (x_36 == 0) +{ +return x_26; +} +else +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_37 = lean_ctor_get(x_26, 0); +x_38 = lean_ctor_get(x_26, 1); +lean_inc(x_38); +lean_inc(x_37); +lean_dec(x_26); +x_39 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_39, 0, x_37); +lean_ctor_set(x_39, 1, x_38); +return x_39; +} +} +} +else +{ +lean_object* x_40; lean_object* x_41; +x_40 = lean_ctor_get(x_17, 1); +lean_inc(x_40); +lean_dec(x_17); +x_41 = l_Lean_Meta_Grind_updateLastTag(x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, 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; +x_42 = lean_ctor_get(x_41, 1); +lean_inc(x_42); +lean_dec(x_41); +lean_inc(x_6); +x_43 = l_Lean_MessageData_ofExpr(x_6); +x_44 = l_Lean_Meta_Grind_propagateForallPropUp___lambda__3___closed__2; +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_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__9; +x_47 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_47, 0, x_45); +lean_ctor_set(x_47, 1, x_46); +lean_inc(x_1); +x_48 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_1, x_47, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_42); +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_Meta_Grind_propagateForallPropUp___lambda__2(x_2, x_3, x_4, x_5, x_6, x_1, x_49, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_50); +lean_dec(x_49); +return x_51; +} +else +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; +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_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_52 = lean_ctor_get(x_41, 0); +lean_inc(x_52); +x_53 = lean_ctor_get(x_41, 1); +lean_inc(x_53); +if (lean_is_exclusive(x_41)) { + lean_ctor_release(x_41, 0); + lean_ctor_release(x_41, 1); + x_54 = x_41; +} else { + lean_dec_ref(x_41); + x_54 = lean_box(0); +} +if (lean_is_scalar(x_54)) { + x_55 = lean_alloc_ctor(1, 2, 0); +} else { + x_55 = x_54; +} +lean_ctor_set(x_55, 0, x_52); +lean_ctor_set(x_55, 1, x_53); +return x_55; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateForallPropUp___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, uint8_t x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +_start: +{ +uint8_t x_17; +x_17 = l_Lean_Expr_hasLooseBVars(x_1); +if (x_17 == 0) +{ +lean_object* x_18; +lean_dec(x_5); +lean_dec(x_4); +x_18 = l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp(x_2, x_3, x_1, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +return x_18; +} +else +{ +lean_object* x_19; +lean_inc(x_3); +x_19 = l_Lean_Meta_Grind_isEqTrue(x_3, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +if (lean_obj_tag(x_19) == 0) +{ +lean_object* x_20; uint8_t x_21; +x_20 = lean_ctor_get(x_19, 0); +lean_inc(x_20); +x_21 = lean_unbox(x_20); +lean_dec(x_20); +if (x_21 == 0) +{ +uint8_t x_22; +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +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_19); +if (x_22 == 0) +{ +lean_object* x_23; lean_object* x_24; +x_23 = lean_ctor_get(x_19, 0); +lean_dec(x_23); +x_24 = lean_box(0); +lean_ctor_set(x_19, 0, x_24); +return x_19; +} +else +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_25 = lean_ctor_get(x_19, 1); +lean_inc(x_25); +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_25); +return x_27; +} +} +else +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_28 = lean_ctor_get(x_19, 1); +lean_inc(x_28); +lean_dec(x_19); +x_29 = lean_box(0); +x_30 = l_Lean_Meta_Grind_propagateForallPropUp___lambda__3(x_4, x_3, x_1, x_5, x_6, x_2, x_29, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_28); +return x_30; +} +} +else +{ +uint8_t x_31; +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_31 = !lean_is_exclusive(x_19); +if (x_31 == 0) +{ +return x_19; +} +else +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_19, 0); +x_33 = lean_ctor_get(x_19, 1); +lean_inc(x_33); +lean_inc(x_32); +lean_dec(x_19); +x_34 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +return x_34; +} +} +} +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateForallPropUp___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("grind", 5, 5); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateForallPropUp___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("debug", 5, 5); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateForallPropUp___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("forallPropagator", 16, 16); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateForallPropUp___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Meta_Grind_propagateForallPropUp___closed__1; +x_2 = l_Lean_Meta_Grind_propagateForallPropUp___closed__2; +x_3 = l_Lean_Meta_Grind_propagateForallPropUp___closed__3; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateForallPropUp(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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) == 7) +{ +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; uint8_t x_18; +x_11 = lean_ctor_get(x_1, 0); +lean_inc(x_11); +x_12 = lean_ctor_get(x_1, 1); +lean_inc(x_12); +x_13 = lean_ctor_get(x_1, 2); +lean_inc(x_13); +x_14 = lean_ctor_get_uint8(x_1, sizeof(void*)*3 + 8); +x_15 = l_Lean_Meta_Grind_propagateForallPropUp___closed__4; +x_16 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_15, x_2, x_3, x_4, x_5, x_6, x_7, 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; +x_19 = lean_ctor_get(x_16, 1); +lean_inc(x_19); +lean_dec(x_16); +x_20 = lean_box(0); +x_21 = l_Lean_Meta_Grind_propagateForallPropUp___lambda__4(x_13, x_1, x_12, x_15, x_11, x_14, x_20, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_19); +return x_21; +} +else +{ +uint8_t x_22; +x_22 = !lean_is_exclusive(x_16); +if (x_22 == 0) +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_16, 1); +x_24 = lean_ctor_get(x_16, 0); +lean_dec(x_24); +x_25 = l_Lean_Meta_Grind_updateLastTag(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_23); +if (lean_obj_tag(x_25) == 0) +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_26 = lean_ctor_get(x_25, 1); +lean_inc(x_26); +lean_dec(x_25); +lean_inc(x_1); +x_27 = l_Lean_MessageData_ofExpr(x_1); +x_28 = l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__9; +lean_ctor_set_tag(x_16, 7); +lean_ctor_set(x_16, 1, x_27); +lean_ctor_set(x_16, 0, x_28); +x_29 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_29, 0, x_16); +lean_ctor_set(x_29, 1, x_28); +x_30 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_15, x_29, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_26); +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_Grind_propagateForallPropUp___lambda__4(x_13, x_1, x_12, x_15, x_11, x_14, x_31, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_32); +lean_dec(x_31); +return x_33; +} +else +{ +uint8_t x_34; +lean_free_object(x_16); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_34 = !lean_is_exclusive(x_25); +if (x_34 == 0) +{ +return x_25; +} +else +{ +lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_35 = lean_ctor_get(x_25, 0); +x_36 = lean_ctor_get(x_25, 1); +lean_inc(x_36); +lean_inc(x_35); +lean_dec(x_25); +x_37 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_37, 0, x_35); +lean_ctor_set(x_37, 1, x_36); +return x_37; +} +} +} +else +{ +lean_object* x_38; lean_object* x_39; +x_38 = lean_ctor_get(x_16, 1); +lean_inc(x_38); +lean_dec(x_16); +x_39 = l_Lean_Meta_Grind_updateLastTag(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_38); +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; lean_object* x_48; +x_40 = lean_ctor_get(x_39, 1); +lean_inc(x_40); +lean_dec(x_39); +lean_inc(x_1); +x_41 = l_Lean_MessageData_ofExpr(x_1); +x_42 = l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__9; +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_43); +lean_ctor_set(x_44, 1, x_42); +x_45 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_15, x_44, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_40); +x_46 = lean_ctor_get(x_45, 0); +lean_inc(x_46); +x_47 = lean_ctor_get(x_45, 1); +lean_inc(x_47); +lean_dec(x_45); +x_48 = l_Lean_Meta_Grind_propagateForallPropUp___lambda__4(x_13, x_1, x_12, x_15, x_11, x_14, x_46, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_47); +lean_dec(x_46); +return x_48; +} +else +{ +lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_49 = lean_ctor_get(x_39, 0); +lean_inc(x_49); +x_50 = lean_ctor_get(x_39, 1); +lean_inc(x_50); +if (lean_is_exclusive(x_39)) { + lean_ctor_release(x_39, 0); + lean_ctor_release(x_39, 1); + x_51 = x_39; +} else { + lean_dec_ref(x_39); + x_51 = lean_box(0); +} +if (lean_is_scalar(x_51)) { + x_52 = lean_alloc_ctor(1, 2, 0); +} else { + x_52 = x_51; +} +lean_ctor_set(x_52, 0, x_49); +lean_ctor_set(x_52, 1, x_50); +return x_52; +} +} +} +} +else +{ +lean_object* x_53; lean_object* x_54; +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_53 = lean_box(0); +x_54 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_54, 0, x_53); +lean_ctor_set(x_54, 1, x_10); +return x_54; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateForallPropUp___lambda__1___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +_start: +{ +lean_object* x_18; +x_18 = l_Lean_Meta_Grind_propagateForallPropUp___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +return x_18; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +_start: +{ +uint8_t x_17; lean_object* x_18; +x_17 = lean_unbox(x_4); +lean_dec(x_4); +x_18 = l_Lean_Meta_Grind_propagateForallPropUp___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_7); +return x_18; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateForallPropUp___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +_start: +{ +uint8_t x_17; lean_object* x_18; +x_17 = lean_unbox(x_5); +lean_dec(x_5); +x_18 = l_Lean_Meta_Grind_propagateForallPropUp___lambda__3(x_1, x_2, x_3, x_4, x_17, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +lean_dec(x_7); +return x_18; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateForallPropUp___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +_start: +{ +uint8_t x_17; lean_object* x_18; +x_17 = lean_unbox(x_6); +lean_dec(x_6); +x_18 = l_Lean_Meta_Grind_propagateForallPropUp___lambda__4(x_1, x_2, x_3, x_4, x_5, x_17, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +lean_dec(x_7); +return x_18; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_isEqTrueHyp_x3f___lambda__1(lean_object* x_1) { +_start: +{ +if (lean_obj_tag(x_1) == 1) +{ +lean_object* x_2; lean_object* x_3; +x_2 = lean_ctor_get(x_1, 0); +lean_inc(x_2); +x_3 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_3, 0, x_2); +return x_3; +} +else +{ +lean_object* x_4; +x_4 = lean_box(0); +return x_4; +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_isEqTrueHyp_x3f___lambda__2(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_box(0); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_isEqTrueHyp_x3f___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_isEqTrueHyp_x3f___lambda__1___boxed), 1, 0); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_isEqTrueHyp_x3f___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_isEqTrueHyp_x3f___lambda__2___boxed), 1, 0); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_isEqTrueHyp_x3f___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_isEqTrueHyp_x3f___closed__2; +x_2 = lean_box(0); +x_3 = lean_apply_1(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_isEqTrueHyp_x3f___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("eq_true", 7, 7); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_isEqTrueHyp_x3f___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_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_isEqTrueHyp_x3f___closed__4; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_isEqTrueHyp_x3f(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; uint8_t x_4; +x_2 = l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_isEqTrueHyp_x3f___closed__1; +x_3 = l_Lean_Expr_cleanupAnnotations(x_1); +x_4 = l_Lean_Expr_isApp(x_3); +if (x_4 == 0) +{ +lean_object* x_5; +lean_dec(x_3); +x_5 = l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_isEqTrueHyp_x3f___closed__3; +return x_5; +} +else +{ +lean_object* x_6; lean_object* x_7; uint8_t x_8; +x_6 = l_Lean_Expr_appArg(x_3, lean_box(0)); +x_7 = l_Lean_Expr_appFnCleanup(x_3, lean_box(0)); +x_8 = l_Lean_Expr_isApp(x_7); +if (x_8 == 0) +{ +lean_object* x_9; +lean_dec(x_7); +lean_dec(x_6); +x_9 = l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_isEqTrueHyp_x3f___closed__3; +return x_9; +} +else +{ +lean_object* x_10; lean_object* x_11; uint8_t x_12; +x_10 = l_Lean_Expr_appFnCleanup(x_7, lean_box(0)); +x_11 = l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_isEqTrueHyp_x3f___closed__5; +x_12 = l_Lean_Expr_isConstOf(x_10, x_11); +lean_dec(x_10); +if (x_12 == 0) +{ +lean_object* x_13; +lean_dec(x_6); +x_13 = l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_isEqTrueHyp_x3f___closed__3; +return x_13; +} +else +{ +lean_object* x_14; +x_14 = lean_apply_1(x_2, x_6); +return x_14; +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_isEqTrueHyp_x3f___lambda__1___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_isEqTrueHyp_x3f___lambda__1(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_isEqTrueHyp_x3f___lambda__2___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_isEqTrueHyp_x3f___lambda__2(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_addLocalEMatchTheorems___lambda__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("issues", 6, 6); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_addLocalEMatchTheorems___lambda__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_propagateForallPropUp___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_addLocalEMatchTheorems___lambda__1___closed__1; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_addLocalEMatchTheorems___lambda__1___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("failed to create E-match local theorem for", 42, 42); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_addLocalEMatchTheorems___lambda__1___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_addLocalEMatchTheorems___lambda__1___closed__3; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_addLocalEMatchTheorems___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; uint8_t x_14; +x_13 = lean_st_ref_get(x_4, x_12); +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; +x_15 = lean_ctor_get(x_13, 0); +x_16 = lean_ctor_get(x_13, 1); +x_17 = lean_ctor_get(x_15, 9); +lean_inc(x_17); +lean_dec(x_15); +x_18 = lean_ctor_get(x_17, 2); +lean_inc(x_18); +lean_dec(x_17); +x_19 = lean_nat_dec_eq(x_18, x_1); +lean_dec(x_18); +if (x_19 == 0) +{ +lean_object* x_20; +lean_dec(x_2); +x_20 = lean_box(0); +lean_ctor_set(x_13, 0, x_20); +return x_13; +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +lean_free_object(x_13); +x_21 = l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_addLocalEMatchTheorems___lambda__1___closed__2; +x_22 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_21, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_16); +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +x_24 = lean_unbox(x_23); +lean_dec(x_23); +if (x_24 == 0) +{ +uint8_t x_25; +lean_dec(x_2); +x_25 = !lean_is_exclusive(x_22); +if (x_25 == 0) +{ +lean_object* x_26; lean_object* x_27; +x_26 = lean_ctor_get(x_22, 0); +lean_dec(x_26); +x_27 = lean_box(0); +lean_ctor_set(x_22, 0, x_27); +return x_22; +} +else +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_28 = lean_ctor_get(x_22, 1); +lean_inc(x_28); +lean_dec(x_22); +x_29 = lean_box(0); +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_28); +return x_30; +} +} +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_31 = lean_ctor_get(x_22, 1); +lean_inc(x_31); +lean_dec(x_22); +x_32 = l_Lean_indentExpr(x_2); +x_33 = l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_addLocalEMatchTheorems___lambda__1___closed__4; +x_34 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_34, 0, x_33); +lean_ctor_set(x_34, 1, x_32); +x_35 = l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__9; +x_36 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_36, 0, x_34); +lean_ctor_set(x_36, 1, x_35); +x_37 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_21, x_36, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_31); +return x_37; +} +} +} +else +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; uint8_t x_42; +x_38 = lean_ctor_get(x_13, 0); +x_39 = lean_ctor_get(x_13, 1); +lean_inc(x_39); +lean_inc(x_38); +lean_dec(x_13); +x_40 = lean_ctor_get(x_38, 9); +lean_inc(x_40); +lean_dec(x_38); +x_41 = lean_ctor_get(x_40, 2); +lean_inc(x_41); +lean_dec(x_40); +x_42 = lean_nat_dec_eq(x_41, x_1); +lean_dec(x_41); +if (x_42 == 0) +{ +lean_object* x_43; lean_object* x_44; +lean_dec(x_2); +x_43 = lean_box(0); +x_44 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_44, 0, x_43); +lean_ctor_set(x_44, 1, x_39); +return x_44; +} +else +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; uint8_t x_48; +x_45 = l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_addLocalEMatchTheorems___lambda__1___closed__2; +x_46 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_45, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_39); +x_47 = lean_ctor_get(x_46, 0); +lean_inc(x_47); +x_48 = lean_unbox(x_47); +lean_dec(x_47); +if (x_48 == 0) +{ +lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; +lean_dec(x_2); +x_49 = lean_ctor_get(x_46, 1); +lean_inc(x_49); +if (lean_is_exclusive(x_46)) { + lean_ctor_release(x_46, 0); + lean_ctor_release(x_46, 1); + x_50 = x_46; +} else { + lean_dec_ref(x_46); + x_50 = lean_box(0); +} +x_51 = lean_box(0); +if (lean_is_scalar(x_50)) { + x_52 = lean_alloc_ctor(0, 2, 0); +} else { + x_52 = x_50; +} +lean_ctor_set(x_52, 0, x_51); +lean_ctor_set(x_52, 1, x_49); +return x_52; +} +else +{ +lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; +x_53 = lean_ctor_get(x_46, 1); +lean_inc(x_53); +lean_dec(x_46); +x_54 = l_Lean_indentExpr(x_2); +x_55 = l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_addLocalEMatchTheorems___lambda__1___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 = l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__9; +x_58 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_58, 0, x_56); +lean_ctor_set(x_58, 1, x_57); +x_59 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_45, x_58, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_53); +return x_59; +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_addLocalEMatchTheorems___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +_start: +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_17 = lean_st_ref_get(x_8, x_16); +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +x_19 = lean_ctor_get(x_17, 1); +lean_inc(x_19); +lean_dec(x_17); +x_20 = lean_ctor_get(x_18, 9); +lean_inc(x_20); +lean_dec(x_18); +x_21 = lean_ctor_get(x_20, 2); +lean_inc(x_21); +lean_dec(x_20); +x_22 = lean_nat_dec_eq(x_21, x_1); +lean_dec(x_21); +if (x_22 == 0) +{ +lean_object* x_23; lean_object* x_24; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_23 = lean_box(0); +x_24 = l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_addLocalEMatchTheorems___lambda__1(x_1, x_2, x_23, x_8, x_9, 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); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +return x_24; +} +else +{ +uint8_t x_25; lean_object* x_26; +x_25 = 5; +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +x_26 = l_Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f(x_3, x_4, x_5, x_25, x_12, x_13, x_14, x_15, x_19); +if (lean_obj_tag(x_26) == 0) +{ +lean_object* x_27; +x_27 = lean_ctor_get(x_26, 0); +lean_inc(x_27); +if (lean_obj_tag(x_27) == 0) +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; +lean_dec(x_6); +x_28 = lean_ctor_get(x_26, 1); +lean_inc(x_28); +lean_dec(x_26); +x_29 = lean_box(0); +x_30 = l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_addLocalEMatchTheorems___lambda__1(x_1, x_2, x_29, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_28); +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); +return x_30; +} +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_26, 1); +lean_inc(x_31); +lean_dec(x_26); +x_32 = lean_ctor_get(x_27, 0); +lean_inc(x_32); +lean_dec(x_27); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +x_33 = l_Lean_Meta_Grind_activateTheorem(x_32, x_6, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_31); +if (lean_obj_tag(x_33) == 0) +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_33, 0); +lean_inc(x_34); +x_35 = lean_ctor_get(x_33, 1); +lean_inc(x_35); +lean_dec(x_33); +x_36 = l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_addLocalEMatchTheorems___lambda__1(x_1, x_2, x_34, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_35); +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_34); +return x_36; +} +else +{ +uint8_t x_37; +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_2); +x_37 = !lean_is_exclusive(x_33); +if (x_37 == 0) +{ +return x_33; +} +else +{ +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_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_41; +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_6); +lean_dec(x_2); +x_41 = !lean_is_exclusive(x_26); +if (x_41 == 0) +{ +return x_26; +} +else +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_42 = lean_ctor_get(x_26, 0); +x_43 = lean_ctor_get(x_26, 1); +lean_inc(x_43); +lean_inc(x_42); +lean_dec(x_26); +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_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_addLocalEMatchTheorems___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +_start: +{ +uint8_t x_17; lean_object* x_18; +x_17 = 4; +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_18 = l_Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f(x_1, x_2, x_3, x_17, x_12, x_13, x_14, x_15, x_16); +if (lean_obj_tag(x_18) == 0) +{ +lean_object* x_19; +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +if (lean_obj_tag(x_19) == 0) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_18, 1); +lean_inc(x_20); +lean_dec(x_18); +x_21 = lean_box(0); +x_22 = l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_addLocalEMatchTheorems___lambda__2(x_4, x_5, x_1, x_2, x_3, x_6, x_21, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_20); +return x_22; +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_18, 1); +lean_inc(x_23); +lean_dec(x_18); +x_24 = lean_ctor_get(x_19, 0); +lean_inc(x_24); +lean_dec(x_19); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_6); +x_25 = l_Lean_Meta_Grind_activateTheorem(x_24, x_6, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_23); +if (lean_obj_tag(x_25) == 0) +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; +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___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_addLocalEMatchTheorems___lambda__2(x_4, x_5, x_1, x_2, x_3, x_6, x_26, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_27); +lean_dec(x_26); +return x_28; +} +else +{ +uint8_t x_29; +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_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_29 = !lean_is_exclusive(x_25); +if (x_29 == 0) +{ +return x_25; +} +else +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_30 = lean_ctor_get(x_25, 0); +x_31 = lean_ctor_get(x_25, 1); +lean_inc(x_31); +lean_inc(x_30); +lean_dec(x_25); +x_32 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_32, 0, x_30); +lean_ctor_set(x_32, 1, x_31); +return x_32; +} +} +} +} +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_9); +lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_33 = !lean_is_exclusive(x_18); +if (x_33 == 0) +{ +return x_18; +} +else +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_18, 0); +x_35 = lean_ctor_get(x_18, 1); +lean_inc(x_35); +lean_inc(x_34); +lean_dec(x_18); +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___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_addLocalEMatchTheorems___lambda__4___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(0); +x_2 = lean_array_mk(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_addLocalEMatchTheorems___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) { +_start: +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_13 = l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__3; +lean_inc(x_1); +x_14 = l_Lean_mkAppB(x_13, x_1, x_2); +x_15 = lean_st_ref_get(x_4, 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 = lean_ctor_get(x_16, 9); +lean_inc(x_18); +lean_dec(x_16); +x_19 = lean_ctor_get(x_18, 2); +lean_inc(x_19); +lean_dec(x_18); +lean_inc(x_1); +x_20 = l_Lean_Meta_Grind_getGeneration(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_17); +if (lean_obj_tag(x_20) == 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, 0); +lean_inc(x_21); +x_22 = lean_ctor_get(x_20, 1); +lean_inc(x_22); +lean_dec(x_20); +x_23 = l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_addLocalEMatchTheorems___lambda__4___closed__1; +x_24 = 3; +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_14); +lean_inc(x_3); +x_25 = l_Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f(x_3, x_23, x_14, x_24, x_8, x_9, x_10, x_11, x_22); +if (lean_obj_tag(x_25) == 0) +{ +lean_object* x_26; +x_26 = lean_ctor_get(x_25, 0); +lean_inc(x_26); +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_25, 1); +lean_inc(x_27); +lean_dec(x_25); +x_28 = lean_box(0); +x_29 = l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_addLocalEMatchTheorems___lambda__3(x_3, x_23, x_14, x_19, x_1, x_21, x_28, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_27); +lean_dec(x_19); +return x_29; +} +else +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_30 = lean_ctor_get(x_25, 1); +lean_inc(x_30); +lean_dec(x_25); +x_31 = lean_ctor_get(x_26, 0); +lean_inc(x_31); +lean_dec(x_26); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_21); +x_32 = l_Lean_Meta_Grind_activateTheorem(x_31, x_21, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_30); +if (lean_obj_tag(x_32) == 0) +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_33 = lean_ctor_get(x_32, 0); +lean_inc(x_33); +x_34 = lean_ctor_get(x_32, 1); +lean_inc(x_34); +lean_dec(x_32); +x_35 = l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_addLocalEMatchTheorems___lambda__3(x_3, x_23, x_14, x_19, x_1, x_21, x_33, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_34); +lean_dec(x_33); +lean_dec(x_19); +return x_35; +} +else +{ +uint8_t x_36; +lean_dec(x_21); +lean_dec(x_19); +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_1); +x_36 = !lean_is_exclusive(x_32); +if (x_36 == 0) +{ +return x_32; +} +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_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; +} +} +} +} +else +{ +uint8_t x_40; +lean_dec(x_21); +lean_dec(x_19); +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_1); +x_40 = !lean_is_exclusive(x_25); +if (x_40 == 0) +{ +return x_25; +} +else +{ +lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_41 = lean_ctor_get(x_25, 0); +x_42 = lean_ctor_get(x_25, 1); +lean_inc(x_42); +lean_inc(x_41); +lean_dec(x_25); +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_44; +lean_dec(x_19); +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_1); +x_44 = !lean_is_exclusive(x_20); +if (x_44 == 0) +{ +return x_20; +} +else +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_45 = lean_ctor_get(x_20, 0); +x_46 = lean_ctor_get(x_20, 1); +lean_inc(x_46); +lean_inc(x_45); +lean_dec(x_20); +x_47 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_47, 0, x_45); +lean_ctor_set(x_47, 1, x_46); +return x_47; +} +} +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_addLocalEMatchTheorems___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("local", 5, 5); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_addLocalEMatchTheorems___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_addLocalEMatchTheorems___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_addLocalEMatchTheorems(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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); +lean_inc(x_2); +lean_inc(x_1); +x_11 = l_Lean_Meta_Grind_mkEqTrueProof(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_11) == 0) +{ +lean_object* x_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); +lean_inc(x_12); +x_14 = l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_isEqTrueHyp_x3f(x_12); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_15 = lean_st_ref_take(x_2, x_13); +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_15, 1); +lean_inc(x_17); +lean_dec(x_15); +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_25; lean_object* x_26; lean_object* x_27; +x_19 = lean_ctor_get(x_16, 19); +x_20 = lean_unsigned_to_nat(1u); +x_21 = lean_nat_add(x_19, x_20); +lean_ctor_set(x_16, 19, x_21); +x_22 = lean_st_ref_set(x_2, x_16, x_17); +x_23 = lean_ctor_get(x_22, 1); +lean_inc(x_23); +lean_dec(x_22); +x_24 = l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_addLocalEMatchTheorems___closed__2; +x_25 = lean_name_append_index_after(x_24, x_19); +x_26 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_26, 0, x_25); +x_27 = l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_addLocalEMatchTheorems___lambda__4(x_1, x_12, x_26, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_23); +return x_27; +} +else +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; 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; lean_object* x_50; lean_object* x_51; 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_28 = lean_ctor_get(x_16, 0); +x_29 = lean_ctor_get(x_16, 1); +x_30 = lean_ctor_get(x_16, 2); +x_31 = lean_ctor_get(x_16, 3); +x_32 = lean_ctor_get(x_16, 4); +x_33 = lean_ctor_get(x_16, 5); +x_34 = lean_ctor_get_uint8(x_16, sizeof(void*)*20); +x_35 = lean_ctor_get(x_16, 6); +x_36 = lean_ctor_get(x_16, 7); +x_37 = lean_ctor_get(x_16, 8); +x_38 = lean_ctor_get(x_16, 9); +x_39 = lean_ctor_get(x_16, 10); +x_40 = lean_ctor_get(x_16, 11); +x_41 = lean_ctor_get(x_16, 12); +x_42 = lean_ctor_get(x_16, 13); +x_43 = lean_ctor_get(x_16, 14); +x_44 = lean_ctor_get(x_16, 15); +x_45 = lean_ctor_get(x_16, 16); +x_46 = lean_ctor_get(x_16, 17); +x_47 = lean_ctor_get(x_16, 18); +x_48 = lean_ctor_get(x_16, 19); +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_inc(x_33); +lean_inc(x_32); +lean_inc(x_31); +lean_inc(x_30); +lean_inc(x_29); +lean_inc(x_28); +lean_dec(x_16); +x_49 = lean_unsigned_to_nat(1u); +x_50 = lean_nat_add(x_48, x_49); +x_51 = lean_alloc_ctor(0, 20, 1); +lean_ctor_set(x_51, 0, x_28); +lean_ctor_set(x_51, 1, x_29); +lean_ctor_set(x_51, 2, x_30); +lean_ctor_set(x_51, 3, x_31); +lean_ctor_set(x_51, 4, x_32); +lean_ctor_set(x_51, 5, x_33); +lean_ctor_set(x_51, 6, x_35); +lean_ctor_set(x_51, 7, x_36); +lean_ctor_set(x_51, 8, x_37); +lean_ctor_set(x_51, 9, x_38); +lean_ctor_set(x_51, 10, x_39); +lean_ctor_set(x_51, 11, x_40); +lean_ctor_set(x_51, 12, x_41); +lean_ctor_set(x_51, 13, x_42); +lean_ctor_set(x_51, 14, x_43); +lean_ctor_set(x_51, 15, x_44); +lean_ctor_set(x_51, 16, x_45); +lean_ctor_set(x_51, 17, x_46); +lean_ctor_set(x_51, 18, x_47); +lean_ctor_set(x_51, 19, x_50); +lean_ctor_set_uint8(x_51, sizeof(void*)*20, x_34); +x_52 = lean_st_ref_set(x_2, x_51, x_17); +x_53 = lean_ctor_get(x_52, 1); +lean_inc(x_53); +lean_dec(x_52); +x_54 = l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_addLocalEMatchTheorems___closed__2; +x_55 = lean_name_append_index_after(x_54, x_48); +x_56 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_56, 0, x_55); +x_57 = l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_addLocalEMatchTheorems___lambda__4(x_1, x_12, x_56, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_53); +return x_57; +} +} +else +{ +uint8_t x_58; +x_58 = !lean_is_exclusive(x_14); +if (x_58 == 0) +{ +lean_object* x_59; +x_59 = l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_addLocalEMatchTheorems___lambda__4(x_1, x_12, x_14, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_13); +return x_59; +} +else +{ +lean_object* x_60; lean_object* x_61; lean_object* x_62; +x_60 = lean_ctor_get(x_14, 0); +lean_inc(x_60); +lean_dec(x_14); +x_61 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_61, 0, x_60); +x_62 = l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_addLocalEMatchTheorems___lambda__4(x_1, x_12, x_61, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_13); +return x_62; +} +} +} +else +{ +uint8_t x_63; +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_63 = !lean_is_exclusive(x_11); +if (x_63 == 0) +{ +return x_11; +} +else +{ +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_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; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_addLocalEMatchTheorems___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_13; +x_13 = l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_addLocalEMatchTheorems___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +return x_13; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_addLocalEMatchTheorems___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +_start: +{ +lean_object* x_17; +x_17 = l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_addLocalEMatchTheorems___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +lean_dec(x_7); +lean_dec(x_1); +return x_17; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_addLocalEMatchTheorems___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +_start: +{ +lean_object* x_17; +x_17 = l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_addLocalEMatchTheorems___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +lean_dec(x_7); +lean_dec(x_4); +return x_17; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateForallPropDown___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("eq_true_of_imp_eq_false", 23, 23); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateForallPropDown___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__1; +x_2 = l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__2; +x_3 = l_Lean_Meta_Grind_propagateForallPropDown___closed__1; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateForallPropDown___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_Grind_propagateForallPropDown___closed__2; +x_3 = l_Lean_Expr_const___override(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateForallPropDown___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("eq_false_of_imp_eq_false", 24, 24); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateForallPropDown___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__1; +x_2 = l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__2; +x_3 = l_Lean_Meta_Grind_propagateForallPropDown___closed__4; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateForallPropDown___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind_propagateForallPropDown___closed__5; +x_3 = l_Lean_Expr_const___override(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateForallPropDown___closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Exists", 6, 6); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateForallPropDown___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_Grind_propagateForallPropDown___closed__7; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateForallPropDown___closed__9() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("of_forall_eq_false", 18, 18); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateForallPropDown___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__1; +x_2 = l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__2; +x_3 = l_Lean_Meta_Grind_propagateForallPropDown___closed__9; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateForallPropDown(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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) == 7) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; lean_object* x_15; +x_11 = lean_ctor_get(x_1, 0); +lean_inc(x_11); +x_12 = lean_ctor_get(x_1, 1); +lean_inc(x_12); +x_13 = lean_ctor_get(x_1, 2); +lean_inc(x_13); +x_14 = lean_ctor_get_uint8(x_1, sizeof(void*)*3 + 8); +lean_inc(x_1); +x_15 = l_Lean_Meta_Grind_isEqFalse(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_15) == 0) +{ +lean_object* x_16; uint8_t x_17; +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +x_17 = lean_unbox(x_16); +lean_dec(x_16); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; +lean_dec(x_12); +lean_dec(x_11); +x_18 = lean_ctor_get(x_15, 1); +lean_inc(x_18); +lean_dec(x_15); +lean_inc(x_1); +x_19 = l_Lean_Meta_Grind_isEqTrue(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_18); +if (lean_obj_tag(x_19) == 0) +{ +lean_object* x_20; uint8_t x_21; +x_20 = lean_ctor_get(x_19, 0); +lean_inc(x_20); +x_21 = lean_unbox(x_20); +lean_dec(x_20); +if (x_21 == 0) +{ +uint8_t x_22; +lean_dec(x_13); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -479,29 +3124,333 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_27 = !lean_is_exclusive(x_15); -if (x_27 == 0) +x_22 = !lean_is_exclusive(x_19); +if (x_22 == 0) { -return x_15; +lean_object* x_23; lean_object* x_24; +x_23 = lean_ctor_get(x_19, 0); +lean_dec(x_23); +x_24 = lean_box(0); +lean_ctor_set(x_19, 0, x_24); +return x_19; } else { -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_object* x_25; lean_object* x_26; lean_object* x_27; +x_25 = lean_ctor_get(x_19, 1); +lean_inc(x_25); +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_25); +return x_27; +} +} +else +{ +uint8_t x_28; +x_28 = !lean_is_exclusive(x_19); +if (x_28 == 0) +{ +lean_object* x_29; lean_object* x_30; uint8_t x_31; +x_29 = lean_ctor_get(x_19, 1); +x_30 = lean_ctor_get(x_19, 0); +lean_dec(x_30); +x_31 = l_Lean_Expr_hasLooseBVars(x_13); +lean_dec(x_13); +if (x_31 == 0) +{ +lean_object* x_32; +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_32 = lean_box(0); +lean_ctor_set(x_19, 0, x_32); +return x_19; +} +else +{ +lean_object* x_33; +lean_free_object(x_19); +x_33 = l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_addLocalEMatchTheorems(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_29); +return x_33; +} +} +else +{ +lean_object* x_34; uint8_t x_35; +x_34 = lean_ctor_get(x_19, 1); +lean_inc(x_34); +lean_dec(x_19); +x_35 = l_Lean_Expr_hasLooseBVars(x_13); +lean_dec(x_13); +if (x_35 == 0) +{ +lean_object* x_36; lean_object* x_37; +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_36 = lean_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_34); +return x_37; +} +else +{ +lean_object* x_38; +x_38 = l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_addLocalEMatchTheorems(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_34); +return x_38; +} +} +} +} +else +{ +uint8_t x_39; +lean_dec(x_13); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_39 = !lean_is_exclusive(x_19); +if (x_39 == 0) +{ +return x_19; +} +else +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_40 = lean_ctor_get(x_19, 0); +x_41 = lean_ctor_get(x_19, 1); +lean_inc(x_41); +lean_inc(x_40); +lean_dec(x_19); +x_42 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_42, 0, x_40); +lean_ctor_set(x_42, 1, x_41); +return x_42; +} +} +} +else +{ +lean_object* x_43; uint8_t x_44; +x_43 = lean_ctor_get(x_15, 1); +lean_inc(x_43); lean_dec(x_15); -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_44 = l_Lean_Expr_hasLooseBVars(x_13); +if (x_44 == 0) +{ +lean_object* x_45; +lean_dec(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); +lean_inc(x_2); +x_45 = l_Lean_Meta_Grind_mkEqFalseProof(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_43); +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; lean_object* x_53; lean_object* x_54; +x_46 = lean_ctor_get(x_45, 0); +lean_inc(x_46); +x_47 = lean_ctor_get(x_45, 1); +lean_inc(x_47); +lean_dec(x_45); +x_48 = l_Lean_Meta_Grind_propagateForallPropDown___closed__3; +lean_inc(x_46); +lean_inc(x_13); +lean_inc(x_12); +x_49 = l_Lean_mkApp3(x_48, x_12, x_13, x_46); +lean_inc(x_12); +x_50 = l_Lean_Meta_Grind_pushEqTrue(x_12, x_49, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_47); +x_51 = lean_ctor_get(x_50, 1); +lean_inc(x_51); +lean_dec(x_50); +x_52 = l_Lean_Meta_Grind_propagateForallPropDown___closed__6; +lean_inc(x_13); +x_53 = l_Lean_mkApp3(x_52, x_12, x_13, x_46); +x_54 = l_Lean_Meta_Grind_pushEqFalse(x_13, x_53, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_51); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_54; +} +else +{ +uint8_t x_55; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_55 = !lean_is_exclusive(x_45); +if (x_55 == 0) +{ +return x_45; +} +else +{ +lean_object* x_56; lean_object* x_57; lean_object* x_58; +x_56 = lean_ctor_get(x_45, 0); +x_57 = lean_ctor_get(x_45, 1); +lean_inc(x_57); +lean_inc(x_56); +lean_dec(x_45); +x_58 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_58, 0, x_56); +lean_ctor_set(x_58, 1, x_57); +return x_58; } } } else { -lean_object* x_31; lean_object* x_32; +lean_object* x_59; +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_12); +x_59 = l_Lean_Meta_getLevel(x_12, x_6, x_7, x_8, x_9, x_43); +if (lean_obj_tag(x_59) == 0) +{ +lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; +x_60 = lean_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 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_63, 0, x_60); +lean_ctor_set(x_63, 1, x_62); +x_64 = l_Lean_Meta_Grind_propagateForallPropDown___closed__8; +lean_inc(x_63); +x_65 = l_Lean_Expr_const___override(x_64, x_63); +lean_inc(x_13); +x_66 = l_Lean_mkNot(x_13); +lean_inc(x_12); +lean_inc(x_11); +x_67 = l_Lean_Expr_lam___override(x_11, x_12, x_66, x_14); +lean_inc(x_12); +x_68 = l_Lean_mkAppB(x_65, x_12, x_67); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_69 = l_Lean_Meta_Grind_mkEqFalseProof(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_61); +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; lean_object* x_76; +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 = l_Lean_Meta_Grind_propagateForallPropDown___closed__10; +x_73 = l_Lean_Expr_const___override(x_72, x_63); +lean_inc(x_12); +x_74 = l_Lean_Expr_lam___override(x_11, x_12, x_13, x_14); +x_75 = l_Lean_mkApp3(x_73, x_12, x_74, x_70); +x_76 = l_Lean_Meta_Grind_getGeneration(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_71); +if (lean_obj_tag(x_76) == 0) +{ +lean_object* x_77; lean_object* x_78; lean_object* x_79; +x_77 = lean_ctor_get(x_76, 0); +lean_inc(x_77); +x_78 = lean_ctor_get(x_76, 1); +lean_inc(x_78); +lean_dec(x_76); +x_79 = l_Lean_Meta_Grind_addNewFact(x_75, x_68, x_77, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, 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_3); +lean_dec(x_2); +return x_79; +} +else +{ +uint8_t x_80; +lean_dec(x_75); +lean_dec(x_68); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_80 = !lean_is_exclusive(x_76); +if (x_80 == 0) +{ +return x_76; +} +else +{ +lean_object* x_81; lean_object* x_82; lean_object* x_83; +x_81 = lean_ctor_get(x_76, 0); +x_82 = lean_ctor_get(x_76, 1); +lean_inc(x_82); +lean_inc(x_81); +lean_dec(x_76); +x_83 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_83, 0, x_81); +lean_ctor_set(x_83, 1, x_82); +return x_83; +} +} +} +else +{ +uint8_t x_84; +lean_dec(x_68); +lean_dec(x_63); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -511,23 +3460,116 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_31 = lean_box(0); -x_32 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_32, 0, x_31); -lean_ctor_set(x_32, 1, x_10); -return x_32; +x_84 = !lean_is_exclusive(x_69); +if (x_84 == 0) +{ +return x_69; } +else +{ +lean_object* x_85; lean_object* x_86; lean_object* x_87; +x_85 = lean_ctor_get(x_69, 0); +x_86 = lean_ctor_get(x_69, 1); +lean_inc(x_86); +lean_inc(x_85); +lean_dec(x_69); +x_87 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_87, 0, x_85); +lean_ctor_set(x_87, 1, x_86); +return x_87; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateForallProp___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { -_start: +} +else +{ +uint8_t x_88; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_88 = !lean_is_exclusive(x_59); +if (x_88 == 0) +{ +return x_59; +} +else +{ +lean_object* x_89; lean_object* x_90; lean_object* x_91; +x_89 = lean_ctor_get(x_59, 0); +x_90 = lean_ctor_get(x_59, 1); +lean_inc(x_90); +lean_inc(x_89); +lean_dec(x_59); +x_91 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_91, 0, x_89); +lean_ctor_set(x_91, 1, x_90); +return x_91; +} +} +} +} +} +else { -uint8_t x_16; lean_object* x_17; -x_16 = lean_unbox(x_4); +uint8_t x_92; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); lean_dec(x_4); -x_17 = l_Lean_Meta_Grind_propagateForallProp___lambda__1(x_1, x_2, x_3, x_16, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_92 = !lean_is_exclusive(x_15); +if (x_92 == 0) +{ +return x_15; +} +else +{ +lean_object* x_93; lean_object* x_94; lean_object* x_95; +x_93 = lean_ctor_get(x_15, 0); +x_94 = lean_ctor_get(x_15, 1); +lean_inc(x_94); +lean_inc(x_93); +lean_dec(x_15); +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_96; lean_object* x_97; +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); lean_dec(x_6); -return x_17; +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_96 = lean_box(0); +x_97 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_97, 0, x_96); +lean_ctor_set(x_97, 1, x_10); +return x_97; +} } } lean_object* initialize_Init_Grind_Lemmas(uint8_t builtin, lean_object*); @@ -551,22 +3593,106 @@ lean_dec_ref(res); res = initialize_Lean_Meta_Tactic_Grind_Simp(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__1 = _init_l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__1(); -lean_mark_persistent(l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__1); -l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__2 = _init_l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__2(); -lean_mark_persistent(l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__2); -l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__3 = _init_l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__3(); -lean_mark_persistent(l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__3); -l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__4 = _init_l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__4(); -lean_mark_persistent(l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__4); -l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__5 = _init_l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__5(); -lean_mark_persistent(l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__5); -l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__6 = _init_l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__6(); -lean_mark_persistent(l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__6); -l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__7 = _init_l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__7(); -lean_mark_persistent(l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__7); -l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__8 = _init_l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__8(); -lean_mark_persistent(l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__8); +l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__1 = _init_l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__1); +l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__2 = _init_l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__2); +l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__3 = _init_l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__3); +l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__4 = _init_l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__4(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__4); +l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__5 = _init_l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__5(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__5); +l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__6 = _init_l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__6(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__6); +l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__7 = _init_l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__7(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__7); +l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__8 = _init_l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__8(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__8); +l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__9 = _init_l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__9(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__9); +l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__10 = _init_l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__10(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__10); +l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__11 = _init_l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__11(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__11); +l_Lean_Meta_Grind_propagateForallPropUp___lambda__1___closed__1 = _init_l_Lean_Meta_Grind_propagateForallPropUp___lambda__1___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateForallPropUp___lambda__1___closed__1); +l_Lean_Meta_Grind_propagateForallPropUp___lambda__1___closed__2 = _init_l_Lean_Meta_Grind_propagateForallPropUp___lambda__1___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateForallPropUp___lambda__1___closed__2); +l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__1 = _init_l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__1); +l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__2 = _init_l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__2); +l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__3 = _init_l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__3); +l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__4 = _init_l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__4(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__4); +l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__5 = _init_l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__5(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__5); +l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__6 = _init_l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__6(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__6); +l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__7 = _init_l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__7(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__7); +l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__8 = _init_l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__8(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__8); +l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__9 = _init_l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__9(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__9); +l_Lean_Meta_Grind_propagateForallPropUp___lambda__3___closed__1 = _init_l_Lean_Meta_Grind_propagateForallPropUp___lambda__3___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateForallPropUp___lambda__3___closed__1); +l_Lean_Meta_Grind_propagateForallPropUp___lambda__3___closed__2 = _init_l_Lean_Meta_Grind_propagateForallPropUp___lambda__3___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateForallPropUp___lambda__3___closed__2); +l_Lean_Meta_Grind_propagateForallPropUp___closed__1 = _init_l_Lean_Meta_Grind_propagateForallPropUp___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateForallPropUp___closed__1); +l_Lean_Meta_Grind_propagateForallPropUp___closed__2 = _init_l_Lean_Meta_Grind_propagateForallPropUp___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateForallPropUp___closed__2); +l_Lean_Meta_Grind_propagateForallPropUp___closed__3 = _init_l_Lean_Meta_Grind_propagateForallPropUp___closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateForallPropUp___closed__3); +l_Lean_Meta_Grind_propagateForallPropUp___closed__4 = _init_l_Lean_Meta_Grind_propagateForallPropUp___closed__4(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateForallPropUp___closed__4); +l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_isEqTrueHyp_x3f___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_isEqTrueHyp_x3f___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_isEqTrueHyp_x3f___closed__1); +l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_isEqTrueHyp_x3f___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_isEqTrueHyp_x3f___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_isEqTrueHyp_x3f___closed__2); +l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_isEqTrueHyp_x3f___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_isEqTrueHyp_x3f___closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_isEqTrueHyp_x3f___closed__3); +l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_isEqTrueHyp_x3f___closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_isEqTrueHyp_x3f___closed__4(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_isEqTrueHyp_x3f___closed__4); +l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_isEqTrueHyp_x3f___closed__5 = _init_l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_isEqTrueHyp_x3f___closed__5(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_isEqTrueHyp_x3f___closed__5); +l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_addLocalEMatchTheorems___lambda__1___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_addLocalEMatchTheorems___lambda__1___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_addLocalEMatchTheorems___lambda__1___closed__1); +l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_addLocalEMatchTheorems___lambda__1___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_addLocalEMatchTheorems___lambda__1___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_addLocalEMatchTheorems___lambda__1___closed__2); +l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_addLocalEMatchTheorems___lambda__1___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_addLocalEMatchTheorems___lambda__1___closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_addLocalEMatchTheorems___lambda__1___closed__3); +l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_addLocalEMatchTheorems___lambda__1___closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_addLocalEMatchTheorems___lambda__1___closed__4(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_addLocalEMatchTheorems___lambda__1___closed__4); +l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_addLocalEMatchTheorems___lambda__4___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_addLocalEMatchTheorems___lambda__4___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_addLocalEMatchTheorems___lambda__4___closed__1); +l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_addLocalEMatchTheorems___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_addLocalEMatchTheorems___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_addLocalEMatchTheorems___closed__1); +l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_addLocalEMatchTheorems___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_addLocalEMatchTheorems___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_ForallProp_0__Lean_Meta_Grind_addLocalEMatchTheorems___closed__2); +l_Lean_Meta_Grind_propagateForallPropDown___closed__1 = _init_l_Lean_Meta_Grind_propagateForallPropDown___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateForallPropDown___closed__1); +l_Lean_Meta_Grind_propagateForallPropDown___closed__2 = _init_l_Lean_Meta_Grind_propagateForallPropDown___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateForallPropDown___closed__2); +l_Lean_Meta_Grind_propagateForallPropDown___closed__3 = _init_l_Lean_Meta_Grind_propagateForallPropDown___closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateForallPropDown___closed__3); +l_Lean_Meta_Grind_propagateForallPropDown___closed__4 = _init_l_Lean_Meta_Grind_propagateForallPropDown___closed__4(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateForallPropDown___closed__4); +l_Lean_Meta_Grind_propagateForallPropDown___closed__5 = _init_l_Lean_Meta_Grind_propagateForallPropDown___closed__5(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateForallPropDown___closed__5); +l_Lean_Meta_Grind_propagateForallPropDown___closed__6 = _init_l_Lean_Meta_Grind_propagateForallPropDown___closed__6(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateForallPropDown___closed__6); +l_Lean_Meta_Grind_propagateForallPropDown___closed__7 = _init_l_Lean_Meta_Grind_propagateForallPropDown___closed__7(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateForallPropDown___closed__7); +l_Lean_Meta_Grind_propagateForallPropDown___closed__8 = _init_l_Lean_Meta_Grind_propagateForallPropDown___closed__8(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateForallPropDown___closed__8); +l_Lean_Meta_Grind_propagateForallPropDown___closed__9 = _init_l_Lean_Meta_Grind_propagateForallPropDown___closed__9(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateForallPropDown___closed__9); +l_Lean_Meta_Grind_propagateForallPropDown___closed__10 = _init_l_Lean_Meta_Grind_propagateForallPropDown___closed__10(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateForallPropDown___closed__10); return lean_io_result_mk_ok(lean_box(0)); } #ifdef __cplusplus diff --git a/stage0/stdlib/Lean/Meta/Tactic/Grind/Internalize.c b/stage0/stdlib/Lean/Meta/Tactic/Grind/Internalize.c index c0a9bbeba790..3a1b248b6ded 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Grind/Internalize.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Grind/Internalize.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Lean.Meta.Tactic.Grind.Internalize -// Imports: Init.Grind.Util Lean.Meta.LitValues Lean.Meta.Tactic.Grind.Types Lean.Meta.Tactic.Grind.Util +// Imports: Init.Grind.Util Init.Grind.Lemmas Lean.Meta.LitValues Lean.Meta.Match.MatcherInfo Lean.Meta.Match.MatchEqsExt Lean.Meta.Tactic.Grind.Types Lean.Meta.Tactic.Grind.Util #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -13,233 +13,314 @@ #ifdef __cplusplus extern "C" { #endif +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); +uint8_t lean_is_matcher(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__2___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_internalize___lambda__3___closed__6; +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAtAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addSplitCandidate___closed__3; +static lean_object* l_Lean_Meta_Grind_internalize___lambda__3___closed__5; extern lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare; static lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_addCongrTable___spec__5___closed__1; +uint8_t l_Lean_PersistentHashMap_contains___at_Lean_NameSSet_contains___spec__2(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_addCongrTable___closed__1; -static lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__3; -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_internalize___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_Grind_internalize___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_PersistentHashMap_find_x3f___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__1___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__8; +static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__3; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__6; lean_object* l_Lean_mkAppN(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__4___boxed(lean_object*, lean_object*, lean_object*); size_t lean_usize_shift_right(size_t, size_t); static lean_object* l_Lean_Meta_Grind_addCongrTable___closed__9; LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___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_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__5(lean_object*, size_t, size_t, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__1___closed__4; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__2; +static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__2; lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__2___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__2; +uint8_t l_Lean_Expr_isDIte(lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__3; lean_object* l_Lean_Meta_Grind_eraseIrrelevantMData___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_usize_dec_le(size_t, size_t); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addSplitCandidate(lean_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_Grind_addCongrTable___closed__4; +LEAN_EXPORT lean_object* l_List_filterTR_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__4(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___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__1; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_addCongrTable___spec__6(lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__2; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__2___boxed(lean_object**); lean_object* l_Lean_Meta_Grind_unfoldReducible___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_filterTR_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__6(lean_object*, lean_object*, lean_object*); size_t lean_uint64_to_usize(uint64_t); lean_object* l_Lean_Meta_Grind_mkENode(lean_object*, 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_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__4; +static lean_object* l_Lean_Meta_Grind_activateTheorem___closed__1; +uint8_t l_Lean_Expr_isApp(lean_object*); +static lean_object* l_Lean_Meta_Grind_internalize___lambda__3___closed__4; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findEntry_x3f___at_Lean_Meta_Grind_addCongrTable___spec__1(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate___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*); uint8_t l_Lean_Meta_Grind_isCongruent(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns___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_PersistentHashMap_insertAux_traverse___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__6(size_t, 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_Grind_internalize___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofList(lean_object*); lean_object* l_Lean_PersistentArray_push___rarg(lean_object*, lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); -LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_contains___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__3(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_internalize___lambda__2___closed__4; +static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__4; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addCongrTable___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*); size_t lean_usize_mul(size_t, size_t); static lean_object* l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__3; +lean_object* l_Lean_PersistentHashMap_insert___at_Lean_NameSSet_insert___spec__2(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__6; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addCongrTable___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_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__1; LEAN_EXPORT lean_object* l_panic___at_Lean_Meta_Grind_internalize___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_mk_array(lean_object*, lean_object*); -static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__2; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addSplitCandidate___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_Grind_activateTheorem___closed__2; uint64_t l_Lean_Meta_Grind_congrHash(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addCongrTable___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_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__4; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns(lean_object*, lean_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_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes; lean_object* lean_array_fget(lean_object*, lean_object*); -static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__9; lean_object* l_Lean_Meta_Grind_normalizeLevels___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__4; +lean_object* l_List_mapTR_loop___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__2(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___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* l_Lean_Meta_Grind_pushEqCore(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_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__1; static lean_object* l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__2; lean_object* l_Lean_Meta_Grind_setENode(lean_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_cleanupAnnotations(lean_object*); lean_object* l_Lean_stringToMessageData(lean_object*); -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___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_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___boxed(lean_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_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addSplitCandidate___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_PersistentHashMap_findEntryAtAux___at_Lean_Meta_Grind_addCongrTable___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_internalize___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_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns___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_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_internalize___lambda__3___closed__1; -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_eraseAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__2(lean_object*, size_t, lean_object*); +uint8_t l_Lean_Meta_Grind_EMatchTheorems_isErased(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__7; -static lean_object* l_Lean_Meta_Grind_internalize___lambda__2___closed__3; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__1___closed__1; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addSplitCandidate___closed__1; lean_object* l_Lean_Meta_Grind_propagateUp(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_isBVar(lean_object*); extern lean_object* l_instInhabitedPUnit; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addCongrTable___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__3(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__4; -static lean_object* l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9___closed__2; lean_object* l_Lean_Meta_Grind_EMatchTheorems_insert(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_activateTheorem(lean_object*, lean_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_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__2___closed__2; +lean_object* l_Lean_mkApp6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_contains___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__1___boxed(lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); static size_t l_Lean_PersistentHashMap_findEntryAux___at_Lean_Meta_Grind_addCongrTable___spec__2___closed__2; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_addCongrTable___spec__5(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_internalize___lambda__2___closed__2; -lean_object* l_Lean_PersistentHashMap_isUnaryNode___rarg(lean_object*); +static lean_object* l_Lean_Meta_Grind_internalize___lambda__3___closed__3; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__4___closed__1; +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_List_elem___at_Lean_addAliasEntry___spec__16(lean_object*, lean_object*); lean_object* l_Lean_Core_transform___at_Lean_Meta_Grind_unfoldReducible___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_contains___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__1(lean_object*, lean_object*); lean_object* lean_st_ref_take(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__8; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns(lean_object*, lean_object*, 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_insertAux_traverse___at_Lean_Meta_Grind_addCongrTable___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___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*); uint8_t lean_expr_eqv(lean_object*, lean_object*); -static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__3; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__2___closed__1; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__8; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__5; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAtAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns___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_Lean_PersistentHashMap_findAtAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addCongrTable___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* l_Lean_Meta_Grind_groundPattern_x3f(lean_object*); -static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__1___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_isMatcher___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns___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_Grind_addCongrTable___closed__5; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate___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_PersistentHashMap_findAtAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__2; +lean_object* l_Lean_Meta_isInductivePredicate(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate___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_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__9(lean_object*, lean_object*); +lean_object* l_Lean_Expr_appArg(lean_object*, lean_object*); lean_object* l_outOfBounds___rarg(lean_object*); -LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__4(lean_object*, size_t, lean_object*); -static lean_object* l_Lean_Meta_Grind_internalize___lambda__2___closed__1; lean_object* lean_st_ref_get(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__4; +lean_object* l_Lean_Meta_reduceMatcher_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___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_Meta_Grind_addCongrTable___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_EXPORT lean_object* l_Lean_Meta_Grind_activateTheorem___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_Expr_appFnCleanup(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_isMatcherApp___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate___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_PersistentHashMap_insert___at_Lean_Meta_Grind_addCongrTable___spec__4(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addCongrTable(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_hasLooseBVars(lean_object*); lean_object* l_Lean_Expr_toHeadIndex(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findEntryAux___at_Lean_Meta_Grind_addCongrTable___spec__2(lean_object*, lean_object*, size_t, lean_object*); -static lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__4; static size_t l_Lean_PersistentHashMap_findEntryAux___at_Lean_Meta_Grind_addCongrTable___spec__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___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_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_levelZero; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__3___closed__2; +static lean_object* l_Lean_Meta_Grind_activateTheorem___closed__4; extern lean_object* l_Lean_instInhabitedExpr; +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_internalize___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_addCongrTable___closed__7; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__11; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_addCongrTable___spec__7(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_PersistentHashMap_insert___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__4(lean_object*, lean_object*, lean_object*); +static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__6; +lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___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*); uint8_t l_Lean_Meta_Grind_isSameExpr_unsafe__1(lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_alreadyInternalized(lean_object*, 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_Meta_Grind_addCongrTable___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__5; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__3___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_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___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*); +uint8_t l_Lean_Expr_isIte(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__2(lean_object*, size_t, lean_object*); static lean_object* l_Lean_Meta_Grind_addCongrTable___closed__6; -static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__7; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs(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_Grind_mkEMatchEqTheorem(lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_Grind_updateLastTag(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_insertAux___at_Lean_Meta_Grind_addCongrTable___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Meta_instMonadMetaM; -lean_object* l_Array_eraseIdx___rarg(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_hasSameType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_activateTheorem___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_usize_to_nat(size_t); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap(lean_object*, 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_Grind_internalize___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofExpr(lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__1___closed__2; +static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__5; +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___boxed(lean_object**); +static lean_object* l_Lean_Meta_Grind_internalize___lambda__4___closed__1; static lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__1; -double l_Float_ofScientific(lean_object*, uint8_t, lean_object*); +lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___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_Expr_withAppAux___at_Lean_Meta_Grind_internalize___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* l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__2(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___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*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate___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_Core_transform___at_Lean_Core_betaReduce___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___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_PersistentHashMap_find_x3f___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__1(lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_mkGroundPattern(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findEntryAtAux___at_Lean_Meta_Grind_addCongrTable___spec__3(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_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_internalize___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_EXPORT lean_object* l_Lean_Meta_Grind_internalize___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__10; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns___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_Grind_addCongrTable___lambda__2___closed__5; LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___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*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_shareCommon(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_containsAtAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___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___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___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___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__3; lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); +uint8_t l_Lean_Meta_isMatcherAppCore(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__1___closed__3; lean_object* l_Lean_Meta_Grind_registerParent(lean_object*, lean_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_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___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* l_Lean_indentExpr(lean_object*); +static lean_object* l_Lean_Meta_Grind_activateTheorem___closed__3; +lean_object* l_Lean_PersistentHashMap_erase___at_Lean_Meta_Grind_EMatchTheorems_retrieve_x3f___spec__1(lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_mkEmptyEntries(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__3___boxed(lean_object**); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__9; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__4___closed__2; uint8_t l_Lean_Expr_isConstOf(lean_object*, lean_object*); static lean_object* l_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__5; lean_object* lean_array_set(lean_object*, lean_object*, lean_object*); -uint64_t l_Lean_Name_hash___override(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_PersistentHashMap_contains___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__3___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__7; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__7; LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___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_panic_fn(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_internalize___lambda__2___closed__6; -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_erase___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__1___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__1; +lean_object* lean_get_match_equations_for(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_PersistentHashMap_findEntryAux___at_Lean_Meta_Grind_addCongrTable___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_getAppFn(lean_object*); -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___boxed(lean_object**); -lean_object* l_Array_indexOfAux___at_Lean_MetavarContext_setMVarUserName___spec__3(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_mkCollisionNode___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_List_mapTR_loop___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__1(lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_getENode(lean_object*, 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_Grind_internalize___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_instantiateMVarsIfMVarApp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___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*); static lean_object* l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__6; lean_object* l_Lean_Meta_Grind_mkENodeCore(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_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_internalize___spec__2___boxed(lean_object**); -LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAtAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_internalize___lambda__2___closed__5; +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__2(lean_object*, size_t, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__3; lean_object* l_List_reverse___rarg(lean_object*); -static double l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9___closed__1; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__12; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___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*); size_t lean_usize_sub(size_t, size_t); -lean_object* lean_array_mk(lean_object*); static lean_object* l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addCongrTable___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_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__6; -lean_object* l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__4(lean_object*, lean_object*); -lean_object* l_Lean_isTracingEnabledForCore(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_internalize___lambda__4___closed__2; +static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__1; size_t lean_usize_add(size_t, size_t); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_eraseAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__2___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9___closed__3; lean_object* lean_array_uget(lean_object*, size_t); size_t lean_array_size(lean_object*); lean_object* l_instInhabitedOfMonad___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__1___boxed(lean_object**); lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern(lean_object*, lean_object*, 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_isMatcher___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_shift_left(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__7(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__1; uint8_t l___private_Lean_HeadIndex_0__Lean_beqHeadIndex____x40_Lean_HeadIndex___hyg_69_(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_addCongrTable___closed__8; +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___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_Meta_Grind_canon(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*); static lean_object* l_Lean_Meta_Grind_addCongrTable___closed__2; static lean_object* l_Lean_Meta_Grind_internalize___lambda__3___closed__2; -LEAN_EXPORT lean_object* l_List_mapM_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_Origin_key(lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_erase___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__1(lean_object*, lean_object*); uint8_t lean_usize_dec_lt(size_t, size_t); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns___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_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__3; lean_object* lean_nat_add(lean_object*, lean_object*); extern lean_object* l_Lean_Meta_Grind_congrPlaceholderProof; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addSplitCandidate___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_PersistentHashMap_getCollisionNodeSize___rarg(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate___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___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint64_t l_Lean_HeadIndex_hash(lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addSplitCandidate___closed__2; static lean_object* l_Lean_Meta_Grind_addCongrTable___closed__3; +static lean_object* l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__8; static lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__2; lean_object* l_Lean_Meta_Grind_unfoldReducible___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_Grind_getConfig___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_uset(lean_object*, size_t, lean_object*); +LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_Meta_Grind_activateTheorem___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofName(lean_object*); lean_object* lean_array_get(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_isLitValue(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__10; lean_object* l_ReaderT_instMonad___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_internalize(lean_object*, lean_object*, 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_Grind_internalize___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_Grind_internalize___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*); size_t lean_usize_land(size_t, size_t); -static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__5; +LEAN_EXPORT lean_object* l_Lean_Meta_isMatcherApp___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate___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___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__5; static lean_object* l_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__4; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findEntryAtAux___at_Lean_Meta_Grind_addCongrTable___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: @@ -527,6 +608,7 @@ LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findEntry_x3f___at_Lean_Meta_G lean_object* x_4; uint64_t x_5; size_t x_6; lean_object* x_7; x_4 = lean_ctor_get(x_1, 1); lean_inc(x_4); +lean_inc(x_3); x_5 = l_Lean_Meta_Grind_congrHash(x_4, x_3); x_6 = lean_uint64_to_usize(x_5); x_7 = l_Lean_PersistentHashMap_findEntryAux___at_Lean_Meta_Grind_addCongrTable___spec__2(x_1, x_2, x_6, x_3); @@ -554,6 +636,7 @@ else lean_object* x_11; lean_object* x_12; uint64_t x_13; size_t x_14; size_t x_15; size_t x_16; size_t x_17; size_t x_18; size_t x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; x_11 = lean_array_fget(x_3, x_6); x_12 = lean_array_fget(x_4, x_6); +lean_inc(x_11); x_13 = l_Lean_Meta_Grind_congrHash(x_8, x_11); x_14 = lean_uint64_to_usize(x_13); x_15 = 1; @@ -1076,6 +1159,7 @@ LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_ad lean_object* x_5; uint64_t x_6; size_t x_7; size_t x_8; lean_object* x_9; x_5 = lean_ctor_get(x_1, 1); lean_inc(x_5); +lean_inc(x_3); x_6 = l_Lean_Meta_Grind_congrHash(x_5, x_3); x_7 = lean_uint64_to_usize(x_6); x_8 = 1; @@ -1083,377 +1167,6 @@ x_9 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_addCongrTable___sp return x_9; } } -LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_11; lean_object* x_12; uint8_t x_13; -x_11 = lean_ctor_get(x_8, 2); -x_12 = l_Lean_isTracingEnabledForCore(x_1, x_11, x_10); -x_13 = !lean_is_exclusive(x_12); -if (x_13 == 0) -{ -return x_12; -} -else -{ -lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_14 = lean_ctor_get(x_12, 0); -x_15 = lean_ctor_get(x_12, 1); -lean_inc(x_15); -lean_inc(x_14); -lean_dec(x_12); -x_16 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_16, 0, x_14); -lean_ctor_set(x_16, 1, x_15); -return x_16; -} -} -} -static double _init_l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9___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_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9___closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("", 0, 0); -return x_1; -} -} -static lean_object* _init_l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9___closed__3() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = lean_box(0); -x_2 = lean_array_mk(x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___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) { -_start: -{ -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; -x_12 = lean_ctor_get(x_9, 5); -x_13 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_2, x_7, x_8, x_9, x_10, x_11); -x_14 = lean_ctor_get(x_13, 0); -lean_inc(x_14); -x_15 = lean_ctor_get(x_13, 1); -lean_inc(x_15); -lean_dec(x_13); -x_16 = lean_st_ref_take(x_10, x_15); -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -x_18 = lean_ctor_get(x_17, 3); -lean_inc(x_18); -x_19 = !lean_is_exclusive(x_16); -if (x_19 == 0) -{ -lean_object* x_20; lean_object* x_21; uint8_t x_22; -x_20 = lean_ctor_get(x_16, 1); -x_21 = lean_ctor_get(x_16, 0); -lean_dec(x_21); -x_22 = !lean_is_exclusive(x_17); -if (x_22 == 0) -{ -lean_object* x_23; uint8_t x_24; -x_23 = lean_ctor_get(x_17, 3); -lean_dec(x_23); -x_24 = !lean_is_exclusive(x_18); -if (x_24 == 0) -{ -lean_object* x_25; double 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; uint8_t x_34; -x_25 = lean_ctor_get(x_18, 0); -x_26 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9___closed__1; -x_27 = 0; -x_28 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9___closed__2; -x_29 = lean_alloc_ctor(0, 2, 17); -lean_ctor_set(x_29, 0, x_1); -lean_ctor_set(x_29, 1, x_28); -lean_ctor_set_float(x_29, sizeof(void*)*2, x_26); -lean_ctor_set_float(x_29, sizeof(void*)*2 + 8, x_26); -lean_ctor_set_uint8(x_29, sizeof(void*)*2 + 16, x_27); -x_30 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9___closed__3; -x_31 = lean_alloc_ctor(9, 3, 0); -lean_ctor_set(x_31, 0, x_29); -lean_ctor_set(x_31, 1, x_14); -lean_ctor_set(x_31, 2, x_30); -lean_inc(x_12); -lean_ctor_set(x_16, 1, x_31); -lean_ctor_set(x_16, 0, x_12); -x_32 = l_Lean_PersistentArray_push___rarg(x_25, x_16); -lean_ctor_set(x_18, 0, x_32); -x_33 = lean_st_ref_set(x_10, x_17, x_20); -x_34 = !lean_is_exclusive(x_33); -if (x_34 == 0) -{ -lean_object* x_35; lean_object* x_36; -x_35 = lean_ctor_get(x_33, 0); -lean_dec(x_35); -x_36 = lean_box(0); -lean_ctor_set(x_33, 0, x_36); -return x_33; -} -else -{ -lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_37 = lean_ctor_get(x_33, 1); -lean_inc(x_37); -lean_dec(x_33); -x_38 = lean_box(0); -x_39 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_39, 0, x_38); -lean_ctor_set(x_39, 1, x_37); -return x_39; -} -} -else -{ -uint64_t x_40; lean_object* x_41; double 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; -x_40 = lean_ctor_get_uint64(x_18, sizeof(void*)*1); -x_41 = lean_ctor_get(x_18, 0); -lean_inc(x_41); -lean_dec(x_18); -x_42 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9___closed__1; -x_43 = 0; -x_44 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9___closed__2; -x_45 = lean_alloc_ctor(0, 2, 17); -lean_ctor_set(x_45, 0, x_1); -lean_ctor_set(x_45, 1, x_44); -lean_ctor_set_float(x_45, sizeof(void*)*2, x_42); -lean_ctor_set_float(x_45, sizeof(void*)*2 + 8, x_42); -lean_ctor_set_uint8(x_45, sizeof(void*)*2 + 16, x_43); -x_46 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9___closed__3; -x_47 = lean_alloc_ctor(9, 3, 0); -lean_ctor_set(x_47, 0, x_45); -lean_ctor_set(x_47, 1, x_14); -lean_ctor_set(x_47, 2, x_46); -lean_inc(x_12); -lean_ctor_set(x_16, 1, x_47); -lean_ctor_set(x_16, 0, x_12); -x_48 = l_Lean_PersistentArray_push___rarg(x_41, x_16); -x_49 = lean_alloc_ctor(0, 1, 8); -lean_ctor_set(x_49, 0, x_48); -lean_ctor_set_uint64(x_49, sizeof(void*)*1, x_40); -lean_ctor_set(x_17, 3, x_49); -x_50 = lean_st_ref_set(x_10, x_17, x_20); -x_51 = lean_ctor_get(x_50, 1); -lean_inc(x_51); -if (lean_is_exclusive(x_50)) { - lean_ctor_release(x_50, 0); - lean_ctor_release(x_50, 1); - x_52 = x_50; -} else { - lean_dec_ref(x_50); - x_52 = lean_box(0); -} -x_53 = lean_box(0); -if (lean_is_scalar(x_52)) { - x_54 = lean_alloc_ctor(0, 2, 0); -} else { - x_54 = x_52; -} -lean_ctor_set(x_54, 0, x_53); -lean_ctor_set(x_54, 1, x_51); -return x_54; -} -} -else -{ -lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; uint64_t x_62; lean_object* x_63; lean_object* x_64; double x_65; uint8_t x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; -x_55 = lean_ctor_get(x_17, 0); -x_56 = lean_ctor_get(x_17, 1); -x_57 = lean_ctor_get(x_17, 2); -x_58 = lean_ctor_get(x_17, 4); -x_59 = lean_ctor_get(x_17, 5); -x_60 = lean_ctor_get(x_17, 6); -x_61 = lean_ctor_get(x_17, 7); -lean_inc(x_61); -lean_inc(x_60); -lean_inc(x_59); -lean_inc(x_58); -lean_inc(x_57); -lean_inc(x_56); -lean_inc(x_55); -lean_dec(x_17); -x_62 = lean_ctor_get_uint64(x_18, sizeof(void*)*1); -x_63 = lean_ctor_get(x_18, 0); -lean_inc(x_63); -if (lean_is_exclusive(x_18)) { - lean_ctor_release(x_18, 0); - x_64 = x_18; -} else { - lean_dec_ref(x_18); - x_64 = lean_box(0); -} -x_65 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9___closed__1; -x_66 = 0; -x_67 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9___closed__2; -x_68 = lean_alloc_ctor(0, 2, 17); -lean_ctor_set(x_68, 0, x_1); -lean_ctor_set(x_68, 1, x_67); -lean_ctor_set_float(x_68, sizeof(void*)*2, x_65); -lean_ctor_set_float(x_68, sizeof(void*)*2 + 8, x_65); -lean_ctor_set_uint8(x_68, sizeof(void*)*2 + 16, x_66); -x_69 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9___closed__3; -x_70 = lean_alloc_ctor(9, 3, 0); -lean_ctor_set(x_70, 0, x_68); -lean_ctor_set(x_70, 1, x_14); -lean_ctor_set(x_70, 2, x_69); -lean_inc(x_12); -lean_ctor_set(x_16, 1, x_70); -lean_ctor_set(x_16, 0, x_12); -x_71 = l_Lean_PersistentArray_push___rarg(x_63, x_16); -if (lean_is_scalar(x_64)) { - x_72 = lean_alloc_ctor(0, 1, 8); -} else { - x_72 = x_64; -} -lean_ctor_set(x_72, 0, x_71); -lean_ctor_set_uint64(x_72, sizeof(void*)*1, x_62); -x_73 = lean_alloc_ctor(0, 8, 0); -lean_ctor_set(x_73, 0, x_55); -lean_ctor_set(x_73, 1, x_56); -lean_ctor_set(x_73, 2, x_57); -lean_ctor_set(x_73, 3, x_72); -lean_ctor_set(x_73, 4, x_58); -lean_ctor_set(x_73, 5, x_59); -lean_ctor_set(x_73, 6, x_60); -lean_ctor_set(x_73, 7, x_61); -x_74 = lean_st_ref_set(x_10, x_73, x_20); -x_75 = lean_ctor_get(x_74, 1); -lean_inc(x_75); -if (lean_is_exclusive(x_74)) { - lean_ctor_release(x_74, 0); - lean_ctor_release(x_74, 1); - x_76 = x_74; -} else { - lean_dec_ref(x_74); - x_76 = lean_box(0); -} -x_77 = lean_box(0); -if (lean_is_scalar(x_76)) { - x_78 = lean_alloc_ctor(0, 2, 0); -} else { - x_78 = x_76; -} -lean_ctor_set(x_78, 0, x_77); -lean_ctor_set(x_78, 1, x_75); -return x_78; -} -} -else -{ -lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; uint64_t x_88; lean_object* x_89; lean_object* x_90; double x_91; uint8_t x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; -x_79 = lean_ctor_get(x_16, 1); -lean_inc(x_79); -lean_dec(x_16); -x_80 = lean_ctor_get(x_17, 0); -lean_inc(x_80); -x_81 = lean_ctor_get(x_17, 1); -lean_inc(x_81); -x_82 = lean_ctor_get(x_17, 2); -lean_inc(x_82); -x_83 = lean_ctor_get(x_17, 4); -lean_inc(x_83); -x_84 = lean_ctor_get(x_17, 5); -lean_inc(x_84); -x_85 = lean_ctor_get(x_17, 6); -lean_inc(x_85); -x_86 = lean_ctor_get(x_17, 7); -lean_inc(x_86); -if (lean_is_exclusive(x_17)) { - lean_ctor_release(x_17, 0); - lean_ctor_release(x_17, 1); - lean_ctor_release(x_17, 2); - lean_ctor_release(x_17, 3); - lean_ctor_release(x_17, 4); - lean_ctor_release(x_17, 5); - lean_ctor_release(x_17, 6); - lean_ctor_release(x_17, 7); - x_87 = x_17; -} else { - lean_dec_ref(x_17); - x_87 = lean_box(0); -} -x_88 = lean_ctor_get_uint64(x_18, sizeof(void*)*1); -x_89 = lean_ctor_get(x_18, 0); -lean_inc(x_89); -if (lean_is_exclusive(x_18)) { - lean_ctor_release(x_18, 0); - x_90 = x_18; -} else { - lean_dec_ref(x_18); - x_90 = lean_box(0); -} -x_91 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9___closed__1; -x_92 = 0; -x_93 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9___closed__2; -x_94 = lean_alloc_ctor(0, 2, 17); -lean_ctor_set(x_94, 0, x_1); -lean_ctor_set(x_94, 1, x_93); -lean_ctor_set_float(x_94, sizeof(void*)*2, x_91); -lean_ctor_set_float(x_94, sizeof(void*)*2 + 8, x_91); -lean_ctor_set_uint8(x_94, sizeof(void*)*2 + 16, x_92); -x_95 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9___closed__3; -x_96 = lean_alloc_ctor(9, 3, 0); -lean_ctor_set(x_96, 0, x_94); -lean_ctor_set(x_96, 1, x_14); -lean_ctor_set(x_96, 2, x_95); -lean_inc(x_12); -x_97 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_97, 0, x_12); -lean_ctor_set(x_97, 1, x_96); -x_98 = l_Lean_PersistentArray_push___rarg(x_89, x_97); -if (lean_is_scalar(x_90)) { - x_99 = lean_alloc_ctor(0, 1, 8); -} else { - x_99 = x_90; -} -lean_ctor_set(x_99, 0, x_98); -lean_ctor_set_uint64(x_99, sizeof(void*)*1, x_88); -if (lean_is_scalar(x_87)) { - x_100 = lean_alloc_ctor(0, 8, 0); -} else { - x_100 = x_87; -} -lean_ctor_set(x_100, 0, x_80); -lean_ctor_set(x_100, 1, x_81); -lean_ctor_set(x_100, 2, x_82); -lean_ctor_set(x_100, 3, x_99); -lean_ctor_set(x_100, 4, x_83); -lean_ctor_set(x_100, 5, x_84); -lean_ctor_set(x_100, 6, x_85); -lean_ctor_set(x_100, 7, x_86); -x_101 = lean_st_ref_set(x_10, x_100, x_79); -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; -} else { - lean_dec_ref(x_101); - x_103 = lean_box(0); -} -x_104 = lean_box(0); -if (lean_is_scalar(x_103)) { - x_105 = lean_alloc_ctor(0, 2, 0); -} else { - x_105 = x_103; -} -lean_ctor_set(x_105, 0, x_104); -lean_ctor_set(x_105, 1, x_102); -return x_105; -} -} -} LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addCongrTable___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { @@ -1682,13 +1395,21 @@ return x_4; static lean_object* _init_l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__5() { _start: { +lean_object* x_1; +x_1 = lean_mk_string_unchecked("", 0, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__6() { +_start: +{ lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9___closed__2; +x_1 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__5; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__6() { +static lean_object* _init_l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__7() { _start: { lean_object* x_1; @@ -1696,11 +1417,11 @@ x_1 = lean_mk_string_unchecked(" = ", 3, 3); return x_1; } } -static lean_object* _init_l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__7() { +static lean_object* _init_l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__8() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__6; +x_1 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__7; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } @@ -1710,7 +1431,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addCongrTable___lambda__2(lean_object { lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; x_13 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__4; -x_14 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +x_14 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); x_15 = lean_ctor_get(x_14, 0); lean_inc(x_15); x_16 = lean_unbox(x_15); @@ -1731,95 +1452,170 @@ uint8_t x_20; x_20 = !lean_is_exclusive(x_14); if (x_20 == 0) { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; +lean_object* x_21; lean_object* x_22; lean_object* x_23; x_21 = lean_ctor_get(x_14, 1); x_22 = lean_ctor_get(x_14, 0); lean_dec(x_22); +x_23 = l_Lean_Meta_Grind_updateLastTag(x_4, x_5, 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; lean_object* x_27; lean_object* x_28; lean_object* x_29; 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_24 = lean_ctor_get(x_23, 1); +lean_inc(x_24); +lean_dec(x_23); lean_inc(x_1); -x_23 = l_Lean_MessageData_ofExpr(x_1); -x_24 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__5; +x_25 = l_Lean_MessageData_ofExpr(x_1); +x_26 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__6; lean_ctor_set_tag(x_14, 7); -lean_ctor_set(x_14, 1, x_23); -lean_ctor_set(x_14, 0, x_24); -x_25 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__7; -x_26 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_26, 0, x_14); -lean_ctor_set(x_26, 1, x_25); -lean_inc(x_2); -x_27 = l_Lean_MessageData_ofExpr(x_2); +lean_ctor_set(x_14, 1, x_25); +lean_ctor_set(x_14, 0, x_26); +x_27 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__8; x_28 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_28, 0, x_26); +lean_ctor_set(x_28, 0, x_14); lean_ctor_set(x_28, 1, x_27); -x_29 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_29, 0, x_28); -lean_ctor_set(x_29, 1, x_24); -x_30 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_13, x_29, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_21); -x_31 = lean_ctor_get(x_30, 0); -lean_inc(x_31); -x_32 = lean_ctor_get(x_30, 1); -lean_inc(x_32); -lean_dec(x_30); -x_33 = l_Lean_Meta_Grind_addCongrTable___lambda__1(x_1, x_2, x_31, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_32); -lean_dec(x_31); -return x_33; +lean_inc(x_2); +x_29 = l_Lean_MessageData_ofExpr(x_2); +x_30 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_30, 0, x_28); +lean_ctor_set(x_30, 1, x_29); +x_31 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_31, 0, x_30); +lean_ctor_set(x_31, 1, x_26); +x_32 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_13, x_31, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_24); +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_Grind_addCongrTable___lambda__1(x_1, x_2, x_33, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_34); +lean_dec(x_33); +return x_35; } else { -lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; 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_14, 1); -lean_inc(x_34); -lean_dec(x_14); -lean_inc(x_1); -x_35 = l_Lean_MessageData_ofExpr(x_1); -x_36 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__5; -x_37 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_37, 0, x_36); -lean_ctor_set(x_37, 1, x_35); -x_38 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__7; -x_39 = lean_alloc_ctor(7, 2, 0); +uint8_t x_36; +lean_free_object(x_14); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_2); +lean_dec(x_1); +x_36 = !lean_is_exclusive(x_23); +if (x_36 == 0) +{ +return x_23; +} +else +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_37 = lean_ctor_get(x_23, 0); +x_38 = lean_ctor_get(x_23, 1); +lean_inc(x_38); +lean_inc(x_37); +lean_dec(x_23); +x_39 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_39, 0, x_37); lean_ctor_set(x_39, 1, x_38); -lean_inc(x_2); -x_40 = l_Lean_MessageData_ofExpr(x_2); -x_41 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_41, 0, x_39); -lean_ctor_set(x_41, 1, x_40); -x_42 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_42, 0, x_41); -lean_ctor_set(x_42, 1, x_36); -x_43 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_13, x_42, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_34); -x_44 = lean_ctor_get(x_43, 0); -lean_inc(x_44); -x_45 = lean_ctor_get(x_43, 1); -lean_inc(x_45); -lean_dec(x_43); -x_46 = l_Lean_Meta_Grind_addCongrTable___lambda__1(x_1, x_2, x_44, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_45); -lean_dec(x_44); -return x_46; -} +return x_39; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addCongrTable___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; -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; -} -} -static lean_object* _init_l_Lean_Meta_Grind_addCongrTable___closed__1() { -_start: +lean_object* x_40; lean_object* x_41; +x_40 = lean_ctor_get(x_14, 1); +lean_inc(x_40); +lean_dec(x_14); +x_41 = l_Lean_Meta_Grind_updateLastTag(x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_40); +if (lean_obj_tag(x_41) == 0) { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("issues", 6, 6); -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; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; +x_42 = lean_ctor_get(x_41, 1); +lean_inc(x_42); +lean_dec(x_41); +lean_inc(x_1); +x_43 = l_Lean_MessageData_ofExpr(x_1); +x_44 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__6; +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_Lean_Meta_Grind_addCongrTable___lambda__2___closed__8; +x_47 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_47, 0, x_45); +lean_ctor_set(x_47, 1, x_46); +lean_inc(x_2); +x_48 = l_Lean_MessageData_ofExpr(x_2); +x_49 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_49, 0, x_47); +lean_ctor_set(x_49, 1, x_48); +x_50 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_50, 0, x_49); +lean_ctor_set(x_50, 1, x_44); +x_51 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_13, x_50, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_42); +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_Grind_addCongrTable___lambda__1(x_1, x_2, x_52, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_53); +lean_dec(x_52); +return x_54; } -static lean_object* _init_l_Lean_Meta_Grind_addCongrTable___closed__2() { +else +{ +lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_2); +lean_dec(x_1); +x_55 = lean_ctor_get(x_41, 0); +lean_inc(x_55); +x_56 = lean_ctor_get(x_41, 1); +lean_inc(x_56); +if (lean_is_exclusive(x_41)) { + lean_ctor_release(x_41, 0); + lean_ctor_release(x_41, 1); + x_57 = x_41; +} else { + lean_dec_ref(x_41); + x_57 = lean_box(0); +} +if (lean_is_scalar(x_57)) { + x_58 = lean_alloc_ctor(1, 2, 0); +} else { + x_58 = x_57; +} +lean_ctor_set(x_58, 0, x_55); +lean_ctor_set(x_58, 1, x_56); +return x_58; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addCongrTable___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; lean_object* x_12; +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; +} +} +static lean_object* _init_l_Lean_Meta_Grind_addCongrTable___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("issues", 6, 6); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_addCongrTable___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -1905,7 +1701,7 @@ lean_inc(x_1); x_16 = l_Lean_PersistentHashMap_findEntry_x3f___at_Lean_Meta_Grind_addCongrTable___spec__1(x_13, x_15, x_1); if (lean_obj_tag(x_16) == 0) { -lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; +lean_object* x_17; lean_object* x_18; lean_object* x_19; 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; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; uint8_t x_43; lean_free_object(x_11); lean_dec(x_9); lean_dec(x_8); @@ -1932,7 +1728,7 @@ x_24 = lean_ctor_get(x_18, 4); lean_inc(x_24); x_25 = lean_ctor_get(x_18, 5); lean_inc(x_25); -x_26 = lean_ctor_get_uint8(x_18, sizeof(void*)*14); +x_26 = lean_ctor_get_uint8(x_18, sizeof(void*)*20); x_27 = lean_ctor_get(x_18, 6); lean_inc(x_27); x_28 = lean_ctor_get(x_18, 7); @@ -1949,272 +1745,389 @@ x_33 = lean_ctor_get(x_18, 12); lean_inc(x_33); x_34 = lean_ctor_get(x_18, 13); lean_inc(x_34); -x_35 = lean_box(0); +x_35 = lean_ctor_get(x_18, 14); +lean_inc(x_35); +x_36 = lean_ctor_get(x_18, 15); +lean_inc(x_36); +x_37 = lean_ctor_get(x_18, 16); +lean_inc(x_37); +x_38 = lean_ctor_get(x_18, 17); +lean_inc(x_38); +x_39 = lean_ctor_get(x_18, 18); +lean_inc(x_39); +x_40 = lean_ctor_get(x_18, 19); +lean_inc(x_40); +x_41 = lean_box(0); lean_inc(x_18); -x_36 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_addCongrTable___spec__4(x_18, x_23, x_1, x_35); -x_37 = !lean_is_exclusive(x_18); -if (x_37 == 0) +x_42 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_addCongrTable___spec__4(x_18, x_23, x_1, x_41); +x_43 = !lean_is_exclusive(x_18); +if (x_43 == 0) { -lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; uint8_t x_53; -x_38 = lean_ctor_get(x_18, 13); -lean_dec(x_38); -x_39 = lean_ctor_get(x_18, 12); -lean_dec(x_39); -x_40 = lean_ctor_get(x_18, 11); -lean_dec(x_40); -x_41 = lean_ctor_get(x_18, 10); -lean_dec(x_41); -x_42 = lean_ctor_get(x_18, 9); -lean_dec(x_42); -x_43 = lean_ctor_get(x_18, 8); -lean_dec(x_43); -x_44 = lean_ctor_get(x_18, 7); +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; uint8_t x_65; +x_44 = lean_ctor_get(x_18, 19); lean_dec(x_44); -x_45 = lean_ctor_get(x_18, 6); +x_45 = lean_ctor_get(x_18, 18); lean_dec(x_45); -x_46 = lean_ctor_get(x_18, 5); +x_46 = lean_ctor_get(x_18, 17); lean_dec(x_46); -x_47 = lean_ctor_get(x_18, 4); +x_47 = lean_ctor_get(x_18, 16); lean_dec(x_47); -x_48 = lean_ctor_get(x_18, 3); +x_48 = lean_ctor_get(x_18, 15); lean_dec(x_48); -x_49 = lean_ctor_get(x_18, 2); +x_49 = lean_ctor_get(x_18, 14); lean_dec(x_49); -x_50 = lean_ctor_get(x_18, 1); +x_50 = lean_ctor_get(x_18, 13); lean_dec(x_50); -x_51 = lean_ctor_get(x_18, 0); +x_51 = lean_ctor_get(x_18, 12); lean_dec(x_51); -lean_ctor_set(x_18, 3, x_36); -x_52 = lean_st_ref_set(x_2, x_18, x_19); +x_52 = lean_ctor_get(x_18, 11); +lean_dec(x_52); +x_53 = lean_ctor_get(x_18, 10); +lean_dec(x_53); +x_54 = lean_ctor_get(x_18, 9); +lean_dec(x_54); +x_55 = lean_ctor_get(x_18, 8); +lean_dec(x_55); +x_56 = lean_ctor_get(x_18, 7); +lean_dec(x_56); +x_57 = lean_ctor_get(x_18, 6); +lean_dec(x_57); +x_58 = lean_ctor_get(x_18, 5); +lean_dec(x_58); +x_59 = lean_ctor_get(x_18, 4); +lean_dec(x_59); +x_60 = lean_ctor_get(x_18, 3); +lean_dec(x_60); +x_61 = lean_ctor_get(x_18, 2); +lean_dec(x_61); +x_62 = lean_ctor_get(x_18, 1); +lean_dec(x_62); +x_63 = lean_ctor_get(x_18, 0); +lean_dec(x_63); +lean_ctor_set(x_18, 3, x_42); +x_64 = lean_st_ref_set(x_2, x_18, x_19); lean_dec(x_2); -x_53 = !lean_is_exclusive(x_52); -if (x_53 == 0) +x_65 = !lean_is_exclusive(x_64); +if (x_65 == 0) { -lean_object* x_54; -x_54 = lean_ctor_get(x_52, 0); -lean_dec(x_54); -lean_ctor_set(x_52, 0, x_35); -return x_52; +lean_object* x_66; +x_66 = lean_ctor_get(x_64, 0); +lean_dec(x_66); +lean_ctor_set(x_64, 0, x_41); +return x_64; } 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_35); -lean_ctor_set(x_56, 1, x_55); -return x_56; +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_alloc_ctor(0, 2, 0); +lean_ctor_set(x_68, 0, x_41); +lean_ctor_set(x_68, 1, x_67); +return x_68; } } else { -lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; +lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_dec(x_18); -x_57 = lean_alloc_ctor(0, 14, 1); -lean_ctor_set(x_57, 0, x_20); -lean_ctor_set(x_57, 1, x_21); -lean_ctor_set(x_57, 2, x_22); -lean_ctor_set(x_57, 3, x_36); -lean_ctor_set(x_57, 4, x_24); -lean_ctor_set(x_57, 5, x_25); -lean_ctor_set(x_57, 6, x_27); -lean_ctor_set(x_57, 7, x_28); -lean_ctor_set(x_57, 8, x_29); -lean_ctor_set(x_57, 9, x_30); -lean_ctor_set(x_57, 10, x_31); -lean_ctor_set(x_57, 11, x_32); -lean_ctor_set(x_57, 12, x_33); -lean_ctor_set(x_57, 13, x_34); -lean_ctor_set_uint8(x_57, sizeof(void*)*14, x_26); -x_58 = lean_st_ref_set(x_2, x_57, x_19); +x_69 = lean_alloc_ctor(0, 20, 1); +lean_ctor_set(x_69, 0, x_20); +lean_ctor_set(x_69, 1, x_21); +lean_ctor_set(x_69, 2, x_22); +lean_ctor_set(x_69, 3, x_42); +lean_ctor_set(x_69, 4, x_24); +lean_ctor_set(x_69, 5, x_25); +lean_ctor_set(x_69, 6, x_27); +lean_ctor_set(x_69, 7, x_28); +lean_ctor_set(x_69, 8, x_29); +lean_ctor_set(x_69, 9, x_30); +lean_ctor_set(x_69, 10, x_31); +lean_ctor_set(x_69, 11, x_32); +lean_ctor_set(x_69, 12, x_33); +lean_ctor_set(x_69, 13, x_34); +lean_ctor_set(x_69, 14, x_35); +lean_ctor_set(x_69, 15, x_36); +lean_ctor_set(x_69, 16, x_37); +lean_ctor_set(x_69, 17, x_38); +lean_ctor_set(x_69, 18, x_39); +lean_ctor_set(x_69, 19, x_40); +lean_ctor_set_uint8(x_69, sizeof(void*)*20, x_26); +x_70 = lean_st_ref_set(x_2, x_69, x_19); lean_dec(x_2); -x_59 = lean_ctor_get(x_58, 1); -lean_inc(x_59); -if (lean_is_exclusive(x_58)) { - lean_ctor_release(x_58, 0); - lean_ctor_release(x_58, 1); - x_60 = x_58; +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_58); - x_60 = lean_box(0); + lean_dec_ref(x_70); + x_72 = lean_box(0); } -if (lean_is_scalar(x_60)) { - x_61 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_72)) { + x_73 = lean_alloc_ctor(0, 2, 0); } else { - x_61 = x_60; + x_73 = x_72; } -lean_ctor_set(x_61, 0, x_35); -lean_ctor_set(x_61, 1, x_59); -return x_61; +lean_ctor_set(x_73, 0, x_41); +lean_ctor_set(x_73, 1, x_71); +return x_73; } } else { -lean_object* x_62; uint8_t x_63; -x_62 = lean_ctor_get(x_16, 0); -lean_inc(x_62); +lean_object* x_74; uint8_t x_75; +x_74 = lean_ctor_get(x_16, 0); +lean_inc(x_74); lean_dec(x_16); -x_63 = !lean_is_exclusive(x_62); -if (x_63 == 0) +x_75 = !lean_is_exclusive(x_74); +if (x_75 == 0) { -lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; uint8_t x_68; -x_64 = lean_ctor_get(x_62, 0); -x_65 = lean_ctor_get(x_62, 1); -lean_dec(x_65); -x_66 = l_Lean_Expr_getAppFn(x_1); -x_67 = l_Lean_Expr_getAppFn(x_64); -x_68 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_66, x_67); -if (x_68 == 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_74, 0); +x_77 = lean_ctor_get(x_74, 1); +lean_dec(x_77); +x_78 = l_Lean_Expr_getAppFn(x_1); +x_79 = l_Lean_Expr_getAppFn(x_76); +x_80 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_78, x_79); +if (x_80 == 0) { -lean_object* x_69; +lean_object* x_81; lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -x_69 = l_Lean_Meta_Grind_hasSameType(x_66, x_67, x_6, x_7, x_8, x_9, x_14); -if (lean_obj_tag(x_69) == 0) +x_81 = l_Lean_Meta_Grind_hasSameType(x_78, x_79, x_6, x_7, x_8, x_9, x_14); +if (lean_obj_tag(x_81) == 0) { -lean_object* x_70; uint8_t x_71; -x_70 = lean_ctor_get(x_69, 0); -lean_inc(x_70); -x_71 = lean_unbox(x_70); -lean_dec(x_70); -if (x_71 == 0) +lean_object* x_82; uint8_t x_83; +x_82 = lean_ctor_get(x_81, 0); +lean_inc(x_82); +x_83 = lean_unbox(x_82); +lean_dec(x_82); +if (x_83 == 0) { -lean_object* x_72; lean_object* x_73; lean_object* x_74; uint8_t x_75; -x_72 = lean_ctor_get(x_69, 1); -lean_inc(x_72); -lean_dec(x_69); -x_73 = l_Lean_Meta_Grind_addCongrTable___closed__2; -x_74 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_73, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_72); -x_75 = !lean_is_exclusive(x_74); -if (x_75 == 0) +lean_object* x_84; lean_object* x_85; lean_object* x_86; uint8_t x_87; +x_84 = lean_ctor_get(x_81, 1); +lean_inc(x_84); +lean_dec(x_81); +x_85 = l_Lean_Meta_Grind_addCongrTable___closed__2; +x_86 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_85, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_84); +x_87 = !lean_is_exclusive(x_86); +if (x_87 == 0) { -lean_object* x_76; lean_object* x_77; lean_object* x_78; uint8_t x_79; -x_76 = lean_ctor_get(x_74, 0); -x_77 = lean_ctor_get(x_74, 1); -x_78 = l_Lean_Meta_Grind_addCongrTable___closed__3; -x_79 = lean_unbox(x_76); -lean_dec(x_76); -if (x_79 == 0) +lean_object* x_88; lean_object* x_89; lean_object* x_90; uint8_t x_91; +x_88 = lean_ctor_get(x_86, 0); +x_89 = lean_ctor_get(x_86, 1); +x_90 = l_Lean_Meta_Grind_addCongrTable___closed__3; +x_91 = lean_unbox(x_88); +lean_dec(x_88); +if (x_91 == 0) { -lean_object* x_80; lean_object* x_81; +lean_object* x_92; lean_object* x_93; +lean_free_object(x_86); lean_free_object(x_74); -lean_free_object(x_62); -lean_dec(x_64); +lean_dec(x_76); lean_free_object(x_11); lean_dec(x_1); -x_80 = lean_box(0); -x_81 = lean_apply_10(x_78, x_80, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_77); -return x_81; +x_92 = lean_box(0); +x_93 = lean_apply_10(x_90, x_92, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_89); +return x_93; } 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; -x_82 = l_Lean_indentExpr(x_1); -x_83 = l_Lean_Meta_Grind_addCongrTable___closed__5; +lean_object* x_94; +x_94 = l_Lean_Meta_Grind_updateLastTag(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_89); +if (lean_obj_tag(x_94) == 0) +{ +lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; +x_95 = lean_ctor_get(x_94, 1); +lean_inc(x_95); +lean_dec(x_94); +x_96 = l_Lean_indentExpr(x_1); +x_97 = l_Lean_Meta_Grind_addCongrTable___closed__5; +lean_ctor_set_tag(x_86, 7); +lean_ctor_set(x_86, 1, x_96); +lean_ctor_set(x_86, 0, x_97); +x_98 = l_Lean_Meta_Grind_addCongrTable___closed__7; lean_ctor_set_tag(x_74, 7); -lean_ctor_set(x_74, 1, x_82); -lean_ctor_set(x_74, 0, x_83); -x_84 = l_Lean_Meta_Grind_addCongrTable___closed__7; -lean_ctor_set_tag(x_62, 7); -lean_ctor_set(x_62, 1, x_84); -lean_ctor_set(x_62, 0, x_74); -x_85 = l_Lean_indentExpr(x_64); +lean_ctor_set(x_74, 1, x_98); +lean_ctor_set(x_74, 0, x_86); +x_99 = l_Lean_indentExpr(x_76); lean_ctor_set_tag(x_11, 7); -lean_ctor_set(x_11, 1, x_85); -lean_ctor_set(x_11, 0, x_62); -x_86 = l_Lean_Meta_Grind_addCongrTable___closed__9; -x_87 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_87, 0, x_11); -lean_ctor_set(x_87, 1, x_86); -x_88 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_73, x_87, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, 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); -lean_dec(x_88); -x_91 = lean_apply_10(x_78, x_89, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_90); -return x_91; +lean_ctor_set(x_11, 1, x_99); +lean_ctor_set(x_11, 0, x_74); +x_100 = l_Lean_Meta_Grind_addCongrTable___closed__9; +x_101 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_101, 0, x_11); +lean_ctor_set(x_101, 1, x_100); +x_102 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_85, x_101, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_95); +x_103 = lean_ctor_get(x_102, 0); +lean_inc(x_103); +x_104 = lean_ctor_get(x_102, 1); +lean_inc(x_104); +lean_dec(x_102); +x_105 = lean_apply_10(x_90, x_103, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_104); +return x_105; +} +else +{ +uint8_t x_106; +lean_free_object(x_86); +lean_free_object(x_74); +lean_dec(x_76); +lean_free_object(x_11); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_106 = !lean_is_exclusive(x_94); +if (x_106 == 0) +{ +return x_94; +} +else +{ +lean_object* x_107; lean_object* x_108; lean_object* x_109; +x_107 = lean_ctor_get(x_94, 0); +x_108 = lean_ctor_get(x_94, 1); +lean_inc(x_108); +lean_inc(x_107); +lean_dec(x_94); +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 { -lean_object* x_92; lean_object* x_93; lean_object* x_94; uint8_t x_95; -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_dec(x_74); -x_94 = l_Lean_Meta_Grind_addCongrTable___closed__3; -x_95 = lean_unbox(x_92); -lean_dec(x_92); -if (x_95 == 0) +lean_object* x_110; lean_object* x_111; lean_object* x_112; uint8_t x_113; +x_110 = lean_ctor_get(x_86, 0); +x_111 = lean_ctor_get(x_86, 1); +lean_inc(x_111); +lean_inc(x_110); +lean_dec(x_86); +x_112 = l_Lean_Meta_Grind_addCongrTable___closed__3; +x_113 = lean_unbox(x_110); +lean_dec(x_110); +if (x_113 == 0) { -lean_object* x_96; lean_object* x_97; -lean_free_object(x_62); -lean_dec(x_64); +lean_object* x_114; lean_object* x_115; +lean_free_object(x_74); +lean_dec(x_76); lean_free_object(x_11); lean_dec(x_1); -x_96 = lean_box(0); -x_97 = lean_apply_10(x_94, x_96, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_93); -return x_97; +x_114 = lean_box(0); +x_115 = lean_apply_10(x_112, x_114, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_111); +return x_115; } else { -lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; -x_98 = l_Lean_indentExpr(x_1); -x_99 = l_Lean_Meta_Grind_addCongrTable___closed__5; -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 = l_Lean_Meta_Grind_addCongrTable___closed__7; -lean_ctor_set_tag(x_62, 7); -lean_ctor_set(x_62, 1, x_101); -lean_ctor_set(x_62, 0, x_100); -x_102 = l_Lean_indentExpr(x_64); +lean_object* x_116; +x_116 = l_Lean_Meta_Grind_updateLastTag(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, 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; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; +x_117 = lean_ctor_get(x_116, 1); +lean_inc(x_117); +lean_dec(x_116); +x_118 = l_Lean_indentExpr(x_1); +x_119 = l_Lean_Meta_Grind_addCongrTable___closed__5; +x_120 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_120, 0, x_119); +lean_ctor_set(x_120, 1, x_118); +x_121 = l_Lean_Meta_Grind_addCongrTable___closed__7; +lean_ctor_set_tag(x_74, 7); +lean_ctor_set(x_74, 1, x_121); +lean_ctor_set(x_74, 0, x_120); +x_122 = l_Lean_indentExpr(x_76); lean_ctor_set_tag(x_11, 7); -lean_ctor_set(x_11, 1, x_102); -lean_ctor_set(x_11, 0, x_62); -x_103 = l_Lean_Meta_Grind_addCongrTable___closed__9; -x_104 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_104, 0, x_11); -lean_ctor_set(x_104, 1, x_103); -x_105 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_73, x_104, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_93); -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_apply_10(x_94, x_106, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_107); -return x_108; +lean_ctor_set(x_11, 1, x_122); +lean_ctor_set(x_11, 0, x_74); +x_123 = l_Lean_Meta_Grind_addCongrTable___closed__9; +x_124 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_124, 0, x_11); +lean_ctor_set(x_124, 1, x_123); +x_125 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_85, x_124, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_117); +x_126 = lean_ctor_get(x_125, 0); +lean_inc(x_126); +x_127 = lean_ctor_get(x_125, 1); +lean_inc(x_127); +lean_dec(x_125); +x_128 = lean_apply_10(x_112, x_126, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_127); +return x_128; +} +else +{ +lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; +lean_free_object(x_74); +lean_dec(x_76); +lean_free_object(x_11); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_129 = lean_ctor_get(x_116, 0); +lean_inc(x_129); +x_130 = lean_ctor_get(x_116, 1); +lean_inc(x_130); +if (lean_is_exclusive(x_116)) { + lean_ctor_release(x_116, 0); + lean_ctor_release(x_116, 1); + x_131 = x_116; +} else { + lean_dec_ref(x_116); + x_131 = lean_box(0); +} +if (lean_is_scalar(x_131)) { + x_132 = lean_alloc_ctor(1, 2, 0); +} else { + x_132 = x_131; +} +lean_ctor_set(x_132, 0, x_129); +lean_ctor_set(x_132, 1, x_130); +return x_132; +} } } } else { -lean_object* x_109; lean_object* x_110; lean_object* x_111; -lean_free_object(x_62); +lean_object* x_133; lean_object* x_134; lean_object* x_135; +lean_free_object(x_74); lean_free_object(x_11); -x_109 = lean_ctor_get(x_69, 1); -lean_inc(x_109); -lean_dec(x_69); -x_110 = lean_box(0); -x_111 = l_Lean_Meta_Grind_addCongrTable___lambda__2(x_1, x_64, x_110, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_109); +x_133 = lean_ctor_get(x_81, 1); +lean_inc(x_133); +lean_dec(x_81); +x_134 = lean_box(0); +x_135 = l_Lean_Meta_Grind_addCongrTable___lambda__2(x_1, x_76, x_134, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_133); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -return x_111; +return x_135; } } else { -uint8_t x_112; -lean_free_object(x_62); -lean_dec(x_64); +uint8_t x_136; +lean_free_object(x_74); +lean_dec(x_76); lean_free_object(x_11); lean_dec(x_9); lean_dec(x_8); @@ -2225,155 +2138,199 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_112 = !lean_is_exclusive(x_69); -if (x_112 == 0) +x_136 = !lean_is_exclusive(x_81); +if (x_136 == 0) { -return x_69; +return x_81; } else { -lean_object* x_113; lean_object* x_114; lean_object* x_115; -x_113 = lean_ctor_get(x_69, 0); -x_114 = lean_ctor_get(x_69, 1); -lean_inc(x_114); -lean_inc(x_113); -lean_dec(x_69); -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_137; lean_object* x_138; lean_object* x_139; +x_137 = lean_ctor_get(x_81, 0); +x_138 = lean_ctor_get(x_81, 1); +lean_inc(x_138); +lean_inc(x_137); +lean_dec(x_81); +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_116; lean_object* x_117; -lean_dec(x_67); -lean_dec(x_66); -lean_free_object(x_62); +lean_object* x_140; lean_object* x_141; +lean_dec(x_79); +lean_dec(x_78); +lean_free_object(x_74); lean_free_object(x_11); -x_116 = lean_box(0); -x_117 = l_Lean_Meta_Grind_addCongrTable___lambda__2(x_1, x_64, x_116, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_14); +x_140 = lean_box(0); +x_141 = l_Lean_Meta_Grind_addCongrTable___lambda__2(x_1, x_76, x_140, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_14); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -return x_117; +return x_141; } } else { -lean_object* x_118; lean_object* x_119; lean_object* x_120; uint8_t x_121; -x_118 = lean_ctor_get(x_62, 0); -lean_inc(x_118); -lean_dec(x_62); -x_119 = l_Lean_Expr_getAppFn(x_1); -x_120 = l_Lean_Expr_getAppFn(x_118); -x_121 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_119, x_120); -if (x_121 == 0) +lean_object* x_142; lean_object* x_143; lean_object* x_144; uint8_t x_145; +x_142 = lean_ctor_get(x_74, 0); +lean_inc(x_142); +lean_dec(x_74); +x_143 = l_Lean_Expr_getAppFn(x_1); +x_144 = l_Lean_Expr_getAppFn(x_142); +x_145 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_143, x_144); +if (x_145 == 0) { -lean_object* x_122; +lean_object* x_146; lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -x_122 = l_Lean_Meta_Grind_hasSameType(x_119, x_120, x_6, x_7, x_8, x_9, x_14); -if (lean_obj_tag(x_122) == 0) +x_146 = l_Lean_Meta_Grind_hasSameType(x_143, x_144, x_6, x_7, x_8, x_9, x_14); +if (lean_obj_tag(x_146) == 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); -lean_dec(x_123); -if (x_124 == 0) +lean_object* x_147; uint8_t x_148; +x_147 = lean_ctor_get(x_146, 0); +lean_inc(x_147); +x_148 = lean_unbox(x_147); +lean_dec(x_147); +if (x_148 == 0) { -lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; uint8_t x_132; -x_125 = lean_ctor_get(x_122, 1); -lean_inc(x_125); -lean_dec(x_122); -x_126 = l_Lean_Meta_Grind_addCongrTable___closed__2; -x_127 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_126, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_125); -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_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; +x_149 = lean_ctor_get(x_146, 1); +lean_inc(x_149); +lean_dec(x_146); +x_150 = l_Lean_Meta_Grind_addCongrTable___closed__2; +x_151 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_150, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_149); +x_152 = lean_ctor_get(x_151, 0); +lean_inc(x_152); +x_153 = lean_ctor_get(x_151, 1); +lean_inc(x_153); +if (lean_is_exclusive(x_151)) { + lean_ctor_release(x_151, 0); + lean_ctor_release(x_151, 1); + x_154 = x_151; } else { - lean_dec_ref(x_127); - x_130 = lean_box(0); + lean_dec_ref(x_151); + x_154 = lean_box(0); } -x_131 = l_Lean_Meta_Grind_addCongrTable___closed__3; -x_132 = lean_unbox(x_128); -lean_dec(x_128); -if (x_132 == 0) +x_155 = l_Lean_Meta_Grind_addCongrTable___closed__3; +x_156 = lean_unbox(x_152); +lean_dec(x_152); +if (x_156 == 0) { -lean_object* x_133; lean_object* x_134; -lean_dec(x_130); -lean_dec(x_118); +lean_object* x_157; lean_object* x_158; +lean_dec(x_154); +lean_dec(x_142); lean_free_object(x_11); lean_dec(x_1); -x_133 = lean_box(0); -x_134 = lean_apply_10(x_131, x_133, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_129); -return x_134; +x_157 = lean_box(0); +x_158 = lean_apply_10(x_155, x_157, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_153); +return x_158; } else { -lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; -x_135 = l_Lean_indentExpr(x_1); -x_136 = l_Lean_Meta_Grind_addCongrTable___closed__5; -if (lean_is_scalar(x_130)) { - x_137 = lean_alloc_ctor(7, 2, 0); +lean_object* x_159; +x_159 = l_Lean_Meta_Grind_updateLastTag(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_153); +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; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; +x_160 = lean_ctor_get(x_159, 1); +lean_inc(x_160); +lean_dec(x_159); +x_161 = l_Lean_indentExpr(x_1); +x_162 = l_Lean_Meta_Grind_addCongrTable___closed__5; +if (lean_is_scalar(x_154)) { + x_163 = lean_alloc_ctor(7, 2, 0); } else { - x_137 = x_130; - lean_ctor_set_tag(x_137, 7); -} -lean_ctor_set(x_137, 0, x_136); -lean_ctor_set(x_137, 1, x_135); -x_138 = l_Lean_Meta_Grind_addCongrTable___closed__7; -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_indentExpr(x_118); + x_163 = x_154; + lean_ctor_set_tag(x_163, 7); +} +lean_ctor_set(x_163, 0, x_162); +lean_ctor_set(x_163, 1, x_161); +x_164 = l_Lean_Meta_Grind_addCongrTable___closed__7; +x_165 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_165, 0, x_163); +lean_ctor_set(x_165, 1, x_164); +x_166 = l_Lean_indentExpr(x_142); lean_ctor_set_tag(x_11, 7); -lean_ctor_set(x_11, 1, x_140); -lean_ctor_set(x_11, 0, x_139); -x_141 = l_Lean_Meta_Grind_addCongrTable___closed__9; -x_142 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_142, 0, x_11); -lean_ctor_set(x_142, 1, x_141); -x_143 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_126, x_142, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_129); -x_144 = lean_ctor_get(x_143, 0); -lean_inc(x_144); -x_145 = lean_ctor_get(x_143, 1); -lean_inc(x_145); -lean_dec(x_143); -x_146 = lean_apply_10(x_131, x_144, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_145); -return x_146; +lean_ctor_set(x_11, 1, x_166); +lean_ctor_set(x_11, 0, x_165); +x_167 = l_Lean_Meta_Grind_addCongrTable___closed__9; +x_168 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_168, 0, x_11); +lean_ctor_set(x_168, 1, x_167); +x_169 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_150, x_168, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_160); +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_apply_10(x_155, x_170, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_171); +return x_172; +} +else +{ +lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; +lean_dec(x_154); +lean_dec(x_142); +lean_free_object(x_11); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_173 = lean_ctor_get(x_159, 0); +lean_inc(x_173); +x_174 = lean_ctor_get(x_159, 1); +lean_inc(x_174); +if (lean_is_exclusive(x_159)) { + lean_ctor_release(x_159, 0); + lean_ctor_release(x_159, 1); + x_175 = x_159; +} else { + lean_dec_ref(x_159); + x_175 = lean_box(0); +} +if (lean_is_scalar(x_175)) { + x_176 = lean_alloc_ctor(1, 2, 0); +} else { + x_176 = x_175; +} +lean_ctor_set(x_176, 0, x_173); +lean_ctor_set(x_176, 1, x_174); +return x_176; +} } } else { -lean_object* x_147; lean_object* x_148; lean_object* x_149; +lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_free_object(x_11); -x_147 = lean_ctor_get(x_122, 1); -lean_inc(x_147); -lean_dec(x_122); -x_148 = lean_box(0); -x_149 = l_Lean_Meta_Grind_addCongrTable___lambda__2(x_1, x_118, x_148, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_147); +x_177 = lean_ctor_get(x_146, 1); +lean_inc(x_177); +lean_dec(x_146); +x_178 = lean_box(0); +x_179 = l_Lean_Meta_Grind_addCongrTable___lambda__2(x_1, x_142, x_178, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_177); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -return x_149; +return x_179; } } else { -lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; -lean_dec(x_118); +lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; +lean_dec(x_142); lean_free_object(x_11); lean_dec(x_9); lean_dec(x_8); @@ -2384,60 +2341,60 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_150 = lean_ctor_get(x_122, 0); -lean_inc(x_150); -x_151 = lean_ctor_get(x_122, 1); -lean_inc(x_151); -if (lean_is_exclusive(x_122)) { - lean_ctor_release(x_122, 0); - lean_ctor_release(x_122, 1); - x_152 = x_122; +x_180 = lean_ctor_get(x_146, 0); +lean_inc(x_180); +x_181 = lean_ctor_get(x_146, 1); +lean_inc(x_181); +if (lean_is_exclusive(x_146)) { + lean_ctor_release(x_146, 0); + lean_ctor_release(x_146, 1); + x_182 = x_146; } else { - lean_dec_ref(x_122); - x_152 = lean_box(0); + lean_dec_ref(x_146); + x_182 = lean_box(0); } -if (lean_is_scalar(x_152)) { - x_153 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_182)) { + x_183 = lean_alloc_ctor(1, 2, 0); } else { - x_153 = x_152; + x_183 = x_182; } -lean_ctor_set(x_153, 0, x_150); -lean_ctor_set(x_153, 1, x_151); -return x_153; +lean_ctor_set(x_183, 0, x_180); +lean_ctor_set(x_183, 1, x_181); +return x_183; } } else { -lean_object* x_154; lean_object* x_155; -lean_dec(x_120); -lean_dec(x_119); +lean_object* x_184; lean_object* x_185; +lean_dec(x_144); +lean_dec(x_143); lean_free_object(x_11); -x_154 = lean_box(0); -x_155 = l_Lean_Meta_Grind_addCongrTable___lambda__2(x_1, x_118, x_154, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_14); +x_184 = lean_box(0); +x_185 = l_Lean_Meta_Grind_addCongrTable___lambda__2(x_1, x_142, x_184, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_14); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -return x_155; +return x_185; } } } } else { -lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; -x_156 = lean_ctor_get(x_11, 0); -x_157 = lean_ctor_get(x_11, 1); -lean_inc(x_157); -lean_inc(x_156); +lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; +x_186 = lean_ctor_get(x_11, 0); +x_187 = lean_ctor_get(x_11, 1); +lean_inc(x_187); +lean_inc(x_186); lean_dec(x_11); -x_158 = lean_ctor_get(x_156, 3); -lean_inc(x_158); +x_188 = lean_ctor_get(x_186, 3); +lean_inc(x_188); lean_inc(x_1); -x_159 = l_Lean_PersistentHashMap_findEntry_x3f___at_Lean_Meta_Grind_addCongrTable___spec__1(x_156, x_158, x_1); -if (lean_obj_tag(x_159) == 0) +x_189 = l_Lean_PersistentHashMap_findEntry_x3f___at_Lean_Meta_Grind_addCongrTable___spec__1(x_186, x_188, x_1); +if (lean_obj_tag(x_189) == 0) { -lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; uint8_t 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_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; uint8_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; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -2445,234 +2402,302 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_160 = lean_st_ref_take(x_2, x_157); -x_161 = lean_ctor_get(x_160, 0); -lean_inc(x_161); -x_162 = lean_ctor_get(x_160, 1); -lean_inc(x_162); -lean_dec(x_160); -x_163 = lean_ctor_get(x_161, 0); -lean_inc(x_163); -x_164 = lean_ctor_get(x_161, 1); -lean_inc(x_164); -x_165 = lean_ctor_get(x_161, 2); -lean_inc(x_165); -x_166 = lean_ctor_get(x_161, 3); -lean_inc(x_166); -x_167 = lean_ctor_get(x_161, 4); -lean_inc(x_167); -x_168 = lean_ctor_get(x_161, 5); -lean_inc(x_168); -x_169 = lean_ctor_get_uint8(x_161, sizeof(void*)*14); -x_170 = lean_ctor_get(x_161, 6); -lean_inc(x_170); -x_171 = lean_ctor_get(x_161, 7); -lean_inc(x_171); -x_172 = lean_ctor_get(x_161, 8); -lean_inc(x_172); -x_173 = lean_ctor_get(x_161, 9); -lean_inc(x_173); -x_174 = lean_ctor_get(x_161, 10); -lean_inc(x_174); -x_175 = lean_ctor_get(x_161, 11); -lean_inc(x_175); -x_176 = lean_ctor_get(x_161, 12); -lean_inc(x_176); -x_177 = lean_ctor_get(x_161, 13); -lean_inc(x_177); -x_178 = lean_box(0); -lean_inc(x_161); -x_179 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_addCongrTable___spec__4(x_161, x_166, x_1, x_178); -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); - lean_ctor_release(x_161, 5); - lean_ctor_release(x_161, 6); - lean_ctor_release(x_161, 7); - lean_ctor_release(x_161, 8); - lean_ctor_release(x_161, 9); - lean_ctor_release(x_161, 10); - lean_ctor_release(x_161, 11); - lean_ctor_release(x_161, 12); - lean_ctor_release(x_161, 13); - x_180 = x_161; +x_190 = lean_st_ref_take(x_2, x_187); +x_191 = lean_ctor_get(x_190, 0); +lean_inc(x_191); +x_192 = lean_ctor_get(x_190, 1); +lean_inc(x_192); +lean_dec(x_190); +x_193 = lean_ctor_get(x_191, 0); +lean_inc(x_193); +x_194 = lean_ctor_get(x_191, 1); +lean_inc(x_194); +x_195 = lean_ctor_get(x_191, 2); +lean_inc(x_195); +x_196 = lean_ctor_get(x_191, 3); +lean_inc(x_196); +x_197 = lean_ctor_get(x_191, 4); +lean_inc(x_197); +x_198 = lean_ctor_get(x_191, 5); +lean_inc(x_198); +x_199 = lean_ctor_get_uint8(x_191, sizeof(void*)*20); +x_200 = lean_ctor_get(x_191, 6); +lean_inc(x_200); +x_201 = lean_ctor_get(x_191, 7); +lean_inc(x_201); +x_202 = lean_ctor_get(x_191, 8); +lean_inc(x_202); +x_203 = lean_ctor_get(x_191, 9); +lean_inc(x_203); +x_204 = lean_ctor_get(x_191, 10); +lean_inc(x_204); +x_205 = lean_ctor_get(x_191, 11); +lean_inc(x_205); +x_206 = lean_ctor_get(x_191, 12); +lean_inc(x_206); +x_207 = lean_ctor_get(x_191, 13); +lean_inc(x_207); +x_208 = lean_ctor_get(x_191, 14); +lean_inc(x_208); +x_209 = lean_ctor_get(x_191, 15); +lean_inc(x_209); +x_210 = lean_ctor_get(x_191, 16); +lean_inc(x_210); +x_211 = lean_ctor_get(x_191, 17); +lean_inc(x_211); +x_212 = lean_ctor_get(x_191, 18); +lean_inc(x_212); +x_213 = lean_ctor_get(x_191, 19); +lean_inc(x_213); +x_214 = lean_box(0); +lean_inc(x_191); +x_215 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_addCongrTable___spec__4(x_191, x_196, x_1, x_214); +if (lean_is_exclusive(x_191)) { + lean_ctor_release(x_191, 0); + lean_ctor_release(x_191, 1); + lean_ctor_release(x_191, 2); + lean_ctor_release(x_191, 3); + lean_ctor_release(x_191, 4); + lean_ctor_release(x_191, 5); + lean_ctor_release(x_191, 6); + lean_ctor_release(x_191, 7); + lean_ctor_release(x_191, 8); + lean_ctor_release(x_191, 9); + lean_ctor_release(x_191, 10); + lean_ctor_release(x_191, 11); + lean_ctor_release(x_191, 12); + lean_ctor_release(x_191, 13); + lean_ctor_release(x_191, 14); + lean_ctor_release(x_191, 15); + lean_ctor_release(x_191, 16); + lean_ctor_release(x_191, 17); + lean_ctor_release(x_191, 18); + lean_ctor_release(x_191, 19); + x_216 = x_191; } else { - lean_dec_ref(x_161); - x_180 = lean_box(0); + lean_dec_ref(x_191); + x_216 = lean_box(0); } -if (lean_is_scalar(x_180)) { - x_181 = lean_alloc_ctor(0, 14, 1); +if (lean_is_scalar(x_216)) { + x_217 = lean_alloc_ctor(0, 20, 1); } else { - x_181 = x_180; -} -lean_ctor_set(x_181, 0, x_163); -lean_ctor_set(x_181, 1, x_164); -lean_ctor_set(x_181, 2, x_165); -lean_ctor_set(x_181, 3, x_179); -lean_ctor_set(x_181, 4, x_167); -lean_ctor_set(x_181, 5, x_168); -lean_ctor_set(x_181, 6, x_170); -lean_ctor_set(x_181, 7, x_171); -lean_ctor_set(x_181, 8, x_172); -lean_ctor_set(x_181, 9, x_173); -lean_ctor_set(x_181, 10, x_174); -lean_ctor_set(x_181, 11, x_175); -lean_ctor_set(x_181, 12, x_176); -lean_ctor_set(x_181, 13, x_177); -lean_ctor_set_uint8(x_181, sizeof(void*)*14, x_169); -x_182 = lean_st_ref_set(x_2, x_181, x_162); -lean_dec(x_2); -x_183 = lean_ctor_get(x_182, 1); -lean_inc(x_183); -if (lean_is_exclusive(x_182)) { - lean_ctor_release(x_182, 0); - lean_ctor_release(x_182, 1); - x_184 = x_182; + x_217 = x_216; +} +lean_ctor_set(x_217, 0, x_193); +lean_ctor_set(x_217, 1, x_194); +lean_ctor_set(x_217, 2, x_195); +lean_ctor_set(x_217, 3, x_215); +lean_ctor_set(x_217, 4, x_197); +lean_ctor_set(x_217, 5, x_198); +lean_ctor_set(x_217, 6, x_200); +lean_ctor_set(x_217, 7, x_201); +lean_ctor_set(x_217, 8, x_202); +lean_ctor_set(x_217, 9, x_203); +lean_ctor_set(x_217, 10, x_204); +lean_ctor_set(x_217, 11, x_205); +lean_ctor_set(x_217, 12, x_206); +lean_ctor_set(x_217, 13, x_207); +lean_ctor_set(x_217, 14, x_208); +lean_ctor_set(x_217, 15, x_209); +lean_ctor_set(x_217, 16, x_210); +lean_ctor_set(x_217, 17, x_211); +lean_ctor_set(x_217, 18, x_212); +lean_ctor_set(x_217, 19, x_213); +lean_ctor_set_uint8(x_217, sizeof(void*)*20, x_199); +x_218 = lean_st_ref_set(x_2, x_217, x_192); +lean_dec(x_2); +x_219 = lean_ctor_get(x_218, 1); +lean_inc(x_219); +if (lean_is_exclusive(x_218)) { + lean_ctor_release(x_218, 0); + lean_ctor_release(x_218, 1); + x_220 = x_218; } else { - lean_dec_ref(x_182); - x_184 = lean_box(0); + lean_dec_ref(x_218); + x_220 = lean_box(0); } -if (lean_is_scalar(x_184)) { - x_185 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_220)) { + x_221 = lean_alloc_ctor(0, 2, 0); } else { - x_185 = x_184; + x_221 = x_220; } -lean_ctor_set(x_185, 0, x_178); -lean_ctor_set(x_185, 1, x_183); -return x_185; +lean_ctor_set(x_221, 0, x_214); +lean_ctor_set(x_221, 1, x_219); +return x_221; } else { -lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; uint8_t x_191; -x_186 = lean_ctor_get(x_159, 0); -lean_inc(x_186); -lean_dec(x_159); -x_187 = lean_ctor_get(x_186, 0); -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; +lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; uint8_t x_227; +x_222 = lean_ctor_get(x_189, 0); +lean_inc(x_222); +lean_dec(x_189); +x_223 = lean_ctor_get(x_222, 0); +lean_inc(x_223); +if (lean_is_exclusive(x_222)) { + lean_ctor_release(x_222, 0); + lean_ctor_release(x_222, 1); + x_224 = x_222; } else { - lean_dec_ref(x_186); - x_188 = lean_box(0); + lean_dec_ref(x_222); + x_224 = lean_box(0); } -x_189 = l_Lean_Expr_getAppFn(x_1); -x_190 = l_Lean_Expr_getAppFn(x_187); -x_191 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_189, x_190); -if (x_191 == 0) +x_225 = l_Lean_Expr_getAppFn(x_1); +x_226 = l_Lean_Expr_getAppFn(x_223); +x_227 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_225, x_226); +if (x_227 == 0) { -lean_object* x_192; +lean_object* x_228; lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -x_192 = l_Lean_Meta_Grind_hasSameType(x_189, x_190, x_6, x_7, x_8, x_9, x_157); -if (lean_obj_tag(x_192) == 0) +x_228 = l_Lean_Meta_Grind_hasSameType(x_225, x_226, x_6, x_7, x_8, x_9, x_187); +if (lean_obj_tag(x_228) == 0) { -lean_object* x_193; uint8_t x_194; -x_193 = lean_ctor_get(x_192, 0); -lean_inc(x_193); -x_194 = lean_unbox(x_193); -lean_dec(x_193); -if (x_194 == 0) +lean_object* x_229; uint8_t x_230; +x_229 = lean_ctor_get(x_228, 0); +lean_inc(x_229); +x_230 = lean_unbox(x_229); +lean_dec(x_229); +if (x_230 == 0) { -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; uint8_t x_202; -x_195 = lean_ctor_get(x_192, 1); -lean_inc(x_195); -lean_dec(x_192); -x_196 = l_Lean_Meta_Grind_addCongrTable___closed__2; -x_197 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_196, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_195); -x_198 = lean_ctor_get(x_197, 0); -lean_inc(x_198); -x_199 = lean_ctor_get(x_197, 1); -lean_inc(x_199); -if (lean_is_exclusive(x_197)) { - lean_ctor_release(x_197, 0); - lean_ctor_release(x_197, 1); - x_200 = x_197; +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; uint8_t x_238; +x_231 = lean_ctor_get(x_228, 1); +lean_inc(x_231); +lean_dec(x_228); +x_232 = l_Lean_Meta_Grind_addCongrTable___closed__2; +x_233 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_232, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, 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); +if (lean_is_exclusive(x_233)) { + lean_ctor_release(x_233, 0); + lean_ctor_release(x_233, 1); + x_236 = x_233; } else { - lean_dec_ref(x_197); - x_200 = lean_box(0); + lean_dec_ref(x_233); + x_236 = lean_box(0); +} +x_237 = l_Lean_Meta_Grind_addCongrTable___closed__3; +x_238 = lean_unbox(x_234); +lean_dec(x_234); +if (x_238 == 0) +{ +lean_object* x_239; lean_object* x_240; +lean_dec(x_236); +lean_dec(x_224); +lean_dec(x_223); +lean_dec(x_1); +x_239 = lean_box(0); +x_240 = lean_apply_10(x_237, x_239, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_235); +return x_240; } -x_201 = l_Lean_Meta_Grind_addCongrTable___closed__3; -x_202 = lean_unbox(x_198); -lean_dec(x_198); -if (x_202 == 0) +else { -lean_object* x_203; lean_object* x_204; -lean_dec(x_200); -lean_dec(x_188); -lean_dec(x_187); -lean_dec(x_1); -x_203 = lean_box(0); -x_204 = lean_apply_10(x_201, x_203, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_199); -return x_204; +lean_object* x_241; +x_241 = l_Lean_Meta_Grind_updateLastTag(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_235); +if (lean_obj_tag(x_241) == 0) +{ +lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; +x_242 = lean_ctor_get(x_241, 1); +lean_inc(x_242); +lean_dec(x_241); +x_243 = l_Lean_indentExpr(x_1); +x_244 = l_Lean_Meta_Grind_addCongrTable___closed__5; +if (lean_is_scalar(x_236)) { + x_245 = lean_alloc_ctor(7, 2, 0); +} else { + x_245 = x_236; + lean_ctor_set_tag(x_245, 7); +} +lean_ctor_set(x_245, 0, x_244); +lean_ctor_set(x_245, 1, x_243); +x_246 = l_Lean_Meta_Grind_addCongrTable___closed__7; +if (lean_is_scalar(x_224)) { + x_247 = lean_alloc_ctor(7, 2, 0); +} else { + x_247 = x_224; + lean_ctor_set_tag(x_247, 7); +} +lean_ctor_set(x_247, 0, x_245); +lean_ctor_set(x_247, 1, x_246); +x_248 = l_Lean_indentExpr(x_223); +x_249 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_249, 0, x_247); +lean_ctor_set(x_249, 1, x_248); +x_250 = l_Lean_Meta_Grind_addCongrTable___closed__9; +x_251 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_251, 0, x_249); +lean_ctor_set(x_251, 1, x_250); +x_252 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_232, x_251, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_242); +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_255 = lean_apply_10(x_237, x_253, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_254); +return x_255; } else { -lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; -x_205 = l_Lean_indentExpr(x_1); -x_206 = l_Lean_Meta_Grind_addCongrTable___closed__5; -if (lean_is_scalar(x_200)) { - x_207 = lean_alloc_ctor(7, 2, 0); +lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; +lean_dec(x_236); +lean_dec(x_224); +lean_dec(x_223); +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_256 = lean_ctor_get(x_241, 0); +lean_inc(x_256); +x_257 = lean_ctor_get(x_241, 1); +lean_inc(x_257); +if (lean_is_exclusive(x_241)) { + lean_ctor_release(x_241, 0); + lean_ctor_release(x_241, 1); + x_258 = x_241; } else { - x_207 = x_200; - lean_ctor_set_tag(x_207, 7); -} -lean_ctor_set(x_207, 0, x_206); -lean_ctor_set(x_207, 1, x_205); -x_208 = l_Lean_Meta_Grind_addCongrTable___closed__7; -if (lean_is_scalar(x_188)) { - x_209 = lean_alloc_ctor(7, 2, 0); + lean_dec_ref(x_241); + x_258 = lean_box(0); +} +if (lean_is_scalar(x_258)) { + x_259 = lean_alloc_ctor(1, 2, 0); } else { - x_209 = x_188; - lean_ctor_set_tag(x_209, 7); + x_259 = x_258; +} +lean_ctor_set(x_259, 0, x_256); +lean_ctor_set(x_259, 1, x_257); +return x_259; } -lean_ctor_set(x_209, 0, x_207); -lean_ctor_set(x_209, 1, x_208); -x_210 = l_Lean_indentExpr(x_187); -x_211 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_211, 0, x_209); -lean_ctor_set(x_211, 1, x_210); -x_212 = l_Lean_Meta_Grind_addCongrTable___closed__9; -x_213 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_213, 0, x_211); -lean_ctor_set(x_213, 1, x_212); -x_214 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_196, x_213, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_199); -x_215 = lean_ctor_get(x_214, 0); -lean_inc(x_215); -x_216 = lean_ctor_get(x_214, 1); -lean_inc(x_216); -lean_dec(x_214); -x_217 = lean_apply_10(x_201, x_215, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_216); -return x_217; } } else { -lean_object* x_218; lean_object* x_219; lean_object* x_220; -lean_dec(x_188); -x_218 = lean_ctor_get(x_192, 1); -lean_inc(x_218); -lean_dec(x_192); -x_219 = lean_box(0); -x_220 = l_Lean_Meta_Grind_addCongrTable___lambda__2(x_1, x_187, x_219, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_218); +lean_object* x_260; lean_object* x_261; lean_object* x_262; +lean_dec(x_224); +x_260 = lean_ctor_get(x_228, 1); +lean_inc(x_260); +lean_dec(x_228); +x_261 = lean_box(0); +x_262 = l_Lean_Meta_Grind_addCongrTable___lambda__2(x_1, x_223, x_261, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_260); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -return x_220; +return x_262; } } else { -lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; -lean_dec(x_188); -lean_dec(x_187); +lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; +lean_dec(x_224); +lean_dec(x_223); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -2682,41 +2707,41 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_221 = lean_ctor_get(x_192, 0); -lean_inc(x_221); -x_222 = lean_ctor_get(x_192, 1); -lean_inc(x_222); -if (lean_is_exclusive(x_192)) { - lean_ctor_release(x_192, 0); - lean_ctor_release(x_192, 1); - x_223 = x_192; +x_263 = lean_ctor_get(x_228, 0); +lean_inc(x_263); +x_264 = lean_ctor_get(x_228, 1); +lean_inc(x_264); +if (lean_is_exclusive(x_228)) { + lean_ctor_release(x_228, 0); + lean_ctor_release(x_228, 1); + x_265 = x_228; } else { - lean_dec_ref(x_192); - x_223 = lean_box(0); + lean_dec_ref(x_228); + x_265 = lean_box(0); } -if (lean_is_scalar(x_223)) { - x_224 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_265)) { + x_266 = lean_alloc_ctor(1, 2, 0); } else { - x_224 = x_223; + x_266 = x_265; } -lean_ctor_set(x_224, 0, x_221); -lean_ctor_set(x_224, 1, x_222); -return x_224; +lean_ctor_set(x_266, 0, x_263); +lean_ctor_set(x_266, 1, x_264); +return x_266; } } else { -lean_object* x_225; lean_object* x_226; -lean_dec(x_190); -lean_dec(x_189); -lean_dec(x_188); -x_225 = lean_box(0); -x_226 = l_Lean_Meta_Grind_addCongrTable___lambda__2(x_1, x_187, x_225, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_157); +lean_object* x_267; lean_object* x_268; +lean_dec(x_226); +lean_dec(x_225); +lean_dec(x_224); +x_267 = lean_box(0); +x_268 = l_Lean_Meta_Grind_addCongrTable___lambda__2(x_1, x_223, x_267, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_187); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -return x_226; +return x_268; } } } @@ -2766,38 +2791,6 @@ x_9 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_addCongrTable___sp return x_9; } } -LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_11; -x_11 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_11; -} -} -LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -lean_object* x_12; -x_12 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -return x_12; -} -} LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addCongrTable___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: { @@ -3628,7 +3621,7 @@ return x_36; } else { -lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; 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_37; 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; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; x_37 = lean_ctor_get(x_12, 1); x_38 = lean_ctor_get(x_14, 0); x_39 = lean_ctor_get(x_14, 1); @@ -3636,7 +3629,7 @@ x_40 = lean_ctor_get(x_14, 2); x_41 = lean_ctor_get(x_14, 3); x_42 = lean_ctor_get(x_14, 4); x_43 = lean_ctor_get(x_14, 5); -x_44 = lean_ctor_get_uint8(x_14, sizeof(void*)*14); +x_44 = lean_ctor_get_uint8(x_14, sizeof(void*)*20); x_45 = lean_ctor_get(x_14, 6); x_46 = lean_ctor_get(x_14, 7); x_47 = lean_ctor_get(x_14, 8); @@ -3645,6 +3638,18 @@ x_49 = lean_ctor_get(x_14, 10); x_50 = lean_ctor_get(x_14, 11); x_51 = lean_ctor_get(x_14, 12); x_52 = lean_ctor_get(x_14, 13); +x_53 = lean_ctor_get(x_14, 14); +x_54 = lean_ctor_get(x_14, 15); +x_55 = lean_ctor_get(x_14, 16); +x_56 = lean_ctor_get(x_14, 17); +x_57 = lean_ctor_get(x_14, 18); +x_58 = lean_ctor_get(x_14, 19); +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); @@ -3661,259 +3666,301 @@ lean_inc(x_39); lean_inc(x_38); lean_dec(x_14); lean_inc(x_42); -x_53 = l_Lean_PersistentHashMap_find_x3f___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__1(x_42, x_11); -if (lean_obj_tag(x_53) == 0) +x_59 = l_Lean_PersistentHashMap_find_x3f___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__1(x_42, x_11); +if (lean_obj_tag(x_59) == 0) { -lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; -x_54 = lean_box(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; +x_60 = lean_box(0); lean_ctor_set_tag(x_12, 1); -lean_ctor_set(x_12, 1, x_54); +lean_ctor_set(x_12, 1, x_60); lean_ctor_set(x_12, 0, x_1); -x_55 = l_Lean_PersistentHashMap_insert___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__4(x_42, x_11, x_12); -x_56 = lean_alloc_ctor(0, 14, 1); -lean_ctor_set(x_56, 0, x_38); -lean_ctor_set(x_56, 1, x_39); -lean_ctor_set(x_56, 2, x_40); -lean_ctor_set(x_56, 3, x_41); -lean_ctor_set(x_56, 4, x_55); -lean_ctor_set(x_56, 5, x_43); -lean_ctor_set(x_56, 6, x_45); -lean_ctor_set(x_56, 7, x_46); -lean_ctor_set(x_56, 8, x_47); -lean_ctor_set(x_56, 9, x_48); -lean_ctor_set(x_56, 10, x_49); -lean_ctor_set(x_56, 11, x_50); -lean_ctor_set(x_56, 12, x_51); -lean_ctor_set(x_56, 13, x_52); -lean_ctor_set_uint8(x_56, sizeof(void*)*14, x_44); -x_57 = lean_st_ref_set(x_2, x_56, x_37); -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; +x_61 = l_Lean_PersistentHashMap_insert___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__4(x_42, x_11, x_12); +x_62 = lean_alloc_ctor(0, 20, 1); +lean_ctor_set(x_62, 0, x_38); +lean_ctor_set(x_62, 1, x_39); +lean_ctor_set(x_62, 2, x_40); +lean_ctor_set(x_62, 3, x_41); +lean_ctor_set(x_62, 4, x_61); +lean_ctor_set(x_62, 5, x_43); +lean_ctor_set(x_62, 6, x_45); +lean_ctor_set(x_62, 7, x_46); +lean_ctor_set(x_62, 8, x_47); +lean_ctor_set(x_62, 9, x_48); +lean_ctor_set(x_62, 10, x_49); +lean_ctor_set(x_62, 11, x_50); +lean_ctor_set(x_62, 12, x_51); +lean_ctor_set(x_62, 13, x_52); +lean_ctor_set(x_62, 14, x_53); +lean_ctor_set(x_62, 15, x_54); +lean_ctor_set(x_62, 16, x_55); +lean_ctor_set(x_62, 17, x_56); +lean_ctor_set(x_62, 18, x_57); +lean_ctor_set(x_62, 19, x_58); +lean_ctor_set_uint8(x_62, sizeof(void*)*20, x_44); +x_63 = lean_st_ref_set(x_2, x_62, x_37); +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_57); - x_59 = lean_box(0); + lean_dec_ref(x_63); + x_65 = lean_box(0); } -x_60 = lean_box(0); -if (lean_is_scalar(x_59)) { - x_61 = lean_alloc_ctor(0, 2, 0); +x_66 = lean_box(0); +if (lean_is_scalar(x_65)) { + x_67 = lean_alloc_ctor(0, 2, 0); } else { - x_61 = x_59; + x_67 = x_65; } -lean_ctor_set(x_61, 0, x_60); -lean_ctor_set(x_61, 1, x_58); -return x_61; +lean_ctor_set(x_67, 0, x_66); +lean_ctor_set(x_67, 1, x_64); +return x_67; } else { -lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; -x_62 = lean_ctor_get(x_53, 0); -lean_inc(x_62); -lean_dec(x_53); +lean_object* x_68; 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_68 = lean_ctor_get(x_59, 0); +lean_inc(x_68); +lean_dec(x_59); lean_ctor_set_tag(x_12, 1); -lean_ctor_set(x_12, 1, x_62); +lean_ctor_set(x_12, 1, x_68); lean_ctor_set(x_12, 0, x_1); -x_63 = l_Lean_PersistentHashMap_insert___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__4(x_42, x_11, x_12); -x_64 = lean_alloc_ctor(0, 14, 1); -lean_ctor_set(x_64, 0, x_38); -lean_ctor_set(x_64, 1, x_39); -lean_ctor_set(x_64, 2, x_40); -lean_ctor_set(x_64, 3, x_41); -lean_ctor_set(x_64, 4, x_63); -lean_ctor_set(x_64, 5, x_43); -lean_ctor_set(x_64, 6, x_45); -lean_ctor_set(x_64, 7, x_46); -lean_ctor_set(x_64, 8, x_47); -lean_ctor_set(x_64, 9, x_48); -lean_ctor_set(x_64, 10, x_49); -lean_ctor_set(x_64, 11, x_50); -lean_ctor_set(x_64, 12, x_51); -lean_ctor_set(x_64, 13, x_52); -lean_ctor_set_uint8(x_64, sizeof(void*)*14, x_44); -x_65 = lean_st_ref_set(x_2, x_64, x_37); -x_66 = lean_ctor_get(x_65, 1); -lean_inc(x_66); -if (lean_is_exclusive(x_65)) { - lean_ctor_release(x_65, 0); - lean_ctor_release(x_65, 1); - x_67 = x_65; +x_69 = l_Lean_PersistentHashMap_insert___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__4(x_42, x_11, x_12); +x_70 = lean_alloc_ctor(0, 20, 1); +lean_ctor_set(x_70, 0, x_38); +lean_ctor_set(x_70, 1, x_39); +lean_ctor_set(x_70, 2, x_40); +lean_ctor_set(x_70, 3, x_41); +lean_ctor_set(x_70, 4, x_69); +lean_ctor_set(x_70, 5, x_43); +lean_ctor_set(x_70, 6, x_45); +lean_ctor_set(x_70, 7, x_46); +lean_ctor_set(x_70, 8, x_47); +lean_ctor_set(x_70, 9, x_48); +lean_ctor_set(x_70, 10, x_49); +lean_ctor_set(x_70, 11, x_50); +lean_ctor_set(x_70, 12, x_51); +lean_ctor_set(x_70, 13, x_52); +lean_ctor_set(x_70, 14, x_53); +lean_ctor_set(x_70, 15, x_54); +lean_ctor_set(x_70, 16, x_55); +lean_ctor_set(x_70, 17, x_56); +lean_ctor_set(x_70, 18, x_57); +lean_ctor_set(x_70, 19, x_58); +lean_ctor_set_uint8(x_70, sizeof(void*)*20, x_44); +x_71 = lean_st_ref_set(x_2, x_70, x_37); +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_65); - x_67 = lean_box(0); + lean_dec_ref(x_71); + x_73 = lean_box(0); } -x_68 = lean_box(0); -if (lean_is_scalar(x_67)) { - x_69 = lean_alloc_ctor(0, 2, 0); +x_74 = lean_box(0); +if (lean_is_scalar(x_73)) { + x_75 = lean_alloc_ctor(0, 2, 0); } else { - x_69 = x_67; + x_75 = x_73; } -lean_ctor_set(x_69, 0, x_68); -lean_ctor_set(x_69, 1, x_66); -return x_69; +lean_ctor_set(x_75, 0, x_74); +lean_ctor_set(x_75, 1, x_72); +return x_75; } } } 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; uint8_t x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; -x_70 = lean_ctor_get(x_12, 0); -x_71 = lean_ctor_get(x_12, 1); -lean_inc(x_71); -lean_inc(x_70); -lean_dec(x_12); -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, 4); -lean_inc(x_76); -x_77 = lean_ctor_get(x_70, 5); +lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; uint8_t x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; 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_76 = lean_ctor_get(x_12, 0); +x_77 = lean_ctor_get(x_12, 1); lean_inc(x_77); -x_78 = lean_ctor_get_uint8(x_70, sizeof(void*)*14); -x_79 = lean_ctor_get(x_70, 6); +lean_inc(x_76); +lean_dec(x_12); +x_78 = lean_ctor_get(x_76, 0); +lean_inc(x_78); +x_79 = lean_ctor_get(x_76, 1); lean_inc(x_79); -x_80 = lean_ctor_get(x_70, 7); +x_80 = lean_ctor_get(x_76, 2); lean_inc(x_80); -x_81 = lean_ctor_get(x_70, 8); +x_81 = lean_ctor_get(x_76, 3); lean_inc(x_81); -x_82 = lean_ctor_get(x_70, 9); +x_82 = lean_ctor_get(x_76, 4); lean_inc(x_82); -x_83 = lean_ctor_get(x_70, 10); +x_83 = lean_ctor_get(x_76, 5); lean_inc(x_83); -x_84 = lean_ctor_get(x_70, 11); -lean_inc(x_84); -x_85 = lean_ctor_get(x_70, 12); +x_84 = lean_ctor_get_uint8(x_76, sizeof(void*)*20); +x_85 = lean_ctor_get(x_76, 6); lean_inc(x_85); -x_86 = lean_ctor_get(x_70, 13); +x_86 = lean_ctor_get(x_76, 7); lean_inc(x_86); -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); - lean_ctor_release(x_70, 7); - lean_ctor_release(x_70, 8); - lean_ctor_release(x_70, 9); - lean_ctor_release(x_70, 10); - lean_ctor_release(x_70, 11); - lean_ctor_release(x_70, 12); - lean_ctor_release(x_70, 13); - x_87 = x_70; +x_87 = lean_ctor_get(x_76, 8); +lean_inc(x_87); +x_88 = lean_ctor_get(x_76, 9); +lean_inc(x_88); +x_89 = lean_ctor_get(x_76, 10); +lean_inc(x_89); +x_90 = lean_ctor_get(x_76, 11); +lean_inc(x_90); +x_91 = lean_ctor_get(x_76, 12); +lean_inc(x_91); +x_92 = lean_ctor_get(x_76, 13); +lean_inc(x_92); +x_93 = lean_ctor_get(x_76, 14); +lean_inc(x_93); +x_94 = lean_ctor_get(x_76, 15); +lean_inc(x_94); +x_95 = lean_ctor_get(x_76, 16); +lean_inc(x_95); +x_96 = lean_ctor_get(x_76, 17); +lean_inc(x_96); +x_97 = lean_ctor_get(x_76, 18); +lean_inc(x_97); +x_98 = lean_ctor_get(x_76, 19); +lean_inc(x_98); +if (lean_is_exclusive(x_76)) { + lean_ctor_release(x_76, 0); + lean_ctor_release(x_76, 1); + lean_ctor_release(x_76, 2); + lean_ctor_release(x_76, 3); + lean_ctor_release(x_76, 4); + lean_ctor_release(x_76, 5); + lean_ctor_release(x_76, 6); + lean_ctor_release(x_76, 7); + lean_ctor_release(x_76, 8); + lean_ctor_release(x_76, 9); + lean_ctor_release(x_76, 10); + lean_ctor_release(x_76, 11); + lean_ctor_release(x_76, 12); + lean_ctor_release(x_76, 13); + lean_ctor_release(x_76, 14); + lean_ctor_release(x_76, 15); + lean_ctor_release(x_76, 16); + lean_ctor_release(x_76, 17); + lean_ctor_release(x_76, 18); + lean_ctor_release(x_76, 19); + x_99 = x_76; } else { - lean_dec_ref(x_70); - x_87 = lean_box(0); + lean_dec_ref(x_76); + x_99 = lean_box(0); } -lean_inc(x_76); -x_88 = l_Lean_PersistentHashMap_find_x3f___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__1(x_76, x_11); -if (lean_obj_tag(x_88) == 0) +lean_inc(x_82); +x_100 = l_Lean_PersistentHashMap_find_x3f___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__1(x_82, x_11); +if (lean_obj_tag(x_100) == 0) { -lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; -x_89 = lean_box(0); -x_90 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_90, 0, x_1); -lean_ctor_set(x_90, 1, x_89); -x_91 = l_Lean_PersistentHashMap_insert___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__4(x_76, x_11, x_90); -if (lean_is_scalar(x_87)) { - x_92 = lean_alloc_ctor(0, 14, 1); +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_101 = lean_box(0); +x_102 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_102, 0, x_1); +lean_ctor_set(x_102, 1, x_101); +x_103 = l_Lean_PersistentHashMap_insert___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__4(x_82, x_11, x_102); +if (lean_is_scalar(x_99)) { + x_104 = lean_alloc_ctor(0, 20, 1); } else { - x_92 = x_87; -} -lean_ctor_set(x_92, 0, x_72); -lean_ctor_set(x_92, 1, x_73); -lean_ctor_set(x_92, 2, x_74); -lean_ctor_set(x_92, 3, x_75); -lean_ctor_set(x_92, 4, x_91); -lean_ctor_set(x_92, 5, x_77); -lean_ctor_set(x_92, 6, x_79); -lean_ctor_set(x_92, 7, x_80); -lean_ctor_set(x_92, 8, x_81); -lean_ctor_set(x_92, 9, x_82); -lean_ctor_set(x_92, 10, x_83); -lean_ctor_set(x_92, 11, x_84); -lean_ctor_set(x_92, 12, x_85); -lean_ctor_set(x_92, 13, x_86); -lean_ctor_set_uint8(x_92, sizeof(void*)*14, x_78); -x_93 = lean_st_ref_set(x_2, x_92, x_71); -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; + x_104 = x_99; +} +lean_ctor_set(x_104, 0, x_78); +lean_ctor_set(x_104, 1, x_79); +lean_ctor_set(x_104, 2, x_80); +lean_ctor_set(x_104, 3, x_81); +lean_ctor_set(x_104, 4, x_103); +lean_ctor_set(x_104, 5, x_83); +lean_ctor_set(x_104, 6, x_85); +lean_ctor_set(x_104, 7, x_86); +lean_ctor_set(x_104, 8, x_87); +lean_ctor_set(x_104, 9, x_88); +lean_ctor_set(x_104, 10, x_89); +lean_ctor_set(x_104, 11, x_90); +lean_ctor_set(x_104, 12, x_91); +lean_ctor_set(x_104, 13, x_92); +lean_ctor_set(x_104, 14, x_93); +lean_ctor_set(x_104, 15, x_94); +lean_ctor_set(x_104, 16, x_95); +lean_ctor_set(x_104, 17, x_96); +lean_ctor_set(x_104, 18, x_97); +lean_ctor_set(x_104, 19, x_98); +lean_ctor_set_uint8(x_104, sizeof(void*)*20, x_84); +x_105 = lean_st_ref_set(x_2, x_104, x_77); +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_93); - x_95 = lean_box(0); + lean_dec_ref(x_105); + x_107 = lean_box(0); } -x_96 = lean_box(0); -if (lean_is_scalar(x_95)) { - x_97 = lean_alloc_ctor(0, 2, 0); +x_108 = lean_box(0); +if (lean_is_scalar(x_107)) { + x_109 = lean_alloc_ctor(0, 2, 0); } else { - x_97 = x_95; + x_109 = x_107; } -lean_ctor_set(x_97, 0, x_96); -lean_ctor_set(x_97, 1, x_94); -return x_97; +lean_ctor_set(x_109, 0, x_108); +lean_ctor_set(x_109, 1, x_106); +return x_109; } else { -lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; -x_98 = lean_ctor_get(x_88, 0); -lean_inc(x_98); -lean_dec(x_88); -x_99 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_99, 0, x_1); -lean_ctor_set(x_99, 1, x_98); -x_100 = l_Lean_PersistentHashMap_insert___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__4(x_76, x_11, x_99); -if (lean_is_scalar(x_87)) { - x_101 = lean_alloc_ctor(0, 14, 1); +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_100, 0); +lean_inc(x_110); +lean_dec(x_100); +x_111 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_111, 0, x_1); +lean_ctor_set(x_111, 1, x_110); +x_112 = l_Lean_PersistentHashMap_insert___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__4(x_82, x_11, x_111); +if (lean_is_scalar(x_99)) { + x_113 = lean_alloc_ctor(0, 20, 1); } else { - x_101 = x_87; -} -lean_ctor_set(x_101, 0, x_72); -lean_ctor_set(x_101, 1, x_73); -lean_ctor_set(x_101, 2, x_74); -lean_ctor_set(x_101, 3, x_75); -lean_ctor_set(x_101, 4, x_100); -lean_ctor_set(x_101, 5, x_77); -lean_ctor_set(x_101, 6, x_79); -lean_ctor_set(x_101, 7, x_80); -lean_ctor_set(x_101, 8, x_81); -lean_ctor_set(x_101, 9, x_82); -lean_ctor_set(x_101, 10, x_83); -lean_ctor_set(x_101, 11, x_84); -lean_ctor_set(x_101, 12, x_85); -lean_ctor_set(x_101, 13, x_86); -lean_ctor_set_uint8(x_101, sizeof(void*)*14, x_78); -x_102 = lean_st_ref_set(x_2, x_101, x_71); -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; + x_113 = x_99; +} +lean_ctor_set(x_113, 0, x_78); +lean_ctor_set(x_113, 1, x_79); +lean_ctor_set(x_113, 2, x_80); +lean_ctor_set(x_113, 3, x_81); +lean_ctor_set(x_113, 4, x_112); +lean_ctor_set(x_113, 5, x_83); +lean_ctor_set(x_113, 6, x_85); +lean_ctor_set(x_113, 7, x_86); +lean_ctor_set(x_113, 8, x_87); +lean_ctor_set(x_113, 9, x_88); +lean_ctor_set(x_113, 10, x_89); +lean_ctor_set(x_113, 11, x_90); +lean_ctor_set(x_113, 12, x_91); +lean_ctor_set(x_113, 13, x_92); +lean_ctor_set(x_113, 14, x_93); +lean_ctor_set(x_113, 15, x_94); +lean_ctor_set(x_113, 16, x_95); +lean_ctor_set(x_113, 17, x_96); +lean_ctor_set(x_113, 18, x_97); +lean_ctor_set(x_113, 19, x_98); +lean_ctor_set_uint8(x_113, sizeof(void*)*20, x_84); +x_114 = lean_st_ref_set(x_2, x_113, x_77); +x_115 = lean_ctor_get(x_114, 1); +lean_inc(x_115); +if (lean_is_exclusive(x_114)) { + lean_ctor_release(x_114, 0); + lean_ctor_release(x_114, 1); + x_116 = x_114; } else { - lean_dec_ref(x_102); - x_104 = lean_box(0); + lean_dec_ref(x_114); + x_116 = lean_box(0); } -x_105 = lean_box(0); -if (lean_is_scalar(x_104)) { - x_106 = lean_alloc_ctor(0, 2, 0); +x_117 = lean_box(0); +if (lean_is_scalar(x_116)) { + x_118 = lean_alloc_ctor(0, 2, 0); } else { - x_106 = x_104; + x_118 = x_116; } -lean_ctor_set(x_106, 0, x_105); -lean_ctor_set(x_106, 1, x_103); -return x_106; +lean_ctor_set(x_118, 0, x_117); +lean_ctor_set(x_118, 1, x_115); +return x_118; } } } @@ -3989,387 +4036,444 @@ lean_dec(x_2); return x_11; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___spec__1(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addSplitCandidate___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_14; -x_14 = lean_usize_dec_lt(x_3, x_2); -if (x_14 == 0) +lean_object* x_12; uint8_t x_13; +x_12 = lean_st_ref_take(x_3, x_11); +x_13 = !lean_is_exclusive(x_12); +if (x_13 == 0) { -lean_object* 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_1); -x_15 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_15, 0, x_4); -lean_ctor_set(x_15, 1, x_13); -return x_15; -} -else +lean_object* x_14; uint8_t x_15; +x_14 = lean_ctor_get(x_12, 0); +x_15 = !lean_is_exclusive(x_14); +if (x_15 == 0) { -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_16 = lean_array_uget(x_4, x_3); -x_17 = lean_unsigned_to_nat(0u); -x_18 = lean_array_uset(x_4, x_3, x_17); -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_1); -x_19 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern(x_16, x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -if (lean_obj_tag(x_19) == 0) +lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; +x_16 = lean_ctor_get(x_12, 1); +x_17 = lean_ctor_get(x_14, 16); +lean_ctor_set_tag(x_12, 1); +lean_ctor_set(x_12, 1, x_17); +lean_ctor_set(x_12, 0, x_1); +lean_ctor_set(x_14, 16, x_12); +x_18 = lean_st_ref_set(x_3, x_14, x_16); +x_19 = !lean_is_exclusive(x_18); +if (x_19 == 0) { -lean_object* x_20; lean_object* x_21; size_t x_22; size_t x_23; lean_object* x_24; -x_20 = lean_ctor_get(x_19, 0); -lean_inc(x_20); -x_21 = lean_ctor_get(x_19, 1); -lean_inc(x_21); -lean_dec(x_19); -x_22 = 1; -x_23 = lean_usize_add(x_3, x_22); -x_24 = lean_array_uset(x_18, x_3, x_20); -x_3 = x_23; -x_4 = x_24; -x_13 = x_21; -goto _start; +lean_object* x_20; lean_object* x_21; +x_20 = lean_ctor_get(x_18, 0); +lean_dec(x_20); +x_21 = lean_box(0); +lean_ctor_set(x_18, 0, x_21); +return x_18; } else { -uint8_t x_26; +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); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_1); -x_26 = !lean_is_exclusive(x_19); -if (x_26 == 0) -{ -return 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_22); +return x_24; +} } else { -lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_27 = lean_ctor_get(x_19, 0); -x_28 = lean_ctor_get(x_19, 1); +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; uint8_t x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; 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_25 = lean_ctor_get(x_12, 1); +x_26 = lean_ctor_get(x_14, 0); +x_27 = lean_ctor_get(x_14, 1); +x_28 = lean_ctor_get(x_14, 2); +x_29 = lean_ctor_get(x_14, 3); +x_30 = lean_ctor_get(x_14, 4); +x_31 = lean_ctor_get(x_14, 5); +x_32 = lean_ctor_get_uint8(x_14, sizeof(void*)*20); +x_33 = lean_ctor_get(x_14, 6); +x_34 = lean_ctor_get(x_14, 7); +x_35 = lean_ctor_get(x_14, 8); +x_36 = lean_ctor_get(x_14, 9); +x_37 = lean_ctor_get(x_14, 10); +x_38 = lean_ctor_get(x_14, 11); +x_39 = lean_ctor_get(x_14, 12); +x_40 = lean_ctor_get(x_14, 13); +x_41 = lean_ctor_get(x_14, 14); +x_42 = lean_ctor_get(x_14, 15); +x_43 = lean_ctor_get(x_14, 16); +x_44 = lean_ctor_get(x_14, 17); +x_45 = lean_ctor_get(x_14, 18); +x_46 = lean_ctor_get(x_14, 19); +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_inc(x_34); +lean_inc(x_33); +lean_inc(x_31); +lean_inc(x_30); +lean_inc(x_29); lean_inc(x_28); lean_inc(x_27); -lean_dec(x_19); -x_29 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_29, 0, x_27); -lean_ctor_set(x_29, 1, x_28); -return x_29; -} -} -} +lean_inc(x_26); +lean_dec(x_14); +lean_ctor_set_tag(x_12, 1); +lean_ctor_set(x_12, 1, x_43); +lean_ctor_set(x_12, 0, x_1); +x_47 = lean_alloc_ctor(0, 20, 1); +lean_ctor_set(x_47, 0, x_26); +lean_ctor_set(x_47, 1, x_27); +lean_ctor_set(x_47, 2, x_28); +lean_ctor_set(x_47, 3, x_29); +lean_ctor_set(x_47, 4, x_30); +lean_ctor_set(x_47, 5, x_31); +lean_ctor_set(x_47, 6, x_33); +lean_ctor_set(x_47, 7, x_34); +lean_ctor_set(x_47, 8, x_35); +lean_ctor_set(x_47, 9, x_36); +lean_ctor_set(x_47, 10, x_37); +lean_ctor_set(x_47, 11, x_38); +lean_ctor_set(x_47, 12, x_39); +lean_ctor_set(x_47, 13, x_40); +lean_ctor_set(x_47, 14, x_41); +lean_ctor_set(x_47, 15, x_42); +lean_ctor_set(x_47, 16, x_12); +lean_ctor_set(x_47, 17, x_44); +lean_ctor_set(x_47, 18, x_45); +lean_ctor_set(x_47, 19, x_46); +lean_ctor_set_uint8(x_47, sizeof(void*)*20, x_32); +x_48 = lean_st_ref_set(x_3, x_47, x_25); +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_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___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) { -_start: -{ -if (lean_obj_tag(x_2) == 5) -{ -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_14 = lean_ctor_get(x_2, 0); -lean_inc(x_14); -x_15 = lean_ctor_get(x_2, 1); -lean_inc(x_15); -lean_dec(x_2); -x_16 = lean_array_set(x_3, x_4, x_15); -x_17 = lean_unsigned_to_nat(1u); -x_18 = lean_nat_sub(x_4, x_17); -lean_dec(x_4); -x_2 = x_14; -x_3 = x_16; -x_4 = x_18; -goto _start; +lean_ctor_set(x_52, 0, x_51); +lean_ctor_set(x_52, 1, x_49); +return x_52; } -else -{ -size_t x_20; size_t x_21; lean_object* x_22; -lean_dec(x_4); -x_20 = lean_array_size(x_3); -x_21 = 0; -x_22 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___spec__1(x_1, x_20, x_21, x_3, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -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; lean_object* x_25; -x_24 = lean_ctor_get(x_22, 0); -x_25 = l_Lean_mkAppN(x_2, x_24); -lean_dec(x_24); -lean_ctor_set(x_22, 0, x_25); -return x_22; } else { -lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_26 = lean_ctor_get(x_22, 0); -x_27 = lean_ctor_get(x_22, 1); -lean_inc(x_27); -lean_inc(x_26); -lean_dec(x_22); -x_28 = l_Lean_mkAppN(x_2, x_26); -lean_dec(x_26); -x_29 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_29, 0, x_28); -lean_ctor_set(x_29, 1, x_27); -return x_29; -} +lean_object* x_53; lean_object* x_54; 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; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; +x_53 = lean_ctor_get(x_12, 0); +x_54 = lean_ctor_get(x_12, 1); +lean_inc(x_54); +lean_inc(x_53); +lean_dec(x_12); +x_55 = lean_ctor_get(x_53, 0); +lean_inc(x_55); +x_56 = lean_ctor_get(x_53, 1); +lean_inc(x_56); +x_57 = lean_ctor_get(x_53, 2); +lean_inc(x_57); +x_58 = lean_ctor_get(x_53, 3); +lean_inc(x_58); +x_59 = lean_ctor_get(x_53, 4); +lean_inc(x_59); +x_60 = lean_ctor_get(x_53, 5); +lean_inc(x_60); +x_61 = lean_ctor_get_uint8(x_53, sizeof(void*)*20); +x_62 = lean_ctor_get(x_53, 6); +lean_inc(x_62); +x_63 = lean_ctor_get(x_53, 7); +lean_inc(x_63); +x_64 = lean_ctor_get(x_53, 8); +lean_inc(x_64); +x_65 = lean_ctor_get(x_53, 9); +lean_inc(x_65); +x_66 = lean_ctor_get(x_53, 10); +lean_inc(x_66); +x_67 = lean_ctor_get(x_53, 11); +lean_inc(x_67); +x_68 = lean_ctor_get(x_53, 12); +lean_inc(x_68); +x_69 = lean_ctor_get(x_53, 13); +lean_inc(x_69); +x_70 = lean_ctor_get(x_53, 14); +lean_inc(x_70); +x_71 = lean_ctor_get(x_53, 15); +lean_inc(x_71); +x_72 = lean_ctor_get(x_53, 16); +lean_inc(x_72); +x_73 = lean_ctor_get(x_53, 17); +lean_inc(x_73); +x_74 = lean_ctor_get(x_53, 18); +lean_inc(x_74); +x_75 = lean_ctor_get(x_53, 19); +lean_inc(x_75); +if (lean_is_exclusive(x_53)) { + lean_ctor_release(x_53, 0); + lean_ctor_release(x_53, 1); + lean_ctor_release(x_53, 2); + lean_ctor_release(x_53, 3); + lean_ctor_release(x_53, 4); + lean_ctor_release(x_53, 5); + lean_ctor_release(x_53, 6); + lean_ctor_release(x_53, 7); + lean_ctor_release(x_53, 8); + lean_ctor_release(x_53, 9); + lean_ctor_release(x_53, 10); + lean_ctor_release(x_53, 11); + lean_ctor_release(x_53, 12); + lean_ctor_release(x_53, 13); + lean_ctor_release(x_53, 14); + lean_ctor_release(x_53, 15); + lean_ctor_release(x_53, 16); + lean_ctor_release(x_53, 17); + lean_ctor_release(x_53, 18); + lean_ctor_release(x_53, 19); + x_76 = x_53; +} else { + lean_dec_ref(x_53); + x_76 = lean_box(0); } -else -{ -uint8_t x_30; -lean_dec(x_2); -x_30 = !lean_is_exclusive(x_22); -if (x_30 == 0) -{ -return x_22; +x_77 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_77, 0, x_1); +lean_ctor_set(x_77, 1, x_72); +if (lean_is_scalar(x_76)) { + x_78 = lean_alloc_ctor(0, 20, 1); +} else { + x_78 = x_76; } -else -{ -lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_31 = lean_ctor_get(x_22, 0); -x_32 = lean_ctor_get(x_22, 1); -lean_inc(x_32); -lean_inc(x_31); -lean_dec(x_22); -x_33 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_33, 0, x_31); -lean_ctor_set(x_33, 1, x_32); -return x_33; +lean_ctor_set(x_78, 0, x_55); +lean_ctor_set(x_78, 1, x_56); +lean_ctor_set(x_78, 2, x_57); +lean_ctor_set(x_78, 3, x_58); +lean_ctor_set(x_78, 4, x_59); +lean_ctor_set(x_78, 5, x_60); +lean_ctor_set(x_78, 6, x_62); +lean_ctor_set(x_78, 7, x_63); +lean_ctor_set(x_78, 8, x_64); +lean_ctor_set(x_78, 9, x_65); +lean_ctor_set(x_78, 10, x_66); +lean_ctor_set(x_78, 11, x_67); +lean_ctor_set(x_78, 12, x_68); +lean_ctor_set(x_78, 13, x_69); +lean_ctor_set(x_78, 14, x_70); +lean_ctor_set(x_78, 15, x_71); +lean_ctor_set(x_78, 16, x_77); +lean_ctor_set(x_78, 17, x_73); +lean_ctor_set(x_78, 18, x_74); +lean_ctor_set(x_78, 19, x_75); +lean_ctor_set_uint8(x_78, sizeof(void*)*20, x_61); +x_79 = lean_st_ref_set(x_3, x_78, x_54); +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); } +x_82 = lean_box(0); +if (lean_is_scalar(x_81)) { + x_83 = lean_alloc_ctor(0, 2, 0); +} else { + x_83 = x_81; } +lean_ctor_set(x_83, 0, x_82); +lean_ctor_set(x_83, 1, x_80); +return x_83; } } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__1() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addSplitCandidate___closed__1() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_levelZero; -x_2 = l_Lean_Expr_sort___override(x_1); -return x_2; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("split", 5, 5); +return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__2() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addSplitCandidate___closed__2() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_unfoldReducible___lambda__2), 6, 0); +x_1 = lean_mk_string_unchecked("candidate", 9, 9); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__3() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addSplitCandidate___closed__3() { _start: { -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_unfoldReducible___lambda__3___boxed), 6, 0); -return x_1; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addSplitCandidate___closed__1; +x_3 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addSplitCandidate___closed__2; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__4() { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addSplitCandidate(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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_alloc_closure((void*)(l_Lean_Meta_Grind_normalizeLevels___lambda__1___boxed), 4, 0); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__5() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_eraseIrrelevantMData___lambda__2___boxed), 4, 0); -return x_1; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, 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 = l_Lean_Expr_isBVar(x_1); -if (x_12 == 0) -{ -lean_object* x_13; uint8_t x_14; -x_13 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare; -x_14 = lean_expr_eqv(x_1, x_13); +lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_11 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addSplitCandidate___closed__3; +x_12 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_11, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_unbox(x_13); +lean_dec(x_13); if (x_14 == 0) { -lean_object* x_15; -x_15 = l_Lean_Meta_Grind_groundPattern_x3f(x_1); -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; -x_16 = lean_unsigned_to_nat(0u); -x_17 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_16); -x_18 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__1; -lean_inc(x_17); -x_19 = lean_mk_array(x_17, x_18); -x_20 = lean_unsigned_to_nat(1u); -x_21 = lean_nat_sub(x_17, x_20); -lean_dec(x_17); -x_22 = l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___spec__2(x_2, x_1, x_19, x_21, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_22; +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_ctor_get(x_12, 1); +lean_inc(x_15); +lean_dec(x_12); +x_16 = lean_box(0); +x_17 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addSplitCandidate___lambda__1(x_1, x_16, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_15); +return x_17; } else { -lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; -lean_dec(x_1); -x_23 = lean_ctor_get(x_15, 0); -lean_inc(x_23); -lean_dec(x_15); -x_24 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__2; -x_25 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__3; -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -x_26 = l_Lean_Core_transform___at_Lean_Meta_Grind_unfoldReducible___spec__1(x_23, x_24, x_25, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_26) == 0) +uint8_t x_18; +x_18 = !lean_is_exclusive(x_12); +if (x_18 == 0) { -lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_12, 1); +x_20 = lean_ctor_get(x_12, 0); +lean_dec(x_20); +x_21 = l_Lean_Meta_Grind_updateLastTag(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_19); +if (lean_obj_tag(x_21) == 0) +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_22 = lean_ctor_get(x_21, 1); +lean_inc(x_22); +lean_dec(x_21); +lean_inc(x_1); +x_23 = l_Lean_MessageData_ofExpr(x_1); +x_24 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__6; +lean_ctor_set_tag(x_12, 7); +lean_ctor_set(x_12, 1, x_23); +lean_ctor_set(x_12, 0, x_24); +x_25 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_25, 0, x_12); +lean_ctor_set(x_25, 1, x_24); +x_26 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_11, x_25, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, 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 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__4; -x_30 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__5; -lean_inc(x_10); -lean_inc(x_9); -x_31 = l_Lean_Core_transform___at_Lean_Core_betaReduce___spec__1(x_27, x_29, x_30, x_9, x_10, x_28); -if (lean_obj_tag(x_31) == 0) -{ -lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_32 = lean_ctor_get(x_31, 0); -lean_inc(x_32); -x_33 = lean_ctor_get(x_31, 1); -lean_inc(x_33); -lean_dec(x_31); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -x_34 = l_Lean_Meta_Grind_canon(x_32, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_33); -if (lean_obj_tag(x_34) == 0) -{ -lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_35 = lean_ctor_get(x_34, 0); -lean_inc(x_35); -x_36 = lean_ctor_get(x_34, 1); -lean_inc(x_36); -lean_dec(x_34); -x_37 = l_Lean_Meta_Grind_shareCommon(x_35, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_36); -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); -lean_inc(x_38); -x_40 = l_Lean_Meta_Grind_internalize(x_38, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_39); -if (lean_obj_tag(x_40) == 0) +x_29 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addSplitCandidate___lambda__1(x_1, x_27, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_28); +lean_dec(x_27); +return x_29; +} +else { -uint8_t x_41; -x_41 = !lean_is_exclusive(x_40); -if (x_41 == 0) +uint8_t x_30; +lean_free_object(x_12); +lean_dec(x_1); +x_30 = !lean_is_exclusive(x_21); +if (x_30 == 0) { -lean_object* x_42; lean_object* x_43; -x_42 = lean_ctor_get(x_40, 0); -lean_dec(x_42); -x_43 = l_Lean_Meta_Grind_mkGroundPattern(x_38); -lean_ctor_set(x_40, 0, x_43); -return x_40; +return x_21; } 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 = l_Lean_Meta_Grind_mkGroundPattern(x_38); -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_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_21, 0); +x_32 = lean_ctor_get(x_21, 1); +lean_inc(x_32); +lean_inc(x_31); +lean_dec(x_21); +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_47; -lean_dec(x_38); -x_47 = !lean_is_exclusive(x_40); -if (x_47 == 0) +lean_object* x_34; lean_object* x_35; +x_34 = lean_ctor_get(x_12, 1); +lean_inc(x_34); +lean_dec(x_12); +x_35 = l_Lean_Meta_Grind_updateLastTag(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_34); +if (lean_obj_tag(x_35) == 0) { -return x_40; +lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_36 = lean_ctor_get(x_35, 1); +lean_inc(x_36); +lean_dec(x_35); +lean_inc(x_1); +x_37 = l_Lean_MessageData_ofExpr(x_1); +x_38 = l_Lean_Meta_Grind_addCongrTable___lambda__2___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_alloc_ctor(7, 2, 0); +lean_ctor_set(x_40, 0, x_39); +lean_ctor_set(x_40, 1, x_38); +x_41 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_11, x_40, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_36); +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___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addSplitCandidate___lambda__1(x_1, x_42, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_43); +lean_dec(x_42); +return x_44; } else { -lean_object* x_48; lean_object* x_49; lean_object* x_50; -x_48 = lean_ctor_get(x_40, 0); -x_49 = lean_ctor_get(x_40, 1); -lean_inc(x_49); -lean_inc(x_48); -lean_dec(x_40); -x_50 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_50, 0, x_48); -lean_ctor_set(x_50, 1, x_49); -return x_50; +lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; +lean_dec(x_1); +x_45 = lean_ctor_get(x_35, 0); +lean_inc(x_45); +x_46 = lean_ctor_get(x_35, 1); +lean_inc(x_46); +if (lean_is_exclusive(x_35)) { + lean_ctor_release(x_35, 0); + lean_ctor_release(x_35, 1); + x_47 = x_35; +} else { + lean_dec_ref(x_35); + x_47 = lean_box(0); } +if (lean_is_scalar(x_47)) { + x_48 = lean_alloc_ctor(1, 2, 0); +} else { + x_48 = x_47; } +lean_ctor_set(x_48, 0, x_45); +lean_ctor_set(x_48, 1, x_46); +return x_48; } -else -{ -uint8_t x_51; -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_51 = !lean_is_exclusive(x_34); -if (x_51 == 0) -{ -return x_34; } -else -{ -lean_object* x_52; lean_object* x_53; lean_object* x_54; -x_52 = lean_ctor_get(x_34, 0); -x_53 = lean_ctor_get(x_34, 1); -lean_inc(x_53); -lean_inc(x_52); -lean_dec(x_34); -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_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addSplitCandidate___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_55; +lean_object* x_12; +x_12 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addSplitCandidate___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -4379,30 +4483,14 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_55 = !lean_is_exclusive(x_31); -if (x_55 == 0) -{ -return x_31; -} -else -{ -lean_object* x_56; lean_object* x_57; lean_object* x_58; -x_56 = lean_ctor_get(x_31, 0); -x_57 = lean_ctor_get(x_31, 1); -lean_inc(x_57); -lean_inc(x_56); -lean_dec(x_31); -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; -} +return x_12; } } -else +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addSplitCandidate___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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_59; -lean_dec(x_10); +lean_object* x_11; +x_11 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addSplitCandidate(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); @@ -4411,2271 +4499,1529 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_59 = !lean_is_exclusive(x_26); -if (x_59 == 0) +return x_11; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__1() { +_start: { -return x_26; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Eq", 2, 2); +return x_1; } -else +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__2() { +_start: { -lean_object* x_60; lean_object* x_61; lean_object* x_62; -x_60 = lean_ctor_get(x_26, 0); -x_61 = lean_ctor_get(x_26, 1); -lean_inc(x_61); -lean_inc(x_60); -lean_dec(x_26); -x_62 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_62, 0, x_60); -lean_ctor_set(x_62, 1, x_61); -return x_62; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } } +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("HEq", 3, 3); +return x_1; } } -else +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__4() { +_start: { -lean_object* x_63; -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_63 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_63, 0, x_1); -lean_ctor_set(x_63, 1, x_11); -return x_63; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__3; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } } -else +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__5() { +_start: { -lean_object* x_64; -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_64 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_64, 0, x_1); -lean_ctor_set(x_64, 1, x_11); -return x_64; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("True", 4, 4); +return x_1; } } +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___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_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__5; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} } -static lean_object* _init_l_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__1() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__7() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_instMonadMetaM; -x_2 = l_ReaderT_instMonad___rarg(x_1); -return x_2; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("False", 5, 5); +return x_1; } } -static lean_object* _init_l_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__2() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__8() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__1; -x_2 = l_ReaderT_instMonad___rarg(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___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__7; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } } -static lean_object* _init_l_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__3() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__9() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__2; -x_2 = l_ReaderT_instMonad___rarg(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___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___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_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__4() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__10() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__3; -x_2 = l_ReaderT_instMonad___rarg(x_1); -return x_2; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__6; +x_2 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___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_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__5() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__4; -x_2 = l_instInhabitedPUnit; -x_3 = l_instInhabitedOfMonad___rarg(x_1, x_2); +x_1 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__4; +x_2 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__10; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_panic___at_Lean_Meta_Grind_internalize___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) { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__12() { _start: { -lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_11 = l_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__5; -x_12 = lean_panic_fn(x_11, x_1); -x_13 = lean_apply_9(x_12, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_13; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__2; +x_2 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__11; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_internalize___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18) { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes() { _start: { -lean_object* x_19; uint8_t x_20; -x_19 = lean_ctor_get(x_5, 1); -x_20 = lean_nat_dec_lt(x_7, x_19); -if (x_20 == 0) +lean_object* x_1; +x_1 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__12; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_isMatcherApp___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate___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_21; +lean_object* x_11; uint8_t x_12; +x_11 = lean_st_ref_get(x_9, x_10); +x_12 = !lean_is_exclusive(x_11); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; +x_13 = lean_ctor_get(x_11, 0); +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +lean_dec(x_13); +x_15 = l_Lean_Meta_isMatcherAppCore(x_14, x_1); +lean_dec(x_14); +x_16 = lean_box(x_15); +lean_ctor_set(x_11, 0, x_16); +return x_11; +} +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; lean_object* x_21; lean_object* x_22; +x_17 = lean_ctor_get(x_11, 0); +x_18 = lean_ctor_get(x_11, 1); +lean_inc(x_18); +lean_inc(x_17); +lean_dec(x_11); +x_19 = lean_ctor_get(x_17, 0); +lean_inc(x_19); lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_15); +x_20 = l_Lean_Meta_isMatcherAppCore(x_19, x_1); +lean_dec(x_19); +x_21 = lean_box(x_20); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_18); +return x_22; +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_13 = l_Lean_Meta_Grind_getConfig___rarg(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_uint8(x_14, sizeof(void*)*4 + 3); lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); +if (x_15 == 0) +{ +uint8_t x_16; lean_dec(x_11); lean_dec(x_10); -lean_dec(x_7); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_2); lean_dec(x_1); +x_16 = !lean_is_exclusive(x_13); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; +x_17 = lean_ctor_get(x_13, 0); +lean_dec(x_17); +x_18 = lean_box(0); +lean_ctor_set(x_13, 0, x_18); +return x_13; +} +else +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_13, 1); +lean_inc(x_19); +lean_dec(x_13); +x_20 = lean_box(0); x_21 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_21, 0, x_6); -lean_ctor_set(x_21, 1, x_18); +lean_ctor_set(x_21, 0, x_20); +lean_ctor_set(x_21, 1, x_19); return x_21; } +} else { lean_object* x_22; lean_object* x_23; -lean_dec(x_6); -x_22 = lean_array_fget(x_3, x_7); -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_22 = lean_ctor_get(x_13, 1); +lean_inc(x_22); +lean_dec(x_13); lean_inc(x_11); lean_inc(x_10); -lean_inc(x_2); -lean_inc(x_22); -x_23 = l_Lean_Meta_Grind_internalize(x_22, x_2, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +lean_inc(x_9); +lean_inc(x_8); +x_23 = l_Lean_Meta_isInductivePredicate(x_1, x_8, x_9, x_10, x_11, 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; -x_24 = lean_ctor_get(x_23, 1); +lean_object* x_24; uint8_t x_25; +x_24 = lean_ctor_get(x_23, 0); lean_inc(x_24); +x_25 = lean_unbox(x_24); +lean_dec(x_24); +if (x_25 == 0) +{ +uint8_t x_26; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_2); +x_26 = !lean_is_exclusive(x_23); +if (x_26 == 0) +{ +lean_object* x_27; lean_object* x_28; +x_27 = lean_ctor_get(x_23, 0); +lean_dec(x_27); +x_28 = lean_box(0); +lean_ctor_set(x_23, 0, x_28); +return x_23; +} +else +{ +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); -lean_inc(x_1); -x_25 = l_Lean_Meta_Grind_registerParent(x_1, x_22, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_24); -lean_dec(x_22); -x_26 = lean_ctor_get(x_25, 1); -lean_inc(x_26); -lean_dec(x_25); -x_27 = lean_ctor_get(x_5, 2); -x_28 = lean_nat_add(x_7, x_27); -lean_dec(x_7); -x_29 = lean_box(0); -x_6 = x_29; -x_7 = x_28; -x_8 = lean_box(0); -x_9 = lean_box(0); -x_18 = x_26; -goto _start; +x_30 = lean_box(0); +x_31 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_31, 0, x_30); +lean_ctor_set(x_31, 1, x_29); +return x_31; +} } else { -uint8_t x_31; -lean_dec(x_22); -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); +lean_object* x_32; lean_object* x_33; +x_32 = lean_ctor_get(x_23, 1); +lean_inc(x_32); +lean_dec(x_23); +x_33 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addSplitCandidate(x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_32); lean_dec(x_11); lean_dec(x_10); -lean_dec(x_7); +lean_dec(x_9); +lean_dec(x_8); +return x_33; +} +} +else +{ +uint8_t x_34; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_2); -lean_dec(x_1); -x_31 = !lean_is_exclusive(x_23); -if (x_31 == 0) +x_34 = !lean_is_exclusive(x_23); +if (x_34 == 0) { return x_23; } else { -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_object* x_35; lean_object* x_36; lean_object* x_37; +x_35 = lean_ctor_get(x_23, 0); +x_36 = lean_ctor_get(x_23, 1); +lean_inc(x_36); +lean_inc(x_35); 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); -return x_34; +x_37 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_37, 0, x_35); +lean_ctor_set(x_37, 1, x_36); +return x_37; } } } } } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___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_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate___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_13; -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_1); -x_13 = l_Lean_Meta_Grind_mkENode(x_1, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_13) == 0) +lean_object* x_12; +x_12 = l_Lean_Expr_getAppFn(x_1); +if (lean_obj_tag(x_12) == 4) { -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_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_1); -x_15 = l_Lean_Meta_Grind_addCongrTable(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_14); -if (lean_obj_tag(x_15) == 0) +lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +lean_dec(x_12); +x_14 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes; +x_15 = l_List_elem___at_Lean_addAliasEntry___spec__16(x_13, x_14); +if (x_15 == 0) { -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_16 = lean_ctor_get(x_15, 1); -lean_inc(x_16); -lean_dec(x_15); -lean_inc(x_1); -x_17 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_16); -x_18 = lean_ctor_get(x_17, 1); -lean_inc(x_18); -lean_dec(x_17); -x_19 = l_Lean_Meta_Grind_propagateUp(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_18); -return x_19; +lean_object* x_16; lean_object* x_17; +x_16 = lean_box(0); +x_17 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate___lambda__1(x_13, x_1, x_16, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +return x_17; } else { -uint8_t x_20; -lean_dec(x_11); +lean_object* x_18; lean_object* x_19; +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_20 = !lean_is_exclusive(x_15); -if (x_20 == 0) -{ -return x_15; -} -else -{ -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_15, 0); -x_22 = lean_ctor_get(x_15, 1); -lean_inc(x_22); -lean_inc(x_21); -lean_dec(x_15); -x_23 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_23, 0, x_21); -lean_ctor_set(x_23, 1, x_22); -return x_23; -} +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_11); +return x_19; } } else { -uint8_t x_24; -lean_dec(x_11); +lean_object* x_20; lean_object* x_21; +lean_dec(x_12); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); lean_dec(x_1); -x_24 = !lean_is_exclusive(x_13); -if (x_24 == 0) -{ -return x_13; -} -else -{ -lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_25 = lean_ctor_get(x_13, 0); -x_26 = lean_ctor_get(x_13, 1); -lean_inc(x_26); -lean_inc(x_25); -lean_dec(x_13); -x_27 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_27, 0, x_25); -lean_ctor_set(x_27, 1, x_26); -return x_27; -} +x_20 = lean_box(0); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_20); +lean_ctor_set(x_21, 1, x_11); +return x_21; } } } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___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_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate___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_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_15 = lean_array_get_size(x_1); -x_16 = lean_unsigned_to_nat(0u); -x_17 = lean_unsigned_to_nat(1u); -x_18 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_18, 0, x_16); -lean_ctor_set(x_18, 1, x_15); -lean_ctor_set(x_18, 2, x_17); -x_19 = lean_box(0); +lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_12 = l_Lean_Meta_Grind_getConfig___rarg(x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_13 = lean_ctor_get(x_12, 0); lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); +x_14 = lean_ctor_get_uint8(x_13, sizeof(void*)*4 + 1); +lean_dec(x_13); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_ctor_get(x_12, 1); +lean_inc(x_15); +lean_dec(x_12); +x_16 = lean_box(0); +x_17 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate___lambda__2(x_1, x_16, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_15); +return x_17; +} +else +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_18 = lean_ctor_get(x_12, 1); +lean_inc(x_18); +lean_dec(x_12); +x_19 = l_Lean_Meta_isMatcherApp___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate___spec__1(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_18); +x_20 = lean_ctor_get(x_19, 0); +lean_inc(x_20); +x_21 = lean_unbox(x_20); +lean_dec(x_20); +if (x_21 == 0) +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_22 = lean_ctor_get(x_19, 1); +lean_inc(x_22); +lean_dec(x_19); +x_23 = lean_box(0); +x_24 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate___lambda__2(x_1, x_23, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_22); +return x_24; +} +else +{ +lean_object* x_25; lean_object* x_26; +x_25 = lean_ctor_get(x_19, 1); +lean_inc(x_25); +lean_dec(x_19); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); -lean_inc(x_6); -x_20 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_internalize___spec__2(x_2, x_3, x_1, x_18, x_18, x_19, x_16, lean_box(0), lean_box(0), x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -lean_dec(x_18); -if (lean_obj_tag(x_20) == 0) +lean_inc(x_1); +x_26 = l_Lean_Meta_reduceMatcher_x3f(x_1, x_7, x_8, x_9, x_10, x_25); +if (lean_obj_tag(x_26) == 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 = lean_apply_10(x_4, x_19, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_21); -return x_22; -} -else +lean_object* x_27; +x_27 = lean_ctor_get(x_26, 0); +lean_inc(x_27); +if (lean_obj_tag(x_27) == 0) { -uint8_t x_23; -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); +uint8_t x_28; +lean_dec(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_4); -x_23 = !lean_is_exclusive(x_20); -if (x_23 == 0) +lean_dec(x_1); +x_28 = !lean_is_exclusive(x_26); +if (x_28 == 0) { -return x_20; +lean_object* x_29; lean_object* x_30; +x_29 = lean_ctor_get(x_26, 0); +lean_dec(x_29); +x_30 = lean_box(0); +lean_ctor_set(x_26, 0, x_30); +return x_26; } else { -lean_object* x_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_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_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_26, 1); +lean_inc(x_31); +lean_dec(x_26); +x_32 = lean_box(0); +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_31); +return x_33; } } -static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__1() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean", 4, 4); -return x_1; -} -} -static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__2() { -_start: +lean_object* x_34; lean_object* x_35; +lean_dec(x_27); +x_34 = lean_ctor_get(x_26, 1); +lean_inc(x_34); +lean_dec(x_26); +x_35 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addSplitCandidate(x_1, x_3, x_4, x_5, x_6, 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_35) == 0) { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Grind", 5, 5); -return x_1; -} -} -static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__3() { -_start: +uint8_t x_36; +x_36 = !lean_is_exclusive(x_35); +if (x_36 == 0) { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("nestedProof", 11, 11); -return x_1; -} +lean_object* x_37; lean_object* x_38; +x_37 = lean_ctor_get(x_35, 0); +lean_dec(x_37); +x_38 = lean_box(0); +lean_ctor_set(x_35, 0, x_38); +return x_35; } -static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__4() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__1; -x_2 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__2; -x_3 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__3; -x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); -return x_4; -} +lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_39 = lean_ctor_get(x_35, 1); +lean_inc(x_39); +lean_dec(x_35); +x_40 = lean_box(0); +x_41 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_41, 0, x_40); +lean_ctor_set(x_41, 1, x_39); +return x_41; } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { -_start: -{ -switch (lean_obj_tag(x_3)) { -case 0: -{ -lean_object* x_15; lean_object* x_16; lean_object* x_28; uint8_t x_29; -lean_dec(x_5); -lean_inc(x_2); -lean_inc(x_1); -x_15 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1___boxed), 12, 2); -lean_closure_set(x_15, 0, x_1); -lean_closure_set(x_15, 1, x_2); -x_28 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__4; -x_29 = l_Lean_Expr_isConstOf(x_3, x_28); -if (x_29 == 0) -{ -lean_object* x_30; -x_30 = lean_box(0); -x_16 = x_30; -goto block_27; } else { -lean_object* x_31; lean_object* x_32; uint8_t x_33; -x_31 = lean_array_get_size(x_4); -x_32 = lean_unsigned_to_nat(2u); -x_33 = lean_nat_dec_eq(x_31, x_32); -if (x_33 == 0) +uint8_t x_42; +x_42 = !lean_is_exclusive(x_35); +if (x_42 == 0) { -lean_object* x_34; -lean_dec(x_31); -x_34 = lean_box(0); -x_16 = x_34; -goto block_27; +return x_35; } else { -lean_object* x_35; uint8_t x_36; -lean_dec(x_15); -lean_dec(x_3); -x_35 = lean_unsigned_to_nat(0u); -x_36 = lean_nat_dec_lt(x_35, x_31); -lean_dec(x_31); -if (x_36 == 0) -{ -lean_object* x_37; lean_object* x_38; lean_object* x_39; -lean_dec(x_4); -x_37 = l_Lean_instInhabitedExpr; -x_38 = l_outOfBounds___rarg(x_37); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_2); -lean_inc(x_38); -x_39 = l_Lean_Meta_Grind_internalize(x_38, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_39) == 0) -{ -lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; -x_40 = lean_ctor_get(x_39, 1); -lean_inc(x_40); -lean_dec(x_39); -lean_inc(x_1); -x_41 = l_Lean_Meta_Grind_registerParent(x_1, x_38, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_40); -lean_dec(x_38); -x_42 = lean_ctor_get(x_41, 0); -lean_inc(x_42); -x_43 = lean_ctor_get(x_41, 1); +lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_43 = lean_ctor_get(x_35, 0); +x_44 = lean_ctor_get(x_35, 1); +lean_inc(x_44); lean_inc(x_43); -lean_dec(x_41); -x_44 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_42, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_43); -lean_dec(x_42); -return x_44; +lean_dec(x_35); +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; -lean_dec(x_38); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); +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_2); lean_dec(x_1); -x_45 = !lean_is_exclusive(x_39); -if (x_45 == 0) +x_46 = !lean_is_exclusive(x_26); +if (x_46 == 0) { -return x_39; +return x_26; } 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_object* x_47; lean_object* x_48; lean_object* x_49; +x_47 = lean_ctor_get(x_26, 0); +x_48 = lean_ctor_get(x_26, 1); +lean_inc(x_48); lean_inc(x_47); -lean_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_dec(x_26); +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_49; lean_object* x_50; -x_49 = lean_array_fget(x_4, x_35); -lean_dec(x_4); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_2); -lean_inc(x_49); -x_50 = l_Lean_Meta_Grind_internalize(x_49, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_50) == 0) -{ -lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; -x_51 = lean_ctor_get(x_50, 1); -lean_inc(x_51); -lean_dec(x_50); -lean_inc(x_1); -x_52 = l_Lean_Meta_Grind_registerParent(x_1, x_49, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_51); -lean_dec(x_49); -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_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_53, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_54); -lean_dec(x_53); -return x_55; } -else +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate___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_56; -lean_dec(x_49); +lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_12 = l_Lean_Meta_Grind_getConfig___rarg(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_ctor_get_uint8(x_13, sizeof(void*)*4 + 2); lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_2); -lean_dec(x_1); -x_56 = !lean_is_exclusive(x_50); -if (x_56 == 0) +if (x_14 == 0) { -return x_50; +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_ctor_get(x_12, 1); +lean_inc(x_15); +lean_dec(x_12); +x_16 = lean_box(0); +x_17 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate___lambda__3(x_1, x_16, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_15); +return x_17; } else { -lean_object* x_57; lean_object* x_58; lean_object* x_59; -x_57 = lean_ctor_get(x_50, 0); -x_58 = lean_ctor_get(x_50, 1); -lean_inc(x_58); -lean_inc(x_57); -lean_dec(x_50); -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; -} -} -} -} -} -block_27: +lean_object* x_18; uint8_t x_19; +x_18 = lean_ctor_get(x_12, 1); +lean_inc(x_18); +lean_dec(x_12); +x_19 = l_Lean_Expr_isIte(x_1); +if (x_19 == 0) { -lean_object* x_17; -lean_dec(x_16); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_2); -lean_inc(x_3); -x_17 = l_Lean_Meta_Grind_internalize(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_17) == 0) +uint8_t x_20; +x_20 = l_Lean_Expr_isDIte(x_1); +if (x_20 == 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, 1); -lean_inc(x_18); -lean_dec(x_17); -lean_inc(x_1); -x_19 = l_Lean_Meta_Grind_registerParent(x_1, x_3, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_18); -lean_dec(x_3); -x_20 = lean_ctor_get(x_19, 0); -lean_inc(x_20); -x_21 = lean_ctor_get(x_19, 1); -lean_inc(x_21); -lean_dec(x_19); -x_22 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(x_4, x_1, x_2, x_15, x_20, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_21); -lean_dec(x_20); -lean_dec(x_4); +lean_object* x_21; lean_object* x_22; +x_21 = lean_box(0); +x_22 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate___lambda__3(x_1, x_21, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_18); return x_22; } else { -uint8_t x_23; -lean_dec(x_15); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); +lean_object* x_23; +x_23 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addSplitCandidate(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_18); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_23 = !lean_is_exclusive(x_17); -if (x_23 == 0) +if (lean_obj_tag(x_23) == 0) { -return x_17; +uint8_t x_24; +x_24 = !lean_is_exclusive(x_23); +if (x_24 == 0) +{ +lean_object* x_25; lean_object* x_26; +x_25 = lean_ctor_get(x_23, 0); +lean_dec(x_25); +x_26 = lean_box(0); +lean_ctor_set(x_23, 0, x_26); +return x_23; } else { -lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_24 = lean_ctor_get(x_17, 0); -x_25 = lean_ctor_get(x_17, 1); -lean_inc(x_25); -lean_inc(x_24); -lean_dec(x_17); -x_26 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_26, 0, x_24); -lean_ctor_set(x_26, 1, x_25); -return x_26; -} -} -} +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_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; } -case 1: -{ -lean_object* x_60; lean_object* x_61; lean_object* x_73; uint8_t x_74; -lean_dec(x_5); -lean_inc(x_2); -lean_inc(x_1); -x_60 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1___boxed), 12, 2); -lean_closure_set(x_60, 0, x_1); -lean_closure_set(x_60, 1, x_2); -x_73 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__4; -x_74 = l_Lean_Expr_isConstOf(x_3, x_73); -if (x_74 == 0) -{ -lean_object* x_75; -x_75 = lean_box(0); -x_61 = x_75; -goto block_72; } else { -lean_object* x_76; lean_object* x_77; uint8_t x_78; -x_76 = lean_array_get_size(x_4); -x_77 = lean_unsigned_to_nat(2u); -x_78 = lean_nat_dec_eq(x_76, x_77); -if (x_78 == 0) +uint8_t x_30; +x_30 = !lean_is_exclusive(x_23); +if (x_30 == 0) { -lean_object* x_79; -lean_dec(x_76); -x_79 = lean_box(0); -x_61 = x_79; -goto block_72; +return x_23; } else { -lean_object* x_80; uint8_t x_81; -lean_dec(x_60); -lean_dec(x_3); -x_80 = lean_unsigned_to_nat(0u); -x_81 = lean_nat_dec_lt(x_80, x_76); -lean_dec(x_76); -if (x_81 == 0) -{ -lean_object* x_82; lean_object* x_83; lean_object* x_84; -lean_dec(x_4); -x_82 = l_Lean_instInhabitedExpr; -x_83 = l_outOfBounds___rarg(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); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_2); -lean_inc(x_83); -x_84 = l_Lean_Meta_Grind_internalize(x_83, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_84) == 0) -{ -lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; -x_85 = lean_ctor_get(x_84, 1); -lean_inc(x_85); -lean_dec(x_84); -lean_inc(x_1); -x_86 = l_Lean_Meta_Grind_registerParent(x_1, x_83, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_85); -lean_dec(x_83); -x_87 = lean_ctor_get(x_86, 0); -lean_inc(x_87); -x_88 = lean_ctor_get(x_86, 1); -lean_inc(x_88); -lean_dec(x_86); -x_89 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_87, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_88); -lean_dec(x_87); -return x_89; +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 { -uint8_t x_90; -lean_dec(x_83); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); +lean_object* x_34; +x_34 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addSplitCandidate(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_18); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_2); -lean_dec(x_1); -x_90 = !lean_is_exclusive(x_84); -if (x_90 == 0) +if (lean_obj_tag(x_34) == 0) { -return x_84; +uint8_t x_35; +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_91; lean_object* x_92; lean_object* x_93; -x_91 = lean_ctor_get(x_84, 0); -x_92 = lean_ctor_get(x_84, 1); -lean_inc(x_92); -lean_inc(x_91); -lean_dec(x_84); -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_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_94; lean_object* x_95; -x_94 = lean_array_fget(x_4, x_80); -lean_dec(x_4); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_2); -lean_inc(x_94); -x_95 = l_Lean_Meta_Grind_internalize(x_94, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_95) == 0) +uint8_t x_41; +x_41 = !lean_is_exclusive(x_34); +if (x_41 == 0) { -lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; -x_96 = lean_ctor_get(x_95, 1); -lean_inc(x_96); -lean_dec(x_95); -lean_inc(x_1); -x_97 = l_Lean_Meta_Grind_registerParent(x_1, x_94, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_96); -lean_dec(x_94); -x_98 = lean_ctor_get(x_97, 0); -lean_inc(x_98); -x_99 = lean_ctor_get(x_97, 1); -lean_inc(x_99); -lean_dec(x_97); -x_100 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_98, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_99); -lean_dec(x_98); -return x_100; +return x_34; } else { -uint8_t x_101; -lean_dec(x_94); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); +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; +} +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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 = l_Lean_Expr_isApp(x_1); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; 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_101 = !lean_is_exclusive(x_95); -if (x_101 == 0) -{ -return x_95; +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_10); +return x_13; } else { -lean_object* x_102; lean_object* x_103; lean_object* x_104; -x_102 = lean_ctor_get(x_95, 0); -x_103 = lean_ctor_get(x_95, 1); -lean_inc(x_103); -lean_inc(x_102); -lean_dec(x_95); -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_14; lean_object* x_15; +x_14 = lean_box(0); +x_15 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate___lambda__4(x_1, x_14, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_15; } } -block_72: -{ -lean_object* x_62; -lean_dec(x_61); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_2); -lean_inc(x_3); -x_62 = l_Lean_Meta_Grind_internalize(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_62) == 0) -{ -lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; -x_63 = lean_ctor_get(x_62, 1); -lean_inc(x_63); -lean_dec(x_62); -lean_inc(x_1); -x_64 = l_Lean_Meta_Grind_registerParent(x_1, x_3, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_63); -lean_dec(x_3); -x_65 = lean_ctor_get(x_64, 0); -lean_inc(x_65); -x_66 = lean_ctor_get(x_64, 1); -lean_inc(x_66); -lean_dec(x_64); -x_67 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(x_4, x_1, x_2, x_60, x_65, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_66); -lean_dec(x_65); -lean_dec(x_4); -return x_67; } -else +LEAN_EXPORT lean_object* l_Lean_Meta_isMatcherApp___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate___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: { -uint8_t x_68; -lean_dec(x_60); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); +lean_object* x_11; +x_11 = l_Lean_Meta_isMatcherApp___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate___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_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_68 = !lean_is_exclusive(x_62); -if (x_68 == 0) -{ -return x_62; -} -else -{ -lean_object* x_69; lean_object* x_70; lean_object* x_71; -x_69 = lean_ctor_get(x_62, 0); -x_70 = lean_ctor_get(x_62, 1); -lean_inc(x_70); -lean_inc(x_69); -lean_dec(x_62); -x_71 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_71, 0, x_69); -lean_ctor_set(x_71, 1, x_70); -return x_71; -} -} +return x_11; } } -case 2: +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate___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_105; lean_object* x_106; lean_object* x_118; uint8_t x_119; +lean_object* x_13; +x_13 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate___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_7); +lean_dec(x_6); lean_dec(x_5); -lean_inc(x_2); -lean_inc(x_1); -x_105 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1___boxed), 12, 2); -lean_closure_set(x_105, 0, x_1); -lean_closure_set(x_105, 1, x_2); -x_118 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__4; -x_119 = l_Lean_Expr_isConstOf(x_3, x_118); -if (x_119 == 0) -{ -lean_object* x_120; -x_120 = lean_box(0); -x_106 = x_120; -goto block_117; +lean_dec(x_4); +lean_dec(x_3); +return x_13; } -else -{ -lean_object* x_121; lean_object* x_122; uint8_t x_123; -x_121 = lean_array_get_size(x_4); -x_122 = lean_unsigned_to_nat(2u); -x_123 = lean_nat_dec_eq(x_121, x_122); -if (x_123 == 0) -{ -lean_object* x_124; -lean_dec(x_121); -x_124 = lean_box(0); -x_106 = x_124; -goto block_117; } -else -{ -lean_object* x_125; uint8_t x_126; -lean_dec(x_105); -lean_dec(x_3); -x_125 = lean_unsigned_to_nat(0u); -x_126 = lean_nat_dec_lt(x_125, x_121); -lean_dec(x_121); -if (x_126 == 0) +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate___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_127; lean_object* x_128; lean_object* x_129; +lean_object* x_12; +x_12 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_6); +lean_dec(x_5); lean_dec(x_4); -x_127 = l_Lean_instInhabitedExpr; -x_128 = l_outOfBounds___rarg(x_127); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_2); -lean_inc(x_128); -x_129 = l_Lean_Meta_Grind_internalize(x_128, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -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; -x_130 = lean_ctor_get(x_129, 1); -lean_inc(x_130); -lean_dec(x_129); -lean_inc(x_1); -x_131 = l_Lean_Meta_Grind_registerParent(x_1, x_128, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_130); -lean_dec(x_128); -x_132 = lean_ctor_get(x_131, 0); -lean_inc(x_132); -x_133 = lean_ctor_get(x_131, 1); -lean_inc(x_133); -lean_dec(x_131); -x_134 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_132, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_133); -lean_dec(x_132); -return x_134; +lean_dec(x_3); +lean_dec(x_2); +return x_12; } -else +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate___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_135; -lean_dec(x_128); -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_object* x_12; +x_12 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate___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_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -x_135 = !lean_is_exclusive(x_129); -if (x_135 == 0) -{ -return x_129; -} -else -{ -lean_object* x_136; lean_object* x_137; lean_object* x_138; -x_136 = lean_ctor_get(x_129, 0); -x_137 = lean_ctor_get(x_129, 1); -lean_inc(x_137); -lean_inc(x_136); -lean_dec(x_129); -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; -} +return x_12; } } -else +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate___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_139; lean_object* x_140; -x_139 = lean_array_fget(x_4, x_125); +lean_object* x_12; +x_12 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_6); +lean_dec(x_5); lean_dec(x_4); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_2); -lean_inc(x_139); -x_140 = l_Lean_Meta_Grind_internalize(x_139, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_140) == 0) -{ -lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; -x_141 = lean_ctor_get(x_140, 1); -lean_inc(x_141); -lean_dec(x_140); -lean_inc(x_1); -x_142 = l_Lean_Meta_Grind_registerParent(x_1, x_139, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_141); -lean_dec(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 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_143, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_144); -lean_dec(x_143); -return x_145; +lean_dec(x_3); +lean_dec(x_2); +return x_12; } -else +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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_146; -lean_dec(x_139); -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_object* x_11; +x_11 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate(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); lean_dec(x_2); -lean_dec(x_1); -x_146 = !lean_is_exclusive(x_140); -if (x_146 == 0) +return x_11; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__1___closed__1() { +_start: { -return x_140; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean", 4, 4); +return x_1; } -else +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__1___closed__2() { +_start: { -lean_object* x_147; lean_object* x_148; lean_object* x_149; -x_147 = lean_ctor_get(x_140, 0); -x_148 = lean_ctor_get(x_140, 1); -lean_inc(x_148); -lean_inc(x_147); -lean_dec(x_140); -x_149 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_149, 0, x_147); -lean_ctor_set(x_149, 1, x_148); -return x_149; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Grind", 5, 5); +return x_1; +} } +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__1___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("eqRecOn_heq", 11, 11); +return x_1; } } +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__1___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__1___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__1___closed__2; +x_3 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__1___closed__3; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; } } -block_117: +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17) { +_start: { -lean_object* x_107; -lean_dec(x_106); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); +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_18 = l_Lean_Expr_constLevels_x21(x_2); +x_19 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__1___closed__4; +x_20 = l_Lean_Expr_const___override(x_19, x_18); lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_2); -lean_inc(x_3); -x_107 = l_Lean_Meta_Grind_internalize(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_107) == 0) -{ -lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; -x_108 = lean_ctor_get(x_107, 1); -lean_inc(x_108); -lean_dec(x_107); -lean_inc(x_1); -x_109 = l_Lean_Meta_Grind_registerParent(x_1, x_3, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_108); -lean_dec(x_3); -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_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(x_4, x_1, x_2, x_105, x_110, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_111); -lean_dec(x_110); -lean_dec(x_4); -return x_112; +x_21 = l_Lean_mkApp6(x_20, x_3, x_4, x_5, x_6, x_7, x_8); +x_22 = 1; +x_23 = l_Lean_Meta_Grind_pushEqCore(x_1, x_8, x_21, x_22, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +return x_23; } -else -{ -uint8_t x_113; -lean_dec(x_105); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_113 = !lean_is_exclusive(x_107); -if (x_113 == 0) -{ -return x_107; } -else +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__2___closed__1() { +_start: { -lean_object* x_114; lean_object* x_115; lean_object* x_116; -x_114 = lean_ctor_get(x_107, 0); -x_115 = lean_ctor_get(x_107, 1); -lean_inc(x_115); -lean_inc(x_114); -lean_dec(x_107); -x_116 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_116, 0, x_114); -lean_ctor_set(x_116, 1, x_115); -return x_116; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("eqNDRec_heq", 11, 11); +return x_1; } } +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__2___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__1___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__1___closed__2; +x_3 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__2___closed__1; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; } } -case 3: -{ -lean_object* x_150; lean_object* x_151; lean_object* x_163; uint8_t x_164; -lean_dec(x_5); -lean_inc(x_2); -lean_inc(x_1); -x_150 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1___boxed), 12, 2); -lean_closure_set(x_150, 0, x_1); -lean_closure_set(x_150, 1, x_2); -x_163 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__4; -x_164 = l_Lean_Expr_isConstOf(x_3, x_163); -if (x_164 == 0) +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17) { +_start: { -lean_object* x_165; -x_165 = lean_box(0); -x_151 = x_165; -goto block_162; +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_18 = l_Lean_Expr_constLevels_x21(x_2); +x_19 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__2___closed__2; +x_20 = l_Lean_Expr_const___override(x_19, x_18); +lean_inc(x_6); +x_21 = l_Lean_mkApp6(x_20, x_3, x_4, x_5, x_6, x_7, x_8); +x_22 = 1; +x_23 = l_Lean_Meta_Grind_pushEqCore(x_1, x_6, x_21, x_22, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +return x_23; } -else -{ -lean_object* x_166; lean_object* x_167; uint8_t x_168; -x_166 = lean_array_get_size(x_4); -x_167 = lean_unsigned_to_nat(2u); -x_168 = lean_nat_dec_eq(x_166, x_167); -if (x_168 == 0) +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__3___closed__1() { +_start: { -lean_object* x_169; -lean_dec(x_166); -x_169 = lean_box(0); -x_151 = x_169; -goto block_162; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("eqRec_heq", 9, 9); +return x_1; } -else +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__3___closed__2() { +_start: { -lean_object* x_170; uint8_t x_171; -lean_dec(x_150); -lean_dec(x_3); -x_170 = lean_unsigned_to_nat(0u); -x_171 = lean_nat_dec_lt(x_170, x_166); -lean_dec(x_166); -if (x_171 == 0) +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__1___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__1___closed__2; +x_3 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__3___closed__1; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___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) { +_start: { -lean_object* x_172; lean_object* x_173; lean_object* x_174; -lean_dec(x_4); -x_172 = l_Lean_instInhabitedExpr; -x_173 = l_outOfBounds___rarg(x_172); -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_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; lean_object* x_23; +x_18 = l_Lean_Expr_constLevels_x21(x_2); +x_19 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__3___closed__2; +x_20 = l_Lean_Expr_const___override(x_19, x_18); lean_inc(x_6); -lean_inc(x_2); -lean_inc(x_173); -x_174 = l_Lean_Meta_Grind_internalize(x_173, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_174) == 0) -{ -lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; -x_175 = lean_ctor_get(x_174, 1); -lean_inc(x_175); -lean_dec(x_174); -lean_inc(x_1); -x_176 = l_Lean_Meta_Grind_registerParent(x_1, x_173, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_175); -lean_dec(x_173); -x_177 = lean_ctor_get(x_176, 0); -lean_inc(x_177); -x_178 = lean_ctor_get(x_176, 1); -lean_inc(x_178); -lean_dec(x_176); -x_179 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_177, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_178); -lean_dec(x_177); -return x_179; +x_21 = l_Lean_mkApp6(x_20, x_3, x_4, x_5, x_6, x_7, x_8); +x_22 = 1; +x_23 = l_Lean_Meta_Grind_pushEqCore(x_1, x_6, x_21, x_22, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +return x_23; } -else -{ -uint8_t x_180; -lean_dec(x_173); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_2); -lean_dec(x_1); -x_180 = !lean_is_exclusive(x_174); -if (x_180 == 0) -{ -return x_174; } -else +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__4___closed__1() { +_start: { -lean_object* x_181; lean_object* x_182; lean_object* x_183; -x_181 = lean_ctor_get(x_174, 0); -x_182 = lean_ctor_get(x_174, 1); -lean_inc(x_182); -lean_inc(x_181); -lean_dec(x_174); -x_183 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_183, 0, x_181); -lean_ctor_set(x_183, 1, x_182); -return x_183; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("cast_heq", 8, 8); +return x_1; } } +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__4___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__4___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } -else +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___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, lean_object* x_15) { +_start: { -lean_object* x_184; lean_object* x_185; -x_184 = lean_array_fget(x_4, x_170); -lean_dec(x_4); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; lean_object* x_21; +x_16 = l_Lean_Expr_constLevels_x21(x_2); +x_17 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__4___closed__2; +x_18 = l_Lean_Expr_const___override(x_17, x_16); lean_inc(x_6); -lean_inc(x_2); -lean_inc(x_184); -x_185 = l_Lean_Meta_Grind_internalize(x_184, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -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; -x_186 = lean_ctor_get(x_185, 1); -lean_inc(x_186); -lean_dec(x_185); -lean_inc(x_1); -x_187 = l_Lean_Meta_Grind_registerParent(x_1, x_184, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_186); -lean_dec(x_184); -x_188 = lean_ctor_get(x_187, 0); -lean_inc(x_188); -x_189 = lean_ctor_get(x_187, 1); -lean_inc(x_189); -lean_dec(x_187); -x_190 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_188, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_189); -lean_dec(x_188); -return x_190; +x_19 = l_Lean_mkApp4(x_18, x_3, x_4, x_5, x_6); +x_20 = 1; +x_21 = l_Lean_Meta_Grind_pushEqCore(x_1, x_6, x_19, x_20, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +return x_21; } -else -{ -uint8_t x_191; -lean_dec(x_184); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_2); -lean_dec(x_1); -x_191 = !lean_is_exclusive(x_185); -if (x_191 == 0) -{ -return x_185; } -else +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__1() { +_start: { -lean_object* x_192; lean_object* x_193; lean_object* x_194; -x_192 = lean_ctor_get(x_185, 0); -x_193 = lean_ctor_get(x_185, 1); -lean_inc(x_193); -lean_inc(x_192); -lean_dec(x_185); -x_194 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_194, 0, x_192); -lean_ctor_set(x_194, 1, x_193); -return x_194; -} +lean_object* x_1; +x_1 = lean_mk_string_unchecked("cast", 4, 4); +return x_1; } } +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } } -block_162: +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__3() { +_start: { -lean_object* x_152; -lean_dec(x_151); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_2); -lean_inc(x_3); -x_152 = l_Lean_Meta_Grind_internalize(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_152) == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("recOn", 5, 5); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__4() { +_start: { -lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; -x_153 = lean_ctor_get(x_152, 1); -lean_inc(x_153); -lean_dec(x_152); -lean_inc(x_1); -x_154 = l_Lean_Meta_Grind_registerParent(x_1, x_3, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_153); -lean_dec(x_3); -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 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(x_4, x_1, x_2, x_150, x_155, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_156); -lean_dec(x_155); -lean_dec(x_4); -return x_157; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__3; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; } -else +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__5() { +_start: { -uint8_t x_158; -lean_dec(x_150); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_158 = !lean_is_exclusive(x_152); -if (x_158 == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("ndrec", 5, 5); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__6() { +_start: { -return x_152; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__5; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; } -else +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__7() { +_start: { -lean_object* x_159; lean_object* x_160; lean_object* x_161; -x_159 = lean_ctor_get(x_152, 0); -x_160 = lean_ctor_get(x_152, 1); -lean_inc(x_160); -lean_inc(x_159); -lean_dec(x_152); -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_1; +x_1 = lean_mk_string_unchecked("rec", 3, 3); +return x_1; } } +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__7; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; } } -case 4: +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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_195; lean_object* x_196; lean_object* x_197; lean_object* x_207; uint8_t x_208; -lean_dec(x_5); -x_195 = lean_ctor_get(x_3, 0); -lean_inc(x_195); -lean_inc(x_2); +lean_object* x_11; uint8_t x_12; lean_inc(x_1); -x_196 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1___boxed), 12, 2); -lean_closure_set(x_196, 0, x_1); -lean_closure_set(x_196, 1, x_2); -x_207 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__4; -x_208 = l_Lean_Expr_isConstOf(x_3, x_207); -lean_dec(x_3); -if (x_208 == 0) +x_11 = l_Lean_Meta_instantiateMVarsIfMVarApp(x_1, x_6, x_7, x_8, x_9, x_10); +x_12 = !lean_is_exclusive(x_11); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_13 = lean_ctor_get(x_11, 0); +x_14 = lean_ctor_get(x_11, 1); +x_15 = l_Lean_Expr_cleanupAnnotations(x_13); +x_16 = l_Lean_Expr_isApp(x_15); +if (x_16 == 0) { -lean_object* x_209; -x_209 = lean_box(0); -x_197 = x_209; -goto block_206; +lean_object* x_17; +lean_dec(x_15); +lean_dec(x_1); +x_17 = lean_box(0); +lean_ctor_set(x_11, 0, x_17); +return x_11; } else { -lean_object* x_210; lean_object* x_211; uint8_t x_212; -x_210 = lean_array_get_size(x_4); -x_211 = lean_unsigned_to_nat(2u); -x_212 = lean_nat_dec_eq(x_210, x_211); -if (x_212 == 0) +lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_18 = l_Lean_Expr_appArg(x_15, lean_box(0)); +x_19 = l_Lean_Expr_appFnCleanup(x_15, lean_box(0)); +x_20 = l_Lean_Expr_isApp(x_19); +if (x_20 == 0) { -lean_object* x_213; -lean_dec(x_210); -x_213 = lean_box(0); -x_197 = x_213; -goto block_206; +lean_object* x_21; +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_1); +x_21 = lean_box(0); +lean_ctor_set(x_11, 0, x_21); +return x_11; } else { -lean_object* x_214; uint8_t x_215; -lean_dec(x_196); -lean_dec(x_195); -x_214 = lean_unsigned_to_nat(0u); -x_215 = lean_nat_dec_lt(x_214, x_210); -lean_dec(x_210); -if (x_215 == 0) -{ -lean_object* x_216; lean_object* x_217; lean_object* x_218; -lean_dec(x_4); -x_216 = l_Lean_instInhabitedExpr; -x_217 = l_outOfBounds___rarg(x_216); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_2); -lean_inc(x_217); -x_218 = l_Lean_Meta_Grind_internalize(x_217, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_218) == 0) +lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_22 = l_Lean_Expr_appArg(x_19, lean_box(0)); +x_23 = l_Lean_Expr_appFnCleanup(x_19, lean_box(0)); +x_24 = l_Lean_Expr_isApp(x_23); +if (x_24 == 0) { -lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; -x_219 = lean_ctor_get(x_218, 1); -lean_inc(x_219); -lean_dec(x_218); -lean_inc(x_1); -x_220 = l_Lean_Meta_Grind_registerParent(x_1, x_217, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_219); -lean_dec(x_217); -x_221 = lean_ctor_get(x_220, 0); -lean_inc(x_221); -x_222 = lean_ctor_get(x_220, 1); -lean_inc(x_222); -lean_dec(x_220); -x_223 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_221, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_222); -lean_dec(x_221); -return x_223; +lean_object* x_25; +lean_dec(x_23); +lean_dec(x_22); +lean_dec(x_18); +lean_dec(x_1); +x_25 = lean_box(0); +lean_ctor_set(x_11, 0, x_25); +return x_11; } else { -uint8_t x_224; -lean_dec(x_217); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_2); -lean_dec(x_1); -x_224 = !lean_is_exclusive(x_218); -if (x_224 == 0) +lean_object* x_26; lean_object* x_27; uint8_t x_28; +x_26 = l_Lean_Expr_appArg(x_23, lean_box(0)); +x_27 = l_Lean_Expr_appFnCleanup(x_23, lean_box(0)); +x_28 = l_Lean_Expr_isApp(x_27); +if (x_28 == 0) { -return x_218; +lean_object* x_29; +lean_dec(x_27); +lean_dec(x_26); +lean_dec(x_22); +lean_dec(x_18); +lean_dec(x_1); +x_29 = lean_box(0); +lean_ctor_set(x_11, 0, x_29); +return x_11; } else { -lean_object* x_225; lean_object* x_226; lean_object* x_227; -x_225 = lean_ctor_get(x_218, 0); -x_226 = lean_ctor_get(x_218, 1); -lean_inc(x_226); -lean_inc(x_225); -lean_dec(x_218); -x_227 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_227, 0, x_225); -lean_ctor_set(x_227, 1, x_226); -return x_227; -} -} +lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; +x_30 = l_Lean_Expr_appArg(x_27, lean_box(0)); +x_31 = l_Lean_Expr_appFnCleanup(x_27, lean_box(0)); +x_32 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__2; +x_33 = l_Lean_Expr_isConstOf(x_31, x_32); +if (x_33 == 0) +{ +uint8_t x_34; +x_34 = l_Lean_Expr_isApp(x_31); +if (x_34 == 0) +{ +lean_object* x_35; +lean_dec(x_31); +lean_dec(x_30); +lean_dec(x_26); +lean_dec(x_22); +lean_dec(x_18); +lean_dec(x_1); +x_35 = lean_box(0); +lean_ctor_set(x_11, 0, x_35); +return x_11; } else { -lean_object* x_228; lean_object* x_229; -x_228 = lean_array_fget(x_4, x_214); -lean_dec(x_4); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_2); -lean_inc(x_228); -x_229 = l_Lean_Meta_Grind_internalize(x_228, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_229) == 0) +lean_object* x_36; lean_object* x_37; uint8_t x_38; +x_36 = l_Lean_Expr_appArg(x_31, lean_box(0)); +x_37 = l_Lean_Expr_appFnCleanup(x_31, lean_box(0)); +x_38 = l_Lean_Expr_isApp(x_37); +if (x_38 == 0) { -lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; -x_230 = lean_ctor_get(x_229, 1); -lean_inc(x_230); -lean_dec(x_229); -lean_inc(x_1); -x_231 = l_Lean_Meta_Grind_registerParent(x_1, x_228, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_230); -lean_dec(x_228); -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_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_232, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_233); -lean_dec(x_232); -return x_234; +lean_object* x_39; +lean_dec(x_37); +lean_dec(x_36); +lean_dec(x_30); +lean_dec(x_26); +lean_dec(x_22); +lean_dec(x_18); +lean_dec(x_1); +x_39 = lean_box(0); +lean_ctor_set(x_11, 0, x_39); +return x_11; } else { -uint8_t x_235; -lean_dec(x_228); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_2); +lean_object* x_40; lean_object* x_41; lean_object* x_42; uint8_t x_43; +x_40 = l_Lean_Expr_appArg(x_37, lean_box(0)); +x_41 = l_Lean_Expr_appFnCleanup(x_37, lean_box(0)); +x_42 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__4; +x_43 = l_Lean_Expr_isConstOf(x_41, x_42); +if (x_43 == 0) +{ +lean_object* x_44; uint8_t x_45; +x_44 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__6; +x_45 = l_Lean_Expr_isConstOf(x_41, x_44); +if (x_45 == 0) +{ +lean_object* x_46; uint8_t x_47; +x_46 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__8; +x_47 = l_Lean_Expr_isConstOf(x_41, x_46); +if (x_47 == 0) +{ +lean_object* x_48; +lean_dec(x_41); +lean_dec(x_40); +lean_dec(x_36); +lean_dec(x_30); +lean_dec(x_26); +lean_dec(x_22); +lean_dec(x_18); lean_dec(x_1); -x_235 = !lean_is_exclusive(x_229); -if (x_235 == 0) +x_48 = lean_box(0); +lean_ctor_set(x_11, 0, x_48); +return x_11; +} +else { -return x_229; +lean_object* x_49; +lean_free_object(x_11); +x_49 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__3(x_1, x_41, x_40, x_36, x_30, x_26, x_22, x_18, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_14); +lean_dec(x_41); +return x_49; +} } else { -lean_object* x_236; lean_object* x_237; lean_object* x_238; -x_236 = lean_ctor_get(x_229, 0); -x_237 = lean_ctor_get(x_229, 1); -lean_inc(x_237); -lean_inc(x_236); -lean_dec(x_229); -x_238 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_238, 0, x_236); -lean_ctor_set(x_238, 1, x_237); -return x_238; +lean_object* x_50; +lean_free_object(x_11); +x_50 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__2(x_1, x_41, x_40, x_36, x_30, x_26, x_22, x_18, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_14); +lean_dec(x_41); +return x_50; } } +else +{ +lean_object* x_51; +lean_free_object(x_11); +x_51 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__1(x_1, x_41, x_40, x_36, x_30, x_26, x_22, x_18, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_14); +lean_dec(x_41); +return x_51; } } } -block_206: -{ -lean_object* x_198; -lean_dec(x_197); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_2); -x_198 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns(x_195, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -lean_dec(x_195); -if (lean_obj_tag(x_198) == 0) -{ -lean_object* x_199; lean_object* x_200; lean_object* x_201; -x_199 = lean_ctor_get(x_198, 0); -lean_inc(x_199); -x_200 = lean_ctor_get(x_198, 1); -lean_inc(x_200); -lean_dec(x_198); -x_201 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(x_4, x_1, x_2, x_196, x_199, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_200); -lean_dec(x_199); -lean_dec(x_4); -return x_201; } else { -uint8_t x_202; -lean_dec(x_196); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_4); -lean_dec(x_2); -lean_dec(x_1); -x_202 = !lean_is_exclusive(x_198); -if (x_202 == 0) -{ -return x_198; +lean_object* x_52; +lean_free_object(x_11); +x_52 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__4(x_1, x_31, x_30, x_26, x_22, x_18, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_14); +lean_dec(x_31); +return x_52; } -else -{ -lean_object* x_203; lean_object* x_204; lean_object* x_205; -x_203 = lean_ctor_get(x_198, 0); -x_204 = lean_ctor_get(x_198, 1); -lean_inc(x_204); -lean_inc(x_203); -lean_dec(x_198); -x_205 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_205, 0, x_203); -lean_ctor_set(x_205, 1, x_204); -return x_205; } } } } -case 5: -{ -lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; -x_239 = lean_ctor_get(x_3, 0); -lean_inc(x_239); -x_240 = lean_ctor_get(x_3, 1); -lean_inc(x_240); -lean_dec(x_3); -x_241 = lean_array_set(x_4, x_5, x_240); -x_242 = lean_unsigned_to_nat(1u); -x_243 = lean_nat_sub(x_5, x_242); -lean_dec(x_5); -x_3 = x_239; -x_4 = x_241; -x_5 = x_243; -goto _start; } -case 6: +else { -lean_object* x_245; lean_object* x_246; lean_object* x_258; uint8_t x_259; -lean_dec(x_5); -lean_inc(x_2); -lean_inc(x_1); -x_245 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1___boxed), 12, 2); -lean_closure_set(x_245, 0, x_1); -lean_closure_set(x_245, 1, x_2); -x_258 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__4; -x_259 = l_Lean_Expr_isConstOf(x_3, x_258); -if (x_259 == 0) +lean_object* x_53; lean_object* x_54; lean_object* x_55; uint8_t x_56; +x_53 = lean_ctor_get(x_11, 0); +x_54 = lean_ctor_get(x_11, 1); +lean_inc(x_54); +lean_inc(x_53); +lean_dec(x_11); +x_55 = l_Lean_Expr_cleanupAnnotations(x_53); +x_56 = l_Lean_Expr_isApp(x_55); +if (x_56 == 0) { -lean_object* x_260; -x_260 = lean_box(0); -x_246 = x_260; -goto block_257; +lean_object* x_57; lean_object* x_58; +lean_dec(x_55); +lean_dec(x_1); +x_57 = lean_box(0); +x_58 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_58, 0, x_57); +lean_ctor_set(x_58, 1, x_54); +return x_58; } else { -lean_object* x_261; lean_object* x_262; uint8_t x_263; -x_261 = lean_array_get_size(x_4); -x_262 = lean_unsigned_to_nat(2u); -x_263 = lean_nat_dec_eq(x_261, x_262); -if (x_263 == 0) +lean_object* x_59; lean_object* x_60; uint8_t x_61; +x_59 = l_Lean_Expr_appArg(x_55, lean_box(0)); +x_60 = l_Lean_Expr_appFnCleanup(x_55, lean_box(0)); +x_61 = l_Lean_Expr_isApp(x_60); +if (x_61 == 0) { -lean_object* x_264; -lean_dec(x_261); -x_264 = lean_box(0); -x_246 = x_264; -goto block_257; +lean_object* x_62; lean_object* x_63; +lean_dec(x_60); +lean_dec(x_59); +lean_dec(x_1); +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_54); +return x_63; } else { -lean_object* x_265; uint8_t x_266; -lean_dec(x_245); -lean_dec(x_3); -x_265 = lean_unsigned_to_nat(0u); -x_266 = lean_nat_dec_lt(x_265, x_261); -lean_dec(x_261); -if (x_266 == 0) +lean_object* x_64; lean_object* x_65; uint8_t x_66; +x_64 = l_Lean_Expr_appArg(x_60, lean_box(0)); +x_65 = l_Lean_Expr_appFnCleanup(x_60, lean_box(0)); +x_66 = l_Lean_Expr_isApp(x_65); +if (x_66 == 0) { -lean_object* x_267; lean_object* x_268; lean_object* x_269; -lean_dec(x_4); -x_267 = l_Lean_instInhabitedExpr; -x_268 = l_outOfBounds___rarg(x_267); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_2); -lean_inc(x_268); -x_269 = l_Lean_Meta_Grind_internalize(x_268, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_269) == 0) -{ -lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; -x_270 = lean_ctor_get(x_269, 1); -lean_inc(x_270); -lean_dec(x_269); -lean_inc(x_1); -x_271 = l_Lean_Meta_Grind_registerParent(x_1, x_268, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_270); -lean_dec(x_268); -x_272 = lean_ctor_get(x_271, 0); -lean_inc(x_272); -x_273 = lean_ctor_get(x_271, 1); -lean_inc(x_273); -lean_dec(x_271); -x_274 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_272, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_273); -lean_dec(x_272); -return x_274; +lean_object* x_67; lean_object* x_68; +lean_dec(x_65); +lean_dec(x_64); +lean_dec(x_59); +lean_dec(x_1); +x_67 = lean_box(0); +x_68 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_68, 0, x_67); +lean_ctor_set(x_68, 1, x_54); +return x_68; } else { -uint8_t x_275; -lean_dec(x_268); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_2); -lean_dec(x_1); -x_275 = !lean_is_exclusive(x_269); -if (x_275 == 0) +lean_object* x_69; lean_object* x_70; uint8_t x_71; +x_69 = l_Lean_Expr_appArg(x_65, lean_box(0)); +x_70 = l_Lean_Expr_appFnCleanup(x_65, lean_box(0)); +x_71 = l_Lean_Expr_isApp(x_70); +if (x_71 == 0) { -return x_269; +lean_object* x_72; lean_object* x_73; +lean_dec(x_70); +lean_dec(x_69); +lean_dec(x_64); +lean_dec(x_59); +lean_dec(x_1); +x_72 = lean_box(0); +x_73 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_73, 0, x_72); +lean_ctor_set(x_73, 1, x_54); +return x_73; } else { -lean_object* x_276; lean_object* x_277; lean_object* x_278; -x_276 = lean_ctor_get(x_269, 0); -x_277 = lean_ctor_get(x_269, 1); -lean_inc(x_277); -lean_inc(x_276); -lean_dec(x_269); -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 +lean_object* x_74; lean_object* x_75; lean_object* x_76; uint8_t x_77; +x_74 = l_Lean_Expr_appArg(x_70, lean_box(0)); +x_75 = l_Lean_Expr_appFnCleanup(x_70, lean_box(0)); +x_76 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__2; +x_77 = l_Lean_Expr_isConstOf(x_75, x_76); +if (x_77 == 0) { -lean_object* x_279; lean_object* x_280; -x_279 = lean_array_fget(x_4, x_265); -lean_dec(x_4); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_2); -lean_inc(x_279); -x_280 = l_Lean_Meta_Grind_internalize(x_279, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_280) == 0) +uint8_t x_78; +x_78 = l_Lean_Expr_isApp(x_75); +if (x_78 == 0) { -lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; -x_281 = lean_ctor_get(x_280, 1); -lean_inc(x_281); -lean_dec(x_280); -lean_inc(x_1); -x_282 = l_Lean_Meta_Grind_registerParent(x_1, x_279, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_281); -lean_dec(x_279); -x_283 = lean_ctor_get(x_282, 0); -lean_inc(x_283); -x_284 = lean_ctor_get(x_282, 1); -lean_inc(x_284); -lean_dec(x_282); -x_285 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_283, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_284); -lean_dec(x_283); -return x_285; +lean_object* x_79; lean_object* x_80; +lean_dec(x_75); +lean_dec(x_74); +lean_dec(x_69); +lean_dec(x_64); +lean_dec(x_59); +lean_dec(x_1); +x_79 = lean_box(0); +x_80 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_80, 0, x_79); +lean_ctor_set(x_80, 1, x_54); +return x_80; } else { -uint8_t x_286; -lean_dec(x_279); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_2); -lean_dec(x_1); -x_286 = !lean_is_exclusive(x_280); -if (x_286 == 0) +lean_object* x_81; lean_object* x_82; uint8_t x_83; +x_81 = l_Lean_Expr_appArg(x_75, lean_box(0)); +x_82 = l_Lean_Expr_appFnCleanup(x_75, lean_box(0)); +x_83 = l_Lean_Expr_isApp(x_82); +if (x_83 == 0) { -return x_280; +lean_object* x_84; lean_object* x_85; +lean_dec(x_82); +lean_dec(x_81); +lean_dec(x_74); +lean_dec(x_69); +lean_dec(x_64); +lean_dec(x_59); +lean_dec(x_1); +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_54); +return x_85; } else { -lean_object* x_287; lean_object* x_288; lean_object* x_289; -x_287 = lean_ctor_get(x_280, 0); -x_288 = lean_ctor_get(x_280, 1); -lean_inc(x_288); -lean_inc(x_287); -lean_dec(x_280); -x_289 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_289, 0, x_287); -lean_ctor_set(x_289, 1, x_288); -return x_289; -} -} -} -} -} -block_257: +lean_object* x_86; lean_object* x_87; lean_object* x_88; uint8_t x_89; +x_86 = l_Lean_Expr_appArg(x_82, lean_box(0)); +x_87 = l_Lean_Expr_appFnCleanup(x_82, lean_box(0)); +x_88 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__4; +x_89 = l_Lean_Expr_isConstOf(x_87, x_88); +if (x_89 == 0) { -lean_object* x_247; -lean_dec(x_246); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_2); -lean_inc(x_3); -x_247 = l_Lean_Meta_Grind_internalize(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_247) == 0) +lean_object* x_90; uint8_t x_91; +x_90 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__6; +x_91 = l_Lean_Expr_isConstOf(x_87, x_90); +if (x_91 == 0) { -lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; -x_248 = lean_ctor_get(x_247, 1); -lean_inc(x_248); -lean_dec(x_247); -lean_inc(x_1); -x_249 = l_Lean_Meta_Grind_registerParent(x_1, x_3, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_248); -lean_dec(x_3); -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_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(x_4, x_1, x_2, x_245, x_250, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_251); -lean_dec(x_250); -lean_dec(x_4); -return x_252; -} -else +lean_object* x_92; uint8_t x_93; +x_92 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__8; +x_93 = l_Lean_Expr_isConstOf(x_87, x_92); +if (x_93 == 0) { -uint8_t x_253; -lean_dec(x_245); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); +lean_object* x_94; lean_object* x_95; +lean_dec(x_87); +lean_dec(x_86); +lean_dec(x_81); +lean_dec(x_74); +lean_dec(x_69); +lean_dec(x_64); +lean_dec(x_59); lean_dec(x_1); -x_253 = !lean_is_exclusive(x_247); -if (x_253 == 0) -{ -return x_247; +x_94 = lean_box(0); +x_95 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_95, 0, x_94); +lean_ctor_set(x_95, 1, x_54); +return x_95; } else { -lean_object* x_254; lean_object* x_255; lean_object* x_256; -x_254 = lean_ctor_get(x_247, 0); -x_255 = lean_ctor_get(x_247, 1); -lean_inc(x_255); -lean_inc(x_254); -lean_dec(x_247); -x_256 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_256, 0, x_254); -lean_ctor_set(x_256, 1, x_255); -return x_256; -} -} -} +lean_object* x_96; +x_96 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__3(x_1, x_87, x_86, x_81, x_74, x_69, x_64, x_59, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_54); +lean_dec(x_87); +return x_96; } -case 7: -{ -lean_object* x_290; lean_object* x_291; lean_object* x_303; uint8_t x_304; -lean_dec(x_5); -lean_inc(x_2); -lean_inc(x_1); -x_290 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1___boxed), 12, 2); -lean_closure_set(x_290, 0, x_1); -lean_closure_set(x_290, 1, x_2); -x_303 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__4; -x_304 = l_Lean_Expr_isConstOf(x_3, x_303); -if (x_304 == 0) -{ -lean_object* x_305; -x_305 = lean_box(0); -x_291 = x_305; -goto block_302; } else { -lean_object* x_306; lean_object* x_307; uint8_t x_308; -x_306 = lean_array_get_size(x_4); -x_307 = lean_unsigned_to_nat(2u); -x_308 = lean_nat_dec_eq(x_306, x_307); -if (x_308 == 0) -{ -lean_object* x_309; -lean_dec(x_306); -x_309 = lean_box(0); -x_291 = x_309; -goto block_302; +lean_object* x_97; +x_97 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__2(x_1, x_87, x_86, x_81, x_74, x_69, x_64, x_59, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_54); +lean_dec(x_87); +return x_97; } -else -{ -lean_object* x_310; uint8_t x_311; -lean_dec(x_290); -lean_dec(x_3); -x_310 = lean_unsigned_to_nat(0u); -x_311 = lean_nat_dec_lt(x_310, x_306); -lean_dec(x_306); -if (x_311 == 0) -{ -lean_object* x_312; lean_object* x_313; lean_object* x_314; -lean_dec(x_4); -x_312 = l_Lean_instInhabitedExpr; -x_313 = l_outOfBounds___rarg(x_312); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_2); -lean_inc(x_313); -x_314 = l_Lean_Meta_Grind_internalize(x_313, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -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; -x_315 = lean_ctor_get(x_314, 1); -lean_inc(x_315); -lean_dec(x_314); -lean_inc(x_1); -x_316 = l_Lean_Meta_Grind_registerParent(x_1, x_313, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_315); -lean_dec(x_313); -x_317 = lean_ctor_get(x_316, 0); -lean_inc(x_317); -x_318 = lean_ctor_get(x_316, 1); -lean_inc(x_318); -lean_dec(x_316); -x_319 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_317, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_318); -lean_dec(x_317); -return x_319; } else { -uint8_t x_320; -lean_dec(x_313); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_2); -lean_dec(x_1); -x_320 = !lean_is_exclusive(x_314); -if (x_320 == 0) -{ -return x_314; +lean_object* x_98; +x_98 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__1(x_1, x_87, x_86, x_81, x_74, x_69, x_64, x_59, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_54); +lean_dec(x_87); +return x_98; } -else -{ -lean_object* x_321; lean_object* x_322; lean_object* x_323; -x_321 = lean_ctor_get(x_314, 0); -x_322 = lean_ctor_get(x_314, 1); -lean_inc(x_322); -lean_inc(x_321); -lean_dec(x_314); -x_323 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_323, 0, x_321); -lean_ctor_set(x_323, 1, x_322); -return x_323; } } } else { -lean_object* x_324; lean_object* x_325; -x_324 = lean_array_fget(x_4, x_310); -lean_dec(x_4); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_2); -lean_inc(x_324); -x_325 = l_Lean_Meta_Grind_internalize(x_324, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_325) == 0) -{ -lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; -x_326 = lean_ctor_get(x_325, 1); -lean_inc(x_326); -lean_dec(x_325); -lean_inc(x_1); -x_327 = l_Lean_Meta_Grind_registerParent(x_1, x_324, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_326); -lean_dec(x_324); -x_328 = lean_ctor_get(x_327, 0); -lean_inc(x_328); -x_329 = lean_ctor_get(x_327, 1); -lean_inc(x_329); -lean_dec(x_327); -x_330 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_328, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_329); -lean_dec(x_328); -return x_330; +lean_object* x_99; +x_99 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__4(x_1, x_75, x_74, x_69, x_64, x_59, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_54); +lean_dec(x_75); +return x_99; } -else -{ -uint8_t x_331; -lean_dec(x_324); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_2); -lean_dec(x_1); -x_331 = !lean_is_exclusive(x_325); -if (x_331 == 0) -{ -return x_325; } -else -{ -lean_object* x_332; lean_object* x_333; lean_object* x_334; -x_332 = lean_ctor_get(x_325, 0); -x_333 = lean_ctor_get(x_325, 1); -lean_inc(x_333); -lean_inc(x_332); -lean_dec(x_325); -x_334 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_334, 0, x_332); -lean_ctor_set(x_334, 1, x_333); -return x_334; } } } } } -block_302: -{ -lean_object* x_292; -lean_dec(x_291); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_2); -lean_inc(x_3); -x_292 = l_Lean_Meta_Grind_internalize(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_292) == 0) -{ -lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; -x_293 = lean_ctor_get(x_292, 1); -lean_inc(x_293); -lean_dec(x_292); -lean_inc(x_1); -x_294 = l_Lean_Meta_Grind_registerParent(x_1, x_3, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_293); -lean_dec(x_3); -x_295 = lean_ctor_get(x_294, 0); -lean_inc(x_295); -x_296 = lean_ctor_get(x_294, 1); -lean_inc(x_296); -lean_dec(x_294); -x_297 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(x_4, x_1, x_2, x_290, x_295, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_296); -lean_dec(x_295); -lean_dec(x_4); -return x_297; } -else +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__1___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +_start: { -uint8_t x_298; -lean_dec(x_290); +lean_object* x_18; +x_18 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_4); -lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -x_298 = !lean_is_exclusive(x_292); -if (x_298 == 0) -{ -return x_292; -} -else -{ -lean_object* x_299; lean_object* x_300; lean_object* x_301; -x_299 = lean_ctor_get(x_292, 0); -x_300 = lean_ctor_get(x_292, 1); -lean_inc(x_300); -lean_inc(x_299); -lean_dec(x_292); -x_301 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_301, 0, x_299); -lean_ctor_set(x_301, 1, x_300); -return x_301; -} -} +return x_18; } } -case 8: -{ -lean_object* x_335; lean_object* x_336; lean_object* x_348; uint8_t x_349; -lean_dec(x_5); -lean_inc(x_2); -lean_inc(x_1); -x_335 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1___boxed), 12, 2); -lean_closure_set(x_335, 0, x_1); -lean_closure_set(x_335, 1, x_2); -x_348 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__4; -x_349 = l_Lean_Expr_isConstOf(x_3, x_348); -if (x_349 == 0) +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__2___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +_start: { -lean_object* x_350; -x_350 = lean_box(0); -x_336 = x_350; -goto block_347; +lean_object* x_18; +x_18 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_2); +return x_18; } -else -{ -lean_object* x_351; lean_object* x_352; uint8_t x_353; -x_351 = lean_array_get_size(x_4); -x_352 = lean_unsigned_to_nat(2u); -x_353 = lean_nat_dec_eq(x_351, x_352); -if (x_353 == 0) -{ -lean_object* x_354; -lean_dec(x_351); -x_354 = lean_box(0); -x_336 = x_354; -goto block_347; } -else -{ -lean_object* x_355; uint8_t x_356; -lean_dec(x_335); -lean_dec(x_3); -x_355 = lean_unsigned_to_nat(0u); -x_356 = lean_nat_dec_lt(x_355, x_351); -lean_dec(x_351); -if (x_356 == 0) -{ -lean_object* x_357; lean_object* x_358; lean_object* x_359; -lean_dec(x_4); -x_357 = l_Lean_instInhabitedExpr; -x_358 = l_outOfBounds___rarg(x_357); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_2); -lean_inc(x_358); -x_359 = l_Lean_Meta_Grind_internalize(x_358, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_359) == 0) +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__3___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +_start: { -lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; -x_360 = lean_ctor_get(x_359, 1); -lean_inc(x_360); -lean_dec(x_359); -lean_inc(x_1); -x_361 = l_Lean_Meta_Grind_registerParent(x_1, x_358, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_360); -lean_dec(x_358); -x_362 = lean_ctor_get(x_361, 0); -lean_inc(x_362); -x_363 = lean_ctor_get(x_361, 1); -lean_inc(x_363); -lean_dec(x_361); -x_364 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_362, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_363); -lean_dec(x_362); -return x_364; +lean_object* x_18; +x_18 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_2); +return x_18; } -else +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +_start: { -uint8_t x_365; -lean_dec(x_358); +lean_object* x_16; +x_16 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); @@ -6683,68 +6029,34 @@ lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_6); lean_dec(x_2); -lean_dec(x_1); -x_365 = !lean_is_exclusive(x_359); -if (x_365 == 0) -{ -return x_359; -} -else -{ -lean_object* x_366; lean_object* x_367; lean_object* x_368; -x_366 = lean_ctor_get(x_359, 0); -x_367 = lean_ctor_get(x_359, 1); -lean_inc(x_367); -lean_inc(x_366); -lean_dec(x_359); -x_368 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_368, 0, x_366); -lean_ctor_set(x_368, 1, x_367); -return x_368; -} +return x_16; } } -else +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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_369; lean_object* x_370; -x_369 = lean_array_fget(x_4, x_355); +lean_object* x_11; +x_11 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); lean_dec(x_4); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_2); -lean_inc(x_369); -x_370 = l_Lean_Meta_Grind_internalize(x_369, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_370) == 0) -{ -lean_object* x_371; lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; -x_371 = lean_ctor_get(x_370, 1); -lean_inc(x_371); -lean_dec(x_370); -lean_inc(x_1); -x_372 = l_Lean_Meta_Grind_registerParent(x_1, x_369, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_371); -lean_dec(x_369); -x_373 = lean_ctor_get(x_372, 0); -lean_inc(x_373); -x_374 = lean_ctor_get(x_372, 1); -lean_inc(x_374); -lean_dec(x_372); -x_375 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_373, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_374); -lean_dec(x_373); -return x_375; +lean_dec(x_3); +lean_dec(x_2); +return x_11; } -else +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___spec__1(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: { -uint8_t x_376; -lean_dec(x_369); -lean_dec(x_13); +uint8_t x_14; +x_14 = lean_usize_dec_lt(x_3, x_2); +if (x_14 == 0) +{ +lean_object* x_15; lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); @@ -6752,35 +6064,19 @@ lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_2); +lean_dec(x_5); lean_dec(x_1); -x_376 = !lean_is_exclusive(x_370); -if (x_376 == 0) -{ -return x_370; +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_4); +lean_ctor_set(x_15, 1, x_13); +return x_15; } else { -lean_object* x_377; lean_object* x_378; lean_object* x_379; -x_377 = lean_ctor_get(x_370, 0); -x_378 = lean_ctor_get(x_370, 1); -lean_inc(x_378); -lean_inc(x_377); -lean_dec(x_370); -x_379 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_379, 0, x_377); -lean_ctor_set(x_379, 1, x_378); -return x_379; -} -} -} -} -} -block_347: -{ -lean_object* x_337; -lean_dec(x_336); -lean_inc(x_13); +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_16 = lean_array_uget(x_4, x_3); +x_17 = lean_unsigned_to_nat(0u); +x_18 = lean_array_uset(x_4, x_3, x_17); lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); @@ -6788,33 +6084,29 @@ lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -lean_inc(x_2); -lean_inc(x_3); -x_337 = l_Lean_Meta_Grind_internalize(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_337) == 0) -{ -lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; -x_338 = lean_ctor_get(x_337, 1); -lean_inc(x_338); -lean_dec(x_337); +lean_inc(x_5); lean_inc(x_1); -x_339 = l_Lean_Meta_Grind_registerParent(x_1, x_3, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_338); -lean_dec(x_3); -x_340 = lean_ctor_get(x_339, 0); -lean_inc(x_340); -x_341 = lean_ctor_get(x_339, 1); -lean_inc(x_341); -lean_dec(x_339); -x_342 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(x_4, x_1, x_2, x_335, x_340, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_341); -lean_dec(x_340); -lean_dec(x_4); -return x_342; +x_19 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern(x_16, x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_19) == 0) +{ +lean_object* x_20; lean_object* x_21; size_t x_22; size_t x_23; lean_object* x_24; +x_20 = lean_ctor_get(x_19, 0); +lean_inc(x_20); +x_21 = lean_ctor_get(x_19, 1); +lean_inc(x_21); +lean_dec(x_19); +x_22 = 1; +x_23 = lean_usize_add(x_3, x_22); +x_24 = lean_array_uset(x_18, x_3, x_20); +x_3 = x_23; +x_4 = x_24; +x_13 = x_21; +goto _start; } else { -uint8_t x_343; -lean_dec(x_335); -lean_dec(x_13); +uint8_t x_26; +lean_dec(x_18); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); @@ -6822,77 +6114,4405 @@ 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_5); lean_dec(x_1); -x_343 = !lean_is_exclusive(x_337); -if (x_343 == 0) +x_26 = !lean_is_exclusive(x_19); +if (x_26 == 0) { -return x_337; +return x_19; } else { -lean_object* x_344; lean_object* x_345; lean_object* x_346; -x_344 = lean_ctor_get(x_337, 0); -x_345 = lean_ctor_get(x_337, 1); -lean_inc(x_345); -lean_inc(x_344); -lean_dec(x_337); -x_346 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_346, 0, x_344); -lean_ctor_set(x_346, 1, x_345); -return x_346; +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_19, 0); +x_28 = lean_ctor_get(x_19, 1); +lean_inc(x_28); +lean_inc(x_27); +lean_dec(x_19); +x_29 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_29, 0, x_27); +lean_ctor_set(x_29, 1, x_28); +return x_29; } } } } -case 9: +} +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___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) { +_start: { -lean_object* x_380; lean_object* x_381; lean_object* x_393; uint8_t x_394; -lean_dec(x_5); -lean_inc(x_2); -lean_inc(x_1); -x_380 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1___boxed), 12, 2); -lean_closure_set(x_380, 0, x_1); -lean_closure_set(x_380, 1, x_2); -x_393 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__4; -x_394 = l_Lean_Expr_isConstOf(x_3, x_393); -if (x_394 == 0) +if (lean_obj_tag(x_2) == 5) { -lean_object* x_395; -x_395 = lean_box(0); -x_381 = x_395; -goto block_392; +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_14 = lean_ctor_get(x_2, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_2, 1); +lean_inc(x_15); +lean_dec(x_2); +x_16 = lean_array_set(x_3, x_4, x_15); +x_17 = lean_unsigned_to_nat(1u); +x_18 = lean_nat_sub(x_4, x_17); +lean_dec(x_4); +x_2 = x_14; +x_3 = x_16; +x_4 = x_18; +goto _start; } else { -lean_object* x_396; lean_object* x_397; uint8_t x_398; -x_396 = lean_array_get_size(x_4); -x_397 = lean_unsigned_to_nat(2u); -x_398 = lean_nat_dec_eq(x_396, x_397); -if (x_398 == 0) +size_t x_20; size_t x_21; lean_object* x_22; +lean_dec(x_4); +x_20 = lean_array_size(x_3); +x_21 = 0; +x_22 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___spec__1(x_1, x_20, x_21, x_3, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_22) == 0) { -lean_object* x_399; -lean_dec(x_396); -x_399 = lean_box(0); -x_381 = x_399; -goto block_392; +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 = l_Lean_mkAppN(x_2, x_24); +lean_dec(x_24); +lean_ctor_set(x_22, 0, x_25); +return x_22; } else { -lean_object* x_400; uint8_t x_401; -lean_dec(x_380); -lean_dec(x_3); -x_400 = lean_unsigned_to_nat(0u); -x_401 = lean_nat_dec_lt(x_400, x_396); -lean_dec(x_396); -if (x_401 == 0) +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 = l_Lean_mkAppN(x_2, x_26); +lean_dec(x_26); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_27); +return x_29; +} +} +else +{ +uint8_t x_30; +lean_dec(x_2); +x_30 = !lean_is_exclusive(x_22); +if (x_30 == 0) +{ +return x_22; +} +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_22, 0); +x_32 = lean_ctor_get(x_22, 1); +lean_inc(x_32); +lean_inc(x_31); +lean_dec(x_22); +x_33 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_33, 0, x_31); +lean_ctor_set(x_33, 1, x_32); +return x_33; +} +} +} +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_levelZero; +x_2 = l_Lean_Expr_sort___override(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_unfoldReducible___lambda__2), 6, 0); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_unfoldReducible___lambda__3___boxed), 6, 0); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_normalizeLevels___lambda__1___boxed), 4, 0); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_eraseIrrelevantMData___lambda__2___boxed), 4, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, 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 = l_Lean_Expr_isBVar(x_1); +if (x_12 == 0) +{ +lean_object* x_13; uint8_t x_14; +x_13 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare; +x_14 = lean_expr_eqv(x_1, x_13); +if (x_14 == 0) +{ +lean_object* x_15; +x_15 = l_Lean_Meta_Grind_groundPattern_x3f(x_1); +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; +x_16 = lean_unsigned_to_nat(0u); +x_17 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_16); +x_18 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__1; +lean_inc(x_17); +x_19 = lean_mk_array(x_17, x_18); +x_20 = lean_unsigned_to_nat(1u); +x_21 = lean_nat_sub(x_17, x_20); +lean_dec(x_17); +x_22 = l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___spec__2(x_2, x_1, x_19, x_21, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +return x_22; +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +lean_dec(x_1); +x_23 = lean_ctor_get(x_15, 0); +lean_inc(x_23); +lean_dec(x_15); +x_24 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__2; +x_25 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__3; +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +x_26 = l_Lean_Core_transform___at_Lean_Meta_Grind_unfoldReducible___spec__1(x_23, x_24, x_25, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_26) == 0) +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_27 = lean_ctor_get(x_26, 0); +lean_inc(x_27); +x_28 = lean_ctor_get(x_26, 1); +lean_inc(x_28); +lean_dec(x_26); +x_29 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__4; +x_30 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__5; +lean_inc(x_10); +lean_inc(x_9); +x_31 = l_Lean_Core_transform___at_Lean_Core_betaReduce___spec__1(x_27, x_29, x_30, x_9, x_10, x_28); +if (lean_obj_tag(x_31) == 0) +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_31, 0); +lean_inc(x_32); +x_33 = lean_ctor_get(x_31, 1); +lean_inc(x_33); +lean_dec(x_31); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +x_34 = l_Lean_Meta_Grind_canon(x_32, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_33); +if (lean_obj_tag(x_34) == 0) +{ +lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_35 = lean_ctor_get(x_34, 0); +lean_inc(x_35); +x_36 = lean_ctor_get(x_34, 1); +lean_inc(x_36); +lean_dec(x_34); +x_37 = l_Lean_Meta_Grind_shareCommon(x_35, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_36); +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); +lean_inc(x_38); +x_40 = l_Lean_Meta_Grind_internalize(x_38, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_39); +if (lean_obj_tag(x_40) == 0) +{ +uint8_t x_41; +x_41 = !lean_is_exclusive(x_40); +if (x_41 == 0) +{ +lean_object* x_42; lean_object* x_43; +x_42 = lean_ctor_get(x_40, 0); +lean_dec(x_42); +x_43 = l_Lean_Meta_Grind_mkGroundPattern(x_38); +lean_ctor_set(x_40, 0, x_43); +return x_40; +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_40, 1); +lean_inc(x_44); +lean_dec(x_40); +x_45 = l_Lean_Meta_Grind_mkGroundPattern(x_38); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_44); +return x_46; +} +} +else +{ +uint8_t x_47; +lean_dec(x_38); +x_47 = !lean_is_exclusive(x_40); +if (x_47 == 0) +{ +return x_40; +} +else +{ +lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_48 = lean_ctor_get(x_40, 0); +x_49 = lean_ctor_get(x_40, 1); +lean_inc(x_49); +lean_inc(x_48); +lean_dec(x_40); +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_51; +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_51 = !lean_is_exclusive(x_34); +if (x_51 == 0) +{ +return x_34; +} +else +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; +x_52 = lean_ctor_get(x_34, 0); +x_53 = lean_ctor_get(x_34, 1); +lean_inc(x_53); +lean_inc(x_52); +lean_dec(x_34); +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_55; +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_55 = !lean_is_exclusive(x_31); +if (x_55 == 0) +{ +return x_31; +} +else +{ +lean_object* x_56; lean_object* x_57; lean_object* x_58; +x_56 = lean_ctor_get(x_31, 0); +x_57 = lean_ctor_get(x_31, 1); +lean_inc(x_57); +lean_inc(x_56); +lean_dec(x_31); +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; +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_59 = !lean_is_exclusive(x_26); +if (x_59 == 0) +{ +return x_26; +} +else +{ +lean_object* x_60; lean_object* x_61; lean_object* x_62; +x_60 = lean_ctor_get(x_26, 0); +x_61 = lean_ctor_get(x_26, 1); +lean_inc(x_61); +lean_inc(x_60); +lean_dec(x_26); +x_62 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_62, 0, x_60); +lean_ctor_set(x_62, 1, x_61); +return x_62; +} +} +} +} +else +{ +lean_object* x_63; +lean_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_63 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_63, 0, x_1); +lean_ctor_set(x_63, 1, x_11); +return x_63; +} +} +else +{ +lean_object* x_64; +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_64 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_64, 0, x_1); +lean_ctor_set(x_64, 1, x_11); +return x_64; +} +} +} +static lean_object* _init_l_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_instMonadMetaM; +x_2 = l_ReaderT_instMonad___rarg(x_1); +return x_2; +} +} +static lean_object* _init_l_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__1; +x_2 = l_ReaderT_instMonad___rarg(x_1); +return x_2; +} +} +static lean_object* _init_l_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__2; +x_2 = l_ReaderT_instMonad___rarg(x_1); +return x_2; +} +} +static lean_object* _init_l_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__3; +x_2 = l_ReaderT_instMonad___rarg(x_1); +return x_2; +} +} +static lean_object* _init_l_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__4; +x_2 = l_instInhabitedPUnit; +x_3 = l_instInhabitedOfMonad___rarg(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_panic___at_Lean_Meta_Grind_internalize___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_11 = l_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__5; +x_12 = lean_panic_fn(x_11, x_1); +x_13 = lean_apply_9(x_12, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_internalize___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18) { +_start: +{ +lean_object* x_19; uint8_t x_20; +x_19 = lean_ctor_get(x_5, 1); +x_20 = lean_nat_dec_lt(x_7, x_19); +if (x_20 == 0) +{ +lean_object* x_21; +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_7); +lean_dec(x_2); +lean_dec(x_1); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_6); +lean_ctor_set(x_21, 1, x_18); +return x_21; +} +else +{ +lean_object* x_22; lean_object* x_23; +lean_dec(x_6); +x_22 = lean_array_fget(x_3, x_7); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_2); +lean_inc(x_22); +x_23 = l_Lean_Meta_Grind_internalize(x_22, x_2, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +if (lean_obj_tag(x_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, 1); +lean_inc(x_24); +lean_dec(x_23); +lean_inc(x_1); +x_25 = l_Lean_Meta_Grind_registerParent(x_1, x_22, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_24); +lean_dec(x_22); +x_26 = lean_ctor_get(x_25, 1); +lean_inc(x_26); +lean_dec(x_25); +x_27 = lean_ctor_get(x_5, 2); +x_28 = lean_nat_add(x_7, x_27); +lean_dec(x_7); +x_29 = lean_box(0); +x_6 = x_29; +x_7 = x_28; +x_8 = lean_box(0); +x_9 = lean_box(0); +x_18 = x_26; +goto _start; +} +else +{ +uint8_t x_31; +lean_dec(x_22); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_7); +lean_dec(x_2); +lean_dec(x_1); +x_31 = !lean_is_exclusive(x_23); +if (x_31 == 0) +{ +return x_23; +} +else +{ +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); +return x_34; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___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) { +_start: +{ +lean_object* x_13; +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_1); +x_13 = l_Lean_Meta_Grind_mkENode(x_1, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_13) == 0) +{ +lean_object* x_14; lean_object* x_15; +x_14 = lean_ctor_get(x_13, 1); +lean_inc(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_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_1); +x_15 = l_Lean_Meta_Grind_addCongrTable(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_14); +if (lean_obj_tag(x_15) == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_16 = lean_ctor_get(x_15, 1); +lean_inc(x_16); +lean_dec(x_15); +lean_inc(x_1); +x_17 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_16); +x_18 = lean_ctor_get(x_17, 1); +lean_inc(x_18); +lean_dec(x_17); +x_19 = l_Lean_Meta_Grind_propagateUp(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_18); +return x_19; +} +else +{ +uint8_t x_20; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_20 = !lean_is_exclusive(x_15); +if (x_20 == 0) +{ +return x_15; +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_15, 0); +x_22 = lean_ctor_get(x_15, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_15); +x_23 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_23, 0, x_21); +lean_ctor_set(x_23, 1, x_22); +return x_23; +} +} +} +else +{ +uint8_t x_24; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_24 = !lean_is_exclusive(x_13); +if (x_24 == 0) +{ +return x_13; +} +else +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_25 = lean_ctor_get(x_13, 0); +x_26 = lean_ctor_get(x_13, 1); +lean_inc(x_26); +lean_inc(x_25); +lean_dec(x_13); +x_27 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +return x_27; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___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) { +_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; +x_15 = lean_array_get_size(x_1); +x_16 = lean_unsigned_to_nat(0u); +x_17 = lean_unsigned_to_nat(1u); +x_18 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_18, 0, x_16); +lean_ctor_set(x_18, 1, x_15); +lean_ctor_set(x_18, 2, x_17); +x_19 = 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); +lean_inc(x_7); +lean_inc(x_6); +x_20 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_internalize___spec__2(x_2, x_3, x_1, x_18, x_18, x_19, x_16, lean_box(0), lean_box(0), x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +lean_dec(x_18); +if (lean_obj_tag(x_20) == 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 = lean_apply_10(x_4, x_19, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_21); +return x_22; +} +else +{ +uint8_t 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_6); +lean_dec(x_4); +x_23 = !lean_is_exclusive(x_20); +if (x_23 == 0) +{ +return x_20; +} +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_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; +} +} +} +} +static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("nestedProof", 11, 11); +return x_1; +} +} +static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__1___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__1___closed__2; +x_3 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__1; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +switch (lean_obj_tag(x_3)) { +case 0: +{ +lean_object* x_15; +lean_dec(x_5); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_1); +x_15 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate(x_1, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_15) == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_16 = lean_ctor_get(x_15, 1); +lean_inc(x_16); +lean_dec(x_15); +lean_inc(x_1); +x_17 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs(x_1, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_16); +x_18 = lean_ctor_get(x_17, 1); +lean_inc(x_18); +lean_dec(x_17); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_3); +x_19 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_18); +if (lean_obj_tag(x_19) == 0) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_34; uint8_t x_35; +x_20 = lean_ctor_get(x_19, 1); +lean_inc(x_20); +lean_dec(x_19); +lean_inc(x_2); +lean_inc(x_1); +x_21 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1___boxed), 12, 2); +lean_closure_set(x_21, 0, x_1); +lean_closure_set(x_21, 1, x_2); +x_34 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__2; +x_35 = l_Lean_Expr_isConstOf(x_3, x_34); +if (x_35 == 0) +{ +lean_object* x_36; +x_36 = lean_box(0); +x_22 = x_36; +goto block_33; +} +else +{ +lean_object* x_37; lean_object* x_38; uint8_t x_39; +x_37 = lean_array_get_size(x_4); +x_38 = lean_unsigned_to_nat(2u); +x_39 = lean_nat_dec_eq(x_37, x_38); +if (x_39 == 0) +{ +lean_object* x_40; +lean_dec(x_37); +x_40 = lean_box(0); +x_22 = x_40; +goto block_33; +} +else +{ +lean_object* x_41; uint8_t x_42; +lean_dec(x_21); +lean_dec(x_3); +x_41 = lean_unsigned_to_nat(0u); +x_42 = lean_nat_dec_lt(x_41, x_37); +lean_dec(x_37); +if (x_42 == 0) +{ +lean_object* x_43; lean_object* x_44; lean_object* x_45; +lean_dec(x_4); +x_43 = l_Lean_instInhabitedExpr; +x_44 = l_outOfBounds___rarg(x_43); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_44); +x_45 = l_Lean_Meta_Grind_internalize(x_44, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_20); +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; +x_46 = lean_ctor_get(x_45, 1); +lean_inc(x_46); +lean_dec(x_45); +lean_inc(x_1); +x_47 = l_Lean_Meta_Grind_registerParent(x_1, x_44, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_46); +lean_dec(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); +lean_dec(x_47); +x_50 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_48, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_49); +lean_dec(x_48); +return x_50; +} +else +{ +uint8_t x_51; +lean_dec(x_44); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_2); +lean_dec(x_1); +x_51 = !lean_is_exclusive(x_45); +if (x_51 == 0) +{ +return x_45; +} +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; +} +} +} +else +{ +lean_object* x_55; lean_object* x_56; +x_55 = lean_array_fget(x_4, x_41); +lean_dec(x_4); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_55); +x_56 = l_Lean_Meta_Grind_internalize(x_55, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_20); +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, 1); +lean_inc(x_57); +lean_dec(x_56); +lean_inc(x_1); +x_58 = l_Lean_Meta_Grind_registerParent(x_1, x_55, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_57); +lean_dec(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); +x_61 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_59, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_60); +lean_dec(x_59); +return x_61; +} +else +{ +uint8_t x_62; +lean_dec(x_55); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_2); +lean_dec(x_1); +x_62 = !lean_is_exclusive(x_56); +if (x_62 == 0) +{ +return x_56; +} +else +{ +lean_object* x_63; lean_object* x_64; lean_object* x_65; +x_63 = lean_ctor_get(x_56, 0); +x_64 = lean_ctor_get(x_56, 1); +lean_inc(x_64); +lean_inc(x_63); +lean_dec(x_56); +x_65 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_65, 0, x_63); +lean_ctor_set(x_65, 1, x_64); +return x_65; +} +} +} +} +} +block_33: +{ +lean_object* x_23; +lean_dec(x_22); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_3); +x_23 = l_Lean_Meta_Grind_internalize(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, 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; +x_24 = lean_ctor_get(x_23, 1); +lean_inc(x_24); +lean_dec(x_23); +lean_inc(x_1); +x_25 = l_Lean_Meta_Grind_registerParent(x_1, x_3, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_24); +lean_dec(x_3); +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_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(x_4, x_1, x_2, x_21, x_26, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_27); +lean_dec(x_26); +lean_dec(x_4); +return x_28; +} +else +{ +uint8_t x_29; +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_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_29 = !lean_is_exclusive(x_23); +if (x_29 == 0) +{ +return x_23; +} +else +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_30 = lean_ctor_get(x_23, 0); +x_31 = lean_ctor_get(x_23, 1); +lean_inc(x_31); +lean_inc(x_30); +lean_dec(x_23); +x_32 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_32, 0, x_30); +lean_ctor_set(x_32, 1, x_31); +return x_32; +} +} +} +} +else +{ +uint8_t x_66; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_66 = !lean_is_exclusive(x_19); +if (x_66 == 0) +{ +return x_19; +} +else +{ +lean_object* x_67; lean_object* x_68; lean_object* x_69; +x_67 = lean_ctor_get(x_19, 0); +x_68 = lean_ctor_get(x_19, 1); +lean_inc(x_68); +lean_inc(x_67); +lean_dec(x_19); +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_70; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_70 = !lean_is_exclusive(x_15); +if (x_70 == 0) +{ +return x_15; +} +else +{ +lean_object* x_71; lean_object* x_72; lean_object* x_73; +x_71 = lean_ctor_get(x_15, 0); +x_72 = lean_ctor_get(x_15, 1); +lean_inc(x_72); +lean_inc(x_71); +lean_dec(x_15); +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; +} +} +} +case 1: +{ +lean_object* x_74; +lean_dec(x_5); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_1); +x_74 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate(x_1, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_74) == 0) +{ +lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; +x_75 = lean_ctor_get(x_74, 1); +lean_inc(x_75); +lean_dec(x_74); +lean_inc(x_1); +x_76 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs(x_1, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_75); +x_77 = lean_ctor_get(x_76, 1); +lean_inc(x_77); +lean_dec(x_76); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_3); +x_78 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_77); +if (lean_obj_tag(x_78) == 0) +{ +lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_93; uint8_t x_94; +x_79 = lean_ctor_get(x_78, 1); +lean_inc(x_79); +lean_dec(x_78); +lean_inc(x_2); +lean_inc(x_1); +x_80 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1___boxed), 12, 2); +lean_closure_set(x_80, 0, x_1); +lean_closure_set(x_80, 1, x_2); +x_93 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__2; +x_94 = l_Lean_Expr_isConstOf(x_3, x_93); +if (x_94 == 0) +{ +lean_object* x_95; +x_95 = lean_box(0); +x_81 = x_95; +goto block_92; +} +else +{ +lean_object* x_96; lean_object* x_97; uint8_t x_98; +x_96 = lean_array_get_size(x_4); +x_97 = lean_unsigned_to_nat(2u); +x_98 = lean_nat_dec_eq(x_96, x_97); +if (x_98 == 0) +{ +lean_object* x_99; +lean_dec(x_96); +x_99 = lean_box(0); +x_81 = x_99; +goto block_92; +} +else +{ +lean_object* x_100; uint8_t x_101; +lean_dec(x_80); +lean_dec(x_3); +x_100 = lean_unsigned_to_nat(0u); +x_101 = lean_nat_dec_lt(x_100, x_96); +lean_dec(x_96); +if (x_101 == 0) +{ +lean_object* x_102; lean_object* x_103; lean_object* x_104; +lean_dec(x_4); +x_102 = l_Lean_instInhabitedExpr; +x_103 = l_outOfBounds___rarg(x_102); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_103); +x_104 = l_Lean_Meta_Grind_internalize(x_103, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_79); +if (lean_obj_tag(x_104) == 0) +{ +lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; +x_105 = lean_ctor_get(x_104, 1); +lean_inc(x_105); +lean_dec(x_104); +lean_inc(x_1); +x_106 = l_Lean_Meta_Grind_registerParent(x_1, x_103, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_105); +lean_dec(x_103); +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 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_107, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_108); +lean_dec(x_107); +return x_109; +} +else +{ +uint8_t x_110; +lean_dec(x_103); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_2); +lean_dec(x_1); +x_110 = !lean_is_exclusive(x_104); +if (x_110 == 0) +{ +return x_104; +} +else +{ +lean_object* x_111; lean_object* x_112; lean_object* x_113; +x_111 = lean_ctor_get(x_104, 0); +x_112 = lean_ctor_get(x_104, 1); +lean_inc(x_112); +lean_inc(x_111); +lean_dec(x_104); +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_114; lean_object* x_115; +x_114 = lean_array_fget(x_4, x_100); +lean_dec(x_4); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_114); +x_115 = l_Lean_Meta_Grind_internalize(x_114, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_79); +if (lean_obj_tag(x_115) == 0) +{ +lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; +x_116 = lean_ctor_get(x_115, 1); +lean_inc(x_116); +lean_dec(x_115); +lean_inc(x_1); +x_117 = l_Lean_Meta_Grind_registerParent(x_1, x_114, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_116); +lean_dec(x_114); +x_118 = lean_ctor_get(x_117, 0); +lean_inc(x_118); +x_119 = lean_ctor_get(x_117, 1); +lean_inc(x_119); +lean_dec(x_117); +x_120 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_118, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_119); +lean_dec(x_118); +return x_120; +} +else +{ +uint8_t x_121; +lean_dec(x_114); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_2); +lean_dec(x_1); +x_121 = !lean_is_exclusive(x_115); +if (x_121 == 0) +{ +return x_115; +} +else +{ +lean_object* x_122; lean_object* x_123; lean_object* x_124; +x_122 = lean_ctor_get(x_115, 0); +x_123 = lean_ctor_get(x_115, 1); +lean_inc(x_123); +lean_inc(x_122); +lean_dec(x_115); +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; +} +} +} +} +} +block_92: +{ +lean_object* x_82; +lean_dec(x_81); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_3); +x_82 = l_Lean_Meta_Grind_internalize(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_79); +if (lean_obj_tag(x_82) == 0) +{ +lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; +x_83 = lean_ctor_get(x_82, 1); +lean_inc(x_83); +lean_dec(x_82); +lean_inc(x_1); +x_84 = l_Lean_Meta_Grind_registerParent(x_1, x_3, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_83); +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_87 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(x_4, x_1, x_2, x_80, x_85, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_86); +lean_dec(x_85); +lean_dec(x_4); +return x_87; +} +else +{ +uint8_t x_88; +lean_dec(x_80); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_88 = !lean_is_exclusive(x_82); +if (x_88 == 0) +{ +return x_82; +} +else +{ +lean_object* x_89; lean_object* x_90; lean_object* x_91; +x_89 = lean_ctor_get(x_82, 0); +x_90 = lean_ctor_get(x_82, 1); +lean_inc(x_90); +lean_inc(x_89); +lean_dec(x_82); +x_91 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_91, 0, x_89); +lean_ctor_set(x_91, 1, x_90); +return x_91; +} +} +} +} +else +{ +uint8_t x_125; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_125 = !lean_is_exclusive(x_78); +if (x_125 == 0) +{ +return x_78; +} +else +{ +lean_object* x_126; lean_object* x_127; lean_object* x_128; +x_126 = lean_ctor_get(x_78, 0); +x_127 = lean_ctor_get(x_78, 1); +lean_inc(x_127); +lean_inc(x_126); +lean_dec(x_78); +x_128 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_128, 0, x_126); +lean_ctor_set(x_128, 1, x_127); +return x_128; +} +} +} +else +{ +uint8_t x_129; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_129 = !lean_is_exclusive(x_74); +if (x_129 == 0) +{ +return x_74; +} +else +{ +lean_object* x_130; lean_object* x_131; lean_object* x_132; +x_130 = lean_ctor_get(x_74, 0); +x_131 = lean_ctor_get(x_74, 1); +lean_inc(x_131); +lean_inc(x_130); +lean_dec(x_74); +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; +} +} +} +case 2: +{ +lean_object* x_133; +lean_dec(x_5); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_1); +x_133 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate(x_1, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_133) == 0) +{ +lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; +x_134 = lean_ctor_get(x_133, 1); +lean_inc(x_134); +lean_dec(x_133); +lean_inc(x_1); +x_135 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs(x_1, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_134); +x_136 = lean_ctor_get(x_135, 1); +lean_inc(x_136); +lean_dec(x_135); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_3); +x_137 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_136); +if (lean_obj_tag(x_137) == 0) +{ +lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_152; uint8_t x_153; +x_138 = lean_ctor_get(x_137, 1); +lean_inc(x_138); +lean_dec(x_137); +lean_inc(x_2); +lean_inc(x_1); +x_139 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1___boxed), 12, 2); +lean_closure_set(x_139, 0, x_1); +lean_closure_set(x_139, 1, x_2); +x_152 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__2; +x_153 = l_Lean_Expr_isConstOf(x_3, x_152); +if (x_153 == 0) +{ +lean_object* x_154; +x_154 = lean_box(0); +x_140 = x_154; +goto block_151; +} +else +{ +lean_object* x_155; lean_object* x_156; uint8_t x_157; +x_155 = lean_array_get_size(x_4); +x_156 = lean_unsigned_to_nat(2u); +x_157 = lean_nat_dec_eq(x_155, x_156); +if (x_157 == 0) +{ +lean_object* x_158; +lean_dec(x_155); +x_158 = lean_box(0); +x_140 = x_158; +goto block_151; +} +else +{ +lean_object* x_159; uint8_t x_160; +lean_dec(x_139); +lean_dec(x_3); +x_159 = lean_unsigned_to_nat(0u); +x_160 = lean_nat_dec_lt(x_159, x_155); +lean_dec(x_155); +if (x_160 == 0) +{ +lean_object* x_161; lean_object* x_162; lean_object* x_163; +lean_dec(x_4); +x_161 = l_Lean_instInhabitedExpr; +x_162 = l_outOfBounds___rarg(x_161); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_162); +x_163 = l_Lean_Meta_Grind_internalize(x_162, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_138); +if (lean_obj_tag(x_163) == 0) +{ +lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; +x_164 = lean_ctor_get(x_163, 1); +lean_inc(x_164); +lean_dec(x_163); +lean_inc(x_1); +x_165 = l_Lean_Meta_Grind_registerParent(x_1, x_162, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_164); +lean_dec(x_162); +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_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_166, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_167); +lean_dec(x_166); +return x_168; +} +else +{ +uint8_t x_169; +lean_dec(x_162); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_2); +lean_dec(x_1); +x_169 = !lean_is_exclusive(x_163); +if (x_169 == 0) +{ +return x_163; +} +else +{ +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_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); +return x_172; +} +} +} +else +{ +lean_object* x_173; lean_object* x_174; +x_173 = lean_array_fget(x_4, x_159); +lean_dec(x_4); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_173); +x_174 = l_Lean_Meta_Grind_internalize(x_173, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_138); +if (lean_obj_tag(x_174) == 0) +{ +lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; +x_175 = lean_ctor_get(x_174, 1); +lean_inc(x_175); +lean_dec(x_174); +lean_inc(x_1); +x_176 = l_Lean_Meta_Grind_registerParent(x_1, x_173, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_175); +lean_dec(x_173); +x_177 = lean_ctor_get(x_176, 0); +lean_inc(x_177); +x_178 = lean_ctor_get(x_176, 1); +lean_inc(x_178); +lean_dec(x_176); +x_179 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_177, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_178); +lean_dec(x_177); +return x_179; +} +else +{ +uint8_t x_180; +lean_dec(x_173); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_2); +lean_dec(x_1); +x_180 = !lean_is_exclusive(x_174); +if (x_180 == 0) +{ +return x_174; +} +else +{ +lean_object* x_181; lean_object* x_182; lean_object* x_183; +x_181 = lean_ctor_get(x_174, 0); +x_182 = lean_ctor_get(x_174, 1); +lean_inc(x_182); +lean_inc(x_181); +lean_dec(x_174); +x_183 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_183, 0, x_181); +lean_ctor_set(x_183, 1, x_182); +return x_183; +} +} +} +} +} +block_151: +{ +lean_object* x_141; +lean_dec(x_140); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_3); +x_141 = l_Lean_Meta_Grind_internalize(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_138); +if (lean_obj_tag(x_141) == 0) +{ +lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; +x_142 = lean_ctor_get(x_141, 1); +lean_inc(x_142); +lean_dec(x_141); +lean_inc(x_1); +x_143 = l_Lean_Meta_Grind_registerParent(x_1, x_3, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_142); +lean_dec(x_3); +x_144 = lean_ctor_get(x_143, 0); +lean_inc(x_144); +x_145 = lean_ctor_get(x_143, 1); +lean_inc(x_145); +lean_dec(x_143); +x_146 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(x_4, x_1, x_2, x_139, x_144, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_145); +lean_dec(x_144); +lean_dec(x_4); +return x_146; +} +else +{ +uint8_t x_147; +lean_dec(x_139); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_147 = !lean_is_exclusive(x_141); +if (x_147 == 0) +{ +return x_141; +} +else +{ +lean_object* x_148; lean_object* x_149; lean_object* x_150; +x_148 = lean_ctor_get(x_141, 0); +x_149 = lean_ctor_get(x_141, 1); +lean_inc(x_149); +lean_inc(x_148); +lean_dec(x_141); +x_150 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_150, 0, x_148); +lean_ctor_set(x_150, 1, x_149); +return x_150; +} +} +} +} +else +{ +uint8_t x_184; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_184 = !lean_is_exclusive(x_137); +if (x_184 == 0) +{ +return x_137; +} +else +{ +lean_object* x_185; lean_object* x_186; lean_object* x_187; +x_185 = lean_ctor_get(x_137, 0); +x_186 = lean_ctor_get(x_137, 1); +lean_inc(x_186); +lean_inc(x_185); +lean_dec(x_137); +x_187 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_187, 0, x_185); +lean_ctor_set(x_187, 1, x_186); +return x_187; +} +} +} +else +{ +uint8_t x_188; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_188 = !lean_is_exclusive(x_133); +if (x_188 == 0) +{ +return x_133; +} +else +{ +lean_object* x_189; lean_object* x_190; lean_object* x_191; +x_189 = lean_ctor_get(x_133, 0); +x_190 = lean_ctor_get(x_133, 1); +lean_inc(x_190); +lean_inc(x_189); +lean_dec(x_133); +x_191 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_191, 0, x_189); +lean_ctor_set(x_191, 1, x_190); +return x_191; +} +} +} +case 3: +{ +lean_object* x_192; +lean_dec(x_5); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_1); +x_192 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate(x_1, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_192) == 0) +{ +lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; +x_193 = lean_ctor_get(x_192, 1); +lean_inc(x_193); +lean_dec(x_192); +lean_inc(x_1); +x_194 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs(x_1, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_193); +x_195 = lean_ctor_get(x_194, 1); +lean_inc(x_195); +lean_dec(x_194); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_3); +x_196 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_195); +if (lean_obj_tag(x_196) == 0) +{ +lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_211; uint8_t x_212; +x_197 = lean_ctor_get(x_196, 1); +lean_inc(x_197); +lean_dec(x_196); +lean_inc(x_2); +lean_inc(x_1); +x_198 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1___boxed), 12, 2); +lean_closure_set(x_198, 0, x_1); +lean_closure_set(x_198, 1, x_2); +x_211 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__2; +x_212 = l_Lean_Expr_isConstOf(x_3, x_211); +if (x_212 == 0) +{ +lean_object* x_213; +x_213 = lean_box(0); +x_199 = x_213; +goto block_210; +} +else +{ +lean_object* x_214; lean_object* x_215; uint8_t x_216; +x_214 = lean_array_get_size(x_4); +x_215 = lean_unsigned_to_nat(2u); +x_216 = lean_nat_dec_eq(x_214, x_215); +if (x_216 == 0) +{ +lean_object* x_217; +lean_dec(x_214); +x_217 = lean_box(0); +x_199 = x_217; +goto block_210; +} +else +{ +lean_object* x_218; uint8_t x_219; +lean_dec(x_198); +lean_dec(x_3); +x_218 = lean_unsigned_to_nat(0u); +x_219 = lean_nat_dec_lt(x_218, x_214); +lean_dec(x_214); +if (x_219 == 0) +{ +lean_object* x_220; lean_object* x_221; lean_object* x_222; +lean_dec(x_4); +x_220 = l_Lean_instInhabitedExpr; +x_221 = l_outOfBounds___rarg(x_220); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_221); +x_222 = l_Lean_Meta_Grind_internalize(x_221, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_197); +if (lean_obj_tag(x_222) == 0) +{ +lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; +x_223 = lean_ctor_get(x_222, 1); +lean_inc(x_223); +lean_dec(x_222); +lean_inc(x_1); +x_224 = l_Lean_Meta_Grind_registerParent(x_1, x_221, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_223); +lean_dec(x_221); +x_225 = lean_ctor_get(x_224, 0); +lean_inc(x_225); +x_226 = lean_ctor_get(x_224, 1); +lean_inc(x_226); +lean_dec(x_224); +x_227 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_225, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_226); +lean_dec(x_225); +return x_227; +} +else +{ +uint8_t x_228; +lean_dec(x_221); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_2); +lean_dec(x_1); +x_228 = !lean_is_exclusive(x_222); +if (x_228 == 0) +{ +return x_222; +} +else +{ +lean_object* x_229; lean_object* x_230; lean_object* x_231; +x_229 = lean_ctor_get(x_222, 0); +x_230 = lean_ctor_get(x_222, 1); +lean_inc(x_230); +lean_inc(x_229); +lean_dec(x_222); +x_231 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_231, 0, x_229); +lean_ctor_set(x_231, 1, x_230); +return x_231; +} +} +} +else +{ +lean_object* x_232; lean_object* x_233; +x_232 = lean_array_fget(x_4, x_218); +lean_dec(x_4); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_232); +x_233 = l_Lean_Meta_Grind_internalize(x_232, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_197); +if (lean_obj_tag(x_233) == 0) +{ +lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; +x_234 = lean_ctor_get(x_233, 1); +lean_inc(x_234); +lean_dec(x_233); +lean_inc(x_1); +x_235 = l_Lean_Meta_Grind_registerParent(x_1, x_232, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_234); +lean_dec(x_232); +x_236 = lean_ctor_get(x_235, 0); +lean_inc(x_236); +x_237 = lean_ctor_get(x_235, 1); +lean_inc(x_237); +lean_dec(x_235); +x_238 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_236, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_237); +lean_dec(x_236); +return x_238; +} +else +{ +uint8_t x_239; +lean_dec(x_232); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_2); +lean_dec(x_1); +x_239 = !lean_is_exclusive(x_233); +if (x_239 == 0) +{ +return x_233; +} +else +{ +lean_object* x_240; lean_object* x_241; lean_object* x_242; +x_240 = lean_ctor_get(x_233, 0); +x_241 = lean_ctor_get(x_233, 1); +lean_inc(x_241); +lean_inc(x_240); +lean_dec(x_233); +x_242 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_242, 0, x_240); +lean_ctor_set(x_242, 1, x_241); +return x_242; +} +} +} +} +} +block_210: +{ +lean_object* x_200; +lean_dec(x_199); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_3); +x_200 = l_Lean_Meta_Grind_internalize(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_197); +if (lean_obj_tag(x_200) == 0) +{ +lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; +x_201 = lean_ctor_get(x_200, 1); +lean_inc(x_201); +lean_dec(x_200); +lean_inc(x_1); +x_202 = l_Lean_Meta_Grind_registerParent(x_1, x_3, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_201); +lean_dec(x_3); +x_203 = lean_ctor_get(x_202, 0); +lean_inc(x_203); +x_204 = lean_ctor_get(x_202, 1); +lean_inc(x_204); +lean_dec(x_202); +x_205 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(x_4, x_1, x_2, x_198, x_203, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_204); +lean_dec(x_203); +lean_dec(x_4); +return x_205; +} +else +{ +uint8_t x_206; +lean_dec(x_198); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_206 = !lean_is_exclusive(x_200); +if (x_206 == 0) +{ +return x_200; +} +else +{ +lean_object* x_207; lean_object* x_208; lean_object* x_209; +x_207 = lean_ctor_get(x_200, 0); +x_208 = lean_ctor_get(x_200, 1); +lean_inc(x_208); +lean_inc(x_207); +lean_dec(x_200); +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 +{ +uint8_t x_243; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_243 = !lean_is_exclusive(x_196); +if (x_243 == 0) +{ +return x_196; +} +else +{ +lean_object* x_244; lean_object* x_245; lean_object* x_246; +x_244 = lean_ctor_get(x_196, 0); +x_245 = lean_ctor_get(x_196, 1); +lean_inc(x_245); +lean_inc(x_244); +lean_dec(x_196); +x_246 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_246, 0, x_244); +lean_ctor_set(x_246, 1, x_245); +return x_246; +} +} +} +else +{ +uint8_t x_247; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_247 = !lean_is_exclusive(x_192); +if (x_247 == 0) +{ +return x_192; +} +else +{ +lean_object* x_248; lean_object* x_249; lean_object* x_250; +x_248 = lean_ctor_get(x_192, 0); +x_249 = lean_ctor_get(x_192, 1); +lean_inc(x_249); +lean_inc(x_248); +lean_dec(x_192); +x_250 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_250, 0, x_248); +lean_ctor_set(x_250, 1, x_249); +return x_250; +} +} +} +case 4: +{ +lean_object* x_251; lean_object* x_252; +lean_dec(x_5); +x_251 = lean_ctor_get(x_3, 0); +lean_inc(x_251); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_1); +x_252 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate(x_1, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_252) == 0) +{ +lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; +x_253 = lean_ctor_get(x_252, 1); +lean_inc(x_253); +lean_dec(x_252); +lean_inc(x_1); +x_254 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs(x_1, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_253); +x_255 = lean_ctor_get(x_254, 1); +lean_inc(x_255); +lean_dec(x_254); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_3); +x_256 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_255); +if (lean_obj_tag(x_256) == 0) +{ +lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_269; uint8_t x_270; +x_257 = lean_ctor_get(x_256, 1); +lean_inc(x_257); +lean_dec(x_256); +lean_inc(x_2); +lean_inc(x_1); +x_258 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1___boxed), 12, 2); +lean_closure_set(x_258, 0, x_1); +lean_closure_set(x_258, 1, x_2); +x_269 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__2; +x_270 = l_Lean_Expr_isConstOf(x_3, x_269); +lean_dec(x_3); +if (x_270 == 0) +{ +lean_object* x_271; +x_271 = lean_box(0); +x_259 = x_271; +goto block_268; +} +else +{ +lean_object* x_272; lean_object* x_273; uint8_t x_274; +x_272 = lean_array_get_size(x_4); +x_273 = lean_unsigned_to_nat(2u); +x_274 = lean_nat_dec_eq(x_272, x_273); +if (x_274 == 0) +{ +lean_object* x_275; +lean_dec(x_272); +x_275 = lean_box(0); +x_259 = x_275; +goto block_268; +} +else +{ +lean_object* x_276; uint8_t x_277; +lean_dec(x_258); +lean_dec(x_251); +x_276 = lean_unsigned_to_nat(0u); +x_277 = lean_nat_dec_lt(x_276, x_272); +lean_dec(x_272); +if (x_277 == 0) +{ +lean_object* x_278; lean_object* x_279; lean_object* x_280; +lean_dec(x_4); +x_278 = l_Lean_instInhabitedExpr; +x_279 = l_outOfBounds___rarg(x_278); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_279); +x_280 = l_Lean_Meta_Grind_internalize(x_279, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_257); +if (lean_obj_tag(x_280) == 0) +{ +lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; +x_281 = lean_ctor_get(x_280, 1); +lean_inc(x_281); +lean_dec(x_280); +lean_inc(x_1); +x_282 = l_Lean_Meta_Grind_registerParent(x_1, x_279, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_281); +lean_dec(x_279); +x_283 = lean_ctor_get(x_282, 0); +lean_inc(x_283); +x_284 = lean_ctor_get(x_282, 1); +lean_inc(x_284); +lean_dec(x_282); +x_285 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_283, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_284); +lean_dec(x_283); +return x_285; +} +else +{ +uint8_t x_286; +lean_dec(x_279); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_2); +lean_dec(x_1); +x_286 = !lean_is_exclusive(x_280); +if (x_286 == 0) +{ +return x_280; +} +else +{ +lean_object* x_287; lean_object* x_288; lean_object* x_289; +x_287 = lean_ctor_get(x_280, 0); +x_288 = lean_ctor_get(x_280, 1); +lean_inc(x_288); +lean_inc(x_287); +lean_dec(x_280); +x_289 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_289, 0, x_287); +lean_ctor_set(x_289, 1, x_288); +return x_289; +} +} +} +else +{ +lean_object* x_290; lean_object* x_291; +x_290 = lean_array_fget(x_4, x_276); +lean_dec(x_4); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_290); +x_291 = l_Lean_Meta_Grind_internalize(x_290, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_257); +if (lean_obj_tag(x_291) == 0) +{ +lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; +x_292 = lean_ctor_get(x_291, 1); +lean_inc(x_292); +lean_dec(x_291); +lean_inc(x_1); +x_293 = l_Lean_Meta_Grind_registerParent(x_1, x_290, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_292); +lean_dec(x_290); +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_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_294, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_295); +lean_dec(x_294); +return x_296; +} +else +{ +uint8_t x_297; +lean_dec(x_290); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_2); +lean_dec(x_1); +x_297 = !lean_is_exclusive(x_291); +if (x_297 == 0) +{ +return x_291; +} +else +{ +lean_object* x_298; lean_object* x_299; lean_object* x_300; +x_298 = lean_ctor_get(x_291, 0); +x_299 = lean_ctor_get(x_291, 1); +lean_inc(x_299); +lean_inc(x_298); +lean_dec(x_291); +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; +} +} +} +} +} +block_268: +{ +lean_object* x_260; +lean_dec(x_259); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +x_260 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns(x_251, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_257); +lean_dec(x_251); +if (lean_obj_tag(x_260) == 0) +{ +lean_object* x_261; lean_object* x_262; lean_object* x_263; +x_261 = lean_ctor_get(x_260, 0); +lean_inc(x_261); +x_262 = lean_ctor_get(x_260, 1); +lean_inc(x_262); +lean_dec(x_260); +x_263 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(x_4, x_1, x_2, x_258, x_261, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_262); +lean_dec(x_261); +lean_dec(x_4); +return x_263; +} +else +{ +uint8_t x_264; +lean_dec(x_258); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_264 = !lean_is_exclusive(x_260); +if (x_264 == 0) +{ +return x_260; +} +else +{ +lean_object* x_265; lean_object* x_266; lean_object* x_267; +x_265 = lean_ctor_get(x_260, 0); +x_266 = lean_ctor_get(x_260, 1); +lean_inc(x_266); +lean_inc(x_265); +lean_dec(x_260); +x_267 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_267, 0, x_265); +lean_ctor_set(x_267, 1, x_266); +return x_267; +} +} +} +} +else +{ +uint8_t x_301; +lean_dec(x_251); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_301 = !lean_is_exclusive(x_256); +if (x_301 == 0) +{ +return x_256; +} +else +{ +lean_object* x_302; lean_object* x_303; lean_object* x_304; +x_302 = lean_ctor_get(x_256, 0); +x_303 = lean_ctor_get(x_256, 1); +lean_inc(x_303); +lean_inc(x_302); +lean_dec(x_256); +x_304 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_304, 0, x_302); +lean_ctor_set(x_304, 1, x_303); +return x_304; +} +} +} +else +{ +uint8_t x_305; +lean_dec(x_251); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_305 = !lean_is_exclusive(x_252); +if (x_305 == 0) +{ +return x_252; +} +else +{ +lean_object* x_306; lean_object* x_307; lean_object* x_308; +x_306 = lean_ctor_get(x_252, 0); +x_307 = lean_ctor_get(x_252, 1); +lean_inc(x_307); +lean_inc(x_306); +lean_dec(x_252); +x_308 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_308, 0, x_306); +lean_ctor_set(x_308, 1, x_307); +return x_308; +} +} +} +case 5: +{ +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_3, 0); +lean_inc(x_309); +x_310 = lean_ctor_get(x_3, 1); +lean_inc(x_310); +lean_dec(x_3); +x_311 = lean_array_set(x_4, x_5, x_310); +x_312 = lean_unsigned_to_nat(1u); +x_313 = lean_nat_sub(x_5, x_312); +lean_dec(x_5); +x_3 = x_309; +x_4 = x_311; +x_5 = x_313; +goto _start; +} +case 6: +{ +lean_object* x_315; +lean_dec(x_5); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_1); +x_315 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate(x_1, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_315) == 0) +{ +lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; +x_316 = lean_ctor_get(x_315, 1); +lean_inc(x_316); +lean_dec(x_315); +lean_inc(x_1); +x_317 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs(x_1, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_316); +x_318 = lean_ctor_get(x_317, 1); +lean_inc(x_318); +lean_dec(x_317); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_3); +x_319 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_318); +if (lean_obj_tag(x_319) == 0) +{ +lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_334; uint8_t x_335; +x_320 = lean_ctor_get(x_319, 1); +lean_inc(x_320); +lean_dec(x_319); +lean_inc(x_2); +lean_inc(x_1); +x_321 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1___boxed), 12, 2); +lean_closure_set(x_321, 0, x_1); +lean_closure_set(x_321, 1, x_2); +x_334 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__2; +x_335 = l_Lean_Expr_isConstOf(x_3, x_334); +if (x_335 == 0) +{ +lean_object* x_336; +x_336 = lean_box(0); +x_322 = x_336; +goto block_333; +} +else +{ +lean_object* x_337; lean_object* x_338; uint8_t x_339; +x_337 = lean_array_get_size(x_4); +x_338 = lean_unsigned_to_nat(2u); +x_339 = lean_nat_dec_eq(x_337, x_338); +if (x_339 == 0) +{ +lean_object* x_340; +lean_dec(x_337); +x_340 = lean_box(0); +x_322 = x_340; +goto block_333; +} +else +{ +lean_object* x_341; uint8_t x_342; +lean_dec(x_321); +lean_dec(x_3); +x_341 = lean_unsigned_to_nat(0u); +x_342 = lean_nat_dec_lt(x_341, x_337); +lean_dec(x_337); +if (x_342 == 0) +{ +lean_object* x_343; lean_object* x_344; lean_object* x_345; +lean_dec(x_4); +x_343 = l_Lean_instInhabitedExpr; +x_344 = l_outOfBounds___rarg(x_343); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_344); +x_345 = l_Lean_Meta_Grind_internalize(x_344, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_320); +if (lean_obj_tag(x_345) == 0) +{ +lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; +x_346 = lean_ctor_get(x_345, 1); +lean_inc(x_346); +lean_dec(x_345); +lean_inc(x_1); +x_347 = l_Lean_Meta_Grind_registerParent(x_1, x_344, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_346); +lean_dec(x_344); +x_348 = lean_ctor_get(x_347, 0); +lean_inc(x_348); +x_349 = lean_ctor_get(x_347, 1); +lean_inc(x_349); +lean_dec(x_347); +x_350 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_348, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_349); +lean_dec(x_348); +return x_350; +} +else +{ +uint8_t x_351; +lean_dec(x_344); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_2); +lean_dec(x_1); +x_351 = !lean_is_exclusive(x_345); +if (x_351 == 0) +{ +return x_345; +} +else +{ +lean_object* x_352; lean_object* x_353; lean_object* x_354; +x_352 = lean_ctor_get(x_345, 0); +x_353 = lean_ctor_get(x_345, 1); +lean_inc(x_353); +lean_inc(x_352); +lean_dec(x_345); +x_354 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_354, 0, x_352); +lean_ctor_set(x_354, 1, x_353); +return x_354; +} +} +} +else +{ +lean_object* x_355; lean_object* x_356; +x_355 = lean_array_fget(x_4, x_341); +lean_dec(x_4); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_355); +x_356 = l_Lean_Meta_Grind_internalize(x_355, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_320); +if (lean_obj_tag(x_356) == 0) +{ +lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; +x_357 = lean_ctor_get(x_356, 1); +lean_inc(x_357); +lean_dec(x_356); +lean_inc(x_1); +x_358 = l_Lean_Meta_Grind_registerParent(x_1, x_355, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_357); +lean_dec(x_355); +x_359 = lean_ctor_get(x_358, 0); +lean_inc(x_359); +x_360 = lean_ctor_get(x_358, 1); +lean_inc(x_360); +lean_dec(x_358); +x_361 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_359, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_360); +lean_dec(x_359); +return x_361; +} +else +{ +uint8_t x_362; +lean_dec(x_355); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_2); +lean_dec(x_1); +x_362 = !lean_is_exclusive(x_356); +if (x_362 == 0) +{ +return x_356; +} +else +{ +lean_object* x_363; lean_object* x_364; lean_object* x_365; +x_363 = lean_ctor_get(x_356, 0); +x_364 = lean_ctor_get(x_356, 1); +lean_inc(x_364); +lean_inc(x_363); +lean_dec(x_356); +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; +} +} +} +} +} +block_333: +{ +lean_object* x_323; +lean_dec(x_322); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_3); +x_323 = l_Lean_Meta_Grind_internalize(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_320); +if (lean_obj_tag(x_323) == 0) +{ +lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; +x_324 = lean_ctor_get(x_323, 1); +lean_inc(x_324); +lean_dec(x_323); +lean_inc(x_1); +x_325 = l_Lean_Meta_Grind_registerParent(x_1, x_3, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_324); +lean_dec(x_3); +x_326 = lean_ctor_get(x_325, 0); +lean_inc(x_326); +x_327 = lean_ctor_get(x_325, 1); +lean_inc(x_327); +lean_dec(x_325); +x_328 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(x_4, x_1, x_2, x_321, x_326, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_327); +lean_dec(x_326); +lean_dec(x_4); +return x_328; +} +else +{ +uint8_t x_329; +lean_dec(x_321); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_329 = !lean_is_exclusive(x_323); +if (x_329 == 0) +{ +return x_323; +} +else +{ +lean_object* x_330; lean_object* x_331; lean_object* x_332; +x_330 = lean_ctor_get(x_323, 0); +x_331 = lean_ctor_get(x_323, 1); +lean_inc(x_331); +lean_inc(x_330); +lean_dec(x_323); +x_332 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_332, 0, x_330); +lean_ctor_set(x_332, 1, x_331); +return x_332; +} +} +} +} +else +{ +uint8_t x_366; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_366 = !lean_is_exclusive(x_319); +if (x_366 == 0) +{ +return x_319; +} +else +{ +lean_object* x_367; lean_object* x_368; lean_object* x_369; +x_367 = lean_ctor_get(x_319, 0); +x_368 = lean_ctor_get(x_319, 1); +lean_inc(x_368); +lean_inc(x_367); +lean_dec(x_319); +x_369 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_369, 0, x_367); +lean_ctor_set(x_369, 1, x_368); +return x_369; +} +} +} +else +{ +uint8_t x_370; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_370 = !lean_is_exclusive(x_315); +if (x_370 == 0) +{ +return x_315; +} +else +{ +lean_object* x_371; lean_object* x_372; lean_object* x_373; +x_371 = lean_ctor_get(x_315, 0); +x_372 = lean_ctor_get(x_315, 1); +lean_inc(x_372); +lean_inc(x_371); +lean_dec(x_315); +x_373 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_373, 0, x_371); +lean_ctor_set(x_373, 1, x_372); +return x_373; +} +} +} +case 7: +{ +lean_object* x_374; +lean_dec(x_5); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_1); +x_374 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate(x_1, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_374) == 0) +{ +lean_object* x_375; lean_object* x_376; lean_object* x_377; lean_object* x_378; +x_375 = lean_ctor_get(x_374, 1); +lean_inc(x_375); +lean_dec(x_374); +lean_inc(x_1); +x_376 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs(x_1, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_375); +x_377 = lean_ctor_get(x_376, 1); +lean_inc(x_377); +lean_dec(x_376); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_3); +x_378 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_377); +if (lean_obj_tag(x_378) == 0) +{ +lean_object* x_379; lean_object* x_380; lean_object* x_381; lean_object* x_393; uint8_t x_394; +x_379 = lean_ctor_get(x_378, 1); +lean_inc(x_379); +lean_dec(x_378); +lean_inc(x_2); +lean_inc(x_1); +x_380 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1___boxed), 12, 2); +lean_closure_set(x_380, 0, x_1); +lean_closure_set(x_380, 1, x_2); +x_393 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__2; +x_394 = l_Lean_Expr_isConstOf(x_3, x_393); +if (x_394 == 0) +{ +lean_object* x_395; +x_395 = lean_box(0); +x_381 = x_395; +goto block_392; +} +else +{ +lean_object* x_396; lean_object* x_397; uint8_t x_398; +x_396 = lean_array_get_size(x_4); +x_397 = lean_unsigned_to_nat(2u); +x_398 = lean_nat_dec_eq(x_396, x_397); +if (x_398 == 0) +{ +lean_object* x_399; +lean_dec(x_396); +x_399 = lean_box(0); +x_381 = x_399; +goto block_392; +} +else +{ +lean_object* x_400; uint8_t x_401; +lean_dec(x_380); +lean_dec(x_3); +x_400 = lean_unsigned_to_nat(0u); +x_401 = lean_nat_dec_lt(x_400, x_396); +lean_dec(x_396); +if (x_401 == 0) +{ +lean_object* x_402; lean_object* x_403; lean_object* x_404; +lean_dec(x_4); +x_402 = l_Lean_instInhabitedExpr; +x_403 = l_outOfBounds___rarg(x_402); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_403); +x_404 = l_Lean_Meta_Grind_internalize(x_403, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_379); +if (lean_obj_tag(x_404) == 0) +{ +lean_object* x_405; lean_object* x_406; lean_object* x_407; lean_object* x_408; lean_object* x_409; +x_405 = lean_ctor_get(x_404, 1); +lean_inc(x_405); +lean_dec(x_404); +lean_inc(x_1); +x_406 = l_Lean_Meta_Grind_registerParent(x_1, x_403, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_405); +lean_dec(x_403); +x_407 = lean_ctor_get(x_406, 0); +lean_inc(x_407); +x_408 = lean_ctor_get(x_406, 1); +lean_inc(x_408); +lean_dec(x_406); +x_409 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_407, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_408); +lean_dec(x_407); +return x_409; +} +else +{ +uint8_t x_410; +lean_dec(x_403); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_2); +lean_dec(x_1); +x_410 = !lean_is_exclusive(x_404); +if (x_410 == 0) +{ +return x_404; +} +else +{ +lean_object* x_411; lean_object* x_412; lean_object* x_413; +x_411 = lean_ctor_get(x_404, 0); +x_412 = lean_ctor_get(x_404, 1); +lean_inc(x_412); +lean_inc(x_411); +lean_dec(x_404); +x_413 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_413, 0, x_411); +lean_ctor_set(x_413, 1, x_412); +return x_413; +} +} +} +else +{ +lean_object* x_414; lean_object* x_415; +x_414 = lean_array_fget(x_4, x_400); +lean_dec(x_4); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_414); +x_415 = l_Lean_Meta_Grind_internalize(x_414, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_379); +if (lean_obj_tag(x_415) == 0) +{ +lean_object* x_416; lean_object* x_417; lean_object* x_418; lean_object* x_419; lean_object* x_420; +x_416 = lean_ctor_get(x_415, 1); +lean_inc(x_416); +lean_dec(x_415); +lean_inc(x_1); +x_417 = l_Lean_Meta_Grind_registerParent(x_1, x_414, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_416); +lean_dec(x_414); +x_418 = lean_ctor_get(x_417, 0); +lean_inc(x_418); +x_419 = lean_ctor_get(x_417, 1); +lean_inc(x_419); +lean_dec(x_417); +x_420 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_418, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_419); +lean_dec(x_418); +return x_420; +} +else +{ +uint8_t x_421; +lean_dec(x_414); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_2); +lean_dec(x_1); +x_421 = !lean_is_exclusive(x_415); +if (x_421 == 0) +{ +return x_415; +} +else +{ +lean_object* x_422; lean_object* x_423; lean_object* x_424; +x_422 = lean_ctor_get(x_415, 0); +x_423 = lean_ctor_get(x_415, 1); +lean_inc(x_423); +lean_inc(x_422); +lean_dec(x_415); +x_424 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_424, 0, x_422); +lean_ctor_set(x_424, 1, x_423); +return x_424; +} +} +} +} +} +block_392: +{ +lean_object* x_382; +lean_dec(x_381); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_3); +x_382 = l_Lean_Meta_Grind_internalize(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_379); +if (lean_obj_tag(x_382) == 0) +{ +lean_object* x_383; lean_object* x_384; lean_object* x_385; lean_object* x_386; lean_object* x_387; +x_383 = lean_ctor_get(x_382, 1); +lean_inc(x_383); +lean_dec(x_382); +lean_inc(x_1); +x_384 = l_Lean_Meta_Grind_registerParent(x_1, x_3, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_383); +lean_dec(x_3); +x_385 = lean_ctor_get(x_384, 0); +lean_inc(x_385); +x_386 = lean_ctor_get(x_384, 1); +lean_inc(x_386); +lean_dec(x_384); +x_387 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(x_4, x_1, x_2, x_380, x_385, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_386); +lean_dec(x_385); +lean_dec(x_4); +return x_387; +} +else +{ +uint8_t x_388; +lean_dec(x_380); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_388 = !lean_is_exclusive(x_382); +if (x_388 == 0) +{ +return x_382; +} +else +{ +lean_object* x_389; lean_object* x_390; lean_object* x_391; +x_389 = lean_ctor_get(x_382, 0); +x_390 = lean_ctor_get(x_382, 1); +lean_inc(x_390); +lean_inc(x_389); +lean_dec(x_382); +x_391 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_391, 0, x_389); +lean_ctor_set(x_391, 1, x_390); +return x_391; +} +} +} +} +else +{ +uint8_t x_425; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_425 = !lean_is_exclusive(x_378); +if (x_425 == 0) +{ +return x_378; +} +else +{ +lean_object* x_426; lean_object* x_427; lean_object* x_428; +x_426 = lean_ctor_get(x_378, 0); +x_427 = lean_ctor_get(x_378, 1); +lean_inc(x_427); +lean_inc(x_426); +lean_dec(x_378); +x_428 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_428, 0, x_426); +lean_ctor_set(x_428, 1, x_427); +return x_428; +} +} +} +else +{ +uint8_t x_429; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_429 = !lean_is_exclusive(x_374); +if (x_429 == 0) +{ +return x_374; +} +else +{ +lean_object* x_430; lean_object* x_431; lean_object* x_432; +x_430 = lean_ctor_get(x_374, 0); +x_431 = lean_ctor_get(x_374, 1); +lean_inc(x_431); +lean_inc(x_430); +lean_dec(x_374); +x_432 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_432, 0, x_430); +lean_ctor_set(x_432, 1, x_431); +return x_432; +} +} +} +case 8: +{ +lean_object* x_433; +lean_dec(x_5); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_1); +x_433 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate(x_1, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_433) == 0) +{ +lean_object* x_434; lean_object* x_435; lean_object* x_436; lean_object* x_437; +x_434 = lean_ctor_get(x_433, 1); +lean_inc(x_434); +lean_dec(x_433); +lean_inc(x_1); +x_435 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs(x_1, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_434); +x_436 = lean_ctor_get(x_435, 1); +lean_inc(x_436); +lean_dec(x_435); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_3); +x_437 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_436); +if (lean_obj_tag(x_437) == 0) +{ +lean_object* x_438; lean_object* x_439; lean_object* x_440; lean_object* x_452; uint8_t x_453; +x_438 = lean_ctor_get(x_437, 1); +lean_inc(x_438); +lean_dec(x_437); +lean_inc(x_2); +lean_inc(x_1); +x_439 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1___boxed), 12, 2); +lean_closure_set(x_439, 0, x_1); +lean_closure_set(x_439, 1, x_2); +x_452 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__2; +x_453 = l_Lean_Expr_isConstOf(x_3, x_452); +if (x_453 == 0) +{ +lean_object* x_454; +x_454 = lean_box(0); +x_440 = x_454; +goto block_451; +} +else +{ +lean_object* x_455; lean_object* x_456; uint8_t x_457; +x_455 = lean_array_get_size(x_4); +x_456 = lean_unsigned_to_nat(2u); +x_457 = lean_nat_dec_eq(x_455, x_456); +if (x_457 == 0) +{ +lean_object* x_458; +lean_dec(x_455); +x_458 = lean_box(0); +x_440 = x_458; +goto block_451; +} +else +{ +lean_object* x_459; uint8_t x_460; +lean_dec(x_439); +lean_dec(x_3); +x_459 = lean_unsigned_to_nat(0u); +x_460 = lean_nat_dec_lt(x_459, x_455); +lean_dec(x_455); +if (x_460 == 0) +{ +lean_object* x_461; lean_object* x_462; lean_object* x_463; +lean_dec(x_4); +x_461 = l_Lean_instInhabitedExpr; +x_462 = l_outOfBounds___rarg(x_461); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_462); +x_463 = l_Lean_Meta_Grind_internalize(x_462, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_438); +if (lean_obj_tag(x_463) == 0) +{ +lean_object* x_464; lean_object* x_465; lean_object* x_466; lean_object* x_467; lean_object* x_468; +x_464 = lean_ctor_get(x_463, 1); +lean_inc(x_464); +lean_dec(x_463); +lean_inc(x_1); +x_465 = l_Lean_Meta_Grind_registerParent(x_1, x_462, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_464); +lean_dec(x_462); +x_466 = lean_ctor_get(x_465, 0); +lean_inc(x_466); +x_467 = lean_ctor_get(x_465, 1); +lean_inc(x_467); +lean_dec(x_465); +x_468 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_466, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_467); +lean_dec(x_466); +return x_468; +} +else +{ +uint8_t x_469; +lean_dec(x_462); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_2); +lean_dec(x_1); +x_469 = !lean_is_exclusive(x_463); +if (x_469 == 0) +{ +return x_463; +} +else +{ +lean_object* x_470; lean_object* x_471; lean_object* x_472; +x_470 = lean_ctor_get(x_463, 0); +x_471 = lean_ctor_get(x_463, 1); +lean_inc(x_471); +lean_inc(x_470); +lean_dec(x_463); +x_472 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_472, 0, x_470); +lean_ctor_set(x_472, 1, x_471); +return x_472; +} +} +} +else +{ +lean_object* x_473; lean_object* x_474; +x_473 = lean_array_fget(x_4, x_459); +lean_dec(x_4); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_473); +x_474 = l_Lean_Meta_Grind_internalize(x_473, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_438); +if (lean_obj_tag(x_474) == 0) +{ +lean_object* x_475; lean_object* x_476; lean_object* x_477; lean_object* x_478; lean_object* x_479; +x_475 = lean_ctor_get(x_474, 1); +lean_inc(x_475); +lean_dec(x_474); +lean_inc(x_1); +x_476 = l_Lean_Meta_Grind_registerParent(x_1, x_473, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_475); +lean_dec(x_473); +x_477 = lean_ctor_get(x_476, 0); +lean_inc(x_477); +x_478 = lean_ctor_get(x_476, 1); +lean_inc(x_478); +lean_dec(x_476); +x_479 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_477, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_478); +lean_dec(x_477); +return x_479; +} +else +{ +uint8_t x_480; +lean_dec(x_473); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_2); +lean_dec(x_1); +x_480 = !lean_is_exclusive(x_474); +if (x_480 == 0) +{ +return x_474; +} +else +{ +lean_object* x_481; lean_object* x_482; lean_object* x_483; +x_481 = lean_ctor_get(x_474, 0); +x_482 = lean_ctor_get(x_474, 1); +lean_inc(x_482); +lean_inc(x_481); +lean_dec(x_474); +x_483 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_483, 0, x_481); +lean_ctor_set(x_483, 1, x_482); +return x_483; +} +} +} +} +} +block_451: +{ +lean_object* x_441; +lean_dec(x_440); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_3); +x_441 = l_Lean_Meta_Grind_internalize(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_438); +if (lean_obj_tag(x_441) == 0) +{ +lean_object* x_442; lean_object* x_443; lean_object* x_444; lean_object* x_445; lean_object* x_446; +x_442 = lean_ctor_get(x_441, 1); +lean_inc(x_442); +lean_dec(x_441); +lean_inc(x_1); +x_443 = l_Lean_Meta_Grind_registerParent(x_1, x_3, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_442); +lean_dec(x_3); +x_444 = lean_ctor_get(x_443, 0); +lean_inc(x_444); +x_445 = lean_ctor_get(x_443, 1); +lean_inc(x_445); +lean_dec(x_443); +x_446 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(x_4, x_1, x_2, x_439, x_444, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_445); +lean_dec(x_444); +lean_dec(x_4); +return x_446; +} +else +{ +uint8_t x_447; +lean_dec(x_439); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_447 = !lean_is_exclusive(x_441); +if (x_447 == 0) +{ +return x_441; +} +else +{ +lean_object* x_448; lean_object* x_449; lean_object* x_450; +x_448 = lean_ctor_get(x_441, 0); +x_449 = lean_ctor_get(x_441, 1); +lean_inc(x_449); +lean_inc(x_448); +lean_dec(x_441); +x_450 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_450, 0, x_448); +lean_ctor_set(x_450, 1, x_449); +return x_450; +} +} +} +} +else +{ +uint8_t x_484; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_484 = !lean_is_exclusive(x_437); +if (x_484 == 0) +{ +return x_437; +} +else +{ +lean_object* x_485; lean_object* x_486; lean_object* x_487; +x_485 = lean_ctor_get(x_437, 0); +x_486 = lean_ctor_get(x_437, 1); +lean_inc(x_486); +lean_inc(x_485); +lean_dec(x_437); +x_487 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_487, 0, x_485); +lean_ctor_set(x_487, 1, x_486); +return x_487; +} +} +} +else +{ +uint8_t x_488; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_488 = !lean_is_exclusive(x_433); +if (x_488 == 0) +{ +return x_433; +} +else +{ +lean_object* x_489; lean_object* x_490; lean_object* x_491; +x_489 = lean_ctor_get(x_433, 0); +x_490 = lean_ctor_get(x_433, 1); +lean_inc(x_490); +lean_inc(x_489); +lean_dec(x_433); +x_491 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_491, 0, x_489); +lean_ctor_set(x_491, 1, x_490); +return x_491; +} +} +} +case 9: +{ +lean_object* x_492; +lean_dec(x_5); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_1); +x_492 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate(x_1, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_492) == 0) +{ +lean_object* x_493; lean_object* x_494; lean_object* x_495; lean_object* x_496; +x_493 = lean_ctor_get(x_492, 1); +lean_inc(x_493); +lean_dec(x_492); +lean_inc(x_1); +x_494 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs(x_1, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_493); +x_495 = lean_ctor_get(x_494, 1); +lean_inc(x_495); +lean_dec(x_494); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_3); +x_496 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_495); +if (lean_obj_tag(x_496) == 0) +{ +lean_object* x_497; lean_object* x_498; lean_object* x_499; lean_object* x_511; uint8_t x_512; +x_497 = lean_ctor_get(x_496, 1); +lean_inc(x_497); +lean_dec(x_496); +lean_inc(x_2); +lean_inc(x_1); +x_498 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1___boxed), 12, 2); +lean_closure_set(x_498, 0, x_1); +lean_closure_set(x_498, 1, x_2); +x_511 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__2; +x_512 = l_Lean_Expr_isConstOf(x_3, x_511); +if (x_512 == 0) +{ +lean_object* x_513; +x_513 = lean_box(0); +x_499 = x_513; +goto block_510; +} +else +{ +lean_object* x_514; lean_object* x_515; uint8_t x_516; +x_514 = lean_array_get_size(x_4); +x_515 = lean_unsigned_to_nat(2u); +x_516 = lean_nat_dec_eq(x_514, x_515); +if (x_516 == 0) +{ +lean_object* x_517; +lean_dec(x_514); +x_517 = lean_box(0); +x_499 = x_517; +goto block_510; +} +else +{ +lean_object* x_518; uint8_t x_519; +lean_dec(x_498); +lean_dec(x_3); +x_518 = lean_unsigned_to_nat(0u); +x_519 = lean_nat_dec_lt(x_518, x_514); +lean_dec(x_514); +if (x_519 == 0) +{ +lean_object* x_520; lean_object* x_521; lean_object* x_522; +lean_dec(x_4); +x_520 = l_Lean_instInhabitedExpr; +x_521 = l_outOfBounds___rarg(x_520); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_521); +x_522 = l_Lean_Meta_Grind_internalize(x_521, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_497); +if (lean_obj_tag(x_522) == 0) +{ +lean_object* x_523; lean_object* x_524; lean_object* x_525; lean_object* x_526; lean_object* x_527; +x_523 = lean_ctor_get(x_522, 1); +lean_inc(x_523); +lean_dec(x_522); +lean_inc(x_1); +x_524 = l_Lean_Meta_Grind_registerParent(x_1, x_521, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_523); +lean_dec(x_521); +x_525 = lean_ctor_get(x_524, 0); +lean_inc(x_525); +x_526 = lean_ctor_get(x_524, 1); +lean_inc(x_526); +lean_dec(x_524); +x_527 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_525, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_526); +lean_dec(x_525); +return x_527; +} +else +{ +uint8_t x_528; +lean_dec(x_521); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_2); +lean_dec(x_1); +x_528 = !lean_is_exclusive(x_522); +if (x_528 == 0) +{ +return x_522; +} +else +{ +lean_object* x_529; lean_object* x_530; lean_object* x_531; +x_529 = lean_ctor_get(x_522, 0); +x_530 = lean_ctor_get(x_522, 1); +lean_inc(x_530); +lean_inc(x_529); +lean_dec(x_522); +x_531 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_531, 0, x_529); +lean_ctor_set(x_531, 1, x_530); +return x_531; +} +} +} +else +{ +lean_object* x_532; lean_object* x_533; +x_532 = lean_array_fget(x_4, x_518); +lean_dec(x_4); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_532); +x_533 = l_Lean_Meta_Grind_internalize(x_532, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_497); +if (lean_obj_tag(x_533) == 0) +{ +lean_object* x_534; lean_object* x_535; lean_object* x_536; lean_object* x_537; lean_object* x_538; +x_534 = lean_ctor_get(x_533, 1); +lean_inc(x_534); +lean_dec(x_533); +lean_inc(x_1); +x_535 = l_Lean_Meta_Grind_registerParent(x_1, x_532, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_534); +lean_dec(x_532); +x_536 = lean_ctor_get(x_535, 0); +lean_inc(x_536); +x_537 = lean_ctor_get(x_535, 1); +lean_inc(x_537); +lean_dec(x_535); +x_538 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_536, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_537); +lean_dec(x_536); +return x_538; +} +else +{ +uint8_t x_539; +lean_dec(x_532); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_2); +lean_dec(x_1); +x_539 = !lean_is_exclusive(x_533); +if (x_539 == 0) +{ +return x_533; +} +else +{ +lean_object* x_540; lean_object* x_541; lean_object* x_542; +x_540 = lean_ctor_get(x_533, 0); +x_541 = lean_ctor_get(x_533, 1); +lean_inc(x_541); +lean_inc(x_540); +lean_dec(x_533); +x_542 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_542, 0, x_540); +lean_ctor_set(x_542, 1, x_541); +return x_542; +} +} +} +} +} +block_510: +{ +lean_object* x_500; +lean_dec(x_499); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_3); +x_500 = l_Lean_Meta_Grind_internalize(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_497); +if (lean_obj_tag(x_500) == 0) +{ +lean_object* x_501; lean_object* x_502; lean_object* x_503; lean_object* x_504; lean_object* x_505; +x_501 = lean_ctor_get(x_500, 1); +lean_inc(x_501); +lean_dec(x_500); +lean_inc(x_1); +x_502 = l_Lean_Meta_Grind_registerParent(x_1, x_3, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_501); +lean_dec(x_3); +x_503 = lean_ctor_get(x_502, 0); +lean_inc(x_503); +x_504 = lean_ctor_get(x_502, 1); +lean_inc(x_504); +lean_dec(x_502); +x_505 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(x_4, x_1, x_2, x_498, x_503, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_504); +lean_dec(x_503); +lean_dec(x_4); +return x_505; +} +else +{ +uint8_t x_506; +lean_dec(x_498); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_506 = !lean_is_exclusive(x_500); +if (x_506 == 0) +{ +return x_500; +} +else +{ +lean_object* x_507; lean_object* x_508; lean_object* x_509; +x_507 = lean_ctor_get(x_500, 0); +x_508 = lean_ctor_get(x_500, 1); +lean_inc(x_508); +lean_inc(x_507); +lean_dec(x_500); +x_509 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_509, 0, x_507); +lean_ctor_set(x_509, 1, x_508); +return x_509; +} +} +} +} +else +{ +uint8_t x_543; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_543 = !lean_is_exclusive(x_496); +if (x_543 == 0) +{ +return x_496; +} +else +{ +lean_object* x_544; lean_object* x_545; lean_object* x_546; +x_544 = lean_ctor_get(x_496, 0); +x_545 = lean_ctor_get(x_496, 1); +lean_inc(x_545); +lean_inc(x_544); +lean_dec(x_496); +x_546 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_546, 0, x_544); +lean_ctor_set(x_546, 1, x_545); +return x_546; +} +} +} +else +{ +uint8_t x_547; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_547 = !lean_is_exclusive(x_492); +if (x_547 == 0) +{ +return x_492; +} +else +{ +lean_object* x_548; lean_object* x_549; lean_object* x_550; +x_548 = lean_ctor_get(x_492, 0); +x_549 = lean_ctor_get(x_492, 1); +lean_inc(x_549); +lean_inc(x_548); +lean_dec(x_492); +x_550 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_550, 0, x_548); +lean_ctor_set(x_550, 1, x_549); +return x_550; +} +} +} +case 10: +{ +lean_object* x_551; +lean_dec(x_5); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_1); +x_551 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate(x_1, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_551) == 0) +{ +lean_object* x_552; lean_object* x_553; lean_object* x_554; lean_object* x_555; +x_552 = lean_ctor_get(x_551, 1); +lean_inc(x_552); +lean_dec(x_551); +lean_inc(x_1); +x_553 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs(x_1, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_552); +x_554 = lean_ctor_get(x_553, 1); +lean_inc(x_554); +lean_dec(x_553); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_3); +x_555 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_554); +if (lean_obj_tag(x_555) == 0) +{ +lean_object* x_556; lean_object* x_557; lean_object* x_558; lean_object* x_570; uint8_t x_571; +x_556 = lean_ctor_get(x_555, 1); +lean_inc(x_556); +lean_dec(x_555); +lean_inc(x_2); +lean_inc(x_1); +x_557 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1___boxed), 12, 2); +lean_closure_set(x_557, 0, x_1); +lean_closure_set(x_557, 1, x_2); +x_570 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__2; +x_571 = l_Lean_Expr_isConstOf(x_3, x_570); +if (x_571 == 0) +{ +lean_object* x_572; +x_572 = lean_box(0); +x_558 = x_572; +goto block_569; +} +else +{ +lean_object* x_573; lean_object* x_574; uint8_t x_575; +x_573 = lean_array_get_size(x_4); +x_574 = lean_unsigned_to_nat(2u); +x_575 = lean_nat_dec_eq(x_573, x_574); +if (x_575 == 0) +{ +lean_object* x_576; +lean_dec(x_573); +x_576 = lean_box(0); +x_558 = x_576; +goto block_569; +} +else +{ +lean_object* x_577; uint8_t x_578; +lean_dec(x_557); +lean_dec(x_3); +x_577 = lean_unsigned_to_nat(0u); +x_578 = lean_nat_dec_lt(x_577, x_573); +lean_dec(x_573); +if (x_578 == 0) +{ +lean_object* x_579; lean_object* x_580; lean_object* x_581; +lean_dec(x_4); +x_579 = l_Lean_instInhabitedExpr; +x_580 = l_outOfBounds___rarg(x_579); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_580); +x_581 = l_Lean_Meta_Grind_internalize(x_580, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_556); +if (lean_obj_tag(x_581) == 0) +{ +lean_object* x_582; lean_object* x_583; lean_object* x_584; lean_object* x_585; lean_object* x_586; +x_582 = lean_ctor_get(x_581, 1); +lean_inc(x_582); +lean_dec(x_581); +lean_inc(x_1); +x_583 = l_Lean_Meta_Grind_registerParent(x_1, x_580, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_582); +lean_dec(x_580); +x_584 = lean_ctor_get(x_583, 0); +lean_inc(x_584); +x_585 = lean_ctor_get(x_583, 1); +lean_inc(x_585); +lean_dec(x_583); +x_586 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_584, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_585); +lean_dec(x_584); +return x_586; +} +else +{ +uint8_t x_587; +lean_dec(x_580); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_2); +lean_dec(x_1); +x_587 = !lean_is_exclusive(x_581); +if (x_587 == 0) +{ +return x_581; +} +else +{ +lean_object* x_588; lean_object* x_589; lean_object* x_590; +x_588 = lean_ctor_get(x_581, 0); +x_589 = lean_ctor_get(x_581, 1); +lean_inc(x_589); +lean_inc(x_588); +lean_dec(x_581); +x_590 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_590, 0, x_588); +lean_ctor_set(x_590, 1, x_589); +return x_590; +} +} +} +else +{ +lean_object* x_591; lean_object* x_592; +x_591 = lean_array_fget(x_4, x_577); +lean_dec(x_4); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_591); +x_592 = l_Lean_Meta_Grind_internalize(x_591, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_556); +if (lean_obj_tag(x_592) == 0) +{ +lean_object* x_593; lean_object* x_594; lean_object* x_595; lean_object* x_596; lean_object* x_597; +x_593 = lean_ctor_get(x_592, 1); +lean_inc(x_593); +lean_dec(x_592); +lean_inc(x_1); +x_594 = l_Lean_Meta_Grind_registerParent(x_1, x_591, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_593); +lean_dec(x_591); +x_595 = lean_ctor_get(x_594, 0); +lean_inc(x_595); +x_596 = lean_ctor_get(x_594, 1); +lean_inc(x_596); +lean_dec(x_594); +x_597 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_595, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_596); +lean_dec(x_595); +return x_597; +} +else +{ +uint8_t x_598; +lean_dec(x_591); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_2); +lean_dec(x_1); +x_598 = !lean_is_exclusive(x_592); +if (x_598 == 0) +{ +return x_592; +} +else +{ +lean_object* x_599; lean_object* x_600; lean_object* x_601; +x_599 = lean_ctor_get(x_592, 0); +x_600 = lean_ctor_get(x_592, 1); +lean_inc(x_600); +lean_inc(x_599); +lean_dec(x_592); +x_601 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_601, 0, x_599); +lean_ctor_set(x_601, 1, x_600); +return x_601; +} +} +} +} +} +block_569: +{ +lean_object* x_559; +lean_dec(x_558); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_3); +x_559 = l_Lean_Meta_Grind_internalize(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_556); +if (lean_obj_tag(x_559) == 0) +{ +lean_object* x_560; lean_object* x_561; lean_object* x_562; lean_object* x_563; lean_object* x_564; +x_560 = lean_ctor_get(x_559, 1); +lean_inc(x_560); +lean_dec(x_559); +lean_inc(x_1); +x_561 = l_Lean_Meta_Grind_registerParent(x_1, x_3, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_560); +lean_dec(x_3); +x_562 = lean_ctor_get(x_561, 0); +lean_inc(x_562); +x_563 = lean_ctor_get(x_561, 1); +lean_inc(x_563); +lean_dec(x_561); +x_564 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(x_4, x_1, x_2, x_557, x_562, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_563); +lean_dec(x_562); +lean_dec(x_4); +return x_564; +} +else +{ +uint8_t x_565; +lean_dec(x_557); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_565 = !lean_is_exclusive(x_559); +if (x_565 == 0) +{ +return x_559; +} +else +{ +lean_object* x_566; lean_object* x_567; lean_object* x_568; +x_566 = lean_ctor_get(x_559, 0); +x_567 = lean_ctor_get(x_559, 1); +lean_inc(x_567); +lean_inc(x_566); +lean_dec(x_559); +x_568 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_568, 0, x_566); +lean_ctor_set(x_568, 1, x_567); +return x_568; +} +} +} +} +else +{ +uint8_t x_602; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_602 = !lean_is_exclusive(x_555); +if (x_602 == 0) +{ +return x_555; +} +else +{ +lean_object* x_603; lean_object* x_604; lean_object* x_605; +x_603 = lean_ctor_get(x_555, 0); +x_604 = lean_ctor_get(x_555, 1); +lean_inc(x_604); +lean_inc(x_603); +lean_dec(x_555); +x_605 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_605, 0, x_603); +lean_ctor_set(x_605, 1, x_604); +return x_605; +} +} +} +else +{ +uint8_t x_606; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_606 = !lean_is_exclusive(x_551); +if (x_606 == 0) { -lean_object* x_402; lean_object* x_403; lean_object* x_404; +return x_551; +} +else +{ +lean_object* x_607; lean_object* x_608; lean_object* x_609; +x_607 = lean_ctor_get(x_551, 0); +x_608 = lean_ctor_get(x_551, 1); +lean_inc(x_608); +lean_inc(x_607); +lean_dec(x_551); +x_609 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_609, 0, x_607); +lean_ctor_set(x_609, 1, x_608); +return x_609; +} +} +} +default: +{ +lean_object* x_610; +lean_dec(x_5); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_1); +x_610 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate(x_1, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_610) == 0) +{ +lean_object* x_611; lean_object* x_612; lean_object* x_613; lean_object* x_614; +x_611 = lean_ctor_get(x_610, 1); +lean_inc(x_611); +lean_dec(x_610); +lean_inc(x_1); +x_612 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs(x_1, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_611); +x_613 = lean_ctor_get(x_612, 1); +lean_inc(x_613); +lean_dec(x_612); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_3); +x_614 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_613); +if (lean_obj_tag(x_614) == 0) +{ +lean_object* x_615; lean_object* x_616; lean_object* x_617; lean_object* x_629; uint8_t x_630; +x_615 = lean_ctor_get(x_614, 1); +lean_inc(x_615); +lean_dec(x_614); +lean_inc(x_2); +lean_inc(x_1); +x_616 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1___boxed), 12, 2); +lean_closure_set(x_616, 0, x_1); +lean_closure_set(x_616, 1, x_2); +x_629 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__2; +x_630 = l_Lean_Expr_isConstOf(x_3, x_629); +if (x_630 == 0) +{ +lean_object* x_631; +x_631 = lean_box(0); +x_617 = x_631; +goto block_628; +} +else +{ +lean_object* x_632; lean_object* x_633; uint8_t x_634; +x_632 = lean_array_get_size(x_4); +x_633 = lean_unsigned_to_nat(2u); +x_634 = lean_nat_dec_eq(x_632, x_633); +if (x_634 == 0) +{ +lean_object* x_635; +lean_dec(x_632); +x_635 = lean_box(0); +x_617 = x_635; +goto block_628; +} +else +{ +lean_object* x_636; uint8_t x_637; +lean_dec(x_616); +lean_dec(x_3); +x_636 = lean_unsigned_to_nat(0u); +x_637 = lean_nat_dec_lt(x_636, x_632); +lean_dec(x_632); +if (x_637 == 0) +{ +lean_object* x_638; lean_object* x_639; lean_object* x_640; lean_dec(x_4); -x_402 = l_Lean_instInhabitedExpr; -x_403 = l_outOfBounds___rarg(x_402); +x_638 = l_Lean_instInhabitedExpr; +x_639 = l_outOfBounds___rarg(x_638); lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); @@ -6902,797 +10522,1204 @@ lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_2); -lean_inc(x_403); -x_404 = l_Lean_Meta_Grind_internalize(x_403, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_404) == 0) +lean_inc(x_639); +x_640 = l_Lean_Meta_Grind_internalize(x_639, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_615); +if (lean_obj_tag(x_640) == 0) +{ +lean_object* x_641; lean_object* x_642; lean_object* x_643; lean_object* x_644; lean_object* x_645; +x_641 = lean_ctor_get(x_640, 1); +lean_inc(x_641); +lean_dec(x_640); +lean_inc(x_1); +x_642 = l_Lean_Meta_Grind_registerParent(x_1, x_639, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_641); +lean_dec(x_639); +x_643 = lean_ctor_get(x_642, 0); +lean_inc(x_643); +x_644 = lean_ctor_get(x_642, 1); +lean_inc(x_644); +lean_dec(x_642); +x_645 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_643, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_644); +lean_dec(x_643); +return x_645; +} +else { -lean_object* x_405; lean_object* x_406; lean_object* x_407; lean_object* x_408; lean_object* x_409; -x_405 = lean_ctor_get(x_404, 1); -lean_inc(x_405); -lean_dec(x_404); +uint8_t x_646; +lean_dec(x_639); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_2); +lean_dec(x_1); +x_646 = !lean_is_exclusive(x_640); +if (x_646 == 0) +{ +return x_640; +} +else +{ +lean_object* x_647; lean_object* x_648; lean_object* x_649; +x_647 = lean_ctor_get(x_640, 0); +x_648 = lean_ctor_get(x_640, 1); +lean_inc(x_648); +lean_inc(x_647); +lean_dec(x_640); +x_649 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_649, 0, x_647); +lean_ctor_set(x_649, 1, x_648); +return x_649; +} +} +} +else +{ +lean_object* x_650; lean_object* x_651; +x_650 = lean_array_fget(x_4, x_636); +lean_dec(x_4); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_650); +x_651 = l_Lean_Meta_Grind_internalize(x_650, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_615); +if (lean_obj_tag(x_651) == 0) +{ +lean_object* x_652; lean_object* x_653; lean_object* x_654; lean_object* x_655; lean_object* x_656; +x_652 = lean_ctor_get(x_651, 1); +lean_inc(x_652); +lean_dec(x_651); lean_inc(x_1); -x_406 = l_Lean_Meta_Grind_registerParent(x_1, x_403, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_405); -lean_dec(x_403); -x_407 = lean_ctor_get(x_406, 0); -lean_inc(x_407); -x_408 = lean_ctor_get(x_406, 1); -lean_inc(x_408); -lean_dec(x_406); -x_409 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_407, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_408); -lean_dec(x_407); -return x_409; +x_653 = l_Lean_Meta_Grind_registerParent(x_1, x_650, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_652); +lean_dec(x_650); +x_654 = lean_ctor_get(x_653, 0); +lean_inc(x_654); +x_655 = lean_ctor_get(x_653, 1); +lean_inc(x_655); +lean_dec(x_653); +x_656 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_654, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_655); +lean_dec(x_654); +return x_656; +} +else +{ +uint8_t x_657; +lean_dec(x_650); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_2); +lean_dec(x_1); +x_657 = !lean_is_exclusive(x_651); +if (x_657 == 0) +{ +return x_651; +} +else +{ +lean_object* x_658; lean_object* x_659; lean_object* x_660; +x_658 = lean_ctor_get(x_651, 0); +x_659 = lean_ctor_get(x_651, 1); +lean_inc(x_659); +lean_inc(x_658); +lean_dec(x_651); +x_660 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_660, 0, x_658); +lean_ctor_set(x_660, 1, x_659); +return x_660; +} +} +} +} +} +block_628: +{ +lean_object* x_618; +lean_dec(x_617); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_3); +x_618 = l_Lean_Meta_Grind_internalize(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_615); +if (lean_obj_tag(x_618) == 0) +{ +lean_object* x_619; lean_object* x_620; lean_object* x_621; lean_object* x_622; lean_object* x_623; +x_619 = lean_ctor_get(x_618, 1); +lean_inc(x_619); +lean_dec(x_618); +lean_inc(x_1); +x_620 = l_Lean_Meta_Grind_registerParent(x_1, x_3, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_619); +lean_dec(x_3); +x_621 = lean_ctor_get(x_620, 0); +lean_inc(x_621); +x_622 = lean_ctor_get(x_620, 1); +lean_inc(x_622); +lean_dec(x_620); +x_623 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(x_4, x_1, x_2, x_616, x_621, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_622); +lean_dec(x_621); +lean_dec(x_4); +return x_623; +} +else +{ +uint8_t x_624; +lean_dec(x_616); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_624 = !lean_is_exclusive(x_618); +if (x_624 == 0) +{ +return x_618; +} +else +{ +lean_object* x_625; lean_object* x_626; lean_object* x_627; +x_625 = lean_ctor_get(x_618, 0); +x_626 = lean_ctor_get(x_618, 1); +lean_inc(x_626); +lean_inc(x_625); +lean_dec(x_618); +x_627 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_627, 0, x_625); +lean_ctor_set(x_627, 1, x_626); +return x_627; +} +} +} +} +else +{ +uint8_t x_661; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_661 = !lean_is_exclusive(x_614); +if (x_661 == 0) +{ +return x_614; +} +else +{ +lean_object* x_662; lean_object* x_663; lean_object* x_664; +x_662 = lean_ctor_get(x_614, 0); +x_663 = lean_ctor_get(x_614, 1); +lean_inc(x_663); +lean_inc(x_662); +lean_dec(x_614); +x_664 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_664, 0, x_662); +lean_ctor_set(x_664, 1, x_663); +return x_664; +} +} +} +else +{ +uint8_t x_665; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_665 = !lean_is_exclusive(x_610); +if (x_665 == 0) +{ +return x_610; +} +else +{ +lean_object* x_666; lean_object* x_667; lean_object* x_668; +x_666 = lean_ctor_get(x_610, 0); +x_667 = lean_ctor_get(x_610, 1); +lean_inc(x_667); +lean_inc(x_666); +lean_dec(x_610); +x_668 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_668, 0, x_666); +lean_ctor_set(x_668, 1, x_667); +return x_668; +} +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_internalize___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +uint8_t x_13; lean_object* x_14; +x_13 = 0; +x_14 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_13, x_13, x_2, 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_Meta_Grind_internalize___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; +x_12 = l_Lean_Meta_Grind_propagateUp(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +return x_12; +} +} +static lean_object* _init_l_Lean_Meta_Grind_internalize___lambda__3___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.Meta.Tactic.Grind.Internalize", 34, 34); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_internalize___lambda__3___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.Meta.Grind.internalize", 27, 27); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_internalize___lambda__3___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("unreachable code has been reached", 33, 33); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_internalize___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; lean_object* x_6; +x_1 = l_Lean_Meta_Grind_internalize___lambda__3___closed__1; +x_2 = l_Lean_Meta_Grind_internalize___lambda__3___closed__2; +x_3 = lean_unsigned_to_nat(144u); +x_4 = lean_unsigned_to_nat(16u); +x_5 = l_Lean_Meta_Grind_internalize___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; +} +} +static lean_object* _init_l_Lean_Meta_Grind_internalize___lambda__3___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("unexpected term during internalization", 38, 38); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_internalize___lambda__3___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_internalize___lambda__3___closed__5; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_internalize___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: +{ +switch (lean_obj_tag(x_1)) { +case 0: +{ +lean_object* x_13; lean_object* x_14; +lean_dec(x_2); +lean_dec(x_1); +x_13 = l_Lean_Meta_Grind_internalize___lambda__3___closed__4; +x_14 = l_panic___at_Lean_Meta_Grind_internalize___spec__1(x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_14; } -else +case 2: { -uint8_t x_410; -lean_dec(x_403); -lean_dec(x_13); -lean_dec(x_12); +lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_15 = l_Lean_Meta_Grind_addCongrTable___closed__2; +x_16 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_15, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +x_18 = lean_unbox(x_17); +lean_dec(x_17); +if (x_18 == 0) +{ +lean_object* x_19; uint8_t x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_16, 1); +lean_inc(x_19); +lean_dec(x_16); +x_20 = 0; +x_21 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_20, x_20, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_19); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_2); -lean_dec(x_1); -x_410 = !lean_is_exclusive(x_404); -if (x_410 == 0) -{ -return x_404; +lean_dec(x_5); +lean_dec(x_4); +return x_21; } else { -lean_object* x_411; lean_object* x_412; lean_object* x_413; -x_411 = lean_ctor_get(x_404, 0); -x_412 = lean_ctor_get(x_404, 1); -lean_inc(x_412); -lean_inc(x_411); -lean_dec(x_404); -x_413 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_413, 0, x_411); -lean_ctor_set(x_413, 1, x_412); -return x_413; -} -} -} -else +uint8_t x_22; +x_22 = !lean_is_exclusive(x_16); +if (x_22 == 0) { -lean_object* x_414; lean_object* x_415; -x_414 = lean_array_fget(x_4, x_400); -lean_dec(x_4); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_2); -lean_inc(x_414); -x_415 = l_Lean_Meta_Grind_internalize(x_414, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_415) == 0) +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_16, 1); +x_24 = lean_ctor_get(x_16, 0); +lean_dec(x_24); +x_25 = l_Lean_Meta_Grind_updateLastTag(x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_23); +if (lean_obj_tag(x_25) == 0) { -lean_object* x_416; lean_object* x_417; lean_object* x_418; lean_object* x_419; lean_object* x_420; -x_416 = lean_ctor_get(x_415, 1); -lean_inc(x_416); -lean_dec(x_415); +lean_object* x_26; 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; +x_26 = lean_ctor_get(x_25, 1); +lean_inc(x_26); +lean_dec(x_25); lean_inc(x_1); -x_417 = l_Lean_Meta_Grind_registerParent(x_1, x_414, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_416); -lean_dec(x_414); -x_418 = lean_ctor_get(x_417, 0); -lean_inc(x_418); -x_419 = lean_ctor_get(x_417, 1); -lean_inc(x_419); -lean_dec(x_417); -x_420 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_418, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_419); -lean_dec(x_418); -return x_420; +x_27 = l_Lean_indentExpr(x_1); +x_28 = l_Lean_Meta_Grind_internalize___lambda__3___closed__6; +lean_ctor_set_tag(x_16, 7); +lean_ctor_set(x_16, 1, x_27); +lean_ctor_set(x_16, 0, x_28); +x_29 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__6; +x_30 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_30, 0, x_16); +lean_ctor_set(x_30, 1, x_29); +x_31 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_15, x_30, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_26); +x_32 = lean_ctor_get(x_31, 1); +lean_inc(x_32); +lean_dec(x_31); +x_33 = 0; +x_34 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_33, x_33, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_32); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_34; } else { -uint8_t x_421; -lean_dec(x_414); -lean_dec(x_13); -lean_dec(x_12); +uint8_t x_35; +lean_free_object(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_2); lean_dec(x_1); -x_421 = !lean_is_exclusive(x_415); -if (x_421 == 0) +x_35 = !lean_is_exclusive(x_25); +if (x_35 == 0) { -return x_415; +return x_25; } else { -lean_object* x_422; lean_object* x_423; lean_object* x_424; -x_422 = lean_ctor_get(x_415, 0); -x_423 = lean_ctor_get(x_415, 1); -lean_inc(x_423); -lean_inc(x_422); -lean_dec(x_415); -x_424 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_424, 0, x_422); -lean_ctor_set(x_424, 1, x_423); -return x_424; -} -} +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; } } } -block_392: +else { -lean_object* x_382; -lean_dec(x_381); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_2); -lean_inc(x_3); -x_382 = l_Lean_Meta_Grind_internalize(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_382) == 0) +lean_object* x_39; lean_object* x_40; +x_39 = lean_ctor_get(x_16, 1); +lean_inc(x_39); +lean_dec(x_16); +x_40 = l_Lean_Meta_Grind_updateLastTag(x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_39); +if (lean_obj_tag(x_40) == 0) { -lean_object* x_383; lean_object* x_384; lean_object* x_385; lean_object* x_386; lean_object* x_387; -x_383 = lean_ctor_get(x_382, 1); -lean_inc(x_383); -lean_dec(x_382); +lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; uint8_t x_49; lean_object* x_50; +x_41 = lean_ctor_get(x_40, 1); +lean_inc(x_41); +lean_dec(x_40); lean_inc(x_1); -x_384 = l_Lean_Meta_Grind_registerParent(x_1, x_3, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_383); -lean_dec(x_3); -x_385 = lean_ctor_get(x_384, 0); -lean_inc(x_385); -x_386 = lean_ctor_get(x_384, 1); -lean_inc(x_386); -lean_dec(x_384); -x_387 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(x_4, x_1, x_2, x_380, x_385, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_386); -lean_dec(x_385); +x_42 = l_Lean_indentExpr(x_1); +x_43 = l_Lean_Meta_Grind_internalize___lambda__3___closed__6; +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_Lean_Meta_Grind_addCongrTable___lambda__2___closed__6; +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_Meta_Grind_updateLastTag___spec__2(x_15, x_46, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_41); +x_48 = lean_ctor_get(x_47, 1); +lean_inc(x_48); +lean_dec(x_47); +x_49 = 0; +x_50 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_49, x_49, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_48); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); lean_dec(x_4); -return x_387; +return x_50; } else { -uint8_t x_388; -lean_dec(x_380); -lean_dec(x_13); -lean_dec(x_12); +lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_dec(x_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_388 = !lean_is_exclusive(x_382); -if (x_388 == 0) -{ -return x_382; +x_51 = lean_ctor_get(x_40, 0); +lean_inc(x_51); +x_52 = lean_ctor_get(x_40, 1); +lean_inc(x_52); +if (lean_is_exclusive(x_40)) { + lean_ctor_release(x_40, 0); + lean_ctor_release(x_40, 1); + x_53 = x_40; +} else { + lean_dec_ref(x_40); + x_53 = lean_box(0); } -else -{ -lean_object* x_389; lean_object* x_390; lean_object* x_391; -x_389 = lean_ctor_get(x_382, 0); -x_390 = lean_ctor_get(x_382, 1); -lean_inc(x_390); -lean_inc(x_389); -lean_dec(x_382); -x_391 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_391, 0, x_389); -lean_ctor_set(x_391, 1, x_390); -return x_391; +if (lean_is_scalar(x_53)) { + x_54 = lean_alloc_ctor(1, 2, 0); +} else { + x_54 = x_53; +} +lean_ctor_set(x_54, 0, x_51); +lean_ctor_set(x_54, 1, x_52); +return x_54; } } } } -case 10: +case 3: { -lean_object* x_425; lean_object* x_426; lean_object* x_438; uint8_t x_439; +lean_object* x_55; lean_object* x_56; +lean_dec(x_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_inc(x_2); -lean_inc(x_1); -x_425 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1___boxed), 12, 2); -lean_closure_set(x_425, 0, x_1); -lean_closure_set(x_425, 1, x_2); -x_438 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__4; -x_439 = l_Lean_Expr_isConstOf(x_3, x_438); -if (x_439 == 0) -{ -lean_object* x_440; -x_440 = lean_box(0); -x_426 = x_440; -goto block_437; +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +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_12); +return x_56; } -else -{ -lean_object* x_441; lean_object* x_442; uint8_t x_443; -x_441 = lean_array_get_size(x_4); -x_442 = lean_unsigned_to_nat(2u); -x_443 = lean_nat_dec_eq(x_441, x_442); -if (x_443 == 0) +case 4: { -lean_object* x_444; -lean_dec(x_441); -x_444 = lean_box(0); -x_426 = x_444; -goto block_437; +lean_object* x_57; +x_57 = l_Lean_Meta_Grind_mkENode(x_1, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_57; } -else -{ -lean_object* x_445; uint8_t x_446; -lean_dec(x_425); -lean_dec(x_3); -x_445 = lean_unsigned_to_nat(0u); -x_446 = lean_nat_dec_lt(x_445, x_441); -lean_dec(x_441); -if (x_446 == 0) +case 5: { -lean_object* x_447; lean_object* x_448; lean_object* x_449; -lean_dec(x_4); -x_447 = l_Lean_instInhabitedExpr; -x_448 = l_outOfBounds___rarg(x_447); -lean_inc(x_13); -lean_inc(x_12); +lean_object* x_58; 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); -lean_inc(x_448); -x_449 = l_Lean_Meta_Grind_internalize(x_448, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_449) == 0) +lean_inc(x_1); +x_58 = l_Lean_Meta_isLitValue(x_1, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_58) == 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_61; lean_object* x_62; 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(x_58, 1); +lean_inc(x_61); +lean_dec(x_58); +x_62 = lean_unsigned_to_nat(0u); +x_63 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_62); +x_64 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__1; +lean_inc(x_63); +x_65 = lean_mk_array(x_63, x_64); +x_66 = lean_unsigned_to_nat(1u); +x_67 = lean_nat_sub(x_63, x_66); +lean_dec(x_63); +lean_inc(x_1); +x_68 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3(x_1, x_2, x_1, x_65, x_67, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_61); +return x_68; +} +else { -lean_object* x_450; lean_object* x_451; lean_object* x_452; lean_object* x_453; lean_object* x_454; -x_450 = lean_ctor_get(x_449, 1); -lean_inc(x_450); -lean_dec(x_449); -lean_inc(x_1); -x_451 = l_Lean_Meta_Grind_registerParent(x_1, x_448, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_450); -lean_dec(x_448); -x_452 = lean_ctor_get(x_451, 0); -lean_inc(x_452); -x_453 = lean_ctor_get(x_451, 1); -lean_inc(x_453); -lean_dec(x_451); -x_454 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_452, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_453); -lean_dec(x_452); -return x_454; +lean_object* x_69; lean_object* x_70; +x_69 = lean_ctor_get(x_58, 1); +lean_inc(x_69); +lean_dec(x_58); +x_70 = l_Lean_Meta_Grind_mkENode(x_1, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_69); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_70; +} } else { -uint8_t x_455; -lean_dec(x_448); -lean_dec(x_13); -lean_dec(x_12); +uint8_t x_71; lean_dec(x_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_455 = !lean_is_exclusive(x_449); -if (x_455 == 0) +x_71 = !lean_is_exclusive(x_58); +if (x_71 == 0) { -return x_449; +return x_58; } else { -lean_object* x_456; lean_object* x_457; lean_object* x_458; -x_456 = lean_ctor_get(x_449, 0); -x_457 = lean_ctor_get(x_449, 1); -lean_inc(x_457); -lean_inc(x_456); -lean_dec(x_449); -x_458 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_458, 0, x_456); -lean_ctor_set(x_458, 1, x_457); -return x_458; +lean_object* x_72; lean_object* x_73; lean_object* x_74; +x_72 = lean_ctor_get(x_58, 0); +x_73 = lean_ctor_get(x_58, 1); +lean_inc(x_73); +lean_inc(x_72); +lean_dec(x_58); +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; +} } } +case 7: +{ +lean_object* x_75; lean_object* x_76; uint8_t x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; uint8_t x_81; lean_object* x_82; lean_object* x_105; +x_75 = lean_ctor_get(x_1, 1); +lean_inc(x_75); +x_76 = lean_ctor_get(x_1, 2); +lean_inc(x_76); +x_77 = 0; +lean_inc(x_2); +lean_inc(x_1); +x_78 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_77, x_77, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +x_79 = lean_ctor_get(x_78, 1); +lean_inc(x_79); +if (lean_is_exclusive(x_78)) { + lean_ctor_release(x_78, 0); + lean_ctor_release(x_78, 1); + x_80 = x_78; +} else { + lean_dec_ref(x_78); + x_80 = lean_box(0); +} +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_75); +x_105 = l_Lean_Meta_isProp(x_75, x_8, x_9, x_10, x_11, x_79); +if (lean_obj_tag(x_105) == 0) +{ +lean_object* x_106; uint8_t x_107; +x_106 = lean_ctor_get(x_105, 0); +lean_inc(x_106); +x_107 = lean_unbox(x_106); +if (x_107 == 0) +{ +lean_object* x_108; uint8_t x_109; +x_108 = lean_ctor_get(x_105, 1); +lean_inc(x_108); +lean_dec(x_105); +x_109 = lean_unbox(x_106); +lean_dec(x_106); +x_81 = x_109; +x_82 = x_108; +goto block_104; } else { -lean_object* x_459; lean_object* x_460; -x_459 = lean_array_fget(x_4, x_445); -lean_dec(x_4); -lean_inc(x_13); -lean_inc(x_12); +lean_object* x_110; lean_object* x_111; +lean_dec(x_106); +x_110 = lean_ctor_get(x_105, 1); +lean_inc(x_110); +lean_dec(x_105); 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); -lean_inc(x_459); -x_460 = l_Lean_Meta_Grind_internalize(x_459, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_460) == 0) -{ -lean_object* x_461; lean_object* x_462; lean_object* x_463; lean_object* x_464; lean_object* x_465; -x_461 = lean_ctor_get(x_460, 1); -lean_inc(x_461); -lean_dec(x_460); lean_inc(x_1); -x_462 = l_Lean_Meta_Grind_registerParent(x_1, x_459, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_461); -lean_dec(x_459); -x_463 = lean_ctor_get(x_462, 0); -lean_inc(x_463); -x_464 = lean_ctor_get(x_462, 1); -lean_inc(x_464); -lean_dec(x_462); -x_465 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_463, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_464); -lean_dec(x_463); -return x_465; +x_111 = l_Lean_Meta_isProp(x_1, x_8, x_9, x_10, x_11, x_110); +if (lean_obj_tag(x_111) == 0) +{ +lean_object* x_112; lean_object* x_113; uint8_t x_114; +x_112 = lean_ctor_get(x_111, 0); +lean_inc(x_112); +x_113 = lean_ctor_get(x_111, 1); +lean_inc(x_113); +lean_dec(x_111); +x_114 = lean_unbox(x_112); +lean_dec(x_112); +x_81 = x_114; +x_82 = x_113; +goto block_104; } else { -uint8_t x_466; -lean_dec(x_459); -lean_dec(x_13); -lean_dec(x_12); +uint8_t x_115; +lean_dec(x_80); +lean_dec(x_76); +lean_dec(x_75); lean_dec(x_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_466 = !lean_is_exclusive(x_460); -if (x_466 == 0) +x_115 = !lean_is_exclusive(x_111); +if (x_115 == 0) { -return x_460; +return x_111; } else { -lean_object* x_467; lean_object* x_468; lean_object* x_469; -x_467 = lean_ctor_get(x_460, 0); -x_468 = lean_ctor_get(x_460, 1); -lean_inc(x_468); -lean_inc(x_467); -lean_dec(x_460); -x_469 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_469, 0, x_467); -lean_ctor_set(x_469, 1, x_468); -return x_469; -} -} +lean_object* x_116; lean_object* x_117; lean_object* x_118; +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); +x_118 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_118, 0, x_116); +lean_ctor_set(x_118, 1, x_117); +return x_118; } } } -block_437: -{ -lean_object* x_427; -lean_dec(x_426); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_2); -lean_inc(x_3); -x_427 = l_Lean_Meta_Grind_internalize(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_427) == 0) -{ -lean_object* x_428; lean_object* x_429; lean_object* x_430; lean_object* x_431; lean_object* x_432; -x_428 = lean_ctor_get(x_427, 1); -lean_inc(x_428); -lean_dec(x_427); -lean_inc(x_1); -x_429 = l_Lean_Meta_Grind_registerParent(x_1, x_3, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_428); -lean_dec(x_3); -x_430 = lean_ctor_get(x_429, 0); -lean_inc(x_430); -x_431 = lean_ctor_get(x_429, 1); -lean_inc(x_431); -lean_dec(x_429); -x_432 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(x_4, x_1, x_2, x_425, x_430, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_431); -lean_dec(x_430); -lean_dec(x_4); -return x_432; } else { -uint8_t x_433; -lean_dec(x_425); -lean_dec(x_13); -lean_dec(x_12); +uint8_t x_119; +lean_dec(x_80); +lean_dec(x_76); +lean_dec(x_75); lean_dec(x_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_433 = !lean_is_exclusive(x_427); -if (x_433 == 0) +x_119 = !lean_is_exclusive(x_105); +if (x_119 == 0) { -return x_427; +return x_105; } else { -lean_object* x_434; lean_object* x_435; lean_object* x_436; -x_434 = lean_ctor_get(x_427, 0); -x_435 = lean_ctor_get(x_427, 1); -lean_inc(x_435); -lean_inc(x_434); -lean_dec(x_427); -x_436 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_436, 0, x_434); -lean_ctor_set(x_436, 1, x_435); -return x_436; -} -} +lean_object* x_120; lean_object* x_121; lean_object* x_122; +x_120 = lean_ctor_get(x_105, 0); +x_121 = lean_ctor_get(x_105, 1); +lean_inc(x_121); +lean_inc(x_120); +lean_dec(x_105); +x_122 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_122, 0, x_120); +lean_ctor_set(x_122, 1, x_121); +return x_122; } } -default: +block_104: { -lean_object* x_470; lean_object* x_471; lean_object* x_483; uint8_t x_484; -lean_dec(x_5); -lean_inc(x_2); -lean_inc(x_1); -x_470 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1___boxed), 12, 2); -lean_closure_set(x_470, 0, x_1); -lean_closure_set(x_470, 1, x_2); -x_483 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__4; -x_484 = l_Lean_Expr_isConstOf(x_3, x_483); -if (x_484 == 0) +if (x_81 == 0) { -lean_object* x_485; -x_485 = lean_box(0); -x_471 = x_485; -goto block_482; +lean_object* x_83; lean_object* x_84; +lean_dec(x_76); +lean_dec(x_75); +lean_dec(x_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_83 = lean_box(0); +if (lean_is_scalar(x_80)) { + x_84 = lean_alloc_ctor(0, 2, 0); +} else { + x_84 = x_80; } -else -{ -lean_object* x_486; lean_object* x_487; uint8_t x_488; -x_486 = lean_array_get_size(x_4); -x_487 = lean_unsigned_to_nat(2u); -x_488 = lean_nat_dec_eq(x_486, x_487); -if (x_488 == 0) -{ -lean_object* x_489; -lean_dec(x_486); -x_489 = lean_box(0); -x_471 = x_489; -goto block_482; +lean_ctor_set(x_84, 0, x_83); +lean_ctor_set(x_84, 1, x_82); +return x_84; } else { -lean_object* x_490; uint8_t x_491; -lean_dec(x_470); -lean_dec(x_3); -x_490 = lean_unsigned_to_nat(0u); -x_491 = lean_nat_dec_lt(x_490, x_486); -lean_dec(x_486); -if (x_491 == 0) -{ -lean_object* x_492; lean_object* x_493; lean_object* x_494; -lean_dec(x_4); -x_492 = l_Lean_instInhabitedExpr; -x_493 = l_outOfBounds___rarg(x_492); -lean_inc(x_13); -lean_inc(x_12); +lean_object* x_85; +lean_dec(x_80); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); lean_inc(x_2); -lean_inc(x_493); -x_494 = l_Lean_Meta_Grind_internalize(x_493, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_494) == 0) +lean_inc(x_75); +x_85 = l_Lean_Meta_Grind_internalize(x_75, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_82); +if (lean_obj_tag(x_85) == 0) { -lean_object* x_495; lean_object* x_496; lean_object* x_497; lean_object* x_498; lean_object* x_499; -x_495 = lean_ctor_get(x_494, 1); -lean_inc(x_495); -lean_dec(x_494); +lean_object* x_86; lean_object* x_87; lean_object* x_88; uint8_t x_89; +x_86 = lean_ctor_get(x_85, 1); +lean_inc(x_86); +lean_dec(x_85); lean_inc(x_1); -x_496 = l_Lean_Meta_Grind_registerParent(x_1, x_493, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_495); -lean_dec(x_493); -x_497 = lean_ctor_get(x_496, 0); -lean_inc(x_497); -x_498 = lean_ctor_get(x_496, 1); -lean_inc(x_498); -lean_dec(x_496); -x_499 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_497, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_498); -lean_dec(x_497); -return x_499; +x_87 = l_Lean_Meta_Grind_registerParent(x_1, x_75, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_86); +lean_dec(x_75); +x_88 = lean_ctor_get(x_87, 1); +lean_inc(x_88); +lean_dec(x_87); +x_89 = l_Lean_Expr_hasLooseBVars(x_76); +if (x_89 == 0) +{ +lean_object* 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); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_76); +x_90 = l_Lean_Meta_Grind_internalize(x_76, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_88); +if (lean_obj_tag(x_90) == 0) +{ +lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; +x_91 = lean_ctor_get(x_90, 1); +lean_inc(x_91); +lean_dec(x_90); +lean_inc(x_1); +x_92 = l_Lean_Meta_Grind_registerParent(x_1, x_76, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_91); +lean_dec(x_76); +x_93 = lean_ctor_get(x_92, 1); +lean_inc(x_93); +lean_dec(x_92); +x_94 = l_Lean_Meta_Grind_propagateUp(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_93); +return x_94; } else { -uint8_t x_500; -lean_dec(x_493); -lean_dec(x_13); -lean_dec(x_12); +uint8_t x_95; +lean_dec(x_76); lean_dec(x_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_5); +lean_dec(x_4); lean_dec(x_1); -x_500 = !lean_is_exclusive(x_494); -if (x_500 == 0) +x_95 = !lean_is_exclusive(x_90); +if (x_95 == 0) { -return x_494; +return x_90; } else { -lean_object* x_501; lean_object* x_502; lean_object* x_503; -x_501 = lean_ctor_get(x_494, 0); -x_502 = lean_ctor_get(x_494, 1); -lean_inc(x_502); -lean_inc(x_501); -lean_dec(x_494); -x_503 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_503, 0, x_501); -lean_ctor_set(x_503, 1, x_502); -return x_503; +lean_object* x_96; lean_object* x_97; lean_object* x_98; +x_96 = lean_ctor_get(x_90, 0); +x_97 = lean_ctor_get(x_90, 1); +lean_inc(x_97); +lean_inc(x_96); +lean_dec(x_90); +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 { -lean_object* x_504; lean_object* x_505; -x_504 = lean_array_fget(x_4, x_490); -lean_dec(x_4); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_2); -lean_inc(x_504); -x_505 = l_Lean_Meta_Grind_internalize(x_504, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_505) == 0) -{ -lean_object* x_506; lean_object* x_507; lean_object* x_508; lean_object* x_509; lean_object* x_510; -x_506 = lean_ctor_get(x_505, 1); -lean_inc(x_506); -lean_dec(x_505); -lean_inc(x_1); -x_507 = l_Lean_Meta_Grind_registerParent(x_1, x_504, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_506); -lean_dec(x_504); -x_508 = lean_ctor_get(x_507, 0); -lean_inc(x_508); -x_509 = lean_ctor_get(x_507, 1); -lean_inc(x_509); -lean_dec(x_507); -x_510 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_508, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_509); -lean_dec(x_508); -return x_510; +lean_object* x_99; +lean_dec(x_76); +lean_dec(x_2); +x_99 = l_Lean_Meta_Grind_propagateUp(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_88); +return x_99; +} } else { -uint8_t x_511; -lean_dec(x_504); -lean_dec(x_13); -lean_dec(x_12); +uint8_t x_100; +lean_dec(x_76); +lean_dec(x_75); lean_dec(x_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_511 = !lean_is_exclusive(x_505); -if (x_511 == 0) +x_100 = !lean_is_exclusive(x_85); +if (x_100 == 0) { -return x_505; +return x_85; } else { -lean_object* x_512; lean_object* x_513; lean_object* x_514; -x_512 = lean_ctor_get(x_505, 0); -x_513 = lean_ctor_get(x_505, 1); -lean_inc(x_513); -lean_inc(x_512); -lean_dec(x_505); -x_514 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_514, 0, x_512); -lean_ctor_set(x_514, 1, x_513); -return x_514; +lean_object* x_101; lean_object* x_102; lean_object* x_103; +x_101 = lean_ctor_get(x_85, 0); +x_102 = lean_ctor_get(x_85, 1); +lean_inc(x_102); +lean_inc(x_101); +lean_dec(x_85); +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; } } } } } -block_482: +case 9: { -lean_object* x_472; -lean_dec(x_471); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_2); -lean_inc(x_3); -x_472 = l_Lean_Meta_Grind_internalize(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_472) == 0) +lean_object* x_123; +x_123 = l_Lean_Meta_Grind_mkENode(x_1, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_123; +} +case 10: { -lean_object* x_473; lean_object* x_474; lean_object* x_475; lean_object* x_476; lean_object* x_477; -x_473 = lean_ctor_get(x_472, 1); -lean_inc(x_473); -lean_dec(x_472); +lean_object* x_124; lean_object* x_125; lean_object* x_126; uint8_t x_127; +x_124 = l_Lean_Meta_Grind_addCongrTable___closed__2; +x_125 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_124, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +x_126 = lean_ctor_get(x_125, 0); +lean_inc(x_126); +x_127 = lean_unbox(x_126); +lean_dec(x_126); +if (x_127 == 0) +{ +lean_object* x_128; uint8_t x_129; lean_object* x_130; +x_128 = lean_ctor_get(x_125, 1); +lean_inc(x_128); +lean_dec(x_125); +x_129 = 0; +x_130 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_129, x_129, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_128); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_130; +} +else +{ +uint8_t x_131; +x_131 = !lean_is_exclusive(x_125); +if (x_131 == 0) +{ +lean_object* x_132; lean_object* x_133; lean_object* x_134; +x_132 = lean_ctor_get(x_125, 1); +x_133 = lean_ctor_get(x_125, 0); +lean_dec(x_133); +x_134 = l_Lean_Meta_Grind_updateLastTag(x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_132); +if (lean_obj_tag(x_134) == 0) +{ +lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; uint8_t x_142; lean_object* x_143; +x_135 = lean_ctor_get(x_134, 1); +lean_inc(x_135); +lean_dec(x_134); lean_inc(x_1); -x_474 = l_Lean_Meta_Grind_registerParent(x_1, x_3, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_473); -lean_dec(x_3); -x_475 = lean_ctor_get(x_474, 0); -lean_inc(x_475); -x_476 = lean_ctor_get(x_474, 1); -lean_inc(x_476); -lean_dec(x_474); -x_477 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(x_4, x_1, x_2, x_470, x_475, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_476); -lean_dec(x_475); +x_136 = l_Lean_indentExpr(x_1); +x_137 = l_Lean_Meta_Grind_internalize___lambda__3___closed__6; +lean_ctor_set_tag(x_125, 7); +lean_ctor_set(x_125, 1, x_136); +lean_ctor_set(x_125, 0, x_137); +x_138 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__6; +x_139 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_139, 0, x_125); +lean_ctor_set(x_139, 1, x_138); +x_140 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_124, x_139, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_135); +x_141 = lean_ctor_get(x_140, 1); +lean_inc(x_141); +lean_dec(x_140); +x_142 = 0; +x_143 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_142, x_142, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_141); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); lean_dec(x_4); -return x_477; +return x_143; } else { -uint8_t x_478; -lean_dec(x_470); -lean_dec(x_13); -lean_dec(x_12); +uint8_t x_144; +lean_free_object(x_125); lean_dec(x_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_478 = !lean_is_exclusive(x_472); -if (x_478 == 0) +x_144 = !lean_is_exclusive(x_134); +if (x_144 == 0) { -return x_472; +return x_134; } else { -lean_object* x_479; lean_object* x_480; lean_object* x_481; -x_479 = lean_ctor_get(x_472, 0); -x_480 = lean_ctor_get(x_472, 1); -lean_inc(x_480); -lean_inc(x_479); -lean_dec(x_472); -x_481 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_481, 0, x_479); -lean_ctor_set(x_481, 1, x_480); -return x_481; -} -} -} -} +lean_object* x_145; lean_object* x_146; lean_object* x_147; +x_145 = lean_ctor_get(x_134, 0); +x_146 = lean_ctor_get(x_134, 1); +lean_inc(x_146); +lean_inc(x_145); +lean_dec(x_134); +x_147 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_147, 0, x_145); +lean_ctor_set(x_147, 1, x_146); +return x_147; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_internalize___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: +else { -uint8_t x_13; lean_object* x_14; -x_13 = 0; -x_14 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_13, x_13, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -return x_14; -} -} -static lean_object* _init_l_Lean_Meta_Grind_internalize___lambda__2___closed__1() { -_start: +lean_object* x_148; lean_object* x_149; +x_148 = lean_ctor_get(x_125, 1); +lean_inc(x_148); +lean_dec(x_125); +x_149 = l_Lean_Meta_Grind_updateLastTag(x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_148); +if (lean_obj_tag(x_149) == 0) { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean.Meta.Tactic.Grind.Internalize", 34, 34); -return x_1; -} +lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; uint8_t x_158; lean_object* x_159; +x_150 = lean_ctor_get(x_149, 1); +lean_inc(x_150); +lean_dec(x_149); +lean_inc(x_1); +x_151 = l_Lean_indentExpr(x_1); +x_152 = l_Lean_Meta_Grind_internalize___lambda__3___closed__6; +x_153 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_153, 0, x_152); +lean_ctor_set(x_153, 1, x_151); +x_154 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__6; +x_155 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_155, 0, x_153); +lean_ctor_set(x_155, 1, x_154); +x_156 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_124, x_155, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_150); +x_157 = lean_ctor_get(x_156, 1); +lean_inc(x_157); +lean_dec(x_156); +x_158 = 0; +x_159 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_158, x_158, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_157); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_159; } -static lean_object* _init_l_Lean_Meta_Grind_internalize___lambda__2___closed__2() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean.Meta.Grind.internalize", 27, 27); -return x_1; +lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; +lean_dec(x_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_160 = lean_ctor_get(x_149, 0); +lean_inc(x_160); +x_161 = lean_ctor_get(x_149, 1); +lean_inc(x_161); +if (lean_is_exclusive(x_149)) { + lean_ctor_release(x_149, 0); + lean_ctor_release(x_149, 1); + x_162 = x_149; +} else { + lean_dec_ref(x_149); + x_162 = lean_box(0); } +if (lean_is_scalar(x_162)) { + x_163 = lean_alloc_ctor(1, 2, 0); +} else { + x_163 = x_162; } -static lean_object* _init_l_Lean_Meta_Grind_internalize___lambda__2___closed__3() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("unreachable code has been reached", 33, 33); -return x_1; +lean_ctor_set(x_163, 0, x_160); +lean_ctor_set(x_163, 1, x_161); +return x_163; } } -static lean_object* _init_l_Lean_Meta_Grind_internalize___lambda__2___closed__4() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_1 = l_Lean_Meta_Grind_internalize___lambda__2___closed__1; -x_2 = l_Lean_Meta_Grind_internalize___lambda__2___closed__2; -x_3 = lean_unsigned_to_nat(76u); -x_4 = lean_unsigned_to_nat(16u); -x_5 = l_Lean_Meta_Grind_internalize___lambda__2___closed__3; -x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); -return x_6; } } -static lean_object* _init_l_Lean_Meta_Grind_internalize___lambda__2___closed__5() { -_start: +case 11: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("unexpected term during internalization", 38, 38); -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_Grind_internalize___lambda__2___closed__6() { -_start: +lean_object* x_164; lean_object* x_165; lean_object* x_166; uint8_t x_167; +x_164 = l_Lean_Meta_Grind_addCongrTable___closed__2; +x_165 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_164, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +x_166 = lean_ctor_get(x_165, 0); +lean_inc(x_166); +x_167 = lean_unbox(x_166); +lean_dec(x_166); +if (x_167 == 0) { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_Grind_internalize___lambda__2___closed__5; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; -} +lean_object* x_168; uint8_t x_169; lean_object* x_170; +x_168 = lean_ctor_get(x_165, 1); +lean_inc(x_168); +lean_dec(x_165); +x_169 = 0; +x_170 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_169, x_169, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_168); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_170; } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_internalize___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -switch (lean_obj_tag(x_1)) { -case 0: +else { -lean_object* x_13; lean_object* x_14; -lean_dec(x_2); -lean_dec(x_1); -x_13 = l_Lean_Meta_Grind_internalize___lambda__2___closed__4; -x_14 = l_panic___at_Lean_Meta_Grind_internalize___spec__1(x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -return x_14; -} -case 2: +uint8_t x_171; +x_171 = !lean_is_exclusive(x_165); +if (x_171 == 0) { -lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; -x_15 = l_Lean_Meta_Grind_addCongrTable___closed__2; -x_16 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_15, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -x_18 = lean_unbox(x_17); -lean_dec(x_17); -if (x_18 == 0) +lean_object* x_172; lean_object* x_173; lean_object* x_174; +x_172 = lean_ctor_get(x_165, 1); +x_173 = lean_ctor_get(x_165, 0); +lean_dec(x_173); +x_174 = l_Lean_Meta_Grind_updateLastTag(x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_172); +if (lean_obj_tag(x_174) == 0) { -lean_object* x_19; uint8_t x_20; lean_object* x_21; -x_19 = lean_ctor_get(x_16, 1); -lean_inc(x_19); -lean_dec(x_16); -x_20 = 0; -x_21 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_20, x_20, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_19); +lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; uint8_t x_182; lean_object* x_183; +x_175 = lean_ctor_get(x_174, 1); +lean_inc(x_175); +lean_dec(x_174); +lean_inc(x_1); +x_176 = l_Lean_indentExpr(x_1); +x_177 = l_Lean_Meta_Grind_internalize___lambda__3___closed__6; +lean_ctor_set_tag(x_165, 7); +lean_ctor_set(x_165, 1, x_176); +lean_ctor_set(x_165, 0, x_177); +x_178 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__6; +x_179 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_179, 0, x_165); +lean_ctor_set(x_179, 1, x_178); +x_180 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_164, x_179, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_175); +x_181 = lean_ctor_get(x_180, 1); +lean_inc(x_181); +lean_dec(x_180); +x_182 = 0; +x_183 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_182, x_182, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_181); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -7701,34 +11728,12 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -return x_21; +return x_183; } else { -uint8_t x_22; -x_22 = !lean_is_exclusive(x_16); -if (x_22 == 0) -{ -lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; lean_object* x_32; -x_23 = lean_ctor_get(x_16, 1); -x_24 = lean_ctor_get(x_16, 0); -lean_dec(x_24); -lean_inc(x_1); -x_25 = l_Lean_indentExpr(x_1); -x_26 = l_Lean_Meta_Grind_internalize___lambda__2___closed__6; -lean_ctor_set_tag(x_16, 7); -lean_ctor_set(x_16, 1, x_25); -lean_ctor_set(x_16, 0, x_26); -x_27 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__5; -x_28 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_28, 0, x_16); -lean_ctor_set(x_28, 1, x_27); -x_29 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_15, x_28, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_23); -x_30 = lean_ctor_get(x_29, 1); -lean_inc(x_30); -lean_dec(x_29); -x_31 = 0; -x_32 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_31, x_31, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_30); +uint8_t x_184; +lean_free_object(x_165); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -7737,30 +11742,57 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -return x_32; +lean_dec(x_2); +lean_dec(x_1); +x_184 = !lean_is_exclusive(x_174); +if (x_184 == 0) +{ +return x_174; } 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; uint8_t x_41; lean_object* x_42; -x_33 = lean_ctor_get(x_16, 1); -lean_inc(x_33); -lean_dec(x_16); +lean_object* x_185; lean_object* x_186; lean_object* x_187; +x_185 = lean_ctor_get(x_174, 0); +x_186 = lean_ctor_get(x_174, 1); +lean_inc(x_186); +lean_inc(x_185); +lean_dec(x_174); +x_187 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_187, 0, x_185); +lean_ctor_set(x_187, 1, x_186); +return x_187; +} +} +} +else +{ +lean_object* x_188; lean_object* x_189; +x_188 = lean_ctor_get(x_165, 1); +lean_inc(x_188); +lean_dec(x_165); +x_189 = l_Lean_Meta_Grind_updateLastTag(x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, 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_195; lean_object* x_196; lean_object* x_197; uint8_t x_198; lean_object* x_199; +x_190 = lean_ctor_get(x_189, 1); +lean_inc(x_190); +lean_dec(x_189); lean_inc(x_1); -x_34 = l_Lean_indentExpr(x_1); -x_35 = l_Lean_Meta_Grind_internalize___lambda__2___closed__6; -x_36 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_36, 0, x_35); -lean_ctor_set(x_36, 1, x_34); -x_37 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__5; -x_38 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_38, 0, x_36); -lean_ctor_set(x_38, 1, x_37); -x_39 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_15, x_38, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_33); -x_40 = lean_ctor_get(x_39, 1); -lean_inc(x_40); -lean_dec(x_39); -x_41 = 0; -x_42 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_41, x_41, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_40); +x_191 = l_Lean_indentExpr(x_1); +x_192 = l_Lean_Meta_Grind_internalize___lambda__3___closed__6; +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_Meta_Grind_addCongrTable___lambda__2___closed__6; +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_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_164, x_195, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_190); +x_197 = lean_ctor_get(x_196, 1); +lean_inc(x_197); +lean_dec(x_196); +x_198 = 0; +x_199 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_198, x_198, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_197); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -7769,13 +11801,11 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -return x_42; -} -} +return x_199; } -case 3: +else { -lean_object* x_43; lean_object* x_44; +lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -7786,73 +11816,126 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_43 = lean_box(0); -x_44 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_44, 0, x_43); -lean_ctor_set(x_44, 1, x_12); -return x_44; +x_200 = lean_ctor_get(x_189, 0); +lean_inc(x_200); +x_201 = lean_ctor_get(x_189, 1); +lean_inc(x_201); +if (lean_is_exclusive(x_189)) { + lean_ctor_release(x_189, 0); + lean_ctor_release(x_189, 1); + x_202 = x_189; +} else { + lean_dec_ref(x_189); + x_202 = lean_box(0); } -case 4: +if (lean_is_scalar(x_202)) { + x_203 = lean_alloc_ctor(1, 2, 0); +} else { + x_203 = x_202; +} +lean_ctor_set(x_203, 0, x_200); +lean_ctor_set(x_203, 1, x_201); +return x_203; +} +} +} +} +default: { -lean_object* x_45; -x_45 = l_Lean_Meta_Grind_mkENode(x_1, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +uint8_t x_204; lean_object* x_205; +x_204 = 0; +x_205 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_204, x_204, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -return x_45; +return x_205; } -case 5: +} +} +} +static lean_object* _init_l_Lean_Meta_Grind_internalize___lambda__4___closed__1() { +_start: { -lean_object* x_46; -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_1); -x_46 = l_Lean_Meta_isLitValue(x_1, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_46) == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("internalize", 11, 11); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_internalize___lambda__4___closed__2() { +_start: { -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); -lean_dec(x_47); -if (x_48 == 0) +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__1; +x_2 = l_Lean_Meta_Grind_internalize___lambda__4___closed__1; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_internalize___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) { +_start: { -lean_object* x_49; lean_object* x_50; 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_46, 1); -lean_inc(x_49); -lean_dec(x_46); -x_50 = lean_unsigned_to_nat(0u); -x_51 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_50); -x_52 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__1; -lean_inc(x_51); -x_53 = lean_mk_array(x_51, x_52); -x_54 = lean_unsigned_to_nat(1u); -x_55 = lean_nat_sub(x_51, x_54); -lean_dec(x_51); -lean_inc(x_1); -x_56 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3(x_1, x_2, x_1, x_53, x_55, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_49); -return x_56; +lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_13 = l_Lean_Meta_Grind_internalize___lambda__4___closed__2; +x_14 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +x_16 = lean_unbox(x_15); +lean_dec(x_15); +if (x_16 == 0) +{ +lean_object* x_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_Grind_internalize___lambda__3(x_1, x_2, x_18, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_17); +return x_19; } else { -lean_object* x_57; lean_object* x_58; -x_57 = lean_ctor_get(x_46, 1); -lean_inc(x_57); -lean_dec(x_46); -x_58 = l_Lean_Meta_Grind_mkENode(x_1, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_57); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -return x_58; -} +uint8_t x_20; +x_20 = !lean_is_exclusive(x_14); +if (x_20 == 0) +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_14, 1); +x_22 = lean_ctor_get(x_14, 0); +lean_dec(x_22); +x_23 = l_Lean_Meta_Grind_updateLastTag(x_4, x_5, 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; 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_23, 1); +lean_inc(x_24); +lean_dec(x_23); +lean_inc(x_1); +x_25 = l_Lean_MessageData_ofExpr(x_1); +x_26 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__6; +lean_ctor_set_tag(x_14, 7); +lean_ctor_set(x_14, 1, x_25); +lean_ctor_set(x_14, 0, x_26); +x_27 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_27, 0, x_14); +lean_ctor_set(x_27, 1, x_26); +x_28 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_13, x_27, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, 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 = l_Lean_Meta_Grind_internalize___lambda__3(x_1, x_2, x_29, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_30); +lean_dec(x_29); +return x_31; } else { -uint8_t x_59; +uint8_t x_32; +lean_free_object(x_14); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -7863,55 +11946,61 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_59 = !lean_is_exclusive(x_46); -if (x_59 == 0) +x_32 = !lean_is_exclusive(x_23); +if (x_32 == 0) { -return x_46; +return x_23; } else { -lean_object* x_60; lean_object* x_61; lean_object* x_62; -x_60 = lean_ctor_get(x_46, 0); -x_61 = lean_ctor_get(x_46, 1); -lean_inc(x_61); -lean_inc(x_60); -lean_dec(x_46); -x_62 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_62, 0, x_60); -lean_ctor_set(x_62, 1, x_61); -return x_62; +lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_33 = lean_ctor_get(x_23, 0); +x_34 = lean_ctor_get(x_23, 1); +lean_inc(x_34); +lean_inc(x_33); +lean_dec(x_23); +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; } } } -case 7: +else { -lean_object* x_63; uint8_t x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; -x_63 = lean_ctor_get(x_1, 1); -lean_inc(x_63); -x_64 = 0; -lean_inc(x_2); -lean_inc(x_1); -x_65 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_64, x_64, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -x_66 = lean_ctor_get(x_65, 1); -lean_inc(x_66); -lean_dec(x_65); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_63); -x_67 = l_Lean_Meta_isProp(x_63, x_8, x_9, x_10, x_11, x_66); -if (lean_obj_tag(x_67) == 0) +lean_object* x_36; lean_object* x_37; +x_36 = lean_ctor_get(x_14, 1); +lean_inc(x_36); +lean_dec(x_14); +x_37 = l_Lean_Meta_Grind_updateLastTag(x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_36); +if (lean_obj_tag(x_37) == 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) +lean_object* x_38; lean_object* x_39; lean_object* x_40; 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, 1); +lean_inc(x_38); +lean_dec(x_37); +lean_inc(x_1); +x_39 = l_Lean_MessageData_ofExpr(x_1); +x_40 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__6; +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 = 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_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_13, x_42, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_38); +x_44 = lean_ctor_get(x_43, 0); +lean_inc(x_44); +x_45 = lean_ctor_get(x_43, 1); +lean_inc(x_45); +lean_dec(x_43); +x_46 = l_Lean_Meta_Grind_internalize___lambda__3(x_1, x_2, x_44, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_45); +lean_dec(x_44); +return x_46; +} +else { -uint8_t x_70; -lean_dec(x_63); +lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -7922,53 +12011,53 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_70 = !lean_is_exclusive(x_67); -if (x_70 == 0) -{ -lean_object* x_71; lean_object* x_72; -x_71 = lean_ctor_get(x_67, 0); -lean_dec(x_71); -x_72 = lean_box(0); -lean_ctor_set(x_67, 0, x_72); -return x_67; +x_47 = lean_ctor_get(x_37, 0); +lean_inc(x_47); +x_48 = lean_ctor_get(x_37, 1); +lean_inc(x_48); +if (lean_is_exclusive(x_37)) { + lean_ctor_release(x_37, 0); + lean_ctor_release(x_37, 1); + x_49 = x_37; +} else { + lean_dec_ref(x_37); + x_49 = lean_box(0); } -else -{ -lean_object* x_73; lean_object* x_74; lean_object* x_75; -x_73 = lean_ctor_get(x_67, 1); -lean_inc(x_73); -lean_dec(x_67); -x_74 = lean_box(0); -x_75 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_75, 0, x_74); -lean_ctor_set(x_75, 1, x_73); -return x_75; +if (lean_is_scalar(x_49)) { + x_50 = lean_alloc_ctor(1, 2, 0); +} else { + x_50 = x_49; +} +lean_ctor_set(x_50, 0, x_47); +lean_ctor_set(x_50, 1, x_48); +return x_50; } } -else +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_internalize(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, 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_76; lean_object* x_77; -x_76 = lean_ctor_get(x_67, 1); -lean_inc(x_76); -lean_dec(x_67); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_1); -x_77 = l_Lean_Meta_isProp(x_1, x_8, x_9, x_10, x_11, x_76); -if (lean_obj_tag(x_77) == 0) +lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_12 = l_Lean_Meta_Grind_alreadyInternalized(x_1, 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_78; uint8_t x_79; -x_78 = lean_ctor_get(x_77, 0); -lean_inc(x_78); -x_79 = lean_unbox(x_78); -lean_dec(x_78); -if (x_79 == 0) +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_ctor_get(x_12, 1); +lean_inc(x_15); +lean_dec(x_12); +x_16 = lean_box(0); +x_17 = l_Lean_Meta_Grind_internalize___lambda__4(x_1, x_2, x_16, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_15); +return x_17; +} +else { -uint8_t x_80; -lean_dec(x_63); -lean_dec(x_11); +uint8_t x_18; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -7976,1036 +12065,1826 @@ 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_80 = !lean_is_exclusive(x_77); -if (x_80 == 0) +x_18 = !lean_is_exclusive(x_12); +if (x_18 == 0) { -lean_object* x_81; lean_object* x_82; -x_81 = lean_ctor_get(x_77, 0); -lean_dec(x_81); -x_82 = lean_box(0); -lean_ctor_set(x_77, 0, x_82); -return x_77; +lean_object* x_19; lean_object* x_20; +x_19 = lean_ctor_get(x_12, 0); +lean_dec(x_19); +x_20 = lean_box(0); +lean_ctor_set(x_12, 0, x_20); +return x_12; } else { -lean_object* x_83; lean_object* x_84; lean_object* x_85; -x_83 = lean_ctor_get(x_77, 1); -lean_inc(x_83); -lean_dec(x_77); -x_84 = lean_box(0); -x_85 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_85, 0, x_84); -lean_ctor_set(x_85, 1, x_83); -return x_85; +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_12, 1); +lean_inc(x_21); +lean_dec(x_12); +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_EXPORT uint8_t l_Lean_PersistentHashMap_containsAtAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___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_86; lean_object* x_87; -x_86 = lean_ctor_get(x_77, 1); -lean_inc(x_86); -lean_dec(x_77); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_63); -x_87 = l_Lean_Meta_Grind_internalize(x_63, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_86); -if (lean_obj_tag(x_87) == 0) +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_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; -x_88 = lean_ctor_get(x_87, 1); -lean_inc(x_88); -lean_dec(x_87); -lean_inc(x_1); -x_89 = l_Lean_Meta_Grind_registerParent(x_1, x_63, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_88); -lean_dec(x_63); -x_90 = lean_ctor_get(x_89, 1); -lean_inc(x_90); -lean_dec(x_89); -x_91 = l_Lean_Meta_Grind_propagateUp(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_90); -return x_91; +uint8_t x_8; +lean_dec(x_4); +x_8 = 0; +return x_8; } else { -uint8_t x_92; -lean_dec(x_63); -lean_dec(x_11); -lean_dec(x_10); +lean_object* x_9; uint8_t x_10; +x_9 = lean_array_fget(x_1, x_4); +x_10 = l___private_Lean_HeadIndex_0__Lean_beqHeadIndex____x40_Lean_HeadIndex___hyg_69_(x_5, x_9); 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_92 = !lean_is_exclusive(x_87); -if (x_92 == 0) +if (x_10 == 0) { -return x_87; +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_93; lean_object* x_94; lean_object* x_95; -x_93 = lean_ctor_get(x_87, 0); -x_94 = lean_ctor_get(x_87, 1); -lean_inc(x_94); -lean_inc(x_93); -lean_dec(x_87); -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; +uint8_t x_14; +lean_dec(x_4); +x_14 = 1; +return x_14; } } } } -else +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__2(lean_object* x_1, size_t x_2, lean_object* x_3) { +_start: { -uint8_t x_96; -lean_dec(x_63); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); +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_findEntryAux___at_Lean_Meta_Grind_addCongrTable___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_7); -lean_dec(x_6); -lean_dec(x_5); lean_dec(x_4); -lean_dec(x_2); -lean_dec(x_1); -x_96 = !lean_is_exclusive(x_77); -if (x_96 == 0) +switch (lean_obj_tag(x_10)) { +case 0: { -return x_77; +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 = l___private_Lean_HeadIndex_0__Lean_beqHeadIndex____x40_Lean_HeadIndex___hyg_69_(x_3, x_11); +lean_dec(x_11); +return x_12; } -else +case 1: { -lean_object* x_97; lean_object* x_98; lean_object* x_99; -x_97 = lean_ctor_get(x_77, 0); -x_98 = lean_ctor_get(x_77, 1); -lean_inc(x_98); -lean_inc(x_97); -lean_dec(x_77); -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_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 { -uint8_t x_100; -lean_dec(x_63); -lean_dec(x_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_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_100 = !lean_is_exclusive(x_67); -if (x_100 == 0) -{ -return x_67; -} -else -{ -lean_object* x_101; lean_object* x_102; lean_object* x_103; -x_101 = lean_ctor_get(x_67, 0); -x_102 = lean_ctor_get(x_67, 1); -lean_inc(x_102); -lean_inc(x_101); -lean_dec(x_67); -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; +x_19 = lean_unsigned_to_nat(0u); +x_20 = l_Lean_PersistentHashMap_containsAtAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__3(x_17, x_18, lean_box(0), x_19, x_3); +lean_dec(x_18); +lean_dec(x_17); +return x_20; } } } -case 9: +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_contains___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__1(lean_object* x_1, lean_object* x_2) { +_start: { -lean_object* x_104; -x_104 = l_Lean_Meta_Grind_mkENode(x_1, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -return x_104; +uint64_t x_3; size_t x_4; uint8_t x_5; +x_3 = l_Lean_HeadIndex_hash(x_2); +x_4 = lean_uint64_to_usize(x_3); +x_5 = l_Lean_PersistentHashMap_containsAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__2(x_1, x_4, x_2); +return x_5; } -case 10: +} +LEAN_EXPORT lean_object* l_List_filterTR_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -lean_object* x_105; lean_object* x_106; lean_object* x_107; uint8_t x_108; -x_105 = l_Lean_Meta_Grind_addCongrTable___closed__2; -x_106 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_105, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -x_107 = lean_ctor_get(x_106, 0); -lean_inc(x_107); -x_108 = lean_unbox(x_107); -lean_dec(x_107); -if (x_108 == 0) +if (lean_obj_tag(x_2) == 0) { -lean_object* x_109; uint8_t x_110; lean_object* x_111; -x_109 = lean_ctor_get(x_106, 1); -lean_inc(x_109); -lean_dec(x_106); -x_110 = 0; -x_111 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_110, x_110, x_2, x_4, x_5, 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); -lean_dec(x_5); -lean_dec(x_4); -return x_111; +lean_object* x_4; +lean_dec(x_1); +x_4 = l_List_reverse___rarg(x_3); +return x_4; } else { -uint8_t x_112; -x_112 = !lean_is_exclusive(x_106); -if (x_112 == 0) +uint8_t x_5; +x_5 = !lean_is_exclusive(x_2); +if (x_5 == 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; uint8_t x_121; lean_object* x_122; -x_113 = lean_ctor_get(x_106, 1); -x_114 = lean_ctor_get(x_106, 0); -lean_dec(x_114); +lean_object* x_6; lean_object* x_7; uint8_t x_8; +x_6 = lean_ctor_get(x_2, 0); +x_7 = lean_ctor_get(x_2, 1); lean_inc(x_1); -x_115 = l_Lean_indentExpr(x_1); -x_116 = l_Lean_Meta_Grind_internalize___lambda__2___closed__6; -lean_ctor_set_tag(x_106, 7); -lean_ctor_set(x_106, 1, x_115); -lean_ctor_set(x_106, 0, x_116); -x_117 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__5; -x_118 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_118, 0, x_106); -lean_ctor_set(x_118, 1, x_117); -x_119 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_105, x_118, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_113); -x_120 = lean_ctor_get(x_119, 1); -lean_inc(x_120); -lean_dec(x_119); -x_121 = 0; -x_122 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_121, x_121, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_120); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); +x_8 = l_Lean_PersistentHashMap_contains___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__1(x_1, x_6); +if (x_8 == 0) +{ +lean_ctor_set(x_2, 1, x_3); +{ +lean_object* _tmp_1 = x_7; +lean_object* _tmp_2 = x_2; +x_2 = _tmp_1; +x_3 = _tmp_2; +} +goto _start; +} +else +{ +lean_free_object(x_2); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -return x_122; +x_2 = x_7; +goto _start; +} } else { -lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; uint8_t x_131; lean_object* x_132; -x_123 = lean_ctor_get(x_106, 1); -lean_inc(x_123); -lean_dec(x_106); +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); lean_inc(x_1); -x_124 = l_Lean_indentExpr(x_1); -x_125 = l_Lean_Meta_Grind_internalize___lambda__2___closed__6; -x_126 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_126, 0, x_125); -lean_ctor_set(x_126, 1, x_124); -x_127 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__5; -x_128 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_128, 0, x_126); -lean_ctor_set(x_128, 1, x_127); -x_129 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_105, x_128, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_123); -x_130 = lean_ctor_get(x_129, 1); -lean_inc(x_130); -lean_dec(x_129); -x_131 = 0; -x_132 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_131, x_131, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_130); +x_13 = l_Lean_PersistentHashMap_contains___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__1(x_1, x_11); +if (x_13 == 0) +{ +lean_object* x_14; +x_14 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_14, 0, x_11); +lean_ctor_set(x_14, 1, x_3); +x_2 = x_12; +x_3 = x_14; +goto _start; +} +else +{ lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -return x_132; +x_2 = x_12; +goto _start; } } } -case 11: -{ -lean_object* x_133; lean_object* x_134; lean_object* x_135; uint8_t x_136; -x_133 = l_Lean_Meta_Grind_addCongrTable___closed__2; -x_134 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_133, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -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) +} +} +static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___lambda__1___closed__1() { +_start: { -lean_object* x_137; uint8_t x_138; lean_object* x_139; -x_137 = lean_ctor_get(x_134, 1); -lean_inc(x_137); -lean_dec(x_134); -x_138 = 0; -x_139 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_138, x_138, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_137); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -return x_139; +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; } -else +} +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___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: { -uint8_t x_140; -x_140 = !lean_is_exclusive(x_134); -if (x_140 == 0) +lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_12 = lean_st_ref_take(x_3, x_11); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_12, 1); +lean_inc(x_14); +lean_dec(x_12); +x_15 = !lean_is_exclusive(x_13); +if (x_15 == 0) { -lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; uint8_t x_149; lean_object* x_150; -x_141 = lean_ctor_get(x_134, 1); -x_142 = lean_ctor_get(x_134, 0); -lean_dec(x_142); -lean_inc(x_1); -x_143 = l_Lean_indentExpr(x_1); -x_144 = l_Lean_Meta_Grind_internalize___lambda__2___closed__6; -lean_ctor_set_tag(x_134, 7); -lean_ctor_set(x_134, 1, x_143); -lean_ctor_set(x_134, 0, x_144); -x_145 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__5; -x_146 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_146, 0, x_134); -lean_ctor_set(x_146, 1, x_145); -x_147 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_133, x_146, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_141); -x_148 = lean_ctor_get(x_147, 1); -lean_inc(x_148); -lean_dec(x_147); -x_149 = 0; -x_150 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_149, x_149, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_148); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -return x_150; +lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; +x_16 = lean_ctor_get(x_13, 10); +x_17 = l_Lean_Meta_Grind_EMatchTheorems_insert(x_16, x_1); +lean_ctor_set(x_13, 10, x_17); +x_18 = lean_st_ref_set(x_3, x_13, x_14); +x_19 = !lean_is_exclusive(x_18); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; +x_20 = lean_ctor_get(x_18, 0); +lean_dec(x_20); +x_21 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___lambda__1___closed__1; +lean_ctor_set(x_18, 0, x_21); +return x_18; } else { -lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; uint8_t x_159; lean_object* x_160; -x_151 = lean_ctor_get(x_134, 1); -lean_inc(x_151); -lean_dec(x_134); -lean_inc(x_1); -x_152 = l_Lean_indentExpr(x_1); -x_153 = l_Lean_Meta_Grind_internalize___lambda__2___closed__6; -x_154 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_154, 0, x_153); -lean_ctor_set(x_154, 1, x_152); -x_155 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__5; -x_156 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_156, 0, x_154); -lean_ctor_set(x_156, 1, x_155); -x_157 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_133, x_156, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_151); -x_158 = lean_ctor_get(x_157, 1); -lean_inc(x_158); -lean_dec(x_157); -x_159 = 0; -x_160 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_159, x_159, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_158); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -return x_160; +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 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___lambda__1___closed__1; +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set(x_24, 1, x_22); +return x_24; } } +else +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; 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_25 = lean_ctor_get(x_13, 0); +x_26 = lean_ctor_get(x_13, 1); +x_27 = lean_ctor_get(x_13, 2); +x_28 = lean_ctor_get(x_13, 3); +x_29 = lean_ctor_get(x_13, 4); +x_30 = lean_ctor_get(x_13, 5); +x_31 = lean_ctor_get_uint8(x_13, sizeof(void*)*20); +x_32 = lean_ctor_get(x_13, 6); +x_33 = lean_ctor_get(x_13, 7); +x_34 = lean_ctor_get(x_13, 8); +x_35 = lean_ctor_get(x_13, 9); +x_36 = lean_ctor_get(x_13, 10); +x_37 = lean_ctor_get(x_13, 11); +x_38 = lean_ctor_get(x_13, 12); +x_39 = lean_ctor_get(x_13, 13); +x_40 = lean_ctor_get(x_13, 14); +x_41 = lean_ctor_get(x_13, 15); +x_42 = lean_ctor_get(x_13, 16); +x_43 = lean_ctor_get(x_13, 17); +x_44 = lean_ctor_get(x_13, 18); +x_45 = lean_ctor_get(x_13, 19); +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_inc(x_34); +lean_inc(x_33); +lean_inc(x_32); +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_dec(x_13); +x_46 = l_Lean_Meta_Grind_EMatchTheorems_insert(x_36, x_1); +x_47 = lean_alloc_ctor(0, 20, 1); +lean_ctor_set(x_47, 0, x_25); +lean_ctor_set(x_47, 1, x_26); +lean_ctor_set(x_47, 2, x_27); +lean_ctor_set(x_47, 3, x_28); +lean_ctor_set(x_47, 4, x_29); +lean_ctor_set(x_47, 5, x_30); +lean_ctor_set(x_47, 6, x_32); +lean_ctor_set(x_47, 7, x_33); +lean_ctor_set(x_47, 8, x_34); +lean_ctor_set(x_47, 9, x_35); +lean_ctor_set(x_47, 10, x_46); +lean_ctor_set(x_47, 11, x_37); +lean_ctor_set(x_47, 12, x_38); +lean_ctor_set(x_47, 13, x_39); +lean_ctor_set(x_47, 14, x_40); +lean_ctor_set(x_47, 15, x_41); +lean_ctor_set(x_47, 16, x_42); +lean_ctor_set(x_47, 17, x_43); +lean_ctor_set(x_47, 18, x_44); +lean_ctor_set(x_47, 19, x_45); +lean_ctor_set_uint8(x_47, sizeof(void*)*20, x_31); +x_48 = lean_st_ref_set(x_3, x_47, x_14); +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); } -default: -{ -uint8_t x_161; lean_object* x_162; -x_161 = 0; -x_162 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_161, x_161, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -return x_162; +x_51 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___lambda__1___closed__1; +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; } } } -static lean_object* _init_l_Lean_Meta_Grind_internalize___lambda__3___closed__1() { +static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("internalize", 11, 11); +x_1 = lean_mk_string_unchecked("ematch", 6, 6); return x_1; } } -static lean_object* _init_l_Lean_Meta_Grind_internalize___lambda__3___closed__2() { +static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__1; -x_2 = l_Lean_Meta_Grind_internalize___lambda__3___closed__1; +x_2 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__1; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_internalize___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) { +static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__3() { _start: { -lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; -x_13 = l_Lean_Meta_Grind_internalize___lambda__3___closed__2; -x_14 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -x_15 = lean_ctor_get(x_14, 0); -lean_inc(x_15); -x_16 = lean_unbox(x_15); -lean_dec(x_15); -if (x_16 == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("reinsert `", 10, 10); +return x_1; +} +} +static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__4() { +_start: { -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_object* x_1; lean_object* x_2; +x_1 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__3; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("`", 1, 1); +return x_1; +} +} +static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__5; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17) { +_start: +{ +if (lean_obj_tag(x_6) == 0) +{ +lean_object* x_18; +lean_dec(x_16); +lean_dec(x_15); lean_dec(x_14); -x_18 = lean_box(0); -x_19 = l_Lean_Meta_Grind_internalize___lambda__2(x_1, x_2, x_18, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_17); -return 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_3); +lean_dec(x_1); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_7); +lean_ctor_set(x_18, 1, x_17); +return x_18; } else { -uint8_t x_20; -x_20 = !lean_is_exclusive(x_14); -if (x_20 == 0) -{ -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_21 = lean_ctor_get(x_14, 1); -x_22 = lean_ctor_get(x_14, 0); -lean_dec(x_22); -lean_inc(x_1); -x_23 = l_Lean_MessageData_ofExpr(x_1); -x_24 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__5; -lean_ctor_set_tag(x_14, 7); -lean_ctor_set(x_14, 1, x_23); -lean_ctor_set(x_14, 0, x_24); -x_25 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_25, 0, x_14); -lean_ctor_set(x_25, 1, x_24); -x_26 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_13, x_25, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_21); +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; uint8_t x_30; +lean_dec(x_7); +x_19 = lean_ctor_get(x_6, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_6, 1); +lean_inc(x_20); +lean_dec(x_6); +x_26 = lean_st_ref_get(x_9, x_17); 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_Grind_internalize___lambda__2(x_1, x_2, x_27, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_28); +x_29 = lean_ctor_get(x_27, 10); +lean_inc(x_29); lean_dec(x_27); -return x_29; +x_30 = !lean_is_exclusive(x_19); +if (x_30 == 0) +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; +x_31 = lean_ctor_get(x_19, 0); +x_32 = lean_ctor_get(x_19, 1); +x_33 = lean_ctor_get(x_19, 2); +x_34 = lean_ctor_get(x_19, 3); +x_35 = lean_ctor_get(x_19, 4); +x_36 = lean_ctor_get(x_19, 5); +x_37 = l_Lean_Meta_Grind_EMatchTheorems_isErased(x_29, x_36); +if (x_37 == 0) +{ +lean_object* x_38; lean_object* x_39; +x_38 = lean_box(0); +lean_inc(x_3); +x_39 = l_List_filterTR_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__4(x_3, x_35, x_38); +lean_inc(x_36); +lean_inc(x_39); +lean_ctor_set(x_19, 4, x_39); +if (lean_obj_tag(x_39) == 0) +{ +lean_object* x_40; +lean_dec(x_36); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_1); +x_40 = l_Lean_Meta_Grind_activateTheorem(x_19, x_1, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_28); +if (lean_obj_tag(x_40) == 0) +{ +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_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___lambda__1___closed__1; +x_21 = x_42; +x_22 = x_41; +goto block_25; } 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; -x_30 = lean_ctor_get(x_14, 1); -lean_inc(x_30); +uint8_t x_43; +lean_dec(x_20); +lean_dec(x_16); +lean_dec(x_15); lean_dec(x_14); -lean_inc(x_1); -x_31 = l_Lean_MessageData_ofExpr(x_1); -x_32 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__5; -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 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_34, 0, x_33); -lean_ctor_set(x_34, 1, x_32); -x_35 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_13, x_34, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, 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_Grind_internalize___lambda__2(x_1, x_2, x_36, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_37); -lean_dec(x_36); -return x_38; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_3); +lean_dec(x_1); +x_43 = !lean_is_exclusive(x_40); +if (x_43 == 0) +{ +return x_40; } +else +{ +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_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; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_internalize(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: +else { -lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_12 = l_Lean_Meta_Grind_alreadyInternalized(x_1, 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) +uint8_t x_47; +x_47 = !lean_is_exclusive(x_39); +if (x_47 == 0) { -lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = lean_ctor_get(x_12, 1); -lean_inc(x_15); -lean_dec(x_12); -x_16 = lean_box(0); -x_17 = l_Lean_Meta_Grind_internalize___lambda__3(x_1, x_2, x_16, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_15); -return x_17; +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_39, 1); +lean_dec(x_48); +x_49 = lean_ctor_get(x_39, 0); +lean_dec(x_49); +x_50 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__2; +x_51 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_50, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_28); +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_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; +lean_free_object(x_39); +lean_dec(x_36); +x_54 = lean_ctor_get(x_51, 1); +lean_inc(x_54); +lean_dec(x_51); +x_55 = lean_box(0); +x_56 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___lambda__1(x_19, x_55, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, 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_21 = x_57; +x_22 = x_58; +goto block_25; } else { -uint8_t x_18; +uint8_t x_59; +x_59 = !lean_is_exclusive(x_51); +if (x_59 == 0) +{ +lean_object* x_60; lean_object* x_61; lean_object* x_62; +x_60 = lean_ctor_get(x_51, 1); +x_61 = lean_ctor_get(x_51, 0); +lean_dec(x_61); +x_62 = l_Lean_Meta_Grind_updateLastTag(x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_60); +if (lean_obj_tag(x_62) == 0) +{ +lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; +x_63 = lean_ctor_get(x_62, 1); +lean_inc(x_63); +lean_dec(x_62); +x_64 = l_Lean_Meta_Grind_Origin_key(x_36); +lean_dec(x_36); +x_65 = l_Lean_MessageData_ofName(x_64); +x_66 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__4; +lean_ctor_set_tag(x_51, 7); +lean_ctor_set(x_51, 1, x_65); +lean_ctor_set(x_51, 0, x_66); +x_67 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__6; +lean_ctor_set_tag(x_39, 7); +lean_ctor_set(x_39, 1, x_67); +lean_ctor_set(x_39, 0, x_51); +x_68 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_50, x_39, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_63); +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_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___lambda__1(x_19, x_69, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_70); +lean_dec(x_69); +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_21 = x_72; +x_22 = x_73; +goto block_25; +} +else +{ +uint8_t x_74; +lean_free_object(x_51); +lean_free_object(x_39); +lean_dec(x_19); +lean_dec(x_36); +lean_dec(x_20); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -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_18 = !lean_is_exclusive(x_12); -if (x_18 == 0) +x_74 = !lean_is_exclusive(x_62); +if (x_74 == 0) { -lean_object* x_19; lean_object* x_20; -x_19 = lean_ctor_get(x_12, 0); -lean_dec(x_19); -x_20 = lean_box(0); -lean_ctor_set(x_12, 0, x_20); -return x_12; +return x_62; +} +else +{ +lean_object* x_75; lean_object* x_76; lean_object* x_77; +x_75 = lean_ctor_get(x_62, 0); +x_76 = lean_ctor_get(x_62, 1); +lean_inc(x_76); +lean_inc(x_75); +lean_dec(x_62); +x_77 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_77, 0, x_75); +lean_ctor_set(x_77, 1, x_76); +return x_77; +} +} +} +else +{ +lean_object* x_78; lean_object* x_79; +x_78 = lean_ctor_get(x_51, 1); +lean_inc(x_78); +lean_dec(x_51); +x_79 = l_Lean_Meta_Grind_updateLastTag(x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, 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; +x_80 = lean_ctor_get(x_79, 1); +lean_inc(x_80); +lean_dec(x_79); +x_81 = l_Lean_Meta_Grind_Origin_key(x_36); +lean_dec(x_36); +x_82 = l_Lean_MessageData_ofName(x_81); +x_83 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__4; +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_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__6; +lean_ctor_set_tag(x_39, 7); +lean_ctor_set(x_39, 1, x_85); +lean_ctor_set(x_39, 0, x_84); +x_86 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_50, x_39, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_80); +x_87 = lean_ctor_get(x_86, 0); +lean_inc(x_87); +x_88 = lean_ctor_get(x_86, 1); +lean_inc(x_88); +lean_dec(x_86); +x_89 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___lambda__1(x_19, x_87, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_88); +lean_dec(x_87); +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_21 = x_90; +x_22 = x_91; +goto block_25; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_12, 1); -lean_inc(x_21); +lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; +lean_free_object(x_39); +lean_dec(x_19); +lean_dec(x_36); +lean_dec(x_20); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); lean_dec(x_12); -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_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_3); +lean_dec(x_1); +x_92 = lean_ctor_get(x_79, 0); +lean_inc(x_92); +x_93 = lean_ctor_get(x_79, 1); +lean_inc(x_93); +if (lean_is_exclusive(x_79)) { + lean_ctor_release(x_79, 0); + lean_ctor_release(x_79, 1); + x_94 = x_79; +} else { + lean_dec_ref(x_79); + x_94 = lean_box(0); } +if (lean_is_scalar(x_94)) { + x_95 = lean_alloc_ctor(1, 2, 0); +} else { + x_95 = x_94; } +lean_ctor_set(x_95, 0, x_92); +lean_ctor_set(x_95, 1, x_93); +return x_95; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_eraseAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___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); -x_5 = 5; -x_6 = l_Lean_PersistentHashMap_findEntryAux___at_Lean_Meta_Grind_addCongrTable___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); -switch (lean_obj_tag(x_10)) { -case 0: +} +} +else { -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_name_eq(x_3, x_11); -lean_dec(x_11); -if (x_12 == 0) +lean_object* x_96; lean_object* x_97; lean_object* x_98; uint8_t x_99; +lean_dec(x_39); +x_96 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__2; +x_97 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_96, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_28); +x_98 = lean_ctor_get(x_97, 0); +lean_inc(x_98); +x_99 = lean_unbox(x_98); +lean_dec(x_98); +if (x_99 == 0) { -lean_dec(x_8); -lean_dec(x_4); -return x_1; +lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; +lean_dec(x_36); +x_100 = lean_ctor_get(x_97, 1); +lean_inc(x_100); +lean_dec(x_97); +x_101 = lean_box(0); +x_102 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___lambda__1(x_19, x_101, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_100); +x_103 = lean_ctor_get(x_102, 0); +lean_inc(x_103); +x_104 = lean_ctor_get(x_102, 1); +lean_inc(x_104); +lean_dec(x_102); +x_21 = x_103; +x_22 = x_104; +goto block_25; } else { -uint8_t x_13; -x_13 = !lean_is_exclusive(x_1); -if (x_13 == 0) +lean_object* x_105; lean_object* x_106; lean_object* x_107; +x_105 = lean_ctor_get(x_97, 1); +lean_inc(x_105); +if (lean_is_exclusive(x_97)) { + lean_ctor_release(x_97, 0); + lean_ctor_release(x_97, 1); + x_106 = x_97; +} else { + lean_dec_ref(x_97); + x_106 = lean_box(0); +} +x_107 = l_Lean_Meta_Grind_updateLastTag(x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_105); +if (lean_obj_tag(x_107) == 0) { -lean_object* x_14; lean_object* x_15; -x_14 = lean_ctor_get(x_1, 0); -lean_dec(x_14); -x_15 = lean_array_set(x_4, x_8, x_9); -lean_dec(x_8); -lean_ctor_set(x_1, 0, x_15); -return x_1; +lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; +x_108 = lean_ctor_get(x_107, 1); +lean_inc(x_108); +lean_dec(x_107); +x_109 = l_Lean_Meta_Grind_Origin_key(x_36); +lean_dec(x_36); +x_110 = l_Lean_MessageData_ofName(x_109); +x_111 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__4; +if (lean_is_scalar(x_106)) { + x_112 = lean_alloc_ctor(7, 2, 0); +} else { + x_112 = x_106; + lean_ctor_set_tag(x_112, 7); +} +lean_ctor_set(x_112, 0, x_111); +lean_ctor_set(x_112, 1, x_110); +x_113 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__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 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_96, x_114, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_108); +x_116 = lean_ctor_get(x_115, 0); +lean_inc(x_116); +x_117 = lean_ctor_get(x_115, 1); +lean_inc(x_117); +lean_dec(x_115); +x_118 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___lambda__1(x_19, x_116, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_117); +lean_dec(x_116); +x_119 = lean_ctor_get(x_118, 0); +lean_inc(x_119); +x_120 = lean_ctor_get(x_118, 1); +lean_inc(x_120); +lean_dec(x_118); +x_21 = x_119; +x_22 = x_120; +goto block_25; } else { -lean_object* x_16; lean_object* x_17; +lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; +lean_dec(x_106); +lean_dec(x_19); +lean_dec(x_36); +lean_dec(x_20); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_3); lean_dec(x_1); -x_16 = lean_array_set(x_4, x_8, x_9); -lean_dec(x_8); -x_17 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_17, 0, x_16); -return x_17; +x_121 = lean_ctor_get(x_107, 0); +lean_inc(x_121); +x_122 = lean_ctor_get(x_107, 1); +lean_inc(x_122); +if (lean_is_exclusive(x_107)) { + lean_ctor_release(x_107, 0); + lean_ctor_release(x_107, 1); + x_123 = x_107; +} else { + lean_dec_ref(x_107); + 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; } -case 1: +} +} +} +} +else { -uint8_t x_18; -x_18 = !lean_is_exclusive(x_1); -if (x_18 == 0) +lean_object* x_125; +lean_free_object(x_19); +lean_dec(x_36); +lean_dec(x_35); +lean_dec(x_34); +lean_dec(x_33); +lean_dec(x_32); +lean_dec(x_31); +x_125 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___lambda__1___closed__1; +x_21 = x_125; +x_22 = x_28; +goto block_25; +} +} +else { -lean_object* x_19; uint8_t x_20; -x_19 = lean_ctor_get(x_1, 0); +lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; uint8_t x_132; +x_126 = lean_ctor_get(x_19, 0); +x_127 = lean_ctor_get(x_19, 1); +x_128 = lean_ctor_get(x_19, 2); +x_129 = lean_ctor_get(x_19, 3); +x_130 = lean_ctor_get(x_19, 4); +x_131 = lean_ctor_get(x_19, 5); +lean_inc(x_131); +lean_inc(x_130); +lean_inc(x_129); +lean_inc(x_128); +lean_inc(x_127); +lean_inc(x_126); lean_dec(x_19); -x_20 = !lean_is_exclusive(x_10); -if (x_20 == 0) +x_132 = l_Lean_Meta_Grind_EMatchTheorems_isErased(x_29, x_131); +if (x_132 == 0) { -lean_object* x_21; lean_object* x_22; size_t x_23; lean_object* x_24; lean_object* x_25; -x_21 = lean_ctor_get(x_10, 0); -x_22 = lean_array_set(x_4, x_8, x_9); -x_23 = lean_usize_shift_right(x_2, x_5); -x_24 = l_Lean_PersistentHashMap_eraseAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__2(x_21, x_23, x_3); -lean_inc(x_24); -x_25 = l_Lean_PersistentHashMap_isUnaryNode___rarg(x_24); -if (lean_obj_tag(x_25) == 0) +lean_object* x_133; lean_object* x_134; lean_object* x_135; +x_133 = lean_box(0); +lean_inc(x_3); +x_134 = l_List_filterTR_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__4(x_3, x_130, x_133); +lean_inc(x_131); +lean_inc(x_134); +x_135 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_135, 0, x_126); +lean_ctor_set(x_135, 1, x_127); +lean_ctor_set(x_135, 2, x_128); +lean_ctor_set(x_135, 3, x_129); +lean_ctor_set(x_135, 4, x_134); +lean_ctor_set(x_135, 5, x_131); +if (lean_obj_tag(x_134) == 0) +{ +lean_object* x_136; +lean_dec(x_131); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_1); +x_136 = l_Lean_Meta_Grind_activateTheorem(x_135, x_1, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_28); +if (lean_obj_tag(x_136) == 0) { -lean_object* x_26; -lean_ctor_set(x_10, 0, x_24); -x_26 = lean_array_set(x_22, x_8, x_10); -lean_dec(x_8); -lean_ctor_set(x_1, 0, x_26); -return x_1; +lean_object* x_137; lean_object* x_138; +x_137 = lean_ctor_get(x_136, 1); +lean_inc(x_137); +lean_dec(x_136); +x_138 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___lambda__1___closed__1; +x_21 = x_138; +x_22 = x_137; +goto block_25; } else { -lean_object* x_27; uint8_t x_28; -lean_dec(x_24); -lean_free_object(x_10); -x_27 = lean_ctor_get(x_25, 0); -lean_inc(x_27); -lean_dec(x_25); -x_28 = !lean_is_exclusive(x_27); -if (x_28 == 0) -{ -lean_object* x_29; -x_29 = lean_array_set(x_22, x_8, x_27); -lean_dec(x_8); -lean_ctor_set(x_1, 0, x_29); -return x_1; +lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; +lean_dec(x_20); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_3); +lean_dec(x_1); +x_139 = lean_ctor_get(x_136, 0); +lean_inc(x_139); +x_140 = lean_ctor_get(x_136, 1); +lean_inc(x_140); +if (lean_is_exclusive(x_136)) { + lean_ctor_release(x_136, 0); + lean_ctor_release(x_136, 1); + x_141 = x_136; +} else { + lean_dec_ref(x_136); + x_141 = lean_box(0); } -else -{ -lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_30 = lean_ctor_get(x_27, 0); -x_31 = lean_ctor_get(x_27, 1); -lean_inc(x_31); -lean_inc(x_30); -lean_dec(x_27); -x_32 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_32, 0, x_30); -lean_ctor_set(x_32, 1, x_31); -x_33 = lean_array_set(x_22, x_8, x_32); -lean_dec(x_8); -lean_ctor_set(x_1, 0, x_33); -return x_1; +if (lean_is_scalar(x_141)) { + x_142 = lean_alloc_ctor(1, 2, 0); +} else { + x_142 = x_141; } +lean_ctor_set(x_142, 0, x_139); +lean_ctor_set(x_142, 1, x_140); +return x_142; } } else { -lean_object* x_34; lean_object* x_35; size_t x_36; lean_object* x_37; lean_object* x_38; -x_34 = lean_ctor_get(x_10, 0); -lean_inc(x_34); -lean_dec(x_10); -x_35 = lean_array_set(x_4, x_8, x_9); -x_36 = lean_usize_shift_right(x_2, x_5); -x_37 = l_Lean_PersistentHashMap_eraseAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__2(x_34, x_36, x_3); -lean_inc(x_37); -x_38 = l_Lean_PersistentHashMap_isUnaryNode___rarg(x_37); -if (lean_obj_tag(x_38) == 0) -{ -lean_object* x_39; lean_object* x_40; -x_39 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_39, 0, x_37); -x_40 = lean_array_set(x_35, x_8, x_39); -lean_dec(x_8); -lean_ctor_set(x_1, 0, x_40); -return x_1; +lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; uint8_t x_147; +if (lean_is_exclusive(x_134)) { + lean_ctor_release(x_134, 0); + lean_ctor_release(x_134, 1); + x_143 = x_134; +} else { + lean_dec_ref(x_134); + x_143 = lean_box(0); +} +x_144 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__2; +x_145 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_144, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_28); +x_146 = lean_ctor_get(x_145, 0); +lean_inc(x_146); +x_147 = lean_unbox(x_146); +lean_dec(x_146); +if (x_147 == 0) +{ +lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; +lean_dec(x_143); +lean_dec(x_131); +x_148 = lean_ctor_get(x_145, 1); +lean_inc(x_148); +lean_dec(x_145); +x_149 = lean_box(0); +x_150 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___lambda__1(x_135, x_149, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_148); +x_151 = lean_ctor_get(x_150, 0); +lean_inc(x_151); +x_152 = lean_ctor_get(x_150, 1); +lean_inc(x_152); +lean_dec(x_150); +x_21 = x_151; +x_22 = x_152; +goto block_25; } 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_dec(x_37); -x_41 = lean_ctor_get(x_38, 0); -lean_inc(x_41); -lean_dec(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_43); -if (lean_is_exclusive(x_41)) { - lean_ctor_release(x_41, 0); - lean_ctor_release(x_41, 1); - x_44 = x_41; +lean_object* x_153; lean_object* x_154; lean_object* x_155; +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_41); - x_44 = lean_box(0); + lean_dec_ref(x_145); + x_154 = lean_box(0); } -if (lean_is_scalar(x_44)) { - x_45 = lean_alloc_ctor(0, 2, 0); +x_155 = l_Lean_Meta_Grind_updateLastTag(x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_153); +if (lean_obj_tag(x_155) == 0) +{ +lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; +x_156 = lean_ctor_get(x_155, 1); +lean_inc(x_156); +lean_dec(x_155); +x_157 = l_Lean_Meta_Grind_Origin_key(x_131); +lean_dec(x_131); +x_158 = l_Lean_MessageData_ofName(x_157); +x_159 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__4; +if (lean_is_scalar(x_154)) { + x_160 = lean_alloc_ctor(7, 2, 0); } else { - x_45 = x_44; -} -lean_ctor_set(x_45, 0, x_42); -lean_ctor_set(x_45, 1, x_43); -x_46 = lean_array_set(x_35, x_8, x_45); -lean_dec(x_8); -lean_ctor_set(x_1, 0, x_46); -return x_1; -} + x_160 = x_154; + lean_ctor_set_tag(x_160, 7); +} +lean_ctor_set(x_160, 0, x_159); +lean_ctor_set(x_160, 1, x_158); +x_161 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__6; +if (lean_is_scalar(x_143)) { + x_162 = lean_alloc_ctor(7, 2, 0); +} else { + x_162 = x_143; + lean_ctor_set_tag(x_162, 7); } +lean_ctor_set(x_162, 0, x_160); +lean_ctor_set(x_162, 1, x_161); +x_163 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_144, x_162, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_156); +x_164 = lean_ctor_get(x_163, 0); +lean_inc(x_164); +x_165 = lean_ctor_get(x_163, 1); +lean_inc(x_165); +lean_dec(x_163); +x_166 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___lambda__1(x_135, x_164, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_165); +lean_dec(x_164); +x_167 = lean_ctor_get(x_166, 0); +lean_inc(x_167); +x_168 = lean_ctor_get(x_166, 1); +lean_inc(x_168); +lean_dec(x_166); +x_21 = x_167; +x_22 = x_168; +goto block_25; } else { -lean_object* x_47; lean_object* x_48; lean_object* x_49; size_t x_50; lean_object* x_51; lean_object* x_52; +lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; +lean_dec(x_154); +lean_dec(x_143); +lean_dec(x_135); +lean_dec(x_131); +lean_dec(x_20); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_3); lean_dec(x_1); -x_47 = lean_ctor_get(x_10, 0); -lean_inc(x_47); -if (lean_is_exclusive(x_10)) { - lean_ctor_release(x_10, 0); - x_48 = x_10; +x_169 = lean_ctor_get(x_155, 0); +lean_inc(x_169); +x_170 = lean_ctor_get(x_155, 1); +lean_inc(x_170); +if (lean_is_exclusive(x_155)) { + lean_ctor_release(x_155, 0); + lean_ctor_release(x_155, 1); + x_171 = x_155; } else { - lean_dec_ref(x_10); - x_48 = lean_box(0); + lean_dec_ref(x_155); + x_171 = lean_box(0); } -x_49 = lean_array_set(x_4, x_8, x_9); -x_50 = lean_usize_shift_right(x_2, x_5); -x_51 = l_Lean_PersistentHashMap_eraseAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__2(x_47, x_50, x_3); -lean_inc(x_51); -x_52 = l_Lean_PersistentHashMap_isUnaryNode___rarg(x_51); -if (lean_obj_tag(x_52) == 0) -{ -lean_object* x_53; lean_object* x_54; lean_object* x_55; -if (lean_is_scalar(x_48)) { - x_53 = lean_alloc_ctor(1, 1, 0); +if (lean_is_scalar(x_171)) { + x_172 = lean_alloc_ctor(1, 2, 0); } else { - x_53 = x_48; + x_172 = x_171; +} +lean_ctor_set(x_172, 0, x_169); +lean_ctor_set(x_172, 1, x_170); +return x_172; +} +} } -lean_ctor_set(x_53, 0, x_51); -x_54 = lean_array_set(x_49, x_8, x_53); -lean_dec(x_8); -x_55 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_55, 0, x_54); -return x_55; } else { -lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; -lean_dec(x_51); -lean_dec(x_48); -x_56 = lean_ctor_get(x_52, 0); -lean_inc(x_56); -lean_dec(x_52); -x_57 = lean_ctor_get(x_56, 0); -lean_inc(x_57); -x_58 = lean_ctor_get(x_56, 1); -lean_inc(x_58); -if (lean_is_exclusive(x_56)) { - lean_ctor_release(x_56, 0); - lean_ctor_release(x_56, 1); - x_59 = x_56; -} else { - lean_dec_ref(x_56); - x_59 = lean_box(0); +lean_object* x_173; +lean_dec(x_131); +lean_dec(x_130); +lean_dec(x_129); +lean_dec(x_128); +lean_dec(x_127); +lean_dec(x_126); +x_173 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___lambda__1___closed__1; +x_21 = x_173; +x_22 = x_28; +goto block_25; } -if (lean_is_scalar(x_59)) { - x_60 = lean_alloc_ctor(0, 2, 0); -} else { - x_60 = x_59; } -lean_ctor_set(x_60, 0, x_57); -lean_ctor_set(x_60, 1, x_58); -x_61 = lean_array_set(x_49, x_8, x_60); -lean_dec(x_8); -x_62 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_62, 0, x_61); -return x_62; +block_25: +{ +lean_object* x_23; +x_23 = lean_ctor_get(x_21, 0); +lean_inc(x_23); +lean_dec(x_21); +x_6 = x_20; +x_7 = x_23; +x_8 = lean_box(0); +x_17 = x_22; +goto _start; } } } -default: +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, 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; +x_12 = lean_st_ref_get(x_3, x_11); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_13, 10); +lean_inc(x_14); +lean_dec(x_13); +x_15 = !lean_is_exclusive(x_12); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_16 = lean_ctor_get(x_12, 1); +x_17 = lean_ctor_get(x_12, 0); +lean_dec(x_17); +x_18 = !lean_is_exclusive(x_14); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_19 = lean_ctor_get(x_14, 0); +x_20 = lean_ctor_get(x_14, 1); +x_21 = lean_ctor_get(x_14, 2); +lean_inc(x_19); +x_22 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__9(x_19, x_1); +if (lean_obj_tag(x_22) == 0) +{ +lean_object* x_23; +lean_free_object(x_14); +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_10); +lean_dec(x_9); lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); lean_dec(x_4); -return x_1; -} -} +lean_dec(x_3); +lean_dec(x_2); +x_23 = lean_box(0); +lean_ctor_set(x_12, 0, x_23); +return x_12; } else { -lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; -x_63 = lean_ctor_get(x_1, 0); -lean_inc(x_63); -x_64 = lean_ctor_get(x_1, 1); -lean_inc(x_64); -x_65 = lean_unsigned_to_nat(0u); -x_66 = l_Array_indexOfAux___at_Lean_MetavarContext_setMVarUserName___spec__3(x_63, x_3, x_65); -if (lean_obj_tag(x_66) == 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; +lean_free_object(x_12); +x_24 = lean_ctor_get(x_22, 0); +lean_inc(x_24); +lean_dec(x_22); +x_25 = l_Lean_PersistentHashMap_erase___at_Lean_Meta_Grind_EMatchTheorems_retrieve_x3f___spec__1(x_19, x_1); +lean_ctor_set(x_14, 0, x_25); +x_26 = lean_st_ref_take(x_3, x_16); +x_27 = lean_ctor_get(x_26, 0); +lean_inc(x_27); +x_28 = lean_ctor_get(x_26, 1); +lean_inc(x_28); +lean_dec(x_26); +x_29 = !lean_is_exclusive(x_27); +if (x_29 == 0) { -lean_dec(x_64); -lean_dec(x_63); -return x_1; -} -else +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_30 = lean_ctor_get(x_27, 10); +lean_dec(x_30); +lean_ctor_set(x_27, 10, x_14); +x_31 = lean_st_ref_set(x_3, x_27, x_28); +x_32 = lean_ctor_get(x_31, 1); +lean_inc(x_32); +lean_dec(x_31); +x_33 = lean_st_ref_get(x_3, 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, 4); +lean_inc(x_36); +lean_dec(x_34); +x_37 = lean_box(0); +x_38 = lean_box(0); +lean_inc(x_24); +x_39 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5(x_2, x_24, x_36, x_37, x_24, x_24, x_38, lean_box(0), x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_35); +lean_dec(x_24); +if (lean_obj_tag(x_39) == 0) { -uint8_t x_67; -x_67 = !lean_is_exclusive(x_1); -if (x_67 == 0) +uint8_t x_40; +x_40 = !lean_is_exclusive(x_39); +if (x_40 == 0) { -lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; -x_68 = lean_ctor_get(x_1, 1); -lean_dec(x_68); -x_69 = lean_ctor_get(x_1, 0); -lean_dec(x_69); -x_70 = lean_ctor_get(x_66, 0); -lean_inc(x_70); -lean_dec(x_66); -lean_inc(x_70); -x_71 = l_Array_eraseIdx___rarg(x_63, x_70, lean_box(0)); -x_72 = l_Array_eraseIdx___rarg(x_64, x_70, lean_box(0)); -lean_ctor_set(x_1, 1, x_72); -lean_ctor_set(x_1, 0, x_71); -return x_1; +lean_object* x_41; +x_41 = lean_ctor_get(x_39, 0); +lean_dec(x_41); +lean_ctor_set(x_39, 0, x_38); +return x_39; } else { -lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; -lean_dec(x_1); -x_73 = lean_ctor_get(x_66, 0); -lean_inc(x_73); -lean_dec(x_66); -lean_inc(x_73); -x_74 = l_Array_eraseIdx___rarg(x_63, x_73, lean_box(0)); -x_75 = l_Array_eraseIdx___rarg(x_64, x_73, lean_box(0)); -x_76 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_76, 0, x_74); -lean_ctor_set(x_76, 1, x_75); -return x_76; -} -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_erase___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__1(lean_object* x_1, lean_object* x_2) { -_start: -{ -uint64_t x_3; size_t x_4; lean_object* x_5; -x_3 = l_Lean_Name_hash___override(x_2); -x_4 = lean_uint64_to_usize(x_3); -x_5 = l_Lean_PersistentHashMap_eraseAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__2(x_1, x_4, x_2); -return x_5; +lean_object* x_42; lean_object* x_43; +x_42 = lean_ctor_get(x_39, 1); +lean_inc(x_42); +lean_dec(x_39); +x_43 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_43, 0, x_38); +lean_ctor_set(x_43, 1, x_42); +return x_43; } } -LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAtAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5(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; 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) +uint8_t x_44; +x_44 = !lean_is_exclusive(x_39); +if (x_44 == 0) { -uint8_t x_8; -lean_dec(x_4); -x_8 = 0; -return x_8; +return x_39; } else { -lean_object* x_9; uint8_t x_10; -x_9 = lean_array_fget(x_1, x_4); -x_10 = l___private_Lean_HeadIndex_0__Lean_beqHeadIndex____x40_Lean_HeadIndex___hyg_69_(x_5, x_9); -lean_dec(x_9); -if (x_10 == 0) +lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_45 = lean_ctor_get(x_39, 0); +x_46 = lean_ctor_get(x_39, 1); +lean_inc(x_46); +lean_inc(x_45); +lean_dec(x_39); +x_47 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_47, 0, x_45); +lean_ctor_set(x_47, 1, x_46); +return x_47; +} +} +} +else +{ +lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; uint8_t x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; 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_48 = lean_ctor_get(x_27, 0); +x_49 = lean_ctor_get(x_27, 1); +x_50 = lean_ctor_get(x_27, 2); +x_51 = lean_ctor_get(x_27, 3); +x_52 = lean_ctor_get(x_27, 4); +x_53 = lean_ctor_get(x_27, 5); +x_54 = lean_ctor_get_uint8(x_27, sizeof(void*)*20); +x_55 = lean_ctor_get(x_27, 6); +x_56 = lean_ctor_get(x_27, 7); +x_57 = lean_ctor_get(x_27, 8); +x_58 = lean_ctor_get(x_27, 9); +x_59 = lean_ctor_get(x_27, 11); +x_60 = lean_ctor_get(x_27, 12); +x_61 = lean_ctor_get(x_27, 13); +x_62 = lean_ctor_get(x_27, 14); +x_63 = lean_ctor_get(x_27, 15); +x_64 = lean_ctor_get(x_27, 16); +x_65 = lean_ctor_get(x_27, 17); +x_66 = lean_ctor_get(x_27, 18); +x_67 = lean_ctor_get(x_27, 19); +lean_inc(x_67); +lean_inc(x_66); +lean_inc(x_65); +lean_inc(x_64); +lean_inc(x_63); +lean_inc(x_62); +lean_inc(x_61); +lean_inc(x_60); +lean_inc(x_59); +lean_inc(x_58); +lean_inc(x_57); +lean_inc(x_56); +lean_inc(x_55); +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_dec(x_27); +x_68 = lean_alloc_ctor(0, 20, 1); +lean_ctor_set(x_68, 0, x_48); +lean_ctor_set(x_68, 1, x_49); +lean_ctor_set(x_68, 2, x_50); +lean_ctor_set(x_68, 3, x_51); +lean_ctor_set(x_68, 4, x_52); +lean_ctor_set(x_68, 5, x_53); +lean_ctor_set(x_68, 6, x_55); +lean_ctor_set(x_68, 7, x_56); +lean_ctor_set(x_68, 8, x_57); +lean_ctor_set(x_68, 9, x_58); +lean_ctor_set(x_68, 10, x_14); +lean_ctor_set(x_68, 11, x_59); +lean_ctor_set(x_68, 12, x_60); +lean_ctor_set(x_68, 13, x_61); +lean_ctor_set(x_68, 14, x_62); +lean_ctor_set(x_68, 15, x_63); +lean_ctor_set(x_68, 16, x_64); +lean_ctor_set(x_68, 17, x_65); +lean_ctor_set(x_68, 18, x_66); +lean_ctor_set(x_68, 19, x_67); +lean_ctor_set_uint8(x_68, sizeof(void*)*20, x_54); +x_69 = lean_st_ref_set(x_3, x_68, x_28); +x_70 = lean_ctor_get(x_69, 1); +lean_inc(x_70); +lean_dec(x_69); +x_71 = lean_st_ref_get(x_3, 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 = lean_ctor_get(x_72, 4); +lean_inc(x_74); +lean_dec(x_72); +x_75 = lean_box(0); +x_76 = lean_box(0); +lean_inc(x_24); +x_77 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5(x_2, x_24, x_74, x_75, x_24, x_24, x_76, lean_box(0), x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_73); +lean_dec(x_24); +if (lean_obj_tag(x_77) == 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; +lean_object* x_78; lean_object* x_79; lean_object* x_80; +x_78 = lean_ctor_get(x_77, 1); +lean_inc(x_78); +if (lean_is_exclusive(x_77)) { + lean_ctor_release(x_77, 0); + lean_ctor_release(x_77, 1); + x_79 = x_77; +} else { + lean_dec_ref(x_77); + x_79 = lean_box(0); +} +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_76); +lean_ctor_set(x_80, 1, x_78); +return x_80; } else { -uint8_t x_14; -lean_dec(x_4); -x_14 = 1; -return x_14; +lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; +x_81 = lean_ctor_get(x_77, 0); +lean_inc(x_81); +x_82 = lean_ctor_get(x_77, 1); +lean_inc(x_82); +if (lean_is_exclusive(x_77)) { + lean_ctor_release(x_77, 0); + lean_ctor_release(x_77, 1); + x_83 = x_77; +} else { + lean_dec_ref(x_77); + 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_EXPORT uint8_t l_Lean_PersistentHashMap_containsAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__4(lean_object* x_1, size_t x_2, lean_object* x_3) { -_start: +else { -if (lean_obj_tag(x_1) == 0) +lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; +x_85 = lean_ctor_get(x_14, 0); +x_86 = lean_ctor_get(x_14, 1); +x_87 = lean_ctor_get(x_14, 2); +lean_inc(x_87); +lean_inc(x_86); +lean_inc(x_85); +lean_dec(x_14); +lean_inc(x_85); +x_88 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__9(x_85, x_1); +if (lean_obj_tag(x_88) == 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_findEntryAux___at_Lean_Meta_Grind_addCongrTable___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_object* x_89; +lean_dec(x_87); +lean_dec(x_86); +lean_dec(x_85); +lean_dec(x_10); +lean_dec(x_9); lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); lean_dec(x_4); -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 = l___private_Lean_HeadIndex_0__Lean_beqHeadIndex____x40_Lean_HeadIndex___hyg_69_(x_3, x_11); -lean_dec(x_11); +lean_dec(x_3); +lean_dec(x_2); +x_89 = lean_box(0); +lean_ctor_set(x_12, 0, x_89); return x_12; } -case 1: +else { -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; +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; 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; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; +lean_free_object(x_12); +x_90 = lean_ctor_get(x_88, 0); +lean_inc(x_90); +lean_dec(x_88); +x_91 = l_Lean_PersistentHashMap_erase___at_Lean_Meta_Grind_EMatchTheorems_retrieve_x3f___spec__1(x_85, x_1); +x_92 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_92, 0, x_91); +lean_ctor_set(x_92, 1, x_86); +lean_ctor_set(x_92, 2, x_87); +x_93 = lean_st_ref_take(x_3, x_16); +x_94 = lean_ctor_get(x_93, 0); +lean_inc(x_94); +x_95 = lean_ctor_get(x_93, 1); +lean_inc(x_95); +lean_dec(x_93); +x_96 = lean_ctor_get(x_94, 0); +lean_inc(x_96); +x_97 = lean_ctor_get(x_94, 1); +lean_inc(x_97); +x_98 = lean_ctor_get(x_94, 2); +lean_inc(x_98); +x_99 = lean_ctor_get(x_94, 3); +lean_inc(x_99); +x_100 = lean_ctor_get(x_94, 4); +lean_inc(x_100); +x_101 = lean_ctor_get(x_94, 5); +lean_inc(x_101); +x_102 = lean_ctor_get_uint8(x_94, sizeof(void*)*20); +x_103 = lean_ctor_get(x_94, 6); +lean_inc(x_103); +x_104 = lean_ctor_get(x_94, 7); +lean_inc(x_104); +x_105 = lean_ctor_get(x_94, 8); +lean_inc(x_105); +x_106 = lean_ctor_get(x_94, 9); +lean_inc(x_106); +x_107 = lean_ctor_get(x_94, 11); +lean_inc(x_107); +x_108 = lean_ctor_get(x_94, 12); +lean_inc(x_108); +x_109 = lean_ctor_get(x_94, 13); +lean_inc(x_109); +x_110 = lean_ctor_get(x_94, 14); +lean_inc(x_110); +x_111 = lean_ctor_get(x_94, 15); +lean_inc(x_111); +x_112 = lean_ctor_get(x_94, 16); +lean_inc(x_112); +x_113 = lean_ctor_get(x_94, 17); +lean_inc(x_113); +x_114 = lean_ctor_get(x_94, 18); +lean_inc(x_114); +x_115 = lean_ctor_get(x_94, 19); +lean_inc(x_115); +if (lean_is_exclusive(x_94)) { + lean_ctor_release(x_94, 0); + lean_ctor_release(x_94, 1); + lean_ctor_release(x_94, 2); + lean_ctor_release(x_94, 3); + lean_ctor_release(x_94, 4); + lean_ctor_release(x_94, 5); + lean_ctor_release(x_94, 6); + lean_ctor_release(x_94, 7); + lean_ctor_release(x_94, 8); + lean_ctor_release(x_94, 9); + lean_ctor_release(x_94, 10); + lean_ctor_release(x_94, 11); + lean_ctor_release(x_94, 12); + lean_ctor_release(x_94, 13); + lean_ctor_release(x_94, 14); + lean_ctor_release(x_94, 15); + lean_ctor_release(x_94, 16); + lean_ctor_release(x_94, 17); + lean_ctor_release(x_94, 18); + lean_ctor_release(x_94, 19); + x_116 = x_94; +} else { + lean_dec_ref(x_94); + x_116 = lean_box(0); } -default: +if (lean_is_scalar(x_116)) { + x_117 = lean_alloc_ctor(0, 20, 1); +} else { + x_117 = x_116; +} +lean_ctor_set(x_117, 0, x_96); +lean_ctor_set(x_117, 1, x_97); +lean_ctor_set(x_117, 2, x_98); +lean_ctor_set(x_117, 3, x_99); +lean_ctor_set(x_117, 4, x_100); +lean_ctor_set(x_117, 5, x_101); +lean_ctor_set(x_117, 6, x_103); +lean_ctor_set(x_117, 7, x_104); +lean_ctor_set(x_117, 8, x_105); +lean_ctor_set(x_117, 9, x_106); +lean_ctor_set(x_117, 10, x_92); +lean_ctor_set(x_117, 11, x_107); +lean_ctor_set(x_117, 12, x_108); +lean_ctor_set(x_117, 13, x_109); +lean_ctor_set(x_117, 14, x_110); +lean_ctor_set(x_117, 15, x_111); +lean_ctor_set(x_117, 16, x_112); +lean_ctor_set(x_117, 17, x_113); +lean_ctor_set(x_117, 18, x_114); +lean_ctor_set(x_117, 19, x_115); +lean_ctor_set_uint8(x_117, sizeof(void*)*20, x_102); +x_118 = lean_st_ref_set(x_3, x_117, x_95); +x_119 = lean_ctor_get(x_118, 1); +lean_inc(x_119); +lean_dec(x_118); +x_120 = lean_st_ref_get(x_3, x_119); +x_121 = lean_ctor_get(x_120, 0); +lean_inc(x_121); +x_122 = lean_ctor_get(x_120, 1); +lean_inc(x_122); +lean_dec(x_120); +x_123 = lean_ctor_get(x_121, 4); +lean_inc(x_123); +lean_dec(x_121); +x_124 = lean_box(0); +x_125 = lean_box(0); +lean_inc(x_90); +x_126 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5(x_2, x_90, x_123, x_124, x_90, x_90, x_125, lean_box(0), x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_122); +lean_dec(x_90); +if (lean_obj_tag(x_126) == 0) { -uint8_t x_16; -x_16 = 0; -return x_16; +lean_object* x_127; lean_object* x_128; lean_object* x_129; +x_127 = lean_ctor_get(x_126, 1); +lean_inc(x_127); +if (lean_is_exclusive(x_126)) { + lean_ctor_release(x_126, 0); + lean_ctor_release(x_126, 1); + x_128 = x_126; +} else { + lean_dec_ref(x_126); + x_128 = lean_box(0); } +if (lean_is_scalar(x_128)) { + x_129 = lean_alloc_ctor(0, 2, 0); +} else { + x_129 = x_128; } +lean_ctor_set(x_129, 0, x_125); +lean_ctor_set(x_129, 1, x_127); +return x_129; } 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___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5(x_17, x_18, lean_box(0), x_19, x_3); -lean_dec(x_18); -lean_dec(x_17); -return x_20; +lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; +x_130 = lean_ctor_get(x_126, 0); +lean_inc(x_130); +x_131 = lean_ctor_get(x_126, 1); +lean_inc(x_131); +if (lean_is_exclusive(x_126)) { + lean_ctor_release(x_126, 0); + lean_ctor_release(x_126, 1); + x_132 = x_126; +} else { + lean_dec_ref(x_126); + x_132 = lean_box(0); } +if (lean_is_scalar(x_132)) { + x_133 = lean_alloc_ctor(1, 2, 0); +} else { + x_133 = x_132; } +lean_ctor_set(x_133, 0, x_130); +lean_ctor_set(x_133, 1, x_131); +return x_133; } -LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_contains___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__3(lean_object* x_1, lean_object* x_2) { -_start: -{ -uint64_t x_3; size_t x_4; uint8_t x_5; -x_3 = l_Lean_HeadIndex_hash(x_2); -x_4 = lean_uint64_to_usize(x_3); -x_5 = l_Lean_PersistentHashMap_containsAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__4(x_1, x_4, x_2); -return x_5; } } -LEAN_EXPORT lean_object* l_List_filterTR_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__6(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_dec(x_1); -x_4 = l_List_reverse___rarg(x_3); -return x_4; } else { -uint8_t x_5; -x_5 = !lean_is_exclusive(x_2); -if (x_5 == 0) -{ -lean_object* x_6; lean_object* x_7; uint8_t x_8; -x_6 = lean_ctor_get(x_2, 0); -x_7 = lean_ctor_get(x_2, 1); -lean_inc(x_1); -x_8 = l_Lean_PersistentHashMap_contains___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__3(x_1, x_6); -if (x_8 == 0) -{ -lean_ctor_set(x_2, 1, x_3); -{ -lean_object* _tmp_1 = x_7; -lean_object* _tmp_2 = x_2; -x_2 = _tmp_1; -x_3 = _tmp_2; -} -goto _start; +lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; +x_134 = lean_ctor_get(x_12, 1); +lean_inc(x_134); +lean_dec(x_12); +x_135 = lean_ctor_get(x_14, 0); +lean_inc(x_135); +x_136 = lean_ctor_get(x_14, 1); +lean_inc(x_136); +x_137 = lean_ctor_get(x_14, 2); +lean_inc(x_137); +if (lean_is_exclusive(x_14)) { + lean_ctor_release(x_14, 0); + lean_ctor_release(x_14, 1); + lean_ctor_release(x_14, 2); + x_138 = x_14; +} else { + lean_dec_ref(x_14); + x_138 = lean_box(0); } -else +lean_inc(x_135); +x_139 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__9(x_135, x_1); +if (lean_obj_tag(x_139) == 0) { -lean_free_object(x_2); +lean_object* x_140; lean_object* x_141; +lean_dec(x_138); +lean_dec(x_137); +lean_dec(x_136); +lean_dec(x_135); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); lean_dec(x_6); -x_2 = x_7; -goto _start; -} +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +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_134); +return x_141; } 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); -lean_inc(x_1); -x_13 = l_Lean_PersistentHashMap_contains___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__3(x_1, x_11); -if (x_13 == 0) -{ -lean_object* x_14; -x_14 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_14, 0, x_11); -lean_ctor_set(x_14, 1, x_3); -x_2 = x_12; -x_3 = x_14; -goto _start; +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; 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; 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_142 = lean_ctor_get(x_139, 0); +lean_inc(x_142); +lean_dec(x_139); +x_143 = l_Lean_PersistentHashMap_erase___at_Lean_Meta_Grind_EMatchTheorems_retrieve_x3f___spec__1(x_135, x_1); +if (lean_is_scalar(x_138)) { + x_144 = lean_alloc_ctor(0, 3, 0); +} else { + x_144 = x_138; +} +lean_ctor_set(x_144, 0, x_143); +lean_ctor_set(x_144, 1, x_136); +lean_ctor_set(x_144, 2, x_137); +x_145 = lean_st_ref_take(x_3, x_134); +x_146 = lean_ctor_get(x_145, 0); +lean_inc(x_146); +x_147 = lean_ctor_get(x_145, 1); +lean_inc(x_147); +lean_dec(x_145); +x_148 = lean_ctor_get(x_146, 0); +lean_inc(x_148); +x_149 = lean_ctor_get(x_146, 1); +lean_inc(x_149); +x_150 = lean_ctor_get(x_146, 2); +lean_inc(x_150); +x_151 = lean_ctor_get(x_146, 3); +lean_inc(x_151); +x_152 = lean_ctor_get(x_146, 4); +lean_inc(x_152); +x_153 = lean_ctor_get(x_146, 5); +lean_inc(x_153); +x_154 = lean_ctor_get_uint8(x_146, sizeof(void*)*20); +x_155 = lean_ctor_get(x_146, 6); +lean_inc(x_155); +x_156 = lean_ctor_get(x_146, 7); +lean_inc(x_156); +x_157 = lean_ctor_get(x_146, 8); +lean_inc(x_157); +x_158 = lean_ctor_get(x_146, 9); +lean_inc(x_158); +x_159 = lean_ctor_get(x_146, 11); +lean_inc(x_159); +x_160 = lean_ctor_get(x_146, 12); +lean_inc(x_160); +x_161 = lean_ctor_get(x_146, 13); +lean_inc(x_161); +x_162 = lean_ctor_get(x_146, 14); +lean_inc(x_162); +x_163 = lean_ctor_get(x_146, 15); +lean_inc(x_163); +x_164 = lean_ctor_get(x_146, 16); +lean_inc(x_164); +x_165 = lean_ctor_get(x_146, 17); +lean_inc(x_165); +x_166 = lean_ctor_get(x_146, 18); +lean_inc(x_166); +x_167 = lean_ctor_get(x_146, 19); +lean_inc(x_167); +if (lean_is_exclusive(x_146)) { + lean_ctor_release(x_146, 0); + lean_ctor_release(x_146, 1); + lean_ctor_release(x_146, 2); + lean_ctor_release(x_146, 3); + lean_ctor_release(x_146, 4); + lean_ctor_release(x_146, 5); + lean_ctor_release(x_146, 6); + lean_ctor_release(x_146, 7); + lean_ctor_release(x_146, 8); + lean_ctor_release(x_146, 9); + lean_ctor_release(x_146, 10); + lean_ctor_release(x_146, 11); + lean_ctor_release(x_146, 12); + lean_ctor_release(x_146, 13); + lean_ctor_release(x_146, 14); + lean_ctor_release(x_146, 15); + lean_ctor_release(x_146, 16); + lean_ctor_release(x_146, 17); + lean_ctor_release(x_146, 18); + lean_ctor_release(x_146, 19); + x_168 = x_146; +} else { + lean_dec_ref(x_146); + x_168 = lean_box(0); +} +if (lean_is_scalar(x_168)) { + x_169 = lean_alloc_ctor(0, 20, 1); +} else { + x_169 = x_168; +} +lean_ctor_set(x_169, 0, x_148); +lean_ctor_set(x_169, 1, x_149); +lean_ctor_set(x_169, 2, x_150); +lean_ctor_set(x_169, 3, x_151); +lean_ctor_set(x_169, 4, x_152); +lean_ctor_set(x_169, 5, x_153); +lean_ctor_set(x_169, 6, x_155); +lean_ctor_set(x_169, 7, x_156); +lean_ctor_set(x_169, 8, x_157); +lean_ctor_set(x_169, 9, x_158); +lean_ctor_set(x_169, 10, x_144); +lean_ctor_set(x_169, 11, x_159); +lean_ctor_set(x_169, 12, x_160); +lean_ctor_set(x_169, 13, x_161); +lean_ctor_set(x_169, 14, x_162); +lean_ctor_set(x_169, 15, x_163); +lean_ctor_set(x_169, 16, x_164); +lean_ctor_set(x_169, 17, x_165); +lean_ctor_set(x_169, 18, x_166); +lean_ctor_set(x_169, 19, x_167); +lean_ctor_set_uint8(x_169, sizeof(void*)*20, x_154); +x_170 = lean_st_ref_set(x_3, x_169, x_147); +x_171 = lean_ctor_get(x_170, 1); +lean_inc(x_171); +lean_dec(x_170); +x_172 = lean_st_ref_get(x_3, 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 = lean_ctor_get(x_173, 4); +lean_inc(x_175); +lean_dec(x_173); +x_176 = lean_box(0); +x_177 = lean_box(0); +lean_inc(x_142); +x_178 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5(x_2, x_142, x_175, x_176, x_142, x_142, x_177, lean_box(0), x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_174); +lean_dec(x_142); +if (lean_obj_tag(x_178) == 0) +{ +lean_object* x_179; lean_object* x_180; lean_object* x_181; +x_179 = lean_ctor_get(x_178, 1); +lean_inc(x_179); +if (lean_is_exclusive(x_178)) { + lean_ctor_release(x_178, 0); + lean_ctor_release(x_178, 1); + x_180 = x_178; +} else { + lean_dec_ref(x_178); + x_180 = lean_box(0); +} +if (lean_is_scalar(x_180)) { + x_181 = lean_alloc_ctor(0, 2, 0); +} else { + x_181 = x_180; +} +lean_ctor_set(x_181, 0, x_177); +lean_ctor_set(x_181, 1, x_179); +return x_181; } else { -lean_dec(x_11); -x_2 = x_12; -goto _start; +lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; +x_182 = lean_ctor_get(x_178, 0); +lean_inc(x_182); +x_183 = lean_ctor_get(x_178, 1); +lean_inc(x_183); +if (lean_is_exclusive(x_178)) { + lean_ctor_release(x_178, 0); + lean_ctor_release(x_178, 1); + x_184 = x_178; +} else { + lean_dec_ref(x_178); + x_184 = lean_box(0); +} +if (lean_is_scalar(x_184)) { + x_185 = lean_alloc_ctor(1, 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; } } } } } -LEAN_EXPORT lean_object* l_List_mapM_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_Meta_Grind_activateTheorem___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: { if (lean_obj_tag(x_2) == 0) @@ -9160,144 +14039,20 @@ if (lean_is_exclusive(x_28)) { lean_dec_ref(x_28); x_35 = lean_box(0); } -if (lean_is_scalar(x_35)) { - x_36 = lean_alloc_ctor(1, 2, 0); -} else { - x_36 = x_35; -} -lean_ctor_set(x_36, 0, x_33); -lean_ctor_set(x_36, 1, x_34); -return x_36; -} -} -} -} -} -static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___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_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___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, 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; -x_12 = lean_st_ref_take(x_3, x_11); -x_13 = lean_ctor_get(x_12, 0); -lean_inc(x_13); -x_14 = lean_ctor_get(x_12, 1); -lean_inc(x_14); -lean_dec(x_12); -x_15 = !lean_is_exclusive(x_13); -if (x_15 == 0) -{ -lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; -x_16 = lean_ctor_get(x_13, 9); -x_17 = l_Lean_PersistentArray_push___rarg(x_16, x_1); -lean_ctor_set(x_13, 9, x_17); -x_18 = lean_st_ref_set(x_3, x_13, x_14); -x_19 = !lean_is_exclusive(x_18); -if (x_19 == 0) -{ -lean_object* x_20; lean_object* x_21; -x_20 = lean_ctor_get(x_18, 0); -lean_dec(x_20); -x_21 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__1___closed__1; -lean_ctor_set(x_18, 0, x_21); -return x_18; -} -else -{ -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_18, 1); -lean_inc(x_22); -lean_dec(x_18); -x_23 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__1___closed__1; -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_23); -lean_ctor_set(x_24, 1, x_22); -return x_24; -} -} -else -{ -lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; 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_25 = lean_ctor_get(x_13, 0); -x_26 = lean_ctor_get(x_13, 1); -x_27 = lean_ctor_get(x_13, 2); -x_28 = lean_ctor_get(x_13, 3); -x_29 = lean_ctor_get(x_13, 4); -x_30 = lean_ctor_get(x_13, 5); -x_31 = lean_ctor_get_uint8(x_13, sizeof(void*)*14); -x_32 = lean_ctor_get(x_13, 6); -x_33 = lean_ctor_get(x_13, 7); -x_34 = lean_ctor_get(x_13, 8); -x_35 = lean_ctor_get(x_13, 9); -x_36 = lean_ctor_get(x_13, 10); -x_37 = lean_ctor_get(x_13, 11); -x_38 = lean_ctor_get(x_13, 12); -x_39 = lean_ctor_get(x_13, 13); -lean_inc(x_39); -lean_inc(x_38); -lean_inc(x_37); -lean_inc(x_36); -lean_inc(x_35); -lean_inc(x_34); -lean_inc(x_33); -lean_inc(x_32); -lean_inc(x_30); -lean_inc(x_29); -lean_inc(x_28); -lean_inc(x_27); -lean_inc(x_26); -lean_inc(x_25); -lean_dec(x_13); -x_40 = l_Lean_PersistentArray_push___rarg(x_35, x_1); -x_41 = lean_alloc_ctor(0, 14, 1); -lean_ctor_set(x_41, 0, x_25); -lean_ctor_set(x_41, 1, x_26); -lean_ctor_set(x_41, 2, x_27); -lean_ctor_set(x_41, 3, x_28); -lean_ctor_set(x_41, 4, x_29); -lean_ctor_set(x_41, 5, x_30); -lean_ctor_set(x_41, 6, x_32); -lean_ctor_set(x_41, 7, x_33); -lean_ctor_set(x_41, 8, x_34); -lean_ctor_set(x_41, 9, x_40); -lean_ctor_set(x_41, 10, x_36); -lean_ctor_set(x_41, 11, x_37); -lean_ctor_set(x_41, 12, x_38); -lean_ctor_set(x_41, 13, x_39); -lean_ctor_set_uint8(x_41, sizeof(void*)*14, x_31); -x_42 = lean_st_ref_set(x_3, x_41, x_14); -x_43 = lean_ctor_get(x_42, 1); -lean_inc(x_43); -if (lean_is_exclusive(x_42)) { - lean_ctor_release(x_42, 0); - lean_ctor_release(x_42, 1); - x_44 = x_42; -} else { - lean_dec_ref(x_42); - x_44 = lean_box(0); -} -x_45 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__1___closed__1; -if (lean_is_scalar(x_44)) { - x_46 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_35)) { + x_36 = lean_alloc_ctor(1, 2, 0); } else { - x_46 = x_44; + x_36 = x_35; } -lean_ctor_set(x_46, 0, x_45); -lean_ctor_set(x_46, 1, x_43); -return x_46; +lean_ctor_set(x_36, 0, x_33); +lean_ctor_set(x_36, 1, x_34); +return x_36; } } } -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___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) { +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_activateTheorem___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; uint8_t x_15; @@ -9311,9 +14066,9 @@ x_15 = !lean_is_exclusive(x_13); if (x_15 == 0) { lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; -x_16 = lean_ctor_get(x_13, 10); -x_17 = l_Lean_Meta_Grind_EMatchTheorems_insert(x_16, x_1); -lean_ctor_set(x_13, 10, x_17); +x_16 = lean_ctor_get(x_13, 9); +x_17 = l_Lean_PersistentArray_push___rarg(x_16, x_1); +lean_ctor_set(x_13, 9, x_17); x_18 = lean_st_ref_set(x_3, x_13, x_14); x_19 = !lean_is_exclusive(x_18); if (x_19 == 0) @@ -9321,7 +14076,7 @@ if (x_19 == 0) lean_object* x_20; lean_object* x_21; x_20 = lean_ctor_get(x_18, 0); lean_dec(x_20); -x_21 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__1___closed__1; +x_21 = lean_box(0); lean_ctor_set(x_18, 0, x_21); return x_18; } @@ -9331,7 +14086,7 @@ 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 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__1___closed__1; +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); @@ -9340,14 +14095,14 @@ return x_24; } else { -lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; 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_25 = lean_ctor_get(x_13, 0); x_26 = lean_ctor_get(x_13, 1); x_27 = lean_ctor_get(x_13, 2); x_28 = lean_ctor_get(x_13, 3); x_29 = lean_ctor_get(x_13, 4); x_30 = lean_ctor_get(x_13, 5); -x_31 = lean_ctor_get_uint8(x_13, sizeof(void*)*14); +x_31 = lean_ctor_get_uint8(x_13, sizeof(void*)*20); x_32 = lean_ctor_get(x_13, 6); x_33 = lean_ctor_get(x_13, 7); x_34 = lean_ctor_get(x_13, 8); @@ -9356,6 +14111,18 @@ x_36 = lean_ctor_get(x_13, 10); x_37 = lean_ctor_get(x_13, 11); x_38 = lean_ctor_get(x_13, 12); x_39 = lean_ctor_get(x_13, 13); +x_40 = lean_ctor_get(x_13, 14); +x_41 = lean_ctor_get(x_13, 15); +x_42 = lean_ctor_get(x_13, 16); +x_43 = lean_ctor_get(x_13, 17); +x_44 = lean_ctor_get(x_13, 18); +x_45 = lean_ctor_get(x_13, 19); +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); @@ -9371,65 +14138,53 @@ lean_inc(x_27); lean_inc(x_26); lean_inc(x_25); lean_dec(x_13); -x_40 = l_Lean_Meta_Grind_EMatchTheorems_insert(x_36, x_1); -x_41 = lean_alloc_ctor(0, 14, 1); -lean_ctor_set(x_41, 0, x_25); -lean_ctor_set(x_41, 1, x_26); -lean_ctor_set(x_41, 2, x_27); -lean_ctor_set(x_41, 3, x_28); -lean_ctor_set(x_41, 4, x_29); -lean_ctor_set(x_41, 5, x_30); -lean_ctor_set(x_41, 6, x_32); -lean_ctor_set(x_41, 7, x_33); -lean_ctor_set(x_41, 8, x_34); -lean_ctor_set(x_41, 9, x_35); -lean_ctor_set(x_41, 10, x_40); -lean_ctor_set(x_41, 11, x_37); -lean_ctor_set(x_41, 12, x_38); -lean_ctor_set(x_41, 13, x_39); -lean_ctor_set_uint8(x_41, sizeof(void*)*14, x_31); -x_42 = lean_st_ref_set(x_3, x_41, x_14); -x_43 = lean_ctor_get(x_42, 1); -lean_inc(x_43); -if (lean_is_exclusive(x_42)) { - lean_ctor_release(x_42, 0); - lean_ctor_release(x_42, 1); - x_44 = x_42; +x_46 = l_Lean_PersistentArray_push___rarg(x_35, x_1); +x_47 = lean_alloc_ctor(0, 20, 1); +lean_ctor_set(x_47, 0, x_25); +lean_ctor_set(x_47, 1, x_26); +lean_ctor_set(x_47, 2, x_27); +lean_ctor_set(x_47, 3, x_28); +lean_ctor_set(x_47, 4, x_29); +lean_ctor_set(x_47, 5, x_30); +lean_ctor_set(x_47, 6, x_32); +lean_ctor_set(x_47, 7, x_33); +lean_ctor_set(x_47, 8, x_34); +lean_ctor_set(x_47, 9, x_46); +lean_ctor_set(x_47, 10, x_36); +lean_ctor_set(x_47, 11, x_37); +lean_ctor_set(x_47, 12, x_38); +lean_ctor_set(x_47, 13, x_39); +lean_ctor_set(x_47, 14, x_40); +lean_ctor_set(x_47, 15, x_41); +lean_ctor_set(x_47, 16, x_42); +lean_ctor_set(x_47, 17, x_43); +lean_ctor_set(x_47, 18, x_44); +lean_ctor_set(x_47, 19, x_45); +lean_ctor_set_uint8(x_47, sizeof(void*)*20, x_31); +x_48 = lean_st_ref_set(x_3, x_47, x_14); +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_42); - x_44 = lean_box(0); + lean_dec_ref(x_48); + x_50 = lean_box(0); } -x_45 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__1___closed__1; -if (lean_is_scalar(x_44)) { - x_46 = lean_alloc_ctor(0, 2, 0); +x_51 = lean_box(0); +if (lean_is_scalar(x_50)) { + x_52 = lean_alloc_ctor(0, 2, 0); } else { - x_46 = x_44; -} -lean_ctor_set(x_46, 0, x_45); -lean_ctor_set(x_46, 1, x_43); -return x_46; -} -} -} -static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("ematch", 6, 6); -return x_1; + x_52 = x_50; } +lean_ctor_set(x_52, 0, x_51); +lean_ctor_set(x_52, 1, x_49); +return x_52; } -static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__1; -x_2 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__1; -x_3 = l_Lean_Name_mkStr2(x_1, x_2); -return x_3; } } -static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__3() { +static lean_object* _init_l_Lean_Meta_Grind_activateTheorem___closed__1() { _start: { lean_object* x_1; @@ -9437,16 +14192,16 @@ x_1 = lean_mk_string_unchecked("activated `", 11, 11); return x_1; } } -static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__4() { +static lean_object* _init_l_Lean_Meta_Grind_activateTheorem___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__3; +x_1 = l_Lean_Meta_Grind_activateTheorem___closed__1; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__5() { +static lean_object* _init_l_Lean_Meta_Grind_activateTheorem___closed__3() { _start: { lean_object* x_1; @@ -9454,672 +14209,816 @@ x_1 = lean_mk_string_unchecked("`, ", 3, 3); return x_1; } } -static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__6() { +static lean_object* _init_l_Lean_Meta_Grind_activateTheorem___closed__4() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__5; +x_1 = l_Lean_Meta_Grind_activateTheorem___closed__3; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__7() { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_activateTheorem(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("reinsert `", 10, 10); -return x_1; +uint8_t x_12; +x_12 = !lean_is_exclusive(x_1); +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; uint8_t x_20; +x_13 = lean_ctor_get(x_1, 0); +x_14 = lean_ctor_get(x_1, 1); +x_15 = lean_ctor_get(x_1, 2); +x_16 = lean_ctor_get(x_1, 3); +x_17 = lean_ctor_get(x_1, 4); +x_18 = lean_ctor_get(x_1, 5); +x_19 = l_Lean_Meta_Grind_shareCommon(x_14, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_20 = !lean_is_exclusive(x_19); +if (x_20 == 0) +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_21 = lean_ctor_get(x_19, 0); +x_22 = lean_ctor_get(x_19, 1); +x_23 = lean_box(0); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_24 = l_List_mapM_loop___at_Lean_Meta_Grind_activateTheorem___spec__1(x_2, x_16, x_23, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_22); +if (lean_obj_tag(x_24) == 0) +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; uint8_t x_30; +x_25 = lean_ctor_get(x_24, 0); +lean_inc(x_25); +x_26 = lean_ctor_get(x_24, 1); +lean_inc(x_26); +lean_dec(x_24); +lean_inc(x_18); +lean_inc(x_25); +lean_ctor_set(x_1, 3, x_25); +lean_ctor_set(x_1, 1, x_21); +x_27 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__2; +x_28 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_27, x_3, x_4, 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_31; lean_object* x_32; lean_object* x_33; +lean_dec(x_25); +lean_free_object(x_19); +lean_dec(x_18); +x_31 = lean_ctor_get(x_28, 1); +lean_inc(x_31); +lean_dec(x_28); +x_32 = lean_box(0); +x_33 = l_Lean_Meta_Grind_activateTheorem___lambda__1(x_1, x_32, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_31); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_33; } +else +{ +uint8_t x_34; +x_34 = !lean_is_exclusive(x_28); +if (x_34 == 0) +{ +lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_35 = lean_ctor_get(x_28, 1); +x_36 = lean_ctor_get(x_28, 0); +lean_dec(x_36); +x_37 = l_Lean_Meta_Grind_updateLastTag(x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_35); +if (lean_obj_tag(x_37) == 0) +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; +x_38 = lean_ctor_get(x_37, 1); +lean_inc(x_38); +lean_dec(x_37); +x_39 = l_Lean_Meta_Grind_Origin_key(x_18); +lean_dec(x_18); +x_40 = l_Lean_MessageData_ofName(x_39); +x_41 = l_Lean_Meta_Grind_activateTheorem___closed__2; +lean_ctor_set_tag(x_28, 7); +lean_ctor_set(x_28, 1, x_40); +lean_ctor_set(x_28, 0, x_41); +x_42 = l_Lean_Meta_Grind_activateTheorem___closed__4; +lean_ctor_set_tag(x_19, 7); +lean_ctor_set(x_19, 1, x_42); +lean_ctor_set(x_19, 0, x_28); +x_43 = l_List_mapTR_loop___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__1(x_25, x_23); +x_44 = l_List_mapTR_loop___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__2(x_43, x_23); +x_45 = l_Lean_MessageData_ofList(x_44); +x_46 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_46, 0, x_19); +lean_ctor_set(x_46, 1, x_45); +x_47 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__6; +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_Grind_updateLastTag___spec__2(x_27, x_48, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_38); +x_50 = lean_ctor_get(x_49, 0); +lean_inc(x_50); +x_51 = lean_ctor_get(x_49, 1); +lean_inc(x_51); +lean_dec(x_49); +x_52 = l_Lean_Meta_Grind_activateTheorem___lambda__1(x_1, x_50, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_51); +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_50); +return x_52; } -static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__8() { -_start: +else { -lean_object* x_1; lean_object* x_2; -x_1 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__7; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; +uint8_t x_53; +lean_free_object(x_28); +lean_dec(x_1); +lean_dec(x_25); +lean_free_object(x_19); +lean_dec(x_18); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_53 = !lean_is_exclusive(x_37); +if (x_53 == 0) +{ +return x_37; +} +else +{ +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_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; } } -static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__9() { -_start: +} +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("`", 1, 1); -return x_1; +lean_object* x_57; lean_object* x_58; +x_57 = lean_ctor_get(x_28, 1); +lean_inc(x_57); +lean_dec(x_28); +x_58 = l_Lean_Meta_Grind_updateLastTag(x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, 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; 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_59 = lean_ctor_get(x_58, 1); +lean_inc(x_59); +lean_dec(x_58); +x_60 = l_Lean_Meta_Grind_Origin_key(x_18); +lean_dec(x_18); +x_61 = l_Lean_MessageData_ofName(x_60); +x_62 = l_Lean_Meta_Grind_activateTheorem___closed__2; +x_63 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_63, 0, x_62); +lean_ctor_set(x_63, 1, x_61); +x_64 = l_Lean_Meta_Grind_activateTheorem___closed__4; +lean_ctor_set_tag(x_19, 7); +lean_ctor_set(x_19, 1, x_64); +lean_ctor_set(x_19, 0, x_63); +x_65 = l_List_mapTR_loop___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__1(x_25, x_23); +x_66 = l_List_mapTR_loop___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__2(x_65, x_23); +x_67 = l_Lean_MessageData_ofList(x_66); +x_68 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_68, 0, x_19); +lean_ctor_set(x_68, 1, x_67); +x_69 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__6; +x_70 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_70, 0, x_68); +lean_ctor_set(x_70, 1, x_69); +x_71 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_27, x_70, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_59); +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_Meta_Grind_activateTheorem___lambda__1(x_1, x_72, x_3, x_4, x_5, x_6, 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); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(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_1); +lean_dec(x_25); +lean_free_object(x_19); +lean_dec(x_18); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_75 = lean_ctor_get(x_58, 0); +lean_inc(x_75); +x_76 = lean_ctor_get(x_58, 1); +lean_inc(x_76); +if (lean_is_exclusive(x_58)) { + lean_ctor_release(x_58, 0); + lean_ctor_release(x_58, 1); + x_77 = x_58; +} else { + lean_dec_ref(x_58); + 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_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__10() { -_start: +else +{ +uint8_t x_79; +lean_free_object(x_19); +lean_dec(x_21); +lean_free_object(x_1); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_15); +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_79 = !lean_is_exclusive(x_24); +if (x_79 == 0) +{ +return x_24; +} +else +{ +lean_object* x_80; lean_object* x_81; lean_object* x_82; +x_80 = lean_ctor_get(x_24, 0); +x_81 = lean_ctor_get(x_24, 1); +lean_inc(x_81); +lean_inc(x_80); +lean_dec(x_24); +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 +{ +lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; +x_83 = lean_ctor_get(x_19, 0); +x_84 = lean_ctor_get(x_19, 1); +lean_inc(x_84); +lean_inc(x_83); +lean_dec(x_19); +x_85 = lean_box(0); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_86 = l_List_mapM_loop___at_Lean_Meta_Grind_activateTheorem___spec__1(x_2, x_16, x_85, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_84); +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; uint8_t x_92; +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_18); +lean_inc(x_87); +lean_ctor_set(x_1, 3, x_87); +lean_ctor_set(x_1, 1, x_83); +x_89 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__2; +x_90 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_89, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_88); +x_91 = lean_ctor_get(x_90, 0); +lean_inc(x_91); +x_92 = lean_unbox(x_91); +lean_dec(x_91); +if (x_92 == 0) +{ +lean_object* x_93; lean_object* x_94; lean_object* x_95; +lean_dec(x_87); +lean_dec(x_18); +x_93 = lean_ctor_get(x_90, 1); +lean_inc(x_93); +lean_dec(x_90); +x_94 = lean_box(0); +x_95 = l_Lean_Meta_Grind_activateTheorem___lambda__1(x_1, x_94, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_93); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_95; +} +else { -lean_object* x_1; lean_object* x_2; -x_1 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__9; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; -} +lean_object* x_96; lean_object* x_97; lean_object* x_98; +x_96 = lean_ctor_get(x_90, 1); +lean_inc(x_96); +if (lean_is_exclusive(x_90)) { + lean_ctor_release(x_90, 0); + lean_ctor_release(x_90, 1); + x_97 = x_90; +} else { + lean_dec_ref(x_90); + x_97 = lean_box(0); } -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17) { -_start: -{ -if (lean_obj_tag(x_6) == 0) +x_98 = l_Lean_Meta_Grind_updateLastTag(x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_96); +if (lean_obj_tag(x_98) == 0) { -lean_object* x_18; -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_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; +x_99 = lean_ctor_get(x_98, 1); +lean_inc(x_99); +lean_dec(x_98); +x_100 = l_Lean_Meta_Grind_Origin_key(x_18); +lean_dec(x_18); +x_101 = l_Lean_MessageData_ofName(x_100); +x_102 = l_Lean_Meta_Grind_activateTheorem___closed__2; +if (lean_is_scalar(x_97)) { + x_103 = lean_alloc_ctor(7, 2, 0); +} else { + x_103 = x_97; + lean_ctor_set_tag(x_103, 7); +} +lean_ctor_set(x_103, 0, x_102); +lean_ctor_set(x_103, 1, x_101); +x_104 = l_Lean_Meta_Grind_activateTheorem___closed__4; +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_List_mapTR_loop___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__1(x_87, x_85); +x_107 = l_List_mapTR_loop___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__2(x_106, x_85); +x_108 = l_Lean_MessageData_ofList(x_107); +x_109 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_109, 0, x_105); +lean_ctor_set(x_109, 1, x_108); +x_110 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__6; +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 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_89, x_111, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_99); +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 = l_Lean_Meta_Grind_activateTheorem___lambda__1(x_1, x_113, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_114); 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_18 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_18, 0, x_7); -lean_ctor_set(x_18, 1, x_17); -return x_18; +lean_dec(x_113); +return x_115; } else { -uint8_t x_19; +lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; +lean_dec(x_97); +lean_dec(x_1); +lean_dec(x_87); +lean_dec(x_18); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_7); -x_19 = !lean_is_exclusive(x_6); -if (x_19 == 0) -{ -lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_27; -x_20 = lean_ctor_get(x_6, 0); -x_21 = lean_ctor_get(x_6, 1); -x_27 = !lean_is_exclusive(x_20); -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; -x_28 = lean_ctor_get(x_20, 0); -x_29 = lean_ctor_get(x_20, 1); -x_30 = lean_ctor_get(x_20, 2); -x_31 = lean_ctor_get(x_20, 3); -x_32 = lean_ctor_get(x_20, 4); -x_33 = lean_ctor_get(x_20, 5); -x_34 = lean_box(0); -lean_inc(x_3); -x_35 = l_List_filterTR_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__6(x_3, x_32, x_34); -if (lean_obj_tag(x_35) == 0) -{ -lean_object* x_36; uint8_t x_37; -x_36 = l_Lean_Meta_Grind_shareCommon(x_29, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); -x_37 = !lean_is_exclusive(x_36); -if (x_37 == 0) -{ -lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_38 = lean_ctor_get(x_36, 0); -x_39 = lean_ctor_get(x_36, 1); -lean_inc(x_16); -lean_inc(x_15); -lean_inc(x_14); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_1); -x_40 = l_List_mapM_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__7(x_1, x_31, x_34, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_39); -if (lean_obj_tag(x_40) == 0) -{ -lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; uint8_t x_46; -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_33); -lean_inc(x_41); -lean_ctor_set(x_20, 4, x_35); -lean_ctor_set(x_20, 3, x_41); -lean_ctor_set(x_20, 1, x_38); -x_43 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__2; -x_44 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_43, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_42); -x_45 = lean_ctor_get(x_44, 0); -lean_inc(x_45); -x_46 = lean_unbox(x_45); -lean_dec(x_45); -if (x_46 == 0) -{ -lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; -lean_dec(x_41); -lean_free_object(x_36); -lean_dec(x_33); -lean_free_object(x_6); -x_47 = lean_ctor_get(x_44, 1); -lean_inc(x_47); -lean_dec(x_44); -x_48 = lean_box(0); -x_49 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__1(x_20, x_48, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_47); -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_22 = x_50; -x_23 = x_51; -goto block_26; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_116 = lean_ctor_get(x_98, 0); +lean_inc(x_116); +x_117 = lean_ctor_get(x_98, 1); +lean_inc(x_117); +if (lean_is_exclusive(x_98)) { + lean_ctor_release(x_98, 0); + lean_ctor_release(x_98, 1); + x_118 = x_98; +} else { + lean_dec_ref(x_98); + x_118 = lean_box(0); } -else -{ -uint8_t x_52; -x_52 = !lean_is_exclusive(x_44); -if (x_52 == 0) -{ -lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; -x_53 = lean_ctor_get(x_44, 1); -x_54 = lean_ctor_get(x_44, 0); -lean_dec(x_54); -x_55 = l_Lean_Meta_Grind_Origin_key(x_33); -lean_dec(x_33); -x_56 = l_Lean_MessageData_ofName(x_55); -x_57 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__4; -lean_ctor_set_tag(x_44, 7); -lean_ctor_set(x_44, 1, x_56); -lean_ctor_set(x_44, 0, x_57); -x_58 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__6; -lean_ctor_set_tag(x_36, 7); -lean_ctor_set(x_36, 1, x_58); -lean_ctor_set(x_36, 0, x_44); -x_59 = l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__3(x_41, x_34); -x_60 = l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__4(x_59, x_34); -x_61 = l_Lean_MessageData_ofList(x_60); -lean_ctor_set_tag(x_6, 7); -lean_ctor_set(x_6, 1, x_61); -lean_ctor_set(x_6, 0, x_36); -x_62 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__5; -x_63 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_63, 0, x_6); -lean_ctor_set(x_63, 1, x_62); -x_64 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_43, x_63, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_53); -x_65 = lean_ctor_get(x_64, 0); -lean_inc(x_65); -x_66 = lean_ctor_get(x_64, 1); -lean_inc(x_66); -lean_dec(x_64); -x_67 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__1(x_20, x_65, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_66); -lean_dec(x_65); -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_22 = x_68; -x_23 = x_69; -goto block_26; +if (lean_is_scalar(x_118)) { + x_119 = lean_alloc_ctor(1, 2, 0); +} else { + x_119 = x_118; } -else -{ -lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; -x_70 = lean_ctor_get(x_44, 1); -lean_inc(x_70); -lean_dec(x_44); -x_71 = l_Lean_Meta_Grind_Origin_key(x_33); -lean_dec(x_33); -x_72 = l_Lean_MessageData_ofName(x_71); -x_73 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__4; -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_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__6; -lean_ctor_set_tag(x_36, 7); -lean_ctor_set(x_36, 1, x_75); -lean_ctor_set(x_36, 0, x_74); -x_76 = l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__3(x_41, x_34); -x_77 = l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__4(x_76, x_34); -x_78 = l_Lean_MessageData_ofList(x_77); -lean_ctor_set_tag(x_6, 7); -lean_ctor_set(x_6, 1, x_78); -lean_ctor_set(x_6, 0, x_36); -x_79 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__5; -x_80 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_80, 0, x_6); -lean_ctor_set(x_80, 1, x_79); -x_81 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_43, x_80, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_70); -x_82 = lean_ctor_get(x_81, 0); -lean_inc(x_82); -x_83 = lean_ctor_get(x_81, 1); -lean_inc(x_83); -lean_dec(x_81); -x_84 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__1(x_20, x_82, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_83); -lean_dec(x_82); -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_26; +lean_ctor_set(x_119, 0, x_116); +lean_ctor_set(x_119, 1, x_117); +return x_119; } } } else { -uint8_t x_87; -lean_free_object(x_36); -lean_dec(x_38); -lean_free_object(x_20); -lean_dec(x_33); -lean_dec(x_30); -lean_dec(x_28); -lean_free_object(x_6); -lean_dec(x_21); -lean_dec(x_16); +lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; +lean_dec(x_83); +lean_free_object(x_1); +lean_dec(x_18); +lean_dec(x_17); lean_dec(x_15); -lean_dec(x_14); lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); lean_dec(x_3); -lean_dec(x_1); -x_87 = !lean_is_exclusive(x_40); -if (x_87 == 0) -{ -return x_40; +x_120 = lean_ctor_get(x_86, 0); +lean_inc(x_120); +x_121 = lean_ctor_get(x_86, 1); +lean_inc(x_121); +if (lean_is_exclusive(x_86)) { + lean_ctor_release(x_86, 0); + lean_ctor_release(x_86, 1); + x_122 = x_86; +} else { + lean_dec_ref(x_86); + x_122 = lean_box(0); } -else -{ -lean_object* x_88; lean_object* x_89; lean_object* x_90; -x_88 = lean_ctor_get(x_40, 0); -x_89 = lean_ctor_get(x_40, 1); -lean_inc(x_89); -lean_inc(x_88); -lean_dec(x_40); -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; +if (lean_is_scalar(x_122)) { + x_123 = lean_alloc_ctor(1, 2, 0); +} else { + x_123 = 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_91; lean_object* x_92; lean_object* x_93; -x_91 = lean_ctor_get(x_36, 0); -x_92 = lean_ctor_get(x_36, 1); -lean_inc(x_92); -lean_inc(x_91); -lean_dec(x_36); -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_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; +x_124 = lean_ctor_get(x_1, 0); +x_125 = lean_ctor_get(x_1, 1); +x_126 = lean_ctor_get(x_1, 2); +x_127 = lean_ctor_get(x_1, 3); +x_128 = lean_ctor_get(x_1, 4); +x_129 = lean_ctor_get(x_1, 5); +lean_inc(x_129); +lean_inc(x_128); +lean_inc(x_127); +lean_inc(x_126); +lean_inc(x_125); +lean_inc(x_124); +lean_dec(x_1); +x_130 = l_Lean_Meta_Grind_shareCommon(x_125, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_131 = lean_ctor_get(x_130, 0); +lean_inc(x_131); +x_132 = lean_ctor_get(x_130, 1); +lean_inc(x_132); +if (lean_is_exclusive(x_130)) { + lean_ctor_release(x_130, 0); + lean_ctor_release(x_130, 1); + x_133 = x_130; +} else { + lean_dec_ref(x_130); + x_133 = lean_box(0); +} +x_134 = lean_box(0); lean_inc(x_10); lean_inc(x_9); -lean_inc(x_1); -x_93 = l_List_mapM_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__7(x_1, x_31, x_34, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_92); -if (lean_obj_tag(x_93) == 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_135 = l_List_mapM_loop___at_Lean_Meta_Grind_activateTheorem___spec__1(x_2, x_127, x_134, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_132); +if (lean_obj_tag(x_135) == 0) { -lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; uint8_t x_99; -x_94 = lean_ctor_get(x_93, 0); -lean_inc(x_94); -x_95 = lean_ctor_get(x_93, 1); -lean_inc(x_95); -lean_dec(x_93); -lean_inc(x_33); -lean_inc(x_94); -lean_ctor_set(x_20, 4, x_35); -lean_ctor_set(x_20, 3, x_94); -lean_ctor_set(x_20, 1, x_91); -x_96 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__2; -x_97 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_96, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_95); -x_98 = lean_ctor_get(x_97, 0); -lean_inc(x_98); -x_99 = lean_unbox(x_98); -lean_dec(x_98); -if (x_99 == 0) +lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; uint8_t x_142; +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); +lean_inc(x_129); +lean_inc(x_136); +x_138 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_138, 0, x_124); +lean_ctor_set(x_138, 1, x_131); +lean_ctor_set(x_138, 2, x_126); +lean_ctor_set(x_138, 3, x_136); +lean_ctor_set(x_138, 4, x_128); +lean_ctor_set(x_138, 5, x_129); +x_139 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__2; +x_140 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_139, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_137); +x_141 = lean_ctor_get(x_140, 0); +lean_inc(x_141); +x_142 = lean_unbox(x_141); +lean_dec(x_141); +if (x_142 == 0) { -lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; -lean_dec(x_94); -lean_dec(x_33); -lean_free_object(x_6); -x_100 = lean_ctor_get(x_97, 1); -lean_inc(x_100); -lean_dec(x_97); -x_101 = lean_box(0); -x_102 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__1(x_20, x_101, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_100); -x_103 = lean_ctor_get(x_102, 0); -lean_inc(x_103); -x_104 = lean_ctor_get(x_102, 1); -lean_inc(x_104); -lean_dec(x_102); -x_22 = x_103; -x_23 = x_104; -goto block_26; +lean_object* x_143; lean_object* x_144; lean_object* x_145; +lean_dec(x_136); +lean_dec(x_133); +lean_dec(x_129); +x_143 = lean_ctor_get(x_140, 1); +lean_inc(x_143); +lean_dec(x_140); +x_144 = lean_box(0); +x_145 = l_Lean_Meta_Grind_activateTheorem___lambda__1(x_138, x_144, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_143); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_145; } else { -lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; -x_105 = lean_ctor_get(x_97, 1); -lean_inc(x_105); -if (lean_is_exclusive(x_97)) { - lean_ctor_release(x_97, 0); - lean_ctor_release(x_97, 1); - x_106 = x_97; +lean_object* x_146; lean_object* x_147; lean_object* x_148; +x_146 = lean_ctor_get(x_140, 1); +lean_inc(x_146); +if (lean_is_exclusive(x_140)) { + lean_ctor_release(x_140, 0); + lean_ctor_release(x_140, 1); + x_147 = x_140; } else { - lean_dec_ref(x_97); - x_106 = lean_box(0); + lean_dec_ref(x_140); + x_147 = lean_box(0); } -x_107 = l_Lean_Meta_Grind_Origin_key(x_33); -lean_dec(x_33); -x_108 = l_Lean_MessageData_ofName(x_107); -x_109 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__4; -if (lean_is_scalar(x_106)) { - x_110 = lean_alloc_ctor(7, 2, 0); +x_148 = l_Lean_Meta_Grind_updateLastTag(x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_146); +if (lean_obj_tag(x_148) == 0) +{ +lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; +x_149 = lean_ctor_get(x_148, 1); +lean_inc(x_149); +lean_dec(x_148); +x_150 = l_Lean_Meta_Grind_Origin_key(x_129); +lean_dec(x_129); +x_151 = l_Lean_MessageData_ofName(x_150); +x_152 = l_Lean_Meta_Grind_activateTheorem___closed__2; +if (lean_is_scalar(x_147)) { + x_153 = lean_alloc_ctor(7, 2, 0); } else { - x_110 = x_106; - lean_ctor_set_tag(x_110, 7); -} -lean_ctor_set(x_110, 0, x_109); -lean_ctor_set(x_110, 1, x_108); -x_111 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__6; -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_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__3(x_94, x_34); -x_114 = l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__4(x_113, x_34); -x_115 = l_Lean_MessageData_ofList(x_114); -lean_ctor_set_tag(x_6, 7); -lean_ctor_set(x_6, 1, x_115); -lean_ctor_set(x_6, 0, x_112); -x_116 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__5; -x_117 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_117, 0, x_6); -lean_ctor_set(x_117, 1, x_116); -x_118 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_96, x_117, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_105); -x_119 = lean_ctor_get(x_118, 0); -lean_inc(x_119); -x_120 = lean_ctor_get(x_118, 1); -lean_inc(x_120); -lean_dec(x_118); -x_121 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__1(x_20, x_119, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_120); -lean_dec(x_119); -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_22 = x_122; -x_23 = x_123; -goto block_26; + x_153 = x_147; + lean_ctor_set_tag(x_153, 7); } +lean_ctor_set(x_153, 0, x_152); +lean_ctor_set(x_153, 1, x_151); +x_154 = l_Lean_Meta_Grind_activateTheorem___closed__4; +if (lean_is_scalar(x_133)) { + x_155 = lean_alloc_ctor(7, 2, 0); +} else { + x_155 = x_133; + lean_ctor_set_tag(x_155, 7); +} +lean_ctor_set(x_155, 0, x_153); +lean_ctor_set(x_155, 1, x_154); +x_156 = l_List_mapTR_loop___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__1(x_136, x_134); +x_157 = l_List_mapTR_loop___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__2(x_156, x_134); +x_158 = l_Lean_MessageData_ofList(x_157); +x_159 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_159, 0, x_155); +lean_ctor_set(x_159, 1, x_158); +x_160 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__6; +x_161 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_161, 0, x_159); +lean_ctor_set(x_161, 1, x_160); +x_162 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_139, x_161, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_149); +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_Meta_Grind_activateTheorem___lambda__1(x_138, x_163, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_164); +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_163); +return x_165; } else { -lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; -lean_dec(x_91); -lean_free_object(x_20); -lean_dec(x_33); -lean_dec(x_30); -lean_dec(x_28); -lean_free_object(x_6); -lean_dec(x_21); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); +lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; +lean_dec(x_147); +lean_dec(x_138); +lean_dec(x_136); +lean_dec(x_133); +lean_dec(x_129); 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_124 = lean_ctor_get(x_93, 0); -lean_inc(x_124); -x_125 = lean_ctor_get(x_93, 1); -lean_inc(x_125); -if (lean_is_exclusive(x_93)) { - lean_ctor_release(x_93, 0); - lean_ctor_release(x_93, 1); - x_126 = x_93; +x_166 = lean_ctor_get(x_148, 0); +lean_inc(x_166); +x_167 = lean_ctor_get(x_148, 1); +lean_inc(x_167); +if (lean_is_exclusive(x_148)) { + lean_ctor_release(x_148, 0); + lean_ctor_release(x_148, 1); + x_168 = x_148; } else { - lean_dec_ref(x_93); - x_126 = lean_box(0); + lean_dec_ref(x_148); + x_168 = lean_box(0); } -if (lean_is_scalar(x_126)) { - x_127 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_168)) { + x_169 = lean_alloc_ctor(1, 2, 0); } else { - x_127 = x_126; + x_169 = x_168; } -lean_ctor_set(x_127, 0, x_124); -lean_ctor_set(x_127, 1, x_125); -return x_127; +lean_ctor_set(x_169, 0, x_166); +lean_ctor_set(x_169, 1, x_167); +return x_169; } } } else { -uint8_t x_128; -lean_free_object(x_6); -lean_inc(x_33); -lean_inc(x_35); -lean_ctor_set(x_20, 4, x_35); -x_128 = !lean_is_exclusive(x_35); -if (x_128 == 0) -{ -lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; uint8_t x_134; -x_129 = lean_ctor_get(x_35, 1); -lean_dec(x_129); -x_130 = lean_ctor_get(x_35, 0); -lean_dec(x_130); -x_131 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__2; -x_132 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_131, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); -x_133 = lean_ctor_get(x_132, 0); -lean_inc(x_133); -x_134 = lean_unbox(x_133); +lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_dec(x_133); -if (x_134 == 0) -{ -lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; -lean_free_object(x_35); -lean_dec(x_33); -x_135 = lean_ctor_get(x_132, 1); -lean_inc(x_135); -lean_dec(x_132); -x_136 = lean_box(0); -x_137 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__2(x_20, x_136, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_135); -x_138 = lean_ctor_get(x_137, 0); -lean_inc(x_138); -x_139 = lean_ctor_get(x_137, 1); -lean_inc(x_139); -lean_dec(x_137); -x_22 = x_138; -x_23 = x_139; -goto block_26; +lean_dec(x_131); +lean_dec(x_129); +lean_dec(x_128); +lean_dec(x_126); +lean_dec(x_124); +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_170 = lean_ctor_get(x_135, 0); +lean_inc(x_170); +x_171 = lean_ctor_get(x_135, 1); +lean_inc(x_171); +if (lean_is_exclusive(x_135)) { + lean_ctor_release(x_135, 0); + lean_ctor_release(x_135, 1); + x_172 = x_135; +} else { + lean_dec_ref(x_135); + x_172 = lean_box(0); } -else -{ -uint8_t x_140; -x_140 = !lean_is_exclusive(x_132); -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; -x_141 = lean_ctor_get(x_132, 1); -x_142 = lean_ctor_get(x_132, 0); -lean_dec(x_142); -x_143 = l_Lean_Meta_Grind_Origin_key(x_33); -lean_dec(x_33); -x_144 = l_Lean_MessageData_ofName(x_143); -x_145 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__8; -lean_ctor_set_tag(x_132, 7); -lean_ctor_set(x_132, 1, x_144); -lean_ctor_set(x_132, 0, x_145); -x_146 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__10; -lean_ctor_set_tag(x_35, 7); -lean_ctor_set(x_35, 1, x_146); -lean_ctor_set(x_35, 0, x_132); -x_147 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_131, x_35, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_141); -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_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__2(x_20, x_148, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_149); -lean_dec(x_148); -x_151 = lean_ctor_get(x_150, 0); -lean_inc(x_151); -x_152 = lean_ctor_get(x_150, 1); -lean_inc(x_152); -lean_dec(x_150); -x_22 = x_151; -x_23 = x_152; -goto block_26; +if (lean_is_scalar(x_172)) { + x_173 = lean_alloc_ctor(1, 2, 0); +} else { + x_173 = x_172; } -else -{ -lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; -x_153 = lean_ctor_get(x_132, 1); -lean_inc(x_153); -lean_dec(x_132); -x_154 = l_Lean_Meta_Grind_Origin_key(x_33); -lean_dec(x_33); -x_155 = l_Lean_MessageData_ofName(x_154); -x_156 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__8; -x_157 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_157, 0, x_156); -lean_ctor_set(x_157, 1, x_155); -x_158 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__10; -lean_ctor_set_tag(x_35, 7); -lean_ctor_set(x_35, 1, x_158); -lean_ctor_set(x_35, 0, x_157); -x_159 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_131, x_35, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_153); -x_160 = lean_ctor_get(x_159, 0); -lean_inc(x_160); -x_161 = lean_ctor_get(x_159, 1); -lean_inc(x_161); -lean_dec(x_159); -x_162 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__2(x_20, x_160, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_161); -lean_dec(x_160); -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_22 = x_163; -x_23 = x_164; -goto block_26; +lean_ctor_set(x_173, 0, x_170); +lean_ctor_set(x_173, 1, x_171); +return x_173; } } } -else +} +LEAN_EXPORT lean_object* l_Lean_Meta_isMatcher___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns___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_165; lean_object* x_166; lean_object* x_167; uint8_t x_168; -lean_dec(x_35); -x_165 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__2; -x_166 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_165, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); -x_167 = lean_ctor_get(x_166, 0); -lean_inc(x_167); -x_168 = lean_unbox(x_167); -lean_dec(x_167); -if (x_168 == 0) +lean_object* x_11; uint8_t x_12; +x_11 = lean_st_ref_get(x_9, x_10); +x_12 = !lean_is_exclusive(x_11); +if (x_12 == 0) { -lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; -lean_dec(x_33); -x_169 = lean_ctor_get(x_166, 1); -lean_inc(x_169); -lean_dec(x_166); -x_170 = lean_box(0); -x_171 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__2(x_20, x_170, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_169); -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_22 = x_172; -x_23 = x_173; -goto block_26; +lean_object* x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; +x_13 = lean_ctor_get(x_11, 0); +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +lean_dec(x_13); +x_15 = lean_is_matcher(x_14, x_1); +x_16 = lean_box(x_15); +lean_ctor_set(x_11, 0, x_16); +return x_11; } else { -lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; -x_174 = lean_ctor_get(x_166, 1); -lean_inc(x_174); -if (lean_is_exclusive(x_166)) { - lean_ctor_release(x_166, 0); - lean_ctor_release(x_166, 1); - x_175 = x_166; -} else { - lean_dec_ref(x_166); - x_175 = lean_box(0); -} -x_176 = l_Lean_Meta_Grind_Origin_key(x_33); -lean_dec(x_33); -x_177 = l_Lean_MessageData_ofName(x_176); -x_178 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__8; -if (lean_is_scalar(x_175)) { - x_179 = lean_alloc_ctor(7, 2, 0); -} else { - x_179 = x_175; - lean_ctor_set_tag(x_179, 7); -} -lean_ctor_set(x_179, 0, x_178); -lean_ctor_set(x_179, 1, x_177); -x_180 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__10; -x_181 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_181, 0, x_179); -lean_ctor_set(x_181, 1, x_180); -x_182 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_165, x_181, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, 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_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__2(x_20, x_183, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_184); -lean_dec(x_183); -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_22 = x_186; -x_23 = x_187; -goto block_26; +lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; lean_object* x_21; lean_object* x_22; +x_17 = lean_ctor_get(x_11, 0); +x_18 = lean_ctor_get(x_11, 1); +lean_inc(x_18); +lean_inc(x_17); +lean_dec(x_11); +x_19 = lean_ctor_get(x_17, 0); +lean_inc(x_19); +lean_dec(x_17); +x_20 = lean_is_matcher(x_19, x_1); +x_21 = lean_box(x_20); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_18); +return x_22; } } } +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, size_t x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +_start: +{ +uint8_t x_17; +x_17 = lean_usize_dec_lt(x_6, x_5); +if (x_17 == 0) +{ +lean_object* x_18; +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_1); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_7); +lean_ctor_set(x_18, 1, x_16); +return x_18; } 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; -x_188 = lean_ctor_get(x_20, 0); -x_189 = lean_ctor_get(x_20, 1); -x_190 = lean_ctor_get(x_20, 2); -x_191 = lean_ctor_get(x_20, 3); -x_192 = lean_ctor_get(x_20, 4); -x_193 = lean_ctor_get(x_20, 5); -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_dec(x_20); -x_194 = lean_box(0); -lean_inc(x_3); -x_195 = l_List_filterTR_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__6(x_3, x_192, x_194); -if (lean_obj_tag(x_195) == 0) +lean_object* x_19; uint8_t x_20; uint8_t x_21; lean_object* x_22; +lean_dec(x_7); +x_19 = lean_array_uget(x_4, x_6); +x_20 = 0; +x_21 = 1; +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +x_22 = l_Lean_Meta_Grind_mkEMatchEqTheorem(x_19, x_20, x_21, x_12, x_13, x_14, x_15, x_16); +if (lean_obj_tag(x_22) == 0) { -lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; -x_196 = l_Lean_Meta_Grind_shareCommon(x_189, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); -x_197 = lean_ctor_get(x_196, 0); -lean_inc(x_197); -x_198 = lean_ctor_get(x_196, 1); -lean_inc(x_198); -if (lean_is_exclusive(x_196)) { - lean_ctor_release(x_196, 0); - lean_ctor_release(x_196, 1); - x_199 = x_196; -} else { - lean_dec_ref(x_196); - x_199 = lean_box(0); -} -lean_inc(x_16); +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +x_24 = lean_ctor_get(x_22, 1); +lean_inc(x_24); +lean_dec(x_22); lean_inc(x_15); lean_inc(x_14); lean_inc(x_13); @@ -10127,125 +15026,26 @@ 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_200 = l_List_mapM_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__7(x_1, x_191, x_194, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_198); -if (lean_obj_tag(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; uint8_t x_207; -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); -lean_inc(x_193); -lean_inc(x_201); -x_203 = lean_alloc_ctor(0, 6, 0); -lean_ctor_set(x_203, 0, x_188); -lean_ctor_set(x_203, 1, x_197); -lean_ctor_set(x_203, 2, x_190); -lean_ctor_set(x_203, 3, x_201); -lean_ctor_set(x_203, 4, x_195); -lean_ctor_set(x_203, 5, x_193); -x_204 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__2; -x_205 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_204, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_202); -x_206 = lean_ctor_get(x_205, 0); -lean_inc(x_206); -x_207 = lean_unbox(x_206); -lean_dec(x_206); -if (x_207 == 0) -{ -lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; -lean_dec(x_201); -lean_dec(x_199); -lean_dec(x_193); -lean_free_object(x_6); -x_208 = lean_ctor_get(x_205, 1); -lean_inc(x_208); -lean_dec(x_205); -x_209 = lean_box(0); -x_210 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__1(x_203, x_209, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, 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_22 = x_211; -x_23 = x_212; -goto block_26; -} -else +x_25 = l_Lean_Meta_Grind_activateTheorem(x_23, x_1, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_24); +if (lean_obj_tag(x_25) == 0) { -lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; -x_213 = lean_ctor_get(x_205, 1); -lean_inc(x_213); -if (lean_is_exclusive(x_205)) { - lean_ctor_release(x_205, 0); - lean_ctor_release(x_205, 1); - x_214 = x_205; -} else { - lean_dec_ref(x_205); - x_214 = lean_box(0); -} -x_215 = l_Lean_Meta_Grind_Origin_key(x_193); -lean_dec(x_193); -x_216 = l_Lean_MessageData_ofName(x_215); -x_217 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__4; -if (lean_is_scalar(x_214)) { - x_218 = lean_alloc_ctor(7, 2, 0); -} else { - x_218 = x_214; - lean_ctor_set_tag(x_218, 7); -} -lean_ctor_set(x_218, 0, x_217); -lean_ctor_set(x_218, 1, x_216); -x_219 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__6; -if (lean_is_scalar(x_199)) { - x_220 = lean_alloc_ctor(7, 2, 0); -} else { - x_220 = x_199; - lean_ctor_set_tag(x_220, 7); -} -lean_ctor_set(x_220, 0, x_218); -lean_ctor_set(x_220, 1, x_219); -x_221 = l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__3(x_201, x_194); -x_222 = l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__4(x_221, x_194); -x_223 = l_Lean_MessageData_ofList(x_222); -lean_ctor_set_tag(x_6, 7); -lean_ctor_set(x_6, 1, x_223); -lean_ctor_set(x_6, 0, x_220); -x_224 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__5; -x_225 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_225, 0, x_6); -lean_ctor_set(x_225, 1, x_224); -x_226 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_204, x_225, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_213); -x_227 = lean_ctor_get(x_226, 0); -lean_inc(x_227); -x_228 = lean_ctor_get(x_226, 1); -lean_inc(x_228); -lean_dec(x_226); -x_229 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__1(x_203, x_227, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_228); -lean_dec(x_227); -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_22 = x_230; -x_23 = x_231; -goto block_26; -} +lean_object* x_26; size_t x_27; size_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 = 1; +x_28 = lean_usize_add(x_6, x_27); +x_29 = lean_box(0); +x_6 = x_28; +x_7 = x_29; +x_16 = x_26; +goto _start; } else { -lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; -lean_dec(x_199); -lean_dec(x_197); -lean_dec(x_193); -lean_dec(x_190); -lean_dec(x_188); -lean_free_object(x_6); -lean_dec(x_21); -lean_dec(x_16); +uint8_t x_31; lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); @@ -10253,489 +15053,271 @@ lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); -lean_dec(x_3); +lean_dec(x_8); lean_dec(x_1); -x_232 = lean_ctor_get(x_200, 0); -lean_inc(x_232); -x_233 = lean_ctor_get(x_200, 1); -lean_inc(x_233); -if (lean_is_exclusive(x_200)) { - lean_ctor_release(x_200, 0); - lean_ctor_release(x_200, 1); - x_234 = x_200; -} else { - lean_dec_ref(x_200); - 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; -} +x_31 = !lean_is_exclusive(x_25); +if (x_31 == 0) +{ +return x_25; } else { -lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; uint8_t x_241; -lean_free_object(x_6); -lean_inc(x_193); -lean_inc(x_195); -x_236 = lean_alloc_ctor(0, 6, 0); -lean_ctor_set(x_236, 0, x_188); -lean_ctor_set(x_236, 1, x_189); -lean_ctor_set(x_236, 2, x_190); -lean_ctor_set(x_236, 3, x_191); -lean_ctor_set(x_236, 4, x_195); -lean_ctor_set(x_236, 5, x_193); -if (lean_is_exclusive(x_195)) { - lean_ctor_release(x_195, 0); - lean_ctor_release(x_195, 1); - x_237 = x_195; -} else { - lean_dec_ref(x_195); - x_237 = lean_box(0); -} -x_238 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__2; -x_239 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_238, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); -x_240 = lean_ctor_get(x_239, 0); -lean_inc(x_240); -x_241 = lean_unbox(x_240); -lean_dec(x_240); -if (x_241 == 0) -{ -lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; -lean_dec(x_237); -lean_dec(x_193); -x_242 = lean_ctor_get(x_239, 1); -lean_inc(x_242); -lean_dec(x_239); -x_243 = lean_box(0); -x_244 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__2(x_236, x_243, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_242); -x_245 = lean_ctor_get(x_244, 0); -lean_inc(x_245); -x_246 = lean_ctor_get(x_244, 1); -lean_inc(x_246); -lean_dec(x_244); -x_22 = x_245; -x_23 = x_246; -goto block_26; -} -else -{ -lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; -x_247 = lean_ctor_get(x_239, 1); -lean_inc(x_247); -if (lean_is_exclusive(x_239)) { - lean_ctor_release(x_239, 0); - lean_ctor_release(x_239, 1); - x_248 = x_239; -} else { - lean_dec_ref(x_239); - x_248 = lean_box(0); -} -x_249 = l_Lean_Meta_Grind_Origin_key(x_193); -lean_dec(x_193); -x_250 = l_Lean_MessageData_ofName(x_249); -x_251 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__8; -if (lean_is_scalar(x_248)) { - x_252 = lean_alloc_ctor(7, 2, 0); -} else { - x_252 = x_248; - lean_ctor_set_tag(x_252, 7); -} -lean_ctor_set(x_252, 0, x_251); -lean_ctor_set(x_252, 1, x_250); -x_253 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__10; -if (lean_is_scalar(x_237)) { - x_254 = lean_alloc_ctor(7, 2, 0); -} else { - x_254 = x_237; - lean_ctor_set_tag(x_254, 7); +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_25, 0); +x_33 = lean_ctor_get(x_25, 1); +lean_inc(x_33); +lean_inc(x_32); +lean_dec(x_25); +x_34 = lean_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_ctor_set(x_254, 0, x_252); -lean_ctor_set(x_254, 1, x_253); -x_255 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_238, x_254, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_247); -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 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__2(x_236, x_256, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_257); -lean_dec(x_256); -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_22 = x_259; -x_23 = x_260; -goto block_26; } } +else +{ +uint8_t x_35; +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_1); +x_35 = !lean_is_exclusive(x_22); +if (x_35 == 0) +{ +return x_22; } -block_26: +else { -lean_object* x_24; -x_24 = lean_ctor_get(x_22, 0); -lean_inc(x_24); +lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_22, 0); +x_37 = lean_ctor_get(x_22, 1); +lean_inc(x_37); +lean_inc(x_36); lean_dec(x_22); -x_6 = x_21; -x_7 = x_24; -x_8 = lean_box(0); -x_17 = x_23; -goto _start; +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_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; -x_261 = lean_ctor_get(x_6, 0); -x_262 = lean_ctor_get(x_6, 1); -lean_inc(x_262); -lean_inc(x_261); -lean_dec(x_6); -x_268 = lean_ctor_get(x_261, 0); -lean_inc(x_268); -x_269 = lean_ctor_get(x_261, 1); -lean_inc(x_269); -x_270 = lean_ctor_get(x_261, 2); -lean_inc(x_270); -x_271 = lean_ctor_get(x_261, 3); -lean_inc(x_271); -x_272 = lean_ctor_get(x_261, 4); -lean_inc(x_272); -x_273 = lean_ctor_get(x_261, 5); -lean_inc(x_273); -if (lean_is_exclusive(x_261)) { - lean_ctor_release(x_261, 0); - lean_ctor_release(x_261, 1); - lean_ctor_release(x_261, 2); - lean_ctor_release(x_261, 3); - lean_ctor_release(x_261, 4); - lean_ctor_release(x_261, 5); - x_274 = x_261; -} else { - lean_dec_ref(x_261); - x_274 = lean_box(0); } -x_275 = lean_box(0); -lean_inc(x_3); -x_276 = l_List_filterTR_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__6(x_3, x_272, x_275); -if (lean_obj_tag(x_276) == 0) -{ -lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; -x_277 = l_Lean_Meta_Grind_shareCommon(x_269, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); -x_278 = lean_ctor_get(x_277, 0); -lean_inc(x_278); -x_279 = lean_ctor_get(x_277, 1); -lean_inc(x_279); -if (lean_is_exclusive(x_277)) { - lean_ctor_release(x_277, 0); - lean_ctor_release(x_277, 1); - x_280 = x_277; -} else { - lean_dec_ref(x_277); - x_280 = lean_box(0); } -lean_inc(x_16); -lean_inc(x_15); +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_13 = lean_st_ref_take(x_4, x_12); +x_14 = lean_ctor_get(x_13, 0); lean_inc(x_14); -lean_inc(x_13); -lean_inc(x_12); +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); +x_17 = lean_ctor_get(x_14, 1); +lean_inc(x_17); +x_18 = lean_ctor_get(x_14, 2); +lean_inc(x_18); +x_19 = lean_ctor_get(x_14, 3); +lean_inc(x_19); +x_20 = lean_ctor_get(x_14, 4); +lean_inc(x_20); +x_21 = lean_ctor_get(x_14, 5); +lean_inc(x_21); +x_22 = lean_ctor_get_uint8(x_14, sizeof(void*)*20); +x_23 = lean_ctor_get(x_14, 6); +lean_inc(x_23); +x_24 = lean_ctor_get(x_14, 7); +lean_inc(x_24); +x_25 = lean_ctor_get(x_14, 8); +lean_inc(x_25); +x_26 = lean_ctor_get(x_14, 9); +lean_inc(x_26); +x_27 = lean_ctor_get(x_14, 10); +lean_inc(x_27); +x_28 = lean_ctor_get(x_14, 11); +lean_inc(x_28); +x_29 = lean_ctor_get(x_14, 12); +lean_inc(x_29); +x_30 = lean_ctor_get(x_14, 13); +lean_inc(x_30); +x_31 = lean_ctor_get(x_14, 14); +lean_inc(x_31); +x_32 = lean_ctor_get(x_14, 15); +lean_inc(x_32); +x_33 = lean_box(0); +lean_inc(x_1); +x_34 = l_Lean_PersistentHashMap_insert___at_Lean_NameSSet_insert___spec__2(x_32, x_1, x_33); +x_35 = lean_ctor_get(x_14, 16); +lean_inc(x_35); +x_36 = lean_ctor_get(x_14, 17); +lean_inc(x_36); +x_37 = lean_ctor_get(x_14, 18); +lean_inc(x_37); +x_38 = lean_ctor_get(x_14, 19); +lean_inc(x_38); +lean_dec(x_14); +x_39 = lean_alloc_ctor(0, 20, 1); +lean_ctor_set(x_39, 0, x_16); +lean_ctor_set(x_39, 1, x_17); +lean_ctor_set(x_39, 2, x_18); +lean_ctor_set(x_39, 3, x_19); +lean_ctor_set(x_39, 4, x_20); +lean_ctor_set(x_39, 5, x_21); +lean_ctor_set(x_39, 6, x_23); +lean_ctor_set(x_39, 7, x_24); +lean_ctor_set(x_39, 8, x_25); +lean_ctor_set(x_39, 9, x_26); +lean_ctor_set(x_39, 10, x_27); +lean_ctor_set(x_39, 11, x_28); +lean_ctor_set(x_39, 12, x_29); +lean_ctor_set(x_39, 13, x_30); +lean_ctor_set(x_39, 14, x_31); +lean_ctor_set(x_39, 15, x_34); +lean_ctor_set(x_39, 16, x_35); +lean_ctor_set(x_39, 17, x_36); +lean_ctor_set(x_39, 18, x_37); +lean_ctor_set(x_39, 19, x_38); +lean_ctor_set_uint8(x_39, sizeof(void*)*20, x_22); +x_40 = lean_st_ref_set(x_4, x_39, x_15); +x_41 = lean_ctor_get(x_40, 1); +lean_inc(x_41); +lean_dec(x_40); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); -lean_inc(x_1); -x_281 = l_List_mapM_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__7(x_1, x_271, x_275, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_279); -if (lean_obj_tag(x_281) == 0) +lean_inc(x_8); +x_42 = lean_get_match_equations_for(x_1, x_8, x_9, x_10, x_11, x_41); +if (lean_obj_tag(x_42) == 0) { -lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; uint8_t x_288; -x_282 = lean_ctor_get(x_281, 0); -lean_inc(x_282); -x_283 = lean_ctor_get(x_281, 1); -lean_inc(x_283); -lean_dec(x_281); -lean_inc(x_273); -lean_inc(x_282); -if (lean_is_scalar(x_274)) { - x_284 = lean_alloc_ctor(0, 6, 0); -} else { - x_284 = x_274; -} -lean_ctor_set(x_284, 0, x_268); -lean_ctor_set(x_284, 1, x_278); -lean_ctor_set(x_284, 2, x_270); -lean_ctor_set(x_284, 3, x_282); -lean_ctor_set(x_284, 4, x_276); -lean_ctor_set(x_284, 5, x_273); -x_285 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__2; -x_286 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_285, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_283); -x_287 = lean_ctor_get(x_286, 0); -lean_inc(x_287); -x_288 = lean_unbox(x_287); -lean_dec(x_287); -if (x_288 == 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; +x_43 = lean_ctor_get(x_42, 0); +lean_inc(x_43); +x_44 = lean_ctor_get(x_42, 1); +lean_inc(x_44); +lean_dec(x_42); +x_45 = lean_ctor_get(x_43, 0); +lean_inc(x_45); +lean_dec(x_43); +x_46 = lean_box(0); +x_47 = lean_array_size(x_45); +x_48 = 0; +x_49 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns___spec__2(x_2, x_45, x_46, x_45, x_47, x_48, x_33, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_44); +lean_dec(x_45); +if (lean_obj_tag(x_49) == 0) { -lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; -lean_dec(x_282); -lean_dec(x_280); -lean_dec(x_273); -x_289 = lean_ctor_get(x_286, 1); -lean_inc(x_289); -lean_dec(x_286); -x_290 = lean_box(0); -x_291 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__1(x_284, x_290, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_289); -x_292 = lean_ctor_get(x_291, 0); -lean_inc(x_292); -x_293 = lean_ctor_get(x_291, 1); -lean_inc(x_293); -lean_dec(x_291); -x_263 = x_292; -x_264 = x_293; -goto block_267; +uint8_t x_50; +x_50 = !lean_is_exclusive(x_49); +if (x_50 == 0) +{ +lean_object* x_51; +x_51 = lean_ctor_get(x_49, 0); +lean_dec(x_51); +lean_ctor_set(x_49, 0, x_33); +return x_49; } else { -lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; -x_294 = lean_ctor_get(x_286, 1); -lean_inc(x_294); -if (lean_is_exclusive(x_286)) { - lean_ctor_release(x_286, 0); - lean_ctor_release(x_286, 1); - x_295 = x_286; -} else { - lean_dec_ref(x_286); - x_295 = lean_box(0); -} -x_296 = l_Lean_Meta_Grind_Origin_key(x_273); -lean_dec(x_273); -x_297 = l_Lean_MessageData_ofName(x_296); -x_298 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__4; -if (lean_is_scalar(x_295)) { - x_299 = lean_alloc_ctor(7, 2, 0); -} else { - x_299 = x_295; - lean_ctor_set_tag(x_299, 7); -} -lean_ctor_set(x_299, 0, x_298); -lean_ctor_set(x_299, 1, x_297); -x_300 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__6; -if (lean_is_scalar(x_280)) { - x_301 = lean_alloc_ctor(7, 2, 0); -} else { - x_301 = x_280; - lean_ctor_set_tag(x_301, 7); -} -lean_ctor_set(x_301, 0, x_299); -lean_ctor_set(x_301, 1, x_300); -x_302 = l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__3(x_282, x_275); -x_303 = l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__4(x_302, x_275); -x_304 = l_Lean_MessageData_ofList(x_303); -x_305 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_305, 0, x_301); -lean_ctor_set(x_305, 1, x_304); -x_306 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__5; -x_307 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_307, 0, x_305); -lean_ctor_set(x_307, 1, x_306); -x_308 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_285, x_307, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_294); -x_309 = lean_ctor_get(x_308, 0); -lean_inc(x_309); -x_310 = lean_ctor_get(x_308, 1); -lean_inc(x_310); -lean_dec(x_308); -x_311 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__1(x_284, x_309, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_310); -lean_dec(x_309); -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_263 = x_312; -x_264 = x_313; -goto block_267; +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_33); +lean_ctor_set(x_53, 1, x_52); +return x_53; } } else { -lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; -lean_dec(x_280); -lean_dec(x_278); -lean_dec(x_274); -lean_dec(x_273); -lean_dec(x_270); -lean_dec(x_268); -lean_dec(x_262); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_3); -lean_dec(x_1); -x_314 = lean_ctor_get(x_281, 0); -lean_inc(x_314); -x_315 = lean_ctor_get(x_281, 1); -lean_inc(x_315); -if (lean_is_exclusive(x_281)) { - lean_ctor_release(x_281, 0); - lean_ctor_release(x_281, 1); - x_316 = x_281; -} else { - lean_dec_ref(x_281); - x_316 = lean_box(0); +uint8_t x_54; +x_54 = !lean_is_exclusive(x_49); +if (x_54 == 0) +{ +return x_49; } -if (lean_is_scalar(x_316)) { - x_317 = lean_alloc_ctor(1, 2, 0); -} else { - x_317 = x_316; +else +{ +lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_55 = lean_ctor_get(x_49, 0); +x_56 = lean_ctor_get(x_49, 1); +lean_inc(x_56); +lean_inc(x_55); +lean_dec(x_49); +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_ctor_set(x_317, 0, x_314); -lean_ctor_set(x_317, 1, x_315); -return x_317; } } else { -lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; uint8_t x_323; -lean_inc(x_273); -lean_inc(x_276); -if (lean_is_scalar(x_274)) { - x_318 = lean_alloc_ctor(0, 6, 0); -} else { - x_318 = x_274; -} -lean_ctor_set(x_318, 0, x_268); -lean_ctor_set(x_318, 1, x_269); -lean_ctor_set(x_318, 2, x_270); -lean_ctor_set(x_318, 3, x_271); -lean_ctor_set(x_318, 4, x_276); -lean_ctor_set(x_318, 5, x_273); -if (lean_is_exclusive(x_276)) { - lean_ctor_release(x_276, 0); - lean_ctor_release(x_276, 1); - x_319 = x_276; -} else { - lean_dec_ref(x_276); - x_319 = lean_box(0); -} -x_320 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__2; -x_321 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_320, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); -x_322 = lean_ctor_get(x_321, 0); -lean_inc(x_322); -x_323 = lean_unbox(x_322); -lean_dec(x_322); -if (x_323 == 0) +uint8_t x_58; +lean_dec(x_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); +x_58 = !lean_is_exclusive(x_42); +if (x_58 == 0) { -lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; -lean_dec(x_319); -lean_dec(x_273); -x_324 = lean_ctor_get(x_321, 1); -lean_inc(x_324); -lean_dec(x_321); -x_325 = lean_box(0); -x_326 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__2(x_318, x_325, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_324); -x_327 = lean_ctor_get(x_326, 0); -lean_inc(x_327); -x_328 = lean_ctor_get(x_326, 1); -lean_inc(x_328); -lean_dec(x_326); -x_263 = x_327; -x_264 = x_328; -goto block_267; +return x_42; } else { -lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; -x_329 = lean_ctor_get(x_321, 1); -lean_inc(x_329); -if (lean_is_exclusive(x_321)) { - lean_ctor_release(x_321, 0); - lean_ctor_release(x_321, 1); - x_330 = x_321; -} else { - lean_dec_ref(x_321); - x_330 = lean_box(0); -} -x_331 = l_Lean_Meta_Grind_Origin_key(x_273); -lean_dec(x_273); -x_332 = l_Lean_MessageData_ofName(x_331); -x_333 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__8; -if (lean_is_scalar(x_330)) { - x_334 = lean_alloc_ctor(7, 2, 0); -} else { - x_334 = x_330; - lean_ctor_set_tag(x_334, 7); -} -lean_ctor_set(x_334, 0, x_333); -lean_ctor_set(x_334, 1, x_332); -x_335 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__10; -if (lean_is_scalar(x_319)) { - x_336 = lean_alloc_ctor(7, 2, 0); -} else { - x_336 = x_319; - lean_ctor_set_tag(x_336, 7); -} -lean_ctor_set(x_336, 0, x_334); -lean_ctor_set(x_336, 1, x_335); -x_337 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_320, x_336, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_329); -x_338 = lean_ctor_get(x_337, 0); -lean_inc(x_338); -x_339 = lean_ctor_get(x_337, 1); -lean_inc(x_339); -lean_dec(x_337); -x_340 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__2(x_318, x_338, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_339); -lean_dec(x_338); -x_341 = lean_ctor_get(x_340, 0); -lean_inc(x_341); -x_342 = lean_ctor_get(x_340, 1); -lean_inc(x_342); -lean_dec(x_340); -x_263 = x_341; -x_264 = x_342; -goto block_267; -} -} -block_267: -{ -lean_object* x_265; -x_265 = lean_ctor_get(x_263, 0); -lean_inc(x_265); -lean_dec(x_263); -x_6 = x_262; -x_7 = x_265; -x_8 = lean_box(0); -x_17 = x_264; -goto _start; -} +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; } } } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -lean_object* x_12; uint8_t x_13; -x_12 = lean_st_ref_get(x_3, x_11); -x_13 = !lean_is_exclusive(x_12); -if (x_13 == 0) +lean_object* x_13; uint8_t x_14; +x_13 = lean_st_ref_get(x_4, x_12); +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_14 = lean_ctor_get(x_12, 0); -x_15 = lean_ctor_get(x_12, 1); -x_16 = lean_ctor_get(x_14, 10); -lean_inc(x_16); -lean_dec(x_14); -x_17 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__2(x_16, x_1); -if (lean_obj_tag(x_17) == 0) +lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_15 = lean_ctor_get(x_13, 0); +x_16 = lean_ctor_get(x_13, 1); +x_17 = lean_ctor_get(x_15, 15); +lean_inc(x_17); +lean_dec(x_15); +x_18 = l_Lean_PersistentHashMap_contains___at_Lean_NameSSet_contains___spec__2(x_17, x_1); +if (x_18 == 0) { -lean_object* x_18; +lean_object* x_19; lean_object* x_20; +lean_free_object(x_13); +x_19 = lean_box(0); +x_20 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns___lambda__1(x_1, x_2, x_19, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_16); +return x_20; +} +else +{ +lean_object* x_21; +lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -10743,229 +15325,149 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); lean_dec(x_2); -x_18 = lean_box(0); -lean_ctor_set(x_12, 0, x_18); -return x_12; +lean_dec(x_1); +x_21 = lean_box(0); +lean_ctor_set(x_13, 0, x_21); +return x_13; +} } else { -lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; -lean_free_object(x_12); -x_19 = lean_ctor_get(x_17, 0); -lean_inc(x_19); -lean_dec(x_17); -x_20 = lean_st_ref_take(x_3, x_15); -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; +x_22 = lean_ctor_get(x_13, 0); +x_23 = lean_ctor_get(x_13, 1); +lean_inc(x_23); lean_inc(x_22); -lean_dec(x_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; lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_24 = lean_ctor_get(x_21, 10); -x_25 = l_Lean_PersistentHashMap_erase___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__1(x_24, x_1); -lean_ctor_set(x_21, 10, x_25); -x_26 = lean_st_ref_set(x_3, x_21, x_22); -x_27 = lean_ctor_get(x_26, 1); -lean_inc(x_27); -lean_dec(x_26); -x_28 = lean_st_ref_get(x_3, 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_ctor_get(x_29, 4); -lean_inc(x_31); -lean_dec(x_29); -x_32 = lean_box(0); -x_33 = lean_box(0); -lean_inc(x_19); -x_34 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8(x_2, x_19, x_31, x_32, x_19, x_19, x_33, lean_box(0), x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_30); -lean_dec(x_19); -if (lean_obj_tag(x_34) == 0) -{ -uint8_t x_35; -x_35 = !lean_is_exclusive(x_34); -if (x_35 == 0) +lean_dec(x_13); +x_24 = lean_ctor_get(x_22, 15); +lean_inc(x_24); +lean_dec(x_22); +x_25 = l_Lean_PersistentHashMap_contains___at_Lean_NameSSet_contains___spec__2(x_24, x_1); +if (x_25 == 0) { -lean_object* x_36; -x_36 = lean_ctor_get(x_34, 0); -lean_dec(x_36); -lean_ctor_set(x_34, 0, x_33); -return x_34; +lean_object* x_26; lean_object* x_27; +x_26 = lean_box(0); +x_27 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns___lambda__1(x_1, x_2, x_26, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_23); +return x_27; } 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_33); -lean_ctor_set(x_38, 1, x_37); -return x_38; +lean_object* x_28; lean_object* x_29; +lean_dec(x_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_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_23); +return x_29; } } -else +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: { -uint8_t x_39; -x_39 = !lean_is_exclusive(x_34); -if (x_39 == 0) +if (lean_obj_tag(x_1) == 4) { -return x_34; -} -else +lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_13 = lean_ctor_get(x_1, 0); +lean_inc(x_13); +lean_dec(x_1); +lean_inc(x_13); +x_14 = l_Lean_Meta_isMatcher___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns___spec__1(x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +x_16 = lean_unbox(x_15); +lean_dec(x_15); +if (x_16 == 0) { -lean_object* x_40; lean_object* x_41; lean_object* x_42; -x_40 = lean_ctor_get(x_34, 0); -x_41 = lean_ctor_get(x_34, 1); -lean_inc(x_41); -lean_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; -} -} +uint8_t x_17; +lean_dec(x_13); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +x_17 = !lean_is_exclusive(x_14); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; +x_18 = lean_ctor_get(x_14, 0); +lean_dec(x_18); +x_19 = lean_box(0); +lean_ctor_set(x_14, 0, x_19); +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; 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; lean_object* x_62; 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_43 = lean_ctor_get(x_21, 0); -x_44 = lean_ctor_get(x_21, 1); -x_45 = lean_ctor_get(x_21, 2); -x_46 = lean_ctor_get(x_21, 3); -x_47 = lean_ctor_get(x_21, 4); -x_48 = lean_ctor_get(x_21, 5); -x_49 = lean_ctor_get_uint8(x_21, sizeof(void*)*14); -x_50 = lean_ctor_get(x_21, 6); -x_51 = lean_ctor_get(x_21, 7); -x_52 = lean_ctor_get(x_21, 8); -x_53 = lean_ctor_get(x_21, 9); -x_54 = lean_ctor_get(x_21, 10); -x_55 = lean_ctor_get(x_21, 11); -x_56 = lean_ctor_get(x_21, 12); -x_57 = lean_ctor_get(x_21, 13); -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_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_21); -x_58 = l_Lean_PersistentHashMap_erase___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__1(x_54, x_1); -x_59 = lean_alloc_ctor(0, 14, 1); -lean_ctor_set(x_59, 0, x_43); -lean_ctor_set(x_59, 1, x_44); -lean_ctor_set(x_59, 2, x_45); -lean_ctor_set(x_59, 3, x_46); -lean_ctor_set(x_59, 4, x_47); -lean_ctor_set(x_59, 5, x_48); -lean_ctor_set(x_59, 6, x_50); -lean_ctor_set(x_59, 7, x_51); -lean_ctor_set(x_59, 8, x_52); -lean_ctor_set(x_59, 9, x_53); -lean_ctor_set(x_59, 10, x_58); -lean_ctor_set(x_59, 11, x_55); -lean_ctor_set(x_59, 12, x_56); -lean_ctor_set(x_59, 13, x_57); -lean_ctor_set_uint8(x_59, sizeof(void*)*14, x_49); -x_60 = lean_st_ref_set(x_3, x_59, x_22); -x_61 = lean_ctor_get(x_60, 1); -lean_inc(x_61); -lean_dec(x_60); -x_62 = lean_st_ref_get(x_3, x_61); -x_63 = lean_ctor_get(x_62, 0); -lean_inc(x_63); -x_64 = lean_ctor_get(x_62, 1); -lean_inc(x_64); -lean_dec(x_62); -x_65 = lean_ctor_get(x_63, 4); -lean_inc(x_65); -lean_dec(x_63); -x_66 = lean_box(0); -x_67 = lean_box(0); -lean_inc(x_19); -x_68 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8(x_2, x_19, x_65, x_66, x_19, x_19, x_67, lean_box(0), x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_64); -lean_dec(x_19); -if (lean_obj_tag(x_68) == 0) -{ -lean_object* x_69; lean_object* x_70; lean_object* x_71; -x_69 = lean_ctor_get(x_68, 1); -lean_inc(x_69); -if (lean_is_exclusive(x_68)) { - lean_ctor_release(x_68, 0); - lean_ctor_release(x_68, 1); - x_70 = x_68; -} else { - lean_dec_ref(x_68); - x_70 = lean_box(0); -} -if (lean_is_scalar(x_70)) { - x_71 = lean_alloc_ctor(0, 2, 0); -} else { - x_71 = x_70; +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_14, 1); +lean_inc(x_20); +lean_dec(x_14); +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_ctor_set(x_71, 0, x_67); -lean_ctor_set(x_71, 1, x_69); -return x_71; } else { -lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; -x_72 = lean_ctor_get(x_68, 0); -lean_inc(x_72); -x_73 = lean_ctor_get(x_68, 1); -lean_inc(x_73); -if (lean_is_exclusive(x_68)) { - lean_ctor_release(x_68, 0); - lean_ctor_release(x_68, 1); - x_74 = x_68; -} else { - lean_dec_ref(x_68); - 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_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_14, 1); +lean_inc(x_23); +lean_dec(x_14); +x_24 = lean_box(0); +x_25 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns___lambda__2(x_13, x_2, x_24, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_23); +return x_25; } -lean_ctor_set(x_75, 0, x_72); -lean_ctor_set(x_75, 1, x_73); -return x_75; } +else +{ +lean_object* x_26; lean_object* x_27; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +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_12); +return x_27; } } } -else +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, 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_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; -x_76 = lean_ctor_get(x_12, 0); -x_77 = lean_ctor_get(x_12, 1); -lean_inc(x_77); -lean_inc(x_76); -lean_dec(x_12); -x_78 = lean_ctor_get(x_76, 10); -lean_inc(x_78); -lean_dec(x_76); -x_79 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__2(x_78, x_1); -if (lean_obj_tag(x_79) == 0) +lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_12 = l_Lean_Meta_Grind_getConfig___rarg(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_ctor_get_uint8(x_13, sizeof(void*)*4); +lean_dec(x_13); +if (x_14 == 0) { -lean_object* x_80; lean_object* x_81; +uint8_t x_15; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -10975,159 +15477,39 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -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_dec(x_1); +x_15 = !lean_is_exclusive(x_12); +if (x_15 == 0) { -lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; uint8_t x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; 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_82 = lean_ctor_get(x_79, 0); -lean_inc(x_82); -lean_dec(x_79); -x_83 = lean_st_ref_take(x_3, x_77); -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_ctor_get(x_84, 0); -lean_inc(x_86); -x_87 = lean_ctor_get(x_84, 1); -lean_inc(x_87); -x_88 = lean_ctor_get(x_84, 2); -lean_inc(x_88); -x_89 = lean_ctor_get(x_84, 3); -lean_inc(x_89); -x_90 = lean_ctor_get(x_84, 4); -lean_inc(x_90); -x_91 = lean_ctor_get(x_84, 5); -lean_inc(x_91); -x_92 = lean_ctor_get_uint8(x_84, sizeof(void*)*14); -x_93 = lean_ctor_get(x_84, 6); -lean_inc(x_93); -x_94 = lean_ctor_get(x_84, 7); -lean_inc(x_94); -x_95 = lean_ctor_get(x_84, 8); -lean_inc(x_95); -x_96 = lean_ctor_get(x_84, 9); -lean_inc(x_96); -x_97 = lean_ctor_get(x_84, 10); -lean_inc(x_97); -x_98 = lean_ctor_get(x_84, 11); -lean_inc(x_98); -x_99 = lean_ctor_get(x_84, 12); -lean_inc(x_99); -x_100 = lean_ctor_get(x_84, 13); -lean_inc(x_100); -if (lean_is_exclusive(x_84)) { - lean_ctor_release(x_84, 0); - lean_ctor_release(x_84, 1); - lean_ctor_release(x_84, 2); - lean_ctor_release(x_84, 3); - lean_ctor_release(x_84, 4); - lean_ctor_release(x_84, 5); - lean_ctor_release(x_84, 6); - lean_ctor_release(x_84, 7); - lean_ctor_release(x_84, 8); - lean_ctor_release(x_84, 9); - lean_ctor_release(x_84, 10); - lean_ctor_release(x_84, 11); - lean_ctor_release(x_84, 12); - lean_ctor_release(x_84, 13); - x_101 = x_84; -} else { - lean_dec_ref(x_84); - x_101 = lean_box(0); +lean_object* x_16; lean_object* x_17; +x_16 = lean_ctor_get(x_12, 0); +lean_dec(x_16); +x_17 = lean_box(0); +lean_ctor_set(x_12, 0, x_17); +return x_12; } -x_102 = l_Lean_PersistentHashMap_erase___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__1(x_97, x_1); -if (lean_is_scalar(x_101)) { - x_103 = lean_alloc_ctor(0, 14, 1); -} else { - x_103 = x_101; -} -lean_ctor_set(x_103, 0, x_86); -lean_ctor_set(x_103, 1, x_87); -lean_ctor_set(x_103, 2, x_88); -lean_ctor_set(x_103, 3, x_89); -lean_ctor_set(x_103, 4, x_90); -lean_ctor_set(x_103, 5, x_91); -lean_ctor_set(x_103, 6, x_93); -lean_ctor_set(x_103, 7, x_94); -lean_ctor_set(x_103, 8, x_95); -lean_ctor_set(x_103, 9, x_96); -lean_ctor_set(x_103, 10, x_102); -lean_ctor_set(x_103, 11, x_98); -lean_ctor_set(x_103, 12, x_99); -lean_ctor_set(x_103, 13, x_100); -lean_ctor_set_uint8(x_103, sizeof(void*)*14, x_92); -x_104 = lean_st_ref_set(x_3, x_103, x_85); -x_105 = lean_ctor_get(x_104, 1); -lean_inc(x_105); -lean_dec(x_104); -x_106 = lean_st_ref_get(x_3, x_105); -x_107 = lean_ctor_get(x_106, 0); -lean_inc(x_107); -x_108 = lean_ctor_get(x_106, 1); -lean_inc(x_108); -lean_dec(x_106); -x_109 = lean_ctor_get(x_107, 4); -lean_inc(x_109); -lean_dec(x_107); -x_110 = lean_box(0); -x_111 = lean_box(0); -lean_inc(x_82); -x_112 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8(x_2, x_82, x_109, x_110, x_82, x_82, x_111, lean_box(0), x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_108); -lean_dec(x_82); -if (lean_obj_tag(x_112) == 0) +else { -lean_object* x_113; lean_object* x_114; lean_object* x_115; -x_113 = lean_ctor_get(x_112, 1); -lean_inc(x_113); -if (lean_is_exclusive(x_112)) { - lean_ctor_release(x_112, 0); - lean_ctor_release(x_112, 1); - x_114 = x_112; -} else { - lean_dec_ref(x_112); - x_114 = lean_box(0); -} -if (lean_is_scalar(x_114)) { - x_115 = lean_alloc_ctor(0, 2, 0); -} else { - x_115 = x_114; +lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_18 = lean_ctor_get(x_12, 1); +lean_inc(x_18); +lean_dec(x_12); +x_19 = lean_box(0); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_19); +lean_ctor_set(x_20, 1, x_18); +return x_20; } -lean_ctor_set(x_115, 0, x_111); -lean_ctor_set(x_115, 1, x_113); -return x_115; } else { -lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; -x_116 = lean_ctor_get(x_112, 0); -lean_inc(x_116); -x_117 = lean_ctor_get(x_112, 1); -lean_inc(x_117); -if (lean_is_exclusive(x_112)) { - lean_ctor_release(x_112, 0); - lean_ctor_release(x_112, 1); - x_118 = x_112; -} else { - lean_dec_ref(x_112); - x_118 = lean_box(0); -} -if (lean_is_scalar(x_118)) { - x_119 = lean_alloc_ctor(1, 2, 0); -} else { - x_119 = x_118; -} -lean_ctor_set(x_119, 0, x_116); -lean_ctor_set(x_119, 1, x_117); -return x_119; -} -} +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_12, 1); +lean_inc(x_21); +lean_dec(x_12); +x_22 = lean_box(0); +x_23 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns___lambda__3(x_1, x_2, x_22, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_21); +return x_23; } } } @@ -11208,13 +15590,13 @@ lean_dec(x_3); return x_13; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_internalize___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_internalize___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_13; -x_13 = l_Lean_Meta_Grind_internalize___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_3); -return x_13; +lean_object* x_12; +x_12 = l_Lean_Meta_Grind_internalize___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_2); +return x_12; } } LEAN_EXPORT lean_object* l_Lean_Meta_Grind_internalize___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) { @@ -11226,31 +15608,20 @@ lean_dec(x_3); return x_13; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_eraseAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_internalize___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -size_t x_4; lean_object* x_5; -x_4 = lean_unbox_usize(x_2); -lean_dec(x_2); -x_5 = l_Lean_PersistentHashMap_eraseAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__2(x_1, x_4, x_3); +lean_object* x_13; +x_13 = l_Lean_Meta_Grind_internalize___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); lean_dec(x_3); -return x_5; -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_erase___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__1___boxed(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; -x_3 = l_Lean_PersistentHashMap_erase___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__1(x_1, x_2); -lean_dec(x_2); -return x_3; +return x_13; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAtAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAtAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___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 = l_Lean_PersistentHashMap_containsAtAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5(x_1, x_2, x_3, x_4, x_5); +x_6 = l_Lean_PersistentHashMap_containsAtAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__3(x_1, x_2, x_3, x_4, x_5); lean_dec(x_5); lean_dec(x_2); lean_dec(x_1); @@ -11258,50 +15629,33 @@ x_7 = lean_box(x_6); return x_7; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___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___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__4(x_1, x_4, x_3); +x_5 = l_Lean_PersistentHashMap_containsAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__2(x_1, x_4, x_3); lean_dec(x_3); x_6 = lean_box(x_5); return x_6; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_contains___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__3___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_contains___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___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___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__3(x_1, x_2); +x_3 = l_Lean_PersistentHashMap_contains___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__1(x_1, x_2); lean_dec(x_2); x_4 = lean_box(x_3); return x_4; } } -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___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, lean_object* x_10, lean_object* x_11) { -_start: -{ -lean_object* x_12; -x_12 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_12; -} -} -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___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) { +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { lean_object* x_12; -x_12 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___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); +x_12 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -11314,7 +15668,7 @@ lean_dec(x_2); return x_12; } } -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___boxed(lean_object** _args) { +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___boxed(lean_object** _args) { lean_object* x_1 = _args[0]; lean_object* x_2 = _args[1]; lean_object* x_3 = _args[2]; @@ -11335,7 +15689,7 @@ lean_object* x_17 = _args[16]; _start: { lean_object* x_18; -x_18 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +x_18 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); @@ -11351,8 +15705,86 @@ lean_dec(x_1); return x_12; } } +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_activateTheorem___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_Meta_Grind_activateTheorem___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_isMatcher___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns___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_Lean_Meta_isMatcher___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns___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_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_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +_start: +{ +size_t x_17; size_t x_18; lean_object* x_19; +x_17 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_18 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_19 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns___spec__2(x_1, x_2, x_3, x_4, x_17, x_18, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_19; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns___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_13; +x_13 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns___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_3); +return x_13; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +x_13 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_3); +return x_13; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +x_13 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_3); +return x_13; +} +} lean_object* initialize_Init_Grind_Util(uint8_t builtin, lean_object*); +lean_object* initialize_Init_Grind_Lemmas(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_LitValues(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Meta_Match_MatcherInfo(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Meta_Match_MatchEqsExt(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_Tactic_Grind_Types(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_Tactic_Grind_Util(uint8_t builtin, lean_object*); static bool _G_initialized = false; @@ -11363,9 +15795,18 @@ _G_initialized = true; res = initialize_Init_Grind_Util(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +res = initialize_Init_Grind_Lemmas(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); res = initialize_Lean_Meta_LitValues(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +res = initialize_Lean_Meta_Match_MatcherInfo(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Lean_Meta_Match_MatchEqsExt(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); res = initialize_Lean_Meta_Tactic_Grind_Types(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); @@ -11376,11 +15817,6 @@ l_Lean_PersistentHashMap_findEntryAux___at_Lean_Meta_Grind_addCongrTable___spec_ l_Lean_PersistentHashMap_findEntryAux___at_Lean_Meta_Grind_addCongrTable___spec__2___closed__2 = _init_l_Lean_PersistentHashMap_findEntryAux___at_Lean_Meta_Grind_addCongrTable___spec__2___closed__2(); l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_addCongrTable___spec__5___closed__1 = _init_l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_addCongrTable___spec__5___closed__1(); lean_mark_persistent(l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_addCongrTable___spec__5___closed__1); -l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9___closed__1 = _init_l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9___closed__1(); -l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9___closed__2 = _init_l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9___closed__2(); -lean_mark_persistent(l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9___closed__2); -l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9___closed__3 = _init_l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9___closed__3(); -lean_mark_persistent(l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9___closed__3); l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__1 = _init_l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__1(); lean_mark_persistent(l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__1); l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__2 = _init_l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__2(); @@ -11395,6 +15831,8 @@ l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__6 = _init_l_Lean_Meta_Grin lean_mark_persistent(l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__6); l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__7 = _init_l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__7(); lean_mark_persistent(l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__7); +l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__8 = _init_l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__8(); +lean_mark_persistent(l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__8); l_Lean_Meta_Grind_addCongrTable___closed__1 = _init_l_Lean_Meta_Grind_addCongrTable___closed__1(); lean_mark_persistent(l_Lean_Meta_Grind_addCongrTable___closed__1); l_Lean_Meta_Grind_addCongrTable___closed__2 = _init_l_Lean_Meta_Grind_addCongrTable___closed__2(); @@ -11413,6 +15851,74 @@ l_Lean_Meta_Grind_addCongrTable___closed__8 = _init_l_Lean_Meta_Grind_addCongrTa lean_mark_persistent(l_Lean_Meta_Grind_addCongrTable___closed__8); l_Lean_Meta_Grind_addCongrTable___closed__9 = _init_l_Lean_Meta_Grind_addCongrTable___closed__9(); lean_mark_persistent(l_Lean_Meta_Grind_addCongrTable___closed__9); +l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addSplitCandidate___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addSplitCandidate___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addSplitCandidate___closed__1); +l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addSplitCandidate___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addSplitCandidate___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addSplitCandidate___closed__2); +l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addSplitCandidate___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addSplitCandidate___closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addSplitCandidate___closed__3); +l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__1); +l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__2); +l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__3); +l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__4(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__4); +l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__5 = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__5(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__5); +l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__6 = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__6(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__6); +l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__7 = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__7(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__7); +l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__8 = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__8(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__8); +l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__9 = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__9(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__9); +l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__10 = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__10(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__10); +l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__11 = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__11(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__11); +l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__12 = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__12(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__12); +l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes); +l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__1___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__1___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__1___closed__1); +l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__1___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__1___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__1___closed__2); +l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__1___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__1___closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__1___closed__3); +l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__1___closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__1___closed__4(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__1___closed__4); +l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__2___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__2___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__2___closed__1); +l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__2___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__2___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__2___closed__2); +l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__3___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__3___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__3___closed__1); +l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__3___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__3___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__3___closed__2); +l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__4___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__4___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__4___closed__1); +l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__4___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__4___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__4___closed__2); +l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__1); +l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__2); +l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__3); +l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__4(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__4); +l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__5 = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__5(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__5); +l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__6 = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__6(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__6); +l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__7 = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__7(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__7); +l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__8 = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__8(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__8); l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__1(); lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__1); l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__2(); @@ -11437,48 +15943,44 @@ l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__1 = lean_mark_persistent(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__1); l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__2 = _init_l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__2(); lean_mark_persistent(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__2); -l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__3 = _init_l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__3(); -lean_mark_persistent(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__3); -l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__4 = _init_l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__4(); -lean_mark_persistent(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__4); -l_Lean_Meta_Grind_internalize___lambda__2___closed__1 = _init_l_Lean_Meta_Grind_internalize___lambda__2___closed__1(); -lean_mark_persistent(l_Lean_Meta_Grind_internalize___lambda__2___closed__1); -l_Lean_Meta_Grind_internalize___lambda__2___closed__2 = _init_l_Lean_Meta_Grind_internalize___lambda__2___closed__2(); -lean_mark_persistent(l_Lean_Meta_Grind_internalize___lambda__2___closed__2); -l_Lean_Meta_Grind_internalize___lambda__2___closed__3 = _init_l_Lean_Meta_Grind_internalize___lambda__2___closed__3(); -lean_mark_persistent(l_Lean_Meta_Grind_internalize___lambda__2___closed__3); -l_Lean_Meta_Grind_internalize___lambda__2___closed__4 = _init_l_Lean_Meta_Grind_internalize___lambda__2___closed__4(); -lean_mark_persistent(l_Lean_Meta_Grind_internalize___lambda__2___closed__4); -l_Lean_Meta_Grind_internalize___lambda__2___closed__5 = _init_l_Lean_Meta_Grind_internalize___lambda__2___closed__5(); -lean_mark_persistent(l_Lean_Meta_Grind_internalize___lambda__2___closed__5); -l_Lean_Meta_Grind_internalize___lambda__2___closed__6 = _init_l_Lean_Meta_Grind_internalize___lambda__2___closed__6(); -lean_mark_persistent(l_Lean_Meta_Grind_internalize___lambda__2___closed__6); l_Lean_Meta_Grind_internalize___lambda__3___closed__1 = _init_l_Lean_Meta_Grind_internalize___lambda__3___closed__1(); lean_mark_persistent(l_Lean_Meta_Grind_internalize___lambda__3___closed__1); l_Lean_Meta_Grind_internalize___lambda__3___closed__2 = _init_l_Lean_Meta_Grind_internalize___lambda__3___closed__2(); lean_mark_persistent(l_Lean_Meta_Grind_internalize___lambda__3___closed__2); -l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__1___closed__1 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__1___closed__1(); -lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__1___closed__1); -l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__1 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__1(); -lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__1); -l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__2 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__2(); -lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__2); -l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__3 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__3(); -lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__3); -l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__4 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__4(); -lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__4); -l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__5 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__5(); -lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__5); -l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__6 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__6(); -lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__6); -l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__7 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__7(); -lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__7); -l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__8 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__8(); -lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__8); -l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__9 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__9(); -lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__9); -l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__10 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__10(); -lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__10); +l_Lean_Meta_Grind_internalize___lambda__3___closed__3 = _init_l_Lean_Meta_Grind_internalize___lambda__3___closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind_internalize___lambda__3___closed__3); +l_Lean_Meta_Grind_internalize___lambda__3___closed__4 = _init_l_Lean_Meta_Grind_internalize___lambda__3___closed__4(); +lean_mark_persistent(l_Lean_Meta_Grind_internalize___lambda__3___closed__4); +l_Lean_Meta_Grind_internalize___lambda__3___closed__5 = _init_l_Lean_Meta_Grind_internalize___lambda__3___closed__5(); +lean_mark_persistent(l_Lean_Meta_Grind_internalize___lambda__3___closed__5); +l_Lean_Meta_Grind_internalize___lambda__3___closed__6 = _init_l_Lean_Meta_Grind_internalize___lambda__3___closed__6(); +lean_mark_persistent(l_Lean_Meta_Grind_internalize___lambda__3___closed__6); +l_Lean_Meta_Grind_internalize___lambda__4___closed__1 = _init_l_Lean_Meta_Grind_internalize___lambda__4___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_internalize___lambda__4___closed__1); +l_Lean_Meta_Grind_internalize___lambda__4___closed__2 = _init_l_Lean_Meta_Grind_internalize___lambda__4___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_internalize___lambda__4___closed__2); +l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___lambda__1___closed__1 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___lambda__1___closed__1(); +lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___lambda__1___closed__1); +l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__1 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__1(); +lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__1); +l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__2 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__2(); +lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__2); +l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__3 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__3(); +lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__3); +l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__4 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__4(); +lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__4); +l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__5 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__5(); +lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__5); +l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__6 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__6(); +lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__6); +l_Lean_Meta_Grind_activateTheorem___closed__1 = _init_l_Lean_Meta_Grind_activateTheorem___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_activateTheorem___closed__1); +l_Lean_Meta_Grind_activateTheorem___closed__2 = _init_l_Lean_Meta_Grind_activateTheorem___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_activateTheorem___closed__2); +l_Lean_Meta_Grind_activateTheorem___closed__3 = _init_l_Lean_Meta_Grind_activateTheorem___closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind_activateTheorem___closed__3); +l_Lean_Meta_Grind_activateTheorem___closed__4 = _init_l_Lean_Meta_Grind_activateTheorem___closed__4(); +lean_mark_persistent(l_Lean_Meta_Grind_activateTheorem___closed__4); return lean_io_result_mk_ok(lean_box(0)); } #ifdef __cplusplus diff --git a/stage0/stdlib/Lean/Meta/Tactic/Grind/Intro.c b/stage0/stdlib/Lean/Meta/Tactic/Grind/Intro.c index 1ccf0922ead3..424e7a401187 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Grind/Intro.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Grind/Intro.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Lean.Meta.Tactic.Grind.Intro -// Imports: Init.Grind.Lemmas Lean.Meta.Tactic.Assert Lean.Meta.Tactic.Grind.Simp Lean.Meta.Tactic.Grind.Types Lean.Meta.Tactic.Grind.Cases Lean.Meta.Tactic.Grind.Injection Lean.Meta.Tactic.Grind.Core +// Imports: Init.Grind.Lemmas Lean.Meta.Tactic.Assert Lean.Meta.Tactic.Grind.Simp Lean.Meta.Tactic.Grind.Types Lean.Meta.Tactic.Grind.Cases Lean.Meta.Tactic.Grind.Injection Lean.Meta.Tactic.Grind.Core Lean.Meta.Tactic.Grind.Combinators #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -15,42 +15,47 @@ extern "C" { #endif lean_object* l_Lean_Expr_bindingName_x21(lean_object*); lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); +lean_object* l_Lean_Meta_Simp_Result_getProof(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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_Grind_injection_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instInhabitedIntroResult; lean_object* l_Lean_mkAppN(lean_object*, lean_object*); lean_object* l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___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___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_mkFreshExprMVarAt(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_isProp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isLet(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_iterate(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_applyInjection_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forM___at_Lean_Meta_Grind_intros_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*); lean_object* l_Lean_MVarId_getTag(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Expr_isLetFun(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_assertAt(lean_object*, lean_object*, lean_object*, 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_Grind_assertNext_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_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_Grind_Intro_0__Lean_Meta_Grind_introNext___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_mkFreshId___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_simp(lean_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_Grind_Intro_0__Lean_Meta_Grind_isCasesCandidate(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_object* l_List_appendTR___rarg(lean_object*, lean_object*); +lean_object* l_Lean_Meta_mkEqMP(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forM___at_Lean_Meta_Grind_intros_go___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext(lean_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_Grind_Intro_0__Lean_Meta_Grind_introNext(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_Grind_add(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_take(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_applyCases_x3f___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forM___at_Lean_Meta_Grind_intros_go___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___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___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_MVarId_getType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_intros_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_Meta_Grind_addNewEq(lean_object*, lean_object*, lean_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___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__2___boxed(lean_object**); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___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_MVarId_withContext___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_assertAll___closed__1; +lean_object* l_Lean_LocalDecl_value(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___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__6___boxed(lean_object**); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_intros___closed__1; lean_object* lean_st_ref_get(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_intros_go___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -59,57 +64,59 @@ lean_object* lean_array_to_list(lean_object*); lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_assertAt___closed__1; lean_object* l_Lean_Meta_Grind_isGrindCasesTarget(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__8___closed__3; LEAN_EXPORT lean_object* l_Lean_mkFreshId___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___spec__2___rarg___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_mapTR_loop___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_applyCases_x3f___spec__1(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__9___boxed(lean_object**); lean_object* l_Lean_FVarId_getType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Std_Queue_dequeue_x3f___rarg(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_intros_go___lambda__2___boxed(lean_object**); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_applyCases_x3f(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*); uint8_t l_Lean_Expr_isArrow(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_Grind_GrindTactic_iterate(lean_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_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_assertAt___closed__2; -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_intros_go___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_MVarId_assert(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_LocalContext_mkLocalDecl(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_intros_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_Lean_Meta_Grind_iterate_go(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_intros_go___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_intros_go(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__8___closed__4; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_isCasesCandidate___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_LocalDecl_userName(lean_object*); lean_object* l_Lean_Meta_Grind_addHypothesis(lean_object*, 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_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__8___closed__2; lean_object* l_Lean_Expr_bindingDomain_x21(lean_object*); lean_object* l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___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_Meta_Grind_intros_go___lambda__4___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_assertAt___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_Grind_shareCommon(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_Grind_intros(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_insert___at_Lean_MVarId_assign___spec__1(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_getLocalInstances(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__2; -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_intros_go___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__8___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___boxed(lean_object**); lean_object* l_Lean_LocalDecl_type(lean_object*); LEAN_EXPORT lean_object* l_List_forM___at_Lean_Meta_Grind_intros_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_EXPORT lean_object* l_Lean_Meta_Grind_intros_go___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_intros_go___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_getAppFn(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_intros_go___lambda__4___boxed(lean_object**); lean_object* l_List_reverse___rarg(lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__3; -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_intros_go___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_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MVarId_byContra_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkFreshId___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___spec__2___rarg(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_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_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_assertAll(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_mk(lean_object*); lean_object* l_Lean_Expr_fvar___override(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__8___boxed(lean_object**); lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_bindingBody_x21(lean_object*); LEAN_EXPORT lean_object* l_List_mapTR_loop___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_applyCases_x3f___spec__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_assign___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_cases(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_intro1Core(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_mkLambdaFVars(lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, 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___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__4; lean_object* l_Lean_Expr_bindingInfo_x21(lean_object*); uint8_t l_Lean_Expr_isForall(lean_object*); LEAN_EXPORT lean_object* l_Lean_mkFreshFVarId___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -118,9 +125,9 @@ lean_object* l_Lean_Expr_mvarId_x21(lean_object*); LEAN_EXPORT lean_object* l_Lean_mkFreshFVarId___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___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_Tactic_Grind_Intro_0__Lean_Meta_Grind_applyCases_x3f___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_assign___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_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_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_assertNext(lean_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_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__2(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_Grind_instInhabitedIntroResult() { _start: { @@ -559,1128 +566,1090 @@ x_10 = l_Lean_FVarId_getDecl(x_1, x_5, x_6, x_7, x_8, x_9); return x_10; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, uint8_t x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19, lean_object* x_20, lean_object* x_21, lean_object* x_22, lean_object* x_23, lean_object* x_24, lean_object* x_25) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___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_26; lean_object* x_27; -x_26 = l_Lean_LocalDecl_type(x_17); -lean_inc(x_24); -lean_inc(x_23); -lean_inc(x_22); -lean_inc(x_21); -lean_inc(x_26); -x_27 = l_Lean_Meta_isProp(x_26, x_21, x_22, x_23, x_24, x_25); -if (lean_obj_tag(x_27) == 0) -{ -lean_object* x_28; uint8_t x_29; -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) -{ -uint8_t x_30; -lean_dec(x_26); -lean_dec(x_24); -lean_dec(x_23); -lean_dec(x_22); -lean_dec(x_21); -x_30 = !lean_is_exclusive(x_27); -if (x_30 == 0) +lean_object* x_10; uint8_t x_11; +x_10 = lean_st_mk_ref(x_1, x_9); +x_11 = !lean_is_exclusive(x_10); +if (x_11 == 0) { -lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_31 = lean_ctor_get(x_27, 0); -lean_dec(x_31); -x_32 = lean_alloc_ctor(0, 14, 1); -lean_ctor_set(x_32, 0, x_1); -lean_ctor_set(x_32, 1, x_2); -lean_ctor_set(x_32, 2, x_3); -lean_ctor_set(x_32, 3, x_4); -lean_ctor_set(x_32, 4, x_5); -lean_ctor_set(x_32, 5, x_6); -lean_ctor_set(x_32, 6, x_8); -lean_ctor_set(x_32, 7, x_9); -lean_ctor_set(x_32, 8, x_10); -lean_ctor_set(x_32, 9, x_11); -lean_ctor_set(x_32, 10, x_12); -lean_ctor_set(x_32, 11, x_13); -lean_ctor_set(x_32, 12, x_14); -lean_ctor_set(x_32, 13, x_15); -lean_ctor_set_uint8(x_32, sizeof(void*)*14, x_7); -x_33 = lean_alloc_ctor(3, 2, 0); -lean_ctor_set(x_33, 0, x_16); -lean_ctor_set(x_33, 1, x_32); -lean_ctor_set(x_27, 0, x_33); -return x_27; +return x_10; } else { -lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; -x_34 = lean_ctor_get(x_27, 1); -lean_inc(x_34); -lean_dec(x_27); -x_35 = lean_alloc_ctor(0, 14, 1); -lean_ctor_set(x_35, 0, x_1); -lean_ctor_set(x_35, 1, x_2); -lean_ctor_set(x_35, 2, x_3); -lean_ctor_set(x_35, 3, x_4); -lean_ctor_set(x_35, 4, x_5); -lean_ctor_set(x_35, 5, x_6); -lean_ctor_set(x_35, 6, x_8); -lean_ctor_set(x_35, 7, x_9); -lean_ctor_set(x_35, 8, x_10); -lean_ctor_set(x_35, 9, x_11); -lean_ctor_set(x_35, 10, x_12); -lean_ctor_set(x_35, 11, x_13); -lean_ctor_set(x_35, 12, x_14); -lean_ctor_set(x_35, 13, x_15); -lean_ctor_set_uint8(x_35, sizeof(void*)*14, x_7); -x_36 = lean_alloc_ctor(3, 2, 0); -lean_ctor_set(x_36, 0, x_16); -lean_ctor_set(x_36, 1, x_35); -x_37 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_37, 0, x_36); -lean_ctor_set(x_37, 1, x_34); -return x_37; +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 +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: { -lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; -x_38 = lean_ctor_get(x_27, 1); -lean_inc(x_38); -lean_dec(x_27); -x_39 = l_Lean_LocalDecl_userName(x_17); -x_40 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_39, x_23, x_24, x_38); -x_41 = lean_ctor_get(x_40, 0); -lean_inc(x_41); -x_42 = lean_ctor_get(x_40, 1); -lean_inc(x_42); -lean_dec(x_40); -x_43 = l_Lean_Expr_fvar___override(x_16); -x_44 = l_Lean_MVarId_assert(x_1, x_41, x_26, x_43, x_21, x_22, x_23, x_24, x_42); -if (lean_obj_tag(x_44) == 0) +lean_object* x_14; +lean_inc(x_5); +x_14 = l_Lean_Meta_Grind_addNewEq(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); +if (lean_obj_tag(x_14) == 0) { -uint8_t x_45; -x_45 = !lean_is_exclusive(x_44); -if (x_45 == 0) +lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_15 = lean_ctor_get(x_14, 1); +lean_inc(x_15); +lean_dec(x_14); +x_16 = lean_st_ref_get(x_5, x_15); +x_17 = !lean_is_exclusive(x_16); +if (x_17 == 0) { -lean_object* x_46; lean_object* x_47; lean_object* x_48; -x_46 = lean_ctor_get(x_44, 0); -x_47 = lean_alloc_ctor(0, 14, 1); -lean_ctor_set(x_47, 0, x_46); -lean_ctor_set(x_47, 1, x_2); -lean_ctor_set(x_47, 2, x_3); -lean_ctor_set(x_47, 3, x_4); -lean_ctor_set(x_47, 4, x_5); -lean_ctor_set(x_47, 5, x_6); -lean_ctor_set(x_47, 6, x_8); -lean_ctor_set(x_47, 7, x_9); -lean_ctor_set(x_47, 8, x_10); -lean_ctor_set(x_47, 9, x_11); -lean_ctor_set(x_47, 10, x_12); -lean_ctor_set(x_47, 11, x_13); -lean_ctor_set(x_47, 12, x_14); -lean_ctor_set(x_47, 13, x_15); -lean_ctor_set_uint8(x_47, sizeof(void*)*14, x_7); -x_48 = lean_alloc_ctor(2, 1, 0); -lean_ctor_set(x_48, 0, x_47); -lean_ctor_set(x_44, 0, x_48); -return x_44; +lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_18 = lean_ctor_get(x_16, 1); +x_19 = lean_st_ref_get(x_5, x_18); +lean_dec(x_5); +x_20 = !lean_is_exclusive(x_19); +if (x_20 == 0) +{ +lean_object* x_21; +x_21 = lean_ctor_get(x_19, 0); +lean_ctor_set(x_16, 1, x_21); +lean_ctor_set(x_19, 0, x_16); +return x_19; } else { -lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; -x_49 = lean_ctor_get(x_44, 0); -x_50 = lean_ctor_get(x_44, 1); -lean_inc(x_50); -lean_inc(x_49); -lean_dec(x_44); -x_51 = lean_alloc_ctor(0, 14, 1); -lean_ctor_set(x_51, 0, x_49); -lean_ctor_set(x_51, 1, x_2); -lean_ctor_set(x_51, 2, x_3); -lean_ctor_set(x_51, 3, x_4); -lean_ctor_set(x_51, 4, x_5); -lean_ctor_set(x_51, 5, x_6); -lean_ctor_set(x_51, 6, x_8); -lean_ctor_set(x_51, 7, x_9); -lean_ctor_set(x_51, 8, x_10); -lean_ctor_set(x_51, 9, x_11); -lean_ctor_set(x_51, 10, x_12); -lean_ctor_set(x_51, 11, x_13); -lean_ctor_set(x_51, 12, x_14); -lean_ctor_set(x_51, 13, x_15); -lean_ctor_set_uint8(x_51, sizeof(void*)*14, x_7); -x_52 = lean_alloc_ctor(2, 1, 0); -lean_ctor_set(x_52, 0, x_51); -x_53 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_53, 0, x_52); -lean_ctor_set(x_53, 1, x_50); -return x_53; +lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_22 = lean_ctor_get(x_19, 0); +x_23 = lean_ctor_get(x_19, 1); +lean_inc(x_23); +lean_inc(x_22); +lean_dec(x_19); +lean_ctor_set(x_16, 1, x_22); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_16); +lean_ctor_set(x_24, 1, x_23); +return x_24; } } else { -uint8_t x_54; -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_6); +lean_object* x_25; lean_object* x_26; 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_16, 0); +x_26 = lean_ctor_get(x_16, 1); +lean_inc(x_26); +lean_inc(x_25); +lean_dec(x_16); +x_27 = lean_st_ref_get(x_5, x_26); lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -x_54 = !lean_is_exclusive(x_44); -if (x_54 == 0) -{ -return x_44; -} -else -{ -lean_object* x_55; lean_object* x_56; lean_object* x_57; -x_55 = lean_ctor_get(x_44, 0); -x_56 = lean_ctor_get(x_44, 1); -lean_inc(x_56); -lean_inc(x_55); -lean_dec(x_44); -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; -} +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; +} else { + lean_dec_ref(x_27); + x_30 = lean_box(0); +} +x_31 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_31, 0, x_25); +lean_ctor_set(x_31, 1, x_28); +if (lean_is_scalar(x_30)) { + x_32 = lean_alloc_ctor(0, 2, 0); +} else { + x_32 = x_30; } +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_29); +return x_32; } } else { -uint8_t x_58; -lean_dec(x_26); -lean_dec(x_24); -lean_dec(x_23); -lean_dec(x_22); -lean_dec(x_21); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_6); +uint8_t x_33; lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_58 = !lean_is_exclusive(x_27); -if (x_58 == 0) +x_33 = !lean_is_exclusive(x_14); +if (x_33 == 0) { -return x_27; +return x_14; } else { -lean_object* x_59; lean_object* x_60; lean_object* x_61; -x_59 = lean_ctor_get(x_27, 0); -x_60 = lean_ctor_get(x_27, 1); -lean_inc(x_60); -lean_inc(x_59); -lean_dec(x_27); -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_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_14, 0); +x_35 = lean_ctor_get(x_14, 1); +lean_inc(x_35); +lean_inc(x_34); +lean_dec(x_14); +x_36 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_36, 0, x_34); +lean_ctor_set(x_36, 1, x_35); +return x_36; } } } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { -lean_object* x_10; -x_10 = l_Lean_Meta_isProp(x_1, x_5, x_6, x_7, x_8, x_9); -return x_10; -} +uint8_t x_10; +x_10 = !lean_is_exclusive(x_1); +if (x_10 == 0) +{ +lean_object* x_11; +x_11 = lean_ctor_get(x_1, 1); +lean_dec(x_11); +lean_ctor_set(x_1, 1, x_9); +return x_1; } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___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: +else { -uint8_t x_11; uint8_t x_12; uint8_t x_13; lean_object* x_14; -x_11 = 0; -x_12 = 1; -x_13 = 1; -x_14 = l_Lean_Meta_mkLambdaFVars(x_1, x_2, x_11, x_12, x_11, x_13, x_6, x_7, x_8, x_9, x_10); -return x_14; +lean_object* x_12; lean_object* x_13; +x_12 = lean_ctor_get(x_1, 0); +lean_inc(x_12); +lean_dec(x_1); +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_9); +return x_13; +} } } static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean", 4, 4); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__4___boxed), 9, 0); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__2() { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, uint8_t x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19, lean_object* x_20, lean_object* x_21, lean_object* x_22, lean_object* x_23, uint8_t x_24, lean_object* x_25, lean_object* x_26, lean_object* x_27, lean_object* x_28, lean_object* x_29, lean_object* x_30, lean_object* x_31, lean_object* x_32, lean_object* x_33, lean_object* x_34) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Grind", 5, 5); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__3() { -_start: +lean_object* x_35; lean_object* x_36; +x_35 = l_Lean_LocalDecl_type(x_26); +lean_inc(x_33); +lean_inc(x_32); +lean_inc(x_31); +lean_inc(x_30); +lean_inc(x_35); +x_36 = l_Lean_Meta_isProp(x_35, x_30, x_31, x_32, x_33, x_34); +if (lean_obj_tag(x_36) == 0) { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("intro_with_eq", 13, 13); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__4() { -_start: +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_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__1; -x_2 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__2; -x_3 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__3; -x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); -return x_4; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, uint8_t x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19, lean_object* x_20, lean_object* x_21, lean_object* x_22, lean_object* x_23, lean_object* x_24, lean_object* x_25, lean_object* x_26, lean_object* x_27, lean_object* x_28, lean_object* x_29, lean_object* x_30, lean_object* x_31) { -_start: +uint8_t x_39; +lean_dec(x_35); +x_39 = !lean_is_exclusive(x_36); +if (x_39 == 0) { -if (lean_obj_tag(x_1) == 0) +lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_40 = lean_ctor_get(x_36, 1); +x_41 = lean_ctor_get(x_36, 0); +lean_dec(x_41); +lean_inc(x_1); +x_42 = lean_alloc_ctor(0, 20, 1); +lean_ctor_set(x_42, 0, x_1); +lean_ctor_set(x_42, 1, x_2); +lean_ctor_set(x_42, 2, x_3); +lean_ctor_set(x_42, 3, x_4); +lean_ctor_set(x_42, 4, x_5); +lean_ctor_set(x_42, 5, x_6); +lean_ctor_set(x_42, 6, x_8); +lean_ctor_set(x_42, 7, x_9); +lean_ctor_set(x_42, 8, x_10); +lean_ctor_set(x_42, 9, x_11); +lean_ctor_set(x_42, 10, x_12); +lean_ctor_set(x_42, 11, x_13); +lean_ctor_set(x_42, 12, x_14); +lean_ctor_set(x_42, 13, x_15); +lean_ctor_set(x_42, 14, x_16); +lean_ctor_set(x_42, 15, x_17); +lean_ctor_set(x_42, 16, x_18); +lean_ctor_set(x_42, 17, x_19); +lean_ctor_set(x_42, 18, x_20); +lean_ctor_set(x_42, 19, x_21); +lean_ctor_set_uint8(x_42, sizeof(void*)*20, x_7); +if (x_24 == 0) { -lean_object* x_32; uint8_t x_33; -lean_dec(x_22); -lean_dec(x_21); -lean_dec(x_20); -lean_dec(x_19); -x_32 = l_Lean_MVarId_assign___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___spec__3(x_2, x_23, x_24, x_25, x_26, x_27, x_28, x_29, x_30, x_31); -x_33 = !lean_is_exclusive(x_32); -if (x_33 == 0) +uint8_t x_113; +x_113 = l_Lean_Expr_isLetFun(x_25); +if (x_113 == 0) { -lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_34 = lean_ctor_get(x_32, 0); -lean_dec(x_34); -x_35 = lean_alloc_ctor(0, 14, 1); -lean_ctor_set(x_35, 0, x_3); -lean_ctor_set(x_35, 1, x_4); -lean_ctor_set(x_35, 2, x_5); -lean_ctor_set(x_35, 3, x_6); -lean_ctor_set(x_35, 4, x_7); -lean_ctor_set(x_35, 5, x_8); -lean_ctor_set(x_35, 6, x_10); -lean_ctor_set(x_35, 7, x_11); -lean_ctor_set(x_35, 8, x_12); -lean_ctor_set(x_35, 9, x_13); -lean_ctor_set(x_35, 10, x_14); -lean_ctor_set(x_35, 11, x_15); -lean_ctor_set(x_35, 12, x_16); -lean_ctor_set(x_35, 13, x_17); -lean_ctor_set_uint8(x_35, sizeof(void*)*14, x_9); -x_36 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_36, 0, x_18); -lean_ctor_set(x_36, 1, x_35); -lean_ctor_set(x_32, 0, x_36); -return x_32; +lean_object* x_114; +lean_dec(x_33); +lean_dec(x_32); +lean_dec(x_31); +lean_dec(x_30); +lean_dec(x_29); +lean_dec(x_28); +lean_dec(x_27); +lean_dec(x_23); +lean_dec(x_1); +x_114 = lean_alloc_ctor(3, 2, 0); +lean_ctor_set(x_114, 0, x_22); +lean_ctor_set(x_114, 1, x_42); +lean_ctor_set(x_36, 0, x_114); +return x_36; } else { -lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_37 = lean_ctor_get(x_32, 1); -lean_inc(x_37); -lean_dec(x_32); -x_38 = lean_alloc_ctor(0, 14, 1); -lean_ctor_set(x_38, 0, x_3); -lean_ctor_set(x_38, 1, x_4); -lean_ctor_set(x_38, 2, x_5); -lean_ctor_set(x_38, 3, x_6); -lean_ctor_set(x_38, 4, x_7); -lean_ctor_set(x_38, 5, x_8); -lean_ctor_set(x_38, 6, x_10); -lean_ctor_set(x_38, 7, x_11); -lean_ctor_set(x_38, 8, x_12); -lean_ctor_set(x_38, 9, x_13); -lean_ctor_set(x_38, 10, x_14); -lean_ctor_set(x_38, 11, x_15); -lean_ctor_set(x_38, 12, x_16); -lean_ctor_set(x_38, 13, x_17); -lean_ctor_set_uint8(x_38, sizeof(void*)*14, x_9); -x_39 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_39, 0, x_18); -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_37); -return x_40; +lean_object* x_115; +lean_free_object(x_36); +x_115 = lean_box(0); +x_43 = x_115; +goto block_112; } } 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; uint8_t x_52; -x_41 = lean_ctor_get(x_1, 0); -x_42 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__4; -lean_inc(x_19); -x_43 = l_Lean_Expr_const___override(x_42, x_19); -x_44 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_44, 0, x_23); -lean_ctor_set(x_44, 1, x_19); -lean_inc(x_41); -x_45 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_45, 0, x_41); -lean_ctor_set(x_45, 1, x_44); -x_46 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_46, 0, x_20); -lean_ctor_set(x_46, 1, x_45); -x_47 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_47, 0, x_21); -lean_ctor_set(x_47, 1, x_46); -x_48 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_48, 0, x_22); -lean_ctor_set(x_48, 1, x_47); -x_49 = lean_array_mk(x_48); -x_50 = l_Lean_mkAppN(x_43, x_49); -lean_dec(x_49); -x_51 = l_Lean_MVarId_assign___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___spec__3(x_2, x_50, x_24, x_25, x_26, x_27, x_28, x_29, x_30, x_31); -x_52 = !lean_is_exclusive(x_51); -if (x_52 == 0) -{ -lean_object* x_53; lean_object* x_54; lean_object* x_55; -x_53 = lean_ctor_get(x_51, 0); -lean_dec(x_53); -x_54 = lean_alloc_ctor(0, 14, 1); -lean_ctor_set(x_54, 0, x_3); -lean_ctor_set(x_54, 1, x_4); -lean_ctor_set(x_54, 2, x_5); -lean_ctor_set(x_54, 3, x_6); -lean_ctor_set(x_54, 4, x_7); -lean_ctor_set(x_54, 5, x_8); -lean_ctor_set(x_54, 6, x_10); -lean_ctor_set(x_54, 7, x_11); -lean_ctor_set(x_54, 8, x_12); -lean_ctor_set(x_54, 9, x_13); -lean_ctor_set(x_54, 10, x_14); -lean_ctor_set(x_54, 11, x_15); -lean_ctor_set(x_54, 12, x_16); -lean_ctor_set(x_54, 13, x_17); -lean_ctor_set_uint8(x_54, sizeof(void*)*14, x_9); -x_55 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_55, 0, x_18); -lean_ctor_set(x_55, 1, x_54); -lean_ctor_set(x_51, 0, x_55); -return x_51; +lean_object* x_116; +lean_free_object(x_36); +x_116 = lean_box(0); +x_43 = x_116; +goto block_112; } -else +block_112: { -lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; -x_56 = lean_ctor_get(x_51, 1); -lean_inc(x_56); -lean_dec(x_51); -x_57 = lean_alloc_ctor(0, 14, 1); -lean_ctor_set(x_57, 0, x_3); -lean_ctor_set(x_57, 1, x_4); -lean_ctor_set(x_57, 2, x_5); -lean_ctor_set(x_57, 3, x_6); -lean_ctor_set(x_57, 4, x_7); -lean_ctor_set(x_57, 5, x_8); -lean_ctor_set(x_57, 6, x_10); -lean_ctor_set(x_57, 7, x_11); -lean_ctor_set(x_57, 8, x_12); -lean_ctor_set(x_57, 9, x_13); -lean_ctor_set(x_57, 10, x_14); -lean_ctor_set(x_57, 11, x_15); -lean_ctor_set(x_57, 12, x_16); -lean_ctor_set(x_57, 13, x_17); -lean_ctor_set_uint8(x_57, sizeof(void*)*14, x_9); -x_58 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_58, 0, x_18); -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_56); -return x_59; -} +lean_object* x_44; +lean_dec(x_43); +lean_inc(x_30); +lean_inc(x_22); +x_44 = l_Lean_FVarId_getDecl(x_22, x_30, x_31, x_32, x_33, x_40); +if (lean_obj_tag(x_44) == 0) +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_45 = lean_ctor_get(x_44, 0); +lean_inc(x_45); +x_46 = lean_ctor_get(x_44, 1); +lean_inc(x_46); +lean_dec(x_44); +x_47 = l_Lean_LocalDecl_value(x_45); +lean_dec(x_45); +lean_inc(x_33); +lean_inc(x_32); +lean_inc(x_31); +lean_inc(x_30); +lean_inc(x_28); +x_48 = l_Lean_Meta_Grind_simp(x_47, x_27, x_28, x_29, x_30, x_31, x_32, x_33, x_46); +if (lean_obj_tag(x_48) == 0) +{ +lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; uint8_t x_53; +x_49 = lean_ctor_get(x_48, 0); +lean_inc(x_49); +x_50 = lean_ctor_get(x_48, 1); +lean_inc(x_50); +lean_dec(x_48); +lean_inc(x_22); +x_51 = l_Lean_Expr_fvar___override(x_22); +x_52 = l_Lean_Meta_Grind_shareCommon(x_51, x_27, x_28, x_29, x_30, x_31, x_32, x_33, x_50); +x_53 = !lean_is_exclusive(x_52); +if (x_53 == 0) +{ +lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_54 = lean_ctor_get(x_52, 0); +x_55 = lean_ctor_get(x_52, 1); +lean_inc(x_33); +lean_inc(x_32); +lean_inc(x_31); +lean_inc(x_30); +lean_inc(x_49); +x_56 = l_Lean_Meta_Simp_Result_getProof(x_49, x_30, x_31, x_32, x_33, 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; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; +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_49, 0); +lean_inc(x_59); +lean_dec(x_49); +x_60 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__2___boxed), 9, 1); +lean_closure_set(x_60, 0, x_42); +x_61 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__3), 13, 4); +lean_closure_set(x_61, 0, x_54); +lean_closure_set(x_61, 1, x_59); +lean_closure_set(x_61, 2, x_57); +lean_closure_set(x_61, 3, x_23); +x_62 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); +lean_closure_set(x_62, 0, x_60); +lean_closure_set(x_62, 1, x_61); +x_63 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__1; +x_64 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); +lean_closure_set(x_64, 0, x_62); +lean_closure_set(x_64, 1, x_63); +x_65 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(x_1, x_64, x_27, x_28, x_29, x_30, x_31, x_32, x_33, x_58); +if (lean_obj_tag(x_65) == 0) +{ +uint8_t x_66; +x_66 = !lean_is_exclusive(x_65); +if (x_66 == 0) +{ +lean_object* x_67; +x_67 = lean_ctor_get(x_65, 0); +lean_ctor_set_tag(x_52, 3); +lean_ctor_set(x_52, 1, x_67); +lean_ctor_set(x_52, 0, x_22); +lean_ctor_set(x_65, 0, x_52); +return x_65; } +else +{ +lean_object* x_68; lean_object* x_69; lean_object* x_70; +x_68 = lean_ctor_get(x_65, 0); +x_69 = lean_ctor_get(x_65, 1); +lean_inc(x_69); +lean_inc(x_68); +lean_dec(x_65); +lean_ctor_set_tag(x_52, 3); +lean_ctor_set(x_52, 1, x_68); +lean_ctor_set(x_52, 0, x_22); +x_70 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_70, 0, x_52); +lean_ctor_set(x_70, 1, x_69); +return x_70; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___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, 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, 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) { -_start: +else { -if (x_18 == 0) +uint8_t x_71; +lean_free_object(x_52); +lean_dec(x_22); +x_71 = !lean_is_exclusive(x_65); +if (x_71 == 0) { -uint8_t x_27; lean_object* x_28; -lean_dec(x_21); -lean_dec(x_20); -lean_dec(x_19); -lean_dec(x_17); -x_27 = 1; -x_28 = l_Lean_Meta_intro1Core(x_1, x_27, x_22, x_23, x_24, x_25, x_26); -if (lean_obj_tag(x_28) == 0) +return x_65; +} +else { -uint8_t x_29; -x_29 = !lean_is_exclusive(x_28); -if (x_29 == 0) +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 { -lean_object* x_30; uint8_t x_31; -x_30 = lean_ctor_get(x_28, 0); -x_31 = !lean_is_exclusive(x_30); -if (x_31 == 0) +uint8_t x_75; +lean_free_object(x_52); +lean_dec(x_54); +lean_dec(x_49); +lean_dec(x_42); +lean_dec(x_33); +lean_dec(x_32); +lean_dec(x_31); +lean_dec(x_30); +lean_dec(x_29); +lean_dec(x_28); +lean_dec(x_27); +lean_dec(x_23); +lean_dec(x_22); +lean_dec(x_1); +x_75 = !lean_is_exclusive(x_56); +if (x_75 == 0) { -lean_object* x_32; lean_object* x_33; -x_32 = lean_ctor_get(x_30, 1); -x_33 = lean_alloc_ctor(0, 14, 1); -lean_ctor_set(x_33, 0, x_32); -lean_ctor_set(x_33, 1, x_2); -lean_ctor_set(x_33, 2, x_3); -lean_ctor_set(x_33, 3, x_4); -lean_ctor_set(x_33, 4, x_5); -lean_ctor_set(x_33, 5, x_6); -lean_ctor_set(x_33, 6, x_8); -lean_ctor_set(x_33, 7, x_9); -lean_ctor_set(x_33, 8, x_10); -lean_ctor_set(x_33, 9, x_11); -lean_ctor_set(x_33, 10, x_12); -lean_ctor_set(x_33, 11, x_13); -lean_ctor_set(x_33, 12, x_14); -lean_ctor_set(x_33, 13, x_15); -lean_ctor_set_uint8(x_33, sizeof(void*)*14, x_7); -lean_ctor_set_tag(x_30, 3); -lean_ctor_set(x_30, 1, x_33); -return x_28; +return x_56; } else { -lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; -x_34 = lean_ctor_get(x_30, 0); -x_35 = lean_ctor_get(x_30, 1); -lean_inc(x_35); -lean_inc(x_34); -lean_dec(x_30); -x_36 = lean_alloc_ctor(0, 14, 1); -lean_ctor_set(x_36, 0, x_35); -lean_ctor_set(x_36, 1, x_2); -lean_ctor_set(x_36, 2, x_3); -lean_ctor_set(x_36, 3, x_4); -lean_ctor_set(x_36, 4, x_5); -lean_ctor_set(x_36, 5, x_6); -lean_ctor_set(x_36, 6, x_8); -lean_ctor_set(x_36, 7, x_9); -lean_ctor_set(x_36, 8, x_10); -lean_ctor_set(x_36, 9, x_11); -lean_ctor_set(x_36, 10, x_12); -lean_ctor_set(x_36, 11, x_13); -lean_ctor_set(x_36, 12, x_14); -lean_ctor_set(x_36, 13, x_15); -lean_ctor_set_uint8(x_36, sizeof(void*)*14, x_7); -x_37 = lean_alloc_ctor(3, 2, 0); -lean_ctor_set(x_37, 0, x_34); -lean_ctor_set(x_37, 1, x_36); -lean_ctor_set(x_28, 0, x_37); -return x_28; +lean_object* x_76; lean_object* x_77; lean_object* x_78; +x_76 = lean_ctor_get(x_56, 0); +x_77 = lean_ctor_get(x_56, 1); +lean_inc(x_77); +lean_inc(x_76); +lean_dec(x_56); +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_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; -x_38 = lean_ctor_get(x_28, 0); -x_39 = lean_ctor_get(x_28, 1); -lean_inc(x_39); -lean_inc(x_38); -lean_dec(x_28); -x_40 = lean_ctor_get(x_38, 0); -lean_inc(x_40); -x_41 = lean_ctor_get(x_38, 1); -lean_inc(x_41); -if (lean_is_exclusive(x_38)) { - lean_ctor_release(x_38, 0); - lean_ctor_release(x_38, 1); - x_42 = x_38; +lean_object* x_79; lean_object* x_80; lean_object* x_81; +x_79 = lean_ctor_get(x_52, 0); +x_80 = lean_ctor_get(x_52, 1); +lean_inc(x_80); +lean_inc(x_79); +lean_dec(x_52); +lean_inc(x_33); +lean_inc(x_32); +lean_inc(x_31); +lean_inc(x_30); +lean_inc(x_49); +x_81 = l_Lean_Meta_Simp_Result_getProof(x_49, x_30, x_31, x_32, x_33, x_80); +if (lean_obj_tag(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; lean_object* x_90; +x_82 = lean_ctor_get(x_81, 0); +lean_inc(x_82); +x_83 = lean_ctor_get(x_81, 1); +lean_inc(x_83); +lean_dec(x_81); +x_84 = lean_ctor_get(x_49, 0); +lean_inc(x_84); +lean_dec(x_49); +x_85 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__2___boxed), 9, 1); +lean_closure_set(x_85, 0, x_42); +x_86 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__3), 13, 4); +lean_closure_set(x_86, 0, x_79); +lean_closure_set(x_86, 1, x_84); +lean_closure_set(x_86, 2, x_82); +lean_closure_set(x_86, 3, x_23); +x_87 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); +lean_closure_set(x_87, 0, x_85); +lean_closure_set(x_87, 1, x_86); +x_88 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__1; +x_89 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); +lean_closure_set(x_89, 0, x_87); +lean_closure_set(x_89, 1, x_88); +x_90 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(x_1, x_89, x_27, x_28, x_29, x_30, x_31, x_32, x_33, x_83); +if (lean_obj_tag(x_90) == 0) +{ +lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; +x_91 = lean_ctor_get(x_90, 0); +lean_inc(x_91); +x_92 = lean_ctor_get(x_90, 1); +lean_inc(x_92); +if (lean_is_exclusive(x_90)) { + lean_ctor_release(x_90, 0); + lean_ctor_release(x_90, 1); + x_93 = x_90; } else { - lean_dec_ref(x_38); - x_42 = lean_box(0); -} -x_43 = lean_alloc_ctor(0, 14, 1); -lean_ctor_set(x_43, 0, x_41); -lean_ctor_set(x_43, 1, x_2); -lean_ctor_set(x_43, 2, x_3); -lean_ctor_set(x_43, 3, x_4); -lean_ctor_set(x_43, 4, x_5); -lean_ctor_set(x_43, 5, x_6); -lean_ctor_set(x_43, 6, x_8); -lean_ctor_set(x_43, 7, x_9); -lean_ctor_set(x_43, 8, x_10); -lean_ctor_set(x_43, 9, x_11); -lean_ctor_set(x_43, 10, x_12); -lean_ctor_set(x_43, 11, x_13); -lean_ctor_set(x_43, 12, x_14); -lean_ctor_set(x_43, 13, x_15); -lean_ctor_set_uint8(x_43, sizeof(void*)*14, x_7); -if (lean_is_scalar(x_42)) { - x_44 = lean_alloc_ctor(3, 2, 0); + lean_dec_ref(x_90); + x_93 = lean_box(0); +} +x_94 = lean_alloc_ctor(3, 2, 0); +lean_ctor_set(x_94, 0, x_22); +lean_ctor_set(x_94, 1, x_91); +if (lean_is_scalar(x_93)) { + x_95 = lean_alloc_ctor(0, 2, 0); } else { - x_44 = x_42; - lean_ctor_set_tag(x_44, 3); -} -lean_ctor_set(x_44, 0, x_40); -lean_ctor_set(x_44, 1, x_43); -x_45 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_45, 0, x_44); -lean_ctor_set(x_45, 1, x_39); -return x_45; + x_95 = x_93; } +lean_ctor_set(x_95, 0, x_94); +lean_ctor_set(x_95, 1, x_92); +return x_95; } else { -uint8_t x_46; -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_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -x_46 = !lean_is_exclusive(x_28); -if (x_46 == 0) -{ -return x_28; +lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; +lean_dec(x_22); +x_96 = lean_ctor_get(x_90, 0); +lean_inc(x_96); +x_97 = lean_ctor_get(x_90, 1); +lean_inc(x_97); +if (lean_is_exclusive(x_90)) { + lean_ctor_release(x_90, 0); + lean_ctor_release(x_90, 1); + x_98 = x_90; +} else { + lean_dec_ref(x_90); + x_98 = lean_box(0); +} +if (lean_is_scalar(x_98)) { + x_99 = lean_alloc_ctor(1, 2, 0); +} else { + x_99 = x_98; +} +lean_ctor_set(x_99, 0, x_96); +lean_ctor_set(x_99, 1, x_97); +return x_99; +} } else { -lean_object* x_47; lean_object* x_48; lean_object* x_49; -x_47 = lean_ctor_get(x_28, 0); -x_48 = lean_ctor_get(x_28, 1); -lean_inc(x_48); -lean_inc(x_47); +lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; +lean_dec(x_79); +lean_dec(x_49); +lean_dec(x_42); +lean_dec(x_33); +lean_dec(x_32); +lean_dec(x_31); +lean_dec(x_30); +lean_dec(x_29); lean_dec(x_28); -x_49 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_49, 0, x_47); -lean_ctor_set(x_49, 1, x_48); -return x_49; +lean_dec(x_27); +lean_dec(x_23); +lean_dec(x_22); +lean_dec(x_1); +x_100 = lean_ctor_get(x_81, 0); +lean_inc(x_100); +x_101 = lean_ctor_get(x_81, 1); +lean_inc(x_101); +if (lean_is_exclusive(x_81)) { + lean_ctor_release(x_81, 0); + lean_ctor_release(x_81, 1); + x_102 = x_81; +} else { + lean_dec_ref(x_81); + 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(x_103, 0, x_100); +lean_ctor_set(x_103, 1, x_101); +return x_103; } } } else { -lean_object* x_50; -lean_inc(x_1); -x_50 = l_Lean_MVarId_getTag(x_1, x_22, x_23, x_24, x_25, x_26); -if (lean_obj_tag(x_50) == 0) +uint8_t x_104; +lean_dec(x_42); +lean_dec(x_33); +lean_dec(x_32); +lean_dec(x_31); +lean_dec(x_30); +lean_dec(x_29); +lean_dec(x_28); +lean_dec(x_27); +lean_dec(x_23); +lean_dec(x_22); +lean_dec(x_1); +x_104 = !lean_is_exclusive(x_48); +if (x_104 == 0) { -lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; -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_Expr_bindingBody_x21(x_16); -lean_inc(x_25); -lean_inc(x_24); -lean_inc(x_23); -lean_inc(x_22); -lean_inc(x_20); -lean_inc(x_17); -x_54 = l_Lean_Meta_Grind_simp(x_17, x_19, x_20, x_21, x_22, x_23, x_24, x_25, x_52); -if (lean_obj_tag(x_54) == 0) +return x_48; +} +else { -lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; uint8_t x_65; uint8_t x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; uint8_t x_71; lean_object* x_72; lean_object* x_73; uint8_t x_74; -x_55 = lean_ctor_get(x_54, 0); -lean_inc(x_55); -x_56 = lean_ctor_get(x_54, 1); -lean_inc(x_56); -lean_dec(x_54); -x_57 = l_Lean_mkFreshFVarId___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___spec__1(x_19, x_20, x_21, x_22, x_23, x_24, x_25, x_56); -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_ctor_get(x_22, 2); -lean_inc(x_60); -x_61 = l_Lean_Expr_bindingName_x21(x_16); -x_62 = lean_ctor_get(x_55, 0); -lean_inc(x_62); -x_63 = lean_ctor_get(x_55, 1); -lean_inc(x_63); -lean_dec(x_55); -x_64 = l_Lean_Expr_bindingInfo_x21(x_16); -x_65 = lean_unbox(x_64); -lean_dec(x_64); -x_66 = 0; -lean_inc(x_62); -lean_inc(x_58); -x_67 = l_Lean_LocalContext_mkLocalDecl(x_60, x_58, x_61, x_62, x_65, x_66); -x_68 = l_Lean_Meta_getLocalInstances(x_22, x_23, x_24, x_25, x_59); -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 = 2; -x_72 = lean_unsigned_to_nat(0u); -lean_inc(x_53); -x_73 = l_Lean_Meta_mkFreshExprMVarAt(x_67, x_69, x_53, x_71, x_51, x_72, x_22, x_23, x_24, x_25, x_70); -x_74 = !lean_is_exclusive(x_73); -if (x_74 == 0) -{ -lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; -x_75 = lean_ctor_get(x_73, 0); -x_76 = lean_ctor_get(x_73, 1); -x_77 = l_Lean_Expr_mvarId_x21(x_75); -lean_inc(x_58); -x_78 = l_Lean_Expr_fvar___override(x_58); -x_79 = lean_box(0); -lean_ctor_set_tag(x_73, 1); -lean_ctor_set(x_73, 1, x_79); -lean_ctor_set(x_73, 0, x_78); -x_80 = lean_array_mk(x_73); -x_81 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__4___boxed), 10, 2); -lean_closure_set(x_81, 0, x_80); -lean_closure_set(x_81, 1, x_75); -x_82 = lean_box(x_7); -lean_inc(x_77); -x_83 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___boxed), 31, 22); -lean_closure_set(x_83, 0, x_63); -lean_closure_set(x_83, 1, x_1); -lean_closure_set(x_83, 2, x_77); -lean_closure_set(x_83, 3, x_2); -lean_closure_set(x_83, 4, x_3); -lean_closure_set(x_83, 5, x_4); -lean_closure_set(x_83, 6, x_5); -lean_closure_set(x_83, 7, x_6); -lean_closure_set(x_83, 8, x_82); -lean_closure_set(x_83, 9, x_8); -lean_closure_set(x_83, 10, x_9); -lean_closure_set(x_83, 11, x_10); -lean_closure_set(x_83, 12, x_11); -lean_closure_set(x_83, 13, x_12); -lean_closure_set(x_83, 14, x_13); -lean_closure_set(x_83, 15, x_14); -lean_closure_set(x_83, 16, x_15); -lean_closure_set(x_83, 17, x_58); -lean_closure_set(x_83, 18, x_79); -lean_closure_set(x_83, 19, x_53); -lean_closure_set(x_83, 20, x_62); -lean_closure_set(x_83, 21, x_17); -x_84 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); -lean_closure_set(x_84, 0, x_81); -lean_closure_set(x_84, 1, x_83); -x_85 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(x_77, x_84, x_19, x_20, x_21, x_22, x_23, x_24, x_25, x_76); -return x_85; +lean_object* x_105; lean_object* x_106; lean_object* x_107; +x_105 = lean_ctor_get(x_48, 0); +x_106 = lean_ctor_get(x_48, 1); +lean_inc(x_106); +lean_inc(x_105); +lean_dec(x_48); +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_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_86 = lean_ctor_get(x_73, 0); -x_87 = lean_ctor_get(x_73, 1); -lean_inc(x_87); -lean_inc(x_86); -lean_dec(x_73); -x_88 = l_Lean_Expr_mvarId_x21(x_86); -lean_inc(x_58); -x_89 = l_Lean_Expr_fvar___override(x_58); -x_90 = lean_box(0); -x_91 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_91, 0, x_89); -lean_ctor_set(x_91, 1, x_90); -x_92 = lean_array_mk(x_91); -x_93 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__4___boxed), 10, 2); -lean_closure_set(x_93, 0, x_92); -lean_closure_set(x_93, 1, x_86); -x_94 = lean_box(x_7); -lean_inc(x_88); -x_95 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___boxed), 31, 22); -lean_closure_set(x_95, 0, x_63); -lean_closure_set(x_95, 1, x_1); -lean_closure_set(x_95, 2, x_88); -lean_closure_set(x_95, 3, x_2); -lean_closure_set(x_95, 4, x_3); -lean_closure_set(x_95, 5, x_4); -lean_closure_set(x_95, 6, x_5); -lean_closure_set(x_95, 7, x_6); -lean_closure_set(x_95, 8, x_94); -lean_closure_set(x_95, 9, x_8); -lean_closure_set(x_95, 10, x_9); -lean_closure_set(x_95, 11, x_10); -lean_closure_set(x_95, 12, x_11); -lean_closure_set(x_95, 13, x_12); -lean_closure_set(x_95, 14, x_13); -lean_closure_set(x_95, 15, x_14); -lean_closure_set(x_95, 16, x_15); -lean_closure_set(x_95, 17, x_58); -lean_closure_set(x_95, 18, x_90); -lean_closure_set(x_95, 19, x_53); -lean_closure_set(x_95, 20, x_62); -lean_closure_set(x_95, 21, x_17); -x_96 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); -lean_closure_set(x_96, 0, x_93); -lean_closure_set(x_96, 1, x_95); -x_97 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(x_88, x_96, x_19, x_20, x_21, x_22, x_23, x_24, x_25, x_87); -return x_97; -} -} -else -{ -uint8_t x_98; -lean_dec(x_53); -lean_dec(x_51); -lean_dec(x_25); -lean_dec(x_24); +uint8_t x_108; +lean_dec(x_42); +lean_dec(x_33); +lean_dec(x_32); +lean_dec(x_31); +lean_dec(x_30); +lean_dec(x_29); +lean_dec(x_28); +lean_dec(x_27); lean_dec(x_23); lean_dec(x_22); -lean_dec(x_21); -lean_dec(x_20); -lean_dec(x_19); -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_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); lean_dec(x_1); -x_98 = !lean_is_exclusive(x_54); -if (x_98 == 0) +x_108 = !lean_is_exclusive(x_44); +if (x_108 == 0) { -return x_54; +return x_44; } else { -lean_object* x_99; lean_object* x_100; lean_object* x_101; -x_99 = lean_ctor_get(x_54, 0); -x_100 = lean_ctor_get(x_54, 1); -lean_inc(x_100); -lean_inc(x_99); -lean_dec(x_54); -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_object* x_109; lean_object* x_110; lean_object* x_111; +x_109 = lean_ctor_get(x_44, 0); +x_110 = lean_ctor_get(x_44, 1); +lean_inc(x_110); +lean_inc(x_109); +lean_dec(x_44); +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_102; -lean_dec(x_25); -lean_dec(x_24); +lean_object* x_117; lean_object* x_118; lean_object* x_119; +x_117 = lean_ctor_get(x_36, 1); +lean_inc(x_117); +lean_dec(x_36); +lean_inc(x_1); +x_118 = lean_alloc_ctor(0, 20, 1); +lean_ctor_set(x_118, 0, x_1); +lean_ctor_set(x_118, 1, x_2); +lean_ctor_set(x_118, 2, x_3); +lean_ctor_set(x_118, 3, x_4); +lean_ctor_set(x_118, 4, x_5); +lean_ctor_set(x_118, 5, x_6); +lean_ctor_set(x_118, 6, x_8); +lean_ctor_set(x_118, 7, x_9); +lean_ctor_set(x_118, 8, x_10); +lean_ctor_set(x_118, 9, x_11); +lean_ctor_set(x_118, 10, x_12); +lean_ctor_set(x_118, 11, x_13); +lean_ctor_set(x_118, 12, x_14); +lean_ctor_set(x_118, 13, x_15); +lean_ctor_set(x_118, 14, x_16); +lean_ctor_set(x_118, 15, x_17); +lean_ctor_set(x_118, 16, x_18); +lean_ctor_set(x_118, 17, x_19); +lean_ctor_set(x_118, 18, x_20); +lean_ctor_set(x_118, 19, x_21); +lean_ctor_set_uint8(x_118, sizeof(void*)*20, x_7); +if (x_24 == 0) +{ +uint8_t x_164; +x_164 = l_Lean_Expr_isLetFun(x_25); +if (x_164 == 0) +{ +lean_object* x_165; lean_object* x_166; +lean_dec(x_33); +lean_dec(x_32); +lean_dec(x_31); +lean_dec(x_30); +lean_dec(x_29); +lean_dec(x_28); +lean_dec(x_27); lean_dec(x_23); -lean_dec(x_22); -lean_dec(x_21); -lean_dec(x_20); -lean_dec(x_19); -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_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); lean_dec(x_1); -x_102 = !lean_is_exclusive(x_50); -if (x_102 == 0) +x_165 = lean_alloc_ctor(3, 2, 0); +lean_ctor_set(x_165, 0, x_22); +lean_ctor_set(x_165, 1, x_118); +x_166 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_166, 0, x_165); +lean_ctor_set(x_166, 1, x_117); +return x_166; +} +else { -return x_50; +lean_object* x_167; +x_167 = lean_box(0); +x_119 = x_167; +goto block_163; +} } else { -lean_object* x_103; lean_object* x_104; lean_object* x_105; -x_103 = lean_ctor_get(x_50, 0); -x_104 = lean_ctor_get(x_50, 1); -lean_inc(x_104); -lean_inc(x_103); -lean_dec(x_50); -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_168; +x_168 = lean_box(0); +x_119 = x_168; +goto block_163; } +block_163: +{ +lean_object* x_120; +lean_dec(x_119); +lean_inc(x_30); +lean_inc(x_22); +x_120 = l_Lean_FVarId_getDecl(x_22, x_30, x_31, x_32, x_33, x_117); +if (lean_obj_tag(x_120) == 0) +{ +lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; +x_121 = lean_ctor_get(x_120, 0); +lean_inc(x_121); +x_122 = lean_ctor_get(x_120, 1); +lean_inc(x_122); +lean_dec(x_120); +x_123 = l_Lean_LocalDecl_value(x_121); +lean_dec(x_121); +lean_inc(x_33); +lean_inc(x_32); +lean_inc(x_31); +lean_inc(x_30); +lean_inc(x_28); +x_124 = l_Lean_Meta_Grind_simp(x_123, x_27, x_28, x_29, x_30, x_31, x_32, x_33, x_122); +if (lean_obj_tag(x_124) == 0) +{ +lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; +x_125 = lean_ctor_get(x_124, 0); +lean_inc(x_125); +x_126 = lean_ctor_get(x_124, 1); +lean_inc(x_126); +lean_dec(x_124); +lean_inc(x_22); +x_127 = l_Lean_Expr_fvar___override(x_22); +x_128 = l_Lean_Meta_Grind_shareCommon(x_127, x_27, x_28, x_29, x_30, x_31, x_32, x_33, x_126); +x_129 = lean_ctor_get(x_128, 0); +lean_inc(x_129); +x_130 = lean_ctor_get(x_128, 1); +lean_inc(x_130); +if (lean_is_exclusive(x_128)) { + lean_ctor_release(x_128, 0); + lean_ctor_release(x_128, 1); + x_131 = x_128; +} else { + lean_dec_ref(x_128); + x_131 = lean_box(0); } +lean_inc(x_33); +lean_inc(x_32); +lean_inc(x_31); +lean_inc(x_30); +lean_inc(x_125); +x_132 = l_Lean_Meta_Simp_Result_getProof(x_125, x_30, x_31, x_32, x_33, x_130); +if (lean_obj_tag(x_132) == 0) +{ +lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; +x_133 = lean_ctor_get(x_132, 0); +lean_inc(x_133); +x_134 = lean_ctor_get(x_132, 1); +lean_inc(x_134); +lean_dec(x_132); +x_135 = lean_ctor_get(x_125, 0); +lean_inc(x_135); +lean_dec(x_125); +x_136 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__2___boxed), 9, 1); +lean_closure_set(x_136, 0, x_118); +x_137 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__3), 13, 4); +lean_closure_set(x_137, 0, x_129); +lean_closure_set(x_137, 1, x_135); +lean_closure_set(x_137, 2, x_133); +lean_closure_set(x_137, 3, x_23); +x_138 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); +lean_closure_set(x_138, 0, x_136); +lean_closure_set(x_138, 1, x_137); +x_139 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__1; +x_140 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); +lean_closure_set(x_140, 0, x_138); +lean_closure_set(x_140, 1, x_139); +x_141 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(x_1, x_140, x_27, x_28, x_29, x_30, x_31, x_32, x_33, x_134); +if (lean_obj_tag(x_141) == 0) +{ +lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; +x_142 = lean_ctor_get(x_141, 0); +lean_inc(x_142); +x_143 = lean_ctor_get(x_141, 1); +lean_inc(x_143); +if (lean_is_exclusive(x_141)) { + lean_ctor_release(x_141, 0); + lean_ctor_release(x_141, 1); + x_144 = x_141; +} else { + lean_dec_ref(x_141); + x_144 = lean_box(0); } +if (lean_is_scalar(x_131)) { + x_145 = lean_alloc_ctor(3, 2, 0); +} else { + x_145 = x_131; + lean_ctor_set_tag(x_145, 3); } +lean_ctor_set(x_145, 0, x_22); +lean_ctor_set(x_145, 1, x_142); +if (lean_is_scalar(x_144)) { + x_146 = lean_alloc_ctor(0, 2, 0); +} else { + x_146 = x_144; } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; 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_1, 0); -lean_inc(x_10); -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 = lean_ctor_get(x_1, 3); -lean_inc(x_13); -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_uint8(x_1, sizeof(void*)*14); -x_17 = lean_ctor_get(x_1, 6); -lean_inc(x_17); -x_18 = lean_ctor_get(x_1, 7); -lean_inc(x_18); -x_19 = lean_ctor_get(x_1, 8); -lean_inc(x_19); -x_20 = lean_ctor_get(x_1, 9); -lean_inc(x_20); -x_21 = lean_ctor_get(x_1, 10); -lean_inc(x_21); -x_22 = lean_ctor_get(x_1, 11); -lean_inc(x_22); -x_23 = lean_ctor_get(x_1, 12); -lean_inc(x_23); -x_24 = lean_ctor_get(x_1, 13); -lean_inc(x_24); -lean_dec(x_1); -lean_inc(x_10); -x_25 = l_Lean_MVarId_getType(x_10, x_5, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_25) == 0) -{ -uint8_t x_26; -x_26 = !lean_is_exclusive(x_25); -if (x_26 == 0) -{ -lean_object* x_27; lean_object* x_28; lean_object* x_29; uint8_t x_46; -x_27 = lean_ctor_get(x_25, 0); -x_28 = lean_ctor_get(x_25, 1); -x_46 = l_Lean_Expr_isArrow(x_27); -if (x_46 == 0) +lean_ctor_set(x_146, 0, x_145); +lean_ctor_set(x_146, 1, x_143); +return x_146; +} +else { -uint8_t x_47; -x_47 = l_Lean_Expr_isLet(x_27); -if (x_47 == 0) +lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; +lean_dec(x_131); +lean_dec(x_22); +x_147 = lean_ctor_get(x_141, 0); +lean_inc(x_147); +x_148 = lean_ctor_get(x_141, 1); +lean_inc(x_148); +if (lean_is_exclusive(x_141)) { + lean_ctor_release(x_141, 0); + lean_ctor_release(x_141, 1); + x_149 = x_141; +} else { + lean_dec_ref(x_141); + x_149 = lean_box(0); +} +if (lean_is_scalar(x_149)) { + x_150 = lean_alloc_ctor(1, 2, 0); +} else { + x_150 = x_149; +} +lean_ctor_set(x_150, 0, x_147); +lean_ctor_set(x_150, 1, x_148); +return x_150; +} +} +else { -uint8_t x_48; -x_48 = l_Lean_Expr_isForall(x_27); +lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; +lean_dec(x_131); +lean_dec(x_129); +lean_dec(x_125); +lean_dec(x_118); +lean_dec(x_33); +lean_dec(x_32); +lean_dec(x_31); +lean_dec(x_30); +lean_dec(x_29); +lean_dec(x_28); lean_dec(x_27); -if (x_48 == 0) -{ -lean_object* x_49; -lean_dec(x_24); lean_dec(x_23); lean_dec(x_22); -lean_dec(x_21); -lean_dec(x_20); -lean_dec(x_19); -lean_dec(x_18); -lean_dec(x_17); -lean_dec(x_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_3); -lean_dec(x_2); -x_49 = lean_box(0); -lean_ctor_set(x_25, 0, x_49); -return x_25; +lean_dec(x_1); +x_151 = lean_ctor_get(x_132, 0); +lean_inc(x_151); +x_152 = lean_ctor_get(x_132, 1); +lean_inc(x_152); +if (lean_is_exclusive(x_132)) { + lean_ctor_release(x_132, 0); + lean_ctor_release(x_132, 1); + x_153 = x_132; +} else { + lean_dec_ref(x_132); + x_153 = lean_box(0); } -else -{ -lean_object* x_50; -lean_free_object(x_25); -x_50 = lean_box(0); -x_29 = x_50; -goto block_45; +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_51; -lean_free_object(x_25); +lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; +lean_dec(x_118); +lean_dec(x_33); +lean_dec(x_32); +lean_dec(x_31); +lean_dec(x_30); +lean_dec(x_29); +lean_dec(x_28); lean_dec(x_27); -x_51 = lean_box(0); -x_29 = x_51; -goto block_45; +lean_dec(x_23); +lean_dec(x_22); +lean_dec(x_1); +x_155 = lean_ctor_get(x_124, 0); +lean_inc(x_155); +x_156 = lean_ctor_get(x_124, 1); +lean_inc(x_156); +if (lean_is_exclusive(x_124)) { + lean_ctor_release(x_124, 0); + lean_ctor_release(x_124, 1); + x_157 = x_124; +} else { + lean_dec_ref(x_124); + x_157 = lean_box(0); +} +if (lean_is_scalar(x_157)) { + x_158 = lean_alloc_ctor(1, 2, 0); +} else { + x_158 = x_157; +} +lean_ctor_set(x_158, 0, x_155); +lean_ctor_set(x_158, 1, x_156); +return x_158; } } 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_free_object(x_25); -x_52 = l_Lean_Expr_bindingDomain_x21(x_27); -lean_inc(x_52); -x_53 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__3___boxed), 9, 1); -lean_closure_set(x_53, 0, x_52); -x_54 = lean_box(x_16); -lean_inc(x_10); -x_55 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__6___boxed), 26, 17); -lean_closure_set(x_55, 0, x_10); -lean_closure_set(x_55, 1, x_11); -lean_closure_set(x_55, 2, x_12); -lean_closure_set(x_55, 3, x_13); -lean_closure_set(x_55, 4, x_14); -lean_closure_set(x_55, 5, x_15); -lean_closure_set(x_55, 6, x_54); -lean_closure_set(x_55, 7, x_17); -lean_closure_set(x_55, 8, x_18); -lean_closure_set(x_55, 9, x_19); -lean_closure_set(x_55, 10, x_20); -lean_closure_set(x_55, 11, x_21); -lean_closure_set(x_55, 12, x_22); -lean_closure_set(x_55, 13, x_23); -lean_closure_set(x_55, 14, x_24); -lean_closure_set(x_55, 15, x_27); -lean_closure_set(x_55, 16, x_52); -x_56 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); -lean_closure_set(x_56, 0, x_53); -lean_closure_set(x_56, 1, x_55); -x_57 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(x_10, x_56, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_28); -return x_57; +lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; +lean_dec(x_118); +lean_dec(x_33); +lean_dec(x_32); +lean_dec(x_31); +lean_dec(x_30); +lean_dec(x_29); +lean_dec(x_28); +lean_dec(x_27); +lean_dec(x_23); +lean_dec(x_22); +lean_dec(x_1); +x_159 = lean_ctor_get(x_120, 0); +lean_inc(x_159); +x_160 = lean_ctor_get(x_120, 1); +lean_inc(x_160); +if (lean_is_exclusive(x_120)) { + lean_ctor_release(x_120, 0); + lean_ctor_release(x_120, 1); + x_161 = x_120; +} else { + lean_dec_ref(x_120); + x_161 = lean_box(0); +} +if (lean_is_scalar(x_161)) { + x_162 = lean_alloc_ctor(1, 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; +} +} } -block_45: +} +else { -uint8_t x_30; lean_object* x_31; +lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_dec(x_29); -x_30 = 1; -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -x_31 = l_Lean_Meta_intro1Core(x_10, x_30, x_5, x_6, x_7, x_8, x_28); -if (lean_obj_tag(x_31) == 0) +lean_dec(x_28); +lean_dec(x_27); +lean_dec(x_23); +x_169 = lean_ctor_get(x_36, 1); +lean_inc(x_169); +lean_dec(x_36); +x_170 = l_Lean_LocalDecl_userName(x_26); +x_171 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_170, x_32, x_33, x_169); +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_Expr_fvar___override(x_22); +x_175 = l_Lean_MVarId_assert(x_1, x_172, x_35, x_174, x_30, x_31, x_32, x_33, x_173); +if (lean_obj_tag(x_175) == 0) +{ +uint8_t x_176; +x_176 = !lean_is_exclusive(x_175); +if (x_176 == 0) +{ +lean_object* x_177; lean_object* x_178; lean_object* x_179; +x_177 = lean_ctor_get(x_175, 0); +x_178 = lean_alloc_ctor(0, 20, 1); +lean_ctor_set(x_178, 0, x_177); +lean_ctor_set(x_178, 1, x_2); +lean_ctor_set(x_178, 2, x_3); +lean_ctor_set(x_178, 3, x_4); +lean_ctor_set(x_178, 4, x_5); +lean_ctor_set(x_178, 5, x_6); +lean_ctor_set(x_178, 6, x_8); +lean_ctor_set(x_178, 7, x_9); +lean_ctor_set(x_178, 8, x_10); +lean_ctor_set(x_178, 9, x_11); +lean_ctor_set(x_178, 10, x_12); +lean_ctor_set(x_178, 11, x_13); +lean_ctor_set(x_178, 12, x_14); +lean_ctor_set(x_178, 13, x_15); +lean_ctor_set(x_178, 14, x_16); +lean_ctor_set(x_178, 15, x_17); +lean_ctor_set(x_178, 16, x_18); +lean_ctor_set(x_178, 17, x_19); +lean_ctor_set(x_178, 18, x_20); +lean_ctor_set(x_178, 19, x_21); +lean_ctor_set_uint8(x_178, sizeof(void*)*20, x_7); +x_179 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_179, 0, x_178); +lean_ctor_set(x_175, 0, x_179); +return x_175; +} +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; -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_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); -lean_inc(x_34); -x_36 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__1___boxed), 9, 1); -lean_closure_set(x_36, 0, x_34); -x_37 = lean_box(x_16); -lean_inc(x_35); -x_38 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__2___boxed), 25, 16); -lean_closure_set(x_38, 0, x_35); -lean_closure_set(x_38, 1, x_11); -lean_closure_set(x_38, 2, x_12); -lean_closure_set(x_38, 3, x_13); -lean_closure_set(x_38, 4, x_14); -lean_closure_set(x_38, 5, x_15); -lean_closure_set(x_38, 6, x_37); -lean_closure_set(x_38, 7, x_17); -lean_closure_set(x_38, 8, x_18); -lean_closure_set(x_38, 9, x_19); -lean_closure_set(x_38, 10, x_20); -lean_closure_set(x_38, 11, x_21); -lean_closure_set(x_38, 12, x_22); -lean_closure_set(x_38, 13, x_23); -lean_closure_set(x_38, 14, x_24); -lean_closure_set(x_38, 15, x_34); -x_39 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); -lean_closure_set(x_39, 0, x_36); -lean_closure_set(x_39, 1, x_38); -x_40 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(x_35, x_39, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_33); -return x_40; +lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; +x_180 = lean_ctor_get(x_175, 0); +x_181 = lean_ctor_get(x_175, 1); +lean_inc(x_181); +lean_inc(x_180); +lean_dec(x_175); +x_182 = lean_alloc_ctor(0, 20, 1); +lean_ctor_set(x_182, 0, x_180); +lean_ctor_set(x_182, 1, x_2); +lean_ctor_set(x_182, 2, x_3); +lean_ctor_set(x_182, 3, x_4); +lean_ctor_set(x_182, 4, x_5); +lean_ctor_set(x_182, 5, x_6); +lean_ctor_set(x_182, 6, x_8); +lean_ctor_set(x_182, 7, x_9); +lean_ctor_set(x_182, 8, x_10); +lean_ctor_set(x_182, 9, x_11); +lean_ctor_set(x_182, 10, x_12); +lean_ctor_set(x_182, 11, x_13); +lean_ctor_set(x_182, 12, x_14); +lean_ctor_set(x_182, 13, x_15); +lean_ctor_set(x_182, 14, x_16); +lean_ctor_set(x_182, 15, x_17); +lean_ctor_set(x_182, 16, x_18); +lean_ctor_set(x_182, 17, x_19); +lean_ctor_set(x_182, 18, x_20); +lean_ctor_set(x_182, 19, x_21); +lean_ctor_set_uint8(x_182, sizeof(void*)*20, x_7); +x_183 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_183, 0, x_182); +x_184 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_184, 0, x_183); +lean_ctor_set(x_184, 1, x_181); +return x_184; +} } else { -uint8_t x_41; -lean_dec(x_24); -lean_dec(x_23); -lean_dec(x_22); +uint8_t x_185; lean_dec(x_21); lean_dec(x_20); lean_dec(x_19); lean_dec(x_18); lean_dec(x_17); +lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); lean_dec(x_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_41 = !lean_is_exclusive(x_31); -if (x_41 == 0) +x_185 = !lean_is_exclusive(x_175); +if (x_185 == 0) { -return x_31; +return x_175; } else { -lean_object* x_42; lean_object* x_43; lean_object* x_44; -x_42 = lean_ctor_get(x_31, 0); -x_43 = lean_ctor_get(x_31, 1); -lean_inc(x_43); -lean_inc(x_42); -lean_dec(x_31); -x_44 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_44, 0, x_42); -lean_ctor_set(x_44, 1, x_43); -return x_44; +lean_object* x_186; lean_object* x_187; lean_object* x_188; +x_186 = lean_ctor_get(x_175, 0); +x_187 = lean_ctor_get(x_175, 1); +lean_inc(x_187); +lean_inc(x_186); +lean_dec(x_175); +x_188 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_188, 0, x_186); +lean_ctor_set(x_188, 1, x_187); +return x_188; } } } } else { -lean_object* x_58; lean_object* x_59; lean_object* x_60; uint8_t x_77; -x_58 = lean_ctor_get(x_25, 0); -x_59 = lean_ctor_get(x_25, 1); -lean_inc(x_59); -lean_inc(x_58); -lean_dec(x_25); -x_77 = l_Lean_Expr_isArrow(x_58); -if (x_77 == 0) -{ -uint8_t x_78; -x_78 = l_Lean_Expr_isLet(x_58); -if (x_78 == 0) -{ -uint8_t x_79; -x_79 = l_Lean_Expr_isForall(x_58); -lean_dec(x_58); -if (x_79 == 0) -{ -lean_object* x_80; lean_object* x_81; -lean_dec(x_24); +uint8_t x_189; +lean_dec(x_35); +lean_dec(x_33); +lean_dec(x_32); +lean_dec(x_31); +lean_dec(x_30); +lean_dec(x_29); +lean_dec(x_28); +lean_dec(x_27); lean_dec(x_23); lean_dec(x_22); lean_dec(x_21); @@ -1688,261 +1657,835 @@ lean_dec(x_20); lean_dec(x_19); lean_dec(x_18); lean_dec(x_17); +lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); lean_dec(x_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_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_59); -return x_81; +lean_dec(x_1); +x_189 = !lean_is_exclusive(x_36); +if (x_189 == 0) +{ +return x_36; } else { -lean_object* x_82; -x_82 = lean_box(0); -x_60 = x_82; -goto block_76; +lean_object* x_190; lean_object* x_191; lean_object* x_192; +x_190 = lean_ctor_get(x_36, 0); +x_191 = lean_ctor_get(x_36, 1); +lean_inc(x_191); +lean_inc(x_190); +lean_dec(x_36); +x_192 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_192, 0, x_190); +lean_ctor_set(x_192, 1, x_191); +return x_192; } } -else +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___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_83; -lean_dec(x_58); -x_83 = lean_box(0); -x_60 = x_83; -goto block_76; +lean_object* x_10; +x_10 = l_Lean_Meta_isProp(x_1, x_5, x_6, x_7, x_8, x_9); +return x_10; } } -else +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___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) { +_start: { -lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; -x_84 = l_Lean_Expr_bindingDomain_x21(x_58); -lean_inc(x_84); -x_85 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__3___boxed), 9, 1); -lean_closure_set(x_85, 0, x_84); -x_86 = lean_box(x_16); -lean_inc(x_10); -x_87 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__6___boxed), 26, 17); -lean_closure_set(x_87, 0, x_10); -lean_closure_set(x_87, 1, x_11); -lean_closure_set(x_87, 2, x_12); -lean_closure_set(x_87, 3, x_13); -lean_closure_set(x_87, 4, x_14); -lean_closure_set(x_87, 5, x_15); -lean_closure_set(x_87, 6, x_86); -lean_closure_set(x_87, 7, x_17); -lean_closure_set(x_87, 8, x_18); -lean_closure_set(x_87, 9, x_19); -lean_closure_set(x_87, 10, x_20); -lean_closure_set(x_87, 11, x_21); -lean_closure_set(x_87, 12, x_22); -lean_closure_set(x_87, 13, x_23); -lean_closure_set(x_87, 14, x_24); -lean_closure_set(x_87, 15, x_58); -lean_closure_set(x_87, 16, x_84); -x_88 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); -lean_closure_set(x_88, 0, x_85); -lean_closure_set(x_88, 1, x_87); -x_89 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(x_10, x_88, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_59); -return x_89; +uint8_t x_11; uint8_t x_12; uint8_t x_13; lean_object* x_14; +x_11 = 0; +x_12 = 1; +x_13 = 1; +x_14 = l_Lean_Meta_mkLambdaFVars(x_1, x_2, x_11, x_12, x_11, x_13, x_6, x_7, x_8, x_9, x_10); +return x_14; +} } -block_76: +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__8___closed__1() { +_start: { -uint8_t x_61; lean_object* x_62; -lean_dec(x_60); -x_61 = 1; -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -x_62 = l_Lean_Meta_intro1Core(x_10, x_61, x_5, x_6, x_7, x_8, x_59); -if (lean_obj_tag(x_62) == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__8___closed__2() { +_start: { -lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; -x_63 = lean_ctor_get(x_62, 0); -lean_inc(x_63); -x_64 = lean_ctor_get(x_62, 1); -lean_inc(x_64); -lean_dec(x_62); -x_65 = lean_ctor_get(x_63, 0); -lean_inc(x_65); -x_66 = lean_ctor_get(x_63, 1); -lean_inc(x_66); -lean_dec(x_63); -lean_inc(x_65); -x_67 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__1___boxed), 9, 1); -lean_closure_set(x_67, 0, x_65); -x_68 = lean_box(x_16); -lean_inc(x_66); -x_69 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__2___boxed), 25, 16); -lean_closure_set(x_69, 0, x_66); -lean_closure_set(x_69, 1, x_11); -lean_closure_set(x_69, 2, x_12); -lean_closure_set(x_69, 3, x_13); -lean_closure_set(x_69, 4, x_14); -lean_closure_set(x_69, 5, x_15); -lean_closure_set(x_69, 6, x_68); -lean_closure_set(x_69, 7, x_17); -lean_closure_set(x_69, 8, x_18); -lean_closure_set(x_69, 9, x_19); -lean_closure_set(x_69, 10, x_20); -lean_closure_set(x_69, 11, x_21); -lean_closure_set(x_69, 12, x_22); -lean_closure_set(x_69, 13, x_23); -lean_closure_set(x_69, 14, x_24); -lean_closure_set(x_69, 15, x_65); -x_70 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); -lean_closure_set(x_70, 0, x_67); -lean_closure_set(x_70, 1, x_69); -x_71 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(x_66, x_70, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_64); -return x_71; -} -else -{ -lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; -lean_dec(x_24); +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Grind", 5, 5); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__8___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("intro_with_eq", 13, 13); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__8___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__8___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__8___closed__2; +x_3 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__8___closed__3; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___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, 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) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_38; uint8_t x_39; +lean_dec(x_28); +lean_dec(x_27); +lean_dec(x_26); +lean_dec(x_25); +x_38 = l_Lean_MVarId_assign___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___spec__3(x_2, x_29, x_30, x_31, x_32, x_33, x_34, x_35, x_36, x_37); +x_39 = !lean_is_exclusive(x_38); +if (x_39 == 0) +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_40 = lean_ctor_get(x_38, 0); +lean_dec(x_40); +x_41 = lean_alloc_ctor(0, 20, 1); +lean_ctor_set(x_41, 0, x_3); +lean_ctor_set(x_41, 1, x_4); +lean_ctor_set(x_41, 2, x_5); +lean_ctor_set(x_41, 3, x_6); +lean_ctor_set(x_41, 4, x_7); +lean_ctor_set(x_41, 5, x_8); +lean_ctor_set(x_41, 6, x_10); +lean_ctor_set(x_41, 7, x_11); +lean_ctor_set(x_41, 8, x_12); +lean_ctor_set(x_41, 9, x_13); +lean_ctor_set(x_41, 10, x_14); +lean_ctor_set(x_41, 11, x_15); +lean_ctor_set(x_41, 12, x_16); +lean_ctor_set(x_41, 13, x_17); +lean_ctor_set(x_41, 14, x_18); +lean_ctor_set(x_41, 15, x_19); +lean_ctor_set(x_41, 16, x_20); +lean_ctor_set(x_41, 17, x_21); +lean_ctor_set(x_41, 18, x_22); +lean_ctor_set(x_41, 19, x_23); +lean_ctor_set_uint8(x_41, sizeof(void*)*20, x_9); +x_42 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_42, 0, x_24); +lean_ctor_set(x_42, 1, x_41); +lean_ctor_set(x_38, 0, x_42); +return x_38; +} +else +{ +lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_43 = lean_ctor_get(x_38, 1); +lean_inc(x_43); +lean_dec(x_38); +x_44 = lean_alloc_ctor(0, 20, 1); +lean_ctor_set(x_44, 0, x_3); +lean_ctor_set(x_44, 1, x_4); +lean_ctor_set(x_44, 2, x_5); +lean_ctor_set(x_44, 3, x_6); +lean_ctor_set(x_44, 4, x_7); +lean_ctor_set(x_44, 5, x_8); +lean_ctor_set(x_44, 6, x_10); +lean_ctor_set(x_44, 7, x_11); +lean_ctor_set(x_44, 8, x_12); +lean_ctor_set(x_44, 9, x_13); +lean_ctor_set(x_44, 10, x_14); +lean_ctor_set(x_44, 11, x_15); +lean_ctor_set(x_44, 12, x_16); +lean_ctor_set(x_44, 13, x_17); +lean_ctor_set(x_44, 14, x_18); +lean_ctor_set(x_44, 15, x_19); +lean_ctor_set(x_44, 16, x_20); +lean_ctor_set(x_44, 17, x_21); +lean_ctor_set(x_44, 18, x_22); +lean_ctor_set(x_44, 19, x_23); +lean_ctor_set_uint8(x_44, sizeof(void*)*20, x_9); +x_45 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_45, 0, x_24); +lean_ctor_set(x_45, 1, x_44); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_43); +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_object* x_56; lean_object* x_57; uint8_t x_58; +x_47 = lean_ctor_get(x_1, 0); +x_48 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__8___closed__4; +lean_inc(x_25); +x_49 = l_Lean_Expr_const___override(x_48, x_25); +x_50 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_50, 0, x_29); +lean_ctor_set(x_50, 1, x_25); +lean_inc(x_47); +x_51 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_51, 0, x_47); +lean_ctor_set(x_51, 1, x_50); +x_52 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_52, 0, x_26); +lean_ctor_set(x_52, 1, x_51); +x_53 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_53, 0, x_27); +lean_ctor_set(x_53, 1, x_52); +x_54 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_54, 0, x_28); +lean_ctor_set(x_54, 1, x_53); +x_55 = lean_array_mk(x_54); +x_56 = l_Lean_mkAppN(x_49, x_55); +lean_dec(x_55); +x_57 = l_Lean_MVarId_assign___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___spec__3(x_2, x_56, x_30, x_31, x_32, x_33, x_34, x_35, x_36, x_37); +x_58 = !lean_is_exclusive(x_57); +if (x_58 == 0) +{ +lean_object* x_59; lean_object* x_60; lean_object* x_61; +x_59 = lean_ctor_get(x_57, 0); +lean_dec(x_59); +x_60 = lean_alloc_ctor(0, 20, 1); +lean_ctor_set(x_60, 0, x_3); +lean_ctor_set(x_60, 1, x_4); +lean_ctor_set(x_60, 2, x_5); +lean_ctor_set(x_60, 3, x_6); +lean_ctor_set(x_60, 4, x_7); +lean_ctor_set(x_60, 5, x_8); +lean_ctor_set(x_60, 6, x_10); +lean_ctor_set(x_60, 7, x_11); +lean_ctor_set(x_60, 8, x_12); +lean_ctor_set(x_60, 9, x_13); +lean_ctor_set(x_60, 10, x_14); +lean_ctor_set(x_60, 11, x_15); +lean_ctor_set(x_60, 12, x_16); +lean_ctor_set(x_60, 13, x_17); +lean_ctor_set(x_60, 14, x_18); +lean_ctor_set(x_60, 15, x_19); +lean_ctor_set(x_60, 16, x_20); +lean_ctor_set(x_60, 17, x_21); +lean_ctor_set(x_60, 18, x_22); +lean_ctor_set(x_60, 19, x_23); +lean_ctor_set_uint8(x_60, sizeof(void*)*20, x_9); +x_61 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_61, 0, x_24); +lean_ctor_set(x_61, 1, x_60); +lean_ctor_set(x_57, 0, x_61); +return x_57; +} +else +{ +lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; +x_62 = lean_ctor_get(x_57, 1); +lean_inc(x_62); +lean_dec(x_57); +x_63 = lean_alloc_ctor(0, 20, 1); +lean_ctor_set(x_63, 0, x_3); +lean_ctor_set(x_63, 1, x_4); +lean_ctor_set(x_63, 2, x_5); +lean_ctor_set(x_63, 3, x_6); +lean_ctor_set(x_63, 4, x_7); +lean_ctor_set(x_63, 5, x_8); +lean_ctor_set(x_63, 6, x_10); +lean_ctor_set(x_63, 7, x_11); +lean_ctor_set(x_63, 8, x_12); +lean_ctor_set(x_63, 9, x_13); +lean_ctor_set(x_63, 10, x_14); +lean_ctor_set(x_63, 11, x_15); +lean_ctor_set(x_63, 12, x_16); +lean_ctor_set(x_63, 13, x_17); +lean_ctor_set(x_63, 14, x_18); +lean_ctor_set(x_63, 15, x_19); +lean_ctor_set(x_63, 16, x_20); +lean_ctor_set(x_63, 17, x_21); +lean_ctor_set(x_63, 18, x_22); +lean_ctor_set(x_63, 19, x_23); +lean_ctor_set_uint8(x_63, sizeof(void*)*20, x_9); +x_64 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_64, 0, x_24); +lean_ctor_set(x_64, 1, x_63); +x_65 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_65, 0, x_64); +lean_ctor_set(x_65, 1, x_62); +return x_65; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___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, 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, uint8_t x_24, lean_object* x_25, lean_object* x_26, lean_object* x_27, lean_object* x_28, lean_object* x_29, lean_object* x_30, lean_object* x_31, lean_object* x_32) { +_start: +{ +if (x_24 == 0) +{ +uint8_t x_33; lean_object* x_34; +lean_dec(x_27); +lean_dec(x_26); +lean_dec(x_25); lean_dec(x_23); -lean_dec(x_22); +x_33 = 1; +x_34 = l_Lean_Meta_intro1Core(x_1, x_33, x_28, x_29, x_30, x_31, x_32); +if (lean_obj_tag(x_34) == 0) +{ +uint8_t x_35; +x_35 = !lean_is_exclusive(x_34); +if (x_35 == 0) +{ +lean_object* x_36; uint8_t x_37; +x_36 = lean_ctor_get(x_34, 0); +x_37 = !lean_is_exclusive(x_36); +if (x_37 == 0) +{ +lean_object* x_38; lean_object* x_39; +x_38 = lean_ctor_get(x_36, 1); +x_39 = lean_alloc_ctor(0, 20, 1); +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_2); +lean_ctor_set(x_39, 2, x_3); +lean_ctor_set(x_39, 3, x_4); +lean_ctor_set(x_39, 4, x_5); +lean_ctor_set(x_39, 5, x_6); +lean_ctor_set(x_39, 6, x_8); +lean_ctor_set(x_39, 7, x_9); +lean_ctor_set(x_39, 8, x_10); +lean_ctor_set(x_39, 9, x_11); +lean_ctor_set(x_39, 10, x_12); +lean_ctor_set(x_39, 11, x_13); +lean_ctor_set(x_39, 12, x_14); +lean_ctor_set(x_39, 13, x_15); +lean_ctor_set(x_39, 14, x_16); +lean_ctor_set(x_39, 15, x_17); +lean_ctor_set(x_39, 16, x_18); +lean_ctor_set(x_39, 17, x_19); +lean_ctor_set(x_39, 18, x_20); +lean_ctor_set(x_39, 19, x_21); +lean_ctor_set_uint8(x_39, sizeof(void*)*20, x_7); +lean_ctor_set_tag(x_36, 3); +lean_ctor_set(x_36, 1, x_39); +return x_34; +} +else +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_40 = lean_ctor_get(x_36, 0); +x_41 = lean_ctor_get(x_36, 1); +lean_inc(x_41); +lean_inc(x_40); +lean_dec(x_36); +x_42 = lean_alloc_ctor(0, 20, 1); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_2); +lean_ctor_set(x_42, 2, x_3); +lean_ctor_set(x_42, 3, x_4); +lean_ctor_set(x_42, 4, x_5); +lean_ctor_set(x_42, 5, x_6); +lean_ctor_set(x_42, 6, x_8); +lean_ctor_set(x_42, 7, x_9); +lean_ctor_set(x_42, 8, x_10); +lean_ctor_set(x_42, 9, x_11); +lean_ctor_set(x_42, 10, x_12); +lean_ctor_set(x_42, 11, x_13); +lean_ctor_set(x_42, 12, x_14); +lean_ctor_set(x_42, 13, x_15); +lean_ctor_set(x_42, 14, x_16); +lean_ctor_set(x_42, 15, x_17); +lean_ctor_set(x_42, 16, x_18); +lean_ctor_set(x_42, 17, x_19); +lean_ctor_set(x_42, 18, x_20); +lean_ctor_set(x_42, 19, x_21); +lean_ctor_set_uint8(x_42, sizeof(void*)*20, x_7); +x_43 = lean_alloc_ctor(3, 2, 0); +lean_ctor_set(x_43, 0, x_40); +lean_ctor_set(x_43, 1, x_42); +lean_ctor_set(x_34, 0, x_43); +return x_34; +} +} +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; +x_44 = lean_ctor_get(x_34, 0); +x_45 = lean_ctor_get(x_34, 1); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_34); +x_46 = lean_ctor_get(x_44, 0); +lean_inc(x_46); +x_47 = lean_ctor_get(x_44, 1); +lean_inc(x_47); +if (lean_is_exclusive(x_44)) { + lean_ctor_release(x_44, 0); + lean_ctor_release(x_44, 1); + x_48 = x_44; +} else { + lean_dec_ref(x_44); + x_48 = lean_box(0); +} +x_49 = lean_alloc_ctor(0, 20, 1); +lean_ctor_set(x_49, 0, x_47); +lean_ctor_set(x_49, 1, x_2); +lean_ctor_set(x_49, 2, x_3); +lean_ctor_set(x_49, 3, x_4); +lean_ctor_set(x_49, 4, x_5); +lean_ctor_set(x_49, 5, x_6); +lean_ctor_set(x_49, 6, x_8); +lean_ctor_set(x_49, 7, x_9); +lean_ctor_set(x_49, 8, x_10); +lean_ctor_set(x_49, 9, x_11); +lean_ctor_set(x_49, 10, x_12); +lean_ctor_set(x_49, 11, x_13); +lean_ctor_set(x_49, 12, x_14); +lean_ctor_set(x_49, 13, x_15); +lean_ctor_set(x_49, 14, x_16); +lean_ctor_set(x_49, 15, x_17); +lean_ctor_set(x_49, 16, x_18); +lean_ctor_set(x_49, 17, x_19); +lean_ctor_set(x_49, 18, x_20); +lean_ctor_set(x_49, 19, x_21); +lean_ctor_set_uint8(x_49, sizeof(void*)*20, x_7); +if (lean_is_scalar(x_48)) { + x_50 = lean_alloc_ctor(3, 2, 0); +} else { + x_50 = x_48; + lean_ctor_set_tag(x_50, 3); +} +lean_ctor_set(x_50, 0, x_46); +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_45); +return x_51; +} +} +else +{ +uint8_t x_52; lean_dec(x_21); lean_dec(x_20); lean_dec(x_19); lean_dec(x_18); lean_dec(x_17); +lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); lean_dec(x_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_72 = lean_ctor_get(x_62, 0); -lean_inc(x_72); -x_73 = lean_ctor_get(x_62, 1); -lean_inc(x_73); -if (lean_is_exclusive(x_62)) { - lean_ctor_release(x_62, 0); - lean_ctor_release(x_62, 1); - x_74 = x_62; -} else { - lean_dec_ref(x_62); - x_74 = lean_box(0); +x_52 = !lean_is_exclusive(x_34); +if (x_52 == 0) +{ +return x_34; } -if (lean_is_scalar(x_74)) { - x_75 = lean_alloc_ctor(1, 2, 0); -} else { - x_75 = x_74; +else +{ +lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_53 = lean_ctor_get(x_34, 0); +x_54 = lean_ctor_get(x_34, 1); +lean_inc(x_54); +lean_inc(x_53); +lean_dec(x_34); +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_ctor_set(x_75, 0, x_72); -lean_ctor_set(x_75, 1, x_73); -return x_75; } } +else +{ +lean_object* x_56; +lean_inc(x_1); +x_56 = l_Lean_MVarId_getTag(x_1, x_28, x_29, x_30, x_31, x_32); +if (lean_obj_tag(x_56) == 0) +{ +lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_57 = lean_ctor_get(x_56, 0); +lean_inc(x_57); +x_58 = lean_ctor_get(x_56, 1); +lean_inc(x_58); +lean_dec(x_56); +x_59 = l_Lean_Expr_bindingBody_x21(x_22); +lean_inc(x_31); +lean_inc(x_30); +lean_inc(x_29); +lean_inc(x_28); +lean_inc(x_26); +lean_inc(x_23); +x_60 = l_Lean_Meta_Grind_simp(x_23, x_25, x_26, x_27, x_28, x_29, x_30, x_31, x_58); +if (lean_obj_tag(x_60) == 0) +{ +lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; 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; lean_object* x_74; lean_object* x_75; lean_object* x_76; uint8_t x_77; lean_object* x_78; lean_object* x_79; uint8_t x_80; +x_61 = lean_ctor_get(x_60, 0); +lean_inc(x_61); +x_62 = lean_ctor_get(x_60, 1); +lean_inc(x_62); +lean_dec(x_60); +x_63 = l_Lean_mkFreshFVarId___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___spec__1(x_25, x_26, x_27, x_28, x_29, x_30, x_31, x_62); +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_28, 2); +lean_inc(x_66); +x_67 = l_Lean_Expr_bindingName_x21(x_22); +x_68 = lean_ctor_get(x_61, 0); +lean_inc(x_68); +x_69 = lean_ctor_get(x_61, 1); +lean_inc(x_69); +lean_dec(x_61); +x_70 = l_Lean_Expr_bindingInfo_x21(x_22); +x_71 = lean_unbox(x_70); +lean_dec(x_70); +x_72 = 0; +lean_inc(x_68); +lean_inc(x_64); +x_73 = l_Lean_LocalContext_mkLocalDecl(x_66, x_64, x_67, x_68, x_71, x_72); +x_74 = l_Lean_Meta_getLocalInstances(x_28, x_29, x_30, x_31, x_65); +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 = 2; +x_78 = lean_unsigned_to_nat(0u); +lean_inc(x_59); +x_79 = l_Lean_Meta_mkFreshExprMVarAt(x_73, x_75, x_59, x_77, x_57, x_78, x_28, x_29, x_30, x_31, x_76); +x_80 = !lean_is_exclusive(x_79); +if (x_80 == 0) +{ +lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; +x_81 = lean_ctor_get(x_79, 0); +x_82 = lean_ctor_get(x_79, 1); +x_83 = l_Lean_Expr_mvarId_x21(x_81); +lean_inc(x_64); +x_84 = l_Lean_Expr_fvar___override(x_64); +x_85 = lean_box(0); +lean_ctor_set_tag(x_79, 1); +lean_ctor_set(x_79, 1, x_85); +lean_ctor_set(x_79, 0, x_84); +x_86 = lean_array_mk(x_79); +x_87 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__7___boxed), 10, 2); +lean_closure_set(x_87, 0, x_86); +lean_closure_set(x_87, 1, x_81); +x_88 = lean_box(x_7); +lean_inc(x_83); +x_89 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__8___boxed), 37, 28); +lean_closure_set(x_89, 0, x_69); +lean_closure_set(x_89, 1, x_1); +lean_closure_set(x_89, 2, x_83); +lean_closure_set(x_89, 3, x_2); +lean_closure_set(x_89, 4, x_3); +lean_closure_set(x_89, 5, x_4); +lean_closure_set(x_89, 6, x_5); +lean_closure_set(x_89, 7, x_6); +lean_closure_set(x_89, 8, x_88); +lean_closure_set(x_89, 9, x_8); +lean_closure_set(x_89, 10, x_9); +lean_closure_set(x_89, 11, x_10); +lean_closure_set(x_89, 12, x_11); +lean_closure_set(x_89, 13, x_12); +lean_closure_set(x_89, 14, x_13); +lean_closure_set(x_89, 15, x_14); +lean_closure_set(x_89, 16, x_15); +lean_closure_set(x_89, 17, x_16); +lean_closure_set(x_89, 18, x_17); +lean_closure_set(x_89, 19, x_18); +lean_closure_set(x_89, 20, x_19); +lean_closure_set(x_89, 21, x_20); +lean_closure_set(x_89, 22, x_21); +lean_closure_set(x_89, 23, x_64); +lean_closure_set(x_89, 24, x_85); +lean_closure_set(x_89, 25, x_59); +lean_closure_set(x_89, 26, x_68); +lean_closure_set(x_89, 27, x_23); +x_90 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); +lean_closure_set(x_90, 0, x_87); +lean_closure_set(x_90, 1, x_89); +x_91 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(x_83, x_90, x_25, x_26, x_27, x_28, x_29, x_30, x_31, x_82); +return x_91; +} +else +{ +lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; +x_92 = lean_ctor_get(x_79, 0); +x_93 = lean_ctor_get(x_79, 1); +lean_inc(x_93); +lean_inc(x_92); +lean_dec(x_79); +x_94 = l_Lean_Expr_mvarId_x21(x_92); +lean_inc(x_64); +x_95 = l_Lean_Expr_fvar___override(x_64); +x_96 = lean_box(0); +x_97 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_97, 0, x_95); +lean_ctor_set(x_97, 1, x_96); +x_98 = lean_array_mk(x_97); +x_99 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__7___boxed), 10, 2); +lean_closure_set(x_99, 0, x_98); +lean_closure_set(x_99, 1, x_92); +x_100 = lean_box(x_7); +lean_inc(x_94); +x_101 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__8___boxed), 37, 28); +lean_closure_set(x_101, 0, x_69); +lean_closure_set(x_101, 1, x_1); +lean_closure_set(x_101, 2, x_94); +lean_closure_set(x_101, 3, x_2); +lean_closure_set(x_101, 4, x_3); +lean_closure_set(x_101, 5, x_4); +lean_closure_set(x_101, 6, x_5); +lean_closure_set(x_101, 7, x_6); +lean_closure_set(x_101, 8, x_100); +lean_closure_set(x_101, 9, x_8); +lean_closure_set(x_101, 10, x_9); +lean_closure_set(x_101, 11, x_10); +lean_closure_set(x_101, 12, x_11); +lean_closure_set(x_101, 13, x_12); +lean_closure_set(x_101, 14, x_13); +lean_closure_set(x_101, 15, x_14); +lean_closure_set(x_101, 16, x_15); +lean_closure_set(x_101, 17, x_16); +lean_closure_set(x_101, 18, x_17); +lean_closure_set(x_101, 19, x_18); +lean_closure_set(x_101, 20, x_19); +lean_closure_set(x_101, 21, x_20); +lean_closure_set(x_101, 22, x_21); +lean_closure_set(x_101, 23, x_64); +lean_closure_set(x_101, 24, x_96); +lean_closure_set(x_101, 25, x_59); +lean_closure_set(x_101, 26, x_68); +lean_closure_set(x_101, 27, x_23); +x_102 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); +lean_closure_set(x_102, 0, x_99); +lean_closure_set(x_102, 1, x_101); +x_103 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(x_94, x_102, x_25, x_26, x_27, x_28, x_29, x_30, x_31, x_93); +return x_103; } } else { -uint8_t x_90; -lean_dec(x_24); +uint8_t x_104; +lean_dec(x_59); +lean_dec(x_57); +lean_dec(x_31); +lean_dec(x_30); +lean_dec(x_29); +lean_dec(x_28); +lean_dec(x_27); +lean_dec(x_26); +lean_dec(x_25); lean_dec(x_23); -lean_dec(x_22); lean_dec(x_21); lean_dec(x_20); lean_dec(x_19); lean_dec(x_18); lean_dec(x_17); +lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); +lean_dec(x_9); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_90 = !lean_is_exclusive(x_25); -if (x_90 == 0) +lean_dec(x_1); +x_104 = !lean_is_exclusive(x_60); +if (x_104 == 0) { -return x_25; +return x_60; } else { -lean_object* x_91; lean_object* x_92; lean_object* x_93; -x_91 = lean_ctor_get(x_25, 0); -x_92 = lean_ctor_get(x_25, 1); -lean_inc(x_92); -lean_inc(x_91); -lean_dec(x_25); -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_105; lean_object* x_106; lean_object* x_107; +x_105 = lean_ctor_get(x_60, 0); +x_106 = lean_ctor_get(x_60, 1); +lean_inc(x_106); +lean_inc(x_105); +lean_dec(x_60); +x_107 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_107, 0, x_105); +lean_ctor_set(x_107, 1, x_106); +return x_107; } -LEAN_EXPORT lean_object* l_Lean_mkFreshId___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___spec__2___rarg___boxed(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; -x_3 = l_Lean_mkFreshId___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___spec__2___rarg(x_1, x_2); -lean_dec(x_1); -return x_3; } } -LEAN_EXPORT lean_object* l_Lean_mkFreshId___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: +else { -lean_object* x_7; -x_7 = l_Lean_mkFreshId___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___spec__2(x_1, x_2, x_3, x_4, x_5, x_6); +uint8_t x_108; +lean_dec(x_31); +lean_dec(x_30); +lean_dec(x_29); +lean_dec(x_28); +lean_dec(x_27); +lean_dec(x_26); +lean_dec(x_25); +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_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -return x_7; +x_108 = !lean_is_exclusive(x_56); +if (x_108 == 0) +{ +return x_56; +} +else +{ +lean_object* x_109; lean_object* x_110; lean_object* x_111; +x_109 = lean_ctor_get(x_56, 0); +x_110 = lean_ctor_get(x_56, 1); +lean_inc(x_110); +lean_inc(x_109); +lean_dec(x_56); +x_111 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_111, 0, x_109); +lean_ctor_set(x_111, 1, x_110); +return x_111; } } -LEAN_EXPORT lean_object* l_Lean_mkFreshFVarId___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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_mkFreshFVarId___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___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); +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; lean_object* x_25; lean_object* x_26; 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_11 = lean_ctor_get(x_1, 0); +lean_inc(x_11); +x_12 = lean_ctor_get(x_1, 1); +lean_inc(x_12); +x_13 = lean_ctor_get(x_1, 2); +lean_inc(x_13); +x_14 = lean_ctor_get(x_1, 3); +lean_inc(x_14); +x_15 = lean_ctor_get(x_1, 4); +lean_inc(x_15); +x_16 = lean_ctor_get(x_1, 5); +lean_inc(x_16); +x_17 = lean_ctor_get_uint8(x_1, sizeof(void*)*20); +x_18 = lean_ctor_get(x_1, 6); +lean_inc(x_18); +x_19 = lean_ctor_get(x_1, 7); +lean_inc(x_19); +x_20 = lean_ctor_get(x_1, 8); +lean_inc(x_20); +x_21 = lean_ctor_get(x_1, 9); +lean_inc(x_21); +x_22 = lean_ctor_get(x_1, 10); +lean_inc(x_22); +x_23 = lean_ctor_get(x_1, 11); +lean_inc(x_23); +x_24 = lean_ctor_get(x_1, 12); +lean_inc(x_24); +x_25 = lean_ctor_get(x_1, 13); +lean_inc(x_25); +x_26 = lean_ctor_get(x_1, 14); +lean_inc(x_26); +x_27 = lean_ctor_get(x_1, 15); +lean_inc(x_27); +x_28 = lean_ctor_get(x_1, 16); +lean_inc(x_28); +x_29 = lean_ctor_get(x_1, 17); +lean_inc(x_29); +x_30 = lean_ctor_get(x_1, 18); +lean_inc(x_30); +x_31 = lean_ctor_get(x_1, 19); +lean_inc(x_31); lean_dec(x_1); -return x_9; -} -} -LEAN_EXPORT lean_object* l_Lean_MVarId_assign___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: +lean_inc(x_11); +x_32 = l_Lean_MVarId_getType(x_11, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_32) == 0) { -lean_object* x_11; -x_11 = l_Lean_MVarId_assign___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___spec__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +uint8_t x_33; +x_33 = !lean_is_exclusive(x_32); +if (x_33 == 0) +{ +lean_object* x_34; lean_object* x_35; uint8_t x_36; +x_34 = lean_ctor_get(x_32, 0); +x_35 = lean_ctor_get(x_32, 1); +x_36 = l_Lean_Expr_isArrow(x_34); +if (x_36 == 0) +{ +uint8_t x_37; lean_object* x_38; +x_37 = l_Lean_Expr_isLet(x_34); +if (x_37 == 0) +{ +uint8_t x_56; +x_56 = l_Lean_Expr_isForall(x_34); +if (x_56 == 0) +{ +uint8_t x_57; +x_57 = l_Lean_Expr_isLetFun(x_34); +if (x_57 == 0) +{ +lean_object* x_58; +lean_dec(x_34); +lean_dec(x_31); +lean_dec(x_30); +lean_dec(x_29); +lean_dec(x_28); +lean_dec(x_27); +lean_dec(x_26); +lean_dec(x_25); +lean_dec(x_24); +lean_dec(x_23); +lean_dec(x_22); +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -1950,78 +2493,121 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -return x_11; -} +lean_dec(x_2); +x_58 = lean_box(0); +lean_ctor_set(x_32, 0, x_58); +return x_32; } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___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_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___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_4); -lean_dec(x_3); -lean_dec(x_2); -return x_10; +lean_object* x_59; +lean_free_object(x_32); +x_59 = lean_box(0); +x_38 = x_59; +goto block_55; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__2___boxed(lean_object** _args) { -lean_object* x_1 = _args[0]; -lean_object* x_2 = _args[1]; -lean_object* x_3 = _args[2]; -lean_object* x_4 = _args[3]; -lean_object* x_5 = _args[4]; -lean_object* x_6 = _args[5]; -lean_object* x_7 = _args[6]; -lean_object* x_8 = _args[7]; -lean_object* x_9 = _args[8]; -lean_object* x_10 = _args[9]; -lean_object* x_11 = _args[10]; -lean_object* x_12 = _args[11]; -lean_object* x_13 = _args[12]; -lean_object* x_14 = _args[13]; -lean_object* x_15 = _args[14]; -lean_object* x_16 = _args[15]; -lean_object* x_17 = _args[16]; -lean_object* x_18 = _args[17]; -lean_object* x_19 = _args[18]; -lean_object* x_20 = _args[19]; -lean_object* x_21 = _args[20]; -lean_object* x_22 = _args[21]; -lean_object* x_23 = _args[22]; -lean_object* x_24 = _args[23]; -lean_object* x_25 = _args[24]; -_start: +else { -uint8_t x_26; lean_object* x_27; -x_26 = lean_unbox(x_7); -lean_dec(x_7); -x_27 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_26, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_23, x_24, x_25); -lean_dec(x_20); -lean_dec(x_19); -lean_dec(x_18); -lean_dec(x_17); -return x_27; +lean_object* x_60; +lean_free_object(x_32); +x_60 = lean_box(0); +x_38 = x_60; +goto block_55; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___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: +else { -lean_object* x_10; -x_10 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_10; +lean_object* x_61; +lean_free_object(x_32); +x_61 = lean_box(0); +x_38 = x_61; +goto block_55; } +block_55: +{ +uint8_t x_39; lean_object* x_40; +lean_dec(x_38); +x_39 = 1; +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +x_40 = l_Lean_Meta_intro1Core(x_11, x_39, x_6, x_7, x_8, x_9, x_35); +if (lean_obj_tag(x_40) == 0) +{ +lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_41 = lean_ctor_get(x_40, 0); +lean_inc(x_41); +x_42 = lean_ctor_get(x_40, 1); +lean_inc(x_42); +lean_dec(x_40); +x_43 = lean_ctor_get(x_41, 0); +lean_inc(x_43); +x_44 = lean_ctor_get(x_41, 1); +lean_inc(x_44); +lean_dec(x_41); +lean_inc(x_43); +x_45 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__1___boxed), 9, 1); +lean_closure_set(x_45, 0, x_43); +x_46 = lean_box(x_17); +x_47 = lean_box(x_37); +lean_inc(x_44); +x_48 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___boxed), 34, 25); +lean_closure_set(x_48, 0, x_44); +lean_closure_set(x_48, 1, x_12); +lean_closure_set(x_48, 2, x_13); +lean_closure_set(x_48, 3, x_14); +lean_closure_set(x_48, 4, x_15); +lean_closure_set(x_48, 5, x_16); +lean_closure_set(x_48, 6, x_46); +lean_closure_set(x_48, 7, x_18); +lean_closure_set(x_48, 8, x_19); +lean_closure_set(x_48, 9, x_20); +lean_closure_set(x_48, 10, x_21); +lean_closure_set(x_48, 11, x_22); +lean_closure_set(x_48, 12, x_23); +lean_closure_set(x_48, 13, x_24); +lean_closure_set(x_48, 14, x_25); +lean_closure_set(x_48, 15, x_26); +lean_closure_set(x_48, 16, x_27); +lean_closure_set(x_48, 17, x_28); +lean_closure_set(x_48, 18, x_29); +lean_closure_set(x_48, 19, x_30); +lean_closure_set(x_48, 20, x_31); +lean_closure_set(x_48, 21, x_43); +lean_closure_set(x_48, 22, x_2); +lean_closure_set(x_48, 23, x_47); +lean_closure_set(x_48, 24, x_34); +x_49 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); +lean_closure_set(x_49, 0, x_45); +lean_closure_set(x_49, 1, x_48); +x_50 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(x_44, x_49, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_42); +return x_50; } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___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: +else { -lean_object* x_11; -x_11 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +uint8_t x_51; +lean_dec(x_34); +lean_dec(x_31); +lean_dec(x_30); +lean_dec(x_29); +lean_dec(x_28); +lean_dec(x_27); +lean_dec(x_26); +lean_dec(x_25); +lean_dec(x_24); +lean_dec(x_23); +lean_dec(x_22); +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -2029,48 +2615,96 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -lean_dec(x_1); -return x_11; +lean_dec(x_2); +x_51 = !lean_is_exclusive(x_40); +if (x_51 == 0) +{ +return x_40; } +else +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; +x_52 = lean_ctor_get(x_40, 0); +x_53 = lean_ctor_get(x_40, 1); +lean_inc(x_53); +lean_inc(x_52); +lean_dec(x_40); +x_54 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_54, 0, x_52); +lean_ctor_set(x_54, 1, x_53); +return x_54; } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___boxed(lean_object** _args) { -lean_object* x_1 = _args[0]; -lean_object* x_2 = _args[1]; -lean_object* x_3 = _args[2]; -lean_object* x_4 = _args[3]; -lean_object* x_5 = _args[4]; -lean_object* x_6 = _args[5]; -lean_object* x_7 = _args[6]; -lean_object* x_8 = _args[7]; -lean_object* x_9 = _args[8]; -lean_object* x_10 = _args[9]; -lean_object* x_11 = _args[10]; -lean_object* x_12 = _args[11]; -lean_object* x_13 = _args[12]; -lean_object* x_14 = _args[13]; -lean_object* x_15 = _args[14]; -lean_object* x_16 = _args[15]; -lean_object* x_17 = _args[16]; -lean_object* x_18 = _args[17]; -lean_object* x_19 = _args[18]; -lean_object* x_20 = _args[19]; -lean_object* x_21 = _args[20]; -lean_object* x_22 = _args[21]; -lean_object* x_23 = _args[22]; -lean_object* x_24 = _args[23]; -lean_object* x_25 = _args[24]; -lean_object* x_26 = _args[25]; -lean_object* x_27 = _args[26]; -lean_object* x_28 = _args[27]; -lean_object* x_29 = _args[28]; -lean_object* x_30 = _args[29]; -lean_object* x_31 = _args[30]; -_start: +} +} +} +else { -uint8_t x_32; lean_object* x_33; -x_32 = lean_unbox(x_9); -lean_dec(x_9); -x_33 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_32, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_23, x_24, x_25, x_26, x_27, x_28, x_29, x_30, x_31); +lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; +lean_free_object(x_32); +lean_dec(x_2); +x_62 = l_Lean_Expr_bindingDomain_x21(x_34); +lean_inc(x_62); +x_63 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__6___boxed), 9, 1); +lean_closure_set(x_63, 0, x_62); +x_64 = lean_box(x_17); +lean_inc(x_11); +x_65 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__9___boxed), 32, 23); +lean_closure_set(x_65, 0, x_11); +lean_closure_set(x_65, 1, x_12); +lean_closure_set(x_65, 2, x_13); +lean_closure_set(x_65, 3, x_14); +lean_closure_set(x_65, 4, x_15); +lean_closure_set(x_65, 5, x_16); +lean_closure_set(x_65, 6, x_64); +lean_closure_set(x_65, 7, x_18); +lean_closure_set(x_65, 8, x_19); +lean_closure_set(x_65, 9, x_20); +lean_closure_set(x_65, 10, x_21); +lean_closure_set(x_65, 11, x_22); +lean_closure_set(x_65, 12, x_23); +lean_closure_set(x_65, 13, x_24); +lean_closure_set(x_65, 14, x_25); +lean_closure_set(x_65, 15, x_26); +lean_closure_set(x_65, 16, x_27); +lean_closure_set(x_65, 17, x_28); +lean_closure_set(x_65, 18, x_29); +lean_closure_set(x_65, 19, x_30); +lean_closure_set(x_65, 20, x_31); +lean_closure_set(x_65, 21, x_34); +lean_closure_set(x_65, 22, x_62); +x_66 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); +lean_closure_set(x_66, 0, x_63); +lean_closure_set(x_66, 1, x_65); +x_67 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(x_11, x_66, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_35); +return x_67; +} +} +else +{ +lean_object* x_68; lean_object* x_69; uint8_t x_70; +x_68 = lean_ctor_get(x_32, 0); +x_69 = lean_ctor_get(x_32, 1); +lean_inc(x_69); +lean_inc(x_68); +lean_dec(x_32); +x_70 = l_Lean_Expr_isArrow(x_68); +if (x_70 == 0) +{ +uint8_t x_71; lean_object* x_72; +x_71 = l_Lean_Expr_isLet(x_68); +if (x_71 == 0) +{ +uint8_t x_90; +x_90 = l_Lean_Expr_isForall(x_68); +if (x_90 == 0) +{ +uint8_t x_91; +x_91 = l_Lean_Expr_isLetFun(x_68); +if (x_91 == 0) +{ +lean_object* x_92; lean_object* x_93; +lean_dec(x_68); +lean_dec(x_31); lean_dec(x_30); lean_dec(x_29); lean_dec(x_28); @@ -2078,591 +2712,970 @@ lean_dec(x_27); lean_dec(x_26); lean_dec(x_25); lean_dec(x_24); -lean_dec(x_1); -return x_33; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__6___boxed(lean_object** _args) { -lean_object* x_1 = _args[0]; -lean_object* x_2 = _args[1]; -lean_object* x_3 = _args[2]; -lean_object* x_4 = _args[3]; -lean_object* x_5 = _args[4]; -lean_object* x_6 = _args[5]; -lean_object* x_7 = _args[6]; -lean_object* x_8 = _args[7]; -lean_object* x_9 = _args[8]; -lean_object* x_10 = _args[9]; -lean_object* x_11 = _args[10]; -lean_object* x_12 = _args[11]; -lean_object* x_13 = _args[12]; -lean_object* x_14 = _args[13]; -lean_object* x_15 = _args[14]; -lean_object* x_16 = _args[15]; -lean_object* x_17 = _args[16]; -lean_object* x_18 = _args[17]; -lean_object* x_19 = _args[18]; -lean_object* x_20 = _args[19]; -lean_object* x_21 = _args[20]; -lean_object* x_22 = _args[21]; -lean_object* x_23 = _args[22]; -lean_object* x_24 = _args[23]; -lean_object* x_25 = _args[24]; -lean_object* x_26 = _args[25]; -_start: -{ -uint8_t x_27; uint8_t x_28; lean_object* x_29; -x_27 = lean_unbox(x_7); -lean_dec(x_7); -x_28 = lean_unbox(x_18); +lean_dec(x_23); +lean_dec(x_22); +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_19); lean_dec(x_18); -x_29 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__6(x_1, x_2, x_3, x_4, x_5, x_6, x_27, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_28, x_19, x_20, x_21, x_22, x_23, x_24, x_25, x_26); lean_dec(x_16); -return x_29; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_isCasesCandidate(lean_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_FVarId_getType(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; lean_object* x_11; -x_9 = lean_ctor_get(x_7, 0); -x_10 = lean_ctor_get(x_7, 1); -x_11 = l_Lean_Expr_getAppFn(x_9); -lean_dec(x_9); -if (lean_obj_tag(x_11) == 4) -{ -lean_object* x_12; lean_object* x_13; -lean_free_object(x_7); -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); -lean_dec(x_11); -x_13 = l_Lean_Meta_Grind_isGrindCasesTarget(x_12, x_4, x_5, x_10); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); lean_dec(x_12); -return x_13; +lean_dec(x_11); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_92 = lean_box(0); +x_93 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_93, 0, x_92); +lean_ctor_set(x_93, 1, x_69); +return x_93; } else { -uint8_t x_14; lean_object* x_15; -lean_dec(x_11); -x_14 = 0; -x_15 = lean_box(x_14); -lean_ctor_set(x_7, 0, x_15); -return x_7; +lean_object* x_94; +x_94 = lean_box(0); +x_72 = x_94; +goto block_89; } } 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 = l_Lean_Expr_getAppFn(x_16); -lean_dec(x_16); -if (lean_obj_tag(x_18) == 4) -{ -lean_object* x_19; lean_object* x_20; -x_19 = lean_ctor_get(x_18, 0); -lean_inc(x_19); -lean_dec(x_18); -x_20 = l_Lean_Meta_Grind_isGrindCasesTarget(x_19, x_4, x_5, x_17); -lean_dec(x_19); -return x_20; +lean_object* x_95; +x_95 = lean_box(0); +x_72 = x_95; +goto block_89; +} } else { -uint8_t x_21; lean_object* x_22; lean_object* x_23; -lean_dec(x_18); -x_21 = 0; -x_22 = lean_box(x_21); -x_23 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_23, 0, x_22); -lean_ctor_set(x_23, 1, x_17); -return x_23; -} +lean_object* x_96; +x_96 = lean_box(0); +x_72 = x_96; +goto block_89; } -} -else +block_89: { -uint8_t x_24; -x_24 = !lean_is_exclusive(x_7); -if (x_24 == 0) +uint8_t x_73; lean_object* x_74; +lean_dec(x_72); +x_73 = 1; +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +x_74 = l_Lean_Meta_intro1Core(x_11, x_73, x_6, x_7, x_8, x_9, x_69); +if (lean_obj_tag(x_74) == 0) { -return x_7; +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_75 = lean_ctor_get(x_74, 0); +lean_inc(x_75); +x_76 = lean_ctor_get(x_74, 1); +lean_inc(x_76); +lean_dec(x_74); +x_77 = lean_ctor_get(x_75, 0); +lean_inc(x_77); +x_78 = lean_ctor_get(x_75, 1); +lean_inc(x_78); +lean_dec(x_75); +lean_inc(x_77); +x_79 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__1___boxed), 9, 1); +lean_closure_set(x_79, 0, x_77); +x_80 = lean_box(x_17); +x_81 = lean_box(x_71); +lean_inc(x_78); +x_82 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___boxed), 34, 25); +lean_closure_set(x_82, 0, x_78); +lean_closure_set(x_82, 1, x_12); +lean_closure_set(x_82, 2, x_13); +lean_closure_set(x_82, 3, x_14); +lean_closure_set(x_82, 4, x_15); +lean_closure_set(x_82, 5, x_16); +lean_closure_set(x_82, 6, x_80); +lean_closure_set(x_82, 7, x_18); +lean_closure_set(x_82, 8, x_19); +lean_closure_set(x_82, 9, x_20); +lean_closure_set(x_82, 10, x_21); +lean_closure_set(x_82, 11, x_22); +lean_closure_set(x_82, 12, x_23); +lean_closure_set(x_82, 13, x_24); +lean_closure_set(x_82, 14, x_25); +lean_closure_set(x_82, 15, x_26); +lean_closure_set(x_82, 16, x_27); +lean_closure_set(x_82, 17, x_28); +lean_closure_set(x_82, 18, x_29); +lean_closure_set(x_82, 19, x_30); +lean_closure_set(x_82, 20, x_31); +lean_closure_set(x_82, 21, x_77); +lean_closure_set(x_82, 22, x_2); +lean_closure_set(x_82, 23, x_81); +lean_closure_set(x_82, 24, x_68); +x_83 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); +lean_closure_set(x_83, 0, x_79); +lean_closure_set(x_83, 1, x_82); +x_84 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(x_78, x_83, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_76); +return x_84; } else { -lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_25 = lean_ctor_get(x_7, 0); -x_26 = lean_ctor_get(x_7, 1); -lean_inc(x_26); -lean_inc(x_25); +lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; +lean_dec(x_68); +lean_dec(x_31); +lean_dec(x_30); +lean_dec(x_29); +lean_dec(x_28); +lean_dec(x_27); +lean_dec(x_26); +lean_dec(x_25); +lean_dec(x_24); +lean_dec(x_23); +lean_dec(x_22); +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_7); -x_27 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_27, 0, x_25); -lean_ctor_set(x_27, 1, x_26); -return x_27; -} -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_isCasesCandidate___boxed(lean_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_Tactic_Grind_Intro_0__Lean_Meta_Grind_isCasesCandidate(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -return x_7; +lean_dec(x_2); +x_85 = lean_ctor_get(x_74, 0); +lean_inc(x_85); +x_86 = lean_ctor_get(x_74, 1); +lean_inc(x_86); +if (lean_is_exclusive(x_74)) { + lean_ctor_release(x_74, 0); + lean_ctor_release(x_74, 1); + x_87 = x_74; +} else { + lean_dec_ref(x_74); + 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_EXPORT lean_object* l_List_mapTR_loop___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_applyCases_x3f___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -if (lean_obj_tag(x_2) == 0) -{ -lean_object* x_4; -x_4 = l_List_reverse___rarg(x_3); -return x_4; +lean_ctor_set(x_88, 0, x_85); +lean_ctor_set(x_88, 1, x_86); +return x_88; } -else -{ -uint8_t x_5; -x_5 = !lean_is_exclusive(x_2); -if (x_5 == 0) -{ -lean_object* x_6; lean_object* x_7; lean_object* x_8; 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_6 = lean_ctor_get(x_2, 0); -x_7 = lean_ctor_get(x_2, 1); -x_8 = lean_ctor_get(x_1, 1); -x_9 = lean_ctor_get(x_1, 2); -x_10 = lean_ctor_get(x_1, 3); -x_11 = lean_ctor_get(x_1, 4); -x_12 = lean_ctor_get(x_1, 5); -x_13 = lean_ctor_get_uint8(x_1, sizeof(void*)*14); -x_14 = lean_ctor_get(x_1, 6); -x_15 = lean_ctor_get(x_1, 7); -x_16 = lean_ctor_get(x_1, 8); -x_17 = lean_ctor_get(x_1, 9); -x_18 = lean_ctor_get(x_1, 10); -x_19 = lean_ctor_get(x_1, 11); -x_20 = lean_ctor_get(x_1, 12); -x_21 = lean_ctor_get(x_1, 13); -lean_inc(x_21); -lean_inc(x_20); -lean_inc(x_19); -lean_inc(x_18); -lean_inc(x_17); -lean_inc(x_16); -lean_inc(x_15); -lean_inc(x_14); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -x_22 = lean_alloc_ctor(0, 14, 1); -lean_ctor_set(x_22, 0, x_6); -lean_ctor_set(x_22, 1, x_8); -lean_ctor_set(x_22, 2, x_9); -lean_ctor_set(x_22, 3, x_10); -lean_ctor_set(x_22, 4, x_11); -lean_ctor_set(x_22, 5, x_12); -lean_ctor_set(x_22, 6, x_14); -lean_ctor_set(x_22, 7, x_15); -lean_ctor_set(x_22, 8, x_16); -lean_ctor_set(x_22, 9, x_17); -lean_ctor_set(x_22, 10, x_18); -lean_ctor_set(x_22, 11, x_19); -lean_ctor_set(x_22, 12, x_20); -lean_ctor_set(x_22, 13, x_21); -lean_ctor_set_uint8(x_22, sizeof(void*)*14, x_13); -lean_ctor_set(x_2, 1, x_3); -lean_ctor_set(x_2, 0, x_22); -{ -lean_object* _tmp_1 = x_7; -lean_object* _tmp_2 = x_2; -x_2 = _tmp_1; -x_3 = _tmp_2; } -goto _start; } else { -lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; -x_24 = lean_ctor_get(x_2, 0); -x_25 = lean_ctor_get(x_2, 1); -lean_inc(x_25); -lean_inc(x_24); +lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_dec(x_2); -x_26 = lean_ctor_get(x_1, 1); -x_27 = lean_ctor_get(x_1, 2); -x_28 = lean_ctor_get(x_1, 3); -x_29 = lean_ctor_get(x_1, 4); -x_30 = lean_ctor_get(x_1, 5); -x_31 = lean_ctor_get_uint8(x_1, sizeof(void*)*14); -x_32 = lean_ctor_get(x_1, 6); -x_33 = lean_ctor_get(x_1, 7); -x_34 = lean_ctor_get(x_1, 8); -x_35 = lean_ctor_get(x_1, 9); -x_36 = lean_ctor_get(x_1, 10); -x_37 = lean_ctor_get(x_1, 11); -x_38 = lean_ctor_get(x_1, 12); -x_39 = lean_ctor_get(x_1, 13); -lean_inc(x_39); -lean_inc(x_38); -lean_inc(x_37); -lean_inc(x_36); -lean_inc(x_35); -lean_inc(x_34); -lean_inc(x_33); -lean_inc(x_32); -lean_inc(x_30); -lean_inc(x_29); -lean_inc(x_28); -lean_inc(x_27); -lean_inc(x_26); -x_40 = lean_alloc_ctor(0, 14, 1); -lean_ctor_set(x_40, 0, x_24); -lean_ctor_set(x_40, 1, x_26); -lean_ctor_set(x_40, 2, x_27); -lean_ctor_set(x_40, 3, x_28); -lean_ctor_set(x_40, 4, x_29); -lean_ctor_set(x_40, 5, x_30); -lean_ctor_set(x_40, 6, x_32); -lean_ctor_set(x_40, 7, x_33); -lean_ctor_set(x_40, 8, x_34); -lean_ctor_set(x_40, 9, x_35); -lean_ctor_set(x_40, 10, x_36); -lean_ctor_set(x_40, 11, x_37); -lean_ctor_set(x_40, 12, x_38); -lean_ctor_set(x_40, 13, x_39); -lean_ctor_set_uint8(x_40, sizeof(void*)*14, x_31); -x_41 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_41, 0, x_40); -lean_ctor_set(x_41, 1, x_3); -x_2 = x_25; -x_3 = x_41; -goto _start; -} +x_97 = l_Lean_Expr_bindingDomain_x21(x_68); +lean_inc(x_97); +x_98 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__6___boxed), 9, 1); +lean_closure_set(x_98, 0, x_97); +x_99 = lean_box(x_17); +lean_inc(x_11); +x_100 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__9___boxed), 32, 23); +lean_closure_set(x_100, 0, x_11); +lean_closure_set(x_100, 1, x_12); +lean_closure_set(x_100, 2, x_13); +lean_closure_set(x_100, 3, x_14); +lean_closure_set(x_100, 4, x_15); +lean_closure_set(x_100, 5, x_16); +lean_closure_set(x_100, 6, x_99); +lean_closure_set(x_100, 7, x_18); +lean_closure_set(x_100, 8, x_19); +lean_closure_set(x_100, 9, x_20); +lean_closure_set(x_100, 10, x_21); +lean_closure_set(x_100, 11, x_22); +lean_closure_set(x_100, 12, x_23); +lean_closure_set(x_100, 13, x_24); +lean_closure_set(x_100, 14, x_25); +lean_closure_set(x_100, 15, x_26); +lean_closure_set(x_100, 16, x_27); +lean_closure_set(x_100, 17, x_28); +lean_closure_set(x_100, 18, x_29); +lean_closure_set(x_100, 19, x_30); +lean_closure_set(x_100, 20, x_31); +lean_closure_set(x_100, 21, x_68); +lean_closure_set(x_100, 22, x_97); +x_101 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); +lean_closure_set(x_101, 0, x_98); +lean_closure_set(x_101, 1, x_100); +x_102 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(x_11, x_101, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_69); +return x_102; } } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_applyCases_x3f___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -lean_object* x_9; -lean_inc(x_4); -lean_inc(x_1); -x_9 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_isCasesCandidate(x_1, x_4, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_9) == 0) -{ -lean_object* x_10; uint8_t x_11; -x_10 = lean_ctor_get(x_9, 0); -lean_inc(x_10); -x_11 = lean_unbox(x_10); -lean_dec(x_10); -if (x_11 == 0) +else { -uint8_t x_12; +uint8_t x_103; +lean_dec(x_31); +lean_dec(x_30); +lean_dec(x_29); +lean_dec(x_28); +lean_dec(x_27); +lean_dec(x_26); +lean_dec(x_25); +lean_dec(x_24); +lean_dec(x_23); +lean_dec(x_22); +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_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_12 = !lean_is_exclusive(x_9); -if (x_12 == 0) +x_103 = !lean_is_exclusive(x_32); +if (x_103 == 0) { -lean_object* x_13; lean_object* x_14; -x_13 = lean_ctor_get(x_9, 0); -lean_dec(x_13); -x_14 = lean_box(0); -lean_ctor_set(x_9, 0, x_14); -return x_9; +return x_32; } else { -lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = lean_ctor_get(x_9, 1); -lean_inc(x_15); -lean_dec(x_9); -x_16 = lean_box(0); -x_17 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_17, 0, x_16); -lean_ctor_set(x_17, 1, x_15); -return x_17; +lean_object* x_104; lean_object* x_105; lean_object* x_106; +x_104 = lean_ctor_get(x_32, 0); +x_105 = lean_ctor_get(x_32, 1); +lean_inc(x_105); +lean_inc(x_104); +lean_dec(x_32); +x_106 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_106, 0, x_104); +lean_ctor_set(x_106, 1, x_105); +return x_106; } } -else -{ -lean_object* x_18; lean_object* x_19; -x_18 = lean_ctor_get(x_9, 1); -lean_inc(x_18); -lean_dec(x_9); -x_19 = l_Lean_Meta_Grind_cases(x_2, x_1, x_4, x_5, x_6, x_7, x_18); -if (lean_obj_tag(x_19) == 0) -{ -uint8_t x_20; -x_20 = !lean_is_exclusive(x_19); -if (x_20 == 0) +} +} +LEAN_EXPORT lean_object* l_Lean_mkFreshId___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___spec__2___rarg___boxed(lean_object* x_1, lean_object* x_2) { +_start: { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_21 = lean_ctor_get(x_19, 0); -x_22 = lean_box(0); -x_23 = l_List_mapTR_loop___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_applyCases_x3f___spec__1(x_3, x_21, x_22); -x_24 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_24, 0, x_23); -lean_ctor_set(x_19, 0, x_24); -return x_19; -} -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_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_box(0); -x_28 = l_List_mapTR_loop___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_applyCases_x3f___spec__1(x_3, x_25, x_27); -x_29 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_29, 0, 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_26); -return x_30; -} +lean_object* x_3; +x_3 = l_Lean_mkFreshId___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___spec__2___rarg(x_1, x_2); +lean_dec(x_1); +return x_3; } -else -{ -uint8_t x_31; -x_31 = !lean_is_exclusive(x_19); -if (x_31 == 0) -{ -return x_19; } -else +LEAN_EXPORT lean_object* l_Lean_mkFreshId___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: { -lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_32 = lean_ctor_get(x_19, 0); -x_33 = lean_ctor_get(x_19, 1); -lean_inc(x_33); -lean_inc(x_32); -lean_dec(x_19); -x_34 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_34, 0, x_32); -lean_ctor_set(x_34, 1, x_33); -return x_34; -} -} +lean_object* x_7; +x_7 = l_Lean_mkFreshId___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___spec__2(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_7; } } -else +LEAN_EXPORT lean_object* l_Lean_mkFreshFVarId___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: { -uint8_t x_35; +lean_object* x_9; +x_9 = l_Lean_mkFreshFVarId___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___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); lean_dec(x_1); -x_35 = !lean_is_exclusive(x_9); -if (x_35 == 0) -{ return x_9; } -else +} +LEAN_EXPORT lean_object* l_Lean_MVarId_assign___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: { -lean_object* x_36; lean_object* x_37; lean_object* x_38; -x_36 = lean_ctor_get(x_9, 0); -x_37 = lean_ctor_get(x_9, 1); -lean_inc(x_37); -lean_inc(x_36); +lean_object* x_11; +x_11 = l_Lean_MVarId_assign___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___spec__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); lean_dec(x_9); -x_38 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_38, 0, x_36); -lean_ctor_set(x_38, 1, x_37); -return x_38; -} -} +lean_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___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_applyCases_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_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { -lean_object* x_8; lean_object* x_9; lean_object* x_10; -x_8 = lean_ctor_get(x_1, 0); -lean_inc(x_8); -lean_inc(x_8); -x_9 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_applyCases_x3f___lambda__1___boxed), 8, 3); -lean_closure_set(x_9, 0, x_2); -lean_closure_set(x_9, 1, x_8); -lean_closure_set(x_9, 2, x_1); -x_10 = l_Lean_MVarId_withContext___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___spec__2___rarg(x_8, x_9, x_3, x_4, x_5, x_6, x_7); +lean_object* x_10; +x_10 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___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_4); +lean_dec(x_3); +lean_dec(x_2); return x_10; } } -LEAN_EXPORT lean_object* l_List_mapTR_loop___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_applyCases_x3f___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___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_4; -x_4 = l_List_mapTR_loop___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_applyCases_x3f___spec__1(x_1, x_2, x_3); -lean_dec(x_1); -return x_4; +lean_object* x_10; +x_10 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_10; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_applyCases_x3f___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { -lean_object* x_9; -x_9 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_applyCases_x3f___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_object* x_10; +x_10 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); lean_dec(x_3); -return x_9; +lean_dec(x_2); +return x_10; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_applyInjection_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_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +lean_object* x_18 = _args[17]; +lean_object* x_19 = _args[18]; +lean_object* x_20 = _args[19]; +lean_object* x_21 = _args[20]; +lean_object* x_22 = _args[21]; +lean_object* x_23 = _args[22]; +lean_object* x_24 = _args[23]; +lean_object* x_25 = _args[24]; +lean_object* x_26 = _args[25]; +lean_object* x_27 = _args[26]; +lean_object* x_28 = _args[27]; +lean_object* x_29 = _args[28]; +lean_object* x_30 = _args[29]; +lean_object* x_31 = _args[30]; +lean_object* x_32 = _args[31]; +lean_object* x_33 = _args[32]; +lean_object* x_34 = _args[33]; _start: { -uint8_t x_8; -x_8 = !lean_is_exclusive(x_1); -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; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_9 = lean_ctor_get(x_1, 0); -x_10 = lean_ctor_get(x_1, 1); -x_11 = lean_ctor_get(x_1, 2); -x_12 = lean_ctor_get(x_1, 3); -x_13 = lean_ctor_get(x_1, 4); -x_14 = lean_ctor_get(x_1, 5); -x_15 = lean_ctor_get(x_1, 6); -x_16 = lean_ctor_get(x_1, 7); -x_17 = lean_ctor_get(x_1, 8); -x_18 = lean_ctor_get(x_1, 9); -x_19 = lean_ctor_get(x_1, 10); -x_20 = lean_ctor_get(x_1, 11); -x_21 = lean_ctor_get(x_1, 12); -x_22 = lean_ctor_get(x_1, 13); -x_23 = l_Lean_Meta_Grind_injection_x3f(x_9, x_2, x_3, x_4, x_5, x_6, x_7); -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) -{ -uint8_t x_25; -lean_free_object(x_1); -lean_dec(x_22); -lean_dec(x_21); -lean_dec(x_20); -lean_dec(x_19); -lean_dec(x_18); -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -x_25 = !lean_is_exclusive(x_23); -if (x_25 == 0) -{ -lean_object* x_26; lean_object* x_27; -x_26 = lean_ctor_get(x_23, 0); +uint8_t x_35; uint8_t x_36; lean_object* x_37; +x_35 = lean_unbox(x_7); +lean_dec(x_7); +x_36 = lean_unbox(x_24); +lean_dec(x_24); +x_37 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5(x_1, x_2, x_3, x_4, x_5, x_6, x_35, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_23, x_36, x_25, x_26, x_27, x_28, x_29, x_30, x_31, x_32, x_33, x_34); lean_dec(x_26); -x_27 = lean_box(0); -lean_ctor_set(x_23, 0, x_27); -return x_23; -} -else -{ -lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_28 = lean_ctor_get(x_23, 1); -lean_inc(x_28); -lean_dec(x_23); -x_29 = lean_box(0); -x_30 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_30, 0, x_29); -lean_ctor_set(x_30, 1, x_28); -return x_30; -} +lean_dec(x_25); +return x_37; } -else -{ -uint8_t x_31; -x_31 = !lean_is_exclusive(x_23); -if (x_31 == 0) -{ -lean_object* x_32; uint8_t x_33; -x_32 = lean_ctor_get(x_23, 0); -lean_dec(x_32); -x_33 = !lean_is_exclusive(x_24); -if (x_33 == 0) -{ -lean_object* x_34; -x_34 = lean_ctor_get(x_24, 0); -lean_ctor_set(x_1, 0, x_34); -lean_ctor_set(x_24, 0, x_1); -return x_23; } -else +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___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) { +_start: { -lean_object* x_35; lean_object* x_36; -x_35 = lean_ctor_get(x_24, 0); -lean_inc(x_35); -lean_dec(x_24); -lean_ctor_set(x_1, 0, x_35); -x_36 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_36, 0, x_1); -lean_ctor_set(x_23, 0, x_36); -return x_23; +lean_object* x_10; +x_10 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__6(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_10; } } -else +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: { -lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; -x_37 = lean_ctor_get(x_23, 1); -lean_inc(x_37); -lean_dec(x_23); -x_38 = lean_ctor_get(x_24, 0); -lean_inc(x_38); -if (lean_is_exclusive(x_24)) { - lean_ctor_release(x_24, 0); - x_39 = x_24; -} else { - lean_dec_ref(x_24); - x_39 = lean_box(0); -} -lean_ctor_set(x_1, 0, x_38); -if (lean_is_scalar(x_39)) { - x_40 = lean_alloc_ctor(1, 1, 0); -} else { - x_40 = x_39; -} -lean_ctor_set(x_40, 0, x_1); -x_41 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_41, 0, x_40); -lean_ctor_set(x_41, 1, x_37); -return x_41; -} +lean_object* x_11; +x_11 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__7(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +return x_11; } } -else -{ -uint8_t x_42; -lean_free_object(x_1); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__8___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +lean_object* x_18 = _args[17]; +lean_object* x_19 = _args[18]; +lean_object* x_20 = _args[19]; +lean_object* x_21 = _args[20]; +lean_object* x_22 = _args[21]; +lean_object* x_23 = _args[22]; +lean_object* x_24 = _args[23]; +lean_object* x_25 = _args[24]; +lean_object* x_26 = _args[25]; +lean_object* x_27 = _args[26]; +lean_object* x_28 = _args[27]; +lean_object* x_29 = _args[28]; +lean_object* x_30 = _args[29]; +lean_object* x_31 = _args[30]; +lean_object* x_32 = _args[31]; +lean_object* x_33 = _args[32]; +lean_object* x_34 = _args[33]; +lean_object* x_35 = _args[34]; +lean_object* x_36 = _args[35]; +lean_object* x_37 = _args[36]; +_start: +{ +uint8_t x_38; lean_object* x_39; +x_38 = lean_unbox(x_9); +lean_dec(x_9); +x_39 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__8(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_38, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_23, x_24, x_25, x_26, x_27, x_28, x_29, x_30, x_31, x_32, x_33, x_34, x_35, x_36, x_37); +lean_dec(x_36); +lean_dec(x_35); +lean_dec(x_34); +lean_dec(x_33); +lean_dec(x_32); +lean_dec(x_31); +lean_dec(x_30); +lean_dec(x_1); +return x_39; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__9___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +lean_object* x_18 = _args[17]; +lean_object* x_19 = _args[18]; +lean_object* x_20 = _args[19]; +lean_object* x_21 = _args[20]; +lean_object* x_22 = _args[21]; +lean_object* x_23 = _args[22]; +lean_object* x_24 = _args[23]; +lean_object* x_25 = _args[24]; +lean_object* x_26 = _args[25]; +lean_object* x_27 = _args[26]; +lean_object* x_28 = _args[27]; +lean_object* x_29 = _args[28]; +lean_object* x_30 = _args[29]; +lean_object* x_31 = _args[30]; +lean_object* x_32 = _args[31]; +_start: +{ +uint8_t x_33; uint8_t x_34; lean_object* x_35; +x_33 = lean_unbox(x_7); +lean_dec(x_7); +x_34 = lean_unbox(x_24); +lean_dec(x_24); +x_35 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__9(x_1, x_2, x_3, x_4, x_5, x_6, x_33, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_23, x_34, x_25, x_26, x_27, x_28, x_29, x_30, x_31, x_32); +lean_dec(x_22); +return x_35; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_isCasesCandidate(lean_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_Expr_getAppFn(x_1); +if (lean_obj_tag(x_7) == 4) +{ +lean_object* x_8; lean_object* x_9; +x_8 = lean_ctor_get(x_7, 0); +lean_inc(x_8); +lean_dec(x_7); +x_9 = l_Lean_Meta_Grind_isGrindCasesTarget(x_8, x_4, x_5, x_6); +lean_dec(x_8); +return x_9; +} +else +{ +uint8_t x_10; lean_object* x_11; lean_object* x_12; +lean_dec(x_7); +x_10 = 0; +x_11 = lean_box(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_6); +return x_12; +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_isCasesCandidate___boxed(lean_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_Tactic_Grind_Intro_0__Lean_Meta_Grind_isCasesCandidate(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_7; +} +} +LEAN_EXPORT lean_object* l_List_mapTR_loop___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_applyCases_x3f___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_4; +x_4 = l_List_reverse___rarg(x_3); +return x_4; +} +else +{ +uint8_t x_5; +x_5 = !lean_is_exclusive(x_2); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_6 = lean_ctor_get(x_2, 0); +x_7 = lean_ctor_get(x_2, 1); +x_8 = lean_ctor_get(x_1, 1); +x_9 = lean_ctor_get(x_1, 2); +x_10 = lean_ctor_get(x_1, 3); +x_11 = lean_ctor_get(x_1, 4); +x_12 = lean_ctor_get(x_1, 5); +x_13 = lean_ctor_get_uint8(x_1, sizeof(void*)*20); +x_14 = lean_ctor_get(x_1, 6); +x_15 = lean_ctor_get(x_1, 7); +x_16 = lean_ctor_get(x_1, 8); +x_17 = lean_ctor_get(x_1, 9); +x_18 = lean_ctor_get(x_1, 10); +x_19 = lean_ctor_get(x_1, 11); +x_20 = lean_ctor_get(x_1, 12); +x_21 = lean_ctor_get(x_1, 13); +x_22 = lean_ctor_get(x_1, 14); +x_23 = lean_ctor_get(x_1, 15); +x_24 = lean_ctor_get(x_1, 16); +x_25 = lean_ctor_get(x_1, 17); +x_26 = lean_ctor_get(x_1, 18); +x_27 = lean_ctor_get(x_1, 19); +lean_inc(x_27); +lean_inc(x_26); +lean_inc(x_25); +lean_inc(x_24); +lean_inc(x_23); +lean_inc(x_22); +lean_inc(x_21); +lean_inc(x_20); +lean_inc(x_19); +lean_inc(x_18); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +x_28 = lean_alloc_ctor(0, 20, 1); +lean_ctor_set(x_28, 0, x_6); +lean_ctor_set(x_28, 1, x_8); +lean_ctor_set(x_28, 2, x_9); +lean_ctor_set(x_28, 3, x_10); +lean_ctor_set(x_28, 4, x_11); +lean_ctor_set(x_28, 5, x_12); +lean_ctor_set(x_28, 6, x_14); +lean_ctor_set(x_28, 7, x_15); +lean_ctor_set(x_28, 8, x_16); +lean_ctor_set(x_28, 9, x_17); +lean_ctor_set(x_28, 10, x_18); +lean_ctor_set(x_28, 11, x_19); +lean_ctor_set(x_28, 12, x_20); +lean_ctor_set(x_28, 13, x_21); +lean_ctor_set(x_28, 14, x_22); +lean_ctor_set(x_28, 15, x_23); +lean_ctor_set(x_28, 16, x_24); +lean_ctor_set(x_28, 17, x_25); +lean_ctor_set(x_28, 18, x_26); +lean_ctor_set(x_28, 19, x_27); +lean_ctor_set_uint8(x_28, sizeof(void*)*20, x_13); +lean_ctor_set(x_2, 1, x_3); +lean_ctor_set(x_2, 0, x_28); +{ +lean_object* _tmp_1 = x_7; +lean_object* _tmp_2 = x_2; +x_2 = _tmp_1; +x_3 = _tmp_2; +} +goto _start; +} +else +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_30 = lean_ctor_get(x_2, 0); +x_31 = lean_ctor_get(x_2, 1); +lean_inc(x_31); +lean_inc(x_30); +lean_dec(x_2); +x_32 = lean_ctor_get(x_1, 1); +x_33 = lean_ctor_get(x_1, 2); +x_34 = lean_ctor_get(x_1, 3); +x_35 = lean_ctor_get(x_1, 4); +x_36 = lean_ctor_get(x_1, 5); +x_37 = lean_ctor_get_uint8(x_1, sizeof(void*)*20); +x_38 = lean_ctor_get(x_1, 6); +x_39 = lean_ctor_get(x_1, 7); +x_40 = lean_ctor_get(x_1, 8); +x_41 = lean_ctor_get(x_1, 9); +x_42 = lean_ctor_get(x_1, 10); +x_43 = lean_ctor_get(x_1, 11); +x_44 = lean_ctor_get(x_1, 12); +x_45 = lean_ctor_get(x_1, 13); +x_46 = lean_ctor_get(x_1, 14); +x_47 = lean_ctor_get(x_1, 15); +x_48 = lean_ctor_get(x_1, 16); +x_49 = lean_ctor_get(x_1, 17); +x_50 = lean_ctor_get(x_1, 18); +x_51 = lean_ctor_get(x_1, 19); +lean_inc(x_51); +lean_inc(x_50); +lean_inc(x_49); +lean_inc(x_48); +lean_inc(x_47); +lean_inc(x_46); +lean_inc(x_45); +lean_inc(x_44); +lean_inc(x_43); +lean_inc(x_42); +lean_inc(x_41); +lean_inc(x_40); +lean_inc(x_39); +lean_inc(x_38); +lean_inc(x_36); +lean_inc(x_35); +lean_inc(x_34); +lean_inc(x_33); +lean_inc(x_32); +x_52 = lean_alloc_ctor(0, 20, 1); +lean_ctor_set(x_52, 0, x_30); +lean_ctor_set(x_52, 1, x_32); +lean_ctor_set(x_52, 2, x_33); +lean_ctor_set(x_52, 3, x_34); +lean_ctor_set(x_52, 4, x_35); +lean_ctor_set(x_52, 5, x_36); +lean_ctor_set(x_52, 6, x_38); +lean_ctor_set(x_52, 7, x_39); +lean_ctor_set(x_52, 8, x_40); +lean_ctor_set(x_52, 9, x_41); +lean_ctor_set(x_52, 10, x_42); +lean_ctor_set(x_52, 11, x_43); +lean_ctor_set(x_52, 12, x_44); +lean_ctor_set(x_52, 13, x_45); +lean_ctor_set(x_52, 14, x_46); +lean_ctor_set(x_52, 15, x_47); +lean_ctor_set(x_52, 16, x_48); +lean_ctor_set(x_52, 17, x_49); +lean_ctor_set(x_52, 18, x_50); +lean_ctor_set(x_52, 19, x_51); +lean_ctor_set_uint8(x_52, sizeof(void*)*20, x_37); +x_53 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_53, 0, x_52); +lean_ctor_set(x_53, 1, x_3); +x_2 = x_31; +x_3 = x_53; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_applyCases_x3f___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; +lean_inc(x_4); +lean_inc(x_1); +x_9 = l_Lean_FVarId_getType(x_1, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_9) == 0) +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +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___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_isCasesCandidate(x_10, x_4, x_5, x_6, x_7, x_11); +lean_dec(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) +{ +uint8_t x_15; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_15 = !lean_is_exclusive(x_12); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; +x_16 = lean_ctor_get(x_12, 0); +lean_dec(x_16); +x_17 = lean_box(0); +lean_ctor_set(x_12, 0, x_17); +return x_12; +} +else +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_18 = lean_ctor_get(x_12, 1); +lean_inc(x_18); +lean_dec(x_12); +x_19 = lean_box(0); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_19); +lean_ctor_set(x_20, 1, x_18); +return x_20; +} +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_12, 1); +lean_inc(x_21); +lean_dec(x_12); +x_22 = l_Lean_Expr_fvar___override(x_1); +x_23 = l_Lean_Meta_Grind_cases(x_2, x_22, x_4, x_5, x_6, x_7, x_21); +if (lean_obj_tag(x_23) == 0) +{ +uint8_t x_24; +x_24 = !lean_is_exclusive(x_23); +if (x_24 == 0) +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_25 = lean_ctor_get(x_23, 0); +x_26 = lean_box(0); +x_27 = l_List_mapTR_loop___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_applyCases_x3f___spec__1(x_3, x_25, x_26); +x_28 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_28, 0, x_27); +lean_ctor_set(x_23, 0, x_28); +return x_23; +} +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; +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_box(0); +x_32 = l_List_mapTR_loop___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_applyCases_x3f___spec__1(x_3, x_29, x_31); +x_33 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_33, 0, x_32); +x_34 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_34, 0, x_33); +lean_ctor_set(x_34, 1, x_30); +return x_34; +} +} +else +{ +uint8_t x_35; +x_35 = !lean_is_exclusive(x_23); +if (x_35 == 0) +{ +return x_23; +} +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; +} +} +} +} +else +{ +uint8_t x_39; +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_39 = !lean_is_exclusive(x_9); +if (x_39 == 0) +{ +return x_9; +} +else +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_40 = lean_ctor_get(x_9, 0); +x_41 = lean_ctor_get(x_9, 1); +lean_inc(x_41); +lean_inc(x_40); +lean_dec(x_9); +x_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_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_applyCases_x3f(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_8 = lean_ctor_get(x_1, 0); +lean_inc(x_8); +lean_inc(x_8); +x_9 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_applyCases_x3f___lambda__1___boxed), 8, 3); +lean_closure_set(x_9, 0, x_2); +lean_closure_set(x_9, 1, x_8); +lean_closure_set(x_9, 2, x_1); +x_10 = l_Lean_MVarId_withContext___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___spec__2___rarg(x_8, x_9, x_3, x_4, x_5, x_6, x_7); +return x_10; +} +} +LEAN_EXPORT lean_object* l_List_mapTR_loop___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_applyCases_x3f___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_List_mapTR_loop___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_applyCases_x3f___spec__1(x_1, x_2, x_3); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_applyCases_x3f___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; +x_9 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_applyCases_x3f___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___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_applyInjection_x3f(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +uint8_t x_8; +x_8 = !lean_is_exclusive(x_1); +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; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; 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_9 = lean_ctor_get(x_1, 0); +x_10 = lean_ctor_get(x_1, 1); +x_11 = lean_ctor_get(x_1, 2); +x_12 = lean_ctor_get(x_1, 3); +x_13 = lean_ctor_get(x_1, 4); +x_14 = lean_ctor_get(x_1, 5); +x_15 = lean_ctor_get(x_1, 6); +x_16 = lean_ctor_get(x_1, 7); +x_17 = lean_ctor_get(x_1, 8); +x_18 = lean_ctor_get(x_1, 9); +x_19 = lean_ctor_get(x_1, 10); +x_20 = lean_ctor_get(x_1, 11); +x_21 = lean_ctor_get(x_1, 12); +x_22 = lean_ctor_get(x_1, 13); +x_23 = lean_ctor_get(x_1, 14); +x_24 = lean_ctor_get(x_1, 15); +x_25 = lean_ctor_get(x_1, 16); +x_26 = lean_ctor_get(x_1, 17); +x_27 = lean_ctor_get(x_1, 18); +x_28 = lean_ctor_get(x_1, 19); +x_29 = l_Lean_Meta_Grind_injection_x3f(x_9, x_2, x_3, x_4, x_5, x_6, x_7); +if (lean_obj_tag(x_29) == 0) +{ +lean_object* x_30; +x_30 = lean_ctor_get(x_29, 0); +lean_inc(x_30); +if (lean_obj_tag(x_30) == 0) +{ +uint8_t x_31; +lean_free_object(x_1); +lean_dec(x_28); +lean_dec(x_27); +lean_dec(x_26); +lean_dec(x_25); +lean_dec(x_24); +lean_dec(x_23); lean_dec(x_22); lean_dec(x_21); lean_dec(x_20); @@ -2676,191 +3689,327 @@ lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); -x_42 = !lean_is_exclusive(x_23); -if (x_42 == 0) +x_31 = !lean_is_exclusive(x_29); +if (x_31 == 0) { -return x_23; +lean_object* x_32; lean_object* x_33; +x_32 = lean_ctor_get(x_29, 0); +lean_dec(x_32); +x_33 = lean_box(0); +lean_ctor_set(x_29, 0, x_33); +return x_29; } 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_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_29, 1); +lean_inc(x_34); +lean_dec(x_29); +x_35 = lean_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; +} +} +else +{ +uint8_t x_37; +x_37 = !lean_is_exclusive(x_29); +if (x_37 == 0) +{ +lean_object* x_38; uint8_t x_39; +x_38 = lean_ctor_get(x_29, 0); +lean_dec(x_38); +x_39 = !lean_is_exclusive(x_30); +if (x_39 == 0) +{ +lean_object* x_40; +x_40 = lean_ctor_get(x_30, 0); +lean_ctor_set(x_1, 0, x_40); +lean_ctor_set(x_30, 0, x_1); +return x_29; +} +else +{ +lean_object* x_41; lean_object* x_42; +x_41 = lean_ctor_get(x_30, 0); +lean_inc(x_41); +lean_dec(x_30); +lean_ctor_set(x_1, 0, x_41); +x_42 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_42, 0, x_1); +lean_ctor_set(x_29, 0, x_42); +return x_29; +} +} +else +{ +lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_43 = lean_ctor_get(x_29, 1); lean_inc(x_43); +lean_dec(x_29); +x_44 = lean_ctor_get(x_30, 0); +lean_inc(x_44); +if (lean_is_exclusive(x_30)) { + lean_ctor_release(x_30, 0); + x_45 = x_30; +} else { + lean_dec_ref(x_30); + x_45 = lean_box(0); +} +lean_ctor_set(x_1, 0, x_44); +if (lean_is_scalar(x_45)) { + x_46 = lean_alloc_ctor(1, 1, 0); +} else { + x_46 = x_45; +} +lean_ctor_set(x_46, 0, x_1); +x_47 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_47, 0, x_46); +lean_ctor_set(x_47, 1, x_43); +return x_47; +} +} +} +else +{ +uint8_t x_48; +lean_free_object(x_1); +lean_dec(x_28); +lean_dec(x_27); +lean_dec(x_26); +lean_dec(x_25); +lean_dec(x_24); 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; +lean_dec(x_22); +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +x_48 = !lean_is_exclusive(x_29); +if (x_48 == 0) +{ +return x_29; +} +else +{ +lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_49 = lean_ctor_get(x_29, 0); +x_50 = lean_ctor_get(x_29, 1); +lean_inc(x_50); +lean_inc(x_49); +lean_dec(x_29); +x_51 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_51, 0, x_49); +lean_ctor_set(x_51, 1, x_50); +return x_51; } } } else { -lean_object* x_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; -x_46 = lean_ctor_get(x_1, 0); -x_47 = lean_ctor_get(x_1, 1); -x_48 = lean_ctor_get(x_1, 2); -x_49 = lean_ctor_get(x_1, 3); -x_50 = lean_ctor_get(x_1, 4); -x_51 = lean_ctor_get(x_1, 5); -x_52 = lean_ctor_get_uint8(x_1, sizeof(void*)*14); -x_53 = lean_ctor_get(x_1, 6); -x_54 = lean_ctor_get(x_1, 7); -x_55 = lean_ctor_get(x_1, 8); -x_56 = lean_ctor_get(x_1, 9); -x_57 = lean_ctor_get(x_1, 10); -x_58 = lean_ctor_get(x_1, 11); -x_59 = lean_ctor_get(x_1, 12); -x_60 = lean_ctor_get(x_1, 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; 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; +x_52 = lean_ctor_get(x_1, 0); +x_53 = lean_ctor_get(x_1, 1); +x_54 = lean_ctor_get(x_1, 2); +x_55 = lean_ctor_get(x_1, 3); +x_56 = lean_ctor_get(x_1, 4); +x_57 = lean_ctor_get(x_1, 5); +x_58 = lean_ctor_get_uint8(x_1, sizeof(void*)*20); +x_59 = lean_ctor_get(x_1, 6); +x_60 = lean_ctor_get(x_1, 7); +x_61 = lean_ctor_get(x_1, 8); +x_62 = lean_ctor_get(x_1, 9); +x_63 = lean_ctor_get(x_1, 10); +x_64 = lean_ctor_get(x_1, 11); +x_65 = lean_ctor_get(x_1, 12); +x_66 = lean_ctor_get(x_1, 13); +x_67 = lean_ctor_get(x_1, 14); +x_68 = lean_ctor_get(x_1, 15); +x_69 = lean_ctor_get(x_1, 16); +x_70 = lean_ctor_get(x_1, 17); +x_71 = lean_ctor_get(x_1, 18); +x_72 = lean_ctor_get(x_1, 19); +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_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_51); -lean_inc(x_50); -lean_inc(x_49); -lean_inc(x_48); -lean_inc(x_47); -lean_inc(x_46); +lean_inc(x_52); lean_dec(x_1); -x_61 = l_Lean_Meta_Grind_injection_x3f(x_46, x_2, x_3, x_4, x_5, x_6, x_7); -if (lean_obj_tag(x_61) == 0) +x_73 = l_Lean_Meta_Grind_injection_x3f(x_52, x_2, x_3, x_4, x_5, x_6, x_7); +if (lean_obj_tag(x_73) == 0) { -lean_object* x_62; -x_62 = lean_ctor_get(x_61, 0); -lean_inc(x_62); -if (lean_obj_tag(x_62) == 0) +lean_object* x_74; +x_74 = lean_ctor_get(x_73, 0); +lean_inc(x_74); +if (lean_obj_tag(x_74) == 0) { -lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; +lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; +lean_dec(x_72); +lean_dec(x_71); +lean_dec(x_70); +lean_dec(x_69); +lean_dec(x_68); +lean_dec(x_67); +lean_dec(x_66); +lean_dec(x_65); +lean_dec(x_64); +lean_dec(x_63); +lean_dec(x_62); +lean_dec(x_61); lean_dec(x_60); lean_dec(x_59); -lean_dec(x_58); lean_dec(x_57); lean_dec(x_56); lean_dec(x_55); lean_dec(x_54); lean_dec(x_53); -lean_dec(x_51); -lean_dec(x_50); -lean_dec(x_49); -lean_dec(x_48); -lean_dec(x_47); -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; +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_61); - x_64 = lean_box(0); + lean_dec_ref(x_73); + x_76 = lean_box(0); } -x_65 = lean_box(0); -if (lean_is_scalar(x_64)) { - x_66 = lean_alloc_ctor(0, 2, 0); +x_77 = lean_box(0); +if (lean_is_scalar(x_76)) { + x_78 = lean_alloc_ctor(0, 2, 0); } else { - x_66 = x_64; + x_78 = x_76; } -lean_ctor_set(x_66, 0, x_65); -lean_ctor_set(x_66, 1, x_63); -return x_66; +lean_ctor_set(x_78, 0, x_77); +lean_ctor_set(x_78, 1, x_75); +return x_78; } else { -lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; -x_67 = lean_ctor_get(x_61, 1); -lean_inc(x_67); -if (lean_is_exclusive(x_61)) { - lean_ctor_release(x_61, 0); - lean_ctor_release(x_61, 1); - x_68 = x_61; +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_79 = lean_ctor_get(x_73, 1); +lean_inc(x_79); +if (lean_is_exclusive(x_73)) { + lean_ctor_release(x_73, 0); + lean_ctor_release(x_73, 1); + x_80 = x_73; } else { - lean_dec_ref(x_61); - x_68 = lean_box(0); + lean_dec_ref(x_73); + x_80 = lean_box(0); } -x_69 = lean_ctor_get(x_62, 0); -lean_inc(x_69); -if (lean_is_exclusive(x_62)) { - lean_ctor_release(x_62, 0); - x_70 = x_62; +x_81 = lean_ctor_get(x_74, 0); +lean_inc(x_81); +if (lean_is_exclusive(x_74)) { + lean_ctor_release(x_74, 0); + x_82 = x_74; } else { - lean_dec_ref(x_62); - x_70 = lean_box(0); -} -x_71 = lean_alloc_ctor(0, 14, 1); -lean_ctor_set(x_71, 0, x_69); -lean_ctor_set(x_71, 1, x_47); -lean_ctor_set(x_71, 2, x_48); -lean_ctor_set(x_71, 3, x_49); -lean_ctor_set(x_71, 4, x_50); -lean_ctor_set(x_71, 5, x_51); -lean_ctor_set(x_71, 6, x_53); -lean_ctor_set(x_71, 7, x_54); -lean_ctor_set(x_71, 8, x_55); -lean_ctor_set(x_71, 9, x_56); -lean_ctor_set(x_71, 10, x_57); -lean_ctor_set(x_71, 11, x_58); -lean_ctor_set(x_71, 12, x_59); -lean_ctor_set(x_71, 13, x_60); -lean_ctor_set_uint8(x_71, sizeof(void*)*14, x_52); -if (lean_is_scalar(x_70)) { - x_72 = lean_alloc_ctor(1, 1, 0); + lean_dec_ref(x_74); + x_82 = lean_box(0); +} +x_83 = lean_alloc_ctor(0, 20, 1); +lean_ctor_set(x_83, 0, x_81); +lean_ctor_set(x_83, 1, x_53); +lean_ctor_set(x_83, 2, x_54); +lean_ctor_set(x_83, 3, x_55); +lean_ctor_set(x_83, 4, x_56); +lean_ctor_set(x_83, 5, x_57); +lean_ctor_set(x_83, 6, x_59); +lean_ctor_set(x_83, 7, x_60); +lean_ctor_set(x_83, 8, x_61); +lean_ctor_set(x_83, 9, x_62); +lean_ctor_set(x_83, 10, x_63); +lean_ctor_set(x_83, 11, x_64); +lean_ctor_set(x_83, 12, x_65); +lean_ctor_set(x_83, 13, x_66); +lean_ctor_set(x_83, 14, x_67); +lean_ctor_set(x_83, 15, x_68); +lean_ctor_set(x_83, 16, x_69); +lean_ctor_set(x_83, 17, x_70); +lean_ctor_set(x_83, 18, x_71); +lean_ctor_set(x_83, 19, x_72); +lean_ctor_set_uint8(x_83, sizeof(void*)*20, x_58); +if (lean_is_scalar(x_82)) { + x_84 = lean_alloc_ctor(1, 1, 0); } else { - x_72 = x_70; + x_84 = x_82; } -lean_ctor_set(x_72, 0, x_71); -if (lean_is_scalar(x_68)) { - x_73 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_84, 0, x_83); +if (lean_is_scalar(x_80)) { + x_85 = lean_alloc_ctor(0, 2, 0); } else { - x_73 = x_68; + x_85 = x_80; } -lean_ctor_set(x_73, 0, x_72); -lean_ctor_set(x_73, 1, x_67); -return x_73; +lean_ctor_set(x_85, 0, x_84); +lean_ctor_set(x_85, 1, x_79); +return x_85; } } else { -lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; +lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; +lean_dec(x_72); +lean_dec(x_71); +lean_dec(x_70); +lean_dec(x_69); +lean_dec(x_68); +lean_dec(x_67); +lean_dec(x_66); +lean_dec(x_65); +lean_dec(x_64); +lean_dec(x_63); +lean_dec(x_62); +lean_dec(x_61); lean_dec(x_60); lean_dec(x_59); -lean_dec(x_58); lean_dec(x_57); lean_dec(x_56); lean_dec(x_55); lean_dec(x_54); lean_dec(x_53); -lean_dec(x_51); -lean_dec(x_50); -lean_dec(x_49); -lean_dec(x_48); -lean_dec(x_47); -x_74 = lean_ctor_get(x_61, 0); -lean_inc(x_74); -x_75 = lean_ctor_get(x_61, 1); -lean_inc(x_75); -if (lean_is_exclusive(x_61)) { - lean_ctor_release(x_61, 0); - lean_ctor_release(x_61, 1); - x_76 = x_61; +x_86 = lean_ctor_get(x_73, 0); +lean_inc(x_86); +x_87 = lean_ctor_get(x_73, 1); +lean_inc(x_87); +if (lean_is_exclusive(x_73)) { + lean_ctor_release(x_73, 0); + lean_ctor_release(x_73, 1); + x_88 = x_73; } else { - lean_dec_ref(x_61); - x_76 = lean_box(0); + lean_dec_ref(x_73); + x_88 = lean_box(0); } -if (lean_is_scalar(x_76)) { - x_77 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_88)) { + x_89 = lean_alloc_ctor(1, 2, 0); } else { - x_77 = x_76; + x_89 = x_88; } -lean_ctor_set(x_77, 0, x_74); -lean_ctor_set(x_77, 1, x_75); -return x_77; +lean_ctor_set(x_89, 0, x_86); +lean_ctor_set(x_89, 1, x_87); +return x_89; } } } @@ -3027,32 +4176,7 @@ return x_22; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_intros_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) { -_start: -{ -lean_object* x_10; uint8_t x_11; -x_10 = lean_st_mk_ref(x_1, x_9); -x_11 = !lean_is_exclusive(x_10); -if (x_11 == 0) -{ -return x_10; -} -else -{ -lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_12 = lean_ctor_get(x_10, 0); -x_13 = lean_ctor_get(x_10, 1); -lean_inc(x_13); -lean_inc(x_12); -lean_dec(x_10); -x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_12); -lean_ctor_set(x_14, 1, x_13); -return x_14; -} -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_intros_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_object* x_11) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_intros_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, lean_object* x_11) { _start: { lean_object* x_12; @@ -3075,166 +4199,545 @@ lean_dec(x_3); x_18 = !lean_is_exclusive(x_17); if (x_18 == 0) { -lean_object* x_19; -x_19 = lean_ctor_get(x_17, 0); -lean_ctor_set(x_14, 1, x_19); -lean_ctor_set(x_17, 0, x_14); -return x_17; +lean_object* x_19; +x_19 = lean_ctor_get(x_17, 0); +lean_ctor_set(x_14, 1, x_19); +lean_ctor_set(x_17, 0, x_14); +return x_17; +} +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_17, 0); +x_21 = lean_ctor_get(x_17, 1); +lean_inc(x_21); +lean_inc(x_20); +lean_dec(x_17); +lean_ctor_set(x_14, 1, x_20); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_14); +lean_ctor_set(x_22, 1, x_21); +return x_22; +} +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_23 = lean_ctor_get(x_14, 0); +x_24 = lean_ctor_get(x_14, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_14); +x_25 = lean_st_ref_get(x_3, x_24); +lean_dec(x_3); +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); +} +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_23); +lean_ctor_set(x_29, 1, x_26); +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; +} +} +else +{ +uint8_t x_31; +lean_dec(x_3); +x_31 = !lean_is_exclusive(x_12); +if (x_31 == 0) +{ +return x_12; +} +else +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_12, 0); +x_33 = lean_ctor_get(x_12, 1); +lean_inc(x_33); +lean_inc(x_32); +lean_dec(x_12); +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_EXPORT lean_object* l_Lean_Meta_Grind_intros_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, 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) { +_start: +{ +lean_object* x_34; +lean_inc(x_32); +lean_inc(x_31); +lean_inc(x_30); +lean_inc(x_29); +lean_inc(x_28); +lean_inc(x_27); +lean_inc(x_26); +lean_inc(x_2); +lean_inc(x_1); +x_34 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext(x_1, x_2, x_26, x_27, x_28, x_29, x_30, x_31, x_32, x_33); +if (lean_obj_tag(x_34) == 0) +{ +lean_object* x_35; +x_35 = lean_ctor_get(x_34, 0); +lean_inc(x_35); +switch (lean_obj_tag(x_35)) { +case 0: +{ +lean_object* x_36; lean_object* x_37; +x_36 = lean_ctor_get(x_34, 1); +lean_inc(x_36); +lean_dec(x_34); +lean_inc(x_32); +lean_inc(x_31); +lean_inc(x_30); +lean_inc(x_29); +x_37 = l_Lean_MVarId_byContra_x3f(x_3, x_29, x_30, x_31, x_32, x_36); +if (lean_obj_tag(x_37) == 0) +{ +lean_object* x_38; +x_38 = lean_ctor_get(x_37, 0); +lean_inc(x_38); +if (lean_obj_tag(x_38) == 0) +{ +lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; uint8_t x_45; +lean_dec(x_32); +lean_dec(x_31); +lean_dec(x_30); +lean_dec(x_29); +lean_dec(x_28); +lean_dec(x_27); +lean_dec(x_26); +lean_dec(x_23); +lean_dec(x_22); +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +x_39 = lean_ctor_get(x_37, 1); +lean_inc(x_39); +lean_dec(x_37); +x_40 = lean_st_ref_take(x_25, x_39); +x_41 = lean_ctor_get(x_40, 0); +lean_inc(x_41); +x_42 = lean_ctor_get(x_40, 1); +lean_inc(x_42); +lean_dec(x_40); +x_43 = lean_array_push(x_41, x_1); +x_44 = lean_st_ref_set(x_25, x_43, x_42); +x_45 = !lean_is_exclusive(x_44); +if (x_45 == 0) +{ +lean_object* x_46; lean_object* x_47; +x_46 = lean_ctor_get(x_44, 0); +lean_dec(x_46); +x_47 = lean_box(0); +lean_ctor_set(x_44, 0, x_47); +return x_44; +} +else +{ +lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_48 = lean_ctor_get(x_44, 1); +lean_inc(x_48); +lean_dec(x_44); +x_49 = lean_box(0); +x_50 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_50, 0, x_49); +lean_ctor_set(x_50, 1, x_48); +return x_50; +} +} +else +{ +lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; +lean_dec(x_1); +x_51 = lean_ctor_get(x_37, 1); +lean_inc(x_51); +lean_dec(x_37); +x_52 = lean_ctor_get(x_38, 0); +lean_inc(x_52); +lean_dec(x_38); +x_53 = lean_alloc_ctor(0, 20, 1); +lean_ctor_set(x_53, 0, x_52); +lean_ctor_set(x_53, 1, x_4); +lean_ctor_set(x_53, 2, x_5); +lean_ctor_set(x_53, 3, x_6); +lean_ctor_set(x_53, 4, x_7); +lean_ctor_set(x_53, 5, x_8); +lean_ctor_set(x_53, 6, x_10); +lean_ctor_set(x_53, 7, x_11); +lean_ctor_set(x_53, 8, x_12); +lean_ctor_set(x_53, 9, x_13); +lean_ctor_set(x_53, 10, x_14); +lean_ctor_set(x_53, 11, x_15); +lean_ctor_set(x_53, 12, x_16); +lean_ctor_set(x_53, 13, x_17); +lean_ctor_set(x_53, 14, x_18); +lean_ctor_set(x_53, 15, x_19); +lean_ctor_set(x_53, 16, x_20); +lean_ctor_set(x_53, 17, x_21); +lean_ctor_set(x_53, 18, x_22); +lean_ctor_set(x_53, 19, x_23); +lean_ctor_set_uint8(x_53, sizeof(void*)*20, x_9); +x_54 = l_Lean_Meta_Grind_intros_go(x_2, x_53, x_25, x_26, x_27, x_28, x_29, x_30, x_31, x_32, x_51); +return x_54; +} +} +else +{ +uint8_t x_55; +lean_dec(x_32); +lean_dec(x_31); +lean_dec(x_30); +lean_dec(x_29); +lean_dec(x_28); +lean_dec(x_27); +lean_dec(x_26); +lean_dec(x_23); +lean_dec(x_22); +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_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_55 = !lean_is_exclusive(x_37); +if (x_55 == 0) +{ +return x_37; } else { -lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_20 = lean_ctor_get(x_17, 0); -x_21 = lean_ctor_get(x_17, 1); -lean_inc(x_21); -lean_inc(x_20); -lean_dec(x_17); -lean_ctor_set(x_14, 1, x_20); -x_22 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_22, 0, x_14); -lean_ctor_set(x_22, 1, x_21); -return x_22; +lean_object* x_56; lean_object* x_57; lean_object* x_58; +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_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 +} +case 1: { -lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_23 = lean_ctor_get(x_14, 0); -x_24 = lean_ctor_get(x_14, 1); -lean_inc(x_24); -lean_inc(x_23); +lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; +lean_dec(x_23); +lean_dec(x_22); +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); lean_dec(x_14); -x_25 = lean_st_ref_get(x_3, x_24); +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_3); -x_26 = lean_ctor_get(x_25, 0); -lean_inc(x_26); -x_27 = lean_ctor_get(x_25, 1); +lean_dec(x_1); +x_59 = lean_ctor_get(x_34, 1); +lean_inc(x_59); +lean_dec(x_34); +x_60 = lean_ctor_get(x_35, 0); +lean_inc(x_60); +x_61 = lean_ctor_get(x_35, 1); +lean_inc(x_61); +lean_dec(x_35); +lean_inc(x_32); +lean_inc(x_31); +lean_inc(x_30); +lean_inc(x_29); +lean_inc(x_60); +lean_inc(x_61); +x_62 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_applyCases_x3f(x_61, x_60, x_29, x_30, x_31, x_32, x_59); +if (lean_obj_tag(x_62) == 0) +{ +lean_object* x_63; +x_63 = lean_ctor_get(x_62, 0); +lean_inc(x_63); +if (lean_obj_tag(x_63) == 0) +{ +lean_object* x_64; lean_object* x_65; +x_64 = lean_ctor_get(x_62, 1); +lean_inc(x_64); +lean_dec(x_62); +lean_inc(x_32); +lean_inc(x_31); +lean_inc(x_30); +lean_inc(x_29); +lean_inc(x_60); +lean_inc(x_61); +x_65 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_applyInjection_x3f(x_61, x_60, x_29, x_30, x_31, x_32, x_64); +if (lean_obj_tag(x_65) == 0) +{ +lean_object* x_66; +x_66 = lean_ctor_get(x_65, 0); +lean_inc(x_66); +if (lean_obj_tag(x_66) == 0) +{ +lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; +x_67 = lean_ctor_get(x_65, 1); +lean_inc(x_67); +lean_dec(x_65); +x_68 = lean_ctor_get(x_61, 0); +lean_inc(x_68); +x_69 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__2___boxed), 9, 1); +lean_closure_set(x_69, 0, x_61); +lean_inc(x_2); +x_70 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_intros_go___lambda__1), 11, 2); +lean_closure_set(x_70, 0, x_60); +lean_closure_set(x_70, 1, x_2); +x_71 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); +lean_closure_set(x_71, 0, x_69); +lean_closure_set(x_71, 1, x_70); +x_72 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__1; +x_73 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); +lean_closure_set(x_73, 0, x_71); +lean_closure_set(x_73, 1, x_72); +lean_inc(x_32); +lean_inc(x_31); +lean_inc(x_30); +lean_inc(x_29); +lean_inc(x_28); 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); -} -x_29 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_29, 0, x_23); -lean_ctor_set(x_29, 1, x_26); -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; -} +lean_inc(x_26); +x_74 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(x_68, x_73, x_26, x_27, x_28, x_29, x_30, x_31, x_32, x_67); +if (lean_obj_tag(x_74) == 0) +{ +lean_object* x_75; lean_object* x_76; lean_object* x_77; +x_75 = lean_ctor_get(x_74, 0); +lean_inc(x_75); +x_76 = lean_ctor_get(x_74, 1); +lean_inc(x_76); +lean_dec(x_74); +x_77 = l_Lean_Meta_Grind_intros_go(x_2, x_75, x_25, x_26, x_27, x_28, x_29, x_30, x_31, x_32, x_76); +return x_77; } else { -uint8_t x_31; -lean_dec(x_3); -x_31 = !lean_is_exclusive(x_12); -if (x_31 == 0) +uint8_t x_78; +lean_dec(x_32); +lean_dec(x_31); +lean_dec(x_30); +lean_dec(x_29); +lean_dec(x_28); +lean_dec(x_27); +lean_dec(x_26); +lean_dec(x_2); +x_78 = !lean_is_exclusive(x_74); +if (x_78 == 0) { -return x_12; +return x_74; } else { -lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_32 = lean_ctor_get(x_12, 0); -x_33 = lean_ctor_get(x_12, 1); -lean_inc(x_33); -lean_inc(x_32); -lean_dec(x_12); -x_34 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_34, 0, x_32); -lean_ctor_set(x_34, 1, x_33); -return x_34; +lean_object* x_79; lean_object* x_80; lean_object* x_81; +x_79 = lean_ctor_get(x_74, 0); +x_80 = lean_ctor_get(x_74, 1); +lean_inc(x_80); +lean_inc(x_79); +lean_dec(x_74); +x_81 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_81, 0, x_79); +lean_ctor_set(x_81, 1, x_80); +return x_81; +} } } +else +{ +lean_object* x_82; lean_object* x_83; lean_object* x_84; +lean_dec(x_61); +lean_dec(x_60); +x_82 = lean_ctor_get(x_65, 1); +lean_inc(x_82); +lean_dec(x_65); +x_83 = lean_ctor_get(x_66, 0); +lean_inc(x_83); +lean_dec(x_66); +x_84 = l_Lean_Meta_Grind_intros_go(x_2, x_83, x_25, x_26, x_27, x_28, x_29, x_30, x_31, x_32, x_82); +return x_84; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_intros_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: +else { -uint8_t x_10; -x_10 = !lean_is_exclusive(x_1); -if (x_10 == 0) +uint8_t x_85; +lean_dec(x_61); +lean_dec(x_60); +lean_dec(x_32); +lean_dec(x_31); +lean_dec(x_30); +lean_dec(x_29); +lean_dec(x_28); +lean_dec(x_27); +lean_dec(x_26); +lean_dec(x_2); +x_85 = !lean_is_exclusive(x_65); +if (x_85 == 0) { -lean_object* x_11; -x_11 = lean_ctor_get(x_1, 1); -lean_dec(x_11); -lean_ctor_set(x_1, 1, x_9); -return x_1; +return x_65; } else { -lean_object* x_12; lean_object* x_13; -x_12 = lean_ctor_get(x_1, 0); -lean_inc(x_12); -lean_dec(x_1); -x_13 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_13, 0, x_12); -lean_ctor_set(x_13, 1, x_9); -return x_13; +lean_object* x_86; lean_object* x_87; lean_object* x_88; +x_86 = lean_ctor_get(x_65, 0); +x_87 = lean_ctor_get(x_65, 1); +lean_inc(x_87); +lean_inc(x_86); +lean_dec(x_65); +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; } } } -static lean_object* _init_l_Lean_Meta_Grind_intros_go___lambda__4___closed__1() { -_start: +else { -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_intros_go___lambda__3___boxed), 9, 0); -return x_1; +lean_object* x_89; lean_object* x_90; lean_object* x_91; +lean_dec(x_61); +lean_dec(x_60); +x_89 = lean_ctor_get(x_62, 1); +lean_inc(x_89); +lean_dec(x_62); +x_90 = lean_ctor_get(x_63, 0); +lean_inc(x_90); +lean_dec(x_63); +x_91 = l_List_forM___at_Lean_Meta_Grind_intros_go___spec__1(x_2, x_90, x_25, x_26, x_27, x_28, x_29, x_30, x_31, x_32, x_89); +return x_91; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_intros_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, 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) { -_start: +else { -lean_object* x_28; -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_1); -x_28 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext(x_1, x_20, x_21, x_22, x_23, x_24, x_25, x_26, x_27); -if (lean_obj_tag(x_28) == 0) +uint8_t x_92; +lean_dec(x_61); +lean_dec(x_60); +lean_dec(x_32); +lean_dec(x_31); +lean_dec(x_30); +lean_dec(x_29); +lean_dec(x_28); +lean_dec(x_27); +lean_dec(x_26); +lean_dec(x_2); +x_92 = !lean_is_exclusive(x_62); +if (x_92 == 0) { -lean_object* x_29; -x_29 = lean_ctor_get(x_28, 0); -lean_inc(x_29); -switch (lean_obj_tag(x_29)) { -case 0: +return x_62; +} +else { -lean_object* x_30; lean_object* x_31; -x_30 = lean_ctor_get(x_28, 1); -lean_inc(x_30); -lean_dec(x_28); -lean_inc(x_26); -lean_inc(x_25); -lean_inc(x_24); -lean_inc(x_23); -x_31 = l_Lean_MVarId_byContra_x3f(x_2, x_23, x_24, x_25, x_26, x_30); -if (lean_obj_tag(x_31) == 0) +lean_object* x_93; lean_object* x_94; lean_object* x_95; +x_93 = lean_ctor_get(x_62, 0); +x_94 = lean_ctor_get(x_62, 1); +lean_inc(x_94); +lean_inc(x_93); +lean_dec(x_62); +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; +} +} +} +case 2: { -lean_object* x_32; -x_32 = lean_ctor_get(x_31, 0); -lean_inc(x_32); -if (lean_obj_tag(x_32) == 0) +lean_object* x_96; lean_object* x_97; lean_object* x_98; +lean_dec(x_23); +lean_dec(x_22); +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_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_96 = lean_ctor_get(x_34, 1); +lean_inc(x_96); +lean_dec(x_34); +x_97 = lean_ctor_get(x_35, 0); +lean_inc(x_97); +lean_dec(x_35); +x_98 = l_Lean_Meta_Grind_intros_go(x_2, x_97, x_25, x_26, x_27, x_28, x_29, x_30, x_31, x_32, x_96); +return x_98; +} +default: { -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_dec(x_26); -lean_dec(x_25); -lean_dec(x_24); +lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_dec(x_23); lean_dec(x_22); lean_dec(x_21); lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); @@ -3243,86 +4746,105 @@ 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_33 = lean_ctor_get(x_31, 1); -lean_inc(x_33); -lean_dec(x_31); -x_34 = lean_st_ref_take(x_19, x_33); -x_35 = lean_ctor_get(x_34, 0); -lean_inc(x_35); -x_36 = lean_ctor_get(x_34, 1); -lean_inc(x_36); +lean_dec(x_1); +x_99 = lean_ctor_get(x_34, 1); +lean_inc(x_99); lean_dec(x_34); -x_37 = lean_array_push(x_35, x_1); -x_38 = lean_st_ref_set(x_19, x_37, x_36); -x_39 = !lean_is_exclusive(x_38); -if (x_39 == 0) +x_100 = lean_ctor_get(x_35, 0); +lean_inc(x_100); +x_101 = lean_ctor_get(x_35, 1); +lean_inc(x_101); +lean_dec(x_35); +lean_inc(x_32); +lean_inc(x_31); +lean_inc(x_30); +lean_inc(x_29); +lean_inc(x_101); +x_102 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_applyCases_x3f(x_101, x_100, x_29, x_30, x_31, x_32, x_99); +if (lean_obj_tag(x_102) == 0) { -lean_object* x_40; lean_object* x_41; -x_40 = lean_ctor_get(x_38, 0); -lean_dec(x_40); -x_41 = lean_box(0); -lean_ctor_set(x_38, 0, x_41); -return x_38; +lean_object* x_103; +x_103 = lean_ctor_get(x_102, 0); +lean_inc(x_103); +if (lean_obj_tag(x_103) == 0) +{ +lean_object* x_104; lean_object* x_105; +x_104 = lean_ctor_get(x_102, 1); +lean_inc(x_104); +lean_dec(x_102); +x_105 = l_Lean_Meta_Grind_intros_go(x_2, x_101, x_25, x_26, x_27, x_28, x_29, x_30, x_31, x_32, x_104); +return x_105; } else { -lean_object* x_42; lean_object* x_43; lean_object* x_44; -x_42 = lean_ctor_get(x_38, 1); -lean_inc(x_42); -lean_dec(x_38); -x_43 = lean_box(0); -x_44 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_44, 0, x_43); -lean_ctor_set(x_44, 1, x_42); -return x_44; +lean_object* x_106; lean_object* x_107; lean_object* x_108; +lean_dec(x_101); +x_106 = lean_ctor_get(x_102, 1); +lean_inc(x_106); +lean_dec(x_102); +x_107 = lean_ctor_get(x_103, 0); +lean_inc(x_107); +lean_dec(x_103); +x_108 = l_List_forM___at_Lean_Meta_Grind_intros_go___spec__2(x_2, x_107, x_25, x_26, x_27, x_28, x_29, x_30, x_31, x_32, x_106); +return x_108; } } else { -lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; -lean_dec(x_1); -x_45 = lean_ctor_get(x_31, 1); -lean_inc(x_45); -lean_dec(x_31); -x_46 = lean_ctor_get(x_32, 0); -lean_inc(x_46); +uint8_t x_109; +lean_dec(x_101); lean_dec(x_32); -x_47 = lean_alloc_ctor(0, 14, 1); -lean_ctor_set(x_47, 0, x_46); -lean_ctor_set(x_47, 1, x_3); -lean_ctor_set(x_47, 2, x_4); -lean_ctor_set(x_47, 3, x_5); -lean_ctor_set(x_47, 4, x_6); -lean_ctor_set(x_47, 5, x_7); -lean_ctor_set(x_47, 6, x_9); -lean_ctor_set(x_47, 7, x_10); -lean_ctor_set(x_47, 8, x_11); -lean_ctor_set(x_47, 9, x_12); -lean_ctor_set(x_47, 10, x_13); -lean_ctor_set(x_47, 11, x_14); -lean_ctor_set(x_47, 12, x_15); -lean_ctor_set(x_47, 13, x_16); -lean_ctor_set_uint8(x_47, sizeof(void*)*14, x_8); -x_48 = l_Lean_Meta_Grind_intros_go(x_17, x_47, x_19, x_20, x_21, x_22, x_23, x_24, x_25, x_26, x_45); -return x_48; +lean_dec(x_31); +lean_dec(x_30); +lean_dec(x_29); +lean_dec(x_28); +lean_dec(x_27); +lean_dec(x_26); +lean_dec(x_2); +x_109 = !lean_is_exclusive(x_102); +if (x_109 == 0) +{ +return x_102; +} +else +{ +lean_object* x_110; lean_object* x_111; lean_object* x_112; +x_110 = lean_ctor_get(x_102, 0); +x_111 = lean_ctor_get(x_102, 1); +lean_inc(x_111); +lean_inc(x_110); +lean_dec(x_102); +x_112 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_112, 0, x_110); +lean_ctor_set(x_112, 1, x_111); +return x_112; +} +} +} } } else { -uint8_t x_49; +uint8_t x_113; +lean_dec(x_32); +lean_dec(x_31); +lean_dec(x_30); +lean_dec(x_29); +lean_dec(x_28); +lean_dec(x_27); lean_dec(x_26); -lean_dec(x_25); -lean_dec(x_24); lean_dec(x_23); lean_dec(x_22); lean_dec(x_21); lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); @@ -3331,467 +4853,440 @@ lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); -lean_dec(x_9); +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); +lean_dec(x_2); lean_dec(x_1); -x_49 = !lean_is_exclusive(x_31); -if (x_49 == 0) +x_113 = !lean_is_exclusive(x_34); +if (x_113 == 0) { -return x_31; +return x_34; } else { -lean_object* x_50; lean_object* x_51; lean_object* x_52; -x_50 = lean_ctor_get(x_31, 0); -x_51 = lean_ctor_get(x_31, 1); -lean_inc(x_51); -lean_inc(x_50); -lean_dec(x_31); -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_114; lean_object* x_115; lean_object* x_116; +x_114 = lean_ctor_get(x_34, 0); +x_115 = lean_ctor_get(x_34, 1); +lean_inc(x_115); +lean_inc(x_114); +lean_dec(x_34); +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; } } } -case 1: +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_intros_go(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: { -lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); +uint8_t x_12; +x_12 = lean_ctor_get_uint8(x_2, sizeof(void*)*20); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_13 = lean_ctor_get(x_2, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_2, 1); +lean_inc(x_14); +x_15 = lean_ctor_get(x_2, 2); +lean_inc(x_15); +x_16 = lean_ctor_get(x_2, 3); +lean_inc(x_16); +x_17 = lean_ctor_get(x_2, 4); +lean_inc(x_17); +x_18 = lean_ctor_get(x_2, 5); +lean_inc(x_18); +x_19 = lean_ctor_get(x_2, 6); +lean_inc(x_19); +x_20 = lean_ctor_get(x_2, 7); +lean_inc(x_20); +x_21 = lean_ctor_get(x_2, 8); +lean_inc(x_21); +x_22 = lean_ctor_get(x_2, 9); +lean_inc(x_22); +x_23 = lean_ctor_get(x_2, 10); +lean_inc(x_23); +x_24 = lean_ctor_get(x_2, 11); +lean_inc(x_24); +x_25 = lean_ctor_get(x_2, 12); +lean_inc(x_25); +x_26 = lean_ctor_get(x_2, 13); +lean_inc(x_26); +x_27 = lean_ctor_get(x_2, 14); +lean_inc(x_27); +x_28 = lean_ctor_get(x_2, 15); +lean_inc(x_28); +x_29 = lean_ctor_get(x_2, 16); +lean_inc(x_29); +x_30 = lean_ctor_get(x_2, 17); +lean_inc(x_30); +x_31 = lean_ctor_get(x_2, 18); +lean_inc(x_31); +x_32 = lean_ctor_get(x_2, 19); +lean_inc(x_32); +x_33 = lean_box(0); +x_34 = l_Lean_Meta_Grind_intros_go___lambda__2(x_2, x_1, x_13, x_14, x_15, x_16, x_17, x_18, x_12, x_19, x_20, x_21, x_22, x_23, x_24, x_25, x_26, x_27, x_28, x_29, x_30, x_31, x_32, x_33, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +return x_34; +} +else +{ +lean_object* x_35; lean_object* x_36; lean_dec(x_10); lean_dec(x_9); +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_53 = lean_ctor_get(x_28, 1); -lean_inc(x_53); -lean_dec(x_28); -x_54 = lean_ctor_get(x_29, 0); -lean_inc(x_54); -x_55 = lean_ctor_get(x_29, 1); -lean_inc(x_55); -lean_dec(x_29); -lean_inc(x_26); -lean_inc(x_25); -lean_inc(x_24); -lean_inc(x_23); -lean_inc(x_54); -lean_inc(x_55); -x_56 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_applyCases_x3f(x_55, x_54, x_23, x_24, x_25, x_26, x_53); -if (lean_obj_tag(x_56) == 0) -{ -lean_object* x_57; -x_57 = lean_ctor_get(x_56, 0); -lean_inc(x_57); -if (lean_obj_tag(x_57) == 0) -{ -lean_object* x_58; lean_object* x_59; -x_58 = lean_ctor_get(x_56, 1); -lean_inc(x_58); -lean_dec(x_56); -lean_inc(x_26); -lean_inc(x_25); -lean_inc(x_24); -lean_inc(x_23); -lean_inc(x_54); -lean_inc(x_55); -x_59 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_applyInjection_x3f(x_55, x_54, x_23, x_24, x_25, x_26, x_58); -if (lean_obj_tag(x_59) == 0) -{ -lean_object* x_60; -x_60 = lean_ctor_get(x_59, 0); -lean_inc(x_60); -if (lean_obj_tag(x_60) == 0) -{ -lean_object* x_61; lean_object* x_62; 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(x_59, 1); -lean_inc(x_61); -lean_dec(x_59); -x_62 = lean_ctor_get(x_55, 0); -lean_inc(x_62); -x_63 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_intros_go___lambda__1___boxed), 9, 1); -lean_closure_set(x_63, 0, x_55); -lean_inc(x_17); -x_64 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_intros_go___lambda__2), 11, 2); -lean_closure_set(x_64, 0, x_54); -lean_closure_set(x_64, 1, x_17); -x_65 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); -lean_closure_set(x_65, 0, x_63); -lean_closure_set(x_65, 1, x_64); -x_66 = l_Lean_Meta_Grind_intros_go___lambda__4___closed__1; -x_67 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); -lean_closure_set(x_67, 0, x_65); -lean_closure_set(x_67, 1, x_66); -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); -x_68 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(x_62, x_67, x_20, x_21, x_22, x_23, x_24, x_25, x_26, x_61); -if (lean_obj_tag(x_68) == 0) -{ -lean_object* x_69; lean_object* x_70; lean_object* x_71; -x_69 = lean_ctor_get(x_68, 0); -lean_inc(x_69); -x_70 = lean_ctor_get(x_68, 1); -lean_inc(x_70); -lean_dec(x_68); -x_71 = l_Lean_Meta_Grind_intros_go(x_17, x_69, x_19, x_20, x_21, x_22, x_23, x_24, x_25, x_26, x_70); -return x_71; +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_11); +return x_36; } -else +} +} +LEAN_EXPORT lean_object* l_List_forM___at_Lean_Meta_Grind_intros_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, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: { -uint8_t x_72; -lean_dec(x_26); -lean_dec(x_25); -lean_dec(x_24); -lean_dec(x_23); -lean_dec(x_22); -lean_dec(x_21); -lean_dec(x_20); -lean_dec(x_17); -x_72 = !lean_is_exclusive(x_68); -if (x_72 == 0) +lean_object* x_12; +x_12 = l_List_forM___at_Lean_Meta_Grind_intros_go___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_3); +return x_12; +} +} +LEAN_EXPORT lean_object* l_List_forM___at_Lean_Meta_Grind_intros_go___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: { -return x_68; +lean_object* x_12; +x_12 = l_List_forM___at_Lean_Meta_Grind_intros_go___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_3); +return x_12; } -else +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_intros_go___lambda__2___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +lean_object* x_18 = _args[17]; +lean_object* x_19 = _args[18]; +lean_object* x_20 = _args[19]; +lean_object* x_21 = _args[20]; +lean_object* x_22 = _args[21]; +lean_object* x_23 = _args[22]; +lean_object* x_24 = _args[23]; +lean_object* x_25 = _args[24]; +lean_object* x_26 = _args[25]; +lean_object* x_27 = _args[26]; +lean_object* x_28 = _args[27]; +lean_object* x_29 = _args[28]; +lean_object* x_30 = _args[29]; +lean_object* x_31 = _args[30]; +lean_object* x_32 = _args[31]; +lean_object* x_33 = _args[32]; +_start: { -lean_object* x_73; lean_object* x_74; lean_object* x_75; -x_73 = lean_ctor_get(x_68, 0); -x_74 = lean_ctor_get(x_68, 1); -lean_inc(x_74); -lean_inc(x_73); -lean_dec(x_68); -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; +uint8_t x_34; lean_object* x_35; +x_34 = lean_unbox(x_9); +lean_dec(x_9); +x_35 = l_Lean_Meta_Grind_intros_go___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_34, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_23, x_24, x_25, x_26, x_27, x_28, x_29, x_30, x_31, x_32, x_33); +lean_dec(x_25); +lean_dec(x_24); +return x_35; +} } +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_intros_go___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; +x_12 = l_Lean_Meta_Grind_intros_go(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_3); +return x_12; } } -else +static lean_object* _init_l_Lean_Meta_Grind_intros___closed__1() { +_start: { -lean_object* x_76; lean_object* x_77; lean_object* x_78; -lean_dec(x_55); -lean_dec(x_54); -x_76 = lean_ctor_get(x_59, 1); -lean_inc(x_76); -lean_dec(x_59); -x_77 = lean_ctor_get(x_60, 0); -lean_inc(x_77); -lean_dec(x_60); -x_78 = l_Lean_Meta_Grind_intros_go(x_17, x_77, x_19, x_20, x_21, x_22, x_23, x_24, x_25, x_26, x_76); -return x_78; +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(0); +x_2 = lean_array_mk(x_1); +return x_2; } } -else +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_intros(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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_79; -lean_dec(x_55); -lean_dec(x_54); -lean_dec(x_26); -lean_dec(x_25); -lean_dec(x_24); -lean_dec(x_23); -lean_dec(x_22); -lean_dec(x_21); -lean_dec(x_20); -lean_dec(x_17); -x_79 = !lean_is_exclusive(x_59); -if (x_79 == 0) +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_11 = l_Lean_Meta_Grind_intros___closed__1; +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_inc(x_14); +lean_dec(x_12); +x_15 = l_Lean_Meta_Grind_intros_go(x_1, x_2, x_13, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_14); +if (lean_obj_tag(x_15) == 0) { -return x_59; -} -else +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 = lean_st_ref_get(x_13, x_16); +lean_dec(x_13); +x_18 = !lean_is_exclusive(x_17); +if (x_18 == 0) { -lean_object* x_80; lean_object* x_81; lean_object* x_82; -x_80 = lean_ctor_get(x_59, 0); -x_81 = lean_ctor_get(x_59, 1); -lean_inc(x_81); -lean_inc(x_80); -lean_dec(x_59); -x_82 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_82, 0, x_80); -lean_ctor_set(x_82, 1, x_81); -return x_82; -} -} +lean_object* x_19; lean_object* x_20; +x_19 = lean_ctor_get(x_17, 0); +x_20 = lean_array_to_list(x_19); +lean_ctor_set(x_17, 0, x_20); +return x_17; } else { -lean_object* x_83; lean_object* x_84; lean_object* x_85; -lean_dec(x_55); -lean_dec(x_54); -x_83 = lean_ctor_get(x_56, 1); -lean_inc(x_83); -lean_dec(x_56); -x_84 = lean_ctor_get(x_57, 0); -lean_inc(x_84); -lean_dec(x_57); -x_85 = l_List_forM___at_Lean_Meta_Grind_intros_go___spec__1(x_17, x_84, x_19, x_20, x_21, x_22, x_23, x_24, x_25, x_26, x_83); -return x_85; +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_21 = lean_ctor_get(x_17, 0); +x_22 = lean_ctor_get(x_17, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_17); +x_23 = lean_array_to_list(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; } } else { -uint8_t x_86; -lean_dec(x_55); -lean_dec(x_54); -lean_dec(x_26); -lean_dec(x_25); -lean_dec(x_24); -lean_dec(x_23); -lean_dec(x_22); -lean_dec(x_21); -lean_dec(x_20); -lean_dec(x_17); -x_86 = !lean_is_exclusive(x_56); -if (x_86 == 0) +uint8_t x_25; +lean_dec(x_13); +x_25 = !lean_is_exclusive(x_15); +if (x_25 == 0) { -return x_56; +return x_15; } else { -lean_object* x_87; lean_object* x_88; lean_object* x_89; -x_87 = lean_ctor_get(x_56, 0); -x_88 = lean_ctor_get(x_56, 1); -lean_inc(x_88); -lean_inc(x_87); -lean_dec(x_56); -x_89 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_89, 0, x_87); -lean_ctor_set(x_89, 1, x_88); -return x_89; +lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_26 = lean_ctor_get(x_15, 0); +x_27 = lean_ctor_get(x_15, 1); +lean_inc(x_27); +lean_inc(x_26); +lean_dec(x_15); +x_28 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_28, 0, x_26); +lean_ctor_set(x_28, 1, x_27); +return x_28; } } } -case 2: -{ -lean_object* x_90; lean_object* x_91; lean_object* x_92; -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_90 = lean_ctor_get(x_28, 1); -lean_inc(x_90); -lean_dec(x_28); -x_91 = lean_ctor_get(x_29, 0); -lean_inc(x_91); -lean_dec(x_29); -x_92 = l_Lean_Meta_Grind_intros_go(x_17, x_91, x_19, x_20, x_21, x_22, x_23, x_24, x_25, x_26, x_90); -return x_92; } -default: +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_assertAt___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: { -lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); +lean_object* x_13; +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_6); +x_13 = l_Lean_Meta_Grind_simp(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_13) == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); lean_dec(x_13); -lean_dec(x_12); -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); -lean_dec(x_1); -x_93 = lean_ctor_get(x_28, 1); -lean_inc(x_93); -lean_dec(x_28); -x_94 = lean_ctor_get(x_29, 0); -lean_inc(x_94); -x_95 = lean_ctor_get(x_29, 1); -lean_inc(x_95); -lean_dec(x_29); -lean_inc(x_26); -lean_inc(x_25); +x_16 = lean_ctor_get(x_14, 0); +lean_inc(x_16); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +x_17 = l_Lean_Meta_Simp_Result_getProof(x_14, x_8, x_9, x_10, x_11, x_15); +if (lean_obj_tag(x_17) == 0) +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +x_19 = lean_ctor_get(x_17, 1); +lean_inc(x_19); +lean_dec(x_17); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +x_20 = l_Lean_Meta_mkEqMP(x_18, x_2, x_8, x_9, x_10, x_11, x_19); +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); +lean_inc(x_4); +x_23 = l_Lean_Meta_Grind_add(x_16, x_21, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_22); +if (lean_obj_tag(x_23) == 0) +{ +lean_object* x_24; lean_object* x_25; uint8_t x_26; +x_24 = lean_ctor_get(x_23, 1); lean_inc(x_24); -lean_inc(x_23); -lean_inc(x_95); -x_96 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_applyCases_x3f(x_95, x_94, x_23, x_24, x_25, x_26, x_93); -if (lean_obj_tag(x_96) == 0) +lean_dec(x_23); +x_25 = lean_st_ref_get(x_4, x_24); +x_26 = !lean_is_exclusive(x_25); +if (x_26 == 0) { -lean_object* x_97; -x_97 = lean_ctor_get(x_96, 0); -lean_inc(x_97); -if (lean_obj_tag(x_97) == 0) +lean_object* x_27; lean_object* x_28; uint8_t x_29; +x_27 = lean_ctor_get(x_25, 1); +x_28 = lean_st_ref_get(x_4, x_27); +lean_dec(x_4); +x_29 = !lean_is_exclusive(x_28); +if (x_29 == 0) { -lean_object* x_98; lean_object* x_99; -x_98 = lean_ctor_get(x_96, 1); -lean_inc(x_98); -lean_dec(x_96); -x_99 = l_Lean_Meta_Grind_intros_go(x_17, x_95, x_19, x_20, x_21, x_22, x_23, x_24, x_25, x_26, x_98); -return x_99; +lean_object* x_30; +x_30 = lean_ctor_get(x_28, 0); +lean_ctor_set(x_25, 1, x_30); +lean_ctor_set(x_28, 0, x_25); +return x_28; } else { -lean_object* x_100; lean_object* x_101; lean_object* x_102; -lean_dec(x_95); -x_100 = lean_ctor_get(x_96, 1); -lean_inc(x_100); -lean_dec(x_96); -x_101 = lean_ctor_get(x_97, 0); -lean_inc(x_101); -lean_dec(x_97); -x_102 = l_List_forM___at_Lean_Meta_Grind_intros_go___spec__2(x_17, x_101, x_19, x_20, x_21, x_22, x_23, x_24, x_25, x_26, x_100); -return x_102; +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_28, 0); +x_32 = lean_ctor_get(x_28, 1); +lean_inc(x_32); +lean_inc(x_31); +lean_dec(x_28); +lean_ctor_set(x_25, 1, x_31); +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_25); +lean_ctor_set(x_33, 1, x_32); +return x_33; } } else { -uint8_t x_103; -lean_dec(x_95); -lean_dec(x_26); +lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_34 = lean_ctor_get(x_25, 0); +x_35 = lean_ctor_get(x_25, 1); +lean_inc(x_35); +lean_inc(x_34); lean_dec(x_25); -lean_dec(x_24); -lean_dec(x_23); -lean_dec(x_22); -lean_dec(x_21); -lean_dec(x_20); -lean_dec(x_17); -x_103 = !lean_is_exclusive(x_96); -if (x_103 == 0) -{ -return x_96; +x_36 = lean_st_ref_get(x_4, x_35); +lean_dec(x_4); +x_37 = lean_ctor_get(x_36, 0); +lean_inc(x_37); +x_38 = lean_ctor_get(x_36, 1); +lean_inc(x_38); +if (lean_is_exclusive(x_36)) { + lean_ctor_release(x_36, 0); + lean_ctor_release(x_36, 1); + x_39 = x_36; +} else { + lean_dec_ref(x_36); + x_39 = lean_box(0); +} +x_40 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_40, 0, x_34); +lean_ctor_set(x_40, 1, x_37); +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; +} } else { -lean_object* x_104; lean_object* x_105; lean_object* x_106; -x_104 = lean_ctor_get(x_96, 0); -x_105 = lean_ctor_get(x_96, 1); -lean_inc(x_105); -lean_inc(x_104); -lean_dec(x_96); -x_106 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_106, 0, x_104); -lean_ctor_set(x_106, 1, x_105); -return x_106; -} +uint8_t x_42; +lean_dec(x_4); +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 { -uint8_t x_107; -lean_dec(x_26); -lean_dec(x_25); -lean_dec(x_24); -lean_dec(x_23); -lean_dec(x_22); -lean_dec(x_21); -lean_dec(x_20); -lean_dec(x_17); +uint8_t x_46; lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); +lean_dec(x_8); lean_dec(x_7); 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_107 = !lean_is_exclusive(x_28); -if (x_107 == 0) +x_46 = !lean_is_exclusive(x_20); +if (x_46 == 0) { -return x_28; +return x_20; } else { -lean_object* x_108; lean_object* x_109; lean_object* x_110; -x_108 = lean_ctor_get(x_28, 0); -x_109 = lean_ctor_get(x_28, 1); -lean_inc(x_109); -lean_inc(x_108); -lean_dec(x_28); -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_47; lean_object* x_48; lean_object* x_49; +x_47 = lean_ctor_get(x_20, 0); +x_48 = lean_ctor_get(x_20, 1); +lean_inc(x_48); +lean_inc(x_47); +lean_dec(x_20); +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_Meta_Grind_intros_go(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -uint8_t x_12; -x_12 = lean_ctor_get_uint8(x_2, sizeof(void*)*14); -if (x_12 == 0) -{ -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_13 = lean_ctor_get(x_2, 0); -lean_inc(x_13); -x_14 = lean_ctor_get(x_2, 1); -lean_inc(x_14); -x_15 = lean_ctor_get(x_2, 2); -lean_inc(x_15); -x_16 = lean_ctor_get(x_2, 3); -lean_inc(x_16); -x_17 = lean_ctor_get(x_2, 4); -lean_inc(x_17); -x_18 = lean_ctor_get(x_2, 5); -lean_inc(x_18); -x_19 = lean_ctor_get(x_2, 6); -lean_inc(x_19); -x_20 = lean_ctor_get(x_2, 7); -lean_inc(x_20); -x_21 = lean_ctor_get(x_2, 8); -lean_inc(x_21); -x_22 = lean_ctor_get(x_2, 9); -lean_inc(x_22); -x_23 = lean_ctor_get(x_2, 10); -lean_inc(x_23); -x_24 = lean_ctor_get(x_2, 11); -lean_inc(x_24); -x_25 = lean_ctor_get(x_2, 12); -lean_inc(x_25); -x_26 = lean_ctor_get(x_2, 13); -lean_inc(x_26); -x_27 = lean_box(0); -x_28 = l_Lean_Meta_Grind_intros_go___lambda__4(x_2, x_13, x_14, x_15, x_16, x_17, x_18, x_12, x_19, x_20, x_21, x_22, x_23, x_24, x_25, x_26, x_1, x_27, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_28; } else { -lean_object* x_29; lean_object* x_30; +uint8_t x_50; +lean_dec(x_16); +lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -3799,39 +5294,34 @@ 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_box(0); -x_30 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_30, 0, x_29); -lean_ctor_set(x_30, 1, x_11); -return x_30; -} -} -} -LEAN_EXPORT lean_object* l_List_forM___at_Lean_Meta_Grind_intros_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, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: +x_50 = !lean_is_exclusive(x_17); +if (x_50 == 0) { -lean_object* x_12; -x_12 = l_List_forM___at_Lean_Meta_Grind_intros_go___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -lean_dec(x_3); -return x_12; -} +return x_17; } -LEAN_EXPORT lean_object* l_List_forM___at_Lean_Meta_Grind_intros_go___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: +else { -lean_object* x_12; -x_12 = l_List_forM___at_Lean_Meta_Grind_intros_go___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -lean_dec(x_3); -return x_12; +lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_51 = lean_ctor_get(x_17, 0); +x_52 = lean_ctor_get(x_17, 1); +lean_inc(x_52); +lean_inc(x_51); +lean_dec(x_17); +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_EXPORT lean_object* l_Lean_Meta_Grind_intros_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: +} +else { -lean_object* x_10; -x_10 = l_Lean_Meta_Grind_intros_go___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +uint8_t x_54; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -3839,231 +5329,353 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -return x_10; -} +x_54 = !lean_is_exclusive(x_13); +if (x_54 == 0) +{ +return x_13; } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_intros_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: +else { -lean_object* x_10; -x_10 = l_Lean_Meta_Grind_intros_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); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_10; +lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_55 = lean_ctor_get(x_13, 0); +x_56 = lean_ctor_get(x_13, 1); +lean_inc(x_56); +lean_inc(x_55); +lean_dec(x_13); +x_57 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_57, 0, x_55); +lean_ctor_set(x_57, 1, x_56); +return x_57; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_intros_go___lambda__4___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_object* x_19 = _args[18]; -lean_object* x_20 = _args[19]; -lean_object* x_21 = _args[20]; -lean_object* x_22 = _args[21]; -lean_object* x_23 = _args[22]; -lean_object* x_24 = _args[23]; -lean_object* x_25 = _args[24]; -lean_object* x_26 = _args[25]; -lean_object* x_27 = _args[26]; -_start: -{ -uint8_t x_28; lean_object* x_29; -x_28 = lean_unbox(x_8); -lean_dec(x_8); -x_29 = l_Lean_Meta_Grind_intros_go___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_28, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_23, x_24, x_25, x_26, x_27); -lean_dec(x_19); -lean_dec(x_18); -return x_29; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_intros_go___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +static lean_object* _init_l_Lean_Meta_Grind_assertAt___closed__1() { _start: { -lean_object* x_12; -x_12 = l_Lean_Meta_Grind_intros_go(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -lean_dec(x_3); -return x_12; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("h", 1, 1); +return x_1; } } -static lean_object* _init_l_Lean_Meta_Grind_intros___closed__1() { +static lean_object* _init_l_Lean_Meta_Grind_assertAt___closed__2() { _start: { -lean_object* x_1; lean_object* x_2; +lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = lean_array_mk(x_1); -return x_2; +x_2 = l_Lean_Meta_Grind_assertAt___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_intros(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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_Grind_assertAt(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_11 = l_Lean_Meta_Grind_intros___closed__1; -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_13; lean_object* x_14; uint8_t x_15; +x_13 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_isCasesCandidate(x_2, x_8, x_9, x_10, x_11, x_12); +x_14 = lean_ctor_get(x_13, 0); lean_inc(x_14); -lean_dec(x_12); -x_15 = l_Lean_Meta_Grind_intros_go(x_2, x_1, x_13, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_14); -if (lean_obj_tag(x_15) == 0) +x_15 = lean_unbox(x_14); +lean_dec(x_14); +if (x_15 == 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 = lean_st_ref_get(x_13, x_16); -lean_dec(x_13); -x_18 = !lean_is_exclusive(x_17); -if (x_18 == 0) +uint8_t x_16; +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; +x_17 = lean_ctor_get(x_13, 1); +x_18 = lean_ctor_get(x_13, 0); +lean_dec(x_18); +x_19 = lean_ctor_get(x_4, 0); +lean_inc(x_19); +x_20 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__2___boxed), 9, 1); +lean_closure_set(x_20, 0, x_4); +x_21 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_assertAt___lambda__1), 12, 3); +lean_closure_set(x_21, 0, x_2); +lean_closure_set(x_21, 1, x_1); +lean_closure_set(x_21, 2, x_3); +x_22 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); +lean_closure_set(x_22, 0, x_20); +lean_closure_set(x_22, 1, x_21); +x_23 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__1; +x_24 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); +lean_closure_set(x_24, 0, x_22); +lean_closure_set(x_24, 1, x_23); +x_25 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(x_19, x_24, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_17); +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_ctor_get_uint8(x_26, sizeof(void*)*20); +if (x_27 == 0) +{ +uint8_t x_28; +x_28 = !lean_is_exclusive(x_25); +if (x_28 == 0) +{ +lean_object* x_29; lean_object* x_30; +x_29 = lean_ctor_get(x_25, 0); +lean_dec(x_29); +x_30 = lean_box(0); +lean_ctor_set_tag(x_13, 1); +lean_ctor_set(x_13, 1, x_30); +lean_ctor_set(x_13, 0, x_26); +lean_ctor_set(x_25, 0, x_13); +return x_25; +} +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_25, 1); +lean_inc(x_31); +lean_dec(x_25); +x_32 = lean_box(0); +lean_ctor_set_tag(x_13, 1); +lean_ctor_set(x_13, 1, x_32); +lean_ctor_set(x_13, 0, x_26); +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_13); +lean_ctor_set(x_33, 1, x_31); +return x_33; +} +} +else +{ +uint8_t x_34; +lean_dec(x_26); +lean_free_object(x_13); +x_34 = !lean_is_exclusive(x_25); +if (x_34 == 0) +{ +lean_object* x_35; lean_object* x_36; +x_35 = lean_ctor_get(x_25, 0); +lean_dec(x_35); +x_36 = lean_box(0); +lean_ctor_set(x_25, 0, x_36); +return x_25; +} +else +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_37 = lean_ctor_get(x_25, 1); +lean_inc(x_37); +lean_dec(x_25); +x_38 = lean_box(0); +x_39 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_37); +return x_39; +} +} +} +else { -lean_object* x_19; lean_object* x_20; -x_19 = lean_ctor_get(x_17, 0); -x_20 = lean_array_to_list(x_19); -lean_ctor_set(x_17, 0, x_20); -return x_17; +uint8_t x_40; +lean_free_object(x_13); +x_40 = !lean_is_exclusive(x_25); +if (x_40 == 0) +{ +return x_25; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_21 = lean_ctor_get(x_17, 0); -x_22 = lean_ctor_get(x_17, 1); -lean_inc(x_22); -lean_inc(x_21); -lean_dec(x_17); -x_23 = lean_array_to_list(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_object* x_41; lean_object* x_42; lean_object* x_43; +x_41 = lean_ctor_get(x_25, 0); +x_42 = lean_ctor_get(x_25, 1); +lean_inc(x_42); +lean_inc(x_41); +lean_dec(x_25); +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_25; +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_44 = lean_ctor_get(x_13, 1); +lean_inc(x_44); lean_dec(x_13); -x_25 = !lean_is_exclusive(x_15); -if (x_25 == 0) +x_45 = lean_ctor_get(x_4, 0); +lean_inc(x_45); +x_46 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__2___boxed), 9, 1); +lean_closure_set(x_46, 0, x_4); +x_47 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_assertAt___lambda__1), 12, 3); +lean_closure_set(x_47, 0, x_2); +lean_closure_set(x_47, 1, x_1); +lean_closure_set(x_47, 2, x_3); +x_48 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); +lean_closure_set(x_48, 0, x_46); +lean_closure_set(x_48, 1, x_47); +x_49 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__1; +x_50 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); +lean_closure_set(x_50, 0, x_48); +lean_closure_set(x_50, 1, x_49); +x_51 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(x_45, x_50, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_44); +if (lean_obj_tag(x_51) == 0) { -return x_15; +lean_object* x_52; uint8_t x_53; +x_52 = lean_ctor_get(x_51, 0); +lean_inc(x_52); +x_53 = lean_ctor_get_uint8(x_52, sizeof(void*)*20); +if (x_53 == 0) +{ +lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; +x_54 = lean_ctor_get(x_51, 1); +lean_inc(x_54); +if (lean_is_exclusive(x_51)) { + lean_ctor_release(x_51, 0); + lean_ctor_release(x_51, 1); + x_55 = x_51; +} else { + lean_dec_ref(x_51); + x_55 = lean_box(0); +} +x_56 = lean_box(0); +x_57 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_57, 0, x_52); +lean_ctor_set(x_57, 1, x_56); +if (lean_is_scalar(x_55)) { + x_58 = lean_alloc_ctor(0, 2, 0); +} else { + x_58 = x_55; +} +lean_ctor_set(x_58, 0, x_57); +lean_ctor_set(x_58, 1, x_54); +return x_58; } else { -lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_26 = lean_ctor_get(x_15, 0); -x_27 = lean_ctor_get(x_15, 1); -lean_inc(x_27); -lean_inc(x_26); -lean_dec(x_15); -x_28 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set(x_28, 1, x_27); -return x_28; +lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; +lean_dec(x_52); +x_59 = lean_ctor_get(x_51, 1); +lean_inc(x_59); +if (lean_is_exclusive(x_51)) { + lean_ctor_release(x_51, 0); + lean_ctor_release(x_51, 1); + x_60 = x_51; +} else { + lean_dec_ref(x_51); + x_60 = lean_box(0); } +x_61 = lean_box(0); +if (lean_is_scalar(x_60)) { + x_62 = lean_alloc_ctor(0, 2, 0); +} else { + x_62 = x_60; } +lean_ctor_set(x_62, 0, x_61); +lean_ctor_set(x_62, 1, x_59); +return x_62; } } -static lean_object* _init_l_Lean_Meta_Grind_assertAt___closed__1() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("h", 1, 1); -return x_1; +lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_63 = lean_ctor_get(x_51, 0); +lean_inc(x_63); +x_64 = lean_ctor_get(x_51, 1); +lean_inc(x_64); +if (lean_is_exclusive(x_51)) { + lean_ctor_release(x_51, 0); + lean_ctor_release(x_51, 1); + x_65 = x_51; +} else { + lean_dec_ref(x_51); + x_65 = lean_box(0); } +if (lean_is_scalar(x_65)) { + x_66 = lean_alloc_ctor(1, 2, 0); +} else { + x_66 = x_65; } -static lean_object* _init_l_Lean_Meta_Grind_assertAt___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_Grind_assertAt___closed__1; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; +lean_ctor_set(x_66, 0, x_63); +lean_ctor_set(x_66, 1, x_64); +return x_66; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_assertAt(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: +} +else { -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; -x_13 = l_Lean_Meta_Grind_assertAt___closed__2; -x_14 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_13, x_10, x_11, x_12); -x_15 = lean_ctor_get(x_14, 0); -lean_inc(x_15); -x_16 = lean_ctor_get(x_14, 1); -lean_inc(x_16); -lean_dec(x_14); -x_17 = !lean_is_exclusive(x_1); -if (x_17 == 0) +lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; uint8_t x_72; +x_67 = lean_ctor_get(x_13, 1); +lean_inc(x_67); +lean_dec(x_13); +x_68 = l_Lean_Meta_Grind_assertAt___closed__2; +x_69 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_68, x_10, x_11, x_67); +x_70 = lean_ctor_get(x_69, 0); +lean_inc(x_70); +x_71 = lean_ctor_get(x_69, 1); +lean_inc(x_71); +lean_dec(x_69); +x_72 = !lean_is_exclusive(x_4); +if (x_72 == 0) { -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; -x_18 = lean_ctor_get(x_1, 0); -x_19 = lean_ctor_get(x_1, 1); -x_20 = lean_ctor_get(x_1, 2); -x_21 = lean_ctor_get(x_1, 3); -x_22 = lean_ctor_get(x_1, 4); -x_23 = lean_ctor_get(x_1, 5); -x_24 = lean_ctor_get(x_1, 6); -x_25 = lean_ctor_get(x_1, 7); -x_26 = lean_ctor_get(x_1, 8); -x_27 = lean_ctor_get(x_1, 9); -x_28 = lean_ctor_get(x_1, 10); -x_29 = lean_ctor_get(x_1, 11); -x_30 = lean_ctor_get(x_1, 12); -x_31 = lean_ctor_get(x_1, 13); +lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; +x_73 = lean_ctor_get(x_4, 0); +x_74 = lean_ctor_get(x_4, 1); +x_75 = lean_ctor_get(x_4, 2); +x_76 = lean_ctor_get(x_4, 3); +x_77 = lean_ctor_get(x_4, 4); +x_78 = lean_ctor_get(x_4, 5); +x_79 = lean_ctor_get(x_4, 6); +x_80 = lean_ctor_get(x_4, 7); +x_81 = lean_ctor_get(x_4, 8); +x_82 = lean_ctor_get(x_4, 9); +x_83 = lean_ctor_get(x_4, 10); +x_84 = lean_ctor_get(x_4, 11); +x_85 = lean_ctor_get(x_4, 12); +x_86 = lean_ctor_get(x_4, 13); +x_87 = lean_ctor_get(x_4, 14); +x_88 = lean_ctor_get(x_4, 15); +x_89 = lean_ctor_get(x_4, 16); +x_90 = lean_ctor_get(x_4, 17); +x_91 = lean_ctor_get(x_4, 18); +x_92 = lean_ctor_get(x_4, 19); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); -x_32 = l_Lean_MVarId_assert(x_18, x_15, x_3, x_2, x_8, x_9, x_10, x_11, x_16); -if (lean_obj_tag(x_32) == 0) +x_93 = l_Lean_MVarId_assert(x_73, x_70, x_2, x_1, x_8, x_9, x_10, x_11, x_71); +if (lean_obj_tag(x_93) == 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_ctor_set(x_1, 0, x_33); -x_35 = l_Lean_Meta_Grind_intros(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_34); -return x_35; +lean_object* x_94; lean_object* x_95; lean_object* x_96; +x_94 = lean_ctor_get(x_93, 0); +lean_inc(x_94); +x_95 = lean_ctor_get(x_93, 1); +lean_inc(x_95); +lean_dec(x_93); +lean_ctor_set(x_4, 0, x_94); +x_96 = l_Lean_Meta_Grind_intros(x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_95); +return x_96; } else { -uint8_t x_36; -lean_free_object(x_1); -lean_dec(x_31); -lean_dec(x_30); -lean_dec(x_29); -lean_dec(x_28); -lean_dec(x_27); -lean_dec(x_26); -lean_dec(x_25); -lean_dec(x_24); -lean_dec(x_23); -lean_dec(x_22); -lean_dec(x_21); -lean_dec(x_20); -lean_dec(x_19); +uint8_t x_97; +lean_free_object(x_4); +lean_dec(x_92); +lean_dec(x_91); +lean_dec(x_90); +lean_dec(x_89); +lean_dec(x_88); +lean_dec(x_87); +lean_dec(x_86); +lean_dec(x_85); +lean_dec(x_84); +lean_dec(x_83); +lean_dec(x_82); +lean_dec(x_81); +lean_dec(x_80); +lean_dec(x_79); +lean_dec(x_78); +lean_dec(x_77); +lean_dec(x_76); +lean_dec(x_75); +lean_dec(x_74); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -4071,108 +5683,132 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); -x_36 = !lean_is_exclusive(x_32); -if (x_36 == 0) +lean_dec(x_3); +x_97 = !lean_is_exclusive(x_93); +if (x_97 == 0) { -return x_32; +return x_93; } 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_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_object* x_98; lean_object* x_99; lean_object* x_100; +x_98 = lean_ctor_get(x_93, 0); +x_99 = lean_ctor_get(x_93, 1); +lean_inc(x_99); +lean_inc(x_98); +lean_dec(x_93); +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_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; uint8_t x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; -x_40 = lean_ctor_get(x_1, 0); -x_41 = lean_ctor_get(x_1, 1); -x_42 = lean_ctor_get(x_1, 2); -x_43 = lean_ctor_get(x_1, 3); -x_44 = lean_ctor_get(x_1, 4); -x_45 = lean_ctor_get(x_1, 5); -x_46 = lean_ctor_get_uint8(x_1, sizeof(void*)*14); -x_47 = lean_ctor_get(x_1, 6); -x_48 = lean_ctor_get(x_1, 7); -x_49 = lean_ctor_get(x_1, 8); -x_50 = lean_ctor_get(x_1, 9); -x_51 = lean_ctor_get(x_1, 10); -x_52 = lean_ctor_get(x_1, 11); -x_53 = lean_ctor_get(x_1, 12); -x_54 = lean_ctor_get(x_1, 13); -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_45); -lean_inc(x_44); -lean_inc(x_43); -lean_inc(x_42); -lean_inc(x_41); -lean_inc(x_40); -lean_dec(x_1); +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; lean_object* x_120; lean_object* x_121; lean_object* x_122; +x_101 = lean_ctor_get(x_4, 0); +x_102 = lean_ctor_get(x_4, 1); +x_103 = lean_ctor_get(x_4, 2); +x_104 = lean_ctor_get(x_4, 3); +x_105 = lean_ctor_get(x_4, 4); +x_106 = lean_ctor_get(x_4, 5); +x_107 = lean_ctor_get_uint8(x_4, sizeof(void*)*20); +x_108 = lean_ctor_get(x_4, 6); +x_109 = lean_ctor_get(x_4, 7); +x_110 = lean_ctor_get(x_4, 8); +x_111 = lean_ctor_get(x_4, 9); +x_112 = lean_ctor_get(x_4, 10); +x_113 = lean_ctor_get(x_4, 11); +x_114 = lean_ctor_get(x_4, 12); +x_115 = lean_ctor_get(x_4, 13); +x_116 = lean_ctor_get(x_4, 14); +x_117 = lean_ctor_get(x_4, 15); +x_118 = lean_ctor_get(x_4, 16); +x_119 = lean_ctor_get(x_4, 17); +x_120 = lean_ctor_get(x_4, 18); +x_121 = lean_ctor_get(x_4, 19); +lean_inc(x_121); +lean_inc(x_120); +lean_inc(x_119); +lean_inc(x_118); +lean_inc(x_117); +lean_inc(x_116); +lean_inc(x_115); +lean_inc(x_114); +lean_inc(x_113); +lean_inc(x_112); +lean_inc(x_111); +lean_inc(x_110); +lean_inc(x_109); +lean_inc(x_108); +lean_inc(x_106); +lean_inc(x_105); +lean_inc(x_104); +lean_inc(x_103); +lean_inc(x_102); +lean_inc(x_101); +lean_dec(x_4); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); -x_55 = l_Lean_MVarId_assert(x_40, x_15, x_3, x_2, x_8, x_9, x_10, x_11, x_16); -if (lean_obj_tag(x_55) == 0) +x_122 = l_Lean_MVarId_assert(x_101, x_70, x_2, x_1, x_8, x_9, x_10, x_11, x_71); +if (lean_obj_tag(x_122) == 0) +{ +lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; +x_123 = lean_ctor_get(x_122, 0); +lean_inc(x_123); +x_124 = lean_ctor_get(x_122, 1); +lean_inc(x_124); +lean_dec(x_122); +x_125 = lean_alloc_ctor(0, 20, 1); +lean_ctor_set(x_125, 0, x_123); +lean_ctor_set(x_125, 1, x_102); +lean_ctor_set(x_125, 2, x_103); +lean_ctor_set(x_125, 3, x_104); +lean_ctor_set(x_125, 4, x_105); +lean_ctor_set(x_125, 5, x_106); +lean_ctor_set(x_125, 6, x_108); +lean_ctor_set(x_125, 7, x_109); +lean_ctor_set(x_125, 8, x_110); +lean_ctor_set(x_125, 9, x_111); +lean_ctor_set(x_125, 10, x_112); +lean_ctor_set(x_125, 11, x_113); +lean_ctor_set(x_125, 12, x_114); +lean_ctor_set(x_125, 13, x_115); +lean_ctor_set(x_125, 14, x_116); +lean_ctor_set(x_125, 15, x_117); +lean_ctor_set(x_125, 16, x_118); +lean_ctor_set(x_125, 17, x_119); +lean_ctor_set(x_125, 18, x_120); +lean_ctor_set(x_125, 19, x_121); +lean_ctor_set_uint8(x_125, sizeof(void*)*20, x_107); +x_126 = l_Lean_Meta_Grind_intros(x_3, x_125, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_124); +return x_126; +} +else { -lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; -x_56 = lean_ctor_get(x_55, 0); -lean_inc(x_56); -x_57 = lean_ctor_get(x_55, 1); -lean_inc(x_57); -lean_dec(x_55); -x_58 = lean_alloc_ctor(0, 14, 1); -lean_ctor_set(x_58, 0, x_56); -lean_ctor_set(x_58, 1, x_41); -lean_ctor_set(x_58, 2, x_42); -lean_ctor_set(x_58, 3, x_43); -lean_ctor_set(x_58, 4, x_44); -lean_ctor_set(x_58, 5, x_45); -lean_ctor_set(x_58, 6, x_47); -lean_ctor_set(x_58, 7, x_48); -lean_ctor_set(x_58, 8, x_49); -lean_ctor_set(x_58, 9, x_50); -lean_ctor_set(x_58, 10, x_51); -lean_ctor_set(x_58, 11, x_52); -lean_ctor_set(x_58, 12, x_53); -lean_ctor_set(x_58, 13, x_54); -lean_ctor_set_uint8(x_58, sizeof(void*)*14, x_46); -x_59 = l_Lean_Meta_Grind_intros(x_58, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_57); -return x_59; -} -else -{ -lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; -lean_dec(x_54); -lean_dec(x_53); -lean_dec(x_52); -lean_dec(x_51); -lean_dec(x_50); -lean_dec(x_49); -lean_dec(x_48); -lean_dec(x_47); -lean_dec(x_45); -lean_dec(x_44); -lean_dec(x_43); -lean_dec(x_42); -lean_dec(x_41); +lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; +lean_dec(x_121); +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_114); +lean_dec(x_113); +lean_dec(x_112); +lean_dec(x_111); +lean_dec(x_110); +lean_dec(x_109); +lean_dec(x_108); +lean_dec(x_106); +lean_dec(x_105); +lean_dec(x_104); +lean_dec(x_103); +lean_dec(x_102); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -4180,39 +5816,40 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); -x_60 = lean_ctor_get(x_55, 0); -lean_inc(x_60); -x_61 = lean_ctor_get(x_55, 1); -lean_inc(x_61); -if (lean_is_exclusive(x_55)) { - lean_ctor_release(x_55, 0); - lean_ctor_release(x_55, 1); - x_62 = x_55; +lean_dec(x_3); +x_127 = lean_ctor_get(x_122, 0); +lean_inc(x_127); +x_128 = lean_ctor_get(x_122, 1); +lean_inc(x_128); +if (lean_is_exclusive(x_122)) { + lean_ctor_release(x_122, 0); + lean_ctor_release(x_122, 1); + x_129 = x_122; } else { - lean_dec_ref(x_55); - x_62 = lean_box(0); + lean_dec_ref(x_122); + x_129 = lean_box(0); } -if (lean_is_scalar(x_62)) { - x_63 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_129)) { + x_130 = lean_alloc_ctor(1, 2, 0); } else { - x_63 = x_62; + x_130 = x_129; } -lean_ctor_set(x_63, 0, x_60); -lean_ctor_set(x_63, 1, x_61); -return x_63; +lean_ctor_set(x_130, 0, x_127); +lean_ctor_set(x_130, 1, x_128); +return x_130; } } } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_assertNext_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_EXPORT lean_object* l_Lean_Meta_Grind_assertNext(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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_1); 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; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; 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_11 = lean_ctor_get(x_1, 0); x_12 = lean_ctor_get(x_1, 1); x_13 = lean_ctor_get(x_1, 2); @@ -4227,11 +5864,23 @@ x_21 = lean_ctor_get(x_1, 10); x_22 = lean_ctor_get(x_1, 11); x_23 = lean_ctor_get(x_1, 12); x_24 = lean_ctor_get(x_1, 13); -x_25 = l_Std_Queue_dequeue_x3f___rarg(x_24); -if (lean_obj_tag(x_25) == 0) +x_25 = lean_ctor_get(x_1, 14); +x_26 = lean_ctor_get(x_1, 15); +x_27 = lean_ctor_get(x_1, 16); +x_28 = lean_ctor_get(x_1, 17); +x_29 = lean_ctor_get(x_1, 18); +x_30 = lean_ctor_get(x_1, 19); +x_31 = l_Std_Queue_dequeue_x3f___rarg(x_25); +if (lean_obj_tag(x_31) == 0) { -lean_object* x_26; lean_object* x_27; +lean_object* x_32; lean_object* x_33; lean_free_object(x_1); +lean_dec(x_30); +lean_dec(x_29); +lean_dec(x_28); +lean_dec(x_27); +lean_dec(x_26); +lean_dec(x_24); lean_dec(x_23); lean_dec(x_22); lean_dec(x_21); @@ -4252,208 +5901,226 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -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_9); -return x_27; +x_32 = lean_box(0); +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_9); +return x_33; } else { -uint8_t x_28; -x_28 = !lean_is_exclusive(x_25); -if (x_28 == 0) +uint8_t x_34; +x_34 = !lean_is_exclusive(x_31); +if (x_34 == 0) { -lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_29 = lean_ctor_get(x_25, 0); -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_ctor_set(x_1, 13, x_31); -x_32 = lean_ctor_get(x_30, 0); -lean_inc(x_32); -x_33 = lean_ctor_get(x_30, 1); -lean_inc(x_33); -x_34 = lean_ctor_get(x_30, 2); -lean_inc(x_34); -lean_dec(x_30); -x_35 = l_Lean_Meta_Grind_assertAt(x_1, x_32, x_33, x_34, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_35) == 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_31, 0); +x_36 = lean_ctor_get(x_35, 0); +lean_inc(x_36); +x_37 = lean_ctor_get(x_35, 1); +lean_inc(x_37); +lean_dec(x_35); +x_38 = lean_ctor_get(x_36, 0); +lean_inc(x_38); +x_39 = lean_ctor_get(x_36, 1); +lean_inc(x_39); +x_40 = lean_ctor_get(x_36, 2); +lean_inc(x_40); +lean_dec(x_36); +lean_ctor_set(x_1, 14, x_37); +x_41 = l_Lean_Meta_Grind_assertAt(x_38, x_39, x_40, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_41) == 0) { -uint8_t x_36; -x_36 = !lean_is_exclusive(x_35); -if (x_36 == 0) +uint8_t x_42; +x_42 = !lean_is_exclusive(x_41); +if (x_42 == 0) { -lean_object* x_37; -x_37 = lean_ctor_get(x_35, 0); -lean_ctor_set(x_25, 0, x_37); -lean_ctor_set(x_35, 0, x_25); -return x_35; +lean_object* x_43; +x_43 = lean_ctor_get(x_41, 0); +lean_ctor_set(x_31, 0, x_43); +lean_ctor_set(x_41, 0, x_31); +return x_41; } else { -lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_38 = lean_ctor_get(x_35, 0); -x_39 = lean_ctor_get(x_35, 1); -lean_inc(x_39); -lean_inc(x_38); -lean_dec(x_35); -lean_ctor_set(x_25, 0, x_38); -x_40 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_40, 0, x_25); -lean_ctor_set(x_40, 1, x_39); -return x_40; +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_41, 0); +x_45 = lean_ctor_get(x_41, 1); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_41); +lean_ctor_set(x_31, 0, x_44); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_31); +lean_ctor_set(x_46, 1, x_45); +return x_46; } } else { -uint8_t x_41; -lean_free_object(x_25); -x_41 = !lean_is_exclusive(x_35); -if (x_41 == 0) -{ -return x_35; -} -else +uint8_t x_47; +lean_free_object(x_31); +x_47 = !lean_is_exclusive(x_41); +if (x_47 == 0) { -lean_object* x_42; lean_object* x_43; lean_object* x_44; -x_42 = lean_ctor_get(x_35, 0); -x_43 = lean_ctor_get(x_35, 1); -lean_inc(x_43); -lean_inc(x_42); -lean_dec(x_35); -x_44 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_44, 0, x_42); -lean_ctor_set(x_44, 1, x_43); -return x_44; -} -} +return x_41; } else { -lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; -x_45 = lean_ctor_get(x_25, 0); -lean_inc(x_45); -lean_dec(x_25); -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); -lean_ctor_set(x_1, 13, x_47); -x_48 = lean_ctor_get(x_46, 0); -lean_inc(x_48); -x_49 = lean_ctor_get(x_46, 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); -x_50 = lean_ctor_get(x_46, 2); -lean_inc(x_50); -lean_dec(x_46); -x_51 = l_Lean_Meta_Grind_assertAt(x_1, x_48, x_49, x_50, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_51) == 0) +lean_inc(x_48); +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 { -lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; +lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_51 = lean_ctor_get(x_31, 0); +lean_inc(x_51); +lean_dec(x_31); x_52 = lean_ctor_get(x_51, 0); lean_inc(x_52); x_53 = lean_ctor_get(x_51, 1); lean_inc(x_53); -if (lean_is_exclusive(x_51)) { - lean_ctor_release(x_51, 0); - lean_ctor_release(x_51, 1); - x_54 = x_51; +lean_dec(x_51); +x_54 = lean_ctor_get(x_52, 0); +lean_inc(x_54); +x_55 = lean_ctor_get(x_52, 1); +lean_inc(x_55); +x_56 = lean_ctor_get(x_52, 2); +lean_inc(x_56); +lean_dec(x_52); +lean_ctor_set(x_1, 14, x_53); +x_57 = l_Lean_Meta_Grind_assertAt(x_54, x_55, x_56, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_57) == 0) +{ +lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; +x_58 = lean_ctor_get(x_57, 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 { - lean_dec_ref(x_51); - x_54 = lean_box(0); + lean_dec_ref(x_57); + x_60 = lean_box(0); } -x_55 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_55, 0, x_52); -if (lean_is_scalar(x_54)) { - x_56 = lean_alloc_ctor(0, 2, 0); +x_61 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_61, 0, x_58); +if (lean_is_scalar(x_60)) { + x_62 = lean_alloc_ctor(0, 2, 0); } else { - x_56 = x_54; + x_62 = x_60; } -lean_ctor_set(x_56, 0, x_55); -lean_ctor_set(x_56, 1, x_53); -return x_56; +lean_ctor_set(x_62, 0, x_61); +lean_ctor_set(x_62, 1, x_59); +return x_62; } else { -lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; -x_57 = lean_ctor_get(x_51, 0); -lean_inc(x_57); -x_58 = lean_ctor_get(x_51, 1); -lean_inc(x_58); -if (lean_is_exclusive(x_51)) { - lean_ctor_release(x_51, 0); - lean_ctor_release(x_51, 1); - x_59 = x_51; +lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_63 = lean_ctor_get(x_57, 0); +lean_inc(x_63); +x_64 = lean_ctor_get(x_57, 1); +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; } else { - lean_dec_ref(x_51); - x_59 = lean_box(0); + lean_dec_ref(x_57); + x_65 = lean_box(0); } -if (lean_is_scalar(x_59)) { - x_60 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_65)) { + x_66 = lean_alloc_ctor(1, 2, 0); } else { - x_60 = x_59; + x_66 = x_65; } -lean_ctor_set(x_60, 0, x_57); -lean_ctor_set(x_60, 1, x_58); -return x_60; +lean_ctor_set(x_66, 0, x_63); +lean_ctor_set(x_66, 1, x_64); +return x_66; } } } } 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; lean_object* x_68; 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_61 = lean_ctor_get(x_1, 0); -x_62 = lean_ctor_get(x_1, 1); -x_63 = lean_ctor_get(x_1, 2); -x_64 = lean_ctor_get(x_1, 3); -x_65 = lean_ctor_get(x_1, 4); -x_66 = lean_ctor_get(x_1, 5); -x_67 = lean_ctor_get_uint8(x_1, sizeof(void*)*14); -x_68 = lean_ctor_get(x_1, 6); -x_69 = lean_ctor_get(x_1, 7); -x_70 = lean_ctor_get(x_1, 8); -x_71 = lean_ctor_get(x_1, 9); -x_72 = lean_ctor_get(x_1, 10); -x_73 = lean_ctor_get(x_1, 11); -x_74 = lean_ctor_get(x_1, 12); -x_75 = lean_ctor_get(x_1, 13); +lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; uint8_t x_73; lean_object* x_74; 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_67 = lean_ctor_get(x_1, 0); +x_68 = lean_ctor_get(x_1, 1); +x_69 = lean_ctor_get(x_1, 2); +x_70 = lean_ctor_get(x_1, 3); +x_71 = lean_ctor_get(x_1, 4); +x_72 = lean_ctor_get(x_1, 5); +x_73 = lean_ctor_get_uint8(x_1, sizeof(void*)*20); +x_74 = lean_ctor_get(x_1, 6); +x_75 = lean_ctor_get(x_1, 7); +x_76 = lean_ctor_get(x_1, 8); +x_77 = lean_ctor_get(x_1, 9); +x_78 = lean_ctor_get(x_1, 10); +x_79 = lean_ctor_get(x_1, 11); +x_80 = lean_ctor_get(x_1, 12); +x_81 = lean_ctor_get(x_1, 13); +x_82 = lean_ctor_get(x_1, 14); +x_83 = lean_ctor_get(x_1, 15); +x_84 = lean_ctor_get(x_1, 16); +x_85 = lean_ctor_get(x_1, 17); +x_86 = lean_ctor_get(x_1, 18); +x_87 = lean_ctor_get(x_1, 19); +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_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_inc(x_71); lean_inc(x_70); lean_inc(x_69); lean_inc(x_68); -lean_inc(x_66); -lean_inc(x_65); -lean_inc(x_64); -lean_inc(x_63); -lean_inc(x_62); -lean_inc(x_61); +lean_inc(x_67); lean_dec(x_1); -x_76 = l_Std_Queue_dequeue_x3f___rarg(x_75); -if (lean_obj_tag(x_76) == 0) -{ -lean_object* x_77; lean_object* x_78; +x_88 = l_Std_Queue_dequeue_x3f___rarg(x_82); +if (lean_obj_tag(x_88) == 0) +{ +lean_object* x_89; lean_object* x_90; +lean_dec(x_87); +lean_dec(x_86); +lean_dec(x_85); +lean_dec(x_84); +lean_dec(x_83); +lean_dec(x_81); +lean_dec(x_80); +lean_dec(x_79); +lean_dec(x_78); +lean_dec(x_77); +lean_dec(x_76); +lean_dec(x_75); lean_dec(x_74); -lean_dec(x_73); lean_dec(x_72); lean_dec(x_71); lean_dec(x_70); lean_dec(x_69); lean_dec(x_68); -lean_dec(x_66); -lean_dec(x_65); -lean_dec(x_64); -lean_dec(x_63); -lean_dec(x_62); -lean_dec(x_61); +lean_dec(x_67); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -4461,333 +6128,123 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -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_9); -return x_78; +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_9); +return x_90; } 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; -x_79 = lean_ctor_get(x_76, 0); -lean_inc(x_79); -if (lean_is_exclusive(x_76)) { - lean_ctor_release(x_76, 0); - x_80 = x_76; -} else { - lean_dec_ref(x_76); - x_80 = lean_box(0); -} -x_81 = lean_ctor_get(x_79, 0); -lean_inc(x_81); -x_82 = lean_ctor_get(x_79, 1); -lean_inc(x_82); -lean_dec(x_79); -x_83 = lean_alloc_ctor(0, 14, 1); -lean_ctor_set(x_83, 0, x_61); -lean_ctor_set(x_83, 1, x_62); -lean_ctor_set(x_83, 2, x_63); -lean_ctor_set(x_83, 3, x_64); -lean_ctor_set(x_83, 4, x_65); -lean_ctor_set(x_83, 5, x_66); -lean_ctor_set(x_83, 6, x_68); -lean_ctor_set(x_83, 7, x_69); -lean_ctor_set(x_83, 8, x_70); -lean_ctor_set(x_83, 9, x_71); -lean_ctor_set(x_83, 10, x_72); -lean_ctor_set(x_83, 11, x_73); -lean_ctor_set(x_83, 12, x_74); -lean_ctor_set(x_83, 13, x_82); -lean_ctor_set_uint8(x_83, sizeof(void*)*14, x_67); -x_84 = lean_ctor_get(x_81, 0); -lean_inc(x_84); -x_85 = lean_ctor_get(x_81, 1); -lean_inc(x_85); -x_86 = lean_ctor_get(x_81, 2); -lean_inc(x_86); -lean_dec(x_81); -x_87 = l_Lean_Meta_Grind_assertAt(x_83, x_84, x_85, x_86, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -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); -if (lean_is_exclusive(x_87)) { - lean_ctor_release(x_87, 0); - lean_ctor_release(x_87, 1); - x_90 = x_87; -} else { - lean_dec_ref(x_87); - x_90 = lean_box(0); -} -if (lean_is_scalar(x_80)) { - x_91 = lean_alloc_ctor(1, 1, 0); -} else { - x_91 = x_80; -} -lean_ctor_set(x_91, 0, x_88); -if (lean_is_scalar(x_90)) { - x_92 = lean_alloc_ctor(0, 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; +x_91 = lean_ctor_get(x_88, 0); +lean_inc(x_91); +if (lean_is_exclusive(x_88)) { + lean_ctor_release(x_88, 0); + x_92 = x_88; } else { - x_92 = x_90; -} -lean_ctor_set(x_92, 0, x_91); -lean_ctor_set(x_92, 1, x_89); -return x_92; + lean_dec_ref(x_88); + x_92 = lean_box(0); } -else -{ -lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; -lean_dec(x_80); -x_93 = lean_ctor_get(x_87, 0); +x_93 = lean_ctor_get(x_91, 0); lean_inc(x_93); -x_94 = lean_ctor_get(x_87, 1); +x_94 = lean_ctor_get(x_91, 1); lean_inc(x_94); -if (lean_is_exclusive(x_87)) { - lean_ctor_release(x_87, 0); - lean_ctor_release(x_87, 1); - x_95 = x_87; +lean_dec(x_91); +x_95 = lean_ctor_get(x_93, 0); +lean_inc(x_95); +x_96 = lean_ctor_get(x_93, 1); +lean_inc(x_96); +x_97 = lean_ctor_get(x_93, 2); +lean_inc(x_97); +lean_dec(x_93); +x_98 = lean_alloc_ctor(0, 20, 1); +lean_ctor_set(x_98, 0, x_67); +lean_ctor_set(x_98, 1, x_68); +lean_ctor_set(x_98, 2, x_69); +lean_ctor_set(x_98, 3, x_70); +lean_ctor_set(x_98, 4, x_71); +lean_ctor_set(x_98, 5, x_72); +lean_ctor_set(x_98, 6, x_74); +lean_ctor_set(x_98, 7, x_75); +lean_ctor_set(x_98, 8, x_76); +lean_ctor_set(x_98, 9, x_77); +lean_ctor_set(x_98, 10, x_78); +lean_ctor_set(x_98, 11, x_79); +lean_ctor_set(x_98, 12, x_80); +lean_ctor_set(x_98, 13, x_81); +lean_ctor_set(x_98, 14, x_94); +lean_ctor_set(x_98, 15, x_83); +lean_ctor_set(x_98, 16, x_84); +lean_ctor_set(x_98, 17, x_85); +lean_ctor_set(x_98, 18, x_86); +lean_ctor_set(x_98, 19, x_87); +lean_ctor_set_uint8(x_98, sizeof(void*)*20, x_73); +x_99 = l_Lean_Meta_Grind_assertAt(x_95, x_96, x_97, x_98, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_99) == 0) +{ +lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; +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_87); - x_95 = lean_box(0); + lean_dec_ref(x_99); + x_102 = lean_box(0); } -if (lean_is_scalar(x_95)) { - x_96 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_92)) { + x_103 = lean_alloc_ctor(1, 1, 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; -} -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_iterate_go(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -if (lean_obj_tag(x_2) == 0) -{ -lean_object* x_12; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_1); -x_12 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_12, 0, x_3); -lean_ctor_set(x_12, 1, x_11); -return x_12; -} -else -{ -uint8_t x_13; -x_13 = !lean_is_exclusive(x_2); -if (x_13 == 0) -{ -lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_14 = lean_ctor_get(x_2, 0); -x_15 = lean_ctor_get(x_2, 1); -lean_inc(x_1); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_14); -x_16 = lean_apply_9(x_1, x_14, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_16) == 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_18; -x_18 = lean_ctor_get(x_16, 1); -lean_inc(x_18); -lean_dec(x_16); -lean_ctor_set(x_2, 1, x_3); -{ -lean_object* _tmp_1 = x_15; -lean_object* _tmp_2 = x_2; -lean_object* _tmp_10 = x_18; -x_2 = _tmp_1; -x_3 = _tmp_2; -x_11 = _tmp_10; -} -goto _start; -} -else -{ -lean_object* x_20; lean_object* x_21; lean_object* x_22; -lean_free_object(x_2); -lean_dec(x_14); -x_20 = lean_ctor_get(x_16, 1); -lean_inc(x_20); -lean_dec(x_16); -x_21 = lean_ctor_get(x_17, 0); -lean_inc(x_21); -lean_dec(x_17); -x_22 = l_List_appendTR___rarg(x_21, x_15); -x_2 = x_22; -x_11 = x_20; -goto _start; -} -} -else -{ -uint8_t x_24; -lean_free_object(x_2); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_1); -x_24 = !lean_is_exclusive(x_16); -if (x_24 == 0) -{ -return x_16; -} -else -{ -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_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; -} -} -} -else -{ -lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_28 = lean_ctor_get(x_2, 0); -x_29 = lean_ctor_get(x_2, 1); -lean_inc(x_29); -lean_inc(x_28); -lean_dec(x_2); -lean_inc(x_1); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_28); -x_30 = lean_apply_9(x_1, x_28, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_30) == 0) -{ -lean_object* x_31; -x_31 = lean_ctor_get(x_30, 0); -lean_inc(x_31); -if (lean_obj_tag(x_31) == 0) -{ -lean_object* x_32; lean_object* x_33; -x_32 = lean_ctor_get(x_30, 1); -lean_inc(x_32); -lean_dec(x_30); -x_33 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_33, 0, x_28); -lean_ctor_set(x_33, 1, x_3); -x_2 = x_29; -x_3 = x_33; -x_11 = x_32; -goto _start; + x_103 = x_92; } -else -{ -lean_object* x_35; lean_object* x_36; lean_object* x_37; -lean_dec(x_28); -x_35 = lean_ctor_get(x_30, 1); -lean_inc(x_35); -lean_dec(x_30); -x_36 = lean_ctor_get(x_31, 0); -lean_inc(x_36); -lean_dec(x_31); -x_37 = l_List_appendTR___rarg(x_36, x_29); -x_2 = x_37; -x_11 = x_35; -goto _start; +lean_ctor_set(x_103, 0, x_100); +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 { -lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; -lean_dec(x_29); -lean_dec(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_3); -lean_dec(x_1); -x_39 = lean_ctor_get(x_30, 0); -lean_inc(x_39); -x_40 = lean_ctor_get(x_30, 1); -lean_inc(x_40); -if (lean_is_exclusive(x_30)) { - lean_ctor_release(x_30, 0); - lean_ctor_release(x_30, 1); - x_41 = x_30; +lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; +lean_dec(x_92); +x_105 = lean_ctor_get(x_99, 0); +lean_inc(x_105); +x_106 = lean_ctor_get(x_99, 1); +lean_inc(x_106); +if (lean_is_exclusive(x_99)) { + lean_ctor_release(x_99, 0); + lean_ctor_release(x_99, 1); + x_107 = x_99; } else { - lean_dec_ref(x_30); - x_41 = lean_box(0); + lean_dec_ref(x_99); + x_107 = lean_box(0); } -if (lean_is_scalar(x_41)) { - x_42 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_107)) { + x_108 = lean_alloc_ctor(1, 2, 0); } else { - x_42 = x_41; -} -lean_ctor_set(x_42, 0, x_39); -lean_ctor_set(x_42, 1, x_40); -return x_42; -} + x_108 = x_107; } +lean_ctor_set(x_108, 0, x_105); +lean_ctor_set(x_108, 1, x_106); +return x_108; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_iterate(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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_box(0); -x_12 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_12, 0, x_1); -lean_ctor_set(x_12, 1, x_11); -x_13 = l_Lean_Meta_Grind_iterate_go(x_2, x_12, x_11, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_13; } } static lean_object* _init_l_Lean_Meta_Grind_assertAll___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_assertNext_x3f), 9, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_assertNext), 9, 0); return x_1; } } @@ -4796,7 +6253,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Grind_assertAll(lean_object* x_1, lean_obje { lean_object* x_10; lean_object* x_11; x_10 = l_Lean_Meta_Grind_assertAll___closed__1; -x_11 = l_Lean_Meta_Grind_iterate(x_1, x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_11 = l_Lean_Meta_Grind_GrindTactic_iterate(x_10, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); return x_11; } } @@ -4807,6 +6264,7 @@ lean_object* initialize_Lean_Meta_Tactic_Grind_Types(uint8_t builtin, lean_objec lean_object* initialize_Lean_Meta_Tactic_Grind_Cases(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_Tactic_Grind_Injection(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_Tactic_Grind_Core(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Meta_Tactic_Grind_Combinators(uint8_t builtin, lean_object*); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Lean_Meta_Tactic_Grind_Intro(uint8_t builtin, lean_object* w) { lean_object * res; @@ -4833,18 +6291,21 @@ lean_dec_ref(res); res = initialize_Lean_Meta_Tactic_Grind_Core(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +res = initialize_Lean_Meta_Tactic_Grind_Combinators(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); l_Lean_Meta_Grind_instInhabitedIntroResult = _init_l_Lean_Meta_Grind_instInhabitedIntroResult(); lean_mark_persistent(l_Lean_Meta_Grind_instInhabitedIntroResult); l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__1(); lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__1); -l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__2(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__2); -l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__3(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__3); -l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__4(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__4); -l_Lean_Meta_Grind_intros_go___lambda__4___closed__1 = _init_l_Lean_Meta_Grind_intros_go___lambda__4___closed__1(); -lean_mark_persistent(l_Lean_Meta_Grind_intros_go___lambda__4___closed__1); +l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__8___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__8___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__8___closed__1); +l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__8___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__8___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__8___closed__2); +l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__8___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__8___closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__8___closed__3); +l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__8___closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__8___closed__4(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__8___closed__4); l_Lean_Meta_Grind_intros___closed__1 = _init_l_Lean_Meta_Grind_intros___closed__1(); lean_mark_persistent(l_Lean_Meta_Grind_intros___closed__1); l_Lean_Meta_Grind_assertAt___closed__1 = _init_l_Lean_Meta_Grind_assertAt___closed__1(); diff --git a/stage0/stdlib/Lean/Meta/Tactic/Grind/Inv.c b/stage0/stdlib/Lean/Meta/Tactic/Grind/Inv.c index dc2a5be0cb31..7b27280575ef 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Grind/Inv.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Grind/Inv.c @@ -16,29 +16,30 @@ extern "C" { LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___closed__1; +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___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_panic___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__1___closed__6; +static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Goal_checkInvariants___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_PersistentHashMap_foldlM___at_Lean_Meta_Grind_checkInvariants___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___closed__2; LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___boxed(lean_object**); +static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__4; size_t lean_usize_shift_right(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_checkInvariants___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(lean_object*, lean_object*); lean_object* l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___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_panic___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__2___closed__1; -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___closed__3; static lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__1___closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Goal_checkInvariants___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_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___closed__3; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux___at_Lean_Meta_Grind_checkInvariants___spec__3(lean_object*, lean_object*, lean_object*); size_t lean_uint64_to_usize(uint64_t); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Goal_checkInvariants___lambda__2(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_Grind_checkInvariants___spec__4___rarg(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isApp(lean_object*); +static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__3; uint8_t l_Lean_Meta_Grind_isCongruent(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_sort___override(lean_object*); -lean_object* l_Lean_PersistentArray_push___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findEntryAux___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__5(lean_object*, lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_Meta_Grind_checkInvariants___spec__5___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_mk_array(lean_object*, lean_object*); @@ -50,27 +51,23 @@ lean_object* lean_array_fget(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Goal_checkInvariants___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Meta_Grind_grind_debug_proofs; -static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__1; -static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__4; -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___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_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___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*); +static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__6; lean_object* l_Lean_Meta_Grind_getEqcs(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_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__5; static lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___lambda__1___closed__4; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Goal_checkInvariants___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_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_stringToMessageData(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__2; -static lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___closed__2; -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___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*); -static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__4; +lean_object* l_Lean_Meta_Grind_isCongrRoot(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_Grind_grind_debug; static lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___closed__8; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents(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_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__2; extern lean_object* l_instInhabitedNat; extern lean_object* l_instInhabitedPUnit; +static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__3; lean_object* l_Lean_Meta_Grind_getRoot(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_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__1; static lean_object* l___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___closed__1; lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_checkInvariants___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -81,41 +78,43 @@ size_t lean_usize_of_nat(lean_object*); LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___closed__3; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forIn___at_Lean_Meta_Grind_checkInvariants___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__3; static lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Goal_checkInvariants(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* lean_st_ref_take(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_checkInvariants___closed__2; static lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___lambda__2___closed__2; +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___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 lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___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_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__2; static lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___lambda__1___closed__2; static lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___closed__1; +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___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* l_Lean_Meta_Grind_getCongrRoot(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_Grind_mkEqHEqProof(lean_object*, 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_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__2; static lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___closed__5; LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___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_object*, lean_object*); -static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__4; static lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__1___closed__7; +static lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__3; lean_object* l_Lean_Meta_check(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__5; lean_object* lean_st_ref_get(lean_object*, lean_object*); lean_object* lean_st_mk_ref(lean_object*, lean_object*); static lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___closed__4; +static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__1; static lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___closed__9; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_Grind_checkInvariants___spec__4___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findEntryAux___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___closed__5; extern lean_object* l_Lean_levelZero; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_checkInvariants(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_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___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*); extern lean_object* l_Lean_instInhabitedExpr; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_Meta_Grind_checkInvariants___spec__5___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_checkInvariants___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_Grind_checkInvariants___closed__1; -LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static size_t l_Lean_PersistentHashMap_findEntryAux___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__5___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static double l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__1; LEAN_EXPORT lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___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_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___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_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__7; +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__3___boxed(lean_object**); uint8_t l_Lean_Meta_Grind_isSameExpr_unsafe__1(lean_object*, lean_object*); uint8_t lean_expr_equal(lean_object*, lean_object*); uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); @@ -123,32 +122,33 @@ static lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Me LEAN_EXPORT lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_getParents(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_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__3; +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___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_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___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* l_Lean_Meta_Grind_updateLastTag(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___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*); extern lean_object* l_Lean_Meta_instMonadMetaM; +LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___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*); static lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___lambda__1___closed__6; lean_object* l_Lean_Meta_Grind_hasSameType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_usize_to_nat(size_t); lean_object* l_Lean_Meta_Grind_getENodes(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*); LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__2___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*); +static lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__1; +lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___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_Meta_Grind_getNext(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___boxed(lean_object**); +static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__5; static lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___closed__7; +static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__6; lean_object* l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___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_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__2; -static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__6; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); +static lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__4; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); -static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__7; -LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___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_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__5; +lean_object* l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___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_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___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_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__1; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_Meta_Grind_checkInvariants___spec__5(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_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___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_object* lean_panic_fn(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_checkInvariants___lambda__1___closed__1; static lean_object* l_Lean_Meta_Grind_Goal_checkInvariants___closed__1; @@ -157,21 +157,20 @@ LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findEntry_x3f___at___private_L lean_object* l_Lean_Expr_getAppFn(lean_object*); LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static size_t l_Lean_PersistentHashMap_findEntryAux___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__5___closed__1; -lean_object* l_Lean_Meta_Grind_getENode(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_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__1; static lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__1___closed__4; static lean_object* l___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___closed__3; +static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__4; static lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___closed__2; static lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__2___closed__2; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findEntryAtAux___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___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*); +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___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*); size_t lean_usize_sub(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* lean_array_mk(lean_object*); +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___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* l_Lean_Meta_Grind_isRoot(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_Grind_Inv_0__Lean_Meta_Grind_checkEqc___closed__1; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forIn___at_Lean_Meta_Grind_checkInvariants___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*); -static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__6; -lean_object* l_Lean_isTracingEnabledForCore(lean_object*, lean_object*, lean_object*); size_t lean_usize_add(size_t, size_t); static lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___closed__6; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_checkInvariants___lambda__1(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -181,39 +180,34 @@ size_t lean_array_size(lean_object*); LEAN_EXPORT lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___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_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_instInhabitedOfMonad___rarg(lean_object*, lean_object*); -lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Goal_checkInvariants___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_shift_left(size_t, size_t); +static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__4; lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(lean_object*, lean_object*, lean_object*); -static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__3; -static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__1; +static lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__5; lean_object* lean_string_append(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc(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_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__8___closed__1; lean_object* lean_array_get_size(lean_object*); -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_infer_type(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___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_PersistentHashMap_findEntryAtAux___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___lambda__2___closed__3; uint8_t lean_usize_dec_lt(size_t, size_t); static lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___lambda__1___closed__5; -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___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_nat_add(lean_object*, lean_object*); -static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__3; -LEAN_EXPORT lean_object* l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___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_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__2; +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___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_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Goal_checkInvariants___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_foldlMUnsafe_fold___at_Lean_Meta_Grind_checkInvariants___spec__4(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_getRoot_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_get(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___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*); -static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__2; lean_object* l_ReaderT_instMonad___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux___at_Lean_Meta_Grind_checkInvariants___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*, lean_object*); size_t lean_usize_land(size_t, size_t); static lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__1___closed__2; -static lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___closed__4; static lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___lambda__1___closed__1; static lean_object* _init_l_panic___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__1___closed__1() { _start: @@ -636,6 +630,7 @@ LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findEntry_x3f___at___private_L lean_object* x_4; uint64_t x_5; size_t x_6; lean_object* x_7; x_4 = lean_ctor_get(x_1, 1); lean_inc(x_4); +lean_inc(x_3); x_5 = l_Lean_Meta_Grind_congrHash(x_4, x_3); x_6 = lean_uint64_to_usize(x_5); x_7 = l_Lean_PersistentHashMap_findEntryAux___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__5(x_1, x_2, x_6, x_3); @@ -1092,7 +1087,7 @@ static lean_object* _init_l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("isSameExpr curr ( __do_lift._@.Lean.Meta.Tactic.Grind.Inv._hyg.34.0 ).cgRoot\n -- If the equivalence class does not have HEq proofs, then the types must be definitionally equal.\n ", 184, 184); +x_1 = lean_mk_string_unchecked("( __do_lift._@.Lean.Meta.Tactic.Grind.Inv._hyg.34.0 )\n -- If the equivalence class does not have HEq proofs, then the types must be definitionally equal.\n ", 161, 161); return x_1; } } @@ -1123,7 +1118,7 @@ static lean_object* _init_l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("isSameExpr e ( __do_lift._@.Lean.Meta.Tactic.Grind.Inv._hyg.31.0 ).cgRoot\n ", 80, 80); +x_1 = lean_mk_string_unchecked("isSameExpr e ( __do_lift._@.Lean.Meta.Tactic.Grind.Inv._hyg.31.0 )\n ", 73, 73); return x_1; } } @@ -1230,26 +1225,23 @@ if (lean_obj_tag(x_46) == 0) { lean_object* x_47; lean_inc(x_29); -x_47 = l_Lean_Meta_Grind_getENode(x_29, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_44); +x_47 = l_Lean_Meta_Grind_isCongrRoot(x_29, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_44); if (lean_obj_tag(x_47) == 0) { -lean_object* x_48; lean_object* x_49; lean_object* x_50; uint8_t x_51; +lean_object* x_48; uint8_t x_49; x_48 = lean_ctor_get(x_47, 0); lean_inc(x_48); -x_49 = lean_ctor_get(x_47, 1); -lean_inc(x_49); -lean_dec(x_47); -x_50 = lean_ctor_get(x_48, 3); -lean_inc(x_50); +x_49 = lean_unbox(x_48); lean_dec(x_48); -x_51 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_29, x_50); -lean_dec(x_50); -if (x_51 == 0) +if (x_49 == 0) { -lean_object* x_52; lean_object* x_53; +lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_dec(x_32); lean_dec(x_29); -x_52 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___closed__6; +x_50 = lean_ctor_get(x_47, 1); +lean_inc(x_50); +lean_dec(x_47); +x_51 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___closed__6; lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); @@ -1258,13 +1250,16 @@ lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -x_53 = l_panic___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__1(x_52, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_49); -x_13 = x_53; +x_52 = l_panic___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__1(x_51, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_50); +x_13 = x_52; goto block_28; } else { -lean_object* x_54; lean_object* x_55; +lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_53 = lean_ctor_get(x_47, 1); +lean_inc(x_53); +lean_dec(x_47); x_54 = lean_box(0); lean_inc(x_11); lean_inc(x_10); @@ -1275,7 +1270,7 @@ lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_2); -x_55 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___lambda__2(x_2, x_32, x_1, x_29, x_54, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_49); +x_55 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___lambda__2(x_2, x_32, x_1, x_29, x_54, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_53); x_13 = x_55; goto block_28; } @@ -1358,27 +1353,24 @@ x_70 = lean_ctor_get(x_64, 1); lean_inc(x_70); lean_dec(x_64); lean_inc(x_29); -x_71 = l_Lean_Meta_Grind_getENode(x_29, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_70); +x_71 = l_Lean_Meta_Grind_getCongrRoot(x_29, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_70); if (lean_obj_tag(x_71) == 0) { -lean_object* x_72; lean_object* x_73; lean_object* x_74; uint8_t x_75; +lean_object* x_72; lean_object* x_73; uint8_t x_74; x_72 = lean_ctor_get(x_71, 0); lean_inc(x_72); x_73 = lean_ctor_get(x_71, 1); lean_inc(x_73); lean_dec(x_71); -x_74 = lean_ctor_get(x_72, 3); -lean_inc(x_74); +x_74 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_61, x_72); lean_dec(x_72); -x_75 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_61, x_74); -lean_dec(x_74); lean_dec(x_61); -if (x_75 == 0) +if (x_74 == 0) { -lean_object* x_76; lean_object* x_77; +lean_object* x_75; lean_object* x_76; lean_dec(x_32); lean_dec(x_29); -x_76 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___closed__9; +x_75 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___closed__9; lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); @@ -1387,14 +1379,14 @@ lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -x_77 = l_panic___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__1(x_76, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_73); -x_13 = x_77; +x_76 = l_panic___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__1(x_75, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_73); +x_13 = x_76; goto block_28; } else { -lean_object* x_78; lean_object* x_79; -x_78 = lean_box(0); +lean_object* x_77; lean_object* x_78; +x_77 = lean_box(0); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); @@ -1404,35 +1396,35 @@ lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_2); -x_79 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___lambda__2(x_2, x_32, x_1, x_29, x_78, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_73); -x_13 = x_79; +x_78 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___lambda__2(x_2, x_32, x_1, x_29, x_77, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_73); +x_13 = x_78; goto block_28; } } else { -uint8_t x_80; +uint8_t x_79; lean_dec(x_61); lean_dec(x_32); lean_dec(x_29); -x_80 = !lean_is_exclusive(x_71); -if (x_80 == 0) +x_79 = !lean_is_exclusive(x_71); +if (x_79 == 0) { x_13 = x_71; goto block_28; } else { -lean_object* x_81; lean_object* x_82; lean_object* x_83; -x_81 = lean_ctor_get(x_71, 0); -x_82 = lean_ctor_get(x_71, 1); -lean_inc(x_82); +lean_object* x_80; lean_object* x_81; lean_object* x_82; +x_80 = lean_ctor_get(x_71, 0); +x_81 = lean_ctor_get(x_71, 1); lean_inc(x_81); +lean_inc(x_80); lean_dec(x_71); -x_83 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_83, 0, x_81); -lean_ctor_set(x_83, 1, x_82); -x_13 = x_83; +x_82 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_82, 0, x_80); +lean_ctor_set(x_82, 1, x_81); +x_13 = x_82; goto block_28; } } @@ -1440,28 +1432,28 @@ goto block_28; } else { -uint8_t x_84; +uint8_t x_83; lean_dec(x_61); lean_dec(x_32); lean_dec(x_29); -x_84 = !lean_is_exclusive(x_64); -if (x_84 == 0) +x_83 = !lean_is_exclusive(x_64); +if (x_83 == 0) { x_13 = x_64; goto block_28; } else { -lean_object* x_85; lean_object* x_86; lean_object* x_87; -x_85 = lean_ctor_get(x_64, 0); -x_86 = lean_ctor_get(x_64, 1); -lean_inc(x_86); +lean_object* x_84; lean_object* x_85; lean_object* x_86; +x_84 = lean_ctor_get(x_64, 0); +x_85 = lean_ctor_get(x_64, 1); lean_inc(x_85); +lean_inc(x_84); lean_dec(x_64); -x_87 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_87, 0, x_85); -lean_ctor_set(x_87, 1, x_86); -x_13 = x_87; +x_86 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_86, 0, x_84); +lean_ctor_set(x_86, 1, x_85); +x_13 = x_86; goto block_28; } } @@ -1471,27 +1463,27 @@ goto block_28; } else { -uint8_t x_88; +uint8_t x_87; lean_dec(x_32); lean_dec(x_29); -x_88 = !lean_is_exclusive(x_33); -if (x_88 == 0) +x_87 = !lean_is_exclusive(x_33); +if (x_87 == 0) { x_13 = x_33; goto block_28; } else { -lean_object* x_89; lean_object* x_90; lean_object* x_91; -x_89 = lean_ctor_get(x_33, 0); -x_90 = lean_ctor_get(x_33, 1); -lean_inc(x_90); +lean_object* x_88; lean_object* x_89; lean_object* x_90; +x_88 = lean_ctor_get(x_33, 0); +x_89 = lean_ctor_get(x_33, 1); lean_inc(x_89); +lean_inc(x_88); lean_dec(x_33); -x_91 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_91, 0, x_89); -lean_ctor_set(x_91, 1, x_90); -x_13 = x_91; +x_90 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_90, 0, x_88); +lean_ctor_set(x_90, 1, x_89); +x_13 = x_90; goto block_28; } } @@ -1973,34 +1965,25 @@ x_13 = lean_apply_9(x_12, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); return x_13; } } -static lean_object* _init_l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_levelZero; -x_2 = l_Lean_Expr_sort___override(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___closed__2() { +static lean_object* _init_l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("( __do_lift._@.Lean.Meta.Tactic.Grind.Inv._hyg.778.0 )\n ", 57, 57); +x_1 = lean_mk_string_unchecked("( __do_lift._@.Lean.Meta.Tactic.Grind.Inv._hyg.782.0 )\n ", 57, 57); return x_1; } } -static lean_object* _init_l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___closed__3() { +static lean_object* _init_l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___lambda__1___closed__1; -x_2 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___closed__2; +x_2 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__1; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___closed__4() { +static lean_object* _init_l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__3() { _start: { lean_object* x_1; @@ -2008,19 +1991,155 @@ x_1 = lean_mk_string_unchecked("_private.Lean.Meta.Tactic.Grind.Inv.0.Lean.Meta. return x_1; } } -static lean_object* _init_l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___closed__5() { +static lean_object* _init_l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___lambda__1___closed__4; -x_2 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___closed__4; -x_3 = lean_unsigned_to_nat(61u); +x_2 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__3; +x_3 = lean_unsigned_to_nat(65u); x_4 = lean_unsigned_to_nat(8u); -x_5 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___closed__3; +x_5 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__2; 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_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__5() { +_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_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +if (x_3 == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_14 = l_Lean_Expr_getAppFn(x_1); +x_15 = l_Lean_Meta_Grind_getRoot_x3f(x_14, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_14); +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_17 = lean_ctor_get(x_15, 1); +lean_inc(x_17); +lean_dec(x_15); +x_18 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__4; +x_19 = l_panic___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__2(x_18, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_17); +return x_19; +} +else +{ +uint8_t x_20; +x_20 = !lean_is_exclusive(x_15); +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_15, 1); +x_22 = lean_ctor_get(x_15, 0); +lean_dec(x_22); +x_23 = lean_ctor_get(x_16, 0); +lean_inc(x_23); +lean_dec(x_16); +x_24 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_23, x_2); +lean_dec(x_23); +if (x_24 == 0) +{ +lean_object* x_25; lean_object* x_26; +lean_free_object(x_15); +x_25 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__4; +x_26 = l_panic___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__2(x_25, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_21); +return x_26; +} +else +{ +lean_object* x_27; +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +x_27 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__5; +lean_ctor_set(x_15, 0, x_27); +return x_15; +} +} +else +{ +lean_object* x_28; lean_object* x_29; uint8_t x_30; +x_28 = lean_ctor_get(x_15, 1); +lean_inc(x_28); +lean_dec(x_15); +x_29 = lean_ctor_get(x_16, 0); +lean_inc(x_29); +lean_dec(x_16); +x_30 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_29, x_2); +lean_dec(x_29); +if (x_30 == 0) +{ +lean_object* x_31; lean_object* x_32; +x_31 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__4; +x_32 = l_panic___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__2(x_31, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_28); +return x_32; +} +else +{ +lean_object* x_33; lean_object* x_34; +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +x_33 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__5; +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; +} +} +} +} +else +{ +lean_object* x_35; lean_object* x_36; +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +x_35 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__5; +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set(x_36, 1, x_13); +return x_36; +} +} +} +static lean_object* _init_l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_levelZero; +x_2 = l_Lean_Expr_sort___override(x_1); +return x_2; +} +} LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { @@ -2101,50 +2220,80 @@ 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; size_t x_33; size_t x_34; uint8_t x_35; lean_object* x_36; lean_object* x_37; uint8_t x_38; +lean_object* x_24; lean_object* x_25; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; size_t x_47; size_t x_48; uint8_t x_49; lean_object* x_50; lean_dec(x_19); x_24 = lean_ctor_get(x_18, 1); lean_inc(x_24); lean_dec(x_18); -x_25 = lean_unsigned_to_nat(0u); -x_26 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_16, x_25); -x_27 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___closed__1; -lean_inc(x_26); -x_28 = lean_mk_array(x_26, x_27); -x_29 = lean_unsigned_to_nat(1u); -x_30 = lean_nat_sub(x_26, x_29); -lean_dec(x_26); +x_39 = lean_unsigned_to_nat(0u); +x_40 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_16, x_39); +x_41 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___closed__1; +lean_inc(x_40); +x_42 = lean_mk_array(x_40, x_41); +x_43 = lean_unsigned_to_nat(1u); +x_44 = lean_nat_sub(x_40, x_43); +lean_dec(x_40); lean_inc(x_16); -x_31 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_16, x_28, x_30); -x_32 = lean_box(0); -x_33 = lean_array_size(x_31); -x_34 = 0; -x_35 = 0; -x_36 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__1(x_1, x_31, x_32, x_31, x_33, x_34, x_35, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_24); -lean_dec(x_31); -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) +x_45 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_16, x_42, x_44); +x_46 = lean_box(0); +x_47 = lean_array_size(x_45); +x_48 = 0; +x_49 = 0; +x_50 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__1(x_1, x_45, x_46, x_45, x_47, x_48, x_49, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_24); +lean_dec(x_45); +if (lean_obj_tag(x_16) == 7) { -lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; -x_39 = lean_ctor_get(x_36, 1); -lean_inc(x_39); -lean_dec(x_36); -x_40 = l_Lean_Expr_getAppFn(x_16); +lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_51 = lean_ctor_get(x_50, 0); +lean_inc(x_51); +x_52 = lean_ctor_get(x_50, 1); +lean_inc(x_52); +lean_dec(x_50); +x_53 = lean_ctor_get(x_16, 1); +lean_inc(x_53); +x_54 = l_Lean_Meta_Grind_getRoot_x3f(x_53, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_52); +lean_dec(x_53); +x_55 = lean_ctor_get(x_54, 0); +lean_inc(x_55); +if (lean_obj_tag(x_55) == 0) +{ +lean_object* x_56; lean_object* x_57; uint8_t x_58; lean_object* x_59; +x_56 = lean_ctor_get(x_54, 1); +lean_inc(x_56); +lean_dec(x_54); +x_57 = lean_box(0); +x_58 = lean_unbox(x_51); +lean_dec(x_51); +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_59 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1(x_16, x_1, x_58, x_57, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_56); lean_dec(x_16); -x_41 = l_Lean_Meta_Grind_getRoot_x3f(x_40, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_39); -lean_dec(x_40); -x_42 = lean_ctor_get(x_41, 0); -lean_inc(x_42); -if (lean_obj_tag(x_42) == 0) +x_25 = x_59; +goto block_38; +} +else { -lean_object* x_43; lean_object* x_44; lean_object* x_45; -x_43 = lean_ctor_get(x_41, 1); -lean_inc(x_43); -lean_dec(x_41); -x_44 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___closed__5; +lean_object* x_60; lean_object* x_61; uint8_t x_62; +x_60 = lean_ctor_get(x_54, 1); +lean_inc(x_60); +lean_dec(x_54); +x_61 = lean_ctor_get(x_55, 0); +lean_inc(x_61); +lean_dec(x_55); +x_62 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_61, x_1); +lean_dec(x_61); +if (x_62 == 0) +{ +lean_object* x_63; uint8_t x_64; lean_object* x_65; +x_63 = lean_box(0); +x_64 = lean_unbox(x_51); +lean_dec(x_51); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); @@ -2153,15 +2302,66 @@ lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -x_45 = l_panic___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__2(x_44, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_43); -if (lean_obj_tag(x_45) == 0) +x_65 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1(x_16, x_1, x_64, x_63, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_60); +lean_dec(x_16); +x_25 = x_65; +goto block_38; +} +else { -lean_object* x_46; -x_46 = lean_ctor_get(x_45, 0); -lean_inc(x_46); -if (lean_obj_tag(x_46) == 0) +uint8_t x_66; lean_object* x_67; lean_object* x_68; +lean_dec(x_51); +x_66 = 1; +x_67 = lean_box(0); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +x_68 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1(x_16, x_1, x_66, x_67, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_60); +lean_dec(x_16); +x_25 = x_68; +goto block_38; +} +} +} +else +{ +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_50, 0); +lean_inc(x_69); +x_70 = lean_ctor_get(x_50, 1); +lean_inc(x_70); +lean_dec(x_50); +x_71 = lean_box(0); +x_72 = lean_unbox(x_69); +lean_dec(x_69); +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_73 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1(x_16, x_1, x_72, x_71, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_70); +lean_dec(x_16); +x_25 = x_73; +goto block_38; +} +block_38: +{ +if (lean_obj_tag(x_25) == 0) +{ +lean_object* x_26; +x_26 = lean_ctor_get(x_25, 0); +lean_inc(x_26); +if (lean_obj_tag(x_26) == 0) { -uint8_t x_47; +uint8_t x_27; lean_dec(x_17); lean_dec(x_11); lean_dec(x_10); @@ -2171,44 +2371,44 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_47 = !lean_is_exclusive(x_45); -if (x_47 == 0) +x_27 = !lean_is_exclusive(x_25); +if (x_27 == 0) { -lean_object* x_48; -x_48 = lean_ctor_get(x_45, 0); -lean_dec(x_48); -return x_45; +lean_object* x_28; +x_28 = lean_ctor_get(x_25, 0); +lean_dec(x_28); +return x_25; } else { -lean_object* x_49; lean_object* x_50; -x_49 = lean_ctor_get(x_45, 1); -lean_inc(x_49); -lean_dec(x_45); -x_50 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_50, 0, x_46); -lean_ctor_set(x_50, 1, x_49); -return x_50; +lean_object* x_29; lean_object* x_30; +x_29 = lean_ctor_get(x_25, 1); +lean_inc(x_29); +lean_dec(x_25); +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_26); +lean_ctor_set(x_30, 1, x_29); +return x_30; } } else { -lean_object* x_51; lean_object* x_52; -x_51 = lean_ctor_get(x_45, 1); -lean_inc(x_51); -lean_dec(x_45); -x_52 = lean_ctor_get(x_46, 0); -lean_inc(x_52); -lean_dec(x_46); +lean_object* x_31; lean_object* x_32; +x_31 = lean_ctor_get(x_25, 1); +lean_inc(x_31); +lean_dec(x_25); +x_32 = lean_ctor_get(x_26, 0); +lean_inc(x_32); +lean_dec(x_26); x_2 = x_17; -x_3 = x_52; -x_12 = x_51; +x_3 = x_32; +x_12 = x_31; goto _start; } } else { -uint8_t x_54; +uint8_t x_34; lean_dec(x_17); lean_dec(x_11); lean_dec(x_10); @@ -2218,163 +2418,31 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_54 = !lean_is_exclusive(x_45); -if (x_54 == 0) +x_34 = !lean_is_exclusive(x_25); +if (x_34 == 0) { -return x_45; +return x_25; } else { -lean_object* x_55; lean_object* x_56; lean_object* x_57; -x_55 = lean_ctor_get(x_45, 0); -x_56 = lean_ctor_get(x_45, 1); -lean_inc(x_56); -lean_inc(x_55); -lean_dec(x_45); -x_57 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_57, 0, x_55); -lean_ctor_set(x_57, 1, x_56); -return x_57; -} -} -} -else -{ -lean_object* x_58; lean_object* x_59; uint8_t x_60; -x_58 = lean_ctor_get(x_41, 1); -lean_inc(x_58); -lean_dec(x_41); -x_59 = lean_ctor_get(x_42, 0); -lean_inc(x_59); -lean_dec(x_42); -x_60 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_59, x_1); -lean_dec(x_59); -if (x_60 == 0) -{ -lean_object* x_61; lean_object* x_62; -x_61 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___closed__5; -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_62 = l_panic___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__2(x_61, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_58); -if (lean_obj_tag(x_62) == 0) -{ -lean_object* x_63; -x_63 = lean_ctor_get(x_62, 0); -lean_inc(x_63); -if (lean_obj_tag(x_63) == 0) -{ -uint8_t x_64; -lean_dec(x_17); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -x_64 = !lean_is_exclusive(x_62); -if (x_64 == 0) -{ -lean_object* x_65; -x_65 = lean_ctor_get(x_62, 0); -lean_dec(x_65); -return x_62; -} -else -{ -lean_object* x_66; lean_object* x_67; -x_66 = lean_ctor_get(x_62, 1); -lean_inc(x_66); -lean_dec(x_62); -x_67 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_67, 0, x_63); -lean_ctor_set(x_67, 1, x_66); -return x_67; -} -} -else -{ -lean_object* x_68; lean_object* x_69; -x_68 = lean_ctor_get(x_62, 1); -lean_inc(x_68); -lean_dec(x_62); -x_69 = lean_ctor_get(x_63, 0); -lean_inc(x_69); -lean_dec(x_63); -x_2 = x_17; -x_3 = x_69; -x_12 = x_68; -goto _start; -} -} -else -{ -uint8_t x_71; -lean_dec(x_17); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -x_71 = !lean_is_exclusive(x_62); -if (x_71 == 0) -{ -return x_62; -} -else -{ -lean_object* x_72; lean_object* x_73; lean_object* x_74; -x_72 = lean_ctor_get(x_62, 0); -x_73 = lean_ctor_get(x_62, 1); -lean_inc(x_73); -lean_inc(x_72); -lean_dec(x_62); -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 -{ -lean_object* x_75; -x_75 = lean_box(0); -x_2 = x_17; -x_3 = x_75; -x_12 = x_58; -goto _start; -} +lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_35 = lean_ctor_get(x_25, 0); +x_36 = lean_ctor_get(x_25, 1); +lean_inc(x_36); +lean_inc(x_35); +lean_dec(x_25); +x_37 = lean_alloc_ctor(1, 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_77; lean_object* x_78; -lean_dec(x_16); -x_77 = lean_ctor_get(x_36, 1); -lean_inc(x_77); -lean_dec(x_36); -x_78 = lean_box(0); -x_2 = x_17; -x_3 = x_78; -x_12 = x_77; -goto _start; } } } else { -uint8_t x_80; +uint8_t x_74; lean_dec(x_17); lean_dec(x_16); lean_dec(x_11); @@ -2385,23 +2453,23 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_80 = !lean_is_exclusive(x_18); -if (x_80 == 0) +x_74 = !lean_is_exclusive(x_18); +if (x_74 == 0) { return x_18; } else { -lean_object* x_81; lean_object* x_82; lean_object* x_83; -x_81 = lean_ctor_get(x_18, 0); -x_82 = lean_ctor_get(x_18, 1); -lean_inc(x_82); -lean_inc(x_81); +lean_object* x_75; lean_object* x_76; lean_object* x_77; +x_75 = lean_ctor_get(x_18, 0); +x_76 = lean_ctor_get(x_18, 1); +lean_inc(x_76); +lean_inc(x_75); lean_dec(x_18); -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; +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; } } } @@ -2411,7 +2479,7 @@ static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Gr _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("( __do_lift._@.Lean.Meta.Tactic.Grind.Inv._hyg.791.0 ).isEmpty\n\n", 64, 64); +x_1 = lean_mk_string_unchecked("( __do_lift._@.Lean.Meta.Tactic.Grind.Inv._hyg.805.0 ).isEmpty\n\n", 64, 64); return x_1; } } @@ -2430,8 +2498,8 @@ static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Gr { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___lambda__1___closed__4; -x_2 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___closed__4; -x_3 = lean_unsigned_to_nat(64u); +x_2 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__3; +x_3 = lean_unsigned_to_nat(68u); x_4 = lean_unsigned_to_nat(4u); x_5 = l___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___closed__2; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -2591,6 +2659,19 @@ lean_dec(x_1); return x_20; } } +LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +uint8_t x_14; lean_object* x_15; +x_14 = lean_unbox(x_3); +lean_dec(x_3); +x_15 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1(x_1, x_2, x_14, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +return x_15; +} +} LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { @@ -2609,7 +2690,17 @@ lean_dec(x_1); return x_11; } } -static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__1() { +LEAN_EXPORT lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_11 = l_panic___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__2___closed__2; +x_12 = lean_panic_fn(x_11, x_1); +x_13 = lean_apply_9(x_12, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_13; +} +} +static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__1() { _start: { lean_object* x_1; @@ -2617,17 +2708,17 @@ x_1 = lean_mk_string_unchecked("!Expr.equal n₁.self n₂.self\n\n", 33, 29); return x_1; } } -static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__2() { +static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___lambda__1___closed__1; -x_2 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__1; +x_2 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__1; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__3() { +static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__3() { _start: { lean_object* x_1; @@ -2635,20 +2726,20 @@ x_1 = lean_mk_string_unchecked("_private.Lean.Meta.Tactic.Grind.Inv.0.Lean.Meta. return x_1; } } -static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__4() { +static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___lambda__1___closed__4; -x_2 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__3; -x_3 = lean_unsigned_to_nat(75u); +x_2 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__3; +x_3 = lean_unsigned_to_nat(79u); x_4 = lean_unsigned_to_nat(6u); -x_5 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__2; +x_5 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__2; 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_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__5() { +static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__5() { _start: { lean_object* x_1; @@ -2656,30 +2747,30 @@ x_1 = lean_mk_string_unchecked("!isSameExpr n₁.self n₂.self\n -- and th return x_1; } } -static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__6() { +static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___lambda__1___closed__1; -x_2 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__5; +x_2 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__5; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__7() { +static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___lambda__1___closed__4; -x_2 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__3; -x_3 = lean_unsigned_to_nat(73u); +x_2 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__3; +x_3 = lean_unsigned_to_nat(77u); x_4 = lean_unsigned_to_nat(6u); -x_5 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__6; +x_5 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__6; 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_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17) { +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___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) { _start: { lean_object* x_18; uint8_t x_19; @@ -2733,7 +2824,7 @@ goto _start; else { lean_object* x_30; lean_object* x_31; -x_30 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__4; +x_30 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__4; lean_inc(x_16); lean_inc(x_15); lean_inc(x_14); @@ -2742,7 +2833,7 @@ lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); -x_31 = l_panic___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__2(x_30, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +x_31 = l_panic___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1(x_30, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); if (lean_obj_tag(x_31) == 0) { lean_object* x_32; @@ -2844,7 +2935,7 @@ else { lean_object* x_48; lean_object* x_49; lean_dec(x_23); -x_48 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__7; +x_48 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__7; lean_inc(x_16); lean_inc(x_15); lean_inc(x_14); @@ -2853,7 +2944,7 @@ lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); -x_49 = l_panic___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__2(x_48, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +x_49 = l_panic___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1(x_48, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); if (lean_obj_tag(x_49) == 0) { lean_object* x_50; @@ -2953,7 +3044,7 @@ return x_65; } } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17) { +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17) { _start: { lean_object* x_18; uint8_t x_19; @@ -2999,7 +3090,7 @@ lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); -x_26 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1(x_1, x_21, x_24, x_24, x_25, x_23, lean_box(0), lean_box(0), x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +x_26 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2(x_1, x_21, x_24, x_24, x_25, x_23, lean_box(0), lean_box(0), x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); lean_dec(x_24); lean_dec(x_21); if (lean_obj_tag(x_26) == 0) @@ -3072,7 +3163,7 @@ lean_ctor_set(x_16, 0, x_14); lean_ctor_set(x_16, 1, x_13); lean_ctor_set(x_16, 2, x_15); x_17 = lean_box(0); -x_18 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2(x_11, x_13, x_16, x_16, x_17, x_14, lean_box(0), lean_box(0), x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_12); +x_18 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__3(x_11, x_13, x_16, x_16, x_17, x_14, lean_box(0), lean_box(0), x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_12); lean_dec(x_16); lean_dec(x_11); if (lean_obj_tag(x_18) == 0) @@ -3123,7 +3214,7 @@ return x_26; } } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___boxed(lean_object** _args) { +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___boxed(lean_object** _args) { lean_object* x_1 = _args[0]; lean_object* x_2 = _args[1]; lean_object* x_3 = _args[2]; @@ -3144,7 +3235,7 @@ lean_object* x_17 = _args[16]; _start: { lean_object* x_18; -x_18 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +x_18 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); @@ -3152,7 +3243,7 @@ lean_dec(x_1); return x_18; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___boxed(lean_object** _args) { +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__3___boxed(lean_object** _args) { lean_object* x_1 = _args[0]; lean_object* x_2 = _args[1]; lean_object* x_3 = _args[2]; @@ -3173,50 +3264,31 @@ lean_object* x_17 = _args[16]; _start: { lean_object* x_18; -x_18 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___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_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); return x_18; } } -LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___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) { +static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__1() { _start: { -lean_object* x_11; lean_object* x_12; uint8_t x_13; -x_11 = lean_ctor_get(x_8, 2); -x_12 = l_Lean_isTracingEnabledForCore(x_1, x_11, x_10); -x_13 = !lean_is_exclusive(x_12); -if (x_13 == 0) -{ -return x_12; -} -else -{ -lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_14 = lean_ctor_get(x_12, 0); -x_15 = lean_ctor_get(x_12, 1); -lean_inc(x_15); -lean_inc(x_14); -lean_dec(x_12); -x_16 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_16, 0, x_14); -lean_ctor_set(x_16, 1, x_15); -return x_16; -} +lean_object* x_1; +x_1 = lean_mk_string_unchecked("checked: ", 9, 9); +return x_1; } } -static double _init_l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__1() { +static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__2() { _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_object* x_1; lean_object* x_2; +x_1 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } } -static lean_object* _init_l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__2() { +static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__3() { _start: { lean_object* x_1; @@ -3224,539 +3296,221 @@ x_1 = lean_mk_string_unchecked("", 0, 0); return x_1; } } -static lean_object* _init_l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__3() { +static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__4() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = lean_box(0); -x_2 = lean_array_mk(x_1); +x_1 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__3; +x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___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_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; -x_12 = lean_ctor_get(x_9, 5); -x_13 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_2, x_7, x_8, x_9, x_10, x_11); -x_14 = lean_ctor_get(x_13, 0); +lean_object* x_13; +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_1); +x_13 = l_Lean_Meta_check(x_1, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_13) == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_14 = lean_ctor_get(x_13, 1); lean_inc(x_14); -x_15 = lean_ctor_get(x_13, 1); -lean_inc(x_15); lean_dec(x_13); -x_16 = lean_st_ref_take(x_10, x_15); -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -x_18 = lean_ctor_get(x_17, 3); -lean_inc(x_18); -x_19 = !lean_is_exclusive(x_16); -if (x_19 == 0) -{ -lean_object* x_20; lean_object* x_21; uint8_t x_22; -x_20 = lean_ctor_get(x_16, 1); -x_21 = lean_ctor_get(x_16, 0); -lean_dec(x_21); -x_22 = !lean_is_exclusive(x_17); -if (x_22 == 0) -{ -lean_object* x_23; uint8_t x_24; -x_23 = lean_ctor_get(x_17, 3); -lean_dec(x_23); -x_24 = !lean_is_exclusive(x_18); -if (x_24 == 0) +lean_inc(x_2); +x_15 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_14); +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_25; double 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; uint8_t x_34; -x_25 = lean_ctor_get(x_18, 0); -x_26 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__1; -x_27 = 0; -x_28 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__2; -x_29 = lean_alloc_ctor(0, 2, 17); -lean_ctor_set(x_29, 0, x_1); -lean_ctor_set(x_29, 1, x_28); -lean_ctor_set_float(x_29, sizeof(void*)*2, x_26); -lean_ctor_set_float(x_29, sizeof(void*)*2 + 8, x_26); -lean_ctor_set_uint8(x_29, sizeof(void*)*2 + 16, x_27); -x_30 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__3; -x_31 = lean_alloc_ctor(9, 3, 0); -lean_ctor_set(x_31, 0, x_29); -lean_ctor_set(x_31, 1, x_14); -lean_ctor_set(x_31, 2, x_30); -lean_inc(x_12); -lean_ctor_set(x_16, 1, x_31); -lean_ctor_set(x_16, 0, x_12); -x_32 = l_Lean_PersistentArray_push___rarg(x_25, x_16); -lean_ctor_set(x_18, 0, x_32); -x_33 = lean_st_ref_set(x_10, x_17, x_20); -x_34 = !lean_is_exclusive(x_33); -if (x_34 == 0) +uint8_t x_18; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_2); +lean_dec(x_1); +x_18 = !lean_is_exclusive(x_15); +if (x_18 == 0) { -lean_object* x_35; lean_object* x_36; -x_35 = lean_ctor_get(x_33, 0); -lean_dec(x_35); -x_36 = lean_box(0); -lean_ctor_set(x_33, 0, x_36); -return x_33; +lean_object* x_19; lean_object* x_20; +x_19 = lean_ctor_get(x_15, 0); +lean_dec(x_19); +x_20 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__5; +lean_ctor_set(x_15, 0, x_20); +return x_15; } else { -lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_37 = lean_ctor_get(x_33, 1); -lean_inc(x_37); -lean_dec(x_33); -x_38 = lean_box(0); -x_39 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_39, 0, x_38); -lean_ctor_set(x_39, 1, x_37); -return x_39; +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_15, 1); +lean_inc(x_21); +lean_dec(x_15); +x_22 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__5; +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 { -uint64_t x_40; lean_object* x_41; double 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; -x_40 = lean_ctor_get_uint64(x_18, sizeof(void*)*1); -x_41 = lean_ctor_get(x_18, 0); -lean_inc(x_41); -lean_dec(x_18); -x_42 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__1; -x_43 = 0; -x_44 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__2; -x_45 = lean_alloc_ctor(0, 2, 17); -lean_ctor_set(x_45, 0, x_1); -lean_ctor_set(x_45, 1, x_44); -lean_ctor_set_float(x_45, sizeof(void*)*2, x_42); -lean_ctor_set_float(x_45, sizeof(void*)*2 + 8, x_42); -lean_ctor_set_uint8(x_45, sizeof(void*)*2 + 16, x_43); -x_46 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__3; -x_47 = lean_alloc_ctor(9, 3, 0); -lean_ctor_set(x_47, 0, x_45); -lean_ctor_set(x_47, 1, x_14); -lean_ctor_set(x_47, 2, x_46); -lean_inc(x_12); -lean_ctor_set(x_16, 1, x_47); -lean_ctor_set(x_16, 0, x_12); -x_48 = l_Lean_PersistentArray_push___rarg(x_41, x_16); -x_49 = lean_alloc_ctor(0, 1, 8); -lean_ctor_set(x_49, 0, x_48); -lean_ctor_set_uint64(x_49, sizeof(void*)*1, x_40); -lean_ctor_set(x_17, 3, x_49); -x_50 = lean_st_ref_set(x_10, x_17, x_20); -x_51 = lean_ctor_get(x_50, 1); -lean_inc(x_51); -if (lean_is_exclusive(x_50)) { - lean_ctor_release(x_50, 0); - lean_ctor_release(x_50, 1); - x_52 = x_50; -} else { - lean_dec_ref(x_50); - x_52 = lean_box(0); -} -x_53 = lean_box(0); -if (lean_is_scalar(x_52)) { - x_54 = lean_alloc_ctor(0, 2, 0); -} else { - x_54 = x_52; -} -lean_ctor_set(x_54, 0, x_53); -lean_ctor_set(x_54, 1, x_51); -return x_54; -} +lean_object* x_24; lean_object* x_25; +x_24 = lean_ctor_get(x_15, 1); +lean_inc(x_24); +lean_dec(x_15); +x_25 = l_Lean_Meta_Grind_updateLastTag(x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_24); +if (lean_obj_tag(x_25) == 0) +{ +lean_object* x_26; lean_object* x_27; +x_26 = lean_ctor_get(x_25, 1); +lean_inc(x_26); +lean_dec(x_25); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +x_27 = lean_infer_type(x_1, x_8, x_9, x_10, x_11, x_26); +if (lean_obj_tag(x_27) == 0) +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; +x_28 = lean_ctor_get(x_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_MessageData_ofExpr(x_28); +x_31 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__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_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___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_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_2, x_34, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_29); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +x_36 = !lean_is_exclusive(x_35); +if (x_36 == 0) +{ +lean_object* x_37; lean_object* x_38; +x_37 = lean_ctor_get(x_35, 0); +lean_dec(x_37); +x_38 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__5; +lean_ctor_set(x_35, 0, x_38); +return x_35; } 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; uint64_t x_62; lean_object* x_63; lean_object* x_64; double x_65; uint8_t x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; -x_55 = lean_ctor_get(x_17, 0); -x_56 = lean_ctor_get(x_17, 1); -x_57 = lean_ctor_get(x_17, 2); -x_58 = lean_ctor_get(x_17, 4); -x_59 = lean_ctor_get(x_17, 5); -x_60 = lean_ctor_get(x_17, 6); -x_61 = lean_ctor_get(x_17, 7); -lean_inc(x_61); -lean_inc(x_60); -lean_inc(x_59); -lean_inc(x_58); -lean_inc(x_57); -lean_inc(x_56); -lean_inc(x_55); -lean_dec(x_17); -x_62 = lean_ctor_get_uint64(x_18, sizeof(void*)*1); -x_63 = lean_ctor_get(x_18, 0); -lean_inc(x_63); -if (lean_is_exclusive(x_18)) { - lean_ctor_release(x_18, 0); - x_64 = x_18; -} else { - lean_dec_ref(x_18); - x_64 = lean_box(0); -} -x_65 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__1; -x_66 = 0; -x_67 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__2; -x_68 = lean_alloc_ctor(0, 2, 17); -lean_ctor_set(x_68, 0, x_1); -lean_ctor_set(x_68, 1, x_67); -lean_ctor_set_float(x_68, sizeof(void*)*2, x_65); -lean_ctor_set_float(x_68, sizeof(void*)*2 + 8, x_65); -lean_ctor_set_uint8(x_68, sizeof(void*)*2 + 16, x_66); -x_69 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__3; -x_70 = lean_alloc_ctor(9, 3, 0); -lean_ctor_set(x_70, 0, x_68); -lean_ctor_set(x_70, 1, x_14); -lean_ctor_set(x_70, 2, x_69); -lean_inc(x_12); -lean_ctor_set(x_16, 1, x_70); -lean_ctor_set(x_16, 0, x_12); -x_71 = l_Lean_PersistentArray_push___rarg(x_63, x_16); -if (lean_is_scalar(x_64)) { - x_72 = lean_alloc_ctor(0, 1, 8); -} else { - x_72 = x_64; -} -lean_ctor_set(x_72, 0, x_71); -lean_ctor_set_uint64(x_72, sizeof(void*)*1, x_62); -x_73 = lean_alloc_ctor(0, 8, 0); -lean_ctor_set(x_73, 0, x_55); -lean_ctor_set(x_73, 1, x_56); -lean_ctor_set(x_73, 2, x_57); -lean_ctor_set(x_73, 3, x_72); -lean_ctor_set(x_73, 4, x_58); -lean_ctor_set(x_73, 5, x_59); -lean_ctor_set(x_73, 6, x_60); -lean_ctor_set(x_73, 7, x_61); -x_74 = lean_st_ref_set(x_10, x_73, x_20); -x_75 = lean_ctor_get(x_74, 1); -lean_inc(x_75); -if (lean_is_exclusive(x_74)) { - lean_ctor_release(x_74, 0); - lean_ctor_release(x_74, 1); - x_76 = x_74; -} else { - lean_dec_ref(x_74); - x_76 = lean_box(0); -} -x_77 = lean_box(0); -if (lean_is_scalar(x_76)) { - x_78 = lean_alloc_ctor(0, 2, 0); -} else { - x_78 = x_76; -} -lean_ctor_set(x_78, 0, x_77); -lean_ctor_set(x_78, 1, x_75); -return x_78; +lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_39 = lean_ctor_get(x_35, 1); +lean_inc(x_39); +lean_dec(x_35); +x_40 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__5; +x_41 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_41, 0, x_40); +lean_ctor_set(x_41, 1, x_39); +return x_41; } } else { -lean_object* x_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; uint64_t x_88; lean_object* x_89; lean_object* x_90; double x_91; uint8_t x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; -x_79 = lean_ctor_get(x_16, 1); -lean_inc(x_79); -lean_dec(x_16); -x_80 = lean_ctor_get(x_17, 0); -lean_inc(x_80); -x_81 = lean_ctor_get(x_17, 1); -lean_inc(x_81); -x_82 = lean_ctor_get(x_17, 2); -lean_inc(x_82); -x_83 = lean_ctor_get(x_17, 4); -lean_inc(x_83); -x_84 = lean_ctor_get(x_17, 5); -lean_inc(x_84); -x_85 = lean_ctor_get(x_17, 6); -lean_inc(x_85); -x_86 = lean_ctor_get(x_17, 7); -lean_inc(x_86); -if (lean_is_exclusive(x_17)) { - lean_ctor_release(x_17, 0); - lean_ctor_release(x_17, 1); - lean_ctor_release(x_17, 2); - lean_ctor_release(x_17, 3); - lean_ctor_release(x_17, 4); - lean_ctor_release(x_17, 5); - lean_ctor_release(x_17, 6); - lean_ctor_release(x_17, 7); - x_87 = x_17; -} else { - lean_dec_ref(x_17); - x_87 = lean_box(0); -} -x_88 = lean_ctor_get_uint64(x_18, sizeof(void*)*1); -x_89 = lean_ctor_get(x_18, 0); -lean_inc(x_89); -if (lean_is_exclusive(x_18)) { - lean_ctor_release(x_18, 0); - x_90 = x_18; -} else { - lean_dec_ref(x_18); - x_90 = lean_box(0); -} -x_91 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__1; -x_92 = 0; -x_93 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__2; -x_94 = lean_alloc_ctor(0, 2, 17); -lean_ctor_set(x_94, 0, x_1); -lean_ctor_set(x_94, 1, x_93); -lean_ctor_set_float(x_94, sizeof(void*)*2, x_91); -lean_ctor_set_float(x_94, sizeof(void*)*2 + 8, x_91); -lean_ctor_set_uint8(x_94, sizeof(void*)*2 + 16, x_92); -x_95 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__3; -x_96 = lean_alloc_ctor(9, 3, 0); -lean_ctor_set(x_96, 0, x_94); -lean_ctor_set(x_96, 1, x_14); -lean_ctor_set(x_96, 2, x_95); -lean_inc(x_12); -x_97 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_97, 0, x_12); -lean_ctor_set(x_97, 1, x_96); -x_98 = l_Lean_PersistentArray_push___rarg(x_89, x_97); -if (lean_is_scalar(x_90)) { - x_99 = lean_alloc_ctor(0, 1, 8); -} else { - x_99 = x_90; -} -lean_ctor_set(x_99, 0, x_98); -lean_ctor_set_uint64(x_99, sizeof(void*)*1, x_88); -if (lean_is_scalar(x_87)) { - x_100 = lean_alloc_ctor(0, 8, 0); -} else { - x_100 = x_87; -} -lean_ctor_set(x_100, 0, x_80); -lean_ctor_set(x_100, 1, x_81); -lean_ctor_set(x_100, 2, x_82); -lean_ctor_set(x_100, 3, x_99); -lean_ctor_set(x_100, 4, x_83); -lean_ctor_set(x_100, 5, x_84); -lean_ctor_set(x_100, 6, x_85); -lean_ctor_set(x_100, 7, x_86); -x_101 = lean_st_ref_set(x_10, x_100, x_79); -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; -} else { - lean_dec_ref(x_101); - x_103 = lean_box(0); -} -x_104 = lean_box(0); -if (lean_is_scalar(x_103)) { - x_105 = lean_alloc_ctor(0, 2, 0); -} else { - x_105 = x_103; -} -lean_ctor_set(x_105, 0, x_104); -lean_ctor_set(x_105, 1, x_102); -return x_105; -} -} -} -static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___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; -} -} -static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("checked: ", 9, 9); -return x_1; -} -} -static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__3() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__2; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; -} -} -static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__4() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__2; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___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) { -_start: -{ -lean_object* x_13; -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_1); -x_13 = l_Lean_Meta_check(x_1, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_13) == 0) -{ -lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; -x_14 = lean_ctor_get(x_13, 1); -lean_inc(x_14); -lean_dec(x_13); -lean_inc(x_2); -x_15 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1(x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_14); -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) -{ -uint8_t x_18; +uint8_t x_42; lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_2); -lean_dec(x_1); -x_18 = !lean_is_exclusive(x_15); -if (x_18 == 0) -{ -lean_object* x_19; lean_object* x_20; -x_19 = lean_ctor_get(x_15, 0); -lean_dec(x_19); -x_20 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__1; -lean_ctor_set(x_15, 0, x_20); -return x_15; -} -else +x_42 = !lean_is_exclusive(x_27); +if (x_42 == 0) { -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_15, 1); -lean_inc(x_21); -lean_dec(x_15); -x_22 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__1; -x_23 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_23, 0, x_22); -lean_ctor_set(x_23, 1, x_21); -return x_23; -} +return x_27; } else { -lean_object* x_24; lean_object* x_25; -x_24 = lean_ctor_get(x_15, 1); -lean_inc(x_24); -lean_dec(x_15); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -x_25 = lean_infer_type(x_1, x_8, x_9, x_10, x_11, x_24); -if (lean_obj_tag(x_25) == 0) -{ -lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; uint8_t x_34; -x_26 = lean_ctor_get(x_25, 0); -lean_inc(x_26); -x_27 = lean_ctor_get(x_25, 1); -lean_inc(x_27); -lean_dec(x_25); -x_28 = l_Lean_MessageData_ofExpr(x_26); -x_29 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__3; -x_30 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_30, 0, x_29); -lean_ctor_set(x_30, 1, x_28); -x_31 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__4; -x_32 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_32, 0, x_30); -lean_ctor_set(x_32, 1, x_31); -x_33 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2(x_2, x_32, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_27); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -x_34 = !lean_is_exclusive(x_33); -if (x_34 == 0) -{ -lean_object* x_35; lean_object* x_36; -x_35 = lean_ctor_get(x_33, 0); -lean_dec(x_35); -x_36 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__1; -lean_ctor_set(x_33, 0, x_36); -return x_33; +lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_43 = lean_ctor_get(x_27, 0); +x_44 = lean_ctor_get(x_27, 1); +lean_inc(x_44); +lean_inc(x_43); +lean_dec(x_27); +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_37; lean_object* x_38; lean_object* x_39; -x_37 = lean_ctor_get(x_33, 1); -lean_inc(x_37); -lean_dec(x_33); -x_38 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__1; -x_39 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_39, 0, x_38); -lean_ctor_set(x_39, 1, x_37); -return x_39; } } else { -uint8_t x_40; +uint8_t x_46; lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_2); -x_40 = !lean_is_exclusive(x_25); -if (x_40 == 0) +lean_dec(x_1); +x_46 = !lean_is_exclusive(x_25); +if (x_46 == 0) { return x_25; } else { -lean_object* x_41; lean_object* x_42; lean_object* x_43; -x_41 = lean_ctor_get(x_25, 0); -x_42 = lean_ctor_get(x_25, 1); -lean_inc(x_42); -lean_inc(x_41); +lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_47 = lean_ctor_get(x_25, 0); +x_48 = lean_ctor_get(x_25, 1); +lean_inc(x_48); +lean_inc(x_47); lean_dec(x_25); -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_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_44; +uint8_t x_50; lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_2); lean_dec(x_1); -x_44 = !lean_is_exclusive(x_13); -if (x_44 == 0) +x_50 = !lean_is_exclusive(x_13); +if (x_50 == 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_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_inc(x_52); +lean_inc(x_51); 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; +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; } } } } -static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__1() { +static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__1() { _start: { lean_object* x_1; @@ -3764,7 +3518,7 @@ x_1 = lean_mk_string_unchecked("grind", 5, 5); return x_1; } } -static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__2() { +static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__2() { _start: { lean_object* x_1; @@ -3772,7 +3526,7 @@ x_1 = lean_mk_string_unchecked("debug", 5, 5); return x_1; } } -static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__3() { +static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__3() { _start: { lean_object* x_1; @@ -3780,18 +3534,18 @@ x_1 = lean_mk_string_unchecked("proofs", 6, 6); return x_1; } } -static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__4() { +static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__1; -x_2 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__2; -x_3 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__3; +x_1 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__1; +x_2 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__2; +x_3 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__3; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__5() { +static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__5() { _start: { lean_object* x_1; @@ -3799,16 +3553,16 @@ x_1 = lean_mk_string_unchecked(" = ", 3, 3); return x_1; } } -static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__6() { +static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__6() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__5; +x_1 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__5; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { _start: { if (lean_obj_tag(x_5) == 0) @@ -3861,8 +3615,8 @@ lean_inc(x_23); x_24 = lean_ctor_get(x_22, 1); lean_inc(x_24); lean_dec(x_22); -x_25 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__4; -x_26 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1(x_25, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_24); +x_25 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__4; +x_26 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_25, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_24); x_27 = lean_ctor_get(x_26, 0); lean_inc(x_27); x_28 = lean_unbox(x_27); @@ -3880,7 +3634,7 @@ lean_inc(x_15); lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); -x_31 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1(x_23, x_25, x_30, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_29); +x_31 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1(x_23, x_25, x_30, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_29); if (lean_obj_tag(x_31) == 0) { lean_object* x_32; @@ -3981,47 +3735,54 @@ uint8_t x_46; x_46 = !lean_is_exclusive(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; lean_object* x_57; lean_object* x_58; +lean_object* x_47; lean_object* x_48; lean_object* x_49; x_47 = lean_ctor_get(x_26, 1); x_48 = lean_ctor_get(x_26, 0); lean_dec(x_48); +x_49 = l_Lean_Meta_Grind_updateLastTag(x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_47); +if (lean_obj_tag(x_49) == 0) +{ +lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_50 = lean_ctor_get(x_49, 1); +lean_inc(x_50); +lean_dec(x_49); lean_inc(x_3); -x_49 = l_Lean_MessageData_ofExpr(x_3); -x_50 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__4; +x_51 = l_Lean_MessageData_ofExpr(x_3); +x_52 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__4; lean_ctor_set_tag(x_26, 7); -lean_ctor_set(x_26, 1, x_49); -lean_ctor_set(x_26, 0, x_50); -x_51 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__6; +lean_ctor_set(x_26, 1, x_51); +lean_ctor_set(x_26, 0, x_52); +x_53 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__6; lean_ctor_set_tag(x_5, 7); -lean_ctor_set(x_5, 1, x_51); +lean_ctor_set(x_5, 1, x_53); lean_ctor_set(x_5, 0, x_26); -x_52 = l_Lean_MessageData_ofExpr(x_19); -x_53 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_53, 0, x_5); -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_50); -x_55 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2(x_25, x_54, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_47); -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_54 = l_Lean_MessageData_ofExpr(x_19); +x_55 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_55, 0, x_5); +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_52); +x_57 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_25, x_56, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_50); +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_15); lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); -x_58 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1(x_23, x_25, x_56, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_57); -lean_dec(x_56); -if (lean_obj_tag(x_58) == 0) +x_60 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1(x_23, x_25, x_58, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_59); +lean_dec(x_58); +if (lean_obj_tag(x_60) == 0) { -lean_object* x_59; -x_59 = lean_ctor_get(x_58, 0); -lean_inc(x_59); -if (lean_obj_tag(x_59) == 0) +lean_object* x_61; +x_61 = lean_ctor_get(x_60, 0); +lean_inc(x_61); +if (lean_obj_tag(x_61) == 0) { -uint8_t x_60; +uint8_t x_62; lean_dec(x_20); lean_dec(x_15); lean_dec(x_14); @@ -4032,52 +3793,52 @@ lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_3); -x_60 = !lean_is_exclusive(x_58); -if (x_60 == 0) +x_62 = !lean_is_exclusive(x_60); +if (x_62 == 0) { -lean_object* x_61; lean_object* x_62; -x_61 = lean_ctor_get(x_58, 0); +lean_object* x_63; lean_object* x_64; +x_63 = lean_ctor_get(x_60, 0); +lean_dec(x_63); +x_64 = lean_ctor_get(x_61, 0); +lean_inc(x_64); lean_dec(x_61); -x_62 = lean_ctor_get(x_59, 0); -lean_inc(x_62); -lean_dec(x_59); -lean_ctor_set(x_58, 0, x_62); -return x_58; +lean_ctor_set(x_60, 0, x_64); +return x_60; } else { -lean_object* x_63; lean_object* x_64; lean_object* x_65; -x_63 = lean_ctor_get(x_58, 1); -lean_inc(x_63); -lean_dec(x_58); -x_64 = lean_ctor_get(x_59, 0); -lean_inc(x_64); -lean_dec(x_59); -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; +lean_object* x_65; lean_object* x_66; lean_object* x_67; +x_65 = lean_ctor_get(x_60, 1); +lean_inc(x_65); +lean_dec(x_60); +x_66 = lean_ctor_get(x_61, 0); +lean_inc(x_66); +lean_dec(x_61); +x_67 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_67, 0, x_66); +lean_ctor_set(x_67, 1, x_65); +return x_67; } } else { -lean_object* x_66; lean_object* x_67; -x_66 = lean_ctor_get(x_58, 1); -lean_inc(x_66); -lean_dec(x_58); -x_67 = lean_ctor_get(x_59, 0); -lean_inc(x_67); -lean_dec(x_59); +lean_object* x_68; lean_object* x_69; +x_68 = lean_ctor_get(x_60, 1); +lean_inc(x_68); +lean_dec(x_60); +x_69 = lean_ctor_get(x_61, 0); +lean_inc(x_69); +lean_dec(x_61); x_5 = x_20; -x_6 = x_67; +x_6 = x_69; x_7 = lean_box(0); -x_16 = x_66; +x_16 = x_68; goto _start; } } else { -uint8_t x_69; +uint8_t x_71; lean_dec(x_20); lean_dec(x_15); lean_dec(x_14); @@ -4088,69 +3849,113 @@ lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_3); -x_69 = !lean_is_exclusive(x_58); -if (x_69 == 0) +x_71 = !lean_is_exclusive(x_60); +if (x_71 == 0) { -return x_58; +return x_60; } else { -lean_object* x_70; lean_object* x_71; lean_object* x_72; -x_70 = lean_ctor_get(x_58, 0); -x_71 = lean_ctor_get(x_58, 1); -lean_inc(x_71); -lean_inc(x_70); -lean_dec(x_58); -x_72 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_72, 0, x_70); -lean_ctor_set(x_72, 1, x_71); -return x_72; +lean_object* x_72; lean_object* x_73; lean_object* x_74; +x_72 = lean_ctor_get(x_60, 0); +x_73 = lean_ctor_get(x_60, 1); +lean_inc(x_73); +lean_inc(x_72); +lean_dec(x_60); +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 { -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_73 = lean_ctor_get(x_26, 1); -lean_inc(x_73); +uint8_t x_75; +lean_free_object(x_26); +lean_dec(x_23); +lean_free_object(x_5); +lean_dec(x_20); +lean_dec(x_19); +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_3); +x_75 = !lean_is_exclusive(x_49); +if (x_75 == 0) +{ +return x_49; +} +else +{ +lean_object* x_76; lean_object* x_77; lean_object* x_78; +x_76 = lean_ctor_get(x_49, 0); +x_77 = lean_ctor_get(x_49, 1); +lean_inc(x_77); +lean_inc(x_76); +lean_dec(x_49); +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_79; lean_object* x_80; +x_79 = lean_ctor_get(x_26, 1); +lean_inc(x_79); lean_dec(x_26); +x_80 = l_Lean_Meta_Grind_updateLastTag(x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_79); +if (lean_obj_tag(x_80) == 0) +{ +lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; +x_81 = lean_ctor_get(x_80, 1); +lean_inc(x_81); +lean_dec(x_80); lean_inc(x_3); -x_74 = l_Lean_MessageData_ofExpr(x_3); -x_75 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__4; -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_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__6; +x_82 = l_Lean_MessageData_ofExpr(x_3); +x_83 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__4; +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_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__6; lean_ctor_set_tag(x_5, 7); -lean_ctor_set(x_5, 1, x_77); -lean_ctor_set(x_5, 0, x_76); -x_78 = l_Lean_MessageData_ofExpr(x_19); -x_79 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_79, 0, x_5); -lean_ctor_set(x_79, 1, x_78); -x_80 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_80, 0, x_79); -lean_ctor_set(x_80, 1, x_75); -x_81 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2(x_25, x_80, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_73); -x_82 = lean_ctor_get(x_81, 0); -lean_inc(x_82); -x_83 = lean_ctor_get(x_81, 1); -lean_inc(x_83); -lean_dec(x_81); +lean_ctor_set(x_5, 1, x_85); +lean_ctor_set(x_5, 0, x_84); +x_86 = l_Lean_MessageData_ofExpr(x_19); +x_87 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_87, 0, x_5); +lean_ctor_set(x_87, 1, x_86); +x_88 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_88, 0, x_87); +lean_ctor_set(x_88, 1, x_83); +x_89 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_25, x_88, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_81); +x_90 = lean_ctor_get(x_89, 0); +lean_inc(x_90); +x_91 = lean_ctor_get(x_89, 1); +lean_inc(x_91); +lean_dec(x_89); lean_inc(x_15); lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); -x_84 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1(x_23, x_25, x_82, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_83); -lean_dec(x_82); -if (lean_obj_tag(x_84) == 0) +x_92 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1(x_23, x_25, x_90, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_91); +lean_dec(x_90); +if (lean_obj_tag(x_92) == 0) { -lean_object* x_85; -x_85 = lean_ctor_get(x_84, 0); -lean_inc(x_85); -if (lean_obj_tag(x_85) == 0) +lean_object* x_93; +x_93 = lean_ctor_get(x_92, 0); +lean_inc(x_93); +if (lean_obj_tag(x_93) == 0) { -lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; +lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_dec(x_20); lean_dec(x_15); lean_dec(x_14); @@ -4161,47 +3966,47 @@ lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_3); -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_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_84); - x_87 = lean_box(0); + lean_dec_ref(x_92); + x_95 = lean_box(0); } -x_88 = lean_ctor_get(x_85, 0); -lean_inc(x_88); -lean_dec(x_85); -if (lean_is_scalar(x_87)) { - x_89 = lean_alloc_ctor(0, 2, 0); +x_96 = lean_ctor_get(x_93, 0); +lean_inc(x_96); +lean_dec(x_93); +if (lean_is_scalar(x_95)) { + x_97 = lean_alloc_ctor(0, 2, 0); } else { - x_89 = x_87; + x_97 = x_95; } -lean_ctor_set(x_89, 0, x_88); -lean_ctor_set(x_89, 1, x_86); -return x_89; +lean_ctor_set(x_97, 0, x_96); +lean_ctor_set(x_97, 1, x_94); +return x_97; } else { -lean_object* x_90; lean_object* x_91; -x_90 = lean_ctor_get(x_84, 1); -lean_inc(x_90); -lean_dec(x_84); -x_91 = lean_ctor_get(x_85, 0); -lean_inc(x_91); -lean_dec(x_85); +lean_object* x_98; lean_object* x_99; +x_98 = lean_ctor_get(x_92, 1); +lean_inc(x_98); +lean_dec(x_92); +x_99 = lean_ctor_get(x_93, 0); +lean_inc(x_99); +lean_dec(x_93); x_5 = x_20; -x_6 = x_91; +x_6 = x_99; x_7 = lean_box(0); -x_16 = x_90; +x_16 = x_98; goto _start; } } else { -lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; +lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_dec(x_20); lean_dec(x_15); lean_dec(x_14); @@ -4212,33 +4017,71 @@ lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_3); -x_93 = lean_ctor_get(x_84, 0); -lean_inc(x_93); -x_94 = lean_ctor_get(x_84, 1); -lean_inc(x_94); -if (lean_is_exclusive(x_84)) { - lean_ctor_release(x_84, 0); - lean_ctor_release(x_84, 1); - x_95 = x_84; +x_101 = lean_ctor_get(x_92, 0); +lean_inc(x_101); +x_102 = lean_ctor_get(x_92, 1); +lean_inc(x_102); +if (lean_is_exclusive(x_92)) { + lean_ctor_release(x_92, 0); + lean_ctor_release(x_92, 1); + x_103 = x_92; } else { - lean_dec_ref(x_84); - x_95 = lean_box(0); + lean_dec_ref(x_92); + x_103 = lean_box(0); } -if (lean_is_scalar(x_95)) { - x_96 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_103)) { + x_104 = lean_alloc_ctor(1, 2, 0); } else { - x_96 = x_95; + x_104 = x_103; } -lean_ctor_set(x_96, 0, x_93); -lean_ctor_set(x_96, 1, x_94); -return x_96; +lean_ctor_set(x_104, 0, x_101); +lean_ctor_set(x_104, 1, x_102); +return x_104; +} +} +else +{ +lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; +lean_dec(x_23); +lean_free_object(x_5); +lean_dec(x_20); +lean_dec(x_19); +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_3); +x_105 = lean_ctor_get(x_80, 0); +lean_inc(x_105); +x_106 = lean_ctor_get(x_80, 1); +lean_inc(x_106); +if (lean_is_exclusive(x_80)) { + lean_ctor_release(x_80, 0); + lean_ctor_release(x_80, 1); + x_107 = x_80; +} else { + lean_dec_ref(x_80); + 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(x_108, 0, x_105); +lean_ctor_set(x_108, 1, x_106); +return x_108; } } } } else { -uint8_t x_97; +uint8_t x_109; lean_free_object(x_5); lean_dec(x_20); lean_dec(x_19); @@ -4251,50 +4094,50 @@ lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_3); -x_97 = !lean_is_exclusive(x_22); -if (x_97 == 0) +x_109 = !lean_is_exclusive(x_22); +if (x_109 == 0) { return x_22; } else { -lean_object* x_98; lean_object* x_99; lean_object* x_100; -x_98 = lean_ctor_get(x_22, 0); -x_99 = lean_ctor_get(x_22, 1); -lean_inc(x_99); -lean_inc(x_98); +lean_object* x_110; lean_object* x_111; lean_object* x_112; +x_110 = lean_ctor_get(x_22, 0); +x_111 = lean_ctor_get(x_22, 1); +lean_inc(x_111); +lean_inc(x_110); lean_dec(x_22); -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; +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_101; +lean_object* x_113; lean_free_object(x_5); lean_dec(x_19); -x_101 = lean_box(0); +x_113 = lean_box(0); x_5 = x_20; -x_6 = x_101; +x_6 = x_113; x_7 = lean_box(0); goto _start; } } else { -lean_object* x_103; lean_object* x_104; uint8_t x_105; -x_103 = lean_ctor_get(x_5, 0); -x_104 = lean_ctor_get(x_5, 1); -lean_inc(x_104); -lean_inc(x_103); +lean_object* x_115; lean_object* x_116; uint8_t x_117; +x_115 = lean_ctor_get(x_5, 0); +x_116 = lean_ctor_get(x_5, 1); +lean_inc(x_116); +lean_inc(x_115); lean_dec(x_5); -x_105 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_3, x_103); -if (x_105 == 0) +x_117 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_3, x_115); +if (x_117 == 0) { -lean_object* x_106; +lean_object* x_118; lean_inc(x_15); lean_inc(x_14); lean_inc(x_13); @@ -4303,45 +4146,45 @@ lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); -lean_inc(x_103); +lean_inc(x_115); lean_inc(x_3); -x_106 = l_Lean_Meta_Grind_mkEqHEqProof(x_3, x_103, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); -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; uint8_t x_112; -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 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__4; -x_110 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1(x_109, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_108); -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_dec(x_103); -x_113 = lean_ctor_get(x_110, 1); -lean_inc(x_113); -lean_dec(x_110); -x_114 = lean_box(0); +x_118 = l_Lean_Meta_Grind_mkEqHEqProof(x_3, x_115, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +if (lean_obj_tag(x_118) == 0) +{ +lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; uint8_t x_124; +x_119 = lean_ctor_get(x_118, 0); +lean_inc(x_119); +x_120 = lean_ctor_get(x_118, 1); +lean_inc(x_120); +lean_dec(x_118); +x_121 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__4; +x_122 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_121, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_120); +x_123 = lean_ctor_get(x_122, 0); +lean_inc(x_123); +x_124 = lean_unbox(x_123); +lean_dec(x_123); +if (x_124 == 0) +{ +lean_object* x_125; lean_object* x_126; lean_object* x_127; +lean_dec(x_115); +x_125 = lean_ctor_get(x_122, 1); +lean_inc(x_125); +lean_dec(x_122); +x_126 = lean_box(0); lean_inc(x_15); lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); -x_115 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1(x_107, x_109, x_114, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_113); -if (lean_obj_tag(x_115) == 0) +x_127 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1(x_119, x_121, x_126, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_125); +if (lean_obj_tag(x_127) == 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_128; +x_128 = lean_ctor_get(x_127, 0); +lean_inc(x_128); +if (lean_obj_tag(x_128) == 0) { -lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; -lean_dec(x_104); +lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; +lean_dec(x_116); lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); @@ -4351,48 +4194,48 @@ lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_3); -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; +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; } else { - lean_dec_ref(x_115); - x_118 = lean_box(0); -} -x_119 = lean_ctor_get(x_116, 0); -lean_inc(x_119); -lean_dec(x_116); -if (lean_is_scalar(x_118)) { - x_120 = lean_alloc_ctor(0, 2, 0); + lean_dec_ref(x_127); + x_130 = lean_box(0); +} +x_131 = lean_ctor_get(x_128, 0); +lean_inc(x_131); +lean_dec(x_128); +if (lean_is_scalar(x_130)) { + x_132 = lean_alloc_ctor(0, 2, 0); } else { - x_120 = x_118; + x_132 = x_130; } -lean_ctor_set(x_120, 0, x_119); -lean_ctor_set(x_120, 1, x_117); -return x_120; +lean_ctor_set(x_132, 0, x_131); +lean_ctor_set(x_132, 1, x_129); +return x_132; } else { -lean_object* x_121; lean_object* x_122; -x_121 = lean_ctor_get(x_115, 1); -lean_inc(x_121); -lean_dec(x_115); -x_122 = lean_ctor_get(x_116, 0); -lean_inc(x_122); -lean_dec(x_116); -x_5 = x_104; -x_6 = x_122; +lean_object* x_133; lean_object* x_134; +x_133 = lean_ctor_get(x_127, 1); +lean_inc(x_133); +lean_dec(x_127); +x_134 = lean_ctor_get(x_128, 0); +lean_inc(x_134); +lean_dec(x_128); +x_5 = x_116; +x_6 = x_134; x_7 = lean_box(0); -x_16 = x_121; +x_16 = x_133; goto _start; } } else { -lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; -lean_dec(x_104); +lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; +lean_dec(x_116); lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); @@ -4402,84 +4245,91 @@ lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_3); -x_124 = lean_ctor_get(x_115, 0); -lean_inc(x_124); -x_125 = lean_ctor_get(x_115, 1); -lean_inc(x_125); -if (lean_is_exclusive(x_115)) { - lean_ctor_release(x_115, 0); - lean_ctor_release(x_115, 1); - x_126 = x_115; +x_136 = lean_ctor_get(x_127, 0); +lean_inc(x_136); +x_137 = lean_ctor_get(x_127, 1); +lean_inc(x_137); +if (lean_is_exclusive(x_127)) { + lean_ctor_release(x_127, 0); + lean_ctor_release(x_127, 1); + x_138 = x_127; } else { - lean_dec_ref(x_115); - x_126 = lean_box(0); + lean_dec_ref(x_127); + x_138 = lean_box(0); } -if (lean_is_scalar(x_126)) { - x_127 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_138)) { + x_139 = lean_alloc_ctor(1, 2, 0); } else { - x_127 = x_126; + x_139 = x_138; } -lean_ctor_set(x_127, 0, x_124); -lean_ctor_set(x_127, 1, x_125); -return x_127; +lean_ctor_set(x_139, 0, x_136); +lean_ctor_set(x_139, 1, x_137); +return x_139; } } else { -lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; -x_128 = lean_ctor_get(x_110, 1); -lean_inc(x_128); -if (lean_is_exclusive(x_110)) { - lean_ctor_release(x_110, 0); - lean_ctor_release(x_110, 1); - x_129 = x_110; +lean_object* x_140; lean_object* x_141; lean_object* x_142; +x_140 = lean_ctor_get(x_122, 1); +lean_inc(x_140); +if (lean_is_exclusive(x_122)) { + lean_ctor_release(x_122, 0); + lean_ctor_release(x_122, 1); + x_141 = x_122; } else { - lean_dec_ref(x_110); - x_129 = lean_box(0); + lean_dec_ref(x_122); + x_141 = lean_box(0); } +x_142 = l_Lean_Meta_Grind_updateLastTag(x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_140); +if (lean_obj_tag(x_142) == 0) +{ +lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; +x_143 = lean_ctor_get(x_142, 1); +lean_inc(x_143); +lean_dec(x_142); lean_inc(x_3); -x_130 = l_Lean_MessageData_ofExpr(x_3); -x_131 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__4; -if (lean_is_scalar(x_129)) { - x_132 = lean_alloc_ctor(7, 2, 0); +x_144 = l_Lean_MessageData_ofExpr(x_3); +x_145 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__4; +if (lean_is_scalar(x_141)) { + x_146 = lean_alloc_ctor(7, 2, 0); } else { - x_132 = x_129; - lean_ctor_set_tag(x_132, 7); + x_146 = x_141; + lean_ctor_set_tag(x_146, 7); } -lean_ctor_set(x_132, 0, x_131); -lean_ctor_set(x_132, 1, x_130); -x_133 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__6; -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_103); -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 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_137, 0, x_136); -lean_ctor_set(x_137, 1, x_131); -x_138 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2(x_109, x_137, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_128); -x_139 = lean_ctor_get(x_138, 0); -lean_inc(x_139); -x_140 = lean_ctor_get(x_138, 1); -lean_inc(x_140); -lean_dec(x_138); +lean_ctor_set(x_146, 0, x_145); +lean_ctor_set(x_146, 1, x_144); +x_147 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__6; +x_148 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_148, 0, x_146); +lean_ctor_set(x_148, 1, x_147); +x_149 = l_Lean_MessageData_ofExpr(x_115); +x_150 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_150, 0, x_148); +lean_ctor_set(x_150, 1, x_149); +x_151 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_151, 0, x_150); +lean_ctor_set(x_151, 1, x_145); +x_152 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_121, x_151, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_143); +x_153 = lean_ctor_get(x_152, 0); +lean_inc(x_153); +x_154 = lean_ctor_get(x_152, 1); +lean_inc(x_154); +lean_dec(x_152); lean_inc(x_15); lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); -x_141 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1(x_107, x_109, x_139, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_140); -lean_dec(x_139); -if (lean_obj_tag(x_141) == 0) +x_155 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1(x_119, x_121, x_153, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_154); +lean_dec(x_153); +if (lean_obj_tag(x_155) == 0) { -lean_object* x_142; -x_142 = lean_ctor_get(x_141, 0); -lean_inc(x_142); -if (lean_obj_tag(x_142) == 0) +lean_object* x_156; +x_156 = lean_ctor_get(x_155, 0); +lean_inc(x_156); +if (lean_obj_tag(x_156) == 0) { -lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; -lean_dec(x_104); +lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; +lean_dec(x_116); lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); @@ -4489,48 +4339,86 @@ lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_3); -x_143 = lean_ctor_get(x_141, 1); -lean_inc(x_143); -if (lean_is_exclusive(x_141)) { - lean_ctor_release(x_141, 0); - lean_ctor_release(x_141, 1); - x_144 = x_141; +x_157 = lean_ctor_get(x_155, 1); +lean_inc(x_157); +if (lean_is_exclusive(x_155)) { + lean_ctor_release(x_155, 0); + lean_ctor_release(x_155, 1); + x_158 = x_155; } else { - lean_dec_ref(x_141); - x_144 = lean_box(0); -} -x_145 = lean_ctor_get(x_142, 0); -lean_inc(x_145); -lean_dec(x_142); -if (lean_is_scalar(x_144)) { - x_146 = lean_alloc_ctor(0, 2, 0); + lean_dec_ref(x_155); + x_158 = lean_box(0); +} +x_159 = lean_ctor_get(x_156, 0); +lean_inc(x_159); +lean_dec(x_156); +if (lean_is_scalar(x_158)) { + x_160 = lean_alloc_ctor(0, 2, 0); } else { - x_146 = x_144; + x_160 = x_158; } -lean_ctor_set(x_146, 0, x_145); -lean_ctor_set(x_146, 1, x_143); -return x_146; +lean_ctor_set(x_160, 0, x_159); +lean_ctor_set(x_160, 1, x_157); +return x_160; } else { -lean_object* x_147; lean_object* x_148; -x_147 = lean_ctor_get(x_141, 1); -lean_inc(x_147); -lean_dec(x_141); -x_148 = lean_ctor_get(x_142, 0); -lean_inc(x_148); -lean_dec(x_142); -x_5 = x_104; -x_6 = x_148; +lean_object* x_161; lean_object* x_162; +x_161 = lean_ctor_get(x_155, 1); +lean_inc(x_161); +lean_dec(x_155); +x_162 = lean_ctor_get(x_156, 0); +lean_inc(x_162); +lean_dec(x_156); +x_5 = x_116; +x_6 = x_162; x_7 = lean_box(0); -x_16 = x_147; +x_16 = x_161; goto _start; } } else { -lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; -lean_dec(x_104); +lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; +lean_dec(x_116); +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_3); +x_164 = lean_ctor_get(x_155, 0); +lean_inc(x_164); +x_165 = lean_ctor_get(x_155, 1); +lean_inc(x_165); +if (lean_is_exclusive(x_155)) { + lean_ctor_release(x_155, 0); + lean_ctor_release(x_155, 1); + x_166 = x_155; +} else { + lean_dec_ref(x_155); + x_166 = lean_box(0); +} +if (lean_is_scalar(x_166)) { + x_167 = 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; +} +} +else +{ +lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; +lean_dec(x_141); +lean_dec(x_119); +lean_dec(x_116); +lean_dec(x_115); lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); @@ -4540,34 +4428,34 @@ lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_3); -x_150 = lean_ctor_get(x_141, 0); -lean_inc(x_150); -x_151 = lean_ctor_get(x_141, 1); -lean_inc(x_151); -if (lean_is_exclusive(x_141)) { - lean_ctor_release(x_141, 0); - lean_ctor_release(x_141, 1); - x_152 = x_141; +x_168 = lean_ctor_get(x_142, 0); +lean_inc(x_168); +x_169 = lean_ctor_get(x_142, 1); +lean_inc(x_169); +if (lean_is_exclusive(x_142)) { + lean_ctor_release(x_142, 0); + lean_ctor_release(x_142, 1); + x_170 = x_142; } else { - lean_dec_ref(x_141); - x_152 = lean_box(0); + lean_dec_ref(x_142); + x_170 = lean_box(0); } -if (lean_is_scalar(x_152)) { - x_153 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_170)) { + x_171 = lean_alloc_ctor(1, 2, 0); } else { - x_153 = x_152; + x_171 = x_170; } -lean_ctor_set(x_153, 0, x_150); -lean_ctor_set(x_153, 1, x_151); -return x_153; +lean_ctor_set(x_171, 0, x_168); +lean_ctor_set(x_171, 1, x_169); +return x_171; } } } else { -lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; -lean_dec(x_104); -lean_dec(x_103); +lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; +lean_dec(x_116); +lean_dec(x_115); lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); @@ -4577,35 +4465,35 @@ lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_3); -x_154 = lean_ctor_get(x_106, 0); -lean_inc(x_154); -x_155 = lean_ctor_get(x_106, 1); -lean_inc(x_155); -if (lean_is_exclusive(x_106)) { - lean_ctor_release(x_106, 0); - lean_ctor_release(x_106, 1); - x_156 = x_106; +x_172 = lean_ctor_get(x_118, 0); +lean_inc(x_172); +x_173 = lean_ctor_get(x_118, 1); +lean_inc(x_173); +if (lean_is_exclusive(x_118)) { + lean_ctor_release(x_118, 0); + lean_ctor_release(x_118, 1); + x_174 = x_118; } else { - lean_dec_ref(x_106); - x_156 = lean_box(0); + lean_dec_ref(x_118); + x_174 = lean_box(0); } -if (lean_is_scalar(x_156)) { - x_157 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_174)) { + x_175 = lean_alloc_ctor(1, 2, 0); } else { - x_157 = x_156; + x_175 = x_174; } -lean_ctor_set(x_157, 0, x_154); -lean_ctor_set(x_157, 1, x_155); -return x_157; +lean_ctor_set(x_175, 0, x_172); +lean_ctor_set(x_175, 1, x_173); +return x_175; } } else { -lean_object* x_158; -lean_dec(x_103); -x_158 = lean_box(0); -x_5 = x_104; -x_6 = x_158; +lean_object* x_176; +lean_dec(x_115); +x_176 = lean_box(0); +x_5 = x_116; +x_6 = x_176; x_7 = lean_box(0); goto _start; } @@ -4613,7 +4501,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___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) { _start: { if (lean_obj_tag(x_4) == 0) @@ -4652,7 +4540,7 @@ lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_1); -x_20 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3(x_1, x_2, x_17, x_1, x_1, x_19, lean_box(0), x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +x_20 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1(x_1, x_2, x_17, x_1, x_1, x_19, lean_box(0), x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); if (lean_obj_tag(x_20) == 0) { lean_object* x_21; @@ -4700,7 +4588,7 @@ return x_26; } } } -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { _start: { if (lean_obj_tag(x_4) == 0) @@ -4739,7 +4627,7 @@ lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc_n(x_17, 2); -x_21 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__4(x_17, x_19, x_17, x_17, x_20, lean_box(0), x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +x_21 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2(x_17, x_19, x_17, x_17, x_20, lean_box(0), x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); lean_dec(x_17); if (lean_obj_tag(x_21) == 0) { @@ -4803,7 +4691,7 @@ lean_dec(x_10); x_13 = lean_box(0); x_14 = lean_box(0); lean_inc(x_11); -x_15 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__5(x_11, x_13, x_11, x_11, x_14, lean_box(0), x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_12); +x_15 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3(x_11, x_13, x_11, x_11, x_14, lean_box(0), x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_12); lean_dec(x_11); if (lean_obj_tag(x_15) == 0) { @@ -4884,43 +4772,11 @@ return x_27; } } } -LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___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_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___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_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_11; -} -} -LEAN_EXPORT lean_object* l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -lean_object* x_12; -x_12 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -return x_12; -} -} -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___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_13; -x_13 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +x_13 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); @@ -4929,32 +4785,32 @@ lean_dec(x_3); return x_13; } } -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { _start: { lean_object* x_17; -x_17 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +x_17 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); return x_17; } } -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { _start: { lean_object* x_16; -x_16 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +x_16 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___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); lean_dec(x_3); lean_dec(x_2); return x_16; } } -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { _start: { lean_object* x_16; -x_16 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__5(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +x_16 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); @@ -5934,7 +5790,7 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_20 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__1; +x_20 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__5; lean_ctor_set(x_14, 0, x_20); return x_14; } @@ -5952,7 +5808,7 @@ if (x_22 == 0) lean_object* x_23; lean_object* x_24; x_23 = lean_ctor_get(x_21, 0); lean_dec(x_23); -x_24 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__1; +x_24 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__5; lean_ctor_set(x_21, 0, x_24); return x_21; } @@ -5962,7 +5818,7 @@ 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_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__1; +x_26 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__5; x_27 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_27, 0, x_26); lean_ctor_set(x_27, 1, x_25); @@ -6016,7 +5872,7 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_35 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__1; +x_35 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__5; x_36 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_36, 0, x_35); lean_ctor_set(x_36, 1, x_32); @@ -6039,7 +5895,7 @@ if (lean_is_exclusive(x_37)) { lean_dec_ref(x_37); x_39 = lean_box(0); } -x_40 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__1; +x_40 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__5; if (lean_is_scalar(x_39)) { x_41 = lean_alloc_ctor(0, 2, 0); } else { @@ -6680,61 +6536,58 @@ l_panic___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParent lean_mark_persistent(l_panic___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__2___closed__1); l_panic___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__2___closed__2 = _init_l_panic___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__2___closed__2(); lean_mark_persistent(l_panic___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__2___closed__2); +l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__1 = _init_l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__1(); +lean_mark_persistent(l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__1); +l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__2 = _init_l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__2(); +lean_mark_persistent(l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__2); +l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__3 = _init_l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__3(); +lean_mark_persistent(l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__3); +l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__4 = _init_l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__4(); +lean_mark_persistent(l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__4); +l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__5 = _init_l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__5(); +lean_mark_persistent(l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__5); l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___closed__1 = _init_l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___closed__1(); lean_mark_persistent(l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___closed__1); -l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___closed__2 = _init_l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___closed__2(); -lean_mark_persistent(l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___closed__2); -l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___closed__3 = _init_l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___closed__3(); -lean_mark_persistent(l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___closed__3); -l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___closed__4 = _init_l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___closed__4(); -lean_mark_persistent(l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___closed__4); -l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___closed__5 = _init_l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___closed__5(); -lean_mark_persistent(l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___closed__5); l___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___closed__1(); lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___closed__1); l___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___closed__2(); lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___closed__2); l___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___closed__3(); lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___closed__3); -l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__1 = _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__1(); -lean_mark_persistent(l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__1); -l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__2 = _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__2(); -lean_mark_persistent(l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__2); -l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__3 = _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__3(); -lean_mark_persistent(l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__3); -l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__4 = _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__4(); -lean_mark_persistent(l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__4); -l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__5 = _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__5(); -lean_mark_persistent(l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__5); -l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__6 = _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__6(); -lean_mark_persistent(l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__6); -l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__7 = _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__7(); -lean_mark_persistent(l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__7); -l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__1 = _init_l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__1(); -l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__2 = _init_l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__2(); -lean_mark_persistent(l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__2); -l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__3 = _init_l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__3(); -lean_mark_persistent(l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__3); -l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__1 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__1(); -lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__1); -l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__2 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__2(); -lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__2); -l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__3 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__3(); -lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__3); -l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__4 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__4(); -lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__4); -l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__1 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__1(); -lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__1); -l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__2 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__2(); -lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__2); -l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__3 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__3(); -lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__3); -l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__4 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__4(); -lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__4); -l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__5 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__5(); -lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__5); -l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__6 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__6(); -lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__6); +l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__1 = _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__1(); +lean_mark_persistent(l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__1); +l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__2 = _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__2(); +lean_mark_persistent(l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__2); +l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__3 = _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__3(); +lean_mark_persistent(l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__3); +l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__4 = _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__4(); +lean_mark_persistent(l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__4); +l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__5 = _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__5(); +lean_mark_persistent(l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__5); +l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__6 = _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__6(); +lean_mark_persistent(l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__6); +l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__7 = _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__7(); +lean_mark_persistent(l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__7); +l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__1 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__1(); +lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__1); +l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__2 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__2(); +lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__2); +l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__3 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__3(); +lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__3); +l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__4 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__4(); +lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__4); +l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__1 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__1(); +lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__1); +l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__2 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__2(); +lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__2); +l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__3 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__3(); +lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__3); +l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__4 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__4(); +lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__4); +l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__5 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__5(); +lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__5); +l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__6 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__6(); +lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__6); l_Lean_Meta_Grind_checkInvariants___lambda__1___closed__1 = _init_l_Lean_Meta_Grind_checkInvariants___lambda__1___closed__1(); lean_mark_persistent(l_Lean_Meta_Grind_checkInvariants___lambda__1___closed__1); l_Lean_Meta_Grind_checkInvariants___closed__1 = _init_l_Lean_Meta_Grind_checkInvariants___closed__1(); diff --git a/stage0/stdlib/Lean/Meta/Tactic/Grind/Main.c b/stage0/stdlib/Lean/Meta/Tactic/Grind/Main.c index 3b38da4bdebb..8077766a5a90 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Grind/Main.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Grind/Main.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Lean.Meta.Tactic.Grind.Main -// Imports: Init.Grind.Lemmas Lean.Meta.Tactic.Util Lean.Meta.Tactic.Grind.RevertAll Lean.Meta.Tactic.Grind.PropagatorAttr Lean.Meta.Tactic.Grind.Proj Lean.Meta.Tactic.Grind.ForallProp Lean.Meta.Tactic.Grind.Util Lean.Meta.Tactic.Grind.Inv Lean.Meta.Tactic.Grind.Intro Lean.Meta.Tactic.Grind.EMatch +// Imports: Init.Grind.Lemmas Lean.Meta.Tactic.Util Lean.Meta.Tactic.Grind.RevertAll Lean.Meta.Tactic.Grind.PropagatorAttr Lean.Meta.Tactic.Grind.Proj Lean.Meta.Tactic.Grind.ForallProp Lean.Meta.Tactic.Grind.Util Lean.Meta.Tactic.Grind.Inv Lean.Meta.Tactic.Grind.Intro Lean.Meta.Tactic.Grind.EMatch Lean.Meta.Tactic.Grind.Split Lean.Meta.Tactic.Grind.SimpUtil #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -13,34 +13,41 @@ #ifdef __cplusplus extern "C" { #endif +static lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore___closed__4; lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkMethods(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkMethods___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_filterMapM_loop___at_Lean_Meta_Grind_main___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_Meta_Grind_mkMethods(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_GrindM_run___rarg___closed__14; lean_object* l_Lean_addTrace___at_Lean_Meta_Grind_simp___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple(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_Meta_Grind_ematchStar(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_Grind_preprocessAndProbe___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___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_Meta_Grind_GrindM_run___rarg___closed__20; +LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Meta_Grind_main___spec__3(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_filterMapM_loop___at_Lean_Meta_Grind_main___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___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___closed__1; static lean_object* l_Lean_Meta_Grind_GrindM_run___rarg___closed__11; extern lean_object* l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedGoal___spec__2; static lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore___closed__2; +lean_object* l_Lean_Meta_Grind_splitNext(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_Grind_Main_0__Lean_Meta_Grind_initCore___closed__3; +lean_object* l_Lean_Meta_Grind_propagateForallPropDown(lean_object*, 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_Grind_mkMethods___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*); size_t lean_uint64_to_usize(uint64_t); +lean_object* l_Lean_Meta_Grind_getSimprocs___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_GrindM_run___rarg___closed__5; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___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_forM___at_Lean_Meta_Grind_preprocessAndProbe___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_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_mkMethods___spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_filterTR_loop___at___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore___spec__2(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_preprocessAndProbe(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_preprocessAndProbe___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_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___lambda__1___closed__4; +LEAN_EXPORT lean_object* l_Lean_MVarId_isAssigned___at_Lean_Meta_Grind_main___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_Grind_mkMethods___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore___closed__1; static lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___closed__3; -LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Meta_Grind_main___spec__1(lean_object*, lean_object*); lean_object* l_Lean_stringToMessageData(lean_object*); lean_object* l_Lean_MVarId_revertAll(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Meta_getSimpCongrTheorems___rarg(lean_object*, lean_object*); +lean_object* l_Lean_Meta_Grind_getSimpContext(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_main___lambda__2___closed__2; lean_object* l_Std_Queue_empty(lean_object*); lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); @@ -52,97 +59,100 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GrindM_run(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_GrindM_run___rarg___closed__10; lean_object* l_Lean_Meta_Grind_Goal_checkInvariants(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* lean_st_ref_take(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___lambda__2___closed__2; uint64_t lean_uint64_shift_right(uint64_t, uint64_t); -lean_object* l_Lean_Meta_Simp_SimprocExtension_getSimprocs(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_filterMapM_loop___at_Lean_Meta_Grind_main___spec__2___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkMethods___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___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_Meta_Grind_main___lambda__2___closed__3; static lean_object* l_Lean_Meta_Grind_main___lambda__2___closed__1; -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkMethods___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_filterMapM_loop___at_Lean_Meta_Grind_main___spec__2___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_propagateProjEq(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_ofFormat(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_main___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_main___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_getTrueExpr___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___closed__2; lean_object* lean_st_ref_get(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_GrindM_run___rarg___closed__13; -LEAN_EXPORT lean_object* l_Lean_Meta_withoutModifyingMCtx___at_Lean_Meta_Grind_preprocessAndProbe___spec__2(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___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_st_mk_ref(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_preprocessAndProbe___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_Grind_preprocessAndProbe___closed__1; +lean_object* l_Lean_Meta_appendTagSuffix(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forM___at___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore___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_Grind_main___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_List_filterMapM_loop___at_Lean_Meta_Grind_main___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_EXPORT lean_object* l_Lean_Meta_Grind_main___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_Grind_mkMethods___closed__1; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___lambda__3___closed__1; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___lambda__1___closed__2; static lean_object* l_Lean_Meta_Grind_GrindM_run___rarg___closed__15; lean_object* l_ShareCommon_mkStateImpl(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkMethods___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 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_Grind_GrindM_run___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Meta_Simp_mkContext(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GrindM_run___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Meta_Grind_builtinPropagatorsRef; -static lean_object* l_Lean_Meta_Grind_GrindM_run___rarg___closed__18; +lean_object* l_Lean_Meta_Grind_GrindTactic_iterate(lean_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_Grind_main___lambda__2___closed__5; lean_object* l_Lean_MVarId_clearAuxDecls(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___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_Grind_GrindM_run___rarg___closed__7; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_GrindM_run___rarg___closed__12; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___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_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkMethods___rarg(lean_object*); -extern lean_object* l_Lean_Meta_Grind_grindNormSimprocExt; -static lean_object* l_Lean_Meta_Grind_main___closed__1; -lean_object* l_Lean_Meta_SimpExtension_getTheorems(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_filterMapM_loop___at_Lean_Meta_Grind_main___spec__2___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_PersistentHashMap_contains___at_Lean_MVarId_isAssigned___spec__1(lean_object*, lean_object*); +lean_object* l_Lean_Meta_Grind_GrindTactic_andThen(lean_object*, lean_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_Grind_GrindM_run___rarg___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_main___lambda__2___closed__4; -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_main(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_main(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_GrindM_run___rarg___closed__6; +LEAN_EXPORT lean_object* l_List_filterMapM_loop___at_Lean_Meta_Grind_main___spec__2___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___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_MVarId_ensureNoMVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_forM___at_Lean_Meta_Grind_preprocessAndProbe___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_Grind_mkMethods___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkMethods___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_intros(lean_object*, 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_Grind_main___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_Grind_main___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_PersistentHashMap_empty___at___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___spec__1; -static lean_object* l_Lean_Meta_Grind_GrindM_run___rarg___closed__19; -static lean_object* l_Lean_Meta_Grind_main___lambda__2___closed__6; +extern lean_object* l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedGoal___spec__3; +LEAN_EXPORT lean_object* l_List_filterMapM_loop___at_Lean_Meta_Grind_main___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*); uint64_t l_Lean_Name_hash___override(lean_object*); -extern lean_object* l_Lean_Meta_Grind_grindNormExt; uint64_t lean_uint64_xor(uint64_t, uint64_t); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_mkMethods___spec__1(lean_object*, lean_object*); lean_object* l_Lean_MVarId_transformTarget(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Meta_resetCache___rarg(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_getAppFn(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___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_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___lambda__2___closed__1; lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_simp___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_Grind_GrindM_run___rarg___closed__17; lean_object* l_Lean_Meta_Grind_mkENodeCore(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* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_withoutModifyingMCtx___at_Lean_Meta_Grind_preprocessAndProbe___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___lambda__1___closed__3; lean_object* l_List_reverse___rarg(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___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_Grind_all(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_sub(size_t, size_t); +lean_object* l_Lean_Meta_Grind_assertAll(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_mk(lean_object*); static lean_object* l_Lean_Meta_Grind_GrindM_run___rarg___closed__4; LEAN_EXPORT lean_object* l_List_foldlM___at_Lean_Meta_Grind_all___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* l_Lean_MVarId_ensureProp(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_Meta_Grind_GrindM_run___rarg___closed__8; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___lambda__1___closed__1; lean_object* lean_array_uget(lean_object*, size_t); -lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_unfoldReducible(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_GrindM_run___rarg___closed__9; -lean_object* l_Lean_Meta_Grind_propagateForallProp(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_Grind_propagateForallPropUp(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_size(lean_object*); extern lean_object* l_Lean_ShareCommon_objectFactory; lean_object* lean_state_sharecommon(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_mkMethods___rarg___closed__1; static lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___closed__1; static lean_object* l_Lean_Meta_Grind_GrindM_run___rarg___closed__16; static lean_object* l_Lean_Meta_Grind_GrindM_run___rarg___closed__2; +lean_object* l_Lean_Meta_Grind_applyToAll(lean_object*, 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_Meta_Grind_main___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_Meta_Grind_ppGoals(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MVarId_betaReduce___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkMethods___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*); +extern lean_object* l_Lean_PersistentHashMap_empty___at_Lean_Meta_Match_instInhabitedMatchEqnsExtState___spec__1; size_t lean_usize_land(size_t, size_t); lean_object* l_Lean_Meta_Grind_getEMatchTheorems___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_mkMethods___spec__1(lean_object* x_1, lean_object* x_2) { @@ -177,7 +187,7 @@ return x_9; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkMethods___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_Grind_mkMethods___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; @@ -190,7 +200,7 @@ lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); -x_12 = l_Lean_Meta_Grind_propagateForallProp(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_12 = l_Lean_Meta_Grind_propagateForallPropUp(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); if (lean_obj_tag(x_12) == 0) { uint8_t x_13; @@ -560,49 +570,60 @@ return x_105; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkMethods___rarg___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkMethods___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { lean_object* x_12; -x_12 = l_Lean_Expr_getAppFn(x_2); -if (lean_obj_tag(x_12) == 4) +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +x_12 = l_Lean_Meta_Grind_propagateForallPropDown(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_12) == 0) { -lean_object* x_13; lean_object* x_14; uint8_t x_15; -x_13 = lean_ctor_get(x_1, 1); -lean_inc(x_13); -lean_dec(x_1); -x_14 = lean_ctor_get(x_12, 0); -lean_inc(x_14); -lean_dec(x_12); -x_15 = !lean_is_exclusive(x_13); -if (x_15 == 0) +uint8_t x_13; +x_13 = !lean_is_exclusive(x_12); +if (x_13 == 0) { -lean_object* x_16; lean_object* x_17; lean_object* x_18; uint64_t x_19; uint64_t x_20; uint64_t x_21; uint64_t x_22; uint64_t x_23; uint64_t x_24; uint64_t x_25; size_t x_26; size_t x_27; size_t x_28; size_t x_29; size_t x_30; lean_object* x_31; lean_object* x_32; -x_16 = lean_ctor_get(x_13, 1); -x_17 = lean_ctor_get(x_13, 0); -lean_dec(x_17); -x_18 = lean_array_get_size(x_16); -x_19 = l_Lean_Name_hash___override(x_14); -x_20 = 32; -x_21 = lean_uint64_shift_right(x_19, x_20); -x_22 = lean_uint64_xor(x_19, x_21); -x_23 = 16; -x_24 = lean_uint64_shift_right(x_22, x_23); -x_25 = lean_uint64_xor(x_22, x_24); -x_26 = lean_uint64_to_usize(x_25); -x_27 = lean_usize_of_nat(x_18); -lean_dec(x_18); -x_28 = 1; -x_29 = lean_usize_sub(x_27, x_28); -x_30 = lean_usize_land(x_26, x_29); -x_31 = lean_array_uget(x_16, x_30); +lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_14 = lean_ctor_get(x_12, 1); +x_15 = lean_ctor_get(x_12, 0); +lean_dec(x_15); +x_16 = l_Lean_Expr_getAppFn(x_2); +if (lean_obj_tag(x_16) == 4) +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint64_t x_21; uint64_t x_22; uint64_t x_23; uint64_t x_24; uint64_t x_25; uint64_t x_26; uint64_t x_27; size_t x_28; size_t x_29; size_t x_30; size_t x_31; size_t x_32; lean_object* x_33; lean_object* x_34; +x_17 = lean_ctor_get(x_1, 1); +x_18 = lean_ctor_get(x_16, 0); +lean_inc(x_18); lean_dec(x_16); -x_32 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_mkMethods___spec__1(x_14, x_31); -lean_dec(x_31); -lean_dec(x_14); -if (lean_obj_tag(x_32) == 0) +x_19 = lean_ctor_get(x_17, 1); +x_20 = lean_array_get_size(x_19); +x_21 = l_Lean_Name_hash___override(x_18); +x_22 = 32; +x_23 = lean_uint64_shift_right(x_21, x_22); +x_24 = lean_uint64_xor(x_21, x_23); +x_25 = 16; +x_26 = lean_uint64_shift_right(x_24, x_25); +x_27 = lean_uint64_xor(x_24, x_26); +x_28 = lean_uint64_to_usize(x_27); +x_29 = lean_usize_of_nat(x_20); +lean_dec(x_20); +x_30 = 1; +x_31 = lean_usize_sub(x_29, x_30); +x_32 = lean_usize_land(x_28, x_31); +x_33 = lean_array_uget(x_19, x_32); +x_34 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_mkMethods___spec__1(x_18, x_33); +lean_dec(x_33); +lean_dec(x_18); +if (lean_obj_tag(x_34) == 0) { -lean_object* x_33; +lean_object* x_35; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -612,50 +633,75 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_33 = lean_box(0); -lean_ctor_set(x_13, 1, x_11); -lean_ctor_set(x_13, 0, x_33); -return x_13; +x_35 = lean_box(0); +lean_ctor_set(x_12, 0, x_35); +return x_12; } else { -lean_object* x_34; lean_object* x_35; -lean_free_object(x_13); -x_34 = lean_ctor_get(x_32, 0); -lean_inc(x_34); -lean_dec(x_32); -x_35 = lean_apply_10(x_34, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_35; +lean_object* x_36; lean_object* x_37; +lean_free_object(x_12); +x_36 = lean_ctor_get(x_34, 0); +lean_inc(x_36); +lean_dec(x_34); +x_37 = lean_apply_10(x_36, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_14); +return x_37; } } else { -lean_object* x_36; lean_object* x_37; uint64_t x_38; uint64_t x_39; uint64_t x_40; uint64_t x_41; uint64_t x_42; uint64_t x_43; uint64_t x_44; size_t x_45; size_t x_46; size_t x_47; size_t x_48; size_t x_49; lean_object* x_50; lean_object* x_51; -x_36 = lean_ctor_get(x_13, 1); -lean_inc(x_36); -lean_dec(x_13); -x_37 = lean_array_get_size(x_36); -x_38 = l_Lean_Name_hash___override(x_14); -x_39 = 32; -x_40 = lean_uint64_shift_right(x_38, x_39); -x_41 = lean_uint64_xor(x_38, x_40); -x_42 = 16; -x_43 = lean_uint64_shift_right(x_41, x_42); -x_44 = lean_uint64_xor(x_41, x_43); -x_45 = lean_uint64_to_usize(x_44); -x_46 = lean_usize_of_nat(x_37); -lean_dec(x_37); -x_47 = 1; -x_48 = lean_usize_sub(x_46, x_47); -x_49 = lean_usize_land(x_45, x_48); -x_50 = lean_array_uget(x_36, x_49); -lean_dec(x_36); -x_51 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_mkMethods___spec__1(x_14, x_50); -lean_dec(x_50); -lean_dec(x_14); -if (lean_obj_tag(x_51) == 0) +lean_object* x_38; +lean_dec(x_16); +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_38 = lean_box(0); +lean_ctor_set(x_12, 0, x_38); +return x_12; +} +} +else { -lean_object* x_52; lean_object* x_53; +lean_object* x_39; lean_object* x_40; +x_39 = lean_ctor_get(x_12, 1); +lean_inc(x_39); +lean_dec(x_12); +x_40 = l_Lean_Expr_getAppFn(x_2); +if (lean_obj_tag(x_40) == 4) +{ +lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; uint64_t x_45; uint64_t x_46; uint64_t x_47; uint64_t x_48; uint64_t x_49; uint64_t x_50; uint64_t x_51; size_t x_52; size_t x_53; size_t x_54; size_t x_55; size_t x_56; lean_object* x_57; lean_object* x_58; +x_41 = lean_ctor_get(x_1, 1); +x_42 = lean_ctor_get(x_40, 0); +lean_inc(x_42); +lean_dec(x_40); +x_43 = lean_ctor_get(x_41, 1); +x_44 = lean_array_get_size(x_43); +x_45 = l_Lean_Name_hash___override(x_42); +x_46 = 32; +x_47 = lean_uint64_shift_right(x_45, x_46); +x_48 = lean_uint64_xor(x_45, x_47); +x_49 = 16; +x_50 = lean_uint64_shift_right(x_48, x_49); +x_51 = lean_uint64_xor(x_48, x_50); +x_52 = lean_uint64_to_usize(x_51); +x_53 = lean_usize_of_nat(x_44); +lean_dec(x_44); +x_54 = 1; +x_55 = lean_usize_sub(x_53, x_54); +x_56 = lean_usize_land(x_52, x_55); +x_57 = lean_array_uget(x_43, x_56); +x_58 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_mkMethods___spec__1(x_42, x_57); +lean_dec(x_57); +lean_dec(x_42); +if (lean_obj_tag(x_58) == 0) +{ +lean_object* x_59; lean_object* x_60; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -665,27 +711,46 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_52 = lean_box(0); -x_53 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_53, 0, x_52); -lean_ctor_set(x_53, 1, x_11); -return x_53; +x_59 = lean_box(0); +x_60 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_60, 0, x_59); +lean_ctor_set(x_60, 1, x_39); +return x_60; } else { -lean_object* x_54; lean_object* x_55; -x_54 = lean_ctor_get(x_51, 0); -lean_inc(x_54); -lean_dec(x_51); -x_55 = lean_apply_10(x_54, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_55; +lean_object* x_61; lean_object* x_62; +x_61 = lean_ctor_get(x_58, 0); +lean_inc(x_61); +lean_dec(x_58); +x_62 = lean_apply_10(x_61, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_39); +return x_62; +} +} +else +{ +lean_object* x_63; lean_object* x_64; +lean_dec(x_40); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_63 = lean_box(0); +x_64 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_64, 0, x_63); +lean_ctor_set(x_64, 1, x_39); +return x_64; } } } else { -lean_object* x_56; lean_object* x_57; -lean_dec(x_12); +uint8_t x_65; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -695,16 +760,28 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -x_56 = lean_box(0); -x_57 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_57, 0, x_56); -lean_ctor_set(x_57, 1, x_11); -return x_57; +x_65 = !lean_is_exclusive(x_12); +if (x_65 == 0) +{ +return x_12; +} +else +{ +lean_object* x_66; lean_object* x_67; lean_object* x_68; +x_66 = lean_ctor_get(x_12, 0); +x_67 = lean_ctor_get(x_12, 1); +lean_inc(x_67); +lean_inc(x_66); +lean_dec(x_12); +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; +} } } } -static lean_object* _init_l_Lean_Meta_Grind_mkMethods___rarg___closed__1() { +static lean_object* _init_l_Lean_Meta_Grind_mkMethods___closed__1() { _start: { lean_object* x_1; @@ -712,86 +789,89 @@ x_1 = l_Lean_Meta_Grind_builtinPropagatorsRef; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkMethods___rarg(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkMethods(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -lean_object* x_2; lean_object* x_3; uint8_t x_4; -x_2 = l_Lean_Meta_Grind_mkMethods___rarg___closed__1; -x_3 = lean_st_ref_get(x_2, x_1); -x_4 = !lean_is_exclusive(x_3); -if (x_4 == 0) +lean_object* x_5; lean_object* x_6; uint8_t x_7; +x_5 = l_Lean_Meta_Grind_mkMethods___closed__1; +x_6 = lean_st_ref_get(x_5, x_4); +x_7 = !lean_is_exclusive(x_6); +if (x_7 == 0) { -lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; -x_5 = lean_ctor_get(x_3, 0); -lean_inc(x_5); -x_6 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_mkMethods___rarg___lambda__1___boxed), 11, 1); -lean_closure_set(x_6, 0, x_5); -x_7 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_mkMethods___rarg___lambda__2), 11, 1); -lean_closure_set(x_7, 0, x_5); -x_8 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_8, 0, x_6); -lean_ctor_set(x_8, 1, x_7); -lean_ctor_set(x_3, 0, x_8); -return x_3; +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_8 = lean_ctor_get(x_6, 0); +lean_inc(x_8); +x_9 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_mkMethods___lambda__1___boxed), 11, 1); +lean_closure_set(x_9, 0, x_8); +x_10 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_mkMethods___lambda__2___boxed), 11, 1); +lean_closure_set(x_10, 0, x_8); +x_11 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_11, 0, x_9); +lean_ctor_set(x_11, 1, x_10); +lean_ctor_set(x_11, 2, x_1); +lean_ctor_set(x_6, 0, x_11); +return x_6; } 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 = lean_ctor_get(x_3, 0); -x_10 = lean_ctor_get(x_3, 1); -lean_inc(x_10); -lean_inc(x_9); -lean_dec(x_3); -lean_inc(x_9); -x_11 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_mkMethods___rarg___lambda__1___boxed), 11, 1); -lean_closure_set(x_11, 0, x_9); -x_12 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_mkMethods___rarg___lambda__2), 11, 1); -lean_closure_set(x_12, 0, x_9); -x_13 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_13, 0, x_11); -lean_ctor_set(x_13, 1, x_12); -x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_13); -lean_ctor_set(x_14, 1, x_10); -return x_14; +lean_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_6, 0); +x_13 = lean_ctor_get(x_6, 1); +lean_inc(x_13); +lean_inc(x_12); +lean_dec(x_6); +lean_inc(x_12); +x_14 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_mkMethods___lambda__1___boxed), 11, 1); +lean_closure_set(x_14, 0, x_12); +x_15 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_mkMethods___lambda__2___boxed), 11, 1); +lean_closure_set(x_15, 0, x_12); +x_16 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_16, 0, x_14); +lean_ctor_set(x_16, 1, x_15); +lean_ctor_set(x_16, 2, x_1); +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_16); +lean_ctor_set(x_17, 1, x_13); +return x_17; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkMethods(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_mkMethods___spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_mkMethods___rarg), 1, 0); +x_3 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_mkMethods___spec__1(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_mkMethods___spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkMethods___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_3; -x_3 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_mkMethods___spec__1(x_1, x_2); -lean_dec(x_2); +lean_object* x_12; +x_12 = l_Lean_Meta_Grind_mkMethods___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_1); -return x_3; +return x_12; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkMethods___rarg___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkMethods___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_Meta_Grind_mkMethods___rarg___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_12 = l_Lean_Meta_Grind_mkMethods___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_1); return x_12; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkMethods___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkMethods___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_Meta_Grind_mkMethods(x_1, x_2); +lean_object* x_5; +x_5 = l_Lean_Meta_Grind_mkMethods(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; } } static lean_object* _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__1() { @@ -874,75 +954,25 @@ static lean_object* _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__9() { _start: { lean_object* x_1; -x_1 = l_Lean_Meta_Grind_grindNormExt; -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__10() { -_start: -{ -lean_object* x_1; -x_1 = l_Lean_Meta_Grind_grindNormSimprocExt; -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__11() { -_start: -{ -lean_object* x_1; lean_object* x_2; uint8_t x_3; uint8_t x_4; uint8_t x_5; lean_object* x_6; -x_1 = l_Lean_Meta_Simp_defaultMaxSteps; -x_2 = lean_unsigned_to_nat(2u); -x_3 = 0; -x_4 = 1; -x_5 = 0; -x_6 = lean_alloc_ctor(0, 2, 19); -lean_ctor_set(x_6, 0, x_1); -lean_ctor_set(x_6, 1, x_2); -lean_ctor_set_uint8(x_6, sizeof(void*)*2, x_3); -lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 1, x_4); -lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 2, x_3); -lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 3, x_4); -lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 4, x_4); -lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 5, x_4); -lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 6, x_5); -lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 7, x_4); -lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 8, x_4); -lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 9, x_3); -lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 10, x_4); -lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 11, x_3); -lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 12, x_4); -lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 13, x_4); -lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 14, x_3); -lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 15, x_3); -lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 16, x_3); -lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 17, x_4); -lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 18, x_4); -return x_6; -} -} -static lean_object* _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__12() { -_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_Grind_GrindM_run___rarg___closed__13() { +static lean_object* _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__10() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__12; +x_1 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__9; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__14() { +static lean_object* _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__11() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__13; +x_1 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__10; x_2 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_2, 0, x_1); lean_ctor_set(x_2, 1, x_1); @@ -950,11 +980,11 @@ lean_ctor_set(x_2, 2, x_1); return x_2; } } -static lean_object* _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__15() { +static lean_object* _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__13; +x_1 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__10; x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -962,7 +992,7 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__16() { +static lean_object* _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__13() { _start: { lean_object* x_1; lean_object* x_2; @@ -971,23 +1001,23 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__17() { +static lean_object* _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__14() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__16; +x_1 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__13; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__18() { +static lean_object* _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__15() { _start: { size_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = 5; -x_2 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__17; -x_3 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__16; +x_2 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__14; +x_3 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__13; 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); @@ -998,12 +1028,12 @@ lean_ctor_set_usize(x_5, 4, x_1); return x_5; } } -static lean_object* _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__19() { +static lean_object* _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__13; -x_2 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__18; +x_1 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__10; +x_2 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__15; x_3 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_1); @@ -1012,23 +1042,22 @@ lean_ctor_set(x_3, 3, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__20() { +static lean_object* _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__15; -x_2 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__19; +x_1 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__12; +x_2 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__16; 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_Grind_GrindM_run___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GrindM_run___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; 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_9 = lean_box(0); +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; x_10 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__5; x_11 = lean_ctor_get(x_10, 0); lean_inc(x_11); @@ -1042,358 +1071,140 @@ lean_inc(x_16); x_17 = lean_ctor_get(x_15, 1); lean_inc(x_17); lean_dec(x_15); -x_18 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__9; -x_19 = l_Lean_Meta_SimpExtension_getTheorems(x_18, x_6, x_7, x_8); -x_20 = lean_ctor_get(x_19, 0); -lean_inc(x_20); -x_21 = lean_ctor_get(x_19, 1); -lean_inc(x_21); -lean_dec(x_19); -x_22 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__10; -x_23 = l_Lean_Meta_Simp_SimprocExtension_getSimprocs(x_22, x_6, x_7, x_21); -x_24 = !lean_is_exclusive(x_23); -if (x_24 == 0) +x_18 = l_Lean_Meta_Grind_getSimprocs___rarg(x_7, x_8, x_9); +if (lean_obj_tag(x_18) == 0) { -lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; -x_25 = lean_ctor_get(x_23, 1); -lean_ctor_set_tag(x_23, 1); -lean_ctor_set(x_23, 1, x_9); -x_26 = lean_array_mk(x_23); -x_27 = l_Lean_Meta_getSimpCongrTheorems___rarg(x_7, x_25); -x_28 = !lean_is_exclusive(x_27); -if (x_28 == 0) -{ -lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; 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_29 = lean_ctor_get(x_27, 0); -x_30 = lean_ctor_get(x_27, 1); -lean_ctor_set_tag(x_27, 1); -lean_ctor_set(x_27, 1, x_9); -lean_ctor_set(x_27, 0, x_20); -x_31 = lean_array_mk(x_27); -x_32 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__11; -x_33 = l_Lean_Meta_Simp_mkContext(x_32, x_31, x_29, x_4, x_5, x_6, x_7, x_30); -x_34 = lean_ctor_get(x_33, 0); -lean_inc(x_34); -x_35 = lean_ctor_get(x_33, 1); +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_18, 1); +lean_inc(x_20); +lean_dec(x_18); +x_21 = l_Lean_Meta_Grind_getSimpContext(x_5, x_6, x_7, x_8, 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_Lean_Meta_Grind_mkMethods(x_4, x_7, x_8, x_23); +x_25 = lean_ctor_get(x_24, 0); +lean_inc(x_25); +x_26 = lean_ctor_get(x_24, 1); +lean_inc(x_26); +lean_dec(x_24); +x_27 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_27, 0, x_22); +lean_ctor_set(x_27, 1, x_19); +lean_ctor_set(x_27, 2, x_2); +lean_ctor_set(x_27, 3, x_3); +x_28 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__11; +x_29 = lean_unsigned_to_nat(1u); +x_30 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__10; +x_31 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__17; +x_32 = lean_box(0); +x_33 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_33, 0, x_28); +lean_ctor_set(x_33, 1, x_17); +lean_ctor_set(x_33, 2, x_29); +lean_ctor_set(x_33, 3, x_30); +lean_ctor_set(x_33, 4, x_31); +lean_ctor_set(x_33, 5, x_16); +lean_ctor_set(x_33, 6, x_11); +lean_ctor_set(x_33, 7, x_32); +x_34 = lean_st_mk_ref(x_33, x_26); +x_35 = lean_ctor_get(x_34, 0); lean_inc(x_35); -lean_dec(x_33); -x_36 = l_Lean_Meta_Grind_mkMethods___rarg(x_35); -x_37 = lean_ctor_get(x_36, 0); -lean_inc(x_37); -x_38 = lean_ctor_get(x_36, 1); -lean_inc(x_38); -lean_dec(x_36); -x_39 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_39, 0, x_34); -lean_ctor_set(x_39, 1, x_26); -lean_ctor_set(x_39, 2, x_2); -lean_ctor_set(x_39, 3, x_3); -x_40 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__14; -x_41 = lean_unsigned_to_nat(1u); -x_42 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__13; -x_43 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__20; -x_44 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_44, 0, x_40); -lean_ctor_set(x_44, 1, x_17); -lean_ctor_set(x_44, 2, x_41); -lean_ctor_set(x_44, 3, x_42); -lean_ctor_set(x_44, 4, x_43); -lean_ctor_set(x_44, 5, x_16); -lean_ctor_set(x_44, 6, x_11); -x_45 = lean_st_mk_ref(x_44, x_38); -x_46 = lean_ctor_get(x_45, 0); -lean_inc(x_46); -x_47 = lean_ctor_get(x_45, 1); -lean_inc(x_47); -lean_dec(x_45); -lean_inc(x_46); -x_48 = lean_apply_8(x_1, x_37, x_39, x_46, x_4, x_5, x_6, x_7, x_47); -if (lean_obj_tag(x_48) == 0) +x_36 = lean_ctor_get(x_34, 1); +lean_inc(x_36); +lean_dec(x_34); +lean_inc(x_35); +x_37 = lean_apply_8(x_1, x_25, x_27, x_35, x_5, x_6, x_7, x_8, x_36); +if (lean_obj_tag(x_37) == 0) { -lean_object* x_49; lean_object* x_50; lean_object* x_51; uint8_t x_52; -x_49 = lean_ctor_get(x_48, 0); -lean_inc(x_49); -x_50 = lean_ctor_get(x_48, 1); -lean_inc(x_50); -lean_dec(x_48); -x_51 = lean_st_ref_get(x_46, x_50); -lean_dec(x_46); -x_52 = !lean_is_exclusive(x_51); -if (x_52 == 0) +lean_object* x_38; lean_object* x_39; lean_object* x_40; uint8_t x_41; +x_38 = lean_ctor_get(x_37, 0); +lean_inc(x_38); +x_39 = lean_ctor_get(x_37, 1); +lean_inc(x_39); +lean_dec(x_37); +x_40 = lean_st_ref_get(x_35, x_39); +lean_dec(x_35); +x_41 = !lean_is_exclusive(x_40); +if (x_41 == 0) { -lean_object* x_53; -x_53 = lean_ctor_get(x_51, 0); -lean_dec(x_53); -lean_ctor_set(x_51, 0, x_49); -return x_51; +lean_object* x_42; +x_42 = lean_ctor_get(x_40, 0); +lean_dec(x_42); +lean_ctor_set(x_40, 0, x_38); +return x_40; } else { -lean_object* x_54; lean_object* x_55; -x_54 = lean_ctor_get(x_51, 1); -lean_inc(x_54); -lean_dec(x_51); -x_55 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_55, 0, x_49); -lean_ctor_set(x_55, 1, x_54); -return x_55; +lean_object* x_43; lean_object* x_44; +x_43 = lean_ctor_get(x_40, 1); +lean_inc(x_43); +lean_dec(x_40); +x_44 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_44, 0, x_38); +lean_ctor_set(x_44, 1, x_43); +return x_44; } } else { -uint8_t x_56; -lean_dec(x_46); -x_56 = !lean_is_exclusive(x_48); -if (x_56 == 0) +uint8_t x_45; +lean_dec(x_35); +x_45 = !lean_is_exclusive(x_37); +if (x_45 == 0) { -return x_48; +return x_37; } else { -lean_object* x_57; lean_object* x_58; lean_object* x_59; -x_57 = lean_ctor_get(x_48, 0); -x_58 = lean_ctor_get(x_48, 1); -lean_inc(x_58); -lean_inc(x_57); -lean_dec(x_48); -x_59 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_59, 0, x_57); -lean_ctor_set(x_59, 1, x_58); -return x_59; +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; } } } else { -lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; -x_60 = lean_ctor_get(x_27, 0); -x_61 = lean_ctor_get(x_27, 1); -lean_inc(x_61); -lean_inc(x_60); -lean_dec(x_27); -x_62 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_62, 0, x_20); -lean_ctor_set(x_62, 1, x_9); -x_63 = lean_array_mk(x_62); -x_64 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__11; -x_65 = l_Lean_Meta_Simp_mkContext(x_64, x_63, x_60, x_4, x_5, x_6, x_7, x_61); -x_66 = lean_ctor_get(x_65, 0); -lean_inc(x_66); -x_67 = lean_ctor_get(x_65, 1); -lean_inc(x_67); -lean_dec(x_65); -x_68 = l_Lean_Meta_Grind_mkMethods___rarg(x_67); -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_alloc_ctor(0, 4, 0); -lean_ctor_set(x_71, 0, x_66); -lean_ctor_set(x_71, 1, x_26); -lean_ctor_set(x_71, 2, x_2); -lean_ctor_set(x_71, 3, x_3); -x_72 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__14; -x_73 = lean_unsigned_to_nat(1u); -x_74 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__13; -x_75 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__20; -x_76 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_76, 0, x_72); -lean_ctor_set(x_76, 1, x_17); -lean_ctor_set(x_76, 2, x_73); -lean_ctor_set(x_76, 3, x_74); -lean_ctor_set(x_76, 4, x_75); -lean_ctor_set(x_76, 5, x_16); -lean_ctor_set(x_76, 6, x_11); -x_77 = lean_st_mk_ref(x_76, x_70); -x_78 = lean_ctor_get(x_77, 0); -lean_inc(x_78); -x_79 = lean_ctor_get(x_77, 1); -lean_inc(x_79); -lean_dec(x_77); -lean_inc(x_78); -x_80 = lean_apply_8(x_1, x_69, x_71, x_78, x_4, x_5, x_6, x_7, x_79); -if (lean_obj_tag(x_80) == 0) +uint8_t x_49; +lean_dec(x_17); +lean_dec(x_16); +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); +lean_dec(x_1); +x_49 = !lean_is_exclusive(x_18); +if (x_49 == 0) { -lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; -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_st_ref_get(x_78, x_82); -lean_dec(x_78); -x_84 = lean_ctor_get(x_83, 1); -lean_inc(x_84); -if (lean_is_exclusive(x_83)) { - lean_ctor_release(x_83, 0); - lean_ctor_release(x_83, 1); - x_85 = x_83; -} else { - lean_dec_ref(x_83); - x_85 = lean_box(0); -} -if (lean_is_scalar(x_85)) { - x_86 = lean_alloc_ctor(0, 2, 0); -} else { - x_86 = x_85; -} -lean_ctor_set(x_86, 0, x_81); -lean_ctor_set(x_86, 1, x_84); -return x_86; +return x_18; } else { -lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; -lean_dec(x_78); -x_87 = lean_ctor_get(x_80, 0); -lean_inc(x_87); -x_88 = lean_ctor_get(x_80, 1); -lean_inc(x_88); -if (lean_is_exclusive(x_80)) { - lean_ctor_release(x_80, 0); - lean_ctor_release(x_80, 1); - x_89 = x_80; -} else { - lean_dec_ref(x_80); - x_89 = lean_box(0); -} -if (lean_is_scalar(x_89)) { - x_90 = lean_alloc_ctor(1, 2, 0); -} else { - x_90 = x_89; -} -lean_ctor_set(x_90, 0, x_87); -lean_ctor_set(x_90, 1, x_88); -return x_90; -} -} -} -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; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; -x_91 = lean_ctor_get(x_23, 0); -x_92 = lean_ctor_get(x_23, 1); -lean_inc(x_92); -lean_inc(x_91); -lean_dec(x_23); -x_93 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_93, 0, x_91); -lean_ctor_set(x_93, 1, x_9); -x_94 = lean_array_mk(x_93); -x_95 = l_Lean_Meta_getSimpCongrTheorems___rarg(x_7, x_92); -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_95); - x_98 = lean_box(0); -} -if (lean_is_scalar(x_98)) { - x_99 = lean_alloc_ctor(1, 2, 0); -} else { - x_99 = x_98; - lean_ctor_set_tag(x_99, 1); -} -lean_ctor_set(x_99, 0, x_20); -lean_ctor_set(x_99, 1, x_9); -x_100 = lean_array_mk(x_99); -x_101 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__11; -x_102 = l_Lean_Meta_Simp_mkContext(x_101, x_100, x_96, x_4, x_5, x_6, x_7, x_97); -x_103 = lean_ctor_get(x_102, 0); -lean_inc(x_103); -x_104 = lean_ctor_get(x_102, 1); -lean_inc(x_104); -lean_dec(x_102); -x_105 = l_Lean_Meta_Grind_mkMethods___rarg(x_104); -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(0, 4, 0); -lean_ctor_set(x_108, 0, x_103); -lean_ctor_set(x_108, 1, x_94); -lean_ctor_set(x_108, 2, x_2); -lean_ctor_set(x_108, 3, x_3); -x_109 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__14; -x_110 = lean_unsigned_to_nat(1u); -x_111 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__13; -x_112 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__20; -x_113 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_113, 0, x_109); -lean_ctor_set(x_113, 1, x_17); -lean_ctor_set(x_113, 2, x_110); -lean_ctor_set(x_113, 3, x_111); -lean_ctor_set(x_113, 4, x_112); -lean_ctor_set(x_113, 5, x_16); -lean_ctor_set(x_113, 6, x_11); -x_114 = lean_st_mk_ref(x_113, x_107); -x_115 = lean_ctor_get(x_114, 0); -lean_inc(x_115); -x_116 = lean_ctor_get(x_114, 1); -lean_inc(x_116); -lean_dec(x_114); -lean_inc(x_115); -x_117 = lean_apply_8(x_1, x_106, x_108, x_115, x_4, x_5, x_6, x_7, x_116); -if (lean_obj_tag(x_117) == 0) -{ -lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; -x_118 = lean_ctor_get(x_117, 0); -lean_inc(x_118); -x_119 = lean_ctor_get(x_117, 1); -lean_inc(x_119); -lean_dec(x_117); -x_120 = lean_st_ref_get(x_115, x_119); -lean_dec(x_115); -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); -} -if (lean_is_scalar(x_122)) { - x_123 = lean_alloc_ctor(0, 2, 0); -} else { - x_123 = x_122; -} -lean_ctor_set(x_123, 0, x_118); -lean_ctor_set(x_123, 1, x_121); -return x_123; -} -else -{ -lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; -lean_dec(x_115); -x_124 = lean_ctor_get(x_117, 0); -lean_inc(x_124); -x_125 = lean_ctor_get(x_117, 1); -lean_inc(x_125); -if (lean_is_exclusive(x_117)) { - lean_ctor_release(x_117, 0); - lean_ctor_release(x_117, 1); - x_126 = x_117; -} else { - lean_dec_ref(x_117); - x_126 = lean_box(0); -} -if (lean_is_scalar(x_126)) { - x_127 = lean_alloc_ctor(1, 2, 0); -} else { - x_127 = x_126; -} -lean_ctor_set(x_127, 0, x_124); -lean_ctor_set(x_127, 1, x_125); -return x_127; +lean_object* x_50; lean_object* x_51; lean_object* x_52; +x_50 = lean_ctor_get(x_18, 0); +x_51 = lean_ctor_get(x_18, 1); +lean_inc(x_51); +lean_inc(x_50); +lean_dec(x_18); +x_52 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_52, 0, x_50); +lean_ctor_set(x_52, 1, x_51); +return x_52; } } } @@ -1402,7 +1213,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GrindM_run(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_GrindM_run___rarg), 8, 0); +x_2 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_GrindM_run___rarg), 9, 0); return x_2; } } @@ -1410,7 +1221,7 @@ static lean_object* _init_l_Lean_PersistentHashMap_empty___at___private_Lean_Met _start: { lean_object* x_1; -x_1 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__13; +x_1 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__10; return x_1; } } @@ -1574,7 +1385,7 @@ return x_1; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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; uint8_t x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; x_10 = l_Lean_Meta_Grind_getTrueExpr___rarg(x_4, x_5, x_6, x_7, x_8, x_9); x_11 = lean_ctor_get(x_10, 0); lean_inc(x_11); @@ -1593,45 +1404,54 @@ lean_inc(x_17); x_18 = lean_ctor_get(x_16, 1); lean_inc(x_18); lean_dec(x_16); -x_19 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__13; -x_20 = l_Lean_PersistentHashMap_empty___at___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___spec__1; -x_21 = l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___closed__1; -x_22 = 0; -x_23 = lean_unsigned_to_nat(0u); -x_24 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__18; -x_25 = l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedGoal___spec__2; -x_26 = l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___closed__2; +x_19 = lean_box(0); +x_20 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__10; +x_21 = l_Lean_PersistentHashMap_empty___at___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___spec__1; +x_22 = l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___closed__1; +x_23 = 0; +x_24 = lean_unsigned_to_nat(0u); +x_25 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__15; +x_26 = l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedGoal___spec__2; +x_27 = l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___closed__2; +x_28 = l_Lean_PersistentHashMap_empty___at_Lean_Meta_Match_instInhabitedMatchEqnsExtState___spec__1; +x_29 = l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedGoal___spec__3; lean_inc(x_1); -x_27 = lean_alloc_ctor(0, 14, 1); -lean_ctor_set(x_27, 0, x_1); -lean_ctor_set(x_27, 1, x_19); -lean_ctor_set(x_27, 2, x_19); -lean_ctor_set(x_27, 3, x_20); -lean_ctor_set(x_27, 4, x_19); -lean_ctor_set(x_27, 5, x_21); -lean_ctor_set(x_27, 6, x_23); -lean_ctor_set(x_27, 7, x_23); -lean_ctor_set(x_27, 8, x_24); -lean_ctor_set(x_27, 9, x_24); -lean_ctor_set(x_27, 10, x_17); -lean_ctor_set(x_27, 11, x_23); -lean_ctor_set(x_27, 12, x_25); -lean_ctor_set(x_27, 13, x_26); -lean_ctor_set_uint8(x_27, sizeof(void*)*14, x_22); -x_28 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___lambda__1___boxed), 9, 1); -lean_closure_set(x_28, 0, x_27); -x_29 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___lambda__2___boxed), 11, 2); -lean_closure_set(x_29, 0, x_14); -lean_closure_set(x_29, 1, x_11); -x_30 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); -lean_closure_set(x_30, 0, x_28); -lean_closure_set(x_30, 1, x_29); -x_31 = l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___closed__3; -x_32 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); -lean_closure_set(x_32, 0, x_30); -lean_closure_set(x_32, 1, x_31); -x_33 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(x_1, x_32, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_18); -return x_33; +x_30 = lean_alloc_ctor(0, 20, 1); +lean_ctor_set(x_30, 0, x_1); +lean_ctor_set(x_30, 1, x_20); +lean_ctor_set(x_30, 2, x_20); +lean_ctor_set(x_30, 3, x_21); +lean_ctor_set(x_30, 4, x_20); +lean_ctor_set(x_30, 5, x_22); +lean_ctor_set(x_30, 6, x_24); +lean_ctor_set(x_30, 7, x_24); +lean_ctor_set(x_30, 8, x_25); +lean_ctor_set(x_30, 9, x_25); +lean_ctor_set(x_30, 10, x_17); +lean_ctor_set(x_30, 11, x_24); +lean_ctor_set(x_30, 12, x_24); +lean_ctor_set(x_30, 13, x_26); +lean_ctor_set(x_30, 14, x_27); +lean_ctor_set(x_30, 15, x_28); +lean_ctor_set(x_30, 16, x_19); +lean_ctor_set(x_30, 17, x_24); +lean_ctor_set(x_30, 18, x_29); +lean_ctor_set(x_30, 19, x_24); +lean_ctor_set_uint8(x_30, sizeof(void*)*20, x_23); +x_31 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___lambda__1___boxed), 9, 1); +lean_closure_set(x_31, 0, x_30); +x_32 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___lambda__2___boxed), 11, 2); +lean_closure_set(x_32, 0, x_14); +lean_closure_set(x_32, 1, x_11); +x_33 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); +lean_closure_set(x_33, 0, x_31); +lean_closure_set(x_33, 1, x_32); +x_34 = l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___closed__3; +x_35 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); +lean_closure_set(x_35, 0, x_33); +lean_closure_set(x_35, 1, x_34); +x_36 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(x_1, x_35, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_18); +return x_36; } } LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___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) { @@ -1773,7 +1593,7 @@ else lean_object* x_4; uint8_t x_5; x_4 = lean_ctor_get(x_1, 0); lean_inc(x_4); -x_5 = lean_ctor_get_uint8(x_4, sizeof(void*)*14); +x_5 = lean_ctor_get_uint8(x_4, sizeof(void*)*20); if (x_5 == 0) { uint8_t x_6; @@ -1836,6 +1656,24 @@ x_1 = lean_alloc_closure((void*)(l_Lean_MVarId_betaReduce___lambda__1___boxed), return x_1; } } +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("grind", 5, 5); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore___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_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore___closed__3; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { @@ -1908,12 +1746,21 @@ lean_inc(x_5); x_25 = l_Lean_MVarId_transformTarget(x_22, x_24, x_5, x_6, x_7, x_8, x_23); if (lean_obj_tag(x_25) == 0) { -lean_object* x_26; lean_object* x_27; lean_object* x_28; +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___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore___closed__4; +lean_inc(x_26); +x_29 = l_Lean_Meta_appendTagSuffix(x_26, x_28, x_5, x_6, x_7, x_8, x_27); +if (lean_obj_tag(x_29) == 0) +{ +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_8); lean_inc(x_7); lean_inc(x_6); @@ -1921,16 +1768,16 @@ lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); -x_28 = l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal(x_26, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_27); -if (lean_obj_tag(x_28) == 0) +x_31 = l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal(x_26, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_30); +if (lean_obj_tag(x_31) == 0) { -lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; -x_29 = lean_ctor_get(x_28, 0); -lean_inc(x_29); -x_30 = lean_ctor_get(x_28, 1); -lean_inc(x_30); -lean_dec(x_28); -x_31 = lean_unsigned_to_nat(0u); +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); +lean_dec(x_31); +x_34 = lean_unsigned_to_nat(0u); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); @@ -1938,72 +1785,102 @@ lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); -x_32 = l_Lean_Meta_Grind_intros(x_29, x_31, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_30); -if (lean_obj_tag(x_32) == 0) -{ -lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_33 = lean_ctor_get(x_32, 0); -lean_inc(x_33); -x_34 = lean_ctor_get(x_32, 1); -lean_inc(x_34); -lean_dec(x_32); -lean_inc(x_33); -x_35 = l_List_forM___at___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore___spec__1(x_33, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_34); +x_35 = l_Lean_Meta_Grind_intros(x_34, x_32, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_33); 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_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_35, 0); +lean_inc(x_36); +x_37 = lean_ctor_get(x_35, 1); +lean_inc(x_37); +lean_dec(x_35); +lean_inc(x_36); +x_38 = l_List_forM___at___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore___spec__1(x_36, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_37); +if (lean_obj_tag(x_38) == 0) { -lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_37 = lean_ctor_get(x_35, 0); -lean_dec(x_37); -x_38 = lean_box(0); -x_39 = l_List_filterTR_loop___at___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore___spec__2(x_33, x_38); -lean_ctor_set(x_35, 0, x_39); -return x_35; +uint8_t x_39; +x_39 = !lean_is_exclusive(x_38); +if (x_39 == 0) +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_40 = lean_ctor_get(x_38, 0); +lean_dec(x_40); +x_41 = lean_box(0); +x_42 = l_List_filterTR_loop___at___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore___spec__2(x_36, x_41); +lean_ctor_set(x_38, 0, x_42); +return x_38; } else { -lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; -x_40 = lean_ctor_get(x_35, 1); -lean_inc(x_40); -lean_dec(x_35); -x_41 = lean_box(0); -x_42 = l_List_filterTR_loop___at___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore___spec__2(x_33, x_41); -x_43 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_43, 0, x_42); -lean_ctor_set(x_43, 1, x_40); -return x_43; +lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_43 = lean_ctor_get(x_38, 1); +lean_inc(x_43); +lean_dec(x_38); +x_44 = lean_box(0); +x_45 = l_List_filterTR_loop___at___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore___spec__2(x_36, 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; } } else { -uint8_t x_44; -lean_dec(x_33); -x_44 = !lean_is_exclusive(x_35); -if (x_44 == 0) +uint8_t x_47; +lean_dec(x_36); +x_47 = !lean_is_exclusive(x_38); +if (x_47 == 0) +{ +return x_38; +} +else +{ +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_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; +} +} +} +else +{ +uint8_t x_51; +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_51 = !lean_is_exclusive(x_35); +if (x_51 == 0) { return x_35; } else { -lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_45 = lean_ctor_get(x_35, 0); -x_46 = lean_ctor_get(x_35, 1); -lean_inc(x_46); -lean_inc(x_45); +lean_object* x_52; lean_object* x_53; lean_object* x_54; +x_52 = lean_ctor_get(x_35, 0); +x_53 = lean_ctor_get(x_35, 1); +lean_inc(x_53); +lean_inc(x_52); lean_dec(x_35); -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_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_48; +uint8_t x_55; lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -2011,29 +1888,30 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_48 = !lean_is_exclusive(x_32); -if (x_48 == 0) +x_55 = !lean_is_exclusive(x_31); +if (x_55 == 0) { -return x_32; +return x_31; } else { -lean_object* x_49; lean_object* x_50; lean_object* x_51; -x_49 = lean_ctor_get(x_32, 0); -x_50 = lean_ctor_get(x_32, 1); -lean_inc(x_50); -lean_inc(x_49); -lean_dec(x_32); -x_51 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_51, 0, x_49); -lean_ctor_set(x_51, 1, x_50); -return x_51; +lean_object* x_56; lean_object* x_57; lean_object* x_58; +x_56 = lean_ctor_get(x_31, 0); +x_57 = lean_ctor_get(x_31, 1); +lean_inc(x_57); +lean_inc(x_56); +lean_dec(x_31); +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_52; +uint8_t x_59; +lean_dec(x_26); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -2041,29 +1919,29 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_52 = !lean_is_exclusive(x_28); -if (x_52 == 0) +x_59 = !lean_is_exclusive(x_29); +if (x_59 == 0) { -return x_28; +return x_29; } else { -lean_object* x_53; lean_object* x_54; lean_object* x_55; -x_53 = lean_ctor_get(x_28, 0); -x_54 = lean_ctor_get(x_28, 1); -lean_inc(x_54); -lean_inc(x_53); -lean_dec(x_28); -x_55 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_55, 0, x_53); -lean_ctor_set(x_55, 1, x_54); -return x_55; +lean_object* x_60; lean_object* x_61; lean_object* x_62; +x_60 = lean_ctor_get(x_29, 0); +x_61 = lean_ctor_get(x_29, 1); +lean_inc(x_61); +lean_inc(x_60); +lean_dec(x_29); +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_56; +uint8_t x_63; lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -2071,29 +1949,29 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_56 = !lean_is_exclusive(x_25); -if (x_56 == 0) +x_63 = !lean_is_exclusive(x_25); +if (x_63 == 0) { return x_25; } else { -lean_object* x_57; lean_object* x_58; lean_object* x_59; -x_57 = lean_ctor_get(x_25, 0); -x_58 = lean_ctor_get(x_25, 1); -lean_inc(x_58); -lean_inc(x_57); +lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_64 = lean_ctor_get(x_25, 0); +x_65 = lean_ctor_get(x_25, 1); +lean_inc(x_65); +lean_inc(x_64); lean_dec(x_25); -x_59 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_59, 0, x_57); -lean_ctor_set(x_59, 1, x_58); -return x_59; +x_66 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_66, 0, x_64); +lean_ctor_set(x_66, 1, x_65); +return x_66; } } } else { -uint8_t x_60; +uint8_t x_67; lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -2101,29 +1979,29 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_60 = !lean_is_exclusive(x_21); -if (x_60 == 0) +x_67 = !lean_is_exclusive(x_21); +if (x_67 == 0) { return x_21; } else { -lean_object* x_61; lean_object* x_62; lean_object* x_63; -x_61 = lean_ctor_get(x_21, 0); -x_62 = lean_ctor_get(x_21, 1); -lean_inc(x_62); -lean_inc(x_61); +lean_object* x_68; lean_object* x_69; lean_object* x_70; +x_68 = lean_ctor_get(x_21, 0); +x_69 = lean_ctor_get(x_21, 1); +lean_inc(x_69); +lean_inc(x_68); lean_dec(x_21); -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_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 { -uint8_t x_64; +uint8_t x_71; lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -2131,29 +2009,29 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_64 = !lean_is_exclusive(x_17); -if (x_64 == 0) +x_71 = !lean_is_exclusive(x_17); +if (x_71 == 0) { return x_17; } else { -lean_object* x_65; lean_object* x_66; lean_object* x_67; -x_65 = lean_ctor_get(x_17, 0); -x_66 = lean_ctor_get(x_17, 1); -lean_inc(x_66); -lean_inc(x_65); +lean_object* x_72; lean_object* x_73; lean_object* x_74; +x_72 = lean_ctor_get(x_17, 0); +x_73 = lean_ctor_get(x_17, 1); +lean_inc(x_73); +lean_inc(x_72); lean_dec(x_17); -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; +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_68; +uint8_t x_75; lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -2161,29 +2039,29 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_68 = !lean_is_exclusive(x_14); -if (x_68 == 0) +x_75 = !lean_is_exclusive(x_14); +if (x_75 == 0) { return x_14; } else { -lean_object* x_69; lean_object* x_70; lean_object* x_71; -x_69 = lean_ctor_get(x_14, 0); -x_70 = lean_ctor_get(x_14, 1); -lean_inc(x_70); -lean_inc(x_69); +lean_object* x_76; lean_object* x_77; lean_object* x_78; +x_76 = lean_ctor_get(x_14, 0); +x_77 = lean_ctor_get(x_14, 1); +lean_inc(x_77); +lean_inc(x_76); lean_dec(x_14); -x_71 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_71, 0, x_69); -lean_ctor_set(x_71, 1, x_70); -return x_71; +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 { -uint8_t x_72; +uint8_t x_79; lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -2192,29 +2070,29 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_72 = !lean_is_exclusive(x_12); -if (x_72 == 0) +x_79 = !lean_is_exclusive(x_12); +if (x_79 == 0) { return x_12; } else { -lean_object* x_73; lean_object* x_74; lean_object* x_75; -x_73 = lean_ctor_get(x_12, 0); -x_74 = lean_ctor_get(x_12, 1); -lean_inc(x_74); -lean_inc(x_73); +lean_object* x_80; lean_object* x_81; lean_object* x_82; +x_80 = lean_ctor_get(x_12, 0); +x_81 = lean_ctor_get(x_12, 1); +lean_inc(x_81); +lean_inc(x_80); lean_dec(x_12); -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; +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_76; +uint8_t x_83; lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -2223,23 +2101,23 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_76 = !lean_is_exclusive(x_10); -if (x_76 == 0) +x_83 = !lean_is_exclusive(x_10); +if (x_83 == 0) { return x_10; } else { -lean_object* x_77; lean_object* x_78; lean_object* x_79; -x_77 = lean_ctor_get(x_10, 0); -x_78 = lean_ctor_get(x_10, 1); -lean_inc(x_78); -lean_inc(x_77); +lean_object* x_84; lean_object* x_85; lean_object* x_86; +x_84 = lean_ctor_get(x_10, 0); +x_85 = lean_ctor_get(x_10, 1); +lean_inc(x_85); +lean_inc(x_84); lean_dec(x_10); -x_79 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_79, 0, x_77); -lean_ctor_set(x_79, 1, x_78); -return x_79; +x_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; } } } @@ -2338,420 +2216,160 @@ x_12 = l_List_foldlM___at_Lean_Meta_Grind_all___spec__1(x_2, x_11, x_1, x_3, x_4 return x_12; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___closed__1() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___lambda__1___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_ematchStar), 9, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_assertAll), 9, 0); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple(lean_object* x_1, lean_object* x_2, lean_object* x_3, 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_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___lambda__1___closed__2() { _start: { -lean_object* x_10; lean_object* x_11; -x_10 = l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___closed__1; -x_11 = l_Lean_Meta_Grind_all(x_1, x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -return x_11; +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_ematchStar), 9, 0); +return x_1; } } -LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Meta_Grind_main___spec__1(lean_object* x_1, lean_object* x_2) { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___lambda__1___closed__3() { _start: { -if (lean_obj_tag(x_1) == 0) -{ -lean_object* x_3; -x_3 = l_List_reverse___rarg(x_2); +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___lambda__1___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___lambda__1___closed__2; +x_3 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_GrindTactic_andThen), 11, 2); +lean_closure_set(x_3, 0, x_1); +lean_closure_set(x_3, 1, 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 = lean_ctor_get(x_5, 0); -lean_inc(x_7); -lean_dec(x_5); -lean_ctor_set(x_1, 1, x_2); -lean_ctor_set(x_1, 0, x_7); -{ -lean_object* _tmp_0 = x_6; -lean_object* _tmp_1 = x_1; -x_1 = _tmp_0; -x_2 = _tmp_1; -} -goto _start; } -else +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___lambda__1___closed__4() { +_start: { -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_9 = lean_ctor_get(x_1, 0); -x_10 = lean_ctor_get(x_1, 1); -lean_inc(x_10); -lean_inc(x_9); -lean_dec(x_1); -x_11 = lean_ctor_get(x_9, 0); -lean_inc(x_11); -lean_dec(x_9); -x_12 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_12, 0, x_11); -lean_ctor_set(x_12, 1, x_2); -x_1 = x_10; -x_2 = x_12; -goto _start; -} -} +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_splitNext), 9, 0); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_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, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___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_11; lean_object* x_12; lean_object* x_13; -x_11 = lean_box(0); -x_12 = l_List_mapTR_loop___at_Lean_Meta_Grind_main___spec__1(x_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_10); -return x_13; +lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_10 = l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___lambda__1___closed__4; +x_11 = l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___lambda__1___closed__3; +x_12 = l_Lean_Meta_Grind_GrindTactic_andThen(x_10, x_11, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_12; } } -static lean_object* _init_l_Lean_Meta_Grind_main___lambda__2___closed__1() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___lambda__2___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("grind", 5, 5); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___lambda__1), 9, 0); return x_1; } } -static lean_object* _init_l_Lean_Meta_Grind_main___lambda__2___closed__2() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___lambda__2___closed__2() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("debug", 5, 5); -return x_1; +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___lambda__2___closed__1; +x_2 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_GrindTactic_iterate), 10, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; } } -static lean_object* _init_l_Lean_Meta_Grind_main___lambda__2___closed__3() { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_10 = l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___lambda__1___closed__2; +x_11 = l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___lambda__2___closed__2; +x_12 = l_Lean_Meta_Grind_GrindTactic_andThen(x_10, x_11, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_12; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___lambda__3___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("final", 5, 5); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___lambda__2), 9, 0); return x_1; } } -static lean_object* _init_l_Lean_Meta_Grind_main___lambda__2___closed__4() { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___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_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Meta_Grind_main___lambda__2___closed__1; -x_2 = l_Lean_Meta_Grind_main___lambda__2___closed__2; -x_3 = l_Lean_Meta_Grind_main___lambda__2___closed__3; -x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); -return x_4; +lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_10 = l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___lambda__1___closed__1; +x_11 = l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___lambda__3___closed__1; +x_12 = l_Lean_Meta_Grind_GrindTactic_andThen(x_10, x_11, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_12; } } -static lean_object* _init_l_Lean_Meta_Grind_main___lambda__2___closed__5() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("", 0, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___lambda__3), 9, 0); return x_1; } } -static lean_object* _init_l_Lean_Meta_Grind_main___lambda__2___closed__6() { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_Grind_main___lambda__2___closed__5; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; +lean_object* x_10; lean_object* x_11; +x_10 = l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___closed__1; +x_11 = l_Lean_Meta_Grind_applyToAll(x_10, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_11; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_main___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_MVarId_isAssigned___at_Lean_Meta_Grind_main___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { -lean_object* x_10; lean_object* x_11; -x_10 = l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___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); -lean_inc(x_2); -x_11 = l_Lean_Meta_Grind_all(x_1, x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_11) == 0) +lean_object* x_10; uint8_t x_11; +x_10 = lean_st_ref_get(x_6, x_9); +x_11 = !lean_is_exclusive(x_10); +if (x_11 == 0) { -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; 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_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; +x_12 = lean_ctor_get(x_10, 0); +x_13 = lean_ctor_get(x_12, 0); lean_inc(x_13); -lean_dec(x_11); -x_14 = l_Lean_Meta_Grind_main___lambda__2___closed__4; -x_15 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_simp___spec__2(x_14, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_13); -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_dec(x_12); +x_14 = lean_ctor_get(x_13, 7); +lean_inc(x_14); +lean_dec(x_13); +x_15 = l_Lean_PersistentHashMap_contains___at_Lean_MVarId_isAssigned___spec__1(x_14, x_1); +x_16 = lean_box(x_15); +lean_ctor_set(x_10, 0, x_16); +return x_10; +} +else { -lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_18 = lean_ctor_get(x_15, 1); +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_17 = lean_ctor_get(x_10, 0); +x_18 = lean_ctor_get(x_10, 1); lean_inc(x_18); -lean_dec(x_15); -x_19 = lean_box(0); -x_20 = l_Lean_Meta_Grind_main___lambda__1(x_12, x_19, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_18); -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_20; -} -else -{ -uint8_t x_21; -x_21 = !lean_is_exclusive(x_15); -if (x_21 == 0) -{ -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_15, 1); -x_23 = lean_ctor_get(x_15, 0); -lean_dec(x_23); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -lean_inc(x_2); -lean_inc(x_12); -x_24 = l_Lean_Meta_Grind_ppGoals(x_12, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_22); -if (lean_obj_tag(x_24) == 0) -{ -lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; -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_MessageData_ofFormat(x_25); -x_28 = l_Lean_Meta_Grind_main___lambda__2___closed__6; -lean_ctor_set_tag(x_15, 7); -lean_ctor_set(x_15, 1, x_27); -lean_ctor_set(x_15, 0, x_28); -x_29 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_29, 0, x_15); -lean_ctor_set(x_29, 1, x_28); -x_30 = l_Lean_addTrace___at_Lean_Meta_Grind_simp___spec__3(x_14, x_29, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_26); -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_Grind_main___lambda__1(x_12, x_31, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_32); -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_31); -return x_33; -} -else -{ -uint8_t x_34; -lean_free_object(x_15); -lean_dec(x_12); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -x_34 = !lean_is_exclusive(x_24); -if (x_34 == 0) -{ -return x_24; -} -else -{ -lean_object* x_35; lean_object* x_36; lean_object* x_37; -x_35 = lean_ctor_get(x_24, 0); -x_36 = lean_ctor_get(x_24, 1); -lean_inc(x_36); -lean_inc(x_35); -lean_dec(x_24); -x_37 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_37, 0, x_35); -lean_ctor_set(x_37, 1, x_36); -return x_37; -} -} -} -else -{ -lean_object* x_38; lean_object* x_39; -x_38 = lean_ctor_get(x_15, 1); -lean_inc(x_38); -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); -lean_inc(x_2); -lean_inc(x_12); -x_39 = l_Lean_Meta_Grind_ppGoals(x_12, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_38); -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; lean_object* x_48; lean_object* x_49; -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_MessageData_ofFormat(x_40); -x_43 = l_Lean_Meta_Grind_main___lambda__2___closed__6; -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 = 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_Lean_addTrace___at_Lean_Meta_Grind_simp___spec__3(x_14, x_45, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_41); -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_Meta_Grind_main___lambda__1(x_12, x_47, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_48); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_47); -return x_49; -} -else -{ -lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* 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); -lean_dec(x_3); -lean_dec(x_2); -x_50 = lean_ctor_get(x_39, 0); -lean_inc(x_50); -x_51 = lean_ctor_get(x_39, 1); -lean_inc(x_51); -if (lean_is_exclusive(x_39)) { - lean_ctor_release(x_39, 0); - lean_ctor_release(x_39, 1); - x_52 = x_39; -} else { - lean_dec_ref(x_39); - x_52 = lean_box(0); -} -if (lean_is_scalar(x_52)) { - x_53 = lean_alloc_ctor(1, 2, 0); -} else { - x_53 = x_52; -} -lean_ctor_set(x_53, 0, x_50); -lean_ctor_set(x_53, 1, x_51); -return x_53; -} -} -} -} -else -{ -uint8_t x_54; -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_54 = !lean_is_exclusive(x_11); -if (x_54 == 0) -{ -return x_11; -} -else -{ -lean_object* x_55; lean_object* x_56; lean_object* x_57; -x_55 = lean_ctor_get(x_11, 0); -x_56 = lean_ctor_get(x_11, 1); -lean_inc(x_56); -lean_inc(x_55); -lean_dec(x_11); -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; -} -} -} -} -static lean_object* _init_l_Lean_Meta_Grind_main___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_main___lambda__2), 9, 0); -return x_1; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_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) { -_start: -{ -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_9 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore), 9, 1); -lean_closure_set(x_9, 0, x_1); -x_10 = l_Lean_Meta_Grind_main___closed__1; -x_11 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); -lean_closure_set(x_11, 0, x_9); -lean_closure_set(x_11, 1, x_10); -x_12 = l_Lean_Meta_Grind_GrindM_run___rarg(x_11, x_3, x_2, x_4, x_5, x_6, x_7, x_8); -return x_12; -} +lean_inc(x_17); +lean_dec(x_10); +x_19 = lean_ctor_get(x_17, 0); +lean_inc(x_19); +lean_dec(x_17); +x_20 = lean_ctor_get(x_19, 7); +lean_inc(x_20); +lean_dec(x_19); +x_21 = l_Lean_PersistentHashMap_contains___at_Lean_MVarId_isAssigned___spec__1(x_20, x_1); +x_22 = lean_box(x_21); +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_18); +return x_23; } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_main___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_11; -x_11 = l_Lean_Meta_Grind_main___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_11; } } -LEAN_EXPORT lean_object* l_List_forM___at_Lean_Meta_Grind_preprocessAndProbe___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_EXPORT lean_object* l_List_filterMapM_loop___at_Lean_Meta_Grind_main___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) { _start: { lean_object* x_11; @@ -2855,394 +2473,577 @@ return x_33; } } } -LEAN_EXPORT lean_object* l_List_forM___at_Lean_Meta_Grind_preprocessAndProbe___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_List_filterMapM_loop___at_Lean_Meta_Grind_main___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) { _start: { -if (lean_obj_tag(x_2) == 0) -{ 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_11 = lean_box(0); +x_11 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_11, 0, x_1); x_12 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_12, 0, x_11); lean_ctor_set(x_12, 1, x_10); return x_12; } -else +} +LEAN_EXPORT lean_object* l_List_filterMapM_loop___at_Lean_Meta_Grind_main___spec__2___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: { -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_13 = lean_ctor_get(x_2, 0); +lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_11 = lean_ctor_get(x_1, 0); +lean_inc(x_11); +x_12 = l_Lean_MVarId_isAssigned___at_Lean_Meta_Grind_main___spec__1(x_11, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_11); +x_13 = lean_ctor_get(x_12, 0); lean_inc(x_13); -x_14 = lean_ctor_get(x_2, 1); -lean_inc(x_14); -lean_dec(x_2); -x_15 = lean_ctor_get(x_13, 0); +x_14 = lean_unbox(x_13); +lean_dec(x_13); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_ctor_get(x_12, 1); lean_inc(x_15); -x_16 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___lambda__1___boxed), 9, 1); -lean_closure_set(x_16, 0, x_13); -lean_inc(x_1); -x_17 = lean_alloc_closure((void*)(l_List_forM___at_Lean_Meta_Grind_preprocessAndProbe___spec__1___lambda__1), 10, 1); -lean_closure_set(x_17, 0, x_1); -x_18 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); -lean_closure_set(x_18, 0, x_16); -lean_closure_set(x_18, 1, x_17); -x_19 = l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___closed__3; -x_20 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); -lean_closure_set(x_20, 0, x_18); -lean_closure_set(x_20, 1, x_19); +lean_dec(x_12); +x_16 = lean_box(0); +x_17 = l_List_filterMapM_loop___at_Lean_Meta_Grind_main___spec__2___lambda__2(x_1, x_16, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_15); +return x_17; +} +else +{ +uint8_t x_18; +lean_dec(x_1); +x_18 = !lean_is_exclusive(x_12); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; +x_19 = lean_ctor_get(x_12, 0); +lean_dec(x_19); +x_20 = lean_box(0); +lean_ctor_set(x_12, 0, x_20); +return x_12; +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_12, 1); +lean_inc(x_21); +lean_dec(x_12); +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_EXPORT lean_object* l_List_filterMapM_loop___at_Lean_Meta_Grind_main___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: +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_12 = lean_ctor_get(x_1, 0); +lean_inc(x_12); +x_13 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___lambda__1___boxed), 9, 1); +lean_closure_set(x_13, 0, x_1); +x_14 = lean_alloc_closure((void*)(l_List_filterMapM_loop___at_Lean_Meta_Grind_main___spec__2___lambda__1), 10, 1); +lean_closure_set(x_14, 0, x_2); +x_15 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); +lean_closure_set(x_15, 0, x_13); +lean_closure_set(x_15, 1, x_14); +x_16 = l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___closed__3; +x_17 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); +lean_closure_set(x_17, 0, x_15); +lean_closure_set(x_17, 1, x_16); +lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -lean_inc(x_3); -x_21 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(x_15, x_20, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_21) == 0) +x_18 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(x_12, x_17, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_18) == 0) { -lean_object* x_22; -x_22 = lean_ctor_get(x_21, 1); -lean_inc(x_22); -lean_dec(x_21); -x_2 = x_14; -x_10 = x_22; -goto _start; +lean_object* x_19; uint8_t x_20; +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_ctor_get_uint8(x_19, sizeof(void*)*20); +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_List_filterMapM_loop___at_Lean_Meta_Grind_main___spec__2___lambda__3(x_19, x_22, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_21); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_23; } else { uint8_t x_24; -lean_dec(x_14); +lean_dec(x_19); +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_24 = !lean_is_exclusive(x_21); +x_24 = !lean_is_exclusive(x_18); if (x_24 == 0) { -return x_21; +lean_object* x_25; lean_object* x_26; +x_25 = lean_ctor_get(x_18, 0); +lean_dec(x_25); +x_26 = lean_box(0); +lean_ctor_set(x_18, 0, x_26); +return x_18; } else { -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_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; -} +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_18, 1); +lean_inc(x_27); +lean_dec(x_18); +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 +{ +uint8_t x_30; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_30 = !lean_is_exclusive(x_18); +if (x_30 == 0) +{ +return x_18; } +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_18, 0); +x_32 = lean_ctor_get(x_18, 1); +lean_inc(x_32); +lean_inc(x_31); +lean_dec(x_18); +x_33 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_33, 0, x_31); +lean_ctor_set(x_33, 1, x_32); +return x_33; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_withoutModifyingMCtx___at_Lean_Meta_Grind_preprocessAndProbe___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_EXPORT lean_object* l_List_filterMapM_loop___at_Lean_Meta_Grind_main___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; -x_7 = lean_st_ref_get(x_3, x_6); -x_8 = lean_ctor_get(x_7, 0); -lean_inc(x_8); -x_9 = lean_ctor_get(x_7, 1); -lean_inc(x_9); +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_12; lean_object* x_13; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_7); -x_10 = lean_ctor_get(x_8, 0); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_12 = l_List_reverse___rarg(x_3); +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_11); +return x_13; +} +else +{ +lean_object* x_14; uint8_t x_15; +x_14 = lean_ctor_get(x_2, 0); +lean_inc(x_14); +x_15 = lean_ctor_get_uint8(x_14, sizeof(void*)*20); +if (x_15 == 0) +{ +uint8_t x_16; +x_16 = !lean_is_exclusive(x_2); +if (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_2, 1); +x_18 = lean_ctor_get(x_2, 0); +lean_dec(x_18); +x_19 = lean_box(0); lean_inc(x_10); -lean_dec(x_8); +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_5(x_1, x_2, x_3, x_4, x_5, x_9); -if (lean_obj_tag(x_11) == 0) +lean_inc(x_1); +x_20 = l_List_filterMapM_loop___at_Lean_Meta_Grind_main___spec__2___lambda__4(x_14, x_1, x_19, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_20) == 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_11, 0); -lean_inc(x_12); -x_13 = lean_ctor_get(x_11, 1); -lean_inc(x_13); -lean_dec(x_11); -x_14 = l_Lean_Meta_resetCache___rarg(x_3, x_4, x_5, x_13); -lean_dec(x_5); -lean_dec(x_4); -x_15 = lean_ctor_get(x_14, 1); -lean_inc(x_15); -lean_dec(x_14); -x_16 = lean_st_ref_take(x_3, x_15); -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -x_18 = lean_ctor_get(x_16, 1); -lean_inc(x_18); -lean_dec(x_16); -x_19 = !lean_is_exclusive(x_17); -if (x_19 == 0) +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_20; lean_object* x_21; uint8_t x_22; -x_20 = lean_ctor_get(x_17, 0); +lean_object* x_22; +lean_free_object(x_2); +x_22 = lean_ctor_get(x_20, 1); +lean_inc(x_22); lean_dec(x_20); -lean_ctor_set(x_17, 0, x_10); -x_21 = lean_st_ref_set(x_3, x_17, x_18); -lean_dec(x_3); -x_22 = !lean_is_exclusive(x_21); -if (x_22 == 0) -{ -lean_object* x_23; -x_23 = lean_ctor_get(x_21, 0); -lean_dec(x_23); -lean_ctor_set(x_21, 0, x_12); -return x_21; +x_2 = x_17; +x_11 = x_22; +goto _start; } else { lean_object* x_24; lean_object* x_25; -x_24 = lean_ctor_get(x_21, 1); +x_24 = lean_ctor_get(x_20, 1); lean_inc(x_24); +lean_dec(x_20); +x_25 = lean_ctor_get(x_21, 0); +lean_inc(x_25); lean_dec(x_21); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_12); -lean_ctor_set(x_25, 1, x_24); -return x_25; +lean_ctor_set(x_2, 1, x_3); +lean_ctor_set(x_2, 0, x_25); +{ +lean_object* _tmp_1 = x_17; +lean_object* _tmp_2 = x_2; +lean_object* _tmp_10 = x_24; +x_2 = _tmp_1; +x_3 = _tmp_2; +x_11 = _tmp_10; +} +goto _start; } } else { -lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_26 = lean_ctor_get(x_17, 1); -x_27 = lean_ctor_get(x_17, 2); -x_28 = lean_ctor_get(x_17, 3); -x_29 = lean_ctor_get(x_17, 4); -lean_inc(x_29); -lean_inc(x_28); -lean_inc(x_27); -lean_inc(x_26); +uint8_t x_27; +lean_free_object(x_2); lean_dec(x_17); -x_30 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_30, 0, x_10); -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_18); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); lean_dec(x_3); -x_32 = lean_ctor_get(x_31, 1); -lean_inc(x_32); -if (lean_is_exclusive(x_31)) { - lean_ctor_release(x_31, 0); - lean_ctor_release(x_31, 1); - x_33 = x_31; -} else { - lean_dec_ref(x_31); - x_33 = lean_box(0); +lean_dec(x_1); +x_27 = !lean_is_exclusive(x_20); +if (x_27 == 0) +{ +return x_20; } -if (lean_is_scalar(x_33)) { - x_34 = lean_alloc_ctor(0, 2, 0); -} else { - x_34 = x_33; +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_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_ctor_set(x_34, 0, x_12); -lean_ctor_set(x_34, 1, x_32); -return x_34; } } else { -lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; uint8_t x_42; -x_35 = lean_ctor_get(x_11, 0); -lean_inc(x_35); -x_36 = lean_ctor_get(x_11, 1); -lean_inc(x_36); -lean_dec(x_11); -x_37 = l_Lean_Meta_resetCache___rarg(x_3, x_4, x_5, x_36); -lean_dec(x_5); -lean_dec(x_4); -x_38 = lean_ctor_get(x_37, 1); -lean_inc(x_38); -lean_dec(x_37); -x_39 = lean_st_ref_take(x_3, x_38); -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_is_exclusive(x_40); -if (x_42 == 0) -{ -lean_object* x_43; lean_object* x_44; uint8_t x_45; -x_43 = lean_ctor_get(x_40, 0); -lean_dec(x_43); -lean_ctor_set(x_40, 0, x_10); -x_44 = lean_st_ref_set(x_3, x_40, x_41); -lean_dec(x_3); -x_45 = !lean_is_exclusive(x_44); -if (x_45 == 0) +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_2, 1); +lean_inc(x_31); +lean_dec(x_2); +x_32 = lean_box(0); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_1); +x_33 = l_List_filterMapM_loop___at_Lean_Meta_Grind_main___spec__2___lambda__4(x_14, x_1, x_32, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_33) == 0) { -lean_object* x_46; -x_46 = lean_ctor_get(x_44, 0); -lean_dec(x_46); -lean_ctor_set_tag(x_44, 1); -lean_ctor_set(x_44, 0, x_35); -return x_44; +lean_object* x_34; +x_34 = lean_ctor_get(x_33, 0); +lean_inc(x_34); +if (lean_obj_tag(x_34) == 0) +{ +lean_object* x_35; +x_35 = lean_ctor_get(x_33, 1); +lean_inc(x_35); +lean_dec(x_33); +x_2 = x_31; +x_11 = x_35; +goto _start; } else { -lean_object* x_47; lean_object* x_48; -x_47 = lean_ctor_get(x_44, 1); -lean_inc(x_47); -lean_dec(x_44); -x_48 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_48, 0, x_35); -lean_ctor_set(x_48, 1, x_47); -return x_48; +lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_37 = lean_ctor_get(x_33, 1); +lean_inc(x_37); +lean_dec(x_33); +x_38 = lean_ctor_get(x_34, 0); +lean_inc(x_38); +lean_dec(x_34); +x_39 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_3); +x_2 = x_31; +x_3 = x_39; +x_11 = x_37; +goto _start; } } 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; -x_49 = lean_ctor_get(x_40, 1); -x_50 = lean_ctor_get(x_40, 2); -x_51 = lean_ctor_get(x_40, 3); -x_52 = lean_ctor_get(x_40, 4); -lean_inc(x_52); -lean_inc(x_51); -lean_inc(x_50); -lean_inc(x_49); -lean_dec(x_40); -x_53 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_53, 0, x_10); -lean_ctor_set(x_53, 1, x_49); -lean_ctor_set(x_53, 2, x_50); -lean_ctor_set(x_53, 3, x_51); -lean_ctor_set(x_53, 4, x_52); -x_54 = lean_st_ref_set(x_3, x_53, x_41); +lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; +lean_dec(x_31); +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_55 = lean_ctor_get(x_54, 1); -lean_inc(x_55); -if (lean_is_exclusive(x_54)) { - lean_ctor_release(x_54, 0); - lean_ctor_release(x_54, 1); - x_56 = x_54; +lean_dec(x_1); +x_41 = lean_ctor_get(x_33, 0); +lean_inc(x_41); +x_42 = lean_ctor_get(x_33, 1); +lean_inc(x_42); +if (lean_is_exclusive(x_33)) { + lean_ctor_release(x_33, 0); + lean_ctor_release(x_33, 1); + x_43 = x_33; } else { - lean_dec_ref(x_54); - x_56 = lean_box(0); + lean_dec_ref(x_33); + x_43 = lean_box(0); } -if (lean_is_scalar(x_56)) { - x_57 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_43)) { + x_44 = lean_alloc_ctor(1, 2, 0); } else { - x_57 = x_56; - lean_ctor_set_tag(x_57, 1); -} -lean_ctor_set(x_57, 0, x_35); -lean_ctor_set(x_57, 1, x_55); -return x_57; + x_44 = x_43; } +lean_ctor_set(x_44, 0, x_41); +lean_ctor_set(x_44, 1, x_42); +return x_44; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_withoutModifyingMCtx___at_Lean_Meta_Grind_preprocessAndProbe___spec__2(lean_object* x_1) { -_start: +else { -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_Meta_withoutModifyingMCtx___at_Lean_Meta_Grind_preprocessAndProbe___spec__2___rarg), 6, 0); -return x_2; +lean_object* x_45; +lean_dec(x_14); +x_45 = lean_ctor_get(x_2, 1); +lean_inc(x_45); +lean_dec(x_2); +x_2 = x_45; +goto _start; +} } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_preprocessAndProbe___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_List_mapTR_loop___at_Lean_Meta_Grind_main___spec__3(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_12; -x_12 = l_List_forM___at_Lean_Meta_Grind_preprocessAndProbe___spec__1(x_1, x_2, x_4, x_5, 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_12); -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_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_3; +x_3 = l_List_reverse___rarg(x_2); +return x_3; } 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; +uint8_t x_4; +x_4 = !lean_is_exclusive(x_1); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_5 = lean_ctor_get(x_1, 0); +x_6 = lean_ctor_get(x_1, 1); +x_7 = lean_ctor_get(x_5, 0); +lean_inc(x_7); +lean_dec(x_5); +lean_ctor_set(x_1, 1, x_2); +lean_ctor_set(x_1, 0, x_7); +{ +lean_object* _tmp_0 = x_6; +lean_object* _tmp_1 = x_1; +x_1 = _tmp_0; +x_2 = _tmp_1; } +goto _start; } else { -uint8_t x_19; -x_19 = !lean_is_exclusive(x_12); -if (x_19 == 0) +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_9 = lean_ctor_get(x_1, 0); +x_10 = lean_ctor_get(x_1, 1); +lean_inc(x_10); +lean_inc(x_9); +lean_dec(x_1); +x_11 = lean_ctor_get(x_9, 0); +lean_inc(x_11); +lean_dec(x_9); +x_12 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_2); +x_1 = x_10; +x_2 = x_12; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_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, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; lean_object* x_13; +x_12 = l_List_mapTR_loop___at_Lean_Meta_Grind_main___spec__3(x_1, x_2); +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_11); +return x_13; +} +} +static lean_object* _init_l_Lean_Meta_Grind_main___lambda__2___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("debug", 5, 5); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_main___lambda__2___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("final", 5, 5); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_main___lambda__2___closed__3() { +_start: { -return x_12; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore___closed__3; +x_2 = l_Lean_Meta_Grind_main___lambda__2___closed__1; +x_3 = l_Lean_Meta_Grind_main___lambda__2___closed__2; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; } -else +} +static lean_object* _init_l_Lean_Meta_Grind_main___lambda__2___closed__4() { +_start: { -lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_20 = lean_ctor_get(x_12, 0); -x_21 = lean_ctor_get(x_12, 1); -lean_inc(x_21); -lean_inc(x_20); -lean_dec(x_12); -x_22 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_22, 0, x_20); -lean_ctor_set(x_22, 1, x_21); -return x_22; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("", 0, 0); +return x_1; } } +static lean_object* _init_l_Lean_Meta_Grind_main___lambda__2___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_main___lambda__2___closed__4; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_preprocessAndProbe___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_main___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_11 = l_Lean_Meta_Grind_main___lambda__2___closed__4; -x_12 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_simp___spec__2(x_11, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_object* x_11; lean_object* x_12; +x_11 = l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___closed__1; +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_12 = l_Lean_Meta_Grind_applyToAll(x_11, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_12) == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; x_13 = lean_ctor_get(x_12, 0); lean_inc(x_13); -x_14 = lean_unbox(x_13); -lean_dec(x_13); -if (x_14 == 0) -{ -lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = lean_ctor_get(x_12, 1); -lean_inc(x_15); +x_14 = lean_ctor_get(x_12, 1); +lean_inc(x_14); lean_dec(x_12); -x_16 = lean_box(0); -x_17 = l_Lean_Meta_Grind_preprocessAndProbe___lambda__1(x_1, x_2, x_16, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_15); -return x_17; +x_15 = lean_box(0); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_16 = l_List_filterMapM_loop___at_Lean_Meta_Grind_main___spec__2(x_1, x_13, x_15, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_14); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +x_18 = lean_ctor_get(x_16, 1); +lean_inc(x_18); +lean_dec(x_16); +x_19 = l_Lean_Meta_Grind_main___lambda__2___closed__3; +x_20 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_simp___spec__2(x_19, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_18); +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; +x_23 = lean_ctor_get(x_20, 1); +lean_inc(x_23); +lean_dec(x_20); +x_24 = lean_box(0); +x_25 = l_Lean_Meta_Grind_main___lambda__1(x_17, x_15, x_24, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_23); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_25; } else { -uint8_t x_18; -x_18 = !lean_is_exclusive(x_12); -if (x_18 == 0) +uint8_t x_26; +x_26 = !lean_is_exclusive(x_20); +if (x_26 == 0) { -lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_19 = lean_ctor_get(x_12, 1); -x_20 = lean_ctor_get(x_12, 0); -lean_dec(x_20); +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_20, 1); +x_28 = lean_ctor_get(x_20, 0); +lean_dec(x_28); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); @@ -3250,38 +3051,46 @@ lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -lean_inc(x_2); -x_21 = l_Lean_Meta_Grind_ppGoals(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_19); -if (lean_obj_tag(x_21) == 0) +lean_inc(x_17); +x_29 = l_Lean_Meta_Grind_ppGoals(x_17, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_27); +if (lean_obj_tag(x_29) == 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; -x_22 = lean_ctor_get(x_21, 0); -lean_inc(x_22); -x_23 = lean_ctor_get(x_21, 1); -lean_inc(x_23); -lean_dec(x_21); -x_24 = l_Lean_MessageData_ofFormat(x_22); -x_25 = l_Lean_Meta_Grind_main___lambda__2___closed__6; -lean_ctor_set_tag(x_12, 7); -lean_ctor_set(x_12, 1, x_24); -lean_ctor_set(x_12, 0, x_25); -x_26 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_26, 0, x_12); -lean_ctor_set(x_26, 1, x_25); -x_27 = l_Lean_addTrace___at_Lean_Meta_Grind_simp___spec__3(x_11, x_26, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_23); -x_28 = lean_ctor_get(x_27, 0); -lean_inc(x_28); -x_29 = lean_ctor_get(x_27, 1); -lean_inc(x_29); -lean_dec(x_27); -x_30 = l_Lean_Meta_Grind_preprocessAndProbe___lambda__1(x_1, x_2, x_28, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_29); -lean_dec(x_28); -return x_30; +lean_object* x_30; lean_object* x_31; lean_object* x_32; 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 = l_Lean_MessageData_ofFormat(x_30); +x_33 = l_Lean_Meta_Grind_main___lambda__2___closed__5; +lean_ctor_set_tag(x_20, 7); +lean_ctor_set(x_20, 1, x_32); +lean_ctor_set(x_20, 0, x_33); +x_34 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_34, 0, x_20); +lean_ctor_set(x_34, 1, x_33); +x_35 = l_Lean_addTrace___at_Lean_Meta_Grind_simp___spec__3(x_19, x_34, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_31); +x_36 = lean_ctor_get(x_35, 0); +lean_inc(x_36); +x_37 = lean_ctor_get(x_35, 1); +lean_inc(x_37); +lean_dec(x_35); +x_38 = l_Lean_Meta_Grind_main___lambda__1(x_17, x_15, x_36, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_37); +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_36); +return x_38; } else { -uint8_t x_31; -lean_free_object(x_12); +uint8_t x_39; +lean_free_object(x_20); +lean_dec(x_17); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -3289,34 +3098,32 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_31 = !lean_is_exclusive(x_21); -if (x_31 == 0) +x_39 = !lean_is_exclusive(x_29); +if (x_39 == 0) { -return x_21; +return x_29; } else { -lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_32 = lean_ctor_get(x_21, 0); -x_33 = lean_ctor_get(x_21, 1); -lean_inc(x_33); -lean_inc(x_32); -lean_dec(x_21); -x_34 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_34, 0, x_32); -lean_ctor_set(x_34, 1, x_33); -return x_34; +lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_40 = lean_ctor_get(x_29, 0); +x_41 = lean_ctor_get(x_29, 1); +lean_inc(x_41); +lean_inc(x_40); +lean_dec(x_29); +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 { -lean_object* x_35; lean_object* x_36; -x_35 = lean_ctor_get(x_12, 1); -lean_inc(x_35); -lean_dec(x_12); +lean_object* x_43; lean_object* x_44; +x_43 = lean_ctor_get(x_20, 1); +lean_inc(x_43); +lean_dec(x_20); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); @@ -3324,37 +3131,45 @@ lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -lean_inc(x_2); -x_36 = l_Lean_Meta_Grind_ppGoals(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_35); -if (lean_obj_tag(x_36) == 0) +lean_inc(x_17); +x_44 = l_Lean_Meta_Grind_ppGoals(x_17, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_43); +if (lean_obj_tag(x_44) == 0) { -lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; -x_37 = lean_ctor_get(x_36, 0); -lean_inc(x_37); -x_38 = lean_ctor_get(x_36, 1); -lean_inc(x_38); -lean_dec(x_36); -x_39 = l_Lean_MessageData_ofFormat(x_37); -x_40 = l_Lean_Meta_Grind_main___lambda__2___closed__6; -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 = 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_addTrace___at_Lean_Meta_Grind_simp___spec__3(x_11, x_42, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_38); -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; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; +x_45 = lean_ctor_get(x_44, 0); lean_inc(x_45); -lean_dec(x_43); -x_46 = l_Lean_Meta_Grind_preprocessAndProbe___lambda__1(x_1, x_2, x_44, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_45); +x_46 = lean_ctor_get(x_44, 1); +lean_inc(x_46); lean_dec(x_44); -return x_46; +x_47 = l_Lean_MessageData_ofFormat(x_45); +x_48 = l_Lean_Meta_Grind_main___lambda__2___closed__5; +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 = 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_addTrace___at_Lean_Meta_Grind_simp___spec__3(x_19, x_50, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_46); +x_52 = lean_ctor_get(x_51, 0); +lean_inc(x_52); +x_53 = lean_ctor_get(x_51, 1); +lean_inc(x_53); +lean_dec(x_51); +x_54 = l_Lean_Meta_Grind_main___lambda__1(x_17, x_15, x_52, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_53); +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_52); +return x_54; } else { -lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; +lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; +lean_dec(x_17); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -3362,75 +3177,177 @@ 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_ctor_get(x_36, 0); -lean_inc(x_47); -x_48 = lean_ctor_get(x_36, 1); -lean_inc(x_48); -if (lean_is_exclusive(x_36)) { - lean_ctor_release(x_36, 0); - lean_ctor_release(x_36, 1); - x_49 = x_36; +x_55 = lean_ctor_get(x_44, 0); +lean_inc(x_55); +x_56 = lean_ctor_get(x_44, 1); +lean_inc(x_56); +if (lean_is_exclusive(x_44)) { + lean_ctor_release(x_44, 0); + lean_ctor_release(x_44, 1); + x_57 = x_44; } else { - lean_dec_ref(x_36); - x_49 = lean_box(0); + lean_dec_ref(x_44); + x_57 = lean_box(0); } -if (lean_is_scalar(x_49)) { - x_50 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_57)) { + x_58 = lean_alloc_ctor(1, 2, 0); } else { - x_50 = x_49; + x_58 = x_57; } -lean_ctor_set(x_50, 0, x_47); -lean_ctor_set(x_50, 1, x_48); -return x_50; +lean_ctor_set(x_58, 0, x_55); +lean_ctor_set(x_58, 1, x_56); +return x_58; +} +} +} +} +else +{ +uint8_t x_59; +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_59 = !lean_is_exclusive(x_16); +if (x_59 == 0) +{ +return x_16; +} +else +{ +lean_object* x_60; lean_object* x_61; lean_object* x_62; +x_60 = lean_ctor_get(x_16, 0); +x_61 = lean_ctor_get(x_16, 1); +lean_inc(x_61); +lean_inc(x_60); +lean_dec(x_16); +x_62 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_62, 0, x_60); +lean_ctor_set(x_62, 1, x_61); +return x_62; +} +} +} +else +{ +uint8_t x_63; +lean_dec(x_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_63 = !lean_is_exclusive(x_12); +if (x_63 == 0) +{ +return x_12; +} +else +{ +lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_64 = lean_ctor_get(x_12, 0); +x_65 = lean_ctor_get(x_12, 1); +lean_inc(x_65); +lean_inc(x_64); +lean_dec(x_12); +x_66 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_66, 0, x_64); +lean_ctor_set(x_66, 1, x_65); +return x_66; } } } } +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_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: +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_10 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore), 9, 1); +lean_closure_set(x_10, 0, x_1); +lean_inc(x_4); +x_11 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_main___lambda__2), 10, 1); +lean_closure_set(x_11, 0, x_4); +x_12 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); +lean_closure_set(x_12, 0, x_10); +lean_closure_set(x_12, 1, x_11); +x_13 = l_Lean_Meta_Grind_GrindM_run___rarg(x_12, x_3, x_2, x_4, x_5, x_6, x_7, x_8, x_9); +return x_13; +} } -static lean_object* _init_l_Lean_Meta_Grind_preprocessAndProbe___closed__1() { +LEAN_EXPORT lean_object* l_Lean_MVarId_isAssigned___at_Lean_Meta_Grind_main___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _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 = lean_unsigned_to_nat(100u); -x_3 = lean_unsigned_to_nat(5u); -x_4 = lean_unsigned_to_nat(1000u); -x_5 = lean_alloc_ctor(0, 4, 1); -lean_ctor_set(x_5, 0, x_2); -lean_ctor_set(x_5, 1, x_3); -lean_ctor_set(x_5, 2, x_3); -lean_ctor_set(x_5, 3, x_4); -lean_ctor_set_uint8(x_5, sizeof(void*)*4, x_1); -return x_5; +lean_object* x_10; +x_10 = l_Lean_MVarId_isAssigned___at_Lean_Meta_Grind_main___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); +lean_dec(x_1); +return x_10; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_preprocessAndProbe(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_List_filterMapM_loop___at_Lean_Meta_Grind_main___spec__2___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _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 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore), 9, 1); -lean_closure_set(x_9, 0, x_1); -x_10 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_preprocessAndProbe___lambda__2), 10, 1); -lean_closure_set(x_10, 0, x_3); -x_11 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); -lean_closure_set(x_11, 0, x_9); -lean_closure_set(x_11, 1, x_10); -x_12 = l_Lean_Meta_Grind_preprocessAndProbe___closed__1; -x_13 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_GrindM_run___rarg), 8, 3); -lean_closure_set(x_13, 0, x_11); -lean_closure_set(x_13, 1, x_2); -lean_closure_set(x_13, 2, x_12); -x_14 = l_Lean_Meta_withoutModifyingMCtx___at_Lean_Meta_Grind_preprocessAndProbe___spec__2___rarg(x_13, x_4, x_5, x_6, x_7, x_8); -return x_14; +lean_object* x_11; +x_11 = l_List_filterMapM_loop___at_Lean_Meta_Grind_main___spec__2___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_11; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_preprocessAndProbe___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_List_filterMapM_loop___at_Lean_Meta_Grind_main___spec__2___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_List_filterMapM_loop___at_Lean_Meta_Grind_main___spec__2___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_11; +} +} +LEAN_EXPORT lean_object* l_List_filterMapM_loop___at_Lean_Meta_Grind_main___spec__2___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; +x_12 = l_List_filterMapM_loop___at_Lean_Meta_Grind_main___spec__2___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_3); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_main___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_Meta_Grind_preprocessAndProbe___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_12 = l_Lean_Meta_Grind_main___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); lean_dec(x_3); return x_12; } @@ -3445,6 +3362,8 @@ lean_object* initialize_Lean_Meta_Tactic_Grind_Util(uint8_t builtin, lean_object lean_object* initialize_Lean_Meta_Tactic_Grind_Inv(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_Tactic_Grind_Intro(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_Tactic_Grind_EMatch(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Meta_Tactic_Grind_Split(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Meta_Tactic_Grind_SimpUtil(uint8_t builtin, lean_object*); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Lean_Meta_Tactic_Grind_Main(uint8_t builtin, lean_object* w) { lean_object * res; @@ -3480,8 +3399,14 @@ lean_dec_ref(res); res = initialize_Lean_Meta_Tactic_Grind_EMatch(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l_Lean_Meta_Grind_mkMethods___rarg___closed__1 = _init_l_Lean_Meta_Grind_mkMethods___rarg___closed__1(); -lean_mark_persistent(l_Lean_Meta_Grind_mkMethods___rarg___closed__1); +res = initialize_Lean_Meta_Tactic_Grind_Split(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Lean_Meta_Tactic_Grind_SimpUtil(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l_Lean_Meta_Grind_mkMethods___closed__1 = _init_l_Lean_Meta_Grind_mkMethods___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_mkMethods___closed__1); l_Lean_Meta_Grind_GrindM_run___rarg___closed__1 = _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__1(); lean_mark_persistent(l_Lean_Meta_Grind_GrindM_run___rarg___closed__1); l_Lean_Meta_Grind_GrindM_run___rarg___closed__2 = _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__2(); @@ -3516,12 +3441,6 @@ l_Lean_Meta_Grind_GrindM_run___rarg___closed__16 = _init_l_Lean_Meta_Grind_Grind lean_mark_persistent(l_Lean_Meta_Grind_GrindM_run___rarg___closed__16); l_Lean_Meta_Grind_GrindM_run___rarg___closed__17 = _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__17(); lean_mark_persistent(l_Lean_Meta_Grind_GrindM_run___rarg___closed__17); -l_Lean_Meta_Grind_GrindM_run___rarg___closed__18 = _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__18(); -lean_mark_persistent(l_Lean_Meta_Grind_GrindM_run___rarg___closed__18); -l_Lean_Meta_Grind_GrindM_run___rarg___closed__19 = _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__19(); -lean_mark_persistent(l_Lean_Meta_Grind_GrindM_run___rarg___closed__19); -l_Lean_Meta_Grind_GrindM_run___rarg___closed__20 = _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__20(); -lean_mark_persistent(l_Lean_Meta_Grind_GrindM_run___rarg___closed__20); l_Lean_PersistentHashMap_empty___at___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___spec__1 = _init_l_Lean_PersistentHashMap_empty___at___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___spec__1(); lean_mark_persistent(l_Lean_PersistentHashMap_empty___at___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___spec__1); l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___closed__1(); @@ -3534,6 +3453,24 @@ l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore___closed__1 lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore___closed__1); l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore___closed__2(); lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore___closed__2); +l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore___closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore___closed__3); +l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore___closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore___closed__4(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore___closed__4); +l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___lambda__1___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___lambda__1___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___lambda__1___closed__1); +l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___lambda__1___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___lambda__1___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___lambda__1___closed__2); +l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___lambda__1___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___lambda__1___closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___lambda__1___closed__3); +l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___lambda__1___closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___lambda__1___closed__4(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___lambda__1___closed__4); +l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___lambda__2___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___lambda__2___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___lambda__2___closed__1); +l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___lambda__2___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___lambda__2___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___lambda__2___closed__2); +l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___lambda__3___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___lambda__3___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___lambda__3___closed__1); l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___closed__1(); lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___closed__1); l_Lean_Meta_Grind_main___lambda__2___closed__1 = _init_l_Lean_Meta_Grind_main___lambda__2___closed__1(); @@ -3546,12 +3483,6 @@ l_Lean_Meta_Grind_main___lambda__2___closed__4 = _init_l_Lean_Meta_Grind_main___ lean_mark_persistent(l_Lean_Meta_Grind_main___lambda__2___closed__4); l_Lean_Meta_Grind_main___lambda__2___closed__5 = _init_l_Lean_Meta_Grind_main___lambda__2___closed__5(); lean_mark_persistent(l_Lean_Meta_Grind_main___lambda__2___closed__5); -l_Lean_Meta_Grind_main___lambda__2___closed__6 = _init_l_Lean_Meta_Grind_main___lambda__2___closed__6(); -lean_mark_persistent(l_Lean_Meta_Grind_main___lambda__2___closed__6); -l_Lean_Meta_Grind_main___closed__1 = _init_l_Lean_Meta_Grind_main___closed__1(); -lean_mark_persistent(l_Lean_Meta_Grind_main___closed__1); -l_Lean_Meta_Grind_preprocessAndProbe___closed__1 = _init_l_Lean_Meta_Grind_preprocessAndProbe___closed__1(); -lean_mark_persistent(l_Lean_Meta_Grind_preprocessAndProbe___closed__1); return lean_io_result_mk_ok(lean_box(0)); } #ifdef __cplusplus diff --git a/stage0/stdlib/Lean/Meta/Tactic/Grind/MarkNestedProofs.c b/stage0/stdlib/Lean/Meta/Tactic/Grind/MarkNestedProofs.c index 2f082515ee08..cf887b54feee 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Grind/MarkNestedProofs.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Grind/MarkNestedProofs.c @@ -14,4872 +14,1913 @@ extern "C" { #endif lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___closed__2; +static lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3___closed__4; lean_object* l_Lean_mkAppN(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markNestedProofs_unsafe__1(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*); lean_object* l_Lean_mkAppB(lean_object*, lean_object*, lean_object*); uint64_t lean_uint64_mix_hash(uint64_t, uint64_t); size_t lean_uint64_to_usize(uint64_t); +LEAN_EXPORT lean_object* l_panic___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__7(lean_object*, 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_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__6(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_panic___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__7___closed__1; lean_object* lean_mk_array(lean_object*, lean_object*); uint8_t lean_usize_dec_eq(size_t, size_t); -lean_object* l_Lean_Expr_mdata___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markNestedProofs(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* lean_array_fget(lean_object*, lean_object*); lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_markNestedProofsImpl_visit___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_panic___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__6(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__6___closed__2; uint8_t l_Lean_Expr_isAppOf(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__8(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__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_x27_loop___at_Lean_Meta_Grind_markNestedProofsImpl_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*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__10___boxed(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___closed__3; -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_markNestedProofsImpl_visit___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_Meta_Grind_markNestedProofsImpl_visit___lambda__6___closed__3; size_t lean_ptr_addr(lean_object*); lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkPtrMap___rarg(lean_object*); size_t lean_usize_of_nat(lean_object*); -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__6(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__6___closed__1; lean_object* lean_st_ref_take(lean_object*, lean_object*); uint64_t lean_uint64_shift_right(uint64_t, uint64_t); uint64_t lean_usize_to_uint64(size_t); +static lean_object* l_panic___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__7___closed__2; lean_object* lean_nat_div(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___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_Meta_Grind_markNestedProofsImpl_visit___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_Meta_Grind_markNestedProofsImpl_visit___lambda__3___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markNestedProofsImpl(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); lean_object* l_outOfBounds___rarg(lean_object*); lean_object* lean_st_ref_get(lean_object*, lean_object*); lean_object* lean_st_mk_ref(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__3___boxed(lean_object*, lean_object*); +uint8_t l_Lean_Expr_hasLooseBVars(lean_object*); +static lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3___closed__2; extern lean_object* l_Lean_levelZero; -static lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___closed__1; extern lean_object* l_Lean_instInhabitedExpr; -static lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2___closed__5; -LEAN_EXPORT uint8_t l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__3(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand_go___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__3(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__7(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_Meta_Grind_markNestedProofsImpl_visit___lambda__2___closed__1; extern lean_object* l_Lean_Meta_instMonadMetaM; LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__10(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_markNestedProofsImpl___closed__1; -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand_go___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__5(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__4(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__4___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__5(lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); +uint8_t l_Lean_Expr_isProj(lean_object*); +static lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__6___closed__5; lean_object* lean_array_set(lean_object*, lean_object*, lean_object*); uint64_t lean_uint64_xor(uint64_t, uint64_t); lean_object* lean_panic_fn(lean_object*, lean_object*); -static lean_object* l_panic___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__1___closed__2; lean_object* lean_nat_sub(lean_object*, lean_object*); lean_object* lean_nat_mul(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2___closed__4; -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__4(lean_object*); +static lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3___closed__3; +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__2(lean_object*); size_t lean_usize_sub(size_t, size_t); -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__6___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__7(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___closed__4; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_uget(lean_object*, size_t); +static lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__6___closed__4; lean_object* l_instInhabitedOfMonad___rarg(lean_object*, lean_object*); lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___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_Grind_markNestedProofsImpl_visit___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_array_get_size(lean_object*); -static lean_object* l_panic___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__1___closed__1; 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_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2___closed__2; -static lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__1___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3___closed__5; lean_object* lean_nat_add(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Expr_isForall(lean_object*); lean_object* lean_array_uset(lean_object*, size_t, lean_object*); -static lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2___closed__3; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_isProof(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_ReaderT_instMonad___rarg(lean_object*); size_t lean_usize_land(size_t, size_t); -static lean_object* _init_l_panic___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__1___closed__1() { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__5___boxed(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_406_(uint8_t, uint8_t); +LEAN_EXPORT uint8_t l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__1(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_instMonadMetaM; -x_2 = l_ReaderT_instMonad___rarg(x_1); -return x_2; -} -} -static lean_object* _init_l_panic___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__1___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_panic___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__1___closed__1; -x_2 = l_Lean_instInhabitedExpr; -x_3 = l_instInhabitedOfMonad___rarg(x_1, x_2); -return x_3; -} -} -LEAN_EXPORT lean_object* l_panic___at_Lean_Meta_Grind_markNestedProofsImpl_visit___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; lean_object* x_10; -x_8 = l_panic___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__1___closed__2; -x_9 = lean_panic_fn(x_8, x_1); -x_10 = lean_apply_6(x_9, x_2, x_3, x_4, x_5, x_6, x_7); -return x_10; -} -} -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -lean_object* x_13; uint8_t x_14; -x_13 = lean_ctor_get(x_2, 1); -x_14 = lean_nat_dec_lt(x_4, x_13); -if (x_14 == 0) -{ -lean_object* x_15; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_4); -x_15 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_15, 0, x_3); -lean_ctor_set(x_15, 1, x_12); -return x_15; -} -else -{ -lean_object* x_16; lean_object* x_17; uint8_t x_23; -x_23 = !lean_is_exclusive(x_3); -if (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_3, 0); -x_25 = lean_ctor_get(x_3, 1); -x_26 = lean_array_get_size(x_24); -x_27 = lean_nat_dec_lt(x_4, x_26); -lean_dec(x_26); -if (x_27 == 0) -{ -lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_28 = l_Lean_instInhabitedExpr; -x_29 = l_outOfBounds___rarg(x_28); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_29); -x_30 = l_Lean_Meta_Grind_markNestedProofsImpl_visit(x_29, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_30) == 0) -{ -lean_object* x_31; lean_object* x_32; size_t x_33; size_t x_34; uint8_t x_35; -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_ptr_addr(x_29); -lean_dec(x_29); -x_34 = lean_ptr_addr(x_31); -x_35 = lean_usize_dec_eq(x_33, x_34); -if (x_35 == 0) -{ -lean_object* x_36; uint8_t x_37; lean_object* x_38; lean_object* x_39; -lean_dec(x_25); -x_36 = lean_array_set(x_24, x_4, x_31); -x_37 = 1; -x_38 = lean_box(x_37); -lean_ctor_set(x_3, 1, x_38); -lean_ctor_set(x_3, 0, x_36); -x_39 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_39, 0, x_3); -x_16 = x_39; -x_17 = x_32; -goto block_22; -} -else -{ -lean_object* x_40; -lean_dec(x_31); -x_40 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_40, 0, x_3); -x_16 = x_40; -x_17 = x_32; -goto block_22; -} -} -else -{ -uint8_t x_41; -lean_dec(x_29); -lean_free_object(x_3); -lean_dec(x_25); -lean_dec(x_24); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_4); -x_41 = !lean_is_exclusive(x_30); -if (x_41 == 0) -{ -return x_30; -} -else -{ -lean_object* x_42; lean_object* x_43; lean_object* x_44; -x_42 = lean_ctor_get(x_30, 0); -x_43 = lean_ctor_get(x_30, 1); -lean_inc(x_43); -lean_inc(x_42); -lean_dec(x_30); -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_45; lean_object* x_46; -x_45 = lean_array_fget(x_24, 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_45); -x_46 = l_Lean_Meta_Grind_markNestedProofsImpl_visit(x_45, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_46) == 0) -{ -lean_object* x_47; lean_object* x_48; size_t x_49; size_t x_50; uint8_t x_51; -x_47 = lean_ctor_get(x_46, 0); -lean_inc(x_47); -x_48 = lean_ctor_get(x_46, 1); -lean_inc(x_48); -lean_dec(x_46); -x_49 = lean_ptr_addr(x_45); -lean_dec(x_45); -x_50 = lean_ptr_addr(x_47); -x_51 = lean_usize_dec_eq(x_49, x_50); -if (x_51 == 0) -{ -lean_object* x_52; uint8_t x_53; lean_object* x_54; lean_object* x_55; -lean_dec(x_25); -x_52 = lean_array_set(x_24, x_4, x_47); -x_53 = 1; -x_54 = lean_box(x_53); -lean_ctor_set(x_3, 1, x_54); -lean_ctor_set(x_3, 0, x_52); -x_55 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_55, 0, x_3); -x_16 = x_55; -x_17 = x_48; -goto block_22; -} -else -{ -lean_object* x_56; -lean_dec(x_47); -x_56 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_56, 0, x_3); -x_16 = x_56; -x_17 = x_48; -goto block_22; -} -} -else -{ -uint8_t x_57; -lean_dec(x_45); -lean_free_object(x_3); -lean_dec(x_25); -lean_dec(x_24); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_4); -x_57 = !lean_is_exclusive(x_46); -if (x_57 == 0) -{ -return x_46; -} -else -{ -lean_object* x_58; lean_object* x_59; lean_object* x_60; -x_58 = lean_ctor_get(x_46, 0); -x_59 = lean_ctor_get(x_46, 1); -lean_inc(x_59); -lean_inc(x_58); -lean_dec(x_46); -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; -x_61 = lean_ctor_get(x_3, 0); -x_62 = lean_ctor_get(x_3, 1); -lean_inc(x_62); -lean_inc(x_61); -lean_dec(x_3); -x_63 = lean_array_get_size(x_61); -x_64 = lean_nat_dec_lt(x_4, x_63); -lean_dec(x_63); -if (x_64 == 0) -{ -lean_object* x_65; lean_object* x_66; lean_object* x_67; -x_65 = l_Lean_instInhabitedExpr; -x_66 = l_outOfBounds___rarg(x_65); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_66); -x_67 = l_Lean_Meta_Grind_markNestedProofsImpl_visit(x_66, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_67) == 0) -{ -lean_object* x_68; lean_object* x_69; size_t x_70; size_t x_71; uint8_t x_72; -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_ptr_addr(x_66); -lean_dec(x_66); -x_71 = lean_ptr_addr(x_68); -x_72 = lean_usize_dec_eq(x_70, x_71); -if (x_72 == 0) -{ -lean_object* x_73; uint8_t x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; -lean_dec(x_62); -x_73 = lean_array_set(x_61, x_4, x_68); -x_74 = 1; -x_75 = lean_box(x_74); -x_76 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_76, 0, x_73); -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_16 = x_77; -x_17 = x_69; -goto block_22; -} -else -{ -lean_object* x_78; lean_object* x_79; -lean_dec(x_68); -x_78 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_78, 0, x_61); -lean_ctor_set(x_78, 1, x_62); -x_79 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_79, 0, x_78); -x_16 = x_79; -x_17 = x_69; -goto block_22; -} -} -else -{ -lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; -lean_dec(x_66); -lean_dec(x_62); -lean_dec(x_61); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_4); -x_80 = lean_ctor_get(x_67, 0); -lean_inc(x_80); -x_81 = lean_ctor_get(x_67, 1); -lean_inc(x_81); -if (lean_is_exclusive(x_67)) { - lean_ctor_release(x_67, 0); - lean_ctor_release(x_67, 1); - x_82 = x_67; -} else { - lean_dec_ref(x_67); - 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; -x_84 = lean_array_fget(x_61, 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_84); -x_85 = l_Lean_Meta_Grind_markNestedProofsImpl_visit(x_84, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_85) == 0) -{ -lean_object* x_86; lean_object* x_87; size_t x_88; size_t x_89; uint8_t x_90; -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_ptr_addr(x_84); -lean_dec(x_84); -x_89 = lean_ptr_addr(x_86); -x_90 = lean_usize_dec_eq(x_88, x_89); -if (x_90 == 0) -{ -lean_object* x_91; uint8_t x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; -lean_dec(x_62); -x_91 = lean_array_set(x_61, x_4, x_86); -x_92 = 1; -x_93 = lean_box(x_92); -x_94 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_94, 0, x_91); -lean_ctor_set(x_94, 1, x_93); -x_95 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_95, 0, x_94); -x_16 = x_95; -x_17 = x_87; -goto block_22; -} -else -{ -lean_object* x_96; lean_object* x_97; -lean_dec(x_86); -x_96 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_96, 0, x_61); -lean_ctor_set(x_96, 1, x_62); -x_97 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_97, 0, x_96); -x_16 = x_97; -x_17 = x_87; -goto block_22; -} -} -else -{ -lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; -lean_dec(x_84); -lean_dec(x_62); -lean_dec(x_61); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_4); -x_98 = lean_ctor_get(x_85, 0); -lean_inc(x_98); -x_99 = lean_ctor_get(x_85, 1); -lean_inc(x_99); -if (lean_is_exclusive(x_85)) { - lean_ctor_release(x_85, 0); - lean_ctor_release(x_85, 1); - x_100 = x_85; -} else { - lean_dec_ref(x_85); - x_100 = lean_box(0); -} -if (lean_is_scalar(x_100)) { - x_101 = lean_alloc_ctor(1, 2, 0); -} else { - x_101 = x_100; -} -lean_ctor_set(x_101, 0, x_98); -lean_ctor_set(x_101, 1, x_99); -return x_101; -} -} -} -block_22: -{ -lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_18 = lean_ctor_get(x_16, 0); -lean_inc(x_18); -lean_dec(x_16); -x_19 = lean_ctor_get(x_2, 2); -x_20 = lean_nat_add(x_4, x_19); -lean_dec(x_4); -x_3 = x_18; -x_4 = x_20; -x_5 = lean_box(0); -x_6 = lean_box(0); -x_12 = x_17; -goto _start; -} -} -} -} -LEAN_EXPORT uint8_t l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__3(lean_object* x_1, lean_object* x_2) { -_start: -{ -if (lean_obj_tag(x_2) == 0) -{ -uint8_t x_3; -x_3 = 0; -return x_3; -} -else -{ -lean_object* x_4; lean_object* x_5; size_t x_6; size_t x_7; uint8_t x_8; -x_4 = lean_ctor_get(x_2, 0); -x_5 = lean_ctor_get(x_2, 2); -x_6 = lean_ptr_addr(x_4); -x_7 = lean_ptr_addr(x_1); -x_8 = lean_usize_dec_eq(x_6, x_7); -if (x_8 == 0) -{ -x_2 = x_5; -goto _start; -} -else -{ -uint8_t x_10; -x_10 = 1; -return x_10; -} -} -} -} -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -if (lean_obj_tag(x_3) == 0) -{ -lean_dec(x_1); -return x_2; -} -else -{ -uint8_t x_4; -x_4 = !lean_is_exclusive(x_3); -if (x_4 == 0) -{ -lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint64_t x_9; uint64_t x_10; uint64_t x_11; uint64_t x_12; uint64_t x_13; uint64_t x_14; uint64_t x_15; size_t x_16; size_t x_17; size_t x_18; size_t x_19; size_t x_20; lean_object* x_21; lean_object* x_22; -x_5 = lean_ctor_get(x_3, 0); -x_6 = lean_ctor_get(x_3, 2); -x_7 = lean_array_get_size(x_2); -lean_inc(x_1); -lean_inc(x_5); -x_8 = lean_apply_1(x_1, x_5); -x_9 = lean_unbox_uint64(x_8); -lean_dec(x_8); -x_10 = 32; -x_11 = lean_uint64_shift_right(x_9, x_10); -x_12 = lean_uint64_xor(x_9, x_11); -x_13 = 16; -x_14 = lean_uint64_shift_right(x_12, x_13); -x_15 = lean_uint64_xor(x_12, x_14); -x_16 = lean_uint64_to_usize(x_15); -x_17 = lean_usize_of_nat(x_7); -lean_dec(x_7); -x_18 = 1; -x_19 = lean_usize_sub(x_17, x_18); -x_20 = lean_usize_land(x_16, x_19); -x_21 = lean_array_uget(x_2, x_20); -lean_ctor_set(x_3, 2, x_21); -x_22 = lean_array_uset(x_2, x_20, x_3); -x_2 = x_22; -x_3 = x_6; -goto _start; -} -else -{ -lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; uint64_t x_29; uint64_t x_30; uint64_t x_31; uint64_t x_32; uint64_t x_33; uint64_t x_34; uint64_t x_35; size_t x_36; size_t x_37; size_t x_38; size_t x_39; size_t x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; -x_24 = lean_ctor_get(x_3, 0); -x_25 = lean_ctor_get(x_3, 1); -x_26 = lean_ctor_get(x_3, 2); -lean_inc(x_26); -lean_inc(x_25); -lean_inc(x_24); -lean_dec(x_3); -x_27 = lean_array_get_size(x_2); -lean_inc(x_1); -lean_inc(x_24); -x_28 = lean_apply_1(x_1, x_24); -x_29 = lean_unbox_uint64(x_28); -lean_dec(x_28); -x_30 = 32; -x_31 = lean_uint64_shift_right(x_29, x_30); -x_32 = lean_uint64_xor(x_29, x_31); -x_33 = 16; -x_34 = lean_uint64_shift_right(x_32, x_33); -x_35 = lean_uint64_xor(x_32, x_34); -x_36 = lean_uint64_to_usize(x_35); -x_37 = lean_usize_of_nat(x_27); -lean_dec(x_27); -x_38 = 1; -x_39 = lean_usize_sub(x_37, x_38); -x_40 = lean_usize_land(x_36, x_39); -x_41 = lean_array_uget(x_2, x_40); -x_42 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_42, 0, x_24); -lean_ctor_set(x_42, 1, x_25); -lean_ctor_set(x_42, 2, x_41); -x_43 = lean_array_uset(x_2, x_40, x_42); -x_2 = x_43; -x_3 = x_26; -goto _start; -} -} -} -} -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__6___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__7(lean_object* x_1, lean_object* x_2) { -_start: -{ -if (lean_obj_tag(x_2) == 0) -{ -return x_1; -} -else -{ -uint8_t x_3; -x_3 = !lean_is_exclusive(x_2); -if (x_3 == 0) -{ -lean_object* x_4; lean_object* x_5; lean_object* x_6; size_t x_7; uint64_t x_8; uint64_t 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; size_t x_17; size_t x_18; size_t x_19; size_t x_20; size_t x_21; lean_object* x_22; lean_object* x_23; -x_4 = lean_ctor_get(x_2, 0); -x_5 = lean_ctor_get(x_2, 2); -x_6 = lean_array_get_size(x_1); -x_7 = lean_ptr_addr(x_4); -x_8 = lean_usize_to_uint64(x_7); -x_9 = 11; -x_10 = lean_uint64_mix_hash(x_8, x_9); -x_11 = 32; -x_12 = lean_uint64_shift_right(x_10, x_11); -x_13 = lean_uint64_xor(x_10, x_12); -x_14 = 16; -x_15 = lean_uint64_shift_right(x_13, x_14); -x_16 = lean_uint64_xor(x_13, x_15); -x_17 = lean_uint64_to_usize(x_16); -x_18 = lean_usize_of_nat(x_6); -lean_dec(x_6); -x_19 = 1; -x_20 = lean_usize_sub(x_18, x_19); -x_21 = lean_usize_land(x_17, x_20); -x_22 = lean_array_uget(x_1, x_21); -lean_ctor_set(x_2, 2, x_22); -x_23 = lean_array_uset(x_1, x_21, x_2); -x_1 = x_23; -x_2 = x_5; -goto _start; -} -else -{ -lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; size_t x_29; uint64_t x_30; uint64_t x_31; uint64_t x_32; uint64_t x_33; uint64_t x_34; uint64_t x_35; uint64_t x_36; uint64_t x_37; uint64_t x_38; size_t x_39; size_t x_40; size_t x_41; size_t x_42; size_t x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; -x_25 = lean_ctor_get(x_2, 0); -x_26 = lean_ctor_get(x_2, 1); -x_27 = lean_ctor_get(x_2, 2); -lean_inc(x_27); -lean_inc(x_26); -lean_inc(x_25); -lean_dec(x_2); -x_28 = lean_array_get_size(x_1); -x_29 = lean_ptr_addr(x_25); -x_30 = lean_usize_to_uint64(x_29); -x_31 = 11; -x_32 = lean_uint64_mix_hash(x_30, x_31); -x_33 = 32; -x_34 = lean_uint64_shift_right(x_32, x_33); -x_35 = lean_uint64_xor(x_32, x_34); -x_36 = 16; -x_37 = lean_uint64_shift_right(x_35, x_36); -x_38 = lean_uint64_xor(x_35, x_37); -x_39 = lean_uint64_to_usize(x_38); -x_40 = lean_usize_of_nat(x_28); -lean_dec(x_28); -x_41 = 1; -x_42 = lean_usize_sub(x_40, x_41); -x_43 = lean_usize_land(x_39, x_42); -x_44 = lean_array_uget(x_1, x_43); -x_45 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_45, 0, x_25); -lean_ctor_set(x_45, 1, x_26); -lean_ctor_set(x_45, 2, x_44); -x_46 = lean_array_uset(x_1, x_43, x_45); -x_1 = x_46; -x_2 = x_27; -goto _start; -} -} -} -} -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand_go___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -lean_object* x_4; uint8_t x_5; -x_4 = lean_array_get_size(x_2); -x_5 = lean_nat_dec_lt(x_1, x_4); -lean_dec(x_4); -if (x_5 == 0) -{ -lean_dec(x_2); -lean_dec(x_1); -return x_3; -} -else -{ -lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; -x_6 = lean_array_fget(x_2, x_1); -x_7 = lean_box(0); -x_8 = lean_array_fset(x_2, x_1, x_7); -x_9 = l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__6___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__7(x_3, x_6); -x_10 = lean_unsigned_to_nat(1u); -x_11 = lean_nat_add(x_1, x_10); -lean_dec(x_1); -x_1 = x_11; -x_2 = x_8; -x_3 = x_9; -goto _start; -} -} -} -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__4(lean_object* x_1) { -_start: -{ -lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; -x_2 = lean_array_get_size(x_1); -x_3 = lean_unsigned_to_nat(2u); -x_4 = lean_nat_mul(x_2, x_3); -lean_dec(x_2); -x_5 = lean_box(0); -x_6 = lean_mk_array(x_4, x_5); -x_7 = lean_unsigned_to_nat(0u); -x_8 = l_Std_DHashMap_Internal_Raw_u2080_expand_go___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__5(x_7, x_1, x_6); -return x_8; -} -} -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__8(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -if (lean_obj_tag(x_3) == 0) -{ -lean_object* x_4; -lean_dec(x_2); -lean_dec(x_1); -x_4 = lean_box(0); -return x_4; -} -else -{ -uint8_t x_5; -x_5 = !lean_is_exclusive(x_3); -if (x_5 == 0) -{ -lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x_9; size_t x_10; uint8_t x_11; -x_6 = lean_ctor_get(x_3, 0); -x_7 = lean_ctor_get(x_3, 1); -x_8 = lean_ctor_get(x_3, 2); -x_9 = lean_ptr_addr(x_6); -x_10 = lean_ptr_addr(x_1); -x_11 = lean_usize_dec_eq(x_9, x_10); -if (x_11 == 0) -{ -lean_object* x_12; -x_12 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__8(x_1, x_2, x_8); -lean_ctor_set(x_3, 2, x_12); -return x_3; -} -else -{ -lean_dec(x_7); -lean_dec(x_6); -lean_ctor_set(x_3, 1, x_2); -lean_ctor_set(x_3, 0, x_1); -return x_3; -} -} -else -{ -lean_object* x_13; lean_object* x_14; lean_object* x_15; size_t x_16; size_t x_17; uint8_t x_18; -x_13 = lean_ctor_get(x_3, 0); -x_14 = lean_ctor_get(x_3, 1); -x_15 = lean_ctor_get(x_3, 2); -lean_inc(x_15); -lean_inc(x_14); -lean_inc(x_13); -lean_dec(x_3); -x_16 = lean_ptr_addr(x_13); -x_17 = lean_ptr_addr(x_1); -x_18 = lean_usize_dec_eq(x_16, x_17); -if (x_18 == 0) -{ -lean_object* x_19; lean_object* x_20; -x_19 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__8(x_1, x_2, x_15); -x_20 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_20, 0, x_13); -lean_ctor_set(x_20, 1, x_14); -lean_ctor_set(x_20, 2, x_19); -return x_20; -} -else -{ -lean_object* x_21; -lean_dec(x_14); -lean_dec(x_13); -x_21 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_21, 0, x_1); -lean_ctor_set(x_21, 1, x_2); -lean_ctor_set(x_21, 2, x_15); -return x_21; -} -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__9(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -switch (lean_obj_tag(x_2)) { -case 0: -{ -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; -lean_dec(x_4); -x_11 = lean_array_get_size(x_3); -x_12 = lean_unsigned_to_nat(0u); -x_13 = lean_unsigned_to_nat(1u); -x_14 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_14, 0, x_12); -lean_ctor_set(x_14, 1, x_11); -lean_ctor_set(x_14, 2, x_13); -x_15 = 0; -x_16 = lean_box(x_15); -x_17 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_17, 0, x_3); -lean_ctor_set(x_17, 1, x_16); -lean_inc(x_5); -x_18 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__2(x_14, x_14, x_17, x_12, lean_box(0), lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_14); -if (lean_obj_tag(x_18) == 0) -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_125; uint8_t x_126; -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_125 = lean_ctor_get(x_19, 1); -lean_inc(x_125); -x_126 = lean_unbox(x_125); -lean_dec(x_125); -if (x_126 == 0) -{ -lean_dec(x_19); -lean_dec(x_2); -lean_inc(x_1); -x_21 = x_1; -goto block_124; -} -else -{ -lean_object* x_127; lean_object* x_128; -x_127 = lean_ctor_get(x_19, 0); -lean_inc(x_127); -lean_dec(x_19); -x_128 = l_Lean_mkAppN(x_2, x_127); -lean_dec(x_127); -x_21 = x_128; -goto block_124; -} -block_124: -{ -lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; -x_22 = lean_st_ref_take(x_5, x_20); -x_23 = lean_ctor_get(x_22, 0); -lean_inc(x_23); -x_24 = lean_ctor_get(x_22, 1); -lean_inc(x_24); -lean_dec(x_22); -x_25 = !lean_is_exclusive(x_23); -if (x_25 == 0) -{ -lean_object* x_26; lean_object* x_27; lean_object* x_28; size_t x_29; uint64_t x_30; uint64_t x_31; uint64_t x_32; uint64_t x_33; uint64_t x_34; uint64_t x_35; uint64_t x_36; uint64_t x_37; uint64_t x_38; size_t x_39; size_t x_40; size_t x_41; size_t x_42; size_t x_43; lean_object* x_44; uint8_t x_45; -x_26 = lean_ctor_get(x_23, 0); -x_27 = lean_ctor_get(x_23, 1); -x_28 = lean_array_get_size(x_27); -x_29 = lean_ptr_addr(x_1); -x_30 = lean_usize_to_uint64(x_29); -x_31 = 11; -x_32 = lean_uint64_mix_hash(x_30, x_31); -x_33 = 32; -x_34 = lean_uint64_shift_right(x_32, x_33); -x_35 = lean_uint64_xor(x_32, x_34); -x_36 = 16; -x_37 = lean_uint64_shift_right(x_35, x_36); -x_38 = lean_uint64_xor(x_35, x_37); -x_39 = lean_uint64_to_usize(x_38); -x_40 = lean_usize_of_nat(x_28); -lean_dec(x_28); -x_41 = 1; -x_42 = lean_usize_sub(x_40, x_41); -x_43 = lean_usize_land(x_39, x_42); -x_44 = lean_array_uget(x_27, x_43); -x_45 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__3(x_1, x_44); -if (x_45 == 0) -{ -lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; uint8_t x_54; -x_46 = lean_nat_add(x_26, x_13); -lean_dec(x_26); -lean_inc(x_21); -x_47 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_47, 0, x_1); -lean_ctor_set(x_47, 1, x_21); -lean_ctor_set(x_47, 2, x_44); -x_48 = lean_array_uset(x_27, x_43, x_47); -x_49 = lean_unsigned_to_nat(4u); -x_50 = lean_nat_mul(x_46, x_49); -x_51 = lean_unsigned_to_nat(3u); -x_52 = lean_nat_div(x_50, x_51); -lean_dec(x_50); -x_53 = lean_array_get_size(x_48); -x_54 = lean_nat_dec_le(x_52, x_53); -lean_dec(x_53); -lean_dec(x_52); -if (x_54 == 0) -{ -lean_object* x_55; lean_object* x_56; uint8_t x_57; -x_55 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__4(x_48); -lean_ctor_set(x_23, 1, x_55); -lean_ctor_set(x_23, 0, x_46); -x_56 = lean_st_ref_set(x_5, x_23, x_24); -lean_dec(x_5); -x_57 = !lean_is_exclusive(x_56); -if (x_57 == 0) -{ -lean_object* x_58; -x_58 = lean_ctor_get(x_56, 0); -lean_dec(x_58); -lean_ctor_set(x_56, 0, x_21); -return x_56; -} -else -{ -lean_object* x_59; lean_object* x_60; -x_59 = lean_ctor_get(x_56, 1); -lean_inc(x_59); -lean_dec(x_56); -x_60 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_60, 0, x_21); -lean_ctor_set(x_60, 1, x_59); -return x_60; -} -} -else -{ -lean_object* x_61; uint8_t x_62; -lean_ctor_set(x_23, 1, x_48); -lean_ctor_set(x_23, 0, x_46); -x_61 = lean_st_ref_set(x_5, x_23, x_24); -lean_dec(x_5); -x_62 = !lean_is_exclusive(x_61); -if (x_62 == 0) -{ -lean_object* x_63; -x_63 = lean_ctor_get(x_61, 0); -lean_dec(x_63); -lean_ctor_set(x_61, 0, x_21); -return x_61; -} -else -{ -lean_object* x_64; lean_object* x_65; -x_64 = lean_ctor_get(x_61, 1); -lean_inc(x_64); -lean_dec(x_61); -x_65 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_65, 0, x_21); -lean_ctor_set(x_65, 1, x_64); -return x_65; -} -} -} -else -{ -lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; uint8_t x_71; -x_66 = lean_box(0); -x_67 = lean_array_uset(x_27, x_43, x_66); -lean_inc(x_21); -x_68 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__8(x_1, x_21, x_44); -x_69 = lean_array_uset(x_67, x_43, x_68); -lean_ctor_set(x_23, 1, x_69); -x_70 = lean_st_ref_set(x_5, x_23, x_24); -lean_dec(x_5); -x_71 = !lean_is_exclusive(x_70); -if (x_71 == 0) +if (lean_obj_tag(x_2) == 0) { -lean_object* x_72; -x_72 = lean_ctor_get(x_70, 0); -lean_dec(x_72); -lean_ctor_set(x_70, 0, x_21); -return x_70; +uint8_t x_3; +x_3 = 0; +return x_3; } 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(0, 2, 0); -lean_ctor_set(x_74, 0, x_21); -lean_ctor_set(x_74, 1, x_73); -return x_74; -} -} -} -else +lean_object* x_4; lean_object* x_5; size_t x_6; size_t x_7; uint8_t x_8; +x_4 = lean_ctor_get(x_2, 0); +x_5 = lean_ctor_get(x_2, 2); +x_6 = lean_ptr_addr(x_4); +x_7 = lean_ptr_addr(x_1); +x_8 = lean_usize_dec_eq(x_6, x_7); +if (x_8 == 0) { -lean_object* x_75; lean_object* x_76; lean_object* x_77; size_t x_78; uint64_t x_79; uint64_t x_80; uint64_t x_81; uint64_t x_82; uint64_t x_83; uint64_t x_84; uint64_t x_85; uint64_t x_86; uint64_t x_87; size_t x_88; size_t x_89; size_t x_90; size_t x_91; size_t x_92; lean_object* x_93; uint8_t x_94; -x_75 = lean_ctor_get(x_23, 0); -x_76 = lean_ctor_get(x_23, 1); -lean_inc(x_76); -lean_inc(x_75); -lean_dec(x_23); -x_77 = lean_array_get_size(x_76); -x_78 = lean_ptr_addr(x_1); -x_79 = lean_usize_to_uint64(x_78); -x_80 = 11; -x_81 = lean_uint64_mix_hash(x_79, x_80); -x_82 = 32; -x_83 = lean_uint64_shift_right(x_81, x_82); -x_84 = lean_uint64_xor(x_81, x_83); -x_85 = 16; -x_86 = lean_uint64_shift_right(x_84, x_85); -x_87 = lean_uint64_xor(x_84, x_86); -x_88 = lean_uint64_to_usize(x_87); -x_89 = lean_usize_of_nat(x_77); -lean_dec(x_77); -x_90 = 1; -x_91 = lean_usize_sub(x_89, x_90); -x_92 = lean_usize_land(x_88, x_91); -x_93 = lean_array_uget(x_76, x_92); -x_94 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__3(x_1, x_93); -if (x_94 == 0) -{ -lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; uint8_t x_103; -x_95 = lean_nat_add(x_75, x_13); -lean_dec(x_75); -lean_inc(x_21); -x_96 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_96, 0, x_1); -lean_ctor_set(x_96, 1, x_21); -lean_ctor_set(x_96, 2, x_93); -x_97 = lean_array_uset(x_76, x_92, x_96); -x_98 = lean_unsigned_to_nat(4u); -x_99 = lean_nat_mul(x_95, x_98); -x_100 = lean_unsigned_to_nat(3u); -x_101 = lean_nat_div(x_99, x_100); -lean_dec(x_99); -x_102 = lean_array_get_size(x_97); -x_103 = lean_nat_dec_le(x_101, x_102); -lean_dec(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; lean_object* x_108; lean_object* x_109; -x_104 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__4(x_97); -x_105 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_105, 0, x_95); -lean_ctor_set(x_105, 1, x_104); -x_106 = lean_st_ref_set(x_5, x_105, x_24); -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(0, 2, 0); -} else { - x_109 = x_108; -} -lean_ctor_set(x_109, 0, x_21); -lean_ctor_set(x_109, 1, x_107); -return x_109; +x_2 = x_5; +goto _start; } else { -lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; -x_110 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_110, 0, x_95); -lean_ctor_set(x_110, 1, x_97); -x_111 = lean_st_ref_set(x_5, x_110, x_24); -lean_dec(x_5); -x_112 = lean_ctor_get(x_111, 1); -lean_inc(x_112); -if (lean_is_exclusive(x_111)) { - lean_ctor_release(x_111, 0); - lean_ctor_release(x_111, 1); - x_113 = x_111; -} else { - lean_dec_ref(x_111); - x_113 = lean_box(0); +uint8_t x_10; +x_10 = 1; +return x_10; } -if (lean_is_scalar(x_113)) { - x_114 = lean_alloc_ctor(0, 2, 0); -} else { - x_114 = x_113; } -lean_ctor_set(x_114, 0, x_21); -lean_ctor_set(x_114, 1, x_112); -return x_114; } } -else +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -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_115 = lean_box(0); -x_116 = lean_array_uset(x_76, x_92, x_115); -lean_inc(x_21); -x_117 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__8(x_1, x_21, x_93); -x_118 = lean_array_uset(x_116, x_92, x_117); -x_119 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_119, 0, x_75); -lean_ctor_set(x_119, 1, x_118); -x_120 = lean_st_ref_set(x_5, x_119, x_24); -lean_dec(x_5); -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); -} -if (lean_is_scalar(x_122)) { - x_123 = lean_alloc_ctor(0, 2, 0); -} else { - x_123 = x_122; -} -lean_ctor_set(x_123, 0, x_21); -lean_ctor_set(x_123, 1, x_121); -return x_123; -} -} -} -} -else +if (lean_obj_tag(x_3) == 0) { -uint8_t x_129; -lean_dec(x_5); -lean_dec(x_2); lean_dec(x_1); -x_129 = !lean_is_exclusive(x_18); -if (x_129 == 0) -{ -return x_18; +return x_2; } else { -lean_object* x_130; lean_object* x_131; lean_object* x_132; -x_130 = lean_ctor_get(x_18, 0); -x_131 = lean_ctor_get(x_18, 1); -lean_inc(x_131); -lean_inc(x_130); -lean_dec(x_18); -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; -} -} -} -case 1: +uint8_t x_4; +x_4 = !lean_is_exclusive(x_3); +if (x_4 == 0) { -lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; uint8_t x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; -lean_dec(x_4); -x_133 = lean_array_get_size(x_3); -x_134 = lean_unsigned_to_nat(0u); -x_135 = lean_unsigned_to_nat(1u); -x_136 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_136, 0, x_134); -lean_ctor_set(x_136, 1, x_133); -lean_ctor_set(x_136, 2, x_135); -x_137 = 0; -x_138 = lean_box(x_137); -x_139 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_139, 0, x_3); -lean_ctor_set(x_139, 1, x_138); -lean_inc(x_5); -x_140 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__2(x_136, x_136, x_139, x_134, lean_box(0), lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_136); -if (lean_obj_tag(x_140) == 0) -{ -lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_247; uint8_t x_248; -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_247 = lean_ctor_get(x_141, 1); -lean_inc(x_247); -x_248 = lean_unbox(x_247); -lean_dec(x_247); -if (x_248 == 0) -{ -lean_dec(x_141); -lean_dec(x_2); +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint64_t 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; size_t x_16; size_t x_17; size_t x_18; size_t x_19; size_t x_20; lean_object* x_21; lean_object* x_22; +x_5 = lean_ctor_get(x_3, 0); +x_6 = lean_ctor_get(x_3, 2); +x_7 = lean_array_get_size(x_2); lean_inc(x_1); -x_143 = x_1; -goto block_246; -} -else -{ -lean_object* x_249; lean_object* x_250; -x_249 = lean_ctor_get(x_141, 0); -lean_inc(x_249); -lean_dec(x_141); -x_250 = l_Lean_mkAppN(x_2, x_249); -lean_dec(x_249); -x_143 = x_250; -goto block_246; -} -block_246: -{ -lean_object* x_144; lean_object* x_145; lean_object* x_146; uint8_t x_147; -x_144 = lean_st_ref_take(x_5, x_142); -x_145 = lean_ctor_get(x_144, 0); -lean_inc(x_145); -x_146 = lean_ctor_get(x_144, 1); -lean_inc(x_146); -lean_dec(x_144); -x_147 = !lean_is_exclusive(x_145); -if (x_147 == 0) -{ -lean_object* x_148; lean_object* x_149; lean_object* x_150; size_t x_151; uint64_t x_152; uint64_t x_153; uint64_t x_154; uint64_t x_155; uint64_t x_156; uint64_t x_157; uint64_t x_158; uint64_t x_159; uint64_t x_160; size_t x_161; size_t x_162; size_t x_163; size_t x_164; size_t x_165; lean_object* x_166; uint8_t x_167; -x_148 = lean_ctor_get(x_145, 0); -x_149 = lean_ctor_get(x_145, 1); -x_150 = lean_array_get_size(x_149); -x_151 = lean_ptr_addr(x_1); -x_152 = lean_usize_to_uint64(x_151); -x_153 = 11; -x_154 = lean_uint64_mix_hash(x_152, x_153); -x_155 = 32; -x_156 = lean_uint64_shift_right(x_154, x_155); -x_157 = lean_uint64_xor(x_154, x_156); -x_158 = 16; -x_159 = lean_uint64_shift_right(x_157, x_158); -x_160 = lean_uint64_xor(x_157, x_159); -x_161 = lean_uint64_to_usize(x_160); -x_162 = lean_usize_of_nat(x_150); -lean_dec(x_150); -x_163 = 1; -x_164 = lean_usize_sub(x_162, x_163); -x_165 = lean_usize_land(x_161, x_164); -x_166 = lean_array_uget(x_149, x_165); -x_167 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__3(x_1, x_166); -if (x_167 == 0) -{ -lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; uint8_t x_176; -x_168 = lean_nat_add(x_148, x_135); -lean_dec(x_148); -lean_inc(x_143); -x_169 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_169, 0, x_1); -lean_ctor_set(x_169, 1, x_143); -lean_ctor_set(x_169, 2, x_166); -x_170 = lean_array_uset(x_149, x_165, x_169); -x_171 = lean_unsigned_to_nat(4u); -x_172 = lean_nat_mul(x_168, x_171); -x_173 = lean_unsigned_to_nat(3u); -x_174 = lean_nat_div(x_172, x_173); -lean_dec(x_172); -x_175 = lean_array_get_size(x_170); -x_176 = lean_nat_dec_le(x_174, x_175); -lean_dec(x_175); -lean_dec(x_174); -if (x_176 == 0) -{ -lean_object* x_177; lean_object* x_178; uint8_t x_179; -x_177 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__4(x_170); -lean_ctor_set(x_145, 1, x_177); -lean_ctor_set(x_145, 0, x_168); -x_178 = lean_st_ref_set(x_5, x_145, x_146); -lean_dec(x_5); -x_179 = !lean_is_exclusive(x_178); -if (x_179 == 0) -{ -lean_object* x_180; -x_180 = lean_ctor_get(x_178, 0); -lean_dec(x_180); -lean_ctor_set(x_178, 0, x_143); -return x_178; -} -else -{ -lean_object* x_181; lean_object* x_182; -x_181 = lean_ctor_get(x_178, 1); -lean_inc(x_181); -lean_dec(x_178); -x_182 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_182, 0, x_143); -lean_ctor_set(x_182, 1, x_181); -return x_182; -} -} -else -{ -lean_object* x_183; uint8_t x_184; -lean_ctor_set(x_145, 1, x_170); -lean_ctor_set(x_145, 0, x_168); -x_183 = lean_st_ref_set(x_5, x_145, x_146); -lean_dec(x_5); -x_184 = !lean_is_exclusive(x_183); -if (x_184 == 0) -{ -lean_object* x_185; -x_185 = lean_ctor_get(x_183, 0); -lean_dec(x_185); -lean_ctor_set(x_183, 0, x_143); -return x_183; -} -else -{ -lean_object* x_186; lean_object* x_187; -x_186 = lean_ctor_get(x_183, 1); -lean_inc(x_186); -lean_dec(x_183); -x_187 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_187, 0, x_143); -lean_ctor_set(x_187, 1, x_186); -return x_187; -} -} -} -else -{ -lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; uint8_t x_193; -x_188 = lean_box(0); -x_189 = lean_array_uset(x_149, x_165, x_188); -lean_inc(x_143); -x_190 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__8(x_1, x_143, x_166); -x_191 = lean_array_uset(x_189, x_165, x_190); -lean_ctor_set(x_145, 1, x_191); -x_192 = lean_st_ref_set(x_5, x_145, x_146); -lean_dec(x_5); -x_193 = !lean_is_exclusive(x_192); -if (x_193 == 0) -{ -lean_object* x_194; -x_194 = lean_ctor_get(x_192, 0); -lean_dec(x_194); -lean_ctor_set(x_192, 0, x_143); -return x_192; -} -else -{ -lean_object* x_195; lean_object* x_196; -x_195 = lean_ctor_get(x_192, 1); -lean_inc(x_195); -lean_dec(x_192); -x_196 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_196, 0, x_143); -lean_ctor_set(x_196, 1, x_195); -return x_196; -} -} -} -else -{ -lean_object* x_197; lean_object* x_198; lean_object* x_199; size_t x_200; uint64_t x_201; uint64_t x_202; uint64_t x_203; uint64_t x_204; uint64_t x_205; uint64_t x_206; uint64_t x_207; uint64_t x_208; uint64_t x_209; size_t x_210; size_t x_211; size_t x_212; size_t x_213; size_t x_214; lean_object* x_215; uint8_t x_216; -x_197 = lean_ctor_get(x_145, 0); -x_198 = lean_ctor_get(x_145, 1); -lean_inc(x_198); -lean_inc(x_197); -lean_dec(x_145); -x_199 = lean_array_get_size(x_198); -x_200 = lean_ptr_addr(x_1); -x_201 = lean_usize_to_uint64(x_200); -x_202 = 11; -x_203 = lean_uint64_mix_hash(x_201, x_202); -x_204 = 32; -x_205 = lean_uint64_shift_right(x_203, x_204); -x_206 = lean_uint64_xor(x_203, x_205); -x_207 = 16; -x_208 = lean_uint64_shift_right(x_206, x_207); -x_209 = lean_uint64_xor(x_206, x_208); -x_210 = lean_uint64_to_usize(x_209); -x_211 = lean_usize_of_nat(x_199); -lean_dec(x_199); -x_212 = 1; -x_213 = lean_usize_sub(x_211, x_212); -x_214 = lean_usize_land(x_210, x_213); -x_215 = lean_array_uget(x_198, x_214); -x_216 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__3(x_1, x_215); -if (x_216 == 0) -{ -lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; uint8_t x_225; -x_217 = lean_nat_add(x_197, x_135); -lean_dec(x_197); -lean_inc(x_143); -x_218 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_218, 0, x_1); -lean_ctor_set(x_218, 1, x_143); -lean_ctor_set(x_218, 2, x_215); -x_219 = lean_array_uset(x_198, x_214, x_218); -x_220 = lean_unsigned_to_nat(4u); -x_221 = lean_nat_mul(x_217, x_220); -x_222 = lean_unsigned_to_nat(3u); -x_223 = lean_nat_div(x_221, x_222); -lean_dec(x_221); -x_224 = lean_array_get_size(x_219); -x_225 = lean_nat_dec_le(x_223, x_224); -lean_dec(x_224); -lean_dec(x_223); -if (x_225 == 0) -{ -lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; -x_226 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__4(x_219); -x_227 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_227, 0, x_217); -lean_ctor_set(x_227, 1, x_226); -x_228 = lean_st_ref_set(x_5, x_227, x_146); -lean_dec(x_5); -x_229 = lean_ctor_get(x_228, 1); -lean_inc(x_229); -if (lean_is_exclusive(x_228)) { - lean_ctor_release(x_228, 0); - lean_ctor_release(x_228, 1); - x_230 = x_228; -} else { - lean_dec_ref(x_228); - x_230 = lean_box(0); -} -if (lean_is_scalar(x_230)) { - x_231 = lean_alloc_ctor(0, 2, 0); -} else { - x_231 = x_230; -} -lean_ctor_set(x_231, 0, x_143); -lean_ctor_set(x_231, 1, x_229); -return x_231; -} -else -{ -lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; -x_232 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_232, 0, x_217); -lean_ctor_set(x_232, 1, x_219); -x_233 = lean_st_ref_set(x_5, x_232, x_146); -lean_dec(x_5); -x_234 = lean_ctor_get(x_233, 1); -lean_inc(x_234); -if (lean_is_exclusive(x_233)) { - lean_ctor_release(x_233, 0); - lean_ctor_release(x_233, 1); - x_235 = x_233; -} else { - lean_dec_ref(x_233); - x_235 = lean_box(0); -} -if (lean_is_scalar(x_235)) { - x_236 = lean_alloc_ctor(0, 2, 0); -} else { - x_236 = x_235; -} -lean_ctor_set(x_236, 0, x_143); -lean_ctor_set(x_236, 1, x_234); -return x_236; -} -} -else -{ -lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; -x_237 = lean_box(0); -x_238 = lean_array_uset(x_198, x_214, x_237); -lean_inc(x_143); -x_239 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__8(x_1, x_143, x_215); -x_240 = lean_array_uset(x_238, x_214, x_239); -x_241 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_241, 0, x_197); -lean_ctor_set(x_241, 1, x_240); -x_242 = lean_st_ref_set(x_5, x_241, x_146); -lean_dec(x_5); -x_243 = lean_ctor_get(x_242, 1); -lean_inc(x_243); -if (lean_is_exclusive(x_242)) { - lean_ctor_release(x_242, 0); - lean_ctor_release(x_242, 1); - x_244 = x_242; -} else { - lean_dec_ref(x_242); - x_244 = lean_box(0); -} -if (lean_is_scalar(x_244)) { - x_245 = lean_alloc_ctor(0, 2, 0); -} else { - x_245 = x_244; -} -lean_ctor_set(x_245, 0, x_143); -lean_ctor_set(x_245, 1, x_243); -return x_245; -} -} -} -} -else -{ -uint8_t x_251; -lean_dec(x_5); -lean_dec(x_2); -lean_dec(x_1); -x_251 = !lean_is_exclusive(x_140); -if (x_251 == 0) -{ -return x_140; -} -else -{ -lean_object* x_252; lean_object* x_253; lean_object* x_254; -x_252 = lean_ctor_get(x_140, 0); -x_253 = lean_ctor_get(x_140, 1); -lean_inc(x_253); -lean_inc(x_252); -lean_dec(x_140); -x_254 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_254, 0, x_252); -lean_ctor_set(x_254, 1, x_253); -return x_254; -} -} -} -case 2: -{ -lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; uint8_t x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; -lean_dec(x_4); -x_255 = lean_array_get_size(x_3); -x_256 = lean_unsigned_to_nat(0u); -x_257 = lean_unsigned_to_nat(1u); -x_258 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_258, 0, x_256); -lean_ctor_set(x_258, 1, x_255); -lean_ctor_set(x_258, 2, x_257); -x_259 = 0; -x_260 = lean_box(x_259); -x_261 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_261, 0, x_3); -lean_ctor_set(x_261, 1, x_260); lean_inc(x_5); -x_262 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__2(x_258, x_258, x_261, x_256, lean_box(0), lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_258); -if (lean_obj_tag(x_262) == 0) -{ -lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_369; uint8_t x_370; -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_369 = lean_ctor_get(x_263, 1); -lean_inc(x_369); -x_370 = lean_unbox(x_369); -lean_dec(x_369); -if (x_370 == 0) -{ -lean_dec(x_263); -lean_dec(x_2); -lean_inc(x_1); -x_265 = x_1; -goto block_368; -} -else -{ -lean_object* x_371; lean_object* x_372; -x_371 = lean_ctor_get(x_263, 0); -lean_inc(x_371); -lean_dec(x_263); -x_372 = l_Lean_mkAppN(x_2, x_371); -lean_dec(x_371); -x_265 = x_372; -goto block_368; -} -block_368: -{ -lean_object* x_266; lean_object* x_267; lean_object* x_268; uint8_t x_269; -x_266 = lean_st_ref_take(x_5, x_264); -x_267 = lean_ctor_get(x_266, 0); -lean_inc(x_267); -x_268 = lean_ctor_get(x_266, 1); -lean_inc(x_268); -lean_dec(x_266); -x_269 = !lean_is_exclusive(x_267); -if (x_269 == 0) -{ -lean_object* x_270; lean_object* x_271; lean_object* x_272; size_t x_273; uint64_t x_274; uint64_t x_275; uint64_t x_276; uint64_t x_277; uint64_t x_278; uint64_t x_279; uint64_t x_280; uint64_t x_281; uint64_t x_282; size_t x_283; size_t x_284; size_t x_285; size_t x_286; size_t x_287; lean_object* x_288; uint8_t x_289; -x_270 = lean_ctor_get(x_267, 0); -x_271 = lean_ctor_get(x_267, 1); -x_272 = lean_array_get_size(x_271); -x_273 = lean_ptr_addr(x_1); -x_274 = lean_usize_to_uint64(x_273); -x_275 = 11; -x_276 = lean_uint64_mix_hash(x_274, x_275); -x_277 = 32; -x_278 = lean_uint64_shift_right(x_276, x_277); -x_279 = lean_uint64_xor(x_276, x_278); -x_280 = 16; -x_281 = lean_uint64_shift_right(x_279, x_280); -x_282 = lean_uint64_xor(x_279, x_281); -x_283 = lean_uint64_to_usize(x_282); -x_284 = lean_usize_of_nat(x_272); -lean_dec(x_272); -x_285 = 1; -x_286 = lean_usize_sub(x_284, x_285); -x_287 = lean_usize_land(x_283, x_286); -x_288 = lean_array_uget(x_271, x_287); -x_289 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__3(x_1, x_288); -if (x_289 == 0) -{ -lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; uint8_t x_298; -x_290 = lean_nat_add(x_270, x_257); -lean_dec(x_270); -lean_inc(x_265); -x_291 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_291, 0, x_1); -lean_ctor_set(x_291, 1, x_265); -lean_ctor_set(x_291, 2, x_288); -x_292 = lean_array_uset(x_271, x_287, x_291); -x_293 = lean_unsigned_to_nat(4u); -x_294 = lean_nat_mul(x_290, x_293); -x_295 = lean_unsigned_to_nat(3u); -x_296 = lean_nat_div(x_294, x_295); -lean_dec(x_294); -x_297 = lean_array_get_size(x_292); -x_298 = lean_nat_dec_le(x_296, x_297); -lean_dec(x_297); -lean_dec(x_296); -if (x_298 == 0) -{ -lean_object* x_299; lean_object* x_300; uint8_t x_301; -x_299 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__4(x_292); -lean_ctor_set(x_267, 1, x_299); -lean_ctor_set(x_267, 0, x_290); -x_300 = lean_st_ref_set(x_5, x_267, x_268); -lean_dec(x_5); -x_301 = !lean_is_exclusive(x_300); -if (x_301 == 0) -{ -lean_object* x_302; -x_302 = lean_ctor_get(x_300, 0); -lean_dec(x_302); -lean_ctor_set(x_300, 0, x_265); -return x_300; -} -else -{ -lean_object* x_303; lean_object* x_304; -x_303 = lean_ctor_get(x_300, 1); -lean_inc(x_303); -lean_dec(x_300); -x_304 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_304, 0, x_265); -lean_ctor_set(x_304, 1, x_303); -return x_304; -} -} -else -{ -lean_object* x_305; uint8_t x_306; -lean_ctor_set(x_267, 1, x_292); -lean_ctor_set(x_267, 0, x_290); -x_305 = lean_st_ref_set(x_5, x_267, x_268); -lean_dec(x_5); -x_306 = !lean_is_exclusive(x_305); -if (x_306 == 0) -{ -lean_object* x_307; -x_307 = lean_ctor_get(x_305, 0); -lean_dec(x_307); -lean_ctor_set(x_305, 0, x_265); -return x_305; +x_8 = lean_apply_1(x_1, x_5); +x_9 = lean_unbox_uint64(x_8); +lean_dec(x_8); +x_10 = 32; +x_11 = lean_uint64_shift_right(x_9, x_10); +x_12 = lean_uint64_xor(x_9, x_11); +x_13 = 16; +x_14 = lean_uint64_shift_right(x_12, x_13); +x_15 = lean_uint64_xor(x_12, x_14); +x_16 = lean_uint64_to_usize(x_15); +x_17 = lean_usize_of_nat(x_7); +lean_dec(x_7); +x_18 = 1; +x_19 = lean_usize_sub(x_17, x_18); +x_20 = lean_usize_land(x_16, x_19); +x_21 = lean_array_uget(x_2, x_20); +lean_ctor_set(x_3, 2, x_21); +x_22 = lean_array_uset(x_2, x_20, x_3); +x_2 = x_22; +x_3 = x_6; +goto _start; } else { -lean_object* x_308; lean_object* x_309; -x_308 = lean_ctor_get(x_305, 1); -lean_inc(x_308); -lean_dec(x_305); -x_309 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_309, 0, x_265); -lean_ctor_set(x_309, 1, x_308); -return x_309; +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; uint64_t x_29; uint64_t x_30; uint64_t x_31; uint64_t x_32; uint64_t x_33; uint64_t x_34; uint64_t x_35; size_t x_36; size_t x_37; size_t x_38; size_t x_39; size_t x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_24 = lean_ctor_get(x_3, 0); +x_25 = lean_ctor_get(x_3, 1); +x_26 = lean_ctor_get(x_3, 2); +lean_inc(x_26); +lean_inc(x_25); +lean_inc(x_24); +lean_dec(x_3); +x_27 = lean_array_get_size(x_2); +lean_inc(x_1); +lean_inc(x_24); +x_28 = lean_apply_1(x_1, x_24); +x_29 = lean_unbox_uint64(x_28); +lean_dec(x_28); +x_30 = 32; +x_31 = lean_uint64_shift_right(x_29, x_30); +x_32 = lean_uint64_xor(x_29, x_31); +x_33 = 16; +x_34 = lean_uint64_shift_right(x_32, x_33); +x_35 = lean_uint64_xor(x_32, x_34); +x_36 = lean_uint64_to_usize(x_35); +x_37 = lean_usize_of_nat(x_27); +lean_dec(x_27); +x_38 = 1; +x_39 = lean_usize_sub(x_37, x_38); +x_40 = lean_usize_land(x_36, x_39); +x_41 = lean_array_uget(x_2, x_40); +x_42 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_42, 0, x_24); +lean_ctor_set(x_42, 1, x_25); +lean_ctor_set(x_42, 2, x_41); +x_43 = lean_array_uset(x_2, x_40, x_42); +x_2 = x_43; +x_3 = x_26; +goto _start; } } } -else +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__4___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__5(lean_object* x_1, lean_object* x_2) { +_start: { -lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; uint8_t x_315; -x_310 = lean_box(0); -x_311 = lean_array_uset(x_271, x_287, x_310); -lean_inc(x_265); -x_312 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__8(x_1, x_265, x_288); -x_313 = lean_array_uset(x_311, x_287, x_312); -lean_ctor_set(x_267, 1, x_313); -x_314 = lean_st_ref_set(x_5, x_267, x_268); -lean_dec(x_5); -x_315 = !lean_is_exclusive(x_314); -if (x_315 == 0) +if (lean_obj_tag(x_2) == 0) { -lean_object* x_316; -x_316 = lean_ctor_get(x_314, 0); -lean_dec(x_316); -lean_ctor_set(x_314, 0, x_265); -return x_314; +return x_1; } else { -lean_object* x_317; lean_object* x_318; -x_317 = lean_ctor_get(x_314, 1); -lean_inc(x_317); -lean_dec(x_314); -x_318 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_318, 0, x_265); -lean_ctor_set(x_318, 1, x_317); -return x_318; -} -} -} -else +uint8_t x_3; +x_3 = !lean_is_exclusive(x_2); +if (x_3 == 0) { -lean_object* x_319; lean_object* x_320; lean_object* x_321; size_t x_322; uint64_t x_323; uint64_t x_324; uint64_t x_325; uint64_t x_326; uint64_t x_327; uint64_t x_328; uint64_t x_329; uint64_t x_330; uint64_t x_331; size_t x_332; size_t x_333; size_t x_334; size_t x_335; size_t x_336; lean_object* x_337; uint8_t x_338; -x_319 = lean_ctor_get(x_267, 0); -x_320 = lean_ctor_get(x_267, 1); -lean_inc(x_320); -lean_inc(x_319); -lean_dec(x_267); -x_321 = lean_array_get_size(x_320); -x_322 = lean_ptr_addr(x_1); -x_323 = lean_usize_to_uint64(x_322); -x_324 = 11; -x_325 = lean_uint64_mix_hash(x_323, x_324); -x_326 = 32; -x_327 = lean_uint64_shift_right(x_325, x_326); -x_328 = lean_uint64_xor(x_325, x_327); -x_329 = 16; -x_330 = lean_uint64_shift_right(x_328, x_329); -x_331 = lean_uint64_xor(x_328, x_330); -x_332 = lean_uint64_to_usize(x_331); -x_333 = lean_usize_of_nat(x_321); -lean_dec(x_321); -x_334 = 1; -x_335 = lean_usize_sub(x_333, x_334); -x_336 = lean_usize_land(x_332, x_335); -x_337 = lean_array_uget(x_320, x_336); -x_338 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__3(x_1, x_337); -if (x_338 == 0) -{ -lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; uint8_t x_347; -x_339 = lean_nat_add(x_319, x_257); -lean_dec(x_319); -lean_inc(x_265); -x_340 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_340, 0, x_1); -lean_ctor_set(x_340, 1, x_265); -lean_ctor_set(x_340, 2, x_337); -x_341 = lean_array_uset(x_320, x_336, x_340); -x_342 = lean_unsigned_to_nat(4u); -x_343 = lean_nat_mul(x_339, x_342); -x_344 = lean_unsigned_to_nat(3u); -x_345 = lean_nat_div(x_343, x_344); -lean_dec(x_343); -x_346 = lean_array_get_size(x_341); -x_347 = lean_nat_dec_le(x_345, x_346); -lean_dec(x_346); -lean_dec(x_345); -if (x_347 == 0) -{ -lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; lean_object* x_353; -x_348 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__4(x_341); -x_349 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_349, 0, x_339); -lean_ctor_set(x_349, 1, x_348); -x_350 = lean_st_ref_set(x_5, x_349, x_268); -lean_dec(x_5); -x_351 = lean_ctor_get(x_350, 1); -lean_inc(x_351); -if (lean_is_exclusive(x_350)) { - lean_ctor_release(x_350, 0); - lean_ctor_release(x_350, 1); - x_352 = x_350; -} else { - lean_dec_ref(x_350); - x_352 = lean_box(0); -} -if (lean_is_scalar(x_352)) { - x_353 = lean_alloc_ctor(0, 2, 0); -} else { - x_353 = x_352; -} -lean_ctor_set(x_353, 0, x_265); -lean_ctor_set(x_353, 1, x_351); -return x_353; +lean_object* x_4; lean_object* x_5; lean_object* x_6; size_t x_7; uint64_t x_8; uint64_t 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; size_t x_17; size_t x_18; size_t x_19; size_t x_20; size_t x_21; lean_object* x_22; lean_object* x_23; +x_4 = lean_ctor_get(x_2, 0); +x_5 = lean_ctor_get(x_2, 2); +x_6 = lean_array_get_size(x_1); +x_7 = lean_ptr_addr(x_4); +x_8 = lean_usize_to_uint64(x_7); +x_9 = 11; +x_10 = lean_uint64_mix_hash(x_8, x_9); +x_11 = 32; +x_12 = lean_uint64_shift_right(x_10, x_11); +x_13 = lean_uint64_xor(x_10, x_12); +x_14 = 16; +x_15 = lean_uint64_shift_right(x_13, x_14); +x_16 = lean_uint64_xor(x_13, x_15); +x_17 = lean_uint64_to_usize(x_16); +x_18 = lean_usize_of_nat(x_6); +lean_dec(x_6); +x_19 = 1; +x_20 = lean_usize_sub(x_18, x_19); +x_21 = lean_usize_land(x_17, x_20); +x_22 = lean_array_uget(x_1, x_21); +lean_ctor_set(x_2, 2, x_22); +x_23 = lean_array_uset(x_1, x_21, x_2); +x_1 = x_23; +x_2 = x_5; +goto _start; } else { -lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; -x_354 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_354, 0, x_339); -lean_ctor_set(x_354, 1, x_341); -x_355 = lean_st_ref_set(x_5, x_354, x_268); -lean_dec(x_5); -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; -} else { - lean_dec_ref(x_355); - x_357 = lean_box(0); +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; size_t x_29; uint64_t x_30; uint64_t x_31; uint64_t x_32; uint64_t x_33; uint64_t x_34; uint64_t x_35; uint64_t x_36; uint64_t x_37; uint64_t x_38; size_t x_39; size_t x_40; size_t x_41; size_t x_42; size_t x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_25 = lean_ctor_get(x_2, 0); +x_26 = lean_ctor_get(x_2, 1); +x_27 = lean_ctor_get(x_2, 2); +lean_inc(x_27); +lean_inc(x_26); +lean_inc(x_25); +lean_dec(x_2); +x_28 = lean_array_get_size(x_1); +x_29 = lean_ptr_addr(x_25); +x_30 = lean_usize_to_uint64(x_29); +x_31 = 11; +x_32 = lean_uint64_mix_hash(x_30, x_31); +x_33 = 32; +x_34 = lean_uint64_shift_right(x_32, x_33); +x_35 = lean_uint64_xor(x_32, x_34); +x_36 = 16; +x_37 = lean_uint64_shift_right(x_35, x_36); +x_38 = lean_uint64_xor(x_35, x_37); +x_39 = lean_uint64_to_usize(x_38); +x_40 = lean_usize_of_nat(x_28); +lean_dec(x_28); +x_41 = 1; +x_42 = lean_usize_sub(x_40, x_41); +x_43 = lean_usize_land(x_39, x_42); +x_44 = lean_array_uget(x_1, x_43); +x_45 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_45, 0, x_25); +lean_ctor_set(x_45, 1, x_26); +lean_ctor_set(x_45, 2, x_44); +x_46 = lean_array_uset(x_1, x_43, x_45); +x_1 = x_46; +x_2 = x_27; +goto _start; } -if (lean_is_scalar(x_357)) { - x_358 = lean_alloc_ctor(0, 2, 0); -} else { - x_358 = x_357; } -lean_ctor_set(x_358, 0, x_265); -lean_ctor_set(x_358, 1, x_356); -return x_358; } } -else +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand_go___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -lean_object* x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; -x_359 = lean_box(0); -x_360 = lean_array_uset(x_320, x_336, x_359); -lean_inc(x_265); -x_361 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__8(x_1, x_265, x_337); -x_362 = lean_array_uset(x_360, x_336, x_361); -x_363 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_363, 0, x_319); -lean_ctor_set(x_363, 1, x_362); -x_364 = lean_st_ref_set(x_5, x_363, x_268); -lean_dec(x_5); -x_365 = lean_ctor_get(x_364, 1); -lean_inc(x_365); -if (lean_is_exclusive(x_364)) { - lean_ctor_release(x_364, 0); - lean_ctor_release(x_364, 1); - x_366 = x_364; -} else { - lean_dec_ref(x_364); - x_366 = lean_box(0); -} -if (lean_is_scalar(x_366)) { - x_367 = lean_alloc_ctor(0, 2, 0); -} else { - x_367 = x_366; -} -lean_ctor_set(x_367, 0, x_265); -lean_ctor_set(x_367, 1, x_365); -return x_367; -} -} -} -} -else +lean_object* x_4; uint8_t x_5; +x_4 = lean_array_get_size(x_2); +x_5 = lean_nat_dec_lt(x_1, x_4); +lean_dec(x_4); +if (x_5 == 0) { -uint8_t x_373; -lean_dec(x_5); lean_dec(x_2); lean_dec(x_1); -x_373 = !lean_is_exclusive(x_262); -if (x_373 == 0) -{ -return x_262; +return x_3; } else { -lean_object* x_374; lean_object* x_375; lean_object* x_376; -x_374 = lean_ctor_get(x_262, 0); -x_375 = lean_ctor_get(x_262, 1); -lean_inc(x_375); -lean_inc(x_374); -lean_dec(x_262); -x_376 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_376, 0, x_374); -lean_ctor_set(x_376, 1, x_375); -return x_376; +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_6 = lean_array_fget(x_2, x_1); +x_7 = lean_box(0); +x_8 = lean_array_fset(x_2, x_1, x_7); +x_9 = l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__4___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__5(x_3, x_6); +x_10 = lean_unsigned_to_nat(1u); +x_11 = lean_nat_add(x_1, x_10); +lean_dec(x_1); +x_1 = x_11; +x_2 = x_8; +x_3 = x_9; +goto _start; } } } -case 3: +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__2(lean_object* x_1) { +_start: { -lean_object* x_377; lean_object* x_378; lean_object* x_379; lean_object* x_380; uint8_t x_381; lean_object* x_382; lean_object* x_383; lean_object* x_384; -lean_dec(x_4); -x_377 = lean_array_get_size(x_3); -x_378 = lean_unsigned_to_nat(0u); -x_379 = lean_unsigned_to_nat(1u); -x_380 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_380, 0, x_378); -lean_ctor_set(x_380, 1, x_377); -lean_ctor_set(x_380, 2, x_379); -x_381 = 0; -x_382 = lean_box(x_381); -x_383 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_383, 0, x_3); -lean_ctor_set(x_383, 1, x_382); -lean_inc(x_5); -x_384 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__2(x_380, x_380, x_383, x_378, lean_box(0), lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_380); -if (lean_obj_tag(x_384) == 0) -{ -lean_object* x_385; lean_object* x_386; lean_object* x_387; lean_object* x_491; uint8_t x_492; -x_385 = lean_ctor_get(x_384, 0); -lean_inc(x_385); -x_386 = lean_ctor_get(x_384, 1); -lean_inc(x_386); -lean_dec(x_384); -x_491 = lean_ctor_get(x_385, 1); -lean_inc(x_491); -x_492 = lean_unbox(x_491); -lean_dec(x_491); -if (x_492 == 0) -{ -lean_dec(x_385); +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_2 = lean_array_get_size(x_1); +x_3 = lean_unsigned_to_nat(2u); +x_4 = lean_nat_mul(x_2, x_3); lean_dec(x_2); -lean_inc(x_1); -x_387 = x_1; -goto block_490; -} -else -{ -lean_object* x_493; lean_object* x_494; -x_493 = lean_ctor_get(x_385, 0); -lean_inc(x_493); -lean_dec(x_385); -x_494 = l_Lean_mkAppN(x_2, x_493); -lean_dec(x_493); -x_387 = x_494; -goto block_490; -} -block_490: -{ -lean_object* x_388; lean_object* x_389; lean_object* x_390; uint8_t x_391; -x_388 = lean_st_ref_take(x_5, x_386); -x_389 = lean_ctor_get(x_388, 0); -lean_inc(x_389); -x_390 = lean_ctor_get(x_388, 1); -lean_inc(x_390); -lean_dec(x_388); -x_391 = !lean_is_exclusive(x_389); -if (x_391 == 0) -{ -lean_object* x_392; lean_object* x_393; lean_object* x_394; size_t x_395; uint64_t x_396; uint64_t x_397; uint64_t x_398; uint64_t x_399; uint64_t x_400; uint64_t x_401; uint64_t x_402; uint64_t x_403; uint64_t x_404; size_t x_405; size_t x_406; size_t x_407; size_t x_408; size_t x_409; lean_object* x_410; uint8_t x_411; -x_392 = lean_ctor_get(x_389, 0); -x_393 = lean_ctor_get(x_389, 1); -x_394 = lean_array_get_size(x_393); -x_395 = lean_ptr_addr(x_1); -x_396 = lean_usize_to_uint64(x_395); -x_397 = 11; -x_398 = lean_uint64_mix_hash(x_396, x_397); -x_399 = 32; -x_400 = lean_uint64_shift_right(x_398, x_399); -x_401 = lean_uint64_xor(x_398, x_400); -x_402 = 16; -x_403 = lean_uint64_shift_right(x_401, x_402); -x_404 = lean_uint64_xor(x_401, x_403); -x_405 = lean_uint64_to_usize(x_404); -x_406 = lean_usize_of_nat(x_394); -lean_dec(x_394); -x_407 = 1; -x_408 = lean_usize_sub(x_406, x_407); -x_409 = lean_usize_land(x_405, x_408); -x_410 = lean_array_uget(x_393, x_409); -x_411 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__3(x_1, x_410); -if (x_411 == 0) -{ -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; uint8_t x_420; -x_412 = lean_nat_add(x_392, x_379); -lean_dec(x_392); -lean_inc(x_387); -x_413 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_413, 0, x_1); -lean_ctor_set(x_413, 1, x_387); -lean_ctor_set(x_413, 2, x_410); -x_414 = lean_array_uset(x_393, x_409, x_413); -x_415 = lean_unsigned_to_nat(4u); -x_416 = lean_nat_mul(x_412, x_415); -x_417 = lean_unsigned_to_nat(3u); -x_418 = lean_nat_div(x_416, x_417); -lean_dec(x_416); -x_419 = lean_array_get_size(x_414); -x_420 = lean_nat_dec_le(x_418, x_419); -lean_dec(x_419); -lean_dec(x_418); -if (x_420 == 0) -{ -lean_object* x_421; lean_object* x_422; uint8_t x_423; -x_421 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__4(x_414); -lean_ctor_set(x_389, 1, x_421); -lean_ctor_set(x_389, 0, x_412); -x_422 = lean_st_ref_set(x_5, x_389, x_390); -lean_dec(x_5); -x_423 = !lean_is_exclusive(x_422); -if (x_423 == 0) -{ -lean_object* x_424; -x_424 = lean_ctor_get(x_422, 0); -lean_dec(x_424); -lean_ctor_set(x_422, 0, x_387); -return x_422; -} -else -{ -lean_object* x_425; lean_object* x_426; -x_425 = lean_ctor_get(x_422, 1); -lean_inc(x_425); -lean_dec(x_422); -x_426 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_426, 0, x_387); -lean_ctor_set(x_426, 1, x_425); -return x_426; +x_5 = lean_box(0); +x_6 = lean_mk_array(x_4, x_5); +x_7 = lean_unsigned_to_nat(0u); +x_8 = l_Std_DHashMap_Internal_Raw_u2080_expand_go___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__3(x_7, x_1, x_6); +return x_8; } } -else +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -lean_object* x_427; uint8_t x_428; -lean_ctor_set(x_389, 1, x_414); -lean_ctor_set(x_389, 0, x_412); -x_427 = lean_st_ref_set(x_5, x_389, x_390); -lean_dec(x_5); -x_428 = !lean_is_exclusive(x_427); -if (x_428 == 0) +if (lean_obj_tag(x_3) == 0) { -lean_object* x_429; -x_429 = lean_ctor_get(x_427, 0); -lean_dec(x_429); -lean_ctor_set(x_427, 0, x_387); -return x_427; +lean_object* x_4; +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(0); +return x_4; } else { -lean_object* x_430; lean_object* x_431; -x_430 = lean_ctor_get(x_427, 1); -lean_inc(x_430); -lean_dec(x_427); -x_431 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_431, 0, x_387); -lean_ctor_set(x_431, 1, x_430); -return x_431; -} -} -} -else +uint8_t x_5; +x_5 = !lean_is_exclusive(x_3); +if (x_5 == 0) { -lean_object* x_432; lean_object* x_433; lean_object* x_434; lean_object* x_435; lean_object* x_436; uint8_t x_437; -x_432 = lean_box(0); -x_433 = lean_array_uset(x_393, x_409, x_432); -lean_inc(x_387); -x_434 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__8(x_1, x_387, x_410); -x_435 = lean_array_uset(x_433, x_409, x_434); -lean_ctor_set(x_389, 1, x_435); -x_436 = lean_st_ref_set(x_5, x_389, x_390); -lean_dec(x_5); -x_437 = !lean_is_exclusive(x_436); -if (x_437 == 0) +lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x_9; size_t x_10; uint8_t x_11; +x_6 = lean_ctor_get(x_3, 0); +x_7 = lean_ctor_get(x_3, 1); +x_8 = lean_ctor_get(x_3, 2); +x_9 = lean_ptr_addr(x_6); +x_10 = lean_ptr_addr(x_1); +x_11 = lean_usize_dec_eq(x_9, x_10); +if (x_11 == 0) { -lean_object* x_438; -x_438 = lean_ctor_get(x_436, 0); -lean_dec(x_438); -lean_ctor_set(x_436, 0, x_387); -return x_436; +lean_object* x_12; +x_12 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__6(x_1, x_2, x_8); +lean_ctor_set(x_3, 2, x_12); +return x_3; } else { -lean_object* x_439; lean_object* x_440; -x_439 = lean_ctor_get(x_436, 1); -lean_inc(x_439); -lean_dec(x_436); -x_440 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_440, 0, x_387); -lean_ctor_set(x_440, 1, x_439); -return x_440; -} +lean_dec(x_7); +lean_dec(x_6); +lean_ctor_set(x_3, 1, x_2); +lean_ctor_set(x_3, 0, x_1); +return x_3; } } else { -lean_object* x_441; lean_object* x_442; lean_object* x_443; size_t x_444; uint64_t x_445; uint64_t x_446; uint64_t x_447; uint64_t x_448; uint64_t x_449; uint64_t x_450; uint64_t x_451; uint64_t x_452; uint64_t x_453; size_t x_454; size_t x_455; size_t x_456; size_t x_457; size_t x_458; lean_object* x_459; uint8_t x_460; -x_441 = lean_ctor_get(x_389, 0); -x_442 = lean_ctor_get(x_389, 1); -lean_inc(x_442); -lean_inc(x_441); -lean_dec(x_389); -x_443 = lean_array_get_size(x_442); -x_444 = lean_ptr_addr(x_1); -x_445 = lean_usize_to_uint64(x_444); -x_446 = 11; -x_447 = lean_uint64_mix_hash(x_445, x_446); -x_448 = 32; -x_449 = lean_uint64_shift_right(x_447, x_448); -x_450 = lean_uint64_xor(x_447, x_449); -x_451 = 16; -x_452 = lean_uint64_shift_right(x_450, x_451); -x_453 = lean_uint64_xor(x_450, x_452); -x_454 = lean_uint64_to_usize(x_453); -x_455 = lean_usize_of_nat(x_443); -lean_dec(x_443); -x_456 = 1; -x_457 = lean_usize_sub(x_455, x_456); -x_458 = lean_usize_land(x_454, x_457); -x_459 = lean_array_uget(x_442, x_458); -x_460 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__3(x_1, x_459); -if (x_460 == 0) -{ -lean_object* x_461; lean_object* x_462; lean_object* x_463; lean_object* x_464; lean_object* x_465; lean_object* x_466; lean_object* x_467; lean_object* x_468; uint8_t x_469; -x_461 = lean_nat_add(x_441, x_379); -lean_dec(x_441); -lean_inc(x_387); -x_462 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_462, 0, x_1); -lean_ctor_set(x_462, 1, x_387); -lean_ctor_set(x_462, 2, x_459); -x_463 = lean_array_uset(x_442, x_458, x_462); -x_464 = lean_unsigned_to_nat(4u); -x_465 = lean_nat_mul(x_461, x_464); -x_466 = lean_unsigned_to_nat(3u); -x_467 = lean_nat_div(x_465, x_466); -lean_dec(x_465); -x_468 = lean_array_get_size(x_463); -x_469 = lean_nat_dec_le(x_467, x_468); -lean_dec(x_468); -lean_dec(x_467); -if (x_469 == 0) -{ -lean_object* x_470; lean_object* x_471; lean_object* x_472; lean_object* x_473; lean_object* x_474; lean_object* x_475; -x_470 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__4(x_463); -x_471 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_471, 0, x_461); -lean_ctor_set(x_471, 1, x_470); -x_472 = lean_st_ref_set(x_5, x_471, x_390); -lean_dec(x_5); -x_473 = lean_ctor_get(x_472, 1); -lean_inc(x_473); -if (lean_is_exclusive(x_472)) { - lean_ctor_release(x_472, 0); - lean_ctor_release(x_472, 1); - x_474 = x_472; -} else { - lean_dec_ref(x_472); - x_474 = lean_box(0); -} -if (lean_is_scalar(x_474)) { - x_475 = lean_alloc_ctor(0, 2, 0); -} else { - x_475 = x_474; -} -lean_ctor_set(x_475, 0, x_387); -lean_ctor_set(x_475, 1, x_473); -return x_475; -} -else +lean_object* x_13; lean_object* x_14; lean_object* x_15; size_t x_16; size_t x_17; uint8_t x_18; +x_13 = lean_ctor_get(x_3, 0); +x_14 = lean_ctor_get(x_3, 1); +x_15 = lean_ctor_get(x_3, 2); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_dec(x_3); +x_16 = lean_ptr_addr(x_13); +x_17 = lean_ptr_addr(x_1); +x_18 = lean_usize_dec_eq(x_16, x_17); +if (x_18 == 0) { -lean_object* x_476; lean_object* x_477; lean_object* x_478; lean_object* x_479; lean_object* x_480; -x_476 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_476, 0, x_461); -lean_ctor_set(x_476, 1, x_463); -x_477 = lean_st_ref_set(x_5, x_476, x_390); -lean_dec(x_5); -x_478 = lean_ctor_get(x_477, 1); -lean_inc(x_478); -if (lean_is_exclusive(x_477)) { - lean_ctor_release(x_477, 0); - lean_ctor_release(x_477, 1); - x_479 = x_477; -} else { - lean_dec_ref(x_477); - x_479 = lean_box(0); -} -if (lean_is_scalar(x_479)) { - x_480 = lean_alloc_ctor(0, 2, 0); -} else { - x_480 = x_479; -} -lean_ctor_set(x_480, 0, x_387); -lean_ctor_set(x_480, 1, x_478); -return x_480; -} +lean_object* x_19; lean_object* x_20; +x_19 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__6(x_1, x_2, x_15); +x_20 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_20, 0, x_13); +lean_ctor_set(x_20, 1, x_14); +lean_ctor_set(x_20, 2, x_19); +return x_20; } else { -lean_object* x_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; -x_481 = lean_box(0); -x_482 = lean_array_uset(x_442, x_458, x_481); -lean_inc(x_387); -x_483 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__8(x_1, x_387, x_459); -x_484 = lean_array_uset(x_482, x_458, x_483); -x_485 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_485, 0, x_441); -lean_ctor_set(x_485, 1, x_484); -x_486 = lean_st_ref_set(x_5, x_485, x_390); -lean_dec(x_5); -x_487 = lean_ctor_get(x_486, 1); -lean_inc(x_487); -if (lean_is_exclusive(x_486)) { - lean_ctor_release(x_486, 0); - lean_ctor_release(x_486, 1); - x_488 = x_486; -} else { - lean_dec_ref(x_486); - x_488 = lean_box(0); -} -if (lean_is_scalar(x_488)) { - x_489 = lean_alloc_ctor(0, 2, 0); -} else { - x_489 = x_488; +lean_object* x_21; +lean_dec(x_14); +lean_dec(x_13); +x_21 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_21, 0, x_1); +lean_ctor_set(x_21, 1, x_2); +lean_ctor_set(x_21, 2, x_15); +return x_21; } -lean_ctor_set(x_489, 0, x_387); -lean_ctor_set(x_489, 1, x_487); -return x_489; } } } } -else -{ -uint8_t x_495; -lean_dec(x_5); -lean_dec(x_2); -lean_dec(x_1); -x_495 = !lean_is_exclusive(x_384); -if (x_495 == 0) +static lean_object* _init_l_panic___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__7___closed__1() { +_start: { -return x_384; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_instMonadMetaM; +x_2 = l_ReaderT_instMonad___rarg(x_1); +return x_2; +} } -else +static lean_object* _init_l_panic___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__7___closed__2() { +_start: { -lean_object* x_496; lean_object* x_497; lean_object* x_498; -x_496 = lean_ctor_get(x_384, 0); -x_497 = lean_ctor_get(x_384, 1); -lean_inc(x_497); -lean_inc(x_496); -lean_dec(x_384); -x_498 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_498, 0, x_496); -lean_ctor_set(x_498, 1, x_497); -return x_498; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_panic___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__7___closed__1; +x_2 = l_Lean_instInhabitedExpr; +x_3 = l_instInhabitedOfMonad___rarg(x_1, x_2); +return x_3; } } +LEAN_EXPORT lean_object* l_panic___at_Lean_Meta_Grind_markNestedProofsImpl_visit___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; +x_8 = l_panic___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__7___closed__2; +x_9 = lean_panic_fn(x_8, x_1); +x_10 = lean_apply_6(x_9, x_2, x_3, x_4, x_5, x_6, x_7); +return x_10; +} } -case 4: +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_markNestedProofsImpl_visit___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: +{ +lean_object* x_13; uint8_t x_14; +x_13 = lean_ctor_get(x_2, 1); +x_14 = lean_nat_dec_lt(x_4, x_13); +if (x_14 == 0) { -lean_object* x_499; lean_object* x_500; lean_object* x_501; lean_object* x_502; uint8_t x_503; lean_object* x_504; lean_object* x_505; lean_object* x_506; +lean_object* x_15; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); lean_dec(x_4); -x_499 = lean_array_get_size(x_3); -x_500 = lean_unsigned_to_nat(0u); -x_501 = lean_unsigned_to_nat(1u); -x_502 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_502, 0, x_500); -lean_ctor_set(x_502, 1, x_499); -lean_ctor_set(x_502, 2, x_501); -x_503 = 0; -x_504 = lean_box(x_503); -x_505 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_505, 0, x_3); -lean_ctor_set(x_505, 1, x_504); -lean_inc(x_5); -x_506 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__2(x_502, x_502, x_505, x_500, lean_box(0), lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_502); -if (lean_obj_tag(x_506) == 0) -{ -lean_object* x_507; lean_object* x_508; lean_object* x_509; lean_object* x_613; uint8_t x_614; -x_507 = lean_ctor_get(x_506, 0); -lean_inc(x_507); -x_508 = lean_ctor_get(x_506, 1); -lean_inc(x_508); -lean_dec(x_506); -x_613 = lean_ctor_get(x_507, 1); -lean_inc(x_613); -x_614 = lean_unbox(x_613); -lean_dec(x_613); -if (x_614 == 0) -{ -lean_dec(x_507); -lean_dec(x_2); -lean_inc(x_1); -x_509 = x_1; -goto block_612; +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_3); +lean_ctor_set(x_15, 1, x_12); +return x_15; } else { -lean_object* x_615; lean_object* x_616; -x_615 = lean_ctor_get(x_507, 0); -lean_inc(x_615); -lean_dec(x_507); -x_616 = l_Lean_mkAppN(x_2, x_615); -lean_dec(x_615); -x_509 = x_616; -goto block_612; -} -block_612: -{ -lean_object* x_510; lean_object* x_511; lean_object* x_512; uint8_t x_513; -x_510 = lean_st_ref_take(x_5, x_508); -x_511 = lean_ctor_get(x_510, 0); -lean_inc(x_511); -x_512 = lean_ctor_get(x_510, 1); -lean_inc(x_512); -lean_dec(x_510); -x_513 = !lean_is_exclusive(x_511); -if (x_513 == 0) -{ -lean_object* x_514; lean_object* x_515; lean_object* x_516; size_t x_517; uint64_t x_518; uint64_t x_519; uint64_t x_520; uint64_t x_521; uint64_t x_522; uint64_t x_523; uint64_t x_524; uint64_t x_525; uint64_t x_526; size_t x_527; size_t x_528; size_t x_529; size_t x_530; size_t x_531; lean_object* x_532; uint8_t x_533; -x_514 = lean_ctor_get(x_511, 0); -x_515 = lean_ctor_get(x_511, 1); -x_516 = lean_array_get_size(x_515); -x_517 = lean_ptr_addr(x_1); -x_518 = lean_usize_to_uint64(x_517); -x_519 = 11; -x_520 = lean_uint64_mix_hash(x_518, x_519); -x_521 = 32; -x_522 = lean_uint64_shift_right(x_520, x_521); -x_523 = lean_uint64_xor(x_520, x_522); -x_524 = 16; -x_525 = lean_uint64_shift_right(x_523, x_524); -x_526 = lean_uint64_xor(x_523, x_525); -x_527 = lean_uint64_to_usize(x_526); -x_528 = lean_usize_of_nat(x_516); -lean_dec(x_516); -x_529 = 1; -x_530 = lean_usize_sub(x_528, x_529); -x_531 = lean_usize_land(x_527, x_530); -x_532 = lean_array_uget(x_515, x_531); -x_533 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__3(x_1, x_532); -if (x_533 == 0) -{ -lean_object* x_534; lean_object* x_535; lean_object* x_536; lean_object* x_537; lean_object* x_538; lean_object* x_539; lean_object* x_540; lean_object* x_541; uint8_t x_542; -x_534 = lean_nat_add(x_514, x_501); -lean_dec(x_514); -lean_inc(x_509); -x_535 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_535, 0, x_1); -lean_ctor_set(x_535, 1, x_509); -lean_ctor_set(x_535, 2, x_532); -x_536 = lean_array_uset(x_515, x_531, x_535); -x_537 = lean_unsigned_to_nat(4u); -x_538 = lean_nat_mul(x_534, x_537); -x_539 = lean_unsigned_to_nat(3u); -x_540 = lean_nat_div(x_538, x_539); -lean_dec(x_538); -x_541 = lean_array_get_size(x_536); -x_542 = lean_nat_dec_le(x_540, x_541); -lean_dec(x_541); -lean_dec(x_540); -if (x_542 == 0) -{ -lean_object* x_543; lean_object* x_544; uint8_t x_545; -x_543 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__4(x_536); -lean_ctor_set(x_511, 1, x_543); -lean_ctor_set(x_511, 0, x_534); -x_544 = lean_st_ref_set(x_5, x_511, x_512); -lean_dec(x_5); -x_545 = !lean_is_exclusive(x_544); -if (x_545 == 0) +lean_object* x_16; lean_object* x_17; uint8_t x_23; +x_23 = !lean_is_exclusive(x_3); +if (x_23 == 0) { -lean_object* x_546; -x_546 = lean_ctor_get(x_544, 0); -lean_dec(x_546); -lean_ctor_set(x_544, 0, x_509); -return x_544; -} -else +lean_object* x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; +x_24 = lean_ctor_get(x_3, 0); +x_25 = lean_ctor_get(x_3, 1); +x_26 = lean_array_get_size(x_24); +x_27 = lean_nat_dec_lt(x_4, x_26); +lean_dec(x_26); +if (x_27 == 0) { -lean_object* x_547; lean_object* x_548; -x_547 = lean_ctor_get(x_544, 1); -lean_inc(x_547); -lean_dec(x_544); -x_548 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_548, 0, x_509); -lean_ctor_set(x_548, 1, x_547); -return x_548; -} -} -else +lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_28 = l_Lean_instInhabitedExpr; +x_29 = l_outOfBounds___rarg(x_28); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_29); +x_30 = l_Lean_Meta_Grind_markNestedProofsImpl_visit(x_29, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_30) == 0) { -lean_object* x_549; uint8_t x_550; -lean_ctor_set(x_511, 1, x_536); -lean_ctor_set(x_511, 0, x_534); -x_549 = lean_st_ref_set(x_5, x_511, x_512); -lean_dec(x_5); -x_550 = !lean_is_exclusive(x_549); -if (x_550 == 0) +lean_object* x_31; lean_object* x_32; size_t x_33; size_t x_34; uint8_t x_35; +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_ptr_addr(x_29); +lean_dec(x_29); +x_34 = lean_ptr_addr(x_31); +x_35 = lean_usize_dec_eq(x_33, x_34); +if (x_35 == 0) { -lean_object* x_551; -x_551 = lean_ctor_get(x_549, 0); -lean_dec(x_551); -lean_ctor_set(x_549, 0, x_509); -return x_549; +lean_object* x_36; uint8_t x_37; lean_object* x_38; lean_object* x_39; +lean_dec(x_25); +x_36 = lean_array_set(x_24, x_4, x_31); +x_37 = 1; +x_38 = lean_box(x_37); +lean_ctor_set(x_3, 1, x_38); +lean_ctor_set(x_3, 0, x_36); +x_39 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_39, 0, x_3); +x_16 = x_39; +x_17 = x_32; +goto block_22; } else { -lean_object* x_552; lean_object* x_553; -x_552 = lean_ctor_get(x_549, 1); -lean_inc(x_552); -lean_dec(x_549); -x_553 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_553, 0, x_509); -lean_ctor_set(x_553, 1, x_552); -return x_553; -} +lean_object* x_40; +lean_dec(x_31); +x_40 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_40, 0, x_3); +x_16 = x_40; +x_17 = x_32; +goto block_22; } } else { -lean_object* x_554; lean_object* x_555; lean_object* x_556; lean_object* x_557; lean_object* x_558; uint8_t x_559; -x_554 = lean_box(0); -x_555 = lean_array_uset(x_515, x_531, x_554); -lean_inc(x_509); -x_556 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__8(x_1, x_509, x_532); -x_557 = lean_array_uset(x_555, x_531, x_556); -lean_ctor_set(x_511, 1, x_557); -x_558 = lean_st_ref_set(x_5, x_511, x_512); -lean_dec(x_5); -x_559 = !lean_is_exclusive(x_558); -if (x_559 == 0) +uint8_t x_41; +lean_dec(x_29); +lean_free_object(x_3); +lean_dec(x_25); +lean_dec(x_24); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_4); +x_41 = !lean_is_exclusive(x_30); +if (x_41 == 0) { -lean_object* x_560; -x_560 = lean_ctor_get(x_558, 0); -lean_dec(x_560); -lean_ctor_set(x_558, 0, x_509); -return x_558; +return x_30; } else { -lean_object* x_561; lean_object* x_562; -x_561 = lean_ctor_get(x_558, 1); -lean_inc(x_561); -lean_dec(x_558); -x_562 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_562, 0, x_509); -lean_ctor_set(x_562, 1, x_561); -return x_562; +lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_42 = lean_ctor_get(x_30, 0); +x_43 = lean_ctor_get(x_30, 1); +lean_inc(x_43); +lean_inc(x_42); +lean_dec(x_30); +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_563; lean_object* x_564; lean_object* x_565; size_t x_566; uint64_t x_567; uint64_t x_568; uint64_t x_569; uint64_t x_570; uint64_t x_571; uint64_t x_572; uint64_t x_573; uint64_t x_574; uint64_t x_575; size_t x_576; size_t x_577; size_t x_578; size_t x_579; size_t x_580; lean_object* x_581; uint8_t x_582; -x_563 = lean_ctor_get(x_511, 0); -x_564 = lean_ctor_get(x_511, 1); -lean_inc(x_564); -lean_inc(x_563); -lean_dec(x_511); -x_565 = lean_array_get_size(x_564); -x_566 = lean_ptr_addr(x_1); -x_567 = lean_usize_to_uint64(x_566); -x_568 = 11; -x_569 = lean_uint64_mix_hash(x_567, x_568); -x_570 = 32; -x_571 = lean_uint64_shift_right(x_569, x_570); -x_572 = lean_uint64_xor(x_569, x_571); -x_573 = 16; -x_574 = lean_uint64_shift_right(x_572, x_573); -x_575 = lean_uint64_xor(x_572, x_574); -x_576 = lean_uint64_to_usize(x_575); -x_577 = lean_usize_of_nat(x_565); -lean_dec(x_565); -x_578 = 1; -x_579 = lean_usize_sub(x_577, x_578); -x_580 = lean_usize_land(x_576, x_579); -x_581 = lean_array_uget(x_564, x_580); -x_582 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__3(x_1, x_581); -if (x_582 == 0) -{ -lean_object* x_583; lean_object* x_584; lean_object* x_585; lean_object* x_586; lean_object* x_587; lean_object* x_588; lean_object* x_589; lean_object* x_590; uint8_t x_591; -x_583 = lean_nat_add(x_563, x_501); -lean_dec(x_563); -lean_inc(x_509); -x_584 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_584, 0, x_1); -lean_ctor_set(x_584, 1, x_509); -lean_ctor_set(x_584, 2, x_581); -x_585 = lean_array_uset(x_564, x_580, x_584); -x_586 = lean_unsigned_to_nat(4u); -x_587 = lean_nat_mul(x_583, x_586); -x_588 = lean_unsigned_to_nat(3u); -x_589 = lean_nat_div(x_587, x_588); -lean_dec(x_587); -x_590 = lean_array_get_size(x_585); -x_591 = lean_nat_dec_le(x_589, x_590); -lean_dec(x_590); -lean_dec(x_589); -if (x_591 == 0) -{ -lean_object* x_592; lean_object* x_593; lean_object* x_594; lean_object* x_595; lean_object* x_596; lean_object* x_597; -x_592 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__4(x_585); -x_593 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_593, 0, x_583); -lean_ctor_set(x_593, 1, x_592); -x_594 = lean_st_ref_set(x_5, x_593, x_512); -lean_dec(x_5); -x_595 = lean_ctor_get(x_594, 1); -lean_inc(x_595); -if (lean_is_exclusive(x_594)) { - lean_ctor_release(x_594, 0); - lean_ctor_release(x_594, 1); - x_596 = x_594; -} else { - lean_dec_ref(x_594); - x_596 = lean_box(0); -} -if (lean_is_scalar(x_596)) { - x_597 = lean_alloc_ctor(0, 2, 0); -} else { - x_597 = x_596; -} -lean_ctor_set(x_597, 0, x_509); -lean_ctor_set(x_597, 1, x_595); -return x_597; -} -else +lean_object* x_45; lean_object* x_46; +x_45 = lean_array_fget(x_24, 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_45); +x_46 = l_Lean_Meta_Grind_markNestedProofsImpl_visit(x_45, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_46) == 0) { -lean_object* x_598; lean_object* x_599; lean_object* x_600; lean_object* x_601; lean_object* x_602; -x_598 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_598, 0, x_583); -lean_ctor_set(x_598, 1, x_585); -x_599 = lean_st_ref_set(x_5, x_598, x_512); -lean_dec(x_5); -x_600 = lean_ctor_get(x_599, 1); -lean_inc(x_600); -if (lean_is_exclusive(x_599)) { - lean_ctor_release(x_599, 0); - lean_ctor_release(x_599, 1); - x_601 = x_599; -} else { - lean_dec_ref(x_599); - x_601 = lean_box(0); -} -if (lean_is_scalar(x_601)) { - x_602 = lean_alloc_ctor(0, 2, 0); -} else { - x_602 = x_601; -} -lean_ctor_set(x_602, 0, x_509); -lean_ctor_set(x_602, 1, x_600); -return x_602; -} +lean_object* x_47; lean_object* x_48; size_t x_49; size_t x_50; uint8_t x_51; +x_47 = lean_ctor_get(x_46, 0); +lean_inc(x_47); +x_48 = lean_ctor_get(x_46, 1); +lean_inc(x_48); +lean_dec(x_46); +x_49 = lean_ptr_addr(x_45); +lean_dec(x_45); +x_50 = lean_ptr_addr(x_47); +x_51 = lean_usize_dec_eq(x_49, x_50); +if (x_51 == 0) +{ +lean_object* x_52; uint8_t x_53; lean_object* x_54; lean_object* x_55; +lean_dec(x_25); +x_52 = lean_array_set(x_24, x_4, x_47); +x_53 = 1; +x_54 = lean_box(x_53); +lean_ctor_set(x_3, 1, x_54); +lean_ctor_set(x_3, 0, x_52); +x_55 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_55, 0, x_3); +x_16 = x_55; +x_17 = x_48; +goto block_22; } else { -lean_object* x_603; lean_object* x_604; lean_object* x_605; lean_object* x_606; lean_object* x_607; lean_object* x_608; lean_object* x_609; lean_object* x_610; lean_object* x_611; -x_603 = lean_box(0); -x_604 = lean_array_uset(x_564, x_580, x_603); -lean_inc(x_509); -x_605 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__8(x_1, x_509, x_581); -x_606 = lean_array_uset(x_604, x_580, x_605); -x_607 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_607, 0, x_563); -lean_ctor_set(x_607, 1, x_606); -x_608 = lean_st_ref_set(x_5, x_607, x_512); -lean_dec(x_5); -x_609 = lean_ctor_get(x_608, 1); -lean_inc(x_609); -if (lean_is_exclusive(x_608)) { - lean_ctor_release(x_608, 0); - lean_ctor_release(x_608, 1); - x_610 = x_608; -} else { - lean_dec_ref(x_608); - x_610 = lean_box(0); -} -if (lean_is_scalar(x_610)) { - x_611 = lean_alloc_ctor(0, 2, 0); -} else { - x_611 = x_610; -} -lean_ctor_set(x_611, 0, x_509); -lean_ctor_set(x_611, 1, x_609); -return x_611; -} -} +lean_object* x_56; +lean_dec(x_47); +x_56 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_56, 0, x_3); +x_16 = x_56; +x_17 = x_48; +goto block_22; } } else { -uint8_t x_617; -lean_dec(x_5); -lean_dec(x_2); -lean_dec(x_1); -x_617 = !lean_is_exclusive(x_506); -if (x_617 == 0) +uint8_t x_57; +lean_dec(x_45); +lean_free_object(x_3); +lean_dec(x_25); +lean_dec(x_24); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_4); +x_57 = !lean_is_exclusive(x_46); +if (x_57 == 0) { -return x_506; +return x_46; } else { -lean_object* x_618; lean_object* x_619; lean_object* x_620; -x_618 = lean_ctor_get(x_506, 0); -x_619 = lean_ctor_get(x_506, 1); -lean_inc(x_619); -lean_inc(x_618); -lean_dec(x_506); -x_620 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_620, 0, x_618); -lean_ctor_set(x_620, 1, x_619); -return x_620; -} +lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_58 = lean_ctor_get(x_46, 0); +x_59 = lean_ctor_get(x_46, 1); +lean_inc(x_59); +lean_inc(x_58); +lean_dec(x_46); +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 5: -{ -lean_object* x_621; lean_object* x_622; lean_object* x_623; lean_object* x_624; lean_object* x_625; -x_621 = lean_ctor_get(x_2, 0); -lean_inc(x_621); -x_622 = lean_ctor_get(x_2, 1); -lean_inc(x_622); -lean_dec(x_2); -x_623 = lean_array_set(x_3, x_4, x_622); -x_624 = lean_unsigned_to_nat(1u); -x_625 = lean_nat_sub(x_4, x_624); -lean_dec(x_4); -x_2 = x_621; -x_3 = x_623; -x_4 = x_625; -goto _start; } -case 6: -{ -lean_object* x_627; lean_object* x_628; lean_object* x_629; lean_object* x_630; uint8_t x_631; lean_object* x_632; lean_object* x_633; lean_object* x_634; -lean_dec(x_4); -x_627 = lean_array_get_size(x_3); -x_628 = lean_unsigned_to_nat(0u); -x_629 = lean_unsigned_to_nat(1u); -x_630 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_630, 0, x_628); -lean_ctor_set(x_630, 1, x_627); -lean_ctor_set(x_630, 2, x_629); -x_631 = 0; -x_632 = lean_box(x_631); -x_633 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_633, 0, x_3); -lean_ctor_set(x_633, 1, x_632); -lean_inc(x_5); -x_634 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__2(x_630, x_630, x_633, x_628, lean_box(0), lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_630); -if (lean_obj_tag(x_634) == 0) -{ -lean_object* x_635; lean_object* x_636; lean_object* x_637; lean_object* x_741; uint8_t x_742; -x_635 = lean_ctor_get(x_634, 0); -lean_inc(x_635); -x_636 = lean_ctor_get(x_634, 1); -lean_inc(x_636); -lean_dec(x_634); -x_741 = lean_ctor_get(x_635, 1); -lean_inc(x_741); -x_742 = lean_unbox(x_741); -lean_dec(x_741); -if (x_742 == 0) -{ -lean_dec(x_635); -lean_dec(x_2); -lean_inc(x_1); -x_637 = x_1; -goto block_740; } else { -lean_object* x_743; lean_object* x_744; -x_743 = lean_ctor_get(x_635, 0); -lean_inc(x_743); -lean_dec(x_635); -x_744 = l_Lean_mkAppN(x_2, x_743); -lean_dec(x_743); -x_637 = x_744; -goto block_740; -} -block_740: -{ -lean_object* x_638; lean_object* x_639; lean_object* x_640; uint8_t x_641; -x_638 = lean_st_ref_take(x_5, x_636); -x_639 = lean_ctor_get(x_638, 0); -lean_inc(x_639); -x_640 = lean_ctor_get(x_638, 1); -lean_inc(x_640); -lean_dec(x_638); -x_641 = !lean_is_exclusive(x_639); -if (x_641 == 0) -{ -lean_object* x_642; lean_object* x_643; lean_object* x_644; size_t x_645; uint64_t x_646; uint64_t x_647; uint64_t x_648; uint64_t x_649; uint64_t x_650; uint64_t x_651; uint64_t x_652; uint64_t x_653; uint64_t x_654; size_t x_655; size_t x_656; size_t x_657; size_t x_658; size_t x_659; lean_object* x_660; uint8_t x_661; -x_642 = lean_ctor_get(x_639, 0); -x_643 = lean_ctor_get(x_639, 1); -x_644 = lean_array_get_size(x_643); -x_645 = lean_ptr_addr(x_1); -x_646 = lean_usize_to_uint64(x_645); -x_647 = 11; -x_648 = lean_uint64_mix_hash(x_646, x_647); -x_649 = 32; -x_650 = lean_uint64_shift_right(x_648, x_649); -x_651 = lean_uint64_xor(x_648, x_650); -x_652 = 16; -x_653 = lean_uint64_shift_right(x_651, x_652); -x_654 = lean_uint64_xor(x_651, x_653); -x_655 = lean_uint64_to_usize(x_654); -x_656 = lean_usize_of_nat(x_644); -lean_dec(x_644); -x_657 = 1; -x_658 = lean_usize_sub(x_656, x_657); -x_659 = lean_usize_land(x_655, x_658); -x_660 = lean_array_uget(x_643, x_659); -x_661 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__3(x_1, x_660); -if (x_661 == 0) -{ -lean_object* x_662; lean_object* x_663; lean_object* x_664; lean_object* x_665; lean_object* x_666; lean_object* x_667; lean_object* x_668; lean_object* x_669; uint8_t x_670; -x_662 = lean_nat_add(x_642, x_629); -lean_dec(x_642); -lean_inc(x_637); -x_663 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_663, 0, x_1); -lean_ctor_set(x_663, 1, x_637); -lean_ctor_set(x_663, 2, x_660); -x_664 = lean_array_uset(x_643, x_659, x_663); -x_665 = lean_unsigned_to_nat(4u); -x_666 = lean_nat_mul(x_662, x_665); -x_667 = lean_unsigned_to_nat(3u); -x_668 = lean_nat_div(x_666, x_667); -lean_dec(x_666); -x_669 = lean_array_get_size(x_664); -x_670 = lean_nat_dec_le(x_668, x_669); -lean_dec(x_669); -lean_dec(x_668); -if (x_670 == 0) -{ -lean_object* x_671; lean_object* x_672; uint8_t x_673; -x_671 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__4(x_664); -lean_ctor_set(x_639, 1, x_671); -lean_ctor_set(x_639, 0, x_662); -x_672 = lean_st_ref_set(x_5, x_639, x_640); -lean_dec(x_5); -x_673 = !lean_is_exclusive(x_672); -if (x_673 == 0) +lean_object* x_61; lean_object* x_62; lean_object* x_63; uint8_t x_64; +x_61 = lean_ctor_get(x_3, 0); +x_62 = lean_ctor_get(x_3, 1); +lean_inc(x_62); +lean_inc(x_61); +lean_dec(x_3); +x_63 = lean_array_get_size(x_61); +x_64 = lean_nat_dec_lt(x_4, x_63); +lean_dec(x_63); +if (x_64 == 0) +{ +lean_object* x_65; lean_object* x_66; lean_object* x_67; +x_65 = l_Lean_instInhabitedExpr; +x_66 = l_outOfBounds___rarg(x_65); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_66); +x_67 = l_Lean_Meta_Grind_markNestedProofsImpl_visit(x_66, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_67) == 0) +{ +lean_object* x_68; lean_object* x_69; size_t x_70; size_t x_71; uint8_t x_72; +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_ptr_addr(x_66); +lean_dec(x_66); +x_71 = lean_ptr_addr(x_68); +x_72 = lean_usize_dec_eq(x_70, x_71); +if (x_72 == 0) { -lean_object* x_674; -x_674 = lean_ctor_get(x_672, 0); -lean_dec(x_674); -lean_ctor_set(x_672, 0, x_637); -return x_672; +lean_object* x_73; uint8_t x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; +lean_dec(x_62); +x_73 = lean_array_set(x_61, x_4, x_68); +x_74 = 1; +x_75 = lean_box(x_74); +x_76 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_76, 0, x_73); +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_16 = x_77; +x_17 = x_69; +goto block_22; } else { -lean_object* x_675; lean_object* x_676; -x_675 = lean_ctor_get(x_672, 1); -lean_inc(x_675); -lean_dec(x_672); -x_676 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_676, 0, x_637); -lean_ctor_set(x_676, 1, x_675); -return x_676; +lean_object* x_78; lean_object* x_79; +lean_dec(x_68); +x_78 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_78, 0, x_61); +lean_ctor_set(x_78, 1, x_62); +x_79 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_79, 0, x_78); +x_16 = x_79; +x_17 = x_69; +goto block_22; } } else { -lean_object* x_677; uint8_t x_678; -lean_ctor_set(x_639, 1, x_664); -lean_ctor_set(x_639, 0, x_662); -x_677 = lean_st_ref_set(x_5, x_639, x_640); -lean_dec(x_5); -x_678 = !lean_is_exclusive(x_677); -if (x_678 == 0) -{ -lean_object* x_679; -x_679 = lean_ctor_get(x_677, 0); -lean_dec(x_679); -lean_ctor_set(x_677, 0, x_637); -return x_677; +lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; +lean_dec(x_66); +lean_dec(x_62); +lean_dec(x_61); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_4); +x_80 = lean_ctor_get(x_67, 0); +lean_inc(x_80); +x_81 = lean_ctor_get(x_67, 1); +lean_inc(x_81); +if (lean_is_exclusive(x_67)) { + lean_ctor_release(x_67, 0); + lean_ctor_release(x_67, 1); + x_82 = x_67; +} else { + lean_dec_ref(x_67); + x_82 = lean_box(0); } -else -{ -lean_object* x_680; lean_object* x_681; -x_680 = lean_ctor_get(x_677, 1); -lean_inc(x_680); -lean_dec(x_677); -x_681 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_681, 0, x_637); -lean_ctor_set(x_681, 1, x_680); -return x_681; +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_682; lean_object* x_683; lean_object* x_684; lean_object* x_685; lean_object* x_686; uint8_t x_687; -x_682 = lean_box(0); -x_683 = lean_array_uset(x_643, x_659, x_682); -lean_inc(x_637); -x_684 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__8(x_1, x_637, x_660); -x_685 = lean_array_uset(x_683, x_659, x_684); -lean_ctor_set(x_639, 1, x_685); -x_686 = lean_st_ref_set(x_5, x_639, x_640); -lean_dec(x_5); -x_687 = !lean_is_exclusive(x_686); -if (x_687 == 0) +lean_object* x_84; lean_object* x_85; +x_84 = lean_array_fget(x_61, 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_84); +x_85 = l_Lean_Meta_Grind_markNestedProofsImpl_visit(x_84, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_85) == 0) { -lean_object* x_688; -x_688 = lean_ctor_get(x_686, 0); -lean_dec(x_688); -lean_ctor_set(x_686, 0, x_637); -return x_686; -} -else +lean_object* x_86; lean_object* x_87; size_t x_88; size_t x_89; uint8_t x_90; +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_ptr_addr(x_84); +lean_dec(x_84); +x_89 = lean_ptr_addr(x_86); +x_90 = lean_usize_dec_eq(x_88, x_89); +if (x_90 == 0) { -lean_object* x_689; lean_object* x_690; -x_689 = lean_ctor_get(x_686, 1); -lean_inc(x_689); -lean_dec(x_686); -x_690 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_690, 0, x_637); -lean_ctor_set(x_690, 1, x_689); -return x_690; -} -} +lean_object* x_91; uint8_t x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; +lean_dec(x_62); +x_91 = lean_array_set(x_61, x_4, x_86); +x_92 = 1; +x_93 = lean_box(x_92); +x_94 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_94, 0, x_91); +lean_ctor_set(x_94, 1, x_93); +x_95 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_95, 0, x_94); +x_16 = x_95; +x_17 = x_87; +goto block_22; } else { -lean_object* x_691; lean_object* x_692; lean_object* x_693; size_t x_694; uint64_t x_695; uint64_t x_696; uint64_t x_697; uint64_t x_698; uint64_t x_699; uint64_t x_700; uint64_t x_701; uint64_t x_702; uint64_t x_703; size_t x_704; size_t x_705; size_t x_706; size_t x_707; size_t x_708; lean_object* x_709; uint8_t x_710; -x_691 = lean_ctor_get(x_639, 0); -x_692 = lean_ctor_get(x_639, 1); -lean_inc(x_692); -lean_inc(x_691); -lean_dec(x_639); -x_693 = lean_array_get_size(x_692); -x_694 = lean_ptr_addr(x_1); -x_695 = lean_usize_to_uint64(x_694); -x_696 = 11; -x_697 = lean_uint64_mix_hash(x_695, x_696); -x_698 = 32; -x_699 = lean_uint64_shift_right(x_697, x_698); -x_700 = lean_uint64_xor(x_697, x_699); -x_701 = 16; -x_702 = lean_uint64_shift_right(x_700, x_701); -x_703 = lean_uint64_xor(x_700, x_702); -x_704 = lean_uint64_to_usize(x_703); -x_705 = lean_usize_of_nat(x_693); -lean_dec(x_693); -x_706 = 1; -x_707 = lean_usize_sub(x_705, x_706); -x_708 = lean_usize_land(x_704, x_707); -x_709 = lean_array_uget(x_692, x_708); -x_710 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__3(x_1, x_709); -if (x_710 == 0) -{ -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; uint8_t x_719; -x_711 = lean_nat_add(x_691, x_629); -lean_dec(x_691); -lean_inc(x_637); -x_712 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_712, 0, x_1); -lean_ctor_set(x_712, 1, x_637); -lean_ctor_set(x_712, 2, x_709); -x_713 = lean_array_uset(x_692, x_708, x_712); -x_714 = lean_unsigned_to_nat(4u); -x_715 = lean_nat_mul(x_711, x_714); -x_716 = lean_unsigned_to_nat(3u); -x_717 = lean_nat_div(x_715, x_716); -lean_dec(x_715); -x_718 = lean_array_get_size(x_713); -x_719 = lean_nat_dec_le(x_717, x_718); -lean_dec(x_718); -lean_dec(x_717); -if (x_719 == 0) -{ -lean_object* x_720; lean_object* x_721; lean_object* x_722; lean_object* x_723; lean_object* x_724; lean_object* x_725; -x_720 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__4(x_713); -x_721 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_721, 0, x_711); -lean_ctor_set(x_721, 1, x_720); -x_722 = lean_st_ref_set(x_5, x_721, x_640); -lean_dec(x_5); -x_723 = lean_ctor_get(x_722, 1); -lean_inc(x_723); -if (lean_is_exclusive(x_722)) { - lean_ctor_release(x_722, 0); - lean_ctor_release(x_722, 1); - x_724 = x_722; -} else { - lean_dec_ref(x_722); - x_724 = lean_box(0); -} -if (lean_is_scalar(x_724)) { - x_725 = lean_alloc_ctor(0, 2, 0); -} else { - x_725 = x_724; +lean_object* x_96; lean_object* x_97; +lean_dec(x_86); +x_96 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_96, 0, x_61); +lean_ctor_set(x_96, 1, x_62); +x_97 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_97, 0, x_96); +x_16 = x_97; +x_17 = x_87; +goto block_22; } -lean_ctor_set(x_725, 0, x_637); -lean_ctor_set(x_725, 1, x_723); -return x_725; } else { -lean_object* x_726; lean_object* x_727; lean_object* x_728; lean_object* x_729; lean_object* x_730; -x_726 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_726, 0, x_711); -lean_ctor_set(x_726, 1, x_713); -x_727 = lean_st_ref_set(x_5, x_726, x_640); -lean_dec(x_5); -x_728 = lean_ctor_get(x_727, 1); -lean_inc(x_728); -if (lean_is_exclusive(x_727)) { - lean_ctor_release(x_727, 0); - lean_ctor_release(x_727, 1); - x_729 = x_727; +lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; +lean_dec(x_84); +lean_dec(x_62); +lean_dec(x_61); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_4); +x_98 = lean_ctor_get(x_85, 0); +lean_inc(x_98); +x_99 = lean_ctor_get(x_85, 1); +lean_inc(x_99); +if (lean_is_exclusive(x_85)) { + lean_ctor_release(x_85, 0); + lean_ctor_release(x_85, 1); + x_100 = x_85; } else { - lean_dec_ref(x_727); - x_729 = lean_box(0); + lean_dec_ref(x_85); + x_100 = lean_box(0); } -if (lean_is_scalar(x_729)) { - x_730 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_100)) { + x_101 = lean_alloc_ctor(1, 2, 0); } else { - x_730 = x_729; -} -lean_ctor_set(x_730, 0, x_637); -lean_ctor_set(x_730, 1, x_728); -return x_730; + x_101 = x_100; } +lean_ctor_set(x_101, 0, x_98); +lean_ctor_set(x_101, 1, x_99); +return x_101; } -else -{ -lean_object* x_731; lean_object* x_732; lean_object* x_733; lean_object* x_734; lean_object* x_735; lean_object* x_736; lean_object* x_737; lean_object* x_738; lean_object* x_739; -x_731 = lean_box(0); -x_732 = lean_array_uset(x_692, x_708, x_731); -lean_inc(x_637); -x_733 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__8(x_1, x_637, x_709); -x_734 = lean_array_uset(x_732, x_708, x_733); -x_735 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_735, 0, x_691); -lean_ctor_set(x_735, 1, x_734); -x_736 = lean_st_ref_set(x_5, x_735, x_640); -lean_dec(x_5); -x_737 = lean_ctor_get(x_736, 1); -lean_inc(x_737); -if (lean_is_exclusive(x_736)) { - lean_ctor_release(x_736, 0); - lean_ctor_release(x_736, 1); - x_738 = x_736; -} else { - lean_dec_ref(x_736); - x_738 = lean_box(0); } -if (lean_is_scalar(x_738)) { - x_739 = lean_alloc_ctor(0, 2, 0); -} else { - x_739 = x_738; } -lean_ctor_set(x_739, 0, x_637); -lean_ctor_set(x_739, 1, x_737); -return x_739; +block_22: +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_18 = lean_ctor_get(x_16, 0); +lean_inc(x_18); +lean_dec(x_16); +x_19 = lean_ctor_get(x_2, 2); +x_20 = lean_nat_add(x_4, x_19); +lean_dec(x_4); +x_3 = x_18; +x_4 = x_20; +x_5 = lean_box(0); +x_6 = lean_box(0); +x_12 = x_17; +goto _start; } } } } -else +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__9(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: { -uint8_t x_745; -lean_dec(x_5); -lean_dec(x_2); -lean_dec(x_1); -x_745 = !lean_is_exclusive(x_634); -if (x_745 == 0) +if (lean_obj_tag(x_2) == 5) { -return x_634; +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_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_array_set(x_3, x_4, x_12); +x_14 = lean_unsigned_to_nat(1u); +x_15 = lean_nat_sub(x_4, x_14); +lean_dec(x_4); +x_2 = x_11; +x_3 = x_13; +x_4 = x_15; +goto _start; } else { -lean_object* x_746; lean_object* x_747; lean_object* x_748; -x_746 = lean_ctor_get(x_634, 0); -x_747 = lean_ctor_get(x_634, 1); -lean_inc(x_747); -lean_inc(x_746); -lean_dec(x_634); -x_748 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_748, 0, x_746); -lean_ctor_set(x_748, 1, x_747); -return x_748; -} -} -} -case 7: -{ -lean_object* x_749; lean_object* x_750; lean_object* x_751; lean_object* x_752; uint8_t x_753; lean_object* x_754; lean_object* x_755; lean_object* x_756; +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_dec(x_4); -x_749 = lean_array_get_size(x_3); -x_750 = lean_unsigned_to_nat(0u); -x_751 = lean_unsigned_to_nat(1u); -x_752 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_752, 0, x_750); -lean_ctor_set(x_752, 1, x_749); -lean_ctor_set(x_752, 2, x_751); -x_753 = 0; -x_754 = lean_box(x_753); -x_755 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_755, 0, x_3); -lean_ctor_set(x_755, 1, x_754); -lean_inc(x_5); -x_756 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__2(x_752, x_752, x_755, x_750, lean_box(0), lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_752); -if (lean_obj_tag(x_756) == 0) -{ -lean_object* x_757; lean_object* x_758; lean_object* x_759; lean_object* x_863; uint8_t x_864; -x_757 = lean_ctor_get(x_756, 0); -lean_inc(x_757); -x_758 = lean_ctor_get(x_756, 1); -lean_inc(x_758); -lean_dec(x_756); -x_863 = lean_ctor_get(x_757, 1); -lean_inc(x_863); -x_864 = lean_unbox(x_863); -lean_dec(x_863); -if (x_864 == 0) -{ -lean_dec(x_757); -lean_dec(x_2); -lean_inc(x_1); -x_759 = x_1; -goto block_862; -} -else +x_17 = lean_array_get_size(x_3); +x_18 = lean_unsigned_to_nat(0u); +x_19 = lean_unsigned_to_nat(1u); +x_20 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_20, 0, x_18); +lean_ctor_set(x_20, 1, x_17); +lean_ctor_set(x_20, 2, x_19); +x_21 = 0; +x_22 = lean_box(x_21); +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_3); +lean_ctor_set(x_23, 1, x_22); +x_24 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__8(x_20, x_20, x_23, x_18, lean_box(0), lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_20); +if (lean_obj_tag(x_24) == 0) { -lean_object* x_865; lean_object* x_866; -x_865 = lean_ctor_get(x_757, 0); -lean_inc(x_865); -lean_dec(x_757); -x_866 = l_Lean_mkAppN(x_2, x_865); -lean_dec(x_865); -x_759 = x_866; -goto block_862; -} -block_862: -{ -lean_object* x_760; lean_object* x_761; lean_object* x_762; uint8_t x_763; -x_760 = lean_st_ref_take(x_5, x_758); -x_761 = lean_ctor_get(x_760, 0); -lean_inc(x_761); -x_762 = lean_ctor_get(x_760, 1); -lean_inc(x_762); -lean_dec(x_760); -x_763 = !lean_is_exclusive(x_761); -if (x_763 == 0) -{ -lean_object* x_764; lean_object* x_765; lean_object* x_766; size_t x_767; uint64_t x_768; uint64_t x_769; uint64_t x_770; uint64_t x_771; uint64_t x_772; uint64_t x_773; uint64_t x_774; uint64_t x_775; uint64_t x_776; size_t x_777; size_t x_778; size_t x_779; size_t x_780; size_t x_781; lean_object* x_782; uint8_t x_783; -x_764 = lean_ctor_get(x_761, 0); -x_765 = lean_ctor_get(x_761, 1); -x_766 = lean_array_get_size(x_765); -x_767 = lean_ptr_addr(x_1); -x_768 = lean_usize_to_uint64(x_767); -x_769 = 11; -x_770 = lean_uint64_mix_hash(x_768, x_769); -x_771 = 32; -x_772 = lean_uint64_shift_right(x_770, x_771); -x_773 = lean_uint64_xor(x_770, x_772); -x_774 = 16; -x_775 = lean_uint64_shift_right(x_773, x_774); -x_776 = lean_uint64_xor(x_773, x_775); -x_777 = lean_uint64_to_usize(x_776); -x_778 = lean_usize_of_nat(x_766); -lean_dec(x_766); -x_779 = 1; -x_780 = lean_usize_sub(x_778, x_779); -x_781 = lean_usize_land(x_777, x_780); -x_782 = lean_array_uget(x_765, x_781); -x_783 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__3(x_1, x_782); -if (x_783 == 0) -{ -lean_object* x_784; lean_object* x_785; lean_object* x_786; lean_object* x_787; lean_object* x_788; lean_object* x_789; lean_object* x_790; lean_object* x_791; uint8_t x_792; -x_784 = lean_nat_add(x_764, x_751); -lean_dec(x_764); -lean_inc(x_759); -x_785 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_785, 0, x_1); -lean_ctor_set(x_785, 1, x_759); -lean_ctor_set(x_785, 2, x_782); -x_786 = lean_array_uset(x_765, x_781, x_785); -x_787 = lean_unsigned_to_nat(4u); -x_788 = lean_nat_mul(x_784, x_787); -x_789 = lean_unsigned_to_nat(3u); -x_790 = lean_nat_div(x_788, x_789); -lean_dec(x_788); -x_791 = lean_array_get_size(x_786); -x_792 = lean_nat_dec_le(x_790, x_791); -lean_dec(x_791); -lean_dec(x_790); -if (x_792 == 0) -{ -lean_object* x_793; lean_object* x_794; uint8_t x_795; -x_793 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__4(x_786); -lean_ctor_set(x_761, 1, x_793); -lean_ctor_set(x_761, 0, x_784); -x_794 = lean_st_ref_set(x_5, x_761, x_762); -lean_dec(x_5); -x_795 = !lean_is_exclusive(x_794); -if (x_795 == 0) +lean_object* x_25; lean_object* x_26; uint8_t x_27; +x_25 = lean_ctor_get(x_24, 0); +lean_inc(x_25); +x_26 = lean_ctor_get(x_25, 1); +lean_inc(x_26); +x_27 = lean_unbox(x_26); +lean_dec(x_26); +if (x_27 == 0) +{ +uint8_t x_28; +lean_dec(x_25); +lean_dec(x_2); +x_28 = !lean_is_exclusive(x_24); +if (x_28 == 0) { -lean_object* x_796; -x_796 = lean_ctor_get(x_794, 0); -lean_dec(x_796); -lean_ctor_set(x_794, 0, x_759); -return x_794; +lean_object* x_29; +x_29 = lean_ctor_get(x_24, 0); +lean_dec(x_29); +lean_ctor_set(x_24, 0, x_1); +return x_24; } else { -lean_object* x_797; lean_object* x_798; -x_797 = lean_ctor_get(x_794, 1); -lean_inc(x_797); -lean_dec(x_794); -x_798 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_798, 0, x_759); -lean_ctor_set(x_798, 1, x_797); -return x_798; +lean_object* x_30; lean_object* x_31; +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_1); +lean_ctor_set(x_31, 1, x_30); +return x_31; } } else { -lean_object* x_799; uint8_t x_800; -lean_ctor_set(x_761, 1, x_786); -lean_ctor_set(x_761, 0, x_784); -x_799 = lean_st_ref_set(x_5, x_761, x_762); -lean_dec(x_5); -x_800 = !lean_is_exclusive(x_799); -if (x_800 == 0) +uint8_t x_32; +lean_dec(x_1); +x_32 = !lean_is_exclusive(x_24); +if (x_32 == 0) { -lean_object* x_801; -x_801 = lean_ctor_get(x_799, 0); -lean_dec(x_801); -lean_ctor_set(x_799, 0, x_759); -return x_799; +lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_33 = lean_ctor_get(x_24, 0); +lean_dec(x_33); +x_34 = lean_ctor_get(x_25, 0); +lean_inc(x_34); +lean_dec(x_25); +x_35 = l_Lean_mkAppN(x_2, x_34); +lean_dec(x_34); +lean_ctor_set(x_24, 0, x_35); +return x_24; } else { -lean_object* x_802; lean_object* x_803; -x_802 = lean_ctor_get(x_799, 1); -lean_inc(x_802); -lean_dec(x_799); -x_803 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_803, 0, x_759); -lean_ctor_set(x_803, 1, x_802); -return x_803; +lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_36 = lean_ctor_get(x_24, 1); +lean_inc(x_36); +lean_dec(x_24); +x_37 = lean_ctor_get(x_25, 0); +lean_inc(x_37); +lean_dec(x_25); +x_38 = l_Lean_mkAppN(x_2, x_37); +lean_dec(x_37); +x_39 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_36); +return x_39; } } } else { -lean_object* x_804; lean_object* x_805; lean_object* x_806; lean_object* x_807; lean_object* x_808; uint8_t x_809; -x_804 = lean_box(0); -x_805 = lean_array_uset(x_765, x_781, x_804); -lean_inc(x_759); -x_806 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__8(x_1, x_759, x_782); -x_807 = lean_array_uset(x_805, x_781, x_806); -lean_ctor_set(x_761, 1, x_807); -x_808 = lean_st_ref_set(x_5, x_761, x_762); -lean_dec(x_5); -x_809 = !lean_is_exclusive(x_808); -if (x_809 == 0) +uint8_t x_40; +lean_dec(x_2); +lean_dec(x_1); +x_40 = !lean_is_exclusive(x_24); +if (x_40 == 0) { -lean_object* x_810; -x_810 = lean_ctor_get(x_808, 0); -lean_dec(x_810); -lean_ctor_set(x_808, 0, x_759); -return x_808; +return x_24; } else { -lean_object* x_811; lean_object* x_812; -x_811 = lean_ctor_get(x_808, 1); -lean_inc(x_811); -lean_dec(x_808); -x_812 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_812, 0, x_759); -lean_ctor_set(x_812, 1, x_811); -return x_812; -} +lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_41 = lean_ctor_get(x_24, 0); +x_42 = lean_ctor_get(x_24, 1); +lean_inc(x_42); +lean_inc(x_41); +lean_dec(x_24); +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_813; lean_object* x_814; lean_object* x_815; size_t x_816; uint64_t x_817; uint64_t x_818; uint64_t x_819; uint64_t x_820; uint64_t x_821; uint64_t x_822; uint64_t x_823; uint64_t x_824; uint64_t x_825; size_t x_826; size_t x_827; size_t x_828; size_t x_829; size_t x_830; lean_object* x_831; uint8_t x_832; -x_813 = lean_ctor_get(x_761, 0); -x_814 = lean_ctor_get(x_761, 1); -lean_inc(x_814); -lean_inc(x_813); -lean_dec(x_761); -x_815 = lean_array_get_size(x_814); -x_816 = lean_ptr_addr(x_1); -x_817 = lean_usize_to_uint64(x_816); -x_818 = 11; -x_819 = lean_uint64_mix_hash(x_817, x_818); -x_820 = 32; -x_821 = lean_uint64_shift_right(x_819, x_820); -x_822 = lean_uint64_xor(x_819, x_821); -x_823 = 16; -x_824 = lean_uint64_shift_right(x_822, x_823); -x_825 = lean_uint64_xor(x_822, x_824); -x_826 = lean_uint64_to_usize(x_825); -x_827 = lean_usize_of_nat(x_815); -lean_dec(x_815); -x_828 = 1; -x_829 = lean_usize_sub(x_827, x_828); -x_830 = lean_usize_land(x_826, x_829); -x_831 = lean_array_uget(x_814, x_830); -x_832 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__3(x_1, x_831); -if (x_832 == 0) -{ -lean_object* x_833; lean_object* x_834; lean_object* x_835; lean_object* x_836; lean_object* x_837; lean_object* x_838; lean_object* x_839; lean_object* x_840; uint8_t x_841; -x_833 = lean_nat_add(x_813, x_751); -lean_dec(x_813); -lean_inc(x_759); -x_834 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_834, 0, x_1); -lean_ctor_set(x_834, 1, x_759); -lean_ctor_set(x_834, 2, x_831); -x_835 = lean_array_uset(x_814, x_830, x_834); -x_836 = lean_unsigned_to_nat(4u); -x_837 = lean_nat_mul(x_833, x_836); -x_838 = lean_unsigned_to_nat(3u); -x_839 = lean_nat_div(x_837, x_838); -lean_dec(x_837); -x_840 = lean_array_get_size(x_835); -x_841 = lean_nat_dec_le(x_839, x_840); -lean_dec(x_840); -lean_dec(x_839); -if (x_841 == 0) -{ -lean_object* x_842; lean_object* x_843; lean_object* x_844; lean_object* x_845; lean_object* x_846; lean_object* x_847; -x_842 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__4(x_835); -x_843 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_843, 0, x_833); -lean_ctor_set(x_843, 1, x_842); -x_844 = lean_st_ref_set(x_5, x_843, x_762); -lean_dec(x_5); -x_845 = lean_ctor_get(x_844, 1); -lean_inc(x_845); -if (lean_is_exclusive(x_844)) { - lean_ctor_release(x_844, 0); - lean_ctor_release(x_844, 1); - x_846 = x_844; -} else { - lean_dec_ref(x_844); - x_846 = lean_box(0); } -if (lean_is_scalar(x_846)) { - x_847 = lean_alloc_ctor(0, 2, 0); -} else { - x_847 = x_846; } -lean_ctor_set(x_847, 0, x_759); -lean_ctor_set(x_847, 1, x_845); -return x_847; } -else +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__10(lean_object* x_1, lean_object* x_2) { +_start: { -lean_object* x_848; lean_object* x_849; lean_object* x_850; lean_object* x_851; lean_object* x_852; -x_848 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_848, 0, x_833); -lean_ctor_set(x_848, 1, x_835); -x_849 = lean_st_ref_set(x_5, x_848, x_762); -lean_dec(x_5); -x_850 = lean_ctor_get(x_849, 1); -lean_inc(x_850); -if (lean_is_exclusive(x_849)) { - lean_ctor_release(x_849, 0); - lean_ctor_release(x_849, 1); - x_851 = x_849; -} else { - lean_dec_ref(x_849); - x_851 = lean_box(0); -} -if (lean_is_scalar(x_851)) { - x_852 = lean_alloc_ctor(0, 2, 0); -} else { - x_852 = x_851; -} -lean_ctor_set(x_852, 0, x_759); -lean_ctor_set(x_852, 1, x_850); -return x_852; -} -} -else +if (lean_obj_tag(x_2) == 0) { -lean_object* x_853; lean_object* x_854; lean_object* x_855; lean_object* x_856; lean_object* x_857; lean_object* x_858; lean_object* x_859; lean_object* x_860; lean_object* x_861; -x_853 = lean_box(0); -x_854 = lean_array_uset(x_814, x_830, x_853); -lean_inc(x_759); -x_855 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__8(x_1, x_759, x_831); -x_856 = lean_array_uset(x_854, x_830, x_855); -x_857 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_857, 0, x_813); -lean_ctor_set(x_857, 1, x_856); -x_858 = lean_st_ref_set(x_5, x_857, x_762); -lean_dec(x_5); -x_859 = lean_ctor_get(x_858, 1); -lean_inc(x_859); -if (lean_is_exclusive(x_858)) { - lean_ctor_release(x_858, 0); - lean_ctor_release(x_858, 1); - x_860 = x_858; -} else { - lean_dec_ref(x_858); - x_860 = lean_box(0); -} -if (lean_is_scalar(x_860)) { - x_861 = lean_alloc_ctor(0, 2, 0); -} else { - x_861 = x_860; -} -lean_ctor_set(x_861, 0, x_759); -lean_ctor_set(x_861, 1, x_859); -return x_861; -} -} -} +lean_object* x_3; +x_3 = lean_box(0); +return x_3; } else { -uint8_t x_867; -lean_dec(x_5); -lean_dec(x_2); -lean_dec(x_1); -x_867 = !lean_is_exclusive(x_756); -if (x_867 == 0) +lean_object* x_4; lean_object* x_5; lean_object* x_6; size_t x_7; size_t x_8; uint8_t x_9; +x_4 = lean_ctor_get(x_2, 0); +x_5 = lean_ctor_get(x_2, 1); +x_6 = lean_ctor_get(x_2, 2); +x_7 = lean_ptr_addr(x_4); +x_8 = lean_ptr_addr(x_1); +x_9 = lean_usize_dec_eq(x_7, x_8); +if (x_9 == 0) { -return x_756; +x_2 = x_6; +goto _start; } else { -lean_object* x_868; lean_object* x_869; lean_object* x_870; -x_868 = lean_ctor_get(x_756, 0); -x_869 = lean_ctor_get(x_756, 1); -lean_inc(x_869); -lean_inc(x_868); -lean_dec(x_756); -x_870 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_870, 0, x_868); -lean_ctor_set(x_870, 1, x_869); -return x_870; +lean_object* x_11; +lean_inc(x_5); +x_11 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_11, 0, x_5); +return x_11; } } } -case 8: -{ -lean_object* x_871; lean_object* x_872; lean_object* x_873; lean_object* x_874; uint8_t x_875; lean_object* x_876; lean_object* x_877; lean_object* x_878; -lean_dec(x_4); -x_871 = lean_array_get_size(x_3); -x_872 = lean_unsigned_to_nat(0u); -x_873 = lean_unsigned_to_nat(1u); -x_874 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_874, 0, x_872); -lean_ctor_set(x_874, 1, x_871); -lean_ctor_set(x_874, 2, x_873); -x_875 = 0; -x_876 = lean_box(x_875); -x_877 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_877, 0, x_3); -lean_ctor_set(x_877, 1, x_876); -lean_inc(x_5); -x_878 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__2(x_874, x_874, x_877, x_872, lean_box(0), lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_874); -if (lean_obj_tag(x_878) == 0) -{ -lean_object* x_879; lean_object* x_880; lean_object* x_881; lean_object* x_985; uint8_t x_986; -x_879 = lean_ctor_get(x_878, 0); -lean_inc(x_879); -x_880 = lean_ctor_get(x_878, 1); -lean_inc(x_880); -lean_dec(x_878); -x_985 = lean_ctor_get(x_879, 1); -lean_inc(x_985); -x_986 = lean_unbox(x_985); -lean_dec(x_985); -if (x_986 == 0) -{ -lean_dec(x_879); -lean_dec(x_2); -lean_inc(x_1); -x_881 = x_1; -goto block_984; } -else +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___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_987; lean_object* x_988; -x_987 = lean_ctor_get(x_879, 0); -lean_inc(x_987); -lean_dec(x_879); -x_988 = l_Lean_mkAppN(x_2, x_987); -lean_dec(x_987); -x_881 = x_988; -goto block_984; -} -block_984: -{ -lean_object* x_882; lean_object* x_883; lean_object* x_884; uint8_t x_885; -x_882 = lean_st_ref_take(x_5, x_880); -x_883 = lean_ctor_get(x_882, 0); -lean_inc(x_883); -x_884 = lean_ctor_get(x_882, 1); -lean_inc(x_884); -lean_dec(x_882); -x_885 = !lean_is_exclusive(x_883); -if (x_885 == 0) -{ -lean_object* x_886; lean_object* x_887; lean_object* x_888; size_t x_889; uint64_t x_890; uint64_t x_891; uint64_t x_892; uint64_t x_893; uint64_t x_894; uint64_t x_895; uint64_t x_896; uint64_t x_897; uint64_t x_898; size_t x_899; size_t x_900; size_t x_901; size_t x_902; size_t x_903; lean_object* x_904; uint8_t x_905; -x_886 = lean_ctor_get(x_883, 0); -x_887 = lean_ctor_get(x_883, 1); -x_888 = lean_array_get_size(x_887); -x_889 = lean_ptr_addr(x_1); -x_890 = lean_usize_to_uint64(x_889); -x_891 = 11; -x_892 = lean_uint64_mix_hash(x_890, x_891); -x_893 = 32; -x_894 = lean_uint64_shift_right(x_892, x_893); -x_895 = lean_uint64_xor(x_892, x_894); -x_896 = 16; -x_897 = lean_uint64_shift_right(x_895, x_896); -x_898 = lean_uint64_xor(x_895, x_897); -x_899 = lean_uint64_to_usize(x_898); -x_900 = lean_usize_of_nat(x_888); -lean_dec(x_888); -x_901 = 1; -x_902 = lean_usize_sub(x_900, x_901); -x_903 = lean_usize_land(x_899, x_902); -x_904 = lean_array_uget(x_887, x_903); -x_905 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__3(x_1, x_904); -if (x_905 == 0) -{ -lean_object* x_906; lean_object* x_907; lean_object* x_908; lean_object* x_909; lean_object* x_910; lean_object* x_911; lean_object* x_912; lean_object* x_913; uint8_t x_914; -x_906 = lean_nat_add(x_886, x_873); -lean_dec(x_886); -lean_inc(x_881); -x_907 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_907, 0, x_1); -lean_ctor_set(x_907, 1, x_881); -lean_ctor_set(x_907, 2, x_904); -x_908 = lean_array_uset(x_887, x_903, x_907); -x_909 = lean_unsigned_to_nat(4u); -x_910 = lean_nat_mul(x_906, x_909); -x_911 = lean_unsigned_to_nat(3u); -x_912 = lean_nat_div(x_910, x_911); -lean_dec(x_910); -x_913 = lean_array_get_size(x_908); -x_914 = lean_nat_dec_le(x_912, x_913); -lean_dec(x_913); -lean_dec(x_912); -if (x_914 == 0) -{ -lean_object* x_915; lean_object* x_916; uint8_t x_917; -x_915 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__4(x_908); -lean_ctor_set(x_883, 1, x_915); -lean_ctor_set(x_883, 0, x_906); -x_916 = lean_st_ref_set(x_5, x_883, x_884); -lean_dec(x_5); -x_917 = !lean_is_exclusive(x_916); -if (x_917 == 0) +lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; +x_9 = lean_st_ref_take(x_3, x_8); +x_10 = lean_ctor_get(x_9, 0); +lean_inc(x_10); +x_11 = lean_ctor_get(x_9, 1); +lean_inc(x_11); +lean_dec(x_9); +x_12 = !lean_is_exclusive(x_10); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; size_t x_16; uint64_t x_17; uint64_t x_18; uint64_t x_19; uint64_t x_20; uint64_t x_21; uint64_t x_22; uint64_t x_23; uint64_t x_24; uint64_t x_25; size_t x_26; size_t x_27; size_t x_28; size_t x_29; size_t x_30; lean_object* x_31; uint8_t x_32; +x_13 = lean_ctor_get(x_10, 0); +x_14 = lean_ctor_get(x_10, 1); +x_15 = lean_array_get_size(x_14); +x_16 = lean_ptr_addr(x_1); +x_17 = lean_usize_to_uint64(x_16); +x_18 = 11; +x_19 = lean_uint64_mix_hash(x_17, x_18); +x_20 = 32; +x_21 = lean_uint64_shift_right(x_19, x_20); +x_22 = lean_uint64_xor(x_19, x_21); +x_23 = 16; +x_24 = lean_uint64_shift_right(x_22, x_23); +x_25 = lean_uint64_xor(x_22, x_24); +x_26 = lean_uint64_to_usize(x_25); +x_27 = lean_usize_of_nat(x_15); +lean_dec(x_15); +x_28 = 1; +x_29 = lean_usize_sub(x_27, x_28); +x_30 = lean_usize_land(x_26, x_29); +x_31 = lean_array_uget(x_14, x_30); +x_32 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__1(x_1, 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; uint8_t x_42; +x_33 = lean_unsigned_to_nat(1u); +x_34 = lean_nat_add(x_13, x_33); +lean_dec(x_13); +lean_inc(x_2); +x_35 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_35, 0, x_1); +lean_ctor_set(x_35, 1, x_2); +lean_ctor_set(x_35, 2, x_31); +x_36 = lean_array_uset(x_14, x_30, x_35); +x_37 = lean_unsigned_to_nat(4u); +x_38 = lean_nat_mul(x_34, x_37); +x_39 = lean_unsigned_to_nat(3u); +x_40 = lean_nat_div(x_38, x_39); +lean_dec(x_38); +x_41 = lean_array_get_size(x_36); +x_42 = lean_nat_dec_le(x_40, x_41); +lean_dec(x_41); +lean_dec(x_40); +if (x_42 == 0) +{ +lean_object* x_43; lean_object* x_44; uint8_t x_45; +x_43 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__2(x_36); +lean_ctor_set(x_10, 1, x_43); +lean_ctor_set(x_10, 0, x_34); +x_44 = lean_st_ref_set(x_3, x_10, x_11); +x_45 = !lean_is_exclusive(x_44); +if (x_45 == 0) { -lean_object* x_918; -x_918 = lean_ctor_get(x_916, 0); -lean_dec(x_918); -lean_ctor_set(x_916, 0, x_881); -return x_916; +lean_object* x_46; +x_46 = lean_ctor_get(x_44, 0); +lean_dec(x_46); +lean_ctor_set(x_44, 0, x_2); +return x_44; } else { -lean_object* x_919; lean_object* x_920; -x_919 = lean_ctor_get(x_916, 1); -lean_inc(x_919); -lean_dec(x_916); -x_920 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_920, 0, x_881); -lean_ctor_set(x_920, 1, x_919); -return x_920; +lean_object* x_47; lean_object* x_48; +x_47 = lean_ctor_get(x_44, 1); +lean_inc(x_47); +lean_dec(x_44); +x_48 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_48, 0, x_2); +lean_ctor_set(x_48, 1, x_47); +return x_48; } } else { -lean_object* x_921; uint8_t x_922; -lean_ctor_set(x_883, 1, x_908); -lean_ctor_set(x_883, 0, x_906); -x_921 = lean_st_ref_set(x_5, x_883, x_884); -lean_dec(x_5); -x_922 = !lean_is_exclusive(x_921); -if (x_922 == 0) +lean_object* x_49; uint8_t x_50; +lean_ctor_set(x_10, 1, x_36); +lean_ctor_set(x_10, 0, x_34); +x_49 = lean_st_ref_set(x_3, x_10, x_11); +x_50 = !lean_is_exclusive(x_49); +if (x_50 == 0) { -lean_object* x_923; -x_923 = lean_ctor_get(x_921, 0); -lean_dec(x_923); -lean_ctor_set(x_921, 0, x_881); -return x_921; +lean_object* x_51; +x_51 = lean_ctor_get(x_49, 0); +lean_dec(x_51); +lean_ctor_set(x_49, 0, x_2); +return x_49; } else { -lean_object* x_924; lean_object* x_925; -x_924 = lean_ctor_get(x_921, 1); -lean_inc(x_924); -lean_dec(x_921); -x_925 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_925, 0, x_881); -lean_ctor_set(x_925, 1, x_924); -return x_925; +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_2); +lean_ctor_set(x_53, 1, x_52); +return x_53; } } } else { -lean_object* x_926; lean_object* x_927; lean_object* x_928; lean_object* x_929; lean_object* x_930; uint8_t x_931; -x_926 = lean_box(0); -x_927 = lean_array_uset(x_887, x_903, x_926); -lean_inc(x_881); -x_928 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__8(x_1, x_881, x_904); -x_929 = lean_array_uset(x_927, x_903, x_928); -lean_ctor_set(x_883, 1, x_929); -x_930 = lean_st_ref_set(x_5, x_883, x_884); -lean_dec(x_5); -x_931 = !lean_is_exclusive(x_930); -if (x_931 == 0) +lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; uint8_t x_59; +x_54 = lean_box(0); +x_55 = lean_array_uset(x_14, x_30, x_54); +lean_inc(x_2); +x_56 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__6(x_1, x_2, x_31); +x_57 = lean_array_uset(x_55, x_30, x_56); +lean_ctor_set(x_10, 1, x_57); +x_58 = lean_st_ref_set(x_3, x_10, x_11); +x_59 = !lean_is_exclusive(x_58); +if (x_59 == 0) { -lean_object* x_932; -x_932 = lean_ctor_get(x_930, 0); -lean_dec(x_932); -lean_ctor_set(x_930, 0, x_881); -return x_930; +lean_object* x_60; +x_60 = lean_ctor_get(x_58, 0); +lean_dec(x_60); +lean_ctor_set(x_58, 0, x_2); +return x_58; } else { -lean_object* x_933; lean_object* x_934; -x_933 = lean_ctor_get(x_930, 1); -lean_inc(x_933); -lean_dec(x_930); -x_934 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_934, 0, x_881); -lean_ctor_set(x_934, 1, x_933); -return x_934; +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(0, 2, 0); +lean_ctor_set(x_62, 0, x_2); +lean_ctor_set(x_62, 1, x_61); +return x_62; } } } else { -lean_object* x_935; lean_object* x_936; lean_object* x_937; size_t x_938; uint64_t x_939; uint64_t x_940; uint64_t x_941; uint64_t x_942; uint64_t x_943; uint64_t x_944; uint64_t x_945; uint64_t x_946; uint64_t x_947; size_t x_948; size_t x_949; size_t x_950; size_t x_951; size_t x_952; lean_object* x_953; uint8_t x_954; -x_935 = lean_ctor_get(x_883, 0); -x_936 = lean_ctor_get(x_883, 1); -lean_inc(x_936); -lean_inc(x_935); -lean_dec(x_883); -x_937 = lean_array_get_size(x_936); -x_938 = lean_ptr_addr(x_1); -x_939 = lean_usize_to_uint64(x_938); -x_940 = 11; -x_941 = lean_uint64_mix_hash(x_939, x_940); -x_942 = 32; -x_943 = lean_uint64_shift_right(x_941, x_942); -x_944 = lean_uint64_xor(x_941, x_943); -x_945 = 16; -x_946 = lean_uint64_shift_right(x_944, x_945); -x_947 = lean_uint64_xor(x_944, x_946); -x_948 = lean_uint64_to_usize(x_947); -x_949 = lean_usize_of_nat(x_937); -lean_dec(x_937); -x_950 = 1; -x_951 = lean_usize_sub(x_949, x_950); -x_952 = lean_usize_land(x_948, x_951); -x_953 = lean_array_uget(x_936, x_952); -x_954 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__3(x_1, x_953); -if (x_954 == 0) -{ -lean_object* x_955; lean_object* x_956; lean_object* x_957; lean_object* x_958; lean_object* x_959; lean_object* x_960; lean_object* x_961; lean_object* x_962; uint8_t x_963; -x_955 = lean_nat_add(x_935, x_873); -lean_dec(x_935); -lean_inc(x_881); -x_956 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_956, 0, x_1); -lean_ctor_set(x_956, 1, x_881); -lean_ctor_set(x_956, 2, x_953); -x_957 = lean_array_uset(x_936, x_952, x_956); -x_958 = lean_unsigned_to_nat(4u); -x_959 = lean_nat_mul(x_955, x_958); -x_960 = lean_unsigned_to_nat(3u); -x_961 = lean_nat_div(x_959, x_960); -lean_dec(x_959); -x_962 = lean_array_get_size(x_957); -x_963 = lean_nat_dec_le(x_961, x_962); -lean_dec(x_962); -lean_dec(x_961); -if (x_963 == 0) -{ -lean_object* x_964; lean_object* x_965; lean_object* x_966; lean_object* x_967; lean_object* x_968; lean_object* x_969; -x_964 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__4(x_957); -x_965 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_965, 0, x_955); -lean_ctor_set(x_965, 1, x_964); -x_966 = lean_st_ref_set(x_5, x_965, x_884); -lean_dec(x_5); -x_967 = lean_ctor_get(x_966, 1); -lean_inc(x_967); -if (lean_is_exclusive(x_966)) { - lean_ctor_release(x_966, 0); - lean_ctor_release(x_966, 1); - x_968 = x_966; +lean_object* x_63; lean_object* x_64; lean_object* x_65; size_t x_66; uint64_t x_67; uint64_t x_68; uint64_t x_69; uint64_t x_70; uint64_t x_71; uint64_t x_72; uint64_t x_73; uint64_t x_74; uint64_t x_75; size_t x_76; size_t x_77; size_t x_78; size_t x_79; size_t x_80; lean_object* x_81; uint8_t x_82; +x_63 = lean_ctor_get(x_10, 0); +x_64 = lean_ctor_get(x_10, 1); +lean_inc(x_64); +lean_inc(x_63); +lean_dec(x_10); +x_65 = lean_array_get_size(x_64); +x_66 = lean_ptr_addr(x_1); +x_67 = lean_usize_to_uint64(x_66); +x_68 = 11; +x_69 = lean_uint64_mix_hash(x_67, x_68); +x_70 = 32; +x_71 = lean_uint64_shift_right(x_69, x_70); +x_72 = lean_uint64_xor(x_69, x_71); +x_73 = 16; +x_74 = lean_uint64_shift_right(x_72, x_73); +x_75 = lean_uint64_xor(x_72, x_74); +x_76 = lean_uint64_to_usize(x_75); +x_77 = lean_usize_of_nat(x_65); +lean_dec(x_65); +x_78 = 1; +x_79 = lean_usize_sub(x_77, x_78); +x_80 = lean_usize_land(x_76, x_79); +x_81 = lean_array_uget(x_64, x_80); +x_82 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__1(x_1, x_81); +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; uint8_t x_92; +x_83 = lean_unsigned_to_nat(1u); +x_84 = lean_nat_add(x_63, x_83); +lean_dec(x_63); +lean_inc(x_2); +x_85 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_85, 0, x_1); +lean_ctor_set(x_85, 1, x_2); +lean_ctor_set(x_85, 2, x_81); +x_86 = lean_array_uset(x_64, x_80, x_85); +x_87 = lean_unsigned_to_nat(4u); +x_88 = lean_nat_mul(x_84, x_87); +x_89 = lean_unsigned_to_nat(3u); +x_90 = lean_nat_div(x_88, x_89); +lean_dec(x_88); +x_91 = lean_array_get_size(x_86); +x_92 = lean_nat_dec_le(x_90, x_91); +lean_dec(x_91); +lean_dec(x_90); +if (x_92 == 0) +{ +lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; +x_93 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__2(x_86); +x_94 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_94, 0, x_84); +lean_ctor_set(x_94, 1, x_93); +x_95 = lean_st_ref_set(x_3, x_94, x_11); +x_96 = lean_ctor_get(x_95, 1); +lean_inc(x_96); +if (lean_is_exclusive(x_95)) { + lean_ctor_release(x_95, 0); + lean_ctor_release(x_95, 1); + x_97 = x_95; } else { - lean_dec_ref(x_966); - x_968 = lean_box(0); + lean_dec_ref(x_95); + x_97 = lean_box(0); } -if (lean_is_scalar(x_968)) { - x_969 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_97)) { + x_98 = lean_alloc_ctor(0, 2, 0); } else { - x_969 = x_968; + x_98 = x_97; } -lean_ctor_set(x_969, 0, x_881); -lean_ctor_set(x_969, 1, x_967); -return x_969; +lean_ctor_set(x_98, 0, x_2); +lean_ctor_set(x_98, 1, x_96); +return x_98; } else { -lean_object* x_970; lean_object* x_971; lean_object* x_972; lean_object* x_973; lean_object* x_974; -x_970 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_970, 0, x_955); -lean_ctor_set(x_970, 1, x_957); -x_971 = lean_st_ref_set(x_5, x_970, x_884); -lean_dec(x_5); -x_972 = lean_ctor_get(x_971, 1); -lean_inc(x_972); -if (lean_is_exclusive(x_971)) { - lean_ctor_release(x_971, 0); - lean_ctor_release(x_971, 1); - x_973 = x_971; +lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; +x_99 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_99, 0, x_84); +lean_ctor_set(x_99, 1, x_86); +x_100 = lean_st_ref_set(x_3, x_99, x_11); +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_971); - x_973 = lean_box(0); + lean_dec_ref(x_100); + x_102 = lean_box(0); } -if (lean_is_scalar(x_973)) { - x_974 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_102)) { + x_103 = lean_alloc_ctor(0, 2, 0); } else { - x_974 = x_973; + x_103 = x_102; } -lean_ctor_set(x_974, 0, x_881); -lean_ctor_set(x_974, 1, x_972); -return x_974; +lean_ctor_set(x_103, 0, x_2); +lean_ctor_set(x_103, 1, x_101); +return x_103; } } else { -lean_object* x_975; lean_object* x_976; lean_object* x_977; lean_object* x_978; lean_object* x_979; lean_object* x_980; lean_object* x_981; lean_object* x_982; lean_object* x_983; -x_975 = lean_box(0); -x_976 = lean_array_uset(x_936, x_952, x_975); -lean_inc(x_881); -x_977 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__8(x_1, x_881, x_953); -x_978 = lean_array_uset(x_976, x_952, x_977); -x_979 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_979, 0, x_935); -lean_ctor_set(x_979, 1, x_978); -x_980 = lean_st_ref_set(x_5, x_979, x_884); -lean_dec(x_5); -x_981 = lean_ctor_get(x_980, 1); -lean_inc(x_981); -if (lean_is_exclusive(x_980)) { - lean_ctor_release(x_980, 0); - lean_ctor_release(x_980, 1); - x_982 = x_980; +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_104 = lean_box(0); +x_105 = lean_array_uset(x_64, x_80, x_104); +lean_inc(x_2); +x_106 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__6(x_1, x_2, x_81); +x_107 = lean_array_uset(x_105, x_80, x_106); +x_108 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_108, 0, x_63); +lean_ctor_set(x_108, 1, x_107); +x_109 = lean_st_ref_set(x_3, x_108, x_11); +x_110 = lean_ctor_get(x_109, 1); +lean_inc(x_110); +if (lean_is_exclusive(x_109)) { + lean_ctor_release(x_109, 0); + lean_ctor_release(x_109, 1); + x_111 = x_109; } else { - lean_dec_ref(x_980); - x_982 = lean_box(0); + lean_dec_ref(x_109); + x_111 = lean_box(0); } -if (lean_is_scalar(x_982)) { - x_983 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_111)) { + x_112 = lean_alloc_ctor(0, 2, 0); } else { - x_983 = x_982; + x_112 = x_111; } -lean_ctor_set(x_983, 0, x_881); -lean_ctor_set(x_983, 1, x_981); -return x_983; +lean_ctor_set(x_112, 0, x_2); +lean_ctor_set(x_112, 1, x_110); +return x_112; } } } } -else +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: { -uint8_t x_989; -lean_dec(x_5); +lean_object* x_14; size_t x_15; size_t x_16; uint8_t x_17; +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_14 = l_Lean_Expr_forallE___override(x_1, x_2, x_3, x_4); +x_15 = lean_ptr_addr(x_2); lean_dec(x_2); -lean_dec(x_1); -x_989 = !lean_is_exclusive(x_878); -if (x_989 == 0) +x_16 = lean_ptr_addr(x_5); +x_17 = lean_usize_dec_eq(x_15, x_16); +if (x_17 == 0) { -return x_878; +lean_object* x_18; lean_object* x_19; +lean_dec(x_14); +lean_dec(x_3); +x_18 = l_Lean_Expr_forallE___override(x_1, x_5, x_7, x_4); +x_19 = lean_apply_7(x_6, x_18, x_8, x_9, x_10, x_11, x_12, x_13); +return x_19; } else { -lean_object* x_990; lean_object* x_991; lean_object* x_992; -x_990 = lean_ctor_get(x_878, 0); -x_991 = lean_ctor_get(x_878, 1); -lean_inc(x_991); -lean_inc(x_990); -lean_dec(x_878); -x_992 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_992, 0, x_990); -lean_ctor_set(x_992, 1, x_991); -return x_992; -} -} -} -case 9: +size_t x_20; size_t x_21; uint8_t x_22; +x_20 = lean_ptr_addr(x_3); +lean_dec(x_3); +x_21 = lean_ptr_addr(x_7); +x_22 = lean_usize_dec_eq(x_20, x_21); +if (x_22 == 0) { -lean_object* x_993; lean_object* x_994; lean_object* x_995; lean_object* x_996; uint8_t x_997; lean_object* x_998; lean_object* x_999; lean_object* x_1000; -lean_dec(x_4); -x_993 = lean_array_get_size(x_3); -x_994 = lean_unsigned_to_nat(0u); -x_995 = lean_unsigned_to_nat(1u); -x_996 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_996, 0, x_994); -lean_ctor_set(x_996, 1, x_993); -lean_ctor_set(x_996, 2, x_995); -x_997 = 0; -x_998 = lean_box(x_997); -x_999 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_999, 0, x_3); -lean_ctor_set(x_999, 1, x_998); -lean_inc(x_5); -x_1000 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__2(x_996, x_996, x_999, x_994, lean_box(0), lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_996); -if (lean_obj_tag(x_1000) == 0) -{ -lean_object* x_1001; lean_object* x_1002; lean_object* x_1003; lean_object* x_1107; uint8_t x_1108; -x_1001 = lean_ctor_get(x_1000, 0); -lean_inc(x_1001); -x_1002 = lean_ctor_get(x_1000, 1); -lean_inc(x_1002); -lean_dec(x_1000); -x_1107 = lean_ctor_get(x_1001, 1); -lean_inc(x_1107); -x_1108 = lean_unbox(x_1107); -lean_dec(x_1107); -if (x_1108 == 0) -{ -lean_dec(x_1001); -lean_dec(x_2); -lean_inc(x_1); -x_1003 = x_1; -goto block_1106; +lean_object* x_23; lean_object* x_24; +lean_dec(x_14); +x_23 = l_Lean_Expr_forallE___override(x_1, x_5, x_7, x_4); +x_24 = lean_apply_7(x_6, x_23, x_8, x_9, x_10, x_11, x_12, x_13); +return x_24; } else { -lean_object* x_1109; lean_object* x_1110; -x_1109 = lean_ctor_get(x_1001, 0); -lean_inc(x_1109); -lean_dec(x_1001); -x_1110 = l_Lean_mkAppN(x_2, x_1109); -lean_dec(x_1109); -x_1003 = x_1110; -goto block_1106; -} -block_1106: -{ -lean_object* x_1004; lean_object* x_1005; lean_object* x_1006; uint8_t x_1007; -x_1004 = lean_st_ref_take(x_5, x_1002); -x_1005 = lean_ctor_get(x_1004, 0); -lean_inc(x_1005); -x_1006 = lean_ctor_get(x_1004, 1); -lean_inc(x_1006); -lean_dec(x_1004); -x_1007 = !lean_is_exclusive(x_1005); -if (x_1007 == 0) -{ -lean_object* x_1008; lean_object* x_1009; lean_object* x_1010; size_t x_1011; uint64_t x_1012; uint64_t x_1013; uint64_t x_1014; uint64_t x_1015; uint64_t x_1016; uint64_t x_1017; uint64_t x_1018; uint64_t x_1019; uint64_t x_1020; size_t x_1021; size_t x_1022; size_t x_1023; size_t x_1024; size_t x_1025; lean_object* x_1026; uint8_t x_1027; -x_1008 = lean_ctor_get(x_1005, 0); -x_1009 = lean_ctor_get(x_1005, 1); -x_1010 = lean_array_get_size(x_1009); -x_1011 = lean_ptr_addr(x_1); -x_1012 = lean_usize_to_uint64(x_1011); -x_1013 = 11; -x_1014 = lean_uint64_mix_hash(x_1012, x_1013); -x_1015 = 32; -x_1016 = lean_uint64_shift_right(x_1014, x_1015); -x_1017 = lean_uint64_xor(x_1014, x_1016); -x_1018 = 16; -x_1019 = lean_uint64_shift_right(x_1017, x_1018); -x_1020 = lean_uint64_xor(x_1017, x_1019); -x_1021 = lean_uint64_to_usize(x_1020); -x_1022 = lean_usize_of_nat(x_1010); -lean_dec(x_1010); -x_1023 = 1; -x_1024 = lean_usize_sub(x_1022, x_1023); -x_1025 = lean_usize_land(x_1021, x_1024); -x_1026 = lean_array_uget(x_1009, x_1025); -x_1027 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__3(x_1, x_1026); -if (x_1027 == 0) -{ -lean_object* x_1028; lean_object* x_1029; lean_object* x_1030; lean_object* x_1031; lean_object* x_1032; lean_object* x_1033; lean_object* x_1034; lean_object* x_1035; uint8_t x_1036; -x_1028 = lean_nat_add(x_1008, x_995); -lean_dec(x_1008); -lean_inc(x_1003); -x_1029 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1029, 0, x_1); -lean_ctor_set(x_1029, 1, x_1003); -lean_ctor_set(x_1029, 2, x_1026); -x_1030 = lean_array_uset(x_1009, x_1025, x_1029); -x_1031 = lean_unsigned_to_nat(4u); -x_1032 = lean_nat_mul(x_1028, x_1031); -x_1033 = lean_unsigned_to_nat(3u); -x_1034 = lean_nat_div(x_1032, x_1033); -lean_dec(x_1032); -x_1035 = lean_array_get_size(x_1030); -x_1036 = lean_nat_dec_le(x_1034, x_1035); -lean_dec(x_1035); -lean_dec(x_1034); -if (x_1036 == 0) -{ -lean_object* x_1037; lean_object* x_1038; uint8_t x_1039; -x_1037 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__4(x_1030); -lean_ctor_set(x_1005, 1, x_1037); -lean_ctor_set(x_1005, 0, x_1028); -x_1038 = lean_st_ref_set(x_5, x_1005, x_1006); -lean_dec(x_5); -x_1039 = !lean_is_exclusive(x_1038); -if (x_1039 == 0) -{ -lean_object* x_1040; -x_1040 = lean_ctor_get(x_1038, 0); -lean_dec(x_1040); -lean_ctor_set(x_1038, 0, x_1003); -return x_1038; -} -else +uint8_t x_25; +x_25 = l___private_Lean_Expr_0__Lean_beqBinderInfo____x40_Lean_Expr___hyg_406_(x_4, x_4); +if (x_25 == 0) { -lean_object* x_1041; lean_object* x_1042; -x_1041 = lean_ctor_get(x_1038, 1); -lean_inc(x_1041); -lean_dec(x_1038); -x_1042 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1042, 0, x_1003); -lean_ctor_set(x_1042, 1, x_1041); -return x_1042; -} +lean_object* x_26; lean_object* x_27; +lean_dec(x_14); +x_26 = l_Lean_Expr_forallE___override(x_1, x_5, x_7, x_4); +x_27 = lean_apply_7(x_6, x_26, x_8, x_9, x_10, x_11, x_12, x_13); +return x_27; } else { -lean_object* x_1043; uint8_t x_1044; -lean_ctor_set(x_1005, 1, x_1030); -lean_ctor_set(x_1005, 0, x_1028); -x_1043 = lean_st_ref_set(x_5, x_1005, x_1006); +lean_object* x_28; +lean_dec(x_7); lean_dec(x_5); -x_1044 = !lean_is_exclusive(x_1043); -if (x_1044 == 0) -{ -lean_object* x_1045; -x_1045 = lean_ctor_get(x_1043, 0); -lean_dec(x_1045); -lean_ctor_set(x_1043, 0, x_1003); -return x_1043; +lean_dec(x_1); +x_28 = lean_apply_7(x_6, x_14, x_8, x_9, x_10, x_11, x_12, x_13); +return x_28; } -else -{ -lean_object* x_1046; lean_object* x_1047; -x_1046 = lean_ctor_get(x_1043, 1); -lean_inc(x_1046); -lean_dec(x_1043); -x_1047 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1047, 0, x_1003); -lean_ctor_set(x_1047, 1, x_1046); -return x_1047; } } } -else -{ -lean_object* x_1048; lean_object* x_1049; lean_object* x_1050; lean_object* x_1051; lean_object* x_1052; uint8_t x_1053; -x_1048 = lean_box(0); -x_1049 = lean_array_uset(x_1009, x_1025, x_1048); -lean_inc(x_1003); -x_1050 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__8(x_1, x_1003, x_1026); -x_1051 = lean_array_uset(x_1049, x_1025, x_1050); -lean_ctor_set(x_1005, 1, x_1051); -x_1052 = lean_st_ref_set(x_5, x_1005, x_1006); -lean_dec(x_5); -x_1053 = !lean_is_exclusive(x_1052); -if (x_1053 == 0) -{ -lean_object* x_1054; -x_1054 = lean_ctor_get(x_1052, 0); -lean_dec(x_1054); -lean_ctor_set(x_1052, 0, x_1003); -return x_1052; } -else +static lean_object* _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3___closed__1() { +_start: { -lean_object* x_1055; lean_object* x_1056; -x_1055 = lean_ctor_get(x_1052, 1); -lean_inc(x_1055); -lean_dec(x_1052); -x_1056 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1056, 0, x_1003); -lean_ctor_set(x_1056, 1, x_1055); -return x_1056; -} +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.Meta.Tactic.Grind.MarkNestedProofs", 39, 39); +return x_1; } } -else +static lean_object* _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3___closed__2() { +_start: { -lean_object* x_1057; lean_object* x_1058; lean_object* x_1059; size_t x_1060; uint64_t x_1061; uint64_t x_1062; uint64_t x_1063; uint64_t x_1064; uint64_t x_1065; uint64_t x_1066; uint64_t x_1067; uint64_t x_1068; uint64_t x_1069; size_t x_1070; size_t x_1071; size_t x_1072; size_t x_1073; size_t x_1074; lean_object* x_1075; uint8_t x_1076; -x_1057 = lean_ctor_get(x_1005, 0); -x_1058 = lean_ctor_get(x_1005, 1); -lean_inc(x_1058); -lean_inc(x_1057); -lean_dec(x_1005); -x_1059 = lean_array_get_size(x_1058); -x_1060 = lean_ptr_addr(x_1); -x_1061 = lean_usize_to_uint64(x_1060); -x_1062 = 11; -x_1063 = lean_uint64_mix_hash(x_1061, x_1062); -x_1064 = 32; -x_1065 = lean_uint64_shift_right(x_1063, x_1064); -x_1066 = lean_uint64_xor(x_1063, x_1065); -x_1067 = 16; -x_1068 = lean_uint64_shift_right(x_1066, x_1067); -x_1069 = lean_uint64_xor(x_1066, x_1068); -x_1070 = lean_uint64_to_usize(x_1069); -x_1071 = lean_usize_of_nat(x_1059); -lean_dec(x_1059); -x_1072 = 1; -x_1073 = lean_usize_sub(x_1071, x_1072); -x_1074 = lean_usize_land(x_1070, x_1073); -x_1075 = lean_array_uget(x_1058, x_1074); -x_1076 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__3(x_1, x_1075); -if (x_1076 == 0) -{ -lean_object* x_1077; lean_object* x_1078; lean_object* x_1079; lean_object* x_1080; lean_object* x_1081; lean_object* x_1082; lean_object* x_1083; lean_object* x_1084; uint8_t x_1085; -x_1077 = lean_nat_add(x_1057, x_995); -lean_dec(x_1057); -lean_inc(x_1003); -x_1078 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1078, 0, x_1); -lean_ctor_set(x_1078, 1, x_1003); -lean_ctor_set(x_1078, 2, x_1075); -x_1079 = lean_array_uset(x_1058, x_1074, x_1078); -x_1080 = lean_unsigned_to_nat(4u); -x_1081 = lean_nat_mul(x_1077, x_1080); -x_1082 = lean_unsigned_to_nat(3u); -x_1083 = lean_nat_div(x_1081, x_1082); -lean_dec(x_1081); -x_1084 = lean_array_get_size(x_1079); -x_1085 = lean_nat_dec_le(x_1083, x_1084); -lean_dec(x_1084); -lean_dec(x_1083); -if (x_1085 == 0) -{ -lean_object* x_1086; lean_object* x_1087; lean_object* x_1088; lean_object* x_1089; lean_object* x_1090; lean_object* x_1091; -x_1086 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__4(x_1079); -x_1087 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1087, 0, x_1077); -lean_ctor_set(x_1087, 1, x_1086); -x_1088 = lean_st_ref_set(x_5, x_1087, x_1006); -lean_dec(x_5); -x_1089 = lean_ctor_get(x_1088, 1); -lean_inc(x_1089); -if (lean_is_exclusive(x_1088)) { - lean_ctor_release(x_1088, 0); - lean_ctor_release(x_1088, 1); - x_1090 = x_1088; -} else { - lean_dec_ref(x_1088); - x_1090 = lean_box(0); -} -if (lean_is_scalar(x_1090)) { - x_1091 = lean_alloc_ctor(0, 2, 0); -} else { - x_1091 = x_1090; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.Meta.Grind.markNestedProofsImpl.visit", 42, 42); +return x_1; } -lean_ctor_set(x_1091, 0, x_1003); -lean_ctor_set(x_1091, 1, x_1089); -return x_1091; } -else +static lean_object* _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3___closed__3() { +_start: { -lean_object* x_1092; lean_object* x_1093; lean_object* x_1094; lean_object* x_1095; lean_object* x_1096; -x_1092 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1092, 0, x_1077); -lean_ctor_set(x_1092, 1, x_1079); -x_1093 = lean_st_ref_set(x_5, x_1092, x_1006); -lean_dec(x_5); -x_1094 = lean_ctor_get(x_1093, 1); -lean_inc(x_1094); -if (lean_is_exclusive(x_1093)) { - lean_ctor_release(x_1093, 0); - lean_ctor_release(x_1093, 1); - x_1095 = x_1093; -} else { - lean_dec_ref(x_1093); - x_1095 = lean_box(0); -} -if (lean_is_scalar(x_1095)) { - x_1096 = lean_alloc_ctor(0, 2, 0); -} else { - x_1096 = x_1095; -} -lean_ctor_set(x_1096, 0, x_1003); -lean_ctor_set(x_1096, 1, x_1094); -return x_1096; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("unreachable code has been reached", 33, 33); +return x_1; } } -else +static lean_object* _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3___closed__4() { +_start: { -lean_object* x_1097; lean_object* x_1098; lean_object* x_1099; lean_object* x_1100; lean_object* x_1101; lean_object* x_1102; lean_object* x_1103; lean_object* x_1104; lean_object* x_1105; -x_1097 = lean_box(0); -x_1098 = lean_array_uset(x_1058, x_1074, x_1097); -lean_inc(x_1003); -x_1099 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__8(x_1, x_1003, x_1075); -x_1100 = lean_array_uset(x_1098, x_1074, x_1099); -x_1101 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1101, 0, x_1057); -lean_ctor_set(x_1101, 1, x_1100); -x_1102 = lean_st_ref_set(x_5, x_1101, x_1006); -lean_dec(x_5); -x_1103 = lean_ctor_get(x_1102, 1); -lean_inc(x_1103); -if (lean_is_exclusive(x_1102)) { - lean_ctor_release(x_1102, 0); - lean_ctor_release(x_1102, 1); - x_1104 = x_1102; -} else { - lean_dec_ref(x_1102); - x_1104 = lean_box(0); -} -if (lean_is_scalar(x_1104)) { - x_1105 = lean_alloc_ctor(0, 2, 0); -} else { - x_1105 = x_1104; +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_Grind_markNestedProofsImpl_visit___lambda__3___closed__1; +x_2 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3___closed__2; +x_3 = lean_unsigned_to_nat(55u); +x_4 = lean_unsigned_to_nat(13u); +x_5 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___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_ctor_set(x_1105, 0, x_1003); -lean_ctor_set(x_1105, 1, x_1103); -return x_1105; } +static lean_object* _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_levelZero; +x_2 = l_Lean_Expr_sort___override(x_1); +return x_2; } } +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___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: +{ +switch (lean_obj_tag(x_1)) { +case 5: +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_9 = lean_unsigned_to_nat(0u); +x_10 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_9); +x_11 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3___closed__5; +lean_inc(x_10); +x_12 = lean_mk_array(x_10, x_11); +x_13 = lean_unsigned_to_nat(1u); +x_14 = lean_nat_sub(x_10, x_13); +lean_dec(x_10); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc_n(x_1, 2); +x_15 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__9(x_1, x_1, x_12, 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; +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_15, 1); +lean_inc(x_17); +lean_dec(x_15); +x_18 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__1(x_1, x_16, x_3, x_4, x_5, x_6, x_7, x_17); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_18; } else { -uint8_t x_1111; +uint8_t x_19; +lean_dec(x_7); +lean_dec(x_6); lean_dec(x_5); -lean_dec(x_2); +lean_dec(x_4); +lean_dec(x_3); lean_dec(x_1); -x_1111 = !lean_is_exclusive(x_1000); -if (x_1111 == 0) +x_19 = !lean_is_exclusive(x_15); +if (x_19 == 0) { -return x_1000; +return x_15; } else { -lean_object* x_1112; lean_object* x_1113; lean_object* x_1114; -x_1112 = lean_ctor_get(x_1000, 0); -x_1113 = lean_ctor_get(x_1000, 1); -lean_inc(x_1113); -lean_inc(x_1112); -lean_dec(x_1000); -x_1114 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_1114, 0, x_1112); -lean_ctor_set(x_1114, 1, x_1113); -return x_1114; +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_15, 0); +x_21 = lean_ctor_get(x_15, 1); +lean_inc(x_21); +lean_inc(x_20); +lean_dec(x_15); +x_22 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_22, 0, x_20); +lean_ctor_set(x_22, 1, x_21); +return x_22; } } } -case 10: +case 7: { -lean_object* x_1115; lean_object* x_1116; lean_object* x_1117; lean_object* x_1118; uint8_t x_1119; lean_object* x_1120; lean_object* x_1121; lean_object* x_1122; -lean_dec(x_4); -x_1115 = lean_array_get_size(x_3); -x_1116 = lean_unsigned_to_nat(0u); -x_1117 = lean_unsigned_to_nat(1u); -x_1118 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_1118, 0, x_1116); -lean_ctor_set(x_1118, 1, x_1115); -lean_ctor_set(x_1118, 2, x_1117); -x_1119 = 0; -x_1120 = lean_box(x_1119); -x_1121 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1121, 0, x_3); -lean_ctor_set(x_1121, 1, x_1120); +lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; lean_object* x_27; lean_object* x_28; +x_23 = lean_ctor_get(x_1, 0); +lean_inc(x_23); +x_24 = lean_ctor_get(x_1, 1); +lean_inc(x_24); +x_25 = lean_ctor_get(x_1, 2); +lean_inc(x_25); +x_26 = lean_ctor_get_uint8(x_1, sizeof(void*)*3 + 8); +x_27 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__1___boxed), 8, 1); +lean_closure_set(x_27, 0, x_1); +lean_inc(x_7); +lean_inc(x_6); lean_inc(x_5); -x_1122 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__2(x_1118, x_1118, x_1121, x_1116, lean_box(0), lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_1118); -if (lean_obj_tag(x_1122) == 0) -{ -lean_object* x_1123; lean_object* x_1124; lean_object* x_1125; lean_object* x_1229; uint8_t x_1230; -x_1123 = lean_ctor_get(x_1122, 0); -lean_inc(x_1123); -x_1124 = lean_ctor_get(x_1122, 1); -lean_inc(x_1124); -lean_dec(x_1122); -x_1229 = lean_ctor_get(x_1123, 1); -lean_inc(x_1229); -x_1230 = lean_unbox(x_1229); -lean_dec(x_1229); -if (x_1230 == 0) -{ -lean_dec(x_1123); -lean_dec(x_2); -lean_inc(x_1); -x_1125 = x_1; -goto block_1228; +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_24); +x_28 = l_Lean_Meta_Grind_markNestedProofsImpl_visit(x_24, x_3, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_28) == 0) +{ +lean_object* x_29; lean_object* x_30; uint8_t x_31; +x_29 = lean_ctor_get(x_28, 0); +lean_inc(x_29); +x_30 = lean_ctor_get(x_28, 1); +lean_inc(x_30); +lean_dec(x_28); +x_31 = l_Lean_Expr_hasLooseBVars(x_25); +if (x_31 == 0) +{ +lean_object* x_32; +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_25); +x_32 = l_Lean_Meta_Grind_markNestedProofsImpl_visit(x_25, x_3, x_4, x_5, x_6, x_7, x_30); +if (lean_obj_tag(x_32) == 0) +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_33 = lean_ctor_get(x_32, 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_Grind_markNestedProofsImpl_visit___lambda__2(x_23, x_24, x_25, x_26, x_29, x_27, x_33, x_3, x_4, x_5, x_6, x_7, x_34); +return x_35; } else { -lean_object* x_1231; lean_object* x_1232; -x_1231 = lean_ctor_get(x_1123, 0); -lean_inc(x_1231); -lean_dec(x_1123); -x_1232 = l_Lean_mkAppN(x_2, x_1231); -lean_dec(x_1231); -x_1125 = x_1232; -goto block_1228; -} -block_1228: -{ -lean_object* x_1126; lean_object* x_1127; lean_object* x_1128; uint8_t x_1129; -x_1126 = lean_st_ref_take(x_5, x_1124); -x_1127 = lean_ctor_get(x_1126, 0); -lean_inc(x_1127); -x_1128 = lean_ctor_get(x_1126, 1); -lean_inc(x_1128); -lean_dec(x_1126); -x_1129 = !lean_is_exclusive(x_1127); -if (x_1129 == 0) -{ -lean_object* x_1130; lean_object* x_1131; lean_object* x_1132; size_t x_1133; uint64_t x_1134; uint64_t x_1135; uint64_t x_1136; uint64_t x_1137; uint64_t x_1138; uint64_t x_1139; uint64_t x_1140; uint64_t x_1141; uint64_t x_1142; size_t x_1143; size_t x_1144; size_t x_1145; size_t x_1146; size_t x_1147; lean_object* x_1148; uint8_t x_1149; -x_1130 = lean_ctor_get(x_1127, 0); -x_1131 = lean_ctor_get(x_1127, 1); -x_1132 = lean_array_get_size(x_1131); -x_1133 = lean_ptr_addr(x_1); -x_1134 = lean_usize_to_uint64(x_1133); -x_1135 = 11; -x_1136 = lean_uint64_mix_hash(x_1134, x_1135); -x_1137 = 32; -x_1138 = lean_uint64_shift_right(x_1136, x_1137); -x_1139 = lean_uint64_xor(x_1136, x_1138); -x_1140 = 16; -x_1141 = lean_uint64_shift_right(x_1139, x_1140); -x_1142 = lean_uint64_xor(x_1139, x_1141); -x_1143 = lean_uint64_to_usize(x_1142); -x_1144 = lean_usize_of_nat(x_1132); -lean_dec(x_1132); -x_1145 = 1; -x_1146 = lean_usize_sub(x_1144, x_1145); -x_1147 = lean_usize_land(x_1143, x_1146); -x_1148 = lean_array_uget(x_1131, x_1147); -x_1149 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__3(x_1, x_1148); -if (x_1149 == 0) -{ -lean_object* x_1150; lean_object* x_1151; lean_object* x_1152; lean_object* x_1153; lean_object* x_1154; lean_object* x_1155; lean_object* x_1156; lean_object* x_1157; uint8_t x_1158; -x_1150 = lean_nat_add(x_1130, x_1117); -lean_dec(x_1130); -lean_inc(x_1125); -x_1151 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1151, 0, x_1); -lean_ctor_set(x_1151, 1, x_1125); -lean_ctor_set(x_1151, 2, x_1148); -x_1152 = lean_array_uset(x_1131, x_1147, x_1151); -x_1153 = lean_unsigned_to_nat(4u); -x_1154 = lean_nat_mul(x_1150, x_1153); -x_1155 = lean_unsigned_to_nat(3u); -x_1156 = lean_nat_div(x_1154, x_1155); -lean_dec(x_1154); -x_1157 = lean_array_get_size(x_1152); -x_1158 = lean_nat_dec_le(x_1156, x_1157); -lean_dec(x_1157); -lean_dec(x_1156); -if (x_1158 == 0) -{ -lean_object* x_1159; lean_object* x_1160; uint8_t x_1161; -x_1159 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__4(x_1152); -lean_ctor_set(x_1127, 1, x_1159); -lean_ctor_set(x_1127, 0, x_1150); -x_1160 = lean_st_ref_set(x_5, x_1127, x_1128); +uint8_t x_36; +lean_dec(x_29); +lean_dec(x_27); +lean_dec(x_25); +lean_dec(x_24); +lean_dec(x_23); +lean_dec(x_7); +lean_dec(x_6); lean_dec(x_5); -x_1161 = !lean_is_exclusive(x_1160); -if (x_1161 == 0) +lean_dec(x_4); +lean_dec(x_3); +x_36 = !lean_is_exclusive(x_32); +if (x_36 == 0) { -lean_object* x_1162; -x_1162 = lean_ctor_get(x_1160, 0); -lean_dec(x_1162); -lean_ctor_set(x_1160, 0, x_1125); -return x_1160; +return x_32; } else { -lean_object* x_1163; lean_object* x_1164; -x_1163 = lean_ctor_get(x_1160, 1); -lean_inc(x_1163); -lean_dec(x_1160); -x_1164 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1164, 0, x_1125); -lean_ctor_set(x_1164, 1, x_1163); -return x_1164; +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_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; } } -else -{ -lean_object* x_1165; uint8_t x_1166; -lean_ctor_set(x_1127, 1, x_1152); -lean_ctor_set(x_1127, 0, x_1150); -x_1165 = lean_st_ref_set(x_5, x_1127, x_1128); -lean_dec(x_5); -x_1166 = !lean_is_exclusive(x_1165); -if (x_1166 == 0) -{ -lean_object* x_1167; -x_1167 = lean_ctor_get(x_1165, 0); -lean_dec(x_1167); -lean_ctor_set(x_1165, 0, x_1125); -return x_1165; } else { -lean_object* x_1168; lean_object* x_1169; -x_1168 = lean_ctor_get(x_1165, 1); -lean_inc(x_1168); -lean_dec(x_1165); -x_1169 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1169, 0, x_1125); -lean_ctor_set(x_1169, 1, x_1168); -return x_1169; -} +lean_object* x_40; +lean_inc(x_25); +x_40 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2(x_23, x_24, x_25, x_26, x_29, x_27, x_25, x_3, x_4, x_5, x_6, x_7, x_30); +return x_40; } } else { -lean_object* x_1170; lean_object* x_1171; lean_object* x_1172; lean_object* x_1173; lean_object* x_1174; uint8_t x_1175; -x_1170 = lean_box(0); -x_1171 = lean_array_uset(x_1131, x_1147, x_1170); -lean_inc(x_1125); -x_1172 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__8(x_1, x_1125, x_1148); -x_1173 = lean_array_uset(x_1171, x_1147, x_1172); -lean_ctor_set(x_1127, 1, x_1173); -x_1174 = lean_st_ref_set(x_5, x_1127, x_1128); +uint8_t x_41; +lean_dec(x_27); +lean_dec(x_25); +lean_dec(x_24); +lean_dec(x_23); +lean_dec(x_7); +lean_dec(x_6); lean_dec(x_5); -x_1175 = !lean_is_exclusive(x_1174); -if (x_1175 == 0) +lean_dec(x_4); +lean_dec(x_3); +x_41 = !lean_is_exclusive(x_28); +if (x_41 == 0) { -lean_object* x_1176; -x_1176 = lean_ctor_get(x_1174, 0); -lean_dec(x_1176); -lean_ctor_set(x_1174, 0, x_1125); -return x_1174; +return x_28; } else { -lean_object* x_1177; lean_object* x_1178; -x_1177 = lean_ctor_get(x_1174, 1); -lean_inc(x_1177); -lean_dec(x_1174); -x_1178 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1178, 0, x_1125); -lean_ctor_set(x_1178, 1, x_1177); -return x_1178; +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_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 +case 11: { -lean_object* x_1179; lean_object* x_1180; lean_object* x_1181; size_t x_1182; uint64_t x_1183; uint64_t x_1184; uint64_t x_1185; uint64_t x_1186; uint64_t x_1187; uint64_t x_1188; uint64_t x_1189; uint64_t x_1190; uint64_t x_1191; size_t x_1192; size_t x_1193; size_t x_1194; size_t x_1195; size_t x_1196; lean_object* x_1197; uint8_t x_1198; -x_1179 = lean_ctor_get(x_1127, 0); -x_1180 = lean_ctor_get(x_1127, 1); -lean_inc(x_1180); -lean_inc(x_1179); -lean_dec(x_1127); -x_1181 = lean_array_get_size(x_1180); -x_1182 = lean_ptr_addr(x_1); -x_1183 = lean_usize_to_uint64(x_1182); -x_1184 = 11; -x_1185 = lean_uint64_mix_hash(x_1183, x_1184); -x_1186 = 32; -x_1187 = lean_uint64_shift_right(x_1185, x_1186); -x_1188 = lean_uint64_xor(x_1185, x_1187); -x_1189 = 16; -x_1190 = lean_uint64_shift_right(x_1188, x_1189); -x_1191 = lean_uint64_xor(x_1188, x_1190); -x_1192 = lean_uint64_to_usize(x_1191); -x_1193 = lean_usize_of_nat(x_1181); -lean_dec(x_1181); -x_1194 = 1; -x_1195 = lean_usize_sub(x_1193, x_1194); -x_1196 = lean_usize_land(x_1192, x_1195); -x_1197 = lean_array_uget(x_1180, x_1196); -x_1198 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__3(x_1, x_1197); -if (x_1198 == 0) -{ -lean_object* x_1199; lean_object* x_1200; lean_object* x_1201; lean_object* x_1202; lean_object* x_1203; lean_object* x_1204; lean_object* x_1205; lean_object* x_1206; uint8_t x_1207; -x_1199 = lean_nat_add(x_1179, x_1117); -lean_dec(x_1179); -lean_inc(x_1125); -x_1200 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1200, 0, x_1); -lean_ctor_set(x_1200, 1, x_1125); -lean_ctor_set(x_1200, 2, x_1197); -x_1201 = lean_array_uset(x_1180, x_1196, x_1200); -x_1202 = lean_unsigned_to_nat(4u); -x_1203 = lean_nat_mul(x_1199, x_1202); -x_1204 = lean_unsigned_to_nat(3u); -x_1205 = lean_nat_div(x_1203, x_1204); -lean_dec(x_1203); -x_1206 = lean_array_get_size(x_1201); -x_1207 = lean_nat_dec_le(x_1205, x_1206); -lean_dec(x_1206); -lean_dec(x_1205); -if (x_1207 == 0) -{ -lean_object* x_1208; lean_object* x_1209; lean_object* x_1210; lean_object* x_1211; lean_object* x_1212; lean_object* x_1213; -x_1208 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__4(x_1201); -x_1209 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1209, 0, x_1199); -lean_ctor_set(x_1209, 1, x_1208); -x_1210 = lean_st_ref_set(x_5, x_1209, x_1128); -lean_dec(x_5); -x_1211 = lean_ctor_get(x_1210, 1); -lean_inc(x_1211); -if (lean_is_exclusive(x_1210)) { - lean_ctor_release(x_1210, 0); - lean_ctor_release(x_1210, 1); - x_1212 = x_1210; -} else { - lean_dec_ref(x_1210); - x_1212 = lean_box(0); -} -if (lean_is_scalar(x_1212)) { - x_1213 = lean_alloc_ctor(0, 2, 0); -} else { - x_1213 = x_1212; -} -lean_ctor_set(x_1213, 0, x_1125); -lean_ctor_set(x_1213, 1, x_1211); -return x_1213; -} -else +lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_45 = lean_ctor_get(x_1, 0); +lean_inc(x_45); +x_46 = lean_ctor_get(x_1, 1); +lean_inc(x_46); +x_47 = lean_ctor_get(x_1, 2); +lean_inc(x_47); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_47); +x_48 = l_Lean_Meta_Grind_markNestedProofsImpl_visit(x_47, x_3, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_48) == 0) +{ +lean_object* x_49; lean_object* x_50; size_t x_51; size_t x_52; uint8_t x_53; +x_49 = lean_ctor_get(x_48, 0); +lean_inc(x_49); +x_50 = lean_ctor_get(x_48, 1); +lean_inc(x_50); +lean_dec(x_48); +x_51 = lean_ptr_addr(x_47); +lean_dec(x_47); +x_52 = lean_ptr_addr(x_49); +x_53 = lean_usize_dec_eq(x_51, x_52); +if (x_53 == 0) { -lean_object* x_1214; lean_object* x_1215; lean_object* x_1216; lean_object* x_1217; lean_object* x_1218; -x_1214 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1214, 0, x_1199); -lean_ctor_set(x_1214, 1, x_1201); -x_1215 = lean_st_ref_set(x_5, x_1214, x_1128); +lean_object* x_54; lean_object* x_55; +x_54 = l_Lean_Expr_proj___override(x_45, x_46, x_49); +x_55 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__1(x_1, x_54, x_3, x_4, x_5, x_6, x_7, x_50); +lean_dec(x_7); +lean_dec(x_6); lean_dec(x_5); -x_1216 = lean_ctor_get(x_1215, 1); -lean_inc(x_1216); -if (lean_is_exclusive(x_1215)) { - lean_ctor_release(x_1215, 0); - lean_ctor_release(x_1215, 1); - x_1217 = x_1215; -} else { - lean_dec_ref(x_1215); - x_1217 = lean_box(0); -} -if (lean_is_scalar(x_1217)) { - x_1218 = lean_alloc_ctor(0, 2, 0); -} else { - x_1218 = x_1217; -} -lean_ctor_set(x_1218, 0, x_1125); -lean_ctor_set(x_1218, 1, x_1216); -return x_1218; -} +lean_dec(x_4); +lean_dec(x_3); +return x_55; } else { -lean_object* x_1219; lean_object* x_1220; lean_object* x_1221; lean_object* x_1222; lean_object* x_1223; lean_object* x_1224; lean_object* x_1225; lean_object* x_1226; lean_object* x_1227; -x_1219 = lean_box(0); -x_1220 = lean_array_uset(x_1180, x_1196, x_1219); -lean_inc(x_1125); -x_1221 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__8(x_1, x_1125, x_1197); -x_1222 = lean_array_uset(x_1220, x_1196, x_1221); -x_1223 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1223, 0, x_1179); -lean_ctor_set(x_1223, 1, x_1222); -x_1224 = lean_st_ref_set(x_5, x_1223, x_1128); +lean_object* x_56; +lean_dec(x_49); +lean_dec(x_46); +lean_dec(x_45); +lean_inc(x_1); +x_56 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__1(x_1, x_1, x_3, x_4, x_5, x_6, x_7, x_50); +lean_dec(x_7); +lean_dec(x_6); lean_dec(x_5); -x_1225 = lean_ctor_get(x_1224, 1); -lean_inc(x_1225); -if (lean_is_exclusive(x_1224)) { - lean_ctor_release(x_1224, 0); - lean_ctor_release(x_1224, 1); - x_1226 = x_1224; -} else { - lean_dec_ref(x_1224); - x_1226 = lean_box(0); -} -if (lean_is_scalar(x_1226)) { - x_1227 = lean_alloc_ctor(0, 2, 0); -} else { - x_1227 = x_1226; -} -lean_ctor_set(x_1227, 0, x_1125); -lean_ctor_set(x_1227, 1, x_1225); -return x_1227; -} -} +lean_dec(x_4); +lean_dec(x_3); +return x_56; } } else { -uint8_t x_1233; +uint8_t x_57; +lean_dec(x_47); +lean_dec(x_46); +lean_dec(x_45); +lean_dec(x_7); +lean_dec(x_6); lean_dec(x_5); -lean_dec(x_2); +lean_dec(x_4); +lean_dec(x_3); lean_dec(x_1); -x_1233 = !lean_is_exclusive(x_1122); -if (x_1233 == 0) +x_57 = !lean_is_exclusive(x_48); +if (x_57 == 0) { -return x_1122; +return x_48; } else { -lean_object* x_1234; lean_object* x_1235; lean_object* x_1236; -x_1234 = lean_ctor_get(x_1122, 0); -x_1235 = lean_ctor_get(x_1122, 1); -lean_inc(x_1235); -lean_inc(x_1234); -lean_dec(x_1122); -x_1236 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_1236, 0, x_1234); -lean_ctor_set(x_1236, 1, x_1235); -return x_1236; +lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_58 = lean_ctor_get(x_48, 0); +x_59 = lean_ctor_get(x_48, 1); +lean_inc(x_59); +lean_inc(x_58); +lean_dec(x_48); +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; } } } default: { -lean_object* x_1237; lean_object* x_1238; lean_object* x_1239; lean_object* x_1240; uint8_t x_1241; lean_object* x_1242; lean_object* x_1243; lean_object* x_1244; -lean_dec(x_4); -x_1237 = lean_array_get_size(x_3); -x_1238 = lean_unsigned_to_nat(0u); -x_1239 = lean_unsigned_to_nat(1u); -x_1240 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_1240, 0, x_1238); -lean_ctor_set(x_1240, 1, x_1237); -lean_ctor_set(x_1240, 2, x_1239); -x_1241 = 0; -x_1242 = lean_box(x_1241); -x_1243 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1243, 0, x_3); -lean_ctor_set(x_1243, 1, x_1242); +lean_object* x_61; lean_object* x_62; +x_61 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3___closed__4; +lean_inc(x_7); +lean_inc(x_6); lean_inc(x_5); -x_1244 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__2(x_1240, x_1240, x_1243, x_1238, lean_box(0), lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_1240); -if (lean_obj_tag(x_1244) == 0) -{ -lean_object* x_1245; lean_object* x_1246; lean_object* x_1247; lean_object* x_1351; uint8_t x_1352; -x_1245 = lean_ctor_get(x_1244, 0); -lean_inc(x_1245); -x_1246 = lean_ctor_get(x_1244, 1); -lean_inc(x_1246); -lean_dec(x_1244); -x_1351 = lean_ctor_get(x_1245, 1); -lean_inc(x_1351); -x_1352 = lean_unbox(x_1351); -lean_dec(x_1351); -if (x_1352 == 0) -{ -lean_dec(x_1245); -lean_dec(x_2); -lean_inc(x_1); -x_1247 = x_1; -goto block_1350; +lean_inc(x_4); +lean_inc(x_3); +x_62 = l_panic___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__7(x_61, x_3, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_62) == 0) +{ +lean_object* x_63; lean_object* x_64; lean_object* x_65; +x_63 = lean_ctor_get(x_62, 0); +lean_inc(x_63); +x_64 = lean_ctor_get(x_62, 1); +lean_inc(x_64); +lean_dec(x_62); +x_65 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__1(x_1, x_63, x_3, x_4, x_5, x_6, x_7, x_64); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_65; } else { -lean_object* x_1353; lean_object* x_1354; -x_1353 = lean_ctor_get(x_1245, 0); -lean_inc(x_1353); -lean_dec(x_1245); -x_1354 = l_Lean_mkAppN(x_2, x_1353); -lean_dec(x_1353); -x_1247 = x_1354; -goto block_1350; -} -block_1350: -{ -lean_object* x_1248; lean_object* x_1249; lean_object* x_1250; uint8_t x_1251; -x_1248 = lean_st_ref_take(x_5, x_1246); -x_1249 = lean_ctor_get(x_1248, 0); -lean_inc(x_1249); -x_1250 = lean_ctor_get(x_1248, 1); -lean_inc(x_1250); -lean_dec(x_1248); -x_1251 = !lean_is_exclusive(x_1249); -if (x_1251 == 0) -{ -lean_object* x_1252; lean_object* x_1253; lean_object* x_1254; size_t x_1255; uint64_t x_1256; uint64_t x_1257; uint64_t x_1258; uint64_t x_1259; uint64_t x_1260; uint64_t x_1261; uint64_t x_1262; uint64_t x_1263; uint64_t x_1264; size_t x_1265; size_t x_1266; size_t x_1267; size_t x_1268; size_t x_1269; lean_object* x_1270; uint8_t x_1271; -x_1252 = lean_ctor_get(x_1249, 0); -x_1253 = lean_ctor_get(x_1249, 1); -x_1254 = lean_array_get_size(x_1253); -x_1255 = lean_ptr_addr(x_1); -x_1256 = lean_usize_to_uint64(x_1255); -x_1257 = 11; -x_1258 = lean_uint64_mix_hash(x_1256, x_1257); -x_1259 = 32; -x_1260 = lean_uint64_shift_right(x_1258, x_1259); -x_1261 = lean_uint64_xor(x_1258, x_1260); -x_1262 = 16; -x_1263 = lean_uint64_shift_right(x_1261, x_1262); -x_1264 = lean_uint64_xor(x_1261, x_1263); -x_1265 = lean_uint64_to_usize(x_1264); -x_1266 = lean_usize_of_nat(x_1254); -lean_dec(x_1254); -x_1267 = 1; -x_1268 = lean_usize_sub(x_1266, x_1267); -x_1269 = lean_usize_land(x_1265, x_1268); -x_1270 = lean_array_uget(x_1253, x_1269); -x_1271 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__3(x_1, x_1270); -if (x_1271 == 0) -{ -lean_object* x_1272; lean_object* x_1273; lean_object* x_1274; lean_object* x_1275; lean_object* x_1276; lean_object* x_1277; lean_object* x_1278; lean_object* x_1279; uint8_t x_1280; -x_1272 = lean_nat_add(x_1252, x_1239); -lean_dec(x_1252); -lean_inc(x_1247); -x_1273 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1273, 0, x_1); -lean_ctor_set(x_1273, 1, x_1247); -lean_ctor_set(x_1273, 2, x_1270); -x_1274 = lean_array_uset(x_1253, x_1269, x_1273); -x_1275 = lean_unsigned_to_nat(4u); -x_1276 = lean_nat_mul(x_1272, x_1275); -x_1277 = lean_unsigned_to_nat(3u); -x_1278 = lean_nat_div(x_1276, x_1277); -lean_dec(x_1276); -x_1279 = lean_array_get_size(x_1274); -x_1280 = lean_nat_dec_le(x_1278, x_1279); -lean_dec(x_1279); -lean_dec(x_1278); -if (x_1280 == 0) -{ -lean_object* x_1281; lean_object* x_1282; uint8_t x_1283; -x_1281 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__4(x_1274); -lean_ctor_set(x_1249, 1, x_1281); -lean_ctor_set(x_1249, 0, x_1272); -x_1282 = lean_st_ref_set(x_5, x_1249, x_1250); +uint8_t x_66; +lean_dec(x_7); +lean_dec(x_6); lean_dec(x_5); -x_1283 = !lean_is_exclusive(x_1282); -if (x_1283 == 0) +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_66 = !lean_is_exclusive(x_62); +if (x_66 == 0) { -lean_object* x_1284; -x_1284 = lean_ctor_get(x_1282, 0); -lean_dec(x_1284); -lean_ctor_set(x_1282, 0, x_1247); -return x_1282; +return x_62; } else { -lean_object* x_1285; lean_object* x_1286; -x_1285 = lean_ctor_get(x_1282, 1); -lean_inc(x_1285); -lean_dec(x_1282); -x_1286 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1286, 0, x_1247); -lean_ctor_set(x_1286, 1, x_1285); -return x_1286; +lean_object* x_67; lean_object* x_68; lean_object* x_69; +x_67 = lean_ctor_get(x_62, 0); +x_68 = lean_ctor_get(x_62, 1); +lean_inc(x_68); +lean_inc(x_67); +lean_dec(x_62); +x_69 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_69, 0, x_67); +lean_ctor_set(x_69, 1, x_68); +return x_69; } } -else -{ -lean_object* x_1287; uint8_t x_1288; -lean_ctor_set(x_1249, 1, x_1274); -lean_ctor_set(x_1249, 0, x_1272); -x_1287 = lean_st_ref_set(x_5, x_1249, x_1250); -lean_dec(x_5); -x_1288 = !lean_is_exclusive(x_1287); -if (x_1288 == 0) -{ -lean_object* x_1289; -x_1289 = lean_ctor_get(x_1287, 0); -lean_dec(x_1289); -lean_ctor_set(x_1287, 0, x_1247); -return x_1287; } -else -{ -lean_object* x_1290; lean_object* x_1291; -x_1290 = lean_ctor_get(x_1287, 1); -lean_inc(x_1290); -lean_dec(x_1287); -x_1291 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1291, 0, x_1247); -lean_ctor_set(x_1291, 1, x_1290); -return x_1291; } } } -else +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___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) { +_start: { -lean_object* x_1292; lean_object* x_1293; lean_object* x_1294; lean_object* x_1295; lean_object* x_1296; uint8_t x_1297; -x_1292 = lean_box(0); -x_1293 = lean_array_uset(x_1253, x_1269, x_1292); -lean_inc(x_1247); -x_1294 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__8(x_1, x_1247, x_1270); -x_1295 = lean_array_uset(x_1293, x_1269, x_1294); -lean_ctor_set(x_1249, 1, x_1295); -x_1296 = lean_st_ref_set(x_5, x_1249, x_1250); -lean_dec(x_5); -x_1297 = !lean_is_exclusive(x_1296); -if (x_1297 == 0) +lean_object* x_9; uint8_t x_10; +x_9 = lean_st_ref_get(x_3, x_8); +x_10 = !lean_is_exclusive(x_9); +if (x_10 == 0) { -lean_object* x_1298; -x_1298 = lean_ctor_get(x_1296, 0); -lean_dec(x_1298); -lean_ctor_set(x_1296, 0, x_1247); -return x_1296; -} -else +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; size_t x_15; uint64_t x_16; uint64_t x_17; uint64_t x_18; uint64_t x_19; uint64_t x_20; uint64_t x_21; uint64_t x_22; uint64_t x_23; uint64_t x_24; size_t x_25; size_t x_26; size_t x_27; size_t x_28; size_t x_29; lean_object* x_30; lean_object* x_31; +x_11 = lean_ctor_get(x_9, 0); +x_12 = lean_ctor_get(x_9, 1); +x_13 = lean_ctor_get(x_11, 1); +lean_inc(x_13); +lean_dec(x_11); +x_14 = lean_array_get_size(x_13); +x_15 = lean_ptr_addr(x_1); +x_16 = lean_usize_to_uint64(x_15); +x_17 = 11; +x_18 = lean_uint64_mix_hash(x_16, x_17); +x_19 = 32; +x_20 = lean_uint64_shift_right(x_18, x_19); +x_21 = lean_uint64_xor(x_18, x_20); +x_22 = 16; +x_23 = lean_uint64_shift_right(x_21, x_22); +x_24 = lean_uint64_xor(x_21, x_23); +x_25 = lean_uint64_to_usize(x_24); +x_26 = lean_usize_of_nat(x_14); +lean_dec(x_14); +x_27 = 1; +x_28 = lean_usize_sub(x_26, x_27); +x_29 = lean_usize_land(x_25, x_28); +x_30 = lean_array_uget(x_13, x_29); +lean_dec(x_13); +x_31 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__10(x_1, x_30); +lean_dec(x_30); +if (lean_obj_tag(x_31) == 0) { -lean_object* x_1299; lean_object* x_1300; -x_1299 = lean_ctor_get(x_1296, 1); -lean_inc(x_1299); -lean_dec(x_1296); -x_1300 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1300, 0, x_1247); -lean_ctor_set(x_1300, 1, x_1299); -return x_1300; -} -} +lean_object* x_32; lean_object* x_33; +lean_free_object(x_9); +x_32 = lean_box(0); +x_33 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3(x_1, x_32, x_3, x_4, x_5, x_6, x_7, x_12); +return x_33; } else { -lean_object* x_1301; lean_object* x_1302; lean_object* x_1303; size_t x_1304; uint64_t x_1305; uint64_t x_1306; uint64_t x_1307; uint64_t x_1308; uint64_t x_1309; uint64_t x_1310; uint64_t x_1311; uint64_t x_1312; uint64_t x_1313; size_t x_1314; size_t x_1315; size_t x_1316; size_t x_1317; size_t x_1318; lean_object* x_1319; uint8_t x_1320; -x_1301 = lean_ctor_get(x_1249, 0); -x_1302 = lean_ctor_get(x_1249, 1); -lean_inc(x_1302); -lean_inc(x_1301); -lean_dec(x_1249); -x_1303 = lean_array_get_size(x_1302); -x_1304 = lean_ptr_addr(x_1); -x_1305 = lean_usize_to_uint64(x_1304); -x_1306 = 11; -x_1307 = lean_uint64_mix_hash(x_1305, x_1306); -x_1308 = 32; -x_1309 = lean_uint64_shift_right(x_1307, x_1308); -x_1310 = lean_uint64_xor(x_1307, x_1309); -x_1311 = 16; -x_1312 = lean_uint64_shift_right(x_1310, x_1311); -x_1313 = lean_uint64_xor(x_1310, x_1312); -x_1314 = lean_uint64_to_usize(x_1313); -x_1315 = lean_usize_of_nat(x_1303); -lean_dec(x_1303); -x_1316 = 1; -x_1317 = lean_usize_sub(x_1315, x_1316); -x_1318 = lean_usize_land(x_1314, x_1317); -x_1319 = lean_array_uget(x_1302, x_1318); -x_1320 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__3(x_1, x_1319); -if (x_1320 == 0) -{ -lean_object* x_1321; lean_object* x_1322; lean_object* x_1323; lean_object* x_1324; lean_object* x_1325; lean_object* x_1326; lean_object* x_1327; lean_object* x_1328; uint8_t x_1329; -x_1321 = lean_nat_add(x_1301, x_1239); -lean_dec(x_1301); -lean_inc(x_1247); -x_1322 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1322, 0, x_1); -lean_ctor_set(x_1322, 1, x_1247); -lean_ctor_set(x_1322, 2, x_1319); -x_1323 = lean_array_uset(x_1302, x_1318, x_1322); -x_1324 = lean_unsigned_to_nat(4u); -x_1325 = lean_nat_mul(x_1321, x_1324); -x_1326 = lean_unsigned_to_nat(3u); -x_1327 = lean_nat_div(x_1325, x_1326); -lean_dec(x_1325); -x_1328 = lean_array_get_size(x_1323); -x_1329 = lean_nat_dec_le(x_1327, x_1328); -lean_dec(x_1328); -lean_dec(x_1327); -if (x_1329 == 0) -{ -lean_object* x_1330; lean_object* x_1331; lean_object* x_1332; lean_object* x_1333; lean_object* x_1334; lean_object* x_1335; -x_1330 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__4(x_1323); -x_1331 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1331, 0, x_1321); -lean_ctor_set(x_1331, 1, x_1330); -x_1332 = lean_st_ref_set(x_5, x_1331, x_1250); +lean_object* x_34; +lean_dec(x_7); +lean_dec(x_6); lean_dec(x_5); -x_1333 = lean_ctor_get(x_1332, 1); -lean_inc(x_1333); -if (lean_is_exclusive(x_1332)) { - lean_ctor_release(x_1332, 0); - lean_ctor_release(x_1332, 1); - x_1334 = x_1332; -} else { - lean_dec_ref(x_1332); - x_1334 = lean_box(0); -} -if (lean_is_scalar(x_1334)) { - x_1335 = lean_alloc_ctor(0, 2, 0); -} else { - x_1335 = x_1334; +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_34 = lean_ctor_get(x_31, 0); +lean_inc(x_34); +lean_dec(x_31); +lean_ctor_set(x_9, 0, x_34); +return x_9; } -lean_ctor_set(x_1335, 0, x_1247); -lean_ctor_set(x_1335, 1, x_1333); -return x_1335; } else { -lean_object* x_1336; lean_object* x_1337; lean_object* x_1338; lean_object* x_1339; lean_object* x_1340; -x_1336 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1336, 0, x_1321); -lean_ctor_set(x_1336, 1, x_1323); -x_1337 = lean_st_ref_set(x_5, x_1336, x_1250); -lean_dec(x_5); -x_1338 = lean_ctor_get(x_1337, 1); -lean_inc(x_1338); -if (lean_is_exclusive(x_1337)) { - lean_ctor_release(x_1337, 0); - lean_ctor_release(x_1337, 1); - x_1339 = x_1337; -} else { - lean_dec_ref(x_1337); - x_1339 = lean_box(0); -} -if (lean_is_scalar(x_1339)) { - x_1340 = lean_alloc_ctor(0, 2, 0); -} else { - x_1340 = x_1339; -} -lean_ctor_set(x_1340, 0, x_1247); -lean_ctor_set(x_1340, 1, x_1338); -return x_1340; -} -} -else +lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; size_t x_39; uint64_t x_40; uint64_t x_41; uint64_t x_42; uint64_t x_43; uint64_t x_44; uint64_t x_45; uint64_t x_46; uint64_t x_47; uint64_t x_48; size_t x_49; size_t x_50; size_t x_51; size_t x_52; size_t x_53; lean_object* x_54; lean_object* x_55; +x_35 = lean_ctor_get(x_9, 0); +x_36 = lean_ctor_get(x_9, 1); +lean_inc(x_36); +lean_inc(x_35); +lean_dec(x_9); +x_37 = lean_ctor_get(x_35, 1); +lean_inc(x_37); +lean_dec(x_35); +x_38 = lean_array_get_size(x_37); +x_39 = lean_ptr_addr(x_1); +x_40 = lean_usize_to_uint64(x_39); +x_41 = 11; +x_42 = lean_uint64_mix_hash(x_40, x_41); +x_43 = 32; +x_44 = lean_uint64_shift_right(x_42, x_43); +x_45 = lean_uint64_xor(x_42, x_44); +x_46 = 16; +x_47 = lean_uint64_shift_right(x_45, x_46); +x_48 = lean_uint64_xor(x_45, x_47); +x_49 = lean_uint64_to_usize(x_48); +x_50 = lean_usize_of_nat(x_38); +lean_dec(x_38); +x_51 = 1; +x_52 = lean_usize_sub(x_50, x_51); +x_53 = lean_usize_land(x_49, x_52); +x_54 = lean_array_uget(x_37, x_53); +lean_dec(x_37); +x_55 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__10(x_1, x_54); +lean_dec(x_54); +if (lean_obj_tag(x_55) == 0) { -lean_object* x_1341; lean_object* x_1342; lean_object* x_1343; lean_object* x_1344; lean_object* x_1345; lean_object* x_1346; lean_object* x_1347; lean_object* x_1348; lean_object* x_1349; -x_1341 = lean_box(0); -x_1342 = lean_array_uset(x_1302, x_1318, x_1341); -lean_inc(x_1247); -x_1343 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__8(x_1, x_1247, x_1319); -x_1344 = lean_array_uset(x_1342, x_1318, x_1343); -x_1345 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1345, 0, x_1301); -lean_ctor_set(x_1345, 1, x_1344); -x_1346 = lean_st_ref_set(x_5, x_1345, x_1250); -lean_dec(x_5); -x_1347 = lean_ctor_get(x_1346, 1); -lean_inc(x_1347); -if (lean_is_exclusive(x_1346)) { - lean_ctor_release(x_1346, 0); - lean_ctor_release(x_1346, 1); - x_1348 = x_1346; -} else { - lean_dec_ref(x_1346); - x_1348 = lean_box(0); -} -if (lean_is_scalar(x_1348)) { - x_1349 = lean_alloc_ctor(0, 2, 0); -} else { - x_1349 = x_1348; -} -lean_ctor_set(x_1349, 0, x_1247); -lean_ctor_set(x_1349, 1, x_1347); -return x_1349; -} -} -} +lean_object* x_56; lean_object* x_57; +x_56 = lean_box(0); +x_57 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3(x_1, x_56, x_3, x_4, x_5, x_6, x_7, x_36); +return x_57; } else { -uint8_t x_1355; +lean_object* x_58; lean_object* x_59; +lean_dec(x_7); +lean_dec(x_6); lean_dec(x_5); -lean_dec(x_2); +lean_dec(x_4); +lean_dec(x_3); lean_dec(x_1); -x_1355 = !lean_is_exclusive(x_1244); -if (x_1355 == 0) -{ -return x_1244; -} -else -{ -lean_object* x_1356; lean_object* x_1357; lean_object* x_1358; -x_1356 = lean_ctor_get(x_1244, 0); -x_1357 = lean_ctor_get(x_1244, 1); -lean_inc(x_1357); -lean_inc(x_1356); -lean_dec(x_1244); -x_1358 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_1358, 0, x_1356); -lean_ctor_set(x_1358, 1, x_1357); -return x_1358; -} -} +x_58 = lean_ctor_get(x_55, 0); +lean_inc(x_58); +lean_dec(x_55); +x_59 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_59, 0, x_58); +lean_ctor_set(x_59, 1, x_36); +return x_59; } } } } -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__10(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { -if (lean_obj_tag(x_2) == 0) +uint8_t x_9; +x_9 = l_Lean_Expr_isApp(x_1); +if (x_9 == 0) { -lean_object* x_3; -x_3 = lean_box(0); -return x_3; -} -else +uint8_t x_10; +x_10 = l_Lean_Expr_isForall(x_1); +if (x_10 == 0) { -lean_object* x_4; lean_object* x_5; lean_object* x_6; size_t x_7; size_t x_8; uint8_t x_9; -x_4 = lean_ctor_get(x_2, 0); -x_5 = lean_ctor_get(x_2, 1); -x_6 = lean_ctor_get(x_2, 2); -x_7 = lean_ptr_addr(x_4); -x_8 = lean_ptr_addr(x_1); -x_9 = lean_usize_dec_eq(x_7, x_8); -if (x_9 == 0) +uint8_t x_11; +x_11 = l_Lean_Expr_isProj(x_1); +if (x_11 == 0) { -x_2 = x_6; -goto _start; +lean_object* x_12; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_1); +lean_ctor_set(x_12, 1, x_8); +return x_12; } else { -lean_object* x_11; -lean_inc(x_5); -x_11 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_11, 0, x_5); -return x_11; -} -} +lean_object* x_13; lean_object* x_14; +x_13 = lean_box(0); +x_14 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__4(x_1, x_13, x_3, x_4, x_5, x_6, x_7, x_8); +return x_14; } } -static lean_object* _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__1___closed__1() { -_start: +else { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_levelZero; -x_2 = l_Lean_Expr_sort___override(x_1); -return x_2; +lean_object* x_15; lean_object* x_16; +x_15 = lean_box(0); +x_16 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__4(x_1, x_15, x_3, x_4, x_5, x_6, x_7, x_8); +return x_16; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: +else { -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_9 = lean_unsigned_to_nat(0u); -x_10 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_9); -x_11 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__1___closed__1; -lean_inc(x_10); -x_12 = lean_mk_array(x_10, x_11); -x_13 = lean_unsigned_to_nat(1u); -x_14 = lean_nat_sub(x_10, x_13); -lean_dec(x_10); -lean_inc(x_1); -x_15 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__9(x_1, x_1, x_12, x_14, x_3, x_4, x_5, x_6, x_7, x_8); -return x_15; +lean_object* x_17; lean_object* x_18; +x_17 = lean_box(0); +x_18 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__4(x_1, x_17, x_3, x_4, x_5, x_6, x_7, x_8); +return x_18; } } -static lean_object* _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2___closed__1() { +} +static lean_object* _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__6___closed__1() { _start: { lean_object* x_1; @@ -4887,7 +1928,7 @@ x_1 = lean_mk_string_unchecked("Lean", 4, 4); return x_1; } } -static lean_object* _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2___closed__2() { +static lean_object* _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__6___closed__2() { _start: { lean_object* x_1; @@ -4895,7 +1936,7 @@ x_1 = lean_mk_string_unchecked("Grind", 5, 5); return x_1; } } -static lean_object* _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2___closed__3() { +static lean_object* _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__6___closed__3() { _start: { lean_object* x_1; @@ -4903,28 +1944,28 @@ x_1 = lean_mk_string_unchecked("nestedProof", 11, 11); return x_1; } } -static lean_object* _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2___closed__4() { +static lean_object* _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__6___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2___closed__1; -x_2 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2___closed__2; -x_3 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2___closed__3; +x_1 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__6___closed__1; +x_2 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__6___closed__2; +x_3 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__6___closed__3; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2___closed__5() { +static lean_object* _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__6___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2___closed__4; +x_2 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__6___closed__4; x_3 = l_Lean_Expr_const___override(x_2, x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { lean_object* x_9; @@ -4938,7 +1979,7 @@ lean_inc(x_10); x_11 = lean_ctor_get(x_9, 1); lean_inc(x_11); lean_dec(x_9); -x_12 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2___closed__5; +x_12 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__6___closed__5; lean_inc(x_1); x_13 = l_Lean_mkAppB(x_12, x_10, x_1); x_14 = lean_st_ref_take(x_3, x_11); @@ -4971,7 +2012,7 @@ x_33 = 1; x_34 = lean_usize_sub(x_32, x_33); x_35 = lean_usize_land(x_31, x_34); x_36 = lean_array_uget(x_19, x_35); -x_37 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__3(x_1, x_36); +x_37 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__1(x_1, x_36); if (x_37 == 0) { lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; uint8_t x_47; @@ -4996,7 +2037,7 @@ lean_dec(x_45); if (x_47 == 0) { lean_object* x_48; lean_object* x_49; uint8_t x_50; -x_48 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__4(x_41); +x_48 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__2(x_41); lean_ctor_set(x_15, 1, x_48); lean_ctor_set(x_15, 0, x_39); x_49 = lean_st_ref_set(x_3, x_15, x_16); @@ -5055,7 +2096,7 @@ lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean x_59 = lean_box(0); x_60 = lean_array_uset(x_19, x_35, x_59); lean_inc(x_13); -x_61 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__8(x_1, x_13, x_36); +x_61 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__6(x_1, x_13, x_36); x_62 = lean_array_uset(x_60, x_35, x_61); lean_ctor_set(x_15, 1, x_62); x_63 = lean_st_ref_set(x_3, x_15, x_16); @@ -5107,7 +2148,7 @@ x_83 = 1; x_84 = lean_usize_sub(x_82, x_83); x_85 = lean_usize_land(x_81, x_84); x_86 = lean_array_uget(x_69, x_85); -x_87 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__3(x_1, x_86); +x_87 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__1(x_1, x_86); if (x_87 == 0) { lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; uint8_t x_97; @@ -5132,7 +2173,7 @@ lean_dec(x_95); if (x_97 == 0) { lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; -x_98 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__4(x_91); +x_98 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__2(x_91); x_99 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_99, 0, x_89); lean_ctor_set(x_99, 1, x_98); @@ -5189,7 +2230,7 @@ lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; x_109 = lean_box(0); x_110 = lean_array_uset(x_69, x_85, x_109); lean_inc(x_13); -x_111 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__8(x_1, x_13, x_86); +x_111 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__6(x_1, x_13, x_86); x_112 = lean_array_uset(x_110, x_85, x_111); x_113 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_113, 0, x_68); @@ -5241,7 +2282,7 @@ return x_121; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___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; uint8_t x_10; @@ -5281,7 +2322,7 @@ if (lean_obj_tag(x_31) == 0) lean_object* x_32; lean_object* x_33; lean_free_object(x_9); x_32 = lean_box(0); -x_33 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2(x_1, x_32, x_3, x_4, x_5, x_6, x_7, x_12); +x_33 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__6(x_1, x_32, x_3, x_4, x_5, x_6, x_7, x_12); return x_33; } else @@ -5335,7 +2376,7 @@ if (lean_obj_tag(x_55) == 0) { lean_object* x_56; lean_object* x_57; x_56 = lean_box(0); -x_57 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2(x_1, x_56, x_3, x_4, x_5, x_6, x_7, x_36); +x_57 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__6(x_1, x_56, x_3, x_4, x_5, x_6, x_7, x_36); return x_57; } else @@ -5357,43 +2398,6 @@ return x_59; } } } -static lean_object* _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean.Meta.Tactic.Grind.MarkNestedProofs", 39, 39); -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean.Meta.Grind.markNestedProofsImpl.visit", 42, 42); -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___closed__3() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("unreachable code has been reached", 33, 33); -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___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_Meta_Grind_markNestedProofsImpl_visit___closed__1; -x_2 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___closed__2; -x_3 = lean_unsigned_to_nat(28u); -x_4 = lean_unsigned_to_nat(20u); -x_5 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___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_Meta_Grind_markNestedProofsImpl_visit(lean_object* x_1, 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: { @@ -5413,393 +2417,34 @@ x_10 = lean_unbox(x_9); lean_dec(x_9); if (x_10 == 0) { -switch (lean_obj_tag(x_1)) { -case 0: -{ lean_object* x_11; lean_object* x_12; lean_object* x_13; -lean_dec(x_1); x_11 = lean_ctor_get(x_8, 1); lean_inc(x_11); lean_dec(x_8); -x_12 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___closed__4; -x_13 = l_panic___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__1(x_12, x_2, x_3, x_4, x_5, x_6, x_11); +x_12 = lean_box(0); +x_13 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__5(x_1, x_12, x_2, x_3, x_4, x_5, x_6, x_11); return x_13; } -case 5: -{ -lean_object* x_14; lean_object* x_15; uint8_t x_16; -x_14 = lean_ctor_get(x_8, 1); -lean_inc(x_14); -lean_dec(x_8); -x_15 = lean_st_ref_get(x_2, x_14); -x_16 = !lean_is_exclusive(x_15); -if (x_16 == 0) -{ -lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; size_t x_21; uint64_t x_22; uint64_t x_23; uint64_t x_24; uint64_t x_25; uint64_t x_26; uint64_t x_27; uint64_t x_28; uint64_t x_29; uint64_t x_30; size_t x_31; size_t x_32; size_t x_33; size_t x_34; size_t x_35; lean_object* x_36; lean_object* x_37; -x_17 = lean_ctor_get(x_15, 0); -x_18 = lean_ctor_get(x_15, 1); -x_19 = lean_ctor_get(x_17, 1); -lean_inc(x_19); -lean_dec(x_17); -x_20 = lean_array_get_size(x_19); -x_21 = lean_ptr_addr(x_1); -x_22 = lean_usize_to_uint64(x_21); -x_23 = 11; -x_24 = lean_uint64_mix_hash(x_22, x_23); -x_25 = 32; -x_26 = lean_uint64_shift_right(x_24, x_25); -x_27 = lean_uint64_xor(x_24, x_26); -x_28 = 16; -x_29 = lean_uint64_shift_right(x_27, x_28); -x_30 = lean_uint64_xor(x_27, x_29); -x_31 = lean_uint64_to_usize(x_30); -x_32 = lean_usize_of_nat(x_20); -lean_dec(x_20); -x_33 = 1; -x_34 = lean_usize_sub(x_32, x_33); -x_35 = lean_usize_land(x_31, x_34); -x_36 = lean_array_uget(x_19, x_35); -lean_dec(x_19); -x_37 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__10(x_1, x_36); -lean_dec(x_36); -if (lean_obj_tag(x_37) == 0) -{ -lean_object* x_38; lean_object* x_39; -lean_free_object(x_15); -x_38 = lean_box(0); -x_39 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__1(x_1, x_38, x_2, x_3, x_4, x_5, x_6, x_18); -return x_39; -} -else -{ -lean_object* x_40; -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_40 = lean_ctor_get(x_37, 0); -lean_inc(x_40); -lean_dec(x_37); -lean_ctor_set(x_15, 0, x_40); -return x_15; -} -} -else -{ -lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; size_t x_45; uint64_t x_46; uint64_t x_47; uint64_t x_48; uint64_t x_49; uint64_t x_50; uint64_t x_51; uint64_t x_52; uint64_t x_53; uint64_t x_54; size_t x_55; size_t x_56; size_t x_57; size_t x_58; size_t x_59; lean_object* x_60; lean_object* x_61; -x_41 = lean_ctor_get(x_15, 0); -x_42 = lean_ctor_get(x_15, 1); -lean_inc(x_42); -lean_inc(x_41); -lean_dec(x_15); -x_43 = lean_ctor_get(x_41, 1); -lean_inc(x_43); -lean_dec(x_41); -x_44 = lean_array_get_size(x_43); -x_45 = lean_ptr_addr(x_1); -x_46 = lean_usize_to_uint64(x_45); -x_47 = 11; -x_48 = lean_uint64_mix_hash(x_46, x_47); -x_49 = 32; -x_50 = lean_uint64_shift_right(x_48, x_49); -x_51 = lean_uint64_xor(x_48, x_50); -x_52 = 16; -x_53 = lean_uint64_shift_right(x_51, x_52); -x_54 = lean_uint64_xor(x_51, x_53); -x_55 = lean_uint64_to_usize(x_54); -x_56 = lean_usize_of_nat(x_44); -lean_dec(x_44); -x_57 = 1; -x_58 = lean_usize_sub(x_56, x_57); -x_59 = lean_usize_land(x_55, x_58); -x_60 = lean_array_uget(x_43, x_59); -lean_dec(x_43); -x_61 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__10(x_1, x_60); -lean_dec(x_60); -if (lean_obj_tag(x_61) == 0) -{ -lean_object* x_62; lean_object* x_63; -x_62 = lean_box(0); -x_63 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__1(x_1, x_62, x_2, x_3, x_4, x_5, x_6, x_42); -return x_63; -} -else -{ -lean_object* x_64; lean_object* x_65; -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_64 = lean_ctor_get(x_61, 0); -lean_inc(x_64); -lean_dec(x_61); -x_65 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_65, 0, x_64); -lean_ctor_set(x_65, 1, x_42); -return x_65; -} -} -} -case 10: -{ -lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; -x_66 = lean_ctor_get(x_8, 1); -lean_inc(x_66); -lean_dec(x_8); -x_67 = lean_ctor_get(x_1, 0); -lean_inc(x_67); -x_68 = lean_ctor_get(x_1, 1); -lean_inc(x_68); -lean_inc(x_68); -x_69 = l_Lean_Meta_Grind_markNestedProofsImpl_visit(x_68, x_2, x_3, x_4, x_5, x_6, x_66); -if (lean_obj_tag(x_69) == 0) -{ -uint8_t x_70; -x_70 = !lean_is_exclusive(x_69); -if (x_70 == 0) -{ -lean_object* x_71; size_t x_72; size_t x_73; uint8_t x_74; -x_71 = lean_ctor_get(x_69, 0); -x_72 = lean_ptr_addr(x_68); -lean_dec(x_68); -x_73 = lean_ptr_addr(x_71); -x_74 = lean_usize_dec_eq(x_72, x_73); -if (x_74 == 0) -{ -lean_object* x_75; -lean_dec(x_1); -x_75 = l_Lean_Expr_mdata___override(x_67, x_71); -lean_ctor_set(x_69, 0, x_75); -return x_69; -} -else -{ -lean_dec(x_71); -lean_dec(x_67); -lean_ctor_set(x_69, 0, x_1); -return x_69; -} -} -else -{ -lean_object* x_76; lean_object* x_77; size_t x_78; size_t x_79; uint8_t x_80; -x_76 = lean_ctor_get(x_69, 0); -x_77 = lean_ctor_get(x_69, 1); -lean_inc(x_77); -lean_inc(x_76); -lean_dec(x_69); -x_78 = lean_ptr_addr(x_68); -lean_dec(x_68); -x_79 = lean_ptr_addr(x_76); -x_80 = lean_usize_dec_eq(x_78, x_79); -if (x_80 == 0) -{ -lean_object* x_81; lean_object* x_82; -lean_dec(x_1); -x_81 = l_Lean_Expr_mdata___override(x_67, x_76); -x_82 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_82, 0, x_81); -lean_ctor_set(x_82, 1, x_77); -return x_82; -} -else -{ -lean_object* x_83; -lean_dec(x_76); -lean_dec(x_67); -x_83 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_83, 0, x_1); -lean_ctor_set(x_83, 1, x_77); -return x_83; -} -} -} -else -{ -uint8_t x_84; -lean_dec(x_68); -lean_dec(x_67); -lean_dec(x_1); -x_84 = !lean_is_exclusive(x_69); -if (x_84 == 0) -{ -return x_69; -} -else -{ -lean_object* x_85; lean_object* x_86; lean_object* x_87; -x_85 = lean_ctor_get(x_69, 0); -x_86 = lean_ctor_get(x_69, 1); -lean_inc(x_86); -lean_inc(x_85); -lean_dec(x_69); -x_87 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_87, 0, x_85); -lean_ctor_set(x_87, 1, x_86); -return x_87; -} -} -} -case 11: -{ -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_8, 1); -lean_inc(x_88); -lean_dec(x_8); -x_89 = lean_ctor_get(x_1, 0); -lean_inc(x_89); -x_90 = lean_ctor_get(x_1, 1); -lean_inc(x_90); -x_91 = lean_ctor_get(x_1, 2); -lean_inc(x_91); -lean_inc(x_91); -x_92 = l_Lean_Meta_Grind_markNestedProofsImpl_visit(x_91, x_2, x_3, x_4, x_5, x_6, x_88); -if (lean_obj_tag(x_92) == 0) -{ -uint8_t x_93; -x_93 = !lean_is_exclusive(x_92); -if (x_93 == 0) -{ -lean_object* x_94; size_t x_95; size_t x_96; uint8_t x_97; -x_94 = lean_ctor_get(x_92, 0); -x_95 = lean_ptr_addr(x_91); -lean_dec(x_91); -x_96 = lean_ptr_addr(x_94); -x_97 = lean_usize_dec_eq(x_95, x_96); -if (x_97 == 0) -{ -lean_object* x_98; -lean_dec(x_1); -x_98 = l_Lean_Expr_proj___override(x_89, x_90, x_94); -lean_ctor_set(x_92, 0, x_98); -return x_92; -} -else -{ -lean_dec(x_94); -lean_dec(x_90); -lean_dec(x_89); -lean_ctor_set(x_92, 0, x_1); -return x_92; -} -} -else -{ -lean_object* x_99; lean_object* x_100; size_t x_101; size_t x_102; uint8_t x_103; -x_99 = lean_ctor_get(x_92, 0); -x_100 = lean_ctor_get(x_92, 1); -lean_inc(x_100); -lean_inc(x_99); -lean_dec(x_92); -x_101 = lean_ptr_addr(x_91); -lean_dec(x_91); -x_102 = lean_ptr_addr(x_99); -x_103 = lean_usize_dec_eq(x_101, x_102); -if (x_103 == 0) -{ -lean_object* x_104; lean_object* x_105; -lean_dec(x_1); -x_104 = l_Lean_Expr_proj___override(x_89, x_90, x_99); -x_105 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_105, 0, x_104); -lean_ctor_set(x_105, 1, x_100); -return x_105; -} -else -{ -lean_object* x_106; -lean_dec(x_99); -lean_dec(x_90); -lean_dec(x_89); -x_106 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_106, 0, x_1); -lean_ctor_set(x_106, 1, x_100); -return x_106; -} -} -} -else -{ -uint8_t x_107; -lean_dec(x_91); -lean_dec(x_90); -lean_dec(x_89); -lean_dec(x_1); -x_107 = !lean_is_exclusive(x_92); -if (x_107 == 0) -{ -return x_92; -} -else -{ -lean_object* x_108; lean_object* x_109; lean_object* x_110; -x_108 = lean_ctor_get(x_92, 0); -x_109 = lean_ctor_get(x_92, 1); -lean_inc(x_109); -lean_inc(x_108); -lean_dec(x_92); -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; -} -} -} -default: -{ -uint8_t x_111; -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -x_111 = !lean_is_exclusive(x_8); -if (x_111 == 0) -{ -lean_object* x_112; -x_112 = lean_ctor_get(x_8, 0); -lean_dec(x_112); -lean_ctor_set(x_8, 0, x_1); -return x_8; -} else { -lean_object* x_113; lean_object* x_114; -x_113 = lean_ctor_get(x_8, 1); -lean_inc(x_113); -lean_dec(x_8); -x_114 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_114, 0, x_1); -lean_ctor_set(x_114, 1, x_113); -return x_114; -} -} -} -} -else -{ -uint8_t x_115; -x_115 = !lean_is_exclusive(x_8); -if (x_115 == 0) +uint8_t x_14; +x_14 = !lean_is_exclusive(x_8); +if (x_14 == 0) { -lean_object* x_116; lean_object* x_117; lean_object* x_118; uint8_t x_119; -x_116 = lean_ctor_get(x_8, 1); -x_117 = lean_ctor_get(x_8, 0); -lean_dec(x_117); -x_118 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2___closed__4; -x_119 = l_Lean_Expr_isAppOf(x_1, x_118); -if (x_119 == 0) +lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_15 = lean_ctor_get(x_8, 1); +x_16 = lean_ctor_get(x_8, 0); +lean_dec(x_16); +x_17 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__6___closed__4; +x_18 = l_Lean_Expr_isAppOf(x_1, x_17); +if (x_18 == 0) { -lean_object* x_120; lean_object* x_121; +lean_object* x_19; lean_object* x_20; lean_free_object(x_8); -x_120 = lean_box(0); -x_121 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3(x_1, x_120, x_2, x_3, x_4, x_5, x_6, x_116); +x_19 = lean_box(0); +x_20 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__7(x_1, x_19, x_2, x_3, x_4, x_5, x_6, x_15); lean_dec(x_2); -return x_121; +return x_20; } else { @@ -5814,85 +2459,85 @@ return x_8; } else { -lean_object* x_122; lean_object* x_123; uint8_t x_124; -x_122 = lean_ctor_get(x_8, 1); -lean_inc(x_122); +lean_object* x_21; lean_object* x_22; uint8_t x_23; +x_21 = lean_ctor_get(x_8, 1); +lean_inc(x_21); lean_dec(x_8); -x_123 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2___closed__4; -x_124 = l_Lean_Expr_isAppOf(x_1, x_123); -if (x_124 == 0) +x_22 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__6___closed__4; +x_23 = l_Lean_Expr_isAppOf(x_1, x_22); +if (x_23 == 0) { -lean_object* x_125; lean_object* x_126; -x_125 = lean_box(0); -x_126 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3(x_1, x_125, x_2, x_3, x_4, x_5, x_6, x_122); +lean_object* x_24; lean_object* x_25; +x_24 = lean_box(0); +x_25 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__7(x_1, x_24, x_2, x_3, x_4, x_5, x_6, x_21); lean_dec(x_2); -return x_126; +return x_25; } else { -lean_object* x_127; +lean_object* x_26; lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_127 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_127, 0, x_1); -lean_ctor_set(x_127, 1, x_122); -return x_127; +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_1); +lean_ctor_set(x_26, 1, x_21); +return x_26; } } } } else { -uint8_t x_128; +uint8_t x_27; 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_128 = !lean_is_exclusive(x_8); -if (x_128 == 0) +x_27 = !lean_is_exclusive(x_8); +if (x_27 == 0) { return x_8; } else { -lean_object* x_129; lean_object* x_130; lean_object* x_131; -x_129 = lean_ctor_get(x_8, 0); -x_130 = lean_ctor_get(x_8, 1); -lean_inc(x_130); -lean_inc(x_129); +lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_28 = lean_ctor_get(x_8, 0); +x_29 = lean_ctor_get(x_8, 1); +lean_inc(x_29); +lean_inc(x_28); lean_dec(x_8); -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; +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_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_13; -x_13 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_markNestedProofsImpl_visit___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); +uint8_t x_3; lean_object* x_4; +x_3 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__1(x_1, x_2); lean_dec(x_2); lean_dec(x_1); -return x_13; +x_4 = lean_box(x_3); +return x_4; } } -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__3___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_markNestedProofsImpl_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: { -uint8_t x_3; lean_object* x_4; -x_3 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__3(x_1, x_2); +lean_object* x_13; +x_13 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_markNestedProofsImpl_visit___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); lean_dec(x_2); lean_dec(x_1); -x_4 = lean_box(x_3); -return x_4; +return x_13; } } LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__10___boxed(lean_object* x_1, lean_object* x_2) { @@ -5910,25 +2555,66 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__ { lean_object* x_9; x_9 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +uint8_t x_14; lean_object* x_15; +x_14 = lean_unbox(x_4); +lean_dec(x_4); +x_15 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2(x_1, x_2, x_3, x_14, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_15; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___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_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); lean_dec(x_2); return x_9; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___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) { _start: { lean_object* x_9; -x_9 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_9 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_2); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___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) { +_start: +{ +lean_object* x_9; +x_9 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__5(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_2); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; +x_9 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__6(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); lean_dec(x_3); lean_dec(x_2); return x_9; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___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_Meta_Grind_markNestedProofsImpl_visit___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_9 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__7(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); lean_dec(x_3); lean_dec(x_2); return x_9; @@ -6049,30 +2735,30 @@ lean_dec_ref(res); res = initialize_Lean_Meta_InferType(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l_panic___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__1___closed__1 = _init_l_panic___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__1___closed__1(); -lean_mark_persistent(l_panic___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__1___closed__1); -l_panic___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__1___closed__2 = _init_l_panic___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__1___closed__2(); -lean_mark_persistent(l_panic___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__1___closed__2); -l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__1___closed__1 = _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__1___closed__1(); -lean_mark_persistent(l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__1___closed__1); -l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2___closed__1 = _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2___closed__1(); -lean_mark_persistent(l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2___closed__1); -l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2___closed__2 = _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2___closed__2(); -lean_mark_persistent(l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2___closed__2); -l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2___closed__3 = _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2___closed__3(); -lean_mark_persistent(l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2___closed__3); -l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2___closed__4 = _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2___closed__4(); -lean_mark_persistent(l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2___closed__4); -l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2___closed__5 = _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2___closed__5(); -lean_mark_persistent(l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2___closed__5); -l_Lean_Meta_Grind_markNestedProofsImpl_visit___closed__1 = _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___closed__1(); -lean_mark_persistent(l_Lean_Meta_Grind_markNestedProofsImpl_visit___closed__1); -l_Lean_Meta_Grind_markNestedProofsImpl_visit___closed__2 = _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___closed__2(); -lean_mark_persistent(l_Lean_Meta_Grind_markNestedProofsImpl_visit___closed__2); -l_Lean_Meta_Grind_markNestedProofsImpl_visit___closed__3 = _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___closed__3(); -lean_mark_persistent(l_Lean_Meta_Grind_markNestedProofsImpl_visit___closed__3); -l_Lean_Meta_Grind_markNestedProofsImpl_visit___closed__4 = _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___closed__4(); -lean_mark_persistent(l_Lean_Meta_Grind_markNestedProofsImpl_visit___closed__4); +l_panic___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__7___closed__1 = _init_l_panic___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__7___closed__1(); +lean_mark_persistent(l_panic___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__7___closed__1); +l_panic___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__7___closed__2 = _init_l_panic___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__7___closed__2(); +lean_mark_persistent(l_panic___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__7___closed__2); +l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3___closed__1 = _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3___closed__1); +l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3___closed__2 = _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3___closed__2); +l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3___closed__3 = _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3___closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3___closed__3); +l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3___closed__4 = _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3___closed__4(); +lean_mark_persistent(l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3___closed__4); +l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3___closed__5 = _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3___closed__5(); +lean_mark_persistent(l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3___closed__5); +l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__6___closed__1 = _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__6___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__6___closed__1); +l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__6___closed__2 = _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__6___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__6___closed__2); +l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__6___closed__3 = _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__6___closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__6___closed__3); +l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__6___closed__4 = _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__6___closed__4(); +lean_mark_persistent(l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__6___closed__4); +l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__6___closed__5 = _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__6___closed__5(); +lean_mark_persistent(l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__6___closed__5); l_Lean_Meta_Grind_markNestedProofsImpl___closed__1 = _init_l_Lean_Meta_Grind_markNestedProofsImpl___closed__1(); lean_mark_persistent(l_Lean_Meta_Grind_markNestedProofsImpl___closed__1); return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Lean/Meta/Tactic/Grind/Proj.c b/stage0/stdlib/Lean/Meta/Tactic/Grind/Proj.c index a7040611d752..5cfa35c8a3b9 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Grind/Proj.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Grind/Proj.c @@ -37,7 +37,6 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateProjEq(lean_object*, lean_ob lean_object* lean_st_ref_get(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_propagateProjEq___lambda__3___closed__3; static lean_object* l_Lean_Meta_Grind_propagateProjEq___lambda__3___closed__4; -lean_object* l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_propagateProjEq___lambda__3___closed__5; LEAN_EXPORT lean_object* l_Lean_getProjectionFnInfo_x3f___at_Lean_Meta_Grind_propagateProjEq___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_getProjectionFnInfo_x3f___at_Lean_Meta_Grind_propagateProjEq___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -45,19 +44,21 @@ lean_object* l_Lean_Meta_Grind_getGeneration(lean_object*, lean_object*, lean_ob uint8_t l_Lean_Meta_Grind_isSameExpr_unsafe__1(lean_object*, lean_object*); lean_object* l_Lean_MapDeclarationExtension_find_x3f___rarg(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_appFn_x21(lean_object*); +lean_object* l_Lean_Meta_Grind_updateLastTag(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_Grind_propagateProjEq___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*); extern lean_object* l_Lean_instInhabitedProjectionFunctionInfo; lean_object* l_Lean_MessageData_ofExpr(lean_object*); +lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___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_Grind_propagateProjEq___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_propagateProjEq___lambda__3___closed__6; lean_object* l_Lean_Expr_app___override(lean_object*, lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_shareCommon(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); +lean_object* l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_mkEqRefl(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*); -lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___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_Meta_Grind_propagateProjEq___lambda__3___closed__2; lean_object* lean_nat_add(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateProjEq___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*); @@ -270,7 +271,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateProjEq___lambda__3(lean_obje { lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; x_13 = l_Lean_Meta_Grind_propagateProjEq___lambda__3___closed__4; -x_14 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +x_14 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); x_15 = lean_ctor_get(x_14, 0); lean_inc(x_15); x_16 = lean_unbox(x_15); @@ -291,53 +292,126 @@ uint8_t x_20; x_20 = !lean_is_exclusive(x_14); if (x_20 == 0) { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +lean_object* x_21; lean_object* x_22; lean_object* x_23; x_21 = lean_ctor_get(x_14, 1); x_22 = lean_ctor_get(x_14, 0); lean_dec(x_22); +x_23 = l_Lean_Meta_Grind_updateLastTag(x_4, x_5, 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; 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_23, 1); +lean_inc(x_24); +lean_dec(x_23); lean_inc(x_3); -x_23 = l_Lean_MessageData_ofExpr(x_3); -x_24 = l_Lean_Meta_Grind_propagateProjEq___lambda__3___closed__6; +x_25 = l_Lean_MessageData_ofExpr(x_3); +x_26 = l_Lean_Meta_Grind_propagateProjEq___lambda__3___closed__6; lean_ctor_set_tag(x_14, 7); -lean_ctor_set(x_14, 1, x_23); -lean_ctor_set(x_14, 0, x_24); -x_25 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_25, 0, x_14); -lean_ctor_set(x_25, 1, x_24); -x_26 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_13, x_25, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_21); -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_Grind_propagateProjEq___lambda__2(x_1, x_2, x_3, x_27, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_28); -lean_dec(x_27); -return x_29; +lean_ctor_set(x_14, 1, x_25); +lean_ctor_set(x_14, 0, x_26); +x_27 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_27, 0, x_14); +lean_ctor_set(x_27, 1, x_26); +x_28 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_13, x_27, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, 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 = l_Lean_Meta_Grind_propagateProjEq___lambda__2(x_1, x_2, x_3, x_29, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_30); +lean_dec(x_29); +return x_31; } else { -lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; -x_30 = lean_ctor_get(x_14, 1); -lean_inc(x_30); +uint8_t x_32; +lean_free_object(x_14); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_3); +x_32 = !lean_is_exclusive(x_23); +if (x_32 == 0) +{ +return x_23; +} +else +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_33 = lean_ctor_get(x_23, 0); +x_34 = lean_ctor_get(x_23, 1); +lean_inc(x_34); +lean_inc(x_33); +lean_dec(x_23); +x_35 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_35, 0, x_33); +lean_ctor_set(x_35, 1, x_34); +return x_35; +} +} +} +else +{ +lean_object* x_36; lean_object* x_37; +x_36 = lean_ctor_get(x_14, 1); +lean_inc(x_36); lean_dec(x_14); +x_37 = l_Lean_Meta_Grind_updateLastTag(x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_36); +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, 1); +lean_inc(x_38); +lean_dec(x_37); lean_inc(x_3); -x_31 = l_Lean_MessageData_ofExpr(x_3); -x_32 = l_Lean_Meta_Grind_propagateProjEq___lambda__3___closed__6; -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 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_34, 0, x_33); -lean_ctor_set(x_34, 1, x_32); -x_35 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_13, x_34, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, 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_Grind_propagateProjEq___lambda__2(x_1, x_2, x_3, x_36, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_37); -lean_dec(x_36); -return x_38; +x_39 = l_Lean_MessageData_ofExpr(x_3); +x_40 = l_Lean_Meta_Grind_propagateProjEq___lambda__3___closed__6; +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 = 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_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_13, x_42, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_38); +x_44 = lean_ctor_get(x_43, 0); +lean_inc(x_44); +x_45 = lean_ctor_get(x_43, 1); +lean_inc(x_45); +lean_dec(x_43); +x_46 = l_Lean_Meta_Grind_propagateProjEq___lambda__2(x_1, x_2, x_3, x_44, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_45); +lean_dec(x_44); +return x_46; +} +else +{ +lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_3); +x_47 = lean_ctor_get(x_37, 0); +lean_inc(x_47); +x_48 = lean_ctor_get(x_37, 1); +lean_inc(x_48); +if (lean_is_exclusive(x_37)) { + lean_ctor_release(x_37, 0); + lean_ctor_release(x_37, 1); + x_49 = x_37; +} else { + lean_dec_ref(x_37); + x_49 = lean_box(0); +} +if (lean_is_scalar(x_49)) { + x_50 = lean_alloc_ctor(1, 2, 0); +} else { + x_50 = x_49; +} +lean_ctor_set(x_50, 0, x_47); +lean_ctor_set(x_50, 1, x_48); +return x_50; +} } } } diff --git a/stage0/stdlib/Lean/Meta/Tactic/Grind/Proof.c b/stage0/stdlib/Lean/Meta/Tactic/Grind/Proof.c index 3378f014e3c2..8a674a81ce00 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Grind/Proof.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Grind/Proof.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Lean.Meta.Tactic.Grind.Proof -// Imports: Lean.Meta.Sorry Lean.Meta.Tactic.Grind.Types +// Imports: Init.Grind.Lemmas Lean.Meta.Tactic.Grind.Types #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -15,6 +15,7 @@ extern "C" { #endif lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon(lean_object*, lean_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_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore___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___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkNestedProofCongr___closed__3; static lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__5___closed__4; @@ -24,51 +25,60 @@ lean_object* l_Lean_mkAppN(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_isCongrDefaultProofTarget_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_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_isEqProof(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__3___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__5; lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_isCongrDefaultProofTarget___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_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_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore(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_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___lambda__1___closed__1; static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore___lambda__1___closed__5; -static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__12; static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkNestedProofCongr___closed__2; +lean_object* l_Lean_mkApp7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_mkEqSymm(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkRefl___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___boxed(lean_object**); +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___spec__1___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore___lambda__1___closed__6; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_flipProof___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__4; uint8_t l_Lean_Expr_isApp(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___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_object*, lean_object*); uint8_t l_Lean_RBNode_isRed___rarg(lean_object*); lean_object* l_Lean_Expr_sort___override(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___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*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__4; LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__7___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkProofTo___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___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_mk_array(lean_object*, lean_object*); static lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__5___closed__1; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__2___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___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_array_fget(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___spec__1___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__6; static lean_object* l_Lean_Meta_Grind_mkEqProofImpl___closed__4; lean_object* l_Lean_Meta_mkHEqOfEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_mkEqProofImpl___closed__2; static lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__5___closed__2; +uint8_t l_Lean_Meta_Grind_hasSameRoot(lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isAppOf(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkProofFrom___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__7___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Expr_cleanupAnnotations(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore___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_Grind_mkEqProofImpl___closed__3; -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof_loop(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___closed__1; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__1; lean_object* l_Lean_Expr_appArg_x21(lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore___lambda__1___closed__3; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___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_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___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*); extern lean_object* l_instInhabitedPUnit; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_isCongrDefaultProofTarget_loop___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkProofFrom___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* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* lean_grind_mk_heq_proof(lean_object*, 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_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__3___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkNestedProofCongr___boxed(lean_object*, lean_object*, lean_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_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkProofFrom___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_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___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_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkTrans_x27(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqOfHEqIfNeeded(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -77,13 +87,18 @@ LEAN_EXPORT lean_object* lean_grind_mk_eq_proof(lean_object*, lean_object*, lean LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__5; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkProofTo___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___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkTrans_x27___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_expr_eqv(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__4___closed__1; static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_isEqProof___closed__1; lean_object* l_Lean_RBNode_setBlack___rarg(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_getFunInfo(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___boxed(lean_object*, lean_object*, lean_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_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkProofFrom___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___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_Meta_mkEqNDRec(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_mkHEqRefl(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*); @@ -91,14 +106,19 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_G lean_object* l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_flipProof___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkProofTo___lambda__1___closed__1; +lean_object* l_Lean_Expr_appArg(lean_object*, lean_object*); lean_object* l_outOfBounds___rarg(lean_object*); +lean_object* lean_st_ref_get(lean_object*, lean_object*); 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_RBNode_find___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__6___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__8; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_mkHEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Expr_appFnCleanup(lean_object*, lean_object*); lean_object* l_Lean_Meta_mkCongrArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__7___lambda__1___closed__2; -static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__11; extern lean_object* l_Lean_levelZero; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_realizeEqProof(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*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__7; extern lean_object* l_Lean_instInhabitedExpr; @@ -107,16 +127,18 @@ static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_ LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_flipProof(lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_panic___at_Lean_Expr_appFn_x21___spec__1(lean_object*); lean_object* l_Lean_Meta_Grind_mkHCongrWithArity(lean_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_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof_loop(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___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_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkProofFrom___lambda__1___closed__3; uint8_t l_Lean_Meta_Grind_isSameExpr_unsafe__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkNestedProofCongr___closed__1; LEAN_EXPORT lean_object* l_Lean_RBNode_find___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__6(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__13; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_isCongrDefaultProofTarget___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___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof(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_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_mkEqProofImpl___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof_loop___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_appFn_x21(lean_object*); extern lean_object* l_Lean_Meta_instMonadMetaM; LEAN_EXPORT lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -124,20 +146,23 @@ static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_ lean_object* l_Lean_Meta_Grind_hasSameType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof_loop___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__8; -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___boxed(lean_object*, lean_object*, lean_object*, 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_ins___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__2(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkTrans___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___closed__1; LEAN_EXPORT lean_object* l_Lean_RBNode_insert___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__1(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__5; lean_object* l_Lean_Meta_mkCongr(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_isCongrDefaultProofTarget___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_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__7___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___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___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkProofFrom___lambda__1___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkProofFrom(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_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Expr_constLevels_x21(lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__6; LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern uint8_t l_Lean_Meta_instInhabitedCongrArgKind; -static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__10; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_realizeEqProof___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); lean_object* l_Lean_mkApp3(lean_object*, lean_object*, lean_object*, lean_object*); @@ -145,21 +170,26 @@ uint8_t lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_isCongrDefaultProofTarget_loop___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_isEqProof___closed__2; static lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkProofFrom___spec__1___closed__1; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__2; lean_object* l_Lean_Meta_mkEqRefl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkNestedProofCongr(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___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_isCongrDefaultProofTarget(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___spec__1(lean_object*); uint8_t l_Lean_Expr_isConstOf(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkProofFrom___boxed(lean_object*, lean_object*, lean_object*, 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_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___lambda__1___closed__1; static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkProofTo___lambda__1___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof(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_panic_fn(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___boxed(lean_object*, lean_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___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore___lambda__1___closed__1; lean_object* l_Lean_Expr_getAppFn(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqOfHEqIfNeeded___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof___closed__1; lean_object* l_Lean_Meta_Grind_getENode(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_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__3; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__7; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__3; static lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__7___lambda__1___closed__4; LEAN_EXPORT lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___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_array_mk(lean_object*); @@ -171,12 +201,16 @@ lean_object* l_Lean_Meta_mkEqOfHEq(lean_object*, lean_object*, lean_object*, lea LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkTrans(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore___lambda__1___closed__4; +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___spec__1(lean_object*); lean_object* l_instInhabitedOfMonad___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof(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_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__7___lambda__1___closed__1; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__2; static lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__9___closed__1; lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkProofTo___boxed(lean_object*, lean_object*, lean_object*, 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_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof___closed__3; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__4; lean_object* lean_string_append(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkRefl(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_mkHEqTrans(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -187,15 +221,14 @@ lean_object* lean_infer_type(lean_object*, lean_object*, lean_object*, lean_obje LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_isCongrDefaultProofTarget_loop(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore___boxed(lean_object*, lean_object*, lean_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_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof_loop___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_mkLambdaFVars(lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Meta_Grind_congrPlaceholderProof; static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof___closed__4; static lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__7___lambda__1___closed__3; lean_object* l_Lean_mkApp5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___lambda__1___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___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_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___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_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__9; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__7; static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore___lambda__1___closed__2; lean_object* l_Lean_Meta_mkHEqSymm(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_ReaderT_instMonad___rarg(lean_object*); @@ -5453,7 +5486,7 @@ lean_dec(x_1); return x_15; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof_loop(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_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_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof_loop(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, 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; uint8_t x_16; @@ -5492,7 +5525,7 @@ lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -x_21 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof_loop(x_1, x_19, x_20, x_15, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +x_21 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof_loop(x_1, x_19, x_20, x_15, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); lean_dec(x_20); lean_dec(x_19); if (lean_obj_tag(x_21) == 0) @@ -5693,7 +5726,7 @@ static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__7___lambda__1___closed__1; x_2 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore___lambda__1___closed__4; -x_3 = lean_unsigned_to_nat(212u); +x_3 = lean_unsigned_to_nat(234u); x_4 = lean_unsigned_to_nat(4u); x_5 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore___lambda__1___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -5706,7 +5739,7 @@ static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__7___lambda__1___closed__1; x_2 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore___lambda__1___closed__4; -x_3 = lean_unsigned_to_nat(215u); +x_3 = lean_unsigned_to_nat(237u); x_4 = lean_unsigned_to_nat(72u); x_5 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__7___lambda__1___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -6224,7 +6257,7 @@ static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__7___lambda__1___closed__1; x_2 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkProofFrom___lambda__1___closed__1; -x_3 = lean_unsigned_to_nat(193u); +x_3 = lean_unsigned_to_nat(215u); x_4 = lean_unsigned_to_nat(35u); x_5 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__7___lambda__1___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -6237,7 +6270,7 @@ static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__7___lambda__1___closed__1; x_2 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkProofFrom___lambda__1___closed__1; -x_3 = lean_unsigned_to_nat(194u); +x_3 = lean_unsigned_to_nat(216u); x_4 = lean_unsigned_to_nat(29u); x_5 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__7___lambda__1___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -7023,167 +7056,11 @@ return x_18; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___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, lean_object* x_11) { -_start: -{ -lean_object* x_12; -x_12 = lean_apply_10(x_1, x_6, x_2, x_3, x_4, x_5, x_7, x_8, x_9, x_10, x_11); -return x_12; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___spec__1___rarg(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { -_start: -{ -lean_object* x_15; lean_object* x_16; -x_15 = lean_alloc_closure((void*)(l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___spec__1___rarg___lambda__1), 11, 5); -lean_closure_set(x_15, 0, x_4); -lean_closure_set(x_15, 1, x_6); -lean_closure_set(x_15, 2, x_7); -lean_closure_set(x_15, 3, x_8); -lean_closure_set(x_15, 4, x_9); -x_16 = l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDeclImp___rarg(x_1, x_2, x_3, x_15, x_5, x_10, x_11, x_12, x_13, x_14); -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; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___spec__1(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___spec__1___rarg___boxed), 14, 0); -return x_2; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___lambda__1___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_levelZero; -x_2 = l_Lean_Expr_sort___override(x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_13 = lean_unsigned_to_nat(0u); -x_14 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_13); -x_15 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___lambda__1___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_1, x_16, x_18); -lean_inc(x_3); -x_20 = l_Lean_mkAppN(x_3, x_19); -lean_dec(x_19); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -x_21 = l_Lean_Meta_mkHEq(x_2, x_20, x_8, x_9, x_10, x_11, x_12); -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; uint8_t x_27; uint8_t x_28; uint8_t x_29; lean_object* x_30; -x_22 = lean_ctor_get(x_21, 0); -lean_inc(x_22); -x_23 = lean_ctor_get(x_21, 1); -lean_inc(x_23); -lean_dec(x_21); -x_24 = lean_box(0); -x_25 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_25, 0, x_3); -lean_ctor_set(x_25, 1, x_24); -x_26 = lean_array_mk(x_25); -x_27 = 0; -x_28 = 1; -x_29 = 1; -x_30 = l_Lean_Meta_mkLambdaFVars(x_26, x_22, x_27, x_28, x_27, x_29, x_8, x_9, 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_26); -return x_30; -} -else -{ -uint8_t x_31; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -x_31 = !lean_is_exclusive(x_21); -if (x_31 == 0) -{ -return x_21; -} -else -{ -lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_32 = lean_ctor_get(x_21, 0); -x_33 = lean_ctor_get(x_21, 1); -lean_inc(x_33); -lean_inc(x_32); -lean_dec(x_21); -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___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("thm.argKinds.size == numArgs\n ", 35, 35); +x_1 = lean_mk_string_unchecked("rhs.getAppNumArgs == numArgs\n ", 33, 33); return x_1; } } @@ -7211,8 +7088,8 @@ static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__7___lambda__1___closed__1; x_2 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__3; -x_3 = lean_unsigned_to_nat(143u); -x_4 = lean_unsigned_to_nat(6u); +x_3 = lean_unsigned_to_nat(182u); +x_4 = lean_unsigned_to_nat(4u); x_5 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__2; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); return x_6; @@ -7222,60 +7099,11 @@ static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_ _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("x", 1, 1); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___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_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__5; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__7() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("rhs.getAppNumArgs == numArgs\n ", 33, 33); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__8() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore___lambda__1___closed__1; -x_2 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__7; -x_3 = lean_string_append(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__9() { -_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_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__7___lambda__1___closed__1; -x_2 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__3; -x_3 = lean_unsigned_to_nat(136u); -x_4 = lean_unsigned_to_nat(4u); -x_5 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__8; -x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); -return x_6; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__10() { -_start: -{ -lean_object* x_1; x_1 = lean_mk_string_unchecked("Lean", 4, 4); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__11() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__6() { _start: { lean_object* x_1; @@ -7283,7 +7111,7 @@ x_1 = lean_mk_string_unchecked("Grind", 5, 5); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__12() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__7() { _start: { lean_object* x_1; @@ -7291,13 +7119,13 @@ x_1 = lean_mk_string_unchecked("nestedProof", 11, 11); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__13() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__10; -x_2 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__11; -x_3 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__12; +x_1 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__5; +x_2 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__6; +x_3 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__7; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } @@ -7305,293 +7133,197 @@ return x_4; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof(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_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_87; uint8_t x_88; +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_58; uint8_t x_59; x_13 = l_Lean_Expr_getAppFn(x_1); x_14 = l_Lean_Expr_getAppFn(x_2); x_15 = lean_unsigned_to_nat(0u); x_16 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_15); -x_87 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_2, x_15); -x_88 = lean_nat_dec_eq(x_87, x_16); -lean_dec(x_87); -if (x_88 == 0) +x_58 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_2, x_15); +x_59 = lean_nat_dec_eq(x_58, x_16); +lean_dec(x_58); +if (x_59 == 0) { -lean_object* x_89; lean_object* x_90; +lean_object* x_60; lean_object* x_61; lean_dec(x_16); lean_dec(x_14); lean_dec(x_13); lean_dec(x_2); lean_dec(x_1); -x_89 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__9; -x_90 = l_panic___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__9(x_89, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -return x_90; +x_60 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__4; +x_61 = l_panic___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__9(x_60, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_61; } else { -lean_object* x_91; uint8_t x_92; -x_91 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__13; -x_92 = l_Lean_Expr_isConstOf(x_13, x_91); -if (x_92 == 0) +lean_object* x_62; uint8_t x_63; +x_62 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__8; +x_63 = l_Lean_Expr_isConstOf(x_13, x_62); +if (x_63 == 0) { -lean_object* x_93; -x_93 = lean_box(0); -x_17 = x_93; -goto block_86; +lean_object* x_64; +x_64 = lean_box(0); +x_17 = x_64; +goto block_57; } else { -uint8_t x_94; -x_94 = l_Lean_Expr_isConstOf(x_14, x_91); -if (x_94 == 0) +uint8_t x_65; +x_65 = l_Lean_Expr_isConstOf(x_14, x_62); +if (x_65 == 0) { -lean_object* x_95; -x_95 = lean_box(0); -x_17 = x_95; -goto block_86; +lean_object* x_66; +x_66 = lean_box(0); +x_17 = x_66; +goto block_57; } else { -lean_object* x_96; uint8_t x_97; -x_96 = lean_unsigned_to_nat(2u); -x_97 = lean_nat_dec_eq(x_16, x_96); -if (x_97 == 0) +lean_object* x_67; uint8_t x_68; +x_67 = lean_unsigned_to_nat(2u); +x_68 = lean_nat_dec_eq(x_16, x_67); +if (x_68 == 0) { -lean_object* x_98; -x_98 = lean_box(0); -x_17 = x_98; -goto block_86; +lean_object* x_69; +x_69 = lean_box(0); +x_17 = x_69; +goto block_57; } else { -lean_object* x_99; +lean_object* x_70; lean_dec(x_16); lean_dec(x_14); lean_dec(x_13); -x_99 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkNestedProofCongr(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_70 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkNestedProofCongr(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_2); lean_dec(x_1); -return x_99; +return x_70; } } } } -block_86: +block_57: { -lean_object* x_18; +lean_object* x_18; uint8_t x_19; lean_dec(x_17); +x_18 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_isEqProof___closed__2; +x_19 = l_Lean_Expr_isConstOf(x_13, x_18); +if (x_19 == 0) +{ +lean_object* x_20; lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); -lean_inc(x_13); -x_18 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_isCongrDefaultProofTarget(x_1, x_2, x_13, x_14, x_16, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_18) == 0) -{ -lean_object* x_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) +x_20 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_isCongrDefaultProofTarget(x_1, x_2, x_13, x_14, x_16, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_16); +lean_dec(x_14); +if (lean_obj_tag(x_20) == 0) { -lean_object* x_21; lean_object* x_22; -x_21 = lean_ctor_get(x_18, 1); +lean_object* x_21; uint8_t x_22; +x_21 = lean_ctor_get(x_20, 0); lean_inc(x_21); -lean_dec(x_18); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_16); -lean_inc(x_13); -x_22 = l_Lean_Meta_Grind_mkHCongrWithArity(x_13, x_16, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_21); -if (lean_obj_tag(x_22) == 0) +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; -x_23 = lean_ctor_get(x_22, 0); +lean_object* x_23; lean_object* x_24; +x_23 = lean_ctor_get(x_20, 1); lean_inc(x_23); -x_24 = lean_ctor_get(x_22, 1); -lean_inc(x_24); -lean_dec(x_22); -x_25 = lean_ctor_get(x_23, 2); +lean_dec(x_20); +x_24 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_23); +return x_24; +} +else +{ +lean_object* x_25; lean_object* x_26; +x_25 = lean_ctor_get(x_20, 1); lean_inc(x_25); -x_26 = lean_array_get_size(x_25); -lean_dec(x_25); -x_27 = lean_nat_dec_eq(x_26, x_16); -lean_dec(x_26); -if (x_27 == 0) +lean_dec(x_20); +x_26 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_25); +lean_dec(x_2); +lean_dec(x_1); +return x_26; +} +} +else { -lean_object* x_28; lean_object* x_29; -lean_dec(x_23); -lean_dec(x_16); -lean_dec(x_14); -lean_dec(x_13); +uint8_t x_27; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_28 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__4; -x_29 = l_panic___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__9(x_28, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_24); -return x_29; +x_27 = !lean_is_exclusive(x_20); +if (x_27 == 0) +{ +return x_20; } else { -lean_object* x_30; -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -x_30 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof_loop(x_23, x_1, x_2, x_16, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_24); -lean_dec(x_16); -lean_dec(x_23); -if (lean_obj_tag(x_30) == 0) -{ -lean_object* x_31; lean_object* x_32; uint8_t 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_Grind_isSameExpr_unsafe__1(x_13, 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; -x_34 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__6; -x_35 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_34, x_10, x_11, x_32); -x_36 = lean_ctor_get(x_35, 0); -lean_inc(x_36); -x_37 = lean_ctor_get(x_35, 1); -lean_inc(x_37); -lean_dec(x_35); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_13); -x_38 = lean_infer_type(x_13, x_8, x_9, x_10, x_11, x_37); -if (lean_obj_tag(x_38) == 0) -{ -lean_object* x_39; lean_object* x_40; lean_object* x_41; uint8_t x_42; uint8_t x_43; lean_object* x_44; -x_39 = lean_ctor_get(x_38, 0); -lean_inc(x_39); -x_40 = lean_ctor_get(x_38, 1); -lean_inc(x_40); -lean_dec(x_38); -x_41 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___lambda__1___boxed), 12, 2); -lean_closure_set(x_41, 0, x_2); -lean_closure_set(x_41, 1, x_1); -x_42 = 0; -x_43 = 0; -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -x_44 = l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___spec__1___rarg(x_36, x_42, x_39, x_41, x_43, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_40); -if (lean_obj_tag(x_44) == 0) +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_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; +} +} +} +else { -lean_object* x_45; lean_object* x_46; uint8_t x_47; lean_object* x_48; -x_45 = lean_ctor_get(x_44, 0); -lean_inc(x_45); -x_46 = lean_ctor_get(x_44, 1); -lean_inc(x_46); -lean_dec(x_44); -x_47 = 0; -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -x_48 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore(x_13, x_14, x_47, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_46); -if (lean_obj_tag(x_48) == 0) +uint8_t x_31; +x_31 = l_Lean_Expr_isConstOf(x_14, x_18); +if (x_31 == 0) { -lean_object* x_49; lean_object* x_50; lean_object* x_51; -x_49 = lean_ctor_get(x_48, 0); -lean_inc(x_49); -x_50 = lean_ctor_get(x_48, 1); -lean_inc(x_50); -lean_dec(x_48); +lean_object* x_32; lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); -x_51 = l_Lean_Meta_mkEqNDRec(x_45, x_31, x_49, x_8, x_9, x_10, x_11, x_50); -if (lean_obj_tag(x_51) == 0) -{ -lean_object* x_52; lean_object* x_53; lean_object* x_54; -x_52 = lean_ctor_get(x_51, 0); -lean_inc(x_52); -x_53 = lean_ctor_get(x_51, 1); -lean_inc(x_53); -lean_dec(x_51); -x_54 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqOfHEqIfNeeded(x_52, x_3, x_8, x_9, x_10, x_11, x_53); -return x_54; -} -else -{ -uint8_t x_55; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -x_55 = !lean_is_exclusive(x_51); -if (x_55 == 0) -{ -return x_51; -} -else -{ -lean_object* x_56; lean_object* x_57; lean_object* x_58; -x_56 = lean_ctor_get(x_51, 0); -x_57 = lean_ctor_get(x_51, 1); -lean_inc(x_57); -lean_inc(x_56); -lean_dec(x_51); -x_58 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_58, 0, x_56); -lean_ctor_set(x_58, 1, x_57); -return x_58; -} -} -} -else +x_32 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_isCongrDefaultProofTarget(x_1, x_2, x_13, x_14, x_16, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_16); +lean_dec(x_14); +if (lean_obj_tag(x_32) == 0) { -uint8_t x_59; -lean_dec(x_45); -lean_dec(x_31); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -x_59 = !lean_is_exclusive(x_48); -if (x_59 == 0) +lean_object* x_33; uint8_t x_34; +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) { -return x_48; +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 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_35); +return x_36; } else { -lean_object* x_60; lean_object* x_61; lean_object* x_62; -x_60 = lean_ctor_get(x_48, 0); -x_61 = lean_ctor_get(x_48, 1); -lean_inc(x_61); -lean_inc(x_60); -lean_dec(x_48); -x_62 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_62, 0, x_60); -lean_ctor_set(x_62, 1, x_61); -return x_62; -} +lean_object* x_37; lean_object* x_38; +x_37 = lean_ctor_get(x_32, 1); +lean_inc(x_37); +lean_dec(x_32); +x_38 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_37); +lean_dec(x_2); +lean_dec(x_1); +return x_38; } } else { -uint8_t x_63; -lean_dec(x_31); -lean_dec(x_14); -lean_dec(x_13); +uint8_t x_39; lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -7600,83 +7332,74 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_63 = !lean_is_exclusive(x_44); -if (x_63 == 0) +lean_dec(x_2); +lean_dec(x_1); +x_39 = !lean_is_exclusive(x_32); +if (x_39 == 0) { -return x_44; +return x_32; } else { -lean_object* x_64; lean_object* x_65; lean_object* x_66; -x_64 = lean_ctor_get(x_44, 0); -x_65 = lean_ctor_get(x_44, 1); -lean_inc(x_65); -lean_inc(x_64); -lean_dec(x_44); -x_66 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_66, 0, x_64); -lean_ctor_set(x_66, 1, x_65); -return x_66; +lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_40 = lean_ctor_get(x_32, 0); +x_41 = lean_ctor_get(x_32, 1); +lean_inc(x_41); +lean_inc(x_40); +lean_dec(x_32); +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_67; -lean_dec(x_36); -lean_dec(x_31); +lean_object* x_43; uint8_t x_44; +x_43 = lean_unsigned_to_nat(3u); +x_44 = lean_nat_dec_eq(x_16, x_43); +if (x_44 == 0) +{ +lean_object* x_45; +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +x_45 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_isCongrDefaultProofTarget(x_1, x_2, x_13, x_14, x_16, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_16); 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_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_2); -lean_dec(x_1); -x_67 = !lean_is_exclusive(x_38); -if (x_67 == 0) +if (lean_obj_tag(x_45) == 0) { -return x_38; -} -else +lean_object* x_46; uint8_t x_47; +x_46 = lean_ctor_get(x_45, 0); +lean_inc(x_46); +x_47 = lean_unbox(x_46); +lean_dec(x_46); +if (x_47 == 0) { -lean_object* x_68; lean_object* x_69; lean_object* x_70; -x_68 = lean_ctor_get(x_38, 0); -x_69 = lean_ctor_get(x_38, 1); -lean_inc(x_69); -lean_inc(x_68); -lean_dec(x_38); -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_48; lean_object* x_49; +x_48 = lean_ctor_get(x_45, 1); +lean_inc(x_48); +lean_dec(x_45); +x_49 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_48); +return x_49; } else { -lean_object* x_71; -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); +lean_object* x_50; lean_object* x_51; +x_50 = lean_ctor_get(x_45, 1); +lean_inc(x_50); +lean_dec(x_45); +x_51 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_50); lean_dec(x_2); lean_dec(x_1); -x_71 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqOfHEqIfNeeded(x_31, x_3, x_8, x_9, x_10, x_11, x_32); -return x_71; +return x_51; } } else { -uint8_t x_72; -lean_dec(x_14); -lean_dec(x_13); +uint8_t x_52; lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -7687,331 +7410,329 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_72 = !lean_is_exclusive(x_30); -if (x_72 == 0) +x_52 = !lean_is_exclusive(x_45); +if (x_52 == 0) { -return x_30; +return x_45; } else { -lean_object* x_73; lean_object* x_74; lean_object* x_75; -x_73 = lean_ctor_get(x_30, 0); -x_74 = lean_ctor_get(x_30, 1); -lean_inc(x_74); -lean_inc(x_73); -lean_dec(x_30); -x_75 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_75, 0, x_73); -lean_ctor_set(x_75, 1, x_74); -return x_75; -} +lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_53 = lean_ctor_get(x_45, 0); +x_54 = lean_ctor_get(x_45, 1); +lean_inc(x_54); +lean_inc(x_53); +lean_dec(x_45); +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_76; +lean_object* x_56; lean_dec(x_16); 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_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_2); -lean_dec(x_1); -x_76 = !lean_is_exclusive(x_22); -if (x_76 == 0) -{ -return x_22; +x_56 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_56; } -else -{ -lean_object* x_77; lean_object* x_78; lean_object* x_79; -x_77 = lean_ctor_get(x_22, 0); -x_78 = lean_ctor_get(x_22, 1); -lean_inc(x_78); -lean_inc(x_77); -lean_dec(x_22); -x_79 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_79, 0, x_77); -lean_ctor_set(x_79, 1, x_78); -return x_79; } } } -else +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___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, lean_object* x_11) { +_start: { -lean_object* x_80; lean_object* x_81; -lean_dec(x_16); -lean_dec(x_14); -lean_dec(x_13); -x_80 = lean_ctor_get(x_18, 1); -lean_inc(x_80); -lean_dec(x_18); -x_81 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_80); -lean_dec(x_2); -lean_dec(x_1); -return x_81; +lean_object* x_12; +x_12 = lean_apply_10(x_1, x_6, x_2, x_3, x_4, x_5, x_7, x_8, x_9, x_10, x_11); +return x_12; } } -else +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___spec__1___rarg(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: { -uint8_t x_82; -lean_dec(x_16); -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_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_2); -lean_dec(x_1); -x_82 = !lean_is_exclusive(x_18); -if (x_82 == 0) +lean_object* x_15; lean_object* x_16; +x_15 = lean_alloc_closure((void*)(l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___spec__1___rarg___lambda__1), 11, 5); +lean_closure_set(x_15, 0, x_4); +lean_closure_set(x_15, 1, x_6); +lean_closure_set(x_15, 2, x_7); +lean_closure_set(x_15, 3, x_8); +lean_closure_set(x_15, 4, x_9); +x_16 = l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDeclImp___rarg(x_1, x_2, x_3, x_15, x_5, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_16) == 0) { -return x_18; +uint8_t x_17; +x_17 = !lean_is_exclusive(x_16); +if (x_17 == 0) +{ +return x_16; } else { -lean_object* x_83; lean_object* x_84; lean_object* x_85; -x_83 = lean_ctor_get(x_18, 0); -x_84 = lean_ctor_get(x_18, 1); -lean_inc(x_84); -lean_inc(x_83); -lean_dec(x_18); -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_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; } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof___closed__1() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Init.Data.Option.BasicAux", 25, 25); -return x_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_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; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof___closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Option.get!", 11, 11); -return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof___closed__3() { +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___spec__1(lean_object* x_1) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("value is none", 13, 13); -return x_1; +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___spec__1___rarg___boxed), 14, 0); +return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof___closed__4() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___lambda__1___closed__1() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_1 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof___closed__1; -x_2 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof___closed__2; -x_3 = lean_unsigned_to_nat(16u); -x_4 = lean_unsigned_to_nat(14u); -x_5 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof___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; lean_object* x_2; +x_1 = l_Lean_levelZero; +x_2 = l_Lean_Expr_sort___override(x_1); +return x_2; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof(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___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___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_13; +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_14 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___lambda__1___closed__1; +lean_inc(x_1); +x_15 = lean_mk_array(x_1, x_14); +x_16 = lean_unsigned_to_nat(1u); +x_17 = lean_nat_sub(x_1, x_16); +lean_dec(x_1); +x_18 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_2, x_15, x_17); +lean_inc(x_4); +x_19 = l_Lean_mkAppN(x_4, x_18); +lean_dec(x_18); +lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); -lean_inc(x_8); -x_13 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof_loop(x_1, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_13) == 0) -{ -lean_object* x_14; -x_14 = lean_ctor_get(x_13, 0); -lean_inc(x_14); -if (lean_obj_tag(x_14) == 0) -{ -uint8_t x_15; -x_15 = !lean_is_exclusive(x_13); -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_13, 1); -x_17 = lean_ctor_get(x_13, 0); -lean_dec(x_17); -x_18 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof___closed__4; -x_19 = l_panic___at_Lean_Expr_appFn_x21___spec__1(x_18); -if (x_3 == 0) +x_20 = l_Lean_Meta_mkHEq(x_3, x_19, x_9, x_10, x_11, x_12, x_13); +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; uint8_t x_26; uint8_t x_27; uint8_t x_28; lean_object* x_29; +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_box(0); +x_24 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_24, 0, x_4); +lean_ctor_set(x_24, 1, x_23); +x_25 = lean_array_mk(x_24); +x_26 = 0; +x_27 = 1; +x_28 = 1; +x_29 = l_Lean_Meta_mkLambdaFVars(x_25, x_21, x_26, x_27, x_26, x_28, x_9, x_10, x_11, x_12, x_22); +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); -lean_dec(x_8); -lean_ctor_set(x_13, 0, x_19); -return x_13; -} -else -{ -lean_object* x_20; -lean_free_object(x_13); -x_20 = l_Lean_Meta_mkHEqOfEq(x_19, x_8, x_9, x_10, x_11, x_16); -return x_20; -} +lean_dec(x_25); +return x_29; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_13, 1); -lean_inc(x_21); -lean_dec(x_13); -x_22 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof___closed__4; -x_23 = l_panic___at_Lean_Expr_appFn_x21___spec__1(x_22); -if (x_3 == 0) -{ -lean_object* x_24; +uint8_t x_30; +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); -lean_dec(x_8); -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_dec(x_4); +x_30 = !lean_is_exclusive(x_20); +if (x_30 == 0) +{ +return x_20; } else { -lean_object* x_25; -x_25 = l_Lean_Meta_mkHEqOfEq(x_23, x_8, x_9, x_10, x_11, x_21); -return x_25; +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_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 -{ -if (x_3 == 0) -{ -uint8_t x_26; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -x_26 = !lean_is_exclusive(x_13); -if (x_26 == 0) -{ -lean_object* x_27; lean_object* x_28; -x_27 = lean_ctor_get(x_13, 0); -lean_dec(x_27); -x_28 = lean_ctor_get(x_14, 0); -lean_inc(x_28); -lean_dec(x_14); -lean_ctor_set(x_13, 0, x_28); -return x_13; } -else +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__1() { +_start: { -lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_29 = lean_ctor_get(x_13, 1); -lean_inc(x_29); -lean_dec(x_13); -x_30 = lean_ctor_get(x_14, 0); -lean_inc(x_30); -lean_dec(x_14); -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_1; +x_1 = lean_mk_string_unchecked("_private.Lean.Meta.Tactic.Grind.Proof.0.Lean.Meta.Grind.mkHCongrProof", 69, 69); +return x_1; } } -else +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__2() { +_start: { -lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_32 = lean_ctor_get(x_13, 1); -lean_inc(x_32); -lean_dec(x_13); -x_33 = lean_ctor_get(x_14, 0); -lean_inc(x_33); -lean_dec(x_14); -x_34 = l_Lean_Meta_mkHEqOfEq(x_33, x_8, x_9, x_10, x_11, x_32); -return x_34; -} -} +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__7___lambda__1___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__1; +x_3 = lean_unsigned_to_nat(135u); +x_4 = lean_unsigned_to_nat(4u); +x_5 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__2; +x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); +return x_6; } -else -{ -uint8_t x_35; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -x_35 = !lean_is_exclusive(x_13); -if (x_35 == 0) -{ -return x_13; } -else +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__3() { +_start: { -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_dec(x_13); -x_38 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_38, 0, x_36); -lean_ctor_set(x_38, 1, x_37); -return x_38; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("thm.argKinds.size == numArgs\n ", 33, 33); +return x_1; } } +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore___lambda__1___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__3; +x_3 = lean_string_append(x_1, x_2); +return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof_loop(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, 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___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__5() { _start: { -uint8_t x_12; -x_12 = l_Lean_Expr_isApp(x_1); -if (x_12 == 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_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__7___lambda__1___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__1; +x_3 = lean_unsigned_to_nat(137u); +x_4 = lean_unsigned_to_nat(4u); +x_5 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__4; +x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); +return x_6; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__6() { +_start: { -lean_object* x_13; lean_object* x_14; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -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_11); -return x_14; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("x", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__6; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof(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_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_13 = l_Lean_Expr_getAppFn(x_1); +x_14 = l_Lean_Expr_getAppFn(x_2); +x_15 = lean_unsigned_to_nat(0u); +x_16 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_15); +x_17 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_2, x_15); +x_18 = lean_nat_dec_eq(x_17, x_16); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_2); +lean_dec(x_1); +x_19 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__2; +x_20 = l_panic___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__9(x_19, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_20; } else { -lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_15 = l_Lean_Expr_appArg_x21(x_1); -x_16 = l_Lean_Expr_appArg_x21(x_2); -x_17 = l_Lean_Expr_appFn_x21(x_1); -x_18 = l_Lean_Expr_appFn_x21(x_2); +lean_object* x_21; +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_16); +lean_inc(x_13); +x_21 = l_Lean_Meta_Grind_mkHCongrWithArity(x_13, x_16, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_21) == 0) +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; +x_22 = lean_ctor_get(x_21, 0); +lean_inc(x_22); +x_23 = lean_ctor_get(x_21, 1); +lean_inc(x_23); +lean_dec(x_21); +x_24 = lean_ctor_get(x_22, 2); +lean_inc(x_24); +x_25 = lean_array_get_size(x_24); +lean_dec(x_24); +x_26 = lean_nat_dec_eq(x_25, x_16); +lean_dec(x_25); +if (x_26 == 0) +{ +lean_object* x_27; lean_object* x_28; +lean_dec(x_22); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_2); +lean_dec(x_1); +x_27 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__5; +x_28 = l_panic___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__9(x_27, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_23); +return x_28; +} +else +{ +lean_object* x_29; +lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); @@ -8019,130 +7740,158 @@ lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -lean_inc(x_3); -x_19 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof_loop(x_17, x_18, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -lean_dec(x_18); -if (lean_obj_tag(x_19) == 0) -{ -lean_object* x_20; -x_20 = lean_ctor_get(x_19, 0); -lean_inc(x_20); -if (lean_obj_tag(x_20) == 0) +x_29 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof_loop(x_22, x_1, x_2, x_16, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_23); +lean_dec(x_16); +lean_dec(x_22); +if (lean_obj_tag(x_29) == 0) { -uint8_t x_21; -x_21 = !lean_is_exclusive(x_19); -if (x_21 == 0) +lean_object* x_30; lean_object* x_31; uint8_t x_32; +x_30 = lean_ctor_get(x_29, 0); +lean_inc(x_30); +x_31 = lean_ctor_get(x_29, 1); +lean_inc(x_31); +lean_dec(x_29); +x_32 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_13, x_14); +if (x_32 == 0) { -lean_object* x_22; lean_object* x_23; uint8_t x_24; -x_22 = lean_ctor_get(x_19, 1); -x_23 = lean_ctor_get(x_19, 0); -lean_dec(x_23); -x_24 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_15, x_16); -if (x_24 == 0) +lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_33 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__7; +x_34 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_33, x_10, x_11, x_31); +x_35 = lean_ctor_get(x_34, 0); +lean_inc(x_35); +x_36 = lean_ctor_get(x_34, 1); +lean_inc(x_36); +lean_dec(x_34); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_13); +x_37 = lean_infer_type(x_13, x_8, x_9, x_10, x_11, x_36); +if (lean_obj_tag(x_37) == 0) { -uint8_t x_25; lean_object* x_26; -lean_free_object(x_19); -x_25 = 0; +lean_object* x_38; lean_object* x_39; lean_object* x_40; uint8_t x_41; uint8_t 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); +lean_dec(x_37); +x_40 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___lambda__1___boxed), 13, 3); +lean_closure_set(x_40, 0, x_17); +lean_closure_set(x_40, 1, x_2); +lean_closure_set(x_40, 2, x_1); +x_41 = 0; +x_42 = 0; +lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); -x_26 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore(x_15, x_16, x_25, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_22); -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_mkCongrArg(x_17, x_27, x_7, x_8, x_9, x_10, x_28); -if (lean_obj_tag(x_29) == 0) +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +x_43 = l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___spec__1___rarg(x_35, x_41, x_38, x_40, x_42, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_39); +if (lean_obj_tag(x_43) == 0) { -uint8_t x_30; -x_30 = !lean_is_exclusive(x_29); -if (x_30 == 0) +lean_object* x_44; lean_object* x_45; uint8_t 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 = 0; +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +x_47 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore(x_13, x_14, x_46, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_45); +if (lean_obj_tag(x_47) == 0) { -lean_object* x_31; lean_object* x_32; -x_31 = lean_ctor_get(x_29, 0); -x_32 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_32, 0, x_31); -lean_ctor_set(x_29, 0, x_32); -return x_29; -} -else +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); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +x_50 = l_Lean_Meta_mkEqNDRec(x_44, x_30, x_48, x_8, x_9, x_10, x_11, x_49); +if (lean_obj_tag(x_50) == 0) { -lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_33 = lean_ctor_get(x_29, 0); -x_34 = lean_ctor_get(x_29, 1); -lean_inc(x_34); -lean_inc(x_33); -lean_dec(x_29); -x_35 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_35, 0, x_33); -x_36 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_36, 0, x_35); -lean_ctor_set(x_36, 1, x_34); -return x_36; -} +lean_object* x_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___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqOfHEqIfNeeded(x_51, x_3, x_8, x_9, x_10, x_11, x_52); +return x_53; } else { -uint8_t x_37; -x_37 = !lean_is_exclusive(x_29); -if (x_37 == 0) +uint8_t x_54; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +x_54 = !lean_is_exclusive(x_50); +if (x_54 == 0) { -return x_29; +return x_50; } else { -lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_38 = lean_ctor_get(x_29, 0); -x_39 = lean_ctor_get(x_29, 1); -lean_inc(x_39); -lean_inc(x_38); -lean_dec(x_29); -x_40 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_40, 0, x_38); -lean_ctor_set(x_40, 1, x_39); -return x_40; +lean_object* x_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; } } } else { -uint8_t x_41; -lean_dec(x_17); +uint8_t x_58; +lean_dec(x_44); +lean_dec(x_30); +lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); -lean_dec(x_7); -x_41 = !lean_is_exclusive(x_26); -if (x_41 == 0) +x_58 = !lean_is_exclusive(x_47); +if (x_58 == 0) { -return x_26; +return x_47; } else { -lean_object* x_42; lean_object* x_43; lean_object* x_44; -x_42 = lean_ctor_get(x_26, 0); -x_43 = lean_ctor_get(x_26, 1); -lean_inc(x_43); -lean_inc(x_42); -lean_dec(x_26); -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_59; lean_object* x_60; lean_object* x_61; +x_59 = lean_ctor_get(x_47, 0); +x_60 = lean_ctor_get(x_47, 1); +lean_inc(x_60); +lean_inc(x_59); +lean_dec(x_47); +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_45; -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_15); +uint8_t x_62; +lean_dec(x_30); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -8150,124 +7899,125 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); -x_45 = lean_box(0); -lean_ctor_set(x_19, 0, x_45); -return x_19; -} +x_62 = !lean_is_exclusive(x_43); +if (x_62 == 0) +{ +return x_43; } else { -lean_object* x_46; uint8_t x_47; -x_46 = lean_ctor_get(x_19, 1); -lean_inc(x_46); -lean_dec(x_19); -x_47 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_15, x_16); -if (x_47 == 0) +lean_object* x_63; lean_object* x_64; lean_object* x_65; +x_63 = lean_ctor_get(x_43, 0); +x_64 = lean_ctor_get(x_43, 1); +lean_inc(x_64); +lean_inc(x_63); +lean_dec(x_43); +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_48; lean_object* x_49; -x_48 = 0; -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -x_49 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore(x_15, x_16, x_48, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_46); -if (lean_obj_tag(x_49) == 0) +uint8_t x_66; +lean_dec(x_35); +lean_dec(x_30); +lean_dec(x_17); +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_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_66 = !lean_is_exclusive(x_37); +if (x_66 == 0) { -lean_object* x_50; lean_object* x_51; lean_object* x_52; -x_50 = lean_ctor_get(x_49, 0); -lean_inc(x_50); -x_51 = lean_ctor_get(x_49, 1); -lean_inc(x_51); -lean_dec(x_49); -x_52 = l_Lean_Meta_mkCongrArg(x_17, x_50, x_7, x_8, x_9, x_10, x_51); -if (lean_obj_tag(x_52) == 0) +return x_37; +} +else { -lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; -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); +lean_object* x_67; lean_object* x_68; lean_object* x_69; +x_67 = lean_ctor_get(x_37, 0); +x_68 = lean_ctor_get(x_37, 1); +lean_inc(x_68); +lean_inc(x_67); +lean_dec(x_37); +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; } -x_56 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_56, 0, 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 { -lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* 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; -} else { - lean_dec_ref(x_52); - x_60 = lean_box(0); -} -if (lean_is_scalar(x_60)) { - x_61 = lean_alloc_ctor(1, 2, 0); -} else { - x_61 = x_60; -} -lean_ctor_set(x_61, 0, x_58); -lean_ctor_set(x_61, 1, x_59); -return x_61; +lean_object* x_70; +lean_dec(x_17); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_70 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqOfHEqIfNeeded(x_30, x_3, x_8, x_9, x_10, x_11, x_31); +return x_70; } } else { -lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; +uint8_t x_71; lean_dec(x_17); +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); -x_62 = lean_ctor_get(x_49, 0); -lean_inc(x_62); -x_63 = lean_ctor_get(x_49, 1); -lean_inc(x_63); -if (lean_is_exclusive(x_49)) { - lean_ctor_release(x_49, 0); - lean_ctor_release(x_49, 1); - x_64 = x_49; -} else { - lean_dec_ref(x_49); - x_64 = lean_box(0); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_71 = !lean_is_exclusive(x_29); +if (x_71 == 0) +{ +return x_29; +} +else +{ +lean_object* x_72; lean_object* x_73; lean_object* x_74; +x_72 = lean_ctor_get(x_29, 0); +x_73 = lean_ctor_get(x_29, 1); +lean_inc(x_73); +lean_inc(x_72); +lean_dec(x_29); +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_64)) { - x_65 = lean_alloc_ctor(1, 2, 0); -} else { - x_65 = x_64; } -lean_ctor_set(x_65, 0, x_62); -lean_ctor_set(x_65, 1, x_63); -return x_65; } } else { -lean_object* x_66; lean_object* x_67; +uint8_t x_75; 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); @@ -8275,387 +8025,1554 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); -x_66 = lean_box(0); -x_67 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_67, 0, x_66); -lean_ctor_set(x_67, 1, x_46); -return x_67; +lean_dec(x_2); +lean_dec(x_1); +x_75 = !lean_is_exclusive(x_21); +if (x_75 == 0) +{ +return x_21; } +else +{ +lean_object* x_76; lean_object* x_77; lean_object* x_78; +x_76 = lean_ctor_get(x_21, 0); +x_77 = lean_ctor_get(x_21, 1); +lean_inc(x_77); +lean_inc(x_76); +lean_dec(x_21); +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 +} +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof___closed__1() { +_start: { -lean_object* x_68; uint8_t x_69; -lean_dec(x_17); -x_68 = lean_ctor_get(x_19, 1); -lean_inc(x_68); -lean_dec(x_19); -x_69 = !lean_is_exclusive(x_20); -if (x_69 == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Init.Data.Option.BasicAux", 25, 25); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof___closed__2() { +_start: { -lean_object* x_70; uint8_t x_71; -x_70 = lean_ctor_get(x_20, 0); -x_71 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_15, x_16); -if (x_71 == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Option.get!", 11, 11); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof___closed__3() { +_start: { -uint8_t x_72; lean_object* x_73; -x_72 = 0; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("value is none", 13, 13); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof___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___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof___closed__2; +x_3 = lean_unsigned_to_nat(16u); +x_4 = lean_unsigned_to_nat(14u); +x_5 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof___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___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof(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_13; +lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); -lean_inc(x_7); -x_73 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore(x_15, x_16, x_72, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_68); -if (lean_obj_tag(x_73) == 0) +x_13 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof_loop(x_1, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_13) == 0) { -lean_object* x_74; lean_object* x_75; lean_object* x_76; -x_74 = lean_ctor_get(x_73, 0); -lean_inc(x_74); -x_75 = lean_ctor_get(x_73, 1); -lean_inc(x_75); -lean_dec(x_73); -x_76 = l_Lean_Meta_mkCongr(x_70, x_74, x_7, x_8, x_9, x_10, x_75); -if (lean_obj_tag(x_76) == 0) +lean_object* x_14; +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +if (lean_obj_tag(x_14) == 0) { -uint8_t x_77; -x_77 = !lean_is_exclusive(x_76); -if (x_77 == 0) +uint8_t x_15; +x_15 = !lean_is_exclusive(x_13); +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_13, 1); +x_17 = lean_ctor_get(x_13, 0); +lean_dec(x_17); +x_18 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof___closed__4; +x_19 = l_panic___at_Lean_Expr_appFn_x21___spec__1(x_18); +if (x_3 == 0) +{ +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_ctor_set(x_13, 0, x_19); +return x_13; +} +else +{ +lean_object* x_20; +lean_free_object(x_13); +x_20 = l_Lean_Meta_mkHEqOfEq(x_19, x_8, x_9, x_10, x_11, x_16); +return x_20; +} +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_13, 1); +lean_inc(x_21); +lean_dec(x_13); +x_22 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof___closed__4; +x_23 = l_panic___at_Lean_Expr_appFn_x21___spec__1(x_22); +if (x_3 == 0) +{ +lean_object* x_24; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +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 +{ +lean_object* x_25; +x_25 = l_Lean_Meta_mkHEqOfEq(x_23, x_8, x_9, x_10, x_11, x_21); +return x_25; +} +} +} +else +{ +if (x_3 == 0) +{ +uint8_t x_26; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +x_26 = !lean_is_exclusive(x_13); +if (x_26 == 0) +{ +lean_object* x_27; lean_object* x_28; +x_27 = lean_ctor_get(x_13, 0); +lean_dec(x_27); +x_28 = lean_ctor_get(x_14, 0); +lean_inc(x_28); +lean_dec(x_14); +lean_ctor_set(x_13, 0, x_28); +return x_13; +} +else +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_29 = lean_ctor_get(x_13, 1); +lean_inc(x_29); +lean_dec(x_13); +x_30 = lean_ctor_get(x_14, 0); +lean_inc(x_30); +lean_dec(x_14); +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_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_13, 1); +lean_inc(x_32); +lean_dec(x_13); +x_33 = lean_ctor_get(x_14, 0); +lean_inc(x_33); +lean_dec(x_14); +x_34 = l_Lean_Meta_mkHEqOfEq(x_33, x_8, x_9, x_10, x_11, x_32); +return x_34; +} +} +} +else +{ +uint8_t x_35; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +x_35 = !lean_is_exclusive(x_13); +if (x_35 == 0) +{ +return x_13; +} +else +{ +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_dec(x_13); +x_38 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_38, 0, x_36); +lean_ctor_set(x_38, 1, x_37); +return x_38; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof_loop(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, 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 = l_Lean_Expr_isApp(x_1); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +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_11); +return x_14; +} +else +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_15 = l_Lean_Expr_appArg_x21(x_1); +x_16 = l_Lean_Expr_appArg_x21(x_2); +x_17 = l_Lean_Expr_appFn_x21(x_1); +x_18 = l_Lean_Expr_appFn_x21(x_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_4); +lean_inc(x_3); +x_19 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof_loop(x_17, x_18, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_18); +if (lean_obj_tag(x_19) == 0) +{ +lean_object* x_20; +x_20 = lean_ctor_get(x_19, 0); +lean_inc(x_20); +if (lean_obj_tag(x_20) == 0) +{ +uint8_t x_21; +x_21 = !lean_is_exclusive(x_19); +if (x_21 == 0) +{ +lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_22 = lean_ctor_get(x_19, 1); +x_23 = lean_ctor_get(x_19, 0); +lean_dec(x_23); +x_24 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_15, x_16); +if (x_24 == 0) +{ +uint8_t x_25; lean_object* x_26; +lean_free_object(x_19); +x_25 = 0; +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +x_26 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore(x_15, x_16, x_25, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_22); +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_mkCongrArg(x_17, x_27, x_7, x_8, x_9, x_10, x_28); +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; +x_31 = lean_ctor_get(x_29, 0); +x_32 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_29, 0, x_32); +return x_29; +} +else +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_33 = lean_ctor_get(x_29, 0); +x_34 = lean_ctor_get(x_29, 1); +lean_inc(x_34); +lean_inc(x_33); +lean_dec(x_29); +x_35 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_35, 0, x_33); +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 +{ +uint8_t x_37; +x_37 = !lean_is_exclusive(x_29); +if (x_37 == 0) +{ +return x_29; +} +else +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_38 = lean_ctor_get(x_29, 0); +x_39 = lean_ctor_get(x_29, 1); +lean_inc(x_39); +lean_inc(x_38); +lean_dec(x_29); +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_17); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +x_41 = !lean_is_exclusive(x_26); +if (x_41 == 0) +{ +return x_26; +} +else +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_42 = lean_ctor_get(x_26, 0); +x_43 = lean_ctor_get(x_26, 1); +lean_inc(x_43); +lean_inc(x_42); +lean_dec(x_26); +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_45; +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_45 = lean_box(0); +lean_ctor_set(x_19, 0, x_45); +return x_19; +} +} +else +{ +lean_object* x_46; uint8_t x_47; +x_46 = lean_ctor_get(x_19, 1); +lean_inc(x_46); +lean_dec(x_19); +x_47 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_15, x_16); +if (x_47 == 0) +{ +uint8_t x_48; lean_object* x_49; +x_48 = 0; +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +x_49 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore(x_15, x_16, x_48, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_46); +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, 0); +lean_inc(x_50); +x_51 = lean_ctor_get(x_49, 1); +lean_inc(x_51); +lean_dec(x_49); +x_52 = l_Lean_Meta_mkCongrArg(x_17, x_50, x_7, x_8, x_9, x_10, x_51); +if (lean_obj_tag(x_52) == 0) +{ +lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; +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(1, 1, 0); +lean_ctor_set(x_56, 0, 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 +{ +lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* 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; +} else { + lean_dec_ref(x_52); + x_60 = lean_box(0); +} +if (lean_is_scalar(x_60)) { + x_61 = lean_alloc_ctor(1, 2, 0); +} else { + x_61 = x_60; +} +lean_ctor_set(x_61, 0, x_58); +lean_ctor_set(x_61, 1, x_59); +return x_61; +} +} +else +{ +lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; +lean_dec(x_17); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +x_62 = lean_ctor_get(x_49, 0); +lean_inc(x_62); +x_63 = lean_ctor_get(x_49, 1); +lean_inc(x_63); +if (lean_is_exclusive(x_49)) { + lean_ctor_release(x_49, 0); + lean_ctor_release(x_49, 1); + x_64 = x_49; +} else { + lean_dec_ref(x_49); + x_64 = lean_box(0); +} +if (lean_is_scalar(x_64)) { + x_65 = lean_alloc_ctor(1, 2, 0); +} else { + x_65 = x_64; +} +lean_ctor_set(x_65, 0, x_62); +lean_ctor_set(x_65, 1, x_63); +return x_65; +} +} +else +{ +lean_object* x_66; lean_object* x_67; +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_66 = lean_box(0); +x_67 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_67, 0, x_66); +lean_ctor_set(x_67, 1, x_46); +return x_67; +} +} +} +else +{ +lean_object* x_68; uint8_t x_69; +lean_dec(x_17); +x_68 = lean_ctor_get(x_19, 1); +lean_inc(x_68); +lean_dec(x_19); +x_69 = !lean_is_exclusive(x_20); +if (x_69 == 0) +{ +lean_object* x_70; uint8_t x_71; +x_70 = lean_ctor_get(x_20, 0); +x_71 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_15, x_16); +if (x_71 == 0) +{ +uint8_t x_72; lean_object* x_73; +x_72 = 0; +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +x_73 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore(x_15, x_16, x_72, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_68); +if (lean_obj_tag(x_73) == 0) +{ +lean_object* x_74; lean_object* x_75; lean_object* x_76; +x_74 = lean_ctor_get(x_73, 0); +lean_inc(x_74); +x_75 = lean_ctor_get(x_73, 1); +lean_inc(x_75); +lean_dec(x_73); +x_76 = l_Lean_Meta_mkCongr(x_70, x_74, x_7, x_8, x_9, x_10, x_75); +if (lean_obj_tag(x_76) == 0) +{ +uint8_t x_77; +x_77 = !lean_is_exclusive(x_76); +if (x_77 == 0) +{ +lean_object* x_78; +x_78 = lean_ctor_get(x_76, 0); +lean_ctor_set(x_20, 0, x_78); +lean_ctor_set(x_76, 0, x_20); +return x_76; +} +else +{ +lean_object* x_79; lean_object* x_80; lean_object* x_81; +x_79 = lean_ctor_get(x_76, 0); +x_80 = lean_ctor_get(x_76, 1); +lean_inc(x_80); +lean_inc(x_79); +lean_dec(x_76); +lean_ctor_set(x_20, 0, x_79); +x_81 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_81, 0, x_20); +lean_ctor_set(x_81, 1, x_80); +return x_81; +} +} +else +{ +uint8_t x_82; +lean_free_object(x_20); +x_82 = !lean_is_exclusive(x_76); +if (x_82 == 0) { -lean_object* x_78; -x_78 = lean_ctor_get(x_76, 0); -lean_ctor_set(x_20, 0, x_78); -lean_ctor_set(x_76, 0, x_20); return x_76; } else { -lean_object* x_79; lean_object* x_80; lean_object* x_81; -x_79 = lean_ctor_get(x_76, 0); -x_80 = lean_ctor_get(x_76, 1); +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 +{ +uint8_t x_86; +lean_free_object(x_20); +lean_dec(x_70); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +x_86 = !lean_is_exclusive(x_73); +if (x_86 == 0) +{ +return x_73; +} +else +{ +lean_object* x_87; lean_object* x_88; lean_object* x_89; +x_87 = lean_ctor_get(x_73, 0); +x_88 = lean_ctor_get(x_73, 1); +lean_inc(x_88); +lean_inc(x_87); +lean_dec(x_73); +x_89 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_89, 0, x_87); +lean_ctor_set(x_89, 1, x_88); +return x_89; +} +} +} +else +{ +lean_object* x_90; +lean_dec(x_16); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_90 = l_Lean_Meta_mkCongrFun(x_70, x_15, x_7, x_8, x_9, x_10, x_68); +if (lean_obj_tag(x_90) == 0) +{ +uint8_t x_91; +x_91 = !lean_is_exclusive(x_90); +if (x_91 == 0) +{ +lean_object* x_92; +x_92 = lean_ctor_get(x_90, 0); +lean_ctor_set(x_20, 0, x_92); +lean_ctor_set(x_90, 0, x_20); +return x_90; +} +else +{ +lean_object* x_93; lean_object* x_94; lean_object* x_95; +x_93 = lean_ctor_get(x_90, 0); +x_94 = lean_ctor_get(x_90, 1); +lean_inc(x_94); +lean_inc(x_93); +lean_dec(x_90); +lean_ctor_set(x_20, 0, x_93); +x_95 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_95, 0, x_20); +lean_ctor_set(x_95, 1, x_94); +return x_95; +} +} +else +{ +uint8_t x_96; +lean_free_object(x_20); +x_96 = !lean_is_exclusive(x_90); +if (x_96 == 0) +{ +return x_90; +} +else +{ +lean_object* x_97; lean_object* x_98; lean_object* x_99; +x_97 = lean_ctor_get(x_90, 0); +x_98 = lean_ctor_get(x_90, 1); +lean_inc(x_98); +lean_inc(x_97); +lean_dec(x_90); +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 +{ +lean_object* x_100; uint8_t x_101; +x_100 = lean_ctor_get(x_20, 0); +lean_inc(x_100); +lean_dec(x_20); +x_101 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_15, x_16); +if (x_101 == 0) +{ +uint8_t x_102; lean_object* x_103; +x_102 = 0; +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +x_103 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore(x_15, x_16, x_102, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_68); +if (lean_obj_tag(x_103) == 0) +{ +lean_object* x_104; lean_object* x_105; lean_object* x_106; +x_104 = lean_ctor_get(x_103, 0); +lean_inc(x_104); +x_105 = lean_ctor_get(x_103, 1); +lean_inc(x_105); +lean_dec(x_103); +x_106 = l_Lean_Meta_mkCongr(x_100, x_104, x_7, x_8, x_9, x_10, x_105); +if (lean_obj_tag(x_106) == 0) +{ +lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; +x_107 = lean_ctor_get(x_106, 0); +lean_inc(x_107); +x_108 = lean_ctor_get(x_106, 1); +lean_inc(x_108); +if (lean_is_exclusive(x_106)) { + lean_ctor_release(x_106, 0); + lean_ctor_release(x_106, 1); + x_109 = x_106; +} else { + lean_dec_ref(x_106); + x_109 = lean_box(0); +} +x_110 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_110, 0, x_107); +if (lean_is_scalar(x_109)) { + x_111 = lean_alloc_ctor(0, 2, 0); +} else { + x_111 = x_109; +} +lean_ctor_set(x_111, 0, x_110); +lean_ctor_set(x_111, 1, x_108); +return x_111; +} +else +{ +lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; +x_112 = lean_ctor_get(x_106, 0); +lean_inc(x_112); +x_113 = lean_ctor_get(x_106, 1); +lean_inc(x_113); +if (lean_is_exclusive(x_106)) { + lean_ctor_release(x_106, 0); + lean_ctor_release(x_106, 1); + x_114 = x_106; +} else { + lean_dec_ref(x_106); + x_114 = lean_box(0); +} +if (lean_is_scalar(x_114)) { + x_115 = lean_alloc_ctor(1, 2, 0); +} else { + x_115 = x_114; +} +lean_ctor_set(x_115, 0, x_112); +lean_ctor_set(x_115, 1, x_113); +return x_115; +} +} +else +{ +lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; +lean_dec(x_100); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +x_116 = lean_ctor_get(x_103, 0); +lean_inc(x_116); +x_117 = lean_ctor_get(x_103, 1); +lean_inc(x_117); +if (lean_is_exclusive(x_103)) { + lean_ctor_release(x_103, 0); + lean_ctor_release(x_103, 1); + x_118 = x_103; +} else { + lean_dec_ref(x_103); + x_118 = lean_box(0); +} +if (lean_is_scalar(x_118)) { + x_119 = lean_alloc_ctor(1, 2, 0); +} else { + x_119 = x_118; +} +lean_ctor_set(x_119, 0, x_116); +lean_ctor_set(x_119, 1, x_117); +return x_119; +} +} +else +{ +lean_object* x_120; +lean_dec(x_16); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_120 = l_Lean_Meta_mkCongrFun(x_100, x_15, x_7, x_8, x_9, x_10, x_68); +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; +x_121 = lean_ctor_get(x_120, 0); +lean_inc(x_121); +x_122 = lean_ctor_get(x_120, 1); +lean_inc(x_122); +if (lean_is_exclusive(x_120)) { + lean_ctor_release(x_120, 0); + lean_ctor_release(x_120, 1); + x_123 = x_120; +} else { + lean_dec_ref(x_120); + x_123 = lean_box(0); +} +x_124 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_124, 0, x_121); +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_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; +x_126 = lean_ctor_get(x_120, 0); +lean_inc(x_126); +x_127 = lean_ctor_get(x_120, 1); +lean_inc(x_127); +if (lean_is_exclusive(x_120)) { + lean_ctor_release(x_120, 0); + lean_ctor_release(x_120, 1); + x_128 = x_120; +} else { + lean_dec_ref(x_120); + x_128 = lean_box(0); +} +if (lean_is_scalar(x_128)) { + x_129 = lean_alloc_ctor(1, 2, 0); +} else { + x_129 = x_128; +} +lean_ctor_set(x_129, 0, x_126); +lean_ctor_set(x_129, 1, x_127); +return x_129; +} +} +} +} +} +else +{ +uint8_t x_130; +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_130 = !lean_is_exclusive(x_19); +if (x_130 == 0) +{ +return x_19; +} +else +{ +lean_object* x_131; lean_object* x_132; lean_object* x_133; +x_131 = lean_ctor_get(x_19, 0); +x_132 = lean_ctor_get(x_19, 1); +lean_inc(x_132); +lean_inc(x_131); +lean_dec(x_19); +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; +} +} +} +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("hasSameRoot enodes a₁ b₂ && hasSameRoot enodes b₁ a₂\n ", 67, 59); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore___lambda__1___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__1; +x_3 = lean_string_append(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("_private.Lean.Meta.Tactic.Grind.Proof.0.Lean.Meta.Grind.mkEqCongrProof", 70, 70); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__7___lambda__1___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__3; +x_3 = lean_unsigned_to_nat(174u); +x_4 = lean_unsigned_to_nat(6u); +x_5 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__2; +x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); +return x_6; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("eq_congr'", 9, 9); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__5; +x_2 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__6; +x_3 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__5; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("eq_congr", 8, 8); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__5; +x_2 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__6; +x_3 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__7; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, uint8_t x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19) { +_start: +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_57; +x_20 = lean_st_ref_get(x_11, x_19); +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, 1); +lean_inc(x_23); +lean_dec(x_21); +x_24 = l_Lean_Expr_constLevels_x21(x_1); +x_57 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_4, x_8); +if (x_57 == 0) +{ +lean_object* x_58; +lean_dec(x_24); +lean_dec(x_23); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_58 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof(x_5, x_6, x_7, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_22); +return x_58; +} +else +{ +uint8_t x_59; +lean_dec(x_6); +lean_dec(x_5); +lean_inc(x_23); +x_59 = l_Lean_Meta_Grind_hasSameRoot(x_23, x_2, x_9); +if (x_59 == 0) +{ +lean_object* x_60; +x_60 = lean_box(0); +x_25 = x_60; +goto block_56; +} +else +{ +uint8_t x_61; +lean_inc(x_23); +x_61 = l_Lean_Meta_Grind_hasSameRoot(x_23, x_3, x_10); +if (x_61 == 0) +{ +lean_object* x_62; +x_62 = lean_box(0); +x_25 = x_62; +goto block_56; +} +else +{ +uint8_t x_63; lean_object* x_64; +lean_dec(x_23); +x_63 = 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); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_9); +lean_inc(x_2); +x_64 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore(x_2, x_9, x_63, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_22); +if (lean_obj_tag(x_64) == 0) +{ +lean_object* x_65; lean_object* x_66; lean_object* x_67; +x_65 = lean_ctor_get(x_64, 0); +lean_inc(x_65); +x_66 = lean_ctor_get(x_64, 1); +lean_inc(x_66); +lean_dec(x_64); +lean_inc(x_10); +lean_inc(x_3); +x_67 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore(x_3, x_10, x_63, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_66); +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; lean_object* x_71; lean_object* x_72; +x_69 = lean_ctor_get(x_67, 0); +x_70 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__8; +x_71 = l_Lean_Expr_const___override(x_70, x_24); +x_72 = l_Lean_mkApp7(x_71, x_4, x_2, x_3, x_9, x_10, x_65, x_69); +lean_ctor_set(x_67, 0, x_72); +return x_67; +} +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; +x_73 = lean_ctor_get(x_67, 0); +x_74 = lean_ctor_get(x_67, 1); +lean_inc(x_74); +lean_inc(x_73); +lean_dec(x_67); +x_75 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__8; +x_76 = l_Lean_Expr_const___override(x_75, x_24); +x_77 = l_Lean_mkApp7(x_76, x_4, x_2, x_3, x_9, x_10, x_65, x_73); +x_78 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_78, 0, x_77); +lean_ctor_set(x_78, 1, x_74); +return x_78; +} +} +else +{ +uint8_t x_79; +lean_dec(x_65); +lean_dec(x_24); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_79 = !lean_is_exclusive(x_67); +if (x_79 == 0) +{ +return x_67; +} +else +{ +lean_object* x_80; lean_object* x_81; lean_object* x_82; +x_80 = lean_ctor_get(x_67, 0); +x_81 = lean_ctor_get(x_67, 1); +lean_inc(x_81); lean_inc(x_80); -lean_inc(x_79); -lean_dec(x_76); -lean_ctor_set(x_20, 0, x_79); -x_81 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_81, 0, x_20); -lean_ctor_set(x_81, 1, x_80); -return x_81; +lean_dec(x_67); +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_83; +lean_dec(x_24); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_83 = !lean_is_exclusive(x_64); +if (x_83 == 0) +{ +return x_64; +} +else +{ +lean_object* x_84; lean_object* x_85; lean_object* x_86; +x_84 = lean_ctor_get(x_64, 0); +x_85 = lean_ctor_get(x_64, 1); +lean_inc(x_85); +lean_inc(x_84); +lean_dec(x_64); +x_86 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_86, 0, x_84); +lean_ctor_set(x_86, 1, x_85); +return x_86; +} +} +} +} } +block_56: +{ +uint8_t x_26; +lean_dec(x_25); +lean_inc(x_23); +x_26 = l_Lean_Meta_Grind_hasSameRoot(x_23, x_2, x_10); +if (x_26 == 0) +{ +lean_object* x_27; lean_object* x_28; +lean_dec(x_24); +lean_dec(x_23); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_27 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__4; +x_28 = l_panic___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__9(x_27, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_22); +return x_28; } else { -uint8_t x_82; -lean_free_object(x_20); -x_82 = !lean_is_exclusive(x_76); -if (x_82 == 0) +uint8_t x_29; +x_29 = l_Lean_Meta_Grind_hasSameRoot(x_23, x_3, x_9); +if (x_29 == 0) { -return x_76; +lean_object* x_30; lean_object* x_31; +lean_dec(x_24); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_30 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__4; +x_31 = l_panic___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__9(x_30, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_22); +return x_31; } 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; +uint8_t x_32; lean_object* x_33; +x_32 = 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); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_2); +x_33 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore(x_2, x_10, x_32, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_22); +if (lean_obj_tag(x_33) == 0) +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_33, 0); +lean_inc(x_34); +x_35 = lean_ctor_get(x_33, 1); +lean_inc(x_35); +lean_dec(x_33); +lean_inc(x_9); +lean_inc(x_3); +x_36 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore(x_3, x_9, x_32, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_35); +if (lean_obj_tag(x_36) == 0) +{ +uint8_t x_37; +x_37 = !lean_is_exclusive(x_36); +if (x_37 == 0) +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_38 = lean_ctor_get(x_36, 0); +x_39 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__6; +x_40 = l_Lean_Expr_const___override(x_39, x_24); +x_41 = l_Lean_mkApp7(x_40, x_4, x_2, x_3, x_9, x_10, x_34, x_38); +lean_ctor_set(x_36, 0, x_41); +return x_36; } +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; +x_42 = lean_ctor_get(x_36, 0); +x_43 = lean_ctor_get(x_36, 1); +lean_inc(x_43); +lean_inc(x_42); +lean_dec(x_36); +x_44 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__6; +x_45 = l_Lean_Expr_const___override(x_44, x_24); +x_46 = l_Lean_mkApp7(x_45, x_4, x_2, x_3, x_9, x_10, x_34, x_42); +x_47 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_47, 0, x_46); +lean_ctor_set(x_47, 1, x_43); +return x_47; } } else { -uint8_t x_86; -lean_free_object(x_20); -lean_dec(x_70); +uint8_t x_48; +lean_dec(x_34); +lean_dec(x_24); lean_dec(x_10); lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -x_86 = !lean_is_exclusive(x_73); -if (x_86 == 0) +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_48 = !lean_is_exclusive(x_36); +if (x_48 == 0) { -return x_73; +return x_36; } else { -lean_object* x_87; lean_object* x_88; lean_object* x_89; -x_87 = lean_ctor_get(x_73, 0); -x_88 = lean_ctor_get(x_73, 1); -lean_inc(x_88); -lean_inc(x_87); -lean_dec(x_73); -x_89 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_89, 0, x_87); -lean_ctor_set(x_89, 1, x_88); -return x_89; +lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_49 = lean_ctor_get(x_36, 0); +x_50 = lean_ctor_get(x_36, 1); +lean_inc(x_50); +lean_inc(x_49); +lean_dec(x_36); +x_51 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_51, 0, x_49); +lean_ctor_set(x_51, 1, x_50); +return x_51; } } } else { -lean_object* x_90; +uint8_t x_52; +lean_dec(x_24); +lean_dec(x_18); +lean_dec(x_17); lean_dec(x_16); -lean_dec(x_6); -lean_dec(x_5); +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_4); lean_dec(x_3); -x_90 = l_Lean_Meta_mkCongrFun(x_70, x_15, x_7, x_8, x_9, x_10, x_68); -if (lean_obj_tag(x_90) == 0) -{ -uint8_t x_91; -x_91 = !lean_is_exclusive(x_90); -if (x_91 == 0) +lean_dec(x_2); +x_52 = !lean_is_exclusive(x_33); +if (x_52 == 0) { -lean_object* x_92; -x_92 = lean_ctor_get(x_90, 0); -lean_ctor_set(x_20, 0, x_92); -lean_ctor_set(x_90, 0, x_20); -return x_90; +return x_33; } else { -lean_object* x_93; lean_object* x_94; lean_object* x_95; -x_93 = lean_ctor_get(x_90, 0); -x_94 = lean_ctor_get(x_90, 1); -lean_inc(x_94); -lean_inc(x_93); -lean_dec(x_90); -lean_ctor_set(x_20, 0, x_93); -x_95 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_95, 0, x_20); -lean_ctor_set(x_95, 1, x_94); -return x_95; +lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_53 = lean_ctor_get(x_33, 0); +x_54 = lean_ctor_get(x_33, 1); +lean_inc(x_54); +lean_inc(x_53); +lean_dec(x_33); +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_96; -lean_free_object(x_20); -x_96 = !lean_is_exclusive(x_90); -if (x_96 == 0) -{ -return x_90; } -else -{ -lean_object* x_97; lean_object* x_98; lean_object* x_99; -x_97 = lean_ctor_get(x_90, 0); -x_98 = lean_ctor_get(x_90, 1); -lean_inc(x_98); -lean_inc(x_97); -lean_dec(x_90); -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 -{ -lean_object* x_100; uint8_t x_101; -x_100 = lean_ctor_get(x_20, 0); -lean_inc(x_100); -lean_dec(x_20); -x_101 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_15, x_16); -if (x_101 == 0) -{ -uint8_t x_102; lean_object* x_103; -x_102 = 0; -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -x_103 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore(x_15, x_16, x_102, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_68); -if (lean_obj_tag(x_103) == 0) -{ -lean_object* x_104; lean_object* x_105; lean_object* x_106; -x_104 = lean_ctor_get(x_103, 0); -lean_inc(x_104); -x_105 = lean_ctor_get(x_103, 1); -lean_inc(x_105); -lean_dec(x_103); -x_106 = l_Lean_Meta_mkCongr(x_100, x_104, x_7, x_8, x_9, x_10, x_105); -if (lean_obj_tag(x_106) == 0) +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__2___closed__1() { +_start: { -lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; -x_107 = lean_ctor_get(x_106, 0); -lean_inc(x_107); -x_108 = lean_ctor_get(x_106, 1); -lean_inc(x_108); -if (lean_is_exclusive(x_106)) { - lean_ctor_release(x_106, 0); - lean_ctor_release(x_106, 1); - x_109 = x_106; -} else { - lean_dec_ref(x_106); - x_109 = lean_box(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_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__7___lambda__1___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__3; +x_3 = lean_unsigned_to_nat(166u); +x_4 = lean_unsigned_to_nat(34u); +x_5 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__7___lambda__1___closed__3; +x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); +return x_6; } -x_110 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_110, 0, x_107); -if (lean_is_scalar(x_109)) { - x_111 = lean_alloc_ctor(0, 2, 0); -} else { - x_111 = x_109; } -lean_ctor_set(x_111, 0, x_110); -lean_ctor_set(x_111, 1, x_108); -return x_111; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; lean_object* x_12; +x_11 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__2___closed__1; +x_12 = l_panic___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__9(x_11, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_12; } -else +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__3___closed__1() { +_start: { -lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; -x_112 = lean_ctor_get(x_106, 0); -lean_inc(x_112); -x_113 = lean_ctor_get(x_106, 1); -lean_inc(x_113); -if (lean_is_exclusive(x_106)) { - lean_ctor_release(x_106, 0); - lean_ctor_release(x_106, 1); - x_114 = x_106; -} else { - lean_dec_ref(x_106); - x_114 = lean_box(0); +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__2___boxed), 10, 0); +return x_1; } -if (lean_is_scalar(x_114)) { - x_115 = lean_alloc_ctor(1, 2, 0); -} else { - x_115 = x_114; } -lean_ctor_set(x_115, 0, x_112); -lean_ctor_set(x_115, 1, x_113); -return x_115; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___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, 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; uint8_t x_19; +x_17 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__3___closed__1; +lean_inc(x_1); +x_18 = l_Lean_Expr_cleanupAnnotations(x_1); +x_19 = l_Lean_Expr_isApp(x_18); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; +lean_dec(x_18); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_2); +lean_dec(x_1); +x_20 = lean_box(0); +x_21 = lean_apply_10(x_17, x_20, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +return x_21; } +else +{ +lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_22 = l_Lean_Expr_appArg(x_18, lean_box(0)); +x_23 = l_Lean_Expr_appFnCleanup(x_18, lean_box(0)); +x_24 = l_Lean_Expr_isApp(x_23); +if (x_24 == 0) +{ +lean_object* x_25; lean_object* x_26; +lean_dec(x_23); +lean_dec(x_22); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_2); +lean_dec(x_1); +x_25 = lean_box(0); +x_26 = lean_apply_10(x_17, x_25, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +return x_26; } else { -lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; -lean_dec(x_100); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); +lean_object* x_27; lean_object* x_28; uint8_t x_29; +x_27 = l_Lean_Expr_appArg(x_23, lean_box(0)); +x_28 = l_Lean_Expr_appFnCleanup(x_23, lean_box(0)); +x_29 = l_Lean_Expr_isApp(x_28); +if (x_29 == 0) +{ +lean_object* x_30; lean_object* x_31; +lean_dec(x_28); +lean_dec(x_27); +lean_dec(x_22); lean_dec(x_7); -x_116 = lean_ctor_get(x_103, 0); -lean_inc(x_116); -x_117 = lean_ctor_get(x_103, 1); -lean_inc(x_117); -if (lean_is_exclusive(x_103)) { - lean_ctor_release(x_103, 0); - lean_ctor_release(x_103, 1); - x_118 = x_103; -} else { - lean_dec_ref(x_103); - x_118 = lean_box(0); -} -if (lean_is_scalar(x_118)) { - x_119 = lean_alloc_ctor(1, 2, 0); -} else { - x_119 = x_118; -} -lean_ctor_set(x_119, 0, x_116); -lean_ctor_set(x_119, 1, x_117); -return x_119; -} +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_2); +lean_dec(x_1); +x_30 = lean_box(0); +x_31 = lean_apply_10(x_17, x_30, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +return x_31; } else { -lean_object* x_120; -lean_dec(x_16); +lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; +x_32 = l_Lean_Expr_appArg(x_28, lean_box(0)); +x_33 = l_Lean_Expr_appFnCleanup(x_28, lean_box(0)); +x_34 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_isEqProof___closed__2; +x_35 = l_Lean_Expr_isConstOf(x_33, x_34); +lean_dec(x_33); +if (x_35 == 0) +{ +lean_object* x_36; lean_object* x_37; +lean_dec(x_32); +lean_dec(x_27); +lean_dec(x_22); +lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -x_120 = l_Lean_Meta_mkCongrFun(x_100, x_15, x_7, x_8, x_9, x_10, x_68); -if (lean_obj_tag(x_120) == 0) +lean_dec(x_2); +lean_dec(x_1); +x_36 = lean_box(0); +x_37 = lean_apply_10(x_17, x_36, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +return x_37; +} +else { -lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; -x_121 = lean_ctor_get(x_120, 0); -lean_inc(x_121); -x_122 = lean_ctor_get(x_120, 1); -lean_inc(x_122); -if (lean_is_exclusive(x_120)) { - lean_ctor_release(x_120, 0); - lean_ctor_release(x_120, 1); - x_123 = x_120; -} else { - lean_dec_ref(x_120); - x_123 = lean_box(0); +lean_object* x_38; +x_38 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1(x_4, x_6, x_7, x_5, x_2, x_1, x_3, x_32, x_27, x_22, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +lean_dec(x_32); +return x_38; } -x_124 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_124, 0, x_121); -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 +} +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__4___closed__1() { +_start: { -lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; -x_126 = lean_ctor_get(x_120, 0); -lean_inc(x_126); -x_127 = lean_ctor_get(x_120, 1); -lean_inc(x_127); -if (lean_is_exclusive(x_120)) { - lean_ctor_release(x_120, 0); - lean_ctor_release(x_120, 1); - x_128 = x_120; -} else { - lean_dec_ref(x_120); - x_128 = lean_box(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_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__7___lambda__1___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__3; +x_3 = lean_unsigned_to_nat(165u); +x_4 = lean_unsigned_to_nat(36u); +x_5 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__7___lambda__1___closed__3; +x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); +return x_6; } -if (lean_is_scalar(x_128)) { - x_129 = lean_alloc_ctor(1, 2, 0); -} else { - x_129 = x_128; } -lean_ctor_set(x_129, 0, x_126); -lean_ctor_set(x_129, 1, x_127); -return x_129; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; lean_object* x_12; +x_11 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__4___closed__1; +x_12 = l_panic___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__9(x_11, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_12; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__4___boxed), 10, 0); +return x_1; } } +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof(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_13; lean_object* x_14; uint8_t x_15; +x_13 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___closed__1; +lean_inc(x_1); +x_14 = l_Lean_Expr_cleanupAnnotations(x_1); +x_15 = l_Lean_Expr_isApp(x_14); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; +lean_dec(x_14); +lean_dec(x_2); +lean_dec(x_1); +x_16 = lean_box(0); +x_17 = lean_apply_10(x_13, x_16, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_17; } +else +{ +lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_18 = l_Lean_Expr_appArg(x_14, lean_box(0)); +x_19 = l_Lean_Expr_appFnCleanup(x_14, lean_box(0)); +x_20 = l_Lean_Expr_isApp(x_19); +if (x_20 == 0) +{ +lean_object* x_21; lean_object* x_22; +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_2); +lean_dec(x_1); +x_21 = lean_box(0); +x_22 = lean_apply_10(x_13, x_21, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_22; } +else +{ +lean_object* x_23; lean_object* x_24; uint8_t x_25; +x_23 = l_Lean_Expr_appArg(x_19, lean_box(0)); +x_24 = l_Lean_Expr_appFnCleanup(x_19, lean_box(0)); +x_25 = l_Lean_Expr_isApp(x_24); +if (x_25 == 0) +{ +lean_object* x_26; lean_object* x_27; +lean_dec(x_24); +lean_dec(x_23); +lean_dec(x_18); +lean_dec(x_2); +lean_dec(x_1); +x_26 = lean_box(0); +x_27 = lean_apply_10(x_13, x_26, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_27; } else { -uint8_t x_130; -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -x_130 = !lean_is_exclusive(x_19); -if (x_130 == 0) +lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; +x_28 = l_Lean_Expr_appArg(x_24, lean_box(0)); +x_29 = l_Lean_Expr_appFnCleanup(x_24, lean_box(0)); +x_30 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_isEqProof___closed__2; +x_31 = l_Lean_Expr_isConstOf(x_29, x_30); +if (x_31 == 0) { -return x_19; +lean_object* x_32; lean_object* x_33; +lean_dec(x_29); +lean_dec(x_28); +lean_dec(x_23); +lean_dec(x_18); +lean_dec(x_2); +lean_dec(x_1); +x_32 = lean_box(0); +x_33 = lean_apply_10(x_13, x_32, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_33; } else { -lean_object* x_131; lean_object* x_132; lean_object* x_133; -x_131 = lean_ctor_get(x_19, 0); -x_132 = lean_ctor_get(x_19, 1); -lean_inc(x_132); -lean_inc(x_131); -lean_dec(x_19); -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_34; +x_34 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__3(x_2, x_1, x_3, x_29, x_28, x_23, x_18, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_29); +return x_34; +} } } } @@ -8673,8 +9590,8 @@ static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__10; -x_2 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__11; +x_1 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__5; +x_2 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__6; x_3 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkNestedProofCongr___closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; @@ -8769,7 +9686,7 @@ static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__7___lambda__1___closed__1; x_2 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkProofTo___lambda__1___closed__1; -x_3 = lean_unsigned_to_nat(181u); +x_3 = lean_unsigned_to_nat(203u); x_4 = lean_unsigned_to_nat(35u); x_5 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__7___lambda__1___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -8782,7 +9699,7 @@ static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__7___lambda__1___closed__1; x_2 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkProofTo___lambda__1___closed__1; -x_3 = lean_unsigned_to_nat(182u); +x_3 = lean_unsigned_to_nat(204u); x_4 = lean_unsigned_to_nat(29u); x_5 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__7___lambda__1___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -9133,11 +10050,11 @@ return x_17; } } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof_loop___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_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_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof_loop___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, 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___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof_loop(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___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof_loop(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); @@ -9201,7 +10118,17 @@ x_17 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_realizeEqProo return x_17; } } -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___spec__1___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, 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___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof(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_Meta_withLocalDecl___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___spec__1___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { _start: { uint8_t x_15; uint8_t x_16; lean_object* x_17; @@ -9209,29 +10136,29 @@ x_15 = lean_unbox(x_2); lean_dec(x_2); x_16 = lean_unbox(x_5); lean_dec(x_5); -x_17 = l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___spec__1___rarg(x_1, x_15, x_3, x_4, x_16, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +x_17 = l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___spec__1___rarg(x_1, x_15, x_3, x_4, x_16, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); return x_17; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___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 lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___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_13; -x_13 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___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_object* x_14; +x_14 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); -return x_13; +return x_14; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_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_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, 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___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof(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_14 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof(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; } } @@ -9257,6 +10184,76 @@ lean_dec(x_1); return x_12; } } +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +lean_object* x_18 = _args[17]; +lean_object* x_19 = _args[18]; +_start: +{ +uint8_t x_20; lean_object* x_21; +x_20 = lean_unbox(x_7); +lean_dec(x_7); +x_21 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_20, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19); +lean_dec(x_8); +lean_dec(x_1); +return x_21; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___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_1); +return x_11; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +_start: +{ +uint8_t x_17; lean_object* x_18; +x_17 = lean_unbox(x_3); +lean_dec(x_3); +x_18 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__3(x_1, x_2, x_17, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +lean_dec(x_4); +return x_18; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___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_11; +x_11 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___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_1); +return x_11; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, 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___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof(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___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkNestedProofCongr___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, 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: { @@ -9296,7 +10293,7 @@ static lean_object* _init_l_Lean_Meta_Grind_mkEqProofImpl___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("( __do_lift._@.Lean.Meta.Tactic.Grind.Proof._hyg.2379.0 )\n ", 60, 60); +x_1 = lean_mk_string_unchecked("( __do_lift._@.Lean.Meta.Tactic.Grind.Proof._hyg.3014.0 )\n ", 60, 60); return x_1; } } @@ -9324,7 +10321,7 @@ static lean_object* _init_l_Lean_Meta_Grind_mkEqProofImpl___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_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__7___lambda__1___closed__1; x_2 = l_Lean_Meta_Grind_mkEqProofImpl___closed__3; -x_3 = lean_unsigned_to_nat(231u); +x_3 = lean_unsigned_to_nat(253u); x_4 = lean_unsigned_to_nat(2u); x_5 = l_Lean_Meta_Grind_mkEqProofImpl___closed__2; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -9415,14 +10412,14 @@ x_13 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore return x_13; } } -lean_object* initialize_Lean_Meta_Sorry(uint8_t builtin, lean_object*); +lean_object* initialize_Init_Grind_Lemmas(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_Tactic_Grind_Types(uint8_t builtin, lean_object*); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Lean_Meta_Tactic_Grind_Proof(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_Sorry(builtin, lean_io_mk_world()); +res = initialize_Init_Grind_Lemmas(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); res = initialize_Lean_Meta_Tactic_Grind_Types(builtin, lean_io_mk_world()); @@ -9476,8 +10473,6 @@ l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkProofFrom___lambda lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkProofFrom___lambda__1___closed__2); l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkProofFrom___lambda__1___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkProofFrom___lambda__1___closed__3(); lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkProofFrom___lambda__1___closed__3); -l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___lambda__1___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___lambda__1___closed__1(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___lambda__1___closed__1); l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__1(); lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__1); l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__2(); @@ -9494,16 +10489,22 @@ l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___close lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__7); l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__8 = _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__8(); lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__8); -l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__9 = _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__9(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__9); -l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__10 = _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__10(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__10); -l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__11 = _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__11(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__11); -l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__12 = _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__12(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__12); -l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__13 = _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__13(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__13); +l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___lambda__1___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___lambda__1___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___lambda__1___closed__1); +l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__1); +l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__2); +l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__3); +l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__4(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__4); +l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__5 = _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__5(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__5); +l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__6 = _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__6(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__6); +l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__7 = _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__7(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__7); l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof___closed__1(); lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof___closed__1); l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof___closed__2(); @@ -9512,6 +10513,30 @@ l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof_ lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof___closed__3); l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof___closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof___closed__4(); lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof___closed__4); +l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__1); +l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__2); +l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__3); +l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__4(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__4); +l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__5 = _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__5(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__5); +l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__6 = _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__6(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__6); +l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__7 = _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__7(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__7); +l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__8 = _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__8(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__8); +l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__2___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__2___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__2___closed__1); +l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__3___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__3___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__3___closed__1); +l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__4___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__4___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__4___closed__1); +l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___closed__1); l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkNestedProofCongr___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkNestedProofCongr___closed__1(); lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkNestedProofCongr___closed__1); l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkNestedProofCongr___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkNestedProofCongr___closed__2(); diff --git a/stage0/stdlib/Lean/Meta/Tactic/Grind/Propagate.c b/stage0/stdlib/Lean/Meta/Tactic/Grind/Propagate.c index 80e456e2dbda..7f380dd9a3ac 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Grind/Propagate.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Grind/Propagate.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Lean.Meta.Tactic.Grind.Propagate -// Imports: Init.Grind Lean.Meta.Tactic.Grind.Proof Lean.Meta.Tactic.Grind.PropagatorAttr +// Imports: Init.Grind Lean.Meta.Tactic.Grind.Proof Lean.Meta.Tactic.Grind.PropagatorAttr Lean.Meta.Tactic.Grind.Simp Lean.Meta.Tactic.Grind.Internalize #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -15,10 +15,12 @@ extern "C" { #endif static lean_object* l_Lean_Meta_Grind_propagateAndUp___lambda__1___closed__2; lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_propagateEqMatchDown___closed__2; static lean_object* l_Lean_Meta_Grind_propagateNotDown___lambda__1___closed__2; -static lean_object* l_Lean_Meta_Grind_propagateNotDown___lambda__1___closed__7; +lean_object* l_Lean_Meta_Simp_Result_getProof(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_propagateEqDown___lambda__1___closed__1; static lean_object* l_Lean_Meta_Grind_propagateNotDown___lambda__1___closed__3; +static lean_object* l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__4; static lean_object* l___regBuiltin_Lean_Meta_Grind_propagateAndUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_383____closed__1; static lean_object* l_Lean_Meta_Grind_propagateOrUp___lambda__1___closed__11; static lean_object* l_Lean_Meta_Grind_propagateOrDown___lambda__1___closed__4; @@ -26,133 +28,171 @@ static lean_object* l_Lean_Meta_Grind_propagateAndDown___lambda__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateAndUp___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_propagateNotDown___lambda__1___closed__1; static lean_object* l_Lean_Meta_Grind_propagateAndUp___lambda__1___closed__5; -static lean_object* l___regBuiltin_Lean_Meta_Grind_propagateNotUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_1426____closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Meta_Grind_propagateAndUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_383_(lean_object*); static lean_object* l_Lean_Meta_Grind_propagateOrUp___lambda__1___closed__10; static lean_object* l_Lean_Meta_Grind_propagateAndDown___lambda__1___closed__5; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateEqUp(lean_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_Meta_Grind_propagateNotUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_1486_(lean_object*); static lean_object* l_Lean_Meta_Grind_propagateEqDown___lambda__1___closed__2; -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Meta_Grind_propagateEqUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2058_(lean_object*); +static lean_object* l_Lean_Meta_Grind_propagateIte___closed__1; static lean_object* l_Lean_Meta_Grind_propagateNotUp___lambda__1___closed__5; static lean_object* l_Lean_Meta_Grind_propagateOrUp___lambda__1___closed__2; static lean_object* l_Lean_Meta_Grind_propagateNotUp___closed__2; +lean_object* l_Lean_mkApp8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_propagateNotDown___lambda__1___closed__6; lean_object* l_Lean_mkAppB(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateNotUp___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_Expr_isApp(lean_object*); -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Meta_Grind_propagateHEqUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2806_(lean_object*); static lean_object* l_Lean_Meta_Grind_propagateAndUp___lambda__1___closed__14; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateAndDown___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_propagateEqUp___lambda__1___closed__5; +static lean_object* l_Lean_Meta_Grind_propagateNotUp___lambda__1___closed__8; static lean_object* l___regBuiltin_Lean_Meta_Grind_propagateOrUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_985____closed__1; static lean_object* l_Lean_Meta_Grind_propagateAndUp___lambda__1___closed__3; static lean_object* l_Lean_Meta_Grind_propagateNotUp___lambda__1___closed__6; static lean_object* l_Lean_Meta_Grind_propagateEqUp___lambda__1___closed__8; +static lean_object* l___regBuiltin_Lean_Meta_Grind_propagateHEqDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2888____closed__1; static lean_object* l_Lean_Meta_Grind_propagateHEqDown___closed__1; lean_object* l_Lean_Meta_Grind_pushEqCore(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_propagateAndDown___lambda__1___closed__1; -static lean_object* l___regBuiltin_Lean_Meta_Grind_propagateHEqDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2546____closed__1; +static lean_object* l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__7; static lean_object* l___regBuiltin_Lean_Meta_Grind_propagateAndDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_604____closed__1; static lean_object* l_Lean_Meta_Grind_propagateAndUp___lambda__1___closed__7; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateAndUp(lean_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_Meta_Grind_propagateHEqDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2888_(lean_object*); static lean_object* l_Lean_Meta_Grind_propagateOrUp___lambda__1___closed__1; static lean_object* l_Lean_Meta_Grind_propagateAndUp___lambda__1___closed__1; lean_object* l_Lean_Expr_cleanupAnnotations(lean_object*); static lean_object* l_Lean_Meta_Grind_propagateAndDown___lambda__1___closed__6; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateEqMatchDown___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_Grind_propagateDIte___lambda__1___closed__1; static lean_object* l_Lean_Meta_Grind_propagateAndUp___closed__2; static lean_object* l_Lean_Meta_Grind_propagateNotUp___lambda__1___closed__1; static lean_object* l_Lean_Meta_Grind_propagateEqUp___lambda__1___closed__6; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateNotDown(lean_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_Grind_propagateEqMatchDown___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateDIte(lean_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_Grind_propagateHEqDown___closed__2; -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Meta_Grind_propagateNotDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_1706_(lean_object*); +lean_object* l_Lean_Meta_Grind_simp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_mkApp6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_propagateEqUp___lambda__1___closed__4; static lean_object* l_Lean_Meta_Grind_propagateOrUp___lambda__1___closed__7; static lean_object* l_Lean_Meta_Grind_propagateOrDown___lambda__1___closed__2; +lean_object* l_Lean_Meta_Grind_markCaseSplitAsResolved(lean_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_Grind_propagateEqUp___lambda__1___closed__7; +static lean_object* l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__3; static lean_object* l_Lean_Meta_Grind_propagateOrUp___lambda__1___closed__12; static lean_object* l_Lean_Meta_Grind_propagateAndUp___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateAndUp___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_Grind_propagateIte___lambda__1___closed__2; lean_object* l_Lean_Meta_Grind_closeGoal(lean_object*, 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_Grind_propagateNotDown___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_Grind_propagateAndDown___lambda__1___closed__4; static lean_object* l_Lean_Meta_Grind_propagateOrUp___lambda__1___closed__8; -static lean_object* l___regBuiltin_Lean_Meta_Grind_propagateHEqUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2806____closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateHEqUp___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_Grind_propagateEqUp___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_Expr_appArg(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_propagateNotUp___lambda__1___closed__9; +static lean_object* l___regBuiltin_Lean_Meta_Grind_propagateIte_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_3530____closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateEqDown___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_propagateOrUp___lambda__1___closed__6; -static lean_object* l___regBuiltin_Lean_Meta_Grind_propagateEqUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2058____closed__1; static lean_object* l_Lean_Meta_Grind_propagateNotUp___lambda__1___closed__2; -static lean_object* l_Lean_Meta_Grind_propagateNotDown___lambda__1___closed__8; static lean_object* l_Lean_Meta_Grind_propagateEqDown___lambda__1___closed__3; lean_object* l_Lean_Expr_appFnCleanup(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Meta_Grind_propagateOrDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_1206_(lean_object*); -static lean_object* l_Lean_Meta_Grind_propagateNotDown___lambda__1___closed__9; +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Meta_Grind_propagateEqUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2118_(lean_object*); static lean_object* l_Lean_Meta_Grind_propagateEqUp___lambda__1___closed__1; static lean_object* l___regBuiltin_Lean_Meta_Grind_propagateOrDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_1206____closed__1; +static lean_object* l___regBuiltin_Lean_Meta_Grind_propagateEqDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2347____closed__1; static lean_object* l_Lean_Meta_Grind_propagateAndUp___lambda__1___closed__11; static lean_object* l_Lean_Meta_Grind_propagateNotDown___lambda__1___closed__4; lean_object* l_Lean_Meta_Grind_pushEqTrue(lean_object*, lean_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_Grind_propagateAndUp___lambda__1___closed__4; +static lean_object* l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__5; static lean_object* l_Lean_Meta_Grind_propagateOrDown___lambda__1___closed__1; static lean_object* l_Lean_Meta_Grind_propagateOrUp___lambda__1___closed__5; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateDIte___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_propagateAndUp___lambda__1___closed__8; lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); +lean_object* l_Lean_Meta_Grind_getGeneration(lean_object*, 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_Grind_propagateEqDown(lean_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_Grind_propagateDIte___lambda__1___closed__2; +static lean_object* l___regBuiltin_Lean_Meta_Grind_propagateHEqUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_3148____closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateIte___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_propagateNotUp___lambda__1___closed__7; static lean_object* l_Lean_Meta_Grind_propagateEqUp___lambda__1___closed__3; +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Meta_Grind_propagateEqMatchDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2629_(lean_object*); static lean_object* l_Lean_Meta_Grind_propagateAndUp___lambda__1___closed__10; -static lean_object* l___regBuiltin_Lean_Meta_Grind_propagateEqDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2287____closed__1; static lean_object* l_Lean_Meta_Grind_propagateNotDown___lambda__1___closed__5; +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Meta_Grind_propagateHEqUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_3148_(lean_object*); +static lean_object* l_Lean_Meta_Grind_propagateIte___lambda__1___closed__1; +static lean_object* l_Lean_Meta_Grind_propagateDIte___closed__2; static lean_object* l_Lean_Meta_Grind_propagateAndUp___lambda__1___closed__9; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateNotUp(lean_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_Grind_propagateOrUp___lambda__1___closed__3; static lean_object* l_Lean_Meta_Grind_propagateAndUp___lambda__1___closed__13; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Meta_Grind_propagateOrUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_985_(lean_object*); -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Meta_Grind_propagateEqDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2287_(lean_object*); +static lean_object* l_Lean_Meta_Grind_propagateIte___closed__2; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Meta_Grind_propagateAndDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_604_(lean_object*); static lean_object* l_Lean_Meta_Grind_propagateOrUp___lambda__1___closed__4; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateOrDown___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_propagateNotUp___lambda__1___closed__4; +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Meta_Grind_propagateIte_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_3530_(lean_object*); lean_object* lean_grind_mk_eq_proof(lean_object*, lean_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_Grind_propagateEqUp___closed__1; lean_object* lean_grind_mk_heq_proof(lean_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_constLevels_x21(lean_object*); +static lean_object* l_Lean_Meta_Grind_propagateDIte___closed__1; +static lean_object* l___regBuiltin_Lean_Meta_Grind_propagateNotUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_1486____closed__1; +static lean_object* l___regBuiltin_Lean_Meta_Grind_propagateEqUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2118____closed__1; +static lean_object* l_Lean_Meta_Grind_propagateIte___lambda__1___closed__4; +lean_object* l_Lean_Expr_app___override(lean_object*, lean_object*); lean_object* l_Lean_mkApp3(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_mkEqFalseProof(lean_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_Grind_propagateOrUp___closed__1; static lean_object* l_Lean_Meta_Grind_propagateAndUp___lambda__1___closed__6; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateDIte___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateOrDown(lean_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_Grind_propagateOrDown___lambda__1___closed__6; static lean_object* l_Lean_Meta_Grind_propagateOrUp___lambda__1___closed__9; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateIte___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateAndDown(lean_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_Grind_propagateEqUp___lambda__1___closed__2; uint8_t l_Lean_Expr_isConstOf(lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_isEqv(lean_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_Tactic_Grind_PropagatorAttr_0__Lean_Meta_Grind_registerBuiltinPropagatorCore(lean_object*, uint8_t, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Meta_Grind_propagateEqDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2347_(lean_object*); lean_object* l_Lean_Meta_Grind_mkEqTrueProof(lean_object*, 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_Grind_propagateAndUp___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___regBuiltin_Lean_Meta_Grind_propagateNotUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_1426_(lean_object*); +static lean_object* l___regBuiltin_Lean_Meta_Grind_propagateEqMatchDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2629____closed__1; +static lean_object* l___regBuiltin_Lean_Meta_Grind_propagateDIte_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_4078____closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateIte(lean_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_Grind_propagateAndDown___lambda__1___closed__3; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateEqMatchDown(lean_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_Grind_propagateOrDown___lambda__1___closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateOrUp(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_Grind_isEqFalse(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_Grind_pushEqFalse(lean_object*, lean_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_Grind_propagateOrDown___lambda__1___closed__5; +static lean_object* l___regBuiltin_Lean_Meta_Grind_propagateNotDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_1766____closed__1; +static lean_object* l_Lean_Meta_Grind_propagateIte___lambda__1___closed__3; lean_object* l_Lean_Meta_Grind_isEqTrue(lean_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_Grind_propagateAndUp___lambda__1___closed__12; static lean_object* l_Lean_Meta_Grind_propagateNotUp___lambda__1___closed__3; -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Meta_Grind_propagateHEqDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2546_(lean_object*); static lean_object* l_Lean_Meta_Grind_propagateOrUp___closed__2; -static lean_object* l___regBuiltin_Lean_Meta_Grind_propagateNotDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_1706____closed__1; static lean_object* l_Lean_Meta_Grind_propagateEqUp___lambda__1___closed__9; static lean_object* l_Lean_Meta_Grind_propagateNotUp___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateOrUp___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_Grind_propagateHEqDown___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_propagateAndUp___closed__3; +static lean_object* l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__6; static lean_object* l_Lean_Meta_Grind_propagateEqUp___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateHEqUp(lean_object*, 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_Grind_propagateHEqDown(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_Grind_internalize(lean_object*, lean_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_Meta_Grind_propagateNotDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_1766_(lean_object*); +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Meta_Grind_propagateDIte_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_4078_(lean_object*); static lean_object* _init_l_Lean_Meta_Grind_propagateAndUp___lambda__1___closed__1() { _start: { @@ -2554,7 +2594,7 @@ static lean_object* _init_l_Lean_Meta_Grind_propagateNotUp___lambda__1___closed_ _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("not_eq_of_eq_true", 17, 17); +x_1 = lean_mk_string_unchecked("false_of_not_eq_self", 20, 20); return x_1; } } @@ -2583,7 +2623,7 @@ static lean_object* _init_l_Lean_Meta_Grind_propagateNotUp___lambda__1___closed_ _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("not_eq_of_eq_false", 18, 18); +x_1 = lean_mk_string_unchecked("not_eq_of_eq_true", 17, 17); return x_1; } } @@ -2608,6 +2648,35 @@ x_3 = l_Lean_Expr_const___override(x_2, x_1); return x_3; } } +static lean_object* _init_l_Lean_Meta_Grind_propagateNotUp___lambda__1___closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("not_eq_of_eq_false", 18, 18); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateNotUp___lambda__1___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Meta_Grind_propagateAndUp___lambda__1___closed__1; +x_2 = l_Lean_Meta_Grind_propagateAndUp___lambda__1___closed__2; +x_3 = l_Lean_Meta_Grind_propagateNotUp___lambda__1___closed__7; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateNotUp___lambda__1___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind_propagateNotUp___lambda__1___closed__8; +x_3 = l_Lean_Expr_const___override(x_2, x_1); +return x_3; +} +} LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateNotUp___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: { @@ -2638,7 +2707,23 @@ x_18 = lean_unbox(x_17); lean_dec(x_17); if (x_18 == 0) { -uint8_t 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); +lean_inc(x_2); +lean_inc(x_1); +x_20 = l_Lean_Meta_Grind_isEqv(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_19); +if (lean_obj_tag(x_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) +{ +uint8_t x_23; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -2649,35 +2734,35 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_19 = !lean_is_exclusive(x_16); -if (x_19 == 0) +x_23 = !lean_is_exclusive(x_20); +if (x_23 == 0) { -lean_object* x_20; lean_object* x_21; -x_20 = lean_ctor_get(x_16, 0); -lean_dec(x_20); -x_21 = lean_box(0); -lean_ctor_set(x_16, 0, x_21); -return x_16; +lean_object* x_24; lean_object* x_25; +x_24 = lean_ctor_get(x_20, 0); +lean_dec(x_24); +x_25 = lean_box(0); +lean_ctor_set(x_20, 0, x_25); +return x_20; } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_16, 1); -lean_inc(x_22); -lean_dec(x_16); -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_26; lean_object* x_27; lean_object* x_28; +x_26 = lean_ctor_get(x_20, 1); +lean_inc(x_26); +lean_dec(x_20); +x_27 = lean_box(0); +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_27); +lean_ctor_set(x_28, 1, x_26); +return x_28; } } else { -lean_object* x_25; lean_object* x_26; -x_25 = lean_ctor_get(x_16, 1); -lean_inc(x_25); -lean_dec(x_16); +lean_object* x_29; lean_object* x_30; +x_29 = lean_ctor_get(x_20, 1); +lean_inc(x_29); +lean_dec(x_20); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); @@ -2687,31 +2772,23 @@ lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); -x_26 = l_Lean_Meta_Grind_mkEqTrueProof(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_25); -if (lean_obj_tag(x_26) == 0) +x_30 = lean_grind_mk_eq_proof(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_29); +if (lean_obj_tag(x_30) == 0) { -lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_27 = lean_ctor_get(x_26, 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_Grind_propagateNotUp___lambda__1___closed__3; -x_30 = l_Lean_mkAppB(x_29, x_2, x_27); -x_31 = l_Lean_Meta_Grind_pushEqFalse(x_1, x_30, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, 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_3); -return x_31; +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_inc(x_32); +lean_dec(x_30); +x_33 = l_Lean_Meta_Grind_propagateNotUp___lambda__1___closed__3; +x_34 = l_Lean_mkAppB(x_33, x_2, x_31); +x_35 = l_Lean_Meta_Grind_closeGoal(x_34, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_32); +return x_35; } else { -uint8_t x_32; +uint8_t x_36; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -2721,31 +2798,30 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -x_32 = !lean_is_exclusive(x_26); -if (x_32 == 0) +x_36 = !lean_is_exclusive(x_30); +if (x_36 == 0) { -return x_26; +return x_30; } else { -lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_33 = lean_ctor_get(x_26, 0); -x_34 = lean_ctor_get(x_26, 1); -lean_inc(x_34); -lean_inc(x_33); -lean_dec(x_26); -x_35 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -return x_35; +lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_37 = lean_ctor_get(x_30, 0); +x_38 = lean_ctor_get(x_30, 1); +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; } } } } else { -uint8_t x_36; +uint8_t x_40; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -2756,32 +2832,32 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_36 = !lean_is_exclusive(x_16); -if (x_36 == 0) +x_40 = !lean_is_exclusive(x_20); +if (x_40 == 0) { -return x_16; +return x_20; } else { -lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_37 = lean_ctor_get(x_16, 0); -x_38 = lean_ctor_get(x_16, 1); -lean_inc(x_38); -lean_inc(x_37); -lean_dec(x_16); -x_39 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_39, 0, x_37); -lean_ctor_set(x_39, 1, x_38); -return x_39; +lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_41 = lean_ctor_get(x_20, 0); +x_42 = lean_ctor_get(x_20, 1); +lean_inc(x_42); +lean_inc(x_41); +lean_dec(x_20); +x_43 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_43, 0, x_41); +lean_ctor_set(x_43, 1, x_42); +return x_43; } } } else { -lean_object* x_40; lean_object* x_41; -x_40 = lean_ctor_get(x_12, 1); -lean_inc(x_40); -lean_dec(x_12); +lean_object* x_44; lean_object* x_45; +x_44 = lean_ctor_get(x_16, 1); +lean_inc(x_44); +lean_dec(x_16); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); @@ -2791,31 +2867,18 @@ lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); -x_41 = l_Lean_Meta_Grind_mkEqFalseProof(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, 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; -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_Grind_propagateNotUp___lambda__1___closed__6; -x_45 = l_Lean_mkAppB(x_44, x_2, x_42); -x_46 = l_Lean_Meta_Grind_pushEqTrue(x_1, x_45, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_43); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -return x_46; -} -else +x_45 = l_Lean_Meta_Grind_mkEqTrueProof(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_44); +if (lean_obj_tag(x_45) == 0) { -uint8_t x_47; +lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_46 = lean_ctor_get(x_45, 0); +lean_inc(x_46); +x_47 = lean_ctor_get(x_45, 1); +lean_inc(x_47); +lean_dec(x_45); +x_48 = l_Lean_Meta_Grind_propagateNotUp___lambda__1___closed__6; +x_49 = l_Lean_mkAppB(x_48, x_2, x_46); +x_50 = l_Lean_Meta_Grind_pushEqFalse(x_1, x_49, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_47); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -2824,29 +2887,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_47 = !lean_is_exclusive(x_41); -if (x_47 == 0) -{ -return x_41; -} -else -{ -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_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_51; @@ -2860,19 +2902,19 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_51 = !lean_is_exclusive(x_12); +x_51 = !lean_is_exclusive(x_45); if (x_51 == 0) { -return x_12; +return x_45; } else { lean_object* x_52; lean_object* x_53; lean_object* x_54; -x_52 = lean_ctor_get(x_12, 0); -x_53 = lean_ctor_get(x_12, 1); +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_12); +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); @@ -2881,35 +2923,10 @@ return x_54; } } } -static lean_object* _init_l_Lean_Meta_Grind_propagateNotUp___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Not", 3, 3); -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_Grind_propagateNotUp___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_Grind_propagateNotUp___closed__1; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateNotUp(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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_inc(x_1); -x_11 = l_Lean_Expr_cleanupAnnotations(x_1); -x_12 = l_Lean_Expr_isApp(x_11); -if (x_12 == 0) +else { -lean_object* x_13; lean_object* x_14; -lean_dec(x_11); +uint8_t x_55; +lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -2919,25 +2936,188 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_13 = lean_box(0); -x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_13); -lean_ctor_set(x_14, 1, x_10); -return x_14; +x_55 = !lean_is_exclusive(x_16); +if (x_55 == 0) +{ +return x_16; } else { -lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; -x_15 = l_Lean_Expr_appArg(x_11, lean_box(0)); -x_16 = l_Lean_Expr_appFnCleanup(x_11, lean_box(0)); -x_17 = l_Lean_Meta_Grind_propagateNotUp___closed__2; -x_18 = l_Lean_Expr_isConstOf(x_16, x_17); +lean_object* x_56; lean_object* x_57; lean_object* x_58; +x_56 = lean_ctor_get(x_16, 0); +x_57 = lean_ctor_get(x_16, 1); +lean_inc(x_57); +lean_inc(x_56); lean_dec(x_16); -if (x_18 == 0) -{ -lean_object* x_19; lean_object* x_20; -lean_dec(x_15); -lean_dec(x_9); +x_58 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_58, 0, x_56); +lean_ctor_set(x_58, 1, x_57); +return x_58; +} +} +} +else +{ +lean_object* x_59; lean_object* x_60; +x_59 = lean_ctor_get(x_12, 1); +lean_inc(x_59); +lean_dec(x_12); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +x_60 = l_Lean_Meta_Grind_mkEqFalseProof(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_59); +if (lean_obj_tag(x_60) == 0) +{ +lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; +x_61 = lean_ctor_get(x_60, 0); +lean_inc(x_61); +x_62 = lean_ctor_get(x_60, 1); +lean_inc(x_62); +lean_dec(x_60); +x_63 = l_Lean_Meta_Grind_propagateNotUp___lambda__1___closed__9; +x_64 = l_Lean_mkAppB(x_63, x_2, x_61); +x_65 = l_Lean_Meta_Grind_pushEqTrue(x_1, x_64, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_62); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_65; +} +else +{ +uint8_t x_66; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_66 = !lean_is_exclusive(x_60); +if (x_66 == 0) +{ +return x_60; +} +else +{ +lean_object* x_67; lean_object* x_68; lean_object* x_69; +x_67 = lean_ctor_get(x_60, 0); +x_68 = lean_ctor_get(x_60, 1); +lean_inc(x_68); +lean_inc(x_67); +lean_dec(x_60); +x_69 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_69, 0, x_67); +lean_ctor_set(x_69, 1, x_68); +return x_69; +} +} +} +} +else +{ +uint8_t x_70; +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_70 = !lean_is_exclusive(x_12); +if (x_70 == 0) +{ +return x_12; +} +else +{ +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_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_Meta_Grind_propagateNotUp___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Not", 3, 3); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateNotUp___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_Grind_propagateNotUp___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateNotUp(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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_inc(x_1); +x_11 = l_Lean_Expr_cleanupAnnotations(x_1); +x_12 = l_Lean_Expr_isApp(x_11); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; +lean_dec(x_11); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_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_10); +return x_14; +} +else +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_15 = l_Lean_Expr_appArg(x_11, lean_box(0)); +x_16 = l_Lean_Expr_appFnCleanup(x_11, lean_box(0)); +x_17 = l_Lean_Meta_Grind_propagateNotUp___closed__2; +x_18 = l_Lean_Expr_isConstOf(x_16, x_17); +lean_dec(x_16); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; +lean_dec(x_15); +lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -2961,7 +3141,7 @@ return x_21; } } } -static lean_object* _init_l___regBuiltin_Lean_Meta_Grind_propagateNotUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_1426____closed__1() { +static lean_object* _init_l___regBuiltin_Lean_Meta_Grind_propagateNotUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_1486____closed__1() { _start: { lean_object* x_1; @@ -2969,13 +3149,13 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_propagateNotUp), 10, 0); return x_1; } } -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Meta_Grind_propagateNotUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_1426_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Meta_Grind_propagateNotUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_1486_(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_Grind_propagateNotUp___closed__2; x_3 = 1; -x_4 = l___regBuiltin_Lean_Meta_Grind_propagateNotUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_1426____closed__1; +x_4 = l___regBuiltin_Lean_Meta_Grind_propagateNotUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_1486____closed__1; x_5 = l___private_Lean_Meta_Tactic_Grind_PropagatorAttr_0__Lean_Meta_Grind_registerBuiltinPropagatorCore(x_2, x_3, x_4, x_1); return x_5; } @@ -2984,7 +3164,7 @@ static lean_object* _init_l_Lean_Meta_Grind_propagateNotDown___lambda__1___close _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("false_of_not_eq_self", 20, 20); +x_1 = lean_mk_string_unchecked("eq_false_of_not_eq_true", 23, 23); return x_1; } } @@ -3013,7 +3193,7 @@ static lean_object* _init_l_Lean_Meta_Grind_propagateNotDown___lambda__1___close _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("eq_false_of_not_eq_true", 23, 23); +x_1 = lean_mk_string_unchecked("eq_true_of_not_eq_false", 23, 23); return x_1; } } @@ -3038,35 +3218,6 @@ x_3 = l_Lean_Expr_const___override(x_2, x_1); return x_3; } } -static lean_object* _init_l_Lean_Meta_Grind_propagateNotDown___lambda__1___closed__7() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("eq_true_of_not_eq_false", 23, 23); -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_Grind_propagateNotDown___lambda__1___closed__8() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Meta_Grind_propagateAndUp___lambda__1___closed__1; -x_2 = l_Lean_Meta_Grind_propagateAndUp___lambda__1___closed__2; -x_3 = l_Lean_Meta_Grind_propagateNotDown___lambda__1___closed__7; -x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); -return x_4; -} -} -static lean_object* _init_l_Lean_Meta_Grind_propagateNotDown___lambda__1___closed__9() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Meta_Grind_propagateNotDown___lambda__1___closed__8; -x_3 = l_Lean_Expr_const___override(x_2, x_1); -return x_3; -} -} LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateNotDown___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: { @@ -3171,13 +3322,9 @@ 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_Grind_propagateNotDown___lambda__1___closed__3; +x_33 = l_Lean_Meta_Grind_propagateNotUp___lambda__1___closed__3; x_34 = l_Lean_mkAppB(x_33, x_2, x_31); x_35 = l_Lean_Meta_Grind_closeGoal(x_34, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_32); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); return x_35; } else @@ -3269,7 +3416,7 @@ lean_inc(x_46); x_47 = lean_ctor_get(x_45, 1); lean_inc(x_47); lean_dec(x_45); -x_48 = l_Lean_Meta_Grind_propagateNotDown___lambda__1___closed__6; +x_48 = l_Lean_Meta_Grind_propagateNotDown___lambda__1___closed__3; lean_inc(x_2); x_49 = l_Lean_mkAppB(x_48, x_2, x_46); x_50 = l_Lean_Meta_Grind_pushEqFalse(x_2, x_49, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_47); @@ -3372,7 +3519,7 @@ lean_inc(x_61); x_62 = lean_ctor_get(x_60, 1); lean_inc(x_62); lean_dec(x_60); -x_63 = l_Lean_Meta_Grind_propagateNotDown___lambda__1___closed__9; +x_63 = l_Lean_Meta_Grind_propagateNotDown___lambda__1___closed__6; lean_inc(x_2); x_64 = l_Lean_mkAppB(x_63, x_2, x_61); x_65 = l_Lean_Meta_Grind_pushEqTrue(x_2, x_64, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_62); @@ -3515,7 +3662,7 @@ return x_21; } } } -static lean_object* _init_l___regBuiltin_Lean_Meta_Grind_propagateNotDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_1706____closed__1() { +static lean_object* _init_l___regBuiltin_Lean_Meta_Grind_propagateNotDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_1766____closed__1() { _start: { lean_object* x_1; @@ -3523,13 +3670,13 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_propagateNotDown), 10, 0); return x_1; } } -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Meta_Grind_propagateNotDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_1706_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Meta_Grind_propagateNotDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_1766_(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_Grind_propagateNotUp___closed__2; x_3 = 0; -x_4 = l___regBuiltin_Lean_Meta_Grind_propagateNotDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_1706____closed__1; +x_4 = l___regBuiltin_Lean_Meta_Grind_propagateNotDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_1766____closed__1; x_5 = l___private_Lean_Meta_Tactic_Grind_PropagatorAttr_0__Lean_Meta_Grind_registerBuiltinPropagatorCore(x_2, x_3, x_4, x_1); return x_5; } @@ -4157,7 +4304,7 @@ return x_30; } } } -static lean_object* _init_l___regBuiltin_Lean_Meta_Grind_propagateEqUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2058____closed__1() { +static lean_object* _init_l___regBuiltin_Lean_Meta_Grind_propagateEqUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2118____closed__1() { _start: { lean_object* x_1; @@ -4165,13 +4312,13 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_propagateEqUp), 10, 0); return x_1; } } -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Meta_Grind_propagateEqUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2058_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Meta_Grind_propagateEqUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2118_(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_Grind_propagateEqUp___closed__2; x_3 = 1; -x_4 = l___regBuiltin_Lean_Meta_Grind_propagateEqUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2058____closed__1; +x_4 = l___regBuiltin_Lean_Meta_Grind_propagateEqUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2118____closed__1; x_5 = l___private_Lean_Meta_Tactic_Grind_PropagatorAttr_0__Lean_Meta_Grind_registerBuiltinPropagatorCore(x_2, x_3, x_4, x_1); return x_5; } @@ -4591,7 +4738,7 @@ return x_63; } } } -static lean_object* _init_l___regBuiltin_Lean_Meta_Grind_propagateEqDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2287____closed__1() { +static lean_object* _init_l___regBuiltin_Lean_Meta_Grind_propagateEqDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2347____closed__1() { _start: { lean_object* x_1; @@ -4599,21 +4746,29 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_propagateEqDown), 10, 0); return x_1; } } -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Meta_Grind_propagateEqDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2287_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Meta_Grind_propagateEqDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2347_(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_Grind_propagateEqUp___closed__2; x_3 = 0; -x_4 = l___regBuiltin_Lean_Meta_Grind_propagateEqDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2287____closed__1; +x_4 = l___regBuiltin_Lean_Meta_Grind_propagateEqDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2347____closed__1; x_5 = l___private_Lean_Meta_Tactic_Grind_PropagatorAttr_0__Lean_Meta_Grind_registerBuiltinPropagatorCore(x_2, x_3, x_4, x_1); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateHEqDown___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_Meta_Grind_propagateEqMatchDown___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_13; +lean_object* x_14; +x_14 = l_Lean_Meta_Grind_markCaseSplitAsResolved(x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; lean_object* x_16; +x_15 = lean_ctor_get(x_14, 1); +lean_inc(x_15); +lean_dec(x_14); +lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); @@ -4621,21 +4776,21 @@ lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -lean_inc(x_4); lean_inc(x_1); -x_13 = l_Lean_Meta_Grind_mkEqTrueProof(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_13) == 0) +x_16 = l_Lean_Meta_Grind_mkEqTrueProof(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_15); +if (lean_obj_tag(x_16) == 0) { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; lean_object* x_19; -x_14 = lean_ctor_get(x_13, 0); -lean_inc(x_14); -x_15 = lean_ctor_get(x_13, 1); -lean_inc(x_15); -lean_dec(x_13); -x_16 = l_Lean_Meta_Grind_propagateEqDown___lambda__1___closed__3; -x_17 = l_Lean_mkAppB(x_16, x_1, x_14); -x_18 = 1; -x_19 = l_Lean_Meta_Grind_pushEqCore(x_2, x_3, x_17, x_18, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_15); +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; 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); +lean_dec(x_16); +x_19 = l_Lean_Meta_Grind_propagateEqDown___lambda__1___closed__3; +x_20 = l_Lean_mkAppB(x_19, x_1, x_17); +x_21 = 0; +x_22 = l_Lean_Meta_Grind_pushEqCore(x_2, x_3, x_20, x_21, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_18); +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -4643,12 +4798,12 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); -return x_19; +return x_22; } else { -uint8_t x_20; +uint8_t x_23; +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -4656,50 +4811,84 @@ 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_20 = !lean_is_exclusive(x_13); -if (x_20 == 0) +x_23 = !lean_is_exclusive(x_16); +if (x_23 == 0) { -return x_13; +return x_16; } else { -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_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_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_16, 0); +x_25 = lean_ctor_get(x_16, 1); +lean_inc(x_25); +lean_inc(x_24); +lean_dec(x_16); +x_26 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_26, 0, x_24); +lean_ctor_set(x_26, 1, x_25); +return x_26; } } } -static lean_object* _init_l_Lean_Meta_Grind_propagateHEqDown___closed__1() { +else +{ +uint8_t x_27; +lean_dec(x_12); +lean_dec(x_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_27 = !lean_is_exclusive(x_14); +if (x_27 == 0) +{ +return x_14; +} +else +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_28 = lean_ctor_get(x_14, 0); +x_29 = lean_ctor_get(x_14, 1); +lean_inc(x_29); +lean_inc(x_28); +lean_dec(x_14); +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_Grind_propagateEqMatchDown___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("HEq", 3, 3); +x_1 = lean_mk_string_unchecked("EqMatch", 7, 7); return x_1; } } -static lean_object* _init_l_Lean_Meta_Grind_propagateHEqDown___closed__2() { +static lean_object* _init_l_Lean_Meta_Grind_propagateEqMatchDown___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_Grind_propagateHEqDown___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; +x_1 = l_Lean_Meta_Grind_propagateAndUp___lambda__1___closed__1; +x_2 = l_Lean_Meta_Grind_propagateAndUp___lambda__1___closed__2; +x_3 = l_Lean_Meta_Grind_propagateEqMatchDown___closed__1; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateHEqDown(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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_Grind_propagateEqMatchDown(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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; @@ -4803,12 +4992,14 @@ return x_11; } else { -lean_object* x_30; uint8_t x_31; -x_30 = l_Lean_Expr_appFnCleanup(x_27, lean_box(0)); -x_31 = l_Lean_Expr_isApp(x_30); -if (x_31 == 0) +lean_object* x_30; lean_object* x_31; uint8_t x_32; +x_30 = l_Lean_Expr_appArg(x_27, lean_box(0)); +x_31 = l_Lean_Expr_appFnCleanup(x_27, lean_box(0)); +x_32 = l_Lean_Expr_isApp(x_31); +if (x_32 == 0) { -lean_object* x_32; +lean_object* x_33; +lean_dec(x_31); lean_dec(x_30); lean_dec(x_26); lean_dec(x_9); @@ -4820,21 +5011,22 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_32 = lean_box(0); -lean_ctor_set(x_11, 0, x_32); +x_33 = lean_box(0); +lean_ctor_set(x_11, 0, x_33); return x_11; } else { -lean_object* x_33; lean_object* x_34; uint8_t x_35; -x_33 = l_Lean_Expr_appArg(x_30, lean_box(0)); -x_34 = l_Lean_Expr_appFnCleanup(x_30, lean_box(0)); -x_35 = l_Lean_Expr_isApp(x_34); -if (x_35 == 0) +lean_object* x_34; lean_object* x_35; uint8_t x_36; +x_34 = l_Lean_Expr_appArg(x_31, lean_box(0)); +x_35 = l_Lean_Expr_appFnCleanup(x_31, lean_box(0)); +x_36 = l_Lean_Expr_isApp(x_35); +if (x_36 == 0) { -lean_object* x_36; +lean_object* x_37; +lean_dec(x_35); lean_dec(x_34); -lean_dec(x_33); +lean_dec(x_30); lean_dec(x_26); lean_dec(x_9); lean_dec(x_8); @@ -4845,21 +5037,22 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_36 = lean_box(0); -lean_ctor_set(x_11, 0, x_36); +x_37 = lean_box(0); +lean_ctor_set(x_11, 0, x_37); return x_11; } else { -lean_object* x_37; lean_object* x_38; uint8_t x_39; -x_37 = l_Lean_Expr_appFnCleanup(x_34, lean_box(0)); -x_38 = l_Lean_Meta_Grind_propagateHEqDown___closed__2; -x_39 = l_Lean_Expr_isConstOf(x_37, x_38); -lean_dec(x_37); -if (x_39 == 0) +lean_object* x_38; lean_object* x_39; uint8_t x_40; +x_38 = l_Lean_Expr_appFnCleanup(x_35, lean_box(0)); +x_39 = l_Lean_Meta_Grind_propagateEqMatchDown___closed__2; +x_40 = l_Lean_Expr_isConstOf(x_38, x_39); +lean_dec(x_38); +if (x_40 == 0) { -lean_object* x_40; -lean_dec(x_33); +lean_object* x_41; +lean_dec(x_34); +lean_dec(x_30); lean_dec(x_26); lean_dec(x_9); lean_dec(x_8); @@ -4870,16 +5063,16 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_40 = lean_box(0); -lean_ctor_set(x_11, 0, x_40); +x_41 = lean_box(0); +lean_ctor_set(x_11, 0, x_41); return x_11; } else { -lean_object* x_41; +lean_object* x_42; lean_free_object(x_11); -x_41 = l_Lean_Meta_Grind_propagateHEqDown___lambda__1(x_1, x_33, x_26, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_21); -return x_41; +x_42 = l_Lean_Meta_Grind_propagateEqMatchDown___lambda__1(x_1, x_34, x_30, x_26, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_21); +return x_42; } } } @@ -4888,17 +5081,17 @@ return x_41; } else { -lean_object* x_42; lean_object* x_43; uint8_t x_44; -x_42 = lean_ctor_get(x_11, 1); -lean_inc(x_42); +lean_object* x_43; lean_object* x_44; uint8_t x_45; +x_43 = lean_ctor_get(x_11, 1); +lean_inc(x_43); lean_dec(x_11); lean_inc(x_1); -x_43 = l_Lean_Expr_cleanupAnnotations(x_1); -x_44 = l_Lean_Expr_isApp(x_43); -if (x_44 == 0) +x_44 = l_Lean_Expr_cleanupAnnotations(x_1); +x_45 = l_Lean_Expr_isApp(x_44); +if (x_45 == 0) { -lean_object* x_45; lean_object* x_46; -lean_dec(x_43); +lean_object* x_46; lean_object* x_47; +lean_dec(x_44); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -4908,23 +5101,23 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -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_42); -return x_46; +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_43); +return x_47; } else { -lean_object* x_47; lean_object* x_48; uint8_t x_49; -x_47 = l_Lean_Expr_appArg(x_43, lean_box(0)); -x_48 = l_Lean_Expr_appFnCleanup(x_43, lean_box(0)); -x_49 = l_Lean_Expr_isApp(x_48); -if (x_49 == 0) +lean_object* x_48; lean_object* x_49; uint8_t x_50; +x_48 = l_Lean_Expr_appArg(x_44, lean_box(0)); +x_49 = l_Lean_Expr_appFnCleanup(x_44, lean_box(0)); +x_50 = l_Lean_Expr_isApp(x_49); +if (x_50 == 0) { -lean_object* x_50; lean_object* x_51; +lean_object* x_51; lean_object* x_52; +lean_dec(x_49); lean_dec(x_48); -lean_dec(x_47); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -4934,22 +5127,24 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -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_42); -return x_51; +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_43); +return x_52; } else { -lean_object* x_52; uint8_t x_53; -x_52 = l_Lean_Expr_appFnCleanup(x_48, lean_box(0)); -x_53 = l_Lean_Expr_isApp(x_52); -if (x_53 == 0) +lean_object* x_53; lean_object* x_54; uint8_t x_55; +x_53 = l_Lean_Expr_appArg(x_49, lean_box(0)); +x_54 = l_Lean_Expr_appFnCleanup(x_49, lean_box(0)); +x_55 = l_Lean_Expr_isApp(x_54); +if (x_55 == 0) { -lean_object* x_54; lean_object* x_55; -lean_dec(x_52); -lean_dec(x_47); +lean_object* x_56; lean_object* x_57; +lean_dec(x_54); +lean_dec(x_53); +lean_dec(x_48); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -4959,24 +5154,25 @@ lean_dec(x_4); 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_42); -return x_55; +x_56 = lean_box(0); +x_57 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_57, 0, x_56); +lean_ctor_set(x_57, 1, x_43); +return x_57; } else { -lean_object* x_56; lean_object* x_57; uint8_t x_58; -x_56 = l_Lean_Expr_appArg(x_52, lean_box(0)); -x_57 = l_Lean_Expr_appFnCleanup(x_52, lean_box(0)); -x_58 = l_Lean_Expr_isApp(x_57); -if (x_58 == 0) +lean_object* x_58; lean_object* x_59; uint8_t x_60; +x_58 = l_Lean_Expr_appArg(x_54, lean_box(0)); +x_59 = l_Lean_Expr_appFnCleanup(x_54, lean_box(0)); +x_60 = l_Lean_Expr_isApp(x_59); +if (x_60 == 0) { -lean_object* x_59; lean_object* x_60; -lean_dec(x_57); -lean_dec(x_56); -lean_dec(x_47); +lean_object* x_61; lean_object* x_62; +lean_dec(x_59); +lean_dec(x_58); +lean_dec(x_53); +lean_dec(x_48); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -4986,24 +5182,25 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_59 = lean_box(0); -x_60 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_60, 0, x_59); -lean_ctor_set(x_60, 1, x_42); -return x_60; +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_43); +return x_62; } else { -lean_object* x_61; lean_object* x_62; uint8_t x_63; -x_61 = l_Lean_Expr_appFnCleanup(x_57, lean_box(0)); -x_62 = l_Lean_Meta_Grind_propagateHEqDown___closed__2; -x_63 = l_Lean_Expr_isConstOf(x_61, x_62); -lean_dec(x_61); -if (x_63 == 0) +lean_object* x_63; lean_object* x_64; uint8_t x_65; +x_63 = l_Lean_Expr_appFnCleanup(x_59, lean_box(0)); +x_64 = l_Lean_Meta_Grind_propagateEqMatchDown___closed__2; +x_65 = l_Lean_Expr_isConstOf(x_63, x_64); +lean_dec(x_63); +if (x_65 == 0) { -lean_object* x_64; lean_object* x_65; -lean_dec(x_56); -lean_dec(x_47); +lean_object* x_66; lean_object* x_67; +lean_dec(x_58); +lean_dec(x_53); +lean_dec(x_48); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -5013,17 +5210,17 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -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_42); -return x_65; +x_66 = lean_box(0); +x_67 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_67, 0, x_66); +lean_ctor_set(x_67, 1, x_43); +return x_67; } else { -lean_object* x_66; -x_66 = l_Lean_Meta_Grind_propagateHEqDown___lambda__1(x_1, x_56, x_47, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_42); -return x_66; +lean_object* x_68; +x_68 = l_Lean_Meta_Grind_propagateEqMatchDown___lambda__1(x_1, x_58, x_53, x_48, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_43); +return x_68; } } } @@ -5034,7 +5231,7 @@ return x_66; } else { -uint8_t x_67; +uint8_t x_69; lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -5044,63 +5241,2061 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_67 = !lean_is_exclusive(x_11); -if (x_67 == 0) +x_69 = !lean_is_exclusive(x_11); +if (x_69 == 0) { return x_11; } else { -lean_object* x_68; lean_object* x_69; lean_object* x_70; -x_68 = lean_ctor_get(x_11, 0); -x_69 = lean_ctor_get(x_11, 1); -lean_inc(x_69); -lean_inc(x_68); +lean_object* x_70; lean_object* x_71; lean_object* x_72; +x_70 = lean_ctor_get(x_11, 0); +x_71 = lean_ctor_get(x_11, 1); +lean_inc(x_71); +lean_inc(x_70); lean_dec(x_11); -x_70 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_70, 0, x_68); -lean_ctor_set(x_70, 1, x_69); -return x_70; +x_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; } } } } -static lean_object* _init_l___regBuiltin_Lean_Meta_Grind_propagateHEqDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2546____closed__1() { +static lean_object* _init_l___regBuiltin_Lean_Meta_Grind_propagateEqMatchDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2629____closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_propagateHEqDown), 10, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_propagateEqMatchDown), 10, 0); return x_1; } } -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Meta_Grind_propagateHEqDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2546_(lean_object* x_1) { -_start: +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Meta_Grind_propagateEqMatchDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2629_(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_Grind_propagateEqMatchDown___closed__2; +x_3 = 0; +x_4 = l___regBuiltin_Lean_Meta_Grind_propagateEqMatchDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2629____closed__1; +x_5 = l___private_Lean_Meta_Tactic_Grind_PropagatorAttr_0__Lean_Meta_Grind_registerBuiltinPropagatorCore(x_2, x_3, x_4, x_1); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateHEqDown___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_1); +x_13 = l_Lean_Meta_Grind_mkEqTrueProof(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_13) == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; lean_object* x_19; +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +x_16 = l_Lean_Meta_Grind_propagateEqDown___lambda__1___closed__3; +x_17 = l_Lean_mkAppB(x_16, x_1, x_14); +x_18 = 1; +x_19 = l_Lean_Meta_Grind_pushEqCore(x_2, x_3, x_17, x_18, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_15); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_19; +} +else +{ +uint8_t x_20; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_20 = !lean_is_exclusive(x_13); +if (x_20 == 0) +{ +return x_13; +} +else +{ +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_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; +} +} +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateHEqDown___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("HEq", 3, 3); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateHEqDown___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_Grind_propagateHEqDown___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateHEqDown(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +lean_inc(x_1); +x_11 = l_Lean_Meta_Grind_isEqTrue(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_11) == 0) +{ +lean_object* x_12; uint8_t x_13; +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +x_13 = lean_unbox(x_12); +lean_dec(x_12); +if (x_13 == 0) +{ +uint8_t x_14; +lean_dec(x_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_14 = !lean_is_exclusive(x_11); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; +x_15 = lean_ctor_get(x_11, 0); +lean_dec(x_15); +x_16 = lean_box(0); +lean_ctor_set(x_11, 0, x_16); +return x_11; +} +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_17 = lean_ctor_get(x_11, 1); +lean_inc(x_17); +lean_dec(x_11); +x_18 = lean_box(0); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_17); +return x_19; +} +} +else +{ +uint8_t x_20; +x_20 = !lean_is_exclusive(x_11); +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_11, 1); +x_22 = lean_ctor_get(x_11, 0); +lean_dec(x_22); +lean_inc(x_1); +x_23 = l_Lean_Expr_cleanupAnnotations(x_1); +x_24 = l_Lean_Expr_isApp(x_23); +if (x_24 == 0) +{ +lean_object* x_25; +lean_dec(x_23); +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_25 = lean_box(0); +lean_ctor_set(x_11, 0, x_25); +return x_11; +} +else +{ +lean_object* x_26; lean_object* x_27; uint8_t x_28; +x_26 = l_Lean_Expr_appArg(x_23, lean_box(0)); +x_27 = l_Lean_Expr_appFnCleanup(x_23, lean_box(0)); +x_28 = l_Lean_Expr_isApp(x_27); +if (x_28 == 0) +{ +lean_object* x_29; +lean_dec(x_27); +lean_dec(x_26); +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_box(0); +lean_ctor_set(x_11, 0, x_29); +return x_11; +} +else +{ +lean_object* x_30; uint8_t x_31; +x_30 = l_Lean_Expr_appFnCleanup(x_27, lean_box(0)); +x_31 = l_Lean_Expr_isApp(x_30); +if (x_31 == 0) +{ +lean_object* x_32; +lean_dec(x_30); +lean_dec(x_26); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_32 = lean_box(0); +lean_ctor_set(x_11, 0, x_32); +return x_11; +} +else +{ +lean_object* x_33; lean_object* x_34; uint8_t x_35; +x_33 = l_Lean_Expr_appArg(x_30, lean_box(0)); +x_34 = l_Lean_Expr_appFnCleanup(x_30, lean_box(0)); +x_35 = l_Lean_Expr_isApp(x_34); +if (x_35 == 0) +{ +lean_object* x_36; +lean_dec(x_34); +lean_dec(x_33); +lean_dec(x_26); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_36 = lean_box(0); +lean_ctor_set(x_11, 0, x_36); +return x_11; +} +else +{ +lean_object* x_37; lean_object* x_38; uint8_t x_39; +x_37 = l_Lean_Expr_appFnCleanup(x_34, lean_box(0)); +x_38 = l_Lean_Meta_Grind_propagateHEqDown___closed__2; +x_39 = l_Lean_Expr_isConstOf(x_37, x_38); +lean_dec(x_37); +if (x_39 == 0) +{ +lean_object* x_40; +lean_dec(x_33); +lean_dec(x_26); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_40 = lean_box(0); +lean_ctor_set(x_11, 0, x_40); +return x_11; +} +else +{ +lean_object* x_41; +lean_free_object(x_11); +x_41 = l_Lean_Meta_Grind_propagateHEqDown___lambda__1(x_1, x_33, x_26, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_21); +return x_41; +} +} +} +} +} +} +else +{ +lean_object* x_42; lean_object* x_43; uint8_t x_44; +x_42 = lean_ctor_get(x_11, 1); +lean_inc(x_42); +lean_dec(x_11); +lean_inc(x_1); +x_43 = l_Lean_Expr_cleanupAnnotations(x_1); +x_44 = l_Lean_Expr_isApp(x_43); +if (x_44 == 0) +{ +lean_object* x_45; lean_object* x_46; +lean_dec(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); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +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_42); +return x_46; +} +else +{ +lean_object* x_47; lean_object* x_48; uint8_t x_49; +x_47 = l_Lean_Expr_appArg(x_43, lean_box(0)); +x_48 = l_Lean_Expr_appFnCleanup(x_43, lean_box(0)); +x_49 = l_Lean_Expr_isApp(x_48); +if (x_49 == 0) +{ +lean_object* x_50; lean_object* x_51; +lean_dec(x_48); +lean_dec(x_47); +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_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_42); +return x_51; +} +else +{ +lean_object* x_52; uint8_t x_53; +x_52 = l_Lean_Expr_appFnCleanup(x_48, lean_box(0)); +x_53 = l_Lean_Expr_isApp(x_52); +if (x_53 == 0) +{ +lean_object* x_54; lean_object* x_55; +lean_dec(x_52); +lean_dec(x_47); +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_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_42); +return x_55; +} +else +{ +lean_object* x_56; lean_object* x_57; uint8_t x_58; +x_56 = l_Lean_Expr_appArg(x_52, lean_box(0)); +x_57 = l_Lean_Expr_appFnCleanup(x_52, lean_box(0)); +x_58 = l_Lean_Expr_isApp(x_57); +if (x_58 == 0) +{ +lean_object* x_59; lean_object* x_60; +lean_dec(x_57); +lean_dec(x_56); +lean_dec(x_47); +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_59 = lean_box(0); +x_60 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_60, 0, x_59); +lean_ctor_set(x_60, 1, x_42); +return x_60; +} +else +{ +lean_object* x_61; lean_object* x_62; uint8_t x_63; +x_61 = l_Lean_Expr_appFnCleanup(x_57, lean_box(0)); +x_62 = l_Lean_Meta_Grind_propagateHEqDown___closed__2; +x_63 = l_Lean_Expr_isConstOf(x_61, x_62); +lean_dec(x_61); +if (x_63 == 0) +{ +lean_object* x_64; lean_object* x_65; +lean_dec(x_56); +lean_dec(x_47); +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_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_42); +return x_65; +} +else +{ +lean_object* x_66; +x_66 = l_Lean_Meta_Grind_propagateHEqDown___lambda__1(x_1, x_56, x_47, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_42); +return x_66; +} +} +} +} +} +} +} +} +else +{ +uint8_t x_67; +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_67 = !lean_is_exclusive(x_11); +if (x_67 == 0) +{ +return x_11; +} +else +{ +lean_object* x_68; lean_object* x_69; lean_object* x_70; +x_68 = lean_ctor_get(x_11, 0); +x_69 = lean_ctor_get(x_11, 1); +lean_inc(x_69); +lean_inc(x_68); +lean_dec(x_11); +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; +} +} +} +} +static lean_object* _init_l___regBuiltin_Lean_Meta_Grind_propagateHEqDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2888____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_propagateHEqDown), 10, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Meta_Grind_propagateHEqDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2888_(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_Grind_propagateHEqDown___closed__2; +x_3 = 0; +x_4 = l___regBuiltin_Lean_Meta_Grind_propagateHEqDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2888____closed__1; +x_5 = l___private_Lean_Meta_Tactic_Grind_PropagatorAttr_0__Lean_Meta_Grind_registerBuiltinPropagatorCore(x_2, x_3, x_4, x_1); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateHEqUp___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +lean_inc(x_3); +lean_inc(x_2); +x_13 = l_Lean_Meta_Grind_isEqv(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_13) == 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) +{ +uint8_t 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_16 = !lean_is_exclusive(x_13); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; +x_17 = lean_ctor_get(x_13, 0); +lean_dec(x_17); +x_18 = lean_box(0); +lean_ctor_set(x_13, 0, x_18); +return x_13; +} +else +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_13, 1); +lean_inc(x_19); +lean_dec(x_13); +x_20 = lean_box(0); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_20); +lean_ctor_set(x_21, 1, x_19); +return x_21; +} +} +else +{ +lean_object* x_22; lean_object* x_23; +x_22 = lean_ctor_get(x_13, 1); +lean_inc(x_22); +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_6); +lean_inc(x_5); +lean_inc(x_4); +x_23 = lean_grind_mk_heq_proof(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, 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; +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +x_25 = lean_ctor_get(x_23, 1); +lean_inc(x_25); +lean_dec(x_23); +x_26 = l_Lean_Meta_Grind_propagateEqUp___lambda__1___closed__3; +lean_inc(x_1); +x_27 = l_Lean_mkAppB(x_26, x_1, x_24); +x_28 = l_Lean_Meta_Grind_pushEqTrue(x_1, x_27, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_25); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_28; +} +else +{ +uint8_t x_29; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_29 = !lean_is_exclusive(x_23); +if (x_29 == 0) +{ +return x_23; +} +else +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_30 = lean_ctor_get(x_23, 0); +x_31 = lean_ctor_get(x_23, 1); +lean_inc(x_31); +lean_inc(x_30); +lean_dec(x_23); +x_32 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_32, 0, x_30); +lean_ctor_set(x_32, 1, x_31); +return x_32; +} +} +} +} +else +{ +uint8_t x_33; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_33 = !lean_is_exclusive(x_13); +if (x_33 == 0) +{ +return x_13; +} +else +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_13, 0); +x_35 = lean_ctor_get(x_13, 1); +lean_inc(x_35); +lean_inc(x_34); +lean_dec(x_13); +x_36 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_36, 0, x_34); +lean_ctor_set(x_36, 1, x_35); +return x_36; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateHEqUp(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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_inc(x_1); +x_11 = l_Lean_Expr_cleanupAnnotations(x_1); +x_12 = l_Lean_Expr_isApp(x_11); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; +lean_dec(x_11); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_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_10); +return x_14; +} +else +{ +lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_15 = l_Lean_Expr_appArg(x_11, lean_box(0)); +x_16 = l_Lean_Expr_appFnCleanup(x_11, lean_box(0)); +x_17 = l_Lean_Expr_isApp(x_16); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; +lean_dec(x_16); +lean_dec(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); +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_10); +return x_19; +} +else +{ +lean_object* x_20; uint8_t x_21; +x_20 = l_Lean_Expr_appFnCleanup(x_16, lean_box(0)); +x_21 = l_Lean_Expr_isApp(x_20); +if (x_21 == 0) +{ +lean_object* x_22; lean_object* x_23; +lean_dec(x_20); +lean_dec(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); +lean_dec(x_1); +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_10); +return x_23; +} +else +{ +lean_object* x_24; lean_object* x_25; uint8_t x_26; +x_24 = l_Lean_Expr_appArg(x_20, lean_box(0)); +x_25 = l_Lean_Expr_appFnCleanup(x_20, lean_box(0)); +x_26 = l_Lean_Expr_isApp(x_25); +if (x_26 == 0) +{ +lean_object* x_27; lean_object* x_28; +lean_dec(x_25); +lean_dec(x_24); +lean_dec(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); +lean_dec(x_1); +x_27 = lean_box(0); +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_27); +lean_ctor_set(x_28, 1, x_10); +return x_28; +} +else +{ +lean_object* x_29; lean_object* x_30; uint8_t x_31; +x_29 = l_Lean_Expr_appFnCleanup(x_25, lean_box(0)); +x_30 = l_Lean_Meta_Grind_propagateHEqDown___closed__2; +x_31 = l_Lean_Expr_isConstOf(x_29, x_30); +lean_dec(x_29); +if (x_31 == 0) +{ +lean_object* x_32; lean_object* x_33; +lean_dec(x_24); +lean_dec(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); +lean_dec(x_1); +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_10); +return x_33; +} +else +{ +lean_object* x_34; +x_34 = l_Lean_Meta_Grind_propagateHEqUp___lambda__1(x_1, x_24, x_15, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_34; +} +} +} +} +} +} +} +static lean_object* _init_l___regBuiltin_Lean_Meta_Grind_propagateHEqUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_3148____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_propagateHEqUp), 10, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Meta_Grind_propagateHEqUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_3148_(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_Grind_propagateHEqDown___closed__2; +x_3 = 1; +x_4 = l___regBuiltin_Lean_Meta_Grind_propagateHEqUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_3148____closed__1; +x_5 = l___private_Lean_Meta_Tactic_Grind_PropagatorAttr_0__Lean_Meta_Grind_registerBuiltinPropagatorCore(x_2, x_3, x_4, x_1); +return x_5; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateIte___lambda__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("ite_cond_eq_false", 17, 17); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateIte___lambda__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind_propagateIte___lambda__1___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateIte___lambda__1___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("ite_cond_eq_true", 16, 16); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateIte___lambda__1___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind_propagateIte___lambda__1___closed__3; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateIte___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +_start: +{ +lean_object* x_17; +lean_inc(x_4); +x_17 = l_Lean_Meta_Grind_isEqTrue(x_4, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +if (lean_obj_tag(x_17) == 0) +{ +lean_object* x_18; uint8_t x_19; +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +x_19 = lean_unbox(x_18); +lean_dec(x_18); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; +x_20 = lean_ctor_get(x_17, 1); +lean_inc(x_20); +lean_dec(x_17); +lean_inc(x_4); +x_21 = l_Lean_Meta_Grind_isEqFalse(x_4, x_8, x_9, 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; uint8_t x_23; +x_22 = lean_ctor_get(x_21, 0); +lean_inc(x_22); +x_23 = lean_unbox(x_22); +lean_dec(x_22); +if (x_23 == 0) +{ +uint8_t x_24; +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_24 = !lean_is_exclusive(x_21); +if (x_24 == 0) +{ +lean_object* x_25; lean_object* x_26; +x_25 = lean_ctor_get(x_21, 0); +lean_dec(x_25); +x_26 = lean_box(0); +lean_ctor_set(x_21, 0, x_26); +return x_21; +} +else +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_21, 1); +lean_inc(x_27); +lean_dec(x_21); +x_28 = lean_box(0); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_27); +return x_29; +} +} +else +{ +lean_object* x_30; lean_object* x_31; +x_30 = lean_ctor_get(x_21, 1); +lean_inc(x_30); +lean_dec(x_21); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_4); +x_31 = l_Lean_Meta_Grind_mkEqFalseProof(x_4, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_30); +if (lean_obj_tag(x_31) == 0) +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; uint8_t x_38; lean_object* x_39; +x_32 = lean_ctor_get(x_31, 0); +lean_inc(x_32); +x_33 = lean_ctor_get(x_31, 1); +lean_inc(x_33); +lean_dec(x_31); +x_34 = l_Lean_Expr_constLevels_x21(x_2); +x_35 = l_Lean_Meta_Grind_propagateIte___lambda__1___closed__2; +x_36 = l_Lean_Expr_const___override(x_35, x_34); +lean_inc(x_7); +x_37 = l_Lean_mkApp6(x_36, x_3, x_4, x_5, x_6, x_7, x_32); +x_38 = 0; +x_39 = l_Lean_Meta_Grind_pushEqCore(x_1, x_7, x_37, x_38, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, 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_9); +lean_dec(x_8); +return x_39; +} +else +{ +uint8_t x_40; +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_40 = !lean_is_exclusive(x_31); +if (x_40 == 0) +{ +return x_31; +} +else +{ +lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_41 = lean_ctor_get(x_31, 0); +x_42 = lean_ctor_get(x_31, 1); +lean_inc(x_42); +lean_inc(x_41); +lean_dec(x_31); +x_43 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_43, 0, x_41); +lean_ctor_set(x_43, 1, x_42); +return x_43; +} +} +} +} +else +{ +uint8_t x_44; +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_44 = !lean_is_exclusive(x_21); +if (x_44 == 0) +{ +return x_21; +} +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_inc(x_46); +lean_inc(x_45); +lean_dec(x_21); +x_47 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_47, 0, x_45); +lean_ctor_set(x_47, 1, x_46); +return x_47; +} +} +} +else +{ +lean_object* x_48; lean_object* x_49; +x_48 = lean_ctor_get(x_17, 1); +lean_inc(x_48); +lean_dec(x_17); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_4); +x_49 = l_Lean_Meta_Grind_mkEqTrueProof(x_4, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_48); +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; uint8_t 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 = l_Lean_Expr_constLevels_x21(x_2); +x_53 = l_Lean_Meta_Grind_propagateIte___lambda__1___closed__4; +x_54 = l_Lean_Expr_const___override(x_53, x_52); +lean_inc(x_6); +x_55 = l_Lean_mkApp6(x_54, x_3, x_4, x_5, x_6, x_7, x_50); +x_56 = 0; +x_57 = l_Lean_Meta_Grind_pushEqCore(x_1, x_6, x_55, x_56, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_51); +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); +return x_57; +} +else +{ +uint8_t x_58; +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_58 = !lean_is_exclusive(x_49); +if (x_58 == 0) +{ +return x_49; +} +else +{ +lean_object* x_59; lean_object* x_60; lean_object* x_61; +x_59 = lean_ctor_get(x_49, 0); +x_60 = lean_ctor_get(x_49, 1); +lean_inc(x_60); +lean_inc(x_59); +lean_dec(x_49); +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_62; +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_62 = !lean_is_exclusive(x_17); +if (x_62 == 0) +{ +return x_17; +} +else +{ +lean_object* x_63; lean_object* x_64; lean_object* x_65; +x_63 = lean_ctor_get(x_17, 0); +x_64 = lean_ctor_get(x_17, 1); +lean_inc(x_64); +lean_inc(x_63); +lean_dec(x_17); +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; +} +} +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateIte___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("ite", 3, 3); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateIte___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_Grind_propagateIte___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateIte(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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_inc(x_1); +x_11 = l_Lean_Expr_cleanupAnnotations(x_1); +x_12 = l_Lean_Expr_isApp(x_11); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; +lean_dec(x_11); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_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_10); +return x_14; +} +else +{ +lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_15 = l_Lean_Expr_appArg(x_11, lean_box(0)); +x_16 = l_Lean_Expr_appFnCleanup(x_11, lean_box(0)); +x_17 = l_Lean_Expr_isApp(x_16); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; +lean_dec(x_16); +lean_dec(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); +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_10); +return x_19; +} +else +{ +lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_20 = l_Lean_Expr_appArg(x_16, lean_box(0)); +x_21 = l_Lean_Expr_appFnCleanup(x_16, lean_box(0)); +x_22 = l_Lean_Expr_isApp(x_21); +if (x_22 == 0) +{ +lean_object* x_23; lean_object* x_24; +lean_dec(x_21); +lean_dec(x_20); +lean_dec(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); +lean_dec(x_1); +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_10); +return x_24; +} +else +{ +lean_object* x_25; lean_object* x_26; uint8_t x_27; +x_25 = l_Lean_Expr_appArg(x_21, lean_box(0)); +x_26 = l_Lean_Expr_appFnCleanup(x_21, lean_box(0)); +x_27 = l_Lean_Expr_isApp(x_26); +if (x_27 == 0) +{ +lean_object* x_28; lean_object* x_29; +lean_dec(x_26); +lean_dec(x_25); +lean_dec(x_20); +lean_dec(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); +lean_dec(x_1); +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_10); +return x_29; +} +else +{ +lean_object* x_30; lean_object* x_31; uint8_t x_32; +x_30 = l_Lean_Expr_appArg(x_26, lean_box(0)); +x_31 = l_Lean_Expr_appFnCleanup(x_26, lean_box(0)); +x_32 = l_Lean_Expr_isApp(x_31); +if (x_32 == 0) +{ +lean_object* x_33; lean_object* x_34; +lean_dec(x_31); +lean_dec(x_30); +lean_dec(x_25); +lean_dec(x_20); +lean_dec(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); +lean_dec(x_1); +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_10); +return x_34; +} +else +{ +lean_object* x_35; lean_object* x_36; lean_object* x_37; uint8_t x_38; +x_35 = l_Lean_Expr_appArg(x_31, lean_box(0)); +x_36 = l_Lean_Expr_appFnCleanup(x_31, lean_box(0)); +x_37 = l_Lean_Meta_Grind_propagateIte___closed__2; +x_38 = l_Lean_Expr_isConstOf(x_36, x_37); +if (x_38 == 0) +{ +lean_object* x_39; lean_object* x_40; +lean_dec(x_36); +lean_dec(x_35); +lean_dec(x_30); +lean_dec(x_25); +lean_dec(x_20); +lean_dec(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); +lean_dec(x_1); +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_10); +return x_40; +} +else +{ +lean_object* x_41; +x_41 = l_Lean_Meta_Grind_propagateIte___lambda__1(x_1, x_36, x_35, x_30, x_25, x_20, x_15, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_36); +return x_41; +} +} +} +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateIte___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +_start: +{ +lean_object* x_17; +x_17 = l_Lean_Meta_Grind_propagateIte___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +lean_dec(x_2); +return x_17; +} +} +static lean_object* _init_l___regBuiltin_Lean_Meta_Grind_propagateIte_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_3530____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_propagateIte), 10, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Meta_Grind_propagateIte_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_3530_(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_Grind_propagateIte___closed__2; +x_3 = 1; +x_4 = l___regBuiltin_Lean_Meta_Grind_propagateIte_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_3530____closed__1; +x_5 = l___private_Lean_Meta_Tactic_Grind_PropagatorAttr_0__Lean_Meta_Grind_registerBuiltinPropagatorCore(x_2, x_3, x_4, x_1); +return x_5; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("of_eq_false", 11, 11); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__2; +x_3 = l_Lean_Expr_const___override(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("dite_cond_eq_false'", 19, 19); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Meta_Grind_propagateAndUp___lambda__1___closed__1; +x_2 = l_Lean_Meta_Grind_propagateAndUp___lambda__1___closed__2; +x_3 = l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__4; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("dite_cond_eq_true'", 18, 18); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Meta_Grind_propagateAndUp___lambda__1___closed__1; +x_2 = l_Lean_Meta_Grind_propagateAndUp___lambda__1___closed__2; +x_3 = l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__6; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateDIte___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +_start: +{ +lean_object* x_17; +lean_inc(x_4); +x_17 = l_Lean_Meta_Grind_isEqTrue(x_4, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +if (lean_obj_tag(x_17) == 0) +{ +lean_object* x_18; uint8_t x_19; +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +x_19 = lean_unbox(x_18); +lean_dec(x_18); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; +x_20 = lean_ctor_get(x_17, 1); +lean_inc(x_20); +lean_dec(x_17); +lean_inc(x_4); +x_21 = l_Lean_Meta_Grind_isEqFalse(x_4, x_8, x_9, 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; uint8_t x_23; +x_22 = lean_ctor_get(x_21, 0); +lean_inc(x_22); +x_23 = lean_unbox(x_22); +lean_dec(x_22); +if (x_23 == 0) +{ +uint8_t x_24; +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_24 = !lean_is_exclusive(x_21); +if (x_24 == 0) +{ +lean_object* x_25; lean_object* x_26; +x_25 = lean_ctor_get(x_21, 0); +lean_dec(x_25); +x_26 = lean_box(0); +lean_ctor_set(x_21, 0, x_26); +return x_21; +} +else +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_21, 1); +lean_inc(x_27); +lean_dec(x_21); +x_28 = lean_box(0); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_27); +return x_29; +} +} +else +{ +lean_object* x_30; lean_object* x_31; +x_30 = lean_ctor_get(x_21, 1); +lean_inc(x_30); +lean_dec(x_21); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_4); +x_31 = l_Lean_Meta_Grind_mkEqFalseProof(x_4, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_30); +if (lean_obj_tag(x_31) == 0) +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; +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_Meta_Grind_propagateDIte___lambda__1___closed__3; +lean_inc(x_32); +lean_inc(x_4); +x_35 = l_Lean_mkAppB(x_34, x_4, x_32); +lean_inc(x_7); +x_36 = l_Lean_Expr_app___override(x_7, x_35); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_10); +x_37 = l_Lean_Meta_Grind_simp(x_36, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_33); +if (lean_obj_tag(x_37) == 0) +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_38 = lean_ctor_get(x_37, 0); +lean_inc(x_38); +x_39 = lean_ctor_get(x_37, 1); +lean_inc(x_39); +lean_dec(x_37); +x_40 = lean_ctor_get(x_38, 0); +lean_inc(x_40); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +x_41 = l_Lean_Meta_Simp_Result_getProof(x_38, x_12, x_13, x_14, x_15, x_39); +if (lean_obj_tag(x_41) == 0) +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_42 = lean_ctor_get(x_41, 0); +lean_inc(x_42); +x_43 = lean_ctor_get(x_41, 1); +lean_inc(x_43); +lean_dec(x_41); +lean_inc(x_1); +x_44 = l_Lean_Meta_Grind_getGeneration(x_1, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_43); +if (lean_obj_tag(x_44) == 0) +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_45 = lean_ctor_get(x_44, 0); +lean_inc(x_45); +x_46 = lean_ctor_get(x_44, 1); +lean_inc(x_46); +lean_dec(x_44); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_40); +x_47 = l_Lean_Meta_Grind_internalize(x_40, x_45, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_46); +if (lean_obj_tag(x_47) == 0) +{ +lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; uint8_t x_53; lean_object* x_54; +x_48 = lean_ctor_get(x_47, 1); +lean_inc(x_48); +lean_dec(x_47); +x_49 = l_Lean_Expr_constLevels_x21(x_2); +x_50 = l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__5; +x_51 = l_Lean_Expr_const___override(x_50, x_49); +lean_inc(x_40); +x_52 = l_Lean_mkApp8(x_51, x_3, x_4, x_5, x_6, x_7, x_40, x_32, x_42); +x_53 = 0; +x_54 = l_Lean_Meta_Grind_pushEqCore(x_1, x_40, x_52, x_53, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_48); +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); +return x_54; +} +else +{ +uint8_t x_55; +lean_dec(x_42); +lean_dec(x_40); +lean_dec(x_32); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_55 = !lean_is_exclusive(x_47); +if (x_55 == 0) +{ +return x_47; +} +else +{ +lean_object* x_56; lean_object* x_57; lean_object* x_58; +x_56 = lean_ctor_get(x_47, 0); +x_57 = lean_ctor_get(x_47, 1); +lean_inc(x_57); +lean_inc(x_56); +lean_dec(x_47); +x_58 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_58, 0, x_56); +lean_ctor_set(x_58, 1, x_57); +return x_58; +} +} +} +else +{ +uint8_t x_59; +lean_dec(x_42); +lean_dec(x_40); +lean_dec(x_32); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_59 = !lean_is_exclusive(x_44); +if (x_59 == 0) +{ +return x_44; +} +else +{ +lean_object* x_60; lean_object* x_61; lean_object* x_62; +x_60 = lean_ctor_get(x_44, 0); +x_61 = lean_ctor_get(x_44, 1); +lean_inc(x_61); +lean_inc(x_60); +lean_dec(x_44); +x_62 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_62, 0, x_60); +lean_ctor_set(x_62, 1, x_61); +return x_62; +} +} +} +else +{ +uint8_t x_63; +lean_dec(x_40); +lean_dec(x_32); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_63 = !lean_is_exclusive(x_41); +if (x_63 == 0) +{ +return x_41; +} +else +{ +lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_64 = lean_ctor_get(x_41, 0); +x_65 = lean_ctor_get(x_41, 1); +lean_inc(x_65); +lean_inc(x_64); +lean_dec(x_41); +x_66 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_66, 0, x_64); +lean_ctor_set(x_66, 1, x_65); +return x_66; +} +} +} +else +{ +uint8_t x_67; +lean_dec(x_32); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_67 = !lean_is_exclusive(x_37); +if (x_67 == 0) +{ +return x_37; +} +else +{ +lean_object* x_68; lean_object* x_69; lean_object* x_70; +x_68 = lean_ctor_get(x_37, 0); +x_69 = lean_ctor_get(x_37, 1); +lean_inc(x_69); +lean_inc(x_68); +lean_dec(x_37); +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 +{ +uint8_t x_71; +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_71 = !lean_is_exclusive(x_31); +if (x_71 == 0) +{ +return x_31; +} +else +{ +lean_object* x_72; lean_object* x_73; lean_object* x_74; +x_72 = lean_ctor_get(x_31, 0); +x_73 = lean_ctor_get(x_31, 1); +lean_inc(x_73); +lean_inc(x_72); +lean_dec(x_31); +x_74 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_74, 0, x_72); +lean_ctor_set(x_74, 1, x_73); +return x_74; +} +} +} +} +else +{ +uint8_t x_75; +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_75 = !lean_is_exclusive(x_21); +if (x_75 == 0) +{ +return x_21; +} +else +{ +lean_object* x_76; lean_object* x_77; lean_object* x_78; +x_76 = lean_ctor_get(x_21, 0); +x_77 = lean_ctor_get(x_21, 1); +lean_inc(x_77); +lean_inc(x_76); +lean_dec(x_21); +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_79; lean_object* x_80; +x_79 = lean_ctor_get(x_17, 1); +lean_inc(x_79); +lean_dec(x_17); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_4); +x_80 = l_Lean_Meta_Grind_mkEqTrueProof(x_4, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_79); +if (lean_obj_tag(x_80) == 0) +{ +lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; +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_Grind_propagateEqDown___lambda__1___closed__3; +lean_inc(x_81); +lean_inc(x_4); +x_84 = l_Lean_mkAppB(x_83, x_4, x_81); +lean_inc(x_6); +x_85 = l_Lean_Expr_app___override(x_6, x_84); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_10); +x_86 = l_Lean_Meta_Grind_simp(x_85, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_82); +if (lean_obj_tag(x_86) == 0) +{ +lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; +x_87 = lean_ctor_get(x_86, 0); +lean_inc(x_87); +x_88 = lean_ctor_get(x_86, 1); +lean_inc(x_88); +lean_dec(x_86); +x_89 = lean_ctor_get(x_87, 0); +lean_inc(x_89); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +x_90 = l_Lean_Meta_Simp_Result_getProof(x_87, x_12, x_13, x_14, x_15, x_88); +if (lean_obj_tag(x_90) == 0) +{ +lean_object* x_91; lean_object* x_92; lean_object* x_93; +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); +lean_inc(x_1); +x_93 = l_Lean_Meta_Grind_getGeneration(x_1, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_92); +if (lean_obj_tag(x_93) == 0) +{ +lean_object* x_94; lean_object* x_95; lean_object* x_96; +x_94 = lean_ctor_get(x_93, 0); +lean_inc(x_94); +x_95 = lean_ctor_get(x_93, 1); +lean_inc(x_95); +lean_dec(x_93); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_89); +x_96 = l_Lean_Meta_Grind_internalize(x_89, x_94, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_95); +if (lean_obj_tag(x_96) == 0) +{ +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; +x_97 = lean_ctor_get(x_96, 1); +lean_inc(x_97); +lean_dec(x_96); +x_98 = l_Lean_Expr_constLevels_x21(x_2); +x_99 = l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__7; +x_100 = l_Lean_Expr_const___override(x_99, x_98); +lean_inc(x_89); +x_101 = l_Lean_mkApp8(x_100, x_3, x_4, x_5, x_6, x_7, x_89, x_81, x_91); +x_102 = 0; +x_103 = l_Lean_Meta_Grind_pushEqCore(x_1, x_89, x_101, x_102, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_97); +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); +return x_103; +} +else +{ +uint8_t x_104; +lean_dec(x_91); +lean_dec(x_89); +lean_dec(x_81); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_104 = !lean_is_exclusive(x_96); +if (x_104 == 0) +{ +return x_96; +} +else +{ +lean_object* x_105; lean_object* x_106; lean_object* x_107; +x_105 = lean_ctor_get(x_96, 0); +x_106 = lean_ctor_get(x_96, 1); +lean_inc(x_106); +lean_inc(x_105); +lean_dec(x_96); +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_108; +lean_dec(x_91); +lean_dec(x_89); +lean_dec(x_81); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_108 = !lean_is_exclusive(x_93); +if (x_108 == 0) +{ +return x_93; +} +else { -lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; -x_2 = l_Lean_Meta_Grind_propagateHEqDown___closed__2; -x_3 = 0; -x_4 = l___regBuiltin_Lean_Meta_Grind_propagateHEqDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2546____closed__1; -x_5 = l___private_Lean_Meta_Tactic_Grind_PropagatorAttr_0__Lean_Meta_Grind_registerBuiltinPropagatorCore(x_2, x_3, x_4, x_1); -return x_5; +lean_object* x_109; lean_object* x_110; lean_object* x_111; +x_109 = lean_ctor_get(x_93, 0); +x_110 = lean_ctor_get(x_93, 1); +lean_inc(x_110); +lean_inc(x_109); +lean_dec(x_93); +x_111 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_111, 0, x_109); +lean_ctor_set(x_111, 1, x_110); +return x_111; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateHEqUp___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -lean_object* x_13; -lean_inc(x_3); -lean_inc(x_2); -x_13 = l_Lean_Meta_Grind_isEqv(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_13) == 0) +} +else { -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); +uint8_t x_112; +lean_dec(x_89); +lean_dec(x_81); +lean_dec(x_15); lean_dec(x_14); -if (x_15 == 0) -{ -uint8_t x_16; +lean_dec(x_13); +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -5110,58 +7305,35 @@ 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_16 = !lean_is_exclusive(x_13); -if (x_16 == 0) +x_112 = !lean_is_exclusive(x_90); +if (x_112 == 0) { -lean_object* x_17; lean_object* x_18; -x_17 = lean_ctor_get(x_13, 0); -lean_dec(x_17); -x_18 = lean_box(0); -lean_ctor_set(x_13, 0, x_18); -return x_13; +return x_90; } else { -lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_19 = lean_ctor_get(x_13, 1); -lean_inc(x_19); -lean_dec(x_13); -x_20 = lean_box(0); -x_21 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_21, 0, x_20); -lean_ctor_set(x_21, 1, x_19); -return x_21; +lean_object* x_113; lean_object* x_114; lean_object* x_115; +x_113 = lean_ctor_get(x_90, 0); +x_114 = lean_ctor_get(x_90, 1); +lean_inc(x_114); +lean_inc(x_113); +lean_dec(x_90); +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_22; lean_object* x_23; -x_22 = lean_ctor_get(x_13, 1); -lean_inc(x_22); +uint8_t x_116; +lean_dec(x_81); +lean_dec(x_15); +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_6); -lean_inc(x_5); -lean_inc(x_4); -x_23 = lean_grind_mk_heq_proof(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, 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; -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -x_25 = lean_ctor_get(x_23, 1); -lean_inc(x_25); -lean_dec(x_23); -x_26 = l_Lean_Meta_Grind_propagateEqUp___lambda__1___closed__3; -lean_inc(x_1); -x_27 = l_Lean_mkAppB(x_26, x_1, x_24); -x_28 = l_Lean_Meta_Grind_pushEqTrue(x_1, x_27, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_25); +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -5170,11 +7342,35 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -return x_28; +lean_dec(x_3); +lean_dec(x_1); +x_116 = !lean_is_exclusive(x_86); +if (x_116 == 0) +{ +return x_86; } else { -uint8_t x_29; +lean_object* x_117; lean_object* x_118; lean_object* x_119; +x_117 = lean_ctor_get(x_86, 0); +x_118 = lean_ctor_get(x_86, 1); +lean_inc(x_118); +lean_inc(x_117); +lean_dec(x_86); +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_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); @@ -5183,31 +7379,36 @@ 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_29 = !lean_is_exclusive(x_23); -if (x_29 == 0) +x_120 = !lean_is_exclusive(x_80); +if (x_120 == 0) { -return x_23; +return x_80; } else { -lean_object* x_30; lean_object* x_31; lean_object* x_32; -x_30 = lean_ctor_get(x_23, 0); -x_31 = lean_ctor_get(x_23, 1); -lean_inc(x_31); -lean_inc(x_30); -lean_dec(x_23); -x_32 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_32, 0, x_30); -lean_ctor_set(x_32, 1, x_31); -return x_32; +lean_object* x_121; lean_object* x_122; lean_object* x_123; +x_121 = lean_ctor_get(x_80, 0); +x_122 = lean_ctor_get(x_80, 1); +lean_inc(x_122); +lean_inc(x_121); +lean_dec(x_80); +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_33; +uint8_t x_124; +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); @@ -5217,30 +7418,47 @@ 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_33 = !lean_is_exclusive(x_13); -if (x_33 == 0) +x_124 = !lean_is_exclusive(x_17); +if (x_124 == 0) { -return x_13; +return x_17; } else { -lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_34 = lean_ctor_get(x_13, 0); -x_35 = lean_ctor_get(x_13, 1); -lean_inc(x_35); -lean_inc(x_34); -lean_dec(x_13); -x_36 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_36, 0, x_34); -lean_ctor_set(x_36, 1, x_35); -return x_36; +lean_object* x_125; lean_object* x_126; lean_object* x_127; +x_125 = lean_ctor_get(x_17, 0); +x_126 = lean_ctor_get(x_17, 1); +lean_inc(x_126); +lean_inc(x_125); +lean_dec(x_17); +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_EXPORT lean_object* l_Lean_Meta_Grind_propagateHEqUp(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +static lean_object* _init_l_Lean_Meta_Grind_propagateDIte___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("dite", 4, 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateDIte___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_Grind_propagateDIte___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateDIte(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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; @@ -5294,12 +7512,14 @@ return x_19; } else { -lean_object* x_20; uint8_t x_21; -x_20 = l_Lean_Expr_appFnCleanup(x_16, lean_box(0)); -x_21 = l_Lean_Expr_isApp(x_20); -if (x_21 == 0) +lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_20 = l_Lean_Expr_appArg(x_16, lean_box(0)); +x_21 = l_Lean_Expr_appFnCleanup(x_16, lean_box(0)); +x_22 = l_Lean_Expr_isApp(x_21); +if (x_22 == 0) { -lean_object* x_22; lean_object* x_23; +lean_object* x_23; lean_object* x_24; +lean_dec(x_21); lean_dec(x_20); lean_dec(x_15); lean_dec(x_9); @@ -5311,23 +7531,24 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -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_10); -return x_23; +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_10); +return x_24; } else { -lean_object* x_24; lean_object* x_25; uint8_t x_26; -x_24 = l_Lean_Expr_appArg(x_20, lean_box(0)); -x_25 = l_Lean_Expr_appFnCleanup(x_20, lean_box(0)); -x_26 = l_Lean_Expr_isApp(x_25); -if (x_26 == 0) +lean_object* x_25; lean_object* x_26; uint8_t x_27; +x_25 = l_Lean_Expr_appArg(x_21, lean_box(0)); +x_26 = l_Lean_Expr_appFnCleanup(x_21, lean_box(0)); +x_27 = l_Lean_Expr_isApp(x_26); +if (x_27 == 0) { -lean_object* x_27; lean_object* x_28; +lean_object* x_28; lean_object* x_29; +lean_dec(x_26); lean_dec(x_25); -lean_dec(x_24); +lean_dec(x_20); lean_dec(x_15); lean_dec(x_9); lean_dec(x_8); @@ -5338,23 +7559,25 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_27 = lean_box(0); -x_28 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_28, 0, x_27); -lean_ctor_set(x_28, 1, x_10); -return x_28; +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_10); +return x_29; } else { -lean_object* x_29; lean_object* x_30; uint8_t x_31; -x_29 = l_Lean_Expr_appFnCleanup(x_25, lean_box(0)); -x_30 = l_Lean_Meta_Grind_propagateHEqDown___closed__2; -x_31 = l_Lean_Expr_isConstOf(x_29, x_30); -lean_dec(x_29); -if (x_31 == 0) +lean_object* x_30; lean_object* x_31; uint8_t x_32; +x_30 = l_Lean_Expr_appArg(x_26, lean_box(0)); +x_31 = l_Lean_Expr_appFnCleanup(x_26, lean_box(0)); +x_32 = l_Lean_Expr_isApp(x_31); +if (x_32 == 0) { -lean_object* x_32; lean_object* x_33; -lean_dec(x_24); +lean_object* x_33; lean_object* x_34; +lean_dec(x_31); +lean_dec(x_30); +lean_dec(x_25); +lean_dec(x_20); lean_dec(x_15); lean_dec(x_9); lean_dec(x_8); @@ -5365,39 +7588,81 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -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_10); -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_10); +return x_34; } else { -lean_object* x_34; -x_34 = l_Lean_Meta_Grind_propagateHEqUp___lambda__1(x_1, x_24, x_15, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_34; +lean_object* x_35; lean_object* x_36; lean_object* x_37; uint8_t x_38; +x_35 = l_Lean_Expr_appArg(x_31, lean_box(0)); +x_36 = l_Lean_Expr_appFnCleanup(x_31, lean_box(0)); +x_37 = l_Lean_Meta_Grind_propagateDIte___closed__2; +x_38 = l_Lean_Expr_isConstOf(x_36, x_37); +if (x_38 == 0) +{ +lean_object* x_39; lean_object* x_40; +lean_dec(x_36); +lean_dec(x_35); +lean_dec(x_30); +lean_dec(x_25); +lean_dec(x_20); +lean_dec(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); +lean_dec(x_1); +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_10); +return x_40; +} +else +{ +lean_object* x_41; +x_41 = l_Lean_Meta_Grind_propagateDIte___lambda__1(x_1, x_36, x_35, x_30, x_25, x_20, x_15, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_36); +return x_41; +} +} +} } } } } } +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateDIte___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +_start: +{ +lean_object* x_17; +x_17 = l_Lean_Meta_Grind_propagateDIte___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +lean_dec(x_2); +return x_17; } } -static lean_object* _init_l___regBuiltin_Lean_Meta_Grind_propagateHEqUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2806____closed__1() { +static lean_object* _init_l___regBuiltin_Lean_Meta_Grind_propagateDIte_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_4078____closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_propagateHEqUp), 10, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_propagateDIte), 10, 0); return x_1; } } -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Meta_Grind_propagateHEqUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2806_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Meta_Grind_propagateDIte_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_4078_(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_Grind_propagateHEqDown___closed__2; +x_2 = l_Lean_Meta_Grind_propagateDIte___closed__2; x_3 = 1; -x_4 = l___regBuiltin_Lean_Meta_Grind_propagateHEqUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2806____closed__1; +x_4 = l___regBuiltin_Lean_Meta_Grind_propagateDIte_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_4078____closed__1; x_5 = l___private_Lean_Meta_Tactic_Grind_PropagatorAttr_0__Lean_Meta_Grind_registerBuiltinPropagatorCore(x_2, x_3, x_4, x_1); return x_5; } @@ -5405,6 +7670,8 @@ return x_5; lean_object* initialize_Init_Grind(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_Tactic_Grind_Proof(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_Tactic_Grind_PropagatorAttr(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Meta_Tactic_Grind_Simp(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Meta_Tactic_Grind_Internalize(uint8_t builtin, lean_object*); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Lean_Meta_Tactic_Grind_Propagate(uint8_t builtin, lean_object* w) { lean_object * res; @@ -5419,6 +7686,12 @@ lean_dec_ref(res); res = initialize_Lean_Meta_Tactic_Grind_PropagatorAttr(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +res = initialize_Lean_Meta_Tactic_Grind_Simp(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Lean_Meta_Tactic_Grind_Internalize(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); l_Lean_Meta_Grind_propagateAndUp___lambda__1___closed__1 = _init_l_Lean_Meta_Grind_propagateAndUp___lambda__1___closed__1(); lean_mark_persistent(l_Lean_Meta_Grind_propagateAndUp___lambda__1___closed__1); l_Lean_Meta_Grind_propagateAndUp___lambda__1___closed__2 = _init_l_Lean_Meta_Grind_propagateAndUp___lambda__1___closed__2(); @@ -5537,13 +7810,19 @@ l_Lean_Meta_Grind_propagateNotUp___lambda__1___closed__5 = _init_l_Lean_Meta_Gri lean_mark_persistent(l_Lean_Meta_Grind_propagateNotUp___lambda__1___closed__5); l_Lean_Meta_Grind_propagateNotUp___lambda__1___closed__6 = _init_l_Lean_Meta_Grind_propagateNotUp___lambda__1___closed__6(); lean_mark_persistent(l_Lean_Meta_Grind_propagateNotUp___lambda__1___closed__6); +l_Lean_Meta_Grind_propagateNotUp___lambda__1___closed__7 = _init_l_Lean_Meta_Grind_propagateNotUp___lambda__1___closed__7(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateNotUp___lambda__1___closed__7); +l_Lean_Meta_Grind_propagateNotUp___lambda__1___closed__8 = _init_l_Lean_Meta_Grind_propagateNotUp___lambda__1___closed__8(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateNotUp___lambda__1___closed__8); +l_Lean_Meta_Grind_propagateNotUp___lambda__1___closed__9 = _init_l_Lean_Meta_Grind_propagateNotUp___lambda__1___closed__9(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateNotUp___lambda__1___closed__9); l_Lean_Meta_Grind_propagateNotUp___closed__1 = _init_l_Lean_Meta_Grind_propagateNotUp___closed__1(); lean_mark_persistent(l_Lean_Meta_Grind_propagateNotUp___closed__1); l_Lean_Meta_Grind_propagateNotUp___closed__2 = _init_l_Lean_Meta_Grind_propagateNotUp___closed__2(); lean_mark_persistent(l_Lean_Meta_Grind_propagateNotUp___closed__2); -l___regBuiltin_Lean_Meta_Grind_propagateNotUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_1426____closed__1 = _init_l___regBuiltin_Lean_Meta_Grind_propagateNotUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_1426____closed__1(); -lean_mark_persistent(l___regBuiltin_Lean_Meta_Grind_propagateNotUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_1426____closed__1); -if (builtin) {res = l___regBuiltin_Lean_Meta_Grind_propagateNotUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_1426_(lean_io_mk_world()); +l___regBuiltin_Lean_Meta_Grind_propagateNotUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_1486____closed__1 = _init_l___regBuiltin_Lean_Meta_Grind_propagateNotUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_1486____closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Meta_Grind_propagateNotUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_1486____closed__1); +if (builtin) {res = l___regBuiltin_Lean_Meta_Grind_propagateNotUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_1486_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); }l_Lean_Meta_Grind_propagateNotDown___lambda__1___closed__1 = _init_l_Lean_Meta_Grind_propagateNotDown___lambda__1___closed__1(); @@ -5558,15 +7837,9 @@ l_Lean_Meta_Grind_propagateNotDown___lambda__1___closed__5 = _init_l_Lean_Meta_G lean_mark_persistent(l_Lean_Meta_Grind_propagateNotDown___lambda__1___closed__5); l_Lean_Meta_Grind_propagateNotDown___lambda__1___closed__6 = _init_l_Lean_Meta_Grind_propagateNotDown___lambda__1___closed__6(); lean_mark_persistent(l_Lean_Meta_Grind_propagateNotDown___lambda__1___closed__6); -l_Lean_Meta_Grind_propagateNotDown___lambda__1___closed__7 = _init_l_Lean_Meta_Grind_propagateNotDown___lambda__1___closed__7(); -lean_mark_persistent(l_Lean_Meta_Grind_propagateNotDown___lambda__1___closed__7); -l_Lean_Meta_Grind_propagateNotDown___lambda__1___closed__8 = _init_l_Lean_Meta_Grind_propagateNotDown___lambda__1___closed__8(); -lean_mark_persistent(l_Lean_Meta_Grind_propagateNotDown___lambda__1___closed__8); -l_Lean_Meta_Grind_propagateNotDown___lambda__1___closed__9 = _init_l_Lean_Meta_Grind_propagateNotDown___lambda__1___closed__9(); -lean_mark_persistent(l_Lean_Meta_Grind_propagateNotDown___lambda__1___closed__9); -l___regBuiltin_Lean_Meta_Grind_propagateNotDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_1706____closed__1 = _init_l___regBuiltin_Lean_Meta_Grind_propagateNotDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_1706____closed__1(); -lean_mark_persistent(l___regBuiltin_Lean_Meta_Grind_propagateNotDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_1706____closed__1); -if (builtin) {res = l___regBuiltin_Lean_Meta_Grind_propagateNotDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_1706_(lean_io_mk_world()); +l___regBuiltin_Lean_Meta_Grind_propagateNotDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_1766____closed__1 = _init_l___regBuiltin_Lean_Meta_Grind_propagateNotDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_1766____closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Meta_Grind_propagateNotDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_1766____closed__1); +if (builtin) {res = l___regBuiltin_Lean_Meta_Grind_propagateNotDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_1766_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); }l_Lean_Meta_Grind_propagateEqUp___lambda__1___closed__1 = _init_l_Lean_Meta_Grind_propagateEqUp___lambda__1___closed__1(); @@ -5591,9 +7864,9 @@ l_Lean_Meta_Grind_propagateEqUp___closed__1 = _init_l_Lean_Meta_Grind_propagateE lean_mark_persistent(l_Lean_Meta_Grind_propagateEqUp___closed__1); l_Lean_Meta_Grind_propagateEqUp___closed__2 = _init_l_Lean_Meta_Grind_propagateEqUp___closed__2(); lean_mark_persistent(l_Lean_Meta_Grind_propagateEqUp___closed__2); -l___regBuiltin_Lean_Meta_Grind_propagateEqUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2058____closed__1 = _init_l___regBuiltin_Lean_Meta_Grind_propagateEqUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2058____closed__1(); -lean_mark_persistent(l___regBuiltin_Lean_Meta_Grind_propagateEqUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2058____closed__1); -if (builtin) {res = l___regBuiltin_Lean_Meta_Grind_propagateEqUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2058_(lean_io_mk_world()); +l___regBuiltin_Lean_Meta_Grind_propagateEqUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2118____closed__1 = _init_l___regBuiltin_Lean_Meta_Grind_propagateEqUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2118____closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Meta_Grind_propagateEqUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2118____closed__1); +if (builtin) {res = l___regBuiltin_Lean_Meta_Grind_propagateEqUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2118_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); }l_Lean_Meta_Grind_propagateEqDown___lambda__1___closed__1 = _init_l_Lean_Meta_Grind_propagateEqDown___lambda__1___closed__1(); @@ -5602,23 +7875,72 @@ l_Lean_Meta_Grind_propagateEqDown___lambda__1___closed__2 = _init_l_Lean_Meta_Gr lean_mark_persistent(l_Lean_Meta_Grind_propagateEqDown___lambda__1___closed__2); l_Lean_Meta_Grind_propagateEqDown___lambda__1___closed__3 = _init_l_Lean_Meta_Grind_propagateEqDown___lambda__1___closed__3(); lean_mark_persistent(l_Lean_Meta_Grind_propagateEqDown___lambda__1___closed__3); -l___regBuiltin_Lean_Meta_Grind_propagateEqDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2287____closed__1 = _init_l___regBuiltin_Lean_Meta_Grind_propagateEqDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2287____closed__1(); -lean_mark_persistent(l___regBuiltin_Lean_Meta_Grind_propagateEqDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2287____closed__1); -if (builtin) {res = l___regBuiltin_Lean_Meta_Grind_propagateEqDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2287_(lean_io_mk_world()); +l___regBuiltin_Lean_Meta_Grind_propagateEqDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2347____closed__1 = _init_l___regBuiltin_Lean_Meta_Grind_propagateEqDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2347____closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Meta_Grind_propagateEqDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2347____closed__1); +if (builtin) {res = l___regBuiltin_Lean_Meta_Grind_propagateEqDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2347_(lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +}l_Lean_Meta_Grind_propagateEqMatchDown___closed__1 = _init_l_Lean_Meta_Grind_propagateEqMatchDown___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateEqMatchDown___closed__1); +l_Lean_Meta_Grind_propagateEqMatchDown___closed__2 = _init_l_Lean_Meta_Grind_propagateEqMatchDown___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateEqMatchDown___closed__2); +l___regBuiltin_Lean_Meta_Grind_propagateEqMatchDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2629____closed__1 = _init_l___regBuiltin_Lean_Meta_Grind_propagateEqMatchDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2629____closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Meta_Grind_propagateEqMatchDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2629____closed__1); +if (builtin) {res = l___regBuiltin_Lean_Meta_Grind_propagateEqMatchDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2629_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); }l_Lean_Meta_Grind_propagateHEqDown___closed__1 = _init_l_Lean_Meta_Grind_propagateHEqDown___closed__1(); lean_mark_persistent(l_Lean_Meta_Grind_propagateHEqDown___closed__1); l_Lean_Meta_Grind_propagateHEqDown___closed__2 = _init_l_Lean_Meta_Grind_propagateHEqDown___closed__2(); lean_mark_persistent(l_Lean_Meta_Grind_propagateHEqDown___closed__2); -l___regBuiltin_Lean_Meta_Grind_propagateHEqDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2546____closed__1 = _init_l___regBuiltin_Lean_Meta_Grind_propagateHEqDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2546____closed__1(); -lean_mark_persistent(l___regBuiltin_Lean_Meta_Grind_propagateHEqDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2546____closed__1); -if (builtin) {res = l___regBuiltin_Lean_Meta_Grind_propagateHEqDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2546_(lean_io_mk_world()); +l___regBuiltin_Lean_Meta_Grind_propagateHEqDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2888____closed__1 = _init_l___regBuiltin_Lean_Meta_Grind_propagateHEqDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2888____closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Meta_Grind_propagateHEqDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2888____closed__1); +if (builtin) {res = l___regBuiltin_Lean_Meta_Grind_propagateHEqDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2888_(lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +}l___regBuiltin_Lean_Meta_Grind_propagateHEqUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_3148____closed__1 = _init_l___regBuiltin_Lean_Meta_Grind_propagateHEqUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_3148____closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Meta_Grind_propagateHEqUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_3148____closed__1); +if (builtin) {res = l___regBuiltin_Lean_Meta_Grind_propagateHEqUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_3148_(lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +}l_Lean_Meta_Grind_propagateIte___lambda__1___closed__1 = _init_l_Lean_Meta_Grind_propagateIte___lambda__1___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateIte___lambda__1___closed__1); +l_Lean_Meta_Grind_propagateIte___lambda__1___closed__2 = _init_l_Lean_Meta_Grind_propagateIte___lambda__1___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateIte___lambda__1___closed__2); +l_Lean_Meta_Grind_propagateIte___lambda__1___closed__3 = _init_l_Lean_Meta_Grind_propagateIte___lambda__1___closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateIte___lambda__1___closed__3); +l_Lean_Meta_Grind_propagateIte___lambda__1___closed__4 = _init_l_Lean_Meta_Grind_propagateIte___lambda__1___closed__4(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateIte___lambda__1___closed__4); +l_Lean_Meta_Grind_propagateIte___closed__1 = _init_l_Lean_Meta_Grind_propagateIte___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateIte___closed__1); +l_Lean_Meta_Grind_propagateIte___closed__2 = _init_l_Lean_Meta_Grind_propagateIte___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateIte___closed__2); +l___regBuiltin_Lean_Meta_Grind_propagateIte_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_3530____closed__1 = _init_l___regBuiltin_Lean_Meta_Grind_propagateIte_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_3530____closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Meta_Grind_propagateIte_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_3530____closed__1); +if (builtin) {res = l___regBuiltin_Lean_Meta_Grind_propagateIte_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_3530_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -}l___regBuiltin_Lean_Meta_Grind_propagateHEqUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2806____closed__1 = _init_l___regBuiltin_Lean_Meta_Grind_propagateHEqUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2806____closed__1(); -lean_mark_persistent(l___regBuiltin_Lean_Meta_Grind_propagateHEqUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2806____closed__1); -if (builtin) {res = l___regBuiltin_Lean_Meta_Grind_propagateHEqUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2806_(lean_io_mk_world()); +}l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__1 = _init_l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__1); +l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__2 = _init_l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__2); +l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__3 = _init_l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__3); +l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__4 = _init_l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__4(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__4); +l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__5 = _init_l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__5(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__5); +l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__6 = _init_l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__6(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__6); +l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__7 = _init_l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__7(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__7); +l_Lean_Meta_Grind_propagateDIte___closed__1 = _init_l_Lean_Meta_Grind_propagateDIte___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateDIte___closed__1); +l_Lean_Meta_Grind_propagateDIte___closed__2 = _init_l_Lean_Meta_Grind_propagateDIte___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateDIte___closed__2); +l___regBuiltin_Lean_Meta_Grind_propagateDIte_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_4078____closed__1 = _init_l___regBuiltin_Lean_Meta_Grind_propagateDIte_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_4078____closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Meta_Grind_propagateDIte_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_4078____closed__1); +if (builtin) {res = l___regBuiltin_Lean_Meta_Grind_propagateDIte_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_4078_(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/Grind/Simp.c b/stage0/stdlib/Lean/Meta/Tactic/Grind/Simp.c index 3e5d9b326863..ba110f27abd9 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Grind/Simp.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Grind/Simp.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Lean.Meta.Tactic.Grind.Simp -// Imports: Init.Grind.Lemmas Lean.Meta.Tactic.Assert Lean.Meta.Tactic.Simp.Main Lean.Meta.Tactic.Grind.Util Lean.Meta.Tactic.Grind.Types Lean.Meta.Tactic.Grind.MarkNestedProofs +// Imports: Init.Grind.Lemmas Lean.Meta.Tactic.Assert Lean.Meta.Tactic.Simp.Main Lean.Meta.Tactic.Grind.Util Lean.Meta.Tactic.Grind.Types Lean.Meta.Tactic.Grind.DoNotSimp Lean.Meta.Tactic.Grind.MarkNestedProofs #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -21,6 +21,7 @@ lean_object* l_Lean_Meta_Grind_eraseIrrelevantMData___lambda__2___boxed(lean_obj lean_object* l_Lean_Meta_Grind_unfoldReducible___lambda__3___boxed(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_Grind_simp___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* l_Lean_Meta_Grind_foldProjs___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_normalizeLevels___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_simp___closed__12; lean_object* l_Lean_stringToMessageData(lean_object*); @@ -42,12 +43,16 @@ lean_object* l_Lean_Meta_transform___at_Lean_Meta_zetaReduce___spec__1(lean_obje lean_object* l_Lean_Meta_Grind_markNestedProofsImpl(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_isTracingEnabledFor___at_Lean_Meta_Grind_simp___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_simp___closed__14; +lean_object* l_Lean_Meta_Grind_foldProjs___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_Grind_eraseDoNotSimp___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofExpr(lean_object*); static lean_object* l_Lean_Meta_Grind_simp___closed__5; double l_Float_ofScientific(lean_object*, uint8_t, lean_object*); lean_object* l_Lean_Core_transform___at_Lean_Core_betaReduce___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_shareCommon(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_simp___closed__3; +static lean_object* l_Lean_Meta_Grind_simp___closed__15; lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at_Lean_Meta_Grind_simp___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_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -61,14 +66,13 @@ lean_object* l_Lean_isTracingEnabledForCore(lean_object*, lean_object*, lean_obj static double l_Lean_addTrace___at_Lean_Meta_Grind_simp___spec__3___closed__1; static lean_object* l_Lean_Meta_Grind_simp___closed__8; lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Meta_Grind_foldProjs___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_simp___closed__10; static lean_object* l_Lean_Meta_Grind_simp___closed__6; static lean_object* l_Lean_Meta_Grind_simp___closed__13; lean_object* l_Lean_Meta_Grind_canon(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_Meta_Grind_simp___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_Grind_eraseDoNotSimp___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_unfoldReducible___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Meta_Grind_foldProjs___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_simpCore(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { @@ -139,13 +143,15 @@ 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_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; 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_32 = lean_ctor_get(x_23, 0); x_33 = lean_ctor_get(x_23, 1); x_34 = lean_ctor_get(x_23, 2); x_35 = lean_ctor_get(x_23, 3); x_36 = lean_ctor_get(x_23, 5); x_37 = lean_ctor_get(x_23, 6); +x_38 = lean_ctor_get(x_23, 7); +lean_inc(x_38); lean_inc(x_37); lean_inc(x_36); lean_inc(x_35); @@ -153,55 +159,56 @@ lean_inc(x_34); lean_inc(x_33); lean_inc(x_32); lean_dec(x_23); -x_38 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_38, 0, x_32); -lean_ctor_set(x_38, 1, x_33); -lean_ctor_set(x_38, 2, x_34); -lean_ctor_set(x_38, 3, x_35); -lean_ctor_set(x_38, 4, x_21); -lean_ctor_set(x_38, 5, x_36); -lean_ctor_set(x_38, 6, x_37); -x_39 = lean_st_ref_set(x_4, x_38, x_24); -x_40 = lean_ctor_get(x_39, 1); -lean_inc(x_40); -if (lean_is_exclusive(x_39)) { - lean_ctor_release(x_39, 0); - lean_ctor_release(x_39, 1); - x_41 = x_39; +x_39 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_39, 0, x_32); +lean_ctor_set(x_39, 1, x_33); +lean_ctor_set(x_39, 2, x_34); +lean_ctor_set(x_39, 3, x_35); +lean_ctor_set(x_39, 4, x_21); +lean_ctor_set(x_39, 5, x_36); +lean_ctor_set(x_39, 6, x_37); +lean_ctor_set(x_39, 7, x_38); +x_40 = lean_st_ref_set(x_4, x_39, x_24); +x_41 = lean_ctor_get(x_40, 1); +lean_inc(x_41); +if (lean_is_exclusive(x_40)) { + lean_ctor_release(x_40, 0); + lean_ctor_release(x_40, 1); + x_42 = x_40; } else { - lean_dec_ref(x_39); - x_41 = lean_box(0); + lean_dec_ref(x_40); + x_42 = lean_box(0); } -if (lean_is_scalar(x_41)) { - x_42 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_42)) { + x_43 = lean_alloc_ctor(0, 2, 0); } else { - x_42 = x_41; + x_43 = x_42; } -lean_ctor_set(x_42, 0, x_20); -lean_ctor_set(x_42, 1, x_40); -return x_42; +lean_ctor_set(x_43, 0, x_20); +lean_ctor_set(x_43, 1, x_41); +return x_43; } } else { -uint8_t x_43; -x_43 = !lean_is_exclusive(x_17); -if (x_43 == 0) +uint8_t x_44; +x_44 = !lean_is_exclusive(x_17); +if (x_44 == 0) { return x_17; } else { -lean_object* x_44; lean_object* x_45; lean_object* x_46; -x_44 = lean_ctor_get(x_17, 0); -x_45 = lean_ctor_get(x_17, 1); +lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_45 = lean_ctor_get(x_17, 0); +x_46 = lean_ctor_get(x_17, 1); +lean_inc(x_46); lean_inc(x_45); -lean_inc(x_44); lean_dec(x_17); -x_46 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_45); -return x_46; +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; } } } @@ -744,7 +751,7 @@ static lean_object* _init_l_Lean_Meta_Grind_simp___closed__5() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_foldProjs___lambda__2___boxed), 6, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_foldProjs___lambda__3___boxed), 6, 0); return x_1; } } @@ -752,7 +759,7 @@ static lean_object* _init_l_Lean_Meta_Grind_simp___closed__6() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_foldProjs___lambda__1), 6, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_foldProjs___lambda__2), 6, 0); return x_1; } } @@ -768,7 +775,7 @@ static lean_object* _init_l_Lean_Meta_Grind_simp___closed__8() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("grind", 5, 5); +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_eraseDoNotSimp___lambda__3), 4, 0); return x_1; } } @@ -776,21 +783,37 @@ static lean_object* _init_l_Lean_Meta_Grind_simp___closed__9() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("simp", 4, 4); +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_eraseDoNotSimp___lambda__4___boxed), 4, 0); return x_1; } } static lean_object* _init_l_Lean_Meta_Grind_simp___closed__10() { _start: { +lean_object* x_1; +x_1 = lean_mk_string_unchecked("grind", 5, 5); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_simp___closed__11() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("simp", 4, 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_simp___closed__12() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_Grind_simp___closed__8; -x_2 = l_Lean_Meta_Grind_simp___closed__9; +x_1 = l_Lean_Meta_Grind_simp___closed__10; +x_2 = l_Lean_Meta_Grind_simp___closed__11; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_Grind_simp___closed__11() { +static lean_object* _init_l_Lean_Meta_Grind_simp___closed__13() { _start: { lean_object* x_1; lean_object* x_2; @@ -799,7 +822,7 @@ x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Lean_Meta_Grind_simp___closed__12() { +static lean_object* _init_l_Lean_Meta_Grind_simp___closed__14() { _start: { lean_object* x_1; @@ -807,11 +830,11 @@ x_1 = lean_mk_string_unchecked("\n===>\n", 6, 6); return x_1; } } -static lean_object* _init_l_Lean_Meta_Grind_simp___closed__13() { +static lean_object* _init_l_Lean_Meta_Grind_simp___closed__15() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_Grind_simp___closed__12; +x_1 = l_Lean_Meta_Grind_simp___closed__14; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } @@ -924,224 +947,237 @@ lean_inc(x_7); x_43 = l_Lean_Core_transform___at_Lean_Core_betaReduce___spec__1(x_40, x_42, x_32, x_7, x_8, x_41); if (lean_obj_tag(x_43) == 0) { -lean_object* x_44; lean_object* x_45; lean_object* x_46; +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; x_44 = lean_ctor_get(x_43, 0); lean_inc(x_44); x_45 = lean_ctor_get(x_43, 1); lean_inc(x_45); lean_dec(x_43); +x_46 = l_Lean_Meta_Grind_simp___closed__8; +x_47 = l_Lean_Meta_Grind_simp___closed__9; +lean_inc(x_8); +lean_inc(x_7); +x_48 = l_Lean_Core_transform___at_Lean_Core_betaReduce___spec__1(x_44, x_46, x_47, x_7, x_8, x_45); +if (lean_obj_tag(x_48) == 0) +{ +lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_49 = lean_ctor_get(x_48, 0); +lean_inc(x_49); +x_50 = lean_ctor_get(x_48, 1); +lean_inc(x_50); +lean_dec(x_48); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -x_46 = l_Lean_Meta_Grind_canon(x_44, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_45); -if (lean_obj_tag(x_46) == 0) -{ -lean_object* x_47; lean_object* x_48; lean_object* x_49; uint8_t x_50; -x_47 = lean_ctor_get(x_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_Meta_Grind_shareCommon(x_47, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_48); -x_50 = !lean_is_exclusive(x_49); -if (x_50 == 0) -{ -lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; uint8_t x_56; -x_51 = lean_ctor_get(x_49, 0); -x_52 = lean_ctor_get(x_49, 1); -x_53 = l_Lean_Meta_Grind_simp___closed__10; -x_54 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_simp___spec__2(x_53, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_52); -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) +x_51 = l_Lean_Meta_Grind_canon(x_49, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_50); +if (lean_obj_tag(x_51) == 0) +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; uint8_t 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 = l_Lean_Meta_Grind_shareCommon(x_52, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_53); +x_55 = !lean_is_exclusive(x_54); +if (x_55 == 0) +{ +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_56 = lean_ctor_get(x_54, 0); +x_57 = lean_ctor_get(x_54, 1); +x_58 = l_Lean_Meta_Grind_simp___closed__12; +x_59 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_simp___spec__2(x_58, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_57); +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_57; lean_object* x_58; lean_object* x_59; -lean_free_object(x_49); +lean_object* x_62; lean_object* x_63; lean_object* x_64; +lean_free_object(x_54); lean_free_object(x_10); lean_dec(x_12); -x_57 = lean_ctor_get(x_54, 1); -lean_inc(x_57); -lean_dec(x_54); -x_58 = lean_box(0); -x_59 = l_Lean_Meta_Grind_simp___lambda__1(x_51, x_18, x_19, x_58, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_57); +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_Meta_Grind_simp___lambda__1(x_56, x_18, x_19, x_63, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_62); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -return x_59; +return x_64; } else { -uint8_t x_60; -x_60 = !lean_is_exclusive(x_54); -if (x_60 == 0) -{ -lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; -x_61 = lean_ctor_get(x_54, 1); -x_62 = lean_ctor_get(x_54, 0); -lean_dec(x_62); -x_63 = l_Lean_MessageData_ofExpr(x_12); -x_64 = l_Lean_Meta_Grind_simp___closed__11; +uint8_t x_65; +x_65 = !lean_is_exclusive(x_59); +if (x_65 == 0) +{ +lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; +x_66 = lean_ctor_get(x_59, 1); +x_67 = lean_ctor_get(x_59, 0); +lean_dec(x_67); +x_68 = l_Lean_MessageData_ofExpr(x_12); +x_69 = l_Lean_Meta_Grind_simp___closed__13; +lean_ctor_set_tag(x_59, 7); +lean_ctor_set(x_59, 1, x_68); +lean_ctor_set(x_59, 0, x_69); +x_70 = l_Lean_Meta_Grind_simp___closed__15; lean_ctor_set_tag(x_54, 7); -lean_ctor_set(x_54, 1, x_63); -lean_ctor_set(x_54, 0, x_64); -x_65 = l_Lean_Meta_Grind_simp___closed__13; -lean_ctor_set_tag(x_49, 7); -lean_ctor_set(x_49, 1, x_65); -lean_ctor_set(x_49, 0, x_54); -lean_inc(x_51); -x_66 = l_Lean_MessageData_ofExpr(x_51); +lean_ctor_set(x_54, 1, x_70); +lean_ctor_set(x_54, 0, x_59); +lean_inc(x_56); +x_71 = l_Lean_MessageData_ofExpr(x_56); lean_ctor_set_tag(x_10, 7); -lean_ctor_set(x_10, 1, x_66); -lean_ctor_set(x_10, 0, x_49); -x_67 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_67, 0, x_10); -lean_ctor_set(x_67, 1, x_64); -x_68 = l_Lean_addTrace___at_Lean_Meta_Grind_simp___spec__3(x_53, x_67, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_61); -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_Grind_simp___lambda__1(x_51, x_18, x_19, x_69, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_70); +lean_ctor_set(x_10, 1, x_71); +lean_ctor_set(x_10, 0, x_54); +x_72 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_72, 0, x_10); +lean_ctor_set(x_72, 1, x_69); +x_73 = l_Lean_addTrace___at_Lean_Meta_Grind_simp___spec__3(x_58, x_72, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_66); +x_74 = lean_ctor_get(x_73, 0); +lean_inc(x_74); +x_75 = lean_ctor_get(x_73, 1); +lean_inc(x_75); +lean_dec(x_73); +x_76 = l_Lean_Meta_Grind_simp___lambda__1(x_56, x_18, x_19, x_74, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_75); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -lean_dec(x_69); -return x_71; +lean_dec(x_74); +return x_76; } 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; -x_72 = lean_ctor_get(x_54, 1); -lean_inc(x_72); -lean_dec(x_54); -x_73 = l_Lean_MessageData_ofExpr(x_12); -x_74 = l_Lean_Meta_Grind_simp___closed__11; -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_Meta_Grind_simp___closed__13; -lean_ctor_set_tag(x_49, 7); -lean_ctor_set(x_49, 1, x_76); -lean_ctor_set(x_49, 0, x_75); -lean_inc(x_51); -x_77 = l_Lean_MessageData_ofExpr(x_51); +lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; +x_77 = lean_ctor_get(x_59, 1); +lean_inc(x_77); +lean_dec(x_59); +x_78 = l_Lean_MessageData_ofExpr(x_12); +x_79 = l_Lean_Meta_Grind_simp___closed__13; +x_80 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_80, 0, x_79); +lean_ctor_set(x_80, 1, x_78); +x_81 = l_Lean_Meta_Grind_simp___closed__15; +lean_ctor_set_tag(x_54, 7); +lean_ctor_set(x_54, 1, x_81); +lean_ctor_set(x_54, 0, x_80); +lean_inc(x_56); +x_82 = l_Lean_MessageData_ofExpr(x_56); lean_ctor_set_tag(x_10, 7); -lean_ctor_set(x_10, 1, x_77); -lean_ctor_set(x_10, 0, x_49); -x_78 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_78, 0, x_10); -lean_ctor_set(x_78, 1, x_74); -x_79 = l_Lean_addTrace___at_Lean_Meta_Grind_simp___spec__3(x_53, x_78, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_72); -x_80 = lean_ctor_get(x_79, 0); -lean_inc(x_80); -x_81 = lean_ctor_get(x_79, 1); -lean_inc(x_81); -lean_dec(x_79); -x_82 = l_Lean_Meta_Grind_simp___lambda__1(x_51, x_18, x_19, x_80, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_81); +lean_ctor_set(x_10, 1, x_82); +lean_ctor_set(x_10, 0, x_54); +x_83 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_83, 0, x_10); +lean_ctor_set(x_83, 1, x_79); +x_84 = l_Lean_addTrace___at_Lean_Meta_Grind_simp___spec__3(x_58, x_83, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_77); +x_85 = lean_ctor_get(x_84, 0); +lean_inc(x_85); +x_86 = lean_ctor_get(x_84, 1); +lean_inc(x_86); +lean_dec(x_84); +x_87 = l_Lean_Meta_Grind_simp___lambda__1(x_56, x_18, x_19, x_85, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_86); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -lean_dec(x_80); -return x_82; +lean_dec(x_85); +return x_87; } } } else { -lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; uint8_t x_88; -x_83 = lean_ctor_get(x_49, 0); -x_84 = lean_ctor_get(x_49, 1); -lean_inc(x_84); -lean_inc(x_83); -lean_dec(x_49); -x_85 = l_Lean_Meta_Grind_simp___closed__10; -x_86 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_simp___spec__2(x_85, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_84); -x_87 = lean_ctor_get(x_86, 0); -lean_inc(x_87); -x_88 = lean_unbox(x_87); -lean_dec(x_87); -if (x_88 == 0) -{ -lean_object* x_89; lean_object* x_90; lean_object* x_91; +lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; uint8_t x_93; +x_88 = lean_ctor_get(x_54, 0); +x_89 = lean_ctor_get(x_54, 1); +lean_inc(x_89); +lean_inc(x_88); +lean_dec(x_54); +x_90 = l_Lean_Meta_Grind_simp___closed__12; +x_91 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_simp___spec__2(x_90, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_89); +x_92 = lean_ctor_get(x_91, 0); +lean_inc(x_92); +x_93 = lean_unbox(x_92); +lean_dec(x_92); +if (x_93 == 0) +{ +lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_free_object(x_10); lean_dec(x_12); -x_89 = lean_ctor_get(x_86, 1); -lean_inc(x_89); -lean_dec(x_86); -x_90 = lean_box(0); -x_91 = l_Lean_Meta_Grind_simp___lambda__1(x_83, x_18, x_19, x_90, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_89); +x_94 = lean_ctor_get(x_91, 1); +lean_inc(x_94); +lean_dec(x_91); +x_95 = lean_box(0); +x_96 = l_Lean_Meta_Grind_simp___lambda__1(x_88, x_18, x_19, x_95, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_94); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -return x_91; +return x_96; } else { -lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; -x_92 = lean_ctor_get(x_86, 1); -lean_inc(x_92); -if (lean_is_exclusive(x_86)) { - lean_ctor_release(x_86, 0); - lean_ctor_release(x_86, 1); - x_93 = x_86; +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_97 = lean_ctor_get(x_91, 1); +lean_inc(x_97); +if (lean_is_exclusive(x_91)) { + lean_ctor_release(x_91, 0); + lean_ctor_release(x_91, 1); + x_98 = x_91; } else { - lean_dec_ref(x_86); - x_93 = lean_box(0); + lean_dec_ref(x_91); + x_98 = lean_box(0); } -x_94 = l_Lean_MessageData_ofExpr(x_12); -x_95 = l_Lean_Meta_Grind_simp___closed__11; -if (lean_is_scalar(x_93)) { - x_96 = lean_alloc_ctor(7, 2, 0); +x_99 = l_Lean_MessageData_ofExpr(x_12); +x_100 = l_Lean_Meta_Grind_simp___closed__13; +if (lean_is_scalar(x_98)) { + x_101 = lean_alloc_ctor(7, 2, 0); } else { - x_96 = x_93; - lean_ctor_set_tag(x_96, 7); -} -lean_ctor_set(x_96, 0, x_95); -lean_ctor_set(x_96, 1, x_94); -x_97 = l_Lean_Meta_Grind_simp___closed__13; -x_98 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_98, 0, x_96); -lean_ctor_set(x_98, 1, x_97); -lean_inc(x_83); -x_99 = l_Lean_MessageData_ofExpr(x_83); + x_101 = x_98; + lean_ctor_set_tag(x_101, 7); +} +lean_ctor_set(x_101, 0, x_100); +lean_ctor_set(x_101, 1, x_99); +x_102 = l_Lean_Meta_Grind_simp___closed__15; +x_103 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_103, 0, x_101); +lean_ctor_set(x_103, 1, x_102); +lean_inc(x_88); +x_104 = l_Lean_MessageData_ofExpr(x_88); lean_ctor_set_tag(x_10, 7); -lean_ctor_set(x_10, 1, x_99); -lean_ctor_set(x_10, 0, x_98); -x_100 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_100, 0, x_10); -lean_ctor_set(x_100, 1, x_95); -x_101 = l_Lean_addTrace___at_Lean_Meta_Grind_simp___spec__3(x_85, x_100, x_2, x_3, x_4, x_5, x_6, x_7, x_8, 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_104 = l_Lean_Meta_Grind_simp___lambda__1(x_83, x_18, x_19, x_102, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_103); +lean_ctor_set(x_10, 1, x_104); +lean_ctor_set(x_10, 0, x_103); +x_105 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_105, 0, x_10); +lean_ctor_set(x_105, 1, x_100); +x_106 = l_Lean_addTrace___at_Lean_Meta_Grind_simp___spec__3(x_90, x_105, x_2, x_3, x_4, x_5, x_6, x_7, x_8, 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 = l_Lean_Meta_Grind_simp___lambda__1(x_88, x_18, x_19, x_107, x_2, x_3, x_4, 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); lean_dec(x_3); -lean_dec(x_102); -return x_104; +lean_dec(x_107); +return x_109; } } } else { -uint8_t x_105; +uint8_t x_110; lean_dec(x_18); lean_free_object(x_10); lean_dec(x_12); @@ -1150,29 +1186,29 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -x_105 = !lean_is_exclusive(x_46); -if (x_105 == 0) +x_110 = !lean_is_exclusive(x_51); +if (x_110 == 0) { -return x_46; +return x_51; } else { -lean_object* x_106; lean_object* x_107; lean_object* x_108; -x_106 = lean_ctor_get(x_46, 0); -x_107 = lean_ctor_get(x_46, 1); -lean_inc(x_107); -lean_inc(x_106); -lean_dec(x_46); -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_object* x_111; lean_object* x_112; lean_object* x_113; +x_111 = lean_ctor_get(x_51, 0); +x_112 = lean_ctor_get(x_51, 1); +lean_inc(x_112); +lean_inc(x_111); +lean_dec(x_51); +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 { -uint8_t x_109; +uint8_t x_114; lean_dec(x_18); lean_free_object(x_10); lean_dec(x_12); @@ -1181,29 +1217,60 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -x_109 = !lean_is_exclusive(x_43); -if (x_109 == 0) +x_114 = !lean_is_exclusive(x_48); +if (x_114 == 0) +{ +return x_48; +} +else +{ +lean_object* x_115; lean_object* x_116; lean_object* x_117; +x_115 = lean_ctor_get(x_48, 0); +x_116 = lean_ctor_get(x_48, 1); +lean_inc(x_116); +lean_inc(x_115); +lean_dec(x_48); +x_117 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_117, 0, x_115); +lean_ctor_set(x_117, 1, x_116); +return x_117; +} +} +} +else +{ +uint8_t x_118; +lean_dec(x_18); +lean_free_object(x_10); +lean_dec(x_12); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +x_118 = !lean_is_exclusive(x_43); +if (x_118 == 0) { return x_43; } else { -lean_object* x_110; lean_object* x_111; lean_object* x_112; -x_110 = lean_ctor_get(x_43, 0); -x_111 = lean_ctor_get(x_43, 1); -lean_inc(x_111); -lean_inc(x_110); +lean_object* x_119; lean_object* x_120; lean_object* x_121; +x_119 = lean_ctor_get(x_43, 0); +x_120 = lean_ctor_get(x_43, 1); +lean_inc(x_120); +lean_inc(x_119); lean_dec(x_43); -x_112 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_112, 0, x_110); -lean_ctor_set(x_112, 1, x_111); -return x_112; +x_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_113; +uint8_t x_122; lean_dec(x_18); lean_free_object(x_10); lean_dec(x_12); @@ -1212,29 +1279,29 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -x_113 = !lean_is_exclusive(x_39); -if (x_113 == 0) +x_122 = !lean_is_exclusive(x_39); +if (x_122 == 0) { return x_39; } else { -lean_object* x_114; lean_object* x_115; lean_object* x_116; -x_114 = lean_ctor_get(x_39, 0); -x_115 = lean_ctor_get(x_39, 1); -lean_inc(x_115); -lean_inc(x_114); +lean_object* x_123; lean_object* x_124; lean_object* x_125; +x_123 = lean_ctor_get(x_39, 0); +x_124 = lean_ctor_get(x_39, 1); +lean_inc(x_124); +lean_inc(x_123); lean_dec(x_39); -x_116 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_116, 0, x_114); -lean_ctor_set(x_116, 1, x_115); -return x_116; +x_125 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_125, 0, x_123); +lean_ctor_set(x_125, 1, x_124); +return x_125; } } } else { -uint8_t x_117; +uint8_t x_126; lean_dec(x_18); lean_free_object(x_10); lean_dec(x_12); @@ -1243,29 +1310,29 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -x_117 = !lean_is_exclusive(x_33); -if (x_117 == 0) +x_126 = !lean_is_exclusive(x_33); +if (x_126 == 0) { return x_33; } else { -lean_object* x_118; lean_object* x_119; lean_object* x_120; -x_118 = lean_ctor_get(x_33, 0); -x_119 = lean_ctor_get(x_33, 1); -lean_inc(x_119); -lean_inc(x_118); +lean_object* x_127; lean_object* x_128; lean_object* x_129; +x_127 = lean_ctor_get(x_33, 0); +x_128 = lean_ctor_get(x_33, 1); +lean_inc(x_128); +lean_inc(x_127); lean_dec(x_33); -x_120 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_120, 0, x_118); -lean_ctor_set(x_120, 1, x_119); -return x_120; +x_129 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_129, 0, x_127); +lean_ctor_set(x_129, 1, x_128); +return x_129; } } } else { -uint8_t x_121; +uint8_t x_130; lean_dec(x_18); lean_free_object(x_10); lean_dec(x_12); @@ -1274,29 +1341,29 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -x_121 = !lean_is_exclusive(x_28); -if (x_121 == 0) +x_130 = !lean_is_exclusive(x_28); +if (x_130 == 0) { return x_28; } else { -lean_object* x_122; lean_object* x_123; lean_object* x_124; -x_122 = lean_ctor_get(x_28, 0); -x_123 = lean_ctor_get(x_28, 1); -lean_inc(x_123); -lean_inc(x_122); +lean_object* x_131; lean_object* x_132; lean_object* x_133; +x_131 = lean_ctor_get(x_28, 0); +x_132 = lean_ctor_get(x_28, 1); +lean_inc(x_132); +lean_inc(x_131); lean_dec(x_28); -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; +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 { -uint8_t x_125; +uint8_t x_134; lean_dec(x_18); lean_free_object(x_10); lean_dec(x_12); @@ -1305,29 +1372,29 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -x_125 = !lean_is_exclusive(x_23); -if (x_125 == 0) +x_134 = !lean_is_exclusive(x_23); +if (x_134 == 0) { return x_23; } else { -lean_object* x_126; lean_object* x_127; lean_object* x_128; -x_126 = lean_ctor_get(x_23, 0); -x_127 = lean_ctor_get(x_23, 1); -lean_inc(x_127); -lean_inc(x_126); +lean_object* x_135; lean_object* x_136; lean_object* x_137; +x_135 = lean_ctor_get(x_23, 0); +x_136 = lean_ctor_get(x_23, 1); +lean_inc(x_136); +lean_inc(x_135); lean_dec(x_23); -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; +x_137 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_137, 0, x_135); +lean_ctor_set(x_137, 1, x_136); +return x_137; } } } else { -uint8_t x_129; +uint8_t x_138; lean_dec(x_18); lean_free_object(x_10); lean_dec(x_12); @@ -1336,29 +1403,29 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -x_129 = !lean_is_exclusive(x_20); -if (x_129 == 0) +x_138 = !lean_is_exclusive(x_20); +if (x_138 == 0) { return x_20; } else { -lean_object* x_130; lean_object* x_131; lean_object* x_132; -x_130 = lean_ctor_get(x_20, 0); -x_131 = lean_ctor_get(x_20, 1); -lean_inc(x_131); -lean_inc(x_130); +lean_object* x_139; lean_object* x_140; lean_object* x_141; +x_139 = lean_ctor_get(x_20, 0); +x_140 = lean_ctor_get(x_20, 1); +lean_inc(x_140); +lean_inc(x_139); lean_dec(x_20); -x_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; +x_141 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_141, 0, x_139); +lean_ctor_set(x_141, 1, x_140); +return x_141; } } } else { -uint8_t x_133; +uint8_t x_142; lean_free_object(x_10); lean_dec(x_12); lean_dec(x_8); @@ -1366,495 +1433,540 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -x_133 = !lean_is_exclusive(x_14); -if (x_133 == 0) +x_142 = !lean_is_exclusive(x_14); +if (x_142 == 0) { return x_14; } else { -lean_object* x_134; lean_object* x_135; lean_object* x_136; -x_134 = lean_ctor_get(x_14, 0); -x_135 = lean_ctor_get(x_14, 1); -lean_inc(x_135); -lean_inc(x_134); +lean_object* x_143; lean_object* x_144; lean_object* x_145; +x_143 = lean_ctor_get(x_14, 0); +x_144 = lean_ctor_get(x_14, 1); +lean_inc(x_144); +lean_inc(x_143); lean_dec(x_14); -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; +x_145 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_145, 0, x_143); +lean_ctor_set(x_145, 1, x_144); +return x_145; } } } else { -lean_object* x_137; lean_object* x_138; lean_object* x_139; -x_137 = lean_ctor_get(x_10, 0); -x_138 = lean_ctor_get(x_10, 1); -lean_inc(x_138); -lean_inc(x_137); +lean_object* x_146; lean_object* x_147; lean_object* x_148; +x_146 = lean_ctor_get(x_10, 0); +x_147 = lean_ctor_get(x_10, 1); +lean_inc(x_147); +lean_inc(x_146); lean_dec(x_10); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_3); -lean_inc(x_137); -x_139 = l_Lean_Meta_Grind_simpCore(x_137, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_138); -if (lean_obj_tag(x_139) == 0) -{ -lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; uint8_t x_144; lean_object* x_145; -x_140 = lean_ctor_get(x_139, 0); -lean_inc(x_140); -x_141 = lean_ctor_get(x_139, 1); -lean_inc(x_141); -lean_dec(x_139); -x_142 = lean_ctor_get(x_140, 0); -lean_inc(x_142); -x_143 = lean_ctor_get(x_140, 1); -lean_inc(x_143); -x_144 = lean_ctor_get_uint8(x_140, sizeof(void*)*2); -lean_dec(x_140); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -x_145 = l_Lean_Meta_Grind_abstractNestedProofs(x_142, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_141); -if (lean_obj_tag(x_145) == 0) -{ -lean_object* x_146; lean_object* x_147; lean_object* x_148; -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); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -x_148 = l_Lean_Meta_Grind_markNestedProofsImpl(x_146, x_5, x_6, x_7, x_8, x_147); +x_148 = l_Lean_Meta_Grind_simpCore(x_146, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_147); if (lean_obj_tag(x_148) == 0) { -lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; +lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; uint8_t x_153; lean_object* x_154; 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 = l_Lean_Meta_Grind_simp___closed__1; -x_152 = l_Lean_Meta_Grind_simp___closed__2; +x_151 = lean_ctor_get(x_149, 0); +lean_inc(x_151); +x_152 = lean_ctor_get(x_149, 1); +lean_inc(x_152); +x_153 = lean_ctor_get_uint8(x_149, sizeof(void*)*2); +lean_dec(x_149); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -x_153 = l_Lean_Core_transform___at_Lean_Meta_Grind_unfoldReducible___spec__1(x_149, x_151, x_152, x_5, x_6, x_7, x_8, x_150); -if (lean_obj_tag(x_153) == 0) +x_154 = l_Lean_Meta_Grind_abstractNestedProofs(x_151, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_150); +if (lean_obj_tag(x_154) == 0) { -lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; -x_154 = lean_ctor_get(x_153, 0); -lean_inc(x_154); -x_155 = lean_ctor_get(x_153, 1); +lean_object* x_155; lean_object* x_156; lean_object* x_157; +x_155 = lean_ctor_get(x_154, 0); lean_inc(x_155); -lean_dec(x_153); -x_156 = l_Lean_Meta_Grind_simp___closed__3; -x_157 = l_Lean_Meta_Grind_simp___closed__4; +x_156 = lean_ctor_get(x_154, 1); +lean_inc(x_156); +lean_dec(x_154); lean_inc(x_8); lean_inc(x_7); -x_158 = l_Lean_Core_transform___at_Lean_Core_betaReduce___spec__1(x_154, x_156, x_157, x_7, x_8, x_155); -if (lean_obj_tag(x_158) == 0) +lean_inc(x_6); +lean_inc(x_5); +x_157 = l_Lean_Meta_Grind_markNestedProofsImpl(x_155, x_5, x_6, x_7, x_8, x_156); +if (lean_obj_tag(x_157) == 0) { -lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; uint8_t x_163; lean_object* x_164; -x_159 = lean_ctor_get(x_158, 0); +lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; +x_158 = lean_ctor_get(x_157, 0); +lean_inc(x_158); +x_159 = lean_ctor_get(x_157, 1); 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_Grind_simp___closed__5; -x_162 = l_Lean_Meta_Grind_simp___closed__6; -x_163 = 0; +lean_dec(x_157); +x_160 = l_Lean_Meta_Grind_simp___closed__1; +x_161 = l_Lean_Meta_Grind_simp___closed__2; lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -x_164 = l_Lean_Meta_transform___at_Lean_Meta_zetaReduce___spec__1(x_159, x_161, x_162, x_163, x_163, x_5, x_6, x_7, x_8, x_160); -if (lean_obj_tag(x_164) == 0) -{ -lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; -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_Meta_Grind_simp___closed__7; +x_162 = l_Lean_Core_transform___at_Lean_Meta_Grind_unfoldReducible___spec__1(x_158, x_160, x_161, x_5, x_6, x_7, x_8, x_159); +if (lean_obj_tag(x_162) == 0) +{ +lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; +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_Meta_Grind_simp___closed__3; +x_166 = l_Lean_Meta_Grind_simp___closed__4; lean_inc(x_8); lean_inc(x_7); -x_168 = l_Lean_Core_transform___at_Lean_Core_betaReduce___spec__1(x_165, x_167, x_157, x_7, x_8, x_166); -if (lean_obj_tag(x_168) == 0) +x_167 = l_Lean_Core_transform___at_Lean_Core_betaReduce___spec__1(x_163, x_165, x_166, x_7, x_8, x_164); +if (lean_obj_tag(x_167) == 0) { -lean_object* x_169; lean_object* x_170; lean_object* x_171; -x_169 = lean_ctor_get(x_168, 0); +lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; uint8_t x_172; lean_object* x_173; +x_168 = lean_ctor_get(x_167, 0); +lean_inc(x_168); +x_169 = lean_ctor_get(x_167, 1); lean_inc(x_169); -x_170 = lean_ctor_get(x_168, 1); -lean_inc(x_170); -lean_dec(x_168); +lean_dec(x_167); +x_170 = l_Lean_Meta_Grind_simp___closed__5; +x_171 = l_Lean_Meta_Grind_simp___closed__6; +x_172 = 0; lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -x_171 = l_Lean_Meta_Grind_canon(x_169, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_170); -if (lean_obj_tag(x_171) == 0) -{ -lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; uint8_t x_181; -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_Grind_shareCommon(x_172, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_173); -x_175 = lean_ctor_get(x_174, 0); +x_173 = l_Lean_Meta_transform___at_Lean_Meta_zetaReduce___spec__1(x_168, x_170, x_171, x_172, x_172, x_5, x_6, x_7, x_8, x_169); +if (lean_obj_tag(x_173) == 0) +{ +lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; +x_174 = lean_ctor_get(x_173, 0); +lean_inc(x_174); +x_175 = lean_ctor_get(x_173, 1); lean_inc(x_175); -x_176 = lean_ctor_get(x_174, 1); -lean_inc(x_176); -if (lean_is_exclusive(x_174)) { - lean_ctor_release(x_174, 0); - lean_ctor_release(x_174, 1); - x_177 = x_174; -} else { - lean_dec_ref(x_174); - x_177 = lean_box(0); -} -x_178 = l_Lean_Meta_Grind_simp___closed__10; -x_179 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_simp___spec__2(x_178, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_176); -x_180 = lean_ctor_get(x_179, 0); -lean_inc(x_180); -x_181 = lean_unbox(x_180); -lean_dec(x_180); -if (x_181 == 0) -{ -lean_object* x_182; lean_object* x_183; lean_object* x_184; +lean_dec(x_173); +x_176 = l_Lean_Meta_Grind_simp___closed__7; +lean_inc(x_8); +lean_inc(x_7); +x_177 = l_Lean_Core_transform___at_Lean_Core_betaReduce___spec__1(x_174, x_176, x_166, x_7, x_8, x_175); +if (lean_obj_tag(x_177) == 0) +{ +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); +x_179 = lean_ctor_get(x_177, 1); +lean_inc(x_179); lean_dec(x_177); -lean_dec(x_137); -x_182 = lean_ctor_get(x_179, 1); -lean_inc(x_182); -lean_dec(x_179); -x_183 = lean_box(0); -x_184 = l_Lean_Meta_Grind_simp___lambda__1(x_175, x_143, x_144, x_183, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_182); +x_180 = l_Lean_Meta_Grind_simp___closed__8; +x_181 = l_Lean_Meta_Grind_simp___closed__9; +lean_inc(x_8); +lean_inc(x_7); +x_182 = l_Lean_Core_transform___at_Lean_Core_betaReduce___spec__1(x_178, x_180, x_181, x_7, x_8, x_179); +if (lean_obj_tag(x_182) == 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); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +x_185 = l_Lean_Meta_Grind_canon(x_183, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_184); +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; lean_object* x_192; lean_object* x_193; lean_object* x_194; uint8_t x_195; +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_Meta_Grind_shareCommon(x_186, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_187); +x_189 = lean_ctor_get(x_188, 0); +lean_inc(x_189); +x_190 = lean_ctor_get(x_188, 1); +lean_inc(x_190); +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); +} +x_192 = l_Lean_Meta_Grind_simp___closed__12; +x_193 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_simp___spec__2(x_192, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_190); +x_194 = lean_ctor_get(x_193, 0); +lean_inc(x_194); +x_195 = lean_unbox(x_194); +lean_dec(x_194); +if (x_195 == 0) +{ +lean_object* x_196; lean_object* x_197; lean_object* x_198; +lean_dec(x_191); +lean_dec(x_146); +x_196 = lean_ctor_get(x_193, 1); +lean_inc(x_196); +lean_dec(x_193); +x_197 = lean_box(0); +x_198 = l_Lean_Meta_Grind_simp___lambda__1(x_189, x_152, x_153, x_197, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_196); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -return x_184; +return x_198; } else { -lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; -x_185 = lean_ctor_get(x_179, 1); -lean_inc(x_185); -if (lean_is_exclusive(x_179)) { - lean_ctor_release(x_179, 0); - lean_ctor_release(x_179, 1); - x_186 = x_179; +lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; +x_199 = lean_ctor_get(x_193, 1); +lean_inc(x_199); +if (lean_is_exclusive(x_193)) { + lean_ctor_release(x_193, 0); + lean_ctor_release(x_193, 1); + x_200 = x_193; } else { - lean_dec_ref(x_179); - x_186 = lean_box(0); + lean_dec_ref(x_193); + x_200 = lean_box(0); } -x_187 = l_Lean_MessageData_ofExpr(x_137); -x_188 = l_Lean_Meta_Grind_simp___closed__11; -if (lean_is_scalar(x_186)) { - x_189 = lean_alloc_ctor(7, 2, 0); +x_201 = l_Lean_MessageData_ofExpr(x_146); +x_202 = l_Lean_Meta_Grind_simp___closed__13; +if (lean_is_scalar(x_200)) { + x_203 = lean_alloc_ctor(7, 2, 0); } else { - x_189 = x_186; - lean_ctor_set_tag(x_189, 7); -} -lean_ctor_set(x_189, 0, x_188); -lean_ctor_set(x_189, 1, x_187); -x_190 = l_Lean_Meta_Grind_simp___closed__13; -if (lean_is_scalar(x_177)) { - x_191 = lean_alloc_ctor(7, 2, 0); + x_203 = x_200; + lean_ctor_set_tag(x_203, 7); +} +lean_ctor_set(x_203, 0, x_202); +lean_ctor_set(x_203, 1, x_201); +x_204 = l_Lean_Meta_Grind_simp___closed__15; +if (lean_is_scalar(x_191)) { + x_205 = lean_alloc_ctor(7, 2, 0); } else { - x_191 = x_177; - lean_ctor_set_tag(x_191, 7); + x_205 = x_191; + lean_ctor_set_tag(x_205, 7); +} +lean_ctor_set(x_205, 0, x_203); +lean_ctor_set(x_205, 1, x_204); +lean_inc(x_189); +x_206 = l_Lean_MessageData_ofExpr(x_189); +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 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_208, 0, x_207); +lean_ctor_set(x_208, 1, x_202); +x_209 = l_Lean_addTrace___at_Lean_Meta_Grind_simp___spec__3(x_192, x_208, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_199); +x_210 = lean_ctor_get(x_209, 0); +lean_inc(x_210); +x_211 = lean_ctor_get(x_209, 1); +lean_inc(x_211); +lean_dec(x_209); +x_212 = l_Lean_Meta_Grind_simp___lambda__1(x_189, x_152, x_153, x_210, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_211); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_210); +return x_212; } -lean_ctor_set(x_191, 0, x_189); -lean_ctor_set(x_191, 1, x_190); -lean_inc(x_175); -x_192 = l_Lean_MessageData_ofExpr(x_175); -x_193 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_193, 0, x_191); -lean_ctor_set(x_193, 1, x_192); -x_194 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_194, 0, x_193); -lean_ctor_set(x_194, 1, x_188); -x_195 = l_Lean_addTrace___at_Lean_Meta_Grind_simp___spec__3(x_178, x_194, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_185); -x_196 = lean_ctor_get(x_195, 0); -lean_inc(x_196); -x_197 = lean_ctor_get(x_195, 1); -lean_inc(x_197); -lean_dec(x_195); -x_198 = l_Lean_Meta_Grind_simp___lambda__1(x_175, x_143, x_144, x_196, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_197); +} +else +{ +lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; +lean_dec(x_152); +lean_dec(x_146); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -lean_dec(x_196); -return x_198; +x_213 = lean_ctor_get(x_185, 0); +lean_inc(x_213); +x_214 = lean_ctor_get(x_185, 1); +lean_inc(x_214); +if (lean_is_exclusive(x_185)) { + lean_ctor_release(x_185, 0); + lean_ctor_release(x_185, 1); + x_215 = x_185; +} else { + lean_dec_ref(x_185); + x_215 = lean_box(0); +} +if (lean_is_scalar(x_215)) { + x_216 = lean_alloc_ctor(1, 2, 0); +} else { + x_216 = x_215; +} +lean_ctor_set(x_216, 0, x_213); +lean_ctor_set(x_216, 1, x_214); +return x_216; } } else { -lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; -lean_dec(x_143); -lean_dec(x_137); +lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; +lean_dec(x_152); +lean_dec(x_146); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -x_199 = lean_ctor_get(x_171, 0); -lean_inc(x_199); -x_200 = lean_ctor_get(x_171, 1); -lean_inc(x_200); -if (lean_is_exclusive(x_171)) { - lean_ctor_release(x_171, 0); - lean_ctor_release(x_171, 1); - x_201 = x_171; +x_217 = lean_ctor_get(x_182, 0); +lean_inc(x_217); +x_218 = lean_ctor_get(x_182, 1); +lean_inc(x_218); +if (lean_is_exclusive(x_182)) { + lean_ctor_release(x_182, 0); + lean_ctor_release(x_182, 1); + x_219 = x_182; } else { - lean_dec_ref(x_171); - x_201 = lean_box(0); + lean_dec_ref(x_182); + x_219 = lean_box(0); } -if (lean_is_scalar(x_201)) { - x_202 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_219)) { + x_220 = lean_alloc_ctor(1, 2, 0); } else { - x_202 = x_201; + x_220 = x_219; } -lean_ctor_set(x_202, 0, x_199); -lean_ctor_set(x_202, 1, x_200); -return x_202; +lean_ctor_set(x_220, 0, x_217); +lean_ctor_set(x_220, 1, x_218); +return x_220; } } else { -lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; -lean_dec(x_143); -lean_dec(x_137); +lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; +lean_dec(x_152); +lean_dec(x_146); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -x_203 = lean_ctor_get(x_168, 0); -lean_inc(x_203); -x_204 = lean_ctor_get(x_168, 1); -lean_inc(x_204); -if (lean_is_exclusive(x_168)) { - lean_ctor_release(x_168, 0); - lean_ctor_release(x_168, 1); - x_205 = x_168; +x_221 = lean_ctor_get(x_177, 0); +lean_inc(x_221); +x_222 = lean_ctor_get(x_177, 1); +lean_inc(x_222); +if (lean_is_exclusive(x_177)) { + lean_ctor_release(x_177, 0); + lean_ctor_release(x_177, 1); + x_223 = x_177; } else { - lean_dec_ref(x_168); - x_205 = lean_box(0); + lean_dec_ref(x_177); + x_223 = lean_box(0); } -if (lean_is_scalar(x_205)) { - x_206 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_223)) { + x_224 = lean_alloc_ctor(1, 2, 0); } else { - x_206 = x_205; + x_224 = x_223; } -lean_ctor_set(x_206, 0, x_203); -lean_ctor_set(x_206, 1, x_204); -return x_206; +lean_ctor_set(x_224, 0, x_221); +lean_ctor_set(x_224, 1, x_222); +return x_224; } } else { -lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; -lean_dec(x_143); -lean_dec(x_137); +lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; +lean_dec(x_152); +lean_dec(x_146); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -x_207 = lean_ctor_get(x_164, 0); -lean_inc(x_207); -x_208 = lean_ctor_get(x_164, 1); -lean_inc(x_208); -if (lean_is_exclusive(x_164)) { - lean_ctor_release(x_164, 0); - lean_ctor_release(x_164, 1); - x_209 = x_164; +x_225 = lean_ctor_get(x_173, 0); +lean_inc(x_225); +x_226 = lean_ctor_get(x_173, 1); +lean_inc(x_226); +if (lean_is_exclusive(x_173)) { + lean_ctor_release(x_173, 0); + lean_ctor_release(x_173, 1); + x_227 = x_173; } else { - lean_dec_ref(x_164); - x_209 = lean_box(0); + lean_dec_ref(x_173); + x_227 = lean_box(0); } -if (lean_is_scalar(x_209)) { - x_210 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_227)) { + x_228 = lean_alloc_ctor(1, 2, 0); } else { - x_210 = x_209; + x_228 = x_227; } -lean_ctor_set(x_210, 0, x_207); -lean_ctor_set(x_210, 1, x_208); -return x_210; +lean_ctor_set(x_228, 0, x_225); +lean_ctor_set(x_228, 1, x_226); +return x_228; } } else { -lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; -lean_dec(x_143); -lean_dec(x_137); +lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; +lean_dec(x_152); +lean_dec(x_146); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -x_211 = lean_ctor_get(x_158, 0); -lean_inc(x_211); -x_212 = lean_ctor_get(x_158, 1); -lean_inc(x_212); -if (lean_is_exclusive(x_158)) { - lean_ctor_release(x_158, 0); - lean_ctor_release(x_158, 1); - x_213 = x_158; +x_229 = lean_ctor_get(x_167, 0); +lean_inc(x_229); +x_230 = lean_ctor_get(x_167, 1); +lean_inc(x_230); +if (lean_is_exclusive(x_167)) { + lean_ctor_release(x_167, 0); + lean_ctor_release(x_167, 1); + x_231 = x_167; } else { - lean_dec_ref(x_158); - x_213 = lean_box(0); + lean_dec_ref(x_167); + x_231 = lean_box(0); } -if (lean_is_scalar(x_213)) { - x_214 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_231)) { + x_232 = lean_alloc_ctor(1, 2, 0); } else { - x_214 = x_213; + x_232 = x_231; } -lean_ctor_set(x_214, 0, x_211); -lean_ctor_set(x_214, 1, x_212); -return x_214; +lean_ctor_set(x_232, 0, x_229); +lean_ctor_set(x_232, 1, x_230); +return x_232; } } else { -lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; -lean_dec(x_143); -lean_dec(x_137); +lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; +lean_dec(x_152); +lean_dec(x_146); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -x_215 = lean_ctor_get(x_153, 0); -lean_inc(x_215); -x_216 = lean_ctor_get(x_153, 1); -lean_inc(x_216); -if (lean_is_exclusive(x_153)) { - lean_ctor_release(x_153, 0); - lean_ctor_release(x_153, 1); - x_217 = x_153; +x_233 = lean_ctor_get(x_162, 0); +lean_inc(x_233); +x_234 = lean_ctor_get(x_162, 1); +lean_inc(x_234); +if (lean_is_exclusive(x_162)) { + lean_ctor_release(x_162, 0); + lean_ctor_release(x_162, 1); + x_235 = x_162; } else { - lean_dec_ref(x_153); - x_217 = lean_box(0); + lean_dec_ref(x_162); + x_235 = lean_box(0); } -if (lean_is_scalar(x_217)) { - x_218 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_235)) { + x_236 = lean_alloc_ctor(1, 2, 0); } else { - x_218 = x_217; + x_236 = x_235; } -lean_ctor_set(x_218, 0, x_215); -lean_ctor_set(x_218, 1, x_216); -return x_218; +lean_ctor_set(x_236, 0, x_233); +lean_ctor_set(x_236, 1, x_234); +return x_236; } } else { -lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; -lean_dec(x_143); -lean_dec(x_137); +lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; +lean_dec(x_152); +lean_dec(x_146); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -x_219 = lean_ctor_get(x_148, 0); -lean_inc(x_219); -x_220 = lean_ctor_get(x_148, 1); -lean_inc(x_220); -if (lean_is_exclusive(x_148)) { - lean_ctor_release(x_148, 0); - lean_ctor_release(x_148, 1); - x_221 = x_148; +x_237 = lean_ctor_get(x_157, 0); +lean_inc(x_237); +x_238 = lean_ctor_get(x_157, 1); +lean_inc(x_238); +if (lean_is_exclusive(x_157)) { + lean_ctor_release(x_157, 0); + lean_ctor_release(x_157, 1); + x_239 = x_157; } else { - lean_dec_ref(x_148); - x_221 = lean_box(0); + lean_dec_ref(x_157); + x_239 = lean_box(0); } -if (lean_is_scalar(x_221)) { - x_222 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_239)) { + x_240 = lean_alloc_ctor(1, 2, 0); } else { - x_222 = x_221; + x_240 = x_239; } -lean_ctor_set(x_222, 0, x_219); -lean_ctor_set(x_222, 1, x_220); -return x_222; +lean_ctor_set(x_240, 0, x_237); +lean_ctor_set(x_240, 1, x_238); +return x_240; } } else { -lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; -lean_dec(x_143); -lean_dec(x_137); +lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; +lean_dec(x_152); +lean_dec(x_146); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -x_223 = lean_ctor_get(x_145, 0); -lean_inc(x_223); -x_224 = lean_ctor_get(x_145, 1); -lean_inc(x_224); -if (lean_is_exclusive(x_145)) { - lean_ctor_release(x_145, 0); - lean_ctor_release(x_145, 1); - x_225 = x_145; +x_241 = lean_ctor_get(x_154, 0); +lean_inc(x_241); +x_242 = lean_ctor_get(x_154, 1); +lean_inc(x_242); +if (lean_is_exclusive(x_154)) { + lean_ctor_release(x_154, 0); + lean_ctor_release(x_154, 1); + x_243 = x_154; } else { - lean_dec_ref(x_145); - x_225 = lean_box(0); + lean_dec_ref(x_154); + x_243 = lean_box(0); } -if (lean_is_scalar(x_225)) { - x_226 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_243)) { + x_244 = lean_alloc_ctor(1, 2, 0); } else { - x_226 = x_225; + x_244 = x_243; } -lean_ctor_set(x_226, 0, x_223); -lean_ctor_set(x_226, 1, x_224); -return x_226; +lean_ctor_set(x_244, 0, x_241); +lean_ctor_set(x_244, 1, x_242); +return x_244; } } else { -lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; -lean_dec(x_137); +lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; +lean_dec(x_146); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -x_227 = lean_ctor_get(x_139, 0); -lean_inc(x_227); -x_228 = lean_ctor_get(x_139, 1); -lean_inc(x_228); -if (lean_is_exclusive(x_139)) { - lean_ctor_release(x_139, 0); - lean_ctor_release(x_139, 1); - x_229 = x_139; +x_245 = lean_ctor_get(x_148, 0); +lean_inc(x_245); +x_246 = lean_ctor_get(x_148, 1); +lean_inc(x_246); +if (lean_is_exclusive(x_148)) { + lean_ctor_release(x_148, 0); + lean_ctor_release(x_148, 1); + x_247 = x_148; } else { - lean_dec_ref(x_139); - x_229 = lean_box(0); + lean_dec_ref(x_148); + x_247 = lean_box(0); } -if (lean_is_scalar(x_229)) { - x_230 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_247)) { + x_248 = lean_alloc_ctor(1, 2, 0); } else { - x_230 = x_229; + x_248 = x_247; } -lean_ctor_set(x_230, 0, x_227); -lean_ctor_set(x_230, 1, x_228); -return x_230; +lean_ctor_set(x_248, 0, x_245); +lean_ctor_set(x_248, 1, x_246); +return x_248; } } } @@ -1937,6 +2049,7 @@ lean_object* initialize_Lean_Meta_Tactic_Assert(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_Tactic_Simp_Main(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_Tactic_Grind_Util(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_Tactic_Grind_Types(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Meta_Tactic_Grind_DoNotSimp(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_Tactic_Grind_MarkNestedProofs(uint8_t builtin, lean_object*); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Lean_Meta_Tactic_Grind_Simp(uint8_t builtin, lean_object* w) { @@ -1958,6 +2071,9 @@ lean_dec_ref(res); res = initialize_Lean_Meta_Tactic_Grind_Types(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +res = initialize_Lean_Meta_Tactic_Grind_DoNotSimp(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); res = initialize_Lean_Meta_Tactic_Grind_MarkNestedProofs(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); @@ -1992,6 +2108,10 @@ l_Lean_Meta_Grind_simp___closed__12 = _init_l_Lean_Meta_Grind_simp___closed__12( lean_mark_persistent(l_Lean_Meta_Grind_simp___closed__12); l_Lean_Meta_Grind_simp___closed__13 = _init_l_Lean_Meta_Grind_simp___closed__13(); lean_mark_persistent(l_Lean_Meta_Grind_simp___closed__13); +l_Lean_Meta_Grind_simp___closed__14 = _init_l_Lean_Meta_Grind_simp___closed__14(); +lean_mark_persistent(l_Lean_Meta_Grind_simp___closed__14); +l_Lean_Meta_Grind_simp___closed__15 = _init_l_Lean_Meta_Grind_simp___closed__15(); +lean_mark_persistent(l_Lean_Meta_Grind_simp___closed__15); return lean_io_result_mk_ok(lean_box(0)); } #ifdef __cplusplus diff --git a/stage0/stdlib/Lean/Meta/Tactic/Grind/SimpUtil.c b/stage0/stdlib/Lean/Meta/Tactic/Grind/SimpUtil.c new file mode 100644 index 000000000000..5a2641e603b2 --- /dev/null +++ b/stage0/stdlib/Lean/Meta/Tactic/Grind/SimpUtil.c @@ -0,0 +1,586 @@ +// Lean compiler output +// Module: Lean.Meta.Tactic.Grind.SimpUtil +// Imports: Lean.Meta.Tactic.Simp.Simproc Lean.Meta.Tactic.Grind.Simp Lean.Meta.Tactic.Grind.DoNotSimp +#include +#if defined(__clang__) +#pragma clang diagnostic ignored "-Wunused-parameter" +#pragma clang diagnostic ignored "-Wunused-label" +#elif defined(__GNUC__) && !defined(__CLANG__) +#pragma GCC diagnostic ignored "-Wunused-parameter" +#pragma GCC diagnostic ignored "-Wunused-label" +#pragma GCC diagnostic ignored "-Wunused-but-set-variable" +#endif +#ifdef __cplusplus +extern "C" { +#endif +lean_object* lean_mk_empty_array_with_capacity(lean_object*); +static lean_object* l_Lean_Meta_Grind_normalizeImp___closed__8; +static lean_object* l_Lean_Meta_Grind_normalizeImp___closed__6; +static lean_object* l_Lean_Meta_Grind_normalizeImp___closed__7; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getSimprocs___rarg(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_getSimpContext___closed__1; +lean_object* l_Lean_Meta_getSimpCongrTheorems___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getSimpContext(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_normalizeImp___closed__5; +lean_object* l_Lean_Meta_Grind_addDoNotSimp(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_normalizeImp___closed__4; +lean_object* l_Lean_Meta_Simp_SimprocExtension_getSimprocs(lean_object*, lean_object*, lean_object*, lean_object*); +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_Lean_Meta_Grind_normalizeImp___closed__3; +lean_object* l_Lean_Meta_Simp_getSEvalSimprocs___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getSimpContext___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_Simp_mkContext(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getSimprocs(lean_object*, lean_object*); +extern lean_object* l_Lean_Meta_Grind_grindNormSimprocExt; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getSimprocs___rarg___boxed(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_Lean_Meta_Grind_getSimprocs___boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_getSimpContext___closed__2; +extern lean_object* l_Lean_Meta_Grind_grindNormExt; +LEAN_EXPORT lean_object* lean_grind_normalize(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_getSimprocs___rarg___closed__1; +lean_object* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); +lean_object* lean_array_mk(lean_object*); +extern lean_object* l_Lean_Meta_Simp_defaultMaxSteps; +static lean_object* l_Lean_Meta_Grind_normalizeImp___closed__1; +static lean_object* l_Lean_Meta_Grind_normalizeImp___closed__2; +static lean_object* _init_l_Lean_Meta_Grind_getSimprocs___rarg___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Meta_Grind_grindNormSimprocExt; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getSimprocs___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 = l_Lean_Meta_Grind_getSimprocs___rarg___closed__1; +x_5 = l_Lean_Meta_Simp_SimprocExtension_getSimprocs(x_4, x_1, x_2, x_3); +x_6 = !lean_is_exclusive(x_5); +if (x_6 == 0) +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_7 = lean_ctor_get(x_5, 0); +x_8 = lean_ctor_get(x_5, 1); +x_9 = l_Lean_Meta_Grind_addDoNotSimp(x_7, x_1, x_2, x_8); +if (lean_obj_tag(x_9) == 0) +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t 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); +lean_dec(x_9); +x_12 = l_Lean_Meta_Simp_getSEvalSimprocs___rarg(x_2, x_11); +x_13 = !lean_is_exclusive(x_12); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_14 = lean_ctor_get(x_12, 0); +x_15 = lean_box(0); +lean_ctor_set_tag(x_5, 1); +lean_ctor_set(x_5, 1, x_15); +lean_ctor_set(x_5, 0, x_14); +x_16 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_16, 0, x_10); +lean_ctor_set(x_16, 1, x_5); +x_17 = lean_array_mk(x_16); +lean_ctor_set(x_12, 0, x_17); +return x_12; +} +else +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_18 = lean_ctor_get(x_12, 0); +x_19 = lean_ctor_get(x_12, 1); +lean_inc(x_19); +lean_inc(x_18); +lean_dec(x_12); +x_20 = lean_box(0); +lean_ctor_set_tag(x_5, 1); +lean_ctor_set(x_5, 1, x_20); +lean_ctor_set(x_5, 0, x_18); +x_21 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_21, 0, x_10); +lean_ctor_set(x_21, 1, x_5); +x_22 = lean_array_mk(x_21); +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_19); +return x_23; +} +} +else +{ +uint8_t x_24; +lean_free_object(x_5); +x_24 = !lean_is_exclusive(x_9); +if (x_24 == 0) +{ +return x_9; +} +else +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; +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_alloc_ctor(1, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +return x_27; +} +} +} +else +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_28 = lean_ctor_get(x_5, 0); +x_29 = lean_ctor_get(x_5, 1); +lean_inc(x_29); +lean_inc(x_28); +lean_dec(x_5); +x_30 = l_Lean_Meta_Grind_addDoNotSimp(x_28, x_1, x_2, x_29); +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; lean_object* x_40; lean_object* x_41; +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_Simp_getSEvalSimprocs___rarg(x_2, 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); +if (lean_is_exclusive(x_33)) { + lean_ctor_release(x_33, 0); + lean_ctor_release(x_33, 1); + x_36 = x_33; +} else { + lean_dec_ref(x_33); + x_36 = lean_box(0); +} +x_37 = lean_box(0); +x_38 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_38, 0, x_34); +lean_ctor_set(x_38, 1, x_37); +x_39 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_39, 0, x_31); +lean_ctor_set(x_39, 1, x_38); +x_40 = lean_array_mk(x_39); +if (lean_is_scalar(x_36)) { + x_41 = lean_alloc_ctor(0, 2, 0); +} else { + x_41 = x_36; +} +lean_ctor_set(x_41, 0, x_40); +lean_ctor_set(x_41, 1, x_35); +return x_41; +} +else +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_42 = lean_ctor_get(x_30, 0); +lean_inc(x_42); +x_43 = lean_ctor_get(x_30, 1); +lean_inc(x_43); +if (lean_is_exclusive(x_30)) { + lean_ctor_release(x_30, 0); + lean_ctor_release(x_30, 1); + x_44 = x_30; +} else { + lean_dec_ref(x_30); + x_44 = lean_box(0); +} +if (lean_is_scalar(x_44)) { + x_45 = lean_alloc_ctor(1, 2, 0); +} else { + x_45 = x_44; +} +lean_ctor_set(x_45, 0, x_42); +lean_ctor_set(x_45, 1, x_43); +return x_45; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getSimprocs(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_getSimprocs___rarg___boxed), 3, 0); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getSimprocs___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Lean_Meta_Grind_getSimprocs___rarg(x_1, x_2, x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getSimprocs___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_Meta_Grind_getSimprocs(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind_getSimpContext___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Meta_Grind_grindNormExt; +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_getSimpContext___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; uint8_t x_3; uint8_t x_4; uint8_t x_5; lean_object* x_6; +x_1 = l_Lean_Meta_Simp_defaultMaxSteps; +x_2 = lean_unsigned_to_nat(2u); +x_3 = 0; +x_4 = 1; +x_5 = 0; +x_6 = lean_alloc_ctor(0, 2, 19); +lean_ctor_set(x_6, 0, x_1); +lean_ctor_set(x_6, 1, x_2); +lean_ctor_set_uint8(x_6, sizeof(void*)*2, x_3); +lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 1, x_4); +lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 2, x_3); +lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 3, x_4); +lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 4, x_4); +lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 5, x_4); +lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 6, x_5); +lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 7, x_4); +lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 8, x_4); +lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 9, x_3); +lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 10, x_4); +lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 11, x_3); +lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 12, x_4); +lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 13, x_4); +lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 14, x_3); +lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 15, x_3); +lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 16, x_3); +lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 17, x_4); +lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 18, x_4); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getSimpContext(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; uint8_t x_11; +x_6 = l_Lean_Meta_Grind_getSimpContext___closed__1; +x_7 = l_Lean_Meta_SimpExtension_getTheorems(x_6, x_3, x_4, 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 = l_Lean_Meta_getSimpCongrTheorems___rarg(x_4, x_9); +x_11 = !lean_is_exclusive(x_10); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_12 = lean_ctor_get(x_10, 0); +x_13 = lean_ctor_get(x_10, 1); +x_14 = lean_box(0); +lean_ctor_set_tag(x_10, 1); +lean_ctor_set(x_10, 1, x_14); +lean_ctor_set(x_10, 0, x_8); +x_15 = lean_array_mk(x_10); +x_16 = l_Lean_Meta_Grind_getSimpContext___closed__2; +x_17 = l_Lean_Meta_Simp_mkContext(x_16, x_15, x_12, x_1, x_2, x_3, x_4, x_13); +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; +x_18 = lean_ctor_get(x_10, 0); +x_19 = lean_ctor_get(x_10, 1); +lean_inc(x_19); +lean_inc(x_18); +lean_dec(x_10); +x_20 = lean_box(0); +x_21 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_21, 0, x_8); +lean_ctor_set(x_21, 1, x_20); +x_22 = lean_array_mk(x_21); +x_23 = l_Lean_Meta_Grind_getSimpContext___closed__2; +x_24 = l_Lean_Meta_Simp_mkContext(x_23, x_22, x_18, x_1, x_2, x_3, x_4, x_19); +return x_24; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getSimpContext___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_Grind_getSimpContext(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_Grind_normalizeImp___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_Grind_normalizeImp___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_normalizeImp___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_Meta_Grind_normalizeImp___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_normalizeImp___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_Meta_Grind_normalizeImp___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(32u); +x_2 = lean_mk_empty_array_with_capacity(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Grind_normalizeImp___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_normalizeImp___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_Meta_Grind_normalizeImp___closed__6() { +_start: +{ +size_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = 5; +x_2 = l_Lean_Meta_Grind_normalizeImp___closed__5; +x_3 = l_Lean_Meta_Grind_normalizeImp___closed__4; +x_4 = lean_unsigned_to_nat(0u); +x_5 = lean_alloc_ctor(0, 4, sizeof(size_t)*1); +lean_ctor_set(x_5, 0, x_2); +lean_ctor_set(x_5, 1, x_3); +lean_ctor_set(x_5, 2, x_4); +lean_ctor_set(x_5, 3, x_4); +lean_ctor_set_usize(x_5, 4, x_1); +return x_5; +} +} +static lean_object* _init_l_Lean_Meta_Grind_normalizeImp___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_normalizeImp___closed__2; +x_2 = l_Lean_Meta_Grind_normalizeImp___closed__6; +x_3 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_1); +lean_ctor_set(x_3, 2, x_1); +lean_ctor_set(x_3, 3, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind_normalizeImp___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_normalizeImp___closed__3; +x_2 = l_Lean_Meta_Grind_normalizeImp___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* lean_grind_normalize(lean_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 = l_Lean_Meta_Grind_getSimpContext(x_2, x_3, x_4, x_5, x_6); +x_8 = lean_ctor_get(x_7, 0); +lean_inc(x_8); +x_9 = lean_ctor_get(x_7, 1); +lean_inc(x_9); +lean_dec(x_7); +x_10 = l_Lean_Meta_Grind_getSimprocs___rarg(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; lean_object* 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 = lean_box(0); +x_14 = l_Lean_Meta_Grind_normalizeImp___closed__8; +x_15 = l_Lean_Meta_simp(x_1, x_8, x_11, x_13, x_14, x_2, x_3, x_4, x_5, x_12); +if (lean_obj_tag(x_15) == 0) +{ +lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +lean_dec(x_16); +x_18 = !lean_is_exclusive(x_15); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; +x_19 = lean_ctor_get(x_15, 0); +lean_dec(x_19); +x_20 = lean_ctor_get(x_17, 0); +lean_inc(x_20); +lean_dec(x_17); +lean_ctor_set(x_15, 0, x_20); +return x_15; +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_15, 1); +lean_inc(x_21); +lean_dec(x_15); +x_22 = lean_ctor_get(x_17, 0); +lean_inc(x_22); +lean_dec(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_21); +return x_23; +} +} +else +{ +uint8_t x_24; +x_24 = !lean_is_exclusive(x_15); +if (x_24 == 0) +{ +return x_15; +} +else +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_25 = lean_ctor_get(x_15, 0); +x_26 = lean_ctor_get(x_15, 1); +lean_inc(x_26); +lean_inc(x_25); +lean_dec(x_15); +x_27 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +return x_27; +} +} +} +else +{ +uint8_t x_28; +lean_dec(x_8); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_28 = !lean_is_exclusive(x_10); +if (x_28 == 0) +{ +return x_10; +} +else +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_29 = lean_ctor_get(x_10, 0); +x_30 = lean_ctor_get(x_10, 1); +lean_inc(x_30); +lean_inc(x_29); +lean_dec(x_10); +x_31 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_31, 0, x_29); +lean_ctor_set(x_31, 1, x_30); +return x_31; +} +} +} +} +lean_object* initialize_Lean_Meta_Tactic_Simp_Simproc(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Meta_Tactic_Grind_Simp(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Meta_Tactic_Grind_DoNotSimp(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Lean_Meta_Tactic_Grind_SimpUtil(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_Tactic_Simp_Simproc(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Lean_Meta_Tactic_Grind_Simp(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Lean_Meta_Tactic_Grind_DoNotSimp(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l_Lean_Meta_Grind_getSimprocs___rarg___closed__1 = _init_l_Lean_Meta_Grind_getSimprocs___rarg___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_getSimprocs___rarg___closed__1); +l_Lean_Meta_Grind_getSimpContext___closed__1 = _init_l_Lean_Meta_Grind_getSimpContext___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_getSimpContext___closed__1); +l_Lean_Meta_Grind_getSimpContext___closed__2 = _init_l_Lean_Meta_Grind_getSimpContext___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_getSimpContext___closed__2); +l_Lean_Meta_Grind_normalizeImp___closed__1 = _init_l_Lean_Meta_Grind_normalizeImp___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_normalizeImp___closed__1); +l_Lean_Meta_Grind_normalizeImp___closed__2 = _init_l_Lean_Meta_Grind_normalizeImp___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_normalizeImp___closed__2); +l_Lean_Meta_Grind_normalizeImp___closed__3 = _init_l_Lean_Meta_Grind_normalizeImp___closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind_normalizeImp___closed__3); +l_Lean_Meta_Grind_normalizeImp___closed__4 = _init_l_Lean_Meta_Grind_normalizeImp___closed__4(); +lean_mark_persistent(l_Lean_Meta_Grind_normalizeImp___closed__4); +l_Lean_Meta_Grind_normalizeImp___closed__5 = _init_l_Lean_Meta_Grind_normalizeImp___closed__5(); +lean_mark_persistent(l_Lean_Meta_Grind_normalizeImp___closed__5); +l_Lean_Meta_Grind_normalizeImp___closed__6 = _init_l_Lean_Meta_Grind_normalizeImp___closed__6(); +lean_mark_persistent(l_Lean_Meta_Grind_normalizeImp___closed__6); +l_Lean_Meta_Grind_normalizeImp___closed__7 = _init_l_Lean_Meta_Grind_normalizeImp___closed__7(); +lean_mark_persistent(l_Lean_Meta_Grind_normalizeImp___closed__7); +l_Lean_Meta_Grind_normalizeImp___closed__8 = _init_l_Lean_Meta_Grind_normalizeImp___closed__8(); +lean_mark_persistent(l_Lean_Meta_Grind_normalizeImp___closed__8); +return lean_io_result_mk_ok(lean_box(0)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Lean/Meta/Tactic/Grind/Split.c b/stage0/stdlib/Lean/Meta/Tactic/Grind/Split.c new file mode 100644 index 000000000000..cec2dde45cfe --- /dev/null +++ b/stage0/stdlib/Lean/Meta/Tactic/Grind/Split.c @@ -0,0 +1,5023 @@ +// Lean compiler output +// Module: Lean.Meta.Tactic.Grind.Split +// Imports: Lean.Meta.Tactic.Grind.Types Lean.Meta.Tactic.Grind.Intro Lean.Meta.Tactic.Grind.Cases Lean.Meta.Tactic.Grind.CasesMatch +#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___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__3; +lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_splitNext___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___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__3___closed__2; +lean_object* l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___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___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_introNewHyp(lean_object*, lean_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_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___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_Meta_Grind_instBEqCaseSplitStatus___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_splitNext(lean_object*, 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*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Expr_isApp(lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__7; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_selectNextSplit_x3f_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___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___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_Grind_splitNext___lambda__4___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___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_panic___at___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___spec__1___closed__5; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__3___closed__3; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___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_Grind_instBEqCaseSplitStatus; +LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Meta_Grind_splitNext___spec__1___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_selectNextSplit_x3f___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Expr_cleanupAnnotations(lean_object*); +static lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___spec__1___closed__1; +lean_object* l_Lean_stringToMessageData(lean_object*); +LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_beqCaseSplitStatus____x40_Lean_Meta_Tactic_Grind_Split___hyg_18_(uint8_t, uint8_t); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_selectNextSplit_x3f(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_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_selectNextSplit_x3f___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__7; +lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); +lean_object* l_List_appendTR___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___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_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__5___closed__3; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__9; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_selectNextSplit_x3f___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_CaseSplitStatus_noConfusion___rarg___lambda__1(lean_object*); +lean_object* lean_st_ref_take(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_CaseSplitStatus_toCtorIdx(uint8_t); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_selectNextSplit_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 lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___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_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___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_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__3; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_beqCaseSplitStatus____x40_Lean_Meta_Tactic_Grind_Split___hyg_18____boxed(lean_object*, lean_object*); +lean_object* l_Lean_MessageData_ofFormat(lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__2; +LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Meta_Grind_splitNext___spec__1(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_splitNext___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* l_Lean_Meta_Grind_checkMaxCaseSplit(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_isInductivePredicate(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__3___closed__1; +lean_object* l_Lean_Expr_appArg(lean_object*, lean_object*); +lean_object* lean_st_ref_get(lean_object*, lean_object*); +lean_object* l_Lean_Meta_Grind_casesMatch(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_selectNextSplit_x3f_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_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___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_Meta_Grind_splitNext___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___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_st_mk_ref(lean_object*, lean_object*); +lean_object* l_Lean_Expr_appFnCleanup(lean_object*, lean_object*); +static lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___spec__1___closed__2; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_CaseSplitStatus_toCtorIdx___boxed(lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__4; +lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); +lean_object* l_Lean_Meta_Grind_getGeneration(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___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__5; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__2___closed__1; +lean_object* l_Lean_Meta_Grind_updateLastTag(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 lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Lean_Meta_Grind_instInhabitedCaseSplitStatus; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__6; +lean_object* l_Lean_MessageData_ofExpr(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___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* l_Lean_Meta_Grind_isResolvedCaseSplit(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_Grind_updateLastTag___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_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__2___closed__3; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__8; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_selectNextSplit_x3f___closed__1; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__8; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__5___closed__4; +lean_object* l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___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_Meta_Grind_CaseSplitStatus_noConfusion___rarg(uint8_t, uint8_t, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t lean_nat_dec_eq(lean_object*, lean_object*); +lean_object* l_Lean_mkApp3(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_Grind_mkEqFalseProof(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t lean_nat_dec_lt(lean_object*, lean_object*); +lean_object* l_Lean_Meta_Grind_intros(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_Grind_updateLastTag___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Expr_isConstOf(lean_object*, lean_object*); +lean_object* lean_panic_fn(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__5___closed__2; +lean_object* l_Lean_Expr_getAppFn(lean_object*); +lean_object* l_Lean_Meta_Grind_mkEqTrueProof(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_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___spec__1___closed__4; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__5___closed__5; +lean_object* l_Lean_Meta_instantiateMVarsIfMVarApp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__2___closed__2; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__2___closed__4; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_CaseSplitStatus_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); +lean_object* l_List_reverse___rarg(lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__4; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_mkEM(lean_object*); +lean_object* l_Lean_Meta_Grind_isEqFalse(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_Grind_Split_0__Lean_Meta_Grind_selectNextSplit_x3f___lambda__2___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_splitNext___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_splitNext___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_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__2; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor(lean_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_Grind_Split_0__Lean_Meta_Grind_selectNextSplit_x3f_go___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_instInhabitedOfMonad___rarg(lean_object*, lean_object*); +lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_CaseSplitStatus_noConfusion(lean_object*); +lean_object* l_Lean_Meta_Grind_isEqTrue(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_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__5___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_splitNext___lambda__4___closed__3; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_splitNext___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_Grind_cases(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___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_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__1; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__9; +lean_object* lean_nat_add(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_splitNext___lambda__4___closed__2; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___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___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__6; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__2___closed__5; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_splitNext___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_splitNext___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_CaseSplitStatus_noConfusion___rarg___lambda__1___boxed(lean_object*); +lean_object* l___private_Init_Data_Repr_0__Nat_reprFast(lean_object*); +static lean_object* l_Lean_Meta_Grind_CaseSplitStatus_noConfusion___rarg___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___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* l_Lean_Meta_Grind_isInconsistent(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_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__5; +lean_object* l_ReaderT_instMonad___rarg(lean_object*); +static lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___spec__1___closed__3; +lean_object* l_Lean_Meta_isMatcherApp___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate___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_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus(lean_object*, 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_Grind_CaseSplitStatus_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_Meta_Grind_CaseSplitStatus_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_Meta_Grind_CaseSplitStatus_toCtorIdx(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_CaseSplitStatus_noConfusion___rarg___lambda__1(lean_object* x_1) { +_start: +{ +lean_inc(x_1); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_CaseSplitStatus_noConfusion___rarg___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_CaseSplitStatus_noConfusion___rarg___lambda__1___boxed), 1, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_CaseSplitStatus_noConfusion___rarg(uint8_t x_1, uint8_t x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Lean_Meta_Grind_CaseSplitStatus_noConfusion___rarg___closed__1; +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_CaseSplitStatus_noConfusion(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_CaseSplitStatus_noConfusion___rarg___boxed), 3, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_CaseSplitStatus_noConfusion___rarg___lambda__1___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Lean_Meta_Grind_CaseSplitStatus_noConfusion___rarg___lambda__1(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_CaseSplitStatus_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_Meta_Grind_CaseSplitStatus_noConfusion___rarg(x_4, x_5, x_3); +return x_6; +} +} +static uint8_t _init_l_Lean_Meta_Grind_instInhabitedCaseSplitStatus() { +_start: +{ +uint8_t x_1; +x_1 = 0; +return x_1; +} +} +LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_beqCaseSplitStatus____x40_Lean_Meta_Tactic_Grind_Split___hyg_18_(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_Meta_Grind_CaseSplitStatus_toCtorIdx(x_1); +x_4 = l_Lean_Meta_Grind_CaseSplitStatus_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_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_beqCaseSplitStatus____x40_Lean_Meta_Tactic_Grind_Split___hyg_18____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_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_beqCaseSplitStatus____x40_Lean_Meta_Tactic_Grind_Split___hyg_18_(x_3, x_4); +x_6 = lean_box(x_5); +return x_6; +} +} +static lean_object* _init_l_Lean_Meta_Grind_instBEqCaseSplitStatus___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_beqCaseSplitStatus____x40_Lean_Meta_Tactic_Grind_Split___hyg_18____boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_instBEqCaseSplitStatus() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Meta_Grind_instBEqCaseSplitStatus___closed__1; +return x_1; +} +} +static lean_object* _init_l_panic___at___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___spec__1___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_instMonadMetaM; +x_2 = l_ReaderT_instMonad___rarg(x_1); +return x_2; +} +} +static lean_object* _init_l_panic___at___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___spec__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_panic___at___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___spec__1___closed__1; +x_2 = l_ReaderT_instMonad___rarg(x_1); +return x_2; +} +} +static lean_object* _init_l_panic___at___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___spec__1___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_panic___at___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___spec__1___closed__2; +x_2 = l_ReaderT_instMonad___rarg(x_1); +return x_2; +} +} +static lean_object* _init_l_panic___at___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___spec__1___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_panic___at___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___spec__1___closed__3; +x_2 = l_ReaderT_instMonad___rarg(x_1); +return x_2; +} +} +static lean_object* _init_l_panic___at___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___spec__1___closed__5() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_panic___at___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___spec__1___closed__4; +x_2 = l_Lean_Meta_Grind_instInhabitedCaseSplitStatus; +x_3 = lean_box(x_2); +x_4 = l_instInhabitedOfMonad___rarg(x_1, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_11 = l_panic___at___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___spec__1___closed__5; +x_12 = lean_panic_fn(x_11, x_1); +x_13 = lean_apply_9(x_12, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_13; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___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_inc(x_1); +x_11 = l_Lean_Meta_Grind_isEqTrue(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_11) == 0) +{ +lean_object* x_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_14; lean_object* x_15; +x_14 = lean_ctor_get(x_11, 1); +lean_inc(x_14); +lean_dec(x_11); +x_15 = l_Lean_Meta_Grind_isEqFalse(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_14); +if (lean_obj_tag(x_15) == 0) +{ +lean_object* x_16; uint8_t x_17; +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +x_17 = lean_unbox(x_16); +lean_dec(x_16); +if (x_17 == 0) +{ +uint8_t x_18; +x_18 = !lean_is_exclusive(x_15); +if (x_18 == 0) +{ +lean_object* x_19; uint8_t x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_15, 0); +lean_dec(x_19); +x_20 = 2; +x_21 = lean_box(x_20); +lean_ctor_set(x_15, 0, x_21); +return x_15; +} +else +{ +lean_object* x_22; uint8_t x_23; lean_object* x_24; lean_object* x_25; +x_22 = lean_ctor_get(x_15, 1); +lean_inc(x_22); +lean_dec(x_15); +x_23 = 2; +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_15); +if (x_26 == 0) +{ +lean_object* x_27; uint8_t x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_15, 0); +lean_dec(x_27); +x_28 = 0; +x_29 = lean_box(x_28); +lean_ctor_set(x_15, 0, x_29); +return x_15; +} +else +{ +lean_object* x_30; uint8_t x_31; lean_object* x_32; lean_object* x_33; +x_30 = lean_ctor_get(x_15, 1); +lean_inc(x_30); +lean_dec(x_15); +x_31 = 0; +x_32 = lean_box(x_31); +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_30); +return x_33; +} +} +} +else +{ +uint8_t x_34; +x_34 = !lean_is_exclusive(x_15); +if (x_34 == 0) +{ +return x_15; +} +else +{ +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_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_38; +lean_dec(x_1); +x_38 = !lean_is_exclusive(x_11); +if (x_38 == 0) +{ +lean_object* x_39; uint8_t x_40; lean_object* x_41; +x_39 = lean_ctor_get(x_11, 0); +lean_dec(x_39); +x_40 = 0; +x_41 = lean_box(x_40); +lean_ctor_set(x_11, 0, x_41); +return x_11; +} +else +{ +lean_object* x_42; uint8_t x_43; lean_object* x_44; lean_object* x_45; +x_42 = lean_ctor_get(x_11, 1); +lean_inc(x_42); +lean_dec(x_11); +x_43 = 0; +x_44 = lean_box(x_43); +x_45 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_45, 0, x_44); +lean_ctor_set(x_45, 1, x_42); +return x_45; +} +} +} +else +{ +uint8_t x_46; +lean_dec(x_1); +x_46 = !lean_is_exclusive(x_11); +if (x_46 == 0) +{ +return x_11; +} +else +{ +lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_47 = lean_ctor_get(x_11, 0); +x_48 = lean_ctor_get(x_11, 1); +lean_inc(x_48); +lean_inc(x_47); +lean_dec(x_11); +x_49 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_49, 0, x_47); +lean_ctor_set(x_49, 1, x_48); +return x_49; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +lean_inc(x_1); +x_13 = l_Lean_Meta_Grind_isEqTrue(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_13) == 0) +{ +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_16; lean_object* x_17; +x_16 = lean_ctor_get(x_13, 1); +lean_inc(x_16); +lean_dec(x_13); +x_17 = l_Lean_Meta_Grind_isEqFalse(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_16); +if (lean_obj_tag(x_17) == 0) +{ +lean_object* x_18; uint8_t x_19; +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +x_19 = lean_unbox(x_18); +lean_dec(x_18); +if (x_19 == 0) +{ +uint8_t x_20; +lean_dec(x_3); +lean_dec(x_2); +x_20 = !lean_is_exclusive(x_17); +if (x_20 == 0) +{ +lean_object* x_21; uint8_t x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_17, 0); +lean_dec(x_21); +x_22 = 1; +x_23 = lean_box(x_22); +lean_ctor_set(x_17, 0, x_23); +return x_17; +} +else +{ +lean_object* x_24; uint8_t x_25; lean_object* x_26; lean_object* x_27; +x_24 = lean_ctor_get(x_17, 1); +lean_inc(x_24); +lean_dec(x_17); +x_25 = 1; +x_26 = lean_box(x_25); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_24); +return x_27; +} +} +else +{ +lean_object* x_28; lean_object* x_29; +x_28 = lean_ctor_get(x_17, 1); +lean_inc(x_28); +lean_dec(x_17); +x_29 = l_Lean_Meta_Grind_isEqFalse(x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_28); +if (lean_obj_tag(x_29) == 0) +{ +lean_object* x_30; uint8_t x_31; +x_30 = lean_ctor_get(x_29, 0); +lean_inc(x_30); +x_31 = lean_unbox(x_30); +lean_dec(x_30); +if (x_31 == 0) +{ +lean_object* x_32; lean_object* x_33; +x_32 = lean_ctor_get(x_29, 1); +lean_inc(x_32); +lean_dec(x_29); +x_33 = l_Lean_Meta_Grind_isEqFalse(x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_32); +if (lean_obj_tag(x_33) == 0) +{ +lean_object* x_34; uint8_t x_35; +x_34 = lean_ctor_get(x_33, 0); +lean_inc(x_34); +x_35 = lean_unbox(x_34); +lean_dec(x_34); +if (x_35 == 0) +{ +uint8_t x_36; +x_36 = !lean_is_exclusive(x_33); +if (x_36 == 0) +{ +lean_object* x_37; uint8_t x_38; lean_object* x_39; +x_37 = lean_ctor_get(x_33, 0); +lean_dec(x_37); +x_38 = 2; +x_39 = lean_box(x_38); +lean_ctor_set(x_33, 0, x_39); +return x_33; +} +else +{ +lean_object* x_40; uint8_t x_41; lean_object* x_42; lean_object* x_43; +x_40 = lean_ctor_get(x_33, 1); +lean_inc(x_40); +lean_dec(x_33); +x_41 = 2; +x_42 = lean_box(x_41); +x_43 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_43, 0, x_42); +lean_ctor_set(x_43, 1, x_40); +return x_43; +} +} +else +{ +uint8_t x_44; +x_44 = !lean_is_exclusive(x_33); +if (x_44 == 0) +{ +lean_object* x_45; uint8_t x_46; lean_object* x_47; +x_45 = lean_ctor_get(x_33, 0); +lean_dec(x_45); +x_46 = 0; +x_47 = lean_box(x_46); +lean_ctor_set(x_33, 0, x_47); +return x_33; +} +else +{ +lean_object* x_48; uint8_t x_49; lean_object* x_50; lean_object* x_51; +x_48 = lean_ctor_get(x_33, 1); +lean_inc(x_48); +lean_dec(x_33); +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); +return x_51; +} +} +} +else +{ +uint8_t x_52; +x_52 = !lean_is_exclusive(x_33); +if (x_52 == 0) +{ +return x_33; +} +else +{ +lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_53 = lean_ctor_get(x_33, 0); +x_54 = lean_ctor_get(x_33, 1); +lean_inc(x_54); +lean_inc(x_53); +lean_dec(x_33); +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_56; +lean_dec(x_3); +x_56 = !lean_is_exclusive(x_29); +if (x_56 == 0) +{ +lean_object* x_57; uint8_t x_58; lean_object* x_59; +x_57 = lean_ctor_get(x_29, 0); +lean_dec(x_57); +x_58 = 0; +x_59 = lean_box(x_58); +lean_ctor_set(x_29, 0, x_59); +return x_29; +} +else +{ +lean_object* x_60; uint8_t x_61; lean_object* x_62; lean_object* x_63; +x_60 = lean_ctor_get(x_29, 1); +lean_inc(x_60); +lean_dec(x_29); +x_61 = 0; +x_62 = lean_box(x_61); +x_63 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_63, 0, x_62); +lean_ctor_set(x_63, 1, x_60); +return x_63; +} +} +} +else +{ +uint8_t x_64; +lean_dec(x_3); +x_64 = !lean_is_exclusive(x_29); +if (x_64 == 0) +{ +return x_29; +} +else +{ +lean_object* x_65; lean_object* x_66; lean_object* x_67; +x_65 = lean_ctor_get(x_29, 0); +x_66 = lean_ctor_get(x_29, 1); +lean_inc(x_66); +lean_inc(x_65); +lean_dec(x_29); +x_67 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_67, 0, x_65); +lean_ctor_set(x_67, 1, x_66); +return x_67; +} +} +} +} +else +{ +uint8_t x_68; +lean_dec(x_3); +lean_dec(x_2); +x_68 = !lean_is_exclusive(x_17); +if (x_68 == 0) +{ +return x_17; +} +else +{ +lean_object* x_69; lean_object* x_70; lean_object* x_71; +x_69 = lean_ctor_get(x_17, 0); +x_70 = lean_ctor_get(x_17, 1); +lean_inc(x_70); +lean_inc(x_69); +lean_dec(x_17); +x_71 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_71, 0, x_69); +lean_ctor_set(x_71, 1, x_70); +return x_71; +} +} +} +else +{ +uint8_t x_72; +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_72 = !lean_is_exclusive(x_13); +if (x_72 == 0) +{ +lean_object* x_73; uint8_t x_74; lean_object* x_75; +x_73 = lean_ctor_get(x_13, 0); +lean_dec(x_73); +x_74 = 0; +x_75 = lean_box(x_74); +lean_ctor_set(x_13, 0, x_75); +return x_13; +} +else +{ +lean_object* x_76; uint8_t x_77; lean_object* x_78; lean_object* x_79; +x_76 = lean_ctor_get(x_13, 1); +lean_inc(x_76); +lean_dec(x_13); +x_77 = 0; +x_78 = lean_box(x_77); +x_79 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_79, 0, x_78); +lean_ctor_set(x_79, 1, x_76); +return x_79; +} +} +} +else +{ +uint8_t x_80; +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_80 = !lean_is_exclusive(x_13); +if (x_80 == 0) +{ +return x_13; +} +else +{ +lean_object* x_81; lean_object* x_82; lean_object* x_83; +x_81 = lean_ctor_get(x_13, 0); +x_82 = lean_ctor_get(x_13, 1); +lean_inc(x_82); +lean_inc(x_81); +lean_dec(x_13); +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_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___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; +lean_inc(x_1); +x_13 = l_Lean_Meta_Grind_isEqTrue(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_13) == 0) +{ +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_16; lean_object* x_17; +lean_dec(x_3); +lean_dec(x_2); +x_16 = lean_ctor_get(x_13, 1); +lean_inc(x_16); +lean_dec(x_13); +x_17 = l_Lean_Meta_Grind_isEqFalse(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_16); +if (lean_obj_tag(x_17) == 0) +{ +lean_object* x_18; uint8_t x_19; +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +x_19 = lean_unbox(x_18); +lean_dec(x_18); +if (x_19 == 0) +{ +uint8_t x_20; +x_20 = !lean_is_exclusive(x_17); +if (x_20 == 0) +{ +lean_object* x_21; uint8_t x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_17, 0); +lean_dec(x_21); +x_22 = 1; +x_23 = lean_box(x_22); +lean_ctor_set(x_17, 0, x_23); +return x_17; +} +else +{ +lean_object* x_24; uint8_t x_25; lean_object* x_26; lean_object* x_27; +x_24 = lean_ctor_get(x_17, 1); +lean_inc(x_24); +lean_dec(x_17); +x_25 = 1; +x_26 = lean_box(x_25); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_24); +return x_27; +} +} +else +{ +uint8_t x_28; +x_28 = !lean_is_exclusive(x_17); +if (x_28 == 0) +{ +lean_object* x_29; uint8_t x_30; lean_object* x_31; +x_29 = lean_ctor_get(x_17, 0); +lean_dec(x_29); +x_30 = 0; +x_31 = lean_box(x_30); +lean_ctor_set(x_17, 0, x_31); +return x_17; +} +else +{ +lean_object* x_32; uint8_t x_33; lean_object* x_34; lean_object* x_35; +x_32 = lean_ctor_get(x_17, 1); +lean_inc(x_32); +lean_dec(x_17); +x_33 = 0; +x_34 = lean_box(x_33); +x_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_34); +lean_ctor_set(x_35, 1, x_32); +return x_35; +} +} +} +else +{ +uint8_t x_36; +x_36 = !lean_is_exclusive(x_17); +if (x_36 == 0) +{ +return x_17; +} +else +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_37 = lean_ctor_get(x_17, 0); +x_38 = lean_ctor_get(x_17, 1); +lean_inc(x_38); +lean_inc(x_37); +lean_dec(x_17); +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_dec(x_1); +x_40 = lean_ctor_get(x_13, 1); +lean_inc(x_40); +lean_dec(x_13); +x_41 = l_Lean_Meta_Grind_isEqTrue(x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_40); +if (lean_obj_tag(x_41) == 0) +{ +lean_object* x_42; uint8_t x_43; +x_42 = lean_ctor_get(x_41, 0); +lean_inc(x_42); +x_43 = lean_unbox(x_42); +lean_dec(x_42); +if (x_43 == 0) +{ +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 = l_Lean_Meta_Grind_isEqTrue(x_3, x_4, 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; uint8_t x_47; +x_46 = lean_ctor_get(x_45, 0); +lean_inc(x_46); +x_47 = lean_unbox(x_46); +lean_dec(x_46); +if (x_47 == 0) +{ +uint8_t x_48; +x_48 = !lean_is_exclusive(x_45); +if (x_48 == 0) +{ +lean_object* x_49; uint8_t x_50; lean_object* x_51; +x_49 = lean_ctor_get(x_45, 0); +lean_dec(x_49); +x_50 = 2; +x_51 = lean_box(x_50); +lean_ctor_set(x_45, 0, x_51); +return x_45; +} +else +{ +lean_object* x_52; uint8_t x_53; lean_object* x_54; lean_object* x_55; +x_52 = lean_ctor_get(x_45, 1); +lean_inc(x_52); +lean_dec(x_45); +x_53 = 2; +x_54 = lean_box(x_53); +x_55 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_55, 0, x_54); +lean_ctor_set(x_55, 1, x_52); +return x_55; +} +} +else +{ +uint8_t x_56; +x_56 = !lean_is_exclusive(x_45); +if (x_56 == 0) +{ +lean_object* x_57; uint8_t x_58; lean_object* x_59; +x_57 = lean_ctor_get(x_45, 0); +lean_dec(x_57); +x_58 = 0; +x_59 = lean_box(x_58); +lean_ctor_set(x_45, 0, x_59); +return x_45; +} +else +{ +lean_object* x_60; uint8_t x_61; lean_object* x_62; lean_object* x_63; +x_60 = lean_ctor_get(x_45, 1); +lean_inc(x_60); +lean_dec(x_45); +x_61 = 0; +x_62 = lean_box(x_61); +x_63 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_63, 0, x_62); +lean_ctor_set(x_63, 1, x_60); +return x_63; +} +} +} +else +{ +uint8_t x_64; +x_64 = !lean_is_exclusive(x_45); +if (x_64 == 0) +{ +return x_45; +} +else +{ +lean_object* x_65; lean_object* x_66; lean_object* x_67; +x_65 = lean_ctor_get(x_45, 0); +x_66 = lean_ctor_get(x_45, 1); +lean_inc(x_66); +lean_inc(x_65); +lean_dec(x_45); +x_67 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_67, 0, x_65); +lean_ctor_set(x_67, 1, x_66); +return x_67; +} +} +} +else +{ +uint8_t x_68; +lean_dec(x_3); +x_68 = !lean_is_exclusive(x_41); +if (x_68 == 0) +{ +lean_object* x_69; uint8_t x_70; lean_object* x_71; +x_69 = lean_ctor_get(x_41, 0); +lean_dec(x_69); +x_70 = 0; +x_71 = lean_box(x_70); +lean_ctor_set(x_41, 0, x_71); +return x_41; +} +else +{ +lean_object* x_72; uint8_t x_73; lean_object* x_74; lean_object* x_75; +x_72 = lean_ctor_get(x_41, 1); +lean_inc(x_72); +lean_dec(x_41); +x_73 = 0; +x_74 = lean_box(x_73); +x_75 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_75, 0, x_74); +lean_ctor_set(x_75, 1, x_72); +return x_75; +} +} +} +else +{ +uint8_t x_76; +lean_dec(x_3); +x_76 = !lean_is_exclusive(x_41); +if (x_76 == 0) +{ +return x_41; +} +else +{ +lean_object* x_77; lean_object* x_78; lean_object* x_79; +x_77 = lean_ctor_get(x_41, 0); +x_78 = lean_ctor_get(x_41, 1); +lean_inc(x_78); +lean_inc(x_77); +lean_dec(x_41); +x_79 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_79, 0, x_77); +lean_ctor_set(x_79, 1, x_78); +return x_79; +} +} +} +} +else +{ +uint8_t x_80; +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_80 = !lean_is_exclusive(x_13); +if (x_80 == 0) +{ +return x_13; +} +else +{ +lean_object* x_81; lean_object* x_82; lean_object* x_83; +x_81 = lean_ctor_get(x_13, 0); +x_82 = lean_ctor_get(x_13, 1); +lean_inc(x_82); +lean_inc(x_81); +lean_dec(x_13); +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_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___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_11; lean_object* x_12; lean_object* x_13; +x_11 = 1; +x_12 = lean_box(x_11); +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_10); +return x_13; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__5___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.Meta.Tactic.Grind.Split", 28, 28); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__5___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("_private.Lean.Meta.Tactic.Grind.Split.0.Lean.Meta.Grind.checkCaseSplitStatus", 76, 76); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__5___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("unreachable code has been reached", 33, 33); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__5___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__5___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__5___closed__2; +x_3 = lean_unsigned_to_nat(58u); +x_4 = lean_unsigned_to_nat(43u); +x_5 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__5___closed__3; +x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); +return x_6; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__5___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__4___boxed), 10, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; +x_12 = l_Lean_Expr_getAppFn(x_1); +if (lean_obj_tag(x_12) == 4) +{ +lean_object* x_13; lean_object* x_14; +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +lean_dec(x_12); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +x_14 = l_Lean_Meta_isInductivePredicate(x_13, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_14) == 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_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +lean_dec(x_1); +x_17 = lean_ctor_get(x_14, 1); +lean_inc(x_17); +lean_dec(x_14); +x_18 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__5___closed__5; +x_19 = lean_box(0); +x_20 = lean_apply_10(x_18, x_19, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_17); +return x_20; +} +else +{ +lean_object* x_21; lean_object* x_22; +x_21 = lean_ctor_get(x_14, 1); +lean_inc(x_21); +lean_dec(x_14); +x_22 = l_Lean_Meta_Grind_isEqTrue(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_21); +if (lean_obj_tag(x_22) == 0) +{ +lean_object* x_23; uint8_t x_24; +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +x_24 = lean_unbox(x_23); +lean_dec(x_23); +if (x_24 == 0) +{ +lean_object* x_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___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__5___closed__5; +x_27 = lean_box(0); +x_28 = lean_apply_10(x_26, x_27, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_25); +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); +x_29 = !lean_is_exclusive(x_22); +if (x_29 == 0) +{ +lean_object* x_30; uint8_t x_31; lean_object* x_32; +x_30 = lean_ctor_get(x_22, 0); +lean_dec(x_30); +x_31 = 2; +x_32 = lean_box(x_31); +lean_ctor_set(x_22, 0, x_32); +return x_22; +} +else +{ +lean_object* x_33; uint8_t x_34; lean_object* x_35; lean_object* x_36; +x_33 = lean_ctor_get(x_22, 1); +lean_inc(x_33); +lean_dec(x_22); +x_34 = 2; +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_33); +return x_36; +} +} +} +else +{ +uint8_t x_37; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_37 = !lean_is_exclusive(x_22); +if (x_37 == 0) +{ +return x_22; +} +else +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_38 = lean_ctor_get(x_22, 0); +x_39 = lean_ctor_get(x_22, 1); +lean_inc(x_39); +lean_inc(x_38); +lean_dec(x_22); +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_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_41 = !lean_is_exclusive(x_14); +if (x_41 == 0) +{ +return x_14; +} +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_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; +} +} +} +else +{ +lean_object* x_45; lean_object* x_46; +lean_dec(x_12); +lean_dec(x_1); +x_45 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__5___closed__4; +x_46 = l_panic___at___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___spec__1(x_45, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +return x_46; +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___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; lean_object* x_13; uint8_t x_14; +x_12 = l_Lean_Meta_isMatcherApp___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate___spec__1(x_1, 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_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_ctor_get(x_12, 1); +lean_inc(x_15); +lean_dec(x_12); +x_16 = lean_box(0); +x_17 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__5(x_1, x_16, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_15); +return x_17; +} +else +{ +uint8_t x_18; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_18 = !lean_is_exclusive(x_12); +if (x_18 == 0) +{ +lean_object* x_19; uint8_t x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_12, 0); +lean_dec(x_19); +x_20 = 2; +x_21 = lean_box(x_20); +lean_ctor_set(x_12, 0, x_21); +return x_12; +} +else +{ +lean_object* x_22; uint8_t x_23; lean_object* x_24; lean_object* x_25; +x_22 = lean_ctor_get(x_12, 1); +lean_inc(x_22); +lean_dec(x_12); +x_23 = 2; +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; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___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) { +_start: +{ +uint8_t x_11; lean_object* x_12; lean_object* x_13; +x_11 = 0; +x_12 = lean_box(x_11); +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_10); +return x_13; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("grind", 5, 5); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("debug", 5, 5); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("split", 5, 5); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__2; +x_3 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__3; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__7___boxed), 10, 0); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("split resolved: ", 16, 16); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__6; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__8() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("", 0, 0); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__8; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_12 = l_Lean_Meta_Grind_isResolvedCaseSplit(x_1, 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_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_ctor_get(x_12, 1); +lean_inc(x_15); +lean_dec(x_12); +x_16 = lean_box(0); +x_17 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__6(x_1, x_16, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_15); +return x_17; +} +else +{ +uint8_t x_18; +x_18 = !lean_is_exclusive(x_12); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; +x_19 = lean_ctor_get(x_12, 1); +x_20 = lean_ctor_get(x_12, 0); +lean_dec(x_20); +x_21 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__4; +x_22 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_21, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_19); +x_23 = !lean_is_exclusive(x_22); +if (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_22, 0); +x_25 = lean_ctor_get(x_22, 1); +x_26 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__5; +x_27 = lean_unbox(x_24); +lean_dec(x_24); +if (x_27 == 0) +{ +lean_object* x_28; lean_object* x_29; +lean_free_object(x_22); +lean_free_object(x_12); +lean_dec(x_1); +x_28 = lean_box(0); +x_29 = lean_apply_10(x_26, x_28, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_25); +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; +x_30 = l_Lean_MessageData_ofExpr(x_1); +x_31 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__7; +lean_ctor_set_tag(x_22, 7); +lean_ctor_set(x_22, 1, x_30); +lean_ctor_set(x_22, 0, x_31); +x_32 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__9; +lean_ctor_set_tag(x_12, 7); +lean_ctor_set(x_12, 1, x_32); +lean_ctor_set(x_12, 0, x_22); +x_33 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_21, x_12, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_25); +x_34 = lean_ctor_get(x_33, 0); +lean_inc(x_34); +x_35 = lean_ctor_get(x_33, 1); +lean_inc(x_35); +lean_dec(x_33); +x_36 = lean_apply_10(x_26, x_34, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_35); +return x_36; +} +} +else +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; uint8_t x_40; +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 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__5; +x_40 = lean_unbox(x_37); +lean_dec(x_37); +if (x_40 == 0) +{ +lean_object* x_41; lean_object* x_42; +lean_free_object(x_12); +lean_dec(x_1); +x_41 = lean_box(0); +x_42 = lean_apply_10(x_39, x_41, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_38); +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; +x_43 = l_Lean_MessageData_ofExpr(x_1); +x_44 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__7; +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___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__9; +lean_ctor_set_tag(x_12, 7); +lean_ctor_set(x_12, 1, x_46); +lean_ctor_set(x_12, 0, x_45); +x_47 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_21, x_12, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_38); +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_apply_10(x_39, x_48, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, 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; uint8_t x_58; +x_51 = lean_ctor_get(x_12, 1); +lean_inc(x_51); +lean_dec(x_12); +x_52 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__4; +x_53 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_52, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_51); +x_54 = lean_ctor_get(x_53, 0); +lean_inc(x_54); +x_55 = lean_ctor_get(x_53, 1); +lean_inc(x_55); +if (lean_is_exclusive(x_53)) { + lean_ctor_release(x_53, 0); + lean_ctor_release(x_53, 1); + x_56 = x_53; +} else { + lean_dec_ref(x_53); + x_56 = lean_box(0); +} +x_57 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__5; +x_58 = lean_unbox(x_54); +lean_dec(x_54); +if (x_58 == 0) +{ +lean_object* x_59; lean_object* x_60; +lean_dec(x_56); +lean_dec(x_1); +x_59 = lean_box(0); +x_60 = lean_apply_10(x_57, x_59, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_55); +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; +x_61 = l_Lean_MessageData_ofExpr(x_1); +x_62 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__7; +if (lean_is_scalar(x_56)) { + x_63 = lean_alloc_ctor(7, 2, 0); +} else { + x_63 = x_56; + lean_ctor_set_tag(x_63, 7); +} +lean_ctor_set(x_63, 0, x_62); +lean_ctor_set(x_63, 1, x_61); +x_64 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__9; +x_65 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_65, 0, x_63); +lean_ctor_set(x_65, 1, x_64); +x_66 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_52, x_65, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_55); +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_apply_10(x_57, x_67, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_68); +return x_69; +} +} +} +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__1___boxed), 10, 0); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("And", 3, 3); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__2; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Or", 2, 2); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___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_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__4; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("dite", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__6; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__8() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("ite", 3, 3); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___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_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__8; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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; +lean_inc(x_1); +x_11 = l_Lean_Meta_instantiateMVarsIfMVarApp(x_1, x_6, x_7, x_8, x_9, x_10); +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +x_13 = lean_ctor_get(x_11, 1); +lean_inc(x_13); +lean_dec(x_11); +x_14 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__1; +x_15 = l_Lean_Expr_cleanupAnnotations(x_12); +x_16 = l_Lean_Expr_isApp(x_15); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; +lean_dec(x_15); +x_17 = lean_box(0); +x_18 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8(x_1, x_17, x_2, x_3, 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; uint8_t x_21; +x_19 = l_Lean_Expr_appArg(x_15, lean_box(0)); +x_20 = l_Lean_Expr_appFnCleanup(x_15, lean_box(0)); +x_21 = l_Lean_Expr_isApp(x_20); +if (x_21 == 0) +{ +lean_object* x_22; lean_object* x_23; +lean_dec(x_20); +lean_dec(x_19); +x_22 = lean_box(0); +x_23 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8(x_1, x_22, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_13); +return x_23; +} +else +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; +x_24 = l_Lean_Expr_appArg(x_20, lean_box(0)); +x_25 = l_Lean_Expr_appFnCleanup(x_20, lean_box(0)); +x_26 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__3; +x_27 = l_Lean_Expr_isConstOf(x_25, x_26); +if (x_27 == 0) +{ +lean_object* x_28; uint8_t x_29; +x_28 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__5; +x_29 = l_Lean_Expr_isConstOf(x_25, x_28); +if (x_29 == 0) +{ +uint8_t x_30; +lean_dec(x_24); +lean_dec(x_19); +x_30 = l_Lean_Expr_isApp(x_25); +if (x_30 == 0) +{ +lean_object* x_31; lean_object* x_32; +lean_dec(x_25); +x_31 = lean_box(0); +x_32 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8(x_1, x_31, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_13); +return x_32; +} +else +{ +lean_object* x_33; uint8_t x_34; +x_33 = l_Lean_Expr_appFnCleanup(x_25, lean_box(0)); +x_34 = l_Lean_Expr_isApp(x_33); +if (x_34 == 0) +{ +lean_object* x_35; lean_object* x_36; +lean_dec(x_33); +x_35 = lean_box(0); +x_36 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8(x_1, x_35, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_13); +return x_36; +} +else +{ +lean_object* x_37; lean_object* x_38; uint8_t x_39; +x_37 = l_Lean_Expr_appArg(x_33, lean_box(0)); +x_38 = l_Lean_Expr_appFnCleanup(x_33, lean_box(0)); +x_39 = l_Lean_Expr_isApp(x_38); +if (x_39 == 0) +{ +lean_object* x_40; lean_object* x_41; +lean_dec(x_38); +lean_dec(x_37); +x_40 = lean_box(0); +x_41 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8(x_1, x_40, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_13); +return x_41; +} +else +{ +lean_object* x_42; lean_object* x_43; uint8_t x_44; +x_42 = l_Lean_Expr_appFnCleanup(x_38, lean_box(0)); +x_43 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__7; +x_44 = l_Lean_Expr_isConstOf(x_42, x_43); +if (x_44 == 0) +{ +lean_object* x_45; uint8_t x_46; +x_45 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__9; +x_46 = l_Lean_Expr_isConstOf(x_42, x_45); +lean_dec(x_42); +if (x_46 == 0) +{ +lean_object* x_47; lean_object* x_48; +lean_dec(x_37); +x_47 = lean_box(0); +x_48 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8(x_1, x_47, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_13); +return x_48; +} +else +{ +lean_object* x_49; +lean_dec(x_1); +x_49 = lean_apply_10(x_14, x_37, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_13); +return x_49; +} +} +else +{ +lean_object* x_50; +lean_dec(x_42); +lean_dec(x_1); +x_50 = lean_apply_10(x_14, x_37, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_13); +return x_50; +} +} +} +} +} +else +{ +lean_object* x_51; +lean_dec(x_25); +x_51 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__3(x_1, x_24, x_19, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_13); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_51; +} +} +else +{ +lean_object* x_52; +lean_dec(x_25); +x_52 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__2(x_1, x_24, x_19, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_13); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_52; +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___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_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_11; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +x_13 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_13; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +x_13 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_13; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___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_11; +x_11 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___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); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_11; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; +x_12 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__5(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_2); +return x_12; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___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) { +_start: +{ +lean_object* x_12; +x_12 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__6(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_2); +return x_12; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__7(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_11; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; +x_12 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_2); +return x_12; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_selectNextSplit_x3f_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, lean_object* x_11) { +_start: +{ +lean_object* x_12; +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_1); +lean_ctor_set(x_12, 1, x_11); +return x_12; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_selectNextSplit_x3f_go(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t 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); +x_13 = lean_st_ref_take(x_4, x_12); +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +x_16 = !lean_is_exclusive(x_14); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_17 = lean_ctor_get(x_14, 16); +lean_dec(x_17); +x_18 = l_List_reverse___rarg(x_3); +lean_ctor_set(x_14, 16, x_18); +x_19 = lean_st_ref_set(x_4, x_14, x_15); +if (lean_obj_tag(x_2) == 0) +{ +uint8_t x_20; +lean_dec(x_4); +x_20 = !lean_is_exclusive(x_19); +if (x_20 == 0) +{ +lean_object* x_21; +x_21 = lean_ctor_get(x_19, 0); +lean_dec(x_21); +lean_ctor_set(x_19, 0, x_2); +return x_19; +} +else +{ +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_2); +lean_ctor_set(x_23, 1, x_22); +return x_23; +} +} +else +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; +x_24 = lean_ctor_get(x_19, 1); +lean_inc(x_24); +lean_dec(x_19); +x_25 = lean_st_ref_take(x_4, x_24); +x_26 = lean_ctor_get(x_25, 0); +lean_inc(x_26); +x_27 = lean_ctor_get(x_25, 1); +lean_inc(x_27); +lean_dec(x_25); +x_28 = !lean_is_exclusive(x_26); +if (x_28 == 0) +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; +x_29 = lean_ctor_get(x_26, 17); +x_30 = lean_ctor_get(x_26, 12); +lean_dec(x_30); +x_31 = lean_unsigned_to_nat(1u); +x_32 = lean_nat_add(x_29, x_31); +lean_dec(x_29); +x_33 = lean_unsigned_to_nat(0u); +lean_ctor_set(x_26, 17, x_32); +lean_ctor_set(x_26, 12, x_33); +x_34 = lean_st_ref_set(x_4, x_26, x_27); +lean_dec(x_4); +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_2); +return x_34; +} +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_2); +lean_ctor_set(x_38, 1, x_37); +return x_38; +} +} +else +{ +lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; 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; 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_39 = lean_ctor_get(x_26, 0); +x_40 = lean_ctor_get(x_26, 1); +x_41 = lean_ctor_get(x_26, 2); +x_42 = lean_ctor_get(x_26, 3); +x_43 = lean_ctor_get(x_26, 4); +x_44 = lean_ctor_get(x_26, 5); +x_45 = lean_ctor_get_uint8(x_26, sizeof(void*)*20); +x_46 = lean_ctor_get(x_26, 6); +x_47 = lean_ctor_get(x_26, 7); +x_48 = lean_ctor_get(x_26, 8); +x_49 = lean_ctor_get(x_26, 9); +x_50 = lean_ctor_get(x_26, 10); +x_51 = lean_ctor_get(x_26, 11); +x_52 = lean_ctor_get(x_26, 13); +x_53 = lean_ctor_get(x_26, 14); +x_54 = lean_ctor_get(x_26, 15); +x_55 = lean_ctor_get(x_26, 16); +x_56 = lean_ctor_get(x_26, 17); +x_57 = lean_ctor_get(x_26, 18); +x_58 = lean_ctor_get(x_26, 19); +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_inc(x_49); +lean_inc(x_48); +lean_inc(x_47); +lean_inc(x_46); +lean_inc(x_44); +lean_inc(x_43); +lean_inc(x_42); +lean_inc(x_41); +lean_inc(x_40); +lean_inc(x_39); +lean_dec(x_26); +x_59 = lean_unsigned_to_nat(1u); +x_60 = lean_nat_add(x_56, x_59); +lean_dec(x_56); +x_61 = lean_unsigned_to_nat(0u); +x_62 = lean_alloc_ctor(0, 20, 1); +lean_ctor_set(x_62, 0, x_39); +lean_ctor_set(x_62, 1, x_40); +lean_ctor_set(x_62, 2, x_41); +lean_ctor_set(x_62, 3, x_42); +lean_ctor_set(x_62, 4, x_43); +lean_ctor_set(x_62, 5, x_44); +lean_ctor_set(x_62, 6, x_46); +lean_ctor_set(x_62, 7, x_47); +lean_ctor_set(x_62, 8, x_48); +lean_ctor_set(x_62, 9, x_49); +lean_ctor_set(x_62, 10, x_50); +lean_ctor_set(x_62, 11, x_51); +lean_ctor_set(x_62, 12, x_61); +lean_ctor_set(x_62, 13, x_52); +lean_ctor_set(x_62, 14, x_53); +lean_ctor_set(x_62, 15, x_54); +lean_ctor_set(x_62, 16, x_55); +lean_ctor_set(x_62, 17, x_60); +lean_ctor_set(x_62, 18, x_57); +lean_ctor_set(x_62, 19, x_58); +lean_ctor_set_uint8(x_62, sizeof(void*)*20, x_45); +x_63 = lean_st_ref_set(x_4, x_62, x_27); +lean_dec(x_4); +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); +} +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_2); +lean_ctor_set(x_66, 1, x_64); +return x_66; +} +} +} +else +{ +lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; uint8_t x_73; 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_67 = lean_ctor_get(x_14, 0); +x_68 = lean_ctor_get(x_14, 1); +x_69 = lean_ctor_get(x_14, 2); +x_70 = lean_ctor_get(x_14, 3); +x_71 = lean_ctor_get(x_14, 4); +x_72 = lean_ctor_get(x_14, 5); +x_73 = lean_ctor_get_uint8(x_14, sizeof(void*)*20); +x_74 = lean_ctor_get(x_14, 6); +x_75 = lean_ctor_get(x_14, 7); +x_76 = lean_ctor_get(x_14, 8); +x_77 = lean_ctor_get(x_14, 9); +x_78 = lean_ctor_get(x_14, 10); +x_79 = lean_ctor_get(x_14, 11); +x_80 = lean_ctor_get(x_14, 12); +x_81 = lean_ctor_get(x_14, 13); +x_82 = lean_ctor_get(x_14, 14); +x_83 = lean_ctor_get(x_14, 15); +x_84 = lean_ctor_get(x_14, 17); +x_85 = lean_ctor_get(x_14, 18); +x_86 = lean_ctor_get(x_14, 19); +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_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_72); +lean_inc(x_71); +lean_inc(x_70); +lean_inc(x_69); +lean_inc(x_68); +lean_inc(x_67); +lean_dec(x_14); +x_87 = l_List_reverse___rarg(x_3); +x_88 = lean_alloc_ctor(0, 20, 1); +lean_ctor_set(x_88, 0, x_67); +lean_ctor_set(x_88, 1, x_68); +lean_ctor_set(x_88, 2, x_69); +lean_ctor_set(x_88, 3, x_70); +lean_ctor_set(x_88, 4, x_71); +lean_ctor_set(x_88, 5, x_72); +lean_ctor_set(x_88, 6, x_74); +lean_ctor_set(x_88, 7, x_75); +lean_ctor_set(x_88, 8, x_76); +lean_ctor_set(x_88, 9, x_77); +lean_ctor_set(x_88, 10, x_78); +lean_ctor_set(x_88, 11, x_79); +lean_ctor_set(x_88, 12, x_80); +lean_ctor_set(x_88, 13, x_81); +lean_ctor_set(x_88, 14, x_82); +lean_ctor_set(x_88, 15, x_83); +lean_ctor_set(x_88, 16, x_87); +lean_ctor_set(x_88, 17, x_84); +lean_ctor_set(x_88, 18, x_85); +lean_ctor_set(x_88, 19, x_86); +lean_ctor_set_uint8(x_88, sizeof(void*)*20, x_73); +x_89 = lean_st_ref_set(x_4, x_88, x_15); +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_90; lean_object* x_91; lean_object* x_92; +lean_dec(x_4); +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); +} +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_2); +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; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; uint8_t x_103; 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_93 = lean_ctor_get(x_89, 1); +lean_inc(x_93); +lean_dec(x_89); +x_94 = lean_st_ref_take(x_4, x_93); +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); +x_98 = lean_ctor_get(x_95, 1); +lean_inc(x_98); +x_99 = lean_ctor_get(x_95, 2); +lean_inc(x_99); +x_100 = lean_ctor_get(x_95, 3); +lean_inc(x_100); +x_101 = lean_ctor_get(x_95, 4); +lean_inc(x_101); +x_102 = lean_ctor_get(x_95, 5); +lean_inc(x_102); +x_103 = lean_ctor_get_uint8(x_95, sizeof(void*)*20); +x_104 = lean_ctor_get(x_95, 6); +lean_inc(x_104); +x_105 = lean_ctor_get(x_95, 7); +lean_inc(x_105); +x_106 = lean_ctor_get(x_95, 8); +lean_inc(x_106); +x_107 = lean_ctor_get(x_95, 9); +lean_inc(x_107); +x_108 = lean_ctor_get(x_95, 10); +lean_inc(x_108); +x_109 = lean_ctor_get(x_95, 11); +lean_inc(x_109); +x_110 = lean_ctor_get(x_95, 13); +lean_inc(x_110); +x_111 = lean_ctor_get(x_95, 14); +lean_inc(x_111); +x_112 = lean_ctor_get(x_95, 15); +lean_inc(x_112); +x_113 = lean_ctor_get(x_95, 16); +lean_inc(x_113); +x_114 = lean_ctor_get(x_95, 17); +lean_inc(x_114); +x_115 = lean_ctor_get(x_95, 18); +lean_inc(x_115); +x_116 = lean_ctor_get(x_95, 19); +lean_inc(x_116); +if (lean_is_exclusive(x_95)) { + lean_ctor_release(x_95, 0); + lean_ctor_release(x_95, 1); + lean_ctor_release(x_95, 2); + lean_ctor_release(x_95, 3); + lean_ctor_release(x_95, 4); + lean_ctor_release(x_95, 5); + lean_ctor_release(x_95, 6); + lean_ctor_release(x_95, 7); + lean_ctor_release(x_95, 8); + lean_ctor_release(x_95, 9); + lean_ctor_release(x_95, 10); + lean_ctor_release(x_95, 11); + lean_ctor_release(x_95, 12); + lean_ctor_release(x_95, 13); + lean_ctor_release(x_95, 14); + lean_ctor_release(x_95, 15); + lean_ctor_release(x_95, 16); + lean_ctor_release(x_95, 17); + lean_ctor_release(x_95, 18); + lean_ctor_release(x_95, 19); + x_117 = x_95; +} else { + lean_dec_ref(x_95); + x_117 = lean_box(0); +} +x_118 = lean_unsigned_to_nat(1u); +x_119 = lean_nat_add(x_114, x_118); +lean_dec(x_114); +x_120 = lean_unsigned_to_nat(0u); +if (lean_is_scalar(x_117)) { + x_121 = lean_alloc_ctor(0, 20, 1); +} else { + x_121 = x_117; +} +lean_ctor_set(x_121, 0, x_97); +lean_ctor_set(x_121, 1, x_98); +lean_ctor_set(x_121, 2, x_99); +lean_ctor_set(x_121, 3, x_100); +lean_ctor_set(x_121, 4, x_101); +lean_ctor_set(x_121, 5, x_102); +lean_ctor_set(x_121, 6, x_104); +lean_ctor_set(x_121, 7, x_105); +lean_ctor_set(x_121, 8, x_106); +lean_ctor_set(x_121, 9, x_107); +lean_ctor_set(x_121, 10, x_108); +lean_ctor_set(x_121, 11, x_109); +lean_ctor_set(x_121, 12, x_120); +lean_ctor_set(x_121, 13, x_110); +lean_ctor_set(x_121, 14, x_111); +lean_ctor_set(x_121, 15, x_112); +lean_ctor_set(x_121, 16, x_113); +lean_ctor_set(x_121, 17, x_119); +lean_ctor_set(x_121, 18, x_115); +lean_ctor_set(x_121, 19, x_116); +lean_ctor_set_uint8(x_121, sizeof(void*)*20, x_103); +x_122 = lean_st_ref_set(x_4, x_121, x_96); +lean_dec(x_4); +x_123 = lean_ctor_get(x_122, 1); +lean_inc(x_123); +if (lean_is_exclusive(x_122)) { + lean_ctor_release(x_122, 0); + lean_ctor_release(x_122, 1); + x_124 = x_122; +} else { + lean_dec_ref(x_122); + x_124 = lean_box(0); +} +if (lean_is_scalar(x_124)) { + x_125 = lean_alloc_ctor(0, 2, 0); +} else { + x_125 = x_124; +} +lean_ctor_set(x_125, 0, x_2); +lean_ctor_set(x_125, 1, x_123); +return x_125; +} +} +} +else +{ +uint8_t x_126; +x_126 = !lean_is_exclusive(x_1); +if (x_126 == 0) +{ +lean_object* x_127; lean_object* x_128; lean_object* x_129; +x_127 = lean_ctor_get(x_1, 0); +x_128 = lean_ctor_get(x_1, 1); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_127); +x_129 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus(x_127, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_129) == 0) +{ +lean_object* x_130; uint8_t x_131; +x_130 = lean_ctor_get(x_129, 0); +lean_inc(x_130); +x_131 = lean_unbox(x_130); +lean_dec(x_130); +switch (x_131) { +case 0: +{ +lean_object* x_132; +lean_free_object(x_1); +lean_dec(x_127); +x_132 = lean_ctor_get(x_129, 1); +lean_inc(x_132); +lean_dec(x_129); +x_1 = x_128; +x_12 = x_132; +goto _start; +} +case 1: +{ +lean_object* x_134; +x_134 = lean_ctor_get(x_129, 1); +lean_inc(x_134); +lean_dec(x_129); +lean_ctor_set(x_1, 1, x_3); +{ +lean_object* _tmp_0 = x_128; +lean_object* _tmp_2 = x_1; +lean_object* _tmp_11 = x_134; +x_1 = _tmp_0; +x_3 = _tmp_2; +x_12 = _tmp_11; +} +goto _start; +} +default: +{ +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_136; lean_object* x_137; +lean_free_object(x_1); +x_136 = lean_ctor_get(x_129, 1); +lean_inc(x_136); +lean_dec(x_129); +x_137 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_137, 0, x_127); +x_1 = x_128; +x_2 = x_137; +x_12 = x_136; +goto _start; +} +else +{ +lean_object* x_139; lean_object* x_140; lean_object* x_141; +x_139 = lean_ctor_get(x_129, 1); +lean_inc(x_139); +lean_dec(x_129); +x_140 = lean_ctor_get(x_2, 0); +lean_inc(x_140); +lean_inc(x_127); +x_141 = l_Lean_Meta_Grind_getGeneration(x_127, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_139); +if (lean_obj_tag(x_141) == 0) +{ +lean_object* x_142; lean_object* x_143; lean_object* x_144; +x_142 = lean_ctor_get(x_141, 0); +lean_inc(x_142); +x_143 = lean_ctor_get(x_141, 1); +lean_inc(x_143); +lean_dec(x_141); +lean_inc(x_140); +x_144 = l_Lean_Meta_Grind_getGeneration(x_140, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_143); +if (lean_obj_tag(x_144) == 0) +{ +lean_object* x_145; lean_object* x_146; uint8_t x_147; +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_nat_dec_lt(x_142, x_145); +lean_dec(x_145); +lean_dec(x_142); +if (x_147 == 0) +{ +lean_dec(x_140); +lean_ctor_set(x_1, 1, x_3); +{ +lean_object* _tmp_0 = x_128; +lean_object* _tmp_2 = x_1; +lean_object* _tmp_11 = x_146; +x_1 = _tmp_0; +x_3 = _tmp_2; +x_12 = _tmp_11; +} +goto _start; +} +else +{ +uint8_t x_149; +x_149 = !lean_is_exclusive(x_2); +if (x_149 == 0) +{ +lean_object* x_150; +x_150 = lean_ctor_get(x_2, 0); +lean_dec(x_150); +lean_ctor_set(x_2, 0, x_127); +lean_ctor_set(x_1, 1, x_3); +lean_ctor_set(x_1, 0, x_140); +{ +lean_object* _tmp_0 = x_128; +lean_object* _tmp_2 = x_1; +lean_object* _tmp_11 = x_146; +x_1 = _tmp_0; +x_3 = _tmp_2; +x_12 = _tmp_11; +} +goto _start; +} +else +{ +lean_object* x_152; +lean_dec(x_2); +x_152 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_152, 0, x_127); +lean_ctor_set(x_1, 1, x_3); +lean_ctor_set(x_1, 0, x_140); +{ +lean_object* _tmp_0 = x_128; +lean_object* _tmp_1 = x_152; +lean_object* _tmp_2 = x_1; +lean_object* _tmp_11 = x_146; +x_1 = _tmp_0; +x_2 = _tmp_1; +x_3 = _tmp_2; +x_12 = _tmp_11; +} +goto _start; +} +} +} +else +{ +uint8_t x_154; +lean_dec(x_142); +lean_dec(x_140); +lean_free_object(x_1); +lean_dec(x_128); +lean_dec(x_127); +lean_dec(x_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_154 = !lean_is_exclusive(x_144); +if (x_154 == 0) +{ +return x_144; +} +else +{ +lean_object* x_155; lean_object* x_156; lean_object* x_157; +x_155 = lean_ctor_get(x_144, 0); +x_156 = lean_ctor_get(x_144, 1); +lean_inc(x_156); +lean_inc(x_155); +lean_dec(x_144); +x_157 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_157, 0, x_155); +lean_ctor_set(x_157, 1, x_156); +return x_157; +} +} +} +else +{ +uint8_t x_158; +lean_dec(x_140); +lean_free_object(x_1); +lean_dec(x_128); +lean_dec(x_127); +lean_dec(x_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_158 = !lean_is_exclusive(x_141); +if (x_158 == 0) +{ +return x_141; +} +else +{ +lean_object* x_159; lean_object* x_160; lean_object* x_161; +x_159 = lean_ctor_get(x_141, 0); +x_160 = lean_ctor_get(x_141, 1); +lean_inc(x_160); +lean_inc(x_159); +lean_dec(x_141); +x_161 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_161, 0, x_159); +lean_ctor_set(x_161, 1, x_160); +return x_161; +} +} +} +} +} +} +else +{ +uint8_t x_162; +lean_free_object(x_1); +lean_dec(x_128); +lean_dec(x_127); +lean_dec(x_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_162 = !lean_is_exclusive(x_129); +if (x_162 == 0) +{ +return x_129; +} +else +{ +lean_object* x_163; lean_object* x_164; lean_object* x_165; +x_163 = lean_ctor_get(x_129, 0); +x_164 = lean_ctor_get(x_129, 1); +lean_inc(x_164); +lean_inc(x_163); +lean_dec(x_129); +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 +{ +lean_object* x_166; lean_object* x_167; lean_object* x_168; +x_166 = lean_ctor_get(x_1, 0); +x_167 = lean_ctor_get(x_1, 1); +lean_inc(x_167); +lean_inc(x_166); +lean_dec(x_1); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_166); +x_168 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus(x_166, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_168) == 0) +{ +lean_object* x_169; uint8_t x_170; +x_169 = lean_ctor_get(x_168, 0); +lean_inc(x_169); +x_170 = lean_unbox(x_169); +lean_dec(x_169); +switch (x_170) { +case 0: +{ +lean_object* x_171; +lean_dec(x_166); +x_171 = lean_ctor_get(x_168, 1); +lean_inc(x_171); +lean_dec(x_168); +x_1 = x_167; +x_12 = x_171; +goto _start; +} +case 1: +{ +lean_object* x_173; lean_object* x_174; +x_173 = lean_ctor_get(x_168, 1); +lean_inc(x_173); +lean_dec(x_168); +x_174 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_174, 0, x_166); +lean_ctor_set(x_174, 1, x_3); +x_1 = x_167; +x_3 = x_174; +x_12 = x_173; +goto _start; +} +default: +{ +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_176; lean_object* x_177; +x_176 = lean_ctor_get(x_168, 1); +lean_inc(x_176); +lean_dec(x_168); +x_177 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_177, 0, x_166); +x_1 = x_167; +x_2 = x_177; +x_12 = x_176; +goto _start; +} +else +{ +lean_object* x_179; lean_object* x_180; lean_object* x_181; +x_179 = lean_ctor_get(x_168, 1); +lean_inc(x_179); +lean_dec(x_168); +x_180 = lean_ctor_get(x_2, 0); +lean_inc(x_180); +lean_inc(x_166); +x_181 = l_Lean_Meta_Grind_getGeneration(x_166, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_179); +if (lean_obj_tag(x_181) == 0) +{ +lean_object* x_182; lean_object* x_183; lean_object* x_184; +x_182 = lean_ctor_get(x_181, 0); +lean_inc(x_182); +x_183 = lean_ctor_get(x_181, 1); +lean_inc(x_183); +lean_dec(x_181); +lean_inc(x_180); +x_184 = l_Lean_Meta_Grind_getGeneration(x_180, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_183); +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_nat_dec_lt(x_182, x_185); +lean_dec(x_185); +lean_dec(x_182); +if (x_187 == 0) +{ +lean_object* x_188; +lean_dec(x_180); +x_188 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_188, 0, x_166); +lean_ctor_set(x_188, 1, x_3); +x_1 = x_167; +x_3 = x_188; +x_12 = x_186; +goto _start; +} +else +{ +lean_object* x_190; lean_object* x_191; lean_object* x_192; +if (lean_is_exclusive(x_2)) { + lean_ctor_release(x_2, 0); + x_190 = x_2; +} else { + lean_dec_ref(x_2); + x_190 = lean_box(0); +} +if (lean_is_scalar(x_190)) { + x_191 = lean_alloc_ctor(1, 1, 0); +} else { + x_191 = x_190; +} +lean_ctor_set(x_191, 0, x_166); +x_192 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_192, 0, x_180); +lean_ctor_set(x_192, 1, x_3); +x_1 = x_167; +x_2 = x_191; +x_3 = x_192; +x_12 = x_186; +goto _start; +} +} +else +{ +lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; +lean_dec(x_182); +lean_dec(x_180); +lean_dec(x_167); +lean_dec(x_166); +lean_dec(x_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_194 = lean_ctor_get(x_184, 0); +lean_inc(x_194); +x_195 = lean_ctor_get(x_184, 1); +lean_inc(x_195); +if (lean_is_exclusive(x_184)) { + lean_ctor_release(x_184, 0); + lean_ctor_release(x_184, 1); + x_196 = x_184; +} else { + lean_dec_ref(x_184); + x_196 = lean_box(0); +} +if (lean_is_scalar(x_196)) { + x_197 = lean_alloc_ctor(1, 2, 0); +} else { + x_197 = x_196; +} +lean_ctor_set(x_197, 0, x_194); +lean_ctor_set(x_197, 1, x_195); +return x_197; +} +} +else +{ +lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; +lean_dec(x_180); +lean_dec(x_167); +lean_dec(x_166); +lean_dec(x_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_198 = lean_ctor_get(x_181, 0); +lean_inc(x_198); +x_199 = lean_ctor_get(x_181, 1); +lean_inc(x_199); +if (lean_is_exclusive(x_181)) { + lean_ctor_release(x_181, 0); + lean_ctor_release(x_181, 1); + x_200 = x_181; +} else { + lean_dec_ref(x_181); + 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; +} +} +} +} +} +else +{ +lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; +lean_dec(x_167); +lean_dec(x_166); +lean_dec(x_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_202 = lean_ctor_get(x_168, 0); +lean_inc(x_202); +x_203 = lean_ctor_get(x_168, 1); +lean_inc(x_203); +if (lean_is_exclusive(x_168)) { + lean_ctor_release(x_168, 0); + lean_ctor_release(x_168, 1); + x_204 = x_168; +} else { + lean_dec_ref(x_168); + x_204 = lean_box(0); +} +if (lean_is_scalar(x_204)) { + x_205 = lean_alloc_ctor(1, 2, 0); +} else { + x_205 = x_204; +} +lean_ctor_set(x_205, 0, x_202); +lean_ctor_set(x_205, 1, x_203); +return x_205; +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_selectNextSplit_x3f_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, lean_object* x_11) { +_start: +{ +lean_object* x_12; +x_12 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_selectNextSplit_x3f_go___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_12; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_selectNextSplit_x3f___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_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; +x_11 = lean_st_ref_get(x_2, 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, 16); +lean_inc(x_14); +lean_dec(x_12); +x_15 = lean_box(0); +x_16 = lean_box(0); +x_17 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_selectNextSplit_x3f_go(x_14, x_15, x_16, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_13); +return x_17; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_selectNextSplit_x3f___lambda__2___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_selectNextSplit_x3f___lambda__1___boxed), 10, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_selectNextSplit_x3f___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_11 = l_Lean_Meta_Grind_checkMaxCaseSplit(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +x_13 = lean_unbox(x_12); +lean_dec(x_12); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_14 = lean_ctor_get(x_11, 1); +lean_inc(x_14); +lean_dec(x_11); +x_15 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_selectNextSplit_x3f___lambda__2___closed__1; +x_16 = lean_box(0); +x_17 = lean_apply_10(x_15, x_16, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_14); +return x_17; +} +else +{ +uint8_t x_18; +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_18 = !lean_is_exclusive(x_11); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; +x_19 = lean_ctor_get(x_11, 0); +lean_dec(x_19); +x_20 = lean_box(0); +lean_ctor_set(x_11, 0, x_20); +return x_11; +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_11, 1); +lean_inc(x_21); +lean_dec(x_11); +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; +} +} +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_selectNextSplit_x3f___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_selectNextSplit_x3f___lambda__2___boxed), 10, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_selectNextSplit_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) { +_start: +{ +lean_object* x_10; lean_object* x_11; uint8_t x_12; +x_10 = l_Lean_Meta_Grind_isInconsistent(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_11 = lean_ctor_get(x_10, 0); +lean_inc(x_11); +x_12 = lean_unbox(x_11); +lean_dec(x_11); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_13 = lean_ctor_get(x_10, 1); +lean_inc(x_13); +lean_dec(x_10); +x_14 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_selectNextSplit_x3f___closed__1; +x_15 = lean_box(0); +x_16 = lean_apply_10(x_14, x_15, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_13); +return x_16; +} +else +{ +uint8_t x_17; +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_17 = !lean_is_exclusive(x_10); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; +x_18 = lean_ctor_get(x_10, 0); +lean_dec(x_18); +x_19 = lean_box(0); +lean_ctor_set(x_10, 0, x_19); +return x_10; +} +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_10, 1); +lean_inc(x_20); +lean_dec(x_10); +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___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_selectNextSplit_x3f___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_selectNextSplit_x3f___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_1); +return x_11; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_selectNextSplit_x3f___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_selectNextSplit_x3f___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_1); +return x_11; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; lean_object* x_12; +x_11 = l_Lean_mkEM(x_1); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_10); +return x_12; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__2___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__2___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Grind", 5, 5); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__2___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("or_of_and_eq_false", 18, 18); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__2___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__2___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__2___closed__2; +x_3 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__2___closed__3; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__2___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__2___closed__4; +x_3 = l_Lean_Expr_const___override(x_2, x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +x_13 = l_Lean_Meta_Grind_mkEqFalseProof(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_13) == 0) +{ +uint8_t x_14; +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_ctor_get(x_13, 0); +x_16 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__2___closed__5; +x_17 = l_Lean_mkApp3(x_16, x_2, x_3, x_15); +lean_ctor_set(x_13, 0, x_17); +return x_13; +} +else +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_18 = lean_ctor_get(x_13, 0); +x_19 = lean_ctor_get(x_13, 1); +lean_inc(x_19); +lean_inc(x_18); +lean_dec(x_13); +x_20 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__2___closed__5; +x_21 = l_Lean_mkApp3(x_20, x_2, x_3, x_18); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_19); +return x_22; +} +} +else +{ +uint8_t x_23; +lean_dec(x_3); +lean_dec(x_2); +x_23 = !lean_is_exclusive(x_13); +if (x_23 == 0) +{ +return x_13; +} +else +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_13, 0); +x_25 = lean_ctor_get(x_13, 1); +lean_inc(x_25); +lean_inc(x_24); +lean_dec(x_13); +x_26 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_26, 0, x_24); +lean_ctor_set(x_26, 1, x_25); +return x_26; +} +} +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__3___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("of_eq_true", 10, 10); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__3___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__3___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__3___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__3___closed__2; +x_3 = l_Lean_Expr_const___override(x_2, x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___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_inc(x_1); +x_12 = l_Lean_Meta_Grind_mkEqTrueProof(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_12) == 0) +{ +uint8_t x_13; +x_13 = !lean_is_exclusive(x_12); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_14 = lean_ctor_get(x_12, 0); +x_15 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__3___closed__3; +x_16 = l_Lean_mkAppB(x_15, x_1, x_14); +lean_ctor_set(x_12, 0, x_16); +return x_12; +} +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_12, 0); +x_18 = lean_ctor_get(x_12, 1); +lean_inc(x_18); +lean_inc(x_17); +lean_dec(x_12); +x_19 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__3___closed__3; +x_20 = l_Lean_mkAppB(x_19, x_1, x_17); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_20); +lean_ctor_set(x_21, 1, x_18); +return x_21; +} +} +else +{ +uint8_t x_22; +lean_dec(x_1); +x_22 = !lean_is_exclusive(x_12); +if (x_22 == 0) +{ +return x_12; +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_12, 0); +x_24 = lean_ctor_get(x_12, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_12); +x_25 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_24); +return x_25; +} +} +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__1___boxed), 10, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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; +lean_inc(x_1); +x_11 = l_Lean_Meta_instantiateMVarsIfMVarApp(x_1, x_6, x_7, x_8, x_9, x_10); +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +x_13 = lean_ctor_get(x_11, 1); +lean_inc(x_13); +lean_dec(x_11); +x_14 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___closed__1; +x_15 = l_Lean_Expr_cleanupAnnotations(x_12); +x_16 = l_Lean_Expr_isApp(x_15); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; +lean_dec(x_15); +x_17 = lean_box(0); +x_18 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__3(x_1, x_17, x_2, x_3, 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; uint8_t x_21; +x_19 = l_Lean_Expr_appArg(x_15, lean_box(0)); +x_20 = l_Lean_Expr_appFnCleanup(x_15, lean_box(0)); +x_21 = l_Lean_Expr_isApp(x_20); +if (x_21 == 0) +{ +lean_object* x_22; lean_object* x_23; +lean_dec(x_20); +lean_dec(x_19); +x_22 = lean_box(0); +x_23 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__3(x_1, x_22, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_13); +return x_23; +} +else +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; +x_24 = l_Lean_Expr_appArg(x_20, lean_box(0)); +x_25 = l_Lean_Expr_appFnCleanup(x_20, lean_box(0)); +x_26 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__3; +x_27 = l_Lean_Expr_isConstOf(x_25, x_26); +if (x_27 == 0) +{ +uint8_t x_28; +lean_dec(x_24); +lean_dec(x_19); +x_28 = l_Lean_Expr_isApp(x_25); +if (x_28 == 0) +{ +lean_object* x_29; lean_object* x_30; +lean_dec(x_25); +x_29 = lean_box(0); +x_30 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__3(x_1, x_29, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_13); +return x_30; +} +else +{ +lean_object* x_31; uint8_t x_32; +x_31 = l_Lean_Expr_appFnCleanup(x_25, lean_box(0)); +x_32 = l_Lean_Expr_isApp(x_31); +if (x_32 == 0) +{ +lean_object* x_33; lean_object* x_34; +lean_dec(x_31); +x_33 = lean_box(0); +x_34 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__3(x_1, x_33, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_13); +return x_34; +} +else +{ +lean_object* x_35; lean_object* x_36; uint8_t x_37; +x_35 = l_Lean_Expr_appArg(x_31, lean_box(0)); +x_36 = l_Lean_Expr_appFnCleanup(x_31, lean_box(0)); +x_37 = l_Lean_Expr_isApp(x_36); +if (x_37 == 0) +{ +lean_object* x_38; lean_object* x_39; +lean_dec(x_36); +lean_dec(x_35); +x_38 = lean_box(0); +x_39 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__3(x_1, x_38, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_13); +return x_39; +} +else +{ +lean_object* x_40; lean_object* x_41; uint8_t x_42; +x_40 = l_Lean_Expr_appFnCleanup(x_36, lean_box(0)); +x_41 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__7; +x_42 = l_Lean_Expr_isConstOf(x_40, x_41); +if (x_42 == 0) +{ +lean_object* x_43; uint8_t x_44; +x_43 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__9; +x_44 = l_Lean_Expr_isConstOf(x_40, x_43); +lean_dec(x_40); +if (x_44 == 0) +{ +lean_object* x_45; lean_object* x_46; +lean_dec(x_35); +x_45 = lean_box(0); +x_46 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__3(x_1, x_45, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_13); +return x_46; +} +else +{ +lean_object* x_47; +lean_dec(x_1); +x_47 = lean_apply_10(x_14, x_35, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_13); +return x_47; +} +} +else +{ +lean_object* x_48; +lean_dec(x_40); +lean_dec(x_1); +x_48 = lean_apply_10(x_14, x_35, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_13); +return x_48; +} +} +} +} +} +else +{ +lean_object* x_49; +lean_dec(x_25); +x_49 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__2(x_1, x_24, x_19, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_13); +return x_49; +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___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_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_11; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___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___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___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_2); +return x_12; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_introNewHyp(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_12; lean_object* x_13; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_12 = l_List_reverse___rarg(x_2); +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_11); +return x_13; +} +else +{ +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_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_16 = l_Lean_Meta_Grind_intros(x_3, x_14, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +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_List_appendTR___rarg(x_17, x_2); +x_1 = x_15; +x_2 = x_19; +x_11 = x_18; +goto _start; +} +else +{ +uint8_t x_21; +lean_dec(x_15); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_21 = !lean_is_exclusive(x_16); +if (x_21 == 0) +{ +return x_16; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_22 = lean_ctor_get(x_16, 0); +x_23 = lean_ctor_get(x_16, 1); +lean_inc(x_23); +lean_inc(x_22); +lean_dec(x_16); +x_24 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_24, 0, x_22); +lean_ctor_set(x_24, 1, x_23); +return x_24; +} +} +} +} +} +LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Meta_Grind_splitNext___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_4; +x_4 = l_List_reverse___rarg(x_3); +return x_4; +} +else +{ +uint8_t x_5; +x_5 = !lean_is_exclusive(x_2); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_6 = lean_ctor_get(x_2, 0); +x_7 = lean_ctor_get(x_2, 1); +x_8 = lean_ctor_get(x_1, 1); +x_9 = lean_ctor_get(x_1, 2); +x_10 = lean_ctor_get(x_1, 3); +x_11 = lean_ctor_get(x_1, 4); +x_12 = lean_ctor_get(x_1, 5); +x_13 = lean_ctor_get_uint8(x_1, sizeof(void*)*20); +x_14 = lean_ctor_get(x_1, 6); +x_15 = lean_ctor_get(x_1, 7); +x_16 = lean_ctor_get(x_1, 8); +x_17 = lean_ctor_get(x_1, 9); +x_18 = lean_ctor_get(x_1, 10); +x_19 = lean_ctor_get(x_1, 11); +x_20 = lean_ctor_get(x_1, 12); +x_21 = lean_ctor_get(x_1, 13); +x_22 = lean_ctor_get(x_1, 14); +x_23 = lean_ctor_get(x_1, 15); +x_24 = lean_ctor_get(x_1, 16); +x_25 = lean_ctor_get(x_1, 17); +x_26 = lean_ctor_get(x_1, 18); +x_27 = lean_ctor_get(x_1, 19); +lean_inc(x_27); +lean_inc(x_26); +lean_inc(x_25); +lean_inc(x_24); +lean_inc(x_23); +lean_inc(x_22); +lean_inc(x_21); +lean_inc(x_20); +lean_inc(x_19); +lean_inc(x_18); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +x_28 = lean_alloc_ctor(0, 20, 1); +lean_ctor_set(x_28, 0, x_6); +lean_ctor_set(x_28, 1, x_8); +lean_ctor_set(x_28, 2, x_9); +lean_ctor_set(x_28, 3, x_10); +lean_ctor_set(x_28, 4, x_11); +lean_ctor_set(x_28, 5, x_12); +lean_ctor_set(x_28, 6, x_14); +lean_ctor_set(x_28, 7, x_15); +lean_ctor_set(x_28, 8, x_16); +lean_ctor_set(x_28, 9, x_17); +lean_ctor_set(x_28, 10, x_18); +lean_ctor_set(x_28, 11, x_19); +lean_ctor_set(x_28, 12, x_20); +lean_ctor_set(x_28, 13, x_21); +lean_ctor_set(x_28, 14, x_22); +lean_ctor_set(x_28, 15, x_23); +lean_ctor_set(x_28, 16, x_24); +lean_ctor_set(x_28, 17, x_25); +lean_ctor_set(x_28, 18, x_26); +lean_ctor_set(x_28, 19, x_27); +lean_ctor_set_uint8(x_28, sizeof(void*)*20, x_13); +lean_ctor_set(x_2, 1, x_3); +lean_ctor_set(x_2, 0, x_28); +{ +lean_object* _tmp_1 = x_7; +lean_object* _tmp_2 = x_2; +x_2 = _tmp_1; +x_3 = _tmp_2; +} +goto _start; +} +else +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_30 = lean_ctor_get(x_2, 0); +x_31 = lean_ctor_get(x_2, 1); +lean_inc(x_31); +lean_inc(x_30); +lean_dec(x_2); +x_32 = lean_ctor_get(x_1, 1); +x_33 = lean_ctor_get(x_1, 2); +x_34 = lean_ctor_get(x_1, 3); +x_35 = lean_ctor_get(x_1, 4); +x_36 = lean_ctor_get(x_1, 5); +x_37 = lean_ctor_get_uint8(x_1, sizeof(void*)*20); +x_38 = lean_ctor_get(x_1, 6); +x_39 = lean_ctor_get(x_1, 7); +x_40 = lean_ctor_get(x_1, 8); +x_41 = lean_ctor_get(x_1, 9); +x_42 = lean_ctor_get(x_1, 10); +x_43 = lean_ctor_get(x_1, 11); +x_44 = lean_ctor_get(x_1, 12); +x_45 = lean_ctor_get(x_1, 13); +x_46 = lean_ctor_get(x_1, 14); +x_47 = lean_ctor_get(x_1, 15); +x_48 = lean_ctor_get(x_1, 16); +x_49 = lean_ctor_get(x_1, 17); +x_50 = lean_ctor_get(x_1, 18); +x_51 = lean_ctor_get(x_1, 19); +lean_inc(x_51); +lean_inc(x_50); +lean_inc(x_49); +lean_inc(x_48); +lean_inc(x_47); +lean_inc(x_46); +lean_inc(x_45); +lean_inc(x_44); +lean_inc(x_43); +lean_inc(x_42); +lean_inc(x_41); +lean_inc(x_40); +lean_inc(x_39); +lean_inc(x_38); +lean_inc(x_36); +lean_inc(x_35); +lean_inc(x_34); +lean_inc(x_33); +lean_inc(x_32); +x_52 = lean_alloc_ctor(0, 20, 1); +lean_ctor_set(x_52, 0, x_30); +lean_ctor_set(x_52, 1, x_32); +lean_ctor_set(x_52, 2, x_33); +lean_ctor_set(x_52, 3, x_34); +lean_ctor_set(x_52, 4, x_35); +lean_ctor_set(x_52, 5, x_36); +lean_ctor_set(x_52, 6, x_38); +lean_ctor_set(x_52, 7, x_39); +lean_ctor_set(x_52, 8, x_40); +lean_ctor_set(x_52, 9, x_41); +lean_ctor_set(x_52, 10, x_42); +lean_ctor_set(x_52, 11, x_43); +lean_ctor_set(x_52, 12, x_44); +lean_ctor_set(x_52, 13, x_45); +lean_ctor_set(x_52, 14, x_46); +lean_ctor_set(x_52, 15, x_47); +lean_ctor_set(x_52, 16, x_48); +lean_ctor_set(x_52, 17, x_49); +lean_ctor_set(x_52, 18, x_50); +lean_ctor_set(x_52, 19, x_51); +lean_ctor_set_uint8(x_52, sizeof(void*)*20, x_37); +x_53 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_53, 0, x_52); +lean_ctor_set(x_53, 1, x_3); +x_2 = x_31; +x_3 = x_53; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_splitNext___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; uint8_t x_11; +x_10 = lean_st_mk_ref(x_1, x_9); +x_11 = !lean_is_exclusive(x_10); +if (x_11 == 0) +{ +return x_10; +} +else +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_12 = lean_ctor_get(x_10, 0); +x_13 = lean_ctor_get(x_10, 1); +lean_inc(x_13); +lean_inc(x_12); +lean_dec(x_10); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_12); +lean_ctor_set(x_14, 1, x_13); +return x_14; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_splitNext___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_12 = lean_st_ref_get(x_3, x_11); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_12, 1); +lean_inc(x_14); +lean_dec(x_12); +x_15 = lean_box(0); +x_16 = l_List_mapTR_loop___at_Lean_Meta_Grind_splitNext___spec__1(x_13, x_2, x_15); +lean_dec(x_13); +x_17 = lean_unsigned_to_nat(1u); +x_18 = lean_nat_add(x_1, x_17); +x_19 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_introNewHyp(x_16, x_15, x_18, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_14); +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; +x_21 = lean_ctor_get(x_19, 0); +x_22 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_19, 0, x_22); +return x_19; +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +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_dec(x_19); +x_25 = lean_alloc_ctor(1, 1, 0); +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_24); +return x_26; +} +} +else +{ +uint8_t x_27; +x_27 = !lean_is_exclusive(x_19); +if (x_27 == 0) +{ +return x_19; +} +else +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_28 = lean_ctor_get(x_19, 0); +x_29 = lean_ctor_get(x_19, 1); +lean_inc(x_29); +lean_inc(x_28); +lean_dec(x_19); +x_30 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_30, 0, x_28); +lean_ctor_set(x_30, 1, x_29); +return x_30; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_splitNext___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; lean_object* x_14; uint8_t x_15; +x_13 = l_Lean_Meta_isMatcherApp___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate___spec__1(x_1, 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_16; lean_object* x_17; +x_16 = lean_ctor_get(x_13, 1); +lean_inc(x_16); +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_6); +lean_inc(x_5); +lean_inc(x_4); +x_17 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_16); +if (lean_obj_tag(x_17) == 0) +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +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_4, x_19); +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); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +x_24 = l_Lean_Meta_Grind_cases(x_23, x_18, x_8, x_9, x_10, x_11, x_22); +if (lean_obj_tag(x_24) == 0) +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_25 = lean_ctor_get(x_24, 0); +lean_inc(x_25); +x_26 = lean_ctor_get(x_24, 1); +lean_inc(x_26); +lean_dec(x_24); +x_27 = l_Lean_Meta_Grind_splitNext___lambda__2(x_2, x_25, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_26); +lean_dec(x_4); +return x_27; +} +else +{ +uint8_t x_28; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_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 +{ +uint8_t x_32; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_32 = !lean_is_exclusive(x_17); +if (x_32 == 0) +{ +return x_17; +} +else +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_33 = lean_ctor_get(x_17, 0); +x_34 = lean_ctor_get(x_17, 1); +lean_inc(x_34); +lean_inc(x_33); +lean_dec(x_17); +x_35 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_35, 0, x_33); +lean_ctor_set(x_35, 1, x_34); +return x_35; +} +} +} +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 = lean_ctor_get(x_13, 1); +lean_inc(x_36); +lean_dec(x_13); +x_37 = lean_st_ref_get(x_4, x_36); +x_38 = lean_ctor_get(x_37, 0); +lean_inc(x_38); +x_39 = lean_ctor_get(x_37, 1); +lean_inc(x_39); +lean_dec(x_37); +x_40 = lean_ctor_get(x_38, 0); +lean_inc(x_40); +lean_dec(x_38); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +x_41 = l_Lean_Meta_Grind_casesMatch(x_40, x_1, x_8, x_9, x_10, x_11, x_39); +if (lean_obj_tag(x_41) == 0) +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_42 = lean_ctor_get(x_41, 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_Grind_splitNext___lambda__2(x_2, x_42, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_43); +lean_dec(x_4); +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); +x_45 = !lean_is_exclusive(x_41); +if (x_45 == 0) +{ +return x_41; +} +else +{ +lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_46 = lean_ctor_get(x_41, 0); +x_47 = lean_ctor_get(x_41, 1); +lean_inc(x_47); +lean_inc(x_46); +lean_dec(x_41); +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_Grind_splitNext___lambda__4___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__3; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind_splitNext___lambda__4___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(", generation: ", 14, 14); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_splitNext___lambda__4___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_splitNext___lambda__4___closed__2; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_splitNext___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; lean_object* x_11; lean_object* x_21; +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_21 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_selectNextSplit_x3f(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_21) == 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_23; lean_object* 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); +lean_dec(x_2); +x_23 = lean_ctor_get(x_21, 1); +lean_inc(x_23); +lean_dec(x_21); +x_24 = lean_box(0); +x_10 = x_24; +x_11 = x_23; +goto block_20; +} +else +{ +lean_object* x_25; uint8_t x_26; +x_25 = lean_ctor_get(x_21, 1); +lean_inc(x_25); +lean_dec(x_21); +x_26 = !lean_is_exclusive(x_22); +if (x_26 == 0) +{ +lean_object* x_27; lean_object* x_28; +x_27 = lean_ctor_get(x_22, 0); +lean_inc(x_27); +x_28 = l_Lean_Meta_Grind_getGeneration(x_27, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_25); +if (lean_obj_tag(x_28) == 0) +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; 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 = l_Lean_Meta_Grind_splitNext___lambda__4___closed__1; +x_32 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_31, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_30); +x_33 = lean_ctor_get(x_32, 0); +lean_inc(x_33); +x_34 = lean_unbox(x_33); +lean_dec(x_33); +if (x_34 == 0) +{ +lean_object* x_35; lean_object* x_36; lean_object* x_37; +lean_free_object(x_22); +x_35 = lean_ctor_get(x_32, 1); +lean_inc(x_35); +lean_dec(x_32); +x_36 = lean_box(0); +lean_inc(x_1); +x_37 = l_Lean_Meta_Grind_splitNext___lambda__3(x_27, x_29, x_36, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_35); +lean_dec(x_29); +if (lean_obj_tag(x_37) == 0) +{ +lean_object* x_38; lean_object* x_39; +x_38 = lean_ctor_get(x_37, 0); +lean_inc(x_38); +x_39 = lean_ctor_get(x_37, 1); +lean_inc(x_39); +lean_dec(x_37); +x_10 = x_38; +x_11 = x_39; +goto block_20; +} +else +{ +uint8_t x_40; +lean_dec(x_1); +x_40 = !lean_is_exclusive(x_37); +if (x_40 == 0) +{ +return x_37; +} +else +{ +lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_41 = lean_ctor_get(x_37, 0); +x_42 = lean_ctor_get(x_37, 1); +lean_inc(x_42); +lean_inc(x_41); +lean_dec(x_37); +x_43 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_43, 0, x_41); +lean_ctor_set(x_43, 1, x_42); +return x_43; +} +} +} +else +{ +uint8_t x_44; +x_44 = !lean_is_exclusive(x_32); +if (x_44 == 0) +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_45 = lean_ctor_get(x_32, 1); +x_46 = lean_ctor_get(x_32, 0); +lean_dec(x_46); +x_47 = l_Lean_Meta_Grind_updateLastTag(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_45); +if (lean_obj_tag(x_47) == 0) +{ +lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; 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_47, 1); +lean_inc(x_48); +lean_dec(x_47); +lean_inc(x_27); +x_49 = l_Lean_MessageData_ofExpr(x_27); +x_50 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__9; +lean_ctor_set_tag(x_32, 7); +lean_ctor_set(x_32, 1, x_49); +lean_ctor_set(x_32, 0, x_50); +x_51 = l_Lean_Meta_Grind_splitNext___lambda__4___closed__3; +x_52 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_52, 0, x_32); +lean_ctor_set(x_52, 1, x_51); +lean_inc(x_29); +x_53 = l___private_Init_Data_Repr_0__Nat_reprFast(x_29); +lean_ctor_set_tag(x_22, 3); +lean_ctor_set(x_22, 0, x_53); +x_54 = l_Lean_MessageData_ofFormat(x_22); +x_55 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_55, 0, x_52); +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_50); +x_57 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_31, x_56, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, 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); +lean_dec(x_57); +lean_inc(x_1); +x_60 = l_Lean_Meta_Grind_splitNext___lambda__3(x_27, x_29, x_58, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_59); +lean_dec(x_58); +lean_dec(x_29); +if (lean_obj_tag(x_60) == 0) +{ +lean_object* x_61; lean_object* x_62; +x_61 = lean_ctor_get(x_60, 0); +lean_inc(x_61); +x_62 = lean_ctor_get(x_60, 1); +lean_inc(x_62); +lean_dec(x_60); +x_10 = x_61; +x_11 = x_62; +goto block_20; +} +else +{ +uint8_t x_63; +lean_dec(x_1); +x_63 = !lean_is_exclusive(x_60); +if (x_63 == 0) +{ +return x_60; +} +else +{ +lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_64 = lean_ctor_get(x_60, 0); +x_65 = lean_ctor_get(x_60, 1); +lean_inc(x_65); +lean_inc(x_64); +lean_dec(x_60); +x_66 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_66, 0, x_64); +lean_ctor_set(x_66, 1, x_65); +return x_66; +} +} +} +else +{ +uint8_t x_67; +lean_free_object(x_32); +lean_dec(x_29); +lean_free_object(x_22); +lean_dec(x_27); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_67 = !lean_is_exclusive(x_47); +if (x_67 == 0) +{ +return x_47; +} +else +{ +lean_object* x_68; lean_object* x_69; lean_object* x_70; +x_68 = lean_ctor_get(x_47, 0); +x_69 = lean_ctor_get(x_47, 1); +lean_inc(x_69); +lean_inc(x_68); +lean_dec(x_47); +x_70 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_70, 0, x_68); +lean_ctor_set(x_70, 1, x_69); +return x_70; +} +} +} +else +{ +lean_object* x_71; lean_object* x_72; +x_71 = lean_ctor_get(x_32, 1); +lean_inc(x_71); +lean_dec(x_32); +x_72 = l_Lean_Meta_Grind_updateLastTag(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_71); +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; 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_73 = lean_ctor_get(x_72, 1); +lean_inc(x_73); +lean_dec(x_72); +lean_inc(x_27); +x_74 = l_Lean_MessageData_ofExpr(x_27); +x_75 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__9; +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_Meta_Grind_splitNext___lambda__4___closed__3; +x_78 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_78, 0, x_76); +lean_ctor_set(x_78, 1, x_77); +lean_inc(x_29); +x_79 = l___private_Init_Data_Repr_0__Nat_reprFast(x_29); +lean_ctor_set_tag(x_22, 3); +lean_ctor_set(x_22, 0, x_79); +x_80 = l_Lean_MessageData_ofFormat(x_22); +x_81 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_81, 0, x_78); +lean_ctor_set(x_81, 1, x_80); +x_82 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_82, 0, x_81); +lean_ctor_set(x_82, 1, x_75); +x_83 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_31, x_82, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_73); +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_1); +x_86 = l_Lean_Meta_Grind_splitNext___lambda__3(x_27, x_29, x_84, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_85); +lean_dec(x_84); +lean_dec(x_29); +if (lean_obj_tag(x_86) == 0) +{ +lean_object* x_87; lean_object* x_88; +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_10 = x_87; +x_11 = x_88; +goto block_20; +} +else +{ +lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; +lean_dec(x_1); +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_86); + 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; +} +} +else +{ +lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; +lean_dec(x_29); +lean_free_object(x_22); +lean_dec(x_27); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_93 = lean_ctor_get(x_72, 0); +lean_inc(x_93); +x_94 = lean_ctor_get(x_72, 1); +lean_inc(x_94); +if (lean_is_exclusive(x_72)) { + lean_ctor_release(x_72, 0); + lean_ctor_release(x_72, 1); + x_95 = x_72; +} else { + lean_dec_ref(x_72); + 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_free_object(x_22); +lean_dec(x_27); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_97 = !lean_is_exclusive(x_28); +if (x_97 == 0) +{ +return x_28; +} +else +{ +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_101; lean_object* x_102; +x_101 = lean_ctor_get(x_22, 0); +lean_inc(x_101); +lean_dec(x_22); +lean_inc(x_101); +x_102 = l_Lean_Meta_Grind_getGeneration(x_101, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_25); +if (lean_obj_tag(x_102) == 0) +{ +lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; uint8_t x_108; +x_103 = lean_ctor_get(x_102, 0); +lean_inc(x_103); +x_104 = lean_ctor_get(x_102, 1); +lean_inc(x_104); +lean_dec(x_102); +x_105 = l_Lean_Meta_Grind_splitNext___lambda__4___closed__1; +x_106 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_105, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_104); +x_107 = lean_ctor_get(x_106, 0); +lean_inc(x_107); +x_108 = lean_unbox(x_107); +lean_dec(x_107); +if (x_108 == 0) +{ +lean_object* x_109; lean_object* x_110; lean_object* x_111; +x_109 = lean_ctor_get(x_106, 1); +lean_inc(x_109); +lean_dec(x_106); +x_110 = lean_box(0); +lean_inc(x_1); +x_111 = l_Lean_Meta_Grind_splitNext___lambda__3(x_101, x_103, x_110, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_109); +lean_dec(x_103); +if (lean_obj_tag(x_111) == 0) +{ +lean_object* x_112; lean_object* x_113; +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_10 = x_112; +x_11 = x_113; +goto block_20; +} +else +{ +lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; +lean_dec(x_1); +x_114 = lean_ctor_get(x_111, 0); +lean_inc(x_114); +x_115 = lean_ctor_get(x_111, 1); +lean_inc(x_115); +if (lean_is_exclusive(x_111)) { + lean_ctor_release(x_111, 0); + lean_ctor_release(x_111, 1); + x_116 = x_111; +} else { + lean_dec_ref(x_111); + x_116 = lean_box(0); +} +if (lean_is_scalar(x_116)) { + x_117 = lean_alloc_ctor(1, 2, 0); +} else { + x_117 = x_116; +} +lean_ctor_set(x_117, 0, x_114); +lean_ctor_set(x_117, 1, x_115); +return x_117; +} +} +else +{ +lean_object* x_118; lean_object* x_119; lean_object* x_120; +x_118 = lean_ctor_get(x_106, 1); +lean_inc(x_118); +if (lean_is_exclusive(x_106)) { + lean_ctor_release(x_106, 0); + lean_ctor_release(x_106, 1); + x_119 = x_106; +} else { + lean_dec_ref(x_106); + x_119 = lean_box(0); +} +x_120 = l_Lean_Meta_Grind_updateLastTag(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_118); +if (lean_obj_tag(x_120) == 0) +{ +lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; +x_121 = lean_ctor_get(x_120, 1); +lean_inc(x_121); +lean_dec(x_120); +lean_inc(x_101); +x_122 = l_Lean_MessageData_ofExpr(x_101); +x_123 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__9; +if (lean_is_scalar(x_119)) { + x_124 = lean_alloc_ctor(7, 2, 0); +} else { + x_124 = x_119; + lean_ctor_set_tag(x_124, 7); +} +lean_ctor_set(x_124, 0, x_123); +lean_ctor_set(x_124, 1, x_122); +x_125 = l_Lean_Meta_Grind_splitNext___lambda__4___closed__3; +x_126 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_126, 0, x_124); +lean_ctor_set(x_126, 1, x_125); +lean_inc(x_103); +x_127 = l___private_Init_Data_Repr_0__Nat_reprFast(x_103); +x_128 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_128, 0, x_127); +x_129 = l_Lean_MessageData_ofFormat(x_128); +x_130 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_130, 0, x_126); +lean_ctor_set(x_130, 1, x_129); +x_131 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_131, 0, x_130); +lean_ctor_set(x_131, 1, x_123); +x_132 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_105, x_131, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_121); +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); +lean_inc(x_1); +x_135 = l_Lean_Meta_Grind_splitNext___lambda__3(x_101, x_103, x_133, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_134); +lean_dec(x_133); +lean_dec(x_103); +if (lean_obj_tag(x_135) == 0) +{ +lean_object* x_136; lean_object* x_137; +x_136 = lean_ctor_get(x_135, 0); +lean_inc(x_136); +x_137 = lean_ctor_get(x_135, 1); +lean_inc(x_137); +lean_dec(x_135); +x_10 = x_136; +x_11 = x_137; +goto block_20; +} +else +{ +lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; +lean_dec(x_1); +x_138 = lean_ctor_get(x_135, 0); +lean_inc(x_138); +x_139 = lean_ctor_get(x_135, 1); +lean_inc(x_139); +if (lean_is_exclusive(x_135)) { + lean_ctor_release(x_135, 0); + lean_ctor_release(x_135, 1); + x_140 = x_135; +} else { + lean_dec_ref(x_135); + x_140 = lean_box(0); +} +if (lean_is_scalar(x_140)) { + x_141 = lean_alloc_ctor(1, 2, 0); +} else { + x_141 = x_140; +} +lean_ctor_set(x_141, 0, x_138); +lean_ctor_set(x_141, 1, x_139); +return x_141; +} +} +else +{ +lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; +lean_dec(x_119); +lean_dec(x_103); +lean_dec(x_101); +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_142 = lean_ctor_get(x_120, 0); +lean_inc(x_142); +x_143 = lean_ctor_get(x_120, 1); +lean_inc(x_143); +if (lean_is_exclusive(x_120)) { + lean_ctor_release(x_120, 0); + lean_ctor_release(x_120, 1); + x_144 = x_120; +} else { + lean_dec_ref(x_120); + 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; +} +} +} +else +{ +lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; +lean_dec(x_101); +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_146 = lean_ctor_get(x_102, 0); +lean_inc(x_146); +x_147 = lean_ctor_get(x_102, 1); +lean_inc(x_147); +if (lean_is_exclusive(x_102)) { + lean_ctor_release(x_102, 0); + lean_ctor_release(x_102, 1); + x_148 = x_102; +} else { + lean_dec_ref(x_102); + x_148 = lean_box(0); +} +if (lean_is_scalar(x_148)) { + x_149 = lean_alloc_ctor(1, 2, 0); +} else { + x_149 = x_148; +} +lean_ctor_set(x_149, 0, x_146); +lean_ctor_set(x_149, 1, x_147); +return x_149; +} +} +} +} +else +{ +uint8_t x_150; +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_150 = !lean_is_exclusive(x_21); +if (x_150 == 0) +{ +return x_21; +} +else +{ +lean_object* x_151; lean_object* x_152; lean_object* x_153; +x_151 = lean_ctor_get(x_21, 0); +x_152 = lean_ctor_get(x_21, 1); +lean_inc(x_152); +lean_inc(x_151); +lean_dec(x_21); +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; +} +} +block_20: +{ +lean_object* x_12; uint8_t x_13; +x_12 = lean_st_ref_get(x_1, x_11); +lean_dec(x_1); +x_13 = !lean_is_exclusive(x_12); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; +x_14 = lean_ctor_get(x_12, 0); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_10); +lean_ctor_set(x_15, 1, x_14); +lean_ctor_set(x_12, 0, x_15); +return x_12; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_16 = lean_ctor_get(x_12, 0); +x_17 = lean_ctor_get(x_12, 1); +lean_inc(x_17); +lean_inc(x_16); +lean_dec(x_12); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_10); +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_17); +return x_19; +} +} +} +} +static lean_object* _init_l_Lean_Meta_Grind_splitNext___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_splitNext___lambda__4), 9, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_splitNext(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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 = lean_ctor_get(x_1, 0); +lean_inc(x_10); +x_11 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_splitNext___lambda__1___boxed), 9, 1); +lean_closure_set(x_11, 0, x_1); +x_12 = l_Lean_Meta_Grind_splitNext___closed__1; +x_13 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); +lean_closure_set(x_13, 0, x_11); +lean_closure_set(x_13, 1, x_12); +x_14 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(x_10, x_13, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_14) == 0) +{ +uint8_t x_15; +x_15 = !lean_is_exclusive(x_14); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; +x_16 = lean_ctor_get(x_14, 0); +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +lean_dec(x_16); +lean_ctor_set(x_14, 0, x_17); +return x_14; +} +else +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_18 = lean_ctor_get(x_14, 0); +x_19 = lean_ctor_get(x_14, 1); +lean_inc(x_19); +lean_inc(x_18); +lean_dec(x_14); +x_20 = lean_ctor_get(x_18, 0); +lean_inc(x_20); +lean_dec(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; +} +} +else +{ +uint8_t x_22; +x_22 = !lean_is_exclusive(x_14); +if (x_22 == 0) +{ +return x_14; +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_14, 0); +x_24 = lean_ctor_get(x_14, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_14); +x_25 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_24); +return x_25; +} +} +} +} +LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Meta_Grind_splitNext___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_List_mapTR_loop___at_Lean_Meta_Grind_splitNext___spec__1(x_1, x_2, x_3); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_splitNext___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l_Lean_Meta_Grind_splitNext___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_splitNext___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_Meta_Grind_splitNext___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_3); +lean_dec(x_1); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_splitNext___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +x_13 = l_Lean_Meta_Grind_splitNext___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_3); +lean_dec(x_2); +return x_13; +} +} +lean_object* initialize_Lean_Meta_Tactic_Grind_Types(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Meta_Tactic_Grind_Intro(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Meta_Tactic_Grind_Cases(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Meta_Tactic_Grind_CasesMatch(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Lean_Meta_Tactic_Grind_Split(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_Tactic_Grind_Types(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Lean_Meta_Tactic_Grind_Intro(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Lean_Meta_Tactic_Grind_Cases(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Lean_Meta_Tactic_Grind_CasesMatch(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l_Lean_Meta_Grind_CaseSplitStatus_noConfusion___rarg___closed__1 = _init_l_Lean_Meta_Grind_CaseSplitStatus_noConfusion___rarg___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_CaseSplitStatus_noConfusion___rarg___closed__1); +l_Lean_Meta_Grind_instInhabitedCaseSplitStatus = _init_l_Lean_Meta_Grind_instInhabitedCaseSplitStatus(); +l_Lean_Meta_Grind_instBEqCaseSplitStatus___closed__1 = _init_l_Lean_Meta_Grind_instBEqCaseSplitStatus___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_instBEqCaseSplitStatus___closed__1); +l_Lean_Meta_Grind_instBEqCaseSplitStatus = _init_l_Lean_Meta_Grind_instBEqCaseSplitStatus(); +lean_mark_persistent(l_Lean_Meta_Grind_instBEqCaseSplitStatus); +l_panic___at___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___spec__1___closed__1 = _init_l_panic___at___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___spec__1___closed__1(); +lean_mark_persistent(l_panic___at___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___spec__1___closed__1); +l_panic___at___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___spec__1___closed__2 = _init_l_panic___at___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___spec__1___closed__2(); +lean_mark_persistent(l_panic___at___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___spec__1___closed__2); +l_panic___at___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___spec__1___closed__3 = _init_l_panic___at___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___spec__1___closed__3(); +lean_mark_persistent(l_panic___at___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___spec__1___closed__3); +l_panic___at___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___spec__1___closed__4 = _init_l_panic___at___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___spec__1___closed__4(); +lean_mark_persistent(l_panic___at___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___spec__1___closed__4); +l_panic___at___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___spec__1___closed__5 = _init_l_panic___at___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___spec__1___closed__5(); +lean_mark_persistent(l_panic___at___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___spec__1___closed__5); +l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__5___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__5___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__5___closed__1); +l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__5___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__5___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__5___closed__2); +l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__5___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__5___closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__5___closed__3); +l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__5___closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__5___closed__4(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__5___closed__4); +l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__5___closed__5 = _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__5___closed__5(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__5___closed__5); +l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__1); +l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__2); +l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__3); +l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__4(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__4); +l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__5 = _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__5(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__5); +l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__6 = _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__6(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__6); +l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__7 = _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__7(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__7); +l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__8 = _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__8(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__8); +l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__9 = _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__9(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__9); +l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__1); +l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__2); +l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__3); +l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__4(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__4); +l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__5 = _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__5(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__5); +l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__6 = _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__6(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__6); +l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__7 = _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__7(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__7); +l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__8 = _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__8(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__8); +l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__9 = _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__9(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__9); +l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_selectNextSplit_x3f___lambda__2___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_selectNextSplit_x3f___lambda__2___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_selectNextSplit_x3f___lambda__2___closed__1); +l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_selectNextSplit_x3f___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_selectNextSplit_x3f___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_selectNextSplit_x3f___closed__1); +l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__2___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__2___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__2___closed__1); +l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__2___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__2___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__2___closed__2); +l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__2___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__2___closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__2___closed__3); +l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__2___closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__2___closed__4(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__2___closed__4); +l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__2___closed__5 = _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__2___closed__5(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__2___closed__5); +l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__3___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__3___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__3___closed__1); +l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__3___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__3___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__3___closed__2); +l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__3___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__3___closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__3___closed__3); +l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___closed__1); +l_Lean_Meta_Grind_splitNext___lambda__4___closed__1 = _init_l_Lean_Meta_Grind_splitNext___lambda__4___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_splitNext___lambda__4___closed__1); +l_Lean_Meta_Grind_splitNext___lambda__4___closed__2 = _init_l_Lean_Meta_Grind_splitNext___lambda__4___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_splitNext___lambda__4___closed__2); +l_Lean_Meta_Grind_splitNext___lambda__4___closed__3 = _init_l_Lean_Meta_Grind_splitNext___lambda__4___closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind_splitNext___lambda__4___closed__3); +l_Lean_Meta_Grind_splitNext___closed__1 = _init_l_Lean_Meta_Grind_splitNext___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_splitNext___closed__1); +return lean_io_result_mk_ok(lean_box(0)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Lean/Meta/Tactic/Grind/Types.c b/stage0/stdlib/Lean/Meta/Tactic/Grind/Types.c index 89336ad917da..cd9abe19d6cb 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Grind/Types.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Grind/Types.c @@ -13,41 +13,60 @@ #ifdef __cplusplus extern "C" { #endif +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__2; +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__52; static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1___closed__1; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_registerParent___spec__7(lean_object*, lean_object*); lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_instInhabitedMethods___closed__2; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__60; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__8; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GoalM_run___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__43; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__36; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instInhabitedMethods; -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__27; LEAN_EXPORT lean_object* l_Lean_RBNode_ins___at_Lean_Meta_Grind_registerParent___spec__2(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_instInhabitedENode___closed__1; +static lean_object* l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__4; +static lean_object* l_Lean_Meta_Grind_updateLastTag___closed__1; +lean_object* l_Lean_Syntax_mkNameLit(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instHashablePreInstance_unsafe__1___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getEqc_go___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pushEq(lean_object*, lean_object*, lean_object*, 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_insert___at_Lean_Meta_Grind_setENode___spec__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instReprENode; +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__21; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_forEachENode___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_markTheoremInstance___spec__4(size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isEqv___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_Meta_Grind_getENode___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_Grind_doElemTrace__goal_x5b___x5d_______closed__23; lean_object* lean_mk_empty_array_with_capacity(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addNewFact___boxed(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_shift_right(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pushEqHEq___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__16; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ENode_isCongrRoot___boxed(lean_object*); +static lean_object* l_Lean_Meta_Grind_markAsInconsistent___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_abstractNestedProofs(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__14; LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___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_Meta_Grind_getTrueExpr___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__47; +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__1; static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_118____closed__4; static size_t l_Lean_Meta_Grind_instInhabitedGoal___closed__3; +LEAN_EXPORT lean_object* l_ReaderT_pure___at_Lean_Meta_Grind_instInhabitedMethods___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_canon___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_Grind_congrHash_goEq___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getTarget_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_updateLastTag___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_Grind_abstractNestedProofs___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getConfig(lean_object*); +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__27; +static lean_object* l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__3; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_markCaseSplitAsResolved___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_registerParent___spec__4(lean_object*, size_t, size_t, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__45; uint8_t lean_usize_dec_le(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedGoal___spec__2; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getMethodsRef(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -58,104 +77,139 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Grind_filterENodes___lambda__1(lean_object* static lean_object* l_Lean_Meta_Grind_instBEqPreInstance___lambda__2___closed__1; uint64_t lean_uint64_of_nat(lean_object*); uint64_t lean_uint64_mix_hash(uint64_t, uint64_t); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__15; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GoalM_run(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_Array_qsort_sort___at_Lean_Meta_Grind_getENodes___spec__4___lambda__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_setENode_unsafe__1___rarg(lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__46; size_t lean_uint64_to_usize(uint64_t); uint64_t lean_uint64_lor(uint64_t, uint64_t); -uint8_t l_Lean_Expr_isAppOfArity(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isTrueExpr___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_Grind_instInhabitedGoal___closed__5; static lean_object* l_Lean_Meta_Grind_getENode___closed__2; +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__22; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkENode(lean_object*, lean_object*, 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_Grind_getMainDeclName___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GoalM_run_x27(lean_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_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__53; +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__24; uint8_t l_Lean_Expr_isApp(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_118_(lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__24; +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__13; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markAsInconsistent___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 uint8_t l_Lean_Meta_Grind_isCongruent(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Syntax_getId(lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__50; uint8_t l_Lean_RBNode_isRed___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isInterpreted(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*); +static lean_object* l_Lean_Meta_Grind_markAsInconsistent___closed__1; lean_object* l_Array_toSubarray___rarg(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__20; +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__7; size_t lean_usize_mul(size_t, size_t); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__7; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_alreadyInternalized___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); lean_object* lean_mk_array(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkHCongrWithArity___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t lean_uint64_dec_lt(uint64_t, uint64_t); +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__5; uint8_t lean_usize_dec_eq(size_t, size_t); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__55; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pushEq___boxed(lean_object*, lean_object*, lean_object*, 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_insertAux___at_Lean_Meta_Grind_markCaseSplitAsResolved___spec__2(lean_object*, size_t, size_t, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__12; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__4(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isSameExpr_unsafe__1___boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__44; LEAN_EXPORT uint64_t l_Lean_Meta_Grind_congrHash(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__3; lean_object* l_Lean_Expr_bvar___override(lean_object*); lean_object* lean_array_fget(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getMaxGeneration___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getFalseExpr___boxed(lean_object*, lean_object*); lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__14; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_grind_debug_proofs; +LEAN_EXPORT lean_object* l_ReaderT_pure___at_Lean_Meta_Grind_instInhabitedMethods___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_Meta_Grind_registerParent___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_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__8; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markCaseSplitAsResolved___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_Grind_doElemTrace__goal_x5b___x5d_______closed__24; +lean_object* l_Lean_MVarId_getTag(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pushEqCore(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_Grind_forEachENode(lean_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_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__10; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateDown(lean_object*, 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___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot___spec__2(lean_object*, size_t, lean_object*); static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78____closed__8; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_forEachEqc(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_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__37; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getCongrRoot___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_PersistentHashMap_containsAux___at_Lean_Meta_Grind_alreadyInternalized___spec__2___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getEqcs(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Std_Queue_enqueue___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instBEqPreInstance___lambda__2(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Lean_Meta_Grind_hasSameRoot(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint64_t l_Lean_Meta_Grind_instHashableCongrTheoremCacheKey(lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_registerParent___spec__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__8; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markAsInconsistent___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_node5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_setENode(lean_object*, 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_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__19; +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*); size_t lean_usize_of_nat(lean_object*); lean_object* l_Nat_nextPowerOfTwo_go(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getMethods___boxed(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_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__51; lean_object* l_Lean_stringToMessageData(lean_object*); static lean_object* l_Lean_Meta_Grind_instInhabitedNewFact___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getMainDeclName___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__13; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_mkHCongrWithArity___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkHCongrWithArity___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__26; +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__36; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getFalseExpr(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_markCaseSplitAsResolved___spec__4(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instInhabitedMethods___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_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__33; +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__41; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__28; LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at_Lean_Meta_Grind_copyParentsTo___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_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_registerParent___spec__6(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_registerParent___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2(lean_object*); +LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___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_Grind_mkENodeCore___boxed(lean_object*, lean_object*, lean_object*, lean_object*, 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_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3(lean_object*, size_t, size_t, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_markTheoremInstance___spec__13___lambda__1(lean_object*, lean_object*, lean_object*); static uint64_t l_Lean_Meta_Grind_hasSameType___closed__1; -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__9; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getCongrRoot___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_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_markCaseSplitAsResolved___spec__3(size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instBEqCongrKey___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instBEqCongrTheoremCacheKey___boxed(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_isResolvedCaseSplit___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Methods_toMethodsRef___boxed(lean_object*); +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__54; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_getParentsAndReset___spec__2___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__38; +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__11; LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_alreadyInternalized___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_setENode___spec__4(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__2; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__13; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getEqc___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_Grind_isCongrRoot(lean_object*, 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_Grind_grind_debug; LEAN_EXPORT lean_object* l_Lean_MVarId_assign___at_Lean_Meta_Grind_closeGoal___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_MethodsRef_toMethods(lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__2; +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__4; lean_object* l_Lean_Expr_appArg_x21(lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_AbstractNestedProofs_visit(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_Grind_propagateUp(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_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__15; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__5; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isResolvedCaseSplit___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__21; static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_118____closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isEqTrue___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint64_t l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instHashablePreInstance___spec__1(lean_object*, lean_object*, lean_object*, size_t, size_t, uint64_t); @@ -167,15 +221,15 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getRoot(lean_object*, lean_object*, l LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEqProof___boxed(lean_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_Queue_empty(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_markTheoremInstance___spec__1(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__52; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_checkMaxEmatchExceeded(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isLambda(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instBEqPreInstance___lambda__3___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Meta_Grind_getENodes___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instBEqPreInstance___boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_updateLastTag___closed__5; size_t lean_ptr_addr(lean_object*); lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_isAssigned___at_Lean_Meta_Grind_closeGoal___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_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pushEqCore___boxed(lean_object*, lean_object*, lean_object*, lean_object*, 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_Grind_GoalM_run___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_Grind_getFalseExpr___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -183,18 +237,22 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Methods_toMethodsRef_unsafe__1___boxe LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getTarget_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_isUnaryNode___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_registerParent___spec__8(lean_object*, size_t, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_checkMaxEmatchExceeded___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_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__48; static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78____closed__1; +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__14; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_setENode_unsafe__1___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__27; +static lean_object* l_Lean_Meta_Grind_markCaseSplitAsResolved___closed__3; LEAN_EXPORT uint8_t l_Lean_Meta_Grind_instBEqPreInstance___lambda__1(lean_object*); -static lean_object* l_Lean_Meta_Grind_getENode___closed__4; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markCaseSplitAsResolved(lean_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_Grind_doElemTrace__goal_x5b___x5d_______closed__15; +static lean_object* l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__17; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markTheoremInstance___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_Grind_closeGoal___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_Grind_getRootENode___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_take(lean_object*, lean_object*); -lean_object* l_Lean_Expr_getRevArg_x21(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_hasSameRoot___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isEqFalse___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__25; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getParentsAndReset___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_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78____closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isCongrRoot___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -202,14 +260,22 @@ LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_Grind_getENode___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_Grind_filterENodes___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_Meta_Grind_instBEqPreInstance___lambda__2___closed__2; -LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hasSameRoot(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2___closed__3; +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__30; +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__1; +static lean_object* l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__6; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instHashablePreInstance_unsafe__2___boxed(lean_object*); uint64_t lean_uint64_shift_right(uint64_t, uint64_t); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getGeneration___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_Grind_getCongrRoot(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_SourceInfo_fromRef(lean_object*, uint8_t); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__10; lean_object* l_Lean_RBNode_setBlack___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addTheoremInstance___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint64_t lean_usize_to_uint64(size_t); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__25; lean_object* lean_nat_to_int(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_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_registerParent___spec__8___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEqHEqProof(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MVarId_getType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -218,69 +284,96 @@ static lean_object* l_Lean_Meta_Grind_abstractNestedProofs___closed__1; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_forEachEqc___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_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pushEqHEq(lean_object*, lean_object*, lean_object*, 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_Grind_getENodes___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_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__16; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_checkMaxInstancesExceeded(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_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__29; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_closeGoal(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MVarId_admit(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_qsort_sort___at_Lean_Meta_Grind_getENodes___spec__4___closed__1; LEAN_EXPORT lean_object* l_Lean_MVarId_assign___at_Lean_Meta_Grind_closeGoal___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___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__13; -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__18; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__18; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getENode_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getConfig___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkHCongrWithArity___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_Grind_checkMaxInstancesExceeded___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_mkHCongrWithArityForConst_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__2; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_setENode___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__2(lean_object*, size_t, size_t, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__17; +static lean_object* l_Lean_Meta_Grind_updateLastTag___closed__4; +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__31; +lean_object* l_Lean_Syntax_getKind(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3___lambda__2___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_checkMaxCaseSplit(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_markTheoremInstance___spec__11___lambda__2(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__14(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78____closed__6; +lean_object* l_Lean_Expr_appArg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markCaseSplitAsResolved___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint64_t l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__5(lean_object*, lean_object*, lean_object*, size_t, size_t, uint64_t); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getENode___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_Grind_getTrueExpr___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_quoteNameMk(lean_object*); +LEAN_EXPORT uint64_t l_Lean_Meta_Grind_congrHash_goEq(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__49; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkENode___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_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__44; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_markTheoremInstance___spec__13___lambda__2___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__58; lean_object* lean_st_ref_get(lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_foldlMAux___at_Lean_MetavarContext_getExprAssignmentDomain___spec__2___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_congrPlaceholderProof___closed__1; -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__48; +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__56; +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_isResolvedCaseSplit___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__28; +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__14; lean_object* lean_st_mk_ref(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_markTheoremInstance___spec__7(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__46; +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__19; +lean_object* l_Lean_Expr_appFnCleanup(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isCongruent___boxed(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*); LEAN_EXPORT lean_object* l_Lean_MVarId_isAssigned___at_Lean_Meta_Grind_closeGoal___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_Grind_getMainDeclName(lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__21; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__35; static lean_object* l_Lean_Meta_Grind_canon___closed__3; -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__20; +lean_object* l_Lean_Syntax_node3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint64_t l_Lean_Meta_Grind_instHashablePreInstance(lean_object*); +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__16; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isCongruent_goEq___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_Grind_getENodes___spec__3(size_t, size_t, lean_object*); +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__33; LEAN_EXPORT uint64_t l_Lean_Meta_Grind_instHashableENodeKey(lean_object*); LEAN_EXPORT uint64_t l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_mkHCongrWithArity___spec__5(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__39; +static lean_object* l_Lean_Meta_Grind_markCaseSplitAsResolved___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_MethodsRef_toMethods_unsafe__1(lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__34; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isSameExpr___boxed(lean_object*, lean_object*); static size_t l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__2___closed__2; +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__39; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getEqc(lean_object*, 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_Grind_getMethodsRef___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_congrHash___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pushEqTrue(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_Meta_Grind_instBEqCongrTheoremCacheKey(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_markTheoremInstance___spec__9(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__11; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_setENode_unsafe__1___rarg___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkHCongrWithArity(lean_object*, 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_Grind_getEqcs___boxed(lean_object*, 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_Grind_mkHCongrWithArity___spec__2___closed__3; +lean_object* l_Lean_addMacroScope(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__50; LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at_Lean_Meta_Grind_copyParentsTo___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_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_registerParent___spec__7___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_canon___closed__1; static lean_object* l_Lean_Meta_Grind_congrHash___closed__1; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__20; +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_Grind_getGeneration(lean_object*, 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_Grind_instBEqENodeKey___boxed(lean_object*, lean_object*); @@ -288,8 +381,8 @@ lean_object* l_Lean_Meta_mkHCongrWithArity(lean_object*, lean_object*, lean_obje LEAN_EXPORT uint8_t l_Lean_Meta_Grind_isSameExpr_unsafe__1(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_118____closed__5; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_alreadyInternalized(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Syntax_node2(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_alreadyInternalized___spec__2(lean_object*, size_t, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__41; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isCongruent_go___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getNext___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_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_registerParent___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -297,81 +390,115 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_Grind_getENodes__ LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instBEqPreInstance___lambda__2___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getMethods(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getMainDeclName___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_filterENodes(lean_object*, 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_containsAtAux___at_Lean_Meta_Grind_markTheoremInstance___spec__13___lambda__1___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__8; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getParents(lean_object*, 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_insertAux_traverse___at_Lean_Meta_Grind_markTheoremInstance___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__5; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__7; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getRoot___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__30; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isInterpreted___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Grind_getENodes___spec__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markTheoremInstance___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_PersistentHashMap_foldlM___at_Lean_Meta_Grind_getENodes___spec__2(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot___spec__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Grind_getENodes___spec__1___closed__1; static lean_object* l_Lean_Meta_Grind_instInhabitedGoal___closed__1; lean_object* l_Lean_Expr_appFn_x21(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_applyFallback(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_Grind_updateLastTag(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__13; +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__3; LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_markTheoremInstance___spec__13(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__19; lean_object* l_Array_eraseIdx___rarg(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__36; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__45; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_hasSameType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkHEqProof___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_Grind_instHashableCongrKey___boxed(lean_object*, lean_object*); LEAN_EXPORT uint64_t l_Lean_Meta_Grind_instHashablePreInstance_unsafe__1(lean_object*); -lean_object* l_Lean_Meta_Grind_Canon_canonImpl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markAsInconsistent___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_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__10; +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_isResolvedCaseSplit___spec__2(lean_object*, size_t, lean_object*); uint8_t l_Lean_PersistentHashMap_contains___at_Lean_MVarId_isAssigned___spec__1(lean_object*, lean_object*); lean_object* l_Lean_Meta_mkFalseElim(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_instInhabitedENode___closed__2; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__48; lean_object* lean_usize_to_nat(size_t); LEAN_EXPORT uint8_t l_Lean_Meta_Grind_isCongruent_go(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__18; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getParents___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_Grind_updateLastTag___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_Grind_alreadyInternalized___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_Grind_getENode___closed__3; LEAN_EXPORT uint64_t l_Lean_Meta_Grind_instHashableCongrKey(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getENodes(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_Grind_Canon_canon(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_MessageData_ofExpr(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_copyParentsTo(lean_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_register___at_Lean_Elab_initFn____x40_Lean_Elab_AutoBound___hyg_6____spec__1(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__14; +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__26; +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__4; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__53; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__52; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__1; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__14___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_instReprENode___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_shareCommon___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_Grind_GoalM_run___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__5; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__49; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_isResolvedCaseSplit___spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot___spec__1(lean_object*, lean_object*); LEAN_EXPORT uint64_t l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__10(lean_object*, lean_object*, lean_object*, size_t, size_t, uint64_t); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isResolvedCaseSplit(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +double l_Float_ofScientific(lean_object*, uint8_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_registerParent___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___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_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_118____closed__3; -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__11; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560_(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_markTheoremInstance___spec__7___lambda__2(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_GoalM_run_x27___closed__1; +static lean_object* l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Goal_admit(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getNext(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_Grind_ENode_isCongrRoot(lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3___lambda__2(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_setENode_unsafe__1(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_abstractNestedProofs___closed__3; -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__23; +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__15; LEAN_EXPORT uint64_t l_Lean_Meta_Grind_instHashablePreInstance_unsafe__2(lean_object*); +static lean_object* l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__16; lean_object* lean_grind_mk_eq_proof(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_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__29; lean_object* lean_grind_mk_heq_proof(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint64_t l_Lean_Meta_Grind_congrHash_go(lean_object*, lean_object*, uint64_t); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__35; +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__6; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instInhabitedMethods___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_PersistentHashMap_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instBEqPreInstance___lambda__1___boxed(lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__54; +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__34; LEAN_EXPORT lean_object* l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___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_Meta_Grind_addNewFact(lean_object*, lean_object*, 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_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__54; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_markTheoremInstance___spec__11___lambda__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_congrPlaceholderProof___closed__3; +LEAN_EXPORT lean_object* l_ReaderT_pure___at_Lean_Meta_Grind_instInhabitedMethods___spec__1(lean_object*); +lean_object* l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__16; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_markTheoremInstance___spec__7___lambda__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_alreadyInternalized___spec__1___boxed(lean_object*, lean_object*); lean_object* lean_string_length(lean_object*); +static lean_object* l_Lean_Meta_Grind_congrHash___closed__3; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__30; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_abstractNestedProofs___boxed(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_Grind_mkEqFalseProof(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -382,46 +509,68 @@ static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_mkHCongrWithArity___spec__3(size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addTheoremInstance(lean_object*, lean_object*, lean_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_Grind_Types_0__Lean_Meta_Grind_MethodsRef_toMethods___boxed(lean_object*); +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__42; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_markTheoremInstance___spec__13___lambda__2(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__5; -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__22; +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__32; +static lean_object* l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__22; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GoalM_run_x27___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_addTrace___at_Lean_Meta_Grind_updateLastTag___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_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_setENode___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_insert___at_Lean_MVarId_assign___spec__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot_unsafe__1___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isInconsistent___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_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__41; lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_instInhabitedMethods___closed__1; +LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___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_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedGoal___spec__3; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getCongrRoot___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_isConstructorAppCore_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__9; +static lean_object* l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getEqc_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_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_registerParent(lean_object*, lean_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_Grind_doElemTrace__goal_x5b___x5d_______closed__7; LEAN_EXPORT uint8_t l_Lean_Meta_Grind_isSameExpr(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_markAsInconsistent___closed__3; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__34; static lean_object* l_Lean_Meta_Grind_congrPlaceholderProof___closed__2; -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__45; +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__59; lean_object* l_Lean_indentExpr(lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__32; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getMaxGeneration___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_isResolvedCaseSplit___spec__3___boxed(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_Meta_Grind_instInhabitedENode; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_isResolvedCaseSplit___spec__2___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_registerParent___spec__3(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__23; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markAsInconsistent(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_containsAtAux___at_Lean_Meta_Grind_markTheoremInstance___spec__13___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_instInhabitedGoal___closed__2; +lean_object* l_Lean_Syntax_node1(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__18; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_markCaseSplitAsResolved___spec__1(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__7; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markTheoremInstance(lean_object*, lean_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_Grind_doElemTrace__goal_x5b___x5d_______closed__21; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkHCongrWithArity___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Expr_isConstOf(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Meta_Grind_instBEqCongrKey(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Grind_getENodes___spec__1___boxed(lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_markTheoremInstance___spec__7___lambda__1(lean_object*, lean_object*, lean_object*); lean_object* lean_array_set(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78____closed__3; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_markTheoremInstance___spec__9___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__31; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Grind_getENodes___spec__1(lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__40; 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_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__50; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__24; +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__2; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_getParentsAndReset___spec__2(lean_object*, size_t, lean_object*); +static lean_object* l_Lean_Meta_Grind_markCaseSplitAsResolved___closed__2; +static lean_object* l_Lean_Meta_Grind_instInhabitedGoal___closed__6; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isEqv(lean_object*, lean_object*, 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_Grind_pushEqFalse___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Expr_0__Lean_reprExpr____x40_Lean_Expr___hyg_3036_(lean_object*, lean_object*); @@ -432,67 +581,88 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Methods_toMethodsRef(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEqTrueProof(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_indexOfAux___at_Lean_Meta_Grind_getParentsAndReset___spec__3(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_indexOfAux___at_Lean_Meta_Grind_getParentsAndReset___spec__3___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__28; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isRoot___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_Grind_setENode___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_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__43; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_copyParentsTo___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint64_t lean_uint64_shift_left(uint64_t, uint64_t); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__6; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Methods_toMethodsRef_unsafe__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pushHEq(lean_object*, lean_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_Array_QSort_0__Array_qpartition___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getENode_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_mkCollisionNode___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getTrueExpr___boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__6; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getENode(lean_object*, 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_Grind_getTrueExpr(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__15; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__6; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkENodeCore(lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__44; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__42; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__11; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GoalM_run_x27___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_forEachENode___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_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isTrueExpr(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_erase_macro_scopes(lean_object*); +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__40; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markCaseSplitAsResolved___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot___spec__2___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__17; +static lean_object* l_Lean_Meta_Grind_congrHash___closed__4; +static lean_object* l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__9; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot___boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__10; +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__47; LEAN_EXPORT uint8_t l_Array_qsort_sort___at_Lean_Meta_Grind_getENodes___spec__4___lambda__1(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__42; +lean_object* l_String_intercalate(lean_object*, lean_object*); size_t lean_usize_sub(size_t, size_t); +static lean_object* l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__20; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instHashableENodeKey_unsafe__1___boxed(lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__22; lean_object* lean_array_mk(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getParentsAndReset(lean_object*, 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_Grind_isRoot(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_isTrue(lean_object*); uint8_t l_Lean_Expr_quickComp(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getRoot_x3f___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_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__12; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isEqFalse(lean_object*, 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_Grind_pushEqFalse(lean_object*, lean_object*, 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_Grind_congrHash_go___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__12; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3___lambda__1___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__17; +lean_object* l_Lean_isTracingEnabledForCore(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint64_t l_Lean_Meta_Grind_instHashableENodeKey_unsafe__1(lean_object*); size_t lean_usize_add(size_t, size_t); +extern lean_object* l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedEMatchTheorems___spec__1; +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__8; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_erase___at_Lean_Meta_Grind_getParentsAndReset___spec__1___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__46; +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__55; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instHashableCongrTheoremCacheKey___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instInhabitedNewFact; -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__26; static lean_object* l_Lean_Meta_Grind_getENode___closed__1; static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78____closed__9; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getMaxGeneration(lean_object*); lean_object* lean_array_uget(lean_object*, size_t); size_t lean_array_size(lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__51; +LEAN_EXPORT uint8_t l_Lean_Meta_Grind_isCongruent_goEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_instInhabitedGoal___closed__4; +static lean_object* l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__19; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pushHEq___boxed(lean_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_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2___closed__1; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_forEachEqc___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__33; lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__23; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at_Lean_Meta_Grind_getENodes___spec__2___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__6(lean_object*, size_t, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__12; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkENode___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hasSameRoot___boxed(lean_object*, lean_object*, lean_object*); size_t lean_usize_shift_left(size_t, size_t); +lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_MethodsRef_toMethods_unsafe__1___boxed(lean_object*); static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_118____closed__1; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Grind_getENodes___spec__1___lambda__1(lean_object*, lean_object*, lean_object*); @@ -500,16 +670,26 @@ LEAN_EXPORT uint64_t l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind static lean_object* l_Lean_Meta_Grind_congrHash___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkHCongrWithArity___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_RBNode_insert___at_Lean_Meta_Grind_registerParent___spec__1(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__3; +extern lean_object* l_Lean_PersistentHashMap_empty___at_Lean_KeyedDeclsAttribute_instInhabitedExtensionState___spec__1; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isEqTrue(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*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__39; +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__37; LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3___lambda__1(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__3; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__37; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_getEqcs___spec__1(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Meta_Grind_getENodes___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pushEqTrue___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_Grind_instBEqPreInstance___lambda__3(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_canon(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_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__47; static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78____closed__5; lean_object* lean_array_get_size(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isFalseExpr___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_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__12; +static lean_object* l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__11; extern lean_object* l_Lean_ShareCommon_objectFactory; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isFalseExpr(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_Grind_getRootENode(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -517,54 +697,68 @@ lean_object* lean_state_sharecommon(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_mkHCongrWithArity___spec__1(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*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__32; uint8_t lean_usize_dec_lt(size_t, size_t); static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78____closed__7; -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__29; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__51; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__40; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__4; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markTheoremInstance___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint64_t l_Lean_Meta_TransparencyMode_toUInt64(uint8_t); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____boxed(lean_object*, lean_object*); lean_object* lean_nat_add(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__38; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_congrPlaceholderProof; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_markTheoremInstance___spec__11___lambda__1___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_markTheoremInstance___spec__11(lean_object*, size_t, lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_alreadyInternalized___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__43; -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__4; -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__31; +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__25; +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__35; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_MethodsRefPointed; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d____; +static lean_object* l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__10; static lean_object* l_Lean_Meta_Grind_canon___closed__2; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__57; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GoalM_run_x27___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_String_toSubstring_x27(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markCaseSplitAsResolved___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_Grind_getConfig___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GoalM_run___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_Lean_Meta_Grind_getFalseExpr___rarg___boxed(lean_object*, lean_object*, 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*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instHashableENodeKey___boxed(lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__38; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getRoot_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT size_t l_Lean_Meta_Grind_instHashableCongrTheoremCacheKey_unsafe__1(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*); lean_object* l_Lean_Meta_isLitValue(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedGoal___spec__1; +static lean_object* l_Lean_Meta_Grind_updateLastTag___closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isInterpreted___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_getEqcs___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*); +static lean_object* l_Lean_Meta_Grind_updateLastTag___closed__2; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__6___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__53; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_markTheoremInstance___spec__11___lambda__2___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_checkMaxCaseSplit___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_Grind_isInconsistent(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_Grind_instInhabitedGoal; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instHashableCongrTheoremCacheKey_unsafe__1___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkENode___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_PersistentHashMap_insertAux___at_Lean_Meta_Grind_setENode___spec__2(lean_object*, size_t, size_t, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_updateLastTag___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_land(size_t, size_t); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__9; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getConfig___boxed(lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__49; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_setENode___spec__3(size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Option_repr___at___private_Lean_Meta_Transform_0__Lean_reprTransformStep____x40_Lean_Meta_Transform___hyg_46____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_mkHCongrWithArity___spec__5___boxed(lean_object*, lean_object*); LEAN_EXPORT uint64_t l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__2(lean_object*, lean_object*, lean_object*, size_t, size_t, uint64_t); +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__9; LEAN_EXPORT uint8_t l_Lean_Meta_Grind_isSameExpr_unsafe__1(lean_object* x_1, lean_object* x_2) { _start: { @@ -1382,13 +1576,15 @@ return x_41; } else { -lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; x_42 = lean_ctor_get(x_33, 0); x_43 = lean_ctor_get(x_33, 1); x_44 = lean_ctor_get(x_33, 3); x_45 = lean_ctor_get(x_33, 4); x_46 = lean_ctor_get(x_33, 5); x_47 = lean_ctor_get(x_33, 6); +x_48 = lean_ctor_get(x_33, 7); +lean_inc(x_48); lean_inc(x_47); lean_inc(x_46); lean_inc(x_45); @@ -1396,57 +1592,58 @@ lean_inc(x_44); lean_inc(x_43); lean_inc(x_42); lean_dec(x_33); -x_48 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_48, 0, x_42); -lean_ctor_set(x_48, 1, x_43); -lean_ctor_set(x_48, 2, x_30); -lean_ctor_set(x_48, 3, x_44); -lean_ctor_set(x_48, 4, x_45); -lean_ctor_set(x_48, 5, x_46); -lean_ctor_set(x_48, 6, x_47); -x_49 = lean_st_ref_set(x_4, x_48, x_34); -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; +x_49 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_49, 0, x_42); +lean_ctor_set(x_49, 1, x_43); +lean_ctor_set(x_49, 2, x_30); +lean_ctor_set(x_49, 3, x_44); +lean_ctor_set(x_49, 4, x_45); +lean_ctor_set(x_49, 5, x_46); +lean_ctor_set(x_49, 6, x_47); +lean_ctor_set(x_49, 7, x_48); +x_50 = lean_st_ref_set(x_4, x_49, x_34); +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_49); - x_51 = lean_box(0); + lean_dec_ref(x_50); + x_52 = lean_box(0); } -if (lean_is_scalar(x_51)) { - x_52 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_52)) { + x_53 = lean_alloc_ctor(0, 2, 0); } else { - x_52 = x_51; + x_53 = x_52; } -lean_ctor_set(x_52, 0, x_25); -lean_ctor_set(x_52, 1, x_50); -return x_52; +lean_ctor_set(x_53, 0, x_25); +lean_ctor_set(x_53, 1, x_51); +return x_53; } } else { -uint8_t x_53; +uint8_t x_54; lean_dec(x_22); lean_dec(x_18); -x_53 = !lean_is_exclusive(x_24); -if (x_53 == 0) +x_54 = !lean_is_exclusive(x_24); +if (x_54 == 0) { return x_24; } else { -lean_object* x_54; lean_object* x_55; lean_object* x_56; -x_54 = lean_ctor_get(x_24, 0); -x_55 = lean_ctor_get(x_24, 1); +lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_55 = lean_ctor_get(x_24, 0); +x_56 = lean_ctor_get(x_24, 1); +lean_inc(x_56); lean_inc(x_55); -lean_inc(x_54); lean_dec(x_24); -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; } } } @@ -1509,7 +1706,7 @@ return x_23; } else { -lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; 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_24 = lean_ctor_get(x_11, 0); x_25 = lean_ctor_get(x_11, 1); x_26 = lean_ctor_get(x_11, 2); @@ -1517,6 +1714,8 @@ x_27 = lean_ctor_get(x_11, 3); x_28 = lean_ctor_get(x_11, 4); x_29 = lean_ctor_get(x_11, 5); x_30 = lean_ctor_get(x_11, 6); +x_31 = lean_ctor_get(x_11, 7); +lean_inc(x_31); lean_inc(x_30); lean_inc(x_29); lean_inc(x_28); @@ -1525,40 +1724,41 @@ lean_inc(x_26); lean_inc(x_25); lean_inc(x_24); lean_dec(x_11); -x_31 = l_Lean_ShareCommon_objectFactory; -x_32 = lean_state_sharecommon(x_31, x_25, x_1); -x_33 = lean_ctor_get(x_32, 0); -lean_inc(x_33); -x_34 = lean_ctor_get(x_32, 1); +x_32 = l_Lean_ShareCommon_objectFactory; +x_33 = lean_state_sharecommon(x_32, x_25, x_1); +x_34 = lean_ctor_get(x_33, 0); lean_inc(x_34); -lean_dec(x_32); -x_35 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_35, 0, x_24); -lean_ctor_set(x_35, 1, x_34); -lean_ctor_set(x_35, 2, x_26); -lean_ctor_set(x_35, 3, x_27); -lean_ctor_set(x_35, 4, x_28); -lean_ctor_set(x_35, 5, x_29); -lean_ctor_set(x_35, 6, x_30); -x_36 = lean_st_ref_set(x_4, x_35, x_12); -x_37 = lean_ctor_get(x_36, 1); -lean_inc(x_37); -if (lean_is_exclusive(x_36)) { - lean_ctor_release(x_36, 0); - lean_ctor_release(x_36, 1); - x_38 = x_36; +x_35 = lean_ctor_get(x_33, 1); +lean_inc(x_35); +lean_dec(x_33); +x_36 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_36, 0, x_24); +lean_ctor_set(x_36, 1, x_35); +lean_ctor_set(x_36, 2, x_26); +lean_ctor_set(x_36, 3, x_27); +lean_ctor_set(x_36, 4, x_28); +lean_ctor_set(x_36, 5, x_29); +lean_ctor_set(x_36, 6, x_30); +lean_ctor_set(x_36, 7, x_31); +x_37 = lean_st_ref_set(x_4, x_36, x_12); +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_36); - x_38 = lean_box(0); + lean_dec_ref(x_37); + x_39 = lean_box(0); } -if (lean_is_scalar(x_38)) { - x_39 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_39)) { + x_40 = lean_alloc_ctor(0, 2, 0); } else { - x_39 = x_38; + x_40 = x_39; } -lean_ctor_set(x_39, 0, x_33); -lean_ctor_set(x_39, 1, x_37); -return x_39; +lean_ctor_set(x_40, 0, x_34); +lean_ctor_set(x_40, 1, x_38); +return x_40; } } } @@ -1628,7 +1828,7 @@ x_16 = lean_st_ref_set(x_4, x_11, x_12); x_17 = lean_ctor_get(x_16, 1); lean_inc(x_17); lean_dec(x_16); -x_18 = l_Lean_Meta_Grind_Canon_canonImpl(x_1, x_14, x_5, x_6, x_7, x_8, x_17); +x_18 = l_Lean_Meta_Grind_Canon_canon(x_1, x_14, x_5, x_6, 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; lean_object* x_25; uint8_t x_26; @@ -1679,13 +1879,15 @@ 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_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; 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_24, 1); x_34 = lean_ctor_get(x_24, 2); x_35 = lean_ctor_get(x_24, 3); x_36 = lean_ctor_get(x_24, 4); x_37 = lean_ctor_get(x_24, 5); x_38 = lean_ctor_get(x_24, 6); +x_39 = lean_ctor_get(x_24, 7); +lean_inc(x_39); lean_inc(x_38); lean_inc(x_37); lean_inc(x_36); @@ -1693,189 +1895,197 @@ lean_inc(x_35); lean_inc(x_34); lean_inc(x_33); lean_dec(x_24); -x_39 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_39, 0, x_22); -lean_ctor_set(x_39, 1, x_33); -lean_ctor_set(x_39, 2, x_34); -lean_ctor_set(x_39, 3, x_35); -lean_ctor_set(x_39, 4, x_36); -lean_ctor_set(x_39, 5, x_37); -lean_ctor_set(x_39, 6, x_38); -x_40 = lean_st_ref_set(x_4, x_39, x_25); -x_41 = lean_ctor_get(x_40, 1); -lean_inc(x_41); -if (lean_is_exclusive(x_40)) { - lean_ctor_release(x_40, 0); - lean_ctor_release(x_40, 1); - x_42 = x_40; +x_40 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_40, 0, x_22); +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_36); +lean_ctor_set(x_40, 5, x_37); +lean_ctor_set(x_40, 6, x_38); +lean_ctor_set(x_40, 7, x_39); +x_41 = lean_st_ref_set(x_4, x_40, x_25); +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_40); - x_42 = lean_box(0); + lean_dec_ref(x_41); + x_43 = lean_box(0); } -if (lean_is_scalar(x_42)) { - x_43 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_43)) { + x_44 = lean_alloc_ctor(0, 2, 0); } else { - x_43 = x_42; + x_44 = x_43; } -lean_ctor_set(x_43, 0, x_21); -lean_ctor_set(x_43, 1, x_41); -return x_43; +lean_ctor_set(x_44, 0, x_21); +lean_ctor_set(x_44, 1, x_42); +return x_44; } } else { -uint8_t x_44; -x_44 = !lean_is_exclusive(x_18); -if (x_44 == 0) +uint8_t x_45; +x_45 = !lean_is_exclusive(x_18); +if (x_45 == 0) { return x_18; } else { -lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_45 = lean_ctor_get(x_18, 0); -x_46 = lean_ctor_get(x_18, 1); +lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_46 = lean_ctor_get(x_18, 0); +x_47 = lean_ctor_get(x_18, 1); +lean_inc(x_47); lean_inc(x_46); -lean_inc(x_45); lean_dec(x_18); -x_47 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_47, 0, x_45); -lean_ctor_set(x_47, 1, x_46); -return x_47; +x_48 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_48, 0, x_46); +lean_ctor_set(x_48, 1, x_47); +return x_48; } } } else { -lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; -x_48 = lean_ctor_get(x_11, 0); -x_49 = lean_ctor_get(x_11, 1); -x_50 = lean_ctor_get(x_11, 2); -x_51 = lean_ctor_get(x_11, 3); -x_52 = lean_ctor_get(x_11, 4); -x_53 = lean_ctor_get(x_11, 5); -x_54 = lean_ctor_get(x_11, 6); +lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; 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_11, 0); +x_50 = lean_ctor_get(x_11, 1); +x_51 = lean_ctor_get(x_11, 2); +x_52 = lean_ctor_get(x_11, 3); +x_53 = lean_ctor_get(x_11, 4); +x_54 = lean_ctor_get(x_11, 5); +x_55 = lean_ctor_get(x_11, 6); +x_56 = lean_ctor_get(x_11, 7); +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_inc(x_49); -lean_inc(x_48); lean_dec(x_11); -x_55 = l_Lean_Meta_Grind_canon___closed__3; -x_56 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_56, 0, x_55); -lean_ctor_set(x_56, 1, x_49); -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); -x_57 = lean_st_ref_set(x_4, x_56, x_12); -x_58 = lean_ctor_get(x_57, 1); -lean_inc(x_58); -lean_dec(x_57); -x_59 = l_Lean_Meta_Grind_Canon_canonImpl(x_1, x_48, x_5, x_6, x_7, x_8, x_58); -if (lean_obj_tag(x_59) == 0) -{ -lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; -x_60 = lean_ctor_get(x_59, 0); +x_57 = l_Lean_Meta_Grind_canon___closed__3; +x_58 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_58, 0, x_57); +lean_ctor_set(x_58, 1, x_50); +lean_ctor_set(x_58, 2, x_51); +lean_ctor_set(x_58, 3, x_52); +lean_ctor_set(x_58, 4, x_53); +lean_ctor_set(x_58, 5, x_54); +lean_ctor_set(x_58, 6, x_55); +lean_ctor_set(x_58, 7, x_56); +x_59 = lean_st_ref_set(x_4, x_58, x_12); +x_60 = lean_ctor_get(x_59, 1); 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); +x_61 = l_Lean_Meta_Grind_Canon_canon(x_1, x_49, x_5, x_6, x_7, x_8, x_60); +if (lean_obj_tag(x_61) == 0) +{ +lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; +x_62 = lean_ctor_get(x_61, 0); lean_inc(x_62); -x_63 = lean_ctor_get(x_60, 1); +x_63 = lean_ctor_get(x_61, 1); lean_inc(x_63); -lean_dec(x_60); -x_64 = lean_st_ref_take(x_4, x_61); -x_65 = lean_ctor_get(x_64, 0); +lean_dec(x_61); +x_64 = lean_ctor_get(x_62, 0); +lean_inc(x_64); +x_65 = lean_ctor_get(x_62, 1); 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, 1); +lean_dec(x_62); +x_66 = lean_st_ref_take(x_4, x_63); +x_67 = lean_ctor_get(x_66, 0); lean_inc(x_67); -x_68 = lean_ctor_get(x_65, 2); +x_68 = lean_ctor_get(x_66, 1); lean_inc(x_68); -x_69 = lean_ctor_get(x_65, 3); +lean_dec(x_66); +x_69 = lean_ctor_get(x_67, 1); lean_inc(x_69); -x_70 = lean_ctor_get(x_65, 4); +x_70 = lean_ctor_get(x_67, 2); lean_inc(x_70); -x_71 = lean_ctor_get(x_65, 5); +x_71 = lean_ctor_get(x_67, 3); lean_inc(x_71); -x_72 = lean_ctor_get(x_65, 6); +x_72 = lean_ctor_get(x_67, 4); lean_inc(x_72); -if (lean_is_exclusive(x_65)) { - lean_ctor_release(x_65, 0); - lean_ctor_release(x_65, 1); - lean_ctor_release(x_65, 2); - lean_ctor_release(x_65, 3); - lean_ctor_release(x_65, 4); - lean_ctor_release(x_65, 5); - lean_ctor_release(x_65, 6); - x_73 = x_65; +x_73 = lean_ctor_get(x_67, 5); +lean_inc(x_73); +x_74 = lean_ctor_get(x_67, 6); +lean_inc(x_74); +x_75 = lean_ctor_get(x_67, 7); +lean_inc(x_75); +if (lean_is_exclusive(x_67)) { + lean_ctor_release(x_67, 0); + lean_ctor_release(x_67, 1); + lean_ctor_release(x_67, 2); + lean_ctor_release(x_67, 3); + lean_ctor_release(x_67, 4); + lean_ctor_release(x_67, 5); + lean_ctor_release(x_67, 6); + lean_ctor_release(x_67, 7); + x_76 = x_67; } else { - lean_dec_ref(x_65); - x_73 = lean_box(0); + lean_dec_ref(x_67); + x_76 = lean_box(0); } -if (lean_is_scalar(x_73)) { - x_74 = lean_alloc_ctor(0, 7, 0); +if (lean_is_scalar(x_76)) { + x_77 = lean_alloc_ctor(0, 8, 0); } else { - x_74 = x_73; -} -lean_ctor_set(x_74, 0, x_63); -lean_ctor_set(x_74, 1, x_67); -lean_ctor_set(x_74, 2, x_68); -lean_ctor_set(x_74, 3, x_69); -lean_ctor_set(x_74, 4, x_70); -lean_ctor_set(x_74, 5, x_71); -lean_ctor_set(x_74, 6, x_72); -x_75 = lean_st_ref_set(x_4, x_74, x_66); -x_76 = lean_ctor_get(x_75, 1); -lean_inc(x_76); -if (lean_is_exclusive(x_75)) { - lean_ctor_release(x_75, 0); - lean_ctor_release(x_75, 1); - x_77 = x_75; + x_77 = x_76; +} +lean_ctor_set(x_77, 0, x_65); +lean_ctor_set(x_77, 1, x_69); +lean_ctor_set(x_77, 2, x_70); +lean_ctor_set(x_77, 3, x_71); +lean_ctor_set(x_77, 4, x_72); +lean_ctor_set(x_77, 5, x_73); +lean_ctor_set(x_77, 6, x_74); +lean_ctor_set(x_77, 7, x_75); +x_78 = lean_st_ref_set(x_4, x_77, x_68); +x_79 = lean_ctor_get(x_78, 1); +lean_inc(x_79); +if (lean_is_exclusive(x_78)) { + lean_ctor_release(x_78, 0); + lean_ctor_release(x_78, 1); + x_80 = x_78; } else { - lean_dec_ref(x_75); - x_77 = lean_box(0); + lean_dec_ref(x_78); + x_80 = lean_box(0); } -if (lean_is_scalar(x_77)) { - x_78 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_80)) { + x_81 = lean_alloc_ctor(0, 2, 0); } else { - x_78 = x_77; + x_81 = x_80; } -lean_ctor_set(x_78, 0, x_62); -lean_ctor_set(x_78, 1, x_76); -return x_78; +lean_ctor_set(x_81, 0, x_64); +lean_ctor_set(x_81, 1, x_79); +return x_81; } else { -lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; -x_79 = lean_ctor_get(x_59, 0); -lean_inc(x_79); -x_80 = lean_ctor_get(x_59, 1); -lean_inc(x_80); -if (lean_is_exclusive(x_59)) { - lean_ctor_release(x_59, 0); - lean_ctor_release(x_59, 1); - x_81 = x_59; +lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; +x_82 = lean_ctor_get(x_61, 0); +lean_inc(x_82); +x_83 = lean_ctor_get(x_61, 1); +lean_inc(x_83); +if (lean_is_exclusive(x_61)) { + lean_ctor_release(x_61, 0); + lean_ctor_release(x_61, 1); + x_84 = x_61; } else { - lean_dec_ref(x_59); - x_81 = lean_box(0); + lean_dec_ref(x_61); + x_84 = lean_box(0); } -if (lean_is_scalar(x_81)) { - x_82 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_84)) { + x_85 = lean_alloc_ctor(1, 2, 0); } else { - x_82 = x_81; + x_85 = x_84; } -lean_ctor_set(x_82, 0, x_79); -lean_ctor_set(x_82, 1, x_80); -return x_82; +lean_ctor_set(x_85, 0, x_82); +lean_ctor_set(x_85, 1, x_83); +return x_85; } } } @@ -2965,7 +3175,7 @@ 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; lean_object* x_39; +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; 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_27 = lean_ctor_get(x_17, 0); x_28 = lean_ctor_get(x_17, 1); x_29 = lean_ctor_get(x_17, 2); @@ -2973,6 +3183,8 @@ x_30 = lean_ctor_get(x_17, 3); x_31 = lean_ctor_get(x_17, 4); x_32 = lean_ctor_get(x_17, 5); x_33 = lean_ctor_get(x_17, 6); +x_34 = lean_ctor_get(x_17, 7); +lean_inc(x_34); lean_inc(x_33); lean_inc(x_32); lean_inc(x_31); @@ -2982,57 +3194,58 @@ lean_inc(x_28); lean_inc(x_27); lean_dec(x_17); lean_inc(x_14); -x_34 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_mkHCongrWithArity___spec__1(x_30, x_3, x_14); -x_35 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_35, 0, x_27); -lean_ctor_set(x_35, 1, x_28); -lean_ctor_set(x_35, 2, x_29); -lean_ctor_set(x_35, 3, x_34); -lean_ctor_set(x_35, 4, x_31); -lean_ctor_set(x_35, 5, x_32); -lean_ctor_set(x_35, 6, x_33); -x_36 = lean_st_ref_set(x_7, x_35, x_18); -x_37 = lean_ctor_get(x_36, 1); -lean_inc(x_37); -if (lean_is_exclusive(x_36)) { - lean_ctor_release(x_36, 0); - lean_ctor_release(x_36, 1); - x_38 = x_36; +x_35 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_mkHCongrWithArity___spec__1(x_30, x_3, x_14); +x_36 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_36, 0, x_27); +lean_ctor_set(x_36, 1, x_28); +lean_ctor_set(x_36, 2, x_29); +lean_ctor_set(x_36, 3, x_35); +lean_ctor_set(x_36, 4, x_31); +lean_ctor_set(x_36, 5, x_32); +lean_ctor_set(x_36, 6, x_33); +lean_ctor_set(x_36, 7, x_34); +x_37 = lean_st_ref_set(x_7, x_36, x_18); +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_36); - x_38 = lean_box(0); + lean_dec_ref(x_37); + x_39 = lean_box(0); } -if (lean_is_scalar(x_38)) { - x_39 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_39)) { + x_40 = lean_alloc_ctor(0, 2, 0); } else { - x_39 = x_38; + x_40 = x_39; } -lean_ctor_set(x_39, 0, x_14); -lean_ctor_set(x_39, 1, x_37); -return x_39; +lean_ctor_set(x_40, 0, x_14); +lean_ctor_set(x_40, 1, x_38); +return x_40; } } else { -uint8_t x_40; +uint8_t x_41; lean_dec(x_3); -x_40 = !lean_is_exclusive(x_13); -if (x_40 == 0) +x_41 = !lean_is_exclusive(x_13); +if (x_41 == 0) { return x_13; } else { -lean_object* x_41; lean_object* x_42; lean_object* x_43; -x_41 = lean_ctor_get(x_13, 0); -x_42 = lean_ctor_get(x_13, 1); +lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_42 = lean_ctor_get(x_13, 0); +x_43 = lean_ctor_get(x_13, 1); +lean_inc(x_43); lean_inc(x_42); -lean_inc(x_41); lean_dec(x_13); -x_43 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_43, 0, x_41); -lean_ctor_set(x_43, 1, x_42); -return x_43; +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; } } } @@ -3070,7 +3283,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; lean_object* x_28; 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; +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -3105,41 +3318,44 @@ x_31 = lean_ctor_get(x_23, 5); lean_inc(x_31); x_32 = lean_ctor_get(x_23, 6); lean_inc(x_32); +x_33 = lean_ctor_get(x_23, 7); +lean_inc(x_33); lean_dec(x_23); -x_33 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_33, 0, x_25); -lean_ctor_set(x_33, 1, x_26); -lean_ctor_set(x_33, 2, x_27); -lean_ctor_set(x_33, 3, x_29); -lean_ctor_set(x_33, 4, x_30); -lean_ctor_set(x_33, 5, x_31); -lean_ctor_set(x_33, 6, x_32); -x_34 = lean_st_ref_set(x_7, x_33, x_24); -x_35 = !lean_is_exclusive(x_34); -if (x_35 == 0) +x_34 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_34, 0, x_25); +lean_ctor_set(x_34, 1, x_26); +lean_ctor_set(x_34, 2, x_27); +lean_ctor_set(x_34, 3, x_29); +lean_ctor_set(x_34, 4, x_30); +lean_ctor_set(x_34, 5, x_31); +lean_ctor_set(x_34, 6, x_32); +lean_ctor_set(x_34, 7, x_33); +x_35 = lean_st_ref_set(x_7, x_34, x_24); +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_21); -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_21); +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_21); -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_21); +lean_ctor_set(x_39, 1, x_38); +return x_39; } } } else { -uint8_t x_39; +uint8_t x_40; lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -3147,32 +3363,32 @@ lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_39 = !lean_is_exclusive(x_15); -if (x_39 == 0) +x_40 = !lean_is_exclusive(x_15); +if (x_40 == 0) { return x_15; } 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_object* x_41; lean_object* x_42; lean_object* x_43; +x_41 = lean_ctor_get(x_15, 0); +x_42 = lean_ctor_get(x_15, 1); +lean_inc(x_42); lean_inc(x_41); -lean_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; +x_43 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_43, 0, x_41); +lean_ctor_set(x_43, 1, x_42); +return x_43; } } } else { -lean_object* x_43; lean_object* x_44; -x_43 = lean_box(0); -x_44 = l_Lean_Meta_Grind_mkHCongrWithArity___lambda__1(x_1, x_2, x_3, x_43, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -return x_44; +lean_object* x_44; lean_object* x_45; +x_44 = lean_box(0); +x_45 = l_Lean_Meta_Grind_mkHCongrWithArity___lambda__1(x_1, x_2, x_3, x_44, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_45; } } } @@ -3395,7 +3611,7 @@ x_1 = l_Lean_Meta_Grind_instInhabitedENode___closed__2; return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__1() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__1() { _start: { lean_object* x_1; @@ -3403,29 +3619,29 @@ x_1 = lean_mk_string_unchecked("self", 4, 4); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__2() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__1; +x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__1; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__3() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__2; +x_2 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__2; x_3 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__4() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__4() { _start: { lean_object* x_1; @@ -3433,29 +3649,29 @@ x_1 = lean_mk_string_unchecked(" := ", 4, 4); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__5() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__5() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__4; +x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__4; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__6() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__3; -x_2 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__5; +x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__3; +x_2 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____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_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__7() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__7() { _start: { lean_object* x_1; lean_object* x_2; @@ -3464,7 +3680,7 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__8() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__8() { _start: { lean_object* x_1; @@ -3472,17 +3688,17 @@ x_1 = lean_mk_string_unchecked(",", 1, 1); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__9() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__9() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__8; +x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____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_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__10() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__10() { _start: { lean_object* x_1; @@ -3490,17 +3706,17 @@ x_1 = lean_mk_string_unchecked("next", 4, 4); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__11() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__11() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__10; +x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____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_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__12() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__12() { _start: { lean_object* x_1; @@ -3508,44 +3724,44 @@ x_1 = lean_mk_string_unchecked("root", 4, 4); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__13() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__13() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__12; +x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____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_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__14() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__14() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("cgRoot", 6, 6); +x_1 = lean_mk_string_unchecked("congr", 5, 5); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__15() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__15() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__14; +x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__14; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__16() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__16() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = lean_unsigned_to_nat(10u); +x_1 = lean_unsigned_to_nat(9u); x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__17() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__17() { _start: { lean_object* x_1; @@ -3553,17 +3769,17 @@ x_1 = lean_mk_string_unchecked("target\?", 7, 7); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__18() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__18() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__17; +x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____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_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__19() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__19() { _start: { lean_object* x_1; lean_object* x_2; @@ -3572,7 +3788,7 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__20() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__20() { _start: { lean_object* x_1; @@ -3580,17 +3796,26 @@ x_1 = lean_mk_string_unchecked("proof\?", 6, 6); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__21() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__21() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__20; +x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____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_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__22() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__22() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(10u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__23() { _start: { lean_object* x_1; @@ -3598,17 +3823,17 @@ x_1 = lean_mk_string_unchecked("flipped", 7, 7); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__23() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__24() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__22; +x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____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_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__24() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__25() { _start: { lean_object* x_1; @@ -3616,17 +3841,17 @@ x_1 = lean_mk_string_unchecked("size", 4, 4); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__25() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__26() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__24; +x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____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_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__26() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__27() { _start: { lean_object* x_1; @@ -3634,17 +3859,17 @@ x_1 = lean_mk_string_unchecked("interpreted", 11, 11); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__27() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__28() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__26; +x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____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_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__28() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__29() { _start: { lean_object* x_1; lean_object* x_2; @@ -3653,7 +3878,7 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__29() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__30() { _start: { lean_object* x_1; @@ -3661,17 +3886,17 @@ x_1 = lean_mk_string_unchecked("ctor", 4, 4); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__30() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__31() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__29; +x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__30; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__31() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__32() { _start: { lean_object* x_1; @@ -3679,17 +3904,17 @@ x_1 = lean_mk_string_unchecked("hasLambdas", 10, 10); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__32() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__33() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__31; +x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__32; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__33() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__34() { _start: { lean_object* x_1; lean_object* x_2; @@ -3698,7 +3923,7 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__34() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__35() { _start: { lean_object* x_1; @@ -3706,17 +3931,17 @@ x_1 = lean_mk_string_unchecked("heqProofs", 9, 9); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__35() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__36() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__34; +x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____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_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__36() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__37() { _start: { lean_object* x_1; lean_object* x_2; @@ -3725,7 +3950,7 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__37() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__38() { _start: { lean_object* x_1; @@ -3733,17 +3958,17 @@ x_1 = lean_mk_string_unchecked("idx", 3, 3); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__38() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__39() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__37; +x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__38; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__39() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__40() { _start: { lean_object* x_1; lean_object* x_2; @@ -3752,7 +3977,7 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__40() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__41() { _start: { lean_object* x_1; @@ -3760,17 +3985,17 @@ x_1 = lean_mk_string_unchecked("generation", 10, 10); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__41() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__42() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__40; +x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____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_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__42() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__43() { _start: { lean_object* x_1; @@ -3778,17 +4003,17 @@ x_1 = lean_mk_string_unchecked("mt", 2, 2); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__43() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__44() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__42; +x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____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_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__44() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__45() { _start: { lean_object* x_1; lean_object* x_2; @@ -3797,7 +4022,7 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__45() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__46() { _start: { lean_object* x_1; @@ -3805,35 +4030,35 @@ x_1 = lean_mk_string_unchecked("{ ", 2, 2); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__46() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__47() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__45; +x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__46; x_2 = lean_string_length(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__47() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__48() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__46; +x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__47; x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__48() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__49() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__45; +x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__46; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__49() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__50() { _start: { lean_object* x_1; @@ -3841,17 +4066,17 @@ x_1 = lean_mk_string_unchecked(" }", 2, 2); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__50() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__51() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__49; +x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__50; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__51() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__52() { _start: { lean_object* x_1; @@ -3859,17 +4084,17 @@ x_1 = lean_mk_string_unchecked("false", 5, 5); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__52() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__53() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__51; +x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__52; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__53() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__54() { _start: { lean_object* x_1; @@ -3877,25 +4102,25 @@ x_1 = lean_mk_string_unchecked("true", 4, 4); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__54() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__55() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__53; +x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__54; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560_(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; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; 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; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* 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; 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_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; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; 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; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; uint8_t x_79; uint8_t x_80; uint8_t x_81; uint8_t x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; x_3 = lean_ctor_get(x_1, 0); lean_inc(x_3); x_4 = lean_unsigned_to_nat(0u); x_5 = l___private_Lean_Expr_0__Lean_reprExpr____x40_Lean_Expr___hyg_3036_(x_3, x_4); -x_6 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__7; +x_6 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____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); @@ -3903,11 +4128,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_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__6; +x_10 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____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_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__9; +x_12 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____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); @@ -3915,11 +4140,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_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__11; +x_16 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____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_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__5; +x_18 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____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); @@ -3941,7 +4166,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_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__13; +x_27 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__13; x_28 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_28, 0, x_26); lean_ctor_set(x_28, 1, x_27); @@ -3966,7 +4191,7 @@ lean_ctor_set(x_35, 1, x_12); x_36 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_36, 0, x_35); lean_ctor_set(x_36, 1, x_14); -x_37 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__15; +x_37 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__15; x_38 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_38, 0, x_36); lean_ctor_set(x_38, 1, x_37); @@ -3976,7 +4201,7 @@ lean_ctor_set(x_39, 1, x_18); x_40 = lean_ctor_get(x_1, 3); lean_inc(x_40); x_41 = l___private_Lean_Expr_0__Lean_reprExpr____x40_Lean_Expr___hyg_3036_(x_40, x_4); -x_42 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__16; +x_42 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__16; x_43 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_43, 0, x_42); lean_ctor_set(x_43, 1, x_41); @@ -3992,7 +4217,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_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__18; +x_48 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____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); @@ -4002,7 +4227,7 @@ lean_ctor_set(x_50, 1, x_18); x_51 = lean_ctor_get(x_1, 4); lean_inc(x_51); x_52 = l_Option_repr___at___private_Lean_Meta_Transform_0__Lean_reprTransformStep____x40_Lean_Meta_Transform___hyg_46____spec__1(x_51, x_4); -x_53 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__19; +x_53 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____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); @@ -4018,7 +4243,7 @@ lean_ctor_set(x_57, 1, x_12); x_58 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_58, 0, x_57); lean_ctor_set(x_58, 1, x_14); -x_59 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__21; +x_59 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__21; x_60 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_60, 0, x_58); lean_ctor_set(x_60, 1, x_59); @@ -4028,345 +4253,346 @@ lean_ctor_set(x_61, 1, x_18); x_62 = lean_ctor_get(x_1, 5); lean_inc(x_62); x_63 = l_Option_repr___at___private_Lean_Meta_Transform_0__Lean_reprTransformStep____x40_Lean_Meta_Transform___hyg_46____spec__1(x_62, x_4); -x_64 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_64, 0, x_42); -lean_ctor_set(x_64, 1, x_63); -x_65 = lean_alloc_ctor(6, 1, 1); +x_64 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__22; +x_65 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_65, 0, x_64); -lean_ctor_set_uint8(x_65, sizeof(void*)*1, x_8); -x_66 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_66, 0, x_61); -lean_ctor_set(x_66, 1, x_65); +lean_ctor_set(x_65, 1, x_63); +x_66 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_66, 0, x_65); +lean_ctor_set_uint8(x_66, sizeof(void*)*1, x_8); x_67 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_67, 0, x_66); -lean_ctor_set(x_67, 1, x_12); +lean_ctor_set(x_67, 0, x_61); +lean_ctor_set(x_67, 1, x_66); x_68 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_68, 0, x_67); -lean_ctor_set(x_68, 1, x_14); -x_69 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__23; -x_70 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_70, 0, x_68); -lean_ctor_set(x_70, 1, x_69); +lean_ctor_set(x_68, 1, x_12); +x_69 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_69, 0, x_68); +lean_ctor_set(x_69, 1, x_14); +x_70 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__24; x_71 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_71, 0, x_70); -lean_ctor_set(x_71, 1, x_18); -x_72 = lean_ctor_get_uint8(x_1, sizeof(void*)*10); -x_73 = lean_ctor_get(x_1, 6); -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(4, 2, 0); -lean_ctor_set(x_76, 0, x_6); -lean_ctor_set(x_76, 1, x_75); -x_77 = lean_alloc_ctor(6, 1, 1); -lean_ctor_set(x_77, 0, x_76); -lean_ctor_set_uint8(x_77, sizeof(void*)*1, x_8); -x_78 = lean_ctor_get_uint8(x_1, sizeof(void*)*10 + 1); -x_79 = lean_ctor_get_uint8(x_1, sizeof(void*)*10 + 2); -x_80 = lean_ctor_get_uint8(x_1, sizeof(void*)*10 + 3); -x_81 = lean_ctor_get_uint8(x_1, sizeof(void*)*10 + 4); -x_82 = lean_ctor_get(x_1, 7); -lean_inc(x_82); -x_83 = l___private_Init_Data_Repr_0__Nat_reprFast(x_82); -x_84 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_84, 0, x_83); -x_85 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__39; -x_86 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_86, 0, x_85); -lean_ctor_set(x_86, 1, x_84); -x_87 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_71, 0, x_69); +lean_ctor_set(x_71, 1, x_70); +x_72 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_72, 0, x_71); +lean_ctor_set(x_72, 1, x_18); +x_73 = lean_ctor_get_uint8(x_1, sizeof(void*)*10); +x_74 = lean_ctor_get(x_1, 6); +lean_inc(x_74); +x_75 = l___private_Init_Data_Repr_0__Nat_reprFast(x_74); +x_76 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_76, 0, x_75); +x_77 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_77, 0, x_6); +lean_ctor_set(x_77, 1, x_76); +x_78 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_78, 0, x_77); +lean_ctor_set_uint8(x_78, sizeof(void*)*1, x_8); +x_79 = lean_ctor_get_uint8(x_1, sizeof(void*)*10 + 1); +x_80 = lean_ctor_get_uint8(x_1, sizeof(void*)*10 + 2); +x_81 = lean_ctor_get_uint8(x_1, sizeof(void*)*10 + 3); +x_82 = lean_ctor_get_uint8(x_1, sizeof(void*)*10 + 4); +x_83 = lean_ctor_get(x_1, 7); +lean_inc(x_83); +x_84 = l___private_Init_Data_Repr_0__Nat_reprFast(x_83); +x_85 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_85, 0, x_84); +x_86 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__40; +x_87 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_87, 0, x_86); -lean_ctor_set_uint8(x_87, sizeof(void*)*1, x_8); -x_88 = lean_ctor_get(x_1, 8); -lean_inc(x_88); -x_89 = l___private_Init_Data_Repr_0__Nat_reprFast(x_88); -x_90 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_90, 0, x_89); -x_91 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__33; -x_92 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_92, 0, x_91); -lean_ctor_set(x_92, 1, x_90); -x_93 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_87, 1, x_85); +x_88 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_88, 0, x_87); +lean_ctor_set_uint8(x_88, sizeof(void*)*1, x_8); +x_89 = lean_ctor_get(x_1, 8); +lean_inc(x_89); +x_90 = l___private_Init_Data_Repr_0__Nat_reprFast(x_89); +x_91 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_91, 0, x_90); +x_92 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__34; +x_93 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_93, 0, x_92); -lean_ctor_set_uint8(x_93, sizeof(void*)*1, x_8); -x_94 = lean_ctor_get(x_1, 9); -lean_inc(x_94); +lean_ctor_set(x_93, 1, x_91); +x_94 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_94, 0, x_93); +lean_ctor_set_uint8(x_94, sizeof(void*)*1, x_8); +x_95 = lean_ctor_get(x_1, 9); +lean_inc(x_95); lean_dec(x_1); -x_95 = l___private_Init_Data_Repr_0__Nat_reprFast(x_94); -x_96 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_96, 0, x_95); -x_97 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__44; -x_98 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_98, 0, x_97); -lean_ctor_set(x_98, 1, x_96); -x_99 = lean_alloc_ctor(6, 1, 1); +x_96 = l___private_Init_Data_Repr_0__Nat_reprFast(x_95); +x_97 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_97, 0, x_96); +x_98 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__45; +x_99 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_99, 0, x_98); -lean_ctor_set_uint8(x_99, sizeof(void*)*1, x_8); -if (x_72 == 0) +lean_ctor_set(x_99, 1, x_97); +x_100 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_100, 0, x_99); +lean_ctor_set_uint8(x_100, sizeof(void*)*1, x_8); +if (x_73 == 0) { -lean_object* x_186; -x_186 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__52; -x_100 = x_186; -goto block_185; +lean_object* x_187; +x_187 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__53; +x_101 = x_187; +goto block_186; } else { -lean_object* x_187; -x_187 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__54; -x_100 = x_187; -goto block_185; -} -block_185: -{ -lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; -x_101 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_101, 0, x_53); -lean_ctor_set(x_101, 1, x_100); -x_102 = lean_alloc_ctor(6, 1, 1); -lean_ctor_set(x_102, 0, x_101); -lean_ctor_set_uint8(x_102, sizeof(void*)*1, x_8); -x_103 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_103, 0, x_71); -lean_ctor_set(x_103, 1, x_102); +lean_object* x_188; +x_188 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__55; +x_101 = x_188; +goto block_186; +} +block_186: +{ +lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; +x_102 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_102, 0, x_53); +lean_ctor_set(x_102, 1, x_101); +x_103 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_103, 0, x_102); +lean_ctor_set_uint8(x_103, sizeof(void*)*1, x_8); x_104 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_104, 0, x_103); -lean_ctor_set(x_104, 1, x_12); +lean_ctor_set(x_104, 0, x_72); +lean_ctor_set(x_104, 1, x_103); x_105 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_105, 0, x_104); -lean_ctor_set(x_105, 1, x_14); -x_106 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__25; -x_107 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_107, 0, x_105); -lean_ctor_set(x_107, 1, x_106); +lean_ctor_set(x_105, 1, x_12); +x_106 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_106, 0, x_105); +lean_ctor_set(x_106, 1, x_14); +x_107 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__26; x_108 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_108, 0, x_107); -lean_ctor_set(x_108, 1, x_18); +lean_ctor_set(x_108, 0, x_106); +lean_ctor_set(x_108, 1, x_107); x_109 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_109, 0, x_108); -lean_ctor_set(x_109, 1, x_77); +lean_ctor_set(x_109, 1, x_18); x_110 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_110, 0, x_109); -lean_ctor_set(x_110, 1, x_12); +lean_ctor_set(x_110, 1, x_78); x_111 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_111, 0, x_110); -lean_ctor_set(x_111, 1, x_14); -x_112 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__27; -x_113 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_113, 0, x_111); -lean_ctor_set(x_113, 1, x_112); +lean_ctor_set(x_111, 1, x_12); +x_112 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_112, 0, x_111); +lean_ctor_set(x_112, 1, x_14); +x_113 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__28; x_114 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_114, 0, x_113); -lean_ctor_set(x_114, 1, x_18); -if (x_78 == 0) +lean_ctor_set(x_114, 0, x_112); +lean_ctor_set(x_114, 1, x_113); +x_115 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_115, 0, x_114); +lean_ctor_set(x_115, 1, x_18); +if (x_79 == 0) { -lean_object* x_183; -x_183 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__52; -x_115 = x_183; -goto block_182; +lean_object* x_184; +x_184 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__53; +x_116 = x_184; +goto block_183; } else { -lean_object* x_184; -x_184 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__54; -x_115 = x_184; -goto block_182; +lean_object* x_185; +x_185 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__55; +x_116 = x_185; +goto block_183; } -block_182: +block_183: { -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_116 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__28; -x_117 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_117, 0, x_116); -lean_ctor_set(x_117, 1, x_115); -x_118 = lean_alloc_ctor(6, 1, 1); +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_117 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__29; +x_118 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_118, 0, x_117); -lean_ctor_set_uint8(x_118, sizeof(void*)*1, x_8); -x_119 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_119, 0, x_114); -lean_ctor_set(x_119, 1, x_118); +lean_ctor_set(x_118, 1, x_116); +x_119 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_119, 0, x_118); +lean_ctor_set_uint8(x_119, sizeof(void*)*1, x_8); x_120 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_120, 0, x_119); -lean_ctor_set(x_120, 1, x_12); +lean_ctor_set(x_120, 0, x_115); +lean_ctor_set(x_120, 1, x_119); x_121 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_121, 0, x_120); -lean_ctor_set(x_121, 1, x_14); -x_122 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__30; -x_123 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_123, 0, x_121); -lean_ctor_set(x_123, 1, x_122); +lean_ctor_set(x_121, 1, x_12); +x_122 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_122, 0, x_121); +lean_ctor_set(x_122, 1, x_14); +x_123 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__31; x_124 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_124, 0, x_123); -lean_ctor_set(x_124, 1, x_18); -if (x_79 == 0) +lean_ctor_set(x_124, 0, x_122); +lean_ctor_set(x_124, 1, x_123); +x_125 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_125, 0, x_124); +lean_ctor_set(x_125, 1, x_18); +if (x_80 == 0) { -lean_object* x_180; -x_180 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__52; -x_125 = x_180; -goto block_179; +lean_object* x_181; +x_181 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__53; +x_126 = x_181; +goto block_180; } else { -lean_object* x_181; -x_181 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__54; -x_125 = x_181; -goto block_179; -} -block_179: -{ -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 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_126, 0, x_6); -lean_ctor_set(x_126, 1, x_125); -x_127 = lean_alloc_ctor(6, 1, 1); -lean_ctor_set(x_127, 0, x_126); -lean_ctor_set_uint8(x_127, sizeof(void*)*1, x_8); -x_128 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_128, 0, x_124); -lean_ctor_set(x_128, 1, x_127); +lean_object* x_182; +x_182 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__55; +x_126 = x_182; +goto block_180; +} +block_180: +{ +lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; +x_127 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_127, 0, x_6); +lean_ctor_set(x_127, 1, x_126); +x_128 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_128, 0, x_127); +lean_ctor_set_uint8(x_128, sizeof(void*)*1, x_8); x_129 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_129, 0, x_128); -lean_ctor_set(x_129, 1, x_12); +lean_ctor_set(x_129, 0, x_125); +lean_ctor_set(x_129, 1, x_128); x_130 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_130, 0, x_129); -lean_ctor_set(x_130, 1, x_14); -x_131 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__32; -x_132 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_132, 0, x_130); -lean_ctor_set(x_132, 1, x_131); +lean_ctor_set(x_130, 1, x_12); +x_131 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_131, 0, x_130); +lean_ctor_set(x_131, 1, x_14); +x_132 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__33; x_133 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_133, 0, x_132); -lean_ctor_set(x_133, 1, x_18); -if (x_80 == 0) +lean_ctor_set(x_133, 0, x_131); +lean_ctor_set(x_133, 1, x_132); +x_134 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_134, 0, x_133); +lean_ctor_set(x_134, 1, x_18); +if (x_81 == 0) { -lean_object* x_177; -x_177 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__52; -x_134 = x_177; -goto block_176; +lean_object* x_178; +x_178 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__53; +x_135 = x_178; +goto block_177; } else { -lean_object* x_178; -x_178 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__54; -x_134 = x_178; -goto block_176; -} -block_176: -{ -lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; -x_135 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_135, 0, x_91); -lean_ctor_set(x_135, 1, x_134); -x_136 = lean_alloc_ctor(6, 1, 1); -lean_ctor_set(x_136, 0, x_135); -lean_ctor_set_uint8(x_136, sizeof(void*)*1, x_8); -x_137 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_137, 0, x_133); -lean_ctor_set(x_137, 1, x_136); +lean_object* x_179; +x_179 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__55; +x_135 = x_179; +goto block_177; +} +block_177: +{ +lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; +x_136 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_136, 0, x_92); +lean_ctor_set(x_136, 1, x_135); +x_137 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_137, 0, x_136); +lean_ctor_set_uint8(x_137, sizeof(void*)*1, x_8); x_138 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_138, 0, x_137); -lean_ctor_set(x_138, 1, x_12); +lean_ctor_set(x_138, 0, x_134); +lean_ctor_set(x_138, 1, x_137); x_139 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_139, 0, x_138); -lean_ctor_set(x_139, 1, x_14); -x_140 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__35; -x_141 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_141, 0, x_139); -lean_ctor_set(x_141, 1, x_140); +lean_ctor_set(x_139, 1, x_12); +x_140 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_140, 0, x_139); +lean_ctor_set(x_140, 1, x_14); +x_141 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__36; x_142 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_142, 0, x_141); -lean_ctor_set(x_142, 1, x_18); -if (x_81 == 0) +lean_ctor_set(x_142, 0, x_140); +lean_ctor_set(x_142, 1, x_141); +x_143 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_143, 0, x_142); +lean_ctor_set(x_143, 1, x_18); +if (x_82 == 0) { -lean_object* x_174; -x_174 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__52; -x_143 = x_174; -goto block_173; +lean_object* x_175; +x_175 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__53; +x_144 = x_175; +goto block_174; } else { -lean_object* x_175; -x_175 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__54; -x_143 = x_175; -goto block_173; -} -block_173: -{ -lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; -x_144 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__36; -x_145 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_145, 0, x_144); -lean_ctor_set(x_145, 1, x_143); -x_146 = lean_alloc_ctor(6, 1, 1); +lean_object* x_176; +x_176 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__55; +x_144 = x_176; +goto block_174; +} +block_174: +{ +lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; +x_145 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__37; +x_146 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_146, 0, x_145); -lean_ctor_set_uint8(x_146, sizeof(void*)*1, x_8); -x_147 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_147, 0, x_142); -lean_ctor_set(x_147, 1, x_146); +lean_ctor_set(x_146, 1, x_144); +x_147 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_147, 0, x_146); +lean_ctor_set_uint8(x_147, sizeof(void*)*1, x_8); x_148 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_148, 0, x_147); -lean_ctor_set(x_148, 1, x_12); +lean_ctor_set(x_148, 0, x_143); +lean_ctor_set(x_148, 1, x_147); x_149 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_149, 0, x_148); -lean_ctor_set(x_149, 1, x_14); -x_150 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__38; -x_151 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_151, 0, x_149); -lean_ctor_set(x_151, 1, x_150); +lean_ctor_set(x_149, 1, x_12); +x_150 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_150, 0, x_149); +lean_ctor_set(x_150, 1, x_14); +x_151 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__39; x_152 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_152, 0, x_151); -lean_ctor_set(x_152, 1, x_18); +lean_ctor_set(x_152, 0, x_150); +lean_ctor_set(x_152, 1, x_151); x_153 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_153, 0, x_152); -lean_ctor_set(x_153, 1, x_87); +lean_ctor_set(x_153, 1, x_18); x_154 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_154, 0, x_153); -lean_ctor_set(x_154, 1, x_12); +lean_ctor_set(x_154, 1, x_88); x_155 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_155, 0, x_154); -lean_ctor_set(x_155, 1, x_14); -x_156 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__41; -x_157 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_157, 0, x_155); -lean_ctor_set(x_157, 1, x_156); +lean_ctor_set(x_155, 1, x_12); +x_156 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_156, 0, x_155); +lean_ctor_set(x_156, 1, x_14); +x_157 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__42; x_158 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_158, 0, x_157); -lean_ctor_set(x_158, 1, x_18); +lean_ctor_set(x_158, 0, x_156); +lean_ctor_set(x_158, 1, x_157); x_159 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_159, 0, x_158); -lean_ctor_set(x_159, 1, x_93); +lean_ctor_set(x_159, 1, x_18); x_160 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_160, 0, x_159); -lean_ctor_set(x_160, 1, x_12); +lean_ctor_set(x_160, 1, x_94); 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_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__43; -x_163 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_163, 0, x_161); -lean_ctor_set(x_163, 1, x_162); +lean_ctor_set(x_161, 1, x_12); +x_162 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_162, 0, x_161); +lean_ctor_set(x_162, 1, x_14); +x_163 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__44; x_164 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_164, 0, x_163); -lean_ctor_set(x_164, 1, x_18); +lean_ctor_set(x_164, 0, x_162); +lean_ctor_set(x_164, 1, x_163); x_165 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_165, 0, x_164); -lean_ctor_set(x_165, 1, x_99); -x_166 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__48; -x_167 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_167, 0, x_166); -lean_ctor_set(x_167, 1, x_165); -x_168 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__50; -x_169 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_169, 0, x_167); -lean_ctor_set(x_169, 1, x_168); -x_170 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__47; -x_171 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_171, 0, x_170); -lean_ctor_set(x_171, 1, x_169); -x_172 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_165, 1, x_18); +x_166 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_166, 0, x_165); +lean_ctor_set(x_166, 1, x_100); +x_167 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__49; +x_168 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_168, 0, x_167); +lean_ctor_set(x_168, 1, x_166); +x_169 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__51; +x_170 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_170, 0, x_168); +lean_ctor_set(x_170, 1, x_169); +x_171 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__48; +x_172 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_172, 0, x_171); -lean_ctor_set_uint8(x_172, sizeof(void*)*1, x_8); -return x_172; +lean_ctor_set(x_172, 1, x_170); +x_173 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_173, 0, x_172); +lean_ctor_set_uint8(x_173, sizeof(void*)*1, x_8); +return x_173; } } } @@ -4374,11 +4600,11 @@ return x_172; } } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532_(x_1, x_2); +x_3 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560_(x_1, x_2); lean_dec(x_2); return x_3; } @@ -4387,7 +4613,7 @@ static lean_object* _init_l_Lean_Meta_Grind_instReprENode___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____boxed), 2, 0); return x_1; } } @@ -4399,6 +4625,26 @@ x_1 = l_Lean_Meta_Grind_instReprENode___closed__1; return x_1; } } +LEAN_EXPORT uint8_t l_Lean_Meta_Grind_ENode_isCongrRoot(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; uint8_t x_4; +x_2 = lean_ctor_get(x_1, 0); +x_3 = lean_ctor_get(x_1, 3); +x_4 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_2, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ENode_isCongrRoot___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = l_Lean_Meta_Grind_ENode_isCongrRoot(x_1); +lean_dec(x_1); +x_3 = lean_box(x_2); +return x_3; +} +} LEAN_EXPORT uint64_t l_Lean_Meta_Grind_instHashableENodeKey_unsafe__1(lean_object* x_1) { _start: { @@ -4733,7 +4979,7 @@ x_4 = lean_box_uint64(x_3); return x_4; } } -LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hasSameRoot(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT uint8_t l_Lean_Meta_Grind_hasSameRoot(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { uint8_t x_4; @@ -4792,17 +5038,50 @@ return x_14; } } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hasSameRoot___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_hasSameRoot___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___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hasSameRoot(x_1, x_2, x_3); +x_4 = l_Lean_Meta_Grind_hasSameRoot(x_1, x_2, x_3); lean_dec(x_3); lean_dec(x_2); x_5 = lean_box(x_4); return x_5; } } +LEAN_EXPORT uint64_t l_Lean_Meta_Grind_congrHash_goEq(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint64_t x_4; uint64_t x_5; uint8_t x_6; +lean_inc(x_1); +x_4 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot(x_1, x_2); +x_5 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot(x_1, x_3); +x_6 = lean_uint64_dec_lt(x_5, x_4); +if (x_6 == 0) +{ +uint64_t x_7; +x_7 = lean_uint64_mix_hash(x_4, x_5); +return x_7; +} +else +{ +uint64_t x_8; +x_8 = lean_uint64_mix_hash(x_5, x_4); +return x_8; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_congrHash_goEq___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint64_t x_4; lean_object* x_5; +x_4 = l_Lean_Meta_Grind_congrHash_goEq(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +x_5 = lean_box_uint64(x_4); +return x_5; +} +} LEAN_EXPORT uint64_t l_Lean_Meta_Grind_congrHash_go(lean_object* x_1, lean_object* x_2, uint64_t x_3) { _start: { @@ -4858,34 +5137,117 @@ x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -LEAN_EXPORT uint64_t l_Lean_Meta_Grind_congrHash(lean_object* x_1, lean_object* x_2) { +static lean_object* _init_l_Lean_Meta_Grind_congrHash___closed__3() { _start: { -lean_object* x_3; lean_object* x_4; uint8_t x_5; -x_3 = l_Lean_Meta_Grind_congrHash___closed__2; -x_4 = lean_unsigned_to_nat(2u); -x_5 = l_Lean_Expr_isAppOfArity(x_2, x_3, x_4); -if (x_5 == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Eq", 2, 2); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_congrHash___closed__4() { +_start: { -uint64_t x_6; uint64_t x_7; -x_6 = 17; -x_7 = l_Lean_Meta_Grind_congrHash_go(x_1, x_2, x_6); -return x_7; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind_congrHash___closed__3; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT uint64_t l_Lean_Meta_Grind_congrHash(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; uint8_t x_4; +lean_inc(x_2); +x_3 = l_Lean_Expr_cleanupAnnotations(x_2); +x_4 = l_Lean_Expr_isApp(x_3); +if (x_4 == 0) +{ +uint64_t x_5; uint64_t x_6; +lean_dec(x_3); +x_5 = 17; +x_6 = l_Lean_Meta_Grind_congrHash_go(x_1, x_2, x_5); +lean_dec(x_2); +return x_6; } 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; uint64_t x_14; -x_8 = lean_unsigned_to_nat(0u); -x_9 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_2, x_8); -x_10 = lean_nat_sub(x_9, x_8); -lean_dec(x_9); -x_11 = lean_unsigned_to_nat(1u); -x_12 = lean_nat_sub(x_10, x_11); -lean_dec(x_10); -x_13 = l_Lean_Expr_getRevArg_x21(x_2, x_12); -x_14 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot(x_1, x_13); +lean_object* x_7; lean_object* x_8; uint8_t x_9; +x_7 = l_Lean_Expr_appArg(x_3, lean_box(0)); +x_8 = l_Lean_Expr_appFnCleanup(x_3, lean_box(0)); +x_9 = l_Lean_Expr_isApp(x_8); +if (x_9 == 0) +{ +uint64_t x_10; uint64_t x_11; +lean_dec(x_8); +lean_dec(x_7); +x_10 = 17; +x_11 = l_Lean_Meta_Grind_congrHash_go(x_1, x_2, x_10); +lean_dec(x_2); +return x_11; +} +else +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_12 = l_Lean_Expr_appArg(x_8, lean_box(0)); +x_13 = l_Lean_Expr_appFnCleanup(x_8, lean_box(0)); +x_14 = l_Lean_Meta_Grind_congrHash___closed__2; +x_15 = l_Lean_Expr_isConstOf(x_13, x_14); +if (x_15 == 0) +{ +uint8_t x_16; +x_16 = l_Lean_Expr_isApp(x_13); +if (x_16 == 0) +{ +uint64_t x_17; uint64_t x_18; lean_dec(x_13); -return x_14; +lean_dec(x_12); +lean_dec(x_7); +x_17 = 17; +x_18 = l_Lean_Meta_Grind_congrHash_go(x_1, x_2, x_17); +lean_dec(x_2); +return x_18; +} +else +{ +lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_19 = l_Lean_Expr_appFnCleanup(x_13, lean_box(0)); +x_20 = l_Lean_Meta_Grind_congrHash___closed__4; +x_21 = l_Lean_Expr_isConstOf(x_19, x_20); +lean_dec(x_19); +if (x_21 == 0) +{ +uint64_t x_22; uint64_t x_23; +lean_dec(x_12); +lean_dec(x_7); +x_22 = 17; +x_23 = l_Lean_Meta_Grind_congrHash_go(x_1, x_2, x_22); +lean_dec(x_2); +return x_23; +} +else +{ +uint64_t x_24; +lean_dec(x_2); +x_24 = l_Lean_Meta_Grind_congrHash_goEq(x_1, x_12, x_7); +lean_dec(x_7); +lean_dec(x_12); +return x_24; +} +} +} +else +{ +uint64_t x_25; +lean_dec(x_13); +lean_dec(x_7); +lean_dec(x_2); +x_25 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot(x_1, x_12); +lean_dec(x_12); +return x_25; +} +} } } } @@ -4894,11 +5256,82 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Grind_congrHash___boxed(lean_object* x_1, l { uint64_t x_3; lean_object* x_4; x_3 = l_Lean_Meta_Grind_congrHash(x_1, x_2); -lean_dec(x_2); x_4 = lean_box_uint64(x_3); return x_4; } } +LEAN_EXPORT uint8_t l_Lean_Meta_Grind_isCongruent_goEq(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_inc(x_1); +x_6 = l_Lean_Meta_Grind_hasSameRoot(x_1, x_2, x_4); +if (x_6 == 0) +{ +uint8_t x_7; +lean_inc(x_1); +x_7 = l_Lean_Meta_Grind_hasSameRoot(x_1, x_2, x_5); +if (x_7 == 0) +{ +uint8_t x_8; +lean_dec(x_1); +x_8 = 0; +return x_8; +} +else +{ +uint8_t x_9; +x_9 = l_Lean_Meta_Grind_hasSameRoot(x_1, x_3, x_4); +return x_9; +} +} +else +{ +uint8_t x_10; +lean_inc(x_1); +x_10 = l_Lean_Meta_Grind_hasSameRoot(x_1, x_3, x_5); +if (x_10 == 0) +{ +uint8_t x_11; +lean_inc(x_1); +x_11 = l_Lean_Meta_Grind_hasSameRoot(x_1, x_2, x_5); +if (x_11 == 0) +{ +uint8_t x_12; +lean_dec(x_1); +x_12 = 0; +return x_12; +} +else +{ +uint8_t x_13; +x_13 = l_Lean_Meta_Grind_hasSameRoot(x_1, x_3, x_4); +return x_13; +} +} +else +{ +uint8_t x_14; +lean_dec(x_1); +x_14 = 1; +return x_14; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isCongruent_goEq___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 = l_Lean_Meta_Grind_isCongruent_goEq(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); +x_7 = lean_box(x_6); +return x_7; +} +} LEAN_EXPORT uint8_t l_Lean_Meta_Grind_isCongruent_go(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { @@ -4907,7 +5340,7 @@ x_4 = l_Lean_Expr_isApp(x_2); if (x_4 == 0) { uint8_t x_5; -x_5 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hasSameRoot(x_1, x_2, x_3); +x_5 = l_Lean_Meta_Grind_hasSameRoot(x_1, x_2, x_3); lean_dec(x_3); lean_dec(x_2); return x_5; @@ -4919,7 +5352,7 @@ x_6 = l_Lean_Expr_isApp(x_3); if (x_6 == 0) { uint8_t x_7; -x_7 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hasSameRoot(x_1, x_2, x_3); +x_7 = l_Lean_Meta_Grind_hasSameRoot(x_1, x_2, x_3); lean_dec(x_3); lean_dec(x_2); return x_7; @@ -4930,7 +5363,7 @@ lean_object* x_8; lean_object* x_9; uint8_t x_10; x_8 = l_Lean_Expr_appArg_x21(x_2); x_9 = l_Lean_Expr_appArg_x21(x_3); lean_inc(x_1); -x_10 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hasSameRoot(x_1, x_8, x_9); +x_10 = l_Lean_Meta_Grind_hasSameRoot(x_1, x_8, x_9); lean_dec(x_9); lean_dec(x_8); if (x_10 == 0) @@ -4969,107 +5402,300 @@ return x_5; LEAN_EXPORT uint8_t l_Lean_Meta_Grind_isCongruent(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 = l_Lean_Meta_Grind_congrHash___closed__2; -x_5 = lean_unsigned_to_nat(2u); -x_6 = l_Lean_Expr_isAppOfArity(x_2, x_4, x_5); -if (x_6 == 0) +lean_object* x_4; uint8_t x_5; +lean_inc(x_2); +x_4 = l_Lean_Expr_cleanupAnnotations(x_2); +x_5 = l_Lean_Expr_isApp(x_4); +if (x_5 == 0) { -uint8_t x_7; -x_7 = l_Lean_Meta_Grind_isCongruent_go(x_1, x_2, x_3); -return x_7; +uint8_t x_6; +lean_dec(x_4); +x_6 = l_Lean_Meta_Grind_isCongruent_go(x_1, x_2, x_3); +return x_6; } else { -uint8_t x_8; -x_8 = l_Lean_Expr_isAppOfArity(x_3, x_4, x_5); -if (x_8 == 0) +lean_object* x_7; lean_object* x_8; uint8_t x_9; +x_7 = l_Lean_Expr_appArg(x_4, lean_box(0)); +x_8 = l_Lean_Expr_appFnCleanup(x_4, lean_box(0)); +x_9 = l_Lean_Expr_isApp(x_8); +if (x_9 == 0) { -uint8_t x_9; -x_9 = l_Lean_Meta_Grind_isCongruent_go(x_1, x_2, x_3); -return x_9; +uint8_t x_10; +lean_dec(x_8); +lean_dec(x_7); +x_10 = l_Lean_Meta_Grind_isCongruent_go(x_1, x_2, x_3); +return x_10; } else { -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; -x_10 = lean_unsigned_to_nat(0u); -x_11 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_2, x_10); -x_12 = lean_nat_sub(x_11, x_10); -lean_dec(x_11); -x_13 = lean_unsigned_to_nat(1u); -x_14 = lean_nat_sub(x_12, x_13); +lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_11 = l_Lean_Expr_appArg(x_8, lean_box(0)); +x_12 = l_Lean_Expr_appFnCleanup(x_8, lean_box(0)); +x_13 = l_Lean_Meta_Grind_congrHash___closed__2; +x_14 = l_Lean_Expr_isConstOf(x_12, x_13); +if (x_14 == 0) +{ +uint8_t x_15; +x_15 = l_Lean_Expr_isApp(x_12); +if (x_15 == 0) +{ +uint8_t x_16; lean_dec(x_12); -x_15 = l_Lean_Expr_getRevArg_x21(x_2, x_14); -lean_dec(x_2); -x_16 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_3, x_10); -x_17 = lean_nat_sub(x_16, x_10); -lean_dec(x_16); -x_18 = lean_nat_sub(x_17, x_13); -lean_dec(x_17); -x_19 = l_Lean_Expr_getRevArg_x21(x_3, x_18); -lean_dec(x_3); -x_20 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hasSameRoot(x_1, x_15, x_19); -lean_dec(x_19); -lean_dec(x_15); -return x_20; -} -} -} +lean_dec(x_11); +lean_dec(x_7); +x_16 = l_Lean_Meta_Grind_isCongruent_go(x_1, x_2, x_3); +return x_16; } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isCongruent___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: +else { -uint8_t x_4; lean_object* x_5; -x_4 = l_Lean_Meta_Grind_isCongruent(x_1, x_2, x_3); -x_5 = lean_box(x_4); -return x_5; -} -} -LEAN_EXPORT uint64_t l_Lean_Meta_Grind_instHashableCongrKey(lean_object* x_1, lean_object* x_2) { -_start: +lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_17 = l_Lean_Expr_appArg(x_12, lean_box(0)); +x_18 = l_Lean_Expr_appFnCleanup(x_12, lean_box(0)); +x_19 = l_Lean_Meta_Grind_congrHash___closed__4; +x_20 = l_Lean_Expr_isConstOf(x_18, x_19); +lean_dec(x_18); +if (x_20 == 0) { -uint64_t x_3; -x_3 = l_Lean_Meta_Grind_congrHash(x_1, x_2); -return x_3; -} +uint8_t x_21; +lean_dec(x_17); +lean_dec(x_11); +lean_dec(x_7); +x_21 = l_Lean_Meta_Grind_isCongruent_go(x_1, x_2, x_3); +return x_21; } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instHashableCongrKey___boxed(lean_object* x_1, lean_object* x_2) { -_start: +else { -uint64_t x_3; lean_object* x_4; -x_3 = l_Lean_Meta_Grind_instHashableCongrKey(x_1, x_2); +lean_object* x_22; uint8_t x_23; +lean_inc(x_3); +x_22 = l_Lean_Expr_cleanupAnnotations(x_3); +x_23 = l_Lean_Expr_isApp(x_22); +if (x_23 == 0) +{ +uint8_t x_24; +lean_dec(x_22); +lean_dec(x_17); +lean_dec(x_11); +lean_dec(x_7); +lean_dec(x_3); lean_dec(x_2); -x_4 = lean_box_uint64(x_3); -return x_4; -} +lean_dec(x_1); +x_24 = 0; +return x_24; } -LEAN_EXPORT uint8_t l_Lean_Meta_Grind_instBEqCongrKey(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: +else { -uint8_t x_4; -x_4 = l_Lean_Meta_Grind_isCongruent(x_1, x_2, x_3); -return x_4; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instBEqCongrKey___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: +lean_object* x_25; lean_object* x_26; uint8_t x_27; +x_25 = l_Lean_Expr_appArg(x_22, lean_box(0)); +x_26 = l_Lean_Expr_appFnCleanup(x_22, lean_box(0)); +x_27 = l_Lean_Expr_isApp(x_26); +if (x_27 == 0) { -uint8_t x_4; lean_object* x_5; -x_4 = l_Lean_Meta_Grind_instBEqCongrKey(x_1, x_2, x_3); -x_5 = lean_box(x_4); -return x_5; -} +uint8_t x_28; +lean_dec(x_26); +lean_dec(x_25); +lean_dec(x_17); +lean_dec(x_11); +lean_dec(x_7); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_28 = 0; +return x_28; } -LEAN_EXPORT uint64_t l_Lean_Meta_Grind_instHashablePreInstance_unsafe__1(lean_object* x_1) { -_start: +else { -lean_object* x_2; size_t x_3; size_t x_4; size_t x_5; uint64_t x_6; -x_2 = lean_ctor_get(x_1, 0); -x_3 = lean_ptr_addr(x_2); -x_4 = 3; -x_5 = lean_usize_shift_right(x_3, x_4); -x_6 = lean_usize_to_uint64(x_5); -return x_6; +lean_object* x_29; lean_object* x_30; uint8_t x_31; +x_29 = l_Lean_Expr_appArg(x_26, lean_box(0)); +x_30 = l_Lean_Expr_appFnCleanup(x_26, lean_box(0)); +x_31 = l_Lean_Expr_isApp(x_30); +if (x_31 == 0) +{ +uint8_t x_32; +lean_dec(x_30); +lean_dec(x_29); +lean_dec(x_25); +lean_dec(x_17); +lean_dec(x_11); +lean_dec(x_7); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_32 = 0; +return x_32; +} +else +{ +lean_object* x_33; lean_object* x_34; uint8_t x_35; +x_33 = l_Lean_Expr_appArg(x_30, lean_box(0)); +x_34 = l_Lean_Expr_appFnCleanup(x_30, lean_box(0)); +x_35 = l_Lean_Expr_isConstOf(x_34, x_19); +lean_dec(x_34); +if (x_35 == 0) +{ +uint8_t x_36; +lean_dec(x_33); +lean_dec(x_29); +lean_dec(x_25); +lean_dec(x_17); +lean_dec(x_11); +lean_dec(x_7); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_36 = 0; +return x_36; +} +else +{ +uint8_t x_37; +x_37 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_17, x_33); +lean_dec(x_33); +lean_dec(x_17); +if (x_37 == 0) +{ +uint8_t x_38; +lean_dec(x_29); +lean_dec(x_25); +lean_dec(x_11); +lean_dec(x_7); +x_38 = l_Lean_Meta_Grind_isCongruent_go(x_1, x_2, x_3); +return x_38; +} +else +{ +uint8_t x_39; +lean_dec(x_3); +lean_dec(x_2); +x_39 = l_Lean_Meta_Grind_isCongruent_goEq(x_1, x_11, x_7, x_29, x_25); +lean_dec(x_25); +lean_dec(x_29); +lean_dec(x_7); +lean_dec(x_11); +return x_39; +} +} +} +} +} +} +} +} +else +{ +lean_object* x_40; uint8_t x_41; +lean_dec(x_12); +lean_dec(x_7); +lean_dec(x_2); +x_40 = l_Lean_Expr_cleanupAnnotations(x_3); +x_41 = l_Lean_Expr_isApp(x_40); +if (x_41 == 0) +{ +uint8_t x_42; +lean_dec(x_40); +lean_dec(x_11); +lean_dec(x_1); +x_42 = 0; +return x_42; +} +else +{ +lean_object* x_43; uint8_t x_44; +x_43 = l_Lean_Expr_appFnCleanup(x_40, lean_box(0)); +x_44 = l_Lean_Expr_isApp(x_43); +if (x_44 == 0) +{ +uint8_t x_45; +lean_dec(x_43); +lean_dec(x_11); +lean_dec(x_1); +x_45 = 0; +return x_45; +} +else +{ +lean_object* x_46; lean_object* x_47; uint8_t x_48; +x_46 = l_Lean_Expr_appArg(x_43, lean_box(0)); +x_47 = l_Lean_Expr_appFnCleanup(x_43, lean_box(0)); +x_48 = l_Lean_Expr_isConstOf(x_47, x_13); +lean_dec(x_47); +if (x_48 == 0) +{ +uint8_t x_49; +lean_dec(x_46); +lean_dec(x_11); +lean_dec(x_1); +x_49 = 0; +return x_49; +} +else +{ +uint8_t x_50; +x_50 = l_Lean_Meta_Grind_hasSameRoot(x_1, x_11, x_46); +lean_dec(x_46); +lean_dec(x_11); +return x_50; +} +} +} +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isCongruent___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_Lean_Meta_Grind_isCongruent(x_1, x_2, x_3); +x_5 = lean_box(x_4); +return x_5; +} +} +LEAN_EXPORT uint64_t l_Lean_Meta_Grind_instHashableCongrKey(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint64_t x_3; +x_3 = l_Lean_Meta_Grind_congrHash(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instHashableCongrKey___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint64_t x_3; lean_object* x_4; +x_3 = l_Lean_Meta_Grind_instHashableCongrKey(x_1, x_2); +x_4 = lean_box_uint64(x_3); +return x_4; +} +} +LEAN_EXPORT uint8_t l_Lean_Meta_Grind_instBEqCongrKey(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; +x_4 = l_Lean_Meta_Grind_isCongruent(x_1, x_2, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instBEqCongrKey___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_Lean_Meta_Grind_instBEqCongrKey(x_1, x_2, x_3); +x_5 = lean_box(x_4); +return x_5; +} +} +LEAN_EXPORT uint64_t l_Lean_Meta_Grind_instHashablePreInstance_unsafe__1(lean_object* x_1) { +_start: +{ +lean_object* x_2; size_t x_3; size_t x_4; size_t x_5; uint64_t x_6; +x_2 = lean_ctor_get(x_1, 0); +x_3 = lean_ptr_addr(x_2); +x_4 = 3; +x_5 = lean_usize_shift_right(x_3, x_4); +x_6 = lean_usize_to_uint64(x_5); +return x_6; } } LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instHashablePreInstance_unsafe__1___boxed(lean_object* x_1) { @@ -5585,6 +6211,14 @@ x_1 = l_Lean_Meta_Grind_canon___closed__2; return x_1; } } +static lean_object* _init_l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedGoal___spec__3() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Meta_Grind_canon___closed__2; +return x_1; +} +} static lean_object* _init_l_Lean_Meta_Grind_instInhabitedGoal___closed__1() { _start: { @@ -5633,6 +6267,19 @@ return x_5; static lean_object* _init_l_Lean_Meta_Grind_instInhabitedGoal___closed__5() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_canon___closed__2; +x_2 = l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedEMatchTheorems___spec__1; +x_3 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +lean_ctor_set(x_3, 2, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind_instInhabitedGoal___closed__6() { +_start: +{ lean_object* x_1; x_1 = l_Std_Queue_empty(lean_box(0)); return x_1; @@ -5641,33 +6288,43 @@ return x_1; static lean_object* _init_l_Lean_Meta_Grind_instInhabitedGoal() { _start: { -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_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; uint8_t x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; x_1 = lean_box(0); -x_2 = l_Lean_Meta_Grind_canon___closed__2; -x_3 = l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedGoal___spec__1; -x_4 = l_Lean_Meta_Grind_instInhabitedGoal___closed__1; -x_5 = 0; -x_6 = lean_unsigned_to_nat(0u); -x_7 = l_Lean_Meta_Grind_instInhabitedGoal___closed__4; -x_8 = l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedGoal___spec__2; +x_2 = lean_box(0); +x_3 = l_Lean_Meta_Grind_canon___closed__2; +x_4 = l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedGoal___spec__1; +x_5 = l_Lean_Meta_Grind_instInhabitedGoal___closed__1; +x_6 = 0; +x_7 = lean_unsigned_to_nat(0u); +x_8 = l_Lean_Meta_Grind_instInhabitedGoal___closed__4; x_9 = l_Lean_Meta_Grind_instInhabitedGoal___closed__5; -x_10 = lean_alloc_ctor(0, 14, 1); -lean_ctor_set(x_10, 0, x_1); -lean_ctor_set(x_10, 1, x_2); -lean_ctor_set(x_10, 2, x_2); -lean_ctor_set(x_10, 3, x_3); -lean_ctor_set(x_10, 4, x_2); -lean_ctor_set(x_10, 5, x_4); -lean_ctor_set(x_10, 6, x_6); -lean_ctor_set(x_10, 7, x_6); -lean_ctor_set(x_10, 8, x_7); -lean_ctor_set(x_10, 9, x_7); -lean_ctor_set(x_10, 10, x_2); -lean_ctor_set(x_10, 11, x_6); -lean_ctor_set(x_10, 12, x_8); -lean_ctor_set(x_10, 13, x_9); -lean_ctor_set_uint8(x_10, sizeof(void*)*14, x_5); -return x_10; +x_10 = l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedGoal___spec__2; +x_11 = l_Lean_Meta_Grind_instInhabitedGoal___closed__6; +x_12 = l_Lean_PersistentHashMap_empty___at_Lean_KeyedDeclsAttribute_instInhabitedExtensionState___spec__1; +x_13 = l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedGoal___spec__3; +x_14 = lean_alloc_ctor(0, 20, 1); +lean_ctor_set(x_14, 0, x_2); +lean_ctor_set(x_14, 1, x_3); +lean_ctor_set(x_14, 2, x_3); +lean_ctor_set(x_14, 3, x_4); +lean_ctor_set(x_14, 4, x_3); +lean_ctor_set(x_14, 5, x_5); +lean_ctor_set(x_14, 6, x_7); +lean_ctor_set(x_14, 7, x_7); +lean_ctor_set(x_14, 8, x_8); +lean_ctor_set(x_14, 9, x_8); +lean_ctor_set(x_14, 10, x_9); +lean_ctor_set(x_14, 11, x_7); +lean_ctor_set(x_14, 12, x_7); +lean_ctor_set(x_14, 13, x_10); +lean_ctor_set(x_14, 14, x_11); +lean_ctor_set(x_14, 15, x_12); +lean_ctor_set(x_14, 16, x_1); +lean_ctor_set(x_14, 17, x_7); +lean_ctor_set(x_14, 18, x_13); +lean_ctor_set(x_14, 19, x_7); +lean_ctor_set_uint8(x_14, sizeof(void*)*20, x_6); +return x_14; } } LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Goal_admit(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { @@ -6113,4942 +6770,4728 @@ lean_dec(x_2); return x_10; } } -LEAN_EXPORT uint64_t l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, uint64_t x_6) { +LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___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: { -uint8_t x_7; -x_7 = lean_usize_dec_lt(x_5, x_4); -if (x_7 == 0) +lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_11 = lean_ctor_get(x_8, 2); +x_12 = l_Lean_isTracingEnabledForCore(x_1, x_11, x_10); +x_13 = !lean_is_exclusive(x_12); +if (x_13 == 0) { -return x_6; +return x_12; } else { -lean_object* x_8; uint64_t x_9; uint64_t x_10; size_t x_11; size_t x_12; -x_8 = lean_array_uget(x_3, x_5); -x_9 = l_Lean_Meta_Grind_instHashablePreInstance_unsafe__2(x_8); -lean_dec(x_8); -x_10 = lean_uint64_mix_hash(x_6, x_9); -x_11 = 1; -x_12 = lean_usize_add(x_5, x_11); -x_5 = x_12; -x_6 = x_10; -goto _start; +lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_14 = lean_ctor_get(x_12, 0); +x_15 = lean_ctor_get(x_12, 1); +lean_inc(x_15); +lean_inc(x_14); +lean_dec(x_12); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_14); +lean_ctor_set(x_16, 1, x_15); +return x_16; } } } -LEAN_EXPORT uint64_t l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, uint64_t x_6) { +static double _init_l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2___closed__1() { _start: { -uint8_t x_7; -x_7 = lean_usize_dec_lt(x_5, x_4); -if (x_7 == 0) -{ -return x_6; +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; } -else -{ -lean_object* x_8; uint64_t x_9; uint64_t x_10; size_t x_11; size_t x_12; -x_8 = lean_array_uget(x_3, x_5); -x_9 = l_Lean_Meta_Grind_instHashablePreInstance_unsafe__2(x_8); -lean_dec(x_8); -x_10 = lean_uint64_mix_hash(x_6, x_9); -x_11 = 1; -x_12 = lean_usize_add(x_5, x_11); -x_5 = x_12; -x_6 = x_10; -goto _start; } +static lean_object* _init_l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("", 0, 0); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_markTheoremInstance___spec__4(size_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +static lean_object* _init_l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2___closed__3() { _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; +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(0); +x_2 = lean_array_mk(x_1); +return x_2; } -else +} +LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: { -lean_object* x_9; lean_object* x_10; uint64_t x_11; lean_object* x_12; lean_object* x_13; size_t x_14; size_t x_15; uint64_t x_16; size_t x_17; size_t x_18; size_t x_19; size_t x_20; size_t x_21; size_t x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; -x_9 = lean_array_fget(x_2, x_5); -x_10 = lean_array_fget(x_3, x_5); -x_11 = l_Lean_Meta_Grind_instHashablePreInstance_unsafe__1(x_9); -x_12 = lean_ctor_get(x_9, 1); -lean_inc(x_12); -x_13 = lean_box(0); -x_14 = lean_array_size(x_12); -x_15 = 0; -x_16 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__5(x_12, x_13, x_12, x_14, x_15, x_11); -lean_dec(x_12); -x_17 = lean_uint64_to_usize(x_16); -x_18 = 1; -x_19 = lean_usize_sub(x_1, x_18); -x_20 = 5; -x_21 = lean_usize_mul(x_20, x_19); -x_22 = lean_usize_shift_right(x_17, x_21); -x_23 = lean_unsigned_to_nat(1u); -x_24 = lean_nat_add(x_5, x_23); -lean_dec(x_5); -x_25 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3(x_6, x_22, x_1, x_9, x_10); -x_4 = lean_box(0); -x_5 = x_24; -x_6 = x_25; -goto _start; -} -} -} -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, size_t x_6, lean_object* x_7) { -_start: -{ -uint8_t x_8; -x_8 = lean_usize_dec_lt(x_6, x_5); -if (x_8 == 0) -{ -lean_dec(x_3); -return x_7; -} -else -{ -lean_object* x_9; uint8_t x_10; -x_9 = lean_array_uget(x_4, x_6); -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; uint8_t x_16; -x_11 = lean_ctor_get(x_7, 1); -x_12 = lean_ctor_get(x_7, 0); -lean_dec(x_12); -x_13 = lean_ctor_get(x_11, 0); -lean_inc(x_13); -x_14 = lean_ctor_get(x_11, 1); +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; +x_12 = lean_ctor_get(x_9, 5); +x_13 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_2, x_7, x_8, x_9, x_10, x_11); +x_14 = lean_ctor_get(x_13, 0); lean_inc(x_14); -x_15 = lean_ctor_get(x_11, 2); +x_15 = lean_ctor_get(x_13, 1); lean_inc(x_15); -x_16 = lean_nat_dec_lt(x_14, x_15); -if (x_16 == 0) -{ -lean_dec(x_15); -lean_dec(x_14); lean_dec(x_13); -lean_dec(x_9); -lean_ctor_set(x_7, 0, x_3); -return x_7; -} -else -{ -uint8_t x_17; -x_17 = !lean_is_exclusive(x_11); -if (x_17 == 0) +x_16 = lean_st_ref_take(x_10, x_15); +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +x_18 = lean_ctor_get(x_17, 3); +lean_inc(x_18); +x_19 = !lean_is_exclusive(x_16); +if (x_19 == 0) { -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; -x_18 = lean_ctor_get(x_11, 2); -lean_dec(x_18); -x_19 = lean_ctor_get(x_11, 1); -lean_dec(x_19); -x_20 = lean_ctor_get(x_11, 0); -lean_dec(x_20); -x_21 = lean_array_fget(x_13, x_14); -x_22 = lean_unsigned_to_nat(1u); -x_23 = lean_nat_add(x_14, x_22); -lean_dec(x_14); -lean_ctor_set(x_11, 1, x_23); -x_24 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_9, x_21); +lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_20 = lean_ctor_get(x_16, 1); +x_21 = lean_ctor_get(x_16, 0); lean_dec(x_21); -lean_dec(x_9); +x_22 = !lean_is_exclusive(x_17); +if (x_22 == 0) +{ +lean_object* x_23; uint8_t x_24; +x_23 = lean_ctor_get(x_17, 3); +lean_dec(x_23); +x_24 = !lean_is_exclusive(x_18); if (x_24 == 0) { -lean_object* x_25; -lean_dec(x_3); -x_25 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1___closed__1; -lean_ctor_set(x_7, 0, x_25); -return x_7; +lean_object* x_25; double 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; uint8_t x_34; +x_25 = lean_ctor_get(x_18, 0); +x_26 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2___closed__1; +x_27 = 0; +x_28 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2___closed__2; +x_29 = lean_alloc_ctor(0, 2, 17); +lean_ctor_set(x_29, 0, x_1); +lean_ctor_set(x_29, 1, x_28); +lean_ctor_set_float(x_29, sizeof(void*)*2, x_26); +lean_ctor_set_float(x_29, sizeof(void*)*2 + 8, x_26); +lean_ctor_set_uint8(x_29, sizeof(void*)*2 + 16, x_27); +x_30 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2___closed__3; +x_31 = lean_alloc_ctor(9, 3, 0); +lean_ctor_set(x_31, 0, x_29); +lean_ctor_set(x_31, 1, x_14); +lean_ctor_set(x_31, 2, x_30); +lean_inc(x_12); +lean_ctor_set(x_16, 1, x_31); +lean_ctor_set(x_16, 0, x_12); +x_32 = l_Lean_PersistentArray_push___rarg(x_25, x_16); +lean_ctor_set(x_18, 0, x_32); +x_33 = lean_st_ref_set(x_10, x_17, x_20); +x_34 = !lean_is_exclusive(x_33); +if (x_34 == 0) +{ +lean_object* x_35; lean_object* x_36; +x_35 = lean_ctor_get(x_33, 0); +lean_dec(x_35); +x_36 = lean_box(0); +lean_ctor_set(x_33, 0, x_36); +return x_33; } else { -size_t x_26; size_t x_27; -lean_inc(x_3); -lean_ctor_set(x_7, 0, x_3); -x_26 = 1; -x_27 = lean_usize_add(x_6, x_26); -x_6 = x_27; -goto _start; +lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_37 = lean_ctor_get(x_33, 1); +lean_inc(x_37); +lean_dec(x_33); +x_38 = lean_box(0); +x_39 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_37); +return x_39; } } else { -lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; -lean_dec(x_11); -x_29 = lean_array_fget(x_13, x_14); -x_30 = lean_unsigned_to_nat(1u); -x_31 = lean_nat_add(x_14, x_30); -lean_dec(x_14); -x_32 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_32, 0, x_13); -lean_ctor_set(x_32, 1, x_31); -lean_ctor_set(x_32, 2, x_15); -x_33 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_9, x_29); -lean_dec(x_29); -lean_dec(x_9); -if (x_33 == 0) -{ -lean_object* x_34; -lean_dec(x_3); -x_34 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1___closed__1; -lean_ctor_set(x_7, 1, x_32); -lean_ctor_set(x_7, 0, x_34); -return x_7; +uint64_t x_40; lean_object* x_41; double 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; +x_40 = lean_ctor_get_uint64(x_18, sizeof(void*)*1); +x_41 = lean_ctor_get(x_18, 0); +lean_inc(x_41); +lean_dec(x_18); +x_42 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2___closed__1; +x_43 = 0; +x_44 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2___closed__2; +x_45 = lean_alloc_ctor(0, 2, 17); +lean_ctor_set(x_45, 0, x_1); +lean_ctor_set(x_45, 1, x_44); +lean_ctor_set_float(x_45, sizeof(void*)*2, x_42); +lean_ctor_set_float(x_45, sizeof(void*)*2 + 8, x_42); +lean_ctor_set_uint8(x_45, sizeof(void*)*2 + 16, x_43); +x_46 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2___closed__3; +x_47 = lean_alloc_ctor(9, 3, 0); +lean_ctor_set(x_47, 0, x_45); +lean_ctor_set(x_47, 1, x_14); +lean_ctor_set(x_47, 2, x_46); +lean_inc(x_12); +lean_ctor_set(x_16, 1, x_47); +lean_ctor_set(x_16, 0, x_12); +x_48 = l_Lean_PersistentArray_push___rarg(x_41, x_16); +x_49 = lean_alloc_ctor(0, 1, 8); +lean_ctor_set(x_49, 0, x_48); +lean_ctor_set_uint64(x_49, sizeof(void*)*1, x_40); +lean_ctor_set(x_17, 3, x_49); +x_50 = lean_st_ref_set(x_10, x_17, x_20); +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 { -size_t x_35; size_t x_36; -lean_inc(x_3); -lean_ctor_set(x_7, 1, x_32); -lean_ctor_set(x_7, 0, x_3); -x_35 = 1; -x_36 = lean_usize_add(x_6, x_35); -x_6 = x_36; -goto _start; +lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; uint64_t x_62; lean_object* x_63; lean_object* x_64; double x_65; uint8_t x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; +x_55 = lean_ctor_get(x_17, 0); +x_56 = lean_ctor_get(x_17, 1); +x_57 = lean_ctor_get(x_17, 2); +x_58 = lean_ctor_get(x_17, 4); +x_59 = lean_ctor_get(x_17, 5); +x_60 = lean_ctor_get(x_17, 6); +x_61 = lean_ctor_get(x_17, 7); +lean_inc(x_61); +lean_inc(x_60); +lean_inc(x_59); +lean_inc(x_58); +lean_inc(x_57); +lean_inc(x_56); +lean_inc(x_55); +lean_dec(x_17); +x_62 = lean_ctor_get_uint64(x_18, sizeof(void*)*1); +x_63 = lean_ctor_get(x_18, 0); +lean_inc(x_63); +if (lean_is_exclusive(x_18)) { + lean_ctor_release(x_18, 0); + x_64 = x_18; +} else { + lean_dec_ref(x_18); + x_64 = lean_box(0); +} +x_65 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2___closed__1; +x_66 = 0; +x_67 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2___closed__2; +x_68 = lean_alloc_ctor(0, 2, 17); +lean_ctor_set(x_68, 0, x_1); +lean_ctor_set(x_68, 1, x_67); +lean_ctor_set_float(x_68, sizeof(void*)*2, x_65); +lean_ctor_set_float(x_68, sizeof(void*)*2 + 8, x_65); +lean_ctor_set_uint8(x_68, sizeof(void*)*2 + 16, x_66); +x_69 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2___closed__3; +x_70 = lean_alloc_ctor(9, 3, 0); +lean_ctor_set(x_70, 0, x_68); +lean_ctor_set(x_70, 1, x_14); +lean_ctor_set(x_70, 2, x_69); +lean_inc(x_12); +lean_ctor_set(x_16, 1, x_70); +lean_ctor_set(x_16, 0, x_12); +x_71 = l_Lean_PersistentArray_push___rarg(x_63, x_16); +if (lean_is_scalar(x_64)) { + x_72 = lean_alloc_ctor(0, 1, 8); +} else { + x_72 = x_64; +} +lean_ctor_set(x_72, 0, x_71); +lean_ctor_set_uint64(x_72, sizeof(void*)*1, x_62); +x_73 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_73, 0, x_55); +lean_ctor_set(x_73, 1, x_56); +lean_ctor_set(x_73, 2, x_57); +lean_ctor_set(x_73, 3, x_72); +lean_ctor_set(x_73, 4, x_58); +lean_ctor_set(x_73, 5, x_59); +lean_ctor_set(x_73, 6, x_60); +lean_ctor_set(x_73, 7, x_61); +x_74 = lean_st_ref_set(x_10, x_73, x_20); +x_75 = lean_ctor_get(x_74, 1); +lean_inc(x_75); +if (lean_is_exclusive(x_74)) { + lean_ctor_release(x_74, 0); + lean_ctor_release(x_74, 1); + x_76 = x_74; +} else { + lean_dec_ref(x_74); + x_76 = lean_box(0); } +x_77 = lean_box(0); +if (lean_is_scalar(x_76)) { + x_78 = lean_alloc_ctor(0, 2, 0); +} else { + x_78 = x_76; } +lean_ctor_set(x_78, 0, x_77); +lean_ctor_set(x_78, 1, x_75); +return x_78; } } else { -lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; uint8_t x_42; -x_38 = lean_ctor_get(x_7, 1); -lean_inc(x_38); -lean_dec(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); -x_41 = lean_ctor_get(x_38, 2); -lean_inc(x_41); -x_42 = lean_nat_dec_lt(x_40, x_41); -if (x_42 == 0) -{ -lean_object* x_43; -lean_dec(x_41); -lean_dec(x_40); -lean_dec(x_39); -lean_dec(x_9); -x_43 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_43, 0, x_3); -lean_ctor_set(x_43, 1, x_38); -return x_43; +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; uint64_t x_88; lean_object* x_89; lean_object* x_90; double x_91; uint8_t x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; +x_79 = lean_ctor_get(x_16, 1); +lean_inc(x_79); +lean_dec(x_16); +x_80 = lean_ctor_get(x_17, 0); +lean_inc(x_80); +x_81 = lean_ctor_get(x_17, 1); +lean_inc(x_81); +x_82 = lean_ctor_get(x_17, 2); +lean_inc(x_82); +x_83 = lean_ctor_get(x_17, 4); +lean_inc(x_83); +x_84 = lean_ctor_get(x_17, 5); +lean_inc(x_84); +x_85 = lean_ctor_get(x_17, 6); +lean_inc(x_85); +x_86 = lean_ctor_get(x_17, 7); +lean_inc(x_86); +if (lean_is_exclusive(x_17)) { + lean_ctor_release(x_17, 0); + lean_ctor_release(x_17, 1); + lean_ctor_release(x_17, 2); + lean_ctor_release(x_17, 3); + lean_ctor_release(x_17, 4); + lean_ctor_release(x_17, 5); + lean_ctor_release(x_17, 6); + lean_ctor_release(x_17, 7); + x_87 = x_17; +} else { + lean_dec_ref(x_17); + x_87 = lean_box(0); +} +x_88 = lean_ctor_get_uint64(x_18, sizeof(void*)*1); +x_89 = lean_ctor_get(x_18, 0); +lean_inc(x_89); +if (lean_is_exclusive(x_18)) { + lean_ctor_release(x_18, 0); + x_90 = x_18; +} else { + lean_dec_ref(x_18); + x_90 = lean_box(0); +} +x_91 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2___closed__1; +x_92 = 0; +x_93 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2___closed__2; +x_94 = lean_alloc_ctor(0, 2, 17); +lean_ctor_set(x_94, 0, x_1); +lean_ctor_set(x_94, 1, x_93); +lean_ctor_set_float(x_94, sizeof(void*)*2, x_91); +lean_ctor_set_float(x_94, sizeof(void*)*2 + 8, x_91); +lean_ctor_set_uint8(x_94, sizeof(void*)*2 + 16, x_92); +x_95 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2___closed__3; +x_96 = lean_alloc_ctor(9, 3, 0); +lean_ctor_set(x_96, 0, x_94); +lean_ctor_set(x_96, 1, x_14); +lean_ctor_set(x_96, 2, x_95); +lean_inc(x_12); +x_97 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_97, 0, x_12); +lean_ctor_set(x_97, 1, x_96); +x_98 = l_Lean_PersistentArray_push___rarg(x_89, x_97); +if (lean_is_scalar(x_90)) { + x_99 = lean_alloc_ctor(0, 1, 8); +} else { + x_99 = x_90; } -else -{ -lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; uint8_t x_49; -if (lean_is_exclusive(x_38)) { - lean_ctor_release(x_38, 0); - lean_ctor_release(x_38, 1); - lean_ctor_release(x_38, 2); - x_44 = x_38; +lean_ctor_set(x_99, 0, x_98); +lean_ctor_set_uint64(x_99, sizeof(void*)*1, x_88); +if (lean_is_scalar(x_87)) { + x_100 = lean_alloc_ctor(0, 8, 0); } else { - lean_dec_ref(x_38); - x_44 = lean_box(0); + x_100 = x_87; +} +lean_ctor_set(x_100, 0, x_80); +lean_ctor_set(x_100, 1, x_81); +lean_ctor_set(x_100, 2, x_82); +lean_ctor_set(x_100, 3, x_99); +lean_ctor_set(x_100, 4, x_83); +lean_ctor_set(x_100, 5, x_84); +lean_ctor_set(x_100, 6, x_85); +lean_ctor_set(x_100, 7, x_86); +x_101 = lean_st_ref_set(x_10, x_100, x_79); +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; +} else { + lean_dec_ref(x_101); + x_103 = lean_box(0); } -x_45 = lean_array_fget(x_39, x_40); -x_46 = lean_unsigned_to_nat(1u); -x_47 = lean_nat_add(x_40, x_46); -lean_dec(x_40); -if (lean_is_scalar(x_44)) { - x_48 = lean_alloc_ctor(0, 3, 0); +x_104 = lean_box(0); +if (lean_is_scalar(x_103)) { + x_105 = lean_alloc_ctor(0, 2, 0); } else { - x_48 = x_44; + x_105 = x_103; } -lean_ctor_set(x_48, 0, x_39); -lean_ctor_set(x_48, 1, x_47); -lean_ctor_set(x_48, 2, x_41); -x_49 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_9, x_45); -lean_dec(x_45); -lean_dec(x_9); -if (x_49 == 0) -{ -lean_object* x_50; lean_object* x_51; -lean_dec(x_3); -x_50 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1___closed__1; -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_ctor_set(x_105, 0, x_104); +lean_ctor_set(x_105, 1, x_102); +return x_105; } -else +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_updateLastTag___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_52; size_t x_53; size_t x_54; -lean_inc(x_3); -x_52 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_52, 0, x_3); -lean_ctor_set(x_52, 1, x_48); -x_53 = 1; -x_54 = lean_usize_add(x_6, x_53); -x_6 = x_54; -x_7 = x_52; -goto _start; +lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_12 = lean_st_ref_take(x_6, x_11); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_12, 1); +lean_inc(x_14); +lean_dec(x_12); +x_15 = !lean_is_exclusive(x_13); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_16 = lean_ctor_get(x_13, 7); +lean_dec(x_16); +lean_ctor_set(x_13, 7, x_1); +x_17 = lean_st_ref_set(x_6, x_13, x_14); +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; +x_24 = lean_ctor_get(x_13, 0); +x_25 = lean_ctor_get(x_13, 1); +x_26 = lean_ctor_get(x_13, 2); +x_27 = lean_ctor_get(x_13, 3); +x_28 = lean_ctor_get(x_13, 4); +x_29 = lean_ctor_get(x_13, 5); +x_30 = lean_ctor_get(x_13, 6); +lean_inc(x_30); +lean_inc(x_29); +lean_inc(x_28); +lean_inc(x_27); +lean_inc(x_26); +lean_inc(x_25); +lean_inc(x_24); +lean_dec(x_13); +x_31 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_31, 0, x_24); +lean_ctor_set(x_31, 1, x_25); +lean_ctor_set(x_31, 2, x_26); +lean_ctor_set(x_31, 3, x_27); +lean_ctor_set(x_31, 4, x_28); +lean_ctor_set(x_31, 5, x_29); +lean_ctor_set(x_31, 6, x_30); +lean_ctor_set(x_31, 7, x_1); +x_32 = lean_st_ref_set(x_6, x_31, x_14); +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_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___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) { +static lean_object* _init_l_Lean_Meta_Grind_updateLastTag___closed__1() { _start: { -uint8_t x_8; -x_8 = lean_usize_dec_lt(x_6, x_5); -if (x_8 == 0) +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78____closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind_updateLastTag___closed__2() { +_start: { -lean_dec(x_3); -return x_7; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("working on goal `", 17, 17); +return x_1; } -else +} +static lean_object* _init_l_Lean_Meta_Grind_updateLastTag___closed__3() { +_start: { -lean_object* x_9; uint8_t x_10; -x_9 = lean_array_uget(x_4, x_6); -x_10 = !lean_is_exclusive(x_7); -if (x_10 == 0) +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_updateLastTag___closed__2; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Grind_updateLastTag___closed__4() { +_start: { -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; -x_11 = lean_ctor_get(x_7, 1); -x_12 = lean_ctor_get(x_7, 0); -lean_dec(x_12); -x_13 = lean_ctor_get(x_11, 0); -lean_inc(x_13); -x_14 = lean_ctor_get(x_11, 1); -lean_inc(x_14); -x_15 = lean_ctor_get(x_11, 2); -lean_inc(x_15); -x_16 = lean_nat_dec_lt(x_14, x_15); -if (x_16 == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("`", 1, 1); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_updateLastTag___closed__5() { +_start: { -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_9); -lean_ctor_set(x_7, 0, x_3); -return x_7; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_updateLastTag___closed__4; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } -else +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_updateLastTag(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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_17; -x_17 = !lean_is_exclusive(x_11); -if (x_17 == 0) +lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_10 = l_Lean_Meta_Grind_updateLastTag___closed__1; +x_11 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_10, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +x_13 = lean_unbox(x_12); +lean_dec(x_12); +if (x_13 == 0) { -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; -x_18 = lean_ctor_get(x_11, 2); -lean_dec(x_18); -x_19 = lean_ctor_get(x_11, 1); -lean_dec(x_19); -x_20 = lean_ctor_get(x_11, 0); -lean_dec(x_20); -x_21 = lean_array_fget(x_13, x_14); -x_22 = lean_unsigned_to_nat(1u); -x_23 = lean_nat_add(x_14, x_22); -lean_dec(x_14); -lean_ctor_set(x_11, 1, x_23); -x_24 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_9, x_21); -lean_dec(x_21); -lean_dec(x_9); -if (x_24 == 0) +uint8_t x_14; +x_14 = !lean_is_exclusive(x_11); +if (x_14 == 0) { -lean_object* x_25; -lean_dec(x_3); -x_25 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1___closed__1; -lean_ctor_set(x_7, 0, x_25); -return x_7; +lean_object* x_15; lean_object* x_16; +x_15 = lean_ctor_get(x_11, 0); +lean_dec(x_15); +x_16 = lean_box(0); +lean_ctor_set(x_11, 0, x_16); +return x_11; } else { -size_t x_26; size_t x_27; -lean_inc(x_3); -lean_ctor_set(x_7, 0, x_3); -x_26 = 1; -x_27 = lean_usize_add(x_6, x_26); -x_6 = x_27; -goto _start; +lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_17 = lean_ctor_get(x_11, 1); +lean_inc(x_17); +lean_dec(x_11); +x_18 = lean_box(0); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_17); +return x_19; } } else { -lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; +lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_20 = lean_ctor_get(x_11, 1); +lean_inc(x_20); lean_dec(x_11); -x_29 = lean_array_fget(x_13, x_14); -x_30 = lean_unsigned_to_nat(1u); -x_31 = lean_nat_add(x_14, x_30); -lean_dec(x_14); -x_32 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_32, 0, x_13); -lean_ctor_set(x_32, 1, x_31); -lean_ctor_set(x_32, 2, x_15); -x_33 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_9, x_29); -lean_dec(x_29); -lean_dec(x_9); -if (x_33 == 0) +x_21 = lean_st_ref_get(x_1, x_20); +x_22 = !lean_is_exclusive(x_21); +if (x_22 == 0) { -lean_object* x_34; -lean_dec(x_3); -x_34 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1___closed__1; -lean_ctor_set(x_7, 1, x_32); -lean_ctor_set(x_7, 0, x_34); -return x_7; +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_23 = lean_ctor_get(x_21, 0); +x_24 = lean_ctor_get(x_21, 1); +x_25 = lean_ctor_get(x_23, 0); +lean_inc(x_25); +lean_dec(x_23); +x_26 = l_Lean_MVarId_getTag(x_25, x_5, x_6, x_7, x_8, x_24); +if (lean_obj_tag(x_26) == 0) +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; uint8_t x_30; +x_27 = lean_ctor_get(x_26, 0); +lean_inc(x_27); +x_28 = lean_ctor_get(x_26, 1); +lean_inc(x_28); +lean_dec(x_26); +x_29 = lean_st_ref_get(x_4, x_28); +x_30 = !lean_is_exclusive(x_29); +if (x_30 == 0) +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; uint8_t x_34; +x_31 = lean_ctor_get(x_29, 0); +x_32 = lean_ctor_get(x_29, 1); +x_33 = lean_ctor_get(x_31, 7); +lean_inc(x_33); +lean_dec(x_31); +x_34 = lean_name_eq(x_27, x_33); +lean_dec(x_33); +if (x_34 == 0) +{ +lean_object* x_35; lean_object* x_36; uint8_t x_37; +lean_free_object(x_29); +x_35 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_10, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_32); +x_36 = lean_ctor_get(x_35, 0); +lean_inc(x_36); +x_37 = lean_unbox(x_36); +lean_dec(x_36); +if (x_37 == 0) +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; +lean_free_object(x_21); +x_38 = lean_ctor_get(x_35, 1); +lean_inc(x_38); +lean_dec(x_35); +x_39 = lean_box(0); +x_40 = l_Lean_Meta_Grind_updateLastTag___lambda__1(x_27, x_39, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_38); +return x_40; } else { -size_t x_35; size_t x_36; -lean_inc(x_3); -lean_ctor_set(x_7, 1, x_32); -lean_ctor_set(x_7, 0, x_3); -x_35 = 1; -x_36 = lean_usize_add(x_6, x_35); -x_6 = x_36; -goto _start; +uint8_t x_41; +x_41 = !lean_is_exclusive(x_35); +if (x_41 == 0) +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_42 = lean_ctor_get(x_35, 1); +x_43 = lean_ctor_get(x_35, 0); +lean_dec(x_43); +lean_inc(x_27); +x_44 = l_Lean_MessageData_ofName(x_27); +x_45 = l_Lean_Meta_Grind_updateLastTag___closed__3; +lean_ctor_set_tag(x_35, 7); +lean_ctor_set(x_35, 1, x_44); +lean_ctor_set(x_35, 0, x_45); +x_46 = l_Lean_Meta_Grind_updateLastTag___closed__5; +lean_ctor_set_tag(x_21, 7); +lean_ctor_set(x_21, 1, x_46); +lean_ctor_set(x_21, 0, x_35); +x_47 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_10, x_21, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, 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 = l_Lean_Meta_Grind_updateLastTag___lambda__1(x_27, x_48, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_49); +lean_dec(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; lean_object* x_57; lean_object* x_58; lean_object* x_59; +x_51 = lean_ctor_get(x_35, 1); +lean_inc(x_51); +lean_dec(x_35); +lean_inc(x_27); +x_52 = l_Lean_MessageData_ofName(x_27); +x_53 = l_Lean_Meta_Grind_updateLastTag___closed__3; +x_54 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_54, 0, x_53); +lean_ctor_set(x_54, 1, x_52); +x_55 = l_Lean_Meta_Grind_updateLastTag___closed__5; +lean_ctor_set_tag(x_21, 7); +lean_ctor_set(x_21, 1, x_55); +lean_ctor_set(x_21, 0, x_54); +x_56 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_10, x_21, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_51); +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_Grind_updateLastTag___lambda__1(x_27, x_57, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_58); +lean_dec(x_57); +return x_59; } } } else { -lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; uint8_t x_42; -x_38 = lean_ctor_get(x_7, 1); -lean_inc(x_38); -lean_dec(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); -x_41 = lean_ctor_get(x_38, 2); -lean_inc(x_41); -x_42 = lean_nat_dec_lt(x_40, x_41); -if (x_42 == 0) +lean_object* x_60; +lean_dec(x_27); +lean_free_object(x_21); +x_60 = lean_box(0); +lean_ctor_set(x_29, 0, x_60); +return x_29; +} +} +else { -lean_object* x_43; -lean_dec(x_41); -lean_dec(x_40); -lean_dec(x_39); -lean_dec(x_9); -x_43 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_43, 0, x_3); -lean_ctor_set(x_43, 1, x_38); -return x_43; +lean_object* x_61; lean_object* x_62; lean_object* x_63; uint8_t x_64; +x_61 = lean_ctor_get(x_29, 0); +x_62 = lean_ctor_get(x_29, 1); +lean_inc(x_62); +lean_inc(x_61); +lean_dec(x_29); +x_63 = lean_ctor_get(x_61, 7); +lean_inc(x_63); +lean_dec(x_61); +x_64 = lean_name_eq(x_27, x_63); +lean_dec(x_63); +if (x_64 == 0) +{ +lean_object* x_65; lean_object* x_66; uint8_t x_67; +x_65 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_10, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_62); +x_66 = lean_ctor_get(x_65, 0); +lean_inc(x_66); +x_67 = lean_unbox(x_66); +lean_dec(x_66); +if (x_67 == 0) +{ +lean_object* x_68; lean_object* x_69; lean_object* x_70; +lean_free_object(x_21); +x_68 = lean_ctor_get(x_65, 1); +lean_inc(x_68); +lean_dec(x_65); +x_69 = lean_box(0); +x_70 = l_Lean_Meta_Grind_updateLastTag___lambda__1(x_27, x_69, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_68); +return x_70; } else { -lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; uint8_t x_49; -if (lean_is_exclusive(x_38)) { - lean_ctor_release(x_38, 0); - lean_ctor_release(x_38, 1); - lean_ctor_release(x_38, 2); - x_44 = x_38; +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_71 = lean_ctor_get(x_65, 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; } else { - lean_dec_ref(x_38); - x_44 = lean_box(0); + lean_dec_ref(x_65); + x_72 = lean_box(0); } -x_45 = lean_array_fget(x_39, x_40); -x_46 = lean_unsigned_to_nat(1u); -x_47 = lean_nat_add(x_40, x_46); -lean_dec(x_40); -if (lean_is_scalar(x_44)) { - x_48 = lean_alloc_ctor(0, 3, 0); +lean_inc(x_27); +x_73 = l_Lean_MessageData_ofName(x_27); +x_74 = l_Lean_Meta_Grind_updateLastTag___closed__3; +if (lean_is_scalar(x_72)) { + x_75 = lean_alloc_ctor(7, 2, 0); } else { - x_48 = x_44; + x_75 = x_72; + lean_ctor_set_tag(x_75, 7); +} +lean_ctor_set(x_75, 0, x_74); +lean_ctor_set(x_75, 1, x_73); +x_76 = l_Lean_Meta_Grind_updateLastTag___closed__5; +lean_ctor_set_tag(x_21, 7); +lean_ctor_set(x_21, 1, x_76); +lean_ctor_set(x_21, 0, x_75); +x_77 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_10, x_21, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_71); +x_78 = lean_ctor_get(x_77, 0); +lean_inc(x_78); +x_79 = lean_ctor_get(x_77, 1); +lean_inc(x_79); +lean_dec(x_77); +x_80 = l_Lean_Meta_Grind_updateLastTag___lambda__1(x_27, x_78, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_79); +lean_dec(x_78); +return x_80; } -lean_ctor_set(x_48, 0, x_39); -lean_ctor_set(x_48, 1, x_47); -lean_ctor_set(x_48, 2, x_41); -x_49 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_9, x_45); -lean_dec(x_45); -lean_dec(x_9); -if (x_49 == 0) -{ -lean_object* x_50; lean_object* x_51; -lean_dec(x_3); -x_50 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1___closed__1; -x_51 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_51, 0, x_50); -lean_ctor_set(x_51, 1, x_48); -return x_51; } else { -lean_object* x_52; size_t x_53; size_t x_54; -lean_inc(x_3); -x_52 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_52, 0, x_3); -lean_ctor_set(x_52, 1, x_48); -x_53 = 1; -x_54 = lean_usize_add(x_6, x_53); -x_6 = x_54; -x_7 = x_52; -goto _start; -} -} -} +lean_object* x_81; lean_object* x_82; +lean_dec(x_27); +lean_free_object(x_21); +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_62); +return x_82; } } } -LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_markTheoremInstance___spec__7___lambda__1(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; lean_object* x_10; lean_object* x_11; size_t x_12; size_t x_13; lean_object* x_14; lean_object* x_15; -x_4 = lean_ctor_get(x_1, 1); -lean_inc(x_4); -lean_dec(x_1); -x_5 = lean_array_get_size(x_4); -x_6 = lean_unsigned_to_nat(0u); -x_7 = l_Array_toSubarray___rarg(x_4, x_6, x_5); -x_8 = lean_ctor_get(x_2, 1); -x_9 = lean_box(0); -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_7); -x_12 = lean_array_size(x_8); -x_13 = 0; -x_14 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__8(x_8, x_9, x_10, x_8, x_12, x_13, x_11); -x_15 = lean_ctor_get(x_14, 0); -lean_inc(x_15); -lean_dec(x_14); -if (lean_obj_tag(x_15) == 0) +uint8_t x_83; +lean_free_object(x_21); +x_83 = !lean_is_exclusive(x_26); +if (x_83 == 0) { -uint8_t x_16; -x_16 = 1; -return x_16; +return x_26; } else { -lean_object* x_17; uint8_t x_18; -x_17 = lean_ctor_get(x_15, 0); -lean_inc(x_17); -lean_dec(x_15); -x_18 = lean_unbox(x_17); -lean_dec(x_17); -return x_18; +lean_object* x_84; lean_object* x_85; lean_object* x_86; +x_84 = lean_ctor_get(x_26, 0); +x_85 = lean_ctor_get(x_26, 1); +lean_inc(x_85); +lean_inc(x_84); +lean_dec(x_26); +x_86 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_86, 0, x_84); +lean_ctor_set(x_86, 1, x_85); +return x_86; } } } -LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_markTheoremInstance___spec__7___lambda__2(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; uint8_t x_8; -x_4 = lean_ctor_get(x_2, 1); -x_5 = lean_array_get_size(x_4); -x_6 = lean_ctor_get(x_1, 1); -lean_inc(x_6); -x_7 = lean_array_get_size(x_6); -lean_dec(x_6); -x_8 = lean_nat_dec_eq(x_5, x_7); -lean_dec(x_7); -lean_dec(x_5); -if (x_8 == 0) +lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; +x_87 = lean_ctor_get(x_21, 0); +x_88 = lean_ctor_get(x_21, 1); +lean_inc(x_88); +lean_inc(x_87); +lean_dec(x_21); +x_89 = lean_ctor_get(x_87, 0); +lean_inc(x_89); +lean_dec(x_87); +x_90 = l_Lean_MVarId_getTag(x_89, x_5, x_6, x_7, x_8, x_88); +if (lean_obj_tag(x_90) == 0) { -uint8_t x_9; -lean_dec(x_1); -x_9 = 0; -return x_9; +lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; uint8_t x_98; +x_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_4, x_92); +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_93); + x_96 = lean_box(0); +} +x_97 = lean_ctor_get(x_94, 7); +lean_inc(x_97); +lean_dec(x_94); +x_98 = lean_name_eq(x_91, x_97); +lean_dec(x_97); +if (x_98 == 0) +{ +lean_object* x_99; lean_object* x_100; uint8_t x_101; +lean_dec(x_96); +x_99 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_10, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_95); +x_100 = lean_ctor_get(x_99, 0); +lean_inc(x_100); +x_101 = lean_unbox(x_100); +lean_dec(x_100); +if (x_101 == 0) +{ +lean_object* x_102; lean_object* x_103; lean_object* x_104; +x_102 = lean_ctor_get(x_99, 1); +lean_inc(x_102); +lean_dec(x_99); +x_103 = lean_box(0); +x_104 = l_Lean_Meta_Grind_updateLastTag___lambda__1(x_91, x_103, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_102); +return x_104; } else { -lean_object* x_10; uint8_t x_11; -x_10 = lean_box(0); -x_11 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_markTheoremInstance___spec__7___lambda__1(x_1, x_2, x_10); -return x_11; +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_105 = lean_ctor_get(x_99, 1); +lean_inc(x_105); +if (lean_is_exclusive(x_99)) { + lean_ctor_release(x_99, 0); + lean_ctor_release(x_99, 1); + x_106 = x_99; +} else { + lean_dec_ref(x_99); + x_106 = lean_box(0); } +lean_inc(x_91); +x_107 = l_Lean_MessageData_ofName(x_91); +x_108 = l_Lean_Meta_Grind_updateLastTag___closed__3; +if (lean_is_scalar(x_106)) { + x_109 = lean_alloc_ctor(7, 2, 0); +} else { + x_109 = x_106; + lean_ctor_set_tag(x_109, 7); } +lean_ctor_set(x_109, 0, x_108); +lean_ctor_set(x_109, 1, x_107); +x_110 = l_Lean_Meta_Grind_updateLastTag___closed__5; +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 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_10, x_111, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_105); +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 = l_Lean_Meta_Grind_updateLastTag___lambda__1(x_91, x_113, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_114); +lean_dec(x_113); +return x_115; } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_markTheoremInstance___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; +lean_object* x_116; lean_object* x_117; +lean_dec(x_91); +x_116 = lean_box(0); +if (lean_is_scalar(x_96)) { + x_117 = lean_alloc_ctor(0, 2, 0); +} else { + x_117 = x_96; +} +lean_ctor_set(x_117, 0, x_116); +lean_ctor_set(x_117, 1, x_95); +return x_117; } } else { -lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; -x_17 = lean_array_fget(x_5, x_2); -x_18 = lean_ctor_get(x_3, 0); -lean_inc(x_18); -x_19 = lean_ctor_get(x_17, 0); -lean_inc(x_19); -x_20 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_18, x_19); -lean_dec(x_19); -lean_dec(x_18); -if (x_20 == 0) +lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; +x_118 = lean_ctor_get(x_90, 0); +lean_inc(x_118); +x_119 = lean_ctor_get(x_90, 1); +lean_inc(x_119); +if (lean_is_exclusive(x_90)) { + lean_ctor_release(x_90, 0); + lean_ctor_release(x_90, 1); + x_120 = x_90; +} else { + lean_dec_ref(x_90); + x_120 = lean_box(0); +} +if (lean_is_scalar(x_120)) { + x_121 = lean_alloc_ctor(1, 2, 0); +} else { + x_121 = x_120; +} +lean_ctor_set(x_121, 0, x_118); +lean_ctor_set(x_121, 1, x_119); +return x_121; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___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_21; lean_object* x_22; -lean_dec(x_17); +lean_object* x_11; +x_11 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___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_9); +lean_dec(x_8); +lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -x_21 = lean_unsigned_to_nat(1u); -x_22 = lean_nat_add(x_2, x_21); +lean_dec(x_4); +lean_dec(x_3); lean_dec(x_2); -x_2 = x_22; -goto _start; +return x_11; } -else -{ -lean_object* x_24; uint8_t x_25; -x_24 = lean_box(0); -x_25 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_markTheoremInstance___spec__7___lambda__2(x_17, x_3, x_24); -if (x_25 == 0) +} +LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: { -lean_object* x_26; lean_object* x_27; +lean_object* x_12; +x_12 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -x_26 = lean_unsigned_to_nat(1u); -x_27 = lean_nat_add(x_2, x_26); -lean_dec(x_2); -x_2 = x_27; -goto _start; +lean_dec(x_4); +lean_dec(x_3); +return x_12; } -else -{ -uint8_t x_29; -x_29 = !lean_is_exclusive(x_1); -if (x_29 == 0) +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_updateLastTag___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_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_30 = lean_ctor_get(x_1, 1); -lean_dec(x_30); -x_31 = lean_ctor_get(x_1, 0); -lean_dec(x_31); -x_32 = lean_array_fset(x_5, x_2, x_3); -x_33 = lean_array_fset(x_6, x_2, x_4); +lean_object* x_12; +x_12 = l_Lean_Meta_Grind_updateLastTag___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); lean_dec(x_2); -lean_ctor_set(x_1, 1, x_33); -lean_ctor_set(x_1, 0, x_32); -return x_1; +return x_12; } -else +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_updateLastTag___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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_34; lean_object* x_35; lean_object* x_36; -lean_dec(x_1); -x_34 = lean_array_fset(x_5, x_2, x_3); -x_35 = lean_array_fset(x_6, x_2, x_4); +lean_object* x_10; +x_10 = l_Lean_Meta_Grind_updateLastTag(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); -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_1); +return x_10; } } +static lean_object* _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("doElemTrace_goal[_]__", 21, 21); +return x_1; } } +static lean_object* _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78____closed__6; +x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78____closed__7; +x_3 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78____closed__8; +x_4 = l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__1; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; } } -LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +static lean_object* _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______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; lean_object* x_11; size_t x_12; size_t x_13; lean_object* x_14; lean_object* x_15; -x_4 = lean_ctor_get(x_1, 1); -lean_inc(x_4); -lean_dec(x_1); -x_5 = lean_array_get_size(x_4); -x_6 = lean_unsigned_to_nat(0u); -x_7 = l_Array_toSubarray___rarg(x_4, x_6, x_5); -x_8 = lean_ctor_get(x_2, 1); -x_9 = lean_box(0); -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_7); -x_12 = lean_array_size(x_8); -x_13 = 0; -x_14 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__6(x_8, x_9, x_10, x_8, x_12, x_13, x_11); -x_15 = lean_ctor_get(x_14, 0); -lean_inc(x_15); -lean_dec(x_14); -if (lean_obj_tag(x_15) == 0) -{ -uint8_t x_16; -x_16 = 1; -return x_16; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("andthen", 7, 7); +return x_1; } -else -{ -lean_object* x_17; uint8_t x_18; -x_17 = lean_ctor_get(x_15, 0); -lean_inc(x_17); -lean_dec(x_15); -x_18 = lean_unbox(x_17); -lean_dec(x_17); -return x_18; } +static lean_object* _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______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_Grind_doElemTrace__goal_x5b___x5d_______closed__3; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } } -LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +static lean_object* _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__5() { _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, 1); -x_5 = lean_array_get_size(x_4); -x_6 = lean_ctor_get(x_1, 1); -lean_inc(x_6); -x_7 = lean_array_get_size(x_6); -lean_dec(x_6); -x_8 = lean_nat_dec_eq(x_5, x_7); -lean_dec(x_7); -lean_dec(x_5); -if (x_8 == 0) -{ -uint8_t x_9; -lean_dec(x_1); -x_9 = 0; -return x_9; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("trace_goal[", 11, 11); +return x_1; } -else -{ -lean_object* x_10; uint8_t x_11; -x_10 = lean_box(0); -x_11 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3___lambda__1(x_1, x_2, x_10); -return x_11; } +static lean_object* _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__5; +x_2 = lean_alloc_ctor(5, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5) { +static lean_object* _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__7() { _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_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___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); +lean_object* x_1; +x_1 = lean_mk_string_unchecked("ident", 5, 5); return x_1; } -else +} +static lean_object* _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__8() { +_start: { -lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = lean_array_fget(x_7, x_12); -x_16 = lean_box(0); -x_17 = lean_array_fset(x_7, x_12, x_16); -switch (lean_obj_tag(x_15)) { -case 0: +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__7; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__9() { +_start: { -uint8_t x_18; -x_18 = !lean_is_exclusive(x_15); -if (x_18 == 0) +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__8; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__10() { +_start: { -lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; -x_19 = lean_ctor_get(x_15, 0); -x_20 = lean_ctor_get(x_15, 1); -x_21 = lean_ctor_get(x_4, 0); -lean_inc(x_21); -x_22 = lean_ctor_get(x_19, 0); -lean_inc(x_22); -x_23 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_21, x_22); -lean_dec(x_22); -lean_dec(x_21); -if (x_23 == 0) +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__4; +x_2 = l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__6; +x_3 = l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__9; +x_4 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__11() { +_start: { -lean_object* x_24; lean_object* x_25; lean_object* x_26; -lean_free_object(x_15); -x_24 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_19, x_20, x_4, x_5); -x_25 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_25, 0, x_24); -x_26 = lean_array_fset(x_17, x_12, x_25); -lean_dec(x_12); -lean_ctor_set(x_1, 0, x_26); +lean_object* x_1; +x_1 = lean_mk_string_unchecked("]", 1, 1); return x_1; } -else +} +static lean_object* _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__12() { +_start: { -uint8_t x_27; -lean_inc(x_19); -x_27 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3___lambda__2(x_19, x_4, x_16); -if (x_27 == 0) +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__11; +x_2 = lean_alloc_ctor(5, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__13() { +_start: { -lean_object* x_28; lean_object* x_29; lean_object* x_30; -lean_free_object(x_15); -x_28 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_19, x_20, x_4, x_5); -x_29 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_29, 0, x_28); -x_30 = lean_array_fset(x_17, x_12, x_29); -lean_dec(x_12); -lean_ctor_set(x_1, 0, x_30); -return x_1; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__4; +x_2 = l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__10; +x_3 = l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__12; +x_4 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; } -else +} +static lean_object* _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__14() { +_start: { -lean_object* x_31; -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_31 = lean_array_fset(x_17, x_12, x_15); -lean_dec(x_12); -lean_ctor_set(x_1, 0, x_31); +lean_object* x_1; +x_1 = lean_mk_string_unchecked("orelse", 6, 6); return x_1; } } -} -else +static lean_object* _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__15() { +_start: { -lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; -x_32 = lean_ctor_get(x_15, 0); -x_33 = lean_ctor_get(x_15, 1); -lean_inc(x_33); -lean_inc(x_32); -lean_dec(x_15); -x_34 = lean_ctor_get(x_4, 0); -lean_inc(x_34); -x_35 = lean_ctor_get(x_32, 0); -lean_inc(x_35); -x_36 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_34, x_35); -lean_dec(x_35); -lean_dec(x_34); -if (x_36 == 0) +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__14; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__16() { +_start: { -lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_37 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_32, x_33, x_4, x_5); -x_38 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_38, 0, x_37); -x_39 = lean_array_fset(x_17, x_12, x_38); -lean_dec(x_12); -lean_ctor_set(x_1, 0, x_39); +lean_object* x_1; +x_1 = lean_mk_string_unchecked("interpolatedStr", 15, 15); return x_1; } -else -{ -uint8_t x_40; -lean_inc(x_32); -x_40 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3___lambda__2(x_32, x_4, x_16); -if (x_40 == 0) +} +static lean_object* _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__17() { +_start: { -lean_object* x_41; lean_object* x_42; lean_object* x_43; -x_41 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_32, x_33, x_4, x_5); -x_42 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_42, 0, x_41); -x_43 = lean_array_fset(x_17, x_12, x_42); -lean_dec(x_12); -lean_ctor_set(x_1, 0, x_43); -return x_1; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__16; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } -else +} +static lean_object* _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__18() { +_start: { -lean_object* x_44; lean_object* x_45; -lean_dec(x_33); -lean_dec(x_32); -x_44 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_44, 0, x_4); -lean_ctor_set(x_44, 1, x_5); -x_45 = lean_array_fset(x_17, x_12, x_44); -lean_dec(x_12); -lean_ctor_set(x_1, 0, x_45); +lean_object* x_1; +x_1 = lean_mk_string_unchecked("term", 4, 4); return x_1; } } +static lean_object* _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__19() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__18; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } } -case 1: +static lean_object* _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__20() { +_start: { -uint8_t x_46; -x_46 = !lean_is_exclusive(x_15); -if (x_46 == 0) +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__19; +x_2 = lean_unsigned_to_nat(0u); +x_3 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__21() { +_start: { -lean_object* x_47; size_t x_48; size_t x_49; lean_object* x_50; lean_object* x_51; -x_47 = lean_ctor_get(x_15, 0); -x_48 = lean_usize_shift_right(x_2, x_9); -x_49 = lean_usize_add(x_3, x_8); -x_50 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3(x_47, x_48, x_49, x_4, x_5); -lean_ctor_set(x_15, 0, x_50); -x_51 = lean_array_fset(x_17, x_12, x_15); -lean_dec(x_12); -lean_ctor_set(x_1, 0, x_51); -return x_1; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__17; +x_2 = l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__20; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; } -else +} +static lean_object* _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__22() { +_start: { -lean_object* x_52; size_t x_53; size_t x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; -x_52 = lean_ctor_get(x_15, 0); -lean_inc(x_52); -lean_dec(x_15); -x_53 = lean_usize_shift_right(x_2, x_9); -x_54 = lean_usize_add(x_3, x_8); -x_55 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3(x_52, x_53, x_54, x_4, x_5); -x_56 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_56, 0, x_55); -x_57 = lean_array_fset(x_17, x_12, x_56); -lean_dec(x_12); -lean_ctor_set(x_1, 0, x_57); -return x_1; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__15; +x_2 = l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__21; +x_3 = l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__20; +x_4 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; } } -default: +static lean_object* _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__23() { +_start: { -lean_object* x_58; lean_object* x_59; -x_58 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_58, 0, x_4); -lean_ctor_set(x_58, 1, x_5); -x_59 = lean_array_fset(x_17, x_12, x_58); -lean_dec(x_12); -lean_ctor_set(x_1, 0, x_59); -return x_1; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__4; +x_2 = l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__13; +x_3 = l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__22; +x_4 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; } } +static lean_object* _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__24() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__2; +x_2 = lean_unsigned_to_nat(1022u); +x_3 = l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__23; +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; } } -else +static lean_object* _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d____() { +_start: { -lean_object* x_60; size_t x_61; size_t x_62; size_t x_63; size_t x_64; lean_object* x_65; lean_object* x_66; uint8_t x_67; -x_60 = lean_ctor_get(x_1, 0); -lean_inc(x_60); -lean_dec(x_1); -x_61 = 1; -x_62 = 5; -x_63 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__2___closed__2; -x_64 = lean_usize_land(x_2, x_63); -x_65 = lean_usize_to_nat(x_64); -x_66 = lean_array_get_size(x_60); -x_67 = lean_nat_dec_lt(x_65, x_66); -lean_dec(x_66); -if (x_67 == 0) +lean_object* x_1; +x_1 = l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__24; +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__1() { +_start: { -lean_object* x_68; -lean_dec(x_65); -lean_dec(x_5); -lean_dec(x_4); -x_68 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_68, 0, x_60); -return x_68; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Parser", 6, 6); +return x_1; } -else +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__2() { +_start: { -lean_object* x_69; lean_object* x_70; lean_object* x_71; -x_69 = lean_array_fget(x_60, x_65); -x_70 = lean_box(0); -x_71 = lean_array_fset(x_60, x_65, x_70); -switch (lean_obj_tag(x_69)) { -case 0: +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Term", 4, 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__3() { +_start: { -lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; uint8_t x_77; -x_72 = lean_ctor_get(x_69, 0); -lean_inc(x_72); -x_73 = lean_ctor_get(x_69, 1); -lean_inc(x_73); -if (lean_is_exclusive(x_69)) { - lean_ctor_release(x_69, 0); - lean_ctor_release(x_69, 1); - x_74 = x_69; -} else { - lean_dec_ref(x_69); - x_74 = lean_box(0); +lean_object* x_1; +x_1 = lean_mk_string_unchecked("doNested", 8, 8); +return x_1; } -x_75 = lean_ctor_get(x_4, 0); -lean_inc(x_75); -x_76 = lean_ctor_get(x_72, 0); -lean_inc(x_76); -x_77 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_75, x_76); -lean_dec(x_76); -lean_dec(x_75); -if (x_77 == 0) +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__4() { +_start: { -lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; -lean_dec(x_74); -x_78 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_72, x_73, x_4, x_5); -x_79 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_79, 0, x_78); -x_80 = lean_array_fset(x_71, x_65, x_79); -lean_dec(x_65); -x_81 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_81, 0, x_80); -return x_81; +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_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78____closed__6; +x_2 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__1; +x_3 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__2; +x_4 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__3; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; } -else +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__5() { +_start: { -uint8_t x_82; -lean_inc(x_72); -x_82 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3___lambda__2(x_72, x_4, x_70); -if (x_82 == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("do", 2, 2); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__6() { +_start: { -lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; -lean_dec(x_74); -x_83 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_72, x_73, x_4, x_5); -x_84 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_84, 0, x_83); -x_85 = lean_array_fset(x_71, x_65, x_84); -lean_dec(x_65); -x_86 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_86, 0, x_85); -return x_86; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("doSeqIndent", 11, 11); +return x_1; } -else +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__7() { +_start: { -lean_object* x_87; lean_object* x_88; lean_object* x_89; -lean_dec(x_73); -lean_dec(x_72); -if (lean_is_scalar(x_74)) { - x_87 = lean_alloc_ctor(0, 2, 0); -} else { - x_87 = x_74; +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_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78____closed__6; +x_2 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__1; +x_3 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__2; +x_4 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__6; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; } -lean_ctor_set(x_87, 0, x_4); -lean_ctor_set(x_87, 1, x_5); -x_88 = lean_array_fset(x_71, x_65, x_87); -lean_dec(x_65); -x_89 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_89, 0, x_88); -return x_89; } +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__8() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("null", 4, 4); +return x_1; } } -case 1: +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__9() { +_start: { -lean_object* x_90; lean_object* x_91; size_t x_92; size_t x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; -x_90 = lean_ctor_get(x_69, 0); -lean_inc(x_90); -if (lean_is_exclusive(x_69)) { - lean_ctor_release(x_69, 0); - x_91 = x_69; -} else { - lean_dec_ref(x_69); - x_91 = lean_box(0); +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__8; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } -x_92 = lean_usize_shift_right(x_2, x_62); -x_93 = lean_usize_add(x_3, x_61); -x_94 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3(x_90, x_92, x_93, x_4, x_5); -if (lean_is_scalar(x_91)) { - x_95 = lean_alloc_ctor(1, 1, 0); -} else { - x_95 = x_91; } -lean_ctor_set(x_95, 0, x_94); -x_96 = lean_array_fset(x_71, x_65, x_95); -lean_dec(x_65); -x_97 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_97, 0, x_96); -return x_97; +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__10() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("doSeqItem", 9, 9); +return x_1; } -default: +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__11() { +_start: { -lean_object* x_98; lean_object* x_99; lean_object* x_100; -x_98 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_98, 0, x_4); -lean_ctor_set(x_98, 1, x_5); -x_99 = lean_array_fset(x_71, x_65, x_98); -lean_dec(x_65); -x_100 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_100, 0, x_99); -return x_100; +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_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78____closed__6; +x_2 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__1; +x_3 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__2; +x_4 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___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_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__12() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("doLet", 5, 5); +return x_1; +} } +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78____closed__6; +x_2 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__1; +x_3 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__2; +x_4 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__12; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; } } -else +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__14() { +_start: { -uint8_t x_101; -x_101 = !lean_is_exclusive(x_1); -if (x_101 == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("let", 3, 3); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__15() { +_start: { -lean_object* x_102; lean_object* x_103; size_t x_104; uint8_t x_105; -x_102 = lean_unsigned_to_nat(0u); -x_103 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_markTheoremInstance___spec__7(x_1, x_102, x_4, x_5); -x_104 = 7; -x_105 = lean_usize_dec_le(x_104, x_3); -if (x_105 == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("letDecl", 7, 7); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__16() { +_start: { -lean_object* x_106; lean_object* x_107; uint8_t x_108; -x_106 = l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(x_103); -x_107 = lean_unsigned_to_nat(4u); -x_108 = lean_nat_dec_lt(x_106, x_107); -lean_dec(x_106); -if (x_108 == 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_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78____closed__6; +x_2 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__1; +x_3 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__2; +x_4 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__15; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__17() { +_start: { -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_inc(x_109); -x_110 = lean_ctor_get(x_103, 1); -lean_inc(x_110); -lean_dec(x_103); -x_111 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__2___closed__3; -x_112 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_markTheoremInstance___spec__4(x_3, x_109, x_110, lean_box(0), x_102, x_111); -lean_dec(x_110); -lean_dec(x_109); -return x_112; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("letIdDecl", 9, 9); +return x_1; } -else +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__18() { +_start: { -return x_103; +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_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78____closed__6; +x_2 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__1; +x_3 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__2; +x_4 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__17; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; } } -else +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__19() { +_start: { -return x_103; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("cls", 3, 3); +return x_1; } } -else +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__20() { +_start: { -lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; size_t x_118; uint8_t x_119; -x_113 = lean_ctor_get(x_1, 0); -x_114 = lean_ctor_get(x_1, 1); -lean_inc(x_114); -lean_inc(x_113); -lean_dec(x_1); -x_115 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_115, 0, x_113); -lean_ctor_set(x_115, 1, x_114); -x_116 = lean_unsigned_to_nat(0u); -x_117 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_markTheoremInstance___spec__7(x_115, x_116, x_4, x_5); -x_118 = 7; -x_119 = lean_usize_dec_le(x_118, x_3); -if (x_119 == 0) +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__19; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__21() { +_start: { -lean_object* x_120; lean_object* x_121; uint8_t x_122; -x_120 = l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(x_117); -x_121 = lean_unsigned_to_nat(4u); -x_122 = lean_nat_dec_lt(x_120, x_121); -lean_dec(x_120); -if (x_122 == 0) +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__19; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__22() { +_start: { -lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; -x_123 = lean_ctor_get(x_117, 0); -lean_inc(x_123); -x_124 = lean_ctor_get(x_117, 1); -lean_inc(x_124); -lean_dec(x_117); -x_125 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__2___closed__3; -x_126 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_markTheoremInstance___spec__4(x_3, x_123, x_124, lean_box(0), x_116, x_125); -lean_dec(x_124); -lean_dec(x_123); -return x_126; +lean_object* x_1; +x_1 = lean_mk_string_unchecked(":=", 2, 2); +return x_1; } -else +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__23() { +_start: { -return x_117; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("doIf", 4, 4); +return x_1; } } -else +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__24() { +_start: { -return x_117; +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_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78____closed__6; +x_2 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__1; +x_3 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__2; +x_4 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__23; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; } } +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__25() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("if", 2, 2); +return x_1; } } +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__26() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("doIfProp", 8, 8); +return x_1; } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_markTheoremInstance___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__27() { _start: { -uint64_t x_4; lean_object* x_5; lean_object* x_6; size_t x_7; size_t x_8; uint64_t x_9; size_t x_10; size_t x_11; lean_object* x_12; -x_4 = l_Lean_Meta_Grind_instHashablePreInstance_unsafe__1(x_2); -x_5 = lean_ctor_get(x_2, 1); -lean_inc(x_5); -x_6 = lean_box(0); -x_7 = lean_array_size(x_5); -x_8 = 0; -x_9 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__2(x_5, x_6, x_5, x_7, x_8, x_4); -lean_dec(x_5); -x_10 = lean_uint64_to_usize(x_9); -x_11 = 1; -x_12 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3(x_1, x_10, x_11, x_2, x_3); -return x_12; +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_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78____closed__6; +x_2 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__1; +x_3 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__2; +x_4 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__26; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; } } -LEAN_EXPORT uint64_t l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__10(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, uint64_t x_6) { +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__28() { _start: { -uint8_t x_7; -x_7 = lean_usize_dec_lt(x_5, x_4); -if (x_7 == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("paren", 5, 5); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__29() { +_start: { -return x_6; +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_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78____closed__6; +x_2 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__1; +x_3 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__2; +x_4 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__28; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; } -else +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__30() { +_start: { -lean_object* x_8; uint64_t x_9; uint64_t x_10; size_t x_11; size_t x_12; -x_8 = lean_array_uget(x_3, x_5); -x_9 = l_Lean_Meta_Grind_instHashablePreInstance_unsafe__2(x_8); -lean_dec(x_8); -x_10 = lean_uint64_mix_hash(x_6, x_9); -x_11 = 1; -x_12 = lean_usize_add(x_5, x_11); -x_5 = x_12; -x_6 = x_10; -goto _start; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("(", 1, 1); +return x_1; } } +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__31() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("liftMethod", 10, 10); +return x_1; } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__12(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, size_t x_6, lean_object* x_7) { +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__32() { _start: { -uint8_t x_8; -x_8 = lean_usize_dec_lt(x_6, x_5); -if (x_8 == 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_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78____closed__6; +x_2 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__1; +x_3 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__2; +x_4 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__31; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__33() { +_start: { -lean_dec(x_3); -return x_7; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("←", 3, 1); +return x_1; } -else +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__34() { +_start: { -lean_object* x_9; uint8_t x_10; -x_9 = lean_array_uget(x_4, x_6); -x_10 = !lean_is_exclusive(x_7); -if (x_10 == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("app", 3, 3); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__35() { +_start: { -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; -x_11 = lean_ctor_get(x_7, 1); -x_12 = lean_ctor_get(x_7, 0); -lean_dec(x_12); -x_13 = lean_ctor_get(x_11, 0); -lean_inc(x_13); -x_14 = lean_ctor_get(x_11, 1); -lean_inc(x_14); -x_15 = lean_ctor_get(x_11, 2); -lean_inc(x_15); -x_16 = lean_nat_dec_lt(x_14, x_15); -if (x_16 == 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_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78____closed__6; +x_2 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__1; +x_3 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__2; +x_4 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__34; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__36() { +_start: { -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_9); -lean_ctor_set(x_7, 0, x_3); -return x_7; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.isTracingEnabledFor", 24, 24); +return x_1; } -else +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__37() { +_start: { -uint8_t x_17; -x_17 = !lean_is_exclusive(x_11); -if (x_17 == 0) +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__36; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__38() { +_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; uint8_t x_24; -x_18 = lean_ctor_get(x_11, 2); -lean_dec(x_18); -x_19 = lean_ctor_get(x_11, 1); -lean_dec(x_19); -x_20 = lean_ctor_get(x_11, 0); -lean_dec(x_20); -x_21 = lean_array_fget(x_13, x_14); -x_22 = lean_unsigned_to_nat(1u); -x_23 = lean_nat_add(x_14, x_22); -lean_dec(x_14); -lean_ctor_set(x_11, 1, x_23); -x_24 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_9, x_21); -lean_dec(x_21); -lean_dec(x_9); -if (x_24 == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("isTracingEnabledFor", 19, 19); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__39() { +_start: { -lean_object* x_25; -lean_dec(x_3); -x_25 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1___closed__1; -lean_ctor_set(x_7, 0, x_25); -return x_7; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78____closed__6; +x_2 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__38; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; } -else +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__40() { +_start: { -size_t x_26; size_t x_27; -lean_inc(x_3); -lean_ctor_set(x_7, 0, x_3); -x_26 = 1; -x_27 = lean_usize_add(x_6, x_26); -x_6 = x_27; -goto _start; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__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; } } -else +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__41() { +_start: { -lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; -lean_dec(x_11); -x_29 = lean_array_fget(x_13, x_14); -x_30 = lean_unsigned_to_nat(1u); -x_31 = lean_nat_add(x_14, x_30); -lean_dec(x_14); -x_32 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_32, 0, x_13); -lean_ctor_set(x_32, 1, x_31); -lean_ctor_set(x_32, 2, x_15); -x_33 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_9, x_29); -lean_dec(x_29); -lean_dec(x_9); -if (x_33 == 0) +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__40; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__42() { +_start: { -lean_object* x_34; -lean_dec(x_3); -x_34 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1___closed__1; -lean_ctor_set(x_7, 1, x_32); -lean_ctor_set(x_7, 0, x_34); -return x_7; +lean_object* x_1; +x_1 = lean_mk_string_unchecked(")", 1, 1); +return x_1; } -else +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__43() { +_start: { -size_t x_35; size_t x_36; -lean_inc(x_3); -lean_ctor_set(x_7, 1, x_32); -lean_ctor_set(x_7, 0, x_3); -x_35 = 1; -x_36 = lean_usize_add(x_6, x_35); -x_6 = x_36; -goto _start; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("then", 4, 4); +return x_1; } } +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__44() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("doExpr", 6, 6); +return x_1; } } -else +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__45() { +_start: { -lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; uint8_t x_42; -x_38 = lean_ctor_get(x_7, 1); -lean_inc(x_38); -lean_dec(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); -x_41 = lean_ctor_get(x_38, 2); -lean_inc(x_41); -x_42 = lean_nat_dec_lt(x_40, x_41); -if (x_42 == 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_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78____closed__6; +x_2 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__1; +x_3 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__2; +x_4 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__44; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__46() { +_start: { -lean_object* x_43; -lean_dec(x_41); -lean_dec(x_40); -lean_dec(x_39); -lean_dec(x_9); -x_43 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_43, 0, x_3); -lean_ctor_set(x_43, 1, x_38); -return x_43; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("updateLastTag", 13, 13); +return x_1; } -else +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__47() { +_start: { -lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; uint8_t x_49; -if (lean_is_exclusive(x_38)) { - lean_ctor_release(x_38, 0); - lean_ctor_release(x_38, 1); - lean_ctor_release(x_38, 2); - x_44 = x_38; -} else { - lean_dec_ref(x_38); - x_44 = lean_box(0); +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__46; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; } -x_45 = lean_array_fget(x_39, x_40); -x_46 = lean_unsigned_to_nat(1u); -x_47 = lean_nat_add(x_40, x_46); -lean_dec(x_40); -if (lean_is_scalar(x_44)) { - x_48 = lean_alloc_ctor(0, 3, 0); -} else { - x_48 = x_44; } -lean_ctor_set(x_48, 0, x_39); -lean_ctor_set(x_48, 1, x_47); -lean_ctor_set(x_48, 2, x_41); -x_49 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_9, x_45); -lean_dec(x_45); -lean_dec(x_9); -if (x_49 == 0) +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__48() { +_start: { -lean_object* x_50; lean_object* x_51; -lean_dec(x_3); -x_50 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1___closed__1; -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_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__46; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } -else +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__49() { +_start: { -lean_object* x_52; size_t x_53; size_t x_54; -lean_inc(x_3); -x_52 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_52, 0, x_3); -lean_ctor_set(x_52, 1, x_48); -x_53 = 1; -x_54 = lean_usize_add(x_6, x_53); -x_6 = x_54; -x_7 = x_52; -goto _start; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78____closed__6; +x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78____closed__7; +x_3 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78____closed__8; +x_4 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__46; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; } } +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__50() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__49; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; } } +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__51() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___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_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___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) { +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__52() { _start: { -uint8_t x_8; -x_8 = lean_usize_dec_lt(x_6, x_5); -if (x_8 == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.addTrace", 13, 13); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__53() { +_start: { -lean_dec(x_3); -return x_7; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__52; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; } -else +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__54() { +_start: { -lean_object* x_9; uint8_t x_10; -x_9 = lean_array_uget(x_4, x_6); -x_10 = !lean_is_exclusive(x_7); -if (x_10 == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("addTrace", 8, 8); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__55() { +_start: { -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; -x_11 = lean_ctor_get(x_7, 1); -x_12 = lean_ctor_get(x_7, 0); -lean_dec(x_12); -x_13 = lean_ctor_get(x_11, 0); -lean_inc(x_13); -x_14 = lean_ctor_get(x_11, 1); -lean_inc(x_14); -x_15 = lean_ctor_get(x_11, 2); -lean_inc(x_15); -x_16 = lean_nat_dec_lt(x_14, x_15); -if (x_16 == 0) -{ -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_9); -lean_ctor_set(x_7, 0, x_3); -return x_7; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78____closed__6; +x_2 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__54; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; } -else -{ -uint8_t x_17; -x_17 = !lean_is_exclusive(x_11); -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; uint8_t x_24; -x_18 = lean_ctor_get(x_11, 2); -lean_dec(x_18); -x_19 = lean_ctor_get(x_11, 1); -lean_dec(x_19); -x_20 = lean_ctor_get(x_11, 0); -lean_dec(x_20); -x_21 = lean_array_fget(x_13, x_14); -x_22 = lean_unsigned_to_nat(1u); -x_23 = lean_nat_add(x_14, x_22); -lean_dec(x_14); -lean_ctor_set(x_11, 1, x_23); -x_24 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_9, x_21); -lean_dec(x_21); -lean_dec(x_9); -if (x_24 == 0) -{ -lean_object* x_25; -lean_dec(x_3); -x_25 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1___closed__1; -lean_ctor_set(x_7, 0, x_25); -return x_7; } -else +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__56() { +_start: { -size_t x_26; size_t x_27; -lean_inc(x_3); -lean_ctor_set(x_7, 0, x_3); -x_26 = 1; -x_27 = lean_usize_add(x_6, x_26); -x_6 = x_27; -goto _start; -} +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__55; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; } -else -{ -lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; -lean_dec(x_11); -x_29 = lean_array_fget(x_13, x_14); -x_30 = lean_unsigned_to_nat(1u); -x_31 = lean_nat_add(x_14, x_30); -lean_dec(x_14); -x_32 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_32, 0, x_13); -lean_ctor_set(x_32, 1, x_31); -lean_ctor_set(x_32, 2, x_15); -x_33 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_9, x_29); -lean_dec(x_29); -lean_dec(x_9); -if (x_33 == 0) -{ -lean_object* x_34; -lean_dec(x_3); -x_34 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1___closed__1; -lean_ctor_set(x_7, 1, x_32); -lean_ctor_set(x_7, 0, x_34); -return x_7; } -else +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__57() { +_start: { -size_t x_35; size_t x_36; -lean_inc(x_3); -lean_ctor_set(x_7, 1, x_32); -lean_ctor_set(x_7, 0, x_3); -x_35 = 1; -x_36 = lean_usize_add(x_6, x_35); -x_6 = x_36; -goto _start; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__56; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; } } +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__58() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("quotedName", 10, 10); +return x_1; } } -else -{ -lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; uint8_t x_42; -x_38 = lean_ctor_get(x_7, 1); -lean_inc(x_38); -lean_dec(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); -x_41 = lean_ctor_get(x_38, 2); -lean_inc(x_41); -x_42 = lean_nat_dec_lt(x_40, x_41); -if (x_42 == 0) +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__59() { +_start: { -lean_object* x_43; -lean_dec(x_41); -lean_dec(x_40); -lean_dec(x_39); -lean_dec(x_9); -x_43 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_43, 0, x_3); -lean_ctor_set(x_43, 1, x_38); -return x_43; +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_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78____closed__6; +x_2 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__1; +x_3 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__2; +x_4 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__58; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; } -else +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__60() { +_start: { -lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; uint8_t x_49; -if (lean_is_exclusive(x_38)) { - lean_ctor_release(x_38, 0); - lean_ctor_release(x_38, 1); - lean_ctor_release(x_38, 2); - x_44 = x_38; -} else { - lean_dec_ref(x_38); - x_44 = lean_box(0); +lean_object* x_1; +x_1 = lean_mk_string_unchecked(".", 1, 1); +return x_1; } -x_45 = lean_array_fget(x_39, x_40); -x_46 = lean_unsigned_to_nat(1u); -x_47 = lean_nat_add(x_40, x_46); -lean_dec(x_40); -if (lean_is_scalar(x_44)) { - x_48 = lean_alloc_ctor(0, 3, 0); -} else { - x_48 = x_44; } -lean_ctor_set(x_48, 0, x_39); -lean_ctor_set(x_48, 1, x_47); -lean_ctor_set(x_48, 2, x_41); -x_49 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_9, x_45); -lean_dec(x_45); -lean_dec(x_9); -if (x_49 == 0) +LEAN_EXPORT lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: { -lean_object* x_50; lean_object* x_51; +lean_object* x_5; uint8_t x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; 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_5 = lean_ctor_get(x_3, 5); +lean_inc(x_5); +x_6 = 0; +x_7 = l_Lean_SourceInfo_fromRef(x_5, x_6); +lean_dec(x_5); +x_8 = lean_ctor_get(x_3, 2); +lean_inc(x_8); +x_9 = lean_ctor_get(x_3, 1); +lean_inc(x_9); lean_dec(x_3); -x_50 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1___closed__1; -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; +x_10 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__5; +lean_inc(x_7); +x_11 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_11, 0, x_7); +lean_ctor_set(x_11, 1, x_10); +x_12 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__14; +lean_inc(x_7); +x_13 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_13, 0, x_7); +lean_ctor_set(x_13, 1, x_12); +x_14 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__9; +x_15 = l_Lean_Meta_Grind_instInhabitedGoal___closed__1; +lean_inc(x_7); +x_16 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_16, 0, x_7); +lean_ctor_set(x_16, 1, x_14); +lean_ctor_set(x_16, 2, x_15); +x_17 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__21; +lean_inc(x_8); +lean_inc(x_9); +x_18 = l_Lean_addMacroScope(x_9, x_17, x_8); +x_19 = lean_box(0); +x_20 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__20; +lean_inc(x_7); +x_21 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_21, 0, x_7); +lean_ctor_set(x_21, 1, x_20); +lean_ctor_set(x_21, 2, x_18); +lean_ctor_set(x_21, 3, x_19); +x_22 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__22; +lean_inc(x_7); +x_23 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_23, 0, x_7); +lean_ctor_set(x_23, 1, x_22); +x_24 = l_Lean_Syntax_getId(x_1); +x_25 = lean_erase_macro_scopes(x_24); +lean_inc(x_25); +x_26 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_19, x_25); +x_27 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__25; +lean_inc(x_7); +x_28 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_28, 0, x_7); +lean_ctor_set(x_28, 1, x_27); +x_29 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__30; +lean_inc(x_7); +x_30 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_30, 0, x_7); +lean_ctor_set(x_30, 1, x_29); +x_31 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__33; +lean_inc(x_7); +x_32 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_32, 0, x_7); +lean_ctor_set(x_32, 1, x_31); +x_33 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__39; +lean_inc(x_8); +lean_inc(x_9); +x_34 = l_Lean_addMacroScope(x_9, x_33, x_8); +x_35 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__37; +x_36 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__41; +lean_inc(x_7); +x_37 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_37, 0, x_7); +lean_ctor_set(x_37, 1, x_35); +lean_ctor_set(x_37, 2, x_34); +lean_ctor_set(x_37, 3, x_36); +lean_inc(x_21); +lean_inc(x_7); +x_38 = l_Lean_Syntax_node1(x_7, x_14, x_21); +x_39 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__35; +lean_inc(x_7); +x_40 = l_Lean_Syntax_node2(x_7, x_39, x_37, x_38); +x_41 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__32; +lean_inc(x_7); +x_42 = l_Lean_Syntax_node2(x_7, x_41, x_32, x_40); +x_43 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__42; +lean_inc(x_7); +x_44 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_44, 0, x_7); +lean_ctor_set(x_44, 1, x_43); +x_45 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__29; +lean_inc(x_7); +x_46 = l_Lean_Syntax_node3(x_7, x_45, x_30, x_42, x_44); +x_47 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__27; +lean_inc(x_16); +lean_inc(x_7); +x_48 = l_Lean_Syntax_node2(x_7, x_47, x_16, x_46); +x_49 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__43; +lean_inc(x_7); +x_50 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_50, 0, x_7); +lean_ctor_set(x_50, 1, x_49); +x_51 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__48; +lean_inc(x_8); +lean_inc(x_9); +x_52 = l_Lean_addMacroScope(x_9, x_51, x_8); +x_53 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__47; +x_54 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__51; +lean_inc(x_7); +x_55 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_55, 0, x_7); +lean_ctor_set(x_55, 1, x_53); +lean_ctor_set(x_55, 2, x_52); +lean_ctor_set(x_55, 3, x_54); +x_56 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__45; +lean_inc(x_7); +x_57 = l_Lean_Syntax_node1(x_7, x_56, x_55); +x_58 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__11; +lean_inc(x_16); +lean_inc(x_7); +x_59 = l_Lean_Syntax_node2(x_7, x_58, x_57, x_16); +x_60 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__55; +x_61 = l_Lean_addMacroScope(x_9, x_60, x_8); +x_62 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__53; +x_63 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__57; +lean_inc(x_7); +x_64 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_64, 0, x_7); +lean_ctor_set(x_64, 1, x_62); +lean_ctor_set(x_64, 2, x_61); +lean_ctor_set(x_64, 3, x_63); +lean_inc(x_21); +lean_inc(x_7); +x_65 = l_Lean_Syntax_node2(x_7, x_14, x_21, x_2); +lean_inc(x_7); +x_66 = l_Lean_Syntax_node2(x_7, x_39, x_64, x_65); +lean_inc(x_7); +x_67 = l_Lean_Syntax_node1(x_7, x_56, x_66); +lean_inc(x_16); +lean_inc(x_7); +x_68 = l_Lean_Syntax_node2(x_7, x_58, x_67, x_16); +lean_inc(x_7); +x_69 = l_Lean_Syntax_node2(x_7, x_14, x_59, x_68); +x_70 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__7; +lean_inc(x_7); +x_71 = l_Lean_Syntax_node1(x_7, x_70, x_69); +x_72 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__24; +lean_inc_n(x_16, 2); +lean_inc(x_7); +x_73 = l_Lean_Syntax_node6(x_7, x_72, x_28, x_48, x_50, x_71, x_16, x_16); +lean_inc(x_16); +lean_inc(x_7); +x_74 = l_Lean_Syntax_node2(x_7, x_58, x_73, x_16); +if (lean_obj_tag(x_26) == 0) +{ +lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; +x_75 = l_Lean_quoteNameMk(x_25); +x_76 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__18; +lean_inc_n(x_16, 2); +lean_inc(x_7); +x_77 = l_Lean_Syntax_node5(x_7, x_76, x_21, x_16, x_16, x_23, x_75); +x_78 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__16; +lean_inc(x_7); +x_79 = l_Lean_Syntax_node1(x_7, x_78, x_77); +x_80 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__13; +lean_inc(x_16); +lean_inc(x_7); +x_81 = l_Lean_Syntax_node3(x_7, x_80, x_13, x_16, x_79); +lean_inc(x_7); +x_82 = l_Lean_Syntax_node2(x_7, x_58, x_81, x_16); +lean_inc(x_7); +x_83 = l_Lean_Syntax_node2(x_7, x_14, x_82, x_74); +lean_inc(x_7); +x_84 = l_Lean_Syntax_node1(x_7, x_70, x_83); +x_85 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__4; +x_86 = l_Lean_Syntax_node2(x_7, x_85, x_11, x_84); +x_87 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_87, 0, x_86); +lean_ctor_set(x_87, 1, x_4); +return x_87; } else { -lean_object* x_52; size_t x_53; size_t x_54; -lean_inc(x_3); -x_52 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_52, 0, x_3); -lean_ctor_set(x_52, 1, x_48); -x_53 = 1; -x_54 = lean_usize_add(x_6, x_53); -x_6 = x_54; -x_7 = x_52; -goto _start; -} +lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; +lean_dec(x_25); +x_88 = lean_ctor_get(x_26, 0); +lean_inc(x_88); +lean_dec(x_26); +x_89 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__60; +x_90 = l_String_intercalate(x_89, x_88); +x_91 = l_Lean_Meta_Grind_updateLastTag___closed__4; +x_92 = lean_string_append(x_91, x_90); +lean_dec(x_90); +x_93 = lean_box(2); +x_94 = l_Lean_Syntax_mkNameLit(x_92, x_93); +x_95 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_95, 0, x_94); +lean_ctor_set(x_95, 1, x_19); +x_96 = lean_array_mk(x_95); +x_97 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__59; +x_98 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_98, 0, x_93); +lean_ctor_set(x_98, 1, x_97); +lean_ctor_set(x_98, 2, x_96); +x_99 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__18; +lean_inc_n(x_16, 2); +lean_inc(x_7); +x_100 = l_Lean_Syntax_node5(x_7, x_99, x_21, x_16, x_16, x_23, x_98); +x_101 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__16; +lean_inc(x_7); +x_102 = l_Lean_Syntax_node1(x_7, x_101, x_100); +x_103 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__13; +lean_inc(x_16); +lean_inc(x_7); +x_104 = l_Lean_Syntax_node3(x_7, x_103, x_13, x_16, x_102); +lean_inc(x_7); +x_105 = l_Lean_Syntax_node2(x_7, x_58, x_104, x_16); +lean_inc(x_7); +x_106 = l_Lean_Syntax_node2(x_7, x_14, x_105, x_74); +lean_inc(x_7); +x_107 = l_Lean_Syntax_node1(x_7, x_70, x_106); +x_108 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__4; +x_109 = l_Lean_Syntax_node2(x_7, x_108, x_11, x_107); +x_110 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_110, 0, x_109); +lean_ctor_set(x_110, 1, x_4); +return x_110; } } } +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("interpolatedStrKind", 19, 19); +return x_1; } } -LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_markTheoremInstance___spec__13___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__2() { _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; size_t x_12; size_t x_13; lean_object* x_14; lean_object* x_15; -x_4 = lean_ctor_get(x_1, 1); -lean_inc(x_4); -lean_dec(x_1); -x_5 = lean_array_get_size(x_4); -x_6 = lean_unsigned_to_nat(0u); -x_7 = l_Array_toSubarray___rarg(x_4, x_6, x_5); -x_8 = lean_ctor_get(x_2, 1); -x_9 = lean_box(0); -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_7); -x_12 = lean_array_size(x_8); -x_13 = 0; -x_14 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__14(x_8, x_9, x_10, x_8, x_12, x_13, x_11); -x_15 = lean_ctor_get(x_14, 0); -lean_inc(x_15); -lean_dec(x_14); -if (lean_obj_tag(x_15) == 0) -{ -uint8_t x_16; -x_16 = 1; -return x_16; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } -else -{ -lean_object* x_17; uint8_t x_18; -x_17 = lean_ctor_get(x_15, 0); -lean_inc(x_17); -lean_dec(x_15); -x_18 = lean_unbox(x_17); -lean_dec(x_17); -return x_18; } +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("typeAscription", 14, 14); +return x_1; } } -LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_markTheoremInstance___spec__13___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__4() { _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, 1); -x_5 = lean_array_get_size(x_4); -x_6 = lean_ctor_get(x_1, 1); -lean_inc(x_6); -x_7 = lean_array_get_size(x_6); -lean_dec(x_6); -x_8 = lean_nat_dec_eq(x_5, x_7); -lean_dec(x_7); -lean_dec(x_5); -if (x_8 == 0) -{ -uint8_t x_9; -lean_dec(x_1); -x_9 = 0; -return x_9; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78____closed__6; +x_2 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__1; +x_3 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__2; +x_4 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__3; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; } -else -{ -lean_object* x_10; uint8_t x_11; -x_10 = lean_box(0); -x_11 = l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_markTheoremInstance___spec__13___lambda__1(x_1, x_2, x_10); -return x_11; } +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(":", 1, 1); +return x_1; } } -LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_markTheoremInstance___spec__13(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_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__6() { _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) -{ -uint8_t x_8; -lean_dec(x_4); -x_8 = 0; -return x_8; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("MessageData", 11, 11); +return x_1; } -else -{ -lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_9 = lean_array_fget(x_1, x_4); -x_10 = lean_ctor_get(x_5, 0); -x_11 = lean_ctor_get(x_9, 0); -lean_inc(x_11); -x_12 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_10, x_11); -lean_dec(x_11); -if (x_12 == 0) +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__7() { +_start: { -lean_object* x_13; lean_object* x_14; -lean_dec(x_9); -x_13 = lean_unsigned_to_nat(1u); -x_14 = lean_nat_add(x_4, x_13); -lean_dec(x_4); -x_3 = lean_box(0); -x_4 = x_14; -goto _start; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__6; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; } -else +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__8() { +_start: { -lean_object* x_16; uint8_t x_17; -x_16 = lean_box(0); -x_17 = l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_markTheoremInstance___spec__13___lambda__2(x_9, x_5, x_16); -if (x_17 == 0) +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__6; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__9() { +_start: { -lean_object* x_18; lean_object* x_19; -x_18 = lean_unsigned_to_nat(1u); -x_19 = lean_nat_add(x_4, x_18); -lean_dec(x_4); -x_3 = lean_box(0); -x_4 = x_19; -goto _start; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78____closed__6; +x_2 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__6; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; } -else +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__10() { +_start: { -uint8_t x_21; -lean_dec(x_4); -x_21 = 1; -return x_21; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__9; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} } +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__9; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; } } +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______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; } } -LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_markTheoremInstance___spec__11___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__13() { _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; size_t x_12; size_t x_13; lean_object* x_14; lean_object* x_15; -x_4 = lean_ctor_get(x_1, 1); -lean_inc(x_4); -lean_dec(x_1); -x_5 = lean_array_get_size(x_4); -x_6 = lean_unsigned_to_nat(0u); -x_7 = l_Array_toSubarray___rarg(x_4, x_6, x_5); -x_8 = lean_ctor_get(x_2, 1); -x_9 = lean_box(0); -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_7); -x_12 = lean_array_size(x_8); -x_13 = 0; -x_14 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__12(x_8, x_9, x_10, x_8, x_12, x_13, x_11); -x_15 = lean_ctor_get(x_14, 0); -lean_inc(x_15); -lean_dec(x_14); -if (lean_obj_tag(x_15) == 0) +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__10; +x_2 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__12; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__14() { +_start: { -uint8_t x_16; -x_16 = 1; -return x_16; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("termM!_", 7, 7); +return x_1; } -else +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__15() { +_start: { -lean_object* x_17; uint8_t x_18; -x_17 = lean_ctor_get(x_15, 0); -lean_inc(x_17); -lean_dec(x_15); -x_18 = lean_unbox(x_17); -lean_dec(x_17); -return x_18; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78____closed__6; +x_2 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__14; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; } } +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__16() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("m!", 2, 2); +return x_1; +} } -LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_markTheoremInstance___spec__11___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______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, 1); -x_5 = lean_array_get_size(x_4); -x_6 = lean_ctor_get(x_1, 1); -lean_inc(x_6); -x_7 = lean_array_get_size(x_6); -lean_dec(x_6); -x_8 = lean_nat_dec_eq(x_5, x_7); -lean_dec(x_7); -lean_dec(x_5); -if (x_8 == 0) +lean_object* x_4; uint8_t x_5; +x_4 = l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__2; +lean_inc(x_1); +x_5 = l_Lean_Syntax_isOfKind(x_1, x_4); +if (x_5 == 0) { -uint8_t x_9; +lean_object* x_6; lean_object* x_7; +lean_dec(x_2); lean_dec(x_1); -x_9 = 0; -return x_9; +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_10; uint8_t x_11; -x_10 = lean_box(0); -x_11 = l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_markTheoremInstance___spec__11___lambda__1(x_1, x_2, x_10); -return x_11; -} +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; uint8_t x_15; +x_8 = lean_unsigned_to_nat(1u); +x_9 = l_Lean_Syntax_getArg(x_1, x_8); +x_10 = lean_unsigned_to_nat(3u); +x_11 = l_Lean_Syntax_getArg(x_1, x_10); +lean_dec(x_1); +lean_inc(x_11); +x_12 = l_Lean_Syntax_getKind(x_11); +x_13 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__2; +x_14 = lean_name_eq(x_12, x_13); +lean_dec(x_12); +if (x_14 == 0) +{ +uint8_t x_54; +x_54 = 0; +x_15 = x_54; +goto block_53; } +else +{ +uint8_t x_55; +x_55 = 1; +x_15 = x_55; +goto block_53; } -LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_markTheoremInstance___spec__11(lean_object* x_1, size_t x_2, lean_object* x_3) { -_start: +block_53: { -if (lean_obj_tag(x_1) == 0) +if (x_15 == 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_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___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; lean_object* x_13; uint8_t x_14; -x_11 = lean_ctor_get(x_10, 0); -lean_inc(x_11); -lean_dec(x_10); -x_12 = lean_ctor_get(x_3, 0); -x_13 = lean_ctor_get(x_11, 0); -lean_inc(x_13); -x_14 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_12, x_13); -lean_dec(x_13); -if (x_14 == 0) +lean_object* x_16; uint8_t x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; +x_16 = lean_ctor_get(x_2, 5); +lean_inc(x_16); +x_17 = 0; +x_18 = l_Lean_SourceInfo_fromRef(x_16, x_17); +lean_dec(x_16); +x_19 = lean_ctor_get(x_2, 2); +lean_inc(x_19); +x_20 = lean_ctor_get(x_2, 1); +lean_inc(x_20); +x_21 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__30; +lean_inc(x_18); +x_22 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_22, 0, x_18); +lean_ctor_set(x_22, 1, x_21); +x_23 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__5; +lean_inc(x_18); +x_24 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_24, 0, x_18); +lean_ctor_set(x_24, 1, x_23); +x_25 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__8; +x_26 = l_Lean_addMacroScope(x_20, x_25, x_19); +x_27 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__7; +x_28 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__13; +lean_inc(x_18); +x_29 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_29, 0, x_18); +lean_ctor_set(x_29, 1, x_27); +lean_ctor_set(x_29, 2, x_26); +lean_ctor_set(x_29, 3, x_28); +x_30 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__9; +lean_inc(x_18); +x_31 = l_Lean_Syntax_node1(x_18, x_30, x_29); +x_32 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__42; +lean_inc(x_18); +x_33 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_33, 0, x_18); +lean_ctor_set(x_33, 1, x_32); +x_34 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__4; +x_35 = l_Lean_Syntax_node5(x_18, x_34, x_22, x_11, x_24, x_31, x_33); +x_36 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1(x_9, x_35, x_2, x_3); +lean_dec(x_9); +x_37 = !lean_is_exclusive(x_36); +if (x_37 == 0) { -uint8_t x_15; -lean_dec(x_11); -x_15 = 0; -return x_15; +return x_36; } else { -lean_object* x_16; uint8_t x_17; -x_16 = lean_box(0); -x_17 = l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_markTheoremInstance___spec__11___lambda__2(x_11, x_3, x_16); -return x_17; +lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_38 = lean_ctor_get(x_36, 0); +x_39 = lean_ctor_get(x_36, 1); +lean_inc(x_39); +lean_inc(x_38); +lean_dec(x_36); +x_40 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_40, 0, x_38); +lean_ctor_set(x_40, 1, x_39); +return x_40; } } -case 1: +else { -lean_object* x_18; size_t x_19; -x_18 = lean_ctor_get(x_10, 0); -lean_inc(x_18); -lean_dec(x_10); -x_19 = lean_usize_shift_right(x_2, x_5); -x_1 = x_18; -x_2 = x_19; -goto _start; +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; uint8_t x_49; +x_41 = lean_ctor_get(x_2, 5); +lean_inc(x_41); +x_42 = 0; +x_43 = l_Lean_SourceInfo_fromRef(x_41, x_42); +lean_dec(x_41); +x_44 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__16; +lean_inc(x_43); +x_45 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +x_46 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__15; +x_47 = l_Lean_Syntax_node2(x_43, x_46, x_45, x_11); +x_48 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1(x_9, x_47, x_2, x_3); +lean_dec(x_9); +x_49 = !lean_is_exclusive(x_48); +if (x_49 == 0) +{ +return x_48; } -default: +else { -uint8_t x_21; -x_21 = 0; -return x_21; +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_dec(x_48); +x_52 = lean_alloc_ctor(0, 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_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; -x_22 = lean_ctor_get(x_1, 0); -lean_inc(x_22); -x_23 = lean_ctor_get(x_1, 1); -lean_inc(x_23); -lean_dec(x_1); -x_24 = lean_unsigned_to_nat(0u); -x_25 = l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_markTheoremInstance___spec__13(x_22, x_23, lean_box(0), x_24, x_3); -lean_dec(x_23); -lean_dec(x_22); -return x_25; } } } -LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_markTheoremInstance___spec__9(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -uint64_t x_3; lean_object* x_4; lean_object* x_5; size_t x_6; size_t x_7; uint64_t x_8; size_t x_9; uint8_t x_10; -x_3 = l_Lean_Meta_Grind_instHashablePreInstance_unsafe__1(x_2); -x_4 = lean_ctor_get(x_2, 1); -x_5 = lean_box(0); -x_6 = lean_array_size(x_4); -x_7 = 0; -x_8 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__10(x_4, x_5, x_4, x_6, x_7, x_3); -x_9 = lean_uint64_to_usize(x_8); -x_10 = l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_markTheoremInstance___spec__11(x_1, x_9, x_2); -return x_10; +lean_object* x_5; +x_5 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1(x_1, x_2, x_3, x_4); +lean_dec(x_1); +return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markTheoremInstance___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 uint64_t l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, uint64_t x_6) { _start: { -lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; -x_12 = lean_st_ref_take(x_3, x_11); -x_13 = lean_ctor_get(x_12, 0); -lean_inc(x_13); -x_14 = lean_ctor_get(x_12, 1); -lean_inc(x_14); -lean_dec(x_12); -x_15 = !lean_is_exclusive(x_13); -if (x_15 == 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_13, 12); -x_17 = lean_box(0); -x_18 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_markTheoremInstance___spec__1(x_16, x_1, x_17); -lean_ctor_set(x_13, 12, x_18); -x_19 = lean_st_ref_set(x_3, x_13, x_14); -x_20 = !lean_is_exclusive(x_19); -if (x_20 == 0) +uint8_t x_7; +x_7 = lean_usize_dec_lt(x_5, x_4); +if (x_7 == 0) { -lean_object* x_21; uint8_t x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_19, 0); -lean_dec(x_21); -x_22 = 1; -x_23 = lean_box(x_22); -lean_ctor_set(x_19, 0, x_23); -return x_19; +return x_6; } else { -lean_object* x_24; uint8_t x_25; lean_object* x_26; lean_object* x_27; -x_24 = lean_ctor_get(x_19, 1); -lean_inc(x_24); -lean_dec(x_19); -x_25 = 1; -x_26 = lean_box(x_25); -x_27 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_27, 0, x_26); -lean_ctor_set(x_27, 1, x_24); -return x_27; +lean_object* x_8; uint64_t x_9; uint64_t x_10; size_t x_11; size_t x_12; +x_8 = lean_array_uget(x_3, x_5); +x_9 = l_Lean_Meta_Grind_instHashablePreInstance_unsafe__2(x_8); +lean_dec(x_8); +x_10 = lean_uint64_mix_hash(x_6, x_9); +x_11 = 1; +x_12 = lean_usize_add(x_5, x_11); +x_5 = x_12; +x_6 = x_10; +goto _start; } } -else -{ -lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; 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; uint8_t x_49; lean_object* x_50; lean_object* x_51; -x_28 = lean_ctor_get(x_13, 0); -x_29 = lean_ctor_get(x_13, 1); -x_30 = lean_ctor_get(x_13, 2); -x_31 = lean_ctor_get(x_13, 3); -x_32 = lean_ctor_get(x_13, 4); -x_33 = lean_ctor_get(x_13, 5); -x_34 = lean_ctor_get_uint8(x_13, sizeof(void*)*14); -x_35 = lean_ctor_get(x_13, 6); -x_36 = lean_ctor_get(x_13, 7); -x_37 = lean_ctor_get(x_13, 8); -x_38 = lean_ctor_get(x_13, 9); -x_39 = lean_ctor_get(x_13, 10); -x_40 = lean_ctor_get(x_13, 11); -x_41 = lean_ctor_get(x_13, 12); -x_42 = lean_ctor_get(x_13, 13); -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_inc(x_33); -lean_inc(x_32); -lean_inc(x_31); -lean_inc(x_30); -lean_inc(x_29); -lean_inc(x_28); -lean_dec(x_13); -x_43 = lean_box(0); -x_44 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_markTheoremInstance___spec__1(x_41, x_1, x_43); -x_45 = lean_alloc_ctor(0, 14, 1); -lean_ctor_set(x_45, 0, x_28); -lean_ctor_set(x_45, 1, x_29); -lean_ctor_set(x_45, 2, x_30); -lean_ctor_set(x_45, 3, x_31); -lean_ctor_set(x_45, 4, x_32); -lean_ctor_set(x_45, 5, x_33); -lean_ctor_set(x_45, 6, x_35); -lean_ctor_set(x_45, 7, x_36); -lean_ctor_set(x_45, 8, x_37); -lean_ctor_set(x_45, 9, x_38); -lean_ctor_set(x_45, 10, x_39); -lean_ctor_set(x_45, 11, x_40); -lean_ctor_set(x_45, 12, x_44); -lean_ctor_set(x_45, 13, x_42); -lean_ctor_set_uint8(x_45, sizeof(void*)*14, x_34); -x_46 = lean_st_ref_set(x_3, x_45, x_14); -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); } -x_49 = 1; -x_50 = lean_box(x_49); -if (lean_is_scalar(x_48)) { - x_51 = lean_alloc_ctor(0, 2, 0); -} else { - x_51 = x_48; +LEAN_EXPORT uint64_t l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, uint64_t x_6) { +_start: +{ +uint8_t x_7; +x_7 = lean_usize_dec_lt(x_5, x_4); +if (x_7 == 0) +{ +return x_6; } -lean_ctor_set(x_51, 0, x_50); -lean_ctor_set(x_51, 1, x_47); -return x_51; +else +{ +lean_object* x_8; uint64_t x_9; uint64_t x_10; size_t x_11; size_t x_12; +x_8 = lean_array_uget(x_3, x_5); +x_9 = l_Lean_Meta_Grind_instHashablePreInstance_unsafe__2(x_8); +lean_dec(x_8); +x_10 = lean_uint64_mix_hash(x_6, x_9); +x_11 = 1; +x_12 = lean_usize_add(x_5, x_11); +x_5 = x_12; +x_6 = x_10; +goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markTheoremInstance(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_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_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_markTheoremInstance___spec__4(size_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { -lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_12 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_12, 0, x_1); -lean_ctor_set(x_12, 1, x_2); -x_13 = lean_st_ref_get(x_3, x_11); -x_14 = !lean_is_exclusive(x_13); -if (x_14 == 0) -{ -lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; -x_15 = lean_ctor_get(x_13, 0); -x_16 = lean_ctor_get(x_13, 1); -x_17 = lean_ctor_get(x_15, 12); -lean_inc(x_17); -lean_dec(x_15); -x_18 = l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_markTheoremInstance___spec__9(x_17, x_12); -if (x_18 == 0) +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_19; lean_object* x_20; -lean_free_object(x_13); -x_19 = lean_box(0); -x_20 = l_Lean_Meta_Grind_markTheoremInstance___lambda__1(x_12, x_19, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_16); -return x_20; +lean_dec(x_5); +return x_6; } else { -uint8_t x_21; lean_object* x_22; +lean_object* x_9; lean_object* x_10; uint64_t x_11; lean_object* x_12; lean_object* x_13; size_t x_14; size_t x_15; uint64_t x_16; size_t x_17; size_t x_18; size_t x_19; size_t x_20; size_t x_21; size_t x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_9 = lean_array_fget(x_2, x_5); +x_10 = lean_array_fget(x_3, x_5); +x_11 = l_Lean_Meta_Grind_instHashablePreInstance_unsafe__1(x_9); +x_12 = lean_ctor_get(x_9, 1); +lean_inc(x_12); +x_13 = lean_box(0); +x_14 = lean_array_size(x_12); +x_15 = 0; +x_16 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__5(x_12, x_13, x_12, x_14, x_15, x_11); lean_dec(x_12); -x_21 = 0; -x_22 = lean_box(x_21); -lean_ctor_set(x_13, 0, x_22); -return x_13; -} -} -else -{ -lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; -x_23 = lean_ctor_get(x_13, 0); -x_24 = lean_ctor_get(x_13, 1); -lean_inc(x_24); -lean_inc(x_23); -lean_dec(x_13); -x_25 = lean_ctor_get(x_23, 12); -lean_inc(x_25); -lean_dec(x_23); -x_26 = l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_markTheoremInstance___spec__9(x_25, x_12); -if (x_26 == 0) -{ -lean_object* x_27; lean_object* x_28; -x_27 = lean_box(0); -x_28 = l_Lean_Meta_Grind_markTheoremInstance___lambda__1(x_12, 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_29; lean_object* x_30; lean_object* x_31; -lean_dec(x_12); -x_29 = 0; -x_30 = lean_box(x_29); -x_31 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_31, 0, x_30); -lean_ctor_set(x_31, 1, x_24); -return x_31; -} +x_17 = lean_uint64_to_usize(x_16); +x_18 = 1; +x_19 = lean_usize_sub(x_1, x_18); +x_20 = 5; +x_21 = lean_usize_mul(x_20, x_19); +x_22 = lean_usize_shift_right(x_17, x_21); +x_23 = lean_unsigned_to_nat(1u); +x_24 = lean_nat_add(x_5, x_23); +lean_dec(x_5); +x_25 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3(x_6, x_22, x_1, x_9, x_10); +x_4 = lean_box(0); +x_5 = x_24; +x_6 = x_25; +goto _start; } } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___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_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__6(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: { -size_t x_7; size_t x_8; uint64_t x_9; uint64_t x_10; lean_object* x_11; -x_7 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_8 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_9 = lean_unbox_uint64(x_6); -lean_dec(x_6); -x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__2(x_1, x_2, x_3, x_7, x_8, x_9); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_11 = lean_box_uint64(x_10); -return x_11; -} -} -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___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_8; +x_8 = lean_usize_dec_lt(x_6, x_5); +if (x_8 == 0) { -size_t x_7; size_t x_8; uint64_t x_9; uint64_t x_10; lean_object* x_11; -x_7 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_8 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_9 = lean_unbox_uint64(x_6); -lean_dec(x_6); -x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__5(x_1, x_2, x_3, x_7, x_8, x_9); lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_11 = lean_box_uint64(x_10); -return x_11; -} +return x_7; } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_markTheoremInstance___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: +else { -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_Grind_markTheoremInstance___spec__4(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_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___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) { -_start: +lean_object* x_9; uint8_t x_10; +x_9 = lean_array_uget(x_4, x_6); +x_10 = !lean_is_exclusive(x_7); +if (x_10 == 0) { -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_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__6(x_1, x_2, x_3, x_4, x_8, x_9, x_7); -lean_dec(x_4); -lean_dec(x_2); -lean_dec(x_1); -return x_10; -} -} -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: +lean_object* x_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_7, 1); +x_12 = lean_ctor_get(x_7, 0); +lean_dec(x_12); +x_13 = lean_ctor_get(x_11, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_11, 1); +lean_inc(x_14); +x_15 = lean_ctor_get(x_11, 2); +lean_inc(x_15); +x_16 = lean_nat_dec_lt(x_14, x_15); +if (x_16 == 0) { -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_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__8(x_1, x_2, x_3, x_4, x_8, x_9, x_7); -lean_dec(x_4); -lean_dec(x_2); -lean_dec(x_1); -return x_10; -} +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_9); +lean_ctor_set(x_7, 0, x_3); +return x_7; } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_markTheoremInstance___spec__7___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: +else { -uint8_t x_4; lean_object* x_5; -x_4 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_markTheoremInstance___spec__7___lambda__1(x_1, x_2, x_3); +uint8_t x_17; +x_17 = !lean_is_exclusive(x_11); +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; uint8_t x_24; +x_18 = lean_ctor_get(x_11, 2); +lean_dec(x_18); +x_19 = lean_ctor_get(x_11, 1); +lean_dec(x_19); +x_20 = lean_ctor_get(x_11, 0); +lean_dec(x_20); +x_21 = lean_array_fget(x_13, x_14); +x_22 = lean_unsigned_to_nat(1u); +x_23 = lean_nat_add(x_14, x_22); +lean_dec(x_14); +lean_ctor_set(x_11, 1, x_23); +x_24 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_9, x_21); +lean_dec(x_21); +lean_dec(x_9); +if (x_24 == 0) +{ +lean_object* x_25; lean_dec(x_3); -lean_dec(x_2); -x_5 = lean_box(x_4); -return x_5; -} +x_25 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1___closed__1; +lean_ctor_set(x_7, 0, x_25); +return x_7; } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_markTheoremInstance___spec__7___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: +else { -uint8_t x_4; lean_object* x_5; -x_4 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_markTheoremInstance___spec__7___lambda__2(x_1, x_2, x_3); -lean_dec(x_3); -lean_dec(x_2); -x_5 = lean_box(x_4); -return x_5; +size_t x_26; size_t x_27; +lean_inc(x_3); +lean_ctor_set(x_7, 0, x_3); +x_26 = 1; +x_27 = lean_usize_add(x_6, x_26); +x_6 = x_27; +goto _start; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: +else { -uint8_t x_4; lean_object* x_5; -x_4 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3___lambda__1(x_1, x_2, x_3); +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; +lean_dec(x_11); +x_29 = lean_array_fget(x_13, x_14); +x_30 = lean_unsigned_to_nat(1u); +x_31 = lean_nat_add(x_14, x_30); +lean_dec(x_14); +x_32 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_32, 0, x_13); +lean_ctor_set(x_32, 1, x_31); +lean_ctor_set(x_32, 2, x_15); +x_33 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_9, x_29); +lean_dec(x_29); +lean_dec(x_9); +if (x_33 == 0) +{ +lean_object* x_34; lean_dec(x_3); -lean_dec(x_2); -x_5 = lean_box(x_4); -return x_5; -} +x_34 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1___closed__1; +lean_ctor_set(x_7, 1, x_32); +lean_ctor_set(x_7, 0, x_34); +return x_7; } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: +else { -uint8_t x_4; lean_object* x_5; -x_4 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3___lambda__2(x_1, x_2, x_3); -lean_dec(x_3); -lean_dec(x_2); -x_5 = lean_box(x_4); -return x_5; +size_t x_35; size_t x_36; +lean_inc(x_3); +lean_ctor_set(x_7, 1, x_32); +lean_ctor_set(x_7, 0, x_3); +x_35 = 1; +x_36 = lean_usize_add(x_6, x_35); +x_6 = x_36; +goto _start; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: -{ -size_t x_6; size_t x_7; lean_object* x_8; -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_Grind_markTheoremInstance___spec__3(x_1, x_6, x_7, x_4, x_5); -return x_8; } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___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: +else { -size_t x_7; size_t x_8; uint64_t x_9; uint64_t x_10; lean_object* x_11; -x_7 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_8 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_9 = lean_unbox_uint64(x_6); -lean_dec(x_6); -x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__10(x_1, x_2, x_3, x_7, x_8, x_9); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_11 = lean_box_uint64(x_10); -return x_11; -} -} -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__12___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; uint8_t x_42; +x_38 = lean_ctor_get(x_7, 1); +lean_inc(x_38); +lean_dec(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); +x_41 = lean_ctor_get(x_38, 2); +lean_inc(x_41); +x_42 = lean_nat_dec_lt(x_40, x_41); +if (x_42 == 0) { -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_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__12(x_1, x_2, x_3, x_4, x_8, x_9, x_7); -lean_dec(x_4); -lean_dec(x_2); -lean_dec(x_1); -return x_10; -} +lean_object* x_43; +lean_dec(x_41); +lean_dec(x_40); +lean_dec(x_39); +lean_dec(x_9); +x_43 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_43, 0, x_3); +lean_ctor_set(x_43, 1, x_38); +return x_43; } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___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) { -_start: +else { -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_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__14(x_1, x_2, x_3, x_4, x_8, x_9, x_7); -lean_dec(x_4); -lean_dec(x_2); -lean_dec(x_1); -return x_10; +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; uint8_t x_49; +if (lean_is_exclusive(x_38)) { + lean_ctor_release(x_38, 0); + lean_ctor_release(x_38, 1); + lean_ctor_release(x_38, 2); + x_44 = x_38; +} else { + lean_dec_ref(x_38); + x_44 = lean_box(0); } +x_45 = lean_array_fget(x_39, x_40); +x_46 = lean_unsigned_to_nat(1u); +x_47 = lean_nat_add(x_40, x_46); +lean_dec(x_40); +if (lean_is_scalar(x_44)) { + x_48 = lean_alloc_ctor(0, 3, 0); +} else { + x_48 = x_44; } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_markTheoremInstance___spec__13___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: +lean_ctor_set(x_48, 0, x_39); +lean_ctor_set(x_48, 1, x_47); +lean_ctor_set(x_48, 2, x_41); +x_49 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_9, x_45); +lean_dec(x_45); +lean_dec(x_9); +if (x_49 == 0) { -uint8_t x_4; lean_object* x_5; -x_4 = l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_markTheoremInstance___spec__13___lambda__1(x_1, x_2, x_3); +lean_object* x_50; lean_object* x_51; lean_dec(x_3); -lean_dec(x_2); -x_5 = lean_box(x_4); -return x_5; -} +x_50 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1___closed__1; +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_EXPORT lean_object* l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_markTheoremInstance___spec__13___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: +else { -uint8_t x_4; lean_object* x_5; -x_4 = l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_markTheoremInstance___spec__13___lambda__2(x_1, x_2, x_3); -lean_dec(x_3); -lean_dec(x_2); -x_5 = lean_box(x_4); -return x_5; +lean_object* x_52; size_t x_53; size_t x_54; +lean_inc(x_3); +x_52 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_52, 0, x_3); +lean_ctor_set(x_52, 1, x_48); +x_53 = 1; +x_54 = lean_usize_add(x_6, x_53); +x_6 = x_54; +x_7 = x_52; +goto _start; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_markTheoremInstance___spec__13___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 = l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_markTheoremInstance___spec__13(x_1, x_2, x_3, x_4, x_5); -lean_dec(x_5); -lean_dec(x_2); -lean_dec(x_1); -x_7 = lean_box(x_6); -return x_7; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_markTheoremInstance___spec__11___lambda__1___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_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_markTheoremInstance___spec__11___lambda__1(x_1, x_2, x_3); -lean_dec(x_3); -lean_dec(x_2); -x_5 = lean_box(x_4); -return x_5; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_markTheoremInstance___spec__11___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___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) { _start: { -uint8_t x_4; lean_object* x_5; -x_4 = l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_markTheoremInstance___spec__11___lambda__2(x_1, x_2, x_3); +uint8_t x_8; +x_8 = lean_usize_dec_lt(x_6, x_5); +if (x_8 == 0) +{ lean_dec(x_3); -lean_dec(x_2); -x_5 = lean_box(x_4); -return x_5; +return x_7; } +else +{ +lean_object* x_9; uint8_t x_10; +x_9 = lean_array_uget(x_4, x_6); +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; uint8_t x_16; +x_11 = lean_ctor_get(x_7, 1); +x_12 = lean_ctor_get(x_7, 0); +lean_dec(x_12); +x_13 = lean_ctor_get(x_11, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_11, 1); +lean_inc(x_14); +x_15 = lean_ctor_get(x_11, 2); +lean_inc(x_15); +x_16 = lean_nat_dec_lt(x_14, x_15); +if (x_16 == 0) +{ +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_9); +lean_ctor_set(x_7, 0, x_3); +return x_7; } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_markTheoremInstance___spec__11___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: +else { -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_Meta_Grind_markTheoremInstance___spec__11(x_1, x_4, x_3); +uint8_t x_17; +x_17 = !lean_is_exclusive(x_11); +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; uint8_t x_24; +x_18 = lean_ctor_get(x_11, 2); +lean_dec(x_18); +x_19 = lean_ctor_get(x_11, 1); +lean_dec(x_19); +x_20 = lean_ctor_get(x_11, 0); +lean_dec(x_20); +x_21 = lean_array_fget(x_13, x_14); +x_22 = lean_unsigned_to_nat(1u); +x_23 = lean_nat_add(x_14, x_22); +lean_dec(x_14); +lean_ctor_set(x_11, 1, x_23); +x_24 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_9, x_21); +lean_dec(x_21); +lean_dec(x_9); +if (x_24 == 0) +{ +lean_object* x_25; lean_dec(x_3); -x_6 = lean_box(x_5); -return x_6; -} +x_25 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1___closed__1; +lean_ctor_set(x_7, 0, x_25); +return x_7; } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_markTheoremInstance___spec__9___boxed(lean_object* x_1, lean_object* x_2) { -_start: +else { -uint8_t x_3; lean_object* x_4; -x_3 = l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_markTheoremInstance___spec__9(x_1, x_2); -lean_dec(x_2); -x_4 = lean_box(x_3); -return x_4; +size_t x_26; size_t x_27; +lean_inc(x_3); +lean_ctor_set(x_7, 0, x_3); +x_26 = 1; +x_27 = lean_usize_add(x_6, x_26); +x_6 = x_27; +goto _start; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markTheoremInstance___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: +else { -lean_object* x_12; -x_12 = l_Lean_Meta_Grind_markTheoremInstance___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_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; +lean_dec(x_11); +x_29 = lean_array_fget(x_13, x_14); +x_30 = lean_unsigned_to_nat(1u); +x_31 = lean_nat_add(x_14, x_30); +lean_dec(x_14); +x_32 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_32, 0, x_13); +lean_ctor_set(x_32, 1, x_31); +lean_ctor_set(x_32, 2, x_15); +x_33 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_9, x_29); +lean_dec(x_29); lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); +if (x_33 == 0) +{ +lean_object* x_34; lean_dec(x_3); -lean_dec(x_2); -return x_12; -} +x_34 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1___closed__1; +lean_ctor_set(x_7, 1, x_32); +lean_ctor_set(x_7, 0, x_34); +return x_7; } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markTheoremInstance___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: +else { -lean_object* x_12; -x_12 = l_Lean_Meta_Grind_markTheoremInstance(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -return x_12; +size_t x_35; size_t x_36; +lean_inc(x_3); +lean_ctor_set(x_7, 1, x_32); +lean_ctor_set(x_7, 0, x_3); +x_35 = 1; +x_36 = lean_usize_add(x_6, x_35); +x_6 = x_36; +goto _start; +} } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addNewFact(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; -x_13 = lean_st_ref_take(x_4, x_12); -x_14 = lean_ctor_get(x_13, 0); -lean_inc(x_14); -x_15 = lean_ctor_get(x_13, 1); -lean_inc(x_15); -lean_dec(x_13); -x_16 = !lean_is_exclusive(x_14); -if (x_16 == 0) -{ -lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; -x_17 = lean_ctor_get(x_14, 13); -x_18 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_18, 0, x_1); -lean_ctor_set(x_18, 1, x_2); -lean_ctor_set(x_18, 2, x_3); -x_19 = l_Std_Queue_enqueue___rarg(x_18, x_17); -lean_ctor_set(x_14, 13, x_19); -x_20 = lean_st_ref_set(x_4, x_14, x_15); -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); -lean_dec(x_22); -x_23 = lean_box(0); -lean_ctor_set(x_20, 0, x_23); -return x_20; } else { -lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_24 = lean_ctor_get(x_20, 1); -lean_inc(x_24); -lean_dec(x_20); -x_25 = lean_box(0); -x_26 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_26, 0, x_25); -lean_ctor_set(x_26, 1, x_24); -return x_26; -} +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; uint8_t x_42; +x_38 = lean_ctor_get(x_7, 1); +lean_inc(x_38); +lean_dec(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); +x_41 = lean_ctor_get(x_38, 2); +lean_inc(x_41); +x_42 = lean_nat_dec_lt(x_40, x_41); +if (x_42 == 0) +{ +lean_object* x_43; +lean_dec(x_41); +lean_dec(x_40); +lean_dec(x_39); +lean_dec(x_9); +x_43 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_43, 0, x_3); +lean_ctor_set(x_43, 1, x_38); +return x_43; } 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; 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; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; -x_27 = lean_ctor_get(x_14, 0); -x_28 = lean_ctor_get(x_14, 1); -x_29 = lean_ctor_get(x_14, 2); -x_30 = lean_ctor_get(x_14, 3); -x_31 = lean_ctor_get(x_14, 4); -x_32 = lean_ctor_get(x_14, 5); -x_33 = lean_ctor_get_uint8(x_14, sizeof(void*)*14); -x_34 = lean_ctor_get(x_14, 6); -x_35 = lean_ctor_get(x_14, 7); -x_36 = lean_ctor_get(x_14, 8); -x_37 = lean_ctor_get(x_14, 9); -x_38 = lean_ctor_get(x_14, 10); -x_39 = lean_ctor_get(x_14, 11); -x_40 = lean_ctor_get(x_14, 12); -x_41 = lean_ctor_get(x_14, 13); -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); -lean_inc(x_32); -lean_inc(x_31); -lean_inc(x_30); -lean_inc(x_29); -lean_inc(x_28); -lean_inc(x_27); -lean_dec(x_14); -x_42 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_42, 0, x_1); -lean_ctor_set(x_42, 1, x_2); -lean_ctor_set(x_42, 2, x_3); -x_43 = l_Std_Queue_enqueue___rarg(x_42, x_41); -x_44 = lean_alloc_ctor(0, 14, 1); -lean_ctor_set(x_44, 0, x_27); -lean_ctor_set(x_44, 1, x_28); -lean_ctor_set(x_44, 2, x_29); -lean_ctor_set(x_44, 3, x_30); -lean_ctor_set(x_44, 4, x_31); -lean_ctor_set(x_44, 5, x_32); -lean_ctor_set(x_44, 6, x_34); -lean_ctor_set(x_44, 7, x_35); -lean_ctor_set(x_44, 8, x_36); -lean_ctor_set(x_44, 9, x_37); -lean_ctor_set(x_44, 10, x_38); -lean_ctor_set(x_44, 11, x_39); -lean_ctor_set(x_44, 12, x_40); -lean_ctor_set(x_44, 13, x_43); -lean_ctor_set_uint8(x_44, sizeof(void*)*14, x_33); -x_45 = lean_st_ref_set(x_4, x_44, x_15); -x_46 = lean_ctor_get(x_45, 1); -lean_inc(x_46); -if (lean_is_exclusive(x_45)) { - lean_ctor_release(x_45, 0); - lean_ctor_release(x_45, 1); - x_47 = x_45; +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; uint8_t x_49; +if (lean_is_exclusive(x_38)) { + lean_ctor_release(x_38, 0); + lean_ctor_release(x_38, 1); + lean_ctor_release(x_38, 2); + x_44 = x_38; } else { - lean_dec_ref(x_45); - x_47 = lean_box(0); + lean_dec_ref(x_38); + x_44 = lean_box(0); } -x_48 = lean_box(0); -if (lean_is_scalar(x_47)) { - x_49 = lean_alloc_ctor(0, 2, 0); +x_45 = lean_array_fget(x_39, x_40); +x_46 = lean_unsigned_to_nat(1u); +x_47 = lean_nat_add(x_40, x_46); +lean_dec(x_40); +if (lean_is_scalar(x_44)) { + x_48 = lean_alloc_ctor(0, 3, 0); } else { - x_49 = x_47; + x_48 = x_44; } -lean_ctor_set(x_49, 0, x_48); -lean_ctor_set(x_49, 1, x_46); -return x_49; +lean_ctor_set(x_48, 0, x_39); +lean_ctor_set(x_48, 1, x_47); +lean_ctor_set(x_48, 2, x_41); +x_49 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_9, x_45); +lean_dec(x_45); +lean_dec(x_9); +if (x_49 == 0) +{ +lean_object* x_50; lean_object* x_51; +lean_dec(x_3); +x_50 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1___closed__1; +x_51 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_51, 0, x_50); +lean_ctor_set(x_51, 1, x_48); +return x_51; +} +else +{ +lean_object* x_52; size_t x_53; size_t x_54; +lean_inc(x_3); +x_52 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_52, 0, x_3); +lean_ctor_set(x_52, 1, x_48); +x_53 = 1; +x_54 = lean_usize_add(x_6, x_53); +x_6 = x_54; +x_7 = x_52; +goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addNewFact___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -lean_object* x_13; -x_13 = l_Lean_Meta_Grind_addNewFact(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -return x_13; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addTheoremInstance(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_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_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_markTheoremInstance___spec__7___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; -x_13 = l_Lean_Meta_Grind_addNewFact(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_14 = lean_ctor_get(x_13, 1); -lean_inc(x_14); -lean_dec(x_13); -x_15 = lean_st_ref_take(x_4, x_14); -x_16 = lean_ctor_get(x_15, 0); -lean_inc(x_16); -x_17 = lean_ctor_get(x_15, 1); -lean_inc(x_17); -lean_dec(x_15); -x_18 = !lean_is_exclusive(x_16); -if (x_18 == 0) -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; -x_19 = lean_ctor_get(x_16, 11); -x_20 = lean_unsigned_to_nat(1u); -x_21 = lean_nat_add(x_19, x_20); -lean_dec(x_19); -lean_ctor_set(x_16, 11, x_21); -x_22 = lean_st_ref_set(x_4, x_16, x_17); -x_23 = !lean_is_exclusive(x_22); -if (x_23 == 0) +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; size_t x_12; size_t x_13; lean_object* x_14; lean_object* x_15; +x_4 = lean_ctor_get(x_1, 1); +lean_inc(x_4); +lean_dec(x_1); +x_5 = lean_array_get_size(x_4); +x_6 = lean_unsigned_to_nat(0u); +x_7 = l_Array_toSubarray___rarg(x_4, x_6, x_5); +x_8 = lean_ctor_get(x_2, 1); +x_9 = lean_box(0); +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_7); +x_12 = lean_array_size(x_8); +x_13 = 0; +x_14 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__8(x_8, x_9, x_10, x_8, x_12, x_13, x_11); +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +lean_dec(x_14); +if (lean_obj_tag(x_15) == 0) { -lean_object* x_24; lean_object* x_25; -x_24 = lean_ctor_get(x_22, 0); -lean_dec(x_24); -x_25 = lean_box(0); -lean_ctor_set(x_22, 0, x_25); -return x_22; +uint8_t x_16; +x_16 = 1; +return x_16; } else { -lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_26 = lean_ctor_get(x_22, 1); -lean_inc(x_26); -lean_dec(x_22); -x_27 = lean_box(0); -x_28 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_28, 0, x_27); -lean_ctor_set(x_28, 1, x_26); -return x_28; +lean_object* x_17; uint8_t x_18; +x_17 = lean_ctor_get(x_15, 0); +lean_inc(x_17); +lean_dec(x_15); +x_18 = lean_unbox(x_17); +lean_dec(x_17); +return x_18; } } -else -{ -lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; -x_29 = lean_ctor_get(x_16, 0); -x_30 = lean_ctor_get(x_16, 1); -x_31 = lean_ctor_get(x_16, 2); -x_32 = lean_ctor_get(x_16, 3); -x_33 = lean_ctor_get(x_16, 4); -x_34 = lean_ctor_get(x_16, 5); -x_35 = lean_ctor_get_uint8(x_16, sizeof(void*)*14); -x_36 = lean_ctor_get(x_16, 6); -x_37 = lean_ctor_get(x_16, 7); -x_38 = lean_ctor_get(x_16, 8); -x_39 = lean_ctor_get(x_16, 9); -x_40 = lean_ctor_get(x_16, 10); -x_41 = lean_ctor_get(x_16, 11); -x_42 = lean_ctor_get(x_16, 12); -x_43 = lean_ctor_get(x_16, 13); -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_34); -lean_inc(x_33); -lean_inc(x_32); -lean_inc(x_31); -lean_inc(x_30); -lean_inc(x_29); -lean_dec(x_16); -x_44 = lean_unsigned_to_nat(1u); -x_45 = lean_nat_add(x_41, x_44); -lean_dec(x_41); -x_46 = lean_alloc_ctor(0, 14, 1); -lean_ctor_set(x_46, 0, x_29); -lean_ctor_set(x_46, 1, x_30); -lean_ctor_set(x_46, 2, x_31); -lean_ctor_set(x_46, 3, x_32); -lean_ctor_set(x_46, 4, x_33); -lean_ctor_set(x_46, 5, x_34); -lean_ctor_set(x_46, 6, x_36); -lean_ctor_set(x_46, 7, x_37); -lean_ctor_set(x_46, 8, x_38); -lean_ctor_set(x_46, 9, x_39); -lean_ctor_set(x_46, 10, x_40); -lean_ctor_set(x_46, 11, x_45); -lean_ctor_set(x_46, 12, x_42); -lean_ctor_set(x_46, 13, x_43); -lean_ctor_set_uint8(x_46, sizeof(void*)*14, x_35); -x_47 = lean_st_ref_set(x_4, x_46, x_17); -x_48 = lean_ctor_get(x_47, 1); -lean_inc(x_48); -if (lean_is_exclusive(x_47)) { - lean_ctor_release(x_47, 0); - lean_ctor_release(x_47, 1); - x_49 = x_47; -} else { - lean_dec_ref(x_47); - x_49 = lean_box(0); } -x_50 = lean_box(0); -if (lean_is_scalar(x_49)) { - x_51 = lean_alloc_ctor(0, 2, 0); -} else { - x_51 = x_49; -} -lean_ctor_set(x_51, 0, x_50); -lean_ctor_set(x_51, 1, x_48); -return x_51; -} -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addTheoremInstance___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_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_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_markTheoremInstance___spec__7___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_13; -x_13 = l_Lean_Meta_Grind_addTheoremInstance(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); +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, 1); +x_5 = lean_array_get_size(x_4); +x_6 = lean_ctor_get(x_1, 1); +lean_inc(x_6); +x_7 = lean_array_get_size(x_6); lean_dec(x_6); +x_8 = lean_nat_dec_eq(x_5, x_7); +lean_dec(x_7); lean_dec(x_5); -lean_dec(x_4); -return x_13; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_checkMaxInstancesExceeded(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_10 = lean_st_ref_get(x_1, 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_Meta_Grind_getConfig___rarg(x_3, x_4, x_5, x_6, x_7, x_8, x_12); -x_14 = !lean_is_exclusive(x_13); -if (x_14 == 0) +if (x_8 == 0) { -lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; lean_object* x_19; -x_15 = lean_ctor_get(x_13, 0); -x_16 = lean_ctor_get(x_15, 3); -lean_inc(x_16); -lean_dec(x_15); -x_17 = lean_ctor_get(x_11, 11); -lean_inc(x_17); -lean_dec(x_11); -x_18 = lean_nat_dec_le(x_16, x_17); -lean_dec(x_17); -lean_dec(x_16); -x_19 = lean_box(x_18); -lean_ctor_set(x_13, 0, x_19); -return x_13; +uint8_t x_9; +lean_dec(x_1); +x_9 = 0; +return x_9; } else { -lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; lean_object* x_25; lean_object* x_26; -x_20 = lean_ctor_get(x_13, 0); -x_21 = lean_ctor_get(x_13, 1); -lean_inc(x_21); -lean_inc(x_20); -lean_dec(x_13); -x_22 = lean_ctor_get(x_20, 3); -lean_inc(x_22); -lean_dec(x_20); -x_23 = lean_ctor_get(x_11, 11); -lean_inc(x_23); -lean_dec(x_11); -x_24 = lean_nat_dec_le(x_22, x_23); -lean_dec(x_23); -lean_dec(x_22); -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_21); -return x_26; +lean_object* x_10; uint8_t x_11; +x_10 = lean_box(0); +x_11 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_markTheoremInstance___spec__7___lambda__1(x_1, x_2, x_10); +return x_11; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_checkMaxInstancesExceeded___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_markTheoremInstance___spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -lean_object* x_10; -x_10 = l_Lean_Meta_Grind_checkMaxInstancesExceeded(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_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); -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_10; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getENode_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: +if (x_8 == 0) { -lean_object* x_11; uint8_t x_12; -x_11 = lean_st_ref_get(x_2, x_10); -x_12 = !lean_is_exclusive(x_11); -if (x_12 == 0) +uint8_t x_9; +lean_dec(x_2); +x_9 = !lean_is_exclusive(x_1); +if (x_9 == 0) { -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_13, 1); -lean_inc(x_14); -lean_dec(x_13); -x_15 = l_Lean_PersistentHashMap_find_x3f___at___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot___spec__1(x_14, x_1); -lean_ctor_set(x_11, 0, x_15); -return x_11; +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_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; -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_ctor_get(x_16, 1); -lean_inc(x_18); -lean_dec(x_16); -x_19 = l_Lean_PersistentHashMap_find_x3f___at___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot___spec__1(x_18, x_1); -x_20 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_17); -return x_20; -} +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; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getENode_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, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: +else { -lean_object* x_11; -x_11 = l_Lean_Meta_Grind_getENode_x3f(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_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_17 = lean_array_fget(x_5, x_2); +x_18 = lean_ctor_get(x_3, 0); +lean_inc(x_18); +x_19 = lean_ctor_get(x_17, 0); +lean_inc(x_19); +x_20 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_18, x_19); +lean_dec(x_19); +lean_dec(x_18); +if (x_20 == 0) +{ +lean_object* x_21; lean_object* x_22; +lean_dec(x_17); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); +x_21 = lean_unsigned_to_nat(1u); +x_22 = lean_nat_add(x_2, x_21); lean_dec(x_2); -lean_dec(x_1); -return x_11; -} +x_2 = x_22; +goto _start; } -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_Grind_getENode___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: +else { -lean_object* x_11; lean_object* x_12; uint8_t x_13; -x_11 = lean_ctor_get(x_8, 5); -x_12 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_1, x_6, x_7, x_8, x_9, x_10); -x_13 = !lean_is_exclusive(x_12); -if (x_13 == 0) +lean_object* x_24; uint8_t x_25; +x_24 = lean_box(0); +x_25 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_markTheoremInstance___spec__7___lambda__2(x_17, x_3, x_24); +if (x_25 == 0) { -lean_object* x_14; lean_object* x_15; -x_14 = lean_ctor_get(x_12, 0); -lean_inc(x_11); -x_15 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_15, 0, x_11); -lean_ctor_set(x_15, 1, x_14); -lean_ctor_set_tag(x_12, 1); -lean_ctor_set(x_12, 0, x_15); -return x_12; +lean_object* x_26; lean_object* x_27; +lean_dec(x_6); +lean_dec(x_5); +x_26 = lean_unsigned_to_nat(1u); +x_27 = lean_nat_add(x_2, x_26); +lean_dec(x_2); +x_2 = x_27; +goto _start; } else { -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_16 = lean_ctor_get(x_12, 0); -x_17 = lean_ctor_get(x_12, 1); -lean_inc(x_17); -lean_inc(x_16); -lean_dec(x_12); -lean_inc(x_11); -x_18 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_18, 0, x_11); -lean_ctor_set(x_18, 1, x_16); -x_19 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_19, 0, x_18); -lean_ctor_set(x_19, 1, x_17); -return x_19; -} -} -} -static lean_object* _init_l_Lean_Meta_Grind_getENode___closed__1() { -_start: +uint8_t x_29; +x_29 = !lean_is_exclusive(x_1); +if (x_29 == 0) { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("internal `grind` error, term has not been internalized", 54, 54); +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_30 = lean_ctor_get(x_1, 1); +lean_dec(x_30); +x_31 = lean_ctor_get(x_1, 0); +lean_dec(x_31); +x_32 = lean_array_fset(x_5, x_2, x_3); +x_33 = lean_array_fset(x_6, x_2, x_4); +lean_dec(x_2); +lean_ctor_set(x_1, 1, x_33); +lean_ctor_set(x_1, 0, x_32); return x_1; } -} -static lean_object* _init_l_Lean_Meta_Grind_getENode___closed__2() { -_start: +else { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_Grind_getENode___closed__1; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; +lean_object* x_34; lean_object* x_35; lean_object* x_36; +lean_dec(x_1); +x_34 = lean_array_fset(x_5, x_2, x_3); +x_35 = lean_array_fset(x_6, x_2, x_4); +lean_dec(x_2); +x_36 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_36, 0, x_34); +lean_ctor_set(x_36, 1, x_35); +return x_36; } } -static lean_object* _init_l_Lean_Meta_Grind_getENode___closed__3() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("", 0, 0); -return x_1; } } -static lean_object* _init_l_Lean_Meta_Grind_getENode___closed__4() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_Grind_getENode___closed__3; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getENode(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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 uint8_t l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_11; uint8_t x_12; -x_11 = lean_st_ref_get(x_2, x_10); -x_12 = !lean_is_exclusive(x_11); -if (x_12 == 0) -{ -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_13 = lean_ctor_get(x_11, 0); -x_14 = lean_ctor_get(x_11, 1); -x_15 = lean_ctor_get(x_13, 1); -lean_inc(x_15); -lean_dec(x_13); -x_16 = l_Lean_PersistentHashMap_find_x3f___at___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot___spec__1(x_15, x_1); -if (lean_obj_tag(x_16) == 0) -{ -lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; -lean_free_object(x_11); -x_17 = l_Lean_indentExpr(x_1); -x_18 = l_Lean_Meta_Grind_getENode___closed__2; -x_19 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_19, 0, x_18); -lean_ctor_set(x_19, 1, x_17); -x_20 = l_Lean_Meta_Grind_getENode___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_Meta_Grind_getENode___spec__1(x_21, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_14); -return x_22; -} -else -{ -lean_object* x_23; +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; size_t x_12; size_t x_13; lean_object* x_14; lean_object* x_15; +x_4 = lean_ctor_get(x_1, 1); +lean_inc(x_4); lean_dec(x_1); -x_23 = lean_ctor_get(x_16, 0); -lean_inc(x_23); -lean_dec(x_16); -lean_ctor_set(x_11, 0, x_23); -return x_11; -} -} -else -{ -lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; -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_26 = lean_ctor_get(x_24, 1); -lean_inc(x_26); -lean_dec(x_24); -x_27 = l_Lean_PersistentHashMap_find_x3f___at___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot___spec__1(x_26, x_1); -if (lean_obj_tag(x_27) == 0) +x_5 = lean_array_get_size(x_4); +x_6 = lean_unsigned_to_nat(0u); +x_7 = l_Array_toSubarray___rarg(x_4, x_6, x_5); +x_8 = lean_ctor_get(x_2, 1); +x_9 = lean_box(0); +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_7); +x_12 = lean_array_size(x_8); +x_13 = 0; +x_14 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__6(x_8, x_9, x_10, x_8, x_12, x_13, x_11); +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +lean_dec(x_14); +if (lean_obj_tag(x_15) == 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 = l_Lean_indentExpr(x_1); -x_29 = l_Lean_Meta_Grind_getENode___closed__2; -x_30 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_30, 0, x_29); -lean_ctor_set(x_30, 1, x_28); -x_31 = l_Lean_Meta_Grind_getENode___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_Lean_Meta_Grind_getENode___spec__1(x_32, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_25); -return x_33; +uint8_t x_16; +x_16 = 1; +return x_16; } else { -lean_object* x_34; lean_object* x_35; -lean_dec(x_1); -x_34 = lean_ctor_get(x_27, 0); -lean_inc(x_34); -lean_dec(x_27); -x_35 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_35, 0, x_34); -lean_ctor_set(x_35, 1, x_25); -return x_35; -} +lean_object* x_17; uint8_t x_18; +x_17 = lean_ctor_get(x_15, 0); +lean_inc(x_17); +lean_dec(x_15); +x_18 = lean_unbox(x_17); +lean_dec(x_17); +return x_18; } } } -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_Grind_getENode___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_11; -x_11 = l_Lean_throwError___at_Lean_Meta_Grind_getENode___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_9); -lean_dec(x_8); -lean_dec(x_7); +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, 1); +x_5 = lean_array_get_size(x_4); +x_6 = lean_ctor_get(x_1, 1); +lean_inc(x_6); +x_7 = lean_array_get_size(x_6); lean_dec(x_6); +x_8 = lean_nat_dec_eq(x_5, x_7); +lean_dec(x_7); lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_11; -} +if (x_8 == 0) +{ +uint8_t x_9; +lean_dec(x_1); +x_9 = 0; +return x_9; } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getENode___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: +else { -lean_object* x_11; -x_11 = l_Lean_Meta_Grind_getENode(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); +lean_object* x_10; uint8_t x_11; +x_10 = lean_box(0); +x_11 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3___lambda__1(x_1, x_2, x_10); return x_11; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getGeneration(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5) { _start: { -lean_object* x_11; -x_11 = l_Lean_Meta_Grind_getENode(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) +if (lean_obj_tag(x_1) == 0) { -uint8_t x_12; -x_12 = !lean_is_exclusive(x_11); -if (x_12 == 0) +uint8_t x_6; +x_6 = !lean_is_exclusive(x_1); +if (x_6 == 0) { -lean_object* x_13; lean_object* x_14; -x_13 = lean_ctor_get(x_11, 0); -x_14 = lean_ctor_get(x_13, 8); -lean_inc(x_14); +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_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___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); -lean_ctor_set(x_11, 0, x_14); -return x_11; -} -else +if (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_11, 0); -x_16 = lean_ctor_get(x_11, 1); -lean_inc(x_16); -lean_inc(x_15); -lean_dec(x_11); -x_17 = lean_ctor_get(x_15, 8); -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_dec(x_12); +lean_dec(x_5); +lean_dec(x_4); +return x_1; } else { -uint8_t x_19; -x_19 = !lean_is_exclusive(x_11); -if (x_19 == 0) -{ -return x_11; -} -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: { -lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_20 = lean_ctor_get(x_11, 0); -x_21 = lean_ctor_get(x_11, 1); +uint8_t x_18; +x_18 = !lean_is_exclusive(x_15); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; +x_19 = lean_ctor_get(x_15, 0); +x_20 = lean_ctor_get(x_15, 1); +x_21 = lean_ctor_get(x_4, 0); lean_inc(x_21); -lean_inc(x_20); -lean_dec(x_11); -x_22 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_22, 0, x_20); -lean_ctor_set(x_22, 1, x_21); -return x_22; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getGeneration___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: +x_22 = lean_ctor_get(x_19, 0); +lean_inc(x_22); +x_23 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_21, x_22); +lean_dec(x_22); +lean_dec(x_21); +if (x_23 == 0) { -lean_object* x_11; -x_11 = l_Lean_Meta_Grind_getGeneration(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_11; -} +lean_object* x_24; lean_object* x_25; lean_object* x_26; +lean_free_object(x_15); +x_24 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_19, x_20, x_4, x_5); +x_25 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_25, 0, x_24); +x_26 = lean_array_fset(x_17, x_12, x_25); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_26); +return x_1; } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isEqTrue(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_11; -x_11 = l_Lean_Meta_Grind_getENode(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_11) == 0) +else { -lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); -x_13 = lean_ctor_get(x_11, 1); -lean_inc(x_13); -lean_dec(x_11); -x_14 = l_Lean_Meta_Grind_getTrueExpr___rarg(x_5, x_6, x_7, x_8, x_9, x_13); -x_15 = !lean_is_exclusive(x_14); -if (x_15 == 0) +uint8_t x_27; +lean_inc(x_19); +x_27 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3___lambda__2(x_19, x_4, x_16); +if (x_27 == 0) { -lean_object* x_16; lean_object* x_17; uint8_t x_18; lean_object* x_19; -x_16 = lean_ctor_get(x_14, 0); -x_17 = lean_ctor_get(x_12, 2); -lean_inc(x_17); +lean_object* x_28; lean_object* x_29; lean_object* x_30; +lean_free_object(x_15); +x_28 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_19, x_20, x_4, x_5); +x_29 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_29, 0, x_28); +x_30 = lean_array_fset(x_17, x_12, x_29); lean_dec(x_12); -x_18 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_17, x_16); -lean_dec(x_16); -lean_dec(x_17); -x_19 = lean_box(x_18); -lean_ctor_set(x_14, 0, x_19); -return x_14; +lean_ctor_set(x_1, 0, x_30); +return x_1; } else { -lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; lean_object* x_24; lean_object* x_25; -x_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_ctor_get(x_12, 2); -lean_inc(x_22); -lean_dec(x_12); -x_23 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_22, x_20); +lean_object* x_31; lean_dec(x_20); -lean_dec(x_22); -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_21); -return x_25; +lean_dec(x_19); +lean_ctor_set(x_15, 1, x_5); +lean_ctor_set(x_15, 0, x_4); +x_31 = lean_array_fset(x_17, x_12, x_15); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_31); +return x_1; +} } } else { -uint8_t x_26; -x_26 = !lean_is_exclusive(x_11); -if (x_26 == 0) +lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; +x_32 = lean_ctor_get(x_15, 0); +x_33 = lean_ctor_get(x_15, 1); +lean_inc(x_33); +lean_inc(x_32); +lean_dec(x_15); +x_34 = lean_ctor_get(x_4, 0); +lean_inc(x_34); +x_35 = lean_ctor_get(x_32, 0); +lean_inc(x_35); +x_36 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_34, x_35); +lean_dec(x_35); +lean_dec(x_34); +if (x_36 == 0) { -return x_11; +lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_37 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_32, x_33, x_4, x_5); +x_38 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_38, 0, x_37); +x_39 = lean_array_fset(x_17, x_12, x_38); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_39); +return x_1; } else { -lean_object* x_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_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; -} +uint8_t x_40; +lean_inc(x_32); +x_40 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3___lambda__2(x_32, x_4, x_16); +if (x_40 == 0) +{ +lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_41 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_32, x_33, x_4, x_5); +x_42 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_42, 0, x_41); +x_43 = lean_array_fset(x_17, x_12, x_42); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_43); +return x_1; } +else +{ +lean_object* x_44; lean_object* x_45; +lean_dec(x_33); +lean_dec(x_32); +x_44 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_44, 0, x_4); +lean_ctor_set(x_44, 1, x_5); +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; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isEqTrue___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_11; -x_11 = l_Lean_Meta_Grind_isEqTrue(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_11; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isEqFalse(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_11; -x_11 = l_Lean_Meta_Grind_getENode(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) +case 1: { -lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); -x_13 = lean_ctor_get(x_11, 1); -lean_inc(x_13); -lean_dec(x_11); -x_14 = l_Lean_Meta_Grind_getFalseExpr___rarg(x_5, x_6, x_7, x_8, x_9, x_13); -x_15 = !lean_is_exclusive(x_14); -if (x_15 == 0) +uint8_t x_46; +x_46 = !lean_is_exclusive(x_15); +if (x_46 == 0) { -lean_object* x_16; lean_object* x_17; uint8_t x_18; lean_object* x_19; -x_16 = lean_ctor_get(x_14, 0); -x_17 = lean_ctor_get(x_12, 2); -lean_inc(x_17); +lean_object* x_47; size_t x_48; size_t x_49; lean_object* x_50; lean_object* x_51; +x_47 = lean_ctor_get(x_15, 0); +x_48 = lean_usize_shift_right(x_2, x_9); +x_49 = lean_usize_add(x_3, x_8); +x_50 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3(x_47, x_48, x_49, x_4, x_5); +lean_ctor_set(x_15, 0, x_50); +x_51 = lean_array_fset(x_17, x_12, x_15); lean_dec(x_12); -x_18 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_17, x_16); -lean_dec(x_16); -lean_dec(x_17); -x_19 = lean_box(x_18); -lean_ctor_set(x_14, 0, x_19); -return x_14; +lean_ctor_set(x_1, 0, x_51); +return x_1; } else { -lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; lean_object* x_24; lean_object* x_25; -x_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_ctor_get(x_12, 2); -lean_inc(x_22); +lean_object* x_52; size_t x_53; size_t x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_52 = lean_ctor_get(x_15, 0); +lean_inc(x_52); +lean_dec(x_15); +x_53 = lean_usize_shift_right(x_2, x_9); +x_54 = lean_usize_add(x_3, x_8); +x_55 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3(x_52, x_53, x_54, x_4, x_5); +x_56 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_56, 0, x_55); +x_57 = lean_array_fset(x_17, x_12, x_56); lean_dec(x_12); -x_23 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_22, x_20); -lean_dec(x_20); -lean_dec(x_22); -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_21); -return x_25; -} +lean_ctor_set(x_1, 0, x_57); +return x_1; } -else -{ -uint8_t x_26; -x_26 = !lean_is_exclusive(x_11); -if (x_26 == 0) -{ -return x_11; } -else +default: { -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_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_object* x_58; lean_object* x_59; +x_58 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_58, 0, x_4); +lean_ctor_set(x_58, 1, x_5); +x_59 = lean_array_fset(x_17, x_12, x_58); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_59); +return x_1; } } } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isEqFalse___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: +else { -lean_object* x_11; -x_11 = l_Lean_Meta_Grind_isEqFalse(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_object* x_60; size_t x_61; size_t x_62; size_t x_63; size_t x_64; lean_object* x_65; lean_object* x_66; uint8_t x_67; +x_60 = lean_ctor_get(x_1, 0); +lean_inc(x_60); +lean_dec(x_1); +x_61 = 1; +x_62 = 5; +x_63 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__2___closed__2; +x_64 = lean_usize_land(x_2, x_63); +x_65 = lean_usize_to_nat(x_64); +x_66 = lean_array_get_size(x_60); +x_67 = lean_nat_dec_lt(x_65, x_66); +lean_dec(x_66); +if (x_67 == 0) +{ +lean_object* x_68; +lean_dec(x_65); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_11; -} +x_68 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_68, 0, x_60); +return x_68; } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isEqv(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, 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 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_1, x_2); -if (x_12 == 0) -{ -lean_object* x_13; -x_13 = l_Lean_Meta_Grind_getENode(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_13) == 0) -{ -lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_14 = lean_ctor_get(x_13, 0); -lean_inc(x_14); -x_15 = lean_ctor_get(x_13, 1); -lean_inc(x_15); -lean_dec(x_13); -x_16 = l_Lean_Meta_Grind_getENode(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_15); -if (lean_obj_tag(x_16) == 0) +else { -uint8_t x_17; -x_17 = !lean_is_exclusive(x_16); -if (x_17 == 0) +lean_object* x_69; lean_object* x_70; lean_object* x_71; +x_69 = lean_array_fget(x_60, x_65); +x_70 = lean_box(0); +x_71 = lean_array_fset(x_60, x_65, x_70); +switch (lean_obj_tag(x_69)) { +case 0: { -lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; lean_object* x_22; -x_18 = lean_ctor_get(x_16, 0); -x_19 = lean_ctor_get(x_14, 2); -lean_inc(x_19); -lean_dec(x_14); -x_20 = lean_ctor_get(x_18, 2); -lean_inc(x_20); -lean_dec(x_18); -x_21 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_19, x_20); -lean_dec(x_20); -lean_dec(x_19); -x_22 = lean_box(x_21); -lean_ctor_set(x_16, 0, x_22); -return x_16; +lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; uint8_t x_77; +x_72 = lean_ctor_get(x_69, 0); +lean_inc(x_72); +x_73 = lean_ctor_get(x_69, 1); +lean_inc(x_73); +if (lean_is_exclusive(x_69)) { + lean_ctor_release(x_69, 0); + lean_ctor_release(x_69, 1); + x_74 = x_69; +} else { + lean_dec_ref(x_69); + x_74 = lean_box(0); } -else +x_75 = lean_ctor_get(x_4, 0); +lean_inc(x_75); +x_76 = lean_ctor_get(x_72, 0); +lean_inc(x_76); +x_77 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_75, x_76); +lean_dec(x_76); +lean_dec(x_75); +if (x_77 == 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; -x_23 = lean_ctor_get(x_16, 0); -x_24 = lean_ctor_get(x_16, 1); -lean_inc(x_24); -lean_inc(x_23); -lean_dec(x_16); -x_25 = lean_ctor_get(x_14, 2); -lean_inc(x_25); -lean_dec(x_14); -x_26 = lean_ctor_get(x_23, 2); -lean_inc(x_26); -lean_dec(x_23); -x_27 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_25, x_26); -lean_dec(x_26); -lean_dec(x_25); -x_28 = lean_box(x_27); -x_29 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_29, 0, x_28); -lean_ctor_set(x_29, 1, x_24); -return x_29; -} +lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; +lean_dec(x_74); +x_78 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_72, x_73, x_4, x_5); +x_79 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_79, 0, x_78); +x_80 = lean_array_fset(x_71, x_65, x_79); +lean_dec(x_65); +x_81 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_81, 0, x_80); +return x_81; } else { -uint8_t x_30; -lean_dec(x_14); -x_30 = !lean_is_exclusive(x_16); -if (x_30 == 0) +uint8_t x_82; +lean_inc(x_72); +x_82 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3___lambda__2(x_72, x_4, x_70); +if (x_82 == 0) { -return x_16; +lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; +lean_dec(x_74); +x_83 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_72, x_73, x_4, x_5); +x_84 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_84, 0, x_83); +x_85 = lean_array_fset(x_71, x_65, x_84); +lean_dec(x_65); +x_86 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_86, 0, x_85); +return x_86; } 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; +lean_object* x_87; lean_object* x_88; lean_object* x_89; +lean_dec(x_73); +lean_dec(x_72); +if (lean_is_scalar(x_74)) { + x_87 = lean_alloc_ctor(0, 2, 0); +} else { + x_87 = x_74; } +lean_ctor_set(x_87, 0, x_4); +lean_ctor_set(x_87, 1, x_5); +x_88 = lean_array_fset(x_71, x_65, x_87); +lean_dec(x_65); +x_89 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_89, 0, x_88); +return x_89; } } -else -{ -uint8_t x_34; -lean_dec(x_2); -x_34 = !lean_is_exclusive(x_13); -if (x_34 == 0) -{ -return x_13; } -else +case 1: { -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_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; +lean_object* x_90; lean_object* x_91; size_t x_92; size_t x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; +x_90 = lean_ctor_get(x_69, 0); +lean_inc(x_90); +if (lean_is_exclusive(x_69)) { + lean_ctor_release(x_69, 0); + x_91 = x_69; +} else { + lean_dec_ref(x_69); + x_91 = lean_box(0); } +x_92 = lean_usize_shift_right(x_2, x_62); +x_93 = lean_usize_add(x_3, x_61); +x_94 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3(x_90, x_92, x_93, x_4, x_5); +if (lean_is_scalar(x_91)) { + x_95 = lean_alloc_ctor(1, 1, 0); +} else { + x_95 = x_91; } +lean_ctor_set(x_95, 0, x_94); +x_96 = lean_array_fset(x_71, x_65, x_95); +lean_dec(x_65); +x_97 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_97, 0, x_96); +return x_97; } -else +default: { -uint8_t x_38; lean_object* x_39; lean_object* x_40; -lean_dec(x_2); -lean_dec(x_1); -x_38 = 1; -x_39 = lean_box(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_11); -return x_40; +lean_object* x_98; lean_object* x_99; lean_object* x_100; +x_98 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_98, 0, x_4); +lean_ctor_set(x_98, 1, x_5); +x_99 = lean_array_fset(x_71, x_65, x_98); +lean_dec(x_65); +x_100 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_100, 0, x_99); +return x_100; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isEqv___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, 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_Meta_Grind_isEqv(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -return x_12; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isRoot(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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; -x_11 = l_Lean_Meta_Grind_getENode_x3f(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); -if (lean_obj_tag(x_12) == 0) +uint8_t x_101; +x_101 = !lean_is_exclusive(x_1); +if (x_101 == 0) { -uint8_t x_13; -x_13 = !lean_is_exclusive(x_11); -if (x_13 == 0) +lean_object* x_102; lean_object* x_103; size_t x_104; uint8_t x_105; +x_102 = lean_unsigned_to_nat(0u); +x_103 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_markTheoremInstance___spec__7(x_1, x_102, x_4, x_5); +x_104 = 7; +x_105 = lean_usize_dec_le(x_104, x_3); +if (x_105 == 0) { -lean_object* x_14; uint8_t x_15; lean_object* x_16; -x_14 = lean_ctor_get(x_11, 0); -lean_dec(x_14); -x_15 = 0; -x_16 = lean_box(x_15); -lean_ctor_set(x_11, 0, x_16); -return x_11; +lean_object* x_106; lean_object* x_107; uint8_t x_108; +x_106 = l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(x_103); +x_107 = lean_unsigned_to_nat(4u); +x_108 = lean_nat_dec_lt(x_106, x_107); +lean_dec(x_106); +if (x_108 == 0) +{ +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_inc(x_109); +x_110 = lean_ctor_get(x_103, 1); +lean_inc(x_110); +lean_dec(x_103); +x_111 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__2___closed__3; +x_112 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_markTheoremInstance___spec__4(x_3, x_109, x_110, lean_box(0), x_102, x_111); +lean_dec(x_110); +lean_dec(x_109); +return x_112; } else { -lean_object* x_17; uint8_t x_18; lean_object* x_19; lean_object* x_20; -x_17 = lean_ctor_get(x_11, 1); -lean_inc(x_17); -lean_dec(x_11); -x_18 = 0; -x_19 = lean_box(x_18); -x_20 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_17); -return x_20; +return x_103; } } else { -uint8_t x_21; -x_21 = !lean_is_exclusive(x_11); -if (x_21 == 0) +return x_103; +} +} +else { -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_11, 0); -lean_dec(x_22); -x_23 = lean_ctor_get(x_12, 0); -lean_inc(x_23); -lean_dec(x_12); -x_24 = lean_ctor_get(x_23, 2); -lean_inc(x_24); -lean_dec(x_23); -x_25 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_24, x_1); -lean_dec(x_24); -x_26 = lean_box(x_25); -lean_ctor_set(x_11, 0, x_26); -return x_11; +lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; size_t x_118; uint8_t x_119; +x_113 = lean_ctor_get(x_1, 0); +x_114 = lean_ctor_get(x_1, 1); +lean_inc(x_114); +lean_inc(x_113); +lean_dec(x_1); +x_115 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_115, 0, x_113); +lean_ctor_set(x_115, 1, x_114); +x_116 = lean_unsigned_to_nat(0u); +x_117 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_markTheoremInstance___spec__7(x_115, x_116, x_4, x_5); +x_118 = 7; +x_119 = lean_usize_dec_le(x_118, x_3); +if (x_119 == 0) +{ +lean_object* x_120; lean_object* x_121; uint8_t x_122; +x_120 = l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(x_117); +x_121 = lean_unsigned_to_nat(4u); +x_122 = lean_nat_dec_lt(x_120, x_121); +lean_dec(x_120); +if (x_122 == 0) +{ +lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; +x_123 = lean_ctor_get(x_117, 0); +lean_inc(x_123); +x_124 = lean_ctor_get(x_117, 1); +lean_inc(x_124); +lean_dec(x_117); +x_125 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__2___closed__3; +x_126 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_markTheoremInstance___spec__4(x_3, x_123, x_124, lean_box(0), x_116, x_125); +lean_dec(x_124); +lean_dec(x_123); +return x_126; } else { -lean_object* x_27; lean_object* x_28; lean_object* x_29; uint8_t x_30; lean_object* x_31; lean_object* x_32; -x_27 = lean_ctor_get(x_11, 1); -lean_inc(x_27); -lean_dec(x_11); -x_28 = lean_ctor_get(x_12, 0); -lean_inc(x_28); -lean_dec(x_12); -x_29 = lean_ctor_get(x_28, 2); -lean_inc(x_29); -lean_dec(x_28); -x_30 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_29, x_1); -lean_dec(x_29); -x_31 = lean_box(x_30); -x_32 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_32, 0, x_31); -lean_ctor_set(x_32, 1, x_27); -return x_32; +return x_117; +} } +else +{ +return x_117; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isRoot___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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_PersistentHashMap_insert___at_Lean_Meta_Grind_markTheoremInstance___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_11; -x_11 = l_Lean_Meta_Grind_isRoot(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); +uint64_t x_4; lean_object* x_5; lean_object* x_6; size_t x_7; size_t x_8; uint64_t x_9; size_t x_10; size_t x_11; lean_object* x_12; +x_4 = l_Lean_Meta_Grind_instHashablePreInstance_unsafe__1(x_2); +x_5 = lean_ctor_get(x_2, 1); +lean_inc(x_5); +x_6 = lean_box(0); +x_7 = lean_array_size(x_5); +x_8 = 0; +x_9 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__2(x_5, x_6, x_5, x_7, x_8, x_4); lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_11; +x_10 = lean_uint64_to_usize(x_9); +x_11 = 1; +x_12 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3(x_1, x_10, x_11, x_2, x_3); +return x_12; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getRoot_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) { +LEAN_EXPORT uint64_t l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__10(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, uint64_t x_6) { _start: { -lean_object* x_11; lean_object* x_12; -x_11 = l_Lean_Meta_Grind_getENode_x3f(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); -if (lean_obj_tag(x_12) == 0) -{ -uint8_t x_13; -x_13 = !lean_is_exclusive(x_11); -if (x_13 == 0) +uint8_t x_7; +x_7 = lean_usize_dec_lt(x_5, x_4); +if (x_7 == 0) { -lean_object* x_14; lean_object* x_15; -x_14 = lean_ctor_get(x_11, 0); -lean_dec(x_14); -x_15 = lean_box(0); -lean_ctor_set(x_11, 0, x_15); -return x_11; +return x_6; } else { -lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_16 = lean_ctor_get(x_11, 1); -lean_inc(x_16); -lean_dec(x_11); -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_8; uint64_t x_9; uint64_t x_10; size_t x_11; size_t x_12; +x_8 = lean_array_uget(x_3, x_5); +x_9 = l_Lean_Meta_Grind_instHashablePreInstance_unsafe__2(x_8); +lean_dec(x_8); +x_10 = lean_uint64_mix_hash(x_6, x_9); +x_11 = 1; +x_12 = lean_usize_add(x_5, x_11); +x_5 = x_12; +x_6 = x_10; +goto _start; } } -else -{ -uint8_t x_19; -x_19 = !lean_is_exclusive(x_11); -if (x_19 == 0) +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__12(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, size_t x_6, lean_object* x_7) { +_start: { -lean_object* x_20; uint8_t x_21; -x_20 = lean_ctor_get(x_11, 0); -lean_dec(x_20); -x_21 = !lean_is_exclusive(x_12); -if (x_21 == 0) +uint8_t x_8; +x_8 = lean_usize_dec_lt(x_6, x_5); +if (x_8 == 0) { -lean_object* x_22; lean_object* x_23; -x_22 = lean_ctor_get(x_12, 0); -x_23 = lean_ctor_get(x_22, 2); -lean_inc(x_23); -lean_dec(x_22); -lean_ctor_set(x_12, 0, x_23); -return x_11; +lean_dec(x_3); +return x_7; } else { -lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_24 = lean_ctor_get(x_12, 0); -lean_inc(x_24); -lean_dec(x_12); -x_25 = lean_ctor_get(x_24, 2); -lean_inc(x_25); -lean_dec(x_24); -x_26 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_26, 0, x_25); -lean_ctor_set(x_11, 0, x_26); -return x_11; -} -} -else +lean_object* x_9; uint8_t x_10; +x_9 = lean_array_uget(x_4, x_6); +x_10 = !lean_is_exclusive(x_7); +if (x_10 == 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; -x_27 = lean_ctor_get(x_11, 1); -lean_inc(x_27); -lean_dec(x_11); -x_28 = lean_ctor_get(x_12, 0); -lean_inc(x_28); -if (lean_is_exclusive(x_12)) { - lean_ctor_release(x_12, 0); - x_29 = x_12; -} else { - lean_dec_ref(x_12); - x_29 = lean_box(0); -} -x_30 = lean_ctor_get(x_28, 2); -lean_inc(x_30); -lean_dec(x_28); -if (lean_is_scalar(x_29)) { - x_31 = lean_alloc_ctor(1, 1, 0); -} else { - x_31 = x_29; -} -lean_ctor_set(x_31, 0, x_30); -x_32 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_32, 0, x_31); -lean_ctor_set(x_32, 1, x_27); -return x_32; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getRoot_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, 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 = lean_ctor_get(x_7, 1); +x_12 = lean_ctor_get(x_7, 0); +lean_dec(x_12); +x_13 = lean_ctor_get(x_11, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_11, 1); +lean_inc(x_14); +x_15 = lean_ctor_get(x_11, 2); +lean_inc(x_15); +x_16 = lean_nat_dec_lt(x_14, x_15); +if (x_16 == 0) { -lean_object* x_11; -x_11 = l_Lean_Meta_Grind_getRoot_x3f(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_11; -} +lean_ctor_set(x_7, 0, x_3); +return x_7; } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getRoot(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: +else { -lean_object* x_11; -x_11 = l_Lean_Meta_Grind_getENode(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_17; +x_17 = !lean_is_exclusive(x_11); +if (x_17 == 0) { -uint8_t x_12; -x_12 = !lean_is_exclusive(x_11); -if (x_12 == 0) +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_18 = lean_ctor_get(x_11, 2); +lean_dec(x_18); +x_19 = lean_ctor_get(x_11, 1); +lean_dec(x_19); +x_20 = lean_ctor_get(x_11, 0); +lean_dec(x_20); +x_21 = lean_array_fget(x_13, x_14); +x_22 = lean_unsigned_to_nat(1u); +x_23 = lean_nat_add(x_14, x_22); +lean_dec(x_14); +lean_ctor_set(x_11, 1, x_23); +x_24 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_9, x_21); +lean_dec(x_21); +lean_dec(x_9); +if (x_24 == 0) { -lean_object* x_13; lean_object* x_14; -x_13 = lean_ctor_get(x_11, 0); -x_14 = lean_ctor_get(x_13, 2); -lean_inc(x_14); -lean_dec(x_13); -lean_ctor_set(x_11, 0, x_14); -return x_11; +lean_object* x_25; +lean_dec(x_3); +x_25 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1___closed__1; +lean_ctor_set(x_7, 0, x_25); +return x_7; } 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_ctor_get(x_15, 2); -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; +size_t x_26; size_t x_27; +lean_inc(x_3); +lean_ctor_set(x_7, 0, x_3); +x_26 = 1; +x_27 = lean_usize_add(x_6, x_26); +x_6 = x_27; +goto _start; } } else { -uint8_t x_19; -x_19 = !lean_is_exclusive(x_11); -if (x_19 == 0) +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; +lean_dec(x_11); +x_29 = lean_array_fget(x_13, x_14); +x_30 = lean_unsigned_to_nat(1u); +x_31 = lean_nat_add(x_14, x_30); +lean_dec(x_14); +x_32 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_32, 0, x_13); +lean_ctor_set(x_32, 1, x_31); +lean_ctor_set(x_32, 2, x_15); +x_33 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_9, x_29); +lean_dec(x_29); +lean_dec(x_9); +if (x_33 == 0) { -return x_11; +lean_object* x_34; +lean_dec(x_3); +x_34 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1___closed__1; +lean_ctor_set(x_7, 1, x_32); +lean_ctor_set(x_7, 0, x_34); +return x_7; } else { -lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_20 = lean_ctor_get(x_11, 0); -x_21 = lean_ctor_get(x_11, 1); -lean_inc(x_21); -lean_inc(x_20); -lean_dec(x_11); -x_22 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_22, 0, x_20); -lean_ctor_set(x_22, 1, x_21); -return x_22; -} -} -} +size_t x_35; size_t x_36; +lean_inc(x_3); +lean_ctor_set(x_7, 1, x_32); +lean_ctor_set(x_7, 0, x_3); +x_35 = 1; +x_36 = lean_usize_add(x_6, x_35); +x_6 = x_36; +goto _start; } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getRoot___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_11; -x_11 = l_Lean_Meta_Grind_getRoot(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_11; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getRootENode(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_11; -x_11 = l_Lean_Meta_Grind_getRoot(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_11) == 0) -{ -lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); -x_13 = lean_ctor_get(x_11, 1); -lean_inc(x_13); -lean_dec(x_11); -x_14 = l_Lean_Meta_Grind_getENode(x_12, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_13); -return x_14; } else { -uint8_t x_15; -x_15 = !lean_is_exclusive(x_11); -if (x_15 == 0) +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; uint8_t x_42; +x_38 = lean_ctor_get(x_7, 1); +lean_inc(x_38); +lean_dec(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); +x_41 = lean_ctor_get(x_38, 2); +lean_inc(x_41); +x_42 = lean_nat_dec_lt(x_40, x_41); +if (x_42 == 0) { -return x_11; +lean_object* x_43; +lean_dec(x_41); +lean_dec(x_40); +lean_dec(x_39); +lean_dec(x_9); +x_43 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_43, 0, x_3); +lean_ctor_set(x_43, 1, x_38); +return x_43; } else { -lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_16 = lean_ctor_get(x_11, 0); -x_17 = lean_ctor_get(x_11, 1); -lean_inc(x_17); -lean_inc(x_16); -lean_dec(x_11); -x_18 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_18, 0, x_16); -lean_ctor_set(x_18, 1, x_17); -return x_18; +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; uint8_t x_49; +if (lean_is_exclusive(x_38)) { + lean_ctor_release(x_38, 0); + lean_ctor_release(x_38, 1); + lean_ctor_release(x_38, 2); + x_44 = x_38; +} else { + lean_dec_ref(x_38); + x_44 = lean_box(0); } +x_45 = lean_array_fget(x_39, x_40); +x_46 = lean_unsigned_to_nat(1u); +x_47 = lean_nat_add(x_40, x_46); +lean_dec(x_40); +if (lean_is_scalar(x_44)) { + x_48 = lean_alloc_ctor(0, 3, 0); +} else { + x_48 = x_44; } -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getRootENode___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_11; -x_11 = l_Lean_Meta_Grind_getRootENode(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_ctor_set(x_48, 0, x_39); +lean_ctor_set(x_48, 1, x_47); +lean_ctor_set(x_48, 2, x_41); +x_49 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_9, x_45); +lean_dec(x_45); lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); +if (x_49 == 0) +{ +lean_object* x_50; lean_object* x_51; lean_dec(x_3); -lean_dec(x_2); -return x_11; +x_50 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1___closed__1; +x_51 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_51, 0, x_50); +lean_ctor_set(x_51, 1, x_48); +return x_51; +} +else +{ +lean_object* x_52; size_t x_53; size_t x_54; +lean_inc(x_3); +x_52 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_52, 0, x_3); +lean_ctor_set(x_52, 1, x_48); +x_53 = 1; +x_54 = lean_usize_add(x_6, x_53); +x_6 = x_54; +x_7 = x_52; +goto _start; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getNext(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___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_11; -x_11 = l_Lean_Meta_Grind_getENode(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) +uint8_t x_8; +x_8 = lean_usize_dec_lt(x_6, x_5); +if (x_8 == 0) { -lean_object* x_13; lean_object* x_14; -x_13 = lean_ctor_get(x_11, 0); -x_14 = lean_ctor_get(x_13, 1); -lean_inc(x_14); -lean_dec(x_13); -lean_ctor_set(x_11, 0, x_14); -return x_11; +lean_dec(x_3); +return x_7; } 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_object* x_9; uint8_t x_10; +x_9 = lean_array_uget(x_4, x_6); +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; uint8_t x_16; +x_11 = lean_ctor_get(x_7, 1); +x_12 = lean_ctor_get(x_7, 0); +lean_dec(x_12); +x_13 = lean_ctor_get(x_11, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_11, 1); +lean_inc(x_14); +x_15 = lean_ctor_get(x_11, 2); lean_inc(x_15); -lean_dec(x_11); -x_17 = lean_ctor_get(x_15, 1); -lean_inc(x_17); +x_16 = lean_nat_dec_lt(x_14, x_15); +if (x_16 == 0) +{ 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_dec(x_14); +lean_dec(x_13); +lean_dec(x_9); +lean_ctor_set(x_7, 0, x_3); +return x_7; } else { -uint8_t x_19; -x_19 = !lean_is_exclusive(x_11); -if (x_19 == 0) -{ -return x_11; -} -else +uint8_t x_17; +x_17 = !lean_is_exclusive(x_11); +if (x_17 == 0) { -lean_object* x_20; lean_object* x_21; lean_object* x_22; +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_18 = lean_ctor_get(x_11, 2); +lean_dec(x_18); +x_19 = lean_ctor_get(x_11, 1); +lean_dec(x_19); x_20 = lean_ctor_get(x_11, 0); -x_21 = lean_ctor_get(x_11, 1); -lean_inc(x_21); -lean_inc(x_20); -lean_dec(x_11); -x_22 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_22, 0, x_20); -lean_ctor_set(x_22, 1, x_21); -return x_22; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getNext___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_11; -x_11 = l_Lean_Meta_Grind_getNext(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_20); +x_21 = lean_array_fget(x_13, x_14); +x_22 = lean_unsigned_to_nat(1u); +x_23 = lean_nat_add(x_14, x_22); +lean_dec(x_14); +lean_ctor_set(x_11, 1, x_23); +x_24 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_9, x_21); +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_4); +if (x_24 == 0) +{ +lean_object* x_25; lean_dec(x_3); -lean_dec(x_2); -return x_11; -} +x_25 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1___closed__1; +lean_ctor_set(x_7, 0, x_25); +return x_7; } -LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_alreadyInternalized___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) +else { -uint8_t x_8; -lean_dec(x_4); -x_8 = 0; -return x_8; +size_t x_26; size_t x_27; +lean_inc(x_3); +lean_ctor_set(x_7, 0, x_3); +x_26 = 1; +x_27 = lean_usize_add(x_6, x_26); +x_6 = x_27; +goto _start; +} } else { -lean_object* x_9; uint8_t x_10; -x_9 = lean_array_fget(x_1, x_4); -x_10 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_5, x_9); +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; +lean_dec(x_11); +x_29 = lean_array_fget(x_13, x_14); +x_30 = lean_unsigned_to_nat(1u); +x_31 = lean_nat_add(x_14, x_30); +lean_dec(x_14); +x_32 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_32, 0, x_13); +lean_ctor_set(x_32, 1, x_31); +lean_ctor_set(x_32, 2, x_15); +x_33 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_9, x_29); +lean_dec(x_29); lean_dec(x_9); -if (x_10 == 0) +if (x_33 == 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; +lean_object* x_34; +lean_dec(x_3); +x_34 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1___closed__1; +lean_ctor_set(x_7, 1, x_32); +lean_ctor_set(x_7, 0, x_34); +return x_7; } else { -uint8_t x_14; -lean_dec(x_4); -x_14 = 1; -return x_14; +size_t x_35; size_t x_36; +lean_inc(x_3); +lean_ctor_set(x_7, 1, x_32); +lean_ctor_set(x_7, 0, x_3); +x_35 = 1; +x_36 = lean_usize_add(x_6, x_35); +x_6 = x_36; +goto _start; } } } } -LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_alreadyInternalized___spec__2(lean_object* x_1, size_t x_2, lean_object* x_3) { -_start: -{ -if (lean_obj_tag(x_1) == 0) +else { -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_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___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_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; uint8_t x_42; +x_38 = lean_ctor_get(x_7, 1); +lean_inc(x_38); +lean_dec(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); +x_41 = lean_ctor_get(x_38, 2); +lean_inc(x_41); +x_42 = lean_nat_dec_lt(x_40, x_41); +if (x_42 == 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 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(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; +lean_object* x_43; +lean_dec(x_41); +lean_dec(x_40); +lean_dec(x_39); +lean_dec(x_9); +x_43 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_43, 0, x_3); +lean_ctor_set(x_43, 1, x_38); +return x_43; } -default: +else { -uint8_t x_16; -x_16 = 0; -return x_16; +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; uint8_t x_49; +if (lean_is_exclusive(x_38)) { + lean_ctor_release(x_38, 0); + lean_ctor_release(x_38, 1); + lean_ctor_release(x_38, 2); + x_44 = x_38; +} else { + lean_dec_ref(x_38); + x_44 = lean_box(0); } +x_45 = lean_array_fget(x_39, x_40); +x_46 = lean_unsigned_to_nat(1u); +x_47 = lean_nat_add(x_40, x_46); +lean_dec(x_40); +if (lean_is_scalar(x_44)) { + x_48 = lean_alloc_ctor(0, 3, 0); +} else { + x_48 = x_44; } +lean_ctor_set(x_48, 0, x_39); +lean_ctor_set(x_48, 1, x_47); +lean_ctor_set(x_48, 2, x_41); +x_49 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_9, x_45); +lean_dec(x_45); +lean_dec(x_9); +if (x_49 == 0) +{ +lean_object* x_50; lean_object* x_51; +lean_dec(x_3); +x_50 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1___closed__1; +x_51 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_51, 0, x_50); +lean_ctor_set(x_51, 1, x_48); +return x_51; } else { -lean_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_Meta_Grind_alreadyInternalized___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_object* x_52; size_t x_53; size_t x_54; +lean_inc(x_3); +x_52 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_52, 0, x_3); +lean_ctor_set(x_52, 1, x_48); +x_53 = 1; +x_54 = lean_usize_add(x_6, x_53); +x_6 = x_54; +x_7 = x_52; +goto _start; } } } -LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_alreadyInternalized___spec__1(lean_object* x_1, lean_object* x_2) { -_start: -{ -uint64_t x_3; size_t x_4; uint8_t x_5; -x_3 = l_Lean_Meta_Grind_instHashableENodeKey_unsafe__1(x_2); -x_4 = lean_uint64_to_usize(x_3); -x_5 = l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_alreadyInternalized___spec__2(x_1, x_4, x_2); -return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_alreadyInternalized(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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 uint8_t l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_markTheoremInstance___spec__13___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_11; uint8_t x_12; -x_11 = lean_st_ref_get(x_2, x_10); -x_12 = !lean_is_exclusive(x_11); -if (x_12 == 0) +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; size_t x_12; size_t x_13; lean_object* x_14; lean_object* x_15; +x_4 = lean_ctor_get(x_1, 1); +lean_inc(x_4); +lean_dec(x_1); +x_5 = lean_array_get_size(x_4); +x_6 = lean_unsigned_to_nat(0u); +x_7 = l_Array_toSubarray___rarg(x_4, x_6, x_5); +x_8 = lean_ctor_get(x_2, 1); +x_9 = lean_box(0); +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_7); +x_12 = lean_array_size(x_8); +x_13 = 0; +x_14 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__14(x_8, x_9, x_10, x_8, x_12, x_13, x_11); +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +lean_dec(x_14); +if (lean_obj_tag(x_15) == 0) { -lean_object* x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; -x_13 = lean_ctor_get(x_11, 0); -x_14 = lean_ctor_get(x_13, 1); -lean_inc(x_14); -lean_dec(x_13); -x_15 = l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_alreadyInternalized___spec__1(x_14, x_1); -x_16 = lean_box(x_15); -lean_ctor_set(x_11, 0, x_16); -return x_11; +uint8_t x_16; +x_16 = 1; +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; -x_17 = lean_ctor_get(x_11, 0); -x_18 = lean_ctor_get(x_11, 1); -lean_inc(x_18); +lean_object* x_17; uint8_t x_18; +x_17 = lean_ctor_get(x_15, 0); lean_inc(x_17); -lean_dec(x_11); -x_19 = lean_ctor_get(x_17, 1); -lean_inc(x_19); +lean_dec(x_15); +x_18 = lean_unbox(x_17); lean_dec(x_17); -x_20 = l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_alreadyInternalized___spec__1(x_19, x_1); -x_21 = lean_box(x_20); -x_22 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_22, 0, x_21); -lean_ctor_set(x_22, 1, x_18); -return x_22; +return x_18; } } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_alreadyInternalized___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_markTheoremInstance___spec__13___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -uint8_t x_6; lean_object* x_7; -x_6 = l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_alreadyInternalized___spec__3(x_1, x_2, x_3, x_4, x_5); +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, 1); +x_5 = lean_array_get_size(x_4); +x_6 = lean_ctor_get(x_1, 1); +lean_inc(x_6); +x_7 = lean_array_get_size(x_6); +lean_dec(x_6); +x_8 = lean_nat_dec_eq(x_5, x_7); +lean_dec(x_7); lean_dec(x_5); -lean_dec(x_2); +if (x_8 == 0) +{ +uint8_t x_9; lean_dec(x_1); -x_7 = lean_box(x_6); -return x_7; -} +x_9 = 0; +return x_9; } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_alreadyInternalized___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: +else { -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_Meta_Grind_alreadyInternalized___spec__2(x_1, x_4, x_3); -lean_dec(x_3); -x_6 = lean_box(x_5); -return x_6; -} +lean_object* x_10; uint8_t x_11; +x_10 = lean_box(0); +x_11 = l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_markTheoremInstance___spec__13___lambda__1(x_1, x_2, x_10); +return x_11; } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_alreadyInternalized___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_Meta_Grind_alreadyInternalized___spec__1(x_1, x_2); -lean_dec(x_2); -x_4 = lean_box(x_3); -return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_alreadyInternalized___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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 uint8_t l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_markTheoremInstance___spec__13(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -lean_object* x_11; -x_11 = l_Lean_Meta_Grind_alreadyInternalized(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_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); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_11; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getTarget_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; -x_11 = l_Lean_Meta_Grind_getENode_x3f(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); -if (lean_obj_tag(x_12) == 0) -{ -uint8_t x_13; -x_13 = !lean_is_exclusive(x_11); -if (x_13 == 0) +if (x_7 == 0) { -lean_object* x_14; lean_object* x_15; -x_14 = lean_ctor_get(x_11, 0); -lean_dec(x_14); -x_15 = lean_box(0); -lean_ctor_set(x_11, 0, x_15); -return x_11; +uint8_t x_8; +lean_dec(x_4); +x_8 = 0; +return x_8; } else { -lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_16 = lean_ctor_get(x_11, 1); -lean_inc(x_16); +lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; +x_9 = lean_array_fget(x_1, x_4); +x_10 = lean_ctor_get(x_5, 0); +x_11 = lean_ctor_get(x_9, 0); +lean_inc(x_11); +x_12 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_10, x_11); lean_dec(x_11); -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; -} +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; +lean_dec(x_9); +x_13 = lean_unsigned_to_nat(1u); +x_14 = lean_nat_add(x_4, x_13); +lean_dec(x_4); +x_3 = lean_box(0); +x_4 = x_14; +goto _start; } else { -uint8_t x_19; -x_19 = !lean_is_exclusive(x_11); -if (x_19 == 0) +lean_object* x_16; uint8_t x_17; +x_16 = lean_box(0); +x_17 = l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_markTheoremInstance___spec__13___lambda__2(x_9, x_5, x_16); +if (x_17 == 0) { -lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_20 = lean_ctor_get(x_11, 0); -lean_dec(x_20); -x_21 = lean_ctor_get(x_12, 0); -lean_inc(x_21); -lean_dec(x_12); -x_22 = lean_ctor_get(x_21, 4); -lean_inc(x_22); -lean_dec(x_21); -lean_ctor_set(x_11, 0, x_22); -return x_11; +lean_object* x_18; lean_object* x_19; +x_18 = lean_unsigned_to_nat(1u); +x_19 = lean_nat_add(x_4, x_18); +lean_dec(x_4); +x_3 = lean_box(0); +x_4 = x_19; +goto _start; } else { -lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_23 = lean_ctor_get(x_11, 1); -lean_inc(x_23); -lean_dec(x_11); -x_24 = lean_ctor_get(x_12, 0); -lean_inc(x_24); -lean_dec(x_12); -x_25 = lean_ctor_get(x_24, 4); -lean_inc(x_25); -lean_dec(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; -} +uint8_t x_21; +lean_dec(x_4); +x_21 = 1; +return x_21; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getTarget_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, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_11; -x_11 = l_Lean_Meta_Grind_getTarget_x3f(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_11; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pushEqCore(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_EXPORT uint8_t l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_markTheoremInstance___spec__11___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; -x_14 = lean_st_ref_take(x_5, x_13); +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; size_t x_12; size_t x_13; lean_object* x_14; lean_object* x_15; +x_4 = lean_ctor_get(x_1, 1); +lean_inc(x_4); +lean_dec(x_1); +x_5 = lean_array_get_size(x_4); +x_6 = lean_unsigned_to_nat(0u); +x_7 = l_Array_toSubarray___rarg(x_4, x_6, x_5); +x_8 = lean_ctor_get(x_2, 1); +x_9 = lean_box(0); +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_7); +x_12 = lean_array_size(x_8); +x_13 = 0; +x_14 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__12(x_8, x_9, x_10, x_8, x_12, x_13, x_11); x_15 = lean_ctor_get(x_14, 0); lean_inc(x_15); -x_16 = lean_ctor_get(x_14, 1); -lean_inc(x_16); lean_dec(x_14); -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; uint8_t x_22; -x_18 = lean_ctor_get(x_15, 5); -x_19 = lean_alloc_ctor(0, 3, 1); -lean_ctor_set(x_19, 0, x_1); -lean_ctor_set(x_19, 1, x_2); -lean_ctor_set(x_19, 2, x_3); -lean_ctor_set_uint8(x_19, sizeof(void*)*3, x_4); -x_20 = lean_array_push(x_18, x_19); -lean_ctor_set(x_15, 5, x_20); -x_21 = lean_st_ref_set(x_5, x_15, x_16); -x_22 = !lean_is_exclusive(x_21); -if (x_22 == 0) +if (lean_obj_tag(x_15) == 0) { -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; +uint8_t x_16; +x_16 = 1; +return x_16; } 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 = lean_box(0); -x_27 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_27, 0, x_26); -lean_ctor_set(x_27, 1, x_25); -return x_27; +lean_object* x_17; uint8_t x_18; +x_17 = lean_ctor_get(x_15, 0); +lean_inc(x_17); +lean_dec(x_15); +x_18 = lean_unbox(x_17); +lean_dec(x_17); +return x_18; } } -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; 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; lean_object* x_50; -x_28 = lean_ctor_get(x_15, 0); -x_29 = lean_ctor_get(x_15, 1); -x_30 = lean_ctor_get(x_15, 2); -x_31 = lean_ctor_get(x_15, 3); -x_32 = lean_ctor_get(x_15, 4); -x_33 = lean_ctor_get(x_15, 5); -x_34 = lean_ctor_get_uint8(x_15, sizeof(void*)*14); -x_35 = lean_ctor_get(x_15, 6); -x_36 = lean_ctor_get(x_15, 7); -x_37 = lean_ctor_get(x_15, 8); -x_38 = lean_ctor_get(x_15, 9); -x_39 = lean_ctor_get(x_15, 10); -x_40 = lean_ctor_get(x_15, 11); -x_41 = lean_ctor_get(x_15, 12); -x_42 = lean_ctor_get(x_15, 13); -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_inc(x_33); -lean_inc(x_32); -lean_inc(x_31); -lean_inc(x_30); -lean_inc(x_29); -lean_inc(x_28); -lean_dec(x_15); -x_43 = lean_alloc_ctor(0, 3, 1); -lean_ctor_set(x_43, 0, x_1); -lean_ctor_set(x_43, 1, x_2); -lean_ctor_set(x_43, 2, x_3); -lean_ctor_set_uint8(x_43, sizeof(void*)*3, x_4); -x_44 = lean_array_push(x_33, x_43); -x_45 = lean_alloc_ctor(0, 14, 1); -lean_ctor_set(x_45, 0, x_28); -lean_ctor_set(x_45, 1, x_29); -lean_ctor_set(x_45, 2, x_30); -lean_ctor_set(x_45, 3, x_31); -lean_ctor_set(x_45, 4, x_32); -lean_ctor_set(x_45, 5, x_44); -lean_ctor_set(x_45, 6, x_35); -lean_ctor_set(x_45, 7, x_36); -lean_ctor_set(x_45, 8, x_37); -lean_ctor_set(x_45, 9, x_38); -lean_ctor_set(x_45, 10, x_39); -lean_ctor_set(x_45, 11, x_40); -lean_ctor_set(x_45, 12, x_41); -lean_ctor_set(x_45, 13, x_42); -lean_ctor_set_uint8(x_45, sizeof(void*)*14, x_34); -x_46 = lean_st_ref_set(x_5, x_45, x_16); -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); } -x_49 = lean_box(0); -if (lean_is_scalar(x_48)) { - x_50 = lean_alloc_ctor(0, 2, 0); -} else { - x_50 = x_48; +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_markTheoremInstance___spec__11___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; uint8_t x_8; +x_4 = lean_ctor_get(x_2, 1); +x_5 = lean_array_get_size(x_4); +x_6 = lean_ctor_get(x_1, 1); +lean_inc(x_6); +x_7 = lean_array_get_size(x_6); +lean_dec(x_6); +x_8 = lean_nat_dec_eq(x_5, x_7); +lean_dec(x_7); +lean_dec(x_5); +if (x_8 == 0) +{ +uint8_t x_9; +lean_dec(x_1); +x_9 = 0; +return x_9; } -lean_ctor_set(x_50, 0, x_49); -lean_ctor_set(x_50, 1, x_47); -return x_50; +else +{ +lean_object* x_10; uint8_t x_11; +x_10 = lean_box(0); +x_11 = l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_markTheoremInstance___spec__11___lambda__1(x_1, x_2, x_10); +return x_11; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pushEqCore___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_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 uint8_t l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_markTheoremInstance___spec__11(lean_object* x_1, size_t x_2, lean_object* x_3) { _start: { -uint8_t x_14; lean_object* x_15; -x_14 = lean_unbox(x_4); +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_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___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); -x_15 = l_Lean_Meta_Grind_pushEqCore(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); -lean_dec(x_12); -lean_dec(x_11); +switch (lean_obj_tag(x_10)) { +case 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); 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_12 = lean_ctor_get(x_3, 0); +x_13 = lean_ctor_get(x_11, 0); +lean_inc(x_13); +x_14 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_12, x_13); +lean_dec(x_13); +if (x_14 == 0) +{ +uint8_t x_15; +lean_dec(x_11); +x_15 = 0; return x_15; } -} -static uint64_t _init_l_Lean_Meta_Grind_hasSameType___closed__1() { -_start: +else { -uint8_t x_1; uint64_t x_2; -x_1 = 1; -x_2 = l_Lean_Meta_TransparencyMode_toUInt64(x_1); -return x_2; +lean_object* x_16; uint8_t x_17; +x_16 = lean_box(0); +x_17 = l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_markTheoremInstance___spec__11___lambda__2(x_11, x_3, x_16); +return x_17; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_hasSameType(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: +case 1: { -uint8_t x_8; -x_8 = !lean_is_exclusive(x_3); -if (x_8 == 0) +lean_object* x_18; size_t x_19; +x_18 = lean_ctor_get(x_10, 0); +lean_inc(x_18); +lean_dec(x_10); +x_19 = lean_usize_shift_right(x_2, x_5); +x_1 = x_18; +x_2 = x_19; +goto _start; +} +default: { -lean_object* x_9; uint8_t x_10; -x_9 = lean_ctor_get(x_3, 0); -x_10 = !lean_is_exclusive(x_9); -if (x_10 == 0) -{ -uint64_t x_11; uint8_t x_12; uint64_t x_13; uint64_t x_14; uint64_t x_15; uint64_t x_16; uint64_t x_17; lean_object* x_18; -x_11 = lean_ctor_get_uint64(x_3, sizeof(void*)*7); -x_12 = 1; -lean_ctor_set_uint8(x_9, 9, x_12); -x_13 = 2; -x_14 = lean_uint64_shift_right(x_11, x_13); -x_15 = lean_uint64_shift_left(x_14, x_13); -x_16 = l_Lean_Meta_Grind_hasSameType___closed__1; -x_17 = lean_uint64_lor(x_15, x_16); -lean_ctor_set_uint64(x_3, sizeof(void*)*7, x_17); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -x_18 = lean_infer_type(x_1, x_3, x_4, x_5, x_6, x_7); -if (lean_obj_tag(x_18) == 0) -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_19 = lean_ctor_get(x_18, 0); -lean_inc(x_19); -x_20 = lean_ctor_get(x_18, 1); -lean_inc(x_20); -lean_dec(x_18); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -x_21 = lean_infer_type(x_2, x_3, x_4, x_5, x_6, x_20); -if (lean_obj_tag(x_21) == 0) +uint8_t x_21; +x_21 = 0; +return x_21; +} +} +} +else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_21, 0); +lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; +x_22 = lean_ctor_get(x_1, 0); lean_inc(x_22); -x_23 = lean_ctor_get(x_21, 1); +x_23 = lean_ctor_get(x_1, 1); lean_inc(x_23); -lean_dec(x_21); -x_24 = l_Lean_Meta_isExprDefEq(x_19, x_22, x_3, x_4, x_5, x_6, x_23); -if (lean_obj_tag(x_24) == 0) -{ -uint8_t x_25; -x_25 = !lean_is_exclusive(x_24); -if (x_25 == 0) -{ -return x_24; +lean_dec(x_1); +x_24 = lean_unsigned_to_nat(0u); +x_25 = l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_markTheoremInstance___spec__13(x_22, x_23, lean_box(0), x_24, x_3); +lean_dec(x_23); +lean_dec(x_22); +return x_25; } -else +} +} +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_markTheoremInstance___spec__9(lean_object* x_1, lean_object* x_2) { +_start: { -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; +uint64_t x_3; lean_object* x_4; lean_object* x_5; size_t x_6; size_t x_7; uint64_t x_8; size_t x_9; uint8_t x_10; +x_3 = l_Lean_Meta_Grind_instHashablePreInstance_unsafe__1(x_2); +x_4 = lean_ctor_get(x_2, 1); +x_5 = lean_box(0); +x_6 = lean_array_size(x_4); +x_7 = 0; +x_8 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__10(x_4, x_5, x_4, x_6, x_7, x_3); +x_9 = lean_uint64_to_usize(x_8); +x_10 = l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_markTheoremInstance___spec__11(x_1, x_9, x_2); +return x_10; } } -else +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markTheoremInstance___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_29; -x_29 = !lean_is_exclusive(x_24); -if (x_29 == 0) +lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_12 = lean_st_ref_take(x_3, x_11); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_12, 1); +lean_inc(x_14); +lean_dec(x_12); +x_15 = !lean_is_exclusive(x_13); +if (x_15 == 0) { -return x_24; -} -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_13, 13); +x_17 = lean_box(0); +x_18 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_markTheoremInstance___spec__1(x_16, x_1, x_17); +lean_ctor_set(x_13, 13, x_18); +x_19 = lean_st_ref_set(x_3, x_13, x_14); +x_20 = !lean_is_exclusive(x_19); +if (x_20 == 0) { -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_object* x_21; uint8_t x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_19, 0); +lean_dec(x_21); +x_22 = 1; +x_23 = lean_box(x_22); +lean_ctor_set(x_19, 0, x_23); +return x_19; } else { -uint8_t x_33; +lean_object* x_24; uint8_t x_25; lean_object* x_26; lean_object* x_27; +x_24 = lean_ctor_get(x_19, 1); +lean_inc(x_24); lean_dec(x_19); -lean_dec(x_3); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -x_33 = !lean_is_exclusive(x_21); -if (x_33 == 0) -{ -return x_21; +x_25 = 1; +x_26 = lean_box(x_25); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_24); +return x_27; +} } else { -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_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; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; 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; +x_28 = lean_ctor_get(x_13, 0); +x_29 = lean_ctor_get(x_13, 1); +x_30 = lean_ctor_get(x_13, 2); +x_31 = lean_ctor_get(x_13, 3); +x_32 = lean_ctor_get(x_13, 4); +x_33 = lean_ctor_get(x_13, 5); +x_34 = lean_ctor_get_uint8(x_13, sizeof(void*)*20); +x_35 = lean_ctor_get(x_13, 6); +x_36 = lean_ctor_get(x_13, 7); +x_37 = lean_ctor_get(x_13, 8); +x_38 = lean_ctor_get(x_13, 9); +x_39 = lean_ctor_get(x_13, 10); +x_40 = lean_ctor_get(x_13, 11); +x_41 = lean_ctor_get(x_13, 12); +x_42 = lean_ctor_get(x_13, 13); +x_43 = lean_ctor_get(x_13, 14); +x_44 = lean_ctor_get(x_13, 15); +x_45 = lean_ctor_get(x_13, 16); +x_46 = lean_ctor_get(x_13, 17); +x_47 = lean_ctor_get(x_13, 18); +x_48 = lean_ctor_get(x_13, 19); +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_inc(x_34); -lean_dec(x_21); -x_36 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_36, 0, x_34); -lean_ctor_set(x_36, 1, x_35); -return x_36; +lean_inc(x_33); +lean_inc(x_32); +lean_inc(x_31); +lean_inc(x_30); +lean_inc(x_29); +lean_inc(x_28); +lean_dec(x_13); +x_49 = lean_box(0); +x_50 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_markTheoremInstance___spec__1(x_42, x_1, x_49); +x_51 = lean_alloc_ctor(0, 20, 1); +lean_ctor_set(x_51, 0, x_28); +lean_ctor_set(x_51, 1, x_29); +lean_ctor_set(x_51, 2, x_30); +lean_ctor_set(x_51, 3, x_31); +lean_ctor_set(x_51, 4, x_32); +lean_ctor_set(x_51, 5, x_33); +lean_ctor_set(x_51, 6, x_35); +lean_ctor_set(x_51, 7, x_36); +lean_ctor_set(x_51, 8, x_37); +lean_ctor_set(x_51, 9, x_38); +lean_ctor_set(x_51, 10, x_39); +lean_ctor_set(x_51, 11, x_40); +lean_ctor_set(x_51, 12, x_41); +lean_ctor_set(x_51, 13, x_50); +lean_ctor_set(x_51, 14, x_43); +lean_ctor_set(x_51, 15, x_44); +lean_ctor_set(x_51, 16, x_45); +lean_ctor_set(x_51, 17, x_46); +lean_ctor_set(x_51, 18, x_47); +lean_ctor_set(x_51, 19, x_48); +lean_ctor_set_uint8(x_51, sizeof(void*)*20, x_34); +x_52 = lean_st_ref_set(x_3, x_51, x_14); +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 = 1; +x_56 = lean_box(x_55); +if (lean_is_scalar(x_54)) { + x_57 = lean_alloc_ctor(0, 2, 0); +} else { + x_57 = x_54; +} +lean_ctor_set(x_57, 0, x_56); +lean_ctor_set(x_57, 1, x_53); +return x_57; } } -else +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markTheoremInstance(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, 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_37; -lean_dec(x_3); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_2); -x_37 = !lean_is_exclusive(x_18); -if (x_37 == 0) +lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_1); +lean_ctor_set(x_12, 1, x_2); +x_13 = lean_st_ref_get(x_3, x_11); +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) { -return x_18; +lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_15 = lean_ctor_get(x_13, 0); +x_16 = lean_ctor_get(x_13, 1); +x_17 = lean_ctor_get(x_15, 13); +lean_inc(x_17); +lean_dec(x_15); +x_18 = l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_markTheoremInstance___spec__9(x_17, x_12); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; +lean_free_object(x_13); +x_19 = lean_box(0); +x_20 = l_Lean_Meta_Grind_markTheoremInstance___lambda__1(x_12, x_19, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_16); +return x_20; } else { -lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_38 = lean_ctor_get(x_18, 0); -x_39 = lean_ctor_get(x_18, 1); -lean_inc(x_39); -lean_inc(x_38); -lean_dec(x_18); -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; -} +uint8_t x_21; lean_object* x_22; +lean_dec(x_12); +x_21 = 0; +x_22 = lean_box(x_21); +lean_ctor_set(x_13, 0, x_22); +return x_13; } } else { -uint64_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; 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; lean_object* x_59; uint64_t x_60; uint64_t x_61; uint64_t x_62; uint64_t x_63; uint64_t x_64; lean_object* x_65; -x_41 = lean_ctor_get_uint64(x_3, sizeof(void*)*7); -x_42 = lean_ctor_get_uint8(x_9, 0); -x_43 = lean_ctor_get_uint8(x_9, 1); -x_44 = lean_ctor_get_uint8(x_9, 2); -x_45 = lean_ctor_get_uint8(x_9, 3); -x_46 = lean_ctor_get_uint8(x_9, 4); -x_47 = lean_ctor_get_uint8(x_9, 5); -x_48 = lean_ctor_get_uint8(x_9, 6); -x_49 = lean_ctor_get_uint8(x_9, 7); -x_50 = lean_ctor_get_uint8(x_9, 8); -x_51 = lean_ctor_get_uint8(x_9, 10); -x_52 = lean_ctor_get_uint8(x_9, 11); -x_53 = lean_ctor_get_uint8(x_9, 12); -x_54 = lean_ctor_get_uint8(x_9, 13); -x_55 = lean_ctor_get_uint8(x_9, 14); -x_56 = lean_ctor_get_uint8(x_9, 15); -x_57 = lean_ctor_get_uint8(x_9, 16); -lean_dec(x_9); -x_58 = 1; -x_59 = lean_alloc_ctor(0, 0, 17); -lean_ctor_set_uint8(x_59, 0, x_42); -lean_ctor_set_uint8(x_59, 1, x_43); -lean_ctor_set_uint8(x_59, 2, x_44); -lean_ctor_set_uint8(x_59, 3, x_45); -lean_ctor_set_uint8(x_59, 4, x_46); -lean_ctor_set_uint8(x_59, 5, x_47); -lean_ctor_set_uint8(x_59, 6, x_48); -lean_ctor_set_uint8(x_59, 7, x_49); -lean_ctor_set_uint8(x_59, 8, x_50); -lean_ctor_set_uint8(x_59, 9, x_58); -lean_ctor_set_uint8(x_59, 10, x_51); -lean_ctor_set_uint8(x_59, 11, x_52); -lean_ctor_set_uint8(x_59, 12, x_53); -lean_ctor_set_uint8(x_59, 13, x_54); -lean_ctor_set_uint8(x_59, 14, x_55); -lean_ctor_set_uint8(x_59, 15, x_56); -lean_ctor_set_uint8(x_59, 16, x_57); -x_60 = 2; -x_61 = lean_uint64_shift_right(x_41, x_60); -x_62 = lean_uint64_shift_left(x_61, x_60); -x_63 = l_Lean_Meta_Grind_hasSameType___closed__1; -x_64 = lean_uint64_lor(x_62, x_63); -lean_ctor_set(x_3, 0, x_59); -lean_ctor_set_uint64(x_3, sizeof(void*)*7, x_64); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -x_65 = lean_infer_type(x_1, x_3, x_4, x_5, x_6, x_7); -if (lean_obj_tag(x_65) == 0) -{ -lean_object* x_66; lean_object* x_67; lean_object* x_68; -x_66 = lean_ctor_get(x_65, 0); -lean_inc(x_66); -x_67 = lean_ctor_get(x_65, 1); -lean_inc(x_67); -lean_dec(x_65); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -x_68 = lean_infer_type(x_2, x_3, x_4, x_5, x_6, x_67); -if (lean_obj_tag(x_68) == 0) -{ -lean_object* x_69; lean_object* x_70; lean_object* x_71; -x_69 = lean_ctor_get(x_68, 0); -lean_inc(x_69); -x_70 = lean_ctor_get(x_68, 1); -lean_inc(x_70); -lean_dec(x_68); -x_71 = l_Lean_Meta_isExprDefEq(x_66, x_69, x_3, x_4, x_5, x_6, x_70); -if (lean_obj_tag(x_71) == 0) +lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; +x_23 = lean_ctor_get(x_13, 0); +x_24 = lean_ctor_get(x_13, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_13); +x_25 = lean_ctor_get(x_23, 13); +lean_inc(x_25); +lean_dec(x_23); +x_26 = l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_markTheoremInstance___spec__9(x_25, x_12); +if (x_26 == 0) { -lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; -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; -} else { - lean_dec_ref(x_71); - 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_72); -lean_ctor_set(x_75, 1, x_73); -return x_75; +lean_object* x_27; lean_object* x_28; +x_27 = lean_box(0); +x_28 = l_Lean_Meta_Grind_markTheoremInstance___lambda__1(x_12, x_27, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_24); +return x_28; } else { -lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; -x_76 = lean_ctor_get(x_71, 0); -lean_inc(x_76); -x_77 = lean_ctor_get(x_71, 1); -lean_inc(x_77); -if (lean_is_exclusive(x_71)) { - lean_ctor_release(x_71, 0); - lean_ctor_release(x_71, 1); - x_78 = x_71; -} else { - lean_dec_ref(x_71); - x_78 = lean_box(0); +uint8_t x_29; lean_object* x_30; lean_object* x_31; +lean_dec(x_12); +x_29 = 0; +x_30 = lean_box(x_29); +x_31 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_31, 0, x_30); +lean_ctor_set(x_31, 1, x_24); +return x_31; } -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 +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: { -lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; -lean_dec(x_66); -lean_dec(x_3); -lean_dec(x_6); -lean_dec(x_5); +size_t x_7; size_t x_8; uint64_t x_9; uint64_t x_10; lean_object* x_11; +x_7 = lean_unbox_usize(x_4); lean_dec(x_4); -x_80 = lean_ctor_get(x_68, 0); -lean_inc(x_80); -x_81 = lean_ctor_get(x_68, 1); -lean_inc(x_81); -if (lean_is_exclusive(x_68)) { - lean_ctor_release(x_68, 0); - lean_ctor_release(x_68, 1); - x_82 = x_68; -} else { - lean_dec_ref(x_68); - x_82 = lean_box(0); +x_8 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_9 = lean_unbox_uint64(x_6); +lean_dec(x_6); +x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__2(x_1, x_2, x_3, x_7, x_8, x_9); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_11 = lean_box_uint64(x_10); +return x_11; } -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; +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +size_t x_7; size_t x_8; uint64_t x_9; uint64_t x_10; lean_object* x_11; +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_9 = lean_unbox_uint64(x_6); +lean_dec(x_6); +x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__5(x_1, x_2, x_3, x_7, x_8, x_9); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_11 = lean_box_uint64(x_10); +return x_11; } } -else +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_markTheoremInstance___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: { -lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; +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_Grind_markTheoremInstance___spec__4(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_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___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) { +_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_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__6(x_1, x_2, x_3, x_4, x_8, x_9, x_7); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___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: +{ +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_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__8(x_1, x_2, x_3, x_4, x_8, x_9, x_7); lean_dec(x_4); lean_dec(x_2); -x_84 = lean_ctor_get(x_65, 0); -lean_inc(x_84); -x_85 = lean_ctor_get(x_65, 1); -lean_inc(x_85); -if (lean_is_exclusive(x_65)) { - lean_ctor_release(x_65, 0); - lean_ctor_release(x_65, 1); - x_86 = x_65; -} else { - lean_dec_ref(x_65); - x_86 = lean_box(0); +lean_dec(x_1); +return x_10; } -if (lean_is_scalar(x_86)) { - x_87 = lean_alloc_ctor(1, 2, 0); -} else { - x_87 = x_86; } -lean_ctor_set(x_87, 0, x_84); -lean_ctor_set(x_87, 1, x_85); -return x_87; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_markTheoremInstance___spec__7___lambda__1___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_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_markTheoremInstance___spec__7___lambda__1(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +x_5 = lean_box(x_4); +return x_5; +} } +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_markTheoremInstance___spec__7___lambda__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 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_markTheoremInstance___spec__7___lambda__2(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +x_5 = lean_box(x_4); +return x_5; } } -else +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -lean_object* x_88; uint64_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; uint8_t x_97; uint8_t 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; uint8_t x_113; uint8_t x_114; lean_object* x_115; uint8_t x_116; lean_object* x_117; uint64_t x_118; uint64_t x_119; uint64_t x_120; uint64_t x_121; uint64_t x_122; lean_object* x_123; lean_object* x_124; -x_88 = lean_ctor_get(x_3, 0); -x_89 = lean_ctor_get_uint64(x_3, sizeof(void*)*7); -x_90 = lean_ctor_get_uint8(x_3, sizeof(void*)*7 + 8); -x_91 = lean_ctor_get(x_3, 1); -x_92 = lean_ctor_get(x_3, 2); -x_93 = lean_ctor_get(x_3, 3); -x_94 = lean_ctor_get(x_3, 4); -x_95 = lean_ctor_get(x_3, 5); -x_96 = lean_ctor_get(x_3, 6); -x_97 = lean_ctor_get_uint8(x_3, sizeof(void*)*7 + 9); -x_98 = lean_ctor_get_uint8(x_3, sizeof(void*)*7 + 10); -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_88); +uint8_t x_4; lean_object* x_5; +x_4 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3___lambda__1(x_1, x_2, x_3); lean_dec(x_3); -x_99 = lean_ctor_get_uint8(x_88, 0); -x_100 = lean_ctor_get_uint8(x_88, 1); -x_101 = lean_ctor_get_uint8(x_88, 2); -x_102 = lean_ctor_get_uint8(x_88, 3); -x_103 = lean_ctor_get_uint8(x_88, 4); -x_104 = lean_ctor_get_uint8(x_88, 5); -x_105 = lean_ctor_get_uint8(x_88, 6); -x_106 = lean_ctor_get_uint8(x_88, 7); -x_107 = lean_ctor_get_uint8(x_88, 8); -x_108 = lean_ctor_get_uint8(x_88, 10); -x_109 = lean_ctor_get_uint8(x_88, 11); -x_110 = lean_ctor_get_uint8(x_88, 12); -x_111 = lean_ctor_get_uint8(x_88, 13); -x_112 = lean_ctor_get_uint8(x_88, 14); -x_113 = lean_ctor_get_uint8(x_88, 15); -x_114 = lean_ctor_get_uint8(x_88, 16); -if (lean_is_exclusive(x_88)) { - x_115 = x_88; -} else { - lean_dec_ref(x_88); - x_115 = lean_box(0); +lean_dec(x_2); +x_5 = lean_box(x_4); +return x_5; } -x_116 = 1; -if (lean_is_scalar(x_115)) { - x_117 = lean_alloc_ctor(0, 0, 17); -} else { - x_117 = x_115; } -lean_ctor_set_uint8(x_117, 0, x_99); -lean_ctor_set_uint8(x_117, 1, x_100); -lean_ctor_set_uint8(x_117, 2, x_101); -lean_ctor_set_uint8(x_117, 3, x_102); -lean_ctor_set_uint8(x_117, 4, x_103); -lean_ctor_set_uint8(x_117, 5, x_104); -lean_ctor_set_uint8(x_117, 6, x_105); -lean_ctor_set_uint8(x_117, 7, x_106); -lean_ctor_set_uint8(x_117, 8, x_107); -lean_ctor_set_uint8(x_117, 9, x_116); -lean_ctor_set_uint8(x_117, 10, x_108); -lean_ctor_set_uint8(x_117, 11, x_109); -lean_ctor_set_uint8(x_117, 12, x_110); -lean_ctor_set_uint8(x_117, 13, x_111); -lean_ctor_set_uint8(x_117, 14, x_112); -lean_ctor_set_uint8(x_117, 15, x_113); -lean_ctor_set_uint8(x_117, 16, x_114); -x_118 = 2; -x_119 = lean_uint64_shift_right(x_89, x_118); -x_120 = lean_uint64_shift_left(x_119, x_118); -x_121 = l_Lean_Meta_Grind_hasSameType___closed__1; -x_122 = lean_uint64_lor(x_120, x_121); -x_123 = lean_alloc_ctor(0, 7, 11); -lean_ctor_set(x_123, 0, x_117); -lean_ctor_set(x_123, 1, x_91); -lean_ctor_set(x_123, 2, x_92); -lean_ctor_set(x_123, 3, x_93); -lean_ctor_set(x_123, 4, x_94); -lean_ctor_set(x_123, 5, x_95); -lean_ctor_set(x_123, 6, x_96); -lean_ctor_set_uint64(x_123, sizeof(void*)*7, x_122); -lean_ctor_set_uint8(x_123, sizeof(void*)*7 + 8, x_90); -lean_ctor_set_uint8(x_123, sizeof(void*)*7 + 9, x_97); -lean_ctor_set_uint8(x_123, sizeof(void*)*7 + 10, x_98); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_123); -x_124 = lean_infer_type(x_1, x_123, x_4, x_5, x_6, x_7); -if (lean_obj_tag(x_124) == 0) -{ -lean_object* x_125; lean_object* x_126; lean_object* x_127; -x_125 = lean_ctor_get(x_124, 0); -lean_inc(x_125); -x_126 = lean_ctor_get(x_124, 1); -lean_inc(x_126); -lean_dec(x_124); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_123); -x_127 = lean_infer_type(x_2, x_123, x_4, x_5, x_6, x_126); -if (lean_obj_tag(x_127) == 0) -{ -lean_object* x_128; lean_object* x_129; lean_object* x_130; -x_128 = lean_ctor_get(x_127, 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_Meta_isExprDefEq(x_125, x_128, x_123, x_4, x_5, x_6, x_129); -if (lean_obj_tag(x_130) == 0) +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; -x_131 = lean_ctor_get(x_130, 0); -lean_inc(x_131); -x_132 = lean_ctor_get(x_130, 1); -lean_inc(x_132); -if (lean_is_exclusive(x_130)) { - lean_ctor_release(x_130, 0); - lean_ctor_release(x_130, 1); - x_133 = x_130; -} else { - lean_dec_ref(x_130); - x_133 = lean_box(0); -} -if (lean_is_scalar(x_133)) { - x_134 = lean_alloc_ctor(0, 2, 0); -} else { - x_134 = x_133; +uint8_t x_4; lean_object* x_5; +x_4 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3___lambda__2(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +x_5 = lean_box(x_4); +return x_5; } -lean_ctor_set(x_134, 0, x_131); -lean_ctor_set(x_134, 1, x_132); -return x_134; } -else +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___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_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; -x_135 = lean_ctor_get(x_130, 0); -lean_inc(x_135); -x_136 = lean_ctor_get(x_130, 1); -lean_inc(x_136); -if (lean_is_exclusive(x_130)) { - lean_ctor_release(x_130, 0); - lean_ctor_release(x_130, 1); - x_137 = x_130; -} else { - lean_dec_ref(x_130); - x_137 = lean_box(0); +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_Grind_markTheoremInstance___spec__3(x_1, x_6, x_7, x_4, x_5); +return x_8; } -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_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___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: +{ +size_t x_7; size_t x_8; uint64_t x_9; uint64_t x_10; lean_object* x_11; +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_9 = lean_unbox_uint64(x_6); +lean_dec(x_6); +x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__10(x_1, x_2, x_3, x_7, x_8, x_9); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_11 = lean_box_uint64(x_10); +return x_11; } } -else +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__12___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: { -lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; -lean_dec(x_125); -lean_dec(x_123); -lean_dec(x_6); +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_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__12(x_1, x_2, x_3, x_4, x_8, x_9, x_7); lean_dec(x_4); -x_139 = lean_ctor_get(x_127, 0); -lean_inc(x_139); -x_140 = lean_ctor_get(x_127, 1); -lean_inc(x_140); -if (lean_is_exclusive(x_127)) { - lean_ctor_release(x_127, 0); - lean_ctor_release(x_127, 1); - x_141 = x_127; -} else { - lean_dec_ref(x_127); - 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(x_142, 0, x_139); -lean_ctor_set(x_142, 1, x_140); -return x_142; +lean_dec(x_2); +lean_dec(x_1); +return x_10; } } -else +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___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) { +_start: { -lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; -lean_dec(x_123); -lean_dec(x_6); +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_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__14(x_1, x_2, x_3, x_4, x_8, x_9, x_7); lean_dec(x_4); lean_dec(x_2); -x_143 = lean_ctor_get(x_124, 0); -lean_inc(x_143); -x_144 = lean_ctor_get(x_124, 1); -lean_inc(x_144); -if (lean_is_exclusive(x_124)) { - lean_ctor_release(x_124, 0); - lean_ctor_release(x_124, 1); - x_145 = x_124; -} else { - lean_dec_ref(x_124); - x_145 = lean_box(0); -} -if (lean_is_scalar(x_145)) { - x_146 = lean_alloc_ctor(1, 2, 0); -} else { - x_146 = x_145; -} -lean_ctor_set(x_146, 0, x_143); -lean_ctor_set(x_146, 1, x_144); -return x_146; -} -} +lean_dec(x_1); +return x_10; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pushEqHEq(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_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_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_markTheoremInstance___spec__13___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_13; -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_2); -lean_inc(x_1); -x_13 = l_Lean_Meta_Grind_hasSameType(x_1, x_2, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_13) == 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_16; uint8_t x_17; lean_object* x_18; -x_16 = lean_ctor_get(x_13, 1); -lean_inc(x_16); -lean_dec(x_13); -x_17 = 1; -x_18 = l_Lean_Meta_Grind_pushEqCore(x_1, x_2, x_3, x_17, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_16); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -return x_18; -} -else -{ -lean_object* x_19; uint8_t x_20; lean_object* x_21; -x_19 = lean_ctor_get(x_13, 1); -lean_inc(x_19); -lean_dec(x_13); -x_20 = 0; -x_21 = l_Lean_Meta_Grind_pushEqCore(x_1, x_2, x_3, x_20, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_19); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -return x_21; -} -} -else -{ -uint8_t x_22; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); +uint8_t x_4; lean_object* x_5; +x_4 = l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_markTheoremInstance___spec__13___lambda__1(x_1, x_2, x_3); lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -x_22 = !lean_is_exclusive(x_13); -if (x_22 == 0) -{ -return x_13; -} -else -{ -lean_object* x_23; lean_object* x_24; lean_object* x_25; -x_23 = lean_ctor_get(x_13, 0); -x_24 = lean_ctor_get(x_13, 1); -lean_inc(x_24); -lean_inc(x_23); -lean_dec(x_13); -x_25 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_25, 0, x_23); -lean_ctor_set(x_25, 1, x_24); -return x_25; -} -} +x_5 = lean_box(x_4); +return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pushEqHEq___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_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_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_markTheoremInstance___spec__13___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_13; -x_13 = l_Lean_Meta_Grind_pushEqHEq(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_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -return x_13; +uint8_t x_4; lean_object* x_5; +x_4 = l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_markTheoremInstance___spec__13___lambda__2(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +x_5 = lean_box(x_4); +return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pushEq(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_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_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_markTheoremInstance___spec__13___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_13; lean_object* x_14; -x_13 = 0; -x_14 = l_Lean_Meta_Grind_pushEqCore(x_1, x_2, x_3, x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -return x_14; +uint8_t x_6; lean_object* x_7; +x_6 = l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_markTheoremInstance___spec__13(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_5); +lean_dec(x_2); +lean_dec(x_1); +x_7 = lean_box(x_6); +return x_7; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pushEq___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_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_PersistentHashMap_containsAux___at_Lean_Meta_Grind_markTheoremInstance___spec__11___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_13; -x_13 = l_Lean_Meta_Grind_pushEq(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -return x_13; +uint8_t x_4; lean_object* x_5; +x_4 = l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_markTheoremInstance___spec__11___lambda__1(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +x_5 = lean_box(x_4); +return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pushHEq(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_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_PersistentHashMap_containsAux___at_Lean_Meta_Grind_markTheoremInstance___spec__11___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -uint8_t x_13; lean_object* x_14; -x_13 = 1; -x_14 = l_Lean_Meta_Grind_pushEqCore(x_1, x_2, x_3, x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -return x_14; +uint8_t x_4; lean_object* x_5; +x_4 = l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_markTheoremInstance___spec__11___lambda__2(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +x_5 = lean_box(x_4); +return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pushHEq___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_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_PersistentHashMap_containsAux___at_Lean_Meta_Grind_markTheoremInstance___spec__11___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_13; -x_13 = l_Lean_Meta_Grind_pushHEq(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -return x_13; +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_Meta_Grind_markTheoremInstance___spec__11(x_1, x_4, x_3); +lean_dec(x_3); +x_6 = lean_box(x_5); +return x_6; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pushEqTrue(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_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_PersistentHashMap_contains___at_Lean_Meta_Grind_markTheoremInstance___spec__9___boxed(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; -x_12 = l_Lean_Meta_Grind_getTrueExpr___rarg(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_ctor_get(x_12, 1); -lean_inc(x_14); -lean_dec(x_12); -x_15 = 0; -x_16 = l_Lean_Meta_Grind_pushEqCore(x_1, x_13, x_2, x_15, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_14); -return x_16; +uint8_t x_3; lean_object* x_4; +x_3 = l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_markTheoremInstance___spec__9(x_1, x_2); +lean_dec(x_2); +x_4 = lean_box(x_3); +return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pushEqTrue___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_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_Grind_markTheoremInstance___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_Meta_Grind_pushEqTrue(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_Meta_Grind_markTheoremInstance___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -11057,29 +11500,15 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); +lean_dec(x_2); return x_12; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pushEqFalse(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, 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; lean_object* x_16; -x_12 = l_Lean_Meta_Grind_getFalseExpr___rarg(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_ctor_get(x_12, 1); -lean_inc(x_14); -lean_dec(x_12); -x_15 = 0; -x_16 = l_Lean_Meta_Grind_pushEqCore(x_1, x_13, x_2, x_15, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_14); -return x_16; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pushEqFalse___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_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_Grind_markTheoremInstance___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, 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_Meta_Grind_pushEqFalse(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_Meta_Grind_markTheoremInstance(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -11091,814 +11520,5589 @@ lean_dec(x_3); return x_12; } } -LEAN_EXPORT lean_object* l_Lean_RBNode_ins___at_Lean_Meta_Grind_registerParent___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addNewFact(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -if (lean_obj_tag(x_1) == 0) +lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_13 = lean_st_ref_take(x_4, x_12); +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +x_16 = !lean_is_exclusive(x_14); +if (x_16 == 0) { -lean_object* x_4; uint8_t x_5; lean_object* x_6; -x_4 = lean_box(0); -x_5 = 0; -x_6 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_6, 0, x_4); -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_uint8(x_6, sizeof(void*)*4, x_5); -return x_6; +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_17 = lean_ctor_get(x_14, 14); +x_18 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_18, 0, x_1); +lean_ctor_set(x_18, 1, x_2); +lean_ctor_set(x_18, 2, x_3); +x_19 = l_Std_Queue_enqueue___rarg(x_18, x_17); +lean_ctor_set(x_14, 14, x_19); +x_20 = lean_st_ref_set(x_4, x_14, x_15); +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); +lean_dec(x_22); +x_23 = lean_box(0); +lean_ctor_set(x_20, 0, x_23); +return x_20; } else { -uint8_t x_7; -x_7 = lean_ctor_get_uint8(x_1, sizeof(void*)*4); -if (x_7 == 0) -{ -uint8_t x_8; -x_8 = !lean_is_exclusive(x_1); -if (x_8 == 0) +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_20, 1); +lean_inc(x_24); +lean_dec(x_20); +x_25 = lean_box(0); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set(x_26, 1, x_24); +return x_26; +} +} +else { -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, 0); -x_10 = lean_ctor_get(x_1, 1); -x_11 = lean_ctor_get(x_1, 2); -x_12 = lean_ctor_get(x_1, 3); -x_13 = l_Lean_Expr_quickComp(x_2, x_10); -switch (x_13) { -case 0: -{ -lean_object* x_14; uint8_t x_15; -x_14 = l_Lean_RBNode_ins___at_Lean_Meta_Grind_registerParent___spec__2(x_9, x_2, x_3); -x_15 = 0; -lean_ctor_set(x_1, 0, x_14); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_15); -return x_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; 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; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; 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_27 = lean_ctor_get(x_14, 0); +x_28 = lean_ctor_get(x_14, 1); +x_29 = lean_ctor_get(x_14, 2); +x_30 = lean_ctor_get(x_14, 3); +x_31 = lean_ctor_get(x_14, 4); +x_32 = lean_ctor_get(x_14, 5); +x_33 = lean_ctor_get_uint8(x_14, sizeof(void*)*20); +x_34 = lean_ctor_get(x_14, 6); +x_35 = lean_ctor_get(x_14, 7); +x_36 = lean_ctor_get(x_14, 8); +x_37 = lean_ctor_get(x_14, 9); +x_38 = lean_ctor_get(x_14, 10); +x_39 = lean_ctor_get(x_14, 11); +x_40 = lean_ctor_get(x_14, 12); +x_41 = lean_ctor_get(x_14, 13); +x_42 = lean_ctor_get(x_14, 14); +x_43 = lean_ctor_get(x_14, 15); +x_44 = lean_ctor_get(x_14, 16); +x_45 = lean_ctor_get(x_14, 17); +x_46 = lean_ctor_get(x_14, 18); +x_47 = lean_ctor_get(x_14, 19); +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_inc(x_34); +lean_inc(x_32); +lean_inc(x_31); +lean_inc(x_30); +lean_inc(x_29); +lean_inc(x_28); +lean_inc(x_27); +lean_dec(x_14); +x_48 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_48, 0, x_1); +lean_ctor_set(x_48, 1, x_2); +lean_ctor_set(x_48, 2, x_3); +x_49 = l_Std_Queue_enqueue___rarg(x_48, x_42); +x_50 = lean_alloc_ctor(0, 20, 1); +lean_ctor_set(x_50, 0, x_27); +lean_ctor_set(x_50, 1, x_28); +lean_ctor_set(x_50, 2, x_29); +lean_ctor_set(x_50, 3, x_30); +lean_ctor_set(x_50, 4, x_31); +lean_ctor_set(x_50, 5, x_32); +lean_ctor_set(x_50, 6, x_34); +lean_ctor_set(x_50, 7, x_35); +lean_ctor_set(x_50, 8, x_36); +lean_ctor_set(x_50, 9, x_37); +lean_ctor_set(x_50, 10, x_38); +lean_ctor_set(x_50, 11, x_39); +lean_ctor_set(x_50, 12, x_40); +lean_ctor_set(x_50, 13, x_41); +lean_ctor_set(x_50, 14, x_49); +lean_ctor_set(x_50, 15, x_43); +lean_ctor_set(x_50, 16, x_44); +lean_ctor_set(x_50, 17, x_45); +lean_ctor_set(x_50, 18, x_46); +lean_ctor_set(x_50, 19, x_47); +lean_ctor_set_uint8(x_50, sizeof(void*)*20, x_33); +x_51 = lean_st_ref_set(x_4, x_50, x_15); +x_52 = lean_ctor_get(x_51, 1); +lean_inc(x_52); +if (lean_is_exclusive(x_51)) { + lean_ctor_release(x_51, 0); + lean_ctor_release(x_51, 1); + x_53 = x_51; +} else { + lean_dec_ref(x_51); + x_53 = lean_box(0); } -case 1: +x_54 = lean_box(0); +if (lean_is_scalar(x_53)) { + x_55 = lean_alloc_ctor(0, 2, 0); +} else { + x_55 = x_53; +} +lean_ctor_set(x_55, 0, x_54); +lean_ctor_set(x_55, 1, x_52); +return x_55; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addNewFact___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, 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_16; +lean_object* x_13; +x_13 = l_Lean_Meta_Grind_addNewFact(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); lean_dec(x_11); lean_dec(x_10); -x_16 = 0; -lean_ctor_set(x_1, 2, x_3); -lean_ctor_set(x_1, 1, x_2); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_16); -return x_1; -} -default: -{ -lean_object* x_17; uint8_t x_18; -x_17 = l_Lean_RBNode_ins___at_Lean_Meta_Grind_registerParent___spec__2(x_12, x_2, x_3); -x_18 = 0; -lean_ctor_set(x_1, 3, x_17); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_18); -return x_1; -} +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; } } -else +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addTheoremInstance(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, 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; uint8_t x_18; +x_13 = l_Lean_Meta_Grind_addNewFact(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_14 = lean_ctor_get(x_13, 1); +lean_inc(x_14); +lean_dec(x_13); +x_15 = lean_st_ref_take(x_4, x_14); +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_15, 1); +lean_inc(x_17); +lean_dec(x_15); +x_18 = !lean_is_exclusive(x_16); +if (x_18 == 0) { lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; -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_inc(x_22); -lean_inc(x_21); -lean_inc(x_20); -lean_inc(x_19); -lean_dec(x_1); -x_23 = l_Lean_Expr_quickComp(x_2, x_20); -switch (x_23) { -case 0: +x_19 = lean_ctor_get(x_16, 11); +x_20 = lean_unsigned_to_nat(1u); +x_21 = lean_nat_add(x_19, x_20); +lean_dec(x_19); +lean_ctor_set(x_16, 11, x_21); +x_22 = lean_st_ref_set(x_4, x_16, x_17); +x_23 = !lean_is_exclusive(x_22); +if (x_23 == 0) { -lean_object* x_24; uint8_t x_25; lean_object* x_26; -x_24 = l_Lean_RBNode_ins___at_Lean_Meta_Grind_registerParent___spec__2(x_19, x_2, x_3); -x_25 = 0; -x_26 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_26, 0, x_24); -lean_ctor_set(x_26, 1, x_20); -lean_ctor_set(x_26, 2, x_21); -lean_ctor_set(x_26, 3, x_22); -lean_ctor_set_uint8(x_26, sizeof(void*)*4, x_25); -return x_26; +lean_object* x_24; lean_object* x_25; +x_24 = lean_ctor_get(x_22, 0); +lean_dec(x_24); +x_25 = lean_box(0); +lean_ctor_set(x_22, 0, x_25); +return x_22; } -case 1: +else { -uint8_t x_27; lean_object* x_28; -lean_dec(x_21); -lean_dec(x_20); -x_27 = 0; -x_28 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_28, 0, x_19); -lean_ctor_set(x_28, 1, x_2); -lean_ctor_set(x_28, 2, x_3); -lean_ctor_set(x_28, 3, x_22); -lean_ctor_set_uint8(x_28, sizeof(void*)*4, x_27); +lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_26 = lean_ctor_get(x_22, 1); +lean_inc(x_26); +lean_dec(x_22); +x_27 = lean_box(0); +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_27); +lean_ctor_set(x_28, 1, x_26); return x_28; } -default: +} +else { -lean_object* x_29; uint8_t x_30; lean_object* x_31; -x_29 = l_Lean_RBNode_ins___at_Lean_Meta_Grind_registerParent___spec__2(x_22, x_2, x_3); -x_30 = 0; -x_31 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_31, 0, x_19); -lean_ctor_set(x_31, 1, x_20); -lean_ctor_set(x_31, 2, x_21); -lean_ctor_set(x_31, 3, x_29); -lean_ctor_set_uint8(x_31, sizeof(void*)*4, x_30); -return x_31; +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; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_29 = lean_ctor_get(x_16, 0); +x_30 = lean_ctor_get(x_16, 1); +x_31 = lean_ctor_get(x_16, 2); +x_32 = lean_ctor_get(x_16, 3); +x_33 = lean_ctor_get(x_16, 4); +x_34 = lean_ctor_get(x_16, 5); +x_35 = lean_ctor_get_uint8(x_16, sizeof(void*)*20); +x_36 = lean_ctor_get(x_16, 6); +x_37 = lean_ctor_get(x_16, 7); +x_38 = lean_ctor_get(x_16, 8); +x_39 = lean_ctor_get(x_16, 9); +x_40 = lean_ctor_get(x_16, 10); +x_41 = lean_ctor_get(x_16, 11); +x_42 = lean_ctor_get(x_16, 12); +x_43 = lean_ctor_get(x_16, 13); +x_44 = lean_ctor_get(x_16, 14); +x_45 = lean_ctor_get(x_16, 15); +x_46 = lean_ctor_get(x_16, 16); +x_47 = lean_ctor_get(x_16, 17); +x_48 = lean_ctor_get(x_16, 18); +x_49 = lean_ctor_get(x_16, 19); +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_34); +lean_inc(x_33); +lean_inc(x_32); +lean_inc(x_31); +lean_inc(x_30); +lean_inc(x_29); +lean_dec(x_16); +x_50 = lean_unsigned_to_nat(1u); +x_51 = lean_nat_add(x_41, x_50); +lean_dec(x_41); +x_52 = lean_alloc_ctor(0, 20, 1); +lean_ctor_set(x_52, 0, x_29); +lean_ctor_set(x_52, 1, x_30); +lean_ctor_set(x_52, 2, x_31); +lean_ctor_set(x_52, 3, x_32); +lean_ctor_set(x_52, 4, x_33); +lean_ctor_set(x_52, 5, x_34); +lean_ctor_set(x_52, 6, x_36); +lean_ctor_set(x_52, 7, x_37); +lean_ctor_set(x_52, 8, x_38); +lean_ctor_set(x_52, 9, x_39); +lean_ctor_set(x_52, 10, x_40); +lean_ctor_set(x_52, 11, x_51); +lean_ctor_set(x_52, 12, x_42); +lean_ctor_set(x_52, 13, x_43); +lean_ctor_set(x_52, 14, x_44); +lean_ctor_set(x_52, 15, x_45); +lean_ctor_set(x_52, 16, x_46); +lean_ctor_set(x_52, 17, x_47); +lean_ctor_set(x_52, 18, x_48); +lean_ctor_set(x_52, 19, x_49); +lean_ctor_set_uint8(x_52, sizeof(void*)*20, x_35); +x_53 = lean_st_ref_set(x_4, x_52, x_17); +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 -{ -uint8_t x_32; -x_32 = !lean_is_exclusive(x_1); -if (x_32 == 0) +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addTheoremInstance___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, 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_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; -x_33 = lean_ctor_get(x_1, 0); -x_34 = lean_ctor_get(x_1, 1); -x_35 = lean_ctor_get(x_1, 2); -x_36 = lean_ctor_get(x_1, 3); -x_37 = l_Lean_Expr_quickComp(x_2, x_34); -switch (x_37) { -case 0: +lean_object* x_13; +x_13 = l_Lean_Meta_Grind_addTheoremInstance(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_checkMaxInstancesExceeded(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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_38; uint8_t x_39; -x_38 = l_Lean_RBNode_ins___at_Lean_Meta_Grind_registerParent___spec__2(x_33, x_2, x_3); -x_39 = lean_ctor_get_uint8(x_38, sizeof(void*)*4); -if (x_39 == 0) +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_10 = lean_st_ref_get(x_1, 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_Meta_Grind_getConfig___rarg(x_3, x_4, x_5, x_6, x_7, x_8, x_12); +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) { -lean_object* x_40; -x_40 = lean_ctor_get(x_38, 0); +lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; lean_object* x_19; +x_15 = lean_ctor_get(x_13, 0); +x_16 = lean_ctor_get(x_15, 3); +lean_inc(x_16); +lean_dec(x_15); +x_17 = lean_ctor_get(x_11, 11); +lean_inc(x_17); +lean_dec(x_11); +x_18 = lean_nat_dec_le(x_16, x_17); +lean_dec(x_17); +lean_dec(x_16); +x_19 = lean_box(x_18); +lean_ctor_set(x_13, 0, x_19); +return x_13; +} +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; lean_object* x_25; lean_object* x_26; +x_20 = lean_ctor_get(x_13, 0); +x_21 = lean_ctor_get(x_13, 1); +lean_inc(x_21); +lean_inc(x_20); +lean_dec(x_13); +x_22 = lean_ctor_get(x_20, 3); +lean_inc(x_22); +lean_dec(x_20); +x_23 = lean_ctor_get(x_11, 11); +lean_inc(x_23); +lean_dec(x_11); +x_24 = lean_nat_dec_le(x_22, x_23); +lean_dec(x_23); +lean_dec(x_22); +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_21); +return x_26; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_checkMaxInstancesExceeded___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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_Grind_checkMaxInstancesExceeded(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); +lean_dec(x_1); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_checkMaxCaseSplit(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_10 = lean_st_ref_get(x_1, 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_Meta_Grind_getConfig___rarg(x_3, x_4, x_5, x_6, x_7, x_8, x_12); +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; lean_object* x_19; +x_15 = lean_ctor_get(x_13, 0); +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +lean_dec(x_15); +x_17 = lean_ctor_get(x_11, 17); +lean_inc(x_17); +lean_dec(x_11); +x_18 = lean_nat_dec_le(x_16, x_17); +lean_dec(x_17); +lean_dec(x_16); +x_19 = lean_box(x_18); +lean_ctor_set(x_13, 0, x_19); +return x_13; +} +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; lean_object* x_25; lean_object* x_26; +x_20 = lean_ctor_get(x_13, 0); +x_21 = lean_ctor_get(x_13, 1); +lean_inc(x_21); +lean_inc(x_20); +lean_dec(x_13); +x_22 = lean_ctor_get(x_20, 0); +lean_inc(x_22); +lean_dec(x_20); +x_23 = lean_ctor_get(x_11, 17); +lean_inc(x_23); +lean_dec(x_11); +x_24 = lean_nat_dec_le(x_22, x_23); +lean_dec(x_23); +lean_dec(x_22); +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_21); +return x_26; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_checkMaxCaseSplit___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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_Grind_checkMaxCaseSplit(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); +lean_dec(x_1); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_checkMaxEmatchExceeded(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_10 = lean_st_ref_get(x_1, 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_Meta_Grind_getConfig___rarg(x_3, x_4, x_5, x_6, x_7, x_8, x_12); +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; lean_object* x_19; +x_15 = lean_ctor_get(x_13, 0); +x_16 = lean_ctor_get(x_15, 1); +lean_inc(x_16); +lean_dec(x_15); +x_17 = lean_ctor_get(x_11, 12); +lean_inc(x_17); +lean_dec(x_11); +x_18 = lean_nat_dec_le(x_16, x_17); +lean_dec(x_17); +lean_dec(x_16); +x_19 = lean_box(x_18); +lean_ctor_set(x_13, 0, x_19); +return x_13; +} +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; lean_object* x_25; lean_object* x_26; +x_20 = lean_ctor_get(x_13, 0); +x_21 = lean_ctor_get(x_13, 1); +lean_inc(x_21); +lean_inc(x_20); +lean_dec(x_13); +x_22 = lean_ctor_get(x_20, 1); +lean_inc(x_22); +lean_dec(x_20); +x_23 = lean_ctor_get(x_11, 12); +lean_inc(x_23); +lean_dec(x_11); +x_24 = lean_nat_dec_le(x_22, x_23); +lean_dec(x_23); +lean_dec(x_22); +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_21); +return x_26; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_checkMaxEmatchExceeded___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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_Grind_checkMaxEmatchExceeded(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); +lean_dec(x_1); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getENode_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; uint8_t x_12; +x_11 = lean_st_ref_get(x_2, x_10); +x_12 = !lean_is_exclusive(x_11); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_13 = lean_ctor_get(x_11, 0); +x_14 = lean_ctor_get(x_13, 1); +lean_inc(x_14); +lean_dec(x_13); +x_15 = l_Lean_PersistentHashMap_find_x3f___at___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot___spec__1(x_14, x_1); +lean_ctor_set(x_11, 0, x_15); +return x_11; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_16 = lean_ctor_get(x_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_ctor_get(x_16, 1); +lean_inc(x_18); +lean_dec(x_16); +x_19 = l_Lean_PersistentHashMap_find_x3f___at___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot___spec__1(x_18, x_1); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_19); +lean_ctor_set(x_20, 1, x_17); +return x_20; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getENode_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, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_Meta_Grind_getENode_x3f(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_Grind_getENode___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_11 = lean_ctor_get(x_8, 5); +x_12 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_1, x_6, x_7, x_8, x_9, x_10); +x_13 = !lean_is_exclusive(x_12); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; +x_14 = lean_ctor_get(x_12, 0); +lean_inc(x_11); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_11); +lean_ctor_set(x_15, 1, x_14); +lean_ctor_set_tag(x_12, 1); +lean_ctor_set(x_12, 0, x_15); +return x_12; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_16 = lean_ctor_get(x_12, 0); +x_17 = lean_ctor_get(x_12, 1); +lean_inc(x_17); +lean_inc(x_16); +lean_dec(x_12); +lean_inc(x_11); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_11); +lean_ctor_set(x_18, 1, x_16); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_17); +return x_19; +} +} +} +static lean_object* _init_l_Lean_Meta_Grind_getENode___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("internal `grind` error, term has not been internalized", 54, 54); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_getENode___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_getENode___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Grind_getENode___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2___closed__2; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getENode(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; uint8_t x_12; +x_11 = lean_st_ref_get(x_2, x_10); +x_12 = !lean_is_exclusive(x_11); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_13 = lean_ctor_get(x_11, 0); +x_14 = lean_ctor_get(x_11, 1); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +x_16 = l_Lean_PersistentHashMap_find_x3f___at___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot___spec__1(x_15, x_1); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +lean_free_object(x_11); +x_17 = l_Lean_indentExpr(x_1); +x_18 = l_Lean_Meta_Grind_getENode___closed__2; +x_19 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_17); +x_20 = l_Lean_Meta_Grind_getENode___closed__3; +x_21 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_21, 0, x_19); +lean_ctor_set(x_21, 1, x_20); +x_22 = l_Lean_throwError___at_Lean_Meta_Grind_getENode___spec__1(x_21, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_14); +return x_22; +} +else +{ +lean_object* x_23; +lean_dec(x_1); +x_23 = lean_ctor_get(x_16, 0); +lean_inc(x_23); +lean_dec(x_16); +lean_ctor_set(x_11, 0, x_23); +return x_11; +} +} +else +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +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_26 = lean_ctor_get(x_24, 1); +lean_inc(x_26); +lean_dec(x_24); +x_27 = l_Lean_PersistentHashMap_find_x3f___at___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot___spec__1(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; +x_28 = l_Lean_indentExpr(x_1); +x_29 = l_Lean_Meta_Grind_getENode___closed__2; +x_30 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_28); +x_31 = l_Lean_Meta_Grind_getENode___closed__3; +x_32 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_32, 0, x_30); +lean_ctor_set(x_32, 1, x_31); +x_33 = l_Lean_throwError___at_Lean_Meta_Grind_getENode___spec__1(x_32, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_25); +return x_33; +} +else +{ +lean_object* x_34; lean_object* x_35; +lean_dec(x_1); +x_34 = lean_ctor_get(x_27, 0); +lean_inc(x_34); +lean_dec(x_27); +x_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_34); +lean_ctor_set(x_35, 1, x_25); +return x_35; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_Grind_getENode___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_Lean_throwError___at_Lean_Meta_Grind_getENode___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_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getENode___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_Meta_Grind_getENode(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getGeneration(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_Meta_Grind_getENode(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_ctor_get(x_13, 8); +lean_inc(x_14); +lean_dec(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_ctor_get(x_15, 8); +lean_inc(x_17); +lean_dec(x_15); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_17); +lean_ctor_set(x_18, 1, x_16); +return x_18; +} +} +else +{ +uint8_t x_19; +x_19 = !lean_is_exclusive(x_11); +if (x_19 == 0) +{ +return x_11; +} +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_11, 0); +x_21 = lean_ctor_get(x_11, 1); +lean_inc(x_21); +lean_inc(x_20); +lean_dec(x_11); +x_22 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_22, 0, x_20); +lean_ctor_set(x_22, 1, x_21); +return x_22; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getGeneration___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_Meta_Grind_getGeneration(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isEqTrue(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_Meta_Grind_getENode(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_11) == 0) +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +x_13 = lean_ctor_get(x_11, 1); +lean_inc(x_13); +lean_dec(x_11); +x_14 = l_Lean_Meta_Grind_getTrueExpr___rarg(x_5, x_6, x_7, x_8, x_9, x_13); +x_15 = !lean_is_exclusive(x_14); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; uint8_t x_18; lean_object* x_19; +x_16 = lean_ctor_get(x_14, 0); +x_17 = lean_ctor_get(x_12, 2); +lean_inc(x_17); +lean_dec(x_12); +x_18 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_17, x_16); +lean_dec(x_16); +lean_dec(x_17); +x_19 = lean_box(x_18); +lean_ctor_set(x_14, 0, x_19); +return x_14; +} +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; lean_object* x_24; lean_object* x_25; +x_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_ctor_get(x_12, 2); +lean_inc(x_22); +lean_dec(x_12); +x_23 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_22, x_20); +lean_dec(x_20); +lean_dec(x_22); +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_21); +return x_25; +} +} +else +{ +uint8_t x_26; +x_26 = !lean_is_exclusive(x_11); +if (x_26 == 0) +{ +return x_11; +} +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_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_EXPORT lean_object* l_Lean_Meta_Grind_isEqTrue___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_Meta_Grind_isEqTrue(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isEqFalse(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_Meta_Grind_getENode(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_11) == 0) +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +x_13 = lean_ctor_get(x_11, 1); +lean_inc(x_13); +lean_dec(x_11); +x_14 = l_Lean_Meta_Grind_getFalseExpr___rarg(x_5, x_6, x_7, x_8, x_9, x_13); +x_15 = !lean_is_exclusive(x_14); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; uint8_t x_18; lean_object* x_19; +x_16 = lean_ctor_get(x_14, 0); +x_17 = lean_ctor_get(x_12, 2); +lean_inc(x_17); +lean_dec(x_12); +x_18 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_17, x_16); +lean_dec(x_16); +lean_dec(x_17); +x_19 = lean_box(x_18); +lean_ctor_set(x_14, 0, x_19); +return x_14; +} +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; lean_object* x_24; lean_object* x_25; +x_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_ctor_get(x_12, 2); +lean_inc(x_22); +lean_dec(x_12); +x_23 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_22, x_20); +lean_dec(x_20); +lean_dec(x_22); +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_21); +return x_25; +} +} +else +{ +uint8_t x_26; +x_26 = !lean_is_exclusive(x_11); +if (x_26 == 0) +{ +return x_11; +} +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_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_EXPORT lean_object* l_Lean_Meta_Grind_isEqFalse___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_Meta_Grind_isEqFalse(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isEqv(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, 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 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_1, x_2); +if (x_12 == 0) +{ +lean_object* x_13; +x_13 = l_Lean_Meta_Grind_getENode(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_13) == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +x_16 = l_Lean_Meta_Grind_getENode(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_15); +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; uint8_t x_21; lean_object* x_22; +x_18 = lean_ctor_get(x_16, 0); +x_19 = lean_ctor_get(x_14, 2); +lean_inc(x_19); +lean_dec(x_14); +x_20 = lean_ctor_get(x_18, 2); +lean_inc(x_20); +lean_dec(x_18); +x_21 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_19, x_20); +lean_dec(x_20); +lean_dec(x_19); +x_22 = lean_box(x_21); +lean_ctor_set(x_16, 0, x_22); +return x_16; +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; lean_object* x_28; lean_object* x_29; +x_23 = lean_ctor_get(x_16, 0); +x_24 = lean_ctor_get(x_16, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_16); +x_25 = lean_ctor_get(x_14, 2); +lean_inc(x_25); +lean_dec(x_14); +x_26 = lean_ctor_get(x_23, 2); +lean_inc(x_26); +lean_dec(x_23); +x_27 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_25, x_26); +lean_dec(x_26); +lean_dec(x_25); +x_28 = lean_box(x_27); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_24); +return x_29; +} +} +else +{ +uint8_t x_30; +lean_dec(x_14); +x_30 = !lean_is_exclusive(x_16); +if (x_30 == 0) +{ +return x_16; +} +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 +{ +uint8_t x_34; +lean_dec(x_2); +x_34 = !lean_is_exclusive(x_13); +if (x_34 == 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_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; +} +} +} +else +{ +uint8_t x_38; lean_object* x_39; lean_object* x_40; +lean_dec(x_2); +lean_dec(x_1); +x_38 = 1; +x_39 = lean_box(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_11); +return x_40; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isEqv___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, 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_Meta_Grind_isEqv(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isRoot(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; lean_object* x_12; +x_11 = l_Lean_Meta_Grind_getENode_x3f(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +if (lean_obj_tag(x_12) == 0) +{ +uint8_t x_13; +x_13 = !lean_is_exclusive(x_11); +if (x_13 == 0) +{ +lean_object* x_14; uint8_t x_15; lean_object* x_16; +x_14 = lean_ctor_get(x_11, 0); +lean_dec(x_14); +x_15 = 0; +x_16 = lean_box(x_15); +lean_ctor_set(x_11, 0, x_16); +return x_11; +} +else +{ +lean_object* x_17; uint8_t x_18; lean_object* x_19; lean_object* x_20; +x_17 = lean_ctor_get(x_11, 1); +lean_inc(x_17); +lean_dec(x_11); +x_18 = 0; +x_19 = lean_box(x_18); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_19); +lean_ctor_set(x_20, 1, x_17); +return x_20; +} +} +else +{ +uint8_t x_21; +x_21 = !lean_is_exclusive(x_11); +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_11, 0); +lean_dec(x_22); +x_23 = lean_ctor_get(x_12, 0); +lean_inc(x_23); +lean_dec(x_12); +x_24 = lean_ctor_get(x_23, 2); +lean_inc(x_24); +lean_dec(x_23); +x_25 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_24, x_1); +lean_dec(x_24); +x_26 = lean_box(x_25); +lean_ctor_set(x_11, 0, x_26); +return x_11; +} +else +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; uint8_t x_30; lean_object* x_31; lean_object* x_32; +x_27 = lean_ctor_get(x_11, 1); +lean_inc(x_27); +lean_dec(x_11); +x_28 = lean_ctor_get(x_12, 0); +lean_inc(x_28); +lean_dec(x_12); +x_29 = lean_ctor_get(x_28, 2); +lean_inc(x_29); +lean_dec(x_28); +x_30 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_29, x_1); +lean_dec(x_29); +x_31 = lean_box(x_30); +x_32 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_27); +return x_32; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isRoot___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_Meta_Grind_isRoot(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getRoot_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; +x_11 = l_Lean_Meta_Grind_getENode_x3f(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +if (lean_obj_tag(x_12) == 0) +{ +uint8_t x_13; +x_13 = !lean_is_exclusive(x_11); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; +x_14 = lean_ctor_get(x_11, 0); +lean_dec(x_14); +x_15 = lean_box(0); +lean_ctor_set(x_11, 0, x_15); +return x_11; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_11, 1); +lean_inc(x_16); +lean_dec(x_11); +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 +{ +uint8_t x_19; +x_19 = !lean_is_exclusive(x_11); +if (x_19 == 0) +{ +lean_object* x_20; uint8_t x_21; +x_20 = lean_ctor_get(x_11, 0); +lean_dec(x_20); +x_21 = !lean_is_exclusive(x_12); +if (x_21 == 0) +{ +lean_object* x_22; lean_object* x_23; +x_22 = lean_ctor_get(x_12, 0); +x_23 = lean_ctor_get(x_22, 2); +lean_inc(x_23); +lean_dec(x_22); +lean_ctor_set(x_12, 0, x_23); +return x_11; +} +else +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_12, 0); +lean_inc(x_24); +lean_dec(x_12); +x_25 = lean_ctor_get(x_24, 2); +lean_inc(x_25); +lean_dec(x_24); +x_26 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set(x_11, 0, x_26); +return x_11; +} +} +else +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_27 = lean_ctor_get(x_11, 1); +lean_inc(x_27); +lean_dec(x_11); +x_28 = lean_ctor_get(x_12, 0); +lean_inc(x_28); +if (lean_is_exclusive(x_12)) { + lean_ctor_release(x_12, 0); + x_29 = x_12; +} else { + lean_dec_ref(x_12); + x_29 = lean_box(0); +} +x_30 = lean_ctor_get(x_28, 2); +lean_inc(x_30); +lean_dec(x_28); +if (lean_is_scalar(x_29)) { + x_31 = lean_alloc_ctor(1, 1, 0); +} else { + x_31 = x_29; +} +lean_ctor_set(x_31, 0, x_30); +x_32 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_27); +return x_32; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getRoot_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, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_Meta_Grind_getRoot_x3f(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getRoot(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_Meta_Grind_getENode(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_ctor_get(x_13, 2); +lean_inc(x_14); +lean_dec(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_ctor_get(x_15, 2); +lean_inc(x_17); +lean_dec(x_15); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_17); +lean_ctor_set(x_18, 1, x_16); +return x_18; +} +} +else +{ +uint8_t x_19; +x_19 = !lean_is_exclusive(x_11); +if (x_19 == 0) +{ +return x_11; +} +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_11, 0); +x_21 = lean_ctor_get(x_11, 1); +lean_inc(x_21); +lean_inc(x_20); +lean_dec(x_11); +x_22 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_22, 0, x_20); +lean_ctor_set(x_22, 1, x_21); +return x_22; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getRoot___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_Meta_Grind_getRoot(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getRootENode(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_Meta_Grind_getRoot(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_11) == 0) +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +x_13 = lean_ctor_get(x_11, 1); +lean_inc(x_13); +lean_dec(x_11); +x_14 = l_Lean_Meta_Grind_getENode(x_12, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_13); +return x_14; +} +else +{ +uint8_t x_15; +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_Lean_Meta_Grind_getRootENode___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_Meta_Grind_getRootENode(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getNext(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_Meta_Grind_getENode(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_ctor_get(x_13, 1); +lean_inc(x_14); +lean_dec(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_ctor_get(x_15, 1); +lean_inc(x_17); +lean_dec(x_15); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_17); +lean_ctor_set(x_18, 1, x_16); +return x_18; +} +} +else +{ +uint8_t x_19; +x_19 = !lean_is_exclusive(x_11); +if (x_19 == 0) +{ +return x_11; +} +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_11, 0); +x_21 = lean_ctor_get(x_11, 1); +lean_inc(x_21); +lean_inc(x_20); +lean_dec(x_11); +x_22 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_22, 0, x_20); +lean_ctor_set(x_22, 1, x_21); +return x_22; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getNext___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_Meta_Grind_getNext(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_11; +} +} +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_alreadyInternalized___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) +{ +uint8_t x_8; +lean_dec(x_4); +x_8 = 0; +return x_8; +} +else +{ +lean_object* x_9; uint8_t x_10; +x_9 = lean_array_fget(x_1, x_4); +x_10 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(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 +{ +uint8_t x_14; +lean_dec(x_4); +x_14 = 1; +return x_14; +} +} +} +} +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_alreadyInternalized___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_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___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 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(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_Meta_Grind_alreadyInternalized___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_Meta_Grind_alreadyInternalized___spec__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint64_t x_3; size_t x_4; uint8_t x_5; +x_3 = l_Lean_Meta_Grind_instHashableENodeKey_unsafe__1(x_2); +x_4 = lean_uint64_to_usize(x_3); +x_5 = l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_alreadyInternalized___spec__2(x_1, x_4, x_2); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_alreadyInternalized(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; uint8_t x_12; +x_11 = lean_st_ref_get(x_2, x_10); +x_12 = !lean_is_exclusive(x_11); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; +x_13 = lean_ctor_get(x_11, 0); +x_14 = lean_ctor_get(x_13, 1); +lean_inc(x_14); +lean_dec(x_13); +x_15 = l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_alreadyInternalized___spec__1(x_14, x_1); +x_16 = lean_box(x_15); +lean_ctor_set(x_11, 0, x_16); +return x_11; +} +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; lean_object* x_21; lean_object* x_22; +x_17 = lean_ctor_get(x_11, 0); +x_18 = lean_ctor_get(x_11, 1); +lean_inc(x_18); +lean_inc(x_17); +lean_dec(x_11); +x_19 = lean_ctor_get(x_17, 1); +lean_inc(x_19); +lean_dec(x_17); +x_20 = l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_alreadyInternalized___spec__1(x_19, x_1); +x_21 = lean_box(x_20); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_18); +return x_22; +} +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_alreadyInternalized___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 = l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_alreadyInternalized___spec__3(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_5); +lean_dec(x_2); +lean_dec(x_1); +x_7 = lean_box(x_6); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_alreadyInternalized___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_Meta_Grind_alreadyInternalized___spec__2(x_1, x_4, x_3); +lean_dec(x_3); +x_6 = lean_box(x_5); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_alreadyInternalized___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_Meta_Grind_alreadyInternalized___spec__1(x_1, x_2); +lean_dec(x_2); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_alreadyInternalized___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_Meta_Grind_alreadyInternalized(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getTarget_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; +x_11 = l_Lean_Meta_Grind_getENode_x3f(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +if (lean_obj_tag(x_12) == 0) +{ +uint8_t x_13; +x_13 = !lean_is_exclusive(x_11); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; +x_14 = lean_ctor_get(x_11, 0); +lean_dec(x_14); +x_15 = lean_box(0); +lean_ctor_set(x_11, 0, x_15); +return x_11; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_11, 1); +lean_inc(x_16); +lean_dec(x_11); +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 +{ +uint8_t x_19; +x_19 = !lean_is_exclusive(x_11); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_11, 0); +lean_dec(x_20); +x_21 = lean_ctor_get(x_12, 0); +lean_inc(x_21); +lean_dec(x_12); +x_22 = lean_ctor_get(x_21, 4); +lean_inc(x_22); +lean_dec(x_21); +lean_ctor_set(x_11, 0, x_22); +return x_11; +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_23 = lean_ctor_get(x_11, 1); +lean_inc(x_23); +lean_dec(x_11); +x_24 = lean_ctor_get(x_12, 0); +lean_inc(x_24); +lean_dec(x_12); +x_25 = lean_ctor_get(x_24, 4); +lean_inc(x_25); +lean_dec(x_24); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set(x_26, 1, x_23); +return x_26; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getTarget_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, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_Meta_Grind_getTarget_x3f(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pushEqCore(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_14 = lean_st_ref_take(x_5, x_13); +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +x_16 = lean_ctor_get(x_14, 1); +lean_inc(x_16); +lean_dec(x_14); +x_17 = !lean_is_exclusive(x_15); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_18 = lean_ctor_get(x_15, 5); +x_19 = lean_alloc_ctor(0, 3, 1); +lean_ctor_set(x_19, 0, x_1); +lean_ctor_set(x_19, 1, x_2); +lean_ctor_set(x_19, 2, x_3); +lean_ctor_set_uint8(x_19, sizeof(void*)*3, x_4); +x_20 = lean_array_push(x_18, x_19); +lean_ctor_set(x_15, 5, x_20); +x_21 = lean_st_ref_set(x_5, x_15, x_16); +x_22 = !lean_is_exclusive(x_21); +if (x_22 == 0) +{ +lean_object* x_23; lean_object* x_24; +x_23 = lean_ctor_get(x_21, 0); +lean_dec(x_23); +x_24 = lean_box(0); +lean_ctor_set(x_21, 0, x_24); +return x_21; +} +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 = lean_box(0); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_25); +return x_27; +} +} +else +{ +lean_object* x_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; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; 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_28 = lean_ctor_get(x_15, 0); +x_29 = lean_ctor_get(x_15, 1); +x_30 = lean_ctor_get(x_15, 2); +x_31 = lean_ctor_get(x_15, 3); +x_32 = lean_ctor_get(x_15, 4); +x_33 = lean_ctor_get(x_15, 5); +x_34 = lean_ctor_get_uint8(x_15, sizeof(void*)*20); +x_35 = lean_ctor_get(x_15, 6); +x_36 = lean_ctor_get(x_15, 7); +x_37 = lean_ctor_get(x_15, 8); +x_38 = lean_ctor_get(x_15, 9); +x_39 = lean_ctor_get(x_15, 10); +x_40 = lean_ctor_get(x_15, 11); +x_41 = lean_ctor_get(x_15, 12); +x_42 = lean_ctor_get(x_15, 13); +x_43 = lean_ctor_get(x_15, 14); +x_44 = lean_ctor_get(x_15, 15); +x_45 = lean_ctor_get(x_15, 16); +x_46 = lean_ctor_get(x_15, 17); +x_47 = lean_ctor_get(x_15, 18); +x_48 = lean_ctor_get(x_15, 19); +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_inc(x_33); +lean_inc(x_32); +lean_inc(x_31); +lean_inc(x_30); +lean_inc(x_29); +lean_inc(x_28); +lean_dec(x_15); +x_49 = lean_alloc_ctor(0, 3, 1); +lean_ctor_set(x_49, 0, x_1); +lean_ctor_set(x_49, 1, x_2); +lean_ctor_set(x_49, 2, x_3); +lean_ctor_set_uint8(x_49, sizeof(void*)*3, x_4); +x_50 = lean_array_push(x_33, x_49); +x_51 = lean_alloc_ctor(0, 20, 1); +lean_ctor_set(x_51, 0, x_28); +lean_ctor_set(x_51, 1, x_29); +lean_ctor_set(x_51, 2, x_30); +lean_ctor_set(x_51, 3, x_31); +lean_ctor_set(x_51, 4, x_32); +lean_ctor_set(x_51, 5, x_50); +lean_ctor_set(x_51, 6, x_35); +lean_ctor_set(x_51, 7, x_36); +lean_ctor_set(x_51, 8, x_37); +lean_ctor_set(x_51, 9, x_38); +lean_ctor_set(x_51, 10, x_39); +lean_ctor_set(x_51, 11, x_40); +lean_ctor_set(x_51, 12, x_41); +lean_ctor_set(x_51, 13, x_42); +lean_ctor_set(x_51, 14, x_43); +lean_ctor_set(x_51, 15, x_44); +lean_ctor_set(x_51, 16, x_45); +lean_ctor_set(x_51, 17, x_46); +lean_ctor_set(x_51, 18, x_47); +lean_ctor_set(x_51, 19, x_48); +lean_ctor_set_uint8(x_51, sizeof(void*)*20, x_34); +x_52 = lean_st_ref_set(x_5, x_51, x_16); +x_53 = lean_ctor_get(x_52, 1); +lean_inc(x_53); +if (lean_is_exclusive(x_52)) { + lean_ctor_release(x_52, 0); + lean_ctor_release(x_52, 1); + x_54 = x_52; +} else { + lean_dec_ref(x_52); + x_54 = lean_box(0); +} +x_55 = lean_box(0); +if (lean_is_scalar(x_54)) { + x_56 = lean_alloc_ctor(0, 2, 0); +} else { + x_56 = x_54; +} +lean_ctor_set(x_56, 0, x_55); +lean_ctor_set(x_56, 1, x_53); +return x_56; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pushEqCore___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, 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_Meta_Grind_pushEqCore(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); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +return x_15; +} +} +static uint64_t _init_l_Lean_Meta_Grind_hasSameType___closed__1() { +_start: +{ +uint8_t x_1; uint64_t x_2; +x_1 = 1; +x_2 = l_Lean_Meta_TransparencyMode_toUInt64(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_hasSameType(lean_object* x_1, 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_3); +if (x_8 == 0) +{ +lean_object* x_9; uint8_t x_10; +x_9 = lean_ctor_get(x_3, 0); +x_10 = !lean_is_exclusive(x_9); +if (x_10 == 0) +{ +uint64_t x_11; uint8_t x_12; uint64_t x_13; uint64_t x_14; uint64_t x_15; uint64_t x_16; uint64_t x_17; lean_object* x_18; +x_11 = lean_ctor_get_uint64(x_3, sizeof(void*)*7); +x_12 = 1; +lean_ctor_set_uint8(x_9, 9, x_12); +x_13 = 2; +x_14 = lean_uint64_shift_right(x_11, x_13); +x_15 = lean_uint64_shift_left(x_14, x_13); +x_16 = l_Lean_Meta_Grind_hasSameType___closed__1; +x_17 = lean_uint64_lor(x_15, x_16); +lean_ctor_set_uint64(x_3, sizeof(void*)*7, x_17); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_18 = lean_infer_type(x_1, x_3, x_4, x_5, x_6, x_7); +if (lean_obj_tag(x_18) == 0) +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_18, 1); +lean_inc(x_20); +lean_dec(x_18); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_21 = lean_infer_type(x_2, x_3, x_4, x_5, x_6, x_20); +if (lean_obj_tag(x_21) == 0) +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_22 = lean_ctor_get(x_21, 0); +lean_inc(x_22); +x_23 = lean_ctor_get(x_21, 1); +lean_inc(x_23); +lean_dec(x_21); +x_24 = l_Lean_Meta_isExprDefEq(x_19, x_22, x_3, x_4, x_5, x_6, x_23); +if (lean_obj_tag(x_24) == 0) +{ +uint8_t x_25; +x_25 = !lean_is_exclusive(x_24); +if (x_25 == 0) +{ +return x_24; +} +else +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_26 = lean_ctor_get(x_24, 0); +x_27 = lean_ctor_get(x_24, 1); +lean_inc(x_27); +lean_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 +{ +uint8_t x_29; +x_29 = !lean_is_exclusive(x_24); +if (x_29 == 0) +{ +return x_24; +} +else +{ +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; +} +} +} +else +{ +uint8_t x_33; +lean_dec(x_19); +lean_dec(x_3); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_33 = !lean_is_exclusive(x_21); +if (x_33 == 0) +{ +return x_21; +} +else +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_21, 0); +x_35 = lean_ctor_get(x_21, 1); +lean_inc(x_35); +lean_inc(x_34); +lean_dec(x_21); +x_36 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_36, 0, x_34); +lean_ctor_set(x_36, 1, x_35); +return x_36; +} +} +} +else +{ +uint8_t x_37; +lean_dec(x_3); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +x_37 = !lean_is_exclusive(x_18); +if (x_37 == 0) +{ +return x_18; +} +else +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_38 = lean_ctor_get(x_18, 0); +x_39 = lean_ctor_get(x_18, 1); +lean_inc(x_39); +lean_inc(x_38); +lean_dec(x_18); +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 +{ +uint64_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; 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; lean_object* x_59; uint64_t x_60; uint64_t x_61; uint64_t x_62; uint64_t x_63; uint64_t x_64; lean_object* x_65; +x_41 = lean_ctor_get_uint64(x_3, sizeof(void*)*7); +x_42 = lean_ctor_get_uint8(x_9, 0); +x_43 = lean_ctor_get_uint8(x_9, 1); +x_44 = lean_ctor_get_uint8(x_9, 2); +x_45 = lean_ctor_get_uint8(x_9, 3); +x_46 = lean_ctor_get_uint8(x_9, 4); +x_47 = lean_ctor_get_uint8(x_9, 5); +x_48 = lean_ctor_get_uint8(x_9, 6); +x_49 = lean_ctor_get_uint8(x_9, 7); +x_50 = lean_ctor_get_uint8(x_9, 8); +x_51 = lean_ctor_get_uint8(x_9, 10); +x_52 = lean_ctor_get_uint8(x_9, 11); +x_53 = lean_ctor_get_uint8(x_9, 12); +x_54 = lean_ctor_get_uint8(x_9, 13); +x_55 = lean_ctor_get_uint8(x_9, 14); +x_56 = lean_ctor_get_uint8(x_9, 15); +x_57 = lean_ctor_get_uint8(x_9, 16); +lean_dec(x_9); +x_58 = 1; +x_59 = lean_alloc_ctor(0, 0, 17); +lean_ctor_set_uint8(x_59, 0, x_42); +lean_ctor_set_uint8(x_59, 1, x_43); +lean_ctor_set_uint8(x_59, 2, x_44); +lean_ctor_set_uint8(x_59, 3, x_45); +lean_ctor_set_uint8(x_59, 4, x_46); +lean_ctor_set_uint8(x_59, 5, x_47); +lean_ctor_set_uint8(x_59, 6, x_48); +lean_ctor_set_uint8(x_59, 7, x_49); +lean_ctor_set_uint8(x_59, 8, x_50); +lean_ctor_set_uint8(x_59, 9, x_58); +lean_ctor_set_uint8(x_59, 10, x_51); +lean_ctor_set_uint8(x_59, 11, x_52); +lean_ctor_set_uint8(x_59, 12, x_53); +lean_ctor_set_uint8(x_59, 13, x_54); +lean_ctor_set_uint8(x_59, 14, x_55); +lean_ctor_set_uint8(x_59, 15, x_56); +lean_ctor_set_uint8(x_59, 16, x_57); +x_60 = 2; +x_61 = lean_uint64_shift_right(x_41, x_60); +x_62 = lean_uint64_shift_left(x_61, x_60); +x_63 = l_Lean_Meta_Grind_hasSameType___closed__1; +x_64 = lean_uint64_lor(x_62, x_63); +lean_ctor_set(x_3, 0, x_59); +lean_ctor_set_uint64(x_3, sizeof(void*)*7, x_64); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_65 = lean_infer_type(x_1, x_3, x_4, x_5, x_6, x_7); +if (lean_obj_tag(x_65) == 0) +{ +lean_object* x_66; lean_object* x_67; lean_object* x_68; +x_66 = lean_ctor_get(x_65, 0); +lean_inc(x_66); +x_67 = lean_ctor_get(x_65, 1); +lean_inc(x_67); +lean_dec(x_65); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_68 = lean_infer_type(x_2, x_3, x_4, x_5, x_6, x_67); +if (lean_obj_tag(x_68) == 0) +{ +lean_object* x_69; lean_object* x_70; lean_object* x_71; +x_69 = lean_ctor_get(x_68, 0); +lean_inc(x_69); +x_70 = lean_ctor_get(x_68, 1); +lean_inc(x_70); +lean_dec(x_68); +x_71 = l_Lean_Meta_isExprDefEq(x_66, x_69, x_3, x_4, x_5, x_6, x_70); +if (lean_obj_tag(x_71) == 0) +{ +lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; +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; +} else { + lean_dec_ref(x_71); + 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_72); +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; +x_76 = lean_ctor_get(x_71, 0); +lean_inc(x_76); +x_77 = lean_ctor_get(x_71, 1); +lean_inc(x_77); +if (lean_is_exclusive(x_71)) { + lean_ctor_release(x_71, 0); + lean_ctor_release(x_71, 1); + x_78 = x_71; +} else { + lean_dec_ref(x_71); + 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 +{ +lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; +lean_dec(x_66); +lean_dec(x_3); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_80 = lean_ctor_get(x_68, 0); +lean_inc(x_80); +x_81 = lean_ctor_get(x_68, 1); +lean_inc(x_81); +if (lean_is_exclusive(x_68)) { + lean_ctor_release(x_68, 0); + lean_ctor_release(x_68, 1); + x_82 = x_68; +} else { + lean_dec_ref(x_68); + 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; lean_object* x_87; +lean_dec(x_3); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +x_84 = lean_ctor_get(x_65, 0); +lean_inc(x_84); +x_85 = lean_ctor_get(x_65, 1); +lean_inc(x_85); +if (lean_is_exclusive(x_65)) { + lean_ctor_release(x_65, 0); + lean_ctor_release(x_65, 1); + x_86 = x_65; +} else { + lean_dec_ref(x_65); + x_86 = lean_box(0); +} +if (lean_is_scalar(x_86)) { + x_87 = lean_alloc_ctor(1, 2, 0); +} else { + x_87 = x_86; +} +lean_ctor_set(x_87, 0, x_84); +lean_ctor_set(x_87, 1, x_85); +return x_87; +} +} +} +else +{ +lean_object* x_88; uint64_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; uint8_t x_97; uint8_t 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; uint8_t x_113; uint8_t x_114; lean_object* x_115; uint8_t x_116; lean_object* x_117; uint64_t x_118; uint64_t x_119; uint64_t x_120; uint64_t x_121; uint64_t x_122; lean_object* x_123; lean_object* x_124; +x_88 = lean_ctor_get(x_3, 0); +x_89 = lean_ctor_get_uint64(x_3, sizeof(void*)*7); +x_90 = lean_ctor_get_uint8(x_3, sizeof(void*)*7 + 8); +x_91 = lean_ctor_get(x_3, 1); +x_92 = lean_ctor_get(x_3, 2); +x_93 = lean_ctor_get(x_3, 3); +x_94 = lean_ctor_get(x_3, 4); +x_95 = lean_ctor_get(x_3, 5); +x_96 = lean_ctor_get(x_3, 6); +x_97 = lean_ctor_get_uint8(x_3, sizeof(void*)*7 + 9); +x_98 = lean_ctor_get_uint8(x_3, sizeof(void*)*7 + 10); +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_88); +lean_dec(x_3); +x_99 = lean_ctor_get_uint8(x_88, 0); +x_100 = lean_ctor_get_uint8(x_88, 1); +x_101 = lean_ctor_get_uint8(x_88, 2); +x_102 = lean_ctor_get_uint8(x_88, 3); +x_103 = lean_ctor_get_uint8(x_88, 4); +x_104 = lean_ctor_get_uint8(x_88, 5); +x_105 = lean_ctor_get_uint8(x_88, 6); +x_106 = lean_ctor_get_uint8(x_88, 7); +x_107 = lean_ctor_get_uint8(x_88, 8); +x_108 = lean_ctor_get_uint8(x_88, 10); +x_109 = lean_ctor_get_uint8(x_88, 11); +x_110 = lean_ctor_get_uint8(x_88, 12); +x_111 = lean_ctor_get_uint8(x_88, 13); +x_112 = lean_ctor_get_uint8(x_88, 14); +x_113 = lean_ctor_get_uint8(x_88, 15); +x_114 = lean_ctor_get_uint8(x_88, 16); +if (lean_is_exclusive(x_88)) { + x_115 = x_88; +} else { + lean_dec_ref(x_88); + x_115 = lean_box(0); +} +x_116 = 1; +if (lean_is_scalar(x_115)) { + x_117 = lean_alloc_ctor(0, 0, 17); +} else { + x_117 = x_115; +} +lean_ctor_set_uint8(x_117, 0, x_99); +lean_ctor_set_uint8(x_117, 1, x_100); +lean_ctor_set_uint8(x_117, 2, x_101); +lean_ctor_set_uint8(x_117, 3, x_102); +lean_ctor_set_uint8(x_117, 4, x_103); +lean_ctor_set_uint8(x_117, 5, x_104); +lean_ctor_set_uint8(x_117, 6, x_105); +lean_ctor_set_uint8(x_117, 7, x_106); +lean_ctor_set_uint8(x_117, 8, x_107); +lean_ctor_set_uint8(x_117, 9, x_116); +lean_ctor_set_uint8(x_117, 10, x_108); +lean_ctor_set_uint8(x_117, 11, x_109); +lean_ctor_set_uint8(x_117, 12, x_110); +lean_ctor_set_uint8(x_117, 13, x_111); +lean_ctor_set_uint8(x_117, 14, x_112); +lean_ctor_set_uint8(x_117, 15, x_113); +lean_ctor_set_uint8(x_117, 16, x_114); +x_118 = 2; +x_119 = lean_uint64_shift_right(x_89, x_118); +x_120 = lean_uint64_shift_left(x_119, x_118); +x_121 = l_Lean_Meta_Grind_hasSameType___closed__1; +x_122 = lean_uint64_lor(x_120, x_121); +x_123 = lean_alloc_ctor(0, 7, 11); +lean_ctor_set(x_123, 0, x_117); +lean_ctor_set(x_123, 1, x_91); +lean_ctor_set(x_123, 2, x_92); +lean_ctor_set(x_123, 3, x_93); +lean_ctor_set(x_123, 4, x_94); +lean_ctor_set(x_123, 5, x_95); +lean_ctor_set(x_123, 6, x_96); +lean_ctor_set_uint64(x_123, sizeof(void*)*7, x_122); +lean_ctor_set_uint8(x_123, sizeof(void*)*7 + 8, x_90); +lean_ctor_set_uint8(x_123, sizeof(void*)*7 + 9, x_97); +lean_ctor_set_uint8(x_123, sizeof(void*)*7 + 10, x_98); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_123); +x_124 = lean_infer_type(x_1, x_123, x_4, x_5, x_6, x_7); +if (lean_obj_tag(x_124) == 0) +{ +lean_object* x_125; lean_object* x_126; lean_object* x_127; +x_125 = lean_ctor_get(x_124, 0); +lean_inc(x_125); +x_126 = lean_ctor_get(x_124, 1); +lean_inc(x_126); +lean_dec(x_124); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_123); +x_127 = lean_infer_type(x_2, x_123, x_4, x_5, x_6, x_126); +if (lean_obj_tag(x_127) == 0) +{ +lean_object* x_128; lean_object* x_129; lean_object* x_130; +x_128 = lean_ctor_get(x_127, 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_Meta_isExprDefEq(x_125, x_128, x_123, x_4, x_5, x_6, x_129); +if (lean_obj_tag(x_130) == 0) +{ +lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; +x_131 = lean_ctor_get(x_130, 0); +lean_inc(x_131); +x_132 = lean_ctor_get(x_130, 1); +lean_inc(x_132); +if (lean_is_exclusive(x_130)) { + lean_ctor_release(x_130, 0); + lean_ctor_release(x_130, 1); + x_133 = x_130; +} else { + lean_dec_ref(x_130); + x_133 = lean_box(0); +} +if (lean_is_scalar(x_133)) { + x_134 = lean_alloc_ctor(0, 2, 0); +} else { + x_134 = x_133; +} +lean_ctor_set(x_134, 0, x_131); +lean_ctor_set(x_134, 1, x_132); +return x_134; +} +else +{ +lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; +x_135 = lean_ctor_get(x_130, 0); +lean_inc(x_135); +x_136 = lean_ctor_get(x_130, 1); +lean_inc(x_136); +if (lean_is_exclusive(x_130)) { + lean_ctor_release(x_130, 0); + lean_ctor_release(x_130, 1); + x_137 = x_130; +} else { + lean_dec_ref(x_130); + 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; +} +} +else +{ +lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; +lean_dec(x_125); +lean_dec(x_123); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_139 = lean_ctor_get(x_127, 0); +lean_inc(x_139); +x_140 = lean_ctor_get(x_127, 1); +lean_inc(x_140); +if (lean_is_exclusive(x_127)) { + lean_ctor_release(x_127, 0); + lean_ctor_release(x_127, 1); + x_141 = x_127; +} else { + lean_dec_ref(x_127); + 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(x_142, 0, x_139); +lean_ctor_set(x_142, 1, x_140); +return x_142; +} +} +else +{ +lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; +lean_dec(x_123); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +x_143 = lean_ctor_get(x_124, 0); +lean_inc(x_143); +x_144 = lean_ctor_get(x_124, 1); +lean_inc(x_144); +if (lean_is_exclusive(x_124)) { + lean_ctor_release(x_124, 0); + lean_ctor_release(x_124, 1); + x_145 = x_124; +} else { + lean_dec_ref(x_124); + x_145 = lean_box(0); +} +if (lean_is_scalar(x_145)) { + x_146 = lean_alloc_ctor(1, 2, 0); +} else { + x_146 = x_145; +} +lean_ctor_set(x_146, 0, x_143); +lean_ctor_set(x_146, 1, x_144); +return x_146; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pushEqHEq(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, 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_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_2); +lean_inc(x_1); +x_13 = l_Lean_Meta_Grind_hasSameType(x_1, x_2, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_13) == 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_16; uint8_t x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_13, 1); +lean_inc(x_16); +lean_dec(x_13); +x_17 = 1; +x_18 = l_Lean_Meta_Grind_pushEqCore(x_1, x_2, x_3, x_17, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_16); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +return x_18; +} +else +{ +lean_object* x_19; uint8_t x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_13, 1); +lean_inc(x_19); +lean_dec(x_13); +x_20 = 0; +x_21 = l_Lean_Meta_Grind_pushEqCore(x_1, x_2, x_3, x_20, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_19); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +return x_21; +} +} +else +{ +uint8_t x_22; +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_22 = !lean_is_exclusive(x_13); +if (x_22 == 0) +{ +return x_13; +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_13, 0); +x_24 = lean_ctor_get(x_13, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_13); +x_25 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_24); +return x_25; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pushEqHEq___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +x_13 = l_Lean_Meta_Grind_pushEqHEq(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_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pushEq(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, 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 = 0; +x_14 = l_Lean_Meta_Grind_pushEqCore(x_1, x_2, x_3, 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_Meta_Grind_pushEq___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +x_13 = l_Lean_Meta_Grind_pushEq(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pushHEq(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, 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 = 1; +x_14 = l_Lean_Meta_Grind_pushEqCore(x_1, x_2, x_3, 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_Meta_Grind_pushHEq___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +x_13 = l_Lean_Meta_Grind_pushHEq(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pushEqTrue(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, 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; lean_object* x_16; +x_12 = l_Lean_Meta_Grind_getTrueExpr___rarg(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_ctor_get(x_12, 1); +lean_inc(x_14); +lean_dec(x_12); +x_15 = 0; +x_16 = l_Lean_Meta_Grind_pushEqCore(x_1, x_13, x_2, x_15, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_14); +return x_16; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pushEqTrue___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, 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_Meta_Grind_pushEqTrue(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pushEqFalse(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, 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; lean_object* x_16; +x_12 = l_Lean_Meta_Grind_getFalseExpr___rarg(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_ctor_get(x_12, 1); +lean_inc(x_14); +lean_dec(x_12); +x_15 = 0; +x_16 = l_Lean_Meta_Grind_pushEqCore(x_1, x_13, x_2, x_15, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_14); +return x_16; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pushEqFalse___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, 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_Meta_Grind_pushEqFalse(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Lean_RBNode_ins___at_Lean_Meta_Grind_registerParent___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_4; uint8_t x_5; lean_object* x_6; +x_4 = lean_box(0); +x_5 = 0; +x_6 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_6, 0, x_4); +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_uint8(x_6, sizeof(void*)*4, x_5); +return x_6; +} +else +{ +uint8_t x_7; +x_7 = lean_ctor_get_uint8(x_1, sizeof(void*)*4); +if (x_7 == 0) +{ +uint8_t x_8; +x_8 = !lean_is_exclusive(x_1); +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_1, 0); +x_10 = lean_ctor_get(x_1, 1); +x_11 = lean_ctor_get(x_1, 2); +x_12 = lean_ctor_get(x_1, 3); +x_13 = l_Lean_Expr_quickComp(x_2, x_10); +switch (x_13) { +case 0: +{ +lean_object* x_14; uint8_t x_15; +x_14 = l_Lean_RBNode_ins___at_Lean_Meta_Grind_registerParent___spec__2(x_9, x_2, x_3); +x_15 = 0; +lean_ctor_set(x_1, 0, x_14); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_15); +return x_1; +} +case 1: +{ +uint8_t x_16; +lean_dec(x_11); +lean_dec(x_10); +x_16 = 0; +lean_ctor_set(x_1, 2, x_3); +lean_ctor_set(x_1, 1, x_2); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_16); +return x_1; +} +default: +{ +lean_object* x_17; uint8_t x_18; +x_17 = l_Lean_RBNode_ins___at_Lean_Meta_Grind_registerParent___spec__2(x_12, x_2, x_3); +x_18 = 0; +lean_ctor_set(x_1, 3, x_17); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_18); +return x_1; +} +} +} +else +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; +x_19 = lean_ctor_get(x_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_inc(x_22); +lean_inc(x_21); +lean_inc(x_20); +lean_inc(x_19); +lean_dec(x_1); +x_23 = l_Lean_Expr_quickComp(x_2, x_20); +switch (x_23) { +case 0: +{ +lean_object* x_24; uint8_t x_25; lean_object* x_26; +x_24 = l_Lean_RBNode_ins___at_Lean_Meta_Grind_registerParent___spec__2(x_19, x_2, x_3); +x_25 = 0; +x_26 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_26, 0, x_24); +lean_ctor_set(x_26, 1, x_20); +lean_ctor_set(x_26, 2, x_21); +lean_ctor_set(x_26, 3, x_22); +lean_ctor_set_uint8(x_26, sizeof(void*)*4, x_25); +return x_26; +} +case 1: +{ +uint8_t x_27; lean_object* x_28; +lean_dec(x_21); +lean_dec(x_20); +x_27 = 0; +x_28 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_28, 0, x_19); +lean_ctor_set(x_28, 1, x_2); +lean_ctor_set(x_28, 2, x_3); +lean_ctor_set(x_28, 3, x_22); +lean_ctor_set_uint8(x_28, sizeof(void*)*4, x_27); +return x_28; +} +default: +{ +lean_object* x_29; uint8_t x_30; lean_object* x_31; +x_29 = l_Lean_RBNode_ins___at_Lean_Meta_Grind_registerParent___spec__2(x_22, x_2, x_3); +x_30 = 0; +x_31 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_31, 0, x_19); +lean_ctor_set(x_31, 1, x_20); +lean_ctor_set(x_31, 2, x_21); +lean_ctor_set(x_31, 3, x_29); +lean_ctor_set_uint8(x_31, sizeof(void*)*4, x_30); +return x_31; +} +} +} +} +else +{ +uint8_t x_32; +x_32 = !lean_is_exclusive(x_1); +if (x_32 == 0) +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; +x_33 = lean_ctor_get(x_1, 0); +x_34 = lean_ctor_get(x_1, 1); +x_35 = lean_ctor_get(x_1, 2); +x_36 = lean_ctor_get(x_1, 3); +x_37 = l_Lean_Expr_quickComp(x_2, x_34); +switch (x_37) { +case 0: +{ +lean_object* x_38; uint8_t x_39; +x_38 = l_Lean_RBNode_ins___at_Lean_Meta_Grind_registerParent___spec__2(x_33, x_2, x_3); +x_39 = lean_ctor_get_uint8(x_38, sizeof(void*)*4); +if (x_39 == 0) +{ +lean_object* x_40; +x_40 = lean_ctor_get(x_38, 0); +lean_inc(x_40); +if (lean_obj_tag(x_40) == 0) +{ +lean_object* x_41; +x_41 = lean_ctor_get(x_38, 3); +lean_inc(x_41); +if (lean_obj_tag(x_41) == 0) +{ +uint8_t x_42; +x_42 = !lean_is_exclusive(x_38); +if (x_42 == 0) +{ +lean_object* x_43; lean_object* x_44; uint8_t x_45; +x_43 = lean_ctor_get(x_38, 3); +lean_dec(x_43); +x_44 = lean_ctor_get(x_38, 0); +lean_dec(x_44); +lean_ctor_set(x_38, 0, x_41); +x_45 = 1; +lean_ctor_set(x_1, 0, x_38); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_45); +return x_1; +} +else +{ +lean_object* x_46; lean_object* x_47; lean_object* x_48; uint8_t x_49; +x_46 = lean_ctor_get(x_38, 1); +x_47 = lean_ctor_get(x_38, 2); +lean_inc(x_47); +lean_inc(x_46); +lean_dec(x_38); +x_48 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_48, 0, x_41); +lean_ctor_set(x_48, 1, x_46); +lean_ctor_set(x_48, 2, x_47); +lean_ctor_set(x_48, 3, x_41); +lean_ctor_set_uint8(x_48, sizeof(void*)*4, x_39); +x_49 = 1; +lean_ctor_set(x_1, 0, x_48); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_49); +return x_1; +} +} +else +{ +uint8_t x_50; +x_50 = lean_ctor_get_uint8(x_41, sizeof(void*)*4); +if (x_50 == 0) +{ +uint8_t x_51; +x_51 = !lean_is_exclusive(x_38); +if (x_51 == 0) +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; uint8_t x_56; +x_52 = lean_ctor_get(x_38, 1); +x_53 = lean_ctor_get(x_38, 2); +x_54 = lean_ctor_get(x_38, 3); +lean_dec(x_54); +x_55 = lean_ctor_get(x_38, 0); +lean_dec(x_55); +x_56 = !lean_is_exclusive(x_41); +if (x_56 == 0) +{ +lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; uint8_t x_61; uint8_t x_62; +x_57 = lean_ctor_get(x_41, 0); +x_58 = lean_ctor_get(x_41, 1); +x_59 = lean_ctor_get(x_41, 2); +x_60 = lean_ctor_get(x_41, 3); +x_61 = 1; +lean_ctor_set(x_41, 3, x_57); +lean_ctor_set(x_41, 2, x_53); +lean_ctor_set(x_41, 1, x_52); +lean_ctor_set(x_41, 0, x_40); +lean_ctor_set_uint8(x_41, sizeof(void*)*4, x_61); +lean_ctor_set(x_38, 3, x_36); +lean_ctor_set(x_38, 2, x_35); +lean_ctor_set(x_38, 1, x_34); +lean_ctor_set(x_38, 0, x_60); +lean_ctor_set_uint8(x_38, sizeof(void*)*4, x_61); +x_62 = 0; +lean_ctor_set(x_1, 3, x_38); +lean_ctor_set(x_1, 2, x_59); +lean_ctor_set(x_1, 1, x_58); +lean_ctor_set(x_1, 0, x_41); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_62); +return x_1; +} +else +{ +lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; uint8_t x_67; lean_object* x_68; uint8_t x_69; +x_63 = lean_ctor_get(x_41, 0); +x_64 = lean_ctor_get(x_41, 1); +x_65 = lean_ctor_get(x_41, 2); +x_66 = lean_ctor_get(x_41, 3); +lean_inc(x_66); +lean_inc(x_65); +lean_inc(x_64); +lean_inc(x_63); +lean_dec(x_41); +x_67 = 1; +x_68 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_68, 0, x_40); +lean_ctor_set(x_68, 1, x_52); +lean_ctor_set(x_68, 2, x_53); +lean_ctor_set(x_68, 3, x_63); +lean_ctor_set_uint8(x_68, sizeof(void*)*4, x_67); +lean_ctor_set(x_38, 3, x_36); +lean_ctor_set(x_38, 2, x_35); +lean_ctor_set(x_38, 1, x_34); +lean_ctor_set(x_38, 0, x_66); +lean_ctor_set_uint8(x_38, sizeof(void*)*4, x_67); +x_69 = 0; +lean_ctor_set(x_1, 3, x_38); +lean_ctor_set(x_1, 2, x_65); +lean_ctor_set(x_1, 1, x_64); +lean_ctor_set(x_1, 0, x_68); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_69); +return x_1; +} +} +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; uint8_t x_77; lean_object* x_78; lean_object* x_79; uint8_t x_80; +x_70 = lean_ctor_get(x_38, 1); +x_71 = lean_ctor_get(x_38, 2); +lean_inc(x_71); +lean_inc(x_70); +lean_dec(x_38); +x_72 = lean_ctor_get(x_41, 0); +lean_inc(x_72); +x_73 = lean_ctor_get(x_41, 1); +lean_inc(x_73); +x_74 = lean_ctor_get(x_41, 2); +lean_inc(x_74); +x_75 = lean_ctor_get(x_41, 3); +lean_inc(x_75); +if (lean_is_exclusive(x_41)) { + lean_ctor_release(x_41, 0); + lean_ctor_release(x_41, 1); + lean_ctor_release(x_41, 2); + lean_ctor_release(x_41, 3); + x_76 = x_41; +} else { + lean_dec_ref(x_41); + x_76 = lean_box(0); +} +x_77 = 1; +if (lean_is_scalar(x_76)) { + x_78 = lean_alloc_ctor(1, 4, 1); +} else { + x_78 = x_76; +} +lean_ctor_set(x_78, 0, x_40); +lean_ctor_set(x_78, 1, x_70); +lean_ctor_set(x_78, 2, x_71); +lean_ctor_set(x_78, 3, x_72); +lean_ctor_set_uint8(x_78, sizeof(void*)*4, x_77); +x_79 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_79, 0, x_75); +lean_ctor_set(x_79, 1, x_34); +lean_ctor_set(x_79, 2, x_35); +lean_ctor_set(x_79, 3, x_36); +lean_ctor_set_uint8(x_79, sizeof(void*)*4, x_77); +x_80 = 0; +lean_ctor_set(x_1, 3, x_79); +lean_ctor_set(x_1, 2, x_74); +lean_ctor_set(x_1, 1, x_73); +lean_ctor_set(x_1, 0, x_78); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_80); +return x_1; +} +} +else +{ +uint8_t x_81; +lean_free_object(x_1); +x_81 = !lean_is_exclusive(x_41); +if (x_81 == 0) +{ +lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; uint8_t x_86; +x_82 = lean_ctor_get(x_41, 3); +lean_dec(x_82); +x_83 = lean_ctor_get(x_41, 2); +lean_dec(x_83); +x_84 = lean_ctor_get(x_41, 1); +lean_dec(x_84); +x_85 = lean_ctor_get(x_41, 0); +lean_dec(x_85); +x_86 = 1; +lean_ctor_set(x_41, 3, x_36); +lean_ctor_set(x_41, 2, x_35); +lean_ctor_set(x_41, 1, x_34); +lean_ctor_set(x_41, 0, x_38); +lean_ctor_set_uint8(x_41, sizeof(void*)*4, x_86); +return x_41; +} +else +{ +uint8_t x_87; lean_object* x_88; +lean_dec(x_41); +x_87 = 1; +x_88 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_88, 0, x_38); +lean_ctor_set(x_88, 1, x_34); +lean_ctor_set(x_88, 2, x_35); +lean_ctor_set(x_88, 3, x_36); +lean_ctor_set_uint8(x_88, sizeof(void*)*4, x_87); +return x_88; +} +} +} +} +else +{ +uint8_t x_89; +x_89 = lean_ctor_get_uint8(x_40, sizeof(void*)*4); +if (x_89 == 0) +{ +uint8_t x_90; +x_90 = !lean_is_exclusive(x_38); +if (x_90 == 0) +{ +lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; uint8_t x_95; +x_91 = lean_ctor_get(x_38, 1); +x_92 = lean_ctor_get(x_38, 2); +x_93 = lean_ctor_get(x_38, 3); +x_94 = lean_ctor_get(x_38, 0); +lean_dec(x_94); +x_95 = !lean_is_exclusive(x_40); +if (x_95 == 0) +{ +uint8_t x_96; uint8_t x_97; +x_96 = 1; +lean_ctor_set_uint8(x_40, sizeof(void*)*4, x_96); +lean_ctor_set(x_38, 3, x_36); +lean_ctor_set(x_38, 2, x_35); +lean_ctor_set(x_38, 1, x_34); +lean_ctor_set(x_38, 0, x_93); +lean_ctor_set_uint8(x_38, sizeof(void*)*4, x_96); +x_97 = 0; +lean_ctor_set(x_1, 3, x_38); +lean_ctor_set(x_1, 2, x_92); +lean_ctor_set(x_1, 1, x_91); +lean_ctor_set(x_1, 0, x_40); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_97); +return x_1; +} +else +{ +lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; uint8_t x_102; lean_object* x_103; uint8_t x_104; +x_98 = lean_ctor_get(x_40, 0); +x_99 = lean_ctor_get(x_40, 1); +x_100 = lean_ctor_get(x_40, 2); +x_101 = lean_ctor_get(x_40, 3); +lean_inc(x_101); +lean_inc(x_100); +lean_inc(x_99); +lean_inc(x_98); +lean_dec(x_40); +x_102 = 1; +x_103 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_103, 0, x_98); +lean_ctor_set(x_103, 1, x_99); +lean_ctor_set(x_103, 2, x_100); +lean_ctor_set(x_103, 3, x_101); +lean_ctor_set_uint8(x_103, sizeof(void*)*4, x_102); +lean_ctor_set(x_38, 3, x_36); +lean_ctor_set(x_38, 2, x_35); +lean_ctor_set(x_38, 1, x_34); +lean_ctor_set(x_38, 0, x_93); +lean_ctor_set_uint8(x_38, sizeof(void*)*4, x_102); +x_104 = 0; +lean_ctor_set(x_1, 3, x_38); +lean_ctor_set(x_1, 2, x_92); +lean_ctor_set(x_1, 1, x_91); +lean_ctor_set(x_1, 0, x_103); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_104); +return x_1; +} +} +else +{ +lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; uint8_t x_113; lean_object* x_114; lean_object* x_115; uint8_t x_116; +x_105 = lean_ctor_get(x_38, 1); +x_106 = lean_ctor_get(x_38, 2); +x_107 = lean_ctor_get(x_38, 3); +lean_inc(x_107); +lean_inc(x_106); +lean_inc(x_105); +lean_dec(x_38); +x_108 = lean_ctor_get(x_40, 0); +lean_inc(x_108); +x_109 = lean_ctor_get(x_40, 1); +lean_inc(x_109); +x_110 = lean_ctor_get(x_40, 2); +lean_inc(x_110); +x_111 = lean_ctor_get(x_40, 3); +lean_inc(x_111); +if (lean_is_exclusive(x_40)) { + lean_ctor_release(x_40, 0); + lean_ctor_release(x_40, 1); + lean_ctor_release(x_40, 2); + lean_ctor_release(x_40, 3); + x_112 = x_40; +} else { + lean_dec_ref(x_40); + x_112 = lean_box(0); +} +x_113 = 1; +if (lean_is_scalar(x_112)) { + x_114 = lean_alloc_ctor(1, 4, 1); +} else { + x_114 = x_112; +} +lean_ctor_set(x_114, 0, x_108); +lean_ctor_set(x_114, 1, x_109); +lean_ctor_set(x_114, 2, x_110); +lean_ctor_set(x_114, 3, x_111); +lean_ctor_set_uint8(x_114, sizeof(void*)*4, x_113); +x_115 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_115, 0, x_107); +lean_ctor_set(x_115, 1, x_34); +lean_ctor_set(x_115, 2, x_35); +lean_ctor_set(x_115, 3, x_36); +lean_ctor_set_uint8(x_115, sizeof(void*)*4, x_113); +x_116 = 0; +lean_ctor_set(x_1, 3, x_115); +lean_ctor_set(x_1, 2, x_106); +lean_ctor_set(x_1, 1, x_105); +lean_ctor_set(x_1, 0, x_114); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_116); +return x_1; +} +} +else +{ +lean_object* x_117; +x_117 = lean_ctor_get(x_38, 3); +lean_inc(x_117); +if (lean_obj_tag(x_117) == 0) +{ +uint8_t x_118; +lean_free_object(x_1); +x_118 = !lean_is_exclusive(x_40); +if (x_118 == 0) +{ +lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; uint8_t x_123; +x_119 = lean_ctor_get(x_40, 3); +lean_dec(x_119); +x_120 = lean_ctor_get(x_40, 2); +lean_dec(x_120); +x_121 = lean_ctor_get(x_40, 1); +lean_dec(x_121); +x_122 = lean_ctor_get(x_40, 0); +lean_dec(x_122); +x_123 = 1; +lean_ctor_set(x_40, 3, x_36); +lean_ctor_set(x_40, 2, x_35); +lean_ctor_set(x_40, 1, x_34); +lean_ctor_set(x_40, 0, x_38); +lean_ctor_set_uint8(x_40, sizeof(void*)*4, x_123); +return x_40; +} +else +{ +uint8_t x_124; lean_object* x_125; +lean_dec(x_40); +x_124 = 1; +x_125 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_125, 0, x_38); +lean_ctor_set(x_125, 1, x_34); +lean_ctor_set(x_125, 2, x_35); +lean_ctor_set(x_125, 3, x_36); +lean_ctor_set_uint8(x_125, sizeof(void*)*4, x_124); +return x_125; +} +} +else +{ +uint8_t x_126; +x_126 = lean_ctor_get_uint8(x_117, sizeof(void*)*4); +if (x_126 == 0) +{ +uint8_t x_127; +lean_free_object(x_1); +x_127 = !lean_is_exclusive(x_38); +if (x_127 == 0) +{ +lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; uint8_t x_132; +x_128 = lean_ctor_get(x_38, 1); +x_129 = lean_ctor_get(x_38, 2); +x_130 = lean_ctor_get(x_38, 3); +lean_dec(x_130); +x_131 = lean_ctor_get(x_38, 0); +lean_dec(x_131); +x_132 = !lean_is_exclusive(x_117); +if (x_132 == 0) +{ +lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; uint8_t x_137; uint8_t x_138; +x_133 = lean_ctor_get(x_117, 0); +x_134 = lean_ctor_get(x_117, 1); +x_135 = lean_ctor_get(x_117, 2); +x_136 = lean_ctor_get(x_117, 3); +x_137 = 1; +lean_inc(x_40); +lean_ctor_set(x_117, 3, x_133); +lean_ctor_set(x_117, 2, x_129); +lean_ctor_set(x_117, 1, x_128); +lean_ctor_set(x_117, 0, x_40); +x_138 = !lean_is_exclusive(x_40); +if (x_138 == 0) +{ +lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; uint8_t x_143; +x_139 = lean_ctor_get(x_40, 3); +lean_dec(x_139); +x_140 = lean_ctor_get(x_40, 2); +lean_dec(x_140); +x_141 = lean_ctor_get(x_40, 1); +lean_dec(x_141); +x_142 = lean_ctor_get(x_40, 0); +lean_dec(x_142); +lean_ctor_set_uint8(x_117, sizeof(void*)*4, x_137); +lean_ctor_set(x_40, 3, x_36); +lean_ctor_set(x_40, 2, x_35); +lean_ctor_set(x_40, 1, x_34); +lean_ctor_set(x_40, 0, x_136); +lean_ctor_set_uint8(x_40, sizeof(void*)*4, x_137); +x_143 = 0; +lean_ctor_set(x_38, 3, x_40); +lean_ctor_set(x_38, 2, x_135); +lean_ctor_set(x_38, 1, x_134); +lean_ctor_set(x_38, 0, x_117); +lean_ctor_set_uint8(x_38, sizeof(void*)*4, x_143); +return x_38; +} +else +{ +lean_object* x_144; uint8_t x_145; +lean_dec(x_40); +lean_ctor_set_uint8(x_117, sizeof(void*)*4, x_137); +x_144 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_144, 0, x_136); +lean_ctor_set(x_144, 1, x_34); +lean_ctor_set(x_144, 2, x_35); +lean_ctor_set(x_144, 3, x_36); +lean_ctor_set_uint8(x_144, sizeof(void*)*4, x_137); +x_145 = 0; +lean_ctor_set(x_38, 3, x_144); +lean_ctor_set(x_38, 2, x_135); +lean_ctor_set(x_38, 1, x_134); +lean_ctor_set(x_38, 0, x_117); +lean_ctor_set_uint8(x_38, sizeof(void*)*4, x_145); +return x_38; +} +} +else +{ +lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; uint8_t x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; uint8_t x_154; +x_146 = lean_ctor_get(x_117, 0); +x_147 = lean_ctor_get(x_117, 1); +x_148 = lean_ctor_get(x_117, 2); +x_149 = lean_ctor_get(x_117, 3); +lean_inc(x_149); +lean_inc(x_148); +lean_inc(x_147); +lean_inc(x_146); +lean_dec(x_117); +x_150 = 1; lean_inc(x_40); -if (lean_obj_tag(x_40) == 0) +x_151 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_151, 0, x_40); +lean_ctor_set(x_151, 1, x_128); +lean_ctor_set(x_151, 2, x_129); +lean_ctor_set(x_151, 3, x_146); +if (lean_is_exclusive(x_40)) { + lean_ctor_release(x_40, 0); + lean_ctor_release(x_40, 1); + lean_ctor_release(x_40, 2); + lean_ctor_release(x_40, 3); + x_152 = x_40; +} else { + lean_dec_ref(x_40); + x_152 = lean_box(0); +} +lean_ctor_set_uint8(x_151, sizeof(void*)*4, x_150); +if (lean_is_scalar(x_152)) { + x_153 = lean_alloc_ctor(1, 4, 1); +} else { + x_153 = x_152; +} +lean_ctor_set(x_153, 0, x_149); +lean_ctor_set(x_153, 1, x_34); +lean_ctor_set(x_153, 2, x_35); +lean_ctor_set(x_153, 3, x_36); +lean_ctor_set_uint8(x_153, sizeof(void*)*4, x_150); +x_154 = 0; +lean_ctor_set(x_38, 3, x_153); +lean_ctor_set(x_38, 2, x_148); +lean_ctor_set(x_38, 1, x_147); +lean_ctor_set(x_38, 0, x_151); +lean_ctor_set_uint8(x_38, sizeof(void*)*4, x_154); +return x_38; +} +} +else +{ +lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; uint8_t x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; uint8_t x_166; lean_object* x_167; +x_155 = lean_ctor_get(x_38, 1); +x_156 = lean_ctor_get(x_38, 2); +lean_inc(x_156); +lean_inc(x_155); +lean_dec(x_38); +x_157 = lean_ctor_get(x_117, 0); +lean_inc(x_157); +x_158 = lean_ctor_get(x_117, 1); +lean_inc(x_158); +x_159 = lean_ctor_get(x_117, 2); +lean_inc(x_159); +x_160 = lean_ctor_get(x_117, 3); +lean_inc(x_160); +if (lean_is_exclusive(x_117)) { + lean_ctor_release(x_117, 0); + lean_ctor_release(x_117, 1); + lean_ctor_release(x_117, 2); + lean_ctor_release(x_117, 3); + x_161 = x_117; +} else { + lean_dec_ref(x_117); + x_161 = lean_box(0); +} +x_162 = 1; +lean_inc(x_40); +if (lean_is_scalar(x_161)) { + x_163 = lean_alloc_ctor(1, 4, 1); +} else { + x_163 = x_161; +} +lean_ctor_set(x_163, 0, x_40); +lean_ctor_set(x_163, 1, x_155); +lean_ctor_set(x_163, 2, x_156); +lean_ctor_set(x_163, 3, x_157); +if (lean_is_exclusive(x_40)) { + lean_ctor_release(x_40, 0); + lean_ctor_release(x_40, 1); + lean_ctor_release(x_40, 2); + lean_ctor_release(x_40, 3); + x_164 = x_40; +} else { + lean_dec_ref(x_40); + x_164 = lean_box(0); +} +lean_ctor_set_uint8(x_163, sizeof(void*)*4, x_162); +if (lean_is_scalar(x_164)) { + x_165 = lean_alloc_ctor(1, 4, 1); +} else { + x_165 = x_164; +} +lean_ctor_set(x_165, 0, x_160); +lean_ctor_set(x_165, 1, x_34); +lean_ctor_set(x_165, 2, x_35); +lean_ctor_set(x_165, 3, x_36); +lean_ctor_set_uint8(x_165, sizeof(void*)*4, x_162); +x_166 = 0; +x_167 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_167, 0, x_163); +lean_ctor_set(x_167, 1, x_158); +lean_ctor_set(x_167, 2, x_159); +lean_ctor_set(x_167, 3, x_165); +lean_ctor_set_uint8(x_167, sizeof(void*)*4, x_166); +return x_167; +} +} +else +{ +uint8_t x_168; +x_168 = !lean_is_exclusive(x_38); +if (x_168 == 0) +{ +lean_object* x_169; lean_object* x_170; uint8_t x_171; +x_169 = lean_ctor_get(x_38, 3); +lean_dec(x_169); +x_170 = lean_ctor_get(x_38, 0); +lean_dec(x_170); +x_171 = !lean_is_exclusive(x_40); +if (x_171 == 0) +{ +uint8_t x_172; +lean_ctor_set_uint8(x_40, sizeof(void*)*4, x_126); +x_172 = 1; +lean_ctor_set(x_1, 0, x_38); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_172); +return x_1; +} +else +{ +lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; uint8_t x_178; +x_173 = lean_ctor_get(x_40, 0); +x_174 = lean_ctor_get(x_40, 1); +x_175 = lean_ctor_get(x_40, 2); +x_176 = lean_ctor_get(x_40, 3); +lean_inc(x_176); +lean_inc(x_175); +lean_inc(x_174); +lean_inc(x_173); +lean_dec(x_40); +x_177 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_177, 0, x_173); +lean_ctor_set(x_177, 1, x_174); +lean_ctor_set(x_177, 2, x_175); +lean_ctor_set(x_177, 3, x_176); +lean_ctor_set_uint8(x_177, sizeof(void*)*4, x_126); +lean_ctor_set(x_38, 0, x_177); +x_178 = 1; +lean_ctor_set(x_1, 0, x_38); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_178); +return x_1; +} +} +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; uint8_t x_188; +x_179 = lean_ctor_get(x_38, 1); +x_180 = lean_ctor_get(x_38, 2); +lean_inc(x_180); +lean_inc(x_179); +lean_dec(x_38); +x_181 = lean_ctor_get(x_40, 0); +lean_inc(x_181); +x_182 = lean_ctor_get(x_40, 1); +lean_inc(x_182); +x_183 = lean_ctor_get(x_40, 2); +lean_inc(x_183); +x_184 = lean_ctor_get(x_40, 3); +lean_inc(x_184); +if (lean_is_exclusive(x_40)) { + lean_ctor_release(x_40, 0); + lean_ctor_release(x_40, 1); + lean_ctor_release(x_40, 2); + lean_ctor_release(x_40, 3); + x_185 = x_40; +} else { + lean_dec_ref(x_40); + x_185 = lean_box(0); +} +if (lean_is_scalar(x_185)) { + x_186 = lean_alloc_ctor(1, 4, 1); +} else { + x_186 = x_185; +} +lean_ctor_set(x_186, 0, x_181); +lean_ctor_set(x_186, 1, x_182); +lean_ctor_set(x_186, 2, x_183); +lean_ctor_set(x_186, 3, x_184); +lean_ctor_set_uint8(x_186, sizeof(void*)*4, x_126); +x_187 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_187, 0, x_186); +lean_ctor_set(x_187, 1, x_179); +lean_ctor_set(x_187, 2, x_180); +lean_ctor_set(x_187, 3, x_117); +lean_ctor_set_uint8(x_187, sizeof(void*)*4, x_39); +x_188 = 1; +lean_ctor_set(x_1, 0, x_187); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_188); +return x_1; +} +} +} +} +} +} +else +{ +uint8_t x_189; +x_189 = 1; +lean_ctor_set(x_1, 0, x_38); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_189); +return x_1; +} +} +case 1: +{ +uint8_t x_190; +lean_dec(x_35); +lean_dec(x_34); +x_190 = 1; +lean_ctor_set(x_1, 2, x_3); +lean_ctor_set(x_1, 1, x_2); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_190); +return x_1; +} +default: +{ +lean_object* x_191; uint8_t x_192; +x_191 = l_Lean_RBNode_ins___at_Lean_Meta_Grind_registerParent___spec__2(x_36, x_2, x_3); +x_192 = lean_ctor_get_uint8(x_191, sizeof(void*)*4); +if (x_192 == 0) +{ +lean_object* x_193; +x_193 = lean_ctor_get(x_191, 0); +lean_inc(x_193); +if (lean_obj_tag(x_193) == 0) +{ +lean_object* x_194; +x_194 = lean_ctor_get(x_191, 3); +lean_inc(x_194); +if (lean_obj_tag(x_194) == 0) +{ +uint8_t x_195; +x_195 = !lean_is_exclusive(x_191); +if (x_195 == 0) +{ +lean_object* x_196; lean_object* x_197; uint8_t x_198; +x_196 = lean_ctor_get(x_191, 3); +lean_dec(x_196); +x_197 = lean_ctor_get(x_191, 0); +lean_dec(x_197); +lean_ctor_set(x_191, 0, x_194); +x_198 = 1; +lean_ctor_set(x_1, 3, x_191); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_198); +return x_1; +} +else +{ +lean_object* x_199; lean_object* x_200; lean_object* x_201; uint8_t x_202; +x_199 = lean_ctor_get(x_191, 1); +x_200 = lean_ctor_get(x_191, 2); +lean_inc(x_200); +lean_inc(x_199); +lean_dec(x_191); +x_201 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_201, 0, x_194); +lean_ctor_set(x_201, 1, x_199); +lean_ctor_set(x_201, 2, x_200); +lean_ctor_set(x_201, 3, x_194); +lean_ctor_set_uint8(x_201, sizeof(void*)*4, x_192); +x_202 = 1; +lean_ctor_set(x_1, 3, x_201); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_202); +return x_1; +} +} +else +{ +uint8_t x_203; +x_203 = lean_ctor_get_uint8(x_194, sizeof(void*)*4); +if (x_203 == 0) +{ +uint8_t x_204; +x_204 = !lean_is_exclusive(x_191); +if (x_204 == 0) +{ +lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; uint8_t x_209; +x_205 = lean_ctor_get(x_191, 1); +x_206 = lean_ctor_get(x_191, 2); +x_207 = lean_ctor_get(x_191, 3); +lean_dec(x_207); +x_208 = lean_ctor_get(x_191, 0); +lean_dec(x_208); +x_209 = !lean_is_exclusive(x_194); +if (x_209 == 0) +{ +lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; uint8_t x_214; uint8_t x_215; +x_210 = lean_ctor_get(x_194, 0); +x_211 = lean_ctor_get(x_194, 1); +x_212 = lean_ctor_get(x_194, 2); +x_213 = lean_ctor_get(x_194, 3); +x_214 = 1; +lean_ctor_set(x_194, 3, x_193); +lean_ctor_set(x_194, 2, x_35); +lean_ctor_set(x_194, 1, x_34); +lean_ctor_set(x_194, 0, x_33); +lean_ctor_set_uint8(x_194, sizeof(void*)*4, x_214); +lean_ctor_set(x_191, 3, x_213); +lean_ctor_set(x_191, 2, x_212); +lean_ctor_set(x_191, 1, x_211); +lean_ctor_set(x_191, 0, x_210); +lean_ctor_set_uint8(x_191, sizeof(void*)*4, x_214); +x_215 = 0; +lean_ctor_set(x_1, 3, x_191); +lean_ctor_set(x_1, 2, x_206); +lean_ctor_set(x_1, 1, x_205); +lean_ctor_set(x_1, 0, x_194); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_215); +return x_1; +} +else +{ +lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; uint8_t x_220; lean_object* x_221; uint8_t x_222; +x_216 = lean_ctor_get(x_194, 0); +x_217 = lean_ctor_get(x_194, 1); +x_218 = lean_ctor_get(x_194, 2); +x_219 = lean_ctor_get(x_194, 3); +lean_inc(x_219); +lean_inc(x_218); +lean_inc(x_217); +lean_inc(x_216); +lean_dec(x_194); +x_220 = 1; +x_221 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_221, 0, x_33); +lean_ctor_set(x_221, 1, x_34); +lean_ctor_set(x_221, 2, x_35); +lean_ctor_set(x_221, 3, x_193); +lean_ctor_set_uint8(x_221, sizeof(void*)*4, x_220); +lean_ctor_set(x_191, 3, x_219); +lean_ctor_set(x_191, 2, x_218); +lean_ctor_set(x_191, 1, x_217); +lean_ctor_set(x_191, 0, x_216); +lean_ctor_set_uint8(x_191, sizeof(void*)*4, x_220); +x_222 = 0; +lean_ctor_set(x_1, 3, x_191); +lean_ctor_set(x_1, 2, x_206); +lean_ctor_set(x_1, 1, x_205); +lean_ctor_set(x_1, 0, x_221); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_222); +return x_1; +} +} +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; uint8_t x_230; lean_object* x_231; lean_object* x_232; uint8_t x_233; +x_223 = lean_ctor_get(x_191, 1); +x_224 = lean_ctor_get(x_191, 2); +lean_inc(x_224); +lean_inc(x_223); +lean_dec(x_191); +x_225 = lean_ctor_get(x_194, 0); +lean_inc(x_225); +x_226 = lean_ctor_get(x_194, 1); +lean_inc(x_226); +x_227 = lean_ctor_get(x_194, 2); +lean_inc(x_227); +x_228 = lean_ctor_get(x_194, 3); +lean_inc(x_228); +if (lean_is_exclusive(x_194)) { + lean_ctor_release(x_194, 0); + lean_ctor_release(x_194, 1); + lean_ctor_release(x_194, 2); + lean_ctor_release(x_194, 3); + x_229 = x_194; +} else { + lean_dec_ref(x_194); + x_229 = lean_box(0); +} +x_230 = 1; +if (lean_is_scalar(x_229)) { + x_231 = lean_alloc_ctor(1, 4, 1); +} else { + x_231 = x_229; +} +lean_ctor_set(x_231, 0, x_33); +lean_ctor_set(x_231, 1, x_34); +lean_ctor_set(x_231, 2, x_35); +lean_ctor_set(x_231, 3, x_193); +lean_ctor_set_uint8(x_231, sizeof(void*)*4, x_230); +x_232 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_232, 0, x_225); +lean_ctor_set(x_232, 1, x_226); +lean_ctor_set(x_232, 2, x_227); +lean_ctor_set(x_232, 3, x_228); +lean_ctor_set_uint8(x_232, sizeof(void*)*4, x_230); +x_233 = 0; +lean_ctor_set(x_1, 3, x_232); +lean_ctor_set(x_1, 2, x_224); +lean_ctor_set(x_1, 1, x_223); +lean_ctor_set(x_1, 0, x_231); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_233); +return x_1; +} +} +else +{ +uint8_t x_234; +lean_free_object(x_1); +x_234 = !lean_is_exclusive(x_194); +if (x_234 == 0) +{ +lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; uint8_t x_239; +x_235 = lean_ctor_get(x_194, 3); +lean_dec(x_235); +x_236 = lean_ctor_get(x_194, 2); +lean_dec(x_236); +x_237 = lean_ctor_get(x_194, 1); +lean_dec(x_237); +x_238 = lean_ctor_get(x_194, 0); +lean_dec(x_238); +x_239 = 1; +lean_ctor_set(x_194, 3, x_191); +lean_ctor_set(x_194, 2, x_35); +lean_ctor_set(x_194, 1, x_34); +lean_ctor_set(x_194, 0, x_33); +lean_ctor_set_uint8(x_194, sizeof(void*)*4, x_239); +return x_194; +} +else +{ +uint8_t x_240; lean_object* x_241; +lean_dec(x_194); +x_240 = 1; +x_241 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_241, 0, x_33); +lean_ctor_set(x_241, 1, x_34); +lean_ctor_set(x_241, 2, x_35); +lean_ctor_set(x_241, 3, x_191); +lean_ctor_set_uint8(x_241, sizeof(void*)*4, x_240); +return x_241; +} +} +} +} +else +{ +uint8_t x_242; +x_242 = lean_ctor_get_uint8(x_193, sizeof(void*)*4); +if (x_242 == 0) +{ +uint8_t x_243; +x_243 = !lean_is_exclusive(x_191); +if (x_243 == 0) +{ +lean_object* x_244; uint8_t x_245; +x_244 = lean_ctor_get(x_191, 0); +lean_dec(x_244); +x_245 = !lean_is_exclusive(x_193); +if (x_245 == 0) +{ +lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; uint8_t x_250; uint8_t x_251; +x_246 = lean_ctor_get(x_193, 0); +x_247 = lean_ctor_get(x_193, 1); +x_248 = lean_ctor_get(x_193, 2); +x_249 = lean_ctor_get(x_193, 3); +x_250 = 1; +lean_ctor_set(x_193, 3, x_246); +lean_ctor_set(x_193, 2, x_35); +lean_ctor_set(x_193, 1, x_34); +lean_ctor_set(x_193, 0, x_33); +lean_ctor_set_uint8(x_193, sizeof(void*)*4, x_250); +lean_ctor_set(x_191, 0, x_249); +lean_ctor_set_uint8(x_191, sizeof(void*)*4, x_250); +x_251 = 0; +lean_ctor_set(x_1, 3, x_191); +lean_ctor_set(x_1, 2, x_248); +lean_ctor_set(x_1, 1, x_247); +lean_ctor_set(x_1, 0, x_193); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_251); +return x_1; +} +else +{ +lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; uint8_t x_256; lean_object* x_257; uint8_t x_258; +x_252 = lean_ctor_get(x_193, 0); +x_253 = lean_ctor_get(x_193, 1); +x_254 = lean_ctor_get(x_193, 2); +x_255 = lean_ctor_get(x_193, 3); +lean_inc(x_255); +lean_inc(x_254); +lean_inc(x_253); +lean_inc(x_252); +lean_dec(x_193); +x_256 = 1; +x_257 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_257, 0, x_33); +lean_ctor_set(x_257, 1, x_34); +lean_ctor_set(x_257, 2, x_35); +lean_ctor_set(x_257, 3, x_252); +lean_ctor_set_uint8(x_257, sizeof(void*)*4, x_256); +lean_ctor_set(x_191, 0, x_255); +lean_ctor_set_uint8(x_191, sizeof(void*)*4, x_256); +x_258 = 0; +lean_ctor_set(x_1, 3, x_191); +lean_ctor_set(x_1, 2, x_254); +lean_ctor_set(x_1, 1, x_253); +lean_ctor_set(x_1, 0, x_257); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_258); +return x_1; +} +} +else +{ +lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; uint8_t x_267; lean_object* x_268; lean_object* x_269; uint8_t x_270; +x_259 = lean_ctor_get(x_191, 1); +x_260 = lean_ctor_get(x_191, 2); +x_261 = lean_ctor_get(x_191, 3); +lean_inc(x_261); +lean_inc(x_260); +lean_inc(x_259); +lean_dec(x_191); +x_262 = lean_ctor_get(x_193, 0); +lean_inc(x_262); +x_263 = lean_ctor_get(x_193, 1); +lean_inc(x_263); +x_264 = lean_ctor_get(x_193, 2); +lean_inc(x_264); +x_265 = lean_ctor_get(x_193, 3); +lean_inc(x_265); +if (lean_is_exclusive(x_193)) { + lean_ctor_release(x_193, 0); + lean_ctor_release(x_193, 1); + lean_ctor_release(x_193, 2); + lean_ctor_release(x_193, 3); + x_266 = x_193; +} else { + lean_dec_ref(x_193); + x_266 = lean_box(0); +} +x_267 = 1; +if (lean_is_scalar(x_266)) { + x_268 = lean_alloc_ctor(1, 4, 1); +} else { + x_268 = x_266; +} +lean_ctor_set(x_268, 0, x_33); +lean_ctor_set(x_268, 1, x_34); +lean_ctor_set(x_268, 2, x_35); +lean_ctor_set(x_268, 3, x_262); +lean_ctor_set_uint8(x_268, sizeof(void*)*4, x_267); +x_269 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_269, 0, x_265); +lean_ctor_set(x_269, 1, x_259); +lean_ctor_set(x_269, 2, x_260); +lean_ctor_set(x_269, 3, x_261); +lean_ctor_set_uint8(x_269, sizeof(void*)*4, x_267); +x_270 = 0; +lean_ctor_set(x_1, 3, x_269); +lean_ctor_set(x_1, 2, x_264); +lean_ctor_set(x_1, 1, x_263); +lean_ctor_set(x_1, 0, x_268); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_270); +return x_1; +} +} +else +{ +lean_object* x_271; +x_271 = lean_ctor_get(x_191, 3); +lean_inc(x_271); +if (lean_obj_tag(x_271) == 0) +{ +uint8_t x_272; +lean_free_object(x_1); +x_272 = !lean_is_exclusive(x_193); +if (x_272 == 0) +{ +lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; uint8_t x_277; +x_273 = lean_ctor_get(x_193, 3); +lean_dec(x_273); +x_274 = lean_ctor_get(x_193, 2); +lean_dec(x_274); +x_275 = lean_ctor_get(x_193, 1); +lean_dec(x_275); +x_276 = lean_ctor_get(x_193, 0); +lean_dec(x_276); +x_277 = 1; +lean_ctor_set(x_193, 3, x_191); +lean_ctor_set(x_193, 2, x_35); +lean_ctor_set(x_193, 1, x_34); +lean_ctor_set(x_193, 0, x_33); +lean_ctor_set_uint8(x_193, sizeof(void*)*4, x_277); +return x_193; +} +else +{ +uint8_t x_278; lean_object* x_279; +lean_dec(x_193); +x_278 = 1; +x_279 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_279, 0, x_33); +lean_ctor_set(x_279, 1, x_34); +lean_ctor_set(x_279, 2, x_35); +lean_ctor_set(x_279, 3, x_191); +lean_ctor_set_uint8(x_279, sizeof(void*)*4, x_278); +return x_279; +} +} +else +{ +uint8_t x_280; +x_280 = lean_ctor_get_uint8(x_271, sizeof(void*)*4); +if (x_280 == 0) +{ +uint8_t x_281; +lean_free_object(x_1); +x_281 = !lean_is_exclusive(x_191); +if (x_281 == 0) +{ +lean_object* x_282; lean_object* x_283; uint8_t x_284; +x_282 = lean_ctor_get(x_191, 3); +lean_dec(x_282); +x_283 = lean_ctor_get(x_191, 0); +lean_dec(x_283); +x_284 = !lean_is_exclusive(x_271); +if (x_284 == 0) +{ +lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; uint8_t x_289; uint8_t x_290; +x_285 = lean_ctor_get(x_271, 0); +x_286 = lean_ctor_get(x_271, 1); +x_287 = lean_ctor_get(x_271, 2); +x_288 = lean_ctor_get(x_271, 3); +x_289 = 1; +lean_inc(x_193); +lean_ctor_set(x_271, 3, x_193); +lean_ctor_set(x_271, 2, x_35); +lean_ctor_set(x_271, 1, x_34); +lean_ctor_set(x_271, 0, x_33); +x_290 = !lean_is_exclusive(x_193); +if (x_290 == 0) +{ +lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; uint8_t x_295; +x_291 = lean_ctor_get(x_193, 3); +lean_dec(x_291); +x_292 = lean_ctor_get(x_193, 2); +lean_dec(x_292); +x_293 = lean_ctor_get(x_193, 1); +lean_dec(x_293); +x_294 = lean_ctor_get(x_193, 0); +lean_dec(x_294); +lean_ctor_set_uint8(x_271, sizeof(void*)*4, x_289); +lean_ctor_set(x_193, 3, x_288); +lean_ctor_set(x_193, 2, x_287); +lean_ctor_set(x_193, 1, x_286); +lean_ctor_set(x_193, 0, x_285); +lean_ctor_set_uint8(x_193, sizeof(void*)*4, x_289); +x_295 = 0; +lean_ctor_set(x_191, 3, x_193); +lean_ctor_set(x_191, 0, x_271); +lean_ctor_set_uint8(x_191, sizeof(void*)*4, x_295); +return x_191; +} +else +{ +lean_object* x_296; uint8_t x_297; +lean_dec(x_193); +lean_ctor_set_uint8(x_271, sizeof(void*)*4, x_289); +x_296 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_296, 0, x_285); +lean_ctor_set(x_296, 1, x_286); +lean_ctor_set(x_296, 2, x_287); +lean_ctor_set(x_296, 3, x_288); +lean_ctor_set_uint8(x_296, sizeof(void*)*4, x_289); +x_297 = 0; +lean_ctor_set(x_191, 3, x_296); +lean_ctor_set(x_191, 0, x_271); +lean_ctor_set_uint8(x_191, sizeof(void*)*4, x_297); +return x_191; +} +} +else +{ +lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; uint8_t x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; uint8_t x_306; +x_298 = lean_ctor_get(x_271, 0); +x_299 = lean_ctor_get(x_271, 1); +x_300 = lean_ctor_get(x_271, 2); +x_301 = lean_ctor_get(x_271, 3); +lean_inc(x_301); +lean_inc(x_300); +lean_inc(x_299); +lean_inc(x_298); +lean_dec(x_271); +x_302 = 1; +lean_inc(x_193); +x_303 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_303, 0, x_33); +lean_ctor_set(x_303, 1, x_34); +lean_ctor_set(x_303, 2, x_35); +lean_ctor_set(x_303, 3, x_193); +if (lean_is_exclusive(x_193)) { + lean_ctor_release(x_193, 0); + lean_ctor_release(x_193, 1); + lean_ctor_release(x_193, 2); + lean_ctor_release(x_193, 3); + x_304 = x_193; +} else { + lean_dec_ref(x_193); + x_304 = lean_box(0); +} +lean_ctor_set_uint8(x_303, sizeof(void*)*4, x_302); +if (lean_is_scalar(x_304)) { + x_305 = lean_alloc_ctor(1, 4, 1); +} else { + x_305 = x_304; +} +lean_ctor_set(x_305, 0, x_298); +lean_ctor_set(x_305, 1, x_299); +lean_ctor_set(x_305, 2, x_300); +lean_ctor_set(x_305, 3, x_301); +lean_ctor_set_uint8(x_305, sizeof(void*)*4, x_302); +x_306 = 0; +lean_ctor_set(x_191, 3, x_305); +lean_ctor_set(x_191, 0, x_303); +lean_ctor_set_uint8(x_191, sizeof(void*)*4, x_306); +return x_191; +} +} +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; uint8_t x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; uint8_t x_318; lean_object* x_319; +x_307 = lean_ctor_get(x_191, 1); +x_308 = lean_ctor_get(x_191, 2); +lean_inc(x_308); +lean_inc(x_307); +lean_dec(x_191); +x_309 = lean_ctor_get(x_271, 0); +lean_inc(x_309); +x_310 = lean_ctor_get(x_271, 1); +lean_inc(x_310); +x_311 = lean_ctor_get(x_271, 2); +lean_inc(x_311); +x_312 = lean_ctor_get(x_271, 3); +lean_inc(x_312); +if (lean_is_exclusive(x_271)) { + lean_ctor_release(x_271, 0); + lean_ctor_release(x_271, 1); + lean_ctor_release(x_271, 2); + lean_ctor_release(x_271, 3); + x_313 = x_271; +} else { + lean_dec_ref(x_271); + x_313 = lean_box(0); +} +x_314 = 1; +lean_inc(x_193); +if (lean_is_scalar(x_313)) { + x_315 = lean_alloc_ctor(1, 4, 1); +} else { + x_315 = x_313; +} +lean_ctor_set(x_315, 0, x_33); +lean_ctor_set(x_315, 1, x_34); +lean_ctor_set(x_315, 2, x_35); +lean_ctor_set(x_315, 3, x_193); +if (lean_is_exclusive(x_193)) { + lean_ctor_release(x_193, 0); + lean_ctor_release(x_193, 1); + lean_ctor_release(x_193, 2); + lean_ctor_release(x_193, 3); + x_316 = x_193; +} else { + lean_dec_ref(x_193); + x_316 = lean_box(0); +} +lean_ctor_set_uint8(x_315, sizeof(void*)*4, x_314); +if (lean_is_scalar(x_316)) { + x_317 = lean_alloc_ctor(1, 4, 1); +} else { + x_317 = x_316; +} +lean_ctor_set(x_317, 0, x_309); +lean_ctor_set(x_317, 1, x_310); +lean_ctor_set(x_317, 2, x_311); +lean_ctor_set(x_317, 3, x_312); +lean_ctor_set_uint8(x_317, sizeof(void*)*4, x_314); +x_318 = 0; +x_319 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_319, 0, x_315); +lean_ctor_set(x_319, 1, x_307); +lean_ctor_set(x_319, 2, x_308); +lean_ctor_set(x_319, 3, x_317); +lean_ctor_set_uint8(x_319, sizeof(void*)*4, x_318); +return x_319; +} +} +else +{ +uint8_t x_320; +x_320 = !lean_is_exclusive(x_191); +if (x_320 == 0) +{ +lean_object* x_321; lean_object* x_322; uint8_t x_323; +x_321 = lean_ctor_get(x_191, 3); +lean_dec(x_321); +x_322 = lean_ctor_get(x_191, 0); +lean_dec(x_322); +x_323 = !lean_is_exclusive(x_193); +if (x_323 == 0) +{ +uint8_t x_324; +lean_ctor_set_uint8(x_193, sizeof(void*)*4, x_280); +x_324 = 1; +lean_ctor_set(x_1, 3, x_191); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_324); +return x_1; +} +else +{ +lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; uint8_t x_330; +x_325 = lean_ctor_get(x_193, 0); +x_326 = lean_ctor_get(x_193, 1); +x_327 = lean_ctor_get(x_193, 2); +x_328 = lean_ctor_get(x_193, 3); +lean_inc(x_328); +lean_inc(x_327); +lean_inc(x_326); +lean_inc(x_325); +lean_dec(x_193); +x_329 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_329, 0, x_325); +lean_ctor_set(x_329, 1, x_326); +lean_ctor_set(x_329, 2, x_327); +lean_ctor_set(x_329, 3, x_328); +lean_ctor_set_uint8(x_329, sizeof(void*)*4, x_280); +lean_ctor_set(x_191, 0, x_329); +x_330 = 1; +lean_ctor_set(x_1, 3, x_191); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_330); +return x_1; +} +} +else +{ +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; uint8_t x_340; +x_331 = lean_ctor_get(x_191, 1); +x_332 = lean_ctor_get(x_191, 2); +lean_inc(x_332); +lean_inc(x_331); +lean_dec(x_191); +x_333 = lean_ctor_get(x_193, 0); +lean_inc(x_333); +x_334 = lean_ctor_get(x_193, 1); +lean_inc(x_334); +x_335 = lean_ctor_get(x_193, 2); +lean_inc(x_335); +x_336 = lean_ctor_get(x_193, 3); +lean_inc(x_336); +if (lean_is_exclusive(x_193)) { + lean_ctor_release(x_193, 0); + lean_ctor_release(x_193, 1); + lean_ctor_release(x_193, 2); + lean_ctor_release(x_193, 3); + x_337 = x_193; +} else { + lean_dec_ref(x_193); + x_337 = lean_box(0); +} +if (lean_is_scalar(x_337)) { + x_338 = lean_alloc_ctor(1, 4, 1); +} else { + x_338 = x_337; +} +lean_ctor_set(x_338, 0, x_333); +lean_ctor_set(x_338, 1, x_334); +lean_ctor_set(x_338, 2, x_335); +lean_ctor_set(x_338, 3, x_336); +lean_ctor_set_uint8(x_338, sizeof(void*)*4, x_280); +x_339 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_339, 0, x_338); +lean_ctor_set(x_339, 1, x_331); +lean_ctor_set(x_339, 2, x_332); +lean_ctor_set(x_339, 3, x_271); +lean_ctor_set_uint8(x_339, sizeof(void*)*4, x_192); +x_340 = 1; +lean_ctor_set(x_1, 3, x_339); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_340); +return x_1; +} +} +} +} +} +} +else +{ +uint8_t x_341; +x_341 = 1; +lean_ctor_set(x_1, 3, x_191); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_341); +return x_1; +} +} +} +} +else +{ +lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; uint8_t x_346; +x_342 = lean_ctor_get(x_1, 0); +x_343 = lean_ctor_get(x_1, 1); +x_344 = lean_ctor_get(x_1, 2); +x_345 = lean_ctor_get(x_1, 3); +lean_inc(x_345); +lean_inc(x_344); +lean_inc(x_343); +lean_inc(x_342); +lean_dec(x_1); +x_346 = l_Lean_Expr_quickComp(x_2, x_343); +switch (x_346) { +case 0: +{ +lean_object* x_347; uint8_t x_348; +x_347 = l_Lean_RBNode_ins___at_Lean_Meta_Grind_registerParent___spec__2(x_342, x_2, x_3); +x_348 = lean_ctor_get_uint8(x_347, sizeof(void*)*4); +if (x_348 == 0) +{ +lean_object* x_349; +x_349 = lean_ctor_get(x_347, 0); +lean_inc(x_349); +if (lean_obj_tag(x_349) == 0) +{ +lean_object* x_350; +x_350 = lean_ctor_get(x_347, 3); +lean_inc(x_350); +if (lean_obj_tag(x_350) == 0) +{ +lean_object* x_351; lean_object* x_352; lean_object* x_353; lean_object* x_354; uint8_t x_355; lean_object* x_356; +x_351 = lean_ctor_get(x_347, 1); +lean_inc(x_351); +x_352 = lean_ctor_get(x_347, 2); +lean_inc(x_352); +if (lean_is_exclusive(x_347)) { + lean_ctor_release(x_347, 0); + lean_ctor_release(x_347, 1); + lean_ctor_release(x_347, 2); + lean_ctor_release(x_347, 3); + x_353 = x_347; +} else { + lean_dec_ref(x_347); + x_353 = lean_box(0); +} +if (lean_is_scalar(x_353)) { + x_354 = lean_alloc_ctor(1, 4, 1); +} else { + x_354 = x_353; +} +lean_ctor_set(x_354, 0, x_350); +lean_ctor_set(x_354, 1, x_351); +lean_ctor_set(x_354, 2, x_352); +lean_ctor_set(x_354, 3, x_350); +lean_ctor_set_uint8(x_354, sizeof(void*)*4, x_348); +x_355 = 1; +x_356 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_356, 0, x_354); +lean_ctor_set(x_356, 1, x_343); +lean_ctor_set(x_356, 2, x_344); +lean_ctor_set(x_356, 3, x_345); +lean_ctor_set_uint8(x_356, sizeof(void*)*4, x_355); +return x_356; +} +else +{ +uint8_t x_357; +x_357 = lean_ctor_get_uint8(x_350, sizeof(void*)*4); +if (x_357 == 0) +{ +lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; uint8_t x_366; lean_object* x_367; lean_object* x_368; uint8_t x_369; lean_object* x_370; +x_358 = lean_ctor_get(x_347, 1); +lean_inc(x_358); +x_359 = lean_ctor_get(x_347, 2); +lean_inc(x_359); +if (lean_is_exclusive(x_347)) { + lean_ctor_release(x_347, 0); + lean_ctor_release(x_347, 1); + lean_ctor_release(x_347, 2); + lean_ctor_release(x_347, 3); + x_360 = x_347; +} else { + lean_dec_ref(x_347); + x_360 = lean_box(0); +} +x_361 = lean_ctor_get(x_350, 0); +lean_inc(x_361); +x_362 = lean_ctor_get(x_350, 1); +lean_inc(x_362); +x_363 = lean_ctor_get(x_350, 2); +lean_inc(x_363); +x_364 = lean_ctor_get(x_350, 3); +lean_inc(x_364); +if (lean_is_exclusive(x_350)) { + lean_ctor_release(x_350, 0); + lean_ctor_release(x_350, 1); + lean_ctor_release(x_350, 2); + lean_ctor_release(x_350, 3); + x_365 = x_350; +} else { + lean_dec_ref(x_350); + x_365 = lean_box(0); +} +x_366 = 1; +if (lean_is_scalar(x_365)) { + x_367 = lean_alloc_ctor(1, 4, 1); +} else { + x_367 = x_365; +} +lean_ctor_set(x_367, 0, x_349); +lean_ctor_set(x_367, 1, x_358); +lean_ctor_set(x_367, 2, x_359); +lean_ctor_set(x_367, 3, x_361); +lean_ctor_set_uint8(x_367, sizeof(void*)*4, x_366); +if (lean_is_scalar(x_360)) { + x_368 = lean_alloc_ctor(1, 4, 1); +} else { + x_368 = x_360; +} +lean_ctor_set(x_368, 0, x_364); +lean_ctor_set(x_368, 1, x_343); +lean_ctor_set(x_368, 2, x_344); +lean_ctor_set(x_368, 3, x_345); +lean_ctor_set_uint8(x_368, sizeof(void*)*4, x_366); +x_369 = 0; +x_370 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_370, 0, x_367); +lean_ctor_set(x_370, 1, x_362); +lean_ctor_set(x_370, 2, x_363); +lean_ctor_set(x_370, 3, x_368); +lean_ctor_set_uint8(x_370, sizeof(void*)*4, x_369); +return x_370; +} +else +{ +lean_object* x_371; uint8_t x_372; lean_object* x_373; +if (lean_is_exclusive(x_350)) { + lean_ctor_release(x_350, 0); + lean_ctor_release(x_350, 1); + lean_ctor_release(x_350, 2); + lean_ctor_release(x_350, 3); + x_371 = x_350; +} else { + lean_dec_ref(x_350); + x_371 = lean_box(0); +} +x_372 = 1; +if (lean_is_scalar(x_371)) { + x_373 = lean_alloc_ctor(1, 4, 1); +} else { + x_373 = x_371; +} +lean_ctor_set(x_373, 0, x_347); +lean_ctor_set(x_373, 1, x_343); +lean_ctor_set(x_373, 2, x_344); +lean_ctor_set(x_373, 3, x_345); +lean_ctor_set_uint8(x_373, sizeof(void*)*4, x_372); +return x_373; +} +} +} +else +{ +uint8_t x_374; +x_374 = lean_ctor_get_uint8(x_349, sizeof(void*)*4); +if (x_374 == 0) +{ +lean_object* x_375; lean_object* x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; lean_object* x_382; lean_object* x_383; uint8_t x_384; lean_object* x_385; lean_object* x_386; uint8_t x_387; lean_object* x_388; +x_375 = lean_ctor_get(x_347, 1); +lean_inc(x_375); +x_376 = lean_ctor_get(x_347, 2); +lean_inc(x_376); +x_377 = lean_ctor_get(x_347, 3); +lean_inc(x_377); +if (lean_is_exclusive(x_347)) { + lean_ctor_release(x_347, 0); + lean_ctor_release(x_347, 1); + lean_ctor_release(x_347, 2); + lean_ctor_release(x_347, 3); + x_378 = x_347; +} else { + lean_dec_ref(x_347); + x_378 = lean_box(0); +} +x_379 = lean_ctor_get(x_349, 0); +lean_inc(x_379); +x_380 = lean_ctor_get(x_349, 1); +lean_inc(x_380); +x_381 = lean_ctor_get(x_349, 2); +lean_inc(x_381); +x_382 = lean_ctor_get(x_349, 3); +lean_inc(x_382); +if (lean_is_exclusive(x_349)) { + lean_ctor_release(x_349, 0); + lean_ctor_release(x_349, 1); + lean_ctor_release(x_349, 2); + lean_ctor_release(x_349, 3); + x_383 = x_349; +} else { + lean_dec_ref(x_349); + x_383 = lean_box(0); +} +x_384 = 1; +if (lean_is_scalar(x_383)) { + x_385 = lean_alloc_ctor(1, 4, 1); +} else { + x_385 = x_383; +} +lean_ctor_set(x_385, 0, x_379); +lean_ctor_set(x_385, 1, x_380); +lean_ctor_set(x_385, 2, x_381); +lean_ctor_set(x_385, 3, x_382); +lean_ctor_set_uint8(x_385, sizeof(void*)*4, x_384); +if (lean_is_scalar(x_378)) { + x_386 = lean_alloc_ctor(1, 4, 1); +} else { + x_386 = x_378; +} +lean_ctor_set(x_386, 0, x_377); +lean_ctor_set(x_386, 1, x_343); +lean_ctor_set(x_386, 2, x_344); +lean_ctor_set(x_386, 3, x_345); +lean_ctor_set_uint8(x_386, sizeof(void*)*4, x_384); +x_387 = 0; +x_388 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_388, 0, x_385); +lean_ctor_set(x_388, 1, x_375); +lean_ctor_set(x_388, 2, x_376); +lean_ctor_set(x_388, 3, x_386); +lean_ctor_set_uint8(x_388, sizeof(void*)*4, x_387); +return x_388; +} +else +{ +lean_object* x_389; +x_389 = lean_ctor_get(x_347, 3); +lean_inc(x_389); +if (lean_obj_tag(x_389) == 0) +{ +lean_object* x_390; uint8_t x_391; lean_object* x_392; +if (lean_is_exclusive(x_349)) { + lean_ctor_release(x_349, 0); + lean_ctor_release(x_349, 1); + lean_ctor_release(x_349, 2); + lean_ctor_release(x_349, 3); + x_390 = x_349; +} else { + lean_dec_ref(x_349); + x_390 = lean_box(0); +} +x_391 = 1; +if (lean_is_scalar(x_390)) { + x_392 = lean_alloc_ctor(1, 4, 1); +} else { + x_392 = x_390; +} +lean_ctor_set(x_392, 0, x_347); +lean_ctor_set(x_392, 1, x_343); +lean_ctor_set(x_392, 2, x_344); +lean_ctor_set(x_392, 3, x_345); +lean_ctor_set_uint8(x_392, sizeof(void*)*4, x_391); +return x_392; +} +else +{ +uint8_t x_393; +x_393 = lean_ctor_get_uint8(x_389, sizeof(void*)*4); +if (x_393 == 0) +{ +lean_object* x_394; lean_object* x_395; lean_object* x_396; lean_object* x_397; lean_object* x_398; lean_object* x_399; lean_object* x_400; lean_object* x_401; uint8_t x_402; lean_object* x_403; lean_object* x_404; lean_object* x_405; uint8_t x_406; lean_object* x_407; +x_394 = lean_ctor_get(x_347, 1); +lean_inc(x_394); +x_395 = lean_ctor_get(x_347, 2); +lean_inc(x_395); +if (lean_is_exclusive(x_347)) { + lean_ctor_release(x_347, 0); + lean_ctor_release(x_347, 1); + lean_ctor_release(x_347, 2); + lean_ctor_release(x_347, 3); + x_396 = x_347; +} else { + lean_dec_ref(x_347); + x_396 = lean_box(0); +} +x_397 = lean_ctor_get(x_389, 0); +lean_inc(x_397); +x_398 = lean_ctor_get(x_389, 1); +lean_inc(x_398); +x_399 = lean_ctor_get(x_389, 2); +lean_inc(x_399); +x_400 = lean_ctor_get(x_389, 3); +lean_inc(x_400); +if (lean_is_exclusive(x_389)) { + lean_ctor_release(x_389, 0); + lean_ctor_release(x_389, 1); + lean_ctor_release(x_389, 2); + lean_ctor_release(x_389, 3); + x_401 = x_389; +} else { + lean_dec_ref(x_389); + x_401 = lean_box(0); +} +x_402 = 1; +lean_inc(x_349); +if (lean_is_scalar(x_401)) { + x_403 = lean_alloc_ctor(1, 4, 1); +} else { + x_403 = x_401; +} +lean_ctor_set(x_403, 0, x_349); +lean_ctor_set(x_403, 1, x_394); +lean_ctor_set(x_403, 2, x_395); +lean_ctor_set(x_403, 3, x_397); +if (lean_is_exclusive(x_349)) { + lean_ctor_release(x_349, 0); + lean_ctor_release(x_349, 1); + lean_ctor_release(x_349, 2); + lean_ctor_release(x_349, 3); + x_404 = x_349; +} else { + lean_dec_ref(x_349); + x_404 = lean_box(0); +} +lean_ctor_set_uint8(x_403, sizeof(void*)*4, x_402); +if (lean_is_scalar(x_404)) { + x_405 = lean_alloc_ctor(1, 4, 1); +} else { + x_405 = x_404; +} +lean_ctor_set(x_405, 0, x_400); +lean_ctor_set(x_405, 1, x_343); +lean_ctor_set(x_405, 2, x_344); +lean_ctor_set(x_405, 3, x_345); +lean_ctor_set_uint8(x_405, sizeof(void*)*4, x_402); +x_406 = 0; +if (lean_is_scalar(x_396)) { + x_407 = lean_alloc_ctor(1, 4, 1); +} else { + x_407 = x_396; +} +lean_ctor_set(x_407, 0, x_403); +lean_ctor_set(x_407, 1, x_398); +lean_ctor_set(x_407, 2, x_399); +lean_ctor_set(x_407, 3, x_405); +lean_ctor_set_uint8(x_407, sizeof(void*)*4, x_406); +return x_407; +} +else +{ +lean_object* x_408; lean_object* x_409; lean_object* x_410; lean_object* x_411; lean_object* x_412; lean_object* x_413; lean_object* x_414; lean_object* x_415; lean_object* x_416; lean_object* x_417; uint8_t x_418; lean_object* x_419; +x_408 = lean_ctor_get(x_347, 1); +lean_inc(x_408); +x_409 = lean_ctor_get(x_347, 2); +lean_inc(x_409); +if (lean_is_exclusive(x_347)) { + lean_ctor_release(x_347, 0); + lean_ctor_release(x_347, 1); + lean_ctor_release(x_347, 2); + lean_ctor_release(x_347, 3); + x_410 = x_347; +} else { + lean_dec_ref(x_347); + x_410 = lean_box(0); +} +x_411 = lean_ctor_get(x_349, 0); +lean_inc(x_411); +x_412 = lean_ctor_get(x_349, 1); +lean_inc(x_412); +x_413 = lean_ctor_get(x_349, 2); +lean_inc(x_413); +x_414 = lean_ctor_get(x_349, 3); +lean_inc(x_414); +if (lean_is_exclusive(x_349)) { + lean_ctor_release(x_349, 0); + lean_ctor_release(x_349, 1); + lean_ctor_release(x_349, 2); + lean_ctor_release(x_349, 3); + x_415 = x_349; +} else { + lean_dec_ref(x_349); + x_415 = lean_box(0); +} +if (lean_is_scalar(x_415)) { + x_416 = lean_alloc_ctor(1, 4, 1); +} else { + x_416 = x_415; +} +lean_ctor_set(x_416, 0, x_411); +lean_ctor_set(x_416, 1, x_412); +lean_ctor_set(x_416, 2, x_413); +lean_ctor_set(x_416, 3, x_414); +lean_ctor_set_uint8(x_416, sizeof(void*)*4, x_393); +if (lean_is_scalar(x_410)) { + x_417 = lean_alloc_ctor(1, 4, 1); +} else { + x_417 = x_410; +} +lean_ctor_set(x_417, 0, x_416); +lean_ctor_set(x_417, 1, x_408); +lean_ctor_set(x_417, 2, x_409); +lean_ctor_set(x_417, 3, x_389); +lean_ctor_set_uint8(x_417, sizeof(void*)*4, x_348); +x_418 = 1; +x_419 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_419, 0, x_417); +lean_ctor_set(x_419, 1, x_343); +lean_ctor_set(x_419, 2, x_344); +lean_ctor_set(x_419, 3, x_345); +lean_ctor_set_uint8(x_419, sizeof(void*)*4, x_418); +return x_419; +} +} +} +} +} +else +{ +uint8_t x_420; lean_object* x_421; +x_420 = 1; +x_421 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_421, 0, x_347); +lean_ctor_set(x_421, 1, x_343); +lean_ctor_set(x_421, 2, x_344); +lean_ctor_set(x_421, 3, x_345); +lean_ctor_set_uint8(x_421, sizeof(void*)*4, x_420); +return x_421; +} +} +case 1: +{ +uint8_t x_422; lean_object* x_423; +lean_dec(x_344); +lean_dec(x_343); +x_422 = 1; +x_423 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_423, 0, x_342); +lean_ctor_set(x_423, 1, x_2); +lean_ctor_set(x_423, 2, x_3); +lean_ctor_set(x_423, 3, x_345); +lean_ctor_set_uint8(x_423, sizeof(void*)*4, x_422); +return x_423; +} +default: +{ +lean_object* x_424; uint8_t x_425; +x_424 = l_Lean_RBNode_ins___at_Lean_Meta_Grind_registerParent___spec__2(x_345, x_2, x_3); +x_425 = lean_ctor_get_uint8(x_424, sizeof(void*)*4); +if (x_425 == 0) +{ +lean_object* x_426; +x_426 = lean_ctor_get(x_424, 0); +lean_inc(x_426); +if (lean_obj_tag(x_426) == 0) +{ +lean_object* x_427; +x_427 = lean_ctor_get(x_424, 3); +lean_inc(x_427); +if (lean_obj_tag(x_427) == 0) +{ +lean_object* x_428; lean_object* x_429; lean_object* x_430; lean_object* x_431; uint8_t x_432; lean_object* x_433; +x_428 = lean_ctor_get(x_424, 1); +lean_inc(x_428); +x_429 = lean_ctor_get(x_424, 2); +lean_inc(x_429); +if (lean_is_exclusive(x_424)) { + lean_ctor_release(x_424, 0); + lean_ctor_release(x_424, 1); + lean_ctor_release(x_424, 2); + lean_ctor_release(x_424, 3); + x_430 = x_424; +} else { + lean_dec_ref(x_424); + x_430 = lean_box(0); +} +if (lean_is_scalar(x_430)) { + x_431 = lean_alloc_ctor(1, 4, 1); +} else { + x_431 = x_430; +} +lean_ctor_set(x_431, 0, x_427); +lean_ctor_set(x_431, 1, x_428); +lean_ctor_set(x_431, 2, x_429); +lean_ctor_set(x_431, 3, x_427); +lean_ctor_set_uint8(x_431, sizeof(void*)*4, x_425); +x_432 = 1; +x_433 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_433, 0, x_342); +lean_ctor_set(x_433, 1, x_343); +lean_ctor_set(x_433, 2, x_344); +lean_ctor_set(x_433, 3, x_431); +lean_ctor_set_uint8(x_433, sizeof(void*)*4, x_432); +return x_433; +} +else +{ +uint8_t x_434; +x_434 = lean_ctor_get_uint8(x_427, sizeof(void*)*4); +if (x_434 == 0) +{ +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; uint8_t x_443; lean_object* x_444; lean_object* x_445; uint8_t x_446; lean_object* x_447; +x_435 = lean_ctor_get(x_424, 1); +lean_inc(x_435); +x_436 = lean_ctor_get(x_424, 2); +lean_inc(x_436); +if (lean_is_exclusive(x_424)) { + lean_ctor_release(x_424, 0); + lean_ctor_release(x_424, 1); + lean_ctor_release(x_424, 2); + lean_ctor_release(x_424, 3); + x_437 = x_424; +} else { + lean_dec_ref(x_424); + x_437 = lean_box(0); +} +x_438 = lean_ctor_get(x_427, 0); +lean_inc(x_438); +x_439 = lean_ctor_get(x_427, 1); +lean_inc(x_439); +x_440 = lean_ctor_get(x_427, 2); +lean_inc(x_440); +x_441 = lean_ctor_get(x_427, 3); +lean_inc(x_441); +if (lean_is_exclusive(x_427)) { + lean_ctor_release(x_427, 0); + lean_ctor_release(x_427, 1); + lean_ctor_release(x_427, 2); + lean_ctor_release(x_427, 3); + x_442 = x_427; +} else { + lean_dec_ref(x_427); + x_442 = lean_box(0); +} +x_443 = 1; +if (lean_is_scalar(x_442)) { + x_444 = lean_alloc_ctor(1, 4, 1); +} else { + x_444 = x_442; +} +lean_ctor_set(x_444, 0, x_342); +lean_ctor_set(x_444, 1, x_343); +lean_ctor_set(x_444, 2, x_344); +lean_ctor_set(x_444, 3, x_426); +lean_ctor_set_uint8(x_444, sizeof(void*)*4, x_443); +if (lean_is_scalar(x_437)) { + x_445 = lean_alloc_ctor(1, 4, 1); +} else { + x_445 = x_437; +} +lean_ctor_set(x_445, 0, x_438); +lean_ctor_set(x_445, 1, x_439); +lean_ctor_set(x_445, 2, x_440); +lean_ctor_set(x_445, 3, x_441); +lean_ctor_set_uint8(x_445, sizeof(void*)*4, x_443); +x_446 = 0; +x_447 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_447, 0, x_444); +lean_ctor_set(x_447, 1, x_435); +lean_ctor_set(x_447, 2, x_436); +lean_ctor_set(x_447, 3, x_445); +lean_ctor_set_uint8(x_447, sizeof(void*)*4, x_446); +return x_447; +} +else +{ +lean_object* x_448; uint8_t x_449; lean_object* x_450; +if (lean_is_exclusive(x_427)) { + lean_ctor_release(x_427, 0); + lean_ctor_release(x_427, 1); + lean_ctor_release(x_427, 2); + lean_ctor_release(x_427, 3); + x_448 = x_427; +} else { + lean_dec_ref(x_427); + x_448 = lean_box(0); +} +x_449 = 1; +if (lean_is_scalar(x_448)) { + x_450 = lean_alloc_ctor(1, 4, 1); +} else { + x_450 = x_448; +} +lean_ctor_set(x_450, 0, x_342); +lean_ctor_set(x_450, 1, x_343); +lean_ctor_set(x_450, 2, x_344); +lean_ctor_set(x_450, 3, x_424); +lean_ctor_set_uint8(x_450, sizeof(void*)*4, x_449); +return x_450; +} +} +} +else { -lean_object* x_41; -x_41 = lean_ctor_get(x_38, 3); -lean_inc(x_41); -if (lean_obj_tag(x_41) == 0) +uint8_t x_451; +x_451 = lean_ctor_get_uint8(x_426, sizeof(void*)*4); +if (x_451 == 0) { -uint8_t x_42; -x_42 = !lean_is_exclusive(x_38); -if (x_42 == 0) +lean_object* x_452; lean_object* x_453; lean_object* x_454; lean_object* x_455; lean_object* x_456; lean_object* x_457; lean_object* x_458; lean_object* x_459; lean_object* x_460; uint8_t x_461; lean_object* x_462; lean_object* x_463; uint8_t x_464; lean_object* x_465; +x_452 = lean_ctor_get(x_424, 1); +lean_inc(x_452); +x_453 = lean_ctor_get(x_424, 2); +lean_inc(x_453); +x_454 = lean_ctor_get(x_424, 3); +lean_inc(x_454); +if (lean_is_exclusive(x_424)) { + lean_ctor_release(x_424, 0); + lean_ctor_release(x_424, 1); + lean_ctor_release(x_424, 2); + lean_ctor_release(x_424, 3); + x_455 = x_424; +} else { + lean_dec_ref(x_424); + x_455 = lean_box(0); +} +x_456 = lean_ctor_get(x_426, 0); +lean_inc(x_456); +x_457 = lean_ctor_get(x_426, 1); +lean_inc(x_457); +x_458 = lean_ctor_get(x_426, 2); +lean_inc(x_458); +x_459 = lean_ctor_get(x_426, 3); +lean_inc(x_459); +if (lean_is_exclusive(x_426)) { + lean_ctor_release(x_426, 0); + lean_ctor_release(x_426, 1); + lean_ctor_release(x_426, 2); + lean_ctor_release(x_426, 3); + x_460 = x_426; +} else { + lean_dec_ref(x_426); + x_460 = lean_box(0); +} +x_461 = 1; +if (lean_is_scalar(x_460)) { + x_462 = lean_alloc_ctor(1, 4, 1); +} else { + x_462 = x_460; +} +lean_ctor_set(x_462, 0, x_342); +lean_ctor_set(x_462, 1, x_343); +lean_ctor_set(x_462, 2, x_344); +lean_ctor_set(x_462, 3, x_456); +lean_ctor_set_uint8(x_462, sizeof(void*)*4, x_461); +if (lean_is_scalar(x_455)) { + x_463 = lean_alloc_ctor(1, 4, 1); +} else { + x_463 = x_455; +} +lean_ctor_set(x_463, 0, x_459); +lean_ctor_set(x_463, 1, x_452); +lean_ctor_set(x_463, 2, x_453); +lean_ctor_set(x_463, 3, x_454); +lean_ctor_set_uint8(x_463, sizeof(void*)*4, x_461); +x_464 = 0; +x_465 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_465, 0, x_462); +lean_ctor_set(x_465, 1, x_457); +lean_ctor_set(x_465, 2, x_458); +lean_ctor_set(x_465, 3, x_463); +lean_ctor_set_uint8(x_465, sizeof(void*)*4, x_464); +return x_465; +} +else { -lean_object* x_43; lean_object* x_44; uint8_t x_45; -x_43 = lean_ctor_get(x_38, 3); -lean_dec(x_43); -x_44 = lean_ctor_get(x_38, 0); -lean_dec(x_44); -lean_ctor_set(x_38, 0, x_41); -x_45 = 1; -lean_ctor_set(x_1, 0, x_38); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_45); -return x_1; +lean_object* x_466; +x_466 = lean_ctor_get(x_424, 3); +lean_inc(x_466); +if (lean_obj_tag(x_466) == 0) +{ +lean_object* x_467; uint8_t x_468; lean_object* x_469; +if (lean_is_exclusive(x_426)) { + lean_ctor_release(x_426, 0); + lean_ctor_release(x_426, 1); + lean_ctor_release(x_426, 2); + lean_ctor_release(x_426, 3); + x_467 = x_426; +} else { + lean_dec_ref(x_426); + x_467 = lean_box(0); +} +x_468 = 1; +if (lean_is_scalar(x_467)) { + x_469 = lean_alloc_ctor(1, 4, 1); +} else { + x_469 = x_467; +} +lean_ctor_set(x_469, 0, x_342); +lean_ctor_set(x_469, 1, x_343); +lean_ctor_set(x_469, 2, x_344); +lean_ctor_set(x_469, 3, x_424); +lean_ctor_set_uint8(x_469, sizeof(void*)*4, x_468); +return x_469; +} +else +{ +uint8_t x_470; +x_470 = lean_ctor_get_uint8(x_466, sizeof(void*)*4); +if (x_470 == 0) +{ +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; uint8_t x_479; lean_object* x_480; lean_object* x_481; lean_object* x_482; uint8_t x_483; lean_object* x_484; +x_471 = lean_ctor_get(x_424, 1); +lean_inc(x_471); +x_472 = lean_ctor_get(x_424, 2); +lean_inc(x_472); +if (lean_is_exclusive(x_424)) { + lean_ctor_release(x_424, 0); + lean_ctor_release(x_424, 1); + lean_ctor_release(x_424, 2); + lean_ctor_release(x_424, 3); + x_473 = x_424; +} else { + lean_dec_ref(x_424); + x_473 = lean_box(0); +} +x_474 = lean_ctor_get(x_466, 0); +lean_inc(x_474); +x_475 = lean_ctor_get(x_466, 1); +lean_inc(x_475); +x_476 = lean_ctor_get(x_466, 2); +lean_inc(x_476); +x_477 = lean_ctor_get(x_466, 3); +lean_inc(x_477); +if (lean_is_exclusive(x_466)) { + lean_ctor_release(x_466, 0); + lean_ctor_release(x_466, 1); + lean_ctor_release(x_466, 2); + lean_ctor_release(x_466, 3); + x_478 = x_466; +} else { + lean_dec_ref(x_466); + x_478 = lean_box(0); +} +x_479 = 1; +lean_inc(x_426); +if (lean_is_scalar(x_478)) { + x_480 = lean_alloc_ctor(1, 4, 1); +} else { + x_480 = x_478; +} +lean_ctor_set(x_480, 0, x_342); +lean_ctor_set(x_480, 1, x_343); +lean_ctor_set(x_480, 2, x_344); +lean_ctor_set(x_480, 3, x_426); +if (lean_is_exclusive(x_426)) { + lean_ctor_release(x_426, 0); + lean_ctor_release(x_426, 1); + lean_ctor_release(x_426, 2); + lean_ctor_release(x_426, 3); + x_481 = x_426; +} else { + lean_dec_ref(x_426); + x_481 = lean_box(0); +} +lean_ctor_set_uint8(x_480, sizeof(void*)*4, x_479); +if (lean_is_scalar(x_481)) { + x_482 = lean_alloc_ctor(1, 4, 1); +} else { + x_482 = x_481; +} +lean_ctor_set(x_482, 0, x_474); +lean_ctor_set(x_482, 1, x_475); +lean_ctor_set(x_482, 2, x_476); +lean_ctor_set(x_482, 3, x_477); +lean_ctor_set_uint8(x_482, sizeof(void*)*4, x_479); +x_483 = 0; +if (lean_is_scalar(x_473)) { + x_484 = lean_alloc_ctor(1, 4, 1); +} else { + x_484 = x_473; +} +lean_ctor_set(x_484, 0, x_480); +lean_ctor_set(x_484, 1, x_471); +lean_ctor_set(x_484, 2, x_472); +lean_ctor_set(x_484, 3, x_482); +lean_ctor_set_uint8(x_484, sizeof(void*)*4, x_483); +return x_484; +} +else +{ +lean_object* x_485; lean_object* x_486; lean_object* x_487; lean_object* x_488; lean_object* x_489; lean_object* x_490; lean_object* x_491; lean_object* x_492; lean_object* x_493; lean_object* x_494; uint8_t x_495; lean_object* x_496; +x_485 = lean_ctor_get(x_424, 1); +lean_inc(x_485); +x_486 = lean_ctor_get(x_424, 2); +lean_inc(x_486); +if (lean_is_exclusive(x_424)) { + lean_ctor_release(x_424, 0); + lean_ctor_release(x_424, 1); + lean_ctor_release(x_424, 2); + lean_ctor_release(x_424, 3); + x_487 = x_424; +} else { + lean_dec_ref(x_424); + x_487 = lean_box(0); +} +x_488 = lean_ctor_get(x_426, 0); +lean_inc(x_488); +x_489 = lean_ctor_get(x_426, 1); +lean_inc(x_489); +x_490 = lean_ctor_get(x_426, 2); +lean_inc(x_490); +x_491 = lean_ctor_get(x_426, 3); +lean_inc(x_491); +if (lean_is_exclusive(x_426)) { + lean_ctor_release(x_426, 0); + lean_ctor_release(x_426, 1); + lean_ctor_release(x_426, 2); + lean_ctor_release(x_426, 3); + x_492 = x_426; +} else { + lean_dec_ref(x_426); + x_492 = lean_box(0); +} +if (lean_is_scalar(x_492)) { + x_493 = lean_alloc_ctor(1, 4, 1); +} else { + x_493 = x_492; +} +lean_ctor_set(x_493, 0, x_488); +lean_ctor_set(x_493, 1, x_489); +lean_ctor_set(x_493, 2, x_490); +lean_ctor_set(x_493, 3, x_491); +lean_ctor_set_uint8(x_493, sizeof(void*)*4, x_470); +if (lean_is_scalar(x_487)) { + x_494 = lean_alloc_ctor(1, 4, 1); +} else { + x_494 = x_487; +} +lean_ctor_set(x_494, 0, x_493); +lean_ctor_set(x_494, 1, x_485); +lean_ctor_set(x_494, 2, x_486); +lean_ctor_set(x_494, 3, x_466); +lean_ctor_set_uint8(x_494, sizeof(void*)*4, x_425); +x_495 = 1; +x_496 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_496, 0, x_342); +lean_ctor_set(x_496, 1, x_343); +lean_ctor_set(x_496, 2, x_344); +lean_ctor_set(x_496, 3, x_494); +lean_ctor_set_uint8(x_496, sizeof(void*)*4, x_495); +return x_496; +} +} } -else -{ -lean_object* x_46; lean_object* x_47; lean_object* x_48; uint8_t x_49; -x_46 = lean_ctor_get(x_38, 1); -x_47 = lean_ctor_get(x_38, 2); -lean_inc(x_47); -lean_inc(x_46); -lean_dec(x_38); -x_48 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_48, 0, x_41); -lean_ctor_set(x_48, 1, x_46); -lean_ctor_set(x_48, 2, x_47); -lean_ctor_set(x_48, 3, x_41); -lean_ctor_set_uint8(x_48, sizeof(void*)*4, x_39); -x_49 = 1; -lean_ctor_set(x_1, 0, x_48); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_49); -return x_1; } } else { -uint8_t x_50; -x_50 = lean_ctor_get_uint8(x_41, sizeof(void*)*4); -if (x_50 == 0) -{ -uint8_t x_51; -x_51 = !lean_is_exclusive(x_38); -if (x_51 == 0) +uint8_t x_497; lean_object* x_498; +x_497 = 1; +x_498 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_498, 0, x_342); +lean_ctor_set(x_498, 1, x_343); +lean_ctor_set(x_498, 2, x_344); +lean_ctor_set(x_498, 3, x_424); +lean_ctor_set_uint8(x_498, sizeof(void*)*4, x_497); +return x_498; +} +} +} +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_RBNode_insert___at_Lean_Meta_Grind_registerParent___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -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_38, 1); -x_53 = lean_ctor_get(x_38, 2); -x_54 = lean_ctor_get(x_38, 3); -lean_dec(x_54); -x_55 = lean_ctor_get(x_38, 0); -lean_dec(x_55); -x_56 = !lean_is_exclusive(x_41); -if (x_56 == 0) +uint8_t x_4; +x_4 = l_Lean_RBNode_isRed___rarg(x_1); +if (x_4 == 0) { -lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; uint8_t x_61; uint8_t x_62; -x_57 = lean_ctor_get(x_41, 0); -x_58 = lean_ctor_get(x_41, 1); -x_59 = lean_ctor_get(x_41, 2); -x_60 = lean_ctor_get(x_41, 3); -x_61 = 1; -lean_ctor_set(x_41, 3, x_57); -lean_ctor_set(x_41, 2, x_53); -lean_ctor_set(x_41, 1, x_52); -lean_ctor_set(x_41, 0, x_40); -lean_ctor_set_uint8(x_41, sizeof(void*)*4, x_61); -lean_ctor_set(x_38, 3, x_36); -lean_ctor_set(x_38, 2, x_35); -lean_ctor_set(x_38, 1, x_34); -lean_ctor_set(x_38, 0, x_60); -lean_ctor_set_uint8(x_38, sizeof(void*)*4, x_61); -x_62 = 0; -lean_ctor_set(x_1, 3, x_38); -lean_ctor_set(x_1, 2, x_59); -lean_ctor_set(x_1, 1, x_58); -lean_ctor_set(x_1, 0, x_41); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_62); -return x_1; +lean_object* x_5; +x_5 = l_Lean_RBNode_ins___at_Lean_Meta_Grind_registerParent___spec__2(x_1, x_2, x_3); +return x_5; } else { -lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; uint8_t x_67; lean_object* x_68; uint8_t x_69; -x_63 = lean_ctor_get(x_41, 0); -x_64 = lean_ctor_get(x_41, 1); -x_65 = lean_ctor_get(x_41, 2); -x_66 = lean_ctor_get(x_41, 3); -lean_inc(x_66); -lean_inc(x_65); -lean_inc(x_64); -lean_inc(x_63); -lean_dec(x_41); -x_67 = 1; -x_68 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_68, 0, x_40); -lean_ctor_set(x_68, 1, x_52); -lean_ctor_set(x_68, 2, x_53); -lean_ctor_set(x_68, 3, x_63); -lean_ctor_set_uint8(x_68, sizeof(void*)*4, x_67); -lean_ctor_set(x_38, 3, x_36); -lean_ctor_set(x_38, 2, x_35); -lean_ctor_set(x_38, 1, x_34); -lean_ctor_set(x_38, 0, x_66); -lean_ctor_set_uint8(x_38, sizeof(void*)*4, x_67); -x_69 = 0; -lean_ctor_set(x_1, 3, x_38); -lean_ctor_set(x_1, 2, x_65); -lean_ctor_set(x_1, 1, x_64); -lean_ctor_set(x_1, 0, x_68); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_69); -return x_1; +lean_object* x_6; lean_object* x_7; +x_6 = l_Lean_RBNode_ins___at_Lean_Meta_Grind_registerParent___spec__2(x_1, x_2, x_3); +x_7 = l_Lean_RBNode_setBlack___rarg(x_6); +return x_7; } } -else +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_registerParent___spec__5(size_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: { -lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; uint8_t x_77; lean_object* x_78; lean_object* x_79; uint8_t x_80; -x_70 = lean_ctor_get(x_38, 1); -x_71 = lean_ctor_get(x_38, 2); -lean_inc(x_71); -lean_inc(x_70); -lean_dec(x_38); -x_72 = lean_ctor_get(x_41, 0); -lean_inc(x_72); -x_73 = lean_ctor_get(x_41, 1); -lean_inc(x_73); -x_74 = lean_ctor_get(x_41, 2); -lean_inc(x_74); -x_75 = lean_ctor_get(x_41, 3); -lean_inc(x_75); -if (lean_is_exclusive(x_41)) { - lean_ctor_release(x_41, 0); - lean_ctor_release(x_41, 1); - lean_ctor_release(x_41, 2); - lean_ctor_release(x_41, 3); - x_76 = x_41; -} else { - lean_dec_ref(x_41); - x_76 = lean_box(0); +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; } -x_77 = 1; -if (lean_is_scalar(x_76)) { - x_78 = lean_alloc_ctor(1, 4, 1); -} else { - x_78 = x_76; +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_Meta_Grind_instHashableENodeKey_unsafe__1(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_Grind_registerParent___spec__4(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_ctor_set(x_78, 0, x_40); -lean_ctor_set(x_78, 1, x_70); -lean_ctor_set(x_78, 2, x_71); -lean_ctor_set(x_78, 3, x_72); -lean_ctor_set_uint8(x_78, sizeof(void*)*4, x_77); -x_79 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_79, 0, x_75); -lean_ctor_set(x_79, 1, x_34); -lean_ctor_set(x_79, 2, x_35); -lean_ctor_set(x_79, 3, x_36); -lean_ctor_set_uint8(x_79, sizeof(void*)*4, x_77); -x_80 = 0; -lean_ctor_set(x_1, 3, x_79); -lean_ctor_set(x_1, 2, x_74); -lean_ctor_set(x_1, 1, x_73); -lean_ctor_set(x_1, 0, x_78); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_80); -return x_1; } } -else +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_registerParent___spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: { -uint8_t x_81; -lean_free_object(x_1); -x_81 = !lean_is_exclusive(x_41); -if (x_81 == 0) +lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; +x_5 = lean_ctor_get(x_1, 0); +lean_inc(x_5); +x_6 = lean_ctor_get(x_1, 1); +lean_inc(x_6); +x_7 = lean_array_get_size(x_5); +x_8 = lean_nat_dec_lt(x_2, x_7); +lean_dec(x_7); +if (x_8 == 0) { -lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; uint8_t x_86; -x_82 = lean_ctor_get(x_41, 3); -lean_dec(x_82); -x_83 = lean_ctor_get(x_41, 2); -lean_dec(x_83); -x_84 = lean_ctor_get(x_41, 1); -lean_dec(x_84); -x_85 = lean_ctor_get(x_41, 0); -lean_dec(x_85); -x_86 = 1; -lean_ctor_set(x_41, 3, x_36); -lean_ctor_set(x_41, 2, x_35); -lean_ctor_set(x_41, 1, x_34); -lean_ctor_set(x_41, 0, x_38); -lean_ctor_set_uint8(x_41, sizeof(void*)*4, x_86); -return x_41; +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 { -uint8_t x_87; lean_object* x_88; -lean_dec(x_41); -x_87 = 1; -x_88 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_88, 0, x_38); -lean_ctor_set(x_88, 1, x_34); -lean_ctor_set(x_88, 2, x_35); -lean_ctor_set(x_88, 3, x_36); -lean_ctor_set_uint8(x_88, sizeof(void*)*4, x_87); -return x_88; -} -} +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 { -uint8_t x_89; -x_89 = lean_ctor_get_uint8(x_40, sizeof(void*)*4); -if (x_89 == 0) -{ -uint8_t x_90; -x_90 = !lean_is_exclusive(x_38); -if (x_90 == 0) -{ -lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; uint8_t x_95; -x_91 = lean_ctor_get(x_38, 1); -x_92 = lean_ctor_get(x_38, 2); -x_93 = lean_ctor_get(x_38, 3); -x_94 = lean_ctor_get(x_38, 0); -lean_dec(x_94); -x_95 = !lean_is_exclusive(x_40); -if (x_95 == 0) +lean_object* x_17; uint8_t x_18; +x_17 = lean_array_fget(x_5, x_2); +x_18 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_3, x_17); +lean_dec(x_17); +if (x_18 == 0) { -uint8_t x_96; uint8_t x_97; -x_96 = 1; -lean_ctor_set_uint8(x_40, sizeof(void*)*4, x_96); -lean_ctor_set(x_38, 3, x_36); -lean_ctor_set(x_38, 2, x_35); -lean_ctor_set(x_38, 1, x_34); -lean_ctor_set(x_38, 0, x_93); -lean_ctor_set_uint8(x_38, sizeof(void*)*4, x_96); -x_97 = 0; -lean_ctor_set(x_1, 3, x_38); -lean_ctor_set(x_1, 2, x_92); -lean_ctor_set(x_1, 1, x_91); -lean_ctor_set(x_1, 0, x_40); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_97); -return x_1; +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 { -lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; uint8_t x_102; lean_object* x_103; uint8_t x_104; -x_98 = lean_ctor_get(x_40, 0); -x_99 = lean_ctor_get(x_40, 1); -x_100 = lean_ctor_get(x_40, 2); -x_101 = lean_ctor_get(x_40, 3); -lean_inc(x_101); -lean_inc(x_100); -lean_inc(x_99); -lean_inc(x_98); -lean_dec(x_40); -x_102 = 1; -x_103 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_103, 0, x_98); -lean_ctor_set(x_103, 1, x_99); -lean_ctor_set(x_103, 2, x_100); -lean_ctor_set(x_103, 3, x_101); -lean_ctor_set_uint8(x_103, sizeof(void*)*4, x_102); -lean_ctor_set(x_38, 3, x_36); -lean_ctor_set(x_38, 2, x_35); -lean_ctor_set(x_38, 1, x_34); -lean_ctor_set(x_38, 0, x_93); -lean_ctor_set_uint8(x_38, sizeof(void*)*4, x_102); -x_104 = 0; -lean_ctor_set(x_1, 3, x_38); -lean_ctor_set(x_1, 2, x_92); -lean_ctor_set(x_1, 1, x_91); -lean_ctor_set(x_1, 0, x_103); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_104); +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_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; uint8_t x_113; lean_object* x_114; lean_object* x_115; uint8_t x_116; -x_105 = lean_ctor_get(x_38, 1); -x_106 = lean_ctor_get(x_38, 2); -x_107 = lean_ctor_get(x_38, 3); -lean_inc(x_107); -lean_inc(x_106); -lean_inc(x_105); -lean_dec(x_38); -x_108 = lean_ctor_get(x_40, 0); -lean_inc(x_108); -x_109 = lean_ctor_get(x_40, 1); -lean_inc(x_109); -x_110 = lean_ctor_get(x_40, 2); -lean_inc(x_110); -x_111 = lean_ctor_get(x_40, 3); -lean_inc(x_111); -if (lean_is_exclusive(x_40)) { - lean_ctor_release(x_40, 0); - lean_ctor_release(x_40, 1); - lean_ctor_release(x_40, 2); - lean_ctor_release(x_40, 3); - x_112 = x_40; -} else { - lean_dec_ref(x_40); - x_112 = lean_box(0); +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; } -x_113 = 1; -if (lean_is_scalar(x_112)) { - x_114 = lean_alloc_ctor(1, 4, 1); -} else { - x_114 = x_112; } -lean_ctor_set(x_114, 0, x_108); -lean_ctor_set(x_114, 1, x_109); -lean_ctor_set(x_114, 2, x_110); -lean_ctor_set(x_114, 3, x_111); -lean_ctor_set_uint8(x_114, sizeof(void*)*4, x_113); -x_115 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_115, 0, x_107); -lean_ctor_set(x_115, 1, x_34); -lean_ctor_set(x_115, 2, x_35); -lean_ctor_set(x_115, 3, x_36); -lean_ctor_set_uint8(x_115, sizeof(void*)*4, x_113); -x_116 = 0; -lean_ctor_set(x_1, 3, x_115); -lean_ctor_set(x_1, 2, x_106); -lean_ctor_set(x_1, 1, x_105); -lean_ctor_set(x_1, 0, x_114); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_116); -return x_1; } } -else +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_registerParent___spec__4(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5) { +_start: { -lean_object* x_117; -x_117 = lean_ctor_get(x_38, 3); -lean_inc(x_117); -if (lean_obj_tag(x_117) == 0) +if (lean_obj_tag(x_1) == 0) { -uint8_t x_118; -lean_free_object(x_1); -x_118 = !lean_is_exclusive(x_40); -if (x_118 == 0) +uint8_t x_6; +x_6 = !lean_is_exclusive(x_1); +if (x_6 == 0) { -lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; uint8_t x_123; -x_119 = lean_ctor_get(x_40, 3); -lean_dec(x_119); -x_120 = lean_ctor_get(x_40, 2); -lean_dec(x_120); -x_121 = lean_ctor_get(x_40, 1); -lean_dec(x_121); -x_122 = lean_ctor_get(x_40, 0); -lean_dec(x_122); -x_123 = 1; -lean_ctor_set(x_40, 3, x_36); -lean_ctor_set(x_40, 2, x_35); -lean_ctor_set(x_40, 1, x_34); -lean_ctor_set(x_40, 0, x_38); -lean_ctor_set_uint8(x_40, sizeof(void*)*4, x_123); -return x_40; -} -else +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_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___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) { -uint8_t x_124; lean_object* x_125; -lean_dec(x_40); -x_124 = 1; -x_125 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_125, 0, x_38); -lean_ctor_set(x_125, 1, x_34); -lean_ctor_set(x_125, 2, x_35); -lean_ctor_set(x_125, 3, x_36); -lean_ctor_set_uint8(x_125, sizeof(void*)*4, x_124); -return x_125; -} +lean_dec(x_12); +lean_dec(x_5); +lean_dec(x_4); +return x_1; } else { -uint8_t x_126; -x_126 = lean_ctor_get_uint8(x_117, sizeof(void*)*4); -if (x_126 == 0) -{ -uint8_t x_127; -lean_free_object(x_1); -x_127 = !lean_is_exclusive(x_38); -if (x_127 == 0) -{ -lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; uint8_t x_132; -x_128 = lean_ctor_get(x_38, 1); -x_129 = lean_ctor_get(x_38, 2); -x_130 = lean_ctor_get(x_38, 3); -lean_dec(x_130); -x_131 = lean_ctor_get(x_38, 0); -lean_dec(x_131); -x_132 = !lean_is_exclusive(x_117); -if (x_132 == 0) +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_array_fget(x_7, x_12); +x_16 = lean_box(0); +x_17 = lean_array_fset(x_7, x_12, x_16); +switch (lean_obj_tag(x_15)) { +case 0: { -lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; uint8_t x_137; uint8_t x_138; -x_133 = lean_ctor_get(x_117, 0); -x_134 = lean_ctor_get(x_117, 1); -x_135 = lean_ctor_get(x_117, 2); -x_136 = lean_ctor_get(x_117, 3); -x_137 = 1; -lean_inc(x_40); -lean_ctor_set(x_117, 3, x_133); -lean_ctor_set(x_117, 2, x_129); -lean_ctor_set(x_117, 1, x_128); -lean_ctor_set(x_117, 0, x_40); -x_138 = !lean_is_exclusive(x_40); -if (x_138 == 0) +uint8_t x_18; +x_18 = !lean_is_exclusive(x_15); +if (x_18 == 0) { -lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; uint8_t x_143; -x_139 = lean_ctor_get(x_40, 3); -lean_dec(x_139); -x_140 = lean_ctor_get(x_40, 2); -lean_dec(x_140); -x_141 = lean_ctor_get(x_40, 1); -lean_dec(x_141); -x_142 = lean_ctor_get(x_40, 0); -lean_dec(x_142); -lean_ctor_set_uint8(x_117, sizeof(void*)*4, x_137); -lean_ctor_set(x_40, 3, x_36); -lean_ctor_set(x_40, 2, x_35); -lean_ctor_set(x_40, 1, x_34); -lean_ctor_set(x_40, 0, x_136); -lean_ctor_set_uint8(x_40, sizeof(void*)*4, x_137); -x_143 = 0; -lean_ctor_set(x_38, 3, x_40); -lean_ctor_set(x_38, 2, x_135); -lean_ctor_set(x_38, 1, x_134); -lean_ctor_set(x_38, 0, x_117); -lean_ctor_set_uint8(x_38, sizeof(void*)*4, x_143); -return x_38; +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 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(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_144; uint8_t x_145; -lean_dec(x_40); -lean_ctor_set_uint8(x_117, sizeof(void*)*4, x_137); -x_144 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_144, 0, x_136); -lean_ctor_set(x_144, 1, x_34); -lean_ctor_set(x_144, 2, x_35); -lean_ctor_set(x_144, 3, x_36); -lean_ctor_set_uint8(x_144, sizeof(void*)*4, x_137); -x_145 = 0; -lean_ctor_set(x_38, 3, x_144); -lean_ctor_set(x_38, 2, x_135); -lean_ctor_set(x_38, 1, x_134); -lean_ctor_set(x_38, 0, x_117); -lean_ctor_set_uint8(x_38, sizeof(void*)*4, x_145); -return x_38; +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_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; uint8_t x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; uint8_t x_154; -x_146 = lean_ctor_get(x_117, 0); -x_147 = lean_ctor_get(x_117, 1); -x_148 = lean_ctor_get(x_117, 2); -x_149 = lean_ctor_get(x_117, 3); -lean_inc(x_149); -lean_inc(x_148); -lean_inc(x_147); -lean_inc(x_146); -lean_dec(x_117); -x_150 = 1; -lean_inc(x_40); -x_151 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_151, 0, x_40); -lean_ctor_set(x_151, 1, x_128); -lean_ctor_set(x_151, 2, x_129); -lean_ctor_set(x_151, 3, x_146); -if (lean_is_exclusive(x_40)) { - lean_ctor_release(x_40, 0); - lean_ctor_release(x_40, 1); - lean_ctor_release(x_40, 2); - lean_ctor_release(x_40, 3); - x_152 = x_40; -} else { - lean_dec_ref(x_40); - x_152 = lean_box(0); +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 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(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; } -lean_ctor_set_uint8(x_151, sizeof(void*)*4, x_150); -if (lean_is_scalar(x_152)) { - x_153 = lean_alloc_ctor(1, 4, 1); -} else { - x_153 = x_152; +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; } -lean_ctor_set(x_153, 0, x_149); -lean_ctor_set(x_153, 1, x_34); -lean_ctor_set(x_153, 2, x_35); -lean_ctor_set(x_153, 3, x_36); -lean_ctor_set_uint8(x_153, sizeof(void*)*4, x_150); -x_154 = 0; -lean_ctor_set(x_38, 3, x_153); -lean_ctor_set(x_38, 2, x_148); -lean_ctor_set(x_38, 1, x_147); -lean_ctor_set(x_38, 0, x_151); -lean_ctor_set_uint8(x_38, sizeof(void*)*4, x_154); -return x_38; } } -else +case 1: { -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; uint8_t x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; uint8_t x_166; lean_object* x_167; -x_155 = lean_ctor_get(x_38, 1); -x_156 = lean_ctor_get(x_38, 2); -lean_inc(x_156); -lean_inc(x_155); -lean_dec(x_38); -x_157 = lean_ctor_get(x_117, 0); -lean_inc(x_157); -x_158 = lean_ctor_get(x_117, 1); -lean_inc(x_158); -x_159 = lean_ctor_get(x_117, 2); -lean_inc(x_159); -x_160 = lean_ctor_get(x_117, 3); -lean_inc(x_160); -if (lean_is_exclusive(x_117)) { - lean_ctor_release(x_117, 0); - lean_ctor_release(x_117, 1); - lean_ctor_release(x_117, 2); - lean_ctor_release(x_117, 3); - x_161 = x_117; -} else { - lean_dec_ref(x_117); - x_161 = lean_box(0); +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_Grind_registerParent___spec__4(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; } -x_162 = 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); -if (lean_is_scalar(x_161)) { - x_163 = lean_alloc_ctor(1, 4, 1); -} else { - x_163 = x_161; +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_Grind_registerParent___spec__4(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; } -lean_ctor_set(x_163, 0, x_40); -lean_ctor_set(x_163, 1, x_155); -lean_ctor_set(x_163, 2, x_156); -lean_ctor_set(x_163, 3, x_157); -if (lean_is_exclusive(x_40)) { - lean_ctor_release(x_40, 0); - lean_ctor_release(x_40, 1); - lean_ctor_release(x_40, 2); - lean_ctor_release(x_40, 3); - x_164 = x_40; -} else { - lean_dec_ref(x_40); - x_164 = lean_box(0); } -lean_ctor_set_uint8(x_163, sizeof(void*)*4, x_162); -if (lean_is_scalar(x_164)) { - x_165 = lean_alloc_ctor(1, 4, 1); -} else { - x_165 = x_164; +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; +} } -lean_ctor_set(x_165, 0, x_160); -lean_ctor_set(x_165, 1, x_34); -lean_ctor_set(x_165, 2, x_35); -lean_ctor_set(x_165, 3, x_36); -lean_ctor_set_uint8(x_165, sizeof(void*)*4, x_162); -x_166 = 0; -x_167 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_167, 0, x_163); -lean_ctor_set(x_167, 1, x_158); -lean_ctor_set(x_167, 2, x_159); -lean_ctor_set(x_167, 3, x_165); -lean_ctor_set_uint8(x_167, sizeof(void*)*4, x_166); -return x_167; } } else { -uint8_t x_168; -x_168 = !lean_is_exclusive(x_38); -if (x_168 == 0) +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_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___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_169; lean_object* x_170; uint8_t x_171; -x_169 = lean_ctor_get(x_38, 3); -lean_dec(x_169); -x_170 = lean_ctor_get(x_38, 0); -lean_dec(x_170); -x_171 = !lean_is_exclusive(x_40); -if (x_171 == 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 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_4, x_60); +if (x_63 == 0) { -uint8_t x_172; -lean_ctor_set_uint8(x_40, sizeof(void*)*4, x_126); -x_172 = 1; -lean_ctor_set(x_1, 0, x_38); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_172); -return x_1; +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_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; uint8_t x_178; -x_173 = lean_ctor_get(x_40, 0); -x_174 = lean_ctor_get(x_40, 1); -x_175 = lean_ctor_get(x_40, 2); -x_176 = lean_ctor_get(x_40, 3); -lean_inc(x_176); -lean_inc(x_175); -lean_inc(x_174); -lean_inc(x_173); -lean_dec(x_40); -x_177 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_177, 0, x_173); -lean_ctor_set(x_177, 1, x_174); -lean_ctor_set(x_177, 2, x_175); -lean_ctor_set(x_177, 3, x_176); -lean_ctor_set_uint8(x_177, sizeof(void*)*4, x_126); -lean_ctor_set(x_38, 0, x_177); -x_178 = 1; -lean_ctor_set(x_1, 0, x_38); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_178); -return x_1; +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; } } -else +case 1: { -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; uint8_t x_188; -x_179 = lean_ctor_get(x_38, 1); -x_180 = lean_ctor_get(x_38, 2); -lean_inc(x_180); -lean_inc(x_179); -lean_dec(x_38); -x_181 = lean_ctor_get(x_40, 0); -lean_inc(x_181); -x_182 = lean_ctor_get(x_40, 1); -lean_inc(x_182); -x_183 = lean_ctor_get(x_40, 2); -lean_inc(x_183); -x_184 = lean_ctor_get(x_40, 3); -lean_inc(x_184); -if (lean_is_exclusive(x_40)) { - lean_ctor_release(x_40, 0); - lean_ctor_release(x_40, 1); - lean_ctor_release(x_40, 2); - lean_ctor_release(x_40, 3); - x_185 = x_40; +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_40); - x_185 = lean_box(0); + lean_dec_ref(x_57); + x_72 = lean_box(0); } -if (lean_is_scalar(x_185)) { - x_186 = lean_alloc_ctor(1, 4, 1); +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_Grind_registerParent___spec__4(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_186 = x_185; + x_76 = x_72; } -lean_ctor_set(x_186, 0, x_181); -lean_ctor_set(x_186, 1, x_182); -lean_ctor_set(x_186, 2, x_183); -lean_ctor_set(x_186, 3, x_184); -lean_ctor_set_uint8(x_186, sizeof(void*)*4, x_126); -x_187 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_187, 0, x_186); -lean_ctor_set(x_187, 1, x_179); -lean_ctor_set(x_187, 2, x_180); -lean_ctor_set(x_187, 3, x_117); -lean_ctor_set_uint8(x_187, sizeof(void*)*4, x_39); -x_188 = 1; -lean_ctor_set(x_1, 0, x_187); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_188); -return x_1; +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; } } } @@ -11906,1590 +17110,1493 @@ return x_1; } else { -uint8_t x_189; -x_189 = 1; -lean_ctor_set(x_1, 0, x_38); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_189); -return x_1; -} -} -case 1: +uint8_t x_82; +x_82 = !lean_is_exclusive(x_1); +if (x_82 == 0) { -uint8_t x_190; -lean_dec(x_35); -lean_dec(x_34); -x_190 = 1; -lean_ctor_set(x_1, 2, x_3); -lean_ctor_set(x_1, 1, x_2); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_190); -return x_1; +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_Grind_registerParent___spec__6(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_Grind_mkHCongrWithArity___spec__2___closed__3; +x_93 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_registerParent___spec__5(x_3, x_90, x_91, lean_box(0), x_83, x_92); +lean_dec(x_91); +lean_dec(x_90); +return x_93; } -default: +else { -lean_object* x_191; uint8_t x_192; -x_191 = l_Lean_RBNode_ins___at_Lean_Meta_Grind_registerParent___spec__2(x_36, x_2, x_3); -x_192 = lean_ctor_get_uint8(x_191, sizeof(void*)*4); -if (x_192 == 0) +return x_84; +} +} +else { -lean_object* x_193; -x_193 = lean_ctor_get(x_191, 0); -lean_inc(x_193); -if (lean_obj_tag(x_193) == 0) +return x_84; +} +} +else { -lean_object* x_194; -x_194 = lean_ctor_get(x_191, 3); -lean_inc(x_194); -if (lean_obj_tag(x_194) == 0) +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_Grind_registerParent___spec__6(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) { -uint8_t x_195; -x_195 = !lean_is_exclusive(x_191); -if (x_195 == 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_196; lean_object* x_197; uint8_t x_198; -x_196 = lean_ctor_get(x_191, 3); -lean_dec(x_196); -x_197 = lean_ctor_get(x_191, 0); -lean_dec(x_197); -lean_ctor_set(x_191, 0, x_194); -x_198 = 1; -lean_ctor_set(x_1, 3, x_191); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_198); -return x_1; +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_Grind_mkHCongrWithArity___spec__2___closed__3; +x_107 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_registerParent___spec__5(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 { -lean_object* x_199; lean_object* x_200; lean_object* x_201; uint8_t x_202; -x_199 = lean_ctor_get(x_191, 1); -x_200 = lean_ctor_get(x_191, 2); -lean_inc(x_200); -lean_inc(x_199); -lean_dec(x_191); -x_201 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_201, 0, x_194); -lean_ctor_set(x_201, 1, x_199); -lean_ctor_set(x_201, 2, x_200); -lean_ctor_set(x_201, 3, x_194); -lean_ctor_set_uint8(x_201, sizeof(void*)*4, x_192); -x_202 = 1; -lean_ctor_set(x_1, 3, x_201); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_202); -return x_1; +return x_98; } } else { -uint8_t x_203; -x_203 = lean_ctor_get_uint8(x_194, sizeof(void*)*4); -if (x_203 == 0) +return x_98; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_registerParent___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -uint8_t x_204; -x_204 = !lean_is_exclusive(x_191); -if (x_204 == 0) +uint64_t x_4; size_t x_5; size_t x_6; lean_object* x_7; +x_4 = l_Lean_Meta_Grind_instHashableENodeKey_unsafe__1(x_2); +x_5 = lean_uint64_to_usize(x_4); +x_6 = 1; +x_7 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_registerParent___spec__4(x_1, x_5, x_6, x_2, x_3); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_registerParent___spec__9(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: { -lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; uint8_t x_209; -x_205 = lean_ctor_get(x_191, 1); -x_206 = lean_ctor_get(x_191, 2); -x_207 = lean_ctor_get(x_191, 3); -lean_dec(x_207); -x_208 = lean_ctor_get(x_191, 0); -lean_dec(x_208); -x_209 = !lean_is_exclusive(x_194); -if (x_209 == 0) +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_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; uint8_t x_214; uint8_t x_215; -x_210 = lean_ctor_get(x_194, 0); -x_211 = lean_ctor_get(x_194, 1); -x_212 = lean_ctor_get(x_194, 2); -x_213 = lean_ctor_get(x_194, 3); -x_214 = 1; -lean_ctor_set(x_194, 3, x_193); -lean_ctor_set(x_194, 2, x_35); -lean_ctor_set(x_194, 1, x_34); -lean_ctor_set(x_194, 0, x_33); -lean_ctor_set_uint8(x_194, sizeof(void*)*4, x_214); -lean_ctor_set(x_191, 3, x_213); -lean_ctor_set(x_191, 2, x_212); -lean_ctor_set(x_191, 1, x_211); -lean_ctor_set(x_191, 0, x_210); -lean_ctor_set_uint8(x_191, sizeof(void*)*4, x_214); -x_215 = 0; -lean_ctor_set(x_1, 3, x_191); -lean_ctor_set(x_1, 2, x_206); -lean_ctor_set(x_1, 1, x_205); -lean_ctor_set(x_1, 0, x_194); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_215); -return x_1; +lean_object* x_8; +lean_dec(x_4); +x_8 = lean_box(0); +return x_8; } else { -lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; uint8_t x_220; lean_object* x_221; uint8_t x_222; -x_216 = lean_ctor_get(x_194, 0); -x_217 = lean_ctor_get(x_194, 1); -x_218 = lean_ctor_get(x_194, 2); -x_219 = lean_ctor_get(x_194, 3); -lean_inc(x_219); -lean_inc(x_218); -lean_inc(x_217); -lean_inc(x_216); -lean_dec(x_194); -x_220 = 1; -x_221 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_221, 0, x_33); -lean_ctor_set(x_221, 1, x_34); -lean_ctor_set(x_221, 2, x_35); -lean_ctor_set(x_221, 3, x_193); -lean_ctor_set_uint8(x_221, sizeof(void*)*4, x_220); -lean_ctor_set(x_191, 3, x_219); -lean_ctor_set(x_191, 2, x_218); -lean_ctor_set(x_191, 1, x_217); -lean_ctor_set(x_191, 0, x_216); -lean_ctor_set_uint8(x_191, sizeof(void*)*4, x_220); -x_222 = 0; -lean_ctor_set(x_1, 3, x_191); -lean_ctor_set(x_1, 2, x_206); -lean_ctor_set(x_1, 1, x_205); -lean_ctor_set(x_1, 0, x_221); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_222); -return x_1; -} +lean_object* x_9; uint8_t x_10; +x_9 = lean_array_fget(x_1, x_4); +x_10 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(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_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; uint8_t x_230; lean_object* x_231; lean_object* x_232; uint8_t x_233; -x_223 = lean_ctor_get(x_191, 1); -x_224 = lean_ctor_get(x_191, 2); -lean_inc(x_224); -lean_inc(x_223); -lean_dec(x_191); -x_225 = lean_ctor_get(x_194, 0); -lean_inc(x_225); -x_226 = lean_ctor_get(x_194, 1); -lean_inc(x_226); -x_227 = lean_ctor_get(x_194, 2); -lean_inc(x_227); -x_228 = lean_ctor_get(x_194, 3); -lean_inc(x_228); -if (lean_is_exclusive(x_194)) { - lean_ctor_release(x_194, 0); - lean_ctor_release(x_194, 1); - lean_ctor_release(x_194, 2); - lean_ctor_release(x_194, 3); - x_229 = x_194; -} else { - lean_dec_ref(x_194); - x_229 = lean_box(0); +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; } -x_230 = 1; -if (lean_is_scalar(x_229)) { - x_231 = lean_alloc_ctor(1, 4, 1); -} else { - x_231 = x_229; } -lean_ctor_set(x_231, 0, x_33); -lean_ctor_set(x_231, 1, x_34); -lean_ctor_set(x_231, 2, x_35); -lean_ctor_set(x_231, 3, x_193); -lean_ctor_set_uint8(x_231, sizeof(void*)*4, x_230); -x_232 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_232, 0, x_225); -lean_ctor_set(x_232, 1, x_226); -lean_ctor_set(x_232, 2, x_227); -lean_ctor_set(x_232, 3, x_228); -lean_ctor_set_uint8(x_232, sizeof(void*)*4, x_230); -x_233 = 0; -lean_ctor_set(x_1, 3, x_232); -lean_ctor_set(x_1, 2, x_224); -lean_ctor_set(x_1, 1, x_223); -lean_ctor_set(x_1, 0, x_231); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_233); -return x_1; } } -else +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_registerParent___spec__8(lean_object* x_1, size_t x_2, lean_object* x_3) { +_start: { -uint8_t x_234; -lean_free_object(x_1); -x_234 = !lean_is_exclusive(x_194); -if (x_234 == 0) +if (lean_obj_tag(x_1) == 0) { -lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; uint8_t x_239; -x_235 = lean_ctor_get(x_194, 3); -lean_dec(x_235); -x_236 = lean_ctor_get(x_194, 2); -lean_dec(x_236); -x_237 = lean_ctor_get(x_194, 1); -lean_dec(x_237); -x_238 = lean_ctor_get(x_194, 0); -lean_dec(x_238); -x_239 = 1; -lean_ctor_set(x_194, 3, x_191); -lean_ctor_set(x_194, 2, x_35); -lean_ctor_set(x_194, 1, x_34); -lean_ctor_set(x_194, 0, x_33); -lean_ctor_set_uint8(x_194, sizeof(void*)*4, x_239); -return x_194; +uint8_t x_4; +x_4 = !lean_is_exclusive(x_1); +if (x_4 == 0) +{ +lean_object* x_5; size_t x_6; size_t x_7; size_t x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_5 = lean_ctor_get(x_1, 0); +x_6 = 5; +x_7 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__2___closed__2; +x_8 = lean_usize_land(x_2, x_7); +x_9 = lean_usize_to_nat(x_8); +x_10 = lean_box(2); +x_11 = lean_array_get(x_10, x_5, x_9); +lean_dec(x_9); +lean_dec(x_5); +switch (lean_obj_tag(x_11)) { +case 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 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_3, x_12); +lean_dec(x_12); +if (x_14 == 0) +{ +lean_object* x_15; +lean_dec(x_13); +lean_free_object(x_1); +x_15 = lean_box(0); +return x_15; } else { -uint8_t x_240; lean_object* x_241; -lean_dec(x_194); -x_240 = 1; -x_241 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_241, 0, x_33); -lean_ctor_set(x_241, 1, x_34); -lean_ctor_set(x_241, 2, x_35); -lean_ctor_set(x_241, 3, x_191); -lean_ctor_set_uint8(x_241, sizeof(void*)*4, x_240); -return x_241; +lean_ctor_set_tag(x_1, 1); +lean_ctor_set(x_1, 0, x_13); +return x_1; +} +} +case 1: +{ +lean_object* x_16; size_t x_17; +lean_free_object(x_1); +x_16 = lean_ctor_get(x_11, 0); +lean_inc(x_16); +lean_dec(x_11); +x_17 = lean_usize_shift_right(x_2, x_6); +x_1 = x_16; +x_2 = x_17; +goto _start; } +default: +{ +lean_object* x_19; +lean_free_object(x_1); +x_19 = lean_box(0); +return x_19; } } } else { -uint8_t x_242; -x_242 = lean_ctor_get_uint8(x_193, sizeof(void*)*4); -if (x_242 == 0) -{ -uint8_t x_243; -x_243 = !lean_is_exclusive(x_191); -if (x_243 == 0) +lean_object* x_20; size_t x_21; size_t x_22; size_t x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_20 = lean_ctor_get(x_1, 0); +lean_inc(x_20); +lean_dec(x_1); +x_21 = 5; +x_22 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__2___closed__2; +x_23 = lean_usize_land(x_2, x_22); +x_24 = lean_usize_to_nat(x_23); +x_25 = lean_box(2); +x_26 = lean_array_get(x_25, x_20, x_24); +lean_dec(x_24); +lean_dec(x_20); +switch (lean_obj_tag(x_26)) { +case 0: { -lean_object* x_244; uint8_t x_245; -x_244 = lean_ctor_get(x_191, 0); -lean_dec(x_244); -x_245 = !lean_is_exclusive(x_193); -if (x_245 == 0) +lean_object* x_27; lean_object* x_28; uint8_t 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_Grind_isSameExpr_unsafe__1(x_3, x_27); +lean_dec(x_27); +if (x_29 == 0) { -lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; uint8_t x_250; uint8_t x_251; -x_246 = lean_ctor_get(x_193, 0); -x_247 = lean_ctor_get(x_193, 1); -x_248 = lean_ctor_get(x_193, 2); -x_249 = lean_ctor_get(x_193, 3); -x_250 = 1; -lean_ctor_set(x_193, 3, x_246); -lean_ctor_set(x_193, 2, x_35); -lean_ctor_set(x_193, 1, x_34); -lean_ctor_set(x_193, 0, x_33); -lean_ctor_set_uint8(x_193, sizeof(void*)*4, x_250); -lean_ctor_set(x_191, 0, x_249); -lean_ctor_set_uint8(x_191, sizeof(void*)*4, x_250); -x_251 = 0; -lean_ctor_set(x_1, 3, x_191); -lean_ctor_set(x_1, 2, x_248); -lean_ctor_set(x_1, 1, x_247); -lean_ctor_set(x_1, 0, x_193); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_251); -return x_1; +lean_object* x_30; +lean_dec(x_28); +x_30 = lean_box(0); +return x_30; } else { -lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; uint8_t x_256; lean_object* x_257; uint8_t x_258; -x_252 = lean_ctor_get(x_193, 0); -x_253 = lean_ctor_get(x_193, 1); -x_254 = lean_ctor_get(x_193, 2); -x_255 = lean_ctor_get(x_193, 3); -lean_inc(x_255); -lean_inc(x_254); -lean_inc(x_253); -lean_inc(x_252); -lean_dec(x_193); -x_256 = 1; -x_257 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_257, 0, x_33); -lean_ctor_set(x_257, 1, x_34); -lean_ctor_set(x_257, 2, x_35); -lean_ctor_set(x_257, 3, x_252); -lean_ctor_set_uint8(x_257, sizeof(void*)*4, x_256); -lean_ctor_set(x_191, 0, x_255); -lean_ctor_set_uint8(x_191, sizeof(void*)*4, x_256); -x_258 = 0; -lean_ctor_set(x_1, 3, x_191); -lean_ctor_set(x_1, 2, x_254); -lean_ctor_set(x_1, 1, x_253); -lean_ctor_set(x_1, 0, x_257); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_258); -return x_1; +lean_object* x_31; +x_31 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_31, 0, x_28); +return x_31; } } -else +case 1: { -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; lean_object* x_269; uint8_t x_270; -x_259 = lean_ctor_get(x_191, 1); -x_260 = lean_ctor_get(x_191, 2); -x_261 = lean_ctor_get(x_191, 3); -lean_inc(x_261); -lean_inc(x_260); -lean_inc(x_259); -lean_dec(x_191); -x_262 = lean_ctor_get(x_193, 0); -lean_inc(x_262); -x_263 = lean_ctor_get(x_193, 1); -lean_inc(x_263); -x_264 = lean_ctor_get(x_193, 2); -lean_inc(x_264); -x_265 = lean_ctor_get(x_193, 3); -lean_inc(x_265); -if (lean_is_exclusive(x_193)) { - lean_ctor_release(x_193, 0); - lean_ctor_release(x_193, 1); - lean_ctor_release(x_193, 2); - lean_ctor_release(x_193, 3); - x_266 = x_193; -} else { - lean_dec_ref(x_193); - x_266 = lean_box(0); +lean_object* x_32; size_t x_33; +x_32 = lean_ctor_get(x_26, 0); +lean_inc(x_32); +lean_dec(x_26); +x_33 = lean_usize_shift_right(x_2, x_21); +x_1 = x_32; +x_2 = x_33; +goto _start; +} +default: +{ +lean_object* x_35; +x_35 = lean_box(0); +return x_35; } -x_267 = 1; -if (lean_is_scalar(x_266)) { - x_268 = lean_alloc_ctor(1, 4, 1); -} else { - x_268 = x_266; } -lean_ctor_set(x_268, 0, x_33); -lean_ctor_set(x_268, 1, x_34); -lean_ctor_set(x_268, 2, x_35); -lean_ctor_set(x_268, 3, x_262); -lean_ctor_set_uint8(x_268, sizeof(void*)*4, x_267); -x_269 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_269, 0, x_265); -lean_ctor_set(x_269, 1, x_259); -lean_ctor_set(x_269, 2, x_260); -lean_ctor_set(x_269, 3, x_261); -lean_ctor_set_uint8(x_269, sizeof(void*)*4, x_267); -x_270 = 0; -lean_ctor_set(x_1, 3, x_269); -lean_ctor_set(x_1, 2, x_264); -lean_ctor_set(x_1, 1, x_263); -lean_ctor_set(x_1, 0, x_268); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_270); -return x_1; } } else { -lean_object* x_271; -x_271 = lean_ctor_get(x_191, 3); -lean_inc(x_271); -if (lean_obj_tag(x_271) == 0) +lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_36 = lean_ctor_get(x_1, 0); +lean_inc(x_36); +x_37 = lean_ctor_get(x_1, 1); +lean_inc(x_37); +lean_dec(x_1); +x_38 = lean_unsigned_to_nat(0u); +x_39 = l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_registerParent___spec__9(x_36, x_37, lean_box(0), x_38, x_3); +lean_dec(x_37); +lean_dec(x_36); +return x_39; +} +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_registerParent___spec__7(lean_object* x_1, lean_object* x_2) { +_start: { -uint8_t x_272; -lean_free_object(x_1); -x_272 = !lean_is_exclusive(x_193); -if (x_272 == 0) +uint64_t x_3; size_t x_4; lean_object* x_5; +x_3 = l_Lean_Meta_Grind_instHashableENodeKey_unsafe__1(x_2); +x_4 = lean_uint64_to_usize(x_3); +x_5 = l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_registerParent___spec__8(x_1, x_4, x_2); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_registerParent(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, 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_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; uint8_t x_277; -x_273 = lean_ctor_get(x_193, 3); -lean_dec(x_273); -x_274 = lean_ctor_get(x_193, 2); -lean_dec(x_274); -x_275 = lean_ctor_get(x_193, 1); -lean_dec(x_275); -x_276 = lean_ctor_get(x_193, 0); -lean_dec(x_276); -x_277 = 1; -lean_ctor_set(x_193, 3, x_191); -lean_ctor_set(x_193, 2, x_35); -lean_ctor_set(x_193, 1, x_34); -lean_ctor_set(x_193, 0, x_33); -lean_ctor_set_uint8(x_193, sizeof(void*)*4, x_277); -return x_193; +lean_object* x_12; lean_object* x_13; +x_12 = l_Lean_Meta_Grind_getRoot_x3f(x_2, 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); +if (lean_obj_tag(x_13) == 0) +{ +uint8_t x_14; +lean_dec(x_1); +x_14 = !lean_is_exclusive(x_12); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; +x_15 = lean_ctor_get(x_12, 0); +lean_dec(x_15); +x_16 = lean_box(0); +lean_ctor_set(x_12, 0, x_16); +return x_12; } else { -uint8_t x_278; lean_object* x_279; -lean_dec(x_193); -x_278 = 1; -x_279 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_279, 0, x_33); -lean_ctor_set(x_279, 1, x_34); -lean_ctor_set(x_279, 2, x_35); -lean_ctor_set(x_279, 3, x_191); -lean_ctor_set_uint8(x_279, sizeof(void*)*4, x_278); -return x_279; +lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_17 = lean_ctor_get(x_12, 1); +lean_inc(x_17); +lean_dec(x_12); +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 { -uint8_t x_280; -x_280 = lean_ctor_get_uint8(x_271, sizeof(void*)*4); -if (x_280 == 0) -{ -uint8_t x_281; -lean_free_object(x_1); -x_281 = !lean_is_exclusive(x_191); -if (x_281 == 0) -{ -lean_object* x_282; lean_object* x_283; uint8_t x_284; -x_282 = lean_ctor_get(x_191, 3); -lean_dec(x_282); -x_283 = lean_ctor_get(x_191, 0); -lean_dec(x_283); -x_284 = !lean_is_exclusive(x_271); -if (x_284 == 0) -{ -lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; uint8_t x_289; uint8_t x_290; -x_285 = lean_ctor_get(x_271, 0); -x_286 = lean_ctor_get(x_271, 1); -x_287 = lean_ctor_get(x_271, 2); -x_288 = lean_ctor_get(x_271, 3); -x_289 = 1; -lean_inc(x_193); -lean_ctor_set(x_271, 3, x_193); -lean_ctor_set(x_271, 2, x_35); -lean_ctor_set(x_271, 1, x_34); -lean_ctor_set(x_271, 0, x_33); -x_290 = !lean_is_exclusive(x_193); -if (x_290 == 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_69; lean_object* x_70; +x_20 = lean_ctor_get(x_12, 1); +lean_inc(x_20); +lean_dec(x_12); +x_21 = lean_ctor_get(x_13, 0); +lean_inc(x_21); +lean_dec(x_13); +x_22 = lean_st_ref_get(x_3, x_20); +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +x_24 = lean_ctor_get(x_22, 1); +lean_inc(x_24); +lean_dec(x_22); +x_25 = lean_ctor_get(x_23, 2); +lean_inc(x_25); +lean_dec(x_23); +x_69 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_registerParent___spec__7(x_25, x_21); +x_70 = lean_st_ref_take(x_3, x_24); +if (lean_obj_tag(x_69) == 0) { -lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; uint8_t x_295; -x_291 = lean_ctor_get(x_193, 3); -lean_dec(x_291); -x_292 = lean_ctor_get(x_193, 2); -lean_dec(x_292); -x_293 = lean_ctor_get(x_193, 1); -lean_dec(x_293); -x_294 = lean_ctor_get(x_193, 0); -lean_dec(x_294); -lean_ctor_set_uint8(x_271, sizeof(void*)*4, x_289); -lean_ctor_set(x_193, 3, x_288); -lean_ctor_set(x_193, 2, x_287); -lean_ctor_set(x_193, 1, x_286); -lean_ctor_set(x_193, 0, x_285); -lean_ctor_set_uint8(x_193, sizeof(void*)*4, x_289); -x_295 = 0; -lean_ctor_set(x_191, 3, x_193); -lean_ctor_set(x_191, 0, x_271); -lean_ctor_set_uint8(x_191, sizeof(void*)*4, x_295); -return x_191; +lean_object* x_71; lean_object* x_72; lean_object* x_73; +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_box(0); +x_26 = x_73; +x_27 = x_71; +x_28 = x_72; +goto block_68; } else { -lean_object* x_296; uint8_t x_297; -lean_dec(x_193); -lean_ctor_set_uint8(x_271, sizeof(void*)*4, x_289); -x_296 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_296, 0, x_285); -lean_ctor_set(x_296, 1, x_286); -lean_ctor_set(x_296, 2, x_287); -lean_ctor_set(x_296, 3, x_288); -lean_ctor_set_uint8(x_296, sizeof(void*)*4, x_289); -x_297 = 0; -lean_ctor_set(x_191, 3, x_296); -lean_ctor_set(x_191, 0, x_271); -lean_ctor_set_uint8(x_191, sizeof(void*)*4, x_297); -return x_191; +lean_object* x_74; lean_object* x_75; lean_object* x_76; +x_74 = lean_ctor_get(x_69, 0); +lean_inc(x_74); +lean_dec(x_69); +x_75 = lean_ctor_get(x_70, 0); +lean_inc(x_75); +x_76 = lean_ctor_get(x_70, 1); +lean_inc(x_76); +lean_dec(x_70); +x_26 = x_74; +x_27 = x_75; +x_28 = x_76; +goto block_68; } +block_68: +{ +uint8_t x_29; +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; uint8_t x_35; +x_30 = lean_ctor_get(x_27, 2); +x_31 = lean_box(0); +x_32 = l_Lean_RBNode_insert___at_Lean_Meta_Grind_registerParent___spec__1(x_26, x_1, x_31); +x_33 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_registerParent___spec__3(x_30, x_21, x_32); +lean_ctor_set(x_27, 2, x_33); +x_34 = lean_st_ref_set(x_3, x_27, x_28); +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_31); +return x_34; } else { -lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; uint8_t x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; uint8_t x_306; -x_298 = lean_ctor_get(x_271, 0); -x_299 = lean_ctor_get(x_271, 1); -x_300 = lean_ctor_get(x_271, 2); -x_301 = lean_ctor_get(x_271, 3); -lean_inc(x_301); -lean_inc(x_300); -lean_inc(x_299); -lean_inc(x_298); -lean_dec(x_271); -x_302 = 1; -lean_inc(x_193); -x_303 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_303, 0, x_33); -lean_ctor_set(x_303, 1, x_34); -lean_ctor_set(x_303, 2, x_35); -lean_ctor_set(x_303, 3, x_193); -if (lean_is_exclusive(x_193)) { - lean_ctor_release(x_193, 0); - lean_ctor_release(x_193, 1); - lean_ctor_release(x_193, 2); - lean_ctor_release(x_193, 3); - x_304 = x_193; -} else { - lean_dec_ref(x_193); - x_304 = lean_box(0); -} -lean_ctor_set_uint8(x_303, sizeof(void*)*4, x_302); -if (lean_is_scalar(x_304)) { - x_305 = lean_alloc_ctor(1, 4, 1); -} else { - x_305 = x_304; -} -lean_ctor_set(x_305, 0, x_298); -lean_ctor_set(x_305, 1, x_299); -lean_ctor_set(x_305, 2, x_300); -lean_ctor_set(x_305, 3, x_301); -lean_ctor_set_uint8(x_305, sizeof(void*)*4, x_302); -x_306 = 0; -lean_ctor_set(x_191, 3, x_305); -lean_ctor_set(x_191, 0, x_303); -lean_ctor_set_uint8(x_191, sizeof(void*)*4, x_306); -return x_191; +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_31); +lean_ctor_set(x_38, 1, x_37); +return x_38; } } 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; uint8_t x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; uint8_t x_318; lean_object* x_319; -x_307 = lean_ctor_get(x_191, 1); -x_308 = lean_ctor_get(x_191, 2); -lean_inc(x_308); -lean_inc(x_307); -lean_dec(x_191); -x_309 = lean_ctor_get(x_271, 0); -lean_inc(x_309); -x_310 = lean_ctor_get(x_271, 1); -lean_inc(x_310); -x_311 = lean_ctor_get(x_271, 2); -lean_inc(x_311); -x_312 = lean_ctor_get(x_271, 3); -lean_inc(x_312); -if (lean_is_exclusive(x_271)) { - lean_ctor_release(x_271, 0); - lean_ctor_release(x_271, 1); - lean_ctor_release(x_271, 2); - lean_ctor_release(x_271, 3); - x_313 = x_271; -} else { - lean_dec_ref(x_271); - x_313 = lean_box(0); -} -x_314 = 1; -lean_inc(x_193); -if (lean_is_scalar(x_313)) { - x_315 = lean_alloc_ctor(1, 4, 1); -} else { - x_315 = x_313; -} -lean_ctor_set(x_315, 0, x_33); -lean_ctor_set(x_315, 1, x_34); -lean_ctor_set(x_315, 2, x_35); -lean_ctor_set(x_315, 3, x_193); -if (lean_is_exclusive(x_193)) { - lean_ctor_release(x_193, 0); - lean_ctor_release(x_193, 1); - lean_ctor_release(x_193, 2); - lean_ctor_release(x_193, 3); - x_316 = x_193; +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; lean_object* x_61; 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_39 = lean_ctor_get(x_27, 0); +x_40 = lean_ctor_get(x_27, 1); +x_41 = lean_ctor_get(x_27, 2); +x_42 = lean_ctor_get(x_27, 3); +x_43 = lean_ctor_get(x_27, 4); +x_44 = lean_ctor_get(x_27, 5); +x_45 = lean_ctor_get_uint8(x_27, sizeof(void*)*20); +x_46 = lean_ctor_get(x_27, 6); +x_47 = lean_ctor_get(x_27, 7); +x_48 = lean_ctor_get(x_27, 8); +x_49 = lean_ctor_get(x_27, 9); +x_50 = lean_ctor_get(x_27, 10); +x_51 = lean_ctor_get(x_27, 11); +x_52 = lean_ctor_get(x_27, 12); +x_53 = lean_ctor_get(x_27, 13); +x_54 = lean_ctor_get(x_27, 14); +x_55 = lean_ctor_get(x_27, 15); +x_56 = lean_ctor_get(x_27, 16); +x_57 = lean_ctor_get(x_27, 17); +x_58 = lean_ctor_get(x_27, 18); +x_59 = lean_ctor_get(x_27, 19); +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_inc(x_50); +lean_inc(x_49); +lean_inc(x_48); +lean_inc(x_47); +lean_inc(x_46); +lean_inc(x_44); +lean_inc(x_43); +lean_inc(x_42); +lean_inc(x_41); +lean_inc(x_40); +lean_inc(x_39); +lean_dec(x_27); +x_60 = lean_box(0); +x_61 = l_Lean_RBNode_insert___at_Lean_Meta_Grind_registerParent___spec__1(x_26, x_1, x_60); +x_62 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_registerParent___spec__3(x_41, x_21, x_61); +x_63 = lean_alloc_ctor(0, 20, 1); +lean_ctor_set(x_63, 0, x_39); +lean_ctor_set(x_63, 1, x_40); +lean_ctor_set(x_63, 2, x_62); +lean_ctor_set(x_63, 3, x_42); +lean_ctor_set(x_63, 4, x_43); +lean_ctor_set(x_63, 5, x_44); +lean_ctor_set(x_63, 6, x_46); +lean_ctor_set(x_63, 7, x_47); +lean_ctor_set(x_63, 8, x_48); +lean_ctor_set(x_63, 9, x_49); +lean_ctor_set(x_63, 10, x_50); +lean_ctor_set(x_63, 11, x_51); +lean_ctor_set(x_63, 12, x_52); +lean_ctor_set(x_63, 13, x_53); +lean_ctor_set(x_63, 14, x_54); +lean_ctor_set(x_63, 15, x_55); +lean_ctor_set(x_63, 16, x_56); +lean_ctor_set(x_63, 17, x_57); +lean_ctor_set(x_63, 18, x_58); +lean_ctor_set(x_63, 19, x_59); +lean_ctor_set_uint8(x_63, sizeof(void*)*20, x_45); +x_64 = lean_st_ref_set(x_3, x_63, x_28); +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_193); - x_316 = lean_box(0); + lean_dec_ref(x_64); + x_66 = lean_box(0); } -lean_ctor_set_uint8(x_315, sizeof(void*)*4, x_314); -if (lean_is_scalar(x_316)) { - x_317 = lean_alloc_ctor(1, 4, 1); +if (lean_is_scalar(x_66)) { + x_67 = lean_alloc_ctor(0, 2, 0); } else { - x_317 = x_316; + x_67 = x_66; } -lean_ctor_set(x_317, 0, x_309); -lean_ctor_set(x_317, 1, x_310); -lean_ctor_set(x_317, 2, x_311); -lean_ctor_set(x_317, 3, x_312); -lean_ctor_set_uint8(x_317, sizeof(void*)*4, x_314); -x_318 = 0; -x_319 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_319, 0, x_315); -lean_ctor_set(x_319, 1, x_307); -lean_ctor_set(x_319, 2, x_308); -lean_ctor_set(x_319, 3, x_317); -lean_ctor_set_uint8(x_319, sizeof(void*)*4, x_318); -return x_319; +lean_ctor_set(x_67, 0, x_60); +lean_ctor_set(x_67, 1, x_65); +return x_67; } } -else -{ -uint8_t x_320; -x_320 = !lean_is_exclusive(x_191); -if (x_320 == 0) -{ -lean_object* x_321; lean_object* x_322; uint8_t x_323; -x_321 = lean_ctor_get(x_191, 3); -lean_dec(x_321); -x_322 = lean_ctor_get(x_191, 0); -lean_dec(x_322); -x_323 = !lean_is_exclusive(x_193); -if (x_323 == 0) -{ -uint8_t x_324; -lean_ctor_set_uint8(x_193, sizeof(void*)*4, x_280); -x_324 = 1; -lean_ctor_set(x_1, 3, x_191); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_324); -return x_1; } -else -{ -lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; uint8_t x_330; -x_325 = lean_ctor_get(x_193, 0); -x_326 = lean_ctor_get(x_193, 1); -x_327 = lean_ctor_get(x_193, 2); -x_328 = lean_ctor_get(x_193, 3); -lean_inc(x_328); -lean_inc(x_327); -lean_inc(x_326); -lean_inc(x_325); -lean_dec(x_193); -x_329 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_329, 0, x_325); -lean_ctor_set(x_329, 1, x_326); -lean_ctor_set(x_329, 2, x_327); -lean_ctor_set(x_329, 3, x_328); -lean_ctor_set_uint8(x_329, sizeof(void*)*4, x_280); -lean_ctor_set(x_191, 0, x_329); -x_330 = 1; -lean_ctor_set(x_1, 3, x_191); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_330); -return x_1; } } -else +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_registerParent___spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: { -lean_object* x_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; uint8_t x_340; -x_331 = lean_ctor_get(x_191, 1); -x_332 = lean_ctor_get(x_191, 2); -lean_inc(x_332); -lean_inc(x_331); -lean_dec(x_191); -x_333 = lean_ctor_get(x_193, 0); -lean_inc(x_333); -x_334 = lean_ctor_get(x_193, 1); -lean_inc(x_334); -x_335 = lean_ctor_get(x_193, 2); -lean_inc(x_335); -x_336 = lean_ctor_get(x_193, 3); -lean_inc(x_336); -if (lean_is_exclusive(x_193)) { - lean_ctor_release(x_193, 0); - lean_ctor_release(x_193, 1); - lean_ctor_release(x_193, 2); - lean_ctor_release(x_193, 3); - x_337 = x_193; -} else { - lean_dec_ref(x_193); - x_337 = lean_box(0); +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_Grind_registerParent___spec__5(x_7, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_3); +lean_dec(x_2); +return x_8; } -if (lean_is_scalar(x_337)) { - x_338 = lean_alloc_ctor(1, 4, 1); -} else { - x_338 = x_337; } -lean_ctor_set(x_338, 0, x_333); -lean_ctor_set(x_338, 1, x_334); -lean_ctor_set(x_338, 2, x_335); -lean_ctor_set(x_338, 3, x_336); -lean_ctor_set_uint8(x_338, sizeof(void*)*4, x_280); -x_339 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_339, 0, x_338); -lean_ctor_set(x_339, 1, x_331); -lean_ctor_set(x_339, 2, x_332); -lean_ctor_set(x_339, 3, x_271); -lean_ctor_set_uint8(x_339, sizeof(void*)*4, x_192); -x_340 = 1; -lean_ctor_set(x_1, 3, x_339); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_340); -return x_1; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_registerParent___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +size_t x_6; size_t x_7; lean_object* x_8; +x_6 = lean_unbox_usize(x_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_Grind_registerParent___spec__4(x_1, x_6, x_7, x_4, x_5); +return x_8; } } +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_registerParent___spec__9___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; +x_6 = l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_registerParent___spec__9(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_Grind_registerParent___spec__8___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_Grind_registerParent___spec__8(x_1, x_4, x_3); +lean_dec(x_3); +return x_5; } } -else +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_registerParent___spec__7___boxed(lean_object* x_1, lean_object* x_2) { +_start: { -uint8_t x_341; -x_341 = 1; -lean_ctor_set(x_1, 3, x_191); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_341); -return x_1; +lean_object* x_3; +x_3 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_registerParent___spec__7(x_1, x_2); +lean_dec(x_2); +return x_3; } } +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_registerParent___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, 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_Meta_Grind_registerParent(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_12; } } -else -{ -lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; uint8_t x_346; -x_342 = lean_ctor_get(x_1, 0); -x_343 = lean_ctor_get(x_1, 1); -x_344 = lean_ctor_get(x_1, 2); -x_345 = lean_ctor_get(x_1, 3); -lean_inc(x_345); -lean_inc(x_344); -lean_inc(x_343); -lean_inc(x_342); -lean_dec(x_1); -x_346 = l_Lean_Expr_quickComp(x_2, x_343); -switch (x_346) { -case 0: -{ -lean_object* x_347; uint8_t x_348; -x_347 = l_Lean_RBNode_ins___at_Lean_Meta_Grind_registerParent___spec__2(x_342, x_2, x_3); -x_348 = lean_ctor_get_uint8(x_347, sizeof(void*)*4); -if (x_348 == 0) +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getParents(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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_349; -x_349 = lean_ctor_get(x_347, 0); -lean_inc(x_349); -if (lean_obj_tag(x_349) == 0) +lean_object* x_11; uint8_t x_12; +x_11 = lean_st_ref_get(x_2, x_10); +x_12 = !lean_is_exclusive(x_11); +if (x_12 == 0) { -lean_object* x_350; -x_350 = lean_ctor_get(x_347, 3); -lean_inc(x_350); -if (lean_obj_tag(x_350) == 0) +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_13, 2); +lean_inc(x_14); +lean_dec(x_13); +x_15 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_registerParent___spec__7(x_14, x_1); +if (lean_obj_tag(x_15) == 0) { -lean_object* x_351; lean_object* x_352; lean_object* x_353; lean_object* x_354; uint8_t x_355; lean_object* x_356; -x_351 = lean_ctor_get(x_347, 1); -lean_inc(x_351); -x_352 = lean_ctor_get(x_347, 2); -lean_inc(x_352); -if (lean_is_exclusive(x_347)) { - lean_ctor_release(x_347, 0); - lean_ctor_release(x_347, 1); - lean_ctor_release(x_347, 2); - lean_ctor_release(x_347, 3); - x_353 = x_347; -} else { - lean_dec_ref(x_347); - x_353 = lean_box(0); +lean_object* x_16; +x_16 = lean_box(0); +lean_ctor_set(x_11, 0, x_16); +return x_11; } -if (lean_is_scalar(x_353)) { - x_354 = lean_alloc_ctor(1, 4, 1); -} else { - x_354 = x_353; +else +{ +lean_object* x_17; +x_17 = lean_ctor_get(x_15, 0); +lean_inc(x_17); +lean_dec(x_15); +lean_ctor_set(x_11, 0, x_17); +return x_11; } -lean_ctor_set(x_354, 0, x_350); -lean_ctor_set(x_354, 1, x_351); -lean_ctor_set(x_354, 2, x_352); -lean_ctor_set(x_354, 3, x_350); -lean_ctor_set_uint8(x_354, sizeof(void*)*4, x_348); -x_355 = 1; -x_356 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_356, 0, x_354); -lean_ctor_set(x_356, 1, x_343); -lean_ctor_set(x_356, 2, x_344); -lean_ctor_set(x_356, 3, x_345); -lean_ctor_set_uint8(x_356, sizeof(void*)*4, x_355); -return x_356; } else { -uint8_t x_357; -x_357 = lean_ctor_get_uint8(x_350, sizeof(void*)*4); -if (x_357 == 0) +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_18 = lean_ctor_get(x_11, 0); +x_19 = lean_ctor_get(x_11, 1); +lean_inc(x_19); +lean_inc(x_18); +lean_dec(x_11); +x_20 = lean_ctor_get(x_18, 2); +lean_inc(x_20); +lean_dec(x_18); +x_21 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_registerParent___spec__7(x_20, x_1); +if (lean_obj_tag(x_21) == 0) { -lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; uint8_t x_366; lean_object* x_367; lean_object* x_368; uint8_t x_369; lean_object* x_370; -x_358 = lean_ctor_get(x_347, 1); -lean_inc(x_358); -x_359 = lean_ctor_get(x_347, 2); -lean_inc(x_359); -if (lean_is_exclusive(x_347)) { - lean_ctor_release(x_347, 0); - lean_ctor_release(x_347, 1); - lean_ctor_release(x_347, 2); - lean_ctor_release(x_347, 3); - x_360 = x_347; -} else { - lean_dec_ref(x_347); - x_360 = lean_box(0); +lean_object* x_22; lean_object* x_23; +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_19); +return x_23; } -x_361 = lean_ctor_get(x_350, 0); -lean_inc(x_361); -x_362 = lean_ctor_get(x_350, 1); -lean_inc(x_362); -x_363 = lean_ctor_get(x_350, 2); -lean_inc(x_363); -x_364 = lean_ctor_get(x_350, 3); -lean_inc(x_364); -if (lean_is_exclusive(x_350)) { - lean_ctor_release(x_350, 0); - lean_ctor_release(x_350, 1); - lean_ctor_release(x_350, 2); - lean_ctor_release(x_350, 3); - x_365 = x_350; -} else { - lean_dec_ref(x_350); - x_365 = lean_box(0); +else +{ +lean_object* x_24; lean_object* x_25; +x_24 = lean_ctor_get(x_21, 0); +lean_inc(x_24); +lean_dec(x_21); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_19); +return x_25; } -x_366 = 1; -if (lean_is_scalar(x_365)) { - x_367 = lean_alloc_ctor(1, 4, 1); -} else { - x_367 = x_365; } -lean_ctor_set(x_367, 0, x_349); -lean_ctor_set(x_367, 1, x_358); -lean_ctor_set(x_367, 2, x_359); -lean_ctor_set(x_367, 3, x_361); -lean_ctor_set_uint8(x_367, sizeof(void*)*4, x_366); -if (lean_is_scalar(x_360)) { - x_368 = lean_alloc_ctor(1, 4, 1); -} else { - x_368 = x_360; } -lean_ctor_set(x_368, 0, x_364); -lean_ctor_set(x_368, 1, x_343); -lean_ctor_set(x_368, 2, x_344); -lean_ctor_set(x_368, 3, x_345); -lean_ctor_set_uint8(x_368, sizeof(void*)*4, x_366); -x_369 = 0; -x_370 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_370, 0, x_367); -lean_ctor_set(x_370, 1, x_362); -lean_ctor_set(x_370, 2, x_363); -lean_ctor_set(x_370, 3, x_368); -lean_ctor_set_uint8(x_370, sizeof(void*)*4, x_369); -return x_370; +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getParents___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_Meta_Grind_getParents(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Array_indexOfAux___at_Lean_Meta_Grind_getParentsAndReset___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; uint8_t x_5; +x_4 = lean_array_get_size(x_1); +x_5 = lean_nat_dec_lt(x_3, x_4); +lean_dec(x_4); +if (x_5 == 0) +{ +lean_object* x_6; +lean_dec(x_3); +x_6 = lean_box(0); +return x_6; } else { -lean_object* x_371; uint8_t x_372; lean_object* x_373; -if (lean_is_exclusive(x_350)) { - lean_ctor_release(x_350, 0); - lean_ctor_release(x_350, 1); - lean_ctor_release(x_350, 2); - lean_ctor_release(x_350, 3); - x_371 = x_350; -} else { - lean_dec_ref(x_350); - x_371 = lean_box(0); +lean_object* x_7; uint8_t x_8; +x_7 = lean_array_fget(x_1, x_3); +x_8 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_7, x_2); +lean_dec(x_7); +if (x_8 == 0) +{ +lean_object* x_9; lean_object* x_10; +x_9 = lean_unsigned_to_nat(1u); +x_10 = lean_nat_add(x_3, x_9); +lean_dec(x_3); +x_3 = x_10; +goto _start; } -x_372 = 1; -if (lean_is_scalar(x_371)) { - x_373 = lean_alloc_ctor(1, 4, 1); -} else { - x_373 = x_371; +else +{ +lean_object* x_12; +x_12 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_12, 0, x_3); +return x_12; } -lean_ctor_set(x_373, 0, x_347); -lean_ctor_set(x_373, 1, x_343); -lean_ctor_set(x_373, 2, x_344); -lean_ctor_set(x_373, 3, x_345); -lean_ctor_set_uint8(x_373, sizeof(void*)*4, x_372); -return x_373; } } } +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_getParentsAndReset___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); +x_5 = 5; +x_6 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___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); +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 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_3, x_11); +lean_dec(x_11); +if (x_12 == 0) +{ +lean_dec(x_8); +lean_dec(x_4); +return x_1; +} else { -uint8_t x_374; -x_374 = lean_ctor_get_uint8(x_349, sizeof(void*)*4); -if (x_374 == 0) +uint8_t x_13; +x_13 = !lean_is_exclusive(x_1); +if (x_13 == 0) { -lean_object* x_375; lean_object* x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; lean_object* x_382; lean_object* x_383; uint8_t x_384; lean_object* x_385; lean_object* x_386; uint8_t x_387; lean_object* x_388; -x_375 = lean_ctor_get(x_347, 1); -lean_inc(x_375); -x_376 = lean_ctor_get(x_347, 2); -lean_inc(x_376); -x_377 = lean_ctor_get(x_347, 3); -lean_inc(x_377); -if (lean_is_exclusive(x_347)) { - lean_ctor_release(x_347, 0); - lean_ctor_release(x_347, 1); - lean_ctor_release(x_347, 2); - lean_ctor_release(x_347, 3); - x_378 = x_347; -} else { - lean_dec_ref(x_347); - x_378 = lean_box(0); +lean_object* x_14; lean_object* x_15; +x_14 = lean_ctor_get(x_1, 0); +lean_dec(x_14); +x_15 = lean_array_set(x_4, x_8, x_9); +lean_dec(x_8); +lean_ctor_set(x_1, 0, x_15); +return x_1; } -x_379 = lean_ctor_get(x_349, 0); -lean_inc(x_379); -x_380 = lean_ctor_get(x_349, 1); -lean_inc(x_380); -x_381 = lean_ctor_get(x_349, 2); -lean_inc(x_381); -x_382 = lean_ctor_get(x_349, 3); -lean_inc(x_382); -if (lean_is_exclusive(x_349)) { - lean_ctor_release(x_349, 0); - lean_ctor_release(x_349, 1); - lean_ctor_release(x_349, 2); - lean_ctor_release(x_349, 3); - x_383 = x_349; -} else { - lean_dec_ref(x_349); - x_383 = lean_box(0); +else +{ +lean_object* x_16; lean_object* x_17; +lean_dec(x_1); +x_16 = lean_array_set(x_4, x_8, x_9); +lean_dec(x_8); +x_17 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_17, 0, x_16); +return x_17; } -x_384 = 1; -if (lean_is_scalar(x_383)) { - x_385 = lean_alloc_ctor(1, 4, 1); -} else { - x_385 = x_383; } -lean_ctor_set(x_385, 0, x_379); -lean_ctor_set(x_385, 1, x_380); -lean_ctor_set(x_385, 2, x_381); -lean_ctor_set(x_385, 3, x_382); -lean_ctor_set_uint8(x_385, sizeof(void*)*4, x_384); -if (lean_is_scalar(x_378)) { - x_386 = lean_alloc_ctor(1, 4, 1); -} else { - x_386 = x_378; } -lean_ctor_set(x_386, 0, x_377); -lean_ctor_set(x_386, 1, x_343); -lean_ctor_set(x_386, 2, x_344); -lean_ctor_set(x_386, 3, x_345); -lean_ctor_set_uint8(x_386, sizeof(void*)*4, x_384); -x_387 = 0; -x_388 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_388, 0, x_385); -lean_ctor_set(x_388, 1, x_375); -lean_ctor_set(x_388, 2, x_376); -lean_ctor_set(x_388, 3, x_386); -lean_ctor_set_uint8(x_388, sizeof(void*)*4, x_387); -return x_388; +case 1: +{ +uint8_t x_18; +x_18 = !lean_is_exclusive(x_1); +if (x_18 == 0) +{ +lean_object* x_19; uint8_t x_20; +x_19 = lean_ctor_get(x_1, 0); +lean_dec(x_19); +x_20 = !lean_is_exclusive(x_10); +if (x_20 == 0) +{ +lean_object* x_21; lean_object* x_22; size_t x_23; lean_object* x_24; lean_object* x_25; +x_21 = lean_ctor_get(x_10, 0); +x_22 = lean_array_set(x_4, x_8, x_9); +x_23 = lean_usize_shift_right(x_2, x_5); +x_24 = l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_getParentsAndReset___spec__2(x_21, x_23, x_3); +lean_inc(x_24); +x_25 = l_Lean_PersistentHashMap_isUnaryNode___rarg(x_24); +if (lean_obj_tag(x_25) == 0) +{ +lean_object* x_26; +lean_ctor_set(x_10, 0, x_24); +x_26 = lean_array_set(x_22, x_8, x_10); +lean_dec(x_8); +lean_ctor_set(x_1, 0, x_26); +return x_1; } else { -lean_object* x_389; -x_389 = lean_ctor_get(x_347, 3); -lean_inc(x_389); -if (lean_obj_tag(x_389) == 0) +lean_object* x_27; uint8_t x_28; +lean_dec(x_24); +lean_free_object(x_10); +x_27 = lean_ctor_get(x_25, 0); +lean_inc(x_27); +lean_dec(x_25); +x_28 = !lean_is_exclusive(x_27); +if (x_28 == 0) { -lean_object* x_390; uint8_t x_391; lean_object* x_392; -if (lean_is_exclusive(x_349)) { - lean_ctor_release(x_349, 0); - lean_ctor_release(x_349, 1); - lean_ctor_release(x_349, 2); - lean_ctor_release(x_349, 3); - x_390 = x_349; -} else { - lean_dec_ref(x_349); - x_390 = lean_box(0); +lean_object* x_29; +x_29 = lean_array_set(x_22, x_8, x_27); +lean_dec(x_8); +lean_ctor_set(x_1, 0, x_29); +return x_1; +} +else +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_30 = lean_ctor_get(x_27, 0); +x_31 = lean_ctor_get(x_27, 1); +lean_inc(x_31); +lean_inc(x_30); +lean_dec(x_27); +x_32 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_32, 0, x_30); +lean_ctor_set(x_32, 1, x_31); +x_33 = lean_array_set(x_22, x_8, x_32); +lean_dec(x_8); +lean_ctor_set(x_1, 0, x_33); +return x_1; } -x_391 = 1; -if (lean_is_scalar(x_390)) { - x_392 = lean_alloc_ctor(1, 4, 1); -} else { - x_392 = x_390; } -lean_ctor_set(x_392, 0, x_347); -lean_ctor_set(x_392, 1, x_343); -lean_ctor_set(x_392, 2, x_344); -lean_ctor_set(x_392, 3, x_345); -lean_ctor_set_uint8(x_392, sizeof(void*)*4, x_391); -return x_392; } else { -uint8_t x_393; -x_393 = lean_ctor_get_uint8(x_389, sizeof(void*)*4); -if (x_393 == 0) +lean_object* x_34; lean_object* x_35; size_t x_36; lean_object* x_37; lean_object* x_38; +x_34 = lean_ctor_get(x_10, 0); +lean_inc(x_34); +lean_dec(x_10); +x_35 = lean_array_set(x_4, x_8, x_9); +x_36 = lean_usize_shift_right(x_2, x_5); +x_37 = l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_getParentsAndReset___spec__2(x_34, x_36, x_3); +lean_inc(x_37); +x_38 = l_Lean_PersistentHashMap_isUnaryNode___rarg(x_37); +if (lean_obj_tag(x_38) == 0) { -lean_object* x_394; lean_object* x_395; lean_object* x_396; lean_object* x_397; lean_object* x_398; lean_object* x_399; lean_object* x_400; lean_object* x_401; uint8_t x_402; lean_object* x_403; lean_object* x_404; lean_object* x_405; uint8_t x_406; lean_object* x_407; -x_394 = lean_ctor_get(x_347, 1); -lean_inc(x_394); -x_395 = lean_ctor_get(x_347, 2); -lean_inc(x_395); -if (lean_is_exclusive(x_347)) { - lean_ctor_release(x_347, 0); - lean_ctor_release(x_347, 1); - lean_ctor_release(x_347, 2); - lean_ctor_release(x_347, 3); - x_396 = x_347; -} else { - lean_dec_ref(x_347); - x_396 = lean_box(0); -} -x_397 = lean_ctor_get(x_389, 0); -lean_inc(x_397); -x_398 = lean_ctor_get(x_389, 1); -lean_inc(x_398); -x_399 = lean_ctor_get(x_389, 2); -lean_inc(x_399); -x_400 = lean_ctor_get(x_389, 3); -lean_inc(x_400); -if (lean_is_exclusive(x_389)) { - lean_ctor_release(x_389, 0); - lean_ctor_release(x_389, 1); - lean_ctor_release(x_389, 2); - lean_ctor_release(x_389, 3); - x_401 = x_389; -} else { - lean_dec_ref(x_389); - x_401 = lean_box(0); +lean_object* x_39; lean_object* x_40; +x_39 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_39, 0, x_37); +x_40 = lean_array_set(x_35, x_8, x_39); +lean_dec(x_8); +lean_ctor_set(x_1, 0, x_40); +return x_1; } -x_402 = 1; -lean_inc(x_349); -if (lean_is_scalar(x_401)) { - x_403 = lean_alloc_ctor(1, 4, 1); +else +{ +lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; +lean_dec(x_37); +x_41 = lean_ctor_get(x_38, 0); +lean_inc(x_41); +lean_dec(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_43); +if (lean_is_exclusive(x_41)) { + lean_ctor_release(x_41, 0); + lean_ctor_release(x_41, 1); + x_44 = x_41; } else { - x_403 = x_401; + lean_dec_ref(x_41); + x_44 = lean_box(0); } -lean_ctor_set(x_403, 0, x_349); -lean_ctor_set(x_403, 1, x_394); -lean_ctor_set(x_403, 2, x_395); -lean_ctor_set(x_403, 3, x_397); -if (lean_is_exclusive(x_349)) { - lean_ctor_release(x_349, 0); - lean_ctor_release(x_349, 1); - lean_ctor_release(x_349, 2); - lean_ctor_release(x_349, 3); - x_404 = x_349; +if (lean_is_scalar(x_44)) { + x_45 = lean_alloc_ctor(0, 2, 0); } else { - lean_dec_ref(x_349); - x_404 = lean_box(0); + x_45 = x_44; } -lean_ctor_set_uint8(x_403, sizeof(void*)*4, x_402); -if (lean_is_scalar(x_404)) { - x_405 = lean_alloc_ctor(1, 4, 1); -} else { - x_405 = x_404; +lean_ctor_set(x_45, 0, x_42); +lean_ctor_set(x_45, 1, x_43); +x_46 = lean_array_set(x_35, x_8, x_45); +lean_dec(x_8); +lean_ctor_set(x_1, 0, x_46); +return x_1; } -lean_ctor_set(x_405, 0, x_400); -lean_ctor_set(x_405, 1, x_343); -lean_ctor_set(x_405, 2, x_344); -lean_ctor_set(x_405, 3, x_345); -lean_ctor_set_uint8(x_405, sizeof(void*)*4, x_402); -x_406 = 0; -if (lean_is_scalar(x_396)) { - x_407 = lean_alloc_ctor(1, 4, 1); -} else { - x_407 = x_396; } -lean_ctor_set(x_407, 0, x_403); -lean_ctor_set(x_407, 1, x_398); -lean_ctor_set(x_407, 2, x_399); -lean_ctor_set(x_407, 3, x_405); -lean_ctor_set_uint8(x_407, sizeof(void*)*4, x_406); -return x_407; } else { -lean_object* x_408; lean_object* x_409; lean_object* x_410; lean_object* x_411; lean_object* x_412; lean_object* x_413; lean_object* x_414; lean_object* x_415; lean_object* x_416; lean_object* x_417; uint8_t x_418; lean_object* x_419; -x_408 = lean_ctor_get(x_347, 1); -lean_inc(x_408); -x_409 = lean_ctor_get(x_347, 2); -lean_inc(x_409); -if (lean_is_exclusive(x_347)) { - lean_ctor_release(x_347, 0); - lean_ctor_release(x_347, 1); - lean_ctor_release(x_347, 2); - lean_ctor_release(x_347, 3); - x_410 = x_347; +lean_object* x_47; lean_object* x_48; lean_object* x_49; size_t x_50; lean_object* x_51; lean_object* x_52; +lean_dec(x_1); +x_47 = lean_ctor_get(x_10, 0); +lean_inc(x_47); +if (lean_is_exclusive(x_10)) { + lean_ctor_release(x_10, 0); + x_48 = x_10; } else { - lean_dec_ref(x_347); - x_410 = lean_box(0); + lean_dec_ref(x_10); + x_48 = lean_box(0); } -x_411 = lean_ctor_get(x_349, 0); -lean_inc(x_411); -x_412 = lean_ctor_get(x_349, 1); -lean_inc(x_412); -x_413 = lean_ctor_get(x_349, 2); -lean_inc(x_413); -x_414 = lean_ctor_get(x_349, 3); -lean_inc(x_414); -if (lean_is_exclusive(x_349)) { - lean_ctor_release(x_349, 0); - lean_ctor_release(x_349, 1); - lean_ctor_release(x_349, 2); - lean_ctor_release(x_349, 3); - x_415 = x_349; +x_49 = lean_array_set(x_4, x_8, x_9); +x_50 = lean_usize_shift_right(x_2, x_5); +x_51 = l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_getParentsAndReset___spec__2(x_47, x_50, x_3); +lean_inc(x_51); +x_52 = l_Lean_PersistentHashMap_isUnaryNode___rarg(x_51); +if (lean_obj_tag(x_52) == 0) +{ +lean_object* x_53; lean_object* x_54; lean_object* x_55; +if (lean_is_scalar(x_48)) { + x_53 = lean_alloc_ctor(1, 1, 0); } else { - lean_dec_ref(x_349); - x_415 = lean_box(0); + x_53 = x_48; } -if (lean_is_scalar(x_415)) { - x_416 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_53, 0, x_51); +x_54 = lean_array_set(x_49, x_8, x_53); +lean_dec(x_8); +x_55 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_55, 0, x_54); +return x_55; +} +else +{ +lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; +lean_dec(x_51); +lean_dec(x_48); +x_56 = lean_ctor_get(x_52, 0); +lean_inc(x_56); +lean_dec(x_52); +x_57 = lean_ctor_get(x_56, 0); +lean_inc(x_57); +x_58 = lean_ctor_get(x_56, 1); +lean_inc(x_58); +if (lean_is_exclusive(x_56)) { + lean_ctor_release(x_56, 0); + lean_ctor_release(x_56, 1); + x_59 = x_56; } else { - x_416 = x_415; + lean_dec_ref(x_56); + x_59 = lean_box(0); } -lean_ctor_set(x_416, 0, x_411); -lean_ctor_set(x_416, 1, x_412); -lean_ctor_set(x_416, 2, x_413); -lean_ctor_set(x_416, 3, x_414); -lean_ctor_set_uint8(x_416, sizeof(void*)*4, x_393); -if (lean_is_scalar(x_410)) { - x_417 = lean_alloc_ctor(1, 4, 1); +if (lean_is_scalar(x_59)) { + x_60 = lean_alloc_ctor(0, 2, 0); } else { - x_417 = x_410; + x_60 = x_59; } -lean_ctor_set(x_417, 0, x_416); -lean_ctor_set(x_417, 1, x_408); -lean_ctor_set(x_417, 2, x_409); -lean_ctor_set(x_417, 3, x_389); -lean_ctor_set_uint8(x_417, sizeof(void*)*4, x_348); -x_418 = 1; -x_419 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_419, 0, x_417); -lean_ctor_set(x_419, 1, x_343); -lean_ctor_set(x_419, 2, x_344); -lean_ctor_set(x_419, 3, x_345); -lean_ctor_set_uint8(x_419, sizeof(void*)*4, x_418); -return x_419; +lean_ctor_set(x_60, 0, x_57); +lean_ctor_set(x_60, 1, x_58); +x_61 = lean_array_set(x_49, x_8, x_60); +lean_dec(x_8); +x_62 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_62, 0, x_61); +return x_62; } } } +default: +{ +lean_dec(x_8); +lean_dec(x_4); +return x_1; +} } } else { -uint8_t x_420; lean_object* x_421; -x_420 = 1; -x_421 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_421, 0, x_347); -lean_ctor_set(x_421, 1, x_343); -lean_ctor_set(x_421, 2, x_344); -lean_ctor_set(x_421, 3, x_345); -lean_ctor_set_uint8(x_421, sizeof(void*)*4, x_420); -return x_421; -} -} -case 1: +lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_63 = lean_ctor_get(x_1, 0); +lean_inc(x_63); +x_64 = lean_ctor_get(x_1, 1); +lean_inc(x_64); +x_65 = lean_unsigned_to_nat(0u); +x_66 = l_Array_indexOfAux___at_Lean_Meta_Grind_getParentsAndReset___spec__3(x_63, x_3, x_65); +if (lean_obj_tag(x_66) == 0) { -uint8_t x_422; lean_object* x_423; -lean_dec(x_344); -lean_dec(x_343); -x_422 = 1; -x_423 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_423, 0, x_342); -lean_ctor_set(x_423, 1, x_2); -lean_ctor_set(x_423, 2, x_3); -lean_ctor_set(x_423, 3, x_345); -lean_ctor_set_uint8(x_423, sizeof(void*)*4, x_422); -return x_423; +lean_dec(x_64); +lean_dec(x_63); +return x_1; } -default: -{ -lean_object* x_424; uint8_t x_425; -x_424 = l_Lean_RBNode_ins___at_Lean_Meta_Grind_registerParent___spec__2(x_345, x_2, x_3); -x_425 = lean_ctor_get_uint8(x_424, sizeof(void*)*4); -if (x_425 == 0) +else { -lean_object* x_426; -x_426 = lean_ctor_get(x_424, 0); -lean_inc(x_426); -if (lean_obj_tag(x_426) == 0) +uint8_t x_67; +x_67 = !lean_is_exclusive(x_1); +if (x_67 == 0) { -lean_object* x_427; -x_427 = lean_ctor_get(x_424, 3); -lean_inc(x_427); -if (lean_obj_tag(x_427) == 0) +lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; +x_68 = lean_ctor_get(x_1, 1); +lean_dec(x_68); +x_69 = lean_ctor_get(x_1, 0); +lean_dec(x_69); +x_70 = lean_ctor_get(x_66, 0); +lean_inc(x_70); +lean_dec(x_66); +lean_inc(x_70); +x_71 = l_Array_eraseIdx___rarg(x_63, x_70, lean_box(0)); +x_72 = l_Array_eraseIdx___rarg(x_64, x_70, lean_box(0)); +lean_ctor_set(x_1, 1, x_72); +lean_ctor_set(x_1, 0, x_71); +return x_1; +} +else { -lean_object* x_428; lean_object* x_429; lean_object* x_430; lean_object* x_431; uint8_t x_432; lean_object* x_433; -x_428 = lean_ctor_get(x_424, 1); -lean_inc(x_428); -x_429 = lean_ctor_get(x_424, 2); -lean_inc(x_429); -if (lean_is_exclusive(x_424)) { - lean_ctor_release(x_424, 0); - lean_ctor_release(x_424, 1); - lean_ctor_release(x_424, 2); - lean_ctor_release(x_424, 3); - x_430 = x_424; -} else { - lean_dec_ref(x_424); - x_430 = lean_box(0); +lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; +lean_dec(x_1); +x_73 = lean_ctor_get(x_66, 0); +lean_inc(x_73); +lean_dec(x_66); +lean_inc(x_73); +x_74 = l_Array_eraseIdx___rarg(x_63, x_73, lean_box(0)); +x_75 = l_Array_eraseIdx___rarg(x_64, x_73, lean_box(0)); +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; } -if (lean_is_scalar(x_430)) { - x_431 = lean_alloc_ctor(1, 4, 1); -} else { - x_431 = x_430; } -lean_ctor_set(x_431, 0, x_427); -lean_ctor_set(x_431, 1, x_428); -lean_ctor_set(x_431, 2, x_429); -lean_ctor_set(x_431, 3, x_427); -lean_ctor_set_uint8(x_431, sizeof(void*)*4, x_425); -x_432 = 1; -x_433 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_433, 0, x_342); -lean_ctor_set(x_433, 1, x_343); -lean_ctor_set(x_433, 2, x_344); -lean_ctor_set(x_433, 3, x_431); -lean_ctor_set_uint8(x_433, sizeof(void*)*4, x_432); -return x_433; } -else -{ -uint8_t x_434; -x_434 = lean_ctor_get_uint8(x_427, sizeof(void*)*4); -if (x_434 == 0) +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_erase___at_Lean_Meta_Grind_getParentsAndReset___spec__1(lean_object* x_1, lean_object* x_2) { +_start: { -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; uint8_t x_443; lean_object* x_444; lean_object* x_445; uint8_t x_446; lean_object* x_447; -x_435 = lean_ctor_get(x_424, 1); -lean_inc(x_435); -x_436 = lean_ctor_get(x_424, 2); -lean_inc(x_436); -if (lean_is_exclusive(x_424)) { - lean_ctor_release(x_424, 0); - lean_ctor_release(x_424, 1); - lean_ctor_release(x_424, 2); - lean_ctor_release(x_424, 3); - x_437 = x_424; -} else { - lean_dec_ref(x_424); - x_437 = lean_box(0); +uint64_t x_3; size_t x_4; lean_object* x_5; +x_3 = l_Lean_Meta_Grind_instHashableENodeKey_unsafe__1(x_2); +x_4 = lean_uint64_to_usize(x_3); +x_5 = l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_getParentsAndReset___spec__2(x_1, x_4, x_2); +return x_5; } -x_438 = lean_ctor_get(x_427, 0); -lean_inc(x_438); -x_439 = lean_ctor_get(x_427, 1); -lean_inc(x_439); -x_440 = lean_ctor_get(x_427, 2); -lean_inc(x_440); -x_441 = lean_ctor_get(x_427, 3); -lean_inc(x_441); -if (lean_is_exclusive(x_427)) { - lean_ctor_release(x_427, 0); - lean_ctor_release(x_427, 1); - lean_ctor_release(x_427, 2); - lean_ctor_release(x_427, 3); - x_442 = x_427; -} else { - lean_dec_ref(x_427); - x_442 = lean_box(0); } -x_443 = 1; -if (lean_is_scalar(x_442)) { - x_444 = lean_alloc_ctor(1, 4, 1); -} else { - x_444 = x_442; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getParentsAndReset(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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; uint8_t x_17; +x_11 = l_Lean_Meta_Grind_getParents(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +x_13 = lean_ctor_get(x_11, 1); +lean_inc(x_13); +lean_dec(x_11); +x_14 = lean_st_ref_take(x_2, x_13); +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +x_16 = lean_ctor_get(x_14, 1); +lean_inc(x_16); +lean_dec(x_14); +x_17 = !lean_is_exclusive(x_15); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_18 = lean_ctor_get(x_15, 2); +x_19 = l_Lean_PersistentHashMap_erase___at_Lean_Meta_Grind_getParentsAndReset___spec__1(x_18, x_1); +lean_ctor_set(x_15, 2, x_19); +x_20 = lean_st_ref_set(x_2, x_15, x_16); +x_21 = !lean_is_exclusive(x_20); +if (x_21 == 0) +{ +lean_object* x_22; +x_22 = lean_ctor_get(x_20, 0); +lean_dec(x_22); +lean_ctor_set(x_20, 0, x_12); +return x_20; } -lean_ctor_set(x_444, 0, x_342); -lean_ctor_set(x_444, 1, x_343); -lean_ctor_set(x_444, 2, x_344); -lean_ctor_set(x_444, 3, x_426); -lean_ctor_set_uint8(x_444, sizeof(void*)*4, x_443); -if (lean_is_scalar(x_437)) { - x_445 = lean_alloc_ctor(1, 4, 1); -} else { - x_445 = x_437; +else +{ +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_12); +lean_ctor_set(x_24, 1, x_23); +return x_24; } -lean_ctor_set(x_445, 0, x_438); -lean_ctor_set(x_445, 1, x_439); -lean_ctor_set(x_445, 2, x_440); -lean_ctor_set(x_445, 3, x_441); -lean_ctor_set_uint8(x_445, sizeof(void*)*4, x_443); -x_446 = 0; -x_447 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_447, 0, x_444); -lean_ctor_set(x_447, 1, x_435); -lean_ctor_set(x_447, 2, x_436); -lean_ctor_set(x_447, 3, x_445); -lean_ctor_set_uint8(x_447, sizeof(void*)*4, x_446); -return x_447; } else { -lean_object* x_448; uint8_t x_449; lean_object* x_450; -if (lean_is_exclusive(x_427)) { - lean_ctor_release(x_427, 0); - lean_ctor_release(x_427, 1); - lean_ctor_release(x_427, 2); - lean_ctor_release(x_427, 3); - x_448 = x_427; +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; 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_25 = lean_ctor_get(x_15, 0); +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); +x_30 = lean_ctor_get(x_15, 5); +x_31 = lean_ctor_get_uint8(x_15, sizeof(void*)*20); +x_32 = lean_ctor_get(x_15, 6); +x_33 = lean_ctor_get(x_15, 7); +x_34 = lean_ctor_get(x_15, 8); +x_35 = lean_ctor_get(x_15, 9); +x_36 = lean_ctor_get(x_15, 10); +x_37 = lean_ctor_get(x_15, 11); +x_38 = lean_ctor_get(x_15, 12); +x_39 = lean_ctor_get(x_15, 13); +x_40 = lean_ctor_get(x_15, 14); +x_41 = lean_ctor_get(x_15, 15); +x_42 = lean_ctor_get(x_15, 16); +x_43 = lean_ctor_get(x_15, 17); +x_44 = lean_ctor_get(x_15, 18); +x_45 = lean_ctor_get(x_15, 19); +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_inc(x_34); +lean_inc(x_33); +lean_inc(x_32); +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_dec(x_15); +x_46 = l_Lean_PersistentHashMap_erase___at_Lean_Meta_Grind_getParentsAndReset___spec__1(x_27, x_1); +x_47 = lean_alloc_ctor(0, 20, 1); +lean_ctor_set(x_47, 0, x_25); +lean_ctor_set(x_47, 1, x_26); +lean_ctor_set(x_47, 2, x_46); +lean_ctor_set(x_47, 3, x_28); +lean_ctor_set(x_47, 4, x_29); +lean_ctor_set(x_47, 5, x_30); +lean_ctor_set(x_47, 6, x_32); +lean_ctor_set(x_47, 7, x_33); +lean_ctor_set(x_47, 8, x_34); +lean_ctor_set(x_47, 9, x_35); +lean_ctor_set(x_47, 10, x_36); +lean_ctor_set(x_47, 11, x_37); +lean_ctor_set(x_47, 12, x_38); +lean_ctor_set(x_47, 13, x_39); +lean_ctor_set(x_47, 14, x_40); +lean_ctor_set(x_47, 15, x_41); +lean_ctor_set(x_47, 16, x_42); +lean_ctor_set(x_47, 17, x_43); +lean_ctor_set(x_47, 18, x_44); +lean_ctor_set(x_47, 19, x_45); +lean_ctor_set_uint8(x_47, sizeof(void*)*20, x_31); +x_48 = lean_st_ref_set(x_2, x_47, x_16); +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_427); - x_448 = lean_box(0); + lean_dec_ref(x_48); + x_50 = lean_box(0); } -x_449 = 1; -if (lean_is_scalar(x_448)) { - x_450 = lean_alloc_ctor(1, 4, 1); +if (lean_is_scalar(x_50)) { + x_51 = lean_alloc_ctor(0, 2, 0); } else { - x_450 = x_448; + x_51 = x_50; } -lean_ctor_set(x_450, 0, x_342); -lean_ctor_set(x_450, 1, x_343); -lean_ctor_set(x_450, 2, x_344); -lean_ctor_set(x_450, 3, x_424); -lean_ctor_set_uint8(x_450, sizeof(void*)*4, x_449); -return x_450; +lean_ctor_set(x_51, 0, x_12); +lean_ctor_set(x_51, 1, x_49); +return x_51; } } } -else +LEAN_EXPORT lean_object* l_Array_indexOfAux___at_Lean_Meta_Grind_getParentsAndReset___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -uint8_t x_451; -x_451 = lean_ctor_get_uint8(x_426, sizeof(void*)*4); -if (x_451 == 0) +lean_object* x_4; +x_4 = l_Array_indexOfAux___at_Lean_Meta_Grind_getParentsAndReset___spec__3(x_1, x_2, x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_getParentsAndReset___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -lean_object* x_452; lean_object* x_453; lean_object* x_454; lean_object* x_455; lean_object* x_456; lean_object* x_457; lean_object* x_458; lean_object* x_459; lean_object* x_460; uint8_t x_461; lean_object* x_462; lean_object* x_463; uint8_t x_464; lean_object* x_465; -x_452 = lean_ctor_get(x_424, 1); -lean_inc(x_452); -x_453 = lean_ctor_get(x_424, 2); -lean_inc(x_453); -x_454 = lean_ctor_get(x_424, 3); -lean_inc(x_454); -if (lean_is_exclusive(x_424)) { - lean_ctor_release(x_424, 0); - lean_ctor_release(x_424, 1); - lean_ctor_release(x_424, 2); - lean_ctor_release(x_424, 3); - x_455 = x_424; -} else { - lean_dec_ref(x_424); - x_455 = lean_box(0); +size_t x_4; lean_object* x_5; +x_4 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_5 = l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_getParentsAndReset___spec__2(x_1, x_4, x_3); +lean_dec(x_3); +return x_5; } -x_456 = lean_ctor_get(x_426, 0); -lean_inc(x_456); -x_457 = lean_ctor_get(x_426, 1); -lean_inc(x_457); -x_458 = lean_ctor_get(x_426, 2); -lean_inc(x_458); -x_459 = lean_ctor_get(x_426, 3); -lean_inc(x_459); -if (lean_is_exclusive(x_426)) { - lean_ctor_release(x_426, 0); - lean_ctor_release(x_426, 1); - lean_ctor_release(x_426, 2); - lean_ctor_release(x_426, 3); - x_460 = x_426; -} else { - lean_dec_ref(x_426); - x_460 = lean_box(0); } -x_461 = 1; -if (lean_is_scalar(x_460)) { - x_462 = lean_alloc_ctor(1, 4, 1); -} else { - x_462 = x_460; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_erase___at_Lean_Meta_Grind_getParentsAndReset___spec__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_PersistentHashMap_erase___at_Lean_Meta_Grind_getParentsAndReset___spec__1(x_1, x_2); +lean_dec(x_2); +return x_3; } -lean_ctor_set(x_462, 0, x_342); -lean_ctor_set(x_462, 1, x_343); -lean_ctor_set(x_462, 2, x_344); -lean_ctor_set(x_462, 3, x_456); -lean_ctor_set_uint8(x_462, sizeof(void*)*4, x_461); -if (lean_is_scalar(x_455)) { - x_463 = lean_alloc_ctor(1, 4, 1); -} else { - x_463 = x_455; } -lean_ctor_set(x_463, 0, x_459); -lean_ctor_set(x_463, 1, x_452); -lean_ctor_set(x_463, 2, x_453); -lean_ctor_set(x_463, 3, x_454); -lean_ctor_set_uint8(x_463, sizeof(void*)*4, x_461); -x_464 = 0; -x_465 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_465, 0, x_462); -lean_ctor_set(x_465, 1, x_457); -lean_ctor_set(x_465, 2, x_458); -lean_ctor_set(x_465, 3, x_463); -lean_ctor_set_uint8(x_465, sizeof(void*)*4, x_464); -return x_465; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getParentsAndReset___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_Meta_Grind_getParentsAndReset(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_11; } -else +} +LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at_Lean_Meta_Grind_copyParentsTo___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) { +_start: { -lean_object* x_466; -x_466 = lean_ctor_get(x_424, 3); -lean_inc(x_466); -if (lean_obj_tag(x_466) == 0) +if (lean_obj_tag(x_1) == 0) { -lean_object* x_467; uint8_t x_468; lean_object* x_469; -if (lean_is_exclusive(x_426)) { - lean_ctor_release(x_426, 0); - lean_ctor_release(x_426, 1); - lean_ctor_release(x_426, 2); - lean_ctor_release(x_426, 3); - x_467 = x_426; -} else { - lean_dec_ref(x_426); - x_467 = lean_box(0); +lean_object* x_12; lean_object* x_13; +x_12 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_12, 0, x_2); +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_11); +return x_13; +} +else +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_14 = lean_ctor_get(x_1, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_1, 1); +lean_inc(x_15); +x_16 = lean_ctor_get(x_1, 3); +lean_inc(x_16); +lean_dec(x_1); +x_17 = l_Lean_RBNode_forIn_visit___at_Lean_Meta_Grind_copyParentsTo___spec__1(x_14, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +x_19 = lean_ctor_get(x_17, 1); +lean_inc(x_19); +lean_dec(x_17); +x_20 = lean_ctor_get(x_18, 0); +lean_inc(x_20); +lean_dec(x_18); +x_21 = lean_box(0); +x_22 = l_Lean_RBNode_insert___at_Lean_Meta_Grind_registerParent___spec__1(x_20, x_15, x_21); +x_1 = x_16; +x_2 = x_22; +x_11 = x_19; +goto _start; } -x_468 = 1; -if (lean_is_scalar(x_467)) { - x_469 = lean_alloc_ctor(1, 4, 1); -} else { - x_469 = x_467; } -lean_ctor_set(x_469, 0, x_342); -lean_ctor_set(x_469, 1, x_343); -lean_ctor_set(x_469, 2, x_344); -lean_ctor_set(x_469, 3, x_424); -lean_ctor_set_uint8(x_469, sizeof(void*)*4, x_468); -return x_469; } -else +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_copyParentsTo(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, 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_470; -x_470 = lean_ctor_get_uint8(x_466, sizeof(void*)*4); -if (x_470 == 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_60; +x_12 = lean_st_ref_get(x_3, x_11); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_12, 1); +lean_inc(x_14); +lean_dec(x_12); +x_15 = lean_ctor_get(x_13, 2); +lean_inc(x_15); +lean_dec(x_13); +x_60 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_registerParent___spec__7(x_15, x_2); +if (lean_obj_tag(x_60) == 0) { -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; uint8_t x_479; lean_object* x_480; lean_object* x_481; lean_object* x_482; uint8_t x_483; lean_object* x_484; -x_471 = lean_ctor_get(x_424, 1); -lean_inc(x_471); -x_472 = lean_ctor_get(x_424, 2); -lean_inc(x_472); -if (lean_is_exclusive(x_424)) { - lean_ctor_release(x_424, 0); - lean_ctor_release(x_424, 1); - lean_ctor_release(x_424, 2); - lean_ctor_release(x_424, 3); - x_473 = x_424; -} else { - lean_dec_ref(x_424); - x_473 = lean_box(0); -} -x_474 = lean_ctor_get(x_466, 0); -lean_inc(x_474); -x_475 = lean_ctor_get(x_466, 1); -lean_inc(x_475); -x_476 = lean_ctor_get(x_466, 2); -lean_inc(x_476); -x_477 = lean_ctor_get(x_466, 3); -lean_inc(x_477); -if (lean_is_exclusive(x_466)) { - lean_ctor_release(x_466, 0); - lean_ctor_release(x_466, 1); - lean_ctor_release(x_466, 2); - lean_ctor_release(x_466, 3); - x_478 = x_466; -} else { - lean_dec_ref(x_466); - x_478 = lean_box(0); -} -x_479 = 1; -lean_inc(x_426); -if (lean_is_scalar(x_478)) { - x_480 = lean_alloc_ctor(1, 4, 1); -} else { - x_480 = x_478; +lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; +x_61 = lean_box(0); +x_62 = l_Lean_RBNode_forIn_visit___at_Lean_Meta_Grind_copyParentsTo___spec__1(x_1, x_61, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_14); +x_63 = lean_ctor_get(x_62, 0); +lean_inc(x_63); +x_64 = lean_ctor_get(x_62, 1); +lean_inc(x_64); +lean_dec(x_62); +x_65 = lean_ctor_get(x_63, 0); +lean_inc(x_65); +lean_dec(x_63); +x_16 = x_65; +x_17 = x_64; +goto block_59; } -lean_ctor_set(x_480, 0, x_342); -lean_ctor_set(x_480, 1, x_343); -lean_ctor_set(x_480, 2, x_344); -lean_ctor_set(x_480, 3, x_426); -if (lean_is_exclusive(x_426)) { - lean_ctor_release(x_426, 0); - lean_ctor_release(x_426, 1); - lean_ctor_release(x_426, 2); - lean_ctor_release(x_426, 3); - x_481 = x_426; -} else { - lean_dec_ref(x_426); - x_481 = lean_box(0); +else +{ +lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; +x_66 = lean_ctor_get(x_60, 0); +lean_inc(x_66); +lean_dec(x_60); +x_67 = l_Lean_RBNode_forIn_visit___at_Lean_Meta_Grind_copyParentsTo___spec__1(x_1, x_66, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_14); +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_16 = x_70; +x_17 = x_69; +goto block_59; } -lean_ctor_set_uint8(x_480, sizeof(void*)*4, x_479); -if (lean_is_scalar(x_481)) { - x_482 = lean_alloc_ctor(1, 4, 1); -} else { - x_482 = x_481; +block_59: +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_18 = lean_st_ref_take(x_3, x_17); +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_18, 1); +lean_inc(x_20); +lean_dec(x_18); +x_21 = !lean_is_exclusive(x_19); +if (x_21 == 0) +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; +x_22 = lean_ctor_get(x_19, 2); +x_23 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_registerParent___spec__3(x_22, x_2, x_16); +lean_ctor_set(x_19, 2, x_23); +x_24 = lean_st_ref_set(x_3, x_19, x_20); +x_25 = !lean_is_exclusive(x_24); +if (x_25 == 0) +{ +lean_object* x_26; lean_object* x_27; +x_26 = lean_ctor_get(x_24, 0); +lean_dec(x_26); +x_27 = lean_box(0); +lean_ctor_set(x_24, 0, x_27); +return x_24; } -lean_ctor_set(x_482, 0, x_474); -lean_ctor_set(x_482, 1, x_475); -lean_ctor_set(x_482, 2, x_476); -lean_ctor_set(x_482, 3, x_477); -lean_ctor_set_uint8(x_482, sizeof(void*)*4, x_479); -x_483 = 0; -if (lean_is_scalar(x_473)) { - x_484 = lean_alloc_ctor(1, 4, 1); -} else { - x_484 = x_473; +else +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_28 = lean_ctor_get(x_24, 1); +lean_inc(x_28); +lean_dec(x_24); +x_29 = lean_box(0); +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_28); +return x_30; } -lean_ctor_set(x_484, 0, x_480); -lean_ctor_set(x_484, 1, x_471); -lean_ctor_set(x_484, 2, x_472); -lean_ctor_set(x_484, 3, x_482); -lean_ctor_set_uint8(x_484, sizeof(void*)*4, x_483); -return x_484; } else { -lean_object* x_485; lean_object* x_486; lean_object* x_487; lean_object* x_488; lean_object* x_489; lean_object* x_490; lean_object* x_491; lean_object* x_492; lean_object* x_493; lean_object* x_494; uint8_t x_495; lean_object* x_496; -x_485 = lean_ctor_get(x_424, 1); -lean_inc(x_485); -x_486 = lean_ctor_get(x_424, 2); -lean_inc(x_486); -if (lean_is_exclusive(x_424)) { - lean_ctor_release(x_424, 0); - lean_ctor_release(x_424, 1); - lean_ctor_release(x_424, 2); - lean_ctor_release(x_424, 3); - x_487 = x_424; -} else { - lean_dec_ref(x_424); - x_487 = lean_box(0); -} -x_488 = lean_ctor_get(x_426, 0); -lean_inc(x_488); -x_489 = lean_ctor_get(x_426, 1); -lean_inc(x_489); -x_490 = lean_ctor_get(x_426, 2); -lean_inc(x_490); -x_491 = lean_ctor_get(x_426, 3); -lean_inc(x_491); -if (lean_is_exclusive(x_426)) { - lean_ctor_release(x_426, 0); - lean_ctor_release(x_426, 1); - lean_ctor_release(x_426, 2); - lean_ctor_release(x_426, 3); - x_492 = x_426; +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; +x_31 = lean_ctor_get(x_19, 0); +x_32 = lean_ctor_get(x_19, 1); +x_33 = lean_ctor_get(x_19, 2); +x_34 = lean_ctor_get(x_19, 3); +x_35 = lean_ctor_get(x_19, 4); +x_36 = lean_ctor_get(x_19, 5); +x_37 = lean_ctor_get_uint8(x_19, sizeof(void*)*20); +x_38 = lean_ctor_get(x_19, 6); +x_39 = lean_ctor_get(x_19, 7); +x_40 = lean_ctor_get(x_19, 8); +x_41 = lean_ctor_get(x_19, 9); +x_42 = lean_ctor_get(x_19, 10); +x_43 = lean_ctor_get(x_19, 11); +x_44 = lean_ctor_get(x_19, 12); +x_45 = lean_ctor_get(x_19, 13); +x_46 = lean_ctor_get(x_19, 14); +x_47 = lean_ctor_get(x_19, 15); +x_48 = lean_ctor_get(x_19, 16); +x_49 = lean_ctor_get(x_19, 17); +x_50 = lean_ctor_get(x_19, 18); +x_51 = lean_ctor_get(x_19, 19); +lean_inc(x_51); +lean_inc(x_50); +lean_inc(x_49); +lean_inc(x_48); +lean_inc(x_47); +lean_inc(x_46); +lean_inc(x_45); +lean_inc(x_44); +lean_inc(x_43); +lean_inc(x_42); +lean_inc(x_41); +lean_inc(x_40); +lean_inc(x_39); +lean_inc(x_38); +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_19); +x_52 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_registerParent___spec__3(x_33, x_2, x_16); +x_53 = lean_alloc_ctor(0, 20, 1); +lean_ctor_set(x_53, 0, x_31); +lean_ctor_set(x_53, 1, x_32); +lean_ctor_set(x_53, 2, x_52); +lean_ctor_set(x_53, 3, x_34); +lean_ctor_set(x_53, 4, x_35); +lean_ctor_set(x_53, 5, x_36); +lean_ctor_set(x_53, 6, x_38); +lean_ctor_set(x_53, 7, x_39); +lean_ctor_set(x_53, 8, x_40); +lean_ctor_set(x_53, 9, x_41); +lean_ctor_set(x_53, 10, x_42); +lean_ctor_set(x_53, 11, x_43); +lean_ctor_set(x_53, 12, x_44); +lean_ctor_set(x_53, 13, x_45); +lean_ctor_set(x_53, 14, x_46); +lean_ctor_set(x_53, 15, x_47); +lean_ctor_set(x_53, 16, x_48); +lean_ctor_set(x_53, 17, x_49); +lean_ctor_set(x_53, 18, x_50); +lean_ctor_set(x_53, 19, x_51); +lean_ctor_set_uint8(x_53, sizeof(void*)*20, x_37); +x_54 = lean_st_ref_set(x_3, x_53, x_20); +x_55 = lean_ctor_get(x_54, 1); +lean_inc(x_55); +if (lean_is_exclusive(x_54)) { + lean_ctor_release(x_54, 0); + lean_ctor_release(x_54, 1); + x_56 = x_54; } else { - lean_dec_ref(x_426); - x_492 = lean_box(0); + lean_dec_ref(x_54); + x_56 = lean_box(0); } -if (lean_is_scalar(x_492)) { - x_493 = lean_alloc_ctor(1, 4, 1); +x_57 = lean_box(0); +if (lean_is_scalar(x_56)) { + x_58 = lean_alloc_ctor(0, 2, 0); } else { - x_493 = x_492; + x_58 = x_56; } -lean_ctor_set(x_493, 0, x_488); -lean_ctor_set(x_493, 1, x_489); -lean_ctor_set(x_493, 2, x_490); -lean_ctor_set(x_493, 3, x_491); -lean_ctor_set_uint8(x_493, sizeof(void*)*4, x_470); -if (lean_is_scalar(x_487)) { - x_494 = lean_alloc_ctor(1, 4, 1); -} else { - x_494 = x_487; +lean_ctor_set(x_58, 0, x_57); +lean_ctor_set(x_58, 1, x_55); +return x_58; } -lean_ctor_set(x_494, 0, x_493); -lean_ctor_set(x_494, 1, x_485); -lean_ctor_set(x_494, 2, x_486); -lean_ctor_set(x_494, 3, x_466); -lean_ctor_set_uint8(x_494, sizeof(void*)*4, x_425); -x_495 = 1; -x_496 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_496, 0, x_342); -lean_ctor_set(x_496, 1, x_343); -lean_ctor_set(x_496, 2, x_344); -lean_ctor_set(x_496, 3, x_494); -lean_ctor_set_uint8(x_496, sizeof(void*)*4, x_495); -return x_496; } } } +LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at_Lean_Meta_Grind_copyParentsTo___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; +x_12 = l_Lean_RBNode_forIn_visit___at_Lean_Meta_Grind_copyParentsTo___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_12; } } -else +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_copyParentsTo___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, 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_497; lean_object* x_498; -x_497 = 1; -x_498 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_498, 0, x_342); -lean_ctor_set(x_498, 1, x_343); -lean_ctor_set(x_498, 2, x_344); -lean_ctor_set(x_498, 3, x_424); -lean_ctor_set_uint8(x_498, sizeof(void*)*4, x_497); -return x_498; -} -} +lean_object* x_12; +x_12 = l_Lean_Meta_Grind_copyParentsTo(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_12; } } +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_setENode_unsafe__1___rarg(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_ctor_get(x_1, 3); +lean_inc(x_2); +return x_2; } } +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_setENode_unsafe__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_setENode_unsafe__1___rarg___boxed), 1, 0); +return x_3; } } -LEAN_EXPORT lean_object* l_Lean_RBNode_insert___at_Lean_Meta_Grind_registerParent___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_setENode_unsafe__1___rarg___boxed(lean_object* x_1) { _start: { -uint8_t x_4; -x_4 = l_Lean_RBNode_isRed___rarg(x_1); -if (x_4 == 0) -{ -lean_object* x_5; -x_5 = l_Lean_RBNode_ins___at_Lean_Meta_Grind_registerParent___spec__2(x_1, x_2, x_3); -return x_5; +lean_object* x_2; +x_2 = l_Lean_Meta_Grind_setENode_unsafe__1___rarg(x_1); +lean_dec(x_1); +return x_2; } -else -{ -lean_object* x_6; lean_object* x_7; -x_6 = l_Lean_RBNode_ins___at_Lean_Meta_Grind_registerParent___spec__2(x_1, x_2, x_3); -x_7 = l_Lean_RBNode_setBlack___rarg(x_6); -return x_7; } +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_setENode_unsafe__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_Meta_Grind_setENode_unsafe__1(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_registerParent___spec__5(size_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_setENode___spec__3(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; @@ -13516,7 +18623,7 @@ 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_Grind_registerParent___spec__4(x_6, x_17, x_1, x_9, x_10); +x_20 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_setENode___spec__2(x_6, x_17, x_1, x_9, x_10); x_4 = lean_box(0); x_5 = x_19; x_6 = x_20; @@ -13524,7 +18631,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_registerParent___spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_setENode___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; @@ -13616,7 +18723,7 @@ return x_29; } } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_registerParent___spec__4(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_setENode___spec__2(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) @@ -13729,7 +18836,7 @@ lean_object* x_35; size_t x_36; size_t x_37; lean_object* x_38; lean_object* x_3 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_Grind_registerParent___spec__4(x_35, x_36, x_37, x_4, x_5); +x_38 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_setENode___spec__2(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); @@ -13744,7 +18851,7 @@ 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_Grind_registerParent___spec__4(x_40, x_41, x_42, x_4, x_5); +x_43 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_setENode___spec__2(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); @@ -13860,7 +18967,7 @@ if (lean_is_exclusive(x_57)) { } 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_Grind_registerParent___spec__4(x_71, x_73, x_74, x_4, x_5); +x_75 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_setENode___spec__2(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 { @@ -13897,7 +19004,7 @@ 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_Grind_registerParent___spec__6(x_1, x_83, x_4, x_5); +x_84 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_setENode___spec__4(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) @@ -13915,483 +19022,227 @@ 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_Grind_mkHCongrWithArity___spec__2___closed__3; -x_93 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_registerParent___spec__5(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_Grind_registerParent___spec__6(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_Grind_mkHCongrWithArity___spec__2___closed__3; -x_107 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_registerParent___spec__5(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_Grind_registerParent___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -uint64_t x_4; size_t x_5; size_t x_6; lean_object* x_7; -x_4 = l_Lean_Meta_Grind_instHashableENodeKey_unsafe__1(x_2); -x_5 = lean_uint64_to_usize(x_4); -x_6 = 1; -x_7 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_registerParent___spec__4(x_1, x_5, x_6, x_2, x_3); -return x_7; -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_registerParent___spec__9(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: -{ -lean_object* x_6; 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 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(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; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_registerParent___spec__8(lean_object* x_1, size_t x_2, lean_object* x_3) { -_start: -{ -if (lean_obj_tag(x_1) == 0) -{ -uint8_t x_4; -x_4 = !lean_is_exclusive(x_1); -if (x_4 == 0) -{ -lean_object* x_5; size_t x_6; size_t x_7; size_t x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; -x_5 = lean_ctor_get(x_1, 0); -x_6 = 5; -x_7 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__2___closed__2; -x_8 = lean_usize_land(x_2, x_7); -x_9 = lean_usize_to_nat(x_8); -x_10 = lean_box(2); -x_11 = lean_array_get(x_10, x_5, x_9); -lean_dec(x_9); -lean_dec(x_5); -switch (lean_obj_tag(x_11)) { -case 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 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_3, x_12); -lean_dec(x_12); -if (x_14 == 0) -{ -lean_object* x_15; -lean_dec(x_13); -lean_free_object(x_1); -x_15 = lean_box(0); -return x_15; -} -else -{ -lean_ctor_set_tag(x_1, 1); -lean_ctor_set(x_1, 0, x_13); -return x_1; -} -} -case 1: -{ -lean_object* x_16; size_t x_17; -lean_free_object(x_1); -x_16 = lean_ctor_get(x_11, 0); -lean_inc(x_16); -lean_dec(x_11); -x_17 = lean_usize_shift_right(x_2, x_6); -x_1 = x_16; -x_2 = x_17; -goto _start; +x_92 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__2___closed__3; +x_93 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_setENode___spec__3(x_3, x_90, x_91, lean_box(0), x_83, x_92); +lean_dec(x_91); +lean_dec(x_90); +return x_93; } -default: +else { -lean_object* x_19; -lean_free_object(x_1); -x_19 = lean_box(0); -return x_19; +return x_84; } } +else +{ +return x_84; +} } else { -lean_object* x_20; size_t x_21; size_t x_22; size_t x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_20 = lean_ctor_get(x_1, 0); -lean_inc(x_20); +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_21 = 5; -x_22 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__2___closed__2; -x_23 = lean_usize_land(x_2, x_22); -x_24 = lean_usize_to_nat(x_23); -x_25 = lean_box(2); -x_26 = lean_array_get(x_25, x_20, x_24); -lean_dec(x_24); -lean_dec(x_20); -switch (lean_obj_tag(x_26)) { -case 0: +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_Grind_setENode___spec__4(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_27; lean_object* x_28; uint8_t 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_Grind_isSameExpr_unsafe__1(x_3, x_27); -lean_dec(x_27); -if (x_29 == 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_30; -lean_dec(x_28); -x_30 = lean_box(0); -return x_30; +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_Grind_mkHCongrWithArity___spec__2___closed__3; +x_107 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_setENode___spec__3(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 { -lean_object* x_31; -x_31 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_31, 0, x_28); -return x_31; -} +return x_98; } -case 1: -{ -lean_object* x_32; size_t x_33; -x_32 = lean_ctor_get(x_26, 0); -lean_inc(x_32); -lean_dec(x_26); -x_33 = lean_usize_shift_right(x_2, x_21); -x_1 = x_32; -x_2 = x_33; -goto _start; } -default: +else { -lean_object* x_35; -x_35 = lean_box(0); -return x_35; -} -} +return x_98; } } -else -{ -lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_36 = lean_ctor_get(x_1, 0); -lean_inc(x_36); -x_37 = lean_ctor_get(x_1, 1); -lean_inc(x_37); -lean_dec(x_1); -x_38 = lean_unsigned_to_nat(0u); -x_39 = l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_registerParent___spec__9(x_36, x_37, lean_box(0), x_38, x_3); -lean_dec(x_37); -lean_dec(x_36); -return x_39; } } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_registerParent___spec__7(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_setENode___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -uint64_t x_3; size_t x_4; lean_object* x_5; -x_3 = l_Lean_Meta_Grind_instHashableENodeKey_unsafe__1(x_2); -x_4 = lean_uint64_to_usize(x_3); -x_5 = l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_registerParent___spec__8(x_1, x_4, x_2); -return x_5; +uint64_t x_4; size_t x_5; size_t x_6; lean_object* x_7; +x_4 = l_Lean_Meta_Grind_instHashableENodeKey_unsafe__1(x_2); +x_5 = lean_uint64_to_usize(x_4); +x_6 = 1; +x_7 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_setENode___spec__2(x_1, x_5, x_6, x_2, x_3); +return x_7; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_registerParent(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_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_Grind_setENode(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -lean_object* x_12; lean_object* x_13; -x_12 = l_Lean_Meta_Grind_getRoot_x3f(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_12 = lean_st_ref_take(x_3, x_11); x_13 = lean_ctor_get(x_12, 0); lean_inc(x_13); -if (lean_obj_tag(x_13) == 0) -{ -uint8_t x_14; -lean_dec(x_1); -x_14 = !lean_is_exclusive(x_12); -if (x_14 == 0) -{ -lean_object* x_15; lean_object* x_16; -x_15 = lean_ctor_get(x_12, 0); -lean_dec(x_15); -x_16 = lean_box(0); -lean_ctor_set(x_12, 0, x_16); -return x_12; -} -else -{ -lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_17 = lean_ctor_get(x_12, 1); -lean_inc(x_17); -lean_dec(x_12); -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_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_63; lean_object* x_64; -x_20 = lean_ctor_get(x_12, 1); -lean_inc(x_20); +x_14 = lean_ctor_get(x_12, 1); +lean_inc(x_14); lean_dec(x_12); -x_21 = lean_ctor_get(x_13, 0); -lean_inc(x_21); -lean_dec(x_13); -x_22 = lean_st_ref_get(x_3, x_20); -x_23 = lean_ctor_get(x_22, 0); -lean_inc(x_23); -x_24 = lean_ctor_get(x_22, 1); -lean_inc(x_24); -lean_dec(x_22); -x_25 = lean_ctor_get(x_23, 2); -lean_inc(x_25); -lean_dec(x_23); -x_63 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_registerParent___spec__7(x_25, x_21); -x_64 = lean_st_ref_take(x_3, x_24); -if (lean_obj_tag(x_63) == 0) -{ -lean_object* x_65; lean_object* x_66; lean_object* x_67; -x_65 = lean_ctor_get(x_64, 0); -lean_inc(x_65); -x_66 = lean_ctor_get(x_64, 1); -lean_inc(x_66); -lean_dec(x_64); -x_67 = lean_box(0); -x_26 = x_67; -x_27 = x_65; -x_28 = x_66; -goto block_62; -} -else -{ -lean_object* x_68; lean_object* x_69; lean_object* x_70; -x_68 = lean_ctor_get(x_63, 0); -lean_inc(x_68); -lean_dec(x_63); -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_26 = x_68; -x_27 = x_69; -x_28 = x_70; -goto block_62; -} -block_62: -{ -uint8_t x_29; -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; uint8_t x_35; -x_30 = lean_ctor_get(x_27, 2); -x_31 = lean_box(0); -x_32 = l_Lean_RBNode_insert___at_Lean_Meta_Grind_registerParent___spec__1(x_26, x_1, x_31); -x_33 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_registerParent___spec__3(x_30, x_21, x_32); -lean_ctor_set(x_27, 2, x_33); -x_34 = lean_st_ref_set(x_3, x_27, x_28); -x_35 = !lean_is_exclusive(x_34); -if (x_35 == 0) +x_15 = !lean_is_exclusive(x_13); +if (x_15 == 0) { -lean_object* x_36; -x_36 = lean_ctor_get(x_34, 0); -lean_dec(x_36); -lean_ctor_set(x_34, 0, x_31); -return x_34; -} -else +lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; +x_16 = lean_ctor_get(x_13, 1); +x_17 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_setENode___spec__1(x_16, x_1, x_2); +lean_ctor_set(x_13, 1, x_17); +x_18 = lean_st_ref_set(x_3, x_13, x_14); +x_19 = !lean_is_exclusive(x_18); +if (x_19 == 0) { -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_31); -lean_ctor_set(x_38, 1, x_37); -return x_38; -} +lean_object* x_20; lean_object* x_21; +x_20 = lean_ctor_get(x_18, 0); +lean_dec(x_20); +x_21 = lean_box(0); +lean_ctor_set(x_18, 0, x_21); +return x_18; } else { -lean_object* x_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; lean_object* x_61; -x_39 = lean_ctor_get(x_27, 0); -x_40 = lean_ctor_get(x_27, 1); -x_41 = lean_ctor_get(x_27, 2); -x_42 = lean_ctor_get(x_27, 3); -x_43 = lean_ctor_get(x_27, 4); -x_44 = lean_ctor_get(x_27, 5); -x_45 = lean_ctor_get_uint8(x_27, sizeof(void*)*14); -x_46 = lean_ctor_get(x_27, 6); -x_47 = lean_ctor_get(x_27, 7); -x_48 = lean_ctor_get(x_27, 8); -x_49 = lean_ctor_get(x_27, 9); -x_50 = lean_ctor_get(x_27, 10); -x_51 = lean_ctor_get(x_27, 11); -x_52 = lean_ctor_get(x_27, 12); -x_53 = lean_ctor_get(x_27, 13); -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_object* x_22; lean_object* x_23; lean_object* x_24; +x_22 = lean_ctor_get(x_18, 1); +lean_inc(x_22); +lean_dec(x_18); +x_23 = lean_box(0); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set(x_24, 1, x_22); +return x_24; +} +} +else +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; 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_25 = lean_ctor_get(x_13, 0); +x_26 = lean_ctor_get(x_13, 1); +x_27 = lean_ctor_get(x_13, 2); +x_28 = lean_ctor_get(x_13, 3); +x_29 = lean_ctor_get(x_13, 4); +x_30 = lean_ctor_get(x_13, 5); +x_31 = lean_ctor_get_uint8(x_13, sizeof(void*)*20); +x_32 = lean_ctor_get(x_13, 6); +x_33 = lean_ctor_get(x_13, 7); +x_34 = lean_ctor_get(x_13, 8); +x_35 = lean_ctor_get(x_13, 9); +x_36 = lean_ctor_get(x_13, 10); +x_37 = lean_ctor_get(x_13, 11); +x_38 = lean_ctor_get(x_13, 12); +x_39 = lean_ctor_get(x_13, 13); +x_40 = lean_ctor_get(x_13, 14); +x_41 = lean_ctor_get(x_13, 15); +x_42 = lean_ctor_get(x_13, 16); +x_43 = lean_ctor_get(x_13, 17); +x_44 = lean_ctor_get(x_13, 18); +x_45 = lean_ctor_get(x_13, 19); +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_dec(x_27); -x_54 = lean_box(0); -x_55 = l_Lean_RBNode_insert___at_Lean_Meta_Grind_registerParent___spec__1(x_26, x_1, x_54); -x_56 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_registerParent___spec__3(x_41, x_21, x_55); -x_57 = lean_alloc_ctor(0, 14, 1); -lean_ctor_set(x_57, 0, x_39); -lean_ctor_set(x_57, 1, x_40); -lean_ctor_set(x_57, 2, x_56); -lean_ctor_set(x_57, 3, x_42); -lean_ctor_set(x_57, 4, x_43); -lean_ctor_set(x_57, 5, x_44); -lean_ctor_set(x_57, 6, x_46); -lean_ctor_set(x_57, 7, x_47); -lean_ctor_set(x_57, 8, x_48); -lean_ctor_set(x_57, 9, x_49); -lean_ctor_set(x_57, 10, x_50); -lean_ctor_set(x_57, 11, x_51); -lean_ctor_set(x_57, 12, x_52); -lean_ctor_set(x_57, 13, x_53); -lean_ctor_set_uint8(x_57, sizeof(void*)*14, x_45); -x_58 = lean_st_ref_set(x_3, x_57, x_28); -x_59 = lean_ctor_get(x_58, 1); -lean_inc(x_59); -if (lean_is_exclusive(x_58)) { - lean_ctor_release(x_58, 0); - lean_ctor_release(x_58, 1); - x_60 = x_58; +lean_inc(x_38); +lean_inc(x_37); +lean_inc(x_36); +lean_inc(x_35); +lean_inc(x_34); +lean_inc(x_33); +lean_inc(x_32); +lean_inc(x_30); +lean_inc(x_29); +lean_inc(x_28); +lean_inc(x_27); +lean_inc(x_26); +lean_inc(x_25); +lean_dec(x_13); +x_46 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_setENode___spec__1(x_26, x_1, x_2); +x_47 = lean_alloc_ctor(0, 20, 1); +lean_ctor_set(x_47, 0, x_25); +lean_ctor_set(x_47, 1, x_46); +lean_ctor_set(x_47, 2, x_27); +lean_ctor_set(x_47, 3, x_28); +lean_ctor_set(x_47, 4, x_29); +lean_ctor_set(x_47, 5, x_30); +lean_ctor_set(x_47, 6, x_32); +lean_ctor_set(x_47, 7, x_33); +lean_ctor_set(x_47, 8, x_34); +lean_ctor_set(x_47, 9, x_35); +lean_ctor_set(x_47, 10, x_36); +lean_ctor_set(x_47, 11, x_37); +lean_ctor_set(x_47, 12, x_38); +lean_ctor_set(x_47, 13, x_39); +lean_ctor_set(x_47, 14, x_40); +lean_ctor_set(x_47, 15, x_41); +lean_ctor_set(x_47, 16, x_42); +lean_ctor_set(x_47, 17, x_43); +lean_ctor_set(x_47, 18, x_44); +lean_ctor_set(x_47, 19, x_45); +lean_ctor_set_uint8(x_47, sizeof(void*)*20, x_31); +x_48 = lean_st_ref_set(x_3, x_47, x_14); +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_58); - x_60 = lean_box(0); + lean_dec_ref(x_48); + x_50 = lean_box(0); } -if (lean_is_scalar(x_60)) { - x_61 = lean_alloc_ctor(0, 2, 0); +x_51 = lean_box(0); +if (lean_is_scalar(x_50)) { + x_52 = lean_alloc_ctor(0, 2, 0); } else { - x_61 = x_60; -} -lean_ctor_set(x_61, 0, x_54); -lean_ctor_set(x_61, 1, x_59); -return x_61; -} + 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_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_registerParent___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_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_setENode___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { size_t x_7; 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_Grind_registerParent___spec__5(x_7, x_2, x_3, x_4, x_5, x_6); +x_8 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_setENode___spec__3(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_Grind_registerParent___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_setENode___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { size_t x_6; size_t x_7; lean_object* x_8; @@ -14399,645 +19250,736 @@ 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_Grind_registerParent___spec__4(x_1, x_6, x_7, x_4, x_5); +x_8 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_setENode___spec__2(x_1, x_6, x_7, x_4, x_5); return x_8; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_registerParent___spec__9___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_Grind_setENode___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, 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_6; -x_6 = l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_registerParent___spec__9(x_1, x_2, x_3, x_4, x_5); +lean_object* x_12; +x_12 = l_Lean_Meta_Grind_setENode(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); lean_dec(x_5); -lean_dec(x_2); -lean_dec(x_1); -return x_6; +lean_dec(x_4); +lean_dec(x_3); +return x_12; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_registerParent___spec__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkENodeCore(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, lean_object* x_12, lean_object* x_13) { _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_Grind_registerParent___spec__8(x_1, x_4, x_3); -lean_dec(x_3); -return x_5; +lean_object* x_14; 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; uint8_t x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; uint8_t x_32; +x_14 = lean_st_ref_get(x_5, x_13); +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +x_16 = lean_ctor_get(x_14, 1); +lean_inc(x_16); +lean_dec(x_14); +x_17 = lean_st_ref_get(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_box(0); +x_21 = l_Lean_Expr_isLambda(x_1); +x_22 = lean_ctor_get(x_18, 7); +lean_inc(x_22); +lean_dec(x_18); +x_23 = lean_ctor_get(x_15, 6); +lean_inc(x_23); +lean_dec(x_15); +x_24 = 0; +x_25 = lean_unsigned_to_nat(1u); +lean_inc_n(x_1, 4); +x_26 = lean_alloc_ctor(0, 10, 5); +lean_ctor_set(x_26, 0, x_1); +lean_ctor_set(x_26, 1, x_1); +lean_ctor_set(x_26, 2, x_1); +lean_ctor_set(x_26, 3, x_1); +lean_ctor_set(x_26, 4, x_20); +lean_ctor_set(x_26, 5, x_20); +lean_ctor_set(x_26, 6, x_25); +lean_ctor_set(x_26, 7, x_22); +lean_ctor_set(x_26, 8, x_4); +lean_ctor_set(x_26, 9, x_23); +lean_ctor_set_uint8(x_26, sizeof(void*)*10, x_24); +lean_ctor_set_uint8(x_26, sizeof(void*)*10 + 1, x_2); +lean_ctor_set_uint8(x_26, sizeof(void*)*10 + 2, x_3); +lean_ctor_set_uint8(x_26, sizeof(void*)*10 + 3, x_21); +lean_ctor_set_uint8(x_26, sizeof(void*)*10 + 4, x_24); +x_27 = l_Lean_Meta_Grind_setENode(x_1, x_26, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_19); +x_28 = lean_ctor_get(x_27, 1); +lean_inc(x_28); +lean_dec(x_27); +x_29 = lean_st_ref_take(x_5, x_28); +x_30 = lean_ctor_get(x_29, 0); +lean_inc(x_30); +x_31 = lean_ctor_get(x_29, 1); +lean_inc(x_31); +lean_dec(x_29); +x_32 = !lean_is_exclusive(x_30); +if (x_32 == 0) +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; +x_33 = lean_ctor_get(x_30, 7); +x_34 = lean_nat_add(x_33, x_25); +lean_dec(x_33); +lean_ctor_set(x_30, 7, x_34); +x_35 = lean_st_ref_set(x_5, x_30, x_31); +x_36 = !lean_is_exclusive(x_35); +if (x_36 == 0) +{ +lean_object* x_37; lean_object* x_38; +x_37 = lean_ctor_get(x_35, 0); +lean_dec(x_37); +x_38 = lean_box(0); +lean_ctor_set(x_35, 0, x_38); +return x_35; +} +else +{ +lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_39 = lean_ctor_get(x_35, 1); +lean_inc(x_39); +lean_dec(x_35); +x_40 = lean_box(0); +x_41 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_41, 0, x_40); +lean_ctor_set(x_41, 1, x_39); +return x_41; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_registerParent___spec__7___boxed(lean_object* x_1, lean_object* x_2) { -_start: +else { -lean_object* x_3; -x_3 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_registerParent___spec__7(x_1, x_2); -lean_dec(x_2); -return x_3; +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; 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; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; +x_42 = lean_ctor_get(x_30, 0); +x_43 = lean_ctor_get(x_30, 1); +x_44 = lean_ctor_get(x_30, 2); +x_45 = lean_ctor_get(x_30, 3); +x_46 = lean_ctor_get(x_30, 4); +x_47 = lean_ctor_get(x_30, 5); +x_48 = lean_ctor_get_uint8(x_30, sizeof(void*)*20); +x_49 = lean_ctor_get(x_30, 6); +x_50 = lean_ctor_get(x_30, 7); +x_51 = lean_ctor_get(x_30, 8); +x_52 = lean_ctor_get(x_30, 9); +x_53 = lean_ctor_get(x_30, 10); +x_54 = lean_ctor_get(x_30, 11); +x_55 = lean_ctor_get(x_30, 12); +x_56 = lean_ctor_get(x_30, 13); +x_57 = lean_ctor_get(x_30, 14); +x_58 = lean_ctor_get(x_30, 15); +x_59 = lean_ctor_get(x_30, 16); +x_60 = lean_ctor_get(x_30, 17); +x_61 = lean_ctor_get(x_30, 18); +x_62 = lean_ctor_get(x_30, 19); +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_inc(x_51); +lean_inc(x_50); +lean_inc(x_49); +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_30); +x_63 = lean_nat_add(x_50, x_25); +lean_dec(x_50); +x_64 = lean_alloc_ctor(0, 20, 1); +lean_ctor_set(x_64, 0, x_42); +lean_ctor_set(x_64, 1, x_43); +lean_ctor_set(x_64, 2, x_44); +lean_ctor_set(x_64, 3, x_45); +lean_ctor_set(x_64, 4, x_46); +lean_ctor_set(x_64, 5, x_47); +lean_ctor_set(x_64, 6, x_49); +lean_ctor_set(x_64, 7, x_63); +lean_ctor_set(x_64, 8, x_51); +lean_ctor_set(x_64, 9, x_52); +lean_ctor_set(x_64, 10, x_53); +lean_ctor_set(x_64, 11, x_54); +lean_ctor_set(x_64, 12, x_55); +lean_ctor_set(x_64, 13, x_56); +lean_ctor_set(x_64, 14, x_57); +lean_ctor_set(x_64, 15, x_58); +lean_ctor_set(x_64, 16, x_59); +lean_ctor_set(x_64, 17, x_60); +lean_ctor_set(x_64, 18, x_61); +lean_ctor_set(x_64, 19, x_62); +lean_ctor_set_uint8(x_64, sizeof(void*)*20, x_48); +x_65 = lean_st_ref_set(x_5, x_64, x_31); +x_66 = lean_ctor_get(x_65, 1); +lean_inc(x_66); +if (lean_is_exclusive(x_65)) { + lean_ctor_release(x_65, 0); + lean_ctor_release(x_65, 1); + x_67 = x_65; +} else { + lean_dec_ref(x_65); + x_67 = lean_box(0); } +x_68 = lean_box(0); +if (lean_is_scalar(x_67)) { + x_69 = lean_alloc_ctor(0, 2, 0); +} else { + x_69 = x_67; } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_registerParent___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +lean_ctor_set(x_69, 0, x_68); +lean_ctor_set(x_69, 1, x_66); +return x_69; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkENodeCore___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { -lean_object* x_12; -x_12 = l_Lean_Meta_Grind_registerParent(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +uint8_t x_14; uint8_t x_15; lean_object* x_16; +x_14 = lean_unbox(x_2); +lean_dec(x_2); +x_15 = lean_unbox(x_3); +lean_dec(x_3); +x_16 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_14, x_15, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_12); +lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_12; +return x_16; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getParents(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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_Grind_mkENode___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -lean_object* x_11; uint8_t x_12; -x_11 = lean_st_ref_get(x_2, x_10); -x_12 = !lean_is_exclusive(x_11); -if (x_12 == 0) -{ -lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_13 = lean_ctor_get(x_11, 0); -x_14 = lean_ctor_get(x_13, 2); +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_13 = l_Lean_Meta_isConstructorAppCore_x3f(x_1, x_8, x_9, x_10, x_11, x_12); +x_14 = lean_ctor_get(x_13, 0); lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); lean_dec(x_13); -x_15 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_registerParent___spec__7(x_14, x_1); -if (lean_obj_tag(x_15) == 0) +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_1); +x_16 = l_Lean_Meta_Grind_isInterpreted(x_1, x_8, x_9, x_10, x_11, x_15); +if (lean_obj_tag(x_14) == 0) { -lean_object* x_16; -x_16 = lean_box(0); -lean_ctor_set(x_11, 0, x_16); -return x_11; -} -else +if (lean_obj_tag(x_16) == 0) { -lean_object* x_17; -x_17 = lean_ctor_get(x_15, 0); +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); -lean_dec(x_15); -lean_ctor_set(x_11, 0, x_17); -return x_11; -} +x_18 = lean_ctor_get(x_16, 1); +lean_inc(x_18); +lean_dec(x_16); +x_19 = 0; +x_20 = lean_unbox(x_17); +lean_dec(x_17); +x_21 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_20, x_19, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_18); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +return x_21; } else { -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_18 = lean_ctor_get(x_11, 0); -x_19 = lean_ctor_get(x_11, 1); -lean_inc(x_19); -lean_inc(x_18); +uint8_t x_22; lean_dec(x_11); -x_20 = lean_ctor_get(x_18, 2); -lean_inc(x_20); -lean_dec(x_18); -x_21 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_registerParent___spec__7(x_20, x_1); -if (lean_obj_tag(x_21) == 0) +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_2); +lean_dec(x_1); +x_22 = !lean_is_exclusive(x_16); +if (x_22 == 0) { -lean_object* x_22; lean_object* x_23; -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_19); -return x_23; +return x_16; } else { -lean_object* x_24; lean_object* x_25; -x_24 = lean_ctor_get(x_21, 0); +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_16, 0); +x_24 = lean_ctor_get(x_16, 1); lean_inc(x_24); -lean_dec(x_21); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_24); -lean_ctor_set(x_25, 1, x_19); +lean_inc(x_23); +lean_dec(x_16); +x_25 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_24); return x_25; } } } +else +{ +lean_dec(x_14); +if (lean_obj_tag(x_16) == 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_16, 0); +lean_inc(x_26); +x_27 = lean_ctor_get(x_16, 1); +lean_inc(x_27); +lean_dec(x_16); +x_28 = 1; +x_29 = lean_unbox(x_26); +lean_dec(x_26); +x_30 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_29, x_28, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_27); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +return x_30; } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getParents___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: +else { -lean_object* x_11; -x_11 = l_Lean_Meta_Grind_getParents(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +uint8_t x_31; +lean_dec(x_11); +lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -return x_11; +x_31 = !lean_is_exclusive(x_16); +if (x_31 == 0) +{ +return x_16; +} +else +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_16, 0); +x_33 = lean_ctor_get(x_16, 1); +lean_inc(x_33); +lean_inc(x_32); +lean_dec(x_16); +x_34 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +return x_34; } } -LEAN_EXPORT lean_object* l_Array_indexOfAux___at_Lean_Meta_Grind_getParentsAndReset___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkENode(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, 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_4; uint8_t x_5; -x_4 = lean_array_get_size(x_1); -x_5 = lean_nat_dec_lt(x_3, x_4); -lean_dec(x_4); -if (x_5 == 0) +lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_12 = l_Lean_Meta_Grind_alreadyInternalized(x_1, 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_6; -lean_dec(x_3); -x_6 = lean_box(0); -return x_6; +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_Meta_Grind_mkENode___lambda__1(x_1, x_2, x_16, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_15); +return x_17; } else { -lean_object* x_7; uint8_t x_8; -x_7 = lean_array_fget(x_1, x_3); -x_8 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_7, x_2); +uint8_t x_18; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_7); -if (x_8 == 0) +lean_dec(x_2); +lean_dec(x_1); +x_18 = !lean_is_exclusive(x_12); +if (x_18 == 0) { -lean_object* x_9; lean_object* x_10; -x_9 = lean_unsigned_to_nat(1u); -x_10 = lean_nat_add(x_3, x_9); -lean_dec(x_3); -x_3 = x_10; -goto _start; +lean_object* x_19; lean_object* x_20; +x_19 = lean_ctor_get(x_12, 0); +lean_dec(x_19); +x_20 = lean_box(0); +lean_ctor_set(x_12, 0, x_20); +return x_12; } else { -lean_object* x_12; -x_12 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_12, 0, x_3); -return x_12; +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_12, 1); +lean_inc(x_21); +lean_dec(x_12); +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_EXPORT lean_object* l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_getParentsAndReset___spec__2(lean_object* x_1, size_t x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkENode___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: { -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); -x_5 = 5; -x_6 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___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); -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 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_3, x_11); -lean_dec(x_11); -if (x_12 == 0) -{ -lean_dec(x_8); +lean_object* x_13; +x_13 = l_Lean_Meta_Grind_mkENode___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_7); +lean_dec(x_6); +lean_dec(x_5); lean_dec(x_4); -return x_1; +lean_dec(x_3); +return x_13; } -else -{ -uint8_t x_13; -x_13 = !lean_is_exclusive(x_1); -if (x_13 == 0) -{ -lean_object* x_14; lean_object* x_15; -x_14 = lean_ctor_get(x_1, 0); -lean_dec(x_14); -x_15 = lean_array_set(x_4, x_8, x_9); -lean_dec(x_8); -lean_ctor_set(x_1, 0, x_15); -return x_1; } -else +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkENode___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, 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_16; lean_object* x_17; -lean_dec(x_1); -x_16 = lean_array_set(x_4, x_8, x_9); -lean_dec(x_8); -x_17 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_17, 0, x_16); -return x_17; -} +lean_object* x_12; +x_12 = l_Lean_Meta_Grind_mkENode(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); +return x_12; } } -case 1: -{ -uint8_t x_18; -x_18 = !lean_is_exclusive(x_1); -if (x_18 == 0) -{ -lean_object* x_19; uint8_t x_20; -x_19 = lean_ctor_get(x_1, 0); -lean_dec(x_19); -x_20 = !lean_is_exclusive(x_10); -if (x_20 == 0) -{ -lean_object* x_21; lean_object* x_22; size_t x_23; lean_object* x_24; lean_object* x_25; -x_21 = lean_ctor_get(x_10, 0); -x_22 = lean_array_set(x_4, x_8, x_9); -x_23 = lean_usize_shift_right(x_2, x_5); -x_24 = l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_getParentsAndReset___spec__2(x_21, x_23, x_3); -lean_inc(x_24); -x_25 = l_Lean_PersistentHashMap_isUnaryNode___rarg(x_24); -if (lean_obj_tag(x_25) == 0) +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isCongrRoot(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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_26; -lean_ctor_set(x_10, 0, x_24); -x_26 = lean_array_set(x_22, x_8, x_10); -lean_dec(x_8); -lean_ctor_set(x_1, 0, x_26); -return x_1; -} -else +lean_object* x_11; +x_11 = l_Lean_Meta_Grind_getENode(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_11) == 0) { -lean_object* x_27; uint8_t x_28; -lean_dec(x_24); -lean_free_object(x_10); -x_27 = lean_ctor_get(x_25, 0); -lean_inc(x_27); -lean_dec(x_25); -x_28 = !lean_is_exclusive(x_27); -if (x_28 == 0) +uint8_t x_12; +x_12 = !lean_is_exclusive(x_11); +if (x_12 == 0) { -lean_object* x_29; -x_29 = lean_array_set(x_22, x_8, x_27); -lean_dec(x_8); -lean_ctor_set(x_1, 0, x_29); -return x_1; +lean_object* x_13; uint8_t x_14; lean_object* x_15; +x_13 = lean_ctor_get(x_11, 0); +x_14 = l_Lean_Meta_Grind_ENode_isCongrRoot(x_13); +lean_dec(x_13); +x_15 = lean_box(x_14); +lean_ctor_set(x_11, 0, x_15); +return x_11; } else { -lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_30 = lean_ctor_get(x_27, 0); -x_31 = lean_ctor_get(x_27, 1); -lean_inc(x_31); -lean_inc(x_30); -lean_dec(x_27); -x_32 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_32, 0, x_30); -lean_ctor_set(x_32, 1, x_31); -x_33 = lean_array_set(x_22, x_8, x_32); -lean_dec(x_8); -lean_ctor_set(x_1, 0, x_33); -return x_1; -} +lean_object* x_16; lean_object* x_17; uint8_t x_18; lean_object* x_19; lean_object* x_20; +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 = l_Lean_Meta_Grind_ENode_isCongrRoot(x_16); +lean_dec(x_16); +x_19 = lean_box(x_18); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_19); +lean_ctor_set(x_20, 1, x_17); +return x_20; } } else { -lean_object* x_34; lean_object* x_35; size_t x_36; lean_object* x_37; lean_object* x_38; -x_34 = lean_ctor_get(x_10, 0); -lean_inc(x_34); -lean_dec(x_10); -x_35 = lean_array_set(x_4, x_8, x_9); -x_36 = lean_usize_shift_right(x_2, x_5); -x_37 = l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_getParentsAndReset___spec__2(x_34, x_36, x_3); -lean_inc(x_37); -x_38 = l_Lean_PersistentHashMap_isUnaryNode___rarg(x_37); -if (lean_obj_tag(x_38) == 0) +uint8_t x_21; +x_21 = !lean_is_exclusive(x_11); +if (x_21 == 0) { -lean_object* x_39; lean_object* x_40; -x_39 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_39, 0, x_37); -x_40 = lean_array_set(x_35, x_8, x_39); -lean_dec(x_8); -lean_ctor_set(x_1, 0, x_40); -return x_1; +return x_11; } 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_dec(x_37); -x_41 = lean_ctor_get(x_38, 0); -lean_inc(x_41); -lean_dec(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_43); -if (lean_is_exclusive(x_41)) { - lean_ctor_release(x_41, 0); - lean_ctor_release(x_41, 1); - x_44 = x_41; -} else { - lean_dec_ref(x_41); - x_44 = lean_box(0); +lean_object* x_22; lean_object* x_23; lean_object* x_24; +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(1, 2, 0); +lean_ctor_set(x_24, 0, x_22); +lean_ctor_set(x_24, 1, x_23); +return x_24; +} } -if (lean_is_scalar(x_44)) { - x_45 = lean_alloc_ctor(0, 2, 0); -} else { - x_45 = x_44; } -lean_ctor_set(x_45, 0, x_42); -lean_ctor_set(x_45, 1, x_43); -x_46 = lean_array_set(x_35, x_8, x_45); -lean_dec(x_8); -lean_ctor_set(x_1, 0, x_46); -return x_1; } +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isCongrRoot___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_Meta_Grind_isCongrRoot(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_11; } } -else +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getCongrRoot___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_47; lean_object* x_48; lean_object* x_49; size_t x_50; lean_object* x_51; lean_object* x_52; +lean_object* x_12; lean_object* x_13; +x_12 = lean_ctor_get(x_1, 3); +lean_inc(x_12); lean_dec(x_1); -x_47 = lean_ctor_get(x_10, 0); -lean_inc(x_47); -if (lean_is_exclusive(x_10)) { - lean_ctor_release(x_10, 0); - x_48 = x_10; -} else { - lean_dec_ref(x_10); - x_48 = lean_box(0); +x_13 = l_Lean_Meta_Grind_getCongrRoot(x_12, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +return x_13; } -x_49 = lean_array_set(x_4, x_8, x_9); -x_50 = lean_usize_shift_right(x_2, x_5); -x_51 = l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_getParentsAndReset___spec__2(x_47, x_50, x_3); -lean_inc(x_51); -x_52 = l_Lean_PersistentHashMap_isUnaryNode___rarg(x_51); -if (lean_obj_tag(x_52) == 0) -{ -lean_object* x_53; lean_object* x_54; lean_object* x_55; -if (lean_is_scalar(x_48)) { - x_53 = lean_alloc_ctor(1, 1, 0); -} else { - x_53 = x_48; } -lean_ctor_set(x_53, 0, x_51); -x_54 = lean_array_set(x_49, x_8, x_53); -lean_dec(x_8); -x_55 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_55, 0, x_54); -return x_55; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getCongrRoot(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +lean_inc(x_1); +x_11 = l_Lean_Meta_Grind_getENode(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_11) == 0) +{ +uint8_t x_12; +x_12 = !lean_is_exclusive(x_11); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_13 = lean_ctor_get(x_11, 0); +x_14 = lean_ctor_get(x_11, 1); +x_15 = lean_ctor_get(x_13, 3); +lean_inc(x_15); +x_16 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_15, x_1); +lean_dec(x_15); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; +lean_free_object(x_11); +lean_dec(x_1); +x_17 = lean_box(0); +x_18 = l_Lean_Meta_Grind_getCongrRoot___lambda__1(x_13, x_17, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_14); +return x_18; } 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_dec(x_51); -lean_dec(x_48); -x_56 = lean_ctor_get(x_52, 0); -lean_inc(x_56); -lean_dec(x_52); -x_57 = lean_ctor_get(x_56, 0); -lean_inc(x_57); -x_58 = lean_ctor_get(x_56, 1); -lean_inc(x_58); -if (lean_is_exclusive(x_56)) { - lean_ctor_release(x_56, 0); - lean_ctor_release(x_56, 1); - x_59 = x_56; -} else { - lean_dec_ref(x_56); - x_59 = lean_box(0); -} -if (lean_is_scalar(x_59)) { - x_60 = lean_alloc_ctor(0, 2, 0); -} else { - x_60 = x_59; -} -lean_ctor_set(x_60, 0, x_57); -lean_ctor_set(x_60, 1, x_58); -x_61 = lean_array_set(x_49, x_8, x_60); -lean_dec(x_8); -x_62 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_62, 0, x_61); -return x_62; +lean_dec(x_13); +lean_ctor_set(x_11, 0, x_1); +return x_11; } } +else +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_19 = lean_ctor_get(x_11, 0); +x_20 = lean_ctor_get(x_11, 1); +lean_inc(x_20); +lean_inc(x_19); +lean_dec(x_11); +x_21 = lean_ctor_get(x_19, 3); +lean_inc(x_21); +x_22 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_21, x_1); +lean_dec(x_21); +if (x_22 == 0) +{ +lean_object* x_23; lean_object* x_24; +lean_dec(x_1); +x_23 = lean_box(0); +x_24 = l_Lean_Meta_Grind_getCongrRoot___lambda__1(x_19, x_23, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_20); +return x_24; } -default: +else { -lean_dec(x_8); -lean_dec(x_4); -return x_1; +lean_object* x_25; +lean_dec(x_19); +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_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; -x_63 = lean_ctor_get(x_1, 0); -lean_inc(x_63); -x_64 = lean_ctor_get(x_1, 1); -lean_inc(x_64); -x_65 = lean_unsigned_to_nat(0u); -x_66 = l_Array_indexOfAux___at_Lean_Meta_Grind_getParentsAndReset___spec__3(x_63, x_3, x_65); -if (lean_obj_tag(x_66) == 0) +uint8_t x_26; +lean_dec(x_1); +x_26 = !lean_is_exclusive(x_11); +if (x_26 == 0) { -lean_dec(x_64); -lean_dec(x_63); -return x_1; +return x_11; } else { -uint8_t x_67; -x_67 = !lean_is_exclusive(x_1); -if (x_67 == 0) -{ -lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; -x_68 = lean_ctor_get(x_1, 1); -lean_dec(x_68); -x_69 = lean_ctor_get(x_1, 0); -lean_dec(x_69); -x_70 = lean_ctor_get(x_66, 0); -lean_inc(x_70); -lean_dec(x_66); -lean_inc(x_70); -x_71 = l_Array_eraseIdx___rarg(x_63, x_70, lean_box(0)); -x_72 = l_Array_eraseIdx___rarg(x_64, x_70, lean_box(0)); -lean_ctor_set(x_1, 1, x_72); -lean_ctor_set(x_1, 0, x_71); -return x_1; +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_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; } -else -{ -lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; -lean_dec(x_1); -x_73 = lean_ctor_get(x_66, 0); -lean_inc(x_73); -lean_dec(x_66); -lean_inc(x_73); -x_74 = l_Array_eraseIdx___rarg(x_63, x_73, lean_box(0)); -x_75 = l_Array_eraseIdx___rarg(x_64, x_73, lean_box(0)); -x_76 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_76, 0, x_74); -lean_ctor_set(x_76, 1, x_75); -return x_76; } } } +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getCongrRoot___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_Meta_Grind_getCongrRoot___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_12; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_erase___at_Lean_Meta_Grind_getParentsAndReset___spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getCongrRoot___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -uint64_t x_3; size_t x_4; lean_object* x_5; -x_3 = l_Lean_Meta_Grind_instHashableENodeKey_unsafe__1(x_2); -x_4 = lean_uint64_to_usize(x_3); -x_5 = l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_getParentsAndReset___spec__2(x_1, x_4, x_2); -return x_5; +lean_object* x_11; +x_11 = l_Lean_Meta_Grind_getCongrRoot(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_11; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getParentsAndReset(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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_Grind_isInconsistent(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; -x_11 = l_Lean_Meta_Grind_getParents(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); -x_13 = lean_ctor_get(x_11, 1); -lean_inc(x_13); -lean_dec(x_11); -x_14 = lean_st_ref_take(x_2, x_13); -x_15 = lean_ctor_get(x_14, 0); -lean_inc(x_15); -x_16 = lean_ctor_get(x_14, 1); -lean_inc(x_16); -lean_dec(x_14); -x_17 = !lean_is_exclusive(x_15); -if (x_17 == 0) -{ -lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; -x_18 = lean_ctor_get(x_15, 2); -x_19 = l_Lean_PersistentHashMap_erase___at_Lean_Meta_Grind_getParentsAndReset___spec__1(x_18, x_1); -lean_ctor_set(x_15, 2, x_19); -x_20 = lean_st_ref_set(x_2, x_15, x_16); -x_21 = !lean_is_exclusive(x_20); -if (x_21 == 0) -{ -lean_object* x_22; -x_22 = lean_ctor_get(x_20, 0); -lean_dec(x_22); -lean_ctor_set(x_20, 0, x_12); -return x_20; -} -else +lean_object* x_10; uint8_t x_11; +x_10 = lean_st_ref_get(x_1, x_9); +x_11 = !lean_is_exclusive(x_10); +if (x_11 == 0) { -lean_object* x_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_12); -lean_ctor_set(x_24, 1, x_23); -return x_24; -} +lean_object* x_12; uint8_t x_13; lean_object* x_14; +x_12 = lean_ctor_get(x_10, 0); +x_13 = lean_ctor_get_uint8(x_12, sizeof(void*)*20); +lean_dec(x_12); +x_14 = lean_box(x_13); +lean_ctor_set(x_10, 0, x_14); +return x_10; } else { -lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; 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_25 = lean_ctor_get(x_15, 0); -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); -x_30 = lean_ctor_get(x_15, 5); -x_31 = lean_ctor_get_uint8(x_15, sizeof(void*)*14); -x_32 = lean_ctor_get(x_15, 6); -x_33 = lean_ctor_get(x_15, 7); -x_34 = lean_ctor_get(x_15, 8); -x_35 = lean_ctor_get(x_15, 9); -x_36 = lean_ctor_get(x_15, 10); -x_37 = lean_ctor_get(x_15, 11); -x_38 = lean_ctor_get(x_15, 12); -x_39 = lean_ctor_get(x_15, 13); -lean_inc(x_39); -lean_inc(x_38); -lean_inc(x_37); -lean_inc(x_36); -lean_inc(x_35); -lean_inc(x_34); -lean_inc(x_33); -lean_inc(x_32); -lean_inc(x_30); -lean_inc(x_29); -lean_inc(x_28); -lean_inc(x_27); -lean_inc(x_26); -lean_inc(x_25); +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_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_ctor_get_uint8(x_15, sizeof(void*)*20); lean_dec(x_15); -x_40 = l_Lean_PersistentHashMap_erase___at_Lean_Meta_Grind_getParentsAndReset___spec__1(x_27, x_1); -x_41 = lean_alloc_ctor(0, 14, 1); -lean_ctor_set(x_41, 0, x_25); -lean_ctor_set(x_41, 1, x_26); -lean_ctor_set(x_41, 2, x_40); -lean_ctor_set(x_41, 3, x_28); -lean_ctor_set(x_41, 4, x_29); -lean_ctor_set(x_41, 5, x_30); -lean_ctor_set(x_41, 6, x_32); -lean_ctor_set(x_41, 7, x_33); -lean_ctor_set(x_41, 8, x_34); -lean_ctor_set(x_41, 9, x_35); -lean_ctor_set(x_41, 10, x_36); -lean_ctor_set(x_41, 11, x_37); -lean_ctor_set(x_41, 12, x_38); -lean_ctor_set(x_41, 13, x_39); -lean_ctor_set_uint8(x_41, sizeof(void*)*14, x_31); -x_42 = lean_st_ref_set(x_2, x_41, x_16); -x_43 = lean_ctor_get(x_42, 1); -lean_inc(x_43); -if (lean_is_exclusive(x_42)) { - lean_ctor_release(x_42, 0); - lean_ctor_release(x_42, 1); - x_44 = x_42; -} else { - lean_dec_ref(x_42); - x_44 = lean_box(0); -} -if (lean_is_scalar(x_44)) { - x_45 = lean_alloc_ctor(0, 2, 0); -} else { - x_45 = x_44; -} -lean_ctor_set(x_45, 0, x_12); -lean_ctor_set(x_45, 1, x_43); -return x_45; +x_18 = lean_box(x_17); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_16); +return x_19; } } } -LEAN_EXPORT lean_object* l_Array_indexOfAux___at_Lean_Meta_Grind_getParentsAndReset___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isInconsistent___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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_Array_indexOfAux___at_Lean_Meta_Grind_getParentsAndReset___spec__3(x_1, x_2, x_3); +lean_object* x_10; +x_10 = l_Lean_Meta_Grind_isInconsistent(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); lean_dec(x_1); -return x_4; +return x_10; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_getParentsAndReset___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEqProof___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, 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_4; lean_object* x_5; -x_4 = lean_unbox_usize(x_2); -lean_dec(x_2); -x_5 = l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_getParentsAndReset___spec__2(x_1, x_4, x_3); -lean_dec(x_3); -return x_5; +lean_object* x_12; +x_12 = lean_grind_mk_eq_proof(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +return x_12; +} } +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkHEqProof___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; +x_12 = lean_grind_mk_heq_proof(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEqHEqProof(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, 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_2); +lean_inc(x_1); +x_12 = l_Lean_Meta_Grind_hasSameType(x_1, x_2, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_12) == 0) +{ +lean_object* x_13; uint8_t x_14; +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_unbox(x_13); +lean_dec(x_13); +if (x_14 == 0) +{ +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_grind_mk_heq_proof(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_15); +return x_16; } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_erase___at_Lean_Meta_Grind_getParentsAndReset___spec__1___boxed(lean_object* x_1, lean_object* x_2) { -_start: +else { -lean_object* x_3; -x_3 = l_Lean_PersistentHashMap_erase___at_Lean_Meta_Grind_getParentsAndReset___spec__1(x_1, x_2); -lean_dec(x_2); -return x_3; +lean_object* x_17; lean_object* x_18; +x_17 = lean_ctor_get(x_12, 1); +lean_inc(x_17); +lean_dec(x_12); +x_18 = lean_grind_mk_eq_proof(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_17); +return x_18; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getParentsAndReset___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: +else { -lean_object* x_11; -x_11 = l_Lean_Meta_Grind_getParentsAndReset(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +uint8_t x_19; +lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -15047,927 +19989,1038 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -return x_11; -} -} -LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at_Lean_Meta_Grind_copyParentsTo___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) { -_start: -{ -if (lean_obj_tag(x_1) == 0) +x_19 = !lean_is_exclusive(x_12); +if (x_19 == 0) { -lean_object* x_12; lean_object* x_13; -x_12 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_12, 0, x_2); -x_13 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_13, 0, x_12); -lean_ctor_set(x_13, 1, x_11); -return x_13; +return x_12; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_14 = lean_ctor_get(x_1, 0); -lean_inc(x_14); -x_15 = lean_ctor_get(x_1, 1); -lean_inc(x_15); -x_16 = lean_ctor_get(x_1, 3); -lean_inc(x_16); -lean_dec(x_1); -x_17 = l_Lean_RBNode_forIn_visit___at_Lean_Meta_Grind_copyParentsTo___spec__1(x_14, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -x_18 = lean_ctor_get(x_17, 0); -lean_inc(x_18); -x_19 = lean_ctor_get(x_17, 1); -lean_inc(x_19); -lean_dec(x_17); -x_20 = lean_ctor_get(x_18, 0); +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_12, 0); +x_21 = lean_ctor_get(x_12, 1); +lean_inc(x_21); lean_inc(x_20); -lean_dec(x_18); -x_21 = lean_box(0); -x_22 = l_Lean_RBNode_insert___at_Lean_Meta_Grind_registerParent___spec__1(x_20, x_15, x_21); -x_1 = x_16; -x_2 = x_22; -x_11 = x_19; -goto _start; +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_EXPORT lean_object* l_Lean_Meta_Grind_copyParentsTo(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_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_Grind_mkEqTrueProof(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_54; -x_12 = lean_st_ref_get(x_3, x_11); -x_13 = lean_ctor_get(x_12, 0); +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_11 = l_Lean_Meta_Grind_getTrueExpr___rarg(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); -x_14 = lean_ctor_get(x_12, 1); -lean_inc(x_14); -lean_dec(x_12); -x_15 = lean_ctor_get(x_13, 2); -lean_inc(x_15); -lean_dec(x_13); -x_54 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_registerParent___spec__7(x_15, x_2); -if (lean_obj_tag(x_54) == 0) -{ -lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; -x_55 = lean_box(0); -x_56 = l_Lean_RBNode_forIn_visit___at_Lean_Meta_Grind_copyParentsTo___spec__1(x_1, x_55, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_14); -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_inc(x_59); -lean_dec(x_57); -x_16 = x_59; -x_17 = x_58; -goto block_53; +lean_dec(x_11); +x_14 = lean_grind_mk_eq_proof(x_1, x_12, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_13); +return x_14; } -else +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEqFalseProof(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; -x_60 = lean_ctor_get(x_54, 0); -lean_inc(x_60); -lean_dec(x_54); -x_61 = l_Lean_RBNode_forIn_visit___at_Lean_Meta_Grind_copyParentsTo___spec__1(x_1, x_60, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_14); -x_62 = lean_ctor_get(x_61, 0); -lean_inc(x_62); -x_63 = lean_ctor_get(x_61, 1); -lean_inc(x_63); -lean_dec(x_61); -x_64 = lean_ctor_get(x_62, 0); -lean_inc(x_64); -lean_dec(x_62); -x_16 = x_64; -x_17 = x_63; -goto block_53; +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_11 = l_Lean_Meta_Grind_getFalseExpr___rarg(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_grind_mk_eq_proof(x_1, x_12, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_13); +return x_14; } -block_53: +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markAsInconsistent___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_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; -x_18 = lean_st_ref_take(x_3, x_17); -x_19 = lean_ctor_get(x_18, 0); -lean_inc(x_19); -x_20 = lean_ctor_get(x_18, 1); -lean_inc(x_20); -lean_dec(x_18); -x_21 = !lean_is_exclusive(x_19); -if (x_21 == 0) +lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_11 = lean_st_ref_take(x_2, x_10); +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +x_13 = lean_ctor_get(x_11, 1); +lean_inc(x_13); +lean_dec(x_11); +x_14 = !lean_is_exclusive(x_12); +if (x_14 == 0) { -lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; -x_22 = lean_ctor_get(x_19, 2); -x_23 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_registerParent___spec__3(x_22, x_2, x_16); -lean_ctor_set(x_19, 2, x_23); -x_24 = lean_st_ref_set(x_3, x_19, x_20); -x_25 = !lean_is_exclusive(x_24); -if (x_25 == 0) +uint8_t x_15; lean_object* x_16; uint8_t x_17; +x_15 = 1; +lean_ctor_set_uint8(x_12, sizeof(void*)*20, x_15); +x_16 = lean_st_ref_set(x_2, x_12, x_13); +x_17 = !lean_is_exclusive(x_16); +if (x_17 == 0) { -lean_object* x_26; lean_object* x_27; -x_26 = lean_ctor_get(x_24, 0); -lean_dec(x_26); -x_27 = lean_box(0); -lean_ctor_set(x_24, 0, x_27); -return x_24; +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_28; lean_object* x_29; lean_object* x_30; -x_28 = lean_ctor_get(x_24, 1); -lean_inc(x_28); -lean_dec(x_24); -x_29 = lean_box(0); -x_30 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_30, 0, x_29); -lean_ctor_set(x_30, 1, x_28); -return x_30; +lean_object* x_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_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; -x_31 = lean_ctor_get(x_19, 0); -x_32 = lean_ctor_get(x_19, 1); -x_33 = lean_ctor_get(x_19, 2); -x_34 = lean_ctor_get(x_19, 3); -x_35 = lean_ctor_get(x_19, 4); -x_36 = lean_ctor_get(x_19, 5); -x_37 = lean_ctor_get_uint8(x_19, sizeof(void*)*14); -x_38 = lean_ctor_get(x_19, 6); -x_39 = lean_ctor_get(x_19, 7); -x_40 = lean_ctor_get(x_19, 8); -x_41 = lean_ctor_get(x_19, 9); -x_42 = lean_ctor_get(x_19, 10); -x_43 = lean_ctor_get(x_19, 11); -x_44 = lean_ctor_get(x_19, 12); -x_45 = lean_ctor_get(x_19, 13); -lean_inc(x_45); -lean_inc(x_44); -lean_inc(x_43); +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; uint8_t x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_23 = lean_ctor_get(x_12, 0); +x_24 = lean_ctor_get(x_12, 1); +x_25 = lean_ctor_get(x_12, 2); +x_26 = lean_ctor_get(x_12, 3); +x_27 = lean_ctor_get(x_12, 4); +x_28 = lean_ctor_get(x_12, 5); +x_29 = lean_ctor_get(x_12, 6); +x_30 = lean_ctor_get(x_12, 7); +x_31 = lean_ctor_get(x_12, 8); +x_32 = lean_ctor_get(x_12, 9); +x_33 = lean_ctor_get(x_12, 10); +x_34 = lean_ctor_get(x_12, 11); +x_35 = lean_ctor_get(x_12, 12); +x_36 = lean_ctor_get(x_12, 13); +x_37 = lean_ctor_get(x_12, 14); +x_38 = lean_ctor_get(x_12, 15); +x_39 = lean_ctor_get(x_12, 16); +x_40 = lean_ctor_get(x_12, 17); +x_41 = lean_ctor_get(x_12, 18); +x_42 = lean_ctor_get(x_12, 19); 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_inc(x_34); lean_inc(x_33); lean_inc(x_32); lean_inc(x_31); -lean_dec(x_19); -x_46 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_registerParent___spec__3(x_33, x_2, x_16); -x_47 = lean_alloc_ctor(0, 14, 1); -lean_ctor_set(x_47, 0, x_31); -lean_ctor_set(x_47, 1, x_32); -lean_ctor_set(x_47, 2, x_46); -lean_ctor_set(x_47, 3, x_34); -lean_ctor_set(x_47, 4, x_35); -lean_ctor_set(x_47, 5, x_36); -lean_ctor_set(x_47, 6, x_38); -lean_ctor_set(x_47, 7, x_39); -lean_ctor_set(x_47, 8, x_40); -lean_ctor_set(x_47, 9, x_41); -lean_ctor_set(x_47, 10, x_42); -lean_ctor_set(x_47, 11, x_43); -lean_ctor_set(x_47, 12, x_44); -lean_ctor_set(x_47, 13, x_45); -lean_ctor_set_uint8(x_47, sizeof(void*)*14, x_37); -x_48 = lean_st_ref_set(x_3, x_47, x_20); -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; +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_dec(x_12); +x_43 = 1; +x_44 = lean_alloc_ctor(0, 20, 1); +lean_ctor_set(x_44, 0, x_23); +lean_ctor_set(x_44, 1, x_24); +lean_ctor_set(x_44, 2, x_25); +lean_ctor_set(x_44, 3, x_26); +lean_ctor_set(x_44, 4, x_27); +lean_ctor_set(x_44, 5, x_28); +lean_ctor_set(x_44, 6, x_29); +lean_ctor_set(x_44, 7, x_30); +lean_ctor_set(x_44, 8, x_31); +lean_ctor_set(x_44, 9, x_32); +lean_ctor_set(x_44, 10, x_33); +lean_ctor_set(x_44, 11, x_34); +lean_ctor_set(x_44, 12, x_35); +lean_ctor_set(x_44, 13, x_36); +lean_ctor_set(x_44, 14, x_37); +lean_ctor_set(x_44, 15, x_38); +lean_ctor_set(x_44, 16, x_39); +lean_ctor_set(x_44, 17, x_40); +lean_ctor_set(x_44, 18, x_41); +lean_ctor_set(x_44, 19, x_42); +lean_ctor_set_uint8(x_44, sizeof(void*)*20, x_43); +x_45 = lean_st_ref_set(x_2, x_44, x_13); +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_48); - x_50 = lean_box(0); + lean_dec_ref(x_45); + x_47 = lean_box(0); } -x_51 = lean_box(0); -if (lean_is_scalar(x_50)) { - x_52 = lean_alloc_ctor(0, 2, 0); +x_48 = lean_box(0); +if (lean_is_scalar(x_47)) { + x_49 = 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_RBNode_forIn_visit___at_Lean_Meta_Grind_copyParentsTo___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -lean_object* x_12; -x_12 = l_Lean_RBNode_forIn_visit___at_Lean_Meta_Grind_copyParentsTo___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -return x_12; + 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_Meta_Grind_copyParentsTo___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, 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_Meta_Grind_copyParentsTo(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -return x_12; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_setENode_unsafe__1___rarg(lean_object* x_1) { +static lean_object* _init_l_Lean_Meta_Grind_markAsInconsistent___closed__1() { _start: { -lean_object* x_2; -x_2 = lean_ctor_get(x_1, 3); -lean_inc(x_2); -return x_2; +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_markAsInconsistent___lambda__1___boxed), 10, 0); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_setENode_unsafe__1(lean_object* x_1, lean_object* x_2) { +static lean_object* _init_l_Lean_Meta_Grind_markAsInconsistent___closed__2() { _start: { -lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_setENode_unsafe__1___rarg___boxed), 1, 0); -return x_3; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("closed `", 8, 8); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_setENode_unsafe__1___rarg___boxed(lean_object* x_1) { +static lean_object* _init_l_Lean_Meta_Grind_markAsInconsistent___closed__3() { _start: { -lean_object* x_2; -x_2 = l_Lean_Meta_Grind_setENode_unsafe__1___rarg(x_1); -lean_dec(x_1); +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_markAsInconsistent___closed__2; +x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_setENode_unsafe__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markAsInconsistent(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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_3; -x_3 = l_Lean_Meta_Grind_setENode_unsafe__1(x_1, x_2); -lean_dec(x_2); -lean_dec(x_1); -return x_3; +lean_object* x_10; lean_object* x_11; uint8_t x_12; +x_10 = lean_st_ref_get(x_1, x_9); +x_11 = lean_ctor_get(x_10, 0); +lean_inc(x_11); +x_12 = lean_ctor_get_uint8(x_11, sizeof(void*)*20); +lean_dec(x_11); +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_10, 1); +lean_inc(x_13); +lean_dec(x_10); +x_14 = l_Lean_Meta_Grind_updateLastTag___closed__1; +x_15 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_14, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_13); +x_16 = !lean_is_exclusive(x_15); +if (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_15, 0); +x_18 = lean_ctor_get(x_15, 1); +x_19 = l_Lean_Meta_Grind_markAsInconsistent___closed__1; +x_20 = lean_unbox(x_17); +lean_dec(x_17); +if (x_20 == 0) +{ +lean_object* x_21; lean_object* x_22; +lean_free_object(x_15); +x_21 = lean_box(0); +x_22 = lean_apply_10(x_19, x_21, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_18); +return x_22; } +else +{ +lean_object* x_23; uint8_t x_24; +x_23 = lean_st_ref_get(x_1, x_18); +x_24 = !lean_is_exclusive(x_23); +if (x_24 == 0) +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_25 = lean_ctor_get(x_23, 0); +x_26 = lean_ctor_get(x_23, 1); +x_27 = lean_ctor_get(x_25, 0); +lean_inc(x_27); +lean_dec(x_25); +x_28 = l_Lean_MVarId_getTag(x_27, x_5, x_6, x_7, x_8, x_26); +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; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_29 = lean_ctor_get(x_28, 0); +lean_inc(x_29); +x_30 = lean_ctor_get(x_28, 1); +lean_inc(x_30); +lean_dec(x_28); +x_31 = l_Lean_MessageData_ofName(x_29); +x_32 = l_Lean_Meta_Grind_markAsInconsistent___closed__3; +lean_ctor_set_tag(x_23, 7); +lean_ctor_set(x_23, 1, x_31); +lean_ctor_set(x_23, 0, x_32); +x_33 = l_Lean_Meta_Grind_updateLastTag___closed__5; +lean_ctor_set_tag(x_15, 7); +lean_ctor_set(x_15, 1, x_33); +lean_ctor_set(x_15, 0, x_23); +x_34 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_14, x_15, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, 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); +x_37 = lean_apply_10(x_19, x_35, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_36); +return x_37; } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_setENode___spec__3(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: +else { -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); +uint8_t x_38; +lean_free_object(x_23); +lean_free_object(x_15); +lean_dec(x_8); lean_dec(x_7); -if (x_8 == 0) -{ +lean_dec(x_6); lean_dec(x_5); -return x_6; +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_38 = !lean_is_exclusive(x_28); +if (x_38 == 0) +{ +return x_28; } 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_Meta_Grind_instHashableENodeKey_unsafe__1(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_Grind_setENode___spec__2(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_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; } } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_setENode___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -lean_object* x_5; lean_object* x_6; lean_object* x_7; 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) +else { -uint8_t x_9; -lean_dec(x_2); -x_9 = !lean_is_exclusive(x_1); -if (x_9 == 0) +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_42 = lean_ctor_get(x_23, 0); +x_43 = lean_ctor_get(x_23, 1); +lean_inc(x_43); +lean_inc(x_42); +lean_dec(x_23); +x_44 = lean_ctor_get(x_42, 0); +lean_inc(x_44); +lean_dec(x_42); +x_45 = l_Lean_MVarId_getTag(x_44, x_5, x_6, x_7, x_8, x_43); +if (lean_obj_tag(x_45) == 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; +lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_46 = lean_ctor_get(x_45, 0); +lean_inc(x_46); +x_47 = lean_ctor_get(x_45, 1); +lean_inc(x_47); +lean_dec(x_45); +x_48 = l_Lean_MessageData_ofName(x_46); +x_49 = l_Lean_Meta_Grind_markAsInconsistent___closed__3; +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_Meta_Grind_updateLastTag___closed__5; +lean_ctor_set_tag(x_15, 7); +lean_ctor_set(x_15, 1, x_51); +lean_ctor_set(x_15, 0, x_50); +x_52 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_14, x_15, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, 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); +lean_dec(x_52); +x_55 = lean_apply_10(x_19, x_53, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_54); +return x_55; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; +lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; +lean_free_object(x_15); +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_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; +x_56 = lean_ctor_get(x_45, 0); +lean_inc(x_56); +x_57 = lean_ctor_get(x_45, 1); +lean_inc(x_57); +if (lean_is_exclusive(x_45)) { + lean_ctor_release(x_45, 0); + lean_ctor_release(x_45, 1); + x_58 = x_45; +} else { + lean_dec_ref(x_45); + x_58 = lean_box(0); +} +if (lean_is_scalar(x_58)) { + x_59 = lean_alloc_ctor(1, 2, 0); +} else { + x_59 = x_58; +} +lean_ctor_set(x_59, 0, x_56); +lean_ctor_set(x_59, 1, x_57); +return x_59; +} +} } } else { -lean_object* x_17; uint8_t x_18; -x_17 = lean_array_fget(x_5, x_2); -x_18 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_3, x_17); -lean_dec(x_17); -if (x_18 == 0) +lean_object* x_60; lean_object* x_61; lean_object* x_62; uint8_t x_63; +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_dec(x_15); +x_62 = l_Lean_Meta_Grind_markAsInconsistent___closed__1; +x_63 = lean_unbox(x_60); +lean_dec(x_60); +if (x_63 == 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; +lean_object* x_64; lean_object* x_65; +x_64 = lean_box(0); +x_65 = lean_apply_10(x_62, x_64, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_61); +return x_65; } else { -uint8_t x_22; -x_22 = !lean_is_exclusive(x_1); -if (x_22 == 0) +lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; +x_66 = lean_st_ref_get(x_1, x_61); +x_67 = lean_ctor_get(x_66, 0); +lean_inc(x_67); +x_68 = lean_ctor_get(x_66, 1); +lean_inc(x_68); +if (lean_is_exclusive(x_66)) { + lean_ctor_release(x_66, 0); + lean_ctor_release(x_66, 1); + x_69 = x_66; +} else { + lean_dec_ref(x_66); + x_69 = lean_box(0); +} +x_70 = lean_ctor_get(x_67, 0); +lean_inc(x_70); +lean_dec(x_67); +x_71 = l_Lean_MVarId_getTag(x_70, x_5, x_6, x_7, x_8, x_68); +if (lean_obj_tag(x_71) == 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; +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_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_MessageData_ofName(x_72); +x_75 = l_Lean_Meta_Grind_markAsInconsistent___closed__3; +if (lean_is_scalar(x_69)) { + x_76 = lean_alloc_ctor(7, 2, 0); +} else { + x_76 = x_69; + lean_ctor_set_tag(x_76, 7); +} +lean_ctor_set(x_76, 0, x_75); +lean_ctor_set(x_76, 1, x_74); +x_77 = l_Lean_Meta_Grind_updateLastTag___closed__5; +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_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_14, x_78, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_73); +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_apply_10(x_62, x_80, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_81); +return x_82; } 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_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; +lean_dec(x_69); +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_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_1); +x_83 = lean_ctor_get(x_71, 0); +lean_inc(x_83); +x_84 = lean_ctor_get(x_71, 1); +lean_inc(x_84); +if (lean_is_exclusive(x_71)) { + lean_ctor_release(x_71, 0); + lean_ctor_release(x_71, 1); + x_85 = x_71; +} else { + lean_dec_ref(x_71); + x_85 = lean_box(0); } +if (lean_is_scalar(x_85)) { + x_86 = lean_alloc_ctor(1, 2, 0); +} else { + x_86 = x_85; } +lean_ctor_set(x_86, 0, x_83); +lean_ctor_set(x_86, 1, x_84); +return x_86; } } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_setENode___spec__2(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_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___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) +} +else { -lean_dec(x_12); +uint8_t x_87; +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -return x_1; +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_87 = !lean_is_exclusive(x_10); +if (x_87 == 0) +{ +lean_object* x_88; lean_object* x_89; +x_88 = lean_ctor_get(x_10, 0); +lean_dec(x_88); +x_89 = lean_box(0); +lean_ctor_set(x_10, 0, x_89); +return x_10; } 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 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(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; +lean_object* x_90; lean_object* x_91; lean_object* x_92; +x_90 = lean_ctor_get(x_10, 1); +lean_inc(x_90); +lean_dec(x_10); +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_90); +return x_92; } -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 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(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_EXPORT lean_object* l_Lean_Meta_Grind_markAsInconsistent___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_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; -} +lean_object* x_11; +x_11 = l_Lean_Meta_Grind_markAsInconsistent___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_11; } } -case 1: +LEAN_EXPORT lean_object* l_Lean_MVarId_isAssigned___at_Lean_Meta_Grind_closeGoal___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: { -uint8_t x_34; -x_34 = !lean_is_exclusive(x_15); -if (x_34 == 0) +lean_object* x_11; uint8_t x_12; +x_11 = lean_st_ref_get(x_7, x_10); +x_12 = !lean_is_exclusive(x_11); +if (x_12 == 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_Grind_setENode___spec__2(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; +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_11, 0); +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +lean_dec(x_13); +x_15 = lean_ctor_get(x_14, 7); +lean_inc(x_15); +lean_dec(x_14); +x_16 = l_Lean_PersistentHashMap_contains___at_Lean_MVarId_isAssigned___spec__1(x_15, x_1); +x_17 = lean_box(x_16); +lean_ctor_set(x_11, 0, x_17); +return x_11; } 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_Grind_setENode___spec__2(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; +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; lean_object* x_23; lean_object* x_24; +x_18 = lean_ctor_get(x_11, 0); +x_19 = lean_ctor_get(x_11, 1); +lean_inc(x_19); +lean_inc(x_18); +lean_dec(x_11); +x_20 = lean_ctor_get(x_18, 0); +lean_inc(x_20); +lean_dec(x_18); +x_21 = lean_ctor_get(x_20, 7); +lean_inc(x_21); +lean_dec(x_20); +x_22 = l_Lean_PersistentHashMap_contains___at_Lean_MVarId_isAssigned___spec__1(x_21, x_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_19); +return x_24; } } -default: +} +LEAN_EXPORT lean_object* l_Lean_MVarId_assign___at_Lean_Meta_Grind_closeGoal___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: { -lean_object* x_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_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_12 = lean_st_ref_take(x_8, x_11); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_12, 1); +lean_inc(x_15); lean_dec(x_12); -lean_ctor_set(x_1, 0, x_47); -return x_1; -} -} -} -} -else +x_16 = !lean_is_exclusive(x_13); +if (x_16 == 0) { -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_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___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_17; uint8_t x_18; +x_17 = lean_ctor_get(x_13, 0); +lean_dec(x_17); +x_18 = !lean_is_exclusive(x_14); +if (x_18 == 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; +lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_19 = lean_ctor_get(x_14, 7); +x_20 = l_Lean_PersistentHashMap_insert___at_Lean_MVarId_assign___spec__1(x_19, x_1, x_2); +lean_ctor_set(x_14, 7, x_20); +x_21 = lean_st_ref_set(x_8, x_13, x_15); +x_22 = !lean_is_exclusive(x_21); +if (x_22 == 0) +{ +lean_object* x_23; lean_object* x_24; +x_23 = lean_ctor_get(x_21, 0); +lean_dec(x_23); +x_24 = lean_box(0); +lean_ctor_set(x_21, 0, x_24); +return x_21; } 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); +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; } -x_63 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(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); +lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_28 = lean_ctor_get(x_14, 0); +x_29 = lean_ctor_get(x_14, 1); +x_30 = lean_ctor_get(x_14, 2); +x_31 = lean_ctor_get(x_14, 3); +x_32 = lean_ctor_get(x_14, 4); +x_33 = lean_ctor_get(x_14, 5); +x_34 = lean_ctor_get(x_14, 6); +x_35 = lean_ctor_get(x_14, 7); +x_36 = lean_ctor_get(x_14, 8); +lean_inc(x_36); +lean_inc(x_35); +lean_inc(x_34); +lean_inc(x_33); +lean_inc(x_32); +lean_inc(x_31); +lean_inc(x_30); +lean_inc(x_29); +lean_inc(x_28); +lean_dec(x_14); +x_37 = l_Lean_PersistentHashMap_insert___at_Lean_MVarId_assign___spec__1(x_35, x_1, x_2); +x_38 = lean_alloc_ctor(0, 9, 0); +lean_ctor_set(x_38, 0, x_28); +lean_ctor_set(x_38, 1, x_29); +lean_ctor_set(x_38, 2, x_30); +lean_ctor_set(x_38, 3, x_31); +lean_ctor_set(x_38, 4, x_32); +lean_ctor_set(x_38, 5, x_33); +lean_ctor_set(x_38, 6, x_34); +lean_ctor_set(x_38, 7, x_37); +lean_ctor_set(x_38, 8, x_36); +lean_ctor_set(x_13, 0, x_38); +x_39 = lean_st_ref_set(x_8, x_13, x_15); +x_40 = lean_ctor_get(x_39, 1); +lean_inc(x_40); +if (lean_is_exclusive(x_39)) { + lean_ctor_release(x_39, 0); + lean_ctor_release(x_39, 1); + x_41 = x_39; } 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; -} + lean_dec_ref(x_39); + x_41 = lean_box(0); } -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; +x_42 = lean_box(0); +if (lean_is_scalar(x_41)) { + x_43 = lean_alloc_ctor(0, 2, 0); } else { - lean_dec_ref(x_57); - x_72 = lean_box(0); + x_43 = x_41; } -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_Grind_setENode___spec__2(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_43, 0, x_42); +lean_ctor_set(x_43, 1, x_40); +return x_43; } -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: +else { -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; -} -} -} +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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_44 = lean_ctor_get(x_13, 1); +x_45 = lean_ctor_get(x_13, 2); +x_46 = lean_ctor_get(x_13, 3); +x_47 = lean_ctor_get(x_13, 4); +lean_inc(x_47); +lean_inc(x_46); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_13); +x_48 = lean_ctor_get(x_14, 0); +lean_inc(x_48); +x_49 = lean_ctor_get(x_14, 1); +lean_inc(x_49); +x_50 = lean_ctor_get(x_14, 2); +lean_inc(x_50); +x_51 = lean_ctor_get(x_14, 3); +lean_inc(x_51); +x_52 = lean_ctor_get(x_14, 4); +lean_inc(x_52); +x_53 = lean_ctor_get(x_14, 5); +lean_inc(x_53); +x_54 = lean_ctor_get(x_14, 6); +lean_inc(x_54); +x_55 = lean_ctor_get(x_14, 7); +lean_inc(x_55); +x_56 = lean_ctor_get(x_14, 8); +lean_inc(x_56); +if (lean_is_exclusive(x_14)) { + lean_ctor_release(x_14, 0); + lean_ctor_release(x_14, 1); + lean_ctor_release(x_14, 2); + lean_ctor_release(x_14, 3); + lean_ctor_release(x_14, 4); + lean_ctor_release(x_14, 5); + lean_ctor_release(x_14, 6); + lean_ctor_release(x_14, 7); + lean_ctor_release(x_14, 8); + x_57 = x_14; +} else { + lean_dec_ref(x_14); + x_57 = lean_box(0); } +x_58 = l_Lean_PersistentHashMap_insert___at_Lean_MVarId_assign___spec__1(x_55, x_1, x_2); +if (lean_is_scalar(x_57)) { + x_59 = lean_alloc_ctor(0, 9, 0); +} else { + x_59 = x_57; } -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_Grind_setENode___spec__4(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_Grind_mkHCongrWithArity___spec__2___closed__3; -x_93 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_setENode___spec__3(x_3, x_90, x_91, lean_box(0), x_83, x_92); -lean_dec(x_91); -lean_dec(x_90); -return x_93; +lean_ctor_set(x_59, 0, x_48); +lean_ctor_set(x_59, 1, x_49); +lean_ctor_set(x_59, 2, x_50); +lean_ctor_set(x_59, 3, x_51); +lean_ctor_set(x_59, 4, x_52); +lean_ctor_set(x_59, 5, x_53); +lean_ctor_set(x_59, 6, x_54); +lean_ctor_set(x_59, 7, x_58); +lean_ctor_set(x_59, 8, x_56); +x_60 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_60, 0, x_59); +lean_ctor_set(x_60, 1, x_44); +lean_ctor_set(x_60, 2, x_45); +lean_ctor_set(x_60, 3, x_46); +lean_ctor_set(x_60, 4, x_47); +x_61 = lean_st_ref_set(x_8, x_60, x_15); +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); } -else -{ -return x_84; +x_64 = lean_box(0); +if (lean_is_scalar(x_63)) { + x_65 = lean_alloc_ctor(0, 2, 0); +} else { + x_65 = x_63; } +lean_ctor_set(x_65, 0, x_64); +lean_ctor_set(x_65, 1, x_62); +return x_65; } -else -{ -return x_84; } } -else +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_closeGoal(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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_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_Grind_setENode___spec__4(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_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); +lean_inc(x_2); +x_11 = l_Lean_Meta_Grind_markAsInconsistent(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_11) == 0) { -lean_object* x_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_12; 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_11, 1); +lean_inc(x_12); +lean_dec(x_11); +x_13 = lean_st_ref_get(x_2, x_12); +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +x_16 = lean_ctor_get(x_14, 0); +lean_inc(x_16); +lean_dec(x_14); +x_17 = l_Lean_MVarId_isAssigned___at_Lean_Meta_Grind_closeGoal___spec__1(x_16, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_15); +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +x_19 = lean_unbox(x_18); +lean_dec(x_18); +if (x_19 == 0) { -lean_object* x_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_Grind_mkHCongrWithArity___spec__2___closed__3; -x_107 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_setENode___spec__3(x_3, x_104, x_105, lean_box(0), x_97, x_106); -lean_dec(x_105); -lean_dec(x_104); -return x_107; +lean_object* x_20; lean_object* x_21; +x_20 = lean_ctor_get(x_17, 1); +lean_inc(x_20); +lean_dec(x_17); +lean_inc(x_16); +x_21 = l_Lean_MVarId_getType(x_16, x_6, x_7, x_8, x_9, x_20); +if (lean_obj_tag(x_21) == 0) +{ +lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_22 = lean_ctor_get(x_21, 0); +lean_inc(x_22); +x_23 = lean_ctor_get(x_21, 1); +lean_inc(x_23); +lean_dec(x_21); +lean_inc(x_22); +x_24 = l_Lean_Expr_isFalse(x_22); +if (x_24 == 0) +{ +lean_object* x_25; +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +x_25 = l_Lean_Meta_mkFalseElim(x_22, x_1, x_6, x_7, x_8, x_9, x_23); +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_28 = l_Lean_MVarId_assign___at_Lean_Meta_Grind_closeGoal___spec__2(x_16, x_26, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_27); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_28; } else { -return x_98; -} +uint8_t x_29; +lean_dec(x_16); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_29 = !lean_is_exclusive(x_25); +if (x_29 == 0) +{ +return x_25; } else { -return x_98; -} -} +lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_30 = lean_ctor_get(x_25, 0); +x_31 = lean_ctor_get(x_25, 1); +lean_inc(x_31); +lean_inc(x_30); +lean_dec(x_25); +x_32 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_32, 0, x_30); +lean_ctor_set(x_32, 1, x_31); +return x_32; } } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_setENode___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: +else { -uint64_t x_4; size_t x_5; size_t x_6; lean_object* x_7; -x_4 = l_Lean_Meta_Grind_instHashableENodeKey_unsafe__1(x_2); -x_5 = lean_uint64_to_usize(x_4); -x_6 = 1; -x_7 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_setENode___spec__2(x_1, x_5, x_6, x_2, x_3); -return x_7; -} +lean_object* x_33; +lean_dec(x_22); +x_33 = l_Lean_MVarId_assign___at_Lean_Meta_Grind_closeGoal___spec__2(x_16, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_23); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_33; } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_setENode(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, 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; -x_12 = lean_st_ref_take(x_3, x_11); -x_13 = lean_ctor_get(x_12, 0); -lean_inc(x_13); -x_14 = lean_ctor_get(x_12, 1); -lean_inc(x_14); -lean_dec(x_12); -x_15 = !lean_is_exclusive(x_13); -if (x_15 == 0) -{ -lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; -x_16 = lean_ctor_get(x_13, 1); -x_17 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_setENode___spec__1(x_16, x_1, x_2); -lean_ctor_set(x_13, 1, x_17); -x_18 = lean_st_ref_set(x_3, x_13, x_14); -x_19 = !lean_is_exclusive(x_18); -if (x_19 == 0) -{ -lean_object* x_20; lean_object* x_21; -x_20 = lean_ctor_get(x_18, 0); -lean_dec(x_20); -x_21 = lean_box(0); -lean_ctor_set(x_18, 0, x_21); -return x_18; } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_18, 1); -lean_inc(x_22); -lean_dec(x_18); -x_23 = lean_box(0); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_23); -lean_ctor_set(x_24, 1, x_22); -return x_24; -} +uint8_t x_34; +lean_dec(x_16); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_34 = !lean_is_exclusive(x_21); +if (x_34 == 0) +{ +return x_21; } else { -lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; 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_25 = lean_ctor_get(x_13, 0); -x_26 = lean_ctor_get(x_13, 1); -x_27 = lean_ctor_get(x_13, 2); -x_28 = lean_ctor_get(x_13, 3); -x_29 = lean_ctor_get(x_13, 4); -x_30 = lean_ctor_get(x_13, 5); -x_31 = lean_ctor_get_uint8(x_13, sizeof(void*)*14); -x_32 = lean_ctor_get(x_13, 6); -x_33 = lean_ctor_get(x_13, 7); -x_34 = lean_ctor_get(x_13, 8); -x_35 = lean_ctor_get(x_13, 9); -x_36 = lean_ctor_get(x_13, 10); -x_37 = lean_ctor_get(x_13, 11); -x_38 = lean_ctor_get(x_13, 12); -x_39 = lean_ctor_get(x_13, 13); -lean_inc(x_39); -lean_inc(x_38); -lean_inc(x_37); +lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_35 = lean_ctor_get(x_21, 0); +x_36 = lean_ctor_get(x_21, 1); lean_inc(x_36); lean_inc(x_35); -lean_inc(x_34); -lean_inc(x_33); -lean_inc(x_32); -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_dec(x_13); -x_40 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_setENode___spec__1(x_26, x_1, x_2); -x_41 = lean_alloc_ctor(0, 14, 1); -lean_ctor_set(x_41, 0, x_25); -lean_ctor_set(x_41, 1, x_40); -lean_ctor_set(x_41, 2, x_27); -lean_ctor_set(x_41, 3, x_28); -lean_ctor_set(x_41, 4, x_29); -lean_ctor_set(x_41, 5, x_30); -lean_ctor_set(x_41, 6, x_32); -lean_ctor_set(x_41, 7, x_33); -lean_ctor_set(x_41, 8, x_34); -lean_ctor_set(x_41, 9, x_35); -lean_ctor_set(x_41, 10, x_36); -lean_ctor_set(x_41, 11, x_37); -lean_ctor_set(x_41, 12, x_38); -lean_ctor_set(x_41, 13, x_39); -lean_ctor_set_uint8(x_41, sizeof(void*)*14, x_31); -x_42 = lean_st_ref_set(x_3, x_41, x_14); -x_43 = lean_ctor_get(x_42, 1); -lean_inc(x_43); -if (lean_is_exclusive(x_42)) { - lean_ctor_release(x_42, 0); - lean_ctor_release(x_42, 1); - x_44 = x_42; -} else { - lean_dec_ref(x_42); - x_44 = lean_box(0); +lean_dec(x_21); +x_37 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_37, 0, x_35); +lean_ctor_set(x_37, 1, x_36); +return x_37; +} } -x_45 = lean_box(0); -if (lean_is_scalar(x_44)) { - x_46 = lean_alloc_ctor(0, 2, 0); -} else { - x_46 = x_44; } -lean_ctor_set(x_46, 0, x_45); -lean_ctor_set(x_46, 1, x_43); -return x_46; +else +{ +uint8_t x_38; +lean_dec(x_16); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_38 = !lean_is_exclusive(x_17); +if (x_38 == 0) +{ +lean_object* x_39; lean_object* x_40; +x_39 = lean_ctor_get(x_17, 0); +lean_dec(x_39); +x_40 = lean_box(0); +lean_ctor_set(x_17, 0, x_40); +return x_17; +} +else +{ +lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_41 = lean_ctor_get(x_17, 1); +lean_inc(x_41); +lean_dec(x_17); +x_42 = lean_box(0); +x_43 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_43, 0, x_42); +lean_ctor_set(x_43, 1, x_41); +return x_43; } } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_setENode___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: +else { -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_Grind_setENode___spec__3(x_7, x_2, x_3, x_4, x_5, x_6); +uint8_t x_44; +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -return x_8; +lean_dec(x_1); +x_44 = !lean_is_exclusive(x_11); +if (x_44 == 0) +{ +return x_11; +} +else +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_45 = lean_ctor_get(x_11, 0); +x_46 = lean_ctor_get(x_11, 1); +lean_inc(x_46); +lean_inc(x_45); +lean_dec(x_11); +x_47 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_47, 0, x_45); +lean_ctor_set(x_47, 1, x_46); +return x_47; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_setENode___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +} +} +LEAN_EXPORT lean_object* l_Lean_MVarId_isAssigned___at_Lean_Meta_Grind_closeGoal___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: { -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_object* x_11; +x_11 = l_Lean_MVarId_isAssigned___at_Lean_Meta_Grind_closeGoal___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_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_8 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_setENode___spec__2(x_1, x_6, x_7, x_4, x_5); -return x_8; +lean_dec(x_2); +lean_dec(x_1); +return x_11; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_setENode___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_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_MVarId_assign___at_Lean_Meta_Grind_closeGoal___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { lean_object* x_12; -x_12 = l_Lean_Meta_Grind_setENode(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_MVarId_assign___at_Lean_Meta_Grind_closeGoal___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -15979,505 +21032,305 @@ lean_dec(x_3); return x_12; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkENodeCore(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, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at_Lean_Meta_Grind_getENodes___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _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; uint8_t x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; uint8_t x_32; -x_14 = lean_st_ref_get(x_5, x_13); -x_15 = lean_ctor_get(x_14, 0); -lean_inc(x_15); -x_16 = lean_ctor_get(x_14, 1); -lean_inc(x_16); -lean_dec(x_14); -x_17 = lean_st_ref_get(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_box(0); -x_21 = l_Lean_Expr_isLambda(x_1); -x_22 = lean_ctor_get(x_18, 7); -lean_inc(x_22); -lean_dec(x_18); -x_23 = lean_ctor_get(x_15, 6); -lean_inc(x_23); -lean_dec(x_15); -x_24 = 0; -x_25 = lean_unsigned_to_nat(1u); -lean_inc_n(x_1, 4); -x_26 = lean_alloc_ctor(0, 10, 5); -lean_ctor_set(x_26, 0, x_1); -lean_ctor_set(x_26, 1, x_1); -lean_ctor_set(x_26, 2, x_1); -lean_ctor_set(x_26, 3, x_1); -lean_ctor_set(x_26, 4, x_20); -lean_ctor_set(x_26, 5, x_20); -lean_ctor_set(x_26, 6, x_25); -lean_ctor_set(x_26, 7, x_22); -lean_ctor_set(x_26, 8, x_4); -lean_ctor_set(x_26, 9, x_23); -lean_ctor_set_uint8(x_26, sizeof(void*)*10, x_24); -lean_ctor_set_uint8(x_26, sizeof(void*)*10 + 1, x_2); -lean_ctor_set_uint8(x_26, sizeof(void*)*10 + 2, x_3); -lean_ctor_set_uint8(x_26, sizeof(void*)*10 + 3, x_21); -lean_ctor_set_uint8(x_26, sizeof(void*)*10 + 4, x_24); -x_27 = l_Lean_Meta_Grind_setENode(x_1, x_26, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_19); -x_28 = lean_ctor_get(x_27, 1); -lean_inc(x_28); -lean_dec(x_27); -x_29 = lean_st_ref_take(x_5, x_28); -x_30 = lean_ctor_get(x_29, 0); -lean_inc(x_30); -x_31 = lean_ctor_get(x_29, 1); -lean_inc(x_31); -lean_dec(x_29); -x_32 = !lean_is_exclusive(x_30); -if (x_32 == 0) -{ -lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; -x_33 = lean_ctor_get(x_30, 7); -x_34 = lean_nat_add(x_33, x_25); -lean_dec(x_33); -lean_ctor_set(x_30, 7, x_34); -x_35 = lean_st_ref_set(x_5, x_30, x_31); -x_36 = !lean_is_exclusive(x_35); -if (x_36 == 0) +lean_object* x_4; +x_4 = l_Lean_PersistentHashMap_foldlMAux___at_Lean_MetavarContext_getExprAssignmentDomain___spec__2___rarg(x_2, x_1, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Grind_getENodes___spec__1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -lean_object* x_37; lean_object* x_38; -x_37 = lean_ctor_get(x_35, 0); -lean_dec(x_37); -x_38 = lean_box(0); -lean_ctor_set(x_35, 0, x_38); -return x_35; +lean_object* x_4; lean_object* x_5; +x_4 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_4, 0, x_2); +lean_ctor_set(x_4, 1, x_3); +x_5 = lean_array_push(x_1, x_4); +return x_5; } -else +} +static lean_object* _init_l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Grind_getENodes___spec__1___closed__1() { +_start: { -lean_object* x_39; lean_object* x_40; lean_object* x_41; -x_39 = lean_ctor_get(x_35, 1); -lean_inc(x_39); -lean_dec(x_35); -x_40 = lean_box(0); -x_41 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_41, 0, x_40); -lean_ctor_set(x_41, 1, x_39); -return x_41; +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Grind_getENodes___spec__1___lambda__1), 3, 0); +return x_1; } } -else +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Grind_getENodes___spec__1(lean_object* x_1) { +_start: { -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; lean_object* x_62; lean_object* x_63; -x_42 = lean_ctor_get(x_30, 0); -x_43 = lean_ctor_get(x_30, 1); -x_44 = lean_ctor_get(x_30, 2); -x_45 = lean_ctor_get(x_30, 3); -x_46 = lean_ctor_get(x_30, 4); -x_47 = lean_ctor_get(x_30, 5); -x_48 = lean_ctor_get_uint8(x_30, sizeof(void*)*14); -x_49 = lean_ctor_get(x_30, 6); -x_50 = lean_ctor_get(x_30, 7); -x_51 = lean_ctor_get(x_30, 8); -x_52 = lean_ctor_get(x_30, 9); -x_53 = lean_ctor_get(x_30, 10); -x_54 = lean_ctor_get(x_30, 11); -x_55 = lean_ctor_get(x_30, 12); -x_56 = lean_ctor_get(x_30, 13); -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_inc(x_49); -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_30); -x_57 = lean_nat_add(x_50, x_25); -lean_dec(x_50); -x_58 = lean_alloc_ctor(0, 14, 1); -lean_ctor_set(x_58, 0, x_42); -lean_ctor_set(x_58, 1, x_43); -lean_ctor_set(x_58, 2, x_44); -lean_ctor_set(x_58, 3, x_45); -lean_ctor_set(x_58, 4, x_46); -lean_ctor_set(x_58, 5, x_47); -lean_ctor_set(x_58, 6, x_49); -lean_ctor_set(x_58, 7, x_57); -lean_ctor_set(x_58, 8, x_51); -lean_ctor_set(x_58, 9, x_52); -lean_ctor_set(x_58, 10, x_53); -lean_ctor_set(x_58, 11, x_54); -lean_ctor_set(x_58, 12, x_55); -lean_ctor_set(x_58, 13, x_56); -lean_ctor_set_uint8(x_58, sizeof(void*)*14, x_48); -x_59 = lean_st_ref_set(x_5, x_58, x_31); -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); +lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_2 = l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Grind_getENodes___spec__1___closed__1; +x_3 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2___closed__3; +x_4 = l_Lean_PersistentHashMap_foldlMAux___at_Lean_MetavarContext_getExprAssignmentDomain___spec__2___rarg(x_2, x_1, x_3); +return x_4; } -x_62 = lean_box(0); -if (lean_is_scalar(x_61)) { - x_63 = lean_alloc_ctor(0, 2, 0); -} else { - x_63 = x_61; } -lean_ctor_set(x_63, 0, x_62); -lean_ctor_set(x_63, 1, x_60); -return x_63; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_Grind_getENodes___spec__3(size_t x_1, size_t x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; +x_4 = lean_usize_dec_lt(x_2, x_1); +if (x_4 == 0) +{ +return x_3; +} +else +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x_9; size_t x_10; lean_object* x_11; +x_5 = lean_array_uget(x_3, x_2); +x_6 = lean_unsigned_to_nat(0u); +x_7 = lean_array_uset(x_3, x_2, x_6); +x_8 = lean_ctor_get(x_5, 1); +lean_inc(x_8); +lean_dec(x_5); +x_9 = 1; +x_10 = lean_usize_add(x_2, x_9); +x_11 = lean_array_uset(x_7, x_2, x_8); +x_2 = x_10; +x_3 = x_11; +goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkENodeCore___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_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 uint8_t l_Array_qsort_sort___at_Lean_Meta_Grind_getENodes___spec__4___lambda__1(lean_object* x_1, lean_object* x_2) { _start: { -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_3); -lean_dec(x_3); -x_16 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_14, x_15, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -return x_16; +lean_object* x_3; lean_object* x_4; uint8_t x_5; +x_3 = lean_ctor_get(x_1, 7); +x_4 = lean_ctor_get(x_2, 7); +x_5 = lean_nat_dec_lt(x_3, x_4); +return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkENode___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) { +static lean_object* _init_l_Array_qsort_sort___at_Lean_Meta_Grind_getENodes___spec__4___closed__1() { _start: { -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_13 = l_Lean_Meta_isConstructorAppCore_x3f(x_1, x_8, x_9, x_10, x_11, x_12); -x_14 = lean_ctor_get(x_13, 0); -lean_inc(x_14); -x_15 = lean_ctor_get(x_13, 1); -lean_inc(x_15); -lean_dec(x_13); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_1); -x_16 = l_Lean_Meta_Grind_isInterpreted(x_1, x_8, x_9, x_10, x_11, x_15); -if (lean_obj_tag(x_14) == 0) +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Array_qsort_sort___at_Lean_Meta_Grind_getENodes___spec__4___lambda__1___boxed), 2, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Meta_Grind_getENodes___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_16) == 0) +uint8_t x_7; +x_7 = lean_nat_dec_lt(x_3, x_4); +if (x_7 == 0) { -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_20 = lean_unbox(x_17); -lean_dec(x_17); -x_21 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_20, x_19, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_18); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -return x_21; +lean_dec(x_3); +return x_2; } else { -uint8_t x_22; -lean_dec(x_11); -lean_dec(x_10); +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; +x_8 = l_Array_qsort_sort___at_Lean_Meta_Grind_getENodes___spec__4___closed__1; +lean_inc(x_3); +x_9 = l___private_Init_Data_Array_QSort_0__Array_qpartition___rarg(x_1, x_2, x_8, x_3, x_4, lean_box(0), lean_box(0)); +x_10 = lean_ctor_get(x_9, 0); +lean_inc(x_10); +x_11 = lean_ctor_get(x_9, 1); +lean_inc(x_11); lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_2); -lean_dec(x_1); -x_22 = !lean_is_exclusive(x_16); -if (x_22 == 0) +x_12 = lean_nat_dec_le(x_4, x_10); +if (x_12 == 0) { -return x_16; +lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_13 = l_Array_qsort_sort___at_Lean_Meta_Grind_getENodes___spec__4(x_1, x_11, x_3, x_10, lean_box(0), lean_box(0)); +x_14 = lean_unsigned_to_nat(1u); +x_15 = lean_nat_add(x_10, x_14); +lean_dec(x_10); +x_2 = x_13; +x_3 = x_15; +x_5 = lean_box(0); +x_6 = lean_box(0); +goto _start; } else { -lean_object* x_23; lean_object* x_24; lean_object* x_25; -x_23 = lean_ctor_get(x_16, 0); -x_24 = lean_ctor_get(x_16, 1); -lean_inc(x_24); -lean_inc(x_23); -lean_dec(x_16); -x_25 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_25, 0, x_23); -lean_ctor_set(x_25, 1, x_24); -return x_25; +lean_dec(x_10); +lean_dec(x_3); +return x_11; } } } -else +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getENodes(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: { -lean_dec(x_14); -if (lean_obj_tag(x_16) == 0) +lean_object* x_10; uint8_t x_11; +x_10 = lean_st_ref_get(x_1, x_9); +x_11 = !lean_is_exclusive(x_10); +if (x_11 == 0) { -lean_object* x_26; lean_object* x_27; uint8_t x_28; uint8_t x_29; lean_object* x_30; -x_26 = lean_ctor_get(x_16, 0); -lean_inc(x_26); -x_27 = lean_ctor_get(x_16, 1); -lean_inc(x_27); -lean_dec(x_16); -x_28 = 1; -x_29 = lean_unbox(x_26); -lean_dec(x_26); -x_30 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_29, x_28, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_27); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -return x_30; -} -else +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; lean_object* x_21; uint8_t x_22; +x_12 = lean_ctor_get(x_10, 0); +x_13 = lean_ctor_get(x_12, 1); +lean_inc(x_13); +lean_dec(x_12); +x_14 = l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Grind_getENodes___spec__1(x_13); +lean_dec(x_13); +x_15 = lean_array_size(x_14); +x_16 = 0; +x_17 = l_Array_mapMUnsafe_map___at_Lean_Meta_Grind_getENodes___spec__3(x_15, x_16, x_14); +x_18 = lean_array_get_size(x_17); +x_19 = lean_unsigned_to_nat(1u); +x_20 = lean_nat_sub(x_18, x_19); +x_21 = lean_unsigned_to_nat(0u); +x_22 = lean_nat_dec_eq(x_18, x_21); +if (x_22 == 0) { -uint8_t x_31; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_2); -lean_dec(x_1); -x_31 = !lean_is_exclusive(x_16); -if (x_31 == 0) +uint8_t x_23; +x_23 = lean_nat_dec_le(x_21, x_20); +if (x_23 == 0) { -return x_16; +lean_object* x_24; +lean_inc(x_20); +x_24 = l_Array_qsort_sort___at_Lean_Meta_Grind_getENodes___spec__4(x_18, x_17, x_20, x_20, lean_box(0), lean_box(0)); +lean_dec(x_20); +lean_dec(x_18); +lean_ctor_set(x_10, 0, x_24); +return x_10; } else { -lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_32 = lean_ctor_get(x_16, 0); -x_33 = lean_ctor_get(x_16, 1); -lean_inc(x_33); -lean_inc(x_32); -lean_dec(x_16); -x_34 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_34, 0, x_32); -lean_ctor_set(x_34, 1, x_33); -return x_34; -} -} -} +lean_object* x_25; +x_25 = l_Array_qsort_sort___at_Lean_Meta_Grind_getENodes___spec__4(x_18, x_17, x_21, x_20, lean_box(0), lean_box(0)); +lean_dec(x_20); +lean_dec(x_18); +lean_ctor_set(x_10, 0, x_25); +return x_10; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkENode(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, 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 = l_Lean_Meta_Grind_alreadyInternalized(x_1, 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) +else { -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_Meta_Grind_mkENode___lambda__1(x_1, x_2, x_16, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_15); -return x_17; +lean_dec(x_20); +lean_dec(x_18); +lean_ctor_set(x_10, 0, x_17); +return x_10; +} } else { -uint8_t x_18; +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; size_t x_30; size_t x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; +x_26 = lean_ctor_get(x_10, 0); +x_27 = lean_ctor_get(x_10, 1); +lean_inc(x_27); +lean_inc(x_26); lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_2); -lean_dec(x_1); -x_18 = !lean_is_exclusive(x_12); -if (x_18 == 0) +x_28 = lean_ctor_get(x_26, 1); +lean_inc(x_28); +lean_dec(x_26); +x_29 = l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Grind_getENodes___spec__1(x_28); +lean_dec(x_28); +x_30 = lean_array_size(x_29); +x_31 = 0; +x_32 = l_Array_mapMUnsafe_map___at_Lean_Meta_Grind_getENodes___spec__3(x_30, x_31, x_29); +x_33 = lean_array_get_size(x_32); +x_34 = lean_unsigned_to_nat(1u); +x_35 = lean_nat_sub(x_33, x_34); +x_36 = lean_unsigned_to_nat(0u); +x_37 = lean_nat_dec_eq(x_33, x_36); +if (x_37 == 0) { -lean_object* x_19; lean_object* x_20; -x_19 = lean_ctor_get(x_12, 0); -lean_dec(x_19); -x_20 = lean_box(0); -lean_ctor_set(x_12, 0, x_20); -return x_12; +uint8_t x_38; +x_38 = lean_nat_dec_le(x_36, x_35); +if (x_38 == 0) +{ +lean_object* x_39; lean_object* x_40; +lean_inc(x_35); +x_39 = l_Array_qsort_sort___at_Lean_Meta_Grind_getENodes___spec__4(x_33, x_32, x_35, x_35, lean_box(0), lean_box(0)); +lean_dec(x_35); +lean_dec(x_33); +x_40 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_40, 0, x_39); +lean_ctor_set(x_40, 1, x_27); +return x_40; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_12, 1); -lean_inc(x_21); -lean_dec(x_12); -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_41; lean_object* x_42; +x_41 = l_Array_qsort_sort___at_Lean_Meta_Grind_getENodes___spec__4(x_33, x_32, x_36, x_35, lean_box(0), lean_box(0)); +lean_dec(x_35); +lean_dec(x_33); +x_42 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_27); +return x_42; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkENode___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: +else { -lean_object* x_13; -x_13 = l_Lean_Meta_Grind_mkENode___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_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -return x_13; +lean_object* x_43; +lean_dec(x_35); +lean_dec(x_33); +x_43 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_43, 0, x_32); +lean_ctor_set(x_43, 1, x_27); +return x_43; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkENode___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, 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_Meta_Grind_mkENode(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); -return x_12; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isCongrRoot(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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_PersistentHashMap_foldlM___at_Lean_Meta_Grind_getENodes___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_11; -lean_inc(x_1); -x_11 = l_Lean_Meta_Grind_getENode(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; uint8_t x_15; lean_object* x_16; -x_13 = lean_ctor_get(x_11, 0); -x_14 = lean_ctor_get(x_13, 3); -lean_inc(x_14); -lean_dec(x_13); -x_15 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_1, x_14); -lean_dec(x_14); +lean_object* x_4; +x_4 = l_Lean_PersistentHashMap_foldlM___at_Lean_Meta_Grind_getENodes___spec__2(x_1, x_2, x_3); lean_dec(x_1); -x_16 = lean_box(x_15); -lean_ctor_set(x_11, 0, x_16); -return x_11; +return x_4; } -else +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Grind_getENodes___spec__1___boxed(lean_object* x_1) { +_start: { -lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; lean_object* x_21; lean_object* x_22; -x_17 = lean_ctor_get(x_11, 0); -x_18 = lean_ctor_get(x_11, 1); -lean_inc(x_18); -lean_inc(x_17); -lean_dec(x_11); -x_19 = lean_ctor_get(x_17, 3); -lean_inc(x_19); -lean_dec(x_17); -x_20 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_1, x_19); -lean_dec(x_19); +lean_object* x_2; +x_2 = l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Grind_getENodes___spec__1(x_1); lean_dec(x_1); -x_21 = lean_box(x_20); -x_22 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_22, 0, x_21); -lean_ctor_set(x_22, 1, x_18); -return x_22; +return x_2; } } -else +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_Grind_getENodes___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -uint8_t x_23; +size_t x_4; size_t x_5; lean_object* x_6; +x_4 = lean_unbox_usize(x_1); lean_dec(x_1); -x_23 = !lean_is_exclusive(x_11); -if (x_23 == 0) -{ -return x_11; -} -else -{ -lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_24 = lean_ctor_get(x_11, 0); -x_25 = lean_ctor_get(x_11, 1); -lean_inc(x_25); -lean_inc(x_24); -lean_dec(x_11); -x_26 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_26, 0, x_24); -lean_ctor_set(x_26, 1, x_25); -return x_26; -} -} +x_5 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_6 = l_Array_mapMUnsafe_map___at_Lean_Meta_Grind_getENodes___spec__3(x_4, x_5, x_3); +return x_6; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isCongrRoot___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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_qsort_sort___at_Lean_Meta_Grind_getENodes___spec__4___lambda__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_11; -x_11 = l_Lean_Meta_Grind_isCongrRoot(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); +uint8_t x_3; lean_object* x_4; +x_3 = l_Array_qsort_sort___at_Lean_Meta_Grind_getENodes___spec__4___lambda__1(x_1, x_2); lean_dec(x_2); -return x_11; +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isInconsistent(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Meta_Grind_getENodes___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { -lean_object* x_10; uint8_t x_11; -x_10 = lean_st_ref_get(x_1, x_9); -x_11 = !lean_is_exclusive(x_10); -if (x_11 == 0) -{ -lean_object* x_12; uint8_t x_13; lean_object* x_14; -x_12 = lean_ctor_get(x_10, 0); -x_13 = lean_ctor_get_uint8(x_12, sizeof(void*)*14); -lean_dec(x_12); -x_14 = lean_box(x_13); -lean_ctor_set(x_10, 0, x_14); -return x_10; -} -else -{ -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_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_ctor_get_uint8(x_15, sizeof(void*)*14); -lean_dec(x_15); -x_18 = lean_box(x_17); -x_19 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_19, 0, x_18); -lean_ctor_set(x_19, 1, x_16); -return x_19; -} +lean_object* x_7; +x_7 = l_Array_qsort_sort___at_Lean_Meta_Grind_getENodes___spec__4(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_4); +lean_dec(x_1); +return x_7; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isInconsistent___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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_Grind_getENodes___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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_Grind_isInconsistent(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_10 = l_Lean_Meta_Grind_getENodes(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); @@ -16489,154 +21342,203 @@ lean_dec(x_1); return x_10; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEqProof___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_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_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_forEachENode___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, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { _start: { -lean_object* x_12; -x_12 = lean_grind_mk_eq_proof(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_12; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkHEqProof___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, 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_17; +x_17 = lean_usize_dec_lt(x_6, x_5); +if (x_17 == 0) { -lean_object* x_12; -x_12 = lean_grind_mk_heq_proof(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_12; -} +lean_object* x_18; +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_1); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_7); +lean_ctor_set(x_18, 1, x_16); +return x_18; } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEqHEqProof(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, 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_19; lean_object* x_20; +lean_dec(x_7); +x_19 = lean_array_uget(x_4, x_6); +lean_inc(x_1); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_2); -lean_inc(x_1); -x_12 = l_Lean_Meta_Grind_hasSameType(x_1, x_2, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_12) == 0) -{ -lean_object* x_13; uint8_t x_14; -x_13 = lean_ctor_get(x_12, 0); -lean_inc(x_13); -x_14 = lean_unbox(x_13); -lean_dec(x_13); -if (x_14 == 0) -{ -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_grind_mk_heq_proof(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_15); -return x_16; -} -else +x_20 = lean_apply_10(x_1, x_19, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +if (lean_obj_tag(x_20) == 0) { -lean_object* x_17; lean_object* x_18; -x_17 = lean_ctor_get(x_12, 1); -lean_inc(x_17); -lean_dec(x_12); -x_18 = lean_grind_mk_eq_proof(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_17); -return x_18; -} +lean_object* x_21; size_t x_22; size_t x_23; lean_object* x_24; +x_21 = lean_ctor_get(x_20, 1); +lean_inc(x_21); +lean_dec(x_20); +x_22 = 1; +x_23 = lean_usize_add(x_6, x_22); +x_24 = lean_box(0); +x_6 = x_23; +x_7 = x_24; +x_16 = x_21; +goto _start; } else { -uint8_t x_19; +uint8_t x_26; +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); lean_dec(x_1); -x_19 = !lean_is_exclusive(x_12); -if (x_19 == 0) +x_26 = !lean_is_exclusive(x_20); +if (x_26 == 0) { -return x_12; +return x_20; } 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_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_object* x_27; lean_object* x_28; lean_object* x_29; +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_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_Grind_mkEqTrueProof(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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_Grind_forEachENode(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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 = l_Lean_Meta_Grind_getTrueExpr___rarg(x_5, x_6, x_7, x_8, x_9, 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; +x_11 = l_Lean_Meta_Grind_getENodes(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); x_12 = lean_ctor_get(x_11, 0); lean_inc(x_12); x_13 = lean_ctor_get(x_11, 1); lean_inc(x_13); lean_dec(x_11); -x_14 = lean_grind_mk_eq_proof(x_1, x_12, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_13); -return x_14; +x_14 = lean_box(0); +x_15 = lean_array_size(x_12); +x_16 = 0; +x_17 = lean_box(0); +x_18 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_forEachENode___spec__1(x_1, x_12, x_14, x_12, x_15, x_16, x_17, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_13); +lean_dec(x_12); +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; +x_20 = lean_ctor_get(x_18, 0); +lean_dec(x_20); +lean_ctor_set(x_18, 0, x_17); +return x_18; } +else +{ +lean_object* x_21; lean_object* x_22; +x_21 = lean_ctor_get(x_18, 1); +lean_inc(x_21); +lean_dec(x_18); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_17); +lean_ctor_set(x_22, 1, x_21); +return x_22; } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEqFalseProof(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +} +else +{ +uint8_t x_23; +x_23 = !lean_is_exclusive(x_18); +if (x_23 == 0) +{ +return x_18; +} +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_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_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_forEachENode___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { _start: { -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_11 = l_Lean_Meta_Grind_getFalseExpr___rarg(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_grind_mk_eq_proof(x_1, x_12, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_13); -return x_14; +size_t x_17; size_t x_18; lean_object* x_19; +x_17 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_18 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_19 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_forEachENode___spec__1(x_1, x_2, x_3, x_4, x_17, x_18, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_19; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markAsInconsistent(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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_Grind_filterENodes___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; -x_10 = lean_st_ref_take(x_1, x_9); -x_11 = lean_ctor_get(x_10, 0); -lean_inc(x_11); -x_12 = lean_ctor_get(x_10, 1); -lean_inc(x_12); -lean_dec(x_10); -x_13 = !lean_is_exclusive(x_11); -if (x_13 == 0) +lean_object* x_13; +lean_inc(x_3); +x_13 = lean_apply_10(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_13) == 0) { -uint8_t x_14; lean_object* x_15; uint8_t x_16; -x_14 = 1; -lean_ctor_set_uint8(x_11, sizeof(void*)*14, x_14); -x_15 = lean_st_ref_set(x_1, x_11, x_12); -x_16 = !lean_is_exclusive(x_15); +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) +{ +uint8_t x_16; +lean_dec(x_3); +x_16 = !lean_is_exclusive(x_13); if (x_16 == 0) { lean_object* x_17; lean_object* x_18; -x_17 = lean_ctor_get(x_15, 0); +x_17 = lean_ctor_get(x_13, 0); lean_dec(x_17); x_18 = lean_box(0); -lean_ctor_set(x_15, 0, x_18); -return x_15; +lean_ctor_set(x_13, 0, x_18); +return x_13; } else { lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_19 = lean_ctor_get(x_15, 1); +x_19 = lean_ctor_get(x_13, 1); lean_inc(x_19); -lean_dec(x_15); +lean_dec(x_13); x_20 = lean_box(0); x_21 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_21, 0, x_20); @@ -16646,501 +21548,383 @@ return x_21; } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; 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; -x_22 = lean_ctor_get(x_11, 0); -x_23 = lean_ctor_get(x_11, 1); -x_24 = lean_ctor_get(x_11, 2); -x_25 = lean_ctor_get(x_11, 3); -x_26 = lean_ctor_get(x_11, 4); -x_27 = lean_ctor_get(x_11, 5); -x_28 = lean_ctor_get(x_11, 6); -x_29 = lean_ctor_get(x_11, 7); -x_30 = lean_ctor_get(x_11, 8); -x_31 = lean_ctor_get(x_11, 9); -x_32 = lean_ctor_get(x_11, 10); -x_33 = lean_ctor_get(x_11, 11); -x_34 = lean_ctor_get(x_11, 12); -x_35 = lean_ctor_get(x_11, 13); -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_inc(x_25); -lean_inc(x_24); -lean_inc(x_23); +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_11); -x_36 = 1; -x_37 = lean_alloc_ctor(0, 14, 1); -lean_ctor_set(x_37, 0, x_22); -lean_ctor_set(x_37, 1, x_23); -lean_ctor_set(x_37, 2, x_24); -lean_ctor_set(x_37, 3, x_25); -lean_ctor_set(x_37, 4, x_26); -lean_ctor_set(x_37, 5, x_27); -lean_ctor_set(x_37, 6, x_28); -lean_ctor_set(x_37, 7, x_29); -lean_ctor_set(x_37, 8, x_30); -lean_ctor_set(x_37, 9, x_31); -lean_ctor_set(x_37, 10, x_32); -lean_ctor_set(x_37, 11, x_33); -lean_ctor_set(x_37, 12, x_34); -lean_ctor_set(x_37, 13, x_35); -lean_ctor_set_uint8(x_37, sizeof(void*)*14, x_36); -x_38 = lean_st_ref_set(x_1, x_37, x_12); -x_39 = lean_ctor_get(x_38, 1); -lean_inc(x_39); -if (lean_is_exclusive(x_38)) { - lean_ctor_release(x_38, 0); - lean_ctor_release(x_38, 1); - x_40 = x_38; -} else { - lean_dec_ref(x_38); - x_40 = lean_box(0); -} -x_41 = lean_box(0); -if (lean_is_scalar(x_40)) { - x_42 = lean_alloc_ctor(0, 2, 0); -} else { - x_42 = x_40; +lean_dec(x_13); +x_23 = lean_st_ref_take(x_2, 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_array_push(x_24, x_3); +x_27 = lean_st_ref_set(x_2, x_26, x_25); +x_28 = !lean_is_exclusive(x_27); +if (x_28 == 0) +{ +return x_27; } -lean_ctor_set(x_42, 0, x_41); -lean_ctor_set(x_42, 1, x_39); -return x_42; +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_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_EXPORT lean_object* l_Lean_Meta_Grind_markAsInconsistent___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: +else { -lean_object* x_10; -x_10 = l_Lean_Meta_Grind_markAsInconsistent(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); +uint8_t x_32; lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_10; -} -} -LEAN_EXPORT lean_object* l_Lean_MVarId_isAssigned___at_Lean_Meta_Grind_closeGoal___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: +x_32 = !lean_is_exclusive(x_13); +if (x_32 == 0) { -lean_object* x_11; uint8_t x_12; -x_11 = lean_st_ref_get(x_7, x_10); -x_12 = !lean_is_exclusive(x_11); -if (x_12 == 0) +return x_13; +} +else { -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_11, 0); -x_14 = lean_ctor_get(x_13, 0); -lean_inc(x_14); +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_dec(x_13); -x_15 = lean_ctor_get(x_14, 7); -lean_inc(x_15); -lean_dec(x_14); -x_16 = l_Lean_PersistentHashMap_contains___at_Lean_MVarId_isAssigned___spec__1(x_15, x_1); -x_17 = lean_box(x_16); -lean_ctor_set(x_11, 0, x_17); -return x_11; +x_35 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_35, 0, x_33); +lean_ctor_set(x_35, 1, x_34); +return x_35; } -else -{ -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; lean_object* x_23; lean_object* x_24; -x_18 = lean_ctor_get(x_11, 0); -x_19 = lean_ctor_get(x_11, 1); -lean_inc(x_19); -lean_inc(x_18); -lean_dec(x_11); -x_20 = lean_ctor_get(x_18, 0); -lean_inc(x_20); -lean_dec(x_18); -x_21 = lean_ctor_get(x_20, 7); -lean_inc(x_21); -lean_dec(x_20); -x_22 = l_Lean_PersistentHashMap_contains___at_Lean_MVarId_isAssigned___spec__1(x_21, x_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_19); -return x_24; } } } -LEAN_EXPORT lean_object* l_Lean_MVarId_assign___at_Lean_Meta_Grind_closeGoal___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_Meta_Grind_filterENodes(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; -x_12 = lean_st_ref_take(x_8, x_11); +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_11 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2___closed__3; +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_13, 0); +x_14 = lean_ctor_get(x_12, 1); lean_inc(x_14); -x_15 = lean_ctor_get(x_12, 1); -lean_inc(x_15); lean_dec(x_12); -x_16 = !lean_is_exclusive(x_13); -if (x_16 == 0) -{ -lean_object* x_17; uint8_t x_18; -x_17 = lean_ctor_get(x_13, 0); -lean_dec(x_17); -x_18 = !lean_is_exclusive(x_14); -if (x_18 == 0) +lean_inc(x_13); +x_15 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_filterENodes___lambda__1___boxed), 12, 2); +lean_closure_set(x_15, 0, x_1); +lean_closure_set(x_15, 1, x_13); +x_16 = l_Lean_Meta_Grind_forEachENode(x_15, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_14); +if (lean_obj_tag(x_16) == 0) { -lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; -x_19 = lean_ctor_get(x_14, 7); -x_20 = l_Lean_PersistentHashMap_insert___at_Lean_MVarId_assign___spec__1(x_19, x_1, x_2); -lean_ctor_set(x_14, 7, x_20); -x_21 = lean_st_ref_set(x_8, x_13, x_15); -x_22 = !lean_is_exclusive(x_21); -if (x_22 == 0) +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 = lean_st_ref_get(x_13, x_17); +lean_dec(x_13); +x_19 = !lean_is_exclusive(x_18); +if (x_19 == 0) { -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; +return x_18; } 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 = lean_box(0); -x_27 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_27, 0, x_26); -lean_ctor_set(x_27, 1, x_25); -return x_27; +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_18, 0); +x_21 = lean_ctor_get(x_18, 1); +lean_inc(x_21); +lean_inc(x_20); +lean_dec(x_18); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_20); +lean_ctor_set(x_22, 1, x_21); +return x_22; } } else { -lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; -x_28 = lean_ctor_get(x_14, 0); -x_29 = lean_ctor_get(x_14, 1); -x_30 = lean_ctor_get(x_14, 2); -x_31 = lean_ctor_get(x_14, 3); -x_32 = lean_ctor_get(x_14, 4); -x_33 = lean_ctor_get(x_14, 5); -x_34 = lean_ctor_get(x_14, 6); -x_35 = lean_ctor_get(x_14, 7); -x_36 = lean_ctor_get(x_14, 8); -lean_inc(x_36); -lean_inc(x_35); -lean_inc(x_34); -lean_inc(x_33); -lean_inc(x_32); -lean_inc(x_31); -lean_inc(x_30); -lean_inc(x_29); -lean_inc(x_28); -lean_dec(x_14); -x_37 = l_Lean_PersistentHashMap_insert___at_Lean_MVarId_assign___spec__1(x_35, x_1, x_2); -x_38 = lean_alloc_ctor(0, 9, 0); -lean_ctor_set(x_38, 0, x_28); -lean_ctor_set(x_38, 1, x_29); -lean_ctor_set(x_38, 2, x_30); -lean_ctor_set(x_38, 3, x_31); -lean_ctor_set(x_38, 4, x_32); -lean_ctor_set(x_38, 5, x_33); -lean_ctor_set(x_38, 6, x_34); -lean_ctor_set(x_38, 7, x_37); -lean_ctor_set(x_38, 8, x_36); -lean_ctor_set(x_13, 0, x_38); -x_39 = lean_st_ref_set(x_8, x_13, x_15); -x_40 = lean_ctor_get(x_39, 1); -lean_inc(x_40); -if (lean_is_exclusive(x_39)) { - lean_ctor_release(x_39, 0); - lean_ctor_release(x_39, 1); - x_41 = x_39; -} else { - lean_dec_ref(x_39); - x_41 = lean_box(0); -} -x_42 = lean_box(0); -if (lean_is_scalar(x_41)) { - x_43 = lean_alloc_ctor(0, 2, 0); -} else { - x_43 = x_41; -} -lean_ctor_set(x_43, 0, x_42); -lean_ctor_set(x_43, 1, x_40); -return x_43; -} +uint8_t x_23; +lean_dec(x_13); +x_23 = !lean_is_exclusive(x_16); +if (x_23 == 0) +{ +return x_16; } else { -lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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_44 = lean_ctor_get(x_13, 1); -x_45 = lean_ctor_get(x_13, 2); -x_46 = lean_ctor_get(x_13, 3); -x_47 = lean_ctor_get(x_13, 4); -lean_inc(x_47); -lean_inc(x_46); -lean_inc(x_45); -lean_inc(x_44); -lean_dec(x_13); -x_48 = lean_ctor_get(x_14, 0); -lean_inc(x_48); -x_49 = lean_ctor_get(x_14, 1); -lean_inc(x_49); -x_50 = lean_ctor_get(x_14, 2); -lean_inc(x_50); -x_51 = lean_ctor_get(x_14, 3); -lean_inc(x_51); -x_52 = lean_ctor_get(x_14, 4); -lean_inc(x_52); -x_53 = lean_ctor_get(x_14, 5); -lean_inc(x_53); -x_54 = lean_ctor_get(x_14, 6); -lean_inc(x_54); -x_55 = lean_ctor_get(x_14, 7); -lean_inc(x_55); -x_56 = lean_ctor_get(x_14, 8); -lean_inc(x_56); -if (lean_is_exclusive(x_14)) { - lean_ctor_release(x_14, 0); - lean_ctor_release(x_14, 1); - lean_ctor_release(x_14, 2); - lean_ctor_release(x_14, 3); - lean_ctor_release(x_14, 4); - lean_ctor_release(x_14, 5); - lean_ctor_release(x_14, 6); - lean_ctor_release(x_14, 7); - lean_ctor_release(x_14, 8); - x_57 = x_14; -} else { - lean_dec_ref(x_14); - x_57 = lean_box(0); -} -x_58 = l_Lean_PersistentHashMap_insert___at_Lean_MVarId_assign___spec__1(x_55, x_1, x_2); -if (lean_is_scalar(x_57)) { - x_59 = lean_alloc_ctor(0, 9, 0); -} else { - x_59 = x_57; -} -lean_ctor_set(x_59, 0, x_48); -lean_ctor_set(x_59, 1, x_49); -lean_ctor_set(x_59, 2, x_50); -lean_ctor_set(x_59, 3, x_51); -lean_ctor_set(x_59, 4, x_52); -lean_ctor_set(x_59, 5, x_53); -lean_ctor_set(x_59, 6, x_54); -lean_ctor_set(x_59, 7, x_58); -lean_ctor_set(x_59, 8, x_56); -x_60 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_60, 0, x_59); -lean_ctor_set(x_60, 1, x_44); -lean_ctor_set(x_60, 2, x_45); -lean_ctor_set(x_60, 3, x_46); -lean_ctor_set(x_60, 4, x_47); -x_61 = lean_st_ref_set(x_8, x_60, x_15); -x_62 = lean_ctor_get(x_61, 1); -lean_inc(x_62); -if (lean_is_exclusive(x_61)) { - lean_ctor_release(x_61, 0); - lean_ctor_release(x_61, 1); - x_63 = x_61; -} else { - lean_dec_ref(x_61); - x_63 = lean_box(0); -} -x_64 = lean_box(0); -if (lean_is_scalar(x_63)) { - x_65 = lean_alloc_ctor(0, 2, 0); -} else { - x_65 = x_63; +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_16, 0); +x_25 = lean_ctor_get(x_16, 1); +lean_inc(x_25); +lean_inc(x_24); +lean_dec(x_16); +x_26 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_26, 0, x_24); +lean_ctor_set(x_26, 1, x_25); +return x_26; } -lean_ctor_set(x_65, 0, x_64); -lean_ctor_set(x_65, 1, x_62); -return x_65; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_closeGoal(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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_Grind_filterENodes___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; -x_11 = l_Lean_Meta_Grind_markAsInconsistent(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -x_12 = lean_ctor_get(x_11, 1); -lean_inc(x_12); -lean_dec(x_11); -x_13 = lean_st_ref_get(x_2, x_12); -x_14 = lean_ctor_get(x_13, 0); -lean_inc(x_14); -x_15 = lean_ctor_get(x_13, 1); -lean_inc(x_15); -lean_dec(x_13); -x_16 = lean_ctor_get(x_14, 0); -lean_inc(x_16); +lean_object* x_13; +x_13 = l_Lean_Meta_Grind_filterENodes___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_2); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_forEachEqc___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, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +_start: +{ +uint8_t x_17; +x_17 = lean_usize_dec_lt(x_6, x_5); +if (x_17 == 0) +{ +lean_object* x_18; +lean_dec(x_15); lean_dec(x_14); -x_17 = l_Lean_MVarId_isAssigned___at_Lean_Meta_Grind_closeGoal___spec__1(x_16, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_15); -x_18 = lean_ctor_get(x_17, 0); -lean_inc(x_18); -x_19 = lean_unbox(x_18); -lean_dec(x_18); -if (x_19 == 0) +lean_dec(x_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_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_7); +lean_ctor_set(x_18, 1, x_16); +return x_18; +} +else { -lean_object* x_20; lean_object* x_21; -x_20 = lean_ctor_get(x_17, 1); +lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; +lean_dec(x_7); +x_19 = lean_array_uget(x_4, x_6); +x_20 = lean_ctor_get(x_19, 0); lean_inc(x_20); -lean_dec(x_17); -lean_inc(x_16); -x_21 = l_Lean_MVarId_getType(x_16, x_6, x_7, x_8, x_9, x_20); -if (lean_obj_tag(x_21) == 0) -{ -lean_object* x_22; lean_object* x_23; uint8_t x_24; -x_22 = lean_ctor_get(x_21, 0); -lean_inc(x_22); -x_23 = lean_ctor_get(x_21, 1); -lean_inc(x_23); +x_21 = lean_ctor_get(x_19, 2); +lean_inc(x_21); +x_22 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_20, x_21); lean_dec(x_21); -lean_inc(x_22); -x_24 = l_Lean_Expr_isFalse(x_22); -if (x_24 == 0) +lean_dec(x_20); +if (x_22 == 0) { -lean_object* x_25; +size_t x_23; size_t x_24; lean_object* x_25; +lean_dec(x_19); +x_23 = 1; +x_24 = lean_usize_add(x_6, x_23); +x_25 = lean_box(0); +x_6 = x_24; +x_7 = x_25; +goto _start; +} +else +{ +lean_object* x_27; +lean_inc(x_1); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -x_25 = l_Lean_Meta_mkFalseElim(x_22, x_1, x_6, x_7, x_8, x_9, x_23); -if (lean_obj_tag(x_25) == 0) +x_27 = lean_apply_10(x_1, x_19, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +if (lean_obj_tag(x_27) == 0) { -lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_26 = lean_ctor_get(x_25, 0); -lean_inc(x_26); -x_27 = lean_ctor_get(x_25, 1); -lean_inc(x_27); -lean_dec(x_25); -x_28 = l_Lean_MVarId_assign___at_Lean_Meta_Grind_closeGoal___spec__2(x_16, x_26, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_27); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -return x_28; +lean_object* x_28; size_t x_29; size_t x_30; lean_object* x_31; +x_28 = lean_ctor_get(x_27, 1); +lean_inc(x_28); +lean_dec(x_27); +x_29 = 1; +x_30 = lean_usize_add(x_6, x_29); +x_31 = lean_box(0); +x_6 = x_30; +x_7 = x_31; +x_16 = x_28; +goto _start; } else { -uint8_t x_29; -lean_dec(x_16); +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_9); lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -x_29 = !lean_is_exclusive(x_25); -if (x_29 == 0) +lean_dec(x_1); +x_33 = !lean_is_exclusive(x_27); +if (x_33 == 0) { -return x_25; +return x_27; } else { -lean_object* x_30; lean_object* x_31; lean_object* x_32; -x_30 = lean_ctor_get(x_25, 0); -x_31 = lean_ctor_get(x_25, 1); -lean_inc(x_31); -lean_inc(x_30); -lean_dec(x_25); -x_32 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_32, 0, x_30); -lean_ctor_set(x_32, 1, x_31); -return x_32; +lean_object* x_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_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_EXPORT lean_object* l_Lean_Meta_Grind_forEachEqc(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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; size_t x_15; size_t x_16; lean_object* x_17; lean_object* x_18; +x_11 = l_Lean_Meta_Grind_getENodes(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +x_13 = lean_ctor_get(x_11, 1); +lean_inc(x_13); +lean_dec(x_11); +x_14 = lean_box(0); +x_15 = lean_array_size(x_12); +x_16 = 0; +x_17 = lean_box(0); +x_18 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_forEachEqc___spec__1(x_1, x_12, x_14, x_12, x_15, x_16, x_17, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_13); +lean_dec(x_12); +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; +x_20 = lean_ctor_get(x_18, 0); +lean_dec(x_20); +lean_ctor_set(x_18, 0, x_17); +return x_18; +} else { -lean_object* x_33; -lean_dec(x_22); -x_33 = l_Lean_MVarId_assign___at_Lean_Meta_Grind_closeGoal___spec__2(x_16, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_23); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -return x_33; +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_17); +lean_ctor_set(x_22, 1, x_21); +return x_22; } } else { -uint8_t x_34; -lean_dec(x_16); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_1); -x_34 = !lean_is_exclusive(x_21); -if (x_34 == 0) +uint8_t x_23; +x_23 = !lean_is_exclusive(x_18); +if (x_23 == 0) { -return x_21; +return x_18; } else { -lean_object* x_35; lean_object* x_36; lean_object* x_37; -x_35 = lean_ctor_get(x_21, 0); -x_36 = lean_ctor_get(x_21, 1); -lean_inc(x_36); -lean_inc(x_35); -lean_dec(x_21); -x_37 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_37, 0, x_35); -lean_ctor_set(x_37, 1, x_36); -return x_37; +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_18, 0); +x_25 = lean_ctor_get(x_18, 1); +lean_inc(x_25); +lean_inc(x_24); +lean_dec(x_18); +x_26 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_26, 0, x_24); +lean_ctor_set(x_26, 1, x_25); +return x_26; } } } -else +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_forEachEqc___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +_start: { -uint8_t x_38; -lean_dec(x_16); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); +size_t x_17; size_t x_18; lean_object* x_19; +x_17 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_18 = lean_unbox_usize(x_6); lean_dec(x_6); -lean_dec(x_1); -x_38 = !lean_is_exclusive(x_17); -if (x_38 == 0) +x_19 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_forEachEqc___spec__1(x_1, x_2, x_3, x_4, x_17, x_18, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_19; +} +} +LEAN_EXPORT lean_object* l_ReaderT_pure___at_Lean_Meta_Grind_instInhabitedMethods___spec__1___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_1); +lean_ctor_set(x_11, 1, x_10); +return x_11; +} +} +LEAN_EXPORT lean_object* l_ReaderT_pure___at_Lean_Meta_Grind_instInhabitedMethods___spec__1(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_ReaderT_pure___at_Lean_Meta_Grind_instInhabitedMethods___spec__1___rarg___boxed), 10, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instInhabitedMethods___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_39; lean_object* x_40; -x_39 = lean_ctor_get(x_17, 0); -lean_dec(x_39); -x_40 = lean_box(0); -lean_ctor_set(x_17, 0, x_40); -return x_17; +lean_object* x_11; lean_object* x_12; +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 +} +static lean_object* _init_l_Lean_Meta_Grind_instInhabitedMethods___closed__1() { +_start: { -lean_object* x_41; lean_object* x_42; lean_object* x_43; -x_41 = lean_ctor_get(x_17, 1); -lean_inc(x_41); -lean_dec(x_17); -x_42 = lean_box(0); -x_43 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_43, 0, x_42); -lean_ctor_set(x_43, 1, x_41); -return x_43; +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_instInhabitedMethods___lambda__1___boxed), 10, 0); +return x_1; } } +static lean_object* _init_l_Lean_Meta_Grind_instInhabitedMethods() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = lean_box(0); +x_2 = lean_alloc_closure((void*)(l_ReaderT_pure___at_Lean_Meta_Grind_instInhabitedMethods___spec__1___rarg___boxed), 10, 1); +lean_closure_set(x_2, 0, x_1); +x_3 = l_Lean_Meta_Grind_instInhabitedMethods___closed__1; +x_4 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_3); +lean_ctor_set(x_4, 2, x_2); +return x_4; } } -LEAN_EXPORT lean_object* l_Lean_MVarId_isAssigned___at_Lean_Meta_Grind_closeGoal___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_ReaderT_pure___at_Lean_Meta_Grind_instInhabitedMethods___spec__1___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { lean_object* x_11; -x_11 = l_Lean_MVarId_isAssigned___at_Lean_Meta_Grind_closeGoal___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_11 = l_ReaderT_pure___at_Lean_Meta_Grind_instInhabitedMethods___spec__1___rarg(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); @@ -17149,16 +21933,14 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); return x_11; } } -LEAN_EXPORT lean_object* l_Lean_MVarId_assign___at_Lean_Meta_Grind_closeGoal___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instInhabitedMethods___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_12; -x_12 = l_Lean_MVarId_assign___at_Lean_Meta_Grind_closeGoal___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -lean_dec(x_10); +lean_object* x_11; +x_11 = l_Lean_Meta_Grind_instInhabitedMethods___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -17166,329 +21948,435 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -return x_12; +lean_dec(x_2); +lean_dec(x_1); +return x_11; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_closeGoal___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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_Grind_Methods_toMethodsRef_unsafe__1(lean_object* x_1) { _start: { -lean_object* x_11; -x_11 = l_Lean_Meta_Grind_closeGoal(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); -lean_dec(x_2); -return x_11; +lean_inc(x_1); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at_Lean_Meta_Grind_getENodes___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Methods_toMethodsRef_unsafe__1___boxed(lean_object* x_1) { _start: { -lean_object* x_4; -x_4 = l_Lean_PersistentHashMap_foldlMAux___at_Lean_MetavarContext_getExprAssignmentDomain___spec__2___rarg(x_2, x_1, x_3); -return x_4; +lean_object* x_2; +x_2 = l_Lean_Meta_Grind_Methods_toMethodsRef_unsafe__1(x_1); +lean_dec(x_1); +return x_2; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Grind_getENodes___spec__1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Methods_toMethodsRef(lean_object* x_1) { _start: { -lean_object* x_4; lean_object* x_5; -x_4 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_4, 0, x_2); -lean_ctor_set(x_4, 1, x_3); -x_5 = lean_array_push(x_1, x_4); -return x_5; +lean_inc(x_1); +return x_1; } } -static lean_object* _init_l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Grind_getENodes___spec__1___closed__1() { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Methods_toMethodsRef___boxed(lean_object* x_1) { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = lean_box(0); -x_2 = lean_array_mk(x_1); +lean_object* x_2; +x_2 = l_Lean_Meta_Grind_Methods_toMethodsRef(x_1); +lean_dec(x_1); return x_2; } } -static lean_object* _init_l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Grind_getENodes___spec__1___closed__2() { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_MethodsRef_toMethods_unsafe__1(lean_object* x_1) { _start: { -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Grind_getENodes___spec__1___lambda__1), 3, 0); +lean_inc(x_1); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Grind_getENodes___spec__1(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_MethodsRef_toMethods_unsafe__1___boxed(lean_object* x_1) { _start: { -lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_2 = l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Grind_getENodes___spec__1___closed__2; -x_3 = l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Grind_getENodes___spec__1___closed__1; -x_4 = l_Lean_PersistentHashMap_foldlMAux___at_Lean_MetavarContext_getExprAssignmentDomain___spec__2___rarg(x_2, x_1, x_3); -return x_4; +lean_object* x_2; +x_2 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_MethodsRef_toMethods_unsafe__1(x_1); +lean_dec(x_1); +return x_2; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_Grind_getENodes___spec__3(size_t x_1, size_t x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_MethodsRef_toMethods(lean_object* x_1) { _start: { -uint8_t x_4; -x_4 = lean_usize_dec_lt(x_2, x_1); -if (x_4 == 0) +lean_inc(x_1); +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_MethodsRef_toMethods___boxed(lean_object* x_1) { +_start: { -return x_3; +lean_object* x_2; +x_2 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_MethodsRef_toMethods(x_1); +lean_dec(x_1); +return x_2; } -else +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getMethods(lean_object* x_1, lean_object* x_2, lean_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_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x_9; size_t x_10; lean_object* x_11; -x_5 = lean_array_uget(x_3, x_2); -x_6 = lean_unsigned_to_nat(0u); -x_7 = lean_array_uset(x_3, x_2, x_6); -x_8 = lean_ctor_get(x_5, 1); -lean_inc(x_8); +lean_object* x_9; +x_9 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_9, 0, x_1); +lean_ctor_set(x_9, 1, x_8); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getMethods___boxed(lean_object* x_1, lean_object* x_2, lean_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_Grind_getMethods(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); -x_9 = 1; -x_10 = lean_usize_add(x_2, x_9); -x_11 = lean_array_uset(x_7, x_2, x_8); -x_2 = x_10; -x_3 = x_11; +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateUp(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; lean_object* x_12; +x_11 = lean_ctor_get(x_3, 0); +lean_inc(x_11); +x_12 = lean_apply_10(x_11, x_1, 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_Meta_Grind_propagateDown(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; lean_object* x_12; +x_11 = lean_ctor_get(x_3, 1); +lean_inc(x_11); +x_12 = lean_apply_10(x_11, x_1, 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_Meta_Grind_applyFallback(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; lean_object* x_11; +x_10 = lean_ctor_get(x_2, 2); +lean_inc(x_10); +x_11 = lean_apply_9(x_10, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getEqc_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: +{ +lean_object* x_13; +lean_inc(x_2); +x_13 = l_Lean_Meta_Grind_getNext(x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_13) == 0) +{ +uint8_t x_14; +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_15 = lean_ctor_get(x_13, 0); +x_16 = lean_ctor_get(x_13, 1); +x_17 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_17, 0, x_2); +lean_ctor_set(x_17, 1, x_3); +x_18 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_1, x_15); +if (x_18 == 0) +{ +lean_free_object(x_13); +x_2 = x_15; +x_3 = x_17; +x_12 = x_16; +goto _start; +} +else +{ +lean_dec(x_15); +lean_ctor_set(x_13, 0, x_17); +return x_13; +} +} +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; +x_20 = lean_ctor_get(x_13, 0); +x_21 = lean_ctor_get(x_13, 1); +lean_inc(x_21); +lean_inc(x_20); +lean_dec(x_13); +x_22 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_22, 0, x_2); +lean_ctor_set(x_22, 1, x_3); +x_23 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_1, x_20); +if (x_23 == 0) +{ +x_2 = x_20; +x_3 = x_22; +x_12 = x_21; goto _start; } +else +{ +lean_object* x_25; +lean_dec(x_20); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_22); +lean_ctor_set(x_25, 1, x_21); +return x_25; +} +} +} +else +{ +uint8_t x_26; +lean_dec(x_3); +lean_dec(x_2); +x_26 = !lean_is_exclusive(x_13); +if (x_26 == 0) +{ +return x_13; +} +else +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_13, 0); +x_28 = lean_ctor_get(x_13, 1); +lean_inc(x_28); +lean_inc(x_27); +lean_dec(x_13); +x_29 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_29, 0, x_27); +lean_ctor_set(x_29, 1, x_28); +return x_29; +} } } -LEAN_EXPORT uint8_t l_Array_qsort_sort___at_Lean_Meta_Grind_getENodes___spec__4___lambda__1(lean_object* x_1, lean_object* x_2) { +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getEqc_go___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -lean_object* x_3; lean_object* x_4; uint8_t x_5; -x_3 = lean_ctor_get(x_1, 7); -x_4 = lean_ctor_get(x_2, 7); -x_5 = lean_nat_dec_lt(x_3, x_4); -return x_5; +lean_object* x_13; +x_13 = l_Lean_Meta_Grind_getEqc_go(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +return x_13; } } -static lean_object* _init_l_Array_qsort_sort___at_Lean_Meta_Grind_getENodes___spec__4___closed__1() { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getEqc(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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_alloc_closure((void*)(l_Array_qsort_sort___at_Lean_Meta_Grind_getENodes___spec__4___lambda__1___boxed), 2, 0); -return x_1; +lean_object* x_11; lean_object* x_12; +x_11 = lean_box(0); +lean_inc(x_1); +x_12 = l_Lean_Meta_Grind_getEqc_go(x_1, x_1, x_11, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_1); +return x_12; } } -LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Meta_Grind_getENodes___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getEqc___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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_7; -x_7 = lean_nat_dec_lt(x_3, x_4); -if (x_7 == 0) -{ -lean_dec(x_3); -return x_2; -} -else -{ -lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_8 = l_Array_qsort_sort___at_Lean_Meta_Grind_getENodes___spec__4___closed__1; -lean_inc(x_3); -x_9 = l___private_Init_Data_Array_QSort_0__Array_qpartition___rarg(x_1, x_2, x_8, x_3, x_4, lean_box(0), lean_box(0)); -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_11; +x_11 = l_Lean_Meta_Grind_getEqc(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); lean_dec(x_9); -x_12 = lean_nat_dec_le(x_4, x_10); -if (x_12 == 0) -{ -lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_13 = l_Array_qsort_sort___at_Lean_Meta_Grind_getENodes___spec__4(x_1, x_11, x_3, x_10, lean_box(0), lean_box(0)); -x_14 = lean_unsigned_to_nat(1u); -x_15 = lean_nat_add(x_10, x_14); -lean_dec(x_10); -x_2 = x_13; -x_3 = x_15; -x_5 = lean_box(0); -x_6 = lean_box(0); -goto _start; -} -else -{ -lean_dec(x_10); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); lean_dec(x_3); +lean_dec(x_2); return x_11; } } -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getENodes(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_getEqcs___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, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { _start: { -lean_object* x_10; uint8_t x_11; -x_10 = lean_st_ref_get(x_1, x_9); -x_11 = !lean_is_exclusive(x_10); -if (x_11 == 0) -{ -lean_object* x_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; lean_object* x_21; uint8_t x_22; -x_12 = lean_ctor_get(x_10, 0); -x_13 = lean_ctor_get(x_12, 1); -lean_inc(x_13); -lean_dec(x_12); -x_14 = l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Grind_getENodes___spec__1(x_13); -lean_dec(x_13); -x_15 = lean_array_size(x_14); -x_16 = 0; -x_17 = l_Array_mapMUnsafe_map___at_Lean_Meta_Grind_getENodes___spec__3(x_15, x_16, x_14); -x_18 = lean_array_get_size(x_17); -x_19 = lean_unsigned_to_nat(1u); -x_20 = lean_nat_sub(x_18, x_19); -x_21 = lean_unsigned_to_nat(0u); -x_22 = lean_nat_dec_eq(x_18, x_21); -if (x_22 == 0) -{ -uint8_t x_23; -x_23 = lean_nat_dec_le(x_21, x_20); -if (x_23 == 0) +uint8_t x_16; +x_16 = lean_usize_dec_lt(x_5, x_4); +if (x_16 == 0) { -lean_object* x_24; -lean_inc(x_20); -x_24 = l_Array_qsort_sort___at_Lean_Meta_Grind_getENodes___spec__4(x_18, x_17, x_20, x_20, lean_box(0), lean_box(0)); -lean_dec(x_20); -lean_dec(x_18); -lean_ctor_set(x_10, 0, x_24); -return x_10; +lean_object* x_17; +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_6); +lean_ctor_set(x_17, 1, x_15); +return x_17; } else { -lean_object* x_25; -x_25 = l_Array_qsort_sort___at_Lean_Meta_Grind_getENodes___spec__4(x_18, x_17, x_21, x_20, lean_box(0), lean_box(0)); -lean_dec(x_20); +lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_18 = lean_array_uget(x_3, x_5); +x_19 = lean_ctor_get(x_18, 2); +lean_inc(x_19); +x_20 = lean_ctor_get(x_18, 0); +lean_inc(x_20); lean_dec(x_18); -lean_ctor_set(x_10, 0, x_25); -return x_10; -} -} -else +x_21 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_19, x_20); +lean_dec(x_19); +if (x_21 == 0) { +size_t x_22; size_t x_23; lean_dec(x_20); -lean_dec(x_18); -lean_ctor_set(x_10, 0, x_17); -return x_10; -} +x_22 = 1; +x_23 = lean_usize_add(x_5, x_22); +x_5 = x_23; +goto _start; } else { -lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; size_t x_30; size_t x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; -x_26 = lean_ctor_get(x_10, 0); -x_27 = lean_ctor_get(x_10, 1); -lean_inc(x_27); +lean_object* x_25; +x_25 = l_Lean_Meta_Grind_getEqc(x_20, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +if (lean_obj_tag(x_25) == 0) +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; size_t x_29; size_t x_30; +x_26 = lean_ctor_get(x_25, 0); lean_inc(x_26); -lean_dec(x_10); -x_28 = lean_ctor_get(x_26, 1); -lean_inc(x_28); -lean_dec(x_26); -x_29 = l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Grind_getENodes___spec__1(x_28); -lean_dec(x_28); -x_30 = lean_array_size(x_29); -x_31 = 0; -x_32 = l_Array_mapMUnsafe_map___at_Lean_Meta_Grind_getENodes___spec__3(x_30, x_31, x_29); -x_33 = lean_array_get_size(x_32); -x_34 = lean_unsigned_to_nat(1u); -x_35 = lean_nat_sub(x_33, x_34); -x_36 = lean_unsigned_to_nat(0u); -x_37 = lean_nat_dec_eq(x_33, x_36); -if (x_37 == 0) +x_27 = lean_ctor_get(x_25, 1); +lean_inc(x_27); +lean_dec(x_25); +x_28 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_28, 0, x_26); +lean_ctor_set(x_28, 1, x_6); +x_29 = 1; +x_30 = lean_usize_add(x_5, x_29); +x_5 = x_30; +x_6 = x_28; +x_15 = x_27; +goto _start; +} +else { -uint8_t x_38; -x_38 = lean_nat_dec_le(x_36, x_35); -if (x_38 == 0) +uint8_t x_32; +lean_dec(x_6); +x_32 = !lean_is_exclusive(x_25); +if (x_32 == 0) { -lean_object* x_39; lean_object* x_40; -lean_inc(x_35); -x_39 = l_Array_qsort_sort___at_Lean_Meta_Grind_getENodes___spec__4(x_33, x_32, x_35, x_35, lean_box(0), lean_box(0)); -lean_dec(x_35); -lean_dec(x_33); -x_40 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_40, 0, x_39); -lean_ctor_set(x_40, 1, x_27); -return x_40; +return x_25; } else { -lean_object* x_41; lean_object* x_42; -x_41 = l_Array_qsort_sort___at_Lean_Meta_Grind_getENodes___spec__4(x_33, x_32, x_36, x_35, lean_box(0), lean_box(0)); -lean_dec(x_35); -lean_dec(x_33); -x_42 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_42, 0, x_41); -lean_ctor_set(x_42, 1, x_27); -return x_42; +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_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_EXPORT lean_object* l_Lean_Meta_Grind_getEqcs(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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; size_t x_15; size_t x_16; lean_object* x_17; +x_10 = lean_box(0); +x_11 = l_Lean_Meta_Grind_getENodes(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +x_13 = lean_ctor_get(x_11, 1); +lean_inc(x_13); +lean_dec(x_11); +x_14 = lean_box(0); +x_15 = lean_array_size(x_12); +x_16 = 0; +x_17 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_getEqcs___spec__1(x_12, x_14, x_12, x_15, x_16, x_10, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_13); +lean_dec(x_12); +if (lean_obj_tag(x_17) == 0) +{ +uint8_t x_18; +x_18 = !lean_is_exclusive(x_17); +if (x_18 == 0) +{ +return x_17; +} else { -lean_object* x_43; -lean_dec(x_35); -lean_dec(x_33); -x_43 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_43, 0, x_32); -lean_ctor_set(x_43, 1, x_27); -return x_43; -} -} +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_17, 0); +x_20 = lean_ctor_get(x_17, 1); +lean_inc(x_20); +lean_inc(x_19); +lean_dec(x_17); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_19); +lean_ctor_set(x_21, 1, x_20); +return x_21; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at_Lean_Meta_Grind_getENodes___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: +else { -lean_object* x_4; -x_4 = l_Lean_PersistentHashMap_foldlM___at_Lean_Meta_Grind_getENodes___spec__2(x_1, x_2, x_3); -lean_dec(x_1); -return x_4; -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Grind_getENodes___spec__1___boxed(lean_object* x_1) { -_start: +uint8_t x_22; +x_22 = !lean_is_exclusive(x_17); +if (x_22 == 0) { -lean_object* x_2; -x_2 = l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Grind_getENodes___spec__1(x_1); -lean_dec(x_1); -return x_2; -} +return x_17; } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_Grind_getENodes___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: +else { -size_t x_4; size_t x_5; lean_object* x_6; -x_4 = lean_unbox_usize(x_1); -lean_dec(x_1); -x_5 = lean_unbox_usize(x_2); -lean_dec(x_2); -x_6 = l_Array_mapMUnsafe_map___at_Lean_Meta_Grind_getENodes___spec__3(x_4, x_5, x_3); -return x_6; +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_17, 0); +x_24 = lean_ctor_get(x_17, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_17); +x_25 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_24); +return x_25; } } -LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Meta_Grind_getENodes___spec__4___lambda__1___boxed(lean_object* x_1, lean_object* x_2) { -_start: -{ -uint8_t x_3; lean_object* x_4; -x_3 = l_Array_qsort_sort___at_Lean_Meta_Grind_getENodes___spec__4___lambda__1(x_1, x_2); -lean_dec(x_2); -lean_dec(x_1); -x_4 = lean_box(x_3); -return x_4; } } -LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Meta_Grind_getENodes___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_getEqcs___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { _start: { -lean_object* x_7; -x_7 = l_Array_qsort_sort___at_Lean_Meta_Grind_getENodes___spec__4(x_1, x_2, x_3, x_4, x_5, x_6); +size_t x_16; size_t x_17; lean_object* x_18; +x_16 = lean_unbox_usize(x_4); lean_dec(x_4); +x_17 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_18 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_getEqcs___spec__1(x_1, x_2, x_3, x_16, x_17, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +lean_dec(x_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_3); +lean_dec(x_2); lean_dec(x_1); -return x_7; +return x_18; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getENodes___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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_Grind_getEqcs___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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_Grind_getENodes(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_10 = l_Lean_Meta_Grind_getEqcs(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); @@ -17500,1001 +22388,1096 @@ lean_dec(x_1); return x_10; } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_forEachENode___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, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_isResolvedCaseSplit___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_17; -x_17 = lean_usize_dec_lt(x_6, x_5); -if (x_17 == 0) -{ -lean_object* x_18; -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_1); -x_18 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_18, 0, x_7); -lean_ctor_set(x_18, 1, x_16); -return x_18; -} -else -{ -lean_object* x_19; lean_object* x_20; -lean_dec(x_7); -x_19 = lean_array_uget(x_4, x_6); -lean_inc(x_1); -lean_inc(x_15); -lean_inc(x_14); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -x_20 = lean_apply_10(x_1, x_19, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); -if (lean_obj_tag(x_20) == 0) +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_21; size_t x_22; size_t x_23; lean_object* x_24; -x_21 = lean_ctor_get(x_20, 1); -lean_inc(x_21); -lean_dec(x_20); -x_22 = 1; -x_23 = lean_usize_add(x_6, x_22); -x_24 = lean_box(0); -x_6 = x_23; -x_7 = x_24; -x_16 = x_21; -goto _start; +uint8_t x_8; +lean_dec(x_4); +x_8 = 0; +return x_8; } else { -uint8_t x_26; -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_object* x_9; uint8_t x_10; +x_9 = lean_array_fget(x_1, x_4); +x_10 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_5, x_9); lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_1); -x_26 = !lean_is_exclusive(x_20); -if (x_26 == 0) +if (x_10 == 0) { -return x_20; +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_27; lean_object* x_28; lean_object* x_29; -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_alloc_ctor(1, 2, 0); -lean_ctor_set(x_29, 0, x_27); -lean_ctor_set(x_29, 1, x_28); -return x_29; -} +uint8_t x_14; +lean_dec(x_4); +x_14 = 1; +return x_14; } } } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_forEachENode(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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 uint8_t l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_isResolvedCaseSplit___spec__2(lean_object* x_1, size_t x_2, lean_object* x_3) { _start: { -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; -x_11 = l_Lean_Meta_Grind_getENodes(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); -x_13 = lean_ctor_get(x_11, 1); -lean_inc(x_13); -lean_dec(x_11); -x_14 = lean_box(0); -x_15 = lean_array_size(x_12); -x_16 = 0; -x_17 = lean_box(0); -x_18 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_forEachENode___spec__1(x_1, x_12, x_14, x_12, x_15, x_16, x_17, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_13); -lean_dec(x_12); -if (lean_obj_tag(x_18) == 0) +if (lean_obj_tag(x_1) == 0) { -uint8_t x_19; -x_19 = !lean_is_exclusive(x_18); -if (x_19 == 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_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___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_20; -x_20 = lean_ctor_get(x_18, 0); -lean_dec(x_20); -lean_ctor_set(x_18, 0, x_17); -return x_18; +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 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_3, x_11); +lean_dec(x_11); +return x_12; } -else +case 1: { -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_17); -lean_ctor_set(x_22, 1, x_21); -return x_22; +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 -{ -uint8_t x_23; -x_23 = !lean_is_exclusive(x_18); -if (x_23 == 0) -{ -return x_18; } 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_inc(x_25); -lean_inc(x_24); +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_Meta_Grind_isResolvedCaseSplit___spec__3(x_17, x_18, lean_box(0), x_19, x_3); 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_17); +return x_20; } } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_forEachENode___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_isResolvedCaseSplit___spec__1(lean_object* x_1, lean_object* x_2) { _start: { -size_t x_17; size_t x_18; lean_object* x_19; -x_17 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_18 = lean_unbox_usize(x_6); -lean_dec(x_6); -x_19 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_forEachENode___spec__1(x_1, x_2, x_3, x_4, x_17, x_18, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_19; +uint64_t x_3; size_t x_4; uint8_t x_5; +x_3 = l_Lean_Meta_Grind_instHashableENodeKey_unsafe__1(x_2); +x_4 = lean_uint64_to_usize(x_3); +x_5 = l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_isResolvedCaseSplit___spec__2(x_1, x_4, x_2); +return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_filterENodes___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_Meta_Grind_isResolvedCaseSplit(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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_inc(x_3); -x_13 = lean_apply_10(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_13) == 0) +lean_object* x_11; uint8_t x_12; +x_11 = lean_st_ref_get(x_2, x_10); +x_12 = !lean_is_exclusive(x_11); +if (x_12 == 0) { -lean_object* x_14; uint8_t x_15; -x_14 = lean_ctor_get(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_11, 0); +x_14 = lean_ctor_get(x_13, 18); lean_inc(x_14); -x_15 = lean_unbox(x_14); -lean_dec(x_14); -if (x_15 == 0) -{ -uint8_t x_16; -lean_dec(x_3); -x_16 = !lean_is_exclusive(x_13); -if (x_16 == 0) -{ -lean_object* x_17; lean_object* x_18; -x_17 = lean_ctor_get(x_13, 0); -lean_dec(x_17); -x_18 = lean_box(0); -lean_ctor_set(x_13, 0, x_18); -return x_13; +lean_dec(x_13); +x_15 = l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_isResolvedCaseSplit___spec__1(x_14, x_1); +x_16 = lean_box(x_15); +lean_ctor_set(x_11, 0, x_16); +return x_11; } else { -lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_19 = lean_ctor_get(x_13, 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; +x_17 = lean_ctor_get(x_11, 0); +x_18 = lean_ctor_get(x_11, 1); +lean_inc(x_18); +lean_inc(x_17); +lean_dec(x_11); +x_19 = lean_ctor_get(x_17, 18); lean_inc(x_19); -lean_dec(x_13); -x_20 = lean_box(0); -x_21 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_21, 0, x_20); -lean_ctor_set(x_21, 1, x_19); -return x_21; +lean_dec(x_17); +x_20 = l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_isResolvedCaseSplit___spec__1(x_19, x_1); +x_21 = lean_box(x_20); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_18); +return x_22; } } -else -{ -lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; -x_22 = lean_ctor_get(x_13, 1); -lean_inc(x_22); -lean_dec(x_13); -x_23 = lean_st_ref_take(x_2, 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_array_push(x_24, x_3); -x_27 = lean_st_ref_set(x_2, x_26, x_25); -x_28 = !lean_is_exclusive(x_27); -if (x_28 == 0) +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_isResolvedCaseSplit___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: { -return x_27; +uint8_t x_6; lean_object* x_7; +x_6 = l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_isResolvedCaseSplit___spec__3(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_5); +lean_dec(x_2); +lean_dec(x_1); +x_7 = lean_box(x_6); +return x_7; } -else +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_isResolvedCaseSplit___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -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_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; +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_Meta_Grind_isResolvedCaseSplit___spec__2(x_1, x_4, x_3); +lean_dec(x_3); +x_6 = lean_box(x_5); +return x_6; } } +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_isResolvedCaseSplit___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_Meta_Grind_isResolvedCaseSplit___spec__1(x_1, x_2); +lean_dec(x_2); +x_4 = lean_box(x_3); +return x_4; +} } -else +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isResolvedCaseSplit___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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_32; +lean_object* x_11; +x_11 = l_Lean_Meta_Grind_isResolvedCaseSplit(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); lean_dec(x_3); -x_32 = !lean_is_exclusive(x_13); -if (x_32 == 0) +lean_dec(x_2); +lean_dec(x_1); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_markCaseSplitAsResolved___spec__3(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: { -return x_13; +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_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_dec(x_13); -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_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_Meta_Grind_instHashableENodeKey_unsafe__1(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_Grind_markCaseSplitAsResolved___spec__2(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_Meta_Grind_filterENodes(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_markCaseSplitAsResolved___spec__4(lean_object* x_1, lean_object* x_2, lean_object* 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; lean_object* x_16; -x_11 = l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Grind_getENodes___spec__1___closed__1; -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_inc(x_14); -lean_dec(x_12); -lean_inc(x_13); -x_15 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_filterENodes___lambda__1___boxed), 12, 2); -lean_closure_set(x_15, 0, x_1); -lean_closure_set(x_15, 1, x_13); -x_16 = l_Lean_Meta_Grind_forEachENode(x_15, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_14); -if (lean_obj_tag(x_16) == 0) +lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; +x_5 = lean_ctor_get(x_1, 0); +lean_inc(x_5); +x_6 = lean_ctor_get(x_1, 1); +lean_inc(x_6); +x_7 = lean_array_get_size(x_5); +x_8 = lean_nat_dec_lt(x_2, x_7); +lean_dec(x_7); +if (x_8 == 0) { -lean_object* x_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 = lean_st_ref_get(x_13, x_17); -lean_dec(x_13); -x_19 = !lean_is_exclusive(x_18); -if (x_19 == 0) +uint8_t x_9; +lean_dec(x_2); +x_9 = !lean_is_exclusive(x_1); +if (x_9 == 0) { -return x_18; +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_20; lean_object* x_21; lean_object* x_22; -x_20 = lean_ctor_get(x_18, 0); -x_21 = lean_ctor_get(x_18, 1); -lean_inc(x_21); -lean_inc(x_20); -lean_dec(x_18); -x_22 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_22, 0, x_20); -lean_ctor_set(x_22, 1, x_21); -return x_22; +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 { -uint8_t x_23; -lean_dec(x_13); -x_23 = !lean_is_exclusive(x_16); -if (x_23 == 0) +lean_object* x_17; uint8_t x_18; +x_17 = lean_array_fget(x_5, x_2); +x_18 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_3, x_17); +lean_dec(x_17); +if (x_18 == 0) { -return x_16; +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 { -lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_24 = lean_ctor_get(x_16, 0); -x_25 = lean_ctor_get(x_16, 1); -lean_inc(x_25); -lean_inc(x_24); -lean_dec(x_16); -x_26 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_26, 0, x_24); -lean_ctor_set(x_26, 1, x_25); -return x_26; +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; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_filterENodes___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_13; -x_13 = l_Lean_Meta_Grind_filterENodes___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_2); -return x_13; } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_forEachEqc___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, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markCaseSplitAsResolved___spec__2(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5) { _start: { -uint8_t x_17; -x_17 = lean_usize_dec_lt(x_6, x_5); -if (x_17 == 0) +if (lean_obj_tag(x_1) == 0) { -lean_object* x_18; -lean_dec(x_15); -lean_dec(x_14); +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_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___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_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_1); -x_18 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_18, 0, x_7); -lean_ctor_set(x_18, 1, x_16); -return x_18; +lean_dec(x_5); +lean_dec(x_4); +return x_1; } else { -lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; -lean_dec(x_7); -x_19 = lean_array_uget(x_4, x_6); -x_20 = lean_ctor_get(x_19, 0); -lean_inc(x_20); -x_21 = lean_ctor_get(x_19, 2); -lean_inc(x_21); -x_22 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_20, x_21); -lean_dec(x_21); -lean_dec(x_20); -if (x_22 == 0) +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_array_fget(x_7, x_12); +x_16 = lean_box(0); +x_17 = lean_array_fset(x_7, x_12, x_16); +switch (lean_obj_tag(x_15)) { +case 0: { -size_t x_23; size_t x_24; lean_object* x_25; -lean_dec(x_19); -x_23 = 1; -x_24 = lean_usize_add(x_6, x_23); -x_25 = lean_box(0); -x_6 = x_24; -x_7 = x_25; -goto _start; +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 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(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_27; -lean_inc(x_1); -lean_inc(x_15); -lean_inc(x_14); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -x_27 = lean_apply_10(x_1, x_19, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); -if (lean_obj_tag(x_27) == 0) -{ -lean_object* x_28; size_t x_29; size_t x_30; lean_object* x_31; -x_28 = lean_ctor_get(x_27, 1); -lean_inc(x_28); -lean_dec(x_27); -x_29 = 1; -x_30 = lean_usize_add(x_6, x_29); -x_31 = lean_box(0); -x_6 = x_30; -x_7 = x_31; -x_16 = x_28; -goto _start; +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 { -uint8_t x_33; +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); -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_33 = !lean_is_exclusive(x_27); -if (x_33 == 0) +x_28 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_4, x_26); +if (x_28 == 0) { -return x_27; +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_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_inc(x_35); -lean_inc(x_34); +lean_object* x_32; lean_object* x_33; 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_EXPORT lean_object* l_Lean_Meta_Grind_forEachEqc(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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; size_t x_15; size_t x_16; lean_object* x_17; lean_object* x_18; -x_11 = l_Lean_Meta_Grind_getENodes(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); -x_13 = lean_ctor_get(x_11, 1); -lean_inc(x_13); -lean_dec(x_11); -x_14 = lean_box(0); -x_15 = lean_array_size(x_12); -x_16 = 0; -x_17 = lean_box(0); -x_18 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_forEachEqc___spec__1(x_1, x_12, x_14, x_12, x_15, x_16, x_17, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_13); +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); -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; -x_20 = lean_ctor_get(x_18, 0); -lean_dec(x_20); -lean_ctor_set(x_18, 0, x_17); -return x_18; +lean_ctor_set(x_1, 0, x_33); +return x_1; } -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_17); -lean_ctor_set(x_22, 1, x_21); -return x_22; } } -else +case 1: { -uint8_t x_23; -x_23 = !lean_is_exclusive(x_18); -if (x_23 == 0) +uint8_t x_34; +x_34 = !lean_is_exclusive(x_15); +if (x_34 == 0) { -return x_18; +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_Grind_markCaseSplitAsResolved___spec__2(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_24; lean_object* x_25; lean_object* x_26; -x_24 = lean_ctor_get(x_18, 0); -x_25 = lean_ctor_get(x_18, 1); -lean_inc(x_25); -lean_inc(x_24); -lean_dec(x_18); -x_26 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_26, 0, x_24); -lean_ctor_set(x_26, 1, x_25); -return x_26; +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_Grind_markCaseSplitAsResolved___spec__2(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; } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_forEachEqc___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { -_start: +} +} +else { -size_t x_17; size_t x_18; lean_object* x_19; -x_17 = lean_unbox_usize(x_5); +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_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___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); -x_18 = lean_unbox_usize(x_6); -lean_dec(x_6); -x_19 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_forEachEqc___spec__1(x_1, x_2, x_3, x_4, x_17, x_18, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_19; -} +x_56 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_56, 0, x_48); +return x_56; } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instInhabitedMethods___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: +else { -lean_object* x_11; lean_object* x_12; -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; -} -} -static lean_object* _init_l_Lean_Meta_Grind_instInhabitedMethods___closed__1() { -_start: +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_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_instInhabitedMethods___lambda__1___boxed), 10, 0); -return x_1; -} +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); } -static lean_object* _init_l_Lean_Meta_Grind_instInhabitedMethods___closed__2() { -_start: +x_63 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_4, x_60); +if (x_63 == 0) { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_Grind_instInhabitedMethods___closed__1; -x_2 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_2, 0, x_1); -lean_ctor_set(x_2, 1, x_1); -return x_2; -} +lean_object* x_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; } -static lean_object* _init_l_Lean_Meta_Grind_instInhabitedMethods() { -_start: +else { -lean_object* x_1; -x_1 = l_Lean_Meta_Grind_instInhabitedMethods___closed__2; -return x_1; +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; } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instInhabitedMethods___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: +} +case 1: { -lean_object* x_11; -x_11 = l_Lean_Meta_Grind_instInhabitedMethods___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_11; +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_Grind_markCaseSplitAsResolved___spec__2(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_EXPORT lean_object* l_Lean_Meta_Grind_Methods_toMethodsRef_unsafe__1(lean_object* x_1) { -_start: +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_inc(x_1); -return x_1; +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; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Methods_toMethodsRef_unsafe__1___boxed(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = l_Lean_Meta_Grind_Methods_toMethodsRef_unsafe__1(x_1); -lean_dec(x_1); -return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Methods_toMethodsRef(lean_object* x_1) { -_start: -{ -lean_inc(x_1); -return x_1; } +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_Grind_markCaseSplitAsResolved___spec__4(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_Grind_mkHCongrWithArity___spec__2___closed__3; +x_93 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_markCaseSplitAsResolved___spec__3(x_3, x_90, x_91, lean_box(0), x_83, x_92); +lean_dec(x_91); +lean_dec(x_90); +return x_93; } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Methods_toMethodsRef___boxed(lean_object* x_1) { -_start: +else { -lean_object* x_2; -x_2 = l_Lean_Meta_Grind_Methods_toMethodsRef(x_1); -lean_dec(x_1); -return x_2; +return x_84; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_MethodsRef_toMethods_unsafe__1(lean_object* x_1) { -_start: +else { -lean_inc(x_1); -return x_1; +return x_84; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_MethodsRef_toMethods_unsafe__1___boxed(lean_object* x_1) { -_start: +else { -lean_object* x_2; -x_2 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_MethodsRef_toMethods_unsafe__1(x_1); +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); -return x_2; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_MethodsRef_toMethods(lean_object* x_1) { -_start: +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_Grind_markCaseSplitAsResolved___spec__4(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_inc(x_1); -return x_1; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_MethodsRef_toMethods___boxed(lean_object* x_1) { -_start: +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_2; -x_2 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_MethodsRef_toMethods(x_1); -lean_dec(x_1); -return x_2; -} +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_Grind_mkHCongrWithArity___spec__2___closed__3; +x_107 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_markCaseSplitAsResolved___spec__3(x_3, x_104, x_105, lean_box(0), x_97, x_106); +lean_dec(x_105); +lean_dec(x_104); +return x_107; } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getMethods(lean_object* x_1, lean_object* x_2, lean_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 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_9, 0, x_1); -lean_ctor_set(x_9, 1, x_8); -return x_9; +return x_98; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getMethods___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: +else { -lean_object* x_9; -x_9 = l_Lean_Meta_Grind_getMethods(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; +return x_98; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateUp(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_11; lean_object* x_12; -x_11 = lean_ctor_get(x_3, 0); -lean_inc(x_11); -x_12 = lean_apply_10(x_11, x_1, 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_Meta_Grind_propagateDown(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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_PersistentHashMap_insert___at_Lean_Meta_Grind_markCaseSplitAsResolved___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_11; lean_object* x_12; -x_11 = lean_ctor_get(x_3, 1); -lean_inc(x_11); -x_12 = lean_apply_10(x_11, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_12; +uint64_t x_4; size_t x_5; size_t x_6; lean_object* x_7; +x_4 = l_Lean_Meta_Grind_instHashableENodeKey_unsafe__1(x_2); +x_5 = lean_uint64_to_usize(x_4); +x_6 = 1; +x_7 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markCaseSplitAsResolved___spec__2(x_1, x_5, x_6, x_2, x_3); +return x_7; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getEqc_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_EXPORT lean_object* l_Lean_Meta_Grind_markCaseSplitAsResolved___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -lean_object* x_13; -lean_inc(x_2); -x_13 = l_Lean_Meta_Grind_getNext(x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_13) == 0) -{ -uint8_t x_14; -x_14 = !lean_is_exclusive(x_13); -if (x_14 == 0) -{ -lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; -x_15 = lean_ctor_get(x_13, 0); -x_16 = lean_ctor_get(x_13, 1); -x_17 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_17, 0, x_2); -lean_ctor_set(x_17, 1, x_3); -x_18 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_1, x_15); -if (x_18 == 0) -{ -lean_free_object(x_13); -x_2 = x_15; -x_3 = x_17; -x_12 = x_16; -goto _start; -} -else -{ -lean_dec(x_15); -lean_ctor_set(x_13, 0, x_17); -return x_13; -} -} -else +lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_12 = lean_st_ref_take(x_3, x_11); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_12, 1); +lean_inc(x_14); +lean_dec(x_12); +x_15 = !lean_is_exclusive(x_13); +if (x_15 == 0) { -lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; -x_20 = lean_ctor_get(x_13, 0); -x_21 = lean_ctor_get(x_13, 1); -lean_inc(x_21); -lean_inc(x_20); -lean_dec(x_13); -x_22 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_22, 0, x_2); -lean_ctor_set(x_22, 1, x_3); -x_23 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_1, x_20); -if (x_23 == 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_13, 18); +x_17 = lean_box(0); +x_18 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_markCaseSplitAsResolved___spec__1(x_16, x_1, x_17); +lean_ctor_set(x_13, 18, x_18); +x_19 = lean_st_ref_set(x_3, x_13, x_14); +x_20 = !lean_is_exclusive(x_19); +if (x_20 == 0) { -x_2 = x_20; -x_3 = x_22; -x_12 = x_21; -goto _start; +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_25; -lean_dec(x_20); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_22); -lean_ctor_set(x_25, 1, x_21); -return x_25; -} -} +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_26; -lean_dec(x_3); -lean_dec(x_2); -x_26 = !lean_is_exclusive(x_13); -if (x_26 == 0) -{ -return x_13; } else { -lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_27 = lean_ctor_get(x_13, 0); -x_28 = lean_ctor_get(x_13, 1); +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; uint8_t x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_24 = lean_ctor_get(x_13, 0); +x_25 = lean_ctor_get(x_13, 1); +x_26 = lean_ctor_get(x_13, 2); +x_27 = lean_ctor_get(x_13, 3); +x_28 = lean_ctor_get(x_13, 4); +x_29 = lean_ctor_get(x_13, 5); +x_30 = lean_ctor_get_uint8(x_13, sizeof(void*)*20); +x_31 = lean_ctor_get(x_13, 6); +x_32 = lean_ctor_get(x_13, 7); +x_33 = lean_ctor_get(x_13, 8); +x_34 = lean_ctor_get(x_13, 9); +x_35 = lean_ctor_get(x_13, 10); +x_36 = lean_ctor_get(x_13, 11); +x_37 = lean_ctor_get(x_13, 12); +x_38 = lean_ctor_get(x_13, 13); +x_39 = lean_ctor_get(x_13, 14); +x_40 = lean_ctor_get(x_13, 15); +x_41 = lean_ctor_get(x_13, 16); +x_42 = lean_ctor_get(x_13, 17); +x_43 = lean_ctor_get(x_13, 18); +x_44 = lean_ctor_get(x_13, 19); +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_inc(x_34); +lean_inc(x_33); +lean_inc(x_32); +lean_inc(x_31); +lean_inc(x_29); lean_inc(x_28); lean_inc(x_27); +lean_inc(x_26); +lean_inc(x_25); +lean_inc(x_24); lean_dec(x_13); -x_29 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_29, 0, x_27); -lean_ctor_set(x_29, 1, x_28); -return x_29; +x_45 = lean_box(0); +x_46 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_markCaseSplitAsResolved___spec__1(x_43, x_1, x_45); +x_47 = lean_alloc_ctor(0, 20, 1); +lean_ctor_set(x_47, 0, x_24); +lean_ctor_set(x_47, 1, x_25); +lean_ctor_set(x_47, 2, x_26); +lean_ctor_set(x_47, 3, x_27); +lean_ctor_set(x_47, 4, x_28); +lean_ctor_set(x_47, 5, x_29); +lean_ctor_set(x_47, 6, x_31); +lean_ctor_set(x_47, 7, x_32); +lean_ctor_set(x_47, 8, x_33); +lean_ctor_set(x_47, 9, x_34); +lean_ctor_set(x_47, 10, x_35); +lean_ctor_set(x_47, 11, x_36); +lean_ctor_set(x_47, 12, x_37); +lean_ctor_set(x_47, 13, x_38); +lean_ctor_set(x_47, 14, x_39); +lean_ctor_set(x_47, 15, x_40); +lean_ctor_set(x_47, 16, x_41); +lean_ctor_set(x_47, 17, x_42); +lean_ctor_set(x_47, 18, x_46); +lean_ctor_set(x_47, 19, x_44); +lean_ctor_set_uint8(x_47, sizeof(void*)*20, x_30); +x_48 = lean_st_ref_set(x_3, x_47, x_14); +x_49 = lean_ctor_get(x_48, 1); +lean_inc(x_49); +if (lean_is_exclusive(x_48)) { + lean_ctor_release(x_48, 0); + lean_ctor_release(x_48, 1); + x_50 = x_48; +} else { + lean_dec_ref(x_48); + x_50 = lean_box(0); +} +if (lean_is_scalar(x_50)) { + x_51 = lean_alloc_ctor(0, 2, 0); +} else { + x_51 = x_50; } +lean_ctor_set(x_51, 0, x_45); +lean_ctor_set(x_51, 1, x_49); +return x_51; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getEqc_go___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +static lean_object* _init_l_Lean_Meta_Grind_markCaseSplitAsResolved___closed__1() { _start: { -lean_object* x_13; -x_13 = l_Lean_Meta_Grind_getEqc_go(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_1); -return x_13; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("split", 5, 5); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getEqc(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +static lean_object* _init_l_Lean_Meta_Grind_markCaseSplitAsResolved___closed__2() { _start: { -lean_object* x_11; lean_object* x_12; -x_11 = lean_box(0); -lean_inc(x_1); -x_12 = l_Lean_Meta_Grind_getEqc_go(x_1, x_1, x_11, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_1); -return x_12; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("resolved", 8, 8); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getEqc___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +static lean_object* _init_l_Lean_Meta_Grind_markCaseSplitAsResolved___closed__3() { _start: { -lean_object* x_11; -x_11 = l_Lean_Meta_Grind_getEqc(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_11; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78____closed__1; +x_2 = l_Lean_Meta_Grind_markCaseSplitAsResolved___closed__1; +x_3 = l_Lean_Meta_Grind_markCaseSplitAsResolved___closed__2; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_getEqcs___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, lean_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_Meta_Grind_markCaseSplitAsResolved(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -uint8_t x_16; -x_16 = lean_usize_dec_lt(x_5, x_4); -if (x_16 == 0) +lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_11 = l_Lean_Meta_Grind_isResolvedCaseSplit(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +x_13 = lean_unbox(x_12); +lean_dec(x_12); +if (x_13 == 0) { -lean_object* x_17; -x_17 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_17, 0, x_6); -lean_ctor_set(x_17, 1, x_15); -return x_17; -} -else +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_11, 1); +lean_inc(x_14); +lean_dec(x_11); +x_15 = l_Lean_Meta_Grind_markCaseSplitAsResolved___closed__3; +x_16 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_15, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_14); +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +x_18 = lean_unbox(x_17); +lean_dec(x_17); +if (x_18 == 0) { -lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; -x_18 = lean_array_uget(x_3, x_5); -x_19 = lean_ctor_get(x_18, 2); +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_16, 1); lean_inc(x_19); -x_20 = lean_ctor_get(x_18, 0); -lean_inc(x_20); -lean_dec(x_18); -x_21 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_19, x_20); -lean_dec(x_19); -if (x_21 == 0) -{ -size_t x_22; size_t x_23; -lean_dec(x_20); -x_22 = 1; -x_23 = lean_usize_add(x_5, x_22); -x_5 = x_23; -goto _start; +lean_dec(x_16); +x_20 = lean_box(0); +x_21 = l_Lean_Meta_Grind_markCaseSplitAsResolved___lambda__1(x_1, x_20, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_19); +return x_21; } else { -lean_object* x_25; -x_25 = l_Lean_Meta_Grind_getEqc(x_20, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +uint8_t x_22; +x_22 = !lean_is_exclusive(x_16); +if (x_22 == 0) +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_16, 1); +x_24 = lean_ctor_get(x_16, 0); +lean_dec(x_24); +x_25 = l_Lean_Meta_Grind_updateLastTag(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_23); if (lean_obj_tag(x_25) == 0) { -lean_object* x_26; lean_object* x_27; lean_object* x_28; size_t x_29; size_t x_30; -x_26 = lean_ctor_get(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_25, 1); lean_inc(x_26); -x_27 = lean_ctor_get(x_25, 1); -lean_inc(x_27); lean_dec(x_25); -x_28 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set(x_28, 1, x_6); -x_29 = 1; -x_30 = lean_usize_add(x_5, x_29); -x_5 = x_30; -x_6 = x_28; -x_15 = x_27; -goto _start; +lean_inc(x_1); +x_27 = l_Lean_MessageData_ofExpr(x_1); +x_28 = l_Lean_Meta_Grind_getENode___closed__3; +lean_ctor_set_tag(x_16, 7); +lean_ctor_set(x_16, 1, x_27); +lean_ctor_set(x_16, 0, x_28); +x_29 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_29, 0, x_16); +lean_ctor_set(x_29, 1, x_28); +x_30 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_15, x_29, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_26); +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_Grind_markCaseSplitAsResolved___lambda__1(x_1, x_31, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_32); +lean_dec(x_31); +return x_33; } else { -uint8_t x_32; -lean_dec(x_6); -x_32 = !lean_is_exclusive(x_25); -if (x_32 == 0) +uint8_t x_34; +lean_free_object(x_16); +lean_dec(x_1); +x_34 = !lean_is_exclusive(x_25); +if (x_34 == 0) { return x_25; } 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_inc(x_34); -lean_inc(x_33); +lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_35 = lean_ctor_get(x_25, 0); +x_36 = lean_ctor_get(x_25, 1); +lean_inc(x_36); +lean_inc(x_35); lean_dec(x_25); -x_35 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -return x_35; -} -} -} +x_37 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_37, 0, x_35); +lean_ctor_set(x_37, 1, x_36); +return x_37; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getEqcs(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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; size_t x_15; size_t x_16; lean_object* x_17; -x_10 = lean_box(0); -x_11 = l_Lean_Meta_Grind_getENodes(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); -x_13 = lean_ctor_get(x_11, 1); -lean_inc(x_13); -lean_dec(x_11); -x_14 = lean_box(0); -x_15 = lean_array_size(x_12); -x_16 = 0; -x_17 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_getEqcs___spec__1(x_12, x_14, x_12, x_15, x_16, x_10, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_13); -lean_dec(x_12); -if (lean_obj_tag(x_17) == 0) +else { -uint8_t x_18; -x_18 = !lean_is_exclusive(x_17); -if (x_18 == 0) +lean_object* x_38; lean_object* x_39; +x_38 = lean_ctor_get(x_16, 1); +lean_inc(x_38); +lean_dec(x_16); +x_39 = l_Lean_Meta_Grind_updateLastTag(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_38); +if (lean_obj_tag(x_39) == 0) { -return x_17; +lean_object* x_40; lean_object* x_41; lean_object* x_42; 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_40 = lean_ctor_get(x_39, 1); +lean_inc(x_40); +lean_dec(x_39); +lean_inc(x_1); +x_41 = l_Lean_MessageData_ofExpr(x_1); +x_42 = l_Lean_Meta_Grind_getENode___closed__3; +x_43 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_43, 0, x_42); +lean_ctor_set(x_43, 1, x_41); +x_44 = 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_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_15, x_44, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_40); +x_46 = lean_ctor_get(x_45, 0); +lean_inc(x_46); +x_47 = lean_ctor_get(x_45, 1); +lean_inc(x_47); +lean_dec(x_45); +x_48 = l_Lean_Meta_Grind_markCaseSplitAsResolved___lambda__1(x_1, x_46, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_47); +lean_dec(x_46); +return x_48; } else { -lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_19 = lean_ctor_get(x_17, 0); -x_20 = lean_ctor_get(x_17, 1); -lean_inc(x_20); -lean_inc(x_19); -lean_dec(x_17); -x_21 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_21, 0, x_19); -lean_ctor_set(x_21, 1, x_20); -return x_21; +lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; +lean_dec(x_1); +x_49 = lean_ctor_get(x_39, 0); +lean_inc(x_49); +x_50 = lean_ctor_get(x_39, 1); +lean_inc(x_50); +if (lean_is_exclusive(x_39)) { + lean_ctor_release(x_39, 0); + lean_ctor_release(x_39, 1); + x_51 = x_39; +} else { + lean_dec_ref(x_39); + x_51 = lean_box(0); +} +if (lean_is_scalar(x_51)) { + x_52 = lean_alloc_ctor(1, 2, 0); +} else { + x_52 = x_51; +} +lean_ctor_set(x_52, 0, x_49); +lean_ctor_set(x_52, 1, x_50); +return x_52; +} +} } } else { -uint8_t x_22; -x_22 = !lean_is_exclusive(x_17); -if (x_22 == 0) +uint8_t x_53; +lean_dec(x_1); +x_53 = !lean_is_exclusive(x_11); +if (x_53 == 0) { -return x_17; +lean_object* x_54; lean_object* x_55; +x_54 = lean_ctor_get(x_11, 0); +lean_dec(x_54); +x_55 = lean_box(0); +lean_ctor_set(x_11, 0, x_55); +return x_11; } else { -lean_object* x_23; lean_object* x_24; lean_object* x_25; -x_23 = lean_ctor_get(x_17, 0); -x_24 = lean_ctor_get(x_17, 1); -lean_inc(x_24); -lean_inc(x_23); -lean_dec(x_17); -x_25 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_25, 0, x_23); -lean_ctor_set(x_25, 1, x_24); -return x_25; +lean_object* x_56; lean_object* x_57; lean_object* x_58; +x_56 = lean_ctor_get(x_11, 1); +lean_inc(x_56); +lean_dec(x_11); +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_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_getEqcs___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_markCaseSplitAsResolved___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { -size_t x_16; size_t x_17; lean_object* x_18; -x_16 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_17 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_18 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_getEqcs___spec__1(x_1, x_2, x_3, x_16, x_17, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); +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_Grind_markCaseSplitAsResolved___spec__3(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_Grind_markCaseSplitAsResolved___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +size_t x_6; size_t x_7; lean_object* x_8; +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_Grind_markCaseSplitAsResolved___spec__2(x_1, x_6, x_7, x_4, x_5); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markCaseSplitAsResolved___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_Meta_Grind_markCaseSplitAsResolved___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -return x_18; +return x_12; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getEqcs___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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_Grind_markCaseSplitAsResolved___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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_Grind_getEqcs(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_Lean_Meta_Grind_markCaseSplitAsResolved(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); @@ -18502,8 +23485,7 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -return x_10; +return x_11; } } lean_object* initialize_Init_Grind_Tactics(uint8_t builtin, lean_object*); @@ -18628,114 +23610,116 @@ l_Lean_Meta_Grind_instInhabitedENode___closed__2 = _init_l_Lean_Meta_Grind_instI lean_mark_persistent(l_Lean_Meta_Grind_instInhabitedENode___closed__2); l_Lean_Meta_Grind_instInhabitedENode = _init_l_Lean_Meta_Grind_instInhabitedENode(); lean_mark_persistent(l_Lean_Meta_Grind_instInhabitedENode); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__1(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__1); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__2(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__2); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__3(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__3); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__4(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__4); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__5 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__5(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__5); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__6 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__6(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__6); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__7 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__7(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__7); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__8 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__8(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__8); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__9 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__9(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__9); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__10 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__10(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__10); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__11 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__11(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__11); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__12 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__12(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__12); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__13 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__13(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__13); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__14 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__14(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__14); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__15 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__15(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__15); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__16 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__16(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__16); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__17 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__17(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__17); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__18 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__18(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__18); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__19 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__19(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__19); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__20 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__20(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__20); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__21 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__21(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__21); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__22 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__22(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__22); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__23 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__23(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__23); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__24 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__24(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__24); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__25 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__25(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__25); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__26 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__26(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__26); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__27 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__27(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__27); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__28 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__28(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__28); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__29 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__29(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__29); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__30 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__30(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__30); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__31 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__31(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__31); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__32 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__32(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__32); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__33 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__33(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__33); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__34 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__34(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__34); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__35 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__35(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__35); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__36 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__36(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__36); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__37 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__37(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__37); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__38 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__38(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__38); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__39 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__39(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__39); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__40 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__40(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__40); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__41 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__41(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__41); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__42 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__42(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__42); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__43 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__43(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__43); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__44 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__44(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__44); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__45 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__45(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__45); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__46 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__46(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__46); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__47 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__47(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__47); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__48 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__48(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__48); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__49 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__49(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__49); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__50 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__50(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__50); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__51 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__51(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__51); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__52 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__52(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__52); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__53 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__53(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__53); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__54 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__54(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__54); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__1); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__2); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__3); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__4(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__4); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__5 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__5(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__5); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__6 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__6(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__6); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__7 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__7(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__7); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__8 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__8(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__8); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__9 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__9(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__9); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__10 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__10(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__10); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__11 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__11(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__11); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__12 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__12(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__12); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__13 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__13(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__13); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__14 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__14(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__14); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__15 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__15(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__15); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__16 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__16(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__16); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__17 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__17(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__17); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__18 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__18(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__18); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__19 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__19(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__19); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__20 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__20(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__20); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__21 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__21(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__21); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__22 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__22(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__22); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__23 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__23(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__23); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__24 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__24(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__24); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__25 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__25(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__25); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__26 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__26(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__26); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__27 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__27(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__27); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__28 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__28(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__28); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__29 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__29(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__29); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__30 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__30(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__30); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__31 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__31(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__31); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__32 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__32(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__32); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__33 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__33(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__33); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__34 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__34(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__34); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__35 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__35(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__35); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__36 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__36(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__36); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__37 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__37(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__37); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__38 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__38(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__38); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__39 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__39(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__39); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__40 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__40(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__40); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__41 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__41(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__41); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__42 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__42(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__42); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__43 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__43(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__43); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__44 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__44(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__44); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__45 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__45(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__45); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__46 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__46(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__46); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__47 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__47(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__47); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__48 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__48(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__48); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__49 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__49(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__49); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__50 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__50(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__50); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__51 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__51(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__51); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__52 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__52(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__52); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__53 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__53(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__53); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__54 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__54(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__54); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__55 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__55(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__55); l_Lean_Meta_Grind_instReprENode___closed__1 = _init_l_Lean_Meta_Grind_instReprENode___closed__1(); lean_mark_persistent(l_Lean_Meta_Grind_instReprENode___closed__1); l_Lean_Meta_Grind_instReprENode = _init_l_Lean_Meta_Grind_instReprENode(); @@ -18744,6 +23728,10 @@ l_Lean_Meta_Grind_congrHash___closed__1 = _init_l_Lean_Meta_Grind_congrHash___cl lean_mark_persistent(l_Lean_Meta_Grind_congrHash___closed__1); l_Lean_Meta_Grind_congrHash___closed__2 = _init_l_Lean_Meta_Grind_congrHash___closed__2(); lean_mark_persistent(l_Lean_Meta_Grind_congrHash___closed__2); +l_Lean_Meta_Grind_congrHash___closed__3 = _init_l_Lean_Meta_Grind_congrHash___closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind_congrHash___closed__3); +l_Lean_Meta_Grind_congrHash___closed__4 = _init_l_Lean_Meta_Grind_congrHash___closed__4(); +lean_mark_persistent(l_Lean_Meta_Grind_congrHash___closed__4); l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1___closed__1 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1___closed__1(); lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1___closed__1); l_Lean_Meta_Grind_instBEqPreInstance___lambda__2___closed__1 = _init_l_Lean_Meta_Grind_instBEqPreInstance___lambda__2___closed__1(); @@ -18758,6 +23746,8 @@ l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedGoal___spec__1 lean_mark_persistent(l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedGoal___spec__1); l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedGoal___spec__2 = _init_l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedGoal___spec__2(); lean_mark_persistent(l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedGoal___spec__2); +l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedGoal___spec__3 = _init_l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedGoal___spec__3(); +lean_mark_persistent(l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedGoal___spec__3); l_Lean_Meta_Grind_instInhabitedGoal___closed__1 = _init_l_Lean_Meta_Grind_instInhabitedGoal___closed__1(); lean_mark_persistent(l_Lean_Meta_Grind_instInhabitedGoal___closed__1); l_Lean_Meta_Grind_instInhabitedGoal___closed__2 = _init_l_Lean_Meta_Grind_instInhabitedGoal___closed__2(); @@ -18767,31 +23757,256 @@ l_Lean_Meta_Grind_instInhabitedGoal___closed__4 = _init_l_Lean_Meta_Grind_instIn lean_mark_persistent(l_Lean_Meta_Grind_instInhabitedGoal___closed__4); l_Lean_Meta_Grind_instInhabitedGoal___closed__5 = _init_l_Lean_Meta_Grind_instInhabitedGoal___closed__5(); lean_mark_persistent(l_Lean_Meta_Grind_instInhabitedGoal___closed__5); +l_Lean_Meta_Grind_instInhabitedGoal___closed__6 = _init_l_Lean_Meta_Grind_instInhabitedGoal___closed__6(); +lean_mark_persistent(l_Lean_Meta_Grind_instInhabitedGoal___closed__6); l_Lean_Meta_Grind_instInhabitedGoal = _init_l_Lean_Meta_Grind_instInhabitedGoal(); lean_mark_persistent(l_Lean_Meta_Grind_instInhabitedGoal); l_Lean_Meta_Grind_GoalM_run_x27___closed__1 = _init_l_Lean_Meta_Grind_GoalM_run_x27___closed__1(); lean_mark_persistent(l_Lean_Meta_Grind_GoalM_run_x27___closed__1); +l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2___closed__1 = _init_l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2___closed__1(); +l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2___closed__2 = _init_l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2___closed__2(); +lean_mark_persistent(l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2___closed__2); +l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2___closed__3 = _init_l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2___closed__3(); +lean_mark_persistent(l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2___closed__3); +l_Lean_Meta_Grind_updateLastTag___closed__1 = _init_l_Lean_Meta_Grind_updateLastTag___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_updateLastTag___closed__1); +l_Lean_Meta_Grind_updateLastTag___closed__2 = _init_l_Lean_Meta_Grind_updateLastTag___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_updateLastTag___closed__2); +l_Lean_Meta_Grind_updateLastTag___closed__3 = _init_l_Lean_Meta_Grind_updateLastTag___closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind_updateLastTag___closed__3); +l_Lean_Meta_Grind_updateLastTag___closed__4 = _init_l_Lean_Meta_Grind_updateLastTag___closed__4(); +lean_mark_persistent(l_Lean_Meta_Grind_updateLastTag___closed__4); +l_Lean_Meta_Grind_updateLastTag___closed__5 = _init_l_Lean_Meta_Grind_updateLastTag___closed__5(); +lean_mark_persistent(l_Lean_Meta_Grind_updateLastTag___closed__5); +l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__1 = _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__1); +l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__2 = _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__2); +l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__3 = _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__3); +l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__4 = _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__4(); +lean_mark_persistent(l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__4); +l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__5 = _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__5(); +lean_mark_persistent(l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__5); +l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__6 = _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__6(); +lean_mark_persistent(l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__6); +l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__7 = _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__7(); +lean_mark_persistent(l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__7); +l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__8 = _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__8(); +lean_mark_persistent(l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__8); +l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__9 = _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__9(); +lean_mark_persistent(l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__9); +l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__10 = _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__10(); +lean_mark_persistent(l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__10); +l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__11 = _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__11(); +lean_mark_persistent(l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__11); +l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__12 = _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__12(); +lean_mark_persistent(l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__12); +l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__13 = _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__13(); +lean_mark_persistent(l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__13); +l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__14 = _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__14(); +lean_mark_persistent(l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__14); +l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__15 = _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__15(); +lean_mark_persistent(l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__15); +l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__16 = _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__16(); +lean_mark_persistent(l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__16); +l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__17 = _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__17(); +lean_mark_persistent(l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__17); +l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__18 = _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__18(); +lean_mark_persistent(l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__18); +l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__19 = _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__19(); +lean_mark_persistent(l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__19); +l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__20 = _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__20(); +lean_mark_persistent(l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__20); +l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__21 = _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__21(); +lean_mark_persistent(l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__21); +l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__22 = _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__22(); +lean_mark_persistent(l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__22); +l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__23 = _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__23(); +lean_mark_persistent(l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__23); +l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__24 = _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__24(); +lean_mark_persistent(l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__24); +l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d____ = _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d____(); +lean_mark_persistent(l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d____); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__1 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__1); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__2 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__2); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__3 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__3); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__4 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__4(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__4); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__5 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__5(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__5); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__6 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__6(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__6); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__7 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__7(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__7); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__8 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__8(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__8); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__9 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__9(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__9); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__10 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__10(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__10); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__11 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__11(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__11); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__12 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__12(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__12); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__13 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__13(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__13); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__14 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__14(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__14); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__15 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__15(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__15); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__16 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__16(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__16); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__17 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__17(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__17); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__18 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__18(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__18); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__19 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__19(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__19); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__20 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__20(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__20); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__21 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__21(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__21); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__22 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__22(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__22); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__23 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__23(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__23); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__24 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__24(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__24); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__25 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__25(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__25); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__26 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__26(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__26); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__27 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__27(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__27); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__28 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__28(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__28); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__29 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__29(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__29); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__30 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__30(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__30); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__31 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__31(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__31); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__32 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__32(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__32); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__33 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__33(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__33); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__34 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__34(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__34); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__35 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__35(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__35); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__36 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__36(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__36); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__37 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__37(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__37); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__38 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__38(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__38); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__39 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__39(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__39); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__40 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__40(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__40); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__41 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__41(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__41); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__42 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__42(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__42); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__43 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__43(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__43); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__44 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__44(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__44); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__45 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__45(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__45); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__46 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__46(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__46); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__47 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__47(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__47); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__48 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__48(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__48); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__49 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__49(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__49); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__50 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__50(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__50); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__51 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__51(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__51); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__52 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__52(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__52); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__53 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__53(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__53); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__54 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__54(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__54); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__55 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__55(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__55); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__56 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__56(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__56); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__57 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__57(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__57); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__58 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__58(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__58); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__59 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__59(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__59); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__60 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__60(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__60); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__1 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__1); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__2 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__2); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__3 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__3); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__4 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__4(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__4); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__5 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__5(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__5); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__6 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__6(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__6); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__7 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__7(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__7); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__8 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__8(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__8); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__9 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__9(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__9); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__10 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__10(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__10); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__11 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__11(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__11); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__12 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__12(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__12); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__13 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__13(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__13); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__14 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__14(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__14); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__15 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__15(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__15); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__16 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__16(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__16); l_Lean_Meta_Grind_getENode___closed__1 = _init_l_Lean_Meta_Grind_getENode___closed__1(); lean_mark_persistent(l_Lean_Meta_Grind_getENode___closed__1); l_Lean_Meta_Grind_getENode___closed__2 = _init_l_Lean_Meta_Grind_getENode___closed__2(); lean_mark_persistent(l_Lean_Meta_Grind_getENode___closed__2); l_Lean_Meta_Grind_getENode___closed__3 = _init_l_Lean_Meta_Grind_getENode___closed__3(); lean_mark_persistent(l_Lean_Meta_Grind_getENode___closed__3); -l_Lean_Meta_Grind_getENode___closed__4 = _init_l_Lean_Meta_Grind_getENode___closed__4(); -lean_mark_persistent(l_Lean_Meta_Grind_getENode___closed__4); l_Lean_Meta_Grind_hasSameType___closed__1 = _init_l_Lean_Meta_Grind_hasSameType___closed__1(); +l_Lean_Meta_Grind_markAsInconsistent___closed__1 = _init_l_Lean_Meta_Grind_markAsInconsistent___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_markAsInconsistent___closed__1); +l_Lean_Meta_Grind_markAsInconsistent___closed__2 = _init_l_Lean_Meta_Grind_markAsInconsistent___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_markAsInconsistent___closed__2); +l_Lean_Meta_Grind_markAsInconsistent___closed__3 = _init_l_Lean_Meta_Grind_markAsInconsistent___closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind_markAsInconsistent___closed__3); l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Grind_getENodes___spec__1___closed__1 = _init_l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Grind_getENodes___spec__1___closed__1(); lean_mark_persistent(l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Grind_getENodes___spec__1___closed__1); -l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Grind_getENodes___spec__1___closed__2 = _init_l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Grind_getENodes___spec__1___closed__2(); -lean_mark_persistent(l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Grind_getENodes___spec__1___closed__2); l_Array_qsort_sort___at_Lean_Meta_Grind_getENodes___spec__4___closed__1 = _init_l_Array_qsort_sort___at_Lean_Meta_Grind_getENodes___spec__4___closed__1(); lean_mark_persistent(l_Array_qsort_sort___at_Lean_Meta_Grind_getENodes___spec__4___closed__1); l_Lean_Meta_Grind_instInhabitedMethods___closed__1 = _init_l_Lean_Meta_Grind_instInhabitedMethods___closed__1(); lean_mark_persistent(l_Lean_Meta_Grind_instInhabitedMethods___closed__1); -l_Lean_Meta_Grind_instInhabitedMethods___closed__2 = _init_l_Lean_Meta_Grind_instInhabitedMethods___closed__2(); -lean_mark_persistent(l_Lean_Meta_Grind_instInhabitedMethods___closed__2); l_Lean_Meta_Grind_instInhabitedMethods = _init_l_Lean_Meta_Grind_instInhabitedMethods(); lean_mark_persistent(l_Lean_Meta_Grind_instInhabitedMethods); +l_Lean_Meta_Grind_markCaseSplitAsResolved___closed__1 = _init_l_Lean_Meta_Grind_markCaseSplitAsResolved___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_markCaseSplitAsResolved___closed__1); +l_Lean_Meta_Grind_markCaseSplitAsResolved___closed__2 = _init_l_Lean_Meta_Grind_markCaseSplitAsResolved___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_markCaseSplitAsResolved___closed__2); +l_Lean_Meta_Grind_markCaseSplitAsResolved___closed__3 = _init_l_Lean_Meta_Grind_markCaseSplitAsResolved___closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind_markCaseSplitAsResolved___closed__3); return lean_io_result_mk_ok(lean_box(0)); } #ifdef __cplusplus diff --git a/stage0/stdlib/Lean/Meta/Tactic/Grind/Util.c b/stage0/stdlib/Lean/Meta/Tactic/Grind/Util.c index 4576028b255a..81c560c88a86 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Grind/Util.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Grind/Util.c @@ -22,6 +22,7 @@ LEAN_EXPORT lean_object* l_Lean_MVarId_byContra_x3f___lambda__1(lean_object*, le lean_object* l_Lean_Meta_throwTacticEx___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkAppN(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_clearAuxDecls___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_Grind_foldProjs___lambda__2___closed__7; static lean_object* l_Lean_Core_transform___at_Lean_Meta_Grind_unfoldReducible___spec__1___closed__3; lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Meta_Grind_normalizeLevels___spec__1(lean_object*, lean_object*); @@ -42,12 +43,15 @@ static lean_object* l_Lean_Meta_Grind_foldProjs___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_unfoldReducible___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_uint64_to_usize(uint64_t); static lean_object* l_Lean_Meta_Grind_eraseIrrelevantMData___closed__2; +static lean_object* l_Lean_Meta_Grind_foldProjs___lambda__2___closed__10; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_eraseIrrelevantMData(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_sort___override(lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_withIncRecDepth___at_Lean_Meta_Grind_unfoldReducible___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_mk_array(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_foldProjs___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_usize_dec_eq(size_t, size_t); lean_object* l_ST_Prim_Ref_get___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_foldProjs___lambda__2___closed__2; lean_object* l_Lean_Expr_mdata___override(lean_object*, lean_object*); lean_object* l_Lean_Meta_abstractNestedProofs(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*); @@ -60,10 +64,12 @@ static lean_object* l_Lean_MVarId_ensureNoMVar___closed__1; lean_object* l_Nat_nextPowerOfTwo_go(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_MVarId_ensureProp___closed__4; static lean_object* l_Lean_MVarId_unfoldReducible___closed__1; +lean_object* l_Lean_stringToMessageData(lean_object*); static lean_object* l_List_forIn_x27_loop___at_Lean_MVarId_clearAuxDecls___spec__6___closed__4; static lean_object* l_Lean_Core_transform___at_Lean_Meta_Grind_unfoldReducible___spec__1___closed__4; lean_object* l_Std_DHashMap_Internal_AssocList_replace___at___private_Lean_MetavarContext_0__Lean_MetavarContext_MkBinding_visit___spec__6(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_MVarId_byContra_x3f___lambda__1___closed__5; +static lean_object* l_Lean_Meta_Grind_foldProjs___lambda__2___closed__9; LEAN_EXPORT lean_object* l_Lean_Core_transform_visit___at_Lean_Meta_Grind_unfoldReducible___spec__2___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_MVarId_byContra_x3f___lambda__1___closed__4; 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*); @@ -80,11 +86,13 @@ static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Meta_Grind_unfoldReducib static lean_object* l_List_forIn_x27_loop___at_Lean_MVarId_clearAuxDecls___spec__6___closed__1; uint64_t lean_uint64_shift_right(uint64_t, uint64_t); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_eraseIrrelevantMData___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_foldProjs___lambda__2___closed__3; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_MVarId_clearAuxDecls___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* l_StateRefT_x27_lift___rarg___boxed(lean_object*, lean_object*); lean_object* lean_nat_div(lean_object*, lean_object*); lean_object* l_Lean_MVarId_getType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_MVarId_betaReduce___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_normalize___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Meta_Grind_unfoldReducible___spec__7___rarg___closed__3; static lean_object* l_Lean_MVarId_byContra_x3f___closed__1; static lean_object* l_Lean_MVarId_ensureProp___closed__1; @@ -101,6 +109,7 @@ uint8_t l_List_isEmpty___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_normalizeLevels___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_mk_ref(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_betaReduce___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_foldProjs___lambda__2___closed__12; static lean_object* l_Lean_MVarId_ensureNoMVar___closed__4; static lean_object* l_Lean_Core_transform___at_Lean_Meta_Grind_unfoldReducible___spec__1___closed__1; static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Meta_Grind_unfoldReducible___spec__7___rarg___closed__6; @@ -110,7 +119,9 @@ extern lean_object* l_Lean_levelZero; static lean_object* l_Lean_MVarId_byContra_x3f___closed__2; static lean_object* l_Lean_MVarId_ensureNoMVar___closed__3; LEAN_EXPORT lean_object* l_Lean_MVarId_betaReduce(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_foldProjs___lambda__2___closed__8; uint64_t l_Lean_Expr_hash(lean_object*); +lean_object* lean_grind_normalize(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_letE___override(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t); lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at___private_Lean_MetavarContext_0__Lean_MetavarContext_MkBinding_visit___spec__1(lean_object*, lean_object*); lean_object* l_Lean_Core_betaReduce___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -128,6 +139,7 @@ LEAN_EXPORT lean_object* l_Lean_MVarId_byContra_x3f___lambda__2(lean_object*, le LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_MVarId_clearAuxDecls___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_LocalDecl_isAuxDecl(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_Grind_unfoldReducible___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_EXPORT lean_object* l_Lean_Meta_Grind_foldProjs___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_transform_visit___at_Lean_Meta_Grind_unfoldReducible___spec__2___lambda__2(size_t, size_t, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_MVarId_ensureProp___closed__2; static lean_object* l_Lean_MVarId_ensureNoMVar___closed__6; @@ -149,6 +161,7 @@ uint8_t lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_clearAuxDecls___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_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); +lean_object* l_Lean_indentExpr(lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_clearAuxDecls___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_unfoldDefinition_x3f(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_transform_visit_visitPost___at_Lean_Meta_Grind_unfoldReducible___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -165,6 +178,7 @@ static lean_object* l_Lean_MVarId_ensureNoMVar___closed__5; LEAN_EXPORT lean_object* l_Lean_MVarId_transformTarget(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_getAppFn(lean_object*); lean_object* lean_nat_mul(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_foldProjs___lambda__2___closed__6; LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at_Lean_MVarId_clearAuxDecls___spec__2___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkArrow(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_MetavarContext_0__Lean_MetavarContext_MkBinding_visit___spec__2(lean_object*, lean_object*); @@ -177,14 +191,18 @@ LEAN_EXPORT lean_object* l_Lean_MVarId_byContra_x3f___lambda__1___boxed(lean_obj LEAN_EXPORT lean_object* l_Lean_MVarId_ensureProp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_MVarId_clearAuxDecls___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*); size_t lean_usize_add(size_t, size_t); +static lean_object* l_Lean_Meta_Grind_foldProjs___lambda__2___closed__5; +static lean_object* l_Lean_Meta_Grind_foldProjs___lambda__2___closed__1; lean_object* lean_array_uget(lean_object*, size_t); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at_Lean_MVarId_clearAuxDecls___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_array_size(lean_object*); static lean_object* l_Lean_MVarId_ensureProp___closed__3; -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_foldProjs___lambda__2___boxed(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*); static lean_object* l_Lean_Meta_Grind_unfoldReducible___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_unfoldReducible(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_foldProjs___lambda__2___closed__11; LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Meta_Grind_unfoldReducible___spec__7(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_object* lean_array_get_size(lean_object*); 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*); @@ -206,9 +224,12 @@ LEAN_EXPORT lean_object* l_Lean_MVarId_betaReduce___lambda__1___boxed(lean_objec static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Meta_Grind_unfoldReducible___spec__7___rarg___closed__5; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_MVarId_clearAuxDecls___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_uset(lean_object*, size_t, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_foldProjs___lambda__1(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_Meta_Grind_foldProjs___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_transform_visit___at_Lean_Meta_Grind_unfoldReducible___spec__2___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l___private_Init_Data_Repr_0__Nat_reprFast(lean_object*); LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_MVarId_clearAuxDecls___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_Meta_Grind_foldProjs___lambda__2___closed__4; static lean_object* l_Lean_Meta_Grind_unfoldReducible___closed__2; size_t lean_usize_land(size_t, size_t); lean_object* l_Lean_Core_betaReduce___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -217,6 +238,7 @@ LEAN_EXPORT lean_object* l_Lean_Core_withIncRecDepth___at_Lean_Meta_Grind_unfold uint8_t l_Lean_Expr_hasExprMVar(lean_object*); static lean_object* l_Lean_Meta_Grind_normalizeLevels___closed__1; LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at_Lean_MVarId_clearAuxDecls___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_foldProjs___lambda__1___boxed(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_406_(uint8_t, uint8_t); static lean_object* _init_l_Lean_MVarId_ensureNoMVar___closed__1() { _start: @@ -5087,7 +5109,7 @@ lean_ctor_set(x_6, 0, x_5); lean_ctor_set(x_6, 1, x_4); return x_6; } -case 7: +case 8: { lean_object* x_7; lean_object* x_8; x_7 = lean_alloc_ctor(0, 1, 0); @@ -5097,42 +5119,32 @@ lean_ctor_set(x_8, 0, x_7); lean_ctor_set(x_8, 1, x_4); return x_8; } -case 8: -{ -lean_object* x_9; lean_object* x_10; -x_9 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_9, 0, x_1); -x_10 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_10, 0, x_9); -lean_ctor_set(x_10, 1, x_4); -return x_10; -} case 10: { -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_11 = lean_ctor_get(x_1, 1); -lean_inc(x_11); +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_9 = lean_ctor_get(x_1, 1); +lean_inc(x_9); lean_dec(x_1); -x_12 = lean_alloc_ctor(1, 1, 0); +x_10 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_10, 0, x_9); +x_11 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_11, 0, x_10); +x_12 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_12, 0, x_11); -x_13 = lean_alloc_ctor(2, 1, 0); -lean_ctor_set(x_13, 0, x_12); -x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_13); -lean_ctor_set(x_14, 1, x_4); -return x_14; +lean_ctor_set(x_12, 1, x_4); +return x_12; } default: { -lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_15, 0, x_1); -x_16 = lean_alloc_ctor(2, 1, 0); -lean_ctor_set(x_16, 0, 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_4); -return x_17; +lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_13 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_13, 0, x_1); +x_14 = lean_alloc_ctor(2, 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_4); +return x_15; } } } @@ -5195,7 +5207,122 @@ lean_dec(x_2); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_foldProjs___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_Meta_Grind_foldProjs___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; lean_object* x_9; +x_8 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_8, 0, x_1); +x_9 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_9, 0, x_8); +lean_ctor_set(x_9, 1, x_7); +return x_9; +} +} +static lean_object* _init_l_Lean_Meta_Grind_foldProjs___lambda__2___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("issues", 6, 6); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_foldProjs___lambda__2___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_MVarId_ensureNoMVar___closed__1; +x_2 = l_Lean_Meta_Grind_foldProjs___lambda__2___closed__1; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind_foldProjs___lambda__2___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("found `Expr.proj` but `", 23, 23); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_foldProjs___lambda__2___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_foldProjs___lambda__2___closed__3; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Grind_foldProjs___lambda__2___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("` is not marked as structure", 28, 28); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_foldProjs___lambda__2___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_foldProjs___lambda__2___closed__5; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Grind_foldProjs___lambda__2___closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("", 0, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_foldProjs___lambda__2___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_foldProjs___lambda__2___closed__7; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Grind_foldProjs___lambda__2___closed__9() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("found `Expr.proj` with invalid field index `", 44, 44); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_foldProjs___lambda__2___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_foldProjs___lambda__2___closed__9; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Grind_foldProjs___lambda__2___closed__11() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("`", 1, 1); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_foldProjs___lambda__2___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_foldProjs___lambda__2___closed__11; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_foldProjs___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: { if (lean_obj_tag(x_1) == 11) @@ -5217,198 +5344,463 @@ x_13 = lean_ctor_get(x_10, 1); x_14 = lean_ctor_get(x_12, 0); lean_inc(x_14); lean_dec(x_12); +lean_inc(x_7); x_15 = l_Lean_getStructureInfo_x3f(x_14, x_7); lean_dec(x_14); if (lean_obj_tag(x_15) == 0) { -lean_object* x_16; +lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; lean_dec(x_9); lean_dec(x_8); +x_16 = l_Lean_Meta_Grind_foldProjs___lambda__2___closed__2; +x_17 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_16, x_2, x_3, x_4, x_5, x_13); +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +x_19 = lean_unbox(x_18); +lean_dec(x_18); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; +lean_free_object(x_10); +lean_dec(x_7); +x_20 = lean_ctor_get(x_17, 1); +lean_inc(x_20); +lean_dec(x_17); +x_21 = lean_box(0); +x_22 = l_Lean_Meta_Grind_foldProjs___lambda__1(x_1, x_21, x_2, x_3, x_4, x_5, x_20); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_16 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_16, 0, x_1); -lean_ctor_set(x_10, 0, x_16); -return x_10; +return x_22; } else { -uint8_t x_17; -x_17 = !lean_is_exclusive(x_15); -if (x_17 == 0) +uint8_t x_23; +x_23 = !lean_is_exclusive(x_17); +if (x_23 == 0) { -lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; -x_18 = lean_ctor_get(x_15, 0); -x_19 = lean_ctor_get(x_18, 1); -lean_inc(x_19); -lean_dec(x_18); -x_20 = lean_array_get_size(x_19); -x_21 = lean_nat_dec_lt(x_8, x_20); -lean_dec(x_20); -if (x_21 == 0) +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_24 = lean_ctor_get(x_17, 1); +x_25 = lean_ctor_get(x_17, 0); +lean_dec(x_25); +x_26 = l_Lean_MessageData_ofName(x_7); +x_27 = l_Lean_Meta_Grind_foldProjs___lambda__2___closed__4; +lean_ctor_set_tag(x_17, 7); +lean_ctor_set(x_17, 1, x_26); +lean_ctor_set(x_17, 0, x_27); +x_28 = l_Lean_Meta_Grind_foldProjs___lambda__2___closed__6; +lean_ctor_set_tag(x_10, 7); +lean_ctor_set(x_10, 1, x_28); +lean_ctor_set(x_10, 0, x_17); +lean_inc(x_1); +x_29 = l_Lean_indentExpr(x_1); +x_30 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_30, 0, x_10); +lean_ctor_set(x_30, 1, x_29); +x_31 = l_Lean_Meta_Grind_foldProjs___lambda__2___closed__8; +x_32 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_32, 0, x_30); +lean_ctor_set(x_32, 1, x_31); +x_33 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_16, x_32, x_2, x_3, x_4, x_5, x_24); +x_34 = lean_ctor_get(x_33, 0); +lean_inc(x_34); +x_35 = lean_ctor_get(x_33, 1); +lean_inc(x_35); +lean_dec(x_33); +x_36 = l_Lean_Meta_Grind_foldProjs___lambda__1(x_1, x_34, x_2, x_3, x_4, x_5, x_35); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_34); +return x_36; +} +else { -lean_dec(x_19); +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; 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_17, 1); +lean_inc(x_37); +lean_dec(x_17); +x_38 = l_Lean_MessageData_ofName(x_7); +x_39 = l_Lean_Meta_Grind_foldProjs___lambda__2___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_Lean_Meta_Grind_foldProjs___lambda__2___closed__6; +lean_ctor_set_tag(x_10, 7); +lean_ctor_set(x_10, 1, x_41); +lean_ctor_set(x_10, 0, x_40); +lean_inc(x_1); +x_42 = l_Lean_indentExpr(x_1); +x_43 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_43, 0, x_10); +lean_ctor_set(x_43, 1, x_42); +x_44 = l_Lean_Meta_Grind_foldProjs___lambda__2___closed__8; +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_Meta_processPostponed_loop___spec__2(x_16, x_45, x_2, x_3, x_4, x_5, x_37); +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_Meta_Grind_foldProjs___lambda__1(x_1, x_47, x_2, x_3, x_4, x_5, x_48); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_47); +return x_49; +} +} +} +else +{ +uint8_t x_50; +lean_dec(x_7); +x_50 = !lean_is_exclusive(x_15); +if (x_50 == 0) +{ +lean_object* x_51; lean_object* x_52; lean_object* x_53; uint8_t x_54; +x_51 = lean_ctor_get(x_15, 0); +x_52 = lean_ctor_get(x_51, 1); +lean_inc(x_52); +lean_dec(x_51); +x_53 = lean_array_get_size(x_52); +x_54 = lean_nat_dec_lt(x_8, x_53); +lean_dec(x_53); +if (x_54 == 0) +{ +lean_object* x_55; lean_object* x_56; lean_object* x_57; uint8_t x_58; +lean_dec(x_52); lean_dec(x_9); +x_55 = l_Lean_Meta_Grind_foldProjs___lambda__2___closed__2; +x_56 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_55, x_2, x_3, x_4, x_5, x_13); +x_57 = lean_ctor_get(x_56, 0); +lean_inc(x_57); +x_58 = lean_unbox(x_57); +lean_dec(x_57); +if (x_58 == 0) +{ +lean_object* x_59; lean_object* x_60; lean_object* x_61; +lean_free_object(x_15); +lean_free_object(x_10); lean_dec(x_8); +x_59 = lean_ctor_get(x_56, 1); +lean_inc(x_59); +lean_dec(x_56); +x_60 = lean_box(0); +x_61 = l_Lean_Meta_Grind_foldProjs___lambda__1(x_1, x_60, x_2, x_3, x_4, x_5, x_59); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -lean_ctor_set_tag(x_15, 0); -lean_ctor_set(x_15, 0, x_1); -lean_ctor_set(x_10, 0, x_15); -return x_10; +return x_61; } else { -lean_object* x_22; lean_object* x_23; +uint8_t x_62; +x_62 = !lean_is_exclusive(x_56); +if (x_62 == 0) +{ +lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; 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_63 = lean_ctor_get(x_56, 1); +x_64 = lean_ctor_get(x_56, 0); +lean_dec(x_64); +x_65 = l___private_Init_Data_Repr_0__Nat_reprFast(x_8); +lean_ctor_set_tag(x_15, 3); +lean_ctor_set(x_15, 0, x_65); +x_66 = l_Lean_MessageData_ofFormat(x_15); +x_67 = l_Lean_Meta_Grind_foldProjs___lambda__2___closed__10; +lean_ctor_set_tag(x_56, 7); +lean_ctor_set(x_56, 1, x_66); +lean_ctor_set(x_56, 0, x_67); +x_68 = l_Lean_Meta_Grind_foldProjs___lambda__2___closed__12; +lean_ctor_set_tag(x_10, 7); +lean_ctor_set(x_10, 1, x_68); +lean_ctor_set(x_10, 0, x_56); +lean_inc(x_1); +x_69 = l_Lean_indentExpr(x_1); +x_70 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_70, 0, x_10); +lean_ctor_set(x_70, 1, x_69); +x_71 = l_Lean_Meta_Grind_foldProjs___lambda__2___closed__8; +x_72 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_72, 0, x_70); +lean_ctor_set(x_72, 1, x_71); +x_73 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_55, x_72, x_2, x_3, x_4, x_5, x_63); +x_74 = lean_ctor_get(x_73, 0); +lean_inc(x_74); +x_75 = lean_ctor_get(x_73, 1); +lean_inc(x_75); +lean_dec(x_73); +x_76 = l_Lean_Meta_Grind_foldProjs___lambda__1(x_1, x_74, x_2, x_3, x_4, x_5, x_75); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_74); +return x_76; +} +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; lean_object* x_88; lean_object* x_89; lean_object* x_90; +x_77 = lean_ctor_get(x_56, 1); +lean_inc(x_77); +lean_dec(x_56); +x_78 = l___private_Init_Data_Repr_0__Nat_reprFast(x_8); +lean_ctor_set_tag(x_15, 3); +lean_ctor_set(x_15, 0, x_78); +x_79 = l_Lean_MessageData_ofFormat(x_15); +x_80 = l_Lean_Meta_Grind_foldProjs___lambda__2___closed__10; +x_81 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_81, 0, x_80); +lean_ctor_set(x_81, 1, x_79); +x_82 = l_Lean_Meta_Grind_foldProjs___lambda__2___closed__12; +lean_ctor_set_tag(x_10, 7); +lean_ctor_set(x_10, 1, x_82); +lean_ctor_set(x_10, 0, x_81); +lean_inc(x_1); +x_83 = l_Lean_indentExpr(x_1); +x_84 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_84, 0, x_10); +lean_ctor_set(x_84, 1, x_83); +x_85 = l_Lean_Meta_Grind_foldProjs___lambda__2___closed__8; +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_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_55, x_86, x_2, x_3, x_4, x_5, x_77); +x_88 = lean_ctor_get(x_87, 0); +lean_inc(x_88); +x_89 = lean_ctor_get(x_87, 1); +lean_inc(x_89); +lean_dec(x_87); +x_90 = l_Lean_Meta_Grind_foldProjs___lambda__1(x_1, x_88, x_2, x_3, x_4, x_5, x_89); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_88); +return x_90; +} +} +} +else +{ +lean_object* x_91; lean_object* x_92; lean_free_object(x_10); lean_dec(x_1); -x_22 = lean_array_fget(x_19, x_8); +x_91 = lean_array_fget(x_52, x_8); lean_dec(x_8); -lean_dec(x_19); -x_23 = l_Lean_Meta_mkProjection(x_9, x_22, x_2, x_3, x_4, x_5, x_13); -if (lean_obj_tag(x_23) == 0) +lean_dec(x_52); +x_92 = l_Lean_Meta_mkProjection(x_9, x_91, x_2, x_3, x_4, x_5, x_13); +if (lean_obj_tag(x_92) == 0) { -uint8_t x_24; -x_24 = !lean_is_exclusive(x_23); -if (x_24 == 0) +uint8_t x_93; +x_93 = !lean_is_exclusive(x_92); +if (x_93 == 0) { -lean_object* x_25; -x_25 = lean_ctor_get(x_23, 0); +lean_object* x_94; +x_94 = lean_ctor_get(x_92, 0); lean_ctor_set_tag(x_15, 0); -lean_ctor_set(x_15, 0, x_25); -lean_ctor_set(x_23, 0, x_15); -return x_23; +lean_ctor_set(x_15, 0, x_94); +lean_ctor_set(x_92, 0, x_15); +return x_92; } else { -lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_26 = lean_ctor_get(x_23, 0); -x_27 = lean_ctor_get(x_23, 1); -lean_inc(x_27); -lean_inc(x_26); -lean_dec(x_23); +lean_object* x_95; lean_object* x_96; lean_object* x_97; +x_95 = lean_ctor_get(x_92, 0); +x_96 = lean_ctor_get(x_92, 1); +lean_inc(x_96); +lean_inc(x_95); +lean_dec(x_92); lean_ctor_set_tag(x_15, 0); -lean_ctor_set(x_15, 0, x_26); -x_28 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_28, 0, x_15); -lean_ctor_set(x_28, 1, x_27); -return x_28; +lean_ctor_set(x_15, 0, x_95); +x_97 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_97, 0, x_15); +lean_ctor_set(x_97, 1, x_96); +return x_97; } } else { -uint8_t x_29; +uint8_t x_98; lean_free_object(x_15); -x_29 = !lean_is_exclusive(x_23); -if (x_29 == 0) +x_98 = !lean_is_exclusive(x_92); +if (x_98 == 0) { -return x_23; +return x_92; } else { -lean_object* x_30; lean_object* x_31; lean_object* x_32; -x_30 = lean_ctor_get(x_23, 0); -x_31 = lean_ctor_get(x_23, 1); -lean_inc(x_31); -lean_inc(x_30); -lean_dec(x_23); -x_32 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_32, 0, x_30); -lean_ctor_set(x_32, 1, x_31); -return x_32; +lean_object* x_99; lean_object* x_100; lean_object* x_101; +x_99 = lean_ctor_get(x_92, 0); +x_100 = lean_ctor_get(x_92, 1); +lean_inc(x_100); +lean_inc(x_99); +lean_dec(x_92); +x_101 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_101, 0, x_99); +lean_ctor_set(x_101, 1, x_100); +return x_101; } } } } else { -lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; -x_33 = lean_ctor_get(x_15, 0); -lean_inc(x_33); +lean_object* x_102; lean_object* x_103; lean_object* x_104; uint8_t x_105; +x_102 = lean_ctor_get(x_15, 0); +lean_inc(x_102); lean_dec(x_15); -x_34 = lean_ctor_get(x_33, 1); -lean_inc(x_34); -lean_dec(x_33); -x_35 = lean_array_get_size(x_34); -x_36 = lean_nat_dec_lt(x_8, x_35); -lean_dec(x_35); -if (x_36 == 0) -{ -lean_object* x_37; -lean_dec(x_34); +x_103 = lean_ctor_get(x_102, 1); +lean_inc(x_103); +lean_dec(x_102); +x_104 = lean_array_get_size(x_103); +x_105 = lean_nat_dec_lt(x_8, x_104); +lean_dec(x_104); +if (x_105 == 0) +{ +lean_object* x_106; lean_object* x_107; lean_object* x_108; uint8_t x_109; +lean_dec(x_103); lean_dec(x_9); +x_106 = l_Lean_Meta_Grind_foldProjs___lambda__2___closed__2; +x_107 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_106, x_2, x_3, x_4, x_5, x_13); +x_108 = lean_ctor_get(x_107, 0); +lean_inc(x_108); +x_109 = lean_unbox(x_108); +lean_dec(x_108); +if (x_109 == 0) +{ +lean_object* x_110; lean_object* x_111; lean_object* x_112; +lean_free_object(x_10); lean_dec(x_8); +x_110 = lean_ctor_get(x_107, 1); +lean_inc(x_110); +lean_dec(x_107); +x_111 = lean_box(0); +x_112 = l_Lean_Meta_Grind_foldProjs___lambda__1(x_1, x_111, x_2, x_3, x_4, x_5, x_110); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_37 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_37, 0, x_1); -lean_ctor_set(x_10, 0, x_37); -return x_10; +return x_112; } else { -lean_object* x_38; lean_object* x_39; +lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; +x_113 = lean_ctor_get(x_107, 1); +lean_inc(x_113); +if (lean_is_exclusive(x_107)) { + lean_ctor_release(x_107, 0); + lean_ctor_release(x_107, 1); + x_114 = x_107; +} else { + lean_dec_ref(x_107); + x_114 = lean_box(0); +} +x_115 = l___private_Init_Data_Repr_0__Nat_reprFast(x_8); +x_116 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_116, 0, x_115); +x_117 = l_Lean_MessageData_ofFormat(x_116); +x_118 = l_Lean_Meta_Grind_foldProjs___lambda__2___closed__10; +if (lean_is_scalar(x_114)) { + x_119 = lean_alloc_ctor(7, 2, 0); +} else { + x_119 = x_114; + lean_ctor_set_tag(x_119, 7); +} +lean_ctor_set(x_119, 0, x_118); +lean_ctor_set(x_119, 1, x_117); +x_120 = l_Lean_Meta_Grind_foldProjs___lambda__2___closed__12; +lean_ctor_set_tag(x_10, 7); +lean_ctor_set(x_10, 1, x_120); +lean_ctor_set(x_10, 0, x_119); +lean_inc(x_1); +x_121 = l_Lean_indentExpr(x_1); +x_122 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_122, 0, x_10); +lean_ctor_set(x_122, 1, x_121); +x_123 = l_Lean_Meta_Grind_foldProjs___lambda__2___closed__8; +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_125 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_106, x_124, x_2, x_3, x_4, x_5, x_113); +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_Meta_Grind_foldProjs___lambda__1(x_1, x_126, x_2, x_3, x_4, x_5, x_127); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_126); +return x_128; +} +} +else +{ +lean_object* x_129; lean_object* x_130; lean_free_object(x_10); lean_dec(x_1); -x_38 = lean_array_fget(x_34, x_8); +x_129 = lean_array_fget(x_103, x_8); lean_dec(x_8); -lean_dec(x_34); -x_39 = l_Lean_Meta_mkProjection(x_9, x_38, x_2, x_3, x_4, x_5, x_13); -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; -x_40 = lean_ctor_get(x_39, 0); -lean_inc(x_40); -x_41 = lean_ctor_get(x_39, 1); -lean_inc(x_41); -if (lean_is_exclusive(x_39)) { - lean_ctor_release(x_39, 0); - lean_ctor_release(x_39, 1); - x_42 = x_39; +lean_dec(x_103); +x_130 = l_Lean_Meta_mkProjection(x_9, x_129, x_2, x_3, x_4, x_5, x_13); +if (lean_obj_tag(x_130) == 0) +{ +lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; +x_131 = lean_ctor_get(x_130, 0); +lean_inc(x_131); +x_132 = lean_ctor_get(x_130, 1); +lean_inc(x_132); +if (lean_is_exclusive(x_130)) { + lean_ctor_release(x_130, 0); + lean_ctor_release(x_130, 1); + x_133 = x_130; } else { - lean_dec_ref(x_39); - x_42 = lean_box(0); + lean_dec_ref(x_130); + x_133 = lean_box(0); } -x_43 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_43, 0, x_40); -if (lean_is_scalar(x_42)) { - x_44 = lean_alloc_ctor(0, 2, 0); +x_134 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_134, 0, x_131); +if (lean_is_scalar(x_133)) { + x_135 = lean_alloc_ctor(0, 2, 0); } else { - x_44 = x_42; + x_135 = x_133; } -lean_ctor_set(x_44, 0, x_43); -lean_ctor_set(x_44, 1, x_41); -return x_44; +lean_ctor_set(x_135, 0, x_134); +lean_ctor_set(x_135, 1, x_132); +return x_135; } else { -lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; -x_45 = lean_ctor_get(x_39, 0); -lean_inc(x_45); -x_46 = lean_ctor_get(x_39, 1); -lean_inc(x_46); -if (lean_is_exclusive(x_39)) { - lean_ctor_release(x_39, 0); - lean_ctor_release(x_39, 1); - x_47 = x_39; +lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; +x_136 = lean_ctor_get(x_130, 0); +lean_inc(x_136); +x_137 = lean_ctor_get(x_130, 1); +lean_inc(x_137); +if (lean_is_exclusive(x_130)) { + lean_ctor_release(x_130, 0); + lean_ctor_release(x_130, 1); + x_138 = x_130; } else { - lean_dec_ref(x_39); - x_47 = lean_box(0); + lean_dec_ref(x_130); + x_138 = lean_box(0); } -if (lean_is_scalar(x_47)) { - x_48 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_138)) { + x_139 = lean_alloc_ctor(1, 2, 0); } else { - x_48 = x_47; + x_139 = x_138; } -lean_ctor_set(x_48, 0, x_45); -lean_ctor_set(x_48, 1, x_46); -return x_48; +lean_ctor_set(x_139, 0, x_136); +lean_ctor_set(x_139, 1, x_137); +return x_139; } } } @@ -5416,136 +5808,263 @@ return x_48; } else { -lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; -x_49 = lean_ctor_get(x_10, 0); -x_50 = lean_ctor_get(x_10, 1); -lean_inc(x_50); -lean_inc(x_49); +lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; +x_140 = lean_ctor_get(x_10, 0); +x_141 = lean_ctor_get(x_10, 1); +lean_inc(x_141); +lean_inc(x_140); lean_dec(x_10); -x_51 = lean_ctor_get(x_49, 0); -lean_inc(x_51); -lean_dec(x_49); -x_52 = l_Lean_getStructureInfo_x3f(x_51, x_7); -lean_dec(x_51); -if (lean_obj_tag(x_52) == 0) +x_142 = lean_ctor_get(x_140, 0); +lean_inc(x_142); +lean_dec(x_140); +lean_inc(x_7); +x_143 = l_Lean_getStructureInfo_x3f(x_142, x_7); +lean_dec(x_142); +if (lean_obj_tag(x_143) == 0) { -lean_object* x_53; lean_object* x_54; +lean_object* x_144; lean_object* x_145; lean_object* x_146; uint8_t x_147; lean_dec(x_9); lean_dec(x_8); +x_144 = l_Lean_Meta_Grind_foldProjs___lambda__2___closed__2; +x_145 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_144, x_2, x_3, x_4, x_5, x_141); +x_146 = lean_ctor_get(x_145, 0); +lean_inc(x_146); +x_147 = lean_unbox(x_146); +lean_dec(x_146); +if (x_147 == 0) +{ +lean_object* x_148; lean_object* x_149; lean_object* x_150; +lean_dec(x_7); +x_148 = lean_ctor_get(x_145, 1); +lean_inc(x_148); +lean_dec(x_145); +x_149 = lean_box(0); +x_150 = l_Lean_Meta_Grind_foldProjs___lambda__1(x_1, x_149, x_2, x_3, x_4, x_5, x_148); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_53 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_53, 0, x_1); -x_54 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_54, 0, x_53); -lean_ctor_set(x_54, 1, x_50); -return x_54; +return x_150; } else { -lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; uint8_t x_59; -x_55 = lean_ctor_get(x_52, 0); -lean_inc(x_55); -if (lean_is_exclusive(x_52)) { - lean_ctor_release(x_52, 0); - x_56 = 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; 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_151 = lean_ctor_get(x_145, 1); +lean_inc(x_151); +if (lean_is_exclusive(x_145)) { + lean_ctor_release(x_145, 0); + lean_ctor_release(x_145, 1); + x_152 = x_145; } else { - lean_dec_ref(x_52); - x_56 = lean_box(0); + lean_dec_ref(x_145); + x_152 = lean_box(0); } -x_57 = lean_ctor_get(x_55, 1); -lean_inc(x_57); -lean_dec(x_55); -x_58 = lean_array_get_size(x_57); -x_59 = lean_nat_dec_lt(x_8, x_58); -lean_dec(x_58); -if (x_59 == 0) +x_153 = l_Lean_MessageData_ofName(x_7); +x_154 = l_Lean_Meta_Grind_foldProjs___lambda__2___closed__4; +if (lean_is_scalar(x_152)) { + x_155 = lean_alloc_ctor(7, 2, 0); +} else { + x_155 = x_152; + lean_ctor_set_tag(x_155, 7); +} +lean_ctor_set(x_155, 0, x_154); +lean_ctor_set(x_155, 1, x_153); +x_156 = l_Lean_Meta_Grind_foldProjs___lambda__2___closed__6; +x_157 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_157, 0, x_155); +lean_ctor_set(x_157, 1, x_156); +lean_inc(x_1); +x_158 = l_Lean_indentExpr(x_1); +x_159 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_159, 0, x_157); +lean_ctor_set(x_159, 1, x_158); +x_160 = l_Lean_Meta_Grind_foldProjs___lambda__2___closed__8; +x_161 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_161, 0, x_159); +lean_ctor_set(x_161, 1, x_160); +x_162 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_144, x_161, x_2, x_3, x_4, x_5, x_151); +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_Meta_Grind_foldProjs___lambda__1(x_1, x_163, x_2, x_3, x_4, x_5, x_164); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_163); +return x_165; +} +} +else { -lean_object* x_60; lean_object* x_61; -lean_dec(x_57); +lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; uint8_t x_170; +lean_dec(x_7); +x_166 = lean_ctor_get(x_143, 0); +lean_inc(x_166); +if (lean_is_exclusive(x_143)) { + lean_ctor_release(x_143, 0); + x_167 = x_143; +} else { + lean_dec_ref(x_143); + x_167 = lean_box(0); +} +x_168 = lean_ctor_get(x_166, 1); +lean_inc(x_168); +lean_dec(x_166); +x_169 = lean_array_get_size(x_168); +x_170 = lean_nat_dec_lt(x_8, x_169); +lean_dec(x_169); +if (x_170 == 0) +{ +lean_object* x_171; lean_object* x_172; lean_object* x_173; uint8_t x_174; +lean_dec(x_168); lean_dec(x_9); +x_171 = l_Lean_Meta_Grind_foldProjs___lambda__2___closed__2; +x_172 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_171, x_2, x_3, x_4, x_5, x_141); +x_173 = lean_ctor_get(x_172, 0); +lean_inc(x_173); +x_174 = lean_unbox(x_173); +lean_dec(x_173); +if (x_174 == 0) +{ +lean_object* x_175; lean_object* x_176; lean_object* x_177; +lean_dec(x_167); lean_dec(x_8); +x_175 = lean_ctor_get(x_172, 1); +lean_inc(x_175); +lean_dec(x_172); +x_176 = lean_box(0); +x_177 = l_Lean_Meta_Grind_foldProjs___lambda__1(x_1, x_176, x_2, x_3, x_4, x_5, x_175); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -if (lean_is_scalar(x_56)) { - x_60 = lean_alloc_ctor(0, 1, 0); +return x_177; +} +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; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; +x_178 = lean_ctor_get(x_172, 1); +lean_inc(x_178); +if (lean_is_exclusive(x_172)) { + lean_ctor_release(x_172, 0); + lean_ctor_release(x_172, 1); + x_179 = x_172; } else { - x_60 = x_56; - lean_ctor_set_tag(x_60, 0); + lean_dec_ref(x_172); + x_179 = lean_box(0); +} +x_180 = l___private_Init_Data_Repr_0__Nat_reprFast(x_8); +if (lean_is_scalar(x_167)) { + x_181 = lean_alloc_ctor(3, 1, 0); +} else { + x_181 = x_167; + lean_ctor_set_tag(x_181, 3); +} +lean_ctor_set(x_181, 0, x_180); +x_182 = l_Lean_MessageData_ofFormat(x_181); +x_183 = l_Lean_Meta_Grind_foldProjs___lambda__2___closed__10; +if (lean_is_scalar(x_179)) { + x_184 = lean_alloc_ctor(7, 2, 0); +} else { + x_184 = x_179; + lean_ctor_set_tag(x_184, 7); +} +lean_ctor_set(x_184, 0, x_183); +lean_ctor_set(x_184, 1, x_182); +x_185 = l_Lean_Meta_Grind_foldProjs___lambda__2___closed__12; +x_186 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_186, 0, x_184); +lean_ctor_set(x_186, 1, x_185); +lean_inc(x_1); +x_187 = l_Lean_indentExpr(x_1); +x_188 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_188, 0, x_186); +lean_ctor_set(x_188, 1, x_187); +x_189 = l_Lean_Meta_Grind_foldProjs___lambda__2___closed__8; +x_190 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_190, 0, x_188); +lean_ctor_set(x_190, 1, x_189); +x_191 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_171, x_190, x_2, x_3, x_4, x_5, x_178); +x_192 = lean_ctor_get(x_191, 0); +lean_inc(x_192); +x_193 = lean_ctor_get(x_191, 1); +lean_inc(x_193); +lean_dec(x_191); +x_194 = l_Lean_Meta_Grind_foldProjs___lambda__1(x_1, x_192, x_2, x_3, x_4, x_5, x_193); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_192); +return x_194; } -lean_ctor_set(x_60, 0, x_1); -x_61 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_61, 0, x_60); -lean_ctor_set(x_61, 1, x_50); -return x_61; } else { -lean_object* x_62; lean_object* x_63; +lean_object* x_195; lean_object* x_196; lean_dec(x_1); -x_62 = lean_array_fget(x_57, x_8); +x_195 = lean_array_fget(x_168, x_8); lean_dec(x_8); -lean_dec(x_57); -x_63 = l_Lean_Meta_mkProjection(x_9, x_62, x_2, x_3, x_4, x_5, x_50); -if (lean_obj_tag(x_63) == 0) -{ -lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; -x_64 = lean_ctor_get(x_63, 0); -lean_inc(x_64); -x_65 = lean_ctor_get(x_63, 1); -lean_inc(x_65); -if (lean_is_exclusive(x_63)) { - lean_ctor_release(x_63, 0); - lean_ctor_release(x_63, 1); - x_66 = x_63; +lean_dec(x_168); +x_196 = l_Lean_Meta_mkProjection(x_9, x_195, x_2, x_3, x_4, x_5, x_141); +if (lean_obj_tag(x_196) == 0) +{ +lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; +x_197 = lean_ctor_get(x_196, 0); +lean_inc(x_197); +x_198 = lean_ctor_get(x_196, 1); +lean_inc(x_198); +if (lean_is_exclusive(x_196)) { + lean_ctor_release(x_196, 0); + lean_ctor_release(x_196, 1); + x_199 = x_196; } else { - lean_dec_ref(x_63); - x_66 = lean_box(0); + lean_dec_ref(x_196); + x_199 = lean_box(0); } -if (lean_is_scalar(x_56)) { - x_67 = lean_alloc_ctor(0, 1, 0); +if (lean_is_scalar(x_167)) { + x_200 = lean_alloc_ctor(0, 1, 0); } else { - x_67 = x_56; - lean_ctor_set_tag(x_67, 0); + x_200 = x_167; + lean_ctor_set_tag(x_200, 0); } -lean_ctor_set(x_67, 0, x_64); -if (lean_is_scalar(x_66)) { - x_68 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_200, 0, x_197); +if (lean_is_scalar(x_199)) { + x_201 = lean_alloc_ctor(0, 2, 0); } else { - x_68 = x_66; + x_201 = x_199; } -lean_ctor_set(x_68, 0, x_67); -lean_ctor_set(x_68, 1, x_65); -return x_68; +lean_ctor_set(x_201, 0, x_200); +lean_ctor_set(x_201, 1, x_198); +return x_201; } else { -lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; -lean_dec(x_56); -x_69 = lean_ctor_get(x_63, 0); -lean_inc(x_69); -x_70 = lean_ctor_get(x_63, 1); -lean_inc(x_70); -if (lean_is_exclusive(x_63)) { - lean_ctor_release(x_63, 0); - lean_ctor_release(x_63, 1); - x_71 = x_63; +lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; +lean_dec(x_167); +x_202 = lean_ctor_get(x_196, 0); +lean_inc(x_202); +x_203 = lean_ctor_get(x_196, 1); +lean_inc(x_203); +if (lean_is_exclusive(x_196)) { + lean_ctor_release(x_196, 0); + lean_ctor_release(x_196, 1); + x_204 = x_196; } else { - lean_dec_ref(x_63); - x_71 = lean_box(0); + lean_dec_ref(x_196); + x_204 = lean_box(0); } -if (lean_is_scalar(x_71)) { - x_72 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_204)) { + x_205 = lean_alloc_ctor(1, 2, 0); } else { - x_72 = x_71; + x_205 = x_204; } -lean_ctor_set(x_72, 0, x_69); -lean_ctor_set(x_72, 1, x_70); -return x_72; +lean_ctor_set(x_205, 0, x_202); +lean_ctor_set(x_205, 1, x_203); +return x_205; } } } @@ -5553,21 +6072,21 @@ return x_72; } else { -lean_object* x_73; lean_object* x_74; +lean_object* x_206; lean_object* x_207; lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_73 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_73, 0, x_1); -x_74 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_74, 0, x_73); -lean_ctor_set(x_74, 1, x_6); -return x_74; +x_206 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_206, 0, x_1); +x_207 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_207, 0, x_206); +lean_ctor_set(x_207, 1, x_6); +return x_207; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_foldProjs___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_foldProjs___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { lean_object* x_7; lean_object* x_8; @@ -5582,7 +6101,7 @@ static lean_object* _init_l_Lean_Meta_Grind_foldProjs___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_foldProjs___lambda__2___boxed), 6, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_foldProjs___lambda__3___boxed), 6, 0); return x_1; } } @@ -5590,7 +6109,7 @@ static lean_object* _init_l_Lean_Meta_Grind_foldProjs___closed__2() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_foldProjs___lambda__1), 6, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_foldProjs___lambda__2), 6, 0); return x_1; } } @@ -5605,11 +6124,24 @@ x_10 = l_Lean_Meta_transform___at_Lean_Meta_zetaReduce___spec__1(x_1, x_7, x_8, return x_10; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_foldProjs___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_EXPORT lean_object* l_Lean_Meta_Grind_foldProjs___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; +x_8 = l_Lean_Meta_Grind_foldProjs___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_foldProjs___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { lean_object* x_7; -x_7 = l_Lean_Meta_Grind_foldProjs___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6); +x_7 = l_Lean_Meta_Grind_foldProjs___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); @@ -5784,6 +6316,14 @@ lean_dec(x_2); return x_5; } } +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_normalize___boxed(lean_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_grind_normalize(x_1, x_2, x_3, x_4, x_5, x_6); +return x_7; +} +} lean_object* initialize_Lean_Meta_AbstractNestedProofs(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_Transform(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_Tactic_Util(uint8_t builtin, lean_object*); @@ -5897,6 +6437,30 @@ l_Lean_Meta_Grind_eraseIrrelevantMData___closed__1 = _init_l_Lean_Meta_Grind_era lean_mark_persistent(l_Lean_Meta_Grind_eraseIrrelevantMData___closed__1); l_Lean_Meta_Grind_eraseIrrelevantMData___closed__2 = _init_l_Lean_Meta_Grind_eraseIrrelevantMData___closed__2(); lean_mark_persistent(l_Lean_Meta_Grind_eraseIrrelevantMData___closed__2); +l_Lean_Meta_Grind_foldProjs___lambda__2___closed__1 = _init_l_Lean_Meta_Grind_foldProjs___lambda__2___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_foldProjs___lambda__2___closed__1); +l_Lean_Meta_Grind_foldProjs___lambda__2___closed__2 = _init_l_Lean_Meta_Grind_foldProjs___lambda__2___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_foldProjs___lambda__2___closed__2); +l_Lean_Meta_Grind_foldProjs___lambda__2___closed__3 = _init_l_Lean_Meta_Grind_foldProjs___lambda__2___closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind_foldProjs___lambda__2___closed__3); +l_Lean_Meta_Grind_foldProjs___lambda__2___closed__4 = _init_l_Lean_Meta_Grind_foldProjs___lambda__2___closed__4(); +lean_mark_persistent(l_Lean_Meta_Grind_foldProjs___lambda__2___closed__4); +l_Lean_Meta_Grind_foldProjs___lambda__2___closed__5 = _init_l_Lean_Meta_Grind_foldProjs___lambda__2___closed__5(); +lean_mark_persistent(l_Lean_Meta_Grind_foldProjs___lambda__2___closed__5); +l_Lean_Meta_Grind_foldProjs___lambda__2___closed__6 = _init_l_Lean_Meta_Grind_foldProjs___lambda__2___closed__6(); +lean_mark_persistent(l_Lean_Meta_Grind_foldProjs___lambda__2___closed__6); +l_Lean_Meta_Grind_foldProjs___lambda__2___closed__7 = _init_l_Lean_Meta_Grind_foldProjs___lambda__2___closed__7(); +lean_mark_persistent(l_Lean_Meta_Grind_foldProjs___lambda__2___closed__7); +l_Lean_Meta_Grind_foldProjs___lambda__2___closed__8 = _init_l_Lean_Meta_Grind_foldProjs___lambda__2___closed__8(); +lean_mark_persistent(l_Lean_Meta_Grind_foldProjs___lambda__2___closed__8); +l_Lean_Meta_Grind_foldProjs___lambda__2___closed__9 = _init_l_Lean_Meta_Grind_foldProjs___lambda__2___closed__9(); +lean_mark_persistent(l_Lean_Meta_Grind_foldProjs___lambda__2___closed__9); +l_Lean_Meta_Grind_foldProjs___lambda__2___closed__10 = _init_l_Lean_Meta_Grind_foldProjs___lambda__2___closed__10(); +lean_mark_persistent(l_Lean_Meta_Grind_foldProjs___lambda__2___closed__10); +l_Lean_Meta_Grind_foldProjs___lambda__2___closed__11 = _init_l_Lean_Meta_Grind_foldProjs___lambda__2___closed__11(); +lean_mark_persistent(l_Lean_Meta_Grind_foldProjs___lambda__2___closed__11); +l_Lean_Meta_Grind_foldProjs___lambda__2___closed__12 = _init_l_Lean_Meta_Grind_foldProjs___lambda__2___closed__12(); +lean_mark_persistent(l_Lean_Meta_Grind_foldProjs___lambda__2___closed__12); l_Lean_Meta_Grind_foldProjs___closed__1 = _init_l_Lean_Meta_Grind_foldProjs___closed__1(); lean_mark_persistent(l_Lean_Meta_Grind_foldProjs___closed__1); l_Lean_Meta_Grind_foldProjs___closed__2 = _init_l_Lean_Meta_Grind_foldProjs___closed__2(); diff --git a/stage0/stdlib/Lean/Meta/Tactic/TryThis.c b/stage0/stdlib/Lean/Meta/Tactic/TryThis.c index f512a2673ff1..cbe7a9067659 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/TryThis.c +++ b/stage0/stdlib/Lean/Meta/Tactic/TryThis.c @@ -54,6 +54,7 @@ lean_object* l_Lean_Json_mkObj(lean_object*); static lean_object* l_Lean_Meta_Tactic_TryThis_addRewriteSuggestion___lambda__2___closed__10; static lean_object* l_Lean_Meta_Tactic_TryThis_addHaveSuggestion___closed__4; LEAN_EXPORT lean_object* l_Lean_Meta_Tactic_TryThis_tryThisProvider___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(lean_object*); static lean_object* l_Lean_Meta_Tactic_TryThis_addHaveSuggestion___closed__7; static lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_Tactic_TryThis_addRewriteSuggestion___spec__1___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_TryThis_0__Lean_Meta_Tactic_TryThis_addSuggestionCore___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -159,7 +160,6 @@ lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_logAt___at_Lean_Elab_Term_reportUnsolvedGoals___spec__2(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Tactic_TryThis_addHaveSuggestion___closed__2; static lean_object* l_Lean_Meta_Tactic_TryThis_addHaveSuggestion___closed__34; -lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(size_t, size_t, lean_object*); static lean_object* l_Lean_Meta_Tactic_TryThis_delabToRefinableSuggestion___closed__2; static lean_object* l_Lean_Meta_Tactic_TryThis_SuggestionStyle_error___closed__12; static lean_object* l_Lean_Meta_Tactic_TryThis_addHaveSuggestion___closed__16; @@ -414,6 +414,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Tactic_TryThis_instCoeHeadTSyntaxConsSyntax lean_object* lean_array_uget(lean_object*, size_t); size_t lean_array_size(lean_object*); double round(double); +lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(size_t, size_t, lean_object*); lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Tactic_TryThis_addHaveSuggestion___closed__17; static lean_object* l___private_Lean_Meta_Tactic_TryThis_0__Lean_Meta_Tactic_TryThis_addSuggestionCore___closed__5; @@ -421,7 +422,6 @@ lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_o static lean_object* l_Lean_Meta_Tactic_TryThis_addRewriteSuggestion___closed__13; static lean_object* l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_608____closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_Tactic_TryThis_addSuggestion___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_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(lean_object*); static lean_object* l_Lean_Meta_Tactic_TryThis_addRewriteSuggestion___closed__5; lean_object* lean_string_append(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Tactic_TryThis_Suggestion_toJsonAndInfoM___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -3869,13 +3869,13 @@ x_42 = lean_ctor_get(x_40, 1); x_43 = lean_ctor_get(x_40, 0); lean_dec(x_43); x_44 = lean_array_size(x_30); -x_45 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_44, x_24, x_30); +x_45 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_44, x_24, x_30); x_46 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_46, 0, x_45); x_47 = l___private_Lean_Meta_Tactic_TryThis_0__Lean_Meta_Tactic_TryThis_addSuggestionCore___closed__1; lean_ctor_set(x_40, 1, x_46); lean_ctor_set(x_40, 0, x_47); -x_48 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_32); +x_48 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_32); x_49 = l___private_Lean_Meta_Tactic_TryThis_0__Lean_Meta_Tactic_TryThis_addSuggestionCore___closed__2; x_50 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_50, 0, x_49); @@ -3959,14 +3959,14 @@ x_79 = lean_ctor_get(x_40, 1); lean_inc(x_79); lean_dec(x_40); x_80 = lean_array_size(x_30); -x_81 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_80, x_24, x_30); +x_81 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_80, x_24, x_30); x_82 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_82, 0, x_81); x_83 = l___private_Lean_Meta_Tactic_TryThis_0__Lean_Meta_Tactic_TryThis_addSuggestionCore___closed__1; x_84 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_84, 0, x_83); lean_ctor_set(x_84, 1, x_82); -x_85 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_32); +x_85 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_32); x_86 = l___private_Lean_Meta_Tactic_TryThis_0__Lean_Meta_Tactic_TryThis_addSuggestionCore___closed__2; x_87 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_87, 0, x_86); @@ -4155,7 +4155,7 @@ if (lean_is_exclusive(x_142)) { x_144 = lean_box(0); } x_145 = lean_array_size(x_131); -x_146 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_145, x_125, x_131); +x_146 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_145, x_125, x_131); x_147 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_147, 0, x_146); x_148 = l___private_Lean_Meta_Tactic_TryThis_0__Lean_Meta_Tactic_TryThis_addSuggestionCore___closed__1; @@ -4166,7 +4166,7 @@ if (lean_is_scalar(x_144)) { } lean_ctor_set(x_149, 0, x_148); lean_ctor_set(x_149, 1, x_147); -x_150 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_133); +x_150 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_133); x_151 = l___private_Lean_Meta_Tactic_TryThis_0__Lean_Meta_Tactic_TryThis_addSuggestionCore___closed__2; x_152 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_152, 0, x_151); diff --git a/stage0/stdlib/Lean/PrettyPrinter/Delaborator/Builtins.c b/stage0/stdlib/Lean/PrettyPrinter/Delaborator/Builtins.c index 740777c4bdeb..f8fd693a7269 100644 --- a/stage0/stdlib/Lean/PrettyPrinter/Delaborator/Builtins.c +++ b/stage0/stdlib/Lean/PrettyPrinter/Delaborator/Builtins.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Lean.PrettyPrinter.Delaborator.Builtins -// Imports: Lean.Parser Lean.PrettyPrinter.Delaborator.Attributes Lean.PrettyPrinter.Delaborator.Basic Lean.PrettyPrinter.Delaborator.SubExpr Lean.PrettyPrinter.Delaborator.TopDownAnalyze Lean.Meta.CoeAttr +// Imports: Lean.PrettyPrinter.Delaborator.Attributes Lean.PrettyPrinter.Delaborator.Basic Lean.PrettyPrinter.Delaborator.SubExpr Lean.PrettyPrinter.Delaborator.TopDownAnalyze Lean.Meta.CoeAttr #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -2390,7 +2390,7 @@ static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabFVar___closed__4 lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_PrettyPrinter_Delaborator_delabFVar___closed__1; x_2 = l_Lean_PrettyPrinter_Delaborator_delabFVar___closed__2; -x_3 = lean_unsigned_to_nat(40u); +x_3 = lean_unsigned_to_nat(39u); x_4 = lean_unsigned_to_nat(35u); 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); @@ -2756,7 +2756,7 @@ static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabBVar___closed__3 lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_PrettyPrinter_Delaborator_delabFVar___closed__1; x_2 = l_Lean_PrettyPrinter_Delaborator_delabBVar___closed__2; -x_3 = lean_unsigned_to_nat(51u); +x_3 = lean_unsigned_to_nat(50u); x_4 = lean_unsigned_to_nat(32u); x_5 = l_Lean_PrettyPrinter_Delaborator_delabFVar___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -3427,7 +3427,7 @@ static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabMVar___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_delabMVar___closed__1; -x_3 = lean_unsigned_to_nat(71u); +x_3 = lean_unsigned_to_nat(70u); x_4 = lean_unsigned_to_nat(30u); 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); @@ -3539,7 +3539,7 @@ static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabSort___closed__2 lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_PrettyPrinter_Delaborator_delabFVar___closed__1; x_2 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__1; -x_3 = lean_unsigned_to_nat(76u); +x_3 = lean_unsigned_to_nat(75u); x_4 = lean_unsigned_to_nat(30u); 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); @@ -5850,7 +5850,7 @@ static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabConst___closed__ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_PrettyPrinter_Delaborator_delabFVar___closed__1; x_2 = l_Lean_PrettyPrinter_Delaborator_delabConst___closed__1; -x_3 = lean_unsigned_to_nat(93u); +x_3 = lean_unsigned_to_nat(92u); x_4 = lean_unsigned_to_nat(35u); 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); @@ -17262,7 +17262,7 @@ static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_ 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(479u); +x_3 = lean_unsigned_to_nat(478u); 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); @@ -26039,7 +26039,7 @@ static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambd 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(710u); +x_3 = lean_unsigned_to_nat(709u); 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); @@ -35371,7 +35371,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(960u); +x_3 = lean_unsigned_to_nat(959u); 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); @@ -36247,7 +36247,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(980u); +x_3 = lean_unsigned_to_nat(979u); 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); @@ -39387,7 +39387,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(1088u); +x_3 = lean_unsigned_to_nat(1087u); 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); @@ -44529,7 +44529,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(1224u); +x_3 = lean_unsigned_to_nat(1223u); 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); @@ -53423,7 +53423,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(1378u); +x_3 = lean_unsigned_to_nat(1377u); 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); @@ -54414,7 +54414,6 @@ x_10 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature(x_9, x_2, x_3, x return x_10; } } -lean_object* initialize_Lean_Parser(uint8_t builtin, lean_object*); lean_object* initialize_Lean_PrettyPrinter_Delaborator_Attributes(uint8_t builtin, lean_object*); lean_object* initialize_Lean_PrettyPrinter_Delaborator_Basic(uint8_t builtin, lean_object*); lean_object* initialize_Lean_PrettyPrinter_Delaborator_SubExpr(uint8_t builtin, lean_object*); @@ -54425,9 +54424,6 @@ LEAN_EXPORT lean_object* initialize_Lean_PrettyPrinter_Delaborator_Builtins(uint lean_object * res; if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); _G_initialized = true; -res = initialize_Lean_Parser(builtin, lean_io_mk_world()); -if (lean_io_result_is_error(res)) return res; -lean_dec_ref(res); res = initialize_Lean_PrettyPrinter_Delaborator_Attributes(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); diff --git a/stage0/stdlib/Lean/Server/CodeActions/Basic.c b/stage0/stdlib/Lean/Server/CodeActions/Basic.c index 0f8b24763c5e..d628a224785e 100644 --- a/stage0/stdlib/Lean/Server/CodeActions/Basic.c +++ b/stage0/stdlib/Lean/Server/CodeActions/Basic.c @@ -41,6 +41,7 @@ lean_object* l_Lean_Json_mkObj(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_initFn____x40_Lean_Server_CodeActions_Basic___hyg_517____lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_CodeAction_getFileSource_x21___closed__2; lean_object* l_Lean_ConstantInfo_type(lean_object*); +lean_object* l_List_flatMapTR_go___at___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Server_CodeActions_Basic_0__Lean_Server_fromJsonCodeActionResolveData____x40_Lean_Server_CodeActions_Basic___hyg_125____closed__11; LEAN_EXPORT lean_object* l_Lean_Server_instFromJsonCodeActionResolveData; static lean_object* l_Lean_Server_initFn____x40_Lean_Server_CodeActions_Basic___hyg_517____closed__20; @@ -92,6 +93,7 @@ LEAN_EXPORT lean_object* l_Lean_Server_handleCodeActionResolve(lean_object*, lea LEAN_EXPORT lean_object* l_Array_mapFinIdxM_map___at_Lean_Server_handleCodeAction___spec__9___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_initFn____x40_Lean_Server_CodeActions_Basic___hyg_458____closed__2; lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_Quotation_withNewLocals___spec__1(lean_object*, size_t, size_t, lean_object*); +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Server_instCodeActionProviderInhabited___closed__1; static lean_object* l_Lean_Server_initFn____x40_Lean_Server_CodeActions_Basic___hyg_517____lambda__2___closed__3; LEAN_EXPORT lean_object* l_Lean_Server_initFn____x40_Lean_Server_CodeActions_Basic___hyg_458_(lean_object*); @@ -145,14 +147,12 @@ static lean_object* l_Lean_Server_initFn____x40_Lean_Server_CodeActions_Basic___ lean_object* lean_st_ref_get(lean_object*, lean_object*); lean_object* l_Lean_throwError___at_Lean_registerTagAttribute___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_instCodeActionProviderInhabited___lambda__1(lean_object*); -lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____spec__1(lean_object*, lean_object*); lean_object* lean_st_mk_ref(lean_object*, lean_object*); static lean_object* l___private_Lean_Server_CodeActions_Basic_0__Lean_Server_fromJsonCodeActionResolveData____x40_Lean_Server_CodeActions_Basic___hyg_125____closed__10; 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_EXPORT lean_object* l_Lean_Server_initFn____x40_Lean_Server_CodeActions_Basic___hyg_517____lambda__3(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_handleCodeAction(lean_object*, lean_object*, lean_object*); -lean_object* l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Server_initFn____x40_Lean_Server_CodeActions_Basic___hyg_517____closed__15; static lean_object* l___private_Lean_Server_CodeActions_Basic_0__Lean_Server_fromJsonCodeActionResolveData____x40_Lean_Server_CodeActions_Basic___hyg_125____closed__7; LEAN_EXPORT lean_object* l_Lean_Server_initFn____x40_Lean_Server_CodeActions_Basic___hyg_412_(lean_object*); @@ -423,7 +423,7 @@ x_24 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_24, 0, x_7); lean_ctor_set(x_24, 1, x_23); x_25 = l___private_Lean_Server_CodeActions_Basic_0__Lean_Server_toJsonCodeActionResolveData____x40_Lean_Server_CodeActions_Basic___hyg_59____closed__5; -x_26 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_24, x_25); +x_26 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____spec__1(x_24, x_25); x_27 = l_Lean_Json_mkObj(x_26); return x_27; } @@ -740,7 +740,7 @@ x_23 = lean_ctor_get(x_14, 0); lean_inc(x_23); lean_dec(x_14); x_24 = l___private_Lean_Server_CodeActions_Basic_0__Lean_Server_toJsonCodeActionResolveData____x40_Lean_Server_CodeActions_Basic___hyg_59____closed__4; -x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____spec__1(x_1, x_24); +x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____spec__1(x_1, x_24); if (lean_obj_tag(x_25) == 0) { uint8_t x_26; diff --git a/stage0/stdlib/Lean/Server/Completion.c b/stage0/stdlib/Lean/Server/Completion.c index db1af91e4c1d..1418b57b0551 100644 --- a/stage0/stdlib/Lean/Server/Completion.c +++ b/stage0/stdlib/Lean/Server/Completion.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Lean.Server.Completion -// Imports: Lean.Server.Completion.CompletionCollectors +// Imports: Lean.Server.Completion.CompletionCollectors Std.Data.HashMap #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -42,7 +42,6 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Co LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___lambda__1___boxed(lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); static lean_object* l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__20; -uint64_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6462_(lean_object*); lean_object* lean_mk_array(lean_object*, lean_object*); uint8_t lean_usize_dec_eq(size_t, size_t); static lean_object* l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__2; @@ -130,6 +129,7 @@ static lean_object* l_Array_groupByKey___at___private_Lean_Server_Completion_0__ lean_object* l_Array_append___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__16___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__17(size_t, size_t, lean_object*); static uint64_t l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__7___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__8___closed__1; +uint64_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6277_(lean_object*); extern lean_object* l_Lean_Lsp_instHashableCompletionItemKind; static lean_object* l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__27; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Server_Completion_find_x3f___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -843,7 +843,7 @@ x_40 = lean_ctor_get(x_34, 0); lean_inc(x_40); lean_dec(x_34); x_41 = 11; -x_42 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6462_(x_40); +x_42 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6277_(x_40); lean_dec(x_40); x_43 = 13; x_44 = lean_uint64_mix_hash(x_42, x_43); @@ -888,7 +888,7 @@ lean_object* x_61; uint64_t x_62; uint64_t x_63; uint64_t x_64; uint64_t x_65; u x_61 = lean_ctor_get(x_34, 0); lean_inc(x_61); lean_dec(x_34); -x_62 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6462_(x_61); +x_62 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6277_(x_61); lean_dec(x_61); x_63 = lean_uint64_mix_hash(x_62, x_55); x_64 = l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__7___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__8___closed__2; @@ -926,7 +926,7 @@ lean_object* x_76; uint64_t x_77; uint64_t x_78; uint64_t x_79; uint64_t x_80; u x_76 = lean_ctor_get(x_34, 0); lean_inc(x_76); lean_dec(x_34); -x_77 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6462_(x_76); +x_77 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6277_(x_76); lean_dec(x_76); x_78 = lean_uint64_mix_hash(x_77, x_55); x_79 = l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__7___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__8___closed__2; @@ -966,7 +966,7 @@ lean_object* x_95; uint64_t x_96; uint64_t x_97; uint64_t x_98; uint64_t x_99; u x_95 = lean_ctor_get(x_34, 0); lean_inc(x_95); lean_dec(x_34); -x_96 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6462_(x_95); +x_96 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6277_(x_95); lean_dec(x_95); x_97 = lean_uint64_mix_hash(x_96, x_55); x_98 = lean_uint64_mix_hash(x_88, x_97); @@ -2193,7 +2193,7 @@ x_146 = lean_ctor_get(x_140, 0); lean_inc(x_146); lean_dec(x_140); x_147 = 11; -x_148 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6462_(x_146); +x_148 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6277_(x_146); lean_dec(x_146); x_149 = 13; x_150 = lean_uint64_mix_hash(x_148, x_149); @@ -2238,7 +2238,7 @@ lean_object* x_167; uint64_t x_168; uint64_t x_169; uint64_t x_170; uint64_t x_1 x_167 = lean_ctor_get(x_140, 0); lean_inc(x_167); lean_dec(x_140); -x_168 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6462_(x_167); +x_168 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6277_(x_167); lean_dec(x_167); x_169 = lean_uint64_mix_hash(x_168, x_161); x_170 = l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__7___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__8___closed__2; @@ -2276,7 +2276,7 @@ lean_object* x_182; uint64_t x_183; uint64_t x_184; uint64_t x_185; uint64_t x_1 x_182 = lean_ctor_get(x_140, 0); lean_inc(x_182); lean_dec(x_140); -x_183 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6462_(x_182); +x_183 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6277_(x_182); lean_dec(x_182); x_184 = lean_uint64_mix_hash(x_183, x_161); x_185 = l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__7___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__8___closed__2; @@ -2316,7 +2316,7 @@ lean_object* x_201; uint64_t x_202; uint64_t x_203; uint64_t x_204; uint64_t x_2 x_201 = lean_ctor_get(x_140, 0); lean_inc(x_201); lean_dec(x_140); -x_202 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6462_(x_201); +x_202 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6277_(x_201); lean_dec(x_201); x_203 = lean_uint64_mix_hash(x_202, x_161); x_204 = lean_uint64_mix_hash(x_194, x_203); @@ -4566,6 +4566,7 @@ return x_13; } } lean_object* initialize_Lean_Server_Completion_CompletionCollectors(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Data_HashMap(uint8_t builtin, lean_object*); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Lean_Server_Completion(uint8_t builtin, lean_object* w) { lean_object * res; @@ -4574,6 +4575,9 @@ _G_initialized = true; res = initialize_Lean_Server_Completion_CompletionCollectors(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +res = initialize_Std_Data_HashMap(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__7___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__8___closed__1 = _init_l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__7___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__8___closed__1(); l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__7___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__8___closed__2 = _init_l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__7___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__8___closed__2(); l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__7___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__8___closed__3 = _init_l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__7___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__8___closed__3(); diff --git a/stage0/stdlib/Lean/Server/FileWorker.c b/stage0/stdlib/Lean/Server/FileWorker.c index ee99f83736a4..46a1a15c7791 100644 --- a/stage0/stdlib/Lean/Server/FileWorker.c +++ b/stage0/stdlib/Lean/Server/FileWorker.c @@ -66,6 +66,7 @@ static lean_object* l_Lean_Server_FileWorker_setupImports___lambda__6___closed__ LEAN_EXPORT lean_object* l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_reportSnapshots_go___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_IO_FS_Stream_readRequestAs___at_Lean_Server_FileWorker_initAndRunWorker___spec__2___closed__25; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_initAndRunWorker___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(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*); lean_object* lean_io_check_canceled(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_setupImports___lambda__7___boxed(lean_object*, lean_object*, lean_object*); @@ -197,6 +198,7 @@ static lean_object* l_IO_FS_Stream_readLspNotificationAs___at_Lean_Server_FileWo LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleRpcConnect___rarg(lean_object*, lean_object*); lean_object* l_IO_CancelToken_new(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_initializeWorker(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Json_opt___at___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_552____spec__1(lean_object*, lean_object*); lean_object* lean_io_mono_ms_now(lean_object*); lean_object* lean_task_pure(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleCancelRequest___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -219,7 +221,6 @@ LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_initAndRunWorker_writeErrorDia static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker___hyg_303____closed__2; static lean_object* l_Lean_Server_FileWorker_handleNotification___closed__5; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleStaleDependency(lean_object*); -lean_object* l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(lean_object*, 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*); @@ -258,6 +259,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Server_FileWorker_0__Lean_Server_FileW LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_parseParams___at_Lean_Server_FileWorker_handleRequest___spec__2(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonPublishDiagnosticsParams____x40_Lean_Data_Lsp_Diagnostics___hyg_2470_(lean_object*); static lean_object* l_Lean_Server_FileWorker_handleRequest___closed__9; +lean_object* l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124_(lean_object*); static lean_object* l_IO_FS_Stream_readRequestAs___at_Lean_Server_FileWorker_initAndRunWorker___spec__2___closed__13; LEAN_EXPORT lean_object* l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_reportSnapshots_go(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Server_FileWorker_handleRpcRelease___spec__1(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*); @@ -500,11 +502,9 @@ LEAN_EXPORT lean_object* l___private_Lean_Server_FileWorker_0__Lean_Server_FileW LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker___hyg_303_(lean_object*); lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLeanFileProgressParams____x40_Lean_Data_Lsp_Extra___hyg_889_(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_reportSnapshots___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(lean_object*); lean_object* l_Std_Mutex_atomically___at_Std_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__56; lean_object* lean_string_append(lean_object*, lean_object*); -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_toStructured_x3f___at_Lean_Server_FileWorker_runRefreshTask___spec__2(lean_object*); lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1472_(lean_object*); static lean_object* l_Lean_Json_toStructured_x3f___at_Lean_Server_FileWorker_runRefreshTask___spec__2___closed__2; @@ -11565,7 +11565,7 @@ LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_parseParams___at_Lean_Server_F { lean_object* x_5; lean_inc(x_1); -x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124_(x_1); +x_5 = l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124_(x_1); if (lean_obj_tag(x_5) == 0) { lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; @@ -15590,7 +15590,7 @@ LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodableDiagnosticWith_enc____x40 lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_165; x_3 = lean_ctor_get(x_1, 0); lean_inc(x_3); -x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_3); +x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_3); x_165 = lean_ctor_get(x_1, 1); lean_inc(x_165); if (lean_obj_tag(x_165) == 0) @@ -15609,7 +15609,7 @@ if (x_167 == 0) { lean_object* x_168; lean_object* x_169; x_168 = lean_ctor_get(x_165, 0); -x_169 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_168); +x_169 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_168); lean_ctor_set(x_165, 0, x_169); x_5 = x_165; x_6 = x_2; @@ -15621,7 +15621,7 @@ lean_object* x_170; lean_object* x_171; lean_object* x_172; x_170 = lean_ctor_get(x_165, 0); lean_inc(x_170); lean_dec(x_165); -x_171 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_170); +x_171 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_170); x_172 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_172, 0, x_171); x_5 = x_172; @@ -25826,7 +25826,7 @@ x_909 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_909, 0, x_907); lean_ctor_set(x_909, 1, x_908); x_910 = l_IO_FS_Stream_readRequestAs___at_Lean_Server_FileWorker_initAndRunWorker___spec__2___closed__24; -x_911 = l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(x_910, x_904); +x_911 = l_Lean_Json_opt___at___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_552____spec__1(x_910, x_904); lean_dec(x_904); switch (lean_obj_tag(x_901)) { case 0: @@ -27222,7 +27222,7 @@ x_385 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_385, 0, x_383); lean_ctor_set(x_385, 1, x_384); x_386 = l_IO_FS_Stream_readRequestAs___at_Lean_Server_FileWorker_initAndRunWorker___spec__2___closed__24; -x_387 = l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(x_386, x_380); +x_387 = l_Lean_Json_opt___at___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_552____spec__1(x_386, x_380); lean_dec(x_380); switch (lean_obj_tag(x_377)) { case 0: diff --git a/stage0/stdlib/Lean/Server/FileWorker/RequestHandling.c b/stage0/stdlib/Lean/Server/FileWorker/RequestHandling.c index 390bd6721b74..fe4458f93206 100644 --- a/stage0/stdlib/Lean/Server/FileWorker/RequestHandling.c +++ b/stage0/stdlib/Lean/Server/FileWorker/RequestHandling.c @@ -83,6 +83,8 @@ static lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_handlePla static lean_object* l_List_mapM_loop___at_Lean_Server_FileWorker_getInteractiveGoals___spec__2___closed__2; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleSemanticTokens_run(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_noHighlightKinds___closed__1; +lean_object* l_List_flatMapTR_go___at___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____spec__1(lean_object*, lean_object*); +uint64_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashPosition____x40_Lean_Data_Lsp_Basic___hyg_180_(lean_object*); static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12951____spec__4___lambda__3___closed__1; LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12951____spec__21___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getHeadInfo(lean_object*); @@ -203,7 +205,6 @@ static lean_object* l_Lean_Server_FileWorker_keywordSemanticTokenMap___closed__1 LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12951____spec__17___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleFoldingRange_addRanges_popRanges___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12951____spec__27___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*); -uint64_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashPosition____x40_Lean_Data_Lsp_Basic___hyg_365_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_collectInfoBasedSemanticTokens___lambda__2(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handlePlainTermGoal___closed__3; static lean_object* l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8962____closed__10; @@ -254,6 +255,7 @@ lean_object* lean_string_utf8_byte_size(lean_object*); static lean_object* l_Lean_Server_FileWorker_noHighlightKinds___closed__4; LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_erase___at_Lean_Server_FileWorker_filterDuplicateSemanticTokens___spec__8___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12951____spec__19(lean_object*, lean_object*, lean_object*); +uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqPosition____x40_Lean_Data_Lsp_Basic___hyg_41_(lean_object*, lean_object*); lean_object* l_Lean_Server_RequestM_findInfoTreeAtPosWithTrailingWhitespace(lean_object*, lean_object*); static lean_object* l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8962____closed__17; LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Server_FileWorker_filterDuplicateSemanticTokens___spec__7(lean_object*, lean_object*, lean_object*); @@ -280,8 +282,10 @@ LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__ lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2889_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleSemanticTokensFull___rarg(lean_object*, lean_object*); lean_object* l_Lean_Server_findModuleRefs(lean_object*, lean_object*, uint8_t, uint8_t); +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155_(lean_object*); static lean_object* l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8962____closed__4; size_t lean_usize_of_nat(lean_object*); +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____spec__1(lean_object*, lean_object*); lean_object* l_Lean_Widget_goalToInteractive(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_LocalContext_pop(lean_object*); static lean_object* l_Lean_Server_FileWorker_instHashableAbsoluteLspSemanticToken___closed__1; @@ -354,7 +358,6 @@ static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorke LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___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_object*, lean_object*, lean_object*); lean_object* l_Lean_Json_getObjValD(lean_object*, lean_object*); lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Server_References_allRefsFor___spec__1(lean_object*, lean_object*); -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326_(lean_object*); static lean_object* l_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___closed__6; lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_toJsonPlainTermGoal____x40_Lean_Data_Lsp_Extra___hyg_1692_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_NamespaceEntry_finish___lambda__1___boxed(lean_object*); @@ -362,7 +365,6 @@ static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_Fil static lean_object* l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8962____closed__20; static lean_object* l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8962____closed__3; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqPosition____x40_Lean_Data_Lsp_Basic___hyg_226_(lean_object*, lean_object*); lean_object* l_Lean_Syntax_getKind(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofFormat(lean_object*); @@ -424,7 +426,6 @@ LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Serve LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_toJsonPlainGoal____x40_Lean_Data_Lsp_Extra___hyg_1344_(lean_object*); lean_object* l_Lean_Syntax_getTrailingSize(lean_object*); -lean_object* l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12951____closed__16; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -485,7 +486,6 @@ lean_object* l_Lean_MapDeclarationExtension_find_x3f___rarg(lean_object*, lean_o LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12951____spec__11(lean_object*); lean_object* l_Lean_Server_locationLinksFromDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8962____closed__19; -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340_(lean_object*); static lean_object* l_Lean_Server_FileWorker_handlePlainTermGoal___closed__2; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__12(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Ord_instDecidableRelLe___rarg(lean_object*, lean_object*, lean_object*); @@ -606,7 +606,6 @@ static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_Fil LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12951____spec__27___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_noHighlightKinds; static lean_object* l_Lean_Server_FileWorker_NamespaceEntry_finish___closed__1; -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406_(lean_object*); static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12951____spec__24___lambda__2___closed__2; lean_object* lean_string_length(lean_object*); lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10625_(lean_object*); @@ -706,7 +705,6 @@ lean_object* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_obj LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12951____spec__37___lambda__2(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12951____spec__10___lambda__1(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_handlePlainGoal___spec__1___boxed(lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_computeAbsoluteLspSemanticTokens___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handlePlainGoal(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_handleWaitForDiagnostics___spec__1(lean_object*, lean_object*, lean_object*); @@ -804,6 +802,7 @@ static lean_object* l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_S LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Server_FileWorker_computeDeltaLspSemanticTokens___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_ReaderT_pure___at_Lean_Server_FileWorker_handleHover___spec__1___rarg(lean_object*, lean_object*, lean_object*); +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141_(lean_object*); lean_object* lean_infer_type(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_Server_FileWorker_getInteractiveGoals___spec__2(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_NamespaceEntry_finish___closed__3; @@ -836,6 +835,7 @@ LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Serve uint8_t l_Lean_Exception_isRuntime(lean_object*); lean_object* l_Lean_Server_RequestError_ofIoError(lean_object*); static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12951____spec__17___lambda__3___closed__1; +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221_(lean_object*); static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12951____closed__28; static lean_object* l_Lean_Server_FileWorker_handleFoldingRange_addRanges___closed__1; static lean_object* l_Lean_Server_FileWorker_filterDuplicateSemanticTokens___closed__1; @@ -23877,7 +23877,7 @@ x_5 = lean_ctor_get_uint8(x_1, sizeof(void*)*2); x_6 = lean_ctor_get(x_2, 0); x_7 = lean_ctor_get(x_2, 1); x_8 = lean_ctor_get_uint8(x_2, sizeof(void*)*2); -x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqPosition____x40_Lean_Data_Lsp_Basic___hyg_226_(x_3, x_6); +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqPosition____x40_Lean_Data_Lsp_Basic___hyg_41_(x_3, x_6); if (x_9 == 0) { uint8_t x_10; @@ -23887,7 +23887,7 @@ return x_10; else { uint8_t x_11; -x_11 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqPosition____x40_Lean_Data_Lsp_Basic___hyg_226_(x_4, x_7); +x_11 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqPosition____x40_Lean_Data_Lsp_Basic___hyg_41_(x_4, x_7); if (x_11 == 0) { uint8_t x_12; @@ -23938,9 +23938,9 @@ x_2 = lean_ctor_get(x_1, 0); x_3 = lean_ctor_get(x_1, 1); x_4 = lean_ctor_get_uint8(x_1, sizeof(void*)*2); x_5 = 0; -x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashPosition____x40_Lean_Data_Lsp_Basic___hyg_365_(x_2); +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashPosition____x40_Lean_Data_Lsp_Basic___hyg_180_(x_2); x_7 = lean_uint64_mix_hash(x_5, x_6); -x_8 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashPosition____x40_Lean_Data_Lsp_Basic___hyg_365_(x_3); +x_8 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashPosition____x40_Lean_Data_Lsp_Basic___hyg_180_(x_3); x_9 = lean_uint64_mix_hash(x_7, x_8); x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8939_(x_4); x_11 = lean_uint64_mix_hash(x_9, x_10); @@ -24200,7 +24200,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Server_FileWorker_RequestHandling_0__L lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8962____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -24238,7 +24238,7 @@ lean_inc(x_12); lean_dec(x_3); x_13 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8962____closed__14; lean_inc(x_1); -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____spec__1(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -24376,7 +24376,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Server_FileWorker_RequestHandling_0__L lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; 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; lean_object* x_22; lean_object* x_23; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406_(x_2); +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221_(x_2); x_4 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8962____closed__1; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); @@ -24387,7 +24387,7 @@ lean_ctor_set(x_7, 0, x_5); lean_ctor_set(x_7, 1, x_6); x_8 = lean_ctor_get(x_1, 1); lean_inc(x_8); -x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406_(x_8); +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221_(x_8); x_10 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8962____closed__14; x_11 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_11, 0, x_10); @@ -24415,7 +24415,7 @@ x_20 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_20, 0, x_7); lean_ctor_set(x_20, 1, x_19); x_21 = l_Lean_Server_FileWorker_locationLinksOfInfo_extractInstances___closed__1; -x_22 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_20, x_21); +x_22 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____spec__1(x_20, x_21); x_23 = l_Lean_Json_mkObj(x_22); return x_23; } @@ -25878,7 +25878,7 @@ x_8 = l_Ord_instDecidableRelLt___rarg(x_7, x_3, x_5); if (x_8 == 0) { uint8_t x_9; -x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqPosition____x40_Lean_Data_Lsp_Basic___hyg_226_(x_3, x_5); +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqPosition____x40_Lean_Data_Lsp_Basic___hyg_41_(x_3, x_5); lean_dec(x_5); lean_dec(x_3); if (x_9 == 0) @@ -31805,7 +31805,7 @@ LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileW { lean_object* x_2; lean_inc(x_1); -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326_(x_1); +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141_(x_1); if (lean_obj_tag(x_2) == 0) { uint8_t x_3; @@ -31918,7 +31918,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_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340_(x_5); +x_8 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155_(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); diff --git a/stage0/stdlib/Lean/Server/FileWorker/WidgetRequests.c b/stage0/stdlib/Lean/Server/FileWorker/WidgetRequests.c index 3e6b14d0ba5a..574afa69119f 100644 --- a/stage0/stdlib/Lean/Server/FileWorker/WidgetRequests.c +++ b/stage0/stdlib/Lean/Server/FileWorker/WidgetRequests.c @@ -47,6 +47,7 @@ static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Wi LEAN_EXPORT lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_58____lambda__1___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1389_(lean_object*); lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_toJsonLineRange____x40_Lean_Data_Lsp_Extra___hyg_2944_(lean_object*); +lean_object* l_List_flatMapTR_go___at___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Widget_instInhabitedMsgToInteractive___closed__1; LEAN_EXPORT lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1945____lambda__3(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_394____spec__1___lambda__2___closed__2; @@ -165,7 +166,6 @@ static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Wi static lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1429____closed__2; lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); static lean_object* l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1945____spec__3___closed__2; -lean_object* l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1389____spec__2(lean_object*, lean_object*); extern lean_object* l_Lean_Lsp_instFromJsonLocationLink; lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcCallParams____x40_Lean_Data_Lsp_Extra___hyg_2014____spec__2(lean_object*, lean_object*); @@ -769,7 +769,7 @@ x_12 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_12, 0, x_7); lean_ctor_set(x_12, 1, x_11); x_13 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_204____closed__1; -x_14 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_12, x_13); +x_14 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____spec__1(x_12, x_13); x_15 = l_Lean_Json_mkObj(x_14); return x_15; } @@ -803,7 +803,7 @@ x_26 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_26, 0, x_21); lean_ctor_set(x_26, 1, x_25); x_27 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_204____closed__1; -x_28 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_26, x_27); +x_28 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____spec__1(x_26, x_27); x_29 = l_Lean_Json_mkObj(x_28); return x_29; } @@ -1710,7 +1710,7 @@ x_17 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_17, 0, x_8); lean_ctor_set(x_17, 1, x_16); x_18 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_204____closed__1; -x_19 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_17, x_18); +x_19 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____spec__1(x_17, x_18); x_20 = l_Lean_Json_mkObj(x_19); return x_20; } @@ -4400,7 +4400,7 @@ x_5 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_5, 0, x_3); lean_ctor_set(x_5, 1, x_4); x_6 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_204____closed__1; -x_7 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_5, x_6); +x_7 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____spec__1(x_5, x_6); x_8 = l_Lean_Json_mkObj(x_7); return x_8; } @@ -4635,7 +4635,7 @@ x_12 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_12, 0, x_7); lean_ctor_set(x_12, 1, x_11); x_13 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_204____closed__1; -x_14 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_12, x_13); +x_14 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____spec__1(x_12, x_13); x_15 = l_Lean_Json_mkObj(x_14); return x_15; } @@ -4669,7 +4669,7 @@ x_26 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_26, 0, x_21); lean_ctor_set(x_26, 1, x_25); x_27 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_204____closed__1; -x_28 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_26, x_27); +x_28 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____spec__1(x_26, x_27); x_29 = l_Lean_Json_mkObj(x_28); return x_29; } diff --git a/stage0/stdlib/Lean/Server/References.c b/stage0/stdlib/Lean/Server/References.c index d6a4bba9ecd5..c959f160b270 100644 --- a/stage0/stdlib/Lean/Server/References.c +++ b/stage0/stdlib/Lean/Server/References.c @@ -35,7 +35,6 @@ LEAN_EXPORT lean_object* l_Lean_isRec___at_Lean_Server_RefInfo_toLspRefInfo___sp LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_combineIdents_buildIdMap___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_IO_throwServerError___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_References_findAt(lean_object*, lean_object*, lean_object*, uint8_t); -uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_627_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_References_allRefs___spec__10(lean_object*, size_t, size_t, lean_object*); static lean_object* l_List_mapTR_loop___at___private_Lean_Server_References_0__Lean_Server_toJsonIlean____x40_Lean_Server_References___hyg_1478____spec__2___closed__3; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Server_combineIdents_useConstRepresentatives___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -216,7 +215,6 @@ LEAN_EXPORT lean_object* l_List_mapTR_loop___at___private_Lean_Server_References static lean_object* l_Lean_Server_instFromJsonIlean___closed__1; LEAN_EXPORT lean_object* l_ReaderT_pure___at_Lean_Server_RefInfo_toLspRefInfo___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_References_definitionOf_x3f___lambda__1___boxed(lean_object*, lean_object*, lean_object*); -uint64_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_701_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_findModuleRefs___lambda__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_References_allRefs___spec__8(lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Lsp_ModuleRefs_findAt___spec__1(lean_object*, lean_object*); @@ -333,6 +331,7 @@ LEAN_EXPORT lean_object* l_Lean_Server_combineIdents_buildIdMap(lean_object*, le static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Server_References_allRefsFor___spec__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Server_References_findAt___lambda__1___boxed(lean_object*); lean_object* l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____rarg(lean_object*, lean_object*, lean_object*); +uint64_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_516_(lean_object*); lean_object* l_StateT_instMonad___rarg(lean_object*); lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_instRpcEncodableArray___spec__2(size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_combineIdents_findCanonicalRepresentative___boxed(lean_object*, lean_object*); @@ -376,6 +375,7 @@ LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Lsp_ModuleRefs_findRang lean_object* l_Lean_Environment_allImportedModuleNames(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_References_definitionOf_x3f___lambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Server_combineIdents_useConstRepresentatives___spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_442_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_foldl___at_Lean_Server_References_removeIlean___spec__5___boxed(lean_object*, lean_object*); lean_object* l_Array_foldrMUnsafe_fold___at_Lean_Lsp_instToJsonModuleRefs___spec__4(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Server_References_referringTo___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -17189,7 +17189,7 @@ lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; x_4 = lean_ctor_get(x_2, 0); x_5 = lean_ctor_get(x_2, 1); x_6 = lean_ctor_get(x_2, 2); -x_7 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_627_(x_4, x_1); +x_7 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_442_(x_4, x_1); if (x_7 == 0) { x_2 = x_6; @@ -17231,7 +17231,7 @@ lean_inc(x_13); lean_dec(x_11); x_14 = lean_ctor_get(x_2, 1); x_15 = lean_array_get_size(x_14); -x_16 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_701_(x_13); +x_16 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_516_(x_13); x_17 = 32; x_18 = lean_uint64_shift_right(x_16, x_17); x_19 = lean_uint64_xor(x_16, x_18); @@ -17880,7 +17880,7 @@ else lean_object* x_4; lean_object* x_5; uint8_t x_6; x_4 = lean_ctor_get(x_2, 0); x_5 = lean_ctor_get(x_2, 2); -x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_627_(x_4, x_1); +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_442_(x_4, x_1); if (x_6 == 0) { x_2 = x_5; @@ -17912,7 +17912,7 @@ lean_object* x_4; lean_object* x_5; lean_object* x_6; uint64_t x_7; uint64_t x_8 x_4 = lean_ctor_get(x_2, 0); x_5 = lean_ctor_get(x_2, 2); x_6 = lean_array_get_size(x_1); -x_7 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_701_(x_4); +x_7 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_516_(x_4); x_8 = 32; x_9 = lean_uint64_shift_right(x_7, x_8); x_10 = lean_uint64_xor(x_7, x_9); @@ -17943,7 +17943,7 @@ lean_inc(x_23); lean_inc(x_22); lean_dec(x_2); x_25 = lean_array_get_size(x_1); -x_26 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_701_(x_22); +x_26 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_516_(x_22); x_27 = 32; x_28 = lean_uint64_shift_right(x_26, x_27); x_29 = lean_uint64_xor(x_26, x_28); @@ -18035,7 +18035,7 @@ lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; x_6 = lean_ctor_get(x_3, 0); x_7 = lean_ctor_get(x_3, 1); x_8 = lean_ctor_get(x_3, 2); -x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_627_(x_6, x_1); +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_442_(x_6, x_1); if (x_9 == 0) { lean_object* x_10; @@ -18062,7 +18062,7 @@ lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); lean_dec(x_3); -x_14 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_627_(x_11, x_1); +x_14 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_442_(x_11, x_1); if (x_14 == 0) { lean_object* x_15; lean_object* x_16; @@ -18127,7 +18127,7 @@ lean_object* x_17; lean_object* x_18; lean_object* x_19; uint64_t x_20; uint64_t x_17 = lean_ctor_get(x_7, 0); x_18 = lean_ctor_get(x_7, 1); x_19 = lean_array_get_size(x_18); -x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_701_(x_15); +x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_516_(x_15); x_21 = 32; x_22 = lean_uint64_shift_right(x_20, x_21); x_23 = lean_uint64_xor(x_20, x_22); @@ -18204,7 +18204,7 @@ lean_inc(x_55); lean_inc(x_54); lean_dec(x_7); x_56 = lean_array_get_size(x_55); -x_57 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_701_(x_15); +x_57 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_516_(x_15); x_58 = 32; x_59 = lean_uint64_shift_right(x_57, x_58); x_60 = lean_uint64_xor(x_57, x_59); @@ -18514,7 +18514,7 @@ goto _start; else { uint8_t x_18; -x_18 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_627_(x_13, x_15); +x_18 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_442_(x_13, x_15); if (x_18 == 0) { x_2 = x_5; @@ -18676,7 +18676,7 @@ lean_inc(x_28); x_29 = lean_ctor_get(x_26, 1); lean_inc(x_29); lean_dec(x_26); -x_30 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_701_(x_29); +x_30 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_516_(x_29); lean_dec(x_29); if (lean_obj_tag(x_28) == 0) { @@ -18854,7 +18854,7 @@ return x_3; else { uint8_t x_21; -x_21 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_627_(x_16, x_18); +x_21 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_442_(x_16, x_18); lean_dec(x_18); lean_dec(x_16); if (x_21 == 0) @@ -18939,7 +18939,7 @@ return x_39; else { uint8_t x_40; -x_40 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_627_(x_34, x_36); +x_40 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_442_(x_34, x_36); lean_dec(x_36); lean_dec(x_34); if (x_40 == 0) @@ -19010,7 +19010,7 @@ goto _start; else { uint8_t x_19; -x_19 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_627_(x_14, x_16); +x_19 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_442_(x_14, x_16); if (x_19 == 0) { x_2 = x_6; @@ -19080,7 +19080,7 @@ x_190 = lean_ctor_get(x_10, 1); lean_inc(x_190); x_191 = lean_array_get_size(x_190); x_192 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_hashRefIdent____x40_Lean_Data_Lsp_Internal___hyg_158_(x_13); -x_193 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_701_(x_14); +x_193 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_516_(x_14); x_194 = 32; x_195 = 16; x_196 = lean_usize_of_nat(x_191); @@ -19213,7 +19213,7 @@ if (lean_is_exclusive(x_10)) { x_23 = lean_array_get_size(x_21); x_24 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_hashRefIdent____x40_Lean_Data_Lsp_Internal___hyg_158_(x_13); lean_dec(x_13); -x_25 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_701_(x_14); +x_25 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_516_(x_14); lean_dec(x_14); x_26 = 32; x_27 = 16; @@ -19366,7 +19366,7 @@ if (lean_is_exclusive(x_10)) { x_80 = lean_array_get_size(x_78); x_81 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_hashRefIdent____x40_Lean_Data_Lsp_Internal___hyg_158_(x_13); lean_dec(x_13); -x_82 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_701_(x_14); +x_82 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_516_(x_14); lean_dec(x_14); x_83 = 32; x_84 = 16; @@ -19535,7 +19535,7 @@ if (lean_is_exclusive(x_10)) { x_141 = lean_array_get_size(x_139); x_142 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_hashRefIdent____x40_Lean_Data_Lsp_Internal___hyg_158_(x_13); lean_dec(x_13); -x_143 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_701_(x_14); +x_143 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_516_(x_14); lean_dec(x_14); x_144 = 32; x_145 = 16; diff --git a/stage0/stdlib/Lean/Server/Watchdog.c b/stage0/stdlib/Lean/Server/Watchdog.c index df9fb0e5efa2..0b97fb0addb0 100644 --- a/stage0/stdlib/Lean/Server/Watchdog.c +++ b/stage0/stdlib/Lean/Server/Watchdog.c @@ -60,6 +60,7 @@ LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_parseParams___at_Lean_Server_Wat static lean_object* l_Lean_Json_toStructured_x3f___at_Lean_Server_Watchdog_initAndRunWatchdogAux___spec__4___closed__3; LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_FileWorker_queuedMsgs___boxed(lean_object*); LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Server_Watchdog_handleWorkspaceSymbol___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86_(lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at_Lean_Server_Watchdog_ImportData_update___spec__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_qsort_sort___at_Lean_Server_Watchdog_handlePrepareCallHierarchy___spec__3___closed__1; static lean_object* l_Lean_Server_Watchdog_FileWorker_queuedMsgs___closed__1; @@ -71,6 +72,7 @@ LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Server_Watchdog_handlePrep static lean_object* l_IO_FS_Stream_readNotificationAs___at_Lean_Server_Watchdog_initAndRunWatchdogAux___spec__2___closed__40; static lean_object* l_Lean_Server_Watchdog_mkLeanServerCapabilities___closed__16; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_Watchdog_handleCallHierarchyOutgoingCalls___spec__2(lean_object*, lean_object*, size_t, size_t, lean_object*); +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(lean_object*); static size_t l_Lean_Server_Watchdog_handleCallHierarchyIncomingCalls_collapseSameIncomingCalls___closed__2; LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_parseParams___at_Lean_Server_Watchdog_handleNotification___spec__14___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_handleCallHierarchyOutgoingCalls_collapseSameOutgoingCalls___boxed(lean_object*); @@ -88,7 +90,6 @@ LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_handleReference(lean_object*, le static lean_object* l_IO_FS_Stream_readNotificationAs___at_Lean_Server_Watchdog_initAndRunWatchdogAux___spec__2___closed__13; LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at_Lean_Server_Watchdog_handleDidSave___spec__2___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_handleRename(lean_object*, lean_object*, lean_object*); -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86_(lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Server_Watchdog_handleDidChangeWatchedFiles___spec__3___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Server_Watchdog_handleCallHierarchyIncomingCalls___spec__5___lambda__1___boxed(lean_object*, lean_object*); lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_toJsonLeanDidOpenTextDocumentParams____x40_Lean_Data_Lsp_Extra___hyg_309_(lean_object*); @@ -167,7 +168,6 @@ LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_Server_Watchdog_handleCanc lean_object* lean_io_getenv(lean_object*, lean_object*); static lean_object* l_IO_FS_Stream_readNotificationAs___at_Lean_Server_Watchdog_initAndRunWatchdogAux___spec__2___closed__24; uint8_t lean_float_decLt(double, double); -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080_(lean_object*); static lean_object* l_IO_FS_Stream_readNotificationAs___at_Lean_Server_Watchdog_initAndRunWatchdogAux___spec__2___closed__15; static lean_object* l_IO_FS_Stream_readRequestAs___at_Lean_Server_Watchdog_initAndRunWatchdog___spec__2___closed__4; lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Server_References_allRefsFor___spec__5(lean_object*, lean_object*); @@ -266,6 +266,7 @@ uint8_t lean_string_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_del___at_Lean_Server_Watchdog_ImportData_update___spec__6(lean_object*, lean_object*); static lean_object* l_Lean_Server_Watchdog_mkLeanServerCapabilities___closed__14; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_Watchdog_handleCallHierarchyIncomingCalls_collapseSameIncomingCalls___spec__12(size_t, size_t, lean_object*); +lean_object* l_Lean_Json_opt___at___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_552____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_FileWorker_waitForProc(lean_object*, lean_object*); lean_object* lean_task_pure(lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_del___at_Lean_Server_Watchdog_FileWorker_erasePendingRequest___spec__2(lean_object*, lean_object*); @@ -294,7 +295,6 @@ static lean_object* l_Lean_Server_Watchdog_handleNotification___closed__2; static lean_object* l___private_Lean_Server_Watchdog_0__Lean_Server_Watchdog_fromJsonCallHierarchyItemData____x40_Lean_Server_Watchdog___hyg_4173____closed__15; lean_object* l_Lean_Json_toStructured_x3f___at_Lean_Lsp_Ipc_shutdown___spec__2(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_parseParams___at_Lean_Server_Watchdog_handleNotification___spec__9(lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_insert___at_Lean_Server_Watchdog_handleRename___spec__11(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Watchdog_0__Lean_Server_Watchdog_forwardMessages_loop___closed__2; static lean_object* l___private_Lean_Server_Watchdog_0__Lean_Server_Watchdog_fromJsonCallHierarchyItemData____x40_Lean_Server_Watchdog___hyg_4173____closed__2; @@ -328,7 +328,6 @@ static lean_object* l_Lean_Server_Watchdog_initAndRunWatchdog___closed__1; static lean_object* l___private_Lean_Server_Watchdog_0__Lean_Server_Watchdog_fromJsonCallHierarchyItemData____x40_Lean_Server_Watchdog___hyg_4173____closed__10; static lean_object* l_Lean_Server_Watchdog_initAndRunWatchdogAux___closed__10; LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_eraseFileWorker(lean_object*, lean_object*, lean_object*); -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326_(lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_del___at_Lean_Server_Watchdog_eraseFileWorker___spec__2___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_CrashOrigin_noConfusion(lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Server_Watchdog_handleCallHierarchyIncomingCalls_collapseSameIncomingCalls___spec__11(lean_object*, lean_object*); @@ -345,6 +344,7 @@ lean_object* lean_io_wait_any(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_erase___at_Lean_Server_Watchdog_handleCallHierarchyIncomingCalls_collapseSameIncomingCalls___spec__8(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_insert___at_Lean_Server_Watchdog_ImportData_update___spec__9(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand_go___at_Lean_Server_Watchdog_handleCallHierarchyOutgoingCalls_collapseSameOutgoingCalls___spec__5(lean_object*, lean_object*, lean_object*); +lean_object* l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124_(lean_object*); static lean_object* l_IO_FS_Stream_readNotificationAs___at_Lean_Server_Watchdog_initAndRunWatchdogAux___spec__2___closed__26; static lean_object* l_Lean_Server_Watchdog_initAndRunWatchdogAux___closed__2; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_Watchdog_handleDidChangeWatchedFiles___spec__8(lean_object*, size_t, size_t, lean_object*); @@ -383,6 +383,7 @@ lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidChangeTex static lean_object* l_Lean_Server_Watchdog_workerCfg___closed__1; lean_object* l_Lean_FileMap_ofString(lean_object*); lean_object* lean_st_mk_ref(lean_object*, lean_object*); +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223_(lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Server_Watchdog_handleRename___spec__13___boxed(lean_object*, lean_object*, lean_object*); static uint8_t l_Lean_Server_Watchdog_tryDischargeQueuedMessages___closed__3; lean_object* lean_array_to_list(lean_object*); @@ -499,7 +500,6 @@ LEAN_EXPORT lean_object* l_IO_FS_Stream_readLspNotificationAs___at_Lean_Server_W static lean_object* l_IO_FS_Stream_readNotificationAs___at_Lean_Server_Watchdog_initAndRunWatchdogAux___spec__2___closed__27; static lean_object* l_Lean_Server_Watchdog_initAndRunWatchdogAux___closed__14; lean_object* l_Lean_Server_References_findRange_x3f(lean_object*, lean_object*, lean_object*, uint8_t); -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408_(lean_object*); static lean_object* l_Lean_Server_Watchdog_mainLoop___closed__3; LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_parseParams___at_Lean_Server_Watchdog_handleRequest___spec__14___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_RBNode_fold___at_Lean_RBTree_diff___spec__1___rarg(lean_object*, lean_object*, lean_object*); @@ -508,7 +508,6 @@ LEAN_EXPORT lean_object* l_Lean_RBNode_del___at_Lean_Server_Watchdog_ImportData_ LEAN_EXPORT lean_object* l_Lean_Json_toStructured_x3f___at_Lean_Server_Watchdog_initAndRunWatchdogAux___spec__4(lean_object*); static lean_object* l___private_Lean_Server_Watchdog_0__Lean_Server_Watchdog_forwardMessages_loop___closed__12; static lean_object* l_Lean_Server_Watchdog_handleNotification___closed__3; -uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordPosition____x40_Lean_Data_Lsp_Basic___hyg_300_(lean_object*, lean_object*); static lean_object* l_Lean_Server_Watchdog_findWorkerPath___closed__2; static lean_object* l_IO_FS_Stream_readNotificationAs___at_Lean_Server_Watchdog_initAndRunWatchdogAux___spec__2___closed__14; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_Watchdog_handleRequest___spec__10(size_t, size_t, lean_object*); @@ -789,12 +788,12 @@ LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_startLoadingReferences(lean_obje static lean_object* l_Lean_Server_Watchdog_tryDischargeQueuedMessages___closed__1; LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_FileWorker_killProcAndWait___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at_Lean_Server_Watchdog_handleDidSave___spec__3___at_Lean_Server_Watchdog_handleDidSave___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordPosition____x40_Lean_Data_Lsp_Basic___hyg_115_(lean_object*, lean_object*); static lean_object* l_Lean_Json_toStructured_x3f___at_Lean_Server_Watchdog_initAndRunWatchdogAux___spec__4___closed__1; static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_Watchdog_handleDidChangeWatchedFiles___spec__7___closed__1; static lean_object* l_Lean_Server_Watchdog_handleRename___closed__2; static lean_object* l_Lean_Server_Watchdog_handleDidOpen___closed__1; static lean_object* l___private_Lean_Server_Watchdog_0__Lean_Server_Watchdog_fromJsonCallHierarchyItemData____x40_Lean_Server_Watchdog___hyg_4173____closed__7; -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(lean_object*); static lean_object* l_Lean_Server_Watchdog_mkLeanServerCapabilities___closed__19; LEAN_EXPORT lean_object* l_Lean_RBNode_ins___at_Lean_Server_Watchdog_forwardRequestToWorker___spec__2(lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_Watchdog_handleCallHierarchyIncomingCalls_collapseSameIncomingCalls___spec__12___closed__1; @@ -804,7 +803,6 @@ lean_object* l___private_Lean_Data_Lsp_Client_0__Lean_Lsp_toJsonRegistrationPara lean_object* lean_string_append(lean_object*, lean_object*); lean_object* l_Lean_Server_References_addIlean(lean_object*, lean_object*, lean_object*); static lean_object* l_IO_FS_Stream_readNotificationAs___at_Lean_Server_Watchdog_initAndRunWatchdogAux___spec__2___closed__55; -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124_(lean_object*); LEAN_EXPORT lean_object* l_IO_FS_Stream_writeLspResponse___at_Lean_Server_Watchdog_handleRequest___spec__9(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_Watchdog_handleRequest___spec__10___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Server_Watchdog_0__Lean_Server_Watchdog_callHierarchyItemOf_x3f(lean_object*, lean_object*, lean_object*, lean_object*); @@ -821,6 +819,7 @@ LEAN_EXPORT lean_object* l_Lean_RBNode_del___at_Lean_Server_Watchdog_ImportData_ extern lean_object* l_Lean_Lsp_instInhabitedCallHierarchyIncomingCall; lean_object* lean_int_neg(lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Server_Watchdog_findDefinitions___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141_(lean_object*); uint8_t lean_nat_dec_le(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Server_Watchdog_handleWorkspaceSymbol___spec__1___lambda__1___boxed(lean_object*, lean_object*); static lean_object* l_IO_FS_Stream_readNotificationAs___at_Lean_Server_Watchdog_initAndRunWatchdogAux___spec__2___closed__64; @@ -870,6 +869,7 @@ LEAN_EXPORT lean_object* l_Lean_RBNode_erase___at_Lean_Server_Watchdog_ImportDat LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_Watchdog_handlePrepareCallHierarchy___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_filterMapM___at_Lean_Server_Watchdog_handleDidChangeWatchedFiles___spec__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_Watchdog_handleCallHierarchyIncomingCalls_collapseSameIncomingCalls___spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895_(lean_object*); LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Server_Watchdog_handleCallHierarchyOutgoingCalls___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_SearchPath_findAllWithExt(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Server_Watchdog_handleRename___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -25937,7 +25937,7 @@ x_9 = lean_ctor_get(x_1, 0); x_10 = lean_ctor_get(x_1, 1); x_11 = lean_ctor_get(x_1, 2); x_12 = lean_ctor_get(x_1, 3); -x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordPosition____x40_Lean_Data_Lsp_Basic___hyg_300_(x_2, x_10); +x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordPosition____x40_Lean_Data_Lsp_Basic___hyg_115_(x_2, x_10); switch (x_13) { case 0: { @@ -25982,7 +25982,7 @@ lean_inc(x_21); lean_inc(x_20); lean_inc(x_19); lean_dec(x_1); -x_23 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordPosition____x40_Lean_Data_Lsp_Basic___hyg_300_(x_2, x_20); +x_23 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordPosition____x40_Lean_Data_Lsp_Basic___hyg_115_(x_2, x_20); switch (x_23) { case 0: { @@ -26038,7 +26038,7 @@ x_33 = lean_ctor_get(x_1, 0); x_34 = lean_ctor_get(x_1, 1); x_35 = lean_ctor_get(x_1, 2); x_36 = lean_ctor_get(x_1, 3); -x_37 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordPosition____x40_Lean_Data_Lsp_Basic___hyg_300_(x_2, x_34); +x_37 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordPosition____x40_Lean_Data_Lsp_Basic___hyg_115_(x_2, x_34); switch (x_37) { case 0: { @@ -27432,7 +27432,7 @@ lean_inc(x_344); lean_inc(x_343); lean_inc(x_342); lean_dec(x_1); -x_346 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordPosition____x40_Lean_Data_Lsp_Basic___hyg_300_(x_2, x_343); +x_346 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordPosition____x40_Lean_Data_Lsp_Basic___hyg_115_(x_2, x_343); switch (x_346) { case 0: { @@ -34377,7 +34377,7 @@ LEAN_EXPORT lean_object* l_Lean_Json_toStructured_x3f___at_Lean_Server_Watchdog_ _start: { lean_object* x_2; uint8_t x_3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86_(x_1); +x_2 = l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86_(x_1); x_3 = !lean_is_exclusive(x_2); if (x_3 == 0) { @@ -37875,7 +37875,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_2, 1); -x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408_(x_5); +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223_(x_5); lean_ctor_set_tag(x_2, 2); lean_ctor_set(x_2, 1, x_6); x_7 = l_IO_FS_Stream_writeLspMessage(x_1, x_2, x_3); @@ -37889,7 +37889,7 @@ x_9 = lean_ctor_get(x_2, 1); lean_inc(x_9); lean_inc(x_8); lean_dec(x_2); -x_10 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408_(x_9); +x_10 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223_(x_9); x_11 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_11, 0, x_8); lean_ctor_set(x_11, 1, x_10); @@ -37984,7 +37984,7 @@ lean_dec(x_14); x_15 = lean_ctor_get(x_4, 0); lean_inc(x_15); lean_dec(x_4); -x_16 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_15); +x_16 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_15); lean_ctor_set_tag(x_2, 2); lean_ctor_set(x_2, 1, x_16); x_17 = l_IO_FS_Stream_writeLspMessage(x_1, x_2, x_3); @@ -37999,7 +37999,7 @@ lean_dec(x_2); x_19 = lean_ctor_get(x_4, 0); lean_inc(x_19); lean_dec(x_4); -x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_19); +x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_19); x_21 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_21, 0, x_18); lean_ctor_set(x_21, 1, x_20); @@ -38479,7 +38479,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_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080_(x_5); +x_8 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895_(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); @@ -38534,7 +38534,7 @@ LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_parseParams___at_Lean_Server_Wat { lean_object* x_4; lean_inc(x_1); -x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326_(x_1); +x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141_(x_1); if (lean_obj_tag(x_4) == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; @@ -40756,7 +40756,7 @@ LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_parseParams___at_Lean_Server_Wat { lean_object* x_4; lean_inc(x_1); -x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124_(x_1); +x_4 = l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124_(x_1); if (lean_obj_tag(x_4) == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; @@ -45012,7 +45012,7 @@ x_371 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_371, 0, x_369); lean_ctor_set(x_371, 1, x_370); x_372 = l_IO_FS_Stream_readNotificationAs___at_Lean_Server_Watchdog_initAndRunWatchdogAux___spec__2___closed__14; -x_373 = l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(x_372, x_366); +x_373 = l_Lean_Json_opt___at___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_552____spec__1(x_372, x_366); lean_dec(x_366); switch (lean_obj_tag(x_363)) { case 0: @@ -52284,7 +52284,7 @@ x_909 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_909, 0, x_907); lean_ctor_set(x_909, 1, x_908); x_910 = l_IO_FS_Stream_readNotificationAs___at_Lean_Server_Watchdog_initAndRunWatchdogAux___spec__2___closed__14; -x_911 = l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(x_910, x_904); +x_911 = l_Lean_Json_opt___at___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_552____spec__1(x_910, x_904); lean_dec(x_904); switch (lean_obj_tag(x_901)) { case 0: diff --git a/stage0/stdlib/Lean/Util/ShareCommon.c b/stage0/stdlib/Lean/Util/ShareCommon.c index aaf582ebde5e..273cdd620545 100644 --- a/stage0/stdlib/Lean/Util/ShareCommon.c +++ b/stage0/stdlib/Lean/Util/ShareCommon.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Lean.Util.ShareCommon -// Imports: Init.ShareCommon Std.Data.HashSet Std.Data.HashMap Lean.Data.PersistentHashMap Lean.Data.PersistentHashSet +// Imports: Init.ShareCommon Std.Data.HashSet.Basic Std.Data.HashMap.Basic Lean.Data.PersistentHashMap Lean.Data.PersistentHashSet #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -1100,8 +1100,8 @@ return x_2; } } lean_object* initialize_Init_ShareCommon(uint8_t builtin, lean_object*); -lean_object* initialize_Std_Data_HashSet(uint8_t builtin, lean_object*); -lean_object* initialize_Std_Data_HashMap(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Data_HashSet_Basic(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Data_HashMap_Basic(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Data_PersistentHashMap(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Data_PersistentHashSet(uint8_t builtin, lean_object*); static bool _G_initialized = false; @@ -1112,10 +1112,10 @@ _G_initialized = true; res = initialize_Init_ShareCommon(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -res = initialize_Std_Data_HashSet(builtin, lean_io_mk_world()); +res = initialize_Std_Data_HashSet_Basic(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -res = initialize_Std_Data_HashMap(builtin, lean_io_mk_world()); +res = initialize_Std_Data_HashMap_Basic(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); res = initialize_Lean_Data_PersistentHashMap(builtin, lean_io_mk_world()); diff --git a/stage0/stdlib/Lean/Util/Trace.c b/stage0/stdlib/Lean/Util/Trace.c index fbda915c6df9..a2f32b862f49 100644 --- a/stage0/stdlib/Lean/Util/Trace.c +++ b/stage0/stdlib/Lean/Util/Trace.c @@ -22,6 +22,7 @@ LEAN_EXPORT lean_object* l_Lean_instMonadAlwaysExceptReaderT___rarg(lean_object* LEAN_EXPORT lean_object* l_Lean_addTrace___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_StateRefT_x27_instMonadExceptOf___rarg(lean_object*); static lean_object* l___private_Lean_Util_Trace_0__Lean_withStartStop___rarg___closed__2; +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__5; LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_addTraceAsMessages___spec__5___at_Lean_addTraceAsMessages___spec__6(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_1644_(lean_object*); lean_object* lean_format_pretty(lean_object*, lean_object*, lean_object*, lean_object*); @@ -29,6 +30,7 @@ static lean_object* l___auto____x40_Lean_Util_Trace___hyg_3014____closed__10; LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_75_(lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_withTraceNode_x27___spec__1___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_addTraceAsMessages___spec__11___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_expandTraceMacro___closed__9; static lean_object* l___auto____x40_Lean_Util_Trace___hyg_3014____closed__1; LEAN_EXPORT lean_object* l_Lean_addTraceAsMessages___rarg___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_instMonadAlwaysExceptEIO___closed__1; @@ -38,22 +40,23 @@ static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__22; static lean_object* l___auto____x40_Lean_Util_Trace___hyg_3014____closed__23; LEAN_EXPORT lean_object* l_Lean_addTrace___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_withTraceNode_x27___spec__1___rarg___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__20; static lean_object* l_Lean_resetTraceState___rarg___lambda__1___closed__2; lean_object* l_IO_getNumHeartbeats___boxed(lean_object*); lean_object* l_instBEqOfDecidableEq___rarg(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__28; LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___rarg___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_addTraceAsMessages___spec__14___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_mk_empty_array_with_capacity(lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__6; +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__34; LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___lambda__6(lean_object*, lean_object*); +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__32; static lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_75____closed__2; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_addTraceAsMessages___spec__11___rarg___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__35; static lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_1276____closed__4; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_printTraces___spec__6___rarg___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_withTraceNode_x27___rarg___closed__1; +LEAN_EXPORT lean_object* l_Lean_expandTraceMacro(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, double, lean_object*); LEAN_EXPORT lean_object* l_Lean_exceptOptionEmoji(lean_object*, lean_object*); static lean_object* l_Lean_addTraceAsMessages___rarg___lambda__2___closed__4; @@ -63,15 +66,14 @@ LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___lambda__9(lean_obje static lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_1276____closed__5; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_addTraceAsMessages___spec__11___rarg___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_StateT_instMonadExceptOf___rarg(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__18; lean_object* l_Lean_PersistentArray_toArray___rarg(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_withStartStop___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_1552____closed__1; static lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___rarg___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___lambda__10(lean_object*, lean_object*); +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__6; LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_withTraceNode_x27___spec__1___rarg___lambda__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Util_Trace___hyg_3014____closed__27; -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__33; static lean_object* l_Array_qsort_sort___at_Lean_addTraceAsMessages___spec__15___closed__1; LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_instExceptToEmojiOption(lean_object*, lean_object*); @@ -87,13 +89,13 @@ uint64_t lean_uint64_of_nat(lean_object*); uint64_t lean_uint64_mix_hash(uint64_t, uint64_t); LEAN_EXPORT lean_object* l_Lean_withTraceNode_x27___rarg___lambda__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__43; LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at_Lean_printTraces___spec__2___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__16; lean_object* l_Lean_MessageData_joinSep(lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__7; size_t lean_uint64_to_usize(uint64_t); LEAN_EXPORT lean_object* l_Lean_exceptBoolEmoji(lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__31; +static lean_object* l_Lean_expandTraceMacro___closed__15; LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_withTraceNode___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_resetTraceState___rarg___lambda__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_modifyTraces___rarg(lean_object*, lean_object*); @@ -102,12 +104,11 @@ LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___rarg LEAN_EXPORT lean_object* l_Lean_instMonadAlwaysExceptStateRefT_x27(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___lambda__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__44; lean_object* l_Lean_Syntax_getId(lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__8; static lean_object* l___auto____x40_Lean_Util_Trace___hyg_3014____closed__6; static lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_1276____closed__1; static lean_object* l___auto____x40_Lean_Util_Trace___hyg_3014____closed__2; +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__51; 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_Array_forIn_x27Unsafe_loop___at_Lean_printTraces___spec__5___rarg___lambda__2(lean_object*, lean_object*); @@ -127,12 +128,10 @@ static lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_1552____closed__3 static lean_object* l___auto____x40_Lean_Util_Trace___hyg_3014____closed__19; LEAN_EXPORT lean_object* l_Lean_getTraces(lean_object*); LEAN_EXPORT lean_object* l_Lean_addTraceAsMessages___rarg___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__22; LEAN_EXPORT lean_object* l_Lean_exceptBoolEmoji___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode_x27___rarg(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* l_Lean_replaceRef(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_printTraces___rarg(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__48; LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_1460_(lean_object*); LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_withTraceNode_x27___spec__3(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instMonadAlwaysExceptMonadCacheT(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -145,15 +144,12 @@ lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Util_Trace___hyg_3014____closed__9; static lean_object* l_Lean_addTraceAsMessages___rarg___lambda__2___closed__2; lean_object* l_Lean_Syntax_getTailPos_x3f(lean_object*, uint8_t); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__41; static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__9; LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___lambda__6___boxed(lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__5; LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_addTraceAsMessages___spec__15(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_printTraces___spec__5(lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___lambda__5___boxed(lean_object**); LEAN_EXPORT lean_object* l_Lean_setTraceState___rarg(lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__39; LEAN_EXPORT lean_object* l_Lean_traceM(lean_object*); static lean_object* l_Lean_crossEmoji___closed__1; static lean_object* l_Lean_addTraceAsMessages___rarg___lambda__2___closed__1; @@ -166,7 +162,6 @@ LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___lambda__4___boxed(l LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___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_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_instInhabitedTraceElem___closed__2; LEAN_EXPORT lean_object* l_Lean_withTraceNode_x27___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__16; LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_withTraceNode_x27___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*); static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__15; LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___rarg___lambda__1(lean_object*); @@ -178,8 +173,8 @@ LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at_Lean_printTraces__ size_t lean_usize_of_nat(lean_object*); lean_object* l_Nat_nextPowerOfTwo_go(lean_object*, lean_object*, lean_object*); static size_t l_Lean_instInhabitedTraceState___closed__4; -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__5; LEAN_EXPORT lean_object* l_Lean_trace_profiler_useHeartbeats; +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__46; LEAN_EXPORT lean_object* l_Lean_addRawTrace(lean_object*); lean_object* l_Lean_stringToMessageData(lean_object*); LEAN_EXPORT lean_object* l_Lean_addTrace___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); @@ -192,7 +187,6 @@ LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_addTraceAsMessag lean_object* l_Lean_Option_register___at_Lean_initFn____x40_Lean_Util_Profile___hyg_40____spec__1(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Exception_toMessageData(lean_object*); lean_object* l_instBEqProd___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__13; LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_withTraceNodeBefore___spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_resetTraceState___rarg(lean_object*); @@ -202,14 +196,13 @@ lean_object* lean_string_push(lean_object*, uint32_t); LEAN_EXPORT lean_object* l_Lean_addTraceAsMessages___rarg___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_instMonadExceptOfEIO(lean_object*); static lean_object* l_Lean_addTraceAsMessages___rarg___lambda__2___closed__5; +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__3; static lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_1368____closed__5; LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___lambda__3___boxed(lean_object**); static lean_object* l_Lean_instInhabitedTraceElem___closed__1; LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at_Lean_addTraceAsMessages___spec__8(lean_object*); static lean_object* l___auto____x40_Lean_Util_Trace___hyg_3014____closed__16; -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__36; LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_withTraceNode_x27___spec__3___rarg(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__2; LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___lambda__13___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_resetTraceState(lean_object*); static lean_object* l_Lean_instInhabitedTraceState___closed__2; @@ -218,6 +211,7 @@ static lean_object* l___auto____x40_Lean_Util_Trace___hyg_3014____closed__13; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Util_Trace_0__Lean_addTraceNode___spec__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_instExceptToEmojiOption___closed__1; lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_expandTraceMacro___closed__10; LEAN_EXPORT lean_object* l_Lean_addRawTrace___rarg___lambda__2(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___lambda__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces(lean_object*); @@ -225,7 +219,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode(lean_o static lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_1460____closed__5; LEAN_EXPORT lean_object* l_Lean_exceptBoolEmoji___rarg___boxed(lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_printTraces___spec__4___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*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__3; static lean_object* l___private_Lean_Util_Trace_0__Lean_checkTraceOption___closed__2; static lean_object* l_Lean_bombEmoji___closed__1; size_t lean_usize_of_nat(lean_object*); @@ -240,8 +233,8 @@ LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___lambda__13(lean_object*, LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_withTraceNode_x27___spec__1___rarg___lambda__5(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, double, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand_go___at_Lean_addTraceAsMessages___spec__4(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___lambda__11___boxed(lean_object**); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__29; static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__2; +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__36; LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Option_register___at_Lean_initFn____x40_Lean_Util_Trace___hyg_1552____spec__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_withTraceNode___rarg___lambda__6___closed__2; @@ -252,7 +245,6 @@ LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_withTraceNode_x27___spec LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_addTraceAsMessages___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*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___lambda__11(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_withTraceNodeBefore___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__3; static lean_object* l_Lean_resetTraceState___rarg___closed__1; static lean_object* l___auto____x40_Lean_Util_Trace___hyg_3014____closed__20; LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_withTraceNode___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*); @@ -266,6 +258,7 @@ LEAN_EXPORT lean_object* l_Lean_exceptEmoji___rarg___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_withTraceNode_x27___spec__1___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_NameSSet_insert___spec__7(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_withStartStop___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__29; LEAN_EXPORT lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1(lean_object*, lean_object*, lean_object*); uint64_t lean_uint64_shift_right(uint64_t, uint64_t); lean_object* l_Lean_SourceInfo_fromRef(lean_object*, uint8_t); @@ -279,12 +272,12 @@ lean_object* l_instDecidableEqPos___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_Lean_withTraceNodeBefore___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_forIn_x27Unsafe_loop___at_Lean_printTraces___spec__5___rarg___lambda__1(lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*); +static lean_object* l_Lean_expandTraceMacro___closed__2; LEAN_EXPORT lean_object* l_Lean_registerTraceClass(lean_object*, uint8_t, lean_object*, lean_object*); static lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_75____closed__3; static lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_1644____closed__3; LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_withTraceNode_x27___spec__1___rarg___lambda__3(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, double, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_withStartStop___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__10; static lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_1460____closed__3; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_printTraces___spec__6___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_withTraceNode___rarg___lambda__6___closed__1; @@ -293,16 +286,16 @@ LEAN_EXPORT lean_object* l_Lean_resetTraceState___rarg___lambda__1___boxed(lean_ lean_object* l_Lean_MessageData_ofFormat(lean_object*); lean_object* l_Lean_PersistentArray_append___rarg(lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Util_Trace___hyg_3014____closed__12; -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__14; static lean_object* l_Lean_withTraceNodeBefore___rarg___lambda__3___closed__2; LEAN_EXPORT lean_object* l_Lean_printTraces(lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__10; +static lean_object* l_Lean_expandTraceMacro___closed__12; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_printTraces___spec__6___rarg___lambda__2(lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_setTraceState___rarg___lambda__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___lambda__5(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, double, lean_object*); LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_1276_(lean_object*); LEAN_EXPORT lean_object* l_Lean_addTraceAsMessages___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_traceM___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__42; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_addTraceAsMessages___spec__10(lean_object*); lean_object* l_Lean_quoteNameMk(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at_Lean_printTraces___spec__2(lean_object*); @@ -318,43 +311,39 @@ LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_withTraceNode_x27___spec LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___lambda__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getTraces___rarg___lambda__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instMonadAlwaysExceptStateT(lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__12; +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__8; uint8_t l_List_isEmpty___rarg(lean_object*); lean_object* lean_st_mk_ref(lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_withTraceNode_x27___spec__1___rarg___lambda__4___boxed(lean_object**); LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_withTraceNodeBefore___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___lambda__8(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, double, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_printTraces___rarg___lambda__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___lambda__14___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_to_list(lean_object*); +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__26; LEAN_EXPORT uint8_t l___private_Lean_Util_Trace_0__Lean_checkTraceOption_go(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__12; +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__45; LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_withTraceNode_x27___spec__1___rarg___lambda__7(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__47; lean_object* l_Lean_Syntax_node3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__50; lean_object* l_Lean_Name_append(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___lambda__9(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_addTraceAsMessages___spec__11___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__51; -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__2; LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___lambda__16___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_withTraceNode_x27___spec__1___rarg___lambda__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Util_Trace___hyg_3014____closed__15; LEAN_EXPORT uint8_t l_Std_DHashMap_Internal_AssocList_contains___at_Lean_addTraceAsMessages___spec__2(lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__13; +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__31; LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor(lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___lambda__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Option_get_x3f___at_Lean_addTraceAsMessages___spec__17___boxed(lean_object*, lean_object*); static lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_1552____closed__2; -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__15; LEAN_EXPORT lean_object* l_Lean_instInhabitedTraceElem; -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__11; LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_withTraceNode_x27___spec__1___rarg___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___lambda__14(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_addTraceAsMessages___spec__14___rarg___closed__1; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_printTraces___spec__4___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__24; +static lean_object* l_Lean_expandTraceMacro___closed__4; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_printTraces___spec__6(lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_withTraceNode_x27___spec__1___rarg___lambda__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_IO_print___at_IO_println___spec__1(lean_object*, lean_object*); @@ -368,8 +357,6 @@ uint8_t lean_name_eq(lean_object*, lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_printTraces___spec__5___rarg___lambda__2___closed__1; LEAN_EXPORT lean_object* l_IO_println___at_Lean_printTraces___spec__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__46; lean_object* l_Lean_Syntax_node2(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at_Lean_addTraceAsMessages___spec__9(lean_object*); @@ -378,24 +365,26 @@ LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_printTraces___sp LEAN_EXPORT lean_object* l_Lean_addTraceAsMessages___rarg___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_trace_profiler_threshold; lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__27; static lean_object* l_Lean_instExceptToEmojiBool___closed__1; LEAN_EXPORT lean_object* l_Lean_withTraceNode_x27___rarg___lambda__2(lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__9; +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__4; +static lean_object* l_Lean_expandTraceMacro___closed__6; +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__38; static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__8; LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__44; static lean_object* l___auto____x40_Lean_Util_Trace___hyg_3014____closed__11; static lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_1552____closed__4; +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__48; +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__15; LEAN_EXPORT lean_object* l_Lean_traceM___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__25; extern lean_object* l_Std_Format_defWidth; LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_withTraceNode_x27___spec__1___rarg___lambda__9(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*, uint8_t, lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__1; LEAN_EXPORT lean_object* l_Lean_trace(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_1276____closed__3; -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__6; static lean_object* l_Lean_addTrace___rarg___lambda__1___closed__2; +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__18; lean_object* lean_register_option(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instMonadAlwaysExceptMonadCacheT___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_addTraceAsMessages___spec__10___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*, size_t, lean_object*); @@ -409,7 +398,6 @@ static lean_object* l_Lean_addTraceAsMessages___rarg___lambda__2___closed__3; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_addTraceAsMessages___spec__12___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_traceM___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__5; -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__49; static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__1; double l_Float_ofScientific(lean_object*, uint8_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___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*, uint8_t, lean_object*); @@ -418,12 +406,14 @@ LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_printTraces___sp static lean_object* l_Lean_instInhabitedTraceState___closed__6; static lean_object* l___private_Lean_Util_Trace_0__Lean_withStartStop___rarg___closed__1; static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__21; -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__4; LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_registerTraceClass___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__14; +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__47; LEAN_EXPORT lean_object* l_Lean_instMonadTraceOfMonadLift___rarg___lambda__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__18; LEAN_EXPORT lean_object* l_Lean_trace___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__2; static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__24; static lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_1368____closed__4; static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__4; @@ -438,7 +428,6 @@ LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at_Lean_addTraceAsMessag LEAN_EXPORT lean_object* l_Lean_printTraces___rarg___lambda__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_1276____closed__2; LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_withTraceNode_x27___spec__1___rarg___lambda__10(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__11; static lean_object* l___auto____x40_Lean_Util_Trace___hyg_3014____closed__17; LEAN_EXPORT lean_object* l_Lean_getTraces___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode_x27___rarg___lambda__2___boxed(lean_object*); @@ -449,12 +438,14 @@ LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_addTraceAsMessages___spec_ static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__10; lean_object* l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__9; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_printTraces___spec__5___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__32; +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__27; +static lean_object* l_Lean_expandTraceMacro___closed__16; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instMonadAlwaysExceptEIO(lean_object*); +static lean_object* l_Lean_expandTraceMacro___closed__14; LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___lambda__7___boxed(lean_object**); +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__10; LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_withStartStop___rarg___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_withTraceNode_x27___spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_addTraceAsMessages___spec__11___rarg___lambda__1(lean_object*, lean_object*, lean_object*); @@ -462,66 +453,73 @@ LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_addTraceAsMessag static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__6; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); static lean_object* l___private_Lean_Util_Trace_0__Lean_checkTraceOption___closed__1; +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__41; LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_addTraceAsMessages___spec__14___rarg___closed__3; LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_replace___at_Lean_addTraceAsMessages___spec__7(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__22; static lean_object* l_Lean_addTraceAsMessages___rarg___lambda__5___closed__1; +static lean_object* l_Lean_expandTraceMacro___closed__1; lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__25; LEAN_EXPORT lean_object* l_Lean_instMonadAlwaysExceptStateT___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___lambda__10(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_withTraceNodeBefore___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__19; static lean_object* l___auto____x40_Lean_Util_Trace___hyg_3014____closed__18; static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__14; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_addTraceAsMessages___spec__14___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___auto____x40_Lean_Util_Trace___hyg_3014_; LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_withStartStop___rarg___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_expandTraceMacro___closed__7; +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__24; LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___lambda__6___boxed(lean_object**); static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__20; LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_withTraceNode_x27___spec__1___rarg___lambda__4(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, double, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at_Lean_printTraces___spec__2___rarg___lambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at_Lean_printTraces___spec__3___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__7; LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___lambda__8___boxed(lean_object**); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_printTraces___spec__4___rarg___lambda__3(lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*); lean_object* l_Lean_Syntax_node1(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__15; static lean_object* l___auto____x40_Lean_Util_Trace___hyg_3014____closed__21; LEAN_EXPORT lean_object* l_Lean_exceptEmoji(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Option_register___at_Lean_initFn____x40_Lean_Util_Trace___hyg_1552____spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___rarg(lean_object*, lean_object*, lean_object*, lean_object*); uint64_t l_Lean_Name_hash___override(lean_object*); -LEAN_EXPORT lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___lambda__8___boxed(lean_object**); +static lean_object* l_Lean_expandTraceMacro___closed__5; static lean_object* l___auto____x40_Lean_Util_Trace___hyg_3014____closed__14; LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___lambda__2(lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__14; uint64_t lean_uint64_xor(uint64_t, uint64_t); LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_addTraceAsMessages___rarg___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__50; lean_object* lean_nat_sub(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_addTraceAsMessages___spec__11___rarg___lambda__3(lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__30; +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__9; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_printTraces___spec__6___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_withTraceNode_x27___spec__1(lean_object*, lean_object*); +static lean_object* l_Lean_expandTraceMacro___closed__11; lean_object* lean_nat_mul(lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__43; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_addTraceAsMessages___spec__16___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___rarg___lambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___lambda__12(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*); static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__3; -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__37; static lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_1460____closed__2; lean_object* l___private_Init_Data_Array_QSort_0__Array_qpartition___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_trace_profiler_threshold_unitAdjusted___boxed(lean_object*); +static lean_object* l_Lean_expandTraceMacro___closed__13; static lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_1644____closed__1; static lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_1368____closed__2; static double l___private_Lean_Util_Trace_0__Lean_withStartStop___rarg___lambda__1___closed__1; -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__23; LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at_Lean_addTraceAsMessages___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* l_BaseIO_toIO___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_printTraces___spec__4(lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_printTraces___spec__4___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_addTraceAsMessages___spec__13(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_addTraceAsMessages___spec__11___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Lean_expandTraceMacro___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_erase_macro_scopes(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_checkTraceOption___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_resetTraceState___rarg___lambda__1___closed__1; @@ -538,19 +536,24 @@ static lean_object* l___auto____x40_Lean_Util_Trace___hyg_3014____closed__8; LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_addTraceAsMessages___spec__5(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_addTraceAsMessages___spec__15___lambda__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_exceptOptionEmoji___rarg(lean_object*); +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__23; LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_checkTraceOption_go___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_isTracingEnabledForCore(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_withTraceNode_x27___spec__1___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_add(size_t, size_t); static lean_object* l___auto____x40_Lean_Util_Trace___hyg_3014____closed__3; LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_withTraceNode_x27___spec__1___rarg___lambda__11(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__17; +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__11; static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__7; static uint64_t l_Lean_instInhabitedTraceState___closed__1; static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__13; LEAN_EXPORT lean_object* l_Lean_trace_profiler_output_pp; LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___lambda__7___boxed(lean_object**); +LEAN_EXPORT lean_object* l_Lean_expandTraceMacro___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_KVMap_findCore(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_addTraceAsMessages___spec__11(lean_object*); +LEAN_EXPORT lean_object* l_Lean_expandTraceMacro___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_uget(lean_object*, size_t); size_t lean_array_size(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_addTraceAsMessages___spec__16(lean_object*, size_t, size_t, lean_object*); @@ -563,22 +566,21 @@ lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_addRawTrace___rarg___lambda__1(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_instInhabitedTraceState___closed__5; -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__7; LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at_Lean_printTraces___spec__3___rarg___lambda__1(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__8; LEAN_EXPORT lean_object* l_Lean_instMonadAlwaysExceptMonadCacheT___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_modifyTraces___rarg___lambda__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_addTraceAsMessages___spec__13___boxed(lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__40; LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_withStartStop(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_addTrace___rarg___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_1368____closed__3; -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__4; lean_object* lean_string_append(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at_Lean_printTraces___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_trace_profiler_output; +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__20; +static lean_object* l_Lean_expandTraceMacro___closed__8; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_printTraces___spec__5___rarg___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__28; LEAN_EXPORT lean_object* l_Lean_instExceptToEmojiBool(lean_object*); static double l_Lean_trace_profiler_threshold_unitAdjusted___closed__2; static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__11; @@ -586,6 +588,7 @@ LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_addTraceAsMessag uint8_t lean_nat_dec_le(lean_object*, lean_object*); lean_object* l_instHashablePos___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_instMonadAlwaysExceptStateRefT_x27___rarg(lean_object*); +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__49; uint8_t lean_usize_dec_lt(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_addRawTrace___rarg___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instMonadTraceOfMonadLift___rarg(lean_object*, lean_object*); @@ -596,46 +599,45 @@ LEAN_EXPORT lean_object* l_Lean_withTraceNode(lean_object*, lean_object*); static lean_object* l_Lean_isTracingEnabledForCore___closed__1; LEAN_EXPORT lean_object* l_Lean_setTraceState(lean_object*); lean_object* lean_nat_add(lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__16; LEAN_EXPORT lean_object* l_Lean_printTraces___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Util_Trace___hyg_3014____closed__24; LEAN_EXPORT lean_object* l_Lean_addTrace(lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__19; +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__33; uint8_t l_Std_DHashMap_Internal_AssocList_contains___at_Lean_NameSSet_insert___spec__6(lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__42; lean_object* l_ReaderT_instMonadExceptOf___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_withTraceNode_x27___spec__1___rarg___lambda__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_addTraceAsMessages___spec__3(lean_object*); +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__13; LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_1368_(lean_object*); static lean_object* l_Lean_resetTraceState___rarg___lambda__1___closed__4; -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__35; lean_object* l_String_toSubstring_x27(lean_object*); lean_object* lean_array_uset(lean_object*, size_t, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__34; +static lean_object* l_Lean_expandTraceMacro___closed__3; +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__40; static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__17; +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__30; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_printTraces___spec__5___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_trace___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_addRawTrace___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_trace___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__12; LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___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*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_format(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__37; +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__39; LEAN_EXPORT lean_object* l_Lean_instMonadAlwaysExceptReaderT(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__45; size_t lean_usize_land(size_t, size_t); LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_withTraceNode_x27___spec__2___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__12; LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_withTraceNode_x27___spec__1___rarg___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__17; -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__21; static lean_object* l_Lean_checkEmoji___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_withStartStop___rarg___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); double lean_float_sub(double, double); +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__16; LEAN_EXPORT lean_object* l_Lean_addTraceAsMessages___rarg___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__21; LEAN_EXPORT lean_object* l_Lean_addTraceAsMessages___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__38; LEAN_EXPORT lean_object* l_Lean_modifyTraces(lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__26; +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__1; lean_object* l_Lean_Option_register___at_Lean_initFn____x40_Lean_Util_Profile___hyg_5____spec__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* _init_l_Lean_instInhabitedTraceElem___closed__1() { _start: @@ -6428,263 +6430,7 @@ x_6 = l_Lean_registerTraceClass(x_1, x_5, x_3, x_4); return x_6; } } -static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("doElemTrace[_]__", 16, 16); -return x_1; -} -} -static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1276____closed__5; -x_2 = l_Lean_doElemTrace_x5b___x5d_______closed__1; -x_3 = l_Lean_Name_mkStr2(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__3() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("andthen", 7, 7); -return x_1; -} -} -static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__4() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_doElemTrace_x5b___x5d_______closed__3; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__5() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("trace[", 6, 6); -return x_1; -} -} -static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__6() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_doElemTrace_x5b___x5d_______closed__5; -x_2 = lean_alloc_ctor(5, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__7() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("ident", 5, 5); -return x_1; -} -} -static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__8() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_doElemTrace_x5b___x5d_______closed__7; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__9() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_doElemTrace_x5b___x5d_______closed__8; -x_2 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__10() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_doElemTrace_x5b___x5d_______closed__4; -x_2 = l_Lean_doElemTrace_x5b___x5d_______closed__6; -x_3 = l_Lean_doElemTrace_x5b___x5d_______closed__9; -x_4 = lean_alloc_ctor(2, 3, 0); -lean_ctor_set(x_4, 0, x_1); -lean_ctor_set(x_4, 1, x_2); -lean_ctor_set(x_4, 2, x_3); -return x_4; -} -} -static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__11() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("]", 1, 1); -return x_1; -} -} -static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__12() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_doElemTrace_x5b___x5d_______closed__11; -x_2 = lean_alloc_ctor(5, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__13() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_doElemTrace_x5b___x5d_______closed__4; -x_2 = l_Lean_doElemTrace_x5b___x5d_______closed__10; -x_3 = l_Lean_doElemTrace_x5b___x5d_______closed__12; -x_4 = lean_alloc_ctor(2, 3, 0); -lean_ctor_set(x_4, 0, x_1); -lean_ctor_set(x_4, 1, x_2); -lean_ctor_set(x_4, 2, x_3); -return x_4; -} -} -static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__14() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("orelse", 6, 6); -return x_1; -} -} -static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__15() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_doElemTrace_x5b___x5d_______closed__14; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__16() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("interpolatedStr", 15, 15); -return x_1; -} -} -static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__17() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_doElemTrace_x5b___x5d_______closed__16; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__18() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("term", 4, 4); -return x_1; -} -} -static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__19() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_doElemTrace_x5b___x5d_______closed__18; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__20() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_doElemTrace_x5b___x5d_______closed__19; -x_2 = lean_unsigned_to_nat(0u); -x_3 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set(x_3, 1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__21() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_doElemTrace_x5b___x5d_______closed__17; -x_2 = l_Lean_doElemTrace_x5b___x5d_______closed__20; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set(x_3, 1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__22() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_doElemTrace_x5b___x5d_______closed__15; -x_2 = l_Lean_doElemTrace_x5b___x5d_______closed__21; -x_3 = l_Lean_doElemTrace_x5b___x5d_______closed__20; -x_4 = lean_alloc_ctor(2, 3, 0); -lean_ctor_set(x_4, 0, x_1); -lean_ctor_set(x_4, 1, x_2); -lean_ctor_set(x_4, 2, x_3); -return x_4; -} -} -static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__23() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_doElemTrace_x5b___x5d_______closed__4; -x_2 = l_Lean_doElemTrace_x5b___x5d_______closed__13; -x_3 = l_Lean_doElemTrace_x5b___x5d_______closed__22; -x_4 = lean_alloc_ctor(2, 3, 0); -lean_ctor_set(x_4, 0, x_1); -lean_ctor_set(x_4, 1, x_2); -lean_ctor_set(x_4, 2, x_3); -return x_4; -} -} -static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__24() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_doElemTrace_x5b___x5d_______closed__2; -x_2 = lean_unsigned_to_nat(1022u); -x_3 = l_Lean_doElemTrace_x5b___x5d_______closed__23; -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_doElemTrace_x5b___x5d____() { -_start: -{ -lean_object* x_1; -x_1 = l_Lean_doElemTrace_x5b___x5d_______closed__24; -return x_1; -} -} -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__1() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__1() { _start: { lean_object* x_1; @@ -6692,19 +6438,19 @@ x_1 = lean_mk_string_unchecked("doNested", 8, 8); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__2() { +static lean_object* _init_l_Lean_expandTraceMacro___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_initFn____x40_Lean_Util_Trace___hyg_1276____closed__5; x_2 = l___auto____x40_Lean_Util_Trace___hyg_3014____closed__1; x_3 = l___auto____x40_Lean_Util_Trace___hyg_3014____closed__13; -x_4 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__1; +x_4 = l_Lean_expandTraceMacro___lambda__1___closed__1; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__3() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__3() { _start: { lean_object* x_1; @@ -6712,7 +6458,7 @@ x_1 = lean_mk_string_unchecked("do", 2, 2); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__4() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__4() { _start: { lean_object* x_1; @@ -6720,19 +6466,19 @@ x_1 = lean_mk_string_unchecked("doSeqIndent", 11, 11); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__5() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1276____closed__5; x_2 = l___auto____x40_Lean_Util_Trace___hyg_3014____closed__1; x_3 = l___auto____x40_Lean_Util_Trace___hyg_3014____closed__13; -x_4 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__4; +x_4 = l_Lean_expandTraceMacro___lambda__1___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___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__6() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__6() { _start: { lean_object* x_1; @@ -6740,19 +6486,19 @@ x_1 = lean_mk_string_unchecked("doSeqItem", 9, 9); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__7() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1276____closed__5; x_2 = l___auto____x40_Lean_Util_Trace___hyg_3014____closed__1; x_3 = l___auto____x40_Lean_Util_Trace___hyg_3014____closed__13; -x_4 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__6; +x_4 = l_Lean_expandTraceMacro___lambda__1___closed__6; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__8() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__8() { _start: { lean_object* x_1; @@ -6760,19 +6506,19 @@ x_1 = lean_mk_string_unchecked("doLet", 5, 5); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__9() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1276____closed__5; x_2 = l___auto____x40_Lean_Util_Trace___hyg_3014____closed__1; x_3 = l___auto____x40_Lean_Util_Trace___hyg_3014____closed__13; -x_4 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__8; +x_4 = l_Lean_expandTraceMacro___lambda__1___closed__8; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__10() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__10() { _start: { lean_object* x_1; @@ -6780,7 +6526,7 @@ x_1 = lean_mk_string_unchecked("let", 3, 3); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__11() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__11() { _start: { lean_object* x_1; @@ -6788,19 +6534,19 @@ x_1 = lean_mk_string_unchecked("letDecl", 7, 7); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__12() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1276____closed__5; x_2 = l___auto____x40_Lean_Util_Trace___hyg_3014____closed__1; x_3 = l___auto____x40_Lean_Util_Trace___hyg_3014____closed__13; -x_4 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__11; +x_4 = l_Lean_expandTraceMacro___lambda__1___closed__11; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__13() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__13() { _start: { lean_object* x_1; @@ -6808,19 +6554,19 @@ x_1 = lean_mk_string_unchecked("letIdDecl", 9, 9); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__14() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1276____closed__5; x_2 = l___auto____x40_Lean_Util_Trace___hyg_3014____closed__1; x_3 = l___auto____x40_Lean_Util_Trace___hyg_3014____closed__13; -x_4 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__13; +x_4 = l_Lean_expandTraceMacro___lambda__1___closed__13; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__15() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__15() { _start: { lean_object* x_1; @@ -6828,26 +6574,26 @@ x_1 = lean_mk_string_unchecked("cls", 3, 3); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__16() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__16() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__15; +x_1 = l_Lean_expandTraceMacro___lambda__1___closed__15; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__17() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__15; +x_2 = l_Lean_expandTraceMacro___lambda__1___closed__15; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__18() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__18() { _start: { lean_object* x_1; @@ -6855,7 +6601,7 @@ x_1 = lean_mk_string_unchecked(":=", 2, 2); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__19() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__19() { _start: { lean_object* x_1; @@ -6863,19 +6609,19 @@ x_1 = lean_mk_string_unchecked("doIf", 4, 4); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__20() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__20() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1276____closed__5; x_2 = l___auto____x40_Lean_Util_Trace___hyg_3014____closed__1; x_3 = l___auto____x40_Lean_Util_Trace___hyg_3014____closed__13; -x_4 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__19; +x_4 = l_Lean_expandTraceMacro___lambda__1___closed__19; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__21() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__21() { _start: { lean_object* x_1; @@ -6883,7 +6629,7 @@ x_1 = lean_mk_string_unchecked("if", 2, 2); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__22() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__22() { _start: { lean_object* x_1; @@ -6891,19 +6637,19 @@ x_1 = lean_mk_string_unchecked("doIfProp", 8, 8); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__23() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__23() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1276____closed__5; x_2 = l___auto____x40_Lean_Util_Trace___hyg_3014____closed__1; x_3 = l___auto____x40_Lean_Util_Trace___hyg_3014____closed__13; -x_4 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__22; +x_4 = l_Lean_expandTraceMacro___lambda__1___closed__22; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__24() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__24() { _start: { lean_object* x_1; @@ -6911,19 +6657,19 @@ x_1 = lean_mk_string_unchecked("paren", 5, 5); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__25() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__25() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1276____closed__5; x_2 = l___auto____x40_Lean_Util_Trace___hyg_3014____closed__1; x_3 = l___auto____x40_Lean_Util_Trace___hyg_3014____closed__13; -x_4 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__24; +x_4 = l_Lean_expandTraceMacro___lambda__1___closed__24; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__26() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__26() { _start: { lean_object* x_1; @@ -6931,7 +6677,7 @@ x_1 = lean_mk_string_unchecked("(", 1, 1); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__27() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__27() { _start: { lean_object* x_1; @@ -6939,19 +6685,19 @@ x_1 = lean_mk_string_unchecked("liftMethod", 10, 10); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__28() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__28() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1276____closed__5; x_2 = l___auto____x40_Lean_Util_Trace___hyg_3014____closed__1; x_3 = l___auto____x40_Lean_Util_Trace___hyg_3014____closed__13; -x_4 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__27; +x_4 = l_Lean_expandTraceMacro___lambda__1___closed__27; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__29() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__29() { _start: { lean_object* x_1; @@ -6959,7 +6705,7 @@ x_1 = lean_mk_string_unchecked("←", 3, 1); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__30() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__30() { _start: { lean_object* x_1; @@ -6967,19 +6713,19 @@ x_1 = lean_mk_string_unchecked("app", 3, 3); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__31() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__31() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1276____closed__5; x_2 = l___auto____x40_Lean_Util_Trace___hyg_3014____closed__1; x_3 = l___auto____x40_Lean_Util_Trace___hyg_3014____closed__13; -x_4 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__30; +x_4 = l_Lean_expandTraceMacro___lambda__1___closed__30; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__32() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__32() { _start: { lean_object* x_1; @@ -6987,16 +6733,16 @@ x_1 = lean_mk_string_unchecked("Lean.isTracingEnabledFor", 24, 24); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__33() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__33() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__32; +x_1 = l_Lean_expandTraceMacro___lambda__1___closed__32; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__34() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__34() { _start: { lean_object* x_1; @@ -7004,41 +6750,41 @@ x_1 = lean_mk_string_unchecked("isTracingEnabledFor", 19, 19); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__35() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__35() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1276____closed__5; -x_2 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__34; +x_2 = l_Lean_expandTraceMacro___lambda__1___closed__34; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__36() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__36() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__35; +x_2 = l_Lean_expandTraceMacro___lambda__1___closed__35; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__37() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__37() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__36; +x_2 = l_Lean_expandTraceMacro___lambda__1___closed__36; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__38() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__38() { _start: { lean_object* x_1; @@ -7046,7 +6792,7 @@ x_1 = lean_mk_string_unchecked(")", 1, 1); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__39() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__39() { _start: { lean_object* x_1; @@ -7054,7 +6800,7 @@ x_1 = lean_mk_string_unchecked("then", 4, 4); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__40() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__40() { _start: { lean_object* x_1; @@ -7062,19 +6808,19 @@ x_1 = lean_mk_string_unchecked("doExpr", 6, 6); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__41() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__41() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1276____closed__5; x_2 = l___auto____x40_Lean_Util_Trace___hyg_3014____closed__1; x_3 = l___auto____x40_Lean_Util_Trace___hyg_3014____closed__13; -x_4 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__40; +x_4 = l_Lean_expandTraceMacro___lambda__1___closed__40; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__42() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__42() { _start: { lean_object* x_1; @@ -7082,16 +6828,16 @@ x_1 = lean_mk_string_unchecked("Lean.addTrace", 13, 13); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__43() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__43() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__42; +x_1 = l_Lean_expandTraceMacro___lambda__1___closed__42; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__44() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__44() { _start: { lean_object* x_1; @@ -7099,41 +6845,41 @@ x_1 = lean_mk_string_unchecked("addTrace", 8, 8); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__45() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__45() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1276____closed__5; -x_2 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__44; +x_2 = l_Lean_expandTraceMacro___lambda__1___closed__44; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__46() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__46() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__45; +x_2 = l_Lean_expandTraceMacro___lambda__1___closed__45; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__47() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__47() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__46; +x_2 = l_Lean_expandTraceMacro___lambda__1___closed__46; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__48() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__48() { _start: { lean_object* x_1; @@ -7141,19 +6887,19 @@ x_1 = lean_mk_string_unchecked("quotedName", 10, 10); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__49() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__49() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1276____closed__5; x_2 = l___auto____x40_Lean_Util_Trace___hyg_3014____closed__1; x_3 = l___auto____x40_Lean_Util_Trace___hyg_3014____closed__13; -x_4 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__48; +x_4 = l_Lean_expandTraceMacro___lambda__1___closed__48; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__50() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__50() { _start: { lean_object* x_1; @@ -7161,7 +6907,7 @@ x_1 = lean_mk_string_unchecked(".", 1, 1); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__51() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__51() { _start: { lean_object* x_1; @@ -7169,7 +6915,7 @@ x_1 = lean_mk_string_unchecked("`", 1, 1); return x_1; } } -LEAN_EXPORT lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_expandTraceMacro___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; uint8_t x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; @@ -7183,12 +6929,12 @@ lean_inc(x_8); x_9 = lean_ctor_get(x_3, 1); lean_inc(x_9); lean_dec(x_3); -x_10 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__3; +x_10 = l_Lean_expandTraceMacro___lambda__1___closed__3; lean_inc(x_7); x_11 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_11, 0, x_7); lean_ctor_set(x_11, 1, x_10); -x_12 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__10; +x_12 = l_Lean_expandTraceMacro___lambda__1___closed__10; lean_inc(x_7); x_13 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_13, 0, x_7); @@ -7200,19 +6946,19 @@ x_16 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_16, 0, x_7); lean_ctor_set(x_16, 1, x_14); lean_ctor_set(x_16, 2, x_15); -x_17 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__17; +x_17 = l_Lean_expandTraceMacro___lambda__1___closed__17; lean_inc(x_8); lean_inc(x_9); x_18 = l_Lean_addMacroScope(x_9, x_17, x_8); x_19 = lean_box(0); -x_20 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__16; +x_20 = l_Lean_expandTraceMacro___lambda__1___closed__16; lean_inc(x_7); x_21 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_21, 0, x_7); lean_ctor_set(x_21, 1, x_20); lean_ctor_set(x_21, 2, x_18); lean_ctor_set(x_21, 3, x_19); -x_22 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__18; +x_22 = l_Lean_expandTraceMacro___lambda__1___closed__18; lean_inc(x_7); x_23 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_23, 0, x_7); @@ -7221,27 +6967,27 @@ x_24 = l_Lean_Syntax_getId(x_1); x_25 = lean_erase_macro_scopes(x_24); lean_inc(x_25); x_26 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_19, x_25); -x_27 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__21; +x_27 = l_Lean_expandTraceMacro___lambda__1___closed__21; lean_inc(x_7); x_28 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_28, 0, x_7); lean_ctor_set(x_28, 1, x_27); -x_29 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__26; +x_29 = l_Lean_expandTraceMacro___lambda__1___closed__26; lean_inc(x_7); x_30 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_30, 0, x_7); lean_ctor_set(x_30, 1, x_29); -x_31 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__29; +x_31 = l_Lean_expandTraceMacro___lambda__1___closed__29; lean_inc(x_7); x_32 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_32, 0, x_7); lean_ctor_set(x_32, 1, x_31); -x_33 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__35; +x_33 = l_Lean_expandTraceMacro___lambda__1___closed__35; lean_inc(x_8); lean_inc(x_9); x_34 = l_Lean_addMacroScope(x_9, x_33, x_8); -x_35 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__33; -x_36 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__37; +x_35 = l_Lean_expandTraceMacro___lambda__1___closed__33; +x_36 = l_Lean_expandTraceMacro___lambda__1___closed__37; lean_inc(x_7); x_37 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_37, 0, x_7); @@ -7251,33 +6997,33 @@ lean_ctor_set(x_37, 3, x_36); lean_inc(x_21); lean_inc(x_7); x_38 = l_Lean_Syntax_node1(x_7, x_14, x_21); -x_39 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__31; +x_39 = l_Lean_expandTraceMacro___lambda__1___closed__31; lean_inc(x_7); x_40 = l_Lean_Syntax_node2(x_7, x_39, x_37, x_38); -x_41 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__28; +x_41 = l_Lean_expandTraceMacro___lambda__1___closed__28; lean_inc(x_7); x_42 = l_Lean_Syntax_node2(x_7, x_41, x_32, x_40); -x_43 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__38; +x_43 = l_Lean_expandTraceMacro___lambda__1___closed__38; lean_inc(x_7); x_44 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_44, 0, x_7); lean_ctor_set(x_44, 1, x_43); -x_45 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__25; +x_45 = l_Lean_expandTraceMacro___lambda__1___closed__25; lean_inc(x_7); x_46 = l_Lean_Syntax_node3(x_7, x_45, x_30, x_42, x_44); -x_47 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__23; +x_47 = l_Lean_expandTraceMacro___lambda__1___closed__23; lean_inc(x_16); lean_inc(x_7); x_48 = l_Lean_Syntax_node2(x_7, x_47, x_16, x_46); -x_49 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__39; +x_49 = l_Lean_expandTraceMacro___lambda__1___closed__39; lean_inc(x_7); x_50 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_50, 0, x_7); lean_ctor_set(x_50, 1, x_49); -x_51 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__45; +x_51 = l_Lean_expandTraceMacro___lambda__1___closed__45; x_52 = l_Lean_addMacroScope(x_9, x_51, x_8); -x_53 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__43; -x_54 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__47; +x_53 = l_Lean_expandTraceMacro___lambda__1___closed__43; +x_54 = l_Lean_expandTraceMacro___lambda__1___closed__47; lean_inc(x_7); x_55 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_55, 0, x_7); @@ -7289,19 +7035,19 @@ lean_inc(x_7); x_56 = l_Lean_Syntax_node2(x_7, x_14, x_21, x_2); lean_inc(x_7); x_57 = l_Lean_Syntax_node2(x_7, x_39, x_55, x_56); -x_58 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__41; +x_58 = l_Lean_expandTraceMacro___lambda__1___closed__41; lean_inc(x_7); x_59 = l_Lean_Syntax_node1(x_7, x_58, x_57); -x_60 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__7; +x_60 = l_Lean_expandTraceMacro___lambda__1___closed__7; lean_inc(x_16); lean_inc(x_7); x_61 = l_Lean_Syntax_node2(x_7, x_60, x_59, x_16); lean_inc(x_7); x_62 = l_Lean_Syntax_node1(x_7, x_14, x_61); -x_63 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__5; +x_63 = l_Lean_expandTraceMacro___lambda__1___closed__5; lean_inc(x_7); x_64 = l_Lean_Syntax_node1(x_7, x_63, x_62); -x_65 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__20; +x_65 = l_Lean_expandTraceMacro___lambda__1___closed__20; lean_inc_n(x_16, 2); lean_inc(x_7); x_66 = l_Lean_Syntax_node6(x_7, x_65, x_28, x_48, x_50, x_64, x_16, x_16); @@ -7312,14 +7058,14 @@ if (lean_obj_tag(x_26) == 0) { lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; x_68 = l_Lean_quoteNameMk(x_25); -x_69 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__14; +x_69 = l_Lean_expandTraceMacro___lambda__1___closed__14; lean_inc_n(x_16, 2); lean_inc(x_7); x_70 = l_Lean_Syntax_node5(x_7, x_69, x_21, x_16, x_16, x_23, x_68); -x_71 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__12; +x_71 = l_Lean_expandTraceMacro___lambda__1___closed__12; lean_inc(x_7); x_72 = l_Lean_Syntax_node1(x_7, x_71, x_70); -x_73 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__9; +x_73 = l_Lean_expandTraceMacro___lambda__1___closed__9; lean_inc(x_16); lean_inc(x_7); x_74 = l_Lean_Syntax_node3(x_7, x_73, x_13, x_16, x_72); @@ -7329,7 +7075,7 @@ lean_inc(x_7); x_76 = l_Lean_Syntax_node2(x_7, x_14, x_75, x_67); lean_inc(x_7); x_77 = l_Lean_Syntax_node1(x_7, x_63, x_76); -x_78 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__2; +x_78 = l_Lean_expandTraceMacro___lambda__1___closed__2; x_79 = l_Lean_Syntax_node2(x_7, x_78, x_11, x_77); x_80 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_80, 0, x_79); @@ -7343,9 +7089,9 @@ lean_dec(x_25); x_81 = lean_ctor_get(x_26, 0); lean_inc(x_81); lean_dec(x_26); -x_82 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__50; +x_82 = l_Lean_expandTraceMacro___lambda__1___closed__50; x_83 = l_String_intercalate(x_82, x_81); -x_84 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__51; +x_84 = l_Lean_expandTraceMacro___lambda__1___closed__51; x_85 = lean_string_append(x_84, x_83); lean_dec(x_83); x_86 = lean_box(2); @@ -7354,19 +7100,19 @@ x_88 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_88, 0, x_87); lean_ctor_set(x_88, 1, x_19); x_89 = lean_array_mk(x_88); -x_90 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__49; +x_90 = l_Lean_expandTraceMacro___lambda__1___closed__49; x_91 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_91, 0, x_86); lean_ctor_set(x_91, 1, x_90); lean_ctor_set(x_91, 2, x_89); -x_92 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__14; +x_92 = l_Lean_expandTraceMacro___lambda__1___closed__14; lean_inc_n(x_16, 2); lean_inc(x_7); x_93 = l_Lean_Syntax_node5(x_7, x_92, x_21, x_16, x_16, x_23, x_91); -x_94 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__12; +x_94 = l_Lean_expandTraceMacro___lambda__1___closed__12; lean_inc(x_7); x_95 = l_Lean_Syntax_node1(x_7, x_94, x_93); -x_96 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__9; +x_96 = l_Lean_expandTraceMacro___lambda__1___closed__9; lean_inc(x_16); lean_inc(x_7); x_97 = l_Lean_Syntax_node3(x_7, x_96, x_13, x_16, x_95); @@ -7376,7 +7122,7 @@ lean_inc(x_7); x_99 = l_Lean_Syntax_node2(x_7, x_14, x_98, x_67); lean_inc(x_7); x_100 = l_Lean_Syntax_node1(x_7, x_63, x_99); -x_101 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__2; +x_101 = l_Lean_expandTraceMacro___lambda__1___closed__2; x_102 = l_Lean_Syntax_node2(x_7, x_101, x_11, x_100); x_103 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_103, 0, x_102); @@ -7385,7 +7131,7 @@ return x_103; } } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__1() { +static lean_object* _init_l_Lean_expandTraceMacro___closed__1() { _start: { lean_object* x_1; @@ -7393,17 +7139,17 @@ x_1 = lean_mk_string_unchecked("interpolatedStrKind", 19, 19); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__2() { +static lean_object* _init_l_Lean_expandTraceMacro___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__1; +x_2 = l_Lean_expandTraceMacro___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__3() { +static lean_object* _init_l_Lean_expandTraceMacro___closed__3() { _start: { lean_object* x_1; @@ -7411,19 +7157,19 @@ x_1 = lean_mk_string_unchecked("typeAscription", 14, 14); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__4() { +static lean_object* _init_l_Lean_expandTraceMacro___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1276____closed__5; x_2 = l___auto____x40_Lean_Util_Trace___hyg_3014____closed__1; x_3 = l___auto____x40_Lean_Util_Trace___hyg_3014____closed__13; -x_4 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__3; +x_4 = l_Lean_expandTraceMacro___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___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__5() { +static lean_object* _init_l_Lean_expandTraceMacro___closed__5() { _start: { lean_object* x_1; @@ -7431,7 +7177,7 @@ x_1 = lean_mk_string_unchecked(":", 1, 1); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__6() { +static lean_object* _init_l_Lean_expandTraceMacro___closed__6() { _start: { lean_object* x_1; @@ -7439,82 +7185,82 @@ x_1 = lean_mk_string_unchecked("MessageData", 11, 11); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__7() { +static lean_object* _init_l_Lean_expandTraceMacro___closed__7() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__6; +x_1 = l_Lean_expandTraceMacro___closed__6; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__8() { +static lean_object* _init_l_Lean_expandTraceMacro___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__6; +x_2 = l_Lean_expandTraceMacro___closed__6; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__9() { +static lean_object* _init_l_Lean_expandTraceMacro___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1276____closed__5; -x_2 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__6; +x_2 = l_Lean_expandTraceMacro___closed__6; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__10() { +static lean_object* _init_l_Lean_expandTraceMacro___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__9; +x_2 = l_Lean_expandTraceMacro___closed__9; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__11() { +static lean_object* _init_l_Lean_expandTraceMacro___closed__11() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__9; +x_1 = l_Lean_expandTraceMacro___closed__9; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__12() { +static lean_object* _init_l_Lean_expandTraceMacro___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__11; +x_2 = l_Lean_expandTraceMacro___closed__11; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__13() { +static lean_object* _init_l_Lean_expandTraceMacro___closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__10; -x_2 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__12; +x_1 = l_Lean_expandTraceMacro___closed__10; +x_2 = l_Lean_expandTraceMacro___closed__12; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__14() { +static lean_object* _init_l_Lean_expandTraceMacro___closed__14() { _start: { lean_object* x_1; @@ -7522,17 +7268,17 @@ x_1 = lean_mk_string_unchecked("termM!_", 7, 7); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__15() { +static lean_object* _init_l_Lean_expandTraceMacro___closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1276____closed__5; -x_2 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__14; +x_2 = l_Lean_expandTraceMacro___closed__14; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__16() { +static lean_object* _init_l_Lean_expandTraceMacro___closed__16() { _start: { lean_object* x_1; @@ -7540,6 +7286,354 @@ x_1 = lean_mk_string_unchecked("m!", 2, 2); return x_1; } } +LEAN_EXPORT lean_object* l_Lean_expandTraceMacro(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; lean_object* x_6; uint8_t x_7; +lean_inc(x_2); +x_5 = l_Lean_Syntax_getKind(x_2); +x_6 = l_Lean_expandTraceMacro___closed__2; +x_7 = lean_name_eq(x_5, x_6); +lean_dec(x_5); +if (x_7 == 0) +{ +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; +x_8 = lean_ctor_get(x_3, 5); +lean_inc(x_8); +x_9 = 0; +x_10 = l_Lean_SourceInfo_fromRef(x_8, x_9); +lean_dec(x_8); +x_11 = lean_ctor_get(x_3, 2); +lean_inc(x_11); +x_12 = lean_ctor_get(x_3, 1); +lean_inc(x_12); +x_13 = l_Lean_expandTraceMacro___lambda__1___closed__26; +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_expandTraceMacro___closed__5; +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_Lean_expandTraceMacro___closed__8; +x_18 = l_Lean_addMacroScope(x_12, x_17, x_11); +x_19 = l_Lean_expandTraceMacro___closed__7; +x_20 = l_Lean_expandTraceMacro___closed__13; +lean_inc(x_10); +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___auto____x40_Lean_Util_Trace___hyg_3014____closed__8; +lean_inc(x_10); +x_23 = l_Lean_Syntax_node1(x_10, x_22, x_21); +x_24 = l_Lean_expandTraceMacro___lambda__1___closed__38; +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_Lean_expandTraceMacro___closed__4; +x_27 = l_Lean_Syntax_node5(x_10, x_26, x_14, x_2, x_16, x_23, x_25); +x_28 = l_Lean_expandTraceMacro___lambda__1(x_1, x_27, x_3, x_4); +return x_28; +} +else +{ +lean_object* x_29; uint8_t x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_29 = lean_ctor_get(x_3, 5); +lean_inc(x_29); +x_30 = 0; +x_31 = l_Lean_SourceInfo_fromRef(x_29, x_30); +lean_dec(x_29); +x_32 = l_Lean_expandTraceMacro___closed__16; +lean_inc(x_31); +x_33 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_33, 0, x_31); +lean_ctor_set(x_33, 1, x_32); +x_34 = l_Lean_expandTraceMacro___closed__15; +x_35 = l_Lean_Syntax_node2(x_31, x_34, x_33, x_2); +x_36 = l_Lean_expandTraceMacro___lambda__1(x_1, x_35, x_3, x_4); +return x_36; +} +} +} +LEAN_EXPORT lean_object* l_Lean_expandTraceMacro___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_expandTraceMacro___lambda__1(x_1, x_2, x_3, x_4); +lean_dec(x_1); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_expandTraceMacro___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_expandTraceMacro(x_1, x_2, x_3, x_4); +lean_dec(x_1); +return x_5; +} +} +static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("doElemTrace[_]__", 16, 16); +return x_1; +} +} +static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1276____closed__5; +x_2 = l_Lean_doElemTrace_x5b___x5d_______closed__1; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("andthen", 7, 7); +return x_1; +} +} +static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_doElemTrace_x5b___x5d_______closed__3; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("trace[", 6, 6); +return x_1; +} +} +static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_doElemTrace_x5b___x5d_______closed__5; +x_2 = lean_alloc_ctor(5, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("ident", 5, 5); +return x_1; +} +} +static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_doElemTrace_x5b___x5d_______closed__7; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_doElemTrace_x5b___x5d_______closed__8; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_doElemTrace_x5b___x5d_______closed__4; +x_2 = l_Lean_doElemTrace_x5b___x5d_______closed__6; +x_3 = l_Lean_doElemTrace_x5b___x5d_______closed__9; +x_4 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__11() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("]", 1, 1); +return x_1; +} +} +static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_doElemTrace_x5b___x5d_______closed__11; +x_2 = lean_alloc_ctor(5, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_doElemTrace_x5b___x5d_______closed__4; +x_2 = l_Lean_doElemTrace_x5b___x5d_______closed__10; +x_3 = l_Lean_doElemTrace_x5b___x5d_______closed__12; +x_4 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__14() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("orelse", 6, 6); +return x_1; +} +} +static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__15() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_doElemTrace_x5b___x5d_______closed__14; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__16() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("interpolatedStr", 15, 15); +return x_1; +} +} +static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__17() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_doElemTrace_x5b___x5d_______closed__16; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__18() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("term", 4, 4); +return x_1; +} +} +static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__19() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_doElemTrace_x5b___x5d_______closed__18; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__20() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_doElemTrace_x5b___x5d_______closed__19; +x_2 = lean_unsigned_to_nat(0u); +x_3 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__21() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_doElemTrace_x5b___x5d_______closed__17; +x_2 = l_Lean_doElemTrace_x5b___x5d_______closed__20; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__22() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_doElemTrace_x5b___x5d_______closed__15; +x_2 = l_Lean_doElemTrace_x5b___x5d_______closed__21; +x_3 = l_Lean_doElemTrace_x5b___x5d_______closed__20; +x_4 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__23() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_doElemTrace_x5b___x5d_______closed__4; +x_2 = l_Lean_doElemTrace_x5b___x5d_______closed__13; +x_3 = l_Lean_doElemTrace_x5b___x5d_______closed__22; +x_4 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__24() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_doElemTrace_x5b___x5d_______closed__2; +x_2 = lean_unsigned_to_nat(1022u); +x_3 = l_Lean_doElemTrace_x5b___x5d_______closed__23; +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_doElemTrace_x5b___x5d____() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_doElemTrace_x5b___x5d_______closed__24; +return x_1; +} +} LEAN_EXPORT lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { @@ -7560,143 +7654,33 @@ return x_7; } else { -lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; uint8_t x_15; +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; x_8 = lean_unsigned_to_nat(1u); x_9 = l_Lean_Syntax_getArg(x_1, x_8); x_10 = lean_unsigned_to_nat(3u); x_11 = l_Lean_Syntax_getArg(x_1, x_10); lean_dec(x_1); -lean_inc(x_11); -x_12 = l_Lean_Syntax_getKind(x_11); -x_13 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__2; -x_14 = lean_name_eq(x_12, x_13); -lean_dec(x_12); -if (x_14 == 0) -{ -uint8_t x_54; -x_54 = 0; -x_15 = x_54; -goto block_53; -} -else -{ -uint8_t x_55; -x_55 = 1; -x_15 = x_55; -goto block_53; -} -block_53: -{ -if (x_15 == 0) -{ -lean_object* x_16; uint8_t x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; -x_16 = lean_ctor_get(x_2, 5); -lean_inc(x_16); -x_17 = 0; -x_18 = l_Lean_SourceInfo_fromRef(x_16, x_17); -lean_dec(x_16); -x_19 = lean_ctor_get(x_2, 2); -lean_inc(x_19); -x_20 = lean_ctor_get(x_2, 1); -lean_inc(x_20); -x_21 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__26; -lean_inc(x_18); -x_22 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_22, 0, x_18); -lean_ctor_set(x_22, 1, x_21); -x_23 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__5; -lean_inc(x_18); -x_24 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_24, 0, x_18); -lean_ctor_set(x_24, 1, x_23); -x_25 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__8; -x_26 = l_Lean_addMacroScope(x_20, x_25, x_19); -x_27 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__7; -x_28 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__13; -lean_inc(x_18); -x_29 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_29, 0, x_18); -lean_ctor_set(x_29, 1, x_27); -lean_ctor_set(x_29, 2, x_26); -lean_ctor_set(x_29, 3, x_28); -x_30 = l___auto____x40_Lean_Util_Trace___hyg_3014____closed__8; -lean_inc(x_18); -x_31 = l_Lean_Syntax_node1(x_18, x_30, x_29); -x_32 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__38; -lean_inc(x_18); -x_33 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_33, 0, x_18); -lean_ctor_set(x_33, 1, x_32); -x_34 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__4; -x_35 = l_Lean_Syntax_node5(x_18, x_34, x_22, x_11, x_24, x_31, x_33); -x_36 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1(x_9, x_35, x_2, x_3); +x_12 = l_Lean_expandTraceMacro(x_9, x_11, x_2, x_3); lean_dec(x_9); -x_37 = !lean_is_exclusive(x_36); -if (x_37 == 0) -{ -return x_36; -} -else -{ -lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_38 = lean_ctor_get(x_36, 0); -x_39 = lean_ctor_get(x_36, 1); -lean_inc(x_39); -lean_inc(x_38); -lean_dec(x_36); -x_40 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_40, 0, x_38); -lean_ctor_set(x_40, 1, x_39); -return x_40; -} -} -else -{ -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; uint8_t x_49; -x_41 = lean_ctor_get(x_2, 5); -lean_inc(x_41); -x_42 = 0; -x_43 = l_Lean_SourceInfo_fromRef(x_41, x_42); -lean_dec(x_41); -x_44 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__16; -lean_inc(x_43); -x_45 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_45, 0, x_43); -lean_ctor_set(x_45, 1, x_44); -x_46 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__15; -x_47 = l_Lean_Syntax_node2(x_43, x_46, x_45, x_11); -x_48 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1(x_9, x_47, x_2, x_3); -lean_dec(x_9); -x_49 = !lean_is_exclusive(x_48); -if (x_49 == 0) +x_13 = !lean_is_exclusive(x_12); +if (x_13 == 0) { -return x_48; +return x_12; } else { -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_dec(x_48); -x_52 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_52, 0, x_50); -lean_ctor_set(x_52, 1, x_51); -return x_52; -} -} -} -} +lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_14 = lean_ctor_get(x_12, 0); +x_15 = lean_ctor_get(x_12, 1); +lean_inc(x_15); +lean_inc(x_14); +lean_dec(x_12); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_14); +lean_ctor_set(x_16, 1, x_15); +return x_16; } } -LEAN_EXPORT lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -lean_object* x_5; -x_5 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1(x_1, x_2, x_3, x_4); -lean_dec(x_1); -return x_5; } } static lean_object* _init_l_Lean_bombEmoji___closed__1() { @@ -11313,6 +11297,140 @@ l_Lean_registerTraceClass___closed__1 = _init_l_Lean_registerTraceClass___closed lean_mark_persistent(l_Lean_registerTraceClass___closed__1); l_Lean_registerTraceClass___closed__2 = _init_l_Lean_registerTraceClass___closed__2(); lean_mark_persistent(l_Lean_registerTraceClass___closed__2); +l_Lean_expandTraceMacro___lambda__1___closed__1 = _init_l_Lean_expandTraceMacro___lambda__1___closed__1(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__1); +l_Lean_expandTraceMacro___lambda__1___closed__2 = _init_l_Lean_expandTraceMacro___lambda__1___closed__2(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__2); +l_Lean_expandTraceMacro___lambda__1___closed__3 = _init_l_Lean_expandTraceMacro___lambda__1___closed__3(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__3); +l_Lean_expandTraceMacro___lambda__1___closed__4 = _init_l_Lean_expandTraceMacro___lambda__1___closed__4(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__4); +l_Lean_expandTraceMacro___lambda__1___closed__5 = _init_l_Lean_expandTraceMacro___lambda__1___closed__5(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__5); +l_Lean_expandTraceMacro___lambda__1___closed__6 = _init_l_Lean_expandTraceMacro___lambda__1___closed__6(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__6); +l_Lean_expandTraceMacro___lambda__1___closed__7 = _init_l_Lean_expandTraceMacro___lambda__1___closed__7(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__7); +l_Lean_expandTraceMacro___lambda__1___closed__8 = _init_l_Lean_expandTraceMacro___lambda__1___closed__8(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__8); +l_Lean_expandTraceMacro___lambda__1___closed__9 = _init_l_Lean_expandTraceMacro___lambda__1___closed__9(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__9); +l_Lean_expandTraceMacro___lambda__1___closed__10 = _init_l_Lean_expandTraceMacro___lambda__1___closed__10(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__10); +l_Lean_expandTraceMacro___lambda__1___closed__11 = _init_l_Lean_expandTraceMacro___lambda__1___closed__11(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__11); +l_Lean_expandTraceMacro___lambda__1___closed__12 = _init_l_Lean_expandTraceMacro___lambda__1___closed__12(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__12); +l_Lean_expandTraceMacro___lambda__1___closed__13 = _init_l_Lean_expandTraceMacro___lambda__1___closed__13(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__13); +l_Lean_expandTraceMacro___lambda__1___closed__14 = _init_l_Lean_expandTraceMacro___lambda__1___closed__14(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__14); +l_Lean_expandTraceMacro___lambda__1___closed__15 = _init_l_Lean_expandTraceMacro___lambda__1___closed__15(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__15); +l_Lean_expandTraceMacro___lambda__1___closed__16 = _init_l_Lean_expandTraceMacro___lambda__1___closed__16(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__16); +l_Lean_expandTraceMacro___lambda__1___closed__17 = _init_l_Lean_expandTraceMacro___lambda__1___closed__17(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__17); +l_Lean_expandTraceMacro___lambda__1___closed__18 = _init_l_Lean_expandTraceMacro___lambda__1___closed__18(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__18); +l_Lean_expandTraceMacro___lambda__1___closed__19 = _init_l_Lean_expandTraceMacro___lambda__1___closed__19(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__19); +l_Lean_expandTraceMacro___lambda__1___closed__20 = _init_l_Lean_expandTraceMacro___lambda__1___closed__20(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__20); +l_Lean_expandTraceMacro___lambda__1___closed__21 = _init_l_Lean_expandTraceMacro___lambda__1___closed__21(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__21); +l_Lean_expandTraceMacro___lambda__1___closed__22 = _init_l_Lean_expandTraceMacro___lambda__1___closed__22(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__22); +l_Lean_expandTraceMacro___lambda__1___closed__23 = _init_l_Lean_expandTraceMacro___lambda__1___closed__23(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__23); +l_Lean_expandTraceMacro___lambda__1___closed__24 = _init_l_Lean_expandTraceMacro___lambda__1___closed__24(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__24); +l_Lean_expandTraceMacro___lambda__1___closed__25 = _init_l_Lean_expandTraceMacro___lambda__1___closed__25(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__25); +l_Lean_expandTraceMacro___lambda__1___closed__26 = _init_l_Lean_expandTraceMacro___lambda__1___closed__26(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__26); +l_Lean_expandTraceMacro___lambda__1___closed__27 = _init_l_Lean_expandTraceMacro___lambda__1___closed__27(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__27); +l_Lean_expandTraceMacro___lambda__1___closed__28 = _init_l_Lean_expandTraceMacro___lambda__1___closed__28(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__28); +l_Lean_expandTraceMacro___lambda__1___closed__29 = _init_l_Lean_expandTraceMacro___lambda__1___closed__29(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__29); +l_Lean_expandTraceMacro___lambda__1___closed__30 = _init_l_Lean_expandTraceMacro___lambda__1___closed__30(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__30); +l_Lean_expandTraceMacro___lambda__1___closed__31 = _init_l_Lean_expandTraceMacro___lambda__1___closed__31(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__31); +l_Lean_expandTraceMacro___lambda__1___closed__32 = _init_l_Lean_expandTraceMacro___lambda__1___closed__32(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__32); +l_Lean_expandTraceMacro___lambda__1___closed__33 = _init_l_Lean_expandTraceMacro___lambda__1___closed__33(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__33); +l_Lean_expandTraceMacro___lambda__1___closed__34 = _init_l_Lean_expandTraceMacro___lambda__1___closed__34(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__34); +l_Lean_expandTraceMacro___lambda__1___closed__35 = _init_l_Lean_expandTraceMacro___lambda__1___closed__35(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__35); +l_Lean_expandTraceMacro___lambda__1___closed__36 = _init_l_Lean_expandTraceMacro___lambda__1___closed__36(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__36); +l_Lean_expandTraceMacro___lambda__1___closed__37 = _init_l_Lean_expandTraceMacro___lambda__1___closed__37(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__37); +l_Lean_expandTraceMacro___lambda__1___closed__38 = _init_l_Lean_expandTraceMacro___lambda__1___closed__38(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__38); +l_Lean_expandTraceMacro___lambda__1___closed__39 = _init_l_Lean_expandTraceMacro___lambda__1___closed__39(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__39); +l_Lean_expandTraceMacro___lambda__1___closed__40 = _init_l_Lean_expandTraceMacro___lambda__1___closed__40(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__40); +l_Lean_expandTraceMacro___lambda__1___closed__41 = _init_l_Lean_expandTraceMacro___lambda__1___closed__41(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__41); +l_Lean_expandTraceMacro___lambda__1___closed__42 = _init_l_Lean_expandTraceMacro___lambda__1___closed__42(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__42); +l_Lean_expandTraceMacro___lambda__1___closed__43 = _init_l_Lean_expandTraceMacro___lambda__1___closed__43(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__43); +l_Lean_expandTraceMacro___lambda__1___closed__44 = _init_l_Lean_expandTraceMacro___lambda__1___closed__44(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__44); +l_Lean_expandTraceMacro___lambda__1___closed__45 = _init_l_Lean_expandTraceMacro___lambda__1___closed__45(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__45); +l_Lean_expandTraceMacro___lambda__1___closed__46 = _init_l_Lean_expandTraceMacro___lambda__1___closed__46(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__46); +l_Lean_expandTraceMacro___lambda__1___closed__47 = _init_l_Lean_expandTraceMacro___lambda__1___closed__47(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__47); +l_Lean_expandTraceMacro___lambda__1___closed__48 = _init_l_Lean_expandTraceMacro___lambda__1___closed__48(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__48); +l_Lean_expandTraceMacro___lambda__1___closed__49 = _init_l_Lean_expandTraceMacro___lambda__1___closed__49(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__49); +l_Lean_expandTraceMacro___lambda__1___closed__50 = _init_l_Lean_expandTraceMacro___lambda__1___closed__50(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__50); +l_Lean_expandTraceMacro___lambda__1___closed__51 = _init_l_Lean_expandTraceMacro___lambda__1___closed__51(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__51); +l_Lean_expandTraceMacro___closed__1 = _init_l_Lean_expandTraceMacro___closed__1(); +lean_mark_persistent(l_Lean_expandTraceMacro___closed__1); +l_Lean_expandTraceMacro___closed__2 = _init_l_Lean_expandTraceMacro___closed__2(); +lean_mark_persistent(l_Lean_expandTraceMacro___closed__2); +l_Lean_expandTraceMacro___closed__3 = _init_l_Lean_expandTraceMacro___closed__3(); +lean_mark_persistent(l_Lean_expandTraceMacro___closed__3); +l_Lean_expandTraceMacro___closed__4 = _init_l_Lean_expandTraceMacro___closed__4(); +lean_mark_persistent(l_Lean_expandTraceMacro___closed__4); +l_Lean_expandTraceMacro___closed__5 = _init_l_Lean_expandTraceMacro___closed__5(); +lean_mark_persistent(l_Lean_expandTraceMacro___closed__5); +l_Lean_expandTraceMacro___closed__6 = _init_l_Lean_expandTraceMacro___closed__6(); +lean_mark_persistent(l_Lean_expandTraceMacro___closed__6); +l_Lean_expandTraceMacro___closed__7 = _init_l_Lean_expandTraceMacro___closed__7(); +lean_mark_persistent(l_Lean_expandTraceMacro___closed__7); +l_Lean_expandTraceMacro___closed__8 = _init_l_Lean_expandTraceMacro___closed__8(); +lean_mark_persistent(l_Lean_expandTraceMacro___closed__8); +l_Lean_expandTraceMacro___closed__9 = _init_l_Lean_expandTraceMacro___closed__9(); +lean_mark_persistent(l_Lean_expandTraceMacro___closed__9); +l_Lean_expandTraceMacro___closed__10 = _init_l_Lean_expandTraceMacro___closed__10(); +lean_mark_persistent(l_Lean_expandTraceMacro___closed__10); +l_Lean_expandTraceMacro___closed__11 = _init_l_Lean_expandTraceMacro___closed__11(); +lean_mark_persistent(l_Lean_expandTraceMacro___closed__11); +l_Lean_expandTraceMacro___closed__12 = _init_l_Lean_expandTraceMacro___closed__12(); +lean_mark_persistent(l_Lean_expandTraceMacro___closed__12); +l_Lean_expandTraceMacro___closed__13 = _init_l_Lean_expandTraceMacro___closed__13(); +lean_mark_persistent(l_Lean_expandTraceMacro___closed__13); +l_Lean_expandTraceMacro___closed__14 = _init_l_Lean_expandTraceMacro___closed__14(); +lean_mark_persistent(l_Lean_expandTraceMacro___closed__14); +l_Lean_expandTraceMacro___closed__15 = _init_l_Lean_expandTraceMacro___closed__15(); +lean_mark_persistent(l_Lean_expandTraceMacro___closed__15); +l_Lean_expandTraceMacro___closed__16 = _init_l_Lean_expandTraceMacro___closed__16(); +lean_mark_persistent(l_Lean_expandTraceMacro___closed__16); l_Lean_doElemTrace_x5b___x5d_______closed__1 = _init_l_Lean_doElemTrace_x5b___x5d_______closed__1(); lean_mark_persistent(l_Lean_doElemTrace_x5b___x5d_______closed__1); l_Lean_doElemTrace_x5b___x5d_______closed__2 = _init_l_Lean_doElemTrace_x5b___x5d_______closed__2(); @@ -11363,140 +11481,6 @@ l_Lean_doElemTrace_x5b___x5d_______closed__24 = _init_l_Lean_doElemTrace_x5b___x lean_mark_persistent(l_Lean_doElemTrace_x5b___x5d_______closed__24); l_Lean_doElemTrace_x5b___x5d____ = _init_l_Lean_doElemTrace_x5b___x5d____(); lean_mark_persistent(l_Lean_doElemTrace_x5b___x5d____); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__1 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__1(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__1); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__2 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__2(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__2); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__3 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__3(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__3); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__4 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__4(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__4); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__5 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__5(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__5); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__6 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__6(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__6); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__7 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__7(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__7); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__8 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__8(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__8); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__9 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__9(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__9); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__10 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__10(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__10); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__11 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__11(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__11); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__12 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__12(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__12); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__13 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__13(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__13); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__14 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__14(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__14); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__15 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__15(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__15); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__16 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__16(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__16); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__17 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__17(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__17); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__18 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__18(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__18); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__19 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__19(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__19); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__20 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__20(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__20); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__21 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__21(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__21); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__22 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__22(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__22); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__23 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__23(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__23); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__24 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__24(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__24); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__25 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__25(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__25); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__26 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__26(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__26); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__27 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__27(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__27); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__28 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__28(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__28); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__29 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__29(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__29); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__30 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__30(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__30); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__31 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__31(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__31); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__32 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__32(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__32); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__33 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__33(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__33); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__34 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__34(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__34); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__35 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__35(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__35); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__36 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__36(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__36); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__37 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__37(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__37); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__38 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__38(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__38); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__39 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__39(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__39); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__40 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__40(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__40); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__41 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__41(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__41); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__42 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__42(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__42); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__43 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__43(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__43); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__44 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__44(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__44); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__45 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__45(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__45); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__46 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__46(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__46); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__47 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__47(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__47); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__48 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__48(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__48); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__49 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__49(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__49); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__50 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__50(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__50); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__51 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__51(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__51); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__1 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__1(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__1); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__2 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__2(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__2); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__3 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__3(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__3); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__4 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__4(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__4); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__5 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__5(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__5); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__6 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__6(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__6); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__7 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__7(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__7); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__8 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__8(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__8); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__9 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__9(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__9); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__10 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__10(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__10); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__11 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__11(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__11); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__12 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__12(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__12); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__13 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__13(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__13); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__14 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__14(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__14); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__15 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__15(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__15); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__16 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__16(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__16); l_Lean_bombEmoji___closed__1 = _init_l_Lean_bombEmoji___closed__1(); lean_mark_persistent(l_Lean_bombEmoji___closed__1); l_Lean_bombEmoji = _init_l_Lean_bombEmoji(); diff --git a/stage0/stdlib/Lean/Widget/InteractiveCode.c b/stage0/stdlib/Lean/Widget/InteractiveCode.c index c8f4c646881a..68e6ce6a0a58 100644 --- a/stage0/stdlib/Lean/Widget/InteractiveCode.c +++ b/stage0/stdlib/Lean/Widget/InteractiveCode.c @@ -72,13 +72,13 @@ lean_object* l_Lean_Json_getStr_x3f(lean_object*); static lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_353____closed__3; static lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_353____closed__29; LEAN_EXPORT lean_object* l_Lean_Widget_DiffTag_toCtorIdx___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_552____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_ppExprTagged___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Widget_CodeWithInfos_mergePosMap___spec__3(lean_object*); static lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_353____closed__21; static lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_353____closed__15; extern lean_object* l_Lean_pp_raw; lean_object* l_Lean_SubExpr_Pos_toString(lean_object*); -lean_object* l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_fromJsonDiffTag____x40_Lean_Widget_InteractiveCode___hyg_68____lambda__2(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_fromJsonDiffTag____x40_Lean_Widget_InteractiveCode___hyg_68____closed__1; lean_object* l_Lean_Json_getObjValD(lean_object*, lean_object*); @@ -99,7 +99,6 @@ LEAN_EXPORT lean_object* l_Lean_Widget_instToJsonDiffTag; LEAN_EXPORT lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_552_(lean_object*); lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); lean_object* l_Lean_PrettyPrinter_Delaborator_delab(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_353____closed__22; LEAN_EXPORT lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_fromJsonDiffTag____x40_Lean_Widget_InteractiveCode___hyg_68____lambda__5(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_353____closed__26; @@ -156,6 +155,7 @@ uint8_t lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_549_; 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_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableSubexprInfo_dec____x40_Lean_Widget_InteractiveCode___hyg_292____spec__2___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_552____spec__1___boxed(lean_object*, lean_object*); extern lean_object* l_Lean_pp_proofs; static lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonDiffTag____x40_Lean_Widget_InteractiveCode___hyg_11____closed__8; static lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_fromJsonDiffTag____x40_Lean_Widget_InteractiveCode___hyg_68____lambda__1___closed__2; @@ -208,6 +208,7 @@ lean_object* lean_array_uset(lean_object*, size_t, lean_object*); static lean_object* l_Lean_Widget_ppExprTagged___lambda__2___closed__5; static lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_353____closed__1; lean_object* l_Lean_Widget_TaggedText_stripTags___rarg(lean_object*); +lean_object* l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_fromJsonDiffTag____x40_Lean_Widget_InteractiveCode___hyg_68____lambda__4___closed__1; static lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_353____closed__31; LEAN_EXPORT lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_fromJsonDiffTag____x40_Lean_Widget_InteractiveCode___hyg_68____lambda__6___boxed(lean_object*, lean_object*, lean_object*); @@ -1410,6 +1411,32 @@ x_1 = l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Widget_Interactiv return x_1; } } +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_552____spec__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 = lean_box(0); +return x_3; +} +else +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_4 = lean_ctor_get(x_2, 0); +lean_inc(x_4); +x_5 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_5, 0, x_1); +lean_ctor_set(x_5, 1, x_4); +x_6 = lean_box(0); +x_7 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_7, 0, x_5); +lean_ctor_set(x_7, 1, x_6); +return x_7; +} +} +} LEAN_EXPORT lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_552_(lean_object* x_1) { _start: { @@ -1435,7 +1462,7 @@ x_11 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_11, 0, x_10); lean_ctor_set(x_11, 1, x_7); x_12 = l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_353____closed__31; -x_13 = l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(x_12, x_4); +x_13 = l_Lean_Json_opt___at___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_552____spec__1(x_12, x_4); x_14 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_14, 0, x_13); lean_ctor_set(x_14, 1, x_7); @@ -1446,11 +1473,20 @@ x_16 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_16, 0, x_8); lean_ctor_set(x_16, 1, x_15); x_17 = l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_fromJsonDiffTag____x40_Lean_Widget_InteractiveCode___hyg_68____closed__1; -x_18 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_16, x_17); +x_18 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_16, x_17); x_19 = l_Lean_Json_mkObj(x_18); return x_19; } } +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_552____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_Widget_InteractiveCode_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_552____spec__1(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} LEAN_EXPORT lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_552____boxed(lean_object* x_1) { _start: { diff --git a/stage0/stdlib/Lean/Widget/InteractiveDiagnostic.c b/stage0/stdlib/Lean/Widget/InteractiveDiagnostic.c index 7306f1166ffd..fd8a9afeea8f 100644 --- a/stage0/stdlib/Lean/Widget/InteractiveDiagnostic.c +++ b/stage0/stdlib/Lean/Widget/InteractiveDiagnostic.c @@ -41,11 +41,11 @@ static lean_object* l_Lean_Widget_instInhabitedMsgEmbed___closed__3; static lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_msgToInteractiveAux_mkContextInfo___closed__1; lean_object* l_Lean_Json_mkObj(lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodableDiagnosticWith_dec____x40_Lean_Widget_InteractiveDiagnostic___hyg_1853____rarg(lean_object*, lean_object*, lean_object*); +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableDiagnosticWith_enc____x40_Lean_Widget_InteractiveDiagnostic___hyg_1853____spec__2___closed__3; static lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveDiagnostic___hyg_769____lambda__2___closed__3; static lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_msgToInteractiveAux_chopUpChildren___closed__2; static lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveDiagnostic___hyg_1938____closed__45; -uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordRange____x40_Lean_Data_Lsp_Basic___hyg_900_(lean_object*, lean_object*); static lean_object* l_Lean_Widget_InteractiveDiagnostic_toDiagnostic_prettyTt___closed__1; static lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveDiagnostic___hyg_1938____closed__37; LEAN_EXPORT lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveDiagnostic___hyg_769____lambda__1(lean_object*, lean_object*, lean_object*); @@ -126,7 +126,6 @@ LEAN_EXPORT uint8_t l_Lean_Lsp_compareByUserVisible___at_Lean_Widget_Interactive lean_object* l_Lean_Json_getStr_x3f(lean_object*); static lean_object* l_Lean_Widget_instRpcEncodableDiagnosticWith_dec____x40_Lean_Widget_InteractiveDiagnostic___hyg_1853____rarg___closed__3; lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); -lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(size_t, size_t, lean_object*); static lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveDiagnostic___hyg_1938____closed__57; lean_object* l_Lean_Widget_goalToInteractive(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveDiagnostic___hyg_1938____closed__3; @@ -135,6 +134,7 @@ static lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget uint8_t lean_string_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_msgToInteractive_fmtToTT___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Widget_msgToInteractive___closed__1; +lean_object* l_Lean_Json_opt___at___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_552____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableMsgEmbed_dec____x40_Lean_Widget_InteractiveDiagnostic___hyg_570____spec__3___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveDiagnostic___hyg_1938____closed__13; static lean_object* l_Lean_Widget_instRpcEncodableDiagnosticWith_enc____x40_Lean_Widget_InteractiveDiagnostic___hyg_1853____rarg___closed__5; @@ -148,11 +148,11 @@ static lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget static lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveDiagnostic___hyg_1938____closed__54; static lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveDiagnostic___hyg_1938____closed__34; static lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_msgToInteractiveAux_go___closed__2; -lean_object* l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(lean_object*, lean_object*); static lean_object* l_Lean_Widget_instRpcEncodableDiagnosticWith_dec____x40_Lean_Widget_InteractiveDiagnostic___hyg_1853____rarg___closed__6; static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_InteractiveDiagnostic___hyg_3275____closed__6; lean_object* lean_nat_to_int(lean_object*); lean_object* l_Subarray_size___rarg(lean_object*); +uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordRange____x40_Lean_Data_Lsp_Basic___hyg_715_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_instTypeNameLazyTraceChildren; static lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_msgToInteractiveAux_go___closed__6; static lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveDiagnostic___hyg_1938____closed__51; @@ -198,7 +198,6 @@ static lean_object* l_Lean_Widget_InteractiveDiagnostic_toDiagnostic_prettyTt___ lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableMsgEmbed_enc____x40_Lean_Widget_InteractiveDiagnostic___hyg_570____spec__4(size_t, size_t, lean_object*, lean_object*); static lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveDiagnostic___hyg_1938____closed__22; -lean_object* l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Widget_instRpcEncodableMsgEmbed_enc____x40_Lean_Widget_InteractiveDiagnostic___hyg_570____lambda__1(lean_object*); lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcCallParams____x40_Lean_Data_Lsp_Extra___hyg_2014____spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableMsgEmbed_dec____x40_Lean_Widget_InteractiveDiagnostic___hyg_570____spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -285,6 +284,7 @@ static lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_DiagnosticWith_ordUserVisible____x40_Lean_Data_Lsp_Diagnostics___hyg_2056____at_Lean_Widget_InteractiveDiagnostic_compareAsDiagnostics___spec__2(lean_object*, lean_object*); static lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_msgToInteractiveAux_chopUpChildren___closed__4; static lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveDiagnostic___hyg_769____lambda__1___closed__1; +lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableDiagnosticWith_dec____x40_Lean_Widget_InteractiveDiagnostic___hyg_1853____spec__6(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_msgToInteractiveAux_go___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_string_length(lean_object*); @@ -367,6 +367,7 @@ lean_object* lean_array_uget(lean_object*, size_t); size_t lean_array_size(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveDiagnostic___hyg_1116_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveDiagnostic___hyg_207_; +lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(size_t, size_t, lean_object*); static lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveDiagnostic___hyg_769____lambda__1___closed__3; LEAN_EXPORT lean_object* l_Lean_Widget_msgToInteractive_fmtToTT___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveDiagnostic___hyg_1938____closed__55; @@ -377,9 +378,7 @@ uint8_t l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_ordDiagnosticSeverity_ static lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveDiagnostic___hyg_1938____closed__11; static lean_object* l_Lean_Widget_msgToInteractiveDiagnostic___closed__10; lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(size_t, size_t, lean_object*); static lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveDiagnostic___hyg_1938____closed__9; -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(lean_object*); static lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveDiagnostic___hyg_1938____closed__8; lean_object* lean_string_append(lean_object*, lean_object*); static lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveDiagnostic___hyg_1938____closed__46; @@ -398,6 +397,7 @@ static lean_object* l_Lean_Widget_instRpcEncodableDiagnosticWith_enc____x40_Lean LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableDiagnosticWith_dec____x40_Lean_Widget_InteractiveDiagnostic___hyg_1853____spec__3(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_msgToInteractiveAux_go___spec__3___boxed(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_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609_(lean_object*); static lean_object* l_Lean_Widget_instRpcEncodableDiagnosticWith_enc____x40_Lean_Widget_InteractiveDiagnostic___hyg_1853____rarg___closed__2; LEAN_EXPORT uint8_t l_Lean_Widget_InteractiveDiagnostic_compareAsDiagnostics(lean_object*, lean_object*); static lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_msgToInteractiveAux_go___closed__8; @@ -428,10 +428,10 @@ static lean_object* l_Lean_Widget_instRpcEncodableMsgEmbed_dec____x40_Lean_Widge lean_object* l_Lean_Elab_ContextInfo_runMetaM___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_infoview_maxTraceChildren; lean_object* l_Lean_Widget_TaggedText_stripTags___rarg(lean_object*); +lean_object* l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableMsgEmbed_enc____x40_Lean_Widget_InteractiveDiagnostic___hyg_570____spec__2(lean_object*, size_t, size_t, lean_object*, lean_object*); static lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveDiagnostic___hyg_1938____closed__58; LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodableStrictOrLazy_enc____x40_Lean_Widget_InteractiveDiagnostic___hyg_4____rarg(lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794_(lean_object*); lean_object* l_Lean_MessageData_format(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_instInhabitedStrictOrLazy___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableMsgEmbed_dec____x40_Lean_Widget_InteractiveDiagnostic___hyg_570____spec__5(size_t, size_t, lean_object*, lean_object*); @@ -2996,7 +2996,7 @@ if (x_8 == 0) lean_object* x_9; size_t x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; x_9 = lean_ctor_get(x_7, 0); x_10 = lean_array_size(x_9); -x_11 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_10, x_6, x_9); +x_11 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_10, x_6, x_9); lean_ctor_set_tag(x_1, 4); lean_ctor_set(x_1, 0, x_11); x_12 = lean_alloc_ctor(0, 1, 0); @@ -3015,7 +3015,7 @@ lean_inc(x_15); lean_inc(x_14); lean_dec(x_7); x_16 = lean_array_size(x_14); -x_17 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_16, x_6, x_14); +x_17 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_16, x_6, x_14); lean_ctor_set_tag(x_1, 4); lean_ctor_set(x_1, 0, x_17); x_18 = lean_alloc_ctor(0, 1, 0); @@ -3050,7 +3050,7 @@ if (lean_is_exclusive(x_24)) { x_27 = lean_box(0); } x_28 = lean_array_size(x_25); -x_29 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_28, x_23, x_25); +x_29 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_28, x_23, x_25); x_30 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_30, 0, x_29); x_31 = lean_alloc_ctor(0, 1, 0); @@ -4047,7 +4047,7 @@ lean_inc(x_10); lean_dec(x_9); x_11 = lean_array_size(x_10); x_12 = 0; -x_13 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_11, x_12, x_10); +x_13 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_11, x_12, x_10); x_14 = !lean_is_exclusive(x_13); if (x_14 == 0) { @@ -4299,7 +4299,7 @@ lean_inc(x_64); lean_dec(x_63); x_65 = lean_array_size(x_64); x_66 = 0; -x_67 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_65, x_66, x_64); +x_67 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_65, x_66, x_64); x_68 = lean_ctor_get(x_67, 0); lean_inc(x_68); if (lean_is_exclusive(x_67)) { @@ -6283,16 +6283,16 @@ lean_ctor_set(x_6, 0, x_4); lean_ctor_set(x_6, 1, x_5); x_7 = lean_ctor_get(x_1, 1); x_8 = l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveDiagnostic___hyg_1938____closed__23; -x_9 = l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(x_8, x_7); +x_9 = l_Lean_Json_opt___at___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_552____spec__1(x_8, x_7); x_10 = lean_ctor_get(x_1, 2); x_11 = l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveDiagnostic___hyg_1938____closed__29; -x_12 = l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(x_11, x_10); +x_12 = l_Lean_Json_opt___at___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_552____spec__1(x_11, x_10); x_13 = lean_ctor_get(x_1, 3); x_14 = l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveDiagnostic___hyg_1938____closed__35; -x_15 = l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(x_14, x_13); +x_15 = l_Lean_Json_opt___at___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_552____spec__1(x_14, x_13); x_16 = lean_ctor_get(x_1, 4); x_17 = l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveDiagnostic___hyg_1938____closed__41; -x_18 = l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(x_17, x_16); +x_18 = l_Lean_Json_opt___at___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_552____spec__1(x_17, x_16); x_19 = lean_ctor_get(x_1, 5); x_20 = l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveDiagnostic___hyg_1938____closed__47; lean_inc(x_19); @@ -6304,13 +6304,13 @@ lean_ctor_set(x_22, 0, x_21); lean_ctor_set(x_22, 1, x_5); x_23 = lean_ctor_get(x_1, 6); x_24 = l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveDiagnostic___hyg_1938____closed__52; -x_25 = l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(x_24, x_23); +x_25 = l_Lean_Json_opt___at___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_552____spec__1(x_24, x_23); x_26 = lean_ctor_get(x_1, 7); x_27 = l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveDiagnostic___hyg_1938____closed__58; -x_28 = l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(x_27, x_26); +x_28 = l_Lean_Json_opt___at___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_552____spec__1(x_27, x_26); x_29 = lean_ctor_get(x_1, 8); x_30 = l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveDiagnostic___hyg_1938____closed__64; -x_31 = l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(x_30, x_29); +x_31 = l_Lean_Json_opt___at___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_552____spec__1(x_30, x_29); x_32 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_32, 0, x_31); lean_ctor_set(x_32, 1, x_5); @@ -6339,7 +6339,7 @@ x_40 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_40, 0, x_6); lean_ctor_set(x_40, 1, x_39); x_41 = l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveDiagnostic___hyg_2479____closed__1; -x_42 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_40, x_41); +x_42 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_40, x_41); x_43 = l_Lean_Json_mkObj(x_42); return x_43; } @@ -6566,7 +6566,7 @@ LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodableDiagnosticWith_enc____x40 lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_165; x_4 = lean_ctor_get(x_2, 0); lean_inc(x_4); -x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_4); +x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_4); x_165 = lean_ctor_get(x_2, 1); lean_inc(x_165); if (lean_obj_tag(x_165) == 0) @@ -6585,7 +6585,7 @@ if (x_167 == 0) { lean_object* x_168; lean_object* x_169; x_168 = lean_ctor_get(x_165, 0); -x_169 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_168); +x_169 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_168); lean_ctor_set(x_165, 0, x_169); x_6 = x_165; x_7 = x_3; @@ -6597,7 +6597,7 @@ lean_object* x_170; lean_object* x_171; lean_object* x_172; x_170 = lean_ctor_get(x_165, 0); lean_inc(x_170); lean_dec(x_165); -x_171 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_170); +x_171 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_170); x_172 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_172, 0, x_171); x_6 = x_172; @@ -6897,7 +6897,7 @@ x_105 = lean_ctor_get(x_103, 1); lean_inc(x_105); lean_dec(x_103); x_106 = lean_array_size(x_104); -x_107 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_106, x_102, x_104); +x_107 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_106, x_102, x_104); x_108 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_108, 0, x_107); lean_ctor_set(x_97, 0, x_108); @@ -6920,7 +6920,7 @@ x_114 = lean_ctor_get(x_112, 1); lean_inc(x_114); lean_dec(x_112); x_115 = lean_array_size(x_113); -x_116 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_115, x_111, x_113); +x_116 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_115, x_111, x_113); x_117 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_117, 0, x_116); x_118 = lean_alloc_ctor(1, 1, 0); @@ -7043,7 +7043,7 @@ if (x_42 == 0) lean_object* x_43; size_t x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; x_43 = lean_ctor_get(x_41, 0); x_44 = lean_array_size(x_43); -x_45 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_44, x_40, x_43); +x_45 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_44, x_40, x_43); x_46 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_46, 0, x_45); lean_ctor_set(x_22, 0, x_46); @@ -7125,7 +7125,7 @@ lean_inc(x_59); lean_inc(x_58); lean_dec(x_41); x_60 = lean_array_size(x_58); -x_61 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_60, x_40, x_58); +x_61 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_60, x_40, x_58); x_62 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_62, 0, x_61); lean_ctor_set(x_22, 0, x_62); @@ -7212,7 +7212,7 @@ if (lean_is_exclusive(x_77)) { x_80 = lean_box(0); } x_81 = lean_array_size(x_78); -x_82 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_81, x_76, x_78); +x_82 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_81, x_76, x_78); x_83 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_83, 0, x_82); x_84 = lean_alloc_ctor(1, 1, 0); @@ -7865,7 +7865,7 @@ if (lean_is_exclusive(x_4)) { } x_7 = lean_ctor_get(x_5, 0); lean_inc(x_7); -x_8 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794_(x_7); +x_8 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609_(x_7); x_9 = l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableInteractiveTermGoal_dec____x40_Lean_Widget_InteractiveGoal___hyg_2227____spec__1(x_8, x_3); if (lean_obj_tag(x_9) == 0) { @@ -7919,7 +7919,7 @@ if (x_2049 == 0) { lean_object* x_2050; lean_object* x_2051; lean_object* x_2052; x_2050 = lean_ctor_get(x_2047, 0); -x_2051 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794_(x_2050); +x_2051 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609_(x_2050); x_2052 = l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableInteractiveTermGoal_dec____x40_Lean_Widget_InteractiveGoal___hyg_2227____spec__1(x_2051, x_3); if (lean_obj_tag(x_2052) == 0) { @@ -7964,7 +7964,7 @@ lean_object* x_2057; lean_object* x_2058; lean_object* x_2059; x_2057 = lean_ctor_get(x_2047, 0); lean_inc(x_2057); lean_dec(x_2047); -x_2058 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794_(x_2057); +x_2058 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609_(x_2057); x_2059 = l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableInteractiveTermGoal_dec____x40_Lean_Widget_InteractiveGoal___hyg_2227____spec__1(x_2058, x_3); if (lean_obj_tag(x_2059) == 0) { @@ -8538,7 +8538,7 @@ lean_inc(x_946); lean_dec(x_931); x_947 = lean_array_size(x_946); x_948 = 0; -x_949 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_947, x_948, x_946); +x_949 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_947, x_948, x_946); x_950 = lean_ctor_get(x_949, 0); lean_inc(x_950); lean_dec(x_949); @@ -8684,7 +8684,7 @@ lean_inc(x_985); lean_dec(x_970); x_986 = lean_array_size(x_985); x_987 = 0; -x_988 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_986, x_987, x_985); +x_988 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_986, x_987, x_985); x_989 = lean_ctor_get(x_988, 0); lean_inc(x_989); lean_dec(x_988); @@ -8974,7 +8974,7 @@ lean_inc(x_1008); lean_dec(x_1006); x_1009 = lean_array_size(x_1008); x_1010 = 0; -x_1011 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_1009, x_1010, x_1008); +x_1011 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_1009, x_1010, x_1008); x_1012 = !lean_is_exclusive(x_1011); if (x_1012 == 0) { @@ -9046,7 +9046,7 @@ x_1053 = lean_ctor_get(x_1052, 0); lean_inc(x_1053); lean_dec(x_1052); x_1054 = lean_array_size(x_1053); -x_1055 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_1054, x_1010, x_1053); +x_1055 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_1054, x_1010, x_1053); x_1056 = lean_ctor_get(x_1055, 0); lean_inc(x_1056); lean_dec(x_1055); @@ -9129,7 +9129,7 @@ x_1070 = lean_ctor_get(x_1069, 0); lean_inc(x_1070); lean_dec(x_1069); x_1071 = lean_array_size(x_1070); -x_1072 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_1071, x_1010, x_1070); +x_1072 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_1071, x_1010, x_1070); x_1073 = lean_ctor_get(x_1072, 0); lean_inc(x_1073); lean_dec(x_1072); @@ -9465,7 +9465,7 @@ x_1115 = lean_ctor_get(x_1113, 0); lean_inc(x_1115); lean_dec(x_1113); x_1116 = lean_array_size(x_1115); -x_1117 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_1116, x_1010, x_1115); +x_1117 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_1116, x_1010, x_1115); x_1118 = lean_ctor_get(x_1117, 0); lean_inc(x_1118); lean_dec(x_1117); @@ -9716,7 +9716,7 @@ lean_inc(x_1150); lean_dec(x_1148); x_1151 = lean_array_size(x_1150); x_1152 = 0; -x_1153 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_1151, x_1152, x_1150); +x_1153 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_1151, x_1152, x_1150); x_1154 = lean_ctor_get(x_1153, 0); lean_inc(x_1154); if (lean_is_exclusive(x_1153)) { @@ -9797,7 +9797,7 @@ x_1184 = lean_ctor_get(x_1182, 0); lean_inc(x_1184); lean_dec(x_1182); x_1185 = lean_array_size(x_1184); -x_1186 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_1185, x_1152, x_1184); +x_1186 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_1185, x_1152, x_1184); x_1187 = lean_ctor_get(x_1186, 0); lean_inc(x_1187); lean_dec(x_1186); @@ -10147,7 +10147,7 @@ lean_inc(x_1260); lean_dec(x_1259); x_1261 = lean_array_size(x_1260); x_1262 = 0; -x_1263 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_1261, x_1262, x_1260); +x_1263 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_1261, x_1262, x_1260); x_1264 = lean_ctor_get(x_1263, 0); lean_inc(x_1264); lean_dec(x_1263); @@ -10231,7 +10231,7 @@ lean_inc(x_1278); lean_dec(x_1277); x_1279 = lean_array_size(x_1278); x_1280 = 0; -x_1281 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_1279, x_1280, x_1278); +x_1281 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_1279, x_1280, x_1278); x_1282 = lean_ctor_get(x_1281, 0); lean_inc(x_1282); lean_dec(x_1281); @@ -10512,7 +10512,7 @@ lean_inc(x_1299); lean_dec(x_1297); x_1300 = lean_array_size(x_1299); x_1301 = 0; -x_1302 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_1300, x_1301, x_1299); +x_1302 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_1300, x_1301, x_1299); x_1303 = !lean_is_exclusive(x_1302); if (x_1303 == 0) { @@ -10585,7 +10585,7 @@ x_1344 = lean_ctor_get(x_1343, 0); lean_inc(x_1344); lean_dec(x_1343); x_1345 = lean_array_size(x_1344); -x_1346 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_1345, x_1301, x_1344); +x_1346 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_1345, x_1301, x_1344); x_1347 = lean_ctor_get(x_1346, 0); lean_inc(x_1347); lean_dec(x_1346); @@ -10670,7 +10670,7 @@ x_1361 = lean_ctor_get(x_1360, 0); lean_inc(x_1361); lean_dec(x_1360); x_1362 = lean_array_size(x_1361); -x_1363 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_1362, x_1301, x_1361); +x_1363 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_1362, x_1301, x_1361); x_1364 = lean_ctor_get(x_1363, 0); lean_inc(x_1364); lean_dec(x_1363); @@ -11011,7 +11011,7 @@ x_1406 = lean_ctor_get(x_1404, 0); lean_inc(x_1406); lean_dec(x_1404); x_1407 = lean_array_size(x_1406); -x_1408 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_1407, x_1301, x_1406); +x_1408 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_1407, x_1301, x_1406); x_1409 = lean_ctor_get(x_1408, 0); lean_inc(x_1409); lean_dec(x_1408); @@ -11266,7 +11266,7 @@ lean_inc(x_1441); lean_dec(x_1439); x_1442 = lean_array_size(x_1441); x_1443 = 0; -x_1444 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_1442, x_1443, x_1441); +x_1444 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_1442, x_1443, x_1441); x_1445 = lean_ctor_get(x_1444, 0); lean_inc(x_1445); if (lean_is_exclusive(x_1444)) { @@ -11348,7 +11348,7 @@ x_1475 = lean_ctor_get(x_1473, 0); lean_inc(x_1475); lean_dec(x_1473); x_1476 = lean_array_size(x_1475); -x_1477 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_1476, x_1443, x_1475); +x_1477 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_1476, x_1443, x_1475); x_1478 = lean_ctor_get(x_1477, 0); lean_inc(x_1478); lean_dec(x_1477); @@ -11666,7 +11666,7 @@ lean_inc(x_1531); lean_dec(x_1529); x_1532 = lean_array_size(x_1531); x_1533 = 0; -x_1534 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_1532, x_1533, x_1531); +x_1534 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_1532, x_1533, x_1531); x_1535 = lean_ctor_get(x_1534, 0); lean_inc(x_1535); lean_dec(x_1534); @@ -11878,7 +11878,7 @@ lean_inc(x_1553); lean_dec(x_1550); x_1554 = lean_array_size(x_1553); x_1555 = 0; -x_1556 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_1554, x_1555, x_1553); +x_1556 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_1554, x_1555, x_1553); x_1557 = lean_ctor_get(x_1556, 0); lean_inc(x_1557); if (lean_is_exclusive(x_1556)) { @@ -11965,7 +11965,7 @@ x_1587 = lean_ctor_get(x_1585, 0); lean_inc(x_1587); lean_dec(x_1585); x_1588 = lean_array_size(x_1587); -x_1589 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_1588, x_1555, x_1587); +x_1589 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_1588, x_1555, x_1587); x_1590 = lean_ctor_get(x_1589, 0); lean_inc(x_1590); lean_dec(x_1589); @@ -12329,7 +12329,7 @@ lean_inc(x_1651); lean_dec(x_1649); x_1652 = lean_array_size(x_1651); x_1653 = 0; -x_1654 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_1652, x_1653, x_1651); +x_1654 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_1652, x_1653, x_1651); x_1655 = lean_ctor_get(x_1654, 0); lean_inc(x_1655); lean_dec(x_1654); @@ -12547,7 +12547,7 @@ lean_inc(x_1673); lean_dec(x_1670); x_1674 = lean_array_size(x_1673); x_1675 = 0; -x_1676 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_1674, x_1675, x_1673); +x_1676 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_1674, x_1675, x_1673); x_1677 = lean_ctor_get(x_1676, 0); lean_inc(x_1677); if (lean_is_exclusive(x_1676)) { @@ -12634,7 +12634,7 @@ x_1707 = lean_ctor_get(x_1705, 0); lean_inc(x_1707); lean_dec(x_1705); x_1708 = lean_array_size(x_1707); -x_1709 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_1708, x_1675, x_1707); +x_1709 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_1708, x_1675, x_1707); x_1710 = lean_ctor_get(x_1709, 0); lean_inc(x_1710); lean_dec(x_1709); @@ -14148,7 +14148,7 @@ lean_inc(x_61); lean_dec(x_60); x_62 = lean_array_size(x_61); x_63 = 0; -x_64 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_62, x_63, x_61); +x_64 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_62, x_63, x_61); x_65 = lean_ctor_get(x_64, 0); lean_inc(x_65); lean_dec(x_64); @@ -14237,7 +14237,7 @@ lean_inc(x_80); lean_dec(x_79); x_81 = lean_array_size(x_80); x_82 = 0; -x_83 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_81, x_82, x_80); +x_83 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_81, x_82, x_80); x_84 = lean_ctor_get(x_83, 0); lean_inc(x_84); lean_dec(x_83); @@ -14523,7 +14523,7 @@ lean_inc(x_102); lean_dec(x_100); x_103 = lean_array_size(x_102); x_104 = 0; -x_105 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_103, x_104, x_102); +x_105 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_103, x_104, x_102); x_106 = !lean_is_exclusive(x_105); if (x_106 == 0) { @@ -14596,7 +14596,7 @@ x_147 = lean_ctor_get(x_146, 0); lean_inc(x_147); lean_dec(x_146); x_148 = lean_array_size(x_147); -x_149 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_148, x_104, x_147); +x_149 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_148, x_104, x_147); x_150 = lean_ctor_get(x_149, 0); lean_inc(x_150); lean_dec(x_149); @@ -14681,7 +14681,7 @@ x_164 = lean_ctor_get(x_163, 0); lean_inc(x_164); lean_dec(x_163); x_165 = lean_array_size(x_164); -x_166 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_165, x_104, x_164); +x_166 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_165, x_104, x_164); x_167 = lean_ctor_get(x_166, 0); lean_inc(x_167); lean_dec(x_166); @@ -15022,7 +15022,7 @@ x_209 = lean_ctor_get(x_207, 0); lean_inc(x_209); lean_dec(x_207); x_210 = lean_array_size(x_209); -x_211 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_210, x_104, x_209); +x_211 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_210, x_104, x_209); x_212 = lean_ctor_get(x_211, 0); lean_inc(x_212); lean_dec(x_211); @@ -15277,7 +15277,7 @@ lean_inc(x_244); lean_dec(x_242); x_245 = lean_array_size(x_244); x_246 = 0; -x_247 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_245, x_246, x_244); +x_247 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_245, x_246, x_244); x_248 = lean_ctor_get(x_247, 0); lean_inc(x_248); if (lean_is_exclusive(x_247)) { @@ -15359,7 +15359,7 @@ x_278 = lean_ctor_get(x_276, 0); lean_inc(x_278); lean_dec(x_276); x_279 = lean_array_size(x_278); -x_280 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_279, x_246, x_278); +x_280 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_279, x_246, x_278); x_281 = lean_ctor_get(x_280, 0); lean_inc(x_281); lean_dec(x_280); @@ -15717,7 +15717,7 @@ lean_inc(x_355); lean_dec(x_354); x_356 = lean_array_size(x_355); x_357 = 0; -x_358 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_356, x_357, x_355); +x_358 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_356, x_357, x_355); x_359 = lean_ctor_get(x_358, 0); lean_inc(x_359); lean_dec(x_358); @@ -15803,7 +15803,7 @@ lean_inc(x_373); lean_dec(x_372); x_374 = lean_array_size(x_373); x_375 = 0; -x_376 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_374, x_375, x_373); +x_376 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_374, x_375, x_373); x_377 = lean_ctor_get(x_376, 0); lean_inc(x_377); lean_dec(x_376); @@ -16088,7 +16088,7 @@ lean_inc(x_394); lean_dec(x_392); x_395 = lean_array_size(x_394); x_396 = 0; -x_397 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_395, x_396, x_394); +x_397 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_395, x_396, x_394); x_398 = !lean_is_exclusive(x_397); if (x_398 == 0) { @@ -16164,7 +16164,7 @@ x_441 = lean_ctor_get(x_440, 0); lean_inc(x_441); lean_dec(x_440); x_442 = lean_array_size(x_441); -x_443 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_442, x_396, x_441); +x_443 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_442, x_396, x_441); x_444 = lean_ctor_get(x_443, 0); lean_inc(x_444); lean_dec(x_443); @@ -16251,7 +16251,7 @@ x_458 = lean_ctor_get(x_457, 0); lean_inc(x_458); lean_dec(x_457); x_459 = lean_array_size(x_458); -x_460 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_459, x_396, x_458); +x_460 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_459, x_396, x_458); x_461 = lean_ctor_get(x_460, 0); lean_inc(x_461); lean_dec(x_460); @@ -16600,7 +16600,7 @@ x_505 = lean_ctor_get(x_503, 0); lean_inc(x_505); lean_dec(x_503); x_506 = lean_array_size(x_505); -x_507 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_506, x_396, x_505); +x_507 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_506, x_396, x_505); x_508 = lean_ctor_get(x_507, 0); lean_inc(x_508); lean_dec(x_507); @@ -16860,7 +16860,7 @@ lean_inc(x_540); lean_dec(x_538); x_541 = lean_array_size(x_540); x_542 = 0; -x_543 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_541, x_542, x_540); +x_543 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_541, x_542, x_540); x_544 = lean_ctor_get(x_543, 0); lean_inc(x_544); if (lean_is_exclusive(x_543)) { @@ -16945,7 +16945,7 @@ x_576 = lean_ctor_get(x_574, 0); lean_inc(x_576); lean_dec(x_574); x_577 = lean_array_size(x_576); -x_578 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_577, x_542, x_576); +x_578 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_577, x_542, x_576); x_579 = lean_ctor_get(x_578, 0); lean_inc(x_579); lean_dec(x_578); @@ -17270,7 +17270,7 @@ lean_inc(x_633); lean_dec(x_631); x_634 = lean_array_size(x_633); x_635 = 0; -x_636 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_634, x_635, x_633); +x_636 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_634, x_635, x_633); x_637 = lean_ctor_get(x_636, 0); lean_inc(x_637); lean_dec(x_636); @@ -17485,7 +17485,7 @@ lean_inc(x_655); lean_dec(x_652); x_656 = lean_array_size(x_655); x_657 = 0; -x_658 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_656, x_657, x_655); +x_658 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_656, x_657, x_655); x_659 = lean_ctor_get(x_658, 0); lean_inc(x_659); if (lean_is_exclusive(x_658)) { @@ -17575,7 +17575,7 @@ x_691 = lean_ctor_get(x_689, 0); lean_inc(x_691); lean_dec(x_689); x_692 = lean_array_size(x_691); -x_693 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_692, x_657, x_691); +x_693 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_692, x_657, x_691); x_694 = lean_ctor_get(x_693, 0); lean_inc(x_694); lean_dec(x_693); @@ -17947,7 +17947,7 @@ lean_inc(x_756); lean_dec(x_754); x_757 = lean_array_size(x_756); x_758 = 0; -x_759 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_757, x_758, x_756); +x_759 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_757, x_758, x_756); x_760 = lean_ctor_get(x_759, 0); lean_inc(x_760); lean_dec(x_759); @@ -18168,7 +18168,7 @@ lean_inc(x_778); lean_dec(x_775); x_779 = lean_array_size(x_778); x_780 = 0; -x_781 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_779, x_780, x_778); +x_781 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_779, x_780, x_778); x_782 = lean_ctor_get(x_781, 0); lean_inc(x_782); if (lean_is_exclusive(x_781)) { @@ -18258,7 +18258,7 @@ x_814 = lean_ctor_get(x_812, 0); lean_inc(x_814); lean_dec(x_812); x_815 = lean_array_size(x_814); -x_816 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_815, x_780, x_814); +x_816 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_815, x_780, x_814); x_817 = lean_ctor_get(x_816, 0); lean_inc(x_817); lean_dec(x_816); @@ -18851,7 +18851,7 @@ lean_inc(x_17); x_18 = lean_ctor_get(x_2, 7); lean_inc(x_18); lean_dec(x_2); -x_19 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordRange____x40_Lean_Data_Lsp_Basic___hyg_900_(x_3, x_11); +x_19 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordRange____x40_Lean_Data_Lsp_Basic___hyg_715_(x_3, x_11); lean_dec(x_11); lean_dec(x_3); x_126 = lean_string_dec_lt(x_8, x_16); @@ -18925,7 +18925,7 @@ lean_dec(x_4); x_123 = lean_ctor_get(x_12, 0); lean_inc(x_123); lean_dec(x_12); -x_124 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordRange____x40_Lean_Data_Lsp_Basic___hyg_900_(x_122, x_123); +x_124 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordRange____x40_Lean_Data_Lsp_Basic___hyg_715_(x_122, x_123); lean_dec(x_123); lean_dec(x_122); x_21 = x_124; diff --git a/stage0/stdlib/Lean/Widget/InteractiveGoal.c b/stage0/stdlib/Lean/Widget/InteractiveGoal.c index 1e0b1177a983..43784d4de622 100644 --- a/stage0/stdlib/Lean/Widget/InteractiveGoal.c +++ b/stage0/stdlib/Lean/Widget/InteractiveGoal.c @@ -37,6 +37,7 @@ LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodableInteractiveHypothesisBund lean_object* l_Lean_Json_mkObj(lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_withGoalCtx___at_Lean_Widget_goalToInteractive___spec__8(lean_object*); static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_2301____closed__16; +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_enc____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__1(size_t, size_t, lean_object*, lean_object*); static lean_object* l_Lean_Widget_instRpcEncodableInteractiveGoals___closed__2; static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_176____closed__38; @@ -122,7 +123,6 @@ static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJ lean_object* l_Lean_Json_getStr_x3f(lean_object*); static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_176____closed__27; LEAN_EXPORT lean_object* l_Lean_Widget_withGoalCtx___rarg___lambda__2(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_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(size_t, size_t, lean_object*); static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_1277____closed__9; static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_1277____closed__6; static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_2301____closed__4; @@ -133,6 +133,7 @@ static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJ LEAN_EXPORT lean_object* l_Lean_Widget_goalToInteractive(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_1277____closed__30; static lean_object* l_Lean_Widget_instRpcEncodableInteractiveGoals___closed__3; +lean_object* l_Lean_Json_opt___at___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_552____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodableInteractiveTermGoal_dec____x40_Lean_Widget_InteractiveGoal___hyg_2227_(lean_object*, lean_object*); static lean_object* l_Lean_Widget_goalToInteractive___lambda__1___closed__1; extern lean_object* l_Lean_Meta_pp_implementationDetailHyps; @@ -141,7 +142,6 @@ static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJ uint8_t l_Lean_Expr_isSort(lean_object*); static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_176____closed__55; LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle; -lean_object* l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(lean_object*, lean_object*); static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_1277____closed__1; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__6(lean_object*, size_t, size_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_CommandContextInfo_saveNoFileMap___at_Lean_Widget_goalToInteractive___spec__7___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -193,7 +193,6 @@ static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJ static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_176____closed__41; LEAN_EXPORT lean_object* l_Lean_Widget_instAppendInteractiveGoals; static lean_object* l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_657____closed__1; -lean_object* l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_176____closed__3; static lean_object* l_Lean_Widget_withGoalCtx___rarg___lambda__2___closed__2; lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcCallParams____x40_Lean_Data_Lsp_Extra___hyg_2014____spec__2(lean_object*, lean_object*); @@ -273,6 +272,7 @@ LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodabl lean_object* l_Lean_throwError___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveGoals_enc____x40_Lean_Widget_InteractiveGoal___hyg_3343____spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_176____closed__17; +lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(size_t, size_t, lean_object*); static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_1277____closed__34; static lean_object* l_Lean_Widget_InteractiveGoalCore_pretty___closed__2; static lean_object* l_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____closed__2; @@ -331,6 +331,7 @@ LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodableInteractiveGoal; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_uget(lean_object*, size_t); size_t lean_array_size(lean_object*); +lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(size_t, size_t, lean_object*); static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_2301____closed__3; LEAN_EXPORT lean_object* l_Lean_Widget_InteractiveGoalCore_pretty(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2(size_t, size_t, lean_object*, lean_object*); @@ -339,8 +340,6 @@ static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJ static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_2301____closed__10; LEAN_EXPORT lean_object* l_Lean_Elab_CommandContextInfo_saveNoFileMap___at_Lean_Widget_goalToInteractive___spec__7(lean_object*); static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_176____closed__40; -lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(size_t, size_t, lean_object*); -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(lean_object*); lean_object* lean_string_append(lean_object*, lean_object*); static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_176____closed__10; static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_176____closed__60; @@ -356,6 +355,7 @@ static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJ static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_176____closed__30; static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_1277____closed__18; uint8_t lean_usize_dec_lt(size_t, size_t); +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609_(lean_object*); lean_object* l_Lean_LocalContext_sanitizeNames(lean_object*, lean_object*); static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_1277____closed__13; static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_176____closed__8; @@ -376,10 +376,10 @@ LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Widget_goalToInt static lean_object* l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_3451____closed__1; static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_176____closed__51; lean_object* l_Lean_Widget_TaggedText_stripTags___rarg(lean_object*); +lean_object* l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_1277_(lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Widget_goalToInteractive___spec__4___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794_(lean_object*); extern lean_object* l_Lean_Widget_instTypeNameTermInfo; LEAN_EXPORT lean_object* l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_2611_; LEAN_EXPORT lean_object* l_Lean_Widget_instInhabitedInteractiveHypothesisBundle; @@ -1228,15 +1228,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_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_176____closed__36; -x_21 = l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(x_20, x_5); +x_21 = l_Lean_Json_opt___at___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_552____spec__1(x_20, x_5); x_22 = l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_176____closed__42; -x_23 = l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(x_22, x_6); +x_23 = l_Lean_Json_opt___at___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_552____spec__1(x_22, x_6); x_24 = l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_176____closed__48; -x_25 = l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(x_24, x_7); +x_25 = l_Lean_Json_opt___at___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_552____spec__1(x_24, x_7); x_26 = l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_176____closed__54; -x_27 = l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(x_26, x_8); +x_27 = l_Lean_Json_opt___at___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_552____spec__1(x_26, x_8); x_28 = l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_176____closed__60; -x_29 = l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(x_28, x_9); +x_29 = l_Lean_Json_opt___at___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_552____spec__1(x_28, x_9); x_30 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_30, 0, x_29); lean_ctor_set(x_30, 1, x_12); @@ -1262,7 +1262,7 @@ x_37 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_37, 0, x_13); lean_ctor_set(x_37, 1, x_36); x_38 = l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_660____closed__1; -x_39 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_37, x_38); +x_39 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_37, x_38); x_40 = l_Lean_Json_mkObj(x_39); return x_40; } @@ -1601,7 +1601,7 @@ x_8 = lean_ctor_get(x_6, 1); lean_inc(x_8); lean_dec(x_6); x_9 = lean_array_size(x_7); -x_10 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_9, x_5, x_7); +x_10 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_9, x_5, x_7); x_11 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_11, 0, x_10); x_12 = lean_ctor_get(x_1, 1); @@ -1614,7 +1614,7 @@ x_16 = lean_ctor_get(x_14, 1); lean_inc(x_16); lean_dec(x_14); x_17 = lean_array_size(x_15); -x_18 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_17, x_5, x_15); +x_18 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_17, x_5, x_15); x_19 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_19, 0, x_18); x_20 = lean_ctor_get(x_1, 2); @@ -3165,7 +3165,7 @@ lean_inc(x_7); lean_dec(x_6); x_8 = lean_array_size(x_7); x_9 = 0; -x_10 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_8, x_9, x_7); +x_10 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_8, x_9, x_7); x_11 = lean_ctor_get(x_10, 0); lean_inc(x_11); lean_dec(x_10); @@ -3207,7 +3207,7 @@ x_19 = lean_ctor_get(x_17, 0); lean_inc(x_19); lean_dec(x_17); x_20 = lean_array_size(x_19); -x_21 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_20, x_9, x_19); +x_21 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_20, x_9, x_19); x_22 = lean_ctor_get(x_21, 0); lean_inc(x_22); lean_dec(x_21); @@ -5592,7 +5592,7 @@ lean_inc(x_418); lean_dec(x_417); x_419 = lean_array_size(x_418); x_420 = 0; -x_421 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_419, x_420, x_418); +x_421 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_419, x_420, x_418); x_422 = lean_ctor_get(x_421, 0); lean_inc(x_422); lean_dec(x_421); @@ -5635,7 +5635,7 @@ x_430 = lean_ctor_get(x_428, 0); lean_inc(x_430); lean_dec(x_428); x_431 = lean_array_size(x_430); -x_432 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_431, x_420, x_430); +x_432 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_431, x_420, x_430); x_433 = lean_ctor_get(x_432, 0); lean_inc(x_433); lean_dec(x_432); @@ -7266,7 +7266,7 @@ 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_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_1277____closed__16; -x_21 = l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(x_20, x_5); +x_21 = l_Lean_Json_opt___at___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_552____spec__1(x_20, x_5); x_22 = l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_1277____closed__22; lean_inc(x_6); x_23 = lean_alloc_ctor(0, 2, 0); @@ -7284,9 +7284,9 @@ x_27 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_27, 0, x_26); lean_ctor_set(x_27, 1, x_12); x_28 = l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_176____closed__54; -x_29 = l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(x_28, x_8); +x_29 = l_Lean_Json_opt___at___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_552____spec__1(x_28, x_8); x_30 = l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_176____closed__60; -x_31 = l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(x_30, x_9); +x_31 = l_Lean_Json_opt___at___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_552____spec__1(x_30, x_9); x_32 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_32, 0, x_31); lean_ctor_set(x_32, 1, x_12); @@ -7312,7 +7312,7 @@ x_39 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_39, 0, x_13); lean_ctor_set(x_39, 1, x_38); x_40 = l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_660____closed__1; -x_41 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_39, x_40); +x_41 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_39, x_40); x_42 = l_Lean_Json_mkObj(x_41); return x_42; } @@ -7394,7 +7394,7 @@ x_9 = lean_ctor_get(x_7, 1); lean_inc(x_9); lean_dec(x_7); x_10 = lean_array_size(x_8); -x_11 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_10, x_6, x_8); +x_11 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_10, x_6, x_8); x_12 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_12, 0, x_11); x_13 = lean_ctor_get(x_3, 1); @@ -7904,7 +7904,7 @@ lean_inc(x_7); lean_dec(x_6); x_8 = lean_array_size(x_7); x_9 = 0; -x_10 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_8, x_9, x_7); +x_10 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_8, x_9, x_7); x_11 = lean_ctor_get(x_10, 0); lean_inc(x_11); lean_dec(x_10); @@ -9188,7 +9188,7 @@ lean_inc(x_250); lean_dec(x_249); x_251 = lean_array_size(x_250); x_252 = 0; -x_253 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_251, x_252, x_250); +x_253 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_251, x_252, x_250); x_254 = lean_ctor_get(x_253, 0); lean_inc(x_254); lean_dec(x_253); @@ -10270,7 +10270,7 @@ x_27 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_27, 0, x_10); lean_ctor_set(x_27, 1, x_26); x_28 = l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_660____closed__1; -x_29 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_27, x_28); +x_29 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_27, x_28); x_30 = l_Lean_Json_mkObj(x_29); return x_30; } @@ -10317,7 +10317,7 @@ x_9 = lean_ctor_get(x_7, 1); lean_inc(x_9); lean_dec(x_7); x_10 = lean_array_size(x_8); -x_11 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_10, x_6, x_8); +x_11 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_10, x_6, x_8); x_12 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_12, 0, x_11); x_13 = lean_ctor_get(x_3, 1); @@ -10342,7 +10342,7 @@ lean_inc(x_23); lean_dec(x_21); x_24 = lean_ctor_get(x_1, 1); lean_inc(x_24); -x_25 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_24); +x_25 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_24); x_26 = lean_ctor_get(x_1, 2); lean_inc(x_26); lean_dec(x_1); @@ -10451,7 +10451,7 @@ lean_inc(x_7); lean_dec(x_6); x_8 = lean_array_size(x_7); x_9 = 0; -x_10 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_8, x_9, x_7); +x_10 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_8, x_9, x_7); x_11 = lean_ctor_get(x_10, 0); lean_inc(x_11); lean_dec(x_10); @@ -10583,7 +10583,7 @@ lean_inc(x_36); lean_dec(x_32); x_37 = lean_ctor_get(x_5, 3); lean_inc(x_37); -x_38 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794_(x_37); +x_38 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609_(x_37); x_39 = l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableInteractiveTermGoal_dec____x40_Lean_Widget_InteractiveGoal___hyg_2227____spec__1(x_38, x_2); if (lean_obj_tag(x_39) == 0) { @@ -10720,7 +10720,7 @@ lean_inc(x_66); lean_dec(x_65); x_67 = lean_array_size(x_66); x_68 = 0; -x_69 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_67, x_68, x_66); +x_69 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_67, x_68, x_66); x_70 = lean_ctor_get(x_69, 0); lean_inc(x_70); lean_dec(x_69); @@ -10856,7 +10856,7 @@ lean_inc(x_95); lean_dec(x_91); x_96 = lean_ctor_get(x_64, 3); lean_inc(x_96); -x_97 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794_(x_96); +x_97 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609_(x_96); x_98 = l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableInteractiveTermGoal_dec____x40_Lean_Widget_InteractiveGoal___hyg_2227____spec__1(x_97, x_2); if (lean_obj_tag(x_98) == 0) { @@ -12007,7 +12007,7 @@ x_6 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_6, 0, x_5); lean_ctor_set(x_6, 1, x_4); x_7 = l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_660____closed__1; -x_8 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_6, x_7); +x_8 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_6, x_7); x_9 = l_Lean_Json_mkObj(x_8); return x_9; } @@ -12076,7 +12076,7 @@ if (x_6 == 0) lean_object* x_7; size_t x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_7 = lean_ctor_get(x_5, 0); x_8 = lean_array_size(x_7); -x_9 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_8, x_4, x_7); +x_9 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_8, x_4, x_7); x_10 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_10, 0, x_9); x_11 = l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_3454_(x_10); @@ -12092,7 +12092,7 @@ lean_inc(x_13); lean_inc(x_12); lean_dec(x_5); x_14 = lean_array_size(x_12); -x_15 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_14, x_4, x_12); +x_15 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_14, x_4, x_12); x_16 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_16, 0, x_15); x_17 = l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_3454_(x_16); @@ -12192,7 +12192,7 @@ lean_inc(x_6); lean_dec(x_5); x_7 = lean_array_size(x_6); x_8 = 0; -x_9 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_7, x_8, x_6); +x_9 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_7, x_8, x_6); x_10 = lean_ctor_get(x_9, 0); lean_inc(x_10); lean_dec(x_9); @@ -12267,7 +12267,7 @@ lean_inc(x_26); lean_dec(x_25); x_27 = lean_array_size(x_26); x_28 = 0; -x_29 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_27, x_28, x_26); +x_29 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_27, x_28, x_26); x_30 = lean_ctor_get(x_29, 0); lean_inc(x_30); lean_dec(x_29); diff --git a/stage0/stdlib/Lean/Widget/UserWidget.c b/stage0/stdlib/Lean/Widget/UserWidget.c index 5d0c9e0cb0c4..e542b76a81c0 100644 --- a/stage0/stdlib/Lean/Widget/UserWidget.c +++ b/stage0/stdlib/Lean/Widget/UserWidget.c @@ -64,8 +64,10 @@ static lean_object* l_Lean_Widget_elabWidgetInstanceSpecAux___closed__18; lean_object* l_Lean_Json_mkObj(lean_object*); LEAN_EXPORT lean_object* l_Lean_ofExcept___at___private_Lean_Widget_UserWidget_0__Lean_Widget_evalUserWidgetDefinitionUnsafe___spec__1___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4203____closed__27; +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____lambda__4(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_ConstantInfo_type(lean_object*); +lean_object* l_List_flatMapTR_go___at___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____lambda__4___closed__9; static lean_object* l_Lean_Widget_showPanelWidgetsCmd___closed__3; extern lean_object* l_Lean_maxRecDepthErrorMessage; @@ -223,7 +225,6 @@ extern lean_object* l_Lean_Lsp_instFromJsonPosition; static lean_object* l_Lean_Widget_instToJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4587____closed__1; static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4203____closed__32; lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); -lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(size_t, size_t, lean_object*); static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonGetWidgetSourceParams____x40_Lean_Widget_UserWidget___hyg_872____closed__1; LEAN_EXPORT lean_object* l_Lean_Widget_instToModuleModule; LEAN_EXPORT lean_object* l_Lean_Widget_getWidgetSource___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -231,6 +232,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Widget_elabShowPanelWi LEAN_EXPORT lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4953_(lean_object*); static lean_object* l_Lean_Widget_showPanelWidgetsCmd___closed__7; size_t lean_usize_of_nat(lean_object*); +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____lambda__3___closed__3; static lean_object* l_Lean_Widget_instRpcEncodablePanelWidgetInstance___closed__1; LEAN_EXPORT lean_object* l_Lean_Widget_addPanelWidgetGlobal(lean_object*); @@ -242,6 +244,7 @@ static lean_object* l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40 static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____lambda__2___closed__3; LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Widget_elabShowPanelWidgetsCmd___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___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4203____closed__19; +lean_object* l_Lean_Json_opt___at___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_552____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_getWidgetSource___lambda__5(uint64_t, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_task_pure(lean_object*); lean_object* l_Lean_FileMap_lspPosToUtf8Pos(lean_object*, lean_object*); @@ -266,7 +269,6 @@ LEAN_EXPORT lean_object* l_Lean_Widget_addPanelWidgetLocal___rarg(lean_object*, lean_object* l_Lean_ResolveName_resolveNamespace(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_ins___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_182____spec__2___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_List_toArray___rarg(lean_object*); -lean_object* l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(lean_object*, lean_object*); lean_object* l_Lean_registerSimpleScopedEnvExtension___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Widget_elabWidgetInstanceSpecAux___closed__10; lean_object* l_Lean_SourceInfo_fromRef(lean_object*, uint8_t); @@ -346,7 +348,6 @@ lean_object* l_Lean_Name_append(lean_object*, lean_object*); static lean_object* l_Lean_Widget_addWidgetSpec___closed__7; LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Widget_elabShowPanelWidgetsCmd___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Widget_elabWidgetInstanceSpecAux___closed__58; -lean_object* l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4203____closed__12; static lean_object* l_Lean_Widget_instRpcEncodableGetWidgetsResponse_dec____x40_Lean_Widget_UserWidget___hyg_4842____closed__1; lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcCallParams____x40_Lean_Data_Lsp_Extra___hyg_2014____spec__2(lean_object*, lean_object*); @@ -489,12 +490,12 @@ LEAN_EXPORT lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJso lean_object* l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_instToModuleUserWidgetDefinition___boxed(lean_object*); static lean_object* l_Lean_Widget_widgetCmd___closed__4; +lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(size_t, size_t, lean_object*); static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Widget_elabShowPanelWidgetsCmd___spec__6___closed__4; static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_182____closed__2; static lean_object* l_Lean_Widget_showPanelWidgetsCmd___closed__8; lean_object* l_Lean_Server_instRpcEncodableOfFromJsonOfToJson___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_logWarning___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406_(lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Widget_elabShowPanelWidgetsCmd___spec__13___lambda__1___closed__1; static lean_object* l_Lean_Widget_showWidgetSpec___closed__5; static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Widget_elabShowPanelWidgetsCmd___spec__6___closed__1; @@ -597,7 +598,6 @@ static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonGe LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Widget_getWidgets___spec__5(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_Widget_UserWidget_0__Lean_Widget_evalUserWidgetDefinitionUnsafe___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodableGetWidgetsResponse; -lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_showWidgetSpec; LEAN_EXPORT lean_object* l_Lean_Widget_widgetInfosAt_x3f___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_registerBuiltinAttribute(lean_object*, lean_object*); @@ -655,6 +655,7 @@ static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRp static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____lambda__4___closed__6; static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____closed__1; static lean_object* l_Lean_Widget_widgetInstanceSpec___closed__6; +lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(size_t, size_t, lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Widget_elabShowPanelWidgetsCmd___spec__13___closed__2; LEAN_EXPORT lean_object* l_Lean_Widget_moduleRegistry; static lean_object* l_Lean_Widget_elabWidgetInstanceSpecAux___closed__9; @@ -677,9 +678,7 @@ static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240 LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Widget_getWidgets___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_Widget_showPanelWidgetsCmd___closed__10; static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_evalModuleUnsafe___closed__3; -lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(size_t, size_t, lean_object*); static lean_object* l_Lean_Widget_widgetInstanceSpec___closed__14; -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(lean_object*); static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonWidgetSource____x40_Lean_Widget_UserWidget___hyg_1038____closed__3; LEAN_EXPORT lean_object* l_Lean_Widget_evalPanelWidgets(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_getWidgetSource___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -704,6 +703,7 @@ static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRp static lean_object* l_Lean_Widget_instToJsonWidgetSource___closed__1; extern lean_object* l_Lean_Elab_unsupportedSyntaxExceptionId; uint8_t lean_usize_dec_lt(size_t, size_t); +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_addPanelWidgetLocal(lean_object*, lean_object*); static lean_object* l_Lean_Widget_elabWidgetInstanceSpecAux___closed__28; static lean_object* l_Lean_Widget_eraseWidgetSpec___closed__1; @@ -717,6 +717,7 @@ static lean_object* l_Lean_Widget_elabWidgetInstanceSpecAux___closed__36; LEAN_EXPORT lean_object* l_Lean_Widget_instToJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4988_; LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Widget_elabShowPanelWidgetsCmd___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_Widget_getWidgetSource___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221_(lean_object*); static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____closed__3; lean_object* l_Lean_Json_pretty(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Widget_elabShowPanelWidgetsCmd___spec__13___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -750,7 +751,6 @@ LEAN_EXPORT lean_object* l_Lean_Widget_instToJsonUserWidgetDefinition; lean_object* l_Lean_Server_Snapshots_Snapshot_endPos(lean_object*); static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____lambda__3___closed__2; LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_1426____spec__1___boxed(lean_object*, lean_object*); -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794_(lean_object*); static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_evalModuleUnsafe___closed__2; static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_evalModuleUnsafe___closed__4; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableGetWidgetsResponse_enc____x40_Lean_Widget_UserWidget___hyg_4842____spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -8186,7 +8186,7 @@ lean_ctor_set(x_8, 1, x_7); x_9 = lean_ctor_get(x_1, 0); lean_inc(x_9); lean_dec(x_1); -x_10 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406_(x_9); +x_10 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221_(x_9); x_11 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonGetWidgetSourceParams____x40_Lean_Widget_UserWidget___hyg_820____closed__2; x_12 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_12, 0, x_11); @@ -8201,7 +8201,7 @@ x_15 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_15, 0, x_8); lean_ctor_set(x_15, 1, x_14); x_16 = l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____lambda__4___closed__9; -x_17 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_15, x_16); +x_17 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____spec__1(x_15, x_16); x_18 = l_Lean_Json_mkObj(x_17); return x_18; } @@ -8411,7 +8411,7 @@ x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); x_13 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonGetWidgetSourceParams____x40_Lean_Widget_UserWidget___hyg_820____closed__2; -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____spec__1(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -8536,7 +8536,7 @@ x_7 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_7, 0, x_6); lean_ctor_set(x_7, 1, x_5); x_8 = l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____lambda__4___closed__9; -x_9 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_7, x_8); +x_9 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____spec__1(x_7, x_8); x_10 = l_Lean_Json_mkObj(x_9); return x_10; } @@ -18894,7 +18894,7 @@ x_14 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_14, 0, x_8); lean_ctor_set(x_14, 1, x_13); x_15 = l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____lambda__4___closed__9; -x_16 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_14, x_15); +x_16 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____spec__1(x_14, x_15); x_17 = l_Lean_Json_mkObj(x_16); return x_17; } @@ -18932,7 +18932,7 @@ x_30 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_30, 0, x_24); lean_ctor_set(x_30, 1, x_29); x_31 = l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____lambda__4___closed__9; -x_32 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_30, x_31); +x_32 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____spec__1(x_30, x_31); x_33 = l_Lean_Json_mkObj(x_32); return x_33; } @@ -19937,9 +19937,9 @@ x_16 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_16, 0, x_15); lean_ctor_set(x_16, 1, x_9); x_17 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4203____closed__25; -x_18 = l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(x_17, x_5); +x_18 = l_Lean_Json_opt___at___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_552____spec__1(x_17, x_5); x_19 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3680____closed__1; -x_20 = l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(x_19, x_6); +x_20 = l_Lean_Json_opt___at___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_552____spec__1(x_19, x_6); x_21 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_21, 0, x_20); lean_ctor_set(x_21, 1, x_9); @@ -19956,7 +19956,7 @@ x_25 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_25, 0, x_10); lean_ctor_set(x_25, 1, x_24); x_26 = l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____lambda__4___closed__9; -x_27 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_25, x_26); +x_27 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____spec__1(x_25, x_26); x_28 = l_Lean_Json_mkObj(x_27); return x_28; } @@ -20155,7 +20155,7 @@ if (x_46 == 0) lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; x_47 = lean_ctor_get(x_13, 0); x_48 = lean_ctor_get(x_14, 0); -x_49 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_48); +x_49 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_48); lean_ctor_set(x_14, 0, x_49); x_50 = lean_ctor_get(x_1, 2); lean_inc(x_50); @@ -20227,7 +20227,7 @@ x_64 = lean_ctor_get(x_13, 0); x_65 = lean_ctor_get(x_14, 0); lean_inc(x_65); lean_dec(x_14); -x_66 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_65); +x_66 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_65); x_67 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_67, 0, x_66); x_68 = lean_ctor_get(x_1, 2); @@ -20298,7 +20298,7 @@ if (lean_is_exclusive(x_14)) { lean_dec_ref(x_14); x_81 = lean_box(0); } -x_82 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_80); +x_82 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_80); if (lean_is_scalar(x_81)) { x_83 = lean_alloc_ctor(1, 1, 0); } else { @@ -20941,7 +20941,7 @@ if (x_62 == 0) { lean_object* x_63; lean_object* x_64; lean_object* x_65; x_63 = lean_ctor_get(x_60, 0); -x_64 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794_(x_63); +x_64 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609_(x_63); x_65 = l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableInteractiveTermGoal_dec____x40_Lean_Widget_InteractiveGoal___hyg_2227____spec__1(x_64, x_2); if (lean_obj_tag(x_65) == 0) { @@ -20985,7 +20985,7 @@ lean_object* x_70; lean_object* x_71; lean_object* x_72; x_70 = lean_ctor_get(x_60, 0); lean_inc(x_70); lean_dec(x_60); -x_71 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794_(x_70); +x_71 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609_(x_70); x_72 = l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableInteractiveTermGoal_dec____x40_Lean_Widget_InteractiveGoal___hyg_2227____spec__1(x_71, x_2); if (lean_obj_tag(x_72) == 0) { @@ -21271,7 +21271,7 @@ if (lean_is_exclusive(x_110)) { lean_dec_ref(x_110); x_113 = lean_box(0); } -x_114 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794_(x_112); +x_114 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609_(x_112); x_115 = l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableInteractiveTermGoal_dec____x40_Lean_Widget_InteractiveGoal___hyg_2227____spec__1(x_114, x_2); if (lean_obj_tag(x_115) == 0) { @@ -21627,7 +21627,7 @@ x_6 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_6, 0, x_5); lean_ctor_set(x_6, 1, x_4); x_7 = l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____lambda__4___closed__9; -x_8 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_6, x_7); +x_8 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____spec__1(x_6, x_7); x_9 = l_Lean_Json_mkObj(x_8); return x_9; } @@ -21696,7 +21696,7 @@ if (x_6 == 0) lean_object* x_7; size_t x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_7 = lean_ctor_get(x_5, 0); x_8 = lean_array_size(x_7); -x_9 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_8, x_4, x_7); +x_9 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_8, x_4, x_7); x_10 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_10, 0, x_9); x_11 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4953_(x_10); @@ -21712,7 +21712,7 @@ lean_inc(x_13); lean_inc(x_12); lean_dec(x_5); x_14 = lean_array_size(x_12); -x_15 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_14, x_4, x_12); +x_15 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_14, x_4, x_12); x_16 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_16, 0, x_15); x_17 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4953_(x_16); @@ -21817,7 +21817,7 @@ lean_inc(x_6); lean_dec(x_5); x_7 = lean_array_size(x_6); x_8 = 0; -x_9 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_7, x_8, x_6); +x_9 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_7, x_8, x_6); x_10 = lean_ctor_get(x_9, 0); lean_inc(x_10); lean_dec(x_9); @@ -21891,7 +21891,7 @@ lean_inc(x_26); lean_dec(x_25); x_27 = lean_array_size(x_26); x_28 = 0; -x_29 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_27, x_28, x_26); +x_29 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_27, x_28, x_26); x_30 = lean_ctor_get(x_29, 0); lean_inc(x_30); lean_dec(x_29); diff --git a/stage0/stdlib/Std.c b/stage0/stdlib/Std.c index c787a0a962bd..9cce1d917063 100644 --- a/stage0/stdlib/Std.c +++ b/stage0/stdlib/Std.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Std -// Imports: Std.Data Std.Sat Std.Sync Std.Time Std.Tactic Std.Internal +// Imports: Std.Data Std.Sat Std.Sync Std.Time Std.Tactic Std.Internal Std.Net #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -19,6 +19,7 @@ lean_object* initialize_Std_Sync(uint8_t builtin, lean_object*); lean_object* initialize_Std_Time(uint8_t builtin, lean_object*); lean_object* initialize_Std_Tactic(uint8_t builtin, lean_object*); lean_object* initialize_Std_Internal(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Net(uint8_t builtin, lean_object*); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Std(uint8_t builtin, lean_object* w) { lean_object * res; @@ -42,6 +43,9 @@ lean_dec_ref(res); res = initialize_Std_Internal(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +res = initialize_Std_Net(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/Std/Net.c b/stage0/stdlib/Std/Net.c new file mode 100644 index 000000000000..20b14b42c1ca --- /dev/null +++ b/stage0/stdlib/Std/Net.c @@ -0,0 +1,29 @@ +// Lean compiler output +// Module: Std.Net +// Imports: Std.Net.Addr +#include +#if defined(__clang__) +#pragma clang diagnostic ignored "-Wunused-parameter" +#pragma clang diagnostic ignored "-Wunused-label" +#elif defined(__GNUC__) && !defined(__CLANG__) +#pragma GCC diagnostic ignored "-Wunused-parameter" +#pragma GCC diagnostic ignored "-Wunused-label" +#pragma GCC diagnostic ignored "-Wunused-but-set-variable" +#endif +#ifdef __cplusplus +extern "C" { +#endif +lean_object* initialize_Std_Net_Addr(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Std_Net(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_Std_Net_Addr(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/Std/Net/Addr.c b/stage0/stdlib/Std/Net/Addr.c new file mode 100644 index 000000000000..15f3beed0173 --- /dev/null +++ b/stage0/stdlib/Std/Net/Addr.c @@ -0,0 +1,1256 @@ +// Lean compiler output +// Module: Std.Net.Addr +// Imports: Init.Data.Vector.Basic +#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_Std_Net_IPAddr_instToString; +static lean_object* l_Std_Net_AddressFamily_noConfusion___rarg___closed__1; +LEAN_EXPORT lean_object* l_Std_Net_IPv4Addr_toString___boxed(lean_object*); +LEAN_EXPORT uint8_t l_Std_Net_instDecidableEqIPv4Addr(lean_object*, lean_object*); +uint8_t lean_uint8_of_nat(lean_object*); +LEAN_EXPORT lean_object* l_Std_Net_instInhabitedSocketAddress; +LEAN_EXPORT lean_object* l_Std_Net_IPv6Addr_instToString; +LEAN_EXPORT lean_object* l_Std_Net_AddressFamily_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Std_Net_Addr_0__Std_Net_decEqSocketAddressV4____x40_Std_Net_Addr___hyg_150_(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Vector_Basic_0__decEqVector____x40_Init_Data_Vector_Basic___hyg_98____at___private_Std_Net_Addr_0__Std_Net_decEqIPv6Addr____x40_Std_Net_Addr___hyg_319____spec__1___rarg___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Net_SocketAddress_family___boxed(lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Net_Addr_0__Std_Net_decEqSocketAddressV4____x40_Std_Net_Addr___hyg_150____boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Net_IPv4Addr_ofString___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Net_IPv4Addr_instCoeIPAddr(lean_object*); +LEAN_EXPORT lean_object* l_Std_Net_SocketAddress_ipAddr(lean_object*); +lean_object* lean_mk_array(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Net_IPv6Addr_ofParts(uint16_t, uint16_t, uint16_t, uint16_t, uint16_t, uint16_t, uint16_t, uint16_t); +LEAN_EXPORT lean_object* l_Std_Net_AddressFamily_toCtorIdx(uint8_t); +LEAN_EXPORT uint8_t l___private_Std_Net_Addr_0__Std_Net_decEqSocketAddressV6____x40_Std_Net_Addr___hyg_440_(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Net_IPAddr_toString___boxed(lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Vector_Basic_0__decEqVector____x40_Init_Data_Vector_Basic___hyg_98____at___private_Std_Net_Addr_0__Std_Net_decEqIPv6Addr____x40_Std_Net_Addr___hyg_319____spec__1(lean_object*); +LEAN_EXPORT lean_object* l_Std_Net_instDecidableEqIPv6Addr___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Net_Addr_0__Std_Net_decEqIPAddr____x40_Std_Net_Addr___hyg_616____boxed(lean_object*, lean_object*); +lean_object* l_instDecidableEqUInt8___boxed(lean_object*, lean_object*); +uint16_t lean_uint16_of_nat(lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Net_Addr_0__Std_Net_decEqSocketAddressV6____x40_Std_Net_Addr___hyg_440____boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Net_Addr_0__Std_Net_decEqIPv4Addr____x40_Std_Net_Addr___hyg_29____boxed(lean_object*, lean_object*); +static lean_object* l___private_Init_Data_Vector_Basic_0__decEqVector____x40_Init_Data_Vector_Basic___hyg_98____at___private_Std_Net_Addr_0__Std_Net_decEqIPv6Addr____x40_Std_Net_Addr___hyg_319____spec__1___rarg___closed__1; +LEAN_EXPORT lean_object* l_Std_Net_instDecidableEqSocketAddress___boxed(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Std_Net_SocketAddress_family(lean_object*); +LEAN_EXPORT lean_object* l_Std_Net_instDecidableEqIPAddr___boxed(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Std_Net_Addr_0__Std_Net_decEqSocketAddress____x40_Std_Net_Addr___hyg_842_(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Std_Net_instDecidableEqIPAddr(lean_object*, lean_object*); +uint8_t lean_nat_dec_eq(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Std_Net_Addr_0__Std_Net_decEqIPv4Addr____x40_Std_Net_Addr___hyg_29_(lean_object*, lean_object*); +static lean_object* l_Std_Net_instInhabitedSocketAddressV6___closed__1; +LEAN_EXPORT lean_object* l_Std_Net_instDecidableEqSocketAddressV4___boxed(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Std_Net_instDecidableEqSocketAddress(lean_object*, lean_object*); +static uint16_t l_Std_Net_instInhabitedSocketAddressV4___closed__1; +LEAN_EXPORT lean_object* l_Std_Net_IPv6Addr_toString___boxed(lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Vector_Basic_0__decEqVector____x40_Init_Data_Vector_Basic___hyg_98____at___private_Std_Net_Addr_0__Std_Net_decEqIPv4Addr____x40_Std_Net_Addr___hyg_29____spec__1___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Net_IPAddr_toString(lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Net_Addr_0__Std_Net_decEqSocketAddress____x40_Std_Net_Addr___hyg_842____boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Net_IPv4Addr_ofParts___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Std_Net_Addr_0__Std_Net_decEqIPAddr____x40_Std_Net_Addr___hyg_616_(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Vector_Basic_0__decEqVector____x40_Init_Data_Vector_Basic___hyg_98____at___private_Std_Net_Addr_0__Std_Net_decEqIPv4Addr____x40_Std_Net_Addr___hyg_29____spec__1(lean_object*); +static uint8_t l_Std_Net_instInhabitedIPv4Addr___closed__1; +LEAN_EXPORT lean_object* l_Std_Net_instDecidableEqSocketAddressV6___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Net_instInhabitedIPv4Addr; +LEAN_EXPORT uint8_t l_Std_Net_IPAddr_family(lean_object*); +LEAN_EXPORT lean_object* l_Std_Net_instInhabitedIPv6Addr; +lean_object* lean_uv_pton_v4(lean_object*); +LEAN_EXPORT lean_object* l_Std_Net_IPv6Addr_ofString___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Net_AddressFamily_noConfusion___rarg(uint8_t, uint8_t, lean_object*); +static lean_object* l_Std_Net_instInhabitedSocketAddressV4___closed__2; +LEAN_EXPORT lean_object* l_Std_Net_IPAddr_family___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Net_instInhabitedIPAddr; +LEAN_EXPORT lean_object* l_Std_Net_SocketAddressV6_instCoeSocketAddress(lean_object*); +static lean_object* l_Std_Net_instInhabitedSocketAddress___closed__1; +uint8_t lean_uint16_dec_eq(uint16_t, uint16_t); +LEAN_EXPORT lean_object* l_Std_Net_SocketAddressV4_instCoeSocketAddress(lean_object*); +LEAN_EXPORT uint8_t l___private_Init_Data_Vector_Basic_0__decEqVector____x40_Init_Data_Vector_Basic___hyg_98____at___private_Std_Net_Addr_0__Std_Net_decEqIPv6Addr____x40_Std_Net_Addr___hyg_319____spec__1___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Vector_Basic_0__decEqVector____x40_Init_Data_Vector_Basic___hyg_98____at___private_Std_Net_Addr_0__Std_Net_decEqIPv4Addr____x40_Std_Net_Addr___hyg_29____spec__1___rarg___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Net_AddressFamily_noConfusion(lean_object*); +LEAN_EXPORT lean_object* l_Std_Net_instDecidableEqIPv4Addr___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Net_IPAddr_instToString___closed__1; +LEAN_EXPORT lean_object* l_Std_Net_instDecidableEqAddressFamily___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Net_AddressFamily_noConfusion___rarg___lambda__1(lean_object*); +LEAN_EXPORT lean_object* l_Std_Net_instInhabitedSocketAddressV6; +LEAN_EXPORT lean_object* l_Std_Net_instInhabitedSocketAddressV4; +LEAN_EXPORT uint8_t l_Std_Net_instDecidableEqSocketAddressV4(lean_object*, lean_object*); +uint8_t lean_nat_dec_eq(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Std_Net_instInhabitedAddressFamily; +LEAN_EXPORT lean_object* l_Std_Net_AddressFamily_noConfusion___rarg___lambda__1___boxed(lean_object*); +lean_object* l_instDecidableEqUInt16___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Net_IPv6Addr_instToString___closed__1; +LEAN_EXPORT lean_object* l_Std_Net_IPv6Addr_instCoeIPAddr(lean_object*); +LEAN_EXPORT lean_object* l_Std_Net_SocketAddress_port___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Net_IPv4Addr_ofParts(uint8_t, uint8_t, uint8_t, uint8_t); +uint8_t l_Array_instDecidableEq___rarg(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Net_IPv4Addr_instToString; +static lean_object* l_Std_Net_instInhabitedIPv4Addr___closed__2; +static lean_object* l_Std_Net_instInhabitedIPv6Addr___closed__1; +LEAN_EXPORT lean_object* l_Std_Net_IPv6Addr_ofParts___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Init_Data_Vector_Basic_0__decEqVector____x40_Init_Data_Vector_Basic___hyg_98____at___private_Std_Net_Addr_0__Std_Net_decEqIPv4Addr____x40_Std_Net_Addr___hyg_29____spec__1___rarg(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Std_Net_instDecidableEqAddressFamily(uint8_t, uint8_t); +lean_object* lean_array_mk(lean_object*); +LEAN_EXPORT uint8_t l_Std_Net_instDecidableEqIPv6Addr(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Std_Net_Addr_0__Std_Net_decEqIPv6Addr____x40_Std_Net_Addr___hyg_319_(lean_object*, lean_object*); +LEAN_EXPORT uint16_t l_Std_Net_SocketAddress_port(lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Vector_Basic_0__decEqVector____x40_Init_Data_Vector_Basic___hyg_98____at___private_Std_Net_Addr_0__Std_Net_decEqIPv6Addr____x40_Std_Net_Addr___hyg_319____spec__1___boxed(lean_object*); +lean_object* lean_uv_pton_v6(lean_object*); +LEAN_EXPORT uint8_t l_Std_Net_instDecidableEqSocketAddressV6(lean_object*, lean_object*); +static lean_object* l_Std_Net_instInhabitedIPAddr___closed__1; +static lean_object* l_Std_Net_IPv4Addr_instToString___closed__1; +lean_object* lean_uv_ntop_v6(lean_object*); +LEAN_EXPORT uint8_t l_Std_Net_AddressFamily_ofNat(lean_object*); +LEAN_EXPORT lean_object* l_Std_Net_AddressFamily_ofNat___boxed(lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Net_Addr_0__Std_Net_decEqIPv6Addr____x40_Std_Net_Addr___hyg_319____boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Net_AddressFamily_toCtorIdx___boxed(lean_object*); +static lean_object* l___private_Init_Data_Vector_Basic_0__decEqVector____x40_Init_Data_Vector_Basic___hyg_98____at___private_Std_Net_Addr_0__Std_Net_decEqIPv4Addr____x40_Std_Net_Addr___hyg_29____spec__1___rarg___closed__1; +lean_object* lean_uv_ntop_v4(lean_object*); +static uint8_t _init_l_Std_Net_instInhabitedIPv4Addr___closed__1() { +_start: +{ +lean_object* x_1; uint8_t x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = lean_uint8_of_nat(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Net_instInhabitedIPv4Addr___closed__2() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; +x_1 = lean_unsigned_to_nat(4u); +x_2 = l_Std_Net_instInhabitedIPv4Addr___closed__1; +x_3 = lean_box(x_2); +x_4 = lean_mk_array(x_1, x_3); +return x_4; +} +} +static lean_object* _init_l_Std_Net_instInhabitedIPv4Addr() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Net_instInhabitedIPv4Addr___closed__2; +return x_1; +} +} +static lean_object* _init_l___private_Init_Data_Vector_Basic_0__decEqVector____x40_Init_Data_Vector_Basic___hyg_98____at___private_Std_Net_Addr_0__Std_Net_decEqIPv4Addr____x40_Std_Net_Addr___hyg_29____spec__1___rarg___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_instDecidableEqUInt8___boxed), 2, 0); +return x_1; +} +} +LEAN_EXPORT uint8_t l___private_Init_Data_Vector_Basic_0__decEqVector____x40_Init_Data_Vector_Basic___hyg_98____at___private_Std_Net_Addr_0__Std_Net_decEqIPv4Addr____x40_Std_Net_Addr___hyg_29____spec__1___rarg(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; uint8_t x_4; +x_3 = l___private_Init_Data_Vector_Basic_0__decEqVector____x40_Init_Data_Vector_Basic___hyg_98____at___private_Std_Net_Addr_0__Std_Net_decEqIPv4Addr____x40_Std_Net_Addr___hyg_29____spec__1___rarg___closed__1; +x_4 = l_Array_instDecidableEq___rarg(x_3, x_1, x_2); +return x_4; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Vector_Basic_0__decEqVector____x40_Init_Data_Vector_Basic___hyg_98____at___private_Std_Net_Addr_0__Std_Net_decEqIPv4Addr____x40_Std_Net_Addr___hyg_29____spec__1(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l___private_Init_Data_Vector_Basic_0__decEqVector____x40_Init_Data_Vector_Basic___hyg_98____at___private_Std_Net_Addr_0__Std_Net_decEqIPv4Addr____x40_Std_Net_Addr___hyg_29____spec__1___rarg___boxed), 2, 0); +return x_2; +} +} +LEAN_EXPORT uint8_t l___private_Std_Net_Addr_0__Std_Net_decEqIPv4Addr____x40_Std_Net_Addr___hyg_29_(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; uint8_t x_4; +x_3 = lean_alloc_closure((void*)(l_instDecidableEqUInt8___boxed), 2, 0); +x_4 = l_Array_instDecidableEq___rarg(x_3, x_1, x_2); +return x_4; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Vector_Basic_0__decEqVector____x40_Init_Data_Vector_Basic___hyg_98____at___private_Std_Net_Addr_0__Std_Net_decEqIPv4Addr____x40_Std_Net_Addr___hyg_29____spec__1___rarg___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l___private_Init_Data_Vector_Basic_0__decEqVector____x40_Init_Data_Vector_Basic___hyg_98____at___private_Std_Net_Addr_0__Std_Net_decEqIPv4Addr____x40_Std_Net_Addr___hyg_29____spec__1___rarg(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_Vector_Basic_0__decEqVector____x40_Init_Data_Vector_Basic___hyg_98____at___private_Std_Net_Addr_0__Std_Net_decEqIPv4Addr____x40_Std_Net_Addr___hyg_29____spec__1___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l___private_Init_Data_Vector_Basic_0__decEqVector____x40_Init_Data_Vector_Basic___hyg_98____at___private_Std_Net_Addr_0__Std_Net_decEqIPv4Addr____x40_Std_Net_Addr___hyg_29____spec__1(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Std_Net_Addr_0__Std_Net_decEqIPv4Addr____x40_Std_Net_Addr___hyg_29____boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l___private_Std_Net_Addr_0__Std_Net_decEqIPv4Addr____x40_Std_Net_Addr___hyg_29_(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT uint8_t l_Std_Net_instDecidableEqIPv4Addr(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; uint8_t x_4; +x_3 = l___private_Init_Data_Vector_Basic_0__decEqVector____x40_Init_Data_Vector_Basic___hyg_98____at___private_Std_Net_Addr_0__Std_Net_decEqIPv4Addr____x40_Std_Net_Addr___hyg_29____spec__1___rarg___closed__1; +x_4 = l_Array_instDecidableEq___rarg(x_3, x_1, x_2); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Net_instDecidableEqIPv4Addr___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Std_Net_instDecidableEqIPv4Addr(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +static uint16_t _init_l_Std_Net_instInhabitedSocketAddressV4___closed__1() { +_start: +{ +lean_object* x_1; uint16_t x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = lean_uint16_of_nat(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Net_instInhabitedSocketAddressV4___closed__2() { +_start: +{ +lean_object* x_1; uint16_t x_2; lean_object* x_3; +x_1 = l_Std_Net_instInhabitedIPv4Addr___closed__2; +x_2 = l_Std_Net_instInhabitedSocketAddressV4___closed__1; +x_3 = lean_alloc_ctor(0, 1, 2); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint16(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Net_instInhabitedSocketAddressV4() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Net_instInhabitedSocketAddressV4___closed__2; +return x_1; +} +} +LEAN_EXPORT uint8_t l___private_Std_Net_Addr_0__Std_Net_decEqSocketAddressV4____x40_Std_Net_Addr___hyg_150_(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; uint16_t x_4; lean_object* x_5; uint16_t x_6; lean_object* x_7; uint8_t x_8; +x_3 = lean_ctor_get(x_1, 0); +x_4 = lean_ctor_get_uint16(x_1, sizeof(void*)*1); +x_5 = lean_ctor_get(x_2, 0); +x_6 = lean_ctor_get_uint16(x_2, sizeof(void*)*1); +x_7 = l___private_Init_Data_Vector_Basic_0__decEqVector____x40_Init_Data_Vector_Basic___hyg_98____at___private_Std_Net_Addr_0__Std_Net_decEqIPv4Addr____x40_Std_Net_Addr___hyg_29____spec__1___rarg___closed__1; +x_8 = l_Array_instDecidableEq___rarg(x_7, x_3, x_5); +if (x_8 == 0) +{ +uint8_t x_9; +x_9 = 0; +return x_9; +} +else +{ +uint8_t x_10; +x_10 = lean_uint16_dec_eq(x_4, x_6); +return x_10; +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Net_Addr_0__Std_Net_decEqSocketAddressV4____x40_Std_Net_Addr___hyg_150____boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l___private_Std_Net_Addr_0__Std_Net_decEqSocketAddressV4____x40_Std_Net_Addr___hyg_150_(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT uint8_t l_Std_Net_instDecidableEqSocketAddressV4(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = l___private_Std_Net_Addr_0__Std_Net_decEqSocketAddressV4____x40_Std_Net_Addr___hyg_150_(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Net_instDecidableEqSocketAddressV4___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Std_Net_instDecidableEqSocketAddressV4(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +static lean_object* _init_l_Std_Net_instInhabitedIPv6Addr___closed__1() { +_start: +{ +lean_object* x_1; uint16_t x_2; lean_object* x_3; lean_object* x_4; +x_1 = lean_unsigned_to_nat(8u); +x_2 = l_Std_Net_instInhabitedSocketAddressV4___closed__1; +x_3 = lean_box(x_2); +x_4 = lean_mk_array(x_1, x_3); +return x_4; +} +} +static lean_object* _init_l_Std_Net_instInhabitedIPv6Addr() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Net_instInhabitedIPv6Addr___closed__1; +return x_1; +} +} +static lean_object* _init_l___private_Init_Data_Vector_Basic_0__decEqVector____x40_Init_Data_Vector_Basic___hyg_98____at___private_Std_Net_Addr_0__Std_Net_decEqIPv6Addr____x40_Std_Net_Addr___hyg_319____spec__1___rarg___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_instDecidableEqUInt16___boxed), 2, 0); +return x_1; +} +} +LEAN_EXPORT uint8_t l___private_Init_Data_Vector_Basic_0__decEqVector____x40_Init_Data_Vector_Basic___hyg_98____at___private_Std_Net_Addr_0__Std_Net_decEqIPv6Addr____x40_Std_Net_Addr___hyg_319____spec__1___rarg(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; uint8_t x_4; +x_3 = l___private_Init_Data_Vector_Basic_0__decEqVector____x40_Init_Data_Vector_Basic___hyg_98____at___private_Std_Net_Addr_0__Std_Net_decEqIPv6Addr____x40_Std_Net_Addr___hyg_319____spec__1___rarg___closed__1; +x_4 = l_Array_instDecidableEq___rarg(x_3, x_1, x_2); +return x_4; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Vector_Basic_0__decEqVector____x40_Init_Data_Vector_Basic___hyg_98____at___private_Std_Net_Addr_0__Std_Net_decEqIPv6Addr____x40_Std_Net_Addr___hyg_319____spec__1(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l___private_Init_Data_Vector_Basic_0__decEqVector____x40_Init_Data_Vector_Basic___hyg_98____at___private_Std_Net_Addr_0__Std_Net_decEqIPv6Addr____x40_Std_Net_Addr___hyg_319____spec__1___rarg___boxed), 2, 0); +return x_2; +} +} +LEAN_EXPORT uint8_t l___private_Std_Net_Addr_0__Std_Net_decEqIPv6Addr____x40_Std_Net_Addr___hyg_319_(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; uint8_t x_4; +x_3 = lean_alloc_closure((void*)(l_instDecidableEqUInt16___boxed), 2, 0); +x_4 = l_Array_instDecidableEq___rarg(x_3, x_1, x_2); +return x_4; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Vector_Basic_0__decEqVector____x40_Init_Data_Vector_Basic___hyg_98____at___private_Std_Net_Addr_0__Std_Net_decEqIPv6Addr____x40_Std_Net_Addr___hyg_319____spec__1___rarg___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l___private_Init_Data_Vector_Basic_0__decEqVector____x40_Init_Data_Vector_Basic___hyg_98____at___private_Std_Net_Addr_0__Std_Net_decEqIPv6Addr____x40_Std_Net_Addr___hyg_319____spec__1___rarg(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_Vector_Basic_0__decEqVector____x40_Init_Data_Vector_Basic___hyg_98____at___private_Std_Net_Addr_0__Std_Net_decEqIPv6Addr____x40_Std_Net_Addr___hyg_319____spec__1___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l___private_Init_Data_Vector_Basic_0__decEqVector____x40_Init_Data_Vector_Basic___hyg_98____at___private_Std_Net_Addr_0__Std_Net_decEqIPv6Addr____x40_Std_Net_Addr___hyg_319____spec__1(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Std_Net_Addr_0__Std_Net_decEqIPv6Addr____x40_Std_Net_Addr___hyg_319____boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l___private_Std_Net_Addr_0__Std_Net_decEqIPv6Addr____x40_Std_Net_Addr___hyg_319_(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT uint8_t l_Std_Net_instDecidableEqIPv6Addr(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; uint8_t x_4; +x_3 = l___private_Init_Data_Vector_Basic_0__decEqVector____x40_Init_Data_Vector_Basic___hyg_98____at___private_Std_Net_Addr_0__Std_Net_decEqIPv6Addr____x40_Std_Net_Addr___hyg_319____spec__1___rarg___closed__1; +x_4 = l_Array_instDecidableEq___rarg(x_3, x_1, x_2); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Net_instDecidableEqIPv6Addr___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Std_Net_instDecidableEqIPv6Addr(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +static lean_object* _init_l_Std_Net_instInhabitedSocketAddressV6___closed__1() { +_start: +{ +lean_object* x_1; uint16_t x_2; lean_object* x_3; +x_1 = l_Std_Net_instInhabitedIPv6Addr___closed__1; +x_2 = l_Std_Net_instInhabitedSocketAddressV4___closed__1; +x_3 = lean_alloc_ctor(0, 1, 2); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint16(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Net_instInhabitedSocketAddressV6() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Net_instInhabitedSocketAddressV6___closed__1; +return x_1; +} +} +LEAN_EXPORT uint8_t l___private_Std_Net_Addr_0__Std_Net_decEqSocketAddressV6____x40_Std_Net_Addr___hyg_440_(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; uint16_t x_4; lean_object* x_5; uint16_t x_6; lean_object* x_7; uint8_t x_8; +x_3 = lean_ctor_get(x_1, 0); +x_4 = lean_ctor_get_uint16(x_1, sizeof(void*)*1); +x_5 = lean_ctor_get(x_2, 0); +x_6 = lean_ctor_get_uint16(x_2, sizeof(void*)*1); +x_7 = l___private_Init_Data_Vector_Basic_0__decEqVector____x40_Init_Data_Vector_Basic___hyg_98____at___private_Std_Net_Addr_0__Std_Net_decEqIPv6Addr____x40_Std_Net_Addr___hyg_319____spec__1___rarg___closed__1; +x_8 = l_Array_instDecidableEq___rarg(x_7, x_3, x_5); +if (x_8 == 0) +{ +uint8_t x_9; +x_9 = 0; +return x_9; +} +else +{ +uint8_t x_10; +x_10 = lean_uint16_dec_eq(x_4, x_6); +return x_10; +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Net_Addr_0__Std_Net_decEqSocketAddressV6____x40_Std_Net_Addr___hyg_440____boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l___private_Std_Net_Addr_0__Std_Net_decEqSocketAddressV6____x40_Std_Net_Addr___hyg_440_(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT uint8_t l_Std_Net_instDecidableEqSocketAddressV6(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = l___private_Std_Net_Addr_0__Std_Net_decEqSocketAddressV6____x40_Std_Net_Addr___hyg_440_(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Net_instDecidableEqSocketAddressV6___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Std_Net_instDecidableEqSocketAddressV6(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +static lean_object* _init_l_Std_Net_instInhabitedIPAddr___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Net_instInhabitedIPv4Addr___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_Std_Net_instInhabitedIPAddr() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Net_instInhabitedIPAddr___closed__1; +return x_1; +} +} +LEAN_EXPORT uint8_t l___private_Std_Net_Addr_0__Std_Net_decEqIPAddr____x40_Std_Net_Addr___hyg_616_(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_3 = lean_ctor_get(x_1, 0); +x_4 = lean_ctor_get(x_2, 0); +x_5 = l___private_Init_Data_Vector_Basic_0__decEqVector____x40_Init_Data_Vector_Basic___hyg_98____at___private_Std_Net_Addr_0__Std_Net_decEqIPv4Addr____x40_Std_Net_Addr___hyg_29____spec__1___rarg___closed__1; +x_6 = l_Array_instDecidableEq___rarg(x_5, x_3, x_4); +return x_6; +} +else +{ +uint8_t x_7; +x_7 = 0; +return x_7; +} +} +else +{ +if (lean_obj_tag(x_2) == 0) +{ +uint8_t x_8; +x_8 = 0; +return x_8; +} +else +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; +x_9 = lean_ctor_get(x_1, 0); +x_10 = lean_ctor_get(x_2, 0); +x_11 = l___private_Init_Data_Vector_Basic_0__decEqVector____x40_Init_Data_Vector_Basic___hyg_98____at___private_Std_Net_Addr_0__Std_Net_decEqIPv6Addr____x40_Std_Net_Addr___hyg_319____spec__1___rarg___closed__1; +x_12 = l_Array_instDecidableEq___rarg(x_11, x_9, x_10); +return x_12; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Net_Addr_0__Std_Net_decEqIPAddr____x40_Std_Net_Addr___hyg_616____boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l___private_Std_Net_Addr_0__Std_Net_decEqIPAddr____x40_Std_Net_Addr___hyg_616_(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT uint8_t l_Std_Net_instDecidableEqIPAddr(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = l___private_Std_Net_Addr_0__Std_Net_decEqIPAddr____x40_Std_Net_Addr___hyg_616_(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Net_instDecidableEqIPAddr___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Std_Net_instDecidableEqIPAddr(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +static lean_object* _init_l_Std_Net_instInhabitedSocketAddress___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Net_instInhabitedSocketAddressV4___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_Std_Net_instInhabitedSocketAddress() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Net_instInhabitedSocketAddress___closed__1; +return x_1; +} +} +LEAN_EXPORT uint8_t l___private_Std_Net_Addr_0__Std_Net_decEqSocketAddress____x40_Std_Net_Addr___hyg_842_(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_3; lean_object* x_4; uint8_t x_5; +x_3 = lean_ctor_get(x_1, 0); +x_4 = lean_ctor_get(x_2, 0); +x_5 = l___private_Std_Net_Addr_0__Std_Net_decEqSocketAddressV4____x40_Std_Net_Addr___hyg_150_(x_3, x_4); +return x_5; +} +else +{ +uint8_t x_6; +x_6 = 0; +return x_6; +} +} +else +{ +if (lean_obj_tag(x_2) == 0) +{ +uint8_t x_7; +x_7 = 0; +return x_7; +} +else +{ +lean_object* x_8; lean_object* x_9; uint8_t x_10; +x_8 = lean_ctor_get(x_1, 0); +x_9 = lean_ctor_get(x_2, 0); +x_10 = l___private_Std_Net_Addr_0__Std_Net_decEqSocketAddressV6____x40_Std_Net_Addr___hyg_440_(x_8, x_9); +return x_10; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Net_Addr_0__Std_Net_decEqSocketAddress____x40_Std_Net_Addr___hyg_842____boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l___private_Std_Net_Addr_0__Std_Net_decEqSocketAddress____x40_Std_Net_Addr___hyg_842_(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT uint8_t l_Std_Net_instDecidableEqSocketAddress(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = l___private_Std_Net_Addr_0__Std_Net_decEqSocketAddress____x40_Std_Net_Addr___hyg_842_(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Net_instDecidableEqSocketAddress___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Std_Net_instDecidableEqSocketAddress(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_Std_Net_AddressFamily_toCtorIdx(uint8_t x_1) { +_start: +{ +if (x_1 == 0) +{ +lean_object* x_2; +x_2 = lean_unsigned_to_nat(0u); +return x_2; +} +else +{ +lean_object* x_3; +x_3 = lean_unsigned_to_nat(1u); +return x_3; +} +} +} +LEAN_EXPORT lean_object* l_Std_Net_AddressFamily_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_Std_Net_AddressFamily_toCtorIdx(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Net_AddressFamily_noConfusion___rarg___lambda__1(lean_object* x_1) { +_start: +{ +lean_inc(x_1); +return x_1; +} +} +static lean_object* _init_l_Std_Net_AddressFamily_noConfusion___rarg___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Net_AddressFamily_noConfusion___rarg___lambda__1___boxed), 1, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Net_AddressFamily_noConfusion___rarg(uint8_t x_1, uint8_t x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Net_AddressFamily_noConfusion___rarg___closed__1; +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Net_AddressFamily_noConfusion(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Net_AddressFamily_noConfusion___rarg___boxed), 3, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Net_AddressFamily_noConfusion___rarg___lambda__1___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Net_AddressFamily_noConfusion___rarg___lambda__1(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Net_AddressFamily_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_Std_Net_AddressFamily_noConfusion___rarg(x_4, x_5, x_3); +return x_6; +} +} +static uint8_t _init_l_Std_Net_instInhabitedAddressFamily() { +_start: +{ +uint8_t x_1; +x_1 = 0; +return x_1; +} +} +LEAN_EXPORT uint8_t l_Std_Net_AddressFamily_ofNat(lean_object* x_1) { +_start: +{ +lean_object* x_2; uint8_t x_3; +x_2 = lean_unsigned_to_nat(0u); +x_3 = lean_nat_dec_eq(x_1, x_2); +if (x_3 == 0) +{ +uint8_t x_4; +x_4 = 1; +return x_4; +} +else +{ +uint8_t x_5; +x_5 = 0; +return x_5; +} +} +} +LEAN_EXPORT lean_object* l_Std_Net_AddressFamily_ofNat___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = l_Std_Net_AddressFamily_ofNat(x_1); +lean_dec(x_1); +x_3 = lean_box(x_2); +return x_3; +} +} +LEAN_EXPORT uint8_t l_Std_Net_instDecidableEqAddressFamily(uint8_t x_1, uint8_t x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; uint8_t x_5; +x_3 = l_Std_Net_AddressFamily_toCtorIdx(x_1); +x_4 = l_Std_Net_AddressFamily_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_Std_Net_instDecidableEqAddressFamily___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_Std_Net_instDecidableEqAddressFamily(x_3, x_4); +x_6 = lean_box(x_5); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Std_Net_IPv4Addr_ofParts(uint8_t x_1, uint8_t x_2, uint8_t x_3, uint8_t x_4) { +_start: +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_5 = lean_box(0); +x_6 = lean_box(x_4); +x_7 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_7, 0, x_6); +lean_ctor_set(x_7, 1, x_5); +x_8 = lean_box(x_3); +x_9 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_9, 0, x_8); +lean_ctor_set(x_9, 1, x_7); +x_10 = lean_box(x_2); +x_11 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_11, 0, x_10); +lean_ctor_set(x_11, 1, x_9); +x_12 = lean_box(x_1); +x_13 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_11); +x_14 = lean_array_mk(x_13); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Std_Net_IPv4Addr_ofParts___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +uint8_t x_5; uint8_t x_6; uint8_t x_7; uint8_t x_8; lean_object* x_9; +x_5 = lean_unbox(x_1); +lean_dec(x_1); +x_6 = lean_unbox(x_2); +lean_dec(x_2); +x_7 = lean_unbox(x_3); +lean_dec(x_3); +x_8 = lean_unbox(x_4); +lean_dec(x_4); +x_9 = l_Std_Net_IPv4Addr_ofParts(x_5, x_6, x_7, x_8); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Std_Net_IPv4Addr_ofString___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_uv_pton_v4(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Net_IPv4Addr_toString___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_uv_ntop_v4(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Net_IPv4Addr_instToString___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Net_IPv4Addr_toString___boxed), 1, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Net_IPv4Addr_instToString() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Net_IPv4Addr_instToString___closed__1; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Net_IPv4Addr_instCoeIPAddr(lean_object* x_1) { +_start: +{ +lean_object* x_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_Std_Net_SocketAddressV4_instCoeSocketAddress(lean_object* x_1) { +_start: +{ +lean_object* x_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_Std_Net_IPv6Addr_ofParts(uint16_t x_1, uint16_t x_2, uint16_t x_3, uint16_t x_4, uint16_t x_5, uint16_t x_6, uint16_t x_7, uint16_t x_8) { +_start: +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_9 = lean_box(0); +x_10 = lean_box(x_8); +x_11 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_11, 0, x_10); +lean_ctor_set(x_11, 1, x_9); +x_12 = lean_box(x_7); +x_13 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_11); +x_14 = lean_box(x_6); +x_15 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_15, 0, x_14); +lean_ctor_set(x_15, 1, x_13); +x_16 = lean_box(x_5); +x_17 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_17, 0, x_16); +lean_ctor_set(x_17, 1, x_15); +x_18 = lean_box(x_4); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_17); +x_20 = lean_box(x_3); +x_21 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_21, 0, x_20); +lean_ctor_set(x_21, 1, x_19); +x_22 = lean_box(x_2); +x_23 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_21); +x_24 = lean_box(x_1); +x_25 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_23); +x_26 = lean_array_mk(x_25); +return x_26; +} +} +LEAN_EXPORT lean_object* l_Std_Net_IPv6Addr_ofParts___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +uint16_t x_9; uint16_t x_10; uint16_t x_11; uint16_t x_12; uint16_t x_13; uint16_t x_14; uint16_t x_15; uint16_t x_16; lean_object* x_17; +x_9 = lean_unbox(x_1); +lean_dec(x_1); +x_10 = lean_unbox(x_2); +lean_dec(x_2); +x_11 = lean_unbox(x_3); +lean_dec(x_3); +x_12 = lean_unbox(x_4); +lean_dec(x_4); +x_13 = lean_unbox(x_5); +lean_dec(x_5); +x_14 = lean_unbox(x_6); +lean_dec(x_6); +x_15 = lean_unbox(x_7); +lean_dec(x_7); +x_16 = lean_unbox(x_8); +lean_dec(x_8); +x_17 = l_Std_Net_IPv6Addr_ofParts(x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +return x_17; +} +} +LEAN_EXPORT lean_object* l_Std_Net_IPv6Addr_ofString___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_uv_pton_v6(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Net_IPv6Addr_toString___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_uv_ntop_v6(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Net_IPv6Addr_instToString___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Net_IPv6Addr_toString___boxed), 1, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Net_IPv6Addr_instToString() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Net_IPv6Addr_instToString___closed__1; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Net_IPv6Addr_instCoeIPAddr(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Net_SocketAddressV6_instCoeSocketAddress(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT uint8_t l_Std_Net_IPAddr_family(lean_object* x_1) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +uint8_t x_2; +x_2 = 0; +return x_2; +} +else +{ +uint8_t x_3; +x_3 = 1; +return x_3; +} +} +} +LEAN_EXPORT lean_object* l_Std_Net_IPAddr_family___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = l_Std_Net_IPAddr_family(x_1); +lean_dec(x_1); +x_3 = lean_box(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Net_IPAddr_toString(lean_object* x_1) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_2; lean_object* x_3; +x_2 = lean_ctor_get(x_1, 0); +x_3 = lean_uv_ntop_v4(x_2); +return x_3; +} +else +{ +lean_object* x_4; lean_object* x_5; +x_4 = lean_ctor_get(x_1, 0); +x_5 = lean_uv_ntop_v6(x_4); +return x_5; +} +} +} +LEAN_EXPORT lean_object* l_Std_Net_IPAddr_toString___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Net_IPAddr_toString(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Net_IPAddr_instToString___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Net_IPAddr_toString___boxed), 1, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Net_IPAddr_instToString() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Net_IPAddr_instToString___closed__1; +return x_1; +} +} +LEAN_EXPORT uint8_t l_Std_Net_SocketAddress_family(lean_object* x_1) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +uint8_t x_2; +x_2 = 0; +return x_2; +} +else +{ +uint8_t x_3; +x_3 = 1; +return x_3; +} +} +} +LEAN_EXPORT lean_object* l_Std_Net_SocketAddress_family___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = l_Std_Net_SocketAddress_family(x_1); +lean_dec(x_1); +x_3 = lean_box(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Net_SocketAddress_ipAddr(lean_object* x_1) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +uint8_t x_2; +x_2 = !lean_is_exclusive(x_1); +if (x_2 == 0) +{ +lean_object* x_3; lean_object* x_4; +x_3 = lean_ctor_get(x_1, 0); +x_4 = lean_ctor_get(x_3, 0); +lean_inc(x_4); +lean_dec(x_3); +lean_ctor_set(x_1, 0, x_4); +return x_1; +} +else +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_5 = lean_ctor_get(x_1, 0); +lean_inc(x_5); +lean_dec(x_1); +x_6 = lean_ctor_get(x_5, 0); +lean_inc(x_6); +lean_dec(x_5); +x_7 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_7, 0, x_6); +return x_7; +} +} +else +{ +uint8_t x_8; +x_8 = !lean_is_exclusive(x_1); +if (x_8 == 0) +{ +lean_object* x_9; lean_object* x_10; +x_9 = lean_ctor_get(x_1, 0); +x_10 = lean_ctor_get(x_9, 0); +lean_inc(x_10); +lean_dec(x_9); +lean_ctor_set(x_1, 0, x_10); +return x_1; +} +else +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_11 = lean_ctor_get(x_1, 0); +lean_inc(x_11); +lean_dec(x_1); +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +lean_dec(x_11); +x_13 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_13, 0, x_12); +return x_13; +} +} +} +} +LEAN_EXPORT uint16_t l_Std_Net_SocketAddress_port(lean_object* x_1) { +_start: +{ +lean_object* x_2; uint16_t x_3; +x_2 = lean_ctor_get(x_1, 0); +x_3 = lean_ctor_get_uint16(x_2, sizeof(void*)*1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Net_SocketAddress_port___boxed(lean_object* x_1) { +_start: +{ +uint16_t x_2; lean_object* x_3; +x_2 = l_Std_Net_SocketAddress_port(x_1); +lean_dec(x_1); +x_3 = lean_box(x_2); +return x_3; +} +} +lean_object* initialize_Init_Data_Vector_Basic(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Std_Net_Addr(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_Vector_Basic(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l_Std_Net_instInhabitedIPv4Addr___closed__1 = _init_l_Std_Net_instInhabitedIPv4Addr___closed__1(); +l_Std_Net_instInhabitedIPv4Addr___closed__2 = _init_l_Std_Net_instInhabitedIPv4Addr___closed__2(); +lean_mark_persistent(l_Std_Net_instInhabitedIPv4Addr___closed__2); +l_Std_Net_instInhabitedIPv4Addr = _init_l_Std_Net_instInhabitedIPv4Addr(); +lean_mark_persistent(l_Std_Net_instInhabitedIPv4Addr); +l___private_Init_Data_Vector_Basic_0__decEqVector____x40_Init_Data_Vector_Basic___hyg_98____at___private_Std_Net_Addr_0__Std_Net_decEqIPv4Addr____x40_Std_Net_Addr___hyg_29____spec__1___rarg___closed__1 = _init_l___private_Init_Data_Vector_Basic_0__decEqVector____x40_Init_Data_Vector_Basic___hyg_98____at___private_Std_Net_Addr_0__Std_Net_decEqIPv4Addr____x40_Std_Net_Addr___hyg_29____spec__1___rarg___closed__1(); +lean_mark_persistent(l___private_Init_Data_Vector_Basic_0__decEqVector____x40_Init_Data_Vector_Basic___hyg_98____at___private_Std_Net_Addr_0__Std_Net_decEqIPv4Addr____x40_Std_Net_Addr___hyg_29____spec__1___rarg___closed__1); +l_Std_Net_instInhabitedSocketAddressV4___closed__1 = _init_l_Std_Net_instInhabitedSocketAddressV4___closed__1(); +l_Std_Net_instInhabitedSocketAddressV4___closed__2 = _init_l_Std_Net_instInhabitedSocketAddressV4___closed__2(); +lean_mark_persistent(l_Std_Net_instInhabitedSocketAddressV4___closed__2); +l_Std_Net_instInhabitedSocketAddressV4 = _init_l_Std_Net_instInhabitedSocketAddressV4(); +lean_mark_persistent(l_Std_Net_instInhabitedSocketAddressV4); +l_Std_Net_instInhabitedIPv6Addr___closed__1 = _init_l_Std_Net_instInhabitedIPv6Addr___closed__1(); +lean_mark_persistent(l_Std_Net_instInhabitedIPv6Addr___closed__1); +l_Std_Net_instInhabitedIPv6Addr = _init_l_Std_Net_instInhabitedIPv6Addr(); +lean_mark_persistent(l_Std_Net_instInhabitedIPv6Addr); +l___private_Init_Data_Vector_Basic_0__decEqVector____x40_Init_Data_Vector_Basic___hyg_98____at___private_Std_Net_Addr_0__Std_Net_decEqIPv6Addr____x40_Std_Net_Addr___hyg_319____spec__1___rarg___closed__1 = _init_l___private_Init_Data_Vector_Basic_0__decEqVector____x40_Init_Data_Vector_Basic___hyg_98____at___private_Std_Net_Addr_0__Std_Net_decEqIPv6Addr____x40_Std_Net_Addr___hyg_319____spec__1___rarg___closed__1(); +lean_mark_persistent(l___private_Init_Data_Vector_Basic_0__decEqVector____x40_Init_Data_Vector_Basic___hyg_98____at___private_Std_Net_Addr_0__Std_Net_decEqIPv6Addr____x40_Std_Net_Addr___hyg_319____spec__1___rarg___closed__1); +l_Std_Net_instInhabitedSocketAddressV6___closed__1 = _init_l_Std_Net_instInhabitedSocketAddressV6___closed__1(); +lean_mark_persistent(l_Std_Net_instInhabitedSocketAddressV6___closed__1); +l_Std_Net_instInhabitedSocketAddressV6 = _init_l_Std_Net_instInhabitedSocketAddressV6(); +lean_mark_persistent(l_Std_Net_instInhabitedSocketAddressV6); +l_Std_Net_instInhabitedIPAddr___closed__1 = _init_l_Std_Net_instInhabitedIPAddr___closed__1(); +lean_mark_persistent(l_Std_Net_instInhabitedIPAddr___closed__1); +l_Std_Net_instInhabitedIPAddr = _init_l_Std_Net_instInhabitedIPAddr(); +lean_mark_persistent(l_Std_Net_instInhabitedIPAddr); +l_Std_Net_instInhabitedSocketAddress___closed__1 = _init_l_Std_Net_instInhabitedSocketAddress___closed__1(); +lean_mark_persistent(l_Std_Net_instInhabitedSocketAddress___closed__1); +l_Std_Net_instInhabitedSocketAddress = _init_l_Std_Net_instInhabitedSocketAddress(); +lean_mark_persistent(l_Std_Net_instInhabitedSocketAddress); +l_Std_Net_AddressFamily_noConfusion___rarg___closed__1 = _init_l_Std_Net_AddressFamily_noConfusion___rarg___closed__1(); +lean_mark_persistent(l_Std_Net_AddressFamily_noConfusion___rarg___closed__1); +l_Std_Net_instInhabitedAddressFamily = _init_l_Std_Net_instInhabitedAddressFamily(); +l_Std_Net_IPv4Addr_instToString___closed__1 = _init_l_Std_Net_IPv4Addr_instToString___closed__1(); +lean_mark_persistent(l_Std_Net_IPv4Addr_instToString___closed__1); +l_Std_Net_IPv4Addr_instToString = _init_l_Std_Net_IPv4Addr_instToString(); +lean_mark_persistent(l_Std_Net_IPv4Addr_instToString); +l_Std_Net_IPv6Addr_instToString___closed__1 = _init_l_Std_Net_IPv6Addr_instToString___closed__1(); +lean_mark_persistent(l_Std_Net_IPv6Addr_instToString___closed__1); +l_Std_Net_IPv6Addr_instToString = _init_l_Std_Net_IPv6Addr_instToString(); +lean_mark_persistent(l_Std_Net_IPv6Addr_instToString); +l_Std_Net_IPAddr_instToString___closed__1 = _init_l_Std_Net_IPAddr_instToString___closed__1(); +lean_mark_persistent(l_Std_Net_IPAddr_instToString___closed__1); +l_Std_Net_IPAddr_instToString = _init_l_Std_Net_IPAddr_instToString(); +lean_mark_persistent(l_Std_Net_IPAddr_instToString); +return lean_io_result_mk_ok(lean_box(0)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Std/Sat/AIG/Basic.c b/stage0/stdlib/Std/Sat/AIG/Basic.c index 62ec80160257..b6c1749bf3be 100644 --- a/stage0/stdlib/Std/Sat/AIG/Basic.c +++ b/stage0/stdlib/Std/Sat/AIG/Basic.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Std.Sat.AIG.Basic -// Imports: Std.Data.HashMap Std.Data.HashSet +// Imports: Std.Data.HashSet #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -6203,16 +6203,12 @@ lean_dec(x_1); return x_6; } } -lean_object* initialize_Std_Data_HashMap(uint8_t builtin, lean_object*); lean_object* initialize_Std_Data_HashSet(uint8_t builtin, lean_object*); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Std_Sat_AIG_Basic(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_Std_Data_HashMap(builtin, lean_io_mk_world()); -if (lean_io_result_is_error(res)) return res; -lean_dec_ref(res); res = initialize_Std_Data_HashSet(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); diff --git a/stage0/stdlib/Std/Tactic/BVDecide/LRAT/Internal/Formula/Implementation.c b/stage0/stdlib/Std/Tactic/BVDecide/LRAT/Internal/Formula/Implementation.c index 12db8fc0a46b..78bcbf2b01c2 100644 --- a/stage0/stdlib/Std/Tactic/BVDecide/LRAT/Internal/Formula/Implementation.c +++ b/stage0/stdlib/Std/Tactic/BVDecide/LRAT/Internal/Formula/Implementation.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Std.Tactic.BVDecide.LRAT.Internal.Formula.Implementation -// Imports: Init.Data.Array Std.Tactic.BVDecide.LRAT.Internal.Formula.Class Std.Tactic.BVDecide.LRAT.Internal.Assignment Std.Sat.CNF.Basic +// Imports: Std.Tactic.BVDecide.LRAT.Internal.Formula.Class Std.Tactic.BVDecide.LRAT.Internal.Assignment Std.Sat.CNF.Basic #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -5748,7 +5748,6 @@ lean_dec(x_1); return x_3; } } -lean_object* initialize_Init_Data_Array(uint8_t builtin, lean_object*); lean_object* initialize_Std_Tactic_BVDecide_LRAT_Internal_Formula_Class(uint8_t builtin, lean_object*); lean_object* initialize_Std_Tactic_BVDecide_LRAT_Internal_Assignment(uint8_t builtin, lean_object*); lean_object* initialize_Std_Sat_CNF_Basic(uint8_t builtin, lean_object*); @@ -5757,9 +5756,6 @@ LEAN_EXPORT lean_object* initialize_Std_Tactic_BVDecide_LRAT_Internal_Formula_Im lean_object * res; if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); _G_initialized = true; -res = initialize_Init_Data_Array(builtin, lean_io_mk_world()); -if (lean_io_result_is_error(res)) return res; -lean_dec_ref(res); res = initialize_Std_Tactic_BVDecide_LRAT_Internal_Formula_Class(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); diff --git a/stage0/stdlib/Std/Time/Date/Unit/Month.c b/stage0/stdlib/Std/Time/Date/Unit/Month.c index 1e368a236137..1238efeca7ad 100644 --- a/stage0/stdlib/Std/Time/Date/Unit/Month.c +++ b/stage0/stdlib/Std/Time/Date/Unit/Month.c @@ -17,8 +17,10 @@ static lean_object* l_Std_Time_Month_Ordinal_toSeconds___closed__1; static lean_object* l_Std_Time_Month_instInhabitedOrdinal___closed__4; static lean_object* l_Std_Time_Month_instOffsetRepr___closed__1; static lean_object* l_Std_Time_Month_Ordinal_days___closed__6; +LEAN_EXPORT uint8_t l_Std_Time_Month_instDecidableLtOffset(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap; static lean_object* l_Std_Time_Month_Quarter_ofMonth___closed__1; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__25; static lean_object* l_Std_Time_Month_Ordinal_december___closed__2; LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_cumulativeDays___boxed(lean_object*, lean_object*); static lean_object* l_Std_Time_Month_Ordinal_october___closed__2; @@ -26,6 +28,7 @@ LEAN_EXPORT lean_object* l_Std_Time_Month_instOrdinalLE; static lean_object* l_Std_Time_Month_instInhabitedOrdinal___closed__3; lean_object* l_instBEqOfDecidableEq___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Time_Month_Ordinal_march___closed__1; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__11; LEAN_EXPORT lean_object* l_Std_Time_Month_Quarter_ofMonth___boxed(lean_object*); static lean_object* l_Std_Time_Month_Ordinal_toSeconds___closed__9; LEAN_EXPORT lean_object* l_Std_Time_Month_instOffsetSub; @@ -35,10 +38,11 @@ lean_object* lean_mk_empty_array_with_capacity(lean_object*); static lean_object* l_Std_Time_Month_Ordinal_february___closed__5; LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_toSeconds(uint8_t, lean_object*); extern lean_object* l_Int_instDiv; -static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__9; +LEAN_EXPORT lean_object* l_Std_Time_Month_instDecidableLeOffset___boxed(lean_object*, lean_object*); static lean_object* l_Std_Time_Month_Ordinal_november___closed__5; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__22; static lean_object* l_Std_Time_Month_Ordinal_april___closed__5; -static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__20; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__7; static lean_object* l_Std_Time_Month_Ordinal_toSeconds___closed__23; static lean_object* l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__11; extern lean_object* l_Int_instInhabited; @@ -47,9 +51,12 @@ LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_toOffset___boxed(lean_object*) static lean_object* l_Std_Time_Month_instInhabitedOrdinal___closed__1; LEAN_EXPORT lean_object* l_Std_Time_Month_instOrdinalRepr; static lean_object* l_Std_Time_Month_Ordinal_february___closed__4; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__23; static lean_object* l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__6; static lean_object* l_Std_Time_Month_Ordinal_march___closed__4; +LEAN_EXPORT lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_256_; LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_days(uint8_t, lean_object*); +static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__4; static lean_object* l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__29; static lean_object* l_Std_Time_Month_Ordinal_august___closed__5; static lean_object* l_Std_Time_Month_Ordinal_september___closed__1; @@ -59,17 +66,17 @@ static lean_object* l_Std_Time_Month_Ordinal_april___closed__3; lean_object* lean_array_push(lean_object*, lean_object*); static lean_object* l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__16; LEAN_EXPORT lean_object* l_Std_Time_Month_Offset_ofInt___boxed(lean_object*); -static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__7; static lean_object* l_Std_Time_Month_Ordinal_october___closed__4; static lean_object* l_Std_Time_Month_Ordinal_may___closed__2; static lean_object* l_Std_Time_Month_Ordinal_august___closed__1; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__12; lean_object* lean_array_fget(lean_object*, lean_object*); lean_object* lean_int_emod(lean_object*, lean_object*); static lean_object* l_Std_Time_Month_Ordinal_toSeconds___closed__26; static lean_object* l_Std_Time_Month_instInhabitedOrdinal___closed__9; -static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__11; static lean_object* l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__5; static lean_object* l_Std_Time_Month_instOffsetBEq___closed__1; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__14; static lean_object* l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__22; static lean_object* l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__20; static lean_object* l_Std_Time_Month_Ordinal_august___closed__2; @@ -77,10 +84,7 @@ static lean_object* l_Std_Time_Month_Ordinal_toSeconds___closed__13; LEAN_EXPORT lean_object* l_Std_Time_Month_instOffsetAdd; static lean_object* l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__17; static lean_object* l_Std_Time_Month_Ordinal_october___closed__1; -static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__15; -static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__19; static lean_object* l_Std_Time_Month_Ordinal_may___closed__5; -static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__16; static lean_object* l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__24; static lean_object* l_Std_Time_Month_Ordinal_april___closed__6; static lean_object* l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__32; @@ -90,9 +94,6 @@ LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_toMinutes(uint8_t, lean_object LEAN_EXPORT lean_object* l_Std_Time_Month_instOffsetRepr; static lean_object* l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__26; static lean_object* l_Std_Time_Month_Ordinal_toSeconds___closed__10; -static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__26; -static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__1; -static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__4; LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_ofInt___boxed(lean_object*, lean_object*); static lean_object* l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__31; static lean_object* l_Std_Time_Month_Ordinal_june___closed__6; @@ -100,13 +101,16 @@ static lean_object* l_Std_Time_Month_Ordinal_december___closed__3; static lean_object* l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__30; static lean_object* l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__8; static lean_object* l_Std_Time_Month_Ordinal_days___closed__3; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__1; static lean_object* l_Std_Time_Month_instOffsetToString___closed__1; lean_object* l_Std_Time_Internal_Bounded_instRepr___rarg___boxed(lean_object*, lean_object*); static lean_object* l_Std_Time_Month_instOrdinalBEq___closed__1; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__15; static lean_object* l_Std_Time_Month_instOfNatOrdinal___closed__1; static lean_object* l_Std_Time_Month_Ordinal_days___closed__5; static lean_object* l_Std_Time_Month_Ordinal_toSeconds___closed__24; -static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__21; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__2; +LEAN_EXPORT uint8_t l_Std_Time_Month_instDecidableLeOffset(lean_object*, lean_object*); static lean_object* l_Std_Time_Month_instInhabitedOrdinal___closed__6; static lean_object* l_Std_Time_Month_Ordinal_november___closed__4; LEAN_EXPORT lean_object* l_Std_Time_Month_instOffsetLE; @@ -116,9 +120,9 @@ static lean_object* l_Std_Time_Month_Ordinal_toSeconds___closed__5; LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_toNat___boxed(lean_object*); extern lean_object* l_Int_instNegInt; static lean_object* l_Std_Time_Month_Ordinal_april___closed__4; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__24; static lean_object* l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__9; lean_object* lean_nat_to_int(lean_object*); -static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__10; static lean_object* l_Std_Time_Month_Ordinal_toSeconds___closed__20; static lean_object* l_Std_Time_Month_Ordinal_september___closed__5; static lean_object* l_Std_Time_Month_Ordinal_december___closed__5; @@ -134,38 +138,35 @@ static lean_object* l_Std_Time_Month_Ordinal_march___closed__5; lean_object* l_outOfBounds___rarg(lean_object*); static lean_object* l_Std_Time_Month_Ordinal_july___closed__6; static lean_object* l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__7; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__9; static lean_object* l_Std_Time_Month_Ordinal_toDays___closed__1; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__16; lean_object* l_Std_Time_Internal_Bounded_LE_instOfNatHAddIntCast(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_july; static lean_object* l_Std_Time_Month_Ordinal_toSeconds___closed__7; LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_toOffset(lean_object*); static lean_object* l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__10; static lean_object* l_Std_Time_Month_Ordinal_may___closed__1; -static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__13; extern lean_object* l_Int_instMul; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__17; LEAN_EXPORT lean_object* l_Std_Time_Month_Offset_ofNat(lean_object*); lean_object* l_instReprInt___boxed(lean_object*, lean_object*); static lean_object* l_Std_Time_Month_Ordinal_december___closed__4; extern lean_object* l_Int_instLTInt; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__10; static lean_object* l_Std_Time_Month_Ordinal_toSeconds___closed__25; LEAN_EXPORT lean_object* l_Std_Time_Month_Offset_ofInt(lean_object*); static lean_object* l_Std_Time_Month_Ordinal_november___closed__3; -static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__22; static lean_object* l_Std_Time_Month_Ordinal_october___closed__6; LEAN_EXPORT lean_object* l_Std_Time_Month_instOrdinalLT; lean_object* l_Std_Internal_Rat_div(lean_object*, lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_toDays(uint8_t, lean_object*); LEAN_EXPORT uint8_t l_Std_Time_Month_instDecidableLeOrdinal(lean_object*, lean_object*); -static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__23; -static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__24; -static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__17; static lean_object* l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__25; static lean_object* l_Std_Time_Month_Ordinal_toSeconds___closed__11; static lean_object* l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__13; LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_march; -static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__2; -static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__18; static lean_object* l_Std_Time_Month_Ordinal_june___closed__3; LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_clipDay(uint8_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Time_Month_instOffsetToString; @@ -177,6 +178,7 @@ static lean_object* l_Std_Time_Month_Ordinal_toNat___closed__1; LEAN_EXPORT lean_object* l_Std_Time_Month_instOffsetDecidableEq___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_may; static lean_object* l_Std_Time_Month_instInhabitedOrdinal___closed__8; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__6; static lean_object* l_Std_Time_Month_Ordinal_june___closed__1; lean_object* lean_int_sub(lean_object*, lean_object*); static lean_object* l_Std_Time_Month_Ordinal_toSeconds___closed__8; @@ -184,24 +186,20 @@ LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_december; static lean_object* l_Std_Time_Month_Ordinal_toDays___closed__3; LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_april; static lean_object* l_Std_Time_Month_Ordinal_april___closed__1; -static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__25; static lean_object* l_Std_Time_Month_Ordinal_november___closed__2; LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_june; static lean_object* l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__12; static lean_object* l_Std_Time_Month_Ordinal_february___closed__2; static lean_object* l_Std_Time_Month_Ordinal_toSeconds___closed__6; -static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__6; -static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__27; -LEAN_EXPORT lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_226_; static lean_object* l_Std_Time_Month_Ordinal_july___closed__3; static lean_object* l_Std_Time_Month_Ordinal_september___closed__2; static lean_object* l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__4; static lean_object* l_Std_Time_Month_Ordinal_may___closed__4; lean_object* lean_nat_abs(lean_object*); -static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__12; static lean_object* l_Std_Time_Month_Ordinal_toSeconds___closed__2; -static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__8; LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_november; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__26; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__20; static lean_object* l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__21; lean_object* lean_int_mul(lean_object*, lean_object*); static lean_object* l_Std_Time_Month_Ordinal_july___closed__4; @@ -221,6 +219,7 @@ LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_february; extern lean_object* l_Int_instAdd; LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_cumulativeDays(uint8_t, lean_object*); LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_august; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__18; static lean_object* l_Std_Time_Month_Ordinal_october___closed__3; static lean_object* l_Std_Time_Month_Ordinal_toSeconds___closed__3; static lean_object* l_Std_Time_Month_Ordinal_june___closed__5; @@ -234,7 +233,6 @@ LEAN_EXPORT lean_object* l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_ static lean_object* l_Std_Time_Month_Ordinal_toSeconds___closed__15; lean_object* l_Int_toNat(lean_object*); static lean_object* l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__19; -static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__5; uint8_t lean_int_dec_lt(lean_object*, lean_object*); static lean_object* l_Std_Time_Month_Ordinal_may___closed__6; static lean_object* l_Std_Time_Month_Ordinal_toSeconds___closed__4; @@ -243,12 +241,17 @@ static lean_object* l_Std_Time_Month_Ordinal_toSeconds___closed__17; static lean_object* l_Std_Time_Month_instOrdinalRepr___closed__1; static lean_object* l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__23; static lean_object* l_Std_Time_Month_Ordinal_july___closed__2; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__8; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__21; lean_object* lean_array_mk(lean_object*); static lean_object* l_Std_Time_Month_Ordinal_toSeconds___closed__19; +LEAN_EXPORT lean_object* l_Std_Time_Month_instDecidableLtOffset___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_september; static lean_object* l_Std_Time_Month_Ordinal_toSeconds___closed__12; static lean_object* l_Std_Time_Month_Ordinal_august___closed__6; static lean_object* l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__14; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__13; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__27; lean_object* lean_int_add(lean_object*, lean_object*); static lean_object* l_Std_Time_Month_Ordinal_february___closed__6; static lean_object* l_Std_Time_Month_Ordinal_days___closed__2; @@ -259,8 +262,8 @@ uint8_t lean_int_dec_eq(lean_object*, lean_object*); static lean_object* l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__2; static lean_object* l_Std_Time_Month_Ordinal_june___closed__2; static lean_object* l_Std_Time_Month_Ordinal_september___closed__3; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__3; LEAN_EXPORT lean_object* l_Std_Time_Month_instOrdinalBEq; -static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__14; static lean_object* l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__27; LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_ofNat(lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); @@ -277,13 +280,13 @@ static lean_object* l_Std_Time_Month_Ordinal_july___closed__1; static lean_object* l_Std_Time_Month_Ordinal_may___closed__3; static lean_object* l_Std_Time_Month_Ordinal_toDays___closed__4; static lean_object* l_Std_Time_Month_Ordinal_toMinutes___closed__1; -static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__3; extern lean_object* l_Int_instLEInt; static lean_object* l_Std_Time_Month_instInhabitedOrdinal___closed__5; static lean_object* l_Std_Time_Month_instInhabitedOrdinal___closed__7; LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_january; extern lean_object* l_Std_Time_Day_instOffsetInhabited; static lean_object* l_Std_Time_Month_Ordinal_september___closed__4; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__5; static lean_object* l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__1; static lean_object* l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__15; static lean_object* l_Std_Time_Month_Ordinal_june___closed__4; @@ -297,6 +300,7 @@ static lean_object* l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordin LEAN_EXPORT lean_object* l_Std_Time_Month_instOffsetMul; LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_days___boxed(lean_object*, lean_object*); static lean_object* l_Std_Time_Month_Ordinal_august___closed__3; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__19; static lean_object* l_Std_Time_Month_Ordinal_toDays___closed__2; static lean_object* _init_l_Std_Time_Month_instOrdinalRepr___closed__1() { _start: @@ -640,6 +644,44 @@ x_4 = lean_box(x_3); return x_4; } } +LEAN_EXPORT uint8_t l_Std_Time_Month_instDecidableLeOffset(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = lean_int_dec_le(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Month_instDecidableLeOffset___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Std_Time_Month_instDecidableLeOffset(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT uint8_t l_Std_Time_Month_instDecidableLtOffset(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = lean_int_dec_lt(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Month_instDecidableLtOffset___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Std_Time_Month_instDecidableLtOffset(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_Std_Time_Month_instOfNatOffset(lean_object* x_1) { _start: { @@ -1473,7 +1515,7 @@ lean_dec(x_1); return x_3; } } -static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__1() { +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__1() { _start: { lean_object* x_1; @@ -1481,7 +1523,7 @@ x_1 = lean_mk_string_unchecked("Lean", 4, 4); return x_1; } } -static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__2() { +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__2() { _start: { lean_object* x_1; @@ -1489,7 +1531,7 @@ x_1 = lean_mk_string_unchecked("Parser", 6, 6); return x_1; } } -static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__3() { +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__3() { _start: { lean_object* x_1; @@ -1497,7 +1539,7 @@ x_1 = lean_mk_string_unchecked("Tactic", 6, 6); return x_1; } } -static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__4() { +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__4() { _start: { lean_object* x_1; @@ -1505,19 +1547,19 @@ x_1 = lean_mk_string_unchecked("tacticSeq", 9, 9); return x_1; } } -static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__5() { +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____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_Std_Time_Date_Unit_Month___hyg_226____closed__1; -x_2 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__2; -x_3 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__3; -x_4 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__4; +x_1 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__1; +x_2 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__2; +x_3 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__3; +x_4 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____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_Std_Time_Date_Unit_Month___hyg_226____closed__6() { +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__6() { _start: { lean_object* x_1; lean_object* x_2; @@ -1526,7 +1568,7 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__7() { +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__7() { _start: { lean_object* x_1; @@ -1534,19 +1576,19 @@ x_1 = lean_mk_string_unchecked("tacticSeq1Indented", 18, 18); return x_1; } } -static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__8() { +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____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_Std_Time_Date_Unit_Month___hyg_226____closed__1; -x_2 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__2; -x_3 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__3; -x_4 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__7; +x_1 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__1; +x_2 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__2; +x_3 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__3; +x_4 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____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_Std_Time_Date_Unit_Month___hyg_226____closed__9() { +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__9() { _start: { lean_object* x_1; @@ -1554,17 +1596,17 @@ x_1 = lean_mk_string_unchecked("null", 4, 4); return x_1; } } -static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__10() { +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____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_Std_Time_Date_Unit_Month___hyg_226____closed__9; +x_2 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__9; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__11() { +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__11() { _start: { lean_object* x_1; @@ -1572,41 +1614,41 @@ x_1 = lean_mk_string_unchecked("decide", 6, 6); return x_1; } } -static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__12() { +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____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_Std_Time_Date_Unit_Month___hyg_226____closed__1; -x_2 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__2; -x_3 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__3; -x_4 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__11; +x_1 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__1; +x_2 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__2; +x_3 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__3; +x_4 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____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_Std_Time_Date_Unit_Month___hyg_226____closed__13() { +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____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_Std_Time_Date_Unit_Month___hyg_226____closed__11; +x_2 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____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_Std_Time_Date_Unit_Month___hyg_226____closed__14() { +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__6; -x_2 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__13; +x_1 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__6; +x_2 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__13; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__15() { +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__15() { _start: { lean_object* x_1; @@ -1614,25 +1656,25 @@ x_1 = lean_mk_string_unchecked("optConfig", 9, 9); return x_1; } } -static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__16() { +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____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_Std_Time_Date_Unit_Month___hyg_226____closed__1; -x_2 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__2; -x_3 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__3; -x_4 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__15; +x_1 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__1; +x_2 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__2; +x_3 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__3; +x_4 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____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_Std_Time_Date_Unit_Month___hyg_226____closed__17() { +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____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_Std_Time_Date_Unit_Month___hyg_226____closed__10; -x_3 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__6; +x_2 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__10; +x_3 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____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); @@ -1640,23 +1682,23 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__18() { +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__6; -x_2 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__17; +x_1 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__6; +x_2 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__17; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__19() { +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__19() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); -x_2 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__16; -x_3 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__18; +x_2 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__16; +x_3 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__18; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -1664,23 +1706,23 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__20() { +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__20() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__14; -x_2 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__19; +x_1 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__14; +x_2 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__19; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__21() { +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__21() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); -x_2 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__12; -x_3 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__20; +x_2 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__12; +x_3 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__20; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -1688,23 +1730,23 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__22() { +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__22() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__6; -x_2 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__21; +x_1 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__6; +x_2 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__21; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__23() { +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__23() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); -x_2 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__10; -x_3 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__22; +x_2 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__10; +x_3 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__22; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -1712,23 +1754,23 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__24() { +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__24() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__6; -x_2 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__23; +x_1 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__6; +x_2 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__23; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__25() { +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__25() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); -x_2 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__8; -x_3 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__24; +x_2 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__8; +x_3 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__24; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -1736,23 +1778,23 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__26() { +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__26() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__6; -x_2 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__25; +x_1 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__6; +x_2 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__25; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__27() { +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____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_Std_Time_Date_Unit_Month___hyg_226____closed__5; -x_3 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__26; +x_2 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__5; +x_3 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____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); @@ -1760,11 +1802,11 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226_() { +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_256_() { _start: { lean_object* x_1; -x_1 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__27; +x_1 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__27; return x_1; } } @@ -3094,62 +3136,62 @@ l_Std_Time_Month_Ordinal_december___closed__6 = _init_l_Std_Time_Month_Ordinal_d lean_mark_persistent(l_Std_Time_Month_Ordinal_december___closed__6); l_Std_Time_Month_Ordinal_december = _init_l_Std_Time_Month_Ordinal_december(); lean_mark_persistent(l_Std_Time_Month_Ordinal_december); -l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__1 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__1(); -lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__1); -l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__2 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__2(); -lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__2); -l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__3 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__3(); -lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__3); -l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__4 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__4(); -lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__4); -l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__5 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__5(); -lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__5); -l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__6 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__6(); -lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__6); -l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__7 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__7(); -lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__7); -l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__8 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__8(); -lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__8); -l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__9 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__9(); -lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__9); -l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__10 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__10(); -lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__10); -l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__11 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__11(); -lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__11); -l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__12 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__12(); -lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__12); -l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__13 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__13(); -lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__13); -l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__14 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__14(); -lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__14); -l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__15 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__15(); -lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__15); -l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__16 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__16(); -lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__16); -l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__17 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__17(); -lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__17); -l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__18 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__18(); -lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__18); -l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__19 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__19(); -lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__19); -l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__20 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__20(); -lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__20); -l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__21 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__21(); -lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__21); -l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__22 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__22(); -lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__22); -l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__23 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__23(); -lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__23); -l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__24 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__24(); -lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__24); -l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__25 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__25(); -lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__25); -l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__26 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__26(); -lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__26); -l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__27 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__27(); -lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__27); -l___auto____x40_Std_Time_Date_Unit_Month___hyg_226_ = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226_(); -lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_226_); +l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__1 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__1(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__1); +l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__2 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__2(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__2); +l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__3 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__3(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__3); +l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__4 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__4(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__4); +l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__5 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__5(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__5); +l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__6 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__6(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__6); +l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__7 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__7(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__7); +l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__8 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__8(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__8); +l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__9 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__9(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__9); +l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__10 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__10(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__10); +l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__11 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__11(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__11); +l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__12 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__12(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__12); +l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__13 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__13(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__13); +l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__14 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__14(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__14); +l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__15 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__15(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__15); +l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__16 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__16(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__16); +l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__17 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__17(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__17); +l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__18 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__18(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__18); +l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__19 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__19(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__19); +l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__20 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__20(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__20); +l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__21 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__21(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__21); +l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__22 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__22(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__22); +l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__23 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__23(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__23); +l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__24 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__24(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__24); +l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__25 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__25(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__25); +l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__26 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__26(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__26); +l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__27 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__27(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_256____closed__27); +l___auto____x40_Std_Time_Date_Unit_Month___hyg_256_ = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_256_(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_256_); l_Std_Time_Month_Ordinal_toNat___closed__1 = _init_l_Std_Time_Month_Ordinal_toNat___closed__1(); lean_mark_persistent(l_Std_Time_Month_Ordinal_toNat___closed__1); l_Std_Time_Month_Ordinal_toSeconds___closed__1 = _init_l_Std_Time_Month_Ordinal_toSeconds___closed__1(); diff --git a/stage0/stdlib/Std/Time/Date/Unit/Week.c b/stage0/stdlib/Std/Time/Date/Unit/Week.c index d05fb893529b..60aa81222a68 100644 --- a/stage0/stdlib/Std/Time/Date/Unit/Week.c +++ b/stage0/stdlib/Std/Time/Date/Unit/Week.c @@ -14,99 +14,102 @@ extern "C" { #endif static lean_object* l_Std_Time_Week_Offset_toHours___closed__1; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__25; static lean_object* l_Std_Time_Week_instOffsetNeg___closed__1; lean_object* lean_mk_empty_array_with_capacity(lean_object*); LEAN_EXPORT lean_object* l_Std_Time_Week_instOrdinalLE; LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_toMilliseconds___boxed(lean_object*); LEAN_EXPORT lean_object* l_Std_Time_Week_Ordinal_ofInt(lean_object*, lean_object*); +static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__11; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__21; LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_ofHours___boxed(lean_object*); LEAN_EXPORT uint8_t l_Std_Time_Week_instDecidableLtOrdinal(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Time_Week_instOffsetLT; static lean_object* l_Std_Time_Week_Offset_toMinutes___closed__1; lean_object* lean_array_push(lean_object*, lean_object*); -static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__1; LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_ofInt___boxed(lean_object*); static lean_object* l_Std_Time_Week_instInhabitedOrdinal___closed__3; LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_ofDays(lean_object*); +static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__20; LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_ofSeconds(lean_object*); -static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__26; static lean_object* l_Std_Time_Week_instInhabitedOrdinal___closed__4; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__7; static lean_object* l_Std_Time_Week_Offset_toDays___closed__1; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__1; lean_object* lean_int_emod(lean_object*, lean_object*); lean_object* l_Std_Time_Internal_UnitVal_instToString___rarg___boxed(lean_object*); LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_toMilliseconds(lean_object*); static lean_object* l_Std_Time_Week_instInhabitedOrdinal___closed__9; LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_ofMinutes(lean_object*); LEAN_EXPORT lean_object* l_Std_Time_Week_instOffsetInhabited; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__2; LEAN_EXPORT lean_object* l_Std_Time_Week_instOfNatOffset(lean_object*); +static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__6; LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_ofHours(lean_object*); +static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__27; uint8_t lean_int_dec_le(lean_object*, lean_object*); static lean_object* l_Std_Time_Week_instOffsetInhabited___closed__1; LEAN_EXPORT lean_object* l_Std_Time_Week_instInhabitedOrdinal; -static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__2; +LEAN_EXPORT uint8_t l_Std_Time_Week_instDecidableLtOffset(lean_object*, lean_object*); lean_object* l_Std_Time_Internal_Bounded_instRepr___rarg___boxed(lean_object*, lean_object*); -static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__13; static lean_object* l_Std_Time_Week_instInhabitedOrdinal___closed__6; LEAN_EXPORT lean_object* l_Std_Time_Week_instOffsetBEq; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__5; +LEAN_EXPORT lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_195_; LEAN_EXPORT lean_object* l_Std_Time_Week_Ordinal_ofNat(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Time_Week_instOrdinalLT; LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_toSeconds___boxed(lean_object*); LEAN_EXPORT lean_object* l_Std_Time_Week_instOffsetToString; -static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__8; LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_ofMilliseconds(lean_object*); lean_object* lean_nat_to_int(lean_object*); -static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__20; -static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__14; -static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__23; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__15; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__22; LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_toDays(lean_object*); lean_object* l_Std_Time_Internal_UnitVal_instNeg___rarg___boxed(lean_object*); -static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__10; LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_ofNanoseconds(lean_object*); static lean_object* l_Std_Time_Week_instOfNatOrdinal___closed__1; -static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__25; lean_object* l_Std_Time_Internal_Bounded_LE_instOfNatHAddIntCast(lean_object*, lean_object*, lean_object*); -static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__4; lean_object* l_Std_Time_Internal_UnitVal_add___rarg___boxed(lean_object*, lean_object*); -static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__7; LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_toNanoseconds(lean_object*); LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_ofDays___boxed(lean_object*); LEAN_EXPORT lean_object* l_Std_Time_Week_instOrdinalBEq; -static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__15; static lean_object* l_Std_Time_Week_instOffsetBEq___closed__1; LEAN_EXPORT lean_object* l_Std_Time_Week_Ordinal_toOffset___boxed(lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); static lean_object* l_Std_Time_Week_Offset_toNanoseconds___closed__1; lean_object* l___private_Std_Time_Internal_UnitVal_0__Std_Time_Internal_beqUnitVal____x40_Std_Time_Internal_UnitVal___hyg_35____rarg___boxed(lean_object*, lean_object*); -static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__3; -static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__9; -static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__24; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__12; static lean_object* l_Std_Time_Week_instInhabitedOrdinal___closed__1; LEAN_EXPORT lean_object* l_Std_Time_Week_instOffsetRepr; LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_toMinutes___boxed(lean_object*); +LEAN_EXPORT uint8_t l_Std_Time_Week_instDecidableLeOffset(lean_object*, lean_object*); static lean_object* l_Std_Time_Week_instOffsetRepr___closed__1; static lean_object* l_Std_Time_Week_instOrdinalBEq___closed__1; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__19; LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_toNanoseconds___boxed(lean_object*); static lean_object* l_Std_Time_Week_instOrdinalRepr___closed__1; LEAN_EXPORT lean_object* l_Std_Time_Week_Ordinal_instOfMonthRepr; lean_object* lean_int_sub(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Time_Week_Ordinal_toOffset(lean_object*); LEAN_EXPORT lean_object* l_Std_Time_Week_instOrdinalRepr; -static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__11; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__3; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__16; LEAN_EXPORT lean_object* l_Std_Time_Week_instOffsetLE; lean_object* l_Std_Time_Internal_UnitVal_sub___rarg___boxed(lean_object*, lean_object*); +static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__17; LEAN_EXPORT lean_object* l_Std_Time_Week_instOffsetSub; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__4; static lean_object* l_Std_Time_Week_instInhabitedOrdinal___closed__2; lean_object* lean_int_mul(lean_object*, lean_object*); +static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__26; LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_ofNanoseconds___boxed(lean_object*); LEAN_EXPORT lean_object* l_Std_Time_Week_Ordinal_ofFin(lean_object*); lean_object* l_Std_Time_Internal_UnitVal_instRepr___rarg___boxed(lean_object*, lean_object*); +static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__14; LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_toDays___boxed(lean_object*); -static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__12; -static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__22; static lean_object* l_Std_Time_Week_Offset_toSeconds___closed__1; LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_ofMilliseconds___boxed(lean_object*); static lean_object* l_Std_Time_Week_instInhabitedOrdinal___closed__5; -static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__16; LEAN_EXPORT lean_object* l_Std_Time_Week_instOffsetAdd; LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_toMinutes(lean_object*); LEAN_EXPORT uint8_t l_Std_Time_Week_instDecidableLeOrdinal(lean_object*, lean_object*); @@ -116,32 +119,33 @@ LEAN_EXPORT lean_object* l_Std_Time_Week_instOfNatOrdinal(lean_object*); LEAN_EXPORT lean_object* l_Std_Time_Week_Ordinal_ofInt___boxed(lean_object*, lean_object*); static lean_object* l_Std_Time_Week_instOffsetToString___closed__1; LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_ofNat(lean_object*); -LEAN_EXPORT lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_149_; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__9; LEAN_EXPORT lean_object* l_Std_Time_Week_instOffsetNeg; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__10; LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_ofInt(lean_object*); +static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__24; static lean_object* l_Std_Time_Week_instInhabitedOrdinal___closed__7; LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_toSeconds(lean_object*); LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_ofMinutes___boxed(lean_object*); -static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__17; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__8; LEAN_EXPORT lean_object* l_Std_Time_Week_instDecidableLeOrdinal___boxed(lean_object*, lean_object*); lean_object* lean_int_add(lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__13; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__18; static lean_object* l_Std_Time_Week_Offset_toMilliseconds___closed__1; -static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__18; +LEAN_EXPORT lean_object* l_Std_Time_Week_instDecidableLeOffset___boxed(lean_object*, lean_object*); lean_object* lean_int_ediv(lean_object*, lean_object*); -static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__21; uint8_t lean_nat_dec_le(lean_object*, lean_object*); -static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__27; LEAN_EXPORT lean_object* l_Std_Time_Week_instDecidableLtOrdinal___boxed(lean_object*, lean_object*); lean_object* l_Std_Time_Internal_Bounded_instBEq___rarg___boxed(lean_object*, lean_object*); -static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__6; static lean_object* l_Std_Time_Week_instInhabitedOrdinal___closed__8; +LEAN_EXPORT lean_object* l_Std_Time_Week_instDecidableLtOffset___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_toHours___boxed(lean_object*); LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_ofSeconds___boxed(lean_object*); -static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__5; static lean_object* l_Std_Time_Week_instOffsetAdd___closed__1; static lean_object* l_Std_Time_Week_instOffsetSub___closed__1; -static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__19; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__23; static lean_object* _init_l_Std_Time_Week_instOrdinalRepr___closed__1() { _start: { @@ -472,6 +476,44 @@ x_1 = l_Std_Time_Week_instOffsetToString___closed__1; return x_1; } } +LEAN_EXPORT uint8_t l_Std_Time_Week_instDecidableLeOffset(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = lean_int_dec_le(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Week_instDecidableLeOffset___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Std_Time_Week_instDecidableLeOffset(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT uint8_t l_Std_Time_Week_instDecidableLtOffset(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = lean_int_dec_lt(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Week_instDecidableLtOffset___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Std_Time_Week_instDecidableLtOffset(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_Std_Time_Week_instOfNatOffset(lean_object* x_1) { _start: { @@ -504,7 +546,7 @@ x_1 = l_Std_Time_Week_instOrdinalRepr___closed__1; return x_1; } } -static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__1() { +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__1() { _start: { lean_object* x_1; @@ -512,7 +554,7 @@ x_1 = lean_mk_string_unchecked("Lean", 4, 4); return x_1; } } -static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__2() { +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__2() { _start: { lean_object* x_1; @@ -520,7 +562,7 @@ x_1 = lean_mk_string_unchecked("Parser", 6, 6); return x_1; } } -static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__3() { +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__3() { _start: { lean_object* x_1; @@ -528,7 +570,7 @@ x_1 = lean_mk_string_unchecked("Tactic", 6, 6); return x_1; } } -static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__4() { +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__4() { _start: { lean_object* x_1; @@ -536,19 +578,19 @@ x_1 = lean_mk_string_unchecked("tacticSeq", 9, 9); return x_1; } } -static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__5() { +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____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_Std_Time_Date_Unit_Week___hyg_149____closed__1; -x_2 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__2; -x_3 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__3; -x_4 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__4; +x_1 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__1; +x_2 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__2; +x_3 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__3; +x_4 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____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_Std_Time_Date_Unit_Week___hyg_149____closed__6() { +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__6() { _start: { lean_object* x_1; lean_object* x_2; @@ -557,7 +599,7 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__7() { +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__7() { _start: { lean_object* x_1; @@ -565,19 +607,19 @@ x_1 = lean_mk_string_unchecked("tacticSeq1Indented", 18, 18); return x_1; } } -static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__8() { +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____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_Std_Time_Date_Unit_Week___hyg_149____closed__1; -x_2 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__2; -x_3 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__3; -x_4 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__7; +x_1 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__1; +x_2 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__2; +x_3 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__3; +x_4 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____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_Std_Time_Date_Unit_Week___hyg_149____closed__9() { +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__9() { _start: { lean_object* x_1; @@ -585,17 +627,17 @@ x_1 = lean_mk_string_unchecked("null", 4, 4); return x_1; } } -static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__10() { +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____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_Std_Time_Date_Unit_Week___hyg_149____closed__9; +x_2 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__9; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__11() { +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__11() { _start: { lean_object* x_1; @@ -603,41 +645,41 @@ x_1 = lean_mk_string_unchecked("decide", 6, 6); return x_1; } } -static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__12() { +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____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_Std_Time_Date_Unit_Week___hyg_149____closed__1; -x_2 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__2; -x_3 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__3; -x_4 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__11; +x_1 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__1; +x_2 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__2; +x_3 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__3; +x_4 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____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_Std_Time_Date_Unit_Week___hyg_149____closed__13() { +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____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_Std_Time_Date_Unit_Week___hyg_149____closed__11; +x_2 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____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_Std_Time_Date_Unit_Week___hyg_149____closed__14() { +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__6; -x_2 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__13; +x_1 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__6; +x_2 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__13; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__15() { +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__15() { _start: { lean_object* x_1; @@ -645,25 +687,25 @@ x_1 = lean_mk_string_unchecked("optConfig", 9, 9); return x_1; } } -static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__16() { +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____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_Std_Time_Date_Unit_Week___hyg_149____closed__1; -x_2 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__2; -x_3 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__3; -x_4 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__15; +x_1 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__1; +x_2 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__2; +x_3 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__3; +x_4 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____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_Std_Time_Date_Unit_Week___hyg_149____closed__17() { +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____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_Std_Time_Date_Unit_Week___hyg_149____closed__10; -x_3 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__6; +x_2 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__10; +x_3 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____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); @@ -671,23 +713,23 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__18() { +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__6; -x_2 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__17; +x_1 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__6; +x_2 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__17; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__19() { +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__19() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); -x_2 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__16; -x_3 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__18; +x_2 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__16; +x_3 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__18; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -695,23 +737,23 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__20() { +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__20() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__14; -x_2 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__19; +x_1 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__14; +x_2 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__19; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__21() { +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__21() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); -x_2 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__12; -x_3 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__20; +x_2 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__12; +x_3 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__20; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -719,23 +761,23 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__22() { +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__22() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__6; -x_2 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__21; +x_1 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__6; +x_2 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__21; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__23() { +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__23() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); -x_2 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__10; -x_3 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__22; +x_2 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__10; +x_3 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__22; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -743,23 +785,23 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__24() { +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__24() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__6; -x_2 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__23; +x_1 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__6; +x_2 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__23; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__25() { +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__25() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); -x_2 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__8; -x_3 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__24; +x_2 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__8; +x_3 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__24; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -767,23 +809,23 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__26() { +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__26() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__6; -x_2 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__25; +x_1 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__6; +x_2 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__25; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__27() { +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____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_Std_Time_Date_Unit_Week___hyg_149____closed__5; -x_3 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__26; +x_2 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__5; +x_3 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____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); @@ -791,11 +833,11 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149_() { +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_195_() { _start: { lean_object* x_1; -x_1 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__27; +x_1 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__27; return x_1; } } @@ -1219,62 +1261,62 @@ l_Std_Time_Week_instOffsetToString = _init_l_Std_Time_Week_instOffsetToString(); lean_mark_persistent(l_Std_Time_Week_instOffsetToString); l_Std_Time_Week_Ordinal_instOfMonthRepr = _init_l_Std_Time_Week_Ordinal_instOfMonthRepr(); lean_mark_persistent(l_Std_Time_Week_Ordinal_instOfMonthRepr); -l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__1 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__1(); -lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__1); -l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__2 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__2(); -lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__2); -l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__3 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__3(); -lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__3); -l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__4 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__4(); -lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__4); -l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__5 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__5(); -lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__5); -l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__6 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__6(); -lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__6); -l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__7 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__7(); -lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__7); -l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__8 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__8(); -lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__8); -l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__9 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__9(); -lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__9); -l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__10 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__10(); -lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__10); -l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__11 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__11(); -lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__11); -l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__12 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__12(); -lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__12); -l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__13 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__13(); -lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__13); -l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__14 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__14(); -lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__14); -l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__15 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__15(); -lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__15); -l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__16 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__16(); -lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__16); -l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__17 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__17(); -lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__17); -l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__18 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__18(); -lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__18); -l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__19 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__19(); -lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__19); -l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__20 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__20(); -lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__20); -l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__21 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__21(); -lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__21); -l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__22 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__22(); -lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__22); -l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__23 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__23(); -lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__23); -l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__24 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__24(); -lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__24); -l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__25 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__25(); -lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__25); -l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__26 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__26(); -lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__26); -l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__27 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__27(); -lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__27); -l___auto____x40_Std_Time_Date_Unit_Week___hyg_149_ = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149_(); -lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_149_); +l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__1 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__1(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__1); +l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__2 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__2(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__2); +l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__3 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__3(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__3); +l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__4 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__4(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__4); +l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__5 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__5(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__5); +l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__6 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__6(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__6); +l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__7 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__7(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__7); +l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__8 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__8(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__8); +l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__9 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__9(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__9); +l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__10 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__10(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__10); +l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__11 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__11(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__11); +l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__12 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__12(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__12); +l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__13 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__13(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__13); +l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__14 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__14(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__14); +l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__15 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__15(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__15); +l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__16 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__16(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__16); +l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__17 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__17(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__17); +l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__18 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__18(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__18); +l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__19 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__19(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__19); +l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__20 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__20(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__20); +l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__21 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__21(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__21); +l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__22 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__22(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__22); +l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__23 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__23(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__23); +l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__24 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__24(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__24); +l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__25 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__25(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__25); +l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__26 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__26(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__26); +l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__27 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__27(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_195____closed__27); +l___auto____x40_Std_Time_Date_Unit_Week___hyg_195_ = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_195_(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_195_); l_Std_Time_Week_Offset_toMilliseconds___closed__1 = _init_l_Std_Time_Week_Offset_toMilliseconds___closed__1(); lean_mark_persistent(l_Std_Time_Week_Offset_toMilliseconds___closed__1); l_Std_Time_Week_Offset_toNanoseconds___closed__1 = _init_l_Std_Time_Week_Offset_toNanoseconds___closed__1(); diff --git a/stage0/stdlib/Std/Time/DateTime/Timestamp.c b/stage0/stdlib/Std/Time/DateTime/Timestamp.c index ea8611e7c670..200a12f05bc6 100644 --- a/stage0/stdlib/Std/Time/DateTime/Timestamp.c +++ b/stage0/stdlib/Std/Time/DateTime/Timestamp.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Std.Time.DateTime.Timestamp -// Imports: Std.Time.Internal Init.Data.Int Init.System.IO Std.Time.Time Std.Time.Date Std.Time.Duration +// Imports: Std.Time.Internal Init.System.IO Std.Time.Time Std.Time.Date Std.Time.Duration #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -2101,7 +2101,6 @@ return x_3; } } lean_object* initialize_Std_Time_Internal(uint8_t builtin, lean_object*); -lean_object* initialize_Init_Data_Int(uint8_t builtin, lean_object*); lean_object* initialize_Init_System_IO(uint8_t builtin, lean_object*); lean_object* initialize_Std_Time_Time(uint8_t builtin, lean_object*); lean_object* initialize_Std_Time_Date(uint8_t builtin, lean_object*); @@ -2114,9 +2113,6 @@ _G_initialized = true; res = initialize_Std_Time_Internal(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -res = initialize_Init_Data_Int(builtin, lean_io_mk_world()); -if (lean_io_result_is_error(res)) return res; -lean_dec_ref(res); res = initialize_Init_System_IO(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); diff --git a/stage0/stdlib/Std/Time/Internal/Bounded.c b/stage0/stdlib/Std/Time/Internal/Bounded.c index 0f71ba77fdca..b2b3b98e2994 100644 --- a/stage0/stdlib/Std/Time/Internal/Bounded.c +++ b/stage0/stdlib/Std/Time/Internal/Bounded.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Std.Time.Internal.Bounded -// Imports: Init.Data.Int +// Imports: Init.Omega #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -1938,13 +1938,13 @@ lean_dec(x_1); return x_3; } } -lean_object* initialize_Init_Data_Int(uint8_t builtin, lean_object*); +lean_object* initialize_Init_Omega(uint8_t builtin, lean_object*); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Std_Time_Internal_Bounded(uint8_t builtin, lean_object* w) { lean_object * res; if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); _G_initialized = true; -res = initialize_Init_Data_Int(builtin, lean_io_mk_world()); +res = initialize_Init_Omega(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); l_Std_Time_Internal_Bounded_instRepr___rarg___closed__1 = _init_l_Std_Time_Internal_Bounded_instRepr___rarg___closed__1(); diff --git a/stage0/stdlib/Std/Time/Time/Unit/Hour.c b/stage0/stdlib/Std/Time/Time/Unit/Hour.c index 6827c8cef051..5d8ed60bcdb4 100644 --- a/stage0/stdlib/Std/Time/Time/Unit/Hour.c +++ b/stage0/stdlib/Std/Time/Time/Unit/Hour.c @@ -37,6 +37,7 @@ static lean_object* l_Std_Time_Hour_Ordinal_shiftTo1BasedHour___closed__2; lean_object* l_Std_Time_Internal_Bounded_instRepr___rarg___boxed(lean_object*, lean_object*); static lean_object* l_Std_Time_Hour_instOffsetNeg___closed__1; lean_object* lean_nat_to_int(lean_object*); +LEAN_EXPORT uint8_t l_Std_Time_Hour_instDecidableLtOffset(lean_object*, lean_object*); static lean_object* l_Std_Time_Hour_Ordinal_toRelative___closed__1; LEAN_EXPORT lean_object* l_Std_Time_Hour_Ordinal_ofNat(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Time_Hour_instOfNatOffset(lean_object*); @@ -68,11 +69,13 @@ LEAN_EXPORT lean_object* l_Std_Time_Hour_Ordinal_toOffset(lean_object*); lean_object* l_Std_Time_Internal_UnitVal_instRepr___rarg___boxed(lean_object*, lean_object*); static lean_object* l_Std_Time_Hour_instInhabitedOrdinal___closed__9; LEAN_EXPORT lean_object* l_Std_Time_Hour_Ordinal_toOffset___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Hour_instOffsetLE; LEAN_EXPORT lean_object* l_Std_Time_Hour_Ordinal_toRelative(lean_object*); LEAN_EXPORT lean_object* l_Std_Time_Hour_instOrdinalRepr; static lean_object* l_Std_Time_Hour_Ordinal_shiftTo1BasedHour___closed__1; static lean_object* l_Std_Time_Hour_instOffsetSub___closed__1; uint8_t lean_int_dec_lt(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Hour_instDecidableLeOffset___boxed(lean_object*, lean_object*); static lean_object* l_Std_Time_Hour_instInhabitedOrdinal___closed__2; LEAN_EXPORT lean_object* l_Std_Time_Hour_instOfNatOrdinal(lean_object*); LEAN_EXPORT lean_object* l_Std_Time_Hour_instOffsetInhabited; @@ -85,8 +88,11 @@ LEAN_EXPORT lean_object* l_Std_Time_Hour_Ordinal_toRelative___boxed(lean_object* lean_object* lean_int_add(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Time_Hour_Ordinal_ofFin(lean_object*); static lean_object* l_Std_Time_Hour_instInhabitedOrdinal___closed__4; +LEAN_EXPORT uint8_t l_Std_Time_Hour_instDecidableLeOffset(lean_object*, lean_object*); lean_object* l_Std_Time_Internal_Bounded_instBEq___rarg___boxed(lean_object*, lean_object*); static lean_object* l_Std_Time_Hour_instOffsetAdd___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Hour_instOffsetLT; +LEAN_EXPORT lean_object* l_Std_Time_Hour_instDecidableLtOffset___boxed(lean_object*, lean_object*); static lean_object* _init_l_Std_Time_Hour_instOrdinalRepr___closed__1() { _start: { @@ -401,6 +407,60 @@ x_1 = l_Std_Time_Hour_instOffsetToString___closed__1; return x_1; } } +static lean_object* _init_l_Std_Time_Hour_instOffsetLT() { +_start: +{ +lean_object* x_1; +x_1 = lean_box(0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Hour_instOffsetLE() { +_start: +{ +lean_object* x_1; +x_1 = lean_box(0); +return x_1; +} +} +LEAN_EXPORT uint8_t l_Std_Time_Hour_instDecidableLeOffset(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = lean_int_dec_le(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Hour_instDecidableLeOffset___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Std_Time_Hour_instDecidableLeOffset(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT uint8_t l_Std_Time_Hour_instDecidableLtOffset(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = lean_int_dec_lt(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Hour_instDecidableLtOffset___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Std_Time_Hour_instDecidableLtOffset(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_Std_Time_Hour_instOfNatOffset(lean_object* x_1) { _start: { @@ -703,6 +763,10 @@ l_Std_Time_Hour_instOffsetToString___closed__1 = _init_l_Std_Time_Hour_instOffse lean_mark_persistent(l_Std_Time_Hour_instOffsetToString___closed__1); l_Std_Time_Hour_instOffsetToString = _init_l_Std_Time_Hour_instOffsetToString(); lean_mark_persistent(l_Std_Time_Hour_instOffsetToString); +l_Std_Time_Hour_instOffsetLT = _init_l_Std_Time_Hour_instOffsetLT(); +lean_mark_persistent(l_Std_Time_Hour_instOffsetLT); +l_Std_Time_Hour_instOffsetLE = _init_l_Std_Time_Hour_instOffsetLE(); +lean_mark_persistent(l_Std_Time_Hour_instOffsetLE); l_Std_Time_Hour_Ordinal_toRelative___closed__1 = _init_l_Std_Time_Hour_Ordinal_toRelative___closed__1(); lean_mark_persistent(l_Std_Time_Hour_Ordinal_toRelative___closed__1); l_Std_Time_Hour_Ordinal_toRelative___closed__2 = _init_l_Std_Time_Hour_Ordinal_toRelative___closed__2(); diff --git a/stage0/stdlib/Std/Time/Time/Unit/Millisecond.c b/stage0/stdlib/Std/Time/Time/Unit/Millisecond.c index eff3dae7cf2f..7e2e8a54ea8a 100644 --- a/stage0/stdlib/Std/Time/Time/Unit/Millisecond.c +++ b/stage0/stdlib/Std/Time/Time/Unit/Millisecond.c @@ -16,6 +16,7 @@ extern "C" { static lean_object* l_Std_Time_Millisecond_instOffsetNeg___closed__1; LEAN_EXPORT lean_object* l_Std_Time_Millisecond_instOffsetLT; static lean_object* l_Std_Time_Millisecond_instOfNatOrdinal___closed__1; +LEAN_EXPORT uint8_t l_Std_Time_Millisecond_instDecidableLtOffset(lean_object*, lean_object*); static lean_object* l_Std_Time_Millisecond_instInhabitedOrdinal___closed__5; LEAN_EXPORT lean_object* l_Std_Time_Millisecond_Ordinal_toOffset(lean_object*); LEAN_EXPORT lean_object* l_Std_Time_Millisecond_Ordinal_ofInt___boxed(lean_object*, lean_object*); @@ -60,7 +61,9 @@ LEAN_EXPORT lean_object* l_Std_Time_Millisecond_instOffsetBEq; static lean_object* l_Std_Time_Millisecond_instOffsetAdd___closed__1; lean_object* lean_int_sub(lean_object*, lean_object*); lean_object* l_Std_Time_Internal_UnitVal_sub___rarg___boxed(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Std_Time_Millisecond_instDecidableLeOffset(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Time_Millisecond_Ordinal_ofFin(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_instDecidableLeOffset___boxed(lean_object*, lean_object*); lean_object* l_Std_Time_Internal_UnitVal_instRepr___rarg___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Time_Millisecond_instOrdinalBEq; static lean_object* l_Std_Time_Millisecond_instOffsetToString___closed__1; @@ -76,6 +79,7 @@ LEAN_EXPORT lean_object* l_Std_Time_Millisecond_instOffsetToString; LEAN_EXPORT lean_object* l_Std_Time_Millisecond_Ordinal_ofNat(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Time_Millisecond_instDecidableLtOrdinal___boxed(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Std_Time_Millisecond_instDecidableLtOrdinal(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_instDecidableLtOffset___boxed(lean_object*, lean_object*); static lean_object* _init_l_Std_Time_Millisecond_instOrdinalRepr___closed__1() { _start: { @@ -406,6 +410,44 @@ x_1 = l_Std_Time_Millisecond_instOffsetToString___closed__1; return x_1; } } +LEAN_EXPORT uint8_t l_Std_Time_Millisecond_instDecidableLeOffset(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = lean_int_dec_le(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_instDecidableLeOffset___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Std_Time_Millisecond_instDecidableLeOffset(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT uint8_t l_Std_Time_Millisecond_instDecidableLtOffset(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = lean_int_dec_lt(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_instDecidableLtOffset___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Std_Time_Millisecond_instDecidableLtOffset(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_Std_Time_Millisecond_instOfNatOffset(lean_object* x_1) { _start: { diff --git a/stage0/stdlib/Std/Time/Time/Unit/Minute.c b/stage0/stdlib/Std/Time/Time/Unit/Minute.c index 6197984455e6..1ae464312f46 100644 --- a/stage0/stdlib/Std/Time/Time/Unit/Minute.c +++ b/stage0/stdlib/Std/Time/Time/Unit/Minute.c @@ -19,7 +19,9 @@ static lean_object* l_Std_Time_Minute_instInhabitedOrdinal___closed__2; static lean_object* l_Std_Time_Minute_instInhabitedOrdinal___closed__6; static lean_object* l_Std_Time_Minute_instInhabitedOrdinal___closed__10; LEAN_EXPORT lean_object* l_Std_Time_Minute_Ordinal_ofInt(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Std_Time_Minute_instDecidableLtOffset(lean_object*, lean_object*); static lean_object* l_Std_Time_Minute_instOfNatOrdinal___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Minute_instDecidableLtOffset___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Time_Minute_instDecidableLtOrdinal___boxed(lean_object*, lean_object*); lean_object* lean_int_emod(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Time_Minute_Ordinal_ofInt___boxed(lean_object*, lean_object*); @@ -32,15 +34,19 @@ uint8_t lean_int_dec_le(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Time_Minute_instOffsetRepr; LEAN_EXPORT lean_object* l_Std_Time_Minute_instOffsetAdd; lean_object* l_Std_Time_Internal_Bounded_instRepr___rarg___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Minute_instDecidableLeOffset___boxed(lean_object*, lean_object*); lean_object* lean_nat_to_int(lean_object*); +LEAN_EXPORT uint8_t l_Std_Time_Minute_instDecidableLeOffset(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Time_Minute_instInhabitedOrdinal; lean_object* l_Std_Time_Internal_UnitVal_instNeg___rarg___boxed(lean_object*); LEAN_EXPORT lean_object* l_Std_Time_Minute_instDecidableLeOrdinal___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Time_Minute_Ordinal_ofFin(lean_object*); LEAN_EXPORT lean_object* l_Std_Time_Minute_instOrdinalBEq; lean_object* l_Std_Time_Internal_Bounded_LE_instOfNatHAddIntCast(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Minute_instOffsetLE; lean_object* l_Std_Time_Internal_UnitVal_add___rarg___boxed(lean_object*, lean_object*); static lean_object* l_Std_Time_Minute_instOffsetToString___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Minute_instOffsetLT; LEAN_EXPORT lean_object* l_Std_Time_Minute_instOrdinalLT; LEAN_EXPORT lean_object* l_Std_Time_Minute_instOrdinalRepr; lean_object* l___private_Std_Time_Internal_UnitVal_0__Std_Time_Internal_beqUnitVal____x40_Std_Time_Internal_UnitVal___hyg_35____rarg___boxed(lean_object*, lean_object*); @@ -388,6 +394,60 @@ x_1 = l_Std_Time_Minute_instOffsetToString___closed__1; return x_1; } } +static lean_object* _init_l_Std_Time_Minute_instOffsetLT() { +_start: +{ +lean_object* x_1; +x_1 = lean_box(0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Minute_instOffsetLE() { +_start: +{ +lean_object* x_1; +x_1 = lean_box(0); +return x_1; +} +} +LEAN_EXPORT uint8_t l_Std_Time_Minute_instDecidableLeOffset(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = lean_int_dec_le(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Minute_instDecidableLeOffset___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Std_Time_Minute_instDecidableLeOffset(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT uint8_t l_Std_Time_Minute_instDecidableLtOffset(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = lean_int_dec_lt(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Minute_instDecidableLtOffset___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Std_Time_Minute_instDecidableLtOffset(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_Std_Time_Minute_instOfNatOffset(lean_object* x_1) { _start: { @@ -547,6 +607,10 @@ l_Std_Time_Minute_instOffsetToString___closed__1 = _init_l_Std_Time_Minute_instO lean_mark_persistent(l_Std_Time_Minute_instOffsetToString___closed__1); l_Std_Time_Minute_instOffsetToString = _init_l_Std_Time_Minute_instOffsetToString(); lean_mark_persistent(l_Std_Time_Minute_instOffsetToString); +l_Std_Time_Minute_instOffsetLT = _init_l_Std_Time_Minute_instOffsetLT(); +lean_mark_persistent(l_Std_Time_Minute_instOffsetLT); +l_Std_Time_Minute_instOffsetLE = _init_l_Std_Time_Minute_instOffsetLE(); +lean_mark_persistent(l_Std_Time_Minute_instOffsetLE); return lean_io_result_mk_ok(lean_box(0)); } #ifdef __cplusplus diff --git a/stage0/stdlib/Std/Time/Time/Unit/Nanosecond.c b/stage0/stdlib/Std/Time/Time/Unit/Nanosecond.c index da8a026f08df..ec6dfbffe67a 100644 --- a/stage0/stdlib/Std/Time/Time/Unit/Nanosecond.c +++ b/stage0/stdlib/Std/Time/Time/Unit/Nanosecond.c @@ -21,6 +21,7 @@ LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_Ordinal_ofFin(lean_object*); LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_instOffsetLT; LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_instOrdinalBEq; LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_instDecidableLeOrdinal___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_instDecidableLtOffset___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_Ordinal_ofInt___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_instOffsetRepr; lean_object* l_Std_Time_Internal_UnitVal_instToString___rarg___boxed(lean_object*); @@ -52,6 +53,7 @@ LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_instOfNatOrdinal(lean_object*); lean_object* l___private_Std_Time_Internal_UnitVal_0__Std_Time_Internal_beqUnitVal____x40_Std_Time_Internal_UnitVal___hyg_35____rarg___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_instOrdinalRepr; LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_Span_toOffset___boxed(lean_object*); +LEAN_EXPORT uint8_t l_Std_Time_Nanosecond_instDecidableLtOffset(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_instOfNatOffset(lean_object*); LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_instOffsetToString; LEAN_EXPORT uint8_t l_Std_Time_Nanosecond_instDecidableLeOffset(lean_object*, lean_object*); @@ -338,6 +340,25 @@ x_4 = lean_box(x_3); return x_4; } } +LEAN_EXPORT uint8_t l_Std_Time_Nanosecond_instDecidableLtOffset(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = lean_int_dec_lt(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_instDecidableLtOffset___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Std_Time_Nanosecond_instDecidableLtOffset(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_Std_Time_Nanosecond_instOfNatOffset(lean_object* x_1) { _start: { diff --git a/stage0/stdlib/Std/Time/Time/Unit/Second.c b/stage0/stdlib/Std/Time/Time/Unit/Second.c index 7a364fa63304..24f160c47d48 100644 --- a/stage0/stdlib/Std/Time/Time/Unit/Second.c +++ b/stage0/stdlib/Std/Time/Time/Unit/Second.c @@ -13,6 +13,7 @@ #ifdef __cplusplus extern "C" { #endif +LEAN_EXPORT lean_object* l_Std_Time_Second_instDecidableLtOffset___boxed(lean_object*, lean_object*); static lean_object* l_Std_Time_Second_instOfNatOrdinal___closed__4; LEAN_EXPORT lean_object* l_Std_Time_Second_instReprOrdinal(uint8_t); LEAN_EXPORT lean_object* l_Std_Time_Second_instReprOrdinal___rarg(lean_object*, lean_object*); @@ -49,6 +50,7 @@ lean_object* l_Int_repr(lean_object*); static lean_object* l_Std_Time_Second_instOfNatOrdinal___closed__3; lean_object* l_Std_Time_Internal_UnitVal_add___rarg___boxed(lean_object*, lean_object*); static lean_object* l_Std_Time_Second_instOffsetToString___closed__1; +LEAN_EXPORT uint8_t l_Std_Time_Second_instDecidableLtOffset(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Time_Second_instOffsetRepr; LEAN_EXPORT lean_object* l_Std_Time_Second_instDecidableLeOrdinal(uint8_t); LEAN_EXPORT lean_object* l_Std_Time_Second_instOffsetLT; @@ -56,11 +58,13 @@ lean_object* l___private_Std_Time_Internal_UnitVal_0__Std_Time_Internal_beqUnitV LEAN_EXPORT lean_object* l_Std_Time_Second_Ordinal_ofFin(uint8_t); static lean_object* l_Std_Time_Second_instOffsetBEq___closed__1; LEAN_EXPORT lean_object* l_Std_Time_Second_Offset_ofInt___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Second_instDecidableLeOffset___boxed(lean_object*, lean_object*); lean_object* lean_int_sub(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Time_Second_Ordinal_toOffset___boxed(lean_object*); LEAN_EXPORT lean_object* l_Std_Time_Second_instOfNatOrdinal(uint8_t, lean_object*); static lean_object* l_Std_Time_Second_instOfNatOrdinal___closed__1; lean_object* l_Std_Time_Internal_UnitVal_sub___rarg___boxed(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Std_Time_Second_instDecidableLeOffset(lean_object*, lean_object*); static lean_object* l_Std_Time_Second_instOfNatOrdinal___closed__2; LEAN_EXPORT lean_object* l_Std_Time_Second_instLEOrdinal___boxed(lean_object*); LEAN_EXPORT lean_object* l_Std_Time_Second_instOfNatOffset(lean_object*); @@ -508,6 +512,44 @@ x_1 = l_Std_Time_Second_instOffsetToString___closed__1; return x_1; } } +LEAN_EXPORT uint8_t l_Std_Time_Second_instDecidableLeOffset(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = lean_int_dec_le(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Second_instDecidableLeOffset___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Std_Time_Second_instDecidableLeOffset(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT uint8_t l_Std_Time_Second_instDecidableLtOffset(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = lean_int_dec_lt(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Second_instDecidableLtOffset___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Std_Time_Second_instDecidableLtOffset(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_Std_Time_Second_instOfNatOffset(lean_object* x_1) { _start: { From 918924c16b82b304857f07dfb784e7e25ac5c7a8 Mon Sep 17 00:00:00 2001 From: Alex Keizer Date: Fri, 10 Jan 2025 23:23:58 +0000 Subject: [PATCH 071/100] feat: `BitVec.{toFin, toInt, msb}_umod` (#6404) This PR adds a `toFin` and `msb` lemma for unsigned bitvector modulus. Similar to #6402, we don't provide a general `toInt_umod` lemmas, but instead choose to provide more specialized rewrites, with extra side-conditions. --------- Co-authored-by: Kim Morrison --- src/Init/Data/BitVec/Lemmas.lean | 60 +++++++++++++++++++++++++++++++ src/Init/Data/Nat/Div/Lemmas.lean | 5 +++ 2 files changed, 65 insertions(+) diff --git a/src/Init/Data/BitVec/Lemmas.lean b/src/Init/Data/BitVec/Lemmas.lean index 8b49a5d8beb0..377ab5cc885f 100644 --- a/src/Init/Data/BitVec/Lemmas.lean +++ b/src/Init/Data/BitVec/Lemmas.lean @@ -11,6 +11,7 @@ import Init.Data.Fin.Lemmas import Init.Data.Nat.Lemmas import Init.Data.Nat.Div.Lemmas import Init.Data.Nat.Mod +import Init.Data.Nat.Div.Lemmas import Init.Data.Int.Bitwise.Lemmas import Init.Data.Int.Pow @@ -99,6 +100,12 @@ theorem ofFin_eq_ofNat : @BitVec.ofFin w (Fin.mk x lt) = BitVec.ofNat w x := by theorem eq_of_toNat_eq {n} : ∀ {x y : BitVec n}, x.toNat = y.toNat → x = y | ⟨_, _⟩, ⟨_, _⟩, rfl => rfl +/-- Prove nonequality of bitvectors in terms of nat operations. -/ +theorem toNat_ne_iff_ne {n} {x y : BitVec n} : x.toNat ≠ y.toNat ↔ x ≠ y := by + constructor + · rintro h rfl; apply h rfl + · intro h h_eq; apply h <| eq_of_toNat_eq h_eq + @[simp] theorem val_toFin (x : BitVec w) : x.toFin.val = x.toNat := rfl @[bv_toNat] theorem toNat_eq {x y : BitVec n} : x = y ↔ x.toNat = y.toNat := @@ -2693,6 +2700,10 @@ theorem umod_def {x y : BitVec n} : theorem toNat_umod {x y : BitVec n} : (x % y).toNat = x.toNat % y.toNat := rfl +@[simp] +theorem toFin_umod {x y : BitVec w} : + (x % y).toFin = x.toFin % y.toFin := rfl + @[simp] theorem umod_zero {x : BitVec n} : x % 0#n = x := by simp [umod_def] @@ -2720,6 +2731,55 @@ theorem umod_eq_and {x y : BitVec 1} : x % y = x &&& (~~~y) := by rcases hy with rfl | rfl <;> rfl +theorem umod_eq_of_lt {x y : BitVec w} (h : x < y) : + x % y = x := by + apply eq_of_toNat_eq + simp [Nat.mod_eq_of_lt h] + +@[simp] +theorem msb_umod {x y : BitVec w} : + (x % y).msb = (x.msb && (x < y || y == 0#w)) := by + rw [msb_eq_decide, toNat_umod] + cases msb_x : x.msb + · suffices x.toNat % y.toNat < 2 ^ (w - 1) by simpa + calc + x.toNat % y.toNat ≤ x.toNat := by apply Nat.mod_le + _ < 2 ^ (w - 1) := by simpa [msb_eq_decide] using msb_x + . by_cases hy : y = 0 + · simp_all [msb_eq_decide] + · suffices 2 ^ (w - 1) ≤ x.toNat % y.toNat ↔ x < y by simp_all + by_cases x_lt_y : x < y + . simp_all [Nat.mod_eq_of_lt x_lt_y, msb_eq_decide] + · suffices x.toNat % y.toNat < 2 ^ (w - 1) by + simpa [x_lt_y] + have y_le_x : y.toNat ≤ x.toNat := by + simpa using x_lt_y + replace hy : y.toNat ≠ 0 := + toNat_ne_iff_ne.mpr hy + by_cases msb_y : y.toNat < 2 ^ (w - 1) + · have : x.toNat % y.toNat < y.toNat := Nat.mod_lt _ (by omega) + omega + · rcases w with _|w + · contradiction + simp only [Nat.add_one_sub_one] + replace msb_y : 2 ^ w ≤ y.toNat := by + simpa using msb_y + have : y.toNat ≤ y.toNat * (x.toNat / y.toNat) := by + apply Nat.le_mul_of_pos_right + apply Nat.div_pos y_le_x + omega + have : x.toNat % y.toNat ≤ x.toNat - y.toNat := by + rw [Nat.mod_eq_sub]; omega + omega + +theorem toInt_umod {x y : BitVec w} : + (x % y).toInt = (x.toNat % y.toNat : Int).bmod (2 ^ w) := by + simp [toInt_eq_toNat_bmod] + +theorem toInt_umod_of_msb {x y : BitVec w} (h : x.msb = false) : + (x % y).toInt = x.toInt % y.toNat := by + simp [toInt_eq_msb_cond, h] + /-! ### smtUDiv -/ theorem smtUDiv_eq (x y : BitVec w) : smtUDiv x y = if y = 0#w then allOnes w else x / y := by diff --git a/src/Init/Data/Nat/Div/Lemmas.lean b/src/Init/Data/Nat/Div/Lemmas.lean index d5a26a0950ba..cf312c6a0fff 100644 --- a/src/Init/Data/Nat/Div/Lemmas.lean +++ b/src/Init/Data/Nat/Div/Lemmas.lean @@ -49,6 +49,11 @@ theorem lt_div_mul_self (h : 0 < k) (w : k ≤ x) : x - k < x / k * k := by have : x % k < k := mod_lt x h omega +theorem div_pos (hba : b ≤ a) (hb : 0 < b) : 0 < a / b := by + cases b + · contradiction + · simp [Nat.pos_iff_ne_zero, div_eq_zero_iff_lt, hba] + theorem div_le_div_left (hcb : c ≤ b) (hc : 0 < c) : a / b ≤ a / c := (Nat.le_div_iff_mul_le hc).2 <| Nat.le_trans (Nat.mul_le_mul_left _ hcb) (Nat.div_mul_le_self a b) From 03081a5b6f9395c03f009b7c0b0b83471631b25b Mon Sep 17 00:00:00 2001 From: David Thrane Christiansen Date: Sat, 11 Jan 2025 01:11:20 +0100 Subject: [PATCH 072/100] doc: update FFI description for Int and signed fixed-width ints (#6599) The FFI description didn't mention Int or signed integers. This PR adds `Int` and signed integers to the FFI document. --- doc/dev/ffi.md | 5 +++-- 1 file changed, 3 insertions(+), 2 deletions(-) diff --git a/doc/dev/ffi.md b/doc/dev/ffi.md index 790d88556e6c..8ab18ca47b04 100644 --- a/doc/dev/ffi.md +++ b/doc/dev/ffi.md @@ -49,8 +49,9 @@ In the case of `@[extern]` all *irrelevant* types are removed first; see next se is represented by the representation of that parameter's type. For example, `{ x : α // p }`, the `Subtype` structure of a value of type `α` and an irrelevant proof, is represented by the representation of `α`. -* `Nat` is represented by `lean_object *`. - Its runtime value is either a pointer to an opaque bignum object or, if the lowest bit of the "pointer" is 1 (`lean_is_scalar`), an encoded unboxed natural number (`lean_box`/`lean_unbox`). + Similarly, the signed integer types `Int8`, ..., `Int64`, `ISize` are also represented by the unsigned C types `uint8_t`, ..., `uint64_t`, `size_t`, respectively, because they have a trivial structure. +* `Nat` and `Int` are represented by `lean_object *`. + Their runtime values is either a pointer to an opaque bignum object or, if the lowest bit of the "pointer" is 1 (`lean_is_scalar`), an encoded unboxed natural number or integer (`lean_box`/`lean_unbox`). * A universe `Sort u`, type constructor `... → Sort u`, or proposition `p : Prop` is *irrelevant* and is either statically erased (see above) or represented as a `lean_object *` with the runtime value `lean_box(0)` * Any other type is represented by `lean_object *`. Its runtime value is a pointer to an object of a subtype of `lean_object` (see the "Inductive types" section below) or the unboxed value `lean_box(cidx)` for the `cidx`th constructor of an inductive type if this constructor does not have any relevant parameters. From 8791a9ce069d6dc87f7cccc4387545b1110c89bd Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Sat, 11 Jan 2025 11:32:43 +1100 Subject: [PATCH 073/100] chore: add lean4-cli to release checklist (#6596) Users have requested toolchain tags on `lean4-cli`, so let's add it to the release checklist to make sure these get added regularly. Previously, `lean4-cli` has used more complicated tags, but going forward we're going to just use the simple `v4.16.0` style tags, with no repository-specific versioning. --------- Co-authored-by: Markus Himmel --- doc/dev/release_checklist.md | 31 ++++++++++++++++++------------- script/release_repos.yml | 7 +++++++ 2 files changed, 25 insertions(+), 13 deletions(-) diff --git a/doc/dev/release_checklist.md b/doc/dev/release_checklist.md index 82b39d029f01..0500032df978 100644 --- a/doc/dev/release_checklist.md +++ b/doc/dev/release_checklist.md @@ -37,16 +37,32 @@ We'll use `v4.6.0` as the intended release version as a running example. - Create the tag `v4.6.0` from `master`/`main` and push it. - Merge the tag `v4.6.0` into the `stable` branch and push it. - We do this for the repositories: - - [lean4checker](https://github.com/leanprover/lean4checker) + - [Batteries](https://github.com/leanprover-community/batteries) - No dependencies - Toolchain bump PR - Create and push the tag - Merge the tag into `stable` - - [Batteries](https://github.com/leanprover-community/batteries) + - [lean4checker](https://github.com/leanprover/lean4checker) - No dependencies - Toolchain bump PR - Create and push the tag - Merge the tag into `stable` + - [doc-gen4](https://github.com/leanprover/doc-gen4) + - Dependencies: exist, but they're not part of the release workflow + - Toolchain bump PR including updated Lake manifest + - Create and push the tag + - There is no `stable` branch; skip this step + - [Verso](https://github.com/leanprover/verso) + - Dependencies: exist, but they're not part of the release workflow + - The `SubVerso` dependency should be compatible with _every_ Lean release simultaneously, rather than following this workflow + - Toolchain bump PR including updated Lake manifest + - Create and push the tag + - There is no `stable` branch; skip this step + - [Cli](https://github.com/leanprover/lean4-cli) + - No dependencies + - Toolchain bump PR + - Create and push the tag + - There is no `stable` branch; skip this step - [ProofWidgets4](https://github.com/leanprover-community/ProofWidgets4) - Dependencies: `Batteries` - Note on versions and branches: @@ -61,17 +77,6 @@ We'll use `v4.6.0` as the intended release version as a running example. - Toolchain bump PR including updated Lake manifest - Create and push the tag - Merge the tag into `stable` - - [doc-gen4](https://github.com/leanprover/doc-gen4) - - Dependencies: exist, but they're not part of the release workflow - - Toolchain bump PR including updated Lake manifest - - Create and push the tag - - There is no `stable` branch; skip this step - - [Verso](https://github.com/leanprover/verso) - - Dependencies: exist, but they're not part of the release workflow - - The `SubVerso` dependency should be compatible with _every_ Lean release simultaneously, rather than following this workflow - - Toolchain bump PR including updated Lake manifest - - Create and push the tag - - There is no `stable` branch; skip this step - [import-graph](https://github.com/leanprover-community/import-graph) - Toolchain bump PR including updated Lake manifest - Create and push the tag diff --git a/script/release_repos.yml b/script/release_repos.yml index 10bc53e85e47..5e805e70f936 100644 --- a/script/release_repos.yml +++ b/script/release_repos.yml @@ -27,6 +27,13 @@ repositories: branch: main dependencies: [] + - name: Cli + url: https://github.com/leanprover/lean4-cli + toolchain-tag: true + stable-branch: false + branch: main + dependencies: [] + - name: ProofWidgets4 url: https://github.com/leanprover-community/ProofWidgets4 toolchain-tag: false From acad5879380bf1c113cf54cd6dbfc989fd224010 Mon Sep 17 00:00:00 2001 From: Leonardo de Moura Date: Sat, 11 Jan 2025 17:29:32 -0800 Subject: [PATCH 074/100] fix: pattern selection for local lemmas (#6606) This PR fixes a bug in the pattern selection in the `grind`. --- src/Init/Grind/Norm.lean | 4 ++-- src/Lean/Meta/Tactic/Grind/EMatchTheorem.lean | 2 +- src/Lean/Meta/Tactic/Grind/ForallProp.lean | 13 ++++++++++--- tests/lean/run/grind_t1.lean | 3 +++ 4 files changed, 16 insertions(+), 6 deletions(-) diff --git a/src/Init/Grind/Norm.lean b/src/Init/Grind/Norm.lean index c3d8b62a5f8d..2673caa837c0 100644 --- a/src/Init/Grind/Norm.lean +++ b/src/Init/Grind/Norm.lean @@ -43,8 +43,8 @@ attribute [grind_norm] not_false_eq_true -- Remark: we disabled the following normalization rule because we want this information when implementing splitting heuristics -- Implication as a clause --- @[grind_norm↓] theorem imp_eq (p q : Prop) : (p → q) = (¬ p ∨ q) := by --- by_cases p <;> by_cases q <;> simp [*] +theorem imp_eq (p q : Prop) : (p → q) = (¬ p ∨ q) := by + by_cases p <;> by_cases q <;> simp [*] -- And @[grind_norm↓] theorem not_and (p q : Prop) : (¬(p ∧ q)) = (¬p ∨ ¬q) := by diff --git a/src/Lean/Meta/Tactic/Grind/EMatchTheorem.lean b/src/Lean/Meta/Tactic/Grind/EMatchTheorem.lean index 1170a8dffa80..3b427eca418b 100644 --- a/src/Lean/Meta/Tactic/Grind/EMatchTheorem.lean +++ b/src/Lean/Meta/Tactic/Grind/EMatchTheorem.lean @@ -499,7 +499,7 @@ def addEMatchEqTheorem (declName : Name) : MetaM Unit := do def getEMatchTheorems : CoreM EMatchTheorems := return ematchTheoremsExt.getState (← getEnv) -private inductive TheoremKind where +inductive TheoremKind where | eqLhs | eqRhs | eqBoth | fwd | bwd | default deriving Inhabited, BEq diff --git a/src/Lean/Meta/Tactic/Grind/ForallProp.lean b/src/Lean/Meta/Tactic/Grind/ForallProp.lean index b8be9355c7f1..cca14b85ebe4 100644 --- a/src/Lean/Meta/Tactic/Grind/ForallProp.lean +++ b/src/Lean/Meta/Tactic/Grind/ForallProp.lean @@ -51,6 +51,13 @@ private def isEqTrueHyp? (proof : Expr) : Option FVarId := Id.run do let .fvar fvarId := p | return none return some fvarId +/-- Similar to `mkEMatchTheoremWithKind?`, but swallow any exceptions. -/ +private def mkEMatchTheoremWithKind'? (origin : Origin) (proof : Expr) (kind : TheoremKind) : MetaM (Option EMatchTheorem) := do + try + mkEMatchTheoremWithKind? origin #[] proof kind + catch _ => + return none + private def addLocalEMatchTheorems (e : Expr) : GoalM Unit := do let proof ← mkEqTrueProof e let origin ← if let some fvarId := isEqTrueHyp? proof then @@ -62,12 +69,12 @@ private def addLocalEMatchTheorems (e : Expr) : GoalM Unit := do let size := (← get).newThms.size let gen ← getGeneration e -- TODO: we should have a flag for collecting all unary patterns in a local theorem - if let some thm ← mkEMatchTheoremWithKind? origin #[] proof .fwd then + if let some thm ← mkEMatchTheoremWithKind'? origin proof .fwd then activateTheorem thm gen - if let some thm ← mkEMatchTheoremWithKind? origin #[] proof .bwd then + if let some thm ← mkEMatchTheoremWithKind'? origin proof .bwd then activateTheorem thm gen if (← get).newThms.size == size then - if let some thm ← mkEMatchTheoremWithKind? origin #[] proof .default then + if let some thm ← mkEMatchTheoremWithKind'? origin proof .default then activateTheorem thm gen if (← get).newThms.size == size then trace[grind.issues] "failed to create E-match local theorem for{indentExpr e}" diff --git a/tests/lean/run/grind_t1.lean b/tests/lean/run/grind_t1.lean index 3bf2a21185e0..b8bb7453c097 100644 --- a/tests/lean/run/grind_t1.lean +++ b/tests/lean/run/grind_t1.lean @@ -210,3 +210,6 @@ set_option trace.grind.ematch.instance true in set_option trace.grind.assert true in example (P Q R : α → Prop) (h₁ : ∀ x, Q x → P x) (h₂ : ∀ x, R x → False = (P x)) : Q a → R a → False := by grind + +example (w : Nat → Type) (h : ∀ n, Subsingleton (w n)) : True := by + grind From 7ea5504af26b8382ab9bd6d96508c7765f9c55b3 Mon Sep 17 00:00:00 2001 From: Leonardo de Moura Date: Sat, 11 Jan 2025 18:25:02 -0800 Subject: [PATCH 075/100] feat: add support for splitting on `<->` to `grind` (#6607) This PR adds support for case-splitting on `<->` (and `@Eq Prop`) in the `grind` tactic. --- src/Init/Grind/Lemmas.lean | 6 ++++++ src/Lean/Meta/Tactic/Grind/Core.lean | 22 ++++++++++++++------- src/Lean/Meta/Tactic/Grind/Internalize.lean | 8 ++++++++ src/Lean/Meta/Tactic/Grind/Split.lean | 10 ++++++++++ tests/lean/run/grind_implies.lean | 20 +++++++++++++++---- tests/lean/run/grind_pre.lean | 17 +++++++++++++++- tests/lean/run/grind_t1.lean | 12 +++++++++++ 7 files changed, 83 insertions(+), 12 deletions(-) diff --git a/src/Init/Grind/Lemmas.lean b/src/Init/Grind/Lemmas.lean index 6707d1b16146..6a6a571abb8a 100644 --- a/src/Init/Grind/Lemmas.lean +++ b/src/Init/Grind/Lemmas.lean @@ -66,6 +66,12 @@ theorem eq_eq_of_eq_true_right {a b : Prop} (h : b = True) : (a = b) = a := by s theorem eq_congr {α : Sort u} {a₁ b₁ a₂ b₂ : α} (h₁ : a₁ = a₂) (h₂ : b₁ = b₂) : (a₁ = b₁) = (a₂ = b₂) := by simp [*] theorem eq_congr' {α : Sort u} {a₁ b₁ a₂ b₂ : α} (h₁ : a₁ = b₂) (h₂ : b₁ = a₂) : (a₁ = b₁) = (a₂ = b₂) := by rw [h₁, h₂, Eq.comm (a := a₂)] +/- The following two helper theorems are used to case-split `a = b` representing `iff`. -/ +theorem of_eq_eq_true {a b : Prop} (h : (a = b) = True) : (¬a ∨ b) ∧ (¬b ∨ a) := by + by_cases a <;> by_cases b <;> simp_all +theorem of_eq_eq_false {a b : Prop} (h : (a = b) = False) : (¬a ∨ ¬b) ∧ (b ∨ a) := by + by_cases a <;> by_cases b <;> simp_all + /-! Forall -/ theorem forall_propagator (p : Prop) (q : p → Prop) (q' : Prop) (h₁ : p = True) (h₂ : q (of_eq_true h₁) = q') : (∀ hp : p, q hp) = q' := by diff --git a/src/Lean/Meta/Tactic/Grind/Core.lean b/src/Lean/Meta/Tactic/Grind/Core.lean index 4aab792f849d..1b0962a229e0 100644 --- a/src/Lean/Meta/Tactic/Grind/Core.lean +++ b/src/Lean/Meta/Tactic/Grind/Core.lean @@ -217,14 +217,22 @@ def add (fact : Expr) (proof : Expr) (generation := 0) : GoalM Unit := do where go (p : Expr) (isNeg : Bool) : GoalM Unit := do match_expr p with - | Eq _ lhs rhs => goEq p lhs rhs isNeg false - | HEq _ lhs _ rhs => goEq p lhs rhs isNeg true - | _ => - internalize p generation - if isNeg then - addEq p (← getFalseExpr) (← mkEqFalse proof) + | Eq α lhs rhs => + if α.isProp then + -- It is morally an iff. + -- We do not use the `goEq` optimization because we want to register `p` as a case-split + goFact p isNeg else - addEq p (← getTrueExpr) (← mkEqTrue proof) + goEq p lhs rhs isNeg false + | HEq _ lhs _ rhs => goEq p lhs rhs isNeg true + | _ => goFact p isNeg + + goFact (p : Expr) (isNeg : Bool) : GoalM Unit := do + internalize p generation + if isNeg then + addEq p (← getFalseExpr) (← mkEqFalse proof) + else + addEq p (← getTrueExpr) (← mkEqTrue proof) goEq (p : Expr) (lhs rhs : Expr) (isNeg : Bool) (isHEq : Bool) : GoalM Unit := do if isNeg then diff --git a/src/Lean/Meta/Tactic/Grind/Internalize.lean b/src/Lean/Meta/Tactic/Grind/Internalize.lean index eaaf18f107e0..5fa727302445 100644 --- a/src/Lean/Meta/Tactic/Grind/Internalize.lean +++ b/src/Lean/Meta/Tactic/Grind/Internalize.lean @@ -53,12 +53,20 @@ private def addSplitCandidate (e : Expr) : GoalM Unit := do -- TODO: add attribute to make this extensible private def forbiddenSplitTypes := [``Eq, ``HEq, ``True, ``False] +/-- Returns `true` if `e` is of the form `@Eq Prop a b` -/ +def isMorallyIff (e : Expr) : Bool := + let_expr Eq α _ _ := e | false + α.isProp + /-- Inserts `e` into the list of case-split candidates if applicable. -/ private def checkAndAddSplitCandidate (e : Expr) : GoalM Unit := do unless e.isApp do return () if (← getConfig).splitIte && (e.isIte || e.isDIte) then addSplitCandidate e return () + if isMorallyIff e then + addSplitCandidate e + return () if (← getConfig).splitMatch then if (← isMatcherApp e) then if let .reduced _ ← reduceMatcher? e then diff --git a/src/Lean/Meta/Tactic/Grind/Split.lean b/src/Lean/Meta/Tactic/Grind/Split.lean index e57a5270e404..4d242cf97486 100644 --- a/src/Lean/Meta/Tactic/Grind/Split.lean +++ b/src/Lean/Meta/Tactic/Grind/Split.lean @@ -39,6 +39,11 @@ private def checkCaseSplitStatus (e : Expr) : GoalM CaseSplitStatus := do return .ready else return .notReady + | Eq _ _ _ => + if (← isEqTrue e <||> isEqFalse e) then + return .ready + else + return .notReady | ite _ c _ _ _ => if (← isEqTrue c <||> isEqFalse c) then return .resolved @@ -94,6 +99,11 @@ private def mkCasesMajor (c : Expr) : GoalM Expr := do | And a b => return mkApp3 (mkConst ``Grind.or_of_and_eq_false) a b (← mkEqFalseProof c) | ite _ c _ _ _ => return mkEM c | dite _ c _ _ _ => return mkEM c + | Eq _ a b => + if (← isEqTrue c) then + return mkApp3 (mkConst ``Grind.of_eq_eq_true) a b (← mkEqTrueProof c) + else + return mkApp3 (mkConst ``Grind.of_eq_eq_false) a b (← mkEqFalseProof c) | _ => return mkApp2 (mkConst ``of_eq_true) c (← mkEqTrueProof c) /-- Introduces new hypotheses in each goal. -/ diff --git a/tests/lean/run/grind_implies.lean b/tests/lean/run/grind_implies.lean index 9e68cc2e5572..8a2ffb56f0f4 100644 --- a/tests/lean/run/grind_implies.lean +++ b/tests/lean/run/grind_implies.lean @@ -28,10 +28,13 @@ example (p q : Prop) : (p → q) → ¬q → ¬p := by grind /-- -info: [grind.internalize] p → q +info: [grind.internalize] (p → q) = r +[grind.internalize] Prop +[grind.internalize] p → q [grind.internalize] p [grind.internalize] q [grind.internalize] r +[grind.eqc] ((p → q) = r) = True [grind.eqc] (p → q) = r [grind.eqc] p = False [grind.eqc] (p → q) = True @@ -43,10 +46,13 @@ example (p q : Prop) : (p → q) = r → ¬p → r := by /-- -info: [grind.internalize] p → q +info: [grind.internalize] (p → q) = r +[grind.internalize] Prop +[grind.internalize] p → q [grind.internalize] p [grind.internalize] q [grind.internalize] r +[grind.eqc] ((p → q) = r) = True [grind.eqc] (p → q) = r [grind.eqc] q = True [grind.eqc] (p → q) = True @@ -57,10 +63,13 @@ example (p q : Prop) : (p → q) = r → q → r := by grind /-- -info: [grind.internalize] p → q +info: [grind.internalize] (p → q) = r +[grind.internalize] Prop +[grind.internalize] p → q [grind.internalize] p [grind.internalize] q [grind.internalize] r +[grind.eqc] ((p → q) = r) = True [grind.eqc] (p → q) = r [grind.eqc] q = False [grind.eqc] r = True @@ -72,10 +81,13 @@ example (p q : Prop) : (p → q) = r → ¬q → r → ¬p := by grind /-- -info: [grind.internalize] p → q +info: [grind.internalize] (p → q) = r +[grind.internalize] Prop +[grind.internalize] p → q [grind.internalize] p [grind.internalize] q [grind.internalize] r +[grind.eqc] ((p → q) = r) = True [grind.eqc] (p → q) = r [grind.eqc] q = False [grind.eqc] r = False diff --git a/tests/lean/run/grind_pre.lean b/tests/lean/run/grind_pre.lean index 1a2bbb378c0a..21e4360b9de6 100644 --- a/tests/lean/run/grind_pre.lean +++ b/tests/lean/run/grind_pre.lean @@ -86,13 +86,28 @@ example (p : Prop) (a b c : Nat) : p → a = 0 → a = b → h a = h c → a = c set_option trace.grind.debug.proof true /-- error: `grind` failed -case grind +case grind.1 +α : Type +a : α +p q r : Prop +h₁ : HEq p a +h₂ : HEq q a +h₃ : p = r +left : ¬p ∨ r +right : ¬r ∨ p +h : ¬r +⊢ False + +case grind.2 α : Type a : α p q r : Prop h₁ : HEq p a h₂ : HEq q a h₃ : p = r +left : ¬p ∨ r +right : ¬r ∨ p +h : p ⊢ False -/ #guard_msgs (error) in diff --git a/tests/lean/run/grind_t1.lean b/tests/lean/run/grind_t1.lean index b8bb7453c097..5b9c88e5210a 100644 --- a/tests/lean/run/grind_t1.lean +++ b/tests/lean/run/grind_t1.lean @@ -213,3 +213,15 @@ example (P Q R : α → Prop) (h₁ : ∀ x, Q x → P x) (h₂ : ∀ x, R x → example (w : Nat → Type) (h : ∀ n, Subsingleton (w n)) : True := by grind + +example {P1 P2 : Prop} : (P1 ∧ P2) ↔ (P2 ∧ P1) := by + grind + +example {P U V W : Prop} (h : P ↔ (V ↔ W)) (w : ¬ U ↔ V) : ¬ P ↔ (U ↔ W) := by + grind (splits := 6) + +example {P Q : Prop} (q : Q) (w : P = (P = ¬ Q)) : False := by + grind + +example (P Q : Prop) : (¬P → ¬Q) ↔ (Q → P) := by + grind From 463609157121d5a86c2bc11eb4eba8ec19683c15 Mon Sep 17 00:00:00 2001 From: Leonardo de Moura Date: Sat, 11 Jan 2025 19:27:13 -0800 Subject: [PATCH 076/100] fix: `simp_arith` (#6608) This PR fixes a bug in the `simp_arith` tactic. See new test. --- .../Meta/Tactic/LinearArith/Nat/Simp.lean | 30 +++++++++++-------- tests/lean/run/grind_t1.lean | 9 ++++++ 2 files changed, 27 insertions(+), 12 deletions(-) diff --git a/src/Lean/Meta/Tactic/LinearArith/Nat/Simp.lean b/src/Lean/Meta/Tactic/LinearArith/Nat/Simp.lean index b048c9c2871b..7e282174ecbc 100644 --- a/src/Lean/Meta/Tactic/LinearArith/Nat/Simp.lean +++ b/src/Lean/Meta/Tactic/LinearArith/Nat/Simp.lean @@ -50,18 +50,24 @@ def simpCnstr? (e : Expr) : MetaM (Option (Expr × Expr)) := do if let some arg := e.not? then let mut eNew? := none let mut thmName := Name.anonymous - if arg.isAppOfArity ``LE.le 4 then - eNew? := some (← mkLE (← mkAdd (arg.getArg! 3) (mkNatLit 1)) (arg.getArg! 2)) - thmName := ``Nat.not_le_eq - else if arg.isAppOfArity ``GE.ge 4 then - eNew? := some (← mkLE (← mkAdd (arg.getArg! 2) (mkNatLit 1)) (arg.getArg! 3)) - thmName := ``Nat.not_ge_eq - else if arg.isAppOfArity ``LT.lt 4 then - eNew? := some (← mkLE (arg.getArg! 3) (arg.getArg! 2)) - thmName := ``Nat.not_lt_eq - else if arg.isAppOfArity ``GT.gt 4 then - eNew? := some (← mkLE (arg.getArg! 2) (arg.getArg! 3)) - thmName := ``Nat.not_gt_eq + match_expr arg with + | LE.le α _ _ _ => + if α.isConstOf ``Nat then + eNew? := some (← mkLE (← mkAdd (arg.getArg! 3) (mkNatLit 1)) (arg.getArg! 2)) + thmName := ``Nat.not_le_eq + | GE.ge α _ _ _ => + if α.isConstOf ``Nat then + eNew? := some (← mkLE (← mkAdd (arg.getArg! 2) (mkNatLit 1)) (arg.getArg! 3)) + thmName := ``Nat.not_ge_eq + | LT.lt α _ _ _ => + if α.isConstOf ``Nat then + eNew? := some (← mkLE (arg.getArg! 3) (arg.getArg! 2)) + thmName := ``Nat.not_lt_eq + | GT.gt α _ _ _ => + if α.isConstOf ``Nat then + eNew? := some (← mkLE (arg.getArg! 2) (arg.getArg! 3)) + thmName := ``Nat.not_gt_eq + | _ => pure () if let some eNew := eNew? then let h₁ := mkApp2 (mkConst thmName) (arg.getArg! 2) (arg.getArg! 3) if let some (eNew', h₂) ← simpCnstrPos? eNew then diff --git a/tests/lean/run/grind_t1.lean b/tests/lean/run/grind_t1.lean index 5b9c88e5210a..71e6c4cf9610 100644 --- a/tests/lean/run/grind_t1.lean +++ b/tests/lean/run/grind_t1.lean @@ -225,3 +225,12 @@ example {P Q : Prop} (q : Q) (w : P = (P = ¬ Q)) : False := by example (P Q : Prop) : (¬P → ¬Q) ↔ (Q → P) := by grind + +example {α} (a b c : α) [LE α] : + ¬(¬a ≤ b ∧ a ≤ c ∨ ¬a ≤ c ∧ a ≤ b) ↔ a ≤ b ∧ a ≤ c ∨ ¬a ≤ c ∧ ¬a ≤ b := by + simp_arith -- should not fail + sorry + +example {α} (a b c : α) [LE α] : + ¬(¬a ≤ b ∧ a ≤ c ∨ ¬a ≤ c ∧ a ≤ b) ↔ a ≤ b ∧ a ≤ c ∨ ¬a ≤ c ∧ ¬a ≤ b := by + grind From 5119528d202c174fe340200a4b96d9f8761bdb26 Mon Sep 17 00:00:00 2001 From: Leonardo de Moura Date: Sat, 11 Jan 2025 20:21:04 -0800 Subject: [PATCH 077/100] feat: improve case-split heuristic used in `grind` (#6609) This PR improves the case-split heuristic used in grind, prioritizing case-splits with fewer cases. --- src/Init/Grind/Tactics.lean | 2 +- src/Lean/Meta/Basic.lean | 21 ++++++--- src/Lean/Meta/Tactic/Grind/Split.lean | 63 +++++++++++++++++---------- tests/lean/run/grind_t1.lean | 2 +- 4 files changed, 55 insertions(+), 33 deletions(-) diff --git a/src/Init/Grind/Tactics.lean b/src/Init/Grind/Tactics.lean index 951d3e2c0c77..5516e51ec597 100644 --- a/src/Init/Grind/Tactics.lean +++ b/src/Init/Grind/Tactics.lean @@ -25,7 +25,7 @@ Passed to `grind` using, for example, the `grind (config := { matchEqs := true } -/ structure Config where /-- Maximum number of case-splits in a proof search branch. It does not include splits performed during normalization. -/ - splits : Nat := 5 + splits : Nat := 8 /-- Maximum number of E-matching (aka heuristic theorem instantiation) rounds before each case split. -/ ematch : Nat := 5 /-- diff --git a/src/Lean/Meta/Basic.lean b/src/Lean/Meta/Basic.lean index d395b62478a6..0da4e5fde124 100644 --- a/src/Lean/Meta/Basic.lean +++ b/src/Lean/Meta/Basic.lean @@ -1964,15 +1964,22 @@ def sortFVarIds (fvarIds : Array FVarId) : MetaM (Array FVarId) := do end Methods -/-- Return `true` if `declName` is an inductive predicate. That is, `inductive` type in `Prop`. -/ -def isInductivePredicate (declName : Name) : MetaM Bool := do +/-- +Return `some info` if `declName` is an inductive predicate where `info : InductiveVal`. +That is, `inductive` type in `Prop`. +-/ +def isInductivePredicate? (declName : Name) : MetaM (Option InductiveVal) := do match (← getEnv).find? declName with - | some (.inductInfo { type := type, ..}) => - forallTelescopeReducing type fun _ type => do + | some (.inductInfo info) => + forallTelescopeReducing info.type fun _ type => do match (← whnfD type) with - | .sort u .. => return u == levelZero - | _ => return false - | _ => return false + | .sort u .. => if u == levelZero then return some info else return none + | _ => return none + | _ => return none + +/-- Return `true` if `declName` is an inductive predicate. That is, `inductive` type in `Prop`. -/ +def isInductivePredicate (declName : Name) : MetaM Bool := do + return (← isInductivePredicate? declName).isSome def isListLevelDefEqAux : List Level → List Level → MetaM Bool | [], [] => return true diff --git a/src/Lean/Meta/Tactic/Grind/Split.lean b/src/Lean/Meta/Tactic/Grind/Split.lean index 4d242cf97486..b7c567f90224 100644 --- a/src/Lean/Meta/Tactic/Grind/Split.lean +++ b/src/Lean/Meta/Tactic/Grind/Split.lean @@ -14,7 +14,7 @@ namespace Lean.Meta.Grind inductive CaseSplitStatus where | resolved | notReady - | ready + | ready (numCases : Nat) (isRec := false) deriving Inhabited, BEq private def checkCaseSplitStatus (e : Expr) : GoalM CaseSplitStatus := do @@ -24,7 +24,7 @@ private def checkCaseSplitStatus (e : Expr) : GoalM CaseSplitStatus := do if (← isEqTrue a <||> isEqTrue b) then return .resolved else - return .ready + return .ready 2 else if (← isEqFalse e) then return .resolved else @@ -36,60 +36,74 @@ private def checkCaseSplitStatus (e : Expr) : GoalM CaseSplitStatus := do if (← isEqFalse a <||> isEqFalse b) then return .resolved else - return .ready + return .ready 2 else return .notReady | Eq _ _ _ => if (← isEqTrue e <||> isEqFalse e) then - return .ready + return .ready 2 else return .notReady | ite _ c _ _ _ => if (← isEqTrue c <||> isEqFalse c) then return .resolved else - return .ready + return .ready 2 | dite _ c _ _ _ => if (← isEqTrue c <||> isEqFalse c) then return .resolved else - return .ready + return .ready 2 | _ => if (← isResolvedCaseSplit e) then trace[grind.debug.split] "split resolved: {e}" return .resolved - if (← isMatcherApp e) then - return .ready + if let some info := isMatcherAppCore? (← getEnv) e then + return .ready info.numAlts let .const declName .. := e.getAppFn | unreachable! - if (← isInductivePredicate declName <&&> isEqTrue e) then - return .ready + if let some info ← isInductivePredicate? declName then + if (← isEqTrue e) then + return .ready info.ctors.length info.isRec return .notReady +private inductive SplitCandidate where + | none + | some (c : Expr) (numCases : Nat) (isRec : Bool) + /-- Returns the next case-split to be performed. It uses a very simple heuristic. -/ -private def selectNextSplit? : GoalM (Option Expr) := do - if (← isInconsistent) then return none - if (← checkMaxCaseSplit) then return none - go (← get).splitCandidates none [] +private def selectNextSplit? : GoalM SplitCandidate := do + if (← isInconsistent) then return .none + if (← checkMaxCaseSplit) then return .none + go (← get).splitCandidates .none [] where - go (cs : List Expr) (c? : Option Expr) (cs' : List Expr) : GoalM (Option Expr) := do + go (cs : List Expr) (c? : SplitCandidate) (cs' : List Expr) : GoalM SplitCandidate := do match cs with | [] => modify fun s => { s with splitCandidates := cs'.reverse } - if c?.isSome then + if let .some _ numCases isRec := c? then + let numSplits := (← get).numSplits + -- We only increase the number of splits if there is more than one case or it is recursive. + let numSplits := if numCases > 1 || isRec then numSplits + 1 else numSplits -- Remark: we reset `numEmatch` after each case split. -- We should consider other strategies in the future. - modify fun s => { s with numSplits := s.numSplits + 1, numEmatch := 0 } + modify fun s => { s with numSplits, numEmatch := 0 } return c? | c::cs => match (← checkCaseSplitStatus c) with | .notReady => go cs c? (c::cs') | .resolved => go cs c? cs' - | .ready => + | .ready numCases isRec => match c? with - | none => go cs (some c) cs' - | some c' => - if (← getGeneration c) < (← getGeneration c') then - go cs (some c) (c'::cs') + | .none => go cs (.some c numCases isRec) cs' + | .some c' numCases' _ => + let isBetter : GoalM Bool := do + if numCases == 1 && !isRec && numCases' > 1 then + return true + if (← getGeneration c) < (← getGeneration c') then + return true + return numCases < numCases' + if (← isBetter) then + go cs (.some c numCases isRec) (c'::cs') else go cs c? (c::cs') @@ -118,9 +132,10 @@ and returns a new list of goals if successful. -/ def splitNext : GrindTactic := fun goal => do let (goals?, _) ← GoalM.run goal do - let some c ← selectNextSplit? + let .some c numCases isRec ← selectNextSplit? | return none let gen ← getGeneration c + let genNew := if numCases > 1 || isRec then gen+1 else gen trace_goal[grind.split] "{c}, generation: {gen}" let mvarIds ← if (← isMatcherApp c) then casesMatch (← get).mvarId c @@ -129,7 +144,7 @@ def splitNext : GrindTactic := fun goal => do cases (← get).mvarId major let goal ← get let goals := mvarIds.map fun mvarId => { goal with mvarId } - let goals ← introNewHyp goals [] (gen+1) + let goals ← introNewHyp goals [] genNew return some goals return goals? diff --git a/tests/lean/run/grind_t1.lean b/tests/lean/run/grind_t1.lean index 71e6c4cf9610..c17d87fcab0e 100644 --- a/tests/lean/run/grind_t1.lean +++ b/tests/lean/run/grind_t1.lean @@ -218,7 +218,7 @@ example {P1 P2 : Prop} : (P1 ∧ P2) ↔ (P2 ∧ P1) := by grind example {P U V W : Prop} (h : P ↔ (V ↔ W)) (w : ¬ U ↔ V) : ¬ P ↔ (U ↔ W) := by - grind (splits := 6) + grind example {P Q : Prop} (q : Q) (w : P = (P = ¬ Q)) : False := by grind From c5c1278315501ed0cfdb780e7123a0a6114ba5ad Mon Sep 17 00:00:00 2001 From: Leonardo de Moura Date: Sat, 11 Jan 2025 21:14:41 -0800 Subject: [PATCH 078/100] fix: bug in the `grind` propagator (#6610) This PR fixes a bug in the `grind` core module responsible for merging equivalence classes and propagating constraints. --- src/Lean/Meta/Tactic/Grind/Core.lean | 14 +++++++------- src/Lean/Meta/Tactic/Grind/Types.lean | 11 ++++++++++- tests/lean/run/grind_propagate_connectives.lean | 2 +- tests/lean/run/grind_t1.lean | 3 +++ 4 files changed, 21 insertions(+), 9 deletions(-) diff --git a/src/Lean/Meta/Tactic/Grind/Core.lean b/src/Lean/Meta/Tactic/Grind/Core.lean index 1b0962a229e0..a1e66ca190cc 100644 --- a/src/Lean/Meta/Tactic/Grind/Core.lean +++ b/src/Lean/Meta/Tactic/Grind/Core.lean @@ -141,6 +141,7 @@ where updateRoots lhs rhsNode.root trace_goal[grind.debug] "{← ppENodeRef lhs} new root {← ppENodeRef rhsNode.root}, {← ppENodeRef (← getRoot lhs)}" reinsertParents parents + propagateEqcDown lhs setENode lhsNode.root { (← getENode lhsRoot.self) with -- We must retrieve `lhsRoot` since it was updated. next := rhsRoot.next } @@ -158,14 +159,13 @@ where updateMT rhsRoot.self updateRoots (lhs : Expr) (rootNew : Expr) : GoalM Unit := do - let rec loop (e : Expr) : GoalM Unit := do - let n ← getENode e - setENode e { n with root := rootNew } + traverseEqc lhs fun n => + setENode n.self { n with root := rootNew } + + propagateEqcDown (lhs : Expr) : GoalM Unit := do + traverseEqc lhs fun n => unless (← isInconsistent) do - propagateDown e - if isSameExpr lhs n.next then return () - loop n.next - loop lhs + propagateDown n.self /-- Ensures collection of equations to be processed is empty. -/ private def resetNewEqs : GoalM Unit := diff --git a/src/Lean/Meta/Tactic/Grind/Types.lean b/src/Lean/Meta/Tactic/Grind/Types.lean index 50e072162a63..a6e5e28bb6ff 100644 --- a/src/Lean/Meta/Tactic/Grind/Types.lean +++ b/src/Lean/Meta/Tactic/Grind/Types.lean @@ -702,6 +702,15 @@ def getENodes : GoalM (Array ENode) := do let nodes := (← get).enodes.toArray.map (·.2) return nodes.qsort fun a b => a.idx < b.idx +/-- Executes `f` to each term in the equivalence class containing `e` -/ +@[inline] def traverseEqc (e : Expr) (f : ENode → GoalM Unit) : GoalM Unit := do + let mut curr := e + repeat + let n ← getENode curr + f n + if isSameExpr n.next e then return () + curr := n.next + def forEachENode (f : ENode → GoalM Unit) : GoalM Unit := do let nodes ← getENodes for n in nodes do @@ -714,7 +723,7 @@ def filterENodes (p : ENode → GoalM Bool) : GoalM (Array ENode) := do ref.modify (·.push n) ref.get -def forEachEqc (f : ENode → GoalM Unit) : GoalM Unit := do +def forEachEqcRoot (f : ENode → GoalM Unit) : GoalM Unit := do let nodes ← getENodes for n in nodes do if isSameExpr n.self n.root then diff --git a/tests/lean/run/grind_propagate_connectives.lean b/tests/lean/run/grind_propagate_connectives.lean index 6cb042ae1bba..b45206bb647f 100644 --- a/tests/lean/run/grind_propagate_connectives.lean +++ b/tests/lean/run/grind_propagate_connectives.lean @@ -6,7 +6,7 @@ def fallback : Fallback := do let falseExprs := (← filterENodes fun e => return e.self.isFVar && (← isEqFalse e.self)).toList.map (·.self) logInfo m!"true: {trueExprs}" logInfo m!"false: {falseExprs}" - forEachEqc fun n => do + forEachEqcRoot fun n => do unless (← isProp n.self) || (← isType n.self) || n.size == 1 do let eqc ← getEqc n.self logInfo eqc diff --git a/tests/lean/run/grind_t1.lean b/tests/lean/run/grind_t1.lean index c17d87fcab0e..3df71a83c520 100644 --- a/tests/lean/run/grind_t1.lean +++ b/tests/lean/run/grind_t1.lean @@ -234,3 +234,6 @@ example {α} (a b c : α) [LE α] : example {α} (a b c : α) [LE α] : ¬(¬a ≤ b ∧ a ≤ c ∨ ¬a ≤ c ∧ a ≤ b) ↔ a ≤ b ∧ a ≤ c ∨ ¬a ≤ c ∧ ¬a ≤ b := by grind + +example (x y : Bool) : ¬(x = true ↔ y = true) ↔ (¬(x = true) ↔ y = true) := by + grind From ce1ff03af049c2eeea0fd33adcb6143dedaf6220 Mon Sep 17 00:00:00 2001 From: Leonardo de Moura Date: Sat, 11 Jan 2025 21:30:41 -0800 Subject: [PATCH 079/100] fix: `checkParents` in `grind` (#6611) This PR fixes one of the sanity check tests used in `grind`. --- src/Lean/Meta/Tactic/Grind/Inv.lean | 5 ++++- 1 file changed, 4 insertions(+), 1 deletion(-) diff --git a/src/Lean/Meta/Tactic/Grind/Inv.lean b/src/Lean/Meta/Tactic/Grind/Inv.lean index 0b374a441f3a..aab448815f5f 100644 --- a/src/Lean/Meta/Tactic/Grind/Inv.lean +++ b/src/Lean/Meta/Tactic/Grind/Inv.lean @@ -58,9 +58,12 @@ private def checkParents (e : Expr) : GoalM Unit := do found := true break -- Recall that we have support for `Expr.forallE` propagation. See `ForallProp.lean`. - if let .forallE _ d _ _ := parent then + if let .forallE _ d b _ := parent then if (← checkChild d) then found := true + unless b.hasLooseBVars do + if (← checkChild b) then + found := true unless found do assert! (← checkChild parent.getAppFn) else From 8b1aabbb1ef46ffe02897dd14504202526b6a42e Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Sun, 12 Jan 2025 21:19:50 +1100 Subject: [PATCH 080/100] feat: lemmas about Array.append (#6612) This PR adds lemmas about `Array.append`, improving alignment with the `List` API. --- src/Init/Data/Array/Basic.lean | 3 +- src/Init/Data/Array/Lemmas.lean | 160 ++++++++++++++++++-------------- src/Init/Data/List/ToArray.lean | 2 +- 3 files changed, 93 insertions(+), 72 deletions(-) diff --git a/src/Init/Data/Array/Basic.lean b/src/Init/Data/Array/Basic.lean index 3ed6171ff685..275bbffe8dfe 100644 --- a/src/Init/Data/Array/Basic.lean +++ b/src/Init/Data/Array/Basic.lean @@ -244,8 +244,7 @@ def ofFn {n} (f : Fin n → α) : Array α := go 0 (mkEmpty n) where def range (n : Nat) : Array Nat := ofFn fun (i : Fin n) => i -def singleton (v : α) : Array α := - mkArray 1 v +@[inline] protected def singleton (v : α) : Array α := #[v] def back! [Inhabited α] (a : Array α) : α := a[a.size - 1]! diff --git a/src/Init/Data/Array/Lemmas.lean b/src/Init/Data/Array/Lemmas.lean index 7348c35f5338..3eee4ef43784 100644 --- a/src/Init/Data/Array/Lemmas.lean +++ b/src/Init/Data/Array/Lemmas.lean @@ -1506,9 +1506,99 @@ theorem filterMap_eq_push_iff {f : α → Option β} {l : Array α} {l' : Array /-! Content below this point has not yet been aligned with `List`. -/ +/-! ### singleton -/ + +@[simp] theorem singleton_def (v : α) : Array.singleton v = #[v] := rfl + +/-! ### append -/ + +@[simp] theorem toArray_eq_append_iff {xs : List α} {as bs : Array α} : + xs.toArray = as ++ bs ↔ xs = as.toList ++ bs.toList := by + cases as + cases bs + simp + +@[simp] theorem append_eq_toArray_iff {as bs : Array α} {xs : List α} : + as ++ bs = xs.toArray ↔ as.toList ++ bs.toList = xs := by + cases as + cases bs + simp + theorem singleton_eq_toArray_singleton (a : α) : #[a] = [a].toArray := rfl -@[simp] theorem singleton_def (v : α) : singleton v = #[v] := rfl +@[simp] theorem empty_append_fun : ((#[] : Array α) ++ ·) = id := by + funext ⟨l⟩ + simp + +@[simp] theorem mem_append {a : α} {s t : Array α} : a ∈ s ++ t ↔ a ∈ s ∨ a ∈ t := by + simp only [mem_def, toList_append, List.mem_append] + +theorem mem_append_left {a : α} {l₁ : Array α} (l₂ : Array α) (h : a ∈ l₁) : a ∈ l₁ ++ l₂ := + mem_append.2 (Or.inl h) + +theorem mem_append_right {a : α} (l₁ : Array α) {l₂ : Array α} (h : a ∈ l₂) : a ∈ l₁ ++ l₂ := + mem_append.2 (Or.inr h) + +@[simp] theorem size_append (as bs : Array α) : (as ++ bs).size = as.size + bs.size := by + simp only [size, toList_append, List.length_append] + +theorem empty_append (as : Array α) : #[] ++ as = as := by simp + +theorem append_empty (as : Array α) : as ++ #[] = as := by simp + +theorem getElem_append {as bs : Array α} (h : i < (as ++ bs).size) : + (as ++ bs)[i] = if h' : i < as.size then as[i] else bs[i - as.size]'(by simp at h; omega) := by + cases as; cases bs + simp [List.getElem_append] + +theorem getElem_append_left {as bs : Array α} {h : i < (as ++ bs).size} (hlt : i < as.size) : + (as ++ bs)[i] = as[i] := by + simp only [← getElem_toList] + have h' : i < (as.toList ++ bs.toList).length := by rwa [← length_toList, toList_append] at h + conv => rhs; rw [← List.getElem_append_left (bs := bs.toList) (h' := h')] + apply List.get_of_eq; rw [toList_append] + +theorem getElem_append_right {as bs : Array α} {h : i < (as ++ bs).size} (hle : as.size ≤ i) : + (as ++ bs)[i] = bs[i - as.size]'(Nat.sub_lt_left_of_lt_add hle (size_append .. ▸ h)) := by + simp only [← getElem_toList] + have h' : i < (as.toList ++ bs.toList).length := by rwa [← length_toList, toList_append] at h + conv => rhs; rw [← List.getElem_append_right (h₁ := hle) (h₂ := h')] + apply List.get_of_eq; rw [toList_append] + +theorem getElem?_append_left {as bs : Array α} {i : Nat} (hn : i < as.size) : + (as ++ bs)[i]? = as[i]? := by + have hn' : i < (as ++ bs).size := Nat.lt_of_lt_of_le hn <| + size_append .. ▸ Nat.le_add_right .. + simp_all [getElem?_eq_getElem, getElem_append] + +theorem getElem?_append_right {as bs : Array α} {i : Nat} (h : as.size ≤ i) : + (as ++ bs)[i]? = bs[i - as.size]? := by + cases as + cases bs + simp at h + simp [List.getElem?_append_right, h] + +theorem getElem?_append {as bs : Array α} {i : Nat} : + (as ++ bs)[i]? = if i < as.size then as[i]? else bs[i - as.size]? := by + split <;> rename_i h + · exact getElem?_append_left h + · exact getElem?_append_right (by simpa using h) + +/-- Variant of `getElem_append_left` useful for rewriting from the small array to the big array. -/ +theorem getElem_append_left' (l₂ : Array α) {l₁ : Array α} {i : Nat} (hi : i < l₁.size) : + l₁[i] = (l₁ ++ l₂)[i]'(by simpa using Nat.lt_add_right l₂.size hi) := by + rw [getElem_append_left] <;> simp + +/-- Variant of `getElem_append_right` useful for rewriting from the small array to the big array. -/ +theorem getElem_append_right' (l₁ : Array α) {l₂ : Array α} {i : Nat} (hi : i < l₂.size) : + l₂[i] = (l₁ ++ l₂)[i + l₁.size]'(by simpa [Nat.add_comm] using Nat.add_lt_add_left hi _) := by + rw [getElem_append_right] <;> simp [*, Nat.le_add_left] + +theorem getElem_of_append {l l₁ l₂ : Array α} (eq : l = l₁.push a ++ l₂) (h : l₁.size = i) : + l[i]'(eq ▸ h ▸ by simp_arith) = a := Option.some.inj <| by + rw [← getElem?_eq_getElem, eq, getElem?_append_left (by simp; omega), ← h] + simp + -- This is a duplicate of `List.toArray_toList`. -- It's confusing to guess which namespace this theorem should live in, @@ -2122,74 +2212,6 @@ theorem getElem?_modify {as : Array α} {i : Nat} {f : α → α} {j : Nat} : theorem size_empty : (#[] : Array α).size = 0 := rfl -/-! ### append -/ - -@[simp] theorem mem_append {a : α} {s t : Array α} : a ∈ s ++ t ↔ a ∈ s ∨ a ∈ t := by - simp only [mem_def, toList_append, List.mem_append] - -theorem mem_append_left {a : α} {l₁ : Array α} (l₂ : Array α) (h : a ∈ l₁) : a ∈ l₁ ++ l₂ := - mem_append.2 (Or.inl h) - -theorem mem_append_right {a : α} (l₁ : Array α) {l₂ : Array α} (h : a ∈ l₂) : a ∈ l₁ ++ l₂ := - mem_append.2 (Or.inr h) - -@[simp] theorem size_append (as bs : Array α) : (as ++ bs).size = as.size + bs.size := by - simp only [size, toList_append, List.length_append] - -theorem empty_append (as : Array α) : #[] ++ as = as := by simp - -theorem append_empty (as : Array α) : as ++ #[] = as := by simp - -theorem getElem_append {as bs : Array α} (h : i < (as ++ bs).size) : - (as ++ bs)[i] = if h' : i < as.size then as[i] else bs[i - as.size]'(by simp at h; omega) := by - cases as; cases bs - simp [List.getElem_append] - -theorem getElem_append_left {as bs : Array α} {h : i < (as ++ bs).size} (hlt : i < as.size) : - (as ++ bs)[i] = as[i] := by - simp only [← getElem_toList] - have h' : i < (as.toList ++ bs.toList).length := by rwa [← length_toList, toList_append] at h - conv => rhs; rw [← List.getElem_append_left (bs := bs.toList) (h' := h')] - apply List.get_of_eq; rw [toList_append] - -theorem getElem_append_right {as bs : Array α} {h : i < (as ++ bs).size} (hle : as.size ≤ i) : - (as ++ bs)[i] = bs[i - as.size]'(Nat.sub_lt_left_of_lt_add hle (size_append .. ▸ h)) := by - simp only [← getElem_toList] - have h' : i < (as.toList ++ bs.toList).length := by rwa [← length_toList, toList_append] at h - conv => rhs; rw [← List.getElem_append_right (h₁ := hle) (h₂ := h')] - apply List.get_of_eq; rw [toList_append] - -theorem getElem?_append_left {as bs : Array α} {i : Nat} (hn : i < as.size) : - (as ++ bs)[i]? = as[i]? := by - have hn' : i < (as ++ bs).size := Nat.lt_of_lt_of_le hn <| - size_append .. ▸ Nat.le_add_right .. - simp_all [getElem?_eq_getElem, getElem_append] - -theorem getElem?_append_right {as bs : Array α} {i : Nat} (h : as.size ≤ i) : - (as ++ bs)[i]? = bs[i - as.size]? := by - cases as - cases bs - simp at h - simp [List.getElem?_append_right, h] - -theorem getElem?_append {as bs : Array α} {i : Nat} : - (as ++ bs)[i]? = if i < as.size then as[i]? else bs[i - as.size]? := by - split <;> rename_i h - · exact getElem?_append_left h - · exact getElem?_append_right (by simpa using h) - -@[simp] theorem toArray_eq_append_iff {xs : List α} {as bs : Array α} : - xs.toArray = as ++ bs ↔ xs = as.toList ++ bs.toList := by - cases as - cases bs - simp - -@[simp] theorem append_eq_toArray_iff {as bs : Array α} {xs : List α} : - as ++ bs = xs.toArray ↔ as.toList ++ bs.toList = xs := by - cases as - cases bs - simp - /-! ### flatten -/ @[simp] theorem toList_flatten {l : Array (Array α)} : diff --git a/src/Init/Data/List/ToArray.lean b/src/Init/Data/List/ToArray.lean index ce2a7814c93d..6b07544c0604 100644 --- a/src/Init/Data/List/ToArray.lean +++ b/src/Init/Data/List/ToArray.lean @@ -46,7 +46,7 @@ theorem toArray_cons (a : α) (l : List α) : (a :: l).toArray = #[a] ++ l.toArr @[simp] theorem isEmpty_toArray (l : List α) : l.toArray.isEmpty = l.isEmpty := by cases l <;> simp [Array.isEmpty] -@[simp] theorem toArray_singleton (a : α) : (List.singleton a).toArray = singleton a := rfl +@[simp] theorem toArray_singleton (a : α) : (List.singleton a).toArray = Array.singleton a := rfl @[simp] theorem back!_toArray [Inhabited α] (l : List α) : l.toArray.back! = l.getLast! := by simp only [back!, size_toArray, Array.get!_eq_getElem!, getElem!_toArray, getLast!_eq_getElem!] From 541902564ba2f9bfbdd62e599d71847ffdb7398a Mon Sep 17 00:00:00 2001 From: Leonardo de Moura Date: Sun, 12 Jan 2025 07:40:36 -0800 Subject: [PATCH 081/100] feat: improve case split heuristic used in `grind` (#6613) This PR improves the case split heuristic used in the `grind` tactic, ensuring it now avoids unnecessary case-splits on `Iff`. --- src/Lean/Meta/Tactic/Grind/Split.lean | 79 ++++++++++++++++++--------- tests/lean/run/grind_t1.lean | 26 +++++++++ 2 files changed, 79 insertions(+), 26 deletions(-) diff --git a/src/Lean/Meta/Tactic/Grind/Split.lean b/src/Lean/Meta/Tactic/Grind/Split.lean index b7c567f90224..3937644300a7 100644 --- a/src/Lean/Meta/Tactic/Grind/Split.lean +++ b/src/Lean/Meta/Tactic/Grind/Split.lean @@ -17,43 +17,70 @@ inductive CaseSplitStatus where | ready (numCases : Nat) (isRec := false) deriving Inhabited, BEq -private def checkCaseSplitStatus (e : Expr) : GoalM CaseSplitStatus := do - match_expr e with - | Or a b => - if (← isEqTrue e) then - if (← isEqTrue a <||> isEqTrue b) then - return .resolved - else - return .ready 2 - else if (← isEqFalse e) then +/-- Given `c`, the condition of an `if-then-else`, check whether we need to case-split on the `if-then-else` or not -/ +private def checkIteCondStatus (c : Expr) : GoalM CaseSplitStatus := do + if (← isEqTrue c <||> isEqFalse c) then + return .resolved + else + return .ready 2 + +/-- +Given `e` of the form `a ∨ b`, check whether we are ready to case-split on `e`. +That is, `e` is `True`, but neither `a` nor `b` is `True`." +-/ +private def checkDisjunctStatus (e a b : Expr) : GoalM CaseSplitStatus := do + if (← isEqTrue e) then + if (← isEqTrue a <||> isEqTrue b) then return .resolved else - return .notReady - | And a b => - if (← isEqTrue e) then + return .ready 2 + else if (← isEqFalse e) then + return .resolved + else + return .notReady + +/-- +Given `e` of the form `a ∧ b`, check whether we are ready to case-split on `e`. +That is, `e` is `False`, but neither `a` nor `b` is `False`. + -/ +private def checkConjunctStatus (e a b : Expr) : GoalM CaseSplitStatus := do + if (← isEqTrue e) then + return .resolved + else if (← isEqFalse e) then + if (← isEqFalse a <||> isEqFalse b) then return .resolved - else if (← isEqFalse e) then - if (← isEqFalse a <||> isEqFalse b) then - return .resolved - else - return .ready 2 else - return .notReady - | Eq _ _ _ => - if (← isEqTrue e <||> isEqFalse e) then return .ready 2 - else - return .notReady - | ite _ c _ _ _ => - if (← isEqTrue c <||> isEqFalse c) then + else + return .notReady + +/-- +Given `e` of the form `@Eq Prop a b`, check whether we are ready to case-split on `e`. +There are two cases: +1- `e` is `True`, but neither both `a` and `b` are `True`, nor both `a` and `b` are `False`. +2- `e` is `False`, but neither `a` is `True` and `b` is `False`, nor `a` is `False` and `b` is `True`. +-/ +private def checkIffStatus (e a b : Expr) : GoalM CaseSplitStatus := do + if (← isEqTrue e) then + if (← (isEqTrue a <&&> isEqTrue b) <||> (isEqFalse a <&&> isEqFalse b)) then return .resolved else return .ready 2 - | dite _ c _ _ _ => - if (← isEqTrue c <||> isEqFalse c) then + else if (← isEqFalse e) then + if (← (isEqTrue a <&&> isEqFalse b) <||> (isEqFalse a <&&> isEqTrue b)) then return .resolved else return .ready 2 + else + return .notReady + +private def checkCaseSplitStatus (e : Expr) : GoalM CaseSplitStatus := do + match_expr e with + | Or a b => checkDisjunctStatus e a b + | And a b => checkConjunctStatus e a b + | Eq _ a b => checkIffStatus e a b + | ite _ c _ _ _ => checkIteCondStatus c + | dite _ c _ _ _ => checkIteCondStatus c | _ => if (← isResolvedCaseSplit e) then trace[grind.debug.split] "split resolved: {e}" diff --git a/tests/lean/run/grind_t1.lean b/tests/lean/run/grind_t1.lean index 3df71a83c520..1752266b8e31 100644 --- a/tests/lean/run/grind_t1.lean +++ b/tests/lean/run/grind_t1.lean @@ -237,3 +237,29 @@ example {α} (a b c : α) [LE α] : example (x y : Bool) : ¬(x = true ↔ y = true) ↔ (¬(x = true) ↔ y = true) := by grind + +/-- +error: `grind` failed +case grind +p q : Prop +a✝¹ : p = q +a✝ : p +⊢ False +-/ +#guard_msgs (error) in +set_option trace.grind.split true in +example (p q : Prop) : (p ↔ q) → p → False := by + grind -- should not split on (p ↔ q) + +/-- +error: `grind` failed +case grind +p q : Prop +a✝¹ : p = ¬q +a✝ : p +⊢ False +-/ +#guard_msgs (error) in +set_option trace.grind.split true in +example (p q : Prop) : ¬(p ↔ q) → p → False := by + grind -- should not split on (p ↔ q) From 349da6cae2dadcb345d667642958ed8b49ef197c Mon Sep 17 00:00:00 2001 From: Leonardo de Moura Date: Sun, 12 Jan 2025 08:51:09 -0800 Subject: [PATCH 082/100] feat: improve `[grind =]` attribute (#6614) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR improves the usability of the `[grind =]` attribute by automatically handling forbidden pattern symbols. For example, consider the following theorem tagged with this attribute: ``` getLast?_eq_some_iff {xs : List α} {a : α} : xs.getLast? = some a ↔ ∃ ys, xs = ys ++ [a] ``` Here, the selected pattern is `xs.getLast? = some a`, but `Eq` is a forbidden pattern symbol. Instead of producing an error, this function converts the pattern into a multi-pattern, allowing the attribute to be used conveniently. --- src/Init/Grind/Norm.lean | 6 ++++ src/Lean/Meta/Tactic/Grind/EMatchTheorem.lean | 35 ++++++++++++++++++- tests/lean/run/grind_eq_pattern.lean | 22 ++++++++++++ 3 files changed, 62 insertions(+), 1 deletion(-) create mode 100644 tests/lean/run/grind_eq_pattern.lean diff --git a/src/Init/Grind/Norm.lean b/src/Init/Grind/Norm.lean index 2673caa837c0..d9ba8051ed0d 100644 --- a/src/Init/Grind/Norm.lean +++ b/src/Init/Grind/Norm.lean @@ -46,6 +46,12 @@ attribute [grind_norm] not_false_eq_true theorem imp_eq (p q : Prop) : (p → q) = (¬ p ∨ q) := by by_cases p <;> by_cases q <;> simp [*] +@[grind_norm] theorem true_imp_eq (p : Prop) : (True → p) = p := by simp +@[grind_norm] theorem false_imp_eq (p : Prop) : (False → p) = True := by simp +@[grind_norm] theorem imp_true_eq (p : Prop) : (p → True) = True := by simp +@[grind_norm] theorem imp_false_eq (p : Prop) : (p → False) = ¬p := by simp +@[grind_norm] theorem imp_self_eq (p : Prop) : (p → p) = True := by simp + -- And @[grind_norm↓] theorem not_and (p q : Prop) : (¬(p ∧ q)) = (¬p ∨ ¬q) := by by_cases p <;> by_cases q <;> simp [*] diff --git a/src/Lean/Meta/Tactic/Grind/EMatchTheorem.lean b/src/Lean/Meta/Tactic/Grind/EMatchTheorem.lean index 3b427eca418b..acf370f11f36 100644 --- a/src/Lean/Meta/Tactic/Grind/EMatchTheorem.lean +++ b/src/Lean/Meta/Tactic/Grind/EMatchTheorem.lean @@ -170,11 +170,43 @@ private builtin_initialize ematchTheoremsExt : SimpleScopedEnvExtension EMatchTh initial := {} } +/-- +Symbols with built-in support in `grind` are unsuitable as pattern candidates for E-matching. +This is because `grind` performs normalization operations and uses specialized data structures +to implement these symbols, which may interfere with E-matching behavior. +-/ -- TODO: create attribute? private def forbiddenDeclNames := #[``Eq, ``HEq, ``Iff, ``And, ``Or, ``Not] private def isForbidden (declName : Name) := forbiddenDeclNames.contains declName +/-- +Auxiliary function to expand a pattern containing forbidden application symbols +into a multi-pattern. + +This function enhances the usability of the `[grind =]` attribute by automatically handling +forbidden pattern symbols. For example, consider the following theorem tagged with this attribute: +``` +getLast?_eq_some_iff {xs : List α} {a : α} : xs.getLast? = some a ↔ ∃ ys, xs = ys ++ [a] +``` +Here, the selected pattern is `xs.getLast? = some a`, but `Eq` is a forbidden pattern symbol. +Instead of producing an error, this function converts the pattern into a multi-pattern, +allowing the attribute to be used conveniently. + +The function recursively expands patterns with forbidden symbols by splitting them +into their sub-components. If the pattern does not contain forbidden symbols, +it is returned as-is. +-/ +partial def splitWhileForbidden (pat : Expr) : List Expr := + match_expr pat with + | Not p => splitWhileForbidden p + | And p₁ p₂ => splitWhileForbidden p₁ ++ splitWhileForbidden p₂ + | Or p₁ p₂ => splitWhileForbidden p₁ ++ splitWhileForbidden p₂ + | Eq _ lhs rhs => splitWhileForbidden lhs ++ splitWhileForbidden rhs + | Iff lhs rhs => splitWhileForbidden lhs ++ splitWhileForbidden rhs + | HEq _ lhs _ rhs => splitWhileForbidden lhs ++ splitWhileForbidden rhs + | _ => [pat] + private def dontCare := mkConst (Name.mkSimple "[grind_dontcare]") def mkGroundPattern (e : Expr) : Expr := @@ -468,7 +500,8 @@ def mkEMatchEqTheoremCore (origin : Origin) (levelParams : Array Name) (proof : | _ => throwError "invalid E-matching equality theorem, conclusion must be an equality{indentExpr type}" let pat := if useLhs then lhs else rhs let pat ← preprocessPattern pat normalizePattern - return (xs.size, [pat.abstract xs]) + let pats := splitWhileForbidden (pat.abstract xs) + return (xs.size, pats) mkEMatchTheoremCore origin levelParams numParams proof patterns /-- diff --git a/tests/lean/run/grind_eq_pattern.lean b/tests/lean/run/grind_eq_pattern.lean new file mode 100644 index 000000000000..179a6d64964f --- /dev/null +++ b/tests/lean/run/grind_eq_pattern.lean @@ -0,0 +1,22 @@ +attribute [grind] List.append_ne_nil_of_left_ne_nil +attribute [grind] List.append_ne_nil_of_right_ne_nil +/-- +info: [grind.ematch.pattern] List.getLast?_eq_some_iff: [@List.getLast? #2 #1, @some ? #0] +-/ +#guard_msgs (info) in +set_option trace.grind.ematch.pattern true in +attribute [grind =] List.getLast?_eq_some_iff + +/-- +info: [grind.assert] xs.getLast? = b? +[grind.assert] b? = some 10 +[grind.assert] xs = [] +[grind.assert] (xs.getLast? = some 10) = ∃ ys, xs = ys ++ [10] +[grind.assert] xs = w ++ [10] +[grind.assert] ¬w = [] → ¬w ++ [10] = [] +[grind.assert] ¬w ++ [10] = [] +-/ +#guard_msgs (info) in +set_option trace.grind.assert true in +example (xs : List Nat) : xs.getLast? = b? → b? = some 10 → xs ≠ [] := by + grind From 0da3624ec9ead9fa26ac1dd5d82964517e558319 Mon Sep 17 00:00:00 2001 From: Parth Shastri <31370288+cppio@users.noreply.github.com> Date: Sun, 12 Jan 2025 12:18:22 -0500 Subject: [PATCH 083/100] fix: allow dot idents to resolve to local names (#6602) This PR allows the dot ident notation to resolve to the current definition, or to any of the other definitions in the same mutual block. Existing code that uses dot ident notation may need to have `nonrec` added if the ident has the same name as the definition. Closes #6601 --- src/Lean/Elab/App.lean | 13 ++++++++----- src/Lean/Parser/Basic.lean | 4 ++-- src/Lean/Server/Requests.lean | 2 +- src/Std/Time/Date/Unit/Month.lean | 2 +- src/Std/Time/Date/Unit/Year.lean | 2 +- tests/lean/6601.lean | 26 ++++++++++++++++++++++++++ tests/lean/6601.lean.expected.out | 0 7 files changed, 39 insertions(+), 10 deletions(-) create mode 100644 tests/lean/6601.lean create mode 100644 tests/lean/6601.lean.expected.out diff --git a/src/Lean/Elab/App.lean b/src/Lean/Elab/App.lean index 23a744a9ff11..566fec4d63af 100644 --- a/src/Lean/Elab/App.lean +++ b/src/Lean/Elab/App.lean @@ -1474,7 +1474,7 @@ where | field::fields, false => .fieldName field field.getId.getString! none fIdent :: toLVals fields false /-- Resolve `(.$id:ident)` using the expected type to infer namespace. -/ -private partial def resolveDotName (id : Syntax) (expectedType? : Option Expr) : TermElabM Name := do +private partial def resolveDotName (id : Syntax) (expectedType? : Option Expr) : TermElabM Expr := do tryPostponeIfNoneOrMVar expectedType? let some expectedType := expectedType? | throwError "invalid dotted identifier notation, expected type must be known" @@ -1489,7 +1489,7 @@ where withForallBody body k else k body - go (resultType : Expr) (expectedType : Expr) (previousExceptions : Array Exception) : TermElabM Name := do + go (resultType : Expr) (expectedType : Expr) (previousExceptions : Array Exception) : TermElabM Expr := do let resultType ← instantiateMVars resultType let resultTypeFn := resultType.cleanupAnnotations.getAppFn try @@ -1497,9 +1497,12 @@ where let .const declName .. := resultTypeFn.cleanupAnnotations | throwError "invalid dotted identifier notation, expected type is not of the form (... → C ...) where C is a constant{indentExpr expectedType}" let idNew := declName ++ id.getId.eraseMacroScopes - unless (← getEnv).contains idNew do + if (← getEnv).contains idNew then + mkConst idNew + else if let some (fvar, []) ← resolveLocalName idNew then + return fvar + else throwError "invalid dotted identifier notation, unknown identifier `{idNew}` from expected type{indentExpr expectedType}" - return idNew catch | ex@(.error ..) => match (← unfoldDefinition? resultType) with @@ -1548,7 +1551,7 @@ private partial def elabAppFn (f : Syntax) (lvals : List LVal) (namedArgs : Arra | `(_) => throwError "placeholders '_' cannot be used where a function is expected" | `(.$id:ident) => addCompletionInfo <| CompletionInfo.dotId f id.getId (← getLCtx) expectedType? - let fConst ← mkConst (← resolveDotName id expectedType?) + let fConst ← resolveDotName id expectedType? let s ← observing do -- Use (force := true) because we want to record the result of .ident resolution even in patterns let fConst ← addTermInfo f fConst expectedType? (force := true) diff --git a/src/Lean/Parser/Basic.lean b/src/Lean/Parser/Basic.lean index 7649171e922b..857b47f40396 100644 --- a/src/Lean/Parser/Basic.lean +++ b/src/Lean/Parser/Basic.lean @@ -1583,8 +1583,8 @@ namespace TokenMap def insert (map : TokenMap α) (k : Name) (v : α) : TokenMap α := match map.find? k with - | none => .insert map k [v] - | some vs => .insert map k (v::vs) + | none => RBMap.insert map k [v] + | some vs => RBMap.insert map k (v::vs) instance : Inhabited (TokenMap α) where default := RBMap.empty diff --git a/src/Lean/Server/Requests.lean b/src/Lean/Server/Requests.lean index e57fd94ce477..ea5f62f0b645 100644 --- a/src/Lean/Server/Requests.lean +++ b/src/Lean/Server/Requests.lean @@ -97,7 +97,7 @@ abbrev RequestT m := ReaderT RequestContext <| ExceptT RequestError m /-- Workers execute request handlers in this monad. -/ abbrev RequestM := ReaderT RequestContext <| EIO RequestError -abbrev RequestTask.pure (a : α) : RequestTask α := .pure (.ok a) +abbrev RequestTask.pure (a : α) : RequestTask α := Task.pure (.ok a) instance : MonadLift IO RequestM where monadLift x := do diff --git a/src/Std/Time/Date/Unit/Month.lean b/src/Std/Time/Date/Unit/Month.lean index f700fbcd65ec..79a8eb8ff594 100644 --- a/src/Std/Time/Date/Unit/Month.lean +++ b/src/Std/Time/Date/Unit/Month.lean @@ -73,7 +73,7 @@ Creates an `Offset` from a natural number. -/ @[inline] def ofNat (data : Nat) : Offset := - .ofNat data + Int.ofNat data /-- Creates an `Offset` from an integer. diff --git a/src/Std/Time/Date/Unit/Year.lean b/src/Std/Time/Date/Unit/Year.lean index 94ac4d67ea49..eef7c2010ca2 100644 --- a/src/Std/Time/Date/Unit/Year.lean +++ b/src/Std/Time/Date/Unit/Year.lean @@ -56,7 +56,7 @@ Creates an `Offset` from a natural number. -/ @[inline] def ofNat (data : Nat) : Offset := - .ofNat data + Int.ofNat data /-- Creates an `Offset` from an integer. diff --git a/tests/lean/6601.lean b/tests/lean/6601.lean new file mode 100644 index 000000000000..4121d4fd5876 --- /dev/null +++ b/tests/lean/6601.lean @@ -0,0 +1,26 @@ +/-! Dot ident notation should resolve mutual definitions -/ + +mutual + +inductive Even + | zero + | succ (n : Odd) + deriving Inhabited + +inductive Odd + | succ (n : Even) + deriving Inhabited + +end + +mutual + +def Even.ofNat : Nat → Even + | 0 => .zero + | n + 1 => .succ (.ofNat n) + +def Odd.ofNat : Nat → Odd + | 0 => panic! "invalid input" + | n + 1 => .succ (.ofNat n) + +end diff --git a/tests/lean/6601.lean.expected.out b/tests/lean/6601.lean.expected.out new file mode 100644 index 000000000000..e69de29bb2d1 From c7939cfb03c78820b069ba9c931900d44b11f58e Mon Sep 17 00:00:00 2001 From: Leonardo de Moura Date: Sun, 12 Jan 2025 12:38:39 -0800 Subject: [PATCH 084/100] feat: offset constraints support for the `grind` tactic (#6603) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR implements support for offset constraints in the `grind` tactic. Several features are still missing, such as constraint propagation and support for offset equalities, but `grind` can already solve examples like the following: ```lean example (a b c : Nat) : a ≤ b → b + 2 ≤ c → a + 1 ≤ c := by grind example (a b c : Nat) : a ≤ b → b ≤ c → a ≤ c := by grind example (a b c : Nat) : a + 1 ≤ b → b + 1 ≤ c → a + 2 ≤ c := by grind example (a b c : Nat) : a + 1 ≤ b → b + 1 ≤ c → a + 1 ≤ c := by grind example (a b c : Nat) : a + 1 ≤ b → b ≤ c + 2 → a ≤ c + 1 := by grind example (a b c : Nat) : a + 2 ≤ b → b ≤ c + 2 → a ≤ c := by grind ``` --------- Co-authored-by: Kim Morrison --- src/Init/Grind/Offset.lean | 197 +++----------- src/Lean/Data/AssocList.lean | 8 +- src/Lean/Meta/AppBuilder.lean | 12 +- src/Lean/Meta/Tactic/FVarSubst.lean | 2 +- src/Lean/Meta/Tactic/Grind.lean | 7 + src/Lean/Meta/Tactic/Grind/Arith.lean | 10 + .../Meta/Tactic/Grind/Arith/Internalize.lean | 33 +++ src/Lean/Meta/Tactic/Grind/Arith/Inv.lean | 14 + src/Lean/Meta/Tactic/Grind/Arith/Main.lean | 34 +++ src/Lean/Meta/Tactic/Grind/Arith/Offset.lean | 173 ++++++++++++ .../Meta/Tactic/Grind/Arith/ProofUtil.lean | 89 ++++++ src/Lean/Meta/Tactic/Grind/Arith/Types.lean | 58 ++++ src/Lean/Meta/Tactic/Grind/Arith/Util.lean | 89 ++++++ src/Lean/Meta/Tactic/Grind/ENodeKey.lean | 30 ++ src/Lean/Meta/Tactic/Grind/Internalize.lean | 2 + src/Lean/Meta/Tactic/Grind/Inv.lean | 2 + src/Lean/Meta/Tactic/Grind/Types.lean | 23 +- tests/lean/run/grind_offset_cnstr.lean | 257 ++++++++++++++++++ tests/lean/run/grind_offset_ineq_thms.lean | 31 --- 19 files changed, 859 insertions(+), 212 deletions(-) create mode 100644 src/Lean/Meta/Tactic/Grind/Arith.lean create mode 100644 src/Lean/Meta/Tactic/Grind/Arith/Internalize.lean create mode 100644 src/Lean/Meta/Tactic/Grind/Arith/Inv.lean create mode 100644 src/Lean/Meta/Tactic/Grind/Arith/Main.lean create mode 100644 src/Lean/Meta/Tactic/Grind/Arith/Offset.lean create mode 100644 src/Lean/Meta/Tactic/Grind/Arith/ProofUtil.lean create mode 100644 src/Lean/Meta/Tactic/Grind/Arith/Types.lean create mode 100644 src/Lean/Meta/Tactic/Grind/Arith/Util.lean create mode 100644 src/Lean/Meta/Tactic/Grind/ENodeKey.lean create mode 100644 tests/lean/run/grind_offset_cnstr.lean delete mode 100644 tests/lean/run/grind_offset_ineq_thms.lean diff --git a/src/Init/Grind/Offset.lean b/src/Init/Grind/Offset.lean index 705ca31f3b73..08d17714aefa 100644 --- a/src/Init/Grind/Offset.lean +++ b/src/Init/Grind/Offset.lean @@ -7,159 +7,44 @@ prelude import Init.Core import Init.Omega -namespace Lean.Grind.Offset - -abbrev Var := Nat -abbrev Context := Lean.RArray Nat - -def fixedVar := 100000000 -- Any big number should work here - -def Var.denote (ctx : Context) (v : Var) : Nat := - bif v == fixedVar then 1 else ctx.get v - -structure Cnstr where - x : Var - y : Var - k : Nat := 0 - l : Bool := true - deriving Repr, DecidableEq, Inhabited - -def Cnstr.denote (c : Cnstr) (ctx : Context) : Prop := - if c.l then - c.x.denote ctx + c.k ≤ c.y.denote ctx - else - c.x.denote ctx ≤ c.y.denote ctx + c.k - -def trivialCnstr : Cnstr := { x := 0, y := 0, k := 0, l := true } - -@[simp] theorem denote_trivial (ctx : Context) : trivialCnstr.denote ctx := by - simp [Cnstr.denote, trivialCnstr] - -def Cnstr.trans (c₁ c₂ : Cnstr) : Cnstr := - if c₁.y = c₂.x then - let { x, k := k₁, l := l₁, .. } := c₁ - let { y, k := k₂, l := l₂, .. } := c₂ - match l₁, l₂ with - | false, false => - { x, y, k := k₁ + k₂, l := false } - | false, true => - if k₁ < k₂ then - { x, y, k := k₂ - k₁, l := true } - else - { x, y, k := k₁ - k₂, l := false } - | true, false => - if k₁ < k₂ then - { x, y, k := k₂ - k₁, l := false } - else - { x, y, k := k₁ - k₂, l := true } - | true, true => - { x, y, k := k₁ + k₂, l := true } - else - trivialCnstr - -@[simp] theorem Cnstr.denote_trans_easy (ctx : Context) (c₁ c₂ : Cnstr) (h : c₁.y ≠ c₂.x) : (c₁.trans c₂).denote ctx := by - simp [*, Cnstr.trans] - -@[simp] theorem Cnstr.denote_trans (ctx : Context) (c₁ c₂ : Cnstr) : c₁.denote ctx → c₂.denote ctx → (c₁.trans c₂).denote ctx := by - by_cases c₁.y = c₂.x - case neg => simp [*] - simp [trans, *] - let { x, k := k₁, l := l₁, .. } := c₁ - let { y, k := k₂, l := l₂, .. } := c₂ - simp_all; split - · simp [denote]; omega - · split <;> simp [denote] <;> omega - · split <;> simp [denote] <;> omega - · simp [denote]; omega - -def Cnstr.isTrivial (c : Cnstr) : Bool := c.x == c.y && c.k == 0 - -theorem Cnstr.of_isTrivial (ctx : Context) (c : Cnstr) : c.isTrivial = true → c.denote ctx := by - cases c; simp [isTrivial]; intros; simp [*, denote] - -def Cnstr.isFalse (c : Cnstr) : Bool := c.x == c.y && c.k != 0 && c.l == true - -theorem Cnstr.of_isFalse (ctx : Context) {c : Cnstr} : c.isFalse = true → ¬c.denote ctx := by - cases c; simp [isFalse]; intros; simp [*, denote]; omega - -def Cnstrs := List Cnstr - -def Cnstrs.denoteAnd' (ctx : Context) (c₁ : Cnstr) (c₂ : Cnstrs) : Prop := - match c₂ with - | [] => c₁.denote ctx - | c::cs => c₁.denote ctx ∧ Cnstrs.denoteAnd' ctx c cs - -theorem Cnstrs.denote'_trans (ctx : Context) (c₁ c : Cnstr) (cs : Cnstrs) : c₁.denote ctx → denoteAnd' ctx c cs → denoteAnd' ctx (c₁.trans c) cs := by - induction cs - next => simp [denoteAnd', *]; apply Cnstr.denote_trans - next c cs ih => simp [denoteAnd']; intros; simp [*] - -def Cnstrs.trans' (c₁ : Cnstr) (c₂ : Cnstrs) : Cnstr := - match c₂ with - | [] => c₁ - | c::c₂ => trans' (c₁.trans c) c₂ - -@[simp] theorem Cnstrs.denote'_trans' (ctx : Context) (c₁ : Cnstr) (c₂ : Cnstrs) : denoteAnd' ctx c₁ c₂ → (trans' c₁ c₂).denote ctx := by - induction c₂ generalizing c₁ - next => intros; simp_all [trans', denoteAnd'] - next c cs ih => simp [denoteAnd']; intros; simp [trans']; apply ih; apply denote'_trans <;> assumption - -def Cnstrs.denoteAnd (ctx : Context) (c : Cnstrs) : Prop := - match c with - | [] => True - | c::cs => denoteAnd' ctx c cs - -def Cnstrs.trans (c : Cnstrs) : Cnstr := - match c with - | [] => trivialCnstr - | c::cs => trans' c cs - -theorem Cnstrs.of_denoteAnd_trans {ctx : Context} {c : Cnstrs} : c.denoteAnd ctx → c.trans.denote ctx := by - cases c <;> simp [*, trans, denoteAnd] <;> intros <;> simp [*] - -def Cnstrs.isFalse (c : Cnstrs) : Bool := - c.trans.isFalse - -theorem Cnstrs.unsat' (ctx : Context) (c : Cnstrs) : c.isFalse = true → ¬ c.denoteAnd ctx := by - simp [isFalse]; intro h₁ h₂ - have := of_denoteAnd_trans h₂ - have := Cnstr.of_isFalse ctx h₁ - contradiction - -/-- `denote ctx [c_1, ..., c_n] C` is `c_1.denote ctx → ... → c_n.denote ctx → C` -/ -def Cnstrs.denote (ctx : Context) (cs : Cnstrs) (C : Prop) : Prop := - match cs with - | [] => C - | c::cs => c.denote ctx → denote ctx cs C - -theorem Cnstrs.not_denoteAnd'_eq (ctx : Context) (c : Cnstr) (cs : Cnstrs) (C : Prop) : (denoteAnd' ctx c cs → C) = denote ctx (c::cs) C := by - simp [denote] - induction cs generalizing c - next => simp [denoteAnd', denote] - next c' cs ih => - simp [denoteAnd', denote, *] - -theorem Cnstrs.not_denoteAnd_eq (ctx : Context) (cs : Cnstrs) (C : Prop) : (denoteAnd ctx cs → C) = denote ctx cs C := by - cases cs - next => simp [denoteAnd, denote] - next c cs => apply not_denoteAnd'_eq - -def Cnstr.isImpliedBy (cs : Cnstrs) (c : Cnstr) : Bool := - cs.trans == c - -/-! Main theorems used by `grind`. -/ - -/-- Auxiliary theorem used by `grind` to prove that a system of offset inequalities is unsatisfiable. -/ -theorem Cnstrs.unsat (ctx : Context) (cs : Cnstrs) : cs.isFalse = true → cs.denote ctx False := by - intro h - rw [← not_denoteAnd_eq] - apply unsat' - assumption - -/-- Auxiliary theorem used by `grind` to prove an implied offset inequality. -/ -theorem Cnstrs.imp (ctx : Context) (cs : Cnstrs) (c : Cnstr) (h : c.isImpliedBy cs = true) : cs.denote ctx (c.denote ctx) := by - rw [← eq_of_beq h] - rw [← not_denoteAnd_eq] - apply of_denoteAnd_trans - -end Lean.Grind.Offset +namespace Lean.Grind +def isLt (x y : Nat) : Bool := x < y + +theorem Nat.le_ro (u w v k : Nat) : u ≤ w → w ≤ v + k → u ≤ v + k := by + omega +theorem Nat.le_lo (u w v k : Nat) : u ≤ w → w + k ≤ v → u + k ≤ v := by + omega +theorem Nat.lo_le (u w v k : Nat) : u + k ≤ w → w ≤ v → u + k ≤ v := by + omega +theorem Nat.lo_lo (u w v k₁ k₂ : Nat) : u + k₁ ≤ w → w + k₂ ≤ v → u + (k₁ + k₂) ≤ v := by + omega +theorem Nat.lo_ro_1 (u w v k₁ k₂ : Nat) : isLt k₂ k₁ = true → u + k₁ ≤ w → w ≤ v + k₂ → u + (k₁ - k₂) ≤ v := by + simp [isLt]; omega +theorem Nat.lo_ro_2 (u w v k₁ k₂ : Nat) : u + k₁ ≤ w → w ≤ v + k₂ → u ≤ v + (k₂ - k₁) := by + omega +theorem Nat.ro_le (u w v k : Nat) : u ≤ w + k → w ≤ v → u ≤ v + k := by + omega +theorem Nat.ro_lo_1 (u w v k₁ k₂ : Nat) : u ≤ w + k₁ → w + k₂ ≤ v → u ≤ v + (k₁ - k₂) := by + omega +theorem Nat.ro_lo_2 (u w v k₁ k₂ : Nat) : isLt k₁ k₂ = true → u ≤ w + k₁ → w + k₂ ≤ v → u + (k₂ - k₁) ≤ v := by + simp [isLt]; omega +theorem Nat.ro_ro (u w v k₁ k₂ : Nat) : u ≤ w + k₁ → w ≤ v + k₂ → u ≤ v + (k₁ + k₂) := by + omega + +theorem Nat.of_le_eq_false (u v : Nat) : ((u ≤ v) = False) → v + 1 ≤ u := by + simp; omega +theorem Nat.of_lo_eq_false_1 (u v : Nat) : ((u + 1 ≤ v) = False) → v ≤ u := by + simp; omega +theorem Nat.of_lo_eq_false (u v k : Nat) : ((u + k ≤ v) = False) → v ≤ u + (k-1) := by + simp; omega +theorem Nat.of_ro_eq_false (u v k : Nat) : ((u ≤ v + k) = False) → v + (k+1) ≤ u := by + simp; omega + +theorem Nat.unsat_le_lo (u v k : Nat) : isLt 0 k = true → u ≤ v → v + k ≤ u → False := by + simp [isLt]; omega +theorem Nat.unsat_lo_lo (u v k₁ k₂ : Nat) : isLt 0 (k₁+k₂) = true → u + k₁ ≤ v → v + k₂ ≤ u → False := by + simp [isLt]; omega +theorem Nat.unsat_lo_ro (u v k₁ k₂ : Nat) : isLt k₂ k₁ = true → u + k₁ ≤ v → v ≤ u + k₂ → False := by + simp [isLt]; omega + +end Lean.Grind diff --git a/src/Lean/Data/AssocList.lean b/src/Lean/Data/AssocList.lean index 91953d5b052b..4ad955e28c3a 100644 --- a/src/Lean/Data/AssocList.lean +++ b/src/Lean/Data/AssocList.lean @@ -24,7 +24,7 @@ abbrev empty : AssocList α β := instance : EmptyCollection (AssocList α β) := ⟨empty⟩ -abbrev insert (m : AssocList α β) (k : α) (v : β) : AssocList α β := +abbrev insertNew (m : AssocList α β) (k : α) (v : β) : AssocList α β := m.cons k v def isEmpty : AssocList α β → Bool @@ -77,6 +77,12 @@ def replace [BEq α] (a : α) (b : β) : AssocList α β → AssocList α β | true => cons a b es | false => cons k v (replace a b es) +def insert [BEq α] (m : AssocList α β) (k : α) (v : β) : AssocList α β := + if m.contains k then + m.replace k v + else + m.insertNew k v + def erase [BEq α] (a : α) : AssocList α β → AssocList α β | nil => nil | cons k v es => match k == a with diff --git a/src/Lean/Meta/AppBuilder.lean b/src/Lean/Meta/AppBuilder.lean index bc6a668a6b57..657543d536c5 100644 --- a/src/Lean/Meta/AppBuilder.lean +++ b/src/Lean/Meta/AppBuilder.lean @@ -569,12 +569,16 @@ def mkLetBodyCongr (a h : Expr) : MetaM Expr := mkAppM ``let_body_congr #[a, h] /-- Return `of_eq_true h` -/ -def mkOfEqTrue (h : Expr) : MetaM Expr := - mkAppM ``of_eq_true #[h] +def mkOfEqTrue (h : Expr) : MetaM Expr := do + match_expr h with + | eq_true _ h => return h + | _ => mkAppM ``of_eq_true #[h] /-- Return `eq_true h` -/ -def mkEqTrue (h : Expr) : MetaM Expr := - mkAppM ``eq_true #[h] +def mkEqTrue (h : Expr) : MetaM Expr := do + match_expr h with + | of_eq_true _ h => return h + | _ => return mkApp2 (mkConst ``eq_true) (← inferType h) h /-- Return `eq_false h` diff --git a/src/Lean/Meta/Tactic/FVarSubst.lean b/src/Lean/Meta/Tactic/FVarSubst.lean index 281ed4f2e753..5c72a63de231 100644 --- a/src/Lean/Meta/Tactic/FVarSubst.lean +++ b/src/Lean/Meta/Tactic/FVarSubst.lean @@ -35,7 +35,7 @@ def insert (s : FVarSubst) (fvarId : FVarId) (v : Expr) : FVarSubst := if s.contains fvarId then s else let map := s.map.mapVal fun e => e.replaceFVarId fvarId v; - { map := map.insert fvarId v } + { map := map.insertNew fvarId v } def erase (s : FVarSubst) (fvarId : FVarId) : FVarSubst := { map := s.map.erase fvarId } diff --git a/src/Lean/Meta/Tactic/Grind.lean b/src/Lean/Meta/Tactic/Grind.lean index 47f4b4911817..4982d12bf8c2 100644 --- a/src/Lean/Meta/Tactic/Grind.lean +++ b/src/Lean/Meta/Tactic/Grind.lean @@ -24,6 +24,7 @@ import Lean.Meta.Tactic.Grind.EMatchTheorem import Lean.Meta.Tactic.Grind.EMatch import Lean.Meta.Tactic.Grind.Main import Lean.Meta.Tactic.Grind.CasesMatch +import Lean.Meta.Tactic.Grind.Arith namespace Lean @@ -42,6 +43,10 @@ builtin_initialize registerTraceClass `grind.simp builtin_initialize registerTraceClass `grind.split builtin_initialize registerTraceClass `grind.split.candidate builtin_initialize registerTraceClass `grind.split.resolved +builtin_initialize registerTraceClass `grind.offset +builtin_initialize registerTraceClass `grind.offset.dist +builtin_initialize registerTraceClass `grind.offset.internalize +builtin_initialize registerTraceClass `grind.offset.internalize.term (inherited := true) /-! Trace options for `grind` developers -/ builtin_initialize registerTraceClass `grind.debug @@ -54,4 +59,6 @@ builtin_initialize registerTraceClass `grind.debug.final builtin_initialize registerTraceClass `grind.debug.forallPropagator builtin_initialize registerTraceClass `grind.debug.split builtin_initialize registerTraceClass `grind.debug.canon +builtin_initialize registerTraceClass `grind.debug.offset +builtin_initialize registerTraceClass `grind.debug.offset.proof end Lean diff --git a/src/Lean/Meta/Tactic/Grind/Arith.lean b/src/Lean/Meta/Tactic/Grind/Arith.lean new file mode 100644 index 000000000000..ad62fd648e18 --- /dev/null +++ b/src/Lean/Meta/Tactic/Grind/Arith.lean @@ -0,0 +1,10 @@ +/- +Copyright (c) 2025 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.Meta.Tactic.Grind.Arith.Util +import Lean.Meta.Tactic.Grind.Arith.Types +import Lean.Meta.Tactic.Grind.Arith.Offset +import Lean.Meta.Tactic.Grind.Arith.Main diff --git a/src/Lean/Meta/Tactic/Grind/Arith/Internalize.lean b/src/Lean/Meta/Tactic/Grind/Arith/Internalize.lean new file mode 100644 index 000000000000..96fac2547f80 --- /dev/null +++ b/src/Lean/Meta/Tactic/Grind/Arith/Internalize.lean @@ -0,0 +1,33 @@ +/- +Copyright (c) 2025 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.Meta.Tactic.Grind.Arith.Offset + +namespace Lean.Meta.Grind.Arith + +namespace Offset + +def internalizeTerm (_e : Expr) (_a : Expr) (_k : Nat) : GoalM Unit := do + -- TODO + return () + +def internalizeCnstr (e : Expr) : GoalM Unit := do + let some c := isNatOffsetCnstr? e | return () + let c := { c with + a := (← mkNode c.a) + b := (← mkNode c.b) + } + trace[grind.offset.internalize] "{e} ↦ {c}" + modify' fun s => { s with + cnstrs := s.cnstrs.insert { expr := e } c + } + +end Offset + +def internalize (e : Expr) : GoalM Unit := do + Offset.internalizeCnstr e + +end Lean.Meta.Grind.Arith diff --git a/src/Lean/Meta/Tactic/Grind/Arith/Inv.lean b/src/Lean/Meta/Tactic/Grind/Arith/Inv.lean new file mode 100644 index 000000000000..359dceef99ef --- /dev/null +++ b/src/Lean/Meta/Tactic/Grind/Arith/Inv.lean @@ -0,0 +1,14 @@ +/- +Copyright (c) 2025 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.Meta.Tactic.Grind.Arith.Offset + +namespace Lean.Meta.Grind.Arith + +def checkInvariants : GoalM Unit := + Offset.checkInvariants + +end Lean.Meta.Grind.Arith diff --git a/src/Lean/Meta/Tactic/Grind/Arith/Main.lean b/src/Lean/Meta/Tactic/Grind/Arith/Main.lean new file mode 100644 index 000000000000..d46388c3ae0b --- /dev/null +++ b/src/Lean/Meta/Tactic/Grind/Arith/Main.lean @@ -0,0 +1,34 @@ +/- +Copyright (c) 2025 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.Meta.Tactic.Grind.PropagatorAttr +import Lean.Meta.Tactic.Grind.Arith.Offset + +namespace Lean.Meta.Grind.Arith + +namespace Offset +def isCnstr? (e : Expr) : GoalM (Option (Cnstr NodeId)) := + return (← get).arith.offset.cnstrs.find? { expr := e } + +def assertTrue (c : Cnstr NodeId) (p : Expr) : GoalM Unit := do + addEdge c.a c.b c.k (← mkOfEqTrue p) + +def assertFalse (c : Cnstr NodeId) (p : Expr) : GoalM Unit := do + let p := mkOfNegEqFalse (← get').nodes c p + let c := c.neg + addEdge c.a c.b c.k p + +end Offset + +builtin_grind_propagator propagateLE ↓LE.le := fun e => do + if (← isEqTrue e) then + if let some c ← Offset.isCnstr? e then + Offset.assertTrue c (← mkEqTrueProof e) + if (← isEqFalse e) then + if let some c ← Offset.isCnstr? e then + Offset.assertFalse c (← mkEqFalseProof e) + +end Lean.Meta.Grind.Arith diff --git a/src/Lean/Meta/Tactic/Grind/Arith/Offset.lean b/src/Lean/Meta/Tactic/Grind/Arith/Offset.lean new file mode 100644 index 000000000000..6a87df5c367b --- /dev/null +++ b/src/Lean/Meta/Tactic/Grind/Arith/Offset.lean @@ -0,0 +1,173 @@ +/- +Copyright (c) 2025 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 Init.Grind.Offset +import Lean.Meta.Tactic.Grind.Types +import Lean.Meta.Tactic.Grind.Arith.ProofUtil + +namespace Lean.Meta.Grind.Arith.Offset +/-! +This module implements a decision procedure for offset constraints of the form: +``` +x + k ≤ y +x ≤ y + k +``` +where `k` is a numeral. +Each constraint is represented as an edge in a weighted graph. +The constraint `x + k ≤ y` is represented as a negative edge. +The shortest path between two nodes in the graph corresponds to an implied inequality. +When adding a new edge, the state is considered unsatisfiable if the new edge creates a negative cycle. +An incremental Floyd-Warshall algorithm is used to find the shortest paths between all nodes. +This module can also handle offset equalities of the form `x + k = y` by representing them with two edges: +``` +x + k ≤ y +y ≤ x + k +``` +The main advantage of this module over a full linear integer arithmetic procedure is +its ability to efficiently detect all implied equalities and inequalities. +-/ + +def get' : GoalM State := do + return (← get).arith.offset + +@[inline] def modify' (f : State → State) : GoalM Unit := do + modify fun s => { s with arith.offset := f s.arith.offset } + +def mkNode (expr : Expr) : GoalM NodeId := do + if let some nodeId := (← get').nodeMap.find? { expr } then + return nodeId + let nodeId : NodeId := (← get').nodes.size + trace[grind.offset.internalize.term] "{expr} ↦ #{nodeId}" + modify' fun s => { s with + nodes := s.nodes.push expr + nodeMap := s.nodeMap.insert { expr } nodeId + sources := s.sources.push {} + targets := s.targets.push {} + proofs := s.proofs.push {} + } + return nodeId + +private def getDist? (u v : NodeId) : GoalM (Option Int) := do + return (← get').targets[u]!.find? v + +private def getProof? (u v : NodeId) : GoalM (Option ProofInfo) := do + return (← get').proofs[u]!.find? v + +partial def extractProof (u v : NodeId) : GoalM Expr := do + go (← getProof? u v).get! +where + go (p : ProofInfo) : GoalM Expr := do + if u == p.w then + return p.proof + else + let p' := (← getProof? u p.w).get! + go (mkTrans (← get').nodes p' p v) + +private def setUnsat (u v : NodeId) (kuv : Int) (huv : Expr) (kvu : Int) : GoalM Unit := do + assert! kuv + kvu < 0 + let hvu ← extractProof v u + let u := (← get').nodes[u]! + let v := (← get').nodes[v]! + if kuv == 0 then + assert! kvu < 0 + closeGoal (mkApp6 (mkConst ``Grind.Nat.unsat_le_lo) u v (toExpr (-kvu).toNat) rfl_true huv hvu) + else if kvu == 0 then + assert! kuv < 0 + closeGoal (mkApp6 (mkConst ``Grind.Nat.unsat_le_lo) v u (toExpr (-kuv).toNat) rfl_true hvu huv) + else if kuv < 0 then + if kvu > 0 then + closeGoal (mkApp7 (mkConst ``Grind.Nat.unsat_lo_ro) u v (toExpr (-kuv).toNat) (toExpr kvu.toNat) rfl_true huv hvu) + else + assert! kvu < 0 + closeGoal (mkApp7 (mkConst ``Grind.Nat.unsat_lo_lo) u v (toExpr (-kuv).toNat) (toExpr (-kvu).toNat) rfl_true huv hvu) + else + assert! kuv > 0 && kvu < 0 + closeGoal (mkApp7 (mkConst ``Grind.Nat.unsat_lo_ro) v u (toExpr (-kvu).toNat) (toExpr kuv.toNat) rfl_true hvu huv) + +private def setDist (u v : NodeId) (k : Int) : GoalM Unit := do + trace[grind.offset.dist] "{({ a := u, b := v, k : Cnstr NodeId})}" + modify' fun s => { s with + targets := s.targets.modify u fun es => es.insert v k + sources := s.sources.modify v fun es => es.insert u k + } + +private def setProof (u v : NodeId) (p : ProofInfo) : GoalM Unit := do + modify' fun s => { s with + proofs := s.proofs.modify u fun es => es.insert v p + } + +@[inline] +private def forEachSourceOf (u : NodeId) (f : NodeId → Int → GoalM Unit) : GoalM Unit := do + (← get').sources[u]!.forM f + +@[inline] +private def forEachTargetOf (u : NodeId) (f : NodeId → Int → GoalM Unit) : GoalM Unit := do + (← get').targets[u]!.forM f + +private def isShorter (u v : NodeId) (k : Int) : GoalM Bool := do + if let some k' ← getDist? u v then + return k < k' + else + return true + +private def updateIfShorter (u v : NodeId) (k : Int) (w : NodeId) : GoalM Unit := do + if (← isShorter u v k) then + setDist u v k + setProof u v (← getProof? w v).get! + +def addEdge (u : NodeId) (v : NodeId) (k : Int) (p : Expr) : GoalM Unit := do + if (← isInconsistent) then return () + if let some k' ← getDist? v u then + if k'+k < 0 then + setUnsat u v k p k' + return () + if (← isShorter u v k) then + setDist u v k + setProof u v { w := u, k, proof := p } + update +where + update : GoalM Unit := do + forEachTargetOf v fun j k₂ => do + /- Check whether new path: `u -(k)-> v -(k₂)-> j` is shorter -/ + updateIfShorter u j (k+k₂) v + forEachSourceOf u fun i k₁ => do + /- Check whether new path: `i -(k₁)-> u -(k)-> v` is shorter -/ + updateIfShorter i v (k₁+k) u + forEachTargetOf v fun j k₂ => do + /- Check whether new path: `i -(k₁)-> u -(k)-> v -(k₂) -> j` is shorter -/ + updateIfShorter i j (k₁+k+k₂) v + +def traceDists : GoalM Unit := do + let s ← get' + for u in [:s.targets.size], es in s.targets.toArray do + for (v, k) in es do + trace[grind.offset.dist] "#{u} -({k})-> #{v}" + +def Cnstr.toExpr (c : Cnstr NodeId) : GoalM Expr := do + let a := (← get').nodes[c.a]! + let b := (← get').nodes[c.b]! + let mk := if c.le then mkNatLE else mkNatEq + if c.k == 0 then + return mk a b + else if c.k < 0 then + return mk (mkNatAdd a (Lean.toExpr ((-c.k).toNat))) b + else + return mk a (mkNatAdd b (Lean.toExpr c.k.toNat)) + +def checkInvariants : GoalM Unit := do + let s ← get' + for u in [:s.targets.size], es in s.targets.toArray do + for (v, k) in es do + let c : Cnstr NodeId := { a := u, b := v, k } + trace[grind.debug.offset] "{c}" + let p ← extractProof u v + trace[grind.debug.offset.proof] "{p} : {← inferType p}" + check p + unless (← withDefault <| isDefEq (← inferType p) (← Cnstr.toExpr c)) do + trace[grind.debug.offset.proof] "failed: {← inferType p} =?= {← Cnstr.toExpr c}" + unreachable! + +end Lean.Meta.Grind.Arith.Offset diff --git a/src/Lean/Meta/Tactic/Grind/Arith/ProofUtil.lean b/src/Lean/Meta/Tactic/Grind/Arith/ProofUtil.lean new file mode 100644 index 000000000000..e281f2b813ab --- /dev/null +++ b/src/Lean/Meta/Tactic/Grind/Arith/ProofUtil.lean @@ -0,0 +1,89 @@ +/- +Copyright (c) 2025 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 Init.Grind.Offset +import Lean.Meta.Tactic.Grind.Types + +namespace Lean.Meta.Grind.Arith + +/-! +Helper functions for constructing proof terms in the arithmetic procedures. +-/ + +namespace Offset + +/-- `Eq.refl true` -/ +def rfl_true : Expr := mkApp2 (mkConst ``Eq.refl [levelOne]) (mkConst ``Bool) (mkConst ``Bool.true) + +open Lean.Grind in +/-- +Assume `pi₁` is `{ w := u, k := k₁, proof := p₁ }` and `pi₂` is `{ w := w, k := k₂, proof := p₂ }` +`p₁` is the proof for edge `u -(k₁) → w` and `p₂` the proof for edge `w -(k₂)-> v`. +Then, this function returns a proof for edge `u -(k₁+k₂) -> v`. +-/ +def mkTrans (nodes : PArray Expr) (pi₁ : ProofInfo) (pi₂ : ProofInfo) (v : NodeId) : ProofInfo := + let { w := u, k := k₁, proof := p₁ } := pi₁ + let { w, k := k₂, proof := p₂ } := pi₂ + let u := nodes[u]! + let w := nodes[w]! + let v := nodes[v]! + let p := if k₁ == 0 then + if k₂ == 0 then + -- u ≤ w, w ≤ v + mkApp5 (mkConst ``Nat.le_trans) u w v p₁ p₂ + else if k₂ > 0 then + -- u ≤ v, w ≤ v + k₂ + mkApp6 (mkConst ``Nat.le_ro) u w v (toExpr k₂.toNat) p₁ p₂ + else + let k₂ := - k₂ + -- u ≤ w, w + k₂ ≤ v + mkApp6 (mkConst ``Nat.le_lo) u w v (toExpr k₂.toNat) p₁ p₂ + else if k₁ < 0 then + let k₁ := -k₁ + if k₂ == 0 then + mkApp6 (mkConst ``Nat.lo_le) u w v (toExpr k₁.toNat) p₁ p₂ + else if k₂ < 0 then + let k₂ := -k₂ + mkApp7 (mkConst ``Nat.lo_lo) u w v (toExpr k₁.toNat) (toExpr k₂.toNat) p₁ p₂ + else + let ke₁ := toExpr k₁.toNat + let ke₂ := toExpr k₂.toNat + if k₁ > k₂ then + mkApp8 (mkConst ``Nat.lo_ro_1) u w v ke₁ ke₂ rfl_true p₁ p₂ + else + mkApp7 (mkConst ``Nat.lo_ro_2) u w v ke₁ ke₂ p₁ p₂ + else + let ke₁ := toExpr k₁.toNat + if k₂ == 0 then + mkApp6 (mkConst ``Nat.ro_le) u w v ke₁ p₁ p₂ + else if k₂ < 0 then + let k₂ := -k₂ + let ke₂ := toExpr k₂.toNat + if k₂ > k₁ then + mkApp8 (mkConst ``Nat.ro_lo_2) u w v ke₁ ke₂ rfl_true p₁ p₂ + else + mkApp7 (mkConst ``Nat.ro_lo_1) u w v ke₁ ke₂ p₁ p₂ + else + let ke₂ := toExpr k₂.toNat + mkApp7 (mkConst ``Nat.ro_ro) u w v ke₁ ke₂ p₁ p₂ + { w := pi₁.w, k := k₁+k₂, proof := p } + +open Lean.Grind in +def mkOfNegEqFalse (nodes : PArray Expr) (c : Cnstr NodeId) (h : Expr) : Expr := + let u := nodes[c.a]! + let v := nodes[c.b]! + if c.k == 0 then + mkApp3 (mkConst ``Nat.of_le_eq_false) u v h + else if c.k == -1 && c.le then + mkApp3 (mkConst ``Nat.of_lo_eq_false_1) u v h + else if c.k < 0 then + mkApp4 (mkConst ``Nat.of_lo_eq_false) u v (toExpr (-c.k).toNat) h + else + mkApp4 (mkConst ``Nat.of_ro_eq_false) u v (toExpr c.k.toNat) h + +end Offset + +end Lean.Meta.Grind.Arith diff --git a/src/Lean/Meta/Tactic/Grind/Arith/Types.lean b/src/Lean/Meta/Tactic/Grind/Arith/Types.lean new file mode 100644 index 000000000000..d4f6d5497b48 --- /dev/null +++ b/src/Lean/Meta/Tactic/Grind/Arith/Types.lean @@ -0,0 +1,58 @@ +/- +Copyright (c) 2025 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.Data.PersistentArray +import Lean.Meta.Tactic.Grind.ENodeKey +import Lean.Meta.Tactic.Grind.Arith.Util + +namespace Lean.Meta.Grind.Arith + +namespace Offset + +abbrev NodeId := Nat + +instance : ToMessageData (Offset.Cnstr NodeId) where + toMessageData c := Offset.toMessageData (α := NodeId) (inst := { toMessageData n := m!"#{n}" }) c + +/-- Auxiliary structure used for proof extraction. -/ +structure ProofInfo where + w : NodeId + k : Int + proof : Expr + deriving Inhabited + +/-- State of the constraint offset procedure. -/ +structure State where + nodes : PArray Expr := {} + nodeMap : PHashMap ENodeKey NodeId := {} + cnstrs : PHashMap ENodeKey (Cnstr NodeId) := {} + /-- + For each node with id `u`, `sources[u]` contains + pairs `(v, k)` s.t. there is a path from `v` to `u` with weight `k`. + -/ + sources : PArray (AssocList NodeId Int) := {} + /-- + For each node with id `u`, `targets[u]` contains + pairs `(v, k)` s.t. there is a path from `u` to `v` with weight `k`. + -/ + targets : PArray (AssocList NodeId Int) := {} + /-- + Proof reconstruction information. For each node with id `u`, `proofs[u]` contains + pairs `(v, { w, proof })` s.t. there is a path from `u` to `v`, and + `w` is the penultimate node in the path, and `proof` is the justification for + the last edge. + -/ + proofs : PArray (AssocList NodeId ProofInfo) := {} + deriving Inhabited + +end Offset + +/-- State for the arithmetic procedures. -/ +structure State where + offset : Offset.State := {} + deriving Inhabited + +end Lean.Meta.Grind.Arith diff --git a/src/Lean/Meta/Tactic/Grind/Arith/Util.lean b/src/Lean/Meta/Tactic/Grind/Arith/Util.lean new file mode 100644 index 000000000000..13e0ae6de46f --- /dev/null +++ b/src/Lean/Meta/Tactic/Grind/Arith/Util.lean @@ -0,0 +1,89 @@ +/- +Copyright (c) 2025 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.Expr +import Lean.Message + +namespace Lean.Meta.Grind.Arith + +/-- Returns `true` if `e` is of the form `Nat` -/ +def isNatType (e : Expr) : Bool := + e.isConstOf ``Nat + +/-- Returns `true` if `e` is of the form `@instHAdd Nat instAddNat` -/ +def isInstAddNat (e : Expr) : Bool := + let_expr instHAdd a b := e | false + isNatType a && b.isConstOf ``instAddNat + +/-- Returns `true` if `e` is `instLENat` -/ +def isInstLENat (e : Expr) : Bool := + e.isConstOf ``instLENat + +/-- +Returns `some (a, b)` if `e` is of the form +``` +@HAdd.hAdd Nat Nat Nat (instHAdd Nat instAddNat) a b +``` +-/ +def isNatAdd? (e : Expr) : Option (Expr × Expr) := + let_expr HAdd.hAdd _ _ _ i a b := e | none + if isInstAddNat i then some (a, b) else none + +/-- Returns `some k` if `e` `@OfNat.ofNat Nat _ (instOfNatNat k)` -/ +def isNatNum? (e : Expr) : Option Nat := Id.run do + let_expr OfNat.ofNat _ _ inst := e | none + let_expr instOfNatNat k := inst | none + let .lit (.natVal k) := k | none + some k + +/-- Returns `some (a, k)` if `e` is of the form `a + k`. -/ +def isNatOffset? (e : Expr) : Option (Expr × Nat) := Id.run do + let some (a, b) := isNatAdd? e | none + let some k := isNatNum? b | none + some (a, k) + +/-- An offset constraint. -/ +structure Offset.Cnstr (α : Type) where + a : α + b : α + k : Int := 0 + le : Bool := true + deriving Inhabited + +def Offset.Cnstr.neg : Cnstr α → Cnstr α + | { a, b, k, le } => { a := b, b := a, le, k := -k - 1 } + +example (c : Offset.Cnstr α) : c.neg.neg = c := by + cases c; simp [Offset.Cnstr.neg]; omega + +def Offset.toMessageData [inst : ToMessageData α] (c : Offset.Cnstr α) : MessageData := + match c.k, c.le with + | .ofNat 0, true => m!"{c.a} ≤ {c.b}" + | .ofNat 0, false => m!"{c.a} = {c.b}" + | .ofNat k, true => m!"{c.a} ≤ {c.b} + {k}" + | .ofNat k, false => m!"{c.a} = {c.b} + {k}" + | .negSucc k, true => m!"{c.a} + {k + 1} ≤ {c.b}" + | .negSucc k, false => m!"{c.a} + {k + 1} = {c.b}" + +instance : ToMessageData (Offset.Cnstr Expr) where + toMessageData c := Offset.toMessageData c + +/-- Returns `some cnstr` if `e` is offset constraint. -/ +def isNatOffsetCnstr? (e : Expr) : Option (Offset.Cnstr Expr) := + match_expr e with + | LE.le _ inst a b => if isInstLENat inst then go a b true else none + | Eq α a b => if isNatType α then go a b false else none + | _ => none +where + go (a b : Expr) (le : Bool) := + if let some (a, k) := isNatOffset? a then + some { a, k := - k, b, le } + else if let some (b, k) := isNatOffset? b then + some { a, b, k := k, le } + else + some { a, b, le } + +end Lean.Meta.Grind.Arith diff --git a/src/Lean/Meta/Tactic/Grind/ENodeKey.lean b/src/Lean/Meta/Tactic/Grind/ENodeKey.lean new file mode 100644 index 000000000000..5a8c34ba9d89 --- /dev/null +++ b/src/Lean/Meta/Tactic/Grind/ENodeKey.lean @@ -0,0 +1,30 @@ +/- +Copyright (c) 2025 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.Expr + +namespace Lean.Meta.Grind + +@[inline] def isSameExpr (a b : Expr) : Bool := + -- It is safe to use pointer equality because we hashcons all expressions + -- inserted into the E-graph + unsafe ptrEq a b + +/-- +Key for the `ENodeMap` and `ParentMap` map. +We use pointer addresses and rely on the fact all internalized expressions +have been hash-consed, i.e., we have applied `shareCommon`. +-/ +structure ENodeKey where + expr : Expr + +instance : Hashable ENodeKey where + hash k := unsafe (ptrAddrUnsafe k.expr).toUInt64 + +instance : BEq ENodeKey where + beq k₁ k₂ := isSameExpr k₁.expr k₂.expr + +end Lean.Meta.Grind diff --git a/src/Lean/Meta/Tactic/Grind/Internalize.lean b/src/Lean/Meta/Tactic/Grind/Internalize.lean index 5fa727302445..b29e3ad7574c 100644 --- a/src/Lean/Meta/Tactic/Grind/Internalize.lean +++ b/src/Lean/Meta/Tactic/Grind/Internalize.lean @@ -11,6 +11,7 @@ import Lean.Meta.Match.MatcherInfo import Lean.Meta.Match.MatchEqsExt import Lean.Meta.Tactic.Grind.Types import Lean.Meta.Tactic.Grind.Util +import Lean.Meta.Tactic.Grind.Arith.Internalize namespace Lean.Meta.Grind @@ -196,6 +197,7 @@ partial def internalize (e : Expr) (generation : Nat) : GoalM Unit := do mkENode e generation addCongrTable e updateAppMap e + Arith.internalize e propagateUp e end diff --git a/src/Lean/Meta/Tactic/Grind/Inv.lean b/src/Lean/Meta/Tactic/Grind/Inv.lean index aab448815f5f..cfebc1653ff1 100644 --- a/src/Lean/Meta/Tactic/Grind/Inv.lean +++ b/src/Lean/Meta/Tactic/Grind/Inv.lean @@ -6,6 +6,7 @@ Authors: Leonardo de Moura prelude import Lean.Meta.Tactic.Grind.Types import Lean.Meta.Tactic.Grind.Proof +import Lean.Meta.Tactic.Grind.Arith.Inv namespace Lean.Meta.Grind @@ -103,6 +104,7 @@ def checkInvariants (expensive := false) : GoalM Unit := do checkEqc node if expensive then checkPtrEqImpliesStructEq + Arith.checkInvariants if expensive && grind.debug.proofs.get (← getOptions) then checkProofs diff --git a/src/Lean/Meta/Tactic/Grind/Types.lean b/src/Lean/Meta/Tactic/Grind/Types.lean index a6e5e28bb6ff..aa68c05ca735 100644 --- a/src/Lean/Meta/Tactic/Grind/Types.lean +++ b/src/Lean/Meta/Tactic/Grind/Types.lean @@ -13,17 +13,14 @@ import Lean.Meta.CongrTheorems import Lean.Meta.AbstractNestedProofs import Lean.Meta.Tactic.Simp.Types import Lean.Meta.Tactic.Util +import Lean.Meta.Tactic.Grind.ENodeKey import Lean.Meta.Tactic.Grind.Canon import Lean.Meta.Tactic.Grind.Attr +import Lean.Meta.Tactic.Grind.Arith.Types import Lean.Meta.Tactic.Grind.EMatchTheorem namespace Lean.Meta.Grind -@[inline] def isSameExpr (a b : Expr) : Bool := - -- It is safe to use pointer equality because we hashcons all expressions - -- inserted into the E-graph - unsafe ptrEq a b - /-- We use this auxiliary constant to mark delayed congruence proofs. -/ def congrPlaceholderProof := mkConst (Name.mkSimple "[congruence]") @@ -224,20 +221,6 @@ structure NewEq where proof : Expr isHEq : Bool -/-- -Key for the `ENodeMap` and `ParentMap` map. -We use pointer addresses and rely on the fact all internalized expressions -have been hash-consed, i.e., we have applied `shareCommon`. --/ -private structure ENodeKey where - expr : Expr - -instance : Hashable ENodeKey where - hash k := unsafe (ptrAddrUnsafe k.expr).toUInt64 - -instance : BEq ENodeKey where - beq k₁ k₂ := isSameExpr k₁.expr k₂.expr - abbrev ENodeMap := PHashMap ENodeKey ENode /-- @@ -368,6 +351,8 @@ structure Goal where gmt : Nat := 0 /-- Next unique index for creating ENodes -/ nextIdx : Nat := 0 + /-- State of arithmetic procedures -/ + arith : Arith.State := {} /-- Active theorems that we have performed ematching at least once. -/ thms : PArray EMatchTheorem := {} /-- Active theorems that we have not performed any round of ematching yet. -/ diff --git a/tests/lean/run/grind_offset_cnstr.lean b/tests/lean/run/grind_offset_cnstr.lean new file mode 100644 index 000000000000..cbad141607c1 --- /dev/null +++ b/tests/lean/run/grind_offset_cnstr.lean @@ -0,0 +1,257 @@ +set_option grind.debug true + +/-- +info: [grind.offset.internalize.term] a1 ↦ #0 +[grind.offset.internalize.term] a2 ↦ #1 +[grind.offset.internalize] a1 + 1 ≤ a2 ↦ #0 + 1 ≤ #1 +[grind.offset.internalize.term] a3 ↦ #2 +[grind.offset.internalize] a2 ≤ a3 + 2 ↦ #1 ≤ #2 + 2 +[grind.offset.internalize.term] a4 ↦ #3 +[grind.offset.internalize] a3 ≤ a4 ↦ #2 ≤ #3 +-/ +#guard_msgs (info) in +set_option trace.grind.offset.internalize true in +example (a1 a2 a3) : + a1 + 1 ≤ a2 → a2 ≤ a3 + 2 → a3 ≤ a4 → False := by + fail_if_success grind + sorry + +/-- +info: [grind.offset.internalize.term] a1 ↦ #0 +[grind.offset.internalize.term] a2 ↦ #1 +[grind.offset.dist] #0 + 1 ≤ #1 +[grind.offset.internalize.term] a3 ↦ #2 +[grind.offset.dist] #1 ≤ #2 +[grind.offset.dist] #0 + 1 ≤ #2 +-/ +#guard_msgs (info) in +set_option trace.grind.offset.internalize.term true in +set_option trace.grind.offset.dist true in +example (a1 a2 a3 : Nat) : + a1 + 1 ≤ a2 → a2 ≤ a3 → False := by + fail_if_success grind + sorry + + +/-- +info: [grind.offset.internalize.term] a1 ↦ #0 +[grind.offset.internalize.term] a2 ↦ #1 +[grind.offset.dist] #0 + 1 ≤ #1 +[grind.offset.internalize.term] a3 ↦ #2 +[grind.offset.dist] #1 + 2 ≤ #2 +[grind.offset.dist] #0 + 3 ≤ #2 +-/ +#guard_msgs (info) in +set_option trace.grind.offset.internalize.term true in +set_option trace.grind.offset.dist true in +example (a1 a2 a3 : Nat) : + a1 + 1 ≤ a2 → a2 + 2 ≤ a3 → False := by + fail_if_success grind + sorry + +/-- +info: [grind.offset.internalize.term] a1 ↦ #0 +[grind.offset.internalize.term] a2 ↦ #1 +[grind.offset.dist] #0 + 1 ≤ #1 +[grind.offset.internalize.term] a3 ↦ #2 +[grind.offset.dist] #1 ≤ #2 + 2 +[grind.offset.dist] #0 ≤ #2 + 1 +-/ +#guard_msgs (info) in +set_option trace.grind.offset.internalize.term true in +set_option trace.grind.offset.dist true in +example (a1 a2 a3 : Nat) : + a1 + 1 ≤ a2 → a2 ≤ a3 + 2 → False := by + fail_if_success grind + sorry + +/-- +info: [grind.offset.internalize.term] a1 ↦ #0 +[grind.offset.internalize.term] a2 ↦ #1 +[grind.offset.dist] #0 ≤ #1 +[grind.offset.internalize.term] a3 ↦ #2 +[grind.offset.dist] #1 ≤ #2 +[grind.offset.dist] #0 ≤ #2 +-/ +#guard_msgs (info) in +set_option trace.grind.offset.internalize.term true in +set_option trace.grind.offset.dist true in +example (a1 a2 a3 : Nat) : + a1 ≤ a2 → a2 ≤ a3 → False := by + fail_if_success grind + sorry + +/-- +info: [grind.offset.internalize.term] a1 ↦ #0 +[grind.offset.internalize.term] a2 ↦ #1 +[grind.offset.dist] #0 ≤ #1 +[grind.offset.internalize.term] a3 ↦ #2 +[grind.offset.dist] #1 + 2 ≤ #2 +[grind.offset.dist] #0 + 2 ≤ #2 +-/ +#guard_msgs (info) in +set_option trace.grind.offset.internalize.term true in +set_option trace.grind.offset.dist true in +example (a1 a2 a3 : Nat) : + a1 ≤ a2 → a2 + 2 ≤ a3 → False := by + fail_if_success grind + sorry + +/-- +info: [grind.offset.internalize.term] a1 ↦ #0 +[grind.offset.internalize.term] a2 ↦ #1 +[grind.offset.dist] #0 ≤ #1 +[grind.offset.internalize.term] a3 ↦ #2 +[grind.offset.dist] #1 ≤ #2 + 5 +[grind.offset.dist] #0 ≤ #2 + 5 +-/ +#guard_msgs (info) in +set_option trace.grind.offset.internalize.term true in +set_option trace.grind.offset.dist true in +example (a1 a2 a3 : Nat) : + a1 ≤ a2 → a2 ≤ a3 + 5 → False := by + fail_if_success grind + sorry + +/-- +info: [grind.offset.internalize.term] a1 ↦ #0 +[grind.offset.internalize.term] a2 ↦ #1 +[grind.offset.dist] #0 ≤ #1 + 5 +[grind.offset.internalize.term] a3 ↦ #2 +[grind.offset.dist] #1 ≤ #2 +[grind.offset.dist] #0 ≤ #2 + 5 +-/ +#guard_msgs (info) in +set_option trace.grind.offset.internalize.term true in +set_option trace.grind.offset.dist true in +example (a1 a2 a3 : Nat) : + a1 ≤ a2 + 5 → a2 ≤ a3 → False := by + fail_if_success grind + sorry + +/-- +info: [grind.offset.internalize.term] a1 ↦ #0 +[grind.offset.internalize.term] a2 ↦ #1 +[grind.offset.dist] #0 ≤ #1 + 5 +[grind.offset.internalize.term] a3 ↦ #2 +[grind.offset.dist] #1 + 2 ≤ #2 +[grind.offset.dist] #0 ≤ #2 + 3 +-/ +#guard_msgs (info) in +set_option trace.grind.offset.internalize.term true in +set_option trace.grind.offset.dist true in +example (a1 a2 a3 : Nat) : + a1 ≤ a2 + 5 → a2 + 2 ≤ a3 → False := by + fail_if_success grind + sorry + +/-- +info: [grind.offset.internalize.term] a1 ↦ #0 +[grind.offset.internalize.term] a2 ↦ #1 +[grind.offset.dist] #0 ≤ #1 + 5 +[grind.offset.internalize.term] a3 ↦ #2 +[grind.offset.dist] #1 ≤ #2 + 2 +[grind.offset.dist] #0 ≤ #2 + 7 +-/ +#guard_msgs (info) in +set_option trace.grind.offset.internalize.term true in +set_option trace.grind.offset.dist true in +example (a1 a2 a3 : Nat) : + a1 ≤ a2 + 5 → a2 ≤ a3 + 2 → False := by + fail_if_success grind + sorry + + +set_option trace.grind.debug.offset.proof true in +example (a1 a2 a3 : Nat) : + a1 ≤ a2 + 5 → a2 ≤ a3 + 2 → False := by + fail_if_success grind + sorry + +/-- +info: [grind.offset.internalize.term] a1 ↦ #0 +[grind.offset.internalize.term] a2 ↦ #1 +[grind.offset.dist] #0 ≤ #1 + 2 +[grind.offset.internalize.term] a3 ↦ #2 +[grind.offset.dist] #1 + 3 ≤ #2 +[grind.offset.dist] #0 + 1 ≤ #2 +-/ +#guard_msgs (info) in +set_option trace.grind.offset.internalize.term true in +set_option trace.grind.offset.dist true in +example (a1 a2 a3 : Nat) : a1 ≤ a2 + 2 → a2 + 3 ≤ a3 → False := by + fail_if_success grind + sorry + +/-- +info: [grind.offset.internalize.term] a2 ↦ #0 +[grind.offset.internalize.term] a1 ↦ #1 +[grind.offset.dist] #1 + 3 ≤ #0 +[grind.offset.internalize.term] a3 ↦ #2 +[grind.offset.dist] #0 + 3 ≤ #2 +[grind.offset.dist] #1 + 6 ≤ #2 +-/ +#guard_msgs (info) in +set_option trace.grind.offset.internalize.term true in +set_option trace.grind.offset.dist true in +example (p : Prop) (a1 a2 a3 : Nat) : (p ↔ a2 ≤ a1 + 2) → ¬p → a2 + 3 ≤ a3 → False := by + fail_if_success grind + sorry + +/-- +info: [grind.offset.internalize.term] a2 ↦ #0 +[grind.offset.internalize.term] a1 ↦ #1 +[grind.offset.dist] #1 ≤ #0 + 1 +[grind.offset.internalize.term] a3 ↦ #2 +[grind.offset.dist] #0 + 3 ≤ #2 +[grind.offset.dist] #1 + 2 ≤ #2 +-/ +#guard_msgs (info) in +set_option trace.grind.offset.internalize.term true in +set_option trace.grind.offset.dist true in +example (p : Prop) (a1 a2 a3 : Nat) : (p ↔ a2 + 2 ≤ a1) → ¬p → a2 + 3 ≤ a3 → False := by + fail_if_success grind + sorry + +/-- +info: [grind.offset.internalize.term] a2 ↦ #0 +[grind.offset.internalize.term] a1 ↦ #1 +[grind.offset.dist] #1 + 1 ≤ #0 +[grind.offset.internalize.term] a3 ↦ #2 +[grind.offset.dist] #0 + 3 ≤ #2 +[grind.offset.dist] #1 + 4 ≤ #2 +-/ +#guard_msgs (info) in +set_option trace.grind.offset.internalize.term true in +set_option trace.grind.offset.dist true in +example (p : Prop) (a1 a2 a3 : Nat) : (p ↔ a2 ≤ a1) → ¬p → a2 + 3 ≤ a3 → False := by + fail_if_success grind + sorry + +/-- +info: [grind.offset.internalize.term] a2 ↦ #0 +[grind.offset.internalize.term] a1 ↦ #1 +[grind.offset.dist] #1 ≤ #0 +[grind.offset.internalize.term] a3 ↦ #2 +[grind.offset.dist] #0 + 3 ≤ #2 +[grind.offset.dist] #1 + 3 ≤ #2 +-/ +#guard_msgs (info) in +set_option trace.grind.offset.internalize.term true in +set_option trace.grind.offset.dist true in +example (p : Prop) (a1 a2 a3 : Nat) : (p ↔ a2 + 1 ≤ a1) → ¬p → a2 + 3 ≤ a3 → False := by + fail_if_success grind + sorry + +example (a b c : Nat) : a ≤ b → b + 2 ≤ c → a + 1 ≤ c := by + grind +example (a b c : Nat) : a ≤ b → b ≤ c → a ≤ c := by + grind +example (a b c : Nat) : a + 1 ≤ b → b + 1 ≤ c → a + 2 ≤ c := by + grind +example (a b c : Nat) : a + 1 ≤ b → b + 1 ≤ c → a + 1 ≤ c := by + grind +example (a b c : Nat) : a + 1 ≤ b → b ≤ c + 2 → a ≤ c + 1 := by + grind +example (a b c : Nat) : a + 2 ≤ b → b ≤ c + 2 → a ≤ c := by + grind diff --git a/tests/lean/run/grind_offset_ineq_thms.lean b/tests/lean/run/grind_offset_ineq_thms.lean deleted file mode 100644 index bb76dcfe205e..000000000000 --- a/tests/lean/run/grind_offset_ineq_thms.lean +++ /dev/null @@ -1,31 +0,0 @@ -import Lean - -elab tk:"#R[" ts:term,* "]" : term => do - let ts : Array Lean.Syntax := ts - let es ← ts.mapM fun stx => Lean.Elab.Term.elabTerm stx none - if h : 0 < es.size then - return (Lean.RArray.toExpr (← Lean.Meta.inferType es[0]!) id (Lean.RArray.ofArray es h)) - else - throwErrorAt tk "RArray cannot be empty" - -open Lean.Grind.Offset - -macro "C[" "#" x:term:max " ≤ " "#" y:term:max "]" : term => `({ x := $x, y := $y : Cnstr }) -macro "C[" "#" x:term:max " + " k:term:max " ≤ " "#" y:term:max "]" : term => `({ x := $x, y := $y, k := $k : Cnstr }) -macro "C[" "#" x:term:max " ≤ " "#"y:term:max " + " k:term:max "]" : term => `({ x := $x, y := $y, k := $k, l := false : Cnstr }) - -example (x y z : Nat) : x + 2 ≤ y → y ≤ z → z + 1 ≤ x → False := - Cnstrs.unsat #R[x, y, z] [ - C[ #0 + 2 ≤ #1 ], - C[ #1 ≤ #2 ], - C[ #2 + 1 ≤ #0 ] - ] rfl - -example (x y z w : Nat) : x + 2 ≤ y → y ≤ z → z ≤ w + 7 → x ≤ w + 5 := - Cnstrs.imp #R[x, y, z, w] [ - C[ #0 + 2 ≤ #1 ], - C[ #1 ≤ #2 ], - C[ #2 ≤ #3 + 7] - ] - C[ #0 ≤ #3 + 5 ] - rfl From af8f3d1ec1cb69c527db61448828bf57070aa248 Mon Sep 17 00:00:00 2001 From: Leonardo de Moura Date: Sun, 12 Jan 2025 15:09:39 -0800 Subject: [PATCH 085/100] feat: avoid some redundant proof terms in `grind` (#6615) This PR adds two auxiliary functions `mkEqTrueCore` and `mkOfEqTrueCore` that avoid redundant proof terms in proofs produced by `grind`. --- src/Init/Grind/Lemmas.lean | 3 + src/Lean/Meta/AppBuilder.lean | 112 ++++++++++-------- .../Meta/Tactic/Grind/Arith/ProofUtil.lean | 5 +- src/Lean/Meta/Tactic/Grind/ForallProp.lean | 4 +- src/Lean/Meta/Tactic/Grind/Propagate.lean | 12 +- src/Lean/Meta/Tactic/Grind/Split.lean | 2 +- tests/lean/run/grind_offset_cnstr.lean | 20 ++++ tests/lean/run/grind_pre.lean | 4 - 8 files changed, 98 insertions(+), 64 deletions(-) diff --git a/src/Init/Grind/Lemmas.lean b/src/Init/Grind/Lemmas.lean index 6a6a571abb8a..af289da274d8 100644 --- a/src/Init/Grind/Lemmas.lean +++ b/src/Init/Grind/Lemmas.lean @@ -12,6 +12,9 @@ import Init.Grind.Util namespace Lean.Grind +theorem rfl_true : true = true := + rfl + theorem intro_with_eq (p p' q : Prop) (he : p = p') (h : p' → q) : p → q := fun hp => h (he.mp hp) diff --git a/src/Lean/Meta/AppBuilder.lean b/src/Lean/Meta/AppBuilder.lean index 657543d536c5..26d333ee796d 100644 --- a/src/Lean/Meta/AppBuilder.lean +++ b/src/Lean/Meta/AppBuilder.lean @@ -11,14 +11,14 @@ import Lean.Meta.DecLevel namespace Lean.Meta -/-- Return `id e` -/ +/-- Returns `id e` -/ def mkId (e : Expr) : MetaM Expr := do let type ← inferType e let u ← getLevel type return mkApp2 (mkConst ``id [u]) type e /-- - Given `e` s.t. `inferType e` is definitionally equal to `expectedType`, return + Given `e` s.t. `inferType e` is definitionally equal to `expectedType`, returns term `@id expectedType e`. -/ def mkExpectedTypeHint (e : Expr) (expectedType : Expr) : MetaM Expr := do let u ← getLevel expectedType @@ -38,13 +38,13 @@ def mkLetFun (x : Expr) (v : Expr) (e : Expr) : MetaM Expr := do let u2 ← getLevel ety return mkAppN (.const ``letFun [u1, u2]) #[α, β, v, f] -/-- Return `a = b`. -/ +/-- Returns `a = b`. -/ def mkEq (a b : Expr) : MetaM Expr := do let aType ← inferType a let u ← getLevel aType return mkApp3 (mkConst ``Eq [u]) aType a b -/-- Return `HEq a b`. -/ +/-- Returns `HEq a b`. -/ def mkHEq (a b : Expr) : MetaM Expr := do let aType ← inferType a let bType ← inferType b @@ -52,7 +52,7 @@ def mkHEq (a b : Expr) : MetaM Expr := do return mkApp4 (mkConst ``HEq [u]) aType a bType b /-- - If `a` and `b` have definitionally equal types, return `Eq a b`, otherwise return `HEq a b`. + If `a` and `b` have definitionally equal types, returns `Eq a b`, otherwise returns `HEq a b`. -/ def mkEqHEq (a b : Expr) : MetaM Expr := do let aType ← inferType a @@ -63,25 +63,25 @@ def mkEqHEq (a b : Expr) : MetaM Expr := do else return mkApp4 (mkConst ``HEq [u]) aType a bType b -/-- Return a proof of `a = a`. -/ +/-- Returns a proof of `a = a`. -/ def mkEqRefl (a : Expr) : MetaM Expr := do let aType ← inferType a let u ← getLevel aType return mkApp2 (mkConst ``Eq.refl [u]) aType a -/-- Return a proof of `HEq a a`. -/ +/-- Returns a proof of `HEq a a`. -/ def mkHEqRefl (a : Expr) : MetaM Expr := do let aType ← inferType a let u ← getLevel aType return mkApp2 (mkConst ``HEq.refl [u]) aType a -/-- Given `hp : P` and `nhp : Not P` returns an instance of type `e`. -/ +/-- Given `hp : P` and `nhp : Not P`, returns an instance of type `e`. -/ def mkAbsurd (e : Expr) (hp hnp : Expr) : MetaM Expr := do let p ← inferType hp let u ← getLevel e return mkApp4 (mkConst ``absurd [u]) p e hp hnp -/-- Given `h : False`, return an instance of type `e`. -/ +/-- Given `h : False`, returns an instance of type `e`. -/ def mkFalseElim (e : Expr) (h : Expr) : MetaM Expr := do let u ← getLevel e return mkApp2 (mkConst ``False.elim [u]) e h @@ -108,7 +108,7 @@ def mkEqSymm (h : Expr) : MetaM Expr := do return mkApp4 (mkConst ``Eq.symm [u]) α a b h | none => throwAppBuilderException ``Eq.symm ("equality proof expected" ++ hasTypeMsg h hType) -/-- Given `h₁ : a = b` and `h₂ : b = c` returns a proof of `a = c`. -/ +/-- Given `h₁ : a = b` and `h₂ : b = c`, returns a proof of `a = c`. -/ def mkEqTrans (h₁ h₂ : Expr) : MetaM Expr := do if h₁.isAppOf ``Eq.refl then return h₂ @@ -185,7 +185,7 @@ def mkHEqOfEq (h : Expr) : MetaM Expr := do return mkApp4 (mkConst ``heq_of_eq [u]) α a b h /-- -If `e` is `@Eq.refl α a`, return `a`. +If `e` is `@Eq.refl α a`, returns `a`. -/ def isRefl? (e : Expr) : Option Expr := do if e.isAppOfArity ``Eq.refl 2 then @@ -194,7 +194,7 @@ def isRefl? (e : Expr) : Option Expr := do none /-- -If `e` is `@congrArg α β a b f h`, return `α`, `f` and `h`. +If `e` is `@congrArg α β a b f h`, returns `α`, `f` and `h`. Also works if `e` can be turned into such an application (e.g. `congrFun`). -/ def congrArg? (e : Expr) : MetaM (Option (Expr × Expr × Expr)) := do @@ -336,13 +336,14 @@ private def withAppBuilderTrace [ToMessageData α] [ToMessageData β] throw ex /-- - Return the application `constName xs`. + Returns the application `constName xs`. It tries to fill the implicit arguments before the last element in `xs`. Remark: ``mkAppM `arbitrary #[α]`` returns `@arbitrary.{u} α` without synthesizing the implicit argument occurring after `α`. - Given a `x : ([Decidable p] → Bool) × Nat`, ``mkAppM `Prod.fst #[x]`` returns `@Prod.fst ([Decidable p] → Bool) Nat x`. + Given a `x : ([Decidable p] → Bool) × Nat`, ``mkAppM `Prod.fst #[x]``, + returns `@Prod.fst ([Decidable p] → Bool) Nat x`. -/ def mkAppM (constName : Name) (xs : Array Expr) : MetaM Expr := do withAppBuilderTrace constName xs do withNewMCtxDepth do @@ -465,8 +466,9 @@ def mkPure (monad : Expr) (e : Expr) : MetaM Expr := mkAppOptM ``Pure.pure #[monad, none, none, e] /-- - `mkProjection s fieldName` returns an expression for accessing field `fieldName` of the structure `s`. - Remark: `fieldName` may be a subfield of `s`. -/ +`mkProjection s fieldName` returns an expression for accessing field `fieldName` of the structure `s`. +Remark: `fieldName` may be a subfield of `s`. +-/ partial def mkProjection (s : Expr) (fieldName : Name) : MetaM Expr := do let type ← inferType s let type ← whnf type @@ -520,11 +522,11 @@ def mkSome (type value : Expr) : MetaM Expr := do let u ← getDecLevel type return mkApp2 (mkConst ``Option.some [u]) type value -/-- Return `Decidable.decide p` -/ +/-- Returns `Decidable.decide p` -/ def mkDecide (p : Expr) : MetaM Expr := mkAppOptM ``Decidable.decide #[p, none] -/-- Return a proof for `p : Prop` using `decide p` -/ +/-- Returns a proof for `p : Prop` using `decide p` -/ def mkDecideProof (p : Expr) : MetaM Expr := do let decP ← mkDecide p let decEqTrue ← mkEq decP (mkConst ``Bool.true) @@ -532,63 +534,75 @@ def mkDecideProof (p : Expr) : MetaM Expr := do let h ← mkExpectedTypeHint h decEqTrue mkAppM ``of_decide_eq_true #[h] -/-- Return `a < b` -/ +/-- Returns `a < b` -/ def mkLt (a b : Expr) : MetaM Expr := mkAppM ``LT.lt #[a, b] -/-- Return `a <= b` -/ +/-- Returns `a <= b` -/ def mkLe (a b : Expr) : MetaM Expr := mkAppM ``LE.le #[a, b] -/-- Return `Inhabited.default α` -/ +/-- Returns `Inhabited.default α` -/ def mkDefault (α : Expr) : MetaM Expr := mkAppOptM ``Inhabited.default #[α, none] -/-- Return `@Classical.ofNonempty α _` -/ +/-- Returns `@Classical.ofNonempty α _` -/ def mkOfNonempty (α : Expr) : MetaM Expr := do mkAppOptM ``Classical.ofNonempty #[α, none] -/-- Return `funext h` -/ +/-- Returns `funext h` -/ def mkFunExt (h : Expr) : MetaM Expr := mkAppM ``funext #[h] -/-- Return `propext h` -/ +/-- Returns `propext h` -/ def mkPropExt (h : Expr) : MetaM Expr := mkAppM ``propext #[h] -/-- Return `let_congr h₁ h₂` -/ +/-- Returns `let_congr h₁ h₂` -/ def mkLetCongr (h₁ h₂ : Expr) : MetaM Expr := mkAppM ``let_congr #[h₁, h₂] -/-- Return `let_val_congr b h` -/ +/-- Returns `let_val_congr b h` -/ def mkLetValCongr (b h : Expr) : MetaM Expr := mkAppM ``let_val_congr #[b, h] -/-- Return `let_body_congr a h` -/ +/-- Returns `let_body_congr a h` -/ def mkLetBodyCongr (a h : Expr) : MetaM Expr := mkAppM ``let_body_congr #[a, h] -/-- Return `of_eq_true h` -/ +/-- Returns `@of_eq_true p h` -/ +def mkOfEqTrueCore (p : Expr) (h : Expr) : Expr := + match_expr h with + | eq_true _ h => h + | _ => mkApp2 (mkConst ``of_eq_true) p h + +/-- Returns `of_eq_true h` -/ def mkOfEqTrue (h : Expr) : MetaM Expr := do match_expr h with | eq_true _ h => return h | _ => mkAppM ``of_eq_true #[h] -/-- Return `eq_true h` -/ +/-- Returns `eq_true h` -/ +def mkEqTrueCore (p : Expr) (h : Expr) : Expr := + match_expr h with + | of_eq_true _ h => h + | _ => mkApp2 (mkConst ``eq_true) p h + +/-- Returns `eq_true h` -/ def mkEqTrue (h : Expr) : MetaM Expr := do match_expr h with | of_eq_true _ h => return h | _ => return mkApp2 (mkConst ``eq_true) (← inferType h) h /-- - Return `eq_false h` + Returns `eq_false h` `h` must have type definitionally equal to `¬ p` in the current reducibility setting. -/ def mkEqFalse (h : Expr) : MetaM Expr := mkAppM ``eq_false #[h] /-- - Return `eq_false' h` + Returns `eq_false' h` `h` must have type definitionally equal to `p → False` in the current reducibility setting. -/ def mkEqFalse' (h : Expr) : MetaM Expr := @@ -606,7 +620,7 @@ def mkImpDepCongrCtx (h₁ h₂ : Expr) : MetaM Expr := def mkForallCongr (h : Expr) : MetaM Expr := mkAppM ``forall_congr #[h] -/-- Return instance for `[Monad m]` if there is one -/ +/-- Returns instance for `[Monad m]` if there is one -/ def isMonad? (m : Expr) : MetaM (Option Expr) := try let monadType ← mkAppM `Monad #[m] @@ -617,52 +631,52 @@ def isMonad? (m : Expr) : MetaM (Option Expr) := catch _ => pure none -/-- Return `(n : type)`, a numeric literal of type `type`. The method fails if we don't have an instance `OfNat type n` -/ +/-- Returns `(n : type)`, a numeric literal of type `type`. The method fails if we don't have an instance `OfNat type n` -/ def mkNumeral (type : Expr) (n : Nat) : MetaM Expr := do let u ← getDecLevel type let inst ← synthInstance (mkApp2 (mkConst ``OfNat [u]) type (mkRawNatLit n)) return mkApp3 (mkConst ``OfNat.ofNat [u]) type (mkRawNatLit n) inst /-- - Return `a op b`, where `op` has name `opName` and is implemented using the typeclass `className`. - This method assumes `a` and `b` have the same type, and typeclass `className` is heterogeneous. - Examples of supported classes: `HAdd`, `HSub`, `HMul`. - We use heterogeneous operators to ensure we have a uniform representation. - -/ +Returns `a op b`, where `op` has name `opName` and is implemented using the typeclass `className`. +This method assumes `a` and `b` have the same type, and typeclass `className` is heterogeneous. +Examples of supported classes: `HAdd`, `HSub`, `HMul`. +We use heterogeneous operators to ensure we have a uniform representation. +-/ private def mkBinaryOp (className : Name) (opName : Name) (a b : Expr) : MetaM Expr := do let aType ← inferType a let u ← getDecLevel aType let inst ← synthInstance (mkApp3 (mkConst className [u, u, u]) aType aType aType) return mkApp6 (mkConst opName [u, u, u]) aType aType aType inst a b -/-- Return `a + b` using a heterogeneous `+`. This method assumes `a` and `b` have the same type. -/ +/-- Returns `a + b` using a heterogeneous `+`. This method assumes `a` and `b` have the same type. -/ def mkAdd (a b : Expr) : MetaM Expr := mkBinaryOp ``HAdd ``HAdd.hAdd a b -/-- Return `a - b` using a heterogeneous `-`. This method assumes `a` and `b` have the same type. -/ +/-- Returns `a - b` using a heterogeneous `-`. This method assumes `a` and `b` have the same type. -/ def mkSub (a b : Expr) : MetaM Expr := mkBinaryOp ``HSub ``HSub.hSub a b -/-- Return `a * b` using a heterogeneous `*`. This method assumes `a` and `b` have the same type. -/ +/-- Returns `a * b` using a heterogeneous `*`. This method assumes `a` and `b` have the same type. -/ def mkMul (a b : Expr) : MetaM Expr := mkBinaryOp ``HMul ``HMul.hMul a b /-- - Return `a r b`, where `r` has name `rName` and is implemented using the typeclass `className`. - This method assumes `a` and `b` have the same type. - Examples of supported classes: `LE` and `LT`. - We use heterogeneous operators to ensure we have a uniform representation. - -/ +Returns `a r b`, where `r` has name `rName` and is implemented using the typeclass `className`. +This method assumes `a` and `b` have the same type. +Examples of supported classes: `LE` and `LT`. +We use heterogeneous operators to ensure we have a uniform representation. +-/ private def mkBinaryRel (className : Name) (rName : Name) (a b : Expr) : MetaM Expr := do let aType ← inferType a let u ← getDecLevel aType let inst ← synthInstance (mkApp (mkConst className [u]) aType) return mkApp4 (mkConst rName [u]) aType inst a b -/-- Return `a ≤ b`. This method assumes `a` and `b` have the same type. -/ +/-- Returns `a ≤ b`. This method assumes `a` and `b` have the same type. -/ def mkLE (a b : Expr) : MetaM Expr := mkBinaryRel ``LE ``LE.le a b -/-- Return `a < b`. This method assumes `a` and `b` have the same type. -/ +/-- Returns `a < b`. This method assumes `a` and `b` have the same type. -/ def mkLT (a b : Expr) : MetaM Expr := mkBinaryRel ``LT ``LT.lt a b -/-- Given `h : a = b`, return a proof for `a ↔ b`. -/ +/-- Given `h : a = b`, returns a proof for `a ↔ b`. -/ def mkIffOfEq (h : Expr) : MetaM Expr := do if h.isAppOfArity ``propext 3 then return h.appArg! diff --git a/src/Lean/Meta/Tactic/Grind/Arith/ProofUtil.lean b/src/Lean/Meta/Tactic/Grind/Arith/ProofUtil.lean index e281f2b813ab..602814cbdfd6 100644 --- a/src/Lean/Meta/Tactic/Grind/Arith/ProofUtil.lean +++ b/src/Lean/Meta/Tactic/Grind/Arith/ProofUtil.lean @@ -5,6 +5,7 @@ Authors: Leonardo de Moura -/ prelude import Init.Grind.Offset +import Init.Grind.Lemmas import Lean.Meta.Tactic.Grind.Types namespace Lean.Meta.Grind.Arith @@ -15,8 +16,8 @@ Helper functions for constructing proof terms in the arithmetic procedures. namespace Offset -/-- `Eq.refl true` -/ -def rfl_true : Expr := mkApp2 (mkConst ``Eq.refl [levelOne]) (mkConst ``Bool) (mkConst ``Bool.true) +/-- Returns a proof for `true = true` -/ +def rfl_true : Expr := mkConst ``Grind.rfl_true open Lean.Grind in /-- diff --git a/src/Lean/Meta/Tactic/Grind/ForallProp.lean b/src/Lean/Meta/Tactic/Grind/ForallProp.lean index cca14b85ebe4..dbf0f234cbd7 100644 --- a/src/Lean/Meta/Tactic/Grind/ForallProp.lean +++ b/src/Lean/Meta/Tactic/Grind/ForallProp.lean @@ -24,7 +24,7 @@ def propagateForallPropUp (e : Expr) : GoalM Unit := do unless (← isEqTrue p) do return trace_goal[grind.debug.forallPropagator] "isEqTrue, {e}" let h₁ ← mkEqTrueProof p - let qh₁ := q.instantiate1 (mkApp2 (mkConst ``of_eq_true) p h₁) + let qh₁ := q.instantiate1 (mkOfEqTrueCore p h₁) let r ← simp qh₁ let q := mkLambda n bi p q let q' := r.expr @@ -65,7 +65,7 @@ private def addLocalEMatchTheorems (e : Expr) : GoalM Unit := do else let idx ← modifyGet fun s => (s.nextThmIdx, { s with nextThmIdx := s.nextThmIdx + 1 }) pure <| .local ((`local).appendIndexAfter idx) - let proof := mkApp2 (mkConst ``of_eq_true) e proof + let proof := mkOfEqTrueCore e proof let size := (← get).newThms.size let gen ← getGeneration e -- TODO: we should have a flag for collecting all unary patterns in a local theorem diff --git a/src/Lean/Meta/Tactic/Grind/Propagate.lean b/src/Lean/Meta/Tactic/Grind/Propagate.lean index bfe1a3d014a5..fd13a47e44f1 100644 --- a/src/Lean/Meta/Tactic/Grind/Propagate.lean +++ b/src/Lean/Meta/Tactic/Grind/Propagate.lean @@ -126,32 +126,32 @@ builtin_grind_propagator propagateEqUp ↑Eq := fun e => do else if (← isEqTrue b) then pushEq e a <| mkApp3 (mkConst ``Lean.Grind.eq_eq_of_eq_true_right) a b (← mkEqTrueProof b) else if (← isEqv a b) then - pushEqTrue e <| mkApp2 (mkConst ``eq_true) e (← mkEqProof a b) + pushEqTrue e <| mkEqTrueCore e (← mkEqProof a b) /-- Propagates `Eq` downwards -/ builtin_grind_propagator propagateEqDown ↓Eq := fun e => do if (← isEqTrue e) then let_expr Eq _ a b := e | return () - pushEq a b <| mkApp2 (mkConst ``of_eq_true) e (← mkEqTrueProof e) + pushEq a b <| mkOfEqTrueCore e (← mkEqTrueProof e) /-- Propagates `EqMatch` downwards -/ builtin_grind_propagator propagateEqMatchDown ↓Grind.EqMatch := fun e => do if (← isEqTrue e) then let_expr Grind.EqMatch _ a b origin := e | return () markCaseSplitAsResolved origin - pushEq a b <| mkApp2 (mkConst ``of_eq_true) e (← mkEqTrueProof e) + pushEq a b <| mkOfEqTrueCore e (← mkEqTrueProof e) /-- Propagates `HEq` downwards -/ builtin_grind_propagator propagateHEqDown ↓HEq := fun e => do if (← isEqTrue e) then let_expr HEq _ a _ b := e | return () - pushHEq a b <| mkApp2 (mkConst ``of_eq_true) e (← mkEqTrueProof e) + pushHEq a b <| mkOfEqTrueCore e (← mkEqTrueProof e) /-- Propagates `HEq` upwards -/ builtin_grind_propagator propagateHEqUp ↑HEq := fun e => do let_expr HEq _ a _ b := e | return () if (← isEqv a b) then - pushEqTrue e <| mkApp2 (mkConst ``eq_true) e (← mkHEqProof a b) + pushEqTrue e <| mkEqTrueCore e (← mkHEqProof a b) /-- Propagates `ite` upwards -/ builtin_grind_propagator propagateIte ↑ite := fun e => do @@ -166,7 +166,7 @@ builtin_grind_propagator propagateDIte ↑dite := fun e => do let_expr f@dite α c h a b := e | return () if (← isEqTrue c) then let h₁ ← mkEqTrueProof c - let ah₁ := mkApp a (mkApp2 (mkConst ``of_eq_true) c h₁) + let ah₁ := mkApp a (mkOfEqTrueCore c h₁) let p ← simp ah₁ let r := p.expr let h₂ ← p.getProof diff --git a/src/Lean/Meta/Tactic/Grind/Split.lean b/src/Lean/Meta/Tactic/Grind/Split.lean index 3937644300a7..15a232bdeb1e 100644 --- a/src/Lean/Meta/Tactic/Grind/Split.lean +++ b/src/Lean/Meta/Tactic/Grind/Split.lean @@ -145,7 +145,7 @@ private def mkCasesMajor (c : Expr) : GoalM Expr := do return mkApp3 (mkConst ``Grind.of_eq_eq_true) a b (← mkEqTrueProof c) else return mkApp3 (mkConst ``Grind.of_eq_eq_false) a b (← mkEqFalseProof c) - | _ => return mkApp2 (mkConst ``of_eq_true) c (← mkEqTrueProof c) + | _ => return mkOfEqTrueCore c (← mkEqTrueProof c) /-- Introduces new hypotheses in each goal. -/ private def introNewHyp (goals : List Goal) (acc : List Goal) (generation : Nat) : GrindM (List Goal) := do diff --git a/tests/lean/run/grind_offset_cnstr.lean b/tests/lean/run/grind_offset_cnstr.lean index cbad141607c1..cb41c2c76242 100644 --- a/tests/lean/run/grind_offset_cnstr.lean +++ b/tests/lean/run/grind_offset_cnstr.lean @@ -255,3 +255,23 @@ example (a b c : Nat) : a + 1 ≤ b → b ≤ c + 2 → a ≤ c + 1 := by grind example (a b c : Nat) : a + 2 ≤ b → b ≤ c + 2 → a ≤ c := by grind + +theorem ex1 (p : Prop) (a1 a2 a3 : Nat) : (p ↔ a2 ≤ a1) → ¬p → a2 + 3 ≤ a3 → (p ↔ a4 ≤ a3 + 2) → a1 ≤ a4 := by + grind + +/-- +info: theorem ex1 : ∀ {a4 : Nat} (p : Prop) (a1 a2 a3 : Nat), + (p ↔ a2 ≤ a1) → ¬p → a2 + 3 ≤ a3 → (p ↔ a4 ≤ a3 + 2) → a1 ≤ a4 := +fun {a4} p a1 a2 a3 => + intro_with_eq (p ↔ a2 ≤ a1) (p = (a2 ≤ a1)) (¬p → a2 + 3 ≤ a3 → (p ↔ a4 ≤ a3 + 2) → a1 ≤ a4) (iff_eq p (a2 ≤ a1)) + fun a a_1 a_2 => + intro_with_eq (p ↔ a4 ≤ a3 + 2) (p = (a4 ≤ a3 + 2)) (a1 ≤ a4) (iff_eq p (a4 ≤ a3 + 2)) fun a_3 => + Classical.byContradiction + (intro_with_eq (¬a1 ≤ a4) (a4 + 1 ≤ a1) False (Nat.not_le_eq a1 a4) fun x => + Nat.unsat_lo_lo a4 a1 1 7 rfl_true x + (Nat.lo_lo a1 a2 a4 1 6 (Nat.of_le_eq_false a2 a1 (Eq.trans (Eq.symm a) (eq_false a_1))) + (Nat.lo_lo a2 a3 a4 3 3 a_2 (Nat.of_ro_eq_false a4 a3 2 (Eq.trans (Eq.symm a_3) (eq_false a_1)))))) +-/ +#guard_msgs (info) in +open Lean Grind in +#print ex1 diff --git a/tests/lean/run/grind_pre.lean b/tests/lean/run/grind_pre.lean index 21e4360b9de6..833f163786ef 100644 --- a/tests/lean/run/grind_pre.lean +++ b/tests/lean/run/grind_pre.lean @@ -9,10 +9,8 @@ case grind.1.2 a b c : Bool p q : Prop left✝ : a = true -right✝ : b = true ∨ c = true left : p right : q -x✝ : b = false ∨ a = false h✝ : b = false h : c = true ⊢ False @@ -94,7 +92,6 @@ h₁ : HEq p a h₂ : HEq q a h₃ : p = r left : ¬p ∨ r -right : ¬r ∨ p h : ¬r ⊢ False @@ -106,7 +103,6 @@ h₁ : HEq p a h₂ : HEq q a h₃ : p = r left : ¬p ∨ r -right : ¬r ∨ p h : p ⊢ False -/ From aa95a1c03f2bc6ef8cb3416c04395b8795485f7a Mon Sep 17 00:00:00 2001 From: Leonardo de Moura Date: Sun, 12 Jan 2025 16:07:48 -0800 Subject: [PATCH 086/100] chore: cleaunp `grind` tests (#6616) Tests using `logInfo` were taking an additional two seconds on my machine. This is a performance issue with the old code generator, where we spend all this time specializing the logging functions for `GoalM`. I have not checked whether the new code generator is also affected by this performance issue. Here is a small example that exposes the issue: ```lean import Lean set_option profiler true open Lean Meta Grind in def test (e : Expr): GoalM Unit := do logInfo e ``` cc @zwarich --- tests/lean/run/grind_canon_insts.lean | 9 ++- tests/lean/run/grind_canon_types.lean | 5 +- tests/lean/run/grind_congr.lean | 11 ++-- tests/lean/run/grind_nested_proofs.lean | 30 ++++----- tests/lean/run/grind_norm_levels.lean | 5 +- .../lean/run/grind_propagate_connectives.lean | 61 ++++++++----------- tests/lean/run/grind_revertAll.lean | 27 -------- 7 files changed, 57 insertions(+), 91 deletions(-) delete mode 100644 tests/lean/run/grind_revertAll.lean diff --git a/tests/lean/run/grind_canon_insts.lean b/tests/lean/run/grind_canon_insts.lean index c38feabfc8fa..266f415363df 100644 --- a/tests/lean/run/grind_canon_insts.lean +++ b/tests/lean/run/grind_canon_insts.lean @@ -53,11 +53,13 @@ theorem left_comm [CommMonoid α] (a b c : α) : a * (b * c) = b * (a * c) := by open Lean Meta Elab Tactic Grind in def fallback : Fallback := do let nodes ← filterENodes fun e => return e.self.isAppOf ``HMul.hMul - logInfo (nodes.toList.map (·.self)) + trace[Meta.debug] "{nodes.toList.map (·.self)}" (← get).mvarId.admit +set_option trace.Meta.debug true + /-- -info: [b * c, a * (b * c), d * (b * c)] +info: [Meta.debug] [b * c, a * (b * c), d * (b * c)] -/ #guard_msgs (info) in example (a b c d : Nat) : b * (a * c) = d * (b * c) → False := by @@ -68,7 +70,8 @@ example (a b c d : Nat) : b * (a * c) = d * (b * c) → False := by set_option pp.notation false in set_option pp.explicit true in /-- -info: [@HMul.hMul Nat Nat Nat (@instHMul Nat instMulNat) b a, @HMul.hMul Nat Nat Nat (@instHMul Nat instMulNat) b d] +info: [Meta.debug] [@HMul.hMul Nat Nat Nat (@instHMul Nat instMulNat) b a, + @HMul.hMul Nat Nat Nat (@instHMul Nat instMulNat) b d] -/ #guard_msgs (info) in example (a b c d : Nat) : b * a = d * b → False := by diff --git a/tests/lean/run/grind_canon_types.lean b/tests/lean/run/grind_canon_types.lean index 08a1e17276a7..8870a41c6271 100644 --- a/tests/lean/run/grind_canon_types.lean +++ b/tests/lean/run/grind_canon_types.lean @@ -6,12 +6,13 @@ def f (a : α) := a open Lean Meta Grind in def fallback : Fallback := do let nodes ← filterENodes fun e => return e.self.isAppOf ``f - logInfo (nodes.toList.map (·.self)) + trace[Meta.debug] "{nodes.toList.map (·.self)}" (← get).mvarId.admit +set_option trace.Meta.debug true set_option pp.explicit true /-- -info: [@f Nat a, @f Nat b] +info: [Meta.debug] [@f Nat a, @f Nat b] -/ #guard_msgs (info) in example (a b c d : Nat) : @f Nat a = b → @f (g Nat) a = c → @f (g Nat) b = d → a = b → False := by diff --git a/tests/lean/run/grind_congr.lean b/tests/lean/run/grind_congr.lean index 1589785ef00f..4f09c4809f06 100644 --- a/tests/lean/run/grind_congr.lean +++ b/tests/lean/run/grind_congr.lean @@ -8,35 +8,36 @@ open Lean Meta Grind in def fallback : Fallback := do let #[n, _] ← filterENodes fun e => return e.self.isAppOf ``f | unreachable! let eqc ← getEqc n.self - logInfo eqc + trace[Meta.debug] eqc (← get).mvarId.admit +set_option trace.Meta.debug true set_option grind.debug true set_option grind.debug.proofs true /-- -info: [d, f b, c, f a] +info: [Meta.debug] [d, f b, c, f a] -/ #guard_msgs (info) in example (a b c d : Nat) : a = b → f a = c → f b = d → False := by grind on_failure fallback /-- -info: [d, f b, c, f a] +info: [Meta.debug] [d, f b, c, f a] -/ #guard_msgs (info) in example (a b c d : Nat) : f a = c → f b = d → a = b → False := by grind on_failure fallback /-- -info: [d, f (g b), c, f (g a)] +info: [Meta.debug] [d, f (g b), c, f (g a)] -/ #guard_msgs (info) in example (a b c d e : Nat) : f (g a) = c → f (g b) = d → a = e → b = e → False := by grind on_failure fallback /-- -info: [d, f (g b), c, f v] +info: [Meta.debug] [d, f (g b), c, f v] -/ #guard_msgs (info) in example (a b c d e v : Nat) : f v = c → f (g b) = d → a = e → b = e → v = g a → False := by diff --git a/tests/lean/run/grind_nested_proofs.lean b/tests/lean/run/grind_nested_proofs.lean index 04b32f7fe742..a394fc7ba6b9 100644 --- a/tests/lean/run/grind_nested_proofs.lean +++ b/tests/lean/run/grind_nested_proofs.lean @@ -5,12 +5,13 @@ def f (α : Type) [Add α] (a : α) := a + a + a open Lean Meta Grind in def fallback : Fallback := do let nodes ← filterENodes fun e => return e.self.isAppOf ``Lean.Grind.nestedProof - logInfo (nodes.toList.map (·.self)) + trace[Meta.debug] "{nodes.toList.map (·.self)}" let nodes ← filterENodes fun e => return e.self.isAppOf ``GetElem.getElem let [_, n, _] := nodes.toList | unreachable! - logInfo (← getEqc n.self) + trace[Meta.debug] "{← getEqc n.self}" (← get).mvarId.admit +set_option trace.Meta.debug true set_option grind.debug true set_option grind.debug.proofs true @@ -20,26 +21,21 @@ The following test relies on `grind` `nestedProof` wrapper to detect equalities between array access terms. -/ -/- -info: [Lean.Grind.nestedProof (i < a.toList.length) (_example.proof_1 i j a b h1 h2), - Lean.Grind.nestedProof (j < a.toList.length) h1, - Lean.Grind.nestedProof (j < b.toList.length) h] ---- -info: [a[i], b[j], a[j]] ---- -warning: declaration uses 'sorry' +/-- +info: [Meta.debug] [Lean.Grind.nestedProof (i < a.toList.length) (_example.proof_1 i j a b h1 h2), + Lean.Grind.nestedProof (j < a.toList.length) h1, + Lean.Grind.nestedProof (j < b.toList.length) h] +[Meta.debug] [a[i], b[j], a[j]] -/ --- #guard_msgs in - +#guard_msgs (info) in example (i j : Nat) (a b : Array Nat) (h1 : j < a.size) (h : j < b.size) (h2 : i ≤ j) : a[i] < a[j] + b[j] → i = j → a = b → False := by grind on_failure fallback /-- -info: [Lean.Grind.nestedProof (i < a.toList.length) (_example.proof_1 i j a b h1 h2), - Lean.Grind.nestedProof (j < a.toList.length) h1, - Lean.Grind.nestedProof (j < b.toList.length) h] ---- -info: [a[i], a[j]] +info: [Meta.debug] [Lean.Grind.nestedProof (i < a.toList.length) (_example.proof_1 i j a b h1 h2), + Lean.Grind.nestedProof (j < a.toList.length) h1, + Lean.Grind.nestedProof (j < b.toList.length) h] +[Meta.debug] [a[i], a[j]] -/ #guard_msgs (info) in example (i j : Nat) (a b : Array Nat) (h1 : j < a.size) (h : j < b.size) (h2 : i ≤ j) : a[i] < a[j] + b[j] → i = j → False := by diff --git a/tests/lean/run/grind_norm_levels.lean b/tests/lean/run/grind_norm_levels.lean index 354b56319dd2..9e68e9267803 100644 --- a/tests/lean/run/grind_norm_levels.lean +++ b/tests/lean/run/grind_norm_levels.lean @@ -5,12 +5,13 @@ def g {α : Sort u} (a : α) := a open Lean Meta Grind in def fallback : Fallback := do let nodes ← filterENodes fun e => return e.self.isAppOf ``g - logInfo (nodes.toList.map (·.self)) + trace[Meta.debug] "{nodes.toList.map (·.self)}" (← get).mvarId.admit -- `grind` final state must contain only two `g`-applications +set_option trace.Meta.debug true in /-- -info: [g (a, b), g (g (a, b))] +info: [Meta.debug] [g (a, b), g (g (a, b))] -/ #guard_msgs (info) in example {β : Type v} {α : Type u} (a c : α) (b d : β) : g.{max u v + 1} (a, b) = (c, d) → g (g.{max (u+1) (v+1)} (a, b)) = (c, d) → False := by diff --git a/tests/lean/run/grind_propagate_connectives.lean b/tests/lean/run/grind_propagate_connectives.lean index b45206bb647f..c7f7178b9504 100644 --- a/tests/lean/run/grind_propagate_connectives.lean +++ b/tests/lean/run/grind_propagate_connectives.lean @@ -1,24 +1,25 @@ import Lean.Meta.Tactic.Grind +set_option trace.Meta.debug true + open Lean Meta Grind in def fallback : Fallback := do let trueExprs := (← filterENodes fun e => return e.self.isFVar && (← isEqTrue e.self)).toList.map (·.self) let falseExprs := (← filterENodes fun e => return e.self.isFVar && (← isEqFalse e.self)).toList.map (·.self) - logInfo m!"true: {trueExprs}" - logInfo m!"false: {falseExprs}" + trace[Meta.debug] "true: {trueExprs}" + trace[Meta.debug] "false: {falseExprs}" forEachEqcRoot fun n => do unless (← isProp n.self) || (← isType n.self) || n.size == 1 do let eqc ← getEqc n.self - logInfo eqc + trace[Meta.debug] eqc (← get).mvarId.admit set_option grind.debug true set_option grind.debug.proofs true /-- -info: true: [q, w] ---- -info: false: [p, r] +info: [Meta.debug] true: [q, w] +[Meta.debug] false: [p, r] -/ #guard_msgs (info) in example : (p ∨ (q ∧ ¬r ∧ w)) → ¬p → False := by @@ -26,9 +27,8 @@ example : (p ∨ (q ∧ ¬r ∧ w)) → ¬p → False := by /-- -info: true: [r] ---- -info: false: [p, q] +info: [Meta.debug] true: [r] +[Meta.debug] false: [p, q] -/ #guard_msgs (info) in example : (p ∨ q ∨ r) → (p ∨ ¬q) → ¬p → False := by @@ -36,72 +36,63 @@ example : (p ∨ q ∨ r) → (p ∨ ¬q) → ¬p → False := by /-- -info: true: [r] ---- -info: false: [p₁, q] +info: [Meta.debug] true: [r] +[Meta.debug] false: [p₁, q] -/ #guard_msgs (info) in example : ((p₁ ∧ p₂) ∨ q ∨ r) → (p₁ ∨ ¬q) → p₁ = False → False := by grind on_failure fallback /-- -info: true: [r] ---- -info: false: [p₂, q] +info: [Meta.debug] true: [r] +[Meta.debug] false: [p₂, q] -/ #guard_msgs (info) in example : ((p₁ ∧ p₂) ∨ q ∨ r) → ((p₂ ∧ p₁) ∨ ¬q) → p₂ = False → False := by grind on_failure fallback /-- -info: true: [q, r] ---- -info: false: [p] +info: [Meta.debug] true: [q, r] +[Meta.debug] false: [p] -/ #guard_msgs (info) in example (p q r : Prop) : p ∨ (q ↔ r) → p = False → q → False := by grind on_failure fallback /-- -info: true: [r] ---- -info: false: [p, s] +info: [Meta.debug] true: [r] +[Meta.debug] false: [p, s] -/ #guard_msgs (info) in example (p q r : Prop) : p ∨ ¬(s ∨ (p ↔ r)) → p = False → False := by grind on_failure fallback /-- -info: true: [p] ---- -info: false: [] ---- -info: [a, b] +info: [Meta.debug] true: [p] +[Meta.debug] false: [] +[Meta.debug] [a, b] -/ #guard_msgs (info) in example (p : Prop) (a : Vector Nat 5) (b : Vector Nat 6) : (p → HEq a b) → p → False := by grind on_failure fallback /-- -info: true: [p, q] ---- -info: false: [r] +info: [Meta.debug] true: [p, q] +[Meta.debug] false: [r] -/ #guard_msgs (info) in example (p q r : Prop) : p ∨ (q ↔ r) → q → ¬r → False := by grind on_failure fallback /-- -info: hello world ---- -info: true: [p, q] ---- -info: false: [r] +info: [Meta.debug] hello world +[Meta.debug] true: [p, q] +[Meta.debug] false: [r] -/ #guard_msgs (info) in example (p q r : Prop) : p ∨ (q ↔ r) → ¬r → q → False := by grind on_failure do - Lean.logInfo "hello world" + trace[Meta.debug] "hello world" fallback example (a b : Nat) : (a = b) = (b = a) := by diff --git a/tests/lean/run/grind_revertAll.lean b/tests/lean/run/grind_revertAll.lean deleted file mode 100644 index 79ee7bde59d5..000000000000 --- a/tests/lean/run/grind_revertAll.lean +++ /dev/null @@ -1,27 +0,0 @@ -import Lean - -open Lean Elab Tactic in -elab "revert_all" : tactic => do - liftMetaTactic1 (·.revertAll) - -open Lean Elab Tactic in -elab "ensure_no_mvar" : tactic => do - liftMetaTactic1 fun mvarId => do - mvarId.ensureNoMVar - return mvarId - -example {α : Type u} [Inhabited α] (a : α) (f : α → α) (h : f a = default) : default = f a := by - revert_all - ensure_no_mvar - guard_target =ₛ∀ {α : Type u} [inst : Inhabited α] (a : α) (f : α → α), f a = default → default = f a - intro α inst a f h - exact h.symm - -example (a b : α) (h₁ : a = b) (f g : α → α) (h₂ : ∀ x, f x = x) (h₃ : ∀ x, g x = f x) : ∃ x : α, f x = a ∧ g x = b := by - apply Exists.intro - revert_all - fail_if_success ensure_no_mvar - intro β a₁ a₂ h f₁ f₂ h' h'' - constructor - · exact h' a₁ - · rw [h'', h]; exact h' a₂ From 603108e34c457a8644aade2da418b1d9213ed441 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Mon, 13 Jan 2025 13:00:49 +1100 Subject: [PATCH 087/100] feat: finish alignment of `List/Array/Vector.append` lemmas (#6617) This PR completes alignment of `List`/`Array`/`Vector` `append` lemmas. --- src/Init/Data/Array/Lemmas.lean | 218 +++++++++++++++++++++++++- src/Init/Data/List/Lemmas.lean | 149 ++++++++++-------- src/Init/Data/List/Zip.lean | 4 +- src/Init/Data/Vector/Lemmas.lean | 261 ++++++++++++++++++++++++++++--- 4 files changed, 535 insertions(+), 97 deletions(-) diff --git a/src/Init/Data/Array/Lemmas.lean b/src/Init/Data/Array/Lemmas.lean index 3eee4ef43784..7e0549addb62 100644 --- a/src/Init/Data/Array/Lemmas.lean +++ b/src/Init/Data/Array/Lemmas.lean @@ -1504,14 +1504,20 @@ theorem filterMap_eq_push_iff {f : α → Option β} {l : Array α} {l' : Array · rintro ⟨⟨l₁⟩, a, ⟨l₂⟩, h₁, h₂, h₃, h₄⟩ refine ⟨l₂.reverse, a, l₁.reverse, by simp_all⟩ -/-! Content below this point has not yet been aligned with `List`. -/ - /-! ### singleton -/ @[simp] theorem singleton_def (v : α) : Array.singleton v = #[v] := rfl /-! ### append -/ +@[simp] theorem size_append (as bs : Array α) : (as ++ bs).size = as.size + bs.size := by + simp only [size, toList_append, List.length_append] + +@[simp] theorem append_push {as bs : Array α} {a : α} : as ++ bs.push a = (as ++ bs).push a := by + cases as + cases bs + simp + @[simp] theorem toArray_eq_append_iff {xs : List α} {as bs : Array α} : xs.toArray = as ++ bs ↔ xs = as.toList ++ bs.toList := by cases as @@ -1539,8 +1545,24 @@ theorem mem_append_left {a : α} {l₁ : Array α} (l₂ : Array α) (h : a ∈ theorem mem_append_right {a : α} (l₁ : Array α) {l₂ : Array α} (h : a ∈ l₂) : a ∈ l₁ ++ l₂ := mem_append.2 (Or.inr h) -@[simp] theorem size_append (as bs : Array α) : (as ++ bs).size = as.size + bs.size := by - simp only [size, toList_append, List.length_append] +theorem not_mem_append {a : α} {s t : Array α} (h₁ : a ∉ s) (h₂ : a ∉ t) : a ∉ s ++ t := + mt mem_append.1 $ not_or.mpr ⟨h₁, h₂⟩ + +/-- +See also `eq_push_append_of_mem`, which proves a stronger version +in which the initial array must not contain the element. +-/ +theorem append_of_mem {a : α} {l : Array α} (h : a ∈ l) : ∃ s t : Array α, l = s.push a ++ t := by + obtain ⟨s, t, w⟩ := List.append_of_mem (l := l.toList) (by simpa using h) + replace w := congrArg List.toArray w + refine ⟨s.toArray, t.toArray, by simp_all⟩ + +theorem mem_iff_append {a : α} {l : Array α} : a ∈ l ↔ ∃ s t : Array α, l = s.push a ++ t := + ⟨append_of_mem, fun ⟨s, t, e⟩ => e ▸ by simp⟩ + +theorem forall_mem_append {p : α → Prop} {l₁ l₂ : Array α} : + (∀ (x) (_ : x ∈ l₁ ++ l₂), p x) ↔ (∀ (x) (_ : x ∈ l₁), p x) ∧ (∀ (x) (_ : x ∈ l₂), p x) := by + simp only [mem_append, or_imp, forall_and] theorem empty_append (as : Array α) : #[] ++ as = as := by simp @@ -1599,6 +1621,194 @@ theorem getElem_of_append {l l₁ l₂ : Array α} (eq : l = l₁.push a ++ l₂ rw [← getElem?_eq_getElem, eq, getElem?_append_left (by simp; omega), ← h] simp +@[simp 1100] theorem append_singleton {a : α} {as : Array α} : as ++ #[a] = as.push a := by + cases as + simp + +theorem append_inj {s₁ s₂ t₁ t₂ : Array α} (h : s₁ ++ t₁ = s₂ ++ t₂) (hl : s₁.size = s₂.size) : + s₁ = s₂ ∧ t₁ = t₂ := by + rcases s₁ with ⟨s₁⟩ + rcases s₂ with ⟨s₂⟩ + rcases t₁ with ⟨t₁⟩ + rcases t₂ with ⟨t₂⟩ + simpa using List.append_inj (by simpa using h) (by simpa using hl) + +theorem append_inj_right {s₁ s₂ t₁ t₂ : Array α} + (h : s₁ ++ t₁ = s₂ ++ t₂) (hl : s₁.size = s₂.size) : t₁ = t₂ := + (append_inj h hl).right + +theorem append_inj_left {s₁ s₂ t₁ t₂ : Array α} + (h : s₁ ++ t₁ = s₂ ++ t₂) (hl : s₁.size = s₂.size) : s₁ = s₂ := + (append_inj h hl).left + +/-- Variant of `append_inj` instead requiring equality of the sizes of the second arrays. -/ +theorem append_inj' {s₁ s₂ t₁ t₂ : Array α} (h : s₁ ++ t₁ = s₂ ++ t₂) (hl : t₁.size = t₂.size) : + s₁ = s₂ ∧ t₁ = t₂ := + append_inj h <| @Nat.add_right_cancel _ t₁.size _ <| by + let hap := congrArg size h; simp only [size_append, ← hl] at hap; exact hap + +/-- Variant of `append_inj_right` instead requiring equality of the sizes of the second arrays. -/ +theorem append_inj_right' {s₁ s₂ t₁ t₂ : Array α} + (h : s₁ ++ t₁ = s₂ ++ t₂) (hl : t₁.size = t₂.size) : t₁ = t₂ := + (append_inj' h hl).right + +/-- Variant of `append_inj_left` instead requiring equality of the sizes of the second arrays. -/ +theorem append_inj_left' {s₁ s₂ t₁ t₂ : Array α} + (h : s₁ ++ t₁ = s₂ ++ t₂) (hl : t₁.size = t₂.size) : s₁ = s₂ := + (append_inj' h hl).left + +theorem append_right_inj {t₁ t₂ : Array α} (s) : s ++ t₁ = s ++ t₂ ↔ t₁ = t₂ := + ⟨fun h => append_inj_right h rfl, congrArg _⟩ + +theorem append_left_inj {s₁ s₂ : Array α} (t) : s₁ ++ t = s₂ ++ t ↔ s₁ = s₂ := + ⟨fun h => append_inj_left' h rfl, congrArg (· ++ _)⟩ + +@[simp] theorem append_left_eq_self {x y : Array α} : x ++ y = y ↔ x = #[] := by + rw [← append_left_inj (s₁ := x), nil_append] + +@[simp] theorem self_eq_append_left {x y : Array α} : y = x ++ y ↔ x = #[] := by + rw [eq_comm, append_left_eq_self] + +@[simp] theorem append_right_eq_self {x y : Array α} : x ++ y = x ↔ y = #[] := by + rw [← append_right_inj (t₁ := y), append_nil] + +@[simp] theorem self_eq_append_right {x y : Array α} : x = x ++ y ↔ y = #[] := by + rw [eq_comm, append_right_eq_self] + +@[simp] theorem append_eq_empty_iff : p ++ q = #[] ↔ p = #[] ∧ q = #[] := by + cases p <;> simp + +@[simp] theorem empty_eq_append_iff : #[] = a ++ b ↔ a = #[] ∧ b = #[] := by + rw [eq_comm, append_eq_empty_iff] + +theorem append_ne_empty_of_left_ne_empty {s : Array α} (h : s ≠ #[]) (t : Array α) : + s ++ t ≠ #[] := by + simp_all + +theorem append_ne_empty_of_right_ne_empty (s : Array α) : t ≠ #[] → s ++ t ≠ #[] := by + simp_all + +theorem append_eq_push_iff {a b c : Array α} {x : α} : + a ++ b = c.push x ↔ (b = #[] ∧ a = c.push x) ∨ (∃ b', b = b'.push x ∧ c = a ++ b') := by + rcases a with ⟨a⟩ + rcases b with ⟨b⟩ + rcases c with ⟨c⟩ + simp only [List.append_toArray, List.push_toArray, mk.injEq, List.append_eq_append_iff, + toArray_eq_append_iff] + constructor + · rintro (⟨a', rfl, rfl⟩ | ⟨b', rfl, h⟩) + · right; exact ⟨⟨a'⟩, by simp⟩ + · rw [List.singleton_eq_append_iff] at h + obtain (⟨rfl, rfl⟩ | ⟨rfl, rfl⟩) := h + · right; exact ⟨#[], by simp⟩ + · left; simp + · rintro (⟨rfl, rfl⟩ | ⟨b', h, rfl⟩) + · right; exact ⟨[x], by simp⟩ + · left; refine ⟨b'.toList, ?_⟩ + replace h := congrArg Array.toList h + simp_all + +theorem push_eq_append_iff {a b c : Array α} {x : α} : + c.push x = a ++ b ↔ (b = #[] ∧ a = c.push x) ∨ (∃ b', b = b'.push x ∧ c = a ++ b') := by + rw [eq_comm, append_eq_push_iff] + +theorem append_eq_singleton_iff {a b : Array α} {x : α} : + a ++ b = #[x] ↔ (a = #[] ∧ b = #[x]) ∨ (a = #[x] ∧ b = #[]) := by + rcases a with ⟨a⟩ + rcases b with ⟨b⟩ + simp only [List.append_toArray, mk.injEq, List.append_eq_singleton_iff, toArray_eq_append_iff] + +theorem singleton_eq_append_iff {a b : Array α} {x : α} : + #[x] = a ++ b ↔ (a = #[] ∧ b = #[x]) ∨ (a = #[x] ∧ b = #[]) := by + rw [eq_comm, append_eq_singleton_iff] + +theorem append_eq_append_iff {a b c d : Array α} : + a ++ b = c ++ d ↔ (∃ a', c = a ++ a' ∧ b = a' ++ d) ∨ ∃ c', a = c ++ c' ∧ d = c' ++ b := by + rcases a with ⟨a⟩ + rcases b with ⟨b⟩ + rcases c with ⟨c⟩ + rcases d with ⟨d⟩ + simp only [List.append_toArray, mk.injEq, List.append_eq_append_iff, toArray_eq_append_iff] + constructor + · rintro (⟨a', rfl, rfl⟩ | ⟨c', rfl, rfl⟩) + · left; exact ⟨⟨a'⟩, by simp⟩ + · right; exact ⟨⟨c'⟩, by simp⟩ + · rintro (⟨a', rfl, rfl⟩ | ⟨c', rfl, rfl⟩) + · left; exact ⟨a'.toList, by simp⟩ + · right; exact ⟨c'.toList, by simp⟩ + +theorem set_append {s t : Array α} {i : Nat} {x : α} (h : i < (s ++ t).size) : + (s ++ t).set i x = + if h' : i < s.size then + s.set i x ++ t + else + s ++ t.set (i - s.size) x (by simp at h; omega) := by + rcases s with ⟨s⟩ + rcases t with ⟨t⟩ + simp only [List.append_toArray, List.set_toArray, List.set_append] + split <;> simp + +@[simp] theorem set_append_left {s t : Array α} {i : Nat} {x : α} (h : i < s.size) : + (s ++ t).set i x (by simp; omega) = s.set i x ++ t := by + simp [set_append, h] + +@[simp] theorem set_append_right {s t : Array α} {i : Nat} {x : α} + (h' : i < (s ++ t).size) (h : s.size ≤ i) : + (s ++ t).set i x = s ++ t.set (i - s.size) x (by simp at h'; omega) := by + rw [set_append, dif_neg (by omega)] + +theorem setIfInBounds_append {s t : Array α} {i : Nat} {x : α} : + (s ++ t).setIfInBounds i x = + if i < s.size then + s.setIfInBounds i x ++ t + else + s ++ t.setIfInBounds (i - s.size) x := by + rcases s with ⟨s⟩ + rcases t with ⟨t⟩ + simp only [List.append_toArray, List.setIfInBounds_toArray, List.set_append] + split <;> simp + +@[simp] theorem setIfInBounds_append_left {s t : Array α} {i : Nat} {x : α} (h : i < s.size) : + (s ++ t).setIfInBounds i x = s.setIfInBounds i x ++ t := by + simp [setIfInBounds_append, h] + +@[simp] theorem setIfInBounds_append_right {s t : Array α} {i : Nat} {x : α} (h : s.size ≤ i) : + (s ++ t).setIfInBounds i x = s ++ t.setIfInBounds (i - s.size) x := by + rw [setIfInBounds_append, if_neg (by omega)] + +theorem filterMap_eq_append_iff {f : α → Option β} : + filterMap f l = L₁ ++ L₂ ↔ ∃ l₁ l₂, l = l₁ ++ l₂ ∧ filterMap f l₁ = L₁ ∧ filterMap f l₂ = L₂ := by + rcases l with ⟨l⟩ + rcases L₁ with ⟨L₁⟩ + rcases L₂ with ⟨L₂⟩ + simp only [size_toArray, List.filterMap_toArray', List.append_toArray, mk.injEq, + List.filterMap_eq_append_iff, toArray_eq_append_iff] + constructor + · rintro ⟨l₁, l₂, rfl, rfl, rfl⟩ + exact ⟨⟨l₁⟩, ⟨l₂⟩, by simp⟩ + · rintro ⟨⟨l₁⟩, ⟨l₂⟩, rfl, h₁, h₂⟩ + exact ⟨l₁, l₂, by simp_all⟩ + +theorem append_eq_filterMap_iff {f : α → Option β} : + L₁ ++ L₂ = filterMap f l ↔ + ∃ l₁ l₂, l = l₁ ++ l₂ ∧ filterMap f l₁ = L₁ ∧ filterMap f l₂ = L₂ := by + rw [eq_comm, filterMap_eq_append_iff] + +@[simp] theorem map_append (f : α → β) (l₁ l₂ : Array α) : + map f (l₁ ++ l₂) = map f l₁ ++ map f l₂ := by + cases l₁ + cases l₂ + simp + +theorem map_eq_append_iff {f : α → β} : + map f l = L₁ ++ L₂ ↔ ∃ l₁ l₂, l = l₁ ++ l₂ ∧ map f l₁ = L₁ ∧ map f l₂ = L₂ := by + rw [← filterMap_eq_map, filterMap_eq_append_iff] + +theorem append_eq_map_iff {f : α → β} : + L₁ ++ L₂ = map f l ↔ ∃ l₁ l₂, l = l₁ ++ l₂ ∧ map f l₁ = L₁ ∧ map f l₂ = L₂ := by + rw [eq_comm, map_eq_append_iff] + +/-! Content below this point has not yet been aligned with `List`. -/ -- This is a duplicate of `List.toArray_toList`. -- It's confusing to guess which namespace this theorem should live in, diff --git a/src/Init/Data/List/Lemmas.lean b/src/Init/Data/List/Lemmas.lean index 934456403a80..095917d5bc9f 100644 --- a/src/Init/Data/List/Lemmas.lean +++ b/src/Init/Data/List/Lemmas.lean @@ -1494,6 +1494,34 @@ theorem filterMap_eq_cons_iff {l} {b} {bs} : @[simp] theorem cons_append_fun (a : α) (as : List α) : (fun bs => ((a :: as) ++ bs)) = fun bs => a :: (as ++ bs) := rfl +@[simp] theorem mem_append {a : α} {s t : List α} : a ∈ s ++ t ↔ a ∈ s ∨ a ∈ t := by + induction s <;> simp_all [or_assoc] + +theorem not_mem_append {a : α} {s t : List α} (h₁ : a ∉ s) (h₂ : a ∉ t) : a ∉ s ++ t := + mt mem_append.1 $ not_or.mpr ⟨h₁, h₂⟩ + +@[deprecated mem_append (since := "2025-01-13")] +theorem mem_append_eq (a : α) (s t : List α) : (a ∈ s ++ t) = (a ∈ s ∨ a ∈ t) := + propext mem_append + +@[deprecated mem_append_left (since := "2024-11-20")] abbrev mem_append_of_mem_left := @mem_append_left +@[deprecated mem_append_right (since := "2024-11-20")] abbrev mem_append_of_mem_right := @mem_append_right + +/-- +See also `eq_append_cons_of_mem`, which proves a stronger version +in which the initial list must not contain the element. +-/ +theorem append_of_mem {a : α} {l : List α} : a ∈ l → ∃ s t : List α, l = s ++ a :: t + | .head l => ⟨[], l, rfl⟩ + | .tail b h => let ⟨s, t, h'⟩ := append_of_mem h; ⟨b::s, t, by rw [h', cons_append]⟩ + +theorem mem_iff_append {a : α} {l : List α} : a ∈ l ↔ ∃ s t : List α, l = s ++ a :: t := + ⟨append_of_mem, fun ⟨s, t, e⟩ => e ▸ by simp⟩ + +theorem forall_mem_append {p : α → Prop} {l₁ l₂ : List α} : + (∀ (x) (_ : x ∈ l₁ ++ l₂), p x) ↔ (∀ (x) (_ : x ∈ l₁), p x) ∧ (∀ (x) (_ : x ∈ l₂), p x) := by + simp only [mem_append, or_imp, forall_and] + theorem getElem_append {l₁ l₂ : List α} (i : Nat) (h : i < (l₁ ++ l₂).length) : (l₁ ++ l₂)[i] = if h' : i < l₁.length then l₁[i] else l₂[i - l₁.length]'(by simp at h h'; exact Nat.sub_lt_left_of_lt_add h' h) := by split <;> rename_i h' @@ -1561,14 +1589,6 @@ theorem get_of_append {l : List α} (eq : l = l₁ ++ a :: l₂) (h : l₁.lengt l.get ⟨i, get_of_append_proof eq h⟩ = a := Option.some.inj <| by rw [← get?_eq_get, eq, get?_append_right (h ▸ Nat.le_refl _), h, Nat.sub_self]; rfl -/-- -See also `eq_append_cons_of_mem`, which proves a stronger version -in which the initial list must not contain the element. --/ -theorem append_of_mem {a : α} {l : List α} : a ∈ l → ∃ s t : List α, l = s ++ a :: t - | .head l => ⟨[], l, rfl⟩ - | .tail b h => let ⟨s, t, h'⟩ := append_of_mem h; ⟨b::s, t, by rw [h', cons_append]⟩ - @[simp 1100] theorem singleton_append : [x] ++ l = x :: l := rfl theorem append_inj : @@ -1585,8 +1605,8 @@ theorem append_inj_left (h : s₁ ++ t₁ = s₂ ++ t₂) (hl : length s₁ = le /-- Variant of `append_inj` instead requiring equality of the lengths of the second lists. -/ theorem append_inj' (h : s₁ ++ t₁ = s₂ ++ t₂) (hl : length t₁ = length t₂) : s₁ = s₂ ∧ t₁ = t₂ := - append_inj h <| @Nat.add_right_cancel _ (length t₁) _ <| by - let hap := congrArg length h; simp only [length_append, ← hl] at hap; exact hap + append_inj h <| @Nat.add_right_cancel _ t₁.length _ <| by + let hap := congrArg length h; simp only [length_append, ← hl] at hap; exact hap /-- Variant of `append_inj_right` instead requiring equality of the lengths of the second lists. -/ theorem append_inj_right' (h : s₁ ++ t₁ = s₂ ++ t₂) (hl : length t₁ = length t₂) : t₁ = t₂ := @@ -1614,9 +1634,6 @@ theorem append_left_inj {s₁ s₂ : List α} (t) : s₁ ++ t = s₂ ++ t ↔ s @[simp] theorem self_eq_append_right {x y : List α} : x = x ++ y ↔ y = [] := by rw [eq_comm, append_right_eq_self] -@[simp] theorem append_eq_nil : p ++ q = [] ↔ p = [] ∧ q = [] := by - cases p <;> simp - theorem getLast_concat {a : α} : ∀ (l : List α), getLast (l ++ [a]) (by simp) = a | [] => rfl | a::t => by @@ -1642,6 +1659,54 @@ theorem get?_append {l₁ l₂ : List α} {n : Nat} (hn : n < l₁.length) : (l₁ ++ l₂).get? n = l₁.get? n := by simp [getElem?_append_left hn] +@[simp] theorem append_eq_nil_iff : p ++ q = [] ↔ p = [] ∧ q = [] := by + cases p <;> simp + +@[deprecated append_eq_nil_iff (since := "2025-01-13")] abbrev append_eq_nil := @append_eq_nil_iff + +@[simp] theorem nil_eq_append_iff : [] = a ++ b ↔ a = [] ∧ b = [] := by + rw [eq_comm, append_eq_nil_iff] + +@[deprecated nil_eq_append_iff (since := "2024-07-24")] abbrev nil_eq_append := @nil_eq_append_iff + +theorem append_ne_nil_of_left_ne_nil {s : List α} (h : s ≠ []) (t : List α) : s ++ t ≠ [] := by simp_all +theorem append_ne_nil_of_right_ne_nil (s : List α) : t ≠ [] → s ++ t ≠ [] := by simp_all + +@[deprecated append_ne_nil_of_left_ne_nil (since := "2024-07-24")] +theorem append_ne_nil_of_ne_nil_left {s : List α} (h : s ≠ []) (t : List α) : s ++ t ≠ [] := by simp_all +@[deprecated append_ne_nil_of_right_ne_nil (since := "2024-07-24")] +theorem append_ne_nil_of_ne_nil_right (s : List α) : t ≠ [] → s ++ t ≠ [] := by simp_all + +theorem append_eq_cons_iff : + a ++ b = x :: c ↔ (a = [] ∧ b = x :: c) ∨ (∃ a', a = x :: a' ∧ c = a' ++ b) := by + cases a with simp | cons a as => ?_ + exact ⟨fun h => ⟨as, by simp [h]⟩, fun ⟨a', ⟨aeq, aseq⟩, h⟩ => ⟨aeq, by rw [aseq, h]⟩⟩ + +@[deprecated append_eq_cons_iff (since := "2024-07-24")] abbrev append_eq_cons := @append_eq_cons_iff + +theorem cons_eq_append_iff : + x :: c = a ++ b ↔ (a = [] ∧ b = x :: c) ∨ (∃ a', a = x :: a' ∧ c = a' ++ b) := by + rw [eq_comm, append_eq_cons_iff] + +@[deprecated cons_eq_append_iff (since := "2024-07-24")] abbrev cons_eq_append := @cons_eq_append_iff + +theorem append_eq_singleton_iff : + a ++ b = [x] ↔ (a = [] ∧ b = [x]) ∨ (a = [x] ∧ b = []) := by + cases a <;> cases b <;> simp + +theorem singleton_eq_append_iff : + [x] = a ++ b ↔ (a = [] ∧ b = [x]) ∨ (a = [x] ∧ b = []) := by + cases a <;> cases b <;> simp [eq_comm] + +theorem append_eq_append_iff {a b c d : List α} : + a ++ b = c ++ d ↔ (∃ a', c = a ++ a' ∧ b = a' ++ d) ∨ ∃ c', a = c ++ c' ∧ d = c' ++ b := by + induction a generalizing c with + | nil => simp_all + | cons a as ih => cases c <;> simp [eq_comm, and_assoc, ih, and_or_left] + +@[deprecated append_inj (since := "2024-07-24")] abbrev append_inj_of_length_left := @append_inj +@[deprecated append_inj' (since := "2024-07-24")] abbrev append_inj_of_length_right := @append_inj' + @[simp] theorem head_append_of_ne_nil {l : List α} {w₁} (w₂) : head (l ++ l') w₁ = head l w₂ := by match l, w₂ with @@ -1691,60 +1756,6 @@ theorem tail_append {l l' : List α} : (l ++ l').tail = if l.isEmpty then l'.tai @[deprecated tail_append_of_ne_nil (since := "2024-07-24")] abbrev tail_append_left := @tail_append_of_ne_nil -theorem nil_eq_append_iff : [] = a ++ b ↔ a = [] ∧ b = [] := by - rw [eq_comm, append_eq_nil] - -@[deprecated nil_eq_append_iff (since := "2024-07-24")] abbrev nil_eq_append := @nil_eq_append_iff - -theorem append_ne_nil_of_left_ne_nil {s : List α} (h : s ≠ []) (t : List α) : s ++ t ≠ [] := by simp_all -theorem append_ne_nil_of_right_ne_nil (s : List α) : t ≠ [] → s ++ t ≠ [] := by simp_all - -@[deprecated append_ne_nil_of_left_ne_nil (since := "2024-07-24")] -theorem append_ne_nil_of_ne_nil_left {s : List α} (h : s ≠ []) (t : List α) : s ++ t ≠ [] := by simp_all -@[deprecated append_ne_nil_of_right_ne_nil (since := "2024-07-24")] -theorem append_ne_nil_of_ne_nil_right (s : List α) : t ≠ [] → s ++ t ≠ [] := by simp_all - -theorem append_eq_cons_iff : - a ++ b = x :: c ↔ (a = [] ∧ b = x :: c) ∨ (∃ a', a = x :: a' ∧ c = a' ++ b) := by - cases a with simp | cons a as => ?_ - exact ⟨fun h => ⟨as, by simp [h]⟩, fun ⟨a', ⟨aeq, aseq⟩, h⟩ => ⟨aeq, by rw [aseq, h]⟩⟩ - -@[deprecated append_eq_cons_iff (since := "2024-07-24")] abbrev append_eq_cons := @append_eq_cons_iff - -theorem cons_eq_append_iff : - x :: c = a ++ b ↔ (a = [] ∧ b = x :: c) ∨ (∃ a', a = x :: a' ∧ c = a' ++ b) := by - rw [eq_comm, append_eq_cons_iff] - -@[deprecated cons_eq_append_iff (since := "2024-07-24")] abbrev cons_eq_append := @cons_eq_append_iff - -theorem append_eq_append_iff {a b c d : List α} : - a ++ b = c ++ d ↔ (∃ a', c = a ++ a' ∧ b = a' ++ d) ∨ ∃ c', a = c ++ c' ∧ d = c' ++ b := by - induction a generalizing c with - | nil => simp_all - | cons a as ih => cases c <;> simp [eq_comm, and_assoc, ih, and_or_left] - -@[deprecated append_inj (since := "2024-07-24")] abbrev append_inj_of_length_left := @append_inj -@[deprecated append_inj' (since := "2024-07-24")] abbrev append_inj_of_length_right := @append_inj' - -@[simp] theorem mem_append {a : α} {s t : List α} : a ∈ s ++ t ↔ a ∈ s ∨ a ∈ t := by - induction s <;> simp_all [or_assoc] - -theorem not_mem_append {a : α} {s t : List α} (h₁ : a ∉ s) (h₂ : a ∉ t) : a ∉ s ++ t := - mt mem_append.1 $ not_or.mpr ⟨h₁, h₂⟩ - -theorem mem_append_eq (a : α) (s t : List α) : (a ∈ s ++ t) = (a ∈ s ∨ a ∈ t) := - propext mem_append - -@[deprecated mem_append_left (since := "2024-11-20")] abbrev mem_append_of_mem_left := @mem_append_left -@[deprecated mem_append_right (since := "2024-11-20")] abbrev mem_append_of_mem_right := @mem_append_right - -theorem mem_iff_append {a : α} {l : List α} : a ∈ l ↔ ∃ s t : List α, l = s ++ a :: t := - ⟨append_of_mem, fun ⟨s, t, e⟩ => e ▸ by simp⟩ - -theorem forall_mem_append {p : α → Prop} {l₁ l₂ : List α} : - (∀ (x) (_ : x ∈ l₁ ++ l₂), p x) ↔ (∀ (x) (_ : x ∈ l₁), p x) ∧ (∀ (x) (_ : x ∈ l₂), p x) := by - simp only [mem_append, or_imp, forall_and] - theorem set_append {s t : List α} : (s ++ t).set i x = if i < s.length then s.set i x ++ t else s ++ t.set (i - s.length) x := by induction s generalizing i with @@ -1974,8 +1985,8 @@ theorem flatten_eq_append_iff {xs : List (List α)} {ys zs : List α} : constructor · induction xs generalizing ys with | nil => - simp only [flatten_nil, nil_eq, append_eq_nil, and_false, cons_append, false_and, exists_const, - exists_false, or_false, and_imp, List.cons_ne_nil] + simp only [flatten_nil, nil_eq, append_eq_nil_iff, and_false, cons_append, false_and, + exists_const, exists_false, or_false, and_imp, List.cons_ne_nil] rintro rfl rfl exact ⟨[], [], by simp⟩ | cons x xs ih => diff --git a/src/Init/Data/List/Zip.lean b/src/Init/Data/List/Zip.lean index 8b0b80808be1..4ffb4b08d300 100644 --- a/src/Init/Data/List/Zip.lean +++ b/src/Init/Data/List/Zip.lean @@ -203,11 +203,11 @@ theorem zipWith_eq_append_iff {f : α → β → γ} {l₁ : List α} {l₂ : Li cases l₂ with | nil => constructor - · simp only [zipWith_nil_right, nil_eq, append_eq_nil, exists_and_left, and_imp] + · simp only [zipWith_nil_right, nil_eq, append_eq_nil_iff, exists_and_left, and_imp] rintro rfl rfl exact ⟨[], x₁ :: l₁, [], by simp⟩ · rintro ⟨w, x, y, z, h₁, _, h₃, rfl, rfl⟩ - simp only [nil_eq, append_eq_nil] at h₃ + simp only [nil_eq, append_eq_nil_iff] at h₃ obtain ⟨rfl, rfl⟩ := h₃ simp | cons x₂ l₂ => diff --git a/src/Init/Data/Vector/Lemmas.lean b/src/Init/Data/Vector/Lemmas.lean index f8fb2eaf2e0c..bce753edba49 100644 --- a/src/Init/Data/Vector/Lemmas.lean +++ b/src/Init/Data/Vector/Lemmas.lean @@ -693,6 +693,24 @@ theorem forall_getElem {l : Vector α n} {p : α → Prop} : rcases l with ⟨l, rfl⟩ simp [Array.forall_getElem] + +/-! ### cast -/ + +@[simp] theorem getElem_cast (a : Vector α n) (h : n = m) (i : Nat) (hi : i < m) : + (a.cast h)[i] = a[i] := by + cases a + simp + +@[simp] theorem getElem?_cast {l : Vector α n} {m : Nat} {w : n = m} {i : Nat} : + (l.cast w)[i]? = l[i]? := by + rcases l with ⟨l, rfl⟩ + simp + +@[simp] theorem mem_cast {a : α} {l : Vector α n} {m : Nat} {w : n = m} : + a ∈ l.cast w ↔ a ∈ l := by + rcases l with ⟨l, rfl⟩ + simp + /-! ### Decidability of bounded quantifiers -/ instance {xs : Vector α n} {p : α → Prop} [DecidablePred p] : @@ -1167,6 +1185,227 @@ theorem map_eq_iff {f : α → β} {l : Vector α n} {l' : Vector β n} : cases as simp +/-! ### singleton -/ + +@[simp] theorem singleton_def (v : α) : Vector.singleton v = #v[v] := rfl + +/-! ### append -/ + +@[simp] theorem append_push {as : Vector α n} {bs : Vector α m} {a : α} : + as ++ bs.push a = (as ++ bs).push a := by + cases as + cases bs + simp + +theorem singleton_eq_toVector_singleton (a : α) : #v[a] = #[a].toVector := rfl + +@[simp] theorem mem_append {a : α} {s : Vector α n} {t : Vector α m} : + a ∈ s ++ t ↔ a ∈ s ∨ a ∈ t := by + cases s + cases t + simp + +theorem mem_append_left {a : α} {s : Vector α n} {t : Vector α m} (h : a ∈ s) : a ∈ s ++ t := + mem_append.2 (Or.inl h) + +theorem mem_append_right {a : α} {s : Vector α n} {t : Vector α m} (h : a ∈ t) : a ∈ s ++ t := + mem_append.2 (Or.inr h) + +theorem not_mem_append {a : α} {s : Vector α n} {t : Vector α m} (h₁ : a ∉ s) (h₂ : a ∉ t) : + a ∉ s ++ t := + mt mem_append.1 $ not_or.mpr ⟨h₁, h₂⟩ + +/-- +See also `eq_push_append_of_mem`, which proves a stronger version +in which the initial array must not contain the element. +-/ +theorem append_of_mem {a : α} {l : Vector α n} (h : a ∈ l) : + ∃ (m k : Nat) (w : m + 1 + k = n) (s : Vector α m) (t : Vector α k), + l = (s.push a ++ t).cast w := by + rcases l with ⟨l, rfl⟩ + obtain ⟨s, t, rfl⟩ := Array.append_of_mem (by simpa using h) + refine ⟨_, _, by simp, s.toVector, t.toVector, by simp_all⟩ + +theorem mem_iff_append {a : α} {l : Vector α n} : + a ∈ l ↔ ∃ (m k : Nat) (w : m + 1 + k = n) (s : Vector α m) (t : Vector α k), + l = (s.push a ++ t).cast w := + ⟨append_of_mem, by rintro ⟨m, k, rfl, s, t, rfl⟩; simp⟩ + +theorem forall_mem_append {p : α → Prop} {l₁ : Vector α n} {l₂ : Vector α m} : + (∀ (x) (_ : x ∈ l₁ ++ l₂), p x) ↔ (∀ (x) (_ : x ∈ l₁), p x) ∧ (∀ (x) (_ : x ∈ l₂), p x) := by + simp only [mem_append, or_imp, forall_and] + +theorem empty_append (as : Vector α n) : (#v[] : Vector α 0) ++ as = as.cast (by omega) := by + rcases as with ⟨as, rfl⟩ + simp + +theorem append_empty (as : Vector α n) : as ++ (#v[] : Vector α 0) = as := by + rw [← toArray_inj, toArray_append, Array.append_nil] + +theorem getElem_append (a : Vector α n) (b : Vector α m) (i : Nat) (hi : i < n + m) : + (a ++ b)[i] = if h : i < n then a[i] else b[i - n] := by + rcases a with ⟨a, rfl⟩ + rcases b with ⟨b, rfl⟩ + simp [Array.getElem_append, hi] + +theorem getElem_append_left {a : Vector α n} {b : Vector α m} {i : Nat} (hi : i < n) : + (a ++ b)[i] = a[i] := by simp [getElem_append, hi] + +theorem getElem_append_right {a : Vector α n} {b : Vector α m} {i : Nat} (h : i < n + m) (hi : n ≤ i) : + (a ++ b)[i] = b[i - n] := by + rw [getElem_append, dif_neg (by omega)] + +theorem getElem?_append_left {as : Vector α n} {bs : Vector α m} {i : Nat} (hn : i < n) : + (as ++ bs)[i]? = as[i]? := by + have hn' : i < n + m := by omega + simp_all [getElem?_eq_getElem, getElem_append] + +theorem getElem?_append_right {as : Vector α n} {bs : Vector α m} {i : Nat} (h : n ≤ i) : + (as ++ bs)[i]? = bs[i - n]? := by + rcases as with ⟨as, rfl⟩ + rcases bs with ⟨bs, rfl⟩ + simp [Array.getElem?_append_right, h] + +theorem getElem?_append {as : Vector α n} {bs : Vector α m} {i : Nat} : + (as ++ bs)[i]? = if i < n then as[i]? else bs[i - n]? := by + split <;> rename_i h + · exact getElem?_append_left h + · exact getElem?_append_right (by simpa using h) + +/-- Variant of `getElem_append_left` useful for rewriting from the small array to the big array. -/ +theorem getElem_append_left' (l₁ : Vector α m) {l₂ : Vector α n} {i : Nat} (hi : i < m) : + l₁[i] = (l₁ ++ l₂)[i] := by + rw [getElem_append_left] <;> simp + +/-- Variant of `getElem_append_right` useful for rewriting from the small array to the big array. -/ +theorem getElem_append_right' (l₁ : Vector α m) {l₂ : Vector α n} {i : Nat} (hi : i < n) : + l₂[i] = (l₁ ++ l₂)[i + m] := by + rw [getElem_append_right] <;> simp [*, Nat.le_add_left] + +theorem getElem_of_append {l : Vector α n} {l₁ : Vector α m} {l₂ : Vector α k} + (w : m + 1 + k = n) (eq : l = (l₁.push a ++ l₂).cast w) : + l[m] = a := Option.some.inj <| by + rw [← getElem?_eq_getElem, eq, getElem?_cast, getElem?_append_left (by simp)] + simp + +@[simp 1100] theorem append_singleton {a : α} {as : Vector α n} : as ++ #v[a] = as.push a := by + cases as + simp + +theorem append_inj {s₁ s₂ : Vector α n} {t₁ t₂ : Vector α m} (h : s₁ ++ t₁ = s₂ ++ t₂) : + s₁ = s₂ ∧ t₁ = t₂ := by + rcases s₁ with ⟨s₁, rfl⟩ + rcases s₂ with ⟨s₂, hs⟩ + rcases t₁ with ⟨t₁, rfl⟩ + rcases t₂ with ⟨t₂, ht⟩ + simpa using Array.append_inj (by simpa using h) (by omega) + +theorem append_inj_right {s₁ s₂ : Vector α n} {t₁ t₂ : Vector α m} + (h : s₁ ++ t₁ = s₂ ++ t₂) : t₁ = t₂ := + (append_inj h).right + +theorem append_inj_left {s₁ s₂ : Vector α n} {t₁ t₂ : Vector α m} + (h : s₁ ++ t₁ = s₂ ++ t₂) : s₁ = s₂ := + (append_inj h).left + +theorem append_right_inj {t₁ t₂ : Vector α m} (s : Vector α n) : s ++ t₁ = s ++ t₂ ↔ t₁ = t₂ := + ⟨fun h => append_inj_right h, congrArg _⟩ + +theorem append_left_inj {s₁ s₂ : Vector α n} (t : Vector α m) : s₁ ++ t = s₂ ++ t ↔ s₁ = s₂ := + ⟨fun h => append_inj_left h, congrArg (· ++ _)⟩ + +theorem append_eq_append_iff {a : Vector α n} {b : Vector α m} {c : Vector α k} {d : Vector α l} + (w : k + l = n + m) : + a ++ b = (c ++ d).cast w ↔ + if h : n ≤ k then + ∃ a' : Vector α (k - n), c = (a ++ a').cast (by omega) ∧ b = (a' ++ d).cast (by omega) + else + ∃ c' : Vector α (n - k), a = (c ++ c').cast (by omega) ∧ d = (c' ++ b).cast (by omega) := by + rcases a with ⟨a, rfl⟩ + rcases b with ⟨b, rfl⟩ + rcases c with ⟨c, rfl⟩ + rcases d with ⟨d, rfl⟩ + simp only [mk_append_mk, Array.append_eq_append_iff, mk_eq, toArray_cast] + constructor + · rintro (⟨a', rfl, rfl⟩ | ⟨c', rfl, rfl⟩) + · rw [dif_pos (by simp)] + exact ⟨a'.toVector.cast (by simp; omega), by simp⟩ + · split <;> rename_i h + · have hc : c'.size = 0 := by simp at h; omega + simp at hc + exact ⟨#v[].cast (by simp; omega), by simp_all⟩ + · exact ⟨c'.toVector.cast (by simp; omega), by simp⟩ + · split <;> rename_i h + · rintro ⟨a', hc, rfl⟩ + left + refine ⟨a'.toArray, hc, rfl⟩ + · rintro ⟨c', ha, rfl⟩ + right + refine ⟨c'.toArray, ha, rfl⟩ + +theorem set_append {s : Vector α n} {t : Vector α m} {i : Nat} {x : α} (h : i < n + m) : + (s ++ t).set i x = + if h' : i < n then + s.set i x ++ t + else + s ++ t.set (i - n) x := by + rcases s with ⟨s, rfl⟩ + rcases t with ⟨t, rfl⟩ + simp only [mk_append_mk, set_mk, Array.set_append] + split <;> simp + +@[simp] theorem set_append_left {s : Vector α n} {t : Vector α m} {i : Nat} {x : α} (h : i < n) : + (s ++ t).set i x = s.set i x ++ t := by + simp [set_append, h] + +@[simp] theorem set_append_right {s : Vector α n} {t : Vector α m} {i : Nat} {x : α} + (h' : i < n + m) (h : n ≤ i) : + (s ++ t).set i x = s ++ t.set (i - n) x := by + rw [set_append, dif_neg (by omega)] + +theorem setIfInBounds_append {s : Vector α n} {t : Vector α m} {i : Nat} {x : α} : + (s ++ t).setIfInBounds i x = + if i < n then + s.setIfInBounds i x ++ t + else + s ++ t.setIfInBounds (i - n) x := by + rcases s with ⟨s, rfl⟩ + rcases t with ⟨t, rfl⟩ + simp only [mk_append_mk, setIfInBounds_mk, Array.setIfInBounds_append] + split <;> simp + +@[simp] theorem setIfInBounds_append_left {s : Vector α n} {t : Vector α m} {i : Nat} {x : α} (h : i < n) : + (s ++ t).setIfInBounds i x = s.setIfInBounds i x ++ t := by + simp [setIfInBounds_append, h] + +@[simp] theorem setIfInBounds_append_right {s : Vector α n} {t : Vector α m} {i : Nat} {x : α} + (h : n ≤ i) : + (s ++ t).setIfInBounds i x = s ++ t.setIfInBounds (i - n) x := by + rw [setIfInBounds_append, if_neg (by omega)] + +@[simp] theorem map_append (f : α → β) (l₁ : Vector α n) (l₂ : Vector α m) : + map f (l₁ ++ l₂) = map f l₁ ++ map f l₂ := by + rcases l₁ with ⟨l₁, rfl⟩ + rcases l₂ with ⟨l₂, rfl⟩ + simp + +theorem map_eq_append_iff {f : α → β} : + map f l = L₁ ++ L₂ ↔ ∃ l₁ l₂, l = l₁ ++ l₂ ∧ map f l₁ = L₁ ∧ map f l₂ = L₂ := by + rcases l with ⟨l, h⟩ + rcases L₁ with ⟨L₁, rfl⟩ + rcases L₂ with ⟨L₂, rfl⟩ + simp only [map_mk, mk_append_mk, eq_mk, Array.map_eq_append_iff, mk_eq, toArray_append, + toArray_map] + constructor + · rintro ⟨l₁, l₂, rfl, rfl, rfl⟩ + exact ⟨l₁.toVector.cast (by simp), l₂.toVector.cast (by simp), by simp⟩ + · rintro ⟨⟨l₁⟩, ⟨l₂⟩, rfl, h₁, h₂⟩ + exact ⟨l₁, l₂, by simp_all⟩ + +theorem append_eq_map_iff {f : α → β} : + L₁ ++ L₂ = map f l ↔ ∃ l₁ l₂, l = l₁ ++ l₂ ∧ map f l₁ = L₁ ∧ map f l₂ = L₂ := by + rw [eq_comm, map_eq_append_iff] + /-! Content below this point has not yet been aligned with `List` and `Array`. -/ @[simp] theorem getElem_ofFn {α n} (f : Fin n → α) (i : Nat) (h : i < n) : @@ -1197,28 +1436,6 @@ defeq issues in the implicit size argument. subst h simp [pop, back, back!, ← Array.eq_push_pop_back!_of_size_ne_zero] -/-! ### append -/ - -theorem getElem_append (a : Vector α n) (b : Vector α m) (i : Nat) (hi : i < n + m) : - (a ++ b)[i] = if h : i < n then a[i] else b[i - n] := by - rcases a with ⟨a, rfl⟩ - rcases b with ⟨b, rfl⟩ - simp [Array.getElem_append, hi] - -theorem getElem_append_left {a : Vector α n} {b : Vector α m} {i : Nat} (hi : i < n) : - (a ++ b)[i] = a[i] := by simp [getElem_append, hi] - -theorem getElem_append_right {a : Vector α n} {b : Vector α m} {i : Nat} (h : i < n + m) (hi : n ≤ i) : - (a ++ b)[i] = b[i - n] := by - rw [getElem_append, dif_neg (by omega)] - -/-! ### cast -/ - -@[simp] theorem getElem_cast (a : Vector α n) (h : n = m) (i : Nat) (hi : i < m) : - (a.cast h)[i] = a[i] := by - cases a - simp - /-! ### extract -/ @[simp] theorem getElem_extract (a : Vector α n) (start stop) (i : Nat) (hi : i < min stop n - start) : From 40efbb9b7a5e5e05e7e4009a8287d7228d18552a Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Mon, 13 Jan 2025 13:29:46 +1100 Subject: [PATCH 088/100] doc: commit conventions and Mathlib CI (#6605) This PR updates the commit conventions documentation to describe the new changelog conventions, and adds brief documentation of integrated Mathlib CI, with a link for further explanation. --- doc/dev/commit_convention.md | 24 ++++++++++++++++-------- doc/dev/index.md | 7 +++++++ 2 files changed, 23 insertions(+), 8 deletions(-) diff --git a/doc/dev/commit_convention.md b/doc/dev/commit_convention.md index 80362c175bec..9793350521a7 100644 --- a/doc/dev/commit_convention.md +++ b/doc/dev/commit_convention.md @@ -33,6 +33,9 @@ Format of the commit message - chore (maintain, ex: travis-ci) - perf (performance improvement, optimization, ...) +Every `feat` or `fix` commit must have a `changelog-*` label, and a commit message +beginning with "This PR " that will be included in the changelog. + ```` has the following constraints: - use imperative, present tense: "change" not "changed" nor "changes" @@ -44,6 +47,7 @@ Format of the commit message - just as in ````, use imperative, present tense - includes motivation for the change and contrasts with previous behavior + - If a `changelog-*` label is present, the body must begin with "This PR ". ``